diff --git a/thys/Akra_Bazzi/Akra_Bazzi_Library.thy b/thys/Akra_Bazzi/Akra_Bazzi_Library.thy --- a/thys/Akra_Bazzi/Akra_Bazzi_Library.thy +++ b/thys/Akra_Bazzi/Akra_Bazzi_Library.thy @@ -1,285 +1,286 @@ (* File: Akra_Bazzi_Library.thy Author: Manuel Eberl Lemma bucket for the Akra-Bazzi theorem. *) section \Auxiliary lemmas\ theory Akra_Bazzi_Library imports Complex_Main "Landau_Symbols.Landau_More" + "Pure-ex.Guess" begin (* TODO: Move? *) lemma ln_mono: "0 < x \ 0 < y \ x \ y \ ln (x::real) \ ln y" by (subst ln_le_cancel_iff) simp_all lemma ln_mono_strict: "0 < x \ 0 < y \ x < y \ ln (x::real) < ln y" by (subst ln_less_cancel_iff) simp_all declare DERIV_powr[THEN DERIV_chain2, derivative_intros] lemma sum_pos': assumes "finite I" assumes "\x\I. f x > (0 :: _ :: linordered_ab_group_add)" assumes "\x. x \ I \ f x \ 0" shows "sum f I > 0" proof- from assms(2) guess x by (elim bexE) note x = this from x have "I = insert x I" by blast also from assms(1) have "sum f ... = f x + sum f (I - {x})" by (rule sum.insert_remove) also from x assms have "... > 0" by (intro add_pos_nonneg sum_nonneg) simp_all finally show ?thesis . qed lemma min_mult_left: assumes "(x::real) > 0" shows "x * min y z = min (x*y) (x*z)" using assms by (auto simp add: min_def algebra_simps) lemma max_mult_left: assumes "(x::real) > 0" shows "x * max y z = max (x*y) (x*z)" using assms by (auto simp add: max_def algebra_simps) lemma DERIV_nonneg_imp_mono: assumes "\t. t\{x..y} \ (f has_field_derivative f' t) (at t)" assumes "\t. t\{x..y} \ f' t \ 0" assumes "(x::real) \ y" shows "(f x :: real) \ f y" proof (cases x y rule: linorder_cases) assume xy: "x < y" hence "\z. x < z \ z < y \ f y - f x = (y - x) * f' z" by (rule MVT2) (insert assms(1), simp) then guess z by (elim exE conjE) note z = this from z(1,2) assms(2) xy have "0 \ (y - x) * f' z" by (intro mult_nonneg_nonneg) simp_all also note z(3)[symmetric] finally show "f x \ f y" by simp qed (insert assms(3), simp_all) lemma eventually_conjE: "eventually (\x. P x \ Q x) F \ (eventually P F \ eventually Q F \ R) \ R" apply (frule eventually_rev_mp[of _ _ P], simp) apply (drule eventually_rev_mp[of _ _ Q], simp) apply assumption done lemma real_natfloor_nat: "x \ \ \ real (nat \x\) = x" by (elim Nats_cases) simp lemma eventually_natfloor: assumes "eventually P (at_top :: nat filter)" shows "eventually (\x. P (nat \x\)) (at_top :: real filter)" proof- from assms obtain N where N: "\n. n \ N \ P n" using eventually_at_top_linorder by blast have "\n\real N. P (nat \n\)" by (intro allI impI N le_nat_floor) simp_all thus ?thesis using eventually_at_top_linorder by blast qed lemma tendsto_0_smallo_1: "f \ o(\x. 1 :: real) \ (f \ 0) at_top" by (drule smalloD_tendsto) simp lemma smallo_1_tendsto_0: "(f \ 0) at_top \ f \ o(\x. 1 :: real)" by (rule smalloI_tendsto) simp_all lemma filterlim_at_top_smallomega_1: assumes "f \ \[F](\x. 1 :: real)" "eventually (\x. f x > 0) F" shows "filterlim f at_top F" proof - from assms have "filterlim (\x. norm (f x / 1)) at_top F" by (intro smallomegaD_filterlim_at_top_norm) (auto elim: eventually_mono) also have "?this \ ?thesis" using assms by (intro filterlim_cong refl) (auto elim!: eventually_mono) finally show ?thesis . qed lemma smallo_imp_abs_less_real: assumes "f \ o[F](g)" "eventually (\x. g x > (0::real)) F" shows "eventually (\x. \f x\ < g x) F" proof - have "1/2 > (0::real)" by simp from landau_o.smallD[OF assms(1) this] assms(2) show ?thesis by eventually_elim auto qed lemma smallo_imp_less_real: assumes "f \ o[F](g)" "eventually (\x. g x > (0::real)) F" shows "eventually (\x. f x < g x) F" using smallo_imp_abs_less_real[OF assms] by eventually_elim simp lemma smallo_imp_le_real: assumes "f \ o[F](g)" "eventually (\x. g x \ (0::real)) F" shows "eventually (\x. f x \ g x) F" using landau_o.smallD[OF assms(1) zero_less_one] assms(2) by eventually_elim simp (* TODO MOVE *) lemma filterlim_at_right: "filterlim f (at_right a) F \ eventually (\x. f x > a) F \ filterlim f (nhds a) F" by (subst filterlim_at) (auto elim!: eventually_mono) (* END TODO *) lemma one_plus_x_powr_approx_ex: assumes x: "abs (x::real) \ 1/2" obtains t where "abs t < 1/2" "(1 + x) powr p = 1 + p * x + p * (p - 1) * (1 + t) powr (p - 2) / 2 * x ^ 2" proof (cases "x = 0") assume x': "x \ 0" let ?f = "\x. (1 + x) powr p" let ?f' = "\x. p * (1 + x) powr (p - 1)" let ?f'' = "\x. p * (p - 1) * (1 + x) powr (p - 2)" let ?fs = "(!) [?f, ?f', ?f'']" have A: "\m t. m < 2 \ t \ -0.5 \ t \ 0.5 \ (?fs m has_real_derivative ?fs (Suc m) t) (at t)" proof (clarify) fix m :: nat and t :: real assume m: "m < 2" and t: "t \ -0.5" "t \ 0.5" thus "(?fs m has_real_derivative ?fs (Suc m) t) (at t)" using m by (cases m) (force intro: derivative_eq_intros algebra_simps)+ qed have "\t. (if x < 0 then x < t \ t < 0 else 0 < t \ t < x) \ (1 + x) powr p = (\m<2. ?fs m 0 / (fact m) * (x - 0)^m) + ?fs 2 t / (fact 2) * (x - 0)\<^sup>2" using assms x' by (intro Taylor[OF _ _ A]) simp_all then guess t by (elim exE conjE) note t = this with assms have "abs t < 1/2" by (auto split: if_split_asm) moreover from t(2) have "(1 + x) powr p = 1 + p * x + p * (p - 1) * (1 + t) powr (p - 2) / 2 * x ^ 2" by (simp add: numeral_2_eq_2 of_nat_Suc) ultimately show ?thesis by (rule that) next assume "x = 0" with that[of 0] show ?thesis by simp qed lemma powr_lower_bound: "\(l::real) > 0; l \ x; x \ u\ \ min (l powr z) (u powr z) \ x powr z" apply (cases "z \ 0") apply (rule order.trans[OF min.cobounded1 powr_mono2], simp_all) [] apply (rule order.trans[OF min.cobounded2 powr_mono2'], simp_all) [] done lemma powr_upper_bound: "\(l::real) > 0; l \ x; x \ u\ \ max (l powr z) (u powr z) \ x powr z" apply (cases "z \ 0") apply (rule order.trans[OF powr_mono2 max.cobounded2], simp_all) [] apply (rule order.trans[OF powr_mono2' max.cobounded1], simp_all) [] done lemma one_plus_x_powr_Taylor2: obtains k where "\x. abs (x::real) \ 1/2 \ abs ((1 + x) powr p - 1 - p*x) \ k*x^2" proof- define k where "k = \p*(p - 1)\ * max ((1/2) powr (p - 2)) ((3/2) powr (p - 2)) / 2" show ?thesis proof (rule that[of k]) fix x :: real assume "abs x \ 1/2" from one_plus_x_powr_approx_ex[OF this, of p] guess t . note t = this from t have "abs ((1 + x) powr p - 1 - p*x) = \p*(p - 1)\ * (1 + t) powr (p - 2)/2 * x\<^sup>2" by (simp add: abs_mult) also from t(1) have "(1 + t) powr (p - 2) \ max ((1/2) powr (p - 2)) ((3/2) powr (p - 2))" by (intro powr_upper_bound) simp_all finally show "abs ((1 + x) powr p - 1 - p*x) \ k*x^2" by (simp add: mult_left_mono mult_right_mono k_def) qed qed lemma one_plus_x_powr_Taylor2_bigo: assumes lim: "(f \ 0) F" shows "(\x. (1 + f x) powr (p::real) - 1 - p * f x) \ O[F](\x. f x ^ 2)" proof - from one_plus_x_powr_Taylor2[of p] guess k . moreover from tendstoD[OF lim, of "1/2"] have "eventually (\x. abs (f x) < 1/2) F" by (simp add: dist_real_def) ultimately have "eventually (\x. norm ((1 + f x) powr p - 1 - p * f x) \ k * norm (f x ^ 2)) F" by (auto elim!: eventually_mono) thus ?thesis by (rule bigoI) qed lemma one_plus_x_powr_Taylor1_bigo: assumes lim: "(f \ 0) F" shows "(\x. (1 + f x) powr (p::real) - 1) \ O[F](\x. f x)" proof - from assms have "(\x. (1 + f x) powr p - 1 - p * f x) \ O[F](\x. (f x)\<^sup>2)" by (rule one_plus_x_powr_Taylor2_bigo) also from assms have "f \ O[F](\_. 1)" by (intro bigoI_tendsto) simp_all from landau_o.big.mult[of f F f, OF _ this] have "(\x. (f x)^2) \ O[F](\x. f x)" by (simp add: power2_eq_square) finally have A: "(\x. (1 + f x) powr p - 1 - p * f x) \ O[F](f)" . have B: "(\x. p * f x) \ O[F](f)" by simp from sum_in_bigo(1)[OF A B] show ?thesis by simp qed lemma x_times_x_minus_1_nonneg: "x \ 0 \ x \ 1 \ (x::_::linordered_idom) * (x - 1) \ 0" proof (elim disjE) assume x: "x \ 0" also have "0 \ x^2" by simp finally show "x * (x - 1) \ 0" by (simp add: power2_eq_square algebra_simps) qed simp lemma x_times_x_minus_1_nonpos: "x \ 0 \ x \ 1 \ (x::_::linordered_idom) * (x - 1) \ 0" by (intro mult_nonneg_nonpos) simp_all lemma powr_mono': assumes "(x::real) > 0" "x \ 1" "a \ b" shows "x powr b \ x powr a" proof- have "inverse x powr a \ inverse x powr b" using assms by (intro powr_mono) (simp_all add: field_simps) hence "inverse (x powr a) \ inverse (x powr b)" using assms by simp with assms show ?thesis by (simp add: field_simps) qed lemma powr_less_mono': assumes "(x::real) > 0" "x < 1" "a < b" shows "x powr b < x powr a" proof- have "inverse x powr a < inverse x powr b" using assms by (intro powr_less_mono) (simp_all add: field_simps) hence "inverse (x powr a) < inverse (x powr b)" using assms by simp with assms show ?thesis by (simp add: field_simps) qed lemma real_powr_at_bot: assumes "(a::real) > 1" shows "((\x. a powr x) \ 0) at_bot" proof- from assms have "filterlim (\x. ln a * x) at_bot at_bot" by (intro filterlim_tendsto_pos_mult_at_bot[OF tendsto_const _ filterlim_ident]) auto hence "((\x. exp (x * ln a)) \ 0) at_bot" by (intro filterlim_compose[OF exp_at_bot]) (simp add: algebra_simps) thus ?thesis using assms unfolding powr_def by simp qed lemma real_powr_at_bot_neg: assumes "(a::real) > 0" "a < 1" shows "filterlim (\x. a powr x) at_top at_bot" proof- from assms have "LIM x at_bot. ln (inverse a) * -x :> at_top" by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] filterlim_uminus_at_top_at_bot) (simp_all add: ln_inverse) with assms have "LIM x at_bot. x * ln a :> at_top" by (subst (asm) ln_inverse) (simp_all add: mult.commute) hence "LIM x at_bot. exp (x * ln a) :> at_top" by (intro filterlim_compose[OF exp_at_top]) simp thus ?thesis using assms unfolding powr_def by simp qed lemma real_powr_at_top_neg: assumes "(a::real) > 0" "a < 1" shows "((\x. a powr x) \ 0) at_top" proof- from assms have "LIM x at_top. ln (inverse a) * x :> at_top" by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const]) (simp_all add: filterlim_ident field_simps) with assms have "LIM x at_top. ln a * x :> at_bot" by (subst filterlim_uminus_at_bot) (simp add: ln_inverse) hence "((\x. exp (x * ln a)) \ 0) at_top" by (intro filterlim_compose[OF exp_at_bot]) (simp_all add: mult.commute) with assms show ?thesis unfolding powr_def by simp qed lemma eventually_nat_real: assumes "eventually P (at_top :: real filter)" shows "eventually (\x. P (real x)) (at_top :: nat filter)" using assms filterlim_real_sequentially unfolding filterlim_def le_filter_def eventually_filtermap by auto end diff --git a/thys/Akra_Bazzi/ROOT b/thys/Akra_Bazzi/ROOT --- a/thys/Akra_Bazzi/ROOT +++ b/thys/Akra_Bazzi/ROOT @@ -1,16 +1,17 @@ chapter AFP session Akra_Bazzi (AFP) = "HOL-Analysis" + options [timeout = 600] sessions + "Pure-ex" "HOL-Decision_Procs" Landau_Symbols theories Akra_Bazzi Master_Theorem Akra_Bazzi_Method Akra_Bazzi_Approximation Master_Theorem_Examples document_files "root.tex" "root.bib" diff --git a/thys/Automatic_Refinement/Lib/Misc.thy b/thys/Automatic_Refinement/Lib/Misc.thy --- a/thys/Automatic_Refinement/Lib/Misc.thy +++ b/thys/Automatic_Refinement/Lib/Misc.thy @@ -1,4507 +1,4507 @@ (* Title: Miscellaneous Definitions and Lemmas Author: Peter Lammich Maintainer: Peter Lammich Thomas Tuerk *) (* CHANGELOG: 2010-05-09: Removed AC, AI locales, they are superseeded by concepts from OrderedGroups 2010-09-22: Merges with ext/Aux 2017-06-12: Added a bunch of lemmas from various other projects *) section \Miscellaneous Definitions and Lemmas\ theory Misc imports Main "HOL-Library.Multiset" "HOL-ex.Quicksort" "HOL-Library.Option_ord" "HOL-Library.Infinite_Set" "HOL-Eisbach.Eisbach" begin text_raw \\label{thy:Misc}\ subsection \Isabelle Distribution Move Proposals\ subsubsection \Pure\ lemma TERMI: "TERM x" unfolding Pure.term_def . subsubsection \HOL\ (* Stronger disjunction elimination rules. *) lemma disjE1: "\ P \ Q; P \ R; \\P;Q\ \ R \ \ R" by metis lemma disjE2: "\ P \ Q; \P; \Q\ \ R; Q \ R \ \ R" by blast lemma imp_mp_iff[simp]: "a \ (a \ b) \ a \ b" "(a \ b) \ a \ a \ b" (* is Inductive.imp_conj_iff, but not in simpset by default *) by blast+ lemma atomize_Trueprop_eq[atomize]: "(Trueprop x \ Trueprop y) \ Trueprop (x=y)" apply rule apply (rule) apply (erule equal_elim_rule1) apply assumption apply (erule equal_elim_rule2) apply assumption apply simp done subsubsection \Set\ lemma inter_compl_diff_conv[simp]: "A \ -B = A - B" by auto lemma pair_set_inverse[simp]: "{(a,b). P a b}\ = {(b,a). P a b}" by auto lemma card_doubleton_eq_2_iff[simp]: "card {a,b} = 2 \ a\b" by auto subsubsection \List\ (* TODO: Move, analogous to List.length_greater_0_conv *) thm List.length_greater_0_conv lemma length_ge_1_conv[iff]: "Suc 0 \ length l \ l\[]" by (cases l) auto \ \Obtains a list from the pointwise characterization of its elements\ lemma obtain_list_from_elements: assumes A: "\ili. P li i)" obtains l where "length l = n" "\il. length l=n \ (\iii l!j" by (simp add: sorted_iff_nth_mono) also from nth_eq_iff_index_eq[OF D] B have "l!i \ l!j" by auto finally show ?thesis . qed lemma distinct_sorted_strict_mono_iff: assumes "distinct l" "sorted l" assumes "i il!j \ i\j" by (metis assms distinct_sorted_strict_mono_iff leD le_less_linear) (* List.thy has: declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] We could, analogously, declare rules for "map _ _ = _@_" as dest!, or use elim!, or declare the _conv-rule as simp *) lemma map_eq_appendE: assumes "map f ls = fl@fl'" obtains l l' where "ls=l@l'" and "map f l=fl" and "map f l' = fl'" using assms proof (induction fl arbitrary: ls thesis) case (Cons x xs) then obtain l ls' where [simp]: "ls = l#ls'" "f l = x" by force with Cons.prems(2) have "map f ls' = xs @ fl'" by simp - from Cons.IH[OF _ this] guess ll ll' . + from Cons.IH[OF _ this] obtain ll ll' where "ls' = ll @ ll'" "map f ll = xs" "map f ll' = fl'" . with Cons.prems(1)[of "l#ll" ll'] show thesis by simp qed simp lemma map_eq_append_conv: "map f ls = fl@fl' \ (\l l'. ls = l@l' \ map f l = fl \ map f l' = fl')" by (auto elim!: map_eq_appendE) lemmas append_eq_mapE = map_eq_appendE[OF sym] lemma append_eq_map_conv: "fl@fl' = map f ls \ (\l l'. ls = l@l' \ map f l = fl \ map f l' = fl')" by (auto elim!: append_eq_mapE) lemma distinct_mapI: "distinct (map f l) \ distinct l" by (induct l) auto lemma map_distinct_upd_conv: "\i \ (map f l)[i := x] = map (f(l!i := x)) l" \ \Updating a mapped distinct list is equal to updating the mapping function\ by (auto simp: nth_eq_iff_index_eq intro: nth_equalityI) lemma zip_inj: "\length a = length b; length a' = length b'; zip a b = zip a' b'\ \ a=a' \ b=b'" proof (induct a b arbitrary: a' b' rule: list_induct2) case Nil then show ?case by (cases a'; cases b'; auto) next case (Cons x xs y ys) then show ?case by (cases a'; cases b'; auto) qed lemma zip_eq_zip_same_len[simp]: "\ length a = length b; length a' = length b' \ \ zip a b = zip a' b' \ a=a' \ b=b'" by (auto dest: zip_inj) lemma upt_merge[simp]: "i\j \ j\k \ [i.. (ys \ [] \ butlast ys = xs \ last ys = x)" by auto (* Case distinction how two elements of a list can be related to each other *) lemma list_match_lel_lel: assumes "c1 @ qs # c2 = c1' @ qs' # c2'" obtains (left) c21' where "c1 = c1' @ qs' # c21'" "c2' = c21' @ qs # c2" | (center) "c1' = c1" "qs' = qs" "c2' = c2" | (right) c21 where "c1' = c1 @ qs # c21" "c2 = c21 @ qs' # c2'" using assms apply (clarsimp simp: append_eq_append_conv2) subgoal for us by (cases us) auto done lemma xy_in_set_cases[consumes 2, case_names EQ XY YX]: assumes A: "x\set l" "y\set l" and C: "!!l1 l2. \ x=y; l=l1@y#l2 \ \ P" "!!l1 l2 l3. \ x\y; l=l1@x#l2@y#l3 \ \ P" "!!l1 l2 l3. \ x\y; l=l1@y#l2@x#l3 \ \ P" shows P proof (cases "x=y") case True with A(1) obtain l1 l2 where "l=l1@y#l2" by (blast dest: split_list) with C(1) True show ?thesis by blast next case False from A(1) obtain l1 l2 where S1: "l=l1@x#l2" by (blast dest: split_list) from A(2) obtain l1' l2' where S2: "l=l1'@y#l2'" by (blast dest: split_list) from S1 S2 have M: "l1@x#l2 = l1'@y#l2'" by simp thus P proof (cases rule: list_match_lel_lel[consumes 1, case_names 1 2 3]) case (1 c) with S1 have "l=l1'@y#c@x#l2" by simp with C(3) False show ?thesis by blast next case 2 with False have False by blast thus ?thesis .. next case (3 c) with S1 have "l=l1@x#c@y#l2'" by simp with C(2) False show ?thesis by blast qed qed lemma list_e_eq_lel[simp]: "[e] = l1@e'#l2 \ l1=[] \ e'=e \ l2=[]" "l1@e'#l2 = [e] \ l1=[] \ e'=e \ l2=[]" apply (cases l1, auto) [] apply (cases l1, auto) [] done lemma list_ee_eq_leel[simp]: "([e1,e2] = l1@e1'#e2'#l2) \ (l1=[] \ e1=e1' \ e2=e2' \ l2=[])" "(l1@e1'#e2'#l2 = [e1,e2]) \ (l1=[] \ e1=e1' \ e2=e2' \ l2=[])" apply (cases l1, auto) [] apply (cases l1, auto) [] done subsubsection \Transitive Closure\ text \A point-free induction rule for elements reachable from an initial set\ lemma rtrancl_reachable_induct[consumes 0, case_names base step]: assumes I0: "I \ INV" assumes IS: "E``INV \ INV" shows "E\<^sup>*``I \ INV" by (metis I0 IS Image_closed_trancl Image_mono subset_refl) lemma acyclic_empty[simp, intro!]: "acyclic {}" by (unfold acyclic_def) auto lemma acyclic_union: "acyclic (A\B) \ acyclic A" "acyclic (A\B) \ acyclic B" by (metis Un_upper1 Un_upper2 acyclic_subset)+ text \Here we provide a collection of miscellaneous definitions and helper lemmas\ subsection "Miscellaneous (1)" text \This stuff is used in this theory itself, and thus occurs in first place or is simply not sorted into any other section of this theory.\ lemma IdD: "(a,b)\Id \ a=b" by simp text \Conversion Tag\ definition [simp]: "CNV x y \ x=y" lemma CNV_I: "CNV x x" by simp lemma CNV_eqD: "CNV x y \ x=y" by simp lemma CNV_meqD: "CNV x y \ x\y" by simp (* TODO: Move. Shouldn't this be detected by simproc? *) lemma ex_b_b_and_simp[simp]: "(\b. b \ Q b) \ Q True" by auto lemma ex_b_not_b_and_simp[simp]: "(\b. \b \ Q b) \ Q False" by auto method repeat_all_new methods m = m;(repeat_all_new \m\)? subsubsection "AC-operators" text \Locale to declare AC-laws as simplification rules\ locale Assoc = fixes f assumes assoc[simp]: "f (f x y) z = f x (f y z)" locale AC = Assoc + assumes commute[simp]: "f x y = f y x" lemma (in AC) left_commute[simp]: "f x (f y z) = f y (f x z)" by (simp only: assoc[symmetric]) simp lemmas (in AC) AC_simps = commute assoc left_commute text \Locale to define functions from surjective, unique relations\ locale su_rel_fun = fixes F and f assumes unique: "\(A,B)\F; (A,B')\F\ \ B=B'" assumes surjective: "\!!B. (A,B)\F \ P\ \ P" assumes f_def: "f A == THE B. (A,B)\F" lemma (in su_rel_fun) repr1: "(A,f A)\F" proof (unfold f_def) obtain B where "(A,B)\F" by (rule surjective) with theI[where P="\B. (A,B)\F", OF this] show "(A, THE x. (A, x) \ F) \ F" by (blast intro: unique) qed lemma (in su_rel_fun) repr2: "(A,B)\F \ B=f A" using repr1 by (blast intro: unique) lemma (in su_rel_fun) repr: "(f A = B) = ((A,B)\F)" using repr1 repr2 by (blast) \ \Contract quantification over two variables to pair\ lemma Ex_prod_contract: "(\a b. P a b) \ (\z. P (fst z) (snd z))" by auto lemma All_prod_contract: "(\a b. P a b) \ (\z. P (fst z) (snd z))" by auto lemma nat_geq_1_eq_neqz: "x\1 \ x\(0::nat)" by auto lemma nat_in_between_eq: "(a b\Suc a) \ b = Suc a" "(a\b \ b b = a" by auto lemma Suc_n_minus_m_eq: "\ n\m; m>1 \ \ Suc (n - m) = n - (m - 1)" by simp lemma Suc_to_right: "Suc n = m \ n = m - Suc 0" by simp lemma Suc_diff[simp]: "\n m. n\m \ m\1 \ Suc (n - m) = n - (m - 1)" by simp lemma if_not_swap[simp]: "(if \c then a else b) = (if c then b else a)" by auto lemma all_to_meta: "Trueprop (\a. P a) \ (\a. P a)" apply rule by auto lemma imp_to_meta: "Trueprop (P\Q) \ (P\Q)" apply rule by auto (* for some reason, there is no such rule in HOL *) lemma iffI2: "\P \ Q; \ P \ \ Q\ \ P \ Q" by metis lemma iffExI: "\ \x. P x \ Q x; \x. Q x \ P x \ \ (\x. P x) \ (\x. Q x)" by metis lemma bex2I[intro?]: "\ (a,b)\S; (a,b)\S \ P a b \ \ \a b. (a,b)\S \ P a b" by blast (* TODO: Move lemma to HOL! *) lemma cnv_conj_to_meta: "(P \ Q \ PROP X) \ (\P;Q\ \ PROP X)" by (rule BNF_Fixpoint_Base.conj_imp_eq_imp_imp) subsection \Sets\ lemma remove_subset: "x\S \ S-{x} \ S" by auto lemma subset_minus_empty: "A\B \ A-B = {}" by auto lemma insert_minus_eq: "x\y \ insert x A - {y} = insert x (A - {y})" by auto lemma set_notEmptyE: "\S\{}; !!x. x\S \ P\ \ P" by (metis equals0I) lemma subset_Collect_conv: "S \ Collect P \ (\x\S. P x)" by auto lemma memb_imp_not_empty: "x\S \ S\{}" by auto lemma disjoint_mono: "\ a\a'; b\b'; a'\b'={} \ \ a\b={}" by auto lemma disjoint_alt_simp1: "A-B = A \ A\B = {}" by auto lemma disjoint_alt_simp2: "A-B \ A \ A\B \ {}" by auto lemma disjoint_alt_simp3: "A-B \ A \ A\B \ {}" by auto lemma disjointI[intro?]: "\ \x. \x\a; x\b\ \ False \ \ a\b={}" by auto lemmas set_simps = subset_minus_empty disjoint_alt_simp1 disjoint_alt_simp2 disjoint_alt_simp3 Un_absorb1 Un_absorb2 lemma set_minus_singleton_eq: "x\X \ X-{x} = X" by auto lemma set_diff_diff_left: "A-B-C = A-(B\C)" by auto lemma image_update[simp]: "x\A \ f(x:=n)`A = f`A" by auto lemma eq_or_mem_image_simp[simp]: "{f l |l. l = a \ l\B} = insert (f a) {f l|l. l\B}" by blast lemma set_union_code [code_unfold]: "set xs \ set ys = set (xs @ ys)" by auto lemma in_fst_imageE: assumes "x \ fst`S" obtains y where "(x,y)\S" using assms by auto lemma in_snd_imageE: assumes "y \ snd`S" obtains x where "(x,y)\S" using assms by auto lemma fst_image_mp: "\fst`A \ B; (x,y)\A \ \ x\B" by (metis Domain.DomainI fst_eq_Domain in_mono) lemma snd_image_mp: "\snd`A \ B; (x,y)\A \ \ y\B" by (metis Range.intros rev_subsetD snd_eq_Range) lemma inter_eq_subsetI: "\ S\S'; A\S' = B\S' \ \ A\S = B\S" by auto text \ Decompose general union over sum types. \ lemma Union_plus: "(\ x \ A <+> B. f x) = (\ a \ A. f (Inl a)) \ (\b \ B. f (Inr b))" by auto lemma Union_sum: "(\x. f (x::'a+'b)) = (\l. f (Inl l)) \ (\r. f (Inr r))" (is "?lhs = ?rhs") proof - have "?lhs = (\x \ UNIV <+> UNIV. f x)" by simp thus ?thesis by (simp only: Union_plus) qed subsubsection \Finite Sets\ lemma card_1_singletonI: "\finite S; card S = 1; x\S\ \ S={x}" proof (safe, rule ccontr, goal_cases) case prems: (1 x') hence "finite (S-{x})" "S-{x} \ {}" by auto hence "card (S-{x}) \ 0" by auto moreover from prems(1-3) have "card (S-{x}) = 0" by auto ultimately have False by simp thus ?case .. qed lemma card_insert_disjoint': "\finite A; x \ A\ \ card (insert x A) - Suc 0 = card A" by (drule (1) card_insert_disjoint) auto lemma card_eq_UNIV[simp]: "card (S::'a::finite set) = card (UNIV::'a set) \ S=UNIV" proof (auto) fix x assume A: "card S = card (UNIV::'a set)" show "x\S" proof (rule ccontr) assume "x\S" hence "S\UNIV" by auto with psubset_card_mono[of UNIV S] have "card S < card (UNIV::'a set)" by auto with A show False by simp qed qed lemma card_eq_UNIV2[simp]: "card (UNIV::'a set) = card (S::'a::finite set) \ S=UNIV" using card_eq_UNIV[of S] by metis lemma card_ge_UNIV[simp]: "card (UNIV::'a::finite set) \ card (S::'a set) \ S=UNIV" using card_mono[of "UNIV::'a::finite set" S, simplified] by auto lemmas length_remdups_card = length_remdups_concat[of "[l]", simplified] for l lemma fs_contract: "fst ` { p | p. f (fst p) (snd p) \ S } = { a . \b. f a b \ S }" by (simp add: image_Collect) lemma finite_Collect: "finite S \ inj f \ finite {a. f a : S}" by(simp add: finite_vimageI vimage_def[symmetric]) \ \Finite sets have an injective mapping to an initial segments of the natural numbers\ (* This lemma is also in the standard library (from Isabelle2009-1 on) as @{thm [source] Finite_Set.finite_imp_inj_to_nat_seg}. However, it is formulated with HOL's \ there rather then with the meta-logic obtain *) lemma finite_imp_inj_to_nat_seg': fixes A :: "'a set" assumes A: "finite A" obtains f::"'a \ nat" and n::"nat" where "f`A = {i. i finite (lists P \ { l. length l = n })" proof (induct n) case 0 thus ?case by auto next case (Suc n) have "lists P \ { l. length l = Suc n } = (\(a,l). a#l) ` (P \ (lists P \ {l. length l = n}))" apply auto apply (case_tac x) apply auto done moreover from Suc have "finite \" by auto ultimately show ?case by simp qed lemma lists_of_len_fin2: "finite P \ finite (lists P \ { l. n = length l })" proof - assume A: "finite P" have S: "{ l. n = length l } = { l. length l = n }" by auto have "finite (lists P \ { l. n = length l }) \ finite (lists P \ { l. length l = n })" by (subst S) simp thus ?thesis using lists_of_len_fin1[OF A] by auto qed lemmas lists_of_len_fin = lists_of_len_fin1 lists_of_len_fin2 (* Try (simp only: cset_fin_simps, fastforce intro: cset_fin_intros) when reasoning about finiteness of collected sets *) lemmas cset_fin_simps = Ex_prod_contract fs_contract[symmetric] image_Collect[symmetric] lemmas cset_fin_intros = finite_imageI finite_Collect inj_onI lemma Un_interval: fixes b1 :: "'a::linorder" assumes "b1\b2" and "b2\b3" shows "{ f i | i. b1\i \ i { f i | i. b2\i \ ii \ i The standard library proves that a generalized union is finite if the index set is finite and if for every index the component set is itself finite. Conversely, we show that every component set must be finite when the union is finite. \ lemma finite_UNION_then_finite: "finite (\(B ` A)) \ a \ A \ finite (B a)" by (metis Set.set_insert UN_insert Un_infinite) lemma finite_if_eq_beyond_finite: "finite S \ finite {s. s - S = s' - S}" proof (rule finite_subset[where B="(\s. s \ (s' - S)) ` Pow S"], clarsimp) fix s have "s = (s \ S) \ (s - S)" by auto also assume "s - S = s' - S" finally show "s \ (\s. s \ (s' - S)) ` Pow S" by blast qed blast lemma distinct_finite_subset: assumes "finite x" shows "finite {ys. set ys \ x \ distinct ys}" (is "finite ?S") proof (rule finite_subset) from assms show "?S \ {ys. set ys \ x \ length ys \ card x}" by clarsimp (metis distinct_card card_mono) from assms show "finite ..." by (rule finite_lists_length_le) qed lemma distinct_finite_set: shows "finite {ys. set ys = x \ distinct ys}" (is "finite ?S") proof (cases "finite x") case False hence "{ys. set ys = x} = {}" by auto thus ?thesis by simp next case True show ?thesis proof (rule finite_subset) show "?S \ {ys. set ys \ x \ length ys \ card x}" using distinct_card by force from True show "finite ..." by (rule finite_lists_length_le) qed qed lemma finite_set_image: assumes f: "finite (set ` A)" and dist: "\xs. xs \ A \ distinct xs" shows "finite A" proof (rule finite_subset) from f show "finite (set -` (set ` A) \ {xs. distinct xs})" proof (induct rule: finite_induct) case (insert x F) have "finite (set -` {x} \ {xs. distinct xs})" apply (simp add: vimage_def) by (metis Collect_conj_eq distinct_finite_set) with insert show ?case apply (subst vimage_insert) apply (subst Int_Un_distrib2) apply (rule finite_UnI) apply simp_all done qed simp moreover from dist show "A \ ..." by (auto simp add: vimage_image_eq) qed subsubsection \Infinite Set\ lemma INFM_nat_inductI: assumes P0: "P (0::nat)" assumes PS: "\i. P i \ \j>i. P j \ Q j" shows "\\<^sub>\i. Q i" proof - have "\i. \j>i. P j \ Q j" proof fix i show "\j>i. P j \ Q j" apply (induction i) using PS[OF P0] apply auto [] by (metis PS Suc_lessI) qed thus ?thesis unfolding INFM_nat by blast qed subsection \Functions\ lemma fun_neq_ext_iff: "m\m' \ (\x. m x \ m' x)" by auto definition "inv_on f A x == SOME y. y\A \ f y = x" lemma inv_on_f_f[simp]: "\inj_on f A; x\A\ \ inv_on f A (f x) = x" by (auto simp add: inv_on_def inj_on_def) lemma f_inv_on_f: "\ y\f`A \ \ f (inv_on f A y) = y" by (auto simp add: inv_on_def intro: someI2) lemma inv_on_f_range: "\ y \ f`A \ \ inv_on f A y \ A" by (auto simp add: inv_on_def intro: someI2) lemma inj_on_map_inv_f [simp]: "\set l \ A; inj_on f A\ \ map (inv_on f A) (map f l) = l" apply (simp) apply (induct l) apply auto done lemma comp_cong_right: "x = y \ f o x = f o y" by (simp) lemma comp_cong_left: "x = y \ x o f = y o f" by (simp) lemma fun_comp_eq_conv: "f o g = fg \ (\x. f (g x) = fg x)" by auto abbreviation comp2 (infixl "oo" 55) where "f oo g \ \x. f o (g x)" abbreviation comp3 (infixl "ooo" 55) where "f ooo g \ \x. f oo (g x)" notation comp2 (infixl "\\" 55) and comp3 (infixl "\\\" 55) definition [code_unfold, simp]: "swap_args2 f x y \ f y x" subsection \Multisets\ (* The following is a syntax extension for multisets. Unfortunately, it depends on a change in the Library/Multiset.thy, so it is commented out here, until it will be incorporated into Library/Multiset.thy by its maintainers. The required change in Library/Multiset.thy is removing the syntax for single: - single :: "'a => 'a multiset" ("{#_#}") + single :: "'a => 'a multiset" And adding the following translations instead: + syntax + "_multiset" :: "args \ 'a multiset" ("{#(_)#}") + translations + "{#x, xs#}" == "{#x#} + {#xs#}" + "{# x #}" == "single x" This translates "{# \ #}" into a sum of singletons, that is parenthesized to the right. ?? Can we also achieve left-parenthesizing ?? *) (* Let's try what happens if declaring AC-rules for multiset union as simp-rules *) (*declare union_ac[simp] -- don't do it !*) lemma count_mset_set_finite_iff: "finite S \ count (mset_set S) a = (if a \ S then 1 else 0)" by simp lemma ex_Melem_conv: "(\x. x \# A) = (A \ {#})" by (simp add: ex_in_conv) subsubsection \Count\ lemma count_ne_remove: "\ x ~= t\ \ count S x = count (S-{#t#}) x" by (auto) lemma mset_empty_count[simp]: "(\p. count M p = 0) = (M={#})" by (auto simp add: multiset_eq_iff) subsubsection \Union, difference and intersection\ lemma size_diff_se: "t \# S \ size S = size (S - {#t#}) + 1" proof (unfold size_multiset_overloaded_eq) let ?SIZE = "sum (count S) (set_mset S)" assume A: "t \# S" from A have SPLITPRE: "finite (set_mset S) & {t}\(set_mset S)" by auto hence "?SIZE = sum (count S) (set_mset S - {t}) + sum (count S) {t}" by (blast dest: sum.subset_diff) hence "?SIZE = sum (count S) (set_mset S - {t}) + count (S) t" by auto moreover with A have "count S t = count (S-{#t#}) t + 1" by auto ultimately have D: "?SIZE = sum (count S) (set_mset S - {t}) + count (S-{#t#}) t + 1" by (arith) moreover have "sum (count S) (set_mset S - {t}) = sum (count (S-{#t#})) (set_mset S - {t})" proof - have "\x\(set_mset S - {t}). count S x = count (S-{#t#}) x" by (auto iff add: count_ne_remove) thus ?thesis by simp qed ultimately have D: "?SIZE = sum (count (S-{#t#})) (set_mset S - {t}) + count (S-{#t#}) t + 1" by (simp) moreover { assume CASE: "t \# S - {#t#}" from CASE have "set_mset S - {t} = set_mset (S - {#t#})" by (auto simp add: in_diff_count split: if_splits) with CASE D have "?SIZE = sum (count (S-{#t#})) (set_mset (S - {#t#})) + 1" by (simp add: not_in_iff) } moreover { assume CASE: "t \# S - {#t#}" from CASE have "t \# S" by (rule in_diffD) with CASE have 1: "set_mset S = set_mset (S-{#t#})" by (auto simp add: in_diff_count split: if_splits) moreover from D have "?SIZE = sum (count (S-{#t#})) (set_mset S - {t}) + sum (count (S-{#t#})) {t} + 1" by simp moreover from SPLITPRE sum.subset_diff have "sum (count (S-{#t#})) (set_mset S) = sum (count (S-{#t#})) (set_mset S - {t}) + sum (count (S-{#t#})) {t}" by (blast) ultimately have "?SIZE = sum (count (S-{#t#})) (set_mset (S-{#t#})) + 1" by simp } ultimately show "?SIZE = sum (count (S-{#t#})) (set_mset (S - {#t#})) + 1" by blast qed (* TODO: Check whether this proof can be done simpler *) lemma mset_union_diff_comm: "t \# S \ T + (S - {#t#}) = (T + S) - {#t#}" proof - assume "t \# S" then obtain S' where S: "S = add_mset t S'" by (metis insert_DiffM) then show ?thesis by auto qed (* lemma mset_diff_diff_left: "A-B-C = A-((B::'a multiset)+C)" proof - have "\e . count (A-B-C) e = count (A-(B+C)) e" by auto thus ?thesis by (simp add: multiset_eq_conv_count_eq) qed lemma mset_diff_commute: "A-B-C = A-C-(B::'a multiset)" proof - have "A-B-C = A-(B+C)" by (simp add: mset_diff_diff_left) also have "\ = A-(C+B)" by (simp add: union_commute) thus ?thesis by (simp add: mset_diff_diff_left) qed lemma mset_diff_same_empty[simp]: "(S::'a multiset) - S = {#}" proof - have "\e . count (S-S) e = 0" by auto hence "\e . ~ (e : set_mset (S-S))" by auto hence "set_mset (S-S) = {}" by blast thus ?thesis by (auto) qed *) lemma mset_right_cancel_union: "\a \# A+B; ~(a \# B)\ \ a\#A" by (simp) lemma mset_left_cancel_union: "\a \# A+B; ~(a \# A)\ \ a\#B" by (simp) lemmas mset_cancel_union = mset_right_cancel_union mset_left_cancel_union lemma mset_right_cancel_elem: "\a \# A+{#b#}; a~=b\ \ a\#A" by simp lemma mset_left_cancel_elem: "\a \# {#b#}+A; a~=b\ \ a\#A" by simp lemmas mset_cancel_elem = mset_right_cancel_elem mset_left_cancel_elem lemma mset_diff_cancel1elem[simp]: "~(a \# B) \ {#a#}-B = {#a#}" by (auto simp add: not_in_iff intro!: multiset_eqI) (* lemma diff_union_inverse[simp]: "A + B - B = (A::'a multiset)" by (auto iff add: multiset_eq_conv_count_eq) lemma diff_union_inverse2[simp]: "B + A - B = (A::'a multiset)" by (auto iff add: multiset_eq_conv_count_eq) *) (*lemma union_diff_assoc_se2: "t \# A \ (A+B)-{#t#} = (A-{#t#}) + B" by (auto iff add: multiset_eq_conv_count_eq) lemmas union_diff_assoc_se = union_diff_assoc_se1 union_diff_assoc_se2*) lemma union_diff_assoc: "C-B={#} \ (A+B)-C = A + (B-C)" by (simp add: multiset_eq_iff) lemmas mset_neutral_cancel1 = union_left_cancel[where N="{#}", simplified] union_right_cancel[where N="{#}", simplified] declare mset_neutral_cancel1[simp] lemma mset_union_2_elem: "{#a, b#} = add_mset c M \ {#a#}=M & b=c | a=c & {#b#}=M" by (auto simp: add_eq_conv_diff) lemma mset_un_cases[cases set, case_names left right]: "\a \# A + B; a \# A \ P; a \# B \ P\ \ P" by (auto) lemma mset_unplusm_dist_cases[cases set, case_names left right]: assumes A: "{#s#}+A = B+C" assumes L: "\B={#s#}+(B-{#s#}); A=(B-{#s#})+C\ \ P" assumes R: "\C={#s#}+(C-{#s#}); A=B+(C-{#s#})\ \ P" shows P proof - from A[symmetric] have "s \# B+C" by simp thus ?thesis proof (cases rule: mset_un_cases) case left hence 1: "B={#s#}+(B-{#s#})" by simp with A have "{#s#}+A = {#s#}+((B-{#s#})+C)" by (simp add: union_ac) hence 2: "A = (B-{#s#})+C" by (simp) from L[OF 1 2] show ?thesis . next case right hence 1: "C={#s#}+(C-{#s#})" by simp with A have "{#s#}+A = {#s#}+(B+(C-{#s#}))" by (simp add: union_ac) hence 2: "A = B+(C-{#s#})" by (simp) from R[OF 1 2] show ?thesis . qed qed lemma mset_unplusm_dist_cases2[cases set, case_names left right]: assumes A: "B+C = {#s#}+A" assumes L: "\B={#s#}+(B-{#s#}); A=(B-{#s#})+C\ \ P" assumes R: "\C={#s#}+(C-{#s#}); A=B+(C-{#s#})\ \ P" shows P using mset_unplusm_dist_cases[OF A[symmetric]] L R by blast lemma mset_single_cases[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "\c'={#s#}+(c'-{#s#}); c={#r'#}+(c-{#r'#}); c-{#r'#} = c'-{#s#} \ \ P" shows "P" proof - { assume CASE: "s=r'" with A have "c=c'" by simp with CASE CASES have ?thesis by auto } moreover { assume CASE: "s\r'" have "s \# {#s#}+c" by simp with A have "s \# {#r'#}+c'" by simp with CASE have "s \# c'" by simp hence 1: "c' = {#s#} + (c' - {#s#})" by simp with A have "{#s#}+c = {#s#}+({#r'#}+(c' - {#s#}))" by (auto simp add: union_ac) hence 2: "c={#r'#}+(c' - {#s#})" by (auto) hence 3: "c-{#r'#} = (c' - {#s#})" by auto from 1 2 3 CASES have ?thesis by auto } ultimately show ?thesis by blast qed lemma mset_single_cases'[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "!!cc. \c'={#s#}+cc; c={#r'#}+cc; c'-{#s#}=cc; c-{#r'#}=cc\ \ P" shows "P" using A CASES by (auto elim!: mset_single_cases) lemma mset_single_cases2[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "\c'=(c'-{#s#})+{#s#}; c=(c-{#r'#})+{#r'#}; c-{#r'#} = c'-{#s#} \ \ P" shows "P" proof - from A have "add_mset s c = add_mset r' c'" by (simp add: union_ac) thus ?thesis using CASES by (cases rule: mset_single_cases) simp_all qed lemma mset_single_cases2'[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "!!cc. \c'=cc+{#s#}; c=cc+{#r'#}; c'-{#s#}=cc; c-{#r'#}=cc\ \ P" shows "P" using A CASES by (auto elim!: mset_single_cases2) lemma mset_un_single_un_cases [consumes 1, case_names left right]: assumes A: "add_mset a A = B + C" and CASES: "a \# B \ A = (B - {#a#}) + C \ P" "a \# C \ A = B + (C - {#a#}) \ P" shows "P" proof - have "a \# A+{#a#}" by simp with A have "a \# B+C" by auto thus ?thesis proof (cases rule: mset_un_cases) case left hence "B=B-{#a#}+{#a#}" by auto with A have "A+{#a#} = (B-{#a#})+C+{#a#}" by (auto simp add: union_ac) hence "A=(B-{#a#})+C" by simp with CASES(1)[OF left] show ?thesis by blast next case right hence "C=C-{#a#}+{#a#}" by auto with A have "A+{#a#} = B+(C-{#a#})+{#a#}" by (auto simp add: union_ac) hence "A=B+(C-{#a#})" by simp with CASES(2)[OF right] show ?thesis by blast qed qed (* TODO: Can this proof be done more automatically ? *) lemma mset_distrib[consumes 1, case_names dist]: assumes A: "(A::'a multiset)+B = M+N" "!!Am An Bm Bn. \A=Am+An; B=Bm+Bn; M=Am+Bm; N=An+Bn\ \ P" shows "P" proof - have BN_MA: "B - N = M - A" by (metis (no_types) add_diff_cancel_right assms(1) union_commute) have H: "A = A\# C + (A - C) \# D" if "A + B = C + D" for A B C D :: "'a multiset" by (metis add.commute diff_intersect_left_idem mset_subset_eq_add_left subset_eq_diff_conv subset_mset.add_diff_inverse subset_mset.inf_absorb1 subset_mset.inf_le1 that) have A': "A = A\# M + (A - M) \# N" using A(1) H by blast moreover have B': "B = (B - N) \# M + B\# N" using A(1) H[of B A N M] by (auto simp: ac_simps) moreover have "M = A \# M + (B - N) \# M" using H[of M N A B] BN_MA[symmetric] A(1) by (metis (no_types) diff_intersect_left_idem diff_union_cancelR multiset_inter_commute subset_mset.diff_add subset_mset.inf.cobounded1 union_commute) moreover have "N = (A - M) \# N + B \# N" by (metis A' assms(1) diff_union_cancelL inter_union_distrib_left inter_union_distrib_right mset_subset_eq_multiset_union_diff_commute subset_mset.inf.cobounded1 subset_mset.inf.commute) ultimately show P using A(2) by blast qed subsubsection \Singleton multisets\ lemma mset_size_le1_cases[case_names empty singleton,consumes 1]: "\ size M \ Suc 0; M={#} \ P; !!m. M={#m#} \ P \ \ P" by (cases M) auto lemma diff_union_single_conv2: "a \# J \ J + I - {#a#} = (J - {#a#}) + I" by simp lemmas diff_union_single_convs = diff_union_single_conv diff_union_single_conv2 lemma mset_contains_eq: "(m \# M) = ({#m#}+(M-{#m#})=M)" using diff_single_trivial by fastforce subsubsection \Pointwise ordering\ (*declare mset_le_trans[trans] Seems to be in there now. Why is this not done in Multiset.thy or order-class ? *) lemma mset_le_incr_right1: "a\#(b::'a multiset) \ a\#b+c" using mset_subset_eq_mono_add[of a b "{#}" c, simplified] . lemma mset_le_incr_right2: "a\#(b::'a multiset) \ a\#c+b" using mset_le_incr_right1 by (auto simp add: union_commute) lemmas mset_le_incr_right = mset_le_incr_right1 mset_le_incr_right2 lemma mset_le_decr_left1: "a+c\#(b::'a multiset) \ a\#b" using mset_le_incr_right1 mset_subset_eq_mono_add_right_cancel by blast lemma mset_le_decr_left2: "c+a\#(b::'a multiset) \ a\#b" using mset_le_decr_left1 by (auto simp add: union_ac) lemma mset_le_add_mset_decr_left1: "add_mset c a\#(b::'a multiset) \ a\#b" by (simp add: mset_subset_eq_insertD subset_mset.dual_order.strict_implies_order) lemma mset_le_add_mset_decr_left2: "add_mset c a\#(b::'a multiset) \ {#c#}\#b" by (simp add: mset_subset_eq_insertD subset_mset.dual_order.strict_implies_order) lemmas mset_le_decr_left = mset_le_decr_left1 mset_le_decr_left2 mset_le_add_mset_decr_left1 mset_le_add_mset_decr_left2 lemma mset_le_subtract: "A\#B \ A-C \# B-(C::'a multiset)" apply (unfold subseteq_mset_def count_diff) apply clarify apply (subgoal_tac "count A a \ count B a") apply arith apply simp done lemma mset_union_subset: "A+B \# C \ A\#C \ B\#(C::'a multiset)" by (auto dest: mset_le_decr_left) lemma mset_le_add_mset: "add_mset x B \# C \ {#x#}\#C \ B\#(C::'a multiset)" by (auto dest: mset_le_decr_left) lemma mset_le_subtract_left: "A+B \# (X::'a multiset) \ B \# X-A \ A\#X" by (auto dest: mset_le_subtract[of "A+B" "X" "A"] mset_union_subset) lemma mset_le_subtract_right: "A+B \# (X::'a multiset) \ A \# X-B \ B\#X" by (auto dest: mset_le_subtract[of "A+B" "X" "B"] mset_union_subset) lemma mset_le_subtract_add_mset_left: "add_mset x B \# (X::'a multiset) \ B \# X-{#x#} \ {#x#}\#X" by (auto dest: mset_le_subtract[of "add_mset x B" "X" "{#x#}"] mset_le_add_mset) lemma mset_le_subtract_add_mset_right: "add_mset x B \# (X::'a multiset) \ {#x#} \# X-B \ B\#X" by (auto dest: mset_le_subtract[of "add_mset x B" "X" "B"] mset_le_add_mset) lemma mset_le_addE: "\ xs \# (ys::'a multiset); !!zs. ys=xs+zs \ P \ \ P" using mset_subset_eq_exists_conv by blast declare subset_mset.diff_add[simp, intro] lemma mset_2dist2_cases: assumes A: "{#a#}+{#b#} \# A+B" assumes CASES: "{#a#}+{#b#} \# A \ P" "{#a#}+{#b#} \# B \ P" "\a \# A; b \# B\ \ P" "\a \# B; b \# A\ \ P" shows "P" proof - { assume C: "a \# A" "b \# A-{#a#}" with mset_subset_eq_mono_add[of "{#a#}" "{#a#}" "{#b#}" "A-{#a#}"] have "{#a#}+{#b#} \# A" by auto } moreover { assume C: "a \# A" "\ (b \# A-{#a#})" with A have "b \# B" by (metis diff_union_single_conv2 mset_le_subtract_left mset_subset_eq_insertD mset_un_cases) } moreover { assume C: "\ (a \# A)" "b \# B-{#a#}" with A have "a \# B" by (auto dest: mset_subset_eqD) with C mset_subset_eq_mono_add[of "{#a#}" "{#a#}" "{#b#}" "B-{#a#}"] have "{#a#}+{#b#} \# B" by auto } moreover { assume C: "\ (a \# A)" "\ (b \# B-{#a#})" with A have "a \# B \ b \# A" apply (intro conjI) apply (auto dest!: mset_subset_eq_insertD simp: insert_union_subset_iff; fail)[] by (metis mset_diff_cancel1elem mset_le_subtract_left multiset_diff_union_assoc single_subset_iff subset_eq_diff_conv) } ultimately show P using CASES by blast qed lemma mset_union_subset_s: "{#a#}+B \# C \ a \# C \ B \# C" by (drule mset_union_subset) simp (* TODO: Check which of these lemmas are already introduced by order-classes ! *) lemma mset_le_single_cases[consumes 1, case_names empty singleton]: "\M\#{#a#}; M={#} \ P; M={#a#} \ P\ \ P" by (induct M) auto lemma mset_le_distrib[consumes 1, case_names dist]: "\(X::'a multiset)\#A+B; !!Xa Xb. \X=Xa+Xb; Xa\#A; Xb\#B\ \ P \ \ P" by (auto elim!: mset_le_addE mset_distrib) lemma mset_le_mono_add_single: "\a \# ys; b \# ws\ \ {#a#} + {#b#} \# ys + ws" by (meson mset_subset_eq_mono_add single_subset_iff) lemma mset_size1elem: "\size P \ 1; q \# P\ \ P={#q#}" by (auto elim: mset_size_le1_cases) lemma mset_size2elem: "\size P \ 2; {#q#}+{#q'#} \# P\ \ P={#q#}+{#q'#}" by (auto elim: mset_le_addE) subsubsection \Image under function\ notation image_mset (infixr "`#" 90) lemma mset_map_split_orig: "!!M1 M2. \f `# P = M1+M2; !!P1 P2. \P=P1+P2; f `# P1 = M1; f `# P2 = M2\ \ Q \ \ Q" by (induct P) (force elim!: mset_un_single_un_cases)+ lemma mset_map_id: "\!!x. f (g x) = x\ \ f `# g `# X = X" by (induct X) auto text \The following is a very specialized lemma. Intuitively, it splits the original multiset by a splitting of some pointwise supermultiset of its image. Application: This lemma came in handy when proving the correctness of a constraint system that collects at most k sized submultisets of the sets of spawned threads. \ lemma mset_map_split_orig_le: assumes A: "f `# P \# M1+M2" and EX: "!!P1 P2. \P=P1+P2; f `# P1 \# M1; f `# P2 \# M2\ \ Q" shows "Q" using A EX by (auto elim: mset_le_distrib mset_map_split_orig) subsection \Lists\ lemma len_greater_imp_nonempty[simp]: "length l > x \ l\[]" by auto lemma list_take_induct_tl2: "\length xs = length ys; \n \ \n < length (tl xs). P ((tl ys) ! n) ((tl xs) ! n)" by (induct xs ys rule: list_induct2) auto lemma not_distinct_split_distinct: assumes "\ distinct xs" obtains y ys zs where "distinct ys" "y \ set ys" "xs = ys@[y]@zs" using assms proof (induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc x xs) thus ?case by (cases "distinct xs") auto qed lemma distinct_length_le: assumes d: "distinct ys" and eq: "set ys = set xs" shows "length ys \ length xs" proof - from d have "length ys = card (set ys)" by (simp add: distinct_card) also from eq List.card_set have "card (set ys) = length (remdups xs)" by simp also have "... \ length xs" by simp finally show ?thesis . qed lemma find_SomeD: "List.find P xs = Some x \ P x" "List.find P xs = Some x \ x\set xs" by (auto simp add: find_Some_iff) lemma in_hd_or_tl_conv[simp]: "l\[] \ x=hd l \ x\set (tl l) \ x\set l" by (cases l) auto lemma length_dropWhile_takeWhile: assumes "x < length (dropWhile P xs)" shows "x + length (takeWhile P xs) < length xs" using assms by (induct xs) auto text \Elim-version of @{thm neq_Nil_conv}.\ lemma neq_NilE: assumes "l\[]" obtains x xs where "l=x#xs" using assms by (metis list.exhaust) lemma length_Suc_rev_conv: "length xs = Suc n \ (\ys y. xs=ys@[y] \ length ys = n)" by (cases xs rule: rev_cases) auto subsubsection \List Destructors\ lemma not_hd_in_tl: "x \ hd xs \ x \ set xs \ x \ set (tl xs)" by (induct xs) simp_all lemma distinct_hd_tl: "distinct xs \ x = hd xs \ x \ set (tl (xs))" by (induct xs) simp_all lemma in_set_tlD: "x \ set (tl xs) \ x \ set xs" by (induct xs) simp_all lemma nth_tl: "xs \ [] \ tl xs ! n = xs ! Suc n" by (induct xs) simp_all lemma tl_subset: "xs \ [] \ set xs \ A \ set (tl xs) \ A" by (metis in_set_tlD rev_subsetD subsetI) lemma tl_last: "tl xs \ [] \ last xs = last (tl xs)" by (induct xs) simp_all lemma tl_obtain_elem: assumes "xs \ []" "tl xs = []" obtains e where "xs = [e]" using assms by (induct xs rule: list_nonempty_induct) simp_all lemma butlast_subset: "xs \ [] \ set xs \ A \ set (butlast xs) \ A" by (metis in_set_butlastD rev_subsetD subsetI) lemma butlast_rev_tl: "xs \ [] \ butlast (rev xs) = rev (tl xs)" by (induct xs rule: rev_induct) simp_all lemma hd_butlast: "length xs > 1 \ hd (butlast xs) = hd xs" by (induct xs) simp_all lemma butlast_upd_last_eq[simp]: "length l \ 2 \ (butlast l) [ length l - 2 := x ] = take (length l - 2) l @ [x]" apply (case_tac l rule: rev_cases) apply simp apply simp apply (case_tac ys rule: rev_cases) apply simp apply simp done lemma distinct_butlast_swap[simp]: "distinct pq \ distinct (butlast (pq[i := last pq]))" apply (cases pq rule: rev_cases) apply (auto simp: list_update_append distinct_list_update split: nat.split) done subsubsection \Splitting list according to structure of other list\ context begin private definition "SPLIT_ACCORDING m l \ length l = length m" private lemma SPLIT_ACCORDINGE: assumes "length m = length l" obtains "SPLIT_ACCORDING m l" unfolding SPLIT_ACCORDING_def using assms by auto private lemma SPLIT_ACCORDING_simp: "SPLIT_ACCORDING m (l1@l2) \ (\m1 m2. m=m1@m2 \ SPLIT_ACCORDING m1 l1 \ SPLIT_ACCORDING m2 l2)" "SPLIT_ACCORDING m (x#l') \ (\y m'. m=y#m' \ SPLIT_ACCORDING m' l')" apply (fastforce simp: SPLIT_ACCORDING_def intro: exI[where x = "take (length l1) m"] exI[where x = "drop (length l1) m"]) apply (cases m;auto simp: SPLIT_ACCORDING_def) done text \Split structure of list @{term m} according to structure of list @{term l}.\ method split_list_according for m :: "'a list" and l :: "'b list" = (rule SPLIT_ACCORDINGE[of m l], (simp; fail), ( simp only: SPLIT_ACCORDING_simp, elim exE conjE, simp only: SPLIT_ACCORDING_def ) ) end subsubsection \\list_all2\\ lemma list_all2_induct[consumes 1, case_names Nil Cons]: assumes "list_all2 P l l'" assumes "Q [] []" assumes "\x x' ls ls'. \ P x x'; list_all2 P ls ls'; Q ls ls' \ \ Q (x#ls) (x'#ls')" shows "Q l l'" using list_all2_lengthD[OF assms(1)] assms apply (induct rule: list_induct2) apply auto done subsubsection \Indexing\ lemma ran_nth_set_encoding_conv[simp]: "ran (\i. if ii l[j:=x]!i = l!i" apply (induction l arbitrary: i j) apply (auto split: nat.splits) done lemma nth_list_update': "l[i:=x]!j = (if i=j \ i length l \ n\0 \ last (take n l) = l!(n - 1)" apply (induction l arbitrary: n) apply (auto simp: take_Cons split: nat.split) done lemma nth_append_first[simp]: "i (l@l')!i = l!i" by (simp add: nth_append) lemma in_set_image_conv_nth: "f x \ f`set l \ (\ii. i f (l!i) = f (l'!i)" shows "f`set l = f`set l'" using assms by (fastforce simp: in_set_conv_nth in_set_image_conv_nth) lemma insert_swap_set_eq: "i insert (l!i) (set (l[i:=x])) = insert x (set l)" by (auto simp: in_set_conv_nth nth_list_update split: if_split_asm) subsubsection \Reverse lists\ lemma neq_Nil_revE: assumes "l\[]" obtains ll e where "l = ll@[e]" using assms by (cases l rule: rev_cases) auto lemma neq_Nil_rev_conv: "l\[] \ (\xs x. l = xs@[x])" by (cases l rule: rev_cases) auto text \Caution: Same order of case variables in snoc-case as @{thm [source] rev_exhaust}, the other way round than @{thm [source] rev_induct} !\ lemma length_compl_rev_induct[case_names Nil snoc]: "\P []; !! l e . \!! ll . length ll <= length l \ P ll\ \ P (l@[e])\ \ P l" apply(induct_tac l rule: length_induct) apply(case_tac "xs" rule: rev_cases) apply(auto) done lemma list_append_eq_Cons_cases[consumes 1]: "\ys@zs = x#xs; \ys=[]; zs=x#xs\ \ P; !!ys'. \ ys=x#ys'; ys'@zs=xs \ \ P \ \ P" by (auto iff add: append_eq_Cons_conv) lemma list_Cons_eq_append_cases[consumes 1]: "\x#xs = ys@zs; \ys=[]; zs=x#xs\ \ P; !!ys'. \ ys=x#ys'; ys'@zs=xs \ \ P \ \ P" by (auto iff add: Cons_eq_append_conv) lemma map_of_rev_distinct[simp]: "distinct (map fst m) \ map_of (rev m) = map_of m" apply (induct m) apply simp apply simp apply (subst map_add_comm) apply force apply simp done \ \Tail-recursive, generalized @{const rev}. May also be used for tail-recursively getting a list with all elements of the two operands, if the order does not matter, e.g. when implementing sets by lists.\ fun revg where "revg [] b = b" | "revg (a#as) b = revg as (a#b)" lemma revg_fun[simp]: "revg a b = rev a @ b" by (induct a arbitrary: b) auto lemma rev_split_conv[simp]: "l \ [] \ rev (tl l) @ [hd l] = rev l" by (induct l) simp_all lemma rev_butlast_is_tl_rev: "rev (butlast l) = tl (rev l)" by (induct l) auto lemma hd_last_singletonI: "\xs \ []; hd xs = last xs; distinct xs\ \ xs = [hd xs]" by (induct xs rule: list_nonempty_induct) auto lemma last_filter: "\xs \ []; P (last xs)\ \ last (filter P xs) = last xs" by (induct xs rule: rev_nonempty_induct) simp_all (* As the following two rules are similar in nature to list_induct2', they are named accordingly. *) lemma rev_induct2' [case_names empty snocl snocr snoclr]: assumes empty: "P [] []" and snocl: "\x xs. P (xs@[x]) []" and snocr: "\y ys. P [] (ys@[y])" and snoclr: "\x xs y ys. P xs ys \ P (xs@[x]) (ys@[y])" shows "P xs ys" proof (induct xs arbitrary: ys rule: rev_induct) case Nil thus ?case using empty snocr by (cases ys rule: rev_exhaust) simp_all next case snoc thus ?case using snocl snoclr by (cases ys rule: rev_exhaust) simp_all qed lemma rev_nonempty_induct2' [case_names single snocl snocr snoclr, consumes 2]: assumes "xs \ []" "ys \ []" assumes single': "\x y. P [x] [y]" and snocl: "\x xs y. xs \ [] \ P (xs@[x]) [y]" and snocr: "\x y ys. ys \ [] \ P [x] (ys@[y])" and snoclr: "\x xs y ys. \P xs ys; xs \ []; ys\[]\ \ P (xs@[x]) (ys@[y])" shows "P xs ys" using assms(1,2) proof (induct xs arbitrary: ys rule: rev_nonempty_induct) case single then obtain z zs where "ys = zs@[z]" by (metis rev_exhaust) thus ?case using single' snocr by (cases "zs = []") simp_all next case (snoc x xs) then obtain z zs where zs: "ys = zs@[z]" by (metis rev_exhaust) thus ?case using snocl snoclr snoc by (cases "zs = []") simp_all qed subsubsection "Folding" text "Ugly lemma about foldl over associative operator with left and right neutral element" lemma foldl_A1_eq: "!!i. \ !! e. f n e = e; !! e. f e n = e; !!a b c. f a (f b c) = f (f a b) c \ \ foldl f i ww = f i (foldl f n ww)" proof (induct ww) case Nil thus ?case by simp next case (Cons a ww i) note IHP[simplified]=this have "foldl f i (a # ww) = foldl f (f i a) ww" by simp also from IHP have "\ = f (f i a) (foldl f n ww)" by blast also from IHP(4) have "\ = f i (f a (foldl f n ww))" by simp also from IHP(1)[OF IHP(2,3,4), where i=a] have "\ = f i (foldl f a ww)" by simp also from IHP(2)[of a] have "\ = f i (foldl f (f n a) ww)" by simp also have "\ = f i (foldl f n (a#ww))" by simp finally show ?case . qed lemmas foldl_conc_empty_eq = foldl_A1_eq[of "(@)" "[]", simplified] lemmas foldl_un_empty_eq = foldl_A1_eq[of "(\)" "{}", simplified, OF Un_assoc[symmetric]] lemma foldl_set: "foldl (\) {} l = \{x. x\set l}" apply (induct l) apply simp_all apply (subst foldl_un_empty_eq) apply auto done lemma (in monoid_mult) foldl_absorb1: "x*foldl (*) 1 zs = foldl (*) x zs" apply (rule sym) apply (rule foldl_A1_eq) apply (auto simp add: mult.assoc) done text \Towards an invariant rule for foldl\ lemma foldl_rule_aux: fixes I :: "'\ \ 'a list \ bool" assumes initial: "I \0 l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ (x#l2) \ \ I (f \ x) l2" shows "I (foldl f \0 l0) []" using initial step apply (induct l0 arbitrary: \0) apply auto done lemma foldl_rule_aux_P: fixes I :: "'\ \ 'a list \ bool" assumes initial: "I \0 l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ (x#l2) \ \ I (f \ x) l2" assumes final: "!!\. I \ [] \ P \" shows "P (foldl f \0 l0)" using foldl_rule_aux[of I \0 l0, OF initial, OF step] final by simp lemma foldl_rule: fixes I :: "'\ \ 'a list \ 'a list \ bool" assumes initial: "I \0 [] l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ l1 (x#l2) \ \ I (f \ x) (l1@[x]) l2" shows "I (foldl f \0 l0) l0 []" using initial step apply (rule_tac I="\\ lr. \ll. l0=ll@lr \ I \ ll lr" in foldl_rule_aux_P) apply auto done text \ Invariant rule for foldl. The invariant is parameterized with the state, the list of items that have already been processed and the list of items that still have to be processed. \ lemma foldl_rule_P: fixes I :: "'\ \ 'a list \ 'a list \ bool" \ \The invariant holds for the initial state, no items processed yet and all items to be processed:\ assumes initial: "I \0 [] l0" \ \The invariant remains valid if one item from the list is processed\ assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ l1 (x#l2) \ \ I (f \ x) (l1@[x]) l2" \ \The proposition follows from the invariant in the final state, i.e. all items processed and nothing to be processed\ assumes final: "!!\. I \ l0 [] \ P \" shows "P (foldl f \0 l0)" using foldl_rule[of I, OF initial step] by (simp add: final) text \Invariant reasoning over @{const foldl} for distinct lists. Invariant rule makes no assumptions about ordering.\ lemma distinct_foldl_invar: "\ distinct S; I (set S) \0; \x it \. \x \ it; it \ set S; I it \\ \ I (it - {x}) (f \ x) \ \ I {} (foldl f \0 S)" proof (induct S arbitrary: \0) case Nil thus ?case by auto next case (Cons x S) note [simp] = Cons.prems(1)[simplified] show ?case apply simp apply (rule Cons.hyps) proof - from Cons.prems(1) show "distinct S" by simp from Cons.prems(3)[of x "set (x#S)", simplified, OF Cons.prems(2)[simplified]] show "I (set S) (f \0 x)" . fix xx it \ assume A: "xx\it" "it \ set S" "I it \" show "I (it - {xx}) (f \ xx)" using A(2) apply (rule_tac Cons.prems(3)) apply (simp_all add: A(1,3)) apply blast done qed qed lemma foldl_length_aux: "foldl (\i x. Suc i) a l = a + length l" by (induct l arbitrary: a) auto lemmas foldl_length[simp] = foldl_length_aux[where a=0, simplified] lemma foldr_length_aux: "foldr (\x i. Suc i) l a = a + length l" by (induct l arbitrary: a rule: rev_induct) auto lemmas foldr_length[simp] = foldr_length_aux[where a=0, simplified] context comp_fun_commute begin lemma foldl_f_commute: "f a (foldl (\a b. f b a) b xs) = foldl (\a b. f b a) (f a b) xs" by(induct xs arbitrary: b)(simp_all add: fun_left_comm) lemma foldr_conv_foldl: "foldr f xs a = foldl (\a b. f b a) a xs" by(induct xs arbitrary: a)(simp_all add: foldl_f_commute) end lemma filter_conv_foldr: "filter P xs = foldr (\x xs. if P x then x # xs else xs) xs []" by(induct xs) simp_all lemma foldr_Cons: "foldr Cons xs [] = xs" by(induct xs) simp_all lemma foldr_snd_zip: "length xs \ length ys \ foldr (\(x, y). f y) (zip xs ys) b = foldr f ys b" proof(induct ys arbitrary: xs) case (Cons y ys) thus ?case by(cases xs) simp_all qed simp lemma foldl_snd_zip: "length xs \ length ys \ foldl (\b (x, y). f b y) b (zip xs ys) = foldl f b ys" proof(induct ys arbitrary: xs b) case (Cons y ys) thus ?case by(cases xs) simp_all qed simp lemma fst_foldl: "fst (foldl (\(a, b) x. (f a x, g a b x)) (a, b) xs) = foldl f a xs" by(induct xs arbitrary: a b) simp_all lemma foldl_foldl_conv_concat: "foldl (foldl f) a xs = foldl f a (concat xs)" by(induct xs arbitrary: a) simp_all lemma foldl_list_update: "n < length xs \ foldl f a (xs[n := x]) = foldl f (f (foldl f a (take n xs)) x) (drop (Suc n) xs)" by(simp add: upd_conv_take_nth_drop) lemma map_by_foldl: fixes l :: "'a list" and f :: "'a \ 'b" shows "foldl (\l x. l@[f x]) [] l = map f l" proof - { fix l' have "foldl (\l x. l@[f x]) l' l = l'@map f l" by (induct l arbitrary: l') auto } thus ?thesis by simp qed subsubsection \Sorting\ lemma sorted_in_between: assumes A: "0\i" "i x" "xk" and "kx" and "xk. i\k \ k l!k\x \ x x") case True from True Suc.hyps have "d = j - (i + 1)" by simp moreover from True have "i+1 < j" by (metis Suc.prems Suc_eq_plus1 Suc_lessI not_less) moreover from True have "0\i+1" by simp ultimately obtain k where "i+1\k" "k x" "xk" "k x" "xsorted l; l\[]\ \ hd l \ last l" by (metis eq_iff hd_Cons_tl last_in_set not_hd_in_tl sorted_wrt.simps(2)) lemma (in linorder) sorted_hd_min: "\xs \ []; sorted xs\ \ \x \ set xs. hd xs \ x" by (induct xs, auto) lemma sorted_append_bigger: "\sorted xs; \x \ set xs. x \ y\ \ sorted (xs @ [y])" proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) then have s: "sorted xs" by (cases xs) simp_all from Cons have a: "\x\set xs. x \ y" by simp from Cons(1)[OF s a] Cons(2-) show ?case by (cases xs) simp_all qed lemma sorted_filter': "sorted l \ sorted (filter P l)" using sorted_filter[where f=id, simplified] . subsubsection \Map\ (* List.thy has: declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] *) lemma map_eq_consE: "\map f ls = fa#fl; !!a l. \ ls=a#l; f a=fa; map f l = fl \ \ P\ \ P" by auto lemma map_fst_mk_snd[simp]: "map fst (map (\x. (x,k)) l) = l" by (induct l) auto lemma map_snd_mk_fst[simp]: "map snd (map (\x. (k,x)) l) = l" by (induct l) auto lemma map_fst_mk_fst[simp]: "map fst (map (\x. (k,x)) l) = replicate (length l) k" by (induct l) auto lemma map_snd_mk_snd[simp]: "map snd (map (\x. (x,k)) l) = replicate (length l) k" by (induct l) auto lemma map_zip1: "map (\x. (x,k)) l = zip l (replicate (length l) k)" by (induct l) auto lemma map_zip2: "map (\x. (k,x)) l = zip (replicate (length l) k) l" by (induct l) auto lemmas map_zip=map_zip1 map_zip2 (* TODO/FIXME: hope nobody changes nth to be underdefined! *) lemma map_eq_nth_eq: assumes A: "map f l = map f l'" shows "f (l!i) = f (l'!i)" proof - from A have "length l = length l'" by (metis length_map) thus ?thesis using A apply (induct arbitrary: i rule: list_induct2) apply simp apply (simp add: nth_def split: nat.split) done qed lemma map_upd_eq: "\i f (l!i) = f x\ \ map f (l[i:=x]) = map f l" by (metis list_update_beyond list_update_id map_update not_le_imp_less) lemma inj_map_inv_f [simp]: "inj f \ map (inv f) (map f l) = l" by (simp) lemma inj_on_map_the: "\D \ dom m; inj_on m D\ \ inj_on (the\m) D" apply (rule inj_onI) apply simp apply (case_tac "m x") apply (case_tac "m y") apply (auto intro: inj_onD) [1] apply (auto intro: inj_onD) [1] apply (case_tac "m y") apply (auto intro: inj_onD) [1] apply simp apply (rule inj_onD) apply assumption apply auto done lemma map_consI: "w=map f ww \ f a#w = map f (a#ww)" "w@l=map f ww@l \ f a#w@l = map f (a#ww)@l" by auto lemma restrict_map_subset_eq: fixes R shows "\m |` R = m'; R'\R\ \ m|` R' = m' |` R'" by (auto simp add: Int_absorb1) lemma restrict_map_self[simp]: "m |` dom m = m" apply (rule ext) apply (case_tac "m x") apply (auto simp add: restrict_map_def) done lemma restrict_map_UNIV[simp]: "f |` UNIV = f" by (auto simp add: restrict_map_def) lemma restrict_map_inv[simp]: "f |` (- dom f) = Map.empty" by (auto simp add: restrict_map_def intro: ext) lemma restrict_map_upd: "(f |` S)(k \ v) = f(k\v) |` (insert k S)" by (auto simp add: restrict_map_def intro: ext) lemma map_upd_eq_restrict: "m (x:=None) = m |` (-{x})" by (auto intro: ext) declare Map.finite_dom_map_of [simp, intro!] lemma dom_const'[simp]: "dom (\x. Some (f x)) = UNIV" by auto lemma restrict_map_eq : "((m |` A) k = None) \ (k \ dom m \ A)" "((m |` A) k = Some v) \ (m k = Some v \ k \ A)" unfolding restrict_map_def by (simp_all add: dom_def) definition "rel_of m P == {(k,v). m k = Some v \ P (k, v)}" lemma rel_of_empty[simp]: "rel_of Map.empty P = {}" by (auto simp add: rel_of_def) lemma remove1_tl: "xs \ [] \ remove1 (hd xs) xs = tl xs" by (cases xs) auto lemma set_oo_map_alt: "(set \\ map) f = (\l. f ` set l)" by auto subsubsection "Filter and Revert" primrec filter_rev_aux where "filter_rev_aux a P [] = a" | "filter_rev_aux a P (x#xs) = ( if P x then filter_rev_aux (x#a) P xs else filter_rev_aux a P xs)" lemma filter_rev_aux_alt: "filter_rev_aux a P l = filter P (rev l) @ a" by (induct l arbitrary: a) auto definition "filter_rev == filter_rev_aux []" lemma filter_rev_alt: "filter_rev P l = filter P (rev l)" unfolding filter_rev_def by (simp add: filter_rev_aux_alt) definition "remove_rev x == filter_rev (Not o (=) x)" lemma remove_rev_alt_def : "remove_rev x xs = (filter (\y. y \ x) (rev xs))" unfolding remove_rev_def apply (simp add: filter_rev_alt comp_def) by metis subsubsection "zip" declare zip_map_fst_snd[simp] lemma pair_list_split: "\ !!l1 l2. \ l = zip l1 l2; length l1=length l2; length l=length l2 \ \ P \ \ P" proof (induct l arbitrary: P) case Nil thus ?case by auto next case (Cons a l) from Cons.hyps obtain l1 l2 where IHAPP: "l=zip l1 l2" "length l1 = length l2" "length l=length l2" . obtain a1 a2 where [simp]: "a=(a1,a2)" by (cases a) auto from IHAPP have "a#l = zip (a1#l1) (a2#l2)" "length (a1#l1) = length (a2#l2)" "length (a#l) = length (a2#l2)" by (simp_all only:) (simp_all (no_asm_use)) with Cons.prems show ?case by blast qed lemma set_zip_cart: "x\set (zip l l') \ x\set l \ set l'" by (auto simp add: set_zip) lemma map_prod_fun_zip: "map (\(x, y). (f x, g y)) (zip xs ys) = zip (map f xs) (map g ys)" proof(induct xs arbitrary: ys) case Nil thus ?case by simp next case (Cons x xs) thus ?case by(cases ys) simp_all qed subsubsection \Generalized Zip\ text \Zip two lists element-wise, where the combination of two elements is specified by a function. Note that this function is underdefined for lists of different length.\ fun zipf :: "('a\'b\'c) \ 'a list \ 'b list \ 'c list" where "zipf f [] [] = []" | "zipf f (a#as) (b#bs) = f a b # zipf f as bs" lemma zipf_zip: "\length l1 = length l2\ \ zipf Pair l1 l2 = zip l1 l2" apply (induct l1 arbitrary: l2) apply auto apply (case_tac l2) apply auto done \ \All quantification over zipped lists\ fun list_all_zip where "list_all_zip P [] [] \ True" | "list_all_zip P (a#as) (b#bs) \ P a b \ list_all_zip P as bs" | "list_all_zip P _ _ \ False" lemma list_all_zip_alt: "list_all_zip P as bs \ length as = length bs \ (\iP as bs rule: list_all_zip.induct) apply auto apply (case_tac i) apply auto done lemma list_all_zip_map1: "list_all_zip P (List.map f as) bs \ list_all_zip (\a b. P (f a) b) as bs" apply (induct as arbitrary: bs) apply (case_tac bs) apply auto [2] apply (case_tac bs) apply auto [2] done lemma list_all_zip_map2: "list_all_zip P as (List.map f bs) \ list_all_zip (\a b. P a (f b)) as bs" apply (induct as arbitrary: bs) apply (case_tac bs) apply auto [2] apply (case_tac bs) apply auto [2] done declare list_all_zip_alt[mono] lemma lazI[intro?]: "\ length a = length b; !!i. i P (a!i) (b!i) \ \ list_all_zip P a b" by (auto simp add: list_all_zip_alt) lemma laz_conj[simp]: "list_all_zip (\x y. P x y \ Q x y) a b \ list_all_zip P a b \ list_all_zip Q a b" by (auto simp add: list_all_zip_alt) lemma laz_len: "list_all_zip P a b \ length a = length b" by (simp add: list_all_zip_alt) lemma laz_eq: "list_all_zip (=) a b \ a=b" apply (induct a arbitrary: b) apply (case_tac b) apply simp apply simp apply (case_tac b) apply simp apply simp done lemma laz_swap_ex: assumes A: "list_all_zip (\a b. \c. P a b c) A B" obtains C where "list_all_zip (\a c. \b. P a b c) A C" "list_all_zip (\b c. \a. P a b c) B C" proof - from A have [simp]: "length A = length B" and IC: "\ici. P (A!i) (B!i) ci" by (auto simp add: list_all_zip_alt) from obtain_list_from_elements[OF IC] obtain C where "length C = length B" "\ia b. P a) A B \ (length A = length B) \ (\a\set A. P a)" by (auto simp add: list_all_zip_alt set_conv_nth) lemma laz_weak_Pb[simp]: "list_all_zip (\a b. P b) A B \ (length A = length B) \ (\b\set B. P b)" by (force simp add: list_all_zip_alt set_conv_nth) subsubsection "Collecting Sets over Lists" definition "list_collect_set f l == \{ f a | a. a\set l }" lemma list_collect_set_simps[simp]: "list_collect_set f [] = {}" "list_collect_set f [a] = f a" "list_collect_set f (a#l) = f a \ list_collect_set f l" "list_collect_set f (l@l') = list_collect_set f l \ list_collect_set f l'" by (unfold list_collect_set_def) auto lemma list_collect_set_map_simps[simp]: "list_collect_set f (map x []) = {}" "list_collect_set f (map x [a]) = f (x a)" "list_collect_set f (map x (a#l)) = f (x a) \ list_collect_set f (map x l)" "list_collect_set f (map x (l@l')) = list_collect_set f (map x l) \ list_collect_set f (map x l')" by simp_all lemma list_collect_set_alt: "list_collect_set f l = \{ f (l!i) | i. i(set (map f l))" by (unfold list_collect_set_def) auto subsubsection \Sorted List with arbitrary Relations\ lemma (in linorder) sorted_wrt_rev_linord [simp] : "sorted_wrt (\) l \ sorted (rev l)" by (simp add: sorted_wrt_rev) lemma (in linorder) sorted_wrt_map_linord [simp] : "sorted_wrt (\(x::'a \ 'b) y. fst x \ fst y) l \ sorted (map fst l)" by (simp add: sorted_wrt_map) lemma (in linorder) sorted_wrt_map_rev_linord [simp] : "sorted_wrt (\(x::'a \ 'b) y. fst x \ fst y) l \ sorted (rev (map fst l))" by (induct l) (auto simp add: sorted_append) subsubsection \Take and Drop\ lemma take_update[simp]: "take n (l[i:=x]) = (take n l)[i:=x]" apply (induct l arbitrary: n i) apply (auto split: nat.split) apply (case_tac n) apply simp_all apply (case_tac n) apply simp_all done lemma take_update_last: "length list>n \ (take (Suc n) list) [n:=x] = take n list @ [x]" by (induct list arbitrary: n) (auto split: nat.split) lemma drop_upd_irrelevant: "m < n \ drop n (l[m:=x]) = drop n l" apply (induct n arbitrary: l m) apply simp apply (case_tac l) apply (auto split: nat.split) done lemma set_drop_conv: "set (drop n l) = { l!i | i. n\i \ i < length l }" (is "?L=?R") proof (intro equalityI subsetI) fix x assume "x\?L" then obtain i where L: "i = l!(n+i)" using L by simp finally show "x\?R" using L by auto next fix x assume "x\?R" then obtain i where L: "n\i" "i?L" by (auto simp add: in_set_conv_nth) qed lemma filter_upt_take_conv: "[i\[n..[n..set (drop n l) \ (\i. n\i \ i x = l!i)" apply (clarsimp simp: in_set_conv_nth) apply safe apply simp apply (metis le_add2 less_diff_conv add.commute) apply (rule_tac x="i-n" in exI) apply auto [] done lemma Union_take_drop_id: "\(set (drop n l)) \ \(set (take n l)) = \(set l)" by (metis Union_Un_distrib append_take_drop_id set_union_code sup_commute) lemma Un_set_drop_extend: "\j\Suc 0; j < length l\ \ l ! (j - Suc 0) \ \(set (drop j l)) = \(set (drop (j - Suc 0) l))" apply safe apply simp_all apply (metis diff_Suc_Suc diff_zero gr0_implies_Suc in_set_drop_conv_nth le_refl less_eq_Suc_le order.strict_iff_order) apply (metis Nat.diff_le_self set_drop_subset_set_drop subset_code(1)) by (metis diff_Suc_Suc gr0_implies_Suc in_set_drop_conv_nth less_eq_Suc_le order.strict_iff_order minus_nat.diff_0) lemma drop_take_drop_unsplit: "i\j \ drop i (take j l) @ drop j l = drop i l" proof - assume "i \ j" then obtain skf where "i + skf = j" by (metis le_iff_add) thus "drop i (take j l) @ drop j l = drop i l" by (metis append_take_drop_id diff_add_inverse drop_drop drop_take add.commute) qed lemma drop_last_conv[simp]: "l\[] \ drop (length l - Suc 0) l = [last l]" by (cases l rule: rev_cases) auto lemma take_butlast_conv[simp]: "take (length l - Suc 0) l = butlast l" by (cases l rule: rev_cases) auto lemma drop_takeWhile: assumes "i\length (takeWhile P l)" shows "drop i (takeWhile P l) = takeWhile P (drop i l)" using assms proof (induction l arbitrary: i) case Nil thus ?case by auto next case (Cons x l) thus ?case by (cases i) auto qed lemma less_length_takeWhile_conv: "i < length (takeWhile P l) \ (i (\j\i. P (l!j)))" apply safe subgoal using length_takeWhile_le less_le_trans by blast subgoal by (metis dual_order.strict_trans2 nth_mem set_takeWhileD takeWhile_nth) subgoal by (meson less_le_trans not_le_imp_less nth_length_takeWhile) done lemma eq_len_takeWhile_conv: "i=length (takeWhile P l) \ i\length l \ (\j (i \P (l!i))" apply safe subgoal using length_takeWhile_le less_le_trans by blast subgoal by (auto simp: less_length_takeWhile_conv) subgoal using nth_length_takeWhile by blast subgoal by (metis length_takeWhile_le nth_length_takeWhile order.order_iff_strict) subgoal by (metis dual_order.strict_trans2 leI less_length_takeWhile_conv linorder_neqE_nat nth_length_takeWhile) done subsubsection \Up-to\ lemma upt_eq_append_conv: "i\j \ [i.. (\k. i\k \ k\j \ [i.. [k..j" thus "\k\i. k \ j \ [i.. [k.. length l \ map (nth l) [M.. is1 = [l.. is2 = [Suc i.. l\i \ ii. i + ofs) [a..Last and butlast\ lemma butlast_upt: "butlast [m.. butlast [n..length l \ take (n - Suc 0) l = butlast (take n l)" by (simp add: butlast_take) lemma butlast_eq_cons_conv: "butlast l = x#xs \ (\xl. l=x#xs@[xl])" by (metis Cons_eq_appendI append_butlast_last_id butlast.simps butlast_snoc eq_Nil_appendI) lemma butlast_eq_consE: assumes "butlast l = x#xs" obtains xl where "l=x#xs@[xl]" using assms by (auto simp: butlast_eq_cons_conv) lemma drop_eq_ConsD: "drop n xs = x # xs' \ drop (Suc n) xs = xs'" by(induct xs arbitrary: n)(simp_all add: drop_Cons split: nat.split_asm) subsubsection \List Slices\ text \Based on Lars Hupel's code.\ definition slice :: "nat \ nat \ 'a list \ 'a list" where "slice from to list = take (to - from) (drop from list)" lemma slice_len[simp]: "\ from \ to; to \ length xs \ \ length (slice from to xs) = to - from" unfolding slice_def by simp lemma slice_head: "\ from < to; to \ length xs \ \ hd (slice from to xs) = xs ! from" unfolding slice_def proof - assume a1: "from < to" assume "to \ length xs" then have "\n. to - (to - n) \ length (take to xs)" by (metis (no_types) slice_def diff_le_self drop_take length_drop slice_len) then show "hd (take (to - from) (drop from xs)) = xs ! from" using a1 by (metis diff_diff_cancel drop_take hd_drop_conv_nth leI le_antisym less_or_eq_imp_le nth_take) qed lemma slice_eq_bounds_empty[simp]: "slice i i xs = []" unfolding slice_def by auto lemma slice_nth: "\ from < to; to \ length xs; i < to - from \ \ slice from to xs ! i = xs ! (from + i)" unfolding slice_def by (induction "to - from" arbitrary: "from" to i) simp+ lemma slice_prepend: "\ i \ k; k \ length xs \ \ let p = length ys in slice i k xs = slice (i + p) (k + p) (ys @ xs)" unfolding slice_def Let_def by force lemma slice_Nil[simp]: "slice begin end [] = []" unfolding slice_def by auto lemma slice_Cons: "slice begin end (x#xs) = (if begin=0 \ end>0 then x#slice begin (end-1) xs else slice (begin - 1) (end - 1) xs)" unfolding slice_def by (auto simp: take_Cons' drop_Cons') lemma slice_complete[simp]: "slice 0 (length xs) xs = xs" unfolding slice_def by simp subsubsection \Miscellaneous\ lemma length_compl_induct[case_names Nil Cons]: "\P []; !! e l . \!! ll . length ll <= length l \ P ll\ \ P (e#l)\ \ P l" apply(induct_tac l rule: length_induct) apply(case_tac "xs") apply(auto) done lemma in_set_list_format: "\ e\set l; !!l1 l2. l=l1@e#l2 \ P \ \ P" proof (induct l arbitrary: P) case Nil thus ?case by auto next case (Cons a l) show ?case proof (cases "a=e") case True with Cons show ?thesis by force next case False with Cons.prems(1) have "e\set l" by auto with Cons.hyps obtain l1 l2 where "l=l1@e#l2" by blast hence "a#l = (a#l1)@e#l2" by simp with Cons.prems(2) show P by blast qed qed lemma in_set_upd_cases: assumes "x\set (l[i:=y])" obtains "iset l" by (metis assms in_set_conv_nth length_list_update nth_list_update_eq nth_list_update_neq) lemma in_set_upd_eq_aux: assumes "iset (l[i:=y]) \ x=y \ (\y. x\set (l[i:=y]))" by (metis in_set_upd_cases assms list_update_overwrite set_update_memI) lemma in_set_upd_eq: assumes "iset (l[i:=y]) \ x=y \ (x\set l \ (\y. x\set (l[i:=y])))" by (metis in_set_upd_cases in_set_upd_eq_aux assms) text \Simultaneous induction over two lists, prepending an element to one of the lists in each step\ lemma list_2pre_induct[case_names base left right]: assumes BASE: "P [] []" and LEFT: "!!e w1' w2. P w1' w2 \ P (e#w1') w2" and RIGHT: "!!e w1 w2'. P w1 w2' \ P w1 (e#w2')" shows "P w1 w2" proof - { \ \The proof is done by induction over the sum of the lengths of the lists\ fix n have "!!w1 w2. \length w1 + length w2 = n; P [] []; !!e w1' w2. P w1' w2 \ P (e#w1') w2; !!e w1 w2'. P w1 w2' \ P w1 (e#w2') \ \ P w1 w2 " apply (induct n) apply simp apply (case_tac w1) apply auto apply (case_tac w2) apply auto done } from this[OF _ BASE LEFT RIGHT] show ?thesis by blast qed lemma list_decomp_1: "length l=1 \ \a. l=[a]" by (case_tac l, auto) lemma list_decomp_2: "length l=2 \ \a b. l=[a,b]" by (case_tac l, auto simp add: list_decomp_1) lemma list_rest_coinc: "\length s2 \ length s1; s1@r1 = s2@r2\ \ \r1p. r2=r1p@r1" by (metis append_eq_append_conv_if) lemma list_tail_coinc: "n1#r1 = n2#r2 \ n1=n2 & r1=r2" by (auto) lemma last_in_set[intro]: "\l\[]\ \ last l \ set l" by (induct l) auto lemma empty_append_eq_id[simp]: "(@) [] = (\x. x)" by auto lemma op_conc_empty_img_id[simp]: "((@) [] ` L) = L" by auto lemma distinct_match: "\ distinct (al@e#bl) \ \ (al@e#bl = al'@e#bl') \ (al=al' \ bl=bl')" proof (rule iffI, induct al arbitrary: al') case Nil thus ?case by (cases al') auto next case (Cons a al) note Cprems=Cons.prems note Chyps=Cons.hyps show ?case proof (cases al') case Nil with Cprems have False by auto thus ?thesis .. next case [simp]: (Cons a' all') with Cprems have [simp]: "a=a'" and P: "al@e#bl = all'@e#bl'" by auto from Cprems(1) have D: "distinct (al@e#bl)" by auto from Chyps[OF D P] have [simp]: "al=all'" "bl=bl'" by auto show ?thesis by simp qed qed simp lemma prop_match: "\ list_all P al; \P e; \P e'; list_all P bl \ \ (al@e#bl = al'@e'#bl') \ (al=al' \ e=e' \ bl=bl')" apply (rule iffI, induct al arbitrary: al') apply (case_tac al', fastforce, fastforce)+ done lemmas prop_matchD = rev_iffD1[OF _ prop_match[where P=P]] for P declare distinct_tl[simp] lemma list_se_match[simp]: "l1 \ [] \ l1@l2 = [a] \ l1 = [a] \ l2 = []" "l2 \ [] \ l1@l2 = [a] \ l1 = [] \ l2 = [a]" "l1 \ [] \ [a] = l1@l2 \ l1 = [a] \ l2 = []" "l2 \ [] \ [a] = l1@l2 \ l1 = [] \ l2 = [a]" apply (cases l1, simp_all) apply (cases l1, simp_all) apply (cases l1, auto) [] apply (cases l1, auto) [] done (* Placed here because it depends on xy_in_set_cases *) lemma distinct_map_eq: "\ distinct (List.map f l); f x = f y; x\set l; y\set l \ \ x=y" by (erule (2) xy_in_set_cases) auto lemma upt_append: assumes "iu'" assumes NP: "\i. u\i \ i \P i" shows "[i\[0..[0..[0..[0..[u..[u..[0.. P (l!i)" proof assume A: "P (l!i)" have "[0..i by (simp add: upt_append) also have "[i..i by (auto simp: upt_conv_Cons) finally have "[k\[0..[Suc i..P (l!i)\ by simp hence "j = last (i#[k\[Suc i.. \ i" proof - have "sorted (i#[k\[Suc i.. last ?l" by (rule sorted_hd_last) simp thus ?thesis by simp qed finally have "i\j" . thus False using \j by simp qed lemma all_set_conv_nth: "(\x\set l. P x) \ (\i *) lemma upt_0_eq_Nil_conv[simp]: "[0.. j=0" by auto lemma filter_eq_snocD: "filter P l = l'@[x] \ x\set l \ P x" proof - assume A: "filter P l = l'@[x]" hence "x\set (filter P l)" by simp thus ?thesis by simp qed lemma lists_image_witness: assumes A: "x\lists (f`Q)" obtains xo where "xo\lists Q" "x=map f xo" proof - have "\ x\lists (f`Q) \ \ \xo\lists Q. x=map f xo" proof (induct x) case Nil thus ?case by auto next case (Cons x xs) then obtain xos where "xos\lists Q" "xs=map f xos" by force moreover from Cons.prems have "x\f`Q" by auto then obtain xo where "xo\Q" "x=f xo" by auto ultimately show ?case by (rule_tac x="xo#xos" in bexI) auto qed thus ?thesis apply (simp_all add: A) apply (erule_tac bexE) apply (rule_tac that) apply assumption+ done qed lemma map_of_None_filterD: "map_of xs x = None \ map_of (filter P xs) x = None" by(induct xs) auto lemma map_of_concat: "map_of (concat xss) = foldr (\xs f. f ++ map_of xs) xss Map.empty" by(induct xss) simp_all lemma map_of_Some_split: "map_of xs k = Some v \ \ys zs. xs = ys @ (k, v) # zs \ map_of ys k = None" proof(induct xs) case (Cons x xs) obtain k' v' where x: "x = (k', v')" by(cases x) show ?case proof(cases "k' = k") case True with \map_of (x # xs) k = Some v\ x have "x # xs = [] @ (k, v) # xs" "map_of [] k = None" by simp_all thus ?thesis by blast next case False with \map_of (x # xs) k = Some v\ x have "map_of xs k = Some v" by simp from \map_of xs k = Some v \ \ys zs. xs = ys @ (k, v) # zs \ map_of ys k = None\[OF this] obtain ys zs where "xs = ys @ (k, v) # zs" "map_of ys k = None" by blast with False x have "x # xs = (x # ys) @ (k, v) # zs" "map_of (x # ys) k = None" by simp_all thus ?thesis by blast qed qed simp lemma map_add_find_left: "g k = None \ (f ++ g) k = f k" by(simp add: map_add_def) lemma map_add_left_None: "f k = None \ (f ++ g) k = g k" by(simp add: map_add_def split: option.split) lemma map_of_Some_filter_not_in: "\ map_of xs k = Some v; \ P (k, v); distinct (map fst xs) \ \ map_of (filter P xs) k = None" apply(induct xs) apply(auto) apply(auto simp add: map_of_eq_None_iff) done lemma distinct_map_fst_filterI: "distinct (map fst xs) \ distinct (map fst (filter P xs))" by(induct xs) auto lemma distinct_map_fstD: "\ distinct (map fst xs); (x, y) \ set xs; (x, z) \ set xs \ \ y = z" by(induct xs)(fastforce elim: notE rev_image_eqI)+ lemma concat_filter_neq_Nil: "concat [ys\xs. ys \ Nil] = concat xs" by(induct xs) simp_all lemma distinct_concat': "\distinct [ys\xs. ys \ Nil]; \ys. ys \ set xs \ distinct ys; \ys zs. \ys \ set xs; zs \ set xs; ys \ zs\ \ set ys \ set zs = {}\ \ distinct (concat xs)" by(erule distinct_concat[of "[ys\xs. ys \ Nil]", unfolded concat_filter_neq_Nil]) auto lemma distinct_idx: assumes "distinct (map f l)" assumes "im. n \ m \ m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" using assms proof(induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc x xs) show ?case proof(cases "P x") case [simp]: False from \n < length (filter P (xs @ [x]))\ have "n < length (filter P xs)" by simp hence "\m\n. m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" by(rule snoc) thus ?thesis by(auto simp add: nth_append) next case [simp]: True show ?thesis proof(cases "n = length (filter P xs)") case False with \n < length (filter P (xs @ [x]))\ have "n < length (filter P xs)" by simp moreover hence "\m\n. m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" by(rule snoc) ultimately show ?thesis by(auto simp add: nth_append) next case [simp]: True hence "filter P (xs @ [x]) ! n = (xs @ [x]) ! length xs" by simp moreover have "length xs < length (xs @ [x])" by simp moreover have "length xs \ n" by simp moreover have "filter P (take (length xs) (xs @ [x])) = take n (filter P (xs @ [x]))" by simp ultimately show ?thesis by blast qed qed qed lemma set_map_filter: "set (List.map_filter g xs) = {y. \x. x \ set xs \ g x = Some y}" by (induct xs) (auto simp add: List.map_filter_def set_eq_iff) subsection \Quicksort by Relation\ text \A functional implementation of quicksort on lists. It it similar to the one in Isabelle/HOL's example directory. However, it uses tail-recursion for append and arbitrary relations.\ fun partition_rev :: "('a \ bool) \ ('a list \ 'a list) \ 'a list \ ('a list \ 'a list)" where "partition_rev P (yes, no) [] = (yes, no)" | "partition_rev P (yes, no) (x # xs) = partition_rev P (if P x then (x # yes, no) else (yes, x # no)) xs" lemma partition_rev_filter_conv : "partition_rev P (yes, no) xs = (rev (filter P xs) @ yes, rev (filter (Not \ P) xs) @ no)" by (induct xs arbitrary: yes no) (simp_all) function quicksort_by_rel :: "('a \ 'a \ bool) \ 'a list \ 'a list \ 'a list" where "quicksort_by_rel R sl [] = sl" | "quicksort_by_rel R sl (x#xs) = (let (xs_s, xs_b) = partition_rev (\y. R y x) ([],[]) xs in quicksort_by_rel R (x # (quicksort_by_rel R sl xs_b)) xs_s)" by pat_completeness simp_all termination by (relation "measure (\(_, _, xs). length xs)") (simp_all add: partition_rev_filter_conv less_Suc_eq_le) lemma quicksort_by_rel_remove_acc : "quicksort_by_rel R sl xs = (quicksort_by_rel R [] xs @ sl)" proof (induct xs arbitrary: sl rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyp1a = ind_hyp[OF length_le(1), of "x # quicksort_by_rel R [] xs2"] note ind_hyp1b = ind_hyp[OF length_le(1), of "x # quicksort_by_rel R [] xs2 @ sl"] note ind_hyp2 = ind_hyp[OF length_le(2), of sl] show ?thesis by (simp add: ind_hyp1a ind_hyp1b ind_hyp2) qed qed lemma quicksort_by_rel_remove_acc_guared : "sl \ [] \ quicksort_by_rel R sl xs = (quicksort_by_rel R [] xs @ sl)" by (metis quicksort_by_rel_remove_acc) lemma quicksort_by_rel_permutes [simp]: "mset (quicksort_by_rel R sl xs) = mset (xs @ sl)" proof (induct xs arbitrary: sl rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have xs'_multi_eq : "mset xs' = mset xs1 + mset xs2" unfolding partition_rev_filter_conv by (simp add: mset_filter) from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyp[OF length_le(1)] ind_hyp[OF length_le(2)] thus ?thesis by (simp add: xs'_multi_eq union_assoc) qed qed lemma set_quicksort_by_rel [simp]: "set (quicksort_by_rel R sl xs) = set (xs @ sl)" unfolding set_mset_comp_mset [symmetric] o_apply by simp lemma sorted_wrt_quicksort_by_rel: fixes R:: "'x \ 'x \ bool" assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (quicksort_by_rel R [] xs)" proof (induct xs rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have xs1_props: "\y. y \ set xs1 \ (R y x)" and xs2_props: "\y. y \ set xs2 \ \(R y x)" unfolding partition_rev_filter_conv by simp_all from xs2_props lin have xs2_props': "\y. y \ set xs2 \ (R x y)" by blast from xs2_props' xs1_props trans_R have xs1_props': "\y1 y2. y1 \ set xs1 \ y2 \ set xs2 \ (R y1 y2)" by metis from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyps = ind_hyp[OF length_le(1)] ind_hyp[OF length_le(2)] thus ?thesis by (simp add: quicksort_by_rel_remove_acc_guared sorted_wrt_append Ball_def xs1_props xs2_props' xs1_props') qed qed lemma sorted_quicksort_by_rel: "sorted (quicksort_by_rel (\) [] xs)" by (rule sorted_wrt_quicksort_by_rel) auto lemma sort_quicksort_by_rel: "sort = quicksort_by_rel (\) []" apply (rule ext, rule properties_for_sort) apply(simp_all add: sorted_quicksort_by_rel) done lemma [code]: "quicksort = quicksort_by_rel (\) []" apply (subst sort_quicksort[symmetric]) by (rule sort_quicksort_by_rel) subsection \Mergesort by Relation\ text \A functional implementation of mergesort on lists. It it similar to the one in Isabelle/HOL's example directory. However, it uses tail-recursion for append and arbitrary relations.\ fun mergesort_by_rel_split :: "('a list \ 'a list) \ 'a list \ ('a list \ 'a list)" where "mergesort_by_rel_split (xs1, xs2) [] = (xs1, xs2)" | "mergesort_by_rel_split (xs1, xs2) [x] = (x # xs1, xs2)" | "mergesort_by_rel_split (xs1, xs2) (x1 # x2 # xs) = mergesort_by_rel_split (x1 # xs1, x2 # xs2) xs" lemma list_induct_first2 [consumes 0, case_names Nil Sing Cons2]: assumes "P []" "\x. P [x]" "\x1 x2 xs. P xs \ P (x1 # x2 #xs)" shows "P xs" proof (induct xs rule: length_induct) case (1 xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis using assms(1) by simp next case (Cons x1 xs') note xs_eq[simp] = this thus ?thesis proof (cases xs') case Nil thus ?thesis using assms(2) by simp next case (Cons x2 xs'') note xs'_eq[simp] = this show ?thesis by (simp add: ind_hyp assms(3)) qed qed qed lemma mergesort_by_rel_split_length : "length (fst (mergesort_by_rel_split (xs1, xs2) xs)) = length xs1 + (length xs div 2) + (length xs mod 2) \ length (snd (mergesort_by_rel_split (xs1, xs2) xs)) = length xs2 + (length xs div 2)" by (induct xs arbitrary: xs1 xs2 rule: list_induct_first2) (simp_all) lemma mset_mergesort_by_rel_split [simp]: "mset (fst (mergesort_by_rel_split (xs1, xs2) xs)) + mset (snd (mergesort_by_rel_split (xs1, xs2) xs)) = mset xs + mset xs1 + mset xs2" apply (induct xs arbitrary: xs1 xs2 rule: list_induct_first2) apply (simp_all add: ac_simps) done fun mergesort_by_rel_merge :: "('a \ 'a \ bool) \ 'a list \ 'a list \ 'a list" where "mergesort_by_rel_merge R (x#xs) (y#ys) = (if R x y then x # mergesort_by_rel_merge R xs (y#ys) else y # mergesort_by_rel_merge R (x#xs) ys)" | "mergesort_by_rel_merge R xs [] = xs" | "mergesort_by_rel_merge R [] ys = ys" declare mergesort_by_rel_merge.simps [simp del] lemma mergesort_by_rel_merge_simps[simp] : "mergesort_by_rel_merge R (x#xs) (y#ys) = (if R x y then x # mergesort_by_rel_merge R xs (y#ys) else y # mergesort_by_rel_merge R (x#xs) ys)" "mergesort_by_rel_merge R xs [] = xs" "mergesort_by_rel_merge R [] ys = ys" apply (simp_all add: mergesort_by_rel_merge.simps) apply (cases ys) apply (simp_all add: mergesort_by_rel_merge.simps) done lemma mergesort_by_rel_merge_induct [consumes 0, case_names Nil1 Nil2 Cons1 Cons2]: assumes "\xs::'a list. P xs []" "\ys::'b list. P [] ys" "\x xs y ys. R x y \ P xs (y # ys) \ P (x # xs) (y # ys)" "\x xs y ys. \(R x y) \ P (x # xs) ys \ P (x # xs) (y # ys)" shows "P xs ys" proof (induct xs arbitrary: ys) case Nil thus ?case using assms(2) by simp next case (Cons x xs) note P_xs = this show ?case proof (induct ys) case Nil thus ?case using assms(1) by simp next case (Cons y ys) note P_x_xs_ys = this show ?case using assms(3,4)[of x y xs ys] P_x_xs_ys P_xs by metis qed qed lemma mset_mergesort_by_rel_merge [simp]: "mset (mergesort_by_rel_merge R xs ys) = mset xs + mset ys" by (induct R xs ys rule: mergesort_by_rel_merge.induct) (simp_all add: ac_simps) lemma set_mergesort_by_rel_merge [simp]: "set (mergesort_by_rel_merge R xs ys) = set xs \ set ys" by (induct R xs ys rule: mergesort_by_rel_merge.induct) auto lemma sorted_wrt_mergesort_by_rel_merge [simp]: assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (mergesort_by_rel_merge R xs ys) \ sorted_wrt R xs \ sorted_wrt R ys" proof (induct xs ys rule: mergesort_by_rel_merge_induct[where R = R]) case Nil1 thus ?case by simp next case Nil2 thus ?case by simp next case (Cons1 x xs y ys) thus ?case by (simp add: Ball_def) (metis trans_R) next case (Cons2 x xs y ys) thus ?case apply (auto simp add: Ball_def) apply (metis lin) apply (metis lin trans_R) done qed function mergesort_by_rel :: "('a \ 'a \ bool) \ 'a list \ 'a list" where "mergesort_by_rel R xs = (if length xs < 2 then xs else (mergesort_by_rel_merge R (mergesort_by_rel R (fst (mergesort_by_rel_split ([], []) xs))) (mergesort_by_rel R (snd (mergesort_by_rel_split ([], []) xs)))))" by auto termination by (relation "measure (\(_, xs). length xs)") (simp_all add: mergesort_by_rel_split_length not_less minus_div_mult_eq_mod [symmetric]) declare mergesort_by_rel.simps [simp del] lemma mergesort_by_rel_simps [simp, code] : "mergesort_by_rel R [] = []" "mergesort_by_rel R [x] = [x]" "mergesort_by_rel R (x1 # x2 # xs) = (let (xs1, xs2) = (mergesort_by_rel_split ([x1], [x2]) xs) in mergesort_by_rel_merge R (mergesort_by_rel R xs1) (mergesort_by_rel R xs2))" apply (simp add: mergesort_by_rel.simps) apply (simp add: mergesort_by_rel.simps) apply (simp add: mergesort_by_rel.simps[of _ "x1 # x2 # xs"] split: prod.splits) done lemma mergesort_by_rel_permutes [simp]: "mset (mergesort_by_rel R xs) = mset xs" proof (induct xs rule: length_induct) case (1 xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x1 xs') note xs_eq[simp] = this show ?thesis proof (cases xs') case Nil thus ?thesis by simp next case (Cons x2 xs'') note xs'_eq[simp] = this have "length (fst (mergesort_by_rel_split ([], []) xs)) < length xs" "length (snd (mergesort_by_rel_split ([], []) xs)) < length xs" by (simp_all add: mergesort_by_rel_split_length) with ind_hyp show ?thesis unfolding mergesort_by_rel.simps[of _ xs] by (simp add: ac_simps) qed qed qed lemma set_mergesort_by_rel [simp]: "set (mergesort_by_rel R xs) = set xs" unfolding set_mset_comp_mset [symmetric] o_apply by simp lemma sorted_wrt_mergesort_by_rel: fixes R:: "'x \ 'x \ bool" assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (mergesort_by_rel R xs)" proof (induct xs rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this thus ?thesis proof (cases xs') case Nil thus ?thesis by simp next case (Cons x2 xs'') note xs'_eq[simp] = this have "length (fst (mergesort_by_rel_split ([], []) xs)) < length xs" "length (snd (mergesort_by_rel_split ([], []) xs)) < length xs" by (simp_all add: mergesort_by_rel_split_length) with ind_hyp show ?thesis unfolding mergesort_by_rel.simps[of _ xs] by (simp add: sorted_wrt_mergesort_by_rel_merge[OF lin trans_R]) qed qed qed lemma sorted_mergesort_by_rel: "sorted (mergesort_by_rel (\) xs)" by (rule sorted_wrt_mergesort_by_rel) auto lemma sort_mergesort_by_rel: "sort = mergesort_by_rel (\)" apply (rule ext, rule properties_for_sort) apply(simp_all add: sorted_mergesort_by_rel) done definition "mergesort = mergesort_by_rel (\)" lemma sort_mergesort: "sort = mergesort" unfolding mergesort_def by (rule sort_mergesort_by_rel) subsubsection \Mergesort with Remdup\ term merge fun merge :: "'a::{linorder} list \ 'a list \ 'a list" where "merge [] l2 = l2" | "merge l1 [] = l1" | "merge (x1 # l1) (x2 # l2) = (if (x1 < x2) then x1 # (merge l1 (x2 # l2)) else (if (x1 = x2) then x1 # (merge l1 l2) else x2 # (merge (x1 # l1) l2)))" lemma merge_correct : assumes l1_OK: "distinct l1 \ sorted l1" assumes l2_OK: "distinct l2 \ sorted l2" shows "distinct (merge l1 l2) \ sorted (merge l1 l2) \ set (merge l1 l2) = set l1 \ set l2" using assms proof (induct l1 arbitrary: l2) case Nil thus ?case by simp next case (Cons x1 l1 l2) note x1_l1_props = Cons(2) note l2_props = Cons(3) from x1_l1_props have l1_props: "distinct l1 \ sorted l1" and x1_nin_l1: "x1 \ set l1" and x1_le: "\x. x \ set l1 \ x1 \ x" by (simp_all add: Ball_def) note ind_hyp_l1 = Cons(1)[OF l1_props] show ?case using l2_props proof (induct l2) case Nil with x1_l1_props show ?case by simp next case (Cons x2 l2) note x2_l2_props = Cons(2) from x2_l2_props have l2_props: "distinct l2 \ sorted l2" and x2_nin_l2: "x2 \ set l2" and x2_le: "\x. x \ set l2 \ x2 \ x" by (simp_all add: Ball_def) note ind_hyp_l2 = Cons(1)[OF l2_props] show ?case proof (cases "x1 < x2") case True note x1_less_x2 = this from ind_hyp_l1[OF x2_l2_props] x1_less_x2 x1_nin_l1 x1_le x2_le show ?thesis apply (auto simp add: Ball_def) apply (metis linorder_not_le) apply (metis linorder_not_less xt1(6) xt1(9)) done next case False note x2_le_x1 = this show ?thesis proof (cases "x1 = x2") case True note x1_eq_x2 = this from ind_hyp_l1[OF l2_props] x1_le x2_le x2_nin_l2 x1_eq_x2 x1_nin_l1 show ?thesis by (simp add: x1_eq_x2 Ball_def) next case False note x1_neq_x2 = this with x2_le_x1 have x2_less_x1 : "x2 < x1" by auto from ind_hyp_l2 x2_le_x1 x1_neq_x2 x2_le x2_nin_l2 x1_le show ?thesis apply (simp add: x2_less_x1 Ball_def) apply (metis linorder_not_le x2_less_x1 xt1(7)) done qed qed qed qed function merge_list :: "'a::{linorder} list list \ 'a list list \ 'a list" where "merge_list [] [] = []" | "merge_list [] [l] = l" | "merge_list (la # acc2) [] = merge_list [] (la # acc2)" | "merge_list (la # acc2) [l] = merge_list [] (l # la # acc2)" | "merge_list acc2 (l1 # l2 # ls) = merge_list ((merge l1 l2) # acc2) ls" by pat_completeness simp_all termination by (relation "measure (\(acc, ls). 3 * length acc + 2 * length ls)") (simp_all) lemma merge_list_correct : assumes ls_OK: "\l. l \ set ls \ distinct l \ sorted l" assumes as_OK: "\l. l \ set as \ distinct l \ sorted l" shows "distinct (merge_list as ls) \ sorted (merge_list as ls) \ set (merge_list as ls) = set (concat (as @ ls))" using assms proof (induct as ls rule: merge_list.induct) case 1 thus ?case by simp next case 2 thus ?case by simp next case 3 thus ?case by simp next case 4 thus ?case by auto next case (5 acc l1 l2 ls) note ind_hyp = 5(1) note l12_l_OK = 5(2) note acc_OK = 5(3) from l12_l_OK acc_OK merge_correct[of l1 l2] have set_merge_eq: "set (merge l1 l2) = set l1 \ set l2" by auto from l12_l_OK acc_OK merge_correct[of l1 l2] have "distinct (merge_list (merge l1 l2 # acc) ls) \ sorted (merge_list (merge l1 l2 # acc) ls) \ set (merge_list (merge l1 l2 # acc) ls) = set (concat ((merge l1 l2 # acc) @ ls))" by (rule_tac ind_hyp) auto with set_merge_eq show ?case by auto qed definition mergesort_remdups where "mergesort_remdups xs = merge_list [] (map (\x. [x]) xs)" lemma mergesort_remdups_correct : "distinct (mergesort_remdups l) \ sorted (mergesort_remdups l) \ (set (mergesort_remdups l) = set l)" proof - let ?l' = "map (\x. [x]) l" { fix xs assume "xs \ set ?l'" then obtain x where xs_eq: "xs = [x]" by auto hence "distinct xs \ sorted xs" by simp } note l'_OK = this from merge_list_correct[of "?l'" "[]", OF l'_OK] show ?thesis unfolding mergesort_remdups_def by simp qed (* TODO: Move *) lemma ex1_eqI: "\\!x. P x; P a; P b\ \ a=b" by blast lemma remdup_sort_mergesort_remdups: "remdups o sort = mergesort_remdups" (is "?lhs=?rhs") proof fix l have "set (?lhs l) = set l" and "sorted (?lhs l)" and "distinct (?lhs l)" by simp_all moreover note mergesort_remdups_correct ultimately show "?lhs l = ?rhs l" by (auto intro!: ex1_eqI[OF finite_sorted_distinct_unique[OF finite_set]]) qed subsection \Native Integers\ lemma int_of_integer_less_iff: "int_of_integer x < int_of_integer y \ x0 \ y\0 \ nat_of_integer x < nat_of_integer y \ xn' < n. \ P n') \ P (n::nat)" shows "\n' \ n. P n'" proof (rule classical) assume contra: "\ (\n'\n. P n')" hence "\n' < n. \ P n'" by auto hence "P n" by (rule hyp) thus "\n'\n. P n'" by auto qed subsubsection \Induction on nat\ lemma nat_compl_induct[case_names 0 Suc]: "\P 0; \n . \nn. nn \ n \ P nn \ P (Suc n)\ \ P n" apply(induct_tac n rule: nat_less_induct) apply(case_tac n) apply(auto) done lemma nat_compl_induct'[case_names 0 Suc]: "\P 0; !! n . \!! nn . nn \ n \ P nn\ \ P (Suc n)\ \ P n" apply(induct_tac n rule: nat_less_induct) apply(case_tac n) apply(auto) done lemma nz_le_conv_less: "0 k \ m \ k - Suc 0 < m" by auto lemma min_Suc_gt[simp]: "a min (Suc a) b = Suc a" "a min b (Suc a) = Suc a" by auto subsection \Integer\ text \Some setup from \int\ transferred to \integer\\ lemma atLeastLessThanPlusOne_atLeastAtMost_integer: "{l.. u \ {(0::integer).. u") apply (subst image_atLeastZeroLessThan_integer, assumption) apply (rule finite_imageI) apply auto done lemma finite_atLeastLessThan_integer [iff]: "finite {l..Functions of type @{typ "bool\bool"}\ lemma boolfun_cases_helper: "g=(\x. False) | g=(\x. x) | g=(\x. True) | g= (\x. \x)" proof - { assume "g False" "g True" hence "g = (\x. True)" by (rule_tac ext, case_tac x, auto) } moreover { assume "g False" "\g True" hence "g = (\x. \x)" by (rule_tac ext, case_tac x, auto) } moreover { assume "\g False" "g True" hence "g = (\x. x)" by (rule_tac ext, case_tac x, auto) } moreover { assume "\g False" "\g True" hence "g = (\x. False)" by (rule_tac ext, case_tac x, auto) } ultimately show ?thesis by fast qed lemma boolfun_cases[case_names False Id True Neg]: "\g=(\x. False) \ P; g=(\x. x) \ P; g=(\x. True) \ P; g=(\x. \x) \ P\ \ P" proof - note boolfun_cases_helper[of g] moreover assume "g=(\x. False) \ P" "g=(\x. x) \ P" "g=(\x. True) \ P" "g=(\x. \x) \ P" ultimately show ?thesis by fast qed subsection \Definite and indefinite description\ text "Combined definite and indefinite description for binary predicate" lemma some_theI: assumes EX: "\a b . P a b" and BUN: "!! b1 b2 . \\a . P a b1; \a . P a b2\ \ b1=b2" shows "P (SOME a . \b . P a b) (THE b . \a . P a b)" proof - from EX have "\b. P (SOME a. \b. P a b) b" by (rule someI_ex) moreover from EX have "\b. \a. P a b" by blast with BUN theI'[of "\b. \a. P a b"] have "\a. P a (THE b. \a. P a b)" by (unfold Ex1_def, blast) moreover note BUN ultimately show ?thesis by (fast) qed lemma some_insert_self[simp]: "S\{} \ insert (SOME x. x\S) S = S" by (auto intro: someI) lemma some_elem[simp]: "S\{} \ (SOME x. x\S) \ S" by (auto intro: someI) subsubsection\Hilbert Choice with option\ definition Eps_Opt where "Eps_Opt P = (if (\x. P x) then Some (SOME x. P x) else None)" lemma some_opt_eq_trivial[simp] : "Eps_Opt (\y. y = x) = Some x" unfolding Eps_Opt_def by simp lemma some_opt_sym_eq_trivial[simp] : "Eps_Opt ((=) x) = Some x" unfolding Eps_Opt_def by simp lemma some_opt_false_trivial[simp] : "Eps_Opt (\_. False) = None" unfolding Eps_Opt_def by simp lemma Eps_Opt_eq_None[simp] : "Eps_Opt P = None \ \(Ex P)" unfolding Eps_Opt_def by simp lemma Eps_Opt_eq_Some_implies : "Eps_Opt P = Some x \ P x" unfolding Eps_Opt_def by (metis option.inject option.simps(2) someI_ex) lemma Eps_Opt_eq_Some : assumes P_prop: "\x'. P x \ P x' \ x' = x" shows "Eps_Opt P = Some x \ P x" using P_prop unfolding Eps_Opt_def by (metis option.inject option.simps(2) someI_ex) subsection \Product Type\ lemma nested_case_prod_simp: "(\(a,b) c. f a b c) x y = (case_prod (\a b. f a b y) x)" by (auto split: prod.split) lemma fn_fst_conv: "(\x. (f (fst x))) = (\(a,_). f a)" by auto lemma fn_snd_conv: "(\x. (f (snd x))) = (\(_,b). f b)" by auto fun pairself where "pairself f (a,b) = (f a, f b)" lemma pairself_image_eq[simp]: "pairself f ` {(a,b). P a b} = {(f a, f b)| a b. P a b}" by force lemma pairself_image_cart[simp]: "pairself f ` (A\B) = f`A \ f`B" by (auto simp: image_def) lemma in_prod_fst_sndI: "fst x \ A \ snd x \ B \ x\A\B" by (cases x) auto lemma inj_Pair[simp]: "inj_on (\x. (x,c x)) S" "inj_on (\x. (c x,x)) S" by (auto intro!: inj_onI) declare Product_Type.swap_inj_on[simp] lemma img_fst [intro]: assumes "(a,b) \ S" shows "a \ fst ` S" by (rule image_eqI[OF _ assms]) simp lemma img_snd [intro]: assumes "(a,b) \ S" shows "b \ snd ` S" by (rule image_eqI[OF _ assms]) simp lemma range_prod: "range f \ (range (fst \ f)) \ (range (snd \ f))" proof fix y assume "y \ range f" then obtain x where y: "y = f x" by auto hence "y = (fst(f x), snd(f x))" by simp thus "y \ (range (fst \ f)) \ (range (snd \ f))" by (fastforce simp add: image_def) qed lemma finite_range_prod: assumes fst: "finite (range (fst \ f))" and snd: "finite (range (snd \ f))" shows "finite (range f)" proof - from fst snd have "finite (range (fst \ f) \ range (snd \ f))" by (rule finite_SigmaI) thus ?thesis by (rule finite_subset[OF range_prod]) qed lemma fstE: "x = (a,b) \ P (fst x) \ P a" by (metis fst_conv) lemma sndE: "x = (a,b) \ P (snd x) \ P b" by (metis snd_conv) subsubsection \Uncurrying\ (* TODO: Move to HOL/Product_Type? Lars H: "It's equal to case_prod, should use an abbreviation"*) definition uncurry :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where "uncurry f \ \(a,b). f a b" lemma uncurry_apply[simp]: "uncurry f (a,b) = f a b" unfolding uncurry_def by simp lemma curry_uncurry_id[simp]: "curry (uncurry f) = f" unfolding uncurry_def by simp lemma uncurry_curry_id[simp]: "uncurry (curry f) = f" unfolding uncurry_def by simp lemma do_curry: "f (a,b) = curry f a b" by simp lemma do_uncurry: "f a b = uncurry f (a,b)" by simp subsection \Sum Type\ lemma map_sum_Inr_conv: "map_sum fl fr s = Inr y \ (\x. s=Inr x \ y = fr x)" by (cases s) auto lemma map_sum_Inl_conv: "map_sum fl fr s = Inl y \ (\x. s=Inl x \ y = fl x)" by (cases s) auto subsection \Directed Graphs and Relations\ subsubsection "Reflexive-Transitive Closure" lemma r_le_rtrancl[simp]: "S\S\<^sup>*" by auto lemma rtrancl_mono_rightI: "S\S' \ S\S'\<^sup>*" by auto lemma trancl_sub: "R \ R\<^sup>+" by auto lemma trancl_single[simp]: "{(a,b)}\<^sup>+ = {(a,b)}" by (auto simp: trancl_insert) text \Pick first non-reflexive step\ lemma converse_rtranclE'[consumes 1, case_names base step]: assumes "(u,v)\R\<^sup>*" obtains "u=v" | vh where "u\vh" and "(u,vh)\R" and "(vh,v)\R\<^sup>*" using assms apply (induct rule: converse_rtrancl_induct) apply auto [] apply (case_tac "y=z") apply auto done lemma in_rtrancl_insert: "x\R\<^sup>* \ x\(insert r R)\<^sup>*" by (metis in_mono rtrancl_mono subset_insertI) lemma rtrancl_apply_insert: "R\<^sup>*``(insert x S) = insert x (R\<^sup>*``(S\R``{x}))" apply (auto) apply (erule converse_rtranclE) apply auto [2] apply (erule converse_rtranclE) apply (auto intro: converse_rtrancl_into_rtrancl) [2] done text \A path in a graph either does not use nodes from S at all, or it has a prefix leading to a node in S and a suffix that does not use nodes in S\ lemma rtrancl_last_visit[cases set, case_names no_visit last_visit_point]: shows "\ (q,q')\R\<^sup>*; (q,q')\(R-UNIV\S)\<^sup>* \ P; !!qt. \ qt\S; (q,qt)\R\<^sup>+; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" proof (induct rule: converse_rtrancl_induct[case_names refl step]) case refl thus ?case by auto next case (step q qh) show P proof (rule step.hyps(3)) assume A: "(qh,q')\(R-UNIV\S)\<^sup>*" show P proof (cases "qh\S") case False with step.hyps(1) A have "(q,q')\(R-UNIV\S)\<^sup>*" by (auto intro: converse_rtrancl_into_rtrancl) with step.prems(1) show P . next case True from step.hyps(1) have "(q,qh)\R\<^sup>+" by auto with step.prems(2) True A show P by blast qed next fix qt assume A: "qt\S" "(qh,qt)\R\<^sup>+" "(qt,q')\(R-UNIV\S)\<^sup>*" with step.hyps(1) have "(q,qt)\R\<^sup>+" by auto with step.prems(2) A(1,3) show P by blast qed qed text \Less general version of \rtrancl_last_visit\, but there's a short automatic proof\ lemma rtrancl_last_visit': "\ (q,q')\R\<^sup>*; (q,q')\(R-UNIV\S)\<^sup>* \ P; !!qt. \ qt\S; (q,qt)\R\<^sup>*; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" by (induct rule: converse_rtrancl_induct) (auto intro: converse_rtrancl_into_rtrancl) lemma rtrancl_last_visit_node: assumes "(s,s')\R\<^sup>*" shows "s\sh \ (s,s')\(R \ (UNIV \ (-{sh})))\<^sup>* \ (s,sh)\R\<^sup>* \ (sh,s')\(R \ (UNIV \ (-{sh})))\<^sup>*" using assms proof (induct rule: converse_rtrancl_induct) case base thus ?case by auto next case (step s st) moreover { assume P: "(st,s')\ (R \ UNIV \ - {sh})\<^sup>*" { assume "st=sh" with step have ?case by auto } moreover { assume "st\sh" with \(s,st)\R\ have "(s,st)\(R \ UNIV \ - {sh})\<^sup>*" by auto also note P finally have ?case by blast } ultimately have ?case by blast } moreover { assume P: "(st, sh) \ R\<^sup>* \ (sh, s') \ (R \ UNIV \ - {sh})\<^sup>*" with step(1) have ?case by (auto dest: converse_rtrancl_into_rtrancl) } ultimately show ?case by blast qed text \Find last point where a path touches a set\ lemma rtrancl_last_touch: "\ (q,q')\R\<^sup>*; q\S; !!qt. \ qt\S; (q,qt)\R\<^sup>*; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" by (erule rtrancl_last_visit') auto text \A path either goes over edge once, or not at all\ lemma trancl_over_edgeE: assumes "(u,w)\(insert (v1,v2) E)\<^sup>+" obtains "(u,w)\E\<^sup>+" | "(u,v1)\E\<^sup>*" and "(v2,w)\E\<^sup>*" using assms proof induct case (base z) thus ?thesis by (metis insertE prod.inject r_into_trancl' rtrancl_eq_or_trancl) next case (step y z) thus ?thesis by (metis (opaque_lifting, no_types) Pair_inject insertE rtrancl.simps trancl.simps trancl_into_rtrancl) qed lemma rtrancl_image_advance: "\q\R\<^sup>* `` Q0; (q,x)\R\ \ x\R\<^sup>* `` Q0" by (auto intro: rtrancl_into_rtrancl) lemma trancl_image_by_rtrancl: "(E\<^sup>+)``Vi \ Vi = (E\<^sup>*)``Vi" by (metis Image_Id Un_Image rtrancl_trancl_reflcl) lemma reachable_mono: "\R\R'; X\X'\ \ R\<^sup>*``X \ R'\<^sup>*``X'" by (metis Image_mono rtrancl_mono) lemma finite_reachable_advance: "\ finite (E\<^sup>*``{v0}); (v0,v)\E\<^sup>* \ \ finite (E\<^sup>*``{v})" by (erule finite_subset[rotated]) auto lemma rtrancl_mono_mp: "U\V \ x\U\<^sup>* \ x\V\<^sup>*" by (metis in_mono rtrancl_mono) lemma trancl_mono_mp: "U\V \ x\U\<^sup>+ \ x\V\<^sup>+" by (metis trancl_mono) lemma rtrancl_mapI: "(a,b)\E\<^sup>* \ (f a, f b)\(pairself f `E)\<^sup>*" apply (induction rule: rtrancl_induct) apply (force intro: rtrancl.intros)+ done lemma rtrancl_image_advance_rtrancl: assumes "q \ R\<^sup>*``Q0" assumes "(q,x) \ R\<^sup>*" shows "x \ R\<^sup>*``Q0" using assms by (metis rtrancl_idemp rtrancl_image_advance) lemma nth_step_trancl: "\n m. \ \ n. n < length xs - 1 \ (xs ! Suc n, xs ! n) \ R \ \ n < length xs \ m < n \ (xs ! n, xs ! m) \ R\<^sup>+" proof (induction xs) case (Cons x xs) hence "\n. n < length xs - 1 \ (xs ! Suc n, xs ! n) \ R" apply clarsimp by (metis One_nat_def diff_Suc_eq_diff_pred nth_Cons_Suc zero_less_diff) note IH = this[THEN Cons.IH] from Cons obtain n' where n': "Suc n' = n" by (cases n) blast+ show ?case proof (cases m) case "0" with Cons have "xs \ []" by auto with "0" Cons.prems(1)[of m] have "(xs ! 0, x) \ R" by simp moreover from IH[where m = 0] have "\n. n < length xs \ n > 0 \ (xs ! n, xs ! 0) \ R\<^sup>+" by simp ultimately have "\n. n < length xs \ (xs ! n, x) \ R\<^sup>+" by (metis trancl_into_trancl gr0I r_into_trancl') with Cons "0" show ?thesis by auto next case (Suc m') with Cons.prems n' have "n' < length xs" "m' < n'" by auto with IH have "(xs ! n', xs ! m') \ R\<^sup>+" by simp with Suc n' show ?thesis by auto qed qed simp lemma Image_empty_trancl_Image_empty: "R `` {v} = {} \ R\<^sup>+ `` {v} = {}" unfolding Image_def by (auto elim: converse_tranclE) lemma Image_empty_rtrancl_Image_id: "R `` {v} = {} \ R\<^sup>* `` {v} = {v}" unfolding Image_def by (auto elim: converse_rtranclE) lemma trans_rtrancl_eq_reflcl: "trans A \ A^* = A^=" by (simp add: rtrancl_trancl_reflcl) lemma refl_on_reflcl_Image: "refl_on B A \ C \ B \ A^= `` C = A `` C" by (auto simp add: Image_def dest: refl_onD) lemma Image_absorb_rtrancl: "\ trans A; refl_on B A; C \ B \ \ A^* `` C = A `` C" by (simp add: trans_rtrancl_eq_reflcl refl_on_reflcl_Image) lemma trancl_Image_unfold_left: "E\<^sup>+``S = E\<^sup>*``E``S" by (auto simp: trancl_unfold_left) lemma trancl_Image_unfold_right: "E\<^sup>+``S = E``E\<^sup>*``S" by (auto simp: trancl_unfold_right) lemma trancl_Image_advance_ss: "(u,v)\E \ E\<^sup>+``{v} \ E\<^sup>+``{u}" by auto lemma rtrancl_Image_advance_ss: "(u,v)\E \ E\<^sup>*``{v} \ E\<^sup>*``{u}" by auto (* FIXME: nicer name *) lemma trancl_union_outside: assumes "(v,w) \ (E\U)\<^sup>+" and "(v,w) \ E\<^sup>+" shows "\x y. (v,x) \ (E\U)\<^sup>* \ (x,y) \ U \ (y,w) \ (E\U)\<^sup>*" using assms proof (induction) case base thus ?case by auto next case (step w x) show ?case proof (cases "(v,w)\E\<^sup>+") case True from step have "(v,w)\(E\U)\<^sup>*" by simp moreover from True step have "(w,x) \ U" by (metis Un_iff trancl.simps) moreover have "(x,x) \ (E\U)\<^sup>*" by simp ultimately show ?thesis by blast next case False with step.IH obtain a b where "(v,a) \ (E\U)\<^sup>*" "(a,b) \ U" "(b,w) \ (E\U)\<^sup>*" by blast moreover with step have "(b,x) \ (E\U)\<^sup>*" by (metis rtrancl_into_rtrancl) ultimately show ?thesis by blast qed qed lemma trancl_restrict_reachable: assumes "(u,v) \ E\<^sup>+" assumes "E``S \ S" assumes "u\S" shows "(u,v) \ (E\S\S)\<^sup>+" using assms by (induction rule: converse_trancl_induct) (auto intro: trancl_into_trancl2) lemma rtrancl_image_unfold_right: "E``E\<^sup>*``V \ E\<^sup>*``V" by (auto intro: rtrancl_into_rtrancl) lemma trancl_Image_in_Range: "R\<^sup>+ `` V \ Range R" by (auto elim: trancl.induct) lemma rtrancl_Image_in_Field: "R\<^sup>* `` V \ Field R \ V" proof - from trancl_Image_in_Range have "R\<^sup>+ `` V \ Field R" unfolding Field_def by fast hence "R\<^sup>+ `` V \ V \ Field R \ V" by blast with trancl_image_by_rtrancl show ?thesis by metis qed lemma rtrancl_sub_insert_rtrancl: "R\<^sup>* \ (insert x R)\<^sup>*" by (auto elim: rtrancl.induct rtrancl_into_rtrancl) lemma trancl_sub_insert_trancl: "R\<^sup>+ \ (insert x R)\<^sup>+" by (auto elim: trancl.induct trancl_into_trancl) lemma Restr_rtrancl_mono: "(v,w) \ (Restr E U)\<^sup>* \ (v,w) \ E\<^sup>*" by (metis inf_le1 rtrancl_mono subsetCE) lemma Restr_trancl_mono: "(v,w) \ (Restr E U)\<^sup>+ \ (v,w) \ E\<^sup>+" by (metis inf_le1 trancl_mono) subsubsection "Converse Relation" lemmas converse_add_simps = converse_Times trancl_converse[symmetric] converse_Un converse_Int lemma dom_ran_disj_comp[simp]: "Domain R \ Range R = {} \ R O R = {}" by auto lemma below_Id_inv[simp]: "R\\Id \ R\Id" by (auto) subsubsection "Cyclicity" lemma cyclicE: "\\acyclic g; !!x. (x,x)\g\<^sup>+ \ P\ \ P" by (unfold acyclic_def) blast lemma acyclic_insert_cyclic: "\acyclic g; \acyclic (insert (x,y) g)\ \ (y,x)\g\<^sup>*" by (unfold acyclic_def) (auto simp add: trancl_insert) text \ This lemma makes a case distinction about a path in a graph where a couple of edges with the same endpoint have been inserted: If there is a path from a to b, then there's such a path in the original graph, or there's a path that uses an inserted edge only once. Originally, this lemma was used to reason about the graph of an updated acquisition history. Any path in this graph is either already contained in the original graph, or passes via an inserted edge. Because all the inserted edges point to the same target node, in the second case, the path can be short-circuited to use exactly one inserted edge. \ lemma trancl_multi_insert[cases set, case_names orig via]: "\ (a,b)\(r\X\{m})\<^sup>+; (a,b)\r\<^sup>+ \ P; !!x. \ x\X; (a,x)\r\<^sup>*; (m,b)\r\<^sup>* \ \ P \ \ P" proof (induct arbitrary: P rule: trancl_induct) case (base b) thus ?case by auto next case (step b c) show ?case proof (rule step.hyps(3)) assume A: "(a,b)\r\<^sup>+" note step.hyps(2) moreover { assume "(b,c)\r" with A have "(a,c)\r\<^sup>+" by auto with step.prems have P by blast } moreover { assume "b\X" "c=m" with A have P by (rule_tac step.prems(2)) simp+ } ultimately show P by auto next fix x assume A: "x \ X" "(a, x) \ r\<^sup>*" "(m, b) \ r\<^sup>*" note step.hyps(2) moreover { assume "(b,c)\r" with A(3) have "(m,c)\r\<^sup>*" by auto with step.prems(2)[OF A(1,2)] have P by blast } moreover { assume "b\X" "c=m" with A have P by (rule_tac step.prems(2)) simp+ } ultimately show P by auto qed qed text \ Version of @{thm [source] trancl_multi_insert} for inserted edges with the same startpoint. \ lemma trancl_multi_insert2[cases set, case_names orig via]: "\(a,b)\(r\{m}\X)\<^sup>+; (a,b)\r\<^sup>+ \ P; !!x. \ x\X; (a,m)\r\<^sup>*; (x,b)\r\<^sup>* \ \ P \ \ P" proof goal_cases case prems: 1 from prems(1) have "(b,a)\((r\{m}\X)\<^sup>+)\" by simp also have "((r\{m}\X)\<^sup>+)\ = (r\\X\{m})\<^sup>+" by (simp add: converse_add_simps) finally have "(b, a) \ (r\ \ X \ {m})\<^sup>+" . thus ?case by (auto elim!: trancl_multi_insert intro: prems(2,3) simp add: trancl_converse rtrancl_converse ) qed lemma cyclic_subset: "\ \ acyclic R; R \ S \ \ \ acyclic S" unfolding acyclic_def by (blast intro: trancl_mono) subsubsection \Wellfoundedness\ lemma wf_min: assumes A: "wf R" "R\{}" "!!m. m\Domain R - Range R \ P" shows P proof - have H: "!!x. wf R \ \y. (x,y)\R \ x\Domain R - Range R \ (\m. m\Domain R - Range R)" by (erule_tac wf_induct_rule[where P="\x. \y. (x,y)\R \ x\Domain R - Range R \ (\m. m\Domain R - Range R)"]) auto from A(2) obtain x y where "(x,y)\R" by auto with A(1,3) H show ?thesis by blast qed lemma finite_wf_eq_wf_converse: "finite R \ wf (R\) \ wf R" by (metis acyclic_converse finite_acyclic_wf finite_acyclic_wf_converse wf_acyclic) lemma wf_max: assumes A: "wf (R\)" "R\{}" and C: "!!m. m\Range R - Domain R \ P" shows "P" proof - from A(2) have NE: "R\\{}" by auto from wf_min[OF A(1) NE] obtain m where "m\Range R - Domain R" by auto thus P by (blast intro: C) qed \ \Useful lemma to show well-foundedness of some process approaching a finite upper bound\ lemma wf_bounded_supset: "finite S \ wf {(Q',Q). Q'\Q \ Q'\ S}" proof - assume [simp]: "finite S" hence [simp]: "!!x. finite (S-x)" by auto have "{(Q',Q). Q\Q' \ Q'\ S} \ inv_image ({(s'::nat,s). s'Q. card (S-Q))" proof (intro subsetI, case_tac x, simp) fix a b assume A: "b\a \ a\S" hence "S-a \ S-b" by blast thus "card (S-a) < card (S-b)" by (auto simp add: psubset_card_mono) qed moreover have "wf ({(s'::nat,s). s' Range R = {} \ wf R" apply (rule wf_no_loop) by simp text \Extend a wf-relation by a break-condition\ definition "brk_rel R \ {((False,x),(False,y)) | x y. (x,y)\R} \ {((True,x),(False,y)) | x y. True}" lemma brk_rel_wf[simp,intro!]: assumes WF[simp]: "wf R" shows "wf (brk_rel R)" proof - have "wf {((False,x),(False,y)) | x y. (x,y)\R}" proof - have "{((False,x),(False,y)) | x y. (x,y)\R} \ inv_image R snd" by auto from wf_subset[OF wf_inv_image[OF WF] this] show ?thesis . qed moreover have "wf {((True,x),(False,y)) | x y. True}" by (rule wf_no_path) auto ultimately show ?thesis unfolding brk_rel_def apply (subst Un_commute) by (blast intro: wf_Un) qed subsubsection \Restrict Relation\ definition rel_restrict :: "('a \ 'a) set \ 'a set \ ('a \ 'a) set" where "rel_restrict R A \ {(v,w). (v,w) \ R \ v \ A \ w \ A}" lemma rel_restrict_alt_def: "rel_restrict R A = R \ (-A) \ (-A)" unfolding rel_restrict_def by auto lemma rel_restrict_empty[simp]: "rel_restrict R {} = R" by (simp add: rel_restrict_def) lemma rel_restrict_notR: assumes "(x,y) \ rel_restrict A R" shows "x \ R" and "y \ R" using assms unfolding rel_restrict_def by auto lemma rel_restrict_sub: "rel_restrict R A \ R" unfolding rel_restrict_def by auto lemma rel_restrict_Int_empty: "A \ Field R = {} \ rel_restrict R A = R" unfolding rel_restrict_def Field_def by auto lemma Domain_rel_restrict: "Domain (rel_restrict R A) \ Domain R - A" unfolding rel_restrict_def by auto lemma Range_rel_restrict: "Range (rel_restrict R A) \ Range R - A" unfolding rel_restrict_def by auto lemma Field_rel_restrict: "Field (rel_restrict R A) \ Field R - A" unfolding rel_restrict_def Field_def by auto lemma rel_restrict_compl: "rel_restrict R A \ rel_restrict R (-A) = {}" unfolding rel_restrict_def by auto lemma finite_rel_restrict: "finite R \ finite (rel_restrict R A)" by (metis finite_subset rel_restrict_sub) lemma R_subset_Field: "R \ Field R \ Field R" unfolding Field_def by auto lemma homo_rel_restrict_mono: "R \ B \ B \ rel_restrict R A \ (B - A) \ (B - A)" proof - assume A: "R \ B \ B" hence "Field R \ B" unfolding Field_def by auto with Field_rel_restrict have "Field (rel_restrict R A) \ B - A" by (metis Diff_mono order_refl order_trans) with R_subset_Field show ?thesis by blast qed lemma rel_restrict_union: "rel_restrict R (A \ B) = rel_restrict (rel_restrict R A) B" unfolding rel_restrict_def by auto lemma rel_restrictI: "x \ R \ y \ R \ (x,y) \ E \ (x,y) \ rel_restrict E R" unfolding rel_restrict_def by auto lemma rel_restrict_lift: "(x,y) \ rel_restrict E R \ (x,y) \ E" unfolding rel_restrict_def by simp lemma rel_restrict_trancl_mem: "(a,b) \ (rel_restrict A R)\<^sup>+ \ (a,b) \ rel_restrict (A\<^sup>+) R" by (induction rule: trancl_induct) (auto simp add: rel_restrict_def) lemma rel_restrict_trancl_sub: "(rel_restrict A R)\<^sup>+ \ rel_restrict (A\<^sup>+) R" by (metis subrelI rel_restrict_trancl_mem) lemma rel_restrict_mono: "A \ B \ rel_restrict A R \ rel_restrict B R" unfolding rel_restrict_def by auto lemma rel_restrict_mono2: "R \ S \ rel_restrict A S \ rel_restrict A R" unfolding rel_restrict_def by auto lemma rel_restrict_Sigma_sub: "rel_restrict ((A\A)\<^sup>+) R \ ((A - R) \ (A - R))\<^sup>+" unfolding rel_restrict_def by auto (metis DiffI converse_tranclE mem_Sigma_iff r_into_trancl tranclE) lemma finite_reachable_restrictedI: assumes F: "finite Q" assumes I: "I\Q" assumes R: "Range E \ Q" shows "finite (E\<^sup>*``I)" proof - from I R have "E\<^sup>*``I \ Q" by (force elim: rtrancl.cases) also note F finally (finite_subset) show ?thesis . qed context begin private lemma rtrancl_restrictI_aux: assumes "(u,v)\(E-UNIV\R)\<^sup>*" assumes "u\R" shows "(u,v)\(rel_restrict E R)\<^sup>* \ v\R" using assms by (induction) (auto simp: rel_restrict_def intro: rtrancl.intros) corollary rtrancl_restrictI: assumes "(u,v)\(E-UNIV\R)\<^sup>*" assumes "u\R" shows "(u,v)\(rel_restrict E R)\<^sup>*" using rtrancl_restrictI_aux[OF assms] .. end lemma E_closed_restr_reach_cases: assumes P: "(u,v)\E\<^sup>*" assumes CL: "E``R \ R" obtains "v\R" | "u\R" "(u,v)\(rel_restrict E R)\<^sup>*" using P proof (cases rule: rtrancl_last_visit[where S=R]) case no_visit show ?thesis proof (cases "u\R") case True with P have "v\R" using rtrancl_reachable_induct[OF _ CL, where I="{u}"] by auto thus ?thesis .. next case False with no_visit have "(u,v)\(rel_restrict E R)\<^sup>*" by (rule rtrancl_restrictI) with False show ?thesis .. qed next case (last_visit_point x) from \(x, v) \ (E - UNIV \ R)\<^sup>*\ have "(x,v)\E\<^sup>*" by (rule rtrancl_mono_mp[rotated]) auto with \x\R\ have "v\R" using rtrancl_reachable_induct[OF _ CL, where I="{x}"] by auto thus ?thesis .. qed lemma rel_restrict_trancl_notR: assumes "(v,w) \ (rel_restrict E R)\<^sup>+" shows "v \ R" and "w \ R" using assms by (metis rel_restrict_trancl_mem rel_restrict_notR)+ lemma rel_restrict_tranclI: assumes "(x,y) \ E\<^sup>+" and "x \ R" "y \ R" and "E `` R \ R" shows "(x,y) \ (rel_restrict E R)\<^sup>+" using assms proof (induct) case base thus ?case by (metis r_into_trancl rel_restrictI) next case (step y z) hence "y \ R" by auto with step show ?case by (metis trancl_into_trancl rel_restrictI) qed subsubsection \Single-Valued Relations\ lemma single_valued_inter1: "single_valued R \ single_valued (R\S)" by (auto intro: single_valuedI dest: single_valuedD) lemma single_valued_inter2: "single_valued R \ single_valued (S\R)" by (auto intro: single_valuedI dest: single_valuedD) lemma single_valued_below_Id: "R\Id \ single_valued R" by (auto intro: single_valuedI) subsubsection \Bijective Relations\ definition "bijective R \ (\x y z. (x,y)\R \ (x,z)\R \ y=z) \ (\x y z. (x,z)\R \ (y,z)\R \ x=y)" lemma bijective_alt: "bijective R \ single_valued R \ single_valued (R\)" unfolding bijective_def single_valued_def by blast lemma bijective_Id[simp, intro!]: "bijective Id" by (auto simp: bijective_def) lemma bijective_Empty[simp, intro!]: "bijective {}" by (auto simp: bijective_def) subsubsection \Miscellaneous\ lemma pair_vimage_is_Image[simp]: "(Pair u -` E) = E``{u}" by auto lemma fst_in_Field: "fst ` R \ Field R" by (simp add: Field_def fst_eq_Domain) lemma snd_in_Field: "snd ` R \ Field R" by (simp add: Field_def snd_eq_Range) lemma ran_map_of: "ran (map_of xs) \ snd ` set (xs)" by (induct xs) (auto simp add: ran_def) lemma Image_subset_snd_image: "A `` B \ snd ` A" unfolding Image_def image_def by force lemma finite_Image_subset: "finite (A `` B) \ C \ A \ finite (C `` B)" by (metis Image_mono order_refl rev_finite_subset) lemma finite_Field_eq_finite[simp]: "finite (Field R) \ finite R" by (metis finite_cartesian_product finite_subset R_subset_Field finite_Field) definition "fun_of_rel R x \ SOME y. (x,y)\R" lemma for_in_RI(*[intro]*): "x\Domain R \ (x,fun_of_rel R x)\R" unfolding fun_of_rel_def by (auto intro: someI) lemma Field_not_elem: "v \ Field R \ \(x,y) \ R. x \ v \ y \ v" unfolding Field_def by auto lemma Sigma_UNIV_cancel[simp]: "(A \ X - A \ UNIV) = {}" by auto lemma same_fst_trancl[simp]: "(same_fst P R)\<^sup>+ = same_fst P (\x. (R x)\<^sup>+)" proof - { fix x y assume "(x,y)\(same_fst P R)\<^sup>+" hence "(x,y)\same_fst P (\x. (R x)\<^sup>+)" by induction (auto simp: same_fst_def) } moreover { fix f f' x y assume "((f,x),(f',y))\same_fst P (\x. (R x)\<^sup>+)" hence [simp]: "f'=f" "P f" and 1: "(x,y)\(R f)\<^sup>+" by (auto simp: same_fst_def) from 1 have "((f,x),(f',y))\(same_fst P R)\<^sup>+" apply induction subgoal by (rule r_into_trancl) auto subgoal by (erule trancl_into_trancl) auto done } ultimately show ?thesis by auto qed subsection \\option\ Datatype\ lemma le_some_optE: "\Some m\x; !!m'. \x=Some m'; m\m'\ \ P\ \ P" by (cases x) auto lemma not_Some_eq2[simp]: "(\x y. v \ Some (x,y)) = (v = None)" by (cases v) auto subsection "Maps" primrec the_default where "the_default _ (Some x) = x" | "the_default x None = x" declare map_add_dom_app_simps[simp] declare map_add_upd_left[simp] lemma ran_add[simp]: "dom f \ dom g = {} \ ran (f++g) = ran f \ ran g" by (fastforce simp add: ran_def map_add_def split: option.split_asm option.split) lemma nempty_dom: "\e\Map.empty; !!m. m\dom e \ P \ \ P" by (subgoal_tac "dom e \ {}") (blast, auto) lemma le_map_dom_mono: "m\m' \ dom m \ dom m'" apply (safe) apply (drule_tac x=x in le_funD) apply simp apply (erule le_some_optE) apply simp done lemma map_add_first_le: fixes m::"'a\('b::order)" shows "\ m\m' \ \ m++n \ m'++n" apply (rule le_funI) apply (auto simp add: map_add_def split: option.split elim: le_funE) done lemma map_add_distinct_le: shows "\ m\m'; n\n'; dom m' \ dom n' = {} \ \ m++n \ m'++n'" apply (rule le_funI) apply (auto simp add: map_add_def split: option.split) apply (fastforce elim: le_funE) apply (drule le_map_dom_mono) apply (drule le_map_dom_mono) apply (case_tac "m x") apply simp apply (force) apply (fastforce dest!: le_map_dom_mono) apply (erule le_funE) apply (erule_tac x=x in le_funE) apply simp done lemma map_add_left_comm: assumes A: "dom A \ dom B = {}" shows "A ++ (B ++ C) = B ++ (A ++ C)" proof - have "A ++ (B ++ C) = (A++B)++C" by simp also have "\ = (B++A)++C" by (simp add: map_add_comm[OF A]) also have "\ = B++(A++C)" by simp finally show ?thesis . qed lemmas map_add_ac = map_add_assoc map_add_comm map_add_left_comm lemma le_map_restrict[simp]: fixes m :: "'a \ ('b::order)" shows "m |` X \ m" by (rule le_funI) (simp add: restrict_map_def) lemma map_of_distinct_upd: "x \ set (map fst xs) \ [x \ y] ++ map_of xs = (map_of xs) (x \ y)" by (induct xs) (auto simp add: fun_upd_twist) lemma map_of_distinct_upd2: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd3: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ (x,y') # ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd4: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ ys) = (map_of (xs @ (x,y) # ys))(x := None)" using assms by (induct xs) (auto simp: map_of_eq_None_iff) lemma map_of_distinct_lookup: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) x = Some y" proof - have "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys)) (x \ y)" using assms map_of_distinct_upd2 by simp thus ?thesis by simp qed lemma ran_distinct: assumes dist: "distinct (map fst al)" shows "ran (map_of al) = snd ` set al" using assms proof (induct al) case Nil then show ?case by simp next case (Cons kv al) then have "ran (map_of al) = snd ` set al" by simp moreover from Cons.prems have "map_of al (fst kv) = None" by (simp add: map_of_eq_None_iff) ultimately show ?case by (simp only: map_of.simps ran_map_upd) simp qed lemma ran_is_image: "ran M = (the \ M) ` (dom M)" unfolding ran_def dom_def image_def by auto lemma map_card_eq_iff: assumes finite: "finite (dom M)" and card_eq: "card (dom M) = card (ran M)" and indom: "x \ dom M" shows "(M x = M y) \ (x = y)" proof - from ran_is_image finite card_eq have *: "inj_on (the \ M) (dom M)" using eq_card_imp_inj_on by metis thus ?thesis proof (cases "y \ dom M") case False with indom show ?thesis by auto next case True with indom have "the (M x) = the (M y) \ (x = y)" using inj_on_eq_iff[OF *] by auto thus ?thesis by auto qed qed lemma map_dom_ran_finite: "finite (dom M) \ finite (ran M)" by (simp add: ran_is_image) lemma map_update_eta_repair[simp]: (* An update operation may get simplified, if it happens to be eta-expanded. This lemma tries to repair some common expressions *) "dom (\x. if x=k then Some v else m x) = insert k (dom m)" "m k = None \ ran (\x. if x=k then Some v else m x) = insert v (ran m)" apply auto [] apply (force simp: ran_def) done lemma map_leI[intro?]: "\\x v. m1 x = Some v \ m2 x = Some v\ \ m1\\<^sub>mm2" unfolding map_le_def by force lemma map_leD: "m1\\<^sub>mm2 \ m1 k = Some v \ m2 k = Some v" unfolding map_le_def by force lemma map_restrict_insert_none_simp: "m x = None \ m|`(-insert x s) = m|`(-s)" by (auto intro!: ext simp:restrict_map_def) (* TODO: Move *) lemma eq_f_restr_conv: "s\dom (f A) \ A = f A |` (-s) \ A \\<^sub>m f A \ s = dom (f A) - dom A" apply auto subgoal by (metis map_leI restrict_map_eq(2)) subgoal by (metis ComplD restrict_map_eq(2)) subgoal by (metis Compl_iff restrict_in) subgoal by (force simp: map_le_def restrict_map_def) done corollary eq_f_restr_ss_eq: "\ s\dom (f A) \ \ A = f A |` (-s) \ A \\<^sub>m f A \ s = dom (f A) - dom A" using eq_f_restr_conv by blast subsubsection \Simultaneous Map Update\ definition "map_mmupd m K v k \ if k\K then Some v else m k" lemma map_mmupd_empty[simp]: "map_mmupd m {} v = m" by (auto simp: map_mmupd_def) lemma mmupd_in_upd[simp]: "k\K \ map_mmupd m K v k = Some v" by (auto simp: map_mmupd_def) lemma mmupd_notin_upd[simp]: "k\K \ map_mmupd m K v k = m k" by (auto simp: map_mmupd_def) lemma map_mmupdE: assumes "map_mmupd m K v k = Some x" obtains "k\K" "m k = Some x" | "k\K" "x=v" using assms by (auto simp: map_mmupd_def split: if_split_asm) lemma dom_mmupd[simp]: "dom (map_mmupd m K v) = dom m \ K" by (auto simp: map_mmupd_def split: if_split_asm) lemma le_map_mmupd_not_dom[simp, intro!]: "m \\<^sub>m map_mmupd m (K-dom m) v" by (auto simp: map_le_def) lemma map_mmupd_update_less: "K\K' \ map_mmupd m (K - dom m) v \\<^sub>m map_mmupd m (K'-dom m) v" by (auto simp: map_le_def map_mmupd_def) subsection\Connection between Maps and Sets of Key-Value Pairs\ definition map_to_set where "map_to_set m = {(k, v) . m k = Some v}" definition set_to_map where "set_to_map S k = Eps_Opt (\v. (k, v) \ S)" lemma set_to_map_simp : assumes inj_on_fst: "inj_on fst S" shows "(set_to_map S k = Some v) \ (k, v) \ S" proof (cases "\v. (k, v) \ S") case True note kv_ex = this then obtain v' where kv'_in: "(k, v') \ S" by blast with inj_on_fst have kv''_in: "\v''. (k, v'') \ S \ v' = v''" unfolding inj_on_def Ball_def by auto show ?thesis unfolding set_to_map_def by (simp add: kv_ex kv''_in) next case False hence kv''_nin: "\v''. (k, v'') \ S" by simp thus ?thesis by (simp add: set_to_map_def) qed lemma inj_on_fst_map_to_set : "inj_on fst (map_to_set m)" unfolding map_to_set_def inj_on_def by simp lemma map_to_set_inverse : "set_to_map (map_to_set m) = m" proof fix k show "set_to_map (map_to_set m) k = m k" proof (cases "m k") case None note mk_eq = this hence "\v. (k, v) \ map_to_set m" unfolding map_to_set_def by simp with set_to_map_simp [OF inj_on_fst_map_to_set, of m k] show ?thesis unfolding mk_eq by auto next case (Some v) note mk_eq = this hence "(k, v) \ map_to_set m" unfolding map_to_set_def by simp with set_to_map_simp [OF inj_on_fst_map_to_set, of m k v] show ?thesis unfolding mk_eq by auto qed qed lemma set_to_map_inverse : assumes inj_on_fst_S: "inj_on fst S" shows "map_to_set (set_to_map S) = S" proof (rule set_eqI) fix kv from set_to_map_simp [OF inj_on_fst_S, of "fst kv" "snd kv"] show "(kv \ map_to_set (set_to_map S)) = (kv \ S)" unfolding map_to_set_def by auto qed lemma map_to_set_empty[simp]: "map_to_set Map.empty = {}" unfolding map_to_set_def by simp lemma set_to_map_empty[simp]: "set_to_map {} = Map.empty" unfolding set_to_map_def[abs_def] by simp lemma map_to_set_empty_iff: "map_to_set m = {} \ m = Map.empty" "{} = map_to_set m \ m = Map.empty" unfolding map_to_set_def by auto lemma set_to_map_empty_iff: "set_to_map S = Map.empty \ S = {}" (is ?T1) "Map.empty = set_to_map S \ S = {}" (is ?T2) proof - show T1: ?T1 apply (simp only: set_eq_iff) apply (simp only: fun_eq_iff) apply (simp add: set_to_map_def) apply auto done from T1 show ?T2 by auto qed lemma map_to_set_upd[simp]: "map_to_set (m (k \ v)) = insert (k, v) (map_to_set m - {(k, v') |v'. True})" unfolding map_to_set_def apply (simp add: set_eq_iff) apply metis done lemma set_to_map_insert: assumes k_nin: "fst kv \ fst ` S" shows "set_to_map (insert kv S) = (set_to_map S) (fst kv \ snd kv)" proof fix k' obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust) from k_nin have k_nin': "\v'. (k, v') \ S" by (auto simp add: image_iff Ball_def) show "set_to_map (insert kv S) k' = (set_to_map S(fst kv \ snd kv)) k'" by (simp add: set_to_map_def k_nin') qed lemma map_to_set_dom : "dom m = fst ` (map_to_set m)" unfolding dom_def map_to_set_def by (auto simp add: image_iff) lemma map_to_set_ran : "ran m = snd ` (map_to_set m)" unfolding ran_def map_to_set_def by (auto simp add: image_iff) lemma set_to_map_dom : "dom (set_to_map S) = fst ` S" unfolding set_to_map_def[abs_def] dom_def by (auto simp add: image_iff Bex_def) lemma set_to_map_ran : "ran (set_to_map S) \ snd ` S" unfolding set_to_map_def[abs_def] ran_def subset_iff by (auto simp add: image_iff Bex_def) (metis Eps_Opt_eq_Some) lemma finite_map_to_set: "finite (map_to_set m) = finite (dom m)" unfolding map_to_set_def map_to_set_dom apply (intro iffI finite_imageI) apply assumption apply (rule finite_imageD[of fst]) apply assumption apply (simp add: inj_on_def) done lemma card_map_to_set : "card (map_to_set m) = card (dom m)" unfolding map_to_set_def map_to_set_dom apply (rule card_image[symmetric]) apply (simp add: inj_on_def) done lemma map_of_map_to_set : "distinct (map fst l) \ map_of l = m \ set l = map_to_set m" proof (induct l arbitrary: m) case Nil thus ?case by (simp add: map_to_set_empty_iff) blast next case (Cons kv l m) obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust) from Cons(2) have dist_l: "distinct (map fst l)" and kv'_nin: "\v'. (k, v') \ set l" by (auto simp add: image_iff) note ind_hyp = Cons(1)[OF dist_l] from kv'_nin have l_eq: "set (kv # l) = map_to_set m \ (set l = map_to_set (m (k := None))) \ m k = Some v" apply (simp add: map_to_set_def restrict_map_def set_eq_iff) apply (auto) apply (metis) apply (metis option.inject) done from kv'_nin have m_eq: "map_of (kv # l) = m \ map_of l = (m (k := None)) \ m k = Some v" apply (simp add: fun_eq_iff restrict_map_def map_of_eq_None_iff image_iff Ball_def) apply metis done show ?case unfolding m_eq l_eq using ind_hyp[of "m (k := None)"] by metis qed lemma map_to_set_map_of : "distinct (map fst l) \ map_to_set (map_of l) = set l" by (metis map_of_map_to_set) subsubsection \Mapping empty set to None\ definition "dflt_None_set S \ if S={} then None else Some S" lemma the_dflt_None_empty[simp]: "dflt_None_set {} = None" unfolding dflt_None_set_def by simp lemma the_dflt_None_nonempty[simp]: "S\{} \ dflt_None_set S = Some S" unfolding dflt_None_set_def by simp lemma the_dflt_None_set[simp]: "the_default {} (dflt_None_set x) = x" unfolding dflt_None_set_def by auto subsection \Orderings\ lemma (in order) min_arg_le[simp]: "n \ min m n \ min m n = n" "m \ min m n \ min m n = m" by (auto simp: min_def) lemma (in linorder) min_arg_not_ge[simp]: "\ min m n < m \ min m n = m" "\ min m n < n \ min m n = n" by (auto simp: min_def) lemma (in linorder) min_eq_arg[simp]: "min m n = m \ m\n" "min m n = n \ n\m" by (auto simp: min_def) lemma min_simps[simp]: "a<(b::'a::order) \ min a b = a" "b<(a::'a::order) \ min a b = b" by (auto simp add: min_def dest: less_imp_le) lemma (in -) min_less_self_conv[simp]: "min a b < a \ b < (a::_::linorder)" "min a b < b \ a < (b::_::linorder)" by (auto simp: min_def) lemma ord_eq_le_eq_trans: "\ a=b; b\c; c=d \ \ a\d" by auto lemma zero_comp_diff_simps[simp]: "(0::'a::linordered_idom) \ a - b \ b \ a" "(0::'a::linordered_idom) < a - b \ b < a" by auto subsubsection \Termination Measures\ text \Lexicographic measure, assuming upper bound for second component\ lemma mlex_fst_decrI: fixes a a' b b' N :: nat assumes "a a*N + N" using \b by linarith also have "\ \ a'*N" using \a by (metis Suc_leI ab_semigroup_add_class.add.commute ab_semigroup_mult_class.mult.commute mult_Suc_right mult_le_mono2) also have "\ \ a'*N + b'" by auto finally show ?thesis by auto qed lemma mlex_leI: fixes a a' b b' N :: nat assumes "a\a'" assumes "b\b'" shows "a*N + b \ a'*N + b'" using assms by (auto intro!: add_mono) lemma mlex_snd_decrI: fixes a a' b b' N :: nat assumes "a=a'" assumes "bCCPOs\ context ccpo begin lemma ccpo_Sup_mono: assumes C: "Complete_Partial_Order.chain (\) A" "Complete_Partial_Order.chain (\) B" assumes B: "\x\A. \y\B. x\y" shows "Sup A \ Sup B" proof (rule ccpo_Sup_least) fix x assume "x\A" with B obtain y where I: "y\B" and L: "x\y" by blast note L also from I ccpo_Sup_upper have "y\Sup B" by (blast intro: C) finally show "x\Sup B" . qed (rule C) lemma fixp_mono: assumes M: "monotone (\) (\) f" "monotone (\) (\) g" assumes LE: "\Z. f Z \ g Z" shows "ccpo_class.fixp f \ ccpo_class.fixp g" unfolding fixp_def[abs_def] apply (rule ccpo_Sup_mono) apply (rule chain_iterates M)+ proof rule fix x assume "x\ccpo_class.iterates f" thus "\y\ccpo_class.iterates g. x\y" proof (induct) case (step x) then obtain y where I: "y\ccpo_class.iterates g" and L: "x\y" by blast hence "g y \ ccpo_class.iterates g" and "f x \ g y" apply - apply (erule iterates.step) apply (rule order_trans) apply (erule monotoneD[OF M(1)]) apply (rule LE) done thus "\y\ccpo_class.iterates g. f x \ y" .. next case (Sup M) define N where "N = {SOME y. y\ccpo_class.iterates g \ x\y | x. x\M}" have N1: "\y\N. y\ccpo_class.iterates g \ (\x\M. x\y)" unfolding N_def apply auto apply (metis (lifting) Sup.hyps(2) tfl_some) by (metis (lifting) Sup.hyps(2) tfl_some) have N2: "\x\M. \y\N. x\y" unfolding N_def apply auto by (metis (lifting) Sup.hyps(2) tfl_some) have "N \ ccpo_class.iterates g" using Sup using N1 by auto hence C_chain: "Complete_Partial_Order.chain (\) N" using chain_iterates[OF M(2)] unfolding chain_def by auto have "Sup N \ ccpo_class.iterates g" and "Sup M \ Sup N" apply - apply (rule iterates.Sup[OF C_chain]) using N1 apply blast apply (rule ccpo_Sup_mono) apply (rule Sup.hyps) apply (rule C_chain) apply (rule N2) done thus ?case by blast qed qed end subsection \Code\ text \Constant for code-abort. If that gets executed, an abort-exception is raised.\ definition [simp]: "CODE_ABORT f = f ()" declare [[code abort: CODE_ABORT]] end diff --git a/thys/Bertrands_Postulate/Bertrand.thy b/thys/Bertrands_Postulate/Bertrand.thy --- a/thys/Bertrands_Postulate/Bertrand.thy +++ b/thys/Bertrands_Postulate/Bertrand.thy @@ -1,1853 +1,1854 @@ (* File: Bertrand.thy Authors: Julian Biendarra, Manuel Eberl , Larry Paulson A proof of Bertrand's postulate (based on John Harrison's HOL Light proof). Uses reflection and the approximation tactic. *) theory Bertrand imports Complex_Main "HOL-Number_Theory.Number_Theory" "HOL-Library.Discrete" "HOL-Decision_Procs.Approximation_Bounds" "HOL-Library.Code_Target_Numeral" Pratt_Certificate.Pratt_Certificate begin subsection \Auxiliary facts\ lemma ln_2_le: "ln 2 \ 355 / (512 :: real)" proof - have "ln 2 \ real_of_float (ub_ln2 12)" by (rule ub_ln2) also have "ub_ln2 12 = Float 5680 (- 13)" by code_simp finally show ?thesis by simp qed lemma ln_2_ge: "ln 2 \ (5677 / 8192 :: real)" proof - have "ln 2 \ real_of_float (lb_ln2 12)" by (rule lb_ln2) also have "lb_ln2 12 = Float 5677 (-13)" by code_simp finally show ?thesis by simp qed lemma ln_2_ge': "ln (2 :: real) \ 2/3" and ln_2_le': "ln (2 :: real) \ 16/23" using ln_2_le ln_2_ge by simp_all lemma of_nat_ge_1_iff: "(of_nat x :: 'a :: linordered_semidom) \ 1 \ x \ 1" using of_nat_le_iff[of 1 x] by (subst (asm) of_nat_1) lemma floor_conv_div_nat: "of_int (floor (real m / real n)) = real (m div n)" by (subst floor_divide_of_nat_eq) simp lemma frac_conv_mod_nat: "frac (real m / real n) = real (m mod n) / real n" by (cases "n = 0") (simp_all add: frac_def floor_conv_div_nat field_simps of_nat_mult [symmetric] of_nat_add [symmetric] del: of_nat_mult of_nat_add) lemma of_nat_prod_mset: "prod_mset (image_mset of_nat A) = of_nat (prod_mset A)" by (induction A) simp_all lemma prod_mset_pos: "(\x :: 'a :: linordered_semidom. x \# A \ x > 0) \ prod_mset A > 0" by (induction A) simp_all lemma ln_msetprod: assumes "\x. x \#I \ x > 0" shows "(\p::nat\#I. ln p) = ln (\p\#I. p)" using assms by (induction I) (simp_all add: of_nat_prod_mset ln_mult prod_mset_pos) lemma ln_fact: "ln (fact n) = (\d=1..n. ln d)" by (induction n) (simp_all add: ln_mult) lemma overpower_lemma: fixes f g :: "real \ real" assumes "f a \ g a" assumes "\x. a \ x \ ((\x. g x - f x) has_real_derivative (d x)) (at x)" assumes "\x. a \ x \ d x \ 0" assumes "a \ x" shows "f x \ g x" proof (cases "a < x") case True with assms have "\z. z > a \ z < x \ g x - f x - (g a - f a) = (x - a) * d z" by (intro MVT2) auto then obtain z where z: "z > a" "z < x" "g x - f x - (g a - f a) = (x - a) * d z" by blast hence "f x = g x + (f a - g a) + (a - x) * d z" by (simp add: algebra_simps) also from assms have "f a - g a \ 0" by (simp add: algebra_simps) also from assms z have "(a - x) * d z \ 0 * d z" by (intro mult_right_mono) simp_all finally show ?thesis by simp qed (insert assms, auto) subsection \Preliminary definitions\ definition primepow_even :: "nat \ bool" where "primepow_even q \ (\ p k. 1 \ k \ prime p \ q = p^(2*k))" definition primepow_odd :: "nat \ bool" where "primepow_odd q \ (\ p k. 1 \ k \ prime p \ q = p^(2*k+1))" abbreviation (input) isprimedivisor :: "nat \ nat \ bool" where "isprimedivisor q p \ prime p \ p dvd q" definition pre_mangoldt :: "nat \ nat" where "pre_mangoldt d = (if primepow d then aprimedivisor d else 1)" definition mangoldt_even :: "nat \ real" where "mangoldt_even d = (if primepow_even d then ln (real (aprimedivisor d)) else 0)" definition mangoldt_odd :: "nat \ real" where "mangoldt_odd d = (if primepow_odd d then ln (real (aprimedivisor d)) else 0)" definition mangoldt_1 :: "nat \ real" where "mangoldt_1 d = (if prime d then ln d else 0)" definition psi :: "nat \ real" where "psi n = (\d=1..n. mangoldt d)" definition psi_even :: "nat \ real" where "psi_even n = (\d=1..n. mangoldt_even d)" definition psi_odd :: "nat \ real" where "psi_odd n = (\d=1..n. mangoldt_odd d)" abbreviation (input) psi_even_2 :: "nat \ real" where "psi_even_2 n \ (\d=2..n. mangoldt_even d)" abbreviation (input) psi_odd_2 :: "nat \ real" where "psi_odd_2 n \ (\d=2..n. mangoldt_odd d)" definition theta :: "nat \ real" where "theta n = (\p=1..n. if prime p then ln (real p) else 0)" subsection \Properties of prime powers\ lemma primepow_even_imp_primepow: assumes "primepow_even n" shows "primepow n" proof - from assms obtain p k where "1 \ k" "prime p" "n = p ^ (2 * k)" unfolding primepow_even_def by blast moreover from \1 \ k\ have "2 * k > 0" by simp ultimately show ?thesis unfolding primepow_def by blast qed lemma primepow_odd_imp_primepow: assumes "primepow_odd n" shows "primepow n" proof - from assms obtain p k where "1 \ k" "prime p" "n = p ^ (2 * k + 1)" unfolding primepow_odd_def by blast moreover from \1 \ k\ have "Suc (2 * k) > 0" by simp ultimately show ?thesis unfolding primepow_def by (auto simp del: power_Suc) qed lemma primepow_odd_altdef: "primepow_odd n \ primepow n \ odd (multiplicity (aprimedivisor n) n) \ multiplicity (aprimedivisor n) n > 1" proof (intro iffI conjI; (elim conjE)?) assume "primepow_odd n" then obtain p k where n: "k \ 1" "prime p" "n = p ^ (2 * k + 1)" by (auto simp: primepow_odd_def) thus "odd (multiplicity (aprimedivisor n) n)" "multiplicity (aprimedivisor n) n > 1" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) next assume A: "primepow n" and B: "odd (multiplicity (aprimedivisor n) n)" and C: "multiplicity (aprimedivisor n) n > 1" from A obtain p k where n: "k \ 1" "prime p" "n = p ^ k" by (auto simp: primepow_def Suc_le_eq) with B C have "odd k" "k > 1" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) then obtain j where j: "k = 2 * j + 1" "j > 0" by (auto elim!: oddE) with n show "primepow_odd n" by (auto simp: primepow_odd_def intro!: exI[of _ p, OF exI[of _ j]]) qed (auto dest: primepow_odd_imp_primepow) lemma primepow_even_altdef: "primepow_even n \ primepow n \ even (multiplicity (aprimedivisor n) n)" proof (intro iffI conjI; (elim conjE)?) assume "primepow_even n" then obtain p k where n: "k \ 1" "prime p" "n = p ^ (2 * k)" by (auto simp: primepow_even_def) thus "even (multiplicity (aprimedivisor n) n)" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) next assume A: "primepow n" and B: "even (multiplicity (aprimedivisor n) n)" from A obtain p k where n: "k \ 1" "prime p" "n = p ^ k" by (auto simp: primepow_def Suc_le_eq) with B have "even k" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) then obtain j where j: "k = 2 * j" by (auto elim!: evenE) from j n have "j \ 0" by (intro notI) simp_all with j n show "primepow_even n" by (auto simp: primepow_even_def intro!: exI[of _ p, OF exI[of _ j]]) qed (auto dest: primepow_even_imp_primepow) lemma primepow_odd_mult: assumes "d > Suc 0" shows "primepow_odd (aprimedivisor d * d) \ primepow_even d" using assms by (auto simp: primepow_odd_altdef primepow_even_altdef primepow_mult_aprimedivisorI aprimedivisor_primepow prime_aprimedivisor' aprimedivisor_dvd' prime_elem_multiplicity_mult_distrib prime_elem_aprimedivisor_nat dest!: primepow_multD) lemma pre_mangoldt_primepow: assumes "primepow n" "aprimedivisor n = p" shows "pre_mangoldt n = p" using assms by (simp add: pre_mangoldt_def) lemma pre_mangoldt_notprimepow: assumes "\primepow n" shows "pre_mangoldt n = 1" using assms by (simp add: pre_mangoldt_def) lemma primepow_cases: "primepow d \ ( primepow_even d \ \ primepow_odd d \ \ prime d) \ (\ primepow_even d \ primepow_odd d \ \ prime d) \ (\ primepow_even d \ \ primepow_odd d \ prime d)" by (auto simp: primepow_even_altdef primepow_odd_altdef multiplicity_aprimedivisor_Suc_0_iff elim!: oddE intro!: Nat.gr0I) subsection \Deriving a recurrence for the psi function\ lemma ln_fact_bounds: assumes "n > 0" shows "abs(ln (fact n) - n * ln n + n) \ 1 + ln n" proof - have "\n\{0<..}. \z>real n. z < real (n + 1) \ real (n + 1) * ln (real (n + 1)) - real n * ln (real n) = (real (n + 1) - real n) * (ln z + 1)" by (intro ballI MVT2) (auto intro!: derivative_eq_intros) hence "\n\{0<..}. \z>real n. z < real (n + 1) \ real (n + 1) * ln (real (n + 1)) - real n * ln (real n) = (ln z + 1)" by (simp add: algebra_simps) from bchoice[OF this] obtain k :: "nat \ real" where lb: "real n < k n" and ub: "k n < real (n + 1)" and mvt: "real (n+1) * ln (real (n+1)) - real n * ln (real n) = ln (k n) + 1" if "n > 0" for n::nat by blast have *: "(n + 1) * ln (n + 1) = (\i=1..n. ln(k i) + 1)" for n::nat proof (induction n) case (Suc n) have "(\i = 1..n+1. ln (k i) + 1) = (\i = 1..n. ln (k i) + 1) + ln (k (n+1)) + 1" by simp also from Suc.IH have "(\i = 1..n. ln (k i) + 1) = real (n+1) * ln (real (n+1))" .. also from mvt[of "n+1"] have "\ = real (n+2) * ln (real (n+2)) - ln (k (n+1)) - 1" by simp finally show ?case by simp qed simp have **: "abs((\i=1..n+1. ln i) - ((n+1) * ln (n+1) - (n+1))) \ 1 + ln(n+1)" for n::nat proof - have "(\i=1..n+1. ln i) \ (\i=1..n. ln i) + ln (n+1)" by simp also have "(\i=1..n. ln i) \ (\i=1..n. ln (k i))" by (intro sum_mono, subst ln_le_cancel_iff) (auto simp: Suc_le_eq dest: lb ub) also have "\ = (\i=1..n. ln (k i) + 1) - n" by (simp add: sum.distrib) also from * have "\ = (n+1) * ln (n+1) - n" by simp finally have a_minus_b: "(\i=1..n+1. ln i) - ((n+1) * ln (n+1) - (n+1)) \ 1 + ln (n+1)" by simp from * have "(n+1) * ln (n+1) - n = (\i=1..n. ln (k i) + 1) - n" by simp also have "\ = (\i=1..n. ln (k i))" by (simp add: sum.distrib) also have "\ \ (\i=1..n. ln (i+1))" by (intro sum_mono, subst ln_le_cancel_iff) (auto simp: Suc_le_eq dest: lb ub) also from sum.shift_bounds_cl_nat_ivl[of "ln" 1 1 n] have "\ = (\i=1+1..n+1. ln i)" .. also have "\ = (\i=1..n+1. ln i)" by (rule sum.mono_neutral_left) auto finally have b_minus_a: "((n+1) * ln (n+1) - (n+1)) - (\i=1..n+1. ln i) \ 1" by simp have "0 \ ln (n+1)" by simp with b_minus_a have "((n+1) * ln (n+1) - (n+1)) - (\i=1..n+1. ln i) \ 1 + ln (n+1)" by linarith with a_minus_b show ?thesis by linarith qed from \n > 0\ have "n \ 1" by simp thus ?thesis proof (induction n rule: dec_induct) case base then show ?case by simp next case (step n) from ln_fact[of "n+1"] **[of n] show ?case by simp qed qed lemma ln_fact_diff_bounds: "abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2) \ 4 * ln (if n = 0 then 1 else n) + 3" proof (cases "n div 2 = 0") case True hence "n \ 1" by simp with ln_le_minus_one[of "2::real"] show ?thesis by (cases n) simp_all next case False then have "n > 1" by simp let ?a = "real n * ln 2" let ?b = "4 * ln (real n) + 3" let ?l1 = "ln (fact (n div 2))" let ?a1 = "real (n div 2) * ln (real (n div 2)) - real (n div 2)" let ?b1 = "1 + ln (real (n div 2))" let ?l2 = "ln (fact n)" let ?a2 = "real n * ln (real n) - real n" let ?b2 = "1 + ln (real n)" have abs_a: "abs(?a - (?a2 - 2 * ?a1)) \ ?b - 2 * ?b1 - ?b2" proof (cases "even n") case True then have "real (2 * (n div 2)) = real n" by simp then have n_div_2: "real (n div 2) = real n / 2" by simp from \n > 1\ have *: "abs(?a - (?a2 - 2 * ?a1)) = 0" by (simp add: n_div_2 ln_div algebra_simps) from \even n\ and \n > 1\ have "0 \ ln (real n) - ln (real (n div 2))" by (auto elim: evenE) also have "2 * \ \ 3 * ln (real n) - 2 * ln (real (n div 2))" using \n > 1\ by (auto intro!: ln_ge_zero) also have "\ = ?b - 2 * ?b1 - ?b2" by simp finally show ?thesis using * by simp next case False then have "real (2 * (n div 2)) = real (n - 1)" by simp with \n > 1\ have n_div_2: "real (n div 2) = (real n - 1) / 2" by simp from \odd n\ \n div 2 \ 0\ have "n \ 3" by presburger have "?a - (?a2 - 2 * ?a1) = real n * ln 2 - real n * ln (real n) + real n + 2 * real (n div 2) * ln (real (n div 2)) - 2* real (n div 2)" by (simp add: algebra_simps) also from n_div_2 have "2 * real (n div 2) = real n - 1" by simp also have "real n * ln 2 - real n * ln (real n) + real n + (real n - 1) * ln (real (n div 2)) - (real n - 1) = real n * (ln (real n - 1) - ln (real n)) - ln (real (n div 2)) + 1" using \n > 1\ by (simp add: algebra_simps n_div_2 ln_div) finally have lhs: "abs(?a - (?a2 - 2 * ?a1)) = abs(real n * (ln (real n - 1) - ln (real n)) - ln (real (n div 2)) + 1)" by simp from \n > 1\ have "real n * (ln (real n - 1) - ln (real n)) \ 0" by (simp add: algebra_simps mult_left_mono) moreover from \n > 1\ have "ln (real (n div 2)) \ ln (real n)" by simp moreover { have "exp 1 \ (3::real)" by (rule exp_le) also from \n \ 3\ have "\ \ exp (ln (real n))" by simp finally have "ln (real n) \ 1" by simp } ultimately have ub: "real n * (ln (real n - 1) - ln (real n)) - ln(real (n div 2)) + 1 \ 3 * ln (real n) - 2 * ln(real (n div 2))" by simp have mon: "real n' * (ln (real n') - ln (real n' - 1)) \ real n * (ln (real n) - ln (real n - 1))" if "n \ 3" "n' \ n" for n n'::nat proof (rule DERIV_nonpos_imp_nonincreasing[where f = "\x. x * (ln x - ln (x - 1))"]) fix t assume t: "real n \ t" "t \ real n'" with that have "1 / (t - 1) \ ln (1 + 1/(t - 1))" by (intro ln_add_one_self_le_self) simp_all also from t that have "ln (1 + 1/(t - 1)) = ln t- ln (t - 1)" by (simp add: ln_div [symmetric] field_simps) finally have "ln t - ln (t - 1) \ 1 / (t - 1)" . with that t show "\y. ((\x. x * (ln x - ln (x - 1))) has_field_derivative y) (at t) \ y \ 0" by (intro exI[of _ "1 / (1 - t) + ln t - ln (t - 1)"]) (force intro!: derivative_eq_intros simp: field_simps)+ qed (use that in simp_all) from \n > 1\ have "ln 2 = ln (real n) - ln (real n / 2)" by (simp add: ln_div) also from \n > 1\ have "\ \ ln (real n) - ln (real (n div 2))" by simp finally have *: "3*ln 2 + ln(real (n div 2)) \ 3* ln(real n) - 2* ln(real (n div 2))" by simp have "- real n * (ln (real n - 1) - ln (real n)) + ln(real (n div 2)) - 1 = real n * (ln (real n) - ln (real n - 1)) - 1 + ln(real (n div 2))" by (simp add: algebra_simps) also have "real n * (ln (real n) - ln (real n - 1)) \ 3 * (ln 3 - ln (3 - 1))" using mon[OF _ \n \ 3\] by simp also { have "Some (Float 3 (-1)) = ub_ln 1 3" by code_simp from ub_ln(1)[OF this] have "ln 3 \ (1.6 :: real)" by simp also have "1.6 - 1 / 3 \ 2 * (2/3 :: real)" by simp also have "2/3 \ ln (2 :: real)" by (rule ln_2_ge') finally have "ln 3 - 1 / 3 \ 2 * ln (2 :: real)" by simp } hence "3 * (ln 3 - ln (3 - 1)) - 1 \ 3 * ln (2 :: real)" by simp also note * finally have "- real n * (ln (real n - 1) - ln (real n)) + ln(real (n div 2)) - 1 \ 3 * ln (real n) - 2 * ln (real (n div 2))" by simp hence lhs': "abs(real n * (ln (real n - 1) - ln (real n)) - ln(real (n div 2)) + 1) \ 3 * ln (real n) - 2 * ln (real (n div 2))" using ub by simp have rhs: "?b - 2 * ?b1 - ?b2 = 3* ln (real n) - 2 * ln (real (n div 2))" by simp from \n > 1\ have "ln (real (n div 2)) \ 3* ln (real n) - 2* ln (real (n div 2))" by simp with rhs lhs lhs' show ?thesis by simp qed then have minus_a: "-?a \ ?b - 2 * ?b1 - ?b2 - (?a2 - 2 * ?a1)" by simp from abs_a have a: "?a \ ?b - 2 * ?b1 - ?b2 + ?a2 - 2 * ?a1" by (simp) from ln_fact_bounds[of "n div 2"] False have abs_l1: "abs(?l1 - ?a1) \ ?b1" by (simp add: algebra_simps) then have minus_l1: "?a1 - ?l1 \ ?b1" by linarith from abs_l1 have l1: "?l1 - ?a1 \ ?b1" by linarith from ln_fact_bounds[of n] False have abs_l2: "abs(?l2 - ?a2) \ ?b2" by (simp add: algebra_simps) then have l2: "?l2 - ?a2 \ ?b2" by simp from abs_l2 have minus_l2: "?a2 - ?l2 \ ?b2" by simp from minus_a minus_l1 l2 have "?l2 - 2 * ?l1 - ?a \ ?b" by simp moreover from a l1 minus_l2 have "- ?l2 + 2 * ?l1 + ?a \ ?b" by simp ultimately have "abs((?l2 - 2*?l1) - ?a) \ ?b" by simp then show ?thesis by simp qed lemma ln_primefact: assumes "n \ (0::nat)" shows "ln n = (\d=1..n. if primepow d \ d dvd n then ln (aprimedivisor d) else 0)" (is "?lhs = ?rhs") proof - have "?rhs = (\d\{x \ {1..n}. primepow x \ x dvd n}. ln (real (aprimedivisor d)))" unfolding primepow_factors_def by (subst sum.inter_filter [symmetric]) simp_all also have "{x \ {1..n}. primepow x \ x dvd n} = primepow_factors n" using assms by (auto simp: primepow_factors_def dest: dvd_imp_le primepow_gt_Suc_0) finally have *: "(\d\primepow_factors n. ln (real (aprimedivisor d))) = ?rhs" .. from in_prime_factors_imp_prime prime_gt_0_nat have pf_pos: "\p. p\#prime_factorization n \ p > 0" by blast from ln_msetprod[of "prime_factorization n", OF pf_pos] assms have "ln n = (\p\#prime_factorization n. ln p)" by (simp add: of_nat_prod_mset) also from * sum_prime_factorization_conv_sum_primepow_factors[of n ln, OF assms(1)] have "\ = ?rhs" by simp finally show ?thesis . qed context begin private lemma divisors: fixes x d::nat assumes "x \ {1..n}" assumes "d dvd x" shows "\k\{1..n div d}. x = d * k" proof - from assms have "x \ n" by simp then have ub: "x div d \ n div d" by (simp add: div_le_mono \x \ n\) from assms have "1 \ x div d" by (auto elim!: dvdE) with ub have "x div d \ {1..n div d}" by simp with \d dvd x\ show ?thesis by (auto intro!: bexI[of _ "x div d"]) qed lemma ln_fact_conv_mangoldt: "ln (fact n) = (\d=1..n. mangoldt d * floor (n / d))" proof - have *: "(\da=1..n. if primepow da \ da dvd d then ln (aprimedivisor da) else 0) = (\(da::nat)=1..d. if primepow da \ da dvd d then ln (aprimedivisor da) else 0)" if d: "d \ {1..n}" for d by (rule sum.mono_neutral_right, insert d) (auto dest: dvd_imp_le) have "(\d=1..n. \da=1..d. if primepow da \ da dvd d then ln (aprimedivisor da) else 0) = (\d=1..n. \da=1..n. if primepow da \ da dvd d then ln (aprimedivisor da) else 0)" by (rule sum.cong) (insert *, simp_all) also have "\ = (\da=1..n. \d=1..n. if primepow da \ da dvd d then ln (aprimedivisor da) else 0)" by (rule sum.swap) also have "\ = sum (\d. mangoldt d * floor (n/d)) {1..n}" proof (rule sum.cong) fix d assume d: "d \ {1..n}" have "(\da = 1..n. if primepow d \ d dvd da then ln (real (aprimedivisor d)) else 0) = (\da = 1..n. if d dvd da then mangoldt d else 0)" by (intro sum.cong) (simp_all add: mangoldt_def) also have "\ = mangoldt d * real (card {x. x \ {1..n} \ d dvd x})" by (subst sum.inter_filter [symmetric]) (simp_all add: algebra_simps) also { have "{x. x \ {1..n} \ d dvd x} = {x. \k \{1..n div d}. x=k*d}" proof safe fix x assume "x \ {1..n}" "d dvd x" thus "\k\{1..n div d}. x = k * d" using divisors[of x n d] by auto next fix x k assume k: "k \ {1..n div d}" from k have "k * d \ n div d * d" by (intro mult_right_mono) simp_all also have "n div d * d \ n div d * d + n mod d" by (rule le_add1) also have "\ = n" by simp finally have "k * d \ n" . thus "k * d \ {1..n}" using d k by auto qed auto also have "\ = (\k. k*d) ` {1..n div d}" by fast also have "card \ = card {1..n div d}" by (rule card_image) (simp add: inj_on_def) also have "\ = n div d" by simp also have "... = \n / d\" by (simp add: floor_divide_of_nat_eq) finally have "real (card {x. x \ {1..n} \ d dvd x}) = real_of_int \n / d\" by force } finally show "(\da = 1..n. if primepow d \ d dvd da then ln (real (aprimedivisor d)) else 0) = mangoldt d * real_of_int \real n / real d\" . qed simp_all finally have "(\d=1..n. \da=1..d. if primepow da \ da dvd d then ln (aprimedivisor da) else 0) = sum (\d. mangoldt d * floor (n/d)) {1..n}" . with ln_primefact have "(\d=1..n. ln d) = (\d=1..n. mangoldt d * floor (n/d))" by simp with ln_fact show ?thesis by simp qed end context begin private lemma div_2_mult_2_bds: fixes n d :: nat assumes "d > 0" shows "0 \ \n / d\ - 2 * \(n div 2) / d\" "\n / d\ - 2 * \(n div 2) / d\ \ 1" proof - have "\2::real\ * \(n div 2) / d\ \ \2 * ((n div 2) / d)\" by (rule le_mult_floor) simp_all also from assms have "\ \ \n / d\" by (intro floor_mono) (simp_all add: field_simps) finally show "0 \ \n / d\ - 2 * \(n div 2) / d\" by (simp add: algebra_simps) next have "real (n div d) \ real (2 * ((n div 2) div d) + 1)" by (subst div_mult2_eq [symmetric], simp only: mult.commute, subst div_mult2_eq) simp thus "\n / d\ - 2 * \(n div 2) / d\ \ 1" unfolding of_nat_add of_nat_mult floor_conv_div_nat [symmetric] by simp_all qed private lemma n_div_d_eq_1: "d \ {n div 2 + 1..n} \ \real n / real d\ = 1" by (cases "n = d") (auto simp: field_simps intro: floor_eq) lemma psi_bounds_ln_fact: shows "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n" "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" proof - fix n::nat let ?k = "n div 2" and ?d = "n mod 2" have *: "\?k / d\ = 0" if "d > ?k" for d proof - from that div_less have "0 = ?k div d" by simp also have "\ = \?k / d\" by (rule floor_divide_of_nat_eq [symmetric]) finally show "\?k / d\ = 0" by simp qed have sum_eq: "(\d=1..2*?k+?d. mangoldt d * \?k / d\) = (\d=1..?k. mangoldt d * \?k / d\)" by (intro sum.mono_neutral_right) (auto simp: *) from ln_fact_conv_mangoldt have "ln (fact n) = (\d=1..n. mangoldt d * \n / d\)" . also have "\ = (\d=1..n. mangoldt d * \(2 * (n div 2) + n mod 2) / d\)" by simp also have "\ \ (\d=1..n. mangoldt d * (2 * \?k / d\ + 1))" using div_2_mult_2_bds(2)[of _ n] by (intro sum_mono mult_left_mono, subst of_int_le_iff) (auto simp: algebra_simps mangoldt_nonneg) also have "\ = 2 * (\d=1..n. mangoldt d * \(n div 2) / d\) + (\d=1..n. mangoldt d)" by (simp add: algebra_simps sum.distrib sum_distrib_left) also have "\ = 2 * (\d=1..2*?k+?d. mangoldt d * \(n div 2) / d\) + (\d=1..n. mangoldt d)" by presburger also from sum_eq have "\ = 2 * (\d=1..?k. mangoldt d * \(n div 2) / d\) + (\d=1..n. mangoldt d)" by presburger also from ln_fact_conv_mangoldt psi_def have "\ = 2 * ln (fact ?k) + psi n" by presburger finally show "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n" by simp next fix n::nat let ?k = "n div 2" and ?d = "n mod 2" from psi_def have "psi n - psi ?k = (\d=1..2*?k+?d. mangoldt d) - (\d=1..?k. mangoldt d)" by presburger also have "\ = sum mangoldt ({1..2 * (n div 2) + n mod 2} - {1..n div 2})" by (subst sum_diff) simp_all also have "\ = (\d\({1..2 * (n div 2) + n mod 2} - {1..n div 2}). (if d \ ?k then 0 else mangoldt d))" by (intro sum.cong) simp_all also have "\ = (\d=1..2*?k+?d. (if d \ ?k then 0 else mangoldt d))" by (intro sum.mono_neutral_left) auto also have "\ = (\d=1..n. (if d \ ?k then 0 else mangoldt d))" by presburger also have "\ = (\d=1..n. (if d \ ?k then mangoldt d * 0 else mangoldt d))" by (intro sum.cong) simp_all also from div_2_mult_2_bds(1) have "\ \ (\d=1..n. (if d \ ?k then mangoldt d * (\n/d\ - 2 * \?k/d\) else mangoldt d))" by (intro sum_mono) (auto simp: algebra_simps mangoldt_nonneg intro!: mult_left_mono simp del: of_int_mult) also from n_div_d_eq_1 have "\ = (\d=1..n. (if d \ ?k then mangoldt d * (\n/d\ - 2 * \?k/d\) else mangoldt d * \n/d\))" by (intro sum.cong refl) auto also have "\ = (\d=1..n. mangoldt d * real_of_int (\real n / real d\) - (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0))" by (intro sum.cong refl) (auto simp: algebra_simps) also have "\ = (\d=1..n. mangoldt d * real_of_int (\real n / real d\)) - (\d=1..n. (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0))" by (rule sum_subtractf) also have "(\d=1..n. (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0)) = (\d=1..?k. (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0))" by (intro sum.mono_neutral_right) auto also have "\ = (\d=1..?k. 2 * mangoldt d * real_of_int \real ?k / real d\)" by (intro sum.cong) simp_all also have "\ = 2 * (\d=1..?k. mangoldt d * real_of_int \real ?k / real d\)" by (simp add: sum_distrib_left mult_ac) also have "(\d = 1..n. mangoldt d * real_of_int \real n / real d\) - \ = ln (fact n) - 2 * ln (fact (n div 2))" by (simp add: ln_fact_conv_mangoldt) finally show "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" . qed end lemma psi_bounds_induct: "real n * ln 2 - (4 * ln (real (if n = 0 then 1 else n)) + 3) \ psi n" "psi n - psi (n div 2) \ real n * ln 2 + (4 * ln (real (if n = 0 then 1 else n)) + 3)" proof - from le_imp_neg_le[OF ln_fact_diff_bounds] have "n * ln 2 - (4 * ln (if n = 0 then 1 else n) + 3) \ n * ln 2 - abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2)" by simp also have "\ \ ln (fact n) - 2 * ln (fact (n div 2))" by simp also from psi_bounds_ln_fact (1) have "\ \ psi n" by simp finally show "real n * ln 2 - (4 * ln (real (if n = 0 then 1 else n)) + 3) \ psi n" . next from psi_bounds_ln_fact (2) have "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" . also have "\ \ n * ln 2 + abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2)" by simp also from ln_fact_diff_bounds [of n] have "abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2) \ (4 * ln (real (if n = 0 then 1 else n)) + 3)" by simp finally show "psi n - psi (n div 2) \ real n * ln 2 + (4 * ln (real (if n = 0 then 1 else n)) + 3)" by simp qed subsection \Bounding the psi function\ text \ In this section, we will first prove the relatively tight estimate @{prop "psi n \ 3 / 2 + ln 2 * n"} for @{term "n \ 128"} and then use the recurrence we have just derived to extend it to @{prop "psi n \ 551 / 256"} for @{term "n \ 1024"}, at which point applying the recurrence can be used to prove the same bound for arbitrarily big numbers. First of all, we will prove the bound for @{term "n <= 128"} using reflection and approximation. \ context begin private lemma Ball_insertD: assumes "\x\insert y A. P x" shows "P y" "\x\A. P x" using assms by auto private lemma meta_eq_TrueE: "PROP A \ Trueprop True \ PROP A" by simp private lemma pre_mangoldt_pos: "pre_mangoldt n > 0" unfolding pre_mangoldt_def by (auto simp: primepow_gt_Suc_0) private lemma psi_conv_pre_mangoldt: "psi n = ln (real (prod pre_mangoldt {1..n}))" by (auto simp: psi_def mangoldt_def pre_mangoldt_def ln_prod primepow_gt_Suc_0 intro!: sum.cong) private lemma eval_psi_aux1: "psi 0 = ln (real (numeral Num.One))" by (simp add: psi_def) private lemma eval_psi_aux2: assumes "psi m = ln (real (numeral x))" "pre_mangoldt n = y" "m + 1 = n" "numeral x * y = z" shows "psi n = ln (real z)" proof - from assms(2) [symmetric] have [simp]: "y > 0" by (simp add: pre_mangoldt_pos) have "psi n = psi (Suc m)" by (simp add: assms(3) [symmetric]) also have "\ = ln (real y * (\x = Suc 0..m. real (pre_mangoldt x)))" using assms(2,3) [symmetric] by (simp add: psi_conv_pre_mangoldt prod.nat_ivl_Suc' mult_ac) also have "\ = ln (real y) + psi m" by (subst ln_mult) (simp_all add: pre_mangoldt_pos prod_pos psi_conv_pre_mangoldt) also have "psi m = ln (real (numeral x))" by fact also have "ln (real y) + \ = ln (real (numeral x * y))" by (simp add: ln_mult) finally show ?thesis by (simp add: assms(4) [symmetric]) qed private lemma Ball_atLeast0AtMost_doubleton: assumes "psi 0 \ 3 / 2 * ln 2 * real 0" assumes "psi 1 \ 3 / 2 * ln 2 * real 1" shows "(\x\{0..1}. psi x \ 3 / 2 * ln 2 * real x)" using assms unfolding One_nat_def atLeast0_atMost_Suc ball_simps by auto private lemma Ball_atLeast0AtMost_insert: assumes "(\x\{0..m}. psi x \ 3 / 2 * ln 2 * real x)" assumes "psi (numeral n) \ 3 / 2 * ln 2 * real (numeral n)" "m = pred_numeral n" shows "(\x\{0..numeral n}. psi x \ 3 / 2 * ln 2 * real x)" using assms by (subst numeral_eq_Suc[of n], subst atLeast0_atMost_Suc, subst ball_simps, simp only: numeral_eq_Suc [symmetric]) private lemma eval_psi_ineq_aux: assumes "psi n = x" "x \ 3 / 2 * ln 2 * n" shows "psi n \ 3 / 2 * ln 2 * n" using assms by simp_all private lemma eval_psi_ineq_aux2: assumes "numeral m ^ 2 \ (2::nat) ^ (3 * n)" shows "ln (real (numeral m)) \ 3 / 2 * ln 2 * real n" proof - have "ln (real (numeral m)) \ 3 / 2 * ln 2 * real n \ 2 * log 2 (real (numeral m)) \ 3 * real n" by (simp add: field_simps log_def) also have "2 * log 2 (real (numeral m)) = log 2 (real (numeral m ^ 2))" by (subst of_nat_power, subst log_nat_power) simp_all also have "\ \ 3 * real n \ real ((numeral m) ^ 2) \ 2 powr real (3 * n)" by (subst Transcendental.log_le_iff) simp_all also have "2 powr (3 * n) = real (2 ^ (3 * n))" by (simp add: powr_realpow [symmetric]) also have "real ((numeral m) ^ 2) \ \ \ numeral m ^ 2 \ (2::nat) ^ (3 * n)" by (rule of_nat_le_iff) finally show ?thesis using assms by blast qed private lemma eval_psi_ineq_aux_mono: assumes "psi n = x" "psi m = x" "psi n \ 3 / 2 * ln 2 * n" "n \ m" shows "psi m \ 3 / 2 * ln 2 * m" proof - from assms have "psi m = psi n" by simp also have "\ \ 3 / 2 * ln 2 * n" by fact also from \n \ m\ have "\ \ 3 / 2 * ln 2 * m" by simp finally show ?thesis . qed lemma not_primepow_1_nat: "\primepow (1 :: nat)" by auto ML_file \bertrand.ML\ (* This should not take more than 1 minute *) local_setup \fn ctxt => let fun tac {context = ctxt, ...} = let val psi_cache = Bertrand.prove_psi ctxt 129 fun prove_psi_ineqs ctxt = let fun tac {context = ctxt, ...} = HEADGOAL (resolve_tac ctxt @{thms eval_psi_ineq_aux2} THEN' Simplifier.simp_tac ctxt) fun prove_by_approx n thm = let val thm = thm RS @{thm eval_psi_ineq_aux} val [prem] = Thm.prems_of thm val prem = Goal.prove ctxt [] [] prem tac in prem RS thm end fun prove_by_mono last_thm last_thm' thm = let val thm = @{thm eval_psi_ineq_aux_mono} OF [last_thm, thm, last_thm'] val [prem] = Thm.prems_of thm val prem = Goal.prove ctxt [] [] prem (K (HEADGOAL (Simplifier.simp_tac ctxt))) in prem RS thm end fun go _ acc [] = acc | go last acc ((n, x, thm) :: xs) = let val thm' = case last of NONE => prove_by_approx n thm | SOME (last_x, last_thm, last_thm') => if last_x = x then prove_by_mono last_thm last_thm' thm else prove_by_approx n thm in go (SOME (x, thm, thm')) (thm' :: acc) xs end in rev o go NONE [] end val psi_ineqs = prove_psi_ineqs ctxt psi_cache fun prove_ball ctxt (thm1 :: thm2 :: thms) = let val thm = @{thm Ball_atLeast0AtMost_doubleton} OF [thm1, thm2] fun solve_prem thm = let fun tac {context = ctxt, ...} = HEADGOAL (Simplifier.simp_tac ctxt) val thm' = Goal.prove ctxt [] [] (Thm.cprem_of thm 1 |> Thm.term_of) tac in thm' RS thm end fun go thm thm' = (@{thm Ball_atLeast0AtMost_insert} OF [thm', thm]) |> solve_prem in fold go thms thm end | prove_ball _ _ = raise Match in HEADGOAL (resolve_tac ctxt [prove_ball ctxt psi_ineqs]) end val thm = Goal.prove @{context} [] [] @{prop "\n\{0..128}. psi n \ 3 / 2 * ln 2 * n"} tac in Local_Theory.note ((@{binding psi_ubound_log_128}, []), [thm]) ctxt |> snd end \ end context begin private lemma psi_ubound_aux: defines "f \ \x::real. (4 * ln x + 3) / (ln 2 * x)" assumes "x \ 2" "x \ y" shows "f x \ f y" using assms(3) proof (rule DERIV_nonpos_imp_nonincreasing, goal_cases) case (1 t) define f' where "f' = (\x. (1 - 4 * ln x) / x^2 / ln 2 :: real)" from 1 assms(2) have "(f has_real_derivative f' t) (at t)" unfolding f_def f'_def by (auto intro!: derivative_eq_intros simp: field_simps power2_eq_square) moreover { from ln_2_ge have "1/4 \ ln (2::real)" by simp also from assms(2) 1 have "\ \ ln t" by simp finally have "ln t \ 1/4" . } with 1 assms(2) have "f' t \ 0" by (simp add: f'_def field_simps) ultimately show ?case by (intro exI[of _ "f' t"]) simp_all qed text \ These next rules are used in combination with @{thm psi_bounds_induct} and @{thm psi_ubound_log_128} to extend the upper bound for @{term "psi"} from values no greater than 128 to values no greater than 1024. The constant factor of the upper bound changes every time, but once we have reached 1024, the recurrence is self-sustaining in the sense that we do not have to adjust the constant factor anymore in order to double the range. \ lemma psi_ubound_log_double_cases': assumes "\n. n \ m \ psi n \ c * ln 2 * real n" "n \ m'" "m' = 2*m" "c \ c'" "c \ 0" "m \ 1" "c' \ 1 + c/2 + (4 * ln (m+1) + 3) / (ln 2 * (m+1))" shows "psi n \ c' * ln 2 * real n" proof (cases "n > m") case False hence "psi n \ c * ln 2 * real n" by (intro assms) simp_all also have "c \ c'" by fact finally show ?thesis by - (simp_all add: mult_right_mono) next case True hence n: "n \ m+1" by simp from psi_bounds_induct(2)[of n] True have "psi n \ real n * ln 2 + 4 * ln (real n) + 3 + psi (n div 2)" by simp also from assms have "psi (n div 2) \ c * ln 2 * real (n div 2)" by (intro assms) simp_all also have "real (n div 2) \ real n / 2" by simp also have "c * ln 2 * \ = c / 2 * ln 2 * real n" by simp also have "real n * ln 2 + 4 * ln (real n) + 3 + \ = (1 + c/2) * ln 2 * real n + (4 * ln (real n) + 3)" by (simp add: field_simps) also { have "(4 * ln (real n) + 3) / (ln 2 * (real n)) \ (4 * ln (m+1) + 3) / (ln 2 * (m+1))" using n assms by (intro psi_ubound_aux) simp_all also from assms have "(4 * ln (m+1) + 3) / (ln 2 * (m+1)) \ c' - 1 - c/2" by (simp add: algebra_simps) finally have "4 * ln (real n) + 3 \ (c' - 1 - c/2) * ln 2 * real n" using n by (simp add: field_simps) } also have "(1 + c / 2) * ln 2 * real n + (c' - 1 - c / 2) * ln 2 * real n = c' * ln 2 * real n" by (simp add: field_simps) finally show ?thesis using \c \ 0\ by (simp_all add: mult_left_mono) qed end lemma psi_ubound_log_double_cases: assumes "\n\m. psi n \ c * ln 2 * real n" "c' \ 1 + c/2 + (4 * ln (m+1) + 3) / (ln 2 * (m+1))" "m' = 2*m" "c \ c'" "c \ 0" "m \ 1" shows "\n\m'. psi n \ c' * ln 2 * real n" using assms(1) by (intro allI impI assms psi_ubound_log_double_cases'[of m c _ m' c']) auto lemma psi_ubound_log_1024: "\n\1024. psi n \ 551 / 256 * ln 2 * real n" proof - from psi_ubound_log_128 have "\n\128. psi n \ 3 / 2 * ln 2 * real n" by simp hence "\n\256. psi n \ 1025 / 512 * ln 2 * real n" proof (rule psi_ubound_log_double_cases, goal_cases) case 1 have "Some (Float 624 (- 7)) = ub_ln 9 129" by code_simp from ub_ln(1)[OF this] and ln_2_ge show ?case by (simp add: field_simps) qed simp_all hence "\n\512. psi n \ 549 / 256 * ln 2 * real n" proof (rule psi_ubound_log_double_cases, goal_cases) case 1 have "Some (Float 180 (- 5)) = ub_ln 7 257" by code_simp from ub_ln(1)[OF this] and ln_2_ge show ?case by (simp add: field_simps) qed simp_all thus "\n\1024. psi n \ 551 / 256 * ln 2 * real n" proof (rule psi_ubound_log_double_cases, goal_cases) case 1 have "Some (Float 203 (- 5)) = ub_ln 7 513" by code_simp from ub_ln(1)[OF this] and ln_2_ge show ?case by (simp add: field_simps) qed simp_all qed lemma psi_bounds_sustained_induct: assumes "4 * ln (1 + 2 ^ j) + 3 \ d * ln 2 * (1 + 2^j)" assumes "4 / (1 + 2^j) \ d * ln 2" assumes "0 \ c" assumes "c / 2 + d + 1 \ c" assumes "j \ k" assumes "\n. n \ 2^k \ psi n \ c * ln 2 * n" assumes "n \ 2^(Suc k)" shows "psi n \ c * ln 2 * n" proof (cases "n \ 2^k") case True with assms(6) show ?thesis . next case False from psi_bounds_induct(2) have "psi n - psi (n div 2) \ real n * ln 2 + (4 * ln (real (if n = 0 then 1 else n)) + 3)" . also from False have "(if n = 0 then 1 else n) = n" by simp finally have "psi n \ real n * ln 2 + (4 * ln (real n) + 3) + psi (n div 2)" by simp also from assms(6,7) have "psi (n div 2) \ c * ln 2 * (n div 2)" by simp also have "real (n div 2) \ real n / 2" by simp also have "real n * ln 2 + (4 * ln (real n) + 3) + c * ln 2 * (n / 2) \ c * ln 2 * real n" proof (rule overpower_lemma[of "\x. x * ln 2 + (4 * ln x + 3) + c * ln 2 * (x / 2)" "1+2^j" "\x. c * ln 2 * x" "\x. c * ln 2 - ln 2 - 4 / x - c / 2 * ln 2" "real n"]) from assms(1) have "4 * ln (1 + 2^j) + 3 \ d * ln 2 * (1 + 2^j)" . also from assms(4) have "d \ c - c/2 - 1" by simp also have "(\) * ln 2 * (1 + 2 ^ j) = c * ln 2 * (1 + 2 ^ j) - c / 2 * ln 2 * (1 + 2 ^ j) - (1 + 2 ^ j) * ln 2" by (simp add: left_diff_distrib) finally have "4 * ln (1 + 2^j) + 3 \ c * ln 2 * (1 + 2 ^ j) - c / 2 * ln 2 * (1 + 2 ^ j) - (1 + 2 ^ j) * ln 2" by (simp add: add_pos_pos) then show "(1 + 2 ^ j) * ln 2 + (4 * ln (1 + 2 ^ j) + 3) + c * ln 2 * ((1 + 2 ^ j) / 2) \ c * ln 2 * (1 + 2 ^ j)" by simp next fix x::real assume x: "1 + 2^j \ x" moreover have "1 + 2 ^ j > (0::real)" by (simp add: add_pos_pos) ultimately have x_pos: "x > 0" by linarith show "((\x. c * ln 2 * x - (x * ln 2 + (4 * ln x + 3) + c * ln 2 * (x / 2))) has_real_derivative c * ln 2 - ln 2 - 4 / x - c / 2 * ln 2) (at x)" by (rule derivative_eq_intros refl | simp add: \0 < x\)+ from \0 < x\ \0 < 1 + 2^j\ have "0 < x * (1 + 2^j)" by (rule mult_pos_pos) have "4 / x \ 4 / (1 + 2^j)" by (intro divide_left_mono mult_pos_pos add_pos_pos x x_pos) simp_all also from assms(2) have "4 / (1 + 2^j) \ d * ln 2" . also from assms(4) have "d \ c - c/2 - 1" by simp also have "\ * ln 2 = c * ln 2 - c/2 * ln 2 - ln 2" by (simp add: algebra_simps) finally show "0 \ c * ln 2 - ln 2 - 4 / x - c / 2 * ln 2" by simp next have "1 + 2^j = real (1 + 2^j)" by simp also from assms(5) have "\ \ real (1 + 2^k)" by simp also from False have "2^k \ n - 1" by simp finally show "1 + 2^j \ real n" using False by simp qed finally show ?thesis using assms by - (simp_all add: mult_left_mono) qed lemma psi_bounds_sustained: assumes "\n. n \ 2^k \ psi n \ c * ln 2 * n" assumes "4 * ln (1 + 2^k) + 3 \ (c/2 - 1) * ln 2 * (1 + 2^k)" assumes "4 / (1 + 2^k) \ (c/2 - 1) * ln 2" assumes "c \ 0" shows "psi n \ c * ln 2 * n" proof - have "psi n \ c * ln 2 * n" if "n \ 2^j" for j n using that proof (induction j arbitrary: n) case 0 with assms(4) 0 show ?case unfolding psi_def mangoldt_def by (cases n) auto next case (Suc j) show ?case proof (cases "k \ j") case True from assms(4) have c_div_2: "c/2 + (c/2 - 1) + 1 \ c" by simp from psi_bounds_sustained_induct[of k "c/2 -1" c j, OF assms(2) assms(3) assms(4) c_div_2 True Suc.IH Suc.prems] show ?thesis by simp next case False then have j_lt_k: "Suc j \ k" by simp from Suc.prems have "n \ 2 ^ Suc j" . also have "(2::nat) ^ Suc j \ 2 ^ k" using power_increasing[of "Suc j" k "2::nat", OF j_lt_k] by simp finally show ?thesis using assms(1) by simp qed qed from less_exp this [of n n] show ?thesis by simp qed lemma psi_ubound_log: "psi n \ 551 / 256 * ln 2 * n" proof (rule psi_bounds_sustained) show "0 \ 551 / (256 :: real)" by simp next fix n :: nat assume "n \ 2 ^ 10" with psi_ubound_log_1024 show "psi n \ 551 / 256 * ln 2 * real n" by auto next have "4 / (1 + 2 ^ 10) \ (551 / 256 / 2 - 1) * (2/3 :: real)" by simp also have "\ \ (551 / 256 / 2 - 1) * ln 2" by (intro mult_left_mono ln_2_ge') simp_all finally show "4 / (1 + 2 ^ 10) \ (551 / 256 / 2 - 1) * ln (2 :: real)" . next have "Some (Float 16 (-1)) = ub_ln 3 1025" by code_simp from ub_ln(1)[OF this] and ln_2_ge have "2048 * ln 1025 + 1536 \ 39975 * (ln 2::real)" by simp thus "4 * ln (1 + 2 ^ 10) + 3 \ (551 / 256 / 2 - 1) * ln 2 * (1 + 2 ^ 10 :: real)" by simp qed lemma psi_ubound_3_2: "psi n \ 3/2 * n" proof - have "(551 / 256) * ln 2 \ (551 / 256) * (16/23 :: real)" by (intro mult_left_mono ln_2_le') auto also have "\ \ 3 / 2" by simp finally have "551 / 256 * ln 2 \ 3/(2::real)" . with of_nat_0_le_iff mult_right_mono have "551 / 256 * ln 2 * n \ 3/2 * n" by blast with psi_ubound_log[of "n"] show ?thesis by linarith qed subsection \Doubling psi and theta\ lemma psi_residues_compare_2: "psi_odd_2 n \ psi_even_2 n" proof - have "psi_odd_2 n = (\d\{d. d \ {2..n} \ primepow_odd d}. mangoldt_odd d)" unfolding mangoldt_odd_def by (rule sum.mono_neutral_right) auto also have "\ = (\d\{d. d \ {2..n} \ primepow_odd d}. ln (real (aprimedivisor d)))" by (intro sum.cong refl) (simp add: mangoldt_odd_def) also have "\ \ (\d\{d. d \ {2..n} \ primepow_even d}. ln (real (aprimedivisor d)))" proof (rule sum_le_included [where i = "\y. y * aprimedivisor y"]; clarify?) fix d :: nat assume "d \ {2..n}" "primepow_odd d" note d = this then obtain p k where d': "k \ 1" "prime p" "d = p ^ (2*k+1)" by (auto simp: primepow_odd_def) from d' have "p ^ (2 * k) \ p ^ (2 * k + 1)" by (subst power_increasing_iff) (auto simp: prime_gt_Suc_0_nat) also from d d' have "\ \ n" by simp finally have "p ^ (2 * k) \ n" . moreover from d' have "p ^ (2 * k) > 1" by (intro one_less_power) (simp_all add: prime_gt_Suc_0_nat) ultimately have "p ^ (2 * k) \ {2..n}" by simp moreover from d' have "primepow_even (p ^ (2 * k))" by (auto simp: primepow_even_def) ultimately show "\y\{d \ {2..n}. primepow_even d}. y * aprimedivisor y = d \ ln (real (aprimedivisor d)) \ ln (real (aprimedivisor y))" using d' by (intro bexI[of _ "p ^ (2 * k)"]) (auto simp: aprimedivisor_prime_power aprimedivisor_primepow) qed (simp_all add: of_nat_ge_1_iff Suc_le_eq) also have "\ = (\d\{d. d \ {2..n} \ primepow_even d}. mangoldt_even d)" by (intro sum.cong refl) (simp add: mangoldt_even_def) also have "\ = psi_even_2 n" unfolding mangoldt_even_def by (rule sum.mono_neutral_left) auto finally show ?thesis . qed lemma psi_residues_compare: "psi_odd n \ psi_even n" proof - have "\ primepow_odd 1" by (simp add: primepow_odd_def) hence *: "mangoldt_odd 1 = 0" by (simp add: mangoldt_odd_def) have "\ primepow_even 1" using primepow_gt_Suc_0[OF primepow_even_imp_primepow, of 1] by auto with mangoldt_even_def have **: "mangoldt_even 1 = 0" by simp from psi_odd_def have "psi_odd n = (\d=1..n. mangoldt_odd d)" by simp also from * have "\ = psi_odd_2 n" by (cases "n \ 1") (simp_all add: eval_nat_numeral sum.atLeast_Suc_atMost) also from psi_residues_compare_2 have "\ \ psi_even_2 n" . also from ** have "\ = psi_even n" by (cases "n \ 1") (simp_all add: eval_nat_numeral sum.atLeast_Suc_atMost psi_even_def) finally show ?thesis . qed lemma primepow_iff_even_sqr: "primepow n \ primepow_even (n^2)" by (cases "n = 0") (auto simp: primepow_even_altdef aprimedivisor_primepow_power primepow_power_iff_nat prime_elem_multiplicity_power_distrib prime_aprimedivisor' prime_imp_prime_elem unit_factor_nat_def primepow_gt_0_nat dest: primepow_gt_Suc_0) lemma psi_sqrt: "psi (Discrete.sqrt n) = psi_even n" proof (induction n) case 0 with psi_def psi_even_def show ?case by simp next case (Suc n) then show ?case proof cases assume asm: "\ m. Suc n = m^2" with sqrt_Suc have sqrt_seq: "Discrete.sqrt(Suc n) = Suc(Discrete.sqrt n)" by simp from asm obtain "m" where " Suc n = m^2" by blast with sqrt_seq have "Suc(Discrete.sqrt n) = m" by simp with \Suc n = m^2\ have suc_sqrt_n_sqrt: "(Suc(Discrete.sqrt n))^2 = Suc n" by simp from sqrt_seq have "psi (Discrete.sqrt (Suc n)) = psi (Suc (Discrete.sqrt n))" by simp also from psi_def have "\ = psi (Discrete.sqrt n) + mangoldt (Suc (Discrete.sqrt n))" by simp also from Suc.IH have "psi (Discrete.sqrt n) = psi_even n" . also have "mangoldt (Suc (Discrete.sqrt n)) = mangoldt_even (Suc n)" proof (cases "primepow (Suc(Discrete.sqrt n))") case True with primepow_iff_even_sqr have True2: "primepow_even ((Suc(Discrete.sqrt n))^2)" by simp from suc_sqrt_n_sqrt have "mangoldt_even (Suc n) = mangoldt_even ((Suc(Discrete.sqrt n))^2)" by simp also from mangoldt_even_def True2 have "\ = ln (aprimedivisor ((Suc (Discrete.sqrt n))^2))" by simp also from True have "aprimedivisor ((Suc (Discrete.sqrt n))^2) = aprimedivisor (Suc (Discrete.sqrt n))" by (simp add: aprimedivisor_primepow_power) also from True have "ln (\) = mangoldt (Suc (Discrete.sqrt n))" by (simp add: mangoldt_def) finally show ?thesis .. next case False with primepow_iff_even_sqr have False2: "\ primepow_even ((Suc(Discrete.sqrt n))^2)" by simp from suc_sqrt_n_sqrt have "mangoldt_even (Suc n) = mangoldt_even ((Suc(Discrete.sqrt n))^2)" by simp also from mangoldt_even_def False2 have "\ = 0" by simp also from False have "\ = mangoldt (Suc (Discrete.sqrt n))" by (simp add: mangoldt_def) finally show ?thesis .. qed also from psi_even_def have "psi_even n + mangoldt_even (Suc n) = psi_even (Suc n)" by simp finally show ?case . next assume asm: "\(\m. Suc n = m^2)" with sqrt_Suc have sqrt_eq: "Discrete.sqrt (Suc n) = Discrete.sqrt n" by simp then have lhs: "psi (Discrete.sqrt (Suc n)) = psi (Discrete.sqrt n)" by simp have "\ primepow_even (Suc n)" proof assume "primepow_even (Suc n)" with primepow_even_def obtain "p" "k" where "1 \ k \ prime p \ Suc n = p ^ (2 * k)" by blast with power_even_eq have "Suc n = (p ^ k)^2" by simp with asm show False by blast qed with psi_even_def mangoldt_even_def have rhs: "psi_even (Suc n) = psi_even n" by simp from Suc.IH lhs rhs show ?case by simp qed qed lemma mangoldt_split: "mangoldt d = mangoldt_1 d + mangoldt_even d + mangoldt_odd d" proof (cases "primepow d") case False thus ?thesis by (auto simp: mangoldt_def mangoldt_1_def mangoldt_even_def mangoldt_odd_def dest: primepow_even_imp_primepow primepow_odd_imp_primepow) next case True thus ?thesis by (auto simp: mangoldt_def mangoldt_1_def mangoldt_even_def mangoldt_odd_def primepow_cases) qed lemma psi_split: "psi n = theta n + psi_even n + psi_odd n" by (induction n) (simp_all add: psi_def theta_def psi_even_def psi_odd_def mangoldt_1_def mangoldt_split) lemma psi_mono: "m \ n \ psi m \ psi n" unfolding psi_def by (intro sum_mono2 mangoldt_nonneg) auto lemma psi_pos: "0 \ psi n" by (auto simp: psi_def intro!: sum_nonneg mangoldt_nonneg) lemma mangoldt_odd_pos: "0 \ mangoldt_odd d" using aprimedivisor_gt_Suc_0[of d] by (auto simp: mangoldt_odd_def of_nat_le_iff[of 1, unfolded of_nat_1] Suc_le_eq intro!: ln_ge_zero dest!: primepow_odd_imp_primepow primepow_gt_Suc_0) lemma psi_odd_mono: "m \ n \ psi_odd m \ psi_odd n" using mangoldt_odd_pos sum_mono2[of "{1..n}" "{1..m}" "mangoldt_odd"] by (simp add: psi_odd_def) lemma psi_odd_pos: "0 \ psi_odd n" by (auto simp: psi_odd_def intro!: sum_nonneg mangoldt_odd_pos) lemma psi_theta: "theta n + psi (Discrete.sqrt n) \ psi n" "psi n \ theta n + 2 * psi (Discrete.sqrt n)" using psi_odd_pos[of n] psi_residues_compare[of n] psi_sqrt[of n] psi_split[of n] by simp_all context begin private lemma sum_minus_one: "(\x \ {1..y}. (- 1 :: real) ^ (x + 1)) = (if odd y then 1 else 0)" by (induction y) simp_all private lemma div_invert: fixes x y n :: nat assumes "x > 0" "y > 0" "y \ n div x" shows "x \ n div y" proof - from assms(1,3) have "y * x \ (n div x) * x" by simp also have "\ \ n" by (simp add: minus_mod_eq_div_mult[symmetric]) finally have "y * x \ n" . with assms(2) show ?thesis using div_le_mono[of "y*x" n y] by simp qed lemma sum_expand_lemma: "(\d=1..n. (-1) ^ (d + 1) * psi (n div d)) = (\d = 1..n. (if odd (n div d) then 1 else 0) * mangoldt d)" proof - have **: "x \ n" if "x \ n div y" for x y using div_le_dividend order_trans that by blast have "(\d=1..n. (-1)^(d+1) * psi (n div d)) = (\d=1..n. (-1)^(d+1) * (\e=1..n div d. mangoldt e))" by (simp add: psi_def) also have "\ = (\d = 1..n. \e = 1..n div d. (-1)^(d+1) * mangoldt e)" by (simp add: sum_distrib_left) also from ** have "\ = (\d = 1..n. \e\{y\{1..n}. y \ n div d}. (-1)^(d+1) * mangoldt e)" by (intro sum.cong) auto also have "\ = (\y = 1..n. \x | x \ {1..n} \ y \ n div x. (- 1) ^ (x + 1) * mangoldt y)" by (rule sum.swap_restrict) simp_all also have "\ = (\y = 1..n. \x | x \ {1..n} \ x \ n div y. (- 1) ^ (x + 1) * mangoldt y)" by (intro sum.cong) (auto intro: div_invert) also from ** have "\ = (\y = 1..n. \x \ {1..n div y}. (- 1) ^ (x + 1) * mangoldt y)" by (intro sum.cong) auto also have "\ = (\y = 1..n. (\x \ {1..n div y}. (- 1) ^ (x + 1)) * mangoldt y)" by (intro sum.cong) (simp_all add: sum_distrib_right) also have "\ = (\y = 1..n. (if odd (n div y) then 1 else 0) * mangoldt y)" by (intro sum.cong refl) (simp_all only: sum_minus_one) finally show ?thesis . qed private lemma floor_half_interval: fixes n d :: nat assumes "d \ 0" shows "real (n div d) - real (2 * ((n div 2) div d)) = (if odd (n div d) then 1 else 0)" proof - have "((n div 2) div d) = (n div (2 * d))" by (rule div_mult2_eq[symmetric]) also have "\ = ((n div d) div 2)" by (simp add: mult_ac div_mult2_eq) also have "real (n div d) - real (2 * \) = (if odd (n div d) then 1 else 0)" by (cases "odd (n div d)", cases "n div d = 0 ", simp_all) finally show ?thesis by simp qed lemma fact_expand_psi: "ln (fact n) - 2 * ln (fact (n div 2)) = (\d=1..n. (-1)^(d+1) * psi (n div d))" proof - have "ln (fact n) - 2 * ln (fact (n div 2)) = (\d=1..n. mangoldt d * \n / d\) - 2 * (\d=1..n div 2. mangoldt d * \(n div 2) / d\)" by (simp add: ln_fact_conv_mangoldt) also have "(\d=1..n div 2. mangoldt d * \real (n div 2) / d\) = (\d=1..n. mangoldt d * \real (n div 2) / d\)" by (rule sum.mono_neutral_left) (auto simp: floor_unique[of 0]) also have "2 * \ = (\d=1..n. mangoldt d * 2 * \real (n div 2) / d\)" by (simp add: sum_distrib_left mult_ac) also have "(\d=1..n. mangoldt d * \n / d\) - \ = (\d=1..n. (mangoldt d * \n / d\ - mangoldt d * 2 * \real (n div 2) / d\))" by (simp add: sum_subtractf) also have "\ = (\d=1..n. mangoldt d * (\n / d\ - 2 * \real (n div 2) / d\))" by (simp add: algebra_simps) also have "\ = (\d=1..n. mangoldt d * (if odd(n div d) then 1 else 0))" by (intro sum.cong refl) (simp_all add: floor_conv_div_nat [symmetric] floor_half_interval [symmetric]) also have "\ = (\d=1..n. (if odd(n div d) then 1 else 0) * mangoldt d)" by (simp add: mult_ac) also from sum_expand_lemma[symmetric] have "\ = (\d=1..n. (-1)^(d+1) * psi (n div d))" . finally show ?thesis . qed end lemma psi_expansion_cutoff: assumes "m \ p" shows "(\d=1..2*m. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*p. (-1)^(d+1) * psi (n div d))" "(\d=1..2*p+1. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*m+1. (-1)^(d+1) * psi (n div d))" using assms proof (induction m rule: inc_induct) case (step k) have "(\d = 1..2 * k. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * Suc k. (-1)^(d + 1) * psi (n div d))" by (simp add: psi_mono div_le_mono2) with step.IH(1) show "(\d = 1..2 * k. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * p. (-1)^(d + 1) * psi (n div d))" by simp from step.IH(2) have "(\d = 1..2 * p + 1. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * Suc k + 1. (-1)^(d + 1) * psi (n div d))" . also have "\ \ (\d = 1..2 * k + 1. (-1)^(d + 1) * psi (n div d))" by (simp add: psi_mono div_le_mono2) finally show "(\d = 1..2 * p + 1. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * k + 1. (-1)^(d + 1) * psi (n div d))" . qed simp_all lemma fact_psi_bound_even: assumes "even k" shows "(\d=1..k. (-1)^(d+1) * psi (n div d)) \ ln (fact n) - 2 * ln (fact (n div 2))" proof - have "(\d=1..k. (-1)^(d+1) * psi (n div d)) \ (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" proof (cases "k \ n") case True with psi_expansion_cutoff(1)[of "k div 2" "n div 2" n] have "(\d=1..2*(k div 2). (-1)^(d+1) * psi (n div d)) \ (\d = 1..2*(n div 2). (- 1) ^ (d + 1) * psi (n div d))" by simp also from assms have "2*(k div 2) = k" by simp also have "(\d = 1..2*(n div 2). (- 1) ^ (d + 1) * psi (n div d)) \ (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" proof (cases "even n") case True then show ?thesis by simp next case False from psi_pos have "(\d = 1..2*(n div 2). (- 1) ^ (d + 1) * psi (n div d)) \ (\d = 1..2*(n div 2) + 1. (- 1) ^ (d + 1) * psi (n div d))" by simp with False show ?thesis by simp qed finally show ?thesis . next case False hence *: "n div 2 \ (k-1) div 2" by simp have "(\d=1..k. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*((k-1) div 2) + 1. (-1)^(d+1) * psi (n div d))" proof (cases "k = 0") case True with psi_pos show ?thesis by simp next case False with sum.cl_ivl_Suc[of "\d. (-1)^(d+1) * psi (n div d)" 1 "k-1"] have "(\d=1..k. (-1)^(d+1) * psi (n div d)) = (\d=1..k-1. (-1)^(d+1) * psi (n div d)) + (-1)^(k+1) * psi (n div k)" by simp also from assms psi_pos have "(-1)^(k+1) * psi (n div k) \ 0" by simp also from assms False have "k-1 = 2*((k-1) div 2) + 1" by presburger finally show ?thesis by simp qed also from * psi_expansion_cutoff(2)[of "n div 2" "(k-1) div 2" n] have "\ \ (\d=1..2*(n div 2) + 1. (-1)^(d+1) * psi (n div d))" by blast also have "\ \ (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" by (cases "even n") (simp_all add: psi_def) finally show ?thesis . qed also from fact_expand_psi have "\ = ln (fact n) - 2 * ln (fact (n div 2))" .. finally show ?thesis . qed lemma fact_psi_bound_odd: assumes "odd k" shows "ln (fact n) - 2 * ln (fact (n div 2)) \ (\d=1..k. (-1)^(d+1) * psi (n div d))" proof - from fact_expand_psi have "ln (fact n) - 2 * ln (fact (n div 2)) = (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" . also have "\ \ (\d=1..k. (-1)^(d+1) * psi (n div d))" proof (cases "k \ n") case True have "(\d=1..n. (-1)^(d+1) * psi (n div d)) \ ( \d=1..2*(n div 2)+1. (-1)^(d+1) * psi (n div d))" by (cases "even n") (simp_all add: psi_pos) also from True assms psi_expansion_cutoff(2)[of "k div 2" "n div 2" n] have "\ \ (\d=1..k. (-1)^(d+1) * psi (n div d))" by simp finally show ?thesis . next case False have "(\d=1..n. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*((n+1) div 2). (-1)^(d+1) * psi (n div d))" by (cases "even n") (simp_all add: psi_def) also from False assms psi_expansion_cutoff(1)[of "(n+1) div 2" "k div 2" n] have "(\d=1..2*((n+1) div 2). (-1)^(d+1) * psi (n div d)) \ (\d=1..2*(k div 2). (-1)^(d+1) * psi (n div d))" by simp also from assms have "\ \ (\d=1..k. (-1)^(d+1) * psi (n div d))" by (auto elim: oddE simp: psi_pos) finally show ?thesis . qed finally show ?thesis . qed lemma fact_psi_bound_2_3: "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n - psi (n div 2) + psi (n div 3)" proof - show "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" by (rule psi_bounds_ln_fact (2)) next from fact_psi_bound_odd[of 3 n] have "ln (fact n) - 2 * ln (fact (n div 2)) \ (\d = 1..3. (- 1) ^ (d + 1) * psi (n div d))" by simp also have "\ = psi n - psi (n div 2) + psi (n div 3)" by (simp add: sum.atLeast_Suc_atMost numeral_2_eq_2) finally show "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n - psi (n div 2) + psi (n div 3)" . qed lemma ub_ln_1200: "ln 1200 \ 57 / (8 :: real)" proof - have "Some (Float 57 (-3)) = ub_ln 8 1200" by code_simp from ub_ln(1)[OF this] show ?thesis by simp qed lemma psi_double_lemma: assumes "n \ 1200" shows "real n / 6 \ psi n - psi (n div 2)" proof - from ln_fact_diff_bounds have "\ln (fact n) - 2 * ln (fact (n div 2)) - real n * ln 2\ \ 4 * ln (real (if n = 0 then 1 else n)) + 3" . with assms have "ln (fact n) - 2 * ln (fact (n div 2)) \ real n * ln 2 - 4 * ln (real n) - 3" by simp moreover have "real n * ln 2 - 4 * ln (real n) - 3 \ 2 / 3 * n" proof (rule overpower_lemma[of "\n. 2/3 * n" 1200]) show "2 / 3 * 1200 \ 1200 * ln 2 - 4 * ln 1200 - (3::real)" using ub_ln_1200 ln_2_ge by linarith next fix x::real assume "1200 \ x" then have "0 < x" by simp show "((\x. x * ln 2 - 4 * ln x - 3 - 2 / 3 * x) has_real_derivative ln 2 - 4 / x - 2 / 3) (at x)" by (rule derivative_eq_intros refl | simp add: \0 < x\)+ next fix x::real assume "1200 \ x" then have "12 / x \ 12 / 1200" by simp then have "0 \ 0.67 - 4 / x - 2 / 3" by simp also have "0.67 \ ln (2::real)" using ln_2_ge by simp finally show "0 \ ln 2 - 4 / x - 2 / 3" by simp next from assms show "1200 \ real n" by simp qed ultimately have "2 / 3 * real n \ ln (fact n) - 2 * ln (fact (n div 2))" by simp with psi_ubound_3_2[of "n div 3"] have "n/6 + psi (n div 3) \ ln (fact n) - 2 * ln (fact (n div 2))" by simp with fact_psi_bound_2_3[of "n"] show ?thesis by simp qed lemma theta_double_lemma: assumes "n \ 1200" shows "theta (n div 2) < theta n" proof - from psi_theta[of "n div 2"] psi_pos[of "Discrete.sqrt (n div 2)"] have theta_le_psi_n_2: "theta (n div 2) \ psi (n div 2)" by simp have "(Discrete.sqrt n * 18)^2 \ 324 * n" by simp from mult_less_cancel2[of "324" "n" "n"] assms have "324 * n < n^2" by (simp add: power2_eq_square) with \(Discrete.sqrt n * 18)^2 \ 324 * n\ have "(Discrete.sqrt n*18)^2 < n^2" by presburger with power2_less_imp_less assms have "Discrete.sqrt n * 18 < n" by blast with psi_ubound_3_2[of "Discrete.sqrt n"] have "2 * psi (Discrete.sqrt n) < n / 6" by simp with psi_theta[of "n"] have psi_lt_theta_n: "psi n - n / 6 < theta n" by simp from psi_double_lemma[OF assms(1)] have "psi (n div 2) \ psi n - n / 6" by simp with theta_le_psi_n_2 psi_lt_theta_n show ?thesis by simp qed subsection \Proof of the main result\ lemma theta_mono: "mono theta" by (auto simp: theta_def [abs_def] intro!: monoI sum_mono2) lemma theta_lessE: assumes "theta m < theta n" "m \ 1" obtains p where "p \ {m<..n}" "prime p" proof - from mono_invE[OF theta_mono assms(1)] have "m \ n" by blast hence "theta n = theta m + (\p\{m<..n}. if prime p then ln (real p) else 0)" unfolding theta_def using assms(2) by (subst sum.union_disjoint [symmetric]) (auto simp: ivl_disj_un) also note assms(1) finally have "(\p\{m<..n}. if prime p then ln (real p) else 0) \ 0" by simp - from sum.not_neutral_contains_not_neutral [OF this] guess p . + then obtain p where "p \ {m<..n}" "(if prime p then ln (real p) else 0) \ 0" + by (rule sum.not_neutral_contains_not_neutral) thus ?thesis using that[of p] by (auto intro!: exI[of _ p] split: if_splits) qed theorem bertrand: fixes n :: nat assumes "n > 1" shows "\p\{n<..<2*n}. prime p" proof cases assume n_less: "n < 600" define prime_constants where "prime_constants = {2, 3, 5, 7, 13, 23, 43, 83, 163, 317, 631::nat}" from \n > 1\ n_less have "\p \ prime_constants. n < p \ p < 2 * n" unfolding bex_simps greaterThanLessThan_iff prime_constants_def by presburger moreover have "\p\prime_constants. prime p" unfolding prime_constants_def ball_simps HOL.simp_thms by (intro conjI; pratt (silent)) ultimately show ?thesis unfolding greaterThanLessThan_def greaterThan_def lessThan_def by blast next assume n: "\(n < 600)" from n have "theta n < theta (2 * n)" using theta_double_lemma[of "2 * n"] by simp with assms obtain p where "p \ {n<..2*n}" "prime p" by (auto elim!: theta_lessE) moreover from assms have "\prime (2*n)" by (auto dest!: prime_product) with \prime p\ have "p \ 2 * n" by auto ultimately show ?thesis by auto qed subsection \Proof of Mertens' first theorem\ text \ The following proof of Mertens' first theorem was ported from John Harrison's HOL Light proof by Larry Paulson: \ lemma sum_integral_ubound_decreasing': fixes f :: "real \ real" assumes "m \ n" and der: "\x. x \ {of_nat m - 1..of_nat n} \ (g has_field_derivative f x) (at x)" and le: "\x y. \real m - 1 \ x; x \ y; y \ real n\ \ f y \ f x" shows "(\k = m..n. f (of_nat k)) \ g (of_nat n) - g (of_nat m - 1)" proof - have "(\k = m..n. f (of_nat k)) \ (\k = m..n. g (of_nat(Suc k) - 1) - g (of_nat k - 1))" proof (rule sum_mono, clarsimp) fix r assume r: "m \ r" "r \ n" hence "\z>real r - 1. z < real r \ g (real r) - g (real r - 1) = (real r - (real r - 1)) * f z" using assms by (intro MVT2) auto hence "\z\{of_nat r - 1..of_nat r}. g (real r) - g (real r - 1) = f z" by auto then obtain u::real where u: "u \ {of_nat r - 1..of_nat r}" and eq: "g r - g (of_nat r - 1) = f u" by blast have "real m \ u + 1" using r u by auto then have "f (of_nat r) \ f u" using r(2) and u by (intro le) auto then show "f (of_nat r) \ g r - g (of_nat r - 1)" by (simp add: eq) qed also have "\ \ g (of_nat n) - g (of_nat m - 1)" using \m \ n\ by (subst sum_Suc_diff) auto finally show ?thesis . qed lemma Mertens_lemma: assumes "n \ 0" shows "\(\d = 1..n. mangoldt d / real d) - ln n\ \ 4" proof - have *: "\abs(s' - nl + n) \ a; abs(s' - s) \ (k - 1) * n - a\ \ abs(s - nl) \ n * k" for s' s k nl a::real by (auto simp: algebra_simps abs_if split: if_split_asm) have le: "\(\d=1..n. mangoldt d * floor (n / d)) - n * ln n + n\ \ 1 + ln n" using ln_fact_bounds ln_fact_conv_mangoldt assms by simp have "\real n * ((\d = 1..n. mangoldt d / real d) - ln n)\ = \((\d = 1..n. real n * mangoldt d / real d) - n * ln n)\" by (simp add: algebra_simps sum_distrib_left) also have "\ \ real n * 4" proof (rule * [OF le]) have "\(\d = 1..n. mangoldt d * \n / d\) - (\d = 1..n. n * mangoldt d / d)\ = \\d = 1..n. mangoldt d * (\n / d\ - n / d)\" by (simp add: sum_subtractf algebra_simps) also have "\ \ psi n" (is "\?sm\ \ ?rhs") proof - have "-?sm = (\d = 1..n. mangoldt d * (n/d - \n/d\))" by (simp add: sum_subtractf algebra_simps) also have "\ \ (\d = 1..n. mangoldt d * 1)" by (intro sum_mono mult_left_mono mangoldt_nonneg) linarith+ finally have "-?sm \ ?rhs" by (simp add: psi_def) moreover have "?sm \ 0" using mangoldt_nonneg by (simp add: mult_le_0_iff sum_nonpos) ultimately show ?thesis by (simp add: abs_if) qed also have "\ \ 3/2 * real n" by (rule psi_ubound_3_2) also have "\\ (4 - 1) * real n - (1 + ln n)" using ln_le_minus_one [of n] assms by (simp add: divide_simps) finally show "\(\d = 1..n. mangoldt d * real_of_int \real n / real d\) - (\d = 1..n. real n * mangoldt d / real d)\ \ (4 - 1) * real n - (1 + ln n)" . qed finally have "\real n * ((\d = 1..n. mangoldt d / real d) - ln n)\ \ real n * 4" . then show ?thesis using assms mult_le_cancel_left_pos by (simp add: abs_mult) qed lemma Mertens_mangoldt_versus_ln: assumes "I \ {1..n}" shows "\(\i\I. mangoldt i / i) - (\p | prime p \ p \ I. ln p / p)\ \ 3" (is "\?lhs\ \ 3") proof (cases "n = 0") case True with assms show ?thesis by simp next case False have "finite I" using assms finite_subset by blast have "0 \ (\i\I. mangoldt i / i - (if prime i then ln i / i else 0))" using mangoldt_nonneg by (intro sum_nonneg) simp_all moreover have "\ \ (\i = 1..n. mangoldt i / i - (if prime i then ln i / i else 0))" using assms by (intro sum_mono2) (auto simp: mangoldt_nonneg) ultimately have *: "\\i\I. mangoldt i / i - (if prime i then ln i / i else 0)\ \ \\i = 1..n. mangoldt i / i - (if prime i then ln i / i else 0)\" by linarith moreover have "?lhs = (\i\I. mangoldt i / i - (if prime i then ln i / i else 0))" "(\i = 1..n. mangoldt i / i - (if prime i then ln i / i else 0)) = (\d = 1..n. mangoldt d / d) - (\p | prime p \ p \ {1..n}. ln p / p)" using sum.inter_restrict [of _ "\i. ln (real i) / i" "Collect prime", symmetric] by (force simp: sum_subtractf \finite I\ intro: sum.cong)+ ultimately have "\?lhs\ \ \(\d = 1..n. mangoldt d / d) - (\p | prime p \ p \ {1..n}. ln p / p)\" by linarith also have "\ \ 3" proof - have eq_sm: "(\i = 1..n. mangoldt i / i) = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 1}. mangoldt i / i)" proof (intro sum.mono_neutral_right ballI, goal_cases) case (3 i) hence "\primepow i" by (auto simp: primepow_def Suc_le_eq) thus ?case by (simp add: mangoldt_def) qed (auto simp: Suc_le_eq prime_gt_0_nat) have "(\i = 1..n. mangoldt i / i) - (\p | prime p \ p \ {1..n}. ln p / p) = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 2}. mangoldt i / i)" proof - have eq: "{p ^ k |p k. prime p \ p ^ k \ n \ 1 \ k} = {p ^ k |p k. prime p \ p ^ k \ n \ 2 \ k} \ {p. prime p \ p \ {1..n}}" (is "?A = ?B \ ?C") proof (intro equalityI subsetI; (elim UnE)?) fix x assume "x \ ?A" then obtain p k where "x = p ^ k" "prime p" "p ^ k \ n" "k \ 1" by auto thus "x \ ?B \ ?C" by (cases "k \ 2") (auto simp: prime_power_iff Suc_le_eq) next fix x assume "x \ ?B" then obtain p k where "x = p ^ k" "prime p" "p ^ k \ n" "k \ 1" by auto thus "x \ ?A" by (auto simp: prime_power_iff Suc_le_eq) next fix x assume "x \ ?C" then obtain p where "x = p ^ 1" "1 \ (1::nat)" "prime p" "p ^ 1 \ n" by auto thus "x \ ?A" by blast qed have eqln: "(\p | prime p \ p \ {1..n}. ln p / p) = (\p | prime p \ p \ {1..n}. mangoldt p / p)" by (rule sum.cong) auto have "(\i \ {p^k |p k. prime p \ p^k \ n \ k \ 1}. mangoldt i / i) = (\i \ {p ^ k |p k. prime p \ p ^ k \ n \ 2 \ k} \ {p. prime p \ p \ {1..n}}. mangoldt i / i)" by (subst eq) simp_all also have "\ = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 2}. mangoldt i / i) + (\p | prime p \ p \ {1..n}. mangoldt p / p)" by (intro sum.union_disjoint) (auto simp: prime_power_iff finite_nat_set_iff_bounded_le) also have "\ = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 2}. mangoldt i / i) + (\p | prime p \ p \ {1..n}. ln p / p)" by (simp only: eqln) finally show ?thesis using eq_sm by auto qed have "(\p | prime p \ p \ {1..n}. ln p / p) \ (\p | prime p \ p \ {1..n}. mangoldt p / p)" using mangoldt_nonneg by (auto intro: sum_mono) also have "\ \ (\i = Suc 0..n. mangoldt i / i)" by (intro sum_mono2) (auto simp: mangoldt_nonneg) finally have "0 \ (\i = 1..n. mangoldt i / i) - (\p | prime p \ p \ {1..n}. ln p / p)" by simp moreover have "(\i = 1..n. mangoldt i / i) - (\p | prime p \ p \ {1..n}. ln p / p) \ 3" (is "?M - ?L \ 3") proof - have *: "\q. \j\{1..n}. prime q \ 1 \ q \ q \ n \ (q ^ j = p ^ k \ mangoldt (p ^ k) / real p ^ k \ ln (real q) / real q ^ j)" if "prime p" "p ^ k \ n" "1 \ k" for p k proof - have "mangoldt (p ^ k) / real p ^ k \ ln p / p ^ k" using that by (simp add: divide_simps) moreover have "p \ n" using that self_le_power[of p k] by (simp add: prime_ge_Suc_0_nat) moreover have "k \ n" proof - have "k < 2^k" using of_nat_less_two_power of_nat_less_numeral_power_cancel_iff by blast also have "\ \ p^k" by (simp add: power_mono prime_ge_2_nat that) also have "\ \ n" by (simp add: that) finally show ?thesis by (simp add: that) qed ultimately show ?thesis using prime_ge_1_nat that by auto (use atLeastAtMost_iff in blast) qed have finite: "finite {p ^ k |p k. prime p \ p ^ k \ n \ 1 \ k}" by (rule finite_subset[of _ "{..n}"]) auto have "?M \ (\(x, k)\{p. prime p \ p \ {1..n}} \ {1..n}. ln (real x) / real x ^ k)" by (subst eq_sm, intro sum_le_included [where i = "\(p,k). p^k"]) (insert * finite, auto) also have "\ = (\p | prime p \ p \ {1..n}. (\k = 1..n. ln p / p^k))" by (subst sum.Sigma) auto also have "\ = ?L + (\p | prime p \ p \ {1..n}. (\k = 2..n. ln p / p^k))" by (simp add: comm_monoid_add_class.sum.distrib sum.atLeast_Suc_atMost numeral_2_eq_2) finally have "?M - ?L \ (\p | prime p \ p \ {1..n}. (\k = 2..n. ln p / p^k))" by (simp add: algebra_simps) also have "\ = (\p | prime p \ p \ {1..n}. ln p * (\k = 2..n. inverse p ^ k))" by (simp add: field_simps sum_distrib_left) also have "\ = (\p | prime p \ p \ {1..n}. ln p * (((inverse p)\<^sup>2 - inverse p ^ Suc n) / (1 - inverse p)))" by (intro sum.cong refl) (simp add: sum_gp) also have "\ \ (\p | prime p \ p \ {1..n}. ln p * inverse (real (p * (p - 1))))" by (intro sum_mono mult_left_mono) (auto simp: divide_simps power2_eq_square of_nat_diff mult_less_0_iff) also have "\ \ (\p = 2..n. ln p * inverse (real (p * (p - 1))))" by (rule sum_mono2) (use prime_ge_2_nat in auto) also have "\ \ (\i = 2..n. ln i / (i - 1)\<^sup>2)" unfolding divide_inverse power2_eq_square mult.assoc by (auto intro: sum_mono mult_left_mono mult_right_mono) also have "\ \ 3" proof (cases "n \ 3") case False then show ?thesis proof (cases "n \ 2") case False then show ?thesis by simp next case True then have "n = 2" using False by linarith with ln_le_minus_one [of 2] show ?thesis by simp qed next case True have "(\i = 3..n. ln (real i) / (real (i - Suc 0))\<^sup>2) \ (ln (of_nat n - 1)) - (ln (of_nat n)) - (ln (of_nat n) / (of_nat n - 1)) + 2 * ln 2" proof - have 1: "((\z. ln (z - 1) - ln z - ln z / (z - 1)) has_field_derivative ln x / (x - 1)\<^sup>2) (at x)" if x: "x \ {2..real n}" for x by (rule derivative_eq_intros | rule refl | (use x in \force simp: power2_eq_square divide_simps\))+ have 2: "ln y / (y - 1)\<^sup>2 \ ln x / (x - 1)\<^sup>2" if xy: "2 \ x" "x \ y" "y \ real n" for x y proof (cases "x = y") case False define f' :: "real \ real" where "f' = (\u. ((u - 1)\<^sup>2 / u - ln u * (2 * u - 2)) / (u - 1) ^ 4)" have f'_altdef: "f' u = inverse u * inverse ((u - 1)\<^sup>2) - 2 * ln u / (u - 1) ^ 3" if u: "u \ {x..y}" for u::real unfolding f'_def using u (* TODO ugly *) by (simp add: eval_nat_numeral divide_simps) (simp add: algebra_simps)? have deriv: "((\z. ln z / (z - 1)\<^sup>2) has_field_derivative f' u) (at u)" if u: "u \ {x..y}" for u::real unfolding f'_def by (rule derivative_eq_intros refl | (use u xy in \force simp: divide_simps\))+ hence "\z>x. z < y \ ln y / (y - 1)\<^sup>2 - ln x / (x - 1)\<^sup>2 = (y - x) * f' z" using xy and \x \ y\ by (intro MVT2) auto then obtain \::real where "x < \" "\ < y" and \: "ln y / (y - 1)\<^sup>2 - ln x / (x - 1)\<^sup>2 = (y - x) * f' \" by blast have "f' \ \ 0" proof - have "2/3 \ ln (2::real)" by (fact ln_2_ge') also have "\ \ ln \" using \x < \\ xy by auto finally have "1 \ 2 * ln \" by simp then have *: "\ \ \ * (2 * ln \)" using \x < \\ xy by auto hence "\ - 1 \ ln \ * 2 * \" by (simp add: algebra_simps) hence "1 / (\ * (\ - 1)\<^sup>2) \ ln \ * 2 / (\ - 1) ^ 3" using xy \x < \\ by (simp add: divide_simps power_eq_if) thus ?thesis using xy \x < \\ \\ < y\ by (subst f'_altdef) (auto simp: divide_simps) qed then have "(ln y / (y - 1)\<^sup>2 - ln x / (x - 1)\<^sup>2) \ 0" using \x \ y\ by (simp add: mult_le_0_iff \) then show ?thesis by simp qed simp_all show ?thesis using sum_integral_ubound_decreasing' [OF \3 \ n\, of "\z. ln(z-1) - ln z - ln z / (z - 1)" "\z. ln z / (z-1)\<^sup>2"] 1 2 \3 \ n\ by (auto simp: in_Reals_norm of_nat_diff) qed also have "\ \ 2" proof - have "ln (real n - 1) - ln n \ 0" "0 \ ln n / (real n - 1)" using \3 \ n\ by auto then have "ln (real n - 1) - ln n - ln n / (real n - 1) \ 0" by linarith with ln_2_less_1 show ?thesis by linarith qed also have "\ \ 3 - ln 2" using ln_2_less_1 by (simp add: algebra_simps) finally show ?thesis using True by (simp add: algebra_simps sum.atLeast_Suc_atMost [of 2 n]) qed finally show ?thesis . qed ultimately show ?thesis by linarith qed finally show ?thesis . qed proposition Mertens: assumes "n \ 0" shows "\(\p | prime p \ p \ n. ln p / of_nat p) - ln n\ \ 7" proof - have "\(\d = 1..n. mangoldt d / real d) - (\p | prime p \ p \ {1..n}. ln (real p) / real p)\ \ 7 - 4" using Mertens_mangoldt_versus_ln [of "{1..n}" n] by simp_all also have "{p. prime p \ p \ {1..n}} = {p. prime p \ p \ n}" using atLeastAtMost_iff prime_ge_1_nat by blast finally have "\(\d = 1..n. mangoldt d / real d) - (\p\\. ln (real p) / real p)\ \ 7 - 4" . moreover from assms have "\(\d = 1..n. mangoldt d / real d) - ln n\ \ 4" by (rule Mertens_lemma) ultimately show ?thesis by linarith qed end diff --git a/thys/DFS_Framework/Examples/Tarjan.thy b/thys/DFS_Framework/Examples/Tarjan.thy --- a/thys/DFS_Framework/Examples/Tarjan.thy +++ b/thys/DFS_Framework/Examples/Tarjan.thy @@ -1,1365 +1,1365 @@ section \Tarjan's Algorithm\ theory Tarjan imports Tarjan_LowLink begin text \We use the DFS Framework to implement Tarjan's algorithm. Note that, currently, we only provide an abstract version, and no refinement to efficient code. \ subsection \Preliminaries\ (* Though this is a general lemma about dropWhile/takeWhile, it is probably only of use for this algorithm. *) lemma tjs_union: fixes tjs u defines "dw \ dropWhile ((\) u) tjs" defines "tw \ takeWhile ((\) u) tjs" assumes "u \ set tjs" shows "set tjs = set (tl dw) \ insert u (set tw)" proof - from takeWhile_dropWhile_id have "set tjs = set (tw@dw)" by (auto simp: dw_def tw_def) hence "set tjs = set tw \ set dw" by (metis set_append) moreover from \u \ set tjs\ dropWhile_eq_Nil_conv have "dw \ []" by (auto simp: dw_def) from hd_dropWhile[OF this[unfolded dw_def]] have "hd dw = u" by (simp add: dw_def) with \dw \ []\ have "set dw = insert u (set (tl dw))" by (cases "dw") auto ultimately show ?thesis by blast qed subsection \Instantiation of the DFS-Framework\ record 'v tarjan_state = "'v state" + sccs :: "'v set set" lowlink :: "'v \ nat" tj_stack :: "'v list" type_synonym 'v tarjan_param = "('v, ('v,unit) tarjan_state_ext) parameterization" abbreviation "the_lowlink s v \ the (lowlink s v)" context timing_syntax begin notation the_lowlink ("\") end locale Tarjan_def = graph_defs G for G :: "('v, 'more) graph_rec_scheme" begin context begin interpretation timing_syntax . definition tarjan_disc :: "'v \ 'v tarjan_state \ ('v,unit) tarjan_state_ext nres" where "tarjan_disc v s = RETURN \ sccs = sccs s, lowlink = (lowlink s)(v \ \ s v), tj_stack = v#tj_stack s\" definition tj_stack_pop :: "'v list \ 'v \ ('v list \ 'v set) nres" where "tj_stack_pop tjs u = RETURN (tl (dropWhile ((\) u) tjs), insert u (set (takeWhile ((\) u) tjs)))" lemma tj_stack_pop_set: "tj_stack_pop tjs u \ SPEC (\(tjs',scc). u \ set tjs \ set tjs = set tjs' \ scc \ u \ scc)" proof - from tjs_union[of u tjs] show ?thesis unfolding tj_stack_pop_def by (refine_vcg) auto qed lemmas tj_stack_pop_set_leof_rule = weaken_SPEC[OF tj_stack_pop_set, THEN leof_lift] definition tarjan_fin :: "'v \ 'v tarjan_state \ ('v,unit) tarjan_state_ext nres" where "tarjan_fin v s = do { let ll = (if stack s = [] then lowlink s else let u = hd (stack s) in (lowlink s)(u \ min (\ s u) (\ s v))); let s' = s\ lowlink := ll \; ASSERT (v \ set (tj_stack s)); ASSERT (distinct (tj_stack s)); if \ s v = \ s v then do { ASSERT (scc_root' E s v (scc_of E v)); (tjs,scc) \ tj_stack_pop (tj_stack s) v; RETURN (state.more (s'\ tj_stack := tjs, sccs := insert scc (sccs s)\)) } else do { ASSERT (\ scc_root' E s v (scc_of E v)); RETURN (state.more s') }}" definition tarjan_back :: "'v \ 'v \ 'v tarjan_state \ ('v,unit) tarjan_state_ext nres" where "tarjan_back u v s = ( if \ s v < \ s u \ v \ set (tj_stack s) then let ul' = min (\ s u) (\ s v) in RETURN (state.more (s\ lowlink := (lowlink s)(u\ul') \)) else NOOP s)" end (* end timing syntax *) definition tarjan_params :: "'v tarjan_param" where "tarjan_params = \ on_init = RETURN \ sccs = {}, lowlink = Map.empty, tj_stack = [] \, on_new_root = tarjan_disc, on_discover = \u. tarjan_disc, on_finish = tarjan_fin, on_back_edge = tarjan_back, on_cross_edge = tarjan_back, is_break = \s. False \" schematic_goal tarjan_params_simps[simp]: "on_init tarjan_params = ?OI" "on_new_root tarjan_params = ?ONR" "on_discover tarjan_params = ?OD" "on_finish tarjan_params = ?OF" "on_back_edge tarjan_params = ?OBE" "on_cross_edge tarjan_params = ?OCE" "is_break tarjan_params = ?IB" unfolding tarjan_params_def gen_parameterization.simps by (rule refl)+ sublocale param_DFS_defs G tarjan_params . end locale Tarjan = Tarjan_def G + param_DFS G tarjan_params for G :: "('v, 'more) graph_rec_scheme" begin lemma [simp]: "sccs (empty_state \sccs = s, lowlink = l, tj_stack = t\) = s" "lowlink (empty_state \sccs = s, lowlink = l, tj_stack = t\) = l" "tj_stack (empty_state \sccs = s, lowlink = l, tj_stack = t\) = t" by (simp_all add: empty_state_def) lemma sccs_more_cong[cong]:"state.more s = state.more s' \ sccs s = sccs s'" by (cases s, cases s') simp lemma lowlink_more_cong[cong]:"state.more s = state.more s' \ lowlink s = lowlink s'" by (cases s, cases s') simp lemma tj_stack_more_cong[cong]:"state.more s = state.more s' \ tj_stack s = tj_stack s'" by (cases s, cases s') simp lemma [simp]: "s\ state.more := \sccs = sc, lowlink = l, tj_stack = t\\ = s\ sccs := sc, lowlink := l, tj_stack := t\" by (cases s) simp end locale Tarjan_invar = Tarjan + DFS_invar where param = tarjan_params context Tarjan_def begin lemma Tarjan_invar_eq[simp]: "DFS_invar G tarjan_params s \ Tarjan_invar G s" (is "?D \ ?T") proof assume ?D then interpret DFS_invar where param=tarjan_params . show ?T .. next assume ?T then interpret Tarjan_invar . show ?D .. qed end subsection \Correctness Proof\ context Tarjan begin lemma i_tj_stack_discovered: "is_invar (\s. set (tj_stack s) \ dom (discovered s))" proof (induct rule: establish_invarI) case (finish s) from finish show ?case apply simp unfolding tarjan_fin_def apply (refine_vcg tj_stack_pop_set_leof_rule) apply auto done qed (auto simp add: tarjan_disc_def tarjan_back_def) lemmas (in Tarjan_invar) tj_stack_discovered = i_tj_stack_discovered[THEN make_invar_thm] lemma i_tj_stack_distinct: "is_invar (\s. distinct (tj_stack s))" proof (induct rule: establish_invarI_ND) case (new_discover s s' v) then interpret Tarjan_invar where s=s by simp from new_discover tj_stack_discovered have "v \ set (tj_stack s)" by auto with new_discover show ?case by (simp add: tarjan_disc_def) next case (finish s) thus ?case apply simp unfolding tarjan_fin_def tj_stack_pop_def apply (refine_vcg) apply (auto intro: distinct_tl) done qed (simp_all add: tarjan_back_def) lemmas (in Tarjan_invar) tj_stack_distinct = i_tj_stack_distinct[THEN make_invar_thm] context begin interpretation timing_syntax . lemma i_tj_stack_incr_disc: "is_invar (\s. \kj s (tj_stack s ! j) > \ s (tj_stack s ! k))" proof (induct rule: establish_invarI_ND) case (new_discover s s' v) then interpret Tarjan_invar where s=s by simp from new_discover tj_stack_discovered have "v \ set (tj_stack s)" by auto moreover { fix k j assume "k < Suc (length (tj_stack s))" "j < k" hence "k - Suc 0 < length (tj_stack s)" by simp hence "tj_stack s ! (k - Suc 0) \ set (tj_stack s)" using nth_mem by metis with tj_stack_discovered timing_less_counter have "\ s (tj_stack s ! (k - Suc 0)) < counter s" by blast } moreover { fix k j define k' where "k' = k - Suc 0" define j' where "j' = j - Suc 0" assume A: "k < Suc (length (tj_stack s))" "j < k" "(v#tj_stack s) ! j \ v" hence gt_0: "j > 0 \ k>0" by (cases "j=0") simp_all moreover with \j < k\ have "j' < k'" by (simp add: j'_def k'_def) moreover from A have "k' < length (tj_stack s)" by (simp add: k'_def) ultimately have "\ s (tj_stack s ! j') > \ s (tj_stack s ! k')" using new_discover by blast with gt_0 have "\ s ((v#tj_stack s) ! j) > \ s (tj_stack s ! k')" unfolding j'_def by (simp add: nth_Cons') } ultimately show ?case using new_discover by (auto simp add: tarjan_disc_def) next case (finish s s' u) { let ?dw = "dropWhile ((\) u) (tj_stack s)" let ?tw = "takeWhile ((\) u) (tj_stack s)" fix a k j assume A: "a = tl ?dw" "k < length a" "j < k" and "u \ set (tj_stack s)" hence "?dw \ []" by auto define j' k' where "j' = Suc j + length ?tw" and "k' = Suc k + length ?tw" with \j < k\ have "j' < k'" by simp have "length (tj_stack s) = length ?tw + length ?dw" by (simp add: length_append[symmetric]) moreover from A have *: "Suc k < length ?dw" and **: "Suc j < length ?dw" by auto ultimately have "k' < length (tj_stack s)" by (simp add: k'_def) with finish \j' have "\ s (tj_stack s ! k') < \ s (tj_stack s ! j')" by simp also from dropWhile_nth[OF *] have "tj_stack s ! k' = ?dw ! Suc k" by (simp add: k'_def) also from dropWhile_nth[OF **] have "tj_stack s ! j' = ?dw ! Suc j" by (simp add: j'_def) also from nth_tl[OF \?dw \ []\] have "?dw ! Suc k = a ! k" by (simp add: A) also from nth_tl[OF \?dw \ []\] have "?dw ! Suc j = a ! j" by (simp add: A) finally have "\ s (a ! k) < \ s (a ! j)" . } note aux = this from finish show ?case apply simp unfolding tarjan_fin_def tj_stack_pop_def apply refine_vcg apply (auto intro!: aux) done qed (simp_all add: tarjan_back_def) end end context Tarjan_invar begin context begin interpretation timing_syntax . lemma tj_stack_incr_disc: assumes "k < length (tj_stack s)" and "j < k" shows "\ s (tj_stack s ! j) > \ s (tj_stack s ! k)" using assms i_tj_stack_incr_disc[THEN make_invar_thm] by blast lemma tjs_disc_dw_tw: fixes u defines "dw \ dropWhile ((\) u) (tj_stack s)" defines "tw \ takeWhile ((\) u) (tj_stack s)" assumes "x \ set dw" "y \ set tw" shows "\ s x < \ s y" proof - from assms obtain k where k: "dw ! k = x" "k < length dw" by (metis in_set_conv_nth) from assms obtain j where j: "tw ! j = y" "j < length tw" by (metis in_set_conv_nth) have "length (tj_stack s) = length tw + length dw" by (simp add: length_append[symmetric] tw_def dw_def) with k j have "\ s (tj_stack s ! (k + length tw)) < \ s (tj_stack s ! j)" by (simp add: tj_stack_incr_disc) also from j takeWhile_nth have "tj_stack s ! j = y" by (metis tw_def) also from dropWhile_nth k have "tj_stack s ! (k + length tw) = x" by (metis tw_def dw_def) finally show ?thesis . qed end end context Tarjan begin context begin interpretation timing_syntax . lemma i_sccs_finished_stack_ss_tj_stack: "is_invar (\s. \(sccs s) \ dom (finished s) \ set (stack s) \ set (tj_stack s))" proof (induct rule: establish_invarI) case (finish s s' u) then interpret Tarjan_invar where s=s by simp let ?tw = "takeWhile ((\) u) (tj_stack s)" let ?dw = "dropWhile ((\) u) (tj_stack s)" { fix x assume A: "x \ u" "x \ set ?tw" "u \ set (tj_stack s)" hence x_tj: "x \ set (tj_stack s)" by (auto dest: set_takeWhileD) have "x \ dom (finished s)" proof (rule ccontr) assume "x \ dom (finished s)" with x_tj tj_stack_discovered discovered_eq_finished_un_stack have "x \ set (stack s)" by blast with \x\u\ finish have "x \ set (tl (stack s))" by (cases "stack s") auto with tl_lt_stack_hd_discover finish have *: "\ s x < \ s u" by simp from A have "?dw \ []" by simp with hd_dropWhile[OF this] hd_in_set have "u \ set ?dw" by metis with tjs_disc_dw_tw \x \ set ?tw\ have "\ s u < \ s x" by simp with * show False by force qed hence "\y. finished s x = Some y" by blast } note aux_scc = this { fix x assume A: "x \ set (tl (stack s))" "u \ set (tj_stack s)" with finish stack_distinct have "x \ u" by (cases "stack s") auto moreover from A have "x \ set (stack s)" by (metis in_set_tlD) with stack_not_finished have "x \ dom (finished s)" by simp with A aux_scc[OF \x \ u\] have "x \ set ?tw" by blast moreover from finish \x \ set (stack s)\ have "x \ set (tj_stack s)" by auto moreover note tjs_union[OF \u \ set (tj_stack s)\] ultimately have "x \ set (tl ?dw)" by blast } note aux_tj = this from finish show ?case apply simp unfolding tarjan_fin_def tj_stack_pop_def apply (refine_vcg) using aux_scc aux_tj apply (auto dest: in_set_tlD) done qed (auto simp add: tarjan_disc_def tarjan_back_def) lemma i_tj_stack_ss_stack_finished: "is_invar (\s. set (tj_stack s) \ set (stack s) \ dom (finished s))" proof (induct rule: establish_invarI) case (finish s) thus ?case apply simp unfolding tarjan_fin_def apply (refine_vcg tj_stack_pop_set_leof_rule) apply ((simp, cases "stack s", simp_all)[])+ done qed (auto simp add: tarjan_disc_def tarjan_back_def) lemma i_finished_ss_sccs_tj_stack: "is_invar (\s. dom (finished s) \ \(sccs s) \ set (tj_stack s))" proof (induction rule: establish_invarI_ND) case (new_discover s s' v) then interpret Tarjan_invar where s=s by simp from new_discover finished_discovered have "v \ dom (finished s)" by auto with new_discover show ?case by (auto simp add: tarjan_disc_def) next case (finish s s' u) then interpret Tarjan_invar where s=s by simp from finish show ?case apply simp unfolding tarjan_fin_def apply (refine_vcg tj_stack_pop_set_leof_rule) apply auto done qed (simp_all add: tarjan_back_def) end end context Tarjan_invar begin lemmas finished_ss_sccs_tj_stack = i_finished_ss_sccs_tj_stack[THEN make_invar_thm] lemmas tj_stack_ss_stack_finished = i_tj_stack_ss_stack_finished[THEN make_invar_thm] lemma sccs_finished: "\(sccs s) \ dom (finished s)" using i_sccs_finished_stack_ss_tj_stack[THEN make_invar_thm] by blast lemma stack_ss_tj_stack: "set (stack s) \ set (tj_stack s)" using i_sccs_finished_stack_ss_tj_stack[THEN make_invar_thm] by blast lemma hd_stack_in_tj_stack: "stack s \ [] \ hd (stack s) \ set (tj_stack s)" using stack_ss_tj_stack hd_in_set by auto end context Tarjan begin context begin interpretation timing_syntax . lemma i_no_finished_root: "is_invar (\s. scc_root s r scc \ r \ dom (finished s) \ (\x \ scc. x \ set (tj_stack s)))" proof (induct rule: establish_invarI_ND_CB) case (new_discover s s' v) then interpret Tarjan_invar where s=s by simp { fix x let ?s = "s'\state.more := x\" assume TRANS: "\\. tarjan_disc v s' \\<^sub>n SPEC \ \ \ x" and inv': "DFS_invar G tarjan_params (s'\state.more := x\)" and r: "scc_root ?s r scc" "r \ dom (finished s')" from inv' interpret s': Tarjan_invar where s="?s" by simp have "tj_stack ?s = v#tj_stack s" by (rule TRANS) (simp add: new_discover tarjan_disc_def) moreover from r s'.scc_root_finished_impl_scc_finished have "scc \ dom (finished ?s)" by auto with new_discover finished_discovered have "v \ scc" by force moreover from r finished_discovered new_discover have "r \ dom (discovered s)" by auto with r inv' new_discover have "scc_root s r scc" apply (intro scc_root_transfer[where s'="?s", THEN iffD2]) apply clarsimp_all done with new_discover r have "\x \ scc. x \ set (tj_stack s')" by simp ultimately have "\x\scc. x \ set (tj_stack ?s)" by (auto simp: new_discover) } with new_discover show ?case by (simp add: pw_leof_iff) next case (cross_back_edge s s' u v) then interpret Tarjan_invar where s=s by simp { fix x let ?s = "s'\state.more := x\" assume TRANS: "\\. tarjan_back u v s' \\<^sub>n SPEC \ \ \ x" and r: "scc_root ?s r scc" "r \ dom (finished s')" with cross_back_edge have "scc_root s r scc" by (simp add: scc_root_transfer'[where s'="?s"]) moreover have "tj_stack ?s = tj_stack s" by (rule TRANS) (simp add: cross_back_edge tarjan_back_def) ultimately have "\x\scc. x \ set (tj_stack ?s)" using cross_back_edge r by simp } with cross_back_edge show ?case by (simp add: pw_leof_iff) next case (finish s s' u) then interpret Tarjan_invar where s=s by simp { fix x let ?s = "s'\state.more := x\" assume TRANS: "\\. tarjan_fin u s' \\<^sub>n SPEC \ \ \ x" and inv': "DFS_invar G tarjan_params (s'\state.more := x\)" and r: "scc_root ?s r scc" "r \ dom (finished s')" from inv' interpret s': Tarjan_invar where s="?s" by simp have "\x\scc. x \ set (tj_stack ?s)" proof (cases "r = u") case False with finish r have "\x\scc. x \ set (tj_stack s)" using scc_root_transfer'[where s'="?s"] by simp moreover have "set (tj_stack ?s) \ set (tj_stack s)" apply (rule TRANS) unfolding tarjan_fin_def apply (refine_vcg tj_stack_pop_set_leof_rule) apply (simp_all add: finish) done ultimately show ?thesis by blast next case True with r s'.scc_root_unique_is_scc have "scc_root ?s u (scc_of E u)" by simp with s'.scc_root_transfer'[where s'=s'] finish have "scc_root s' u (scc_of E u)" by simp moreover hence [simp]: "tj_stack ?s = tl (dropWhile ((\) u) (tj_stack s))" apply (rule_tac TRANS) unfolding tarjan_fin_def tj_stack_pop_def apply (refine_vcg) apply (simp_all add: finish) done { let ?dw = "dropWhile ((\) u) (tj_stack s)" let ?tw = "takeWhile ((\) u) (tj_stack s)" fix x define j::nat where "j = 0" assume x: "x \ set (tj_stack ?s)" then obtain i where i: "i < length (tj_stack ?s)" "tj_stack ?s ! i = x" by (metis in_set_conv_nth) have "length (tj_stack s) = length ?tw + length ?dw" by (simp add: length_append[symmetric]) with i have "\ s (tj_stack s ! (Suc i + length ?tw)) < \ s (tj_stack s ! length ?tw)" by (simp add: tj_stack_incr_disc) also from hd_stack_in_tj_stack finish have ne: "?dw \ []" and "length ?dw > 0" by simp_all from hd_dropWhile[OF ne] hd_conv_nth[OF ne] have "?dw ! 0 = u" by simp with dropWhile_nth[OF \length ?dw > 0\] have "tj_stack s ! length ?tw = u" by simp also from i have "?dw ! Suc i = x" "Suc i < length ?dw" by (simp_all add: nth_tl[OF ne]) with dropWhile_nth[OF this(2)] have "tj_stack s ! (Suc i + length ?tw) = x" by simp finally have "\ ?s x < \ ?s u" by (simp add: finish) moreover from x s'.tj_stack_discovered have "x \ dom (discovered ?s)" by auto ultimately have "x \ scc" using s'.scc_root_disc_le r True by force } thus ?thesis by metis qed } with finish show ?case by (simp add: pw_leof_iff) qed simp_all end end context Tarjan_invar begin lemma no_finished_root: assumes "scc_root s r scc" and "r \ dom (finished s)" and "x \ scc" shows "x \ set (tj_stack s)" using assms using i_no_finished_root[THEN make_invar_thm] by blast context begin interpretation timing_syntax . lemma tj_stack_reach_stack: assumes "u \ set (tj_stack s)" shows "\v \ set (stack s). (u,v) \ E\<^sup>* \ \ s v \ \ s u" proof - have u_scc: "u \ scc_of E u" by simp from assms tj_stack_discovered have u_disc: "u \ dom (discovered s)" by auto with scc_root_of_node_exists obtain r where r: "scc_root s r (scc_of E u)" by blast have "r \ set (stack s)" proof (rule ccontr) assume "r \ set (stack s)" with r[unfolded scc_root_def] stack_set_def have "r \ dom (finished s)" by simp with u_scc have "u \ set (tj_stack s)" using no_finished_root r by blast with assms show False by contradiction qed moreover from r scc_reach_scc_root u_scc u_disc have "(u,r) \ E\<^sup>*" by blast moreover from r scc_root_disc_le u_scc u_disc have "\ s r \ \ s u" by blast ultimately show ?thesis by metis qed lemma tj_stack_reach_hd_stack: assumes "v \ set (tj_stack s)" shows "(v, hd (stack s)) \ E\<^sup>*" proof - from tj_stack_reach_stack assms obtain r where r: "r \ set (stack s)" "(v,r) \ E\<^sup>*" by blast hence "r = hd (stack s) \ r \ set (tl (stack s))" by (cases "stack s") auto thus ?thesis proof assume "r = hd (stack s)" with r show ?thesis by simp next from r have ne :"stack s \ []" by auto assume "r \ set (tl (stack s))" with tl_stack_hd_tree_path ne have "(r,hd (stack s)) \ (tree_edges s)\<^sup>+" by simp with trancl_mono_mp tree_edges_ssE have "(r,hd (stack s))\E\<^sup>*" by (metis rtrancl_eq_or_trancl) with \(v,r)\E\<^sup>*\ show ?thesis by (metis rtrancl_trans) qed qed lemma empty_stack_imp_empty_tj_stack: assumes "stack s = []" shows "tj_stack s = []" proof (rule ccontr) assume ne: "tj_stack s \ []" then obtain x where x: "x \ set (tj_stack s)" by auto with tj_stack_reach_stack obtain r where "r \ set (stack s)" by auto with assms show False by simp qed lemma stacks_eq_iff: "stack s = [] \ tj_stack s = []" using empty_stack_imp_empty_tj_stack stack_ss_tj_stack by auto end end context Tarjan begin context begin interpretation timing_syntax . lemma i_sccs_are_sccs: "is_invar (\s. \scc \ sccs s. is_scc E scc)" proof (induction rule: establish_invarI) case (finish s s' u) then interpret Tarjan_invar where s=s by simp from finish have EQ[simp]: "finished s' = (finished s)(u \ counter s)" "discovered s' = discovered s" "tree_edges s' = tree_edges s" "sccs s' = sccs s" "tj_stack s' = tj_stack s" by simp_all { fix x let ?s = "s'\state.more := x\" assume TRANS: "\\. tarjan_fin u s' \\<^sub>n SPEC \ \ \ x" and inv': "DFS_invar G tarjan_params (s'\state.more := x\)" then interpret s': Tarjan_invar where s="?s" by simp from finish hd_in_set stack_set_def have u_disc: "u \ dom (discovered s)" and u_n_fin: "u \ dom (finished s)" by blast+ have "\scc \ sccs ?s. is_scc E scc" proof (cases "scc_root s' u (scc_of E u)") case False have "sccs ?s = sccs s" apply (rule TRANS) unfolding tarjan_fin_def tj_stack_pop_def by (refine_vcg) (simp_all add: False) thus ?thesis by (simp add: finish) next case True let ?dw = "dropWhile ((\) u) (tj_stack s)" let ?tw = "takeWhile ((\) u) (tj_stack s)" let ?tw' = "insert u (set ?tw)" have [simp]: "sccs ?s = insert ?tw' (sccs s)" apply (rule TRANS) unfolding tarjan_fin_def tj_stack_pop_def by (refine_vcg) (simp_all add: True) have [simp]: "tj_stack ?s = tl ?dw" apply (rule TRANS) unfolding tarjan_fin_def tj_stack_pop_def by (refine_vcg) (simp_all add: True) from True scc_root_transfer'[where s'=s'] have "scc_root s u (scc_of E u)" by simp with inv' scc_root_transfer[where s'="?s"] u_disc have u_root: "scc_root ?s u (scc_of E u)" by simp have "?tw' \ scc_of E u" proof fix v assume v: "v \ ?tw'" show "v \ scc_of E u" proof cases assume "v \ u" with v have v: "v \ set ?tw" by auto hence v_tj: "v \ set (tj_stack s)" by (auto dest: set_takeWhileD) with tj_stack_discovered have v_disc: "v \ dom (discovered s)" by auto from hd_stack_in_tj_stack finish have "?dw \ []" by simp with hd_dropWhile[OF this] hd_in_set have "u \ set ?dw" by metis with v have "\ s v > \ s u" using tjs_disc_dw_tw by blast moreover have "v \ dom (finished s)" proof (rule ccontr) assume "v \ dom (finished s)" with v_disc stack_set_def have "v \ set (stack s)" by auto with \v\u\ finish have "v \ set (tl (stack s))" by (cases "stack s") auto with tl_lt_stack_hd_discover finish have "\ s v < \ s u" by simp with \\ s v > \ s u\ show False by force qed ultimately have "(u,v) \ (tree_edges s)\<^sup>+" using parenthesis_impl_tree_path_not_finished[OF u_disc] u_n_fin by force with trancl_mono_mp tree_edges_ssE have "(u,v)\E\<^sup>*" by (metis rtrancl_eq_or_trancl) moreover from tj_stack_reach_hd_stack v_tj finish have "(v,u)\E\<^sup>*" by simp moreover have "is_scc E (scc_of E u)" "u \ scc_of E u" by simp_all ultimately show ?thesis using is_scc_closed by metis qed simp qed moreover have "scc_of E u \ ?tw'" proof fix v assume v: "v \ scc_of E u" moreover note u_root moreover have "u \ dom (finished ?s)" by simp ultimately have "v \ dom (finished ?s)" "v \ set (tj_stack ?s)" using s'.scc_root_finished_impl_scc_finished s'.no_finished_root by auto with s'.finished_ss_sccs_tj_stack have "v \ \(sccs ?s)" by blast hence "v \ \(sccs s) \ v \ ?tw'" by auto thus "v \ ?tw'" proof assume "v \ \(sccs s)" then obtain scc where scc: "v \ scc" "scc \ sccs s" by auto moreover with finish have "is_scc E scc" by simp moreover have "is_scc E (scc_of E u)" by simp moreover note v ultimately have "scc = scc_of E u" using is_scc_unique by metis hence "u \ scc" by simp with scc sccs_finished have "u \ dom (finished s)" by auto with u_n_fin show ?thesis by contradiction qed simp qed ultimately have "?tw' = scc_of E u" by auto hence "is_scc E ?tw'" by simp with finish show ?thesis by auto qed } thus ?case by (auto simp: pw_leof_iff finish) qed (simp_all add: tarjan_back_def tarjan_disc_def) end lemmas (in Tarjan_invar) sccs_are_sccs = i_sccs_are_sccs[THEN make_invar_thm] context begin interpretation timing_syntax . lemma i_lowlink_eq_LowLink: "is_invar (\s. \x \ dom (discovered s). \ s x = LowLink s x)" proof - { fix s s' :: "'v tarjan_state" fix v w fix x let ?s = "s'\state.more := x\" assume pre_ll_sub_rev: "\w. \Tarjan_invar G ?s; w \ dom (discovered ?s); w \ v\ \ lowlink_set ?s w \ lowlink_set s w \ {v}" assume tree_sub : "tree_edges s' = tree_edges s \ (\u. u \ v \ tree_edges s' = tree_edges s \ {(u,v)})" assume "Tarjan_invar G s" assume [simp]: "discovered s' = (discovered s)(v \ counter s)" "finished s' = finished s" "lowlink s' = lowlink s" "cross_edges s' = cross_edges s" "back_edges s' = back_edges s" assume v_n_disc: "v \ dom (discovered s)" assume IH: "\w. w\dom (discovered s) \ \ s w = LowLink s w" assume TRANS: "\\. tarjan_disc v s' \\<^sub>n SPEC \ \ \ x" and INV: "DFS_invar G tarjan_params ?s" and w_disc: "w \ dom (discovered ?s)" interpret Tarjan_invar where s=s by fact from INV interpret s':Tarjan_invar where s="?s" by simp have [simp]: "lowlink ?s = (lowlink s)(v \ counter s)" by (rule TRANS) (auto simp: tarjan_disc_def) from v_n_disc edge_imp_discovered have "edges s `` {v} = {}" by auto with tree_sub tree_edge_imp_discovered have "edges ?s `` {v} = {}" by auto with s'.no_path_imp_no_lowlink_path have "\w. \(\p. lowlink_path ?s v p w)" by metis hence ll_v: "lowlink_set ?s v = {v}" unfolding lowlink_set_def by auto have "\ ?s w = LowLink ?s w" proof (cases "w=v") case True with ll_v show ?thesis by simp next case False hence "\ ?s w = \ s w" by simp also from IH have "\ s w = LowLink s w" using w_disc False by simp also have "LowLink s w = LowLink ?s w" proof (rule LowLink_eqI[OF INV]) from v_n_disc show "discovered s \\<^sub>m discovered ?s" by (simp add: map_le_def) from tree_sub show "lowlink_set s w \ lowlink_set ?s w" unfolding lowlink_set_def lowlink_path_def by auto show "lowlink_set ?s w \ lowlink_set s w \ {v}" proof (cases "w = v") case True with ll_v show ?thesis by auto next case False thus ?thesis using pre_ll_sub_rev w_disc INV by simp qed show "w \ dom (discovered s)" using w_disc False by simp fix ll assume "ll \ {v}" with timing_less_counter lowlink_set_discovered have "\x. x\\ s`lowlink_set s w \ x < \ ?s ll" by simp force moreover from Min_in lowlink_set_finite lowlink_set_not_empty w_disc False have "LowLink s w \ \ s`lowlink_set s w " by auto ultimately show "LowLink s w \ \ ?s ll" by force qed finally show ?thesis . qed } note tarjan_disc_aux = this show ?thesis proof (induct rule: establish_invarI_CB) case (new_root s s' v0) { fix w x let ?s = "new_root v0 s\state.more := x\" have "lowlink_set ?s w \ lowlink_set s w \ {v0}" unfolding lowlink_set_def lowlink_path_def by auto } note * = this from new_root show ?case using tarjan_disc_aux[OF *] by (auto simp add: pw_leof_iff) next case (discover s s' u v) then interpret Tarjan_invar where s=s by simp let ?s' = "discover (hd (stack s)) v (s\pending := pending s - {(hd (stack s),v)}\)" { fix w x let ?s = "?s'\state.more := x\" assume INV: "Tarjan_invar G ?s" and d: "w \ dom (discovered ?s')" and "w\v" interpret s': Tarjan_invar where s="?s" by fact have "lowlink_set ?s w \ lowlink_set s w \ {v}" proof fix ll assume ll: "ll \ lowlink_set ?s w" hence "ll = w \ (\p. lowlink_path ?s w p ll)" by (auto simp add: lowlink_set_def) thus "ll \ lowlink_set s w \ {v}" (is "ll \ ?L") proof assume "ll = w" with d show ?thesis by (auto simp add: lowlink_set_def) next assume "\p. lowlink_path ?s w p ll" - then guess p .. note p = this + then obtain p where p: "lowlink_path ?s w p ll" .. hence [simp]: "p\[]" by (simp add: lowlink_path_def) from p have "hd p = w" by (auto simp add: lowlink_path_def path_hd) show ?thesis proof (rule tri_caseE) assume "v\ll" "v \ set p" hence "lowlink_path s w p ll" using p by (auto simp add: lowlink_path_def) with ll show ?thesis by (auto simp add: lowlink_set_def) next assume "v = ll" thus ?thesis by simp next assume "v \ set p" "v \ ll" then obtain i where i: "i < length p" "p!i = v" by (metis in_set_conv_nth) have "False" proof (cases i) case "0" with i have "hd p = v" by (simp add: hd_conv_nth) with \hd p = w\ \w \ v\ show False by simp next case (Suc n) with i s'.lowlink_path_finished[OF p, where j=i] have "v \ dom (finished ?s)" by simp with finished_discovered discover show False by auto qed thus ?thesis .. qed qed qed } note * = this from discover hd_in_set stack_set_def have "v \ u" by auto with discover have **: "tree_edges ?s' = tree_edges s \ (\u. u \ v \ tree_edges ?s' = tree_edges s \ {(u,v)})" by auto from discover show ?case using tarjan_disc_aux[OF * **] by (auto simp: pw_leof_iff) next case (cross_back_edge s s' u v) then interpret Tarjan_invar where s=s by simp from cross_back_edge have [simp]: "discovered s' = discovered s" "finished s' = finished s" "tree_edges s' = tree_edges s" "lowlink s' = lowlink s" by simp_all { fix w :: "'v" fix x let ?s = "s'\state.more := x\" let ?L = "\ s ` lowlink_set s w" let ?L' = "\ ?s ` lowlink_set ?s w" assume TRANS: "\\. tarjan_back u v s' \\<^sub>n SPEC \ \ \ x" and inv': "DFS_invar G tarjan_params ?s" and w_disc': "w \ dom (discovered ?s)" from inv' interpret s':Tarjan_invar where s="?s" by simp have ll_sub: "lowlink_set s w \ lowlink_set ?s w" unfolding lowlink_set_def lowlink_path_def by (auto simp: cross_back_edge) have ll_sub_rev: "lowlink_set ?s w \ lowlink_set s w \ {v}" unfolding lowlink_set_def lowlink_path_def by (auto simp: cross_back_edge) from w_disc' have w_disc: "w \ dom (discovered s)" by simp with LowLink_le_disc have LLw: "LowLink s w \ \ s w" by simp from cross_back_edge hd_in_set have u_n_fin: "u \ dom (finished s)" using stack_not_finished by auto { assume *: "v \ lowlink_set ?s w \ LowLink s w \ \ ?s v" have "LowLink s w = LowLink ?s w" proof (rule LowLink_eqI[OF inv' _ ll_sub ll_sub_rev w_disc]) show "discovered s \\<^sub>m discovered ?s" by simp fix ll assume "ll \ {v}" "ll \ lowlink_set ?s w" with * show "LowLink s w \ \ ?s ll" by simp qed } note LL_eqI = this have "\ ?s w = LowLink ?s w" proof (cases "w=u") case True show ?thesis proof (cases "(\ s v < \ s w \ v \ set (tj_stack s) \ \ s v < \ s w)") case False note all_False = this with \w = u\ have "\ ?s w = \ s w" by (rule_tac TRANS) (auto simp add: tarjan_back_def cross_back_edge) also from cross_back_edge w_disc have \w: "... = LowLink s w" by simp also have "LowLink s w = LowLink ?s w" proof (rule LL_eqI) assume v: "v \ lowlink_set ?s w" show "LowLink s w \ \ ?s v" proof (cases "\ s v < \ s w \ \ s v < \ s w") case False with \LowLink s w \ \ s w\ \w show ?thesis by auto next case True with all_False have v_n_tj: "v \ set (tj_stack s)" by simp from v have e: "(v,u) \ E\<^sup>*" "(u,v) \ E\<^sup>*" unfolding lowlink_set_def by (auto simp add: \w=u\) from v_n_tj have "v \ set (stack s)" using stack_ss_tj_stack by auto with cross_back_edge have "v \ dom (finished s)" by (auto simp add: stack_set_def) with finished_ss_sccs_tj_stack v_n_tj sccs_are_sccs obtain scc where scc: "v \ scc" "scc \ sccs s" "is_scc E scc" by blast with is_scc_closed e have "u \ scc" by metis with scc sccs_finished u_n_fin have False by blast thus ?thesis .. qed qed finally show ?thesis . next case True note all_True = this with \w=u\ have "\ ?s w = \ s v" by (rule_tac TRANS) (simp add: tarjan_back_def cross_back_edge) also from True cross_back_edge w_disc have "\ s v < LowLink s w" by simp with lowlink_set_finite lowlink_set_not_empty w_disc have "\ s v = Min (?L \ {\ s v})" by simp also have "v \ lowlink_set ?s w" proof - have cb: "(u,v) \ cross_edges ?s \ back_edges ?s" by (simp add: cross_back_edge) with s'.lowlink_path_single have "lowlink_path ?s u [u] v" by auto moreover from cb s'.cross_edges_ssE s'.back_edges_ssE have "(u,v) \ E" by blast hence "(u,v) \ E\<^sup>*" .. moreover from all_True tj_stack_reach_hd_stack have "(v,u) \ E\<^sup>*" by (simp add: cross_back_edge) moreover note \v \ dom (discovered s)\ ultimately show ?thesis by (auto intro: s'.lowlink_setI simp: \w=u\) qed with ll_sub ll_sub_rev have "lowlink_set ?s w = lowlink_set s w \ {v}" by auto hence "Min (?L \ {\ s v}) = LowLink ?s w" by simp finally show ?thesis . qed next case False \ \\w \ u\\ hence "\ ?s w = \ s w" by (rule_tac TRANS) (simp add: tarjan_back_def cross_back_edge) also have "\ s w = LowLink s w" using w_disc False by (simp add: cross_back_edge) also have "LowLink s w = LowLink ?s w" proof (rule LL_eqI) assume v: "v \ lowlink_set ?s w" thus "LowLink s w \ \ ?s v" using LLw proof cases assume "v \ w" with v obtain p where p: "lowlink_path ?s w p v" "p\[]" by (auto simp add: lowlink_set_def lowlink_path_def) hence "hd p = w" by (auto simp add: lowlink_path_def path_hd) show ?thesis proof (cases "u \ set p") case False with last_in_set p cross_back_edge have "last p \ hd (stack s)" by force with p have "lowlink_path s w p v" by (auto simp: cross_back_edge lowlink_path_def) with v have "v \ lowlink_set s w" by (auto intro: lowlink_setI simp: lowlink_set_def cross_back_edge) thus ?thesis by simp next case True then obtain i where i: "i < length p" "p!i = u" by (metis in_set_conv_nth) have "False" proof (cases i) case "0" with i have "hd p = u" by (simp add: hd_conv_nth) with \hd p = w\ \w \ u\ show False by simp next case (Suc n) with i s'.lowlink_path_finished[OF p(1), where j=i] have "u \ dom (finished ?s)" by simp with u_n_fin show ?thesis by simp qed thus ?thesis .. qed qed simp qed finally show ?thesis . qed } note aux = this with cross_back_edge show ?case by (auto simp: pw_leof_iff) next case (finish s s' u) then interpret Tarjan_invar where s=s by simp from finish have [simp]: "discovered s' = discovered s" "finished s' = (finished s)(u\counter s)" "tree_edges s' = tree_edges s" "back_edges s' = back_edges s" "cross_edges s' = cross_edges s" "lowlink s' = lowlink s" "tj_stack s' = tj_stack s" by simp_all from finish hd_in_set stack_discovered have u_disc: "u \ dom (discovered s)" by blast { fix w :: "'v" fix x let ?s = "s'\state.more := x\" let ?L = "\ s ` lowlink_set s w" let ?Lu = "\ s ` lowlink_set s u" let ?L' = "\ s ` lowlink_set ?s w" assume TRANS: "\\. tarjan_fin u s' \\<^sub>n SPEC \ \ \ x" and inv': "DFS_invar G tarjan_params ?s" and w_disc: "w \ dom (discovered ?s)" from inv' interpret s':Tarjan_invar where s="?s" by simp have ll_sub: "lowlink_set s w \ lowlink_set ?s w" unfolding lowlink_set_def lowlink_path_def by auto have ll_sub_rev: "lowlink_set ?s w \ lowlink_set s w \ lowlink_set s u" proof fix ll assume ll: "ll \ lowlink_set ?s w" hence "ll = w \ (\p. lowlink_path ?s w p ll)" by (auto simp add: lowlink_set_def) thus "ll \ lowlink_set s w \ lowlink_set s u" proof (rule disjE1) assume "ll = w" with w_disc show ?thesis by (auto simp add: lowlink_set_def) next assume "ll \ w" assume "\p. lowlink_path ?s w p ll" - then guess p .. note p = this + then obtain p where p: "lowlink_path ?s w p ll" .. hence [simp]: "p\[]" by (simp add: lowlink_path_def) from p have "hd p = w" by (auto simp add: lowlink_path_def path_hd) show ?thesis proof (cases "u \ set p") case False hence "lowlink_path s w p ll" using p by (auto simp add: lowlink_path_def) with ll show ?thesis by (auto simp add: lowlink_set_def) next case True then obtain i where i: "i < length p" "p!i = u" by (metis in_set_conv_nth) moreover let ?dp = "drop i p" from i have "?dp \ []" by simp from i have "hd ?dp = u" by (simp add: hd_drop_conv_nth) moreover from i have "last ?dp = last p" by simp moreover { fix k assume "1 < length ?dp" and "k < length ?dp - 1" hence l: "1 < length p" "k+i < length p - 1" by (auto) with p have "(p!(k+i), p!Suc (k+i)) \ tree_edges s" by (auto simp add: lowlink_path_def) moreover from l have i: "i+k \ length p" "i+Suc k \ length p" by simp_all ultimately have "(?dp!k,?dp!Suc k) \ tree_edges s" by (simp add: add.commute) } note aux = this moreover { assume *: "1 < length ?dp" hence l: "1 + i < length p" by simp with s'.lowlink_path_finished[OF p] have "p ! (1+i) \ dom (finished ?s)" by auto moreover from l have "i+1\length p" by simp ultimately have "?dp!1 \ dom (finished ?s)" by simp moreover from aux[of 0] * have "(?dp!0,?dp!Suc 0) \ tree_edges s" by simp with \hd ?dp = u\ hd_conv_nth[of "?dp"] * have "(u,?dp!Suc 0) \ tree_edges s" by simp with no_self_loop_in_tree have "?dp!1 \ u" by auto ultimately have "?dp!1 \ dom (finished s)" by simp } moreover from p have P: "path E w p ll" by (simp add: lowlink_path_def) have "p = (take i p)@?dp" by simp with P path_conc_conv obtain x where p': "path E x ?dp ll" "path E w (take i p) x" by metis with \?dp \ []\ path_hd have "hd ?dp = x" by metis with \hd ?dp = u\ p' have u_path: "path E u ?dp ll" and path_u: "path E w (take i p) u" by metis+ ultimately have "lowlink_path s u ?dp ll" using p by (simp add: lowlink_path_def) moreover from u_path path_is_trancl \?dp \ []\ have "(u,ll) \ E\<^sup>+" by force moreover { from ll \ll \ w\ have "(ll,w) \ E\<^sup>+" by (auto simp add: lowlink_set_def) also from path_u path_is_rtrancl have "(w,u) \ E\<^sup>*" by metis finally have "(ll,u)\E\<^sup>+" . } moreover note ll u_disc ultimately have "ll \ lowlink_set s u" unfolding lowlink_set_def by auto thus ?thesis by auto qed qed qed hence ll_sub_rev': "?L' \ ?L \ ?Lu" by auto have ref_ne: "stack ?s \ [] \ lowlink ?s = (lowlink s)(hd (stack ?s) \ min (\ s (hd (stack ?s))) (\ s u))" apply (rule TRANS) unfolding tarjan_fin_def tj_stack_pop_def by refine_vcg simp_all have ref_e: "stack ?s = [] \ lowlink ?s = lowlink s" apply (rule TRANS) unfolding tarjan_fin_def tj_stack_pop_def by refine_vcg simp_all have ref_tj: "\ s u \ \ s u \ tj_stack ?s = tj_stack s" apply (rule TRANS) unfolding tarjan_fin_def tj_stack_pop_def by refine_vcg simp_all have "\ ?s w = LowLink ?s w" proof (cases "w = hd (stack ?s) \ stack ?s \ []") case True note all_True = this with ref_ne have *: "\ ?s w = min (\ s w) (\ s u)" by simp show ?thesis proof (cases "\ s u < \ s w") case False with * finish w_disc have "\ ?s w = LowLink s w" by simp also have "LowLink s w = LowLink ?s w" proof (rule LowLink_eqI[OF inv' _ ll_sub ll_sub_rev]) from w_disc show "w \ dom (discovered s)" by simp fix ll assume "ll \ lowlink_set s u" hence "LowLink s u \ \ s ll" by simp moreover from False finish w_disc u_disc have "LowLink s w \ LowLink s u" by simp ultimately show "LowLink s w \ \ ?s ll" by simp qed simp finally show ?thesis . next case True note \rel = this have "LowLink s u \ ?L'" proof - from all_True finish have w_tl: "w\set (tl (stack s))" by auto obtain ll where ll: "ll \ lowlink_set s u" "\ s ll = LowLink s u" using Min_in[of ?Lu] lowlink_set_finite lowlink_set_not_empty u_disc by fastforce have "ll \ lowlink_set ?s w" proof (cases "\ s u = \ s u") case True moreover from w_tl finish tl_lt_stack_hd_discover have "\ s w < \ s u" by simp moreover from w_disc have "LowLink s w \ \ s w" by (simp add: LowLink_le_disc) with w_disc finish have "\ s w \ \ s w" by simp moreover note \rel ultimately have False by force thus ?thesis .. next case False with u_disc finish ll have "u \ ll" by auto with ll have e: "(ll,u) \ E\<^sup>+" "(u,ll) \ E\<^sup>+" and p: "\p. lowlink_path s u p ll" and ll_disc: "ll \ dom (discovered s)" by (auto simp: lowlink_set_def) from p have p': "\p. lowlink_path ?s u p ll" unfolding lowlink_path_def by auto from w_tl tl_stack_hd_tree_path finish have T: "(w,u) \ (tree_edges ?s)\<^sup>+" by simp with s'.lowlink_path_tree_prepend all_True p' have "\p. lowlink_path ?s w p ll" by blast moreover from T trancl_mono_mp[OF s'.tree_edges_ssE] have "(w,u) \ E\<^sup>+" by blast with e have "(w,ll) \ E\<^sup>+" by simp moreover { note e(1) also from finish False ref_tj have "tj_stack ?s = tj_stack s" by simp with hd_in_set finish stack_ss_tj_stack have "u \ set (tj_stack ?s)" by auto with s'.tj_stack_reach_stack obtain x where x: "x \ set (stack ?s)" "(u,x) \ E\<^sup>*" by blast note this(2) also have "(x,w) \ E\<^sup>*" proof (rule rtrancl_eq_or_trancl[THEN iffD2], safe) assume "x \ w" with all_True x have "x \ set (tl (stack ?s))" by (cases "stack ?s") auto with s'.tl_stack_hd_tree_path all_True have "(x,w) \ (tree_edges s)\<^sup>+" by auto with trancl_mono_mp[OF tree_edges_ssE] show "(x,w) \ E\<^sup>+" by simp qed finally have "(ll,w) \ E\<^sup>+" . } moreover note ll_disc ultimately show ?thesis by (simp add: lowlink_set_def) qed hence "\ s ll \ ?L'" by auto with ll show ?thesis by simp qed hence "LowLink ?s w \ LowLink s u" using Min_le_iff[of ?L'] s'.lowlink_set_not_empty w_disc s'.lowlink_set_finite by fastforce also from True u_disc w_disc finish have "LowLink s u < LowLink s w" by simp hence "Min (?L \ ?Lu) = LowLink s u" using Min_Un[of ?L ?Lu] lowlink_set_finite lowlink_set_not_empty u_disc w_disc by simp hence "LowLink s u \ LowLink ?s w" using Min_antimono[OF ll_sub_rev'] lowlink_set_finite s'.lowlink_set_not_empty w_disc by auto also from True u_disc finish * have "LowLink s u = \ ?s w" by simp finally show ?thesis .. qed next case False note all_False = this have "\ ?s w = \ s w" proof (cases "stack ?s = []") case True with ref_e show ?thesis by simp next case False with all_False have "w \ hd (stack ?s)" by simp with False ref_ne show ?thesis by simp qed also from finish have "\ s w = LowLink s w" using w_disc by simp also { fix v assume "v \ lowlink_set s u" and *: "v \ lowlink_set s w" hence "v \ w" "w\u" by (auto simp add: lowlink_set_def) have "v \ lowlink_set ?s w" proof (rule notI) assume v: "v \ lowlink_set ?s w" hence e: "(v,w) \ E\<^sup>*" "(w,v) \ E\<^sup>*" and v_disc: "v \ dom (discovered s)" by (auto simp add: lowlink_set_def) from v \v\w\ obtain p where p: "lowlink_path ?s w p v" by (auto simp add: lowlink_set_def) hence [simp]: "p\[]" by (simp add: lowlink_path_def) from p have "hd p = w" by (auto simp add: lowlink_path_def path_hd) show False proof (cases "u \ set p") case False hence "lowlink_path s w p v" using p by (auto simp add: lowlink_path_def) with e v_disc have "v \ lowlink_set s w" by (auto intro: lowlink_setI) with * show False .. next case True then obtain i where i: "i < length p" "p!i = u" by (metis in_set_conv_nth) show "False" proof (cases i) case "0" with i have "hd p = u" by (simp add: hd_conv_nth) with \hd p = w\ \w \ u\ show False by simp next case (Suc n) with i p have *: "(p!n,u) \ tree_edges s" "n < length p" unfolding lowlink_path_def by auto with tree_edge_imp_discovered have "p!n \ dom (discovered s)" by auto moreover from finish hd_in_set stack_not_finished have "u \ dom (finished s)" by auto with * have pn_n_fin: "p!n \ dom (finished s)" by (metis tree_edge_impl_parenthesis) moreover from * no_self_loop_in_tree have "p!n \ u" by blast ultimately have "p!n \ set (stack ?s)" using stack_set_def finish by (cases "stack s") auto hence s_ne: "stack ?s \ []" by auto with all_False have "w \ hd (stack ?s)" by simp from stack_is_tree_path finish obtain v0 where "path (tree_edges s) v0 (rev (stack ?s)) u" by auto with s_ne have "(hd (stack ?s), u) \ tree_edges s" by (auto simp: neq_Nil_conv path_simps) with * tree_eq_rule have **: "hd (stack ?s) = p!n" by simp show ?thesis proof (cases n) case "0" with * have "hd p = p!n" by (simp add: hd_conv_nth) with \hd p = w\ ** have "w = hd (stack ?s)" by simp with \w\hd (stack ?s)\ show False .. next case (Suc m) with * ** s'.lowlink_path_finished[OF p, where j=n] have "hd (stack ?s) \ dom (finished ?s)" by simp with hd_in_set[OF s_ne] s'.stack_not_finished show ?thesis by blast qed qed qed qed } with ll_sub ll_sub_rev have "lowlink_set ?s w = lowlink_set s w" by auto hence "LowLink s w = LowLink ?s w" by simp finally show ?thesis . qed } with finish show ?case by (auto simp: pw_leof_iff) qed simp_all qed end end context Tarjan_invar begin context begin interpretation timing_syntax . lemmas lowlink_eq_LowLink = i_lowlink_eq_LowLink[THEN make_invar_thm, rule_format] lemma lowlink_eq_disc_iff_scc_root: assumes "v \ dom (finished s) \ (stack s \ [] \ v = hd (stack s) \ pending s `` {v} = {})" shows "\ s v = \ s v \ scc_root s v (scc_of E v)" proof - from assms have "v \ dom (discovered s)" using finished_discovered hd_in_set stack_discovered by blast hence "\ s v = LowLink s v" using lowlink_eq_LowLink by simp with LowLink_eq_disc_iff_scc_root[OF assms] show ?thesis by simp qed lemma nc_sccs_eq_reachable: assumes NC: "\ cond s" shows "reachable = \(sccs s)" proof from nc_finished_eq_reachable NC have [simp]: "reachable = dom (finished s)" by simp with sccs_finished show "\(sccs s) \ reachable" by simp from NC have "stack s = []" by (simp add: cond_alt) with stacks_eq_iff have "tj_stack s = []" by simp with finished_ss_sccs_tj_stack show "reachable \ \(sccs s)" by simp qed end end context Tarjan begin lemma tarjan_fin_nofail: assumes "pre_on_finish u s'" shows "nofail (tarjan_fin u s')" proof - from assms obtain s where s: "DFS_invar G tarjan_params s" "stack s \ []" "u = hd (stack s)" "s' = finish u s" "cond s" "pending s `` {u} = {}" by (auto simp: pre_on_finish_def) then interpret Tarjan_invar where s=s by simp from s hd_stack_in_tj_stack have "u \ set (tj_stack s')" by simp moreover from s tj_stack_distinct have "distinct (tj_stack s')" by simp moreover have "the (lowlink s' u) = the (discovered s' u) \ scc_root s' u (scc_of E u)" proof - from s have "the (lowlink s' u) = the (discovered s' u) \ the (lowlink s u) = the (discovered s u)" by simp also from s lowlink_eq_disc_iff_scc_root have "... \ scc_root s u (scc_of E u)" by blast also from s scc_root_transfer'[where s'=s'] have "... \ scc_root s' u (scc_of E u)" by simp finally show ?thesis . qed ultimately show ?thesis unfolding tarjan_fin_def tj_stack_pop_def by simp qed sublocale DFS G tarjan_params by unfold_locales (simp_all add: tarjan_disc_def tarjan_back_def tarjan_fin_nofail) end interpretation tarjan: Tarjan_def for G . subsection \Interface\ definition "tarjan G \ do { ASSERT (fb_graph G); s \ tarjan.it_dfs TYPE('a) G; RETURN (sccs s) }" definition "tarjan_spec G \ do { ASSERT (fb_graph G); SPEC (\sccs. (\scc \ sccs. is_scc (g_E G) scc) \ \sccs = tarjan.reachable TYPE('a) G)}" lemma tarjan_correct: "tarjan G \ tarjan_spec G" unfolding tarjan_def tarjan_spec_def proof (refine_vcg le_ASSERTI order_trans[OF DFS.it_dfs_correct]) assume "fb_graph G" then interpret fb_graph G . interpret Tarjan .. show "DFS G (tarjan.tarjan_params TYPE('b) G)" .. next fix s assume C: "DFS_invar G (tarjan.tarjan_params TYPE('b) G) s \ \ tarjan.cond TYPE('b) G s" then interpret Tarjan_invar G s by simp from sccs_are_sccs show "\scc\sccs s. is_scc (g_E G) scc" . from nc_sccs_eq_reachable C show "\(sccs s) = tarjan.reachable TYPE('b) G" by simp qed end diff --git a/thys/Dirichlet_Series/Dirichlet_Misc.thy b/thys/Dirichlet_Series/Dirichlet_Misc.thy --- a/thys/Dirichlet_Series/Dirichlet_Misc.thy +++ b/thys/Dirichlet_Series/Dirichlet_Misc.thy @@ -1,158 +1,159 @@ (* File: Dirichlet_Misc.thy Author: Manuel Eberl, TU München *) section \Miscellaneous auxiliary facts\ theory Dirichlet_Misc imports Main "HOL-Number_Theory.Number_Theory" begin lemma fixes a k :: nat assumes "a > 1" "k > 0" shows geometric_sum_nat_aux: "(a - 1) * (\iiiiiii n > 0 \ n div d > (0::nat)" by (auto elim: dvdE) lemma Set_filter_insert: "Set.filter P (insert x A) = (if P x then insert x (Set.filter P A) else Set.filter P A)" by (auto simp: Set.filter_def) lemma Set_filter_union: "Set.filter P (A \ B) = Set.filter P A \ Set.filter P B" by (auto simp: Set.filter_def) lemma Set_filter_empty [simp]: "Set.filter P {} = {}" by (auto simp: Set.filter_def) lemma Set_filter_image: "Set.filter P (f ` A) = f ` Set.filter (P \ f) A" by (auto simp: Set.filter_def) lemma Set_filter_cong [cong]: "(\x. x \ A \ P x \ Q x) \ A = B \ Set.filter P A = Set.filter Q B" by (auto simp: Set.filter_def) lemma finite_Set_filter: "finite A \ finite (Set.filter P A)" by (auto simp: Set.filter_def) lemma inj_on_insert': "(\B. B \ A \ x \ B) \ inj_on (insert x) A" by (auto simp: inj_on_def insert_eq_iff) lemma assumes "finite A" "A \ {}" shows card_even_subset_aux: "card {B. B \ A \ even (card B)} = 2 ^ (card A - 1)" and card_odd_subset_aux: "card {B. B \ A \ odd (card B)} = 2 ^ (card A - 1)" and card_even_odd_subset: "card {B. B \ A \ even (card B)} = card {B. B \ A \ odd (card B)}" proof - from assms have *: "2 * card (Set.filter (even \ card) (Pow A)) = 2 ^ card A" proof (induction A rule: finite_ne_induct) case (singleton x) hence "Pow {x} = {{}, {x}}" by auto thus ?case by (simp add: Set_filter_insert) next case (insert x A) note fin = finite_subset[OF _ \finite A\] have "Pow (insert x A) = Pow A \ insert x ` Pow A" by (rule Pow_insert) have "Set.filter (even \ card) (Pow (insert x A)) = Set.filter (even \ card) (Pow A) \ insert x ` Set.filter (even \ card \ insert x) (Pow A)" unfolding Pow_insert Set_filter_union Set_filter_image by blast also have "Set.filter (even \ card \ insert x) (Pow A) = Set.filter (odd \ card) (Pow A)" unfolding o_def by (intro Set_filter_cong refl, subst card_insert_disjoint) (insert insert.hyps, auto dest: finite_subset) also have "card (Set.filter (even \ card) (Pow A) \ insert x ` \) = card (Set.filter (even \ card) (Pow A)) + card (insert x ` \)" (is "card (?A \ ?B) = _") by (intro card_Un_disjoint finite_Set_filter finite_imageI) (auto simp: insert.hyps) also have "card ?B = card (Set.filter (odd \ card) (Pow A))" using insert.hyps by (intro card_image inj_on_insert') auto also have "Set.filter (odd \ card) (Pow A) = Pow A - Set.filter (even \ card) (Pow A)" by auto also have "card \ = card (Pow A) - card (Set.filter (even \ card) (Pow A))" using insert.hyps by (subst card_Diff_subset) (auto simp: finite_Set_filter) also have "card (Set.filter (even \ card) (Pow A)) + \ = card (Pow A)" by (intro add_diff_inverse_nat, subst not_less, rule card_mono) (insert insert.hyps, auto) also have "2 * \ = 2 ^ card (insert x A)" using insert.hyps by (simp add: card_Pow) finally show ?case . qed from * show A: "card {B. B \ A \ even (card B)} = 2 ^ (card A - 1)" by (cases "card A") (simp_all add: Set.filter_def) have "Set.filter (odd \ card) (Pow A) = Pow A - Set.filter (even \ card) (Pow A)" by auto also have "2 * card \ = 2 * 2 ^ card A - 2 * card (Set.filter (even \ card) (Pow A))" using assms by (subst card_Diff_subset) (auto intro!: finite_Set_filter simp: card_Pow) also note * also have "2 * 2 ^ card A - 2 ^ card A = (2 ^ card A :: nat)" by simp finally show B: "card {B. B \ A \ odd (card B)} = 2 ^ (card A - 1)" by (cases "card A") (simp_all add: Set.filter_def) from A and B show "card {B. B \ A \ even (card B)} = card {B. B \ A \ odd (card B)}" by simp qed lemma bij_betw_prod_divisors_coprime: assumes "coprime a (b :: nat)" shows "bij_betw (\x. fst x * snd x) ({d. d dvd a} \ {d. d dvd b}) {k. k dvd a * b}" unfolding bij_betw_def proof from assms show "inj_on (\x. fst x * snd x) ({d. d dvd a} \ {d. d dvd b})" by (auto simp: inj_on_def coprime_crossproduct_nat coprime_divisors) show "(\x. fst x * snd x) ` ({d. d dvd a} \ {d. d dvd b}) = {k. k dvd a * b}" proof safe fix x assume "x dvd a * b" - from division_decomp[OF this] guess b' c' by (elim exE conjE) + then obtain b' c' where "x = b' * c'" "b' dvd a" "c' dvd b" + using division_decomp by blast thus "x \ (\x. fst x * snd x) ` ({d. d dvd a} \ {d. d dvd b})" by force qed (insert assms, auto intro: mult_dvd_mono) qed lemma bij_betw_prime_power_divisors: assumes "prime (p :: nat)" shows "bij_betw ((^) p) {..k} {d. d dvd p ^ k}" unfolding bij_betw_def proof from assms have *: "p > 1" by (simp add: prime_gt_Suc_0_nat) show "inj_on ((^) p) {..k}" using assms by (auto simp: inj_on_def prime_gt_Suc_0_nat power_inject_exp[OF *]) show "(^) p ` {..k} = {d. d dvd p ^ k}" using assms by (auto simp: le_imp_power_dvd divides_primepow_nat) qed lemma power_diff': assumes "m \ n" "x \ 0" shows "x ^ (m - n) = (x ^ m div x ^ n :: 'a :: unique_euclidean_semiring)" proof - from assms have "x ^ m = x ^ (m - n) * x ^ n" by (subst power_add [symmetric]) simp also from assms have "\ div x ^ n = x ^ (m - n)" by simp finally show ?thesis .. qed lemma sum_divisors_coprime_mult: assumes "coprime a (b :: nat)" shows "(\d | d dvd a * b. f d) = (\r | r dvd a. \s | s dvd b. f (r * s))" proof - have "(\r | r dvd a. \s | s dvd b. f (r * s)) = (\z\{r. r dvd a} \ {s. s dvd b}. f (fst z * snd z))" by (subst sum.cartesian_product) (simp add: case_prod_unfold) also have "\ = (\d | d dvd a * b. f d)" by (intro sum.reindex_bij_betw bij_betw_prod_divisors_coprime assms) finally show ?thesis .. qed end diff --git a/thys/E_Transcendental/E_Transcendental.thy b/thys/E_Transcendental/E_Transcendental.thy --- a/thys/E_Transcendental/E_Transcendental.thy +++ b/thys/E_Transcendental/E_Transcendental.thy @@ -1,687 +1,687 @@ (* File: E_Transcendental.thy Author: Manuel Eberl A proof that e (Euler's number) is transcendental. Could possibly be extended to a transcendence proof for pi or the very general Lindemann-Weierstrass theorem. *) section \Proof of the Transcendence of $e$\ theory E_Transcendental imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Number_Theory.Number_Theory" "HOL-Computational_Algebra.Polynomial" begin (* TODO: Lots of stuff to move to the distribution *) subsection \Various auxiliary facts\ lemma fact_dvd_pochhammer: assumes "m \ n + 1" shows "fact m dvd pochhammer (int n - int m + 1) m" proof - have "(real n gchoose m) * fact m = of_int (pochhammer (int n - int m + 1) m)" by (simp add: gbinomial_pochhammer' pochhammer_of_int [symmetric]) also have "(real n gchoose m) * fact m = of_int (int (n choose m) * fact m)" by (simp add: binomial_gbinomial) finally have "int (n choose m) * fact m = pochhammer (int n - int m + 1) m" by (subst (asm) of_int_eq_iff) from this [symmetric] show ?thesis by simp qed lemma of_nat_eq_1_iff [simp]: "of_nat x = (1 :: 'a :: semiring_char_0) \ x = 1" by (fact of_nat_eq_1_iff) lemma prime_elem_int_not_dvd_neg1_power: "prime_elem (p :: int) \ \p dvd (-1) ^ n" by (rule notI, frule (1) prime_elem_dvd_power, cases "p \ 0") (auto simp: prime_elem_def) lemma nat_fact [simp]: "nat (fact n) = fact n" by (subst of_nat_fact [symmetric]) (rule nat_int) lemma prime_dvd_fact_iff_int: "p dvd fact n \ p \ int n" if "prime p" using that prime_dvd_fact_iff [of "nat \p\" n] by auto (simp add: prime_ge_0_int) lemma filterlim_minus_nat_at_top: "filterlim (\n. n - k :: nat) at_top at_top" proof - have "sequentially = filtermap (\n. n + k) at_top" by (auto simp: filter_eq_iff eventually_filtermap) also have "filterlim (\n. n - k :: nat) at_top \" by (simp add: filterlim_filtermap filterlim_ident) finally show ?thesis . qed lemma power_over_fact_tendsto_0: "(\n. (x :: real) ^ n / fact n) \ 0" using summable_exp[of x] by (intro summable_LIMSEQ_zero) (simp add: sums_iff field_simps) lemma power_over_fact_tendsto_0': "(\n. c * (x :: real) ^ n / fact n) \ 0" using tendsto_mult[OF tendsto_const[of c] power_over_fact_tendsto_0[of x]] by simp subsection \Lifting integer polynomials\ lift_definition of_int_poly :: "int poly \ 'a :: comm_ring_1 poly" is "\g x. of_int (g x)" by (auto elim: eventually_mono) lemma coeff_of_int_poly [simp]: "coeff (of_int_poly p) n = of_int (coeff p n)" by transfer' simp lemma of_int_poly_0 [simp]: "of_int_poly 0 = 0" by transfer (simp add: fun_eq_iff) lemma of_int_poly_pCons [simp]: "of_int_poly (pCons c p) = pCons (of_int c) (of_int_poly p)" by transfer' (simp add: fun_eq_iff split: nat.splits) lemma of_int_poly_smult [simp]: "of_int_poly (smult c p) = smult (of_int c) (of_int_poly p)" by transfer simp lemma of_int_poly_1 [simp]: "of_int_poly 1 = 1" by (simp add: one_pCons) lemma of_int_poly_add [simp]: "of_int_poly (p + q) = of_int_poly p + of_int_poly q" by transfer' (simp add: fun_eq_iff) lemma of_int_poly_mult [simp]: "of_int_poly (p * q) = (of_int_poly p * of_int_poly q)" by (induction p) simp_all lemma of_int_poly_sum [simp]: "of_int_poly (sum f A) = sum (\x. of_int_poly (f x)) A" by (induction A rule: infinite_finite_induct) simp_all lemma of_int_poly_prod [simp]: "of_int_poly (prod f A) = prod (\x. of_int_poly (f x)) A" by (induction A rule: infinite_finite_induct) simp_all lemma of_int_poly_power [simp]: "of_int_poly (p ^ n) = of_int_poly p ^ n" by (induction n) simp_all lemma of_int_poly_monom [simp]: "of_int_poly (monom c n) = monom (of_int c) n" by transfer (simp add: fun_eq_iff) lemma poly_of_int_poly [simp]: "poly (of_int_poly p) (of_int x) = of_int (poly p x)" by (induction p) simp_all lemma poly_of_int_poly_of_nat [simp]: "poly (of_int_poly p) (of_nat x) = of_int (poly p (int x))" by (induction p) simp_all lemma poly_of_int_poly_0 [simp]: "poly (of_int_poly p) 0 = of_int (poly p 0)" by (induction p) simp_all lemma poly_of_int_poly_1 [simp]: "poly (of_int_poly p) 1 = of_int (poly p 1)" by (induction p) simp_all lemma poly_of_int_poly_of_real [simp]: "poly (of_int_poly p) (of_real x) = of_real (poly (of_int_poly p) x)" by (induction p) simp_all lemma of_int_poly_eq_iff [simp]: "of_int_poly p = (of_int_poly q :: 'a :: {comm_ring_1, ring_char_0} poly) \ p = q" by (simp add: poly_eq_iff) lemma of_int_poly_eq_0_iff [simp]: "of_int_poly p = (0 :: 'a :: {comm_ring_1, ring_char_0} poly) \ p = 0" using of_int_poly_eq_iff[of p 0] by (simp del: of_int_poly_eq_iff) lemma degree_of_int_poly [simp]: "degree (of_int_poly p :: 'a :: {comm_ring_1, ring_char_0} poly) = degree p" by (simp add: degree_def) lemma pderiv_of_int_poly [simp]: "pderiv (of_int_poly p) = of_int_poly (pderiv p)" by (induction p) (simp_all add: pderiv_pCons) lemma higher_pderiv_of_int_poly [simp]: "(pderiv ^^ n) (of_int_poly p) = of_int_poly ((pderiv ^^ n) p)" by (induction n) simp_all lemma int_polyE: assumes "\n. coeff (p :: 'a :: {comm_ring_1, ring_char_0} poly) n \ \" obtains p' where "p = of_int_poly p'" proof - from assms have "\n. \c. coeff p n = of_int c" by (auto simp: Ints_def) hence "\c. \n. of_int (c n) = coeff p n" by (simp add: choice_iff eq_commute) then obtain c where c: "of_int (c n) = coeff p n" for n by blast have [simp]: "coeff (Abs_poly c) = c" proof (rule poly.Abs_poly_inverse, clarify) have "eventually (\n. n > degree p) at_top" by (rule eventually_gt_at_top) hence "eventually (\n. coeff p n = 0) at_top" by eventually_elim (simp add: coeff_eq_0) thus "eventually (\n. c n = 0) cofinite" by (simp add: c [symmetric] cofinite_eq_sequentially) qed have "p = of_int_poly (Abs_poly c)" by (rule poly_eqI) (simp add: c) thus ?thesis by (rule that) qed subsection \General facts about polynomials\ lemma pderiv_power: "pderiv (p ^ n) = smult (of_nat n) (p ^ (n - 1) * pderiv p)" by (cases n) (simp_all add: pderiv_power_Suc del: power_Suc) lemma degree_prod_sum_eq: "(\x. x \ A \ f x \ 0) \ degree (prod f A :: 'a :: idom poly) = (\x\A. degree (f x))" by (induction A rule: infinite_finite_induct) (auto simp: degree_mult_eq) lemma pderiv_monom: "pderiv (monom c n) = monom (of_nat n * c) (n - 1)" by (cases n) (simp_all add: monom_altdef pderiv_power_Suc pderiv_smult pderiv_pCons mult_ac del: power_Suc) lemma power_poly_const [simp]: "[:c:] ^ n = [:c ^ n:]" by (induction n) (simp_all add: power_commutes) lemma monom_power: "monom c n ^ k = monom (c ^ k) (n * k)" by (induction k) (simp_all add: mult_monom) lemma coeff_higher_pderiv: "coeff ((pderiv ^^ m) f) n = pochhammer (of_nat (Suc n)) m * coeff f (n + m)" by (induction m arbitrary: n) (simp_all add: coeff_pderiv pochhammer_rec algebra_simps) lemma higher_pderiv_add: "(pderiv ^^ n) (p + q) = (pderiv ^^ n) p + (pderiv ^^ n) q" by (induction n arbitrary: p q) (simp_all del: funpow.simps add: funpow_Suc_right pderiv_add) lemma higher_pderiv_smult: "(pderiv ^^ n) (smult c p) = smult c ((pderiv ^^ n) p)" by (induction n arbitrary: p) (simp_all del: funpow.simps add: funpow_Suc_right pderiv_smult) lemma higher_pderiv_0 [simp]: "(pderiv ^^ n) 0 = 0" by (induction n) simp_all lemma higher_pderiv_monom: "m \ n + 1 \ (pderiv ^^ m) (monom c n) = monom (pochhammer (int n - int m + 1) m * c) (n - m)" proof (induction m arbitrary: c n) case (Suc m) thus ?case by (cases n) (simp_all del: funpow.simps add: funpow_Suc_right pderiv_monom pochhammer_rec' Suc.IH) qed simp_all lemma higher_pderiv_monom_eq_zero: "m > n + 1 \ (pderiv ^^ m) (monom c n) = 0" proof (induction m arbitrary: c n) case (Suc m) thus ?case by (cases n) (simp_all del: funpow.simps add: funpow_Suc_right pderiv_monom pochhammer_rec' Suc.IH) qed simp_all lemma higher_pderiv_sum: "(pderiv ^^ n) (sum f A) = (\x\A. (pderiv ^^ n) (f x))" by (induction A rule: infinite_finite_induct) (simp_all add: higher_pderiv_add) lemma fact_dvd_higher_pderiv: "[:fact n :: int:] dvd (pderiv ^^ n) p" proof - have "[:fact n:] dvd (pderiv ^^ n) (monom c k)" for c :: int and k :: nat by (cases "n \ k + 1") (simp_all add: higher_pderiv_monom higher_pderiv_monom_eq_zero fact_dvd_pochhammer const_poly_dvd_iff) hence "[:fact n:] dvd (pderiv ^^ n) (\k\degree p. monom (coeff p k) k)" by (simp_all add: higher_pderiv_sum dvd_sum) thus ?thesis by (simp add: poly_as_sum_of_monoms) qed lemma fact_dvd_poly_higher_pderiv_aux: "(fact n :: int) dvd poly ((pderiv ^^ n) p) x" proof - have "[:fact n:] dvd (pderiv ^^ n) p" by (rule fact_dvd_higher_pderiv) then obtain q where "(pderiv ^^ n) p = [:fact n:] * q" by (erule dvdE) thus ?thesis by simp qed lemma fact_dvd_poly_higher_pderiv_aux': "m \ n \ (fact m :: int) dvd poly ((pderiv ^^ n) p) x" by (rule dvd_trans[OF fact_dvd fact_dvd_poly_higher_pderiv_aux]) simp_all lemma algebraicE': assumes "algebraic (x :: 'a :: field_char_0)" obtains p where "p \ 0" "poly (of_int_poly p) x = 0" proof - from assms obtain q where "\i. coeff q i \ \" "q \ 0" "poly q x = 0" by (erule algebraicE) moreover from this(1) obtain q' where "q = of_int_poly q'" by (erule int_polyE) ultimately show ?thesis by (intro that[of q']) simp_all qed lemma algebraicE'_nonzero: assumes "algebraic (x :: 'a :: field_char_0)" "x \ 0" obtains p where "p \ 0" "coeff p 0 \ 0" "poly (of_int_poly p) x = 0" proof - from assms(1) obtain p where p: "p \ 0" "poly (of_int_poly p) x = 0" by (erule algebraicE') define n :: nat where "n = order 0 p" have "monom 1 n dvd p" by (simp add: monom_1_dvd_iff p n_def) then obtain q where q: "p = monom 1 n * q" by (erule dvdE) from p have "q \ 0" "poly (of_int_poly q) x = 0" by (auto simp: q poly_monom assms(2)) moreover from this have "order 0 p = n + order 0 q" by (simp add: q order_mult) hence "order 0 q = 0" by (simp add: n_def) with \q \ 0\ have "poly q 0 \ 0" by (simp add: order_root) ultimately show ?thesis using that[of q] by (auto simp: poly_0_coeff_0) qed lemma algebraic_of_real_iff [simp]: "algebraic (of_real x :: 'a :: {real_algebra_1,field_char_0}) \ algebraic x" proof assume "algebraic (of_real x :: 'a)" then obtain p where "p \ 0" "poly (of_int_poly p) (of_real x :: 'a) = 0" by (erule algebraicE') hence "(of_int_poly p :: real poly) \ 0" "poly (of_int_poly p :: real poly) x = 0" by simp_all thus "algebraic x" by (intro algebraicI[of "of_int_poly p"]) simp_all next assume "algebraic x" then obtain p where "p \ 0" "poly (of_int_poly p) x = 0" by (erule algebraicE') hence "of_int_poly p \ (0 :: 'a poly)" "poly (of_int_poly p) (of_real x :: 'a) = 0" by simp_all thus "algebraic (of_real x)" by (intro algebraicI[of "of_int_poly p"]) simp_all qed subsection \Main proof\ lemma lindemann_weierstrass_integral: fixes u :: complex and f :: "complex poly" defines "df \ \n. (pderiv ^^ n) f" defines "m \ degree f" defines "I \ \f u. exp u * (\j\degree f. poly ((pderiv ^^ j) f) 0) - (\j\degree f. poly ((pderiv ^^ j) f) u)" shows "((\t. exp (u - t) * poly f t) has_contour_integral I f u) (linepath 0 u)" proof - note [derivative_intros] = exp_scaleR_has_vector_derivative_right vector_diff_chain_within let ?g = "\t. 1 - t" and ?f = "\t. -exp (t *\<^sub>R u)" have "((\t. exp ((1 - t) *\<^sub>R u) * u) has_integral (?f \ ?g) 1 - (?f \ ?g) 0) {0..1}" by (rule fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp del: o_apply) hence aux_integral: "((\t. exp (u - t *\<^sub>R u) * u) has_integral exp u - 1) {0..1}" by (simp add: algebra_simps) have "((\t. exp (u - t *\<^sub>R u) * u * poly f (t *\<^sub>R u)) has_integral I f u) {0..1}" unfolding df_def m_def proof (induction "degree f" arbitrary: f) case 0 then obtain c where c: "f = [:c:]" by (auto elim: degree_eq_zeroE) have "((\t. c * (exp (u - t *\<^sub>R u) * u)) has_integral c * (exp u - 1)) {0..1}" using aux_integral by (rule has_integral_mult_right) with c show ?case by (simp add: algebra_simps I_def) next case (Suc m) define df where "df = (\j. (pderiv ^^ j) f)" show ?case proof (rule integration_by_parts[OF bounded_bilinear_mult]) fix t :: real assume "t \ {0..1}" have "((?f \ ?g) has_vector_derivative exp (u - t *\<^sub>R u) * u) (at t)" by (auto intro!: derivative_eq_intros simp: algebra_simps simp del: o_apply) thus "((\t. -exp (u - t *\<^sub>R u)) has_vector_derivative exp (u - t *\<^sub>R u) * u) (at t)" by (simp add: algebra_simps o_def) next fix t :: real assume "t \ {0..1}" have "(poly f \ (\t. t *\<^sub>R u) has_vector_derivative u * poly (pderiv f) (t *\<^sub>R u)) (at t)" by (rule field_vector_diff_chain_at) (auto intro!: derivative_eq_intros) thus "((\t. poly f (t *\<^sub>R u)) has_vector_derivative u * poly (pderiv f) (t *\<^sub>R u)) (at t)" by (simp add: o_def) next from Suc(2) have m: "m = degree (pderiv f)" by (simp add: degree_pderiv) from Suc(1)[OF this] this have "((\t. exp (u - t *\<^sub>R u) * u * poly (pderiv f) (t *\<^sub>R u)) has_integral exp u * (\j=0..m. poly (df (Suc j)) 0) - (\j=0..m. poly (df (Suc j)) u)) {0..1}" by (simp add: df_def funpow_swap1 atMost_atLeast0 I_def) also have "(\j=0..m. poly (df (Suc j)) 0) = (\j=Suc 0..Suc m. poly (df j) 0)" by (rule sum.shift_bounds_cl_Suc_ivl [symmetric]) also have "\ = (\j=0..Suc m. poly (df j) 0) - poly f 0" by (subst (2) sum.atLeast_Suc_atMost) (simp_all add: df_def) also have "(\j=0..m. poly (df (Suc j)) u) = (\j=Suc 0..Suc m. poly (df j) u)" by (rule sum.shift_bounds_cl_Suc_ivl [symmetric]) also have "\ = (\j=0..Suc m. poly (df j) u) - poly f u" by (subst (2) sum.atLeast_Suc_atMost) (simp_all add: df_def) finally have "((\t. - (exp (u - t *\<^sub>R u) * u * poly (pderiv f) (t *\<^sub>R u))) has_integral -(exp u * ((\j = 0..Suc m. poly (df j) 0) - poly f 0) - ((\j = 0..Suc m. poly (df j) u) - poly f u))) {0..1}" (is "(_ has_integral ?I) _") by (rule has_integral_neg) also have "?I = - exp (u - 1 *\<^sub>R u) * poly f (1 *\<^sub>R u) - - exp (u - 0 *\<^sub>R u) * poly f (0 *\<^sub>R u) - I f u" by (simp add: df_def algebra_simps Suc(2) atMost_atLeast0 I_def) finally show "((\t. - exp (u - t *\<^sub>R u) * (u * poly (pderiv f) (t *\<^sub>R u))) has_integral \) {0..1}" by (simp add: algebra_simps) qed (auto intro!: continuous_intros) qed thus ?thesis by (simp add: has_contour_integral_linepath algebra_simps) qed locale lindemann_weierstrass_aux = fixes f :: "complex poly" begin definition I :: "complex \ complex" where "I u = exp u * (\j\degree f. poly ((pderiv ^^ j) f) 0) - (\j\degree f. poly ((pderiv ^^ j) f) u)" lemma lindemann_weierstrass_integral_bound: fixes u :: complex assumes "C \ 0" "\t. t \ closed_segment 0 u \ norm (poly f t) \ C" shows "norm (I u) \ norm u * exp (norm u) * C" proof - have "I u = contour_integral (linepath 0 u) (\t. exp (u - t) * poly f t)" using contour_integral_unique[OF lindemann_weierstrass_integral[of u f]] unfolding I_def .. also have "norm \ \ exp (norm u) * C * norm (u - 0)" proof (intro contour_integral_bound_linepath) fix t assume t: "t \ closed_segment 0 u" then obtain s where s: "s \ {0..1}" "t = s *\<^sub>R u" by (auto simp: closed_segment_def) hence "s * norm u \ 1 * norm u" by (intro mult_right_mono) simp_all with s have norm_t: "norm t \ norm u" by auto from s have "Re u - Re t = (1 - s) * Re u" by (simp add: algebra_simps) also have "\ \ norm u" proof (cases "Re u \ 0") case True with \s \ {0..1}\ have "(1 - s) * Re u \ 1 * Re u" by (intro mult_right_mono) simp_all also have "Re u \ norm u" by (rule complex_Re_le_cmod) finally show ?thesis by simp next case False with \s \ {0..1}\ have "(1 - s) * Re u \ 0" by (intro mult_nonneg_nonpos) simp_all also have "\ \ norm u" by simp finally show ?thesis . qed finally have "exp (Re u - Re t) \ exp (norm u)" by simp hence "exp (Re u - Re t) * norm (poly f t) \ exp (norm u) * C" using assms t norm_t by (intro mult_mono) simp_all thus "norm (exp (u - t) * poly f t) \ exp (norm u) * C" by (simp add: norm_mult exp_diff norm_divide field_simps) qed (auto simp: intro!: mult_nonneg_nonneg contour_integrable_continuous_linepath continuous_intros assms) finally show ?thesis by (simp add: mult_ac) qed end lemma poly_higher_pderiv_aux1: fixes c :: "'a :: idom" assumes "k < n" shows "poly ((pderiv ^^ k) ([:-c, 1:] ^ n * p)) c = 0" using assms proof (induction k arbitrary: n p) case (Suc k n p) from Suc.prems obtain n' where n: "n = Suc n'" by (cases n) auto from Suc.prems n have "k < n'" by simp have "(pderiv ^^ Suc k) ([:- c, 1:] ^ n * p) = (pderiv ^^ k) ([:- c, 1:] ^ n * pderiv p + [:- c, 1:] ^ n' * smult (of_nat n) p)" by (simp only: funpow_Suc_right o_def pderiv_mult n pderiv_power_Suc, simp only: n [symmetric]) (simp add: pderiv_pCons mult_ac) also from Suc.prems \k < n'\ have "poly \ c = 0" by (simp add: higher_pderiv_add Suc.IH del: mult_smult_right) finally show ?case . qed simp_all lemma poly_higher_pderiv_aux1': fixes c :: "'a :: idom" assumes "k < n" "[:-c, 1:] ^ n dvd p" shows "poly ((pderiv ^^ k) p) c = 0" proof - from assms(2) obtain q where "p = [:-c, 1:] ^ n * q" by (elim dvdE) also from assms(1) have "poly ((pderiv ^^ k) \) c = 0" by (rule poly_higher_pderiv_aux1) finally show ?thesis . qed lemma poly_higher_pderiv_aux2: fixes c :: "'a :: {idom, semiring_char_0}" shows "poly ((pderiv ^^ n) ([:-c, 1:] ^ n * p)) c = fact n * poly p c" proof (induction n arbitrary: p) case (Suc n p) have "(pderiv ^^ Suc n) ([:- c, 1:] ^ Suc n * p) = (pderiv ^^ n) ([:- c, 1:] ^ Suc n * pderiv p) + (pderiv ^^ n) ([:- c, 1:] ^ n * smult (1 + of_nat n) p)" by (simp del: funpow.simps power_Suc add: funpow_Suc_right pderiv_mult pderiv_power_Suc higher_pderiv_add pderiv_pCons mult_ac) also have "[:- c, 1:] ^ Suc n * pderiv p = [:- c, 1:] ^ n * ([:-c, 1:] * pderiv p)" by (simp add: algebra_simps) finally show ?case by (simp add: Suc.IH del: mult_smult_right power_Suc) qed simp_all lemma poly_higher_pderiv_aux3: fixes c :: "'a :: {idom,semiring_char_0}" assumes "k \ n" shows "\q. poly ((pderiv ^^ k) ([:-c, 1:] ^ n * p)) c = fact n * poly q c" using assms proof (induction k arbitrary: n p) case (Suc k n p) show ?case proof (cases n) fix n' assume n: "n = Suc n'" have "poly ((pderiv ^^ Suc k) ([:-c, 1:] ^ n * p)) c = poly ((pderiv ^^ k) ([:- c, 1:] ^ n * pderiv p)) c + of_nat n * poly ((pderiv ^^ k) ([:-c, 1:] ^ n' * p)) c" by (simp del: funpow.simps power_Suc add: funpow_Suc_right pderiv_power_Suc pderiv_mult n pderiv_pCons higher_pderiv_add mult_ac higher_pderiv_smult) also have "\q1. poly ((pderiv ^^ k) ([:-c, 1:] ^ n * pderiv p)) c = fact n * poly q1 c" using Suc.prems Suc.IH[of n "pderiv p"] by (cases "n' = k") (auto simp: n poly_higher_pderiv_aux1 simp del: power_Suc of_nat_Suc intro: exI[of _ "0::'a poly"]) then obtain q1 where "poly ((pderiv ^^ k) ([:-c, 1:] ^ n * pderiv p)) c = fact n * poly q1 c" .. also from Suc.IH[of n' p] Suc.prems obtain q2 where "poly ((pderiv ^^ k) ([:-c, 1:] ^ n' * p)) c = fact n' * poly q2 c" by (auto simp: n) finally show ?case by (auto intro!: exI[of _ "q1 + q2"] simp: n algebra_simps) qed auto qed auto lemma poly_higher_pderiv_aux3': fixes c :: "'a :: {idom, semiring_char_0}" assumes "k \ n" "[:-c, 1:] ^ n dvd p" shows "fact n dvd poly ((pderiv ^^ k) p) c" proof - from assms(2) obtain q where "p = [:-c, 1:] ^ n * q" by (elim dvdE) with poly_higher_pderiv_aux3[OF assms(1), of c q] show ?thesis by auto qed lemma e_transcendental_aux_bound: obtains C where "C \ 0" "\x. x \ closed_segment 0 (of_nat n) \ norm (\k\{1..n}. (x - of_nat k :: complex)) \ C" proof - let ?f = "\x. (\k\{1..n}. (x - of_nat k))" define C where "C = max 0 (Sup (cmod ` ?f ` closed_segment 0 (of_nat n)))" have "C \ 0" by (simp add: C_def) moreover { fix x :: complex assume "x \ closed_segment 0 (of_nat n)" hence "cmod (?f x) \ Sup ((cmod \ ?f) ` closed_segment 0 (of_nat n))" by (intro cSup_upper bounded_imp_bdd_above compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros) also have "\ \ C" by (simp add: C_def image_comp) finally have "cmod (?f x) \ C" . } ultimately show ?thesis by (rule that) qed theorem e_transcendental_complex: "\ algebraic (exp 1 :: complex)" proof assume "algebraic (exp 1 :: complex)" then obtain q :: "int poly" where q: "q \ 0" "coeff q 0 \ 0" "poly (of_int_poly q) (exp 1 :: complex) = 0" by (elim algebraicE'_nonzero) simp_all define n :: nat where "n = degree q" from q have [simp]: "n \ 0" by (intro notI) (auto simp: n_def elim!: degree_eq_zeroE) define qmax where "qmax = Max (insert 0 (abs ` set (coeffs q)))" have qmax_nonneg [simp]: "qmax \ 0" by (simp add: qmax_def) have qmax: "\coeff q k\ \ qmax" for k by (cases "k \ degree q") (auto simp: qmax_def coeff_eq_0 coeffs_def simp del: upt_Suc intro: Max.coboundedI) obtain C where C: "C \ 0" "\x. x \ closed_segment 0 (of_nat n) \ norm (\k\{1..n}. (x - of_nat k :: complex)) \ C" by (erule e_transcendental_aux_bound) define E where "E = (1 + real n) * real_of_int qmax * real n * exp (real n) / real n" define F where "F = real n * C" have ineq: "fact (p - 1) \ E * F ^ p" if p: "prime p" "p > n" "p > abs (coeff q 0)" for p proof - from p(1) have p_pos: "p > 0" by (simp add: prime_gt_0_nat) define f :: "int poly" where "f = monom 1 (p - 1) * (\k\{1..n}. [:-of_nat k, 1:] ^ p)" have poly_f: "poly (of_int_poly f) x = x ^ (p - 1) * (\k\{1..n}. (x - of_nat k)) ^ p" for x :: complex by (simp add: f_def poly_prod poly_monom prod_power_distrib) define m :: nat where "m = degree f" from p_pos have m: "m = (n + 1) * p - 1" by (simp add: m_def f_def degree_mult_eq degree_monom_eq degree_prod_sum_eq degree_linear_power) define M :: int where "M = (- 1) ^ (n * p) * fact n ^ p" with p have p_not_dvd_M: "\int p dvd M" by (auto simp: M_def prime_elem_int_not_dvd_neg1_power prime_dvd_power_iff prime_gt_0_nat prime_dvd_fact_iff_int prime_dvd_mult_iff) interpret lindemann_weierstrass_aux "of_int_poly f" . define J :: complex where "J = (\k\n. of_int (coeff q k) * I (of_nat k))" define idxs where "idxs = ({..n}\{..m}) - {(0, p - 1)}" hence "J = (\k\n. of_int (coeff q k) * exp 1 ^ k) * (\n\m. of_int (poly ((pderiv ^^ n) f) 0)) - of_int (\k\n. \n\m. coeff q k * poly ((pderiv ^^ n) f) (int k))" by (simp add: J_def I_def algebra_simps sum_subtractf sum_distrib_left m_def exp_of_nat_mult [symmetric]) also have "(\k\n. of_int (coeff q k) * exp 1 ^ k) = poly (of_int_poly q) (exp 1 :: complex)" by (simp add: poly_altdef n_def) also have "\ = 0" by fact finally have "J = of_int (-(\(k,n)\{..n}\{..m}. coeff q k * poly ((pderiv ^^ n) f) (int k)))" by (simp add: sum.cartesian_product) also have "{..n}\{..m} = insert (0, p - 1) idxs" by (auto simp: m idxs_def) also have "-(\(k,n)\\. coeff q k * poly ((pderiv ^^ n) f) (int k)) = - (coeff q 0 * poly ((pderiv ^^ (p - 1)) f) 0) - (\(k, n)\idxs. coeff q k * poly ((pderiv ^^ n) f) (of_nat k))" by (subst sum.insert) (simp_all add: idxs_def) also have "coeff q 0 * poly ((pderiv ^^ (p - 1)) f) 0 = coeff q 0 * M * fact (p - 1)" proof - have "f = [:-0, 1:] ^ (p - 1) * (\k = 1..n. [:- of_nat k, 1:] ^ p)" by (simp add: f_def monom_altdef) also have "poly ((pderiv ^^ (p - 1)) \) 0 = fact (p - 1) * poly (\k = 1..n. [:- of_nat k, 1:] ^ p) 0" by (rule poly_higher_pderiv_aux2) also have "poly (\k = 1..n. [:- of_nat k :: int, 1:] ^ p) 0 = (-1)^(n*p) * fact n ^ p" by (induction n) (simp_all add: prod.nat_ivl_Suc' power_mult_distrib mult_ac power_minus' power_add del: of_nat_Suc) finally show ?thesis by (simp add: mult_ac M_def) qed - also have "\N. (\(k, n)\idxs. coeff q k * poly ((pderiv ^^ n) f) (int k)) = fact p * N" + also obtain N where "(\(k, n)\idxs. coeff q k * poly ((pderiv ^^ n) f) (int k)) = fact p * N" proof - have "\(k, n)\idxs. fact p dvd poly ((pderiv ^^ n) f) (of_nat k)" proof clarify fix k j assume idxs: "(k, j) \ idxs" then consider "k = 0" "j < p - 1" | "k = 0" "j > p - 1" | "k \ 0" "j < p" | "k \ 0" "j \ p" by (fastforce simp: idxs_def) thus "fact p dvd poly ((pderiv ^^ j) f) (of_nat k)" proof cases case 1 thus ?thesis by (simp add: f_def poly_higher_pderiv_aux1' monom_altdef) next case 2 thus ?thesis by (simp add: f_def poly_higher_pderiv_aux3' monom_altdef fact_dvd_poly_higher_pderiv_aux') next case 3 thus ?thesis unfolding f_def by (subst poly_higher_pderiv_aux1'[of _ p]) (insert idxs, auto simp: idxs_def intro!: dvd_mult) next case 4 thus ?thesis unfolding f_def by (intro poly_higher_pderiv_aux3') (insert idxs, auto intro!: dvd_mult simp: idxs_def) qed qed hence "fact p dvd (\(k, n)\idxs. coeff q k * poly ((pderiv ^^ n) f) (int k))" by (auto intro!: dvd_sum dvd_mult simp del: of_int_fact) - thus ?thesis by (blast elim: dvdE) + with that show thesis + by blast qed - then guess N .. note N = this also from p have "- (coeff q 0 * M * fact (p - 1)) - fact p * N = - fact (p - 1) * (coeff q 0 * M + p * N)" by (subst fact_reduce[of p]) (simp_all add: algebra_simps) finally have J: "J = -of_int (fact (p - 1) * (coeff q 0 * M + p * N))" by simp from p q(2) have "\p dvd coeff q 0 * M + p * N" by (auto simp: dvd_add_left_iff p_not_dvd_M prime_dvd_fact_iff_int prime_dvd_mult_iff dest: dvd_imp_le_int) hence "coeff q 0 * M + p * N \ 0" by (intro notI) simp_all hence "abs (coeff q 0 * M + p * N) \ 1" by simp hence "norm (of_int (coeff q 0 * M + p * N) :: complex) \ 1" by (simp only: norm_of_int) hence "fact (p - 1) * \ \ fact (p - 1) * 1" by (intro mult_left_mono) simp_all hence J_lower: "norm J \ fact (p - 1)" unfolding J norm_minus_cancel of_int_mult of_int_fact by (simp add: norm_mult) have "norm J \ (\k\n. norm (of_int (coeff q k) * I (of_nat k)))" unfolding J_def by (rule norm_sum) also have "\ \ (\k\n. of_int qmax * (real n * exp (real n) * real n ^ (p - 1) * C ^ p))" proof (intro sum_mono) fix k assume k: "k \ {..n}" have "n > 0" by (rule ccontr) simp { fix x :: complex assume x: "x \ closed_segment 0 (of_nat k)" then obtain t where t: "t \ 0" "t \ 1" "x = of_real t * of_nat k" by (auto simp: closed_segment_def scaleR_conv_of_real) hence "norm x = t * real k" by (simp add: norm_mult) also from \t \ 1\ k have *: "\ \ 1 * real n" by (intro mult_mono) simp_all finally have x': "norm x \ real n" by simp from t \n > 0\ * have x'': "x \ closed_segment 0 (of_nat n)" by (auto simp: closed_segment_def scaleR_conv_of_real field_simps intro!: exI[of _ "t * real k / real n"] ) have "norm (poly (of_int_poly f) x) = norm x ^ (p - 1) * cmod (\i = 1..n. x - i) ^ p" by (simp add: poly_f norm_mult norm_power) also from x x' x'' have "\ \ of_nat n ^ (p - 1) * C ^ p" by (intro mult_mono C power_mono) simp_all finally have "norm (poly (of_int_poly f) x) \ real n ^ (p - 1) * C ^ p" . } note A = this have "norm (I (of_nat k)) \ cmod (of_nat k) * exp (cmod (of_nat k)) * (of_nat n ^ (p - 1) * C ^ p)" by (intro lindemann_weierstrass_integral_bound[OF _ A] C mult_nonneg_nonneg zero_le_power) auto also have "\ \ cmod (of_nat n) * exp (cmod (of_nat n)) * (of_nat n ^ (p - 1) * C ^ p)" using k by (intro mult_mono zero_le_power mult_nonneg_nonneg C) simp_all finally show "cmod (of_int (coeff q k) * I (of_nat k)) \ of_int qmax * (real n * exp (real n) * real n ^ (p - 1) * C ^ p)" unfolding norm_mult by (intro mult_mono) (simp_all add: qmax of_int_abs [symmetric] del: of_int_abs) qed also have "\ = E * F ^ p" using p_pos by (simp add: power_diff power_mult_distrib E_def F_def) finally show "fact (p - 1) \ E * F ^ p" using J_lower by linarith qed have "(\n. E * F * F ^ (n - 1) / fact (n - 1)) \ 0" (is ?P) by (intro filterlim_compose[OF power_over_fact_tendsto_0' filterlim_minus_nat_at_top]) also have "?P \ (\n. E * F ^ n / fact (n - 1)) \ 0" by (intro filterlim_cong refl eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) (auto simp: power_Suc [symmetric] simp del: power_Suc) finally have "eventually (\n. E * F ^ n / fact (n - 1) < 1) at_top" by (rule order_tendstoD) simp_all hence "eventually (\n. E * F ^ n < fact (n - 1)) at_top" by eventually_elim simp then obtain P where P: "\n. n \ P \ E * F ^ n < fact (n - 1)" by (auto simp: eventually_at_top_linorder) have "\p. prime p \ p > Max {nat (abs (coeff q 0)), n, P}" by (rule bigger_prime) then obtain p where "prime p" "p > Max {nat (abs (coeff q 0)), n, P}" by blast hence "int p > abs (coeff q 0)" "p > n" "p \ P" by auto with ineq[of p] \prime p\ have "fact (p - 1) \ E * F ^ p" by simp moreover from \p \ P\ have "fact (p - 1) > E * F ^ p" by (rule P) ultimately show False by linarith qed corollary e_transcendental_real: "\ algebraic (exp 1 :: real)" proof - have "\algebraic (exp 1 :: complex)" by (rule e_transcendental_complex) also have "(exp 1 :: complex) = of_real (exp 1)" using exp_of_real[of 1] by simp also have "algebraic \ \ algebraic (exp 1 :: real)" by simp finally show ?thesis . qed end diff --git a/thys/Error_Function/Error_Function.thy b/thys/Error_Function/Error_Function.thy --- a/thys/Error_Function/Error_Function.thy +++ b/thys/Error_Function/Error_Function.thy @@ -1,655 +1,657 @@ (* File: Error_Function.thy Author: Manuel Eberl, TU München The complex and real error function (most results are on the real error function though) *) section \The complex and real error function\ theory Error_Function imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Library.Landau_Symbols" begin subsection \Auxiliary Facts\ lemma tendsto_sandwich_mono: assumes "(\n. f (real n)) \ (c::real)" assumes "eventually (\x. \y z. x \ y \ y \ z \ f y \ f z) at_top" shows "(f \ c) at_top" proof (rule tendsto_sandwich) from assms(2) obtain C where C: "\x y. C \ x \ x \ y \ f x \ f y" by (auto simp: eventually_at_top_linorder) show "eventually (\x. f (real (nat \x\)) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) show "eventually (\x. f (real (Suc (nat \x\))) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) qed (auto intro!: filterlim_compose[OF assms(1)] filterlim_compose[OF filterlim_nat_sequentially] filterlim_compose[OF filterlim_Suc] filterlim_floor_sequentially simp del: of_nat_Suc) lemma tendsto_sandwich_antimono: assumes "(\n. f (real n)) \ (c::real)" assumes "eventually (\x. \y z. x \ y \ y \ z \ f y \ f z) at_top" shows "(f \ c) at_top" proof (rule tendsto_sandwich) from assms(2) obtain C where C: "\x y. C \ x \ x \ y \ f x \ f y" by (auto simp: eventually_at_top_linorder) show "eventually (\x. f (real (nat \x\)) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) show "eventually (\x. f (real (Suc (nat \x\))) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) qed (auto intro!: filterlim_compose[OF assms(1)] filterlim_compose[OF filterlim_nat_sequentially] filterlim_compose[OF filterlim_Suc] filterlim_floor_sequentially simp del: of_nat_Suc) lemma has_bochner_integral_completion [intro]: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "has_bochner_integral M f I \ has_bochner_integral (completion M) f I" by (auto simp: has_bochner_integral_iff integrable_completion integral_completion borel_measurable_integrable) lemma has_bochner_integral_imp_has_integral: "has_bochner_integral lebesgue (\x. indicator S x *\<^sub>R f x) I \ (f has_integral (I :: 'b :: euclidean_space)) S" using has_integral_set_lebesgue[of S f] by (simp add: has_bochner_integral_iff set_integrable_def set_lebesgue_integral_def) lemma has_bochner_integral_imp_has_integral': "has_bochner_integral lborel (\x. indicator S x *\<^sub>R f x) I \ (f has_integral (I :: 'b :: euclidean_space)) S" by (intro has_bochner_integral_imp_has_integral has_bochner_integral_completion) lemma has_bochner_integral_erf_aux: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2)) (sqrt pi / 2)" proof - let ?pI = "\f. (\\<^sup>+s. f (s::real) * indicator {0..} s \lborel)" let ?gauss = "\x. exp (- x\<^sup>2)" let ?I = "indicator {0<..} :: real \ real" let ?ff= "\x s. x * exp (- x\<^sup>2 * (1 + s\<^sup>2)) :: real" have *: "?pI ?gauss = (\\<^sup>+x. ?gauss x * ?I x \lborel)" by (intro nn_integral_cong_AE AE_I[where N="{0}"]) (auto split: split_indicator) have "?pI ?gauss * ?pI ?gauss = (\\<^sup>+x. \\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \lborel \lborel)" by (auto simp: nn_integral_cmult[symmetric] nn_integral_multc[symmetric] * ennreal_mult[symmetric] intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+x. \\<^sup>+s. ?ff x s * ?I s * ?I x \lborel \lborel)" proof (rule nn_integral_cong, cases) fix x :: real assume "x \ 0" then show "(\\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \lborel) = (\\<^sup>+s. ?ff x s * ?I s * ?I x \lborel)" by (subst nn_integral_real_affine[where t="0" and c="x"]) (auto simp: mult_exp_exp nn_integral_cmult[symmetric] field_simps zero_less_mult_iff ennreal_mult[symmetric] intro!: nn_integral_cong split: split_indicator) qed simp also have "... = \\<^sup>+s. \\<^sup>+x. ?ff x s * ?I s * ?I x \lborel \lborel" by (rule lborel_pair.Fubini'[symmetric]) auto also have "... = ?pI (\s. ?pI (\x. ?ff x s))" by (rule nn_integral_cong_AE) (auto intro!: nn_integral_cong_AE AE_I[where N="{0}"] split: split_indicator_asm) also have "\ = ?pI (\s. ennreal (1 / (2 * (1 + s\<^sup>2))))" proof (intro nn_integral_cong ennreal_mult_right_cong) fix s :: real show "?pI (\x. ?ff x s) = ennreal (1 / (2 * (1 + s\<^sup>2)))" proof (subst nn_integral_FTC_atLeast) have "((\a. - (exp (- ((1 + s\<^sup>2) * a\<^sup>2)) / (2 + 2 * s\<^sup>2))) \ (- (0 / (2 + 2 * s\<^sup>2)))) at_top" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_compose[OF filterlim_uminus_at_bot_at_top] filterlim_tendsto_pos_mult_at_top) (auto intro!: filterlim_at_top_mult_at_top[OF filterlim_ident filterlim_ident] simp: add_pos_nonneg power2_eq_square add_nonneg_eq_0_iff) then show "((\a. - (exp (- a\<^sup>2 - s\<^sup>2 * a\<^sup>2) / (2 + 2 * s\<^sup>2))) \ 0) at_top" by (simp add: field_simps) qed (auto intro!: derivative_eq_intros simp: field_simps add_nonneg_eq_0_iff) qed also have "... = ennreal (pi / 4)" proof (subst nn_integral_FTC_atLeast) show "((\a. arctan a / 2) \ (pi / 2) / 2 ) at_top" by (intro tendsto_intros) (simp_all add: tendsto_arctan_at_top) qed (auto intro!: derivative_eq_intros simp: add_nonneg_eq_0_iff field_simps power2_eq_square) finally have "?pI ?gauss^2 = pi / 4" by (simp add: power2_eq_square) then have "?pI ?gauss = sqrt (pi / 4)" using power_eq_iff_eq_base[of 2 "enn2real (?pI ?gauss)" "sqrt (pi / 4)"] by (cases "?pI ?gauss") (auto simp: power2_eq_square ennreal_mult[symmetric] ennreal_top_mult) also have "?pI ?gauss = (\\<^sup>+x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2) \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also have "sqrt (pi / 4) = sqrt pi / 2" by (simp add: real_sqrt_divide) finally show ?thesis by (rule has_bochner_integral_nn_integral[rotated 3]) auto qed lemma has_integral_erf_aux: "((\t::real. exp (-t\<^sup>2)) has_integral (sqrt pi / 2)) {0..}" by (intro has_bochner_integral_imp_has_integral' has_bochner_integral_erf_aux) lemma contour_integrable_on_linepath_neg_exp_squared [simp, intro]: "(\t. exp (-(t^2))) contour_integrable_on linepath 0 z" by (auto intro!: contour_integrable_continuous_linepath continuous_intros) lemma holomorphic_on_chain: "g holomorphic_on t \ f holomorphic_on s \ f ` s \ t \ (\x. g (f x)) holomorphic_on s" using holomorphic_on_compose_gen[of f s g t] by (simp add: o_def) lemma holomorphic_on_chain_UNIV: "g holomorphic_on UNIV \ f holomorphic_on s\ (\x. g (f x)) holomorphic_on s" using holomorphic_on_compose_gen[of f s g UNIV] by (simp add: o_def) lemmas holomorphic_on_exp' [holomorphic_intros] = holomorphic_on_exp [THEN holomorphic_on_chain_UNIV] lemma leibniz_rule_field_derivative_real: fixes f::"'a::{real_normed_field, banach} \ real \ 'a" assumes fx: "\x t. x \ U \ t \ {a..b} \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on {a..b}" assumes cont_fx: "continuous_on (U \ {a..b}) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral {a..b} (f x)) has_field_derivative integral {a..b} (fx x0)) (at x0 within U)" using leibniz_rule_field_derivative[of U a b f fx x0] assms by simp lemma has_vector_derivative_linepath_within [derivative_intros]: assumes [derivative_intros]: "(f has_vector_derivative f') (at x within S)" "(g has_vector_derivative g') (at x within S)" "(h has_real_derivative h') (at x within S)" shows "((\x. linepath (f x) (g x) (h x)) has_vector_derivative (1 - h x) *\<^sub>R f' + h x *\<^sub>R g' - h' *\<^sub>R (f x - g x)) (at x within S)" unfolding linepath_def [abs_def] by (auto intro!: derivative_eq_intros simp: field_simps scaleR_diff_right) lemma has_field_derivative_linepath_within [derivative_intros]: assumes [derivative_intros]: "(f has_field_derivative f') (at x within S)" "(g has_field_derivative g') (at x within S)" "(h has_real_derivative h') (at x within S)" shows "((\x. linepath (f x) (g x) (h x)) has_field_derivative (1 - h x) *\<^sub>R f' + h x *\<^sub>R g' - h' *\<^sub>R (f x - g x)) (at x within S)" unfolding linepath_def [abs_def] by (auto intro!: derivative_eq_intros simp: field_simps scaleR_diff_right) lemma continuous_on_linepath' [continuous_intros]: assumes [continuous_intros]: "continuous_on A f" "continuous_on A g" "continuous_on A h" shows "continuous_on A (\x. linepath (f x) (g x) (h x))" using assms unfolding linepath_def by (auto intro!: continuous_intros) lemma contour_integral_has_field_derivative: assumes A: "open A" "convex A" "a \ A" "z \ A" assumes integrable: "\z. z \ A \ f contour_integrable_on linepath a z" assumes holo: "f holomorphic_on A" shows "((\z. contour_integral (linepath a z) f) has_field_derivative f z) (at z within B)" proof - have "(f has_field_derivative deriv f z) (at z)" if "z \ A" for z using that assms by (auto intro!: holomorphic_derivI) note [derivative_intros] = DERIV_chain2[OF this] note [continuous_intros] = continuous_on_compose2[OF holomorphic_on_imp_continuous_on [OF holo]] continuous_on_compose2[OF holomorphic_on_imp_continuous_on [OF holomorphic_deriv[OF holo]]] have [derivative_intros]: "((\x. linepath a x t) has_field_derivative of_real t) (at x within A)" for t x by (auto simp: linepath_def scaleR_conv_of_real intro!: derivative_eq_intros) have *: "linepath a b t \ A" if "a \ A" "b \ A" "t \ {0..1}" for a b t using that linepath_in_convex_hull[of a A b t] \convex A\ by (simp add: hull_same) have "((\z. integral {0..1} (\x. f (linepath a z x)) * (z - a)) has_field_derivative integral {0..1} (\t. deriv f (linepath a z t) * of_real t) * (z - a) + integral {0..1} (\x. f (linepath a z x))) (at z within A)" (is "(_ has_field_derivative ?I) _") by (rule derivative_eq_intros leibniz_rule_field_derivative_real)+ (insert assms, auto intro!: derivative_eq_intros leibniz_rule_field_derivative_real integrable_continuous_real continuous_intros simp: split_beta scaleR_conv_of_real *) also have "(\z. integral {0..1} (\x. f (linepath a z x)) * (z - a)) = (\z. contour_integral (linepath a z) f)" by (simp add: contour_integral_integral) also have "?I = integral {0..1} (\x. deriv f (linepath a z x) * of_real x * (z - a) + f (linepath a z x))" (is "_ = integral _ ?g") by (subst integral_mult_left [symmetric], subst integral_add [symmetric]) (insert assms, auto intro!: integrable_continuous_real continuous_intros simp: *) also have "(?g has_integral of_real 1 * f (linepath a z 1) - of_real 0 * f (linepath a z 0)) {0..1}" using * A by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: linepath_def scaleR_conv_of_real) hence "integral {0..1} ?g = f (linepath a z 1)" by (simp add: has_integral_iff) also have "linepath a z 1 = z" by (simp add: linepath_def) also from \z \ A\ and \open A\ have "at z within A = at z" by (rule at_within_open) finally show ?thesis by (rule DERIV_subset) simp_all qed subsection \Definition of the error function\ definition erf_coeffs :: "nat \ real" where "erf_coeffs n = (if odd n then 2 / sqrt pi * (-1) ^ (n div 2) / (of_nat n * fact (n div 2)) else 0)" lemma summable_erf: fixes z :: "'a :: {real_normed_div_algebra, banach}" shows "summable (\n. of_real (erf_coeffs n) * z ^ n)" proof - define b where "b = (\n. 2 / sqrt pi * (if odd n then inverse (fact (n div 2)) else 0))" show ?thesis proof (rule summable_comparison_test[OF exI[of _ 1]], clarify) fix n :: nat assume n: "n \ 1" hence "norm (of_real (erf_coeffs n) * z ^ n) \ b n * norm z ^ n" unfolding norm_mult norm_power erf_coeffs_def b_def by (intro mult_right_mono) (auto simp: field_simps norm_divide abs_mult) thus "norm (of_real (erf_coeffs n) * z ^ n) \ b n * norm z ^ n" by (simp add: mult_ac) next have "summable (\n. (norm z * 2 / sqrt pi) * (inverse (fact n) * norm z ^ (2*n)))" (is "summable ?c") unfolding power_mult by (intro summable_mult summable_exp) also have "?c = (\n. b (2*n+1) * norm z ^ (2*n+1))" unfolding b_def by (auto simp: fun_eq_iff b_def) also have "summable \ \ summable (\n. b n * norm z ^ n)" using summable_mono_reindex [of "\n. 2*n+1"] by (intro summable_mono_reindex [of "\n. 2*n+1"]) (auto elim!: oddE simp: strict_mono_def b_def) finally show \ . qed qed definition erf :: "('a :: {real_normed_field, banach}) \ 'a" where "erf z = (\n. of_real (erf_coeffs n) * z ^ n)" lemma erf_converges: "(\n. of_real (erf_coeffs n) * z ^ n) sums erf z" using summable_erf by (simp add: sums_iff erf_def) lemma erf_0 [simp]: "erf 0 = 0" unfolding erf_def powser_zero by (simp add: erf_coeffs_def) lemma erf_minus [simp]: "erf (-z) = - erf z" unfolding erf_def by (subst suminf_minus [OF summable_erf, symmetric], rule suminf_cong) (simp_all add: erf_coeffs_def) lemma erf_of_real [simp]: "erf (of_real x) = of_real (erf x)" unfolding erf_def using summable_erf[of x] by (subst suminf_of_real) (simp_all add: summable_erf) lemma of_real_erf_numeral [simp]: "of_real (erf (numeral n)) = erf (numeral n)" by (simp only: erf_of_real [symmetric] of_real_numeral) lemma of_real_erf_1 [simp]: "of_real (erf 1) = erf 1" by (simp only: erf_of_real [symmetric] of_real_1) lemma erf_has_field_derivative: "(erf has_field_derivative of_real (2 / sqrt pi) * exp (-(z^2))) (at z within A)" proof - define a' where "a' = (\n. 2 / sqrt pi * (if even n then (- 1) ^ (n div 2) / fact (n div 2) else 0))" have "(erf has_field_derivative (\n. diffs (\n. of_real (erf_coeffs n)) n * z ^ n)) (at z)" using summable_erf unfolding erf_def by (rule termdiffs_strong_converges_everywhere) also have "diffs (\n. of_real (erf_coeffs n)) = (\n. of_real (a' n) :: 'a)" by (simp add: erf_coeffs_def a'_def diffs_def fun_eq_iff del: of_nat_Suc) hence "(\n. diffs (\n. of_real (erf_coeffs n)) n * z ^ n) = (\n. of_real (a' n) * z ^ n)" by simp also have "\ = (\n. of_real (a' (2*n)) * z ^ (2*n))" by (intro suminf_mono_reindex [symmetric]) (auto simp: strict_mono_def a'_def elim!: evenE) also have "(\n. of_real (a' (2*n)) * z ^ (2*n)) = (\n. of_real (2 / sqrt pi) * (inverse (fact n) * (-(z^2))^n))" by (simp add: fun_eq_iff power_mult [symmetric] a'_def field_simps power_minus') also have "suminf \ = of_real (2 / sqrt pi) * exp (-(z^2))" by (subst suminf_mult, intro summable_exp) (auto simp: field_simps scaleR_conv_of_real exp_def) finally show ?thesis by (rule DERIV_subset) simp_all qed lemmas erf_has_field_derivative' [derivative_intros] = erf_has_field_derivative [THEN DERIV_chain2] lemma erf_continuous_on: "continuous_on A erf" by (rule DERIV_continuous_on erf_has_field_derivative)+ lemma continuous_on_compose2_UNIV: "continuous_on UNIV g \ continuous_on s f \ continuous_on s (\x. g (f x))" by (rule continuous_on_compose2[of UNIV g s f]) simp_all lemmas erf_continuous_on' [continuous_intros] = erf_continuous_on [THEN continuous_on_compose2_UNIV] lemma erf_continuous [continuous_intros]: "continuous (at x within A) erf" by (rule continuous_within_subset[OF _ subset_UNIV]) (insert erf_continuous_on[of UNIV], auto simp: continuous_on_eq_continuous_at) lemmas erf_continuous' [continuous_intros] = continuous_within_compose2[OF _ erf_continuous] lemmas tendsto_erf [tendsto_intros] = isCont_tendsto_compose[OF erf_continuous] lemma erf_cnj [simp]: "erf (cnj z) = cnj (erf z)" proof - interpret bounded_linear cnj by (rule bounded_linear_cnj) from suminf[OF summable_erf] show ?thesis by (simp add: erf_def erf_coeffs_def) qed lemma integral_exp_minus_squared_real: assumes "a \ b" shows "((\t. exp (-(t^2))) has_integral (sqrt pi / 2 * (erf b - erf a))) {a..b}" proof - have "((\t. exp (-(t^2))) has_integral (sqrt pi / 2 * erf b - sqrt pi / 2 * erf a)) {a..b}" using assms by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric]) thus ?thesis by (simp add: field_simps) qed lemma erf_real_altdef_nonneg: "x \ 0 \ erf (x::real) = 2 / sqrt pi * integral {0..x} (\t. exp (-(t^2)))" using integral_exp_minus_squared_real[of 0 x] by (simp add: has_integral_iff field_simps) lemma erf_real_altdef_nonpos: "x \ 0 \ erf (x::real) = -2 / sqrt pi * integral {0..-x} (\t. exp (-(t^2)))" using erf_real_altdef_nonneg[of "-x"] by simp lemma less_imp_erf_real_less: assumes "a < (b::real)" shows "erf a < erf b" proof - from assms have "\z. z > a \ z < b \ erf b - erf a = (b - a) * (2 / sqrt pi * exp (- z\<^sup>2))" by (intro MVT2) (auto intro!: derivative_eq_intros simp: field_simps) - then guess z by (elim exE conjE) note z = this - note z(3) + then obtain z where z: "a < z" "z < b" + and erf: "erf b - erf a = (b - a) * (2 / sqrt pi * exp (- z\<^sup>2))" + by blast + note erf also from assms have "(b - a) * (2 / sqrt pi * exp (- z\<^sup>2)) > 0" by (intro mult_pos_pos divide_pos_pos) simp_all finally show ?thesis by simp qed lemma le_imp_erf_real_le: "a \ (b::real) \ erf a \ erf b" by (cases "a < b") (auto dest: less_imp_erf_real_less) lemma erf_real_less_cancel [simp]: "(erf (a :: real) < erf b) \ a < b" using less_imp_erf_real_less[of a b] less_imp_erf_real_less[of b a] by (cases a b rule: linorder_cases) simp_all lemma erf_real_eq_iff [simp]: "erf (a::real) = erf b \ a = b" by (cases a b rule: linorder_cases) (auto dest: less_imp_erf_real_less) lemma erf_real_le_cancel [simp]: "(erf (a :: real) \ erf b) \ a \ b" by (cases a b rule: linorder_cases) (auto simp: less_eq_real_def) lemma inj_on_erf_real [intro]: "inj_on (erf :: real \ real) A" by (auto simp: inj_on_def) lemma strict_mono_erf_real [intro]: "strict_mono (erf :: real \ real)" by (auto simp: strict_mono_def) lemma mono_erf_real [intro]: "mono (erf :: real \ real)" by (auto simp: mono_def) lemma erf_real_ge_0_iff [simp]: "erf (x::real) \ 0 \ x \ 0" using erf_real_le_cancel[of 0 x] unfolding erf_0 . lemma erf_real_le_0_iff [simp]: "erf (x::real) \ 0 \ x \ 0" using erf_real_le_cancel[of x 0] unfolding erf_0 . lemma erf_real_gt_0_iff [simp]: "erf (x::real) > 0 \ x > 0" using erf_real_less_cancel[of 0 x] unfolding erf_0 . lemma erf_real_less_0_iff [simp]: "erf (x::real) < 0 \ x < 0" using erf_real_less_cancel[of x 0] unfolding erf_0 . lemma erf_at_top [tendsto_intros]: "((erf :: real \ real) \ 1) at_top" proof - have *: "(\n. {0..real n}) = {0..}" by (auto intro!: real_nat_ceiling_ge) let ?f = "\t::real. exp (-t\<^sup>2)" have "(\n. set_lebesgue_integral lborel {0..real n} ?f) \ set_lebesgue_integral lborel (\n. {0..real n}) ?f" using has_bochner_integral_erf_aux by (intro set_integral_cont_up ) (insert *, auto simp: incseq_def has_bochner_integral_iff set_integrable_def) also note * also have "(\n. set_lebesgue_integral lborel {0..real n} ?f) = (\n. integral {0..real n} ?f)" proof - have "\n. set_integrable lborel {0..real n} (\x. exp (- x\<^sup>2))" unfolding set_integrable_def by (intro borel_integrable_compact) (auto intro!: continuous_intros) then show ?thesis by (intro set_borel_integral_eq_integral ext) qed also have "\ = (\n. sqrt pi / 2 * erf (real n))" by (simp add: erf_real_altdef_nonneg) also have "set_lebesgue_integral lborel {0..} ?f = sqrt pi / 2" using has_bochner_integral_erf_aux by (simp add: has_bochner_integral_iff set_lebesgue_integral_def) finally have "(\n. 2 / sqrt pi * (sqrt pi / 2 * erf (real n))) \ (2 / sqrt pi) * (sqrt pi / 2)" by (intro tendsto_intros) hence "(\n. erf (real n)) \ 1" by simp thus ?thesis by (rule tendsto_sandwich_mono) auto qed lemma erf_at_bot [tendsto_intros]: "((erf :: real \ real) \ -1) at_bot" by (simp add: filterlim_at_bot_mirror tendsto_minus_cancel_left erf_at_top) lemmas tendsto_erf_at_top [tendsto_intros] = filterlim_compose[OF erf_at_top] lemmas tendsto_erf_at_bot [tendsto_intros] = filterlim_compose[OF erf_at_bot] subsection \The complimentary error function\ definition erfc where "erfc z = 1 - erf z" lemma erf_conv_erfc: "erf z = 1 - erfc z" by (simp add: erfc_def) lemma erfc_0 [simp]: "erfc 0 = 1" by (simp add: erfc_def) lemma erfc_minus: "erfc (-z) = 2 - erfc z" by (simp add: erfc_def) lemma erfc_of_real [simp]: "erfc (of_real x) = of_real (erfc x)" by (simp add: erfc_def) lemma of_real_erfc_numeral [simp]: "of_real (erfc (numeral n)) = erfc (numeral n)" by (simp add: erfc_def) lemma of_real_erfc_1 [simp]: "of_real (erfc 1) = erfc 1" by (simp add: erfc_def) lemma less_imp_erfc_real_less: "a < (b::real) \ erfc a > erfc b" by (simp add: erfc_def) lemma le_imp_erfc_real_le: "a \ (b::real) \ erfc a \ erfc b" by (simp add: erfc_def) lemma erfc_real_less_cancel [simp]: "(erfc (a :: real) < erfc b) \ a > b" by (simp add: erfc_def) lemma erfc_real_eq_iff [simp]: "erfc (a::real) = erfc b \ a = b" by (simp add: erfc_def) lemma erfc_real_le_cancel [simp]: "(erfc (a :: real) \ erfc b) \ a \ b" by (simp add: erfc_def) lemma inj_on_erfc_real [intro]: "inj_on (erfc :: real \ real) A" by (auto simp: inj_on_def) lemma antimono_erfc_real [intro]: "antimono (erfc :: real \ real)" by (auto simp: antimono_def) lemma erfc_real_ge_0_iff [simp]: "erfc (x::real) \ 1 \ x \ 0" by (simp add: erfc_def) lemma erfc_real_le_0_iff [simp]: "erfc (x::real) \ 1 \ x \ 0" by (simp add: erfc_def) lemma erfc_real_gt_0_iff [simp]: "erfc (x::real) > 1 \ x < 0" by (simp add: erfc_def) lemma erfc_real_less_0_iff [simp]: "erfc (x::real) < 1 \ x > 0" by (simp add: erfc_def) lemma erfc_has_field_derivative: "(erfc has_field_derivative -of_real (2 / sqrt pi) * exp (-(z^2))) (at z within A)" unfolding erfc_def [abs_def] by (auto intro!: derivative_eq_intros) lemmas erfc_has_field_derivative' [derivative_intros] = erfc_has_field_derivative [THEN DERIV_chain2] lemma erfc_continuous_on: "continuous_on A erfc" by (rule DERIV_continuous_on erfc_has_field_derivative)+ lemmas erfc_continuous_on' [continuous_intros] = erfc_continuous_on [THEN continuous_on_compose2_UNIV] lemma erfc_continuous [continuous_intros]: "continuous (at x within A) erfc" by (rule continuous_within_subset[OF _ subset_UNIV]) (insert erfc_continuous_on[of UNIV], auto simp: continuous_on_eq_continuous_at) lemmas erfc_continuous' [continuous_intros] = continuous_within_compose2[OF _ erfc_continuous] lemmas tendsto_erfc [tendsto_intros] = isCont_tendsto_compose[OF erfc_continuous] lemma erfc_at_top [tendsto_intros]: "((erfc :: real \ real) \ 0) at_top" unfolding erfc_def [abs_def] by (auto intro!: tendsto_eq_intros) lemma erfc_at_bot [tendsto_intros]: "((erfc :: real \ real) \ 2) at_bot" unfolding erfc_def [abs_def] by (auto intro!: tendsto_eq_intros) lemmas tendsto_erfc_at_top [tendsto_intros] = filterlim_compose[OF erfc_at_top] lemmas tendsto_erfc_at_bot [tendsto_intros] = filterlim_compose[OF erfc_at_bot] lemma integrable_exp_minus_squared: assumes "A \ {0..}" "A \ sets lborel" shows "set_integrable lborel A (\t::real. exp (-t\<^sup>2))" (is ?thesis1) and "(\t::real. exp (-t\<^sup>2)) integrable_on A" (is ?thesis2) proof - show ?thesis1 by (rule set_integrable_subset[of _ "{0..}"]) (insert assms has_bochner_integral_erf_aux, auto simp: has_bochner_integral_iff set_integrable_def) thus ?thesis2 by (rule set_borel_integral_eq_integral) qed lemma assumes "x \ 0" shows erfc_real_altdef_nonneg: "erfc x = 2 / sqrt pi * integral {x..} (\t. exp (-t\<^sup>2))" and has_integral_erfc: "((\t. exp (-t\<^sup>2)) has_integral (sqrt pi / 2 * erfc x)) {x..}" proof - let ?f = "\t::real. exp (-t\<^sup>2)" have int: "set_integrable lborel {0..} ?f" using has_bochner_integral_erf_aux by (simp add: has_bochner_integral_iff set_integrable_def) from assms have *: "{(0::real)..} = {0..x} \ {x..}" by auto have "set_lebesgue_integral lborel ({0..x} \ {x..}) ?f = set_lebesgue_integral lborel {0..x} ?f + set_lebesgue_integral lborel {x..} ?f" by (subst set_integral_Un_AE; (rule set_integrable_subset[OF int])?) (insert assms AE_lborel_singleton[of x], auto elim!: eventually_mono) also note * [symmetric] also have "set_lebesgue_integral lborel {0..} ?f = sqrt pi / 2" using has_bochner_integral_erf_aux by (simp add: has_bochner_integral_iff set_lebesgue_integral_def) also have "set_lebesgue_integral lborel {0..x} ?f = sqrt pi / 2 * erf x" by (subst set_borel_integral_eq_integral(2)[OF set_integrable_subset[OF int]]) (insert assms, auto simp: erf_real_altdef_nonneg) also have "set_lebesgue_integral lborel {x..} ?f = integral {x..} ?f" by (subst set_borel_integral_eq_integral(2)[OF set_integrable_subset[OF int]]) (insert assms, auto) finally show "erfc x = 2 / sqrt pi * integral {x..} ?f" by (simp add: field_simps erfc_def) with integrable_exp_minus_squared(2)[of "{x..}"] assms show "(?f has_integral (sqrt pi / 2 * erfc x)) {x..}" by (simp add: has_integral_iff) qed lemma erfc_real_gt_0 [simp, intro]: "erfc (x::real) > 0" proof (cases "x \ 0") case True have "0 < integral {x..x+1} (\t. exp (-(x+1)\<^sup>2))" by simp also from True have "\ \ integral {x..x+1} (\t. exp (-t\<^sup>2))" by (intro integral_le) (auto intro!: integrable_continuous_real continuous_intros power_mono) also have "\ \ sqrt pi / 2 * erfc x" by (rule has_integral_subset_le[OF _ integrable_integral has_integral_erfc]) (auto intro!: integrable_continuous_real continuous_intros True) finally have "sqrt pi / 2 * erfc x > 0" . hence "\ * (2 / sqrt pi) > 0" by (rule mult_pos_pos) simp_all thus "erfc x > 0" by simp next case False have "0 \ (1::real)" by simp also from False have "\ < erfc x" by simp finally show ?thesis . qed lemma erfc_real_less_2 [intro]: "erfc (x::real) < 2" using erfc_real_gt_0[of "-x"] unfolding erfc_minus by simp lemma erf_real_gt_neg1 [intro]: "erf (x::real) > -1" using erfc_real_less_2[of x] unfolding erfc_def by simp lemma erf_real_less_1 [intro]: "erf (x::real) < 1" using erfc_real_gt_0[of x] unfolding erfc_def by simp lemma erfc_cnj [simp]: "erfc (cnj z) = cnj (erfc z)" by (simp add: erfc_def) subsection \Specific facts about the complex case\ lemma erf_complex_altdef: "erf z = of_real (2 / sqrt pi) * contour_integral (linepath 0 z) (\t. exp (-(t^2)))" proof - define A where "A = (\z. contour_integral (linepath 0 z) (\t. exp (-(t^2))))" have [derivative_intros]: "(A has_field_derivative exp (-(z^2))) (at z)" for z :: complex unfolding A_def by (rule contour_integral_has_field_derivative[where A = UNIV]) (auto intro!: holomorphic_intros) have "erf z - 2 / sqrt pi * A z = erf 0 - 2 / sqrt pi * A 0" by (rule has_derivative_zero_unique [where f = "\z. erf z - 2 / sqrt pi * A z" and s = UNIV]) (auto intro!: has_field_derivative_imp_has_derivative derivative_eq_intros) also have "A 0 = 0" by (simp only: A_def contour_integral_trivial) finally show ?thesis unfolding A_def by (simp add: algebra_simps) qed lemma erf_holomorphic_on: "erf holomorphic_on A" by (auto simp: holomorphic_on_def field_differentiable_def intro!: erf_has_field_derivative) lemmas erf_holomorphic_on' [holomorphic_intros] = erf_holomorphic_on [THEN holomorphic_on_chain_UNIV] lemma erf_analytic_on: "erf analytic_on A" by (auto simp: analytic_on_def) (auto intro!: exI[of _ 1] holomorphic_intros) lemma erf_analytic_on' [analytic_intros]: assumes "f analytic_on A" shows "(\x. erf (f x)) analytic_on A" proof - from assms and erf_analytic_on have "erf \ f analytic_on A" by (rule analytic_on_compose_gen) auto thus ?thesis by (simp add: o_def) qed lemma erfc_holomorphic_on: "erfc holomorphic_on A" by (auto simp: holomorphic_on_def field_differentiable_def intro!: erfc_has_field_derivative) lemmas erfc_holomorphic_on' [holomorphic_intros] = erfc_holomorphic_on [THEN holomorphic_on_chain_UNIV] lemma erfc_analytic_on: "erfc analytic_on A" by (auto simp: analytic_on_def) (auto intro!: exI[of _ 1] holomorphic_intros) lemma erfc_analytic_on' [analytic_intros]: assumes "f analytic_on A" shows "(\x. erfc (f x)) analytic_on A" proof - from assms and erfc_analytic_on have "erfc \ f analytic_on A" by (rule analytic_on_compose_gen) auto thus ?thesis by (simp add: o_def) qed end diff --git a/thys/Euler_MacLaurin/Euler_MacLaurin.thy b/thys/Euler_MacLaurin/Euler_MacLaurin.thy --- a/thys/Euler_MacLaurin/Euler_MacLaurin.thy +++ b/thys/Euler_MacLaurin/Euler_MacLaurin.thy @@ -1,1617 +1,1617 @@ section \The Euler--MacLaurin summation formula\ theory Euler_MacLaurin imports "HOL-Complex_Analysis.Complex_Analysis" Bernoulli.Periodic_Bernpoly Bernoulli.Bernoulli_FPS begin subsection \Auxiliary facts\ (* TODO Move? *) lemma pbernpoly_of_int [simp]: "pbernpoly n (of_int a) = bernoulli n" by (simp add: pbernpoly_def) lemma continuous_on_bernpoly' [continuous_intros]: assumes "continuous_on A f" shows "continuous_on A (\x. bernpoly n (f x) :: 'a :: real_normed_algebra_1)" using continuous_on_compose2[OF continuous_on_bernpoly assms, of UNIV n] by auto lemma sum_atLeastAtMost_int_last: assumes "a < (b :: int)" shows "sum f {a..b} = sum f {a.. = sum f {a.. = f a + sum f {a<..b}" by (subst sum.insert) auto finally show ?thesis . qed lemma not_in_nonpos_Reals_imp_add_nonzero: assumes "z \ \\<^sub>\\<^sub>0" "x \ 0" shows "z + of_real x \ 0" using assms by (auto simp: add_eq_0_iff2) (* END TODO *) lemma negligible_atLeastAtMostI: "b \ a \ negligible {a..(b::real)}" by (cases "b < a") auto lemma integrable_on_negligible: "negligible A \ (f :: 'n :: euclidean_space \ 'a :: banach) integrable_on A" by (subst integrable_spike_set_eq[of _ "{}"]) (simp_all add: integrable_on_empty) lemma Union_atLeastAtMost_real_of_int: assumes "a < b" shows "(\n\{a.. {real_of_int a..real_of_int b}" thus "x \ (\n\{a.. real_of_int a" "x < real_of_int b" by simp_all hence "x \ of_int \x\" "x \ of_int \x\ + 1" by linarith+ moreover from x have "\x\ \ a" "\x\ < b" by linarith+ ultimately have "\n\{a.. {of_int n..of_int (n + 1)}" by (intro bexI[of _ "\x\"]) simp_all thus ?thesis by blast qed qed auto subsection \The remainder terms\ text \ The following describes the remainder term in the classical version of the Euler--MacLaurin formula. \ definition EM_remainder' :: "nat \ (real \ 'a :: banach) \ real \ real \ 'a" where "EM_remainder' n f a b = ((-1) ^ Suc n / fact n) *\<^sub>R integral {a..b} (\t. pbernpoly n t *\<^sub>R f t)" text \ Next, we define the remainder term that occurs when one lets the right bound of summation in the Euler--MacLaurin formula tend to infinity. \ definition EM_remainder_converges :: "nat \ (real \ 'a :: banach) \ int \ bool" where "EM_remainder_converges n f a \ (\L. ((\x. EM_remainder' n f a (of_int x)) \ L) at_top)" definition EM_remainder :: "nat \ (real \ 'a :: banach) \ int \ 'a" where "EM_remainder n f a = (if EM_remainder_converges n f a then Lim at_top (\x. EM_remainder' n f a (of_int x)) else 0)" text \ The following lemmas are fairly obvious -- but tedious to prove -- properties of the remainder terms. \ lemma EM_remainder_eqI: fixes L assumes "((\x. EM_remainder' n f b (of_int x)) \ L) at_top" shows "EM_remainder n f b = L" using assms by (auto simp: EM_remainder_def EM_remainder_converges_def intro!: tendsto_Lim) lemma integrable_EM_remainder'_int: fixes a b :: int and f :: "real \ 'a :: banach" assumes "continuous_on {of_int a..of_int b} f" shows "(\t. pbernpoly n t *\<^sub>R f t) integrable_on {a..b}" proof - have [continuous_intros]: "continuous_on A f" if "A \ {of_int a..of_int b}" for A using continuous_on_subset[OF assms that] . consider "a > b" | "a = b" | "a < b" "n = 1" | "a < b" "n \ 1" by (cases a b rule: linorder_cases) auto thus ?thesis proof cases assume "a < b" and "n \ 1" thus ?thesis by (intro integrable_continuous_real continuous_intros) auto next assume ab: "a < b" and [simp]: "n = 1" let ?A = "(\n. {real_of_int n..real_of_int (n+1)}) ` {a.. ?A" then obtain i where i: "i \ {a..t. pbernpoly n t *\<^sub>R f t) integrable_on I" proof (rule integrable_spike) show "(\t. (t - of_int i - 1/2) *\<^sub>R f t) integrable_on I" using i by (auto intro!: integrable_continuous_real continuous_intros) next fix x assume "x \ I - {of_int (i + 1)}" with i have "of_int i \ x" "x < of_int i + 1" by simp_all hence "floor x = i" by linarith thus "pbernpoly n x *\<^sub>R f x = (x - of_int i - 1 / 2) *\<^sub>R f x" by (simp add: pbernpoly_def bernpoly_def frac_def) qed simp_all qed qed (simp_all add: integrable_on_negligible) qed lemma integrable_EM_remainder': fixes a b :: real and f :: "real \ 'a :: banach" assumes "continuous_on {a..b} f" shows "(\t. pbernpoly n t *\<^sub>R f t) integrable_on {a..b}" proof (cases "\a\ \ \b\") case True define a' b' where "a' = \a\" and "b' = \b\" from True have *: "a' \ b'" "a' \ a" "b' \ b" by (auto simp: a'_def b'_def) from * have A: "(\t. pbernpoly n t *\<^sub>R f t) integrable_on ({a'..b'})" by (intro integrable_EM_remainder'_int continuous_on_subset[OF assms]) auto have B: "(\t. pbernpoly n t *\<^sub>R f t) integrable_on ({a..a'})" proof (rule integrable_spike) show "pbernpoly n x *\<^sub>R f x = bernpoly n (x - of_int (floor a)) *\<^sub>R f x" if "x \ {a..real_of_int a'} - {real_of_int a'}"for x proof - have "x \ a" "x t. pbernpoly n t *\<^sub>R f t) integrable_on ({b'..b})" proof (rule integrable_spike) show "pbernpoly n x *\<^sub>R f x = bernpoly n (x - of_int b') *\<^sub>R f x" if "x \ {real_of_int b'..b} - {real_of_int b'}" for x proof - have "x \ b" "x > real_of_int b'" using that by auto with True have "floor x = b'" unfolding b'_def by (intro floor_unique; linarith) thus ?thesis by (simp add: pbernpoly_def frac_def) qed qed (insert *, auto intro!: continuous_intros integrable_continuous_real continuous_on_subset[OF assms]) have "(\t. pbernpoly n t *\<^sub>R f t) integrable_on ({a..a'} \ {a'..b'} \ {b'..b})" using * A B C by (intro integrable_Un; (subst ivl_disj_un)?) (auto simp: ivl_disj_un max_def min_def) also have "{a..a'} \ {a'..b'} \ {b'..b} = {a..b}" using * by auto finally show ?thesis . next assume *: "\ceiling a \ floor b" show ?thesis proof (rule integrable_spike) show "(\t. bernpoly n (t - floor a) *\<^sub>R f t) integrable_on {a..b}" using * by (auto intro!: integrable_continuous_real continuous_intros assms) next show "pbernpoly n x *\<^sub>R f x = bernpoly n (x - floor a) *\<^sub>R f x" if "x \ {a..b} - {}" for x proof - from * have **: "b < floor a + 1" unfolding ceiling_altdef by (auto split: if_splits simp: le_floor_iff) from that have x: "x \ a" "x \ b" by simp_all with * ** have "floor x = floor a" by linarith thus ?thesis by (simp add: pbernpoly_def frac_def) qed qed simp_all qed lemma EM_remainder'_bounded_linear: assumes "bounded_linear h" assumes "continuous_on {a..b} f" shows "EM_remainder' n (\x. h (f x)) a b = h (EM_remainder' n f a b)" proof - have "integral {a..b} (\t. pbernpoly n t *\<^sub>R h (f t)) = integral {a..b} (\t. h (pbernpoly n t *\<^sub>R f t))" using assms by (simp add: linear_simps) also have "\ = h (integral {a..b} (\t. pbernpoly n t *\<^sub>R f t))" by (subst integral_linear [OF _ assms(1), symmetric]) (auto intro!: integrable_EM_remainder' assms(2) simp: o_def) finally show ?thesis using assms(1) by (simp add: EM_remainder'_def linear_simps) qed lemma EM_remainder_converges_of_real: assumes "EM_remainder_converges n f a" "continuous_on {of_int a..} f" shows "EM_remainder_converges n (\x. of_real (f x)) a" proof - from assms obtain L where L: "((\b. EM_remainder' n f (real_of_int a) (real_of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have "((\b. EM_remainder' n (\x. of_real (f x)) (of_int a) (of_int b)) \ of_real L) at_top" proof (rule Lim_transform_eventually) show "eventually (\b. of_real (EM_remainder' n f (of_int a) (of_int b)) = EM_remainder' n (\x. of_real (f x)) (of_int a) (of_int b)) at_top" using eventually_ge_at_top[of a] by eventually_elim (intro EM_remainder'_bounded_linear [OF bounded_linear_of_real, symmetric] continuous_on_subset[OF assms(2)], auto) qed (intro tendsto_intros L) thus ?thesis unfolding EM_remainder_converges_def .. qed lemma EM_remainder_converges_of_real_iff: fixes f :: "real \ real" assumes "continuous_on {of_int a..} f" shows "EM_remainder_converges n (\x. of_real (f x) :: 'a :: {banach, real_normed_algebra_1, real_inner}) a \ EM_remainder_converges n f a" proof assume "EM_remainder_converges n (\x. of_real (f x) :: 'a) a" then obtain L :: 'a where L: "((\b. EM_remainder' n (\x. of_real (f x)) (of_int a) (of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have "((\b. EM_remainder' n f (of_int a) (of_int b)) \ L \ 1) at_top" proof (rule Lim_transform_eventually) show "eventually (\b. EM_remainder' n (\x. of_real (f x) :: 'a) (of_int a) (of_int b) \ 1 = EM_remainder' n f (of_int a) (of_int b)) at_top" using eventually_ge_at_top[of a] by eventually_elim (subst EM_remainder'_bounded_linear [OF bounded_linear_of_real], auto intro!: continuous_on_subset[OF assms]) qed (intro tendsto_intros L) thus "EM_remainder_converges n f a" unfolding EM_remainder_converges_def .. qed (intro EM_remainder_converges_of_real assms) lemma EM_remainder_of_real: assumes "continuous_on {a..} f" shows "EM_remainder n (\x. of_real (f x) :: 'a :: {banach, real_normed_algebra_1, real_inner}) a = of_real (EM_remainder n f a)" proof - have eq: "EM_remainder' n (\x. of_real (f x) :: 'a) (real_of_int a) = (\x::int. of_real (EM_remainder' n f a x))" by (intro ext EM_remainder'_bounded_linear[OF bounded_linear_of_real] continuous_on_subset[OF assms]) auto show ?thesis proof (cases "EM_remainder_converges n f a") case False with EM_remainder_converges_of_real_iff[OF assms, of n] show ?thesis by (auto simp: EM_remainder_def) next case True then obtain L where L: "((\x. EM_remainder' n f a (real_of_int x)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have L': "((\x. EM_remainder' n (\x. of_real (f x) :: 'a) a (real_of_int x)) \ of_real L) at_top" unfolding eq by (intro tendsto_of_real L) from L L' tendsto_Lim[OF _ L] tendsto_Lim[OF _ L'] show ?thesis by (auto simp: EM_remainder_def EM_remainder_converges_def) qed qed lemma EM_remainder'_cong: assumes "\x. x \ {a..b} \ f x = g x" "n = n'" "a = a'" "b = b'" shows "EM_remainder' n f a b = EM_remainder' n' g a' b'" proof - have "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) = integral {a'..b'} (\t. pbernpoly n' t *\<^sub>R g t)" unfolding assms using assms by (intro integral_cong) auto with assms show ?thesis by (simp add: EM_remainder'_def) qed lemma EM_remainder_converges_cong: assumes "\x. x \ of_int a \ f x = g x" "n = n'" "a = a'" shows "EM_remainder_converges n f a = EM_remainder_converges n' g a'" unfolding EM_remainder_converges_def by (subst EM_remainder'_cong[OF _ refl refl refl, of _ _ f g]) (use assms in auto) lemma EM_remainder_cong: assumes "\x. x \ of_int a \ f x = g x" "n = n'" "a = a'" shows "EM_remainder n f a = EM_remainder n' g a'" proof - have *: "EM_remainder_converges n f a = EM_remainder_converges n' g a'" using assms by (intro EM_remainder_converges_cong) auto show ?thesis unfolding EM_remainder_def by (subst EM_remainder'_cong[OF _ refl refl refl, of _ _ f g]) (use assms * in auto) qed lemma EM_remainder_converges_cnj: assumes "continuous_on {a..} f" and "EM_remainder_converges n f a" shows "EM_remainder_converges n (\x. cnj (f x)) a" proof - interpret bounded_linear cnj by (rule bounded_linear_cnj) obtain L where L: "((\x. EM_remainder' n f (real_of_int a) (real_of_int x)) \ L) at_top" using assms unfolding EM_remainder_converges_def by blast note tendsto_cnj [OF this] also have "(\x. cnj (EM_remainder' n f (real_of_int a) (real_of_int x))) = (\x. EM_remainder' n (\x. cnj (f x)) (real_of_int a) (real_of_int x))" by (subst EM_remainder'_bounded_linear [OF bounded_linear_cnj]) (rule continuous_on_subset [OF assms(1)], auto) finally have L': "(\ \ cnj L) at_top" . thus "EM_remainder_converges n (\x. cnj (f x)) a" by (auto simp: EM_remainder_converges_def) qed lemma EM_remainder_converges_cnj_iff: assumes "continuous_on {of_int a..} f" shows "EM_remainder_converges n (\x. cnj (f x)) a \ EM_remainder_converges n f a" proof assume "EM_remainder_converges n (\x. cnj (f x)) a" hence "EM_remainder_converges n (\x. cnj (cnj (f x))) a" by (rule EM_remainder_converges_cnj [rotated]) (auto intro: continuous_intros assms) thus "EM_remainder_converges n f a" by simp qed (intro EM_remainder_converges_cnj assms) lemma EM_remainder_cnj: assumes "continuous_on {a..} f" shows "EM_remainder n (\x. cnj (f x)) a = cnj (EM_remainder n f a)" proof (cases "EM_remainder_converges n f a") case False hence "\EM_remainder_converges n (\x. cnj (f x)) a" by (subst EM_remainder_converges_cnj_iff [OF assms]) with False show ?thesis by (simp add: EM_remainder_def) next case True then obtain L where L: "((\x. EM_remainder' n f (real_of_int a) (real_of_int x)) \ L) at_top" unfolding EM_remainder_converges_def by blast note tendsto_cnj [OF this] also have "(\x. cnj (EM_remainder' n f (real_of_int a) (real_of_int x))) = (\x. EM_remainder' n (\x. cnj (f x)) (real_of_int a) (real_of_int x))" by (subst EM_remainder'_bounded_linear [OF bounded_linear_cnj]) (rule continuous_on_subset [OF assms(1)], auto) finally have L': "(\ \ cnj L) at_top" . moreover from assms and L have "EM_remainder n f a = L" by (intro EM_remainder_eqI) ultimately show "EM_remainder n (\x. cnj (f x)) a = cnj (EM_remainder n f a)" using L' by (intro EM_remainder_eqI) simp_all qed lemma EM_remainder'_combine: fixes f :: "real \ 'a :: banach" assumes [continuous_intros]: "continuous_on {a..c} f" assumes "a \ b" "b \ c" shows "EM_remainder' n f a b + EM_remainder' n f b c = EM_remainder' n f a c" proof - have "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) + integral {b..c} (\t. pbernpoly n t *\<^sub>R f t) = integral {a..c} (\t. pbernpoly n t *\<^sub>R f t)" by (intro Henstock_Kurzweil_Integration.integral_combine assms integrable_EM_remainder') from this [symmetric] show ?thesis by (simp add: EM_remainder'_def algebra_simps) qed lemma uniformly_convergent_EM_remainder': fixes f :: "'a \ real \ 'b :: {banach,real_normed_algebra}" assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes integrable: "\a' b y. y \ A \ a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniformly_convergent_on A (\b s. EM_remainder' n (f s) a b)" proof - interpret bounded_linear "\x::'b. ((- 1) ^ Suc n / fact n) *\<^sub>R x" by (rule bounded_linear_scaleR_right) - from bounded_pbernpoly[of n] guess C . note C = this + from bounded_pbernpoly obtain C where C: "\x. norm (pbernpoly n x) \ C" by auto from C[of 0] have [simp]: "C \ 0" by simp show ?thesis unfolding EM_remainder'_def proof (intro uniformly_convergent_on uniformly_convergent_improper_integral') fix x assume "x \ a" thus "((\x. C * G x) has_real_derivative C * g x) (at x within {a..})" by (intro DERIV_cmult deriv) next fix y a' b assume "y \ A" "a \ a'" "a' \ b" thus "(\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" by (rule integrable) next from conv obtain L where "(\y. G (real y)) \ L" by (auto simp: convergent_def) from tendsto_mult[OF tendsto_const[of C] this] show "convergent (\y. C * G (real y))" by (auto simp: convergent_def) next show "\\<^sub>F x in at_top. \y\A. norm (pbernpoly n x *\<^sub>R f y x) \ C * g x" using C unfolding norm_scaleR by (intro eventually_mono[OF bound] ballI mult_mono) auto qed qed lemma uniform_limit_EM_remainder: fixes f :: "'a \ real \ 'b :: {banach,real_normed_algebra}" assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes integrable: "\a' b y. y \ A \ a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniform_limit A (\b s. EM_remainder' n (f s) a b) (\s. EM_remainder n (f s) a) sequentially" proof - have *: "uniformly_convergent_on A (\b s. EM_remainder' n (f s) a b)" by (rule uniformly_convergent_EM_remainder'[OF assms]) also have "?this \ ?thesis" unfolding uniformly_convergent_uniform_limit_iff proof (intro uniform_limit_cong refl always_eventually allI ballI) fix s assume "s \ A" with * have **: "convergent (\b. EM_remainder' n (f s) a b)" by (rule uniformly_convergent_imp_convergent) show "lim (\b. EM_remainder' n (f s) a b) = EM_remainder n (f s) a" proof (rule sym, rule EM_remainder_eqI) have "((\x. EM_remainder' n (f s) (real_of_int a) (real x)) \ lim (\x. EM_remainder' n (f s) (real_of_int a) (real x))) at_top" (is "(_ \ ?L) _") using ** unfolding convergent_LIMSEQ_iff by blast hence "((\x. EM_remainder' n (f s) (real_of_int a) (real (nat x))) \ ?L) at_top" by (rule filterlim_compose) (fact filterlim_nat_sequentially) thus "((\x. EM_remainder' n (f s) (real_of_int a) (real_of_int x)) \ ?L) at_top" by (rule Lim_transform_eventually) (auto intro: eventually_mono[OF eventually_ge_at_top[of 0]]) qed qed finally show \ . qed lemma tendsto_EM_remainder: fixes f :: "real \ 'b :: {banach,real_normed_algebra}" assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes integrable: "\a' b . a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. norm (f x) \ g x) at_top" shows "filterlim (\b. EM_remainder' n f a b) (nhds (EM_remainder n f a)) sequentially" proof - have "uniform_limit {()} (\b s. EM_remainder' n f a b) (\s. EM_remainder n f a) sequentially" using assms by (intro uniform_limit_EM_remainder[where G = G and g = g]) auto moreover have "() \ {()}" by simp ultimately show ?thesis by (rule tendsto_uniform_limitI) qed lemma EM_remainder_0 [simp]: "EM_remainder n (\x. 0) a = 0" by (rule EM_remainder_eqI) (simp add: EM_remainder'_def) lemma holomorphic_EM_remainder': assumes deriv: "\z t. z \ U \ t \ {a..x} \ ((\z. f z t) has_field_derivative f' z t) (at z within U)" assumes int: "\b c z e. a \ b \ c \ x \ z \ U \ (\t. of_real (bernpoly n (t - e)) * f z t) integrable_on {b..c}" assumes cont: "continuous_on (U \ {a..x}) (\(z, t). f' z t)" assumes "convex U" shows "(\s. EM_remainder' n (f s) a x) holomorphic_on U" unfolding EM_remainder'_def scaleR_conv_of_real proof (intro holomorphic_intros) have holo: "(\z. integral (cbox b c) (\t. of_real (bernpoly n (t - e)) * f z t)) holomorphic_on U" if "b \ a" "c \ x" for b c e :: real proof (rule leibniz_rule_holomorphic) fix z t assume "z \ U" "t \ cbox b c" thus "((\z. complex_of_real (bernpoly n (t - e)) * f z t) has_field_derivative complex_of_real (bernpoly n (t - e)) * f' z t) (at z within U)" using that by (intro DERIV_cmult deriv) auto next fix z assume "z \ U" thus "(\t. complex_of_real (bernpoly n (t - e)) * f z t) integrable_on cbox b c" using that int[of b c z] by auto next have "continuous_on (U \ {b..c}) (\(z, t). f' z t)" using cont by (rule continuous_on_subset) (insert that, auto) thus "continuous_on (U \ cbox b c) (\(z, t). complex_of_real (bernpoly n (t - e)) * f' z t)" by (auto simp: case_prod_unfold intro!: continuous_intros) qed fact+ consider "a > x" | "a \ x" "floor x \ a" | "a \ x" "floor x > a" by force hence "(\z. integral (cbox a x) (\t. of_real (pbernpoly n t) * f z t)) holomorphic_on U" (is "?f a x holomorphic_on _") proof cases case 2 have "(\z. integral (cbox a x) (\t. of_real (bernpoly n (t - of_int \x\)) * f z t)) holomorphic_on U" by (intro holo) auto also have "(\z. integral (cbox a x) (\t. of_real (bernpoly n (t - of_int \x\)) * f z t)) = ?f a x" proof (intro ext integral_cong, goal_cases) case (1 z t) hence "t \ a" "t \ x" by auto hence "floor t = floor x" using 2 by linarith thus ?case by (simp add: pbernpoly_def frac_def) qed finally show ?thesis . next case 3 define N :: "int set" where "N = {\a\..<\x\}" define A where "A = insert {a..of_int \a\} (insert {of_int \x\..x} ((\n. {of_int n..of_int n + 1}) ` N))" { fix X assume "X \ A" then consider "X = {a..of_int \a\}" | "X = {of_int \x\..x}" | n where "X = {of_int n..of_int n + 1}" "n \ N" by (auto simp: A_def) } note A_cases = this have division: "A division_of {a..x}" proof (rule division_ofI) show "finite A" by (auto simp: A_def N_def) fix K assume K: "K \ A" from 3 have "of_int \a\ \ x" using ceiling_le[of a "floor x"] by linarith moreover from 3 have "of_int \x\ \ a" by linarith ultimately show "K \ {a..x}" using K 3 by (auto simp: A_def N_def) linarith+ from K show "K \ {}" and "\a b. K = cbox a b" by (auto simp: A_def) next fix K1 K2 assume K: "K1 \ A" "K2 \ A" "K1 \ K2" have F1: "interior {a..\a\} \ interior {\x\..x} = {}" using 3 ceiling_le[of a "floor x"] by (auto simp: min_def max_def) hence F2: "interior {\x\..x} \ interior {a..\a\} = {}" by simp have F3: "interior {a..\a\} \ interior {of_int n..of_int n+1} = {}" "interior {\x\..x} \ interior {of_int n..of_int n+1} = {}" "interior {of_int n..of_int n+1} \ interior {a..\a\} = {}" "interior {of_int n..of_int n+1} \ interior {\x\..x} = {}"if "n \ N" for n using 3 ceiling_le[of a "floor x"] that by (auto simp: min_def max_def N_def) have F4: "interior {real_of_int n..of_int n+1} \ interior {of_int m..of_int m+1} = {}" if "{real_of_int n..of_int n+1} \ {of_int m..of_int m+1}" for m n proof - from that have "n \ m" by auto thus ?thesis by simp qed from F1 F2 F3 F4 K show "interior K1 \ interior K2 = {}" by (elim A_cases) (simp_all only: not_False_eq_True) next show "\A = {a..x}" proof (cases "\a\ = \x\") case True thus ?thesis using 3 by (auto simp: A_def N_def intro: order.trans) linarith+ next case False with 3 have *: "\a\ < \x\" by linarith have "\A = {a..of_int \a\} \ (\n\N. {of_int n..of_int (n + 1)}) \ {of_int \x\..x}" by (simp add: A_def Un_ac) also have "(\n\N. {of_int n..of_int (n + 1)}) = {of_int \a\..real_of_int \x\}" using * unfolding N_def by (intro Union_atLeastAtMost_real_of_int) also have "{a..of_int \a\} \ \ = {a..real_of_int \x\}" using 3 * by (intro ivl_disj_un) auto also have "\ \ {of_int \x\..x} = {a..x}" using 3 * by (intro ivl_disj_un) auto finally show ?thesis . qed qed have "(\z. \X\A. integral X (\t. of_real (bernpoly n (t - \Inf X\)) * f z t)) holomorphic_on U" proof (intro holomorphic_on_sum holo, goal_cases) case (1 X) from 1 and division have subset: "X \ {a..x}" by (auto simp: division_of_def) from 1 obtain b c where [simp]: "X = cbox b c" "b \ c" by (auto simp: A_def) from subset have "b \ a" "c \ x" by auto hence "(\x. integral (cbox b c) (\t. of_real (bernpoly n (t - \Inf {b..c}\)) * f x t)) holomorphic_on U" by (intro holo) auto thus ?case by simp qed also have "?this \ (\z. integral {a..x} (\t. of_real (pbernpoly n t) * f z t)) holomorphic_on U" proof (intro holomorphic_cong refl, goal_cases) case (1 z) have "((\t. of_real (pbernpoly n t) * f z t) has_integral (\X\A. integral X (\t. of_real (bernpoly n (t - \Inf X\)) * f z t))) {a..x}" using division proof (rule has_integral_combine_division) fix X assume X: "X \ A" then obtain b c where X': "X = {b..c}" "b \ c" by (elim A_cases) auto from X and division have "X \ {a..x}" by (auto simp: division_of_def) with X' have bc: "b \ a" "c \ x" by auto have "((\t. of_real (bernpoly n (t - of_int \Inf X\)) * f z t) has_integral integral X (\t. of_real (bernpoly n (t - of_int \Inf X\)) * f z t)) X" unfolding X' using \z \ U\ bc by (intro integrable_integral int) also have "?this \ ((\t. of_real (pbernpoly n t) * f z t) has_integral integral X (\t. of_real (bernpoly n (t - of_int \Inf X\)) * f z t)) X" proof (rule has_integral_spike_eq[of "{Sup X}"], goal_cases) case (2 t) note t = this from \X \ A\ have "\t\ = \Inf X\" proof (cases rule: A_cases [consumes 1]) case 1 with t show ?thesis by (intro floor_unique) (auto simp: ceiling_altdef split: if_splits, (linarith+)?) next case 2 with t show ?thesis by (intro floor_unique) (auto simp: ceiling_altdef split: if_splits, (linarith+)?) next case 3 with t show ?thesis by (intro floor_unique) (auto simp: ceiling_altdef N_def split: if_splits) qed thus ?case by (simp add: pbernpoly_def frac_def) qed auto finally show \ . qed thus ?case by (simp add: has_integral_iff) qed finally show ?thesis by simp qed auto thus "(\z. integral {a..x} (\t. of_real (pbernpoly n t) * f z t)) holomorphic_on U" by simp qed lemma assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes deriv': "\z t x. z \ U \ x \ a \ t \ {a..x} \ ((\z. f z t) has_field_derivative f' z t) (at z within U)" assumes cont: "continuous_on (U \ {of_int a..}) (\(z, t). f' z t)" assumes int: "\b c z e. a \ b \ z \ U \ (\t. of_real (bernpoly n (t - e)) * f z t) integrable_on {b..c}" assumes int': "\a' b y. y \ U \ a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. \y\U. norm (f y x) \ g x) at_top" assumes "open U" shows analytic_EM_remainder: "(\s::complex. EM_remainder n (f s) a) analytic_on U" and holomorphic_EM_remainder: "(\s::complex. EM_remainder n (f s) a) holomorphic_on U" proof - show "(\s::complex. EM_remainder n (f s) a) analytic_on U" unfolding analytic_on_def proof fix z assume "z \ U" from \z \ U\ and \open U\ obtain \ where \: "\ > 0" "ball z \ \ U" by (auto simp: open_contains_ball) have "(\s. EM_remainder n (f s) a) holomorphic_on ball z \" proof (rule holomorphic_uniform_sequence) fix x :: nat show "(\s. EM_remainder' n (f s) a x) holomorphic_on ball z \" proof (rule holomorphic_EM_remainder', goal_cases) fix s t assume "s \ ball z \" "t \ {real_of_int a..real x}" thus "((\z. f z t) has_field_derivative f' s t) (at s within ball z \)" using \ by (intro DERIV_subset[OF deriv'[of _ x]]) auto next case (2 b c s e) with \ have "s \ U" by blast with 2 show ?case using \ int[of b s e c] by (cases "a \ x") auto next from cont show "continuous_on (ball z \ \ {real_of_int a..real x}) (\(z, t). f' z t)" by (rule continuous_on_subset) (insert \, auto) qed (auto) next fix s assume s: "s \ ball z \" have "open (ball z \)" by simp with s obtain \ where \: "\ > 0" "cball s \ \ ball z \" unfolding open_contains_cball by blast moreover have bound': "eventually (\x. \y\cball s \. norm (f y x) \ g x) at_top" by (intro eventually_mono [OF bound]) (insert \ \, auto) have "uniform_limit (cball s \) (\x s. EM_remainder' n (f s) (real_of_int a) (real x)) (\s. EM_remainder n (f s) a) sequentially" by (rule uniform_limit_EM_remainder[OF deriv int' conv bound']) (insert \ \ s, auto) ultimately show "\\>0. cball s \ \ ball z \ \ uniform_limit (cball s \) (\x s. EM_remainder' n (f s) (real_of_int a) (real x)) (\s. EM_remainder n (f s) a) sequentially" by blast qed auto with \ show "\\>0. (\s. EM_remainder n (f s) a) holomorphic_on ball z \" by blast qed thus "(\s::complex. EM_remainder n (f s) a) holomorphic_on U" by (rule analytic_imp_holomorphic) qed text \ The following lemma is the first step in the proof of the Euler--MacLaurin formula: We show the relationship between the first-order remainder term and the difference of the integral and the sum. \ context fixes f f' :: "real \ 'a :: banach" fixes a b :: int and I S :: 'a fixes Y :: "real set" assumes "a \ b" assumes fin: "finite Y" assumes cont: "continuous_on {real_of_int a..real_of_int b} f" assumes deriv [derivative_intros]: "\x::real. x \ {a..b} - Y \ (f has_vector_derivative f' x) (at x)" defines S_def: "S \ (\i\{a<..b}. f i)" and I_def: "I \ integral {a..b} f" begin lemma diff_sum_integral_has_integral_int: "((\t. (frac t - 1/2) *\<^sub>R f' t) has_integral (S - I - (f b - f a) /\<^sub>R 2)) {a..b}" proof (cases "a = b") case True thus ?thesis by (simp_all add: S_def I_def has_integral_refl) next case False with \a \ b\ have ab: "a < b" by simp let ?A = "(\n. {real_of_int n..real_of_int (n+1)}) ` {a.. {of_int a..of_int b}" for A using continuous_on_subset[OF cont that] . define d where "d = (\x. (f x + f (x + 1)) /\<^sub>R 2 - integral {x..x+1} f)" have "((\t. (frac t - 1/2) *\<^sub>R f' t) has_integral d i) {of_int i..of_int (i+1)}" if i: "i \ {a..R f' x = (x - of_int i - 1 / 2) *\<^sub>R f' x" if "x \ {of_int i..of_int (i + 1)} - {of_int (i + 1)}" for x proof - have "x \ of_int i" "x < of_int (i + 1)" using that by auto hence "floor x = of_int i" by (subst floor_unique) auto thus ?thesis by (simp add: frac_def) qed next define h where "h = (\x::real. (x - of_int i - 1 / 2) *\<^sub>R f' x)" define g where "g = (\x::real. (x - of_int i - 1/2) *\<^sub>R f x - integral {of_int i..x} f)" have *: "((\x. integral {real_of_int i..x} f) has_vector_derivative f x) (at x within {i..i+1})" if "x \ {of_int i<..x. integral {real_of_int i..x} f) has_vector_derivative f x) (at x)" if "x \ {of_int i<..I. I\?A \ ((\x. (frac x - 1 / 2) *\<^sub>R f' x) has_integral d (\Inf I\)) I" by (auto simp: add_ac) have "((\x::real. (frac x - 1 / 2) *\<^sub>R f' x) has_integral (\I\?A. d (\Inf I\))) (\?A)" by (intro has_integral_Union * finite_imageI) (force intro!: negligible_atLeastAtMostI pairwiseI)+ also have "\?A = {of_int a..of_int b}" by (intro Union_atLeastAtMost_real_of_int ab) also have "(\I\?A. d (\Inf I\)) = (\i=a.. = (1 / 2) *\<^sub>R ((\i = a..i = a..i = a..R (?S1 + ?S2) - ?S3") by (simp add: d_def algebra_simps sum.distrib sum_subtractf scaleR_sum_right) also have "?S1 = (\i = a..b. f (real_of_int i)) - f b" unfolding S_def using ab by (subst sum_atLeastAtMost_int_last) auto also have "(\i = a..b. f (real_of_int i)) = S + f a" unfolding S_def using ab by (subst sum_atLeastAtMost_int_head) auto also have "?S2 = S" unfolding S_def by (intro sum.reindex_bij_witness[of _ "\i. i-1" "\i. i+1"]) auto also have "(1 / 2) *\<^sub>R (S + f a - f b + S) = (1/2) *\<^sub>R S + (1/2) *\<^sub>R S - (f b - f a) /\<^sub>R 2" by (simp add: algebra_simps) also have "(1/2) *\<^sub>R S + (1/2) *\<^sub>R S = S" by (simp add: scaleR_add_right [symmetric]) also have "?S3 = (\I\?A. integral I f)" by (subst sum.reindex) (auto simp: inj_on_def add_ac) also have "\ = I" unfolding I_def by (intro integral_combine_division_topdown [symmetric] division integrable_continuous_real continuous_intros) simp_all finally show ?thesis by (simp add: algebra_simps) qed lemma diff_sum_integral_has_integral_int': "((\t. pbernpoly 1 t *\<^sub>R f' t) has_integral (S - I - (f b - f a) /\<^sub>R 2 )) {a..b}" using diff_sum_integral_has_integral_int by (simp add: pbernpoly_def bernpoly_def) lemma EM_remainder'_Suc_0: "EM_remainder' (Suc 0) f' a b = S - I - (f b - f a) /\<^sub>R 2" using diff_sum_integral_has_integral_int' by (simp add: has_integral_iff EM_remainder'_def) end text \ Next, we show that the $n$-th-order remainder can be expressed in terms of the $n+1$-th-order remainder term. Iterating this essentially yields the Euler--MacLaurin formula. \ context fixes f f' :: "real \ 'a :: banach" and a b :: int and n :: nat and A :: "real set" assumes ab: "a \ b" and n: "n > 0" assumes fin: "finite A" assumes cont: "continuous_on {of_int a..of_int b} f" assumes cont': "continuous_on {of_int a..of_int b} f'" assumes deriv: "\x. x \ {of_int a<.. (f has_vector_derivative f' x) (at x)" begin lemma EM_remainder'_integral_conv_Suc: shows "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) = (bernoulli (Suc n) / real (Suc n)) *\<^sub>R (f b - f a) - integral {a..b} (\t. pbernpoly (Suc n) t *\<^sub>R f' t) /\<^sub>R real (Suc n)" unfolding EM_remainder'_def proof - let ?h = "\i. (pbernpoly (Suc n) (real_of_int i) / real (Suc n)) *\<^sub>R f (real_of_int i)" define T where "T = integral {a..b} (\t. (pbernpoly (Suc n) t / real (Suc n)) *\<^sub>R f' t)" note [derivative_intros] = has_field_derivative_pbernpoly_Suc' let ?A = "real_of_int ` {a..b} \ A" have "((\t. pbernpoly n t *\<^sub>R f t) has_integral (-T + (?h b - ?h a))) {a..b}" proof (rule integration_by_parts_interior_strong[OF bounded_bilinear_scaleR]) from fin show "finite ?A" by simp from \n > 0\ show "continuous_on {of_int a..of_int b} (\t. pbernpoly (Suc n) t / real (Suc n))" by (intro continuous_intros) auto show "continuous_on {of_int a..of_int b} f" by fact show "(f has_vector_derivative f' t) (at t)" if "t \ {of_int a<..t. pbernpoly (Suc n) t *\<^sub>R f' t) integrable_on {a..b}" by (intro integrable_EM_remainder' cont') hence "(\t. (1 / real (Suc n)) *\<^sub>R pbernpoly (Suc n) t *\<^sub>R f' t) integrable_on {a..b}" by (rule integrable_cmul) also have "(\t. (1 / real (Suc n)) *\<^sub>R pbernpoly (Suc n) t *\<^sub>R f' t) = (\t. (pbernpoly (Suc n) t / real (Suc n)) *\<^sub>R f' t)" by (rule ext) (simp add: algebra_simps) finally show "((\t. (pbernpoly (Suc n) t / real (Suc n)) *\<^sub>R f' t) has_integral ?h b - ?h a - (- T + (?h b - ?h a))) {a..b}" using integrable_EM_remainder'[of a b f' "Suc n"] by (simp add: has_integral_iff T_def) qed (insert ab n, auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric] not_le elim!: Ints_cases) also have "?h b - ?h a = (bernoulli (Suc n) / real (Suc n)) *\<^sub>R (f b - f a)" using n by (simp add: algebra_simps bernoulli'_def) finally have "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) = \ - T" by (simp add: has_integral_iff) also have "T = integral {a..b} (\t. (1 / real (Suc n)) *\<^sub>R (pbernpoly (Suc n) t) *\<^sub>R f' t)" by (simp add: T_def) also have "\ = integral {a..b} (\t. pbernpoly (Suc n) t *\<^sub>R f' t) /\<^sub>R real (Suc n)" by (subst integral_cmul) (simp_all add: divide_simps) finally show ?thesis . qed lemma EM_remainder'_conv_Suc: "EM_remainder' n f a b = ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (f b - f a) + EM_remainder' (Suc n) f' a b" by (simp add: EM_remainder'_def EM_remainder'_integral_conv_Suc scaleR_diff_right scaleR_add_right field_simps del: of_nat_Suc) end context fixes f f' :: "real \ 'a :: banach" and a :: int and n :: nat and A :: "real set" and C assumes n: "n > 0" assumes fin: "finite A" assumes cont: "continuous_on {of_int a..} f" assumes cont': "continuous_on {of_int a..} f'" assumes lim: "(f \ C) at_top" assumes deriv: "\x. x \ {of_int a<..} - A \ (f has_vector_derivative f' x) (at x)" begin lemma shows EM_remainder_converges_iff_Suc_converges: "EM_remainder_converges n f a \ EM_remainder_converges (Suc n) f' a" and EM_remainder_conv_Suc: "EM_remainder_converges n f a \ EM_remainder n f a = ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (C - f a) + EM_remainder (Suc n) f' a" proof (rule iffI) define g where "g = (\x. ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (f x - f a))" define G where "G = ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (C - f a)" have limit_g: "(g \ G) at_top" unfolding g_def G_def by (intro tendsto_intros lim) have *: "eventually (\x. EM_remainder' n f (real_of_int a) (real_of_int x) = g x + EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x)) at_top" using eventually_ge_at_top[of a] proof eventually_elim case (elim b) thus ?case using EM_remainder'_conv_Suc[OF elim n fin continuous_on_subset[OF cont] continuous_on_subset[OF cont'] deriv] by (auto simp: g_def) qed { assume "EM_remainder_converges n f a" then obtain L where L: "((\b. EM_remainder' n f (real_of_int a) (real_of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have *: "((\b. EM_remainder' (Suc n) f' (real_of_int a) (real_of_int b)) \ L - G) at_top" proof (rule Lim_transform_eventually) show "\\<^sub>F x in at_top. EM_remainder' n f (real_of_int a) (real_of_int x) - g x = EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x)" using * by (simp add: algebra_simps) show "((\x. EM_remainder' n f (real_of_int a) (real_of_int x) - g x) \ L - G) at_top" by (intro tendsto_intros filterlim_compose[OF limit_g] L) qed from * show "EM_remainder_converges (Suc n) f' a" unfolding EM_remainder_converges_def .. from * have "EM_remainder (Suc n) f' a = L - G" by (rule EM_remainder_eqI) moreover from L have "EM_remainder n f a = L" by (rule EM_remainder_eqI) ultimately show "EM_remainder n f a = G + EM_remainder (Suc n) f' a" by (simp add: G_def) } { assume "EM_remainder_converges (Suc n) f' a" then obtain L where L: "((\b. EM_remainder' (Suc n) f' (real_of_int a) (real_of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have *: "((\b. EM_remainder' n f (real_of_int a) (real_of_int b)) \ G + L) at_top" proof (rule Lim_transform_eventually) show "\\<^sub>F x in at_top. g x + EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x) = EM_remainder' n f (real_of_int a) (real_of_int x)" using * by (subst eq_commute) show "((\x. g x + EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x)) \ G + L) at_top" by (intro tendsto_intros filterlim_compose[OF limit_g] L) qed thus "EM_remainder_converges n f a" unfolding EM_remainder_converges_def .. } qed end subsection \The conventional version of the Euler--MacLaurin formula\ text \ The following theorems are the classic Euler--MacLaurin formula that can be found, with slight variations, in many sources (e.\,g.\ \cite{AS_HMF, apostol99, GKP_CM}). \ context fixes f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a b :: int assumes ab: "a \ b" fixes N :: nat assumes N: "N > 0" fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ N \ continuous_on {real_of_int a..real_of_int b} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k < N \ x \ {a..b} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" begin theorem euler_maclaurin_raw_strong_int: defines "S \ (\i\{a<..b}. f (of_int i))" defines "I \ integral {of_int a..of_int b} f" defines "c' \ \k. (bernoulli' (Suc k) / fact (Suc k)) *\<^sub>R (fs k b - fs k a)" shows "S - I = (\k 'a" where "c = (\k. ((-1) ^ (Suc k) * bernoulli (Suc k) / fact (Suc k)) *\<^sub>R (fs k b - fs k a))" have "S - I = (\k 1" "m \ N" for m using that proof (induction m rule: dec_induct) case base with ab fin fs_cont[of 0] show ?case using fs_deriv[of 0] N unfolding One_nat_def by (subst EM_remainder'_Suc_0[of _ _ Y f]) (simp_all add: algebra_simps S_def I_def c_def) next case (step n) from step.prems have "S - I = (\kkkR (fs n b - fs n a))" (is "_ = _ + ?c") by (simp add: EM_remainder'_Suc_0 c_def) also have "\ + EM_remainder' n (fs n) a b = (\k b" "0 < N" "finite Y" "fs 0 = f" "(\k. k \ N \ continuous_on {real a..real b} (fs k))" "(\k x. k < N \ x \ {real a..real b} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x))" shows "(\i\{a<..b}. f (real i)) - integral {real a..real b} f = (\kR (fs k (real b) - fs k (real a))) + EM_remainder' N (fs N) (real a) (real b)" proof - have "(\i\{int a<..int b}. f (real_of_int i)) - integral {real_of_int (int a)..real_of_int (int b)} f = (\kR (fs k (real_of_int (int b)) - fs k (real_of_int (int a)))) + EM_remainder' N (fs N) (real_of_int (int a)) (real_of_int (int b))" using assms by (intro euler_maclaurin_raw_strong_int[where Y = Y] assms) simp_all also have "(\i\{int a<..int b}. f (real_of_int i)) = (\i\{a<..b}. f (real i))" by (intro sum.reindex_bij_witness[of _ int nat]) auto finally show ?thesis by simp qed subsection \The ``Concrete Mathematics'' version of the Euler--MacLaurin formula\ text \ As explained in \textit{Concrete Mathematics}~\cite{GKP_CM}, the above form of the formula has some drawbacks: When applying it to determine the asymptotics of some concrete function, one is usually left with several different unwieldy constant terms that are difficult to get rid of. There is no general way to determine what these constant terms are, but in concrete applications, they can often be determined or estimated by other means. We can therefore simply group all the constant terms into a single constant and have the user provide a proof of what it is. \ locale euler_maclaurin_int = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: int fixes N :: nat assumes N: "N > 0" fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ N \ continuous_on {real_of_int a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k < N \ x \ {of_int a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {of_int a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {of_int a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (of_int b) - (\iR fs i (of_int b))) \ C) at_top" begin context fixes C' T defines "C' \ -f a + F a + C + (\kR (fs k (of_int a)))" and "T \ (\x. \iR fs i x)" begin lemma euler_maclaurin_strong_int_aux: assumes ab: "a \ b" defines "S \ (\k=a..b. f (of_int k))" shows "S - F (of_int b) - T (of_int b) = EM_remainder' N (fs N) (of_int a) (of_int b) + (C - C')" proof (cases "a = b") case True thus ?thesis unfolding C'_def by (simp add: S_def EM_remainder'_def T_def) next case False with assms have ab: "a < b" by simp define T' where "T' = (\kR (fs k (of_int a)))" have "(\i\{a<..b}. f (of_int i)) - integral {of_int a..of_int b} f = (\kR (fs k (of_int b) - fs k (of_int a))) + EM_remainder' N (fs N) (of_int a) (of_int b)" using ab by (intro euler_maclaurin_raw_strong_int [where Y = Y] N fin fs_0 continuous_on_subset[OF fs_cont] fs_deriv) auto also have "(f has_integral (F b - F a)) {of_int a..of_int b}" using ab by (intro fundamental_theorem_of_calculus_strong[OF fin]) (auto intro!: continuous_on_subset[OF F_cont] derivative_intros) hence "integral {of_int a..of_int b} f = F (of_int b) - F (of_int a)" by (simp add: has_integral_iff) also have "(\kR (fs k (of_int b) - fs k (of_int a))) = T (of_int b) - T'" by (simp add: T_def T'_def algebra_simps sum_subtractf) also have "(\i\{a<..b}. f (of_int i)) = S - f (of_int a)" unfolding S_def using ab by (subst sum_atLeastAtMost_int_head) auto finally show ?thesis by (simp add: algebra_simps C'_def T'_def) qed lemma EM_remainder_limit: assumes ab: "a \ b" defines "D \ EM_remainder' N (fs N) (of_int a) (of_int b)" shows "EM_remainder N (fs N) b = C' - D" and EM_remainder_converges: "EM_remainder_converges N (fs N) b" proof - note limit also have "((\b. (\k = a..b. f (of_int k)) - F (of_int b) - (\iR fs i (of_int b))) \ C) at_top = ((\b. (\k = a..b. f (of_int k)) - F (of_int b) - T (of_int b)) \ C) at_top" unfolding T_def .. also have "eventually (\x. (\k=a..x. f k) - F (of_int x) - T (of_int x) = EM_remainder' N (fs N) (of_int a) (of_int x) + (C - C')) at_top" (is "eventually (\x. ?f x = ?g x) _") using eventually_gt_at_top[of b] by eventually_elim (rule euler_maclaurin_strong_int_aux, insert ab, simp_all) hence "(?f \ C) at_top \ (?g \ C) at_top" by (intro filterlim_cong refl) finally have "((\x. ?g x - (C - C')) \ (C - (C - C'))) at_top" by (rule tendsto_diff[OF _ tendsto_const]) hence *: "((\x. EM_remainder' N (fs N) (of_int a) (of_int x)) \ C') at_top" by simp have "((\x. EM_remainder' N (fs N) (of_int a) (of_int x) - D) \ C' - D) at_top" by (intro tendsto_intros *) also have "eventually (\x. EM_remainder' N (fs N) (of_int a) (of_int x) - D = EM_remainder' N (fs N) (of_int b) (of_int x)) at_top" (is "eventually (\x. ?f x = ?g x) _") using eventually_ge_at_top[of b] proof eventually_elim case (elim x) have "EM_remainder' N (fs N) (of_int a) (of_int x) = D + EM_remainder' N (fs N) (of_int b) (of_int x)" using elim ab unfolding D_def by (intro EM_remainder'_combine [symmetric] continuous_on_subset[OF fs_cont]) auto thus ?case by simp qed hence "(?f \ C' - D) at_top \ (?g \ C' - D) at_top" by (intro filterlim_cong refl) finally have *: \ . from * show "EM_remainder_converges N (fs N) b" unfolding EM_remainder_converges_def .. from * show "EM_remainder N (fs N) b = C' - D" by (rule EM_remainder_eqI) qed theorem euler_maclaurin_strong_int: assumes ab: "a \ b" defines "S \ (\k=a..b. f (of_int k))" shows "S = F (of_int b) + C + T (of_int b) - EM_remainder N (fs N) b" proof - have "S = F (of_int b) + T (of_int b) + - (C' - EM_remainder' N (fs N) (of_int a) (of_int b)) + C" using euler_maclaurin_strong_int_aux[OF ab] by (simp add: algebra_simps S_def) also have "C' - EM_remainder' N (fs N) (of_int a) (of_int b) = EM_remainder N (fs N) b" using ab by (rule EM_remainder_limit(1) [symmetric]) finally show ?thesis by simp qed end end text \ The following version of the formula removes all the terms where the associated Bernoulli numbers vanish. \ locale euler_maclaurin_int' = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: int fixes N :: nat fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ 2*N+1 \ continuous_on {real_of_int a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k \ 2*N \ x \ {of_int a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {of_int a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {of_int a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (of_int b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b))) \ C) at_top" begin sublocale euler_maclaurin_int F f fs a "2*N+1" C Y by standard (insert fin fs_0 fs_cont fs_deriv F_cont F_deriv limit, simp_all) theorem euler_maclaurin_strong_int': assumes "a \ b" shows "(\k=a..b. f (of_int k)) = F (of_int b) + C + (1 / 2) *\<^sub>R f (of_int b) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (of_int b)) - EM_remainder (2*N+1) (fs (2*N+1)) b" proof - have "(\k=a..b. f (real_of_int k)) = F (of_int b) + C + (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b)) - EM_remainder (2*N+1) (fs (2*N+1)) b" by (rule euler_maclaurin_strong_int) (simp_all only: lessThan_Suc_atMost Suc_eq_plus1 [symmetric] assms) also have "{..<2*N+1} = insert 0 {1..2*N}" by auto also have "(\i\\. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b)) = (1 / 2) *\<^sub>R f (of_int b) + (\i\{1..2*N}. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b))" by (subst sum.insert) (auto simp: assms bernoulli'_def) also have "(\i\{1..2*N}. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b)) = (\i\{1..N}. (bernoulli' (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (of_int b))" proof (rule sym, rule sum.reindex_bij_witness_not_neutral) fix i assume "i \ {1..2*N} - {i\{1..2*N}. even i}" thus "2 * ((i + 1) div 2) - 1 = i" "(i + 1) div 2 \ {1..N} - {}" by (auto elim!: oddE) qed (auto simp: bernoulli_odd_eq_0 bernoulli'_def algebra_simps) also have "\ = (\i\{1..N}. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (of_int b))" by (intro sum.cong refl) (auto simp: bernoulli'_def) finally show ?thesis by (simp only: add_ac) qed end text \ For convenience, we also offer a version where the sum ranges over natural numbers instead of integers. \ lemma sum_atLeastAtMost_of_int_nat_transfer: "(\k=int a..int b. f (of_int k)) = (\k=a..b. f (of_nat k))" by (intro sum.reindex_bij_witness[of _ int nat]) auto lemma euler_maclaurin_nat_int_transfer: fixes F and f :: "real \ 'a :: real_normed_vector" assumes "((\b. (\k=a..b. f (real k)) - F (real b) - T (real b)) \ C) at_top" shows "((\b. (\k=int a..b. f (of_int k)) - F (of_int b) - T (of_int b)) \ C) at_top" proof - have *: "((\b. (\k=int a..int b. f (of_int k)) - F (of_int (int b)) - T (of_int (int b))) \ C) at_top" using assms by (subst sum_atLeastAtMost_of_int_nat_transfer) simp thus ?thesis by (rule filterlim_int_of_nat_at_topD) qed locale euler_maclaurin_nat = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: nat fixes N :: nat assumes N: "N > 0" fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ N \ continuous_on {real a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k < N \ x \ {real a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {real a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {real a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (real b) - (\iR fs i (real b))) \ C) at_top" begin sublocale euler_maclaurin_int F f fs "int a" N C Y by standard (insert N fin fs_cont fs_deriv F_cont F_deriv euler_maclaurin_nat_int_transfer[OF limit], simp_all) theorem euler_maclaurin_strong_nat: assumes ab: "a \ b" defines "S \ (\k=a..b. f (real k))" shows "S = F (real b) + C + (\iR fs i (real b)) - EM_remainder N (fs N) (int b)" using euler_maclaurin_strong_int[of "int b"] by (simp add: assms sum_atLeastAtMost_of_int_nat_transfer) end locale euler_maclaurin_nat' = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: nat fixes N :: nat fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ 2*N+1 \ continuous_on {real a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k \ 2*N \ x \ {real a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {real a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {real a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (real b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (real b))) \ C) at_top" begin sublocale euler_maclaurin_int' F f fs "int a" N C Y by standard (insert fin fs_cont fs_deriv F_cont F_deriv euler_maclaurin_nat_int_transfer[OF limit], simp_all) theorem euler_maclaurin_strong_nat': assumes "a \ b" shows "(\k=a..b. f (real k)) = F (real b) + C + (1 / 2) *\<^sub>R f (real b) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (real b)) - EM_remainder (2*N+1) (fs (2*N+1)) b" using euler_maclaurin_strong_int'[of b] by (simp add: assms sum_atLeastAtMost_of_int_nat_transfer) end subsection \Bounds on the remainder term\ text \ The following theorems provide some simple means to bound the remainder terms. In practice, better bounds can often be obtained e.\,g. for the $n$-th remainder term by expanding it to the sum of the first discarded term in the expansion and the $n+1$-th remainder term. \ lemma fixes f :: "real \ 'a :: {real_normed_field, banach}" and g g' :: "real \ real" assumes fin: "finite Y" assumes pbernpoly_bound: "\x. \pbernpoly n x\ \ D" assumes cont_f: "continuous_on {a..} f" assumes cont_g: "continuous_on {a..} g" assumes cont_g': "continuous_on {a..} g'" assumes limit_g: "(g \ C) at_top" assumes f_bound: "\x. x \ a \ norm (f x) \ g' x" assumes deriv: "\x. x \ {a..} - Y \ (g has_field_derivative g' x) (at x)" shows norm_EM_remainder_le_strong_int: "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * (C - g x)" and norm_EM_remainder_le_strong_nat: "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * (C - g x)" proof - from pbernpoly_bound have D: "norm (pbernpoly n x) \ D" for x by auto from this[of 0] have D_nonneg: "D \ 0" by simp define D' where "D' = D / fact n" from D_nonneg have D'_nonneg: "D' \ 0" by (simp add: D'_def) have bound: "norm (EM_remainder' n f x y) \ D' * (g y - g x)" if xy: "x \ a" "x \ y" for x y :: real proof - have "norm (EM_remainder' n f x y) = norm (integral {x..y} (\t. pbernpoly n t *\<^sub>R f t)) / fact n" by (simp add: EM_remainder'_def) also have "(\t. D * g' t) integrable_on {x..y}" using xy by (intro integrable_continuous_real continuous_intros continuous_on_subset[OF cont_g']) auto hence "norm (integral {x..y} (\t. pbernpoly n t *\<^sub>R f t)) \ integral {x..y} (\t. D * g' t)" using D D_nonneg xy by (intro integral_norm_bound_integral integrable_EM_remainder' continuous_on_subset[OF cont_f]) (auto intro!: mult_mono f_bound) also have "\ = D * integral {x..y} g'" by simp also have "(g' has_integral (g y - g x)) {x..y}" using xy by (intro fundamental_theorem_of_calculus_strong[OF fin] continuous_on_subset[OF cont_g]) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] intro!: deriv) hence "integral {x..y} g' = g y - g x" by (simp add: has_integral_iff) finally show ?thesis by (simp add: D'_def divide_simps) qed have lim: "((\y. EM_remainder' n f x (of_int y)) \ EM_remainder n f x) at_top" if x: "x \ a" for x :: int proof - have "(\n. g (real n)) \ C" by (rule filterlim_compose[OF limit_g filterlim_real_sequentially]) hence Cauchy: "Cauchy (\n. g (real n))" using convergent_eq_Cauchy by blast have "Cauchy (\y. EM_remainder' n f x (int y))" proof (rule CauchyI', goal_cases) case (1 \) define \' where "\' = (if D' = 0 then 1 else \ / (2*D'))" from \\ > 0\ D'_nonneg have \': "\' > 0" by (simp add: \'_def divide_simps) from CauchyD[OF Cauchy this] obtain M where M: "\m n. m \ M \ n \ M \ norm (g (real m) - g (real n)) < \'" by blast show ?case proof (intro CauchyI' exI[of _ "max (max 0 M) (nat x)"] allI impI, goal_cases) case (1 k l) have "EM_remainder' n f x k + EM_remainder' n f k l = EM_remainder' n f x l" using 1 x by (intro EM_remainder'_combine continuous_on_subset[OF cont_f]) auto hence "EM_remainder' n f x l - EM_remainder' n f x k = EM_remainder' n f k l" by (simp add: algebra_simps) also from 1 x have "norm \ \ D' * (g l - g k)" by (intro bound) auto also have "g l - g k \ norm (g l - g k)" by simp also from 1 have "\ \ \'" using M[of l k] by auto also from \\ > 0\ have "D' * \' \ \ / 2" by (simp add: \'_def) also from \\ > 0\ have "\ < \" by simp finally show ?case by (simp add: D'_nonneg mult_left_mono dist_norm norm_minus_commute) qed qed then obtain L where "(\y. EM_remainder' n f x (int y)) \ L" by (auto simp: Cauchy_convergent_iff convergent_def) from filterlim_int_of_nat_at_topD[OF this] have "((\y. EM_remainder' n f x (of_int y)) \ L) at_top" by simp moreover from this have "EM_remainder n f x = L" by (rule EM_remainder_eqI) ultimately show "((\y. EM_remainder' n f x (of_int y)) \ EM_remainder n f x) at_top" by simp qed have *: "norm (EM_remainder n f x) \ D' * (C - g x)" if x: "x \ a" for x :: int proof (rule tendsto_le) show "((\y. D' * (g (of_int y) - g (of_int x))) \ D' * (C - g (of_int x))) at_top" by (intro tendsto_intros filterlim_compose[OF limit_g]) show "((\y. norm (EM_remainder' n f x (of_int y))) \ norm (EM_remainder n f x)) at_top" using x by (intro tendsto_norm lim) show "eventually (\y. norm (EM_remainder' n f (of_int x) (of_int y)) \ D' * (g (of_int y) - g (of_int x))) at_top" using eventually_ge_at_top[of x] by eventually_elim (rule bound, insert x, simp_all) qed simp_all thus "\x. of_int x \ a \ norm (EM_remainder n f x) \ D' * (C - g x)" by blast have "norm (EM_remainder n f x) \ D' * (C - g x)" if x: "x \ a" for x :: nat using x *[of "int x"] by simp thus "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D' * (C - g x)" by blast qed lemma fixes f :: "real \ 'a :: {real_normed_field, banach}" and g g' :: "real \ real" assumes fin: "finite Y" assumes pbernpoly_bound: "\x. \pbernpoly n x\ \ D" assumes cont_f: "continuous_on {a..} f" assumes cont_g: "continuous_on {a..} g" assumes cont_g': "continuous_on {a..} g'" assumes limit_g: "(g \ 0) at_top" assumes f_bound: "\x. x \ a \ norm (f x) \ g' x" assumes deriv: "\x. x \ {a..} - Y \ (g has_field_derivative -g' x) (at x)" shows norm_EM_remainder_le_strong_int': "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * g x" and norm_EM_remainder_le_strong_nat': "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * g x" proof - have "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * (0 - (-g x))" using assms by (intro norm_EM_remainder_le_strong_int[OF fin pbernpoly_bound _ _ cont_g']) (auto intro!: continuous_intros tendsto_eq_intros derivative_eq_intros) thus "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * g x" by auto next have "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * (0 - (-g x))" using assms by (intro norm_EM_remainder_le_strong_nat[OF fin pbernpoly_bound _ _ cont_g']) (auto intro!: continuous_intros tendsto_eq_intros derivative_eq_intros) thus "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * g x" by auto qed lemma norm_EM_remainder'_le: fixes f :: "real \ 'a :: {real_normed_field, banach}" and g g' :: "real \ real" assumes cont_f: "continuous_on {a..} f" assumes cont_g': "continuous_on {a..} g'" assumes f_bigo: "eventually (\x. norm (f x) \ g' x) at_top" assumes deriv: "eventually (\x. (g has_field_derivative g' x) (at x)) at_top" obtains C D where "eventually (\x. norm (EM_remainder' n f a x) \ C + D * g x) at_top" proof - note cont = continuous_on_subset[OF cont_f] continuous_on_subset[OF cont_g'] from bounded_pbernpoly[of n] obtain D where D: "\x. norm (pbernpoly n x) \ D" by blast from this[of 0] have D_nonneg: "D \ 0" by simp from eventually_conj[OF f_bigo eventually_conj[OF deriv eventually_ge_at_top[of a]]] obtain x0 where x0: "x0 \ a" "\x. x \ x0 \ norm (f x) \ g' x" "\x. x \ x0 \ (g has_field_derivative g' x) (at x)" by (auto simp: eventually_at_top_linorder) define C where "C = (norm (integral {a..x0} (\t. pbernpoly n t *\<^sub>R f t)) - D * g x0) / fact n" have "eventually (\x. norm (EM_remainder' n f a x) \ C + D / fact n * g x) at_top" using eventually_ge_at_top[of x0] proof eventually_elim case (elim x) have "integral {a..x} (\t. pbernpoly n t *\<^sub>R f t) = integral {a..x0} (\t. pbernpoly n t *\<^sub>R f t) + integral {x0..x} (\t. pbernpoly n t *\<^sub>R f t)" (is "_ = ?I1 + ?I2") using elim x0(1) by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric] integrable_EM_remainder' cont) auto also have "norm \ \ norm ?I1 + norm ?I2" by (rule norm_triangle_ineq) also have "norm ?I2 \ integral {x0..x} (\t. D * g' t)" using x0 D D_nonneg by (intro integral_norm_bound_integral integrable_EM_remainder') (auto intro!: integrable_continuous_real continuous_intros cont mult_mono) also have "\ = D * integral {x0..x} g'" by simp also from elim have "(g' has_integral (g x - g x0)) {x0..x}" by (intro fundamental_theorem_of_calculus) (auto intro!: has_field_derivative_at_within[OF x0(3)] simp: has_field_derivative_iff_has_vector_derivative [symmetric]) hence "integral {x0..x} g' = g x - g x0" by (simp add: has_integral_iff) finally have "norm (integral {a..x} (\t. pbernpoly n t *\<^sub>R f t)) \ norm ?I1 + D * (g x - g x0)" by simp_all thus "norm (EM_remainder' n f a x) \ C + D / fact n * g x" by (simp add: EM_remainder'_def field_simps C_def) qed thus ?thesis by (rule that) qed subsection \Application to harmonic numbers\ text \ As a first application, we can apply the machinery we have developed to the harmonic numbers. \ definition harm_remainder :: "nat \ nat \ real" where "harm_remainder N n = EM_remainder (2*N+1) (\x. -fact (2*N+1) / x ^ (2*N+2)) (int n)" lemma harm_expansion: assumes n: "n > 0" and N: "N > 0" shows "harm n = ln n + euler_mascheroni + 1 / (2*n) - (\i=1..N. bernoulli (2*i) / ((2*i) * n ^ (2*i))) - harm_remainder N n" proof - define fs where "fs = (\k x. (-1) ^ k * fact k / x ^ (Suc k) :: real)" interpret euler_maclaurin_nat' ln "\x. 1/x" fs 1 N euler_mascheroni "{}" proof fix k x assume "k \ 2*N" "x \ {real 1..} - {}" thus "(fs k has_vector_derivative fs (Suc k) x) (at x)" by (cases "k = 0") (auto intro!: derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric] field_simps power_diff) next have "(\b. harm b - ln (real b) - (\i<2*N+1. bernoulli' (Suc i) * (- 1) ^ i / (real (Suc i) * (real b ^ Suc i)))) \ (euler_mascheroni - (\i<2*N+1. 0))" by (intro tendsto_diff euler_mascheroni_LIMSEQ tendsto_sum real_tendsto_divide_at_top[OF tendsto_const] filterlim_tendsto_pos_mult_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_real_sequentially) auto thus "(\b. (\k = 1..b. 1 / real k) - ln (real b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (real b))) \ euler_mascheroni" by (simp add: harm_def divide_simps fs_def) qed (insert n N, auto intro!: continuous_intros derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric]) have "harm n = (\k=1..n. 1 / real k)" by (simp add: harm_def divide_simps) also have "\ = ln (real n) + euler_mascheroni + (1/2) *\<^sub>R (1 / real n) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (real n)) - EM_remainder (2*N+1) (fs (2*N+1)) (int n)" using n N using n by (intro euler_maclaurin_strong_nat') simp_all also have "(\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R (fs (2*i-1) (real n))) = (\i=1..N. -(bernoulli (2*i) / (real (2*i) * real n ^ (2*i))))" by (intro sum.cong refl) (simp_all add: fs_def divide_simps fact_reduce del: of_nat_Suc power_Suc) also have "\ = -(\i=1..N. bernoulli (2*i) / (real (2*i) * real n ^ (2*i)))" by (simp add: sum_negf) finally show ?thesis unfolding fs_def by (simp add: harm_remainder_def) qed lemma of_nat_ge_1_iff: "of_nat x \ (1 :: 'a :: linordered_semidom) \ x \ 1" using of_nat_le_iff[of 1 x] by (simp del: of_nat_le_iff) lemma harm_remainder_bound: fixes N :: nat assumes N: "N > 0" shows "\C. \n\1. norm (harm_remainder N n) \ C / real n ^ (2*N+1)" proof - from bounded_pbernpoly[of "2*N+1"] obtain D where D: "\x. \pbernpoly (2*N+1) x\ \ D" by auto have "\x. 1 \ real x \ norm (harm_remainder N x) \ D / fact (2*N+1) * (fact (2*N) / x ^ (2*N+1))" unfolding harm_remainder_def of_int_of_nat_eq proof (rule norm_EM_remainder_le_strong_nat'[of "{}"]) fix x :: real assume x: "x \ 1" show "norm (-fact (2*N+1) / x ^ (2 * N + 2)) \ fact (2*N+1) / x ^ (2*N+2)" using x by simp next show "((\x::real. fact (2 * N) / x ^ (2 * N + 1)) \ 0) at_top" by (intro real_tendsto_divide_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_ident) simp_all qed (insert N D, auto intro!: derivative_eq_intros continuous_intros simp: field_simps power_diff) hence "\x. 1 \ x \ norm (harm_remainder N x) \ D / (2*N+1) / real x ^ (2*N+1)" by simp thus ?thesis by blast qed subsection \Application to sums of inverse squares\ text \ In the same vein, we can derive the asymptotics of the partial sum of inverse squares. \ lemma sum_inverse_squares_expansion: assumes n: "n > 0" and N: "N > 0" shows "(\k=1..n. 1 / real k ^ 2) = pi ^ 2 / 6 - 1 / real n + 1 / (2 * real n ^ 2) - (\i=1..N. bernoulli (2*i) / n ^ (2*i+1)) - EM_remainder (2*N+1) (\x. -fact (2*N+2) / x ^ (2*N+3)) (int n)" proof - have 3: "3 = Suc (Suc (Suc 0))" by (simp add: eval_nat_numeral) define fs where "fs = (\k x. (-1) ^ k * fact (Suc k) / x ^ (k+2) :: real)" interpret euler_maclaurin_nat' "\x. -1/x" "\x. 1/x^2" fs 1 N "pi^2/6" "{}" proof fix k x assume "k \ 2*N" "x \ {real 1..} - {}" thus "(fs k has_vector_derivative fs (Suc k) x) (at x)" by (cases "k = 0") (auto intro!: derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric] field_simps power_diff) next from inverse_squares_sums have "(\n. \k pi\<^sup>2 / 6" by (simp add: sums_def) also have "(\n. \kn. \k=1..n. 1 / real k ^ 2)" by (intro ext sum.reindex_bij_witness[of _ "\n. n - 1" Suc]) auto finally have "(\b. (\k = 1..b. 1 / real k^2) + 1 / real b - (\i<2*N+1. bernoulli' (Suc i) * (- 1) ^ i / (real b ^ (i+2)))) \ (pi^2/6 + 0 - (\i<2*N+1. 0))" by (intro tendsto_diff tendsto_add real_tendsto_divide_at_top[OF tendsto_const] filterlim_tendsto_pos_mult_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_real_sequentially tendsto_sum) auto thus "(\b. (\k = 1..b. 1 / real k^2) - (- 1 / real b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (real b))) \ pi^2/6" by (simp add: harm_def field_simps fs_def del: power_Suc of_nat_Suc) qed (insert n N, auto intro!: continuous_intros derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric] power2_eq_square) have "(\k=1..n. 1 / real k ^ 2) = - 1 / real n + pi^2/6 + (1/2) *\<^sub>R (1 / real n^2) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (real n)) - EM_remainder (2*N+1) (fs (2*N+1)) (int n)" using n N using n by (intro euler_maclaurin_strong_nat') simp_all also have "(\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R (fs (2*i-1) (real n))) = (\i=1..N. -(bernoulli (2*i) / (real n ^ (2*i+1))))" by (intro sum.cong refl) (simp_all add: fs_def divide_simps fact_reduce del: of_nat_Suc power_Suc) also have "\ = -(\i=1..N. bernoulli (2*i) / real n ^ (2*i+1))" by (simp add: sum_negf) finally show ?thesis unfolding fs_def by (simp add: fs_def 3) qed lemma sum_inverse_squares_remainder_bound: fixes N :: nat assumes N: "N > 0" defines "R \ (\n. EM_remainder (2*N+1) (\x. -fact (2*N+2) / x ^ (2*N+3)) (int n))" shows "\C. \n\1. norm (R n) \ C / real n ^ (2*N+2)" proof - have 3: "3 = Suc (Suc (Suc 0))" by simp from bounded_pbernpoly[of "2*N+1"] obtain D where D: "\x. \pbernpoly (2*N+1) x\ \ D" by auto have "\x. 1 \ real x \ norm (R x) \ D / fact (2*N+1) * (fact (2*N+1) / x ^ (2*N+2))" unfolding R_def of_int_of_nat_eq proof (rule norm_EM_remainder_le_strong_nat'[of "{}"]) fix x :: real assume x: "x \ 1" show "norm (-fact (2*N+2) / x ^ (2*N+3)) \ fact (2*N+2) / x ^ (2*N+3)" using x by simp next show "((\x::real. fact (2*N+1) / x ^ (2*N+2)) \ 0) at_top" by (intro real_tendsto_divide_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_ident) simp_all qed (insert N D, auto intro!: derivative_eq_intros continuous_intros simp: field_simps power_diff 3) hence "\x\1. norm (R x) \ D / real x ^ (2 * N + 2)" by simp thus ?thesis by blast qed end diff --git a/thys/Euler_MacLaurin/Euler_MacLaurin_Landau.thy b/thys/Euler_MacLaurin/Euler_MacLaurin_Landau.thy --- a/thys/Euler_MacLaurin/Euler_MacLaurin_Landau.thy +++ b/thys/Euler_MacLaurin/Euler_MacLaurin_Landau.thy @@ -1,170 +1,174 @@ section \Connection of Euler--MacLaurin summation to Landau symbols\ theory Euler_MacLaurin_Landau imports Euler_MacLaurin Landau_Symbols.Landau_More begin subsection \$O$-bound for the remainder term\ text \ Landau symbols allow us to state the bounds on the remainder terms from the Euler--MacLaurin formula a bit more nicely. \ lemma fixes f :: "real \ 'a :: {real_normed_field, banach}" and g g' :: "real \ real" assumes fin: "finite Y" assumes cont_f: "continuous_on {a..} f" assumes cont_g: "continuous_on {a..} g" assumes cont_g': "continuous_on {a..} g'" assumes limit_g: "(g \ 0) at_top" assumes f_bound: "\x. x \ a \ norm (f x) \ g' x" assumes deriv: "\x. x \ {a..} - Y \ (g has_field_derivative -g' x) (at x)" shows EM_remainder_strong_bigo_int: "(\x::int. norm (EM_remainder n f x)) \ O(g)" and EM_remainder_strong_bigo_nat: "(\x::nat. norm (EM_remainder n f x)) \ O(g)" proof - from bounded_pbernpoly[of n] obtain D where D: "\x. \pbernpoly n x\ \ D" by auto from norm_EM_remainder_le_strong_int'[OF fin D assms(2-)] have *: "\x. x \ a \ norm (EM_remainder n f x) \ D / fact n * g x" by auto have **: "eventually (\x::int. norm (EM_remainder n f x) \ abs (D / fact n) * abs (g x)) at_top" using eventually_ge_at_top[of "ceiling a"] proof eventually_elim case (elim x) with *[of x] have "norm (EM_remainder n f x) \ D / fact n * g x" by (simp add: ceiling_le_iff) also have "\ \ abs (D / fact n * g x)" by (rule abs_ge_self) also have "\ = abs (D / fact n) * abs (g x)" by (simp add: abs_mult) finally show ?case . qed thus "(\x::int. norm (EM_remainder n f x)) \ O(g)" by (intro bigoI[of _ "abs D / fact n"]) (auto elim!: eventually_mono) hence "(\x::nat. norm (EM_remainder n f (int x))) \ O(\x. g (of_int (int x)))" by (rule landau_o.big.compose) (fact filterlim_int_sequentially) thus "(\x::nat. norm (EM_remainder n f x)) \ O(g)" by simp qed subsection \Asymptotic expansion of the harmonic numbers\ text \ We can now show the asymptotic expansion \[H_n = \ln n + \gamma + \frac{1}{2n} - \sum_{i=1}^m \frac{B_{2i}}{2i} n^{-2i} + O(n^{-2m-2})\] \ lemma harm_remainder_bigo: assumes "N > 0" shows "harm_remainder N \ O(\n. 1 / real n ^ (2 * N + 1))" proof - - from harm_remainder_bound[OF assms] guess C .. + from harm_remainder_bound[OF assms] + obtain C where "\n\1. norm (harm_remainder N n) \ C / real n ^ (2 * N + 1)" .. thus ?thesis by (intro bigoI[of _ C] eventually_mono[OF eventually_ge_at_top[of 1]]) auto qed lemma harm_expansion_bigo: fixes N :: nat defines "T \ \n. ln n + euler_mascheroni + 1 / (2*n) - (\i=1..N. bernoulli (2*i) / ((2*i) * n ^ (2*i)))" defines "S \ (\n. bernoulli (2*(Suc N)) / ((2*Suc N) * real n ^ (2*Suc N)))" shows "(\n. harm n - T n) \ O(\n. 1 / real n ^ (2 * N + 2))" proof - have "(\n. harm n - T n) \ \(\n. -S n - harm_remainder (Suc N) n)" by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) (auto simp: T_def harm_expansion[of _ "Suc N"] S_def) also have "(\n. -S n - harm_remainder (Suc N) n) \ O(\n. 1 / real n ^ (2 * N + 2))" proof (intro sum_in_bigo) show "(\x. - S x) \ O(\n. 1 / real n ^ (2 * N + 2))" unfolding S_def by (rule landau_o.big.compose[OF _ filterlim_real_sequentially]) simp have "harm_remainder (Suc N) \ O(\n. 1 / real n ^ (2 * Suc N + 1))" by (rule harm_remainder_bigo) simp_all also have "(\n. 1 / real n ^ (2 * Suc N + 1)) \ O(\n. 1 / real n ^ (2 * N + 2))" by (rule landau_o.big.compose[OF _ filterlim_real_sequentially]) simp finally show "harm_remainder (Suc N) \ \" . qed finally show ?thesis . qed lemma harm_expansion_bigo_simple1: "(\n. harm n - (ln n + euler_mascheroni + 1 / (2 * n))) \ O(\n. 1 / n ^ 2)" using harm_expansion_bigo[of 0] by (simp add: power2_eq_square) lemma harm_expansion_bigo_simple2: "(\n. harm n - (ln n + euler_mascheroni)) \ O(\n. 1 / n)" proof - have "(\n. harm n - (ln n + euler_mascheroni + 1 / (2 * n)) + 1 / (2 * n)) \ O(\n. 1 / n)" proof (rule sum_in_bigo) have "(\n. harm n - (ln n + euler_mascheroni + 1 / (2 * n))) \ O(\n. 1 / real n ^ 2)" using harm_expansion_bigo_simple1 by simp also have "(\n. 1 / real n ^ 2) \ O(\n. 1 / real n)" by (rule landau_o.big.compose[OF _ filterlim_real_sequentially]) simp_all finally show "(\n. harm n - (ln n + euler_mascheroni + 1 / (2 * n))) \ O(\n. 1 / n)" by simp qed simp_all thus ?thesis by (simp add: algebra_simps) qed lemma harm_expansion_bigo_simple': "harm =o (\n. ln n + euler_mascheroni + 1 / (2 * n)) +o O(\n. 1 / n ^ 2)" using harm_expansion_bigo_simple1 by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def) subsection \Asymptotic expansion of the sum of inverse squares\ text \ Similarly to before, we show \[\sum_{i=1}^n \frac{1}{i^2} = \frac{\pi^2}{6} - \frac{1}{n} + \frac{1}{2n^2} - \sum_{i=1}^m B_{2i} n^{-2i-1} + O(n^{-2m-3})\] \ context fixes R :: "nat \ nat \ real" defines "R \ (\N n. EM_remainder (2*N+1) (\x. -fact (2*N+2) / x ^ (2*N+3)) (int n))" begin lemma sum_inverse_squares_remainder_bigo: assumes "N > 0" shows "R N \ O(\n. 1 / real n ^ (2 * N + 2))" proof - - from sum_inverse_squares_remainder_bound[OF assms] guess C .. + from sum_inverse_squares_remainder_bound[OF assms] + obtain C + where "\n\1. norm (EM_remainder (2 * N + 1) (\x. - fact (2 * N + 2) / x ^ (2 * N + 3)) (int n)) + \ C / real n ^ (2 * N + 2)" .. thus ?thesis by (intro bigoI[of _ C] eventually_mono[OF eventually_ge_at_top[of 1]]) (auto simp: R_def) qed lemma sum_inverse_squares_expansion_bigo: fixes N :: nat defines "T \ \n. pi ^ 2 / 6 - 1 / n + 1 / (2*n ^ 2) - (\i=1..N. bernoulli (2*i) / (n ^ (2*i+1)))" defines "S \ (\n. bernoulli (2*(Suc N)) / (real n ^ (2*N+3)))" shows "(\n. (\i=1..n. 1 / real i ^ 2) - T n) \ O(\n. 1 / real n ^ (2 * N + 3))" proof - have 3: "3 = Suc (Suc (Suc 0))" by simp have "(\n. (\i=1..n. 1 / real i ^ 2) - T n) \ \(\n. -S n - R (Suc N) n)" unfolding R_def by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) (auto simp: T_def sum_inverse_squares_expansion[of _ "Suc N"] S_def 3 simp del: One_nat_def) also have "(\n. -S n - R (Suc N) n) \ O(\n. 1 / real n ^ (2 * N + 3))" proof (intro sum_in_bigo) show "(\x. - S x) \ O(\n. 1 / real n ^ (2 * N + 3))" unfolding S_def by (rule landau_o.big.compose[OF _ filterlim_real_sequentially]) simp have "R (Suc N) \ O(\n. 1 / real n ^ (2 * Suc N + 2))" by (rule sum_inverse_squares_remainder_bigo) simp_all also have "2 * Suc N + 2 = 2 * N + 4" by simp also have "(\n. 1 / real n ^ (2 * N + 4)) \ O(\n. 1 / real n ^ (2 * N + 3))" by (rule landau_o.big.compose[OF _ filterlim_real_sequentially]) simp finally show "R (Suc N) \ \" . qed finally show ?thesis . qed lemma sum_inverse_squares_expansion_bigo_simple: "(\n. (\i=1..n. 1 / real i ^ 2) - (pi ^ 2 / 6 - 1 / n + 1 / (2*n^2))) \ O(\n. 1 / n ^ 3)" using sum_inverse_squares_expansion_bigo[of 0] by (simp add: power2_eq_square) lemma sum_inverse_squares_expansion_bigo_simple': "(\n. (\i=1..n. 1 / real i ^ 2)) =o (\n. pi ^ 2 / 6 - 1 / n + 1 / (2*n^2)) +o O(\n. 1 / n^3)" using sum_inverse_squares_expansion_bigo_simple by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def) end end diff --git a/thys/Fishburn_Impossibility/Fishburn_Impossibility.thy b/thys/Fishburn_Impossibility/Fishburn_Impossibility.thy --- a/thys/Fishburn_Impossibility/Fishburn_Impossibility.thy +++ b/thys/Fishburn_Impossibility/Fishburn_Impossibility.thy @@ -1,346 +1,347 @@ (* File: Fishburn_Impossibility.thy Author: Manuel Eberl, TU München Author: Christian Saile, TU München Incompatibility of Pareto Optimality and Fishburn-Strategy-Proofness *) section \Main impossibility result\ theory Fishburn_Impossibility imports Social_Choice_Functions begin subsection \Setting of the base case\ text \ Suppose we have an anonymous, Fishburn-strategyproof, and Pareto-efficient SCF for three agents $A1$ to $A3$ and three alternatives $a$, $b$, and $c$. We will derive a contradiction from this. \ locale fb_impossibility_3_3 = strategyproof_anonymous_scf agents alts scf Fishb + pareto_efficient_scf agents alts scf for agents :: "'agent set" and alts :: "'alt set" and scf A1 A2 A3 a b c + assumes agents_eq: "agents = {A1, A2, A3}" assumes alts_eq: "alts = {a, b, c}" assumes distinct_agents: "distinct [A1, A2, A3]" assumes distinct_alts: "distinct [a, b, c]" begin text \ We first give some simple rules that will allow us to break down the strategyproofness and support conditions more easily later. \ lemma agents_neq [simp]: "A1 \ A2" "A2 \ A1" "A1 \ A3" "A3 \ A1" "A2 \ A3" "A3 \ A2" using distinct_agents by auto lemma alts_neq [simp]: "a \ b" "a \ c" "b \ c" "b \ a" "c \ a" "c \ b" using distinct_alts by auto lemma agent_in_agents [simp]: "A1 \ agents" "A2 \ agents" "A3 \ agents" by (simp_all add: agents_eq) lemma alt_in_alts [simp]: "a \ alts" "b \ alts" "c \ alts" by (simp_all add: alts_eq) lemma Bex_alts: "(\x\alts. P x) \ P a \ P b \ P c" by (simp add: alts_eq) lemma eval_pareto_loser_aux: assumes "is_pref_profile R" shows "x \ pareto_losers R \ (\y\{a,b,c}. x \[Pareto(R)] y)" proof - interpret pref_profile_wf agents alts R by fact have *: "y \ {a,b,c}" if "x \[Pareto(R)] y" for y using Pareto.strict_not_outside[of x y] that by (simp add: alts_eq) show ?thesis by (auto simp: pareto_losers_def dest: *) qed lemma eval_Pareto: assumes "is_pref_profile R" shows "x \[Pareto(R)] y \ (\i\{A1,A2,A3}. x \[R i] y) \ (\i\{A1,A2,A3}. \x \[R i] y)" proof - interpret R: pref_profile_wf agents alts by fact show ?thesis unfolding R.Pareto_strict_iff by (auto simp: strongly_preferred_def agents_eq) qed lemmas eval_pareto = eval_pareto_loser_aux eval_Pareto lemma pareto_efficiency: "is_pref_profile R \ x \ pareto_losers R \ x \ scf R" using pareto_efficient[of R] by blast lemma Ball_scf: assumes "is_pref_profile R" shows "(\x\scf R. P x) \ (a \ scf R \ P a) \ (b \ scf R \ P b) \ (c \ scf R \ P c)" using scf_alts[OF assms] unfolding alts_eq by blast lemma Ball_scf_diff: assumes "is_pref_profile R1" "is_pref_profile R2" shows "(\x\scf R1 - scf R2. P x) \ (a \ scf R2 \ a \ scf R1 \ P a) \ (b \ scf R2 \ b \ scf R1 \ P b) \ (c \ scf R2 \ c \ scf R1 \ P c)" using assms[THEN scf_alts] unfolding alts_eq by blast lemma scf_nonempty': assumes "is_pref_profile R" shows "\x\alts. x \ scf R" using scf_nonempty[OF assms] scf_alts[OF assms] by blast subsection \Definition of Preference Profiles and Fact Gathering\ text \ We now define the 21 preference profile that will lead to the impossibility result. \ preference_profile agents: agents alts: alts where R1 = A1: [a, c], b A2: [a, c], b A3: b, c, a and R2 = A1: c, [a, b] A2: b, c, a A3: c, b, a and R3 = A1: [a, c], b A2: b, c, a A3: c, b, a and R4 = A1: [a, c], b A2: a, b, c A3: b, c, a and R5 = A1: c, [a, b] A2: a, b, c A3: b, c, a and R6 = A1: b, [a, c] A2: c, [a, b] A3: b, c, a and R7 = A1: [a, c], b A2: b, [a, c] A3: b, c, a and R8 = A1: [b, c], a A2: a, [b, c] A3: a, c, b and R9 = A1: [b, c], a A2: b, [a, c] A3: a, b, c and R10 = A1: c, [a, b] A2: a, b, c A3: c, b, a and R11 = A1: [a, c], b A2: a, b, c A3: c, b, a and R12 = A1: c, [a, b] A2: b, a, c A3: c, b, a and R13 = A1: [a, c], b A2: b, a, c A3: c, b, a and R14 = A1: a, [b, c] A2: c, [a, b] A3: a, c, b and R15 = A1: [b, c], a A2: a, [b, c] A3: a, b, c and R16 = A1: [a, b], c A2: c, [a, b] A3: a, b, c and R17 = A1: a, [b, c] A2: a, b, c A3: b, c, a and R18 = A1: [a, c], b A2: b, [a, c] A3: b, a, c and R19 = A1: a, [b, c] A2: c, [a, b] A3: a, b, c and R20 = A1: b, [a, c] A2: a, b, c A3: b, a, c and R21 = A1: [b, c], a A2: a, b, c A3: b, c, a by (simp_all add: agents_eq alts_eq) lemmas R_wfs = R1.wf R2.wf R3.wf R4.wf R5.wf R6.wf R7.wf R8.wf R9.wf R10.wf R11.wf R12.wf R13.wf R14.wf R15.wf R16.wf R17.wf R18.wf R19.wf R20.wf R21.wf lemmas R_evals = R1.eval R2.eval R3.eval R4.eval R5.eval R6.eval R7.eval R8.eval R9.eval R10.eval R11.eval R12.eval R13.eval R14.eval R15.eval R16.eval R17.eval R18.eval R19.eval R20.eval R21.eval lemmas nonemptiness = R_wfs [THEN scf_nonempty', unfolded Bex_alts] text \ We show the support conditions from Pareto efficiency \ lemma [simp]: "a \ scf R1" by (rule pareto_efficiency) (simp_all add: eval_pareto R1.eval) lemma [simp]: "a \ scf R2" by (rule pareto_efficiency) (simp_all add: eval_pareto R2.eval) lemma [simp]: "a \ scf R3" by (rule pareto_efficiency) (simp_all add: eval_pareto R3.eval) lemma [simp]: "a \ scf R6" by (rule pareto_efficiency) (simp_all add: eval_pareto R6.eval) lemma [simp]: "a \ scf R7" by (rule pareto_efficiency) (simp_all add: eval_pareto R7.eval) lemma [simp]: "b \ scf R8" by (rule pareto_efficiency) (simp_all add: eval_pareto R8.eval) lemma [simp]: "c \ scf R9" by (rule pareto_efficiency) (simp_all add: eval_pareto R9.eval) lemma [simp]: "a \ scf R12" by (rule pareto_efficiency) (simp_all add: eval_pareto R12.eval) lemma [simp]: "b \ scf R14" by (rule pareto_efficiency) (simp_all add: eval_pareto R14.eval) lemma [simp]: "c \ scf R15" by (rule pareto_efficiency) (simp_all add: eval_pareto R15.eval) lemma [simp]: "b \ scf R16" by (rule pareto_efficiency) (simp_all add: eval_pareto R16.eval) lemma [simp]: "c \ scf R17" by (rule pareto_efficiency) (simp_all add: eval_pareto R17.eval) lemma [simp]: "c \ scf R18" by (rule pareto_efficiency) (simp_all add: eval_pareto R18.eval) lemma [simp]: "b \ scf R19" by (rule pareto_efficiency) (simp_all add: eval_pareto R19.eval) lemma [simp]: "c \ scf R20" by (rule pareto_efficiency) (simp_all add: eval_pareto R20.eval) lemma [simp]: "c \ scf R21" by (rule pareto_efficiency) (simp_all add: eval_pareto R21.eval) text \ We derive the strategyproofness conditions: \ lemma s41: "\ scf R4 \[Fishb(R1 A2)] scf R1" by (intro strategyproof'[where j = A2]) (simp_all add: R4.eval R1.eval) lemma s32: "\ scf R3 \[Fishb(R2 A1)] scf R2" by (intro strategyproof'[where j = A1]) (simp_all add: R3.eval R2.eval) lemma s122: "\ scf R12 \[Fishb(R2 A2)] scf R2" by (intro strategyproof'[where j = A2]) (simp_all add: R12.eval R2.eval) lemma s133: "\ scf R13 \[Fishb(R3 A2)] scf R3" by (intro strategyproof'[where j = A2]) (simp_all add: R13.eval R3.eval) lemma s102: "\ scf R10 \[Fishb(R2 A2)] scf R2" by (intro strategyproof'[where j = A2]) (simp_all add: R10.eval R2.eval) lemma s13: "\ scf R1 \[Fishb(R3 A3)] scf R3" by (intro strategyproof'[where j = A2]) (simp_all add: R1.eval R3.eval) lemma s54: "\ scf R5 \[Fishb(R4 A1)] scf R4" by (intro strategyproof'[where j = A1]) (simp_all add: R5.eval R4.eval) lemma s174: "\ scf R17 \[Fishb(R4 A1)] scf R4" by (intro strategyproof'[where j = A1]) (simp_all add: R17.eval R4.eval) lemma s74: "\ scf R7 \[Fishb(R4 A2)] scf R4" by (intro strategyproof'[where j = A2]) (simp_all add: R7.eval R4.eval) lemma s114: "\ scf R11 \[Fishb(R4 A3)] scf R4" by (intro strategyproof'[where j = A3]) (simp_all add: R11.eval R4.eval) lemma s45: "\ scf R4 \[Fishb(R5 A1)] scf R5" by (intro strategyproof'[where j = A1]) (simp_all add: R4.eval R5.eval) lemma s65: "\ scf R6 \[Fishb(R5 A2)] scf R5" by (intro strategyproof'[where j = A1]) (simp_all add: insert_commute R5.eval R6.eval) lemma s105: "\ scf R10 \[Fishb(R5 A3)] scf R5" by (intro strategyproof'[where j = A3]) (simp_all add: R10.eval R5.eval) lemma s67: "\ scf R6 \[Fishb(R7 A1)] scf R7" by (intro strategyproof'[where j = A2]) (simp_all add: insert_commute R6.eval R7.eval) lemma s187: "\ scf R18 \[Fishb(R7 A3)] scf R7" by (intro strategyproof'[where j = A3]) (simp_all add: insert_commute R7.eval R18.eval) lemma s219: "\ scf R21 \[Fishb(R9 A2)] scf R9" by (intro strategyproof'[where j = A3]) (simp_all add: insert_commute R9.eval R21.eval) lemma s1011: "\ scf R10 \[Fishb(R11 A1)] scf R11" by (intro strategyproof'[where j = A1]) (simp_all add: insert_commute R10.eval R11.eval) lemma s1012: "\ scf R10 \[Fishb(R12 A2)] scf R12" by (intro strategyproof'[where j = A2]) (simp_all add: insert_commute R10.eval R12.eval) lemma s1213: "\ scf R12 \[Fishb(R13 A1)] scf R13" by (intro strategyproof'[where j = A1]) (simp_all add: insert_commute R12.eval R13.eval) lemma s1113: "\ scf R11 \[Fishb(R13 A2)] scf R13" by (intro strategyproof'[where j = A2]) (simp_all add: insert_commute R11.eval R13.eval) lemma s1813: "\ scf R18 \[Fishb(R13 A3)] scf R13" by (intro strategyproof'[where j = A2]) (simp_all add: insert_commute R18.eval R13.eval) lemma s814: "\ scf R8 \[Fishb(R14 A2)] scf R14" by (intro strategyproof'[where j = A1]) (simp_all add: insert_commute R8.eval R14.eval) lemma s1914: "\ scf R19 \[Fishb(R14 A3)] scf R14" by (intro strategyproof'[where j = A3]) (simp_all add: insert_commute R19.eval R14.eval) lemma s1715: "\ scf R17 \[Fishb(R15 A1)] scf R15" by (intro strategyproof'[where j = A3]) (simp_all add: insert_commute R17.eval R15.eval) lemma s815: "\ scf R8 \[Fishb(R15 A3)] scf R15" by (intro strategyproof'[where j = A3]) (simp_all add: insert_commute R8.eval R15.eval) lemma s516: "\ scf R5 \[Fishb(R16 A1)] scf R16" by (intro strategyproof'[where j = A3]) (simp_all add: insert_commute R5.eval R16.eval) lemma s517: "\ scf R5 \[Fishb(R17 A1)] scf R17" by (intro strategyproof'[where j = A1]) (simp_all add: insert_commute R5.eval R17.eval) lemma s1619: "\ scf R16 \[Fishb(R19 A1)] scf R19" by (intro strategyproof'[where j = A1]) (simp_all add: insert_commute R16.eval R19.eval) lemma s1820: "\ scf R18 \[Fishb(R20 A2)] scf R20" by (intro strategyproof'[where j = A1]) (simp_all add: insert_commute R18.eval R20.eval) lemma s920: "\ scf R9 \[Fishb(R20 A3)] scf R20" by (intro strategyproof'[where j = A1]) (simp_all add: insert_commute R20.eval R9.eval) lemma s521: "\ scf R5 \[Fishb(R21 A1)] scf R21" by (intro strategyproof'[where j = A1]) (simp_all add: insert_commute R21.eval R5.eval) lemma s421: "\ scf R4 \[Fishb(R21 A1)] scf R21" by (intro strategyproof'[where j = A1]) (simp_all add: insert_commute R21.eval R4.eval) lemmas sp = s41 s32 s122 s102 s133 s13 s54 s174 s54 s74 s114 s45 s65 s105 s67 s187 s219 s1011 s1012 s1213 s1113 s1813 s814 s1914 s1715 s815 s516 s517 s1619 s1820 s920 s521 s421 text \ We now use the simplifier to break down the strategyproofness conditions into SAT formulae. This takes a few seconds, so we use some low-level ML code to at least do the simplification in parallel. \ local_setup \fn lthy => let val lthy' = lthy addsimps @{thms Fishb_strict_iff Ball_scf Ball_scf_diff R_evals} val thms = Par_List.map (Simplifier.asm_full_simplify lthy') @{thms sp} in Local_Theory.notes [((@{binding sp'}, []), [(thms, [])])] lthy |> snd end \ text \ We show that the strategyproofness conditions, the non-emptiness conditions (i.\,e.\ every SCF must return at least one winner), and the efficiency conditions are not satisfiable together, which means that the SCF whose existence we assumed simply cannot exist. \ theorem absurd: False using sp' and nonemptiness [simplified] by satx end subsection \Lifting to more than 3 agents and alternatives\ text \ We now employ the standard lifting argument outlined before to lift this impossibility from 3 agents and alternatives to any setting with at least 3 agents and alternatives. \ locale fb_impossibility = strategyproof_anonymous_scf agents alts scf Fishb + pareto_efficient_scf agents alts scf for agents :: "'agent set" and alts :: "'alt set" and scf + assumes card_agents_ge: "card agents \ 3" and card_alts_ge: "card alts \ 3" begin (* TODO: Move? *) lemma finite_list': assumes "finite A" obtains xs where "A = set xs" "distinct xs" "length xs = card A" proof - from assms obtain xs where "set xs = A" using finite_list by blast thus ?thesis using distinct_card[of "remdups xs"] by (intro that[of "remdups xs"]) simp_all qed lemma finite_list_subset: assumes "finite A" "card A \ n" obtains xs where "set xs \ A" "distinct xs" "length xs = n" proof - obtain xs where "A = set xs" "distinct xs" "length xs = card A" using finite_list'[OF assms(1)] by blast with assms show ?thesis by (intro that[of "take n xs"]) (simp_all add: set_take_subset) qed lemma card_ge_3E: assumes "finite A" "card A \ 3" obtains a b c where "distinct [a,b,c]" "{a,b,c} \ A" proof - - from finite_list_subset[OF assms] guess xs . - moreover then obtain a b c where "xs = [a, b, c]" + from finite_list_subset[OF assms] + obtain xs where xs: "set xs \ A" "distinct xs" "length xs = 3" by auto + then obtain a b c where "xs = [a, b, c]" by (auto simp: eval_nat_numeral length_Suc_conv) - ultimately show ?thesis by (intro that[of a b c]) simp_all + with xs show ?thesis by (intro that[of a b c]) simp_all qed theorem absurd: False proof - - from card_ge_3E[OF finite_agents card_agents_ge] guess A1 A2 A3 . - note agents = this + from card_ge_3E[OF finite_agents card_agents_ge] obtain A1 A2 A3 + where agents: "distinct [A1, A2, A3]" "{A1, A2, A3} \ agents" . let ?agents' = "{A1, A2, A3}" - from card_ge_3E[OF finite_alts card_alts_ge] guess a b c . - note alts = this + from card_ge_3E[OF finite_alts card_alts_ge] + obtain a b c where alts: "distinct [a, b, c]" "{a, b, c} \ alts" . let ?alts' = "{a, b, c}" interpret scf_lowering_anonymous agents alts scf ?agents' ?alts' by standard (use agents alts in auto) interpret liftable_set_extension ?alts' alts Fishb by (intro Fishburn_liftable alts) interpret scf_lowering_strategyproof agents alts scf ?agents' ?alts' Fishb .. interpret fb_impossibility_3_3 ?agents' ?alts' lowered A1 A2 A3 a b c by standard (use agents alts in simp_all) from absurd show False . qed end end \ No newline at end of file diff --git a/thys/Fishburn_Impossibility/Social_Choice_Functions.thy b/thys/Fishburn_Impossibility/Social_Choice_Functions.thy --- a/thys/Fishburn_Impossibility/Social_Choice_Functions.thy +++ b/thys/Fishburn_Impossibility/Social_Choice_Functions.thy @@ -1,684 +1,685 @@ (* Title: Social_Choice_Functions.thy Author: Manuel Eberl, TU München Definitions of Social Choice Functions, their properties, and related concepts *) section \Social Choice Functions\ theory Social_Choice_Functions imports "Randomised_Social_Choice.Preference_Profile_Cmd" begin subsection \Weighted majority graphs\ definition supporters :: "('agent, 'alt) pref_profile \ 'alt \ 'alt \ 'agent set" where supporters_auxdef: "supporters R x y = {i. x \[R i] y}" definition weighted_majority :: "('agent, 'alt) pref_profile \ 'alt \ 'alt \ int" where "weighted_majority R x y = int (card (supporters R x y)) - int (card (supporters R y x))" lemma weighted_majority_refl [simp]: "weighted_majority R x x = 0" by (simp add: weighted_majority_def) lemma weighted_majority_swap: "weighted_majority R x y = -weighted_majority R y x" by (simp add: weighted_majority_def) lemma eval_set_filter: "Set.filter P {} = {}" "P x \ Set.filter P (insert x A) = insert x (Set.filter P A)" "\P x \ Set.filter P (insert x A) = Set.filter P A" unfolding Set.filter_def by auto context election begin lemma supporters_def: assumes "is_pref_profile R" shows "supporters R x y = {i\agents. x \[R i] y}" proof - interpret pref_profile_wf agents alts R by fact show ?thesis using not_outside unfolding supporters_auxdef by blast qed lemma supporters_nonagent1: assumes "is_pref_profile R" "x \ alts" shows "supporters R x y = {}" proof - interpret pref_profile_wf agents alts R by fact from assms show ?thesis by (auto simp: supporters_def dest: not_outside) qed lemma supporters_nonagent2: assumes "is_pref_profile R" "y \ alts" shows "supporters R x y = {}" proof - interpret pref_profile_wf agents alts R by fact from assms show ?thesis by (auto simp: supporters_def dest: not_outside) qed lemma weighted_majority_nonagent1: assumes "is_pref_profile R" "x \ alts" shows "weighted_majority R x y = 0" using assms by (simp add: weighted_majority_def supporters_nonagent1 supporters_nonagent2) lemma weighted_majority_nonagent2: assumes "is_pref_profile R" "y \ alts" shows "weighted_majority R x y = 0" using assms by (simp add: weighted_majority_def supporters_nonagent1 supporters_nonagent2) lemma weighted_majority_eq_iff: assumes "is_pref_profile R1" "is_pref_profile R2" shows "weighted_majority R1 = weighted_majority R2 \ (\x\alts. \y\alts. weighted_majority R1 x y = weighted_majority R2 x y)" proof (intro iffI ext) fix x y :: 'alt assume "\x\alts. \y\alts. weighted_majority R1 x y = weighted_majority R2 x y" with assms show "weighted_majority R1 x y = weighted_majority R2 x y" by (cases "x \ alts"; cases "y \ alts") (auto simp: fun_eq_iff weighted_majority_nonagent1 weighted_majority_nonagent2) qed auto end subsection \Definition of Social Choice Functions\ locale social_choice_function = election agents alts for agents :: "'agent set" and alts :: "'alt set" + fixes scf :: "('agent, 'alt) pref_profile \ 'alt set" assumes scf_nonempty: "is_pref_profile R \ scf R \ {}" assumes scf_alts: "is_pref_profile R \ scf R \ alts" subsection \Anonymity\ text \ An SCF is anonymous if permuting the agents in the input does not change the result. \ locale anonymous_scf = social_choice_function agents alts scf for agents :: "'agent set" and alts :: "'alt set" and scf + assumes anonymous: "\ permutes agents \ is_pref_profile R \ scf (R \ \) = scf R" begin lemma anonymous': assumes "anonymous_profile R1 = anonymous_profile R2" assumes "is_pref_profile R1" "is_pref_profile R2" shows "scf R1 = scf R2" proof - from anonymous_profile_agent_permutation[OF assms finite_agents] obtain \ where "\ permutes agents" "R1 = R2 \ \" by blast with anonymous[of \ R2] assms show ?thesis by simp qed lemma anonymity_prefs_from_table: assumes "prefs_from_table_wf agents alts xs" "prefs_from_table_wf agents alts ys" assumes "mset (map snd xs) = mset (map snd ys)" shows "scf (prefs_from_table xs) = scf (prefs_from_table ys)" proof - - from prefs_from_table_agent_permutation[OF assms] guess \ . + from assms obtain \ where "\ permutes agents" "prefs_from_table xs \ \ = prefs_from_table ys" + by (rule prefs_from_table_agent_permutation) with anonymous[of \, of "prefs_from_table xs"] assms(1) show ?thesis by (simp add: pref_profile_from_tableI) qed context begin qualified lemma anonymity_prefs_from_table_aux: assumes "R1 = prefs_from_table xs" "prefs_from_table_wf agents alts xs" assumes "R2 = prefs_from_table ys" "prefs_from_table_wf agents alts ys" assumes "mset (map snd xs) = mset (map snd ys)" shows "scf R1 = scf R2" unfolding assms(1,3) by (rule anonymity_prefs_from_table) (simp_all add: assms del: mset_map) end end subsection \Neutrality\ text \ An SCF is neutral if permuting the alternatives in the input does not change the result, modulo the equivalent permutation in the output lottery. \ locale neutral_scf = social_choice_function agents alts scf for agents :: "'agent set" and alts :: "'alt set" and scf + assumes neutral: "\ permutes alts \ is_pref_profile R \ scf (permute_profile \ R) = \ ` scf R" begin text \ Alternative formulation of neutrality that shows that our definition is equivalent to that in the paper. \ lemma neutral': assumes "\ permutes alts" assumes "is_pref_profile R" assumes "a \ alts" shows "\ a \ scf (permute_profile \ R) \ a \ scf R" proof - have *: "x = y" if "\ x = \ y" for x y using permutes_inj[OF assms(1)] that by (auto dest: injD) from assms show ?thesis by (auto simp: neutral dest!: *) qed end locale an_scf = anonymous_scf agents alts scf + neutral_scf agents alts scf for agents :: "'agent set" and alts :: "'alt set" and scf begin lemma scf_anonymous_neutral: assumes perm: "\ permutes alts" and wf: "is_pref_profile R1" "is_pref_profile R2" assumes eq: "anonymous_profile R1 = image_mset (map (\A. \ ` A)) (anonymous_profile R2)" shows "scf R1 = \ ` scf R2" proof - interpret R1: pref_profile_wf agents alts R1 by fact interpret R2: pref_profile_wf agents alts R2 by fact from perm have wf': "is_pref_profile (permute_profile \ R2)" by (rule R2.wf_permute_alts) from eq perm have "anonymous_profile R1 = anonymous_profile (permute_profile \ R2)" by (simp add: R2.anonymous_profile_permute) from anonymous_profile_agent_permutation[OF this wf(1) wf'] obtain \ where "\ permutes agents" "permute_profile \ R2 \ \ = R1" by auto have "scf (permute_profile \ R2 \ \) = scf (permute_profile \ R2)" by (rule anonymous) fact+ also have "\ = \ ` scf R2" by (rule neutral) fact+ also have "permute_profile \ R2 \ \ = R1" by fact finally show ?thesis . qed lemma scf_anonymous_neutral': assumes perm: "\ permutes alts" and wf: "is_pref_profile R1" "is_pref_profile R2" assumes eq: "anonymous_profile R1 = image_mset (map (\A. \ ` A)) (anonymous_profile R2)" shows "\ x \ scf R1 \ x \ scf R2" proof - have "scf R1 = \ ` scf R2" by (intro scf_anonymous_neutral) fact+ also have "\ x \ \ \ x \ scf R2" by (blast dest: injD[OF permutes_inj[OF perm]]) finally show ?thesis . qed lemma scf_automorphism: assumes perm: "\ permutes alts" and wf: "is_pref_profile R" assumes eq: "image_mset (map (\A. \ ` A)) (anonymous_profile R) = anonymous_profile R" shows "\ ` scf R = scf R" using scf_anonymous_neutral[OF perm wf wf eq [symmetric]] .. end lemma an_scf_automorphism_aux: assumes wf: "prefs_from_table_wf agents alts yss" "R \ prefs_from_table yss" assumes an: "an_scf agents alts scf" assumes eq: "mset (map ((map (\A. permutation_of_list xs ` A)) \ snd) yss) = mset (map snd yss)" assumes perm: "set (map fst xs) \ alts" "set (map snd xs) = set (map fst xs)" "distinct (map fst xs)" and x: "x \ alts" "y = permutation_of_list xs x" shows "x \ scf R \ y \ scf R" proof - note perm = list_permutesI[OF perm] let ?\ = "permutation_of_list xs" note perm' = permutation_of_list_permutes [OF perm] from wf have wf': "pref_profile_wf agents alts R" by (simp add: pref_profile_from_tableI) then interpret R: pref_profile_wf agents alts R . from perm' interpret R': pref_profile_wf agents alts "permute_profile ?\ R" by (simp add: R.wf_permute_alts) from an interpret an_scf agents alts scf . from eq wf have eq': "image_mset (map (\A. ?\ ` A)) (anonymous_profile R) = anonymous_profile R" by (simp add: anonymise_prefs_from_table mset_map multiset.map_comp) have "x \ scf R \ ?\ x \ ?\ ` scf R" by (blast dest: injD[OF permutes_inj[OF perm']]) also from eq' x wf' perm' have "?\ ` scf R = scf R" by (intro scf_automorphism) (simp_all add: R.anonymous_profile_permute pref_profile_from_tableI) finally show ?thesis using x by simp qed subsection \Weighted-majoritarian SCFs\ locale pairwise_scf = social_choice_function agents alts scf for agents :: "'agent set" and alts :: "'alt set" and scf + assumes pairwise: "is_pref_profile R1 \ is_pref_profile R2 \ weighted_majority R1 = weighted_majority R2 \ scf R1 = scf R2" subsection \Pareto efficiency\ locale pareto_efficient_scf = social_choice_function agents alts scf for agents :: "'agent set" and alts :: "'alt set" and scf + assumes pareto_efficient: "is_pref_profile R \ scf R \ pareto_losers R = {}" begin lemma pareto_efficient': assumes "is_pref_profile R" "y \[Pareto(R)] x" shows "x \ scf R" using pareto_efficient[of R] assms by (auto simp: pareto_losers_def) lemma pareto_efficient'': assumes "is_pref_profile R" "i \ agents" "\i\agents. y \[R i] x" "\y \[R i] x" shows "x \ scf R" proof - from assms(1) interpret pref_profile_wf agents alts R . from assms(2-) show ?thesis by (intro pareto_efficient'[OF assms(1), of _ y]) (auto simp: Pareto_iff strongly_preferred_def) qed end subsection \Set extensions\ type_synonym 'alt set_extension = "'alt relation \ 'alt set relation" definition Kelly :: "'alt set_extension" where "A \[Kelly(R)] B \ (\a\A. \b\B. a \[R] b)" lemma Kelly_strict_iff: "A \[Kelly(R)] B \ ((\a\A. \b\B. R b a) \ \ (\a\B. \b\A. R b a))" unfolding strongly_preferred_def Kelly_def .. lemmas Kelly_eval = Kelly_def Kelly_strict_iff definition Fishb :: "'alt set_extension" where "A \[Fishb(R)] B \ (\a\A. \b\B-A. a \[R] b) \ (\a\A-B. \b\B. a \[R] b)" lemma Fishb_strict_iff: "A \[Fishb(R)] B \ ((\a\A. \b\B - A. R b a) \ (\a\A - B. \b\B. R b a)) \ \ ((\a\B. \b\A - B. R b a) \ (\a\B - A. \b\A. R b a))" unfolding strongly_preferred_def Fishb_def .. lemmas Fishb_eval = Fishb_def Fishb_strict_iff subsection \Strategyproofness\ locale strategyproof_scf = social_choice_function agents alts scf for agents :: "'agent set" and alts :: "'alt set" and scf + fixes set_ext :: "'alt set_extension" assumes strategyproof: "is_pref_profile R \ total_preorder_on alts Ri' \ i \ agents \ \ scf (R(i := Ri')) \[set_ext(R i)] scf R" locale strategyproof_anonymous_scf = anonymous_scf agents alts scf + strategyproof_scf agents alts scf set_ext for agents :: "'agent set" and alts :: "'alt set" and scf and set_ext begin lemma strategyproof': assumes "is_pref_profile R1" "is_pref_profile R2" "i \ agents" "j \ agents" assumes "anonymous_profile R2 = anonymous_profile R1 - {#weak_ranking (R1 i)#} + {#weak_ranking (R2 j)#}" shows "\scf R2 \[set_ext (R1 i)] scf R1" proof - let ?R3 = "R1(i := R2 j)" interpret R1: pref_profile_wf agents alts R1 by fact interpret R2: pref_profile_wf agents alts R2 by fact from \j \ agents\ have "total_preorder_on alts (R2 j)" by simp from strategyproof[OF assms(1) this \i \ agents\] have "\scf ?R3 \[set_ext (R1 i)] scf R1" . also from assms have "scf ?R3 = scf R2" by (intro anonymous') (simp_all add: R1.anonymous_profile_update) finally show ?thesis . qed end context preorder_on begin lemma strict_not_outside: assumes "x \[le] y" shows "x \ carrier" "y \ carrier" using assms not_outside[of x y] unfolding strongly_preferred_def by blast+ end subsection \Lifting preferences\ text \ Preference profiles can be lifted to a setting with more agents and alternatives by padding them with dummy agents and alternatives in such a way that every original agent prefers and original alternative over any dummy alternative and is indifferent between the dummy alternatives. Moreover, every dummy agent is completely indifferent. \ definition lift_prefs :: "'alt set \ 'alt set \ 'alt relation \ 'alt relation" where "lift_prefs alts alts' R = (\x y. x \ alts' \ y \ alts' \ (x = y \ x \ alts \ (y \ alts \ R x y)))" lemma lift_prefs_wf: assumes "total_preorder_on alts R" "alts \ alts'" shows "total_preorder_on alts' (lift_prefs alts alts' R)" proof - interpret R: total_preorder_on alts R by fact show ?thesis by standard (insert R.total, auto dest: R.trans simp: lift_prefs_def) qed definition lift_pref_profile :: "'agent set \ 'alt set \ 'agent set \ 'alt set \ ('agent, 'alt) pref_profile \ ('agent, 'alt) pref_profile" where "lift_pref_profile agents alts agents' alts' R = (\i x y. x \ alts' \ y \ alts' \ i \ agents' \ (x = y \ x \ alts \ i \ agents \ (y \ alts \ R i x y)))" lemma lift_pref_profile_conv_vector: assumes "i \ agents" "i \ agents'" shows "lift_pref_profile agents alts agents' alts' R i = lift_prefs alts alts' (R i)" using assms by (auto simp: lift_pref_profile_def lift_prefs_def) lemma lift_pref_profile_wf: assumes "pref_profile_wf agents alts R" assumes "agents \ agents'" "alts \ alts'" "finite alts'" defines "R' \ lift_pref_profile agents alts agents' alts' R" shows "pref_profile_wf agents' alts' R'" proof - from assms interpret R: pref_profile_wf agents alts by simp have "finite_total_preorder_on alts' (R' i)" if i: "i \ agents'" for i proof (cases "i \ agents") case True then interpret finite_total_preorder_on alts "R i" by simp from True assms show ?thesis by unfold_locales (auto simp: lift_pref_profile_def dest: total intro: trans) next case False with assms i show ?thesis by unfold_locales (simp_all add: lift_pref_profile_def) qed moreover have "R' i = (\_ _. False)" if "i \ agents'" for i unfolding lift_pref_profile_def R'_def using that by simp ultimately show ?thesis unfolding pref_profile_wf_def using assms by auto qed lemma lift_pref_profile_permute_agents: assumes "\ permutes agents" "agents \ agents'" shows "lift_pref_profile agents alts agents' alts' (R \ \) = lift_pref_profile agents alts agents' alts' R \ \" using assms permutes_subset[OF assms] by (auto simp add: lift_pref_profile_def o_def permutes_in_image) lemma lift_pref_profile_permute_alts: assumes "\ permutes alts" "alts \ alts'" shows "lift_pref_profile agents alts agents' alts' (permute_profile \ R) = permute_profile \ (lift_pref_profile agents alts agents' alts' R)" proof - from assms have inv: "inv \ permutes alts" by (intro permutes_inv) from this assms(2) have "inv \ permutes alts'" by (rule permutes_subset) with inv show ?thesis using assms permutes_inj[OF \inv \ permutes alts\] by (fastforce simp add: lift_pref_profile_def permutes_in_image permute_profile_def fun_eq_iff dest: injD) qed context fixes agents alts R agents' alts' R' assumes R_wf: "pref_profile_wf agents alts R" assumes election: "agents \ agents'" "alts \ alts'" "alts \ {}" "agents \ {}" "finite alts'" defines "R' \ lift_pref_profile agents alts agents' alts' R" begin interpretation R: pref_profile_wf agents alts R by fact interpretation R': pref_profile_wf agents' alts' R' using election R_wf by (simp add: R'_def lift_pref_profile_wf) lemma lift_pref_profile_strict_iff: "x \[lift_pref_profile agents alts agents' alts' R i] y \ i \ agents \ ((y \ alts \ x \ alts' - alts) \ x \[R i] y)" proof (cases "i \ agents") case True then interpret total_preorder_on alts "R i" by simp show ?thesis using not_outside election by (auto simp: lift_pref_profile_def strongly_preferred_def) qed (simp_all add: lift_pref_profile_def strongly_preferred_def) lemma preferred_alts_lift_pref_profile: assumes i: "i \ agents'" and x: "x \ alts'" shows "preferred_alts (R' i) x = (if i \ agents \ x \ alts then preferred_alts (R i) x else alts')" proof (cases "i \ agents") assume i: "i \ agents" then interpret Ri: total_preorder_on alts "R i" by simp show ?thesis using i x election Ri.not_outside by (auto simp: preferred_alts_def R'_def lift_pref_profile_def Ri.refl) qed (auto simp: preferred_alts_def R'_def lift_pref_profile_def i x) lemma lift_pref_profile_Pareto_iff: "x \[Pareto(R')] y \ x \ alts' \ y \ alts' \ (x \ alts \ x \[Pareto(R)] y)" proof - from R.nonempty_agents obtain i where i: "i \ agents" by blast then interpret Ri: finite_total_preorder_on alts "R i" by simp show ?thesis unfolding R'.Pareto_iff R.Pareto_iff unfolding R'_def lift_pref_profile_def using election i by (auto simp: preorder_on.refl[OF R.in_dom] simp del: R.nonempty_alts R.nonempty_agents intro: Ri.not_outside) qed lemma lift_pref_profile_Pareto_strict_iff: "x \[Pareto(R')] y \ x \ alts' \ y \ alts' \ (x \ alts \ y \ alts \ x \[Pareto(R)] y)" by (auto simp: strongly_preferred_def lift_pref_profile_Pareto_iff R.Pareto.not_outside) lemma pareto_losers_lift_pref_profile: shows "pareto_losers R' = pareto_losers R \ (alts' - alts)" proof - have A: "x \ alts" "y \ alts" if "x \[Pareto(R)] y" for x y using that R.Pareto.not_outside unfolding strongly_preferred_def by auto have B: "x \ alts'" if "x \ alts" for x using election that by blast from R.nonempty_alts obtain x where x: "x \ alts" by blast thus ?thesis unfolding pareto_losers_def lift_pref_profile_Pareto_strict_iff [abs_def] by (auto dest: A B) qed end subsection \Lowering SCFs\ text \ Using the preference lifting, we can now \emph{lower} an SCF to a setting with fewer agents and alternatives under mild conditions to the original SCF. This preserves many nice properties, such as anonymity, neutrality, and strategyproofness. \ locale scf_lowering = pareto_efficient_scf agents alts scf for agents :: "'agent set" and alts :: "'alt set" and scf + fixes agents' alts' assumes agents'_subset: "agents' \ agents" and alts'_subset: "alts' \ alts" and agents'_nonempty [simp]: "agents' \ {}" and alts'_nonempty [simp]: "alts' \ {}" begin lemma finite_agents' [simp]: "finite agents'" using agents'_subset finite_agents by (rule finite_subset) lemma finite_alts' [simp]: "finite alts'" using alts'_subset finite_alts by (rule finite_subset) abbreviation lift :: "('agent, 'alt) pref_profile \ ('agent, 'alt) pref_profile" where "lift \ lift_pref_profile agents' alts' agents alts" definition lowered :: "('agent, 'alt) pref_profile \ 'alt set" where "lowered = scf \ lift" lemma lift_wf [simp, intro]: "pref_profile_wf agents' alts' R \ is_pref_profile (lift R)" using alts'_subset agents'_subset by (intro lift_pref_profile_wf) simp_all sublocale lowered: election agents' alts' by unfold_locales simp_all lemma preferred_alts_lift: "lowered.is_pref_profile R \ i \ agents \ x \ alts \ preferred_alts (lift R i) x = (if i \ agents' \ x \ alts' then preferred_alts (R i) x else alts)" using alts'_subset agents'_subset by (intro preferred_alts_lift_pref_profile) simp_all lemma pareto_losers_lift: "lowered.is_pref_profile R \ pareto_losers (lift R) = pareto_losers R \ (alts - alts')" using agents'_subset alts'_subset by (intro pareto_losers_lift_pref_profile) simp_all sublocale lowered: social_choice_function agents' alts' lowered proof fix R assume R_wf: "pref_profile_wf agents' alts' R" from R_wf have R'_wf: "pref_profile_wf agents alts (lift R)" by (rule lift_wf) show "lowered R \ alts'" proof safe fix x assume "x \ lowered R" hence "x \ scf (lift R)" by (auto simp: o_def lowered_def) moreover have "scf (lift R) \ pareto_losers (lift R) = {}" by (intro pareto_efficient R'_wf) ultimately show "x \ alts'" using scf_alts[of "lift R"] by (auto simp: pareto_losers_lift R_wf R'_wf) qed show "lowered R \ {}" using R'_wf by (auto simp: lowered_def scf_nonempty) qed sublocale lowered: pareto_efficient_scf agents' alts' lowered proof fix R assume R_wf: "pref_profile_wf agents' alts' R" from R_wf have R'_wf: "pref_profile_wf agents alts (lift R)" by (rule lift_wf) have "lowered R \ pareto_losers (lift R) = {}" unfolding lowered_def o_def by (intro pareto_efficient R'_wf) with R_wf show "lowered R \ pareto_losers R = {}" by (auto simp: pareto_losers_lift) qed end locale scf_lowering_anonymous = anonymous_scf agents alts scf + scf_lowering agents alts scf agents' alts' for agents :: "'agent set" and alts :: "'alt set" and scf agents' alts' begin sublocale lowered: anonymous_scf agents' alts' lowered proof fix \ R assume "\ permutes agents'" and "lowered.is_pref_profile R" thus "lowered (R \ \) = lowered R" using agents'_subset by (auto simp: lowered_def lift_pref_profile_permute_agents anonymous permutes_subset) qed end locale scf_lowering_neutral = neutral_scf agents alts scf + scf_lowering agents alts scf agents' alts' for agents :: "'agent set" and alts :: "'alt set" and scf agents' alts' begin sublocale lowered: neutral_scf agents' alts' lowered proof fix \ R assume "\ permutes alts'" and "lowered.is_pref_profile R" thus "lowered (permute_profile \ R) = \ ` lowered R" using alts'_subset by (auto simp: lowered_def lift_pref_profile_permute_alts neutral permutes_subset) qed end text \ The following is a technical condition that we need from a set extensions in order for strategyproofness to survive the lowering. The condition could probably be weakened a bit, but it is good enough for our purposes the way it is. \ locale liftable_set_extension = fixes alts' alts :: "'alt set" and set_ext :: "'alt relation \ 'alt set relation" assumes set_ext_strong_lift: "total_preorder_on alts' R \ A \ {} \ B \ {} \ A \ alts' \ B \ alts' \ A \[set_ext R] B \ A \[set_ext (lift_prefs alts' alts R)] B" lemma liftable_set_extensionI_weak: assumes "\R A B. total_preorder_on alts' R \ A \ {} \ B \ {} \ A \ alts' \ B \ alts' \ A \[set_ext R] B \ A \[set_ext (lift_prefs alts' alts R)] B" shows "liftable_set_extension alts' alts set_ext" proof (standard, goal_cases) case (1 R A B) from assms[of R A B] and assms[of R B A] and 1 show ?case by (auto simp: strongly_preferred_def) qed lemma Kelly_liftable: assumes "alts' \ alts" shows "liftable_set_extension alts' alts Kelly" proof (rule liftable_set_extensionI_weak, goal_cases) case (1 R A B) interpret R: total_preorder_on alts' R by fact from 1(2-5) show ?case using assms R.refl by (force simp: Kelly_def lift_prefs_def) qed lemma Fishburn_liftable: assumes "alts' \ alts" shows "liftable_set_extension alts' alts Fishb" proof (rule liftable_set_extensionI_weak, goal_cases) case (1 R A B) interpret R: total_preorder_on alts' R by fact have conj_cong: "P1 \ Q1 \ P2 \ Q2" if "P1 \ P2" "Q1 \ Q2" for P1 P2 Q1 Q2 using that by blast from 1(2-5) show ?case using assms unfolding Fishb_def lift_prefs_def by (intro conj_cong ball_cong refl) auto qed locale scf_lowering_strategyproof = strategyproof_scf agents alts scf set_ext + liftable_set_extension alts' alts set_ext + scf_lowering agents alts scf agents' alts' for agents :: "'agent set" and alts :: "'alt set" and scf agents' alts' set_ext begin sublocale lowered: strategyproof_scf agents' alts' lowered proof fix R Ri' i assume R_wf: "lowered.is_pref_profile R" and Ri'_wf: "total_preorder_on alts' Ri'" and i: "i \ agents'" interpret R: pref_profile_wf agents' alts' R by fact interpret Ri': total_preorder_on alts' Ri' by fact from R_wf have R'_wf: "is_pref_profile (lift R)" by (rule lift_wf) \ \We lift the alternative preference for the agent @{term i} in @{term R} to preferences in the lifted profile. \ define Ri'' where "Ri'' = lift_prefs alts' alts Ri'" have "\scf (lift R) \[set_ext (lift R i)] scf ((lift R)(i := Ri''))" using i agents'_subset alts'_subset unfolding Ri''_def by (intro strategyproof R'_wf Ri'_wf lift_prefs_wf) auto also have "(lift R)(i := Ri'') = lift (R(i := Ri'))" using i agents'_subset by (auto simp: fun_eq_iff Ri''_def lift_pref_profile_def lift_prefs_def) finally have not_less: "\scf (lift R) \[set_ext (lift R i)] scf (lift (R(i := Ri')))" . show "\lowered R \[set_ext (R i)] lowered (R(i := Ri'))" proof assume "lowered R \[set_ext (R i)] lowered (R(i := Ri'))" hence "lowered R \[set_ext (lift_prefs alts' alts (R i))] lowered (R(i := Ri'))" by (intro set_ext_strong_lift R.prefs_wf'(1) i lowered.scf_nonempty lowered.scf_alts R.wf_update R_wf Ri'_wf) also have "lift_prefs alts' alts (R i) = lift R i" using agents'_subset i by (subst lift_pref_profile_conv_vector) auto finally show False using not_less unfolding lowered_def o_def by contradiction qed qed end end \ No newline at end of file diff --git a/thys/Functional-Automata/MaxPrefix.thy b/thys/Functional-Automata/MaxPrefix.thy --- a/thys/Functional-Automata/MaxPrefix.thy +++ b/thys/Functional-Automata/MaxPrefix.thy @@ -1,92 +1,90 @@ (* Author: Tobias Nipkow Copyright 1998 TUM *) section "Maximal prefix" theory MaxPrefix imports "HOL-Library.Sublist" begin definition is_maxpref :: "('a list \ bool) \ 'a list \ 'a list \ bool" where "is_maxpref P xs ys = (prefix xs ys \ (xs=[] \ P xs) \ (\zs. prefix zs ys \ P zs \ prefix zs xs))" type_synonym 'a splitter = "'a list \ 'a list * 'a list" definition is_maxsplitter :: "('a list \ bool) \ 'a splitter \ bool" where "is_maxsplitter P f = (\xs ps qs. f xs = (ps,qs) = (xs=ps@qs \ is_maxpref P ps xs))" fun maxsplit :: "('a list \ bool) \ 'a list * 'a list \ 'a list \ 'a splitter" where "maxsplit P res ps [] = (if P ps then (ps,[]) else res)" | "maxsplit P res ps (q#qs) = maxsplit P (if P ps then (ps,q#qs) else res) (ps@[q]) qs" declare if_split[split del] lemma maxsplit_lemma: "(maxsplit P res ps qs = (xs,ys)) = (if \us. prefix us qs \ P(ps@us) then xs@ys=ps@qs \ is_maxpref P xs (ps@qs) else (xs,ys)=res)" proof (induction P res ps qs rule: maxsplit.induct) case 1 thus ?case by (auto simp: is_maxpref_def split: if_splits) next case (2 P res ps q qs) show ?case proof (cases "\us. prefix us qs \ P ((ps @ [q]) @ us)") - case True - note ex1 = this - then guess us by (elim exE conjE) note us = this + case ex1: True + then obtain us where "prefix us qs" "P ((ps @ [q]) @ us)" by blast hence ex2: "\us. prefix us (q # qs) \ P (ps @ us)" by (intro exI[of _ "q#us"]) auto with ex1 and 2 show ?thesis by simp next - case False - note ex1 = this + case ex1: False show ?thesis proof (cases "\us. prefix us (q#qs) \ P (ps @ us)") case False from 2 show ?thesis by (simp only: ex1 False) (insert ex1 False, auto simp: prefix_Cons) next case True note ex2 = this show ?thesis proof (cases "P ps") case True with 2 have "(maxsplit P (ps, q # qs) (ps @ [q]) qs = (xs, ys)) \ (xs = ps \ ys = q # qs)" by (simp only: ex1 ex2) simp_all also have "\ \ (xs @ ys = ps @ q # qs \ is_maxpref P xs (ps @ q # qs))" using ex1 True by (auto simp: is_maxpref_def prefix_append prefix_Cons append_eq_append_conv2) finally show ?thesis using True by (simp only: ex1 ex2) simp_all next case False with 2 have "(maxsplit P res (ps @ [q]) qs = (xs, ys)) \ ((xs, ys) = res)" by (simp only: ex1 ex2) simp also have "\ \ (xs @ ys = ps @ q # qs \ is_maxpref P xs (ps @ q # qs))" using ex1 ex2 False by (auto simp: append_eq_append_conv2 is_maxpref_def prefix_Cons) finally show ?thesis using False by (simp only: ex1 ex2) simp qed qed qed qed declare if_split[split] lemma is_maxpref_Nil[simp]: "\(\us. prefix us xs \ P us) \ is_maxpref P ps xs = (ps = [])" by (auto simp: is_maxpref_def) lemma is_maxsplitter_maxsplit: "is_maxsplitter P (\xs. maxsplit P ([],xs) [] xs)" by (auto simp: maxsplit_lemma is_maxsplitter_def) lemmas maxsplit_eq = is_maxsplitter_maxsplit[simplified is_maxsplitter_def] end diff --git a/thys/Goodstein_Lambda/Goodstein_Lambda.thy b/thys/Goodstein_Lambda/Goodstein_Lambda.thy --- a/thys/Goodstein_Lambda/Goodstein_Lambda.thy +++ b/thys/Goodstein_Lambda/Goodstein_Lambda.thy @@ -1,755 +1,756 @@ section \Specification\ theory Goodstein_Lambda imports Main begin subsection \Hereditary base representation\ text \We define a data type of trees and an evaluation function that sums siblings and exponentiates with respect to the given base on nesting.\ datatype C = C (unC: "C list") fun evalC where "evalC b (C []) = 0" | "evalC b (C (x # xs)) = b^evalC b x + evalC b (C xs)" value "evalC 2 (C [])" \ \$0$\ value "evalC 2 (C [C []])" \ \$2^0 = 1$\ value "evalC 2 (C [C [C []]])" \ \$2^1 = 2$\ value "evalC 2 (C [C [], C []])" \ \$2^0 + 2^0 = 2^0 \cdot 2 = 2$; not in hereditary base $2$\ text \The hereditary base representation is characterized as trees (i.e., nested lists) whose lists have monotonically increasing evaluations, with fewer than @{term "b"} repetitions for each value. We will show later that this representation is unique.\ inductive_set hbase for b where "C [] \ hbase b" | "i \ 0 \ i < b \ n \ hbase b \ C ms \ hbase b \ (\m'. m' \ set ms \ evalC b n < evalC b m') \ C (replicate i n @ ms) \ hbase b" text \We can convert to and from natural numbers as follows.\ definition H2N where "H2N b n = evalC b n" text \As we will show later, @{term "H2N b"} restricted to @{term "hbase n"} is bijective if @{prop "b \ (2 :: nat)"}, so we can convert from natural numbers by taking the inverse.\ definition N2H where "N2H b n = inv_into (hbase b) (H2N b) n" subsection \The Goodstein function\ text \We define a function that computes the length of the Goodstein sequence whose $c$-th element is $g_c = n$. Termination will be shown later, thereby establishing Goodstein's theorem.\ function (sequential) goodstein :: "nat \ nat \ nat" where "goodstein 0 n = 0" \ \we start counting at 1; also note that the initial base is @{term "c+1 :: nat"} and\ \ \hereditary base 1 makes no sense, so we have to avoid this case\ | "goodstein c 0 = c" | "goodstein c n = goodstein (c+1) (H2N (c+2) (N2H (c+1) n) - 1)" by pat_completeness auto abbreviation \ where "\ n \ goodstein (Suc 0) n" section \Ordinals\ text \The following type contains countable ordinals, by the usual case distinction into 0, successor ordinal, or limit ordinal; limit ordinals are given by their fundamental sequence. Hereditary base @{term "b"} representations carry over to such ordinals by replacing each occurrence of the base by @{term "\"}.\ datatype Ord = Z | S Ord | L "nat \ Ord" text \Note that the following arithmetic operations are not correct for all ordinals. However, they will only be used in cases where they actually correspond to the ordinal arithmetic operations.\ primrec addO where "addO n Z = n" | "addO n (S m) = S (addO n m)" | "addO n (L f) = L (\i. addO n (f i))" primrec mulO where "mulO n Z = Z" | "mulO n (S m) = addO (mulO n m) n" | "mulO n (L f) = L (\i. mulO n (f i))" definition \ where "\ = L (\n. (S ^^ n) Z)" primrec exp\ where "exp\ Z = S Z" | "exp\ (S n) = mulO (exp\ n) \" | "exp\ (L f) = L (\i. exp\ (f i))" subsection \Evaluation\ text \Evaluating an ordinal number at base $b$ is accomplished by taking the $b$-th element of all fundamental sequences and interpreting zero and successor over the natural numbers.\ primrec evalO where "evalO b Z = 0" | "evalO b (S n) = Suc (evalO b n)" | "evalO b (L f) = evalO b (f b)" subsection \Goodstein function and sequence\ text \We can define the Goodstein function very easily, but proving correctness will take a while.\ primrec goodsteinO where "goodsteinO c Z = c" | "goodsteinO c (S n) = goodsteinO (c+1) n" | "goodsteinO c (L f) = goodsteinO c (f (c+2))" primrec stepO where "stepO c Z = Z" | "stepO c (S n) = n" | "stepO c (L f) = stepO c (f (c+2))" text \We can compute a few values of the Goodstein sequence starting at $4$.\ definition g4O where "g4O n = fold stepO [1.. ^^ 3) Z)" value "map (\n. evalO (n+2) (g4O n)) [0..<10]" \ \@{value "[4, 26, 41, 60, 83, 109, 139, 173, 211, 253] :: nat list"}\ subsection \Properties of evaluation\ lemma evalO_addO [simp]: "evalO b (addO n m) = evalO b n + evalO b m" by (induct m) auto lemma evalO_mulO [simp]: "evalO b (mulO n m) = evalO b n * evalO b m" by (induct m) auto lemma evalO_n [simp]: "evalO b ((S ^^ n) Z) = n" by (induct n) auto lemma evalO_\ [simp]: "evalO b \ = b" by (auto simp: \_def) lemma evalO_exp\ [simp]: "evalO b (exp\ n) = b^(evalO b n)" by (induct n) auto text \Note that evaluation is useful for proving that @{type "Ord"} values are distinct:\ notepad begin have "addO n (exp\ m) \ n" for n m by (auto dest: arg_cong[of _ _ "evalO 1"]) end subsection \Arithmetic properties\ lemma addO_Z [simp]: "addO Z n = n" by (induct n) auto lemma addO_assoc [simp]: "addO n (addO m p) = addO (addO n m) p" by (induct p) auto lemma mul0_distrib [simp]: "mulO n (addO p q) = addO (mulO n p) (mulO n q)" by (induct q) auto lemma mulO_assoc [simp]: "mulO n (mulO m p) = mulO (mulO n m) p" by (induct p) auto lemma exp\_addO [simp]: "exp\ (addO n m) = mulO (exp\ n) (exp\ m)" by (induct m) auto section \Cantor normal form\ text \The previously introduced tree type @{type C} can be used to represent Cantor normal forms; they are trees (evaluated at base @{term \}) such that siblings are in non-decreasing order. One can think of this as hereditary base @{term \}. The plan is to mirror selected operations on ordinals in Cantor normal forms.\ subsection \Conversion to and from the ordinal type @{type Ord}\ fun C2O where "C2O (C []) = Z" | "C2O (C (n # ns)) = addO (C2O (C ns)) (exp\ (C2O n))" definition O2C where "O2C = inv C2O" text \We show that @{term C2O} is injective, meaning the inverse is unique.\ lemma addO_exp\_inj: assumes "addO n (exp\ m) = addO n' (exp\ m')" shows "n = n'" and "m = m'" proof - have "addO n (exp\ m) = addO n' (exp\ m') \ n = n'" by (induct m arbitrary: m'; case_tac m'; force simp: \_def dest!: fun_cong[of _ _ 1]) moreover have "addO n (exp\ m) = addO n (exp\ m') \ m = m'" apply (induct m arbitrary: n m'; case_tac m') apply (auto 0 3 simp: \_def intro: rangeI dest: arg_cong[of _ _ "evalO 1"] fun_cong[of _ _ 0] fun_cong[of _ _ 1])[8] (* 1 left *) by simp (meson ext rangeI) ultimately show "n = n'" and "m = m'" using assms by simp_all qed lemma C2O_inj: "C2O n = C2O m \ n = m" by (induct n arbitrary: m rule: C2O.induct; case_tac m rule: C2O.cases) (auto dest: addO_exp\_inj arg_cong[of _ _ "evalO 1"]) lemma O2C_C2O [simp]: "O2C (C2O n) = n" by (auto intro!: inv_f_f simp: O2C_def inj_def C2O_inj) lemma O2C_Z [simp]: "O2C Z = C []" using O2C_C2O[of "C []", unfolded C2O.simps] . lemma C2O_replicate: "C2O (C (replicate i n)) = mulO (exp\ (C2O n)) ((S ^^ i) Z)" by (induct i) auto lemma C2O_app: "C2O (C (xs @ ys)) = addO (C2O (C ys)) (C2O (C xs))" by (induct xs arbitrary: ys) auto subsection \Evaluation\ lemma evalC_def': "evalC b n = evalO b (C2O n)" by (induct n rule: C2O.induct) auto lemma evalC_app [simp]: "evalC b (C (ns @ ms)) = evalC b (C ns) + evalC b (C ms)" by (induct ns) auto lemma evalC_replicate [simp]: "evalC b (C (replicate c n)) = c * evalC b (C [n])" by (induct c) auto subsection \Transfer of the @{type Ord} induction principle to @{type C}\ fun funC where \ \@{term funC} computes the fundamental sequence on @{type C}\ "funC (C []) = (\i. [C []])" | "funC (C (C [] # ns)) = (\i. replicate i (C ns))" | "funC (C (n # ns)) = (\i. [C (funC n i @ ns)])" lemma C2O_cons: "C2O (C (n # ns)) = (if n = C [] then S (C2O (C ns)) else L (\i. C2O (C (funC n i @ ns))))" by (induct n arbitrary: ns rule: funC.induct) (simp_all add: \_def C2O_replicate C2O_app flip: exp\_addO) lemma C_Ord_induct: assumes "P (C [])" and "\ns. P (C ns) \ P (C (C [] # ns))" and "\n ns ms. (\i. P (C (funC (C (n # ns)) i @ ms))) \ P (C (C (n # ns) # ms))" shows "P n" proof - have "\n. C2O n = m \ P n" for m by (induct m; intro allI; case_tac n rule: funC.cases) (auto simp: C2O_cons simp del: C2O.simps(2) intro: assms) then show ?thesis by simp qed subsection \Goodstein function and sequence on @{type C}\ function (domintros) goodsteinC where "goodsteinC c (C []) = c" | "goodsteinC c (C (C [] # ns)) = goodsteinC (c+1) (C ns)" | "goodsteinC c (C (C (n # ns) # ms)) = goodsteinC c (C (funC (C (n # ns)) (c+2) @ ms))" by pat_completeness auto termination proof - have "goodsteinC_dom (c, n)" for c n by (induct n arbitrary: c rule: C_Ord_induct) (auto intro: goodsteinC.domintros) then show ?thesis by simp qed lemma goodsteinC_def': "goodsteinC c n = goodsteinO c (C2O n)" by (induct c n rule: goodsteinC.induct) (simp_all add: C2O_cons del: C2O.simps(2)) function (domintros) stepC where "stepC c (C []) = C []" | "stepC c (C (C [] # ns)) = C ns" | "stepC c (C (C (n # ns) # ms)) = stepC c (C (funC (C (n # ns)) (Suc (Suc c)) @ ms))" by pat_completeness auto termination proof - have "stepC_dom (c, n)" for c n by (induct n arbitrary: c rule: C_Ord_induct) (auto intro: stepC.domintros) then show ?thesis by simp qed definition g4C where "g4C n = fold stepC [1..n. evalC (n+2) (g4C n)) [0..<10]" \ \@{value "[4, 26, 41, 60, 83, 109, 139, 173, 211, 253] :: nat list"}\ subsection \Properties\ lemma stepC_def': "stepC c n = O2C (stepO c (C2O n))" by (induct c n rule: stepC.induct) (simp_all add: C2O_cons del: C2O.simps(2)) lemma funC_ne [simp]: "funC m (Suc n) \ []" by (cases m rule: funC.cases) simp_all lemma evalC_funC [simp]: "evalC b (C (funC n b)) = evalC b (C [n])" by (induct n rule: funC.induct) simp_all lemma stepC_app [simp]: "n \ C [] \ stepC c (C (unC n @ ns)) = C (unC (stepC c n) @ ns)" by (induct n arbitrary: ns rule: stepC.induct) simp_all lemma stepC_cons [simp]: "ns \ [] \ stepC c (C (n # ns)) = C (unC (stepC c (C [n])) @ ns)" using stepC_app[of "C[n]" c ns] by simp lemma stepC_dec: "n \ C [] \ Suc (evalC (Suc (Suc c)) (stepC c n)) = evalC (Suc (Suc c)) n" by (induct c n rule: stepC.induct) simp_all lemma stepC_dec': "n \ C [] \ evalC (c+3) (stepC c n) < evalC (c+3) n" proof (induct c n rule: stepC.induct) case (3 c n ns ms) have "evalC (c+3) (C (funC (C (n # ns)) (Suc (Suc c)))) \ (c+3) ^ ((c+3) ^ evalC (c+3) n + evalC (c+3) (C ns))" by (induct n rule: funC.induct) (simp_all add: distrib_right) then show ?case using 3 by simp qed simp_all section \Hereditary base @{term b} representation\ text \We now turn to properties of the @{term "hbase b"} subset of trees.\ subsection \Uniqueness\ text \We show uniqueness of the hereditary base representation by showing that @{term "evalC b"} restricted to @{term "hbase b"} is injective.\ lemma hbaseI2: "i < b \ n \ hbase b \ C m \ hbase b \ (\m'. m' \ set m \ evalC b n < evalC b m') \ C (replicate i n @ m) \ hbase b" by (cases i) (auto intro: hbase.intros simp del: replicate.simps(2)) lemmas hbase_singletonI = hbase.intros(2)[of 1 "Suc (Suc b)" for b, OF _ _ _ hbase.intros(1), simplified] lemma hbase_hd: "C ns \ hbase b \ ns \ [] \ hd ns \ hbase b" by (cases rule: hbase.cases) auto lemmas hbase_hd' [dest] = hbase_hd[of "n # ns" for n ns, simplified] lemma hbase_tl: "C ns \ hbase b \ ns \ [] \ C (tl ns) \ hbase b" by (cases "C ns" b rule: hbase.cases) (auto intro: hbaseI2) lemmas hbase_tl' [dest] = hbase_tl[of "n # ns" for n ns, simplified] lemma hbase_elt [dest]: "C ns \ hbase b \ n \ set ns \ n \ hbase b" by (induct ns) auto lemma evalC_sum_list: "evalC b (C ns) = sum_list (map (\n. b^evalC b n) ns)" by (induct ns) auto lemma sum_list_replicate: "sum_list (replicate n x) = n * x" by (induct n) auto lemma base_red: fixes b :: nat assumes n: "\n'. n' \ set ns \ n < n'" "i < b" "i \ 0" and m: "\m'. m' \ set ms \ m < m'" "j < b" "j \ 0" and s: "i * b^n + sum_list (map (\n. b^n) ns) = j * b^m + sum_list (map (\n. b^n) ms)" shows "i = j \ n = m" using n(1) m(1) s proof (induct n arbitrary: m ns ms) { fix ns ms :: "nat list" and i j m :: nat assume n': "\n'. n' \ set ns \ 0 < n'" "i < b" "i \ 0" assume m': "\m'. m' \ set ms \ m < m'" "j < b" "j \ 0" assume s': "i * b^0 + sum_list (map (\n. b^n) ns) = j * b^m + sum_list (map (\n. b^n) ms)" obtain x where [simp]: "sum_list (map ((^) b) ns) = x*b" using n'(1) by (intro that[of "sum_list (map (\n. b^(n-1)) ns)"]) (simp add: ac_simps flip: sum_list_const_mult power_Suc cong: map_cong) obtain y where [simp]: "sum_list (map ((^) b) ms) = y*b" using order.strict_trans1[OF le0 m'(1)] by (intro that[of "sum_list (map (\n. b^(n-1)) ms)"]) (simp add: ac_simps flip: sum_list_const_mult power_Suc cong: map_cong) have [simp]: "m = 0" using s' n'(2,3) by (cases m, simp_all) (metis Groups.mult_ac(2) Groups.mult_ac(3) Suc_pred div_less mod_div_mult_eq mod_mult_self2 mod_mult_self2_is_0 mult_zero_right nat.simps(3)) have "i = j \ 0 = m" using s' n'(2,3) m'(2,3) by simp (metis div_less mod_div_mult_eq mod_mult_self1) } note BASE = this { case 0 show ?case by (rule BASE; fact) next case (Suc n m') have "j = i \ 0 = Suc n" if "m' = 0" using Suc(2-4) by (intro BASE[of ms j ns "Suc n" i]) (simp_all add: ac_simps that n(2,3) m(2,3)) then obtain m where m' [simp]: "m' = Suc m" by (cases m') auto obtain ns' where [simp]: "ns = map Suc ns'" "\n'. n' \ set ns' \ n < n'" using Suc(2) less_trans[OF zero_less_Suc Suc(2)] by (intro that[of "map (\n. n-1) ns"]; force cong: map_cong) obtain ms' where [simp]: "ms = map Suc ms'" "\m'. m' \ set ms' \ m < m'" using Suc(3)[unfolded m'] less_trans[OF zero_less_Suc Suc(3)[unfolded m']] by (intro that[of "map (\n. n-1) ms"]; force cong: map_cong) have *: "b * x = b * y \ x = y" for x y using n(2) by simp have "i = j \ n = m" proof (rule Suc(1)[of "map (\n. n-1) ns" "map (\n. n-1) ms" m, OF _ _ *], goal_cases) case 3 show ?case using Suc(4) unfolding add_mult_distrib2 by (simp add: comp_def ac_simps flip: sum_list_const_mult) qed simp_all then show ?case by simp } qed lemma evalC_inj_on_hbase: "n \ hbase b \ m \ hbase b \ evalC b n = evalC b m \ n = m" proof (induct n arbitrary: m rule: hbase.induct) case 1 then show ?case by (cases m rule: hbase.cases) simp_all next case (2 i n ns m') obtain j m ms where [simp]: "m' = C (replicate j m @ ms)" and m: "j \ 0" "j < b" "m \ hbase b" "C ms \ hbase b" "\m'. m' \ set ms \ evalC b m < evalC b m'" using 2(8,1,2,9) by (cases m' rule: hbase.cases) simp_all have "i = j \ evalC b n = evalC b m" using 2(1,2,7,9) m(1,2,5) by (intro base_red[of "map (evalC b) ns" _ _ b "map (evalC b) ms"]) (auto simp: comp_def evalC_sum_list sum_list_replicate) then show ?case using 2(4)[OF m(3)] 2(6)[OF m(4)] 2(9) by simp qed subsection \Correctness of @{const stepC}\ text \We show that @{term "stepC c"} preserves hereditary base @{term "c + 2 :: nat"} representations. In order to cover intermediate results produced by @{const stepC}, we extend the hereditary base representation to allow the least significant digit to be equal to @{term b}, which essentially means that we may have an extra sibling in front on every level.\ inductive_set hbase_ext for b where "n \ hbase b \ n \ hbase_ext b" | "n \ hbase_ext b \ C m \ hbase b \ (\m'. m' \ set m \ evalC b n \ evalC b m') \ C (n # m) \ hbase_ext b" lemma hbase_ext_hd' [dest]: "C (n # ns) \ hbase_ext b \ n \ hbase_ext b" by (cases rule: hbase_ext.cases) (auto intro: hbase_ext.intros(1)) lemma hbase_ext_tl: "C ns \ hbase_ext b \ ns \ [] \ C (tl ns) \ hbase b" by (cases "C ns" b rule: hbase_ext.cases; cases ns) (simp_all add: hbase_tl') lemmas hbase_ext_tl' [dest] = hbase_ext_tl[of "n # ns" for n ns, simplified] lemma hbase_funC: "c \ 0 \ C (n # ns) \ hbase_ext (Suc c) \ C (funC n (Suc c) @ ns) \ hbase_ext (Suc c)" proof (induct n arbitrary: ns rule: funC.induct) case (2 ms) have [simp]: "evalC (Suc c) (C ms) < evalC (Suc c) m'" if "m' \ set ns" for m' using 2(2) proof (cases rule: hbase_ext.cases) case 1 then show ?thesis using that by (cases rule: hbase.cases, case_tac i) (auto intro: Suc_lessD) qed (auto simp: Suc_le_eq that) show ?case using 2 by (auto 0 4 intro: hbase_ext.intros hbase.intros(2) order.strict_implies_order) next case (3 m ms ms') show ?case unfolding funC.simps append_Cons append_Nil proof (rule hbase_ext.intros(2), goal_cases 31 32 33) case (33 m') show ?case using 3(3) proof (cases rule: hbase_ext.cases) case 1 show ?thesis using 1 3(1,2) 33 by (cases rule: hbase.cases, case_tac i) (auto intro: less_or_eq_imp_le) qed (insert 33, simp) qed (insert 3, blast+) qed auto lemma stepC_sound: "n \ hbase_ext (Suc (Suc c)) \ stepC c n \ hbase (Suc (Suc c))" proof (induct c n rule: stepC.induct) case (3 c n ns ms) show ?case using 3(2,1) by (cases rule: hbase_ext.cases; unfold stepC.simps) (auto intro: hbase_funC) qed (auto intro: hbase.intros) subsection \Surjectivity of @{const evalC}\ text \Note that the base must be at least @{term "2 :: nat"}.\ lemma evalC_surjective: "\n' \ hbase (Suc (Suc b)). evalC (Suc (Suc b)) n' = n" proof (induct n) case 0 then show ?case by (auto intro: bexI[of _ "C []"] hbase.intros) next have [simp]: "Suc x \ Suc (Suc b)^x" for x by (induct x) auto case (Suc n) - then guess n' by (rule bexE) + then obtain n' where "n' \ hbase (Suc (Suc b))" "evalC (Suc (Suc b)) n' = n" by blast then obtain n' j where n': "Suc n \ j" "j = evalC (Suc (Suc b)) n'" "n' \ hbase (Suc (Suc b))" by (intro that[of _ "C [n']"]) (auto intro!: intro: hbase.intros(1) dest!: hbaseI2[of 1 "b+2" n' "[]", simplified]) then show ?case proof (induct rule: inc_induct) case (step m) - guess n' using step(3)[OF step(4,5)] by (rule bexE) + obtain n' where "n' \ hbase (Suc (Suc b))" "evalC (Suc (Suc b)) n' = Suc m" + using step(3)[OF step(4,5)] by blast then show ?case using stepC_dec[of n' "b"] by (cases n' rule: C2O.cases) (auto intro: stepC_sound hbase_ext.intros(1)) qed blast qed subsection \Monotonicity of @{const hbase}\ text \Here we show that every hereditary base @{term "b :: nat"} number is also a valid hereditary base @{term "b+1 :: nat"} number. This is not immediate because we have to show that monotonicity of siblings is preserved.\ lemma hbase_evalC_mono: assumes "n \ hbase b" "m \ hbase b" "evalC b n < evalC b m" shows "evalC (Suc b) n < evalC (Suc b) m" proof (cases "b < 2") case True show ?thesis using assms(2,3) True by (cases rule: hbase.cases) simp_all next case False then obtain b' where [simp]: "b = Suc (Suc b')" by (auto simp: numeral_2_eq_2 not_less_eq dest: less_imp_Suc_add) show ?thesis using assms(3,1,2) proof (induct "evalC b n" "evalC b m" arbitrary: n m rule: less_Suc_induct) case 1 then show ?case using stepC_sound[of m b', OF hbase_ext.intros(1)] stepC_dec[of m b'] stepC_dec'[of m b'] evalC_inj_on_hbase by (cases m rule: C2O.cases) (fastforce simp: eval_nat_numeral)+ next case (2 j) then show ?case using evalC_surjective[of b' j] less_trans by fastforce qed qed lemma hbase_mono: "n \ hbase b \ n \ hbase (Suc b)" by (induct n rule: hbase.induct) (auto 0 3 intro: hbase.intros hbase_evalC_mono) subsection \Conversion to and from @{type nat}\ text \We have previously defined @{term "H2N b = evalC b"} and @{term "N2H b"} as its inverse. So we can use the injectivity and surjectivity of @{term "evalC b"} for simplification.\ lemma N2H_inv: "n \ hbase b \ N2H b (H2N b n) = n" using evalC_inj_on_hbase by (auto simp: N2H_def H2N_def[abs_def] inj_on_def intro!: inv_into_f_f) lemma H2N_inv: "H2N (Suc (Suc b)) (N2H (Suc (Suc b)) n) = n" using evalC_surjective[of "b" n] by (auto simp: N2H_def H2N_def[abs_def] intro: f_inv_into_f) lemma N2H_eqI: "n \ hbase (Suc (Suc b)) \ H2N (Suc (Suc b)) n = m \ N2H (Suc (Suc b)) m = n" using N2H_inv by blast lemma N2H_neI: "n \ hbase (Suc (Suc b)) \ H2N (Suc (Suc b)) n \ m \ N2H (Suc (Suc b)) m \ n" using H2N_inv by blast lemma N2H_0 [simp]: "N2H (Suc (Suc c)) 0 = C []" using H2N_def N2H_inv hbase.intros(1) by fastforce lemma N2H_nz [simp]: "0 < n \ N2H (Suc (Suc c)) n \ C []" by (metis N2H_0 H2N_inv neq0_conv) section \The Goodstein function revisited\ text \We are now ready to prove termination of the Goodstein function @{const goodstein} as well as its relation to @{const goodsteinC} and @{const goodsteinO}.\ lemma goodstein_aux: "goodsteinC (Suc c) (N2H (Suc (Suc c)) (Suc n)) = goodsteinC (c+2) (N2H (c+3) (H2N (c+3) (N2H (c+2) (n+1)) - 1))" proof - have [simp]: "n \ C [] \ goodsteinC c n = goodsteinC (c+1) (stepC c n)" for c n by (induct c n rule: stepC.induct) simp_all have [simp]: "stepC (Suc c) (N2H (Suc (Suc c)) (Suc n)) \ hbase (Suc (Suc (Suc c)))" by (metis H2N_def N2H_inv evalC_surjective hbase_ext.intros(1) hbase_mono stepC_sound) show ?thesis using arg_cong[OF stepC_dec[of "N2H (c+2) (n+1)" "c+1", folded H2N_def], of "\n. N2H (c+3) (n-1)"] by (simp add: eval_nat_numeral N2H_inv) qed termination goodstein proof (relation "measure (\(c, n). goodsteinC c (N2H (c+1) n) - c)", goal_cases _ 1) case (1 c n) have *: "goodsteinC c n \ c" for c n by (induct c n rule: goodsteinC.induct) simp_all show ?case by (simp add: goodstein_aux eval_nat_numeral) (meson Suc_le_eq diff_less_mono2 lessI *) qed simp lemma goodstein_def': "c \ 0 \ goodstein c n = goodsteinC c (N2H (c+1) n)" by (induct c n rule: goodstein.induct) (simp_all add: goodstein_aux eval_nat_numeral) lemma goodstein_impl: "c \ 0 \ goodstein c n = goodsteinO c (C2O (N2H (c+1) n))" \ \but note that @{term N2H} is not executable as currently defined\ using goodstein_def'[unfolded goodsteinC_def'] . lemma goodstein_16: "\ 16 = goodsteinO 1 (exp\ (exp\ (exp\ (exp\ Z))))" proof - have "N2H (Suc (Suc 0)) 16 = C [C [C [C [C []]]]]" by (auto simp: H2N_def intro!: N2H_eqI hbase_singletonI hbase.intros(1)) then show ?thesis by (simp add: goodstein_impl) qed section \Translation to $\lambda$-calculus\ text \We define Church encodings for @{type nat} and @{type Ord}. Note that we are basically in a Hindley-Milner type system, so we cannot use a proper polymorphic type. We can still express Church encodings as folds over values of the original type.\ abbreviation Z\<^sub>N where "Z\<^sub>N \ (\s z. z)" abbreviation S\<^sub>N where "S\<^sub>N \ (\n s z. s (n s z))" primrec fold_nat ("\_\\<^sub>N") where "\0\\<^sub>N = Z\<^sub>N" | "\Suc n\\<^sub>N = S\<^sub>N \n\\<^sub>N" lemma one\<^sub>N: "\1\\<^sub>N = (\x. x)" by simp abbreviation Z\<^sub>O where "Z\<^sub>O \ (\z s l. z)" abbreviation S\<^sub>O where "S\<^sub>O \ (\n z s l. s (n z s l))" abbreviation L\<^sub>O where "L\<^sub>O \ (\f z s l. l (\i. f i z s l))" primrec fold_Ord ("\_\\<^sub>O") where "\Z\\<^sub>O = Z\<^sub>O" | "\S n\\<^sub>O = S\<^sub>O \n\\<^sub>O" | "\L f\\<^sub>O = L\<^sub>O (\i. \f i\\<^sub>O)" text \The following abbreviations and lemmas show how to implement the arithmetic functions and the Goodstein function on a Church-encoded @{type Ord} in lambda calculus.\ abbreviation (input) add\<^sub>O where "add\<^sub>O n m \ (\z s l. m (n z s l) s l)" lemma add\<^sub>O: "\addO n m\\<^sub>O = add\<^sub>O \n\\<^sub>O \m\\<^sub>O" by (induct m) simp_all abbreviation (input) mul\<^sub>O where "mul\<^sub>O n m \ (\z s l. m z (\m. n m s l) l)" lemma mul\<^sub>O: "\mulO n m\\<^sub>O = mul\<^sub>O \n\\<^sub>O \m\\<^sub>O" by (induct m) (simp_all add: add\<^sub>O) abbreviation (input) \\<^sub>O where "\\<^sub>O \ (\z s l. l (\n. \n\\<^sub>N s z))" lemma \\<^sub>O: "\\\\<^sub>O = \\<^sub>O" proof - have [simp]: "\(S ^^ i) Z\\<^sub>O z s l = \i\\<^sub>N s z" for i z s l by (induct i) simp_all show ?thesis by (simp add: \_def) qed abbreviation (input) exp\\<^sub>O where "exp\\<^sub>O n \ (\z s l. n s (\x z. l (\n. \n\\<^sub>N x z)) (\f z. l (\n. f n z)) z)" lemma exp\\<^sub>O: "\exp\ n\\<^sub>O = exp\\<^sub>O \n\\<^sub>O" by (induct n) (simp_all add: mul\<^sub>O \\<^sub>O) abbreviation (input) goodstein\<^sub>O where "goodstein\<^sub>O \ (\c n. n (\x. x) (\n m. n (m + 1)) (\f m. f (m + 2) m) c)" lemma goodstein\<^sub>O: "goodsteinO c n = goodstein\<^sub>O c \n\\<^sub>O" by (induct n arbitrary: c) simp_all text \Note that modeling Church encodings with folds is still limited. For example, the meaningful expression @{text "\n\\<^sub>N exp\\<^sub>O Z\<^sub>O"} cannot be typed in Isabelle/HOL, as that would require rank-2 polymorphism.\ subsection \Alternative: free theorems\ text \The following is essentially the free theorem for Church-encoded @{type Ord} values.\ lemma freeOrd: assumes "\n. h (s n) = s' (h n)" and "\f. h (l f) = l' (\i. h (f i))" shows "h (\n\\<^sub>O z s l) = \n\\<^sub>O (h z) s' l'" by (induct n) (simp_all add: assms) text \Each of the following proofs first states a naive definition of the corresponding function (which is proved correct by induction), from which we then derive the optimized version using the free theorem, by (conditional) rewriting (without induction).\ lemma add\<^sub>O': "\addO n m\\<^sub>O = add\<^sub>O \n\\<^sub>O \m\\<^sub>O" proof - have [simp]: "\addO n m\\<^sub>O = \m\\<^sub>O \n\\<^sub>O S\<^sub>O L\<^sub>O" by (induct m) simp_all show ?thesis by (intro ext) (simp add: freeOrd[where h = "\n. n _ _ _"]) qed lemma mul\<^sub>O': "\mulO n m\\<^sub>O = mul\<^sub>O \n\\<^sub>O \m\\<^sub>O" proof - have [simp]: "\mulO n m\\<^sub>O = \m\\<^sub>O Z\<^sub>O (\m. add\<^sub>O m \n\\<^sub>O) L\<^sub>O" by (induct m) (simp_all add: add\<^sub>O) show ?thesis by (intro ext) (simp add: freeOrd[where h = "\n. n _ _ _"]) qed lemma exp\\<^sub>O': "\exp\ n\\<^sub>O = exp\\<^sub>O \n\\<^sub>O" proof - have [simp]: "\exp\ n\\<^sub>O = \n\\<^sub>O (S\<^sub>O Z\<^sub>O) (\m. mul\<^sub>O m \\<^sub>O) L\<^sub>O" by (induct n) (simp_all add: mul\<^sub>O \\<^sub>O) show ?thesis by (intro ext) (simp add: fun_cong[OF freeOrd[where h = "\n z. n z _ _"]]) qed end \ No newline at end of file diff --git a/thys/Hoare_Time/SepLogK_VCG.thy b/thys/Hoare_Time/SepLogK_VCG.thy --- a/thys/Hoare_Time/SepLogK_VCG.thy +++ b/thys/Hoare_Time/SepLogK_VCG.thy @@ -1,403 +1,407 @@ theory SepLogK_VCG imports SepLogK_Hoare begin lemmas conseqS = conseq[where k=1, simplified] datatype acom = Askip ("SKIP") | Aassign vname aexp ("(_ ::= _)" [1000, 61] 61) | Aseq acom acom ("_;;/ _" [60, 61] 60) | Aif bexp acom acom ("(IF _/ THEN _/ ELSE _)" [0, 0, 61] 61) | Awhile assn2 bexp acom ("({_}/ WHILE _/ DO _)" [0, 0, 61] 61) notation com.SKIP ("SKIP") fun strip :: "acom \ com" where "strip SKIP = SKIP" | "strip (x ::= a) = (x ::= a)" | "strip (C\<^sub>1;; C\<^sub>2) = (strip C\<^sub>1;; strip C\<^sub>2)" | "strip (IF b THEN C\<^sub>1 ELSE C\<^sub>2) = (IF b THEN strip C\<^sub>1 ELSE strip C\<^sub>2)" | "strip ({_} WHILE b DO C) = (WHILE b DO strip C)" fun pre :: "acom \ assn2 \ assn2" where "pre SKIP Q = ($1 ** Q)" | "pre (x ::= a) Q = ((\(ps,t). x\dom ps \ vars a \ dom ps \ Q (ps(x\(paval a ps)),t) ) ** $1)" | "pre (C\<^sub>1;; C\<^sub>2) Q = pre C\<^sub>1 (pre C\<^sub>2 Q)" | "pre (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q = ( $1 ** (\(ps,n). vars b \ dom ps \ (if pbval b ps then pre C\<^sub>1 Q (ps,n) else pre C\<^sub>2 Q (ps,n) )))" | "pre ({I} WHILE b DO C) Q = (I ** $1)" fun vc :: "acom \ assn2 \ bool" where "vc SKIP Q = True" | "vc (x ::= a) Q = True" | "vc (C\<^sub>1 ;; C\<^sub>2) Q = ((vc C\<^sub>1 (pre C\<^sub>2 Q)) \ (vc C\<^sub>2 Q) )" | "vc (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q = (vc C\<^sub>1 Q \ vc C\<^sub>2 Q)" | "vc ({I} WHILE b DO C) Q = ( (\s. (I s \ vars b \ dom (fst s) ) \ ((\(s,n). I (s,n) \ lmaps_to_axpr b True s) s \ pre C (I ** $ 1) s) \ ( (\(s,n). I (s,n) \ lmaps_to_axpr b False s) s \ Q s)) \ vc C (I ** $ 1))" lemma dollar0_left: "($ 0 \* Q) = Q" apply rule unfolding dollar_def sep_conj_def by force lemma vc_sound: "vc C Q \ \\<^sub>3\<^sub>' {pre C Q} strip C { Q }" proof (induct C arbitrary: Q) case Askip then show ?case apply simp apply(rule conseqS[OF Frame[OF Skip]]) by (auto simp: dollar0_left) next case (Aassign x1 x2) then show ?case apply simp apply(rule conseqS) apply(rule Assign4) apply auto done next case (Aseq C1 C2) then show ?case apply (auto intro: Seq) done next case (Aif b C1 C2) then have Aif1: "\\<^sub>3\<^sub>' {pre C1 Q} strip C1 {Q}" and Aif2: "\\<^sub>3\<^sub>' {pre C2 Q} strip C2 {Q}" by auto show ?case apply simp apply (rule conseqS) apply(rule If[where P="%(ps,n). (if pbval b ps then pre C1 Q (ps,n) else pre C2 Q (ps,n))" and Q="Q"]) subgoal apply simp apply (rule conseqS) apply(fact Aif1) by auto subgoal apply simp apply (rule conseqS) apply(fact Aif2) by auto apply (auto simp: sep_conj_ac) unfolding sep_conj_def by blast next case (Awhile I b C) then have dom : "\s. (I s \ vars b \ dom (fst s) )" and i: "\s. (\(s,n). I (s,n) \ lmaps_to_axpr b True s) s \ pre C (I ** $ 1) s" and ii: "\s. (\(s,n). I (s,n) \ lmaps_to_axpr b False s) s \ Q s" and C: "\\<^sub>3\<^sub>' {pre C (I ** $ 1)} strip C {I ** $ 1}" by fastforce+ show ?case apply simp apply(rule conseqS) apply(rule While[where P=I]) apply(rule conseqS) apply(rule C) subgoal using i by auto subgoal apply simp using dom unfolding sep_conj_def by force subgoal apply simp using dom unfolding sep_conj_def by force subgoal using ii apply auto done done qed lemma vc2valid: "vc C Q \ \s. P s \ pre C Q s \ \\<^sub>3\<^sub>' {P} strip C { Q}" using hoareT_sound2_part weakenpre vc_sound by metis lemma pre_mono: assumes "\s. P s \ Q s" shows "\s. pre C P s \ pre C Q s" using assms proof(induct C arbitrary: P Q) case Askip then show ?case apply (auto simp: sep_conj_def dollar_def) by force next case (Aassign x1 x2) then show ?case by (auto simp: sep_conj_def dollar_def) next case (Aseq C1 C2) then show ?case by auto next case (Aif b C1 C2) then show ?case apply (auto simp: sep_conj_def dollar_def) subgoal for ps n apply(rule exI[where x=0]) apply(rule exI[where x=1]) apply(rule exI[where x=ps]) by auto done next case (Awhile x1 x2 C) then show ?case by auto qed lemma vc_mono: assumes "\s. P s \ Q s" shows "vc C P \ vc C Q" using assms proof(induct C arbitrary: P Q) case Askip then show ?case by auto next case (Aassign x1 x2) then show ?case by auto next case (Aseq C1 C2 P Q) then have i: "vc C1 (pre C2 P)" and ii: "vc C2 P" by auto from pre_mono[OF ] Aseq(4) have iii: "\s. pre C2 P s \ pre C2 Q s" by blast show ?case apply auto using Aseq(1)[OF i iii] Aseq(2)[OF ii Aseq(4)] by auto next case (Aif x1 C1 C2) then show ?case by auto next case (Awhile I b C P Q) then show ?case by auto qed lemma vc_sound': "vc C Q \ (\s n. P' (s, n) \ pre C Q (s, k * n)) \ (\s n. Q (s, n) \ Q' (s, n div k)) \ 0 < k \ \\<^sub>3\<^sub>' {P'} strip C { Q'}" using conseq vc_mono vc_sound by metis lemma pre_Frame: "(\s. P s \ pre C Q s) \ vc C Q \ (\C'. strip C = strip C' \ vc C' (Q ** F) \ (\s. (P ** F) s \ pre C' (Q ** F) s) )" proof (induct C arbitrary: P Q) case Askip show ?case proof (rule exI[where x="Askip"], safe) fix a b assume "(P \* F) (a, b)" then obtain ps1 ps2 n1 n2 where A: "ps1##ps2" "a=ps1+ps2" "b=n1+n2" and P: "P (ps1,n1)" and F: "F (ps2,n2)" unfolding sep_conj_def by auto from P Askip have p: "($ (Suc 0) \* Q) (ps1, n1)" by auto from p A F have "(($ (Suc 0) \* Q) \* F) (a, b)" apply(subst (2) sep_conj_def) by auto then show "pre SKIP (Q \* F) (a, b)" by (simp add: sep_conj_ac) qed simp next case (Aassign x a) show ?case proof (rule exI[where x="Aassign x a"], safe) fix ps n assume "(P \* F) (ps,n)" then obtain ps1 ps2 n1 n2 where o: "ps1##ps2" "ps=ps1+ps2" "n=n1+n2" and P: "P (ps1,n1)" and F: "F (ps2,n2)" unfolding sep_conj_def by auto from P Aassign(1) have z: "((\(ps, t). x \ dom ps \ vars a \ dom ps \ Q (ps(x \ paval a ps), t)) \* $ (Suc 0)) (ps1, n1)" by auto with o F show " pre (x ::= a) (Q \* F) (ps,n)" apply auto unfolding sep_conj_def dollar_def apply (auto) subgoal by(simp add: plus_fun_def) subgoal by(auto simp add: plus_fun_def) subgoal by (smt add_update_distrib dom_fun_upd domain_conv insert_dom option.simps(3) paval_extend sep_disj_fun_def) done qed auto next case (Aseq C1 C2) from Aseq(3) have pre: "\s. P s \ pre C1 (pre C2 Q) s" by auto from Aseq(4) have vc1: "vc C1 (pre C2 Q)" and vc2: "vc C2 Q" by auto from Aseq(1)[OF pre vc1] obtain C1' where S1: "strip C1 = strip C1'" and vc1': "vc C1' (pre C2 Q \* F)" and I1: "(\s. (P \* F) s \ pre C1' (pre C2 Q \* F) s)" by blast from Aseq(2)[of "pre C2 Q" Q, OF _ vc2] obtain C2' where S2: "strip C2 = strip C2'" and vc2': "vc C2' (Q \* F)" and I2: "(\s. (pre C2 Q \* F) s \ pre C2' (Q \* F) s) " by blast show ?case apply(rule exI[where x="Aseq C1' C2'"]) apply safe subgoal using S1 S2 by auto subgoal apply simp apply safe subgoal using vc_mono[OF I2 vc1'] . subgoal by (fact vc2') done subgoal using I1 I2 pre_mono by force done next case (Aif b C1 C2) from Aif(3) have i: "\s. P s \ ($ (Suc 0) \* (\(ps, n). vars b \ dom ps \ (if pbval b ps then pre C1 Q (ps, n) else pre C2 Q (ps, n)))) s" by simp from Aif(4) have vc1: "vc C1 Q" and vc2: "vc C2 Q" by auto from Aif(1)[where P="pre C1 Q" and Q="Q", OF _ vc1] obtain C1' where s1: "strip C1 = strip C1'" and v1: "vc C1' (Q \* F)" and p1: "(\s. (pre C1 Q \* F) s \ pre C1' (Q \* F) s)" by auto from Aif(2)[where P="pre C2 Q" and Q="Q", OF _ vc2] obtain C2' where s2: "strip C2 = strip C2'" and v2: "vc C2' (Q \* F)" and p2: "(\s. (pre C2 Q \* F) s \ pre C2' (Q \* F) s)" by auto show ?case apply(rule exI[where x="Aif b C1' C2'"]) proof safe fix ps n assume "(P \* F) (ps, n)" then obtain ps1 ps2 n1 n2 where o: "ps1##ps2" "ps=ps1+ps2" "n=n1+n2" and P: "P (ps1,n1)" and F: "F (ps2,n2)" unfolding sep_conj_def by auto from P i have P': "($ (Suc 0) \* (\(ps, n). vars b \ dom ps \ (if pbval b ps then pre C1 Q (ps, n) else pre C2 Q (ps, n)))) (ps1,n1)" by auto have PF: "(($ (Suc 0) \* (\(ps, n). vars b \ dom ps \ (if pbval b ps then pre C1 Q (ps, n) else pre C2 Q (ps, n)))) ** F) (ps,n)" apply(subst (2) sep_conj_def) apply(rule exI[where x="(ps1,n1)"]) apply(rule exI[where x="(ps2,n2)"]) using F P' o by auto from this[simplified sep_conj_assoc] obtain ps1 ps2 n1 n2 where o: "ps1##ps2" "ps=ps1+ps2" "n=n1+n2" and P: "($ (Suc 0)) (ps1,n1)" and F: "((\(ps, n). vars b \ dom ps \ (if pbval b ps then pre C1 Q (ps, n) else pre C2 Q (ps, n))) \* F) (ps2,n2)" unfolding sep_conj_def apply auto by fast then have "((\(ps, n). vars b \ dom ps \ (if pbval b ps then (pre C1 Q \* F) (ps, n) else (pre C2 Q \* F) (ps, n)))) (ps2,n2)" unfolding sep_conj_def apply auto apply (metis contra_subsetD domD map_add_dom_app_simps(1) plus_fun_conv sep_add_commute) using pbval_extend apply auto[1] apply (metis contra_subsetD domD map_add_dom_app_simps(1) plus_fun_conv sep_add_commute) using pbval_extend apply auto[1] done then have "((\(ps, n). vars b \ dom ps \ (if pbval b ps then (pre C1' (Q \* F)) (ps, n) else (pre C2' (Q \* F)) (ps, n)))) (ps2,n2)" using p1 p2 by auto with o P show "pre (IF b THEN C1' ELSE C2') (Q \* F) (ps, n)" apply auto apply(subst sep_conj_def) by force qed (auto simp: s1 s2 v1 v2) next case (Awhile I b C) from Awhile(2) have pre: "\s. P s \ (I ** $1) s" by auto from Awhile(3) have dom: "\ps n. I (ps, n) \ vars b \ dom ps" and tB: "\s. I s \ vars b \ dom (fst s) \ pbval b (fst s) \ pre C (I \* $ (Suc 0)) s" and fB: "\ps n. I (ps, n) \ vars b \ dom ps \ \ pbval b ps \ Q (ps, n)" and vcB: "vc C (I \* $(Suc 0))" by auto from Awhile(1)[OF tB vcB] obtain C' where st: "strip C = strip C'" and vc': "vc C' ((I \* $ (Suc 0)) \* F)" and pre': "(\s. ((\a. I a \ vars b \ dom (fst a) \ pbval b (fst a)) \* F) s \ pre C' ((I \* $ (Suc 0)) \* F) s)" by auto show ?case apply(rule exI[where x="Awhile (I**F) b C'"]) apply safe subgoal using st by simp subgoal apply simp apply safe subgoal using dom unfolding sep_conj_def apply auto by (metis domD sep_substate_disj_add subState subsetCE) subgoal using pre' apply(auto simp: sep_conj_ac) apply(subst (asm) sep_conj_def) apply(subst (asm) sep_conj_def) apply auto by (metis dom pbval_extend sep_add_commute sep_disj_commuteI) subgoal using fB unfolding sep_conj_def apply auto using dom pbval_extend by fastforce subgoal using vc' apply(auto simp: sep_conj_ac) done done subgoal apply simp using pre unfolding sep_conj_def apply auto by (smt semiring_normalization_rules(23) sep_add_assoc sep_add_commute sep_add_disjD sep_add_disjI1) done qed lemma vc_complete: "\\<^sub>3\<^sub>a {P} c { Q } \ (\C. vc C Q \ (\s. P s \ pre C Q s) \ strip C = c)" proof(induct rule: hoare3a.induct) case Skip then show ?case apply(rule exI[where x="Askip"]) by auto next case (Assign4 x a Q) then show ?case apply(rule exI[where x="Aassign x a"]) by auto next case (If P b c\<^sub>1 Q c\<^sub>2) from If(2) obtain C1 where A1: "vc C1 Q" "strip C1 = c\<^sub>1" and A2: "\ps n. (P (ps, n) \ lmaps_to_axpr b True ps) \ pre C1 Q (ps,n)" by blast from If(4) obtain C2 where B1: "vc C2 Q" "strip C2 = c\<^sub>2" and B2: "\ps n. (P (ps, n) \ lmaps_to_axpr b False ps) \ pre C2 Q (ps,n)" by blast show ?case apply(rule exI[where x="Aif b C1 C2"]) using A1 B1 apply auto subgoal for ps n unfolding sep_conj_def dollar_def apply auto apply(rule exI[where x="0"]) apply(rule exI[where x="1"]) apply(rule exI[where x="ps"]) using A2 B2 by auto done next case (Frame P C Q F) then obtain C' where vc: "vc C' Q" and pre: "(\s. P s \ pre C' Q s)" and strip: "strip C' = C" by auto show ?case using pre_Frame[OF pre vc] strip by metis next case (Seq P c\<^sub>1 Q c\<^sub>2 R) from Seq(2) obtain C1 where A1: "vc C1 Q" "strip C1 = c\<^sub>1" and A2: "\s. P s \ pre C1 Q s" by blast from Seq(4) obtain C2 where B1: "vc C2 R" "strip C2 = c\<^sub>2" and B2: "\s. Q s \ pre C2 R s" by blast show ?case apply(rule exI[where x="Aseq C1 C2"]) using B1 A1 apply auto subgoal using vc_mono B2 by auto subgoal apply(rule pre_mono[where P=Q]) using B2 apply auto using A2 by auto done next case (While I b c) then obtain C where 1: "vc C ((\(s, n). I (s, n) \ vars b \ dom s) \* $ 1)" "strip C = c" and 2: "\ps n. (I (ps, n) \ lmaps_to_axpr b True ps) \ pre C ((\(s, n). I (s, n) \ vars b \ dom s) \* $ 1) (ps,n) " by blast show ?case apply(rule exI[where x="Awhile (\(s, n). I (s, n) \ vars b \ dom s) b C"]) using 1 2 by auto next case (conseqS P c Q P' Q') then obtain C' where C': "vc C' Q" "(\s. P s \ pre C' Q s)" "strip C' = c" by blast show ?case apply(rule exI[where x=C']) using C' conseqS(3,4) pre_mono vc_mono by force qed theorem vc_completeness: assumes "\\<^sub>3\<^sub>' {P} c { Q}" shows "\C k. vc C (Q ** sep_true) \ (\ps n. P (ps, n) \ pre C (\(ps, n). (Q ** sep_true) (ps, n div k)) (ps, k * n)) \ strip C = c" proof - let ?QG = "\k (ps,n). (Q ** sep_true) (ps,n div k)" from assms obtain k where k[simp]: "k>0" and p: "\ps n. P (ps, n) \ wp\<^sub>3\<^sub>' c (\(ps, n). (Q ** sep_true) (ps, n div k)) (ps, k * n)" using valid_wp by blast from wpT_is_pre have R: "\\<^sub>3\<^sub>a {wp\<^sub>3\<^sub>' c (?QG k)} c {?QG k}" by auto have z: "(\s. (\(ps, n). (Q \* (\s. True)) (ps, n div k)) s) \ (\s. (\(ps, n). (Q \* (\s. True)) (ps, n)) s)" by (metis (no_types) case_prod_conv k neq0_conv nonzero_mult_div_cancel_left old.prod.exhaust) have z: "\ps n. ((Q \* (\s. True)) (ps, n div k) \ (Q \* (\s. True)) (ps, n))" proof - fix ps n assume "(Q \* (\s. True)) (ps, n div k) " - then guess ps1 n1 ps2 n2 unfolding sep_conj_def by auto - note o = this + then obtain ps1 n1 ps2 n2 + where o: "ps1 ## ps2" "ps = ps1 + ps2" "Q (ps1, n1)" "n div k = n1 + n2" + unfolding sep_conj_def by auto from o(4) have nn1: "n\n1" using k by (metis (full_types) add_leE div_le_dividend) show "(Q \* (\s. True)) (ps, n)" unfolding sep_conj_def apply(rule exI[where x="(ps1, n1)"]) apply(rule exI[where x="(ps2, n - n1)"]) using o nn1 by auto qed then have z': "\s. ((Q \* (\s. True)) (fst s, (snd s) div k) \ (Q \* (\s. True)) s)" by (metis prod.collapse) - from vc_complete[OF R] guess C by auto - note o = this + from vc_complete[OF R] obtain C + where o: "vc C (\(ps, n). (Q \* (\s. True)) (ps, n div k))" + "\a b. wp\<^sub>3\<^sub>' (strip C) (\(ps, n). (Q \* (\s. True)) (ps, n div k)) (a, b) \ + pre C (\(ps, n). (Q \* (\s. True)) (ps, n div k)) (a, b)" + "c = strip C" by auto have y: "\ps n. P (ps, n) \ pre C (\(ps, n). (Q \* (\s. True)) (ps, n div k)) (ps, k * n)" using o p by metis show ?thesis apply(rule exI[where x=C]) apply(rule exI[where x=k]) apply safe subgoal apply(rule vc_mono[OF _ o(1)]) using z by blast subgoal using y by blast subgoal using o by simp done qed end \ No newline at end of file diff --git a/thys/IFC_Tracking/IFC.thy b/thys/IFC_Tracking/IFC.thy --- a/thys/IFC_Tracking/IFC.thy +++ b/thys/IFC_Tracking/IFC.thy @@ -1,4580 +1,4578 @@ section \Definitions\ text \ This section contains all necessary definitions of this development. Section~\ref{sec:pm} contains the structural definition of our program model which includes the security specification as well as abstractions of control flow and data. Executions of our program model are defined in section~\ref{sec:ex}. Additional well-formedness properties are defined in section~\ref{sec:wf}. Our security property is defined in section~\ref{sec:sec}. Our characterisation of how information is propagated by executions of our program model is defined in section~\ref{sec:char-cp}, for which the correctness result can be found in section~\ref{sec:cor-cp}. Section~\ref{sec:char-scp} contains an additional approximation of this characterisation whose correctness result can be found in section~\ref{sec:cor-scp}. \ theory IFC -imports Main + imports Main begin subsection \Program Model\ text_raw \\label{sec:pm}\ text \Our program model contains all necessary components for the remaining development and consists of:\ record ('n, 'var, 'val, 'obs) ifc_problem = \ \A set of nodes representing program locations:\ nodes :: \'n set\ \ \An initial node where all executions start:\ entry :: \'n\ \ \A final node where executions can terminate:\ return :: \'n\ \ \An abstraction of control flow in the form of an edge relation:\ edges :: \('n \ 'n) set\ \ \An abstraction of variables written at program locations:\ writes :: \'n \ 'var set\ \ \An abstraction of variables read at program locations:\ reads :: \'n \ 'var set\ \ \A set of variables containing the confidential information in the initial state:\ hvars :: \'var set\ \ \A step function on location state pairs:\ step :: \('n \ ('var \ 'val)) \ ('n \ ('var \ 'val))\ \ \An attacker model producing observations based on the reached state at certain locations:\ att :: \'n \ (('var \ 'val) \ 'obs)\ text \We fix a program in the following in order to define the central concepts. The necessary well-formedness assumptions will be made in section~\ref{sec:wf}.\ locale IFC_def = fixes prob :: \('n, 'var, 'val, 'obs) ifc_problem\ begin text \Some short hands to the components of the program which we will utilise exclusively in the following.\ definition nodes where \nodes = ifc_problem.nodes prob\ definition entry where \entry = ifc_problem.entry prob\ definition return where \return = ifc_problem.return prob\ definition edges where \edges = ifc_problem.edges prob\ definition writes where \writes = ifc_problem.writes prob\ definition reads where \reads = ifc_problem.reads prob\ definition hvars where \hvars = ifc_problem.hvars prob\ definition step where \step = ifc_problem.step prob\ definition att where \att = ifc_problem.att prob\ text \The components of the step function for convenience.\ definition suc where \suc n \ = fst (step (n, \))\ definition sem where \sem n \ = snd (step (n, \))\ lemma step_suc_sem: \step (n,\) = (suc n \, sem n \)\ unfolding suc_def sem_def by auto subsubsection \Executions\ text \\label{sec:ex}\ text \In order to define what it means for a program to be well-formed, we first require concepts of executions and program paths.\ text \The sequence of nodes visited by the execution corresponding to an input state.\ definition path where \path \ k= fst ((step^^k) (entry,\))\ text \The sequence of states visited by the execution corresponding to an input state.\ definition kth_state ( \_\<^bsup>_\<^esup>\ [111,111] 110) where \\\<^bsup>k\<^esup> = snd ((step^^k) (entry,\))\ text \A predicate asserting that a sequence of nodes is a valid program path according to the control flow graph.\ definition is_path where \is_path \ = (\ n. (\ n, \ (Suc n)) \ edges)\ end subsubsection \Well-formed Programs\ text_raw \\label{sec:wf}\ text \The following assumptions define our notion of valid programs.\ locale IFC = IFC_def \prob\ for prob:: \('n, 'var, 'val, 'out) ifc_problem\ + assumes ret_is_node[simp,intro]: \return \ nodes\ and entry_is_node[simp,intro]: \entry \ nodes\ and writes: \\ v n. (\\. \ v \ sem n \ v) \ v \ writes n\ and writes_return: \writes return = {}\ and uses_writes: \\ n \ \'. (\ v \ reads n. \ v = \' v) \ \ v \ writes n. sem n \ v = sem n \' v\ and uses_suc: \\ n \ \'. (\ v \ reads n. \ v = \' v) \ suc n \ = suc n \'\ and uses_att: \\ n f \ \'. att n = Some f \ (\ v \ reads n. \ v = \' v) \ f \ = f \'\ and edges_complete[intro,simp]: \\m \. m \ nodes \ (m,suc m \) \ edges\ and edges_return : \\x. (return,x) \ edges \ x = return \ and edges_nodes: \edges \ nodes \ nodes\ and reaching_ret: \\ x. x \ nodes \ \ \ n. is_path \ \ \ 0 = x \ \ n = return\ subsection \Security\ text_raw \\label{sec:sec}\ text \We define our notion of security, which corresponds to what Bohannon et al.~\cite{Bohannon:2009:RN:1653662.1653673} refer to as indistinguishable security. In order to do so we require notions of observations made by the attacker, termination and equivalence of input states.\ context IFC_def begin subsubsection \Observations\ text_raw \\label{sec:obs}\ text \The observation made at a given index within an execution.\ definition obsp where \obsp \ k = (case att(path \ k) of Some f \ Some (f (\\<^bsup>k\<^esup>)) | None \ None)\ text \The indices within a path where an observation is made.\ definition obs_ids :: \(nat \ 'n) \ nat set\ where \obs_ids \ = {k. att (\ k) \ None}\ text \A predicate relating an observable index to the number of observations made before.\ definition is_kth_obs :: \(nat \ 'n) \ nat \ nat \ bool\where \is_kth_obs \ k i = (card (obs_ids \ \ {.. att (\ i) \ None)\ text \The final sequence of observations made for an execution.\ definition obs where \obs \ k = (if (\i. is_kth_obs (path \) k i) then obsp \ (THE i. is_kth_obs (path \) k i) else None)\ text \Comparability of observations.\ definition obs_prefix :: \(nat \ 'obs option) \ (nat \ 'obs option) \ bool\ (infix \\\ 50) where \a \ b \ \ i. a i \ None \ a i = b i\ definition obs_comp (infix \\\ 50) where \a \ b \ a \ b \ b \ a\ subsubsection \Low equivalence of input states\ definition restrict (infix \\\ 100 ) where \f\U = (\ n. if n \ U then f n else undefined)\ text \Two input states are low equivalent if they coincide on the non high variables.\ definition loweq (infix \=\<^sub>L\ 50) where \\ =\<^sub>L \' = (\\(-hvars) = \'\(-hvars))\ subsubsection \Termination\ text \An execution terminates iff it reaches the terminal node at any point.\ definition terminates where \terminates \ \ \ i. path \ i = return\ subsubsection \Security Property\ text \The fixed program is secure if and only if for all pairs of low equivalent inputs the observation sequences are comparable and if the execution for an input state terminates then the observation sequence is not missing any observations.\ definition secure where \secure \ \ \ \'. \ =\<^sub>L \' \ (obs \ \ obs \' \ (terminates \ \ obs \' \ obs \))\ subsection \Characterisation of Information Flows\ text \We now define our characterisation of information flows which tracks data and control dependencies within executions. To do so we first require some additional concepts.\ subsubsection \Post Dominance\ text \We utilise the post dominance relation in order to define control dependence.\ text \The basic post dominance relation.\ definition is_pd (infix \pd\\ 50) where \y pd\ x \ x \ nodes \ (\ \ n. is_path \ \ \ (0::nat) = x \ \ n = return \ (\k\n. \ k = y))\ text \The immediate post dominance relation.\ definition is_ipd (infix \ipd\\ 50)where \y ipd\ x \ x \ y \ y pd\ x \ (\ z. z\x \ z pd\ x \ z pd\ y)\ definition ipd where \ipd x = (THE y. y ipd\ x)\ text \The post dominance tree.\ definition pdt where \pdt = {(x,y). x\y \ y pd\ x}\ subsubsection \Control Dependence\ text \An index on an execution path is control dependent upon another if the path does not visit the immediate post domiator of the node reached by the smaller index.\ definition is_cdi (\_ cd\<^bsup>_\<^esup>\ _\ [51,51,51]50) where \i cd\<^bsup>\\<^esup>\ k \ is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ j \ ipd (\ k))\ text \The largest control dependency of an index is the immediate control dependency.\ definition is_icdi (\_ icd\<^bsup>_\<^esup>\ _\ [51,51,51]50) where \n icd\<^bsup>\\<^esup>\ n' \ is_path \ \ n cd\<^bsup>\\<^esup>\ n' \ (\ m \ {n'<.. n cd\<^bsup>\\<^esup>\ m)\ text \For the definition of the control slice, which we will define next, we require the uniqueness of the immediate control dependency.\ lemma icd_uniq: assumes \m icd\<^bsup>\\<^esup>\ n\ \ m icd\<^bsup>\\<^esup>\ n'\ shows \n = n'\ proof - { fix n n' assume *: \m icd\<^bsup>\\<^esup>\ n\ \ m icd\<^bsup>\\<^esup>\ n'\ \n < n'\ have \n' using * unfolding is_icdi_def is_cdi_def by auto hence \\ m cd\<^bsup>\\<^esup>\ n'\ using * unfolding is_icdi_def by auto with *(2) have \False\ unfolding is_icdi_def by auto } thus ?thesis using assms by (metis linorder_neqE_nat) qed subsubsection \Control Slice\ text \We utilise the control slice, that is the sequence of nodes visited by the control dependencies of an index, to match indices between executions.\ function cs:: \(nat \ 'n) \ nat \ 'n list\ (\cs\<^bsup>_\<^esup> _\ [51,70] 71) where \cs\<^bsup>\\<^esup> n = (if (\ m. n icd\<^bsup>\\<^esup>\ m) then (cs \ (THE m. n icd\<^bsup>\\<^esup>\ m))@[\ n] else [\ n])\ by pat_completeness auto termination \cs\ proof show \wf (measure snd)\ by simp fix \ n define m where \m == (The (is_icdi n \))\ assume \Ex (is_icdi n \)\ hence \n icd\<^bsup>\\<^esup>\ m\ unfolding m_def by (metis (full_types) icd_uniq theI') hence \m < n\ unfolding is_icdi_def is_cdi_def by simp thus \((\, The (is_icdi n \)), \, n) \ measure snd\ by (metis in_measure m_def snd_conv) qed inductive cs_less (infix \\\ 50) where \length xs < length ys \ take (length xs) ys = xs \ xs \ ys\ definition cs_select (infix \\\ 50) where \\\xs = (THE k. cs\<^bsup>\\<^esup> k = xs)\ subsubsection \Data Dependence\ text \Data dependence is defined straight forward. An index is data dependent upon another, if the index reads a variable written by the earlier index and the variable in question has not been written by any index in between.\ definition is_ddi (\_ dd\<^bsup>_,_\<^esup>\ _\ [51,51,51,51] 50) where \n dd\<^bsup>\,v\<^esup>\ m \ is_path \ \ m < n \ v \ reads (\ n) \ (writes (\ m)) \ (\ l \ {m<.. writes (\ l))\ subsubsection \Characterisation via Critical Paths\ text_raw \\label{sec:char-cp}\ text \With the above we define the set of critical paths which as we will prove characterise the matching points in executions where diverging data is read.\ inductive_set cp where \ \Any pair of low equivalent input states and indices where a diverging high variable is first read is critical.\ \\\ =\<^sub>L \'; cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'; h \ reads(path \ n); (\\<^bsup>n\<^esup>) h \ (\'\<^bsup>n'\<^esup>) h; \ kwrites(path \ k); \ k'writes(path \' k') \ \ ((\,n),(\',n')) \ cp\ | \ \If from a pair of critical indices in two executions there exist data dependencies from both indices to a pair of matching indices where the variable diverges, the later pair of indices is critical.\ \\((\,k),(\',k')) \ cp; n dd\<^bsup>path \,v\<^esup>\ k; n' dd\<^bsup>path \',v\<^esup>\ k'; cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'; (\\<^bsup>n\<^esup>) v \ (\'\<^bsup>n'\<^esup>) v \ \ ((\,n),(\',n')) \ cp\ | \ \If from a pair of critical indices the executions take different branches and one of the critical indices is a control dependency of an index that is data dependency of a matched index where diverging data is read and the variable in question is not written by the other execution after the executions first reached matching indices again, then the later matching pair of indices is critical.\ \\((\,k),(\',k')) \ cp; n dd\<^bsup>path \,v\<^esup>\ l; l cd\<^bsup>path \\<^esup>\ k; cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'; path \ (Suc k) \ path \' (Suc k'); (\\<^bsup>n\<^esup>) v \ (\'\<^bsup>n'\<^esup>) v; \j'\{(LEAST i'. k' < i' \ (\i. cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i'))..writes (path \' j') \ \ ((\,n),(\',n')) \ cp\ | \ \The relation is symmetric.\ \\((\,k),(\',k')) \ cp\ \ ((\',k'),(\,k)) \ cp\ text \Based on the set of critical paths, the critical observable paths are those that either directly reach observable nodes or are diverging control dependencies of an observable index.\ inductive_set cop where \\((\,n),(\',n')) \ cp; path \ n \ dom att \ \ ((\,n),(\',n')) \ cop\ | \\((\,k),(\',k')) \ cp; n cd\<^bsup>path \\<^esup>\ k; path \ (Suc k) \ path \' (Suc k'); path \ n \ dom att \ \ ((\,n),(\',k')) \ cop\ subsubsection \Approximation via Single Critical Paths\ text_raw \\label{sec:char-scp}\ text \For applications we also define a single execution approximation.\ definition is_dcdi_via (\_ dcd\<^bsup>_,_\<^esup>\ _ via _ _\ [51,51,51,51,51,51] 50) where \n dcd\<^bsup>\,v\<^esup>\ m via \' m' = (is_path \ \ m < n \ (\ l' n'. cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m' \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n' \ n' dd\<^bsup>\',v\<^esup>\ l' \ l' cd\<^bsup>\'\<^esup>\ m') \ (\ l \ {m.. writes(\ l)))\ inductive_set scp where \\h \ hvars; h \ reads (path \ n); (\ k writes(path \ k))\ \ (path \,n) \ scp\ | \\(\,m) \ scp; n cd\<^bsup>\\<^esup>\ m\ \ (\,n) \ scp\| \\(\,m) \ scp; n dd\<^bsup>\,v\<^esup>\ m\ \ (\,n) \ scp\| \\(\,m) \ scp; (\',m') \ scp; n dcd\<^bsup>\,v\<^esup>\ m via \' m'\ \ (\,n) \ scp\ inductive_set scop where \\(\,n) \ scp; \ n \ dom att\ \ (\,n) \ scop\ subsubsection \Further Definitions\ text \The following concepts are utilised by the proofs.\ inductive contradicts (infix \\\ 50) where \\cs\<^bsup>\'\<^esup> k' \ cs\<^bsup>\\<^esup> k ; \ = path \; \' = path \' ; \ (Suc (\\cs\<^bsup>\'\<^esup> k')) \ \' (Suc k')\ \ (\', k') \ (\, k)\| \\cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k ; \ = path \; \' = path \' ; \\<^bsup>k\<^esup> \ (reads (\ k)) \ \'\<^bsup>k'\<^esup> \ (reads (\ k))\ \ (\',k') \ (\,k)\ definition path_shift (infixl \\\ 51) where [simp]: \\\m = (\ n. \ (m+n))\ definition path_append :: \(nat \ 'n) \ nat \ (nat \ 'n) \ (nat \ 'n)\ (\_ @\<^bsup>_\<^esup> _\ [0,0,999] 51) where [simp]: \\ @\<^bsup>m\<^esup> \' = (\n.(if n \ m then \ n else \' (n-m)))\ definition eq_up_to :: \(nat \ 'n) \ nat \ (nat \ 'n) \ bool\ (\_ =\<^bsub>_\<^esub> _\ [55,55,55] 50) where \\ =\<^bsub>k\<^esub> \' = (\ i \ k. \ i = \' i)\ end (* End of locale IFC_def *) section \Proofs\ text_raw \\label{sec:proofs}\ subsection \Miscellaneous Facts\ lemma option_neq_cases: assumes \x \ y\ obtains (none1) a where \x = None\ \y = Some a\ | (none2) a where \x = Some a\ \y = None\ | (some) a b where \x = Some a\ \y = Some b\ \a \ b\ using assms by fastforce lemmas nat_sym_cases[case_names less sym eq] = linorder_less_wlog lemma mod_bound_instance: assumes \j < (i::nat)\ obtains j' where \k < j'\ and \j' mod i = j\ proof - have \k < Suc k * i + j\ using assms less_imp_Suc_add by fastforce moreover have \(Suc k * i + j) mod i = j\ by (metis assms mod_less mod_mult_self3) ultimately show \thesis\ using that by auto qed lemma list_neq_prefix_cases: assumes \ls \ ls'\ and \ls \ Nil\ and \ls' \ Nil\ obtains (diverge) xs x x' ys ys' where \ls = xs@[x]@ys\ \ls' = xs@[x']@ys'\ \x \ x'\ | (prefix1) xs where \ls = ls'@xs\ and \xs \ Nil\ | (prefix2) xs where \ls@xs = ls'\ and \xs \ Nil\ using assms proof (induct \length ls\ arbitrary: \ls\ \ls'\ rule: less_induct) case (less ls ls') obtain z zs z' zs' where lz: \ls = z#zs\ \ls' = z'#zs'\ by (metis list.exhaust less(6,7)) show \?case\ proof cases assume zz: \z = z'\ hence zsz: \zs \ zs'\ using less(5) lz by auto have lenz: \length zs < length ls\ using lz by auto show \?case\ proof(cases \zs = Nil\) assume zs: \zs = Nil\ hence \zs' \ Nil\ using zsz by auto moreover have \ls@zs' = ls'\ using zs lz zz by auto ultimately show \thesis\ using less(4) by blast next assume zs: \zs \ Nil\ show \thesis\ proof (cases \zs' = Nil\) assume \zs' = Nil\ hence \ls = ls'@zs\ using lz zz by auto thus \thesis\ using zs less(3) by blast next assume zs': \zs' \ Nil\ { fix xs x ys x' ys' assume \zs = xs @ [x] @ ys\ \zs' = xs @ [x'] @ ys'\ and xx: \x \ x'\ hence \ls = (z#xs) @ [x] @ ys\ \ls' = (z#xs) @ [x'] @ ys'\ using lz zz by auto hence \thesis\ using less(2) xx by blast } note * = this { fix xs assume \zs = zs' @ xs\ and xs: \xs \ []\ hence \ls = ls' @ xs\ using lz zz by auto hence \thesis\ using xs less(3) by blast } note ** = this { fix xs assume \zs@xs = zs'\ and xs: \xs \ []\ hence \ls@xs = ls'\ using lz zz by auto hence \thesis\ using xs less(4) by blast } note *** = this have \(\xs x ys x' ys'. zs = xs @ [x] @ ys \ zs' = xs @ [x'] @ ys' \ x \ x' \ thesis) \ (\xs. zs = zs' @ xs \ xs \ [] \ thesis) \ (\xs. zs @ xs = zs' \ xs \ [] \ thesis) \ thesis\ using less(1)[OF lenz _ _ _ zsz zs zs' ] . thus \thesis\ using * ** *** by blast qed qed next assume \z \ z'\ moreover have \ls = []@[z]@zs\ \ls' = []@[z']@zs'\ using lz by auto ultimately show \thesis\ using less(2) by blast qed qed lemma three_cases: assumes \A \ B \ C\ obtains \A\ | \B\ | \C\ using assms by auto lemma insort_greater: \\ x \ set ls. x < y \ insort y ls = ls@[y]\ by (induction \ls\,auto) lemma insort_append_first: assumes \\ y \ set ys. x \ y\ shows \insort x (xs@ys) = insort x xs @ ys\ using assms by (induction \xs\,auto,metis insort_is_Cons) lemma sorted_list_of_set_append: assumes \finite xs\ \finite ys\ \\ x \ xs. \ y \ ys. x < y\ shows \sorted_list_of_set (xs \ ys) = sorted_list_of_set xs @ (sorted_list_of_set ys)\ using assms(1,3) proof (induction \xs\) case empty thus \?case\ by simp next case (insert x xs) hence iv: \sorted_list_of_set (xs \ ys) = sorted_list_of_set xs @ sorted_list_of_set ys\ by blast have le: \\ y \ set (sorted_list_of_set ys). x < y\ using insert(4) assms(2) sorted_list_of_set by auto have \sorted_list_of_set (insert x xs \ ys) = sorted_list_of_set (insert x (xs \ ys))\ by auto also have \\ = insort x (sorted_list_of_set (xs \ ys))\ by (metis Un_iff assms(2) finite_Un insert.hyps(1) insert.hyps(2) insert.prems insertI1 less_irrefl sorted_list_of_set_insert) also have \\ = insort x (sorted_list_of_set xs @ sorted_list_of_set ys)\ using iv by simp also have \\ = insort x (sorted_list_of_set xs) @ sorted_list_of_set ys\ by (metis le insort_append_first less_le_not_le) also have \\ = sorted_list_of_set (insert x xs) @ sorted_list_of_set ys\ using sorted_list_of_set_insert[OF insert(1),of \x\] insert(2) by auto finally show \?case\ . qed lemma filter_insort: \sorted xs \ filter P (insort x xs) = (if P x then insort x (filter P xs) else filter P xs)\ by (induction \xs\, simp) (metis filter_insort filter_insort_triv map_ident) lemma filter_sorted_list_of_set: assumes \finite xs\ shows \filter P (sorted_list_of_set xs) = sorted_list_of_set {x \ xs. P x}\ using assms proof(induction \xs\) case empty thus \?case\ by simp next case (insert x xs) have *: \set (sorted_list_of_set xs) = xs\ \sorted (sorted_list_of_set xs)\ \distinct (sorted_list_of_set xs)\ by (auto simp add: insert.hyps(1)) have **: \P x \ {y \ insert x xs. P y} = insert x {y \ xs. P y}\ by auto have ***: \\ P x \ {y \ insert x xs. P y} = {y \ xs. P y}\ by auto note filter_insort[OF *(2),of \P\ \x\] sorted_list_of_set_insert[OF insert(1), of \x\] insert(2,3) ** *** thus \?case\ by (metis (mono_tags) "*"(1) List.finite_set distinct_filter distinct_insort distinct_sorted_list_of_set set_filter sorted_list_of_set_insert) qed lemma unbounded_nat_set_infinite: assumes \\ (i::nat). \ j\i. j \ A\ shows \\ finite A\ using assms by (metis finite_nat_set_iff_bounded_le not_less_eq_eq) lemma infinite_ascending: assumes nf: \\ finite (A::nat set)\ obtains f where \range f = A\ \\ i. f i < f (Suc i)\ proof let \?f\ = \\ i. (LEAST a. a \ A \ card (A \ {.. { fix i obtain a where \a \ A\ \card (A \ {.. proof (induction \i\ arbitrary: \thesis\) case 0 let \?a0\ = \(LEAST a. a \ A)\ have \?a0 \ A\ by (metis LeastI empty_iff finite.emptyI nf set_eq_iff) moreover have \\b. b \ A \ ?a0 \ b\ by (metis Least_le) hence \card (A \ {.. by force ultimately show \?case\ using 0 by blast next case (Suc i) obtain a where aa: \a \ A\ and card: \card (A \ {.. using Suc.IH by metis have nf': \~ finite (A - {..a})\ using nf by auto let \?b\ = \LEAST b. b \ A - {..a}\ have bin: \?b \ A-{..a}\ by (metis LeastI empty_iff finite.emptyI nf' set_eq_iff) have le: \\c. c \ A-{..a} \ ?b \ c\ by (metis Least_le) have ab: \a < ?b\ using bin by auto have \\ c. c \ A \ c < ?b \ c \ a\ using le by force hence \A \ {.. {.. using bin ab aa by force hence \card (A \{.. using card by auto thus \?case\ using Suc.prems bin by auto qed note \\ thesis. ((\a. a \ A \ card (A \ {.. thesis) \ thesis)\ } note ex = this { fix i obtain a where a: \a \ A \ card (A \{.. using ex by blast have ina: \?f i \ A\ and card: \card (A \{.. using LeastI[of \\ a. a \ A \ card (A \{.. \a\, OF a] by auto obtain b where b: \b \ A \ card (A \{.. using ex by blast have inab: \?f (Suc i) \ A\ and cardb: \card (A \{.. using LeastI[of \\ a. a \ A \ card (A \{.. \b\, OF b] by auto have \?f i < ?f (Suc i)\ proof (rule ccontr) assume \\ ?f i < ?f (Suc i)\ hence \A \{.. A \{.. by auto moreover have \finite (A \{.. by auto ultimately have \card(A \{.. card (A \{.. by (metis (erased, lifting) card_mono) thus \False\ using card cardb by auto qed note this ina } note b = this thus \\ i. ?f i < ?f (Suc i)\ by auto have *: \range ?f \ A\ using b by auto moreover { fix a assume ina: \a \ A\ let \?i\ = \card (A \ {.. obtain b where b: \b \ A \ card (A \{.. using ex by blast have inab: \?f ?i \ A\ and cardb: \card (A \{.. using LeastI[of \\ a. a \ A \ card (A \{.. \b\, OF b] by auto have le: \?f ?i \ a\ using Least_le[of \\ a. a \ A \ card (A \{.. \a\] ina by auto have \a = ?f ?i\ proof (rule ccontr) have fin: \finite (A \ {.. by auto assume \a \ ?f ?i\ hence \?f ?i < a\ using le by simp hence \?f ?i \ A \ {.. using inab by auto moreover have \A \ {.. A \ {.. using le by auto hence \A \ {.. {.. using cardb card_subset_eq[OF fin] by auto ultimately show \False\ by auto qed hence \a \ range ?f\ by auto } hence \A \ range ?f\ by auto ultimately show \range ?f = A\ by auto qed lemma mono_ge_id: \\ i. f i < f (Suc i) \ i \ f i\ apply (induction \i\,auto) by (metis not_le not_less_eq_eq order_trans) lemma insort_map_mono: assumes mono: \\ n m. n < m \ f n < f m\ shows \map f (insort n ns) = insort (f n) (map f ns)\ apply (induction \ns\) apply auto apply (metis not_less not_less_iff_gr_or_eq mono) apply (metis antisym_conv1 less_imp_le mono) apply (metis mono not_less) by (metis mono not_less) lemma sorted_list_of_set_map_mono: assumes mono: \\ n m. n < m \ f n < f m\ and fin: \finite A\ shows \map f (sorted_list_of_set A) = sorted_list_of_set (f`A)\ using fin proof (induction) case empty thus \?case\ by simp next case (insert x A) have [simp]:\sorted_list_of_set (insert x A) = insort x (sorted_list_of_set A)\ using insert sorted_list_of_set_insert by simp have \f ` insert x A = insert (f x) (f ` A)\ by auto moreover have \f x \ f`A\ apply (rule ccontr) using insert(2) mono apply auto by (metis insert.hyps(2) mono neq_iff) ultimately have \sorted_list_of_set (f ` insert x A) = insort (f x) (sorted_list_of_set (f`A))\ using insert(1) sorted_list_of_set_insert by simp also have \\ = insort (f x) (map f (sorted_list_of_set A))\ using insert.IH by auto also have \\ = map f (insort x (sorted_list_of_set A))\ using insort_map_mono[OF mono] by auto finally show \map f (sorted_list_of_set (insert x A)) = sorted_list_of_set (f ` insert x A)\ by simp qed lemma GreatestIB: fixes n :: \nat\ and P assumes a:\\k\n. P k\ shows GreatestBI: \P (GREATEST k. k\n \ P k)\ and GreatestB: \(GREATEST k. k\n \ P k) \ n\ proof - show \P (GREATEST k. k\n \ P k)\ using GreatestI_ex_nat[OF assms] by auto show \(GREATEST k. k\n \ P k) \ n\ using GreatestI_ex_nat[OF assms] by auto qed lemma GreatestB_le: fixes n :: \nat\ assumes \x\n\ and \P x\ shows \x \ (GREATEST k. k\n \ P k)\ proof - have *: \\ y. y\n \ P y \ y by auto then show \x \ (GREATEST k. k\n \ P k)\ using assms by (blast intro: Greatest_le_nat) qed lemma LeastBI_ex: assumes \\k \ n. P k\ shows \P (LEAST k::'c::wellorder. P k)\ and \(LEAST k. P k) \ n\ proof - - from assms guess k .. - hence k: \k \ n\ \P k\ by auto + from assms obtain k where k: "k \ n" "P k" by blast thus \P (LEAST k. P k)\ using LeastI[of \P\ \k\] by simp show \(LEAST k. P k) \ n\ using Least_le[of \P\ \k\] k by auto qed lemma allB_atLeastLessThan_lower: assumes \(i::nat) \ j\ \\ x\{i.. shows \\ x\{j.. proof fix x assume \x\{j.. hence \x\{i.. using assms(1) by simp thus \P x\ using assms(2) by auto qed subsection \Facts about Paths\ context IFC begin lemma path0: \path \ 0 = entry\ unfolding path_def by auto lemma path_in_nodes[intro]: \path \ k \ nodes\ proof (induction \k\) case (Suc k) hence \\ \'. (path \ k, suc (path \ k) \') \ edges\ by auto hence \(path \ k, path \ (Suc k)) \ edges\ unfolding path_def by (metis suc_def comp_apply funpow.simps(2) prod.collapse) thus \?case\ using edges_nodes by force qed (auto simp add: path_def) lemma path_is_path[simp]: \is_path (path \)\ unfolding is_path_def path_def using step_suc_sem apply auto by (metis path_def suc_def edges_complete path_in_nodes prod.collapse) lemma term_path_stable: assumes \is_path \\ \\ i = return\ and le: \i \ j\ shows \\ j = return\ using le proof (induction \j\) case (Suc j) show \?case\ proof cases assume \i\j\ hence \\ j = return\ using Suc by simp hence \(return, \ (Suc j)) \ edges\ using assms(1) unfolding is_path_def by metis thus \\ (Suc j) = return\ using edges_return by auto next assume \\ i \ j\ hence \Suc j = i\ using Suc by auto thus \?thesis\ using assms(2) by auto qed next case 0 thus \?case\ using assms by simp qed lemma path_path_shift: assumes \is_path \\ shows \is_path (\\m)\ using assms unfolding is_path_def by simp lemma path_cons: assumes \is_path \\ \is_path \'\ \\ m = \' 0\ shows \is_path (\ @\<^bsup>m\<^esup> \')\ unfolding is_path_def proof(rule,cases) fix n assume \m < n\ thus \((\ @\<^bsup>m\<^esup> \') n, (\ @\<^bsup>m\<^esup> \') (Suc n)) \ edges\ using assms(2) unfolding is_path_def path_append_def by (auto,metis Suc_diff_Suc diff_Suc_Suc less_SucI) next fix n assume *: \\ m < n\ thus \((\ @\<^bsup>m\<^esup> \') n, (\ @\<^bsup>m\<^esup> \') (Suc n)) \ edges\ proof cases assume [simp]: \n = m\ thus \?thesis\ using assms unfolding is_path_def path_append_def by force next assume \n \ m\ hence \Suc n \ m\ \n\ m\ using * by auto with assms(1) show \?thesis\ unfolding is_path_def by auto qed qed lemma is_path_loop: assumes \is_path \\ \0 < i\ \\ i = \ 0\ shows \is_path (\ n. \ (n mod i))\ unfolding is_path_def proof (rule,cases) fix n assume \0 < Suc n mod i\ hence \Suc n mod i = Suc (n mod i)\ by (metis mod_Suc neq0_conv) moreover have \(\ (n mod i), \ (Suc (n mod i))) \ edges\ using assms(1) unfolding is_path_def by auto ultimately show \(\ (n mod i), \ (Suc n mod i)) \ edges\ by simp next fix n assume \\ 0 < Suc n mod i\ hence \Suc n mod i = 0\ by auto moreover hence \n mod i = i - 1\ using assms(2) by (metis Zero_neq_Suc diff_Suc_1 mod_Suc) ultimately show \(\(n mod i), \ (Suc n mod i)) \ edges\ using assms(1) unfolding is_path_def by (metis assms(3) mod_Suc) qed lemma path_nodes: \is_path \ \ \ k \ nodes\ unfolding is_path_def using edges_nodes by force lemma direct_path_return': assumes \is_path \ \ \\ 0 = x\ \x \ return\ \\ n = return\ obtains \' n' where \is_path \'\ \\' 0 = x\ \\' n' = return\ \\ i> 0. \' i \ x\ using assms proof (induction \n\ arbitrary: \\\ rule: less_induct) case (less n \) hence ih: \\ n' \'. n' < n \ is_path \' \ \' 0 = x \ \' n' = return \ thesis\ using assms by auto show \thesis\ proof cases assume \\ i>0. \ i \ x\ thus \thesis\ using less by auto next assume \\ (\ i>0. \ i \ x)\ then obtain i where \0 \\ i = x\ by auto hence \(\\i) 0 = x\ by auto moreover have \i < n\ using less(3,5,6) \\ i = x\ by (metis linorder_neqE_nat term_path_stable less_imp_le) hence \(\\i) (n-i) = return\ using less(6) by auto moreover have \is_path (\\i)\ using less(3) by (metis path_path_shift) moreover have \n - i < n\ using \0 \i < n\ by auto ultimately show \thesis\ using ih by auto qed qed lemma direct_path_return: assumes \x \ nodes\ \x \ return\ obtains \ n where \is_path \\ \\ 0 = x\ \\ n = return\ \\ i> 0. \ i \ x\ using direct_path_return'[of _ \x\] reaching_ret[OF assms(1)] assms(2) by blast lemma path_append_eq_up_to: \(\ @\<^bsup>k\<^esup> \') =\<^bsub>k\<^esub> \\ unfolding eq_up_to_def by auto lemma eq_up_to_le: assumes \k \ n\ \\ =\<^bsub>n\<^esub> \'\ shows \\ =\<^bsub>k\<^esub> \'\ using assms unfolding eq_up_to_def by auto lemma eq_up_to_refl: shows \\ =\<^bsub>k\<^esub> \\ unfolding eq_up_to_def by auto lemma eq_up_to_sym: assumes \\ =\<^bsub>k\<^esub> \'\ shows \\' =\<^bsub>k\<^esub> \\ using assms unfolding eq_up_to_def by auto lemma eq_up_to_apply: assumes \\ =\<^bsub>k\<^esub> \'\ \j \ k\ shows \\ j = \' j\ using assms unfolding eq_up_to_def by auto lemma path_swap_ret: assumes \is_path \\ obtains \' n where \is_path \'\ \\ =\<^bsub>k\<^esub> \'\ \\' n = return\ proof - have nd: \\ k \ nodes\ using assms path_nodes by simp obtain \' n where *: \is_path \'\ \\' 0 = \ k\ \\' n = return\ using reaching_ret[OF nd] by blast have \\ =\<^bsub>k\<^esub> (\@\<^bsup>k\<^esup> \')\ by (metis eq_up_to_sym path_append_eq_up_to) moreover have \is_path (\@\<^bsup>k\<^esup> \')\ using assms * path_cons by metis moreover have \(\@\<^bsup>k\<^esup> \') (k + n) = return\ using * by auto ultimately show \thesis\ using that by blast qed lemma path_suc: \path \ (Suc k) = fst (step (path \ k, \\<^bsup>k\<^esup>))\ by (induction \k\, auto simp: path_def kth_state_def) lemma kth_state_suc: \\\<^bsup>Suc k\<^esup> = snd (step (path \ k, \\<^bsup>k\<^esup>))\ by (induction \k\, auto simp: path_def kth_state_def) subsection \Facts about Post Dominators\ lemma pd_trans: assumes 1: \y pd\ x\ and 2: \z pd\y\ shows \z pd\x\ proof - { fix \ n assume 3[simp]: \is_path \\ \\ 0 = x\ \\ n = return\ then obtain k where \\ k = y\ and 7: \k \ n\ using 1 unfolding is_pd_def by blast then have \(\\k) 0 = y\ and \(\\k) (n-k) = return\ by auto moreover have \is_path (\\k)\ by(metis 3(1) path_path_shift) ultimately obtain k' where 8: \(\\k) k' = z\ and \k' \ n-k\ using 2 unfolding is_pd_def by blast hence \k+k'\n\ and \\ (k+ k') = z\ using 7 by auto hence \\k\n. \ k = z\ using path_nodes by auto } thus \?thesis\ using 1 unfolding is_pd_def by blast qed lemma pd_path: assumes \y pd\ x\ obtains \ n k where \is_path \\ and \\ 0 = x\ and \\ n = return\ and \\ k = y\ and \k \ n\ using assms unfolding is_pd_def using reaching_ret[of \x\] by blast lemma pd_antisym: assumes xpdy: \x pd\ y\ and ypdx: \y pd\ x\ shows \x = y\ proof - obtain \ n where path: \is_path \\ and \0: \\ 0 = x\ and \n: \\ n = return\ using pd_path[OF ypdx] by metis hence kex: \\k\n. \ k = y\ using ypdx unfolding is_pd_def by auto obtain k where k: \k = (GREATEST k. k\n \ \ k = y)\ by simp have \k: \\ k = y\ and kn: \k \ n\ using k kex by (auto intro: GreatestIB) have kpath: \is_path (\\k)\ by (metis path_path_shift path) moreover have k0: \(\\k) 0 = y\ using \k by simp moreover have kreturn: \(\\k) (n-k) = return\ using kn \n by simp ultimately have ky': \\k'\(n-k).(\\k) k' = x\ using xpdy unfolding is_pd_def by simp obtain k' where k': \k' = (GREATEST k'. k'\(n-k) \ (\\k) k' = x)\ by simp with ky' have \k': \(\\k) k' = x\ and kn': \k' \ (n-k)\ by (auto intro: GreatestIB) have k'path: \is_path (\\k\k')\ using kpath by(metis path_path_shift) moreover have k'0: \(\\k\k') 0 = x\ using \k' by simp moreover have k'return: \(\\k\k') (n-k-k') = return\ using kn' kreturn by (metis path_shift_def le_add_diff_inverse) ultimately have ky'': \\k''\(n-k-k').(\\k\k') k'' = y\ using ypdx unfolding is_pd_def by blast obtain k'' where k'': \k''= (GREATEST k''. k''\(n-k-k') \ (\\k\k') k'' = y)\ by simp with ky'' have \k'': \(\\k\k') k'' = y\ and kn'': \k'' \ (n-k-k')\ by (auto intro: GreatestIB) from this(1) have \\ (k + k' + k'') = y\ by (metis path_shift_def add.commute add.left_commute) moreover have \k + k' +k'' \ n\ using kn'' kn' kn by simp ultimately have \k + k' + k''\ k\ using k by(auto simp: GreatestB_le) hence \k' = 0\ by simp with k0 \k' show \x = y\ by simp qed lemma pd_refl[simp]: \x \ nodes \ x pd\ x\ unfolding is_pd_def by blast lemma pdt_trans_in_pdt: \(x,y) \ pdt\<^sup>+ \ (x,y) \ pdt\ proof (induction rule: trancl_induct) case base thus \?case\ by simp next case (step y z) show \?case\ unfolding pdt_def proof (simp) have *: \y pd\ x\ \z pd\ y\ using step unfolding pdt_def by auto hence [simp]: \z pd\ x\ using pd_trans[where x=\x\ and y=\y\ and z=\z\] by simp have \x\z\ proof assume \x = z\ hence \z pd\ y\ \y pd\ z\ using * by auto hence \z = y\ using pd_antisym by auto thus \False\ using step(2) unfolding pdt_def by simp qed thus \x \ z \ z pd\ x\ by auto qed qed lemma pdt_trancl_pdt: \pdt\<^sup>+ = pdt\ using pdt_trans_in_pdt by fast lemma trans_pdt: \trans pdt\ by (metis pdt_trancl_pdt trans_trancl) definition [simp]: \pdt_inv = pdt\\ lemma wf_pdt_inv: \wf (pdt_inv)\ proof (rule ccontr) assume \\ wf (pdt_inv)\ then obtain f where \\i. (f (Suc i), f i) \ pdt\\ using wf_iff_no_infinite_down_chain by force hence *: \\ i. (f i, f (Suc i)) \ pdt\ by simp have **:\\ i. \ j>i. (f i, f j) \ pdt\ proof(rule,rule,rule) fix i j assume \i < (j::nat)\ thus \(f i, f j) \ pdt\ proof (induction \j\ rule: less_induct) case (less k) show \?case\ proof (cases \Suc i < k\) case True hence k:\k-1 < k\ \i < k-1\ and sk: \Suc (k-1) = k\ by auto show \?thesis\ using less(1)[OF k] *[rule_format,of \k-1\,unfolded sk] trans_pdt[unfolded trans_def] by blast next case False hence \Suc i = k\ using less(2) by auto then show \?thesis\ using * by auto qed qed qed hence ***:\\ i. \ j > i. f j pd\ f i\ \\ i. \ j > i. f i \ f j\ unfolding pdt_def by auto hence ****:\\ i>0. f i pd\ f 0\ by simp hence \f 0 \ nodes\ using * is_pd_def by fastforce then obtain \ n where \:\is_path \\ \\ 0 = f 0\ \\ n = return\ using reaching_ret by blast hence \\ i>0. \ k\n. \ k = f i\ using ***(1) \f 0 \ nodes\ unfolding is_pd_def by blast hence \f:\\ i. \ k\n. \ k = f i\ using \(2) by (metis le0 not_gr_zero) have \range f \ \ ` {..n}\ proof(rule subsetI) fix x assume \x \ range f\ then obtain i where \x = f i\ by auto then obtain k where \x = \ k\ \k \ n\ using \f by metis thus \x \ \ ` {..n}\ by simp qed hence f:\finite (range f)\ using finite_surj by auto hence fi:\\ i. infinite {j. f j = f i}\ using pigeonhole_infinite[OF _ f] by auto obtain i where \infinite {j. f j = f i}\ using fi .. thus \False\ by (metis (mono_tags, lifting) "***"(2) bounded_nat_set_is_finite gt_ex mem_Collect_eq nat_neq_iff) qed lemma return_pd: assumes \x \ nodes\ shows \return pd\ x\ unfolding is_pd_def using assms by blast lemma pd_total: assumes xz: \x pd\ z\ and yz: \y pd\ z\ shows \x pd\ y \ y pd\x\ proof - obtain \ n where path: \is_path \\ and \0: \\ 0 = z\ and \n: \\ n = return\ using xz reaching_ret unfolding is_pd_def by force have *: \\ k\n. (\ k = x \ \ k = y)\ (is \\ k\n. ?P k\) using path \0 \n xz yz unfolding is_pd_def by auto obtain k where k: \k = (LEAST k. \ k = x \ \ k = y)\ by simp hence kn: \k\n\ and \k: \\ k = x \ \ k = y\ using LeastBI_ex[OF *] by auto note k_le = Least_le[where P = \?P\] show \?thesis\ proof (cases) assume kx: \\ k = x\ have k_min: \\ k'. \ k' = y \ k \ k'\ using k_le unfolding k by auto { fix \' and n' :: \nat\ assume path': \is_path \'\ and \'0: \\' 0 = x\ and \'n': \\' n' = return\ have path'': \is_path (\ @\<^bsup>k\<^esup> \')\ using path_cons[OF path path'] kx \'0 by auto have \''0: \(\ @\<^bsup>k\<^esup> \') 0 = z\ using \0 by simp have \''n: \(\ @\<^bsup>k\<^esup> \') (k+n') = return\ using \'n' kx \'0 by auto obtain k' where k': \k' \ k + n'\ \(\ @\<^bsup>k\<^esup> \') k' = y\ using yz path'' \''0 \''n unfolding is_pd_def by blast have **: \k \ k'\ proof (rule ccontr) assume \\ k \ k'\ hence \k' < k\ by simp moreover hence \\ k' = y\ using k' by auto ultimately show \False\ using k_min by force qed hence \\' (k' - k) = y\ using k' \'0 kx by auto moreover have \(k' - k) \ n'\ using k' by auto ultimately have \\ k\ n'. \' k = y\ by auto } hence \y pd\ x\ using kx path_nodes path unfolding is_pd_def by auto thus \?thesis\ .. next \ \This is analogous argument\ assume kx: \\ k \ x\ hence ky: \\ k = y\ using \k by auto have k_min: \\ k'. \ k' = x \ k \ k'\ using k_le unfolding k by auto { fix \' and n' :: \nat\ assume path': \is_path \'\ and \'0: \\' 0 = y\ and \'n': \\' n' = return\ have path'': \is_path (\ @\<^bsup>k\<^esup> \')\ using path_cons[OF path path'] ky \'0 by auto have \''0: \(\ @\<^bsup>k\<^esup> \') 0 = z\ using \0 by simp have \''n: \(\ @\<^bsup>k\<^esup> \') (k+n') = return\ using \'n' ky \'0 by auto obtain k' where k': \k' \ k + n'\ \(\ @\<^bsup>k\<^esup> \') k' = x\ using xz path'' \''0 \''n unfolding is_pd_def by blast have **: \k \ k'\ proof (rule ccontr) assume \\ k \ k'\ hence \k' < k\ by simp moreover hence \\ k' = x\ using k' by auto ultimately show \False\ using k_min by force qed hence \\' (k' - k) = x\ using k' \'0 ky by auto moreover have \(k' - k) \ n'\ using k' by auto ultimately have \\ k\ n'. \' k = x\ by auto } hence \x pd\ y\ using ky path_nodes path unfolding is_pd_def by auto thus \?thesis\ .. qed qed lemma pds_finite: \finite {y . (x,y) \ pdt}\ proof cases assume \x \ nodes\ then obtain \ n where \:\is_path \\ \\ 0 = x\ \\ n = return\ using reaching_ret by blast have *: \\ y \ {y. (x,y)\ pdt}. y pd\ x\ using pdt_def by auto have \\ y \ {y. (x,y)\ pdt}. \ k \ n. \ k = y\ using * \ is_pd_def by blast hence \{y. (x,y)\ pdt} \ \ ` {..n}\ by auto then show \?thesis\ using finite_surj by blast next assume \\ x\ nodes\ hence \{y. (x,y)\pdt} = {}\ unfolding pdt_def is_pd_def using path_nodes reaching_ret by fastforce then show \?thesis\ by simp qed lemma ipd_exists: assumes node: \x \ nodes\ and not_ret: \x\return\ shows \\y. y ipd\ x\ proof - let \?Q\ = \{y. x\y \ y pd\ x}\ have *: \return \ ?Q\ using assms return_pd by simp hence **: \\ x. x\ ?Q\ by auto have fin: \finite ?Q\ using pds_finite unfolding pdt_def by auto have tot: \\ y z. y\?Q \ z \ ?Q \ z pd\ y \ y pd\ z\ using pd_total by auto obtain y where ymax: \y\ ?Q\ \\ z\?Q. z = y \ z pd\ y\ using fin ** tot proof (induct) case empty then show \?case\ by auto next case (insert x F) show \thesis\ proof (cases \F = {}\) assume \F = {}\ thus \thesis\ using insert(4)[of \x\] by auto next assume \F \ {}\ hence \\ x. x\ F\ by auto have \\y. y \ F \ \z\F. z = y \ z pd\ y \ thesis\ proof - fix y assume a: \y \ F\ \\z\F. z = y \ z pd\ y\ have \x \ y\ using insert a by auto have \x pd\ y \ y pd\ x\ using insert(6) a(1) by auto thus \thesis\ proof assume \x pd\ y\ hence \\z\insert x F. z = y \ z pd\ y\ using a(2) by blast thus \thesis\ using a(1) insert(4) by blast next assume \y pd\ x\ have \\z\insert x F. z = x \ z pd\ x\ proof fix z assume \z\ insert x F\ thus \z = x \ z pd\ x\ proof(rule,simp) assume \z\F\ hence \z = y \ z pd\ y\ using a(2) by auto thus \z = x \ z pd\ x\ proof(rule,simp add: \y pd\ x\) assume \z pd\ y\ show \z = x \ z pd\ x\ using \y pd\ x\ \z pd\ y\ pd_trans by blast qed qed qed then show \thesis\ using insert by blast qed qed then show \thesis\ using insert by blast qed qed hence ***: \y pd\ x\ \x\y\ by auto have \\ z. z \ x \ z pd\ x \ z pd\ y\ proof (rule,rule) fix z assume a: \ z \ x \ z pd\ x\ hence b: \z \ ?Q\ by auto have \y pd\ z \ z pd\ y\ using pd_total ***(1) a by auto thus \z pd\ y\ proof assume c: \y pd\ z\ hence \y = z\ using b ymax pdt_def pd_antisym by auto thus \z pd\ y\ using c by simp qed simp qed with *** have \y ipd\ x\ unfolding is_ipd_def by simp thus \?thesis\ by blast qed lemma ipd_unique: assumes yipd: \y ipd\ x\ and y'ipd: \y' ipd\ x\ shows \y = y'\ proof - have 1: \y pd\ y'\ and 2: \y' pd\ y\ using yipd y'ipd unfolding is_ipd_def by auto show \?thesis\ using pd_antisym[OF 1 2] . qed lemma ipd_is_ipd: assumes \x \ nodes\ and \x\return\ shows \ipd x ipd\ x\ proof - from assms obtain y where \y ipd\ x\ using ipd_exists by auto moreover hence \\ z. z ipd\x \ z = y\ using ipd_unique by simp ultimately show \?thesis\ unfolding ipd_def by (auto intro: theI2) qed lemma is_ipd_in_pdt: \y ipd\ x \ (x,y) \ pdt\ unfolding is_ipd_def pdt_def by auto lemma ipd_in_pdt: \x \ nodes \ x\return \ (x,ipd x) \ pdt\ by (metis ipd_is_ipd is_ipd_in_pdt) lemma no_pd_path: assumes \x \ nodes\ and \\ y pd\ x\ obtains \ n where \is_path \\ and \\ 0 = x\ and \\ n = return\ and \\ k \ n. \ k \ y\ proof (rule ccontr) assume \\ thesis\ hence \\ \ n. is_path \ \ \ 0 = x \ \ n = return \ (\ k\n . \ k = y)\ using that by force thus \False\ using assms unfolding is_pd_def by auto qed lemma pd_pd_ipd: assumes \x \ nodes\ \x\return\ \y\x\ \y pd\ x\ shows \y pd\ ipd x\ proof - have \ipd x pd\ x\ by (metis assms(1,2) ipd_is_ipd is_ipd_def) hence \y pd\ ipd x \ ipd x pd\ y\ by (metis assms(4) pd_total) thus \?thesis\ proof have 1: \ipd x ipd\ x\ by (metis assms(1,2) ipd_is_ipd) moreover assume \ipd x pd\ y\ ultimately show \y pd\ ipd x\ unfolding is_ipd_def using assms(3,4) by auto qed auto qed lemma pd_nodes: assumes \y pd\ x\ shows pd_node1: \y \ nodes\ and pd_node2: \x \ nodes\ proof - obtain \ k where \is_path \\ \\ k = y\ using assms unfolding is_pd_def using reaching_ret by force thus \y \ nodes\ using path_nodes by auto show \x \ nodes\ using assms unfolding is_pd_def by simp qed lemma pd_ret_is_ret: \x pd\ return \ x = return\ by (metis pd_antisym pd_node1 return_pd) lemma ret_path_none_pd: assumes \x \ nodes\ \x\return\ obtains \ n where \is_path \\ \\ 0 = x\ \\ n = return\ \\ i>0. \ x pd\ \ i\ proof(rule ccontr) assume \\thesis\ hence *: \\ \ n. \is_path \; \ 0 = x; \ n = return\ \ \i>0. x pd\ \ i\ using that by blast obtain \ n where **: \is_path \\ \\ 0 = x\ \\ n = return\ \\ i>0. \ i \ x\ using direct_path_return[OF assms] by metis then obtain i where ***: \i>0\ \x pd\ \ i\ using * by blast hence \\ i \ return\ using pd_ret_is_ret assms(2) by auto hence \i < n\ using assms(2) term_path_stable ** by (metis linorder_neqE_nat less_imp_le) hence \(\\i)(n-i) = return\ using **(3) by auto moreover have \(\\i) (0) = \ i\ by simp moreover have \is_path (\\i)\ using **(1) path_path_shift by metis ultimately obtain k where \(\\i) k = x\ using ***(2) unfolding is_pd_def by metis hence \\ (i + k) = x\ by auto thus \False\ using **(4) \i>0\ by auto qed lemma path_pd_ipd0': assumes \is_path \\ and \\ n \ return\ \\ n \ \ 0\ and \\ n pd\ \ 0\ obtains k where \k \ n\ and \\ k = ipd(\ 0)\ proof(rule ccontr) have *: \\ n pd\ ipd (\ 0)\ by (metis is_pd_def assms(3,4) pd_pd_ipd pd_ret_is_ret) obtain \' n' where **: \is_path \'\ \\' 0 = \ n\ \\' n' = return\ \\ i>0. \ \ n pd\ \' i\ by (metis assms(2) assms(4) pd_node1 ret_path_none_pd) hence \\ i>0. \' i \ ipd (\ 0)\ using * by metis moreover assume \\ thesis\ hence \\ k\n. \ k \ ipd (\ 0)\ using that by blast ultimately have \\ i. (\@\<^bsup>n\<^esup> \') i \ ipd (\ 0)\ by (metis diff_is_0_eq neq0_conv path_append_def) moreover have \(\@\<^bsup>n\<^esup> \') (n + n') = return\ by (metis \\' 0 = \ n\ \\' n' = return\ add_diff_cancel_left' assms(2) diff_is_0_eq path_append_def) moreover have \(\@\<^bsup>n\<^esup> \') 0 = \ 0\ by (metis le0 path_append_def) moreover have \is_path (\@\<^bsup>n\<^esup> \')\ by (metis \\' 0 = \ n\ \is_path \'\ assms(1) path_cons) moreover have \ipd (\ 0) pd\ \ 0\ by (metis **(2,3,4) assms(2) assms(4) ipd_is_ipd is_ipd_def neq0_conv pd_node2) moreover have \\ 0 \ nodes\ by (metis assms(1) path_nodes) ultimately show \False\ unfolding is_pd_def by blast qed lemma path_pd_ipd0: assumes \is_path \\ and \\ 0 \ return\ \\ n \ \ 0\ and \\ n pd\ \ 0\ obtains k where \k \ n\ and \\ k = ipd(\ 0)\ proof cases assume *: \\ n = return\ have \ipd (\ 0) pd\ (\ 0)\ by (metis is_ipd_def is_pd_def assms(2,4) ipd_is_ipd) with assms(1,2,3) * show \thesis\ unfolding is_pd_def by (metis that) next assume \\ n \ return\ from path_pd_ipd0' [OF assms(1) this assms(3,4)] that show \thesis\ by auto qed lemma path_pd_ipd: assumes \is_path \\ and \\ k \ return\ \\ n \ \ k\ and \\ n pd\ \ k\ and kn: \k < n\ obtains l where \k < l\ and \l \ n\ and \\ l = ipd(\ k)\ proof - have \is_path (\ \ k)\ \(\ \ k) 0 \ return\ \(\ \ k) (n - k) \ (\ \ k) 0\ \(\ \ k) (n - k) pd\ (\ \ k) 0\ using assms path_path_shift by auto with path_pd_ipd0[of \\\k\ \n-k\] obtain ka where \ka \ n - k\ \(\ \ k) ka = ipd ((\ \ k) 0)\ . hence \k + ka \ n\ \\ (k + ka) = ipd (\ k)\ using kn by auto moreover hence \\ (k + ka) ipd\ \ k\ by (metis assms(1) assms(2) ipd_is_ipd path_nodes) hence \k < k + ka\ unfolding is_ipd_def by (metis nat_neq_iff not_add_less1) ultimately show \thesis\ using that[of \k+ka\] by auto qed lemma path_ret_ipd: assumes \is_path \\ and \\ k \ return\ \\ n = return\ obtains l where \k < l\ and \l \ n\ and \\ l = ipd(\ k)\ proof - have \\ n \ \ k\ using assms by auto moreover have \k \ n\ apply (rule ccontr) using term_path_stable assms by auto hence \k < n\ by (metis assms(2,3) dual_order.order_iff_strict) moreover have \\ n pd\ \ k\ by (metis assms(1,3) path_nodes return_pd) ultimately obtain l where \k < l\ \l \ n\ \\ l = ipd (\ k)\ using assms path_pd_ipd by blast thus \thesis\ using that by auto qed lemma pd_intro: assumes \l pd\ k\ \is_path \\ \\ 0 = k\ \\ n = return\ obtains i where \i \ n\ \\ i = l\ using assms unfolding is_pd_def by metis lemma path_pd_pd0: assumes path: \is_path \\ and lpdn: \\ l pd\ n\ and npd0: \n pd\ \ 0\ obtains k where \k \ l\ \\ k = n\ proof (rule ccontr) assume \\ thesis\ hence notn: \\ k. k \ l \ \ k \ n\ using that by blast have nret: \\ l \ return\ by (metis is_pd_def assms(1,3) notn) obtain \' n' where path': \is_path \'\ and \0': \\' 0 = \ l\ and \n': \\' n' = return\ and nonepd: \\ i>0. \ \ l pd\ \' i\ using nret path path_nodes ret_path_none_pd by metis have \\ l \ n\ using notn by simp hence \\ i. \' i \ n\ using nonepd \0' lpdn by (metis neq0_conv) hence notn': \\ i. (\@\<^bsup>l\<^esup> \') i \ n\ using notn \0' by auto have \is_path (\@\<^bsup>l\<^esup> \')\ using path path' by (metis \0' path_cons) moreover have \(\@\<^bsup>l\<^esup> \') 0 = \ 0\ by simp moreover have \(\@\<^bsup>l\<^esup> \') (n' + l) = return\ using \0' \n' by auto ultimately show \False\ using notn' npd0 unfolding is_pd_def by blast qed subsection \Facts about Control Dependencies\ lemma icd_imp_cd: \n icd\<^bsup>\\<^esup>\ k \ n cd\<^bsup>\\<^esup>\ k\ by (metis is_icdi_def) lemma ipd_impl_not_cd: assumes \j \ {k..i}\ and \\ j = ipd (\ k)\ shows \\ i cd\<^bsup>\\<^esup>\ k\ by (metis assms(1) assms(2) is_cdi_def) lemma cd_not_ret: assumes \i cd\<^bsup>\\<^esup>\ k \ shows \\ k \ return\ by (metis is_cdi_def assms nat_less_le term_path_stable) lemma cd_path_shift: assumes \j \ k\ \is_path \ \ shows \(i cd\<^bsup>\\<^esup>\ k) = (i - j cd\<^bsup>\\j\<^esup>\ k-j)\ proof assume a: \i cd\<^bsup>\\<^esup>\ k\ hence b: \k < i\ by (metis is_cdi_def) hence \is_path (\ \ j)\ \k - j < i - j\ using assms apply (metis path_path_shift) by (metis assms(1) b diff_less_mono) moreover have c: \\ j \ {k..i}. \ j \ ipd (\ k)\ by (metis a ipd_impl_not_cd) hence \\ ja \ {k - j..i - j}. (\ \ j) ja \ ipd ((\ \ j) (k - j))\ using b assms by auto fastforce moreover have \j < i\ using assms(1) b by auto hence \(\\j) (i - j) \ return\ using a unfolding is_cdi_def by auto ultimately show \i - j cd\<^bsup>\\j\<^esup>\ k-j\ unfolding is_cdi_def by simp next assume a: \i - j cd\<^bsup>\\j\<^esup>\ k-j\ hence b: \k - j < i-j\ by (metis is_cdi_def) moreover have c: \\ ja \ {k - j..i - j}. (\ \ j) ja \ ipd ((\ \ j) (k - j))\ by (metis a ipd_impl_not_cd) have \\ j \ {k..i}. \ j \ ipd (\ k)\ proof (rule,goal_cases) case (1 n) hence \n-j \ {k-j..i-j}\ using assms by auto hence \\ (j + (n-j)) \ ipd(\ (j + (k-j)))\ by (metis c path_shift_def) thus \?case\ using 1 assms(1) by auto qed moreover have \j < i\ using assms(1) b by auto hence \\ i \ return\ using a unfolding is_cdi_def by auto ultimately show \i cd\<^bsup>\\<^esup>\k\ unfolding is_cdi_def by (metis assms(1) assms(2) diff_is_0_eq' le_diff_iff nat_le_linear nat_less_le) qed lemma cd_path_shift0: assumes \is_path \\ shows \(i cd\<^bsup>\\<^esup>\ k) = (i-k cd\<^bsup>\\k\<^esup>\0)\ using cd_path_shift[OF _ assms] by (metis diff_self_eq_0 le_refl) lemma icd_path_shift: assumes \l \ k\ \is_path \\ shows \(i icd\<^bsup>\\<^esup>\ k) = (i - l icd\<^bsup>\\l\<^esup>\ k - l)\ proof - have \is_path (\\l)\ using path_path_shift assms(2) by auto moreover have \(i cd\<^bsup>\\<^esup>\ k) = (i - l cd\<^bsup>\\l\<^esup>\ k - l)\ using assms cd_path_shift by auto moreover have \(\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m) = (\ m \ {k - l<.. i - l cd\<^bsup>\ \ l\<^esup>\ m)\ proof - {fix m assume *: \\ m \ {k - l<.. i - l cd\<^bsup>\ \ l\<^esup>\ m\ \m \ {k<.. hence \m-l \ {k-l<.. using assms(1) by auto hence \\ i - l cd\<^bsup>\\l\<^esup>\(m-l)\ using * by blast moreover have \l \ m\ using * assms by auto ultimately have \\ i cd\<^bsup>\\<^esup>\m\ using assms(2) cd_path_shift by blast } moreover {fix m assume *: \\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m\ \m-l \ {k-l<.. hence \m \ {k<.. using assms(1) by auto hence \\ i cd\<^bsup>\\<^esup>\m\ using * by blast moreover have \l \ m\ using * assms by auto ultimately have \\ i - l cd\<^bsup>\\l\<^esup>\(m-l)\ using assms(2) cd_path_shift by blast } ultimately show \?thesis\ by auto (metis diff_add_inverse) qed ultimately show \?thesis\ unfolding is_icdi_def using assms by blast qed lemma icd_path_shift0: assumes \is_path \\ shows \(i icd\<^bsup>\\<^esup>\ k) = (i-k icd\<^bsup>\\k\<^esup>\0)\ using icd_path_shift[OF _ assms] by (metis diff_self_eq_0 le_refl) lemma cdi_path_swap: assumes \is_path \'\ \j cd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>j\<^esub> \'\ shows \j cd\<^bsup>\'\<^esup>\k\ using assms unfolding eq_up_to_def is_cdi_def by auto lemma cdi_path_swap_le: assumes \is_path \'\ \j cd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>n\<^esub> \'\ \j \ n\ shows \j cd\<^bsup>\'\<^esup>\k\ by (metis assms cdi_path_swap eq_up_to_le) lemma not_cd_impl_ipd: assumes \is_path \\ and \k < i\ and \\ i cd\<^bsup>\\<^esup>\ k\ and \\ i \ return\ obtains j where \j \ {k..i}\ and \\ j = ipd (\ k)\ by (metis assms(1) assms(2) assms(3) assms(4) is_cdi_def) lemma icd_is_the_icd: assumes \i icd\<^bsup>\\<^esup>\ k\ shows \k = (THE k. i icd\<^bsup>\\<^esup>\ k)\ using assms icd_uniq by (metis the1_equality) lemma all_ipd_imp_ret: assumes \is_path \\ and \\ i. \ i \ return \ (\ j>i. \ j = ipd (\ i))\ shows \\j. \ j = return\ proof - { fix x assume *: \\ 0 = x\ have \?thesis\ using wf_pdt_inv * assms proof(induction \x\ arbitrary: \\\ rule: wf_induct_rule ) case (less x \) show \?case\ proof (cases \x = return\) case True thus \?thesis\ using less(2) by auto next assume not_ret: \x \ return\ moreover then obtain k where k_ipd: \\ k = ipd x\ using less(2,4) by auto moreover have \x \ nodes\ using less(2,3) by (metis path_nodes) ultimately have \(x, \ k) \ pdt\ by (metis ipd_in_pdt) hence a: \(\ k, x) \ pdt_inv\ unfolding pdt_inv_def by simp have b: \is_path (\ \ k)\ by (metis less.prems(2) path_path_shift) have c: \\ i. (\\k) i \ return \ (\j>i. (\\k) j = ipd ((\\k) i))\ using less(4) apply auto by (metis (full_types) ab_semigroup_add_class.add_ac(1) less_add_same_cancel1 less_imp_add_positive) from less(1)[OF a _ b c] have \\j. (\\k) j = return\ by auto thus \\j. \ j = return\ by auto qed qed } thus \?thesis\ by simp qed lemma loop_has_cd: assumes \is_path \\ \0 < i\ \\ i = \ 0\ \\ 0 \ return\ shows \\ k < i. i cd\<^bsup>\\<^esup>\ k\ proof (rule ccontr) let \?\\ = \(\ n. \ (n mod i))\ assume \\ (\k\\<^esup>\ k)\ hence \\ k i cd\<^bsup>\\<^esup>\ k\ by blast hence *: \\ kj \ {k..i}. \ j = ipd (\ k))\ using assms(1,3,4) not_cd_impl_ipd by metis have \\ k. (\ j > k. ?\ j = ipd (?\ k))\ proof fix k have \k mod i < i\ using assms(2) by auto with * obtain j where \j \ {(k mod i)..i}\ \\ j = ipd (\ (k mod i))\ by auto then obtain j' where 1: \j' < i\ \\ j' = ipd (\ (k mod i))\ by (cases \j = i\, auto ,metis assms(2) assms(3),metis le_neq_implies_less) then obtain j'' where 2: \j'' > k\ \j'' mod i = j'\ by (metis mod_bound_instance) hence \?\ j'' = ipd (?\ k)\ using 1 by auto with 2(1) show \\ j > k. ?\ j = ipd (?\ k)\ by auto qed moreover have \is_path ?\\ by (metis assms(1) assms(2) assms(3) is_path_loop) ultimately obtain k where \?\ k = return\ by (metis (lifting) all_ipd_imp_ret) moreover have \k mod i < i\ by (simp add: assms(2)) ultimately have \\ i = return\ by (metis assms(1) term_path_stable less_imp_le) thus \False\ by (metis assms(3) assms(4)) qed lemma loop_has_cd': assumes \is_path \\ \j < i\ \\ i = \ j\ \\ j \ return\ shows \\ k \ {j..\\<^esup>\ k\ proof - have \\ k'< i-j. i-j cd\<^bsup>\\j\<^esup>\k'\ apply(rule loop_has_cd) apply (metis assms(1) path_path_shift) apply (auto simp add: assms less_imp_le) done then obtain k where k: \k \i-j cd\<^bsup>\\j\<^esup>\k\ by auto hence k': \(k+j) < i\ \i-j cd\<^bsup>\\j\<^esup>\ (k+j)-j\ by auto note cd_path_shift[OF _ assms(1)] hence \i cd\<^bsup>\\<^esup>\ k+j\ using k'(2) by (metis le_add1 add.commute) with k'(1) show \?thesis\ by force qed lemma claim'': assumes path\: \is_path \\ and path\': \is_path \'\ and \i: \\ i = \' i'\ and \j: \\ j = \' j'\ and not_cd: \\ k. \ j cd\<^bsup>\\<^esup>\ k\ \\ k. \ i' cd\<^bsup>\'\<^esup>\ k\ and nret: \\ i \ return\ and ilj: \i < j\ shows \i' < j'\ proof (rule ccontr) assume \\ i' < j'\ hence jlei: \j' \ i'\ by auto show \False\ proof (cases) assume j'li': \j' < i'\ define \'' where \\'' \ (\@\<^bsup>j\<^esup>(\'\j'))\i\ note \''_def[simp] have \\ j = (\' \ j') 0\ by (metis path_shift_def Nat.add_0_right \j) hence \is_path \''\ using path\ path\' \''_def path_path_shift path_cons by presburger moreover have \\'' (j-i+(i'-j')) = \'' 0\ using ilj jlei \i \j by (auto, metis add_diff_cancel_left' le_antisym le_diff_conv le_eq_less_or_eq) moreover have \\'' 0 \ return\ by (simp add: ilj less_or_eq_imp_le nret) moreover have \0 < j-i+(i'-j')\ by (metis add_is_0 ilj neq0_conv zero_less_diff) ultimately obtain k where k: \k < j-i+(i'-j')\ \j-i+(i'-j') cd\<^bsup>\''\<^esup>\ k\ by (metis loop_has_cd) hence *: \\ l \ {k..j-i+(i'-j')}. \'' l \ ipd (\'' k)\ by (metis is_cdi_def) show \False\ proof (cases \k < j-i\) assume a: \k < j - i\ hence b: \\'' k = \ (i + k)\ by auto have \\ l \ {i+k..j}. \ l \ ipd (\ (i+k))\ proof fix l assume l: \l \ {i + k..j}\ hence \\ l = \'' (l - i)\ by auto moreover from a l have \l-i \ {k .. j-i + (i'-j')}\ by force ultimately show \\ l \ ipd (\ (i + k))\ using * b by auto qed moreover have \i + k < j\ using a by simp moreover have \\ j \ return\ by (metis \i \j j'li' nret path\' term_path_stable less_imp_le) ultimately have \j cd\<^bsup>\\<^esup>\ i+k\ by (metis not_cd_impl_ipd path\) thus \False\ by (metis not_cd(1)) next assume \\ k < j - i\ hence a: \j - i \ k\ by simp hence b: \\'' k = \' (j' + (i + k) - j)\ unfolding \''_def path_shift_def path_append_def using ilj by(auto,metis \j add_diff_cancel_left' le_antisym le_diff_conv add.commute) have \\ l \ {j' + (i+k) - j..i'}. \' l \ ipd (\' (j' + (i+k) - j))\ proof fix l assume l: \l \ {j' + (i+k) - j..i'}\ hence \\' l = \'' (j + l - i - j')\ unfolding \''_def path_shift_def path_append_def using ilj by (auto, metis Nat.diff_add_assoc \j a add.commute add_diff_cancel_left' add_leD1 le_antisym le_diff_conv) moreover from a l have \j + l - i - j' \ {k .. j-i + (i'-j')}\ by force ultimately show \\' l \ ipd (\' (j' + (i + k) - j))\ using * b by auto qed moreover have \j' + (i+k) - j < i'\ using a j'li' ilj k(1) by linarith moreover have \\' i' \ return\ by (metis \i nret) ultimately have \i' cd\<^bsup>\'\<^esup>\ j' + (i+k) - j\ by (metis not_cd_impl_ipd path\') thus \False\ by (metis not_cd(2)) qed next assume \\ j' < i'\ hence \j' = i'\ by (metis \\ i' < j'\ linorder_cases) hence \\ i = \ j\ by (metis \i \j) thus \False\ by (metis ilj loop_has_cd' not_cd(1) nret path\) qed qed lemma other_claim': assumes path: \is_path \\ and eq: \\ i = \ j\ and \\ i \ return\ and icd: \\ k. \ i cd\<^bsup>\\<^esup>\ k\ and \\ k. \ j cd\<^bsup>\\<^esup>\ k\ shows \i = j\ proof (rule ccontr,cases) assume \i < j\ thus \False\ using assms claim'' by blast next assume \\ i < j\ \i \ j\ hence \j < i\ by auto thus \False\ using assms claim'' by (metis loop_has_cd') qed lemma icd_no_cd_path_shift: assumes \i icd\<^bsup>\\<^esup>\ 0\ shows \(\ k. \ i - 1 cd\<^bsup>\\1\<^esup>\ k)\ proof (rule,rule ccontr,goal_cases) case (1 k) hence *: \i - 1 cd\<^bsup>\ \ 1\<^esup>\ k\ by simp have **: \1 \ k + 1\ by simp have ***: \is_path \\ by (metis assms is_icdi_def) hence \i cd\<^bsup>\\<^esup>\ k+1\ using cd_path_shift[OF ** ***] * by auto moreover hence \k+1 < i\ unfolding is_cdi_def by simp moreover have \0 < k + 1\ by simp ultimately show \False\ using assms[unfolded is_icdi_def] by auto qed lemma claim': assumes path\: \is_path \\ and path\': \is_path \'\ and \i: \\ i = \' i'\ and \j: \\ j = \' j'\ and not_cd: \i icd\<^bsup>\\<^esup>\ 0\ \j icd\<^bsup>\\<^esup>\ 0\ \i' icd\<^bsup>\'\<^esup>\ 0\ \j' icd\<^bsup>\'\<^esup>\ 0\ and ilj: \i < j\ and nret: \\ i \ return\ shows \i' < j'\ proof - have g0: \0 < i\ \0 < j\ \0 < i'\ \0 < j'\using not_cd[unfolded is_icdi_def is_cdi_def] by auto have \(\ \ 1) (i - 1) = (\' \ 1) (i' - 1)\ \(\ \ 1) (j - 1) = (\' \ 1) (j' - 1)\ using \i \j g0 by auto moreover have \\ k. \ (j - 1) cd\<^bsup>\\1\<^esup>\ k\ \\ k. \ (i' - 1) cd\<^bsup>\'\1\<^esup>\ k\ by (metis icd_no_cd_path_shift not_cd(2)) (metis icd_no_cd_path_shift not_cd(3)) moreover have \is_path (\\1)\ \is_path (\'\1)\ using path\ path\' path_path_shift by blast+ moreover have \(\\1) (i - 1) \ return\ using g0 nret by auto moreover have \i - 1 < j - 1\ using g0 ilj by auto ultimately have \i' - 1 < j' - 1\ using claim'' by blast thus \i' by auto qed lemma other_claim: assumes path: \is_path \\ and eq: \\ i = \ j\ and \\ i \ return\ and icd: \i icd\<^bsup>\\<^esup>\ 0\ and \j icd\<^bsup>\\<^esup>\ 0\ shows \i = j\ proof (rule ccontr,cases) assume \i < j\ thus \False\ using assms claim' by blast next assume \\ i < j\ \i \ j\ hence \j < i\ by auto thus \False\ using assms claim' by (metis less_not_refl) qed lemma cd_trans0: assumes \j cd\<^bsup>\\<^esup>\ 0\ and \k cd\<^bsup>\\<^esup>\j\ shows \k cd\<^bsup>\\<^esup>\ 0\ proof (rule ccontr) have path: \is_path \\ and ij: \0 < j\ and jk: \j < k\ and nret: \\ j \ return\ \\ k \ return\ and noipdi: \\ l \ {0..j}. \ l \ ipd (\ 0)\ and noipdj: \\ l \ {j..k}. \ l \ ipd (\ j)\ using assms unfolding is_cdi_def by auto assume \\ k cd\<^bsup>\\<^esup>\ 0\ hence \\l \ {0..k}. \ l = ipd (\ 0)\ unfolding is_cdi_def using path ij jk nret by force then obtain l where \l \ {0..k}\ and l: \\ l = ipd (\ 0)\ by auto hence jl: \j and lk: \l\k\ using noipdi ij by auto have pdj: \ipd (\ 0) pd\ \ j\ proof (rule ccontr) have \\ j \ nodes\ using path by (metis path_nodes) moreover assume \\ ipd (\ 0) pd\ \ j\ ultimately obtain \' n where *: \is_path \'\ \\' 0 = \ j\ \\' n = return\ \\ k\n. \' k \ ipd(\ 0)\ using no_pd_path by metis hence path': \is_path (\ @\<^bsup>j\<^esup> \')\ by (metis path path_cons) moreover have \\ k \ j + n. (\@\<^bsup>j\<^esup> \') k \ ipd (\ 0)\ using noipdi *(4) by auto moreover have \(\@\<^bsup>j\<^esup> \') 0 = \ 0\ by auto moreover have \(\@\<^bsup>j\<^esup> \') (j + n) = return\ using *(2,3) by auto ultimately have \\ ipd (\ 0) pd\ \ 0\ unfolding is_pd_def by metis thus \False\ by (metis is_ipd_def ij ipd_is_ipd nret(1) path path_nodes term_path_stable less_imp_le) qed hence \(\\j) (l-j) pd\ (\\j) 0\ using jl l by auto moreover have \is_path (\\j)\ by (metis path path_path_shift) moreover have \\ l \ return\ by (metis lk nret(2) path term_path_stable) hence \(\\j) (l-j) \ return\ using jl by auto moreover have \\ j \ ipd (\ 0)\ using noipdi by force hence \(\\j) (l-j) \ (\\j) 0\ using jl l by auto ultimately obtain k' where \k' \ l-j\ and \(\\j) k' = ipd ((\\j) 0)\ using path_pd_ipd0' by blast hence \j + k' \ {j..k}\ \\ (j+k') = ipd (\ j)\ using jl lk by auto thus \False\ using noipdj by auto qed lemma cd_trans: assumes \j cd\<^bsup>\\<^esup>\ i\ and \k cd\<^bsup>\\<^esup>\j\ shows \k cd\<^bsup>\\<^esup>\ i\ proof - have path: \is_path \\ using assms is_cdi_def by auto have ij: \i using assms is_cdi_def by auto let \?\\ = \\\i\ have \j-i cd\<^bsup>?\\<^esup>\ 0\ using assms(1) cd_path_shift0 path by auto moreover have \k-i cd\<^bsup>?\\<^esup>\j-i\ by (metis assms(2) cd_path_shift is_cdi_def ij less_imp_le_nat) ultimately have \k-i cd\<^bsup>?\\<^esup>\ 0\ using cd_trans0 by auto thus \k cd\<^bsup>\\<^esup>\ i\ using path cd_path_shift0 by auto qed lemma excd_impl_exicd: assumes \\ k. i cd\<^bsup>\\<^esup>\k\ shows \\ k. i icd\<^bsup>\\<^esup>\k\ using assms proof(induction \i\ arbitrary: \\\ rule: less_induct) case (less i) then obtain k where k: \i cd\<^bsup>\\<^esup>\k\ by auto hence ip: \is_path \\ unfolding is_cdi_def by auto show \?case\ proof (cases) assume *: \\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m\ hence \i icd\<^bsup>\\<^esup>\k\ using k ip unfolding is_icdi_def by auto thus \?case\ by auto next assume \\ (\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m)\ then obtain m where m: \m \ {k<.. \i cd\<^bsup>\\<^esup>\ m\ by blast hence \i - m cd\<^bsup>\\m\<^esup>\ 0\ by (metis cd_path_shift0 is_cdi_def) moreover have \i - m < i\ using m by auto ultimately obtain k' where k': \i - m icd\<^bsup>\\m\<^esup>\ k'\ using less(1) by blast hence \i icd\<^bsup>\\<^esup>\ k' + m\ using ip by (metis add.commute add_diff_cancel_right' icd_path_shift le_add1) thus \?case\ by auto qed qed lemma cd_split: assumes \i cd\<^bsup>\\<^esup>\ k\ and \\ i icd\<^bsup>\\<^esup>\ k\ obtains m where \i icd\<^bsup>\\<^esup>\ m\ and \m cd\<^bsup>\\<^esup>\ k\ proof - have ki: \k < i\ using assms is_cdi_def by auto obtain m where m: \i icd\<^bsup>\\<^esup>\ m\ using assms(1) by (metis excd_impl_exicd) hence \k \ m\ unfolding is_icdi_def using ki assms(1) by force hence km: \k < m\using m assms(2) by (metis le_eq_less_or_eq) moreover have \\ m \ return\ using m unfolding is_icdi_def is_cdi_def by (simp, metis term_path_stable less_imp_le) moreover have \m using m unfolding is_cdi_def is_icdi_def by auto ultimately have \m cd\<^bsup>\\<^esup>\ k\ using assms(1) unfolding is_cdi_def by auto with m that show \thesis\ by auto qed lemma cd_induct[consumes 1, case_names base IS]: assumes prem: \i cd\<^bsup>\\<^esup>\ k\ and base: \\ i. i icd\<^bsup>\\<^esup>\k \ P i\ and IH: \\ k' i'. k' cd\<^bsup>\\<^esup>\ k \ P k' \ i' icd\<^bsup>\\<^esup>\ k' \ P i'\ shows \P i\ using prem IH proof (induction \i\ rule: less_induct,cases) case (less i) assume \i icd\<^bsup>\\<^esup>\ k\ thus \P i\ using base by simp next case (less i') assume \\ i' icd\<^bsup>\\<^esup>\ k\ then obtain k' where k': \ i' icd\<^bsup>\\<^esup>\ k'\ \k' cd\<^bsup>\\<^esup>\ k\ using less cd_split by blast hence icdk: \i' cd\<^bsup>\\<^esup>\ k'\ using is_icdi_def by auto note ih=less(3)[OF k'(2) _ k'(1)] have ki: \k' < i'\ using k' is_icdi_def is_cdi_def by auto have \P k'\ using less(1)[OF ki k'(2) ] less(3) by auto thus \P i'\ using ih by simp qed lemma cdi_prefix: \n cd\<^bsup>\\<^esup>\ m \ m < n' \ n' \ n \ n' cd\<^bsup>\\<^esup>\ m\ unfolding is_cdi_def by (simp, metis term_path_stable) lemma cr_wn': assumes 1: \n cd\<^bsup>\\<^esup>\ m\ and nc: \\ m' cd\<^bsup>\\<^esup>\ m\ and 3: \m < m'\ shows \n < m'\ proof (rule ccontr) assume \\ n < m'\ hence \m' \ n\ by simp hence \m' cd\<^bsup>\\<^esup>\ m\ by (metis 1 3 cdi_prefix) thus \False\ using nc by simp qed lemma cr_wn'': assumes \i cd\<^bsup>\\<^esup>\ m\ and \j cd\<^bsup>\\<^esup>\ n\ and \\ m cd\<^bsup>\\<^esup>\ n\ and \i \ j\ shows \m \ n\ proof (rule ccontr) assume \\m\n\ hence nm: \n < m\ by auto moreover have \m using assms(1) assms(4) unfolding is_cdi_def by auto ultimately have \m cd\<^bsup>\\<^esup>\ n\ using assms(2) cdi_prefix by auto thus \False\ using assms(3) by auto qed lemma ret_no_cd: assumes \\ n = return\ shows \\ n cd\<^bsup>\\<^esup>\ k\ by (metis assms is_cdi_def) lemma ipd_not_self: assumes \x \ nodes\ \x\ return\ shows \x \ ipd x\ by (metis is_ipd_def assms ipd_is_ipd) lemma icd_cs: assumes \l icd\<^bsup>\\<^esup>\k\ shows \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ [\ l]\ proof - from assms have \k = (THE k. l icd\<^bsup>\\<^esup>\ k)\ by (metis icd_is_the_icd) with assms show \?thesis\ by auto qed lemma cd_not_pd: assumes \l cd\<^bsup>\\<^esup>\ k\ \\ l \ \ k\ shows \\ \ l pd\ \ k\ proof assume pd: \\ l pd\ \ k\ have nret: \\ k \ return\ by (metis assms(1) pd pd_ret_is_ret ret_no_cd) have kl: \k < l\ by (metis is_cdi_def assms(1)) have path: \is_path \\ by (metis is_cdi_def assms(1)) from path_pd_ipd[OF path nret assms(2) pd kl] obtain n where \k < n\ \n \ l\ \\ n = ipd (\ k)\ . thus \False\ using assms(1) unfolding is_cdi_def by auto qed lemma cd_ipd_is_cd: assumes \k \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ and mcdj: \m cd\<^bsup>\\<^esup>\ j\ shows \k cd\<^bsup>\\<^esup>\ j\ proof cases assume \j < k\ thus \k cd\<^bsup>\\<^esup>\ j\ by (metis mcdj assms(1) cdi_prefix less_imp_le_nat) next assume \\ j < k\ hence kj: \k \ j\ by simp have \k < j\ apply (rule ccontr) using kj assms mcdj by (auto, metis is_cdi_def is_ipd_def cd_not_pd ipd_is_ipd path_nodes term_path_stable less_imp_le) moreover have \j < m\ using mcdj is_cdi_def by auto hence \\ n \ {k..j}. \ n \ ipd(\ k)\ using assms(3) by force ultimately have \j cd\<^bsup>\\<^esup>\ k\ by (metis mcdj is_cdi_def term_path_stable less_imp_le) hence \m cd\<^bsup>\\<^esup>\ k\ by (metis mcdj cd_trans) hence \False\ by (metis is_cdi_def is_ipd_def assms(2) cd_not_pd ipd_is_ipd path_nodes term_path_stable less_imp_le) thus \?thesis\ by simp qed lemma ipd_pd_cd0: assumes lcd: \n cd\<^bsup>\\<^esup>\ 0\ shows \ipd (\ 0) pd\ (\ n)\ proof - obtain k l where \0: \\ 0 = k\ and \n: \\ n = l\ and cdi: \n cd\<^bsup>\\<^esup>\ 0\ using lcd unfolding is_cdi_def by blast have nret: \k \ return\ by (metis is_cdi_def \0 cdi term_path_stable less_imp_le) have path: \is_path \\ and ipd: \\ i\n. \ i \ ipd k\ using cdi unfolding is_cdi_def \0 by auto { fix \' n' assume path': \is_path \'\ and \'0: \\' 0 = l\ and ret: \\' n' = return\ have \is_path (\ @\<^bsup>n\<^esup> \')\ using path path' \n \'0 by (metis path_cons) moreover have \(\ @\<^bsup>n\<^esup> \') (n+n') = return\ using ret \n \'0 by auto moreover have \(\ @\<^bsup>n\<^esup> \') 0 = k\ using \0 by auto moreover have \ipd k pd\ k\ by (metis is_ipd_def path \0 ipd_is_ipd nret path_nodes) ultimately obtain k' where k': \k' \ n+n'\ \(\ @\<^bsup>n\<^esup> \') k' = ipd k\ by (metis pd_intro) have \\ k'\ n\ proof assume \k' \ n\ hence \(\ @\<^bsup>n\<^esup> \') k' = \ k'\ by auto thus \False\ using k'(2) ipd by (metis \k' \ n\) qed hence \(\ @\<^bsup>n\<^esup> \') k' = \' (k' - n)\ by auto moreover have \(k' - n) \ n'\ using k' by simp ultimately have \\ k'\n'. \' k' = ipd k\ unfolding k' by auto } moreover have \l \ nodes\ by (metis \n path path_nodes) ultimately show \ipd (\ 0) pd\ (\ n)\ unfolding is_pd_def by (simp add: \0 \n) qed lemma ipd_pd_cd: assumes lcd: \l cd\<^bsup>\\<^esup>\ k\ shows \ipd (\ k) pd\ (\ l)\ proof - have \l-k cd\<^bsup>\\k\<^esup>\0\ using lcd cd_path_shift0 is_cdi_def by blast moreover note ipd_pd_cd0[OF this] moreover have \(\ \ k) 0 = \ k\ by auto moreover have \k < l\ using lcd unfolding is_cdi_def by simp then have \(\ \ k) (l - k) = \ l\ by simp ultimately show \?thesis\ by simp qed lemma cd_is_cd_ipd: assumes km: \k and ipd: \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ and cdj: \k cd\<^bsup>\\<^esup>\ j\ and nipdj: \ipd (\ j) \ \ m\ shows \m cd\<^bsup>\\<^esup>\ j\ proof - have path: \is_path \\ and jk: \j < k\ and nretj: \\ k \ return\ and nipd: \\ l \ {j..k}. \ l \ ipd (\ j)\ using cdj is_cdi_def by auto have pd: \ipd (\ j) pd\ \ m\ by (metis atLeastAtMost_iff cdj ipd(1) ipd_pd_cd jk le_refl less_imp_le nipd nretj path path_nodes pd_pd_ipd) have nretm: \\ m \ return\ by (metis nipdj pd pd_ret_is_ret) have jm: \j < m\ using jk km by simp show \m cd\<^bsup>\\<^esup>\ j\ proof (rule ccontr) assume ncdj: \\ m cd\<^bsup>\\<^esup>\ j\ hence \\ l \ {j..m}. \ l = ipd (\ j)\ unfolding is_cdi_def by (metis jm nretm path) then obtain l where jl: \j \ l\ and \l \ m\ and lipd: \\ l = ipd (\ j)\ by force hence lm: \l < m\ using nipdj by (metis le_eq_less_or_eq) have npd: \\ ipd (\ k) pd\ \ l\ by (metis ipd(1) lipd nipdj pd pd_antisym) have nd: \\ l \ nodes\ using path path_nodes by simp from no_pd_path[OF nd npd] obtain \' n where path': \is_path \'\ and \'0: \\' 0 = \ l\ and \'n: \\' n = return\ and nipd: \\ ka\n. \' ka \ ipd (\ k)\ . let \?\\ = \(\@\<^bsup>l\<^esup> \') \ k\ have path'': \is_path ?\\ by (metis \'0 path path' path_cons path_path_shift) moreover have kl: \k < l\ using lipd cdj jl unfolding is_cdi_def by fastforce have \?\ 0 = \ k\ using kl by auto moreover have \?\ (l + n - k) = return\ using \'n \'0 kl by auto moreover have \ipd (\ k) pd\ \ k\ by (metis is_ipd_def ipd_is_ipd nretj path path_nodes) ultimately obtain l' where l': \l' \ (l + n - k)\ \?\ l' = ipd (\ k)\ unfolding is_pd_def by blast show \False\ proof (cases ) assume *: \k + l' \ l\ hence \\ (k + l') = ipd (\ k)\ using l' by auto moreover have \k + l' < m\ by (metis "*" dual_order.strict_trans2 lm) ultimately show \False\ using ipd(2) by simp next assume \\ k + l' \ l\ hence \\' (k + l' - l) = ipd (\ k)\ using l' by auto moreover have \k + l' - l \ n\ using l' kl by linarith ultimately show \False\ using nipd by auto qed qed qed lemma ipd_icd_greatest_cd_not_ipd: assumes ipd: \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ and km: \k < m\ and icdj: \m icd\<^bsup>\\<^esup>\ j\ shows \j = (GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m)\ proof - let \?j\ = \GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m\ have kcdj: \k cd\<^bsup>\\<^esup>\ j\ using assms cd_ipd_is_cd is_icdi_def by blast have nipd: \ipd (\ j) \ \ m\ using icdj unfolding is_icdi_def is_cdi_def by auto have bound: \\ j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m \ j \ k\ unfolding is_cdi_def by simp have exists: \k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m\ (is \?P j\) using kcdj nipd by auto note GreatestI_nat[of \?P\ _ \k\, OF exists] Greatest_le_nat[of \?P\ \j\ \k\, OF exists] hence kcdj': \k cd\<^bsup>\\<^esup>\ ?j\ and ipd': \ipd (\ ?j) \ \ m\ and jj: \j \ ?j\ using bound by auto hence mcdj': \m cd\<^bsup>\\<^esup>\ ?j\ using ipd km cd_is_cd_ipd by auto show \j = ?j\ proof (rule ccontr) assume \j \ ?j\ hence jlj: \j < ?j\ using jj by simp moreover have \?j < m\ using kcdj' km unfolding is_cdi_def by auto ultimately show \False\ using icdj mcdj' unfolding is_icdi_def by auto qed qed lemma cd_impl_icd_cd: assumes \i cd\<^bsup>\\<^esup>\ l\ and \i icd\<^bsup>\\<^esup>\ k\ and \\ i icd\<^bsup>\\<^esup>\ l\ shows \k cd\<^bsup>\\<^esup>\ l\ using assms cd_split icd_uniq by metis lemma cdi_is_cd_icdi: assumes \k icd\<^bsup>\\<^esup>\ j\ shows \k cd\<^bsup>\\<^esup>\ i \ j cd\<^bsup>\\<^esup>\ i \ i = j\ by (metis assms cd_impl_icd_cd cd_trans icd_imp_cd icd_uniq) lemma same_ipd_stable: assumes \k cd\<^bsup>\\<^esup>\ i\ \k cd\<^bsup>\\<^esup>\ j\ \i \ipd (\ i) = ipd (\ k)\ shows \ipd (\ j) = ipd (\ k)\ proof - have jcdi: \j cd\<^bsup>\\<^esup>\ i\ by (metis is_cdi_def assms(1,2,3) cr_wn' le_antisym less_imp_le_nat) have 1: \ipd (\ j) pd\ \ k \ by (metis assms(2) ipd_pd_cd) have 2: \ipd (\ k) pd\ \ j \ by (metis assms(4) ipd_pd_cd jcdi) have 3: \ipd (\ k) pd\ (ipd (\ j))\ by (metis 2 IFC_def.is_cdi_def assms(1,2,4) atLeastAtMost_iff jcdi less_imp_le pd_node2 pd_pd_ipd) have 4: \ipd (\ j) pd\ (ipd (\ k))\ by (metis 1 2 IFC_def.is_ipd_def assms(2) cd_not_pd ipd_is_ipd jcdi pd_node2 ret_no_cd) show \?thesis\ using 3 4 pd_antisym by simp qed lemma icd_pd_intermediate': assumes icd: \i icd\<^bsup>\\<^esup>\ k\ and j: \k < j\ \j < i\ shows \\ i pd\ (\ j)\ using j proof (induction \i - j\ arbitrary: \j\ rule: less_induct) case (less j) have \\ i cd\<^bsup>\\<^esup>\ j\ using less.prems icd unfolding is_icdi_def by force moreover have \is_path \\ using icd by (metis is_icdi_def) moreover have \\ i \ return\ using icd by (metis is_icdi_def ret_no_cd) ultimately have \\ l. j \ l \ l \ i \ \ l = ipd (\ j)\ unfolding is_cdi_def using less.prems by auto then obtain l where l: \j \ l\ \l \ i\ \\ l = ipd (\ j)\ by blast hence lpd: \\ l pd\ (\ j)\ by (metis is_ipd_def \\ i \ return\ \is_path \\ ipd_is_ipd path_nodes term_path_stable) show \?case\ proof (cases) assume \l = i\ thus \?case\ using lpd by auto next assume \l \ i\ hence \l < i\ using l by simp moreover have \j \ l\ using l by (metis is_ipd_def \\ i \ return\ \is_path \\ ipd_is_ipd path_nodes term_path_stable) hence \j < l\ using l by simp moreover hence \i - l < i - j\ by (metis diff_less_mono2 less.prems(2)) moreover have \k < l\ by (metis l(1) less.prems(1) linorder_neqE_nat not_le order.strict_trans) ultimately have \\ i pd\ (\ l)\ using less.hyps by auto thus \?case\ using lpd by (metis pd_trans) qed qed lemma icd_pd_intermediate: assumes icd: \i icd\<^bsup>\\<^esup>\ k\ and j: \k < j\ \j \ i\ shows \\ i pd\ (\ j)\ using assms icd_pd_intermediate'[OF assms(1,2)] apply (cases \j < i\,metis) by (metis is_icdi_def le_neq_trans path_nodes pd_refl) lemma no_icd_pd: assumes path: \is_path \\ and noicd: \\ l\n. \ k icd\<^bsup>\\<^esup>\ l\ and nk: \n \ k\ shows \\ k pd\ \ n\ proof cases assume \\ k = return\ thus \?thesis\ by (metis path path_nodes return_pd) next assume nret: \\ k \ return\ have nocd: \\ l. n\l \ \ k cd\<^bsup>\\<^esup>\ l\ proof fix l assume kcd: \k cd\<^bsup>\\<^esup>\ l\ and nl: \n \ l\ hence \(k - n) cd\<^bsup>\\n\<^esup>\ (l - n)\ using cd_path_shift[OF nl path] by simp hence \\ l. (k - n) icd\<^bsup>\\n\<^esup>\ l\ using excd_impl_exicd by blast - then guess l' .. + then obtain l' where "k - n icd\<^bsup>\ \ n\<^esup>\ l'" .. hence \k icd\<^bsup>\\<^esup>\ (l' + n)\ using icd_path_shift[of \n\ \l' + n\ \\\ \k\] path by auto thus \False\ using noicd by auto qed hence \\l. n \ l \ l \ j \ {l..k}. \ j = ipd (\ l)\ using path nret unfolding is_cdi_def by auto thus \?thesis\ using nk proof (induction \k - n\ arbitrary: \n\ rule: less_induct,cases) case (less n) assume \n = k\ thus \?case\ using pd_refl path path_nodes by auto next case (less n) assume \n \ k\ hence nk: \n < k\ using less(3) by auto with less(2) obtain j where jnk: \j \ {n..k}\ and ipdj: \\ j = ipd (\ n)\ by blast have nretn: \\ n \ return\ using nk nret term_path_stable path by auto with ipd_is_ipd path path_nodes is_ipd_def ipdj have jpdn: \\ j pd\ \ n\ by auto show \?case\ proof cases assume \j = k\ thus \?case\ using jpdn by simp next assume \j \ k\ hence jk: \j < k\ using jnk by auto have \j \ n\ using ipdj by (metis ipd_not_self nretn path path_nodes) hence nj: \n < j\ using jnk by auto have *: \k - j < k - n\ using jk nj by auto with less(1)[OF *] less(2) jk nj have \\ k pd\ \ j\ by auto thus \?thesis\ using jpdn pd_trans by metis qed qed qed lemma first_pd_no_cd: assumes path: \is_path \\ and pd: \\ n pd\ \ 0\ and first: \\ l < n. \ l \ \ n\ shows \\ l. \ n cd\<^bsup>\\<^esup>\ l\ proof (rule ccontr, goal_cases) case 1 then obtain l where ncdl: \n cd\<^bsup>\\<^esup>\ l\ by blast hence ln: \l < n\ using is_cdi_def by auto have \\ \ n pd\ \ l\ using ncdl cd_not_pd by (metis ln first) then obtain \' n' where path': \is_path \'\ and \0: \\' 0 = \ l\ and \n: \\' n' = return\ and not\n: \\ j\ n'. \' j \ \ n\ unfolding is_pd_def using path path_nodes by auto let \?\\ = \\@\<^bsup>l\<^esup> \'\ have \is_path ?\\ by (metis \0 path path' path_cons) moreover have \?\ 0 = \ 0\ by auto moreover have \?\ (n' + l) = return\ using \0 \n by auto ultimately obtain j where j: \j \ n' + l\ and jn: \?\ j = \ n\ using pd unfolding is_pd_def by blast show \False\ proof cases assume \j \ l\ thus \False\ using jn first ln by auto next assume \\ j \ l\ thus \False\ using j jn not\n by auto qed qed lemma first_pd_no_icd: assumes path: \is_path \\ and pd: \\ n pd\ \ 0\ and first: \\ l < n. \ l \ \ n\ shows \\ l. \ n icd\<^bsup>\\<^esup>\ l\ by (metis first first_pd_no_cd icd_imp_cd path pd) lemma path_nret_ex_nipd: assumes \is_path \\ \\ i. \ i \ return\ shows \\ i. (\ j\i. (\ k>j. \ k \ ipd (\ j)))\ proof(rule, rule ccontr) fix i assume \\ (\j\i. \ k>j. \ k \ ipd (\ j))\ hence *: \\ j\i. (\k>j. \ k = ipd (\ j))\ by blast have \\ j. (\k>j. (\\i) k = ipd ((\\i) j))\ proof fix j have \i + j \ i\ by auto then obtain k where k: \k>i+j\ \\ k = ipd (\ (i+j))\ using * by blast hence \(\\i) (k - i) = ipd ((\\i) j)\ by auto moreover have \k - i > j\ using k by auto ultimately show \\k>j. (\\i) k = ipd ((\\i) j)\ by auto qed moreover have \is_path (\\i)\ using assms(1) path_path_shift by simp ultimately obtain k where \(\\i) k = return\ using all_ipd_imp_ret by blast thus \False\ using assms(2) by auto qed lemma path_nret_ex_all_cd: assumes \is_path \\ \\ i. \ i \ return\ shows \\ i. (\ j\i. (\ k>j. k cd\<^bsup>\\<^esup>\ j))\ unfolding is_cdi_def using assms path_nret_ex_nipd[OF assms] by (metis atLeastAtMost_iff ipd_not_self linorder_neqE_nat not_le path_nodes) lemma path_nret_inf_all_cd: assumes \is_path \\ \\ i. \ i \ return\ shows \\ finite {j. \ k>j. k cd\<^bsup>\\<^esup>\ j}\ using unbounded_nat_set_infinite path_nret_ex_all_cd[OF assms] by auto lemma path_nret_inf_icd_seq: assumes path: \is_path \\ and nret: \\ i. \ i \ return\ obtains f where \\ i. f (Suc i) icd\<^bsup>\\<^esup>\ f i\ \range f = {i. \ j>i. j cd\<^bsup>\\<^esup>\ i}\ \\ (\i. f 0 cd\<^bsup>\\<^esup>\ i)\ proof - note path_nret_inf_all_cd[OF assms] then obtain f where ran: \range f = {j. \ k>j. k cd\<^bsup>\\<^esup>\ j}\ and asc: \\ i. f i < f (Suc i)\ using infinite_ascending by blast have mono: \\ i j. i < j \ f i < f j\ using asc by (metis lift_Suc_mono_less) { fix i have cd: \f (Suc i) cd\<^bsup>\\<^esup>\ f i\ using ran asc by auto have \f (Suc i) icd\<^bsup>\\<^esup>\ f i\ proof (rule ccontr) assume \\ f (Suc i) icd\<^bsup>\\<^esup>\ f i\ then obtain m where im: \f i < m\ and mi: \ m < f (Suc i)\ and cdm: \f (Suc i) cd\<^bsup>\\<^esup>\ m\ unfolding is_icdi_def using assms(1) cd by auto have \\ k>m. k cd\<^bsup>\\<^esup>\m\ proof (rule,rule,cases) fix k assume \f (Suc i) < k\ hence \k cd\<^bsup>\\<^esup>\ f (Suc i)\ using ran by auto thus \k cd\<^bsup>\\<^esup>\ m\ using cdm cd_trans by metis next fix k assume mk: \m < k\ and \\ f (Suc i) < k\ hence ik: \k \ f (Suc i)\ by simp thus \k cd\<^bsup>\\<^esup>\ m\ using cdm by (metis cdi_prefix mk) qed hence \m \ range f\ using ran by blast then obtain j where m: \m = f j\ by blast show \False\ using im mi mono unfolding m by (metis Suc_lessI le_less not_le) qed } moreover { fix m assume cdm: \f 0 cd\<^bsup>\\<^esup>\ m\ have \\ k>m. k cd\<^bsup>\\<^esup>\m\ proof (rule,rule,cases) fix k assume \f 0 < k\ hence \k cd\<^bsup>\\<^esup>\ f 0\ using ran by auto thus \k cd\<^bsup>\\<^esup>\ m\ using cdm cd_trans by metis next fix k assume mk: \m < k\ and \\ f 0 < k\ hence ik: \k \ f 0\ by simp thus \k cd\<^bsup>\\<^esup>\ m\ using cdm by (metis cdi_prefix mk) qed hence \m \ range f\ using ran by blast then obtain j where m: \m = f j\ by blast hence fj0: \f j < f 0\ using cdm m is_cdi_def by auto hence \0 < j\ by (metis less_irrefl neq0_conv) hence \False\ using fj0 mono by fastforce } ultimately show \thesis\ using that ran by blast qed lemma cdi_iff_no_strict_pd: \i cd\<^bsup>\\<^esup>\ k \ is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ (\ k, \ j) \ pdt)\ proof assume cd:\i cd\<^bsup>\\<^esup>\ k\ have 1: \is_path \ \ k < i \ \ i \ return\ using cd unfolding is_cdi_def by auto have 2: \\ j \ {k..i}. \ (\ k, \ j) \ pdt\ proof (rule ccontr) assume \ \ (\j\{k..i}. (\ k, \ j) \ pdt)\ then obtain j where \j \ {k..i}\ and \(\ k, \ j) \ pdt\ by auto hence \\ j \ \ k\ and \\ j pd\ \ k\ unfolding pdt_def by auto thus \False\ using path_pd_ipd by (metis \j \ {k..i}\ atLeastAtMost_iff cd cd_not_pd cdi_prefix le_eq_less_or_eq) qed show \is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ (\ k, \ j) \ pdt)\ using 1 2 by simp next assume \is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ (\ k, \ j) \ pdt)\ thus \i cd\<^bsup>\\<^esup>\ k\ by (metis ipd_in_pdt term_path_stable less_or_eq_imp_le not_cd_impl_ipd path_nodes) qed subsection \Facts about Control Slices\ lemma last_cs: \last (cs\<^bsup>\\<^esup> i) = \ i\ by auto lemma cs_not_nil: \cs\<^bsup>\\<^esup> n \ []\ by (auto) lemma cs_return: assumes \\ n = return\ shows \cs\<^bsup>\\<^esup> n = [\ n]\ by (metis assms cs.elims icd_imp_cd ret_no_cd) lemma cs_0[simp]: \cs\<^bsup>\\<^esup> 0 = [\ 0]\ using is_icdi_def is_cdi_def by auto lemma cs_inj: assumes \is_path \\ \\ n \ return\ \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> n'\ shows \n = n'\ using assms proof (induction \cs\<^bsup>\\<^esup> n\ arbitrary: \\\ \n\ \n'\ rule:rev_induct) case Nil hence \False\ using cs_not_nil by metis thus \?case\ by simp next case (snoc x xs \ n n') show \?case\ proof (cases \xs\) case Nil hence *: \\ (\ k. n icd\<^bsup>\\<^esup>\k)\ using snoc(2) cs_not_nil by (auto,metis append1_eq_conv append_Nil cs_not_nil) moreover have \[x] = cs\<^bsup>\\<^esup> n'\ using Nil snoc by auto hence **: \\ (\ k. n' icd\<^bsup>\\<^esup>\k)\ using cs_not_nil by (auto,metis append1_eq_conv append_Nil cs_not_nil) ultimately have \\ k. \ n cd\<^bsup>\\<^esup>\ k\ \\ k. \ n' cd\<^bsup>\\<^esup>\ k\ using excd_impl_exicd by auto blast+ moreover hence \\ n = \ n'\ using snoc(5,2) by auto (metis * ** list.inject) ultimately show \n = n'\ using other_claim' snoc by blast next case (Cons y ys) hence *: \\ k. n icd\<^bsup>\\<^esup>\k\ using snoc(2) by auto (metis append_is_Nil_conv list.distinct(1) list.inject) then obtain k where k: \n icd\<^bsup>\\<^esup>\k\ by auto have \k = (THE k . n icd\<^bsup>\\<^esup>\ k)\ using k by (metis icd_is_the_icd) hence xsk: \xs = cs\<^bsup>\\<^esup> k\ using * k snoc(2) unfolding cs.simps[of \\\ \n\] by auto have **: \\ k. n' icd\<^bsup>\\<^esup>\k\ using snoc(2)[unfolded snoc(5)] by auto (metis Cons append1_eq_conv append_Nil list.distinct(1)) then obtain k' where k': \n' icd\<^bsup>\\<^esup>\ k'\ by auto hence \k' = (THE k . n' icd\<^bsup>\\<^esup>\ k)\ using k' by (metis icd_is_the_icd) hence xsk': \xs = cs\<^bsup>\\<^esup> k'\ using ** k' snoc(5,2) unfolding cs.simps[of \\\ \n'\] by auto hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> k'\ using xsk by simp moreover have kn: \k < n\ using k by (metis is_icdi_def is_cdi_def) hence \\ k \ return\ using snoc by (metis term_path_stable less_imp_le) ultimately have kk'[simp]: \k' = k\ using snoc(1) xsk snoc(3) by metis have nk0: \n - k icd\<^bsup>\\k\<^esup>\ 0\ \n' - k icd\<^bsup>\\k\<^esup>\ 0\ using k k' icd_path_shift0 snoc(3) by auto moreover have nkr: \(\\k)(n-k) \ return\ using snoc(4) kn by auto moreover have \is_path (\\k)\ by (metis path_path_shift snoc.prems(1)) moreover have kn': \k < n'\ using k' kk' by (metis is_icdi_def is_cdi_def) have \\ n = \ n'\ using snoc(5) * ** by auto hence \(\\k)(n-k) = (\\k)(n'-k)\ using kn kn' by auto ultimately have \n - k = n' - k\ using other_claim by auto thus \n = n'\ using kn kn' by auto qed qed lemma cs_cases: fixes \ i obtains (base) \cs\<^bsup>\\<^esup> i = [\ i]\ and \\ k. \ i cd\<^bsup>\\<^esup>\ k\ | (depend) k where \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ and \i icd\<^bsup>\\<^esup>\ k\ proof cases assume *: \\ k. i icd\<^bsup>\\<^esup>\ k\ then obtain k where k: \i icd\<^bsup>\\<^esup>\ k\ .. hence \k = (THE k. i icd\<^bsup>\\<^esup>\ k)\ by (metis icd_is_the_icd) hence \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ using * by auto with k that show \thesis\ by simp next assume *: \\ (\ k. i icd\<^bsup>\\<^esup>\ k)\ hence \\ k. \ i cd\<^bsup>\\<^esup>\ k\ by (metis excd_impl_exicd) moreover have \cs\<^bsup>\\<^esup> i = [\ i]\ using * by auto ultimately show \thesis\ using that by simp qed lemma cs_length_one: assumes \length (cs\<^bsup>\\<^esup> i) = 1\ shows \cs\<^bsup>\\<^esup> i = [\ i]\ and \\ k. \ i cd\<^bsup>\\<^esup>\ k\ apply (cases \i\ \\\ rule: cs_cases) using assms cs_not_nil apply auto apply (cases \i\ \\\ rule: cs_cases) using assms cs_not_nil by auto lemma cs_length_g_one: assumes \length (cs\<^bsup>\\<^esup> i) \ 1\ obtains k where \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ and \i icd\<^bsup>\\<^esup>\ k\ apply (cases \i\ \\\ rule: cs_cases) using assms cs_not_nil by auto lemma claim: assumes path: \is_path \\ \is_path \'\ and ii: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ and jj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and bl: \butlast (cs\<^bsup>\\<^esup> i) = butlast (cs\<^bsup>\\<^esup> j)\ and nret: \\ i \ return\ and ilj: \i < j\ shows \i' < j'\ proof (cases ) assume *: \length (cs\<^bsup>\\<^esup> i) = 1\ hence **: \length (cs\<^bsup>\\<^esup> i) = 1\ \length (cs\<^bsup>\\<^esup> j) = 1\ \length (cs\<^bsup>\'\<^esup> i') = 1\ \length (cs\<^bsup>\'\<^esup> j') = 1\ apply metis apply (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil) apply (metis "*" ii) by (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil jj) then obtain \cs\<^bsup>\\<^esup> i = [\ i]\ \cs\<^bsup>\\<^esup> j = [\ j]\ \cs\<^bsup>\'\<^esup> j' = [\' j']\ \cs\<^bsup>\'\<^esup> i'= [\' i']\ \\ k. \ j cd\<^bsup>\\<^esup>\ k\ \\ k. \ i' cd\<^bsup>\'\<^esup>\ k\ \\ k. \ j' cd\<^bsup>\'\<^esup>\ k\ by (metis cs_length_one ** ) moreover hence \\ i = \' i'\ \\ j = \' j'\ using assms by auto ultimately show \i' < j'\ using nret ilj path claim'' by blast next assume *: \length (cs\<^bsup>\\<^esup> i) \ 1\ hence **: \length (cs\<^bsup>\\<^esup> i) \ 1\ \length (cs\<^bsup>\\<^esup> j) \ 1\ \length (cs\<^bsup>\'\<^esup> i') \ 1\ \length (cs\<^bsup>\'\<^esup> j') \ 1\ apply metis apply (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil) apply (metis "*" ii) by (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil jj) obtain k l k' l' where ***: \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ \cs\<^bsup>\\<^esup> j = (cs\<^bsup>\\<^esup> l)@[\ j]\ \cs\<^bsup>\'\<^esup> i' = (cs\<^bsup>\'\<^esup> k')@[\' i']\ \cs\<^bsup>\'\<^esup> j' = (cs\<^bsup>\'\<^esup> l')@[\' j']\ and icds: \i icd\<^bsup>\\<^esup>\ k\ \j icd\<^bsup>\\<^esup>\ l\ \i' icd\<^bsup>\'\<^esup>\ k'\ \j' icd\<^bsup>\'\<^esup>\ l'\ by (metis ** cs_length_g_one) hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> l\ \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\'\<^esup> l'\ using assms by auto moreover have \\ k \ return\ \\' k' \ return\ using nret apply (metis is_icdi_def icds(1) is_cdi_def term_path_stable less_imp_le) by (metis is_cdi_def is_icdi_def icds(3) term_path_stable less_imp_le) ultimately have lk[simp]: \l = k\ \l' = k'\ using path cs_inj by auto let \?\\ = \\ \ k\ let \?\'\ = \\'\k'\ have \i-k icd\<^bsup>?\\<^esup>\ 0\ \j-k icd\<^bsup>?\\<^esup>\ 0\ \i'-k' icd\<^bsup>?\'\<^esup>\ 0\ \j'-k' icd\<^bsup>?\'\<^esup>\ 0\ using icd_path_shift0 path icds by auto moreover have ki: \k < i\ using icds by (metis is_icdi_def is_cdi_def) hence \i-k < j-k\ by (metis diff_is_0_eq diff_less_mono ilj nat_le_linear order.strict_trans) moreover have \i: \\ i = \' i'\ \\ j = \' j'\ using assms *** by auto have \k' < i'\ \k' < j'\ using icds unfolding lk by (metis is_cdi_def is_icdi_def)+ hence \?\ (i-k) = ?\' (i'-k')\ \?\ (j-k) = ?\' (j'-k')\ using \i ki ilj by auto moreover have \?\ (i-k) \ return\ using nret ki by auto moreover have \is_path ?\\ \is_path ?\'\ using path path_path_shift by auto ultimately have \i'-k' < j' - k'\ using claim' by blast thus \i' < j'\ by (metis diff_is_0_eq diff_less_mono less_nat_zero_code linorder_neqE_nat nat_le_linear) qed lemma cs_split': assumes \cs\<^bsup>\\<^esup> i = xs@[x,x']@ys\ shows \\ m. cs\<^bsup>\\<^esup> m = xs@[x] \ i cd\<^bsup>\\<^esup>\ m\ using assms proof (induction \ys\ arbitrary: \i\ rule:rev_induct ) case (snoc y ys) hence \length (cs\<^bsup>\\<^esup> i) \ 1\ by auto then obtain i' where \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> i') @ [\ i]\ and *: \i icd\<^bsup>\\<^esup>\ i'\ using cs_length_g_one[of \\\ \i\] by metis hence \cs\<^bsup>\\<^esup> i' = xs@[x,x']@ys\ using snoc(2) by (metis append1_eq_conv append_assoc) then obtain m where **: \cs\<^bsup>\\<^esup> m = xs @ [x]\ and \i' cd\<^bsup>\\<^esup>\ m\ using snoc(1) by blast hence \i cd\<^bsup>\\<^esup>\ m\ using * cd_trans by (metis is_icdi_def) with ** show \?case\ by blast next case Nil hence \length (cs\<^bsup>\\<^esup> i) \ 1\ by auto then obtain i' where a: \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> i') @ [\ i]\ and *: \i icd\<^bsup>\\<^esup>\ i'\ using cs_length_g_one[of \\\ \i\] by metis have \cs\<^bsup>\\<^esup> i = (xs@[x])@[x']\ using Nil by auto hence \cs\<^bsup>\\<^esup> i' = xs@[x]\ using append1_eq_conv a by metis thus \?case\ using * unfolding is_icdi_def by blast qed lemma cs_split: assumes \cs\<^bsup>\\<^esup> i = xs@[x]@ys@[\ i]\ shows \\ m. cs\<^bsup>\\<^esup> m = xs@[x] \ i cd\<^bsup>\\<^esup>\ m\ proof - obtain x' ys' where \ys@[\ i] = [x']@ys'\ by (metis append_Cons append_Nil neq_Nil_conv) thus \?thesis\ using cs_split'[of \\\ \i\ \xs\ \x\ \x'\ \ys'\] assms by auto qed lemma cs_less_split: assumes \xs \ ys\ obtains a as where \ys = xs@a#as\ using assms unfolding cs_less.simps apply auto by (metis Cons_nth_drop_Suc append_take_drop_id) lemma cs_select_is_cs: assumes \is_path \\ \xs \ Nil\ \xs \ cs\<^bsup>\\<^esup> k\ shows \cs\<^bsup>\\<^esup> (\\xs) = xs\ \k cd\<^bsup>\\<^esup>\ (\\xs)\proof - obtain b bs where b: \cs\<^bsup>\\<^esup> k = xs@b#bs\ using assms cs_less_split by blast obtain a as where a: \xs = as@[a]\ using assms by (metis rev_exhaust) have \cs\<^bsup>\\<^esup> k = as@[a,b]@bs\ using a b by auto then obtain k' where csk: \cs\<^bsup>\\<^esup> k' = xs\ and is_cd: \k cd\<^bsup>\\<^esup>\ k'\ using cs_split' a by blast hence nret: \\ k' \ return\ by (metis is_cdi_def term_path_stable less_imp_le) show a: \cs\<^bsup>\\<^esup> (\\xs) = xs\ unfolding cs_select_def using cs_inj[OF assms(1) nret] csk the_equality[of _ \k'\] by (metis (mono_tags)) show \k cd\<^bsup>\\<^esup>\ (\\xs)\ unfolding cs_select_def by (metis a assms(1) cs_inj cs_select_def csk is_cd nret) qed lemma cd_in_cs: assumes \n cd\<^bsup>\\<^esup>\ m\ shows \\ ns. cs\<^bsup>\\<^esup> n = (cs\<^bsup>\\<^esup> m) @ ns @[\ n]\ using assms proof (induction rule: cd_induct) case (base n) thus \?case\ by (metis append_Nil cs.simps icd_is_the_icd) next case (IS k n) hence \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> k @ [\ n]\ by (metis cs.simps icd_is_the_icd) thus \?case\ using IS by force qed lemma butlast_cs_not_cd: assumes \butlast (cs\<^bsup>\\<^esup> m) = butlast (cs\<^bsup>\\<^esup> n)\ shows \\ m cd\<^bsup>\\<^esup>\n\ by (metis append_Cons append_Nil append_assoc assms cd_in_cs cs_not_nil list.distinct(1) self_append_conv snoc_eq_iff_butlast) lemma wn_cs_butlast: assumes \butlast (cs\<^bsup>\\<^esup> m) = butlast (cs\<^bsup>\\<^esup> n)\ \i cd\<^bsup>\\<^esup>\ m\ \j cd\<^bsup>\\<^esup>\ n\ \m shows \i proof (rule ccontr) assume \\ i < j\ moreover have \\ n cd\<^bsup>\\<^esup>\ m\ by (metis assms(1) butlast_cs_not_cd) ultimately have \n \ m\ using assms(2,3) cr_wn'' by auto thus \False\ using assms(4) by auto qed text \This is the central theorem making the control slice suitable for matching indices between executions.\ theorem cs_order: assumes path: \is_path \\ \is_path \'\ and csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and nret: \\ i \ return\ and ilj: \i < j\ shows \i' proof - have \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\\<^esup> j\ using cs_inj[OF path(1) nret] ilj by blast moreover have \cs\<^bsup>\\<^esup> i \ Nil\ \cs\<^bsup>\\<^esup> j \ Nil\ by (metis cs_not_nil)+ ultimately show \?thesis\ proof (cases rule: list_neq_prefix_cases) case (diverge xs x x' ys ys') note csx = \cs\<^bsup>\\<^esup> i = xs @ [x] @ ys\ note csx' = \cs\<^bsup>\\<^esup> j = xs @ [x'] @ ys'\ note xx = \x \ x'\ show \i' < j'\ proof (cases \ys\) assume ys: \ys = Nil\ show \?thesis\ proof (cases \ys'\) assume ys': \ys' = Nil\ have cs: \cs\<^bsup>\\<^esup> i = xs @ [x]\ \cs\<^bsup>\\<^esup> j = xs @ [x']\ by (metis append_Nil2 csx ys, metis append_Nil2 csx' ys') hence bl: \butlast (cs\<^bsup>\\<^esup> i) = butlast (cs\<^bsup>\\<^esup> j)\ by auto show \i' < j'\ using claim[OF path csi csj bl nret ilj] . next fix y' zs' assume ys': \ys' = y'#zs'\ have cs: \cs\<^bsup>\\<^esup> i = xs @ [x]\ \cs\<^bsup>\\<^esup> j = xs @ [x',y']@ zs'\ by (metis append_Nil2 csx ys, metis append_Cons append_Nil csx' ys') obtain n where n: \cs\<^bsup>\\<^esup> n = xs@[x']\ and jn: \j cd\<^bsup>\\<^esup>\ n\ using cs cs_split' by blast obtain n' where n': \cs\<^bsup>\'\<^esup> n' = xs@[x']\ and jn': \j' cd\<^bsup>\'\<^esup>\ n'\ using cs cs_split' unfolding csj by blast have csn : \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and bl: \butlast (cs\<^bsup>\\<^esup> i) = butlast (cs\<^bsup>\\<^esup> n)\ using n n' cs by auto hence bl': \butlast (cs\<^bsup>\'\<^esup> i') = butlast (cs\<^bsup>\'\<^esup> n')\ using csi by auto have notcd: \\ i cd\<^bsup>\\<^esup>\ n\ by (metis butlast_cs_not_cd bl) have nin: \i \ n\ using cs n xx by auto have iln: \i < n\ apply (rule ccontr) using cr_wn'[OF jn notcd] nin ilj by auto note claim[OF path csi csn bl nret iln] hence \i' < n'\ . thus \i' < j'\ using jn' unfolding is_cdi_def by auto qed next fix y zs assume ys: \ys = y#zs\ show \?thesis\ proof (cases \ys'\) assume ys' : \ys' = Nil\ have cs: \cs\<^bsup>\\<^esup> i = xs @ [x,y]@zs\ \cs\<^bsup>\\<^esup> j = xs @ [x']\ by (metis append_Cons append_Nil csx ys, metis append_Nil2 csx' ys') obtain n where n: \cs\<^bsup>\\<^esup> n = xs@[x]\ and jn: \i cd\<^bsup>\\<^esup>\ n\ using cs cs_split' by blast obtain n' where n': \cs\<^bsup>\'\<^esup> n' = xs@[x]\ and jn': \i' cd\<^bsup>\'\<^esup>\ n'\ using cs cs_split' unfolding csi by blast have csn : \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and bl: \butlast (cs\<^bsup>\\<^esup> n) = butlast (cs\<^bsup>\\<^esup> j)\ using n n' cs by auto hence bl': \butlast (cs\<^bsup>\'\<^esup> j') = butlast (cs\<^bsup>\'\<^esup> n')\ using csj by auto have notcd: \\ j' cd\<^bsup>\'\<^esup>\ n'\ by (metis butlast_cs_not_cd bl') have nin: \n < i\ using jn unfolding is_cdi_def by auto have nlj: \n < j\ using nin ilj by auto note claim[OF path csn csj bl _ nlj] hence nj': \n' < j'\ using term_path_stable[OF path(1) _] less_imp_le nin nret by auto show \i' < j'\ apply(rule ccontr) using cdi_prefix[OF jn' nj'] notcd by auto next fix y' zs' assume ys' : \ys' = y'#zs'\ have cs: \cs\<^bsup>\\<^esup> i = xs@[x,y]@zs\ \cs\<^bsup>\\<^esup> j = xs@[x',y']@zs'\ by (metis append_Cons append_Nil csx ys,metis append_Cons append_Nil csx' ys') have neq: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\\<^esup> j\ using cs_inj path nret ilj by blast obtain m where m: \cs\<^bsup>\\<^esup> m = xs@[x]\ and im: \i cd\<^bsup>\\<^esup>\ m\ using cs cs_split' by blast obtain n where n: \cs\<^bsup>\\<^esup> n = xs@[x']\ and jn: \j cd\<^bsup>\\<^esup>\ n\ using cs cs_split' by blast obtain m' where m': \cs\<^bsup>\'\<^esup> m' = xs@[x]\ and im': \i' cd\<^bsup>\'\<^esup>\ m'\ using cs cs_split' unfolding csi by blast obtain n' where n': \cs\<^bsup>\'\<^esup> n' = xs@[x']\ and jn': \j' cd\<^bsup>\'\<^esup>\ n'\ using cs cs_split' unfolding csj by blast have \m \ n\ using ilj m n wn_cs_butlast[OF _ jn im] by force moreover have \m \ n\ using m n xx by (metis last_snoc) ultimately have mn: \m < n\ by auto moreover have \\ m \ return\ by (metis last_cs last_snoc m mn n path(1) term_path_stable xx less_imp_le) moreover have \butlast (cs\<^bsup>\\<^esup> m) = butlast (cs\<^bsup>\\<^esup> n)\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ using m n n' m' by auto ultimately have \m' < n'\ using claim path by blast thus \i' < j'\ using m' n' im' jn' wn_cs_butlast by (metis butlast_snoc) qed qed next case (prefix1 xs) note pfx = \cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> j @ xs\ note xs = \xs \ []\ obtain a as where \xs = a#as\ using xs by (metis list.exhaust) moreover obtain bs b where bj: \cs\<^bsup>\\<^esup> j = bs@[b]\ using cs_not_nil by (metis rev_exhaust) ultimately have \cs\<^bsup>\\<^esup> i = bs@[b,a]@as\ using pfx by auto then obtain m where \cs\<^bsup>\\<^esup> m = bs@[b]\ and cdep: \i cd\<^bsup>\\<^esup>\ m\ using cs_split' by blast hence mi: \m = j\ using bj cs_inj by (metis is_cdi_def term_path_stable less_imp_le) hence \i cd\<^bsup>\\<^esup>\ j\ using cdep by auto hence \False\ using ilj unfolding is_cdi_def by auto thus \i' < j'\ .. next case (prefix2 xs) have pfx : \cs\<^bsup>\'\<^esup> i' @ xs = cs\<^bsup>\'\<^esup> j'\ using prefix2 csi csj by auto note xs = \xs \ []\ obtain a as where \xs = a#as\ using xs by (metis list.exhaust) moreover obtain bs b where bj: \cs\<^bsup>\'\<^esup> i' = bs@[b]\ using cs_not_nil by (metis rev_exhaust) ultimately have \cs\<^bsup>\'\<^esup> j' = bs@[b,a]@as\ using pfx by auto then obtain m where \cs\<^bsup>\'\<^esup> m = bs@[b]\ and cdep: \j' cd\<^bsup>\'\<^esup>\ m\ using cs_split' by blast hence mi: \m = i'\ using bj cs_inj by (metis is_cdi_def term_path_stable less_imp_le) hence \j' cd\<^bsup>\'\<^esup>\ i'\ using cdep by auto thus \i' < j'\ unfolding is_cdi_def by auto qed qed lemma cs_order_le: assumes path: \is_path \\ \is_path \'\ and csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and nret: \\ i \ return\ and ilj: \i \ j\ shows \i'\j'\ proof cases assume \i < j\ with cs_order[OF assms(1,2,3,4,5)] show \?thesis\ by simp next assume \\ i < j\ hence \i = j\ using ilj by simp hence csij: \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\'\<^esup> j'\ using csi csj by simp have nret': \\' i' \ return\ using nret last_cs csi by metis show \?thesis\ using cs_inj[OF path(2) nret' csij] by simp qed lemmas cs_induct[case_names cs] = cs.induct lemma icdi_path_swap: assumes \is_path \'\ \j icd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>j\<^esub> \'\ shows \j icd\<^bsup>\'\<^esup>\k\ using assms unfolding eq_up_to_def is_icdi_def is_cdi_def by auto lemma icdi_path_swap_le: assumes \is_path \'\ \j icd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>n\<^esub> \'\ \j \ n\ shows \j icd\<^bsup>\'\<^esup>\k\ by (metis assms icdi_path_swap eq_up_to_le) lemma cs_path_swap: assumes \is_path \\ \is_path \'\ \\ =\<^bsub>k\<^esub> \'\ shows \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k\ using assms(1,3) proof (induction \\\ \k\ rule:cs_induct,cases) case (cs \ k) let \?l\ = \(THE l. k icd\<^bsup>\\<^esup>\ l)\ assume *: \\l. k icd\<^bsup>\\<^esup>\ l\ have kicd: \k icd\<^bsup>\\<^esup>\ ?l\ by (metis "*" icd_is_the_icd) hence \?l < k\ unfolding is_cdi_def[of \k\ \\\ \?l\] is_icdi_def[of \k\ \\\ \?l\] by auto hence \\ i\?l. \ i = \' i\ using cs(2,3) unfolding eq_up_to_def by auto hence csl: \cs\<^bsup>\\<^esup> ?l = cs\<^bsup>\'\<^esup> ?l\ using cs(1,2) * unfolding eq_up_to_def by auto have kicd: \k icd\<^bsup>\\<^esup>\ ?l\ by (metis "*" icd_is_the_icd) hence csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> ?l @ [\ k]\ using kicd by auto have kicd': \k icd\<^bsup>\'\<^esup>\ ?l\ using kicd icdi_path_swap[OF assms(2) _ cs(3)] by simp hence \?l = (THE l. k icd\<^bsup>\'\<^esup>\ l)\ by (metis icd_is_the_icd) hence csk': \cs\<^bsup>\'\<^esup> k = cs\<^bsup>\'\<^esup> ?l @ [\' k]\ using kicd' by auto have \\' k = \ k\ using cs(3) unfolding eq_up_to_def by auto with csl csk csk' show \?case\ by auto next case (cs \ k) assume *: \\ (\l. k icd\<^bsup>\\<^esup>\ l)\ hence csk: \cs\<^bsup>\\<^esup> k = [\ k]\ by auto have \\ (\l. k icd\<^bsup>\'\<^esup>\ l)\ apply (rule ccontr) using * icdi_path_swap_le[OF cs(2) _, of \k\ \\'\] cs(3) by (metis eq_up_to_sym le_refl) hence csk': \cs\<^bsup>\'\<^esup> k = [\' k]\ by auto with csk show \?case\ using cs(3) eq_up_to_apply by auto qed lemma cs_path_swap_le: assumes \is_path \\ \is_path \'\ \\ =\<^bsub>n\<^esub> \'\ \k \ n\ shows \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k\ by (metis assms cs_path_swap eq_up_to_le) lemma cs_path_swap_cd: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and \n cd\<^bsup>\\<^esup>\ k\ obtains k' where \n' cd\<^bsup>\'\<^esup>\ k'\ and \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ proof - from cd_in_cs[OF assms(4)] obtain ns where *: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> k @ ns @ [\ n]\ by blast obtain xs x where csk: \cs\<^bsup>\\<^esup> k = xs @ [x]\ by (metis cs_not_nil rev_exhaust) have \\ n = \' n'\ using assms(3) last_cs by metis hence **: \cs\<^bsup>\'\<^esup> n' = xs@[x]@ns@[\' n']\ using * assms(3) csk by auto from cs_split[OF **] obtain k' where \cs\<^bsup>\'\<^esup> k' = xs @ [x]\ \n' cd\<^bsup>\'\<^esup>\ k'\ by blast thus \thesis\ using that csk by auto qed lemma path_ipd_swap: assumes \is_path \\ \\ k \ return\ \k < n\ obtains \' m where \is_path \'\ \\ =\<^bsub>n\<^esub> \'\ \k < m\ \\' m = ipd (\' k)\ \\ l \ {k..' l \ ipd (\' k)\ proof - obtain \' r where *: \\' 0 = \ n\ \is_path \'\ \\' r = return\ by (metis assms(1) path_nodes reaching_ret) let \?\\ = \\@\<^bsup>n\<^esup> \'\ have path: \is_path ?\\ and ret: \?\ (n + r) = return\ and equpto: \?\ =\<^bsub>n\<^esub> \\ using assms path_cons * path_append_eq_up_to by auto have \k: \?\ k = \ k\ by (metis assms(3) less_imp_le_nat path_append_def) obtain j where j: \k < j \ j \ (n + r) \ ?\ j = ipd (\ k)\ (is \?P j\ )by (metis \k assms(2) path path_ret_ipd ret) define m where m: \m \ LEAST m . ?P m\ have Pm: \?P m\ using LeastI[of \?P\ \j\] j m by auto hence km: \k < m\ \m \ (n + r)\ \?\ m = ipd (\ k)\ by auto have le: \\l. ?P l \ m \ l\ using Least_le[of \?P\] m by blast have \knipd: \?\ k \ ipd (\ k)\ by (metis \k assms(1) assms(2) ipd_not_self path_nodes) have nipd': \\l. k < l \ l < m \ ?\ l \ ipd (\ k)\ apply (rule ccontr) using le km(2) by force have \\ l \ {k.. l \ ipd(\ k)\ using \knipd nipd' by(auto, metis le_eq_less_or_eq,metis le_eq_less_or_eq) thus \thesis\ using that by (metis \k eq_up_to_sym km(1) km(3) path path_append_eq_up_to) qed lemma cs_sorted_list_of_cd': \cs\<^bsup>\\<^esup> k = map \ (sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i}) @ [\ k]\ proof (induction \\\ \k\ rule: cs.induct, cases) case (1 \ k) assume \\ j. k icd\<^bsup>\\<^esup>\ j\ - then guess j .. - note j = this + then obtain j where j: "k icd\<^bsup>\\<^esup>\ j" .. hence csj: \cs\<^bsup>\\<^esup> j = map \ (sorted_list_of_set {i. j cd\<^bsup>\\<^esup>\ i}) @ [\ j]\ by (metis "1.IH" icd_is_the_icd) have \{i. k cd\<^bsup>\\<^esup>\ i} = insert j {i. j cd\<^bsup>\\<^esup>\ i}\ using cdi_is_cd_icdi[OF j] by auto moreover have f: \finite {i. j cd\<^bsup>\\<^esup>\ i}\ unfolding is_cdi_def by auto moreover have \j \ {i. j cd\<^bsup>\\<^esup>\ i}\ unfolding is_cdi_def by auto ultimately have \sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i} = insort j (sorted_list_of_set { i . j cd\<^bsup>\\<^esup>\ i})\ using sorted_list_of_set_insert by auto moreover have \\ x \ {i. j cd\<^bsup>\\<^esup>\ i}. x < j\ unfolding is_cdi_def by auto hence \\ x \ set (sorted_list_of_set {i. j cd\<^bsup>\\<^esup>\ i}). x < j\ by (simp add: f) ultimately have \sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i} = (sorted_list_of_set { i . j cd\<^bsup>\\<^esup>\ i})@[j]\ using insort_greater by auto hence \cs\<^bsup>\\<^esup> j = map \ (sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i})\ using csj by auto thus \?case\ by (metis icd_cs j) next case (1 \ k) assume *: \\ (\ j. k icd\<^bsup>\\<^esup>\ j)\ hence \cs\<^bsup>\\<^esup> k = [\ k]\ by (metis cs_cases) moreover have \{ i . k cd\<^bsup>\\<^esup>\ i} = {}\ by (auto, metis * excd_impl_exicd) ultimately show \?case\ by (metis append_Nil list.simps(8) sorted_list_of_set_empty) qed lemma cs_sorted_list_of_cd: \cs\<^bsup>\\<^esup> k = map \ (sorted_list_of_set ({ i . k cd\<^bsup>\\<^esup>\ i} \ {k}))\ proof - have le: \\ x \ {i. k cd\<^bsup>\\<^esup>\i}.\ y \ {k}. x < y\ unfolding is_cdi_def by auto have fin: \finite {i. k cd\<^bsup>\\<^esup>\i}\ \finite {k}\ unfolding is_cdi_def by auto show \?thesis\ unfolding cs_sorted_list_of_cd'[of \\\ \k\] sorted_list_of_set_append[OF fin le] by auto qed lemma cs_not_ipd: assumes \k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd (\ k)\ (is \?Q j\) shows \cs\<^bsup>\\<^esup> (GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd (\ k)) = [n\cs\<^bsup>\\<^esup> k . ipd n \ ipd (\ k)]\ (is \cs\<^bsup>\\<^esup> ?j = filter ?P _\) proof - have csk: \cs\<^bsup>\\<^esup> k = map \ (sorted_list_of_set ({ i . k cd\<^bsup>\\<^esup>\ i } \ {k}))\ by (metis cs_sorted_list_of_cd) have csj: \cs\<^bsup>\\<^esup> ?j = map \ (sorted_list_of_set ({i. ?j cd\<^bsup>\\<^esup>\ i } \ {?j}))\ by (metis cs_sorted_list_of_cd) have bound: \\ j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd(\ k) \ j \ k\ unfolding is_cdi_def by simp have kcdj: \k cd\<^bsup>\\<^esup>\ ?j\ and ipd': \ipd (\ ?j) \ ipd(\ k)\ using GreatestI_nat[of \?Q\ \j\ \k\, OF assms] bound by auto have greatest: \\ j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd (\ k) \ j \ ?j\ using Greatest_le_nat[of \?Q\ _ \k\] bound by auto have less_not_ipdk: \\ j. k cd\<^bsup>\\<^esup>\ j \ j < ?j \ ipd (\ j) \ ipd (\ k)\ by (metis (lifting) ipd' kcdj same_ipd_stable) hence le_not_ipdk: \\ j. k cd\<^bsup>\\<^esup>\ j \ j \ ?j \ ipd (\ j) \ ipd (\ k)\ using kcdj ipd' by (case_tac \j = ?j\,auto) have *: \{j \ {i. k cd\<^bsup>\\<^esup>\i} \ {k}. ?P (\ j)} = insert ?j { i . ?j cd\<^bsup>\\<^esup>\ i} \ apply auto apply (metis (lifting, no_types) greatest cr_wn'' kcdj le_antisym le_refl) apply (metis kcdj) apply (metis ipd') apply (metis (full_types) cd_trans kcdj) apply (subgoal_tac \k cd\<^bsup>\\<^esup>\ x\) apply (metis (lifting, no_types) is_cdi_def less_not_ipdk) by (metis (full_types) cd_trans kcdj) have \finite ({i . k cd\<^bsup>\\<^esup>\ i} \ {k})\ unfolding is_cdi_def by auto note filter_sorted_list_of_set[OF this, of \?P o \\] hence \[n\cs\<^bsup>\\<^esup> k . ipd n \ ipd(\ k)] = map \ (sorted_list_of_set {j \ {i. k cd\<^bsup>\\<^esup>\i} \ {k}. ?P (\ j)})\ unfolding csk filter_map by auto also have \\ = map \ (sorted_list_of_set (insert ?j { i . ?j cd\<^bsup>\\<^esup>\ i}))\ unfolding * by auto also have \\ = cs\<^bsup>\\<^esup> ?j\ using csj by auto finally show \?thesis\ by metis qed lemma cs_ipd: assumes ipd: \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ and km: \k < m\ shows \cs\<^bsup>\\<^esup> m = [n\cs\<^bsup>\\<^esup> k . ipd n \ \ m] @ [\ m]\ proof cases assume \\ j. m icd\<^bsup>\\<^esup>\ j\ then obtain j where jicd: \m icd\<^bsup>\\<^esup>\ j\ by blast hence *: \cs\<^bsup>\\<^esup> m = cs\<^bsup>\\<^esup> j @ [\ m]\ by (metis icd_cs) have j: \j = (GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m)\ using jicd assms ipd_icd_greatest_cd_not_ipd by blast moreover have \ipd (\ j) \ ipd (\ k)\ by (metis is_cdi_def is_icdi_def is_ipd_def cd_not_pd ipd(1) ipd_is_ipd jicd path_nodes less_imp_le term_path_stable) moreover have \k cd\<^bsup>\\<^esup>\ j\ unfolding j by (metis (lifting, no_types) assms(3) cd_ipd_is_cd icd_imp_cd ipd(1) ipd(2) j jicd) ultimately have \cs\<^bsup>\\<^esup> j = [n\cs\<^bsup>\\<^esup> k . ipd n \ \ m]\ using cs_not_ipd[of \k\ \\\ \j\] ipd(1) by metis thus \?thesis\ using * by metis next assume noicd: \\ (\ j. m icd\<^bsup>\\<^esup>\ j)\ hence csm: \cs\<^bsup>\\<^esup> m = [\ m]\ by auto have \\j. k cd\<^bsup>\\<^esup>\j \ ipd(\ j) = \ m\ using cd_is_cd_ipd[OF km ipd] by (metis excd_impl_exicd noicd) hence *: \{j \ {i. k cd\<^bsup>\\<^esup>\ i} \ {k}. ipd (\ j) \ \ m} = {}\ using ipd(1) by auto have **: \((\n. ipd n \ \ m) o \) = (\n. ipd (\ n) \ \ m)\ by auto have fin: \finite ({i. k cd\<^bsup>\\<^esup>\ i} \ {k})\ unfolding is_cdi_def by auto note csk = cs_sorted_list_of_cd[of \\\ \k\] hence \[n\cs\<^bsup>\\<^esup> k . ipd n \ \ m] = [n\ (map \ (sorted_list_of_set ({i. k cd\<^bsup>\\<^esup>\ i} \ {k}))) . ipd n \ \ m]\ by simp also have \\ = map \ [n <- sorted_list_of_set ({i. k cd\<^bsup>\\<^esup>\ i} \ {k}). ipd (\ n) \ \ m]\ by (auto simp add: filter_map **) also have \\ = []\ unfolding * filter_sorted_list_of_set[OF fin, of \\n. ipd (\ n) \ \ m\] by auto finally show \?thesis\ using csm by (metis append_Nil) qed lemma converged_ipd_same_icd: assumes path: \is_path \\ \is_path \'\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ and csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and icd: \l icd\<^bsup>\\<^esup>\ k\ and suc: \\ (Suc k) = \' (Suc k')\ and ipd: \\' m' = ipd (\ k)\ \\ n \ {k'..' n \ ipd (\ k)\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof cases assume l: \l = Suc k\ hence \Suc k cd\<^bsup>\\<^esup>\ k\ using icd by (metis is_icdi_def) hence \\ (Suc k) \ ipd (\ k)\ unfolding is_cdi_def by auto hence \\' (Suc k') \ ipd (\' k')\ by (metis csk last_cs suc) moreover have \\' (Suc k') \ return\ by (metis \Suc k cd\<^bsup>\\<^esup>\ k\ ret_no_cd suc) ultimately have \Suc k' cd\<^bsup>\'\<^esup>\ k'\ unfolding is_cdi_def using path(2) apply auto by (metis ipd_not_self le_Suc_eq le_antisym path_nodes term_path_stable) hence \Suc k' icd\<^bsup>\'\<^esup>\ k'\ unfolding is_icdi_def using path(2) by fastforce hence \cs\<^bsup>\'\<^esup> (Suc k') = cs\<^bsup>\'\<^esup> k' @[\' (Suc k')]\ using icd_cs by auto moreover have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ [\ l]\ using icd icd_cs by auto ultimately have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> (Suc k')\ by (metis csk l suc) thus \?thesis\ by blast next assume nsuck: \l \ Suc k\ have kk'[simp]: \\' k' = \ k\ by (metis csk last_cs) have kl: \k < l\ using icd unfolding is_icdi_def is_cdi_def by auto hence skl: \Suc k < l\ by (metis Suc_lessI nsuck) hence lpd: \\ l pd\ \ (Suc k)\ using icd icd_pd_intermediate by auto have km: \k < m\ by (metis converge(1) kl order.strict_trans) have lcd: \l cd\<^bsup>\\<^esup>\ k\ using icd is_icdi_def by auto hence ipdk_pdl: \ipd (\ k) pd\ (\ l)\ by (metis ipd_pd_cd) have *: \ipd (\ k) \ nodes\ by (metis ipdk_pdl pd_node1) have nretk: \\ k \ return\ by (metis kl lcd path(1) ret_no_cd term_path_stable less_imp_le) have **: \\ (\ l) pd\ ipd (\ k)\ proof assume a: \\ l pd\ ipd (\ k)\ hence \\ l pd\ (\ k)\ by (metis is_ipd_def \k < l\ ipd_is_ipd ipdk_pdl path(1) path_nodes pd_antisym term_path_stable less_imp_le) moreover have \\ l \ (\ k)\ by (metis "*" a ipd_not_self ipdk_pdl lcd pd_antisym ret_no_cd) ultimately show \False\ using lcd cd_not_pd by auto qed have km': \k' < m'\ using cs_order[OF path csk converge(2) nretk km] . obtain \'' n'' where path'': \is_path \''\ and \''0: \\'' 0 = ipd (\ k)\ and \''n: \\'' n'' = return\ and not\l: \\ i\n''. \'' i \ \ l\ using no_pd_path[OF * **] . let \?\'\ = \(\' @\<^bsup>m'\<^esup> \'') \ Suc k'\ have \is_path ?\'\ by (metis \''0 ipd(1) path'' path(2) path_cons path_path_shift) moreover have \?\' 0 = \ (Suc k)\ using km' suc by auto moreover have \?\' (m' - Suc k' + n'') = return\ using \''n km' \''0 ipd(1) by auto ultimately obtain l'' where l'': \l'' \ m' - Suc k' + n''\ \?\' l'' = \ l\ using lpd unfolding is_pd_def by blast have l''m: \l'' \ m' - Suc k'\ apply (rule ccontr) using l'' not\l km' by (cases \Suc (k' + l'') \ m'\, auto) let \?l'\ = \Suc ( k' + l'')\ have lm': \?l' \ m'\ using l''m km' by auto \ \Now we have found our desired l'\ have 1: \\' ?l' = \ l\ using l'' l''m lm' by auto have 2: \k' < ?l'\ by simp have 3: \?l' < m'\ apply (rule ccontr) using lm' by (simp, metis "**" 1 ipd(1) ipdk_pdl) \ \Need the least such l'\ let \?P\ = \\ l'. \' l' = \ l \ k' < l' \ l' < m'\ have *: \?P ?l'\ using 1 2 3 by blast define l' where l': \l' == LEAST l'. ?P l'\ have \l': \\' l' = \ l\ using l' 1 2 3 LeastI[of \?P\] by blast have kl': \k' < l'\ using l' 1 2 3 LeastI[of \?P\] by blast have lm': \l' < m'\ using l' 1 2 3 LeastI[of \?P\] by blast have nretl': \\' l' \ return\ by (metis \''n \l' le_refl not\l) have nipd': \\ j \ {k'..l'}. \' j \ ipd (\' k')\ using lm' kk' ipd(2) kl' by force have lcd': \l' cd\<^bsup>\'\<^esup>\ k'\ by (metis is_cdi_def kl' nipd' nretl' path(2)) have licd: \l' icd\<^bsup>\'\<^esup>\ k'\ proof - have \\ m \ {k'<.. l' cd\<^bsup>\'\<^esup>\ m\ proof (rule ccontr) assume \\ (\ m \ {k'<.. l' cd\<^bsup>\'\<^esup>\ m)\ then obtain j' where kj': \k' < j'\ and jl': \j' < l'\ and lcdj': \l' cd\<^bsup>\'\<^esup>\ j'\ by force have jm': \j' by (metis jl' lm' order.strict_trans) have \\' j' \ \ l\ apply (rule ccontr) using l' kj' jm' jl' Least_le[of \?P\ \j'\] by auto hence \\ \' l' pd\ \' j'\ using cd_not_pd lcdj' \l' by metis moreover have \\' j' \ nodes\ using path(2) path_nodes by auto ultimately obtain \\<^sub>1 n\<^sub>1 where path\<^sub>1: \is_path \\<^sub>1\ and \0\<^sub>1: \\\<^sub>1 0 = \' j'\ and \n\<^sub>1: \\\<^sub>1 n\<^sub>1 = return\ and nl': \\ l \n\<^sub>1. \\<^sub>1 l \ \' l'\ unfolding is_pd_def by blast let \?\''\ = \(\'@\<^bsup>j'\<^esup> \\<^sub>1) \ Suc k'\ have \is_path ?\''\ by (metis \0\<^sub>1 path(2) path\<^sub>1 path_cons path_path_shift) moreover have \?\'' 0 = \ (Suc k)\ by (simp, metis kj' less_eq_Suc_le suc) moreover have kj': \Suc k' \ j'\ by (metis kj' less_eq_Suc_le) hence \?\'' (j' - Suc k' + n\<^sub>1) = return\ by (simp, metis \0\<^sub>1 \n\<^sub>1) ultimately obtain l'' where *: \?\'' l'' = \ l\ and **: \l'' \j' - Suc k' + n\<^sub>1\ using lpd is_pd_def by blast show \False\ proof (cases) assume a: \l'' \ j' - Suc k'\ hence \\' (l'' + Suc k') = \ l\ using * kj' by(simp, metis Nat.le_diff_conv2 add_Suc diff_add_inverse le_add1 le_add_diff_inverse2) moreover have \l'' + Suc k' < l'\ by (metis a jl' add_diff_cancel_right' kj' le_add_diff_inverse less_imp_diff_less ordered_cancel_comm_monoid_diff_class.le_diff_conv2) moreover have \l'' + Suc k' < m'\ by (metis Suc_lessD calculation(2) less_trans_Suc lm') moreover have \k' < l'' + Suc k'\ by simp ultimately show \False\ using Least_le[of \?P\ \l'' + Suc k'\] l' by auto next assume a: \\ l'' \ j' - Suc k'\ hence \\ Suc (k' + l'') \ j'\ by simp hence \\\<^sub>1 (Suc (k' + l'') - j') = \ l\ using * kj' by simp moreover have \Suc (k' + l'') - j' \ n\<^sub>1\ using ** kj' by simp ultimately show \False\ using nl' by (metis \l') qed qed thus \?thesis\ unfolding is_icdi_def using lcd' path(2) by simp qed hence \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\'\<^esup> k' @ [\' l']\ by (metis icd_cs) hence \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l\ by (metis \l' csk icd icd_cs) thus \?thesis\ by metis qed lemma converged_same_icd: assumes path: \is_path \\ \is_path \'\ and converge: \l < n\ \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and icd: \l icd\<^bsup>\\<^esup>\ k\ and suc: \\ (Suc k) = \' (Suc k')\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof - have nret: \\ k \ return\ using icd unfolding is_icdi_def is_cdi_def using term_path_stable less_imp_le by metis have kl: \k < l\ using icd unfolding is_icdi_def is_cdi_def by auto have kn: \k < n\ using converge kl by simp from path_ipd_swap[OF path(1) nret kn] obtain \ m where path\: \is_path \\ and \\: \\ =\<^bsub>n\<^esub> \\ and km: \k < m\ and ipd: \\ m = ipd (\ k)\ \\ l \ {k.. l \ ipd (\ k)\ . have csk1: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> k\ using cs_path_swap_le path path\ \\ kn by auto have suc\: \\ (Suc k) = \ (Suc k)\ by (metis \\ eq_up_to_def kn less_eq_Suc_le) have nret': \\' k' \ return\ by (metis csk last_cs nret) have kn': \k' < n'\ using cs_order[OF path csk converge(2) nret kn] . from path_ipd_swap[OF path(2) nret' kn'] obtain \' m' where path\': \is_path \'\ and \\': \\' =\<^bsub>n'\<^esub> \'\ and km': \k' < m'\ and ipd': \\' m' = ipd (\' k')\ \\ l \ {k'..' l \ ipd (\' k')\ . have csk1': \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\'\<^esup> k'\ using cs_path_swap_le path path\' \\' kn' by auto have suc\': \\' (Suc k') = \' (Suc k')\ by (metis \\' eq_up_to_def kn' less_eq_Suc_le) have icd\: \l icd\<^bsup>\\<^esup>\ k\ using icdi_path_swap_le[OF path\ icd \\] converge by simp have lm: \l < m\ using ipd(1) icd\ km unfolding is_icdi_def is_cdi_def by auto have csk': \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ using csk1 csk1' csk by auto hence kk': \\' k' = \ k\ using last_cs by metis have suc': \\ (Suc k) = \' (Suc k')\ using suc suc\ suc\' by auto have mm': \\' m' = \ m\ using ipd(1) ipd'(1) kk' by auto from cs_ipd[OF ipd km] cs_ipd[OF ipd' km',unfolded mm', folded csk'] have csm: \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ by metis from converged_ipd_same_icd[OF path\ path\' lm csm csk' icd\ suc' ipd'[unfolded kk']] obtain l' where csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ by blast have csl\: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> l\ using \\ converge(1) cs_path_swap_le less_imp_le_nat path(1) path\ by blast have nretl: \\ l \ return\ by (metis icd\ icd_imp_cd ret_no_cd) have csn': \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ using converge(2) cs_path_swap path path\ path\' \\ \\' by auto have ln': \l' < n'\ using cs_order[OF path\ path\' csl csn' nretl converge(1)] . have csl\': \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\'\<^esup> l'\ using cs_path_swap_le[OF path(2) path\' \\'] ln' by auto have csl': \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using csl\ csl\' csl by auto thus \?thesis\ by blast qed lemma cd_is_cs_less: assumes \l cd\<^bsup>\\<^esup>\ k\ shows \cs\<^bsup>\\<^esup> k \ cs\<^bsup>\\<^esup> l\ proof - obtain xs where csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ xs @[\ l]\ using cd_in_cs[OF assms] by blast hence len: \length(cs\<^bsup>\\<^esup> k) < length (cs\<^bsup>\\<^esup> l)\ by auto have take: \take (length (cs\<^bsup>\\<^esup> k)) (cs\<^bsup>\\<^esup> l) = cs\<^bsup>\\<^esup> k\ using csl by auto show \?thesis\ using cs_less.intros[OF len take] . qed lemma cs_select_id: assumes \is_path \\ \\ k \ return\ shows \\\cs\<^bsup>\\<^esup> k = k\ (is \?k = k\) proof - have *: \\ i . cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> k \ i = k\ using cs_inj[OF assms] by metis hence \cs\<^bsup>\\<^esup> ?k = cs\<^bsup>\\<^esup> k\ unfolding cs_select_def using theI[of \\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> k\ \k\] by auto thus \?k = k\ using * by auto qed lemma cs_single_nocd: assumes \cs\<^bsup>\\<^esup> i = [x]\ shows \\ k. \ i cd\<^bsup>\\<^esup>\ k\ proof - have \\ (\ k. i icd\<^bsup>\\<^esup>\ k)\ apply (rule ccontr) using assms cs_not_nil by auto hence \\ (\ k. i cd\<^bsup>\\<^esup>\ k)\ by (metis excd_impl_exicd) thus \?thesis\ by blast qed lemma cs_single_pd_intermed: assumes \is_path \\ \cs\<^bsup>\\<^esup> n = [\ n]\ \k \ n\ shows \\ n pd\ \ k\ proof - have \\ l. \ n icd\<^bsup>\\<^esup>\ l\ by (metis assms(2) cs_single_nocd icd_imp_cd) thus \?thesis\ by (metis assms(1) assms(3) no_icd_pd) qed lemma cs_first_pd: assumes path: \is_path \\ and pd: \\ n pd\ \ 0\ and first: \\ l < n. \ l \ \ n\ shows \cs\<^bsup>\\<^esup> n = [\ n]\ by (metis cs_cases first first_pd_no_cd icd_imp_cd path pd) lemma converged_pd_cs_single: assumes path: \is_path \\ \is_path \'\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ and \0: \\ 0 = \' 0\ and mpdl: \\ m pd\ \ l\ and csl: \cs\<^bsup>\\<^esup> l = [\ l]\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof - have *: \\ l pd\ \' 0\ using cs_single_pd_intermed[OF path(1) csl] \0[symmetric] by auto have \m: \\ m = \' m'\ by (metis converge(2) last_cs) hence **: \\' m' pd\ \ l\ using mpdl by metis obtain l' where lm': \l' \ m'\ and \l: \\' l' = \ l\ (is \?P l'\) using path_pd_pd0[OF path(2) ** *] . let \?l\ = \(LEAST l'. \' l' = \ l)\ have \l': \\' ?l = \ l\ using LeastI[of \?P\,OF \l] . moreover have \\ i ' i \ \ l\ using Least_le[of \?P\] by (metis not_less) hence \\ i ' i \ \' ?l\ using \l' by metis moreover have \\' ?l pd\ \' 0\ using * \l' by metis ultimately have \cs\<^bsup>\'\<^esup> ?l = [\' ?l]\ using cs_first_pd[OF path(2)] by metis thus \?thesis\ using csl \l' by metis qed lemma converged_cs_single: assumes path: \is_path \\ \is_path \'\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ and \0: \\ 0 = \' 0\ and csl: \cs\<^bsup>\\<^esup> l = [\ l]\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof cases assume *: \\ l = return\ hence \\ m = return\ by (metis converge(1) path(1) term_path_stable less_imp_le) hence \cs\<^bsup>\\<^esup> m = [return]\ using cs_return by auto hence \cs\<^bsup>\'\<^esup> m' = [return]\ using converge by simp moreover have \cs\<^bsup>\\<^esup> l = [return]\ using * cs_return by auto ultimately show \?thesis\ by metis next assume nret: \\ l \ return\ have \m: \\ m = \' m'\ by (metis converge(2) last_cs) obtain \\<^sub>1 n where path1: \is_path \\<^sub>1\ and upto: \\ =\<^bsub>m\<^esub> \\<^sub>1\ and \n: \\\<^sub>1 n = return\ using path(1) path_swap_ret by blast obtain \\<^sub>1' n' where path1': \is_path \\<^sub>1'\ and upto': \\' =\<^bsub>m'\<^esub> \\<^sub>1'\ and \n': \\\<^sub>1' n' = return\ using path(2) path_swap_ret by blast have \1l: \\\<^sub>1 l = \ l\ using upto converge(1) by (metis eq_up_to_def nat_less_le) have cs1l: \cs\<^bsup>\\<^sub>1\<^esup> l = cs\<^bsup>\\<^esup> l\ using cs_path_swap_le upto path1 path(1) converge(1) by auto have csl1: \cs\<^bsup>\\<^sub>1\<^esup> l = [\\<^sub>1 l]\ by (metis \1l cs1l csl) have converge1: \cs\<^bsup>\\<^sub>1\<^esup> n = cs\<^bsup>\\<^sub>1'\<^esup> n'\ using \n \n' cs_return by auto have ln: \l < n\ using nret \n \1l term_path_stable[OF path1 \n] by (auto, metis linorder_neqE_nat less_imp_le) have \01: \\\<^sub>1 0 = \\<^sub>1' 0\ using \0 eq_up_to_apply[OF upto] eq_up_to_apply[OF upto'] by auto have pd: \\\<^sub>1 n pd\ \\<^sub>1 l\ using \n by (metis path1 path_nodes return_pd) obtain l' where csl: \cs\<^bsup>\\<^sub>1\<^esup> l = cs\<^bsup>\\<^sub>1'\<^esup> l'\ using converged_pd_cs_single[OF path1 path1' ln converge1 \01 pd csl1] by blast have cs1m: \cs\<^bsup>\\<^sub>1\<^esup> m = cs\<^bsup>\\<^esup> m\ using cs_path_swap upto path1 path(1) by auto have cs1m': \cs\<^bsup>\\<^sub>1'\<^esup> m' = cs\<^bsup>\'\<^esup> m'\ using cs_path_swap upto' path1' path(2) by auto hence converge1: \cs\<^bsup>\\<^sub>1\<^esup> m = cs\<^bsup>\\<^sub>1'\<^esup> m'\ using converge(2) cs1m by metis have nret1: \\\<^sub>1 l \ return\ using nret \1l by auto have lm': \l' < m'\ using cs_order[OF path1 path1' csl converge1 nret1 converge(1)] . have \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^sub>1'\<^esup> l'\ using cs_path_swap_le[OF path(2) path1' upto'] lm' by auto moreover have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^sub>1\<^esup> l\ using cs_path_swap_le[OF path(1) path1 upto] converge(1) by auto ultimately have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using csl by auto thus \?thesis\ by blast qed lemma converged_cd_same_suc: assumes path: \is_path \\ \is_path \'\ and init: \\ 0 = \' 0\ and cd_suc: \\ k k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\\<^esup>\ k \ \ (Suc k) = \' (Suc k')\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using path init cd_suc converge proof (induction \\\ \l\ rule: cs_induct,cases) case (cs \ l) assume *: \\k. l icd\<^bsup>\\<^esup>\ k\ let \?k\ = \THE k. l icd\<^bsup>\\<^esup>\ k\ have icd: \l icd\<^bsup>\\<^esup>\ ?k\ by (metis "*" icd_is_the_icd) hence lcdk: \l cd\<^bsup>\\<^esup>\ ?k\ by (metis is_icdi_def) hence kl: \?k using is_cdi_def by metis have \\ j. ?k cd\<^bsup>\\<^esup>\ j \ l cd\<^bsup>\\<^esup>\ j\ using icd cd_trans is_icdi_def by fast hence suc': \\ j j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j' \ ?k cd\<^bsup>\\<^esup>\ j \ \ (Suc j) = \' (Suc j')\ using cs.prems(4) by blast from cs.IH[OF * cs(2) path(2) cs(4) suc'] cs.prems kl have \\k'. cs\<^bsup>\\<^esup> (THE k. l icd\<^bsup>\\<^esup>\ k) = cs\<^bsup>\'\<^esup> k'\ by (metis Suc_lessD less_trans_Suc) then obtain k' where csk: \cs\<^bsup>\\<^esup> ?k = cs\<^bsup>\'\<^esup> k'\ by blast have suc2: \\ (Suc ?k) = \' (Suc k')\ using cs.prems(4) lcdk csk by auto have km: \?k < m\ using kl cs.prems(5) by simp from converged_same_icd[OF cs(2) path(2) cs.prems(5) cs.prems(6) csk icd suc2] show \?case\ . next case (cs \ l) assume \\ (\k. l icd\<^bsup>\\<^esup>\ k)\ hence \cs\<^bsup>\\<^esup> l = [\ l]\ by auto with cs converged_cs_single show \?case\ by metis qed lemma converged_cd_diverge: assumes path: \is_path \\ \is_path \'\ and init: \\ 0 = \' 0\ and notin: \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ obtains k k' where \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ \l cd\<^bsup>\\<^esup>\ k\ \\ (Suc k) \ \' (Suc k')\ using assms converged_cd_same_suc by blast lemma converged_cd_same_suc_return: assumes path: \is_path \\ \is_path \'\ and \0: \\ 0 = \' 0\ and cd_suc: \\ k k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\\<^esup>\ k \ \ (Suc k) = \' (Suc k')\ and ret: \\' n' = return\ shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\proof cases assume \\ l = return\ hence \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> n'\ using ret cs_return by presburger thus \?thesis\ by blast next assume nretl: \\ l \ return\ have \\ l \ nodes\ using path path_nodes by auto then obtain \l n where ipl: \is_path \l\ and \l: \\ l = \l 0\ and retn: \\l n = return\ and notl: \\ i>0. \l i \ \ l\ by (metis direct_path_return nretl) hence ip: \is_path (\@\<^bsup>l\<^esup> \l)\ and l: \(\@\<^bsup>l\<^esup> \l) l = \ l\ and retl: \(\@\<^bsup>l\<^esup> \l) (l + n) = return\ and nl: \\ i>l. (\@\<^bsup>l\<^esup> \l) i \ \ l\ using path_cons[OF path(1) ipl \l] by auto have \0': \(\@\<^bsup>l\<^esup> \l) 0 = \' 0\ unfolding cs_0 using \l \0 by auto have csn: \cs\<^bsup>\@\<^bsup>l\<^esup> \l\<^esup> (l+n) = cs\<^bsup>\'\<^esup> n'\ using ret retl cs_return by metis have eql: \(\@\<^bsup>l\<^esup> \l) =\<^bsub>l\<^esub> \\ by (metis path_append_eq_up_to) have csl': \cs\<^bsup>\@\<^bsup>l\<^esup> \l\<^esup> l = cs\<^bsup>\\<^esup> l\ using eql cs_path_swap ip path(1) by metis have \0 < n\ using nretl[unfolded \l] retn by (metis neq0_conv) hence ln: \l < l + n\ by simp have *: \\ k k'. cs\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup>\ k \ (\ @\<^bsup>l\<^esup> \l) (Suc k) = \' (Suc k')\ proof (rule,rule,rule) fix k k' assume *: \cs\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup>\ k\ hence kl: \k < l\ using is_cdi_def by auto hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\\<^esup>\ k\ using eql * cs_path_swap_le[OF ip path(1) eql,of \k\] cdi_path_swap_le[OF path(1) _ eql,of \l\ \k\] by auto hence \\ (Suc k) = \' (Suc k')\ using cd_suc by blast then show \(\ @\<^bsup>l\<^esup> \l) (Suc k) = \' (Suc k')\ using cs_path_swap_le[OF ip path(1) eql,of \Suc k\] kl by auto qed obtain l' where \cs\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using converged_cd_same_suc[OF ip path(2) \0' * ln csn] by blast moreover have \cs\<^bsup>\@\<^bsup>l\<^esup> \l\<^esup> l = cs\<^bsup>\\<^esup> l\ using eql by (metis cs_path_swap ip path(1)) ultimately show \?thesis\ by metis qed lemma converged_cd_diverge_return: assumes path: \is_path \\ \is_path \'\ and init: \\ 0 = \' 0\ and notin: \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ and ret: \\' m' = return\ obtains k k' where \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ \l cd\<^bsup>\\<^esup>\ k\ \\ (Suc k) \ \' (Suc k')\ using converged_cd_same_suc_return[OF path init _ ret, of \l\] notin by blast lemma returned_missing_cd_or_loop: assumes path: \is_path \\ \is_path \'\ and \0: \\ 0 = \' 0\ and notin': \\(\ k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ and nret: \\ n'. \' n' \ return\ and ret: \\ n = return\ obtains i i' where \i \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ \\ (Suc i) \ \' (Suc i')\ \k cd\<^bsup>\\<^esup>\ i \ (\ j'> i'. j' cd\<^bsup>\'\<^esup>\ i')\ proof - obtain f where icdf: \\ i'. f (Suc i') icd\<^bsup>\'\<^esup>\ f i'\ and ran: \range f = {i'. \ j'>i'. j' cd\<^bsup>\'\<^esup>\ i'}\ and icdf0: \\ (\i'. f 0 cd\<^bsup>\'\<^esup>\ i')\ using path(2) path_nret_inf_icd_seq nret by blast show \thesis\ proof cases assume \\ j. \ (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j))\ then obtain j where ni\: \\ (\ i. cs\<^bsup>\'\<^esup> (f j) = cs\<^bsup>\\<^esup> i)\ by metis note converged_cd_diverge_return[OF path(2,1) \0[symmetric] ni\ ret] that then obtain i k' where csk: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> k'\ and cdj: \f j cd\<^bsup>\'\<^esup>\ k'\ and div: \\ (Suc i) \ \' (Suc k')\ by metis have \k' \ range f\ using cdj proof (induction \j\) case 0 thus \?case\ using icdf0 by blast next case (Suc j) have icdfj: \f (Suc j) icd\<^bsup>\'\<^esup>\ f j\ using icdf by auto show \?case\ proof cases assume \f (Suc j) icd\<^bsup>\'\<^esup>\ k'\ hence \k' = f j\ using icdfj by (metis icd_uniq) thus \?case\ by auto next assume \\ f (Suc j) icd\<^bsup>\'\<^esup>\ k'\ hence \f j cd\<^bsup>\'\<^esup>\ k'\ using cd_impl_icd_cd[OF Suc.prems icdfj] by auto thus \?case\ using Suc.IH by auto qed qed hence alldep: \\ i'>k'. i' cd\<^bsup>\'\<^esup>\ k'\ using ran by auto show \thesis\ proof cases assume \i < k\ with alldep that[OF _ csk div] show \thesis\ by blast next assume \\ i < k\ hence ki: \k\i\ by auto have \k \ i\ using notin' csk by auto hence ki': \k using ki by auto obtain ka k' where \cs\<^bsup>\\<^esup> ka = cs\<^bsup>\'\<^esup> k'\ \k cd\<^bsup>\\<^esup>\ ka\ \\ (Suc ka) \ \' (Suc k')\ using converged_cd_diverge[OF path \0 notin' ki' csk] by blast moreover hence \ka < k\ unfolding is_cdi_def by auto ultimately show \?thesis\ using that by blast qed next assume \\(\ j. \ (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j)))\ hence allin: \\ j. (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j))\ by blast define f' where f': \f' \ \ j. (SOME i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j))\ have \\ i. f' i < f' (Suc i)\ proof fix i have csi: \cs\<^bsup>\'\<^esup> (f i) = cs\<^bsup>\\<^esup> (f' i)\ unfolding f' using allin by (metis (mono_tags) someI_ex) have cssuci: \cs\<^bsup>\'\<^esup> (f (Suc i)) = cs\<^bsup>\\<^esup> (f' (Suc i))\ unfolding f' using allin by (metis (mono_tags) someI_ex) have fi: \f i < f (Suc i)\ using icdf unfolding is_icdi_def is_cdi_def by auto have \f (Suc i) cd\<^bsup>\'\<^esup>\ f i\ using icdf unfolding is_icdi_def by blast hence nreti: \\' (f i) \ return\ by (metis cd_not_ret) show \f' i < f' (Suc i)\ using cs_order[OF path(2,1) csi cssuci nreti fi] . qed hence kle: \k < f' (Suc k)\ using mono_ge_id[of \f'\ \Suc k\] by auto have cssk: \cs\<^bsup>\\<^esup> (f' (Suc k)) = cs\<^bsup>\'\<^esup> (f (Suc k))\ unfolding f' using allin by (metis (mono_tags) someI_ex) obtain ka k' where \cs\<^bsup>\\<^esup> ka = cs\<^bsup>\'\<^esup> k'\ \k cd\<^bsup>\\<^esup>\ ka\ \\ (Suc ka) \ \' (Suc k')\ using converged_cd_diverge[OF path \0 notin' kle cssk] by blast moreover hence \ka < k\ unfolding is_cdi_def by auto ultimately show \?thesis\ using that by blast qed qed lemma missing_cd_or_loop: assumes path: \is_path \\ \is_path \'\ and \0: \\ 0 = \' 0\ and notin': \\(\ k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ obtains i i' where \i < k\ \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ \\ (Suc i) \ \' (Suc i')\ \k cd\<^bsup>\\<^esup>\ i \ (\ j'> i'. j' cd\<^bsup>\'\<^esup>\ i')\ proof cases assume \\ n'. \' n' = return\ then obtain n' where retn: \\' n' = return\ by blast note converged_cd_diverge_return[OF path \0 notin' retn] then obtain ka k' where \cs\<^bsup>\\<^esup> ka = cs\<^bsup>\'\<^esup> k'\ \k cd\<^bsup>\\<^esup>\ ka\ \\ (Suc ka) \ \' (Suc k')\ by blast moreover hence \ka < k\ unfolding is_cdi_def by auto ultimately show \thesis\ using that by simp next assume \\ (\ n'. \' n' = return)\ hence notret: \\ n'. \' n' \ return\ by auto then obtain \l n where ipl: \is_path \l\ and \l: \\ k = \l 0\ and retn: \\l n = return\ using reaching_ret path(1) path_nodes by metis hence ip: \is_path (\@\<^bsup>k\<^esup>\l)\ and l: \(\@\<^bsup>k\<^esup>\l) k = \ k\ and retl: \(\@\<^bsup>k\<^esup>\l) (k + n) = return\ using path_cons[OF path(1) ipl \l] by auto have \0': \(\@\<^bsup>k\<^esup>\l) 0 = \' 0\ unfolding cs_0 using \l \0 by auto have eql: \(\@\<^bsup>k\<^esup>\l) =\<^bsub>k\<^esub> \\ by (metis path_append_eq_up_to) have csl': \cs\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup> k = cs\<^bsup>\\<^esup> k\ using eql cs_path_swap ip path(1) by metis hence notin: \\(\ k'. cs\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup> k = cs\<^bsup>\'\<^esup> k')\ using notin' by auto obtain i i' where *: \i < k\ and csi: \cs\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup> i = cs\<^bsup>\'\<^esup> i'\ and suci: \(\ @\<^bsup>k\<^esup> \l) (Suc i) \ \' (Suc i')\ and cdloop: \k cd\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup>\ i \ (\ j'>i'. j' cd\<^bsup>\'\<^esup>\ i')\ using returned_missing_cd_or_loop[OF ip path(2) \0' notin notret retl] by blast have \i \ k\ using notin csi by auto hence ik: \i < k\ using * by auto hence \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using csi cs_path_swap_le[OF ip path(1) eql] by auto moreover have \\ (Suc i) \ \' (Suc i')\ using ik eq_up_to_apply[OF eql, of \Suc i\] suci by auto moreover have \k cd\<^bsup>\\<^esup>\ i \ (\ j'>i'. j' cd\<^bsup>\'\<^esup>\ i')\ using cdloop cdi_path_swap_le[OF path(1) _ eql, of \k\ \i\] by auto ultimately show \thesis\ using that[OF *] by blast qed lemma path_shift_set_cd: assumes \is_path \\ shows \{k + j| j . n cd\<^bsup>\\k\<^esup>\ j } = {i. (k+n) cd\<^bsup>\\<^esup>\ i \ k \ i }\ proof - { fix i assume \i\{k+j | j . n cd\<^bsup>\\k\<^esup>\ j }\ then obtain j where \i = k+j\ \n cd\<^bsup>\\k\<^esup>\ j\ by auto hence \k+n cd\<^bsup>\\<^esup>\ i \ k \ i\ using cd_path_shift[OF _ assms, of \k\ \k+j\ \k+n\] by simp hence \i\{ i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i }\ by blast } moreover { fix i assume \i\{ i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i }\ hence *: \k+n cd\<^bsup>\\<^esup>\ i \ k \ i\ by blast then obtain j where i: \i = k+j\ by (metis le_Suc_ex) hence \k+n cd\<^bsup>\\<^esup>\ k+j\ using * by auto hence \n cd\<^bsup>\\k\<^esup>\ j\ using cd_path_shift[OF _ assms, of \k\ \k+j\ \k+n\] by simp hence \i\{k+j | j . n cd\<^bsup>\\k\<^esup>\ j }\ using i by simp } ultimately show \?thesis\ by blast qed lemma cs_path_shift_set_cd: assumes path: \is_path \\ shows \cs\<^bsup>\\k\<^esup> n = map \ (sorted_list_of_set {i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i }) @ [\ (k+n)]\ proof - have mono:\\n m. n < m \ k + n < k + m\ by auto have fin: \finite {i. n cd\<^bsup>\ \ k\<^esup>\ i}\ unfolding is_cdi_def by auto have *: \(\ x. k+x)`{i. n cd\<^bsup>\\k\<^esup>\ i} = {k + i | i. n cd\<^bsup>\\k\<^esup>\ i}\ by auto have \cs\<^bsup>\\k\<^esup> n = map (\\k) (sorted_list_of_set {i. n cd\<^bsup>\\k\<^esup>\ i}) @ [(\\k) n]\ using cs_sorted_list_of_cd' by blast also have \\ = map \ (map (\ x. k+x) (sorted_list_of_set{i. n cd\<^bsup>\\k\<^esup>\ i})) @ [\ (k+n)]\ by auto also have \\ = map \ (sorted_list_of_set ((\ x. k+x)`{i. n cd\<^bsup>\\k\<^esup>\ i})) @ [\ (k+n)]\ using sorted_list_of_set_map_mono[OF mono fin] by auto also have \\ = map \ (sorted_list_of_set ({k + i | i. n cd\<^bsup>\\k\<^esup>\ i})) @ [\ (k+n)]\ using * by auto also have \\ = map \ (sorted_list_of_set ({i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ (k+n)]\ using path_shift_set_cd[OF path] by auto finally show \?thesis\ . qed lemma cs_split_shift_cd: assumes \n cd\<^bsup>\\<^esup>\ j\ and \j < k\ and \k < n\ and \\j'\\<^esup>\ j' \ j' \ j\ shows \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\\k\<^esup> (n-k)\ proof - have path: \is_path \\ using assms unfolding is_cdi_def by auto have 1: \{i. n cd\<^bsup>\\<^esup>\ i} = {i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}\ by auto have le: \\ i\ {i. n cd\<^bsup>\\<^esup>\ i \ i < k}. \ j\ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}. i < j\ by auto have 2: \{i. n cd\<^bsup>\\<^esup>\ i \ i < k} = {i . j cd\<^bsup>\\<^esup>\ i} \ {j}\ proof - { fix i assume \i\{i. n cd\<^bsup>\\<^esup>\ i \ i < k}\ hence cd: \n cd\<^bsup>\\<^esup>\ i\ and ik:\i < k\ by auto have \i\{i . j cd\<^bsup>\\<^esup>\ i} \ {j}\ proof cases assume \i < j\ hence \j cd\<^bsup>\\<^esup>\ i\ by (metis is_cdi_def assms(1) cd cdi_prefix nat_less_le) thus \?thesis\ by simp next assume \\ i < j\ moreover have \i \ j\ using assms(4) ik cd by auto ultimately show \?thesis\ by auto qed } moreover { fix i assume \i\{i . j cd\<^bsup>\\<^esup>\ i} \ {j}\ hence \j cd\<^bsup>\\<^esup>\ i \ i = j\ by auto hence \i\{i. n cd\<^bsup>\\<^esup>\ i \ i < k}\ using assms(1,2) cd_trans[OF _ assms(1)] apply auto unfolding is_cdi_def by (metis (poly_guards_query) diff_diff_cancel diff_is_0_eq le_refl le_trans nat_less_le) } ultimately show \?thesis\ by blast qed have \cs\<^bsup>\\<^esup> n = map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i}) @ [\ n]\ using cs_sorted_list_of_cd' by simp also have \\ = map \ (sorted_list_of_set ({i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ using 1 by metis also have \\ = map \ ((sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ i < k}) @ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ using sorted_list_of_set_append[OF _ _ le] is_cdi_def by auto also have \\ = (map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ i < k})) @ (map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ by auto also have \\ = cs\<^bsup>\\<^esup> j @ (map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ unfolding 2 using cs_sorted_list_of_cd by auto also have \\ = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\\k\<^esup> (n-k)\ using cs_path_shift_set_cd[OF path, of \k\ \n-k\] assms(3) by auto finally show \?thesis\ . qed lemma cs_split_shift_nocd: assumes \is_path \\ and \k < n\ and \\j. n cd\<^bsup>\\<^esup>\ j \ k \ j\ shows \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\k\<^esup> (n-k)\ proof - have path: \is_path \\ using assms by auto have 1: \{i. n cd\<^bsup>\\<^esup>\ i} = {i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}\ by auto have le: \\ i\ {i. n cd\<^bsup>\\<^esup>\ i \ i < k}. \ j\ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}. i < j\ by auto have 2: \{i. n cd\<^bsup>\\<^esup>\ i \ i < k} = {}\ using assms by auto have \cs\<^bsup>\\<^esup> n = map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i}) @ [\ n]\ using cs_sorted_list_of_cd' by simp also have \\ = map \ (sorted_list_of_set ({i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ using 1 by metis also have \\ = map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}) @ [\ n]\ unfolding 2 by auto also have \\ = cs\<^bsup>\\k\<^esup> (n-k)\ using cs_path_shift_set_cd[OF path, of \k\ \n-k\] assms(2) by auto finally show \?thesis\ . qed lemma shifted_cs_eq_is_eq: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and \cs\<^bsup>\\k\<^esup> n = cs\<^bsup>\'\k'\<^esup> n'\ shows \cs\<^bsup>\\<^esup> (k+n) = cs\<^bsup>\'\<^esup> (k'+n')\ proof (rule ccontr) note path = assms(1,2) note csk = assms(3) note csn = assms(4) assume ne: \cs\<^bsup>\\<^esup> (k+n) \ cs\<^bsup>\'\<^esup> (k'+n')\ have nretkn:\\ (k+n) \ return\ proof assume 1:\\ (k+n) = return\ hence \(\\k) n = return\ by auto hence \(\'\k') n' = return\ using last_cs assms(4) by metis hence \\' (k' + n') = return\ by auto thus \False\ using ne 1 cs_return by auto qed hence nretk: \\ k \ return\ using term_path_stable[OF assms(1), of \k\ \k +n\] by auto have nretkn': \\' (k'+n') \ return\ proof assume 1:\\' (k'+n') = return\ hence \(\'\k') n' = return\ by auto hence \(\\k) n = return\ using last_cs assms(4) by metis hence \\ (k + n) = return\ by auto thus \False\ using ne 1 cs_return by auto qed hence nretk': \\' k' \ return\ using term_path_stable[OF assms(2), of \k'\ \k' +n'\] by auto have n0:\n > 0\ proof (rule ccontr) assume *: \\ 0 < n\ hence 1:\cs\<^bsup>\\k\<^esup> 0 = cs\<^bsup>\'\k'\<^esup> n'\ using assms(3,4) by auto have \(\\k) 0 = (\'\k') 0\ using assms(3) last_cs path_shift_def by (metis monoid_add_class.add.right_neutral) hence \cs\<^bsup>\'\k'\<^esup> 0 = cs\<^bsup>\'\k'\<^esup> n'\ using 1 cs_0 by metis hence n0': \n' = 0\ using cs_inj[of \\'\k'\ \0\ \n'\ ] * assms(2) by (metis path_shift_def assms(4) last_cs nretkn path_path_shift) thus \False\ using ne * assms(3) by fastforce qed have n0':\n' > 0\ proof (rule ccontr) assume *: \\ 0 < n'\ hence 1:\cs\<^bsup>\'\k'\<^esup> 0 = cs\<^bsup>\\k\<^esup> n\ using assms(3,4) by auto have \(\'\k') 0 = (\\k) 0\ using assms(3) last_cs path_shift_def by (metis monoid_add_class.add.right_neutral) hence \cs\<^bsup>\\k\<^esup> 0 = cs\<^bsup>\\k\<^esup> n\ using 1 cs_0 by metis hence n0: \n = 0\ using cs_inj[of \\\k\ \0\ \n\ ] * assms(1) by (metis path_shift_def assms(4) last_cs nretkn path_path_shift) thus \False\ using ne * assms(3) by fastforce qed have cdleswap': \\ j'\'\<^esup>\ j' \ (\j\\<^esup>\ j \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ proof (rule,rule,rule, rule ccontr) fix j' assume jk': \j' and ncdj': \(k'+n') cd\<^bsup>\'\<^esup>\ j'\ and ne: \\ (\j\\<^esup>\ j \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ hence kcdj': \k' cd\<^bsup>\'\<^esup>\ j'\ using cr_wn' by blast then obtain j where kcdj: \k cd\<^bsup>\\<^esup>\ j\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using csk cs_path_swap_cd path by metis hence jk: \j < k\ unfolding is_cdi_def by auto have ncdn: \\ (k+n) cd\<^bsup>\\<^esup>\ j\ using ne csj jk by blast obtain l' where lnocd': \l' = n' \ n' cd\<^bsup>\'\k'\<^esup>\ l'\ and cslsing': \cs\<^bsup>\'\k'\<^esup> l' = [(\'\k') l']\ proof cases assume \cs\<^bsup>\'\k'\<^esup> n' = [(\'\k') n']\ thus \thesis\ using that[of \n'\] by auto next assume *: \cs\<^bsup>\'\k'\<^esup> n' \ [(\'\k') n']\ then obtain x ys where \cs\<^bsup>\'\k'\<^esup> n' = [x]@ys@[(\'\k') n']\ by (metis append_Cons append_Nil cs_length_g_one cs_length_one(1) neq_Nil_conv) then obtain l' where \cs\<^bsup>\'\k'\<^esup> l' = [x]\ and cdl': \n' cd\<^bsup>\'\k'\<^esup>\ l'\ using cs_split[of \\'\k'\ \n'\ \Nil\ \x\ \ys\] by auto hence \cs\<^bsup>\'\k'\<^esup> l' = [(\'\k') l']\ using last_cs by (metis last.simps) thus \thesis\ using that cdl' by auto qed hence ln': \l'\n'\ unfolding is_cdi_def by auto hence lcdj': \k'+l' cd\<^bsup>\'\<^esup>\ j'\ using jk' ncdj' by (metis add_le_cancel_left cdi_prefix trans_less_add1) obtain l where lnocd: \l = n \ n cd\<^bsup>\\k\<^esup>\ l\ and csl: \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using lnocd' proof assume \l' = n'\ thus \thesis\ using csn that[of \n\] by auto next assume \n' cd\<^bsup>\'\k'\<^esup>\ l'\ then obtain l where \n cd\<^bsup>\\k\<^esup>\ l\ \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using cs_path_swap_cd path csn by (metis path_path_shift) thus \thesis\ using that by auto qed have cslsing: \cs\<^bsup>\\k\<^esup> l = [(\\k) l]\ using cslsing' last_cs csl last.simps by metis have ln: \l\n\ using lnocd unfolding is_cdi_def by auto hence nretkl: \\ (k + l) \ return\ using term_path_stable[of \\\ \k+l\ \k+n\] nretkn path(1) by auto have *: \n cd\<^bsup>\\k\<^esup>\ l \ k+n cd\<^bsup>\\<^esup>\ k+l\ using cd_path_shift[of \k\ \k+l\ \\\ \k+n\] path(1) by auto have ncdl: \\ (k+l) cd\<^bsup>\\<^esup>\ j\ apply rule using lnocd apply rule using ncdn apply blast using cd_trans ncdn * by blast hence \\ i\ {j..k+l}. \ i = ipd (\ j)\ unfolding is_cdi_def using path(1) jk nretkl by auto hence \\ i\ {k<..k+l}. \ i = ipd (\ j)\ using kcdj unfolding is_cdi_def by force then obtain i where ki: \k < i\ and il: \i \ k+l\ and ipdi: \\ i = ipd (\ j)\ by force hence \(\\k) (i-k) = ipd (\ j)\ \i-k \ l\ by auto hence pd: \(\\k) l pd\ ipd (\ j)\ using cs_single_pd_intermed[OF _ cslsing] path(1) path_path_shift by metis moreover have \(\\k) l = \' (k' + l')\ using csl last_cs by (metis path_shift_def) moreover have \\ j = \' j'\ using csj last_cs by metis ultimately have \\' (k'+l') pd\ ipd (\' j')\ by simp moreover have \ipd (\' j') pd\ \' (k'+l')\ using ipd_pd_cd[OF lcdj'] . ultimately have \\' (k'+l') = ipd (\' j')\ using pd_antisym by auto thus \False\ using lcdj' unfolding is_cdi_def by force qed \ \Symmetric version of the above statement\ have cdleswap: \\ j\\<^esup>\ j \ (\j'\'\<^esup>\ j' \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ proof (rule,rule,rule, rule ccontr) fix j assume jk: \j and ncdj: \(k+n) cd\<^bsup>\\<^esup>\ j\ and ne: \\ (\j'\'\<^esup>\ j' \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ hence kcdj: \k cd\<^bsup>\\<^esup>\ j\ using cr_wn' by blast then obtain j' where kcdj': \k' cd\<^bsup>\'\<^esup>\ j'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using csk cs_path_swap_cd path by metis hence jk': \j' < k'\ unfolding is_cdi_def by auto have ncdn': \\ (k'+n') cd\<^bsup>\'\<^esup>\ j'\ using ne csj jk' by blast obtain l where lnocd: \l = n \ n cd\<^bsup>\\k\<^esup>\ l\ and cslsing: \cs\<^bsup>\\k\<^esup> l = [(\\k) l]\ proof cases assume \cs\<^bsup>\\k\<^esup> n = [(\\k) n]\ thus \thesis\ using that[of \n\] by auto next assume *: \cs\<^bsup>\\k\<^esup> n \ [(\\k) n]\ then obtain x ys where \cs\<^bsup>\\k\<^esup> n = [x]@ys@[(\\k) n]\ by (metis append_Cons append_Nil cs_length_g_one cs_length_one(1) neq_Nil_conv) then obtain l where \cs\<^bsup>\\k\<^esup> l = [x]\ and cdl: \n cd\<^bsup>\\k\<^esup>\ l\ using cs_split[of \\\k\ \n\ \Nil\ \x\ \ys\] by auto hence \cs\<^bsup>\\k\<^esup> l = [(\\k) l]\ using last_cs by (metis last.simps) thus \thesis\ using that cdl by auto qed hence ln: \l\n\ unfolding is_cdi_def by auto hence lcdj: \k+l cd\<^bsup>\\<^esup>\ j\ using jk ncdj by (metis add_le_cancel_left cdi_prefix trans_less_add1) obtain l' where lnocd': \l' = n' \ n' cd\<^bsup>\'\k'\<^esup>\ l'\ and csl: \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using lnocd proof assume \l = n\ thus \thesis\ using csn that[of \n'\] by auto next assume \n cd\<^bsup>\\k\<^esup>\ l\ then obtain l' where \n' cd\<^bsup>\'\k'\<^esup>\ l'\ \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using cs_path_swap_cd path csn by (metis path_path_shift) thus \thesis\ using that by auto qed have cslsing': \cs\<^bsup>\'\k'\<^esup> l' = [(\'\k') l']\ using cslsing last_cs csl last.simps by metis have ln': \l'\n'\ using lnocd' unfolding is_cdi_def by auto hence nretkl': \\' (k' + l') \ return\ using term_path_stable[of \\'\ \k'+l'\ \k'+n'\] nretkn' path(2) by auto have *: \n' cd\<^bsup>\'\k'\<^esup>\ l' \ k'+n' cd\<^bsup>\'\<^esup>\ k'+l'\ using cd_path_shift[of \k'\ \k'+l'\ \\'\ \k'+n'\] path(2) by auto have ncdl': \\ (k'+l') cd\<^bsup>\'\<^esup>\ j'\ apply rule using lnocd' apply rule using ncdn' apply blast using cd_trans ncdn' * by blast hence \\ i'\ {j'..k'+l'}. \' i' = ipd (\' j')\ unfolding is_cdi_def using path(2) jk' nretkl' by auto hence \\ i'\ {k'<..k'+l'}. \' i' = ipd (\' j')\ using kcdj' unfolding is_cdi_def by force then obtain i' where ki': \k' < i'\ and il': \i' \ k'+l'\ and ipdi: \\' i' = ipd (\' j')\ by force hence \(\'\k') (i'-k') = ipd (\' j')\ \i'-k' \ l'\ by auto hence pd: \(\'\k') l' pd\ ipd (\' j')\ using cs_single_pd_intermed[OF _ cslsing'] path(2) path_path_shift by metis moreover have \(\'\k') l' = \ (k + l)\ using csl last_cs by (metis path_shift_def) moreover have \\' j' = \ j\ using csj last_cs by metis ultimately have \\ (k+l) pd\ ipd (\ j)\ by simp moreover have \ipd (\ j) pd\ \ (k+l)\ using ipd_pd_cd[OF lcdj] . ultimately have \\ (k+l) = ipd (\ j)\ using pd_antisym by auto thus \False\ using lcdj unfolding is_cdi_def by force qed have cdle: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ (is \\ j. ?P j\) proof (rule ccontr) assume \\ (\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k)\ hence allge: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ k \ j\ by auto have allge': \\j'. (k'+n') cd\<^bsup>\'\<^esup>\ j' \ k' \ j'\ proof (rule, rule, rule ccontr) fix j' assume *: \k' + n' cd\<^bsup>\'\<^esup>\ j'\ and \\ k' \ j'\ then obtain j where \j \(k+n) cd\<^bsup>\\<^esup>\ j\ using cdleswap' by (metis le_neq_implies_less nat_le_linear) thus \False\ using allge by auto qed have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_nocd[OF assms(1) _ allge] n0 by auto moreover have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_nocd[OF assms(2) _ allge'] n0' by auto ultimately show \False\ using ne assms(4) by auto qed define j where \j == GREATEST j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ have cdj:\(k+n) cd\<^bsup>\\<^esup>\ j\ and jk: \j < k\ and jge:\\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ proof - have bound: \\ y. ?P y \ y \ k\ by auto show \(k+n) cd\<^bsup>\\<^esup>\ j\ using GreatestI_nat[of \?P\] j_def bound cdle by blast show \j < k\ using GreatestI_nat[of \?P\] bound j_def cdle by blast show \\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ using Greatest_le_nat[of \?P\] bound j_def by blast qed obtain j' where cdj':\(k'+n') cd\<^bsup>\'\<^esup>\ j'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and jk': \j' < k'\ using cdleswap cdj jk by blast have jge':\\ i'< k'. (k'+n') cd\<^bsup>\'\<^esup>\ i' \ i' \ j'\ proof(rule,rule,rule) fix i' assume ik': \i' < k'\ and cdi': \k' + n' cd\<^bsup>\'\<^esup>\ i'\ then obtain i where cdi:\(k+n) cd\<^bsup>\\<^esup>\ i\ and csi: \ cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ and ik: \i using cdleswap' by force have ij: \i \ j\ using jge cdi ik by auto show \i' \ j'\ using cs_order_le[OF assms(1,2) csi[symmetric] csj _ ij] cd_not_ret[OF cdi] by simp qed have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_cd[OF cdj jk _ jge] n0 by auto moreover have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\'\<^esup> j' @ cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_cd[OF cdj' jk' _ jge'] n0' by auto ultimately have \cs\<^bsup>\\<^esup> (k+n) = cs\<^bsup>\'\<^esup> (k'+n')\ using csj assms(4) by auto thus \False\ using ne by simp qed lemma cs_eq_is_eq_shifted: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and \cs\<^bsup>\\<^esup> (k+n) = cs\<^bsup>\'\<^esup> (k'+n')\ shows \cs\<^bsup>\\k\<^esup> n = cs\<^bsup>\'\k'\<^esup> n'\ proof (rule ccontr) assume ne: \cs\<^bsup>\ \ k\<^esup> n \ cs\<^bsup>\' \ k'\<^esup> n'\ have nretkn:\\ (k+n) \ return\ proof assume 1:\\ (k+n) = return\ hence 2:\\' (k'+n') = return\ using assms(4) last_cs by metis hence \(\\k) n = return\ \(\'\k') n' = return\ using 1 by auto hence \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_return by metis thus \False\ using ne by simp qed hence nretk: \\ k \ return\ using term_path_stable[OF assms(1), of \k\ \k +n\] by auto have nretkn': \\' (k'+n') \ return\ proof assume 1:\\' (k'+n') = return\ hence 2:\\ (k+n) = return\ using assms(4) last_cs by metis hence \(\\k) n = return\ \(\'\k') n' = return\ using 1 by auto hence \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_return by metis thus \False\ using ne by simp qed hence nretk': \\' k' \ return\ using term_path_stable[OF assms(2), of \k'\ \k' +n'\] by auto have n0:\n > 0\ proof (rule ccontr) assume *: \\ 0 < n\ hence \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\'\<^esup> (k'+ n')\ using assms(3,4) by auto hence n0: \n = 0\ \n' = 0\ using cs_inj[OF assms(2) nretkn', of \k'\] * by auto have \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ unfolding n0 cs_0 by (auto , metis last_cs assms(3)) thus \False\ using ne by simp qed have n0':\n' > 0\ proof (rule ccontr) assume *: \\ 0 < n'\ hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> (k+ n)\ using assms(3,4) by auto hence n0: \n = 0\ \n' = 0\ using cs_inj[OF assms(1) nretkn, of \k\] * by auto have \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ unfolding n0 cs_0 by (auto , metis last_cs assms(3)) thus \False\ using ne by simp qed have cdle: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ (is \\ j. ?P j\) proof (rule ccontr) assume \\ (\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k)\ hence allge: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ k \ j\ by auto have allge': \\j'. (k'+n') cd\<^bsup>\'\<^esup>\ j' \ k' \ j'\ proof (rule, rule) fix j' assume *: \k' + n' cd\<^bsup>\'\<^esup>\ j'\ obtain j where cdj: \k+n cd\<^bsup>\\<^esup>\ j\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using cs_path_swap_cd[OF assms(2,1) assms(4)[symmetric] *] by metis hence kj:\k \ j\ using allge by auto thus kj': \k' \ j'\ using cs_order_le[OF assms(1,2,3) csj nretk] by simp qed have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_nocd[OF assms(1) _ allge] n0 by auto moreover have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_nocd[OF assms(2) _ allge'] n0' by auto ultimately show \False\ using ne assms(4) by auto qed define j where \j == GREATEST j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ have cdj:\(k+n) cd\<^bsup>\\<^esup>\ j\ and jk: \j < k\ and jge:\\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ proof - have bound: \\ y. ?P y \ y \ k\ by auto show \(k+n) cd\<^bsup>\\<^esup>\ j\ using GreatestI_nat[of \?P\] bound j_def cdle by blast show \j < k\ using GreatestI_nat[of \?P\] bound j_def cdle by blast show \\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ using Greatest_le_nat[of \?P\] bound j_def by blast qed obtain j' where cdj':\(k'+n') cd\<^bsup>\'\<^esup>\ j'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using cs_path_swap_cd assms cdj by blast have jge':\\ i'< k'. (k'+n') cd\<^bsup>\'\<^esup>\ i' \ i' \ j'\ proof(rule,rule,rule) fix i' assume ik': \i' < k'\ and cdi': \k' + n' cd\<^bsup>\'\<^esup>\ i'\ then obtain i where cdi:\(k+n) cd\<^bsup>\\<^esup>\ i\ and csi: \ cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ using cs_path_swap_cd[OF assms(2,1) assms(4)[symmetric]] by blast have nreti': \\' i' \ return\ by (metis cd_not_ret cdi') have ik: \i < k\ using cs_order[OF assms(2,1) csi _ nreti' ik'] assms(3) by auto have ij: \i \ j\ using jge cdi ik by auto show \i' \ j'\ using cs_order_le[OF assms(1,2) csi[symmetric] csj _ ij] cd_not_ret[OF cdi] by simp qed have jk': \j' < k'\ using cs_order[OF assms(1,2) csj assms(3) cd_not_ret[OF cdj] jk] . have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_cd[OF cdj jk _ jge] n0 by auto moreover have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\'\<^esup> j' @ cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_cd[OF cdj' jk' _ jge'] n0' by auto ultimately have \cs\<^bsup>\\k\<^esup> n = cs\<^bsup>\'\k'\<^esup> n'\ using csj assms(4) by auto thus \False\ using ne by simp qed lemma converged_cd_diverge_cs: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and \j and \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ and \l < m\ and \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ obtains k k' where \j\k\ \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and \l cd\<^bsup>\\<^esup>\ k\ and \\ (Suc k) \ \' (Suc k')\ proof - have \is_path (\\j)\ \is_path (\'\j')\ using assms(1,2) path_path_shift by auto moreover have \(\\j) 0 = (\'\j') 0\ using assms(3) last_cs by (metis path_shift_def add.right_neutral) moreover have \\(\l'. cs\<^bsup>\\j\<^esup> (l-j) = cs\<^bsup>\'\j'\<^esup> l')\ proof assume \\l'. cs\<^bsup>\ \ j\<^esup> (l - j) = cs\<^bsup>\' \ j'\<^esup> l'\ then obtain l' where csl: \cs\<^bsup>\\j\<^esup> (l - j) = cs\<^bsup>\'\j'\<^esup> l'\ by blast have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> (j' + l')\ using shifted_cs_eq_is_eq[OF assms(1,2,3) csl] assms(4) by auto thus \False\ using assms(5) by blast qed moreover have \l-j < m-j\ using assms by auto moreover have \\ j \ return\ using cs_return assms(1-5) term_path_stable by (metis nat_less_le) hence \j' using cs_order[OF assms(1,2,3,7)] assms by auto hence \cs\<^bsup>\\j\<^esup> (m-j) = cs\<^bsup>\'\j'\<^esup> (m'-j')\ using cs_eq_is_eq_shifted[OF assms(1,2,3),of \m-j\ \m'-j'\] assms(4,6,7) by auto ultimately obtain k k' where csk: \cs\<^bsup>\\j\<^esup> k = cs\<^bsup>\'\j'\<^esup> k'\ and lcdk: \l-j cd\<^bsup>\\j\<^esup>\ k\ and suc:\(\\j) (Suc k) \ (\'\j') (Suc k')\ using converged_cd_diverge by blast have \cs\<^bsup>\\<^esup> (j+k) = cs\<^bsup>\'\<^esup> (j'+k')\ using shifted_cs_eq_is_eq[OF assms(1-3) csk] . moreover have \l cd\<^bsup>\\<^esup>\ j+k\ using lcdk assms(1,2,4) by (metis add.commute add_diff_cancel_right' cd_path_shift le_add1) moreover have \\ (Suc (j+k)) \ \' (Suc (j'+ k'))\ using suc by auto moreover have \j \ j+k\ by auto ultimately show \thesis\ using that[of \j+k\ \j'+k'\] by auto qed lemma cs_ipd_conv: assumes csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and ipd: \\ l = ipd (\ k)\ \\' l' = ipd(\' k')\ and nipd: \\n\{k.. n \ ipd (\ k)\ \\n'\{k'..' n' \ ipd (\' k')\ and kl: \k < l\ \k' < l'\ shows \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using cs_ipd[OF ipd(1) nipd(1) kl(1)] cs_ipd[OF ipd(2) nipd(2) kl(2)] csk ipd by (metis (no_types) last_cs) lemma cp_eq_cs: assumes \((\,k),(\',k'))\cp\ shows \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ using assms apply(induction rule: cp.induct) apply blast+ apply simp done lemma cd_cs_swap: assumes \l cd\<^bsup>\\<^esup>\ k\ \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ shows \l' cd\<^bsup>\'\<^esup>\ k'\ proof - have \\ i. l icd\<^bsup>\\<^esup>\ i\ using assms(1) excd_impl_exicd by blast hence \cs\<^bsup>\\<^esup> l \ [\ l]\ by auto hence \cs\<^bsup>\'\<^esup> l' \ [\' l']\ using assms last_cs by metis hence \\ i'. l' icd\<^bsup>\'\<^esup>\ i'\ by (metis cs_cases) hence path': \is_path \'\ unfolding is_icdi_def is_cdi_def by auto from cd_in_cs[OF assms(1)] obtain ys where csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ ys @ [\ l]\ by blast obtain xs where csk: \cs\<^bsup>\\<^esup> k = xs@[\ k]\ by (metis append_butlast_last_id cs_not_nil last_cs) have \l: \\ l = \' l'\ using assms last_cs by metis have csl': \cs\<^bsup>\'\<^esup> l' = xs@[\ k]@ys@[\' l']\ by (metis \l append_eq_appendI assms(2) csk csl) from cs_split[of \\'\ \l'\ \xs\ \\ k\ \ys\] obtain m where csm: \cs\<^bsup>\'\<^esup> m = xs @ [\ k]\ and lcdm: \l' cd\<^bsup>\'\<^esup>\ m\ using csl' by metis have csm': \cs\<^bsup>\'\<^esup> m = cs\<^bsup>\'\<^esup> k'\ by (metis assms(3) csk csm) have \\' m \ return\ using lcdm unfolding is_cdi_def using term_path_stable by (metis nat_less_le) hence \m = k'\ using cs_inj path' csm' by auto thus \?thesis\ using lcdm by auto qed subsection \Facts about Observations\ lemma kth_obs_not_none: assumes \is_kth_obs (path \) k i\ obtains a where \obsp \ i = Some a\ using assms unfolding is_kth_obs_def obsp_def by auto lemma kth_obs_unique: \is_kth_obs \ k i \ is_kth_obs \ k j \ i = j\ proof (induction \i\ \j\ rule: nat_sym_cases) case sym thus \?case\ by simp next case eq thus \?case\ by simp next case (less i j) have \obs_ids \ \ {.. obs_ids \ \ {.. using less(1) by auto moreover have \i \ obs_ids \ \ {.. using less unfolding is_kth_obs_def obs_ids_def by auto moreover have \i \ obs_ids \ \ {.. by auto moreover have \card (obs_ids \ \ {.. \ {.. using less.prems unfolding is_kth_obs_def by auto moreover have \finite (obs_ids \ \ {.. \finite (obs_ids \ \ {.. by auto ultimately have \False\ by (metis card_subset_eq) thus \?case\ .. qed lemma obs_none_no_kth_obs: assumes \obs \ k = None\ shows \\ (\ i. is_kth_obs (path \) k i)\ apply rule using assms unfolding obs_def obsp_def apply (auto split: option.split_asm) by (metis assms kth_obs_not_none kth_obs_unique obs_def option.distinct(2) the_equality) lemma obs_some_kth_obs : assumes \obs \ k \ None\ obtains i where \is_kth_obs (path \) k i\ by (metis obs_def assms) lemma not_none_is_obs: assumes \att(\ i) \ None\ shows \is_kth_obs \ (card (obs_ids \ \ {.. unfolding is_kth_obs_def using assms by auto lemma in_obs_ids_is_kth_obs: assumes \i \ obs_ids \\ obtains k where \is_kth_obs \ k i\ proof have \att (\ i) \ None\ using assms obs_ids_def by auto thus \is_kth_obs \ (card (obs_ids \ \ {.. using not_none_is_obs by auto qed lemma kth_obs_stable: assumes \is_kth_obs \ l j\ \k < l\ shows \\ i. is_kth_obs \ k i\ using assms proof (induction \l\ arbitrary: \j\ rule: less_induct ) case (less l j) have cardl: \card (obs_ids \ \ {.. using less is_kth_obs_def by auto then obtain i where ex: \i \ obs_ids \ \ {.. (is \?P i\) using less(3) by (metis card.empty empty_iff less_irrefl subsetI subset_antisym zero_diff zero_less_diff) have bound: \\ i. i \ obs_ids \ \ {.. i \ j\ by auto let \?i\ = \GREATEST i. i \ obs_ids \ \ {.. have *: \?i < j\ \?i \ obs_ids \\ using GreatestI_nat[of \?P\ \i\ \j\] ex bound by auto have **: \\ i. i \ obs_ids \ \ i i \ ?i\ using Greatest_le_nat[of \?P\ _ \j\] ex bound by auto have \(obs_ids \ \ {.. {?i} = obs_ids \ \ {.. apply rule apply auto using *[simplified] apply simp+ using **[simplified] by auto moreover have \?i \ (obs_ids \ \ {.. by auto ultimately have \Suc (card (obs_ids \ \ {.. using cardl by (metis Un_empty_right Un_insert_right card_insert_disjoint finite_Int finite_lessThan) hence \card (obs_ids \ \ {.. by auto hence iko: \is_kth_obs \ (l - 1) ?i\ using *(2) unfolding is_kth_obs_def obs_ids_def by auto have ll: \l - 1 < l\ by (metis One_nat_def diff_Suc_less less.prems(2) not_gr0 not_less0) note IV=less(1)[OF ll iko] show \?thesis\ proof cases assume \k < l - 1\ thus \?thesis\ using IV by simp next assume \\ k < l - 1\ hence \k = l - 1\ using less by auto thus \?thesis\ using iko by blast qed qed lemma kth_obs_mono: assumes \is_kth_obs \ k i\ \is_kth_obs \ l j\ \k < l\ shows \i < j\ proof (rule ccontr) assume \\ i < j\ hence \{.. {.. by auto hence \obs_ids \ \ {.. obs_ids \ \ {.. by auto moreover have \finite (obs_ids \ \ {.. by auto ultimately have \card (obs_ids \ \ {.. card (obs_ids \ \ {.. by (metis card_mono) thus \False\ using assms unfolding is_kth_obs_def by auto qed lemma kth_obs_le_iff: assumes \is_kth_obs \ k i\ \is_kth_obs \ l j\ shows \k < l \ i < j\ by (metis assms kth_obs_unique kth_obs_mono not_less_iff_gr_or_eq) lemma ret_obs_all_obs: assumes path: \is_path \\ and iki: \is_kth_obs \ k i\ and ret: \\ i = return\ and kl: \k < l\ obtains j where \is_kth_obs \ l j\ proof- show \thesis\ using kl iki ret proof (induction \l - k\ arbitrary: \k\ \i\ rule: less_induct) case (less k i) note kl = \k < l\ note iki = \is_kth_obs \ k i\ note ret = \\ i = return\ have card: \card (obs_ids \ \ {.. and att_ret: \att return \ None\using iki ret unfolding is_kth_obs_def by auto have rets: \\ (Suc i) = return\ using path ret term_path_stable by auto hence attsuc: \att (\ (Suc i)) \ None\ using att_ret by auto hence *: \i \ obs_ids \\ using att_ret ret unfolding obs_ids_def by auto have \{..< Suc i} = insert i {.. by auto hence a: \obs_ids \ \ {..< Suc i} = insert i (obs_ids \ \ {.. using * by auto have b: \i \ obs_ids \ \ {.. by auto have \finite (obs_ids \ \ {.. by auto hence \card (obs_ids \ \ {.. by (metis card card_insert_disjoint a b) hence iksuc: \is_kth_obs \ (Suc k) (Suc i)\ using attsuc unfolding is_kth_obs_def by auto have suckl: \Suc k \ l\ using kl by auto note less thus \thesis\ proof (cases \Suc k < l\) assume skl: \Suc k < l\ from less(1)[OF _ skl iksuc rets] skl show \thesis\ by auto next assume \\ Suc k < l\ hence \Suc k = l\ using suckl by auto thus \thesis\ using iksuc that by auto qed qed qed lemma no_kth_obs_missing_cs: assumes path: \is_path \\ \is_path \'\ and iki: \is_kth_obs \ k i\ and not_in_\': \\(\i'. is_kth_obs \' k i')\ obtains l j where \is_kth_obs \ l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ proof (rule ccontr) assume \\ thesis\ hence all_in_\': \\ l j. is_kth_obs \ l j \ (\ j' . cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ using that by blast then obtain i' where csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using assms by blast hence \att(\' i') \ None\ using iki by (metis is_kth_obs_def last_cs) then obtain k' where ik': \is_kth_obs \' k' i'\ by (metis not_none_is_obs) hence kk': \k' < k\ using not_in_\' kth_obs_stable by (auto, metis not_less_iff_gr_or_eq) show \False\ proof (cases \\ i = return\) assume \\ i \ return\ thus \False\ using kk' ik' csi iki proof (induction \k\ arbitrary: \i\ \i'\ \k'\ ) case 0 thus \?case\ by simp next case (Suc k i i' k') then obtain j where ikj: \is_kth_obs \ k j\ by (metis kth_obs_stable lessI) then obtain j' where csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using all_in_\' by blast hence \att(\' j') \ None\ using ikj by (metis is_kth_obs_def last_cs) then obtain k2 where ik2: \is_kth_obs \' k2 j'\ by (metis not_none_is_obs) have ji: \j < i\ using kth_obs_mono [OF ikj \is_kth_obs \ (Suc k) i\] by auto hence nretj: \\ j \ return\ using Suc(2) term_path_stable less_imp_le path(1) by metis have ji': \j' < i'\ using cs_order[OF path _ _ nretj, of \j'\ \i\ \i'\] csj \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ ji by auto have \k2 \ k'\ using ik2 Suc(4) ji' kth_obs_unique[of \\'\ \k'\ \i'\ \j'\] by (metis less_irrefl) hence k2k': \k2 < k'\ using kth_obs_mono[OF \is_kth_obs \' k' i'\ ik2] ji' by (metis not_less_iff_gr_or_eq) hence k2k: \k2 < k\ using Suc by auto from Suc.IH[OF nretj k2k ik2 csj ikj] show \False\ . qed next assume \\ i = return\ hence reti': \\' i' = return\ by (metis csi last_cs) from ret_obs_all_obs[OF path(2) ik' reti' kk', of \False\] not_in_\' show \False\ by blast qed qed lemma kth_obs_cs_missing_cs: assumes path: \is_path \\ \is_path \'\ and iki: \is_kth_obs \ k i\ and iki': \is_kth_obs \' k i'\ and csi: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ obtains l j where \j \ i\ \is_kth_obs \ l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ | l' j' where \j' \ i'\ \is_kth_obs \' l' j'\ \\ (\ j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ proof (rule ccontr) assume nt: \\ thesis\ show \False\ using iki iki' csi that proof (induction \k\ arbitrary: \i\ \i'\ rule: less_induct) case (less k i i') hence all_in_\': \\ l j. j\i \ is_kth_obs \ l j \ (\ j' . cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ and all_in_\: \\ l' j'. j' \ i' \ is_kth_obs \' l' j' \ (\ j . cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ by (metis nt) (metis nt less(6)) obtain j j' where csji: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> i'\ and csij: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> j'\ using all_in_\ all_in_\' less by blast then obtain l l' where ilj: \is_kth_obs \ l j\ and ilj': \is_kth_obs \' l' j'\ by (metis is_kth_obs_def last_cs less.prems(1,2)) have lnk: \l \ k\ using ilj csji less(2) less(4) kth_obs_unique by auto have lnk': \l' \ k\ using ilj' csij less(3) less(4) kth_obs_unique by auto have cseq: \\ l j j'. l < k \ is_kth_obs \ l j \ is_kth_obs \' l j' \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ proof - { fix t p p' assume tk: \t < k\ and ikp: \is_kth_obs \ t p\ and ikp': \is_kth_obs \' t p'\ hence pi: \p < i\ and pi': \p' < i'\ by (metis kth_obs_mono less.prems(1)) (metis kth_obs_mono less.prems(2) tk ikp') have *: \\j l. j \ p \ is_kth_obs \ l j \ \j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using pi all_in_\' by auto have **: \\j' l'. j' \ p' \ is_kth_obs \' l' j' \ \j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using pi' all_in_\ by auto have \cs\<^bsup>\\<^esup> p = cs\<^bsup>\'\<^esup> p'\ apply(rule ccontr) using less(1)[OF tk ikp ikp'] * ** by blast } thus \?thesis\ by blast qed have ii'nret: \\ i \ return \ \' i' \ return\ using less cs_return by auto have a: \k < l \ k < l'\ proof (rule ccontr) assume \\(k < l \ k < l')\ hence *: \l < k\ \l' < k\ using lnk lnk' by auto hence ji: \j < i\ and ji': \j' < i'\ using ilj ilj' less(2,3) kth_obs_mono by auto show \False\ using ii'nret proof assume nreti: \\ i \ return\ hence nretj': \\' j' \ return\ using last_cs csij by metis show \False\ using cs_order[OF path(2,1) csij[symmetric] csji[symmetric] nretj' ji'] ji by simp next assume nreti': \\' i' \ return\ hence nretj': \\ j \ return\ using last_cs csji by metis show \False\ using cs_order[OF path csji csij nretj' ji] ji' by simp qed qed have \l < k \ l' < k\ proof (rule ccontr) assume \\ (l< k \ l' < k)\ hence \k < l\ \k < l'\ using lnk lnk' by auto hence ji: \i < j\ and ji': \i' < j'\ using ilj ilj' less(2,3) kth_obs_mono by auto show \False\ using ii'nret proof assume nreti: \\ i \ return\ show \False\ using cs_order[OF path csij csji nreti ji] ji' by simp next assume nreti': \\' i' \ return\ show \False\ using cs_order[OF path(2,1) csji[symmetric] csij[symmetric] nreti' ji'] ji by simp qed qed hence \k < l \ l' < k \ k < l' \ l < k\ using a by auto thus \False\ proof assume \k < l \ l' < k\ hence kl: \k < l\ and lk': \l' < k\ by auto hence ij: \i < j\ and ji': \j' < i'\ using less(2,3) ilj ilj' kth_obs_mono by auto have nreti: \\ i \ return\ by (metis csji ii'nret ij last_cs path(1) term_path_stable less_imp_le) obtain h where ilh: \is_kth_obs \ l' h\ using ji' all_in_\ ilj' no_kth_obs_missing_cs path(1) path(2) by (metis kl lk' ilj kth_obs_stable) hence \cs\<^bsup>\\<^esup> h = cs\<^bsup>\'\<^esup> j'\ using cseq lk' ilj' by blast hence \cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> h\ using csij by auto hence hi: \h = i\ using cs_inj nreti path(1) by metis have \l' = k\ using less(2) ilh unfolding hi by (metis is_kth_obs_def) thus \False\ using lk' by simp next assume \k < l' \ l < k\ hence kl': \k < l'\ and lk: \l < k\ by auto hence ij': \i' < j'\ and ji: \j < i\ using less(2,3) ilj ilj' kth_obs_mono by auto have nreti': \\' i' \ return\ by (metis csij ii'nret ij' last_cs path(2) term_path_stable less_imp_le) obtain h' where ilh': \is_kth_obs \' l h'\ using all_in_\' ilj no_kth_obs_missing_cs path(1) path(2) kl' lk ilj' kth_obs_stable by metis hence \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> h'\ using cseq lk ilj by blast hence \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\'\<^esup> h'\ using csji by auto hence hi: \h' = i'\ using cs_inj nreti' path(2) by metis have \l = k\ using less(3) ilh' unfolding hi by (metis is_kth_obs_def) thus \False\ using lk by simp qed qed qed subsection \Facts about Data\ lemma reads_restrict1: \\ \ (reads n) = \' \ (reads n) \ \ x \ reads n. \ x = \' x\ by (metis restrict_def) lemma reads_restrict2: \\ x \ reads n. \ x = \' x \ \ \ (reads n) = \' \ (reads n)\ unfolding restrict_def by auto lemma reads_restrict: \(\ \ (reads n) = \' \ (reads n)) = (\ x \ reads n. \ x = \' x)\ using reads_restrict1 reads_restrict2 by metis lemma reads_restr_suc: \\ \ (reads n) = \' \ (reads n) \ suc n \ = suc n \'\ by (metis reads_restrict uses_suc) lemma reads_restr_sem: \\ \ (reads n) = \' \ (reads n) \ \ v \ writes n. sem n \ v = sem n \' v\ by (metis reads_restrict1 uses_writes) lemma reads_obsp: assumes \path \ k = path \' k'\ \\\<^bsup>k\<^esup> \ (reads (path \ k)) = \'\<^bsup>k'\<^esup> \ (reads (path \ k))\ shows \obsp \ k = obsp \' k'\ using assms(2) uses_att unfolding obsp_def assms(1) reads_restrict apply (cases \att (path \' k')\) by auto lemma no_writes_unchanged0: assumes \\ l writes(path \ l)\ shows \(\\<^bsup>k\<^esup>) v = \ v\ using assms proof (induction \k\) case 0 thus \?case\ by(auto simp add: kth_state_def) next case (Suc k) hence \(\\<^bsup>k\<^esup>) v = \ v\ by auto moreover have \\\<^bsup>Suc k\<^esup> = snd ( step (path \ k,\\<^bsup>k\<^esup>))\ by (metis kth_state_suc) hence \\\<^bsup>Suc k\<^esup> = sem (path \ k) (\\<^bsup>k\<^esup>)\ by (metis step_suc_sem snd_conv) moreover have \v \ writes (path \ k)\ using Suc.prems by blast ultimately show \?case\ using writes by metis qed lemma written_read_dd: assumes \is_path \\ \v \ reads (\ k) \ \v \ writes (\ j)\ \j obtains l where \k dd\<^bsup>\,v\<^esup>\ l\ proof - let \?l\ = \GREATEST l. l < k \ v \ writes (\ l)\ have \?l < k\ by (metis (no_types, lifting) GreatestI_ex_nat assms(3) assms(4) less_or_eq_imp_le) moreover have \v \ writes (\ ?l)\ by (metis (no_types, lifting) GreatestI_nat assms(3) assms(4) less_or_eq_imp_le) hence \v \ reads (\ k) \ writes (\ ?l)\ using assms(2) by auto moreover note is_ddi_def have \\ l \ {?l<.. writes (\ l)\ by (auto, metis (lifting, no_types) Greatest_le_nat le_antisym nat_less_le) ultimately have \k dd\<^bsup>\,v\<^esup>\ ?l\ using assms(1) unfolding is_ddi_def by blast thus \thesis\ using that by simp qed lemma no_writes_unchanged: assumes \k \ l\ \\ j \ {k.. writes(path \ j)\ shows \(\\<^bsup>l\<^esup>) v = (\\<^bsup>k\<^esup>) v\ using assms proof (induction \l - k\ arbitrary: \l\) case 0 thus \?case\ by(auto) next case (Suc lk l) hence kl: \k < l\ by auto then obtain l' where lsuc: \l = Suc l'\ using lessE by blast hence \lk = l' - k\ using Suc by auto moreover have \\ j \ {k.. writes (path \ j)\ using Suc(4) lsuc by auto ultimately have \(\\<^bsup>l'\<^esup>) v = (\\<^bsup>k\<^esup>) v\ using Suc(1)[of \l'\] lsuc kl by fastforce moreover have \\\<^bsup>l\<^esup> = snd ( step (path \ l',\\<^bsup>l'\<^esup>))\ by (metis kth_state_suc lsuc) hence \\\<^bsup>l\<^esup> = sem (path \ l') (\\<^bsup>l'\<^esup>)\ by (metis step_suc_sem snd_conv) moreover have \l' < l\ \k \ l'\ using kl lsuc by auto hence \v \ writes (path \ l')\ using Suc.prems(2) by auto ultimately show \?case\ using writes by metis qed lemma ddi_value: assumes \l dd\<^bsup>(path \),v\<^esup>\ k\ shows \(\\<^bsup>l\<^esup>) v = (\\<^bsup>Suc k\<^esup> ) v\ using assms no_writes_unchanged[of \Suc k\ \l\ \v\ \\\] unfolding is_ddi_def by auto lemma written_value: assumes \path \ l = path \' l'\ \\\<^bsup>l\<^esup> \ reads (path \ l) = \'\<^bsup>l'\<^esup> \ reads (path \ l)\ \v \ writes (path \ l)\ shows \(\\<^bsup>Suc l\<^esup> ) v = (\'\<^bsup>Suc l'\<^esup> ) v\ by (metis assms reads_restr_sem snd_conv step_suc_sem kth_state_suc) subsection \Facts about Contradicting Paths\ lemma obsp_contradict: assumes csk: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ and obs: \obsp \ k \ obsp \' k'\ shows \(\', k') \ (\, k)\ proof - have pk: \path \ k = path \' k'\ using assms last_cs by metis hence \\\<^bsup>k\<^esup>\(reads (path \ k)) \ \'\<^bsup>k'\<^esup>\(reads (path \ k))\ using obs reads_obsp[OF pk] by auto thus \(\',k') \ (\,k)\ using contradicts.intros(2)[OF csk[symmetric]] by auto qed lemma missing_cs_contradicts: assumes notin: \\(\ k'. cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k')\ and converge: \k \cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'\ shows \\ j'. (\', j') \ (\, k)\ proof - let \?\\ = \path \\ let \?\'\ = \path \'\ have init: \?\ 0 = ?\' 0\ unfolding path_def by auto have path: \is_path ?\\ \is_path ?\'\ using path_is_path by auto obtain j j' where csj: \cs\<^bsup>?\\<^esup> j = cs\<^bsup>?\'\<^esup> j'\ and cd: \k cd\<^bsup>?\\<^esup>\j\ and suc: \?\ (Suc j) \ ?\' (Suc j')\ using converged_cd_diverge[OF path init notin converge] . have less: \cs\<^bsup>?\\<^esup> j \ cs\<^bsup>?\\<^esup> k\ using cd cd_is_cs_less by auto have nretj: \?\ j \ return\ by (metis cd is_cdi_def term_path_stable less_imp_le) have cs: \?\ \ cs\<^bsup>?\'\<^esup> j' = j\ using csj cs_select_id nretj path_is_path by metis have \(\',j') \ (\,k)\ using contradicts.intros(1)[of \?\'\ \j'\ \?\\ \k\ \\\ \\'\,unfolded cs] less suc csj by metis thus \?thesis\ by blast qed theorem obs_neq_contradicts_term: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes ret: \\ n = return\ \\' n' = return\ and obsne: \obs \ \ obs \'\ shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ proof - have path: \is_path \\ \is_path \'\ using \ \' path_is_path by auto obtain k1 where neq: \obs \ k1 \ obs \' k1\ using obsne ext[of \obs \\ \obs \'\] by blast hence \(\k i i'. is_kth_obs \ k i \ is_kth_obs \' k i' \ obsp \ i \ obsp \' i' \ cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i') \ (\ k i. is_kth_obs \ k i \ \ (\ i'. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')) \ (\ k i'. is_kth_obs \' k i' \ \ (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))\ proof(cases rule: option_neq_cases) case (none2 x) have notin\': \\ (\ l. is_kth_obs \' k1 l)\ using none2(2) \' obs_none_no_kth_obs by auto obtain i where in\: \is_kth_obs \ k1 i\ using obs_some_kth_obs[of \\\ \k1\] none2(1) \ by auto obtain l j where \is_kth_obs \ l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ using path in\ notin\' by (metis no_kth_obs_missing_cs) thus \?thesis\ by blast next case (none1 x) have notin\: \\ (\ l. is_kth_obs \ k1 l)\ using none1(1) \ obs_none_no_kth_obs by auto obtain i' where in\': \is_kth_obs \' k1 i'\ using obs_some_kth_obs[of \\'\ \k1\] none1(2) \' by auto obtain l j where \is_kth_obs \' l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j' = cs\<^bsup>\'\<^esup> j)\ using path in\' notin\ by (metis no_kth_obs_missing_cs) thus \?thesis\ by blast next case (some x y) obtain i where in\: \is_kth_obs \ k1 i\ using obs_some_kth_obs[of \\\ \k1\] some \ by auto obtain i' where in\': \is_kth_obs \' k1 i'\ using obs_some_kth_obs[of \\'\ \k1\] some \' by auto show \?thesis\ proof (cases) assume *: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ have \obsp \ i = obs \ k1\ by (metis obs_def \ in\ kth_obs_unique the_equality) moreover have \obsp \' i' = obs \' k1\ by (metis obs_def \' in\' kth_obs_unique the_equality) ultimately have \obsp \ i \ obsp \' i'\ using neq by auto thus \?thesis\ using * in\ in\' by blast next assume *: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ note kth_obs_cs_missing_cs[OF path in\ in\' *] thus \?thesis\ by metis qed qed thus \?thesis\ proof (cases rule: three_cases) case 1 then obtain k i i' where iki: \is_kth_obs \ k i\ \is_kth_obs \' k i'\ and obsne: \obsp \ i \ obsp \' i'\ and csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ by auto note obsp_contradict[OF csi[unfolded \ \'] obsne] moreover have \\ i \ dom att\ using iki unfolding is_kth_obs_def by auto ultimately show \?thesis\ by blast next case 2 then obtain k i where iki: \is_kth_obs \ k i\ and notin\': \\ (\i'. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')\ by auto let \?n\ = \Suc (max n i)\ have nn: \n < ?n\ by auto have iln: \i < ?n\ by auto have retn: \\ ?n = return\ using ret term_path_stable path by auto hence \cs\<^bsup>\\<^esup> ?n = cs\<^bsup>\'\<^esup> n'\ using ret(2) cs_return by auto then obtain i' where \(\',i') \ (\,i)\ using missing_cs_contradicts[OF notin\'[unfolded \ \'] iln] \ \' by auto moreover have \\ i \ dom att\ using iki is_kth_obs_def by auto ultimately show \?thesis\ by blast next case 3 then obtain k i' where iki: \is_kth_obs \' k i'\ and notin\': \\ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')\ by auto let \?n\ = \Suc (max n' i')\ have nn: \n' < ?n\ by auto have iln: \i' < ?n\ by auto have retn: \\' ?n = return\ using ret term_path_stable path by auto hence \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> ?n\ using ret(1) cs_return by auto then obtain i where \(\,i) \ (\',i')\ using missing_cs_contradicts notin\' iln \ \' by metis moreover have \\' i' \ dom att\ using iki is_kth_obs_def by auto ultimately show \?thesis\ by blast qed qed lemma obs_neq_some_contradicts': fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes obsnecs: \obsp \ i \ obsp \' i' \ cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ and iki: \is_kth_obs \ k i\ and iki': \is_kth_obs \' k i'\ shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom att) \ ((\, k) \ (\' ,k') \ \' k' \ dom att)\ using obsnecs iki iki' proof (induction \k\ arbitrary: \i\ \i'\ rule: less_induct ) case (less k i i') note iki = \is_kth_obs \ k i\ and iki' = \is_kth_obs \' k i'\ have domi: \\ i \ dom att\ by (metis is_kth_obs_def domIff iki) have domi': \\' i' \ dom att\ by (metis is_kth_obs_def domIff iki') note obsnecs = \obsp \ i \ obsp \' i' \ cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ show \?thesis\ proof cases assume csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ hence *: \obsp \ i \ obsp \' i'\ using obsnecs by auto note obsp_contradict[OF _ *] csi domi \ \' thus \?thesis\ by blast next assume ncsi: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ have path: \is_path \\ \is_path \'\ using \ \' path_is_path by auto have \0: \\ 0 = \' 0\ unfolding \ \' path_def by auto note kth_obs_cs_missing_cs[of \\\ \\'\ \k\ \i\ \i'\] \ \' path_is_path iki iki' ncsi hence \(\ l j .j \ i \ is_kth_obs \ l j \ \ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')) \ (\ l' j'. j' \ i' \ is_kth_obs \' l' j' \ \ (\ j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'))\ by metis thus \?thesis\ proof assume \\l j. j \ i \ is_kth_obs \ l j \ \ (\j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ then obtain l j where ji: \j\i\ and iobs: \is_kth_obs \ l j\ and notin: \\ (\j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ by blast have dom: \\ j \ dom att\ using iobs is_kth_obs_def by auto obtain n n' where nj: \n < j\ and csn: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and sucn: \\ (Suc n) \ \' (Suc n')\ and cdloop: \j cd\<^bsup>\\<^esup>\ n \ (\ j'> n'. j' cd\<^bsup>\'\<^esup>\ n')\ using missing_cd_or_loop[OF path \0 notin] by blast show \?thesis\ using cdloop proof assume cdjn: \j cd\<^bsup>\\<^esup>\ n\ hence csnj: \cs\<^bsup>\'\<^esup> n' \ cs\<^bsup>\\<^esup> j\ using csn by (metis cd_is_cs_less) have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> n')) = \ (Suc n)\ using csn by (metis cdjn cd_not_ret cs_select_id path(1)) have \(\',n') \ (\,j)\ using csnj apply(rule contradicts.intros(1)) using cssel \ \' sucn by auto thus \?thesis\ using dom by auto next assume loop: \\ j'>n'. j' cd\<^bsup>\'\<^esup>\ n'\ show \?thesis\ proof cases assume in': \i' \ n'\ have nreti': \\' i' \ return\ by( metis le_eq_less_or_eq lessI loop not_le path(2) ret_no_cd term_path_stable) show \?thesis\ proof cases assume \\ \. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> \\ then obtain \ where cs\: \cs\<^bsup>\\<^esup> \ = cs\<^bsup>\'\<^esup> i'\ by metis have \n: \\ \ n\ using cs_order_le[OF path(2,1) cs\[symmetric] csn[symmetric] nreti' in'] . hence \i: \\ < i\ using nj ji by auto have dom\: \\ \ \ dom att\ using domi' cs\ last_cs by metis obtain \ where i\\: \is_kth_obs \ \ \\ using dom\ by (metis is_kth_obs_def domIff) hence \k: \\ < k\ using \i iki by (metis kth_obs_le_iff) obtain \' where i\\': \is_kth_obs \' \ \'\ using \k iki' by (metis kth_obs_stable) have \\' < i'\ using \k iki' i\\' by (metis kth_obs_le_iff) hence cs\': \cs\<^bsup>\\<^esup> \ \ cs\<^bsup>\'\<^esup> \'\ unfolding cs\ using cs_inj[OF path(2) nreti', of \\'\] by blast thus \?thesis\ using less(1)[OF \k _ i\\ i\\'] by auto next assume notin'': \\(\ \. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> \)\ obtain \ \' where \i': \\' < i'\ and cs\: \cs\<^bsup>\\<^esup> \ = cs\<^bsup>\'\<^esup> \'\ and suc\: \\ (Suc \) \ \' (Suc \')\ and cdloop': \i' cd\<^bsup>\'\<^esup>\ \' \ (\ j>\. j cd\<^bsup>\\<^esup>\ \)\ using missing_cd_or_loop[OF path(2,1) \0[symmetric] notin''] by metis show \?thesis\ using cdloop' proof assume cdjn: \i' cd\<^bsup>\'\<^esup>\ \'\ hence csnj: \cs\<^bsup>\\<^esup> \ \ cs\<^bsup>\'\<^esup> i'\ using cs\ by (metis cd_is_cs_less) have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> \)) = \' (Suc \')\ using cs\ by (metis cdjn cd_not_ret cs_select_id path(2)) have \(\,\) \ (\',i')\ using csnj apply(rule contradicts.intros(1)) using cssel \ \' suc\ by auto thus \?thesis\ using domi' by auto next assume loop': \\ j>\. j cd\<^bsup>\\<^esup>\ \\ have \n': \\' < n'\ using in' \i' by auto have nret\': \\' \' \ return\ by (metis cs\ last_cs le_eq_less_or_eq lessI path(1) path(2) suc\ term_path_stable) have \\ < n\ using cs_order[OF path(2,1) cs\[symmetric] csn[symmetric] nret\' \n'] . hence \\ < i\ using nj ji by auto hence cdi\: \i cd\<^bsup>\\<^esup>\ \\ using loop' by auto hence cs\i: \cs\<^bsup>\'\<^esup> \' \ cs\<^bsup>\\<^esup> i\ using cs\ by (metis cd_is_cs_less) have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> \')) = \ (Suc \)\ using cs\ by (metis cdi\ cd_not_ret cs_select_id path(1)) have \(\',\') \ (\,i)\ using cs\i apply(rule contradicts.intros(1)) using cssel \ \' suc\ by auto thus \?thesis\ using domi by auto qed qed next assume \\ i' \ n'\ hence ni': \n'< i'\ by simp hence cdin: \i' cd\<^bsup>\'\<^esup>\ n'\ using loop by auto hence csni: \cs\<^bsup>\\<^esup> n \ cs\<^bsup>\'\<^esup> i'\ using csn by (metis cd_is_cs_less) have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> n)) = \' (Suc n')\ using csn by (metis cdin cd_not_ret cs_select_id path(2)) have \(\,n) \ (\',i')\ using csni apply(rule contradicts.intros(1)) using cssel \ \' sucn by auto thus \?thesis\ using domi' by auto qed qed next \ \Symmetric case as above, indices might be messy.\ assume \\l j. j \ i' \ is_kth_obs \' l j \ \ (\j'. cs\<^bsup>\\<^esup> j' = cs\<^bsup>\'\<^esup> j)\ then obtain l j where ji': \j\i'\ and iobs: \is_kth_obs \' l j\ and notin: \\ (\j'. cs\<^bsup>\'\<^esup> j = cs\<^bsup>\\<^esup> j')\ by metis have dom: \\' j \ dom att\ using iobs is_kth_obs_def by auto obtain n n' where nj: \n < j\ and csn: \cs\<^bsup>\'\<^esup> n = cs\<^bsup>\\<^esup> n'\ and sucn: \\' (Suc n) \ \ (Suc n')\ and cdloop: \j cd\<^bsup>\'\<^esup>\ n \ (\ j'> n'. j' cd\<^bsup>\\<^esup>\ n')\ using missing_cd_or_loop[OF path(2,1) \0[symmetric] ] notin by metis show \?thesis\ using cdloop proof assume cdjn: \j cd\<^bsup>\'\<^esup>\ n\ hence csnj: \cs\<^bsup>\\<^esup> n' \ cs\<^bsup>\'\<^esup> j\ using csn by (metis cd_is_cs_less) have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> n')) = \' (Suc n)\ using csn by (metis cdjn cd_not_ret cs_select_id path(2)) have \(\,n') \ (\',j)\ using csnj apply(rule contradicts.intros(1)) using cssel \' \ sucn by auto thus \?thesis\ using dom by auto next assume loop: \\ j'>n'. j' cd\<^bsup>\\<^esup>\ n'\ show \?thesis\ proof cases assume in': \i \ n'\ have nreti: \\ i \ return\ by (metis le_eq_less_or_eq lessI loop not_le path(1) ret_no_cd term_path_stable) show \?thesis\ proof cases assume \\ \. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> \\ then obtain \ where cs\: \cs\<^bsup>\'\<^esup> \ = cs\<^bsup>\\<^esup> i\ by metis have \n: \\ \ n\ using cs_order_le[OF path cs\[symmetric] csn[symmetric] nreti in'] . hence \i': \\ < i'\ using nj ji' by auto have dom\: \\' \ \ dom att\ using domi cs\ last_cs by metis obtain \ where i\\: \is_kth_obs \' \ \\ using dom\ by (metis is_kth_obs_def domIff) hence \k: \\ < k\ using \i' iki' by (metis kth_obs_le_iff) obtain \' where i\\': \is_kth_obs \ \ \'\ using \k iki by (metis kth_obs_stable) have \\' < i\ using \k iki i\\' by (metis kth_obs_le_iff) hence cs\': \cs\<^bsup>\'\<^esup> \ \ cs\<^bsup>\\<^esup> \'\ unfolding cs\ using cs_inj[OF path(1) nreti, of \\'\] by blast thus \?thesis\ using less(1)[OF \k _ i\\' i\\] by auto next assume notin'': \\(\ \. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> \)\ obtain \ \' where \i: \\' < i\ and cs\: \cs\<^bsup>\'\<^esup> \ = cs\<^bsup>\\<^esup> \'\ and suc\: \\' (Suc \) \ \ (Suc \')\ and cdloop': \i cd\<^bsup>\\<^esup>\ \' \ (\ j>\. j cd\<^bsup>\'\<^esup>\ \)\ using missing_cd_or_loop[OF path \0 notin''] by metis show \?thesis\ using cdloop' proof assume cdjn: \i cd\<^bsup>\\<^esup>\ \'\ hence csnj: \cs\<^bsup>\'\<^esup> \ \ cs\<^bsup>\\<^esup> i\ using cs\ by (metis cd_is_cs_less) have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> \)) = \ (Suc \')\ using cs\ by (metis cdjn cd_not_ret cs_select_id path(1)) have \(\',\) \ (\,i)\ using csnj apply(rule contradicts.intros(1)) using cssel \' \ suc\ by auto thus \?thesis\ using domi by auto next assume loop': \\ j>\. j cd\<^bsup>\'\<^esup>\ \\ have \n': \\' < n'\ using in' \i by auto have nret\': \\ \' \ return\ by (metis cs\ last_cs le_eq_less_or_eq lessI path(1) path(2) suc\ term_path_stable) have \\ < n\ using cs_order[OF path cs\[symmetric] csn[symmetric] nret\' \n'] . hence \\ < i'\ using nj ji' by auto hence cdi\: \i' cd\<^bsup>\'\<^esup>\ \\ using loop' by auto hence cs\i': \cs\<^bsup>\\<^esup> \' \ cs\<^bsup>\'\<^esup> i'\ using cs\ by (metis cd_is_cs_less) have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> \')) = \' (Suc \)\ using cs\ by (metis cdi\ cd_not_ret cs_select_id path(2)) have \(\,\') \ (\',i')\ using cs\i' apply(rule contradicts.intros(1)) using cssel \' \ suc\ by auto thus \?thesis\ using domi' by auto qed qed next assume \\ i \ n'\ hence ni: \n'< i\ by simp hence cdin: \i cd\<^bsup>\\<^esup>\ n'\ using loop by auto hence csni': \cs\<^bsup>\'\<^esup> n \ cs\<^bsup>\\<^esup> i\ using csn by (metis cd_is_cs_less) have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> n)) = \ (Suc n')\ using csn by (metis cdin cd_not_ret cs_select_id path(1)) have \(\',n) \ (\,i)\ using csni' apply(rule contradicts.intros(1)) using cssel \' \ sucn by auto thus \?thesis\ using domi by auto qed qed qed qed qed theorem obs_neq_some_contradicts: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes obsne: \obs \ k \ obs \' k\ and not_none: \obs \ k \ None\ \obs \' k \ None\ shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom att) \ ((\, k) \ (\' ,k') \ \' k' \ dom att)\ proof - obtain i where iki: \is_kth_obs \ k i\ using not_none(1) by (metis \ obs_some_kth_obs) obtain i' where iki': \is_kth_obs \' k i'\ using not_none(2) by (metis \' obs_some_kth_obs) have \obsp \ i = obs \ k\ by (metis \ iki kth_obs_unique obs_def the_equality) moreover have \obsp \' i' = obs \' k\ by (metis \' iki' kth_obs_unique obs_def the_equality) ultimately have obspne: \obsp \ i \ obsp \' i'\ using obsne by auto show \?thesis\ using obs_neq_some_contradicts'[OF _ iki[unfolded \] iki'[unfolded \']] using obspne \ \' by metis qed theorem obs_neq_ret_contradicts: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes ret: \\ n = return\ and obsne: \obs \' i \ obs \ i\ and obs:\obs \' i \ None\ shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ proof (cases \\ j k'. is_kth_obs \' j k' \ (\ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\) case True obtain l k' where jk': \is_kth_obs \' l k'\ and unmatched: \(\ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ using True by blast have \0: \\ 0 = \' 0\ using \ \' path0 by auto obtain j j' where csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and cd: \k' cd\<^bsup>\'\<^esup>\j'\ and suc: \\ (Suc j) \ \' (Suc j')\ using converged_cd_diverge_return[of \\'\ \\\ \k'\ \n\] ret unmatched path_is_path \ \' \0 by metis hence *: \(\, j) \ (\' ,k')\ using contradicts.intros(1)[of \\\ \j\ \\'\ \k'\ \\'\ \\\, unfolded csj] \ \' using cd_is_cs_less cd_not_ret cs_select_id by auto have \\' k' \ dom(att)\ using jk' by (meson domIff is_kth_obs_def) thus \?thesis\ using * by blast next case False hence *: \\ j k'. is_kth_obs \' j k' \ \ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ by auto obtain k' where k': \is_kth_obs \' i k'\ using obs \' obs_some_kth_obs by blast obtain l where \is_kth_obs \ i l\ using * \ \' k' no_kth_obs_missing_cs path_is_path by metis thus \?thesis\ using \ \' obs obs_neq_some_contradicts obs_none_no_kth_obs obsne by metis qed subsection \Facts about Critical Observable Paths\ lemma contradicting_in_cp: assumes leq:\\ =\<^sub>L \'\ and cseq: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ and readv: \v\reads(path \ k)\ and vneq: \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ shows \((\,k),(\',k')) \ cp\ using cseq readv vneq proof(induction \k+k'\ arbitrary: \k\ \k'\ \v\ rule: less_induct) fix k k' v assume csk: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ assume vread: \v \ reads (path \ k)\ assume vneq: \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ assume IH: \\ka k'a v. ka + k'a < k + k' \ cs\<^bsup>path \\<^esup> ka = cs\<^bsup>path \'\<^esup> k'a \ v \ reads (path \ ka) \ (\\<^bsup>ka\<^esup>) v \ (\'\<^bsup>k'a\<^esup>) v \ ((\, ka), \', k'a) \ cp\ define \ where \\ \ path \\ define \' where \\' \ path \'\ have path: \\ = path \\ \\' = path \'\ using \_def \'_def path_is_path by auto have ip: \is_path \\ \is_path \'\ using path path_is_path by auto have \0: \\' 0 = \ 0\ unfolding path path_def by auto have vread': \v \ reads (path \' k')\ using csk vread by (metis last_cs) have cseq: \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k\ using csk path by simp show \((\, k), \', k') \ cp\ proof cases assume vnw: \\ l < k. v\writes (\ l)\ hence \v: \(\\<^bsup>k\<^esup>) v = \ v\ by (metis no_writes_unchanged0 path(1)) show \?thesis\ proof cases assume vnw': \\ l < k'. v\writes (\' l)\ hence \v': \(\'\<^bsup>k'\<^esup>) v = \' v\ by (metis no_writes_unchanged0 path(2)) with \v vneq have \\ v \ \' v\ by auto hence vhigh: \v \ hvars\ using leq unfolding loweq_def restrict_def by (auto,metis) thus \?thesis\ using cp.intros(1)[OF leq csk vread vneq] vnw vnw' path by simp next assume \\(\ l < k'. v\writes (\' l))\ then obtain l' where kddl': \k' dd\<^bsup>\',v\<^esup>\ l'\ using path(2) path_is_path written_read_dd vread' by blast hence lv': \v \ writes (\' l')\ unfolding is_ddi_def by auto have lk': \l' < k'\ by (metis is_ddi_def kddl') have nret: \\' l' \ return\ using lv' writes_return by auto have notin\: \\ (\l. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l)\ proof assume \\l. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l\ - then guess l .. + then obtain l where "cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l" .. note csl = \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l\ have lk: \l < k\ using lk' cseq ip cs_order[of \\'\ \\\ \l'\ \l\ \k'\ \k\] csl nret path by force have \v \ writes (\ l)\ using csl lv' last_cs by metis thus \False\ using lk vnw by blast qed from converged_cd_diverge[OF ip(2,1) \0 notin\ lk' cseq] obtain i i' where csi: \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ and lcdi: \l' cd\<^bsup>\'\<^esup>\ i'\ and div: \\' (Suc i') \ \ (Suc i)\ . have 1: \\ (Suc i) = suc (\ i) (\\<^bsup>i\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc i') = suc (\' i') (\'\<^bsup>i'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' i' = \ i\ using csi last_cs by metis have nreads: \\\<^bsup>i\<^esup> \ reads (\ i) \ \'\<^bsup>i'\<^esup> \ reads (\ i)\ by (metis 1 2 3 div reads_restr_suc) then obtain v' where v'read: \v'\ reads(path \ i)\ \(\\<^bsup>i\<^esup>) v' \ (\'\<^bsup>i'\<^esup>) v'\ unfolding path by (metis reads_restrict) have nreti: \\' i' \ return\ by (metis csi div ip(1) ip(2) last_cs lessI term_path_stable less_imp_le) have ik': \i' < k'\ using lcdi lk' is_cdi_def by auto have ik: \i < k\ using cs_order[OF ip(2,1) csi cseq nreti ik'] . have cpi: \((\, i), (\', i')) \ cp\ using IH[of \i\ \i'\] v'read csi ik ik' path by auto hence cpi': \((\', i'), (\, i)) \ cp\ using cp.intros(4) by blast have nwvi: \\j'\{LEAST i'. i < i' \ (\i. cs\<^bsup>path \'\<^esup> i = cs\<^bsup>path \\<^esup> i').. writes (path \ j')\ using vnw[unfolded path] by (metis (poly_guards_query) atLeastLessThan_iff) from cp.intros(3)[OF cpi' kddl'[unfolded path] lcdi[unfolded path] csk[symmetric] div[unfolded path] vneq[symmetric] nwvi] show \?thesis\ using cp.intros(4) by simp qed next assume wv: \\ (\ l writes (\ l))\ then obtain l where kddl: \k dd\<^bsup>\,v\<^esup>\ l\ using path(1) path_is_path written_read_dd vread by blast hence lv: \v \ writes (\ l)\ unfolding is_ddi_def by auto have lk: \l < k\ by (metis is_ddi_def kddl) have nret: \\ l \ return\ using lv writes_return by auto have nwb: \\ i \ {Suc l..< k}. v\writes(\ i)\ using kddl unfolding is_ddi_def by auto have \vk: \(\\<^bsup>k\<^esup>) v = (\\<^bsup>Suc l\<^esup> ) v\ using kddl ddi_value path(1) by auto show \?thesis\ proof cases assume vnw': \\ l < k'. v\writes (\' l)\ hence \v': \(\'\<^bsup>k'\<^esup>) v = \' v\ by (metis no_writes_unchanged0 path(2)) have notin\': \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ proof assume \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ - then guess l' .. + then obtain l' where "cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'" .. note csl = \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ have lk: \l' < k'\ using lk cseq ip cs_order[of \\\ \\'\ \l\ \l'\ \k\ \k'\] csl nret by metis have \v \ writes (\' l')\ using csl lv last_cs by metis thus \False\ using lk vnw' by blast qed from converged_cd_diverge[OF ip(1,2) \0[symmetric] notin\' lk cseq[symmetric]] obtain i i' where csi: \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ and lcdi: \l cd\<^bsup>\\<^esup>\ i\ and div: \\ (Suc i) \ \' (Suc i')\ by metis have 1: \\ (Suc i) = suc (\ i) (\\<^bsup>i\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc i') = suc (\' i') (\'\<^bsup>i'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' i' = \ i\ using csi last_cs by metis have nreads: \\\<^bsup>i\<^esup> \ reads (\ i) \ \'\<^bsup>i'\<^esup> \ reads (\ i)\ by (metis 1 2 3 div reads_restr_suc) have contri: \(\',i') \ (\,i)\ using contradicts.intros(2)[OF csi path nreads] . have nreti: \\ i \ return\ by (metis csi div ip(1) ip(2) last_cs lessI term_path_stable less_imp_le) have ik: \i < k\ using lcdi lk is_cdi_def by auto have ik': \i' < k'\ using cs_order[OF ip(1,2) csi[symmetric] cseq[symmetric] nreti ik] . have nreads: \\\<^bsup>i\<^esup> \ reads (\ i) \ \'\<^bsup>i'\<^esup> \ reads (\ i)\ by (metis 1 2 3 div reads_restr_suc) then obtain v' where v'read: \v'\ reads(path \ i)\ \(\\<^bsup>i\<^esup>) v' \ (\'\<^bsup>i'\<^esup>) v'\ unfolding path by (metis reads_restrict) have cpi: \((\, i), (\', i')) \ cp\ using IH[of \i\ \i'\] v'read csi ik ik' path by auto hence cpi': \((\', i'), (\, i)) \ cp\ using cp.intros(4) by blast have vnwi: \\j'\{LEAST i'a. i' < i'a \ (\i. cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i'a).. writes (path \' j')\ using vnw'[unfolded path] by (metis (poly_guards_query) atLeastLessThan_iff) from cp.intros(3)[OF cpi kddl[unfolded path] lcdi[unfolded path] csk div[unfolded path] vneq vnwi] show \?thesis\ using cp.intros(4) by simp next assume \\ (\ l writes (\' l))\ then obtain l' where kddl': \k' dd\<^bsup>\',v\<^esup>\ l'\ using path(2) path_is_path written_read_dd vread' by blast hence lv': \v \ writes (\' l')\ unfolding is_ddi_def by auto have lk': \l' < k'\ by (metis is_ddi_def kddl') have nretl': \\' l' \ return\ using lv' writes_return by auto have nwb': \\ i' \ {Suc l'..< k'}. v\writes(\' i')\ using kddl' unfolding is_ddi_def by auto have \vk': \(\'\<^bsup>k'\<^esup>) v = (\'\<^bsup>Suc l'\<^esup> ) v\ using kddl' ddi_value path(2) by auto show \?thesis\ proof cases assume csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ hence \l: \\ l = \' l'\ by (metis last_cs) have \vls: \(\\<^bsup>Suc l\<^esup> ) v \ (\'\<^bsup>Suc l'\<^esup> ) v\ by (metis \vk \vk' vneq) have r\: \\\<^bsup>l\<^esup> \ reads (\ l) \ \'\<^bsup>l'\<^esup> \ reads (\ l)\ using path \l \vls written_value lv by blast then obtain v' where v'read: \v'\ reads(path \ l)\ \(\\<^bsup>l\<^esup>) v' \ (\'\<^bsup>l'\<^esup>) v'\ unfolding path by (metis reads_restrict) have cpl: \((\, l), (\', l')) \ cp\ using IH[of \l\ \l'\] v'read csl lk lk' path by auto show \((\, k), (\', k')) \ cp\ using cp.intros(2)[OF cpl kddl[unfolded path] kddl'[unfolded path] csk vneq] . next assume csl: \cs\<^bsup>\\<^esup> l \ cs\<^bsup>\'\<^esup> l'\ show \?thesis\ proof cases assume \\ i'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> i'\ then obtain i' where csli': \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> i'\ by blast have ilne': \i' \ l'\ using csl csli' by auto have ij': \i' < k'\ using cs_order[OF ip csli' cseq[symmetric] nret lk] . have iv': \v \ writes(\' i')\ using lv csli' last_cs by metis have il': \i' < l'\ using kddl' ilne' ij' iv' unfolding is_ddi_def by auto have nreti': \\' i' \ return\ using csli' nret last_cs by metis have l'notin\: \\(\i. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> i )\ proof assume \\i. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> i\ then obtain i where csil: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l'\ by metis have ik: \i < k\ using cs_order[OF ip(2,1) csil[symmetric] cseq nretl' lk'] . have li: \l < i\ using cs_order[OF ip(2,1) csli'[symmetric] csil[symmetric] nreti' il'] . have iv: \v \ writes(\ i)\ using lv' csil last_cs by metis show \False\ using kddl ik li iv is_ddi_def by auto qed obtain n n' where csn: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ and sucn: \\ (Suc n) \ \' (Suc n')\ and in': \i' \ n'\ using converged_cd_diverge_cs [OF ip(2,1) csli'[symmetric] il' l'notin\ lk' cseq] by metis \ \Can apply the IH to n and n'\ have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' n' = \ n\ using csn last_cs by metis have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 sucn reads_restr_suc) then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) moreover have nl': \n' < l'\ using lcdn' is_cdi_def by auto have nk': \n' < k'\ using nl' lk' by simp have nretn': \\' n' \ return\ by (metis ip(2) nl' nretl' term_path_stable less_imp_le) have nk: \n < k\ using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . hence lenn: \n+n' < k+k'\ using nk' by auto ultimately have \((\, n), (\', n')) \ cp\ using IH csn path by auto hence ncp: \((\', n'), (\, n)) \ cp\ using cp.intros(4) by auto have nles: \n < (LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i'))\ (is \_ < (LEAST i. ?P i)\) using nk cseq LeastI[of \?P\ \k\] by metis moreover have ln: \l \ n\ using cs_order_le[OF ip(2,1) csli'[symmetric] csn[symmetric] nreti' in'] . ultimately have lles: \Suc l \ (LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i'))\ by auto have nwcseq: \\j'\{LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i').. writes (\ j')\ proof fix j' assume *: \j' \ {LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i').. hence \(LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i')) \ j'\ by (metis (poly_guards_query) atLeastLessThan_iff) hence \Suc l \ j'\ using lles by auto moreover have \j' < k\ using * by (metis (poly_guards_query) atLeastLessThan_iff) ultimately have \j'\ {Suc l.. by (metis (poly_guards_query) atLeastLessThan_iff) thus \v\writes (\ j')\ using nwb by auto qed from cp.intros(3)[OF ncp,folded path,OF kddl' lcdn' cseq sucn[symmetric] vneq[symmetric] nwcseq] have \((\', k'), \, k) \ cp\ . thus \((\, k), (\', k')) \ cp\ using cp.intros(4) by auto next assume lnotin\': \\ (\i'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> i')\ show \?thesis\ proof cases assume \\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l'\ then obtain i where csli: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l'\ by blast have ilne: \i \ l\ using csl csli by auto have ij: \i < k\ using cs_order[OF ip(2,1) csli[symmetric] cseq nretl' lk'] . have iv: \v \ writes(\ i)\ using lv' csli last_cs by metis have il: \i < l\ using kddl ilne ij iv unfolding is_ddi_def by auto have nreti: \\ i \ return\ using csli nretl' last_cs by metis obtain n n' where csn: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn: \l cd\<^bsup>\\<^esup>\ n\ and sucn: \\ (Suc n) \ \' (Suc n')\ and ilen: \i \ n\ using converged_cd_diverge_cs [OF ip csli il lnotin\' lk cseq[symmetric]] by metis \ \Can apply the IH to n and n'\ have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' n' = \ n\ using csn last_cs by metis have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 sucn reads_restr_suc) then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) moreover have nl: \n < l\ using lcdn is_cdi_def by auto have nk: \n < k\ using nl lk by simp have nretn: \\ n \ return\ by (metis ip(1) nl nret term_path_stable less_imp_le) have nk': \n' < k'\ using cs_order[OF ip csn cseq[symmetric] nretn nk] . hence lenn: \n+n' < k+k'\ using nk by auto ultimately have ncp: \((\, n), (\', n')) \ cp\ using IH csn path by auto have nles': \n' < (LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))\ (is \_ < (LEAST i. ?P i)\) using nk' cseq LeastI[of \?P\ \k'\] by metis moreover have ln': \l' \ n'\ using cs_order_le[OF ip csli csn nreti ilen] . ultimately have lles': \Suc l' \ (LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))\ by auto have nwcseq': \\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')).. writes (\' j')\ proof fix j' assume *: \j' \ {(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')).. hence \(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')) \ j'\ by (metis (poly_guards_query) atLeastLessThan_iff) hence \Suc l' \ j'\ using lles' by auto moreover have \j' < k'\ using * by (metis (poly_guards_query) atLeastLessThan_iff) ultimately have \j'\ {Suc l'.. by (metis (poly_guards_query) atLeastLessThan_iff) thus \v\writes (\' j')\ using nwb' by auto qed from cp.intros(3)[OF ncp,folded path, OF kddl lcdn cseq[symmetric] sucn vneq nwcseq'] show \((\, k), (\', k')) \ cp\ . next assume l'notin\: \\ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l')\ define m where \m \ 0::nat\ define m' where \m' \ 0::nat\ have csm: \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ unfolding m_def m'_def cs_0 by (metis \0) have ml: \m m' using csm csl unfolding m_def m'_def by (metis neq0_conv) have \\ n n'. cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n' \ \ (Suc n) \ \' (Suc n') \ (l cd\<^bsup>\\<^esup>\ n \ (\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j')) \ l' cd\<^bsup>\'\<^esup>\ n' \ (\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)))\ using csm ml proof (induction \k+k'-(m+m')\ arbitrary: \m\ \m'\ rule: less_induct) case (less m m') note csm = \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ note lm = \m < l \ m' < l'\ note IH = \\ n n'. k + k' - (n + n') < k + k' - (m + m') \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n' \ n < l \ n' < l' \ ?thesis\ show \?thesis\ using lm proof assume ml: \m < l\ obtain n n' where mn: \m \ n\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn: \l cd\<^bsup>\\<^esup>\ n\ and suc: \\ (Suc n) \ \' (Suc n')\ using converged_cd_diverge_cs[OF ip csm ml lnotin\' lk cseq[symmetric]] . have nl: \n < l\ using lcdn is_cdi_def by auto hence nk: \n using lk by auto have nretn: \\ n \ return\ using lcdn by (metis cd_not_ret) have nk': \n' using cs_order[OF ip csn cseq[symmetric] nretn nk] . show \?thesis\ proof cases assume \\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j')\ thus \?thesis\ using lcdn csn suc by blast next assume \\(\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j'))\ then obtain j' where jin': \j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')).. and vwrite: \v\writes (\' j')\ by blast define i' where \i' \ LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')\ have Pk': \n' < k' \ (\ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ (is \?P k'\) using nk' cseq[symmetric] by blast have ni': \n' < i'\ using LeastI[of \?P\, OF Pk'] i'_def by auto obtain i where csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using LeastI[of \?P\, OF Pk'] i'_def by blast have ij': \i'\j'\ using jin'[folded i'_def] by auto have jk': \j' using jin'[folded i'_def] by auto have jl': \j' \ l'\ using kddl' jk' vwrite unfolding is_ddi_def by auto have nretn': \\' n' \ return\ using nretn csn last_cs by metis have iln: \n using cs_order[OF ip(2,1) csn[symmetric] csi[symmetric] nretn' ni'] . hence mi: \m < i\ using mn by auto have nretm: \\ m \ return\ by (metis ip(1) mn nretn term_path_stable) have mi': \m' using cs_order[OF ip csm csi nretm mi] . have ik': \i' < k'\ using ij' jk' by auto have nreti': \\' i' \ return\ by (metis ij' jl' nretl' ip(2) term_path_stable) have ik: \i < k\ using cs_order[OF ip(2,1) csi[symmetric] cseq nreti' ik'] . show \?thesis\ proof cases assume il:\i < l\ have le: \k + k' - (i +i') < k+k' - (m+m')\ using mi mi' ik ik' by auto show \?thesis\ using IH[OF le] using csi il by blast next assume \\ i < l\ hence li: \l \ i\ by auto have \i' \ l'\ using ij' jl' by auto hence il': \i' < l'\ using csi l'notin\ by fastforce obtain n n' where in': \i' \ n'\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ and suc: \\ (Suc n) \ \' (Suc n')\ using converged_cd_diverge_cs[OF ip(2,1) csi[symmetric] il' _ lk' cseq] l'notin\ by metis have nk': \n' < k'\ using lcdn' is_cdi_def lk' by auto have nretn': \\' n' \ return\ by (metis cd_not_ret lcdn') have nk: \n < k\ using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . define j where \j \ LEAST j. n < j \ (\j'. cs\<^bsup>\'\<^esup> j' = cs\<^bsup>\\<^esup> j)\ have Pk: \n < k \ (\j'. cs\<^bsup>\'\<^esup> j' = cs\<^bsup>\\<^esup> k)\ (is \?P k\) using nk cseq by blast have nj: \n using LeastI[of \?P\, OF Pk] j_def by auto have ilen: \i \ n\ using cs_order_le[OF ip(2,1) csi[symmetric] csn[symmetric] nreti' in'] . hence lj: \l using li nj by simp have \\l\{l<.. writes (\ l)\ using kddl unfolding is_ddi_def by simp hence nw: \\l\{j.. writes (\ l)\ using lj by auto show \?thesis\ using csn lcdn' suc nw[unfolded j_def] by blast qed qed next assume ml': \m' < l'\ obtain n n' where mn': \m' \ n'\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ and suc: \\ (Suc n) \ \' (Suc n')\ using converged_cd_diverge_cs[OF ip(2,1) csm[symmetric] ml' _ lk' cseq] l'notin\ by metis have nl': \n' < l'\ using lcdn' is_cdi_def by auto hence nk': \n' using lk' by auto have nretn': \\' n' \ return\ using lcdn' by (metis cd_not_ret) have nk: \n using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . show \?thesis\ proof cases assume \\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)\ thus \?thesis\ using lcdn' csn suc by blast next assume \\(\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j))\ then obtain j where jin: \j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i)).. and vwrite: \v\writes (\ j)\ by blast define i where \i \ LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i)\ have Pk: \n < k \ (\ k'. cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k)\ (is \?P k\) using nk cseq by blast have ni: \n < i\ using LeastI[of \?P\, OF Pk] i_def by auto obtain i' where csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using LeastI[of \?P\, OF Pk] i_def by metis have ij: \i\j\ using jin[folded i_def] by auto have jk: \j using jin[folded i_def] by auto have jl: \j \ l\ using kddl jk vwrite unfolding is_ddi_def by auto have nretn: \\ n \ return\ using nretn' csn last_cs by metis have iln': \n' using cs_order[OF ip csn csi nretn ni] . hence mi': \m' < i'\ using mn' by auto have nretm': \\' m' \ return\ by (metis ip(2) mn' nretn' term_path_stable) have mi: \m using cs_order[OF ip(2,1) csm[symmetric] csi[symmetric] nretm' mi'] . have ik: \i < k\ using ij jk by auto have nreti: \\ i \ return\ by (metis ij ip(1) jl nret term_path_stable) have ik': \i' < k'\ using cs_order[OF ip csi cseq[symmetric] nreti ik] . show \?thesis\ proof cases assume il':\i' < l'\ have le: \k + k' - (i +i') < k+k' - (m+m')\ using mi mi' ik ik' by auto show \?thesis\ using IH[OF le] using csi il' by blast next assume \\ i' < l'\ hence li': \l' \ i'\ by auto have \i \ l\ using ij jl by auto hence il: \i < l\ using csi lnotin\' by fastforce obtain n n' where ilen: \i \ n\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn: \l cd\<^bsup>\\<^esup>\ n\ and suc: \\ (Suc n) \ \' (Suc n')\ using converged_cd_diverge_cs[OF ip csi il _ lk cseq[symmetric]] lnotin\' by metis have nk: \n < k\ using lcdn is_cdi_def lk by auto have nretn: \\ n \ return\ by (metis cd_not_ret lcdn) have nk': \n' < k'\ using cs_order[OF ip csn cseq[symmetric] nretn nk] . define j' where \j' \ LEAST j'. n' < j' \ (\j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ have Pk': \n' < k' \ (\j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> k')\ (is \?P k'\) using nk' cseq[symmetric] by blast have nj': \n' using LeastI[of \?P\, OF Pk'] j'_def by auto have in': \i' \ n'\ using cs_order_le[OF ip csi csn nreti ilen] . hence lj': \l' using li' nj' by simp have \\l\{l'<.. writes (\' l)\ using kddl' unfolding is_ddi_def by simp hence nw': \\l\{j'.. writes (\' l)\ using lj' by auto show \?thesis\ using csn lcdn suc nw'[unfolded j'_def] by blast qed qed qed qed then obtain n n' where csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and suc: \\ (Suc n) \ \' (Suc n')\ and cdor: \(l cd\<^bsup>\\<^esup>\ n \ (\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j')) \ l' cd\<^bsup>\'\<^esup>\ n' \ (\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)))\ by blast show \?thesis\ using cdor proof assume *: \l cd\<^bsup>\\<^esup>\ n \ (\j'\{LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i').. local.writes (\' j'))\ hence lcdn: \l cd\<^bsup>\\<^esup>\ n\ by blast have nowrite: \\j'\{LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i').. local.writes (\' j')\ using * by blast show \?thesis\ proof (rule cp.intros(3)[of \\\ \n\ \\'\ \n'\,folded path]) show \l cd\<^bsup>\\<^esup>\ n\ using lcdn . show \k dd\<^bsup>\,v\<^esup>\ l\ using kddl . show \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ using cseq by simp show \\ (Suc n) \ \' (Suc n')\ using suc by simp show \\j'\{LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i').. local.writes (\' j')\ using nowrite . show \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ using vneq . have nk: \n < k\ using lcdn lk is_cdi_def by auto have nretn: \\ n \ return\ using cd_not_ret lcdn by metis have nk': \n' < k'\ using cs_order[OF ip csn cseq[symmetric] nretn nk] . hence le: \n + n' < k + k'\ using nk by auto moreover have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' n' = \ n\ using csn last_cs by metis have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 suc reads_restr_suc) then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) ultimately show \((\, n), (\', n')) \ cp\ using IH csn path by auto qed next assume *: \l' cd\<^bsup>\'\<^esup>\ n' \ (\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j))\ hence lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ by blast have nowrite: \\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)\ using * by blast show \?thesis\ proof (rule cp.intros(4), rule cp.intros(3)[of \\'\ \n'\ \\\ \n\,folded path]) show \l' cd\<^bsup>\'\<^esup>\ n'\ using lcdn' . show \k' dd\<^bsup>\',v\<^esup>\ l'\ using kddl' . show \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k\ using cseq . show \\' (Suc n') \ \ (Suc n)\ using suc by simp show \\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)\ using nowrite . show \(\'\<^bsup>k'\<^esup>) v \ (\\<^bsup>k\<^esup>) v\ using vneq by simp have nk': \n' < k'\ using lcdn' lk' is_cdi_def by auto have nretn': \\' n' \ return\ using cd_not_ret lcdn' by metis have nk: \n < k\ using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . hence le: \n + n' < k + k'\ using nk' by auto moreover have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) have 3: \\' n' = \ n\ using csn last_cs by metis have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 suc reads_restr_suc) then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) ultimately have \((\, n), (\', n')) \ cp\ using IH csn path by auto thus \((\', n'), \, n) \ cp\ using cp.intros(4) by simp qed qed qed qed qed qed qed qed theorem contradicting_in_cop: assumes \\ =\<^sub>L \'\ and \(\',k') \ (\,k)\ and \path \ k \ dom att\ shows \((\,k),\',k') \ cop\ using assms(2) proof(cases) case (1 \' \) define j where \j \ \ \ cs\<^bsup>\'\<^esup> k'\ have csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> k'\ unfolding j_def using 1 by (metis cs_not_nil cs_select_is_cs(1) path_is_path) have suc: \\ (Suc j) \ \' (Suc k')\ using 1 j_def by simp have kcdj: \k cd\<^bsup>\\<^esup>\ j\ by (metis cs_not_nil cs_select_is_cs(2) 1(1,2) j_def path_is_path) obtain v where readv: \v\reads(path \ j)\ and vneq: \(\\<^bsup>j\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ using suc csj unfolding 1 by (metis IFC_def.suc_def 1(2) 1(3) last_cs path_suc reads_restr_suc reads_restrict) have \((\,j),\',k') \ cp\ apply (rule contradicting_in_cp[OF assms(1)]) using readv vneq csj 1 by auto thus \((\,k),\',k') \ cop\ using kcdj suc assms(3) cop.intros(2) unfolding 1 by auto next case (2 \' \) obtain v where readv: \v\reads(path \ k)\ and vneq: \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ using 2(2-4) by (metis reads_restrict) have \((\,k),\',k') \ cp\ apply (rule contradicting_in_cp[OF assms(1)]) using readv vneq 2 by auto thus \((\,k),\',k') \ cop\ using assms(3) cop.intros(1) unfolding 2 by auto qed theorem cop_correct_term: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes ret: \\ n = return\ \\' n' = return\ and obsne: \obs \ \ obs \'\ and leq: \\ =\<^sub>L \'\ shows \\ k k'. ((\,k),\',k')\ cop \ ((\',k'),\,k)\ cop\ proof - have *: \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ using obs_neq_contradicts_term ret obsne \ \' by auto have leq' :\\' =\<^sub>L \\ using leq unfolding loweq_def by auto from * contradicting_in_cop[OF leq] contradicting_in_cop[OF leq'] show \?thesis\ unfolding \ \' by metis qed theorem cop_correct_ret: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes ret: \\ n = return\ and obsne: \obs \ i \ obs \' i\ and obs: \obs \' i \ None\ and leq: \\ =\<^sub>L \'\ shows \\ k k'. ((\,k),\',k')\ cop \ ((\',k'),\,k)\ cop\ proof - have *: \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ by (metis (no_types, lifting) \ \' obs obs_neq_ret_contradicts obsne ret) have leq' :\\' =\<^sub>L \\ using leq unfolding loweq_def by auto from * contradicting_in_cop[OF leq] contradicting_in_cop[OF leq'] show \?thesis\ unfolding \ \' by metis qed theorem cop_correct_nterm: assumes obsne: \obs \ k \ obs \' k\ \obs \ k \ None\ \obs \' k \ None\ and leq: \\ =\<^sub>L \'\ shows \\ k k'. ((\,k),\',k')\ cop \ ((\',k'),\,k)\ cop\ proof - obtain k k' where \((\', k') \ (\ ,k) \ path \ k \ dom att) \ ((\, k) \ (\' ,k') \ path \' k' \ dom att)\ using obs_neq_some_contradicts[OF obsne] by metis thus \?thesis\ proof assume *: \(\', k') \ (\ ,k) \ path \ k \ dom att\ hence \((\,k),\',k') \ cop\ using leq by (metis contradicting_in_cop) thus \?thesis\ using * by blast next assume *: \(\, k) \ (\' ,k') \ path \' k' \ dom att\ hence \((\',k'),\,k) \ cop\ using leq by (metis contradicting_in_cop loweq_def) thus \?thesis\ using * by blast qed qed subsection \Correctness of the Characterisation\ text_raw \\label{sec:cor-cp}\ text \The following is our main correctness result. If there exist no critical observable paths, then the program is secure.\ theorem cop_correct: assumes \cop = empty\ shows \secure\ proof (rule ccontr) assume \\ secure\ then obtain \ \' where leq: \ \ =\<^sub>L \'\ and **: \\ obs \ \ obs \' \ (terminates \ \ \ obs \' \ obs \)\ unfolding secure_def by blast show \False\ using ** proof assume \\ obs \ \ obs \'\ then obtain k where \obs \ k \ obs \' k \ obs \ k \ None \ obs \' k \ None\ unfolding obs_comp_def obs_prefix_def by (metis kth_obs_stable linorder_neqE_nat obs_none_no_kth_obs obs_some_kth_obs) thus \False\ using cop_correct_nterm leq assms by auto next assume *: \terminates \ \ \ obs \' \ obs \\ then obtain n where ret: \path \ n = return\ unfolding terminates_def by auto obtain k where \obs \ k \ obs \' k \ obs \' k \ None\ using * unfolding obs_prefix_def by metis thus \False\ using cop_correct_ret ret leq assms by (metis empty_iff) qed qed text \Our characterisation is not only correct, it is also precise in the way that \cp\ characterises exactly the matching indices in executions for low equivalent input states where diverging data is read. This follows easily as the inverse implication to lemma \contradicting_in_cp\ can be shown by simple induction.\ theorem cp_iff_reads_contradict: \((\,k),(\',k')) \ cp \ \ =\<^sub>L \' \ cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k' \ (\ v\reads(path \ k). (\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v)\ proof assume \\ =\<^sub>L \' \ cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k' \ (\v\reads (path \ k). (\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v)\ thus \((\, k), \', k') \ cp\ using contradicting_in_cp by blast next assume \((\, k), \', k') \ cp\ thus \\ =\<^sub>L \' \ cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k' \ (\v\reads (path \ k). (\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v)\ proof (induction) case (1 \ \' n n' h) then show \?case\ by blast next case (2 \ k \' k' n v n') have \v\reads (path \ n)\ using 2(2) unfolding is_ddi_def by auto then show \?case\ using 2 by auto next case (3 \ k \' k' n v l n') have \v\reads (path \ n)\ using 3(2) unfolding is_ddi_def by auto then show \?case\ using 3(4,6,8) by auto next case (4 \ k \' k') hence \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ by simp hence \path \' k' = path \ k\ by (metis last_cs) moreover have \\' =\<^sub>L \\ using 4(2) unfolding loweq_def by simp ultimately show \?case\ using 4 by metis qed qed text \In the same way the inverse implication to \contradicting_in_cop\ follows easily such that we obtain the following characterisation of \cop\.\ theorem cop_iff_contradicting: \((\,k),(\',k')) \ cop \ \ =\<^sub>L \' \ (\',k') \ (\,k) \ path \ k \ dom att\ proof assume \\ =\<^sub>L \' \ (\', k') \ (\, k) \ path \ k \ dom att\ thus \((\,k),(\',k')) \ cop\ using contradicting_in_cop by simp next assume \((\,k),(\',k')) \ cop\ thus \ \ =\<^sub>L \' \ (\',k') \ (\,k) \ path \ k \ dom att\ proof (cases rule: cop.cases) case 1 then show \?thesis\ using cp_iff_reads_contradict contradicts.simps by (metis (full_types) reads_restrict1) next case (2 k) then show \?thesis\ using cp_iff_reads_contradict contradicts.simps by (metis cd_is_cs_less cd_not_ret contradicts.intros(1) cs_select_id path_is_path) qed qed subsection \Correctness of the Single Path Approximation\ text_raw \\label{sec:cor-scp}\ theorem cp_in_scp: assumes \((\,k),(\',k'))\cp\ shows \(path \,k)\scp \ (path \',k')\scp\ using assms proof(induction \\\ \k\ \\'\ \k'\ rule:cp.induct[case_names read_high dd dcd sym]) case (read_high \ \' k k' h) have \\ h = (\\<^bsup>k\<^esup>) h\ using read_high(5) by (simp add: no_writes_unchanged0) moreover have \\' h = (\'\<^bsup>k'\<^esup>) h\ using read_high(6) by (simp add: no_writes_unchanged0) ultimately have \\ h \ \' h\ using read_high(4) by simp hence *: \h\hvars\ using read_high(1) unfolding loweq_def by (metis Compl_iff IFC_def.restrict_def) have 1: \(path \,k)\scp\ using scp.intros(1) read_high(3,5) * by auto have \path \ k = path \' k'\ using read_high(2) by (metis last_cs) hence \(path \',k')\scp\ using scp.intros(1) read_high(3,6) * by auto thus \?case\ using 1 by auto next case dd show \?case\ using scp.intros(3) dd by auto next case sym thus \?case\ by blast next case (dcd \ k \' k' n v l n') note scp.intros(4) is_dcdi_via_def cd_cs_swap cs_ipd have 1: \(path \, n)\scp\ using dcd.IH dcd.hyps(2) dcd.hyps(3) scp.intros(2) scp.intros(3) by blast have csk: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ using cp_eq_cs[OF dcd(1)] . have kn: \k and kl: \k and ln: \l using dcd(2,3) unfolding is_ddi_def is_cdi_def by auto have nret: \path \ k \ return\ using cd_not_ret dcd.hyps(3) by auto have \k' < n'\ using kn csk dcd(4) cs_order nret path_is_path last_cs by blast have 2: \(path \', n')\scp\ proof cases assume j'ex: \\j'\{k'.. writes (path \' j')\ hence \\j'. j'\{k'.. v \ writes (path \' j')\ by auto note * = GreatestI_ex_nat[OF this] define j' where \j' == GREATEST j'. j'\{k'.. v \ writes (path \' j')\ note ** = *[of \j'\,folded j'_def] have \k' \ j'\ \j' and j'write: \v \ writes (path \' j')\ using "*" atLeastLessThan_iff j'_def nat_less_le by auto have nowrite: \\ i'\{j'<.. writes(path \' i')\ proof (rule, rule ccontr) fix i' assume \i' \ {j'<.. \\ v \ local.writes (path \' i')\ hence \i' \ {k'.. v \ local.writes (path \' i')\ using \k' \ j'\ by auto hence \i' \ j'\ using Greatest_le_nat by (metis (no_types, lifting) atLeastLessThan_iff j'_def nat_less_le) thus \False\ using \i' \ {j'<.. by auto qed have \path \' n' = path \ n\ using dcd(4) last_cs by metis hence \v\reads(path \' n')\ using dcd(2) unfolding is_ddi_def by auto hence nddj': \n' dd\<^bsup>path \',v\<^esup>\ j'\ using dcd(2) unfolding is_ddi_def using nowrite \j' j'write by auto show \?thesis\ proof cases assume \j' cd\<^bsup>path \'\<^esup>\ k'\ thus \(path \',n') \ scp\ using scp.intros(2) scp.intros(3) dcd.IH nddj' by fast next assume jcdk': \\ j' cd\<^bsup>path \'\<^esup>\ k'\ show \?thesis\ proof cases assume \j' = k'\ thus \?thesis\ using scp.intros(3) dcd.IH nddj' by fastforce next assume \j' \ k'\ hence \k' < j'\ using \k' \ j'\ by auto have \path \' j' \ return\ using j'write writes_return by auto hence ipdex':\\j. j \{k'..j'} \ path \' j = ipd (path \' k') \ using path_is_path \k' < j'\ jcdk' is_cdi_def by blast define i' where \i' == LEAST j. j\ {k'..j'} \ path \' j = ipd (path \' k')\ have iipd': \i'\ {k'..j'}\ \path \' i' = ipd (path \' k')\ unfolding i'_def using LeastI_ex[OF ipdex'] by simp_all have *:\\ i \ {k'..' i \ ipd (path \' k')\ proof (rule, rule ccontr) fix i assume *: \i \ {k'.. \\ path \' i \ ipd (path \' k')\ hence **: \i \{k'..j'} \ path \' i = ipd (path \' k')\ (is \?P i\) using iipd'(1) by auto thus \False\ using Least_le[of \?P\ \i\] i'_def * by auto qed have \i' \ k'\ using iipd'(2) by (metis csk last_cs nret path_in_nodes ipd_not_self) hence \k' using iipd'(1) by simp hence csi': \cs\<^bsup>path \'\<^esup> i' = [n\cs\<^bsup>path \'\<^esup> k' . ipd n \ path \' i'] @ [path \' i']\using cs_ipd[OF iipd'(2) *] by fast have ncdk': \\ n' cd\<^bsup>path \'\<^esup>\ k'\ using \j' < n'\ \k' < j'\ cdi_prefix jcdk' less_imp_le_nat by blast hence ncdk: \\ n cd\<^bsup>path \\<^esup>\ k\ using cd_cs_swap csk dcd(4) by blast have ipdex: \\i. i\{k..n} \ path \ i = ipd (path \ k)\ (is \\i. ?P i\) proof cases assume *:\path \ n = return\ from path_ret_ipd[of \path \\ \k\ \n\,OF path_is_path nret *] obtain i where \?P i\ by fastforce thus \?thesis\ by auto next assume *:\path \ n \ return\ show \?thesis\ using not_cd_impl_ipd [of \path \\ \k\ \n\, OF path_is_path \k ncdk *] by auto qed define i where \i == LEAST j. j\ {k..n} \ path \ j = ipd (path \ k)\ have iipd: \i\ {k..n}\ \path \ i = ipd (path \ k)\ unfolding i_def using LeastI_ex[OF ipdex] by simp_all have **:\\ i' \ {k.. i' \ ipd (path \ k)\ proof (rule, rule ccontr) fix i' assume *: \i' \ {k.. \\ path \ i' \ ipd (path \ k)\ hence **: \i' \{k..n} \ path \ i' = ipd (path \ k)\ (is \?P i'\) using iipd(1) by auto thus \False\ using Least_le[of \?P\ \i'\] i_def * by auto qed have \i \ k\ using iipd(2) by (metis nret path_in_nodes ipd_not_self) hence \k using iipd(1) by simp hence \cs\<^bsup>path \\<^esup> i = [n\cs\<^bsup>path \\<^esup> k . ipd n \ path \ i] @ [path \ i]\using cs_ipd[OF iipd(2) **] by fast hence csi: \cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i'\ using csi' csk unfolding iipd'(2) iipd(2) by (metis last_cs) hence \(LEAST i'. k' < i' \ (\i. cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i')) \ i'\ (is \(LEAST x. ?P x) \ _\) using \k' < i'\ Least_le[of \?P\ \i'\] by blast hence nw: \\j'\{i'.. writes (path \' j')\ using dcd(7) allB_atLeastLessThan_lower by blast moreover have \v \ writes (path \' j')\ using nddj' unfolding is_ddi_def by auto moreover have \i' \ j'\ using iipd'(1) by auto ultimately have \False\ using \j' < n'\ by auto thus \?thesis\ .. qed qed next assume \\ (\j'\{k'.. writes (path \' j'))\ hence \n' dcd\<^bsup>path \',v\<^esup>\ k' via (path \) k\ unfolding is_dcdi_via_def using dcd(2-4) csk \k' path_is_path by metis thus \?thesis\ using dcd.IH scp.intros(4) by blast qed with 1 show \?case\ .. qed theorem cop_in_scop: assumes \((\,k),(\',k'))\cop\ shows \(path \,k)\scop \ (path \',k')\scp\ using assms apply (induct rule: cop.induct) apply (simp add: cp_in_scp) using cp_in_scp scop.intros scp.intros(2) apply blast using cp_in_scp scop.intros scp.intros(2) apply blast done text \The main correctness result for out single execution approximation follows directly.\ theorem scop_correct: assumes \scop = empty\ shows \secure\ using cop_correct assms cop_in_scop by fast end end \ No newline at end of file diff --git a/thys/IMP2/doc/Examples.thy b/thys/IMP2/doc/Examples.thy --- a/thys/IMP2/doc/Examples.thy +++ b/thys/IMP2/doc/Examples.thy @@ -1,1872 +1,1873 @@ theory Examples imports "../IMP2" "../lib/IMP2_Aux_Lemmas" begin section \Examples\ lemmas nat_distribs = nat_add_distrib nat_diff_distrib Suc_diff_le nat_mult_distrib nat_div_distrib subsection \Common Loop Patterns\ subsubsection \Count Up\ text \ Counter \c\ counts from \0\ to \n\, such that loop is executed \n\ times. The result is computed in an accumulator \a\. \ text \The invariant states that we have computed the function for the counter value \c\\ text \The variant is the difference between \n\ and \c\, i.e., the number of loop iterations that we still have to do\ program_spec exp_count_up assumes "0\n" ensures "a = 2^nat n\<^sub>0" defines \ a = 1; c = 0; while (cn-c\ @invariant \0\c \ c\n \ a=2^nat c\ { G_par = a; scope { a = G_par; a=2*a; G_ret = a }; a = G_ret; c=c+1 } \ apply vcg by (auto simp: algebra_simps nat_distribs) program_spec sum_prog assumes "n \ 0" ensures "s = \{0..n\<^sub>0}" defines \ s = 0; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n\<^sub>0 = n \ 0 \ i \ i \ n \ s = \{0..i}\ { i = i + 1; s = s + i } \ apply vcg_cs by (simp_all add: intvs_incdec) program_spec sq_prog assumes "n \ 0" ensures "a = n\<^sub>0 * n\<^sub>0" defines \ a = 0; z = 1; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n\<^sub>0 = n \ 0 \ i \ i \ n \ a = i * i \ z = 2 * i + 1\ { a = a + z; z = z + 2; i = i + 1 } \ by vcg_cs (simp add: algebra_simps) fun factorial :: "int \ int" where "factorial i = (if i \ 0 then 1 else i * factorial (i - 1))" program_spec factorial_prog assumes "n \ 0" ensures "a = factorial n\<^sub>0" defines \ a = 1; i = 1; while (i \ n) @variant \n\<^sub>0 + 1 - i\ @invariant \n\<^sub>0 = n \ 1 \ i \ i \ n + 1 \ a = factorial (i - 1)\ { a = a * i; i = i + 1 } \ by vcg (simp add: antisym_conv)+ fun fib :: "int \ int" where "fib i = (if i \ 0 then 0 else if i = 1 then 1 else fib (i - 2) + fib (i - 1))" lemma fib_simps[simp]: "i \ 0 \ fib i = 0" "i = 1 \ fib i = 1" "i > 1 \ fib i = fib (i - 2) + fib (i - 1)" by simp+ lemmas [simp del] = fib.simps text \With precondition\ program_spec fib_prog assumes "n \ 0" ensures "a = fib n" defines \ a = 0; b = 1; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n = n\<^sub>0 \ 0 \ i \ i \ n \ a = fib i \ b = fib (i + 1)\ { c = b; b = a + b; a = c; i = i + 1 } \ by vcg_cs (simp add: algebra_simps) text \Without precondition, returning \0\ for negative numbers\ program_spec fib_prog' assumes True ensures "a = fib n\<^sub>0" defines \ a = 0; b = 1; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n = n\<^sub>0 \ (0 \ i \ i \ n \ n\<^sub>0 < 0 \ i = 0) \ a = fib i \ b = fib (i+ 1)\ { c = b; b = a + b; a = c; i = i + 1 } \ by vcg_cs (auto simp: algebra_simps) subsubsection \Count down\ text \Essentially the same as count up, but we (ab)use the input variable as a counter.\ text \The invariant is the same as for count-up. Only that we have to compute the actual number of loop iterations by \n\<^sub>0 - n\. We locally introduce the name \c\ for that. \ program_spec exp_count_down assumes "0\n" ensures "a = 2^nat n\<^sub>0" defines \ a = 1; while (n>0) @variant \n\ @invariant \let c = n\<^sub>0-n in 0\n \ n\n\<^sub>0 \ a=2^nat c\ { a=2*a; n=n-1 } \ apply vcg_cs by (auto simp: algebra_simps nat_distribs) subsubsection \Approximate from Below\ text \Used to invert a monotonic function. We count up, until we overshoot the desired result, then we subtract one. \ text \The invariant states that the \r-1\ is not too big. When the loop terminates, \r-1\ is not too big, but \r\ is already too big, so \r-1\ is the desired value (rounding down). \ text \The variant measures the gap that we have to the correct result. Note that the loop will do a final iteration, when the result has been reached exactly. We account for that by adding one, such that the measure also decreases in this case. \ program_spec sqr_approx_below assumes "0\n" ensures "0\r \ r\<^sup>2 \ n\<^sub>0 \ n\<^sub>0 < (r+1)\<^sup>2" defines \ r=1; while (r*r \ n) @variant \n + 1 - r*r\ @invariant \0\r \ (r-1)\<^sup>2 \ n\<^sub>0\ { r = r + 1 }; r = r - 1 \ apply vcg by (auto simp: algebra_simps power2_eq_square) subsubsection \Bisection\ text \A more efficient way of inverting monotonic functions is by bisection, that is, one keeps track of a possible interval for the solution, and halfs the interval in each step. The program will need $O(\log n)$ iterations, and is thus very efficient in practice. Although the final algorithm looks quite simple, getting it right can be quite tricky. \ text \The invariant is surprisingly easy, just stating that the solution is in the interval \l... \ lemma "\h l n\<^sub>0 :: int. \\''invar-final''; 0 \ n\<^sub>0; \ 1 + l < h; 0 \ l; l < h; l * l \ n\<^sub>0; n\<^sub>0 < h * h\ \ n\<^sub>0 < 1 + (l * l + l * 2)" by (smt mult.commute semiring_normalization_rules(3)) program_spec sqr_bisect assumes "0\n" ensures "r\<^sup>2\n\<^sub>0 \ n\<^sub>0<(r+1)\<^sup>2" defines \ l=0; h=n+1; while (l+1 < h) @variant \h-l\ @invariant \0\l \ l l\<^sup>2\n \ n < h\<^sup>2\ { m = (l + h) / 2; if (m*m \ n) l=m else h=m }; r=l \ apply vcg text \We use quick-and-dirty apply style proof to discharge the VCs\ apply (auto simp: power2_eq_square algebra_simps add_pos_pos) apply (smt not_sum_squares_lt_zero) by (smt mult.commute semiring_normalization_rules(3)) subsection \Debugging\ subsubsection \Testing Programs\ text \Stepwise\ schematic_goal "Map.empty: (sqr_approx_below,<''n'':=\_. 4>) \ ?s" unfolding sqr_approx_below_def apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' done text \Or all steps at once\ schematic_goal "Map.empty: (sqr_bisect,<''n'':=\_. 4900000001>) \ ?s" unfolding sqr_bisect_def by big_step subsection \More Numeric Algorithms\ subsubsection \Euclid's Algorithm (with subtraction)\ (* Crucial Lemmas *) thm gcd.commute gcd_diff1 program_spec euclid1 assumes "a>0 \ b>0" ensures "a = gcd a\<^sub>0 b\<^sub>0" defines \ while (a\b) @invariant \gcd a b = gcd a\<^sub>0 b\<^sub>0 \ (a>0 \ b>0)\ @variant \a+b\ { if (a apply vcg_cs apply (metis gcd.commute gcd_diff1) apply (metis gcd_diff1) done subsubsection \Euclid's Algorithm (with mod)\ (* Crucial Lemmas *) thm gcd_red_int[symmetric] program_spec euclid2 assumes "a>0 \ b>0" ensures "a = gcd a\<^sub>0 b\<^sub>0" defines \ while (b\0) @invariant \gcd a b = gcd a\<^sub>0 b\<^sub>0 \ b\0 \ a>0\ @variant \b\ { t = a; a = b; b = t mod b } \ apply vcg_cs by (simp add: gcd_red_int[symmetric]) subsubsection \Extended Euclid's Algorithm\ (* Provided by Simon Wimmer. *) locale extended_euclid_aux_lemmas begin lemma aux2: fixes a b :: int assumes "b = t * b\<^sub>0 + s * a\<^sub>0" "q = a div b" "gcd a b = gcd a\<^sub>0 b\<^sub>0" shows "gcd b (a - (a\<^sub>0 * (s * q) + b\<^sub>0 * (t * q))) = gcd a\<^sub>0 b\<^sub>0" proof - have "a - (a\<^sub>0 * (s * q) + b\<^sub>0 * (t * q)) = a - b * q" unfolding \b = _\ by (simp add: algebra_simps) also have "a - b * q = a mod b" unfolding \q = _\ by (simp add: algebra_simps) finally show ?thesis unfolding \gcd a b = _\[symmetric] by (simp add: gcd_red_int[symmetric]) qed lemma aux3: fixes a b :: int assumes "b = t * b\<^sub>0 + s * a\<^sub>0" "q = a div b" "b > 0" shows "t * (b\<^sub>0 * q) + s * (a\<^sub>0 * q) \ a" proof - have "t * (b\<^sub>0 * q) + s * (a\<^sub>0 * q) = q * b" unfolding \b = _\ by (simp add: algebra_simps) then show ?thesis using \b > 0\ by (simp add: algebra_simps \q = _\) (smt Euclidean_Division.pos_mod_sign cancel_div_mod_rules(1) mult.commute) qed end text \The following is a direct translation of the pseudocode for the Extended Euclidean algorithm as described by the English version of Wikipedia (\<^url>\https://en.wikipedia.org/wiki/Extended_Euclidean_algorithm\): \ program_spec euclid_extended assumes "a>0 \ b>0" ensures "old_r = gcd a\<^sub>0 b\<^sub>0 \ gcd a\<^sub>0 b\<^sub>0 = a\<^sub>0 * old_s + b\<^sub>0 * old_t" defines \ s = 0; old_s = 1; t = 1; old_t = 0; r = b; old_r = a; while (r\0) @invariant \ gcd old_r r = gcd a\<^sub>0 b\<^sub>0 \ r\0 \ old_r>0 \ a\<^sub>0 * old_s + b\<^sub>0 * old_t = old_r \ a\<^sub>0 * s + b\<^sub>0 * t = r \ @variant \r\ { quotient = old_r / r; temp = old_r; old_r = r; r = temp - quotient * r; temp = old_s; old_s = s; s = temp - quotient * s; temp = old_t; old_t = t; t = temp - quotient * t } \ proof - interpret extended_euclid_aux_lemmas . show ?thesis apply vcg_cs apply (simp add: algebra_simps) apply (simp add: aux2 aux3 minus_div_mult_eq_mod)+ done qed text \Non-Wikipedia version\ context extended_euclid_aux_lemmas begin lemma aux: fixes a b q x y:: int assumes "a = old_y * b\<^sub>0 + old_x * a\<^sub>0" "b = y * b\<^sub>0 + x * a\<^sub>0" "q = a div b" shows "a mod b + (a\<^sub>0 * (x * q) + b\<^sub>0 * (y * q)) = a" proof - have *: "a\<^sub>0 * (x * q) + b\<^sub>0 * (y * q) = q * b" unfolding \b = _\ by (simp add: algebra_simps) show ?thesis unfolding * unfolding \q = _\ by simp qed end program_spec euclid_extended' assumes "a>0 \ b>0" ensures "a = gcd a\<^sub>0 b\<^sub>0 \ gcd a\<^sub>0 b\<^sub>0 = a\<^sub>0 * x + b\<^sub>0 * y" defines \ x = 0; y = 1; old_x = 1; old_y = 0; while (b\0) @invariant \ gcd a b = gcd a\<^sub>0 b\<^sub>0 \ b\0 \ a>0 \ a = a\<^sub>0 * old_x + b\<^sub>0 * old_y \ b = a\<^sub>0 * x + b\<^sub>0 * y \ @variant \b\ { q = a / b; t = a; a = b; b = t mod b; t = x; x = old_x - q * x; old_x = t; t = y; y = old_y - q * y; old_y = t }; x = old_x; y = old_y \ proof - interpret extended_euclid_aux_lemmas . show ?thesis apply vcg_cs apply (simp add: gcd_red_int[symmetric]) apply (simp add: algebra_simps) apply (rule aux; simp add: algebra_simps) done qed subsubsection \Exponentiation by Squaring\ (* Contributed by Simon Wimmer *) lemma ex_by_sq_aux: fixes x :: int and n :: nat assumes "n mod 2 = 1" shows "x * (x * x) ^ (n div 2) = x ^ n" proof - have "n > 0" using assms by presburger have "2 * (n div 2) = n - 1" using assms by presburger then have "(x * x) ^ (n div 2) = x ^ (n - 1)" by (simp add: semiring_normalization_rules) with \0 < n\ show ?thesis by simp (metis Suc_pred power.simps(2)) qed text \A classic algorithm for computing \x\<^sup>n\ works by repeated squaring, using the following recurrence: \<^item> \x\<^sup>n = x * x\<^bsup>(n-1)/2\<^esup>\ if \n\ is odd \<^item> \x\<^sup>n = x\<^bsup>n/2\<^esup>\ if \n\ is even \ program_spec ex_by_sq assumes "n\0" ensures "r = x\<^sub>0 ^ nat n\<^sub>0" defines \ r = 1; while (n\0) @invariant \ n \ 0 \ r * x ^ nat n = x\<^sub>0 ^ nat n\<^sub>0 \ @variant \n\ { if (n mod 2 == 1) { r = r * x }; x = x * x; n = n / 2 } \ apply vcg_cs subgoal apply (subst (asm) ex_by_sq_aux[symmetric]) apply (smt Suc_1 nat_1 nat_2 nat_mod_distrib nat_one_as_int) by (simp add: div_int_pos_iff int_div_less_self nat_div_distrib) apply (simp add: algebra_simps semiring_normalization_rules nat_mult_distrib; fail) apply (simp add: algebra_simps semiring_normalization_rules nat_mult_distrib; fail) done subsubsection \Power-Tower of 2s\ fun tower2 where "tower2 0 = 1" | "tower2 (Suc n) = 2 ^ tower2 n" definition "tower2' n = int (tower2 (nat n))" program_spec tower2_imp assumes \m>0\ ensures \a = tower2' m\<^sub>0\ defines \ a=1; while (m>0) @variant \m\ @invariant \0\m \ m\m\<^sub>0 \ a = tower2' (m\<^sub>0-m)\ { n=a; a = 1; while (n>0) @variant \n\ @invariant \True\ \ \This will get ugly, there is no \n\<^sub>0\ that we could use!\ { a=2*a; n=n-1 }; m=m-1 } \ oops text \We prove the inner loop separately instead! (It happens to be exactly our \<^const>\exp_count_down\ program.) \ program_spec tower2_imp assumes \m>0\ ensures \a = tower2' m\<^sub>0\ defines \ a=1; while (m>0) @variant \m\ @invariant \0\m \ m\m\<^sub>0 \ a = tower2' (m\<^sub>0-m)\ { n=a; inline exp_count_down; m=m-1 } \ apply vcg_cs by (auto simp: algebra_simps tower2'_def nat_distribs) subsection \Array Algorithms\ (* Presentation in lecture: introduce array syntax and semantics CONVENTIONS: * Every variable is an array. * Arrays are modeled as int\int. Negative indexes allowed. * Syntactic sugar to use it at index 0 by default V vname \ Vidx vname aexp Syntax: x[i] abbreviation "V x = Vidx x 0" Assign vname aexp \ AssignIdx vname aexp aexp Syntax: x[i] = a abbreviation: Assign x a = AssignIdx x 0 a NEW COMMAND: ArrayCpy vname vname Syntax: x[] = y copy two arrays as a whole. Assign x (V y) only copies index 0 semantics: Obvious aval (Vidx x i) s = s x (aval i s) (x[i] = a, s) \ s(x := (s x)(aval i s := aval a s)) (x[] = y, s) \ s(x := s y) weakest preconditions: Also obvious \ Reasoning with arrays: Usually: Model as int\int is appropriate. Common idioms: * {l..i Warning: Sledgehammer does not perform well on these! Data Refinement Sometimes, abstraction to list may be useful. Idea: Model an algorithm on lists first. Prove it correct. Then implement algorithm. Show that it implements the model. By transitivity, get: Algorithm is correct! Sometimes, abstraction to list may be useful *) subsubsection \Summation\ program_spec array_sum assumes "l\h" ensures "r = (\i=l\<^sub>0..0. a\<^sub>0 i)" defines \ r = 0; while (ll\<^sub>0\l \ l\h \ r = (\i=l\<^sub>0..0 i)\ @variant \h-l\ { r = r + a[l]; l=l+1 } \ apply vcg_cs by (auto simp: intvs_incdec) (* Exercise: Remove l\h precondition! *) subsubsection \Finding Least Index of Element\ program_spec find_least_idx assumes \l\h\ ensures \if l=h\<^sub>0 then x\<^sub>0 \ a\<^sub>0`{l\<^sub>0..0} else l\{l\<^sub>0..0} \ a\<^sub>0 l = x\<^sub>0 \ x\<^sub>0\a\<^sub>0`{l\<^sub>0.. defines \ while (l a[l] \ x) @variant \h-l\ @invariant \l\<^sub>0\l \ l\h \ x\a`{l\<^sub>0.. l=l+1 \ apply vcg_cs by (smt atLeastLessThan_iff imageI) subsubsection \Check for Sortedness\ term ran_sorted program_spec check_sorted assumes \l\h\ ensures \r\0 \ ran_sorted a\<^sub>0 l\<^sub>0 h\<^sub>0\ defines \ if (l==h) r=1 else { l=l+1; while (l a[l-1] \ a[l]) @variant \h-l\ @invariant \l\<^sub>0 l\h \ ran_sorted a l\<^sub>0 l\ l=l+1; if (l==h) r=1 else r=0 } \ apply vcg_cs apply (auto simp: ran_sorted_def) by (smt atLeastLessThan_iff) subsubsection \Find Equilibrium Index\ definition "is_equil a l h i \ l\i \ i (\j=l..j=i..l\h\ ensures \is_equil a l h i \ i=h \ (\i. is_equil a l h i)\ defines \ usum=0; i=l; while (ih-i\ @invariant \l\i \ i\h \ usum = (\j=l.. { usum = usum + a[i]; i=i+1 }; i=l; lsum=0; while (usum \ lsum \ ih-i\ @invariant \l\i \ i\h \ lsum=(\j=l.. usum=(\j=i.. (\jis_equil a l h j) \ { lsum = lsum + a[i]; usum = usum - a[i]; i=i+1 } \ apply vcg_cs apply (auto simp: intvs_incdec is_equil_def) apply (metis atLeastLessThan_iff eq_iff finite_atLeastLessThan_int sum_diff1) apply force by force subsubsection \Rotate Right\ program_spec rotate_right assumes "0i\{0..0 ((i-1) mod n)" defines \ i = 0; prev = a[n - 1]; while (i < n) @invariant \ 0 \ i \ i \ n \ (\j\{0..0 ((j-1) mod n)) \ (\j\{i..0 j) \ prev = a\<^sub>0 ((i-1) mod n) \ @variant \n - i\ { temp = a[i]; a[i] = prev; prev = temp; i = i + 1 } \ apply vcg_cs by (simp add: zmod_minus1) subsubsection \Binary Search, Leftmost Element\ text \We first specify the pre- and postcondition\ definition "bin_search_pre a l h \ l\h \ ran_sorted a l h" definition "bin_search_post a l h x i \ l\i \ i\h \ (\i\{l.. (\i\{i.. a i)" text \Then we prove that the program is correct\ program_spec binsearch assumes \bin_search_pre a l h\ ensures \bin_search_post a\<^sub>0 l\<^sub>0 h\<^sub>0 x\<^sub>0 l\ defines \ while (l < h) @variant \h-l\ @invariant \l\<^sub>0\l \ l\h \ h\h\<^sub>0 \ (\i\{l\<^sub>0.. (\i\{h..0}. x \ a i)\ { m = (l + h) / 2; if (a[m] < x) l = m + 1 else h = m } \ apply vcg_cs apply (auto simp: algebra_simps bin_search_pre_def bin_search_post_def) txt \Driving sledgehammer to its limits ...\ apply (smt div_add_self1 even_succ_div_two odd_two_times_div_two_succ ran_sorted_alt) by (smt div_add_self1 even_succ_div_two odd_two_times_div_two_succ ran_sorted_alt) text \Next, we show that our postcondition (which was easy to prove) implies the expected properties of the algorithm. \ lemma assumes "bin_search_pre a l h" "bin_search_post a l h x i" shows bin_search_decide_membership: "x\a`{l.. (i x = a i)" and bin_search_leftmost: "x\a`{l..Naive String Search\ (* By Simon Wimmer *) program_spec match_string assumes "l1 \ h1" ensures "(\j \ {0.. (i < h1 - l1 \ a (l + i) \ b (l1 + i)) \ 0 \ i \ i \ h1 - l1" defines \ i = 0; while (l1 + i < h1 \ a[l + i] == b[l1 + i]) @invariant \(\j \ {0.. 0 \ i \ i \ h1 - l1\ @variant \(h1 - (l1 + i))\ { i = i + 1 } \ by vcg_cs auto lemma lran_eq_iff': "lran a l1 (l1 + (h - l)) = lran a' l h \ (\i. 0 \ i \ i < h - l \ a (l1 + i) = a' (l + i))" if "l \ h" using that proof (induction "nat (h - l)" arbitrary: h) case 0 then show ?case by auto next case (Suc x) then have *: "x = nat (h - 1 - l)" "l \ h - 1" by auto note IH = Suc.hyps(1)[OF *] from * have 1: "lran a l1 (l1 + (h - l)) = lran a l1 (l1 + (h - 1 - l)) @ [a (l1 + (h - 1 - l))]" "lran a' l h = lran a' l (h - 1) @ [a' (h - 1)]" by (auto simp: lran_bwd_simp algebra_simps lran_butlast[symmetric]) from IH * Suc.hyps(2) Suc.prems show ?case unfolding 1 apply auto subgoal for i by (cases "i = h - 1 - l") auto done qed program_spec match_string' assumes "l1 \ h1" ensures "i = h1 - l1 \ lran a l (l + (h1 - l1)) = lran b l1 h1" for i h1 l1 l a[] b[] defines \inline match_string\ by vcg_cs (auto simp: lran_eq_iff') program_spec substring assumes "l \ h \ l1 \ h1" ensures "match = 1 \ (\ j \ {l\<^sub>0..0}. lran a j (j + (h1 - l1)) = lran b l1 h1)" for a[] b[] defines \ match = 0; while (l < h \ match == 0) @invariant\l\<^sub>0 \ l \ l \ h \ match \ {0,1} \ (if match = 1 then lran a l (l + (h1 - l1)) = lran b l1 h1 \ l < h else (\j \ {l\<^sub>0.. lran b l1 h1))\ @variant\(h - l) * (1 - match)\ { inline match_string'; if (i == h1 - l1) {match = 1} else {l = l + 1} } \ by vcg_cs auto program_spec substring' assumes "l \ h \ l1 \ h1" ensures "match = 1 \ (\ j \ {l\<^sub>0..h\<^sub>0-(h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1)" for a[] b[] defines \ match = 0; if (l + (h1 - l1) \ h) { h = h - (h1 - l1) + 1; inline substring } \ by vcg_cs auto (* Or combined: *) program_spec substring'' assumes "l \ h \ l1 \ h1" ensures "match = 1 \ (\ j \ {l\<^sub>0..0-(h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1)" for a[] b[] defines \ match = 0; if (l + (h1 - l1) \ h) { while (l + (h1 - l1) < h \ match == 0) @invariant\l\<^sub>0 \ l \ l \ h - (h1 - l1) \ match \ {0,1} \ (if match = 1 then lran a l (l + (h1 - l1)) = lran b l1 h1 \ l < h - (h1 - l1) else (\j \ {l\<^sub>0.. lran b l1 h1))\ @variant\(h - l) * (1 - match)\ { inline match_string'; if (i == h1 - l1) {match = 1} else {l = l + 1} } } \ by vcg_cs auto lemma lran_split: "lran a l h = lran a l p @ lran a p h" if "l \ p" "p \ h" using that by (induction p; simp; simp add: lran.simps) lemma lran_eq_append_iff: "lran a l h = as @ bs \ (\ i. l \ i \ i \ h \ as = lran a l i \ bs = lran a i h)" if "l \ h" apply safe subgoal using that proof (induction as arbitrary: l) case Nil then show ?case by auto next case (Cons x as) from this(2-) have "lran a (l + 1) h = as @ bs" "a l = x" "l + 1 \ h" apply - subgoal by simp subgoal by (smt append_Cons list.inject lran.simps lran_append1) subgoal using add1_zle_eq by fastforce done - from Cons.IH[OF this(1,3)] guess i by safe - note IH = this + from Cons.IH[OF this(1,3)] obtain i + where IH: "l + 1 \ i" "i \ h" "as = lran a (l + 1) i" "bs = lran a i h" + by auto with \a l = x\ show ?case apply (intro exI[where x = i]) apply auto by (smt IH(3) lran_prepend1) qed apply (subst lran_split; simp) done lemma lran_split': "(\j\{l..h - (h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1) = (\as bs. lran a l h = as @ lran b l1 h1 @ bs)" if "l \ h" "l1 \ h1" proof safe fix j assume j: "j \ {l..h - (h1 - l1)}" and match: "lran a j (j + (h1 - l1)) = lran b l1 h1" with \l1 \ h1\ have "lran a l h = lran a l j @ lran a j (j + (h1 - l1)) @ lran a (j + (h1 - l1)) h" apply (subst lran_split[where p = j], simp, simp) apply (subst (2) lran_split[where p = "j + (h1 - l1)"]; simp) done then show "\as bs. lran a l h = as @ lran b l1 h1 @ bs" by (auto simp: match) next fix as bs assume "lran a l h = as @ lran b l1 h1 @ bs" with that lran_eq_append_iff[of l h a as "lran b l1 h1 @ bs"] obtain i where "as = lran a l i" "lran a i h = lran b l1 h1 @ bs" "l \ i" "i \ h" by auto with lran_eq_append_iff[of i h a "lran b l1 h1" bs] obtain j where j: "lran b l1 h1 = lran a i j" "bs = lran a j h" "i \ j" "j \ h" by auto moreover have "j = i + (h1 - l1)" proof - have "length (lran b l1 h1) = nat (h1 - l1)" "length (lran a i j) = nat (j - i)" by auto with j(1,3) that show ?thesis by auto qed ultimately show "\j\{l..h - (h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1" using \l \ i\ by auto qed program_spec substring_final assumes "l \ h \ 0 \ len" ensures "match = 1 \ (\as bs. lran a l\<^sub>0 h\<^sub>0 = as @ lran b 0 len @ bs)" for l h len match a[] b[] defines \l1 = 0; h1 = len; inline substring'\ supply [simp] = lran_split'[symmetric] apply vcg_cs done subsubsection \Insertion Sort\ text \ We first prove the inner loop. The specification here specifies what the algorithm does as closely as possible, such that it becomes easier to prove. In this case, sortedness is not a precondition for the inner loop to move the key element backwards over all greater elements. \ definition "insort_insert_post l j a\<^sub>0 a i \ (let key = a\<^sub>0 j in i\{l-1.. \\i\ is in range\ \ \Content of new array\ \ (\k\{l..i}. a k = a\<^sub>0 k) \ a (i+1) = key \ (\k\{i+2..j}. a k = a\<^sub>0 (k-1)) \ a = a\<^sub>0 on -{l..j} \ \Placement of \key\\ \ (i\l \ a i \ key) \ \Element at \i\ smaller than \key\, if it exists \ \ (\k\{i+2..j}. a k > key) \ \Elements \\i+2\ greater than \key\\ ) " for l j i :: int and a\<^sub>0 a :: "int \ int" program_spec insort_insert assumes "l0 a i" defines \ key = a[j]; i = j-1; while (i\l \ a[i]>key) @variant \i-l+1\ @invariant \l-1\i \ i (\k\{l..i}. a k = a\<^sub>0 k) \ (\k\{i+2..j}. a k > key \ a k = a\<^sub>0 (k-1)) \ a = a\<^sub>0 on -{l..j} \ { a[i+1] = a[i]; i=i-1 }; a[i+1] = key \ unfolding insort_insert_post_def Let_def apply vcg apply auto by (smt atLeastAtMost_iff) text \Next, we show that our specification that was easy to prove implies the specification that the outer loop expects:\ text \Invoking \insort_insert\ will sort in the element\ lemma insort_insert_sorted: assumes "lInvoking \insort_insert\ will only mutate the elements\ lemma insort_insert_ran1: assumes "li+1" "i+1\j" unfolding insort_insert_post_def Let_def by auto have ranshift: "mset_ran (a o (+)(-1)) {i+2..j} = mset_ran a {i+1..j-1}" by (simp add: mset_ran_shift algebra_simps) have "mset_ran a' {l..j} = mset_ran a' {l..i} + {# a' (i+1) #} + mset_ran a' {i+2..j}" using \l \l\i+1\ \i+1\j\ apply (simp add: mset_ran_combine) by (auto intro: arg_cong[where f="mset_ran a'"]) also have "\ = mset_ran a {l..i} + {# a j #} + mset_ran (a o (+)(-1)) {i+2..j}" using EQS(1,3)[THEN mset_ran_cong] EQS(2) by simp also have "\ = mset_ran a {l..i} + mset_ran a {i+1..j-1} + {# a j #}" by (simp add: mset_ran_shift algebra_simps) also have "\ = mset_ran a {l..j}" using \l \l\i+1\ \i+1\j\ apply (simp add: mset_ran_combine) by (auto intro: arg_cong[where f="mset_ran a"]) finally show ?thesis . qed text \The property @{thm insort_insert_ran1} extends to the whole array to be sorted\ lemma insort_insert_ran2: assumes "linsort_insert_post l j a a' i\ have "a' = a on {j<..Finally, we specify and prove correct the outer loop\ program_spec insort assumes "l mset_ran a {l..0 {l.. a=a\<^sub>0 on -{l.. j = l + 1; while (jh-j\ @invariant \ l j\h \ \\j\ in range\ \ ran_sorted a l j \ \Array is sorted up to \j\\ \ mset_ran a {l..0 {l.. \Elements in range only permuted\ \ a=a\<^sub>0 on -{l.. { inline insort_insert; j=j+1 } \ apply vcg_cs apply (intro conjI) subgoal by (rule insort_insert_sorted) subgoal using insort_insert_ran2(1) by auto subgoal apply (frule (2) insort_insert_ran2(2)) by (auto simp: eq_on_def) done subsubsection \Quicksort\ procedure_spec partition_aux(a,l,h,p) returns (a,i) assumes "l\h" ensures "mset_ran a\<^sub>0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ (\j\{l\<^sub>0..0) \ (\j\{i..0}. a j \ p\<^sub>0) \ l\<^sub>0\i \ i\h\<^sub>0 \ a\<^sub>0 = a on -{l\<^sub>0..0} " defines \ i=l; j=l; while (j l\i \ i\j \ j\h \ mset_ran a\<^sub>0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ (\k\{l.. (\k\{i.. p) \ (\k\{j..0 k = a k) \ a\<^sub>0 = a on -{l\<^sub>0..0} \ @variant \(h-j)\ { if (a[j] supply lran_eq_iff[simp] lran_tail[simp del] apply vcg_cs subgoal by (simp add: mset_ran_swap[unfolded swap_def]) subgoal by auto done procedure_spec partition(a,l,h,p) returns (a,i) assumes "l0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ (\j\{l\<^sub>0.. (\j\{i..0}. a j \ a i) \ l\<^sub>0\i \ i0 \ a\<^sub>0 (h\<^sub>0-1) = a i \ a\<^sub>0 = a on -{l\<^sub>0..0} " defines \ p = a[h-1]; (a,i) = partition_aux(a,l,h-1,p); a[h-1] = a[i]; a[i] = p \ apply vcg_cs apply (auto simp: eq_on_def mset_ran_swap[unfolded swap_def] intvs_incdec intro: mset_ran_combine_eq_diff) done lemma quicksort_sorted_aux: assumes BOUNDS: "l \ i" "i < h" assumes LESS: "\j\{l..1 j < a\<^sub>1 i" assumes GEQ: "\j\{i..1 i \ a\<^sub>1 j" assumes R1: "mset_ran a\<^sub>1 {l..2 {l..1 = a\<^sub>2 on - {l..2 l i" assumes R2: "mset_ran a\<^sub>2 {i + 1..3 {i + 1..2 = a\<^sub>3 on - {i + 1..3 (i + 1) h" shows "ran_sorted a\<^sub>3 l h" proof - have [simp]: "{l.. - {i + 1..1 i = a\<^sub>3 i" using E1 E2 by (auto simp: eq_on_def) note X1 = mset_ran_xfer_pointwise[where P = \\x. x < p\ for p, OF R1, simplified] note X2 = eq_on_xfer_pointwise[where P = \\x. x < p\ for p, OF E2, of "{l..j\{l..3 j < a\<^sub>3 i" by (simp add: X1 X2) from GEQ have GEQ1: "\j\{i+1..1 i \ a\<^sub>1 j" by auto have [simp]: "{i + 1.. - {l..\x. x \ p\ for p, OF E1, of "{i+1..\x. x \ p\ for p, OF R2, simplified] from GEQ1 have GEQ': "\j\{i+1..3 i \ a\<^sub>3 j" by (simp add: X3 X4) from SL eq_on_xfer_ran_sorted[OF E2, of l i] have SL': "ran_sorted a\<^sub>3 l i" by simp show ?thesis using combine_sorted_pivot[OF BOUNDS SL' SH LESS' GEQ'] . qed lemma quicksort_mset_aux: assumes B: "l\<^sub>0\i" "i0" assumes R1: "mset_ran a {l\<^sub>0..0..0..0} = mset_ran ab {i + 1..0}" assumes E2: "aa = ab on - {i + 1..0}" shows "mset_ran a {l\<^sub>0..0} = mset_ran ab {l\<^sub>0..0}" apply (rule trans) apply (rule mset_ran_eq_extend[OF R1 E1]) using B apply auto [2] apply (rule mset_ran_eq_extend[OF R2 E2]) using B apply auto [2] done recursive_spec quicksort(a,l,h) returns a assumes "True" ensures "ran_sorted a l\<^sub>0 h\<^sub>0 \ mset_ran a\<^sub>0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ a\<^sub>0=a on -{l\<^sub>0..0}" variant "h-l" defines \ if (l apply (vcg_cs; (intro conjI)?) subgoal using quicksort_sorted_aux by metis subgoal using quicksort_mset_aux by metis subgoal by (smt ComplD ComplI atLeastLessThan_iff eq_on_def) subgoal by (auto simp: ran_sorted_def) done subsection \Data Refinement\ subsubsection \Filtering\ program_spec array_filter_negative assumes "l\h" ensures "lran a l\<^sub>0 i = filter (\x. x\0) (lran a\<^sub>0 l\<^sub>0 h\<^sub>0)" defines \ i=l; j=l; while (j l\i \ i\j \ j\h \ lran a l i = filter (\x. x\0) (lran a\<^sub>0 l j) \ lran a j h = lran a\<^sub>0 j h \ @variant \h-j\ { if (a[j]\0) {a[i] = a[j]; i=i+1}; j=j+1 } \ supply lran_eq_iff[simp] lran_tail[simp del] apply vcg_cs done (* Exercise: Use swap to modify this algorithm to partitioning! *) subsubsection \Merge Two Sorted Lists\ text \ We define the merge function abstractly first, as a functional program on lists. \ fun merge where "merge [] ys = ys" | "merge xs [] = xs" | "merge (x#xs) (y#ys) = (if xIt's straightforward to show that this produces a sorted list with the same elements.\ lemma merge_sorted: assumes "sorted xs" "sorted ys" shows "sorted (merge xs ys) \ set (merge xs ys) = set xs \ set ys" using assms apply (induction xs ys rule: merge.induct) apply auto done lemma merge_mset: "mset (merge xs ys) = mset xs + mset ys" by (induction xs ys rule: merge.induct) auto text \Next, we prove an equation that characterizes one step of the while loop, on the list level.\ lemma merge_eq: "xs\[] \ ys\[] \ merge xs ys = ( if ys=[] \ (xs\[] \ hd xs < hd ys) then hd xs # merge (tl xs) ys else hd ys # merge xs (tl ys) )" by (cases xs; cases ys; simp) text \ We do a first proof that our merge implementation on the arrays and indexes behaves like the functional merge on the corresponding lists. The annotations use the @{const lran} function to map from the implementation level to the list level. Moreover, the invariant of the implementation, \l\h\, is carried through explicitly. \ program_spec merge_imp' assumes "l1\h1 \ l2\h2" ensures "let ms = lran m 0 j; xs\<^sub>0 = lran a1\<^sub>0 l1\<^sub>0 h1\<^sub>0; ys\<^sub>0 = lran a2\<^sub>0 l2\<^sub>0 h2\<^sub>0 in j\0 \ ms = merge xs\<^sub>0 ys\<^sub>0" defines \ j=0; while (l1\h1 \ l2\h2) @variant \h1 + h2 - l1 - l2\ @invariant \let xs=lran a1 l1 h1; ys = lran a2 l2 h2; ms = lran m 0 j; xs\<^sub>0 = lran a1\<^sub>0 l1\<^sub>0 h1\<^sub>0; ys\<^sub>0 = lran a2\<^sub>0 l2\<^sub>0 h2\<^sub>0 in l1\h1 \ l2\h2 \ 0\j \ merge xs\<^sub>0 ys\<^sub>0 = ms@merge xs ys \ { if (l2==h2 \ (l1\h1 \ a1[l1] < a2[l2])) { m[j] = a1[l1]; l1=l1+1 } else { m[j] = a2[l2]; l2=l2+1 }; j=j+1 } \ text \Given the @{thm [source] merge_eq} theorem, which captures the essence of a loop step, and the theorems @{thm lran_append1}, @{thm lran_tail}, and @{thm hd_lran}, which convert from the operations on arrays and indexes to operations on lists, the proof is straightforward \ apply vcg_cs subgoal apply (subst merge_eq) by auto subgoal by linarith subgoal apply (subst merge_eq) by auto done text \ In a next step, we refine our proof to combine it with the abstract properties we have proved about merge. The program does not change (we simply inline the original one here). \ procedure_spec merge_imp (a1,l1,h1,a2,l2,h2) returns (m,j) assumes "l1\h1 \ l2\h2 \ sorted (lran a1 l1 h1) \ sorted (lran a2 l2 h2)" ensures "let ms = lran m 0 j in j\0 \ sorted ms \ mset ms = mset (lran a1\<^sub>0 l1\<^sub>0 h1\<^sub>0) + mset (lran a2\<^sub>0 l2\<^sub>0 h2\<^sub>0)" for l1 h1 l2 h2 a1[] a2[] m[] j defines \inline merge_imp'\ apply vcg_cs apply (auto simp: Let_def merge_mset dest: merge_sorted) done thm merge_imp_spec thm merge_imp_def lemma [named_ss vcg_bb]: "UNIV \ a = UNIV" "a \ UNIV = UNIV" by auto lemma merge_msets_aux: "\l\m; m\h\ \ mset (lran a l m) + mset (lran a m h) = mset (lran a l h)" by (auto simp: mset_lran mset_ran_combine ivl_disj_un_two) recursive_spec mergesort (a,l,h) returns (b,j) assumes "l\h" ensures \0\j \ sorted (lran b 0 j) \ mset (lran b 0 j) = mset (lran a\<^sub>0 l\<^sub>0 h\<^sub>0)\ variant \h-l\ for a[] b[] defines \ if (l==h) j=0 else if (l+1==h) { b[0] = a[l]; j=1 } else { m = (h+l) / 2; (a1,h1) = rec mergesort (a,l,m); (a2,h2) = rec mergesort (a,m,h); (b,j) = merge_imp (a1,0,h1,a2,0,h2) } \ apply vcg apply auto [] apply (auto simp: lran.simps) [] apply auto [] apply auto [] apply auto [] apply (auto simp: Let_def merge_msets_aux) [] done print_theorems subsubsection \Remove Duplicates from Array, using Bitvector Set\ text \We use an array to represent a set of integers. If we only insert elements in range \{0.., this representation is called bit-vector (storing a single bit per index is enough). \ definition set_of :: "(int \ int) \ int set" where "set_of a \ {i. a i \ 0}" context notes [simp] = set_of_def begin lemma set_of_empty[simp]: "set_of (\_. 0) = {}" by auto lemma set_of_insert[simp]: "x\0 \ set_of (a(i:=x)) = insert i (set_of a)" by auto lemma set_of_remove[simp]: "set_of (a(i:=0)) = set_of a - {i}" by auto lemma set_of_mem[simp]: "i\set_of a \ a i \ 0" by auto end program_spec dedup assumes \l\h\ ensures \set (lran a l i) = set (lran a\<^sub>0 l h) \ distinct (lran a l i)\ defines \ i=l; j=l; clear b[]; while (jh-j\ @invariant \l\i \ i\j \ j\h \ set (lran a l i) = set (lran a\<^sub>0 l j) \ distinct (lran a l i) \ lran a j h = lran a\<^sub>0 j h \ set_of b = set (lran a l i) \ { if (b[a[j]] == 0) { a[i] = a[j]; i=i+1; b[a[j]] = 1 }; j=j+1 } \ apply vcg_cs apply (auto simp: lran_eq_iff lran_upd_inside intro: arg_cong[where f=tl]) [] apply (auto simp: lran_eq_iff) [] done procedure_spec bv_init () returns b assumes True ensures \set_of b = {}\ defines \clear b[]\ by vcg_cs procedure_spec bv_insert (x, b) returns b assumes True ensures \set_of b = insert x\<^sub>0 (set_of b\<^sub>0)\ defines \b[x] = 1\ by vcg_cs procedure_spec bv_remove (x, b) returns b assumes True ensures \set_of b = set_of b\<^sub>0 - {x\<^sub>0}\ defines \b[x] = 0\ by vcg_cs procedure_spec bv_elem (x,b) returns r assumes True ensures \r\0 \ x\<^sub>0\set_of b\<^sub>0\ defines \r = b[x]\ by vcg_cs procedure_spec dedup' (a,l,h) returns (a,l,i) assumes \l\h\ ensures \set (lran a l i) = set (lran a\<^sub>0 l\<^sub>0 h\<^sub>0) \ distinct (lran a l i) \ for b[] defines \ b = bv_init(); i=l; j=l; while (jh-j\ @invariant \l\i \ i\j \ j\h \ set (lran a l i) = set (lran a\<^sub>0 l j) \ distinct (lran a l i) \ lran a j h = lran a\<^sub>0 j h \ set_of b = set (lran a l i) \ { mem = bv_elem (a[j],b); if (mem == 0) { a[i] = a[j]; i=i+1; b = bv_insert(a[j],b) }; j=j+1 } \ apply vcg_cs apply (auto simp: lran_eq_iff lran_upd_inside intro: arg_cong[where f=tl]) done subsection \Recursion\ subsubsection \Recursive Fibonacci\ recursive_spec fib_imp (i) returns r assumes True ensures \r = fib i\<^sub>0\ variant \i\ defines \ if (i\0) r=0 else if (i==1) r=1 else { r1=rec fib_imp (i-2); r2=rec fib_imp (i-1); r = r1+r2 } \ by vcg_cs subsubsection \Homeier's Cycling Termination\ text \A contrived example from Homeier's thesis. Only the termination proof is done.\ (* TODO: Also show correctness: pedal (a,b) will multiply its arguments! correctness proof may be a good test how to handle specs with logical vars! *) recursive_spec pedal (n,m) returns () assumes \n\0 \ m\0\ ensures True variant \n+m\ defines \ if (n\0 \ m\0) { G = G + m; if (n and coast (n,m) returns () assumes \n\0 \ m\0\ ensures True variant \n+m+1\ defines \ G = G + n; if (n by vcg_cs subsubsection \Ackermann\ fun ack :: "nat \ nat \ nat" where "ack 0 n = n+1" | "ack m 0 = ack (m-1) 1" | "ack m n = ack (m-1) (ack m (n-1))" lemma ack_add_simps[simp]: "m\0 \ ack m 0 = ack (m-1) 1" "\m\0; n\0\ \ ack m n = ack (m-1) (ack m (n-1))" subgoal by (cases m) auto subgoal by (cases "(m,n)" rule: ack.cases) (auto) done recursive_spec relation "less_than <*lex*> less_than" ack_imp (m,n) returns r assumes "m\0 \ n\0" ensures "r=int (ack (nat m\<^sub>0) (nat n\<^sub>0))" variant "(nat m, nat n)" defines \ if (m==0) r = n+1 else if (n==0) r = rec ack_imp (m-1,1) else { t = rec ack_imp (m,n-1); r = rec ack_imp (m-1,t) } \ supply nat_distribs[simp] by vcg_cs subsubsection \McCarthy's 91 Function\ text \A standard benchmark for verification of recursive functions. We use Homeier's version with a global variable.\ recursive_spec p91(y) assumes True ensures "if 1000 then G = y\<^sub>0-10 else G = 91" variant "101-y" for G defines \ if (100 apply vcg_cs apply (auto split: if_splits) done subsubsection \Odd/Even\ recursive_spec odd_imp (a) returns b assumes True ensures "b\0 \ odd a\<^sub>0" variant "\a\" defines \ if (a==0) b=0 else if (a<0) b = rec even_imp (a+1) else b = rec even_imp (a-1) \ and even_imp (a) returns b assumes True ensures "b\0 \ even a\<^sub>0" variant "\a\" defines \ if (a==0) b=1 else if (a<0) b = rec odd_imp (a+1) else b = rec odd_imp (a-1) \ apply vcg apply auto done thm even_imp_spec subsubsection \Pandya and Joseph's Product Producers\ text \Again, taking the version from Homeier's thesis, but with a modification to also terminate for negative \y\. \ recursive_spec relation \measure nat <*lex*> less_than\ product () assumes True ensures \GZ = GZ\<^sub>0 + GX\<^sub>0*GY\<^sub>0\ variant "(\GY\,1::nat)" for GX GY GZ defines \ e = even_imp (GY); if (e\0) rec evenproduct() else rec oddproduct() \ and oddproduct() assumes \odd GY\ ensures \GZ = GZ\<^sub>0 + GX\<^sub>0*GY\<^sub>0\ variant "(\GY\,0::nat)" for GX GY GZ defines \ if (GY<0) { GY = GY + 1; GZ = GZ - GX } else { GY = GY - 1; GZ = GZ + GX }; rec evenproduct() \ and evenproduct() assumes \even GY\ ensures \GZ = GZ\<^sub>0 + GX\<^sub>0*GY\<^sub>0\ variant "(\GY\,0::nat)" for GX GY GZ defines \ if (GY\0) { GX = 2*GX; GY = GY / 2; rec product() } \ apply vcg_cs apply (auto simp: algebra_simps) apply presburger+ done (* TODO: * Infrastructure to prove different specification for same program (DONE) We can use inline, and just give program a new name. * ghost variables. Keep them existentially qualified, with naming tag. ABS_GHOST ''name'' (\x. P x) \ \x. P x GHOST ''name'' value = True assumes ABS_GHOST name (\x. P x) obtains x where "GHOST name x" "P x" assumes "GHOST name x" and "P x" shows "ABS_GHOST name (\x. P x)" Let_Ghost name (\x s. C x s) = skip assumes "\x. C x s" assumes "\x. \ GHOST name x; C x s \ \ Q s" shows "wp (Let_Ghost name (\x s. C x s)) Q s" Let specification command also abstract over ghost variables in program. *) (* IDEAS: Extend arrays to multiple dimensions: int list \ int Vidx vname "aexp list" AssignIdx vname "aexp list" aexp ArrayCpy vname vname -- copy all dimensions (it's simpler) ArrayInit vname -- Init all dimensions plain values on dimension []! then we can easily encode matrices and adjacency lists! *) subsection \Graph Algorithms\ subsubsection \DFS\ (* By Simon Wimmer *) text \A graph is stored as an array of integers. Each node is an index into this array, pointing to a size-prefixed list of successors. Example for node \i\, which has successors \s1... sn\: \<^verbatim>\ Indexes: ... | i | i+1 | ... | i+n | ... Data: ... | n | s1 | ... | sn | ... \ \ definition succs where "succs a i \ a ` {i+1.. int" definition Edges where "Edges a \ {(i, j). j \ succs a i}" procedure_spec push' (x, stack, ptr) returns (stack, ptr) assumes "ptr \ 0" ensures \lran stack 0 ptr = lran stack\<^sub>0 0 ptr\<^sub>0 @ [x\<^sub>0] \ ptr = ptr\<^sub>0 + 1\ defines \stack[ptr] = x; ptr = ptr + 1\ by vcg_cs procedure_spec push (x, stack, ptr) returns (stack, ptr) assumes "ptr \ 0" ensures \stack ` {0..0} \ stack\<^sub>0 ` {0..0} \ ptr = ptr\<^sub>0 + 1\ for stack[] defines \stack[ptr] = x; ptr = ptr + 1\ by vcg_cs (auto simp: fun_upd_image) program_spec get_succs assumes "j \ stop \ stop = a (j - 1) \ 0 \ i" ensures " stack ` {0..0 - 1, x) \ Edges a \ x \ set_of visited} \ stack\<^sub>0 ` {0..0} \ i \ i\<^sub>0" for i j stop stack[] a[] visited[] defines \ while (j < stop) @invariant \stack ` {0.. a ` {j\<^sub>0.. x \ set_of visited} \ stack\<^sub>0 ` {0..0} \ j \ stop \ i\<^sub>0 \ i \ j\<^sub>0 \ j \ @variant \(stop - j)\ { succ = a[j]; is_elem = bv_elem(succ,visited); if (is_elem == 0) { (stack, i) = push (succ, stack, i) }; j = j + 1 } \ by vcg_cs (auto simp: intvs_incr_h Edges_def succs_def) procedure_spec pop (stack, ptr) returns (x, ptr) assumes "ptr \ 1" ensures \stack\<^sub>0 ` {0..0} = stack\<^sub>0 ` {0.. {x} \ ptr\<^sub>0 = ptr + 1\ for stack[] defines \ptr = ptr - 1; x = stack[ptr]\ by vcg_cs (simp add: intvs_upper_decr) procedure_spec stack_init () returns i assumes True ensures \i = 0\ defines \i = 0\ by vcg_cs lemma Edges_empty: "Edges a `` {i} = {}" if "i + 1 \ a i" using that unfolding Edges_def succs_def by auto text \This is one of the main insights of the algorithm: if a set of visited states is closed w.r.t.\ to the edge relation, then it is guaranteed to contain all the states that are reachable from any state within the set.\ lemma reachability_invariant: assumes reachable: "(s, x) \ (Edges a)\<^sup>*" and closed: "\v\visited. Edges a `` {v} \ visited" and start: "s \ visited" shows "x \ visited" using reachable start closed by induction auto program_spec (partial) dfs assumes "0 \ x \ 0 \ s" ensures "b = 1 \ x \ (Edges a)\<^sup>* `` {s}" defines \ b = 0; clear stack[]; i = stack_init(); (stack, i) = push (s, stack, i); clear visited[]; while (b == 0 \ i \ 0) @invariant \0 \ i \ (s \ stack ` {0.. s \ set_of visited) \ (b = 0 \ b = 1) \ ( if b = 0 then stack ` {0.. (Edges a)\<^sup>* `` {s} \ (\v \ set_of visited. (Edges a) `` {v} \ set_of visited \ stack ` {0.. (x \ set_of visited) else x \ (Edges a)\<^sup>* `` {s}) \ { (next, i) = pop(stack, i); \\Take the top most element from the stack.\ visited = bv_insert(next, visited); \\Mark it as visited,\ if (next == x) { b = 1 \\If it is the target, we are done.\ } else { \\Else, put its successors on the stack if they are not yet visited.\ stop = a[next]; j = next + 1; if (j \ stop) { inline get_succs } } } \ apply vcg_cs subgoal by (auto simp: set_of_def) subgoal using intvs_lower_incr by (auto simp: Edges_empty) subgoal by auto (fastforce simp: set_of_def dest!: reachability_invariant) done text \Assuming that the input graph is finite, we can also prove that the algorithm terminates. We will thus use an \Isabelle context\ to fix a certain finite graph and a start state: \ context fixes start :: int and edges assumes finite_graph[intro!]: "finite ((Edges edges)\<^sup>* `` {start})" begin lemma sub_insert_same_iff: "s \ insert x s \ x\s" by auto program_spec dfs1 assumes "0 \ x \ 0 \ s \ start = s \ edges = a" ensures "b = 1 \ x \ (Edges a)\<^sup>* `` {s}" for visited[] defines \ b = 0; \\\i\ will point to the next free space in the stack (i.e. it is the size of the stack)\ i = 1; \\Initially, we put \s\ on the stack.\ stack[0] = s; visited = bv_init(); while (b == 0 \ i \ 0) @invariant \ 0 \ i \ (s \ stack ` {0.. s \ set_of visited) \ (b = 0 \ b = 1) \ set_of visited \ (Edges edges)\<^sup>* `` {start} \ ( if b = 0 then stack ` {0.. (Edges a)\<^sup>* `` {s} \ (\v \ set_of visited. (Edges a) `` {v} \ set_of visited \ stack ` {0.. (x \ set_of visited) else x \ (Edges a)\<^sup>* `` {s}) \ @relation \finite_psupset ((Edges edges)\<^sup>* `` {start}) <*lex*> less_than\ @variant \ (set_of visited, nat i)\ { \\Take the top most element from the stack.\ (next, i) = pop(stack, i); if (next == x) { \\If it is the target, we are done.\ visited = bv_insert(next, visited); b = 1 } else { is_elem = bv_elem(next, visited); if (is_elem == 0) { visited = bv_insert(next, visited); \\Else, put its successors on the stack if they are not yet visited.\ stop = a[next]; j = next + 1; if (j \ stop) { inline get_succs } } } } \ apply vcg_cs subgoal by auto subgoal by (auto simp add: image_constant_conv) subgoal by (clarsimp simp: finite_psupset_def sub_insert_same_iff) subgoal by (auto simp: set_of_def) subgoal by (clarsimp simp: finite_psupset_def sub_insert_same_iff) subgoal by (auto simp: Edges_empty) subgoal by (clarsimp simp: finite_psupset_def sub_insert_same_iff) subgoal by (auto simp: set_of_def) subgoal by auto (fastforce simp: set_of_def dest!: reachability_invariant) done end end \ No newline at end of file diff --git a/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy b/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy --- a/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy +++ b/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy @@ -1,1056 +1,1057 @@ theory LinuxRouter_OpenFlow_Translation imports IP_Addresses.CIDR_Split Automatic_Refinement.Misc (*TODO@Peter: rename and make available at better place :)*) Simple_Firewall.Generic_SimpleFw Semantics_OpenFlow OpenFlow_Matches OpenFlow_Action Routing.Linux_Router + "Pure-ex.Guess" begin hide_const Misc.uncurry hide_fact Misc.uncurry_def (* For reference: iiface :: "iface" --"in-interface" oiface :: "iface" --"out-interface" src :: "(ipv4addr \ nat) " --"source IP address" dst :: "(ipv4addr \ nat) " --"destination" proto :: "protocol" sports :: "(16 word \ 16 word)" --"source-port first:last" dports :: "(16 word \ 16 word)" --"destination-port first:last" p_iiface :: string p_oiface :: string p_src :: ipv4addr p_dst :: ipv4addr p_proto :: primitive_protocol p_sport :: "16 word" p_dport :: "16 word" p_tcp_flags :: "tcp_flag set" p_payload :: string *) definition "route2match r = \iiface = ifaceAny, oiface = ifaceAny, src = (0,0), dst=(pfxm_prefix (routing_match r),pfxm_length (routing_match r)), proto=ProtoAny, sports=(0,- 1), ports=(0,- 1)\" definition toprefixmatch where "toprefixmatch m \ (let pm = PrefixMatch (fst m) (snd m) in if pm = PrefixMatch 0 0 then None else Some pm)" lemma prefix_match_semantics_simple_match: assumes some: "toprefixmatch m = Some pm" assumes vld: "valid_prefix pm" shows "prefix_match_semantics pm = simple_match_ip m" using some by(cases m) (clarsimp simp add: toprefixmatch_def ipset_from_cidr_def pfxm_mask_def fun_eq_iff prefix_match_semantics_ipset_from_netmask[OF vld] NOT_mask_shifted_lenword[symmetric] split: if_splits) definition simple_match_to_of_match_single :: "(32, 'a) simple_match_scheme \ char list option \ protocol \ (16 word \ 16 word) option \ (16 word \ 16 word) option \ of_match_field set" where "simple_match_to_of_match_single m iif prot sport dport \ uncurry L4Src ` option2set sport \ uncurry L4Dst ` option2set dport \ IPv4Proto ` (case prot of ProtoAny \ {} | Proto p \ {p}) \ \protocol is an 8 word option anyway...\ \ IngressPort ` option2set iif \ IPv4Src ` option2set (toprefixmatch (src m)) \ IPv4Dst ` option2set (toprefixmatch (dst m)) \ {EtherType 0x0800}" (* okay, we need to make sure that no packets are output on the interface they were input on. So for rules that don't have an input interface, we'd need to do a product over all interfaces, if we stay naive. The more smart way would be to insert a rule with the same match condition that additionally matches the input interface and drops. However, I'm afraid this is going to be very tricky to verify\ *) definition simple_match_to_of_match :: "32 simple_match \ string list \ of_match_field set list" where "simple_match_to_of_match m ifs \ (let npm = (\p. fst p = 0 \ snd p = - 1); sb = (\p. (if npm p then [None] else if fst p \ snd p then map (Some \ (\pfx. (pfxm_prefix pfx, NOT (pfxm_mask pfx)))) (wordinterval_CIDR_split_prefixmatch (WordInterval (fst p) (snd p))) else [])) in [simple_match_to_of_match_single m iif (proto m) sport dport. iif \ (if iiface m = ifaceAny then [None] else [Some i. i \ ifs, match_iface (iiface m) i]), sport \ sb (sports m), dport \ sb (dports m)] )" (* I wonder\ should I check whether list_all (match_iface (iiface m)) ifs instead of iiface m = ifaceAny? It would be pretty stupid if that wasn't the same, but you know\ *) lemma smtoms_eq_hlp: "simple_match_to_of_match_single r a b c d = simple_match_to_of_match_single r f g h i \ (a = f \ b = g \ c = h \ d = i)" (* In case this proof breaks: there are two alternate proofs in the repo. They are of similar quality, though. Good luck. *) proof(rule iffI,goal_cases) case 1 thus ?case proof(intro conjI) have *: "\P z x. \\x :: of_match_field. P x; z = Some x\ \ P (IngressPort x)" by simp show "a = f" using 1 by(cases a; cases f) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ next have *: "\P z x. \\x :: of_match_field. P x; z = Proto x\ \ P (IPv4Proto x)" by simp show "b = g" using 1 by(cases b; cases g) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ next have *: "\P z x. \\x :: of_match_field. P x; z = Some x\ \ P (uncurry L4Src x)" by simp show "c = h" using 1 by(cases c; cases h) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ next have *: "\P z x. \\x :: of_match_field. P x; z = Some x\ \ P (uncurry L4Dst x)" by simp show "d = i" using 1 by(cases d; cases i) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ qed qed simp lemma simple_match_to_of_match_generates_prereqs: "simple_match_valid m \ r \ set (simple_match_to_of_match m ifs) \ all_prerequisites r" unfolding simple_match_to_of_match_def Let_def proof(clarsimp, goal_cases) case (1 xiface xsrcp xdstp) note o = this show ?case unfolding simple_match_to_of_match_single_def all_prerequisites_def unfolding ball_Un proof((intro conjI; ((simp;fail)| - )), goal_cases) case 1 have e: "(fst (sports m) = 0 \ snd (sports m) = - 1) \ proto m = Proto TCP \ proto m = Proto UDP \ proto m = Proto L4_Protocol.SCTP" using o(1) unfolding simple_match_valid_alt Let_def by(clarsimp split: if_splits) show ?case using o(3) e by(elim disjE; simp add: option2set_def split: if_splits prod.splits uncurry_splits) next case 2 have e: "(fst (dports m) = 0 \ snd (dports m) = - 1) \ proto m = Proto TCP \ proto m = Proto UDP \ proto m = Proto L4_Protocol.SCTP" using o(1) unfolding simple_match_valid_alt Let_def by(clarsimp split: if_splits) show ?case using o(4) e by(elim disjE; simp add: option2set_def split: if_splits prod.splits uncurry_splits) qed qed lemma and_assoc: "a \ b \ c \ (a \ b) \ c" by simp lemmas custom_simpset = Let_def set_concat set_map map_map comp_def concat_map_maps set_maps UN_iff fun_app_def Set.image_iff abbreviation "simple_fw_prefix_to_wordinterval \ prefix_to_wordinterval \ uncurry PrefixMatch" lemma simple_match_port_alt: "simple_match_port m p \ p \ wordinterval_to_set (uncurry WordInterval m)" by(simp split: uncurry_splits) lemma simple_match_src_alt: "simple_match_valid r \ simple_match_ip (src r) p \ prefix_match_semantics (PrefixMatch (fst (src r)) (snd (src r))) p" by(cases "(src r)") (simp add: prefix_match_semantics_ipset_from_netmask2 prefix_to_wordset_ipset_from_cidr simple_match_valid_def valid_prefix_fw_def) lemma simple_match_dst_alt: "simple_match_valid r \ simple_match_ip (dst r) p \ prefix_match_semantics (PrefixMatch (fst (dst r)) (snd (dst r))) p" by(cases "(dst r)") (simp add: prefix_match_semantics_ipset_from_netmask2 prefix_to_wordset_ipset_from_cidr simple_match_valid_def valid_prefix_fw_def) lemma "x \ set (wordinterval_CIDR_split_prefixmatch w) \ valid_prefix x" using wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1] . lemma simple_match_to_of_matchI: assumes mv: "simple_match_valid r" assumes mm: "simple_matches r p" assumes ii: "p_iiface p \ set ifs" assumes ippkt: "p_l2type p = 0x800" shows eq: "\gr \ set (simple_match_to_of_match r ifs). OF_match_fields gr p = Some True" proof - let ?npm = "\p. fst p = 0 \ snd p = - 1" let ?sb = "\p r. (if ?npm p then None else Some r)" obtain si where si: "case si of Some ssi \ p_sport p \ prefix_to_wordset ssi | None \ True" "case si of None \ True | Some ssi \ ssi \ set ( wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r)))" "si = None \ ?npm (sports r)" proof(cases "?npm (sports r)", goal_cases) case 1 (* True *) hence "(case None of None \ True | Some ssi \ p_sport p \ prefix_to_wordset ssi) \ (case None of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r))))" by simp with 1 show ?thesis by blast next case 2 (* False *) from mm have "p_sport p \ wordinterval_to_set (uncurry WordInterval (sports r))" by(simp only: simple_matches.simps simple_match_port_alt) then obtain ssi where ssi: "ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r)))" "p_sport p \ prefix_to_wordset ssi" using wordinterval_CIDR_split_existential by fast hence "(case Some ssi of None \ True | Some ssi \ p_sport p \ prefix_to_wordset ssi) \ (case Some ssi of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r))))" by simp with 2 show ?thesis by blast qed obtain di where di: "case di of Some ddi \ p_dport p \ prefix_to_wordset ddi | None \ True" "case di of None \ True | Some ddi \ ddi \ set ( wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r)))" "di = None \ ?npm (dports r)" proof(cases "?npm (dports r)", goal_cases) case 1 hence "(case None of None \ True | Some ssi \ p_dport p \ prefix_to_wordset ssi) \ (case None of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r))))" by simp with 1 show ?thesis by blast next case 2 from mm have "p_dport p \ wordinterval_to_set (uncurry WordInterval (dports r))" by(simp only: simple_matches.simps simple_match_port_alt) then obtain ddi where ddi: "ddi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r)))" "p_dport p \ prefix_to_wordset ddi" using wordinterval_CIDR_split_existential by fast hence "(case Some ddi of None \ True | Some ssi \ p_dport p \ prefix_to_wordset ssi) \ (case Some ddi of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r))))" by simp with 2 show ?thesis by blast qed show ?thesis proof let ?mf = "map_option (apsnd (wordNOT \ mask \ (-) 16) \ prefix_match_dtor)" let ?gr = "simple_match_to_of_match_single r (if iiface r = ifaceAny then None else Some (p_iiface p)) (if proto r = ProtoAny then ProtoAny else Proto (p_proto p)) (?mf si) (?mf di)" note mfu = simple_match_port.simps[of "fst (sports r)" "snd (sports r)", unfolded surjective_pairing[of "sports r",symmetric]] simple_match_port.simps[of "fst (dports r)" "snd (dports r)", unfolded surjective_pairing[of "dports r",symmetric]] note u = mm[unfolded simple_matches.simps mfu ord_class.atLeastAtMost_iff simple_packet_unext_def simple_packet.simps] note of_safe_unsafe_match_eq[OF simple_match_to_of_match_generates_prereqs] from u have ple: "fst (sports r) \ snd (sports r)" "fst (dports r) \ snd (dports r)" by force+ show eg: "?gr \ set (simple_match_to_of_match r ifs)" unfolding simple_match_to_of_match_def unfolding custom_simpset unfolding smtoms_eq_hlp proof(intro bexI, (intro conjI; ((rule refl)?)), goal_cases) case 2 thus ?case using ple(2) di apply(simp add: pfxm_mask_def prefix_match_dtor_def Set.image_iff split: option.splits prod.splits uncurry_splits) apply(erule bexI[rotated]) apply(simp split: prefix_match.splits) done next case 3 thus ?case using ple(1) si apply(simp add: pfxm_mask_def prefix_match_dtor_def Set.image_iff split: option.splits prod.splits uncurry_splits) apply(erule bexI[rotated]) apply(simp split: prefix_match.splits) done next case 4 thus ?case using u ii by(clarsimp simp: set_maps split: if_splits) next case 1 thus ?case using ii u by simp_all (metis match_proto.elims(2)) qed have dpm: "di = Some (PrefixMatch x1 x2) \ p_dport p && ~~ (mask (16 - x2)) = x1" for x1 x2 proof - have *: "di = Some (PrefixMatch x1 x2) \ prefix_match_semantics (the di) (p_dport p) \ p_dport p && ~~ (mask (16 - x2)) = x1" by(clarsimp simp: prefix_match_semantics_def pfxm_mask_def word_bw_comms;fail) have **: "pfx \ set (wordinterval_CIDR_split_prefixmatch ra) \ prefix_match_semantics pfx a = (a \ prefix_to_wordset pfx)" for pfx ra and a :: "16 word" by (fact prefix_match_semantics_wordset[OF wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1]]) have "\di = Some (PrefixMatch x1 x2); p_dport p \ prefix_to_wordset (PrefixMatch x1 x2); PrefixMatch x1 x2 \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r)))\ \ p_dport p && ~~ (mask (16 - x2)) = x1" using di(1,2) using * ** by auto thus "di = Some (PrefixMatch x1 x2) \ p_dport p && ~~ (mask (16 - x2)) = x1" using di(1,2) by auto qed have spm: "si = Some (PrefixMatch x1 x2) \ p_sport p && ~~ (mask (16 - x2)) = x1" for x1 x2 using si proof - have *: "si = Some (PrefixMatch x1 x2) \ prefix_match_semantics (the si) (p_sport p) \ p_sport p && ~~ (mask (16 - x2)) = x1" by(clarsimp simp: prefix_match_semantics_def pfxm_mask_def word_bw_comms;fail) have **: "pfx \ set (wordinterval_CIDR_split_prefixmatch ra) \ prefix_match_semantics pfx a = (a \ prefix_to_wordset pfx)" for pfx ra and a :: "16 word" by (fact prefix_match_semantics_wordset[OF wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1]]) have "\si = Some (PrefixMatch x1 x2); p_sport p \ prefix_to_wordset (PrefixMatch x1 x2); PrefixMatch x1 x2 \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r)))\ \ p_sport p && ~~ (mask (16 - x2)) = x1" using si(1,2) using * ** by auto thus "si = Some (PrefixMatch x1 x2) \ p_sport p && ~~ (mask (16 - x2)) = x1" using si(1,2) by auto qed show "OF_match_fields ?gr p = Some True" unfolding of_safe_unsafe_match_eq[OF simple_match_to_of_match_generates_prereqs[OF mv eg]] by(cases si; cases di) (simp_all add: simple_match_to_of_match_single_def OF_match_fields_unsafe_def spm option2set_def u ippkt prefix_match_dtor_def toprefixmatch_def dpm simple_match_dst_alt[OF mv, symmetric] simple_match_src_alt[OF mv, symmetric] split: prefix_match.splits) qed qed lemma prefix_match_00[simp,intro!]: "prefix_match_semantics (PrefixMatch 0 0) p" by (simp add: valid_prefix_def zero_prefix_match_all) lemma simple_match_to_of_matchD: assumes eg: "gr \ set (simple_match_to_of_match r ifs)" assumes mo: "OF_match_fields gr p = Some True" assumes me: "match_iface (oiface r) (p_oiface p)" assumes mv: "simple_match_valid r" shows "simple_matches r p" proof - from mv have validpfx: "valid_prefix (uncurry PrefixMatch (src r))" "valid_prefix (uncurry PrefixMatch (dst r))" "\pm. toprefixmatch (src r) = Some pm \ valid_prefix pm" "\pm. toprefixmatch (dst r) = Some pm \ valid_prefix pm" unfolding simple_match_valid_def valid_prefix_fw_def toprefixmatch_def by(simp_all split: uncurry_splits if_splits) from mo have mo: "OF_match_fields_unsafe gr p" unfolding of_safe_unsafe_match_eq[OF simple_match_to_of_match_generates_prereqs[OF mv eg]] by simp note this[unfolded OF_match_fields_unsafe_def] note eg[unfolded simple_match_to_of_match_def simple_match_to_of_match_single_def custom_simpset option2set_def] then guess x .. moreover from this(2) guess xa .. moreover from this(2) guess xb .. note xx = calculation(1,3) this { fix a b xc xa fix pp :: "16 word" have "\pp && ~~ (pfxm_mask xc) = pfxm_prefix xc\ \ prefix_match_semantics xc (pp)" for xc by(simp add: prefix_match_semantics_def word_bw_comms;fail) moreover have "pp \ wordinterval_to_set (WordInterval a b) \ a \ pp \ pp \ b" by simp moreover have "xc \ set (wordinterval_CIDR_split_prefixmatch (WordInterval a b)) \ pp \ prefix_to_wordset xc \ pp \ wordinterval_to_set (WordInterval a b)" by(subst wordinterval_CIDR_split_prefixmatch) blast moreover have "\xc \ set (wordinterval_CIDR_split_prefixmatch (WordInterval a b)); xa = Some (pfxm_prefix xc, ~~ (pfxm_mask xc)); prefix_match_semantics xc (pp)\ \ pp \ prefix_to_wordset xc" apply(subst(asm)(1) prefix_match_semantics_wordset) apply(erule wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1];fail) apply assumption done ultimately have "\xc \ set (wordinterval_CIDR_split_prefixmatch (WordInterval a b)); xa = Some (pfxm_prefix xc, ~~ (pfxm_mask xc)); pp && ~~ (pfxm_mask xc) = pfxm_prefix xc\ \ a \ pp \ pp \ b" by metis } note l4port_logic = this show ?thesis unfolding simple_matches.simps proof(unfold and_assoc, (rule)+) show "match_iface (iiface r) (p_iiface p)" apply(cases "iiface r = ifaceAny") apply (simp add: match_ifaceAny) using xx(1) mo unfolding xx(4) OF_match_fields_unsafe_def apply(simp only: if_False set_maps UN_iff) apply(clarify) apply(rename_tac a; subgoal_tac "match_iface (iiface r) a") apply(clarsimp simp add: option2set_def;fail) apply(rule ccontr,simp;fail) done next show "match_iface (oiface r) (p_oiface p)" using me . next show "simple_match_ip (src r) (p_src p)" using mo unfolding xx(4) OF_match_fields_unsafe_def toprefixmatch_def by(clarsimp simp add: simple_packet_unext_def option2set_def validpfx simple_match_src_alt[OF mv] toprefixmatch_def split: if_splits) next show "simple_match_ip (dst r) (p_dst p)" using mo unfolding xx(4) OF_match_fields_unsafe_def toprefixmatch_def by(clarsimp simp add: simple_packet_unext_def option2set_def validpfx simple_match_dst_alt[OF mv] toprefixmatch_def split: if_splits) next show "match_proto (proto r) (p_proto p)" using mo unfolding xx(4) OF_match_fields_unsafe_def using xx(1) by(clarsimp simp add: singleton_iff simple_packet_unext_def option2set_def prefix_match_semantics_simple_match ball_Un split: if_splits protocol.splits) next show "simple_match_port (sports r) (p_sport p)" using mo xx(2) unfolding xx(4) OF_match_fields_unsafe_def by(cases "sports r") (clarsimp simp add: l4port_logic simple_packet_unext_def option2set_def prefix_match_semantics_simple_match split: if_splits) next show "simple_match_port (dports r) (p_dport p)" using mo xx(3) unfolding xx(4) OF_match_fields_unsafe_def by(cases "dports r") (clarsimp simp add: l4port_logic simple_packet_unext_def option2set_def prefix_match_semantics_simple_match split: if_splits) qed qed primrec annotate_rlen where "annotate_rlen [] = []" | "annotate_rlen (a#as) = (length as, a) # annotate_rlen as" lemma "annotate_rlen ''asdf'' = [(3, CHR ''a''), (2, CHR ''s''), (1, CHR ''d''), (0, CHR ''f'')]" by simp lemma fst_annotate_rlen_le: "(k, a) \ set (annotate_rlen l) \ k < length l" by(induction l arbitrary: k; simp; force) lemma distinct_fst_annotate_rlen: "distinct (map fst (annotate_rlen l))" using fst_annotate_rlen_le by(induction l) (simp, fastforce) lemma distinct_annotate_rlen: "distinct (annotate_rlen l)" using distinct_fst_annotate_rlen unfolding distinct_map by blast lemma in_annotate_rlen: "(a,x) \ set (annotate_rlen l) \ x \ set l" by(induction l) (simp_all, blast) lemma map_snd_annotate_rlen: "map snd (annotate_rlen l) = l" by(induction l) simp_all lemma "sorted_descending (map fst (annotate_rlen l))" by(induction l; clarsimp) (force dest: fst_annotate_rlen_le) lemma "annotate_rlen l = zip (rev [0.. *) primrec annotate_rlen_code where "annotate_rlen_code [] = (0,[])" | "annotate_rlen_code (a#as) = (case annotate_rlen_code as of (r,aas) \ (Suc r, (r, a) # aas))" lemma annotate_rlen_len: "fst (annotate_rlen_code r) = length r" by(induction r) (clarsimp split: prod.splits)+ lemma annotate_rlen_code[code]: "annotate_rlen s = snd (annotate_rlen_code s)" proof(induction s) case (Cons s ss) thus ?case using annotate_rlen_len[of ss] by(clarsimp split: prod.split) qed simp lemma suc2plus_inj_on: "inj_on (of_nat :: nat \ ('l :: len) word) {0..unat (max_word :: 'l word)}" proof(rule inj_onI) let ?mmw = "(max_word :: 'l word)" let ?mstp = "(of_nat :: nat \ 'l word)" fix x y :: nat assume "x \ {0..unat ?mmw}" "y \ {0..unat ?mmw}" hence se: "x \ unat ?mmw" "y \ unat ?mmw" by simp_all assume eq: "?mstp x = ?mstp y" note f = le_unat_uoi[OF se(1)] le_unat_uoi[OF se(2)] show "x = y" using eq le_unat_uoi se by metis qed lemma distinct_of_nat_list: (* TODO: Move to CaesarWordLemmaBucket *) "distinct l \ \e \ set l. e \ unat (max_word :: ('l::len) word) \ distinct (map (of_nat :: nat \ 'l word) l)" proof(induction l) let ?mmw = "(max_word :: 'l word)" let ?mstp = "(of_nat :: nat \ 'l word)" case (Cons a as) have "distinct as" "\e\set as. e \ unat ?mmw" using Cons.prems by simp_all note mIH = Cons.IH[OF this] moreover have "?mstp a \ ?mstp ` set as" proof have representable_set: "set as \ {0..unat ?mmw}" using \\e\set (a # as). e \ unat max_word\ by fastforce have a_reprbl: "a \ {0..unat ?mmw}" using \\e\set (a # as). e \ unat max_word\ by simp assume "?mstp a \ ?mstp ` set as" with inj_on_image_mem_iff[OF suc2plus_inj_on a_reprbl representable_set] have "a \ set as" by simp with \distinct (a # as)\ show False by simp qed ultimately show ?case by simp qed simp lemma annotate_first_le_hlp: "length l < unat (max_word :: ('l :: len) word) \ \e\set (map fst (annotate_rlen l)). e \ unat (max_word :: 'l word)" by(clarsimp) (meson fst_annotate_rlen_le less_trans nat_less_le) lemmas distinct_of_prio_hlp = distinct_of_nat_list[OF distinct_fst_annotate_rlen annotate_first_le_hlp] (* don't need these right now, but maybe later? *) lemma fst_annotate_rlen: "map fst (annotate_rlen l) = rev [0.. (of_nat :: nat \ ('l :: len) word)" assumes "length l \ unat (max_word :: 'l word)" shows "sorted_descending (map won (rev [0.. unat (max_word :: ('l :: len) word)" shows "sorted_descending (map fst (map (apfst (of_nat :: nat \ 'l word)) (annotate_rlen l)))" proof - let ?won = "(of_nat :: nat \ 'l word)" have "sorted_descending (map ?won (rev [0..l3 device to l2 forwarding\ definition "lr_of_tran_s3 ifs ard = ( [(p, b, case a of simple_action.Accept \ [Forward c] | simple_action.Drop \ []). (p,r,(c,a)) \ ard, b \ simple_match_to_of_match r ifs])" definition "oif_ne_iif_p1 ifs \ [(simple_match_any\oiface := Iface oif, iiface := Iface iif\, simple_action.Accept). oif \ ifs, iif \ ifs, oif \ iif]" definition "oif_ne_iif_p2 ifs = [(simple_match_any\oiface := Iface i, iiface := Iface i\, simple_action.Drop). i \ ifs]" definition "oif_ne_iif ifs = oif_ne_iif_p2 ifs @ oif_ne_iif_p1 ifs" (* order irrelephant *) (*value "oif_ne_iif [''a'', ''b'']"*) (* I first tried something like "oif_ne_iif ifs \ [(simple_match_any\oiface := Iface oi, iiface := Iface ii\, if oi = ii then simple_action.Drop else simple_action.Accept). oi \ ifs, ii \ ifs]", but making the statement I wanted with that was really tricky. Much easier to have the second element constant and do it separately. *) definition "lr_of_tran_s4 ard ifs \ generalized_fw_join ard (oif_ne_iif ifs)" definition "lr_of_tran_s1 rt = [(route2match r, output_iface (routing_action r)). r \ rt]" definition "lr_of_tran_fbs rt fw ifs \ let gfw = map simple_rule_dtor fw; \ \generalized simple fw, hopefully for FORWARD\ frt = lr_of_tran_s1 rt; \ \rt as fw\ prd = generalized_fw_join frt gfw in prd " definition "pack_OF_entries ifs ard \ (map (split3 OFEntry) (lr_of_tran_s3 ifs ard))" definition "no_oif_match \ list_all (\m. oiface (match_sel m) = ifaceAny)" definition "lr_of_tran rt fw ifs \ if \ (no_oif_match fw \ has_default_policy fw \ simple_fw_valid fw \ valid_prefixes rt \ has_default_route rt \ distinct ifs) then Inl ''Error in creating OpenFlow table: prerequisites not satisifed'' else ( let nrd = lr_of_tran_fbs rt fw ifs; ard = map (apfst of_nat) (annotate_rlen nrd) \ \give them a priority\ in if length nrd < unat (- 1 :: 16 word) then Inr (pack_OF_entries ifs ard) else Inl ''Error in creating OpenFlow table: priority number space exhausted'') " definition "is_iface_name i \ i \ [] \ \Iface.iface_name_is_wildcard i" definition "is_iface_list ifs \ distinct ifs \ list_all is_iface_name ifs" lemma max_16_word_max[simp]: "(a :: 16 word) \ 0xffff" proof - have "0xFFFF = (- 1 :: 16 word)" by simp then show ?thesis by (simp only:) simp qed lemma replicate_FT_hlp: "x \ 16 \ y \ 16 \ replicate (16 - x) False @ replicate x True = replicate (16 - y) False @ replicate y True \ x = y" proof - let ?ns = "{0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15,16}" assume "x \ 16 \ y \ 16" hence "x \ ?ns" "y \ ?ns" by(simp; presburger)+ moreover assume "replicate (16 - x) False @ replicate x True = replicate (16 - y) False @ replicate y True" ultimately show "x = y" by simp (elim disjE; simp_all add: numeral_eq_Suc) (* that's only 289 subgoals after the elim *) qed lemma mask_inj_hlp1: "inj_on (mask :: nat \ 16 word) {0..16}" proof(intro inj_onI, goal_cases) case (1 x y) from 1(3) have oe: "of_bl (replicate (16 - x) False @ replicate x True) = (of_bl (replicate (16 - y) False @ replicate y True) :: 16 word)" unfolding mask_bl of_bl_rep_False . have "\z. z \ 16 \ length (replicate (16 - z) False @ replicate z True) = 16" by auto with 1(1,2) have ps: "replicate (16 - x) False @ replicate x True \ {bl. length bl = LENGTH(16)}" " replicate (16 - y) False @ replicate y True \ {bl. length bl = LENGTH(16)}" by simp_all from inj_onD[OF word_bl.Abs_inj_on, OF oe ps] show ?case using 1(1,2) by(fastforce intro: replicate_FT_hlp) qed lemma distinct_simple_match_to_of_match_portlist_hlp: fixes ps :: "(16 word \ 16 word)" shows "distinct ifs \ distinct (if fst ps = 0 \ snd ps = max_word then [None] else if fst ps \ snd ps then map (Some \ (\pfx. (pfxm_prefix pfx, ~~ (pfxm_mask pfx)))) (wordinterval_CIDR_split_prefixmatch (WordInterval (fst ps) (snd ps))) else [])" proof - assume di: "distinct ifs" { define wis where "wis = set (wordinterval_CIDR_split_prefixmatch (WordInterval (fst ps) (snd ps)))" fix x y :: "16 prefix_match" obtain xm xn ym yn where xyd[simp]: "x = PrefixMatch xm xn" "y = PrefixMatch ym yn" by(cases x; cases y) assume iw: "x \ wis" "y \ wis" and et: "(pfxm_prefix x, ~~ (pfxm_mask x)) = (pfxm_prefix y, ~~ (pfxm_mask y))" hence le16: "xn \ 16" "yn \ 16" unfolding wis_def using wordinterval_CIDR_split_prefixmatch_all_valid_Ball[unfolded Ball_def, THEN spec, THEN mp] by force+ with et have "16 - xn = 16 - yn" unfolding pfxm_mask_def by(auto intro: mask_inj_hlp1[THEN inj_onD]) hence "x = y" using et le16 using diff_diff_cancel by simp } note * = this show ?thesis apply(clarsimp simp add: smtoms_eq_hlp distinct_map wordinterval_CIDR_split_distinct) apply(subst comp_inj_on_iff[symmetric]; intro inj_onI) using * by simp_all qed lemma distinct_simple_match_to_of_match: "distinct ifs \ distinct (simple_match_to_of_match m ifs)" apply(unfold simple_match_to_of_match_def Let_def) apply(rule distinct_3lcomprI) subgoal by(induction ifs; clarsimp) subgoal by(fact distinct_simple_match_to_of_match_portlist_hlp) subgoal by(fact distinct_simple_match_to_of_match_portlist_hlp) subgoal by(simp_all add: smtoms_eq_hlp) done lemma inj_inj_on: "inj F \ inj_on F A" using subset_inj_on by auto (* TODO: include Word_Lib *) lemma no_overlaps_lroft_hlp2: "distinct (map fst amr) \ (\r. distinct (fm r)) \ distinct (concat (map (\(p, r, c, a). map (\b. (p, b, fs a c)) (fm r)) amr))" by(induction amr; force intro: injI inj_onI simp add: distinct_map split: prod.splits) lemma distinct_lroft_s3: "\distinct (map fst amr); distinct ifs\ \ distinct (lr_of_tran_s3 ifs amr)" unfolding lr_of_tran_s3_def by(erule no_overlaps_lroft_hlp2, simp add: distinct_simple_match_to_of_match) lemma no_overlaps_lroft_hlp3: "distinct (map fst amr) \ (aa, ab, ac) \ set (lr_of_tran_s3 ifs amr) \ (ba, bb, bc) \ set (lr_of_tran_s3 ifs amr) \ ac \ bc \ aa \ ba" apply(unfold lr_of_tran_s3_def) apply(clarsimp) apply(clarsimp split: simple_action.splits) apply(metis map_of_eq_Some_iff old.prod.inject option.inject) apply(metis map_of_eq_Some_iff old.prod.inject option.inject simple_action.distinct(2))+ done lemma no_overlaps_lroft_s3_hlp_hlp: (* I hlps *) "\distinct (map fst amr); OF_match_fields_unsafe ab p; ab \ ad \ ba \ bb; OF_match_fields_unsafe ad p; (ac, ab, ba) \ set (lr_of_tran_s3 ifs amr); (ac, ad, bb) \ set (lr_of_tran_s3 ifs amr)\ \ False" proof(elim disjE, goal_cases) case 1 have 4: "\distinct (map fst amr); (ac, ab, x1, x2) \ set amr; (ac, bb, x4, x5) \ set amr; ab \ bb\ \ False" for ab x1 x2 bb x4 x5 by (meson distinct_map_fstD old.prod.inject) have conjunctSomeProtoAnyD: "Some ProtoAny = simple_proto_conjunct a (Proto b) \ False" for a b using conjunctProtoD by force have 5: "\OF_match_fields_unsafe am p; OF_match_fields_unsafe bm p; am \ bm; am \ set (simple_match_to_of_match ab ifs); bm \ set (simple_match_to_of_match bb ifs); \ ab \ bb\ \ False" for ab bb am bm by(clarify | unfold simple_match_to_of_match_def smtoms_eq_hlp Let_def set_concat set_map de_Morgan_conj not_False_eq_True)+ (auto dest: conjunctSomeProtoAnyD cidrsplit_no_overlaps simp add: OF_match_fields_unsafe_def simple_match_to_of_match_single_def option2set_def comp_def split: if_splits cong: smtoms_eq_hlp) (*1min*) from 1 show ?case using 4 5 by(clarsimp simp add: lr_of_tran_s3_def) blast qed(metis no_overlaps_lroft_hlp3) lemma no_overlaps_lroft_s3_hlp: "distinct (map fst amr) \ distinct ifs \ no_overlaps OF_match_fields_unsafe (map (split3 OFEntry) (lr_of_tran_s3 ifs amr))" apply(rule no_overlapsI[rotated]) apply(subst distinct_map, rule conjI) subgoal by(erule (1) distinct_lroft_s3) subgoal apply(rule inj_inj_on) apply(rule injI) apply(rename_tac x y, case_tac x, case_tac y) apply(simp add: split3_def;fail) done subgoal apply(unfold check_no_overlap_def) apply(clarify) apply(unfold set_map) apply(clarify) apply(unfold split3_def prod.simps flow_entry_match.simps flow_entry_match.sel de_Morgan_conj) apply(clarsimp simp only:) apply(erule (1) no_overlaps_lroft_s3_hlp_hlp) apply simp apply assumption apply assumption apply simp done done lemma lr_of_tran_no_overlaps: assumes "distinct ifs" shows "Inr t = (lr_of_tran rt fw ifs) \ no_overlaps OF_match_fields_unsafe t" apply(unfold lr_of_tran_def Let_def pack_OF_entries_def) apply(simp split: if_splits) apply(thin_tac "t = _") apply(drule distinct_of_prio_hlp) apply(rule no_overlaps_lroft_s3_hlp[rotated]) subgoal by(simp add: assms) subgoal by(simp add: o_assoc) done lemma sorted_lr_of_tran_s3_hlp: "\x\set f. fst x \ a \ b \ set (lr_of_tran_s3 s f) \ fst b \ a" by(auto simp add: lr_of_tran_s3_def) lemma lr_of_tran_s3_Cons: "lr_of_tran_s3 ifs (a#ard) = ( [(p, b, case a of simple_action.Accept \ [Forward c] | simple_action.Drop \ []). (p,r,(c,a)) \ [a], b \ simple_match_to_of_match r ifs]) @ lr_of_tran_s3 ifs ard" by(clarsimp simp: lr_of_tran_s3_def) lemma sorted_lr_of_tran_s3: "sorted_descending (map fst f) \ sorted_descending (map fst (lr_of_tran_s3 s f))" apply(induction f) subgoal by(simp add: lr_of_tran_s3_def) apply(clarsimp simp: lr_of_tran_s3_Cons map_concat comp_def) apply(unfold sorted_descending_append) apply(simp add: sorted_descending_alt rev_map sorted_lr_of_tran_s3_hlp sorted_const) done lemma sorted_lr_of_tran_hlp: "(ofe_prio \ split3 OFEntry) = fst" by(simp add: fun_eq_iff comp_def split3_def) lemma lr_of_tran_sorted_descending: "Inr r = lr_of_tran rt fw ifs \ sorted_descending (map ofe_prio r)" apply(unfold lr_of_tran_def Let_def) apply(simp split: if_splits) apply(thin_tac "r = _") apply(unfold sorted_lr_of_tran_hlp pack_OF_entries_def split3_def[abs_def] fun_app_def map_map comp_def prod.case_distrib) apply(simp add: fst_def[symmetric]) apply(rule sorted_lr_of_tran_s3) apply(drule sorted_annotated[OF less_or_eq_imp_le, OF disjI1]) apply(simp add: o_assoc) done lemma lr_of_tran_s1_split: "lr_of_tran_s1 (a # rt) = (route2match a, output_iface (routing_action a)) # lr_of_tran_s1 rt" by(unfold lr_of_tran_s1_def list.map, rule) lemma route2match_correct: "valid_prefix (routing_match a) \ prefix_match_semantics (routing_match a) (p_dst p) \ simple_matches (route2match a) (p)" by(simp add: route2match_def simple_matches.simps match_ifaceAny match_iface_refl ipset_from_cidr_0 prefix_match_semantics_ipset_from_netmask2) lemma s1_correct: "valid_prefixes rt \ has_default_route (rt::('i::len) prefix_routing) \ \rm ra. generalized_sfw (lr_of_tran_s1 rt) p = Some (rm,ra) \ ra = output_iface (routing_table_semantics rt (p_dst p))" apply(induction rt) apply(simp;fail) apply(drule valid_prefixes_split) apply(clarsimp) apply(erule disjE) subgoal for a rt apply(case_tac a) apply(rename_tac routing_m metric routing_action) apply(case_tac routing_m) apply(simp add: valid_prefix_def pfxm_mask_def prefix_match_semantics_def generalized_sfw_def lr_of_tran_s1_def route2match_def simple_matches.simps match_ifaceAny match_iface_refl ipset_from_cidr_0 max_word_mask[where 'a = 'i, symmetric, simplified]) done subgoal apply(rule conjI) apply(simp add: generalized_sfw_def lr_of_tran_s1_def route2match_correct;fail) apply(simp add: route2match_def simple_matches.simps prefix_match_semantics_ipset_from_netmask2 lr_of_tran_s1_split generalized_sfw_simps) done done definition "to_OF_action a \ (case a of (p,d) \ (case d of simple_action.Accept \ [Forward p] | simple_action.Drop \ []))" definition "from_OF_action a = (case a of [] \ ('''',simple_action.Drop) | [Forward p] \ (p, simple_action.Accept))" lemma OF_match_linear_not_noD: "OF_match_linear \ oms p \ NoAction \ \ome. ome \ set oms \ \ (ofe_fields ome) p" apply(induction oms) apply(simp) apply(simp split: if_splits) apply blast+ done lemma s3_noaction_hlp: "\simple_match_valid ac; \simple_matches ac p; match_iface (oiface ac) (p_oiface p)\ \ OF_match_linear OF_match_fields_safe (map (\x. split3 OFEntry (x1, x, case ba of simple_action.Accept \ [Forward ad] | simple_action.Drop \ [])) (simple_match_to_of_match ac ifs)) p = NoAction" apply(rule ccontr) apply(drule OF_match_linear_not_noD) apply(clarsimp) apply(rename_tac x) apply(subgoal_tac "all_prerequisites x") apply(drule simple_match_to_of_matchD) apply(simp add: split3_def) apply(subst(asm) of_match_fields_safe_eq2) apply(simp;fail)+ using simple_match_to_of_match_generates_prereqs by blast lemma aux: \v = Some x \ the v = x\ by simp lemma s3_correct: assumes vsfwm: "list_all simple_match_valid (map (fst \ snd) ard)" assumes ippkt: "p_l2type p = 0x800" assumes iiifs: "p_iiface p \ set ifs" assumes oiifs: "list_all (\m. oiface (fst (snd m)) = ifaceAny) ard" shows "OF_match_linear OF_match_fields_safe (pack_OF_entries ifs ard) p = Action ao \ (\r af. generalized_sfw (map snd ard) p = (Some (r,af)) \ (if snd af = simple_action.Drop then ao = [] else ao = [Forward (fst af)]))" unfolding pack_OF_entries_def lr_of_tran_s3_def fun_app_def using vsfwm oiifs apply(induction ard) subgoal by(simp add: generalized_sfw_simps) apply simp apply(clarsimp simp add: generalized_sfw_simps split: prod.splits) apply(intro conjI) (* make two subgoals from one *) subgoal for ard x1 ac ad ba apply(clarsimp simp add: OF_match_linear_append split: prod.splits) apply(drule simple_match_to_of_matchI[rotated]) apply(rule iiifs) apply(rule ippkt) apply blast apply(clarsimp simp add: comp_def) apply(drule OF_match_linear_match_allsameaction[where \=OF_match_fields_safe and pri = x1 and oms = "simple_match_to_of_match ac ifs" and act = "case ba of simple_action.Accept \ [Forward ad] | simple_action.Drop \ []"]) apply(unfold OF_match_fields_safe_def comp_def) apply(erule aux) apply(clarsimp) apply(intro iffI) subgoal apply(rule exI[where x = ac]) apply(rule exI[where x = ad]) apply(rule exI[where x = ba]) apply(clarsimp simp: split3_def split: simple_action.splits flowtable_behavior.splits if_splits) done subgoal apply(clarsimp) apply(rename_tac b) apply(case_tac b) apply(simp_all) done done subgoal for ard x1 ac ad ba apply(simp add: OF_match_linear_append OF_match_fields_safe_def comp_def) apply(clarify) apply(subgoal_tac "OF_match_linear OF_match_fields_safe (map (\x. split3 OFEntry (x1, x, case ba of simple_action.Accept \ [Forward ad] | simple_action.Drop \ [])) (simple_match_to_of_match ac ifs)) p = NoAction") apply(simp;fail) apply(erule (1) s3_noaction_hlp) apply(simp add: match_ifaceAny;fail) done done context notes valid_prefix_00[simp, intro!] begin lemma lr_of_tran_s1_valid: "valid_prefixes rt \ gsfw_valid (lr_of_tran_s1 rt)" unfolding lr_of_tran_s1_def route2match_def gsfw_valid_def list_all_iff apply(clarsimp simp: simple_match_valid_def valid_prefix_fw_def) apply(intro conjI) apply force apply(simp add: valid_prefixes_alt_def) done end lemma simple_match_valid_fbs_rlen: "\valid_prefixes rt; simple_fw_valid fw; (a, aa, ab, b) \ set (annotate_rlen (lr_of_tran_fbs rt fw ifs))\ \ simple_match_valid aa" proof(goal_cases) case 1 note 1[unfolded lr_of_tran_fbs_def Let_def] have "gsfw_valid (map simple_rule_dtor fw)" using gsfw_validI 1 by blast moreover have "gsfw_valid (lr_of_tran_s1 rt)" using 1 lr_of_tran_s1_valid by blast ultimately have "gsfw_valid (generalized_fw_join (lr_of_tran_s1 rt) (map simple_rule_dtor fw))" using gsfw_join_valid by blast moreover have "(aa, ab, b) \ set (lr_of_tran_fbs rt fw ifs)" using 1 using in_annotate_rlen by fast ultimately show ?thesis unfolding lr_of_tran_fbs_def Let_def gsfw_valid_def list_all_iff by fastforce qed lemma simple_match_valid_fbs: "\valid_prefixes rt; simple_fw_valid fw\ \ list_all simple_match_valid (map fst (lr_of_tran_fbs rt fw ifs))" proof(goal_cases) case 1 note 1[unfolded lr_of_tran_fbs_def Let_def] have "gsfw_valid (map simple_rule_dtor fw)" using gsfw_validI 1 by blast moreover have "gsfw_valid (lr_of_tran_s1 rt)" using 1 lr_of_tran_s1_valid by blast ultimately have "gsfw_valid (generalized_fw_join (lr_of_tran_s1 rt) (map simple_rule_dtor fw))" using gsfw_join_valid by blast thus ?thesis unfolding lr_of_tran_fbs_def Let_def gsfw_valid_def list_all_iff by fastforce qed lemma lr_of_tran_prereqs: "valid_prefixes rt \ simple_fw_valid fw \ lr_of_tran rt fw ifs = Inr oft \ list_all (all_prerequisites \ ofe_fields) oft" unfolding lr_of_tran_def pack_OF_entries_def lr_of_tran_s3_def Let_def apply(simp add: map_concat comp_def prod.case_distrib split3_def split: if_splits) apply(simp add: list_all_iff) apply(clarsimp) apply(drule simple_match_valid_fbs_rlen[rotated]) apply(simp add: list_all_iff;fail) apply(simp add: list_all_iff;fail) apply(rule simple_match_to_of_match_generates_prereqs; assumption) done (* TODO: move. where? *) lemma OF_unsafe_safe_match3_eq: " list_all (all_prerequisites \ ofe_fields) oft \ OF_priority_match OF_match_fields_unsafe oft = OF_priority_match OF_match_fields_safe oft" unfolding OF_priority_match_def[abs_def] proof(goal_cases) case 1 from 1 have "\packet. [f\oft . OF_match_fields_unsafe (ofe_fields f) packet] = [f\oft . OF_match_fields_safe (ofe_fields f) packet]" apply(clarsimp simp add: list_all_iff of_match_fields_safe_eq) using of_match_fields_safe_eq by(metis (mono_tags, lifting) filter_cong) thus ?case by metis qed lemma OF_unsafe_safe_match_linear_eq: " list_all (all_prerequisites \ ofe_fields) oft \ OF_match_linear OF_match_fields_unsafe oft = OF_match_linear OF_match_fields_safe oft" unfolding fun_eq_iff by(induction oft) (clarsimp simp add: list_all_iff of_match_fields_safe_eq)+ lemma simple_action_ne[simp]: "b \ simple_action.Accept \ b = simple_action.Drop" "b \ simple_action.Drop \ b = simple_action.Accept" using simple_action.exhaust by blast+ lemma map_snd_apfst: "map snd (map (apfst x) l) = map snd l" unfolding map_map comp_def snd_apfst .. lemma match_ifaceAny_eq: "oiface m = ifaceAny \ simple_matches m p = simple_matches m (p\p_oiface := any\)" by(cases m) (simp add: simple_matches.simps match_ifaceAny) lemma no_oif_matchD: "no_oif_match fw \ simple_fw fw p = simple_fw fw (p\p_oiface := any\)" by(induction fw) (auto simp add: no_oif_match_def simple_fw_alt dest: match_ifaceAny_eq) lemma lr_of_tran_fbs_acceptD: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" shows "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Accept) \ simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\)" proof(goal_cases) case 1 note 1[unfolded lr_of_tran_fbs_def Let_def, THEN generalized_fw_joinD] then guess r1 .. then guess r2 .. note r12 = this note s1_correct[OF s1, of p] then guess rm .. then guess ra .. note rmra = this from r12 rmra have oifra: "oif = ra" by simp from r12 have sfw: "simple_fw fw p = Decision FinalAllow" using simple_fw_iff_generalized_fw_accept by blast note ifupdateirrel = no_oif_matchD[OF s2, where any = " output_iface (routing_table_semantics rt (p_dst p))" and p = p, symmetric] show ?case unfolding simple_linux_router_nol12_def by(simp add: Let_def ifupdateirrel sfw oifra rmra split: Option.bind_splits option.splits) qed lemma lr_of_tran_fbs_acceptI: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" "has_default_policy fw" shows "simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\) \ \r. generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Accept)" proof(goal_cases) from s2 have nud: "\p. simple_fw fw p \ Undecided" by (metis has_default_policy state.distinct(1)) note ifupdateirrel = no_oif_matchD[OF s2(1), symmetric] case 1 from 1 have "simple_fw fw p = Decision FinalAllow" by(simp add: simple_linux_router_nol12_def Let_def nud ifupdateirrel split: Option.bind_splits state.splits final_decision.splits) then obtain r where r: "generalized_sfw (map simple_rule_dtor fw) p = Some (r, simple_action.Accept)" using simple_fw_iff_generalized_fw_accept by blast have oif_def: "oif = output_iface (routing_table_semantics rt (p_dst p))" using 1 by(cases p) (simp add: simple_linux_router_nol12_def Let_def nud ifupdateirrel split: Option.bind_splits state.splits final_decision.splits) note s1_correct[OF s1, of p] then guess rm .. then guess ra .. note rmra = this show ?case unfolding lr_of_tran_fbs_def Let_def apply(rule exI) apply(rule generalized_fw_joinI) unfolding oif_def using rmra apply simp apply(rule r) done qed lemma lr_of_tran_fbs_dropD: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" shows "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop) \ simple_linux_router_nol12 rt fw p = None" proof(goal_cases) note ifupdateirrel = no_oif_matchD[OF s2(1), symmetric] case 1 from 1[unfolded lr_of_tran_fbs_def Let_def, THEN generalized_fw_joinD] obtain rr fr where "generalized_sfw (lr_of_tran_s1 rt) p = Some (rr, oif) \ generalized_sfw (map simple_rule_dtor fw) p = Some (fr, simple_action.Drop) \ Some r = simple_match_and rr fr" by presburger hence fd: "\u. simple_fw fw (p\p_oiface := u\) = Decision FinalDeny" unfolding ifupdateirrel using simple_fw_iff_generalized_fw_drop by blast show ?thesis by(clarsimp simp: simple_linux_router_nol12_def Let_def fd split: Option.bind_splits) qed lemma lr_of_tran_fbs_dropI: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" "has_default_policy fw" shows "simple_linux_router_nol12 rt fw p = None \ \r oif. generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop)" proof(goal_cases) from s2 have nud: "\p. simple_fw fw p \ Undecided" by (metis has_default_policy state.distinct(1)) note ifupdateirrel = no_oif_matchD[OF s2(1), symmetric] case 1 from 1 have "simple_fw fw p = Decision FinalDeny" by(simp add: simple_linux_router_nol12_def Let_def nud ifupdateirrel split: Option.bind_splits state.splits final_decision.splits) then obtain r where r: "generalized_sfw (map simple_rule_dtor fw) p = Some (r, simple_action.Drop)" using simple_fw_iff_generalized_fw_drop by blast note s1_correct[OF s1, of p] then guess rm .. then guess ra .. note rmra = this show ?case unfolding lr_of_tran_fbs_def Let_def apply(rule exI) apply(rule exI[where x = ra]) apply(rule generalized_fw_joinI) using rmra apply simp apply(rule r) done qed lemma no_oif_match_fbs: "no_oif_match fw \ list_all (\m. oiface (fst (snd m)) = ifaceAny) (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))" proof(goal_cases) case 1 have c: "\mr ar mf af f a. \(mr, ar) \ set (lr_of_tran_s1 rt); (mf, af) \ simple_rule_dtor ` set fw; simple_match_and mr mf = Some a\ \ oiface a = ifaceAny" proof(goal_cases) case (1 mr ar mf af f a) have "oiface mr = ifaceAny" using 1(1) unfolding lr_of_tran_s1_def route2match_def by(clarsimp simp add: Set.image_iff) moreover have "oiface mf = ifaceAny" using 1(2) \no_oif_match fw\ unfolding no_oif_match_def simple_rule_dtor_def[abs_def] by(clarsimp simp: list_all_iff split: simple_rule.splits) fastforce ultimately show ?case using 1(3) by(cases a; cases mr; cases mf) (simp add: iface_conjunct_ifaceAny split: option.splits) qed have la: "list_all (\m. oiface (fst m) = ifaceAny) (lr_of_tran_fbs rt fw ifs)" unfolding lr_of_tran_fbs_def Let_def list_all_iff apply(clarify) apply(subst(asm) generalized_sfw_join_set) apply(clarsimp) using c by blast thus ?case proof(goal_cases) case 1 have *: "(\m. oiface (fst (snd m)) = ifaceAny) = (\m. oiface (fst m) = ifaceAny) \ snd" unfolding comp_def .. show ?case unfolding * list_all_map[symmetric] map_snd_apfst map_snd_annotate_rlen using la . qed qed lemma lr_of_tran_correct: fixes p :: "(32, 'a) simple_packet_ext_scheme" assumes nerr: "lr_of_tran rt fw ifs = Inr oft" and ippkt: "p_l2type p = 0x800" and ifvld: "p_iiface p \ set ifs" shows "OF_priority_match OF_match_fields_safe oft p = Action [Forward oif] \ simple_linux_router_nol12 rt fw p = (Some (p\p_oiface := oif\))" "OF_priority_match OF_match_fields_safe oft p = Action [] \ simple_linux_router_nol12 rt fw p = None" (* fun stuff: *) "OF_priority_match OF_match_fields_safe oft p \ NoAction" "OF_priority_match OF_match_fields_safe oft p \ Undefined" "OF_priority_match OF_match_fields_safe oft p = Action ls \ length ls \ 1" "\ls. length ls \ 1 \ OF_priority_match OF_match_fields_safe oft p = Action ls" proof - have s1: "valid_prefixes rt" "has_default_route rt" and s2: "has_default_policy fw" "simple_fw_valid fw" "no_oif_match fw" and difs: "distinct ifs" using nerr unfolding lr_of_tran_def by(simp_all split: if_splits) have "no_oif_match fw" using nerr unfolding lr_of_tran_def by(simp split: if_splits) note s2 = s2 this have unsafe_safe_eq: "OF_priority_match OF_match_fields_unsafe oft = OF_priority_match OF_match_fields_safe oft" "OF_match_linear OF_match_fields_unsafe oft = OF_match_linear OF_match_fields_safe oft" apply(subst OF_unsafe_safe_match3_eq; (rule lr_of_tran_prereqs s1 s2 nerr refl)+) apply(subst OF_unsafe_safe_match_linear_eq; (rule lr_of_tran_prereqs s1 s2 nerr refl)+) done have lin: "OF_priority_match OF_match_fields_safe oft = OF_match_linear OF_match_fields_safe oft" using OF_eq[OF lr_of_tran_no_overlaps lr_of_tran_sorted_descending, OF difs nerr[symmetric] nerr[symmetric]] unfolding fun_eq_iff unsafe_safe_eq by metis let ?ard = "map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs))" have oft_def: "oft = pack_OF_entries ifs ?ard" using nerr unfolding lr_of_tran_def Let_def by(simp split: if_splits) have vld: "list_all simple_match_valid (map (fst \ snd) ?ard)" unfolding fun_app_def map_map[symmetric] snd_apfst map_snd_apfst map_snd_annotate_rlen using simple_match_valid_fbs[OF s1(1) s2(2)] . have *: "list_all (\m. oiface (fst (snd m)) = ifaceAny) ?ard" using no_oif_match_fbs[OF s2(3)] . have not_undec: "\p. simple_fw fw p \ Undecided" by (metis has_default_policy s2(1) state.simps(3)) have w1_1: "\oif. OF_match_linear OF_match_fields_safe oft p = Action [Forward oif] \ simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\) \ oif = output_iface (routing_table_semantics rt (p_dst p))" proof(intro conjI, goal_cases) case (1 oif) note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD1, unfolded oft_def[symmetric], OF 1] hence "\r. generalized_sfw (map snd (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))) p = Some (r, (oif, simple_action.Accept))" by(clarsimp split: if_splits) then obtain r where "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, (oif, simple_action.Accept))" unfolding map_map comp_def snd_apfst map_snd_annotate_rlen by blast thus ?case using lr_of_tran_fbs_acceptD[OF s1 s2(3)] by metis thus "oif = output_iface (routing_table_semantics rt (p_dst p))" by(cases p) (clarsimp simp: simple_linux_router_nol12_def Let_def not_undec split: Option.bind_splits state.splits final_decision.splits) qed have w1_2: "\oif. simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\) \ OF_match_linear OF_match_fields_safe oft p = Action [Forward oif]" proof(goal_cases) case (1 oif) note lr_of_tran_fbs_acceptI[OF s1 s2(3) s2(1) this, of ifs] then guess r .. note r = this hence "generalized_sfw (map snd (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))) p = Some (r, (oif, simple_action.Accept))" unfolding map_snd_apfst map_snd_annotate_rlen . moreover note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD2, unfolded oft_def[symmetric], of "[Forward oif]"] ultimately show ?case by simp qed show w1: "\oif. (OF_priority_match OF_match_fields_safe oft p = Action [Forward oif]) = (simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\))" unfolding lin using w1_1 w1_2 by blast show w2: "(OF_priority_match OF_match_fields_safe oft p = Action []) = (simple_linux_router_nol12 rt fw p = None)" unfolding lin proof(rule iffI, goal_cases) case 1 note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD1, unfolded oft_def[symmetric], OF 1] then obtain r oif where roif: "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop)" unfolding map_snd_apfst map_snd_annotate_rlen by(clarsimp split: if_splits) note lr_of_tran_fbs_dropD[OF s1 s2(3) this] thus ?case . next case 2 note lr_of_tran_fbs_dropI[OF s1 s2(3) s2(1) this, of ifs] then obtain r oif where "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop)" by blast hence "generalized_sfw (map snd (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))) p = Some (r, oif, simple_action.Drop)" unfolding map_snd_apfst map_snd_annotate_rlen . moreover note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD2, unfolded oft_def[symmetric], of "[]"] ultimately show ?case by force qed have lr_determ: "\a. simple_linux_router_nol12 rt fw p = Some a \ a = p\p_oiface := output_iface (routing_table_semantics rt (p_dst p))\" by(clarsimp simp: simple_linux_router_nol12_def Let_def not_undec split: Option.bind_splits state.splits final_decision.splits) show notno: "OF_priority_match OF_match_fields_safe oft p \ NoAction" apply(cases "simple_linux_router_nol12 rt fw p") using w2 apply(simp) using w1[of "output_iface (routing_table_semantics rt (p_dst p))"] apply(simp) apply(drule lr_determ) apply(simp) done show notub: "OF_priority_match OF_match_fields_safe oft p \ Undefined" unfolding lin using OF_match_linear_ne_Undefined . show notmult: "\ls. OF_priority_match OF_match_fields_safe oft p = Action ls \ length ls \ 1" apply(cases "simple_linux_router_nol12 rt fw p") using w2 apply(simp) using w1[of "output_iface (routing_table_semantics rt (p_dst p))"] apply(simp) apply(drule lr_determ) apply(clarsimp) done show "\ls. length ls \ 1 \ OF_priority_match OF_match_fields_safe oft p = Action ls" apply(cases "OF_priority_match OF_match_fields_safe oft p") using notmult apply blast using notno apply blast using notub apply blast done qed end diff --git a/thys/LOFT/ROOT b/thys/LOFT/ROOT --- a/thys/LOFT/ROOT +++ b/thys/LOFT/ROOT @@ -1,27 +1,29 @@ chapter AFP session LOFT (AFP) = Iptables_Semantics + options [timeout = 1200] + sessions + "Pure-ex" directories "Examples/OF_conv_test" "Examples/RFC2544" theories [document = false] OpenFlow_Helpers (* just some helper lemmas *) List_Group Sort_Descending theories OpenFlow_Action OpenFlow_Matches OpenFlow_Serialize Semantics_OpenFlow Featherweight_OpenFlow_Comparison LinuxRouter_OpenFlow_Translation "Examples/OF_conv_test/OF_conv_test" "Examples/RFC2544/RFC2544" OpenFlow_Documentation document_files "root.tex" "chap3.tex" "root.bib" "moeptikz.sty" "bench.csv" 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,1486 @@ (* 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 + then obtain c1 c2 :: real + where c: + "c1 > 0" "\\<^sub>F x in F. norm (h x) \ c1 * norm (f x * g x)" + "c2 > 0" "\\<^sub>F x in F. c2 * norm (f x * g x) \ norm (h x)" + unfolding bigtheta_def by (blast elim: landau_o.bigE) 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_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_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_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_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/Linear_Recurrences/RatFPS.thy b/thys/Linear_Recurrences/RatFPS.thy --- a/thys/Linear_Recurrences/RatFPS.thy +++ b/thys/Linear_Recurrences/RatFPS.thy @@ -1,948 +1,966 @@ (* File: RatFPS.thy Author: Manuel Eberl, TU München *) section \Rational formal power series\ theory RatFPS imports Complex_Main "HOL-Computational_Algebra.Computational_Algebra" "HOL-Computational_Algebra.Polynomial_Factorial" begin subsection \Some auxiliary\ abbreviation constant_term :: "'a poly \ 'a::zero" where "constant_term p \ coeff p 0" lemma coeff_0_mult: "coeff (p * q) 0 = coeff p 0 * coeff q 0" by (simp add: coeff_mult) lemma coeff_0_div: assumes "coeff p 0 \ 0" assumes "(q :: 'a :: field poly) dvd p" shows "coeff (p div q) 0 = coeff p 0 div coeff q 0" proof (cases "q = 0") case False from assms have "p = p div q * q" by simp also have "coeff \ 0 = coeff (p div q) 0 * coeff q 0" by (simp add: coeff_0_mult) finally show ?thesis using assms by auto qed simp_all lemma coeff_0_add_fract_nonzero: assumes "coeff (snd (quot_of_fract x)) 0 \ 0" "coeff (snd (quot_of_fract y)) 0 \ 0" shows "coeff (snd (quot_of_fract (x + y))) 0 \ 0" proof - define num where "num = fst (quot_of_fract x) * snd (quot_of_fract y) + snd (quot_of_fract x) * fst (quot_of_fract y)" define denom where "denom = snd (quot_of_fract x) * snd (quot_of_fract y)" define z where "z = (num, denom)" from assms have "snd z \ 0" by (auto simp: denom_def z_def) - from normalize_quotE'[OF this] guess d . note d = this + then obtain d where d: + "fst z = fst (normalize_quot z) * d" + "snd z = snd (normalize_quot z) * d" + "d dvd fst z" + "d dvd snd z" + "d \ 0" + by (rule normalize_quotE') from assms have z: "coeff (snd z) 0 \ 0" by (simp add: z_def denom_def coeff_0_mult) have "coeff (snd (quot_of_fract (x + y))) 0 = coeff (snd (normalize_quot z)) 0" by (simp add: quot_of_fract_add Let_def case_prod_unfold z_def num_def denom_def) also from z have "\ \ 0" using d by (simp add: d coeff_0_mult) finally show ?thesis . qed lemma coeff_0_normalize_quot_nonzero [simp]: assumes "coeff (snd x) 0 \ 0" shows "coeff (snd (normalize_quot x)) 0 \ 0" proof - from assms have "snd x \ 0" by auto - from normalize_quotE'[OF this] guess d . + then obtain d where + "fst x = fst (normalize_quot x) * d" + "snd x = snd (normalize_quot x) * d" + "d dvd fst x" + "d dvd snd x" + "d \ 0" + by (rule normalize_quotE') with assms show ?thesis by (auto simp: coeff_0_mult) qed abbreviation numerator :: "'a fract \ 'a::{ring_gcd,idom_divide,semiring_gcd_mult_normalize}" where "numerator x \ fst (quot_of_fract x)" abbreviation denominator :: "'a fract \ 'a::{ring_gcd,idom_divide,semiring_gcd_mult_normalize}" where "denominator x \ snd (quot_of_fract x)" declare unit_factor_snd_quot_of_fract [simp] normalize_snd_quot_of_fract [simp] lemma constant_term_denominator_nonzero_imp_constant_term_denominator_div_gcd_nonzero: "constant_term (denominator x div gcd a (denominator x)) \ 0" if "constant_term (denominator x) \ 0" using that coeff_0_normalize_quot_nonzero [of "(a, denominator x)"] normalize_quot_proj(2) [of "denominator x" a] by simp subsection \The type of rational formal power series\ typedef (overloaded) 'a :: field_gcd ratfps = "{x :: 'a poly fract. constant_term (denominator x) \ 0}" by (rule exI [of _ 0]) simp setup_lifting type_definition_ratfps instantiation ratfps :: (field_gcd) idom begin lift_definition zero_ratfps :: "'a ratfps" is "0" by simp lift_definition one_ratfps :: "'a ratfps" is "1" by simp lift_definition uminus_ratfps :: "'a ratfps \ 'a ratfps" is "uminus" by (simp add: quot_of_fract_uminus case_prod_unfold Let_def) lift_definition plus_ratfps :: "'a ratfps \ 'a ratfps \ 'a ratfps" is "(+)" by (rule coeff_0_add_fract_nonzero) lift_definition minus_ratfps :: "'a ratfps \ 'a ratfps \ 'a ratfps" is "(-)" by (simp only: diff_conv_add_uminus, rule coeff_0_add_fract_nonzero) (simp_all add: quot_of_fract_uminus Let_def case_prod_unfold) lift_definition times_ratfps :: "'a ratfps \ 'a ratfps \ 'a ratfps" is "(*)" by (simp add: quot_of_fract_mult Let_def case_prod_unfold coeff_0_mult constant_term_denominator_nonzero_imp_constant_term_denominator_div_gcd_nonzero) instance by (standard; transfer) (simp_all add: ring_distribs) end fun ratfps_nth_aux :: "('a::field) poly \ nat \ 'a" where "ratfps_nth_aux p 0 = inverse (coeff p 0)" | "ratfps_nth_aux p n = - inverse (coeff p 0) * sum (\i. coeff p i * ratfps_nth_aux p (n - i)) {1..n}" lemma ratfps_nth_aux_correct: "ratfps_nth_aux p n = natfun_inverse (fps_of_poly p) n" by (induction p n rule: ratfps_nth_aux.induct) simp_all lift_definition ratfps_nth :: "'a :: field_gcd ratfps \ nat \ 'a" is "\x n. let (a,b) = quot_of_fract x in (\i = 0..n. coeff a i * ratfps_nth_aux b (n - i))" . lift_definition ratfps_subdegree :: "'a :: field_gcd ratfps \ nat" is "\x. poly_subdegree (fst (quot_of_fract x))" . context includes lifting_syntax begin lemma RatFPS_parametric: "(rel_prod (=) (=) ===> (=)) (\(p,q). if coeff q 0 = 0 then 0 else quot_to_fract (p, q)) (\(p,q). if coeff q 0 = 0 then 0 else quot_to_fract (p, q))" by transfer_prover end lemma normalize_quot_quot_of_fract [simp]: "normalize_quot (quot_of_fract x) = quot_of_fract x" by (rule normalize_quot_id, rule quot_of_fract_in_normalized_fracts) context assumes "SORT_CONSTRAINT('a::field_gcd)" begin lift_definition quot_of_ratfps :: "'a ratfps \ ('a poly \ 'a poly)" is "quot_of_fract :: 'a poly fract \ ('a poly \ 'a poly)" . lift_definition quot_to_ratfps :: "('a poly \ 'a poly) \ 'a ratfps" is "\(x,y). let (x',y') = normalize_quot (x,y) in if coeff y' 0 = 0 then 0 else quot_to_fract (x',y')" by (simp add: case_prod_unfold Let_def quot_of_fract_quot_to_fract) lemma quot_to_ratfps_quot_of_ratfps [code abstype]: "quot_to_ratfps (quot_of_ratfps x) = x" by transfer (simp add: case_prod_unfold Let_def) lemma coeff_0_snd_quot_of_ratfps_nonzero [simp]: "coeff (snd (quot_of_ratfps x)) 0 \ 0" by transfer simp lemma quot_of_ratfps_quot_to_ratfps: "coeff (snd x) 0 \ 0 \ x \ normalized_fracts \ quot_of_ratfps (quot_to_ratfps x) = x" by transfer (simp add: Let_def case_prod_unfold coeff_0_normalize_quot_nonzero quot_of_fract_quot_to_fract normalize_quot_id) lemma quot_of_ratfps_0 [simp, code abstract]: "quot_of_ratfps 0 = (0, 1)" by transfer simp_all lemma quot_of_ratfps_1 [simp, code abstract]: "quot_of_ratfps 1 = (1, 1)" by transfer simp_all lift_definition ratfps_of_poly :: "'a poly \ 'a ratfps" is "to_fract :: 'a poly \ _" by transfer simp lemma ratfps_of_poly_code [code abstract]: "quot_of_ratfps (ratfps_of_poly p) = (p, 1)" by transfer' simp lemmas zero_ratfps_code = quot_of_ratfps_0 lemmas one_ratfps_code = quot_of_ratfps_1 lemma uminus_ratfps_code [code abstract]: "quot_of_ratfps (- x) = (let (a, b) = quot_of_ratfps x in (-a, b))" by transfer (rule quot_of_fract_uminus) lemma plus_ratfps_code [code abstract]: "quot_of_ratfps (x + y) = (let (a,b) = quot_of_ratfps x; (c,d) = quot_of_ratfps y in normalize_quot (a * d + b * c, b * d))" by transfer' (rule quot_of_fract_add) lemma minus_ratfps_code [code abstract]: "quot_of_ratfps (x - y) = (let (a,b) = quot_of_ratfps x; (c,d) = quot_of_ratfps y in normalize_quot (a * d - b * c, b * d))" by transfer' (rule quot_of_fract_diff) definition ratfps_cutoff :: "nat \ 'a :: field_gcd ratfps \ 'a poly" where "ratfps_cutoff n x = poly_of_list (map (ratfps_nth x) [0.. 'a :: field_gcd ratfps \ 'a ratfps" where "ratfps_shift n x = (let (a, b) = quot_of_ratfps (x - ratfps_of_poly (ratfps_cutoff n x)) in quot_to_ratfps (poly_shift n a, b))" lemma times_ratfps_code [code abstract]: "quot_of_ratfps (x * y) = (let (a,b) = quot_of_ratfps x; (c,d) = quot_of_ratfps y; (e,f) = normalize_quot (a,d); (g,h) = normalize_quot (c,b) in (e*g, f*h))" by transfer' (rule quot_of_fract_mult) lemma ratfps_nth_code [code]: "ratfps_nth x n = (let (a,b) = quot_of_ratfps x in \i = 0..n. coeff a i * ratfps_nth_aux b (n - i))" by transfer' simp lemma ratfps_subdegree_code [code]: "ratfps_subdegree x = poly_subdegree (fst (quot_of_ratfps x))" by transfer simp end instantiation ratfps :: ("field_gcd") inverse begin lift_definition inverse_ratfps :: "'a ratfps \ 'a ratfps" is "\x. let (a,b) = quot_of_fract x in if coeff a 0 = 0 then 0 else inverse x" by (auto simp: case_prod_unfold Let_def quot_of_fract_inverse) lift_definition divide_ratfps :: "'a ratfps \ 'a ratfps \ 'a ratfps" is "\f g. (if g = 0 then 0 else let n = ratfps_subdegree g; h = ratfps_shift n g in ratfps_shift n (f * inverse h))" . instance .. end lemma ratfps_inverse_code [code abstract]: "quot_of_ratfps (inverse x) = (let (a,b) = quot_of_ratfps x in if coeff a 0 = 0 then (0, 1) else let u = unit_factor a in (b div u, a div u))" by transfer' (simp_all add: Let_def case_prod_unfold quot_of_fract_inverse) instantiation ratfps :: (equal) equal begin definition equal_ratfps :: "'a ratfps \ 'a ratfps \ bool" where [simp]: "equal_ratfps x y \ x = y" instance by standard simp end lemma quot_of_fract_eq_iff [simp]: "quot_of_fract x = quot_of_fract y \ x = y" by transfer (auto simp: normalize_quot_eq_iff) lemma equal_ratfps_code [code]: "HOL.equal x y \ quot_of_ratfps x = quot_of_ratfps y" unfolding equal_ratfps_def by transfer simp lemma fps_of_poly_quot_normalize_quot [simp]: "fps_of_poly (fst (normalize_quot x)) / fps_of_poly (snd (normalize_quot x)) = fps_of_poly (fst x) / fps_of_poly (snd x)" if "(snd x :: 'a :: field_gcd poly) \ 0" proof - from that obtain d where "fst x = fst (normalize_quot x) * d" and "snd x = snd (normalize_quot x) * d" and "d \ 0" by (rule normalize_quotE') then show ?thesis by (simp add: fps_of_poly_mult) qed lemma fps_of_poly_quot_normalize_quot' [simp]: "fps_of_poly (fst (normalize_quot x)) / fps_of_poly (snd (normalize_quot x)) = fps_of_poly (fst x) / fps_of_poly (snd x)" if "coeff (snd x) 0 \ (0 :: 'a :: field_gcd)" using that by (auto intro: fps_of_poly_quot_normalize_quot) lift_definition fps_of_ratfps :: "'a :: field_gcd ratfps \ 'a fps" is "\x. fps_of_poly (numerator x) / fps_of_poly (denominator x)" . lemma fps_of_ratfps_altdef: "fps_of_ratfps x = (case quot_of_ratfps x of (a, b) \ fps_of_poly a / fps_of_poly b)" by transfer (simp add: case_prod_unfold) lemma fps_of_ratfps_ratfps_of_poly [simp]: "fps_of_ratfps (ratfps_of_poly p) = fps_of_poly p" by transfer simp lemma fps_of_ratfps_0 [simp]: "fps_of_ratfps 0 = 0" by transfer simp lemma fps_of_ratfps_1 [simp]: "fps_of_ratfps 1 = 1" by transfer simp lemma fps_of_ratfps_uminus [simp]: "fps_of_ratfps (-x) = - fps_of_ratfps x" by transfer (simp add: quot_of_fract_uminus case_prod_unfold Let_def fps_of_poly_simps dvd_neg_div) lemma fps_of_ratfps_add [simp]: "fps_of_ratfps (x + y) = fps_of_ratfps x + fps_of_ratfps y" by transfer (simp add: quot_of_fract_add Let_def case_prod_unfold fps_of_poly_simps) lemma fps_of_ratfps_diff [simp]: "fps_of_ratfps (x - y) = fps_of_ratfps x - fps_of_ratfps y" by transfer (simp add: quot_of_fract_diff Let_def case_prod_unfold fps_of_poly_simps) lemma is_unit_div_div_commute: "is_unit b \ is_unit c \ a div b div c = a div c div b" by (metis is_unit_div_mult2_eq mult.commute) lemma fps_of_ratfps_mult [simp]: "fps_of_ratfps (x * y) = fps_of_ratfps x * fps_of_ratfps y" proof (transfer, goal_cases) case (1 x y) moreover define x' y' where "x' = quot_of_fract x" and "y' = quot_of_fract y" ultimately have assms: "coeff (snd x') 0 \ 0" "coeff (snd y') 0 \ 0" by simp_all moreover define w z where "w = normalize_quot (fst x', snd y')" and "z = normalize_quot (fst y', snd x')" ultimately have unit: "coeff (snd x') 0 \ 0" "coeff (snd y') 0 \ 0" "coeff (snd w) 0 \ 0" "coeff (snd z) 0 \ 0" by (simp_all add: coeff_0_normalize_quot_nonzero) have "fps_of_poly (fst w * fst z) / fps_of_poly (snd w * snd z) = (fps_of_poly (fst w) / fps_of_poly (snd w)) * (fps_of_poly (fst z) / fps_of_poly (snd z))" (is "_ = ?A * ?B") by (simp add: is_unit_div_mult2_eq fps_of_poly_mult unit_div_mult_swap unit_div_commute unit) also have "\ = (fps_of_poly (fst x') / fps_of_poly (snd x')) * (fps_of_poly (fst y') / fps_of_poly (snd y'))" using unit by (simp add: w_def z_def unit_div_commute unit_div_mult_swap is_unit_div_div_commute) finally show ?case by (simp add: w_def z_def x'_def y'_def Let_def case_prod_unfold quot_of_fract_mult mult_ac) qed lemma div_const_unit_poly: "is_unit c \ p div [:c:] = smult (1 div c) p" by (simp add: is_unit_const_poly_iff unit_eq_div1) lemma normalize_field: "normalize (x :: 'a :: {normalization_semidom,field}) = (if x = 0 then 0 else 1)" by (auto simp: normalize_1_iff dvd_field_iff) lemma unit_factor_field [simp]: "unit_factor (x :: 'a :: {normalization_semidom,field}) = x" using unit_factor_mult_normalize[of x] normalize_field[of x] by (simp split: if_splits) lemma fps_of_poly_normalize_field: "fps_of_poly (normalize (p :: 'a :: {field, normalization_semidom} poly)) = fps_of_poly p * fps_const (inverse (lead_coeff p))" by (cases "p = 0") (simp_all add: normalize_poly_def div_const_unit_poly divide_simps dvd_field_iff) lemma unit_factor_poly_altdef: "unit_factor p = monom (unit_factor (lead_coeff p)) 0" by (simp add: unit_factor_poly_def monom_altdef) lemma div_const_poly: "p div [:c::'a::field:] = smult (inverse c) p" by (cases "c = 0") (simp_all add: unit_eq_div1 is_unit_triv) lemma fps_of_ratfps_inverse [simp]: "fps_of_ratfps (inverse x) = inverse (fps_of_ratfps x)" proof (transfer, goal_cases) case (1 x) hence "smult (lead_coeff (fst (quot_of_fract x))) (snd (quot_of_fract x)) div unit_factor (fst (quot_of_fract x)) = snd (quot_of_fract x)" if "fst (quot_of_fract x) \ 0" using that by (simp add: unit_factor_poly_altdef monom_0 div_const_poly) with 1 show ?case by (auto simp: Let_def case_prod_unfold fps_divide_unit fps_inverse_mult quot_of_fract_inverse mult_ac fps_of_poly_simps fps_const_inverse fps_of_poly_normalize_field div_smult_left [symmetric]) qed context includes fps_notation begin lemma ratfps_nth_altdef: "ratfps_nth x n = fps_of_ratfps x $ n" by transfer (simp_all add: case_prod_unfold fps_divide_unit fps_times_def fps_inverse_def ratfps_nth_aux_correct Let_def) lemma fps_of_ratfps_is_unit: "fps_of_ratfps a $ 0 \ 0 \ ratfps_nth a 0 \ 0" by (simp add: ratfps_nth_altdef) lemma ratfps_nth_0 [simp]: "ratfps_nth 0 n = 0" by (simp add: ratfps_nth_altdef) lemma fps_of_ratfps_cases: obtains p q where "coeff q 0 \ 0" "fps_of_ratfps f = fps_of_poly p / fps_of_poly q" by (rule that[of "snd (quot_of_ratfps f)" "fst (quot_of_ratfps f)"]) (simp_all add: fps_of_ratfps_altdef case_prod_unfold) lemma fps_of_ratfps_cutoff [simp]: "fps_of_poly (ratfps_cutoff n x) = fps_cutoff n (fps_of_ratfps x)" by (simp add: fps_eq_iff ratfps_cutoff_def nth_default_def ratfps_nth_altdef) lemma subdegree_fps_of_ratfps: "subdegree (fps_of_ratfps x) = ratfps_subdegree x" by transfer (simp_all add: case_prod_unfold subdegree_div_unit poly_subdegree_def) lemma ratfps_subdegree_altdef: "ratfps_subdegree x = subdegree (fps_of_ratfps x)" using subdegree_fps_of_ratfps .. end code_datatype fps_of_ratfps lemma fps_zero_code [code]: "0 = fps_of_ratfps 0" by simp lemma fps_one_code [code]: "1 = fps_of_ratfps 1" by simp lemma fps_const_code [code]: "fps_const c = fps_of_poly [:c:]" by simp lemma fps_of_poly_code [code]: "fps_of_poly p = fps_of_ratfps (ratfps_of_poly p)" by simp lemma fps_X_code [code]: "fps_X = fps_of_ratfps (ratfps_of_poly [:0,1:])" by simp lemma fps_nth_code [code]: "fps_nth (fps_of_ratfps x) n = ratfps_nth x n" by (simp add: ratfps_nth_altdef) lemma fps_uminus_code [code]: "- fps_of_ratfps x = fps_of_ratfps (-x)" by simp lemma fps_add_code [code]: "fps_of_ratfps x + fps_of_ratfps y = fps_of_ratfps (x + y)" by simp lemma fps_diff_code [code]: "fps_of_ratfps x - fps_of_ratfps y = fps_of_ratfps (x - y)" by simp lemma fps_mult_code [code]: "fps_of_ratfps x * fps_of_ratfps y = fps_of_ratfps (x * y)" by simp lemma fps_inverse_code [code]: "inverse (fps_of_ratfps x) = fps_of_ratfps (inverse x)" by simp lemma fps_cutoff_code [code]: "fps_cutoff n (fps_of_ratfps x) = fps_of_poly (ratfps_cutoff n x)" by simp lemmas subdegree_code [code] = subdegree_fps_of_ratfps lemma fractrel_normalize_quot: "fractrel p p \ fractrel q q \ fractrel (normalize_quot p) (normalize_quot q) \ fractrel p q" by (subst fractrel_normalize_quot_left fractrel_normalize_quot_right, simp)+ (rule refl) lemma fps_of_ratfps_eq_iff [simp]: "fps_of_ratfps p = fps_of_ratfps q \ p = q" proof - { fix p q :: "'a poly fract" assume "fractrel (quot_of_fract p) (quot_of_fract q)" hence "p = q" by transfer (simp only: fractrel_normalize_quot) } note A = this show ?thesis by transfer (auto simp: case_prod_unfold unit_eq_div1 unit_eq_div2 unit_div_commute intro: A) qed lemma fps_of_ratfps_eq_zero_iff [simp]: "fps_of_ratfps p = 0 \ p = 0" by (simp del: fps_of_ratfps_0 add: fps_of_ratfps_0 [symmetric]) lemma unit_factor_snd_quot_of_ratfps [simp]: "unit_factor (snd (quot_of_ratfps x)) = 1" by transfer simp lemma poly_shift_times_monom_le: "n \ m \ poly_shift n (monom c m * p) = monom c (m - n) * p" by (intro poly_eqI) (auto simp: coeff_monom_mult coeff_poly_shift) lemma poly_shift_times_monom_ge: "n \ m \ poly_shift n (monom c m * p) = smult c (poly_shift (n - m) p)" by (intro poly_eqI) (auto simp: coeff_monom_mult coeff_poly_shift) lemma poly_shift_times_monom: "poly_shift n (monom c n * p) = smult c p" by (intro poly_eqI) (auto simp: coeff_monom_mult coeff_poly_shift) lemma monom_times_poly_shift: assumes "poly_subdegree p \ n" shows "monom c n * poly_shift n p = smult c p" (is "?lhs = ?rhs") proof (intro poly_eqI) fix k show "coeff ?lhs k = coeff ?rhs k" proof (cases "k < n") case True with assms have "k < poly_subdegree p" by simp hence "coeff p k = 0" by (simp add: coeff_less_poly_subdegree) thus ?thesis by (auto simp: coeff_monom_mult coeff_poly_shift) qed (auto simp: coeff_monom_mult coeff_poly_shift) qed lemma monom_times_poly_shift': assumes "poly_subdegree p \ n" shows "monom (1 :: 'a :: comm_semiring_1) n * poly_shift n p = p" by (simp add: monom_times_poly_shift[OF assms]) lemma subdegree_minus_cutoff_ge: assumes "f - fps_cutoff n (f :: 'a :: ab_group_add fps) \ 0" shows "subdegree (f - fps_cutoff n f) \ n" using assms by (rule subdegree_geI) simp_all lemma fps_shift_times_X_power'': "fps_shift n (fps_X ^ n * f :: 'a :: comm_ring_1 fps) = f" using fps_shift_times_fps_X_power'[of n f] by (simp add: mult.commute) lemma ratfps_shift_code [code abstract]: "quot_of_ratfps (ratfps_shift n x) = (let (a, b) = quot_of_ratfps (x - ratfps_of_poly (ratfps_cutoff n x)) in (poly_shift n a, b))" (is "?lhs1 = ?rhs1") and fps_of_ratfps_shift [simp]: "fps_of_ratfps (ratfps_shift n x) = fps_shift n (fps_of_ratfps x)" proof - include fps_notation define x' where "x' = ratfps_of_poly (ratfps_cutoff n x)" define y where "y = quot_of_ratfps (x - x')" have "coprime (fst y) (snd y)" unfolding y_def by transfer (rule coprime_quot_of_fract) also have fst_y: "fst y = monom 1 n * poly_shift n (fst y)" proof (cases "x = x'") case False have "poly_subdegree (fst y) = subdegree (fps_of_poly (fst y))" by (simp add: poly_subdegree_def) also have "\ = subdegree (fps_of_poly (fst y) / fps_of_poly (snd y))" by (subst subdegree_div_unit) (simp_all add: y_def) also have "fps_of_poly (fst y) / fps_of_poly (snd y) = fps_of_ratfps (x - x')" unfolding y_def by transfer (simp add: case_prod_unfold) also from False have "subdegree \ \ n" proof (intro subdegree_geI) fix k assume "k < n" thus "fps_of_ratfps (x - x') $ k = 0" by (simp add: x'_def) qed simp_all finally show ?thesis by (rule monom_times_poly_shift' [symmetric]) qed (simp_all add: y_def) finally have coprime: "coprime (poly_shift n (fst y)) (snd y)" by simp have "quot_of_ratfps (ratfps_shift n x) = quot_of_ratfps (quot_to_ratfps (poly_shift n (fst y), snd y))" by (simp add: ratfps_shift_def Let_def case_prod_unfold x'_def y_def) also from coprime have "\ = (poly_shift n (fst y), snd y)" by (intro quot_of_ratfps_quot_to_ratfps) (simp_all add: y_def normalized_fracts_def) also have "\ = ?rhs1" by (simp add: case_prod_unfold Let_def y_def x'_def) finally show eq: "?lhs1 = ?rhs1" . have "fps_shift n (fps_of_ratfps x) = fps_shift n (fps_of_ratfps (x - x'))" by (intro fps_ext) (simp_all add: x'_def) also have "fps_of_ratfps (x - x') = fps_of_poly (fst y) / fps_of_poly (snd y)" by (simp add: fps_of_ratfps_altdef y_def case_prod_unfold) also have "fps_shift n \ = fps_of_ratfps (ratfps_shift n x)" by (subst fst_y, subst fps_of_poly_mult, subst unit_div_mult_swap [symmetric]) (simp_all add: y_def fps_of_poly_monom fps_shift_times_X_power'' eq fps_of_ratfps_altdef case_prod_unfold Let_def x'_def) finally show "fps_of_ratfps (ratfps_shift n x) = fps_shift n (fps_of_ratfps x)" .. qed lemma fps_shift_code [code]: "fps_shift n (fps_of_ratfps x) = fps_of_ratfps (ratfps_shift n x)" by simp instantiation fps :: (equal) equal begin definition equal_fps :: "'a fps \ 'a fps \ bool" where [simp]: "equal_fps f g \ f = g" instance by standard simp_all end lemma equal_fps_code [code]: "HOL.equal (fps_of_ratfps f) (fps_of_ratfps g) \ f = g" by simp lemma fps_of_ratfps_divide [simp]: "fps_of_ratfps (f div g) = fps_of_ratfps f div fps_of_ratfps g" unfolding fps_divide_def Let_def by transfer' (simp add: Let_def ratfps_subdegree_altdef) lemma ratfps_eqI: "fps_of_ratfps x = fps_of_ratfps y \ x = y" by simp instance ratfps :: ("field_gcd") algebraic_semidom by standard (auto intro: ratfps_eqI) lemma fps_of_ratfps_dvd [simp]: "fps_of_ratfps x dvd fps_of_ratfps y \ x dvd y" proof assume "fps_of_ratfps x dvd fps_of_ratfps y" hence "fps_of_ratfps y = fps_of_ratfps y div fps_of_ratfps x * fps_of_ratfps x" by simp also have "\ = fps_of_ratfps (y div x * x)" by simp finally have "y = y div x * x" by (subst (asm) fps_of_ratfps_eq_iff) thus "x dvd y" by (intro dvdI[of _ _ "y div x"]) (simp add: mult_ac) next assume "x dvd y" hence "y = y div x * x" by simp also have "fps_of_ratfps \ = fps_of_ratfps (y div x) * fps_of_ratfps x" by simp finally show "fps_of_ratfps x dvd fps_of_ratfps y" by (simp del: fps_of_ratfps_divide) qed lemma is_unit_ratfps_iff [simp]: "is_unit x \ ratfps_nth x 0 \ 0" proof assume "is_unit x" then obtain y where "1 = x * y" by (auto elim!: dvdE) hence "1 = fps_of_ratfps (x * y)" by (simp del: fps_of_ratfps_mult) also have "\ = fps_of_ratfps x * fps_of_ratfps y" by simp finally have "is_unit (fps_of_ratfps x)" by (rule dvdI[of _ _ "fps_of_ratfps y"]) thus "ratfps_nth x 0 \ 0" by (simp add: ratfps_nth_altdef) next assume "ratfps_nth x 0 \ 0" hence "fps_of_ratfps (x * inverse x) = 1" by (simp add: ratfps_nth_altdef inverse_mult_eq_1') also have "\ = fps_of_ratfps 1" by simp finally have "x * inverse x = 1" by (subst (asm) fps_of_ratfps_eq_iff) thus "is_unit x" by (intro dvdI[of _ _ "inverse x"]) simp_all qed instantiation ratfps :: ("field_gcd") normalization_semidom begin definition unit_factor_ratfps :: "'a ratfps \ 'a ratfps" where "unit_factor x = ratfps_shift (ratfps_subdegree x) x" definition normalize_ratfps :: "'a ratfps \ 'a ratfps" where "normalize x = (if x = 0 then 0 else ratfps_of_poly (monom 1 (ratfps_subdegree x)))" lemma fps_of_ratfps_unit_factor [simp]: "fps_of_ratfps (unit_factor x) = unit_factor (fps_of_ratfps x)" unfolding unit_factor_ratfps_def by (simp add: ratfps_subdegree_altdef) lemma fps_of_ratfps_normalize [simp]: "fps_of_ratfps (normalize x) = normalize (fps_of_ratfps x)" unfolding normalize_ratfps_def by (simp add: fps_of_poly_monom ratfps_subdegree_altdef) instance proof show "unit_factor x * normalize x = x" "normalize (0 :: 'a ratfps) = 0" "unit_factor (0 :: 'a ratfps) = 0" for x :: "'a ratfps" by (rule ratfps_eqI, simp add: ratfps_subdegree_code del: fps_of_ratfps_eq_iff fps_unit_factor_def fps_normalize_def)+ show "is_unit (unit_factor a)" if "a \ 0" for a :: "'a ratfps" using that by (auto simp: ratfps_nth_altdef) fix a b :: "'a ratfps" assume "is_unit a" thus "unit_factor (a * b) = a * unit_factor b" by (intro ratfps_eqI, unfold fps_of_ratfps_unit_factor fps_of_ratfps_mult, subst unit_factor_mult_unit_left) (auto simp: ratfps_nth_altdef) show "unit_factor a = a" if "is_unit a" for a :: "'a ratfps" by (rule ratfps_eqI) (insert that, auto simp: fps_of_ratfps_is_unit) qed end instance ratfps :: ("field_gcd") normalization_semidom_multiplicative proof show "unit_factor (a * b) = unit_factor a * unit_factor b" for a b :: "'a ratfps" by (rule ratfps_eqI, insert unit_factor_mult[of "fps_of_ratfps a" "fps_of_ratfps b"]) (simp del: fps_of_ratfps_eq_iff) qed instantiation ratfps :: ("field_gcd") semidom_modulo begin lift_definition modulo_ratfps :: "'a ratfps \ 'a ratfps \ 'a ratfps" is "\f g. if g = 0 then f else let n = ratfps_subdegree g; h = ratfps_shift n g in ratfps_of_poly (ratfps_cutoff n (f * inverse h)) * h" . lemma fps_of_ratfps_mod [simp]: "fps_of_ratfps (f mod g :: 'a ratfps) = fps_of_ratfps f mod fps_of_ratfps g" unfolding fps_mod_def by transfer' (simp add: Let_def ratfps_subdegree_altdef) instance by standard (auto intro: ratfps_eqI) end instantiation ratfps :: ("field_gcd") euclidean_ring begin definition euclidean_size_ratfps :: "'a ratfps \ nat" where "euclidean_size_ratfps x = (if x = 0 then 0 else 2 ^ ratfps_subdegree x)" lemma fps_of_ratfps_euclidean_size [simp]: "euclidean_size x = euclidean_size (fps_of_ratfps x)" unfolding euclidean_size_ratfps_def fps_euclidean_size_def by (simp add: ratfps_subdegree_altdef) instance proof show "euclidean_size (0 :: 'a ratfps) = 0" by simp show "euclidean_size (a mod b) < euclidean_size b" "euclidean_size a \ euclidean_size (a * b)" if "b \ 0" for a b :: "'a ratfps" using that by (simp_all add: mod_size_less size_mult_mono) qed end instantiation ratfps :: ("field_gcd") euclidean_ring_cancel begin instance by standard (auto intro: ratfps_eqI) end lemma quot_of_ratfps_eq_iff [simp]: "quot_of_ratfps x = quot_of_ratfps y \ x = y" by transfer simp lemma ratfps_eq_0_code: "x = 0 \ fst (quot_of_ratfps x) = 0" proof assume "fst (quot_of_ratfps x) = 0" moreover have "coprime (fst (quot_of_ratfps x)) (snd (quot_of_ratfps x))" by transfer (simp add: coprime_quot_of_fract) moreover have "normalize (snd (quot_of_ratfps x)) = snd (quot_of_ratfps x)" by (simp add: div_unit_factor [symmetric] del: div_unit_factor) ultimately have "quot_of_ratfps x = (0,1)" by (simp add: prod_eq_iff normalize_idem_imp_is_unit_iff) also have "\ = quot_of_ratfps 0" by simp finally show "x = 0" by (subst (asm) quot_of_ratfps_eq_iff) qed simp_all lemma fps_dvd_code [code_unfold]: "x dvd y \ y = 0 \ ((x::'a::field_gcd fps) \ 0 \ subdegree x \ subdegree y)" using fps_dvd_iff[of x y] by (cases "x = 0") auto lemma ratfps_dvd_code [code_unfold]: "x dvd y \ y = 0 \ (x \ 0 \ ratfps_subdegree x \ ratfps_subdegree y)" using fps_dvd_code [of "fps_of_ratfps x" "fps_of_ratfps y"] by (simp add: ratfps_subdegree_altdef) instance ratfps :: ("field_gcd") normalization_euclidean_semiring .. instantiation ratfps :: ("field_gcd") euclidean_ring_gcd begin definition "gcd_ratfps = (Euclidean_Algorithm.gcd :: 'a ratfps \ _)" definition "lcm_ratfps = (Euclidean_Algorithm.lcm :: 'a ratfps \ _)" definition "Gcd_ratfps = (Euclidean_Algorithm.Gcd :: 'a ratfps set \ _)" definition "Lcm_ratfps = (Euclidean_Algorithm.Lcm:: 'a ratfps set \ _)" instance by standard (simp_all add: gcd_ratfps_def lcm_ratfps_def Gcd_ratfps_def Lcm_ratfps_def) end lemma ratfps_eq_0_iff: "x = 0 \ fps_of_ratfps x = 0" using fps_of_ratfps_eq_iff[of x 0] unfolding fps_of_ratfps_0 by simp lemma ratfps_of_poly_eq_0_iff: "ratfps_of_poly x = 0 \ x = 0" by (auto simp: ratfps_eq_0_iff) lemma ratfps_gcd: assumes [simp]: "f \ 0" "g \ 0" shows "gcd f g = ratfps_of_poly (monom 1 (min (ratfps_subdegree f) (ratfps_subdegree g)))" by (rule sym, rule gcdI) (auto simp: ratfps_subdegree_altdef ratfps_dvd_code subdegree_fps_of_poly ratfps_of_poly_eq_0_iff normalize_ratfps_def) lemma ratfps_gcd_altdef: "gcd (f :: 'a :: field_gcd ratfps) g = (if f = 0 \ g = 0 then 0 else if f = 0 then ratfps_of_poly (monom 1 (ratfps_subdegree g)) else if g = 0 then ratfps_of_poly (monom 1 (ratfps_subdegree f)) else ratfps_of_poly (monom 1 (min (ratfps_subdegree f) (ratfps_subdegree g))))" by (simp add: ratfps_gcd normalize_ratfps_def) lemma ratfps_lcm: assumes [simp]: "f \ 0" "g \ 0" shows "lcm f g = ratfps_of_poly (monom 1 (max (ratfps_subdegree f) (ratfps_subdegree g)))" by (rule sym, rule lcmI) (auto simp: ratfps_subdegree_altdef ratfps_dvd_code subdegree_fps_of_poly ratfps_of_poly_eq_0_iff normalize_ratfps_def) lemma ratfps_lcm_altdef: "lcm (f :: 'a :: field_gcd ratfps) g = (if f = 0 \ g = 0 then 0 else ratfps_of_poly (monom 1 (max (ratfps_subdegree f) (ratfps_subdegree g))))" by (simp add: ratfps_lcm) lemma ratfps_Gcd: assumes "A - {0} \ {}" shows "Gcd A = ratfps_of_poly (monom 1 (INF f\A-{0}. ratfps_subdegree f))" proof (rule sym, rule GcdI) fix f assume "f \ A" thus "ratfps_of_poly (monom 1 (INF f\A - {0}. ratfps_subdegree f)) dvd f" by (cases "f = 0") (auto simp: ratfps_dvd_code ratfps_of_poly_eq_0_iff ratfps_subdegree_altdef subdegree_fps_of_poly intro!: cINF_lower) next fix d assume d: "\f. f \ A \ d dvd f" from assms obtain f where "f \ A - {0}" by auto with d[of f] have [simp]: "d \ 0" by auto from d assms have "ratfps_subdegree d \ (INF f\A-{0}. ratfps_subdegree f)" by (intro cINF_greatest) (auto simp: ratfps_dvd_code) with d assms show "d dvd ratfps_of_poly (monom 1 (INF f\A-{0}. ratfps_subdegree f))" by (simp add: ratfps_dvd_code ratfps_subdegree_altdef subdegree_fps_of_poly) qed (simp_all add: ratfps_subdegree_altdef subdegree_fps_of_poly normalize_ratfps_def) lemma ratfps_Gcd_altdef: "Gcd (A :: 'a :: field_gcd ratfps set) = (if A \ {0} then 0 else ratfps_of_poly (monom 1 (INF f\A-{0}. ratfps_subdegree f)))" using ratfps_Gcd by auto lemma ratfps_Lcm: assumes "A \ {}" "0 \ A" "bdd_above (ratfps_subdegree`A)" shows "Lcm A = ratfps_of_poly (monom 1 (SUP f\A. ratfps_subdegree f))" proof (rule sym, rule LcmI) fix f assume "f \ A" moreover from assms(3) have "bdd_above (ratfps_subdegree ` A)" by auto ultimately show "f dvd ratfps_of_poly (monom 1 (SUP f\A. ratfps_subdegree f))" using assms(2) by (cases "f = 0") (auto simp: ratfps_dvd_code ratfps_of_poly_eq_0_iff subdegree_fps_of_poly ratfps_subdegree_altdef [abs_def] intro!: cSUP_upper) next fix d assume d: "\f. f \ A \ f dvd d" from assms obtain f where f: "f \ A" "f \ 0" by auto show "ratfps_of_poly (monom 1 (SUP f\A. ratfps_subdegree f)) dvd d" proof (cases "d = 0") assume "d \ 0" moreover from d have "\f. f \ A \ f \ 0 \ f dvd d" by blast ultimately have "ratfps_subdegree d \ (SUP f\A. ratfps_subdegree f)" using assms by (intro cSUP_least) (auto simp: ratfps_dvd_code) with \d \ 0\ show ?thesis by (simp add: ratfps_dvd_code ratfps_of_poly_eq_0_iff ratfps_subdegree_altdef subdegree_fps_of_poly) qed simp_all qed (simp_all add: ratfps_subdegree_altdef subdegree_fps_of_poly normalize_ratfps_def) lemma ratfps_Lcm_altdef: "Lcm (A :: 'a :: field_gcd ratfps set) = (if 0 \ A \ \bdd_above (ratfps_subdegree`A) then 0 else if A = {} then 1 else ratfps_of_poly (monom 1 (SUP f\A. ratfps_subdegree f)))" proof (cases "bdd_above (ratfps_subdegree`A)") assume unbounded: "\bdd_above (ratfps_subdegree`A)" have "Lcm A = 0" proof (rule ccontr) assume "Lcm A \ 0" from unbounded obtain f where f: "f \ A" "ratfps_subdegree (Lcm A) < ratfps_subdegree f" unfolding bdd_above_def by (auto simp: not_le) moreover from this and \Lcm A \ 0\ have "ratfps_subdegree f \ ratfps_subdegree (Lcm A)" using dvd_Lcm[of f A] by (auto simp: ratfps_dvd_code) ultimately show False by simp qed with unbounded show ?thesis by simp qed (simp_all add: ratfps_Lcm Lcm_eq_0_I) lemma fps_of_ratfps_quot_to_ratfps: "coeff y 0 \ 0 \ fps_of_ratfps (quot_to_ratfps (x,y)) = fps_of_poly x / fps_of_poly y" proof (transfer, goal_cases) case (1 y x) define x' y' where "x' = fst (normalize_quot (x,y))" and "y' = snd (normalize_quot (x,y))" from 1 have nz: "y \ 0" by auto have eq: "normalize_quot (x', y') = (x', y')" by (simp add: x'_def y'_def) - from normalize_quotE[OF nz, of x] guess d . note d = this [folded x'_def y'_def] + from normalize_quotE[OF nz, of x] obtain d where + "x = fst (normalize_quot (x, y)) * d" + "y = snd (normalize_quot (x, y)) * d" + "d dvd x" + "d dvd y" + "d \ 0" . + note d [folded x'_def y'_def] = this have "(case quot_of_fract (if coeff y' 0 = 0 then 0 else quot_to_fract (x', y')) of (a, b) \ fps_of_poly a / fps_of_poly b) = fps_of_poly x / fps_of_poly y" using d eq 1 by (auto simp: case_prod_unfold fps_of_poly_simps quot_of_fract_quot_to_fract Let_def coeff_0_mult) thus ?case by (auto simp add: Let_def case_prod_unfold x'_def y'_def) qed lemma fps_of_ratfps_quot_to_ratfps_code_post1: "fps_of_ratfps (quot_to_ratfps (x,pCons 1 y)) = fps_of_poly x / fps_of_poly (pCons 1 y)" "fps_of_ratfps (quot_to_ratfps (x,pCons (-1) y)) = fps_of_poly x / fps_of_poly (pCons (-1) y)" by (simp_all add: fps_of_ratfps_quot_to_ratfps) lemma fps_of_ratfps_quot_to_ratfps_code_post2: "fps_of_ratfps (quot_to_ratfps (x'::'a::{field_char_0,field_gcd} poly,pCons (numeral n) y')) = fps_of_poly x' / fps_of_poly (pCons (numeral n) y')" "fps_of_ratfps (quot_to_ratfps (x'::'a::{field_char_0,field_gcd} poly,pCons (-numeral n) y')) = fps_of_poly x' / fps_of_poly (pCons (-numeral n) y')" by (simp_all add: fps_of_ratfps_quot_to_ratfps) lemmas fps_of_ratfps_quot_to_ratfps_code_post [code_post] = fps_of_ratfps_quot_to_ratfps_code_post1 fps_of_ratfps_quot_to_ratfps_code_post2 lemma fps_dehorner: fixes a b c :: "'a :: semiring_1 fps" and d e f :: "'b :: ring_1 fps" shows "(b + c) * fps_X = b * fps_X + c * fps_X" "(a * fps_X) * fps_X = a * fps_X ^ 2" "a * fps_X ^ m * fps_X = a * fps_X ^ (Suc m)" "a * fps_X * fps_X ^ m = a * fps_X ^ (Suc m)" "a * fps_X^m * fps_X^n = a * fps_X^(m+n)" "a + (b + c) = a + b + c" "a * 1 = a" "1 * a = a" "d + - e = d - e" "(-d) * e = - (d * e)" "d + (e - f) = d + e - f" "(d - e) * fps_X = d * fps_X - e * fps_X" "fps_X * fps_X = fps_X^2" "fps_X * fps_X^m = fps_X^(Suc m)" "fps_X^m * fps_X = fps_X^Suc m" "fps_X^m * fps_X^n = fps_X^(m + n)" by (simp_all add: algebra_simps power2_eq_square power_add power_commutes) lemma fps_divide_1: "(a :: 'a :: field fps) / 1 = a" by simp lemmas fps_of_poly_code_post [code_post] = fps_of_poly_simps fps_const_0_eq_0 fps_const_1_eq_1 numeral_fps_const [symmetric] fps_const_neg [symmetric] fps_const_divide [symmetric] fps_dehorner Suc_numeral arith_simps fps_divide_1 context includes term_syntax begin definition valterm_ratfps :: "'a ::{field_gcd, typerep} poly \ (unit \ Code_Evaluation.term) \ 'a poly \ (unit \ Code_Evaluation.term) \ 'a ratfps \ (unit \ Code_Evaluation.term)" where [code_unfold]: "valterm_ratfps k l = Code_Evaluation.valtermify (/) {\} (Code_Evaluation.valtermify ratfps_of_poly {\} k) {\} (Code_Evaluation.valtermify ratfps_of_poly {\} l)" end instantiation ratfps :: ("{field_gcd,random}") random begin context includes state_combinator_syntax term_syntax begin definition "Quickcheck_Random.random i = Quickcheck_Random.random i \\ (\num::'a poly \ (unit \ term). Quickcheck_Random.random i \\ (\denom::'a poly \ (unit \ term). Pair (let denom = (if fst denom = 0 then Code_Evaluation.valtermify 1 else denom) in valterm_ratfps num denom)))" instance .. end end instantiation ratfps :: ("{field,factorial_ring_gcd,exhaustive}") exhaustive begin definition "exhaustive_ratfps f d = Quickcheck_Exhaustive.exhaustive (\num. Quickcheck_Exhaustive.exhaustive (\denom. f ( let denom = if denom = 0 then 1 else denom in ratfps_of_poly num / ratfps_of_poly denom)) d) d" instance .. end instantiation ratfps :: ("{field_gcd,full_exhaustive}") full_exhaustive begin definition "full_exhaustive_ratfps f d = Quickcheck_Exhaustive.full_exhaustive (\num::'a poly \ (unit \ term). Quickcheck_Exhaustive.full_exhaustive (\denom::'a poly \ (unit \ term). f (let denom = if fst denom = 0 then Code_Evaluation.valtermify 1 else denom in valterm_ratfps num denom)) d) d" instance .. end quickcheck_generator fps constructors: fps_of_ratfps end diff --git a/thys/Linear_Recurrences/Rational_FPS_Asymptotics.thy b/thys/Linear_Recurrences/Rational_FPS_Asymptotics.thy --- a/thys/Linear_Recurrences/Rational_FPS_Asymptotics.thy +++ b/thys/Linear_Recurrences/Rational_FPS_Asymptotics.thy @@ -1,398 +1,404 @@ (* File: Rational_FPS_Asymptotics.thy Author: Manuel Eberl, TU München *) theory Rational_FPS_Asymptotics imports "HOL-Library.Landau_Symbols" "Polynomial_Factorization.Square_Free_Factorization" "HOL-Real_Asymp.Real_Asymp" "Count_Complex_Roots.Count_Complex_Roots" Linear_Homogenous_Recurrences Linear_Inhomogenous_Recurrences RatFPS Rational_FPS_Solver "HOL-Library.Code_Target_Numeral" + begin lemma poly_asymp_equiv: assumes "p \ 0" and "F \ at_infinity" shows "poly p \[F] (\x. lead_coeff p * x ^ degree p)" proof - have poly_pCons': "poly (pCons a q) = (\x. a + x * poly q x)" for a :: 'a and q by (simp add: fun_eq_iff) show ?thesis using assms(1) proof (induction p) case (pCons a p) define n where "n = Suc (degree p)" show ?case proof (cases "p = 0") case [simp]: False hence *: "poly p \[F] (\x. lead_coeff p * x ^ degree p)" by (intro pCons.IH) have "poly (pCons a p) = (\x. a + x * poly p x)" by (simp add: poly_pCons') moreover have "\ \[F] (\x. lead_coeff p * x ^ n)" proof (subst asymp_equiv_add_left) have "(\x. x * poly p x) \[F] (\x. x * (lead_coeff p * x ^ degree p))" by (intro asymp_equiv_intros *) also have "\ = (\x. lead_coeff p * x ^ n)" by (simp add: n_def mult_ac) finally show "(\x. x * poly p x) \[F] \" . next have "filterlim (\x. x) at_infinity F" by (simp add: filterlim_def assms) hence "(\x. x ^ n) \ \[F](\_. 1 :: 'a)" unfolding smallomega_1_conv_filterlim by (intro Limits.filterlim_power_at_infinity filterlim_ident) (auto simp: n_def) hence "(\x. a) \ o[F](\x. x ^ n)" unfolding smallomega_iff_smallo[symmetric] by (cases "a = 0") auto thus "(\x. a) \ o[F](\x. lead_coeff p * x ^ n)" by simp qed ultimately show ?thesis by (simp add: n_def) qed auto qed auto qed lemma poly_bigtheta: assumes "p \ 0" and "F \ at_infinity" shows "poly p \ \[F](\x. x ^ degree p)" proof - have "poly p \[F] (\x. lead_coeff p * x ^ degree p)" by (intro poly_asymp_equiv assms) thus ?thesis using assms by (auto dest!: asymp_equiv_imp_bigtheta) qed lemma poly_bigo: assumes "F \ at_infinity" and "degree p \ k" shows "poly p \ O[F](\x. x ^ k)" proof (cases "p = 0") case True hence "poly p = (\_. 0)" by (auto simp: fun_eq_iff) thus ?thesis by simp next case False have *: "(\x. x ^ (k - degree p)) \ \[F](\x. 1)" proof (cases "k = degree p") case False hence "(\x. x ^ (k - degree p)) \ \[F](\_. 1)" unfolding smallomega_1_conv_filterlim using assms False by (intro Limits.filterlim_power_at_infinity filterlim_ident) (auto simp: filterlim_def) thus ?thesis by (rule landau_omega.small_imp_big) qed auto have "poly p \ \[F](\x. x ^ degree p * 1)" using poly_bigtheta[OF False assms(1)] by simp also have "(\x. x ^ degree p * 1) \ O[F](\x. x ^ degree p * x ^ (k - degree p))" using * by (intro landau_o.big.mult landau_o.big_refl) (auto simp: bigomega_iff_bigo) also have "(\x::'a. x ^ degree p * x ^ (k - degree p)) = (\x. x ^ k)" using assms by (simp add: power_add [symmetric]) finally show ?thesis . qed lemma reflect_poly_dvdI: fixes p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "p dvd q" shows "reflect_poly p dvd reflect_poly q" using assms by (auto simp: reflect_poly_mult) lemma smult_altdef: "smult c p = [:c:] * p" by (induction p) (auto simp: mult_ac) lemma smult_power: "smult (c ^ n) (p ^ n) = (smult c p) ^ n" proof - have "smult (c ^ n) (p ^ n) = [:c ^ n:] * p ^ n" by simp also have "[:c:] ^ n = [:c ^ n:]" by (induction n) (auto simp: mult_ac) hence "[:c ^ n:] = [:c:] ^ n" .. also have "\ * p ^ n = ([:c:] * p) ^ n" by (rule power_mult_distrib [symmetric]) also have "\ = (smult c p) ^ n" by simp finally show ?thesis . qed lemma order_reflect_poly_ge: fixes c :: "'a :: field" assumes "c \ 0" and "p \ 0" shows "order c (reflect_poly p) \ order (1 / c) p" proof - have "reflect_poly ([:-(1 / c), 1:] ^ order (1 / c) p) dvd reflect_poly p" by (intro reflect_poly_dvdI, subst order_divides) auto also have "reflect_poly ([:-(1 / c), 1:] ^ order (1 / c) p) = smult ((-1 / c) ^ order (1 / c) p) ([:-c, 1:] ^ order (1 / c) p)" using assms by (simp add: reflect_poly_power reflect_poly_pCons smult_power) finally have "([:-c, 1:] ^ order (1 / c) p) dvd reflect_poly p" by (rule smult_dvd_cancel) with \p \ 0\ show ?thesis by (subst (asm) order_divides) auto qed lemma order_reflect_poly: fixes c :: "'a :: field" assumes "c \ 0" and "coeff p 0 \ 0" shows "order c (reflect_poly p) = order (1 / c) p" proof (rule antisym) from assms show "order c (reflect_poly p) \ order (1 / c) p" by (intro order_reflect_poly_ge) auto next from assms have "order (1 / (1 / c)) (reflect_poly p) \ order (1 / c) (reflect_poly (reflect_poly p))" by (intro order_reflect_poly_ge) auto with assms show "order c (reflect_poly p) \ order (1 / c) p" by simp qed lemma poly_reflect_eq_0_iff: "poly (reflect_poly p) (x :: 'a :: field) = 0 \ p = 0 \ x \ 0 \ poly p (1 / x) = 0" by (cases "x = 0") (auto simp: poly_reflect_poly_nz inverse_eq_divide) theorem ratfps_nth_bigo: fixes q :: "complex poly" assumes "R > 0" assumes roots1: "\z. z \ ball 0 (1 / R) \ poly q z \ 0" assumes roots2: "\z. z \ sphere 0 (1 / R) \ poly q z = 0 \ order z q \ Suc k" shows "fps_nth (fps_of_poly p / fps_of_poly q) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof - define q' where "q' = reflect_poly q" from roots1[of 0] and \R > 0\ have [simp]: "coeff q 0 \ 0" "q \ 0" by (auto simp: poly_0_coeff_0) - from ratfps_closed_form_exists[OF this(1), of p] guess r rs . note closed_form = this + from ratfps_closed_form_exists[OF this(1), of p] + obtain r rs where closed_form: + "\n. (fps_of_poly p / fps_of_poly q) $ n = + coeff r n + (\c | poly (reflect_poly q) c = 0. poly (rs c) (of_nat n) * c ^ n)" + "\z. poly (reflect_poly q) z = 0 \ degree (rs z) \ order z (reflect_poly q) - 1" + by blast have "fps_nth (fps_of_poly p / fps_of_poly q) = (\n. coeff r n + (\c | poly q' c = 0. poly (rs c) (of_nat n) * c ^ n))" by (intro ext, subst closed_form) (simp_all add: q'_def) also have "\ \ O(\n. of_nat n ^ k * of_real R ^ n)" proof (intro sum_in_bigo big_sum_in_bigo) have "eventually (\n. coeff r n = 0) at_top" using MOST_nat coeff_eq_0 cofinite_eq_sequentially by force hence "coeff r \ \(\_. 0)" by (rule bigthetaI_cong) also have "(\_. 0 :: complex) \ O(\n. of_nat n ^ k * of_real R ^ n)" by simp finally show "coeff r \ O(\n. of_nat n ^ k * of_real R ^ n)" . next fix c assume c: "c \ {c. poly q' c = 0}" hence [simp]: "c \ 0" by (auto simp: q'_def) show "(\n. poly (rs c) n * c ^ n) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof (cases "norm c = R") case True \\The case of a root at the border of the disc\ show ?thesis proof (intro landau_o.big.mult landau_o.big.compose[OF poly_bigo tendsto_of_nat]) have "degree (rs c) \ order c (reflect_poly q) - 1" using c by (intro closed_form(2)) (auto simp: q'_def) also have "order c (reflect_poly q) = order (1 / c) q" using c by (intro order_reflect_poly) (auto simp: q'_def) also { have "order (1 / c) q \ Suc k" using \R > 0\ and True and c by (intro roots2) (auto simp: q'_def norm_divide poly_reflect_eq_0_iff) moreover have "order (1 / c) q \ 0" using order_root[of q "1 / c"] c by (auto simp: q'_def poly_reflect_eq_0_iff) ultimately have "order (1 / c) q - 1 \ k" by simp } finally show "degree (rs c) \ k" . next have "(\n. norm (c ^ n)) \ O(\n. norm (complex_of_real R ^ n))" using True and \R > 0\ by (simp add: norm_power) thus "(\n. c ^ n) \ O(\n. complex_of_real R ^ n)" by (subst (asm) landau_o.big.norm_iff) qed auto next case False \\The case of a root in the interior of the disc\ hence "norm c < R" using c and roots1[of "1/c"] and \R > 0\ by (cases "norm c" R rule: linorder_cases) (auto simp: q'_def poly_reflect_eq_0_iff norm_divide field_simps) define l where "l = degree (rs c)" have "(\n. poly (rs c) (of_nat n) * c ^ n) \ O(\n. of_nat n ^ l * c ^ n)" by (intro landau_o.big.mult landau_o.big.compose[OF poly_bigo tendsto_of_nat]) (auto simp: l_def) also have "(\n. of_nat n ^ l * c ^ n) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof (subst landau_o.big.norm_iff [symmetric]) have "(\n. real n ^ l) \ O(\n. real n ^ k * (R / norm c) ^ n)" using \norm c < R\ and \R > 0\ by real_asymp hence "(\n. real n ^ l * norm c ^ n) \ O(\n. real n ^ k * R ^ n)" by (simp add: power_divide landau_o.big.divide_eq1) thus "(\x. norm (of_nat x ^ l * c ^ x)) \ O(\x. norm (of_nat x ^ k * complex_of_real R ^ x))" unfolding norm_power norm_mult using \R > 0\ by simp qed finally show ?thesis . qed qed finally show ?thesis . qed lemma order_power: "p \ 0 \ order c (p ^ n) = n * order c p" by (induction n) (auto simp: order_mult) lemma same_root_imp_not_coprime: assumes "poly p x = 0" and "poly q (x :: 'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize}) = 0" shows "\coprime p q" proof assume "coprime p q" from assms have "[:-x, 1:] dvd p" and "[:-x, 1:] dvd q" by (simp_all add: poly_eq_0_iff_dvd) hence "[:-x, 1:] dvd gcd p q" by (simp add: poly_eq_0_iff_dvd) also from \coprime p q\ have "gcd p q = 1" by (rule coprime_imp_gcd_eq_1) finally show False by (elim is_unit_polyE) auto qed lemma ratfps_nth_bigo_square_free_factorization: fixes p :: "complex poly" assumes "square_free_factorization q (b, cs)" assumes "q \ 0" and "R > 0" assumes roots1: "\c l. (c, l) \ set cs \ \x\ball 0 (1 / R). poly c x \ 0" assumes roots2: "\c l. (c, l) \ set cs \ l > k \ \x\sphere 0 (1 / R). poly c x \ 0" shows "fps_nth (fps_of_poly p / fps_of_poly q) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof - from assms(1) have q: "q = smult b (\(c, l)\set cs. c ^ Suc l)" unfolding square_free_factorization_def prod.case by blast with \q \ 0\ have [simp]: "b \ 0" by auto from assms(1) have [simp]: "(0, x) \ set cs" for x by (auto simp: square_free_factorization_def) from assms(1) have coprime: "c1 = c2" "m = n" if "\coprime c1 c2" "(c1, m) \ set cs" "(c2, n) \ set cs" for c1 c2 m n using that by (auto simp: square_free_factorization_def case_prod_unfold) show ?thesis proof (rule ratfps_nth_bigo) fix z :: complex assume z: "z \ ball 0 (1 / R)" show "poly q z \ 0" proof assume "poly q z = 0" then obtain c l where cl: "(c, l) \ set cs" and "poly c z = 0" by (auto simp: q poly_prod image_iff) with roots1[of c l] and z show False by auto qed next fix z :: complex assume z: "z \ sphere 0 (1 / R)" have order: "order z q = order z (\(c, l)\set cs. c ^ Suc l)" by (simp add: order_smult q) also have "\ = (\x\set cs. order z (case x of (c, l) \ c ^ Suc l))" by (subst order_prod) (auto dest: coprime) also have "\ = (\(c, l)\set cs. Suc l * order z c)" unfolding case_prod_unfold by (intro sum.cong refl, subst order_power) auto finally have "order z q = \" . show "order z q \ Suc k" proof (cases "\c0 l0. (c0, l0) \ set cs \ poly c0 z = 0") case False have "order z q = (\(c, l)\set cs. Suc l * order z c)" by fact also have "order z c = 0" if "(c, l) \ set cs" for c l using False that by (auto simp: order_root) hence "(\(c, l)\set cs. Suc l * order z c) = 0" by (intro sum.neutral) auto finally show "order z q \ Suc k" by simp next case True \\The order of a root is determined by the unique polynomial in the square-free factorisation that contains it.\ then obtain c0 l0 where cl0: "(c0, l0) \ set cs" "poly c0 z = 0" by blast have "order z q = (\(c, l)\set cs. Suc l * order z c)" by fact also have "\ = Suc l0 * order z c0 + (\(c, l) \ set cs - {(c0, l0)}. Suc l * order z c)" using cl0 by (subst sum.remove[of _ "(c0, l0)"]) auto also have "(\(c, l) \ set cs - {(c0, l0)}. Suc l * order z c) = 0" proof (intro sum.neutral ballI, goal_cases) case (1 cl) then obtain c l where [simp]: "cl = (c, l)" and cl: "(c, l) \ set cs" "(c0, l0) \ (c, l)" by (cases cl) auto from cl and cl0 and coprime[of c c0 l l0] have "coprime c c0" by auto with same_root_imp_not_coprime[of c z c0] and cl0 have "poly c z \ 0" by auto thus ?case by (auto simp: order_root) qed also have "square_free c0" using cl0 assms(1) by (auto simp: square_free_factorization_def) hence "rsquarefree c0" by (rule square_free_rsquarefree) with cl0 have "order z c0 = 1" by (auto simp: rsquarefree_def' order_root intro: antisym) finally have "order z q = Suc l0" by simp also from roots2[of c0 l0] cl0 z have "l0 \ k" by (cases l0 k rule: linorder_cases) auto finally show "order z q \ Suc k" by simp qed qed fact+ qed find_consts name:"Count_Complex" term proots_ball_card term proots_sphere_card lemma proots_within_card_zero_iff: assumes "p \ (0 :: 'a :: idom poly)" shows "card (proots_within p A) = 0 \ (\x\A. poly p x \ 0)" using assms by (subst card_0_eq) (auto intro: finite_proots) lemma ratfps_nth_bigo_square_free_factorization': fixes p :: "complex poly" assumes "square_free_factorization q (b, cs)" assumes "q \ 0" and "R > 0" assumes roots1: "list_all (\cl. proots_ball_card (fst cl) 0 (1 / R) = 0) cs" assumes roots2: "list_all (\cl. proots_sphere_card (fst cl) 0 (1 / R) = 0) (filter (\cl. snd cl > k) cs)" shows "fps_nth (fps_of_poly p / fps_of_poly q) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof (rule ratfps_nth_bigo_square_free_factorization[OF assms(1)]) from assms(1) have q: "q = smult b (\(c, l)\set cs. c ^ Suc l)" unfolding square_free_factorization_def prod.case by blast with \q \ 0\ have [simp]: "b \ 0" by auto from assms(1) have [simp]: "(0, x) \ set cs" for x by (auto simp: square_free_factorization_def) show "\x\ball 0 (1 / R). poly c x \ 0" if "(c, l) \ set cs" for c l proof - from roots1 that have "card (proots_within c (ball 0 (1 / R))) = 0" by (auto simp: proots_ball_card_def list_all_def) with that show ?thesis by (subst (asm) proots_within_card_zero_iff) auto qed show "\x\sphere 0 (1 / R). poly c x \ 0" if "(c, l) \ set cs" "l > k" for c l proof - from roots2 that have "card (proots_within c (sphere 0 (1 / R))) = 0" by (auto simp: proots_sphere_card_def list_all_def) with that show ?thesis by (subst (asm) proots_within_card_zero_iff) auto qed qed fact+ definition ratfps_has_asymptotics where "ratfps_has_asymptotics q k R \ q \ 0 \ R > 0 \ (let cs = snd (yun_factorization gcd q) in list_all (\cl. proots_ball_card (fst cl) 0 (1 / R) = 0) cs \ list_all (\cl. proots_sphere_card (fst cl) 0 (1 / R) = 0) (filter (\cl. snd cl > k) cs))" lemma ratfps_has_asymptotics_correct: assumes "ratfps_has_asymptotics q k R" shows "fps_nth (fps_of_poly p / fps_of_poly q) \ O(\n. of_nat n ^ k * of_real R ^ n)" proof (rule ratfps_nth_bigo_square_free_factorization') show "square_free_factorization q (fst (yun_factorization gcd q), snd (yun_factorization gcd q))" by (rule yun_factorization) simp qed (insert assms, auto simp: ratfps_has_asymptotics_def Let_def list_all_def) value "map (fps_nth (fps_of_poly [:0, 1:] / fps_of_poly [:1, -1, -1 :: real:])) [0..<5]" method ratfps_bigo = (rule ratfps_has_asymptotics_correct; eval) lemma "fps_nth (fps_of_poly [:0, 1:] / fps_of_poly [:1, -1, -1 :: complex:]) \ O(\n. of_nat n ^ 0 * complex_of_real 1.618034 ^ n)" by ratfps_bigo lemma "fps_nth (fps_of_poly 1 / fps_of_poly [:1, -3, 3, -1 :: complex:]) \ O(\n. of_nat n ^ 2 * complex_of_real 1 ^ n)" by ratfps_bigo lemma "fps_nth (fps_of_poly f / fps_of_poly [:5, 4, 3, 2, 1 :: complex:]) \ O(\n. of_nat n ^ 0 * complex_of_real 0.69202 ^ n)" by ratfps_bigo end diff --git a/thys/Liouville_Numbers/Liouville_Numbers.thy b/thys/Liouville_Numbers/Liouville_Numbers.thy --- a/thys/Liouville_Numbers/Liouville_Numbers.thy +++ b/thys/Liouville_Numbers/Liouville_Numbers.thy @@ -1,384 +1,388 @@ (* File: Liouville_Numbers.thy Author: Manuel Eberl The definition of Liouville numbers and their standard construction, plus the proof that any Liouville number is transcendental. *) theory Liouville_Numbers imports Complex_Main "HOL-Computational_Algebra.Polynomial" Liouville_Numbers_Misc begin (* TODO: Move definition of algebraic numbers out of Algebraic_Numbers to reduce unnecessary dependencies. *) text \ A Liouville number is a real number that can be approximated well -- but not perfectly -- by a sequence of rational numbers. ``Well``, in this context, means that the error of the $n$-th rational in the sequence is bounded by the $n$-th power of its denominator. Our approach will be the following: \begin{itemize} \item Liouville numbers cannot be rational. \item Any irrational algebraic number cannot be approximated in the Liouville sense \item Therefore, all Liouville numbers are transcendental. \item The standard construction fulfils all the properties of Liouville numbers. \end{itemize} \ subsection \Definition of Liouville numbers\ text \ The following definitions and proofs are largely adapted from those in the Wikipedia article on Liouville numbers.~\cite{wikipedia} \ text \ A Liouville number is a real number that can be approximated well -- but not perfectly -- by a sequence of rational numbers. The error of the $n$-th term $\frac{p_n}{q_n}$ is at most $q_n^{-n}$, where $p_n\in\isasymint$ and $q_n \in\isasymint_{\geq 2}$. We will say that such a number can be approximated in the Liouville sense. \ locale liouville = fixes x :: real and p q :: "nat \ int" assumes approx_int_pos: "abs (x - p n / q n) > 0" and denom_gt_1: "q n > 1" and approx_int: "abs (x - p n / q n) < 1 / of_int (q n) ^ n" text \ First, we show that any Liouville number is irrational. \ lemma (in liouville) irrational: "x \ \" proof assume "x \ \" then obtain c d :: int where d: "x = of_int c / of_int d" "d > 0" by (elim Rats_cases') simp define n where "n = Suc (nat \log 2 d\)" note q_gt_1 = denom_gt_1[of n] have neq: "c * q n \ d * p n" proof assume "c * q n = d * p n" hence "of_int (c * q n) = of_int (d * p n)" by (simp only: ) with approx_int_pos[of n] d q_gt_1 show False by (auto simp: field_simps) qed hence abs_pos: "0 < abs (c * q n - d * p n)" by simp from q_gt_1 d have "of_int d = 2 powr log 2 d" by (subst powr_log_cancel) simp_all also from d have "log 2 (of_int d) \ log 2 1" by (subst log_le_cancel_iff) simp_all hence "2 powr log 2 d \ 2 powr of_nat (nat \log 2 d\)" by (intro powr_mono) simp_all also have "\ = of_int (2 ^ nat \log 2 d\)" by (subst powr_realpow) simp_all finally have "d \ 2 ^ nat \log 2 (of_int d)\" by (subst (asm) of_int_le_iff) also have "\ * q n \ q n ^ Suc (nat \log 2 (of_int d)\)" by (subst power_Suc, subst mult.commute, intro mult_left_mono power_mono, insert q_gt_1) simp_all also from q_gt_1 have "\ = q n ^ n" by (simp add: n_def) finally have "1 / of_int (q n ^ n) \ 1 / real_of_int (d * q n)" using q_gt_1 d by (intro divide_left_mono Rings.mult_pos_pos of_int_pos, subst of_int_le_iff) simp_all also have "\ \ of_int (abs (c * q n - d * p n)) / real_of_int (d * q n)" using q_gt_1 d abs_pos by (intro divide_right_mono) (linarith, simp) also have "\ = abs (x - of_int (p n) / of_int (q n))" using q_gt_1 d(2) by (simp_all add: divide_simps d(1) mult_ac) finally show False using approx_int[of n] by simp qed text \ Next, any irrational algebraic number cannot be approximated with rational numbers in the Liouville sense. \ lemma liouville_irrational_algebraic: fixes x :: real assumes irrationsl: "x \ \" and "algebraic x" obtains c :: real and n :: nat where "c > 0" and "\(p::int) (q::int). q > 0 \ abs (x - p / q) > c / of_int q ^ n" proof - - from \algebraic x\ guess p by (elim algebraicE) note p = this + from \algebraic x\ obtain p where p: "\i. coeff p i \ \" "p \ 0" "poly p x = 0" + by (elim algebraicE) blast define n where "n = degree p" \ \The derivative of @{term p} is bounded within @{term "{x-1..x+1}"}.\ let ?f = "\t. \poly (pderiv p) t\" define M where "M = (SUP t\{x-1..x+1}. ?f t)" define roots where "roots = {x. poly p x = 0} - {x}" define A_set where "A_set = {1, 1/M} \ {abs (x' - x) |x'. x' \ roots}" define A' where "A' = Min A_set" define A where "A = A' / 2" \ \We define @{term A} to be a positive real number that is less than @{term 1}, @{term "1/M"} and any distance from @{term x} to another root of @{term p}.\ \ \Properties of @{term M}, our upper bound for the derivative of @{term p}:\ have "\s\{x-1..x+1}. \y\{x-1..x+1}. ?f s \ ?f y" by (intro continuous_attains_sup) (auto intro!: continuous_intros) hence bdd: "bdd_above (?f ` {x-1..x+1})" by auto have M_pos: "M > 0" proof - from p have "pderiv p \ 0" by (auto dest!: pderiv_iszero) hence "finite {x. poly (pderiv p) x = 0}" using poly_roots_finite by blast moreover have "\finite {x-1..x+1}" by (simp add: infinite_Icc) ultimately have "\finite ({x-1..x+1} - {x. poly (pderiv p) x = 0})" by simp hence "{x-1..x+1} - {x. poly (pderiv p) x = 0} \ {}" by (intro notI) simp then obtain t where t: "t \ {x-1..x+1}" and "poly (pderiv p) t \ 0" by blast hence "0 < ?f t" by simp also from t and bdd have "\ \ M" unfolding M_def by (rule cSUP_upper) finally show "M > 0" . qed have M_sup: "?f t \ M" if "t \ {x-1..x+1}" for t proof - from that and bdd show "?f t \ M" unfolding M_def by (rule cSUP_upper) qed \ \Properties of @{term A}:\ from p poly_roots_finite[of p] have "finite A_set" unfolding A_set_def roots_def by simp have "x \ roots" unfolding roots_def by auto hence "A' > 0" using Min_gr_iff[OF \finite A_set\, folded A'_def, of 0] by (auto simp: A_set_def M_pos) hence A_pos: "A > 0" by (simp add: A_def) from \A' > 0\ have "A < A'" by (simp add: A_def) moreover from Min_le[OF \finite A_set\, folded A'_def] have "A' \ 1" "A' \ 1/M" "\x'. x' \ x \ poly p x' = 0 \ A' \ abs (x' - x)" unfolding A_set_def roots_def by auto ultimately have A_less: "A < 1" "A < 1/M" "\x'. x' \ x \ poly p x' = 0 \ A < abs (x' - x)" by fastforce+ \ \Finally: no rational number can approximate @{term x} ``well enough''.\ have "\(a::int) (b :: int). b > 0 \ \x - of_int a / of_int b\ > A / of_int b ^ n" proof (intro allI impI, rule ccontr) fix a b :: int assume b: "b > 0" and "\(A / of_int b ^ n < \x - of_int a / of_int b\)" hence ab: "abs (x - of_int a / of_int b) \ A / of_int b ^ n" by simp also from A_pos b have "A / of_int b ^ n \ A / 1" by (intro divide_left_mono) simp_all finally have ab': "abs (x - a / b) \ A" by simp also have "\ \ 1" using A_less by simp finally have ab'': "a / b \ {x-1..x+1}" by auto have no_root: "poly p (a / b) \ 0" proof assume "poly p (a / b) = 0" moreover from \x \ \\ have "x \ a / b" by auto ultimately have "A < abs (a / b - x)" using A_less(3) by simp with ab' show False by simp qed have "\x0\{x-1..x+1}. poly p (a / b) - poly p x = (a / b - x) * poly (pderiv p) x0" using ab'' by (intro poly_MVT') (auto simp: min_def max_def) with p obtain x0 :: real where x0: "x0 \ {x-1..x+1}" "poly p (a / b) = (a / b - x) * poly (pderiv p) x0" by auto from x0(2) no_root have deriv_pos: "poly (pderiv p) x0 \ 0" by auto from b p no_root have p_ab: "abs (poly p (a / b)) \ 1 / of_int b ^ n" unfolding n_def by (intro int_poly_rat_no_root_ge) note ab also from A_less b have "A / of_int b ^ n < (1/M) / of_int b ^ n" by (intro divide_strict_right_mono) simp_all also have "\ = (1 / b ^ n) / M" by simp also from M_pos ab p_ab have "\ \ abs (poly p (a / b)) / M" by (intro divide_right_mono) simp_all also from deriv_pos M_pos x0(1) have "\ \ abs (poly p (a / b)) / abs (poly (pderiv p) x0)" by (intro divide_left_mono M_sup) simp_all also have "\poly p (a / b)\ = \a / b - x\ * \poly (pderiv p) x0\" by (subst x0(2)) (simp add: abs_mult) with deriv_pos have "\poly p (a / b)\ / \poly (pderiv p) x0\ = \x - a / b\" by (subst nonzero_divide_eq_eq) simp_all finally show False by simp qed with A_pos show ?thesis using that[of A n] by blast qed text \ Since Liouville numbers are irrational, but can be approximated well by rational numbers in the Liouville sense, they must be transcendental. \ lemma (in liouville) transcendental: "\algebraic x" proof assume "algebraic x" - from liouville_irrational_algebraic[OF irrational this] guess c n . note cn = this + from liouville_irrational_algebraic[OF irrational this] + obtain c n where cn: + "c > 0" "\p q. q > 0 \ c / real_of_int q ^ n < \x - real_of_int p / real_of_int q\" + by auto define r where "r = nat \log 2 (1 / c)\" define m where "m = n + r" from cn(1) have "(1/c) = 2 powr log 2 (1/c)" by (subst powr_log_cancel) simp_all also have "log 2 (1/c) \ of_nat (nat \log 2 (1/c)\)" by linarith hence "2 powr log 2 (1/c) \ 2 powr \" by (intro powr_mono) simp_all also have "\ = 2 ^ r" unfolding r_def by (simp add: powr_realpow) finally have "1 / (2^r) \ c" using cn(1) by (simp add: field_simps) have "abs (x - p m / q m) < 1 / of_int (q m) ^ m" by (rule approx_int) also have "\ = (1 / of_int (q m) ^ r) * (1 / real_of_int (q m) ^ n)" by (simp add: m_def power_add) also from denom_gt_1[of m] have "1 / real_of_int (q m) ^ r \ 1 / 2 ^ r" by (intro divide_left_mono power_mono) simp_all also have "\ \ c" by fact finally have "abs (x - p m / q m) < c / of_int (q m) ^ n" using denom_gt_1[of m] by - (simp_all add: divide_right_mono) with cn(2)[of "q m" "p m"] denom_gt_1[of m] show False by simp qed subsection \Standard construction for Liouville numbers\ text \ We now define the standard construction for Liouville numbers. \ definition standard_liouville :: "(nat \ int) \ int \ real" where "standard_liouville p q = (\k. p k / of_int q ^ fact (Suc k))" lemma standard_liouville_summable: fixes p :: "nat \ int" and q :: int assumes "q > 1" "range p \ {0..k. p k / of_int q ^ fact (Suc k))" proof (rule summable_comparison_test') from assms(1) show "summable (\n. (1 / q) ^ n)" by (intro summable_geometric) simp_all next fix n :: nat from assms have "p n \ {0.. q / of_int q ^ (fact (Suc n))" by (auto simp: field_simps) also from assms(1) have "\ = 1 / of_int q ^ (fact (Suc n) - 1)" by (subst power_diff) (auto simp del: fact_Suc) also have "Suc n \ fact (Suc n)" by (rule fact_ge_self) with assms(1) have "1 / real_of_int q ^ (fact (Suc n) - 1) \ 1 / of_int q ^ n" by (intro divide_left_mono power_increasing) (auto simp del: fact_Suc simp add: algebra_simps) finally show "norm (p n / of_int q ^ fact (Suc n)) \ (1 / q) ^ n" by (simp add: power_divide) qed lemma standard_liouville_sums: assumes "q > 1" "range p \ {0..k. p k / of_int q ^ fact (Suc k)) sums standard_liouville p q" using standard_liouville_summable[OF assms] unfolding standard_liouville_def by (simp add: summable_sums) text \ Now we prove that the standard construction indeed yields Liouville numbers. \ lemma standard_liouville_is_liouville: assumes "q > 1" "range p \ {0..n. p n \ 0) sequentially" defines "b \ \n. q ^ fact (Suc n)" defines "a \ \n. (\k\n. p k * q ^ (fact (Suc n) - fact (Suc k)))" shows "liouville (standard_liouville p q) a b" proof fix n :: nat from assms(1) have "1 < q ^ 1" by simp also from assms(1) have "\ \ b n" unfolding b_def by (intro power_increasing) (simp_all del: fact_Suc) finally show "b n > 1" . note summable = standard_liouville_summable[OF assms(1,2)] let ?S = "\k. p (k + n + 1) / of_int q ^ (fact (Suc (k + n + 1)))" let ?C = "(q - 1) / of_int q ^ (fact (n+2))" have "a n / b n = (\k\n. p k * (of_int q ^ (fact (n+1) - fact (k+1)) / of_int q ^ (fact (n+1))))" by (simp add: a_def b_def sum_divide_distrib of_int_sum) also have "\ = (\k\n. p k / of_int q ^ (fact (Suc k)))" by (intro sum.cong refl, subst inverse_divide [symmetric], subst power_diff [symmetric]) (insert assms(1), simp_all add: divide_simps fact_mono_nat del: fact_Suc) also have "standard_liouville p q - \ = ?S" unfolding standard_liouville_def by (subst diff_eq_eq) (intro suminf_split_initial_segment' summable) finally have "abs (standard_liouville p q - a n / b n) = abs ?S" by (simp only: ) moreover from assms have "?S \ 0" by (intro suminf_nonneg allI divide_nonneg_pos summable_ignore_initial_segment summable) force+ ultimately have eq: "abs (standard_liouville p q - a n / b n) = ?S" by simp also have "?S \ (\k. ?C * (1 / q) ^ k)" proof (intro suminf_le allI summable_ignore_initial_segment summable) from assms show "summable (\k. ?C * (1 / q) ^ k)" by (intro summable_mult summable_geometric) simp_all next fix k :: nat from assms have "p (k + n + 1) \ q - 1" by force with \q > 1\ have "p (k + n + 1) / of_int q ^ fact (Suc (k + n + 1)) \ (q - 1) / of_int q ^ (fact (Suc (k + n + 1)))" by (intro divide_right_mono) (simp_all del: fact_Suc) also from \q > 1\ have "\ \ (q - 1) / of_int q ^ (fact (n+2) + k)" using fact_ineq[of "n+2" k] by (intro divide_left_mono power_increasing) (simp_all add: add_ac) also have "\ = ?C * (1 / q) ^ k" by (simp add: field_simps power_add del: fact_Suc) finally show "p (k + n + 1) / of_int q ^ fact (Suc (k + n + 1)) \ \" . qed also from assms have "\ = ?C * (\k. (1 / q) ^ k)" by (intro suminf_mult summable_geometric) simp_all also from assms have "(\k. (1 / q) ^ k) = 1 / (1 - 1 / q)" by (intro suminf_geometric) simp_all also from assms(1) have "?C * \ = of_int q ^ 1 / of_int q ^ fact (n + 2)" by (simp add: field_simps) also from assms(1) have "\ \ of_int q ^ fact (n + 1) / of_int q ^ fact (n + 2)" by (intro divide_right_mono power_increasing) (simp_all add: field_simps del: fact_Suc) also from assms(1) have "\ = 1 / (of_int q ^ (fact (n + 2) - fact (n + 1)))" by (subst power_diff) simp_all also have "fact (n + 2) - fact (n + 1) = (n + 1) * fact (n + 1)" by (simp add: algebra_simps) also from assms(1) have "1 / (of_int q ^ \) < 1 / (real_of_int q ^ (fact (n + 1) * n))" by (intro divide_strict_left_mono power_increasing mult_right_mono) simp_all also have "\ = 1 / of_int (b n) ^ n" by (simp add: power_mult b_def power_divide del: fact_Suc) finally show "\standard_liouville p q - a n / b n\ < 1 / of_int (b n) ^ n" . from assms(3) obtain k where k: "k \ n + 1" "p k \ 0" by (auto simp: frequently_def eventually_at_top_linorder) define k' where "k' = k - n - 1" from assms k have "p k \ 0" by force with k assms have k': "p (k' + n + 1) > 0" unfolding k'_def by force with assms(1,2) have "?S > 0" by (intro suminf_pos2[of _ k'] summable_ignore_initial_segment summable) (force intro!: divide_pos_pos divide_nonneg_pos)+ with eq show "\standard_liouville p q - a n / b n\ > 0" by (simp only: ) qed text \ We can now show our main result: any standard Liouville number is transcendental. \ theorem transcendental_standard_liouville: assumes "q > 1" "range p \ {0..k. p k \ 0) sequentially" shows "\algebraic (standard_liouville p q)" proof - from assms interpret liouville "standard_liouville p q" "\n. (\k\n. p k * q ^ (fact (Suc n) - fact (Suc k)))" "\n. q ^ fact (Suc n)" by (rule standard_liouville_is_liouville) from transcendental show ?thesis . qed text \ In particular: The the standard construction for constant sequences, such as the ``classic'' Liouville constant $\sum_{n=1}^\infty 10^{-n!} = 0.11000100\ldots$, are transcendental. This shows that Liouville numbers exists and therefore gives a concrete and elementary proof that transcendental numbers exist. \ corollary transcendental_standard_standard_liouville: "a \ {0<.. \algebraic (standard_liouville (\_. int a) (int b))" by (intro transcendental_standard_liouville) auto corollary transcendental_liouville_constant: "\algebraic (standard_liouville (\_. 1) 10)" by (intro transcendental_standard_liouville) auto end diff --git a/thys/Markov_Models/Classifying_Markov_Chain_States.thy b/thys/Markov_Models/Classifying_Markov_Chain_States.thy --- a/thys/Markov_Models/Classifying_Markov_Chain_States.thy +++ b/thys/Markov_Models/Classifying_Markov_Chain_States.thy @@ -1,2340 +1,2344 @@ section \Classifying Markov Chain States\ theory Classifying_Markov_Chain_States imports "HOL-Computational_Algebra.Group_Closure" Discrete_Time_Markov_Chain begin lemma eventually_mult_Gcd: fixes S :: "nat set" assumes S: "\s t. s \ S \ t \ S \ s + t \ S" assumes s: "s \ S" "s > 0" shows "eventually (\m. m * Gcd S \ S) sequentially" proof - define T where "T = insert 0 (int ` S)" with s S have "int s \ T" "0 \ T" and T: "r \ T \ t \ T \ r + t \ T" for r t by (auto simp del: of_nat_add simp add: of_nat_add [symmetric]) have "Gcd T \ group_closure T" by (rule Gcd_in_group_closure) also have "group_closure T = {s - t | s t. s \ T \ t \ T}" proof (auto intro: group_closure.base group_closure.diff) fix x assume "x \ group_closure T" then show "\s t. x = s - t \ s \ T \ t \ T" proof induction case (base x) with \0 \ T\ show ?case apply (rule_tac x=x in exI) apply (rule_tac x=0 in exI) apply auto done next case (diff x y) then obtain a b c d where "a \ T" "b \ T" "x = a - b" "c \ T" "d \ T" "y = c - d" by auto then show ?case apply (rule_tac x="a + d" in exI) apply (rule_tac x="b + c" in exI) apply (auto intro: T) done qed qed finally obtain s' t' :: int where "s' \ T" "t' \ T" "Gcd T = s' - t'" by blast moreover define s and t where "s = nat s'" and "t = nat t'" moreover have "int (Gcd S) = - int t \ S \ {0} \ t = 0" by auto (metis Gcd_dvd_nat dvd_0_right dvd_antisym nat_int nat_zminus_int) ultimately have st: "s = 0 \ s \ S" "t = 0 \ t \ S" and Gcd_S: "Gcd S = s - t" using T_def by safe simp_all with s have "t < s" by (rule_tac ccontr) auto { fix s n have "0 < n \ s \ S \ n * s \ S" proof (induct n) case (Suc n) then show ?case by (cases n) (auto intro: S) qed simp } note cmult_S = this show ?thesis unfolding eventually_sequentially proof cases assume "s = 0 \ t = 0" with st Gcd_S s have *: "Gcd S \ S" by (auto simp: int_eq_iff) then show "\N. \n\N. n * Gcd S \ S" by (auto intro!: exI[of _ 1] cmult_S) next assume "\ (s = 0 \ t = 0)" with st have "s \ S" "t \ S" "t \ 0" by auto then have "Gcd S dvd t" by auto then obtain a where a: "t = Gcd S * a" .. with \t \ 0\ have "0 < a" by auto show "\N. \n\N. n * Gcd S \ S" proof (safe intro!: exI[of _ "a * a"]) fix n define m where "m = (n - a * a) div a" define r where "r = (n - a * a) mod a" with \0 < a\ have "r < a" by simp moreover define am where "am = a + m" ultimately have "r < am" by simp assume "a * a \ n" then have n: "n = a * a + (m * a + r)" unfolding m_def r_def by simp have "n * Gcd S = am * t + r * Gcd S" unfolding n a by (simp add: field_simps am_def) also have "\ = r * s + (am - r) * t" unfolding \Gcd S = s - t\ using \t < s\ \r < am\ by (simp add: field_simps diff_mult_distrib2) also have "\ \ S" using \s \ S\ \t \ S\ \r < am\ by (cases "r = 0") (auto intro!: cmult_S S) finally show "n * Gcd S \ S" . qed qed qed context MC_syntax begin subsection \Expected number of visits\ definition "G s t = (\\<^sup>+\. scount (HLD {t}) (s ## \) \T s)" lemma G_eq: "G s t = (\\<^sup>+\. emeasure (count_space UNIV) {i. (s ## \) !! i = t} \T s)" by (simp add: G_def scount_eq_emeasure HLD_iff) definition "p s t n = \

(\ in T s. (s ## \) !! n = t)" definition "gf_G s t z = (\n. p s t n *\<^sub>R z ^ n)" definition "convergence_G s t z \ summable (\n. p s t n * norm z ^ n)" lemma p_nonneg[simp]: "0 \ p x y n" by (simp add: p_def) lemma p_le_1: "p x y n \ 1" by (simp add: p_def) lemma p_x_x_0[simp]: "p x x 0 = 1" by (simp add: p_def T.prob_space del: space_T) lemma p_0: "p x y 0 = (if x = y then 1 else 0)" by (simp add: p_def T.prob_space del: space_T) lemma p_in_reachable: assumes "(x, y) \ (SIGMA x:UNIV. K x)\<^sup>*" shows "p x y n = 0" unfolding p_def proof (rule T.prob_eq_0_AE) from AE_T_reachable show "AE \ in T x. (x ## \) !! n \ y" proof eventually_elim fix \ assume "alw (HLD ((SIGMA \:UNIV. K \)\<^sup>* `` {x})) \" then have "alw (HLD (- {y})) \" using assms by (auto intro: alw_mono simp: HLD_iff) then show "(x ## \) !! n \ y" using assms by (cases n) (auto simp: alw_HLD_iff_streams streams_iff_snth) qed qed lemma p_Suc: "ennreal (p x y (Suc n)) = (\\<^sup>+ w. p w y n \K x)" unfolding p_def T.emeasure_eq_measure[symmetric] by (subst emeasure_Collect_T) simp_all lemma p_Suc': "p x y (Suc n) = (\x'. p x' y n \K x)" using p_Suc[of x y n] by (subst (asm) nn_integral_eq_integral) (auto simp: p_le_1 intro!: measure_pmf.integrable_const_bound[where B=1]) lemma p_add: "p x y (n + m) = (\\<^sup>+ w. p x w n * p w y m \count_space UNIV)" proof (induction n arbitrary: x) case 0 have [simp]: "\w. (if x = w then 1 else 0) * p w y m = ennreal (p x y m) * indicator {x} w" by auto show ?case by (simp add: p_0 one_ennreal_def[symmetric] max_def) next case (Suc n) define X where "X = (SIGMA x:UNIV. K x)\<^sup>* `` K x" then have X: "countable X" by (blast intro: countable_Image countable_reachable countable_set_pmf) then interpret X: sigma_finite_measure "count_space X" by (rule sigma_finite_measure_count_space_countable) interpret XK: pair_sigma_finite "K x" "count_space X" by unfold_locales have "ennreal (p x y (Suc n + m)) = (\\<^sup>+t. (\\<^sup>+w. p t w n * p w y m \count_space UNIV) \K x)" by (simp add: p_Suc Suc) also have "\ = (\\<^sup>+t. (\\<^sup>+w. ennreal (p t w n * p w y m) * indicator X w \count_space UNIV) \K x)" by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff AE_count_space Image_iff p_in_reachable X_def split: split_indicator) also have "\ = (\\<^sup>+t. (\\<^sup>+w. p t w n * p w y m \count_space X) \K x)" by (subst nn_integral_restrict_space[symmetric]) (simp_all add: restrict_count_space) also have "\ = (\\<^sup>+w. (\\<^sup>+t. p t w n * p w y m \K x) \count_space X)" apply (rule XK.Fubini'[symmetric]) unfolding measurable_split_conv apply (rule measurable_compose_countable'[OF _ measurable_snd X]) apply (rule measurable_compose[OF measurable_fst]) apply simp done also have "\ = (\\<^sup>+w. (\\<^sup>+t. ennreal (p t w n * p w y m) * indicator X w \K x) \count_space UNIV)" by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space nn_integral_multc) also have "\ = (\\<^sup>+w. (\\<^sup>+t. ennreal (p t w n * p w y m) \K x) \count_space UNIV)" by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff AE_count_space Image_iff p_in_reachable X_def split: split_indicator) also have "\ = (\\<^sup>+w. (\\<^sup>+t. p t w n \K x) * p w y m \count_space UNIV)" by (simp add: nn_integral_multc[symmetric] ennreal_mult) finally show ?case by (simp add: ennreal_mult p_Suc) qed lemma prob_reachable_le: assumes [simp]: "m \ n" shows "p x y m * p y w (n - m) \ p x w n" proof - have "p x y m * p y w (n - m) = (\\<^sup>+y'. ennreal (p x y m * p y w (n - m)) * indicator {y} y' \count_space UNIV)" by simp also have "\ \ p x w (m + (n - m))" by (subst p_add) (auto intro!: nn_integral_mono split: split_indicator simp del: nn_integral_indicator_singleton) finally show ?thesis by simp qed lemma G_eq_suminf: "G x y = (\i. ennreal (p x y i))" proof - have *: "\i \. indicator {\ \ space S. (x ## \) !! i = y} \ = indicator {i. (x ## \) !! i = y} i" by (auto simp: space_stream_space split: split_indicator) have "G x y = (\\<^sup>+ \. (\i. indicator {\\space (T x). (x ## \) !! i = y} \) \T x)" unfolding G_eq by (simp add: nn_integral_count_space_nat[symmetric] *) also have "\ = (\i. ennreal (p x y i))" by (simp add: T.emeasure_eq_measure[symmetric] p_def nn_integral_suminf) finally show ?thesis . qed lemma G_eq_real_suminf: "convergence_G x y (1::real) \ G x y = ennreal (\i. p x y i)" unfolding G_eq_suminf by (intro suminf_ennreal ennreal_suminf_neq_top p_nonneg) (auto simp: convergence_G_def p_def) lemma convergence_norm_G: "convergence_G x y z \ summable (\n. p x y n * norm z ^ n)" unfolding convergence_G_def . lemma convergence_G: "convergence_G x y (z::'a::{banach, real_normed_div_algebra}) \ summable (\n. p x y n *\<^sub>R z ^ n)" unfolding convergence_G_def by (rule summable_norm_cancel) (simp add: abs_mult norm_power) lemma convergence_G_less_1: fixes z :: "_ :: {banach, real_normed_field}" assumes z: "norm z < 1" shows "convergence_G x y z" unfolding convergence_G_def proof (rule summable_comparison_test) have "\n. p x y n * norm (z ^ n) \ 1 * norm (z ^ n)" by (intro mult_right_mono p_le_1) simp_all then show "\N. \n\N. norm (p x y n * norm z ^ n) \ norm z ^ n" by (simp add: norm_power) qed (simp add: z summable_geometric) lemma lim_gf_G: "((\z. ennreal (gf_G x y z)) \ G x y) (at_left (1::real))" unfolding gf_G_def G_eq_suminf real_scaleR_def by (intro power_series_tendsto_at_left p_nonneg p_le_1 summable_power_series) subsection \Reachability probability\ definition "u x y n = \

(\ in T x. ev_at (HLD {y}) n \)" definition "U s t = \

(\ in T s. ev (HLD {t}) \)" definition "gf_U x y z = (\n. u x y n *\<^sub>R z ^ Suc n)" definition "f x y n = \

(\ in T x. ev_at (HLD {y}) n (x ## \))" definition "F s t = \

(\ in T s. ev (HLD {t}) (s ## \))" definition "gf_F x y z = (\n. f x y n * z ^ n)" lemma f_Suc: "x \ y \ f x y (Suc n) = u x y n" by (simp add: u_def f_def) lemma f_Suc_eq: "f x x (Suc n) = 0" by (simp add: f_def) lemma f_0: "f x y 0 = (if x = y then 1 else 0)" using T.prob_space by (simp add: f_def) lemma shows u_nonneg: "0 \ u x y n" and u_le_1: "u x y n \ 1" by (simp_all add: u_def) lemma shows f_nonneg: "0 \ f x y n" and f_le_1: "f x y n \ 1" by (simp_all add: f_def) lemma U_nonneg[simp]: "0 \ U x y" by (simp add: U_def) lemma U_le_1: "U s t \ 1" by (auto simp add: U_def intro!: antisym) lemma U_cases: "U s s = 1 \ U s s < 1" by (auto simp add: U_def intro!: antisym) lemma u_sums_U: "u x y sums U x y" unfolding u_def[abs_def] U_def ev_iff_ev_at by (intro T.prob_sums) (auto intro: ev_at_unique) lemma gf_U_eq_U: "gf_U x y 1 = U x y" using u_sums_U[THEN sums_unique] by (simp add: gf_U_def U_def) lemma f_sums_F: "f x y sums F x y" unfolding f_def[abs_def] F_def ev_iff_ev_at by (intro T.prob_sums) (auto intro: ev_at_unique) lemma F_nonneg[simp]: "0 \ F x y" by (auto simp: F_def) lemma F_le_1: "F x y \ 1" by (simp add: F_def) lemma gf_F_eq_F: "gf_F x y 1 = F x y" using f_sums_F[THEN sums_unique] by (simp add: gf_F_def F_def) lemma gf_F_le_1: fixes z :: real assumes z: "0 \ z" "z \ 1" shows "gf_F x y z \ 1" proof - have "gf_F x y z \ gf_F x y 1" using z unfolding gf_F_def by (intro suminf_le[OF _ summable_comparison_test[OF _ sums_summable[OF f_sums_F[of x y]]]] mult_left_mono allI f_nonneg) (simp_all add: power_le_one f_nonneg mult_right_le_one_le f_le_1 sums_summable[OF f_sums_F[of x y]]) also have "\ \ 1" by (simp add: gf_F_eq_F F_def) finally show ?thesis . qed lemma u_le_p: "u x y n \ p x y (Suc n)" unfolding u_def p_def by (auto intro!: T.finite_measure_mono dest: ev_at_HLD_imp_snth) lemma f_le_p: "f x y n \ p x y n" unfolding f_def p_def by (auto intro!: T.finite_measure_mono dest: ev_at_HLD_imp_snth) lemma convergence_norm_U: fixes z :: "_ :: real_normed_div_algebra" assumes z: "convergence_G x y z" shows "summable (\n. u x y n * norm z ^ Suc n)" using summable_ignore_initial_segment[OF convergence_norm_G[OF z], of 1] by (rule summable_comparison_test[rotated]) (auto simp add: u_nonneg abs_mult intro!: exI[of _ 0] mult_right_mono u_le_p) lemma convergence_norm_F: fixes z :: "_ :: real_normed_div_algebra" assumes z: "convergence_G x y z" shows "summable (\n. f x y n * norm z ^ n)" using convergence_norm_G[OF z] by (rule summable_comparison_test[rotated]) (auto simp add: f_nonneg abs_mult intro!: exI[of _ 0] mult_right_mono f_le_p) lemma gf_G_nonneg: fixes z :: real shows "0 \ z \ z < 1 \ 0 \ gf_G x y z" unfolding gf_G_def by (intro suminf_nonneg convergence_G convergence_G_less_1) simp_all lemma gf_F_nonneg: fixes z :: real shows "0 \ z \ z < 1 \ 0 \ gf_F x y z" unfolding gf_F_def using convergence_norm_F[OF convergence_G_less_1, of z x y] by (intro suminf_nonneg) (simp_all add: f_nonneg) lemma convergence_U: fixes z :: "_ :: banach" shows "convergence_G x y z \ summable (\n. u x y n * z ^ Suc n)" by (rule summable_norm_cancel) (auto simp add: abs_mult u_nonneg power_abs dest!: convergence_norm_U) lemma p_eq_sum_p_u: "p x y (Suc n) = (\i\n. p y y (n - i) * u x y i)" proof - have "\\. \ !! n = y \ (\i. i \ n \ ev_at (HLD {y}) i \)" proof (induction n) case (Suc n) then obtain i where "i \ n" "ev_at (HLD {y}) i (stl \)" by auto then show ?case by (auto intro!: exI[of _ "if HLD {y} \ then 0 else Suc i"]) qed (simp add: HLD_iff) then have "p x y (Suc n) = (\i\n. \

(\ in T x. ev_at (HLD {y}) i \ \ \ !! n = y))" unfolding p_def by (intro T.prob_sum) (auto intro: ev_at_unique) also have "\ = (\i\n. p y y (n - i) * u x y i)" proof (intro sum.cong refl) fix i assume i: "i \ {.. n}" then have "\\. (Suc i \ n \ \ !! (n - Suc i) = y) \ ((y ## \) !! (n - i) = y)" by (auto simp: Stream_snth diff_Suc split: nat.split) from i have "i \ n" by auto then have "\

(\ in T x. ev_at (HLD {y}) i \ \ \ !! n = y) = (\\'. \

(\ in T y. (y ## \) !! (n - i) = y) * indicator {\'\space (T x). ev_at (HLD {y}) i \' } \' \T x)" by (subst prob_T_split[where n="Suc i"]) (auto simp: ev_at_shift ev_at_HLD_single_imp_snth shift_snth diff_Suc split: split_indicator nat.split intro!: Bochner_Integration.integral_cong arg_cong2[where f=measure] simp del: stake.simps integral_mult_right_zero) then show "\

(\ in T x. ev_at (HLD {y}) i \ \ \ !! n = y) = p y y (n - i) * u x y i" by (simp add: p_def u_def) qed finally show ?thesis . qed lemma p_eq_sum_p_f: "p x y n = (\i\n. p y y (n - i) * f x y i)" by (cases n) (simp_all del: sum.atMost_Suc add: f_0 p_0 p_eq_sum_p_u atMost_Suc_eq_insert_0 zero_notin_Suc_image sum.reindex f_Suc f_Suc_eq) lemma gf_G_eq_gf_F: assumes z: "norm z < 1" shows "gf_G x y z = gf_F x y z * gf_G y y z" proof - have "gf_G x y z = (\n. \i\n. p y y (n - i) * f x y i * z^n)" by (simp add: gf_G_def p_eq_sum_p_f[of x y] sum_distrib_right) also have "\ = (\n. \i\n. (f x y i * z^i) * (p y y (n - i) * z^(n - i)))" by (intro arg_cong[where f=suminf] sum.cong ext atLeast0AtMost[symmetric]) (simp_all add: power_add[symmetric]) also have "\ = (\n. f x y n * z^n) * (\n. p y y n * z^n)" using convergence_norm_F[OF convergence_G_less_1[OF z]] convergence_norm_G[OF convergence_G_less_1[OF z]] by (intro Cauchy_product[symmetric]) (auto simp: f_nonneg abs_mult power_abs) also have "\ = gf_F x y z * gf_G y y z" by (simp add: gf_F_def gf_G_def) finally show ?thesis . qed lemma gf_G_eq_gf_U: fixes z :: "'z :: {banach, real_normed_field}" assumes z: "convergence_G x x z" shows "gf_G x x z = 1 / (1 - gf_U x x z)" "gf_U x x z \ 1" proof - { fix n have "p x x (Suc n) *\<^sub>R z^Suc n = (\i\n. (p x x (n - i) * u x x i) *\<^sub>R z^Suc n)" unfolding scaleR_sum_left[symmetric] by (simp add: p_eq_sum_p_u) also have "\ = (\i\n. (u x x i *\<^sub>R z^Suc i) * (p x x (n - i) *\<^sub>R z^(n - i)))" by (intro sum.cong refl) (simp add: field_simps power_diff cong: disj_cong) finally have "p x x (Suc n) *\<^sub>R z^(Suc n) = (\i\n. (u x x i *\<^sub>R z^Suc i) * (p x x (n - i) *\<^sub>R z^(n - i)))" unfolding atLeast0AtMost . } note gfs_Suc_eq = this have "gf_G x x z = 1 + (\n. p x x (Suc n) *\<^sub>R z^(Suc n))" unfolding gf_G_def by (subst suminf_split_initial_segment[OF convergence_G[OF z], of 1]) simp also have "\ = 1 + (\n. \i\n. (u x x i *\<^sub>R z^Suc i) * (p x x (n - i) *\<^sub>R z^(n - i)))" unfolding gfs_Suc_eq .. also have "\ = 1 + gf_U x x z * gf_G x x z" unfolding gf_U_def gf_G_def by (subst Cauchy_product) (auto simp: u_nonneg norm_power simp del: power_Suc intro!: z convergence_norm_G convergence_norm_U) finally show "gf_G x x z = 1 / (1 - gf_U x x z)" "gf_U x x z \ 1" apply - apply (cases "gf_U x x z = 1") apply (auto simp add: field_simps) done qed lemma gf_U: "(gf_U x y \ U x y) (at_left 1)" proof - have "((\z. ennreal (\n. u x y n * z ^ n)) \ (\n. ennreal (u x y n))) (at_left 1)" using u_le_1 u_nonneg by (intro power_series_tendsto_at_left summable_power_series) also have "(\n. ennreal (u x y n)) = ennreal (suminf (u x y))" by (intro u_nonneg suminf_ennreal ennreal_suminf_neq_top sums_summable[OF u_sums_U]) also have "suminf (u x y) = U x y" using u_sums_U by (rule sums_unique[symmetric]) finally have "((\z. \n. u x y n * z ^ n) \ U x y) (at_left 1)" by (rule tendsto_ennrealD) (auto simp: u_nonneg u_le_1 intro!: suminf_nonneg summable_power_series eventually_at_left_1) then have "((\z. z * (\n. u x y n * z ^ n)) \ 1 * U x y) (at_left 1)" by (intro tendsto_intros) simp then have "((\z. \n. u x y n * z ^ Suc n) \ 1 * U x y) (at_left 1)" apply (rule filterlim_cong[OF refl refl, THEN iffD1, rotated]) apply (rule eventually_at_left_1) apply (subst suminf_mult[symmetric]) apply (auto intro!: summable_power_series u_le_1 u_nonneg) apply (simp add: field_simps) done then show ?thesis by (simp add: gf_U_def[abs_def] U_def) qed lemma gf_U_le_1: assumes z: "0 < z" "z < 1" shows "gf_U x y z \ (1::real)" proof - note u = u_sums_U[of x y, THEN sums_summable] have "gf_U x y z \ gf_U x y 1" using z unfolding gf_U_def real_scaleR_def by (intro suminf_le allI mult_mono power_mono summable_comparison_test_ev[OF _ u] always_eventually) (auto simp: u_nonneg intro!: mult_left_le mult_le_one power_le_one) also have "\ \ 1" unfolding gf_U_eq_U by (rule U_le_1) finally show ?thesis . qed lemma gf_F: "(gf_F x y \ F x y) (at_left 1)" proof - have "((\z. ennreal (\n. f x y n * z ^ n)) \ (\n. ennreal (f x y n))) (at_left 1)" using f_le_1 f_nonneg by (intro power_series_tendsto_at_left summable_power_series) also have "(\n. ennreal (f x y n)) = ennreal (suminf (f x y))" by (intro f_nonneg suminf_ennreal ennreal_suminf_neq_top sums_summable[OF f_sums_F]) also have "suminf (f x y) = F x y" using f_sums_F by (rule sums_unique[symmetric]) finally have "((\z. \n. f x y n * z ^ n) \ F x y) (at_left 1)" by (rule tendsto_ennrealD) (auto simp: f_nonneg f_le_1 intro!: suminf_nonneg summable_power_series eventually_at_left_1) then show ?thesis by (simp add: gf_F_def[abs_def] F_def) qed lemma U_bounded: "0 \ U x y" "U x y \ 1" unfolding U_def by simp_all subsection \Recurrent states\ definition recurrent :: "'s \ bool" where "recurrent s \ (AE \ in T s. ev (HLD {s}) \)" lemma recurrent_iff_U_eq_1: "recurrent s \ U s s = 1" unfolding recurrent_def U_def by (subst T.prob_Collect_eq_1) simp_all definition "H s t = \

(\ in T s. alw (ev (HLD {t})) \)" lemma H_eq: "recurrent s \ H s s = 1" "\ recurrent s \ H s s = 0" "H s t = U s t * H t t" proof - define H' where "H' t n = {\\space S. enat n \ scount (HLD {t::'s}) \}" for t n have [measurable]: "\y n. H' y n \ sets S" by (simp add: H'_def) let ?H' = "\s t n. measure (T s) (H' t n)" { fix x y :: 's and \ have "Suc 0 \ scount (HLD {y}) \ \ ev (HLD {y}) \" using scount_eq_0_iff[of "HLD {y}" \] by (cases "scount (HLD {y}) \" rule: enat_coexhaust) (auto simp: not_ev_iff[symmetric] eSuc_enat[symmetric] enat_0 HLD_iff[abs_def]) } then have H'_1: "\x y. ?H' x y 1 = U x y" unfolding H'_def U_def by simp { fix n and x y :: 's let ?U = "(not (HLD {y}) suntil (HLD {y} aand nxt (\\. enat n \ scount (HLD {y}) \)))" { fix \ have "enat (Suc n) \ scount (HLD {y}) \ \ ?U \" proof assume "enat (Suc n) \ scount (HLD {y}) \" with scount_eq_0_iff[of "HLD {y}" \] have "ev (HLD {y}) \" "enat (Suc n) \ scount (HLD {y}) \" by (auto simp add: not_ev_iff[symmetric] eSuc_enat[symmetric]) then show "?U \" by (induction rule: ev_induct_strong) (auto simp: scount_simps eSuc_enat[symmetric] intro: suntil.intros) next assume "?U \" then show "enat (Suc n) \ scount (HLD {y}) \" by induction (auto simp: scount_simps eSuc_enat[symmetric]) qed } then have "emeasure (T x) (H' y (Suc n)) = emeasure (T x) {\\space (T x). ?U \}" by (simp add: H'_def) also have "\ = U x y * ?H' y y n" by (subst emeasure_suntil_HLD) (simp_all add: T.emeasure_eq_measure U_def H'_def ennreal_mult) finally have "?H' x y (Suc n) = U x y * ?H' y y n" by (simp add: T.emeasure_eq_measure) } note H'_Suc = this { fix m and x :: 's have "?H' x x (Suc m) = U x x^Suc m" using H'_1 H'_Suc by (induct m) auto } note H'_eq = this { fix x y have "?H' x y \ measure (T x) (\i. H' y i)" apply (rule T.finite_Lim_measure_decseq) apply safe apply simp apply (auto simp add: decseq_Suc_iff subset_eq H'_def eSuc_enat[symmetric] intro: ile_eSuc order_trans) done also have "(\i. H' y i) = {\\space (T x). alw (ev (HLD {y})) \}" by (auto simp: H'_def scount_infinite_iff[symmetric]) (metis Suc_ile_eq enat.exhaust neq_iff) finally have "?H' x y \ H x y" unfolding H_def . } note H'_lim = this from H'_lim[of s s, THEN LIMSEQ_Suc] have "(\n. U s s ^ Suc n) \ H s s" by (simp add: H'_eq) then have lim_H: "(\n. U s s ^ n) \ H s s" by (rule LIMSEQ_imp_Suc) have "U s s < 1 \ (\n. U s s ^ n) \ 0" by (rule LIMSEQ_realpow_zero) (simp_all add: U_def) with lim_H have "U s s < 1 \ H s s = 0" by (blast intro: LIMSEQ_unique) moreover have "U s s = 1 \ (\n. U s s ^ n) \ 1" by simp with lim_H have "U s s = 1 \ H s s = 1" by (blast intro: LIMSEQ_unique) moreover note recurrent_iff_U_eq_1 U_cases ultimately show "recurrent s \ H s s = 1" "\ recurrent s \ H s s = 0" by (metis one_neq_zero)+ from H'_lim[of s t, THEN LIMSEQ_Suc] H'_Suc[of s] have "(\n. U s t * ?H' t t n) \ H s t" by simp moreover have "(\n. U s t * ?H' t t n) \ U s t * H t t" by (intro tendsto_intros H'_lim) ultimately show "H s t = U s t * H t t" by (blast intro: LIMSEQ_unique) qed lemma recurrent_iff_G_infinite: "recurrent x \ G x x = \" proof - have "((\z. ennreal (gf_G x x z)) \ G x x) (at_left 1)" by (rule lim_gf_G) then have G: "((\z. ennreal (1 / (1 - gf_U x x z))) \ G x x) (at_left (1::real))" apply (rule filterlim_cong[OF refl refl, THEN iffD1, rotated]) apply (rule eventually_at_left_1) apply (subst gf_G_eq_gf_U) apply (rule convergence_G_less_1) apply simp apply simp done { fix z :: real assume z: "0 < z" "z < 1" have 1: "summable (u x x)" using u_sums_U by (rule sums_summable) have "gf_U x x z \ 1" using gf_G_eq_gf_U[OF convergence_G_less_1[of z]] z by simp moreover have "gf_U x x z \ U x x" unfolding gf_U_def gf_U_eq_U[symmetric] using z by (intro suminf_le) (auto simp add: 1 convergence_U convergence_G_less_1 u_nonneg simp del: power_Suc intro!: mult_right_le_one_le power_le_one) ultimately have "gf_U x x z < 1" using U_bounded[of x x] by simp } note strict = this { assume "U x x = 1" moreover have "((\xa. 1 - gf_U x x xa :: real) \ 1 - U x x) (at_left 1)" by (intro tendsto_intros gf_U) moreover have "eventually (\z. gf_U x x z < 1) (at_left (1::real))" by (auto intro!: eventually_at_left_1 strict simp: \U x x = 1\ gf_U_eq_U) ultimately have "((\z. ennreal (1 / (1 - gf_U x x z))) \ top) (at_left 1)" unfolding ennreal_tendsto_top_eq_at_top by (intro LIM_at_top_divide[where a=1] tendsto_const zero_less_one) (auto simp: field_simps) with G have "G x x = top" by (rule tendsto_unique[rotated]) simp } moreover { assume "U x x < 1" then have "((\xa. ennreal (1 / (1 - gf_U x x xa))) \ 1 / (1 - U x x)) (at_left 1)" by (intro tendsto_intros gf_U tendsto_ennrealI) simp from tendsto_unique[OF _ G this] have "G x x \ \" by simp } ultimately show ?thesis using U_cases recurrent_iff_U_eq_1 by auto qed definition communicating :: "('s \ 's) set" where "communicating = acc \ acc\" definition essential_class :: "'s set \ bool" where "essential_class C \ C \ UNIV // communicating \ acc `` C \ C" lemma accI_U: assumes "0 < U x y" shows "(x, y) \ acc" proof (rule ccontr) assume *: "(x, y) \ acc" { fix \ assume "ev (HLD {y}) \" "alw (HLD (acc `` {x})) \" from this * have False by induction (auto simp: HLD_iff) } with AE_T_reachable[of x] have "U x y = 0" unfolding U_def by (intro T.prob_eq_0_AE) auto with \0 < U x y\ show False by auto qed lemma accD_pos: assumes "(x, y) \ acc" shows "\n. 0 < p x y n" using assms proof induction case base with T.prob_space[of x] show ?case by (auto intro!: exI[of _ 0]) next have [simp]: "\x y. (if x = y then 1 else 0::real) = indicator {y} x" by simp case (step w y) then obtain n where "0 < p x w n" and "0 < pmf (K w) y" by (auto simp: set_pmf_iff less_le) then have "0 < p x w n * pmf (K w) y" by (intro mult_pos_pos) also have "\ \ p x w n * p w y (Suc 0)" by (simp add: p_Suc' p_0 pmf.rep_eq) also have "\ \ p x y (Suc n)" using prob_reachable_le[of n "Suc n" x w y] by simp finally show ?case .. qed lemma accI_pos: "0 < p x y n \ (x, y) \ acc" proof (induct n arbitrary: x) case (Suc n) then have less: "0 < (\x'. p x' y n \K x)" by (simp add: p_Suc') have "\x'\K x. 0 < p x' y n" proof (rule ccontr) assume "\ ?thesis" then have "AE x' in K x. p x' y n = 0" by (simp add: AE_measure_pmf_iff less_le) then have "(\x'. p x' y n \K x) = (\x'. 0 \K x)" by (intro integral_cong_AE) simp_all with less show False by simp qed with Suc show ?case by (auto intro: converse_rtrancl_into_rtrancl) qed (simp add: p_0 split: if_split_asm) lemma recurrent_iffI_communicating: assumes "(x, y) \ communicating" shows "recurrent x \ recurrent y" proof - from assms obtain n m where "0 < p x y n" "0 < p y x m" by (force simp: communicating_def dest: accD_pos) moreover { fix x y n m assume "0 < p x y n" "0 < p y x m" "G y y = \" then have "\ = ennreal (p x y n * p y x m) * G y y" by (auto intro: mult_pos_pos simp: ennreal_mult_top) also have "ennreal (p x y n * p y x m) * G y y = (\i. ennreal (p x y n * p y x m) * p y y i)" unfolding G_eq_suminf by (rule ennreal_suminf_cmult[symmetric]) also have "\ \ (\i. ennreal (p x x (n + i + m)))" proof (intro suminf_le allI) fix i have "(p x y n * p y y ((n + i) - n)) * p y x ((n + i + m) - (n + i)) \ p x y (n + i) * p y x ((n + i + m) - (n + i))" by (intro mult_right_mono prob_reachable_le) simp_all also have "\ \ p x x (n + i + m)" by (intro prob_reachable_le) simp_all finally show "ennreal (p x y n * p y x m) * p y y i \ ennreal (p x x (n + i + m))" by (simp add: ac_simps ennreal_mult'[symmetric]) qed auto also have "\ \ (\i. ennreal (p x x (i + (n + m))))" by (simp add: ac_simps) also have "\ \ (\i. ennreal (p x x i))" by (subst suminf_offset[of "\i. ennreal (p x x i)" "n + m"]) auto also have "\ \ G x x" unfolding G_eq_suminf by (auto intro!: suminf_le_pos) finally have "G x x = \" by (simp add: top_unique) } ultimately show ?thesis using recurrent_iff_G_infinite by blast qed lemma recurrent_acc: assumes "recurrent x" "(x, y) \ acc" shows "U y x = 1" "H y x = 1" "recurrent y" "(x, y) \ communicating" proof - { fix w y assume step: "(x, w) \ acc" "y \ K w" "U w x = 1" "H w x = 1" "recurrent w" "x \ y" have "measure (K w) UNIV = U w x" using step measure_pmf.prob_space[of "K w"] by simp also have "\ = (\v. indicator {x} v + U v x * indicator (- {x}) v \K w)" unfolding U_def by (subst prob_T) (auto intro!: Bochner_Integration.integral_cong arg_cong2[where f=measure] AE_I2 simp: ev_Stream T.prob_eq_1 split: split_indicator) also have "\ = measure (K w) {x} + (\v. U v x * indicator (- {x}) v \K w)" by (subst Bochner_Integration.integral_add) (auto intro!: measure_pmf.integrable_const_bound[where B=1] simp: abs_mult mult_le_one U_bounded(2) measure_pmf.emeasure_eq_measure) finally have "measure (K w) UNIV - measure (K w) {x} = (\v. U v x * indicator (- {x}) v \K w)" by simp also have "measure (K w) UNIV - measure (K w) {x} = measure (K w) (UNIV - {x})" by (subst measure_pmf.finite_measure_Diff) auto finally have "0 = (\v. indicator (- {x}) v \K w) - (\v. U v x * indicator (- {x}) v \K w)" by (simp add: measure_pmf.emeasure_eq_measure Compl_eq_Diff_UNIV) also have "\ = (\v. (1 - U v x) * indicator (- {x}) v \K w)" by (subst Bochner_Integration.integral_diff[symmetric]) (auto intro!: measure_pmf.integrable_const_bound[where B=1] Bochner_Integration.integral_cong simp: abs_mult mult_le_one U_bounded(2) split: split_indicator) also have "\ \ (\v. (1 - U y x) * indicator {y} v \K w)" (is "_ \ ?rhs") using \recurrent x\ by (intro integral_mono measure_pmf.integrable_const_bound[where B=1]) (auto simp: abs_mult mult_le_one U_bounded(2) recurrent_iff_U_eq_1 field_simps split: split_indicator) also (xtrans) have "?rhs = (1 - U y x) * pmf (K w) y" by (simp add: measure_pmf.emeasure_eq_measure pmf.rep_eq) finally have "(1 - U y x) * pmf (K w) y = 0" by (auto intro!: antisym simp: U_bounded(2) mult_le_0_iff) with \y \ K w\ have "U y x = 1" by (simp add: set_pmf_iff) then have "U y x = 1" "H y x = 1" using H_eq(3)[of y x] H_eq(1)[of x] by (simp_all add: \recurrent x\) then have "(y, x) \ acc" by (intro accI_U) auto with step have "(x, y) \ communicating" by (auto simp add: communicating_def intro: rtrancl_trans) with \recurrent x\ have "recurrent y" by (simp add: recurrent_iffI_communicating) note this \U y x = 1\ \H y x = 1\ \(x, y) \ communicating\ } note enabled = this from \(x, y) \ acc\ show "U y x = 1" "H y x = 1" "recurrent y" "(x, y) \ communicating" proof induction case base then show "U x x = 1" "H x x = 1" "recurrent x" "(x, x) \ communicating" using \recurrent x\ H_eq(1)[of x] by (auto simp: recurrent_iff_U_eq_1 communicating_def) next case (step w y) with enabled[of w y] \recurrent x\ H_eq(1)[of x] have "U y x = 1 \ H y x = 1 \ recurrent y \ (x, y) \ communicating" by (cases "x = y") (auto simp: recurrent_iff_U_eq_1 communicating_def) then show "U y x = 1" "H y x = 1" "recurrent y" "(x, y) \ communicating" by auto qed qed lemma equiv_communicating: "equiv UNIV communicating" by (auto simp: equiv_def sym_def communicating_def refl_on_def trans_def) lemma recurrent_class: assumes "recurrent x" shows "acc `` {x} = communicating `` {x}" using recurrent_acc(4)[OF \recurrent x\] by (auto simp: communicating_def) lemma irreduccible_recurrent_class: assumes "recurrent x" shows "acc `` {x} \ UNIV // communicating" unfolding recurrent_class[OF \recurrent x\] by (rule quotientI) simp lemma essential_classI: assumes C: "C \ UNIV // communicating" assumes eq: "\x y. x \ C \ (x, y) \ acc \ y \ C" shows "essential_class C" by (auto simp: essential_class_def intro: C) (metis eq) lemma essential_recurrent_class: assumes "recurrent x" shows "essential_class (communicating `` {x})" unfolding recurrent_class[OF \recurrent x\, symmetric] apply (rule essential_classI) apply (rule irreduccible_recurrent_class[OF assms]) apply (auto simp: communicating_def) done lemma essential_classD2: "essential_class C \ x \ C \ (x, y) \ acc \ y \ C" unfolding essential_class_def by auto lemma essential_classD3: "essential_class C \ x \ C \ y \ C \ (x, y) \ communicating" unfolding essential_class_def by (auto elim!: quotientE simp: communicating_def) lemma AE_acc: shows "AE \ in T x. \m. (x, (x ## \) !! m) \ acc" using AE_T_reachable by eventually_elim (auto simp: alw_HLD_iff_streams streams_iff_snth Stream_snth split: nat.splits) lemma finite_essential_class_imp_recurrent: assumes C: "essential_class C" "finite C" and x: "x \ C" shows "recurrent x" proof - have "AE \ in T x. \y\C. alw (ev (HLD {y})) \" using AE_T_reachable proof eventually_elim fix \ assume "alw (HLD (acc `` {x})) \" then have "alw (HLD C) \" by (rule alw_mono) (auto simp: HLD_iff intro: assms essential_classD2) then show "\y\C. alw (ev (HLD {y})) \" by (rule pigeonhole_stream) fact qed then have "1 = \

(\ in T x. \y\C. alw (ev (HLD {y})) \)" by (subst (asm) T.prob_Collect_eq_1[symmetric]) (auto simp: \finite C\) also have "\ = measure (T x) (\y\C. {\\space (T x). alw (ev (HLD {y})) \})" by (intro arg_cong2[where f=measure]) auto also have "\ \ (\y\C. H x y)" unfolding H_def using \finite C\ by (rule T.finite_measure_subadditive_finite) auto also have "\ = (\y\C. U x y * H y y)" by (auto intro!: sum.cong H_eq) finally have "\y\C. recurrent y" by (rule_tac ccontr) (simp add: H_eq(2)) - then guess y .. + then obtain y where "y \ C" "recurrent y" .. from essential_classD3[OF C(1) x this(1)] recurrent_acc(3)[OF this(2)] show "recurrent x" by (simp add: communicating_def) qed lemma irreducibleD: "C \ UNIV // communicating \ a \ C \ b \ C \ (a, b) \ communicating" by (auto elim!: quotientE simp: communicating_def) lemma irreducibleD2: "C \ UNIV // communicating \ a \ C \ (a, b) \ communicating \ b \ C" by (auto elim!: quotientE simp: communicating_def) lemma essential_class_iff_recurrent: "finite C \ C \ UNIV // communicating \ essential_class C \ (\x\C. recurrent x)" by (metis finite_essential_class_imp_recurrent irreducibleD2 recurrent_acc(4) essential_classI) definition "U' x y = (\\<^sup>+\. eSuc (sfirst (HLD {y}) \) \T x)" lemma U'_neq_zero[simp]: "U' x y \ 0" unfolding U'_def by (simp add: nn_integral_add) definition "gf_U' x y z = (\n. u x y n * Suc n * z ^ n)" definition "pos_recurrent x \ recurrent x \ U' x x \ \" lemma summable_gf_U': assumes z: "norm z < 1" shows "summable (\n. u x y n * Suc n * z ^ n)" proof - have "summable (\n. n * \z\ ^ n)" proof (rule root_test_convergence) have "(\n. root n n * \z\) \ 1 * \z\" by (intro tendsto_intros LIMSEQ_root) then show "(\n. root n (norm (n * \z\ ^ n))) \ \z\" by (rule filterlim_cong[THEN iffD1, rotated 3]) (auto intro!: exI[of _ 1] simp add: abs_mult u_nonneg real_root_mult power_abs eventually_sequentially real_root_power) qed (insert z, simp add: abs_less_iff) note summable_mult[OF this, of "1 / \z\"] from summable_ignore_initial_segment[OF this, of 1] show "summable (\n. u x y n * Suc n * z ^ n)" apply (rule summable_comparison_test[rotated]) using z apply (auto intro!: exI[of _ 1] simp: abs_mult u_nonneg power_abs Suc_le_eq gr0_conv_Suc field_simps le_divide_eq u_le_1 simp del: of_nat_Suc) done qed lemma gf_U'_nonneg[simp]: "0 < z \ z < 1 \ 0 \ gf_U' x y z" unfolding gf_U'_def by (intro suminf_nonneg summable_gf_U') (auto simp: u_nonneg) lemma DERIV_gf_U: fixes z :: real assumes z: "0 < z" "z < 1" shows "DERIV (gf_U x y) z :> gf_U' x y z" unfolding gf_U_def[abs_def] gf_U'_def real_scaleR_def u_def[symmetric] using z by (intro DERIV_power_series'[where R=1] summable_gf_U') auto lemma sfirst_finiteI_recurrent: "recurrent x \ (x, y) \ acc \ AE \ in T x. sfirst (HLD {y}) \ < \" using recurrent_acc(1)[of y x] recurrent_acc[of x y] T.AE_prob_1[of x "{\\space (T x). ev (HLD {y}) \}"] unfolding sfirst_finite U_def by (simp add: space_stream_space communicating_def) lemma U'_eq_suminf: assumes x: "recurrent x" "(x, y) \ acc" shows "U' x y = (\i. ennreal (u x y i * Suc i))" proof - have "(\\<^sup>+\. eSuc (sfirst (HLD {y}) \) \T x) = (\\<^sup>+\. (\i. ennreal (Suc i) * indicator {\\space (T y). ev_at (HLD {y}) i \} \) \T x)" using sfirst_finiteI_recurrent[OF x] proof (intro nn_integral_cong_AE, eventually_elim) fix \ assume "sfirst (HLD {y}) \ < \" then obtain n :: nat where [simp]: "sfirst (HLD {y}) \ = n" by auto show "eSuc (sfirst (HLD {y}) \) = (\i. ennreal (Suc i) * indicator {\\space (T y). ev_at (HLD {y}) i \} \)" by (subst suminf_cmult_indicator[where i=n]) (auto simp: disjoint_family_on_def ev_at_unique space_stream_space sfirst_eq_enat_iff[symmetric] ennreal_of_nat_eq_real_of_nat split: split_indicator) qed also have "\ = (\i. ennreal (Suc i) * emeasure (T x) {\\space (T x). ev_at (HLD {y}) i \})" by (subst nn_integral_suminf) (auto intro!: arg_cong[where f=suminf] nn_integral_cmult_indicator simp: fun_eq_iff) finally show ?thesis by (simp add: U'_def u_def T.emeasure_eq_measure mult_ac ennreal_mult) qed lemma gf_U'_tendsto_U': assumes x: "recurrent x" "(x, y) \ acc" shows "((\z. ennreal (gf_U' x y z)) \ U' x y) (at_left 1)" unfolding U'_eq_suminf[OF x] gf_U'_def by (auto intro!: power_series_tendsto_at_left summable_gf_U' mult_nonneg_nonneg u_nonneg simp del: of_nat_Suc) lemma one_le_integral_t: assumes x: "recurrent x" shows "1 \ U' x x" by (simp add: nn_integral_add T.emeasure_space_1 U'_def del: space_T) lemma gf_U'_pos: fixes z :: real assumes z: "0 < z" "z < 1" and "U x y \ 0" shows "0 < gf_U' x y z" unfolding gf_U'_def proof (subst suminf_pos_iff) show "summable (\n. u x y n * real (Suc n) * z ^ n)" using z by (intro summable_gf_U') simp show pos: "\n. 0 \ u x y n * real (Suc n) * z ^ n" using u_nonneg z by auto show "\n. 0 < u x y n * real (Suc n) * z ^ n" proof (rule ccontr) assume "\ (\n. 0 < u x y n * real (Suc n) * z ^ n)" with pos have "\n. u x y n * real (Suc n) * z ^ n = 0" by (intro antisym allI) (simp_all add: not_less) with z have "u x y = (\n. 0)" by (intro ext) simp with u_sums_U[of x y, THEN sums_unique] \U x y \ 0\ show False by simp qed qed lemma inverse_gf_U'_tendsto: assumes "recurrent y" shows "((\x. - 1 / - gf_U' y y x) \ enn2real (1 / U' y y)) (at_left (1::real))" proof cases assume inf: "U' y y = \" with gf_U'_tendsto_U'[of y y] \recurrent y\ have "LIM z (at_left 1). gf_U' y y z :> at_top" by (auto simp: ennreal_tendsto_top_eq_at_top U'_def) then have "LIM z (at_left 1). gf_U' y y z :> at_infinity" by (rule filterlim_mono) (auto simp: at_top_le_at_infinity) with inf show ?thesis by (auto intro!: tendsto_divide_0) next assume fin: "U' y y \ \" then obtain r where r: "U' y y = ennreal r" and [simp]: "0 \ r" by (cases "U' y y") (auto simp: U'_def) then have eq: "enn2real (1 / U' y y) = - 1 / - r" and "1 \ r" using one_le_integral_t[OF \recurrent y\] by (auto simp add: ennreal_1[symmetric] divide_ennreal simp del: ennreal_1) have "((\z. ennreal (gf_U' y y z)) \ ennreal r) (at_left 1)" using gf_U'_tendsto_U'[OF \recurrent y\, of y] r by simp then have gf_U': "(gf_U' y y \ r) (at_left (1::real))" by (rule tendsto_ennrealD) (insert summable_gf_U', auto intro!: eventually_at_left_1 suminf_nonneg simp: gf_U'_def u_nonneg) show ?thesis using \1 \ r\ unfolding eq by (intro tendsto_intros gf_U') simp qed lemma gf_G_pos: fixes z :: real assumes z: "0 < z" "z < 1" and *: "(x, y) \ acc" shows "0 < gf_G x y z" unfolding gf_G_def proof (subst suminf_pos_iff) show "summable (\n. p x y n *\<^sub>R z ^ n)" using z by (intro convergence_G convergence_G_less_1) simp show pos: "\n. 0 \ p x y n *\<^sub>R z ^ n" using z by (auto intro!: mult_nonneg_nonneg p_nonneg) show "\n. 0 < p x y n *\<^sub>R z ^ n" proof (rule ccontr) assume "\ (\n. 0 < p x y n *\<^sub>R z ^ n)" with pos have "\n. p x y n * z ^ n = 0" by (intro antisym allI) (simp_all add: not_less) with z have "\n. p x y n = 0" by simp with *[THEN accD_pos] show False by simp qed qed lemma pos_recurrentI_communicating: assumes y: "pos_recurrent y" and x: "(y, x) \ communicating" shows "pos_recurrent x" proof - from y x have recurrent: "recurrent y" "recurrent x" and fin: "U' y y \ \" by (auto simp: pos_recurrent_def recurrent_iffI_communicating nn_integral_add) have pos: "0 < enn2real (1 / U' y y)" using one_le_integral_t[OF \recurrent y\] fin by (auto simp: U'_def enn2real_positive_iff less_top[symmetric] ennreal_zero_less_divide ennreal_divide_eq_top_iff) from fin obtain r where r: "U' y y = ennreal r" and [simp]: "0 \ r" by (cases "U' y y") (auto simp: U'_def) from x obtain n m where "0 < p x y n" "0 < p y x m" by (auto dest!: accD_pos simp: communicating_def) let ?L = "at_left (1::real)" have le: "eventually (\z. p x y n * p y x m * z^(n + m) \ (1 - gf_U y y z) / (1 - gf_U x x z)) ?L" proof (rule eventually_at_left_1) fix z :: real assume z: "0 < z" "z < 1" then have conv: "\x. convergence_G x x z" by (intro convergence_G_less_1) simp have sums: "(\i. (p x y n * p y x m * z^(n + m)) * (p y y i * z^i)) sums ((p x y n * p y x m * z^(n + m)) * gf_G y y z)" unfolding gf_G_def by (intro sums_mult summable_sums) (auto intro!: conv convergence_G[where 'a=real, simplified]) have "(\i. (p x y n * p y x m * z^(n + m)) * (p y y i * z^i)) \ (\i. p x x (i + (n + m)) * z^(i + (n + m)))" proof (intro allI suminf_le sums_summable[OF sums] summable_ignore_initial_segment convergence_G[where 'a=real, simplified] convergence_G_less_1) show "norm z < 1" using z by simp fix i have "(p x y n * p y y ((n + i) - n)) * p y x ((n + i + m) - (n + i)) \ p x y (n + i) * p y x ((n + i + m) - (n + i))" by (intro mult_right_mono prob_reachable_le) simp_all also have "\ \ p x x (n + i + m)" by (intro prob_reachable_le) simp_all finally show "p x y n * p y x m * z ^ (n + m) * (p y y i * z ^ i) \ p x x (i + (n + m)) * z ^ (i + (n + m))" using z by (auto simp add: ac_simps power_add intro!: mult_left_mono) qed also have "\ \ gf_G x x z" unfolding gf_G_def using z apply (subst (2) suminf_split_initial_segment[where k="n + m"]) apply (intro convergence_G conv) apply (simp add: sum_nonneg) done finally have "(p x y n * p y x m * z^(n + m)) * gf_G y y z \ gf_G x x z" using sums_unique[OF sums] by simp then have "(p x y n * p y x m * z^(n + m)) \ gf_G x x z / gf_G y y z" using z gf_G_pos[of z y y] by (simp add: field_simps) also have "\ = (1 - gf_U y y z) / (1 - gf_U x x z)" unfolding gf_G_eq_gf_U[OF conv] using gf_G_eq_gf_U(2)[OF conv] by (simp add: field_simps ) finally show "p x y n * p y x m * z^(n + m) \ (1 - gf_U y y z) / (1 - gf_U x x z)" . qed have "U' x x \ \" proof assume "U' x x = \" have "((\z. (1 - gf_U y y z) / (1 - gf_U x x z)) \ 0) ?L" proof (rule lhopital_left) show "((\z. 1 - gf_U y y z) \ 0) ?L" using gf_U[of y] recurrent_iff_U_eq_1[of y] \recurrent y\ by (auto intro!: tendsto_eq_intros) show "((\z. 1 - gf_U x x z) \ 0) ?L" using gf_U[of x] recurrent_iff_U_eq_1[of x] \recurrent x\ by (auto intro!: tendsto_eq_intros) show "eventually (\z. 1 - gf_U x x z \ 0) ?L" by (auto intro!: eventually_at_left_1 simp: gf_G_eq_gf_U(2) convergence_G_less_1) show "eventually (\z. - gf_U' x x z \ 0) ?L" using gf_U'_pos[of _ x x] recurrent_iff_U_eq_1[of x] \recurrent x\ by (auto intro!: eventually_at_left_1) (metis less_le) show "eventually (\z. DERIV (\xa. 1 - gf_U x x xa) z :> - gf_U' x x z) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros DERIV_gf_U) show "eventually (\z. DERIV (\xa. 1 - gf_U y y xa) z :> - gf_U' y y z) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros DERIV_gf_U) have "(gf_U' y y \ U' y y) ?L" using \recurrent y\ by (rule gf_U'_tendsto_U') simp then have *: "(gf_U' y y \ r) ?L" by (auto simp add: r eventually_at_left_1 dest!: tendsto_ennrealD) moreover have "(gf_U' x x \ U' x x) ?L" using \recurrent x\ by (rule gf_U'_tendsto_U') simp then have "LIM z ?L. - gf_U' x x z :> at_bot" by (simp add: ennreal_tendsto_top_eq_at_top \U' x x = \\ filterlim_uminus_at_top del: ennreal_of_enat_eSuc) then have "LIM z ?L. - gf_U' x x z :> at_infinity" by (rule filterlim_mono) (auto simp: at_bot_le_at_infinity) ultimately show "((\z. - gf_U' y y z / - gf_U' x x z) \ 0) ?L" by (intro tendsto_divide_0[where c="- r"] tendsto_intros) qed moreover have "((\z. p x y n * p y x m * z^(n + m)) \ p x y n * p y x m) ?L" by (auto intro!: tendsto_eq_intros) ultimately have "p x y n * p y x m \ 0" using le by (rule tendsto_le[OF trivial_limit_at_left_real]) with \0 < p x y n\ \0 < p y x m\ show False by (auto simp add: mult_le_0_iff) qed with \recurrent x\ show ?thesis by (simp add: pos_recurrent_def nn_integral_add) qed lemma pos_recurrent_iffI_communicating: "(y, x) \ communicating \ pos_recurrent y \ pos_recurrent x" using pos_recurrentI_communicating[of x y] pos_recurrentI_communicating[of y x] by (auto simp add: communicating_def) lemma U_le_F: "U x y \ F x y" by (auto simp: U_def F_def intro!: T.finite_measure_mono) lemma not_empty_irreducible: "C \ UNIV // communicating \ C \ {}" by (auto simp: quotient_def Image_def communicating_def) subsection \Stationary distribution\ definition stat :: "'s set \ 's measure" where "stat C = point_measure UNIV (\x. indicator C x / U' x x)" lemma sets_stat[simp]: "sets (stat C) = sets (count_space UNIV)" by (simp add: stat_def sets_point_measure) lemma space_stat[simp]: "space (stat C) = UNIV" by (simp add: stat_def space_point_measure) lemma stat_subprob: assumes C: "essential_class C" and "countable C" and pos: "\c\C. pos_recurrent c" shows "emeasure (stat C) C \ 1" proof - let ?L = "at_left (1::real)" - from finite_sequence_to_countable_set[OF \countable C\] guess A . note A = this + from finite_sequence_to_countable_set[OF \countable C\] + obtain A where A: "\i. A i \ C" "\i. A i \ A (Suc i)" "\i. finite (A i)" "\ (range A) = C" + by blast then have "(\n. emeasure (stat C) (A n)) \ emeasure (stat C) (\i. A i)" by (intro Lim_emeasure_incseq) (auto simp: incseq_Suc_iff) then have "emeasure (stat C) (\i. A i) \ 1" proof (rule LIMSEQ_le[OF _ tendsto_const], intro exI allI impI) fix n from A(1,3) have A_n: "finite (A n)" by auto from C have "C \ {}" by (simp add: essential_class_def not_empty_irreducible) then obtain x where "x \ C" by auto have "((\z. (\y\A n. gf_F x y z * ((1 - z) / (1 - gf_U y y z)))) \ (\y\A n. F x y * enn2real (1 / U' y y))) ?L" proof (intro tendsto_intros gf_F, rule lhopital_left) fix y assume "y \ A n" with \A n \ C\ have "y \ C" by auto show "((-) 1 \ 0) ?L" by (intro tendsto_eq_intros) simp_all have "recurrent y" using pos[THEN bspec, OF \y\C\] by (simp add: pos_recurrent_def) then have "U y y = 1" by (simp add: recurrent_iff_U_eq_1) show "((\x. 1 - gf_U y y x) \ 0) ?L" using gf_U[of y y] \U y y = 1\ by (intro tendsto_eq_intros) auto show "eventually (\x. 1 - gf_U y y x \ 0) ?L" using gf_G_eq_gf_U(2)[OF convergence_G_less_1, where 'z=real] by (auto intro!: eventually_at_left_1) have "eventually (\x. 0 < gf_U' y y x) ?L" by (intro eventually_at_left_1 gf_U'_pos) (simp_all add: \U y y = 1\) then show "eventually (\x. - gf_U' y y x \ 0) ?L" by eventually_elim simp show "eventually (\x. DERIV (\x. 1 - gf_U y y x) x :> - gf_U' y y x) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros DERIV_gf_U) show "eventually (\x. DERIV ((-) 1) x :> - 1) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros) show "((\x. - 1 / - gf_U' y y x) \ enn2real (1 / U' y y)) ?L" using \recurrent y\ by (rule inverse_gf_U'_tendsto) qed also have "(\y\A n. F x y * enn2real (1 / U' y y)) = (\y\A n. enn2real (1 / U' y y))" proof (intro sum.cong refl) fix y assume "y \ A n" with \A n \ C\ have "y \ C" by auto with \x \ C\ have "(x, y) \ communicating" by (rule essential_classD3[OF C]) with \y\C\ have "recurrent y" "(y, x) \ acc" using pos[THEN bspec, of y] by (auto simp add: pos_recurrent_def communicating_def) then have "U x y = 1" by (rule recurrent_acc) with F_le_1[of x y] U_le_F[of x y] have "F x y = 1" by simp then show "F x y * enn2real (1 / U' y y) = enn2real (1 / U' y y)" by simp qed finally have le: "(\y\A n. enn2real (1 / U' y y)) \ 1" proof (rule tendsto_le[OF trivial_limit_at_left_real tendsto_const], intro eventually_at_left_1) fix z :: real assume z: "0 < z" "z < 1" with \x \ C\ have "norm z < 1" by auto then have conv: "\x y. convergence_G x y z" by (simp add: convergence_G_less_1) have "(\y\A n. gf_F x y z / (1 - gf_U y y z)) = (\y\A n. gf_G x y z)" using \norm z < 1\ apply (intro sum.cong refl) apply (subst gf_G_eq_gf_F) apply assumption apply (subst gf_G_eq_gf_U(1)[OF conv]) apply auto done also have "\ = (\y\A n. \n. p x y n * z^n)" by (simp add: gf_G_def) also have "\ = (\i. \y\A n. p x y i *\<^sub>R z^i)" by (subst suminf_sum[OF convergence_G[OF conv]]) simp also have "\ \ (\i. z^i)" proof (intro suminf_le summable_sum convergence_G conv summable_geometric allI) fix l have "(\y\A n. p x y l *\<^sub>R z ^ l) = (\y\A n. p x y l) * z ^ l" by (simp add: sum_distrib_right) also have "\ \ z ^ l" proof (intro mult_left_le_one_le) have "(\y\A n. p x y l) = \

(\ in T x. (x ## \) !! l \ A n)" unfolding p_def using \finite (A n)\ by (subst T.finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def intro!: arg_cong2[where f=measure]) then show "(\y\A n. p x y l) \ 1" by simp qed (insert z, auto simp: sum_nonneg) finally show "(\y\A n. p x y l *\<^sub>R z ^ l) \ z ^ l" . qed fact also have "\ = 1 / (1 - z)" using sums_unique[OF geometric_sums, OF \norm z < 1\] .. finally have "(\y\A n. gf_F x y z / (1 - gf_U y y z)) \ 1 / (1 - z)" . then have "(\y\A n. gf_F x y z / (1 - gf_U y y z)) * (1 - z) \ 1" using z by (simp add: field_simps) then have "(\y\A n. gf_F x y z / (1 - gf_U y y z) * (1 - z)) \ 1" by (simp add: sum_distrib_right) then show "(\y\A n. gf_F x y z * ((1 - z) / (1 - gf_U y y z))) \ 1" by simp qed from A_n have "emeasure (stat C) (A n) = (\y\A n. emeasure (stat C) {y})" by (intro emeasure_eq_sum_singleton) simp_all also have "\ = (\y\A n. inverse (U' y y))" unfolding stat_def U'_def using A(1)[of n] apply (intro sum.cong refl) apply (subst emeasure_point_measure_finite2) apply (auto simp: divide_ennreal_def Collect_conv_if) done also have "\ = ennreal (\y\A n. enn2real (1 / U' y y))" apply (subst sum_ennreal[symmetric], simp) proof (intro sum.cong refl) fix y assume "y \ A n" with \A n \ C\ pos have "pos_recurrent y" by auto with one_le_integral_t[of y] obtain r where "U' y y = ennreal r" "1 \ U' y y" and [simp]: "0 \ r" by (cases "U' y y") (auto simp: pos_recurrent_def nn_integral_add) then show "inverse (U' y y) = ennreal (enn2real (1 / U' y y))" by (simp add: ennreal_1[symmetric] divide_ennreal inverse_ennreal inverse_eq_divide del: ennreal_1) qed also have "\ \ 1" using le by simp finally show "emeasure (stat C) (A n) \ 1" . qed with A show ?thesis by simp qed lemma emeasure_stat_not_C: assumes "y \ C" shows "emeasure (stat C) {y} = 0" unfolding stat_def using \y \ C\ by (subst emeasure_point_measure_finite2) auto definition stationary_distribution :: "'s pmf \ bool" where "stationary_distribution N \ N = bind_pmf N K" lemma stationary_distributionI: assumes le: "\y. (\x. pmf (K x) y \measure_pmf N) \ pmf N y" shows "stationary_distribution N" unfolding stationary_distribution_def proof (rule pmf_eqI antisym)+ fix i show "pmf (bind_pmf N K) i \ pmf N i" by (simp add: pmf_bind le) define \ where "\ = N \ (\i\N. set_pmf (K i))" then have \: "countable \" by (auto intro: countable_set_pmf) then interpret N: sigma_finite_measure "count_space \" by (rule sigma_finite_measure_count_space_countable) interpret pN: pair_sigma_finite N "count_space \" by unfold_locales have measurable_pmf[measurable]: "(\(x, y). pmf (K x) y) \ borel_measurable (N \\<^sub>M count_space \)" unfolding measurable_split_conv apply (rule measurable_compose_countable'[OF _ measurable_snd]) apply (rule measurable_compose[OF measurable_fst]) apply (simp_all add: \) done { assume *: "(\y. pmf (K y) i \N) < pmf N i" have "0 \ (\y. pmf (K y) i \N)" by (intro integral_nonneg_AE) simp with * have i: "i \ set_pmf N" "i \ \" by (auto simp: set_pmf_iff \_def not_le[symmetric]) from * have "0 < pmf N i - (\y. pmf (K y) i \N)" by (simp add: field_simps) also have "\ = (\t. (pmf N i - (\y. pmf (K y) i \N)) * indicator {i} t \count_space \)" by (simp add: i) also have "\ \ (\t. pmf N t - \y. pmf (K y) t \N \count_space \)" using le by (intro integral_mono integrable_diff) (auto simp: i pmf_bind[symmetric] integrable_pmf field_simps split: split_indicator) also have "\ = (\t. pmf N t \count_space \) - (\t. \y. pmf (K y) t \N \count_space \)" by (subst Bochner_Integration.integral_diff) (auto intro!: integrable_pmf simp: pmf_bind[symmetric]) also have "(\t. \y. pmf (K y) t \N \count_space \) = (\y. \t. pmf (K y) t \count_space \ \N)" apply (intro pN.Fubini_integral integrable_iff_bounded[THEN iffD2] conjI) apply (auto simp add: N.nn_integral_fst[symmetric] nn_integral_eq_integral integrable_pmf) unfolding less_top[symmetric] unfolding infinity_ennreal_def[symmetric] apply (intro integrableD) apply (auto intro!: measure_pmf.integrable_const_bound[where B=1] simp: AE_measure_pmf_iff integral_nonneg_AE integral_pmf) done also have "(\y. \t. pmf (K y) t \count_space \ \N) = (\y. 1 \N)" by (intro integral_cong_AE) (auto simp: AE_measure_pmf_iff integral_pmf \_def intro!: measure_pmf.prob_eq_1[THEN iffD2]) finally have False using measure_pmf.prob_space[of N] by (simp add: integral_pmf field_simps not_le[symmetric]) } then show "pmf N i \ pmf (bind_pmf N K) i" by (auto simp: pmf_bind not_le[symmetric]) qed lemma stationary_distribution_iterate: assumes N: "stationary_distribution N" shows "ennreal (pmf N y) = (\\<^sup>+x. p x y n \N)" proof (induct n arbitrary: y) have [simp]: "\x y. ennreal (if x = y then 1 else 0) = indicator {y} x" by simp case 0 then show ?case by (simp add: p_0 pmf.rep_eq measure_pmf.emeasure_eq_measure) next case (Suc n) with N show ?case apply (simp add: nn_integral_eq_integral[symmetric] p_le_1 p_Suc' measure_pmf.integrable_const_bound[where B=1]) apply (subst nn_integral_bind[symmetric, where B="count_space UNIV"]) apply (auto simp: stationary_distribution_def measure_pmf_bind[symmetric] simp del: measurable_pmf_measure1) done qed lemma stationary_distribution_iterate': assumes "stationary_distribution N" shows "measure N {y} = (\x. p x y n \N)" using stationary_distribution_iterate[OF assms] by (subst (asm) nn_integral_eq_integral) (auto intro!: measure_pmf.integrable_const_bound[where B=1] simp: p_le_1 pmf.rep_eq) lemma stationary_distributionD: assumes C: "essential_class C" "countable C" assumes N: "stationary_distribution N" "N \ C" shows "\x\C. pos_recurrent x" "measure_pmf N = stat C" proof - have integrable_K: "\f x. integrable N (\s. pmf (K s) (f x))" by (rule measure_pmf.integrable_const_bound[where B=1]) (simp_all add: pmf_le_1) have measure_C: "measure N C = 1" and ae_C: "AE x in N. x \ C" using N C measure_pmf.prob_eq_1[of C] by (auto simp: AE_measure_pmf_iff) have integrable_p: "\n y. integrable N (\x. p x y n)" by (rule measure_pmf.integrable_const_bound[where B=1]) (simp_all add: p_le_1) { fix e :: real assume "0 < e" then have [simp]: "0 \ e" by simp have "\A\C. finite A \ 1 - e < measure N A" proof (rule ccontr) assume contr: "\ (\A \ C. finite A \ 1 - e < measure N A)" - from finite_sequence_to_countable_set[OF \countable C\] guess F . note F = this + from finite_sequence_to_countable_set[OF \countable C\] + obtain F where F: "\i. F i \ C" "\i. F i \ F (Suc i)" "\i. finite (F i)" "\ (range F) = C" + by blast then have *: "(\n. measure N (F n)) \ measure N (\i. F i)" by (intro measure_pmf.finite_Lim_measure_incseq) (auto simp: incseq_Suc_iff) with F contr have "measure N (\i. F i) \ 1 - e" by (intro LIMSEQ_le[OF * tendsto_const]) (auto simp: not_less) with F \0 < e\ show False by (simp add: measure_C) qed then obtain A where "A \ C" "finite A" and e: "1 - e < measure N A" by auto { fix y n assume "y \ C" from N(1) have "measure N {y} = (\x. p x y n \N)" by (rule stationary_distribution_iterate') also have "\ \ (\x. p x y n * indicator A x + indicator (C - A) x \N)" using ae_C \A \ C\ by (intro integral_mono_AE) (auto elim!: eventually_mono intro!: integral_add integral_indicator p_le_1 integrable_real_mult_indicator integrable_add split: split_indicator simp: integrable_p less_top[symmetric] top_unique) also have "\ = (\x. p x y n * indicator A x \N) + measure N (C - A)" using ae_C \A \ C\ apply (subst Bochner_Integration.integral_add) apply (auto elim!: eventually_mono intro!: integral_add integral_indicator p_le_1 integrable_real_mult_indicator split: split_indicator simp: integrable_p less_top[symmetric] top_unique) done also have "\ \ (\x. p x y n * indicator A x \N) + e" using e \A \ C\ by (simp add: measure_pmf.finite_measure_Diff measure_C) finally have "measure N {y} \ (\x. p x y n * indicator A x \N) + e" . then have "emeasure N {y} \ ennreal (\x. p x y n * indicator A x \N) + e" by (simp add: measure_pmf.emeasure_eq_measure ennreal_plus[symmetric] del: ennreal_plus) also have "\ = (\\<^sup>+x. ennreal (p x y n) * indicator A x \N) + e" by (subst nn_integral_eq_integral[symmetric]) (auto intro!: measure_pmf.integrable_const_bound[where B=1] simp: abs_mult p_le_1 mult_le_one ennreal_indicator ennreal_mult) finally have "emeasure N {y} \ (\\<^sup>+x. ennreal (p x y n) * indicator A x \N) + e" . } note v_le = this { fix y and z :: real assume y: "y \ C" and z: "0 < z" "z < 1" have summable_int_p: "summable (\n. (\ x. p x y n * indicator A x \N) * (1 - z) * z ^ n)" using \y\C\ z \A \ C\ by (auto intro!: summable_comparison_test[OF _ summable_mult[OF summable_geometric[of z], of 1]] exI[of _ 0] mult_le_one measure_pmf.integral_le_const integrable_real_mult_indicator integrable_p AE_I2 p_le_1 simp: abs_mult integral_nonneg_AE) from y z have sums_y: "(\n. measure N {y} * (1 - z) * z ^ n) sums measure N {y}" using sums_mult[OF geometric_sums[of z], of "measure N {y} * (1 - z)"] by simp then have "emeasure N {y} = ennreal (\n. (measure N {y} * (1 - z)) * z ^ n)" by (auto simp add: sums_unique[symmetric] measure_pmf.emeasure_eq_measure) also have "\ = (\n. emeasure N {y} * (1 - z) * z ^ n)" using z summable_mult[OF summable_geometric[of z], of "measure_pmf.prob N {y} * (1 - z)"] by (subst suminf_ennreal[symmetric]) (auto simp: measure_pmf.emeasure_eq_measure ennreal_mult[symmetric] ennreal_suminf_neq_top) also have "\ \ (\n. ((\\<^sup>+x. ennreal (p x y n) * indicator A x \N) + e) * (1 - z) * z ^ n)" using \y\C\ z \A \ C\ by (intro suminf_le mult_right_mono v_le allI) (auto simp: measure_pmf.emeasure_eq_measure) also have "\ = (\n. (\\<^sup>+x. ennreal (p x y n) * indicator A x \N) * (1 - z) * z ^ n) + e" using \0 < e\ z sums_mult[OF geometric_sums[of z], of "e * (1 - z)"] \0 \z<1\ by (simp add: distrib_right suminf_add[symmetric] ennreal_suminf_cmult[symmetric] ennreal_mult[symmetric] suminf_ennreal_eq sums_unique[symmetric] del: ennreal_suminf_cmult) also have "\ = (\n. ennreal (1 - z) * ((\\<^sup>+x. ennreal (p x y n) * indicator A x \N) * z ^ n)) + e" by (simp add: ac_simps) also have "\ = ennreal (1 - z) * (\n. ((\\<^sup>+x. ennreal (p x y n) * indicator A x \N) * z ^ n)) + e" using z by (subst ennreal_suminf_cmult) simp_all also have "(\n. ((\\<^sup>+x. ennreal (p x y n) * indicator A x \N) * z ^ n)) = (\n. (\\<^sup>+x. ennreal (p x y n * z ^ n) * indicator A x \N))" using z by (simp add: ac_simps nn_integral_cmult[symmetric] ennreal_mult) also have "\ = (\\<^sup>+x. ennreal (gf_G x y z) * indicator A x \N)" using z apply (subst nn_integral_suminf[symmetric]) apply (auto simp add: gf_G_def simp del: suminf_ennreal intro!: ennreal_mult_right_cong suminf_ennreal2 nn_integral_cong) apply (intro summable_comparison_test[OF _ summable_mult[OF summable_geometric[of z], of 1]] impI) apply (simp_all add: abs_mult p_le_1 mult_le_one power_le_one split: split_indicator) done also have "\ = (\\<^sup>+x. ennreal (gf_F x y z * gf_G y y z) * indicator A x \N)" using z by (intro nn_integral_cong) (simp add: gf_G_eq_gf_F[symmetric]) also have "\ = ennreal (gf_G y y z) * (\\<^sup>+x. ennreal (gf_F x y z) * indicator A x \N)" using z by (subst nn_integral_cmult[symmetric]) (simp_all add: gf_G_nonneg gf_F_nonneg ac_simps ennreal_mult) also have "\ = ennreal (1 / (1 - gf_U y y z)) * (\\<^sup>+x. ennreal (gf_F x y z) * indicator A x \N)" using z \y \ C\ by (subst gf_G_eq_gf_U) (auto intro!: convergence_G_less_1) finally have "emeasure N {y} \ ennreal ((1 - z) / (1 - gf_U y y z)) * (\\<^sup>+x. gf_F x y z * indicator A x \N) + e" using z by (subst (asm) mult.assoc[symmetric]) (simp add: ennreal_indicator[symmetric] ennreal_mult'[symmetric] gf_F_nonneg) then have "measure N {y} \ (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e" using z by (subst (asm) nn_integral_eq_integral[OF measure_pmf.integrable_const_bound[where B=1]]) (auto simp: gf_F_nonneg gf_U_le_1 gf_F_le_1 measure_pmf.emeasure_eq_measure mult_le_one ennreal_mult''[symmetric] ennreal_plus[symmetric] simp del: ennreal_plus) } then have "\A \ C. finite A \ (\y\C. \z. 0 < z \ z < 1 \ measure N {y} \ (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e)" using \A \ C\ \finite A\ by auto } note eps = this { fix y A assume "y \ C" "finite A" "A \ C" then have "((\z. \x. gf_F x y z * indicator A x \N) \ \x. F x y * indicator A x \N) (at_left 1)" by (subst (1 2) integral_measure_pmf[of A]) (auto intro!: tendsto_intros gf_F simp: indicator_eq_0_iff) } note int_gf_F = this have all_recurrent: "\y\C. recurrent y" proof (rule ccontr) assume "\ (\y\C. recurrent y)" then obtain x where "x \ C" "\ recurrent x" by auto then have transient: "\x. x \ C \ \ recurrent x" using C by (auto simp: essential_class_def recurrent_iffI_communicating[symmetric] elim!: quotientE) { fix y assume "y \ C" with transient have "U y y < 1" by (metis recurrent_iff_U_eq_1 U_cases) have "measure N {y} \ 0" proof (rule dense_ge) fix e :: real assume "0 < e" from eps[OF this] \y \ C\ obtain A where A: "finite A" "A \ C" and le: "\z. 0 < z \ z < 1 \ measure N {y} \ (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e" by auto have "((\z. (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e) \ (1 - 1) / (1 - U y y) * (\x. F x y * indicator A x \N) + e) (at_left (1::real))" using A \U y y < 1\ \y \ C\ by (intro tendsto_intros gf_U int_gf_F) auto then have 1: "((\z. (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e) \ e) (at_left (1::real))" by simp with le show "measure N {y} \ e" by (intro tendsto_le[OF trivial_limit_at_left_real _ tendsto_const]) (auto simp: eventually_at_left_1) qed then have "measure N {y} = 0" by (intro antisym measure_nonneg) } then have "emeasure N C = 0" by (subst emeasure_countable_singleton) (auto simp: measure_pmf.emeasure_eq_measure nn_integral_0_iff_AE ae_C C) then show False using \measure N C = 1\ by (simp add: measure_pmf.emeasure_eq_measure) qed then have "\x. x \ C \ U x x = 1" by (metis recurrent_iff_U_eq_1) { fix y assume "y \ C" then have "U y y = 1" "recurrent y" using \y \ C \ U y y = 1\ all_recurrent by auto have "measure N {y} \ enn2real (1 / U' y y)" proof (rule field_le_epsilon) fix e :: real assume "0 < e" from eps[OF \0 < e\] \y \ C\ obtain A where A: "finite A" "A \ C" and le: "\z. 0 < z \ z < 1 \ measure N {y} \ (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e" by auto let ?L = "at_left (1::real)" have "((\z. (1 - z) / (1 - gf_U y y z) * (\x. gf_F x y z * indicator A x \N) + e) \ enn2real (1 / U' y y) * (\x. F x y * indicator A x \N) + e) ?L" proof (intro tendsto_add tendsto_const tendsto_mult int_gf_F, rule lhopital_left[where f'="\x. - 1" and g'="\z. - gf_U' y y z"]) show "((-) 1 \ 0) ?L" "((\x. 1 - gf_U y y x) \ 0) ?L" using gf_U[of y y] by (auto intro!: tendsto_eq_intros simp: \U y y = 1\) show "y \ C" "finite A" "A \ C" by fact+ show "eventually (\x. 1 - gf_U y y x \ 0) ?L" using gf_G_eq_gf_U(2)[OF convergence_G_less_1, where 'z=real] by (auto intro!: eventually_at_left_1) show "((\x. - 1 / - gf_U' y y x) \ enn2real (1 / U' y y)) ?L" using \recurrent y\ by (rule inverse_gf_U'_tendsto) have "eventually (\x. 0 < gf_U' y y x) ?L" by (intro eventually_at_left_1 gf_U'_pos) (simp_all add: \U y y = 1\) then show "eventually (\x. - gf_U' y y x \ 0) ?L" by eventually_elim simp show "eventually (\x. DERIV (\x. 1 - gf_U y y x) x :> - gf_U' y y x) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros DERIV_gf_U) show "eventually (\x. DERIV ((-) 1) x :> - 1) ?L" by (auto intro!: eventually_at_left_1 derivative_eq_intros) qed then have "measure N {y} \ enn2real (1 / U' y y) * (\x. F x y * indicator A x \N) + e" by (rule tendsto_le[OF trivial_limit_at_left_real _ tendsto_const]) (intro eventually_at_left_1 le) then have "measure N {y} - e \ enn2real (1 / U' y y) * (\x. F x y * indicator A x \N)" by simp also have "\ \ enn2real (1 / U' y y)" using A by (intro mult_left_le measure_pmf.integral_le_const measure_pmf.integrable_const_bound[where B=1]) (auto simp: mult_le_one F_le_1 U'_def) finally show "measure N {y} \ enn2real (1 / U' y y) + e" by simp qed } note measure_y_le = this show pos: "\y\C. pos_recurrent y" proof (rule ccontr) assume "\ (\y\C. pos_recurrent y)" then obtain x where x: "x \ C" "\ pos_recurrent x" by auto { fix y assume "y \ C" with x have "\ pos_recurrent y" using C by (auto simp: essential_class_def pos_recurrent_iffI_communicating[symmetric] elim!: quotientE) with all_recurrent \y \ C\ have "enn2real (1 / U' y y) = 0" by (simp add: pos_recurrent_def nn_integral_add) with measure_y_le[OF \y \ C\] have "measure N {y} = 0" by (auto intro!: antisym simp: pos_recurrent_def) } then have "emeasure N C = 0" by (subst emeasure_countable_singleton) (auto simp: C ae_C measure_pmf.emeasure_eq_measure nn_integral_0_iff_AE) then show False using \measure N C = 1\ by (simp add: measure_pmf.emeasure_eq_measure) qed { fix A :: "'s set" assume [simp]: "countable A" have "emeasure N A = (\\<^sup>+x. emeasure N {x} \count_space A)" by (intro emeasure_countable_singleton) auto also have "\ \ (\\<^sup>+x. emeasure (stat C) {x} \count_space A)" proof (intro nn_integral_mono) fix y assume "y \ space (count_space A)" show "emeasure N {y} \ emeasure (stat C) {y}" proof cases assume "y \ C" with pos have "pos_recurrent y" by auto with one_le_integral_t[of y] obtain r where r: "U' y y = ennreal r" "1 \ U' y y" and [simp]: "0 \ r" by (cases "U' y y") (auto simp: pos_recurrent_def nn_integral_add) from measure_y_le[OF \y \ C\] have "emeasure N {y} \ ennreal (enn2real (1 / U' y y))" by (simp add: measure_pmf.emeasure_eq_measure) also have "\ = emeasure (stat C) {y}" unfolding stat_def using \y \ C\ r by (subst emeasure_point_measure_finite2) (auto simp add: ennreal_1[symmetric] divide_ennreal inverse_ennreal inverse_eq_divide ennreal_mult[symmetric] simp del: ennreal_1) finally show "emeasure N {y} \ emeasure (stat C) {y}" by simp next assume "y \ C" with ae_C have "emeasure N {y} = 0" by (subst AE_iff_measurable[symmetric, where P="\x. x \ y"]) (auto elim!: eventually_mono) moreover have "emeasure (stat C) {y} = 0" using emeasure_stat_not_C[OF \y \ C\] . ultimately show ?thesis by simp qed qed also have "\ = emeasure (stat C) A" by (intro emeasure_countable_singleton[symmetric]) auto finally have "emeasure N A \ emeasure (stat C) A" . } note N_le_C = this from stat_subprob[OF C(1) \countable C\ pos] N_le_C[OF \countable C\] \measure N C = 1\ have stat_C_eq_1: "emeasure (stat C) C = 1" by (auto simp add: measure_pmf.emeasure_eq_measure one_ennreal_def) moreover have "emeasure (stat C) (UNIV - C) = 0" by (subst AE_iff_measurable[symmetric, where P="\x. x \ C"]) (auto simp: stat_def AE_point_measure sets_point_measure space_point_measure split: split_indicator cong del: AE_cong) ultimately have "emeasure (stat C) (space (stat C)) = 1" using plus_emeasure[of C "stat C" "UNIV - C"] by (simp add: Un_absorb1) interpret stat: prob_space "stat C" by standard fact show "measure_pmf N = stat C" proof (rule measure_eqI_countable_AE) show "sets N = UNIV" "sets (stat C) = UNIV" by auto show "countable C" "AE x in N. x \ C" and ae_stat: "AE x in stat C. x \ C" using C ae_C stat_C_eq_1 by (auto intro!: stat.AE_prob_1 simp: stat.emeasure_eq_measure) { assume "\x. emeasure N {x} \ emeasure (stat C) {x}" then obtain x where [simp]: "emeasure N {x} \ emeasure (stat C) {x}" by auto with N_le_C[of "{x}"] have x: "emeasure N {x} < emeasure (stat C) {x}" by (auto simp: less_le) have "1 = emeasure N {x} + emeasure N (C - {x})" using ae_C by (subst plus_emeasure) (auto intro!: measure_pmf.emeasure_eq_1_AE) also have "\ < emeasure (stat C) {x} + emeasure (stat C) (C - {x})" using x N_le_C[of "C - {x}"] C ae_C by (simp add: stat.emeasure_eq_measure measure_pmf.emeasure_eq_measure ennreal_plus[symmetric] ennreal_less_iff del: ennreal_plus) also have "\ = 1" using ae_stat by (subst plus_emeasure) (auto intro!: stat.emeasure_eq_1_AE) finally have False by simp } then show "\x. emeasure N {x} = emeasure (stat C) {x}" by auto qed qed lemma measure_point_measure_singleton: "x \ A \ measure (point_measure A X) {x} = enn2real (X x)" unfolding measure_def by (subst emeasure_point_measure_finite2) auto lemma stationary_distribution_imp_int_t: assumes C: "essential_class C" "countable C" "stationary_distribution N" "N \ C" assumes x: "x \ C" shows "U' x x = 1 / ennreal (pmf N x)" proof - from stationary_distributionD[OF C] have "measure_pmf N = stat C" and *: "\x\C. pos_recurrent x" by auto show ?thesis unfolding \measure_pmf N = stat C\ pmf.rep_eq stat_def using *[THEN bspec, OF x] x apply (simp add: measure_point_measure_singleton) apply (cases "U' x x") subgoal for r by (cases "r = 0") (simp_all add: divide_ennreal_def inverse_ennreal) apply simp done qed definition "period_set x = {i. 0 < i \ 0 < p x x i }" definition "period C = (SOME d. \x\C. d = Gcd (period_set x))" lemma Gcd_period_set_invariant: assumes c: "(x, y) \ communicating" shows "Gcd (period_set x) = Gcd (period_set y)" proof - { fix x y n assume c: "(x, y) \ communicating" "x \ y" and n: "n \ period_set x" from c obtain l k where "0 < p x y l" "0 < p y x k" by (auto simp: communicating_def dest!: accD_pos) moreover with \x \ y\ have "l \ 0 \ k \ 0" by (intro notI conjI) (auto simp: p_0) ultimately have pos: "0 < l" "0 < k" and l: "0 < p x y l" and k: "0 < p y x k" by auto from mult_pos_pos[OF k l] prob_reachable_le[of k "k + l" y x y] c have k_l: "0 < p y y (k + l)" by simp then have "Gcd (period_set y) dvd k + l" using pos by (auto intro!: Gcd_dvd_nat simp: period_set_def) moreover from n have "0 < p x x n" "0 < n" by (auto simp: period_set_def) from mult_pos_pos[OF k this(1)] prob_reachable_le[of k "k + n" y x x] c have "0 < p y x (k + n)" by simp from mult_pos_pos[OF this(1) l] prob_reachable_le[of "k + n" "(k + n) + l" y x y] c have "0 < p y y (k + n + l)" by simp then have "Gcd (period_set y) dvd (k + l) + n" using pos by (auto intro!: Gcd_dvd_nat simp: period_set_def ac_simps) ultimately have "Gcd (period_set y) dvd n" by (metis dvd_add_left_iff add.commute) } note this[of x y] this[of y x] c moreover have "(y, x) \ communicating" using c by (simp add: communicating_def) ultimately show ?thesis by (auto intro: dvd_antisym Gcd_greatest Gcd_dvd) qed lemma period_eq: assumes "C \ UNIV // communicating" "x \ C" shows "period C = Gcd (period_set x)" unfolding period_def using assms by (rule_tac someI2[where a="Gcd (period_set x)"]) (auto intro!: Gcd_period_set_invariant irreducibleD) definition "aperiodic C \ C \ UNIV // communicating \ period C = 1" definition "not_ephemeral C \ C \ UNIV // communicating \ \ (\x. C = {x} \ p x x 1 = 0)" lemma not_ephemeralD: assumes C: "not_ephemeral C" "x \ C" shows "\n>0. 0 < p x x n" proof cases assume "\x. C = {x}" with \x \ C\ have "C = {x}" by auto with C p_nonneg[of x x 1] have "0 < p x x 1" by (auto simp: not_ephemeral_def less_le) with \C = {x}\ show ?thesis by auto next from C have irr: "C \ UNIV // communicating" by (auto simp: not_ephemeral_def) assume "\(\x. C = {x})" then have "\x. C \ {x}" by auto with \x \ C\ obtain y where "y \ C" "x \ y" by blast with irreducibleD[OF irr, of x y] C \x \ C\ have c: "(x, y) \ communicating" by auto with accD_pos[of x y] accD_pos[of y x] obtain k l where pos: "0 < p x y k" "0 < p y x l" by (auto simp: communicating_def) with \x \ y\ have "l \ 0" by (intro notI) (auto simp: p_0) have "0 < p x y k * p y x (k + l - k)" using pos by auto also have "p x y k * p y x (k + l - k) \ p x x (k + l)" using prob_reachable_le[of "k" "k + l" x y x] c by auto finally show ?thesis using \l \ 0\ \x \ C\ by (auto intro!: exI[of _ "k + l"]) qed lemma not_ephemeralD_pos_period: assumes C: "not_ephemeral C" shows "0 < period C" proof - from C not_empty_irreducible[of C] obtain x where "x \ C" by (auto simp: not_ephemeral_def) from not_ephemeralD[OF C this] obtain n where n: "0 < p x x n" "0 < n" by auto have C': "C \ UNIV // communicating" using C by (auto simp: not_ephemeral_def) have "period C \ 0" unfolding period_eq [OF C' \x \ C\] using n by (auto simp: period_set_def) then show ?thesis by auto qed lemma period_posD: assumes C: "C \ UNIV // communicating" and "0 < period C" "x \ C" shows "\n>0. 0 < p x x n" proof - from \0 < period C\ have "period C \ 0" by auto then show ?thesis unfolding period_eq [OF C \x \ C\] unfolding period_set_def by auto qed lemma not_ephemeralD_pos_period': assumes C: "C \ UNIV // communicating" shows "not_ephemeral C \ 0 < period C" proof (auto dest!: not_ephemeralD_pos_period intro: C) from C not_empty_irreducible[of C] obtain x where "x \ C" by (auto simp: not_ephemeral_def) assume "0 < period C" then show "not_ephemeral C" apply (auto simp: not_ephemeral_def C) oops \ \should be easy to finish\ lemma eventually_periodic: assumes C: "C \ UNIV // communicating" "0 < period C" "x \ C" shows "eventually (\m. 0 < p x x (m * period C)) sequentially" proof - from period_posD[OF assms] obtain n where n: "0 < p x x n" "0 < n" by auto have C': "C \ UNIV // communicating" using C by auto have "period C \ 0" unfolding period_eq [OF C' \x \ C\] using n by (auto simp: period_set_def) have "eventually (\m. m * Gcd (period_set x) \ (period_set x)) sequentially" proof (rule eventually_mult_Gcd) show "n > 0" "n \ period_set x" using n by (auto simp add: period_set_def) fix k l assume "k \ period_set x" "l \ period_set x" then have "0 < p x x k * p x x l" "0 < l" "0 < k" by (auto simp: period_set_def) moreover have "p x x k * p x x l \ p x x (k + l)" using prob_reachable_le[of k "k + l" x x x] \x \ C\ by auto ultimately show "k + l \ period_set x" using \0 < l\ by (auto simp: period_set_def) qed with eventually_ge_at_top[of 1] show "eventually (\m. 0 < p x x (m * period C)) sequentially" by eventually_elim (insert \period C \ 0\ period_eq[OF C' \x \ C\, symmetric], auto simp: period_set_def) qed lemma aperiodic_eventually_recurrent: "aperiodic C \ C \ UNIV // communicating \ (\x\C. eventually (\m. 0 < p x x m) sequentially)" proof safe fix x assume "x \ C" "aperiodic C" with eventually_periodic[of C x] show "eventually (\m. 0 < p x x m) sequentially" by (auto simp add: aperiodic_def) next assume "\x\C. eventually (\m. 0 < p x x m) sequentially" and C: "C \ UNIV // communicating" moreover from not_empty_irreducible[OF C] obtain x where "x \ C" by auto ultimately obtain N where "\M. M\N \ 0 < p x x M" by (auto simp: eventually_sequentially) then have "{N <..} \ period_set x" by (auto simp: period_set_def) from C show "aperiodic C" unfolding period_eq [OF C \x \ C\] aperiodic_def proof show "Gcd (period_set x) = 1" proof (rule Gcd_eqI) from one_dvd show "1 dvd q" for q :: nat . fix m assume "\q. q \ period_set x \ m dvd q" moreover from \{N <..} \ period_set x\ have "{Suc N, Suc (Suc N)} \ period_set x" by auto ultimately have "m dvd Suc (Suc N)" and "m dvd Suc N" by auto then have "m dvd Suc (Suc N) - Suc N" by (rule dvd_diff_nat) then show "is_unit m" by simp qed simp qed qed (simp add: aperiodic_def) lemma stationary_distributionD_emeasure: assumes N: "stationary_distribution N" shows "emeasure N A = (\\<^sup>+s. emeasure (K s) A \N)" proof - have "prob_space (measure_pmf N)" by intro_locales then interpret subprob_space "measure_pmf N" by (rule prob_space_imp_subprob_space) show ?thesis unfolding measure_pmf.emeasure_eq_measure apply (subst N[unfolded stationary_distribution_def]) apply (simp add: measure_pmf_bind) apply (subst measure_pmf.measure_bind[where N="count_space UNIV"]) apply (rule measurable_compose[OF _ measurable_measure_pmf]) apply (auto intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1]) done qed lemma communicatingD1: "C \ UNIV // communicating \ (a, b) \ communicating \ a \ C \ b \ C" by (auto elim!: quotientE) (auto simp add: communicating_def) lemma communicatingD2: "C \ UNIV // communicating \ (a, b) \ communicating \ b \ C \ a \ C" by (auto elim!: quotientE) (auto simp add: communicating_def) lemma acc_iff: "(x, y) \ acc \ (\n. 0 < p x y n)" by (blast intro: accD_pos accI_pos) lemma communicating_iff: "(x, y) \ communicating \ (\n. 0 < p x y n) \ (\n. 0 < p y x n)" by (auto simp add: acc_iff communicating_def) end context MC_pair begin lemma p_eq_p1_p2: "p (x1, x2) (y1, y2) n = K1.p x1 y1 n * K2.p x2 y2 n" unfolding p_def K1.p_def K2.p_def by (subst prod_eq_prob_T) (auto intro!: arg_cong2[where f=measure] split: nat.splits simp: Stream_snth) lemma P_accD: assumes "((x1, x2), (y1, y2)) \ acc"shows "(x1, y1) \ K1.acc" "(x2, y2) \ K2.acc" using assms by (auto simp: acc_iff K1.acc_iff K2.acc_iff p_eq_p1_p2 zero_less_mult_iff not_le[of 0, symmetric] cong: conj_cong) lemma aperiodicI_pair: assumes C1: "K1.aperiodic C1" and C2: "K2.aperiodic C2" shows "aperiodic (C1 \ C2)" unfolding aperiodic_eventually_recurrent proof safe from C1[unfolded K1.aperiodic_eventually_recurrent] C2[unfolded K2.aperiodic_eventually_recurrent] have C1: "C1 \ UNIV // K1.communicating" and C2: "C2 \ UNIV // K2.communicating" and ev: "\x. x \ C1 \ eventually (\m. 0 < K1.p x x m) sequentially" "\x. x \ C2 \ eventually (\m. 0 < K2.p x x m) sequentially" by auto { fix x1 x2 assume x: "x1 \ C1" "x2 \ C2" from ev(1)[OF x(1)] ev(2)[OF x(2)] show "eventually (\m. 0 < p (x1, x2) (x1, x2) m) sequentially" by eventually_elim (simp add: p_eq_p1_p2 x) } { fix x1 x2 y1 y2 assume acc: "(x1, y1) \ K1.acc" "(x2, y2) \ K2.acc" "x1 \ C1" "y1 \ C1" "x2 \ C2" "y2 \ C2" then obtain k l where "0 < K1.p x1 y1 l" "0 < K2.p x2 y2 k" by (auto dest!: K1.accD_pos K2.accD_pos) with acc ev(1)[of y1] ev(2)[of y2] have "eventually (\m. 0 < K1.p x1 y1 l * K1.p y1 y1 m \ 0 < K2.p x2 y2 k * K2.p y2 y2 m) sequentially" by (auto elim: eventually_elim2) then have "eventually (\m. 0 < K1.p x1 y1 (m + l) \ 0 < K2.p x2 y2 (m + k)) sequentially" proof eventually_elim fix m assume "0 < K1.p x1 y1 l * K1.p y1 y1 m \ 0 < K2.p x2 y2 k * K2.p y2 y2 m" with acc K1.prob_reachable_le[of l "l + m" x1 y1 y1] K2.prob_reachable_le[of k "k + m" x2 y2 y2] show "0 < K1.p x1 y1 (m + l) \ 0 < K2.p x2 y2 (m + k)" by (auto simp add: ac_simps) qed then have "eventually (\m. 0 < K1.p x1 y1 m \ 0 < K2.p x2 y2 m) sequentially" unfolding eventually_conj_iff by (subst (asm) (1 2) eventually_sequentially_seg) (auto elim: eventually_elim2) then obtain N where "0 < K1.p x1 y1 N" "0 < K2.p x2 y2 N" by (auto simp: eventually_sequentially) with acc have "0 < p (x1, x2) (y1, y2) N" by (auto simp add: p_eq_p1_p2) with acc have "((x1, x2), (y1, y2)) \ acc" by (auto intro!: accI_pos) } note 1 = this { fix x1 x2 y1 y2 assume acc:"((x1, x2), (y1, y2)) \ acc" moreover from acc obtain k where "0 < p (x1, x2) (y1, y2) k" by (auto dest!: accD_pos) ultimately have "(x1, y1) \ K1.acc \ (x2, y2) \ K2.acc" by (subst (asm) p_eq_p1_p2) (auto intro!: K1.accI_pos K2.accI_pos simp: zero_less_mult_iff not_le[of 0, symmetric]) } note 2 = this from K1.not_empty_irreducible[OF C1] K2.not_empty_irreducible[OF C2] obtain x1 x2 where xC: "x1 \ C1" "x2 \ C2" by auto show "C1 \ C2 \ UNIV // communicating" apply (simp add: quotient_def Image_def) apply (safe intro!: exI[of _ x1] exI[of _ x2]) proof - fix y1 y2 assume yC: "y1 \ C1" "y2 \ C2" from K1.irreducibleD[OF C1 \x1 \ C1\ \y1 \ C1\] K2.irreducibleD[OF C2 \x2 \ C2\ \y2 \ C2\] show "((x1, x2), (y1, y2)) \ communicating" using 1[of x1 y1 x2 y2] 1[of y1 x1 y2 x2] xC yC by (auto simp: communicating_def K1.communicating_def K2.communicating_def) next fix y1 y2 assume "((x1, x2), (y1, y2)) \ communicating" with 2[of x1 x2 y1 y2] 2[of y1 y2 x1 x2] have "(x1, y1) \ K1.communicating" "(x2, y2) \ K2.communicating" by (auto simp: communicating_def K1.communicating_def K2.communicating_def) with xC show "y1 \ C1" "y2 \ C2" using K1.communicatingD1[OF C1] K2.communicatingD1[OF C2] by auto qed qed lemma stationary_distributionI_pair: assumes N1: "K1.stationary_distribution N1" assumes N2: "K2.stationary_distribution N2" shows "stationary_distribution (pair_pmf N1 N2)" unfolding stationary_distribution_def unfolding Kp_def pair_pmf_def apply (subst N1[unfolded K1.stationary_distribution_def]) apply (subst N2[unfolded K2.stationary_distribution_def]) apply (simp add: bind_assoc_pmf bind_return_pmf) apply (subst bind_commute_pmf[of N2]) apply simp done end context MC_syntax begin lemma stationary_distribution_imp_limit: assumes C: "aperiodic C" "essential_class C" "countable C" and N: "stationary_distribution N" "N \ C" assumes [simp]: "y \ C" shows "(\n. \x. \p y x n - pmf N x\ \count_space C) \ 0" (is "?L \ 0") proof - from \essential_class C\ have C_comm: "C \ UNIV // communicating" by (simp add: essential_class_def) define K' where "K' = (\Some x \ map_pmf Some (K x) | None \ map_pmf Some N)" interpret K2: MC_syntax K' . interpret KN: MC_pair K K' . from stationary_distributionD[OF C(2,3) N] have pos: "\x. x \ C \ pos_recurrent x" and "measure_pmf N = stat C" by auto have pos: "\x. x \ C \ 0 < emeasure N {x}" using pos unfolding stat_def \measure_pmf N = stat C\ by (subst emeasure_point_measure_finite2) (auto simp: U'_def pos_recurrent_def nn_integral_add ennreal_zero_less_divide less_top) then have rpos: "\x. x \ C \ 0 < pmf N x" by (simp add: measure_pmf.emeasure_eq_measure pmf.rep_eq) have eq: "\x y. (if x = y then 1 else 0) = indicator {y} x" by auto have intK: "\f x. (\x. (f x :: real) \K' (Some x)) = (\x. f (Some x) \K x)" by (simp add: K'_def integral_distr map_pmf_rep_eq) { fix m and x y :: 's have "K2.p (Some x) (Some y) m = p x y m" by (induct m arbitrary: x) (auto intro!: integral_cong simp add: K2.p_Suc' p_Suc' intK K2.p_0 p_0) } note K_p_eq = this { fix n and x :: 's have "K2.p (Some x) None n = 0" by (induct n arbitrary: x) (auto simp: K2.p_Suc' K2.p_0 intK cong: integral_cong) } note K_S_None = this from not_empty_irreducible[OF C_comm] obtain c0 where c0: "c0 \ C" by auto have K2_acc: "\x y. (Some x, y) \ K2.acc \ (\z. y = Some z \ (x, z) \ acc)" apply (auto simp: K2.acc_iff acc_iff K_p_eq) apply (case_tac y) apply (auto simp: K_p_eq K_S_None) done have K2_communicating: "\c x. c \ C \ (Some c, x) \ K2.communicating \ (\c'\C. x = Some c')" proof safe fix x c assume "c \ C" "(Some c, x) \ K2.communicating" then show "\c'\C. x = Some c'" by (cases x) (auto simp: communicating_iff K2.communicating_iff K_p_eq K_S_None intro!: irreducibleD2[OF C_comm \c\C\]) next fix c c' x assume "c \ C" "c' \ C" with irreducibleD[OF C_comm this] show "(Some c, Some c') \ K2.communicating" by (auto simp: K2.communicating_iff communicating_iff K_p_eq) qed have "Some ` C \ UNIV // K2.communicating" by (auto simp add: quotient_def Image_def c0 K2_communicating intro!: exI[of _ "Some c0"]) then have "K2.essential_class (Some ` C)" by (rule K2.essential_classI) (auto simp: K2_acc essential_classD2[OF \essential_class C\]) have "K2.aperiodic (Some ` C)" unfolding K2.aperiodic_eventually_recurrent proof safe fix x assume "x \ C" then show "eventually (\m. 0 < K2.p (Some x) (Some x) m) sequentially" using \aperiodic C\ unfolding aperiodic_eventually_recurrent by (auto elim!: eventually_mono simp: K_p_eq) qed fact then have aperiodic: "KN.aperiodic (C \ Some ` C)" by (rule KN.aperiodicI_pair[OF \aperiodic C\]) have KN_essential: "KN.essential_class (C \ Some ` C)" proof (rule KN.essential_classI) show "C \ Some ` C \ UNIV // KN.communicating" using aperiodic by (simp add: KN.aperiodic_def) next fix x y assume "x \ C \ Some ` C" "(x, y) \ KN.acc" with KN.P_accD[of "fst x" "snd x" "fst y" "snd y"] show "y \ C \ Some ` C" by (cases x y rule: prod.exhaust[case_product prod.exhaust]) (auto simp: K2_acc essential_classD2[OF \essential_class C\]) qed { fix n and x y :: 's have "measure N {y} = \

(\ in K2.T None. (None ## \) !! (Suc n) = Some y)" unfolding stationary_distribution_iterate'[OF N(1), of y n] apply (subst K2.p_def[symmetric]) apply (subst K2.p_Suc') apply (subst K'_def) apply (simp add: map_pmf_rep_eq integral_distr K_p_eq) done then have "measure N {y} = \

(\ in K2.T None. \ !! n = Some y)" by simp } note measure_y_eq = this define D where "D = {x::'s \ 's option. Some (fst x) = snd x}" have [measurable]: "\P::('s \ 's option \ bool). P \ measurable (count_space UNIV) (count_space UNIV)" by simp { fix n and x :: 's have "\

(\ in KN.T (y, None). \i !! n) = Some x \ ev_at (HLD D) i \) = (\i(\ in KN.T (y, None). snd (\ !! n) = Some x \ ev_at (HLD D) i \))" by (subst KN.T.finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def intro!: arg_cong2[where f=measure] dest: ev_at_unique) also have "\ = (\i(\ in KN.T (y, None). fst (\ !! n) = x \ ev_at (HLD D) i \))" proof (intro sum.cong refl) fix i assume i: "i \ {..< n}" show "\

(\ in KN.T (y, None). snd (\ !! n) = Some x \ ev_at (HLD D) i \) = \

(\ in KN.T (y, None). fst (\ !! n) = x \ ev_at (HLD D) i \)" apply (subst (1 2) KN.prob_T_split[where n="Suc i"]) apply (simp_all add: ev_at_shift snth_Stream del: stake.simps KN.space_T) unfolding ev_at_shift snth_Stream proof (intro Bochner_Integration.integral_cong refl) fix \ :: "('s \ 's option) stream" let ?s = "\\'. stake (Suc i) \ @- \'" show "\

(\' in KN.T (\ !! i). snd (?s \' !! n) = Some x \ ev_at (HLD D) i \) = \

(\' in KN.T (\ !! i). fst (?s \' !! n) = x \ ev_at (HLD D) i \)" proof cases assume "ev_at (HLD D) i \" from ev_at_imp_snth[OF this] have eq: "snd (\ !! i) = Some (fst (\ !! i))" by (simp add: D_def HLD_iff) have "\

(\' in KN.T (\ !! i). fst (\' !! (n - Suc i)) = x) = \

(\' in T (fst (\ !! i)). \' !! (n - Suc i) = x) * \

(\' in K2.T (snd (\ !! i)). True)" by (subst KN.prod_eq_prob_T) simp_all also have "\ = p (fst (\ !! i)) x (Suc (n - Suc i))" using K2.T.prob_space by (simp add: p_def) also have "\ = K2.p (snd (\ !! i)) (Some x) (Suc (n - Suc i))" by (simp add: K_p_eq eq) also have "\ = \

(\' in T (fst (\ !! i)). True) * \

(\' in K2.T (snd (\ !! i)). \' !! (n - Suc i) = Some x)" using T.prob_space by (simp add: K2.p_def) also have "\ = \

(\' in KN.T (\ !! i). snd (\' !! (n - Suc i)) = Some x)" by (subst KN.prod_eq_prob_T) simp_all finally show ?thesis using \ev_at (HLD D) i \\ i by (simp del: stake.simps) qed simp qed qed also have "\ = \

(\ in KN.T (y, None). (\i !! n) = x \ ev_at (HLD D) i \))" by (subst KN.T.finite_measure_finite_Union[symmetric]) (auto simp add: disjoint_family_on_def dest: ev_at_unique intro!: arg_cong2[where f=measure]) finally have eq: "\

(\ in KN.T (y, None). (\i !! n) = Some x \ ev_at (HLD D) i \)) = \

(\ in KN.T (y, None). (\i !! n) = x \ ev_at (HLD D) i \))" . have "p y x (Suc n) - measure N {x} = \

(\ in T y. \ !! n = x) - \

(\ in K2.T None. \ !! n = Some x)" unfolding p_def by (subst measure_y_eq) simp_all also have "\

(\ in T y. \ !! n = x) = \

(\ in T y. \ !! n = x) * \

(\ in K2.T None. True)" using K2.T.prob_space by simp also have "\ = \

(\ in KN.T (y, None). fst (\ !! n) = x)" by (subst KN.prod_eq_prob_T) auto also have "\ = \

(\ in KN.T (y, None). (\i !! n) = x \ ev_at (HLD D) i \)) + \

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i))" by (subst KN.T.finite_measure_Union[symmetric]) (auto intro!: arg_cong2[where f=measure]) also have "\

(\ in K2.T None. \ !! n = Some x) = \

(\ in T y. True) * \

(\ in K2.T None. \ !! n = Some x)" using T.prob_space by simp also have "\ = \

(\ in KN.T (y, None). snd (\ !! n) = Some x)" by (subst KN.prod_eq_prob_T) auto also have "\ = \

(\ in KN.T (y, None). (\i !! n) = Some x \ ev_at (HLD D) i \)) + \

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i))" by (subst KN.T.finite_measure_Union[symmetric]) (auto intro!: arg_cong2[where f=measure]) finally have "\ p y x (Suc n) - measure N {x} \ = \ \

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i)) - \

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i)) \" unfolding eq by (simp add: field_simps) also have "\ \ \ \

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i)) \ + \ \

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i)) \" by (rule abs_triangle_ineq4) also have "\ \ \

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i)) + \

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i))" by simp finally have "\ p y x (Suc n) - measure N {x} \ \ \" . } note mono = this { fix n :: nat have "(\\<^sup>+x. \ p y x (Suc n) - measure N {x} \ \count_space C) \ (\\<^sup>+x. ennreal (\

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i))) + ennreal (\

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i))) \count_space C)" using mono by (intro nn_integral_mono) (simp add: ennreal_plus[symmetric] del: ennreal_plus) also have "\ = (\\<^sup>+x. \

(\ in KN.T (y, None). fst (\ !! n) = x \ \ (\i)) \count_space C) + (\\<^sup>+x. \

(\ in KN.T (y, None). snd (\ !! n) = Some x \ \ (\i)) \count_space C)" by (subst nn_integral_add) auto also have "\ = emeasure (KN.T (y, None)) (\x\C. {\\space (KN.T (y, None)). fst (\ !! n) = x \ \ (\i)}) + emeasure (KN.T (y, None)) (\x\C. {\\space (KN.T (y, None)). snd (\ !! n) = Some x \ \ (\i)})" by (subst (1 2) emeasure_UN_countable) (auto simp add: disjoint_family_on_def KN.T.emeasure_eq_measure C) also have "\ \ ennreal (\

(\ in KN.T (y, None). \ (\i))) + ennreal (\

(\ in KN.T (y, None). \ (\i)))" unfolding KN.T.emeasure_eq_measure by (intro add_mono) (auto intro!: KN.T.finite_measure_mono) also have "\ \ 2 * \

(\ in KN.T (y, None). \ (\i))" by (simp add: ennreal_plus[symmetric] del: ennreal_plus) finally have "?L (Suc n) \ 2 * \

(\ in KN.T (y, None). \ (\i))" by (auto intro!: integral_real_bounded simp add: pmf.rep_eq) } note le_2 = this have c0_D: "(c0, Some c0) \ D" by (simp add: D_def c0) let ?N' = "map_pmf Some N" interpret NP: pair_prob_space N ?N' .. have pos_recurrent: "\x\C \ Some ` C. KN.pos_recurrent x" proof (rule KN.stationary_distributionD(1)[OF KN_essential _ KN.stationary_distributionI_pair[OF N(1)]]) show "K2.stationary_distribution ?N'" unfolding K2.stationary_distribution_def by (subst N(1)[unfolded stationary_distribution_def]) (auto intro!: bind_pmf_cong simp: K'_def map_pmf_def bind_assoc_pmf bind_return_pmf) show "countable (C \ Some`C)" using C by auto show "set_pmf (pair_pmf N (map_pmf Some N)) \ C \ Some ` C" using \N \ C\ by auto qed from c0_D have "\

(\ in KN.T (y, None). alw (not (HLD D)) \) \ \

(\ in KN.T (y, None). alw (not (HLD {(c0, Some c0)})) \)" apply (auto intro!: KN.T.finite_measure_mono) apply (rule alw_mono, assumption) apply (auto simp: HLD_iff) done also have "\ = 0" apply (rule KN.T.prob_eq_0_AE) apply (simp add: not_ev_iff[symmetric]) apply (subst KN.AE_T_iff) apply simp proof fix t assume t: "t \ KN.Kp (y, None)" then obtain a b where t_eq: "t = (a, Some b)" "a \ K y" "b \ N" unfolding KN.Kp_def by (auto simp: K'_def) with \y \ C\ have "a \ C" using essential_classD2[OF \essential_class C\ \y \ C\] by auto have "b \ C" using \N \ C\ \b \ N\ by auto from pos_recurrent[THEN bspec, of "(c0, Some c0)"] have recurrent_c0: "KN.recurrent (c0, Some c0)" by (simp add: KN.pos_recurrent_def c0) have "C \ Some ` C \ UNIV // KN.communicating" using aperiodic by (simp add: KN.aperiodic_def) then have "((c0, Some c0), t) \ KN.communicating" by (rule KN.irreducibleD) (simp_all add: t_eq c0 \b \ C\ \a \ C\) then have "((c0, Some c0), t) \ KN.acc" by (simp add: KN.communicating_def) then have "KN.U t (c0, Some c0) = 1" by (rule KN.recurrent_acc(1)[OF recurrent_c0]) then show "AE \ in KN.T t. ev (HLD {(c0, Some c0)}) (t ## \)" unfolding KN.U_def by (subst (asm) KN.T.prob_Collect_eq_1) (auto simp add: ev_Stream) qed finally have "\

(\ in KN.T (y, None). alw (not (HLD D)) \) = 0" by (intro antisym measure_nonneg) have "(\n. \

(\ in KN.T (y, None). \ (\i))) \ measure (KN.T (y, None)) (\n. {\\space (KN.T (y, None)). \ (\i)})" by (rule KN.T.finite_Lim_measure_decseq) (auto simp: decseq_def) also have "(\n. {\\space (KN.T (y, None)). \ (\i)}) = {\\space (KN.T (y, None)). alw (not (HLD D)) \}" by (auto simp: not_ev_iff[symmetric] ev_iff_ev_at) also have "\

(\ in KN.T (y, None). alw (not (HLD D)) \) = 0" by fact finally have *: "(\n. 2 * \

(\ in KN.T (y, None). \ (\i))) \ 0" by (intro tendsto_eq_intros) auto show ?thesis apply (rule LIMSEQ_imp_Suc) apply (rule tendsto_sandwich[OF _ _ tendsto_const *]) using le_2 apply (simp_all add: integral_nonneg_AE) done qed lemma stationary_distribution_imp_p_limit: assumes "aperiodic C" "essential_class C" and [simp]: "countable C" assumes N: "stationary_distribution N" "N \ C" assumes [simp]: "x \ C" "y \ C" shows "p x y \ pmf N y" proof - define D where "D y n = \p x y n - pmf N y\" for y n from stationary_distribution_imp_limit[OF assms(1,2,3,4,5,6)] have INT: "(\n. \y. D y n \count_space C) \ 0" unfolding D_def . { fix n have "D y n \ (\z. D y n * indicator {y} z \count_space C)" by simp also have "\ \ (\y. D y n \count_space C)" by (intro integral_mono) (auto split: split_indicator simp: D_def p_def disjoint_family_on_def intro!: Bochner_Integration.integrable_diff integrable_pmf T.integrable_measure) finally have "D y n \ (\y. D y n \count_space C)" . } note * = this have D_nonneg: "\n. 0 \ D y n" by (simp add: D_def) have "D y \ 0" by (rule tendsto_sandwich[OF _ _ tendsto_const INT]) (auto simp: eventually_sequentially * D_nonneg) then show ?thesis using Lim_null[where l="pmf N y" and net=sequentially and f="p x y"] by (simp add: D_def [abs_def] tendsto_rabs_zero_iff) qed end lemma (in MC_syntax) essential_classI2: assumes "X \ {}" assumes accI: "\x y. x \ X \ y \ X \ (x, y) \ acc" assumes ED: "\x y. x \ X \ y \ set_pmf (K x) \ y \ X" shows "essential_class X" proof (rule essential_classI) { fix x y assume "(x, y) \ acc" "x \ X" then show "y \ X" by induct (auto dest: ED)} note accD = this from \X \ {}\ obtain x where "x \ X" by auto from \x \ X\ show "X \ UNIV // communicating" by (auto simp add: quotient_def Image_def communicating_def accI dest: accD intro!: exI[of _ x]) qed end diff --git a/thys/Minsky_Machines/Minsky.thy b/thys/Minsky_Machines/Minsky.thy --- a/thys/Minsky_Machines/Minsky.thy +++ b/thys/Minsky_Machines/Minsky.thy @@ -1,835 +1,835 @@ section \Minsky machines\ theory Minsky - imports Recursive_Inseparability "Abstract-Rewriting.Abstract_Rewriting" + imports Recursive_Inseparability "Abstract-Rewriting.Abstract_Rewriting" "Pure-ex.Guess" begin text \We formalize Minksy machines, and relate them to recursive functions. In our flavor of Minsky machines, a machine has a set of registers and a set of labels, and a program is a set of labeled operations. There are two operations, @{term Inc} and @{term Dec}; the former takes a register and a label, and the latter takes a register and two labels. When an @{term Inc} instruction is executed, the register is incremented and execution continues at the provided label. The @{term Dec} instruction checks the register. If it is non-zero, the register and continues execution at the first label. Otherwise, the register remains at zero and execution continues at the second label. We continue to show that Minksy machines can implement any primitive recursive function. Based on that, we encode recursively enumerable sets as Minsky machines, and finally show that \begin{enumerate} \item The set of Minsky configurations such that from state $1$, state $0$ can be reached, is undecidable; \item There is a deterministic Minsky machine $U$ such that the set of values $x$ such that $(2, \lambda n.\, \text{if $n = 0$ then $x$ else $0$})$ reach state $0$ is recursively inseparable from those that reach state $1$; and \item As a corollary, the set of Minsky configurations that reach state $0$ but not state $1$ is recursively inseparable from the configurations that reach state $1$ but not state $0$. \end{enumerate}\ subsection \Deterministic relations\ text \A relation $\to$ is \emph{deterministic} if $t \from s \to u'$ implies $t = u$. This abstract rewriting notion is useful for talking about deterministic Minsky machines.\ definition "deterministic R \ R\ O R \ Id" lemma deterministicD: "deterministic R \ (x, y) \ R \ (x, z) \ R \ y = z" by (auto simp: deterministic_def) lemma deterministic_empty [simp]: "deterministic {}" by (auto simp: deterministic_def) lemma deterministic_singleton [simp]: "deterministic {p}" by (auto simp: deterministic_def) lemma deterministic_imp_weak_diamond [intro]: "deterministic R \ w\ R" by (auto simp: weak_diamond_def deterministic_def) lemmas deterministic_imp_CR = deterministic_imp_weak_diamond[THEN weak_diamond_imp_CR] lemma deterministic_union: "fst ` S \ fst ` R = {} \ deterministic S \ deterministic R \ deterministic (S \ R)" by (fastforce simp add: deterministic_def disjoint_iff_not_equal) lemma deterministic_map: "inj_on f (fst ` R) \ deterministic R \ deterministic (map_prod f g ` R)" by (auto simp add: deterministic_def dest!: inj_onD; force) subsection \Minsky machine definition\ text \A Minsky operation either decrements a register (testing for zero, with two possible successor states), or increments a register (with one successor state). A Minsky machine is a set of pairs of states and operations.\ datatype ('s, 'v) Op = Dec (op_var: 'v) 's 's | Inc (op_var: 'v) 's type_synonym ('s, 'v) minsky = "('s \ ('s, 'v) Op) set" text \Semantics: A Minsky machine operates on pairs consisting of a state and an assignment of the registers; in each step, either a register is incremented, or a register is decremented, provided it is non-zero. We write $\alpha$ for assignments; $\alpha[v]$ for the value of the register $v$ in $\alpha$ and $\alpha[v := n]$ for the update of $v$ to $n$. Thus, the semantics is as follows: \begin{enumerate} \item if $(s, Inc\ v\ s') \in M$ then $(s, \alpha) \to (s', \alpha[v := \alpha[v] + 1]$); \item if $(s, Dec\ v\ s_n\ s_z) \in M$ and $\alpha[v] > 0$ then $(s, \alpha) \to (s_n, \alpha[v := \alpha[v] - 1]$); and \item if $(s, Dec\ v\ s_n\ s_z) \in M$ and $\alpha[v] = 0$ then $(s, \alpha) \to (s_z, \alpha)$. \end{enumerate} A state is finite if there is no operation associated with it.\ inductive_set step :: "('s, 'v) minsky \ ('s \ ('v \ nat)) rel" for M :: "('s, 'v) minsky" where inc: "(s, Inc v s') \ M \ ((s, vs), (s', \x. if x = v then Suc (vs v) else vs x)) \ step M" | decn: "(s, Dec v sn sz) \ M \ vs v = Suc n \ ((s, vs), (sn, \x. if x = v then n else vs x)) \ step M" | decz: "(s, Dec v sn sz) \ M \ vs v = 0 \ ((s, vs), (sz, vs)) \ step M" lemma step_mono: "M \ M' \ step M \ step M'" by (auto elim: step.cases intro: step.intros) lemmas steps_mono = rtrancl_mono[OF step_mono] text \A Minsky machine has deterministic steps if its defining relation between states and operations is deterministic.\ lemma deterministic_stepI [intro]: assumes "deterministic M" shows "deterministic (step M)" proof - { fix s vs s1 vs1 s2 vs2 assume s: "((s, vs), (s1, vs1)) \ step M" "((s, vs), (s2, vs2)) \ step M" have "(s1, vs1) = (s2, vs2)" using deterministicD[OF assms] by (cases rule: step.cases[OF s(1)]; cases rule: step.cases[OF s(2)]) fastforce+ } then show ?thesis by (auto simp: deterministic_def) qed text \A Minksy machine halts when it reaches a state with no associated operation.\ lemma NF_stepI [intro]: "s \ fst ` M \ (s, vs) \ NF (step M)" by (auto intro!: no_step elim!: step.cases simp: rev_image_eqI) text \Deterministic Minsky machines enjoy unique normal forms.\ lemmas deterministic_minsky_UN = join_NF_imp_eq[OF CR_divergence_imp_join[OF deterministic_imp_CR[OF deterministic_stepI]] NF_stepI NF_stepI] text \We will rename states and variables.\ definition map_minsky where "map_minsky f g M = map_prod f (map_Op f g) ` M" lemma map_minsky_id: "map_minsky id id M = M" by (simp add: map_minsky_def Op.map_id0 map_prod.id) lemma map_minsky_comp: "map_minsky f g (map_minsky f' g' M) = map_minsky (f \ f') (g \ g') M" unfolding map_minsky_def image_comp Op.map_comp map_prod.comp comp_def[of "map_Op _ _"] .. text \When states and variables are renamed, computations carry over from the original machine, provided that variables are renamed injectively.\ lemma map_step: assumes "inj g" "vs = vs' \ g" "((s, vs), (t, ws)) \ step M" shows "((f s, vs'), (f t, \x. if x \ range g then ws (inv g x) else vs' x)) \ step (map_minsky f g M)" using assms(3) proof (cases rule: step.cases) case (inc v) note [simp] = inc(1) let ?ws' = "\w. if w = g v then Suc (vs' (g v)) else vs' w" have "((f s, vs'), (f t, ?ws')) \ step (map_minsky f g M)" using inc(2) step.inc[of "f s" "g v" "f t" "map_minsky f g M" vs'] by (force simp: map_minsky_def) moreover have "(\x. if x \ range g then ws (inv g x) else vs' x) = ?ws'" using assms(1,2) by (auto intro!: ext simp: injD image_def) ultimately show ?thesis by auto next case (decn v sz n) note [simp] = decn(1) let ?ws' = "\x. if x = g v then n else vs' x" have "((f s, vs'), (f t, ?ws')) \ step (map_minsky f g M)" using assms(2) decn(2-) step.decn[of "f s" "g v" "f t" "f sz" "map_minsky f g M" vs' n] by (force simp: map_minsky_def) moreover have "(\x. if x \ range g then ws (inv g x) else vs' x) = ?ws'" using assms(1,2) by (auto intro!: ext simp: injD image_def) ultimately show ?thesis by auto next case (decz v sn) note [simp] = decz(1) have "((f s, vs'), (f t, vs')) \ step (map_minsky f g M)" using assms(2) decz(2-) step.decz[of "f s" "g v" "f sn" "f t" "map_minsky f g M" vs'] by (force simp: map_minsky_def) moreover have "(\x. if x \ range g then ws (inv g x) else vs' x) = vs'" using assms(1,2) by (auto intro!: ext simp: injD image_def) ultimately show ?thesis by auto qed lemma map_steps: assumes "inj g" "vs = ws \ g" "((s, vs), (t, vs')) \ (step M)\<^sup>*" shows "((f s, ws), (f t, \x. if x \ range g then vs' (inv g x) else ws x)) \ (step (map_minsky f g M))\<^sup>*" using assms(3,2) proof (induct "(s, vs)" arbitrary: s vs ws rule: converse_rtrancl_induct) case base then have "(\x. if x \ range g then vs' (inv g x) else ws x) = ws" using assms(1) by (auto intro!: ext simp: injD image_def) then show ?case by auto next case (step y) have "snd y = (\x. if x \ range g then snd y (inv g x) else ws x) \ g" (is "_ = ?ys' \ _") using assms(1) by auto then show ?case using map_step[OF assms(1) step(4), of s "fst y" "snd y" M f] step(1) step(3)[OF prod.collapse[symmetric], of ?ys'] by (auto cong: if_cong) qed subsection \Concrete Minsky machines\ text \The following definition expresses when a Minsky machine $M$ implements a specification $P$. We adopt the convention that computations always start out in state $1$ and end in state $0$, which must be a final state. The specification $P$ relates initial assignments to final assignments.\ definition mk_minsky_wit :: "(nat, nat) minsky \ ((nat \ nat) \ (nat \ nat) \ bool) \ bool" where "mk_minsky_wit M P \ finite M \ deterministic M \ 0 \ fst ` M \ (\vs. \vs'. ((Suc 0, vs), (0, vs')) \ (step M)\<^sup>* \ P vs vs')" abbreviation mk_minsky :: "((nat \ nat) \ (nat \ nat) \ bool) \ bool" where "mk_minsky P \ \M. mk_minsky_wit M P" lemmas mk_minsky_def = mk_minsky_wit_def lemma mk_minsky_mono: shows "mk_minsky P \ (\vs vs'. P vs vs' \ Q vs vs') \ mk_minsky Q" unfolding mk_minsky_def by meson lemma mk_minsky_sound: assumes "mk_minsky_wit M P" "((Suc 0, vs), (0, vs')) \ (step M)\<^sup>*" shows "P vs vs'" proof - have M: "deterministic M" "0 \ fst ` M" "\vs. \vs'. ((Suc 0, vs), 0, vs') \ (step M)\<^sup>* \ P vs vs'" using assms(1) by (auto simp: mk_minsky_wit_def) obtain vs'' where vs'': "((Suc 0, vs), (0, vs'')) \ (step M)\<^sup>*" "P vs vs''" using M(3) by blast have "(0 :: nat, vs') = (0, vs'')" using M(1,2) by (intro deterministic_minsky_UN[OF _ assms(2) vs''(1)]) then show ?thesis using vs''(2) by simp qed text \Realizability of $n$-ary functions for $n = 1 \dots 3$. Here we use the convention that the arguments are passed in registers $1 \dots 3$, and the result is stored in register $0$.\ abbreviation mk_minsky1 where "mk_minsky1 f \ mk_minsky (\vs vs'. vs' 0 = f (vs 1))" abbreviation mk_minsky2 where "mk_minsky2 f \ mk_minsky (\vs vs'. vs' 0 = f (vs 1) (vs 2))" abbreviation mk_minsky3 where "mk_minsky3 f \ mk_minsky (\vs vs'. vs' 0 = f (vs 1) (vs 2) (vs 3))" subsection \Trivial building blocks\ text \We can increment and decrement any register.\ lemma mk_minsky_inc: shows "mk_minsky (\vs vs'. vs' = (\x. if x = v then Suc (vs v) else vs x))" using step.inc[of "Suc 0" v 0] by (auto simp: deterministic_def mk_minsky_def intro!: exI[of _ "{(1, Inc v 0)} :: (nat, nat) minsky"]) lemma mk_minsky_dec: shows "mk_minsky (\vs vs'. vs' = (\x. if x = v then vs v - 1 else vs x))" proof - let ?M = "{(1, Dec v 0 0)} :: (nat, nat) minsky" show ?thesis unfolding mk_minsky_def proof (intro exI[of _ ?M] allI conjI, goal_cases) case (4 vs) have [simp]: "vs v = 0 \ (\x. if x = v then 0 else vs x) = vs" by auto show ?case using step.decz[of "Suc 0" v 0 0 ?M] step.decn[of "Suc 0" v 0 0 ?M] by (cases "vs v") (auto cong: if_cong) qed auto qed subsection \Sequential composition\ text \The following lemma has two useful corollaries (which we prove simultaneously because they share much of the proof structure): First, if $P$ and $Q$ are realizable, then so is $P \circ Q$. Secondly, if we rename variables by an injective function $f$ in a Minksy machine, then the variables outside the range of $f$ remain unchanged.\ lemma mk_minsky_seq_map: assumes "mk_minsky P" "mk_minsky Q" "inj g" "\vs vs' vs''. P vs vs' \ Q vs' vs'' \ R vs vs''" shows "mk_minsky (\vs vs'. R (vs \ g) (vs' \ g) \ (\x. x \ range g \ vs x = vs' x))" proof - obtain M where M: "finite M" "deterministic M" "0 \ fst ` M" "\vs. \vs'. ((Suc 0, vs), 0, vs') \ (step M)\<^sup>* \ P vs vs'" using assms(1) by (auto simp: mk_minsky_def) obtain N where N: "finite N" "deterministic N" "0 \ fst ` N" "\vs. \vs'. ((Suc 0, vs), 0, vs') \ (step N)\<^sup>* \ Q vs vs'" using assms(2) by (auto simp: mk_minsky_def) let ?fM = "\s. if s = 0 then 2 else if s = 1 then 1 else 2 * s + 1" \ \M: from 1 to 2\ let ?fN = "\s. 2 * s" \ \N: from 2 to 0\ let ?M = "map_minsky ?fM g M \ map_minsky ?fN g N" show ?thesis unfolding mk_minsky_def proof (intro exI[of _ ?M] conjI allI, goal_cases) case 1 show ?case using M(1) N(1) by (auto simp: map_minsky_def) next case 2 show ?case using M(2,3) N(2) unfolding map_minsky_def by (intro deterministic_union deterministic_map) (auto simp: inj_on_def rev_image_eqI Suc_double_not_eq_double split: if_splits) next case 3 show ?case using N(3) by (auto simp: rev_image_eqI map_minsky_def split: if_splits) next case (4 vs) obtain vsM where M': "((Suc 0, vs \ g), 0, vsM) \ (step M)\<^sup>*" "P (vs \ g) vsM" using M(4) by blast obtain vsN where N': "((Suc 0, vsM), 0, vsN) \ (step N)\<^sup>*" "Q vsM vsN" using N(4) by blast note * = subsetD[OF steps_mono, of _ ?M] map_steps[OF _ _ M'(1), of g vs ?fM, simplified] map_steps[OF _ _ N'(1), of g _ ?fN, simplified] show ?case using assms(3,4) M'(2) N'(2) rtrancl_trans[OF *(1)[OF _ *(2)] *(1)[OF _ *(3)]] by (auto simp: comp_def) qed qed text \Sequential composition.\ lemma mk_minsky_seq: assumes "mk_minsky P" "mk_minsky Q" "\vs vs' vs''. P vs vs' \ Q vs' vs'' \ R vs vs''" shows "mk_minsky R" using mk_minsky_seq_map[OF assms(1,2), of id] assms(3) by simp lemma mk_minsky_seq': assumes "mk_minsky P" "mk_minsky Q" shows "mk_minsky (\vs vs''. (\vs'. P vs vs' \ Q vs' vs''))" by (intro mk_minsky_seq[OF assms]) blast text \We can do nothing (besides transitioning from state 1 to state 0).\ lemma mk_minsky_nop: "mk_minsky (\vs vs'. vs = vs')" by (intro mk_minsky_seq[OF mk_minsky_inc mk_minsky_dec]) auto text \Renaming variables.\ lemma mk_minsky_map: assumes "mk_minsky P" "inj f" shows "mk_minsky (\vs vs'. P (vs \ f) (vs' \ f) \ (\x. x \ range f \ vs x = vs' x))" using mk_minsky_seq_map[OF assms(1) mk_minsky_nop assms(2)] by simp lemma inj_shift [simp]: fixes a b :: nat assumes "a < b" shows "inj (\x. if x = 0 then a else x + b)" using assms by (auto simp: inj_on_def) subsection \Bounded loop\ text \In the following lemma, $P$ is the specification of a loop body, and $Q$ the specification of the loop itself (a loop invariant). The loop variable is $v$. $Q$ can be realized provided that \begin{enumerate} \item $P$ can be realized; \item $P$ ensures that the loop variable is not changed by the loop body; and \item $Q$ follows by induction on the loop variable: \begin{enumerate} \item $\alpha\,Q\,\alpha$ holds when $\alpha[v] = 0$; and \item $\alpha[v := n]\,P\,\alpha'$ and $\alpha'\,Q\,\alpha''$ imply $\alpha\,Q\,alpha''$ when $\alpha[v] = n+1$. \end{enumerate} \end{enumerate}\ lemma mk_minsky_loop: assumes "mk_minsky P" "\vs vs'. P vs vs' \ vs' v = vs v" "\vs. vs v = 0 \ Q vs vs" "\n vs vs' vs''. vs v = Suc n \ P (\x. if x = v then n else vs x) vs' \ Q vs' vs'' \ Q vs vs''" shows "mk_minsky Q" proof - obtain M where M: "finite M" "deterministic M" "0 \ fst ` M" "\vs. \vs'. ((Suc 0, vs), 0, vs') \ (step M)\<^sup>* \ P vs vs'" using assms(1) by (auto simp: mk_minsky_def) let ?M = "{(1, Dec v 2 0)} \ map_minsky Suc id M" show ?thesis unfolding mk_minsky_def proof (intro exI[of _ ?M] conjI allI, goal_cases) case 1 show ?case using M(1) by (auto simp: map_minsky_def) next case 2 show ?case using M(2,3) unfolding map_minsky_def by (intro deterministic_union deterministic_map) (auto simp: rev_image_eqI) next case 3 show ?case by (auto simp: map_minsky_def) next case (4 vs) show ?case proof (induct "vs v" arbitrary: vs) case 0 then show ?case using assms(3)[of vs] step.decz[of 1 v 2 0 ?M vs] by (auto simp: id_def) next case (Suc n) obtain vs' where M': "((Suc 0, \x. if x = v then n else vs x), 0, vs') \ (step M)\<^sup>*" "P (\x. if x = v then n else vs x) vs'" using M(4) by blast obtain vs'' where D: "((Suc 0, vs'), 0, vs'') \ (step ?M)\<^sup>*" "Q vs' vs''" using Suc(1)[of vs'] assms(2)[OF M'(2)] by auto note * = subsetD[OF steps_mono, of _ ?M] r_into_rtrancl[OF decn[of "Suc 0" v 2 0 ?M vs n]] map_steps[OF _ _ M'(1), of id _ Suc, simplified, OF refl, simplified, folded numeral_2_eq_2] show ?case using rtrancl_trans[OF rtrancl_trans, OF *(2) *(1)[OF _ *(3)] D(1)] D(2) Suc(2) assms(4)[OF _ M'(2), of vs''] by auto qed qed qed subsection \Copying values\ text \We work up to copying values in several steps. \begin{enumerate} \item Clear a register. This is a loop that decrements the register until it reaches $0$. \item Add a register to another one. This is a loop that decrements one register, and increments the other register, until the first register reaches 0. \item Add a register to two others. This is the same, except that two registers are incremented. \item Move a register: set a register to 0, then add another register to it. \item Copy a register destructively: clear two registers, then add another register to them. \end{enumerate}\ lemma mk_minsky_zero: shows "mk_minsky (\vs vs'. vs' = (\x. if x = v then 0 else vs x))" by (intro mk_minsky_loop[where v = v, OF \ \ while v[v]$--$: \ mk_minsky_nop]) auto \ \ pass \ lemma mk_minsky_add1: assumes "v \ w" shows "mk_minsky (\vs vs'. vs' = (\x. if x = v then 0 else if x = w then vs v + vs w else vs x))" using assms by (intro mk_minsky_loop[where v = v, OF \ \ while v[v]$--$: \ mk_minsky_inc[of w]]) auto \ \ v[w]++ \ lemma mk_minsky_add2: assumes "u \ v" "u \ w" "v \ w" shows "mk_minsky (\vs vs'. vs' = (\x. if x = u then 0 else if x = v then vs u + vs v else if x = w then vs u + vs w else vs x))" using assms by (intro mk_minsky_loop[where v = u, OF mk_minsky_seq'[OF \ \ while v[u]$--$: \ mk_minsky_inc[of v] \ \ v[v]++ \ mk_minsky_inc[of w]]]) auto \ \ v[w]++ \ lemma mk_minsky_copy1: assumes "v \ w" shows "mk_minsky (\vs vs'. vs' = (\x. if x = v then 0 else if x = w then vs v else vs x))" using assms by (intro mk_minsky_seq[OF mk_minsky_zero[of w] \ \ v[w] := 0 \ mk_minsky_add1[of v w]]) auto \ \ v[w] := v[w] + v[v], v[v] := 0 \ lemma mk_minsky_copy2: assumes "u \ v" "u \ w" "v \ w" shows "mk_minsky (\vs vs'. vs' = (\x. if x = u then 0 else if x = v then vs u else if x = w then vs u else vs x))" using assms by (intro mk_minsky_seq[OF mk_minsky_seq', OF mk_minsky_zero[of v] \ \ v[v] := 0 \ mk_minsky_zero[of w] \ \ v[w] := 0 \ mk_minsky_add2[of u v w]]) auto \ \ v[v] := v[v] + v[u], v[w] := v[w] + v[u], v[u] := 0 \ lemma mk_minsky_copy: assumes "u \ v" "u \ w" "v \ w" shows "mk_minsky (\vs vs'. vs' = (\x. if x = v then vs u else if x = w then 0 else vs x))" using assms by (intro mk_minsky_seq[OF mk_minsky_copy2[of u v w] \ \ v[v] := v[u], v[w] := v[u], v[u] := 0 \ mk_minsky_copy1[of w u]]) auto \ \ v[u] := v[w], v[w] := 0 \ subsection \Primitive recursive functions\ text \Nondestructive apply: compute $f$ on arguments $\alpha[u]$, $\alpha[v]$, $\alpha[w]$, storing the result in $\alpha[t]$ and preserving all other registers below $k$. This is easy now that we can copy values.\ lemma mk_minsky_apply3: assumes "mk_minsky3 f" "t < k" "u < k" "v < k" "w < k" shows "mk_minsky (\vs vs'. \x < k. vs' x = (if x = t then f (vs u) (vs v) (vs w) else vs x))" using assms(2-) by (intro mk_minsky_seq[OF mk_minsky_seq'[OF mk_minsky_seq'], OF mk_minsky_copy[of u "1 + k" k] \ \ v[1+k] := v[u] \ mk_minsky_copy[of v "2 + k" k] \ \ v[2+k] := v[v] \ mk_minsky_copy[of w "3 + k" k] \ \ v[3+k] := v[w] \ mk_minsky_map[OF assms(1), of "\x. if x = 0 then t else x + k"]]) (auto 0 2) \ \ v[t] := f v[1+k] v[2+k] v[3+k] \ text \Composition is just four non-destructive applies.\ lemma mk_minsky_comp3_3: assumes "mk_minsky3 f" "mk_minsky3 g" "mk_minsky3 h" "mk_minsky3 k" shows "mk_minsky3 (\x y z. f (g x y z) (h x y z) (k x y z))" by (rule mk_minsky_seq[OF mk_minsky_seq'[OF mk_minsky_seq'], OF mk_minsky_apply3[OF assms(2), of 4 7 1 2 3] \ \ v[4] := g v[1] v[2] v[3] \ mk_minsky_apply3[OF assms(3), of 5 7 1 2 3] \ \ v[5] := h v[1] v[2] v[3] \ mk_minsky_apply3[OF assms(4), of 6 7 1 2 3] \ \ v[6] := k v[1] v[2] v[3] \ mk_minsky_apply3[OF assms(1), of 0 7 4 5 6]]) auto \ \ v[0] := f v[4] v[5] v[6] \ text \Primitive recursion is a non-destructive apply followed by a loop with another non-destructive apply. The key to the proof is the loop invariant, which we can specify as part of composing the various \mk_minsky_*\ lemmas.\ lemma mk_minsky_prim_rec: assumes "mk_minsky1 g" "mk_minsky3 h" shows "mk_minsky2 (PrimRecOp g h)" by (intro mk_minsky_seq[OF mk_minsky_seq', OF mk_minsky_apply3[OF assms(1), of 0 4 2 2 2] \ \ v[0] := g v[2] \ mk_minsky_zero[of 3] \ \ v[3] := 0 \ mk_minsky_loop[where v = 1, OF mk_minsky_seq', OF \ \ while v[1]$--$: \ mk_minsky_apply3[OF assms(2), of 0 4 3 0 2] \ \ v[0] := h v[3] v[0] v[2] \ mk_minsky_inc[of 3], \ \ v[3]++ \ of "\vs vs'. vs 0 = PrimRecOp g h (vs 3) (vs 2) \ vs' 0 = PrimRecOp g h (vs 3 + vs 1) (vs 2)" ]]) auto text \With these building blocks we can easily show that all primitive recursive functions can be realized by a Minsky machine.\ lemma mk_minsky_PrimRec: "f \ PrimRec1 \ mk_minsky1 f" "g \ PrimRec2 \ mk_minsky2 g" "h \ PrimRec3 \ mk_minsky3 h" proof (goal_cases) have *: "(f \ PrimRec1 \ mk_minsky1 f) \ (g \ PrimRec2 \ mk_minsky2 g) \ (h \ PrimRec3 \ mk_minsky3 h)" proof (induction rule: PrimRec1_PrimRec2_PrimRec3.induct) case zero show ?case by (intro mk_minsky_mono[OF mk_minsky_zero]) auto next case suc show ?case by (intro mk_minsky_seq[OF mk_minsky_copy1[of 1 0] mk_minsky_inc[of 0]]) auto next case id1_1 show ?case by (intro mk_minsky_mono[OF mk_minsky_copy1[of 1 0]]) auto next case id2_1 show ?case by (intro mk_minsky_mono[OF mk_minsky_copy1[of 1 0]]) auto next case id2_2 show ?case by (intro mk_minsky_mono[OF mk_minsky_copy1[of 2 0]]) auto next case id3_1 show ?case by (intro mk_minsky_mono[OF mk_minsky_copy1[of 1 0]]) auto next case id3_2 show ?case by (intro mk_minsky_mono[OF mk_minsky_copy1[of 2 0]]) auto next case id3_3 show ?case by (intro mk_minsky_mono[OF mk_minsky_copy1[of 3 0]]) auto next case (comp1_1 f g) then show ?case using mk_minsky_comp3_3 by fast next case (comp1_2 f g) then show ?case using mk_minsky_comp3_3 by fast next case (comp1_3 f g) then show ?case using mk_minsky_comp3_3 by fast next case (comp2_1 f g h) then show ?case using mk_minsky_comp3_3 by fast next case (comp3_1 f g h k) then show ?case using mk_minsky_comp3_3 by fast next case (comp2_2 f g h) then show ?case using mk_minsky_comp3_3 by fast next case (comp2_3 f g h) then show ?case using mk_minsky_comp3_3 by fast next case (comp3_2 f g h k) then show ?case using mk_minsky_comp3_3 by fast next case (comp3_3 f g h k) then show ?case using mk_minsky_comp3_3 by fast next case (prim_rec g h) then show ?case using mk_minsky_prim_rec by blast qed { case 1 thus ?case using * by blast next case 2 thus ?case using * by blast next case 3 thus ?case using * by blast } qed subsection \Recursively enumerable sets as Minsky machines\ text \The following is the most complicated lemma of this theory: Given two r.e.\ sets $A$ and $B$ we want to construct a Minsky machine that reaches the final state $0$ for input $x$ if $x \in A$ and final state $1$ if $x \in B$, and never reaches either of these states if $x \notin A \cup B$. (If $x \in A \cap B$, then either state $0$ or state $1$ may be reached.) We consider two r.e.\ sets rather than one because we target recursive inseparability. For the r.e.\ set $A$, there is a primitive recursive function $f$ such that $x \in A \iff \exists y.\, f(x,y) = 0$. Similarly there is a primitive recursive function $g$ for $B$ such that $x \in B \iff \exists y.\, f(x,y) = 0$. Our Minsky machine takes $x$ in register $0$ and $y$ in register $1$ (initially $0$) and works as follows. \begin{enumerate} \item evaluate $f(x,y)$; if the result is $0$, transition to state $0$; otherwise, \item evaluate $g(x,y)$; if the result is $0$, transition to state $1$; otherwise, \item increment $y$ and start over. \end{enumerate}\ lemma ce_set_pair_by_minsky: assumes "A \ ce_sets" "B \ ce_sets" obtains M :: "(nat, nat) minsky" where "finite M" "deterministic M" "0 \ fst ` M" "Suc 0 \ fst ` M" "\x vs. vs 0 = x \ vs 1 = 0 \ x \ A \ B \ \vs'. ((2, vs), (0, vs')) \ (step M)\<^sup>* \ ((2, vs), (Suc 0, vs')) \ (step M)\<^sup>*" "\x vs vs'. vs 0 = x \ vs 1 = 0 \ ((2, vs), (0, vs')) \ (step M)\<^sup>* \ x \ A" "\x vs vs'. vs 0 = x \ vs 1 = 0 \ ((2, vs), (Suc 0, vs')) \ (step M)\<^sup>* \ x \ B" proof - obtain g where g: "g \ PrimRec2" "\x. x \ A \ (\y. g x y = 0)" using assms(1) by (auto simp: ce_sets_def fn_to_set_def) obtain h where h: "h \ PrimRec2" "\x. x \ B \ (\y. h x y = 0)" using assms(2) by (auto simp: ce_sets_def fn_to_set_def) have "mk_minsky (\vs vs'. vs' 0 = vs 0 \ vs' 1 = vs 1 \ vs' 2 = g (vs 0) (vs 1))" using mk_minsky_seq[OF mk_minsky_apply3[OF mk_minsky_PrimRec(2)[OF g(1)], of 2 3 0 1 0] \ \ v[2] := g v[0] v[1] \ mk_minsky_nop] by auto \ \ pass \ then obtain M :: "(nat, nat) minsky" where M: "finite M" "deterministic M" "0 \ fst ` M" "\vs. \vs'. ((Suc 0, vs), 0, vs') \ (step M)\<^sup>* \ vs' 0 = vs 0 \ vs' 1 = vs 1 \ vs' 2 = g (vs 0) (vs 1)" unfolding mk_minsky_def by blast have "mk_minsky (\vs vs'. vs' 0 = vs 0 \ vs' 1 = vs 1 + 1 \ vs' 2 = h (vs 0) (vs 1))" using mk_minsky_seq[OF mk_minsky_apply3[OF mk_minsky_PrimRec(2)[OF h(1)], of 2 3 0 1 0] \ \ v[2] := h v[0] v[1] \ mk_minsky_inc[of 1]] by auto \ \ v[1] := v[1] + 1 \ then obtain N :: "(nat, nat) minsky" where N: "finite N" "deterministic N" "0 \ fst ` N" "\vs. \vs'. ((Suc 0, vs), 0, vs') \ (step N)\<^sup>* \ vs' 0 = vs 0 \ vs' 1 = vs 1 + 1 \ vs' 2 = h (vs 0) (vs 1)" unfolding mk_minsky_def by blast let ?f = "\s. if s = 0 then 3 else 2 * s" \ \M: from state 4 to state 3\ let ?g = "\s. 2 * s + 5" \ \N: from state 7 to state 5\ define X where "X = map_minsky ?f id M \ map_minsky ?g id N \ {(3, Dec 2 7 0)} \ {(5, Dec 2 2 1)}" have MX: "map_minsky ?f id M \ X" by (auto simp: X_def) have NX: "map_minsky ?g id N \ X" by (auto simp: X_def) have DX: "(3, Dec 2 7 0) \ X" "(5, Dec 2 2 1) \ X" by (auto simp: X_def) have X1: "finite X" using M(1) N(1) by (auto simp: map_minsky_def X_def) have X2: "deterministic X" unfolding X_def using M(2,3) N(2,3) apply (intro deterministic_union) by (auto simp: map_minsky_def rev_image_eqI inj_on_def split: if_splits intro!: deterministic_map) presburger+ have X3: "0 \ fst ` X" "Suc 0 \ fst ` X" using M(3) N(3) by (auto simp: X_def map_minsky_def split: if_splits) have X4: "\vs'. g (vs 0) (vs 1) = 0 \ ((2, vs), (0, vs')) \ (step X)\<^sup>* \ h (vs 0) (vs 1) = 0 \ ((2, vs), (1, vs')) \ (step X)\<^sup>* \ g (vs 0) (vs 1) \ 0 \ h (vs 0) (vs 1) \ 0 \ vs' 0 = vs 0 \ vs' 1 = vs 1 + 1 \ ((2, vs), (2, vs')) \ (step X)\<^sup>+" for vs proof - guess vs' using M(4)[of vs] by (elim exE conjE) note vs' = this have 1: "((2, vs), (3, vs')) \ (step X)\<^sup>*" using subsetD[OF steps_mono[OF MX], OF map_steps[OF _ _ vs'(1), of id vs ?f]] by simp show ?thesis proof (cases "vs' 2") case 0 then show ?thesis using decz[OF DX(1), of vs'] vs' 1 by (auto intro: rtrancl_into_rtrancl) next case (Suc n) note Suc' = Suc let ?vs = "\x. if x = 2 then n else vs' x" have 2: "((2, vs), (7, ?vs)) \ (step X)\<^sup>*" using 1 decn[OF DX(1), of vs'] Suc by (auto intro: rtrancl_into_rtrancl) guess vs'' using N(4)[of "?vs"] by (elim exE conjE) note vs'' = this have 3: "((2, vs), (5, vs'')) \ (step X)\<^sup>*" using 2 subsetD[OF steps_mono[OF NX], OF map_steps[OF _ _ vs''(1), of id ?vs ?g]] by simp show ?thesis proof (cases "vs'' 2") case 0 then show ?thesis using 3 decz[OF DX(2), of vs''] vs''(2-) vs'(2-) by (auto intro: rtrancl_into_rtrancl) next case (Suc m) let ?vs = "\x. if x = 2 then m else vs'' x" have 4: "((2, vs), (2, ?vs)) \ (step X)\<^sup>+" using 3 decn[OF DX(2), of vs'' m] Suc by auto then show ?thesis using vs''(2-) vs'(2-) Suc Suc' by (auto intro!: exI[of _ ?vs]) qed qed qed have *: "vs 1 \ y \ g (vs 0) y = 0 \ h (vs 0) y = 0 \ \vs'. ((2, vs), (0, vs')) \ (step X)\<^sup>* \ ((2, vs), (1, vs')) \ (step X)\<^sup>*" for vs y proof (induct "vs 1" arbitrary: vs rule: inc_induct, goal_cases base step) case (base vs) then show ?case using X4[of vs] by auto next case (step vs) guess vs' using X4[of vs] by (elim exE) then show ?case unfolding ex_disj_distrib using step(4) step(3)[of vs'] by (auto dest!: trancl_into_rtrancl) (meson rtrancl_trans)+ qed have **: "((s, vs), (t, ws)) \ (step X)\<^sup>* \ t \ {0, 1} \ ((s, vs), (2, ws')) \ (step X)\<^sup>* \ \y. if t = 0 then g (ws' 0) y = 0 else h (ws' 0) y = 0" for s t vs ws' ws proof (induct arbitrary: ws' rule: converse_rtrancl_induct2) case refl show ?case using refl(1) NF_not_suc[OF refl(2) NF_stepI] X3 by auto next case (step s vs s' vs') show ?case using step(5) proof (cases rule: converse_rtranclE[case_names base' step']) case base' note *** = deterministic_minsky_UN[OF X2 _ _ X3] show ?thesis using X4[of ws'] proof (elim exE disjE conjE, goal_cases) case (1 vs'') then show ?case using step(1,2,4) ***[of "(2,ws')" vs'' ws] by (auto simp: base' intro: converse_rtrancl_into_rtrancl) next case (2 vs'') then show ?case using step(1,2,4) ***[of "(2,ws')" ws vs''] by (auto simp: base' intro: converse_rtrancl_into_rtrancl) next case (3 vs'') then show ?case using step(2) step(3)[of vs'', OF step(4)] deterministicD[OF deterministic_stepI[OF X2], OF _ step(1)] by (auto simp: base' if_bool_eq_conj trancl_unfold_left) qed next case (step' y) then show ?thesis by (metis deterministicD[OF deterministic_stepI[OF X2]] step(1) step(3)[OF step(4)]) qed qed show ?thesis proof (intro that[of X] X1 X2 X3, goal_cases) case (1 x vs) then show ?case using *[of vs] by (auto simp: g(2) h(2)) next case (2 x vs vs') then show ?case using **[of 2 vs 0 vs' vs] by (auto simp: g(2) h(2)) next case (3 x vs vs') then show ?case using **[of 2 vs 1 vs' vs] by (auto simp: g(2) h(2)) qed qed text \For r.e.\ sets we obtain the following lemma as a special case (taking $B = \varnothing$, and swapping states $1$ and $2$).\ lemma ce_set_by_minsky: assumes "A \ ce_sets" obtains M :: "(nat, nat) minsky" where "finite M" "deterministic M" "0 \ fst ` M" "\x vs. vs 0 = x \ vs 1 = 0 \ x \ A \ \vs'. ((Suc 0, vs), (0, vs')) \ (step M)\<^sup>*" "\x vs vs'. vs 0 = x \ vs 1 = 0 \ ((Suc 0, vs), (0, vs')) \ (step M)\<^sup>* \ x \ A" proof - guess M using ce_set_pair_by_minsky[OF assms(1) ce_empty] . note M = this let ?f = "\s. if s = 1 then 2 else if s = 2 then 1 else s" \ \swap states 1 and 2\ have "?f \ ?f = id" by auto define N where "N = map_minsky ?f id M" have M_def: "M = map_minsky ?f id N" unfolding N_def map_minsky_comp \?f \ ?f = id\ map_minsky_id o_id .. show ?thesis using M(1-3) proof (intro that[of N], goal_cases) case (4 x vs) show ?case using M(5)[OF 4(4,5)] 4(6) M(7)[OF 4(4,5)] map_steps[of id vs vs 2 0 _ M ?f] by (auto simp: N_def) next case (5 x vs vs') show ?case using M(6)[OF 5(4,5)] 5(6) map_steps[of id vs vs 1 0 _ N ?f] by (auto simp: M_def) qed (auto simp: N_def map_minsky_def inj_on_def rev_image_eqI deterministic_map split: if_splits) qed subsection \Encoding of Minsky machines\ text \So far, Minsky machines have been sets of pairs of states and operations. We now provide an encoding of Minsky machines as natural numbers, so that we can talk about them as r.e.\ or computable sets. First we encode operations.\ primrec encode_Op :: "(nat, nat) Op \ nat" where "encode_Op (Dec v s s') = c_pair 0 (c_pair v (c_pair s s'))" | "encode_Op (Inc v s) = c_pair 1 (c_pair v s)" definition decode_Op :: "nat \ (nat, nat) Op" where "decode_Op n = (if c_fst n = 0 then Dec (c_fst (c_snd n)) (c_fst (c_snd (c_snd n))) (c_snd (c_snd (c_snd n))) else Inc (c_fst (c_snd n)) (c_snd (c_snd n)))" lemma encode_Op_inv [simp]: "decode_Op (encode_Op x) = x" by (cases x) (auto simp: decode_Op_def) text \Minsky machines are encoded via lists of pairs of states and operations.\ definition encode_minsky :: "(nat \ (nat, nat) Op) list \ nat" where "encode_minsky M = list_to_nat (map (\x. c_pair (fst x) (encode_Op (snd x))) M)" definition decode_minsky :: "nat \ (nat \ (nat, nat) Op) list" where "decode_minsky n = map (\n. (c_fst n, decode_Op (c_snd n))) (nat_to_list n)" lemma encode_minsky_inv [simp]: "decode_minsky (encode_minsky M) = M" by (auto simp: encode_minsky_def decode_minsky_def comp_def) text \Assignments are stored as lists (starting with register 0).\ definition decode_regs :: "nat \ (nat \ nat)" where "decode_regs n = (\i. let xs = nat_to_list n in if i < length xs then nat_to_list n ! i else 0)" text \The undecidability results talk about Minsky configurations (pairs of Minsky machines and assignments). This means that we do not have to construct any recursive functions that modify Minsky machines (for example in order to initialize variables), keeping the proofs simple.\ definition decode_minsky_state :: "nat \ ((nat, nat) minsky \ (nat \ nat))" where "decode_minsky_state n = (set (decode_minsky (c_fst n)), (decode_regs (c_snd n)))" subsection \Undecidablity results\ text \We conclude with some undecidability results. First we show that it is undecidable whether a Minksy machine starting at state 1 terminates in state 0.\ definition minsky_reaching_0 where "minsky_reaching_0 = {n |n M vs vs'. (M, vs) = decode_minsky_state n \ ((Suc 0, vs), (0, vs')) \ (step M)\<^sup>*}" lemma minsky_reaching_0_not_computable: "\ computable minsky_reaching_0" proof - guess U using ce_set_by_minsky[OF univ_is_ce] . note U = this obtain us where [simp]: "set us = U" using finite_list[OF U(1)] by blast let ?f = "\n. c_pair (encode_minsky us) (c_cons n 0)" have "?f \ PrimRec1" using comp2_1[OF c_pair_is_pr const_is_pr comp2_1[OF c_cons_is_pr id1_1 const_is_pr]] by simp moreover have "?f x \ minsky_reaching_0 \ x \ univ_ce" for x using U(4,5)[of "\i. if i = 0 then x else 0"] by (auto simp: minsky_reaching_0_def decode_minsky_state_def decode_regs_def c_cons_def cong: if_cong) ultimately have "many_reducible_to univ_ce minsky_reaching_0" by (auto simp: many_reducible_to_def many_reducible_to_via_def dest: pr_is_total_rec) then show ?thesis by (rule many_reducible_lm_1) qed text \The remaining results are resursive inseparability results. We start be showing that there is a Minksy machine $U$ with final states $0$ and $1$ such that it is not possible to recursively separate inputs reaching state $0$ from inputs reaching state $1$.\ lemma rec_inseparable_0not1_1not0: "rec_inseparable {p. 0 \ nat_to_ce_set p \ 1 \ nat_to_ce_set p} {p. 0 \ nat_to_ce_set p \ 1 \ nat_to_ce_set p}" proof - obtain n where n: "nat_to_ce_set n = {0}" using nat_to_ce_set_srj[OF ce_finite[of "{0}"]] by auto obtain m where m: "nat_to_ce_set m = {1}" using nat_to_ce_set_srj[OF ce_finite[of "{1}"]] by auto show ?thesis by (rule rec_inseparable_mono[OF Rice_rec_inseparable[of n m]]) (auto simp: n m) qed lemma ce_sets_containing_n_ce: "{p. n \ nat_to_ce_set p} \ ce_sets" using ce_set_lm_5[OF univ_is_ce comp2_1[OF c_pair_is_pr id1_1 const_is_pr[of n]]] by (auto simp: univ_ce_lm_1) lemma rec_inseparable_fixed_minsky_reaching_0_1: obtains U :: "(nat, nat) minsky" where "finite U" "deterministic U" "0 \ fst ` U" "1 \ fst ` U" "rec_inseparable {x |x vs'. ((2, (\n. if n = 0 then x else 0)), (0, vs')) \ (step U)\<^sup>*} {x |x vs'. ((2, (\n. if n = 0 then x else 0)), (1, vs')) \ (step U)\<^sup>*}" proof - guess U using ce_set_pair_by_minsky[OF ce_sets_containing_n_ce ce_sets_containing_n_ce, of 0 1] . from this(1-4) this(5-7)[of "\n. if n = 0 then _ else 0"] show ?thesis by (auto 0 0 intro!: that[of U] rec_inseparable_mono[OF rec_inseparable_0not1_1not0] pr_is_total_rec simp: rev_image_eqI cong: if_cong) meson+ qed text \Consequently, it is impossible to separate Minsky configurations with determistic machines and final states $0$ and $1$ that reach state $0$ from those that reach state $1$.\ definition minsky_reaching_s where "minsky_reaching_s s = {m |M m vs vs'. (M, vs) = decode_minsky_state m \ deterministic M \ 0 \ fst ` M \ 1 \ fst ` M \ ((2, vs), (s, vs')) \ (step M)\<^sup>*}" lemma rec_inseparable_minsky_reaching_0_1: "rec_inseparable (minsky_reaching_s 0) (minsky_reaching_s 1)" proof - guess U using rec_inseparable_fixed_minsky_reaching_0_1 . note U = this obtain us where [simp]: "set us = U" using finite_list[OF U(1)] by blast let ?f = "\n. c_pair (encode_minsky us) (c_cons n 0)" have "?f \ PrimRec1" using comp2_1[OF c_pair_is_pr const_is_pr comp2_1[OF c_cons_is_pr id1_1 const_is_pr]] by simp then show ?thesis using U(1-4) rec_inseparable_many_reducible[of ?f, OF _ rec_inseparable_mono[OF U(5)]] by (auto simp: pr_is_total_rec minsky_reaching_s_def decode_minsky_state_def rev_image_eqI decode_regs_def c_cons_def cong: if_cong) qed text \As a corollary, it is impossible to separate Minsky configurations that reach state 0 but not state 1 from those that reach state 1 but not state 0.\ definition minsky_reaching_s_not_t where "minsky_reaching_s_not_t s t = {m |M m vs vs'. (M, vs) = decode_minsky_state m \ ((2, vs), (s, vs')) \ (step M)\<^sup>* \ ((2, vs), (t, vs')) \ (step M)\<^sup>*}" lemma minsky_reaching_s_imp_minsky_reaching_s_not_t: assumes "s \ {0,1}" "t \ {0,1}" "s \ t" shows "minsky_reaching_s s \ minsky_reaching_s_not_t s t" proof - have [dest!]: "((2, vs), (0, vs')) \ (step M)\<^sup>* \ ((2, vs), (1, vs')) \ (step M)\<^sup>*" if "deterministic M" "0 \ fst ` M" "1 \ fst ` M" for M :: "(nat, nat) minsky" and vs vs' using deterministic_minsky_UN[OF that(1) _ _ that(2,3)] by auto show ?thesis using assms by (auto simp: minsky_reaching_s_def minsky_reaching_s_not_t_def rev_image_eqI) qed lemma rec_inseparable_minsky_reaching_0_not_1_1_not_0: "rec_inseparable (minsky_reaching_s_not_t 0 1) (minsky_reaching_s_not_t 1 0)" by (intro rec_inseparable_mono[OF rec_inseparable_minsky_reaching_0_1] minsky_reaching_s_imp_minsky_reaching_s_not_t) simp_all end diff --git a/thys/Minsky_Machines/ROOT b/thys/Minsky_Machines/ROOT --- a/thys/Minsky_Machines/ROOT +++ b/thys/Minsky_Machines/ROOT @@ -1,12 +1,13 @@ chapter AFP session "Minsky_Machines" (AFP) = "Abstract-Rewriting" + options [timeout = 600] sessions + "Pure-ex" "Recursion-Theory-I" theories Recursive_Inseparability Minsky document_files "root.tex" "root.bib" diff --git a/thys/Myhill-Nerode/Non_Regular_Languages.thy b/thys/Myhill-Nerode/Non_Regular_Languages.thy --- a/thys/Myhill-Nerode/Non_Regular_Languages.thy +++ b/thys/Myhill-Nerode/Non_Regular_Languages.thy @@ -1,269 +1,272 @@ (* File: Non_Regular_Languages.thy Author: Manuel Eberl This file provides some tools for showing the non-regularity of a language, either via an infinite set of equivalence classes or via the Pumping Lemma. *) section \Tools for showing non-regularity of a language\ theory Non_Regular_Languages imports Myhill begin subsection \Auxiliary material\ lemma bij_betw_image_quotient: "bij_betw (\y. f -` {y}) (f ` A) (A // {(a,b). f a = f b})" by (force simp: bij_betw_def inj_on_def image_image quotient_def) lemma regular_Derivs_finite: fixes r :: "'a :: finite rexp" shows "finite (range (\w. Derivs w (lang r)))" proof - have "?thesis \ finite (UNIV // \lang r)" unfolding str_eq_conv_Derivs by (rule bij_betw_finite bij_betw_image_quotient)+ also have "\" by (subst Myhill_Nerode [symmetric]) auto finally show ?thesis . qed lemma Nil_in_Derivs_iff: "[] \ Derivs w A \ w \ A" by (auto simp: Derivs_def) (* TODO: Move to distribution? *) text \ The following operation repeats a list $n$ times (usually written as $w ^ n$). \ primrec repeat :: "nat \ 'a list \ 'a list" where "repeat 0 xs = []" | "repeat (Suc n) xs = xs @ repeat n xs" lemma repeat_Cons_left: "repeat (Suc n) xs = xs @ repeat n xs" by simp lemma repeat_Cons_right: "repeat (Suc n) xs = repeat n xs @ xs" by (induction n) simp_all lemma repeat_Cons_append_commute [simp]: "repeat n xs @ xs = xs @ repeat n xs" by (subst repeat_Cons_right [symmetric]) simp lemma repeat_Cons_add [simp]: "repeat (m + n) xs = repeat m xs @ repeat n xs" by (induction m) simp_all lemma repeat_Nil [simp]: "repeat n [] = []" by (induction n) simp_all lemma repeat_conv_replicate: "repeat n xs = concat (replicate n xs)" by (induction n) simp_all (* TODO: Move to distribution? *) lemma nth_prefixes [simp]: "n \ length xs \ prefixes xs ! n = take n xs" by (induction xs arbitrary: n) (auto simp: nth_Cons split: nat.splits) lemma nth_suffixes [simp]: "n \ length xs \ suffixes xs ! n = drop (length xs - n) xs" by (subst suffixes_conv_prefixes) (simp_all add: rev_take) lemma length_take_prefixes: assumes "xs \ set (take n (prefixes ys))" shows "length xs < n" proof (cases "n \ Suc (length ys)") case True with assms obtain i where "i < n" "xs = take i ys" by (subst (asm) nth_image [symmetric]) auto thus ?thesis by simp next case False with assms have "prefix xs ys" by simp hence "length xs \ length ys" by (rule prefix_length_le) also from False have "\ < n" by simp finally show ?thesis . qed subsection \Non-regularity by giving an infinite set of equivalence classes\ text \ Non-regularity can be shown by giving an infinite set of non-equivalent words (w.r.t. the Myhill--Nerode relation). \ lemma not_regular_langI: assumes "infinite B" "\x y. x \ B \ y \ B \ x \ y \ \w. \(x @ w \ A \ y @ w \ A)" shows "\regular_lang (A :: 'a :: finite list set)" proof - have "(\w. Derivs w A) ` B \ range (\w. Derivs w A)" by blast moreover from assms(2) have "inj_on (\w. Derivs w A) B" by (auto simp: inj_on_def Derivs_def) with assms(1) have "infinite ((\w. Derivs w A) ` B)" by (blast dest: finite_imageD) ultimately have "infinite (range (\w. Derivs w A))" by (rule infinite_super) with regular_Derivs_finite show ?thesis by blast qed lemma not_regular_langI': assumes "infinite B" "\x y. x \ B \ y \ B \ x \ y \ \w. \(f x @ w \ A \ f y @ w \ A)" shows "\regular_lang (A :: 'a :: finite list set)" proof (rule not_regular_langI) from assms(2) have "inj_on f B" by (force simp: inj_on_def) with \infinite B\ show "infinite (f ` B)" by (simp add: finite_image_iff) qed (insert assms, auto) subsection \The Pumping Lemma\ text \ The Pumping lemma can be shown very easily from the Myhill--Nerode theorem: if we have a word whose length is more than the (finite) number of equivalence classes, then it must have two different prefixes in the same class and the difference between these two prefixes can then be ``pumped''. \ lemma pumping_lemma_aux: fixes A :: "'a list set" defines "\ \ \w. Derivs w A" defines "n \ card (range \)" assumes "z \ A" "finite (range \)" "length z \ n" shows "\u v w. z = u @ v @ w \ length (u @ v) \ n \ v \ [] \ (\i. u @ repeat i v @ w \ A)" proof - define P where "P = set (take (Suc n) (prefixes z))" from \length z \ n\ have [simp]: "card P = Suc n" unfolding P_def by (subst distinct_card) (auto intro!: distinct_take) have length_le: "length y \ n" if "y \ P" for y using length_take_prefixes[OF that [unfolded P_def]] by simp have "card (\ ` P) \ card (range \)" by (intro card_mono assms) auto also from assms have "\ < card P" by simp finally have "\inj_on \ P" by (rule pigeonhole) then obtain a b where ab: "a \ P" "b \ P" "a \ b" "Derivs a A = Derivs b A" by (auto simp: inj_on_def \_def) from ab have prefix_ab: "prefix a z" "prefix b z" by (auto simp: P_def dest: in_set_takeD) from ab have length_ab: "length a \ n" "length b \ n" by (simp_all add: length_le) have *: ?thesis if uz': "prefix u z'" "prefix z' z" "length z' \ n" "u \ z'" "Derivs z' A = Derivs u A" for u z' proof - from \prefix u z'\ and \u \ z'\ obtain v where v [simp]: "z' = u @ v" "v \ []" by (auto simp: prefix_def) from \prefix z' z\ obtain w where [simp]: "z = u @ v @ w" by (auto simp: prefix_def) hence [simp]: "Derivs (repeat i v) (Derivs u A) = Derivs u A" for i by (induction i) (use uz' in simp_all) have "Derivs z A = Derivs (u @ repeat i v @ w) A" for i using uz' by simp with \z \ A\ and uz' have "\i. u @ repeat i v @ w \ A" by (simp add: Nil_in_Derivs_iff [of _ A, symmetric]) moreover have "z = u @ v @ w" by simp moreover from \length z' \ n\ have "length (u @ v) \ n" by simp ultimately show ?thesis using \v \ []\ by blast qed from prefix_ab have "prefix a b \ prefix b a" by (rule prefix_same_cases) with *[of a b] and *[of b a] and ab and prefix_ab and length_ab show ?thesis by blast qed theorem pumping_lemma: fixes r :: "'a :: finite rexp" obtains n where "\z. z \ lang r \ length z \ n \ \u v w. z = u @ v @ w \ length (u @ v) \ n \ v \ [] \ (\i. u @ repeat i v @ w \ lang r)" proof - let ?n = "card (range (\w. Derivs w (lang r)))" have "\u v w. z = u @ v @ w \ length (u @ v) \ ?n \ v \ [] \ (\i. u @ repeat i v @ w \ lang r)" if "z \ lang r" and "length z \ ?n" for z by (intro pumping_lemma_aux[of z] that regular_Derivs_finite) thus ?thesis by (rule that) qed corollary pumping_lemma_not_regular_lang: fixes A :: "'a :: finite list set" assumes "\n. length (z n) \ n" and "\n. z n \ A" assumes "\n u v w. z n = u @ v @ w \ length (u @ v) \ n \ v \ [] \ u @ repeat (i n u v w) v @ w \ A" shows "\regular_lang A" proof assume "regular_lang A" then obtain r where r: "lang r = A" by blast - from pumping_lemma[of r] guess n . - from this[of "z n"] and assms[of n] obtain u v w + from pumping_lemma[of r] obtain n + where "z n \ lang r \ n \ length (z n) \ + \u v w. z n = u @ v @ w \ length (u @ v) \ n \ v \ [] \ (\i. u @ repeat i v @ w \ lang r)" + by metis + from this and assms[of n] obtain u v w where "z n = u @ v @ w" and "length (u @ v) \ n" and "v \ []" and "\i. u @ repeat i v @ w \ lang r" by (auto simp: r) with assms(3)[of n u v w] show False by (auto simp: r) qed subsection \Examples\ text \ The language of all words containing the same number of $a$s and $b$s is not regular. \ lemma "\regular_lang {w. length (filter id w) = length (filter Not w)}" (is "\regular_lang ?A") proof (rule not_regular_langI') show "infinite (UNIV :: nat set)" by simp next fix m n :: nat assume "m \ n" hence "replicate m True @ replicate m False \ ?A" and "replicate n True @ replicate m False \ ?A" by simp_all thus "\w. \(replicate m True @ w \ ?A \ replicate n True @ w \ ?A)" by blast qed text \ The language $\{a^i b^i\ |\ i \in \mathbb{N}\}$ is not regular. \ lemma eq_replicate_iff: "xs = replicate n x \ set xs \ {x} \ length xs = n" using replicate_length_same[of xs x] by (subst eq_commute) auto lemma replicate_eq_appendE: assumes "xs @ ys = replicate n x" obtains i j where "n = i + j" "xs = replicate i x" "ys = replicate j x" proof - have "n = length (replicate n x)" by simp also note assms [symmetric] finally have "n = length xs + length ys" by simp moreover have "xs = replicate (length xs) x" and "ys = replicate (length ys) x" using assms by (auto simp: eq_replicate_iff) ultimately show ?thesis using that[of "length xs" "length ys"] by auto qed lemma "\regular_lang (range (\i. replicate i True @ replicate i False))" (is "\regular_lang ?A") proof (rule pumping_lemma_not_regular_lang) fix n :: nat show "length (replicate n True @ replicate n False) \ n" by simp show "replicate n True @ replicate n False \ ?A" by simp next fix n :: nat and u v w :: "bool list" assume decomp: "replicate n True @ replicate n False = u @ v @ w" and length_le: "length (u @ v) \ n" and v_ne: "v \ []" define w1 w2 where "w1 = take (n - length (u@v)) w" and "w2 = drop (n - length (u@v)) w" have w_decomp: "w = w1 @ w2" by (simp add: w1_def w2_def) have "replicate n True = take n (replicate n True @ replicate n False)" by simp also note decomp also have "take n (u @ v @ w) = u @ v @ w1" using length_le by (simp add: w1_def) finally have "u @ v @ w1 = replicate n True" by simp then obtain i j k where uvw1: "n = i + j + k" "u = replicate i True" "v = replicate j True" "w1 = replicate k True" by (elim replicate_eq_appendE) auto have "replicate n False = drop n (replicate n True @ replicate n False)" by simp also note decomp finally have [simp]: "w2 = replicate n False" using length_le by (simp add: w2_def) have "u @ repeat (Suc (Suc 0)) v @ w = replicate (n + j) True @ replicate n False" by (simp add: uvw1 w_decomp replicate_add [symmetric]) also have "\ \ ?A" proof safe fix m assume *: "replicate (n + j) True @ replicate n False = replicate m True @ replicate m False" (is "?lhs = ?rhs") have "n = length (filter Not ?lhs)" by simp also note * also have "length (filter Not ?rhs) = m" by simp finally have [simp]: "m = n" by simp from * have "v = []" by (simp add: uvw1) with \v \ []\ show False by contradiction qed finally show "u @ repeat (Suc (Suc 0)) v @ w \ ?A" . qed end diff --git a/thys/Ordinary_Differential_Equations/Numerics/Refine_Rigorous_Numerics_Aform.thy b/thys/Ordinary_Differential_Equations/Numerics/Refine_Rigorous_Numerics_Aform.thy --- a/thys/Ordinary_Differential_Equations/Numerics/Refine_Rigorous_Numerics_Aform.thy +++ b/thys/Ordinary_Differential_Equations/Numerics/Refine_Rigorous_Numerics_Aform.thy @@ -1,1740 +1,1744 @@ theory Refine_Rigorous_Numerics_Aform imports Refine_Rigorous_Numerics "HOL-Types_To_Sets.Types_To_Sets" begin lemma Joints_ne_empty[simp]: "Joints xs \ {}" "{} \ Joints xs" by (auto simp: Joints_def valuate_def) lemma Inf_aform_le_Affine: "x \ Affine X \ Inf_aform X \ x" by (auto simp: Affine_def valuate_def intro!: Inf_aform) lemma Sup_aform_ge_Affine: "x \ Affine X \ x \ Sup_aform X" by (auto simp: Affine_def valuate_def intro!: Sup_aform) lemmas Inf_aform'_Affine_le = order_trans[OF Inf_aform' Inf_aform_le_Affine] lemmas Sup_aform'_Affine_ge = order_trans[OF Sup_aform_ge_Affine Sup_aform'] fun approx_form_aform :: "nat \ form \ real aform list \ bool" where "approx_form_aform prec (Less a b) bs = (case (approx_floatariths prec [a - b] bs) of Some [r] \ Sup_aform' prec r < 0 | _ \ False)" | "approx_form_aform prec (LessEqual a b) bs = (case (approx_floatariths prec [a - b] bs) of Some [r] \ Sup_aform' prec r \ 0 | _ \ False)" | "approx_form_aform prec (AtLeastAtMost a b c) bs = (case (approx_floatariths prec [a - b, a - c] bs) of Some [r, s] \ 0 \ Inf_aform' prec r \ Sup_aform' prec s \ 0 | _ \ False)" | "approx_form_aform prec (Bound a b c d) bs = False" | "approx_form_aform prec (Assign a b c) bs = False" | "approx_form_aform prec (Conj a b) bs \ approx_form_aform prec a bs \ approx_form_aform prec b bs" | "approx_form_aform prec (Disj a b) bs \ approx_form_aform prec a bs \ approx_form_aform prec b bs" lemma in_Joints_intervalD: "x \ {Inf_aform' p X .. Sup_aform' p X} \ xs \ Joints XS" if "x#xs \ Joints (X#XS)" using that by (auto simp: Joints_def valuate_def Affine_def intro!: Inf_aform'_Affine_le Sup_aform'_Affine_ge) lemma approx_form_aform: "interpret_form f vs" if "approx_form_aform p f VS" "vs \ Joints VS" using that by (induction f) (auto split: option.splits list.splits simp: algebra_simps dest!: approx_floatariths_outer dest!: in_Joints_intervalD[where p=p]) fun msum_aforms::"nat \ real aform list \ real aform list \ real aform list" where "msum_aforms d (x#xs) (y#ys) = msum_aform d x y#msum_aforms d xs ys" | "msum_aforms d _ _ = []" definition [simp]: "degree_aforms_real = (degree_aforms::real aform list \ nat)" abbreviation "msum_aforms' \ \X Y. msum_aforms (degree_aforms_real (X @ Y)) X Y" lemma aform_val_msum_aforms: assumes "degree_aforms xs \ d" shows "aform_vals e (msum_aforms d xs ys) = List.map2 (+) (aform_vals e xs) (aform_vals (\i. e (i + d)) ys)" using assms proof (induction xs ys rule: msum_aforms.induct) case (1 d x xs y ys) from 1 have "degree_aforms xs \ d" by (auto simp: degrees_def) from 1(1)[OF this] 1 have "aform_vals e (msum_aforms d xs ys) = List.map2 (+) (aform_vals e xs) (aform_vals (\i. e (i + d)) ys)" by simp then show ?case using 1 by (simp add: aform_vals_def aform_val_msum_aform degrees_def) qed (auto simp: aform_vals_def) lemma Joints_msum_aforms: assumes "degree_aforms xs \ d" assumes "degree_aforms ys \ d" shows "Joints (msum_aforms d xs ys) = {List.map2 (+) a b |a b. a \ Joints xs \ b \ Joints ys}" apply (auto simp: Joints_def valuate_def aform_vals_def[symmetric] aform_val_msum_aforms assms) apply force subgoal for e e' apply (rule image_eqI[where x="\i. if i < d then e i else e' (i - d)"]) apply (auto simp: Pi_iff) proof goal_cases case 1 have "(aform_vals e xs) = aform_vals (\i. if i < d then e i else e' (i - d)) xs" using assms by (auto intro!: simp: aform_vals_def aform_val_def degrees_def intro!: pdevs_val_degree_cong) then show ?case by simp qed done definition "split_aforms_largest_uncond_take n X = (let (i, x) = max_pdev (abssum_of_pdevs_list (map snd (take n X))) in split_aforms X i)" text \intersection with plane\ definition "project_coord x b n = (\i\Basis_list. (if i = b then n else if i = -b then -n else x \ i) *\<^sub>R i)" lemma inner_eq_abs_times_sgn: "u \ b = abs u \ b * sgn (u \ b)" if "b \ Basis" for b::"'a::ordered_euclidean_space" by (subst sgn_mult_abs[symmetric]) (auto simp: that abs_inner ) lemma inner_Basis_eq_zero_absI: "x \ Basis \ abs u \ Basis \ x \ \u\ \ x \ 0 \ u \ x = 0" for x::"'a::ordered_euclidean_space" apply (subst euclidean_inner) apply (auto simp: inner_Basis if_distribR if_distrib cong: if_cong) apply (subst inner_eq_abs_times_sgn) by (auto simp: inner_Basis) lemma abs_in_BasisE: fixes u::"'a::ordered_euclidean_space" assumes "abs u \ Basis" obtains "abs u = u" | "abs u = - u" proof (cases "u \ abs u = 1") case True have "abs u = (\i\Basis. (if i = abs u then (abs u \ i) *\<^sub>R i else 0))" using assms by (auto simp: ) also have "\ = (\i\Basis. (if i = abs u then (u \ i) *\<^sub>R i else 0))" by (rule sum.cong) (auto simp: True) also have "\ = (\i\Basis. (u \ i) *\<^sub>R i)" by (rule sum.cong) (auto simp: inner_Basis_eq_zero_absI assms) also have "\ = u" by (simp add: euclidean_representation) finally show ?thesis .. next case False then have F: "u \ \u\ = -1" apply (subst inner_eq_abs_times_sgn[OF assms]) apply (subst (asm) inner_eq_abs_times_sgn[OF assms]) using assms apply (auto simp: inner_Basis sgn_if) by (metis (full_types) abs_0_eq euclidean_all_zero_iff inner_Basis_eq_zero_absI nonzero_Basis) have "abs u = (\i\Basis. (if i = abs u then (abs u \ i) *\<^sub>R i else 0))" using assms by (auto simp: ) also have "\ = (\i\Basis. (if i = abs u then (- u \ i) *\<^sub>R i else 0))" by (rule sum.cong) (auto simp: F) also have "\ = (\i\Basis. (- u \ i) *\<^sub>R i)" by (rule sum.cong) (auto simp: inner_Basis_eq_zero_absI assms) also have "\ = - u" by (subst euclidean_representation) simp finally show ?thesis .. qed lemma abs_in_Basis_iff: fixes u::"'a::ordered_euclidean_space" shows "abs u \ Basis \ u \ Basis \ - u \ Basis" proof - have u: "u = (\i\Basis. (u \ i) *\<^sub>R i)" by (simp add: euclidean_representation) have u': "- u = (\i\Basis. (- (u \ i)) *\<^sub>R i)" by (subst u) (simp add: sum_negf) have au: "abs u = (\i\Basis. \u \ i\ *\<^sub>R i)" by (simp add: eucl_abs[where 'a='a]) have "(u \ Basis \ - u \ Basis)" if "(\u\ \ Basis)" apply (rule abs_in_BasisE[OF that]) using that by auto show ?thesis apply safe subgoal premises prems using prems(1) apply (rule abs_in_BasisE) using prems by simp_all subgoal apply (subst au) apply (subst (asm) u) apply (subst sum.cong[OF refl]) apply (subst abs_inner[symmetric]) apply auto using u apply auto[1] done subgoal apply (subst au) apply (subst (asm) u') apply (subst sum.cong[OF refl]) apply (subst abs_inner[symmetric]) apply auto using u' apply auto by (metis Basis_nonneg abs_of_nonpos inner_minus_left neg_0_le_iff_le scaleR_left.minus) done qed lemma abs_inner_Basis: fixes u v::"'a::ordered_euclidean_space" assumes "abs u \ Basis" "abs v \ Basis" shows "inner u v = (if u = v then 1 else if u = -v then -1 else 0)" proof - define i where "i \ if u \ Basis then u else -u" define j where "j \ if v \ Basis then v else -v" have u: "u = (if u \ Basis then i else - i)" and v: "v = (if v \ Basis then j else - j)" by (auto simp: i_def j_def) have "i \ Basis" "j \ Basis" using assms by (auto simp: i_def j_def abs_in_Basis_iff) then show ?thesis apply (subst u) apply (subst v) apply (auto simp: inner_Basis) apply (auto simp: j_def i_def split: if_splits) using Basis_nonneg abs_of_nonpos by fastforce qed lemma project_coord_inner_Basis: assumes "i \ Basis" shows "(project_coord x b n) \ i = (if i = b then n else if i = -b then -n else x \ i)" proof - have "project_coord x b n \ i = (\j\Basis. (if j = b then n else if j = -b then -n else x \ j) * (if j = i then 1 else 0))" using assms by (auto simp: project_coord_def inner_sum_left inner_Basis) also have "\ = (\j\Basis. (if j = i then (if j = b then n else if j = -b then -n else x \ j) else 0))" by (rule sum.cong) auto also have "\ = (if i = b then n else if i = -b then -n else x \ i)" using assms by (subst sum.delta) auto finally show ?thesis by simp qed lemma project_coord_inner: assumes "abs i \ Basis" shows "(project_coord x b n) \ i = (if i = b then n else if i = -b then -n else x \ i)" proof - consider "i \ Basis" | "- i \ Basis" using assms by (auto simp: abs_in_Basis_iff) then show ?thesis proof cases case 1 then show ?thesis by (simp add: project_coord_inner_Basis) next case 2 have "project_coord x b n \ i = - (project_coord x b n \ - i)" by simp also have "\ = - (if - i = b then n else if - i = -b then -n else x \ - i)" using 2 by (subst project_coord_inner_Basis) simp_all also have "\ = (if i = b then n else if i = -b then -n else x \ i)" using 2 by auto (metis Basis_nonneg abs_le_zero_iff abs_of_nonneg neg_le_0_iff_le nonzero_Basis) finally show ?thesis . qed qed lift_definition project_coord_pdevs::"'a::executable_euclidean_space sctn \ 'a pdevs \ 'a pdevs" is "\sctn xs i. project_coord (xs i) (normal sctn) 0" by (rule finite_subset) (auto simp: project_coord_def cong: if_cong) definition "project_coord_aform sctn cxs = (project_coord (fst cxs) (normal sctn) (pstn sctn), project_coord_pdevs sctn (snd cxs))" definition project_coord_aform_lv :: "real aform list \ nat \ real \ real aform list" where "project_coord_aform_lv xs b n = xs[b:=(n, zero_pdevs)]" definition "project_ncoord_aform x b n = project_coord_aform (Sctn (Basis_list ! b) n) x" lemma finite_sum_subset_aux_lemma: assumes "finite s" shows " {i. (\x\s. f x i) \ 0} \ {i. \x\s. f x i \ 0}" proof - have "{i. (\x\s. f x i) \ 0} = - {i. (\x\s. f x i) = 0}" by auto also have "\ \ - {i. \x \ s. f x i = 0}" by auto also have "\ = {i. \x \ s. f x i \ 0}" by auto finally show ?thesis by simp qed lift_definition sum_pdevs::"('i \ 'a::comm_monoid_add pdevs) \ 'i set \ 'a pdevs" is "\f X. if finite X then (\i. \x\X. f x i) else (\_. 0)" apply auto apply (rule finite_subset) apply (rule finite_sum_subset_aux_lemma) apply auto done lemma pdevs_apply_sum_pdevs[simp]: "pdevs_apply (sum_pdevs f X) i = (\x\X. pdevs_apply (f x) i)" by transfer auto lemma sum_pdevs_empty[simp]: "sum_pdevs f {} = zero_pdevs" by transfer auto lemma sum_pdevs_insert[simp]: "finite xs \ sum_pdevs f (insert a xs) = (if a \ xs then sum_pdevs f xs else add_pdevs (f a) (sum_pdevs f xs))" by (auto simp: insert_absorb intro!: pdevs_eqI) lemma sum_pdevs_infinite[simp]: "infinite X \ sum_pdevs f X = zero_pdevs" by transfer auto lemma compute_sum_pdevs[code]: "sum_pdevs f (set XS) = foldr (\a b. add_pdevs (f a) b) (remdups XS) zero_pdevs" by (induction XS) (auto simp: ) lemma degree_sum_pdevs_le: "degree (sum_pdevs f X) \ Max (degree ` f ` X)" apply (rule degree_le) apply auto apply (cases "X = {}") subgoal by (simp add: ) subgoal by (cases "finite X") simp_all done lemma pdevs_val_sum_pdevs: "pdevs_val e (sum_pdevs f X) = (\x\X. pdevs_val e (f x))" apply (cases "finite X") subgoal apply (subst pdevs_val_sum_le) apply (rule degree_sum_pdevs_le) apply (auto simp: scaleR_sum_right) apply (subst sum.swap) apply (rule sum.cong[OF refl]) apply (subst pdevs_val_sum_le[where d="Max (degree ` f ` X)"]) apply (auto simp: Max_ge_iff) done subgoal by simp done definition eucl_of_list_aform :: "(real \ real pdevs) list \ 'a::executable_euclidean_space aform" where "eucl_of_list_aform xs = (eucl_of_list (map fst xs), sum_pdevs (\i. pdevs_scaleR (snd (xs ! index Basis_list i)) i) Basis)" definition lv_aforms_rel::"(((real \ real pdevs) list) \ ('a::executable_euclidean_space aform)) set" where "lv_aforms_rel = br eucl_of_list_aform (\xs. length xs = DIM('a))" definition "inner_aforms' p X Y = (let fas = [inner_floatariths (map floatarith.Var [0..d x y. Some (F d x y)) f" assumes "[x, y] \ Joints [X, Y]" assumes "d \ degree_aform X" assumes "d \ degree_aform Y" shows "f x y \ Affine (F d X Y)" proof - from assms(2) obtain e where e: "e \ funcset UNIV {-1 .. 1}" "x = aform_val e X" "y = aform_val e Y" by (auto simp: Joints_def valuate_def) from affine_extension2E[OF assms(1) refl e(1) assms(3) assms(4)] obtain e' where "e' \ funcset UNIV {- 1..1}" "f (aform_val e X) (aform_val e Y) = aform_val e' (F d X Y)" "\n. n < d \ e' n = e n" "aform_val e X = aform_val e' X" "aform_val e Y = aform_val e' Y" by auto then show ?thesis by (auto simp: Affine_def valuate_def e) qed lemma aform_val_zero_pdevs[simp]: "aform_val e (x, zero_pdevs) = x" by (auto simp: aform_val_def) lemma neg_equal_zero_eucl[simp]: "-a = a \ a = 0" for a::"'a::euclidean_space" by (auto simp: algebra_simps euclidean_eq_iff[where 'a='a]) context includes autoref_syntax begin lemma Option_bind_param[param, autoref_rules]: "((\), (\)) \ \S\option_rel \ (S \ \R\option_rel) \ \R\option_rel" unfolding Option.bind_def by parametricity lemma zip_Basis_list_pat[autoref_op_pat_def]: "\(b, m)\zip Basis_list ms. m *\<^sub>R b \ OP eucl_of_list $ ms" proof (rule eq_reflection) have z: "zip ms (Basis_list::'a list) = map (\(x, y). (y, x)) (zip Basis_list ms)" by (subst zip_commute) simp show "(\(b, m)\zip (Basis_list::'a list) ms. m *\<^sub>R b) = OP eucl_of_list $ ms" unfolding eucl_of_list_def autoref_tag_defs apply (subst z) apply (subst map_map) apply (auto cong: map_cong simp: split_beta') done qed lemma length_aforms_of_ivls: "length (aforms_of_ivls a b) = min (length a) (length b)" by (auto simp: aforms_of_ivls_def) lemma length_lv_rel: "(ys, y::'a) \ lv_rel \ length ys = DIM('a::executable_euclidean_space)" by (auto simp: lv_rel_def br_def) lemma lv_rel_nth[simp]: "(xs, x::'a::executable_euclidean_space) \ lv_rel \ b \ Basis \ xs ! index (Basis_list) b = x \ b" unfolding lv_rel_def by (auto simp: br_def eucl_of_list_inner) lemma aform_of_ivl_autoref[autoref_rules]: "(aforms_of_ivls, aform_of_ivl) \ lv_rel \ lv_rel \ lv_aforms_rel" apply (auto simp: lv_aforms_rel_def br_def eucl_of_list_aform_def length_aforms_of_ivls length_lv_rel) subgoal for xs x ys y apply (auto simp: aform_of_ivl_def aforms_of_ivls_def o_def eucl_of_list_inner inner_simps pdevs_apply_pdevs_of_ivl length_lv_rel intro!: euclidean_eqI[where 'a='a] pdevs_eqI) by (metis index_Basis_list_nth inner_not_same_Basis length_Basis_list nth_Basis_list_in_Basis) done lemma bound_intersect_2d_ud[autoref_rules]: shows "(bound_intersect_2d_ud, bound_intersect_2d_ud) \ nat_rel \ ((rnv_rel \\<^sub>r rnv_rel) \\<^sub>r Id) \ rnv_rel \ \rnv_rel \\<^sub>r rnv_rel\option_rel" by auto lemma eucl_of_list_autoref[autoref_rules]: includes autoref_syntax assumes "SIDE_PRECOND (length xs = DIM('a::executable_euclidean_space))" assumes "(xsi, xs) \ \rnv_rel\list_rel" shows "(xsi, eucl_of_list $ xs::'a) \ lv_rel" using assms unfolding lv_rel_def by (auto simp: br_def) definition "inner2s x b c = (inner_lv_rel x b, inner_lv_rel x c)" lemma inner_lv_rel_autoref[autoref_rules]: "(inner_lv_rel, (\)) \ lv_rel \ lv_rel \ rnv_rel" using lv_rel_inner[unfolded inner_lv_rel_def[symmetric]] by auto lemma inner_lv_rel_eq: "\length xs = DIM('a::executable_euclidean_space); (xa, x'a) \ lv_rel\ \ inner_lv_rel xs xa = eucl_of_list xs \ (x'a::'a)" using inner_lv_rel_autoref[param_fo, of "xs" "eucl_of_list xs" xa x'a] unfolding lv_rel_def by (auto simp: br_def) definition "inner_pdevs xs ys = sum_pdevs (\i. scaleR_pdevs (ys ! i) (xs ! i)) {..xs a b. (inner2s (map fst xs) a b, pdevs_inner2s (map snd xs) a b)), inner2_aform) \ lv_aforms_rel \ lv_rel \ lv_rel \ ((rnv_rel \\<^sub>r rnv_rel)\\<^sub>rId)" unfolding inner2_aform_def apply (auto simp: lv_aforms_rel_def br_def eucl_of_list_aform_def inner2_aform_def ) apply (auto simp: inner2s_def inner2_def inner_lv_rel_eq pdevs_inner2s_def inner_pdevs_def sum_Basis_sum_nth_Basis_list inner_sum_left inner_Basis mult.commute intro!: pdevs_eqI) subgoal for a b c d e f apply (rule sum.cong) apply force subgoal for n by (auto dest!: lv_rel_nth[where b="Basis_list ! n"] simp: inner_commute) done subgoal for a b c d e f apply (rule sum.cong) apply force subgoal for n by (auto dest!: lv_rel_nth[where b="Basis_list ! n"] simp: inner_commute) done done lemma RETURN_inter_aform_plane_ortho_annot: "RETURN (inter_aform_plane_ortho p (Z::'a aform) n g) = (case those (map (\b. bound_intersect_2d_ud p (inner2_aform Z n b) g) Basis_list) of Some mMs \ do { ASSERT (length mMs = DIM('a::executable_euclidean_space)); 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); RETURN (Some (aform_of_ivl l u)) } | None \ RETURN None)" apply (auto simp: inter_aform_plane_ortho_def split: option.splits) subgoal premises prems for x2 proof - have "length x2 = DIM('a)" using prems using map_eq_imp_length_eq by (force simp: those_eq_Some_map_Some_iff) then show ?thesis using prems by auto qed done definition "inter_aform_plane_ortho_nres p Z n g = RETURN (inter_aform_plane_ortho p Z n g)" schematic_goal inter_aform_plane_ortho_lv: fixes Z::"'a::executable_euclidean_space aform" assumes [autoref_rules_raw]: "DIM_precond TYPE('a) D" assumes [autoref_rules]: "(pp, p) \ nat_rel" "(Zi, Z) \ lv_aforms_rel" "(ni, n) \ lv_rel" "(gi, g) \ rnv_rel" notes [autoref_rules] = those_param param_map shows "(nres_of (?f::?'a dres), inter_aform_plane_ortho_nres $ p $ Z $ n $ g) \ ?R" unfolding autoref_tag_defs unfolding RETURN_inter_aform_plane_ortho_annot inter_aform_plane_ortho_nres_def including art by autoref_monadic concrete_definition inter_aform_plane_ortho_lv for pp Zi ni gi uses inter_aform_plane_ortho_lv lemmas [autoref_rules] = inter_aform_plane_ortho_lv.refine lemma project_coord_lv[autoref_rules]: assumes "(xs, x::'a::executable_euclidean_space aform) \ lv_aforms_rel" "(ni, n) \ nat_rel" assumes "SIDE_PRECOND (n < DIM('a))" shows "(project_coord_aform_lv xs ni, project_ncoord_aform $ x $ n) \ rnv_rel \ lv_aforms_rel" using assms apply (auto simp: project_coord_aform_lv_def project_ncoord_aform_def lv_rel_def project_coord_aform_def eucl_of_list_aform_def lv_aforms_rel_def br_def) subgoal apply (auto intro!: euclidean_eqI[where 'a='a]) apply (subst project_coord_inner_Basis) apply (auto simp: eucl_of_list_inner ) apply (subst nth_list_update) apply (auto simp: ) using in_Basis_index_Basis_list apply force using inner_not_same_Basis nth_Basis_list_in_Basis apply fastforce apply (auto simp: nth_list_update) done subgoal for i apply (auto intro!: pdevs_eqI simp: project_coord_pdevs.rep_eq) apply (auto intro!: euclidean_eqI[where 'a='a]) apply (subst project_coord_inner_Basis) apply (auto simp: eucl_of_list_inner ) apply (subst nth_list_update) apply (auto simp: ) using inner_not_same_Basis nth_Basis_list_in_Basis apply fastforce apply (auto simp: nth_list_update) done done definition inter_aform_plane where "inter_aform_plane prec Xs (sctn::'a sctn) = do { cxs \ inter_aform_plane_ortho_nres (prec) (Xs) (normal sctn) (pstn sctn); case cxs of Some cxs \ (if normal sctn \ set Basis_list then do { ASSERT (index Basis_list (normal sctn) < DIM('a::executable_euclidean_space)); RETURN ((project_ncoord_aform cxs (index Basis_list (normal sctn)) (pstn sctn))) } else if - normal sctn \ set Basis_list then do { ASSERT (index Basis_list (-normal sctn) < DIM('a)); RETURN ((project_ncoord_aform cxs (index Basis_list (-normal sctn)) (- pstn sctn))) } else SUCCEED) | None \ SUCCEED }" lemma [autoref_rules]: assumes [THEN GEN_OP_D, param]: "GEN_OP (=) (=) (A \ A \ bool_rel)" shows "(index, index) \ \A\list_rel \ A \ nat_rel" unfolding index_def find_index_def by parametricity schematic_goal inter_aform_plane_lv: fixes Z::"'a::executable_euclidean_space aform" assumes [autoref_rules_raw]: "DIM_precond TYPE('a) D" assumes [autoref_rules]: "(preci, prec) \ nat_rel" "(Zi, Z) \ lv_aforms_rel" "(sctni, sctn) \ \lv_rel\sctn_rel" notes [autoref_rules] = those_param param_map shows "(nres_of (?f::?'a dres), inter_aform_plane prec Z sctn) \ ?R" unfolding autoref_tag_defs unfolding inter_aform_plane_def including art by (autoref_monadic ) concrete_definition inter_aform_plane_lv for preci Zi sctni uses inter_aform_plane_lv lemmas [autoref_rules] = inter_aform_plane_lv.refine[autoref_higher_order_rule(1)] end lemma Basis_not_uminus: fixes i::"'a::euclidean_space" shows "i \ Basis \ - i \ Basis" by (metis inner_Basis inner_minus_left one_neq_neg_one zero_neq_neg_one) lemma pdevs_apply_project_coord_pdevs: assumes "normal sctn \ Basis \ - normal sctn \ Basis" assumes "i \ Basis" shows "pdevs_apply (project_coord_pdevs sctn cxs) x \ i = (if i = abs (normal sctn) then 0 else pdevs_apply cxs x \ i)" using assms[unfolded abs_in_Basis_iff[symmetric]] apply (transfer fixing: sctn) apply (auto simp: project_coord_inner Basis_not_uminus) using Basis_nonneg apply force using Basis_nonneg assms(1) by force lemma pdevs_domain_project_coord_pdevs_subset: "pdevs_domain (project_coord_pdevs sctn X) \ pdevs_domain X" apply (auto simp: ) apply (auto simp: euclidean_eq_iff[where 'a='a] ) by (metis add.inverse_neutral project_coord_inner_Basis project_coord_pdevs.rep_eq) lemma pdevs_val_project_coord_pdevs_inner_Basis: assumes "b \ Basis" shows "pdevs_val e (project_coord_pdevs sctn X) \ b = (if b = abs (normal sctn) then 0 else pdevs_val e X \ b)" using assms apply (auto simp: ) apply (simp add: pdevs_val_pdevs_domain inner_sum_left ) apply (subst sum.cong[OF refl]) apply (subst pdevs_apply_project_coord_pdevs) apply (simp add: abs_in_Basis_iff) apply simp apply (rule refl) apply auto unfolding pdevs_val_pdevs_domain inner_sum_left apply (rule sum.mono_neutral_cong_left) apply simp apply (rule pdevs_domain_project_coord_pdevs_subset) apply auto apply (metis Basis_nonneg abs_minus_cancel abs_of_nonneg euclidean_all_zero_iff project_coord_inner_Basis project_coord_pdevs.rep_eq) apply (metis Basis_nonneg abs_minus_cancel abs_of_nonneg project_coord_inner_Basis project_coord_pdevs.rep_eq) done lemma inner_in_Sum: "b \ Basis \ x \ b = (\i\Basis. x \ i * (i \ b))" apply (subst euclidean_representation[of x, symmetric]) unfolding inner_sum_left by (auto simp: intro!: ) lemma aform_val_project_coord_aform: "aform_val e (project_coord_aform sctn X) = project_coord (aform_val e X) (normal sctn) (pstn sctn)" apply (auto simp: aform_val_def project_coord_def project_coord_aform_def) apply (rule euclidean_eqI) unfolding inner_add_left inner_sum_left apply (subst pdevs_val_project_coord_pdevs_inner_Basis) apply (auto simp: ) apply (rule sum.cong) apply auto apply (metis abs_in_Basis_iff abs_inner abs_inner_Basis abs_zero inner_commute) apply (subst inner_in_Sum[where x="pdevs_val e (snd X)"], force) unfolding sum.distrib[symmetric] apply (rule sum.cong) apply auto apply (metis Basis_nonneg abs_inner_Basis abs_of_nonneg abs_of_nonpos inner_commute neg_0_le_iff_le) apply (metis Basis_nonneg abs_inner_Basis abs_of_nonneg abs_of_nonpos neg_0_le_iff_le) apply (auto simp: inner_Basis) done lemma project_coord_on_plane: assumes "x \ plane_of (Sctn n c)" shows "project_coord x n c = x" using assms by (auto simp: project_coord_def plane_of_def intro!: euclidean_eqI[where 'a='a]) lemma mem_Affine_project_coord_aformI: assumes "x \ Affine X" assumes "x \ plane_of sctn" shows "x \ Affine (project_coord_aform sctn X)" proof - have "x = project_coord x (normal sctn) (pstn sctn)" using assms by (intro project_coord_on_plane[symmetric]) auto also from assms obtain e where "e \ funcset UNIV {-1 .. 1}" "x = aform_val e X" by (auto simp: Affine_def valuate_def) note this(2) also have "project_coord (aform_val e X) (normal sctn) (pstn sctn) = aform_val e (project_coord_aform sctn X)" by (simp add: aform_val_project_coord_aform) finally have "x = aform_val e (project_coord_aform sctn X)" unfolding \x = aform_val e X\ . then show ?thesis using \e \ _\ by (auto simp: Affine_def valuate_def) qed lemma mem_Affine_project_coord_aformD: assumes "x \ Affine (project_coord_aform sctn X)" assumes "abs (normal sctn) \ Basis" shows "x \ plane_of sctn" using assms by (auto simp: Affine_def valuate_def aform_val_project_coord_aform plane_of_def project_coord_inner) definition "reduce_aform prec t X = summarize_aforms (prec) (collect_threshold (prec) 0 t) (degree_aforms X) X" definition "correct_girard p m N (X::real aform list) = (let Xs = map snd X; M = pdevs_mapping Xs; D = domain_pdevs Xs; diff = m - card D in if 0 < diff then (\_ _. True) else let Ds = sorted_list_of_set D; ortho_indices = map fst (take (diff + N) (sort_key (\(i, r). r) (map (\i. let xs = M i in (i, sum_list' p (map abs xs) - fold max (map abs xs) 0)) Ds))); _ = () in (\i (xs::real list). i \ set ortho_indices))" definition "reduce_aforms prec C X = summarize_aforms (prec) C (degree_aforms X) X" definition "pdevs_of_real x = (x, zero_pdevs)" definition aform_inf_inner where "aform_inf_inner prec X n = (case inner_aforms' (prec) X (map pdevs_of_real n) of Some Xn \ Inf_aform' (prec) (hd Xn))" definition aform_sup_inner where "aform_sup_inner prec X n = (case inner_aforms' (prec) X (map pdevs_of_real n) of Some Xn \ Sup_aform' (prec) (hd Xn))" text \cannot fail\ lemma approx_un_ne_None: "approx_un p (\ivl. Some (f ivl)) (Some r) \ None" by (auto simp: approx_un_def split_beta') lemma approx_un_eq_Some: "approx_un p (\ivl. Some (f ivl)) (Some r) = Some s \ s = ivl_err (real_interval (f (ivl_of_aform_err p r)))" using approx_un_ne_None by (auto simp: approx_un_def split_beta') lemma is_float_uminus: "is_float aa \ is_float (- aa)" by (auto simp: is_float_def) lemma is_float_uminus_iff[simp]: "is_float (- aa) = is_float aa" using is_float_uminus[of aa] is_float_uminus[of "-aa"] by (auto simp: is_float_def) lemma is_float_simps[simp]: "is_float 0" "is_float 1" "is_float (real_of_float x)" "is_float (truncate_down p X)" "is_float (truncate_up p X)" "is_float (eucl_truncate_down p X)" "is_float (eucl_truncate_up p X)" by (auto simp: is_float_def eucl_truncate_down_def eucl_truncate_up_def truncate_down_def truncate_up_def) lemma is_float_sum_list'[simp]: "is_float (sum_list' p xs)" by (induction xs) (auto simp: is_float_def) lemma is_float_inverse_aform': "is_float (fst (fst (inverse_aform' p X)))" unfolding inverse_aform'_def by (simp add: Let_def trunc_bound_eucl_def) lemma is_float_min_range: "min_range_antimono p F D E X Y = Some ((a, b), c) \ is_float a" "min_range_antimono p F D E X Y = Some ((a, b), c) \ is_float c" "min_range_mono p F D E X Y = Some ((a, b), c) \ is_float a" "min_range_mono p F D E X Y = Some ((a, b), c) \ is_float c" by (auto simp: min_range_antimono_def min_range_mono_def Let_def bind_eq_Some_conv prod_eq_iff trunc_bound_eucl_def mid_err_def affine_unop_def split: prod.splits) lemma is_float_ivl_err: assumes "ivl_err x = ((a, b), c)" "is_float (lower x)" "is_float (upper x)" shows "is_float a" "is_float c" proof - define x1 where "x1 = lower x" define x2 where "x2 = upper x" have "a = (x1 + x2) / 2" "c = (x2 - x1) / 2" using assms by (auto simp: ivl_err_def x1_def x2_def) moreover have "(x1 + x2) / 2 \ float" "(x2 - x1) / 2 \ float" using assms by (auto simp: is_float_def x1_def x2_def) ultimately show "is_float a" "is_float c" unfolding is_float_def by blast+ qed lemma is_float_trunc_bound_eucl[simp]: "is_float (fst (trunc_bound_eucl p X))" by (auto simp: trunc_bound_eucl_def Let_def) lemma add_eq_times_2_in_float: "a + b = c * 2 \ is_float a \ is_float b \ is_float c" proof goal_cases case 1 then have "c = (a + b) / 2" by simp also have "\ \ float" using 1(3,2) by (simp add: is_float_def) finally show ?case by (simp add: is_float_def) qed lemma approx_floatarith_Some_is_float: "approx_floatarith p fa XS = Some ((a, b), ba) \ list_all (\((a, b), c). is_float a \ is_float c) XS \ is_float a \ is_float ba" apply (induction fa arbitrary: a b ba) subgoal by (auto simp: bind_eq_Some_conv add_aform'_def Let_def ) subgoal by (auto simp: bind_eq_Some_conv uminus_aform_def Let_def) subgoal by (auto simp: bind_eq_Some_conv mult_aform'_def Let_def ) subgoal by (auto simp: bind_eq_Some_conv inverse_aform_err_def map_aform_err_def prod_eq_iff is_float_inverse_aform' acc_err_def aform_to_aform_err_def aform_err_to_aform_def inverse_aform_def Let_def split: if_splits) subgoal by (auto simp: bind_eq_Some_conv cos_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv arctan_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal apply (auto simp: bind_eq_Some_conv uminus_aform_def Let_def is_float_min_range is_float_ivl_err set_of_eq real_interval_abs split: if_splits prod.splits) apply (metis is_float_ivl_err(1) is_float_simps(3) lower_real_interval real_interval_abs upper_real_interval) by (metis is_float_ivl_err(2) is_float_simps(3) lower_real_interval real_interval_abs upper_real_interval) subgoal by (auto simp: bind_eq_Some_conv max_aform_err_def Let_def is_float_min_range is_float_ivl_err max_def split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv min_aform_err_def Let_def is_float_min_range is_float_ivl_err min_def split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv arctan_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv sqrt_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv exp_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv powr_aform_err_def approx_bin_def exp_aform_err_def Let_def is_float_min_range is_float_ivl_err mult_aform'_def split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv ln_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv power_aform_err_def Let_def is_float_min_range is_float_ivl_err mid_err_def dest!: add_eq_times_2_in_float split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv cos_aform_err_def Let_def is_float_min_range approx_un_def is_float_ivl_err split: if_splits) subgoal by (auto simp: bind_eq_Some_conv cos_aform_err_def Let_def is_float_min_range list_all_length split: if_splits) subgoal by (auto simp: bind_eq_Some_conv num_aform_def Let_def) done lemma plain_floatarith_not_None: assumes "plain_floatarith N fa" "N \ length XS" "list_all (\((a, b), c). is_float a \ is_float c) XS" shows "approx_floatarith p fa XS \ None" using assms by (induction fa) (auto simp: Let_def split_beta' approx_un_eq_Some prod_eq_iff approx_floatarith_Some_is_float) lemma plain_floatarith_slp_not_None: assumes "\fa. fa \ set fas \ plain_floatarith N fa" "N \ length XS" "list_all (\((a, b), c). is_float a \ is_float c) XS" shows "approx_slp p fas XS \ None" using assms apply (induction fas arbitrary: XS) subgoal by simp subgoal for fa fas XS using plain_floatarith_not_None[of N fa XS p] by (auto simp: Let_def split_beta' approx_un_eq_Some prod_eq_iff bind_eq_Some_conv approx_floatarith_Some_is_float) done lemma plain_floatarithE: assumes "plain_floatarith N fa" "N \ length XS" "list_all (\((a, b), c). is_float a \ is_float c) XS" obtains R where "approx_floatarith p fa XS = Some R" using plain_floatarith_not_None[OF assms] by force lemma approx_slp_outer_eq_None_iff: "approx_slp_outer p a b XS = None \ approx_slp p b ((map (\x. (x, 0)) XS)) = None" by (auto simp: approx_slp_outer_def Let_def bind_eq_None_conv) lemma approx_slp_sing_eq_None_iff: "approx_slp p [b] xs = None \ approx_floatarith p b xs = None" by (auto simp: approx_slp_outer_def Let_def bind_eq_None_conv) lemma plain_inner_floatariths: "plain_floatarith N (inner_floatariths xs ys)" if "list_all (plain_floatarith N) xs" "list_all (plain_floatarith N) ys" using that by (induction xs ys rule: inner_floatariths.induct) auto lemma aiN: "approx_floatarith p (inner_floatariths xs ys) zs \ None" if "\x. x \ set xs \ approx_floatarith p x zs \ None" and "\x. x \ set ys \ approx_floatarith p x zs \ None" using that apply (induction xs ys rule: inner_floatariths.induct) apply (auto simp: Let_def bind_eq_Some_conv) by (metis old.prod.exhaust) lemma aiVN: "approx_floatarith p (inner_floatariths (map floatarith.Var [0..x. (x, 0)) a @ map (\x. (x, 0)) b) \ None" by (rule aiN) (auto simp: nth_append) lemma iaN: "inner_aforms' p a b \ None" unfolding inner_aforms'_def Let_def approx_slp_outer_def using aiVN[of p a b] by (auto simp: Let_def bind_eq_Some_conv) definition "support_aform prec b X = (let ia = inner_aform X b in fst X \ b + tdev' (prec) (snd ia))" definition "width_aforms prec X = (let t = tdev' (prec) ((abssum_of_pdevs_list (map snd X))) in t)" definition "inf_aforms prec xs = map (Inf_aform' (prec)) xs" definition "sup_aforms prec xs = map (Sup_aform' (prec)) xs" definition "fresh_index_aforms xs = Max (insert 0 (degree_aform ` set xs))" definition "independent_aforms x env = (msum_aform (fresh_index_aforms env) (0, zero_pdevs) x#env)" definition "aform_form_ivl prec form xs = dRETURN (approx_form prec form (ivls_of_aforms prec xs) [])" definition "aform_form prec form xs = dRETURN (approx_form_aform prec form xs)" definition "aform_slp prec n slp xs = dRETURN (approx_slp_outer prec n slp xs)" definition "aform_isFDERIV prec n ns fas xs = dRETURN (isFDERIV_approx prec n ns fas (ivls_of_aforms prec xs))" definition aform_rel_internal: "aform_rel R = br Affine top O \R\set_rel" lemma aform_rel_def: "\rnv_rel\aform_rel = br Affine top" unfolding relAPP_def by (auto simp: aform_rel_internal) definition "aforms_rel = br Joints top" definition aform_rell :: "((real \ real pdevs) list \ real list set) set" where "aform_rell = aforms_rel" definition aforms_relp_internal: "aforms_relp R = aforms_rel O \R\set_rel" lemma aforms_relp_def: "\R\aforms_relp = aforms_rel O \R\set_rel" by (auto simp: aforms_relp_internal relAPP_def) definition "product_aforms x y = x @ msum_aforms (degree_aforms (x)) (replicate (length y) (0, zero_pdevs)) y" lemma eucl_of_list_mem_eucl_of_list_aform: assumes "x \ Joints a" assumes "length a = DIM('a::executable_euclidean_space)" shows "eucl_of_list x \ Affine (eucl_of_list_aform a::'a aform)" using assms by (auto simp: Joints_def Affine_def valuate_def eucl_of_list_aform_def aform_val_def pdevs_val_sum_pdevs inner_simps eucl_of_list_inner intro!: euclidean_eqI[where 'a='a]) lemma in_image_eucl_of_list_eucl_of_list_aform: assumes "length x = DIM('a::executable_euclidean_space)" "xa \ Affine (eucl_of_list_aform x::'a aform)" shows "xa \ eucl_of_list ` Joints x" using assms by (auto intro!: nth_equalityI image_eqI[where x="list_of_eucl xa"] simp: aform_val_def eucl_of_list_aform_def Affine_def valuate_def Joints_def inner_simps pdevs_val_sum_pdevs eucl_of_list_inner) lemma eucl_of_list_image_Joints: assumes "length x = DIM('a::executable_euclidean_space)" shows "eucl_of_list ` Joints x = Affine (eucl_of_list_aform x::'a aform)" using assms by (auto intro!: eucl_of_list_mem_eucl_of_list_aform in_image_eucl_of_list_eucl_of_list_aform) definition "aform_ops = \ appr_of_ivl = aforms_of_ivls, product_appr = product_aforms, msum_appr = msum_aforms', inf_of_appr = \optns. inf_aforms (precision optns), sup_of_appr = \optns. sup_aforms (precision optns), split_appr = split_aforms_largest_uncond_take, appr_inf_inner = \optns. aform_inf_inner (precision optns), appr_sup_inner = \optns. aform_sup_inner (precision optns), inter_appr_plane = \optns xs sctn. inter_aform_plane_lv (length xs) (precision optns) xs sctn, reduce_appr = \optns. reduce_aforms (precision optns), width_appr = \optns. width_aforms (precision optns), approx_slp_dres = \optns. aform_slp (precision optns), approx_euclarithform = \optns. aform_form (precision optns), approx_isFDERIV = \optns. aform_isFDERIV (precision optns) \" lemma Affine_eq_permI: assumes "fst X = fst Y" assumes "map snd (list_of_pdevs (snd X)) <~~> map snd (list_of_pdevs (snd Y))" (is "?perm X Y") shows "Affine X = Affine Y" proof - { fix X Y and e::"nat \ real" assume perm: "?perm X Y" and e: "e \ funcset UNIV {- 1..1}" from pdevs_val_of_list_of_pdevs2[OF e, of "snd X"] obtain e' where e': "pdevs_val e (snd X) = pdevs_val e' (pdevs_of_list (map snd (list_of_pdevs (snd X))))" "e' \ funcset UNIV {- 1..1}" by auto note e'(1) also from pdevs_val_perm[OF perm e'(2)] obtain e'' where e'': "e'' \ funcset UNIV {- 1..1}" "pdevs_val e' (pdevs_of_list (map snd (list_of_pdevs (snd X)))) = pdevs_val e'' (pdevs_of_list (map snd (list_of_pdevs (snd Y))))" by auto note e''(2) also from pdevs_val_of_list_of_pdevs[OF e''(1), of "snd Y", simplified] obtain e''' where e''': "pdevs_val e'' (pdevs_of_list (map snd (list_of_pdevs (snd Y)))) = pdevs_val e''' (snd Y)" "e''' \ funcset UNIV {- 1..1}" by auto note e'''(1) finally have "\e' \ funcset UNIV {-1 .. 1}. pdevs_val e (snd X) = pdevs_val e' (snd Y)" using e'''(2) by auto } note E = this note e1 = E[OF assms(2)] and e2 = E[OF perm_sym[OF assms(2)]] show ?thesis by (auto simp: Affine_def valuate_def aform_val_def assms dest: e1 e2) qed context includes autoref_syntax begin lemma aform_of_ivl_refine: "x \ y \ (aform_of_ivl x y, atLeastAtMost x y) \ \rnv_rel\aform_rel" by (auto simp: aform_rel_def br_def Affine_aform_of_ivl) lemma aforms_of_ivl_leI1: fixes en::real assumes "-1 \ en" "xsn \ ysn" shows "xsn \ (xsn + ysn) / 2 + (ysn - xsn) * en / 2" proof - have "xsn * (1 + en) \ ysn * (1 + en)" using assms mult_right_mono by force then show ?thesis by (auto simp: algebra_simps divide_simps) qed lemma aforms_of_ivl_leI2: fixes en::real assumes "1 \ en" "xsn \ ysn" shows "(xsn + ysn) / 2 + (ysn - xsn) * en / 2 \ ysn" proof - have "xsn * (1 - en) \ ysn * (1 - en)" using assms mult_right_mono by force then show ?thesis by (auto simp: algebra_simps divide_simps) qed lemma Joints_aforms_of_ivlsD1: "zs \ Joints (aforms_of_ivls xs ys) \ list_all2 (\) xs ys \ list_all2 (\) xs zs" by (auto simp: Joints_def valuate_def aforms_of_ivls_def aform_val_def Pi_iff list_all2_conv_all_nth intro!: list_all2_all_nthI aforms_of_ivl_leI1) lemma Joints_aforms_of_ivlsD2: "zs \ Joints (aforms_of_ivls xs ys) \ list_all2 (\) xs ys \ list_all2 (\) zs ys" by (auto simp: Joints_def valuate_def aforms_of_ivls_def aform_val_def Pi_iff list_all2_conv_all_nth intro!: list_all2_all_nthI aforms_of_ivl_leI2) lemma aforms_of_ivls_refine: "list_all2 (\) xrs yrs \ (xri, xrs) \ \rnv_rel\list_rel \ (yri, yrs) \ \rnv_rel\list_rel \ (aforms_of_ivls xri yri, lv_ivl xrs yrs) \ aforms_rel" apply (auto simp: aforms_rel_def br_def list_all2_lengthD lv_ivl_def Joints_aforms_of_ivlsD1 Joints_aforms_of_ivlsD2 intro!: aforms_of_ivls) subgoal by (simp add: list_all2_conv_all_nth) subgoal by (simp add: list_all2_conv_all_nth) done lemma degrees_zero_pdevs[simp]: "degrees (replicate n zero_pdevs) = 0" by (auto simp: degrees_def intro!: Max_eqI) lemma Joints_product_aforms: "Joints (product_aforms a b) = product_listset (Joints a) (Joints b)" apply (auto simp: Joints_def valuate_def product_listset_def product_aforms_def aform_vals_def[symmetric] aform_val_msum_aforms) subgoal for e apply (rule image_eqI[where x="(aform_vals e a, List.map2 (+) (aform_vals e (replicate (length b) (0, zero_pdevs))) (aform_vals (\i. e (i + degree_aforms a)) b))"]) apply (auto simp: split_beta') apply (auto simp: aform_vals_def intro!: nth_equalityI image_eqI[where x="\i. e (i + degree_aforms a)"]) done subgoal for e1 e2 apply (rule image_eqI[where x="\i. if i < degree_aforms a then e1 i else e2 (i - degree_aforms a)"]) apply (auto simp: aform_vals_def aform_val_def Pi_iff intro!: nth_equalityI pdevs_val_degree_cong) subgoal premises prems for x xs k proof - from prems have "degree xs \ degree_aforms a" by (auto simp: degrees_def Max_gr_iff) then show ?thesis using prems by auto qed done done lemma product_aforms_refine: "(product_aforms, product_listset) \ aforms_rel \ aforms_rel \ aforms_rel" by (auto simp: aforms_rel_def br_def Joints_product_aforms) lemma mem_lv_rel_set_rel_iff: fixes z::"'a::executable_euclidean_space set" shows "(y, z) \ \lv_rel\set_rel \ (z = eucl_of_list ` y \ (\x \ y. length x = DIM('a)))" unfolding lv_rel_def by (auto simp: set_rel_def br_def) lemma eucl_of_list_mem_lv_rel: "length x = DIM('a::executable_euclidean_space) \ (x, eucl_of_list x::'a) \ lv_rel" unfolding lv_rel_def by (auto simp: br_def) lemma mem_Joints_msum_aforms'I: "a \ Joints x \ b \ Joints y \ List.map2 (+) a b \ Joints (msum_aforms' x y)" by (auto simp: Joints_msum_aforms degrees_def) lemma mem_Joints_msum_aforms'E: assumes "xa \ Joints (msum_aforms' x y)" obtains a b where "xa = List.map2 (+) a b" "a \ Joints x" "b \ Joints y" using assms by (auto simp: Joints_msum_aforms degrees_def) lemma msum_aforms'_refine_raw: shows "(msum_aforms' x y, {List.map2 (+) a b|a b. a \ Joints x \ b \ Joints y}) \ aforms_rel" unfolding aforms_rel_def br_def by (safe elim!: mem_Joints_msum_aforms'E intro!: mem_Joints_msum_aforms'I) (auto simp: Joints_imp_length_eq) lemma aforms_relD: "(a, b) \ aforms_rel \ b = Joints a" by (auto simp: aforms_rel_def br_def) lemma msum_aforms'_refine: "(msum_aforms', \xs ys. {List.map2 (+) x y |x y. x \ xs \ y \ ys}) \ aforms_rel \ aforms_rel \ aforms_rel" by (safe dest!: aforms_relD intro!: msum_aforms'_refine_raw) lemma length_inf_aforms[simp]: "length (inf_aforms optns x) = length x" and length_sup_aforms[simp]: "length (sup_aforms optns x) = length x" by (auto simp: inf_aforms_def sup_aforms_def) lemma inf_aforms_refine: "(xi, x) \ aforms_rel \ length xi = d \ (RETURN (inf_aforms optns xi), Inf_specs d x) \ \rl_rel\nres_rel" unfolding o_def apply (auto simp: br_def aforms_rel_def mem_lv_rel_set_rel_iff nres_rel_def Inf_specs_def intro!: RETURN_SPEC_refine) unfolding lv_rel_def by (auto simp: aforms_rel_def br_def Joints_imp_length_eq eucl_of_list_inner inf_aforms_def nth_in_AffineI intro!: Inf_aform'_Affine_le list_all2_all_nthI) lemma sup_aforms_refine: "(xi, x) \ aforms_rel \ length xi = d \ (RETURN (sup_aforms optns xi), Sup_specs d x) \ \rl_rel\nres_rel" unfolding o_def apply (auto simp: aforms_relp_def mem_lv_rel_set_rel_iff nres_rel_def Sup_specs_def intro!: RETURN_SPEC_refine) unfolding lv_rel_def by (auto simp: aforms_rel_def br_def Joints_imp_length_eq eucl_of_list_inner sup_aforms_def nth_in_AffineI intro!: Sup_aform'_Affine_ge list_all2_all_nthI) lemma nres_of_THE_DRES_le: assumes "\v. x = Some v \ RETURN v \ y" shows "nres_of (THE_DRES x) \ y" using assms by (auto simp: THE_DRES_def split: option.split) lemma degree_le_fresh_index: "a \ set A \ degree_aform a \ fresh_index_aforms A" by (auto simp: fresh_index_aforms_def intro!: Max_ge) lemma zero_in_JointsI: "xs \ Joints XS \ z = (0, zero_pdevs) \ 0 # xs \ Joints (z # XS)" by (auto simp: Joints_def valuate_def) lemma cancel_nonneg_pos_add_multI: "0 \ c + c * x" if "c \ 0" "1 + x \ 0" for c x::real proof - have "0 \ c + c * x \ 0 \ c * (1 + x)" by (auto simp: algebra_simps) also have "\ \ 0 \ 1 + x" using that by (auto simp: zero_le_mult_iff) finally show ?thesis using that by simp qed lemma Joints_map_split_aform: fixes x::"real aform list" shows "Joints x \ Joints (map (\a. fst (split_aform a i)) x) \ Joints (map (\b. snd (split_aform b i)) x)" unfolding subset_iff apply (intro allI impI) proof goal_cases case (1 xs) then obtain e where e: "e \ funcset UNIV {-1 .. 1}" "xs = map (aform_val e) x" by (auto simp: Joints_def valuate_def) consider "e i \ 0" | "e i \ 0" by arith then show ?case proof cases case 1 let ?e = "e(i:= 2 * e i - 1)" note e(2) also have "map (aform_val e) x = map (aform_val ?e) (map (\b. snd (split_aform b i)) x)" by (auto simp: aform_val_def split_aform_def Let_def divide_simps algebra_simps) also have "\ \ Joints (map (\b. snd (split_aform b i)) x)" using e \0 \ e i\ by (auto simp: Joints_def valuate_def Pi_iff intro!: image_eqI[where x = ?e]) finally show ?thesis by simp next case 2 let ?e = "e(i:= 2 * e i + 1)" note e(2) also have "map (aform_val e) x = map (aform_val ?e) (map (\b. fst (split_aform b i)) x)" by (auto simp: aform_val_def split_aform_def Let_def divide_simps algebra_simps) also have "\ \ Joints (map (\b. fst (split_aform b i)) x)" using e \0 \ e i\ by (auto simp: Joints_def valuate_def Pi_iff real_0_le_add_iff intro!: image_eqI[where x = ?e] cancel_nonneg_pos_add_multI) finally show ?thesis by simp qed qed lemma split_aforms_lemma: fixes x::"real aform list" assumes "(x, y) \ aforms_rel" assumes "z \ y" assumes "l = (length x)" shows "\a b. (split_aforms x i, a, b) \ aforms_rel \\<^sub>r aforms_rel \ env_len a l \ env_len b l \ z \ a \ b" using assms apply (auto simp: split_aforms_def o_def) apply (rule exI[where x="Joints (map (\x. fst (split_aform x i)) x)"]) apply auto subgoal by (auto intro!: brI simp: aforms_rel_def) apply (rule exI[where x="Joints (map (\x. snd (split_aform x i)) x)"]) apply (rule conjI) subgoal by (auto intro!: brI simp: aforms_rel_def) subgoal using Joints_map_split_aform[of x i] by (auto simp: br_def aforms_rel_def env_len_def Joints_imp_length_eq) done lemma length_summarize_pdevs_list[simp]: "length (summarize_pdevs_list a b c d) = length d" by (auto simp: summarize_pdevs_list_def) lemma length_summarize_aforms[simp]: "length (summarize_aforms a b e d) = length d" by (auto simp: summarize_aforms_def Let_def) lemma split_aform_largest_take_refine: "(ni, n) \ nat_rel \ (xi::real aform list, x) \ aforms_rel \ length xi = d \ (RETURN (split_aforms_largest_uncond_take ni xi), split_spec_params d n x) \ \aforms_rel \\<^sub>r aforms_rel\nres_rel" apply (auto simp: split_spec_params_def nres_rel_def aforms_relp_def mem_lv_rel_set_rel_iff split_aforms_largest_uncond_take_def Let_def comps split: prod.splits intro!: RETURN_SPEC_refine) apply (rule split_aforms_lemma) by (auto simp add: aforms_rel_def) lemma aform_val_pdevs_of_real[simp]: "aform_val e (pdevs_of_real x) = x" by (auto simp: pdevs_of_real_def) lemma degree_aform_pdevs_of_real[simp]: "degree_aform (pdevs_of_real x) = 0" by (auto simp: pdevs_of_real_def) lemma interpret_floatarith_inner_eq: shows "interpret_floatarith (inner_floatariths xs ys) vs = (\(x, y) \ (zip xs ys). (interpret_floatarith x vs) * (interpret_floatarith y vs))" by (induction xs ys rule: inner_floatariths.induct) auto lemma approx_slp_outer_sing: "approx_slp_outer p (Suc 0) [fa] XS = Some R \ (\Y. approx_floatarith p fa (map (\x. (x, 0)) XS) = Some Y \ [aform_err_to_aform Y (max (degree_aforms XS) (degree_aform_err Y))] = R)" by (auto simp: approx_slp_outer_def bind_eq_Some_conv degrees_def) lemma aforms_err_aform_valsI: assumes "vs = aform_vals e XS" shows "vs \ aforms_err e (map (\x. (x, 0)) (XS))" by (auto simp: assms aforms_err_def o_def aform_err_def aform_vals_def) lemma aform_val_degree_cong: "b = d \ (\i. i < degree_aform d \ a i = c i) \ aform_val a b = aform_val c d" by (auto simp: aform_val_def intro!: pdevs_val_degree_cong) lemma mem_degree_aformD: "x \ set XS \ degree_aform x \ degree_aforms XS" by (auto simp: degrees_def) lemma degrees_append_leD1: "(degrees xs) \ degrees (xs @ ys)" unfolding degrees_def by (rule Max_mono) (auto simp: degrees_def min_def max_def Max_ge_iff image_Un Max_gr_iff) lemma inner_aforms': assumes "xs \ Joints XS" assumes "inner_aforms' p XS (map pdevs_of_real rs) = Some R" shows "(\(x, y) \ (zip xs rs). x * y) \ Affine (hd R)" (is ?th1) "length R = 1" (is ?th2) proof - from assms obtain e where "e \ funcset UNIV {-1 .. 1}" "xs = aform_vals e XS" by (auto simp: Joints_def valuate_def aform_vals_def) then have e: "xs @ rs = aform_vals e (XS @ map pdevs_of_real rs)" "xs @ rs \ Joints (XS @ map pdevs_of_real rs)" "length xs = length XS" "e \ funcset UNIV {- 1..1}" by (auto simp: aform_vals_def o_def degrees_def Joints_def valuate_def) have "approx_slp_outer p (Suc 0) ([inner_floatariths (map floatarith.Var [0..x. (x, 0)) (XS @ map pdevs_of_real rs)) = Some Y" "R = [aform_err_to_aform Y (max (degree_aforms (XS @ map pdevs_of_real rs)) (degree_aform_err Y))]" unfolding approx_slp_outer_sing by auto let ?m = "(max (degree_aforms (XS @ map pdevs_of_real rs)) (degree_aform_err Y))" from approx_floatarith_Elem[OF Y(1) e(4) aforms_err_aform_valsI[OF e(1)]] have "interpret_floatarith (inner_floatariths (map floatarith.Var [0.. aform_err e Y" "degree_aform_err Y \ max (degree_aforms (XS @ map pdevs_of_real rs)) (degree_aform_err Y)" by auto from aform_err_to_aformE[OF this] obtain err where err: "interpret_floatarith (inner_floatariths (map floatarith.Var [0.. err" "err \ 1" by auto let ?e' = "(e(max (degrees (map snd XS @ map (snd \ pdevs_of_real) rs)) (degree_aform_err Y) := err))" from e(1) have e': "xs @ rs = aform_vals ?e' (XS @ map pdevs_of_real rs)" apply (auto simp: aform_vals_def intro!: aform_val_degree_cong) apply (frule mem_degree_aformD) apply (frule le_less_trans[OF degrees_append_leD1]) apply (auto simp: o_def) done from err have "interpret_floatarith (inner_floatariths (map floatarith.Var [0.. Joints (R @ XS @ map pdevs_of_real rs)" using e(1,3,4) e' apply (auto simp: valuate_def Joints_def intro!: nth_equalityI image_eqI[where x="?e'"]) apply (simp add: Y aform_vals_def; fail) apply (simp add: Y o_def) apply (simp add: nth_append nth_Cons) apply (auto split: nat.splits simp: nth_append nth_Cons aform_vals_def) done then have "(\(x, y)\zip (map floatarith.Var [0.. Joints (R @ XS @ map pdevs_of_real rs)" apply (subst (asm)interpret_floatarith_inner_eq ) apply (auto simp: ) done also have "(\(x, y)\zip (map floatarith.Var [0..(x, y)\zip xs rs. x * y)" by (auto simp: sum_list_sum_nth assms e(3) nth_append intro!: sum.cong) finally show ?th1 ?th2 by (auto simp: Affine_def valuate_def Joints_def Y) qed lemma inner_aforms'_inner_lv_rel: "(a, a') \ aforms_rel \ inner_aforms' prec a (map pdevs_of_real a'a) = Some R \ x \ a' \ inner_lv_rel x a'a \ Affine (hd R)" unfolding mem_lv_rel_set_rel_iff unfolding lv_rel_def aforms_rel_def apply (auto simp: br_def) apply (subst arg_cong2[where f="(\)", OF _ refl]) defer apply (rule inner_aforms') apply (auto simp: br_def Joints_imp_length_eq inner_lv_rel_def) done lemma aform_inf_inner_refine: "(RETURN o2 aform_inf_inner optns, Inf_inners) \ aforms_rel \ rl_rel \ \rnv_rel\nres_rel" by (auto simp: aforms_relp_def nres_rel_def Inf_inners_def aform_inf_inner_def[abs_def] comp2_def iaN intro!: Inf_aform'_Affine_le inner_aforms'_inner_lv_rel split: option.splits list.splits) lemma aform_sup_inner_refine: "(RETURN o2 aform_sup_inner optns, Sup_inners) \ aforms_rel \ rl_rel \ \rnv_rel\nres_rel" by (auto simp: aforms_relp_def nres_rel_def Sup_inners_def aform_sup_inner_def[abs_def] comp2_def iaN intro!: Sup_aform'_Affine_ge inner_aforms'_inner_lv_rel split: option.splits list.splits) lemma lv_aforms_rel_comp_br_Affine_le: "lv_aforms_rel O br Affine top \ \lv_rel\aforms_relp" apply (auto simp: lv_aforms_rel_def aforms_relp_def br_def) apply (rule relcompI) apply (auto simp: aforms_rel_def intro!: brI ) by (auto simp: mem_lv_rel_set_rel_iff Joints_imp_length_eq intro!: eucl_of_list_mem_eucl_of_list_aform intro!: in_image_eucl_of_list_eucl_of_list_aform) lemma bijective_lv_rel[relator_props]: "bijective lv_rel" unfolding lv_rel_def bijective_def apply (auto simp: br_def) by (metis eucl_of_list_inj) lemma sv_lv_rel_inverse[relator_props]: "single_valued (lv_rel\)" using bijective_lv_rel by (rule bijective_imp_sv) lemma list_of_eucl_image_lv_rel_inverse: "(x, list_of_eucl ` x) \ \lv_rel\\set_rel" unfolding set_rel_sv[OF sv_lv_rel_inverse] apply (auto simp: ) apply (rule ImageI) apply (rule converseI) apply (rule lv_relI) apply auto apply (rule image_eqI) prefer 2 apply assumption unfolding lv_rel_def apply (auto simp: br_def) subgoal for y apply (rule exI[where x="list_of_eucl y"]) apply auto done done lemma lv_rel_comp_lv_rel_inverse: "((lv_rel::(_\'a::executable_euclidean_space) set) O lv_rel\) = {(x, y). x = y \ length x = DIM('a)}" apply (auto simp: intro!: lv_relI) unfolding lv_rel_def by (auto simp: br_def intro!: eucl_of_list_inj) lemma inter_aform_plane_refine_aux: "d = CARD('n::enum) \ (xi, x) \ aforms_rel \ (si, s) \ \rl_rel\sctn_rel \ length xi = d \ length (normal si) = d \ (nres_of (inter_aform_plane_lv (length xi) optns xi si), inter_sctn_specs d x s) \ \aforms_rel\nres_rel" proof (goal_cases) case 1 from 1 have sis: "si = s" by (cases si; cases s) (auto simp: sctn_rel_def) have Dp: "DIM_precond TYPE('n rvec) CARD('n)" by auto have a: "(xi, eucl_of_list_aform xi::'n rvec aform) \ lv_aforms_rel" by (auto simp: lv_aforms_rel_def br_def aforms_relp_def aforms_rel_def mem_lv_rel_set_rel_iff 1 Joints_imp_length_eq) have b: "(si, (Sctn (eucl_of_list (normal si)) (pstn si))::'n rvec sctn) \ \lv_rel\sctn_rel" using 1 by (cases si) (auto simp: lv_aforms_rel_def br_def aforms_relp_def aforms_rel_def mem_lv_rel_set_rel_iff Joints_imp_length_eq sctn_rel_def intro!: lv_relI) have a: "(nres_of (inter_aform_plane_lv CARD('n) optns xi si), inter_aform_plane optns (eucl_of_list_aform xi::'n rvec aform) (Sctn (eucl_of_list (normal si)) (pstn si))) \ \lv_aforms_rel\nres_rel" (is "(_, inter_aform_plane _ ?ea ?se) \ _") using inter_aform_plane_lv.refine[OF Dp IdI a b, of optns] by simp have b: "(inter_aform_plane optns ?ea ?se, inter_sctn_spec (Affine ?ea) ?se) \ \br Affine top\nres_rel" using 1 apply (auto simp: inter_sctn_spec_def nres_rel_def inter_aform_plane_def project_ncoord_aform_def inter_aform_plane_ortho_nres_def split: option.splits intro!: RETURN_SPEC_refine dest!: inter_inter_aform_plane_ortho) apply (auto simp: aform_rel_def br_def nres_rel_def comps bind_eq_Some_conv inter_sctn_spec_def inter_aform_plane_def plane_of_def aforms_relp_def aforms_rel_def mem_lv_rel_set_rel_iff intro!: RETURN_SPEC_refine nres_of_THE_DRES_le mem_Affine_project_coord_aformI dest!: inter_inter_aform_plane_ortho split: if_splits) apply (auto dest!: mem_Affine_project_coord_aformD simp: abs_in_Basis_iff plane_of_def) done from relcompI[OF a b] have "(nres_of (inter_aform_plane_lv CARD('n) optns xi si), inter_sctn_spec (Affine ?ea) ?se) \ \lv_aforms_rel\nres_rel O \br Affine top\nres_rel" by auto also have "\lv_aforms_rel\nres_rel O \br Affine top\nres_rel \ \\lv_rel\aforms_relp\nres_rel" unfolding nres_rel_comp apply (rule nres_rel_mono) apply (rule lv_aforms_rel_comp_br_Affine_le) done finally have step1: "(nres_of (inter_aform_plane_lv CARD('n) optns xi si), inter_sctn_spec (Affine (eucl_of_list_aform xi)::'n rvec set) (Sctn (eucl_of_list (normal si)) (pstn si))) \ \\lv_rel\aforms_relp\nres_rel" by simp have step2: "(inter_sctn_spec (Affine (eucl_of_list_aform xi)::'n rvec set) (Sctn (eucl_of_list (normal si)) (pstn si)), inter_sctn_specs CARD('n) (Joints xi) si) \ \\lv_rel\\set_rel\nres_rel" apply (auto simp: inter_sctn_specs_def inter_sctn_spec_def simp: nres_rel_def intro!: SPEC_refine) subgoal for x apply (rule exI[where x="list_of_eucl ` x"]) apply (auto simp: env_len_def plane_ofs_def intro: list_of_eucl_image_lv_rel_inverse) subgoal for y apply (rule image_eqI[where x="eucl_of_list y"]) apply (subst list_of_eucl_eucl_of_list) apply (auto simp: 1 Joints_imp_length_eq) apply (rule subsetD, assumption) apply (auto simp: intro!: eucl_of_list_mem_eucl_of_list_aform 1) using inner_lv_rel_autoref[where 'a="'n rvec", param_fo, OF lv_relI lv_relI, of y "normal si"] by (auto simp: plane_of_def 1 Joints_imp_length_eq) subgoal for y apply (drule subsetD, assumption) using inner_lv_rel_autoref[where 'a="'n rvec", param_fo, OF lv_relI lv_relI, of "list_of_eucl y" "normal si"] by (auto simp: plane_of_def 1 Joints_imp_length_eq) done done from relcompI[OF step1 step2] have "(nres_of (inter_aform_plane_lv CARD('n) optns xi si), inter_sctn_specs CARD('n) (Joints xi) si) \ \\lv_rel::(real list \ 'n rvec)set\aforms_relp\nres_rel O \\lv_rel\\set_rel\nres_rel" by simp also have "\\lv_rel::(real list \ 'n rvec)set\aforms_relp\nres_rel O \\lv_rel\\set_rel\nres_rel \ \aforms_rel O Id\nres_rel" unfolding nres_rel_comp O_assoc aforms_relp_def apply (rule nres_rel_mono) apply (rule relcomp_mono[OF order_refl]) unfolding set_rel_sv[OF sv_lv_rel_inverse] set_rel_sv[OF lv_rel_sv] apply (auto simp: length_lv_rel) unfolding relcomp_Image[symmetric] lv_rel_comp_lv_rel_inverse apply (auto simp: Basis_not_uminus length_lv_rel) unfolding lv_rel_def subgoal for a b c d apply (auto dest!: brD simp: length_lv_rel) using eucl_of_list_inj[where 'a="'n rvec", of d b] by auto done finally show ?case using 1 by (simp add: aforms_rel_def br_def sis) qed end setup \Sign.add_const_constraint (@{const_name "enum_class.enum"}, SOME @{typ "'a list"})\ setup \Sign.add_const_constraint (@{const_name "enum_class.enum_all"}, SOME @{typ "('a \ bool) \ bool"})\ setup \Sign.add_const_constraint (@{const_name "enum_class.enum_ex"}, SOME @{typ "('a \ bool) \ bool"})\ lemmas inter_aform_plane_refine_unoverloaded0 = inter_aform_plane_refine_aux[internalize_sort "'n::enum", unoverload enum_class.enum, unoverload enum_class.enum_all, unoverload enum_class.enum_ex] theorem inter_aform_plane_refine_unoverloaded: "class.enum (enum::'a list) enum_all enum_ex \ d = CARD('a) \ (xi, x) \ aforms_rel \ (si, s) \ \rl_rel\sctn_rel \ length xi = d \ length (normal si) = d \ (nres_of (inter_aform_plane_lv (length xi) optns xi si), inter_sctn_specs d x s) \ \aforms_rel\nres_rel" by (rule inter_aform_plane_refine_unoverloaded0) auto setup \Sign.add_const_constraint (@{const_name "enum_class.enum"}, SOME @{typ "'a::enum list"})\ setup \Sign.add_const_constraint (@{const_name "enum_class.enum_all"}, SOME @{typ "('a::enum \ bool) \ bool"})\ setup \Sign.add_const_constraint (@{const_name "enum_class.enum_ex"}, SOME @{typ "('a::enum \ bool) \ bool"})\ context includes autoref_syntax begin text \TODO: this is a special case of \Cancel_Card_Constraint\ from \AFP/Perron_Frobenius\!\ lemma type_impl_card_n_enum: assumes "\(Rep :: 'a \ nat) Abs. type_definition Rep Abs {0 ..< n :: nat}" obtains enum enum_all enum_ex where "class.enum (enum::'a list) enum_all enum_ex \ n = CARD('a)" proof - from assms obtain rep :: "'a \ nat" and abs :: "nat \ 'a" where t: "type_definition rep abs {0 ..{0 .. . have "card (UNIV :: 'a set) = card {0 ..< n}" by (rule card) also have "\ = n" by auto finally have bn: "CARD ('a) = n" . let ?enum = "(map abs [0.. n = CARD('a)" by simp then show ?thesis .. qed lemma inter_aform_plane_refine_ex_typedef: "(xi, x) \ aforms_rel \ (si, s) \ \rl_rel\sctn_rel \ length xi = d \ length (normal si) = d \ (nres_of (inter_aform_plane_lv (length xi) optns xi si), inter_sctn_specs d x s) \ \aforms_rel\nres_rel" if "\(Rep :: 'a \ nat) Abs. type_definition Rep Abs {0 ..< d :: nat}" by (rule type_impl_card_n_enum[OF that]) (rule inter_aform_plane_refine_unoverloaded; assumption) lemma inter_aform_plane_refine: "0 < d \ (xi, x) \ aforms_rel \ (si, s) \ \Id\sctn_rel \ length xi = d \ length (normal si) = d \ (nres_of (inter_aform_plane_lv (length xi) optns xi si), inter_sctn_specs d x s) \ \aforms_rel\nres_rel" by (rule inter_aform_plane_refine_ex_typedef[cancel_type_definition, simplified]) lemma Joints_reduce_aforms: "x \ Joints X \ x \ Joints (reduce_aforms prec t X)" proof (auto simp: reduce_aforms_def summarize_threshold_def[abs_def] Joints_def valuate_def aform_val_def, goal_cases) case (1 e) - from summarize_aformsE[OF \e \ _\ order_refl, of "X" "prec" t] - guess e' . + from summarize_aformsE[OF \e \ _\ order_refl] + obtain e' where + "aform_vals e X = aform_vals e' (summarize_aforms prec t (degree_aforms X) X)" + "\i. i < degree_aforms X \ e i = e' i" + "e' \ funcset UNIV {- 1..1}" + by blast thus ?case by (auto intro!: image_eqI[where x=e'] simp: aform_vals_def) qed lemma length_reduce_aform[simp]: "length (reduce_aforms optns a x) = length x" by (auto simp: reduce_aforms_def) lemma reduce_aform_refine: "(xi, x) \ aforms_rel \ length xi = d \ (RETURN (reduce_aforms prec C xi), reduce_specs d r x) \ \aforms_rel\nres_rel" apply (auto simp: reduce_specs_def nres_rel_def comps aforms_relp_def mem_lv_rel_set_rel_iff aforms_rel_def env_len_def intro!: RETURN_SPEC_refine) apply (auto simp: br_def env_len_def) by (auto simp: mem_lv_rel_set_rel_iff Joints_imp_length_eq intro!: in_image_eucl_of_list_eucl_of_list_aform Joints_reduce_aforms eucl_of_list_mem_eucl_of_list_aform) lemma aform_euclarithform_refine: "(nres_of o2 aform_form optns, approx_form_spec) \ Id \ aforms_rel \ \bool_rel\nres_rel" by (auto simp: approx_form_spec_def nres_rel_def comps aform_form_def aforms_rel_def br_def approx_form_aform dest!: approx_form_aux intro!: ivls_of_aforms) lemma aform_isFDERIV: "(\N xs fas vs. nres_of (aform_isFDERIV optns N xs fas vs), isFDERIV_spec) \ nat_rel \ \nat_rel\list_rel \ \Id\list_rel \ aforms_rel \ \bool_rel\nres_rel" by (auto simp: isFDERIV_spec_def nres_rel_def comps aform_isFDERIV_def aforms_rel_def br_def dest!: approx_form_aux intro!: ivls_of_aforms isFDERIV_approx) lemma approx_slp_lengthD: "approx_slp p slp a = Some xs \ length xs = length slp + length a" by (induction slp arbitrary: xs a) (auto simp: bind_eq_Some_conv) lemma approx_slp_outer_lengthD: "approx_slp_outer p d slp a = Some xs \ length xs = min d (length slp + length a)" by (auto simp: approx_slp_outer_def Let_def bind_eq_Some_conv o_def aforms_err_to_aforms_def aform_err_to_aform_def dest!: approx_slp_lengthD) lemma approx_slp_refine: "(nres_of o3 aform_slp prec, approx_slp_spec fas) \ nat_rel \ fas_rel \ aforms_rel \ \\aforms_rel\option_rel\nres_rel" apply (auto simp: approx_slp_spec_def comps aform_slp_def nres_rel_def intro!: RETURN_SPEC_refine ASSERT_refine_right) subgoal for a b apply (rule exI[where x = "map_option Joints (approx_slp_outer prec (length fas) (slp_of_fas fas) a)"]) apply (auto simp: option.splits aforms_rel_def br_def env_len_def Joints_imp_length_eq) apply (auto dest!: approx_slp_outer_lengthD)[] using length_slp_of_fas_le trans_le_add1 approx_slp_outer_lengthD apply blast using approx_slp_outer_plain by blast done lemma fresh_index_aforms_Nil[simp]: "fresh_index_aforms [] = 0" by (auto simp: fresh_index_aforms_def) lemma independent_aforms_Nil[simp]: "independent_aforms x [] = [x]" by (auto simp: independent_aforms_def) lemma mem_Joints_zero_iff[simp]: "x # xs \ Joints ((0, zero_pdevs) # XS) \ (x = 0 \ xs \ Joints XS)" by (auto simp: Joints_def valuate_def) lemma Joints_independent_aforms_eq: "Joints (independent_aforms x xs) = set_Cons (Affine x) (Joints xs)" by (simp add: independent_aforms_def Joints_msum_aform degree_le_fresh_index set_Cons_def) lemma independent_aforms_refine: "(independent_aforms, set_Cons) \ \rnv_rel\aform_rel \ aforms_rel \ aforms_rel" by (auto simp: aforms_rel_def br_def aform_rel_def Joints_independent_aforms_eq) end locale aform_approximate_sets = approximate_sets aform_ops Joints aforms_rel begin lemma Joints_in_lv_rel_set_relD: "(Joints xs, X) \ \lv_rel\set_rel \ X = Affine (eucl_of_list_aform xs)" unfolding lv_rel_def set_rel_br by (auto simp: br_def Joints_imp_length_eq eucl_of_list_image_Joints[symmetric]) lemma ncc_precond: "ncc_precond TYPE('a::executable_euclidean_space)" unfolding ncc_precond_def ncc_def appr_rel_def by (auto simp: aforms_rel_def compact_Affine convex_Affine dest!: Joints_in_lv_rel_set_relD brD) lemma fst_eucl_of_list_aform_map: "fst (eucl_of_list_aform (map (\x. (fst x, asdf x)) x)) = fst (eucl_of_list_aform x)" by (auto simp: eucl_of_list_aform_def o_def) lemma Affine_pdevs_of_list:\ \TODO: move!\ "Affine (fst x, pdevs_of_list (map snd (list_of_pdevs (snd x)))) = Affine x" by (auto simp: Affine_def valuate_def aform_val_def elim: pdevs_val_of_list_of_pdevs2[where X = "snd x"] pdevs_val_of_list_of_pdevs[where X = "snd x"]) end lemma aform_approximate_sets: "aform_approximate_sets prec" apply (unfold_locales) unfolding aform_ops_def approximate_set_ops.simps subgoal unfolding relAPP_def aforms_rel_def . subgoal by (force simp: aforms_of_ivls_refine) subgoal by (rule product_aforms_refine) subgoal by (rule msum_aforms'_refine) subgoal by (rule inf_aforms_refine) subgoal by (rule sup_aforms_refine) subgoal by (rule split_aform_largest_take_refine) subgoal by (rule aform_inf_inner_refine) subgoal by (rule aform_sup_inner_refine) subgoal by (rule inter_aform_plane_refine) simp_all subgoal by (auto split: option.splits intro!: reduce_aform_refine) subgoal by (force simp: width_spec_def nres_rel_def) subgoal by (rule approx_slp_refine) subgoal by (rule aform_euclarithform_refine) subgoal by (rule aform_isFDERIV) subgoal by simp subgoal by (auto simp: Joints_imp_length_eq) subgoal by (force simp: Affine_def Joints_def valuate_def intro!:) subgoal by (force simp: Affine_def Joints_def valuate_def intro!:) subgoal by (auto simp: Joints_imp_length_eq) done end diff --git a/thys/Pratt_Certificate/Pratt_Certificate.thy b/thys/Pratt_Certificate/Pratt_Certificate.thy --- a/thys/Pratt_Certificate/Pratt_Certificate.thy +++ b/thys/Pratt_Certificate/Pratt_Certificate.thy @@ -1,884 +1,889 @@ section \Pratt's Primality Certificates\ text_raw \\label{sec:pratt}\ theory Pratt_Certificate imports Complex_Main Lehmer.Lehmer begin text \ This work formalizes Pratt's proof system as described in his article ``Every Prime has a Succinct Certificate''\cite{pratt1975certificate}. The proof system makes use of two types of predicates: \begin{itemize} \item $\text{Prime}(p)$: $p$ is a prime number \item $(p, a, x)$: \\q \ prime_factors(x). [a^((p - 1) div q) \ 1] (mod p)\ \end{itemize} We represent these predicates with the following datatype: \ datatype pratt = Prime nat | Triple nat nat nat text \ Pratt describes an inference system consisting of the axiom $(p, a, 1)$ and the following inference rules: \begin{itemize} \item R1: If we know that $(p, a, x)$ and \[a^((p - 1) div q) \ 1] (mod p)\ hold for some prime number $q$ we can conclude $(p, a, qx)$ from that. \item R2: If we know that $(p, a, p - 1)$ and \[a^(p - 1) = 1] (mod p)\ hold, we can infer $\text{Prime}(p)$. \end{itemize} Both rules follow from Lehmer's theorem as we will show later on. A list of predicates (i.e., values of type @{type pratt}) is a \emph{certificate}, if it is built according to the inference system described above. I.e., a list @{term "x # xs :: pratt list"} is a certificate if @{term "xs :: pratt list"} is a certificate and @{term "x :: pratt"} is either an axiom or all preconditions of @{term "x :: pratt"} occur in @{term "xs :: pratt list"}. We call a certificate @{term "xs :: pratt list"} a \emph{certificate for @{term p}}, if @{term "Prime p"} occurs in @{term "xs :: pratt list"}. The function \valid_cert\ checks whether a list is a certificate. \ fun valid_cert :: "pratt list \ bool" where "valid_cert [] = True" | R2: "valid_cert (Prime p#xs) \ 1 < p \ valid_cert xs \ (\ a . [a^(p - 1) = 1] (mod p) \ Triple p a (p - 1) \ set xs)" | R1: "valid_cert (Triple p a x # xs) \ p > 1 \ 0 < x \ valid_cert xs \ (x=1 \ (\q y. x = q * y \ Prime q \ set xs \ Triple p a y \ set xs \ [a^((p - 1) div q) \ 1] (mod p)))" text \ We define a function @{term size_cert} to measure the size of a certificate, assuming a binary encoding of numbers. We will use this to show that there is a certificate for a prime number $p$ such that the size of the certificate is polynomially bounded in the size of the binary representation of $p$. \ fun size_pratt :: "pratt \ real" where "size_pratt (Prime p) = log 2 p" | "size_pratt (Triple p a x) = log 2 p + log 2 a + log 2 x" fun size_cert :: "pratt list \ real" where "size_cert [] = 0" | "size_cert (x # xs) = 1 + size_pratt x + size_cert xs" subsection \Soundness\ text \ In Section \ref{sec:pratt} we introduced the predicates $\text{Prime}(p)$ and $(p, a, x)$. In this section we show that for a certificate every predicate occurring in this certificate holds. In particular, if $\text{Prime}(p)$ occurs in a certificate, $p$ is prime. \ lemma prime_factors_one [simp]: shows "prime_factors (Suc 0) = {}" using prime_factorization_1 [where ?'a = nat] by simp lemma prime_factors_of_prime: fixes p :: nat assumes "prime p" shows "prime_factors p = {p}" using assms by (fact prime_prime_factors) definition pratt_triple :: "nat \ nat \ nat \ bool" where "pratt_triple p a x \ x > 0 \ (\q\prime_factors x. [a ^ ((p - 1) div q) \ 1] (mod p))" lemma pratt_triple_1: "p > 1 \ x = 1 \ pratt_triple p a x" by (auto simp: pratt_triple_def) lemma pratt_triple_extend: assumes "prime q" "pratt_triple p a y" "p > 1" "x > 0" "x = q * y" "[a ^ ((p - 1) div q) \ 1] (mod p)" shows "pratt_triple p a x" proof - have "prime_factors x = insert q (prime_factors y)" using assms by (simp add: prime_factors_product prime_prime_factors) also have "\r\\. [a ^ ((p - 1) div r) \ 1] (mod p)" using assms by (auto simp: pratt_triple_def) finally show ?thesis using assms unfolding pratt_triple_def by blast qed lemma pratt_triple_imp_prime: assumes "pratt_triple p a x" "p > 1" "x = p - 1" "[a ^ (p - 1) = 1] (mod p)" shows "prime p" using lehmers_theorem[of p a] assms by (auto simp: pratt_triple_def) theorem pratt_sound: assumes 1: "valid_cert c" assumes 2: "t \ set c" shows "(t = Prime p \ prime p) \ (t = Triple p a x \ ((\q \ prime_factors x . [a^((p - 1) div q) \ 1] (mod p)) \ 0 q \ prime_factors x . [a^((p - 1) div q) \ 1] (mod p)) \ 00" using Cons.prems by auto obtain q z where "x=q*z" "Prime q \ set ys \ Triple p a z \ set ys" and cong:"[a^((p - 1) div q) \ 1] (mod p)" using Cons.prems x_y by auto then have factors_IH:"(\ r \ prime_factors z . [a^((p - 1) div r) \ 1] (mod p))" "prime q" "z>0" using Cons.IH Cons.prems \x>0\ \y=Triple p a x\ by force+ then have "prime_factors x = prime_factors z \ {q}" using \x =q*z\ \x>0\ by (simp add: prime_factors_product prime_factors_of_prime) then have "(\ q \ prime_factors x . [a^((p - 1) div q) \ 1] (mod p)) \ 0 < x" using factors_IH cong by (simp add: \x>0\) } ultimately have y_Triple:"y=Triple p a x \ (\ q \ prime_factors x . [a^((p - 1) div q) \ 1] (mod p)) \ 02" then obtain a where a:"[a^(p - 1) = 1] (mod p)" "Triple p a (p - 1) \ set ys" using Cons.prems by auto then have Bier:"(\q\prime_factors (p - 1). [a^((p - 1) div q) \ 1] (mod p))" using Cons.IH Cons.prems(1) by (simp add:y(1)) then have "prime p" using lehmers_theorem[OF _ _a(1)] \p>2\ by fastforce } moreover { assume "y=Prime p" "p=2" hence "prime p" by simp } moreover { assume "y=Prime p" then have "p>1" using Cons.prems by simp } ultimately have y_Prime:"y = Prime p \ prime p" by linarith show ?case proof (cases "t \ set ys") case True show ?thesis using Cons.IH[OF _ True] Cons.prems(1) by (cases y) auto next case False thus ?thesis using Cons.prems(2) y_Prime y_Triple by force qed qed corollary pratt_primeI: assumes "valid_cert xs" "Prime p \ set xs" shows "prime p" using pratt_sound[OF assms] by simp subsection \Completeness\ text \ In this section we show completeness of Pratt's proof system, i.e., we show that for every prime number $p$ there exists a certificate for $p$. We also give an upper bound for the size of a minimal certificate The prove we give is constructive. We assume that we have certificates for all prime factors of $p - 1$ and use these to build a certificate for $p$ from that. It is important to note that certificates can be concatenated. \ lemma valid_cert_appendI: assumes "valid_cert r" assumes "valid_cert s" shows "valid_cert (r @ s)" using assms proof (induction r) case (Cons y ys) then show ?case by (cases y) auto qed simp lemma valid_cert_concatI: "(\x \ set xs . valid_cert x) \ valid_cert (concat xs)" by (induction xs) (auto simp add: valid_cert_appendI) lemma size_pratt_le: fixes d::real assumes "\ x \ set c. size_pratt x \ d" shows "size_cert c \ length c * (1 + d)" using assms by (induction c) (simp_all add: algebra_simps) fun build_fpc :: "nat \ nat \ nat \ nat list \ pratt list" where "build_fpc p a r [] = [Triple p a r]" | "build_fpc p a r (y # ys) = Triple p a r # build_fpc p a (r div y) ys" text \ The function @{term build_fpc} helps us to construct a certificate for $p$ from the certificates for the prime factors of $p - 1$. Called as @{term "build_fpc p a (p - 1) qs"} where $@{term "qs"} = q_1 \ldots q_n$ is prime decomposition of $p - 1$ such that $q_1 \cdot \dotsb \cdot q_n = @{term "p - 1 :: nat"}$, it returns the following list of predicates: \[ (p,a,p-1), (p,a,\frac{p - 1}{q_1}), (p,a,\frac{p - 1}{q_1 q_2}), \ldots, (p,a,\frac{p-1}{q_1 \ldots q_n}) = (p,a,1) \] I.e., if there is an appropriate $a$ and and a certificate @{term rs} for all prime factors of $p$, then we can construct a certificate for $p$ as @{term [display] "Prime p # build_fpc p a (p - 1) qs @ rs"} \ text \ The following lemma shows that \build_fpc\ extends a certificate that satisfies the preconditions described before to a correct certificate. \ lemma correct_fpc: assumes "valid_cert xs" "p > 1" assumes "prod_list qs = r" "r \ 0" assumes "\ q \ set qs . Prime q \ set xs" assumes "\ q \ set qs . [a^((p - 1) div q) \ 1] (mod p)" shows "valid_cert (build_fpc p a r qs @ xs)" using assms proof (induction qs arbitrary: r) case Nil thus ?case by auto next case (Cons y ys) have "prod_list ys = r div y" using Cons.prems by auto then have T_in: "Triple p a (prod_list ys) \ set (build_fpc p a (r div y) ys @ xs)" by (cases ys) auto have "valid_cert (build_fpc p a (r div y) ys @ xs)" using Cons.prems by (intro Cons.IH) auto then have "valid_cert (Triple p a r # build_fpc p a (r div y) ys @ xs)" using \r \ 0\ T_in Cons.prems by auto then show ?case by simp qed lemma length_fpc: "length (build_fpc p a r qs) = length qs + 1" by (induction qs arbitrary: r) auto lemma div_gt_0: fixes m n :: nat assumes "m \ n" "0 < m" shows "0 < n div m" proof - have "0 < m div m" using \0 < m\ div_self by auto also have "m div m \ n div m" using \m \ n\ by (rule div_le_mono) finally show ?thesis . qed lemma size_pratt_fpc: assumes "a \ p" "r \ p" "0 < a" "0 < r" "0 < p" "prod_list qs = r" shows "\x \ set (build_fpc p a r qs) . size_pratt x \ 3 * log 2 p" using assms proof (induction qs arbitrary: r) case Nil then have "log 2 a \ log 2 p" "log 2 r \ log 2 p" by auto then show ?case by simp next case (Cons q qs) then have "log 2 a \ log 2 p" "log 2 r \ log 2 p" by auto then have "log 2 a + log 2 r \ 2 * log 2 p" by arith moreover have "r div q > 0" using Cons.prems by (fastforce intro: div_gt_0) moreover hence "prod_list qs = r div q" using Cons.prems(6) by auto moreover have "r div q \ p" using \r\p\ div_le_dividend[of r q] by linarith ultimately show ?case using Cons by simp qed lemma concat_set: assumes "\ q \ qs . \ c \ set cs . Prime q \ set c" shows "\ q \ qs . Prime q \ set (concat cs)" using assms by (induction cs) auto lemma p_in_prime_factorsE: fixes n :: nat assumes "p \ prime_factors n" "0 < n" obtains "2 \ p" "p \ n" "p dvd n" "prime p" proof from assms show "prime p" by auto then show "2 \ p" by (auto dest: prime_gt_1_nat) from assms show "p dvd n" by auto then show "p \ n" using \0 < n\ by (rule dvd_imp_le) qed lemma prime_factors_list_prime: fixes n :: nat assumes "prime n" shows "\ qs. prime_factors n = set qs \ prod_list qs = n \ length qs = 1" using assms by (auto simp add: prime_factorization_prime intro: exI [of _ "[n]"]) lemma prime_factors_list: fixes n :: nat assumes "3 < n" "\ prime n" shows "\ qs. prime_factors n = set qs \ prod_list qs = n \ length qs \ 2" using assms proof (induction n rule: less_induct) case (less n) obtain p where "p \ prime_factors n" using \n > 3\ prime_factors_elem by force then have p':"2 \ p" "p \ n" "p dvd n" "prime p" using \3 < n\ by (auto elim: p_in_prime_factorsE) { assume "n div p > 3" "\ prime (n div p)" then obtain qs where "prime_factors (n div p) = set qs" "prod_list qs = (n div p)" "length qs \ 2" using p' by atomize_elim (auto intro: less simp: div_gt_0) moreover have "prime_factors (p * (n div p)) = insert p (prime_factors (n div p))" using \3 < n\ \2 \ p\ \p \ n\ \prime p\ by (auto simp: prime_factors_product div_gt_0 prime_factors_of_prime) ultimately have "prime_factors n = set (p # qs)" "prod_list (p # qs) = n" "length (p#qs) \ 2" using \p dvd n\ by simp_all hence ?case by blast } moreover { assume "prime (n div p)" then obtain qs where "prime_factors (n div p) = set qs" "prod_list qs = (n div p)" "length qs = 1" using prime_factors_list_prime by blast moreover have "prime_factors (p * (n div p)) = insert p (prime_factors (n div p))" using \3 < n\ \2 \ p\ \p \ n\ \prime p\ by (auto simp: prime_factors_product div_gt_0 prime_factors_of_prime) ultimately have "prime_factors n = set (p # qs)" "prod_list (p # qs) = n" "length (p#qs) \ 2" using \p dvd n\ by simp_all hence ?case by blast } note case_prime = this moreover { assume "n div p = 1" hence "n = p" using \n>3\ using One_leq_div[OF \p dvd n\] p'(2) by force hence ?case using \prime p\ \\ prime n\ by auto } moreover { assume "n div p = 2" hence ?case using case_prime by force } moreover { assume "n div p = 3" hence ?case using p' case_prime by force } ultimately show ?case using p' div_gt_0[of p n] case_prime by fastforce qed lemma prod_list_ge: fixes xs::"nat list" assumes "\ x \ set xs . x \ 1" shows "prod_list xs \ 1" using assms by (induction xs) auto lemma sum_list_log: fixes b::real fixes xs::"nat list" assumes b: "b > 0" "b \ 1" assumes xs:"\ x \ set xs . x \ b" shows "(\x\xs. log b x) = log b (prod_list xs)" using assms proof (induction xs) case Nil thus ?case by simp next case (Cons y ys) have "real (prod_list ys) > 0" using prod_list_ge Cons.prems by fastforce thus ?case using log_mult[OF Cons.prems(1-2)] Cons by force qed lemma concat_length_le: fixes g :: "nat \ real" assumes "\ x \ set xs . real (length (f x)) \ g x" shows "length (concat (map f xs)) \ (\x\xs. g x)" using assms by (induction xs) force+ lemma prime_gt_3_impl_p_minus_one_not_prime: fixes p::nat assumes "prime p" "p>3" shows "\ prime (p - 1)" proof assume "prime (p - 1)" have "\ even p" using assms by (simp add: prime_odd_nat) hence "2 dvd (p - 1)" by presburger then obtain q where "p - 1 = 2 * q" .. then have "2 \ prime_factors (p - 1)" using \p>3\ by (auto simp: prime_factorization_times_prime) thus False using prime_factors_of_prime \p>3\ \prime (p - 1)\ by auto qed text \ We now prove that Pratt's proof system is complete and derive upper bounds for the length and the size of the entries of a minimal certificate. \ theorem pratt_complete': assumes "prime p" shows "\c. Prime p \ set c \ valid_cert c \ length c \ 6*log 2 p - 4 \ (\ x \ set c. size_pratt x \ 3 * log 2 p)" using assms proof (induction p rule: less_induct) case (less p) from \prime p\ have "p > 1" by (rule prime_gt_1_nat) then consider "p = 2" | " p = 3" | "p > 3" by force thus ?case proof cases assume [simp]: "p = 2" have "Prime p \ set [Prime 2, Triple 2 1 1]" by simp thus ?case by fastforce next assume [simp]: "p = 3" let ?cert = "[Prime 3, Triple 3 2 2, Triple 3 2 1, Prime 2, Triple 2 1 1]" have "length ?cert \ 6*log 2 p - 4 \ 3 \ 2 * log 2 3" by simp also have "2 * log 2 3 = log 2 (3 ^ 2 :: real)" by (subst log_nat_power) simp_all also have "\ = log 2 9" by simp also have "3 \ log 2 9 \ True" by (subst le_log_iff) simp_all finally show ?case by (intro exI[where x = "?cert"]) (simp add: cong_def) next assume "p > 3" have qlp: "\q \ prime_factors (p - 1) . q < p" using \prime p\ by (metis One_nat_def Suc_pred le_imp_less_Suc lessI less_trans p_in_prime_factorsE prime_gt_1_nat zero_less_diff) hence factor_certs:"\q \ prime_factors (p - 1) . (\c . ((Prime q \ set c) \ (valid_cert c) \ length c \ 6*log 2 q - 4) \ (\ x \ set c. size_pratt x \ 3 * log 2 q))" by (auto intro: less.IH) obtain a where a:"[a^(p - 1) = 1] (mod p) \ (\ q. q \ prime_factors (p - 1) \ [a^((p - 1) div q) \ 1] (mod p))" and a_size: "a > 0" "a < p" using converse_lehmer[OF \prime p\] by blast have "\ prime (p - 1)" using \p>3\ prime_gt_3_impl_p_minus_one_not_prime \prime p\ by auto have "p \ 4" using \prime p\ by auto hence "p - 1 > 3" using \p > 3\ by auto then obtain qs where prod_qs_eq:"prod_list qs = p - 1" and qs_eq:"set qs = prime_factors (p - 1)" and qs_length_eq: "length qs \ 2" using prime_factors_list[OF _ \\ prime (p - 1)\] by auto obtain f where f:"\q \ prime_factors (p - 1) . \ c. f q = c \ ((Prime q \ set c) \ (valid_cert c) \ length c \ 6*log 2 q - 4) \ (\ x \ set c. size_pratt x \ 3 * log 2 q)" using factor_certs by metis let ?cs = "map f qs" have cs: "\q \ prime_factors (p - 1) . (\c \ set ?cs . (Prime q \ set c) \ (valid_cert c) \ length c \ 6*log 2 q - 4 \ (\ x \ set c. size_pratt x \ 3 * log 2 q))" using f qs_eq by auto have cs_cert_size: "\c \ set ?cs . \ x \ set c. size_pratt x \ 3 * log 2 p" proof fix c assume "c \ set (map f qs)" then obtain q where "c = f q" and "q \ set qs" by auto hence *:"\ x \ set c. size_pratt x \ 3 * log 2 q" using f qs_eq by blast have "q < p" "q > 0" using qlp \q \ set qs\ qs_eq prime_factors_gt_0_nat by auto show "\ x \ set c. size_pratt x \ 3 * log 2 p" proof fix x assume "x \ set c" hence "size_pratt x \ 3 * log 2 q" using * by fastforce also have "\ \ 3 * log 2 p" using \q < p\ \q > 0\ \p > 3\ by simp finally show "size_pratt x \ 3 * log 2 p" . qed qed have cs_valid_all: "\c \ set ?cs . valid_cert c" using f qs_eq by fastforce have "\x \ set (build_fpc p a (p - 1) qs). size_pratt x \ 3 * log 2 p" using cs_cert_size a_size \p > 3\ prod_qs_eq by (intro size_pratt_fpc) auto hence "\x \ set (build_fpc p a (p - 1) qs @ concat ?cs) . size_pratt x \ 3 * log 2 p" using cs_cert_size by auto moreover have "Triple p a (p - 1) \ set (build_fpc p a (p - 1) qs @ concat ?cs)" by (cases qs) auto moreover have "valid_cert ((build_fpc p a (p - 1) qs)@ concat ?cs)" proof (rule correct_fpc) show "valid_cert (concat ?cs)" using cs_valid_all by (auto simp: valid_cert_concatI) show "prod_list qs = p - 1" by (rule prod_qs_eq) show "p - 1 \ 0" using prime_gt_1_nat[OF \prime p\] by arith show "\ q \ set qs . Prime q \ set (concat ?cs)" using concat_set[of "prime_factors (p - 1)"] cs qs_eq by blast show "\ q \ set qs . [a^((p - 1) div q) \ 1] (mod p)" using qs_eq a by auto qed (insert \p > 3\, simp_all) moreover { let ?k = "length qs" have qs_ge_2:"\q \ set qs . q \ 2" using qs_eq by (auto intro: prime_ge_2_nat) have "\x\set qs. real (length (f x)) \ 6 * log 2 (real x) - 4" using f qs_eq by blast hence "length (concat ?cs) \ (\q\qs. 6*log 2 q - 4)" using concat_length_le by fast hence "length (Prime p # ((build_fpc p a (p - 1) qs)@ concat ?cs)) \ ((\q\(map real qs). 6*log 2 q - 4) + ?k + 2)" by (simp add: o_def length_fpc) also have "\ = (6*(\q\(map real qs). log 2 q) + (-4 * real ?k) + ?k + 2)" by (simp add: o_def sum_list_subtractf sum_list_triv sum_list_const_mult) also have "\ \ 6*log 2 (p - 1) - 4" using \?k\2\ prod_qs_eq sum_list_log[of 2 qs] qs_ge_2 by force also have "\ \ 6*log 2 p - 4" using log_le_cancel_iff[of 2 "p - 1" p] \p>3\ by force ultimately have "length (Prime p # ((build_fpc p a (p - 1) qs)@ concat ?cs)) \ 6*log 2 p - 4" by linarith } ultimately obtain c where c:"Triple p a (p - 1) \ set c" "valid_cert c" "length (Prime p #c) \ 6*log 2 p - 4" "(\ x \ set c. size_pratt x \ 3 * log 2 p)" by blast hence "Prime p \ set (Prime p # c)" "valid_cert (Prime p # c)" "(\ x \ set (Prime p # c). size_pratt x \ 3 * log 2 p)" using a \prime p\ by (auto simp: Primes.prime_gt_Suc_0_nat) thus ?case using c by blast qed qed text \ We now recapitulate our results. A number $p$ is prime if and only if there is a certificate for $p$. Moreover, for a prime $p$ there always is a certificate whose size is polynomially bounded in the logarithm of $p$. \ corollary pratt: "prime p \ (\c. Prime p \ set c \ valid_cert c)" using pratt_complete' pratt_sound(1) by blast corollary pratt_size: assumes "prime p" shows "\c. Prime p \ set c \ valid_cert c \ size_cert c \ (6 * log 2 p - 4) * (1 + 3 * log 2 p)" proof - obtain c where c: "Prime p \ set c" "valid_cert c" and len: "length c \ 6*log 2 p - 4" and "(\ x \ set c. size_pratt x \ 3 * log 2 p)" using pratt_complete' assms by blast hence "size_cert c \ length c * (1 + 3 * log 2 p)" by (simp add: size_pratt_le) also have "\ \ (6*log 2 p - 4) * (1 + 3 * log 2 p)" using len by simp finally show ?thesis using c by blast qed subsection \Efficient modular exponentiation\ locale efficient_power = fixes f :: "'a \ 'a \ 'a" assumes f_assoc: "\x z. f x (f x z) = f (f x x) z" begin function efficient_power :: "'a \ 'a \ nat \ 'a" where "efficient_power y x 0 = y" | "efficient_power y x (Suc 0) = f x y" | "n \ 0 \ even n \ efficient_power y x n = efficient_power y (f x x) (n div 2)" | "n \ 1 \ odd n \ efficient_power y x n = efficient_power (f x y) (f x x) (n div 2)" by force+ termination by (relation "measure (snd \ snd)") (auto elim: oddE) lemma efficient_power_code: "efficient_power y x n = (if n = 0 then y else if n = 1 then f x y else if even n then efficient_power y (f x x) (n div 2) else efficient_power (f x y) (f x x) (n div 2))" by (induction y x n rule: efficient_power.induct) auto lemma efficient_power_correct: "efficient_power y x n = (f x ^^ n) y" proof - have [simp]: "f ^^ 2 = (\x. f (f x))" for f :: "'a \ 'a" by (simp add: eval_nat_numeral o_def) show ?thesis by (induction y x n rule: efficient_power.induct) (auto elim!: evenE oddE simp: funpow_mult [symmetric] funpow_Suc_right f_assoc simp del: funpow.simps(2)) qed end interpretation mod_exp_nat: efficient_power "\x y :: nat. (x * y) mod m" by standard (simp add: mod_mult_left_eq mod_mult_right_eq mult_ac) definition mod_exp_nat_aux where "mod_exp_nat_aux = mod_exp_nat.efficient_power" lemma mod_exp_nat_aux_code [code]: "mod_exp_nat_aux m y x n = (if n = 0 then y else if n = 1 then (x * y) mod m else if even n then mod_exp_nat_aux m y ((x * x) mod m) (n div 2) else mod_exp_nat_aux m ((x * y) mod m) ((x * x) mod m) (n div 2))" unfolding mod_exp_nat_aux_def by (rule mod_exp_nat.efficient_power_code) lemma mod_exp_nat_aux_correct: "mod_exp_nat_aux m y x n mod m = (x ^ n * y) mod m" proof - have "mod_exp_nat_aux m y x n = ((\y. x * y mod m) ^^ n) y" by (simp add: mod_exp_nat_aux_def mod_exp_nat.efficient_power_correct) also have "((\y. x * y mod m) ^^ n) y mod m = (x ^ n * y) mod m" proof (induction n) case (Suc n) hence "x * ((\y. x * y mod m) ^^ n) y mod m = x * x ^ n * y mod m" by (metis mod_mult_right_eq mult.assoc) thus ?case by auto qed auto finally show ?thesis . qed definition mod_exp_nat :: "nat \ nat \ nat \ nat" where [code_abbrev]: "mod_exp_nat b e m = (b ^ e) mod m" lemma mod_exp_nat_code [code]: "mod_exp_nat b e m = mod_exp_nat_aux m 1 b e mod m" by (simp add: mod_exp_nat_def mod_exp_nat_aux_correct) lemmas [code_unfold] = cong_def lemma eval_mod_exp_nat_aux [simp]: "mod_exp_nat_aux m y x 0 = y" "mod_exp_nat_aux m y x (Suc 0) = (x * y) mod m" "mod_exp_nat_aux m y x (numeral (num.Bit0 n)) = mod_exp_nat_aux m y (x\<^sup>2 mod m) (numeral n)" "mod_exp_nat_aux m y x (numeral (num.Bit1 n)) = mod_exp_nat_aux m ((x * y) mod m) (x\<^sup>2 mod m) (numeral n)" proof - define n' where "n' = (numeral n :: nat)" have [simp]: "n' \ 0" by (auto simp: n'_def) show "mod_exp_nat_aux m y x 0 = y" and "mod_exp_nat_aux m y x (Suc 0) = (x * y) mod m" by (simp_all add: mod_exp_nat_aux_def) have "numeral (num.Bit0 n) = (2 * n')" by (subst numeral.numeral_Bit0) (simp del: arith_simps add: n'_def) also have "mod_exp_nat_aux m y x \ = mod_exp_nat_aux m y (x^2 mod m) n'" by (subst mod_exp_nat_aux_code) (simp_all add: power2_eq_square) finally show "mod_exp_nat_aux m y x (numeral (num.Bit0 n)) = mod_exp_nat_aux m y (x\<^sup>2 mod m) (numeral n)" by (simp add: n'_def) have "numeral (num.Bit1 n) = Suc (2 * n')" by (subst numeral.numeral_Bit1) (simp del: arith_simps add: n'_def) also have "mod_exp_nat_aux m y x \ = mod_exp_nat_aux m ((x * y) mod m) (x^2 mod m) n'" by (subst mod_exp_nat_aux_code) (simp_all add: power2_eq_square) finally show "mod_exp_nat_aux m y x (numeral (num.Bit1 n)) = mod_exp_nat_aux m ((x * y) mod m) (x\<^sup>2 mod m) (numeral n)" by (simp add: n'_def) qed lemma eval_mod_exp [simp]: "mod_exp_nat b' 0 m' = 1 mod m'" "mod_exp_nat b' 1 m' = b' mod m'" "mod_exp_nat b' (Suc 0) m' = b' mod m'" "mod_exp_nat b' e' 0 = b' ^ e'" "mod_exp_nat b' e' 1 = 0" "mod_exp_nat b' e' (Suc 0) = 0" "mod_exp_nat 0 1 m' = 0" "mod_exp_nat 0 (Suc 0) m' = 0" "mod_exp_nat 0 (numeral e) m' = 0" "mod_exp_nat 1 e' m' = 1 mod m'" "mod_exp_nat (Suc 0) e' m' = 1 mod m'" "mod_exp_nat (numeral b) (numeral e) (numeral m) = mod_exp_nat_aux (numeral m) 1 (numeral b) (numeral e) mod numeral m" by (simp_all add: mod_exp_nat_def mod_exp_nat_aux_correct) subsection \Executable certificate checker\ lemmas [code] = valid_cert.simps(1) context begin lemma valid_cert_Cons1 [code]: "valid_cert (Prime p # xs) \ p > 1 \ (\t\set xs. case t of Prime _ \ False | Triple p' a x \ p' = p \ x = p - 1 \ mod_exp_nat a (p-1) p = 1 ) \ valid_cert xs" (is "?lhs = ?rhs") proof assume ?lhs thus ?rhs by (auto simp: mod_exp_nat_def cong_def split: pratt.splits) next assume ?rhs hence "p > 1" "valid_cert xs" by blast+ moreover from \?rhs\ obtain t where "t \ set xs" "case t of Prime _ \ False | Triple p' a x \ p' = p \ x = p - 1 \ [a^(p-1) = 1] (mod p)" by (auto simp: cong_def mod_exp_nat_def cong: pratt.case_cong) ultimately show ?lhs by (cases t) auto qed private lemma Suc_0_mod_eq_Suc_0_iff: "Suc 0 mod n = Suc 0 \ n \ Suc 0" proof - consider "n = 0" | "n = Suc 0" | "n > 1" by (cases n) auto thus ?thesis by cases auto qed private lemma Suc_0_eq_Suc_0_mod_iff: "Suc 0 = Suc 0 mod n \ n \ Suc 0" using Suc_0_mod_eq_Suc_0_iff by (simp add: eq_commute) lemma valid_cert_Cons2 [code]: "valid_cert (Triple p a x # xs) \ x > 0 \ p > 1 \ (x = 1 \ ( (\t\set xs. case t of Prime _ \ False | Triple p' a' y \ p' = p \ a' = a \ y dvd x \ (let q = x div y in Prime q \ set xs \ mod_exp_nat a ((p-1) div q) p \ 1)))) \ valid_cert xs" (is "?lhs = ?rhs") proof assume ?lhs from \?lhs\ have pos: "x > 0" and gt_1: "p > 1" and valid: "valid_cert xs" by simp_all show ?rhs proof (cases "x = 1") case True with \?lhs\ show ?thesis by auto next case False with \?lhs\ have "(\q y. x = q * y \ Prime q \ set xs \ Triple p a y \ set xs \ [a^((p - 1) div q) \ 1] (mod p))" by auto - then guess q y by (elim exE conjE) note qy = this + then obtain q y where qy: + "x = q * y" + "Prime q \ set xs" + "Triple p a y \ set xs" + "[a ^ ((p - 1) div q) \ 1] (mod p)" + by blast hence "(\t\set xs. case t of Prime _ \ False | Triple p' a' y \ p' = p \ a' = a \ y dvd x \ (let q = x div y in Prime q \ set xs \ mod_exp_nat a ((p-1) div q) p \ 1))" using pos gt_1 by (intro bexI [of _ "Triple p a y"]) (auto simp: Suc_0_mod_eq_Suc_0_iff Suc_0_eq_Suc_0_mod_iff cong_def mod_exp_nat_def) with pos gt_1 valid show ?thesis by blast qed next assume ?rhs hence pos: "x > 0" and gt_1: "p > 1" and valid: "valid_cert xs" by simp_all show ?lhs proof (cases "x = 1") case True with \?rhs\ show ?thesis by auto next case False with \?rhs\ obtain t where t: "t \ set xs" "case t of Prime x \ False | Triple p' a' y \ p' = p \ a' = a \ y dvd x \ (let q = x div y in Prime q \ set xs \ mod_exp_nat a ((p - 1) div q) p \ 1)" by auto then obtain y where y: "t = Triple p a y" "y dvd x" "let q = x div y in Prime q \ set xs \ mod_exp_nat a ((p - 1) div q) p \ 1" by (cases t rule: pratt.exhaust) auto with gt_1 have y': "let q = x div y in Prime q \ set xs \ [a^((p - 1) div q) \ 1] (mod p)" by (auto simp: cong_def Let_def mod_exp_nat_def Suc_0_mod_eq_Suc_0_iff Suc_0_eq_Suc_0_mod_iff) define q where "q = x div y" have "\q y. x = q * y \ Prime q \ set xs \ Triple p a y \ set xs \ [a^((p - 1) div q) \ 1] (mod p)" by (rule exI[of _ q], rule exI[of _ y]) (insert t y y', auto simp: Let_def q_def) with pos gt_1 valid show ?thesis by simp qed qed declare valid_cert.simps(2,3) [simp del] lemmas eval_valid_cert = valid_cert.simps(1) valid_cert_Cons1 valid_cert_Cons2 end text \ The following alternative tree representation of certificates is better suited for efficient checking. \ datatype pratt_tree = Pratt_Node "nat \ nat \ pratt_tree list" fun pratt_tree_number where "pratt_tree_number (Pratt_Node (n, _, _)) = n" text \ The following function checks that a given list contains all the prime factors of the given number. \ fun check_prime_factors_subset :: "nat \ nat list \ bool" where "check_prime_factors_subset n [] \ n = 1" | "check_prime_factors_subset n (p # ps) \ (if n = 0 then False else (if p > 1 \ p dvd n then check_prime_factors_subset (n div p) (p # ps) else check_prime_factors_subset n ps))" lemma check_prime_factors_subset_0 [simp]: "\check_prime_factors_subset 0 ps" by (induction ps) auto lemmas [simp del] = check_prime_factors_subset.simps(2) lemma check_prime_factors_subset_Cons [simp]: "check_prime_factors_subset (Suc 0) (p # ps) \ check_prime_factors_subset (Suc 0) ps" "check_prime_factors_subset 1 (p # ps) \ check_prime_factors_subset 1 ps" "p > 1 \ p dvd numeral n \ check_prime_factors_subset (numeral n) (p # ps) \ check_prime_factors_subset (numeral n div p) (p # ps)" "p \ 1 \ \p dvd numeral n \ check_prime_factors_subset (numeral n) (p # ps) \ check_prime_factors_subset (numeral n) ps" by (subst check_prime_factors_subset.simps; force)+ lemma check_prime_factors_subset_correct: assumes "check_prime_factors_subset n ps" "list_all prime ps" shows "prime_factors n \ set ps" using assms proof (induction n ps rule: check_prime_factors_subset.induct) case (2 n p ps) note * = this from "2.prems" have "prime p" and "p > 1" by (auto simp: prime_gt_Suc_0_nat) consider "n = 0" | "n > 0" "p dvd n" | "n > 0" "\(p dvd n)" by blast thus ?case proof cases case 2 hence "n div p > 0" by auto hence "prime_factors ((n div p) * p) = insert p (prime_factors (n div p))" using \p > 1\ \prime p\ by (auto simp: prime_factors_product prime_prime_factors) also have "(n div p) * p = n" using 2 by auto finally show ?thesis using 2 \p > 1\ * by (auto simp: check_prime_factors_subset.simps(2)[of n]) next case 3 with * and \p > 1\ show ?thesis by (auto simp: check_prime_factors_subset.simps(2)[of n]) qed auto qed auto fun valid_pratt_tree where "valid_pratt_tree (Pratt_Node (n, a, ts)) \ n \ 2 \ check_prime_factors_subset (n - 1) (map pratt_tree_number ts) \ [a ^ (n - 1) = 1] (mod n) \ (\t\set ts. [a ^ ((n - 1) div pratt_tree_number t) \ 1] (mod n)) \ (\t\set ts. valid_pratt_tree t)" lemma valid_pratt_tree_code [code]: "valid_pratt_tree (Pratt_Node (n, a, ts)) \ n \ 2 \ check_prime_factors_subset (n - 1) (map pratt_tree_number ts) \ mod_exp_nat a (n - 1) n = 1 \ (\t\set ts. mod_exp_nat a ((n - 1) div pratt_tree_number t) n \ 1) \ (\t\set ts. valid_pratt_tree t)" by (simp add: mod_exp_nat_def cong_def) lemma valid_pratt_tree_imp_prime: assumes "valid_pratt_tree t" shows "prime (pratt_tree_number t)" using assms proof (induction t rule: valid_pratt_tree.induct) case (1 n a ts) from 1 have "prime_factors (n - 1) \ set (map pratt_tree_number ts)" by (intro check_prime_factors_subset_correct) (auto simp: list.pred_set) with 1 show ?case by (intro lehmers_theorem[where a = a]) auto qed lemma valid_pratt_tree_imp_prime': assumes "PROP (Trueprop (valid_pratt_tree (Pratt_Node (n, a, ts)))) \ PROP (Trueprop True)" shows "prime n" proof - have "valid_pratt_tree (Pratt_Node (n, a, ts))" by (subst assms) auto from valid_pratt_tree_imp_prime[OF this] show ?thesis by simp qed subsection \Proof method setup\ theorem lehmers_theorem': fixes p :: nat assumes "list_all prime ps" "a \ a" "n \ n" assumes "list_all (\p. mod_exp_nat a ((n - 1) div p) n \ 1) ps" "mod_exp_nat a (n - 1) n = 1" assumes "check_prime_factors_subset (n - 1) ps" "2 \ n" shows "prime n" using assms check_prime_factors_subset_correct[OF assms(6,1)] by (intro lehmers_theorem[where a = a]) (auto simp: cong_def mod_exp_nat_def list.pred_set) lemma list_all_ConsI: "P x \ list_all P xs \ list_all P (x # xs)" by simp ML_file \pratt.ML\ method_setup pratt = \ Scan.lift (Pratt.tac_config_parser -- Scan.option Pratt.cert_cartouche) >> (fn (config, cert) => fn ctxt => SIMPLE_METHOD (HEADGOAL (Pratt.tac config cert ctxt))) \ "Prove primality of natural numbers using Pratt certificates." text \ The proof method replays a given Pratt certificate to prove the primality of a given number. If no certificate is given, the method attempts to compute one. The computed certificate is then also printed with a prompt to insert it into the proof document so that it does not have to be recomputed the next time. The format of the certificates is compatible with those generated by Mathematica. Therefore, for larger numbers, certificates generated by Mathematica can be used with this method directly. \ lemma "prime (47 :: nat)" by (pratt (silent)) lemma "prime (2503 :: nat)" by pratt lemma "prime (7919 :: nat)" by pratt lemma "prime (131059 :: nat)" by (pratt \{131059, 2, {2, {3, 2, {2}}, {809, 3, {2, {101, 2, {2, {5, 2, {2}}}}}}}}\) 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,2042 @@ 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: 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 . + from assms obtain a b where "cbox a b \ A" "x \ box a b" "\i\Basis. a \ i < b \ i" + by (rule open_contains_cbox) 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 of_bool_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/Probabilistic_Timed_Automata/ROOT b/thys/Probabilistic_Timed_Automata/ROOT --- a/thys/Probabilistic_Timed_Automata/ROOT +++ b/thys/Probabilistic_Timed_Automata/ROOT @@ -1,17 +1,18 @@ chapter AFP session Probabilistic_Timed_Automata (AFP) = "Markov_Models" + options [timeout = 1200] sessions + "Pure-ex" "Timed_Automata" "HOL-Eisbach" directories "library" theories [document = false] "library/Lib" theories "PTA" "PTA_Reachability" document_files "root.tex" "root.bib" diff --git a/thys/Probabilistic_Timed_Automata/library/Graphs.thy b/thys/Probabilistic_Timed_Automata/library/Graphs.thy --- a/thys/Probabilistic_Timed_Automata/library/Graphs.thy +++ b/thys/Probabilistic_Timed_Automata/library/Graphs.thy @@ -1,1362 +1,1365 @@ (* Author: Simon Wimmer *) theory Graphs imports More_List Stream_More "HOL-Library.Rewrite" Instantiate_Existentials begin section \Graphs\ subsection \Basic Definitions and Theorems\ locale Graph_Defs = fixes E :: "'a \ 'a \ bool" begin inductive steps where Single: "steps [x]" | Cons: "steps (x # y # xs)" if "E x y" "steps (y # xs)" lemmas [intro] = steps.intros lemma steps_append: "steps (xs @ tl ys)" if "steps xs" "steps ys" "last xs = hd ys" using that by induction (auto 4 4 elim: steps.cases) lemma steps_append': "steps xs" if "steps as" "steps bs" "last as = hd bs" "as @ tl bs = xs" using steps_append that by blast coinductive run where "run (x ## y ## xs)" if "E x y" "run (y ## xs)" lemmas [intro] = run.intros lemma steps_appendD1: "steps xs" if "steps (xs @ ys)" "xs \ []" using that proof (induction xs) case Nil then show ?case by auto next case (Cons a xs) then show ?case by - (cases xs; auto elim: steps.cases) qed lemma steps_appendD2: "steps ys" if "steps (xs @ ys)" "ys \ []" using that by (induction xs) (auto elim: steps.cases) lemma steps_appendD3: "steps (xs @ [x]) \ E x y" if "steps (xs @ [x, y])" using that proof (induction xs) case Nil then show ?case by (auto elim!: steps.cases) next case prems: (Cons a xs) then show ?case by (cases xs) (auto elim: steps.cases) qed lemma steps_ConsD: "steps xs" if "steps (x # xs)" "xs \ []" using that by (auto elim: steps.cases) lemmas stepsD = steps_ConsD steps_appendD1 steps_appendD2 lemma steps_alt_induct[consumes 1, case_names Single Snoc]: assumes "steps x" "(\x. P [x])" "\y x xs. E y x \ steps (xs @ [y]) \ P (xs @ [y]) \ P (xs @ [y,x])" shows "P x" using assms(1) proof (induction rule: rev_induct) case Nil then show ?case by (auto elim: steps.cases) next case prems: (snoc x xs) then show ?case by (cases xs rule: rev_cases) (auto intro: assms(2,3) dest!: steps_appendD3) qed lemma steps_appendI: "steps (xs @ [x, y])" if "steps (xs @ [x])" "E x y" using that proof (induction xs) case Nil then show ?case by auto next case (Cons a xs) then show ?case by (cases xs; auto elim: steps.cases) qed lemma steps_append_single: assumes "steps xs" "E (last xs) x" "xs \ []" shows "steps (xs @ [x])" using assms(3,1,2) by (induction xs rule: list_nonempty_induct) (auto 4 4 elim: steps.cases) lemma extend_run: assumes "steps xs" "E (last xs) x" "run (x ## ys)" "xs \ []" shows "run (xs @- x ## ys)" using assms(4,1-3) by (induction xs rule: list_nonempty_induct) (auto 4 3 elim: steps.cases) lemma run_cycle: assumes "steps xs" "E (last xs) (hd xs)" "xs \ []" shows "run (cycle xs)" using assms proof (coinduction arbitrary: xs) case run then show ?case apply (rewrite at \cycle xs\ stream.collapse[symmetric]) apply (rewrite at \stl (cycle xs)\ stream.collapse[symmetric]) apply clarsimp apply (erule steps.cases) subgoal for x apply (rule conjI) apply (simp; fail) apply (rule disjI1) apply (inst_existentials xs) apply (simp, metis cycle_Cons[of x "[]", simplified]) by auto subgoal for x y xs' apply (rule conjI) apply (simp; fail) apply (rule disjI1) apply (inst_existentials "y # xs' @ [x]") using steps_append_single[of "y # xs'" x] apply (auto elim: steps.cases split: if_split_asm) apply (subst (2) cycle_Cons, simp) (* XXX Automate forward reasoning *) apply (subst cycle_Cons, simp) done done qed lemma run_stl: "run (stl xs)" if "run xs" using that by (auto elim: run.cases) lemma run_sdrop: "run (sdrop n xs)" if "run xs" using that by (induction n arbitrary: xs) (auto intro: run_stl) lemma run_reachable': assumes "run (x ## xs)" "E\<^sup>*\<^sup>* x\<^sub>0 x" shows "pred_stream (\ x. E\<^sup>*\<^sup>* x\<^sub>0 x) xs" using assms by (coinduction arbitrary: x xs) (auto 4 3 elim: run.cases) lemma run_reachable: assumes "run (x\<^sub>0 ## xs)" shows "pred_stream (\ x. E\<^sup>*\<^sup>* x\<^sub>0 x) xs" by (rule run_reachable'[OF assms]) blast lemma run_decomp: assumes "run (xs @- ys)" "xs \ []" shows "steps xs \ run ys \ E (last xs) (shd ys)" using assms(2,1) proof (induction xs rule: list_nonempty_induct) case (single x) then show ?case by (auto elim: run.cases) next case (cons x xs) then show ?case by (cases xs; auto 4 4 elim: run.cases) qed lemma steps_decomp: assumes "steps (xs @ ys)" "xs \ []" "ys \ []" shows "steps xs \ steps ys \ E (last xs) (hd ys)" using assms(2,1,3) proof (induction xs rule: list_nonempty_induct) case (single x) then show ?case by (auto elim: steps.cases) next case (cons x xs) then show ?case by (cases xs; auto 4 4 elim: steps.cases) qed lemma steps_rotate: assumes "steps (x # xs @ y # ys @ [x])" shows "steps (y # ys @ x # xs @ [y])" proof - from steps_decomp[of "x # xs" "y # ys @ [x]"] assms have "steps (x # xs)" "steps (y # ys @ [x])" "E (last (x # xs)) y" by auto then have "steps ((x # xs) @ [y])" by (blast intro: steps_append_single) from steps_append[OF \steps (y # ys @ [x])\ this] show ?thesis by auto qed lemma run_shift_coinduct[case_names run_shift, consumes 1]: assumes "R w" and "\ w. R w \ \ u v x y. w = u @- x ## y ## v \ steps (u @ [x]) \ E x y \ R (y ## v)" shows "run w" using assms(2)[OF \R w\] proof (coinduction arbitrary: w) case (run w) then obtain u v x y where "w = u @- x ## y ## v" "steps (u @ [x])" "E x y" "R (y ## v)" by auto then show ?case apply - apply (drule assms(2)) apply (cases u) apply force subgoal for z zs apply (cases zs) subgoal apply simp apply safe apply (force elim: steps.cases) subgoal for u' v' x' y' by (inst_existentials "x # u'") (cases u'; auto) done subgoal for a as apply simp apply safe apply (force elim: steps.cases) subgoal for u' v' x' y' apply (inst_existentials "a # as @ x # u'") using steps_append[of "a # as @ [x, y]" "u' @ [x']"] apply simp apply (drule steps_appendI[of "a # as" x, rotated]) by (cases u'; force elim: steps.cases)+ done done done qed lemma run_flat_coinduct[case_names run_shift, consumes 1]: assumes "R xss" and "\ xs ys xss. R (xs ## ys ## xss) \ xs \ [] \ steps xs \ E (last xs) (hd ys) \ R (ys ## xss)" shows "run (flat xss)" proof - obtain xs ys xss' where "xss = xs ## ys ## xss'" by (metis stream.collapse) with assms(2)[OF assms(1)[unfolded this]] show ?thesis proof (coinduction arbitrary: xs ys xss' xss rule: run_shift_coinduct) case (run_shift xs ys xss' xss) from run_shift show ?case apply (cases xss') apply clarify apply (drule assms(2)) apply (inst_existentials "butlast xs" "tl ys @- flat xss'" "last xs" "hd ys") apply (cases ys) apply (simp; fail) subgoal premises prems for x1 x2 z zs proof (cases "xs = []") case True with prems show ?thesis by auto next case False then have "xs = butlast xs @ [last xs]" by auto then have "butlast xs @- last xs ## tail = xs @- tail" for tail by (metis shift.simps(1,2) shift_append) with prems show ?thesis by simp qed apply (simp; fail) apply assumption subgoal for ws wss by (inst_existentials ys ws wss) (cases ys, auto) done qed qed lemma steps_non_empty[simp]: "\ steps []" by (auto elim: steps.cases) lemma steps_non_empty'[simp]: "xs \ []" if "steps xs" using that by auto (* XXX Generalize *) lemma steps_replicate: "steps (hd xs # concat (replicate n (tl xs)))" if "last xs = hd xs" "steps xs" "n > 0" using that proof (induction n) case 0 then show ?case by simp next case (Suc n) show ?case proof (cases n) case 0 with Suc.prems show ?thesis by (cases xs; auto) next case prems: (Suc nat) from Suc.prems have [simp]: "hd xs # tl xs @ ys = xs @ ys" for ys by (cases xs; auto) from Suc.prems have **: "tl xs @ ys = tl (xs @ ys)" for ys by (cases xs; auto) from prems Suc show ?thesis by (fastforce intro: steps_append') qed qed notation E ("_ \ _" [100, 100] 40) abbreviation reaches ("_ \* _" [100, 100] 40) where "reaches x y \ E\<^sup>*\<^sup>* x y" abbreviation reaches1 ("_ \\<^sup>+ _" [100, 100] 40) where "reaches1 x y \ E\<^sup>+\<^sup>+ x y" lemma steps_reaches: "hd xs \* last xs" if "steps xs" using that by (induction xs) auto lemma steps_reaches': "x \* y" if "steps xs" "hd xs = x" "last xs = y" using that steps_reaches by auto lemma reaches_steps: "\ xs. hd xs = x \ last xs = y \ steps xs" if "x \* y" using that apply (induction) apply force apply clarsimp subgoal for z xs by (inst_existentials "xs @ [z]", (cases xs; simp), auto intro: steps_append_single) done lemma reaches_steps_iff: "x \* y \ (\ xs. hd xs = x \ last xs = y \ steps xs)" using steps_reaches reaches_steps by fast lemma steps_reaches1: "x \\<^sup>+ y" if "steps (x # xs @ [y])" by (metis list.sel(1,3) rtranclp_into_tranclp2 snoc_eq_iff_butlast steps.cases steps_reaches that) lemma stepsI: "steps (x # xs)" if "x \ hd xs" "steps xs" using that by (cases xs) auto lemma reaches1_steps: "\ xs. steps (x # xs @ [y])" if "x \\<^sup>+ y" proof - from that obtain z where "x \ z" "z \* y" by atomize_elim (simp add: tranclpD) from reaches_steps[OF this(2)] obtain xs where *: "hd xs = z" "last xs = y" "steps xs" by auto then obtain xs' where [simp]: "xs = xs' @ [y]" by atomize_elim (auto 4 3 intro: append_butlast_last_id[symmetric]) with \x \ z\ * show ?thesis by (auto intro: stepsI) qed lemma reaches1_steps_iff: "x \\<^sup>+ y \ (\ xs. steps (x # xs @ [y]))" using steps_reaches1 reaches1_steps by fast lemma reaches1_reaches_iff1: "x \\<^sup>+ y \ (\ z. x \ z \ z \* y)" by (auto dest: tranclpD) lemma reaches1_reaches_iff2: "x \\<^sup>+ y \ (\ z. x \* z \ z \ y)" apply safe apply (metis Nitpick.rtranclp_unfold tranclp.cases) by auto lemma "x \\<^sup>+ z" if "x \* y" "y \\<^sup>+ z" using that by auto lemma "x \\<^sup>+ z" if "x \\<^sup>+ y" "y \* z" using that by auto lemma steps_append2: "steps (xs @ x # ys)" if "steps (xs @ [x])" "steps (x # ys)" using that by (auto dest: steps_append) lemma reaches1_steps_append: assumes "a \\<^sup>+ b" "steps xs" "hd xs = b" shows "\ ys. steps (a # ys @ xs)" using assms by (fastforce intro: steps_append' dest: reaches1_steps) lemma steps_last_step: "\ a. a \ last xs" if "steps xs" "length xs > 1" using that by induction auto lemmas graphI = steps.intros steps_append_single steps_reaches' stepsI end (* Graph Defs *) subsection \Graphs with a Start Node\ locale Graph_Start_Defs = Graph_Defs + fixes s\<^sub>0 :: 'a begin definition reachable where "reachable = E\<^sup>*\<^sup>* s\<^sub>0" lemma start_reachable[intro!, simp]: "reachable s\<^sub>0" unfolding reachable_def by auto lemma reachable_step: "reachable b" if "reachable a" "E a b" using that unfolding reachable_def by auto lemma reachable_reaches: "reachable b" if "reachable a" "a \* b" using that(2,1) by induction (auto intro: reachable_step) lemma reachable_steps_append: assumes "reachable a" "steps xs" "hd xs = a" "last xs = b" shows "reachable b" using assms by (auto intro: graphI reachable_reaches) lemmas steps_reachable = reachable_steps_append[of s\<^sub>0, simplified] lemma reachable_steps_elem: "reachable y" if "reachable x" "steps xs" "y \ set xs" "hd xs = x" proof - from \y \ set xs\ obtain as bs where [simp]: "xs = as @ y # bs" by (auto simp: in_set_conv_decomp) show ?thesis proof (cases "as = []") case True with that show ?thesis by simp next case False (* XXX *) from \steps xs\ have "steps (as @ [y])" by (auto intro: stepsD) with \as \ []\ \hd xs = x\ \reachable x\ show ?thesis by (auto 4 3 intro: reachable_reaches graphI) qed qed lemma reachable_steps: "\ xs. steps xs \ hd xs = s\<^sub>0 \ last xs = x" if "reachable x" using that unfolding reachable_def proof induction case base then show ?case by (inst_existentials "[s\<^sub>0]"; force) next case (step y z) - from step.IH guess xs by clarify + from step.IH obtain xs where "steps xs" "s\<^sub>0 = hd xs" "y = last xs" + by auto with step.hyps show ?case apply (inst_existentials "xs @ [z]") apply (force intro: graphI) by (cases xs; auto)+ qed lemma reachable_cycle_iff: "reachable x \ x \\<^sup>+ x \ (\ ws xs. steps (s\<^sub>0 # ws @ [x] @ xs @ [x]))" proof (safe, goal_cases) case (2 ws) then show ?case by (auto intro: steps_reachable stepsD) next case (3 ws xs) then show ?case by (auto intro: stepsD steps_reaches1) next case prems: 1 from \reachable x\ prems(2) have "s\<^sub>0 \\<^sup>+ x" unfolding reachable_def by auto with \x \\<^sup>+ x\ show ?case by (fastforce intro: steps_append' dest: reaches1_steps) qed lemma reachable_induct[consumes 1, case_names start step, induct pred: reachable]: assumes "reachable x" and "P s\<^sub>0" and "\ a b. reachable a \ P a \ a \ b \ P b" shows "P x" using assms(1) unfolding reachable_def by induction (auto intro: assms(2-)[unfolded reachable_def]) lemmas graphI_aggressive = tranclp_into_rtranclp rtranclp.rtrancl_into_rtrancl tranclp.trancl_into_trancl rtranclp_into_tranclp2 lemmas graphI_aggressive1 = graphI_aggressive steps_append' lemmas graphI_aggressive2 = graphI_aggressive stepsD steps_reaches1 steps_reachable lemmas graphD = reaches1_steps lemmas graphD_aggressive = tranclpD lemmas graph_startI = reachable_reaches start_reachable end (* Graph Start Defs *) subsection \Subgraphs\ subsubsection \Edge-induced Subgraphs\ locale Subgraph_Defs = G: Graph_Defs + fixes E' :: "'a \ 'a \ bool" begin sublocale G': Graph_Defs E' . end (* Subgraph Defs *) locale Subgraph_Start_Defs = G: Graph_Start_Defs + fixes E' :: "'a \ 'a \ bool" begin sublocale G': Graph_Start_Defs E' s\<^sub>0 . end (* Subgraph Start Defs *) locale Subgraph = Subgraph_Defs + assumes subgraph[intro]: "E' a b \ E a b" begin lemma non_subgraph_cycle_decomp: "\ c d. G.reaches a c \ E c d \ \ E' c d \ G.reaches d b" if "G.reaches1 a b" "\ G'.reaches1 a b" for a b using that proof induction case (base y) then show ?case by auto next case (step y z) show ?case proof (cases "E' y z") case True with step have "\ G'.reaches1 a y" by (auto intro: tranclp.trancl_into_trancl) (* XXX *) with step obtain c d where "G.reaches a c" "E c d" "\ E' c d" "G.reaches d y" by auto with \E' y z\ show ?thesis by (blast intro: rtranclp.rtrancl_into_rtrancl) (* XXX *) next case False with step show ?thesis by (intro exI conjI) auto qed qed lemma reaches: "G.reaches a b" if "G'.reaches a b" using that by induction (auto intro: rtranclp.intros(2)) lemma reaches1: "G.reaches1 a b" if "G'.reaches1 a b" using that by induction (auto intro: tranclp.intros(2)) end (* Subgraph *) locale Subgraph_Start = Subgraph_Start_Defs + Subgraph begin lemma reachable_subgraph[intro]: "G.reachable b" if \G.reachable a\ \G'.reaches a b\ for a b using that by (auto intro: G.graph_startI mono_rtranclp[rule_format, of E']) lemma reachable: "G.reachable x" if "G'.reachable x" using that by (fastforce simp: G.reachable_def G'.reachable_def) end (* Subgraph Start *) subsubsection \Node-induced Subgraphs\ locale Subgraph_Node_Defs = Graph_Defs + fixes V :: "'a \ bool" begin definition E' where "E' x y \ E x y \ V x \ V y" sublocale Subgraph E E' by standard (auto simp: E'_def) lemma subgraph': "E' x y" if "E x y" "V x" "V y" using that unfolding E'_def by auto lemma E'_V1: "V x" if "E' x y" using that unfolding E'_def by auto lemma E'_V2: "V y" if "E' x y" using that unfolding E'_def by auto lemma G'_reaches_V: "V y" if "G'.reaches x y" "V x" using that by (cases) (auto intro: E'_V2) lemma G'_steps_V_all: "list_all V xs" if "G'.steps xs" "V (hd xs)" using that by induction (auto intro: E'_V2) lemma G'_steps_V_last: "V (last xs)" if "G'.steps xs" "V (hd xs)" using that by induction (auto dest: E'_V2) lemmas subgraphI = E'_V1 E'_V2 G'_reaches_V lemmas subgraphD = E'_V1 E'_V2 G'_reaches_V end (* Subgraph Node *) locale Subgraph_Node_Defs_Notation = Subgraph_Node_Defs begin no_notation E ("_ \ _" [100, 100] 40) notation E' ("_ \ _" [100, 100] 40) no_notation reaches ("_ \* _" [100, 100] 40) notation G'.reaches ("_ \* _" [100, 100] 40) no_notation reaches1 ("_ \\<^sup>+ _" [100, 100] 40) notation G'.reaches1 ("_ \\<^sup>+ _" [100, 100] 40) end (* Subgraph_Node_Defs_Notation *) subsubsection \The Reachable Subgraph\ context Graph_Start_Defs begin interpretation Subgraph_Node_Defs_Notation E reachable . sublocale reachable_subgraph: Subgraph_Node_Defs E reachable . lemma reachable_supgraph: "x \ y" if "E x y" "reachable x" using that unfolding E'_def by (auto intro: graph_startI) lemma reachable_reaches_equiv: "reaches x y \ x \* y" if "reachable x" for x y apply standard subgoal premises prems using prems \reachable x\ by induction (auto dest: reachable_supgraph intro: graph_startI graphI_aggressive) subgoal premises prems using prems \reachable x\ by induction (auto dest: subgraph) done lemma reachable_reaches1_equiv: "reaches1 x y \ x \\<^sup>+ y" if "reachable x" for x y apply standard subgoal premises prems using prems \reachable x\ by induction (auto dest: reachable_supgraph intro: graph_startI graphI_aggressive) subgoal premises prems using prems \reachable x\ by induction (auto dest: subgraph) done lemma reachable_steps_equiv: "steps (x # xs) \ G'.steps (x # xs)" if "reachable x" apply standard subgoal premises prems using prems \reachable x\ by (induction "x # xs" arbitrary: x xs) (auto dest: reachable_supgraph intro: graph_startI) subgoal premises prems using prems by induction auto done end (* Graph Start Defs *) subsection \Bundles\ bundle graph_automation begin lemmas [intro] = Graph_Defs.graphI Graph_Start_Defs.graph_startI lemmas [dest] = Graph_Start_Defs.graphD end (* Bundle *) bundle reaches_steps_iff = Graph_Defs.reaches1_steps_iff [iff] Graph_Defs.reaches_steps_iff [iff] bundle graph_automation_aggressive begin unbundle graph_automation lemmas [intro] = Graph_Start_Defs.graphI_aggressive lemmas [dest] = Graph_Start_Defs.graphD_aggressive end (* Bundle *) bundle subgraph_automation begin unbundle graph_automation lemmas [intro] = Subgraph_Node_Defs.subgraphI lemmas [dest] = Subgraph_Node_Defs.subgraphD end (* Bundle *) subsection \Simulations and Bisimulations\ locale Graph_Invariant = Graph_Defs + fixes P :: "'a \ bool" assumes invariant: "P a \ a \ b \ P b" begin lemma invariant_steps: "list_all P as" if "steps (a # as)" "P a" using that by (induction "a # as" arbitrary: as a) (auto intro: invariant) lemma invariant_reaches: "P b" if "a \* b" "P a" using that by (induction; blast intro: invariant) lemma invariant_run: assumes run: "run (x ## xs)" and P: "P x" shows "pred_stream P (x ## xs)" using run P by (coinduction arbitrary: x xs) (auto 4 3 elim: invariant run.cases) end (* Graph Invariant *) locale Graph_Invariants = Graph_Defs + fixes P Q :: "'a \ bool" assumes invariant: "P a \ a \ b \ Q b" and Q_P: "Q a \ P a" begin sublocale Pre: Graph_Invariant E P by standard (blast intro: invariant Q_P) sublocale Post: Graph_Invariant E Q by standard (blast intro: invariant Q_P) lemma invariant_steps: "list_all Q as" if "steps (a # as)" "P a" using that by (induction "a # as" arbitrary: as a) (auto intro: invariant Q_P) lemma invariant_run: assumes run: "run (x ## xs)" and P: "P x" shows "pred_stream Q xs" using run P by (coinduction arbitrary: x xs) (auto 4 4 elim: invariant run.cases intro: Q_P) lemma invariant_reaches1: "Q b" if "a \\<^sup>+ b" "P a" using that by (induction; blast intro: invariant Q_P) end (* Graph Invariants *) locale Graph_Invariant_Start = Graph_Start_Defs + Graph_Invariant + assumes P_s\<^sub>0: "P s\<^sub>0" begin lemma invariant_steps: "list_all P as" if "steps (s\<^sub>0 # as)" using that P_s\<^sub>0 by (rule invariant_steps) lemma invariant_reaches: "P b" if "s\<^sub>0 \* b" using invariant_reaches[OF that P_s\<^sub>0] . lemmas invariant_run = invariant_run[OF _ P_s\<^sub>0] end (* Graph Invariant Start *) locale Graph_Invariant_Strong = Graph_Defs + fixes P :: "'a \ bool" assumes invariant: "a \ b \ P b" begin sublocale inv: Graph_Invariant by standard (rule invariant) lemma P_invariant_steps: "list_all P as" if "steps (a # as)" using that by (induction "a # as" arbitrary: as a) (auto intro: invariant) lemma steps_last_invariant: "P (last xs)" if "steps (x # xs)" "xs \ []" using steps_last_step[of "x # xs"] that by (auto intro: invariant) lemmas invariant_reaches = inv.invariant_reaches lemma invariant_reaches1: "P b" if "a \\<^sup>+ b" using that by (induction; blast intro: invariant) end (* Graph Invariant *) locale Simulation_Defs = fixes A :: "'a \ 'a \ bool" and B :: "'b \ 'b \ bool" and sim :: "'a \ 'b \ bool" (infixr "\" 60) begin sublocale A: Graph_Defs A . sublocale B: Graph_Defs B . end (* Simulation Defs *) locale Simulation = Simulation_Defs + assumes A_B_step: "\ a b a'. A a b \ a \ a' \ (\ b'. B a' b' \ b \ b')" begin lemma simulation_reaches: "\ b'. B\<^sup>*\<^sup>* b b' \ a' \ b'" if "A\<^sup>*\<^sup>* a a'" "a \ b" using that by (induction rule: rtranclp_induct) (auto intro: rtranclp.intros(2) dest: A_B_step) lemma simulation_reaches1: "\ b'. B\<^sup>+\<^sup>+ b b' \ a' \ b'" if "A\<^sup>+\<^sup>+ a a'" "a \ b" using that by (induction rule: tranclp_induct) (auto 4 3 intro: tranclp.intros(2) dest: A_B_step) lemma simulation_steps: "\ bs. B.steps (b # bs) \ list_all2 (\ a b. a \ b) as bs" if "A.steps (a # as)" "a \ b" using that apply (induction "a # as" arbitrary: a b as) apply force apply (frule A_B_step, auto) done lemma simulation_run: "\ ys. B.run (y ## ys) \ stream_all2 (\) xs ys" if "A.run (x ## xs)" "x \ y" proof - let ?ys = "sscan (\ a' b. SOME b'. B b b' \ a' \ b') xs y" have "B.run (y ## ?ys)" using that by (coinduction arbitrary: x y xs) (force dest!: someI_ex A_B_step elim: A.run.cases) moreover have "stream_all2 (\) xs ?ys" using that by (coinduction arbitrary: x y xs) (force dest!: someI_ex A_B_step elim: A.run.cases) ultimately show ?thesis by blast qed end (* Simulation *) lemma (in Subgraph) Subgraph_Simulation: "Simulation E' E (=)" by standard auto locale Simulation_Invariant = Simulation_Defs + fixes PA :: "'a \ bool" and PB :: "'b \ bool" assumes A_B_step: "\ a b a'. A a b \ PA a \ PB a' \ a \ a' \ (\ b'. B a' b' \ b \ b')" assumes A_invariant[intro]: "\ a b. PA a \ A a b \ PA b" assumes B_invariant[intro]: "\ a b. PB a \ B a b \ PB b" begin definition "equiv' \ \ a b. a \ b \ PA a \ PB b" sublocale Simulation A B equiv' by standard (auto dest: A_B_step simp: equiv'_def) sublocale PA_invariant: Graph_Invariant A PA by standard blast lemma simulation_reaches: "\ b'. B\<^sup>*\<^sup>* b b' \ a' \ b' \ PA a' \ PB b'" if "A\<^sup>*\<^sup>* a a'" "a \ b" "PA a" "PB b" using simulation_reaches[of a a' b] that unfolding equiv'_def by simp lemma simulation_steps: "\ bs. B.steps (b # bs) \ list_all2 (\ a b. a \ b \ PA a \ PB b) as bs" if "A.steps (a # as)" "a \ b" "PA a" "PB b" using simulation_steps[of a as b] that unfolding equiv'_def by simp lemma simulation_steps': "\ bs. B.steps (b # bs) \ list_all2 (\ a b. a \ b) as bs \ list_all PA as \ list_all PB bs" if "A.steps (a # as)" "a \ b" "PA a" "PB b" using simulation_steps[OF that] by (metis (mono_tags, lifting) list_all2_conv_all_nth list_all_length) context fixes f assumes eq: "a \ b \ b = f a" begin lemma simulation_steps'_map: "\ bs. B.steps (b # bs) \ bs = map f as \ list_all2 (\ a b. a \ b) as bs \ list_all PA as \ list_all PB bs" if "A.steps (a # as)" "a \ b" "PA a" "PB b" proof - - from simulation_steps'[OF that] guess bs by clarify - note guessed = this - from this(2) have "bs = map f as" + from simulation_steps'[OF that] + obtain bs where bs: "B.steps (b # bs)" "list_all2 (\) as bs" "list_all PA as" "list_all PB bs" + by auto + from bs(2) have "bs = map f as" by (induction; simp add: eq) - with guessed show ?thesis + with bs show ?thesis by auto qed end (* Context for Equality Relation *) end (* Simulation Invariant *) locale Simulation_Invariants = Simulation_Defs + fixes PA QA :: "'a \ bool" and PB QB :: "'b \ bool" assumes A_B_step: "\ a b a'. A a b \ PA a \ PB a' \ a \ a' \ (\ b'. B a' b' \ b \ b')" assumes A_invariant[intro]: "\ a b. PA a \ A a b \ QA b" assumes B_invariant[intro]: "\ a b. PB a \ B a b \ QB b" assumes PA_QA[intro]: "\ a. QA a \ PA a" and PB_QB[intro]: "\ a. QB a \ PB a" begin sublocale Pre: Simulation_Invariant A B "(\)" PA PB by standard (auto intro: A_B_step) sublocale Post: Simulation_Invariant A B "(\)" QA QB by standard (auto intro: A_B_step) sublocale A_invs: Graph_Invariants A PA QA by standard auto sublocale B_invs: Graph_Invariants B PB QB by standard auto lemma simulation_reaches1: "\ b2. B.reaches1 b1 b2 \ a2 \ b2 \ QB b2" if "A.reaches1 a1 a2" "a1 \ b1" "PA a1" "PB b1" using that by - (drule Pre.simulation_reaches1, auto intro: B_invs.invariant_reaches1 simp: Pre.equiv'_def) lemma reaches1_unique: assumes unique: "\ b2. a \ b2 \ QB b2 \ b2 = b" and that: "A.reaches1 a a" "a \ b" "PA a" "PB b" shows "B.reaches1 b b" using that by (auto dest: unique simulation_reaches1) end (* Simualation Invariants *) locale Bisimulation = Simulation_Defs + assumes A_B_step: "\ a b a'. A a b \ a \ a' \ (\ b'. B a' b' \ b \ b')" assumes B_A_step: "\ a a' b'. B a' b' \ a \ a' \ (\ b. A a b \ b \ b')" begin sublocale A_B: Simulation A B "(\)" by standard (rule A_B_step) sublocale B_A: Simulation B A "\ x y. y \ x" by standard (rule B_A_step) lemma A_B_reaches: "\ b'. B\<^sup>*\<^sup>* b b' \ a' \ b'" if "A\<^sup>*\<^sup>* a a'" "a \ b" using A_B.simulation_reaches[OF that] . lemma B_A_reaches: "\ b'. A\<^sup>*\<^sup>* b b' \ b' \ a'" if "B\<^sup>*\<^sup>* a a'" "b \ a" using B_A.simulation_reaches[OF that] . end (* Bisim *) locale Bisimulation_Invariant = Simulation_Defs + fixes PA :: "'a \ bool" and PB :: "'b \ bool" assumes A_B_step: "\ a b a'. A a b \ a \ a' \ PA a \ PB a' \ (\ b'. B a' b' \ b \ b')" assumes B_A_step: "\ a a' b'. B a' b' \ a \ a' \ PA a \ PB a' \ (\ b. A a b \ b \ b')" assumes A_invariant[intro]: "\ a b. PA a \ A a b \ PA b" assumes B_invariant[intro]: "\ a b. PB a \ B a b \ PB b" begin sublocale PA_invariant: Graph_Invariant A PA by standard blast sublocale PB_invariant: Graph_Invariant B PB by standard blast lemmas B_steps_invariant[intro] = PB_invariant.invariant_reaches definition "equiv' \ \ a b. a \ b \ PA a \ PB b" sublocale bisim: Bisimulation A B equiv' by standard (clarsimp simp add: equiv'_def, frule A_B_step B_A_step, assumption; auto)+ sublocale A_B: Simulation_Invariant A B "(\)" PA PB by (standard; blast intro: A_B_step B_A_step) sublocale B_A: Simulation_Invariant B A "\ x y. y \ x" PB PA by (standard; blast intro: A_B_step B_A_step) context fixes f assumes eq: "a \ b \ b = f a" and inj: "\ a b. PB (f a) \ PA b \ f a = f b \ a = b" begin lemma list_all2_inj_map_eq: "as = bs" if "list_all2 (\a b. a = f b) (map f as) bs" "list_all PB (map f as)" "list_all PA bs" using that inj by (induction "map f as" bs arbitrary: as rule: list_all2_induct) (auto simp: inj_on_def) lemma steps_map_equiv: "A.steps (a # as) \ B.steps (b # map f as)" if "a \ b" "PA a" "PB b" using A_B.simulation_steps'_map[of f a as b] B_A.simulation_steps'[of b "map f as" a] that eq by (auto dest: list_all2_inj_map_eq) lemma steps_map: "\ as. bs = map f as" if "B.steps (f a # bs)" "PA a" "PB (f a)" proof - have "a \ f a" unfolding eq .. - from B_A.simulation_steps'[OF that(1) this \PB _\ \PA _\] guess as by clarify - from this(2) show ?thesis + from B_A.simulation_steps'[OF that(1) this \PB _\ \PA _\] + obtain as where "list_all2 (\a b. b \ a) bs as" by auto + then show ?thesis unfolding eq by (inst_existentials as, induction rule: list_all2_induct, auto) qed lemma reaches_equiv: "A.reaches a a' \ B.reaches (f a) (f a')" if "PA a" "PB (f a)" apply safe apply (drule A_B.simulation_reaches[of a a' "f a"]; simp add: eq that) apply (drule B_A.simulation_reaches) defer apply (rule that | clarsimp simp: eq | metis inj)+ done end (* Context for Equality Relation *) lemma equiv'_D: "a \ b" if "A_B.equiv' a b" using that unfolding A_B.equiv'_def by auto lemma equiv'_rotate_1: "B_A.equiv' b a" if "A_B.equiv' a b" using that by (auto simp: B_A.equiv'_def A_B.equiv'_def) lemma equiv'_rotate_2: "A_B.equiv' a b" if "B_A.equiv' b a" using that by (auto simp: B_A.equiv'_def A_B.equiv'_def) lemma stream_all2_equiv'_D: "stream_all2 (\) xs ys" if "stream_all2 A_B.equiv' xs ys" using stream_all2_weaken[OF that equiv'_D] by fast lemma stream_all2_equiv'_D2: "stream_all2 B_A.equiv' ys xs \ stream_all2 (\)\\ ys xs" by (coinduction arbitrary: xs ys) (auto simp: B_A.equiv'_def) lemma stream_all2_rotate_1: "stream_all2 B_A.equiv' ys xs \ stream_all2 A_B.equiv' xs ys" by (coinduction arbitrary: xs ys) (auto simp: B_A.equiv'_def A_B.equiv'_def) lemma stream_all2_rotate_2: "stream_all2 A_B.equiv' xs ys \ stream_all2 B_A.equiv' ys xs" by (coinduction arbitrary: xs ys) (auto simp: B_A.equiv'_def A_B.equiv'_def) end (* Bisim Invariant *) locale Bisimulation_Invariants = Simulation_Defs + fixes PA QA :: "'a \ bool" and PB QB :: "'b \ bool" assumes A_B_step: "\ a b a'. A a b \ a \ a' \ PA a \ PB a' \ (\ b'. B a' b' \ b \ b')" assumes B_A_step: "\ a a' b'. B a' b' \ a \ a' \ PA a \ PB a' \ (\ b. A a b \ b \ b')" assumes A_invariant[intro]: "\ a b. PA a \ A a b \ QA b" assumes B_invariant[intro]: "\ a b. PB a \ B a b \ QB b" assumes PA_QA[intro]: "\ a. QA a \ PA a" and PB_QB[intro]: "\ a. QB a \ PB a" begin sublocale PA_invariant: Graph_Invariant A PA by standard blast sublocale PB_invariant: Graph_Invariant B PB by standard blast sublocale QA_invariant: Graph_Invariant A QA by standard blast sublocale QB_invariant: Graph_Invariant B QB by standard blast sublocale Pre_Bisim: Bisimulation_Invariant A B "(\)" PA PB by standard (auto intro: A_B_step B_A_step) sublocale Post_Bisim: Bisimulation_Invariant A B "(\)" QA QB by standard (auto intro: A_B_step B_A_step) sublocale A_B: Simulation_Invariants A B "(\)" PA QA PB QB by standard (blast intro: A_B_step)+ sublocale B_A: Simulation_Invariants B A "\ x y. y \ x" PB QB PA QA by standard (blast intro: B_A_step)+ context fixes f assumes eq[simp]: "a \ b \ b = f a" and inj: "\ a b. QB (f a) \ QA b \ f a = f b \ a = b" begin lemmas list_all2_inj_map_eq = Post_Bisim.list_all2_inj_map_eq[OF eq inj] lemmas steps_map_equiv' = Post_Bisim.steps_map_equiv[OF eq inj] lemma list_all2_inj_map_eq': "as = bs" if "list_all2 (\a b. a = f b) (map f as) bs" "list_all QB (map f as)" "list_all QA bs" using that by (rule list_all2_inj_map_eq) lemma steps_map_equiv: "A.steps (a # as) \ B.steps (b # map f as)" if "a \ b" "PA a" "PB b" proof assume "A.steps (a # as)" then show "B.steps (b # map f as)" proof cases case Single then show ?thesis by auto next case prems: (Cons a' xs) from A_B_step[OF \A a a'\ \a \ b\ \PA a\ \PB b\] obtain b' where "B b b'" "a' \ b'" by auto with steps_map_equiv'[OF \a' \ b'\, of xs] prems that show ?thesis by auto qed next assume "B.steps (b # map f as)" then show "A.steps (a # as)" proof cases case Single then show ?thesis by auto next case prems: (Cons b' xs) from B_A_step[OF \B b b'\ \a \ b\ \PA a\ \PB b\] obtain a' where "A a a'" "a' \ b'" by auto with that prems have "QA a'" "QB b'" by auto with \A a a'\ \a' \ b'\ steps_map_equiv'[OF \a' \ b'\, of "tl as"] prems that show ?thesis apply clarsimp subgoal for z zs using inj[rule_format, of z a'] by auto done qed qed lemma steps_map: "\ as. bs = map f as" if "B.steps (f a # bs)" "PA a" "PB (f a)" using that proof cases case Single then show ?thesis by simp next case prems: (Cons b' xs) from B_A_step[OF \B _ b'\ _ \PA a\ \PB (f a)\] obtain a' where "A a a'" "a' \ b'" by auto with that prems have "QA a'" "QB b'" by auto with Post_Bisim.steps_map[OF eq inj, of a' xs] prems \a' \ b'\ obtain ys where "xs = map f ys" by auto with \bs = _\ \a' \ b'\ show ?thesis by (inst_existentials "a' # ys") auto qed text \ @{thm Post_Bisim.reaches_equiv} cannot be lifted directly: injectivity cannot be applied for the reflexive case. \ lemma reaches1_equiv: "A.reaches1 a a' \ B.reaches1 (f a) (f a')" if "PA a" "PB (f a)" proof safe assume "A.reaches1 a a'" then obtain a'' where prems: "A a a''" "A.reaches a'' a'" including graph_automation_aggressive by blast from A_B_step[OF \A a _\ _ that] obtain b where "B (f a) b" "a'' \ b" by auto with that prems have "QA a''" "QB b" by auto with Post_Bisim.reaches_equiv[OF eq inj, of a'' a'] prems \B (f a) b\ \a'' \ b\ show "B.reaches1 (f a) (f a')" by auto next assume "B.reaches1 (f a) (f a')" then obtain b where prems: "B (f a) b" "B.reaches b (f a')" including graph_automation_aggressive by blast from B_A_step[OF \B _ b\ _ \PA a\ \PB (f a)\] obtain a'' where "A a a''" "a'' \ b" by auto with that prems have "QA a''" "QB b" by auto with Post_Bisim.reaches_equiv[OF eq inj, of a'' a'] prems \A a a''\ \a'' \ b\ show "A.reaches1 a a'" by auto qed end (* Context for Equality Relation *) end (* Bisim Invariant *) lemma Bisimulation_Invariant_composition: assumes "Bisimulation_Invariant A B sim1 PA PB" "Bisimulation_Invariant B C sim2 PB PC" shows "Bisimulation_Invariant A C (\ a c. \ b. PB b \ sim1 a b \ sim2 b c) PA PC" proof - interpret A: Bisimulation_Invariant A B sim1 PA PB by (rule assms(1)) interpret B: Bisimulation_Invariant B C sim2 PB PC by (rule assms(2)) show ?thesis by (standard; (blast dest: A.A_B_step B.A_B_step | blast dest: A.B_A_step B.B_A_step)) qed lemma Bisimulation_Invariant_filter: assumes "Bisimulation_Invariant A B sim PA PB" "\ a b. sim a b \ PA a \ PB b \ FA a \ FB b" "\ a b. A a b \ FA b \ A' a b" "\ a b. B a b \ FB b \ B' a b" shows "Bisimulation_Invariant A' B' sim PA PB" proof - interpret Bisimulation_Invariant A B sim PA PB by (rule assms(1)) have unfold: "A' = (\ a b. A a b \ FA b)" "B' = (\ a b. B a b \ FB b)" using assms(3,4) by auto show ?thesis unfolding unfold apply standard using assms(2) apply (blast dest: A_B_step) using assms(2) apply (blast dest: B_A_step) by blast+ qed lemma Bisimulation_Invariants_filter: assumes "Bisimulation_Invariants A B sim PA QA PB QB" "\ a b. QA a \ QB b \ FA a \ FB b" "\ a b. A a b \ FA b \ A' a b" "\ a b. B a b \ FB b \ B' a b" shows "Bisimulation_Invariants A' B' sim PA QA PB QB" proof - interpret Bisimulation_Invariants A B sim PA QA PB QB by (rule assms(1)) have unfold: "A' = (\ a b. A a b \ FA b)" "B' = (\ a b. B a b \ FB b)" using assms(3,4) by auto show ?thesis unfolding unfold apply standard using assms(2) apply (blast dest: A_B_step) using assms(2) apply (blast dest: B_A_step) by blast+ qed lemma Bisimulation_Invariants_composition: assumes "Bisimulation_Invariants A B sim1 PA QA PB QB" "Bisimulation_Invariants B C sim2 PB QB PC QC" shows "Bisimulation_Invariants A C (\ a c. \ b. PB b \ sim1 a b \ sim2 b c) PA QA PC QC" proof - interpret A: Bisimulation_Invariants A B sim1 PA QA PB QB by (rule assms(1)) interpret B: Bisimulation_Invariants B C sim2 PB QB PC QC by (rule assms(2)) show ?thesis by (standard; (blast dest: A.A_B_step B.A_B_step | blast dest: A.B_A_step B.B_A_step)) qed lemma Bisimulation_Invariant_Invariants_composition: assumes "Bisimulation_Invariant A B sim1 PA PB" "Bisimulation_Invariants B C sim2 PB QB PC QC" shows "Bisimulation_Invariants A C (\ a c. \ b. PB b \ sim1 a b \ sim2 b c) PA PA PC QC" proof - interpret Bisimulation_Invariant A B sim1 PA PB by (rule assms(1)) interpret B: Bisimulation_Invariants B C sim2 PB QB PC QC by (rule assms(2)) interpret A: Bisimulation_Invariants A B sim1 PA PA PB QB by (standard; blast intro: A_B_step B_A_step)+ show ?thesis by (standard; (blast dest: A.A_B_step B.A_B_step | blast dest: A.B_A_step B.B_A_step)) qed lemma Bisimulation_Invariant_Bisimulation_Invariants: assumes "Bisimulation_Invariant A B sim PA PB" shows "Bisimulation_Invariants A B sim PA PA PB PB" proof - interpret Bisimulation_Invariant A B sim PA PB by (rule assms) show ?thesis by (standard; blast intro: A_B_step B_A_step) qed lemma Bisimulation_Invariant_strengthen_post: assumes "Bisimulation_Invariant A B sim PA PB" "\ a b. PA' a \ PA b \ A a b \ PA' b" "\ a. PA' a \ PA a" shows "Bisimulation_Invariant A B sim PA' PB" proof - interpret Bisimulation_Invariant A B sim PA PB by (rule assms) show ?thesis by (standard; blast intro: A_B_step B_A_step assms) qed lemma Bisimulation_Invariant_strengthen_post': assumes "Bisimulation_Invariant A B sim PA PB" "\ a b. PB' a \ PB b \ B a b \ PB' b" "\ a. PB' a \ PB a" shows "Bisimulation_Invariant A B sim PA PB'" proof - interpret Bisimulation_Invariant A B sim PA PB by (rule assms) show ?thesis by (standard; blast intro: A_B_step B_A_step assms) qed lemma Simulation_Invariant_strengthen_post: assumes "Simulation_Invariant A B sim PA PB" "\ a b. PA a \ PA b \ A a b \ PA' b" "\ a. PA' a \ PA a" shows "Simulation_Invariant A B sim PA' PB" proof - interpret Simulation_Invariant A B sim PA PB by (rule assms) show ?thesis by (standard; blast intro: A_B_step assms) qed lemma Simulation_Invariant_strengthen_post': assumes "Simulation_Invariant A B sim PA PB" "\ a b. PB a \ PB b \ B a b \ PB' b" "\ a. PB' a \ PB a" shows "Simulation_Invariant A B sim PA PB'" proof - interpret Simulation_Invariant A B sim PA PB by (rule assms) show ?thesis by (standard; blast intro: A_B_step assms) qed lemma Simulation_Invariants_strengthen_post: assumes "Simulation_Invariants A B sim PA QA PB QB" "\ a b. PA a \ QA b \ A a b \ QA' b" "\ a. QA' a \ QA a" shows "Simulation_Invariants A B sim PA QA' PB QB" proof - interpret Simulation_Invariants A B sim PA QA PB QB by (rule assms) show ?thesis by (standard; blast intro: A_B_step assms) qed lemma Simulation_Invariants_strengthen_post': assumes "Simulation_Invariants A B sim PA QA PB QB" "\ a b. PB a \ QB b \ B a b \ QB' b" "\ a. QB' a \ QB a" shows "Simulation_Invariants A B sim PA QA PB QB'" proof - interpret Simulation_Invariants A B sim PA QA PB QB by (rule assms) show ?thesis by (standard; blast intro: A_B_step assms) qed lemma Bisimulation_Invariant_sim_replace: assumes "Bisimulation_Invariant A B sim PA PB" and "\ a b. PA a \ PB b \ sim a b \ sim' a b" shows "Bisimulation_Invariant A B sim' PA PB" proof - interpret Bisimulation_Invariant A B sim PA PB by (rule assms(1)) show ?thesis apply standard using assms(2) apply (blast dest: A_B_step) using assms(2) apply (blast dest: B_A_step) by blast+ qed end (* Theory *) diff --git a/thys/Probabilistic_Timed_Automata/library/More_List.thy b/thys/Probabilistic_Timed_Automata/library/More_List.thy --- a/thys/Probabilistic_Timed_Automata/library/More_List.thy +++ b/thys/Probabilistic_Timed_Automata/library/More_List.thy @@ -1,301 +1,308 @@ (* Author: Simon Wimmer *) theory More_List imports Main Instantiate_Existentials begin section \First and Last Elements of Lists\ lemma (in -) hd_butlast_last_id: "hd xs # tl (butlast xs) @ [last xs] = xs" if "length xs > 1" using that by (cases xs) auto section \@{term list_all}\ lemma (in -) list_all_map: assumes inv: "\ x. P x \ \ y. f y = x" and all: "list_all P as" shows "\ as'. map f as' = as" using all apply (induction as) apply (auto dest!: inv) subgoal for as' a by (inst_existentials "a # as'") simp done section \@{term list_all2}\ lemma list_all2_op_map_iff: "list_all2 (\ a b. b = f a) xs ys \ map f xs = ys" unfolding list_all2_iff proof (induction xs arbitrary: ys) case Nil then show ?case by auto next case (Cons a xs ys) then show ?case by (cases ys) auto qed lemma list_all2_last: "R (last xs) (last ys)" if "list_all2 R xs ys" "xs \ []" using that unfolding list_all2_iff proof (induction xs arbitrary: ys) case Nil then show ?case by simp next case (Cons a xs ys) then show ?case by (cases ys) auto qed lemma list_all2_set1: "\x\set xs. \xa\set as. P x xa" if "list_all2 P xs as" using that proof (induction xs arbitrary: as) case Nil then show ?case by auto next case (Cons a xs as) then show ?case by (cases as) auto qed lemma list_all2_swap: "list_all2 P xs ys \ list_all2 (\ x y. P y x) ys xs" unfolding list_all2_iff by (fastforce simp: in_set_zip)+ lemma list_all2_set2: "\x\set as. \xa\set xs. P xa x" if "list_all2 P xs as" using that by - (rule list_all2_set1, subst (asm) list_all2_swap) section \Distinct lists\ (* XXX Duplication with Floyd_Warshall.thy *) lemma distinct_length_le:"finite s \ set xs \ s \ distinct xs \ length xs \ card s" by (metis card_mono distinct_card) section \@{term filter}\ lemma filter_eq_appendD: "\ xs' ys'. filter P xs' = xs \ filter P ys' = ys \ as = xs' @ ys'" if "filter P as = xs @ ys" using that proof (induction xs arbitrary: as) case Nil then show ?case by (inst_existentials "[] :: 'a list" as) auto next case (Cons a xs) from filter_eq_ConsD[OF Cons.prems[simplified]] obtain us vs where "as = us @ a # vs" "\u\set us. \ P u" "P a" "filter P vs = xs @ ys" by auto moreover from Cons.IH[OF \_ = xs @ ys\] obtain xs' ys where "filter P xs' = xs" "vs = xs' @ ys" by auto ultimately show ?case by (inst_existentials "us @ [a] @ xs'" ys) auto qed lemma list_all2_elem_filter: assumes "list_all2 P xs us" "x \ set xs" shows "length (filter (P x) us) \ 1" using assms by (induction xs arbitrary: us) (auto simp: list_all2_Cons1) lemma list_all2_replicate_elem_filter: assumes "list_all2 P (concat (replicate n xs)) ys" "x \ set xs" shows "length (filter (P x) ys) \ n" using assms by (induction n arbitrary: ys; fastforce dest: list_all2_elem_filter simp: list_all2_append1) section \Sublists\ lemma nths_split: "nths xs (A \ B) = nths xs A @ nths xs B" if "\ i \ A. \ j \ B. i < j" using that proof (induction xs arbitrary: A B) case Nil then show ?case by simp next case (Cons a xs) let ?A = "{j. Suc j \ A}" and ?B = "{j. Suc j \ B}" from Cons.prems have *: "\i\?A. \a\?B. i < a" by auto have [simp]: "{j. Suc j \ A \ Suc j \ B} = ?A \ ?B" by auto show ?case unfolding nths_Cons proof (clarsimp, safe, goal_cases) case 2 with Cons.prems have "A = {}" by auto with Cons.IH[OF *] show ?case by auto qed (use Cons.prems Cons.IH[OF *] in auto) qed lemma nths_nth: "nths xs {i} = [xs ! i]" if "i < length xs" using that proof (induction xs arbitrary: i) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases i) (auto simp: nths_Cons) qed lemma nths_shift: "nths (xs @ ys) S = nths ys {x - length xs | x. x \ S}" if "\ i \ S. length xs \ i" using that proof (induction xs arbitrary: S) case Nil then show ?case by auto next case (Cons a xs) have [simp]: "{x - length xs |x. Suc x \ S} = {x - Suc (length xs) |x. x \ S}" if "0 \ S" using that apply safe apply force subgoal for x x' by (cases x') auto done from Cons.prems show ?case by (simp, subst nths_Cons, subst Cons.IH; auto) qed lemma nths_eq_ConsD: assumes "nths xs I = x # as" shows "\ ys zs. xs = ys @ x # zs \ length ys \ I \ (\ i \ I. i \ length ys) \ nths zs ({i - length ys - 1 | i. i \ I \ i > length ys}) = as" using assms proof (induction xs arbitrary: I x as) case Nil then show ?case by simp next case (Cons a xs) from Cons.prems show ?case unfolding nths_Cons apply (auto split: if_split_asm) subgoal by (inst_existentials "[] :: 'a list" xs; force intro: arg_cong2[of xs xs _ _ nths]) subgoal apply (drule Cons.IH) apply safe subgoal for ys zs apply (inst_existentials "a # ys" zs) apply simp+ apply standard subgoal for i by (cases i; auto) apply (rule arg_cong2[of zs zs _ _ nths]) apply simp apply safe subgoal for _ i by (cases i; auto) by force done done qed lemma nths_out_of_bounds: "nths xs I = []" if "\i \ I. i \ length xs" using that (* Found by sledgehammer *) proof - have "\N as. (\n. n \ N \ \ length (as::'a list) \ n) \ (\asa. nths (as @ asa) N = nths asa {n - length as |n. n \ N})" using nths_shift by blast then obtain nn :: "nat set \ 'a list \ nat" where "\N as. nn N as \ N \ \ length as \ nn N as \ (\asa. nths (as @ asa) N = nths asa {n - length as |n. n \ N})" by moura then have "\as. nths as {n - length xs |n. n \ I} = nths (xs @ as) I \ nths (xs @ []) I = []" using that by fastforce then have "nths (xs @ []) I = []" by (metis (no_types) nths_nil) then show ?thesis by simp qed lemma nths_eq_appendD: assumes "nths xs I = as @ bs" shows "\ ys zs. xs = ys @ zs \ nths ys I = as \ nths zs {i - length ys | i. i \ I \ i \ length ys} = bs" using assms proof (induction as arbitrary: xs I) case Nil then show ?case by (inst_existentials "[] :: 'a list" "nths bs") auto next case (Cons a ys xs) - from nths_eq_ConsD[of xs I a "ys @ bs"] Cons.prems obtain ys' zs' where - "xs = ys' @ a # zs'" "length ys' \ I" "\i \ I. i \ length ys'" - "nths zs' {i - length ys' - 1 |i. i \ I \ i > length ys'} = ys @ bs" + from nths_eq_ConsD[of xs I a "ys @ bs"] Cons.prems + obtain ys' zs' where + "xs = ys' @ a # zs'" + "length ys' \ I" + "\i \ I. i \ length ys'" + "nths zs' {i - length ys' - 1 |i. i \ I \ i > length ys'} = ys @ bs" by auto - moreover from Cons.IH[OF \nths zs' _ = _\] guess ys'' zs'' by clarify + moreover from Cons.IH[OF \nths zs' _ = _\] obtain ys'' zs'' where + "zs' = ys'' @ zs''" + "ys = nths ys'' {i - length ys' - 1 |i. i \ I \ length ys' < i}" + "bs = nths zs'' {i - length ys'' |i. i \ {i - length ys' - 1 |i. i \ I \ length ys' < i} \ length ys'' \ i}" + by auto ultimately show ?case apply (inst_existentials "ys' @ a # ys''" zs'') apply (simp; fail) subgoal by (simp add: nths_out_of_bounds nths_append nths_Cons) (rule arg_cong2[of ys'' ys'' _ _ nths]; force) subgoal by safe (rule arg_cong2[of zs'' zs'' _ _ nths]; force) (* Slow *) done qed lemma filter_nths_length: "length (filter P (nths xs I)) \ length (filter P xs)" proof (induction xs arbitrary: I) case Nil then show ?case by simp next case Cons then show ?case (* Found by sledgehammer *) proof - fix a :: 'a and xsa :: "'a list" and Ia :: "nat set" assume a1: "\I. length (filter P (nths xsa I)) \ length (filter P xsa)" have f2: "\b bs N. if 0 \ N then nths ((b::'a) # bs) N = [b] @ nths bs {n. Suc n \ N} else nths (b # bs) N = [] @ nths bs {n. Suc n \ N}" by (simp add: nths_Cons) have f3: "nths (a # xsa) Ia = [] @ nths xsa {n. Suc n \ Ia} \ length (filter P (nths (a # xsa) Ia)) \ length (filter P xsa)" using a1 by (metis append_Nil) have f4: "length (filter P (nths xsa {n. Suc n \ Ia})) + 0 \ length (filter P xsa) + 0" using a1 by simp have f5: "Suc (length (filter P (nths xsa {n. Suc n \ Ia})) + 0) = length (a # filter P (nths xsa {n. Suc n \ Ia}))" by force have f6: "Suc (length (filter P xsa) + 0) = length (a # filter P xsa)" by simp { assume "\ length (filter P (nths (a # xsa) Ia)) \ length (filter P (a # xsa))" { assume "nths (a # xsa) Ia \ [a] @ nths xsa {n. Suc n \ Ia}" moreover { assume "nths (a # xsa) Ia = [] @ nths xsa {n. Suc n \ Ia} \ length (filter P (a # xsa)) \ length (filter P xsa)" then have "length (filter P (nths (a # xsa) Ia)) \ length (filter P (a # xsa))" using a1 by (metis (no_types) append_Nil filter.simps(2) impossible_Cons) } ultimately have "length (filter P (nths (a # xsa) Ia)) \ length (filter P (a # xsa))" using f3 f2 by (meson dual_order.trans le_cases) } then have "length (filter P (nths (a # xsa) Ia)) \ length (filter P (a # xsa))" using f6 f5 f4 a1 by (metis Suc_le_mono append_Cons append_Nil filter.simps(2)) } then show "length (filter P (nths (a # xsa) Ia)) \ length (filter P (a # xsa))" by meson qed qed end (* Theory *) diff --git a/thys/Probabilistic_Timed_Automata/library/Stream_More.thy b/thys/Probabilistic_Timed_Automata/library/Stream_More.thy --- a/thys/Probabilistic_Timed_Automata/library/Stream_More.thy +++ b/thys/Probabilistic_Timed_Automata/library/Stream_More.thy @@ -1,331 +1,358 @@ theory Stream_More imports Sequence_LTL Instantiate_Existentials "HOL-Library.Rewrite" begin (* (* TODO: can be replaced by stream.rel_cases *) lemma stream_all2_Cons1: "stream_all2 P (y ## ys) xs \ (\ x xs'. xs = x ## xs' \ P y x \ stream_all2 P ys xs')" by (cases xs) auto lemma stream_all2_Cons2: "stream_all2 P xs (y ## ys) \ (\ x xs'. xs = x ## xs' \ P x y \ stream_all2 P xs' ys)" by (cases xs) auto (* TODO: inline? *) lemma filter_list_all_iff: "filter P xs = [] \ list_all (Not \ P) xs" unfolding filter_empty_conv list.pred_set by simp (* TODO: inline? *) lemma append_single_shift: "(xs @ [x]) @- ys = xs @- x ## ys" by simp (* TODO: inline? *) lemma sset_finite_decomp: assumes "finite (sset w)" obtains u v a w' where "w = u @- a ## v @- a ## w'" using sdistinct_infinite_sset not_sdistinct_decomp assms by metis *) lemma list_all_stake_least: "list_all (Not \ P) (stake (LEAST n. P (xs !! n)) xs)" (is "?G") if "\ n. P (xs !! n)" proof (rule ccontr) let ?n = "LEAST n. P (xs !! n)" assume "\ ?G" then have "\ x \ set (stake ?n xs). P x" unfolding list_all_iff by auto then obtain n' where "n' < ?n" "P (xs !! n')" using set_stake_snth by metis with Least_le[of "\ n. P (xs !! n)" n'] show False by auto qed lemma alw_stream_all2_mono: assumes "stream_all2 P xs ys" "alw Q xs" "\ xs ys. stream_all2 P xs ys \ Q xs \ R ys" shows "alw R ys" using assms stream.rel_sel by (coinduction arbitrary: xs ys) (blast) lemma alw_ev_HLD_cycle: assumes "stream_all2 (\) xs (cycle as)" "a \ set as" shows "infs a xs" using assms(1) proof coinduct case (infs xs) have 1: "as \ []" using assms(2) by auto have 2: "list_all2 (\) (stake (length as) xs) (stake (length as) (cycle as))" "stream_all2 (\) (sdrop (length as) xs) (sdrop (length as) (cycle as))" using infs stream_rel_shift stake_sdrop length_stake by metis+ have 3: "stake (length as) (cycle as) = as" using 1 by simp have 4: "sdrop (length as) (cycle as) = cycle as" using sdrop_cycle_eq 1 by this have 5: "set (stake (length as) xs) \ a \ {}" using assms(2) 2(1) unfolding list.in_rel 3 by (auto) (metis IntI empty_iff mem_Collect_eq set_zip_leftD split_conv subsetCE zip_map_fst_snd) show ?case using 2 5 unfolding 4 by force qed lemma alw_ev_mono: assumes "alw (ev \) xs" and "\ xs. \ xs \ \ xs" shows "alw (ev \) xs" by (rule alw_mp[OF assms(1)]) (auto intro: ev_mono assms(2) simp: alw_iff_sdrop) lemma alw_ev_lockstep: assumes "alw (ev (holds P)) xs" "stream_all2 Q xs as" "\ x a. P x \ Q x a \ R a" shows "alw (ev (holds R)) as" using assms(1,2) apply (coinduction arbitrary: xs as) apply auto subgoal by (metis alw.cases assms(3) ev_holds_sset stream_all2_sset1) subgoal by (meson alw.cases stream.rel_sel) done subsection \sfilter, wait, nxt\ text \Useful?\ lemma nxt_holds_iff_snth: "(nxt ^^ y) (holds P) xs \ P (xs !! y)" by (induction y arbitrary: xs; simp) text \Useful?\ lemma wait_LEAST: "wait (holds P) xs = (LEAST n. P (xs !! n))" unfolding wait_def nxt_holds_iff_snth .. lemma sfilter_SCons_decomp: assumes "sfilter P xs = x ## zs" "ev (holds P) xs" shows "\ ys' zs'. xs = ys' @- x ## zs' \ list_all (Not o P) ys' \ P x \ sfilter P zs' = zs" proof - from ev_imp_shift[OF assms(2)] obtain as bs where "xs = as @- bs" "holds P bs" by auto then have "P (shd bs)" by auto with \xs = _\ have "\ n. P (xs !! n)" using assms(2) sdrop_wait by fastforce from sdrop_while_sdrop_LEAST[OF this] have *: "sdrop_while (Not \ P) xs = sdrop (LEAST n. P (xs !! n)) xs" . let ?xs = "sdrop_while (Not \ P) xs" let ?n = "LEAST n. P (xs !! n)" from assms(1) have "x = shd ?xs" "zs = sfilter P (stl ?xs)" by (subst (asm) sfilter.ctr; simp)+ have "xs = stake ?n xs @- sdrop ?n xs" by simp moreover have "P x" using assms(1) unfolding sfilter_eq[OF assms(2)] .. moreover from \\ n. P _\ have "list_all (Not o P) (stake ?n xs)" by (rule list_all_stake_least) ultimately show ?thesis using \x = _\ \zs = _\ *[symmetric] by (inst_existentials "stake ?n xs" "stl ?xs") auto qed lemma sfilter_SCons_decomp': assumes "sfilter P xs = x ## zs" "ev (holds P) xs" shows "list_all (Not o P) (stake (wait (holds P) xs) xs)" (is "?G1") "P x" "\ zs'. xs = stake (wait (holds P) xs) xs @- x ## zs' \ sfilter P zs' = zs" (is "?G2") proof - from ev_imp_shift[OF assms(2)] obtain as bs where "xs = as @- bs" "holds P bs" by auto then have "P (shd bs)" by auto with \xs = _\ have "\ n. P (xs !! n)" using assms(2) sdrop_wait by fastforce thm sdrop_wait from sdrop_while_sdrop_LEAST[OF this] have *: "sdrop_while (Not \ P) xs = sdrop (LEAST n. P (xs !! n)) xs" . let ?xs = "sdrop_while (Not \ P) xs" let ?n = "wait (holds P) xs" from assms(1) have "x = shd ?xs" "zs = sfilter P (stl ?xs)" by (subst (asm) sfilter.ctr; simp)+ have "xs = stake ?n xs @- sdrop ?n xs" by simp moreover show "P x" using assms(1) unfolding sfilter_eq[OF assms(2)] .. moreover from \\ n. P _\ show "list_all (Not o P) (stake ?n xs)" by (auto intro: list_all_stake_least simp: wait_LEAST) ultimately show ?G2 using \x = _\ \zs = _\ *[symmetric] by (inst_existentials "stl ?xs") (auto simp: wait_LEAST) qed lemma sfilter_shift_decomp: assumes "sfilter P xs = ys @- zs" "alw (ev (holds P)) xs" shows "\ ys' zs'. xs = ys' @- zs' \ filter P ys' = ys \ sfilter P zs' = zs" using assms(1,2) proof (induction ys arbitrary: xs) case Nil then show ?case by (inst_existentials "[] :: 'a list" xs; simp) next case (Cons y ys) from alw_ev_imp_ev_alw[OF \alw (ev _) xs\] have "ev (holds P) xs" by (auto elim: ev_mono) - with Cons.prems(1) sfilter_SCons_decomp[of P xs y "ys @- zs"] guess ys' zs' by clarsimp - note guessed = this + with Cons.prems(1) sfilter_SCons_decomp[of P xs y "ys @- zs"] + obtain ys' zs' where *: + "xs = ys' @- y ## zs'" + "list_all (Not \ P) ys'" + "P y" + "sfilter P zs' = ys @- zs" + by auto then have "sfilter P zs' = ys @- zs" by auto from \alw (ev _) xs\ \xs = _\ have "alw (ev (holds P)) zs'" by (metis ev.intros(2) ev_shift not_alw_iff stream.sel(2)) - from Cons.IH[OF \sfilter P zs' = _\ this] guess zs1 zs2 by clarsimp - with guessed show ?case + from Cons.IH[OF \sfilter P zs' = _\ this] + obtain zs1 zs2 where + "zs' = zs1 @- zs2" + "ys = filter P zs1" + "zs = sfilter P zs2" + by auto + with * show ?case by (inst_existentials "ys' @ y # zs1" zs2; simp add: list.pred_set) qed lemma finite_sset_sfilter_decomp: assumes "finite (sset (sfilter P xs))" "alw (ev (holds P)) xs" obtains x ws ys zs where "xs = ws @- x ## ys @- x ## zs" "P x" proof atomize_elim let ?xs = "sfilter P xs" have 1: "\ sdistinct (sfilter P xs)" using sdistinct_infinite_sset assms(1) by auto - from not_sdistinct_decomp[OF 1] guess ws ys x zs . - note guessed1 = this - from sfilter_shift_decomp[OF this assms(2)] guess ys' zs' by clarsimp - note guessed2 = this + from not_sdistinct_decomp[OF 1] + obtain ws ys x zs where 1: "sfilter P xs = ws @- x ## ys @- x ## zs" . + from sfilter_shift_decomp[OF this assms(2)] + obtain ys' zs' where 2: "xs = ys' @- zs'" "sfilter P zs' = x ## ys @- x ## zs" by auto then have "ev (holds P) zs'" using alw_shift assms(2) by blast - from sfilter_SCons_decomp[OF guessed2(2) this] guess zs1 zs2 by clarsimp - note guessed3 = this + from sfilter_SCons_decomp[OF 2(2) this] + obtain zs1 zs2 where 3: + "zs' = zs1 @- x ## zs2" "list_all (Not \ P) zs1" "P x" "sfilter P zs2 = ys @- x ## zs" + by auto have "alw (ev (holds P)) zs2" - by (metis alw_ev_stl alw_shift assms(2) guessed2(1) guessed3(1) stream.sel(2)) - from sfilter_shift_decomp[OF guessed3(4) this] guess zs3 zs4 by clarsimp - note guessed4 = this + by (metis alw_ev_stl alw_shift assms(2) 2(1) 3(1) stream.sel(2)) + from sfilter_shift_decomp[OF 3(4) this] + obtain zs3 zs4 where 4: "zs2 = zs3 @- zs4" "sfilter P zs4 = x ## zs" by auto have "ev (holds P) zs4" - using \alw (ev (holds P)) zs2\ alw_shift guessed4(1) by blast - from sfilter_SCons_decomp[OF guessed4(2) this] guess zs5 zs6 by clarsimp - with guessed1 guessed2 guessed3 guessed4 show "\ws x ys zs. xs = ws @- x ## ys @- x ## zs \ P x" + using \alw (ev (holds P)) zs2\ alw_shift 4(1) by blast + from sfilter_SCons_decomp[OF 4(2) this] + obtain zs5 zs6 where "zs4 = zs5 @- x ## zs6" "list_all (Not \ P) zs5" "P x" "zs = sfilter P zs6" + by auto + with 1 2 3 4 show "\ws x ys zs. xs = ws @- x ## ys @- x ## zs \ P x" by (inst_existentials "ys' @ zs1" x "zs3 @ zs5" zs6; force) qed text \Useful?\ lemma sfilter_shd_LEAST: "shd (sfilter P xs) = xs !! (LEAST n. P (xs !! n))" if "ev (holds P) xs" proof - from sdrop_wait[OF \ev _ xs\] have "\ n. P (xs !! n)" by auto from sdrop_while_sdrop_LEAST[OF this] show ?thesis by simp qed lemma alw_nxt_holds_cong: "(nxt ^^ n) (holds (\x. P x \ Q x)) xs = (nxt ^^ n) (holds Q) xs" if "alw (holds P) xs" using that unfolding nxt_holds_iff_snth alw_iff_sdrop by simp lemma alw_wait_holds_cong: "wait (holds (\x. P x \ Q x)) xs = wait (holds Q) xs" if "alw (holds P) xs" unfolding wait_def alw_nxt_holds_cong[OF that] .. lemma alw_sfilter: "sfilter (\ x. P x \ Q x) xs = sfilter Q xs" if "alw (holds P) xs" "alw (ev (holds Q)) xs" using that proof (coinduction arbitrary: xs) case prems: stream_eq from prems(3,4) have ev_one: "ev (holds (\x. P x \ Q x)) xs" by (subst ev_cong[of _ _ _ "holds Q"]) (assumption | auto)+ from prems have "a = shd (sfilter (\x. P x \ Q x) xs)" "b = shd (sfilter Q xs)" by (metis stream.sel(1))+ with prems(3,4) have "a = xs !! (LEAST n. P (xs !! n) \ Q (xs !! n))" "b = xs !! (LEAST n. Q (xs !! n))" using ev_one by (auto 4 3 dest: sfilter_shd_LEAST) with alw_wait_holds_cong[unfolded wait_LEAST, OF \alw (holds P) xs\] have "a = b" by simp - from sfilter_SCons_decomp'[OF prems(1)[symmetric], OF ev_one] guess u2 by clarsimp - note guessed_a = this + from sfilter_SCons_decomp'[OF prems(1)[symmetric], OF ev_one] + obtain u2 where u2: + "xs = stake (wait (holds (\x. P x \ Q x)) xs) xs @- a ## u2" + "u = sfilter (\x. P x \ Q x) u2" + by auto have "ev (holds Q) xs" using prems(4) by blast - from sfilter_SCons_decomp'[OF prems(2)[symmetric], OF this] guess v2 by clarsimp - with guessed_a \a = b\ show ?case + from sfilter_SCons_decomp'[OF prems(2)[symmetric], OF this] + obtain v2 where + "list_all (Not \ Q) (stake (wait (holds Q) xs) xs)" + "xs = stake (wait (holds Q) xs) xs @- b ## v2" + "v = sfilter Q v2" + by auto + with u2 \a = b\ show ?case apply (intro conjI exI) apply assumption+ apply (simp add: alw_wait_holds_cong[OF prems(3)], metis shift_left_inj stream.inject) by (metis alw.cases alw_shift prems(3,4) stream.sel(2))+ qed lemma alw_ev_holds_mp: "alw (holds P) xs \ ev (holds Q) xs \ ev (holds (\x. P x \ Q x)) xs" by (subst ev_cong, assumption) auto lemma alw_ev_conjI: "alw (ev (holds (\ x. P x \ Q x))) xs" if "alw (holds P) xs" "alw (ev (holds Q)) xs" using that(2,1) by - (erule alw_mp, coinduction arbitrary: xs, auto intro: alw_ev_holds_mp) subsection \Useful?\ lemma alw_holds_pred_stream_iff: "alw (holds P) xs \ pred_stream P xs" by (simp add: alw_iff_sdrop stream_pred_snth) lemma alw_holds_sset: "alw (holds P) xs = (\ x \ sset xs. P x)" by (simp add: alw_holds_pred_stream_iff stream.pred_set) lemma pred_stream_sfilter: assumes alw_ev: "alw (ev (holds P)) xs" shows "pred_stream P (sfilter P xs)" using alw_ev proof (coinduction arbitrary: xs) case (stream_pred xs) then have "ev (holds P) xs" by auto have "sfilter P xs = shd (sfilter P xs) ## stl (sfilter P xs)" by (cases "sfilter P xs") auto - from sfilter_SCons_decomp[OF this \ev (holds P) xs\] guess ys' zs' by clarsimp + from sfilter_SCons_decomp[OF this \ev (holds P) xs\] + obtain ys' zs' where + "xs = ys' @- shd (sdrop_while (Not \ P) xs) ## zs'" + "P (shd (sdrop_while (Not \ P) xs))" + "sfilter P zs' = sfilter P (stl (sdrop_while (Not \ P) xs))" + by auto then show ?case apply (inst_existentials zs') apply (metis sfilter.simps(1) stream.sel(1) stream_pred(1)) apply (metis scons_eq sfilter.simps(2) stream_pred(1)) apply (metis alw_ev_stl alw_shift stream.sel(2) stream_pred(2)) done qed lemma alw_ev_sfilter_mono: assumes alw_ev: "alw (ev (holds P)) xs" and mono: "\ x. P x \ Q x" shows "pred_stream Q (sfilter P xs)" using stream.pred_mono[of P Q] assms pred_stream_sfilter by blast lemma sset_sfilter: "sset (sfilter P xs) \ sset xs" if "alw (ev (holds P)) xs" proof - have "alw (holds (\ x. x \ sset xs)) xs" by (simp add: alw_iff_sdrop) with \alw (ev _) _\ alw_sfilter[OF this \alw (ev _) _\, symmetric] have "pred_stream (\ x. x \ sset xs) (sfilter P xs)" by (simp) (rule alw_ev_sfilter_mono; auto intro: alw_ev_conjI) then have "\ x \ sset (sfilter P xs). x \ sset xs" unfolding stream.pred_set by this then show ?thesis by blast qed lemma stream_all2_weaken: "stream_all2 Q xs ys" if "stream_all2 P xs ys" "\ x y. P x y \ Q x y" using that by (coinduction arbitrary: xs ys) auto lemma stream_all2_SCons1: "stream_all2 P (x ## xs) ys = (\z zs. ys = z ## zs \ P x z \ stream_all2 P xs zs)" by (subst (3) stream.collapse[symmetric], simp del: stream.collapse, force) lemma stream_all2_SCons2: "stream_all2 P xs (y ## ys) = (\z zs. xs = z ## zs \ P z y \ stream_all2 P zs ys)" by (subst stream.collapse[symmetric], simp del: stream.collapse, force) lemma stream_all2_combine: "stream_all2 R xs zs" if "stream_all2 P xs ys" "stream_all2 Q ys zs" "\ x y z. P x y \ Q y z \ R x z" using that(1,2) by (coinduction arbitrary: xs ys zs) (auto intro: that(3) simp: stream_all2_SCons1 stream_all2_SCons2) lemma stream_all2_shift1: "stream_all2 P (xs1 @- xs2) ys = (\ ys1 ys2. ys = ys1 @- ys2 \ list_all2 P xs1 ys1 \ stream_all2 P xs2 ys2)" apply (induction xs1 arbitrary: ys) apply (simp; fail) apply (simp add: stream_all2_SCons1 list_all2_Cons1) apply safe subgoal for a xs1 ys z zs ys1 ys2 by (inst_existentials "z # ys1" ys2; simp) subgoal for a xs1 ys ys1 ys2 z zs by (inst_existentials z "zs @- ys2" zs "ys2"; simp) done lemma stream_all2_shift2: "stream_all2 P ys (xs1 @- xs2) = (\ ys1 ys2. ys = ys1 @- ys2 \ list_all2 P ys1 xs1 \ stream_all2 P ys2 xs2)" by (meson list.rel_flip stream.rel_flip stream_all2_shift1) lemma stream_all2_bisim: assumes "stream_all2 (\) xs as" "stream_all2 (\) ys as" "sset as \ S" shows "stream_all2 (\ x y. \ a. x \ a \ y \ a \ a \ S) xs ys" using assms apply (coinduction arbitrary: as xs ys) subgoal for a u b v as xs ys apply (rule conjI) apply (inst_existentials "shd as", auto simp: stream_all2_SCons1; fail) apply (inst_existentials "stl as", auto 4 3 simp: stream_all2_SCons1; fail) done done end diff --git a/thys/Propositional_Proof_Systems/Compactness.thy b/thys/Propositional_Proof_Systems/Compactness.thy --- a/thys/Propositional_Proof_Systems/Compactness.thy +++ b/thys/Propositional_Proof_Systems/Compactness.thy @@ -1,312 +1,312 @@ subsection\Compactness\ theory Compactness imports Sema begin lemma fin_sat_extend: "fin_sat S \ fin_sat (insert F S) \ fin_sat (insert (\<^bold>\F) S)" proof (rule ccontr) assume fs: "fin_sat S" assume nfs: "\ (fin_sat (insert F S) \ fin_sat (insert (\<^bold>\ F) S))" from nfs obtain s1 where s1: "s1 \ insert F S" "finite s1" "\sat s1" unfolding fin_sat_def by blast from nfs obtain s2 where s2: "s2 \ insert (Not F) S" "finite s2" "\sat s2" unfolding fin_sat_def by blast let ?u = "(s1 - {F}) \ (s2 - {Not F})" have "?u \ S" "finite ?u" using s1 s2 by auto hence "sat ?u" using fs unfolding fin_sat_def by blast then obtain A where A: "\F \ ?u. A \ F" unfolding sat_def by blast have "A \ F \ A \ \<^bold>\F" by simp hence "sat s1 \ sat s2" using A unfolding sat_def by(fastforce intro!: exI[where x=A]) thus False using s1(3) s2(3) by simp qed context begin lemma fin_sat_antimono: "fin_sat F \ G \ F \ fin_sat G" unfolding fin_sat_def by simp lemmas fin_sat_insert = fin_sat_antimono[OF _ subset_insertI] primrec extender :: "nat \ ('a :: countable) formula set \ 'a formula set" where "extender 0 S = S" | "extender (Suc n) S = ( let r = extender n S; rt = insert (from_nat n) r; rf = insert (\<^bold>\(from_nat n)) r in if fin_sat rf then rf else rt )" private lemma extender_fin_sat: "fin_sat S \ fin_sat (extender n S)" proof(induction n arbitrary: S) case (Suc n) note mIH = Suc.IH[OF Suc.prems] show ?case proof(cases "fin_sat (insert (Not (from_nat n)) (extender n S))") case True thus ?thesis by simp next case False hence "fin_sat (insert ((from_nat n)) (extender n S))" using mIH fin_sat_extend by auto thus ?thesis by(simp add: Let_def) qed qed simp definition "extended S = \{extender n S|n. True}" lemma extended_max: "F \ extended S \ Not F \ extended S" proof - obtain n where [simp]: "F = from_nat n" by (metis from_nat_to_nat) have "F \ extender (Suc n) S \ Not F \ extender (Suc n) S" by(simp add: Let_def) thus ?thesis unfolding extended_def by blast qed private lemma extender_Sucset: "extender k S \ extender (Suc k) S" by(force simp add: Let_def) private lemma extender_deeper: "F \ extender k S \ k \ l \ F \ extender l S" using extender_Sucset le_Suc_eq by(induction l) (auto simp del: extender.simps) private lemma extender_subset: "S \ extender k S" proof - from extender_deeper[OF _ le0] have "F \ extender 0 Sa \ F \ extender l Sa" for Sa l F . thus ?thesis by auto qed lemma extended_fin_sat: assumes "fin_sat S" shows "fin_sat (extended S)" proof - have assm: "\s \ extender n S; finite s\ \ sat s" for s n using extender_fin_sat[OF assms] unfolding fin_sat_def by presburger hence "sat s" if su: "s \ \{extender n S |n. True}" and fin: "finite s" for s proof - { fix x assume e: "x \ s" with su have "x \ \{extender n S |n. True}" by blast hence "\n. x \ extender n S" unfolding Union_eq by blast } hence "\x \ s. \n. x \ extender n S" by blast from finite_set_choice[OF fin this] obtain f where cf: "\x\s. x \ extender (f x) S" .. have "\k. s \ \{extender n S |n. n \ k}" proof(intro exI subsetI) fix x assume e: "x \ s" with cf have "x \ extender (f x) S" .. hence "x \ extender (Max (f ` s)) S" by(elim extender_deeper; simp add: e fin) thus "x \ \{extender n S |n. n \ Max (f ` s)}" by blast qed moreover have "\{extender n S |n. n \ k} = extender k S" for k proof(induction k) case (Suc k) moreover have "\{extender n S |n. n \ Suc k} = \{extender n S |n. n \ k} \ extender (Suc k) S" unfolding Union_eq le_Suc_eq using le_Suc_eq by(auto simp del: extender.simps) ultimately show ?case using extender_Sucset by(force simp del: extender.simps) qed simp ultimately show "sat s" using assm fin by auto qed thus ?thesis unfolding extended_def fin_sat_def by presburger qed lemma extended_superset: "S \ extended S" unfolding extended_def using extender.simps(1) by blast lemma extended_complem: assumes fs: "fin_sat S" shows "(F \ extended S) \ (Not F \ extended S)" proof - note fs = fs[THEN extended_fin_sat] show ?thesis proof(cases "F \ extended S") case False with extended_max show ?thesis by blast next case True have "Not F \ extended S" proof assume False: "Not F \ extended S" with True have "{F, Not F} \ extended S" by blast moreover have "finite {F, Not F}" by simp ultimately have "sat {F, Not F}" using fs unfolding fin_sat_def by blast thus False unfolding sat_def by simp qed with True show ?thesis by blast qed qed lemma not_fin_sat_extended_UNIV: fixes S :: "'a :: countable formula set" assumes "\fin_sat S" shows "extended S = UNIV" text\Note that this crucially depends on the fact that we check \emph{first} whether adding @{term "\<^bold>\F"} makes the set not satisfiable, and add @{term F} otherwise \emph{without any further checks}. The proof of compactness does (to the best of my knowledge) depend on neither of these two facts.\ proof - from assms[unfolded fin_sat_def, simplified] obtain s :: "'a :: countable formula set" where "finite s" "\ sat s" by clarify from this(2)[unfolded sat_def, simplified] have "\x\s. \ A \ x" for A .. have nfs: "\fin_sat (insert x (extender n S))" for n x apply(rule notI) apply(drule fin_sat_insert) apply(drule fin_sat_antimono) apply(rule extender_subset) apply(erule notE[rotated]) apply(fact assms) done have "x \ \{extender n S |n. True}" for x proof cases assume "x \ S" thus ?thesis by (metis extended_def extended_superset insert_absorb insert_subset) next assume "x \ S" have "x \ extender (Suc (to_nat x)) S" unfolding extender.simps Let_def unfolding if_not_P[OF nfs] by simp thus ?thesis by blast qed thus ?thesis unfolding extended_def by auto qed lemma extended_tran: "S \ T \ extended S \ extended T" text\This lemma doesn't hold: think of making S empty and inserting a formula into T s.t. it can never be satisfied simultaneously with the first non-tautological formula in the extension S. Showing that this is possible is not worth the effort, since we can't influence the ordering of formulae. But we showed it anyway.\ oops lemma extended_not_increasing: "\S T. fin_sat S \ fin_sat T \ \ (S \ T \ extended S \ extended (T :: 'a :: countable formula set))" proof - have ex_then_min: "\x :: nat. P x \ P (LEAST x. P x)" for P using LeastI2_wellorder by auto define P where "P x = (let F = (from_nat x :: 'a formula) in (\A. \ A \ F) \ (\ A. A \ F))" for x define x where "x = (LEAST n. P n)" hence "\n. P n" unfolding P_def Let_def by(auto intro!: exI[where x="to_nat (Atom undefined :: 'a formula)"]) from ex_then_min[OF this] have Px: "P x" unfolding x_def . have lessx: "n < x \ \ P n" for n unfolding x_def using not_less_Least by blast let ?S = "{} :: 'a formula set" let ?T = "{from_nat x :: 'a formula}" have s: "fin_sat ?S" "fin_sat ?T" using Px unfolding P_def fin_sat_def sat_def Let_def by fastforce+ have reject: "Q A \ \A. \ Q A \ False" for A Q by simp have "y \ x \ F \ extender y ?S \ \ F" for F y proof(induction y arbitrary: F) case (Suc y) have *: "F \ extender y {} \ \ F" for F :: "'a formula" using Suc.IH Suc.prems(1) by auto let ?Y = "from_nat y :: 'a formula" have ex: "(\A. \ A \ ?Y) \ \ ?Y" unfolding formula_semantics.simps by (meson P_def Suc.prems(1) Suc_le_lessD lessx) have 1: "\A. \ A \ ?Y" if "fin_sat (Not ?Y \ extender y ?S)" proof - note[[show_types]] from that have "\A. A \ Not ?Y" unfolding fin_sat_def sat_def by(elim allE[where x="{Not ?Y}"]) simp hence "\\ ?Y" by simp hence "\A. \ A \ ?Y" using ex by argo thus ?thesis by simp qed have 2: "\ fin_sat (Not ?Y \ extender y ?S) \ \ ?Y" proof(erule contrapos_np) assume "\ \ ?Y" hence "\A. \ A \ ?Y" using ex by argo hence "\ \<^bold>\ ?Y" by simp thus "fin_sat (\<^bold>\ ?Y \ extender y ?S)" unfolding fin_sat_def sat_def by(auto intro!: exI[where x="\_ :: 'a. False"] dest!: rev_subsetD[rotated] *) qed show ?case using Suc.prems(2) by(simp add: Let_def split: if_splits; elim disjE; simp add: * 1 2) qed simp hence "fin_sat (\<^bold>\ (from_nat x) \ extender x ?S)" using Px unfolding P_def Let_def by (clarsimp simp: fin_sat_def sat_def) (insert formula_semantics.simps(3), blast) hence "Not (from_nat x) \ extender (Suc x) ?S" by(simp) hence "Not (from_nat x) \ extended ?S" unfolding extended_def by blast moreover have "Not (from_nat x) \ extended ?T" using extended_complem extended_superset s(2) by blast ultimately show ?thesis using s by blast qed private lemma not_in_extended_FE: "fin_sat S \ (\sat (insert (Not F) G)) \ F \ extended S \ G \ extended S \ finite G \ False" proof(goal_cases) case 1 hence "Not F \ extended S" using extended_max by blast thus False using 1 extended_fin_sat fin_sat_def by (metis Diff_eq_empty_iff finite.insertI insert_Diff_if) qed lemma extended_id: "extended (extended S) = extended S" using extended_complem extended_fin_sat extended_max extended_superset not_fin_sat_extended_UNIV by(intro equalityI[rotated] extended_superset) blast (* This would be nicer, though\ inductive_set extended_set :: "form set \ form set" for S where "F \ S \ F \ extended_set S" | "fin_sat (insert F (extended_set S)) \ F \ extended_set S \ F \ extended_set S" but it can't work, as extended is not increasing (?) *) lemma ext_model: assumes r: "fin_sat S" shows "(\k. Atom k \ extended S) \ F \ F \ extended S" proof - note fs = r[THEN extended_fin_sat] have Elim: "F \ S \ G \ S \ {F,G} \ S" "F \ S \ {F} \ S" for F G S by simp+ show ?thesis proof(induction F) case Atom thus ?case by(simp) next case Bot have False if "\ \ extended S" proof - have "finite {\}" by simp moreover from that have "{\} \ extended S" by simp ultimately have "\A. A \ \" using fs unfolding fin_sat_def sat_def by(elim allE[of _ "{\}"]) simp thus False by simp qed thus ?case by auto next case (Not F) moreover have "A \ F \ A \ \<^bold>\F" for A F by simp ultimately show ?case using extended_complem[OF r] by blast next case (And F G) have "(F \ extended S \ G \ extended S) = (F \<^bold>\ G \ extended S)" proof - have *: "\sat {\<^bold>\ (F \<^bold>\ G), F, G}" "\sat {\<^bold>\ F, (F \<^bold>\ G)}" "\sat {\<^bold>\ G, (F \<^bold>\ G)}" unfolding sat_def by auto show ?thesis by(intro iffI; rule ccontr) (auto intro: *[THEN not_in_extended_FE[OF r]]) qed thus ?case by(simp add: And) next case (Or F G) have "(F \ extended S \ G \ extended S) = (F \<^bold>\ G \ extended S)" proof - have "\sat {\<^bold>\(F \<^bold>\ G), F}" "\sat {\<^bold>\(F \<^bold>\ G), G}" unfolding sat_def by auto from this[THEN not_in_extended_FE[OF r]] have 1: "\F \ extended S \ G \ extended S; F \<^bold>\ G \ extended S\ \ False" by auto have "\sat {\<^bold>\F, \<^bold>\G, F \<^bold>\ G}" unfolding sat_def by auto hence 2: "\F \<^bold>\ G \ extended S; F \ extended S; G \ extended S\ \ False" using extended_max not_in_extended_FE[OF r] by fastforce show ?thesis by(intro iffI; rule ccontr) (auto intro: 1 2) qed thus ?case by(simp add: Or) next case (Imp F G) have "(F \ extended S \ G \ extended S) = (F \<^bold>\ G \ extended S)" proof - have "\sat {\<^bold>\G, F, F \<^bold>\ G}" unfolding sat_def by auto hence 1: "\F \<^bold>\ G \ extended S; F \ extended S; G \ extended S\ \ False" using extended_max not_in_extended_FE[OF r] by blast have "\sat {\<^bold>\F,\<^bold>\(F \<^bold>\ G)}" unfolding sat_def by auto hence 2: "\F \<^bold>\ G \ extended S; F \ extended S\ \ False" using extended_max not_in_extended_FE[OF r] by blast have "\sat {\<^bold>\(F \<^bold>\ G),G}" unfolding sat_def by auto hence 3: "\F \<^bold>\ G \ extended S; G \ extended S\ \ False" using extended_max not_in_extended_FE[OF r] by blast show ?thesis by(intro iffI; rule ccontr) (auto intro: 1 2 3) qed thus ?case by(simp add: Imp) qed qed theorem compactness: fixes S :: "'a :: countable formula set" shows "sat S \ fin_sat S" (is "?l = ?r") proof assume ?l thus ?r unfolding sat_def fin_sat_def by blast next assume r: ?r note ext_model[OF r, THEN iffD2] hence "\F\S. (\k. Atom k \ extended S) \ F" using extended_superset by blast thus ?l unfolding sat_def by blast qed corollary compact_entailment: fixes F :: "'a :: countable formula" assumes fent: "\ \ F" shows "\\'. finite \' \ \' \ \ \ \' \ F" proof - have ND_sem: "\ \ F \ \sat (insert (\<^bold>\F) \)" for \ F unfolding sat_def entailment_def by auto obtain \' where 0: "finite \'" "\' \ F" "\' \ \" proof(goal_cases) from fent[unfolded ND_sem compactness] have "\ fin_sat (insert (\<^bold>\ F) \)" . from this[unfolded fin_sat_def] obtain s where s: "s \ insert(\<^bold>\F) \" "finite s" "\sat s" by blast have 2: "finite (s - {\<^bold>\ F})" using s by simp have 3: "s - {\<^bold>\ F} \ F" unfolding ND_sem using s(3) unfolding sat_def by blast have 4: "s - {\<^bold>\ F} \ \" using s by blast case 1 from 2 3 4 show ?case by(intro 1[of "s - {Not F}"]) qed thus ?thesis by blast qed corollary compact_to_formula: fixes F :: "'a :: countable formula" assumes fent: "\ \ F" obtains \' where "set \' \ \" "\ (\<^bold>\\') \<^bold>\ F" proof goal_cases case 1 - note compact_entailment[OF assms] - then guess \' .. - moreover then obtain \'' where "\' = set \''" using finite_list by auto - ultimately show thesis by(intro 1) (blast, simp add: entailment_def) + from compact_entailment[OF assms] + obtain \' where \': "finite \' \ \' \ \ \ \' \ F" .. + then obtain \'' where "\' = set \''" using finite_list by auto + with \' show thesis by(intro 1) (blast, simp add: entailment_def) qed end end diff --git a/thys/Propositional_Proof_Systems/Consistency.thy b/thys/Propositional_Proof_Systems/Consistency.thy --- a/thys/Propositional_Proof_Systems/Consistency.thy +++ b/thys/Propositional_Proof_Systems/Consistency.thy @@ -1,359 +1,358 @@ subsection\Consistency\ text\We follow the proofs by Melvin Fitting~\cite{fitting1990first}.\ theory Consistency imports Sema begin definition "Hintikka S \ ( \ \ S \ (\k. Atom k \ S \ \<^bold>\ (Atom k) \ S \ False) \ (\F G. F \<^bold>\ G \ S \ F \ S \ G \ S) \ (\F G. F \<^bold>\ G \ S \ F \ S \ G \ S) \ (\F G. F \<^bold>\ G \ S \ \<^bold>\F \ S \ G \ S) \ (\F. \<^bold>\ (\<^bold>\F) \ S \ F \ S) \ (\F G. \<^bold>\(F \<^bold>\ G) \ S \ \<^bold>\ F \ S \ \<^bold>\ G \ S) \ (\F G. \<^bold>\(F \<^bold>\ G) \ S \ \<^bold>\ F \ S \ \<^bold>\ G \ S) \ (\F G. \<^bold>\(F \<^bold>\ G) \ S \ F \ S \ \<^bold>\ G \ S) )" lemma "Hintikka {Atom 0 \<^bold>\ ((\<^bold>\ (Atom 1)) \<^bold>\ Atom 2), ((\<^bold>\ (Atom 1)) \<^bold>\ Atom 2), Atom 0, \<^bold>\(\<^bold>\ (Atom 1)), Atom 1}" unfolding Hintikka_def by simp theorem Hintikkas_lemma: assumes H: "Hintikka S" shows "sat S" proof - from H[unfolded Hintikka_def] have H': "\ \ S" "Atom k \ S \ \<^bold>\ (Atom k) \ S \ False" "F \<^bold>\ G \ S \ F \ S \ G \ S" "F \<^bold>\ G \ S \ F \ S \ G \ S" "F \<^bold>\ G \ S \ \<^bold>\F \ S \ G \ S" "\<^bold>\ (\<^bold>\ F) \ S \ F \ S" "\<^bold>\ (F \<^bold>\ G) \ S \ \<^bold>\ F \ S \ \<^bold>\ G \ S" "\<^bold>\ (F \<^bold>\ G) \ S \ \<^bold>\ F \ S \ \<^bold>\ G \ S" "\<^bold>\ (F \<^bold>\ G) \ S \ F \ S \ \<^bold>\ G \ S" for k F G by blast+ let ?M = "\k. Atom k \ S" have "(F \ S \ (?M \ F)) \ (\<^bold>\ F \ S \ (\(?M \ F)))" for F by(induction F) (auto simp: H'(1) dest!: H'(2-)) thus ?thesis unfolding sat_def by blast qed definition "pcp C \ (\S \ C. \ \ S \ (\k. Atom k \ S \ \<^bold>\ (Atom k) \ S \ False) \ (\F G. F \<^bold>\ G \ S \ F \ G \ S \ C) \ (\F G. F \<^bold>\ G \ S \ F \ S \ C \ G \ S \ C) \ (\F G. F \<^bold>\ G \ S \ \<^bold>\F \ S \ C \ G \ S \ C) \ (\F. \<^bold>\ (\<^bold>\F) \ S \ F \ S \ C) \ (\F G. \<^bold>\(F \<^bold>\ G) \ S \ \<^bold>\ F \ S \ C \ \<^bold>\ G \ S \ C) \ (\F G. \<^bold>\(F \<^bold>\ G) \ S \ \<^bold>\ F \ \<^bold>\ G \ S \ C) \ (\F G. \<^bold>\(F \<^bold>\ G) \ S \ F \ \<^bold>\ G \ S \ C) )" (* just some examples *) lemma "pcp {}" "pcp {{}}" "pcp {{Atom 0}}" by (simp add: pcp_def)+ lemma "pcp {{(\<^bold>\ (Atom 1)) \<^bold>\ Atom 2}, {((\<^bold>\ (Atom 1)) \<^bold>\ Atom 2), \<^bold>\(\<^bold>\ (Atom 1))}, {((\<^bold>\ (Atom 1)) \<^bold>\ Atom 2), \<^bold>\(\<^bold>\ (Atom 1)), Atom 1}}" by (auto simp add: pcp_def) text\Fitting uses uniform notation~\cite{smullyan1963unifying} for the definition of @{const pcp}. We try to mimic this, more to see whether it works than because it is ultimately necessary.\ (* It does help a bit, occasionally. *) inductive Con :: "'a formula => 'a formula => 'a formula => bool" where "Con (And F G) F G" | "Con (Not (Or F G)) (Not F) (Not G)" | "Con (Not (Imp F G)) F (Not G)" | "Con (Not (Not F)) F F" inductive Dis :: "'a formula => 'a formula => 'a formula => bool" where "Dis (Or F G) F G" | "Dis (Imp F G) (Not F) G" | "Dis (Not (And F G)) (Not F) (Not G)" | "Dis (Not (Not F)) F F" (* note that *) lemma "Con (Not (Not F)) F F" "Dis (Not (Not F)) F F" by(intro Con.intros Dis.intros)+ (* i.e. \<^bold>\\<^bold>\ is both Conjunctive and Disjunctive. I saw no reason to break this symmetry. *) lemma con_dis_simps: "Con a1 a2 a3 = (a1 = a2 \<^bold>\ a3 \ (\F G. a1 = \<^bold>\ (F \<^bold>\ G) \ a2 = \<^bold>\ F \ a3 = \<^bold>\ G) \ (\G. a1 = \<^bold>\ (a2 \<^bold>\ G) \ a3 = \<^bold>\ G) \ a1 = \<^bold>\ (\<^bold>\ a2) \ a3 = a2)" "Dis a1 a2 a3 = (a1 = a2 \<^bold>\ a3 \ (\F G. a1 = F \<^bold>\ G \ a2 = \<^bold>\ F \ a3 = G) \ (\F G. a1 = \<^bold>\ (F \<^bold>\ G) \ a2 = \<^bold>\ F \ a3 = \<^bold>\ G) \ a1 = \<^bold>\ (\<^bold>\ a2) \ a3 = a2)" by(simp_all add: Con.simps Dis.simps) lemma Hintikka_alt: "Hintikka S = ( \ \ S \ (\k. Atom k \ S \ \<^bold>\ (Atom k) \ S \ False) \ (\F G H. Con F G H \ F \ S \ G \ S \ H \ S) \ (\F G H. Dis F G H \ F \ S \ G \ S \ H \ S) )" apply(simp add: Hintikka_def con_dis_simps) apply(rule iffI) (* simp only: \ loops. :( *) subgoal by blast subgoal by safe metis+ done lemma pcp_alt: "pcp C = (\S \ C. \ \ S \ (\k. Atom k \ S \ \<^bold>\ (Atom k) \ S \ False) \ (\F G H. Con F G H \ F \ S \ G \ H \ S \ C) \ (\F G H. Dis F G H \ F \ S \ G \ S \ C \ H \ S \ C) )" apply(simp add: pcp_def con_dis_simps) apply(rule iffI; unfold Ball_def; elim all_forward) by (auto simp add: insert_absorb split: formula.splits) definition "subset_closed C \ (\S \ C. \s\S. s \ C)" definition "finite_character C \ (\S. S \ C \ (\s \ S. finite s \ s \ C))" lemma ex1: "pcp C \ \C'. C \ C' \ pcp C' \ subset_closed C'" proof(intro exI[of _ "{s . \S \ C. s \ S}"] conjI) let ?E = "{s. \S\C. s \ S}" show "C \ ?E" by blast show "subset_closed ?E" unfolding subset_closed_def by blast assume C: \pcp C\ show "pcp ?E" using C unfolding pcp_alt by (intro ballI conjI; simp; meson insertI1 rev_subsetD subset_insertI subset_insertI2) qed lemma sallI: "(\s. s \ S \ P s) \ \s \ S. P s" by simp lemma ex2: assumes fc: "finite_character C" shows "subset_closed C" unfolding subset_closed_def proof (intro ballI sallI) fix s S assume e: \S \ C\ and s: \s \ S\ hence *: "t \ s \ t \ S" for t by simp from fc have "t \ S \ finite t \ t \ C" for t unfolding finite_character_def using e by blast hence "t \ s \ finite t \ t \ C" for t using * by simp with fc show \s \ C\ unfolding finite_character_def by blast qed lemma assumes C: "pcp C" assumes S: "subset_closed C" shows ex3: "\C'. C \ C' \ pcp C' \ finite_character C'" proof(intro exI[of _ "C \ {S. \s \ S. finite s \ s \ C}"] conjI) let ?E = " {S. \s \ S. finite s \ s \ C}" show "C \ C \ ?E" by blast from S show "finite_character (C \ ?E)" unfolding finite_character_def subset_closed_def by blast note C'' = C[unfolded pcp_alt, THEN bspec] (* uniform notation. what did I learn? only slightly more elegant\ *) have CON: "G \ H \ S \ C \ {S. \s\S. finite s \ s \ C}" if si: "\s. \s\S; finite s\ \ s \ C" and un: "Con F G H" and el: " F \ S" for F G H S proof - have k: "\s \ S. finite s \ F \ s \ G \ H \ s \ C" using si un C'' by simp have "G \ H \ S \ ?E" unfolding mem_Collect_eq Un_iff proof safe fix s assume "s \ G \ H \ S" and f: "finite s" hence "F \ (s - {G,H}) \ S" using el by blast with k f have "G \ H \ F \ (s - {G,H}) \ C" by simp hence "F \ G \ H \ s \ C" using insert_absorb by fastforce thus "s \ C" using S unfolding subset_closed_def by fast qed thus "G \ H \ S \ C \ ?E" by simp qed have DIS: "G \ S \ C \ {S. \s\S. finite s \ s \ C} \ H \ S \ C \ {S. \s\S. finite s \ s \ C}" if si: "\s. s\S \ finite s \ s \ C" and un: "Dis F G H" and el: "F \ S" for F G H S proof - have l: "\I\{G, H}. I \ s1 \ C \ I \ s2 \ C" if "s1 \ S" "finite s1" "F \ s1" "s2 \ S" "finite s2" "F \ s2" for s1 s2 proof - let ?s = "s1 \ s2" have "?s \ S" "finite ?s" using that by simp_all with si have "?s \ C" by simp moreover have "F \ ?s" using that by simp ultimately have "\I\{G,H}. I \ ?s \ C" using C'' un by simp thus "\I\{G,H}. I \ s1 \ C \ I \ s2 \ C" by (meson S[unfolded subset_closed_def, THEN bspec] insert_mono sup.cobounded2 sup_ge1) qed have m: "\s1 \ S; finite s1; F \ s1; G \ s1 \ C; s2 \ S; finite s2; F \ s2; H \ s2 \ C\ \ False" for s1 s2 using l by blast have "False" if "s1 \ S" "finite s1" "G \ s1 \ C" "s2 \ S" "finite s2" "H \ s2 \ C" for s1 s2 proof - have *: "F \ s1 \ S" "finite (F \ s1)" "F \ F \ s1" if "s1 \ S" "finite s1" for s1 using that el by simp_all have "G \ F \ s1 \ C" "H \ F \ s2 \ C" by (meson S insert_mono subset_closed_def subset_insertI that(3,6))+ from m[OF *[OF that(1-2)] this(1) *[OF that(4-5)] this(2)] show False . qed hence "G \ S \ ?E \ H \ S \ ?E" unfolding mem_Collect_eq Un_iff by (metis (no_types, lifting) finite_Diff insert_Diff si subset_insert_iff) thus "G \ S \ C \ ?E \ H \ S \ C \ ?E" by blast qed (* this proof does benefit from uniform notation a bit. Before it was introduced, the subclaims CON, DIS had to be stated as *) have CON': "\f2 g2 h2 F2 G2 S2. \\s. \s \ C; h2 F2 G2 \ s\ \ f2 F2 \ s \ C \ g2 G2 \ s \ C; \s\S2. finite s \ s \ C; h2 F2 G2 \ S2; False\ \ f2 F2 \ S2 \ C \ {S. \s\S. finite s \ s \ C} \ g2 G2 \ S2 \ C \ {S. \s\S. finite s \ s \ C}" by fast (* (without the False, obviously). The proof of the subclaim does not change gravely, but the f2 g2 h2 have to be instantiated manually upon use (with, e.g., id Not (\F G. \<^bold>\(F \<^bold>\ G))), and there were multiple working instantiations. Generally not as beautiful. *) show "pcp (C \ ?E)" unfolding pcp_alt apply(intro ballI conjI; elim UnE; (unfold mem_Collect_eq)?) subgoal using C'' by blast subgoal using C'' by blast subgoal using C'' by (simp;fail) subgoal by (meson C'' empty_subsetI finite.emptyI finite_insert insert_subset subset_insertI) subgoal using C'' by simp subgoal using CON by simp subgoal using C'' by blast subgoal using DIS by simp done qed primrec pcp_seq where "pcp_seq C S 0 = S" | "pcp_seq C S (Suc n) = ( let Sn = pcp_seq C S n; Sn1 = from_nat n \ Sn in if Sn1 \ C then Sn1 else Sn )" lemma pcp_seq_in: "pcp C \ S \ C \ pcp_seq C S n \ C" proof(induction n) case (Suc n) hence "pcp_seq C S n \ C" by simp thus ?case by(simp add: Let_def) qed simp lemma pcp_seq_mono: "n \ m \ pcp_seq C S n \ pcp_seq C S m" proof(induction m) case (Suc m) thus ?case by(cases "n = Suc m"; simp add: Let_def; blast) qed simp lemma pcp_seq_UN: "\{pcp_seq C S n|n. n \ m} = pcp_seq C S m" proof(induction m) case (Suc m) have "{f n |n. n \ Suc m} = f (Suc m) \ {f n |n. n \ m}" for f using le_Suc_eq by auto hence "{pcp_seq C S n |n. n \ Suc m} = pcp_seq C S (Suc m) \ {pcp_seq C S n |n. n \ m}" . hence "\{pcp_seq C S n |n. n \ Suc m} = \{pcp_seq C S n |n. n \ m} \ pcp_seq C S (Suc m)" by auto thus ?case using Suc pcp_seq_mono by blast qed simp lemma wont_get_added: "(F :: ('a :: countable) formula) \ pcp_seq C S (Suc (to_nat F)) \ F \ pcp_seq C S (Suc (to_nat F) + n)" text\We don't necessarily have @{term "n = to_nat (from_nat n)"}, so this doesn't hold.\ oops definition "pcp_lim C S \ \{pcp_seq C S n|n. True}" lemma pcp_seq_sub: "pcp_seq C S n \ pcp_lim C S" unfolding pcp_lim_def by(induction n; blast) lemma pcp_lim_inserted_at_ex: "x \ pcp_lim C S \ \k. x \ pcp_seq C S k" unfolding pcp_lim_def by blast lemma pcp_lim_in: assumes c: "pcp C" and el: "S \ C" and sc: "subset_closed C" and fc: "finite_character C" shows "pcp_lim C S \ C" (is "?cl \ C") proof - from pcp_seq_in[OF c el, THEN allI] have "\n. pcp_seq C S n \ C" . hence "\m. \{pcp_seq C S n|n. n \ m} \ C" unfolding pcp_seq_UN . have "\s \ ?cl. finite s \ s \ C" proof safe fix s :: "'a formula set" have "pcp_seq C S (Suc (Max (to_nat ` s))) \ pcp_lim C S" using pcp_seq_sub by blast assume \finite s\ \s \ pcp_lim C S\ hence "\k. s \ pcp_seq C S k" proof(induction s rule: finite_induct) case (insert x s) - hence "\k. s \ pcp_seq C S k" by fast - then guess k1 .. + then obtain k1 where "s \ pcp_seq C S k1" by blast moreover obtain k2 where "x \ pcp_seq C S k2" by (meson pcp_lim_inserted_at_ex insert.prems insert_subset) ultimately have "x \ s \ pcp_seq C S (max k1 k2)" by (meson pcp_seq_mono dual_order.trans insert_subset max.bounded_iff order_refl subsetCE) thus ?case by blast qed simp with pcp_seq_in[OF c el] sc show "s \ C" unfolding subset_closed_def by blast qed thus "?cl \ C" using fc unfolding finite_character_def by blast qed lemma cl_max: assumes c: "pcp C" assumes sc: "subset_closed C" assumes el: "K \ C" assumes su: "pcp_lim C S \ K" shows "pcp_lim C S = K" (is ?e) proof (rule ccontr) assume \\?e\ with su have "pcp_lim C S \ K" by simp then obtain F where e: "F \ K" and ne: "F \ pcp_lim C S" by blast from ne have "F \ pcp_seq C S (Suc (to_nat F))" using pcp_seq_sub by fast hence 1: "F \ pcp_seq C S (to_nat F) \ C" by (simp add: Let_def split: if_splits) have "F \ pcp_seq C S (to_nat F) \ K" using pcp_seq_sub e su by blast hence "F \ pcp_seq C S (to_nat F) \ C" using sc unfolding subset_closed_def using el by blast with 1 show False .. qed lemma cl_max': assumes c: "pcp C" assumes sc: "subset_closed C" shows "F \ pcp_lim C S \ C \ F \ pcp_lim C S" "F \ G \ pcp_lim C S \ C \ F \ pcp_lim C S \ G \ pcp_lim C S" using cl_max[OF assms] by blast+ lemma pcp_lim_Hintikka: assumes c: "pcp C" assumes sc: "subset_closed C" assumes fc: "finite_character C" assumes el: "S \ C" shows "Hintikka (pcp_lim C S)" proof - let ?cl = "pcp_lim C S" have "?cl \ C" using pcp_lim_in[OF c el sc fc] . from c[unfolded pcp_alt, THEN bspec, OF this] have d: "\ \ ?cl" "Atom k \ ?cl \ \<^bold>\ (Atom k) \ ?cl \ False" "Con F G H \ F \ ?cl \ G \ H \ ?cl \ C" "Dis F G H \ F \ ?cl \ G \ ?cl \ C \ H \ ?cl \ C" for k F G H by blast+ have "Con F G H \ F \ ?cl \ G \ ?cl \ H \ ?cl" "Dis F G H \ F \ ?cl \ G \ ?cl \ H \ ?cl" for F G H by(auto dest: d(3-) cl_max'[OF c sc]) with d(1,2) show ?thesis unfolding Hintikka_alt by fast qed theorem pcp_sat: \ \model existence theorem\ fixes S :: "'a :: countable formula set" assumes c: "pcp C" assumes el: "S \ C" shows "sat S" proof - note [[show_types]] from c obtain Ce where "C \ Ce" "pcp Ce" "subset_closed Ce" "finite_character Ce" using ex1[where 'a='a] ex2[where 'a='a] ex3[where 'a='a] by (meson dual_order.trans ex2) have "S \ Ce" using \C \ Ce\ el .. with pcp_lim_Hintikka \pcp Ce\ \subset_closed Ce\ \finite_character Ce\ have "Hintikka (pcp_lim Ce S)" . with Hintikkas_lemma have "sat (pcp_lim Ce S)" . moreover have "S \ pcp_lim Ce S" using pcp_seq.simps(1) pcp_seq_sub by fast ultimately show ?thesis unfolding sat_def by fast qed (* This and Hintikka's lemma are the only two where we need semantics. Still, I don't think it's meaningful to separate those two into an extra theory. *) end diff --git a/thys/Propositional_Proof_Systems/LSC_Resolution.thy b/thys/Propositional_Proof_Systems/LSC_Resolution.thy --- a/thys/Propositional_Proof_Systems/LSC_Resolution.thy +++ b/thys/Propositional_Proof_Systems/LSC_Resolution.thy @@ -1,232 +1,233 @@ subsection\Converting between Resolution and SC proofs\ theory LSC_Resolution imports LSC Resolution begin lemma literal_subset_sandwich: assumes "is_lit_plus F" "cnf F = {C}" "R \ C" shows "R = \ \ R = C" using assms by(cases F rule: is_lit_plus.cases; simp) blast+ (* proof somewhat strange internally\ *) text\Proof following Gallier~\cite{gallier2015logic}.\ theorem CSC_Resolution_pre: "\ \\<^sub>n \ \\ \ set_mset \. is_cnf \ \ (\(cnf ` set_mset \)) \ \" proof(induction rule: LSC.induct) case (Ax k \) let ?s = "\(cnf ` set_mset (\<^bold>\ (Atom k), Atom k, \))" have "?s \ {k\<^sup>+}" "?s \ {k\}" using Resolution.Ass[where 'a='a] by simp_all from Resolution.R[OF this, of k] have "?s \ \" by simp thus ?case by simp next case (BotL \) thus ?case by(simp add: Ass) next case (AndL F G \) hence "\(cnf ` set_mset (F, G, \)) \ \" by simp thus ?case by(simp add: Un_left_commute sup.assoc) next (* The idea for the whole trickery with F being a literal (is_disj has to only allows right deep formulas) and the sandwiching is from Gallier, but mentioned there only in one little sentence. *) case (OrL F \ G) hence "is_cnf (F \<^bold>\ G)" by simp hence d: "is_disj (F \<^bold>\ G)" by simp hence db: "is_disj F" "is_lit_plus F" "is_disj G" by (-, cases F) simp_all hence "is_cnf F \ is_cnf G" by(cases F; cases G; simp) with OrL have IH: "(\(cnf ` set_mset (F, \))) \ \" "(\(cnf ` set_mset (G, \))) \ \" by simp_all let ?\ = "(\(cnf ` set_mset \))" from IH have IH_readable: "cnf F \ ?\ \ \" "cnf G \ ?\ \ \" by auto show ?case proof(cases "cnf F = {} \ cnf G = {}") case True hence "cnf (F \<^bold>\ G) = {}" by auto thus ?thesis using True IH by auto next case False then obtain S T where ST: "cnf F = {S}" "cnf G = {T}" using cnf_disj_ex db(1,3) (* try applying meson here. It's weird. and sledgehammer even suggests it. *) by metis (* hint: card S \ 1 *) hence R: "cnf (F \<^bold>\ G) = { S \ T }" by simp have "\S\?\ \ \; T\?\ \ \\ \ S \ T\ ?\ \ \" proof - assume s: "S\?\ \ \" and t: "T\?\ \ \" hence s_w: "S \ S \ T \ ?\ \ \" using Resolution_weaken by (metis insert_commute insert_is_Un) note Resolution_taint_assumptions[of "{T}" ?\ "\" S] t then obtain R where R: "S \ T \ \(cnf ` set_mset \) \ R" "R\S" by (auto simp: Un_commute) show ?thesis using literal_subset_sandwich[OF db(2) ST(1) R(2)] proof assume "R = \" thus ?thesis using R(1) by blast next from Resolution_unnecessary[where T="{_}", simplified] R(1) have "(R \ S \ T \ ?\ \ \) = (S \ T \ ?\ \ \)" . moreover assume "R = S" ultimately show ?thesis using s_w by simp qed qed thus ?thesis using IH ST R by simp qed hence case_readable: "cnf (F \<^bold>\ G) \ ?\ \ \" by auto qed auto corollary LSC_Resolution: assumes "\ \\<^sub>n" shows "(\(cnf ` nnf ` set_mset \)) \ \" proof - from assms have "image_mset nnf \ \\<^sub>n" by (simp add: LSC_NNF) from LSC_cnf[OF this] have "image_mset (cnf_form_of \ nnf) \ \\<^sub>n" by(simp add: image_mset.compositionality is_nnf_nnf) moreover have "\\ \ set_mset (image_mset (cnf_form_of \ nnf) \). is_cnf \" using cnf_form_of_is[where 'a='a, OF is_nnf_nnf] by simp moreover note CSC_Resolution_pre ultimately have "\(cnf ` set_mset (image_mset (cnf_form_of \ nnf) \)) \ \" by blast hence "\((\F. cnf (cnf_form_of (nnf F))) ` set_mset \) \ \" by simp thus ?thesis unfolding cnf_cnf[OF is_nnf_nnf] by simp qed corollary SC_Resolution: assumes "\ \ {#}" shows "(\(cnf ` nnf ` set_mset \)) \ \" proof - from assms have "image_mset nnf \ \\<^sub>n" by (simp add: LSC_NNF SC_LSC) hence "\(cnf` nnf ` set_mset (image_mset nnf \)) \ \" using LSC_Resolution by blast thus ?thesis using is_nnf_nnf_id[where 'a='a] is_nnf_nnf[where 'a='a] by auto qed (* Gallier just goes "Any resolution refutation can be transformed to a derivation in SCNF'" But we don't know what a resolution refutation is, inductively speaking. *) (* would have been a bit nicer with orderings on formula to convert mset \ list *) lemma Resolution_LSC_pre: assumes "S \ R" assumes "finite R" assumes "finite S" "Ball S finite" shows "\S' R'. \\. set R' = R \ set (map set S') = S \ (disj_of_clause R', {#disj_of_clause c. c \# mset S'#} + \ \\<^sub>n \ {#disj_of_clause c. c \# mset S'#} + \ \\<^sub>n)" (* order of quantifiers is important here\ *) using assms proof(induction S R rule: Resolution.induct) case (Ass F S) (* Idea: we don't just obtain an S', we obtain an S' that contains R (and not just some reordering of it) *) define Sm where "Sm = S - {F}" hence Sm: "S = F\Sm" "F \ Sm" using Ass by fast+ (* try doing this with obtain\ *) with Ass have fsm: "finite Sm" "Ball Sm finite" by auto then obtain Sm' where "Sm = set (map set Sm')" by (metis (full_types) ex_map_conv finite_list) moreover obtain R' where [simp]: "F = set R'" using Ass finite_list by blast ultimately have S: "S = set (map set (R'#Sm'))" unfolding Sm by simp show ?case using LSC_Contract[where 'a='a] by(intro exI[where x="R'#Sm'"] exI[where x=R']) (simp add: S add_ac) (* this was the base case. ugh. *) next case (R S F G k) from R.prems have fin: "finite F" "finite G" by simp_all from R.IH(1)[OF fin(1) R.prems(2,3)] obtain FR FS where IHF: "set FR = F" "set (map set FS) = S" "\\ GS. (disj_of_clause FR, image_mset disj_of_clause (mset (FS@GS)) + \ \\<^sub>n \ image_mset disj_of_clause (mset (FS@GS)) + \ \\<^sub>n)" by simp (metis add.assoc) from R.IH(2)[OF fin(2) R.prems(2,3)] obtain GR GS where IHG: "set GR = G" "set (map set GS) = S" "\\ HS. (disj_of_clause GR, image_mset disj_of_clause (mset (GS@HS)) + \ \\<^sub>n \ image_mset disj_of_clause (mset (GS@HS)) + \ \\<^sub>n)" by simp (metis add.assoc) have IH: "image_mset disj_of_clause (mset (FS @ GS)) + \ \\<^sub>n" if "disj_of_clause FR, disj_of_clause GR, image_mset disj_of_clause (mset (FS @ GS)) + \ \\<^sub>n" for \ using IHF(3)[of GS \] IHG(3)[of FS "disj_of_clause FR, \"] that by(simp add: add_mset_commute add_ac) show ?case apply(intro exI[where x="FS @ GS"] exI[where x="removeAll (k\<^sup>+) FR @ removeAll (k\) GR"] allI impI conjI) apply(simp add: IHF IHG;fail) apply(insert IHF IHG; simp;fail) apply(intro IH) apply(auto dest!: LSC_Sim_resolution_la simp add: IHF IHG R.hyps) done qed lemma Resolution_LSC_pre_nodisj: (* I've tried showing this directly instead of Resolution_LSC_pre, but that is surprisingly painful. *) assumes "S \ R" assumes "finite R" assumes "finite S" "Ball S finite" shows "\S' R'. \\. is_nnf_mset \ \ is_disj R' \ is_nnf S' \ cnf R' = {R} \ cnf S' \ S \ (R', S', \ \\<^sub>n \ S', \ \\<^sub>n)" proof - have mehorder: "F, \<^bold>\G, \ = \<^bold>\G, F, \" for F G \ by(simp add: add_ac) from Resolution_LSC_pre[where 'a='a,OF assms] obtain S' R' where o: "\\. is_nnf_mset \ \ set R' = R \ set (map set S') = S \ (disj_of_clause R', image_mset disj_of_clause (mset S') + \ \\<^sub>n \ image_mset disj_of_clause (mset S') + \ \\<^sub>n)" by blast hence p: "is_nnf_mset \ \ (disj_of_clause R', image_mset disj_of_clause (mset S') + \ \\<^sub>n \ image_mset disj_of_clause (mset S') + \ \\<^sub>n)" for \ by blast show ?thesis apply(rule exI[where x="\<^bold>\map disj_of_clause S'"]) apply(rule exI[where x="disj_of_clause R'"]) apply safe apply(intro disj_of_clause_is;fail) apply(simp add: cnf_disj o; fail)+ subgoal using o by(fastforce simp add: cnf_BigAnd cnf_disj) subgoal for \ apply(frule p) apply(unfold mehorder) apply(drule LSC_BigAndL_inv) apply(simp;fail)+ by (simp add: LSC_BigAndL) done qed corollary Resolution_LSC1: assumes "S \ \" shows "\F. is_nnf F \ cnf F \ S \ {#F#} \\<^sub>n" proof - have *: "{f \ g |f g. f \ F \ g \ G} = {\} \ F = {\}" for F G proof (rule ccontr) assume m: "{f \ g |f g. f \ F \ g \ G} = {\}" assume "F \ {\}" hence "F = {} \ (\E. E \ F \ E \ \)" by blast thus False proof assume "F = {}" with m show False by simp next - assume "\E. E \ F \ E \ \" then guess E .. note E = this + assume "\E. E \ F \ E \ \" + then obtain E where E: "E \ F \ E \ \" .. show False proof cases assume "G = {}" with m show False by simp next assume "G \ {}" then obtain D where "D \ G" by blast with E have "E \ D \ {f \ g |f g. f \ F \ g \ G}" by blast with m E show False by simp qed qed qed have *: "F = {\} \ G = {\}" if "{f \ g |f g. f \ F \ g \ G} = {\}" for F G proof (intro conjI) show "G = {\}" apply(rule *[of G F]) apply(subst that[symmetric]) by blast qed (rule *[OF that]) have *: "is_nnf F \ is_nnf_mset \ \ cnf F = {\} \ F,\ \\<^sub>n" for F \ apply(induction F rule: cnf.induct; simp) apply blast apply (metis LSC.LSC.AndL LSC_weaken add_mset_commute singleton_Un_iff) apply(drule *; simp add: LSC.LSC.OrL) done from Resolution_useless_infinite[OF assms] obtain S' where su: "S'\S" and fin: "finite S'" "Ball S' finite" and pr: "(S' \ \)" by blast from Resolution_LSC_pre_nodisj[OF pr finite.emptyI fin] obtain S'' where "is_nnf S''" "cnf S'' \ S'" "{# S'' #} \\<^sub>n" using * [OF disj_is_nnf, of _ \{#}\] by (metis LSC_weaken add_mset_commute empty_iff set_mset_empty) with su show ?thesis by blast qed corollary Resolution_SC1: assumes "S \ \" shows "\F. cnf (nnf F) \ S \ {#F#} \ {#}" apply(insert Resolution_LSC1[OF assms]) apply(elim ex_forward) apply(elim conjE; intro conjI) subgoal by(subst is_nnf_nnf_id; assumption) apply(unfold SC_LSC) subgoal by (simp;fail) done (* really, these proofs have to do with sets, multisets and lists. If I'd introduce finite sets somehow, the chaos would be perfect. I want out. *) end diff --git a/thys/Propositional_Proof_Systems/ND_Compl_Truthtable_Compact.thy b/thys/Propositional_Proof_Systems/ND_Compl_Truthtable_Compact.thy --- a/thys/Propositional_Proof_Systems/ND_Compl_Truthtable_Compact.thy +++ b/thys/Propositional_Proof_Systems/ND_Compl_Truthtable_Compact.thy @@ -1,17 +1,17 @@ theory ND_Compl_Truthtable_Compact imports ND_Compl_Truthtable Compactness begin theorem fixes \ :: "'a :: countable formula set" shows "\ \ F \ \ \ F" proof - assume \\ \ F\ - with compact_to_formula guess G . + then obtain G where "set G \ \" "\ \<^bold>\G \<^bold>\ F" by (rule compact_to_formula) from ND_complete \\ \<^bold>\G \<^bold>\ F\ have \{} \ \<^bold>\G \<^bold>\ F\ . with AssmBigAnd have \set G \ F\ .. with Weaken show ?thesis using \set G \ \\ . qed end diff --git a/thys/Propositional_Proof_Systems/ND_FiniteAssms.thy b/thys/Propositional_Proof_Systems/ND_FiniteAssms.thy --- a/thys/Propositional_Proof_Systems/ND_FiniteAssms.thy +++ b/thys/Propositional_Proof_Systems/ND_FiniteAssms.thy @@ -1,81 +1,89 @@ theory ND_FiniteAssms imports ND begin lemma ND_finite_assms: "\ \ F \ \\'. \' \ \ \ finite \' \ (\' \ F)" proof(induction rule: ND.induct) case (Ax F \) thus ?case by(intro exI[of _ "{F}"]) (simp add: ND.Ax) next case (AndI \ F G) - from AndI.IH(1) guess \1 .. moreover - from AndI.IH(2) guess \2 .. ultimately - show ?case by(intro exI[where x="\1\\2"]) (force elim: Weaken intro!: ND.AndI) + from AndI.IH obtain \1 \2 + where "\1 \ \ \ finite \1 \ (\1 \ F)" + and "\2 \ \ \ finite \2 \ (\2 \ G)" + by blast + then show ?case by(intro exI[where x="\1\\2"]) (force elim: Weaken intro!: ND.AndI) next case (CC F \) - from CC.IH guess \' .. note \' = this + from CC.IH obtain \' where \': "\' \ \<^bold>\ F \ \ \ finite \' \ (\' \ \)" .. thus ?case proof(cases "Not F \ \'") text\case distinction: Did we actually use @{term "\<^bold>\F"}?\ case False hence "\' \ \" using \' by blast with \' show ?thesis using BotE by(intro exI[where x="\'"]) fast next case True then obtain \'' where "\' = \<^bold>\ F\\''" "\<^bold>\ F \ \''" by (meson Set.set_insert) hence "\'' \ \" "finite \''" "\<^bold>\ F\\'' \ \" using \' by auto thus ?thesis using ND.CC by auto qed next case AndE1 thus ?case by(blast dest: ND.AndE1) next case AndE2 thus ?case by(blast dest: ND.AndE2) next case OrI1 thus ?case by(blast dest: ND.OrI1) next case OrI2 thus ?case by(blast dest: ND.OrI2) next case (OrE \ F G H) - from OrE.IH(1) guess \1 .. moreover - from OrE.IH(2) guess \2 .. moreover - from OrE.IH(3) guess \3 .. - note IH = calculation this + from OrE.IH obtain \1 \2 \3 + where IH: + "\1 \ \ \ finite \1 \ (\1 \ F \<^bold>\ G)" + "\2 \ F \ \ \ finite \2 \ (\2 \ H)" + "\3 \ G \ \ \ finite \3 \ (\3 \ H)" + by blast let ?w = "\1 \ (\2 - {F}) \ (\3 - {G})" from IH have "?w \ F \<^bold>\ G" using Weaken[OF _ sup_ge1] by metis moreover from IH have "F\?w \ H" "G\?w \ H" using Weaken by (metis Un_commute Un_insert_right Un_upper1 Weaken insert_Diff_single)+ ultimately have "?w \ H" using ND.OrE by blast thus ?case using IH by(intro exI[where x="?w"]) auto text\Clever evasion of the case distinction made for CC.\ next case (ImpI F \ G) - from ImpI.IH guess \' .. + from ImpI.IH obtain \' where "\' \ F \ \ \ finite \' \ (\' \ G)" .. thus ?case by (intro exI[where x="\' - {F}"]) (force elim: Weaken intro!: ND.ImpI) next case (ImpE \ F G) - from ImpE.IH(1) guess \1 .. moreover - from ImpE.IH(2) guess \2 .. ultimately - show ?case by(intro exI[where x="\1 \ \2"]) (force elim: Weaken intro: ND.ImpE[where F=F]) + from ImpE.IH obtain \1 \2 where + "\1 \ \ \ finite \1 \ (\1 \ F \<^bold>\ G)" + "\2 \ \ \ finite \2 \ (\2 \ F)" + by blast + then show ?case by(intro exI[where x="\1 \ \2"]) (force elim: Weaken intro: ND.ImpE[where F=F]) next case (NotE \ F) - from NotE.IH(1) guess \1 .. moreover - from NotE.IH(2) guess \2 .. ultimately - show ?case by(intro exI[where x="\1 \ \2"]) (force elim: Weaken intro: ND.NotE[where F=F]) + from NotE.IH obtain \1 \2 where + "\1 \ \ \ finite \1 \ (\1 \ \<^bold>\ F)" + "\2 \ \ \ finite \2 \ (\2 \ F)" + by blast + then show ?case by(intro exI[where x="\1 \ \2"]) (force elim: Weaken intro: ND.NotE[where F=F]) next case (NotI F \) - from NotI.IH guess \' .. + from NotI.IH obtain \' where "\' \ F \ \ \ finite \' \ (\' \ \)" .. thus ?case by(intro exI[where x="\' - {F}"]) (force elim: Weaken intro: ND.NotI[where F=F]) qed text\We thought that a lemma like this would be necessary for the ND completeness by SC completeness proof (this lemma shows that if we made an ND proof, we can always limit ourselves to a finite set of assumptions -- and thus put all the assumptions into one formula). That is not the case, since in the completeness proof, we assume a valid entailment and have to show (the existence of) a derivation. The author hopes that his misunderstanding can help the reader's understanding.\ corollary ND_no_assms: assumes "\ \ F" obtains \' where "set \' \ \ \ ({} \ \<^bold>\\' \<^bold>\ F)" proof(goal_cases) case 1 from ND_finite_assms[OF assms] obtain \' where "\'\\" "finite \'" "\' \ F" by blast from \finite \'\ obtain G where \'[simp]: "\' = set G" using finite_list by blast with \\'\\\ have "set G \ \" by clarify moreover from \\' \ F\ have "{} \ \<^bold>\ G \<^bold>\ F" unfolding \' AssmBigAnd . ultimately show ?case by(intro 1[where \'=G] conjI) qed end diff --git a/thys/Propositional_Proof_Systems/SC_Depth.thy b/thys/Propositional_Proof_Systems/SC_Depth.thy --- a/thys/Propositional_Proof_Systems/SC_Depth.thy +++ b/thys/Propositional_Proof_Systems/SC_Depth.thy @@ -1,592 +1,593 @@ theory SC_Depth imports SC begin text\ Many textbook arguments about SC use the depth of the derivation tree as basis for inductions. We had originally thought that this is mandatory for the proof of contraction, but found out it is not. It remains unclear to us whether there is any proof on SC that requires an argument using depth. \ text\ We keep our formalization of SC with depth for didactic reasons: we think that arguments about depth do not need much meta-explanation, but structural induction and rule induction usually need extra explanation for students unfamiliar with Isabelle/HOL. They are also a lot harder to execute. We dare the reader to write down (a few of) the cases for, e.g. \AndL_inv'\, by hand. \ (* Warning: This theory may have name collisions with SC.thy\ *) inductive SCc :: "'a formula multiset \ 'a formula multiset \ nat \ bool" ("((_ \/ _) \ _)" [53,53] 53) where BotL: "\ \# \ \ \\\ \ Suc n" | Ax: "Atom k \# \ \ Atom k \# \ \ \\\ \ Suc n" | NotL: "\ \ F,\ \ n \ Not F, \ \ \ \ Suc n" | NotR: "F,\ \ \ \ n \ \ \ Not F, \ \ Suc n" | AndL: "F,G,\ \ \ \ n \ And F G, \ \ \ \ Suc n" | AndR: "\ \ \ F,\ \ n; \ \ G,\ \ n \ \ \ \ And F G, \ \ Suc n" | OrL: "\ F,\ \ \ \ n; G,\ \ \ \ n \ \ Or F G, \ \ \ \ Suc n" | OrR: "\ \ F,G,\ \ n \ \ \ Or F G, \ \ Suc n" | ImpL: "\ \ \ F,\ \ n; G,\ \ \ \ n \ \ Imp F G, \ \ \ \ Suc n" | ImpR: "F,\ \ G,\ \ n \ \ \ Imp F G, \ \ Suc n" lemma shows BotL_canonical'[intro!]: "\,\\\ \ Suc n" and Ax_canonical'[intro!]: "Atom k,\ \ Atom k,\ \ Suc n" by (meson SCc.intros union_single_eq_member)+ lemma deeper: "\ \ \ \ n \ \ \ \ \ n + m" by(induction rule: SCc.induct; insert SCc.intros; auto) lemma deeper_suc: "\ \ \ \ n \ \ \ \ \ Suc n" (* this version is actually required more often. It can be declared as an elim!, but I want to make its usage explicit *) thm deeper[unfolded Suc_eq_plus1[symmetric]] by(drule deeper[where m=1]) simp text\The equivalence is obvious.\ theorem SC_SCp_eq: fixes \ \ :: "'a formula multiset" shows "(\n. \ \ \ \ n) \ \ \ \" (is "?c \ ?p") proof - assume ?c then guess n .. + assume ?c + then obtain n where "\ \ \ \ n" .. thus ?p by(induction rule: SCc.induct; simp add: SCp.intros) next have deeper_max: "\ \ \ \ max m n" "\ \ \ \ max n m" if "\ \ \ \ n" for n m and \ \ :: "'a formula multiset" proof - have "n \ m \ \k. m = n + k" by presburger with that[THEN deeper] that show "\ \ \ \ max n m" unfolding max_def by clarsimp thus "\ \ \ \ max m n" by (simp add: max.commute) qed assume ?p thus ?c by(induction rule: SCp.induct) (insert SCc.intros[where 'a='a] deeper_max; metis)+ qed (* it makes a difference whether we say Ax is 0 or any: with Ax \ 0 we could not prove the deepening rules. Also, it is important to allow only atoms in Ax, otherwise the inv-rules are not depth preserving. Making Ax/BotL start from \1 saves proving the base-cases twice *) lemma no_0_SC[dest!]: "\ \ \ \ 0 \ False" by(cases rule: SCc.cases[of \ \ 0]) assumption lemma inR1': "\ \ G, H, \ \ n \ \ \ H, G, \ \ n" by(simp add: add_mset_commute) lemma inL1': "G, H, \ \ \ \ n \ H, G, \ \ \ \ n" by(simp add: add_mset_commute) lemma inR2': "\ \ F, G, H, \ \ n \ \ \ G, H, F, \ \ n" by(simp add: add_mset_commute) lemma inL2': "F, G, H, \ \ \ \ n \ G, H, F, \ \ \ \ n" by(simp add: add_mset_commute) lemma inR3': "\ \ F, G, H, \ \ n \ \ \ H, F, G, \ \ n" by(simp add: add_mset_commute) lemma inR4': "\ \ F, G, H, I, \ \ n \ \ \ H, I, F, G, \ \ n" by(simp add: add_mset_commute) lemma inL3': "F, G, H, \ \ \ \ n \ H, F, G, \ \ \ \ n" by(simp add: add_mset_commute) lemma inL4': "F, G, H, I, \ \ \ \ n \ H, I, F, G, \ \ \ \ n" by(simp add: add_mset_commute) lemmas SC_swap_applies[intro,elim!] = inL1' inL2' inL3' inL4' inR1' inR2' inR3' inR4' (* these are here because they can be used for more careful reasoning with intro *) lemma "Atom C \<^bold>\ Atom D \<^bold>\ Atom E, Atom k \<^bold>\ Atom C \<^bold>\ Atom D, Atom k, {#} \ {# Atom E #} \ Suc (Suc (Suc (Suc (Suc 0))))" by(auto intro!: SCc.intros(3-) intro: SCc.intros(1,2)) lemma Bot_delR': "\ \ \ \ n \ \ \ \-{#\#} \ n" proof(induction rule: SCc.induct) case BotL thus ?case by(rule SCc.BotL; simp) (* SC.BotL refers to SC.SCp.BotL. Even more interestingly, there used to be a SCc.SC.BotL (which was not reffered to). I'll just leave this here and see how long till someone breaks it. *) next case (Ax k) thus ?case by(intro SCc.Ax[of k]; simp; metis diff_single_trivial formula.distinct(1) insert_DiffM lem1) next case NotL thus ?case using SCc.NotL by (metis add_mset_remove_trivial diff_single_trivial diff_union_swap insert_DiffM) next case NotR thus ?case using SCc.NotR by (metis diff_union_swap formula.distinct(11)) next case AndR thus ?case using SCc.AndR by (metis diff_single_trivial diff_union_swap diff_union_swap2 formula.distinct(13)) next case OrR thus ?case using SCc.OrR by (metis diff_single_trivial diff_union_swap2 formula.distinct(15) insert_iff set_mset_add_mset_insert) next case ImpL thus ?case using SCc.ImpL by (metis diff_single_trivial diff_union_swap2) next case ImpR thus ?case using SCc.ImpR by (metis diff_single_trivial diff_union_swap diff_union_swap2 formula.distinct(17)) qed (simp_all add: SCc.intros) lemma NotL_inv': "Not F, \ \ \ \ n \ \ \ F,\ \ n" proof(induction "Not F, \" \ n arbitrary: \ rule: SCc.induct) case (NotL \' G \ n) thus ?case by(cases "Not F = Not G") (auto intro!: SCc.intros(3-) dest!: multi_member_split simp: lem1 lem2 deeper_suc) qed (auto intro!: SCc.intros(3-) dest!: multi_member_split simp: SCp.intros lem1 lem2) lemma AndL_inv': "And F G, \ \ \ \ n \ F,G,\ \ \ \ n" proof(induction "And F G, \" \ n arbitrary: \ rule: SCc.induct) case (AndL F' G' \' \) thus ?case by(cases "And F G = And F' G'"; auto intro!: SCc.intros(3-) dest!: multi_member_split simp: lem1 lem2 deeper_suc; metis add_mset_commute) qed (auto intro!: SCc.intros(3-) dest!: multi_member_split simp: SCc.intros lem1 lem2 inL2') lemma OrL_inv': assumes "Or F G, \ \ \ \ n" shows "F,\ \ \ \ n \ G,\ \ \ \ n" proof(insert assms, induction "Or F G, \" \ n arbitrary: \ rule: SCc.induct) case (OrL F' \' \ n G') thus ?case by(cases "Or F G = Or F' G'"; auto intro!: SCc.intros(3-) dest!: multi_member_split simp: lem1 lem2 deeper_suc; blast) qed (fastforce intro!: SCc.intros(3-) dest!: multi_member_split simp: SCc.intros lem1 lem2)+ lemma ImpL_inv': assumes "Imp F G, \ \ \ \ n" shows "\ \ F,\ \ n \ G,\ \ \ \ n" proof(insert assms, induction "Imp F G, \" \ n arbitrary: \ rule: SCc.induct) case (ImpL \' F' \ n G') thus ?case by(cases "Or F G = Or F' G'"; (* oops, I didn't pay attention and used the wrong constructor\ doesn't matter! *) auto intro!: SCc.intros(3-) dest!: multi_member_split simp: lem1 lem2 deeper_suc; blast) qed (fastforce intro!: SCc.intros(3-) dest!: multi_member_split simp: SCc.intros lem1 lem2)+ lemma ImpR_inv': assumes "\ \ Imp F G,\ \ n" shows "F,\ \ G,\ \ n" proof(insert assms, induction \ "Imp F G, \" n arbitrary: \ rule: SCc.induct) case (ImpR F' \ G' \') thus ?case by(cases "Or F G = Or F' G'"; auto intro!: SCc.intros(3-) dest!: multi_member_split simp: lem1 lem2 deeper_suc; blast) qed (fastforce intro!: SCc.intros(3-) dest!: multi_member_split simp: SCc.intros lem1 lem2)+ lemma AndR_inv': shows "\ \ And F G, \ \ n \ \ \ F, \ \ n \ \ \ G, \ \ n" proof(induction \ "And F G, \" n arbitrary: \ rule: SCc.induct) case (AndR \ F' \' n G') thus ?case by(cases "Or F G = Or F' G'"; auto intro!: SCc.intros(3-) dest!: multi_member_split simp: lem1 lem2 deeper_suc; blast) qed (fastforce intro!: SCc.intros(3-) dest!: multi_member_split simp: SCc.intros lem1 lem2)+ lemma OrR_inv': "\ \ Or F G, \ \ n \ \ \ F,G,\ \ n" proof(induction \ "Or F G, \" n arbitrary: \ rule: SCc.induct) case (OrR \ F' G' \') thus ?case by(cases "Or F G = Or F' G'"; auto intro!: SCc.intros(3-) dest!: multi_member_split simp: lem1 lem2 deeper_suc; metis add_mset_commute) qed (fastforce intro!: SCc.intros(3-) dest!: multi_member_split simp: SCc.intros lem1 lem2)+ lemma NotR_inv': "\ \ Not F, \ \ n \ F,\ \ \ \ n" proof(induction \ "Not F, \" n arbitrary: \ rule: SCc.induct) case (NotR G \ \') thus ?case by(cases "Not F = Not G"; auto intro!: SCc.intros(3-) dest!: multi_member_split simp: lem1 lem2 deeper_suc; metis add_mset_commute) qed (fastforce intro!: SCc.intros(3-) dest!: multi_member_split simp: SCc.intros lem1 lem2)+ lemma weakenL': "\ \ \ \ n \ F,\ \ \ \ n" by(induction rule: SCc.induct) (auto intro!: SCc.intros(3-) intro: SCc.intros(1,2)) lemma weakenR': "\ \ \ \ n \ \ \ F,\ \ n" by(induction rule: SCc.induct) (auto intro!: SCc.intros(3-) intro: SCc.intros(1,2)) lemma contract': "(F,F,\ \ \ \ n \ F,\ \ \ \ n) \ (\ \ F,F,\ \ n \ \ \ F,\ \ n)" proof(induction n arbitrary: F \ \) case (Suc n) hence IH: "F, F, \ \ \ \ n \ F, \ \ \ \ n" "\ \ F, F, \ \ n \ \ \ F, \ \ n" for F :: "'a formula" and \ \ by blast+ show ?case proof(intro conjI allI impI, goal_cases) case 1 let ?ffs = "\\. \ - {# F #} - {# F #}" from 1 show ?case proof(insert 1; cases rule: SCc.cases[of "F,F,\" \ "Suc n"]) case (NotL \' G) show ?thesis proof(cases) assume e: "F = \<^bold>\G" with NotL have \': "\' = \<^bold>\G, \" by simp from NotL_inv' NotL(2) have "\ \ G, G, \ \ n" unfolding \' . with IH(2) have "\ \ G, \ \ n" . with SCc.NotL show ?thesis unfolding e . next assume e: "F \ \<^bold>\G" have ?thesis using NotL(2) IH(1)[of F "?ffs \'" "G, \"] SCc.NotL[of "F, \' - {# F #} - {# F #}" G \ n] using e NotL(1) by (metis (no_types, lifting) insert_DiffM lem2) from e NotL(1) have \: "\ = \<^bold>\ G, ?ffs \'" by (meson lem1) with NotL(1) have \': "F,F,?ffs \' = \'" by simp show ?thesis using NotL(2) IH(1)[of F "?ffs \'" "G, \"] SCc.NotL[of "F, ?ffs \'" G \ n] \F, \ \ \ \ Suc n\ by blast qed next case (AndL G H \') show ?thesis proof cases assume e: "F = And G H" with AndL(1) have \': "\' = And G H, \" by simp have "G \<^bold>\ H, G, H, \ \ \ \ n" using AndL(2) unfolding \' by auto hence "G, H, G, H, \ \ \ \ n" by(rule AndL_inv') hence "G, H, \ \ \ \ n" using IH(1) by (metis inL1' inL3') thus "F, \ \ \ \ Suc n" using e SCc.AndL by simp next assume ne: "F \ And G H" with AndL(1) have \: "\ = And G H, ?ffs \'" by (metis (no_types, lifting) diff_diff_add lem2) have "F, F, G, H, ?ffs \' \ \ \ n" using AndL(2) using \ inL4' local.AndL(1) by auto hence "G, H, F, ?ffs \' \ \ \ n" using IH(1) inL2 by blast thus ?thesis using SCc.AndL unfolding \ using inL1 by blast qed next case (OrL G \' H) show ?thesis proof cases assume e: "F = Or G H" with OrL(1) have \': "\' = Or G H, \" by simp have "Or G H, G, \ \ \ \ n" "Or G H, H, \ \ \ \ n" using OrL(2,3) unfolding \' by simp_all hence "G, G, \ \ \ \ n" "H, H, \ \ \ \ n" using OrL_inv' by blast+ hence "G, \ \ \ \ n" "H, \ \ \ \ n" using IH(1) by presburger+ thus "F, \ \ \ \ Suc n" unfolding e using SCc.OrL by blast next assume ne: "F \ Or G H" with OrL(1) have \: "\ = Or G H, ?ffs \'" by (metis (no_types, lifting) diff_diff_add lem2) have "F, F, G, ?ffs \' \ \ \ n" "F, F, H,?ffs \' \ \ \ n" using OrL \ by auto hence "G, F, ?ffs \' \ \ \ n" "H, F, ?ffs \' \ \ \ n" using IH(1) by(metis add_mset_commute)+ thus ?thesis using SCc.OrL unfolding \ by auto qed next case (ImpL \' G H) show ?thesis proof cases assume e: "F = Imp G H" with ImpL(1) have \': "\' = Imp G H, \" by simp have "H, \ \ \ \ n" "\ \ G,\ \ n" using IH ImpL_inv' ImpL(2,3) unfolding \' by (metis add_mset_commute)+ thus ?thesis unfolding e using SCc.ImpL[where 'a='a] by simp next assume ne: "F \ Imp G H" with ImpL(1) have \: "\ = Imp G H, ?ffs \'" by (metis (no_types, lifting) diff_diff_add lem2) have "F, F, ?ffs \' \ G, \ \ n" "F, F, H, ?ffs \' \ \ \ n" using ImpL \ by auto thus ?thesis using SCc.ImpL IH unfolding \ by (metis inL1') qed next case ImpR thus ?thesis by (simp add: IH(1) SCc.intros(10) add_mset_commute) next case (NotR G \') thus ?thesis using IH(1) by (simp add: SCc.NotR add_mset_commute) qed (auto intro: IH SCc.intros(1,2) intro!: SCc.intros(3-10)) next case 2 let ?ffs = "\\. \ - {# F #} - {# F #}" have not_principal[dest]: "\F \ f G H; F, F, \ = f G H, \'\ \ \ = f G H, ?ffs \' \ F, F, ?ffs \' = \'" for f G H \' proof(intro conjI, goal_cases) case 2 from 2 have "F \# \'" by(blast dest: lem1[THEN iffD1]) then obtain \'' where \': "\' = F,\''" by (metis insert_DiffM) with 2(2) have "F, \ = f G H, \''" by(simp add: add_mset_commute) hence "F \# \''" using 2(1) by(blast dest: lem1[THEN iffD1]) then obtain \''' where \'': "\'' = F,\'''" by (metis insert_DiffM) show ?case unfolding \' \'' by simp case 1 show ?case using 1(2) unfolding \' \'' by(simp add: add_mset_commute) qed have principal[dest]: "F, F, \ = f G H, \' \ F = f G H \ \' = f G H, \" for f G H \' by simp from 2 show ?case proof(cases rule: SCc.cases[of \ "F,F,\" "Suc n"]) case (ImpR G H \') thus ?thesis proof cases assume e[simp]: "F = Imp G H" with ImpR(1) have \'[simp]: "\' = Imp G H, \" by simp have "G, \ \ Imp G H, H, \ \ n" using ImpR(2) by simp hence "G, G, \ \ H, H, \ \ n" by(rule ImpR_inv') hence "G, \ \ H, \ \ n" using IH by fast thus "\ \ F, \ \ Suc n" using SCc.ImpR by simp next assume a: "F \ Imp G H" with ImpR(1) have \: "\ = Imp G H, ?ffs \'" by (metis (no_types, lifting) diff_diff_add lem2) have "G,\ \ F, F, H, ?ffs \' \ n" using ImpR \ by fastforce thus ?thesis using SCc.ImpR IH unfolding \ by (metis inR1') qed next case (AndR G \' H) thus ?thesis proof(cases "F = And G H") case True thus ?thesis using AndR by(auto intro!: SCc.intros(3-) dest!: AndR_inv' intro: IH) next case False hence e: "\ = And G H, ?ffs \'" using AndR(1) using diff_diff_add lem2 by blast hence "G \<^bold>\ H, F, F, ?ffs \' = G \<^bold>\ H, \'" using AndR(1) by simp hence "\ \ F, F, G, ?ffs \' \ n" "\ \ F, F, H, ?ffs \' \ n" using AndR(2,3) using add_left_imp_eq inR2 by fastforce+ hence "\ \ G, F, ?ffs \' \ n" "\ \ H, F, ?ffs \' \ n" using IH(2) by blast+ thus ?thesis unfolding e by(intro SCc.AndR[THEN inR1']) qed next case (OrR G H \') thus ?thesis proof cases assume a: "F = Or G H" hence \': "\' = G \<^bold>\ H, \" using OrR(1) by(intro principal) hence "\ \ G, H, G, H, \ \ n" using inR3'[THEN OrR_inv'] OrR(2) by auto hence "\ \ H, G, \ \ n" using IH(2)[of \ G "H,H,\"] IH(2)[of \ H "G,\"] unfolding add_ac(3)[of "{#H#}" "{#G#}"] using inR2 by blast hence "\ \ G, H, \ \ n" by(elim SC_swap_applies) thus ?thesis unfolding a by (simp add: SCc.OrR) next assume a: "F \ Or G H" with not_principal have np: "\ = G \<^bold>\ H, ?ffs \' \ F, F, ?ffs \' = \'" using OrR(1) . with OrR(2) have "\ \ G, H, F, ?ffs \' \ n" using IH(2) by (metis inR2' inR4') hence "\ \ F, G \<^bold>\ H, ?ffs \' \ Suc n" by(intro SCc.OrR[THEN inR1']) thus ?thesis using np by simp qed next case (NotR G \') thus ?thesis proof(cases "F = Not G") case True with principal NotR(1) have "\' = \<^bold>\ G, \" . with NotR_inv' NotR(2) have "G, G, \ \ \ \ n" by blast with IH(1) have "G, \ \ \ \ n" . thus "\ \ F, \ \ Suc n" unfolding True by(intro SCc.NotR) next case False with not_principal have np: "\ = \<^bold>\ G, \' - (F, {#F#}) \ F, F, \' - (F, {#F#}) = \'" using NotR(1) by auto hence "G, \ \ F, F, ?ffs \' \ n" using NotR(2) by simp hence "G, \ \ F, ?ffs \' \ n" by(elim IH(2)) thus ?thesis using np SCc.NotR inR1 by auto qed next case BotL thus ?thesis by(elim SCc.BotL) next case (Ax k) thus ?thesis by(intro SCc.Ax[where k=k]) simp_all next case NotL thus ?thesis by (simp add: SCc.NotL Suc.IH add_mset_commute) next case AndL thus ?thesis using SCc.AndL Suc.IH by blast next case OrL thus ?thesis using SCc.OrL Suc.IH by blast next case ImpL thus ?thesis by (metis SCc.ImpL Suc.IH add_mset_commute) qed qed qed blast (* depth for cut? *) lemma Cut_Atom_depth: "Atom k,\ \ \ \ n \ \ \ Atom k,\ \ m \ \ \ \ \ n + m" proof(induction "Atom k,\" "\" n arbitrary: \ m rule: SCc.induct) case (BotL \) hence "\ \# \" by simp thus ?case using SCc.BotL by auto next case (Ax l \) show ?case proof cases assume "l = k" with \Atom l \# \\ obtain \' where "\ = Atom k, \'" by (meson multi_member_split) with \\ \ Atom k, \ \ m\ have "\ \ \ \ m" using contract' by blast thus ?thesis by (metis add.commute deeper) next assume "l \ k" with \Atom l \# Atom k, \\ have "Atom l \# \" by simp with \Atom l \# \\ show ?thesis using SCc.Ax[of l] by simp qed next case (NotL \ F \) obtain \' where \: "\ = Not F, \'" by (meson NotL.hyps(3) add_eq_conv_ex formula.simps(9)) show ?case unfolding \ apply(unfold plus_nat.add_Suc) apply(intro SCc.NotL) apply(intro NotL.hyps (* IH *)) subgoal using NotL \ by (simp add: lem2) subgoal using \ NotL.prems NotL_inv' by blast done next case (NotR F \) then show ?case by(auto intro!: SCc.NotR dest!: NotR_inv') next case (AndL F G \ \) obtain \' where \: "\ = And F G, \'" by (metis AndL.hyps(3) add_eq_conv_diff formula.distinct(5)) show ?case unfolding \ apply(unfold plus_nat.add_Suc) apply(intro SCc.AndL) apply(intro AndL.hyps (* IH *)) subgoal using AndL \ by (simp add: lem2) subgoal using \ AndL.prems AndL_inv' by blast done next case (AndR F \ G) then show ?case using AndR_inv' SCc.AndR by (metis add_Suc inR1') next case (OrL F \' \ n G) obtain \'' where \: "\ = Or F G, \''" by (meson OrL.hyps(5) add_eq_conv_ex formula.simps(13)) have ihm: "F, \' = Atom k, F, \''" "G, \' = Atom k, G, \''" using OrL \ by (simp_all add: lem2) show ?case unfolding \ apply(unfold plus_nat.add_Suc) apply(intro SCc.OrL OrL.hyps(2)[OF ihm(1)] OrL.hyps(4)[OF ihm(2)]) subgoal using \ OrL.prems OrL_inv' by blast subgoal using \ OrL.prems OrL_inv' by blast done next case (OrR F G \) then show ?case by(auto intro!: SCc.intros(3-) dest!: OrR_inv') next case (ImpL \' F \ n G) obtain \'' where \: "\ = Imp F G, \''" by (metis ImpL.hyps(5) add_eq_conv_ex formula.simps) show ?case unfolding \ apply(unfold plus_nat.add_Suc) apply(intro SCc.ImpL ImpL.hyps(2) ImpL.hyps(4)) subgoal using ImpL \ by (simp add: lem2) subgoal using \ ImpL.prems by(auto dest!: ImpL_inv') subgoal using ImpL \ by (simp add: lem2) subgoal using \ ImpL.prems ImpL_inv' by blast done next case (ImpR F G \) then show ?case by (auto dest!: ImpR_inv' intro!: SCc.ImpR) qed primrec cut_bound :: "nat \ nat \ 'a formula \ nat" where "cut_bound n m (Atom _) = m + n" | "cut_bound n m Bot = n" | "cut_bound n m (Not F) = cut_bound m n F" | "cut_bound n m (And F G) = cut_bound n (cut_bound n m F) G" | "cut_bound n m (Or F G) = cut_bound (cut_bound n m F) m G" | "cut_bound n m (Imp F G) = cut_bound (cut_bound m n F) m G" theorem cut_bound: "\ \ F,\ \ n \ F,\ \ \ \ m \ \ \ \ \ cut_bound n m F" proof(induction F arbitrary: \ \ n m) case (Atom k) thus ?case using Cut_Atom_depth by simp fast next case Bot thus ?case using Bot_delR' by fastforce next case Not from Not.prems show ?case by(auto dest!: NotL_inv' NotR_inv' intro!: Not.IH elim!: weakenL) next case (And F G) from And.prems show ?case by(auto dest!: AndL_inv' AndR_inv' intro!: And.IH elim!: weakenR' weakenL') next case (Or F G) from Or.prems show ?case by(auto dest: OrL_inv' OrR_inv' intro!: Or.IH elim!: weakenR' weakenL') next case (Imp F G) from ImpR_inv' \\ \ F \<^bold>\ G, \ \ n\ have R: "F, \ \ G, \ \ n" by blast from ImpL_inv' \F \<^bold>\ G, \ \ \ \ m\ have L: "\ \ F, \ \ m" "G, \ \ \ \ m" by blast+ from L(1) have "\ \ F, G, \ \ m" using weakenR' by blast from Imp.IH(1)[OF this R] have "\ \ G, \ \ cut_bound m n F" . from Imp.IH(2)[OF this L(2)] have "\ \ \ \ cut_bound (cut_bound m n F) m G" . thus "\ \ \ \ cut_bound n m (F \<^bold>\ G)" by simp qed context begin private primrec cut_bound' :: "nat \ 'a formula \ nat" where "cut_bound' n (Atom _) = 2*n" | "cut_bound' n Bot = n" | "cut_bound' n (Not F) = cut_bound' n F" | "cut_bound' n (And F G) = cut_bound' (cut_bound' n F) G" | "cut_bound' n (Or F G) = cut_bound' (cut_bound' n F) G" | "cut_bound' n (Imp F G) = cut_bound' (cut_bound' n F) G" private lemma cut_bound'_mono: "a \ b \ cut_bound' a F \ cut_bound' b F" by(induction F arbitrary: a b; simp) private lemma cut_bound_mono: "a \ c \ b \ d \ cut_bound a b F \ cut_bound c d F" by(induction F arbitrary: a b c d; simp) private lemma cut_bound_max: "max n (cut_bound' (max n m) F) = cut_bound' (max n m) F" by(induction F arbitrary: n m; simp; metis) private lemma cut_bound_max': "max n (cut_bound' n F) = cut_bound' n F" by(induction F arbitrary: n ; simp; metis max.assoc) private lemma cut_bound_': "cut_bound n m F \ cut_bound' (max n m) F" proof(induction F arbitrary: n m) case (Not F) then show ?case by simp (metis max.commute) next case (And F1 F2) from And.IH(1) have 1: "cut_bound n (cut_bound n m F1) F2 \ cut_bound n (cut_bound' (max n m) F1) F2" by(rule cut_bound_mono[OF order.refl]) also from And.IH(2) have "\ \ cut_bound' (max n (cut_bound' (max n m) F1)) F2" by simp also have "\ = cut_bound' (cut_bound' (max n m) F1) F2" by (simp add: cut_bound_max) finally show ?case by simp next case (Or F1 F2) from Or.IH(1) have 1: "cut_bound (cut_bound n m F1) m F2 \ cut_bound (cut_bound' (max n m) F1) m F2" by(rule cut_bound_mono[OF _ order.refl]) also from Or.IH(2)[of "cut_bound' (max n m) F1"] have "\ \ cut_bound' (max (cut_bound' (max n m) F1) m) F2" by simp also have "\ = cut_bound' (cut_bound' (max n m) F1) F2" by (simp add: cut_bound_max max.commute) finally show ?case by simp next case (Imp F1 F2) from Imp.IH(1) have 1: "cut_bound (cut_bound m n F1) m F2 \ cut_bound (cut_bound' (max m n) F1) m F2" by(rule cut_bound_mono[OF _ order.refl]) also from Imp.IH(2)[of "cut_bound' (max m n) F1"] have "\ \ cut_bound' (max (cut_bound' (max m n) F1) m) F2" by simp also have "\ = cut_bound' (cut_bound' (max n m) F1) F2" by (simp add: cut_bound_max max.commute) finally show ?case by simp qed simp_all primrec depth :: "'a formula \ nat" where "depth (Atom _) = 0" | "depth Bot = 0" | "depth (Not F) = Suc (depth F)" | "depth (And F G) = Suc (max (depth F) (depth G))" | "depth (Or F G) = Suc (max (depth F) (depth G))" | "depth (Imp F G) = Suc (max (depth F) (depth G))" private primrec (* primrec works, but fun would give me a nicer induction rule *) cbnd where "cbnd k 0 = 2*k" | "cbnd k (Suc n) = cbnd (cbnd k n) n" private lemma cbnd_grow: "(k :: nat) \ cbnd k d" by(induction d arbitrary: k; simp) (insert le_trans, blast) private lemma cbnd_mono: assumes "b \ d" shows "cbnd (a::nat) b \ cbnd a d" proof - have "cbnd (a::nat) b \ cbnd a (b + d)" for b d by(induction d arbitrary: a b; simp) (insert le_trans cbnd_grow, blast) thus ?thesis using assms using le_Suc_ex by blast qed private lemma cut_bound'_cbnd: "cut_bound' n F \ cbnd n (depth F)" proof(induction F arbitrary: n) next case (Not F) then show ?case using cbnd_grow dual_order.trans by fastforce next case (And F1 F2) let ?md = "max (depth F1) (depth F2)" have "cut_bound' (cut_bound' n F1) F2 \ cut_bound' (cbnd n (depth F1)) F2" by (simp add: And.IH(1) cut_bound'_mono) also have "... \ cut_bound' (cbnd n ?md) F2" by (simp add: cbnd_mono cut_bound'_mono) also have "... \ cbnd (cbnd n ?md) (depth F2)" using And.IH(2) by blast also have "... \ cbnd (cbnd n ?md) ?md" by (simp add: cbnd_mono) finally show ?case by simp next case (Imp F1 F2) case (Or F1 F2) text\analogous\ (*<*) let ?md = "max (depth F1) (depth F2)" have "cut_bound' (cut_bound' n F1) F2 \ cut_bound' (cbnd n (depth F1)) F2" by (simp add: Or.IH(1) cut_bound'_mono) also have "... \ cut_bound' (cbnd n ?md) F2" by (simp add: cbnd_mono cut_bound'_mono) also have "... \ cbnd (cbnd n ?md) (depth F2)" using Or.IH(2) by blast also have "... \ cbnd (cbnd n ?md) ?md" by (simp add: cbnd_mono) finally show ?case by simp next case (Imp F1 F2) let ?md = "max (depth F1) (depth F2)" have "cut_bound' (cut_bound' n F1) F2 \ cut_bound' (cbnd n (depth F1)) F2" by (simp add: Imp.IH(1) cut_bound'_mono) also have "... \ cut_bound' (cbnd n ?md) F2" by (simp add: cbnd_mono cut_bound'_mono) also have "... \ cbnd (cbnd n ?md) (depth F2)" using Imp.IH(2) by blast also have "... \ cbnd (cbnd n ?md) ?md" by (simp add: cbnd_mono) finally show ?case by simp (*>*) qed simp_all value "map (cbnd (0::int)) [0,1,2,3,4]" value "map (cbnd (1::int)) [0,1,2,3,4]" (* int, because that's easier to read\ *) value "map (cbnd (2::int)) [0,1,2,3,4]" value "map (cbnd (3::int)) [0,1,2,3,4]" value [nbe] "map (int \ (\n. n div 3) \ cut_bound 3 3 \ (\n. ((\F. And F F) ^^ n) (Atom 0))) [0,1,2,3,4,5,6,7]" value [nbe] "map (int \ (\n. n div 3) \ cut_bound' 3 \ (\n. ((\F. And F F) ^^ n) (Atom 0))) [0,1,2,3,4]" (* TODO: hm *) value [nbe] "map (int \ (\n. n div 3) \ cut_bound 3 3 \ (\n. ((\F. Imp (Or F F) (And F F)) ^^ n) (Atom 0))) [0,1,2]" value [nbe] "map (int \ (\n. n div 3) \ cut_bound' 3 \ (\n. ((\F. Imp (Or F F) (And F F)) ^^ n) (Atom 0))) [0,1,2]" (* Hmhm *) value [nbe] "(\F. And (Or F F) (Or F F)) ^^ 2" lemma "n + ((n + m) * 2 ^ (size F - Suc 0) + (n + (n + m + (n + m) * 2 ^ (size F - Suc 0))) * 2 ^ (size G - Suc 0)) \ (n + (m :: nat)) * 2 ^ (size F + size G)" oops (* so the proof below won't work. *) lemma "cut_bound (n :: nat) m F \ (n + m) * (2 ^ (size F - 1) + 1)" proof(induction F arbitrary: n m) next case (Not F) show ?case unfolding cut_bound.simps by(rule le_trans[OF Not]) (simp add: add.commute) next have "1 \ size F" for F :: "'a formula" by(cases F; simp) case (And F G) from And(2) have "cut_bound n (cut_bound n m F) G \ (n + (cut_bound n m F)) * (2 ^ (size G - 1) + 1)" by simp also from And(1) have "\ \ (n + (n + m) * (2 ^ (size F - 1) + 1)) * (2 ^ (size G - 1) + 1)" by (meson add_le_cancel_left mult_le_mono1) also have "\ \ (n + m) * (2 ^ (size (F \<^bold>\ G) - 1) + 1)" apply simp oops (* Oh! Look: *) private lemma cbnd_comm: "cbnd (l * k::nat) n = l * cbnd (k::nat) n" by(induction n arbitrary: k; simp) private lemma cbnd_closed: "cbnd (k::nat) n = k * 2 ^ (2 ^ n)" by(induction n arbitrary: k;simp add: semiring_normalization_rules(26)) theorem cut': assumes "\ \ F,\ \ n" "F,\ \ \ \ n" shows "\ \ \ \ n * 2 ^ (2 ^ depth F)" proof - from cut_bound[OF assms] have c: "\ \ \ \ cut_bound n n F" . have d: "cut_bound n n F \ max n n * 2 ^ (2 ^ depth F)" using cut_bound_' cut_bound'_cbnd cbnd_closed by (metis order_trans) show ?thesis using c d le_Suc_ex deeper unfolding max.idem by metis qed end end diff --git a/thys/Propositional_Proof_Systems/SC_Gentzen.thy b/thys/Propositional_Proof_Systems/SC_Gentzen.thy --- a/thys/Propositional_Proof_Systems/SC_Gentzen.thy +++ b/thys/Propositional_Proof_Systems/SC_Gentzen.thy @@ -1,250 +1,257 @@ subsubsection\Mimicking the original\ theory SC_Gentzen imports SC_Depth SC_Cut begin text\This system attempts to mimic the original sequent calculus (``Reihen von Formeln, durch Kommata getrennt'', translates roughly to sequences of formulas, separated by a comma)~\cite{gentzen1935untersuchungen}.\ inductive SCg :: "'a formula list \ 'a formula list \ bool" (infix "\" 30) where Anfang: "[\

] \ [\
]" | FalschA: "[\] \ []" | VerduennungA: "\ \ \ \ \
#\ \ \" | VerduennungS: "\ \ \ \ \ \ \
#\" | ZusammenziehungA: "\
#\
#\ \ \ \ \
#\ \ \" | ZusammenziehungS: "\ \ \
#\
#\ \ \ \ \
#\" | VertauschungA: "\@\
#\#\ \ \ \ \@\#\
#\ \ \" | VertauschungS: "\ \ \@\#\
#\ \ \ \ \@\
#\#\" | Schnitt: "\\ \ \
#\; \
#\ \ \\ \ \@\ \ \@\" | UES: "\\ \ \#\; \ \ \#\\ \ \ \ \\<^bold>\\#\" | UEA1: "\#\ \ \ \ \\<^bold>\\#\ \ \" | UEA2: "\#\ \ \ \ \\<^bold>\\#\ \ \" | OEA: "\\#\ \ \; \#\ \ \\ \ \\<^bold>\\#\ \ \" | OES1: "\ \ \#\ \ \ \ \\<^bold>\\#\" | OES2: "\ \ \#\ \ \ \ \\<^bold>\\#\" | FES: "\#\ \ \#\ \ \ \ \\<^bold>\\#\" | FEA: "\\ \ \#\; \#\ \ \\ \ \\<^bold>\\#\@\ \ \@\" | NES: "\#\ \ \ \ \ \ \<^bold>\\#\" | NEA: "\ \ \#\ \ \<^bold>\\#\ \ \" text\Nota bene: E here stands for ``Einführung'', which is introduction and not elemination.\ text\The rule for @{term \} is not part of the original calculus. Its addition is necessary to show equivalence to our @{const SCp}.\ text\Note that we purposefully did not recreate the fact that Gentzen sometimes puts his principal formulas on end and sometimes on the beginning of the list.\ lemma AnfangTauschA: "\
#\@\ \ \ \ \@\
#\ \ \" by(induction \ arbitrary: \ rule: List.rev_induct) (simp_all add: VertauschungA) lemma AnfangTauschS: "\ \ \
#\@\ \ \ \ \@\
#\" by(induction \ arbitrary: \ rule: List.rev_induct) (simp_all add: VertauschungS) lemma MittenTauschA: "\@\
#\ \ \ \ \
#\@\ \ \" by(induction \ arbitrary: \ rule: List.rev_induct) (simp_all add: VertauschungA) lemma MittenTauschS: "\ \ \@\
#\ \ \ \ \
#\@\" by(induction \ arbitrary: \ rule: List.rev_induct) (simp_all add: VertauschungS) lemma BotLe: "\\set \ \ \\\" proof - have A: "\#\\[]" for \ by(induction \) (simp_all add: FalschA VerduennungA VertauschungA[where \=Nil, simplified]) have *: "\#\\\" for \ by(induction \) (simp_all add: A VerduennungS) assume "\\set \" then obtain \1 \2 where \: "\ = \1@\#\2" by (meson split_list) show ?thesis unfolding \ using AnfangTauschA * by blast qed lemma Axe: "A \ set \ \ A \ set \ \ \ \ \" proof - have A: "A#\ \ [A]" for \ by(induction \) (simp_all add: Anfang VertauschungA[where \=Nil, simplified] VerduennungA) have S: "A#\ \ A#\" for \ \ by(induction \) (simp_all add: A Anfang VertauschungS[where \=Nil, simplified] VerduennungS) assume "A \ set \" "A \ set \" thus ?thesis apply(-) apply(drule split_list)+ apply(clarify) apply(intro AnfangTauschA AnfangTauschS) apply(rule S) done qed lemma VerduennungListeA: "\ \ \ \ \@\ \ \" proof - have "\ \ \ \ \\''. \=\''@\' \ \'@\ \ \" for \' proof(induction \') case (Cons a as) - from \\\''. \ = \'' @ a # as\ guess \'' .. + then obtain \'' where "\ = \'' @ a # as" by blast hence "\\''. \ = \'' @ as" by(intro exI[where x="\'' @ [a]"]) simp from Cons.IH[OF Cons.prems(1) this] have "as @ \ \ \" . thus ?case using VerduennungA by simp qed simp thus "\ \ \ \ \@\ \ \" by simp qed lemma VerduennungListeS: "\ \ \ \ \ \ \@\" proof - have "\ \ \ \ \\''. \=\''@\' \ \ \ \'@\" for \' proof(induction \') case (Cons a as) - from \\\''. \ = \'' @ a # as\ guess \'' .. + then obtain \'' where "\ = \'' @ a # as" by blast hence "\\''. \ = \'' @ as" by(intro exI[where x="\'' @ [a]"]) simp from Cons.IH[OF Cons.prems(1) this] have " \ \ as @ \" . thus ?case using VerduennungS by simp qed simp thus "\ \ \ \ \ \ \@\" by simp qed (* wak, those weren't the droids I was looking for. *) lemma ZusammenziehungListeA: "\@\ \ \ \ \ \ \" proof - have "\'@\ \ \ \ \\''. \=\''@\' \ \ \ \" for \' proof(induction \') case (Cons a \') - from \\\''. \ = \'' @ a # \'\ guess \'' .. note \'' = this + then obtain \'' where \'': "\ = \'' @ a # \'" by blast then obtain \1 \2 where \: "\ = \1 @ a # \2" by blast from \'' have **: "\\''. \ = \'' @ \'" by(intro exI[where x="\'' @ [a]"]) simp from Cons.prems(1) have "a # (a # \') @ \1 @ \2 \ \" unfolding \ using MittenTauschA by (metis append_assoc) hence "(a # \') @ \1 @ \2 \ \" using ZusammenziehungA by auto hence "\' @ \ \ \" unfolding \ using AnfangTauschA by (metis append_Cons append_assoc) from Cons.IH[OF this **] show "\ \ \" . qed simp thus "\@\ \ \ \ \ \ \" by simp qed lemma ZusammenziehungListeS: "\ \ \@\ \ \ \ \" proof - have "\ \ \'@\ \ \\''. \=\''@\' \ \ \ \" for \' proof(induction \') case (Cons a \') - from \\\''. \ = \'' @ a # \'\ guess \'' .. note \'' = this + then obtain \'' where \'': "\ = \'' @ a # \'" by blast then obtain \1 \2 where \: "\ = \1 @ a # \2" by blast from \'' have **: "\\''. \ = \'' @ \'" by(intro exI[where x="\'' @ [a]"]) simp from Cons.prems(1) have "\ \ a # (a # \') @ \1 @ \2" unfolding \ using MittenTauschS by (metis append_assoc) hence "\ \ (a # \') @ \1 @ \2" using ZusammenziehungS by auto hence "\ \ \' @ \" unfolding \ using AnfangTauschS by (metis append_Cons append_assoc) from Cons.IH[OF this **] show "\ \ \" . qed simp thus "\ \ \@\ \ \ \\" by simp qed theorem gentzen_sc_eq: "mset \ \ mset \ \ \ \ \" proof assume "mset \ \ mset \" then obtain n where "mset \ \ mset \ \ n" unfolding SC_SCp_eq[symmetric] .. thus "\ \ \" (* structural induction not necessary at all. I still don't get it. *) proof(induction n arbitrary: \ \ rule: nat.induct) case (Suc n) have sr: "\\1 \2. \ = \1 @ F # \2 \ \' = mset (\1@\2)" (is "?s") if "mset \ = F, \'" for \ \' F proof - from that obtain \1 \2 where \: "\ = \1 @ F # \2" by (metis split_list add.commute ex_mset list.set_intros(1) mset.simps(2) set_mset_mset) hence \': "\' = mset (\1@\2)" using that by auto show ?s using \ \' by blast qed from Suc.prems show ?case proof(cases rule: SCc.cases) case BotL thus ?thesis using BotLe by simp next case Ax thus ?thesis using Axe by simp next case (NotL \' F) from \mset \ = \<^bold>\ F, \'\ obtain \1 \2 where \: "\ = \1 @ \<^bold>\ F # \2" by (metis split_list add.commute ex_mset list.set_intros(1) mset.simps(2) set_mset_mset) hence \': "\' = mset (\1@\2)" using NotL(1) by simp from \\' \ F, mset \ \ n\ have "mset (\1@\2) \ mset (F#\) \ n" unfolding \' by (simp add: add.commute) from Suc.IH[OF this] show ?thesis unfolding \ using AnfangTauschA NEA by blast next case (NotR F \') - from sr[OF NotR(1)] guess \1 .. then guess \2 .. note \ = this + from sr[OF NotR(1)] obtain \1 \2 where \: "\ = \1 @ \<^bold>\ F # \2 \ \' = mset (\1 @ \2)" + by blast with NotR have "mset (F#\) \ mset (\1@\2) \ n" by (simp add: add.commute) from Suc.IH[OF this] show ?thesis using \ using AnfangTauschS NES by blast next case (AndR F \' G) - from sr[OF AndR(1)] guess \1 .. then guess \2 .. note \ = this + from sr[OF AndR(1)] obtain \1 \2 where \: "\ = \1 @ F \<^bold>\ G # \2 \ \' = mset (\1 @ \2)" + by blast with AndR have "mset \ \ mset (F # \1@\2) \ n" "mset \ \ mset (G # \1@\2) \ n" by (simp add: add.commute)+ from this[THEN Suc.IH] show ?thesis using \ using AnfangTauschS UES by blast next case (OrR F G \') - from sr[OF OrR(1)] guess \1 .. then guess \2 .. note \ = this + from sr[OF OrR(1)] obtain \1 \2 where \: "\ = \1 @ F \<^bold>\ G # \2 \ \' = mset (\1 @ \2)" + by blast with OrR have "mset \ \ mset (G # F # \1@\2) \ n" by (simp add: add.commute add.left_commute add_mset_commute) from this[THEN Suc.IH] have "\ \ G # F # \1 @ \2" . with OES2 have "\ \ F \<^bold>\ G # F # \1 @ \2" . with VertauschungS[where \=Nil, simplified] have "\ \ F # F \<^bold>\ G # \1 @ \2" . with OES1 have "\ \ F \<^bold>\ G # F \<^bold>\ G # \1 @ \2" . hence "\ \ F \<^bold>\ G # \1 @ \2" using ZusammenziehungS by fast thus ?thesis unfolding \[THEN conjunct1] using AnfangTauschS by blast next case (ImpR F G \') - from sr[OF ImpR(1)] guess \1 .. then guess \2 .. note \ = this + from sr[OF ImpR(1)] obtain \1 \2 where \: "\ = \1 @ F \<^bold>\ G # \2 \ \' = mset (\1 @ \2)" + by blast with ImpR have "mset (F#\) \ mset (G # \1@\2) \ n" by (simp add: add.commute) from this[THEN Suc.IH] show ?thesis using \ using AnfangTauschS FES by blast next case (AndL F G \') - from sr[OF this(1)] guess \1 .. then guess \2 .. note \ = this + from sr[OF this(1)] obtain \1 \2 where \: "\ = \1 @ F \<^bold>\ G # \2 \ \' = mset (\1 @ \2)" + by blast with AndL have "mset (G # F # \1@\2) \ mset \ \ n" by (simp add: add.commute add.left_commute add_mset_commute) from this[THEN Suc.IH] have "G # F # \1 @ \2 \ \" . with UEA2 have "F \<^bold>\ G # F # \1 @ \2 \ \" . with VertauschungA[where \=Nil, simplified] have "F # F \<^bold>\ G # \1 @ \2 \ \" . with UEA1 have "F \<^bold>\ G # F \<^bold>\ G # \1 @ \2 \ \" . hence "F \<^bold>\ G # \1 @ \2 \ \" using ZusammenziehungA by fast thus ?thesis unfolding \[THEN conjunct1] using AnfangTauschA by blast next case (OrL F \' G) - from sr[OF this(1)] guess \1 .. then guess \2 .. note \ = this + from sr[OF this(1)] obtain \1 \2 where \: "\ = \1 @ F \<^bold>\ G # \2 \ \' = mset (\1 @ \2)" + by blast with OrL have "mset (F # \1@\2) \ mset \ \ n" "mset (G # \1@\2) \ mset \ \ n" by (simp add: add.commute)+ from this[THEN Suc.IH] show ?thesis using \ using AnfangTauschA OEA by blast next case (ImpL \' F G) - from sr[OF this(1)] guess \1 .. then guess \2 .. note \ = this + from sr[OF this(1)] obtain \1 \2 where \: "\ = \1 @ F \<^bold>\ G # \2 \ \' = mset (\1 @ \2)" + by blast with ImpL have "mset (\1@\2) \ mset (F#\) \ n" "mset (G # \1@\2) \ mset \ \ n" by (simp add: add.commute)+ from this[THEN Suc.IH] have "\1 @ \2 \ F # \" "G # \1 @ \2 \ \" . from FEA[OF this] have "F \<^bold>\ G # (\1 @ \2) @ (\1 @ \2) \ \ @ \" . hence "F \<^bold>\ G # (\1 @ \2) @ (F \<^bold>\ G # \1 @ \2) \ \ @ \" using AnfangTauschA VerduennungA by blast (* given the form of ZusammenziehungListeA, using this intermediate step is just easier. maybe a different from for ZusammenziehungListeA would be nice? *) hence "F \<^bold>\ G # (\1 @ \2) \ \ @ \" using ZusammenziehungListeA[where \="F \<^bold>\ G # (\1 @ \2)"] by simp thus ?thesis unfolding \[THEN conjunct1] by(intro AnfangTauschA; elim ZusammenziehungListeS) qed qed blast next have mset_Cons[simp]: "mset (A # S) = A, mset S" for A::"'a formula" and S by (simp add: add.commute) note mset.simps(2)[simp del] show "\ \ \ \ mset \ \ mset \" proof(induction rule: SCg.induct) (*sed -n '/^\([^\\]*\):/ {N;s/_\\\?//g; s/let.*$//; s/\n//; s/[\\:]//g; s/fix //; s/ / /g; s/^.*$/case (&) thus ?case sorry/; p;}' ${print_cases} *) case (Anfang \
) thus ?case using extended_Ax SC_SCp_eq by force next case (FalschA) thus ?case using SCp.BotL by force next case (VerduennungA \ \ \
) thus ?case by (simp add: SC.weakenL) next case (VerduennungS \ \ \
) thus ?case by (simp add: SC.weakenR) next case (ZusammenziehungA \
\ \) thus ?case using contractL by force next case (ZusammenziehungS \ \
\) thus ?case using contract by force next case (VertauschungA \ \
\ \ \) thus ?case by fastforce next case (VertauschungS \ \ \ \
\) thus ?case by fastforce next case (Schnitt \ \
\ \ \) hence "mset \ \ \
,mset \" "\
,mset \ \ mset \" using SC_SCp_eq by auto from cut_cf[OF this] show ?case unfolding SC_SCp_eq by simp next case (UES \ \ \ \) thus ?case using SCp.AndR by (simp add: SC_SCp_eq) next case (UEA1 \ \ \ \) from \mset (\ # \) \ mset \\ have "\,\,mset \ \ mset \" using SC.weakenL by auto thus ?case using SCp.AndL by force next case (UEA2 \ \ \ \) from \mset (\ # \) \ mset \\ have "\,\,mset \ \ mset \" using SC.weakenL by auto thus ?case using SCp.AndL by force next case (OEA \ \ \ \) thus ?case unfolding SC_SCp_eq by (simp add: SCp.OrL) next case (OES1 \ \ \ \) thus ?case using SC.weakenR[where 'a='a] by(auto intro!: SCp.intros(3-)) next case (OES2 \ \ \ \) thus ?case by (simp add: SC.weakenR SCp.OrR) next case (FES \ \ \ \) thus ?case using weakenR unfolding SC_SCp_eq by (simp add: SCp.ImpR) next case (FEA \ \ \ \ \ \) from \mset \ \ mset (\ # \)\[THEN weakenL_set, THEN weakenR_set, of "mset \" "mset \"] have S: "mset (\@\) \ \,mset (\@\)" unfolding mset_append mset_Cons by (simp add: add_ac) (* sledgehammer comes up with some funny proof using cut\ *) from FEA obtain m where "mset (\ # \) \ mset \" by blast hence "mset \ + mset (\ # \) \ mset \ + mset \" using weakenL_set weakenR_set by fast hence A: "\,mset (\@\) \ mset (\@\)" by (simp add: add.left_commute) show ?case using S A SC_SCp_eq SCp.ImpL unfolding mset_Cons by blast next case (NES \ \ \) thus ?case using SCp.NotR by(simp add: SC_SCp_eq) next case (NEA \ \ \) thus ?case using SCp.NotL by(simp add: SC_SCp_eq) qed qed end diff --git a/thys/Randomised_Social_Choice/Preference_Profiles.thy b/thys/Randomised_Social_Choice/Preference_Profiles.thy --- a/thys/Randomised_Social_Choice/Preference_Profiles.thy +++ b/thys/Randomised_Social_Choice/Preference_Profiles.thy @@ -1,837 +1,839 @@ (* Title: Preference_Profiles.thy Author: Manuel Eberl, TU München Definition of (weak) preference profiles and functions for building and manipulating them *) section \Preference Profiles\ theory Preference_Profiles imports Main Order_Predicates "HOL-Library.Multiset" "HOL-Library.Disjoint_Sets" begin text \The type of preference profiles\ type_synonym ('agent, 'alt) pref_profile = "'agent \ 'alt relation" locale preorder_family = fixes dom :: "'a set" and carrier :: "'b set" and R :: "'a \ 'b relation" assumes nonempty_dom: "dom \ {}" assumes in_dom [simp]: "i \ dom \ preorder_on carrier (R i)" assumes not_in_dom [simp]: "i \ dom \ \R i x y" begin lemma not_in_dom': "i \ dom \ R i = (\_ _. False)" by (simp add: fun_eq_iff) end locale pref_profile_wf = fixes agents :: "'agent set" and alts :: "'alt set" and R :: "('agent, 'alt) pref_profile" assumes nonempty_agents [simp]: "agents \ {}" and nonempty_alts [simp]: "alts \ {}" assumes prefs_wf [simp]: "i \ agents \ finite_total_preorder_on alts (R i)" assumes prefs_undefined [simp]: "i \ agents \ \R i x y" begin lemma finite_alts [simp]: "finite alts" proof - from nonempty_agents obtain i where "i \ agents" by blast then interpret finite_total_preorder_on alts "R i" by simp show ?thesis by (rule finite_carrier) qed lemma prefs_wf' [simp]: "i \ agents \ total_preorder_on alts (R i)" "i \ agents \ preorder_on alts (R i)" using prefs_wf[of i] by (simp_all add: finite_total_preorder_on_def total_preorder_on_def del: prefs_wf) lemma not_outside: assumes "x \[R i] y" shows "i \ agents" "x \ alts" "y \ alts" proof - from assms show "i \ agents" by (cases "i \ agents") auto then interpret preorder_on alts "R i" by simp from assms show "x \ alts" "y \ alts" by (simp_all add: not_outside) qed sublocale preorder_family agents alts R by (intro preorder_family.intro) simp_all lemmas prefs_undefined' = not_in_dom' lemma wf_update: assumes "i \ agents" "total_preorder_on alts Ri'" shows "pref_profile_wf agents alts (R(i := Ri'))" proof - interpret total_preorder_on alts Ri' by fact from finite_alts have "finite_total_preorder_on alts Ri'" by unfold_locales with assms show ?thesis by (auto intro!: pref_profile_wf.intro split: if_splits) qed lemma wf_permute_agents: assumes "\ permutes agents" shows "pref_profile_wf agents alts (R \ \)" unfolding o_def using permutes_in_image[OF assms(1)] by (intro pref_profile_wf.intro prefs_wf) simp_all lemma (in -) pref_profile_eqI: assumes "pref_profile_wf agents alts R1" "pref_profile_wf agents alts R2" assumes "\x. x \ agents \ R1 x = R2 x" shows "R1 = R2" proof interpret R1: pref_profile_wf agents alts R1 by fact interpret R2: pref_profile_wf agents alts R2 by fact fix x show "R1 x = R2 x" by (cases "x \ agents"; intro ext) (simp_all add: assms(3)) qed end text \ Permutes a preference profile w.r.t. alternatives in the way described in the paper. This is needed for the definition of neutrality. \ definition permute_profile where "permute_profile \ R = (\i x y. R i (inv \ x) (inv \ y))" lemma permute_profile_map_relation: "permute_profile \ R = (\i. map_relation (inv \) (R i))" by (simp add: permute_profile_def map_relation_def) lemma permute_profile_compose [simp]: "permute_profile \ (R \ \) = permute_profile \ R \ \" by (auto simp: fun_eq_iff permute_profile_def o_def) lemma permute_profile_id [simp]: "permute_profile id R = R" by (simp add: permute_profile_def) lemma permute_profile_o: assumes "bij f" "bij g" shows "permute_profile f (permute_profile g R) = permute_profile (f \ g) R" using assms by (simp add: permute_profile_def o_inv_distrib) lemma (in pref_profile_wf) wf_permute_alts: assumes "\ permutes alts" shows "pref_profile_wf agents alts (permute_profile \ R)" proof (rule pref_profile_wf.intro) fix i assume "i \ agents" with assms interpret R: finite_total_preorder_on alts "R i" by simp from assms have [simp]: "inv \ x \ alts \ x \ alts" for x by (simp add: permutes_in_image permutes_inv) show "finite_total_preorder_on alts (permute_profile \ R i)" proof fix x y assume "permute_profile \ R i x y" thus "x \ alts" "y \ alts" using R.not_outside[of "inv \ x" "inv \ y"] by (auto simp: permute_profile_def) next fix x y z assume "permute_profile \ R i x y" "permute_profile \ R i y z" thus "permute_profile \ R i x z" using R.trans[of "inv \ x" "inv \ y" "inv \ z"] by (simp_all add: permute_profile_def) qed (insert R.total R.refl R.finite_carrier, simp_all add: permute_profile_def) qed (insert assms, simp_all add: permute_profile_def pref_profile_wf_def) text \ This shows that the above definition is equivalent to that in the paper. \ lemma permute_profile_iff [simp]: fixes R :: "('agent, 'alt) pref_profile" assumes "\ permutes alts" "x \ alts" "y \ alts" defines "R' \ permute_profile \ R" shows "\ x \[R' i] \ y \ x \[R i] y" using assms by (simp add: permute_profile_def permutes_inverses) subsection \Pareto dominance\ definition Pareto :: "('agent \ 'alt relation) \ 'alt relation" where "x \[Pareto(R)] y \ (\j. x \[R j] x) \ (\i. x \[R i] x \ x \[R i] y)" text \ A Pareto loser is an alternative that is Pareto-dominated by some other alternative. \ definition pareto_losers :: "('agent, 'alt) pref_profile \ 'alt set" where "pareto_losers R = {x. \y. y \[Pareto(R)] x}" lemma pareto_losersI [intro?, simp]: "y \[Pareto(R)] x \ x \ pareto_losers R" by (auto simp: pareto_losers_def) context preorder_family begin lemma Pareto_iff: "x \[Pareto(R)] y \ (\i\dom. x \[R i] y)" proof assume A: "x \[Pareto(R)] y" then obtain j where j: "x \[R j] x" by (auto simp: Pareto_def) hence j': "j \ dom" by (cases "j \ dom") auto then interpret preorder_on carrier "R j" by simp from j have "x \ carrier" by (auto simp: carrier_eq) with A preorder_on.refl[OF in_dom] show "(\i\dom. x \[R i] y)" by (auto simp: Pareto_def) next assume A: "(\i\dom. x \[R i] y)" from nonempty_dom obtain j where j: "j \ dom" by blast then interpret preorder_on carrier "R j" by simp from j A have "x \[R j] y" by simp hence "x \[R j] x" using not_outside refl by blast with A show "x \[Pareto(R)] y" by (auto simp: Pareto_def) qed lemma Pareto_strict_iff: "x \[Pareto(R)] y \ (\i\dom. x \[R i] y) \ (\i\dom. x \[R i] y)" by (auto simp: strongly_preferred_def Pareto_iff nonempty_dom) lemma Pareto_strictI: assumes "\i. i \ dom \ x \[R i] y" "i \ dom" "x \[R i] y" shows "x \[Pareto(R)] y" using assms by (auto simp: Pareto_strict_iff) lemma Pareto_strictI': assumes "\i. i \ dom \ x \[R i] y" "i \ dom" "\x \[R i] y" shows "x \[Pareto(R)] y" proof - from assms interpret preorder_on carrier "R i" by simp from assms have "x \[R i] y" by (simp add: strongly_preferred_def) with assms show ?thesis by (auto simp: Pareto_strict_iff ) qed sublocale Pareto: preorder_on carrier "Pareto(R)" proof - have "preorder_on carrier (R i)" if "i \ dom" for i using that by simp_all note A = preorder_on.not_outside[OF this(1)] preorder_on.refl[OF this(1)] preorder_on.trans[OF this(1)] from nonempty_dom obtain i where i: "i \ dom" by blast show "preorder_on carrier (Pareto R)" proof fix x y assume "x \[Pareto(R)] y" with A(1,2)[OF i] i show "x \ carrier" "y \ carrier" by (auto simp: Pareto_iff) qed (auto simp: Pareto_iff intro: A) qed lemma pareto_loser_in_alts: assumes "x \ pareto_losers R" shows "x \ carrier" proof - from assms obtain y i where "i \ dom" "x \[R i] y" by (auto simp: pareto_losers_def Pareto_strict_iff) then interpret preorder_on carrier "R i" by simp from \x \[R i] y\ have "x \[R i] y" by (simp add: strongly_preferred_def) thus "x \ carrier" using not_outside by simp qed lemma pareto_losersE: assumes "x \ pareto_losers R" obtains y where "y \ carrier" "y \[Pareto(R)] x" proof - from assms obtain y where y: "y \[Pareto(R)] x" unfolding pareto_losers_def by blast with Pareto.not_outside[of x y] have "y \ carrier" by (simp add: strongly_preferred_def) with y show ?thesis using that by blast qed end subsection \Preferred alternatives\ context pref_profile_wf begin lemma preferred_alts_subset_alts: "preferred_alts (R i) x \ alts" (is ?A) and finite_preferred_alts [simp,intro!]: "finite (preferred_alts (R i) x)" (is ?B) proof - have "?A \ ?B" proof (cases "i \ agents") assume "i \ agents" then interpret total_preorder_on alts "R i" by simp have "preferred_alts (R i) x \ alts" using not_outside by (auto simp: preferred_alts_def) thus ?thesis by (auto dest: finite_subset) qed (auto simp: preferred_alts_def) thus ?A ?B by blast+ qed lemma preferred_alts_altdef: "i \ agents \ preferred_alts (R i) x = {y\alts. y \[R i] x}" by (simp add: preorder_on.preferred_alts_altdef) end subsection \Favourite alternatives\ definition favorites :: "('agent, 'alt) pref_profile \ 'agent \ 'alt set" where "favorites R i = Max_wrt (R i)" definition favorite :: "('agent, 'alt) pref_profile \ 'agent \ 'alt" where "favorite R i = the_elem (favorites R i)" definition has_unique_favorites :: "('agent, 'alt) pref_profile \ bool" where "has_unique_favorites R \ (\i. favorites R i = {} \ is_singleton (favorites R i))" context pref_profile_wf begin lemma favorites_altdef: "favorites R i = Max_wrt_among (R i) alts" proof (cases "i \ agents") assume "i \ agents" then interpret total_preorder_on alts "R i" by simp show ?thesis by (simp add: favorites_def Max_wrt_total_preorder Max_wrt_among_total_preorder) qed (simp_all add: favorites_def Max_wrt_def Max_wrt_among_def pref_profile_wf_def) lemma favorites_no_agent [simp]: "i \ agents \ favorites R i = {}" by (auto simp: favorites_def Max_wrt_def Max_wrt_among_def) lemma favorites_altdef': "favorites R i = {x\alts. \y\alts. x \[R i] y}" proof (cases "i \ agents") assume "i \ agents" then interpret finite_total_preorder_on alts "R i" by simp show ?thesis using Max_wrt_among_nonempty[of alts] Max_wrt_among_subset[of alts] by (auto simp: favorites_altdef Max_wrt_among_total_preorder) qed simp_all lemma favorites_subset_alts: "favorites R i \ alts" by (auto simp: favorites_altdef') lemma finite_favorites [simp, intro]: "finite (favorites R i)" using favorites_subset_alts finite_alts by (rule finite_subset) lemma favorites_nonempty: "i \ agents \ favorites R i \ {}" proof - assume "i \ agents" then interpret finite_total_preorder_on alts "R i" by simp show ?thesis unfolding favorites_def by (intro Max_wrt_nonempty) simp_all qed lemma favorites_permute: assumes i: "i \ agents" and perm: "\ permutes alts" shows "favorites (permute_profile \ R) i = \ ` favorites R i" proof - from i interpret finite_total_preorder_on alts "R i" by simp from perm show ?thesis unfolding favorites_def by (subst Max_wrt_map_relation_bij) (simp_all add: permute_profile_def map_relation_def permutes_bij) qed lemma has_unique_favorites_altdef: "has_unique_favorites R \ (\i\agents. is_singleton (favorites R i))" proof safe fix i assume "has_unique_favorites R" "i \ agents" thus "is_singleton (favorites R i)" using favorites_nonempty[of i] by (auto simp: has_unique_favorites_def) next assume "\i\agents. is_singleton (favorites R i)" hence "is_singleton (favorites R i) \ favorites R i = {}" for i by (cases "i \ agents") (simp add: favorites_nonempty, simp add: favorites_altdef') thus "has_unique_favorites R" by (auto simp: has_unique_favorites_def) qed end locale pref_profile_unique_favorites = pref_profile_wf agents alts R for agents :: "'agent set" and alts :: "'alt set" and R + assumes unique_favorites': "has_unique_favorites R" begin lemma unique_favorites: "i \ agents \ favorites R i = {favorite R i}" using unique_favorites' by (auto simp: favorite_def has_unique_favorites_altdef is_singleton_the_elem) lemma favorite_in_alts: "i \ agents \ favorite R i \ alts" using favorites_subset_alts[of i] by (simp add: unique_favorites) end subsection \Anonymous profiles\ type_synonym ('agent, 'alt) apref_profile = "'alt set list multiset" definition anonymous_profile :: "('agent, 'alt) pref_profile \ ('agent, 'alt) apref_profile" where anonymous_profile_auxdef: "anonymous_profile R = image_mset (weak_ranking \ R) (mset_set {i. R i \ (\_ _. False)})" lemma (in pref_profile_wf) agents_eq: "agents = {i. R i \ (\_ _. False)}" proof safe fix i assume i: "i \ agents" and Ri: "R i = (\_ _. False)" from i interpret preorder_on alts "R i" by simp from carrier_eq Ri nonempty_alts show False by simp next fix i assume "R i \ (\_ _. False)" thus "i \ agents" using prefs_undefined'[of i] by (cases "i \ agents") auto qed lemma (in pref_profile_wf) anonymous_profile_def: "anonymous_profile R = image_mset (weak_ranking \ R) (mset_set agents)" by (simp only: agents_eq anonymous_profile_auxdef) lemma (in pref_profile_wf) anonymous_profile_permute: assumes "\ permutes alts" "finite agents" shows "anonymous_profile (permute_profile \ R) = image_mset (map ((`) \)) (anonymous_profile R)" proof - from assms(1) interpret R': pref_profile_wf agents alts "permute_profile \ R" by (rule wf_permute_alts) have "anonymous_profile (permute_profile \ R) = {#weak_ranking (map_relation (inv \) (R x)). x \# mset_set agents#}" unfolding R'.anonymous_profile_def by (simp add: multiset.map_comp permute_profile_map_relation o_def) also from assms have "\ = {#map ((`) \) (weak_ranking (R x)). x \# mset_set agents#}" by (intro image_mset_cong) (simp add: finite_total_preorder_on.weak_ranking_permute[of alts]) also have "\ = image_mset (map ((`) \)) (anonymous_profile R)" by (simp add: anonymous_profile_def multiset.map_comp o_def) finally show ?thesis . qed lemma (in pref_profile_wf) anonymous_profile_update: assumes i: "i \ agents" and fin [simp]: "finite agents" and "total_preorder_on alts Ri'" shows "anonymous_profile (R(i := Ri')) = anonymous_profile R - {#weak_ranking (R i)#} + {#weak_ranking Ri'#}" proof - from assms interpret R': pref_profile_wf agents alts "R(i := Ri')" by (simp add: finite_total_preorder_on_iff wf_update) have "anonymous_profile (R(i := Ri')) = {#weak_ranking (if x = i then Ri' else R x). x \# mset_set agents#}" by (simp add: R'.anonymous_profile_def o_def) also have "\ = {#if x = i then weak_ranking Ri' else weak_ranking (R x). x \# mset_set agents#}" by (intro image_mset_cong) simp_all also have "\ = {#weak_ranking Ri'. x \# mset_set {x \ agents. x = i}#} + {#weak_ranking (R x). x \# mset_set {x \ agents. x \ i}#}" by (subst image_mset_If) ((subst filter_mset_mset_set, simp)+, rule refl) also from i have "{x \ agents. x = i} = {i}" by auto also have "{x \ agents. x \ i} = agents - {i}" by auto also have "{#weak_ranking Ri'. x \# mset_set {i}#} = {#weak_ranking Ri'#}" by simp also from i have "mset_set (agents - {i}) = mset_set agents - {#i#}" by (simp add: mset_set_Diff) also from i have "{#weak_ranking (R x). x \# \#} = {#weak_ranking (R x). x \# mset_set agents#} - {#weak_ranking (R i)#}" by (subst image_mset_Diff) (simp_all add: in_multiset_in_set mset_subset_eq_single) also have "{#weak_ranking Ri'#} + \ = anonymous_profile R - {#weak_ranking (R i)#} + {#weak_ranking Ri'#}" by (simp add: anonymous_profile_def add_ac o_def) finally show ?thesis . qed subsection \Preference profiles from lists\ definition prefs_from_table :: "('agent \ 'alt set list) list \ ('agent, 'alt) pref_profile" where "prefs_from_table xss = (\i. case_option (\_ _. False) of_weak_ranking (map_of xss i))" definition prefs_from_table_wf where "prefs_from_table_wf agents alts xss \ agents \ {} \ alts \ {} \ distinct (map fst xss) \ set (map fst xss) = agents \ (\xs\set (map snd xss). \(set xs) = alts \ is_finite_weak_ranking xs)" lemma prefs_from_table_wfI: assumes "agents \ {}" "alts \ {}" "distinct (map fst xss)" assumes "set (map fst xss) = agents" assumes "\xs. xs \ set (map snd xss) \ \(set xs) = alts" assumes "\xs. xs \ set (map snd xss) \ is_finite_weak_ranking xs" shows "prefs_from_table_wf agents alts xss" using assms unfolding prefs_from_table_wf_def by auto lemma prefs_from_table_wfD: assumes "prefs_from_table_wf agents alts xss" shows "agents \ {}" "alts \ {}" "distinct (map fst xss)" and "set (map fst xss) = agents" and "\xs. xs \ set (map snd xss) \ \(set xs) = alts" and "\xs. xs \ set (map snd xss) \ is_finite_weak_ranking xs" using assms unfolding prefs_from_table_wf_def by auto lemma pref_profile_from_tableI: "prefs_from_table_wf agents alts xss \ pref_profile_wf agents alts (prefs_from_table xss)" proof (intro pref_profile_wf.intro) assume wf: "prefs_from_table_wf agents alts xss" fix i assume i: "i \ agents" with wf have "i \ set (map fst xss)" by (simp add: prefs_from_table_wf_def) then obtain xs where xs: "xs \ set (map snd xss)" "prefs_from_table xss i = of_weak_ranking xs" by (cases "map_of xss i") (fastforce dest: map_of_SomeD simp: prefs_from_table_def map_of_eq_None_iff)+ with wf show "finite_total_preorder_on alts (prefs_from_table xss i)" by (auto simp: prefs_from_table_wf_def intro!: finite_total_preorder_of_weak_ranking) next assume wf: "prefs_from_table_wf agents alts xss" fix i x y assume i: "i \ agents" with wf have "i \ set (map fst xss)" by (simp add: prefs_from_table_wf_def) hence "map_of xss i = None" by (simp add: map_of_eq_None_iff) thus "\prefs_from_table xss i x y" by (simp add: prefs_from_table_def) qed (simp_all add: prefs_from_table_wf_def) lemma prefs_from_table_eqI: assumes "distinct (map fst xs)" "distinct (map fst ys)" "set xs = set ys" shows "prefs_from_table xs = prefs_from_table ys" proof - from assms have "map_of xs = map_of ys" by (subst map_of_inject_set) simp_all thus ?thesis by (simp add: prefs_from_table_def) qed lemma prefs_from_table_undef: assumes "prefs_from_table_wf agents alts xss" "i \ agents" shows "prefs_from_table xss i = (\_ _. False)" proof - from assms have "i \ fst ` set xss" by (simp add: prefs_from_table_wf_def) hence "map_of xss i = None" by (simp add: map_of_eq_None_iff) thus ?thesis by (simp add: prefs_from_table_def) qed lemma prefs_from_table_map_of: assumes "prefs_from_table_wf agents alts xss" "i \ agents" shows "prefs_from_table xss i = of_weak_ranking (the (map_of xss i))" using assms by (auto simp: prefs_from_table_def map_of_eq_None_iff prefs_from_table_wf_def split: option.splits) lemma prefs_from_table_update: fixes x xs assumes "i \ set (map fst xs)" defines "xs' \ map (\(j,y). if j = i then (j, x) else (j, y)) xs" shows "(prefs_from_table xs)(i := of_weak_ranking x) = prefs_from_table xs'" (is "?lhs = ?rhs") proof have xs': "set (map fst xs') = set (map fst xs)" by (force simp: xs'_def) fix k consider "k = i" | "k \ set (map fst xs)" | "k \ i" "k \ set (map fst xs)" by blast thus "?lhs k = ?rhs k" proof cases assume k: "k = i" moreover from k have "y = x" if "(i, y) \ set xs'" for y using that by (auto simp: xs'_def split: if_splits) ultimately show ?thesis using assms(1) k xs' by (auto simp add: prefs_from_table_def map_of_eq_None_iff dest!: map_of_SomeD split: option.splits) next assume k: "k \ set (map fst xs)" with assms(1) have k': "k \ i" by auto with k xs' have "map_of xs k = None" "map_of xs' k = None" by (simp_all add: map_of_eq_None_iff) thus ?thesis by (simp add: prefs_from_table_def k') next assume k: "k \ i" "k \ set (map fst xs)" with k(1) have "map_of xs k = map_of xs' k" unfolding xs'_def by (induction xs) fastforce+ with k show ?thesis by (simp add: prefs_from_table_def) qed qed lemma prefs_from_table_swap: "x \ y \ prefs_from_table ((x,x')#(y,y')#xs) = prefs_from_table ((y,y')#(x,x')#xs)" by (intro ext) (auto simp: prefs_from_table_def) lemma permute_prefs_from_table: assumes "\ permutes fst ` set xs" shows "prefs_from_table xs \ \ = prefs_from_table (map (\(x,y). (inv \ x, y)) xs)" proof fix i have "(prefs_from_table xs \ \) i = (case map_of xs (\ i) of None \ \_ _. False | Some x \ of_weak_ranking x)" by (simp add: prefs_from_table_def o_def) also have "map_of xs (\ i) = map_of (map (\(x,y). (inv \ x, y)) xs) i" using map_of_permute[OF assms] by (simp add: o_def fun_eq_iff) finally show "(prefs_from_table xs \ \) i = prefs_from_table (map (\(x,y). (inv \ x, y)) xs) i" by (simp only: prefs_from_table_def) qed lemma permute_profile_from_table: assumes wf: "prefs_from_table_wf agents alts xss" assumes perm: "\ permutes alts" shows "permute_profile \ (prefs_from_table xss) = prefs_from_table (map (\(x,y). (x, map ((`) \) y)) xss)" (is "?f = ?g") proof fix i have wf': "prefs_from_table_wf agents alts (map (\(x, y). (x, map ((`) \) y)) xss)" proof (intro prefs_from_table_wfI, goal_cases) case (5 xs) then obtain y where "y \ set xss" "xs = map ((`) \) (snd y)" by (auto simp add: o_def case_prod_unfold) with assms show ?case by (simp add: image_Union [symmetric] prefs_from_table_wf_def permutes_image o_def case_prod_unfold) next case (6 xs) then obtain y where "y \ set xss" "xs = map ((`) \) (snd y)" by (auto simp add: o_def case_prod_unfold) with assms show ?case by (auto simp: is_finite_weak_ranking_def is_weak_ranking_iff prefs_from_table_wf_def distinct_map permutes_inj_on inj_on_image intro!: disjoint_image) qed (insert assms, simp_all add: image_Union [symmetric] prefs_from_table_wf_def permutes_image o_def case_prod_unfold) show "?f i = ?g i" proof (cases "i \ agents") assume "i \ agents" with assms wf' show ?thesis by (simp add: permute_profile_def prefs_from_table_undef) next assume i: "i \ agents" define xs where "xs = the (map_of xss i)" from i wf have xs: "map_of xss i = Some xs" by (cases "map_of xss i") (auto simp: prefs_from_table_wf_def xs_def) have xs_in_xss: "xs \ snd ` set xss" using xs by (force dest!: map_of_SomeD) with wf have set_xs: "\(set xs) = alts" by (simp add: prefs_from_table_wfD) from i have "prefs_from_table (map (\(x,y). (x, map ((`) \) y)) xss) i = of_weak_ranking (the (map_of (map (\(x,y). (x, map ((`) \) y)) xss) i))" using wf' by (intro prefs_from_table_map_of) simp_all also have "\ = of_weak_ranking (map ((`) \) xs)" by (subst map_of_map) (simp add: xs) also have "\ = (\a b. of_weak_ranking xs (inv \ a) (inv \ b))" by (intro ext) (simp add: of_weak_ranking_permute map_relation_def set_xs perm) also have "\ = permute_profile \ (prefs_from_table xss) i" by (simp add: prefs_from_table_def xs permute_profile_def) finally show ?thesis .. qed qed subsection \Automatic evaluation of preference profiles\ lemma eval_prefs_from_table [simp]: "prefs_from_table []i = (\_ _. False)" "prefs_from_table ((i, y) # xs) i = of_weak_ranking y" "i \ j \ prefs_from_table ((j, y) # xs) i = prefs_from_table xs i" by (simp_all add: prefs_from_table_def) lemma eval_of_weak_ranking [simp]: "a \ \(set xs) \ \of_weak_ranking xs a b" "b \ x \ a \ \(set (x#xs)) \ of_weak_ranking (x # xs) a b" "b \ x \ of_weak_ranking (x # xs) a b \ of_weak_ranking xs a b" by (induction xs) (simp_all add: of_weak_ranking_Cons) lemma prefs_from_table_cong [cong]: assumes "prefs_from_table xs = prefs_from_table ys" shows "prefs_from_table (x#xs) = prefs_from_table (x#ys)" proof fix i show "prefs_from_table (x # xs) i = prefs_from_table (x # ys) i" using assms by (cases x, cases "i = fst x") simp_all qed definition of_weak_ranking_Collect_ge where "of_weak_ranking_Collect_ge xs x = {y. of_weak_ranking xs y x}" lemma eval_Collect_of_weak_ranking: "Collect (of_weak_ranking xs x) = of_weak_ranking_Collect_ge (rev xs) x" by (simp add: of_weak_ranking_Collect_ge_def) lemma of_weak_ranking_Collect_ge_empty [simp]: "of_weak_ranking_Collect_ge [] x = {}" by (simp add: of_weak_ranking_Collect_ge_def) lemma of_weak_ranking_Collect_ge_Cons [simp]: "y \ x \ of_weak_ranking_Collect_ge (x#xs) y = \(set (x#xs))" "y \ x \ of_weak_ranking_Collect_ge (x#xs) y = of_weak_ranking_Collect_ge xs y" by (auto simp: of_weak_ranking_Cons of_weak_ranking_Collect_ge_def) lemma of_weak_ranking_Collect_ge_Cons': "of_weak_ranking_Collect_ge (x#xs) = (\y. (if y \ x then \(set (x#xs)) else of_weak_ranking_Collect_ge xs y))" by (auto simp: of_weak_ranking_Cons of_weak_ranking_Collect_ge_def fun_eq_iff) lemma anonymise_prefs_from_table: assumes "prefs_from_table_wf agents alts xs" shows "anonymous_profile (prefs_from_table xs) = mset (map snd xs)" proof - from assms interpret pref_profile_wf agents alts "prefs_from_table xs" by (simp add: pref_profile_from_tableI) from assms have agents: "agents = fst ` set xs" by (simp add: prefs_from_table_wf_def) hence [simp]: "finite agents" by auto have "anonymous_profile (prefs_from_table xs) = {#weak_ranking (prefs_from_table xs x). x \# mset_set agents#}" by (simp add: o_def anonymous_profile_def) also from assms have "\ = {#the (map_of xs i). i \# mset_set agents#}" proof (intro image_mset_cong) fix i assume i: "i \# mset_set agents" from i assms have "weak_ranking (prefs_from_table xs i) = weak_ranking (of_weak_ranking (the (map_of xs i))) " by (simp add: prefs_from_table_map_of) also from assms i have "\ = the (map_of xs i)" by (intro weak_ranking_of_weak_ranking) (auto simp: prefs_from_table_wf_def) finally show "weak_ranking (prefs_from_table xs i) = the (map_of xs i)" . qed also from agents have "mset_set agents = mset_set (set (map fst xs))" by simp also from assms have "\ = mset (map fst xs)" by (intro mset_set_set) (simp_all add: prefs_from_table_wf_def) also from assms have "{#the (map_of xs i). i \# mset (map fst xs)#} = mset (map snd xs)" by (intro image_mset_map_of) (simp_all add: prefs_from_table_wf_def) finally show ?thesis . qed lemma prefs_from_table_agent_permutation: assumes wf: "prefs_from_table_wf agents alts xs" "prefs_from_table_wf agents alts ys" assumes mset_eq: "mset (map snd xs) = mset (map snd ys)" obtains \ where "\ permutes agents" "prefs_from_table xs \ \ = prefs_from_table ys" proof - from wf(1) have agents: "agents = set (map fst xs)" by (simp_all add: prefs_from_table_wf_def) from wf(2) have agents': "agents = set (map fst ys)" by (simp_all add: prefs_from_table_wf_def) from agents agents' wf(1) wf(2) have "mset (map fst xs) = mset (map fst ys)" by (subst set_eq_iff_mset_eq_distinct [symmetric]) (simp_all add: prefs_from_table_wfD) hence same_length: "length xs = length ys" by (auto dest: mset_eq_length simp del: mset_map) from \mset (map fst xs) = mset (map fst ys)\ obtain g where g: "g permutes {.. x < length xs" "f x < length ys \ x < length ys" for x by (simp_all add: same_length) define idx unidx where "idx = index (map fst xs)" and "unidx i = map fst xs ! i" for i from wf(1) have "bij_betw idx agents {0.. agents" for x using that by (simp add: idx_def agents) have [simp]: "unidx i \ agents" if "i < length xs" for i using that by (simp add: agents unidx_def) have unidx_idx: "unidx (idx x) = x" if x: "x \ agents" for x using x unfolding idx_def unidx_def using nth_index[of x "map fst xs"] by (simp add: agents set_map [symmetric] nth_map [symmetric] del: set_map) have idx_unidx: "idx (unidx i) = i" if i: "i < length xs" for i unfolding idx_def unidx_def using wf(1) index_nth_id[of "map fst xs" i] i by (simp add: prefs_from_table_wfD(3)) define \ where "\ x = (if x \ agents then (unidx \ f \ idx) x else x)" for x define \' where "\' x = (if x \ agents then (unidx \ inv f \ idx) x else x)" for x have "bij_betw (unidx \ f \ idx) agents agents" (is "?P") unfolding unidx_def by (rule bij_betw_trans bij_betw_idx permutes_imp_bij f g bij_betw_nth)+ (insert wf(1) g, simp_all add: prefs_from_table_wf_def same_length) also have "?P \ bij_betw \ agents agents" by (intro bij_betw_cong) (simp add: \_def) finally have perm: "\ permutes agents" by (intro bij_imp_permutes) (simp_all add: \_def) define h where "h = g \ f" from f g have h: "h permutes {..: "inv \ = \'" proof (rule permutes_invI[OF perm]) fix x assume "x \ agents" with f(1) show "\' (\ x) = x" by (simp add: \_def \'_def idx_unidx unidx_idx inv_f_f permutes_inj) qed (simp add: \_def \'_def) with perm have inv_\': "inv \' = \" by (auto simp: inv_inv_eq permutes_bij) from wf h have "prefs_from_table ys = prefs_from_table (permute_list h ys)" by (intro prefs_from_table_eqI) (simp_all add: prefs_from_table_wfD permute_list_map [symmetric]) also have "permute_list h ys = permute_list h (zip (map fst ys) (map snd ys))" by (simp add: zip_map_fst_snd) also from same_length f g have "permute_list h (zip (map fst ys) (map snd ys)) = zip (permute_list f (map fst xs)) (map snd xs)" by (subst permute_list_zip[OF h]) (simp_all add: h_def permute_list_compose) also { fix i assume i: "i < length xs" from i have "permute_list f (map fst xs) ! i = unidx (f i)" using permutes_in_image[OF f(1)] f(1) by (subst permute_list_nth) (simp_all add: same_length unidx_def) also from i have "\ = \ (unidx i)" by (simp add: \_def idx_unidx) also from i have "\ = map \ (map fst xs) ! i" by (simp add: unidx_def) finally have "permute_list f (map fst xs) ! i = map \ (map fst xs) ! i" . } hence "permute_list f (map fst xs) = map \ (map fst xs)" by (intro nth_equalityI) simp_all also have "zip (map \ (map fst xs)) (map snd xs) = map (\(x,y). (inv \' x, y)) xs" by (induction xs) (simp_all add: case_prod_unfold inv_\') also from permutes_inv[OF perm] inv_\ have "prefs_from_table \ = prefs_from_table xs \ \'" by (intro permute_prefs_from_table [symmetric]) (simp_all add: agents) finally have "prefs_from_table xs \ \' = prefs_from_table ys" .. with that[of \'] permutes_inv[OF perm] inv_\ show ?thesis by auto qed lemma permute_list_distinct: assumes "f ` {.. {..x. xs ! f (index xs x)) xs" using assms by (intro nth_equalityI) (auto simp: index_nth_id permute_list_def) lemma image_mset_eq_permutation: assumes "{#f x. x \# mset_set A#} = {#g x. x \# mset_set A#}" "finite A" obtains \ where "\ permutes A" "\x. x \ A \ g (\ x) = f x" proof - from assms(2) obtain xs where xs: "A = set xs" "distinct xs" using finite_distinct_list by blast with assms have "mset (map f xs) = mset (map g xs)" by (simp add: mset_set_set) from mset_eq_permutation[OF this] obtain \ where \: "\ permutes {0.. (map g xs) = map f xs" by (auto simp: atLeast0LessThan) define \' where "\' x = (if x \ A then ((!) xs \ \ \ index xs) x else x)" for x have "bij_betw ((!) xs \ \ \ index xs) A A" (is "?P") by (rule bij_betw_trans bij_betw_index xs refl permutes_imp_bij \ bij_betw_nth)+ (simp_all add: atLeast0LessThan xs) also have "?P \ bij_betw \' A A" by (intro bij_betw_cong) (simp_all add: \'_def) finally have "\' permutes A" by (rule bij_imp_permutes) (simp_all add: \'_def) moreover from \ xs(1)[symmetric] xs(2) have "g (\' x) = f x" if "x \ A" for x by (simp add: permute_list_map permute_list_distinct permutes_image \'_def that atLeast0LessThan) ultimately show ?thesis by (rule that) qed lemma anonymous_profile_agent_permutation: assumes eq: "anonymous_profile R1 = anonymous_profile R2" assumes wf: "pref_profile_wf agents alts R1" "pref_profile_wf agents alts R2" assumes fin: "finite agents" obtains \ where "\ permutes agents" "R2 \ \ = R1" proof - interpret R1: pref_profile_wf agents alts R1 by fact interpret R2: pref_profile_wf agents alts R2 by fact from eq have "{#weak_ranking (R1 x). x \# mset_set agents#} = {#weak_ranking (R2 x). x \# mset_set agents#}" by (simp add: R1.anonymous_profile_def R2.anonymous_profile_def o_def) - from image_mset_eq_permutation[OF this fin] guess \ . note \ = this + from image_mset_eq_permutation[OF this fin] obtain \ + where \: "\ permutes agents" + "\x. x \ agents \ weak_ranking (R2 (\ x)) = weak_ranking (R1 x)" by auto from \ have wf': "pref_profile_wf agents alts (R2 \ \)" by (intro R2.wf_permute_agents) then interpret R2': pref_profile_wf agents alts "R2 \ \" . have "R2 \ \ = R1" proof (intro pref_profile_eqI[OF wf' wf(1)]) fix x assume x: "x \ agents" with \ have "weak_ranking ((R2 o \) x) = weak_ranking (R1 x)" by simp with wf' wf(1) x show "(R2 \ \) x = R1 x" by (intro weak_ranking_eqD[of alts] R2'.prefs_wf) simp_all qed from \(1) and this show ?thesis by (rule that) qed end diff --git a/thys/Randomised_Social_Choice/SDS_Lowering.thy b/thys/Randomised_Social_Choice/SDS_Lowering.thy --- a/thys/Randomised_Social_Choice/SDS_Lowering.thy +++ b/thys/Randomised_Social_Choice/SDS_Lowering.thy @@ -1,409 +1,410 @@ (* Title: SDS_Lowering.thy Author: Manuel Eberl, TU München Allows to lower an SDS, i.e. take an existing ex-post efficient SDS and construct from it an SDS for fewer agents and alternatives. (which is also ex-post efficient) The standard properties (anonymity, neutrality, SD efficiency, strategyproofness) are preserved by this construction. *) section \Lowering Social Decision Schemes\ theory SDS_Lowering imports Social_Decision_Schemes begin definition lift_pref_profile :: "'agent set \ 'alt set \ 'agent set \ 'alt set \ ('agent, 'alt) pref_profile \ ('agent, 'alt) pref_profile" where "lift_pref_profile agents alts agents' alts' R = (\i x y. x \ alts' \ y \ alts' \ i \ agents' \ (x = y \ x \ alts \ i \ agents \ (y \ alts \ R i x y)))" lemma lift_pref_profile_wf: assumes "pref_profile_wf agents alts R" assumes "agents \ agents'" "alts \ alts'" "finite alts'" defines "R' \ lift_pref_profile agents alts agents' alts' R" shows "pref_profile_wf agents' alts' R'" proof - from assms interpret R: pref_profile_wf agents alts by simp have "finite_total_preorder_on alts' (R' i)" if i: "i \ agents'" for i proof (cases "i \ agents") case True then interpret finite_total_preorder_on alts "R i" by simp from True assms show ?thesis by unfold_locales (auto simp: lift_pref_profile_def dest: total intro: trans) next case False with assms i show ?thesis by unfold_locales (simp_all add: lift_pref_profile_def) qed moreover have "R' i = (\_ _. False)" if "i \ agents'" for i unfolding lift_pref_profile_def R'_def using that by simp ultimately show ?thesis unfolding pref_profile_wf_def using assms by auto qed lemma lift_pref_profile_permute_agents: assumes "\ permutes agents" "agents \ agents'" shows "lift_pref_profile agents alts agents' alts' (R \ \) = lift_pref_profile agents alts agents' alts' R \ \" using assms permutes_subset[OF assms] by (auto simp add: lift_pref_profile_def o_def permutes_in_image) lemma lift_pref_profile_permute_alts: assumes "\ permutes alts" "alts \ alts'" shows "lift_pref_profile agents alts agents' alts' (permute_profile \ R) = permute_profile \ (lift_pref_profile agents alts agents' alts' R)" proof - from assms have inv: "inv \ permutes alts" by (intro permutes_inv) from this assms(2) have "inv \ permutes alts'" by (rule permutes_subset) with inv show ?thesis using assms permutes_inj[OF \inv \ permutes alts\] by (fastforce simp add: lift_pref_profile_def permutes_in_image permute_profile_def fun_eq_iff dest: injD) qed lemma lotteries_on_subset: "A \ B \ p \ lotteries_on A \ p \ lotteries_on B" unfolding lotteries_on_def by blast lemma lottery_prob_carrier: "p \ lotteries_on A \ measure_pmf.prob p A = 1" by (auto simp: measure_pmf.prob_eq_1 lotteries_on_def AE_measure_pmf_iff) context fixes agents alts R agents' alts' R' assumes R_wf: "pref_profile_wf agents alts R" assumes election: "agents \ agents'" "alts \ alts'" "alts \ {}" "agents \ {}" "finite alts'" defines "R' \ lift_pref_profile agents alts agents' alts' R" begin interpretation R: pref_profile_wf agents alts R by fact interpretation R': pref_profile_wf agents' alts' R' using election R_wf by (simp add: R'_def lift_pref_profile_wf) lemma lift_pref_profile_strict_iff: "x \[lift_pref_profile agents alts agents' alts' R i] y \ i \ agents \ ((y \ alts \ x \ alts' - alts) \ x \[R i] y)" proof (cases "i \ agents") case True then interpret total_preorder_on alts "R i" by simp show ?thesis using not_outside election by (auto simp: lift_pref_profile_def strongly_preferred_def) qed (simp_all add: lift_pref_profile_def strongly_preferred_def) lemma preferred_alts_lift_pref_profile: assumes i: "i \ agents'" and x: "x \ alts'" shows "preferred_alts (R' i) x = (if i \ agents \ x \ alts then preferred_alts (R i) x else alts')" proof (cases "i \ agents") assume i: "i \ agents" then interpret Ri: total_preorder_on alts "R i" by simp show ?thesis using i x election Ri.not_outside by (auto simp: preferred_alts_def R'_def lift_pref_profile_def Ri.refl) qed (auto simp: preferred_alts_def R'_def lift_pref_profile_def i x) lemma lift_pref_profile_Pareto_iff: "x \[Pareto(R')] y \ x \ alts' \ y \ alts' \ (x \ alts \ x \[Pareto(R)] y)" proof - from R.nonempty_agents obtain i where i: "i \ agents" by blast then interpret Ri: finite_total_preorder_on alts "R i" by simp show ?thesis unfolding R'.Pareto_iff R.Pareto_iff unfolding R'_def lift_pref_profile_def using election i by (auto simp: preorder_on.refl[OF R.in_dom] simp del: R.nonempty_alts R.nonempty_agents intro: Ri.not_outside) qed lemma lift_pref_profile_Pareto_strict_iff: "x \[Pareto(R')] y \ x \ alts' \ y \ alts' \ (x \ alts \ y \ alts \ x \[Pareto(R)] y)" by (auto simp: strongly_preferred_def lift_pref_profile_Pareto_iff R.Pareto.not_outside) lemma pareto_losers_lift_pref_profile: shows "pareto_losers R' = pareto_losers R \ (alts' - alts)" proof - have A: "x \ alts" "y \ alts" if "x \[Pareto(R)] y" for x y using that R.Pareto.not_outside unfolding strongly_preferred_def by auto have B: "x \ alts'" if "x \ alts" for x using election that by blast from R.nonempty_alts obtain x where x: "x \ alts" by blast thus ?thesis unfolding pareto_losers_def lift_pref_profile_Pareto_strict_iff [abs_def] by (auto dest: A B) qed context begin private lemma lift_SD_iff_agent: assumes "p \ lotteries_on alts" "q \ lotteries_on alts" and i: "i \ agents" shows "p \[SD(R' i)] q \ p \[SD(R i)] q" proof - from i interpret Ri: preorder_on alts "R i" by simp from i election have i': "i \ agents'" by blast then interpret R'i: preorder_on alts' "R' i" by simp from assms election have "p \ lotteries_on alts'" "q \ lotteries_on alts'" by (auto intro: lotteries_on_subset) with assms election i' show ?thesis by (auto simp: Ri.SD_preorder R'i.SD_preorder preferred_alts_lift_pref_profile lottery_prob_carrier) qed private lemma lift_SD_iff_nonagent: assumes "p \ lotteries_on alts" "q \ lotteries_on alts" and i: "i \ agents' - agents" shows "p \[SD(R' i)] q" proof - from i election have i': "i \ agents'" by blast then interpret R'i: preorder_on alts' "R' i" by simp from assms election have "p \ lotteries_on alts'" "q \ lotteries_on alts'" by (auto intro: lotteries_on_subset) with assms election i' show ?thesis by (auto simp: R'i.SD_preorder preferred_alts_lift_pref_profile lottery_prob_carrier) qed lemmas lift_SD_iff = lift_SD_iff_agent lift_SD_iff_nonagent lemma lift_SD_iff': "p \ lotteries_on alts \ q \ lotteries_on alts \ i \ agents' \ p \[SD(R' i)] q \ i \ agents \ p \[SD(R i)] q" by (cases "i \ agents") (simp_all add: lift_SD_iff) end lemma lift_SD_strict_iff: assumes "p \ lotteries_on alts" "q \ lotteries_on alts" and i: "i \ agents" shows "p \[SD(R' i)] q \ p \[SD(R i)] q" using assms by (simp add: strongly_preferred_def lift_SD_iff) lemma lift_Pareto_SD_iff: assumes "p \ lotteries_on alts" "q \ lotteries_on alts" shows "p \[Pareto(SD \ R')] q \ p \[Pareto(SD \ R)] q" using assms election by (auto simp: R.SD.Pareto_iff R'.SD.Pareto_iff lift_SD_iff') lemma lift_Pareto_SD_strict_iff: assumes "p \ lotteries_on alts" "q \ lotteries_on alts" shows "p \[Pareto(SD \ R')] q \ p \[Pareto(SD \ R)] q" using assms by (simp add: strongly_preferred_def lift_Pareto_SD_iff) lemma lift_SD_efficient_iff: assumes p: "p \ lotteries_on alts" shows "SD_efficient R' p \ SD_efficient R p" proof assume eff: "SD_efficient R' p" have "\(q \[Pareto(SD \ R)] p)" if q: "q \ lotteries_on alts" for q proof - from q election have q': "q \ lotteries_on alts'" by (blast intro: lotteries_on_subset) with eff have "\(q \[Pareto(SD \ R')] p)" by (simp add: R'.SD_efficient_def) with p q show ?thesis by (simp add: lift_Pareto_SD_strict_iff) qed thus "SD_efficient R p" by (simp add: R.SD_efficient_def) next assume eff: "SD_efficient R p" have "\(q \[Pareto(SD \ R')] p)" if q: "q \ lotteries_on alts'" for q proof assume less: "q \[Pareto(SD \ R')] p" - from R'.SD_efficient_lottery_exists[OF q] guess q' . note q' = this + from R'.SD_efficient_lottery_exists[OF q] + obtain q' where q': "q' \ lotteries_on alts'" "Pareto (SD \ R') q q'" "SD_efficient R' q'" . have "x \ set_pmf q'" if x: "x \ alts' - alts" for x proof - from x have "x \ pareto_losers R'" by (simp add: pareto_losers_lift_pref_profile) with R'.SD_efficient_no_pareto_loser[OF q'(3,1)] show "x \ set_pmf q'" by blast qed with q' have "q' \ lotteries_on alts" by (auto simp: lotteries_on_def) moreover from q' less have "q' \[Pareto(SD \ R')] p" by (auto intro: R'.SD.Pareto.strict_weak_trans) with \q' \ lotteries_on alts\ p have "q' \[Pareto(SD \ R)] p" by (subst (asm) lift_Pareto_SD_strict_iff) ultimately have "\SD_efficient R p" by (auto simp: R.SD_efficient_def) with eff show False by contradiction qed thus "SD_efficient R' p" by (simp add: R'.SD_efficient_def) qed end locale sds_lowering = ex_post_efficient_sds agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds + fixes agents' alts' assumes agents'_subset: "agents' \ agents" and alts'_subset: "alts' \ alts" and agents'_nonempty [simp]: "agents' \ {}" and alts'_nonempty [simp]: "alts' \ {}" begin lemma finite_agents' [simp]: "finite agents'" using agents'_subset finite_agents by (rule finite_subset) lemma finite_alts' [simp]: "finite alts'" using alts'_subset finite_alts by (rule finite_subset) abbreviation lift :: "('agent, 'alt) pref_profile \ ('agent, 'alt) pref_profile" where "lift \ lift_pref_profile agents' alts' agents alts" definition lowered :: "('agent, 'alt) pref_profile \ 'alt lottery" where "lowered = sds \ lift" lemma lift_wf [simp, intro]: "pref_profile_wf agents' alts' R \ is_pref_profile (lift R)" using alts'_subset agents'_subset by (intro lift_pref_profile_wf) simp_all sublocale lowered: election agents' alts' by unfold_locales simp_all lemma preferred_alts_lift: "lowered.is_pref_profile R \ i \ agents \ x \ alts \ preferred_alts (lift R i) x = (if i \ agents' \ x \ alts' then preferred_alts (R i) x else alts)" using alts'_subset agents'_subset by (intro preferred_alts_lift_pref_profile) simp_all lemma pareto_losers_lift: "lowered.is_pref_profile R \ pareto_losers (lift R) = pareto_losers R \ (alts - alts')" using agents'_subset alts'_subset by (intro pareto_losers_lift_pref_profile) simp_all lemma lowered_lotteries: "lowered.lotteries \ lotteries" unfolding lotteries_on_def using alts'_subset by blast sublocale lowered: social_decision_scheme agents' alts' lowered proof fix R assume R_wf: "pref_profile_wf agents' alts' R" from R_wf have R'_wf: "pref_profile_wf agents alts (lift R)" by (rule lift_wf) show "lowered R \ lowered.lotteries" unfolding lotteries_on_def proof safe fix x assume "x \ set_pmf (lowered R)" hence x: "x \ set_pmf (sds (lift R))" by (simp add: lowered_def) with ex_post_efficient[OF R'_wf] have "x \ pareto_losers (lift R)" by blast with pareto_losers_lift[OF R_wf] have "x \ alts - alts'" by blast moreover from x have "x \ alts" using sds_wf[OF R'_wf] by (auto simp: lotteries_on_def) ultimately show "x \ alts'" by simp qed qed sublocale ex_post_efficient_sds agents' alts' lowered proof fix R assume R_wf: "lowered.is_pref_profile R" hence "is_pref_profile (lift R)" by simp have "set_pmf (lowered R) \ pareto_losers (lift R) = {}" unfolding lowered_def o_def by (intro ex_post_efficient lift_wf R_wf) also have "pareto_losers (lift R) = pareto_losers R \ (alts - alts')" by (intro pareto_losers_lift R_wf) finally show "set_pmf (lowered R) \ pareto_losers R = {}" by blast qed lemma lowered_in_lotteries [simp]: "lowered.is_pref_profile R \ lowered R \ lotteries" using lowered.sds_wf[of R] lowered_lotteries by blast end locale sds_lowering_anonymous = anonymous_sds agents alts sds + sds_lowering agents alts sds agents' alts' for agents :: "'agent set" and alts :: "'alt set" and sds agents' alts' begin sublocale lowered: anonymous_sds agents' alts' lowered proof fix \ R assume perm: "\ permutes agents'" and R_wf: "lowered.is_pref_profile R" from perm have "lift (R \ \) = lift R \ \" using agents'_subset by (rule lift_pref_profile_permute_agents) hence "sds (lift (R \ \)) = sds (lift R \ \)" by simp also from perm R_wf have "\ permutes agents" "is_pref_profile (lift R)" using agents'_subset by (auto dest: permutes_subset) from anonymous[OF this] have "sds (lift R \ \) = sds (lift R)" by (simp add: lowered_def) finally show "lowered (R \ \) = lowered R" unfolding lowered_def o_def . qed end locale sds_lowering_neutral = neutral_sds agents alts sds + sds_lowering agents alts sds agents' alts' for agents :: "'agent set" and alts :: "'alt set" and sds agents' alts' begin sublocale lowered: neutral_sds agents' alts' lowered proof fix \ R assume perm: "\ permutes alts'" and R_wf: "lowered.is_pref_profile R" from perm alts'_subset have "lift (permute_profile \ R) = permute_profile \ (lift R)" by (rule lift_pref_profile_permute_alts) hence "sds (lift (permute_profile \ R)) = sds (permute_profile \ (lift R))" by simp also from R_wf perm have "is_pref_profile (lift R)" by simp with perm alts'_subset have "sds (permute_profile \ (lift R)) = map_pmf \ (sds (lift R))" by (intro neutral) (auto intro: permutes_subset) finally show "lowered (permute_profile \ R) = map_pmf \ (lowered R)" by (simp add: lowered_def o_def) qed end locale sds_lowering_sd_efficient = sd_efficient_sds agents alts sds + sds_lowering agents alts sds agents' alts' for agents :: "'agent set" and alts :: "'alt set" and sds agents' alts' begin sublocale sd_efficient_sds agents' alts' lowered proof fix R assume R_wf: "lowered.is_pref_profile R" interpret R: pref_profile_wf agents' alts' R by fact from R_wf agents'_subset alts'_subset show "SD_efficient R (lowered R)" unfolding lowered_def o_def by (subst lift_SD_efficient_iff [symmetric]) (insert SD_efficient R_wf lowered.sds_wf[OF R_wf], auto simp: lowered_def) qed end locale sds_lowering_strategyproof = strategyproof_sds agents alts sds + sds_lowering agents alts sds agents' alts' for agents :: "'agent set" and alts :: "'alt set" and sds agents' alts' begin sublocale strategyproof_sds agents' alts' lowered proof (unfold_locales, safe) fix R i Ri' assume R_wf: "lowered.is_pref_profile R" and i: "i \ agents'" assume Ri': "total_preorder_on alts' Ri'" assume manipulable: "lowered.manipulable_profile R i Ri'" from i agents'_subset have i': "i \ agents" by blast interpret R: pref_profile_wf agents' alts' R by fact from R_wf interpret liftR: pref_profile_wf agents alts "lift R" by simp define lift_Ri' where "lift_Ri' x y \ x \ alts \ y \ alts \ (x = y \ x \ alts' \ (y \ alts' \ Ri' x y))" for x y define S where "S = (lift R)(i := lift_Ri')" from Ri' interpret Ri': total_preorder_on alts' Ri' . have wf_lift_Ri': "total_preorder_on alts lift_Ri'" using Ri'.total by unfold_locales (auto simp: lift_Ri'_def intro: Ri'.trans) from agents'_subset i have S_altdef: "S = lift (R(i := Ri'))" by (auto simp: fun_eq_iff lift_pref_profile_def lift_Ri'_def S_def) have "lowered (R(i := Ri')) \ lowered.lotteries" by (intro lowered.sds_wf R.wf_update i Ri') hence sds_S_wf: "sds S \ lowered.lotteries" by (simp add: S_altdef lowered_def) from manipulable have "lowered R \[SD (R i)] sds (lift (R(i := Ri')))" unfolding lowered.manipulable_profile_def by (simp add: lowered_def) also note S_altdef [symmetric] finally have "lowered R \[SD (lift R i)] sds S" using R_wf i lowered.sds_wf[OF R_wf] sds_S_wf by (subst lift_SD_strict_iff) (simp_all add: agents'_subset alts'_subset) hence "manipulable_profile (lift R) i lift_Ri'" by (simp add: manipulable_profile_def lowered_def S_def) with strategyproof[OF lift_wf[OF R_wf] i' wf_lift_Ri'] show False by contradiction qed end locale sds_lowering_anonymous_neutral_sdeff_stratproof = sds_lowering_anonymous + sds_lowering_neutral + sds_lowering_sd_efficient + sds_lowering_strategyproof end diff --git a/thys/Randomised_Social_Choice/SD_Efficiency.thy b/thys/Randomised_Social_Choice/SD_Efficiency.thy --- a/thys/Randomised_Social_Choice/SD_Efficiency.thy +++ b/thys/Randomised_Social_Choice/SD_Efficiency.thy @@ -1,392 +1,393 @@ (* Title: SD_Efficiency.thy Author: Manuel Eberl, TU München Characterisation of SD-efficiency. *) theory SD_Efficiency imports Complex_Main Preference_Profiles Lotteries Stochastic_Dominance begin (* TODO: Perhaps a general concept of "efficiency" can be introduced, parametrised over the way in which two lotteries are compared. *) context pref_profile_wf begin lemma SD_inefficient_support_subset: assumes inefficient: "\SD_efficient R p'" assumes support: "set_pmf p' \ set_pmf p" assumes lotteries: "p \ lotteries_on alts" shows "\SD_efficient R p" proof - from assms have p'_wf: "p' \ lotteries_on alts" by (simp add: lotteries_on_def) from inefficient obtain q' i where q': "q' \ lotteries_on alts" "i \ agents" "\i. i \ agents \ q' \[SD(R i)] p'" "q' \[SD(R i)] p'" unfolding SD_efficient_def' by blast have subset: "{x. pmf p' x > pmf q' x} \ set_pmf p'" by (auto simp: set_pmf_eq) also have "\ \ set_pmf p" by fact also have "\ \ alts" using lotteries by (simp add: lotteries_on_def) finally have finite: "finite {x. pmf p' x > pmf q' x}" using finite_alts by (rule finite_subset) define \ where "\ = Min (insert 1 {pmf p x / (pmf p' x - pmf q' x) |x. pmf p' x > pmf q' x})" define supp where "supp = set_pmf p \ set_pmf q'" from lotteries finite_alts q'(1) have finite_supp: "finite supp" by (auto simp: lotteries_on_def supp_def dest: finite_subset) from support have [simp]: "pmf p x = 0" "pmf p' x = 0" "pmf q' x = 0" if "x \ supp" for x using that by (auto simp: supp_def set_pmf_eq) from finite support subset have \: "\ > 0" unfolding \_def by (auto simp: field_simps set_pmf_eq') have nonneg: "pmf p x + \ * (pmf q' x - pmf p' x) \ 0" for x proof (cases "pmf p' x > pmf q' x") case True with finite have "\ \ pmf p x / (pmf p' x - pmf q' x)" unfolding \_def by (intro Min_le) auto with True show ?thesis by (simp add: field_simps) next case False with pmf_nonneg[of p x] \ show ?thesis by simp qed define q where "q = embed_pmf (\x. pmf p x + \ * (pmf q' x - pmf p' x))" have "(\\<^sup>+ x. ennreal (pmf p x + \ * (pmf q' x - pmf p' x)) \count_space UNIV) = 1" proof (subst nn_integral_count_space') have "(\x\supp. ennreal (pmf p x + \ * (pmf q' x - pmf p' x))) = ennreal ((\x\supp. pmf p x) + \ * ((\x\supp. pmf q' x) - (\x\supp. pmf p' x)))" by (subst sum_ennreal[OF nonneg], rule ennreal_cong) (auto simp: sum_subtractf ring_distribs sum.distrib sum_distrib_left) also from finite_supp support have "\ = 1" by (subst (1 2 3) sum_pmf_eq_1) (auto simp: supp_def) finally show "(\x\supp. ennreal (pmf p x + \ * (pmf q' x - pmf p' x))) = 1" . qed (insert nonneg finite_supp, simp_all) with nonneg have pmf_q: "pmf q x = pmf p x + \ * (pmf q' x - pmf p' x)" for x unfolding q_def by (intro pmf_embed_pmf) simp_all with support have support_q: "set_pmf q \ supp" unfolding supp_def by (auto simp: set_pmf_eq) with lotteries support q'(1) have q_wf: "q \ lotteries_on alts" by (auto simp add: lotteries_on_def supp_def) from support_q support have expected_utility: "measure_pmf.expectation q u = measure_pmf.expectation p u + \ * (measure_pmf.expectation q' u - measure_pmf.expectation p' u)" for u by (subst (1 2 3 4) integral_measure_pmf[OF finite_supp]) (auto simp: pmf_q supp_def sum.distrib sum_distrib_left sum_distrib_right sum_subtractf algebra_simps) have "q \[SD(R i)] p" if i: "i \ agents" for i proof - from i interpret finite_total_preorder_on alts "R i" by simp from i lotteries q'(1) q'(3)[OF i] q_wf p'_wf \ show ?thesis by (fastforce simp: SD_iff_expected_utilities_le expected_utility) qed moreover from \i \ agents\ interpret finite_total_preorder_on alts "R i" by simp from lotteries q'(1,4) q_wf p'_wf \ have "q \[SD(R i)] p" by (force simp: SD_iff_expected_utilities_le expected_utility not_le strongly_preferred_def) ultimately show ?thesis using q_wf \i \ agents\ unfolding SD_efficient_def' by blast qed lemma SD_efficient_support_subset: assumes "SD_efficient R p" "set_pmf p' \ set_pmf p" "p \ lotteries_on alts" shows "SD_efficient R p'" using SD_inefficient_support_subset[OF _ assms(2,3)] assms(1) by blast lemma SD_efficient_same_support: assumes "set_pmf p = set_pmf p'" "p \ lotteries_on alts" shows "SD_efficient R p \ SD_efficient R p'" using SD_inefficient_support_subset[of p p'] SD_inefficient_support_subset[of p' p] assms by (auto simp: lotteries_on_def) lemma SD_efficient_iff: assumes "p \ lotteries_on alts" shows "SD_efficient R p \ SD_efficient R (pmf_of_set (set_pmf p))" using assms finite_alts by (intro SD_efficient_same_support) (simp, subst set_pmf_of_set, auto simp: set_pmf_not_empty lotteries_on_def intro: finite_subset[OF _ finite_alts]) lemma SD_efficient_no_pareto_loser: assumes efficient: "SD_efficient R p" and p_wf: "p \ lotteries_on alts" shows "set_pmf p \ pareto_losers R = {}" proof - have "x \ pareto_losers R" if x: "x \ set_pmf p" for x proof - from x have "set_pmf (return_pmf x) \ set_pmf p" by auto from efficient this p_wf have "SD_efficient R (return_pmf x)" by (rule SD_efficient_support_subset) moreover from assms x have "x \ alts" by (auto simp: lotteries_on_def) ultimately show "x \ pareto_losers R" by (simp add: SD_efficient_singleton_iff) qed thus ?thesis by blast qed text \ Given two lotteries with the same support where one is strictly Pareto-SD-preferred to the other, one can construct a third lottery that is weakly Pareto-SD-preferred to the better lottery (and therefore strictly Pareto-SD-preferred to the worse lottery) and whose support is a strict subset of the original supports. \ lemma improve_lottery_support_subset: assumes "p \ lotteries_on alts" "q \ lotteries_on alts" "q \[Pareto(SD \ R)] p" "set_pmf p = set_pmf q" obtains r where "r \ lotteries_on alts" "r \[Pareto(SD \ R)] q" "set_pmf r \ set_pmf p" proof - have subset: "{x. pmf p x > pmf q x} \ set_pmf p" by (auto simp: set_pmf_eq) also have "\ \ alts" using assms by (simp add: lotteries_on_def) finally have finite: "finite {x. pmf p x > pmf q x}" using finite_alts by (rule finite_subset) from assms have "q \ p" by (auto simp: strongly_preferred_def) hence ex_less: "\x. pmf p x > pmf q x" by (rule pmf_neq_exists_less) define \ where "\ = Min {pmf p x / (pmf p x - pmf q x) |x. pmf p x > pmf q x}" define supp where "supp = set_pmf p" from assms finite_alts have finite_supp: "finite supp" by (auto simp: lotteries_on_def supp_def dest: finite_subset) from assms have [simp]: "pmf p x = 0" "pmf q x = 0" if "x \ supp" for x using that by (auto simp: supp_def set_pmf_eq) from finite subset ex_less have \: "\ \ 1" unfolding \_def by (intro Min.boundedI) (auto simp: field_simps pmf_nonneg) have nonneg: "pmf p x + \ * (pmf q x - pmf p x) \ 0" for x proof (cases "pmf p x > pmf q x") case True with finite have "\ \ pmf p x / (pmf p x - pmf q x)" unfolding \_def by (intro Min_le) auto with True show ?thesis by (simp add: field_simps) next case False with pmf_nonneg[of p x] \ show ?thesis by simp qed define r where "r = embed_pmf (\x. pmf p x + \ * (pmf q x - pmf p x))" have "(\\<^sup>+ x. ennreal (pmf p x + \ * (pmf q x - pmf p x)) \count_space UNIV) = 1" proof (subst nn_integral_count_space') have "(\x\supp. ennreal (pmf p x + \ * (pmf q x - pmf p x))) = ennreal ((\x\supp. pmf p x) + \ * ((\x\supp. pmf q x) - (\x\supp. pmf p x)))" by (subst sum_ennreal[OF nonneg], rule ennreal_cong) (auto simp: sum_subtractf ring_distribs sum.distrib sum_distrib_left) also from finite_supp have "\ = 1" by (subst (1 2 3) sum_pmf_eq_1) (auto simp: supp_def assms) finally show "(\x\supp. ennreal (pmf p x + \ * (pmf q x - pmf p x))) = 1" . qed (insert nonneg finite_supp, simp_all) with nonneg have pmf_r: "pmf r x = pmf p x + \ * (pmf q x - pmf p x)" for x unfolding r_def by (intro pmf_embed_pmf) simp_all with assms have "set_pmf r \ supp" unfolding supp_def by (auto simp: set_pmf_eq) from finite ex_less have "\ \ {pmf p x / (pmf p x - pmf q x) |x. pmf p x > pmf q x}" unfolding \_def by (intro Min_in) auto then obtain x where "\ = pmf p x / (pmf p x - pmf q x)" "pmf p x > pmf q x" by blast hence "pmf r x = 0" by (simp add: pmf_r field_simps) moreover from \pmf p x > pmf q x\ pmf_nonneg[of q x] have "pmf p x > 0" by linarith ultimately have "x \ set_pmf p - set_pmf r" by (auto simp: set_pmf_iff) with \set_pmf r \ supp\ have support_r: "set_pmf r \ set_pmf p" unfolding supp_def by blast from this assms have r_wf: "r \ lotteries_on alts" by (simp add: lotteries_on_def) have "r \[Pareto(SD\R)] q" unfolding SD.Pareto_iff unfolding o_def proof fix i assume i: "i \ agents" then interpret finite_total_preorder_on alts "R i" by simp show "r \[SD(R i)] q" proof (subst SD_iff_expected_utilities_le; safe?) fix u assume u: "is_vnm_utility u" from support_r have expected_utility_r: "measure_pmf.expectation r u = measure_pmf.expectation p u + \ * (measure_pmf.expectation q u - measure_pmf.expectation p u)" by (subst (1 2 3 4) integral_measure_pmf[OF finite_supp]) (auto simp: supp_def assms pmf_r sum.distrib sum_distrib_left sum_distrib_right sum_subtractf algebra_simps) from assms i have "q \[SD(R i)] p" by (simp add: SD.Pareto_strict_iff) with assms u have "measure_pmf.expectation q u \ measure_pmf.expectation p u" by (simp add: SD_iff_expected_utilities_le r_wf) hence "(\ - 1) * measure_pmf.expectation p u \ (\ - 1) * measure_pmf.expectation q u" using \ by (intro mult_left_mono) simp_all thus "measure_pmf.expectation q u \ measure_pmf.expectation r u" by (simp add: algebra_simps expected_utility_r) qed fact+ qed from that[OF r_wf this support_r] show ?thesis . qed subsection \Existence of SD-efficient lotteries\ text \ In this section, we will show that any lottery can be `improved' to an SD-efficient lottery, i.e. for any lottery, there exists an SD-efficient lottery that is weakly SD-preferred to the original one by all agents. \ context fixes p :: "'alt lottery" assumes lott: "p \ lotteries_on alts" begin private definition improve_lottery :: "'alt lottery \ 'alt lottery" where "improve_lottery q = (let A = {r\lotteries_on alts. r \[Pareto(SD\R)] q} in (SOME r. r \ A \ \(\r'\A. set_pmf r' \ set_pmf r)))" private lemma improve_lottery: assumes "\SD_efficient R q" defines "r \ improve_lottery q" shows "r \ lotteries_on alts" "r \[Pareto(SD\R)] q" "\r'. r' \ lotteries_on alts \ r' \[Pareto(SD\R)] q \ \(set_pmf r' \ set_pmf r)" proof - define A where "A = {r\lotteries_on alts. r \[Pareto(SD\R)] q}" have subset_alts: "X \ alts" if "X \ set_pmf`A" for X using that by (auto simp: A_def lotteries_on_def) have r_altdef: "r = (SOME r. r \ A \ \(\r'\A. set_pmf r' \ set_pmf r))" unfolding r_def improve_lottery_def Let_def A_def by simp from assms have nonempty: "A \ {}" by (auto simp: A_def SD_efficient_def) hence nonempty': "set_pmf`A \ {}" by simp have "set_pmf ` A \ Pow alts" by (auto simp: A_def lotteries_on_def) from finite_alts have wf: "wf {(X,Y). X \ Y \ Y \ alts}" by (rule finite_subset_wf) obtain X where "X \ set_pmf`A" "\Y. Y \ X \ X \ alts \ Y \ set_pmf ` A" by (rule wfE_min'[OF wf nonempty']) simp_all hence "\r. r \ A \ \(\r'\A. set_pmf r' \ set_pmf r)" by (auto simp: subset_alts[of X]) from someI_ex[OF this, folded r_altdef] show "r \ lotteries_on alts" "r \[Pareto(SD\R)] q" "\r'. r' \ lotteries_on alts \ r' \[Pareto(SD\R)] q \ \(set_pmf r' \ set_pmf r)" unfolding A_def by blast+ qed private fun sd_chain :: "nat \ 'alt lottery option" where "sd_chain 0 = Some p" | "sd_chain (Suc n) = (case sd_chain n of None \ None | Some p \ if SD_efficient R p then None else Some (improve_lottery p))" private lemma sd_chain_None_propagate: "m \ n \ sd_chain n = None \ sd_chain m = None" by (induction rule: inc_induct) simp_all private lemma sd_chain_Some_propagate: "m \ n \ sd_chain m = Some q \ \q'. sd_chain n = Some q'" by (cases "sd_chain n") (auto simp: sd_chain_None_propagate) private lemma sd_chain_NoneD: "sd_chain n = None \ \n p. sd_chain n = Some p \ SD_efficient R p" by (induction n) (auto split: option.splits if_splits) private lemma sd_chain_lottery: "sd_chain n = Some q \ q \ lotteries_on alts" by (induction n) (insert lott, auto split: option.splits if_splits simp: improve_lottery) private lemma sd_chain_Suc: assumes "sd_chain m = Some q" assumes "sd_chain (Suc m) = Some r" shows "q \[Pareto(SD\R)] r" using assms by (auto split: if_splits simp: improve_lottery) private lemma sd_chain_strictly_preferred: assumes "m < n" assumes "sd_chain m = Some q" assumes "sd_chain n = Some s" shows "q \[Pareto(SD\R)] s" using assms proof (induction arbitrary: q rule: strict_inc_induct) case (base k q) with sd_chain_Suc[of k q s] show ?case by (simp del: sd_chain.simps add: o_def) next case (step k q) from step.hyps have "Suc k \ n" by simp from sd_chain_Some_propagate[OF this, of s] step.prems obtain r where r: "sd_chain (Suc k) = Some r" by (auto simp del: sd_chain.simps) with step.prems have "q \[Pareto (SD \ R)] r" by (intro sd_chain_Suc) moreover from r step.prems have "r \[Pareto (SD \ R)] s" by (intro step.IH) simp_all ultimately show ?case by (rule SD.Pareto.strict_trans) qed private lemma sd_chain_preferred: assumes "m \ n" assumes "sd_chain m = Some q" assumes "sd_chain n = Some s" shows "q \[Pareto(SD\R)] s" proof (cases "m < n") case True from sd_chain_strictly_preferred[OF this assms(2,3)] show ?thesis by (simp add: strongly_preferred_def) next case False with assms show ?thesis by (auto intro: SD.Pareto.refl sd_chain_lottery) qed lemma SD_efficient_lottery_exists: obtains q where "q \ lotteries_on alts" "q \[Pareto(SD\R)] p" "SD_efficient R q" proof - consider "\n. sd_chain n = None" | "\n. \q. sd_chain n = Some q" using option.exhaust by metis thus ?thesis proof cases case 1 define m where "m = (LEAST m. sd_chain m = None)" define k where "k = m - 1" from LeastI_ex[OF 1] have m: "sd_chain m = None" by (simp add: m_def) from m have nz: "m \ 0" by (intro notI) simp_all from nz have m_altdef: "m = Suc k" by (simp add: k_def) from nz Least_le[of "\m. sd_chain m = None" "m - 1", folded m_def] obtain q where q: "sd_chain k = Some q" by (cases "sd_chain (m - 1)") (auto simp: k_def) from sd_chain_preferred[OF _ sd_chain.simps(1) this] have "q \[Pareto(SD\R)] p" by simp moreover from q have "q \ lotteries_on alts" by (simp add: sd_chain_lottery) moreover from q m have "SD_efficient R q" by (auto split: if_splits simp: m_altdef) ultimately show ?thesis using that[of q] by blast next case 2 have "range (set_pmf \ the \ sd_chain) \ Pow alts" unfolding o_def proof safe fix n x assume A: "x \ set_pmf (the (sd_chain n))" from 2 obtain q where "sd_chain n = Some q" by auto with sd_chain_lottery[of n q] have "set_pmf (the (sd_chain n)) \ alts" by (simp add: lotteries_on_def) with A show "x \ alts" by blast qed hence "finite (range (set_pmf \ the \ sd_chain))" by (rule finite_subset) simp_all from pigeonhole_infinite[OF infinite_UNIV_nat this] obtain m where "infinite {n. set_pmf (the (sd_chain n)) = set_pmf (the (sd_chain m))}" by auto hence "infinite ({n. set_pmf (the (sd_chain n)) = set_pmf (the (sd_chain m))} - {k. \(k > m)})" by (simp add: not_less) hence "({n. set_pmf (the (sd_chain n)) = set_pmf (the (sd_chain m))} - {k. \(k > m)}) \ {}" by (intro notI) simp_all then obtain n where mn: "n > m" "set_pmf (the (sd_chain n)) = set_pmf (the (sd_chain m))" by blast from 2 obtain p q where pq: "sd_chain m = Some p" "sd_chain n = Some q" by blast from mn pq have supp_eq: "set_pmf p = set_pmf q" by simp from mn(1) pq have less: "p \[Pareto(SD\R)] q" by (rule sd_chain_strictly_preferred) from \m < n\ have "n > 0" by simp with \sd_chain n = Some q\ sd_chain.simps(2)[of "n - 1"] obtain r where r: "\SD_efficient R r" "q = improve_lottery r" by (auto simp del: sd_chain.simps split: if_splits option.splits) from pq have "p \ lotteries_on alts" "q \ lotteries_on alts" by (simp_all add: sd_chain_lottery) - from improve_lottery_support_subset[OF this less supp_eq] guess s . note s = this + from improve_lottery_support_subset[OF this less supp_eq] + obtain s where s: "s \ lotteries_on alts" "Pareto (SD \ R) q s" "set_pmf s \ set_pmf p" . from improve_lottery(2)[of r] r s have "s \[Pareto(SD\R)] r" by (auto intro: SD.Pareto.strict_weak_trans) from improve_lottery(3)[OF r(1) s(1) this] supp_eq r have "\set_pmf s \ set_pmf p" by simp with s(3) show ?thesis by contradiction qed qed end lemma assumes "p \ lotteries_on alts" shows "\q\lotteries_on alts. q \[Pareto(SD\R)] p \ SD_efficient R q" using SD_efficient_lottery_exists[OF assms] by blast end end diff --git a/thys/Randomised_Social_Choice/Social_Decision_Schemes.thy b/thys/Randomised_Social_Choice/Social_Decision_Schemes.thy --- a/thys/Randomised_Social_Choice/Social_Decision_Schemes.thy +++ b/thys/Randomised_Social_Choice/Social_Decision_Schemes.thy @@ -1,411 +1,412 @@ (* Title: Social Decision Schemes.thy Author: Manuel Eberl, TU München Definitions of Social Decision Schemes, their properties, and related concepts *) section \Social Decision Schemes\ theory Social_Decision_Schemes imports Complex_Main "HOL-Probability.Probability" Preference_Profiles Elections Order_Predicates Stochastic_Dominance SD_Efficiency begin subsection \Basic Social Choice definitions\ context election begin text \ The set of lotteries, i.e. the probability mass functions on the type @{typ "'alt"} whose support is a subset of the alternative set. \ abbreviation lotteries where "lotteries \ lotteries_on alts" text \ The probability that a lottery returns an alternative that is in the given set \ abbreviation lottery_prob :: "'alt lottery \ 'alt set \ real" where "lottery_prob \ measure_pmf.prob" lemma lottery_prob_alts_superset: assumes "p \ lotteries" "alts \ A" shows "lottery_prob p A = 1" using assms by (subst measure_pmf.prob_eq_1) (auto simp: AE_measure_pmf_iff lotteries_on_def) lemma lottery_prob_alts: "p \ lotteries \ lottery_prob p alts = 1" by (rule lottery_prob_alts_superset) simp_all end text \ In the context of an election, a preference profile is a function that assigns to each agent her preference relation (which is a total preorder) \ subsection \Social Decision Schemes\ text \ In the context of an election, a Social Decision Scheme (SDS) is a function that maps preference profiles to lotteries on the alternatives. \ locale social_decision_scheme = election agents alts for agents :: "'agent set" and alts :: "'alt set" + fixes sds :: "('agent, 'alt) pref_profile \ 'alt lottery" assumes sds_wf: "is_pref_profile R \ sds R \ lotteries" subsection \Anonymity\ text \ An SDS is anonymous if permuting the agents in the input does not change the result. \ locale anonymous_sds = social_decision_scheme agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds + assumes anonymous: "\ permutes agents \ is_pref_profile R \ sds (R \ \) = sds R" begin lemma anonymity_prefs_from_table: assumes "prefs_from_table_wf agents alts xs" "prefs_from_table_wf agents alts ys" assumes "mset (map snd xs) = mset (map snd ys)" shows "sds (prefs_from_table xs) = sds (prefs_from_table ys)" proof - - from prefs_from_table_agent_permutation[OF assms] guess \ . + from assms obtain \ where "\ permutes agents" "prefs_from_table xs \ \ = prefs_from_table ys" + by (rule prefs_from_table_agent_permutation) with anonymous[of \, of "prefs_from_table xs"] assms(1) show ?thesis by (simp add: pref_profile_from_tableI) qed context begin qualified lemma anonymity_prefs_from_table_aux: assumes "R1 = prefs_from_table xs" "prefs_from_table_wf agents alts xs" assumes "R2 = prefs_from_table ys" "prefs_from_table_wf agents alts ys" assumes "mset (map snd xs) = mset (map snd ys)" shows "sds R1 = sds R2" unfolding assms(1,3) by (rule anonymity_prefs_from_table) (simp_all add: assms del: mset_map) end end subsection \Neutrality\ text \ An SDS is neutral if permuting the alternatives in the input does not change the result, modulo the equivalent permutation in the output lottery. \ locale neutral_sds = social_decision_scheme agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds + assumes neutral: "\ permutes alts \ is_pref_profile R \ sds (permute_profile \ R) = map_pmf \ (sds R)" begin text \ Alternative formulation of neutrality that shows that our definition is equivalent to that in the paper. \ lemma neutral': assumes "\ permutes alts" assumes "is_pref_profile R" assumes "a \ alts" shows "pmf (sds (permute_profile \ R)) (\ a) = pmf (sds R) a" proof - from assms have A: "set_pmf (sds R) \ alts" using sds_wf by (simp add: lotteries_on_def) from assms(1,2) have "pmf (sds (permute_profile \ R)) (\ a) = pmf (map_pmf \ (sds R)) (\ a)" by (subst neutral) simp_all also from assms have "\ = pmf (sds R) a" by (intro pmf_map_inj') (simp_all add: permutes_inj) finally show ?thesis . qed end locale an_sds = anonymous_sds agents alts sds + neutral_sds agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds begin lemma sds_anonymous_neutral: assumes perm: "\ permutes alts" and wf: "is_pref_profile R1" "is_pref_profile R2" assumes eq: "anonymous_profile R1 = image_mset (map ((`) \)) (anonymous_profile R2)" shows "sds R1 = map_pmf \ (sds R2)" proof - interpret R1: pref_profile_wf agents alts R1 by fact interpret R2: pref_profile_wf agents alts R2 by fact from perm have wf': "is_pref_profile (permute_profile \ R2)" by (rule R2.wf_permute_alts) from eq perm have "anonymous_profile R1 = anonymous_profile (permute_profile \ R2)" by (simp add: R2.anonymous_profile_permute) from anonymous_profile_agent_permutation[OF this wf(1) wf'] obtain \ where "\ permutes agents" "permute_profile \ R2 \ \ = R1" by auto have "sds (permute_profile \ R2 \ \) = sds (permute_profile \ R2)" by (rule anonymous) fact+ also have "\ = map_pmf \ (sds R2)" by (rule neutral) fact+ also have "permute_profile \ R2 \ \ = R1" by fact finally show ?thesis . qed lemma sds_anonymous_neutral': assumes perm: "\ permutes alts" and wf: "is_pref_profile R1" "is_pref_profile R2" assumes eq: "anonymous_profile R1 = image_mset (map ((`) \)) (anonymous_profile R2)" shows "pmf (sds R1) (\ x) = pmf (sds R2) x" proof - have "sds R1 = map_pmf \ (sds R2)" by (intro sds_anonymous_neutral) fact+ also have "pmf \ (\ x) = pmf (sds R2) x" by (intro pmf_map_inj' permutes_inj[OF perm]) finally show ?thesis . qed lemma sds_automorphism: assumes perm: "\ permutes alts" and wf: "is_pref_profile R" assumes eq: "image_mset (map ((`) \)) (anonymous_profile R) = anonymous_profile R" shows "map_pmf \ (sds R) = sds R" using sds_anonymous_neutral[OF perm wf wf eq [symmetric]] .. end lemma an_sds_automorphism_aux: assumes wf: "prefs_from_table_wf agents alts yss" "R \ prefs_from_table yss" assumes an: "an_sds agents alts sds" assumes eq: "mset (map ((map ((`) (permutation_of_list xs))) \ snd) yss) = mset (map snd yss)" assumes perm: "set (map fst xs) \ alts" "set (map snd xs) = set (map fst xs)" "distinct (map fst xs)" and x: "x \ alts" "y = permutation_of_list xs x" shows "pmf (sds R) x = pmf (sds R) y" proof - note perm = list_permutesI[OF perm] let ?\ = "permutation_of_list xs" note perm' = permutation_of_list_permutes [OF perm] from wf have wf': "pref_profile_wf agents alts R" by (simp add: pref_profile_from_tableI) then interpret R: pref_profile_wf agents alts R . from perm' interpret R': pref_profile_wf agents alts "permute_profile ?\ R" by (simp add: R.wf_permute_alts) from an interpret an_sds agents alts sds . from eq wf have eq': "image_mset (map ((`) ?\)) (anonymous_profile R) = anonymous_profile R" by (simp add: anonymise_prefs_from_table mset_map multiset.map_comp) from perm' x have "pmf (sds R) x = pmf (map_pmf ?\ (sds R)) (?\ x)" by (simp add: pmf_map_inj' permutes_inj) also from eq' x wf' perm' have "map_pmf ?\ (sds R) = sds R" by (intro sds_automorphism) (simp_all add: R.anonymous_profile_permute pref_profile_from_tableI) finally show ?thesis using x by simp qed subsection \Ex-post efficiency\ locale ex_post_efficient_sds = social_decision_scheme agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds + assumes ex_post_efficient: "is_pref_profile R \ set_pmf (sds R) \ pareto_losers R = {}" begin lemma ex_post_efficient': assumes "is_pref_profile R" "y \[Pareto(R)] x" shows "pmf (sds R) x = 0" using ex_post_efficient[of R] assms by (auto simp: set_pmf_eq pareto_losers_def) lemma ex_post_efficient'': assumes "is_pref_profile R" "i \ agents" "\i\agents. y \[R i] x" "\y \[R i] x" shows "pmf (sds R) x = 0" proof - from assms(1) interpret pref_profile_wf agents alts R . from assms(2-) show ?thesis by (intro ex_post_efficient'[OF assms(1), of _ y]) (auto simp: Pareto_iff strongly_preferred_def) qed end subsection \SD efficiency\ text \ An SDS is SD-efficient if it returns an SD-efficient lottery for every preference profile, i.e. if the SDS outputs a lottery, it is never the case that there is another lottery that is weakly preferred by all agents an strictly preferred by at least one agent. \ locale sd_efficient_sds = social_decision_scheme agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds + assumes SD_efficient: "is_pref_profile R \ SD_efficient R (sds R)" begin text \ An alternative formulation of SD-efficiency that is somewhat more convenient to use. \ lemma SD_efficient': assumes "is_pref_profile R" "q \ lotteries" assumes "\i. i \ agents \ q \[SD(R i)] sds R" "i \ agents" "q \[SD(R i)] sds R" shows P proof - interpret pref_profile_wf agents alts R by fact show ?thesis using SD_efficient[of R] sds_wf[OF assms(1)] assms unfolding SD_efficient_def' by blast qed text \ Any SD-efficient SDS is also ex-post efficient. \ sublocale ex_post_efficient_sds proof unfold_locales fix R :: "('agent, 'alt) pref_profile" assume R_wf: "is_pref_profile R" interpret pref_profile_wf agents alts R by fact from R_wf show "set_pmf (sds R) \ pareto_losers R = {}" by (intro SD_efficient_no_pareto_loser SD_efficient sds_wf) qed text \ The following rule can be used to derive facts from inefficient supports: If a set of alternatives is an inefficient support, at least one of the alternatives in it must receive probability 0. \ lemma SD_inefficient_support: assumes A: "A \ {}" "A \ alts" and inefficient: "\SD_efficient R (pmf_of_set A)" assumes wf: "is_pref_profile R" shows "\x\A. pmf (sds R) x = 0" proof (rule ccontr) interpret pref_profile_wf agents alts R by fact assume "\(\x\A. pmf (sds R) x = 0)" with A have "set_pmf (pmf_of_set A) \ set_pmf (sds R)" by (subst set_pmf_of_set) (auto simp: set_pmf_eq intro: finite_subset[OF _ finite_alts]) from inefficient and this have "\SD_efficient R (sds R)" by (rule SD_inefficient_support_subset) (simp add: wf sds_wf) moreover from SD_efficient wf have "SD_efficient R (sds R)" . ultimately show False by contradiction qed lemma SD_inefficient_support': assumes wf: "is_pref_profile R" assumes A: "A \ {}" "A \ alts" and wit: "p \ lotteries" "\i\agents. p \[SD(R i)] pmf_of_set A" "i \ agents" "\p \[SD(R i)] pmf_of_set A" shows "\x\A. pmf (sds R) x = 0" proof (rule SD_inefficient_support) from wf interpret pref_profile_wf agents alts R . from wit show "\SD_efficient R (pmf_of_set A)" by (intro SD_inefficientI') (auto intro!: bexI[of _ i] simp: strongly_preferred_def) qed fact+ end subsection \Weak strategyproofness\ context social_decision_scheme begin text \ The SDS is said to be manipulable for a particular preference profile, a particular agent, and a particular alternative preference ordering for that agent if the lottery obtained if the agent submits the alternative preferences strictly SD-dominates that obtained if the original preferences are submitted. (SD-dominated w.r.t. the original preferences) \ definition manipulable_profile :: "('agent, 'alt) pref_profile \ 'agent \ 'alt relation \ bool" where "manipulable_profile R i Ri' \ sds (R(i := Ri')) \[SD (R i)] sds R" end text \ An SDS is weakly strategyproof (or just strategyproof) if it is not manipulable for any combination of preference profiles, agents, and alternative preference relations. \ locale strategyproof_sds = social_decision_scheme agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds + assumes strategyproof: "is_pref_profile R \ i \ agents \ total_preorder_on alts Ri' \ \manipulable_profile R i Ri'" subsection \Strong strategyproofness\ context social_decision_scheme begin text \ The SDS is said to be strongly strategyproof for a particular preference profile, a particular agent, and a particular alternative preference ordering for that agent if the lottery obtained if the agent submits the alternative preferences is SD-dominated by the one obtained if the original preferences are submitted. (SD-dominated w.r.t. the original preferences) In other words: the SDS is strategyproof w.r.t the preference profile $R$ and the agent $i$ and the alternative preference relation $R_i'$ if the lottery for obtained for $R$ is at least as good for $i$ as the lottery obtained when $i$ misrepresents her preferences as $R_i'$. \ definition strongly_strategyproof_profile :: "('agent, 'alt) pref_profile \ 'agent \ 'alt relation \ bool" where "strongly_strategyproof_profile R i Ri' \ sds R \[SD (R i)] sds (R(i := Ri'))" lemma strongly_strategyproof_profileI [intro]: assumes "is_pref_profile R" "total_preorder_on alts Ri'" "i \ agents" assumes "\x. x \ alts \ lottery_prob (sds (R(i := Ri'))) (preferred_alts (R i) x) \ lottery_prob (sds R) (preferred_alts (R i) x)" shows "strongly_strategyproof_profile R i Ri'" proof - interpret pref_profile_wf agents alts R by fact show ?thesis unfolding strongly_strategyproof_profile_def by rule (auto intro!: sds_wf assms pref_profile_wf.wf_update) qed lemma strongly_strategyproof_imp_not_manipulable: assumes "strongly_strategyproof_profile R i Ri'" shows "\manipulable_profile R i Ri'" using assms unfolding strongly_strategyproof_profile_def manipulable_profile_def by (auto simp: strongly_preferred_def) end text \ An SDS is strongly strategyproof if it is strongly strategyproof for all combinations of preference profiles, agents, and alternative preference relations. \ locale strongly_strategyproof_sds = social_decision_scheme agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds + assumes strongly_strategyproof: "is_pref_profile R \ i \ agents \ total_preorder_on alts Ri' \ strongly_strategyproof_profile R i Ri'" begin text \ Any SDS that is strongly strategyproof is also weakly strategyproof. \ sublocale strategyproof_sds by unfold_locales (simp add: strongly_strategyproof_imp_not_manipulable strongly_strategyproof) end locale strategyproof_an_sds = strategyproof_sds agents alts sds + an_sds agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds end diff --git a/thys/Randomised_Social_Choice/Stochastic_Dominance.thy b/thys/Randomised_Social_Choice/Stochastic_Dominance.thy --- a/thys/Randomised_Social_Choice/Stochastic_Dominance.thy +++ b/thys/Randomised_Social_Choice/Stochastic_Dominance.thy @@ -1,530 +1,537 @@ section \Stochastic Dominance\ theory Stochastic_Dominance imports Complex_Main "HOL-Probability.Probability" Lotteries Preference_Profiles Utility_Functions begin subsection \Definition of Stochastic Dominance\ text \ This is the definition of stochastic dominance. It lifts a preference relation on alternatives to the stochastic dominance ordering on lotteries. \ definition SD :: "'alt relation \ 'alt lottery relation" where "p \[SD(R)] q \ p \ lotteries_on {x. R x x} \ q \ lotteries_on {x. R x x} \ (\x. R x x \ measure_pmf.prob p {y. y \[R] x} \ measure_pmf.prob q {y. y \[R] x})" lemma SD_empty [simp]: "SD (\_ _. False) = (\_ _. False)" by (auto simp: fun_eq_iff SD_def lotteries_on_def set_pmf_not_empty) text \ Stochastic dominance over any relation is a preorder. \ lemma SD_refl: "p \[SD(R)] p \ p \ lotteries_on {x. R x x}" by (simp add: SD_def) lemma SD_trans [simp, trans]: "p \[SD(R)] q \ q \[SD(R)] r \ p \[SD(R)] r" unfolding SD_def by (auto intro: order.trans) lemma SD_is_preorder: "preorder_on (lotteries_on {x. R x x}) (SD R)" by unfold_locales (auto simp: SD_def intro: order.trans) context preorder_on begin lemma SD_preorder: "p \[SD(le)] q \ p \ lotteries_on carrier \ q \ lotteries_on carrier \ (\x\carrier. measure_pmf.prob p (preferred_alts le x) \ measure_pmf.prob q (preferred_alts le x))" by (simp add: SD_def preferred_alts_def carrier_eq) lemma SD_preorderI [intro?]: assumes "p \ lotteries_on carrier" "q \ lotteries_on carrier" assumes "\x. x \ carrier \ measure_pmf.prob p (preferred_alts le x) \ measure_pmf.prob q (preferred_alts le x)" shows "p \[SD(le)] q" using assms by (simp add: SD_preorder) lemma SD_preorderD: assumes "p \[SD(le)] q" shows "p \ lotteries_on carrier" "q \ lotteries_on carrier" and "\x. x \ carrier \ measure_pmf.prob p (preferred_alts le x) \ measure_pmf.prob q (preferred_alts le x)" using assms unfolding SD_preorder by simp_all lemma SD_refl' [simp]: "p \[SD(le)] p \ p \ lotteries_on carrier" by (simp add: SD_def carrier_eq) lemma SD_is_preorder': "preorder_on (lotteries_on carrier) (SD(le))" using SD_is_preorder[of le] by (simp add: carrier_eq) lemma SD_singleton_left: assumes "x \ carrier" "q \ lotteries_on carrier" shows "return_pmf x \[SD(le)] q \ (\y\set_pmf q. x \[le] y)" proof assume SD: "return_pmf x \[SD(le)] q" from assms SD_preorderD(3)[OF SD, of x] have "measure_pmf.prob (return_pmf x) (preferred_alts le x) \ measure_pmf.prob q (preferred_alts le x)" by simp also from assms have "measure_pmf.prob (return_pmf x) (preferred_alts le x) = 1" by (simp add: indicator_def) finally have "AE y in q. y \[le] x" by (simp add: measure_pmf.measure_ge_1_iff measure_pmf.prob_eq_1 preferred_alts_def) thus "\y\set_pmf q. y \[le] x" by (simp add: AE_measure_pmf_iff) next assume A: "\y\set_pmf q. x \[le] y" show "return_pmf x \[SD(le)] q" proof (rule SD_preorderI) fix y assume y: "y \ carrier" show "measure_pmf.prob (return_pmf x) (preferred_alts le y) \ measure_pmf.prob q (preferred_alts le y)" proof (cases "y \[le] x") case True from True A have "measure_pmf.prob q (preferred_alts le y) = 1" by (auto simp: AE_measure_pmf_iff measure_pmf.prob_eq_1 preferred_alts_def intro: trans) thus ?thesis by simp qed (simp_all add: preferred_alts_def indicator_def measure_nonneg) qed (insert assms, simp_all add: lotteries_on_def) qed lemma SD_singleton_right: assumes x: "x \ carrier" and q: "q \ lotteries_on carrier" shows "q \[SD(le)] return_pmf x \ (\y\set_pmf q. y \[le] x)" proof safe fix y assume SD: "q \[SD(le)] return_pmf x" and y: "y \ set_pmf q" from y assms have [simp]: "y \ carrier" by (auto simp: lotteries_on_def) from y have "0 < measure_pmf.prob q (preferred_alts le y)" by (rule measure_pmf_posI) simp_all also have "\ \ measure_pmf.prob (return_pmf x) (preferred_alts le y)" by (rule SD_preorderD(3)[OF SD]) simp_all finally show "y \[le] x" by (auto simp: indicator_def preferred_alts_def split: if_splits) next assume A: "\y\set_pmf q. y \[le] x" show "q \[SD(le)] return_pmf x" proof (rule SD_preorderI) fix y assume y: "y \ carrier" show "measure_pmf.prob q (preferred_alts le y) \ measure_pmf.prob (return_pmf x) (preferred_alts le y)" proof (cases "y \[le] x") case False with A show ?thesis by (auto simp: preferred_alts_def indicator_def measure_le_0_iff measure_pmf.prob_eq_0 AE_measure_pmf_iff intro: trans) qed (simp_all add: indicator_def preferred_alts_def) qed (insert assms, simp_all add: lotteries_on_def) qed lemma SD_strict_singleton_left: assumes "x \ carrier" "q \ lotteries_on carrier" shows "return_pmf x \[SD(le)] q \ (\y\set_pmf q. x \[le] y) \ (\y\set_pmf q. (x \[le] y))" using assms by (auto simp add: strongly_preferred_def SD_singleton_left SD_singleton_right) lemma SD_strict_singleton_right: assumes "x \ carrier" "q \ lotteries_on carrier" shows "q \[SD(le)] return_pmf x \ (\y\set_pmf q. y \[le] x) \ (\y\set_pmf q. (y \[le] x))" using assms by (auto simp add: strongly_preferred_def SD_singleton_left SD_singleton_right) lemma SD_singleton [simp]: "x \ carrier \ y \ carrier \ return_pmf x \[SD(le)] return_pmf y \ x \[le] y" by (subst SD_singleton_left) (simp_all add: lotteries_on_def) lemma SD_strict_singleton [simp]: "x \ carrier \ y \ carrier \ return_pmf x \[SD(le)] return_pmf y \ x \[le] y" by (simp add: strongly_preferred_def) end context pref_profile_wf begin context fixes i assumes i: "i \ agents" begin interpretation Ri: preorder_on alts "R i" by (simp add: i) lemmas SD_singleton_left = Ri.SD_singleton_left lemmas SD_singleton_right = Ri.SD_singleton_right lemmas SD_strict_singleton_left = Ri.SD_strict_singleton_left lemmas SD_strict_singleton_right = Ri.SD_strict_singleton_right lemmas SD_singleton = Ri.SD_singleton lemmas SD_strict_singleton = Ri.SD_strict_singleton end end lemmas (in pref_profile_wf) [simp] = SD_singleton SD_strict_singleton subsection \Stochastic Dominance for preference profiles\ context pref_profile_wf begin lemma SD_pref_profile: assumes "i \ agents" shows "p \[SD(R i)] q \ p \ lotteries_on alts \ q \ lotteries_on alts \ (\x\alts. measure_pmf.prob p (preferred_alts (R i) x) \ measure_pmf.prob q (preferred_alts (R i) x))" proof - from assms interpret total_preorder_on alts "R i" by simp have "preferred_alts (R i) x = {y. y \[R i] x}" for x using not_outside by (auto simp: preferred_alts_def) thus ?thesis by (simp add: SD_preorder preferred_alts_def) qed lemma SD_pref_profileI [intro?]: assumes "i \ agents" "p \ lotteries_on alts" "q \ lotteries_on alts" assumes "\x. x \ alts \ measure_pmf.prob p (preferred_alts (R i) x) \ measure_pmf.prob q (preferred_alts (R i) x)" shows "p \[SD(R i)] q" using assms by (simp add: SD_pref_profile) lemma SD_pref_profileD: assumes "i \ agents" "p \[SD(R i)] q" shows "p \ lotteries_on alts" "q \ lotteries_on alts" and "\x. x \ alts \ measure_pmf.prob p (preferred_alts (R i) x) \ measure_pmf.prob q (preferred_alts (R i) x)" using assms by (simp_all add: SD_pref_profile) end subsection \SD efficient lotteries\ definition SD_efficient :: "('agent, 'alt) pref_profile \ 'alt lottery \ bool" where SD_efficient_auxdef: "SD_efficient R p \ \(\q\lotteries_on {x. \i. R i x x}. q \[Pareto (SD \ R)] p)" context pref_profile_wf begin sublocale SD: preorder_family agents "lotteries_on alts" "SD \ R" unfolding o_def by (intro preorder_family.intro SD_is_preorder) (simp_all add: preorder_on.SD_is_preorder' prefs_undefined') text \ A lottery is considered SD-efficient if there is no other lottery such that all agents weakly prefer the other lottery (w.r.t. stochastic dominance) and at least one agent strongly prefers the other lottery. \ lemma SD_efficient_def: "SD_efficient R p \ \(\q\lotteries_on alts. q \[Pareto (SD \ R)] p)" proof - have "SD_efficient R p \ \(\q\lotteries_on {x. \i. R i x x}. q \[Pareto (SD \ R)] p)" unfolding SD_efficient_auxdef .. also from nonempty_agents obtain i where i: "i \ agents" by blast with preorder_on.refl[of alts "R i"] have "{x. \i. R i x x} = alts" by (auto intro!: exI[of _ i] not_outside) finally show ?thesis . qed lemma SD_efficient_def': "SD_efficient R p \ \(\q\lotteries_on alts. (\i\agents. q \[SD(R i)] p) \ (\i\agents. q \[SD(R i)] p))" unfolding SD_efficient_def SD.Pareto_iff strongly_preferred_def [abs_def] by auto lemma SD_inefficientI: assumes "q \ lotteries_on alts" "\i. i \ agents \ q \[SD(R i)] p" "i \ agents" "q \[SD(R i)] p" shows "\SD_efficient R p" using assms unfolding SD_efficient_def' by blast lemma SD_inefficientI': assumes "q \ lotteries_on alts" "\i. i \ agents \ q \[SD(R i)] p" "\i \ agents. q \[SD(R i)] p" shows "\SD_efficient R p" using assms unfolding SD_efficient_def' by blast lemma SD_inefficientE: assumes "\SD_efficient R p" obtains q i where "q \ lotteries_on alts" "\i. i \ agents \ q \[SD(R i)] p" "i \ agents" "q \[SD(R i)] p" using assms unfolding SD_efficient_def' by blast lemma SD_efficientD: assumes "SD_efficient R p" "q \ lotteries_on alts" and "\i. i \ agents \ q \[SD(R i)] p" "\i\agents. \(q \[SD(R i)] p)" shows False using assms unfolding SD_efficient_def' strongly_preferred_def by blast lemma SD_efficient_singleton_iff: assumes [simp]: "x \ alts" shows "SD_efficient R (return_pmf x) \ x \ pareto_losers R" proof - { - assume x: "x \ pareto_losers R" - from pareto_losersE[OF x] guess y . note y = this - from y have "\SD_efficient R (return_pmf x)" + assume "x \ pareto_losers R" + then obtain y where "y \ alts" "x \[Pareto R] y" + by (rule pareto_losersE) + then have "\SD_efficient R (return_pmf x)" by (intro SD_inefficientI'[of "return_pmf y"]) (force simp: Pareto_strict_iff)+ } moreover { assume "\SD_efficient R (return_pmf x)" - from SD_inefficientE[OF this] guess q i . note q = this - moreover from q obtain y where "y \ set_pmf q" "y \[R i] x" + from SD_inefficientE[OF this] obtain q i + where q: + "q \ lotteries_on alts" + "\i. i \ agents \ SD (R i) (return_pmf x) q" + "i \ agents" + "return_pmf x \[SD (R i)] q" + by blast + from q obtain y where "y \ set_pmf q" "y \[R i] x" by (auto simp: SD_strict_singleton_left) - ultimately have "y \[Pareto(R)] x" + with q have "y \[Pareto(R)] x" by (fastforce simp: Pareto_strict_iff SD_singleton_left) hence "x \ pareto_losers R" by simp } ultimately show ?thesis by blast qed end subsection \Equivalence proof\ text \ We now show that a lottery is preferred w.r.t. Stochastic Dominance iff it yields more expected utility for all compatible utility functions. \ context finite_total_preorder_on begin abbreviation "is_vnm_utility \ vnm_utility carrier le" lemma utility_weak_ranking_index: "is_vnm_utility (\x. real (length (weak_ranking le) - weak_ranking_index x))" proof fix x y assume xy: "x \ carrier" "y \ carrier" with this[THEN nth_weak_ranking_index(1)] this[THEN nth_weak_ranking_index(2)] show "(real (length (weak_ranking le) - weak_ranking_index x) \ real (length (weak_ranking le) - weak_ranking_index y)) \ le x y" by (simp add: le_diff_iff') qed (* TODO: one direction could probably be generalised to weakly consistent utility functions *) lemma SD_iff_expected_utilities_le: assumes "p \ lotteries_on carrier" "q \ lotteries_on carrier" shows "p \[SD(le)] q \ (\u. is_vnm_utility u \ measure_pmf.expectation p u \ measure_pmf.expectation q u)" proof safe fix u assume SD: "p \[SD(le)] q" and is_utility: "is_vnm_utility u" from is_utility interpret vnm_utility carrier le u . define xs where "xs = weak_ranking le" have le: "is_weak_ranking xs" "le = of_weak_ranking xs" by (simp_all add: xs_def weak_ranking_total_preorder) let ?pref = "\p x. measure_pmf.prob p {y. x \[le] y}" and ?pref' = "\p x. measure_pmf.prob p {y. x \[le] y}" define f where "f i = (SOME x. x \ xs ! i)" for i have xs_wf: "is_weak_ranking xs" by (simp add: xs_def weak_ranking_total_preorder) hence f: "f i \ xs ! i" if "i < length xs" for i using that unfolding f_def is_weak_ranking_def by (intro someI_ex[of "\x. x \ xs ! i"]) (auto simp: set_conv_nth) have f': "f i \ carrier" if "i < length xs" for i using that f weak_ranking_Union unfolding xs_def by (auto simp: set_conv_nth) define n where "n = length xs - 1" from assms weak_ranking_Union have carrier_nonempty: "carrier \ {}" and "xs \ []" by (auto simp: xs_def lotteries_on_def set_pmf_not_empty) hence n: "length xs = Suc n" and xs_nonempty: "xs \ []" by (auto simp add: n_def) have SD': "?pref p (f i) \ ?pref q (f i)" if "i < length xs" for i using f'[OF that] SD by (auto simp: SD_preorder preferred_alts_def) have f_le: "le (f i) (f j) \ i \ j" if "i < length xs" "j < length xs" for i j using that weak_ranking_index_unique[OF xs_wf that(1) _ f] weak_ranking_index_unique[OF xs_wf that(2) _ f] by (auto simp add: le intro: f elim!: of_weak_ranking.cases intro!: of_weak_ranking.intros) have "measure_pmf.expectation p u = (\i lotteries_on carrier" for p proof - from p have "measure_pmf.expectation p u = (\i = (\i {..[le] y} - {y. f i \[le] y})" by (subst measure_pmf.finite_measure_Diff [symmetric]) (auto simp: strongly_preferred_def) also have "{y. f i \[le] y} - {y. f i \[le] y} = {y. f i \[le] y \ y \[le] f i}" (is "_ = ?A") by (auto simp: strongly_preferred_def) also have "\ = xs ! i" proof safe fix x assume le: "le (f i) x" "le x (f i)" from i f show "x \ xs ! i" by (intro weak_ranking_eqclass2[OF _ _ le]) (auto simp: xs_def) next fix x assume "x \ xs ! i" from weak_ranking_eqclass1[OF _ this f] weak_ranking_eqclass1[OF _ f this] i show "le x (f i)" "le (f i) x" by (simp_all add: xs_def) qed finally show "u (f i) * measure_pmf.prob p (xs ! i) = u (f i) * (?pref p (f i) - ?pref' p (f i))" by simp qed also have "\ = (\iiiiiiii {..[le] y \ f i \[le] y" for y proof (cases "y \ carrier") assume "y \ carrier" with weak_ranking_Union obtain j where j: "j < length xs" "y \ xs ! j" by (auto simp: set_conv_nth xs_def) with weak_ranking_eqclass1[OF _ f j(2)] weak_ranking_eqclass1[OF _ j(2) f] have iff: "le y z \ le (f j) z" "le z y \ le z (f j)" for z by (auto intro: trans simp: xs_def) thus ?thesis using i j unfolding n_def by (auto simp: iff f_le strongly_preferred_def) qed (insert not_outside, auto simp: strongly_preferred_def) thus "u (f (Suc i)) * ?pref' p (f (Suc i)) = u (f (Suc i)) * ?pref p (f i)" by simp qed also have "?sum - (\ + u (f 0) * ?pref' p (f 0)) = (\i[le] y} = {}" using hd_weak_ranking[of "f 0"] xs_nonempty f not_outside by (auto simp: hd_conv_nth xs_def strongly_preferred_def) also have "{y. le (f n) y} = carrier" using last_weak_ranking[of "f n"] xs_nonempty f not_outside by (auto simp: last_conv_nth xs_def n_def) also from p have "measure_pmf.prob p carrier = 1" by (subst measure_pmf.prob_eq_1) (auto simp: AE_measure_pmf_iff lotteries_on_def) finally show ?thesis by simp qed from assms[THEN this] show "measure_pmf.expectation p u \ measure_pmf.expectation q u" unfolding SD_preorder n_def using f' by (auto intro!: sum_mono mult_left_mono SD' simp: utility_le_iff f_le) next assume "\u. is_vnm_utility u \ measure_pmf.expectation p u \ measure_pmf.expectation q u" hence expected_utility_le: "measure_pmf.expectation p u \ measure_pmf.expectation q u" if "is_vnm_utility u" for u using that by blast define xs where "xs = weak_ranking le" have le: "le = of_weak_ranking xs" and [simp]: "is_weak_ranking xs" by (simp_all add: xs_def weak_ranking_total_preorder) have carrier: "carrier = \(set xs)" by (simp add: xs_def weak_ranking_Union) from assms have carrier_nonempty: "carrier \ {}" by (auto simp: lotteries_on_def set_pmf_not_empty) { fix x assume x: "x \ carrier" let ?idx = "\y. length xs - weak_ranking_index y" have preferred_subset_carrier: "{y. le x y} \ carrier" using not_outside x by auto have "measure_pmf.prob p {y. le x y} / real (length xs) \ measure_pmf.prob q {y. le x y} / real (length xs)" proof (rule field_le_epsilon) fix \ :: real assume \: "\ > 0" define u where "u y = indicator {y. y \[le] x} y + \ * ?idx y" for y have is_utility: "is_vnm_utility u" unfolding u_def xs_def proof (intro vnm_utility.add_left vnm_utility.scaled utility_weak_ranking_index) fix y z assume "le y z" thus "indicator {y. y \[le] x} y \ (indicator {y. y \[le] x} z :: real)" by (auto intro: trans simp: indicator_def split: if_splits) qed fact+ have "(\y|le x y. pmf p y) \ (\y|le x y. pmf p y) + \ * (\y\carrier. ?idx y * pmf p y)" using \ by (auto intro!: mult_nonneg_nonneg sum_nonneg pmf_nonneg) also from expected_utility_le is_utility have "measure_pmf.expectation p u \ measure_pmf.expectation q u" . with assms have "(\a\carrier. u a * pmf p a) \ (\a\carrier. u a * pmf q a)" by (subst (asm) (1 2) integral_measure_pmf[OF finite_carrier]) (auto simp: lotteries_on_def set_pmf_eq ac_simps) hence "(\y|le x y. pmf p y) + \ * (\y\carrier. ?idx y * pmf p y) \ (\y|le x y. pmf q y) + \ * (\y\carrier. ?idx y * pmf q y)" using x preferred_subset_carrier not_outside by (simp add: u_def sum.distrib finite_carrier algebra_simps sum_distrib_left Int_absorb1 cong: rev_conj_cong) also have "(\y\carrier. ?idx y * pmf q y) \ (\y\carrier. length xs * pmf q y)" by (intro sum_mono mult_right_mono) (simp_all add: pmf_nonneg) also have "\ = measure_pmf.expectation q (\_. length xs)" using assms by (subst integral_measure_pmf[OF finite_carrier]) (auto simp: lotteries_on_def set_pmf_eq ac_simps) also have "\ = length xs" by simp also have "(\y | le x y. pmf p y) = measure_pmf.prob p {y. le x y}" using finite_subset[OF preferred_subset_carrier finite_carrier] by (simp add: measure_measure_pmf_finite) also have "(\y | le x y. pmf q y) = measure_pmf.prob q {y. le x y}" using finite_subset[OF preferred_subset_carrier finite_carrier] by (simp add: measure_measure_pmf_finite) finally show "measure_pmf.prob p {y. le x y} / length xs \ measure_pmf.prob q {y. le x y} / length xs + \" using \ by (simp add: divide_simps) qed moreover from carrier_nonempty carrier have "xs \ []" by auto ultimately have "measure_pmf.prob p {y. le x y} \ measure_pmf.prob q {y. le x y}" by (simp add: field_simps) } with assms show "p \[SD(le)] q" unfolding SD_preorder preferred_alts_def by blast qed lemma not_strict_SD_iff: assumes "p \ lotteries_on carrier" "q \ lotteries_on carrier" shows "\(p \[SD(le)] q) \ (\u. is_vnm_utility u \ measure_pmf.expectation q u \ measure_pmf.expectation p u)" proof let ?E = "measure_pmf.expectation :: 'a pmf \ _ \ real" assume "\u. is_vnm_utility u \ ?E p u \ ?E q u" then obtain u where u: "is_vnm_utility u" "?E p u \ ?E q u" by blast interpret u: vnm_utility carrier le u by fact show "\ p \[SD le] q" proof assume less: "p \[SD le] q" with assms have pq: "?E p u \ ?E q u" if "is_vnm_utility u" for u using that by (auto simp: SD_iff_expected_utilities_le strongly_preferred_def) with u have u_eq: "?E p u = ?E q u" by (intro antisym) simp_all from less assms obtain u' where u': "is_vnm_utility u'" "?E p u' < ?E q u'" by (auto simp: SD_iff_expected_utilities_le strongly_preferred_def not_le) interpret u': vnm_utility carrier le u' by fact have "\\>0. is_vnm_utility (\x. u x - \ * u' x)" by (intro u.diff_epsilon antisym u'.utility_le) - then guess \ by (elim exE conjE) note \ = this + then obtain \ where \: "\ > 0" "is_vnm_utility (\x. u x - \ * u' x)" by auto define u'' where "u'' x = u x - \ * u' x" for x interpret u'': vnm_utility carrier le u'' unfolding u''_def by fact have exp_u'': "?E p u'' = ?E p u - \ * ?E p u'" if "p \ lotteries_on carrier" for p using that by (subst (1 2 3) integral_measure_pmf[of carrier]) (auto simp: lotteries_on_def u''_def algebra_simps sum_subtractf sum_distrib_left) from assms \ have "?E p u'' > ?E q u''" by (simp_all add: exp_u'' algebra_simps u_eq u') with pq[OF u''.vnm_utility_axioms] show False by simp qed qed (insert assms utility_weak_ranking_index, auto simp: strongly_preferred_def SD_iff_expected_utilities_le not_le not_less intro: antisym) lemma strict_SD_iff: assumes "p \ lotteries_on carrier" "q \ lotteries_on carrier" shows "(p \[SD(le)] q) \ (\u. is_vnm_utility u \ measure_pmf.expectation p u < measure_pmf.expectation q u)" using not_strict_SD_iff[OF assms] by auto end end diff --git a/thys/Routing/ROOT b/thys/Routing/ROOT --- a/thys/Routing/ROOT +++ b/thys/Routing/ROOT @@ -1,12 +1,14 @@ chapter AFP session Routing (AFP) = Simple_Firewall + options [timeout = 300] + sessions + "Pure-ex" theories Routing_Table Linux_Router IpRoute_Parser Linorder_Helper (*Dependencies: RoutingIpAssmt depends on Iptables_Semantics and is thus in another session*) document_files "root.tex" diff --git a/thys/Routing/Routing_Table.thy b/thys/Routing/Routing_Table.thy --- a/thys/Routing/Routing_Table.thy +++ b/thys/Routing/Routing_Table.thy @@ -1,503 +1,504 @@ section\Routing Table\ theory Routing_Table imports IP_Addresses.Prefix_Match IP_Addresses.IPv4 IP_Addresses.IPv6 Linorder_Helper IP_Addresses.Prefix_Match_toString + "Pure-ex.Guess" begin text\This section makes the necessary definitions to work with a routing table using longest prefix matching.\ subsection\Definition\ record(overloaded) 'i routing_action = output_iface :: string next_hop :: "'i word option" (* no next hop iff locally attached *) (* Routing rule matching ip route unicast type *) record(overloaded) 'i routing_rule = routing_match :: "('i::len) prefix_match" (* done on the dst *) metric :: nat routing_action :: "'i routing_action" text\This definition is engineered to model routing tables on packet forwarding devices. It eludes, e.g., the source address hint, which is only relevant for packets originating from the device itself.\ (* See also: http://linux-ip.net/html/routing-saddr-selection.html *) context begin definition "default_metric = 0" type_synonym 'i prefix_routing = "('i routing_rule) list" abbreviation "routing_oiface a \ output_iface (routing_action a)" (* I needed this a lot... *) abbreviation "routing_prefix r \ pfxm_length (routing_match r)" definition valid_prefixes where "valid_prefixes r = foldr conj (map (\rr. valid_prefix (routing_match rr)) r) True" lemma valid_prefixes_split: "valid_prefixes (r#rs) \ valid_prefix (routing_match r) \ valid_prefixes rs" using valid_prefixes_def by force lemma foldr_True_set: "foldr (\x. (\) (f x)) l True = (\x \ set l. f x)" by (induction l) simp_all lemma valid_prefixes_alt_def: "valid_prefixes r = (\e \ set r. valid_prefix (routing_match e))" unfolding valid_prefixes_def unfolding foldr_map unfolding comp_def unfolding foldr_True_set .. fun has_default_route :: "('i::len) prefix_routing \ bool" where "has_default_route (r#rs) = (((pfxm_length (routing_match r)) = 0) \ has_default_route rs)" | "has_default_route Nil = False" lemma has_default_route_alt: "has_default_route rt \ (\r \ set rt. pfxm_length (routing_match r) = 0)" by(induction rt) simp_all subsection\Single Packet Semantics\ fun routing_table_semantics :: "('i::len) prefix_routing \ 'i word \ 'i routing_action" where "routing_table_semantics [] _ = routing_action (undefined::'i routing_rule)" | "routing_table_semantics (r#rs) p = (if prefix_match_semantics (routing_match r) p then routing_action r else routing_table_semantics rs p)" lemma routing_table_semantics_ports_from_table: "valid_prefixes rtbl \ has_default_route rtbl \ routing_table_semantics rtbl packet = r \ r \ routing_action ` set rtbl" proof(induction rtbl) case (Cons r rs) note v_pfxs = valid_prefixes_split[OF Cons.prems(1)] show ?case proof(cases "pfxm_length (routing_match r) = 0") case True note zero_prefix_match_all[OF conjunct1[OF v_pfxs] True] Cons.prems(3) then show ?thesis by simp next case False hence "has_default_route rs" using Cons.prems(2) by simp from Cons.IH[OF conjunct2[OF v_pfxs] this] Cons.prems(3) show ?thesis by force qed qed simp subsection\Longest Prefix Match\ text\We can abuse @{const LinordHelper} to sort.\ definition "routing_rule_sort_key \ \r. LinordHelper (0 - (of_nat :: nat \ int) (pfxm_length (routing_match r))) (metric r)" text\There is actually a slight design choice here. We can choose to sort based on @{thm less_eq_prefix_match_def} (thus including the address) or only the prefix length (excluding it). Which is taken does not matter gravely, since the bits of the prefix can't matter. They're either eqal or the rules don't overlap and the metric decides. (It does matter for the resulting list though.) Ignoring the prefix and taking only its length is slightly easier.\ (*example: get longest prefix match by sorting by pfxm_length*) definition "rr_ctor m l a nh me \ \ routing_match = PrefixMatch (ipv4addr_of_dotdecimal m) l, metric = me, routing_action =\output_iface = a, next_hop = (map_option ipv4addr_of_dotdecimal nh)\ \" value "sort_key routing_rule_sort_key [ rr_ctor (0,0,0,1) 3 '''' None 0, rr_ctor (0,0,0,2) 8 [] None 0, rr_ctor (0,0,0,3) 4 [] None 13, rr_ctor (0,0,0,3) 4 [] None 42]" definition "is_longest_prefix_routing \ sorted \ map routing_rule_sort_key" definition correct_routing :: "('i::len) prefix_routing \ bool" where "correct_routing r \ is_longest_prefix_routing r \ valid_prefixes r" text\Many proofs and functions around routing require at least parts of @{const correct_routing} as an assumption. Obviously, @{const correct_routing} is not given for arbitrary routing tables. Therefore, @{const correct_routing} is made to be executable and should be checked for any routing table after parsing. Note: @{const correct_routing} used to also require @{const has_default_route}, but none of the proofs require it anymore and it is not given for any routing table.\ lemma is_longest_prefix_routing_rule_exclusion: assumes "is_longest_prefix_routing (r1 # rn # rss)" shows "is_longest_prefix_routing (r1 # rss)" using assms by(case_tac rss) (auto simp add: is_longest_prefix_routing_def) lemma int_of_nat_less: "int_of_nat a < int_of_nat b \ a < b" by (simp add: int_of_nat_def) lemma is_longest_prefix_routing_sorted_by_length: assumes "is_longest_prefix_routing r" and "r = r1 # rs @ r2 # rss" shows "(pfxm_length (routing_match r1) \ pfxm_length (routing_match r2))" using assms proof(induction rs arbitrary: r) case (Cons rn rs) let ?ro = "r1 # rs @ r2 # rss" have "is_longest_prefix_routing ?ro" using Cons.prems is_longest_prefix_routing_rule_exclusion[of r1 rn "rs @ r2 # rss"] by simp from Cons.IH[OF this] show ?case by simp next case Nil thus ?case by(auto simp add: is_longest_prefix_routing_def routing_rule_sort_key_def linord_helper_less_eq1_def less_eq_linord_helper_def int_of_nat_def) qed definition "sort_rtbl :: ('i::len) routing_rule list \ 'i routing_rule list \ sort_key routing_rule_sort_key" lemma is_longest_prefix_routing_sort: "is_longest_prefix_routing (sort_rtbl r)" unfolding sort_rtbl_def is_longest_prefix_routing_def by simp definition "unambiguous_routing rtbl \ (\rt1 rt2 rr ra. rtbl = rt1 @ rr # rt2 \ ra \ set (rt1 @ rt2) \ routing_match rr = routing_match ra \ routing_rule_sort_key rr \ routing_rule_sort_key ra)" lemma unambiguous_routing_Cons: "unambiguous_routing (r # rtbl) \ unambiguous_routing rtbl" unfolding unambiguous_routing_def by(clarsimp) (metis append_Cons in_set_conv_decomp) lemma "unambiguous_routing (rr # rtbl) \ is_longest_prefix_routing (rr # rtbl) \ ra \ set rtbl \ routing_match rr = routing_match ra \ routing_rule_sort_key rr < routing_rule_sort_key ra" unfolding is_longest_prefix_routing_def unambiguous_routing_def by(fastforce) primrec unambiguous_routing_code where "unambiguous_routing_code [] = True" | "unambiguous_routing_code (rr#rtbl) = (list_all (\ra. routing_match rr \ routing_match ra \ routing_rule_sort_key rr \ routing_rule_sort_key ra) rtbl \ unambiguous_routing_code rtbl)" lemma unambiguous_routing_code[code_unfold]: "unambiguous_routing rtbl \ unambiguous_routing_code rtbl" proof(induction rtbl) case (Cons rr rtbl) show ?case (is "?l \ ?r") proof assume l: ?l with unambiguous_routing_Cons Cons.IH have "unambiguous_routing_code rtbl" by blast moreover have "list_all (\ra. routing_match rr \ routing_match ra \ routing_rule_sort_key rr \ routing_rule_sort_key ra) rtbl" using l unfolding unambiguous_routing_def by(fastforce simp add: list_all_iff) ultimately show ?r by simp next assume r: ?r with Cons.IH have "unambiguous_routing rtbl" by simp from r have *: "list_all (\ra. routing_match rr \ routing_match ra \ routing_rule_sort_key rr \ routing_rule_sort_key ra) rtbl" by simp have False if "rr # rtbl = rt1 @ rra # rt2" "ra \ set (rt1 @ rt2)" "routing_rule_sort_key rra = routing_rule_sort_key ra \ routing_match rra = routing_match ra" for rt1 rt2 rra ra proof(cases "rt1 = []") case True thus ?thesis using that * by(fastforce simp add: list_all_iff) next case False with that(1) have rtbl: "rtbl = tl rt1 @ rra # rt2" by (metis list.sel(3) tl_append2) show ?thesis proof(cases "ra = hd rt1") (* meh case split\ *) case False hence "ra \ set (tl rt1 @ rt2)" using that by(cases rt1; simp) with \unambiguous_routing rtbl\ show ?thesis using that(3) rtbl unfolding unambiguous_routing_def by fast next case True hence "rr = ra" using that \rt1 \ []\ by(cases rt1; simp) thus ?thesis using that * unfolding rtbl by(fastforce simp add: list_all_iff) qed qed thus ?l unfolding unambiguous_routing_def by blast qed qed(simp add: unambiguous_routing_def) lemma unambigous_prefix_routing_weak_mono: assumes lpfx: "is_longest_prefix_routing (rr#rtbl)" assumes e:"rr' \ set rtbl" shows "routing_rule_sort_key rr' \ routing_rule_sort_key rr" using assms by(simp add: is_longest_prefix_routing_def) lemma unambigous_prefix_routing_strong_mono: assumes lpfx: "is_longest_prefix_routing (rr#rtbl)" assumes uam: "unambiguous_routing (rr#rtbl)" assumes e:"rr' \ set rtbl" assumes ne: "routing_match rr' = routing_match rr" shows "routing_rule_sort_key rr' > routing_rule_sort_key rr" proof - from uam e ne have "routing_rule_sort_key rr \ routing_rule_sort_key rr'" by(fastforce simp add: unambiguous_routing_def) moreover from unambigous_prefix_routing_weak_mono lpfx e have "routing_rule_sort_key rr \ routing_rule_sort_key rr'" . ultimately show ?thesis by simp qed lemma "routing_rule_sort_key (rr_ctor (0,0,0,0) 8 [] None 0) > routing_rule_sort_key (rr_ctor (0,0,0,0) 24 [] None 0)" by eval (* get the inequality right\ bigger means lower priority *) text\In case you don't like that formulation of @{const is_longest_prefix_routing} over sorting, this is your lemma.\ theorem existential_routing: "valid_prefixes rtbl \ is_longest_prefix_routing rtbl \ has_default_route rtbl \ unambiguous_routing rtbl \ routing_table_semantics rtbl addr = act \ (\rr \ set rtbl. prefix_match_semantics (routing_match rr) addr \ routing_action rr = act \ (\ra \ set rtbl. routing_rule_sort_key ra < routing_rule_sort_key rr \ \prefix_match_semantics (routing_match ra) addr))" proof(induction rtbl) case Nil thus ?case by simp next case (Cons rr rtbl) show ?case proof(cases "prefix_match_semantics (routing_match rr) addr") case False hence [simp]: "routing_table_semantics (rr # rtbl) addr = routing_table_semantics (rr # rtbl) addr" by simp show ?thesis proof(cases "routing_prefix rr = 0") case True text\Need special treatment, rtbl won't have a default route, so the IH is not usable.\ have "valid_prefix (routing_match rr)" using Cons.prems valid_prefixes_split by blast with True False have False using zero_prefix_match_all by blast thus ?thesis .. next case False with Cons.prems have mprems: "valid_prefixes rtbl" "is_longest_prefix_routing rtbl" "has_default_route rtbl" "unambiguous_routing rtbl" by(simp_all add: valid_prefixes_split unambiguous_routing_Cons is_longest_prefix_routing_def) show ?thesis using Cons.IH[OF mprems] False \\ prefix_match_semantics (routing_match rr) addr\ by simp qed next case True from True have [simp]: "routing_table_semantics (rr # rtbl) addr = routing_action rr" by simp show ?thesis (is "?l \ ?r") proof assume ?l hence [simp]: "act = routing_action rr" by(simp add: True) have *: "(\ra\set (rr # rtbl). routing_rule_sort_key rr \ routing_rule_sort_key ra)" using \is_longest_prefix_routing (rr # rtbl)\ by(clarsimp simp: is_longest_prefix_routing_def) thus ?r by(fastforce simp add: True) next assume ?r then guess rr' .. note rr' = this have "rr' = rr" proof(rule ccontr) assume C: "rr' \ rr" from C have e: "rr' \ set rtbl" using rr' by simp show False proof cases assume eq: "routing_match rr' = routing_match rr" with e have "routing_rule_sort_key rr < routing_rule_sort_key rr'" using unambigous_prefix_routing_strong_mono[OF Cons.prems(2,4) _ eq] by simp with True rr' show False by simp next assume ne: "routing_match rr' \ routing_match rr" from rr' Cons.prems have "valid_prefix (routing_match rr)" "valid_prefix (routing_match rr')" "prefix_match_semantics (routing_match rr') addr" by(auto simp add: valid_prefixes_alt_def) note same_length_prefixes_distinct[OF this(1,2) ne[symmetric] _ True this(3)] moreover have "routing_prefix rr = routing_prefix rr'" (is ?pe) proof - have "routing_rule_sort_key rr < routing_rule_sort_key rr' \ \ prefix_match_semantics (routing_match rr) addr" using rr' by simp with unambigous_prefix_routing_weak_mono[OF Cons.prems(2) e] True have "routing_rule_sort_key rr = routing_rule_sort_key rr'" by simp thus ?pe by(simp add: routing_rule_sort_key_def int_of_nat_def) qed ultimately show False . qed qed thus ?l using rr' by simp qed qed qed subsection\Printing\ definition "routing_rule_32_toString (rr::32 routing_rule) \ prefix_match_32_toString (routing_match rr) @ (case next_hop (routing_action rr) of Some nh \ '' via '' @ ipv4addr_toString nh | _ \ []) @ '' dev '' @ routing_oiface rr @ '' metric '' @ string_of_nat (metric rr)" definition "routing_rule_128_toString (rr::128 routing_rule) \ prefix_match_128_toString (routing_match rr) @ (case next_hop (routing_action rr) of Some nh \ '' via '' @ ipv6addr_toString nh | _ \ []) @ '' dev '' @ routing_oiface rr @ '' metric '' @ string_of_nat (metric rr)" lemma "map routing_rule_32_toString [rr_ctor (42,0,0,0) 7 ''eth0'' None 808, rr_ctor (0,0,0,0) 0 ''eth1'' (Some (222,173,190,239)) 707] = [''42.0.0.0/7 dev eth0 metric 808'', ''0.0.0.0/0 via 222.173.190.239 dev eth1 metric 707'']" by eval section\Routing table to Relation\ text\Walking through a routing table splits the (remaining) IP space when traversing a routing table into a pair of sets: the pair contains the IPs concerned by the current rule and those left alone.\ private definition ipset_prefix_match where "ipset_prefix_match pfx rg = (let pfxrg = prefix_to_wordset pfx in (rg \ pfxrg, rg - pfxrg))" private lemma ipset_prefix_match_m[simp]: "fst (ipset_prefix_match pfx rg) = rg \ (prefix_to_wordset pfx)" by(simp only: Let_def ipset_prefix_match_def, simp) private lemma ipset_prefix_match_nm[simp]: "snd (ipset_prefix_match pfx rg) = rg - (prefix_to_wordset pfx)" by(simp only: Let_def ipset_prefix_match_def, simp) private lemma ipset_prefix_match_distinct: "rpm = ipset_prefix_match pfx rg \ (fst rpm) \ (snd rpm) = {}" by force private lemma ipset_prefix_match_complete: "rpm = ipset_prefix_match pfx rg \ (fst rpm) \ (snd rpm) = rg" by force private lemma rpm_m_dup_simp: "rg \ fst (ipset_prefix_match (routing_match r) rg) = fst (ipset_prefix_match (routing_match r) rg)" by simp private definition range_prefix_match :: "'i::len prefix_match \ 'i wordinterval \ 'i wordinterval \ 'i wordinterval" where "range_prefix_match pfx rg \ (let pfxrg = prefix_to_wordinterval pfx in (wordinterval_intersection rg pfxrg, wordinterval_setminus rg pfxrg))" private lemma range_prefix_match_set_eq: "(\(r1,r2). (wordinterval_to_set r1, wordinterval_to_set r2)) (range_prefix_match pfx rg) = ipset_prefix_match pfx (wordinterval_to_set rg)" unfolding range_prefix_match_def ipset_prefix_match_def Let_def using wordinterval_intersection_set_eq wordinterval_setminus_set_eq prefix_to_wordinterval_set_eq by auto private lemma range_prefix_match_sm[simp]: "wordinterval_to_set (fst (range_prefix_match pfx rg)) = fst (ipset_prefix_match pfx (wordinterval_to_set rg))" by (metis fst_conv ipset_prefix_match_m wordinterval_intersection_set_eq prefix_to_wordinterval_set_eq range_prefix_match_def) private lemma range_prefix_match_snm[simp]: "wordinterval_to_set (snd (range_prefix_match pfx rg)) = snd (ipset_prefix_match pfx (wordinterval_to_set rg))" by (metis snd_conv ipset_prefix_match_nm wordinterval_setminus_set_eq prefix_to_wordinterval_set_eq range_prefix_match_def) subsection\Wordintervals for Ports by Routing\ text\This split, although rather trivial, can be used to construct the sets (or rather: the intervals) of IPs that are actually matched by an entry in a routing table.\ private fun routing_port_ranges :: "'i prefix_routing \ 'i wordinterval \ (string \ ('i::len) wordinterval) list" where "routing_port_ranges [] lo = (if wordinterval_empty lo then [] else [(routing_oiface (undefined::'i routing_rule),lo)])" | (* insert default route to nirvana. has to match what routing_table_semantics does. *) "routing_port_ranges (a#as) lo = ( let rpm = range_prefix_match (routing_match a) lo; m = fst rpm; nm = snd rpm in ( (routing_oiface a,m) # routing_port_ranges as nm))" private lemma routing_port_ranges_subsets: "(a1, b1) \ set (routing_port_ranges tbl s) \ wordinterval_to_set b1 \ wordinterval_to_set s" by(induction tbl arbitrary: s; fastforce simp add: Let_def split: if_splits) private lemma routing_port_ranges_sound: "e \ set (routing_port_ranges tbl s) \ k \ wordinterval_to_set (snd e) \ valid_prefixes tbl \ fst e = output_iface (routing_table_semantics tbl k)" proof(induction tbl arbitrary: s) case (Cons a as) note s = Cons.prems(1)[unfolded routing_port_ranges.simps Let_def list.set] note vpfx = valid_prefixes_split[OF Cons.prems(3)] show ?case (is ?kees) proof(cases "e = (routing_oiface a, fst (range_prefix_match (routing_match a) s))") case False hence es: "e \ set (routing_port_ranges as (snd (range_prefix_match (routing_match a) s)))" using s by blast note eq = Cons.IH[OF this Cons.prems(2) conjunct2[OF vpfx]] have "\prefix_match_semantics (routing_match a) k" (is ?nom) proof - from routing_port_ranges_subsets[of "fst e" "snd e", unfolded prod.collapse, OF es] have *: "wordinterval_to_set (snd e) \ wordinterval_to_set (snd (range_prefix_match (routing_match a) s))" . show ?nom unfolding prefix_match_semantics_wordset[OF conjunct1[OF vpfx]] using * Cons.prems(2) unfolding wordinterval_subset_set_eq by(auto simp add: range_prefix_match_def Let_def) qed thus ?kees using eq by simp next case True hence fe: "fst e = routing_oiface a" by simp from True have "k \ wordinterval_to_set (fst (range_prefix_match (routing_match a) s))" using Cons.prems(2) by(simp) hence "prefix_match_semantics (routing_match a) k" unfolding prefix_match_semantics_wordset[OF conjunct1, OF vpfx] unfolding range_prefix_match_def Let_def by simp thus ?kees by(simp add: fe) qed qed (simp split: if_splits) private lemma routing_port_ranges_disjoined: assumes vpfx: "valid_prefixes tbl" and ins: "(a1, b1) \ set (routing_port_ranges tbl s)" "(a2, b2) \ set (routing_port_ranges tbl s)" and nemp: "wordinterval_to_set b1 \ {}" shows "b1 \ b2 \ wordinterval_to_set b1 \ wordinterval_to_set b2 = {}" using assms proof(induction tbl arbitrary: s) case (Cons r rs) have vpfx: "valid_prefix (routing_match r)" "valid_prefixes rs" using Cons.prems(1) using valid_prefixes_split by blast+ { (* In case that one of b1 b2 stems from r and one does not, we assume it is b1 WLOG. *) fix a1 b1 a2 b2 assume one: "b1 = fst (range_prefix_match (routing_match r) s)" assume two: "(a2, b2) \ set (routing_port_ranges rs (snd (range_prefix_match (routing_match r) s)))" have dc: "wordinterval_to_set (snd (range_prefix_match (routing_match r) s)) \ wordinterval_to_set (fst (range_prefix_match (routing_match r) s)) = {}" by force hence "wordinterval_to_set b1 \ wordinterval_to_set b2 = {}" unfolding one using two[THEN routing_port_ranges_subsets] by fast } note * = this show ?case using \(a1, b1) \ set (routing_port_ranges (r # rs) s)\ \(a2, b2) \ set (routing_port_ranges (r # rs) s)\ nemp Cons.IH[OF vpfx(2)] * by(fastforce simp add: Let_def) qed (simp split: if_splits) private lemma routing_port_rangesI: "valid_prefixes tbl \ output_iface (routing_table_semantics tbl k) = output_port \ k \ wordinterval_to_set wi \ (\ip_range. (output_port, ip_range) \ set (routing_port_ranges tbl wi) \ k \ wordinterval_to_set ip_range)" proof(induction tbl arbitrary: wi) case Nil thus ?case by simp blast next case (Cons r rs) from Cons.prems(1) have vpfx: "valid_prefix (routing_match r)" and vpfxs: "valid_prefixes rs" by(simp_all add: valid_prefixes_split) show ?case proof(cases "prefix_match_semantics (routing_match r) k") case True thus ?thesis using Cons.prems(2) using vpfx \k \ wordinterval_to_set wi\ by (intro exI[where x = "fst (range_prefix_match (routing_match r) wi)"]) (simp add: prefix_match_semantics_wordset Let_def) next case False with \k \ wordinterval_to_set wi\ have ksnd: "k \ wordinterval_to_set (snd (range_prefix_match (routing_match r) wi))" by (simp add: prefix_match_semantics_wordset vpfx) have mpr: "output_iface (routing_table_semantics rs k) = output_port" using Cons.prems False by simp note Cons.IH[OF vpfxs mpr ksnd] thus ?thesis by(fastforce simp: Let_def) qed qed subsection\Reduction\ text\So far, one entry in the list would be generated for each routing table entry. This next step reduces it to one for each port. The resulting list will represent a function from port to IP wordinterval. (It can also be understood as a function from IP (interval) to port (where the intervals don't overlap).\ definition "reduce_range_destination l \ let ps = remdups (map fst l) in let c = \s. (wordinterval_Union \ map snd \ filter (((=) s) \ fst)) l in [(p, c p). p \ ps] " definition "routing_ipassmt_wi tbl \ reduce_range_destination (routing_port_ranges tbl wordinterval_UNIV)" lemma routing_ipassmt_wi_distinct: "distinct (map fst (routing_ipassmt_wi tbl))" unfolding routing_ipassmt_wi_def reduce_range_destination_def by(simp add: comp_def) private lemma routing_port_ranges_superseted: "(a1,b1) \ set (routing_port_ranges tbl wordinterval_UNIV) \ \b2. (a1,b2) \ set (routing_ipassmt_wi tbl) \ wordinterval_to_set b1 \ wordinterval_to_set b2" unfolding routing_ipassmt_wi_def reduce_range_destination_def by(force simp add: Set.image_iff wordinterval_Union) private lemma routing_ipassmt_wi_subsetted: "(a1,b1) \ set (routing_ipassmt_wi tbl) \ (a1,b2) \ set (routing_port_ranges tbl wordinterval_UNIV) \ wordinterval_to_set b2 \ wordinterval_to_set b1" unfolding routing_ipassmt_wi_def reduce_range_destination_def by(fastforce simp add: Set.image_iff wordinterval_Union comp_def) text\This lemma should hold without the @{const valid_prefixes} assumption, but that would break the semantic argument and make the proof a lot harder.\ lemma routing_ipassmt_wi_disjoint: assumes vpfx: "valid_prefixes (tbl::('i::len) prefix_routing)" and dif: "a1 \ a2" and ins: "(a1, b1) \ set (routing_ipassmt_wi tbl)" "(a2, b2) \ set (routing_ipassmt_wi tbl)" shows "wordinterval_to_set b1 \ wordinterval_to_set b2 = {}" proof(rule ccontr) note iuf = ins[unfolded routing_ipassmt_wi_def reduce_range_destination_def Let_def, simplified, unfolded Set.image_iff comp_def, simplified] assume "(wordinterval_to_set b1 \ wordinterval_to_set b2 \ {})" hence "wordinterval_to_set b1 \ wordinterval_to_set b2 \ {}" by simp text\If the intervals are not disjoint, there exists a witness of that.\ then obtain x where x[simp]: "x \ wordinterval_to_set b1" "x \ wordinterval_to_set b2" by blast text\This witness has to have come from some entry in the result of @{const routing_port_ranges}, for both of @{term b1} and @{term b2}.\ hence "\b1g. x \ wordinterval_to_set b1g \ wordinterval_to_set b1g \ wordinterval_to_set b1 \ (a1, b1g) \ set (routing_port_ranges tbl wordinterval_UNIV)" using iuf(1) by(fastforce simp add: wordinterval_Union) (* strangely, this doesn't work with obtain *) then obtain b1g where b1g: "x \ wordinterval_to_set b1g" "wordinterval_to_set b1g \ wordinterval_to_set b1" "(a1, b1g) \ set (routing_port_ranges tbl wordinterval_UNIV)" by clarsimp from x have "\b2g. x \ wordinterval_to_set b2g \ wordinterval_to_set b2g \ wordinterval_to_set b2 \ (a2, b2g) \ set (routing_port_ranges tbl wordinterval_UNIV)" using iuf(2) by(fastforce simp add: wordinterval_Union) then obtain b2g where b2g: "x \ wordinterval_to_set b2g" "wordinterval_to_set b2g \ wordinterval_to_set b2" "(a2, b2g) \ set (routing_port_ranges tbl wordinterval_UNIV)" by clarsimp text\Soudness tells us that the both @{term a1} and @{term a2} have to be the result of routing @{term x}.\ note routing_port_ranges_sound[OF b1g(3), unfolded fst_conv snd_conv, OF b1g(1) vpfx] routing_port_ranges_sound[OF b2g(3), unfolded fst_conv snd_conv, OF b2g(1) vpfx] text\A contradiction follows from @{thm dif}.\ with dif show False by simp qed lemma routing_ipassmt_wi_sound: assumes vpfx: "valid_prefixes tbl" and ins: "(ea,eb) \ set (routing_ipassmt_wi tbl)" and x: "k \ wordinterval_to_set eb" shows "ea = output_iface (routing_table_semantics tbl k)" proof - note iuf = ins[unfolded routing_ipassmt_wi_def reduce_range_destination_def Let_def, simplified, unfolded Set.image_iff comp_def, simplified] from x have "\b1g. k \ wordinterval_to_set b1g \ wordinterval_to_set b1g \ wordinterval_to_set eb \ (ea, b1g) \ set (routing_port_ranges tbl wordinterval_UNIV)" using iuf(1) by(fastforce simp add: wordinterval_Union) then obtain b1g where b1g: "k \ wordinterval_to_set b1g" "wordinterval_to_set b1g \ wordinterval_to_set eb" "(ea, b1g) \ set (routing_port_ranges tbl wordinterval_UNIV)" by clarsimp note routing_port_ranges_sound[OF b1g(3), unfolded fst_conv snd_conv, OF b1g(1) vpfx] thus ?thesis . qed theorem routing_ipassmt_wi: assumes vpfx: "valid_prefixes tbl" shows "output_iface (routing_table_semantics tbl k) = output_port \ (\ip_range. k \ wordinterval_to_set ip_range \ (output_port, ip_range) \ set (routing_ipassmt_wi tbl))" proof (intro iffI, goal_cases) case 2 with vpfx routing_ipassmt_wi_sound show ?case by blast next case 1 then obtain ip_range where "(output_port, ip_range) \ set (routing_port_ranges tbl wordinterval_UNIV) \ k \ wordinterval_to_set ip_range" using routing_port_rangesI[where wi = wordinterval_UNIV, OF vpfx] by auto thus ?case unfolding routing_ipassmt_wi_def reduce_range_destination_def unfolding Let_def comp_def by(force simp add: Set.image_iff wordinterval_Union) qed (* this was not given for the old reduced_range_destination *) lemma routing_ipassmt_wi_has_all_interfaces: assumes in_tbl: "r \ set tbl" shows "\s. (routing_oiface r,s) \ set (routing_ipassmt_wi tbl)" proof - from in_tbl have "\s. (routing_oiface r,s) \ set (routing_port_ranges tbl S)" for S proof(induction tbl arbitrary: S) case (Cons l ls) show ?case proof(cases "r = l") case True thus ?thesis using Cons.prems by(auto simp: Let_def) next case False with Cons.prems have "r \ set ls" by simp from Cons.IH[OF this] show ?thesis by(simp add: Let_def) blast qed qed simp thus ?thesis unfolding routing_ipassmt_wi_def reduce_range_destination_def by(force simp add: Set.image_iff) qed end end diff --git a/thys/SDS_Impossibility/SDS_Impossibility.thy b/thys/SDS_Impossibility/SDS_Impossibility.thy --- a/thys/SDS_Impossibility/SDS_Impossibility.thy +++ b/thys/SDS_Impossibility/SDS_Impossibility.thy @@ -1,815 +1,816 @@ (* File: SDS_Impossibility.thy Author: Manuel Eberl The proof that there exists no anonymous and neutral SDS for at least four voters and alternatives that satisfies SD-Efficiency and SD-Strategy-Proofness. *) section \Incompatibility of SD-Efficiency and SD-Strategy-Proofness\ theory SDS_Impossibility imports Randomised_Social_Choice.SDS_Automation Randomised_Social_Choice.Randomised_Social_Choice begin subsection \Preliminary Definitions\ locale sds_impossibility = anonymous_sds agents alts sds + neutral_sds agents alts sds + sd_efficient_sds agents alts sds + strategyproof_sds agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds + assumes agents_ge_4: "card agents \ 4" and alts_ge_4: "card alts \ 4" locale sds_impossibility_4_4 = sds_impossibility agents alts sds for agents :: "'agent set" and alts :: "'alt set" and sds + fixes A1 A2 A3 A4 :: 'agent and a b c d :: 'alt assumes distinct_agents: "distinct [A1, A2, A3, A4]" and distinct_alts: "distinct [a, b, c, d]" and agents: "agents = {A1, A2, A3, A4}" and alts: "alts = {a, b, c, d}" begin lemma an_sds: "an_sds agents alts sds" by unfold_locales lemma ex_post_efficient_sds: "ex_post_efficient_sds agents alts sds" by unfold_locales lemma sd_efficient_sds: "sd_efficient_sds agents alts sds" by unfold_locales lemma strategyproof_an_sds: "strategyproof_an_sds agents alts sds" by unfold_locales lemma distinct_agents' [simp]: "A1 \ A2" "A1 \ A3" "A1 \ A4" "A2 \ A1" "A2 \ A3" "A2 \ A4" "A3 \ A1" "A3 \ A2" "A3 \ A4" "A4 \ A1" "A4 \ A2" "A4 \ A3" using distinct_agents by auto lemma distinct_alts' [simp]: "a \ b" "a \ c" "a \ d" "b \ a" "b \ c" "b \ d" "c \ a" "c \ b" "c \ d" "d \ a" "d \ b" "d \ c" using distinct_alts by auto lemma card_agents [simp]: "card agents = 4" and card_alts [simp]: "card alts = 4" using distinct_agents distinct_alts by (simp_all add: agents alts) lemma in_agents [simp]: "A1 \ agents" "A2 \ agents" "A3 \ agents" "A4 \ agents" by (simp_all add: agents) lemma in_alts [simp]: "a \ alts" "b \ alts" "c \ alts" "d \ alts" by (simp_all add: alts) lemma agent_iff: "x \ agents \ x \ {A1, A2, A3, A4}" "(\x\agents. P x) \ P A1 \ P A2 \ P A3 \ P A4" "(\x\agents. P x) \ P A1 \ P A2 \ P A3 \ P A4" by (auto simp add: agents) lemma alt_iff: "x \ alts \ x \ {a,b,c,d}" "(\x\alts. P x) \ P a \ P b \ P c \ P d" "(\x\alts. P x) \ P a \ P b \ P c \ P d" by (auto simp add: alts) subsection \Definition of Preference Profiles and Fact Gathering\ preference_profile agents: agents alts: alts where R1 = A1: [c, d], [a, b] A2: [b, d], a, c A3: a, b, [c, d] A4: [a, c], [b, d] and R2 = A1: [a, c], [b, d] A2: [c, d], a, b A3: [b, d], a, c A4: a, b, [c, d] and R3 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: d, [a, b], c A4: c, a, [b, d] and R4 = A1: [a, b], [c, d] A2: [a, d], [b, c] A3: c, [a, b], d A4: d, c, [a, b] and R5 = A1: [c, d], [a, b] A2: [a, b], [c, d] A3: [a, c], d, b A4: d, [a, b], c and R6 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: [a, c], [b, d] A4: d, b, a, c and R7 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: a, c, d, b A4: d, [a, b], c and R8 = A1: [a, b], [c, d] A2: [a, c], [b, d] A3: d, [a, b], c A4: d, c, [a, b] and R9 = A1: [a, b], [c, d] A2: [a, d], c, b A3: d, c, [a, b] A4: [a, b, c], d and R10 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: [a, c], d, b A4: [b, d], a, c and R11 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: d, [a, b], c A4: c, a, b, d and R12 = A1: [c, d], [a, b] A2: [a, b], [c, d] A3: [a, c], d, b A4: [a, b, d], c and R13 = A1: [a, c], [b, d] A2: [c, d], a, b A3: [b, d], a, c A4: a, b, d, c and R14 = A1: [a, b], [c, d] A2: d, c, [a, b] A3: [a, b, c], d A4: a, d, c, b and R15 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: [b, d], a, c A4: a, c, d, b and R16 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: a, c, d, b A4: [a, b, d], c and R17 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: [a, c], [b, d] A4: d, [a, b], c and R18 = A1: [a, b], [c, d] A2: [a, d], [b, c] A3: [a, b, c], d A4: d, c, [a, b] and R19 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: [b, d], a, c A4: [a, c], [b, d] and R20 = A1: [b, d], a, c A2: b, a, [c, d] A3: a, c, [b, d] A4: d, c, [a, b] and R21 = A1: [a, d], c, b A2: d, c, [a, b] A3: c, [a, b], d A4: a, b, [c, d] and R22 = A1: [a, c], d, b A2: d, c, [a, b] A3: d, [a, b], c A4: a, b, [c, d] and R23 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: [a, c], [b, d] A4: [a, b, d], c and R24 = A1: [c, d], [a, b] A2: d, b, a, c A3: c, a, [b, d] A4: b, a, [c, d] and R25 = A1: [c, d], [a, b] A2: [b, d], a, c A3: a, b, [c, d] A4: a, c, [b, d] and R26 = A1: [b, d], [a, c] A2: [c, d], [a, b] A3: a, b, [c, d] A4: a, c, [b, d] and R27 = A1: [a, b], [c, d] A2: [b, d], a, c A3: [a, c], [b, d] A4: [c, d], a, b and R28 = A1: [c, d], a, b A2: [b, d], a, c A3: a, b, [c, d] A4: a, c, [b, d] and R29 = A1: [a, c], d, b A2: [b, d], a, c A3: a, b, [c, d] A4: d, c, [a, b] and R30 = A1: [a, d], c, b A2: d, c, [a, b] A3: c, [a, b], d A4: [a, b], d, c and R31 = A1: [b, d], a, c A2: [a, c], d, b A3: c, d, [a, b] A4: [a, b], c, d and R32 = A1: [a, c], d, b A2: d, c, [a, b] A3: d, [a, b], c A4: [a, b], d, c and R33 = A1: [c, d], [a, b] A2: [a, c], d, b A3: a, b, [c, d] A4: d, [a, b], c and R34 = A1: [a, b], [c, d] A2: a, c, d, b A3: b, [a, d], c A4: c, d, [a, b] and R35 = A1: [a, d], c, b A2: a, b, [c, d] A3: [a, b, c], d A4: d, c, [a, b] and R36 = A1: [c, d], [a, b] A2: [a, c], d, b A3: [b, d], a, c A4: a, b, [c, d] and R37 = A1: [a, c], [b, d] A2: [b, d], [a, c] A3: a, b, [c, d] A4: c, d, [a, b] and R38 = A1: [c, d], a, b A2: [b, d], a, c A3: a, b, [c, d] A4: [a, c], b, d and R39 = A1: [a, c], d, b A2: [b, d], a, c A3: a, b, [c, d] A4: [c, d], a, b and R40 = A1: [a, d], c, b A2: [a, b], c, d A3: [a, b, c], d A4: d, c, [a, b] and R41 = A1: [a, d], c, b A2: [a, b], d, c A3: [a, b, c], d A4: d, c, [a, b] and R42 = A1: [c, d], [a, b] A2: [a, b], [c, d] A3: d, b, a, c A4: c, a, [b, d] and R43 = A1: [a, b], [c, d] A2: [c, d], [a, b] A3: d, [a, b], c A4: a, [c, d], b and R44 = A1: [c, d], [a, b] A2: [a, c], d, b A3: [a, b], d, c A4: [a, b, d], c and R45 = A1: [a, c], d, b A2: [b, d], a, c A3: [a, b], c, d A4: [c, d], b, a and R46 = A1: [b, d], a, c A2: d, c, [a, b] A3: [a, c], [b, d] A4: b, a, [c, d] and R47 = A1: [a, b], [c, d] A2: [a, d], c, b A3: d, c, [a, b] A4: c, [a, b], d by (simp_all add: agents alts) derive_orbit_equations (an_sds) R10 R26 R27 R28 R29 R43 R45 by simp_all prove_inefficient_supports (ex_post_efficient_sds sd_efficient_sds) R3 [b] and R4 [b] and R5 [b] and R7 [b] and R8 [b] and R9 [b] and R11 [b] and R12 [b] and R14 [b] and R16 [b] and R17 [b] and R18 [b] and R21 [b] and R22 [b] and R23 [b] and R30 [b] and R32 [b] and R33 [b] and R35 [b] and R40 [b] and R41 [b] and R43 [b] and R44 [b] and R47 [b] and R10 [c, b] witness: [a: 1 / 2, b: 0, c: 0, d: 1 / 2] and R15 [c, b] witness: [a: 1 / 2, b: 0, c: 0, d: 1 / 2] and R19 [c, b] witness: [a: 1 / 2, b: 0, c: 0, d: 1 / 2] and R25 [b, c] witness: [c: 0, d: 1 / 2, a: 1 / 2, b: 0] and R26 [c, b] witness: [b: 0, d: 1 / 2, a: 1 / 2, c: 0] and R27 [c, b] witness: [a: 1 / 2, b: 0, c: 0, d: 1 / 2] and R28 [b, c] witness: [c: 0, d: 1 / 2, a: 1 / 2, b: 0] and R29 [b, c] witness: [a: 1 / 2, c: 0, d: 1 / 2, b: 0] and R39 [b, c] witness: [a: 1 / 2, c: 0, d: 1 / 2, b: 0] by (simp_all add: agent_iff alt_iff) derive_strategyproofness_conditions (strategyproof_an_sds) distance: 2 R1 R2 R3 R4 R5 R6 R7 R8 R9 R10 R11 R12 R13 R14 R15 R16 R17 R18 R19 R20 R21 R22 R23 R24 R25 R26 R27 R28 R29 R30 R31 R32 R33 R34 R35 R36 R37 R38 R39 R40 R41 R42 R43 R44 R45 R46 R47 by (simp_all add: agent_iff alt_iff) lemma lottery_conditions: assumes "is_pref_profile R" shows "pmf (sds R) a \ 0" "pmf (sds R) b \ 0" "pmf (sds R) c \ 0" "pmf (sds R) d \ 0" "pmf (sds R) a + pmf (sds R) b + pmf (sds R) c + pmf (sds R) d = 1" using lottery_prob_alts[OF sds_wf[OF assms]] by (simp_all add: alts pmf_nonneg measure_measure_pmf_finite) subsection \Main Proof\ lemma R45 [simp]: "pmf (sds R45) a = 1/4" "pmf (sds R45) b = 1/4" "pmf (sds R45) c = 1/4" "pmf (sds R45) d = 1/4" using R45.orbits lottery_conditions[OF R45.wf] by simp_all lemma R10_bc [simp]: "pmf (sds R10) b = 0" "pmf (sds R10) c = 0" using R10.support R10.orbits by auto lemma R10_ad [simp]: "pmf (sds R10) a = 1/2" "pmf (sds R10) d = 1/2" using lottery_conditions[OF R10.wf] R10_bc R10.orbits by simp_all lemma R26_bc [simp]: "pmf (sds R26) b = 0" "pmf (sds R26) c = 0" using R26.support R26.orbits by auto lemma R26_d [simp]: "pmf (sds R26) d = 1 - pmf (sds R26) a" using lottery_conditions[OF R26.wf] R26_bc by simp lemma R27_bc [simp]: "pmf (sds R27) b = 0" "pmf (sds R27) c = 0" using R27.support R27.orbits by auto lemma R27_d [simp]: "pmf (sds R27) d = 1 - pmf (sds R27) a" using lottery_conditions[OF R27.wf] R27_bc by simp lemma R28_bc [simp]: "pmf (sds R28) b = 0" "pmf (sds R28) c = 0" using R28.support R28.orbits by auto lemma R28_d [simp]: "pmf (sds R28) d = 1 - pmf (sds R28) a" using lottery_conditions[OF R28.wf] R28_bc by simp lemma R29_bc [simp]: "pmf (sds R29) b = 0" "pmf (sds R29) c = 0" using R29.support R29.orbits by auto lemma R29_ac [simp]: "pmf (sds R29) a = 1/2" "pmf (sds R29) d = 1/2" using lottery_conditions[OF R29.wf] R29_bc R29.orbits by simp_all lemmas R43_bc [simp] = R43.support lemma R43_ad [simp]: "pmf (sds R43) a = 1/2" "pmf (sds R43) d = 1/2" using lottery_conditions[OF R43.wf] R43_bc R43.orbits by simp_all lemma R39_b [simp]: "pmf (sds R39) b = 0" proof - { assume [simp]: "pmf (sds R39) c = 0" with R29_R39.strategyproofness(1) have "pmf (sds R39) d \ 1/2" by auto with R39_R29.strategyproofness(1) lottery_conditions[OF R39.wf] have "pmf (sds R39) b = 0" by auto } with R39.support show ?thesis by blast qed lemma R36_a [simp]: "pmf (sds R36) a = 1/2" and R36_b [simp]: "pmf (sds R36) b = 0" proof - from R10_R36.strategyproofness(1) lottery_conditions[OF R36.wf] have "pmf (sds R36) a + pmf (sds R36) b \ 1/2" by auto with R36_R10.strategyproofness(1) lottery_conditions[OF R36.wf] show "pmf (sds R36) a = 1/2" "pmf (sds R36) b = 0" by auto qed lemma R36_d [simp]: "pmf (sds R36) d = 1/2 - pmf (sds R36) c" using lottery_conditions[OF R36.wf] by simp lemma R39_a [simp]: "pmf (sds R39) a = 1/2" proof - from R36_R39.strategyproofness(1) lottery_conditions[OF R39.wf] have "pmf (sds R39) a \ 1/2" by auto with R39_R36.strategyproofness(1) lottery_conditions[OF R39.wf] show ?thesis by auto qed lemma R39_d [simp]: "pmf (sds R39) d = 1/2 - pmf (sds R39) c" using lottery_conditions[OF R39.wf] by simp lemmas R12_b [simp] = R12.support lemma R12_c [simp]: "pmf (sds R12) c = 0" using R12_R10.strategyproofness(1) lottery_conditions[OF R12.wf] by (auto simp del: pmf_nonneg) lemma R12_d [simp]: "pmf (sds R12) d = 1 - pmf (sds R12) a" using lottery_conditions[OF R12.wf] by simp lemma R12_a_ge_one_half: "pmf (sds R12) a \ 1/2" using R10_R12.strategyproofness(1) lottery_conditions[OF R12.wf] by auto lemma R44 [simp]: "pmf (sds R44) a = pmf (sds R12) a" "pmf (sds R44) d = 1 - pmf (sds R12) a" "pmf (sds R44) b = 0" "pmf (sds R44) c = 0" proof - from R12_R44.strategyproofness(1) R44.support have "pmf (sds R44) a \ pmf (sds R12) a" by simp with R44_R12.strategyproofness(1) R44.support lottery_conditions[OF R44.wf] show "pmf (sds R44) a = pmf (sds R12) a" "pmf (sds R44) c = 0" "pmf (sds R44) d = 1 - pmf (sds R12) a" by (auto simp del: pmf_nonneg) qed (insert R44.support, simp_all) lemma R9_a [simp]: "pmf (sds R9) a = pmf (sds R35) a" proof - from R9_R35.strategyproofness(1) R35.support R9.support have "pmf (sds R35) a \ pmf (sds R9) a" by simp with R35_R9.strategyproofness(1) R9.support R35.support show ?thesis by simp qed lemma R18_c [simp]: "pmf (sds R18) c = pmf (sds R9) c" proof - from R18_R9.strategyproofness(1) R18.support R9.support have "pmf (sds R18) d + pmf (sds R18) a \ pmf (sds R9) d + pmf (sds R9) a" by auto with R9_R18.strategyproofness(1) R18.support R9.support lottery_conditions[OF R9.wf] lottery_conditions[OF R18.wf] show ?thesis by auto qed lemma R5_d_ge_one_half: "pmf (sds R5) d \ 1/2" using R5_R10.strategyproofness(1) R5.support lottery_conditions[OF R5.wf] by auto lemma R7 [simp]: "pmf (sds R7) a = 1/2" "pmf (sds R7) b = 0" "pmf (sds R7) c = 0" "pmf (sds R7) d = 1/2" proof - from R5_d_ge_one_half have "1/2 \ pmf (sds R5) d" by simp also from R5_R17.strategyproofness(1) R17.support lottery_conditions[OF R5.wf] lottery_conditions[OF R17.wf] have "\ \ pmf (sds R17) d" by (auto simp del: pmf_nonneg) also from R17_R7.strategyproofness(1) lottery_conditions[OF R7.wf] lottery_conditions[OF R17.wf] R7.support have "pmf (sds R17) d \ pmf (sds R7) d" by (auto simp del: pmf_nonneg) finally have "pmf (sds R7) d \ 1/2" . with R7_R43.strategyproofness(1) lottery_conditions[OF R7.wf] R7.support show "pmf (sds R7) a = 1/2" "pmf (sds R7) b = 0" "pmf (sds R7) c = 0" "pmf (sds R7) d = 1/2" by auto qed lemma R5 [simp]: "pmf (sds R5) a = 1/2" "pmf (sds R5) b = 0" "pmf (sds R5) c = 0" "pmf (sds R5) d = 1/2" proof - from R5_R7.strategyproofness(1) lottery_conditions[OF R5.wf] R5.support have "pmf (sds R5) d \ 1/2" by auto with R5_d_ge_one_half show d: "pmf (sds R5) d = 1 / 2" by simp with R5_R10.strategyproofness(1) lottery_conditions[OF R5.wf] R5.support show "pmf (sds R5) c = 0" "pmf (sds R5) a = 1/2" by simp_all qed (simp_all add: R5.support) lemma R15 [simp]: "pmf (sds R15) a = 1/2" "pmf (sds R15) b = 0" "pmf (sds R15) c = 0" "pmf (sds R15) d = 1/2" proof - { assume "pmf (sds R15) b = 0" with R10_R15.strategyproofness(1) lottery_conditions[OF R15.wf] have "pmf (sds R15) a + pmf (sds R15) c \ 1/2" by auto with R15_R10.strategyproofness(1) lottery_conditions[OF R15.wf] have "pmf (sds R15) c = 0" by auto } with R15.support show [simp]: "pmf (sds R15) c = 0" by blast with R15_R5.strategyproofness(1) lottery_conditions[OF R15.wf] have "pmf (sds R15) a \ 1/2" by auto moreover from R15_R7.strategyproofness(1) lottery_conditions[OF R15.wf] have "pmf (sds R15) b + pmf (sds R15) d \ 1/2" by auto ultimately show "pmf (sds R15) a = 1/2" using lottery_conditions[OF R15.wf] by auto with R15_R5.strategyproofness(1) lottery_conditions[OF R15.wf] show "pmf (sds R15) d = 1/2" "pmf (sds R15) b = 0" by auto qed lemma R13_aux: "pmf (sds R13) b = 0" "pmf (sds R13) c = 0" "pmf (sds R13) d = 1 - pmf (sds R13) a" and R27_R13 [simp]: "pmf (sds R27) a = pmf (sds R13) a" using R27_R13.strategyproofness(1) R13_R27.strategyproofness(1) lottery_conditions[OF R13.wf] by auto lemma R13 [simp]: "pmf (sds R13) a = 1/2" "pmf (sds R13) b = 0" "pmf (sds R13) c = 0" "pmf (sds R13) d = 1/2" using R15_R13.strategyproofness(1) R13_R15.strategyproofness(1) R13_aux by simp_all lemma R27 [simp]: "pmf (sds R27) a = 1/2" "pmf (sds R27) b = 0" "pmf (sds R27) c = 0" "pmf (sds R27) d = 1/2" by simp_all lemma R19 [simp]: "pmf (sds R19) a = 1/2" "pmf (sds R19) b = 0" "pmf (sds R19) c = 0" "pmf (sds R19) d = 1/2" proof - have "pmf (sds R19) a = 1/2 \ pmf (sds R19) b = 0 \ pmf (sds R19) c = 0 \ pmf (sds R19) d = 1/2" proof (rule disjE[OF R19.support]; safe) assume [simp]: "pmf (sds R19) b = 0" from R10_R19.strategyproofness(1) lottery_conditions[OF R19.wf] have "pmf (sds R19) a + pmf (sds R19) c \ 1/2" by auto moreover from R19_R10.strategyproofness(1) have "pmf (sds R19) a + pmf (sds R19) c \ 1/2" by simp ultimately show "pmf (sds R19) d = 1/2" using lottery_conditions[OF R19.wf] by simp with R27_R19.strategyproofness(1) lottery_conditions[OF R19.wf] show "pmf (sds R19) a = 1/2" "pmf (sds R19) c = 0" by auto next assume [simp]: "pmf (sds R19) c = 0" from R19_R10.strategyproofness(1) have "pmf (sds R19) a \ 1/2" by auto moreover from R19_R27.strategyproofness(1) have "pmf (sds R19) d \ 1/2" by auto ultimately show "pmf (sds R19) a = 1/2" "pmf (sds R19) d = 1/2" "pmf (sds R19) b = 0" using lottery_conditions[OF R19.wf] by (auto simp del: pmf_nonneg) qed thus "pmf (sds R19) a = 1/2" "pmf (sds R19) b = 0" "pmf (sds R19) c = 0" "pmf (sds R19) d = 1/2" by blast+ qed lemma R1 [simp]: "pmf (sds R1) a = 1/2" "pmf (sds R1) b = 0" proof - from R19_R1.strategyproofness(1) lottery_conditions[OF R1.wf] have "pmf (sds R1) a + pmf (sds R1) b \ 1/2" by simp with R1_R19.strategyproofness(1) lottery_conditions[OF R1.wf] show "pmf (sds R1) a = 1/2" "pmf (sds R1) b = 0" by auto qed lemma R22 [simp]: "pmf (sds R22) a = 1/2" "pmf (sds R22) b = 0" "pmf (sds R22) c = 0" "pmf (sds R22) d = 1/2" proof - from R33_R5.strategyproofness(1) R33.support have "1/2 \ pmf (sds R33) a" by auto also from R33_R22.strategyproofness(1) R22.support R33.support lottery_conditions[OF R22.wf] lottery_conditions[OF R33.wf] have "\ \ pmf (sds R22) a" by simp finally show "pmf (sds R22) a = 1/2" "pmf (sds R22) b = 0" "pmf (sds R22) c = 0" "pmf (sds R22) d = 1/2" using R22_R29.strategyproofness(1) lottery_conditions[OF R22.wf] by (auto simp del: pmf_nonneg) qed lemma R28 [simp]: "pmf (sds R28) a = 1/2" "pmf (sds R28) b = 0" "pmf (sds R28) c = 0" "pmf (sds R28) d = 1/2" proof - have "pmf (sds R28) a \ pmf (sds R32) d" using R32_R28.strategyproofness(1) lottery_conditions[OF R32.wf] by auto hence R32_d: "pmf (sds R32) d = pmf (sds R28) a" using R28_R32.strategyproofness(1) lottery_conditions[OF R32.wf] by auto from R22_R32.strategyproofness(1) lottery_conditions[OF R32.wf] R32.support have "pmf (sds R32) a \ 1/2" by auto with R32_R22.strategyproofness(1) lottery_conditions[OF R32.wf] R32.support show "pmf (sds R28) a = 1/2" "pmf (sds R28) b = 0" "pmf (sds R28) c = 0" "pmf (sds R28) d = 1/2" by (auto simp: R32_d simp del: pmf_nonneg) qed lemma R39 [simp]: "pmf (sds R39) a = 1/2" "pmf (sds R39) b = 0" "pmf (sds R39) c = 0" "pmf (sds R39) d = 1/2" proof - from R28_R39.strategyproofness(1) show "pmf (sds R39) c = 0" by simp thus "pmf (sds R39) a = 1/2" "pmf (sds R39) b = 0" "pmf (sds R39) d = 1/2" by simp_all qed lemma R2 [simp]: "pmf (sds R2) a = 1/2" "pmf (sds R2) b = 0" "pmf (sds R2) c = 0" "pmf (sds R2) d = 1/2" proof - from R1_R2.strategyproofness(1) R2_R1.strategyproofness(1) lottery_conditions[OF R2.wf] lottery_conditions[OF R1.wf] have "pmf (sds R2) a = 1/2" "pmf (sds R2) c + pmf (sds R2) d = 1/2" by (auto simp: algebra_simps simp del: pmf_nonneg) with R39_R2.strategyproofness(1) lottery_conditions[OF R2.wf] show "pmf (sds R2) a = 1/2" "pmf (sds R2) b = 0" "pmf (sds R2) c = 0" "pmf (sds R2) d = 1/2" by auto qed lemma R42 [simp]: "pmf (sds R42) a = 0" "pmf (sds R42) b = 0" "pmf (sds R42) c = 1/2" "pmf (sds R42) d = 1/2" proof - from R17_R5.strategyproofness(1) lottery_conditions[OF R17.wf] R17.support have "pmf (sds R17) d \ 1/2" by auto moreover from R5_R17.strategyproofness(1) R17.support lottery_conditions[OF R17.wf] have "pmf (sds R17) d \ 1/2" by auto ultimately have R17_d: "pmf (sds R17) d = 1/2" by simp from R6_R42.strategyproofness(1) have "pmf (sds R42) a + pmf (sds R42) c \ pmf (sds R6) a + pmf (sds R6) c" by simp also from R6_R19.strategyproofness(1) lottery_conditions[OF R6.wf] have "pmf (sds R6) a + pmf (sds R6) c \ 1/2" by (auto simp del: pmf_nonneg) finally have "pmf (sds R42) a + pmf (sds R42) c \ 1 / 2" . moreover from R17_R11.strategyproofness(1) R11.support R17.support lottery_conditions[OF R11.wf] lottery_conditions[OF R17.wf] have "pmf (sds R11) d \ 1/2" by (auto simp: R17_d) ultimately have "pmf (sds R42) a + pmf (sds R42) c \ pmf (sds R11) d" by simp with R42_R11.strategyproofness(1) R11.support have E: "pmf (sds R11) d \ pmf (sds R42) c" by auto with \pmf (sds R11) d \ 1/2\ have "pmf (sds R42) c \ 1/2" by simp moreover from R17_R3.strategyproofness(1) R3.support R17.support lottery_conditions[OF R17.wf] lottery_conditions[OF R3.wf] have "pmf (sds R3) d \ 1/2" by (auto simp: R17_d) ultimately show "pmf (sds R42) a = 0" "pmf (sds R42) b = 0" "pmf (sds R42) c = 1/2" "pmf (sds R42) d = 1/2" using R42_R3.strategyproofness(1) lottery_conditions[OF R3.wf] lottery_conditions[OF R42.wf] by linarith+ qed lemma R37 [simp]: "pmf (sds R37) a = 1/2" "pmf (sds R37) b = 0" "pmf (sds R37) c = 1/2" "pmf (sds R37) d = 0" proof - from R37_R42.strategyproofness(1) lottery_conditions[OF R37.wf] have "pmf (sds R37) a = 1/2 \ pmf (sds R37) a + pmf (sds R37) b > 1/2" by (auto simp del: pmf_nonneg) moreover from R37_R42.strategyproofness(2) lottery_conditions[OF R37.wf] have "pmf (sds R37) c = 1/2 \ pmf (sds R37) c + pmf (sds R37) d > 1/2" by (auto simp del: pmf_nonneg) ultimately show "pmf (sds R37) a = 1/2" "pmf (sds R37) b = 0" "pmf (sds R37) c = 1/2" "pmf (sds R37) d = 0" using lottery_conditions[OF R37.wf] by (auto simp del: pmf_nonneg) qed lemma R24 [simp]: "pmf (sds R24) a = 0" "pmf (sds R24) b = 0" "pmf (sds R24) d = 1 - pmf (sds R24) c" using R42_R24.strategyproofness(1) lottery_conditions[OF R24.wf] by (auto simp del: pmf_nonneg) lemma R34 [simp]: "pmf (sds R34) a = 1 - pmf (sds R24) c" "pmf (sds R34) b = pmf (sds R24) c" "pmf (sds R34) c = 0" "pmf (sds R34) d = 0" proof - from R24_R34.strategyproofness(1) lottery_conditions[OF R34.wf] have "pmf (sds R34) b \ pmf (sds R24) c" by (auto simp del: pmf_nonneg) moreover from R34_R24.strategyproofness(1) lottery_conditions[OF R34.wf] have "pmf (sds R34) b \ pmf (sds R24) c" by auto ultimately show bc: "pmf (sds R34) b = pmf (sds R24) c" by simp from R34_R24.strategyproofness(1) bc lottery_conditions[OF R34.wf] show "pmf (sds R34) c = 0" by auto moreover from R24_R34.strategyproofness(1) bc show "pmf (sds R34) d = 0" by simp ultimately show "pmf (sds R34) a = 1 - pmf (sds R24) c" using bc lottery_conditions[OF R34.wf] by auto qed lemma R14 [simp]: "pmf (sds R14) b = 0" "pmf (sds R14) d = 0" "pmf (sds R14) c = 1 - pmf (sds R14) a" using R14_R34.strategyproofness(1) R14.support lottery_conditions[OF R14.wf] by (auto simp del: pmf_nonneg) lemma R46 [simp]: "pmf (sds R46) a = 0" "pmf (sds R46) c = 0" "pmf (sds R46) d = 1 - pmf (sds R46) b" using R46_R37.strategyproofness(1) lottery_conditions[OF R46.wf] by auto lemma R20 [simp]: "pmf (sds R20) a = 0" "pmf (sds R20) c = 0" "pmf (sds R20) d = 1 - pmf (sds R20) b" using R46_R20.strategyproofness(1) lottery_conditions[OF R20.wf] by (auto simp del: pmf_nonneg) lemma R21 [simp]: "pmf (sds R21) d = 1 - pmf (sds R21) a" "pmf (sds R21) b = 0" "pmf (sds R21) c = 0" using R20_R21.strategyproofness(1) lottery_conditions[OF R21.wf] by auto lemma R16_R12: "pmf (sds R16) c + pmf (sds R16) a \ pmf (sds R12) a" using R12_R16.strategyproofness(1) R16.support lottery_conditions[OF R16.wf] by auto lemma R16 [simp]: "pmf (sds R16) b = 0" "pmf (sds R16) c = 0" "pmf (sds R16) d = 1 - pmf (sds R16) a" proof - from R16_R12 have "pmf (sds R16) c + pmf (sds R16) a \ pmf (sds R12) a" by simp also from R44_R40.strategyproofness(1) lottery_conditions[OF R40.wf] R40.support have "pmf (sds R12) a \ pmf (sds R40) a" by auto also from R9_R40.strategyproofness(1) R9.support R40.support have "pmf (sds R40) a \ pmf (sds R9) a" by auto finally have "pmf (sds R16) c + pmf (sds R16) a \ pmf (sds R9) a" by simp moreover from R14_R16.strategyproofness(1) R16.support lottery_conditions[OF R16.wf] have "pmf (sds R16) a \ pmf (sds R14) a" by auto ultimately have "pmf (sds R16) c \ pmf (sds R9) a - pmf (sds R14) a" by simp also from R14_R9.strategyproofness(1) R9.support lottery_conditions[OF R9.wf] have "pmf (sds R9) a - pmf (sds R14) a \ 0" by (auto simp del: pmf_nonneg) finally show "pmf (sds R16) b = 0" "pmf (sds R16) c = 0" "pmf (sds R16) d = 1 - pmf (sds R16) a" using lottery_conditions[OF R16.wf] R16.support by auto qed lemma R12_R14: "pmf (sds R14) a \ pmf (sds R12) a" using R14_R16.strategyproofness(1) R16_R12 by auto lemma R12_a [simp]: "pmf (sds R12) a = pmf (sds R9) a" proof - from R44_R40.strategyproofness(1) R40.support lottery_conditions[OF R40.wf] have "pmf (sds R12) a \ pmf (sds R40) a" by auto also from R9_R40.strategyproofness(1) R9.support R40.support have "pmf (sds R40) a \ pmf (sds R9) a" by auto finally have B: "pmf (sds R12) a \ pmf (sds R9) a" by simp moreover from R14_R9.strategyproofness(1) lottery_conditions[OF R9.wf] R9.support have "pmf (sds R9) a \ pmf (sds R14) a" by (auto simp del: pmf_nonneg) with R12_R14 have "pmf (sds R9) a \ pmf (sds R12) a" by simp ultimately show "pmf (sds R12) a = pmf (sds R9) a" by simp qed lemma R9 [simp]: "pmf (sds R9) b = 0" "pmf (sds R9) d = 0" "pmf (sds R14) a = pmf (sds R35) a" "pmf (sds R9) c = 1 - pmf (sds R35) a" using R12_R14 R14_R9.strategyproofness(1) lottery_conditions[OF R9.wf] R9.support by auto lemma R23 [simp]: "pmf (sds R23) b = 0" "pmf (sds R23) c = 0" "pmf (sds R23) d = 1 - pmf (sds R23) a" using R23_R19.strategyproofness(1) lottery_conditions[OF R23.wf] R23.support by (auto simp del: pmf_nonneg) lemma R35 [simp]: "pmf (sds R35) a = pmf (sds R21) a" "pmf (sds R35) b = 0" "pmf (sds R35) c = 0" "pmf (sds R35) d = 1 - pmf (sds R21) a" proof - from R35_R21.strategyproofness(1) R35.support have "pmf (sds R21) a \ pmf (sds R35) a + pmf (sds R35) c" by auto with R21_R35.strategyproofness(1) R35.support lottery_conditions[OF R35.wf] show "pmf (sds R35) a = pmf (sds R21) a" "pmf (sds R35) b = 0" "pmf (sds R35) c = 0" "pmf (sds R35) d = 1 - pmf (sds R21) a" by simp_all qed lemma R18 [simp]: "pmf (sds R18) a = pmf (sds R14) a" "pmf (sds R18) b = 0" "pmf (sds R18) d = 0" "pmf (sds R18) c = 1 - pmf (sds R14) a" proof - from R23_R12.strategyproofness(1) have R21_R23: "pmf (sds R21) a \ pmf (sds R23) a" by simp from R23_R18.strategyproofness(1) have "pmf (sds R18) d \ pmf (sds R21) a - pmf (sds R23) a" by simp also from R21_R23 have "\ \ 0" by simp finally show "pmf (sds R18) d = 0" by simp with lottery_conditions[OF R18.wf] R18.support show "pmf (sds R18) a = pmf (sds R14) a" "pmf (sds R18) c = 1 - pmf (sds R14) a" by auto qed (insert R18.support, simp_all) lemma R4 [simp]: "pmf (sds R4) a = pmf (sds R21) a" "pmf (sds R4) b = 0" "pmf (sds R4) c = 1 - pmf (sds R4) a" "pmf (sds R4) d = 0" proof - from R30_R21.strategyproofness(1) R30.support lottery_conditions[OF R30.wf] have "pmf (sds R4) c + pmf (sds R21) a \ pmf (sds R4) c + pmf (sds R30) a" by auto also { have "pmf (sds R30) a \ pmf (sds R47) a" using R47_R30.strategyproofness(1) R30.support R47.support lottery_conditions[OF R4.wf] lottery_conditions[OF R47.wf] by auto moreover from R4_R47.strategyproofness(1) R4.support R47.support lottery_conditions[OF R4.wf] lottery_conditions[OF R47.wf] have "pmf (sds R4) c \ pmf (sds R47) c" by simp ultimately have "pmf (sds R4) c + pmf (sds R30) a \ 1 - pmf (sds R47) d" using lottery_conditions[OF R47.wf] R47.support by simp } finally have "pmf (sds R4) c + pmf (sds R14) a \ 1" using lottery_conditions[OF R47.wf] by (auto simp del: pmf_nonneg) with R4_R18.strategyproofness(1) lottery_conditions[OF R4.wf] R4.support show "pmf (sds R4) a = pmf (sds R21) a" "pmf (sds R4) b = 0" "pmf (sds R4) c = 1 - pmf (sds R4) a" "pmf (sds R4) d = 0" by auto qed lemma R8_d [simp]: "pmf (sds R8) d = 1 - pmf (sds R8) a" and R8_c [simp]: "pmf (sds R8) c = 0" and R26_a [simp]: "pmf (sds R26) a = 1 - pmf (sds R8) a" proof - from R8_R26.strategyproofness(2) R8.support lottery_conditions[OF R8.wf] have "pmf (sds R26) a \ pmf (sds R8) d" by auto with R26_R8.strategyproofness(2) R8.support lottery_conditions[OF R8.wf] have "pmf (sds R26) a = pmf (sds R8) d" by auto with R8_R26.strategyproofness(2) R8.support lottery_conditions[OF R8.wf] show "pmf (sds R8) c = 0" "pmf (sds R8) d = 1 - pmf (sds R8) a" "pmf (sds R26) a = 1 - pmf (sds R8) a" by (auto simp del: pmf_nonneg) qed lemma R21_R47: "pmf (sds R21) d \ pmf (sds R47) c" using R4_R47.strategyproofness(1) R4.support R47.support lottery_conditions[OF R4.wf] lottery_conditions[OF R47.wf] by auto lemma R30 [simp]: "pmf (sds R30) a = pmf (sds R47) a" "pmf (sds R30) b = 0" "pmf (sds R30) c = 0" "pmf (sds R30) d = 1 - pmf (sds R47) a" proof - have A: "pmf (sds R30) a \ pmf (sds R47) a" using R47_R30.strategyproofness(1) R30.support R47.support lottery_conditions[OF R4.wf] lottery_conditions[OF R47.wf] by auto with R21_R47 R30_R21.strategyproofness(1) lottery_conditions[OF R30.wf] lottery_conditions[OF R47.wf] show "pmf (sds R30) a = pmf (sds R47) a" "pmf (sds R30) b = 0" "pmf (sds R30) c = 0" "pmf (sds R30) d = 1 - pmf (sds R47) a" by (auto simp: R30.support R47.support simp del: pmf_nonneg) (* tricky step! *) qed lemma R31_c_ge_one_half: "pmf (sds R31) c \ 1/2" proof - from R25.support have "pmf (sds R25) a \ 1/2" proof assume "pmf (sds R25) c = 0" with R25_R36.strategyproofness(1) lottery_conditions[OF R36.wf] show "pmf (sds R25) a \ 1/2" by (auto simp del: pmf_nonneg) next assume [simp]: "pmf (sds R25) b = 0" from R36_R25.strategyproofness(1) lottery_conditions[OF R25.wf] have "pmf (sds R25) c + pmf (sds R25) a \ pmf (sds R36) c + 1 / 2" by auto with R25_R36.strategyproofness(1) show "pmf (sds R25) a \ 1/2" by auto qed hence "pmf (sds R26) a \ 1/2" using R25_R26.strategyproofness(1) lottery_conditions[OF R25.wf] by (auto simp del: pmf_nonneg) with lottery_conditions[OF R47.wf] have "1/2 \ pmf (sds R26) a + pmf (sds R47) d" by (simp del: pmf_nonneg) also have "\ = 1 - pmf (sds R8) a + pmf (sds R47) d" by simp also from R4_R8.strategyproofness(1) have "1 - pmf (sds R8) a \ pmf (sds R21) d" by auto also note R21_R47 also from R30_R41.strategyproofness(1) R41.support lottery_conditions[OF R41.wf] lottery_conditions[OF R47.wf] have "pmf (sds R47) c + pmf (sds R47) d \ pmf (sds R41) d" by (auto simp del: pmf_nonneg) also from R41_R31.strategyproofness(1) R41.support lottery_conditions[OF R31.wf] lottery_conditions[OF R41.wf] have "pmf (sds R41) d \ pmf (sds R31) c" by auto finally show "pmf (sds R31) c \ 1/2" by simp qed lemma R31: "pmf (sds R31) a = 0" "pmf (sds R31) c = 1/2" "pmf (sds R31) b + pmf (sds R31) d = 1/2" proof - from R2_R38.strategyproofness(1) lottery_conditions[OF R38.wf] have A: "pmf (sds R38) b + pmf (sds R38) d \ 1/2" by auto with R31_c_ge_one_half R31_R38.strategyproofness(1) lottery_conditions[OF R31.wf] lottery_conditions[OF R38.wf] have "pmf (sds R38) b + pmf (sds R38) d = pmf (sds R31) d + pmf (sds R31) b" by auto with R31_c_ge_one_half A lottery_conditions[OF R31.wf] lottery_conditions[OF R38.wf] show "pmf (sds R31) a = 0" "pmf (sds R31) c = 1/2" "pmf (sds R31) b + pmf (sds R31) d = 1/2" by linarith+ qed lemma absurd: False using R31 R45_R31.strategyproofness(2) by simp (* TODO (Re-)move *) (* This is just to output a list of all the Strategy-Proofness conditions used in the proof *) (* ML_val \ let val thms = @{thms R1_R2.strategyproofness(1) R1_R19.strategyproofness(1) R2_R1.strategyproofness(1) R2_R38.strategyproofness(1) R4_R8.strategyproofness(1) R4_R18.strategyproofness(1) R4_R47.strategyproofness(1) R5_R7.strategyproofness(1) R5_R10.strategyproofness(1) R5_R17.strategyproofness(1) R6_R19.strategyproofness(1) R6_R42.strategyproofness(1) R7_R43.strategyproofness(1) R8_R26.strategyproofness(2) R9_R18.strategyproofness(1) R9_R35.strategyproofness(1) R9_R40.strategyproofness(1) R10_R12.strategyproofness(1) R10_R15.strategyproofness(1) R10_R19.strategyproofness(1) R10_R36.strategyproofness(1) R12_R10.strategyproofness(1) R12_R16.strategyproofness(1) R12_R44.strategyproofness(1) R13_R15.strategyproofness(1) R13_R27.strategyproofness(1) R14_R9.strategyproofness(1) R14_R16.strategyproofness(1) R14_R34.strategyproofness(1) R15_R5.strategyproofness(1) R15_R7.strategyproofness(1) R15_R10.strategyproofness(1) R15_R13.strategyproofness(1) R17_R3.strategyproofness(1) R17_R5.strategyproofness(1) R17_R7.strategyproofness(1) R17_R11.strategyproofness(1) R18_R9.strategyproofness(1) R19_R1.strategyproofness(1) R19_R10.strategyproofness(1) R19_R27.strategyproofness(1) R20_R21.strategyproofness(1) R21_R35.strategyproofness(1) R22_R29.strategyproofness(1) R22_R32.strategyproofness(1) R23_R12.strategyproofness(1) R23_R18.strategyproofness(1) R23_R19.strategyproofness(1) R24_R34.strategyproofness(1) R25_R26.strategyproofness(1) R25_R36.strategyproofness(1) R26_R8.strategyproofness(2) R27_R13.strategyproofness(1) R27_R19.strategyproofness(1) R28_R32.strategyproofness(1) R28_R39.strategyproofness(1) R29_R39.strategyproofness(1) R30_R21.strategyproofness(1) R30_R41.strategyproofness(1) R31_R38.strategyproofness(1) R32_R22.strategyproofness(1) R32_R28.strategyproofness(1) R33_R5.strategyproofness(1) R33_R22.strategyproofness(1) R34_R24.strategyproofness(1) R35_R9.strategyproofness(1) R35_R21.strategyproofness(1) R36_R10.strategyproofness(1) R36_R25.strategyproofness(1) R36_R39.strategyproofness(1) R37_R42.strategyproofness(1) R37_R42.strategyproofness(2) R39_R2.strategyproofness(1) R39_R29.strategyproofness(1) R39_R36.strategyproofness(1) R41_R31.strategyproofness(1) R42_R3.strategyproofness(1) R42_R11.strategyproofness(1) R42_R24.strategyproofness(1) R44_R12.strategyproofness(1) R44_R40.strategyproofness(1) R45_R31.strategyproofness(2) R46_R20.strategyproofness(1) R46_R37.strategyproofness(1) R47_R30.strategyproofness(1) }; in thms |> map (Pretty.quote o Pretty.str o Pretty.unformatted_string_of o Syntax.pretty_term @{context} o Thm.prop_of) |> Pretty.list "[" "]" |> (fn x => Pretty.block [Pretty.str "thms = ", x]) |> Pretty.string_of |> writeln end \*) end subsection \Lifting to more than 4 agents and alternatives\ (* TODO: Move? *) lemma finite_list': assumes "finite A" obtains xs where "A = set xs" "distinct xs" "length xs = card A" proof - from assms obtain xs where "set xs = A" using finite_list by blast thus ?thesis using distinct_card[of "remdups xs"] by (intro that[of "remdups xs"]) simp_all qed lemma finite_list_subset: assumes "finite A" "card A \ n" obtains xs where "set xs \ A" "distinct xs" "length xs = n" proof - obtain xs where "A = set xs" "distinct xs" "length xs = card A" using finite_list'[OF assms(1)] by blast with assms show ?thesis by (intro that[of "take n xs"]) (simp_all add: set_take_subset) qed lemma card_ge_4E: assumes "finite A" "card A \ 4" obtains a b c d where "distinct [a,b,c,d]" "{a,b,c,d} \ A" proof - - from finite_list_subset[OF assms] guess xs . - moreover then obtain a b c d where "xs = [a, b, c, d]" + from assms obtain xs where xs: "set xs \ A" "distinct xs" "length xs = 4" + by (rule finite_list_subset) + then obtain a b c d where "xs = [a, b, c, d]" by (auto simp: eval_nat_numeral length_Suc_conv) - ultimately show ?thesis by (intro that[of a b c d]) simp_all + with xs show ?thesis by (intro that[of a b c d]) simp_all qed context sds_impossibility begin lemma absurd: False proof - - from card_ge_4E[OF finite_agents agents_ge_4] guess A1 A2 A3 A4 . - note agents = this - from card_ge_4E[OF finite_alts alts_ge_4] guess a b c d . - note alts = this + from card_ge_4E[OF finite_agents agents_ge_4] + obtain A1 A2 A3 A4 where agents: "distinct [A1, A2, A3, A4]" "{A1, A2, A3, A4} \ agents" . + from card_ge_4E[OF finite_alts alts_ge_4] + obtain a b c d where alts: "distinct [a, b, c, d]" "{a, b, c, d} \ alts" . define agents' alts' where "agents' = {A1,A2,A3,A4}" and "alts' = {a,b,c,d}" from agents alts interpret sds_lowering_anonymous_neutral_sdeff_stratproof agents alts sds agents' alts' unfolding agents'_def alts'_def by unfold_locales simp_all from agents alts interpret sds_impossibility_4_4 agents' alts' lowered A1 A2 A3 A4 a b c d by unfold_locales (simp_all add: agents'_def alts'_def) from absurd show False . qed end end diff --git a/thys/Stirling_Formula/Gamma_Asymptotics.thy b/thys/Stirling_Formula/Gamma_Asymptotics.thy --- a/thys/Stirling_Formula/Gamma_Asymptotics.thy +++ b/thys/Stirling_Formula/Gamma_Asymptotics.thy @@ -1,1891 +1,1893 @@ (* File: Gamma_Asymptotics.thy Author: Manuel Eberl The complete asymptotics of the real and complex logarithmic Gamma functions. Also of the real Polygamma functions (could be extended to the complex ones fairly easily if needed). *) section \Complete asymptotics of the logarithmic Gamma function\ theory Gamma_Asymptotics imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Real_Asymp.Real_Asymp" Bernoulli.Bernoulli_FPS Bernoulli.Periodic_Bernpoly Stirling_Formula begin subsection \Auxiliary Facts\ (* TODO: could be automated with Laurent series expansions in the future *) lemma stirling_limit_aux1: "((\y. Ln (1 + z * of_real y) / of_real y) \ z) (at_right 0)" for z :: complex proof (cases "z = 0") case True then show ?thesis by simp next case False have "((\y. ln (1 + z * of_real y)) has_vector_derivative 1 * z) (at 0)" by (rule has_vector_derivative_real_field) (auto intro!: derivative_eq_intros) then have "(\y. (Ln (1 + z * of_real y) - of_real y * z) / of_real \y\) \0\ 0" by (auto simp add: has_vector_derivative_def has_derivative_def netlimit_at scaleR_conv_of_real field_simps) then have "((\y. (Ln (1 + z * of_real y) - of_real y * z) / of_real \y\) \ 0) (at_right 0)" by (rule filterlim_mono[OF _ _ at_le]) simp_all also have "?this \ ((\y. Ln (1 + z * of_real y) / (of_real y) - z) \ 0) (at_right 0)" using eventually_at_right_less[of "0::real"] by (intro filterlim_cong refl) (auto elim!: eventually_mono simp: field_simps) finally show ?thesis by (simp only: LIM_zero_iff) qed lemma stirling_limit_aux2: "((\y. y * Ln (1 + z / of_real y)) \ z) at_top" for z :: complex using stirling_limit_aux1[of z] by (subst filterlim_at_top_to_right) (simp add: field_simps) lemma Union_atLeastAtMost: assumes "N > 0" shows "(\n\{0.. {0..real N}" thus "x \ (\n\{0.. 0" "x < real N" by simp_all hence "x \ real (nat \x\)" "x \ real (nat \x\ + 1)" by linarith+ moreover from x have "nat \x\ < N" by linarith ultimately have "\n\{0.. {real n..real (n + 1)}" by (intro bexI[of _ "nat \x\"]) simp_all thus ?thesis by blast qed qed auto subsection \Cones in the complex plane\ definition complex_cone :: "real \ real \ complex set" where "complex_cone a b = {z. \y\{a..b}. z = rcis (norm z) y}" abbreviation complex_cone' :: "real \ complex set" where "complex_cone' a \ complex_cone (-a) a" lemma zero_in_complex_cone [simp, intro]: "a \ b \ 0 \ complex_cone a b" by (auto simp: complex_cone_def) lemma complex_coneE: assumes "z \ complex_cone a b" obtains r \ where "r \ 0" "\ \ {a..b}" "z = rcis r \" proof - from assms obtain y where "y \ {a..b}" "z = rcis (norm z) y" unfolding complex_cone_def by auto thus ?thesis using that[of "norm z" y] by auto qed lemma arg_cis [simp]: assumes "x \ {-pi<..pi}" shows "Arg (cis x) = x" using assms by (intro cis_Arg_unique) auto lemma arg_mult_of_real_left [simp]: assumes "r > 0" shows "Arg (of_real r * z) = Arg z" proof (cases "z = 0") case False thus ?thesis using Arg_bounded[of z] assms by (intro cis_Arg_unique) (auto simp: sgn_mult sgn_of_real cis_Arg) qed auto lemma arg_mult_of_real_right [simp]: assumes "r > 0" shows "Arg (z * of_real r) = Arg z" by (subst mult.commute, subst arg_mult_of_real_left) (simp_all add: assms) lemma arg_rcis [simp]: assumes "x \ {-pi<..pi}" "r > 0" shows "Arg (rcis r x) = x" using assms by (simp add: rcis_def) lemma rcis_in_complex_cone [intro]: assumes "\ \ {a..b}" "r \ 0" shows "rcis r \ \ complex_cone a b" using assms by (auto simp: complex_cone_def) lemma arg_imp_in_complex_cone: assumes "Arg z \ {a..b}" shows "z \ complex_cone a b" proof - have "z = rcis (norm z) (Arg z)" by (simp add: rcis_cmod_Arg) also have "\ \ complex_cone a b" using assms by auto finally show ?thesis . qed lemma complex_cone_altdef: assumes "-pi < a" "a \ b" "b \ pi" shows "complex_cone a b = insert 0 {z. Arg z \ {a..b}}" proof (intro equalityI subsetI) fix z assume "z \ complex_cone a b" then obtain r \ where *: "r \ 0" "\ \ {a..b}" "z = rcis r \" by (auto elim: complex_coneE) have "Arg z \ {a..b}" if [simp]: "z \ 0" proof - have "r > 0" using that * by (subst (asm) *) auto hence "\ \ {a..b}" using *(1,2) assms by (auto simp: *(1)) moreover from assms *(2) have "\ \ {-pi<..pi}" by auto ultimately show ?thesis using *(3) \r > 0\ by (subst *) auto qed thus "z \ insert 0 {z. Arg z \ {a..b}}" by auto qed (use assms in \auto intro: arg_imp_in_complex_cone\) lemma nonneg_of_real_in_complex_cone [simp, intro]: assumes "x \ 0" "a \ 0" "0 \ b" shows "of_real x \ complex_cone a b" proof - from assms have "rcis x 0 \ complex_cone a b" by (intro rcis_in_complex_cone) auto thus ?thesis by simp qed lemma one_in_complex_cone [simp, intro]: "a \ 0 \ 0 \ b \ 1 \ complex_cone a b" using nonneg_of_real_in_complex_cone[of 1] by (simp del: nonneg_of_real_in_complex_cone) lemma of_nat_in_complex_cone [simp, intro]: "a \ 0 \ 0 \ b \ of_nat n \ complex_cone a b" using nonneg_of_real_in_complex_cone[of "real n"] by (simp del: nonneg_of_real_in_complex_cone) subsection \Another integral representation of the Beta function\ lemma complex_cone_inter_nonpos_Reals: assumes "-pi < a" "a \ b" "b < pi" shows "complex_cone a b \ \\<^sub>\\<^sub>0 = {0}" proof (safe elim!: nonpos_Reals_cases) fix x :: real assume "complex_of_real x \ complex_cone a b" "x \ 0" hence "\(x < 0)" using assms by (intro notI) (auto simp: complex_cone_altdef) with \x \ 0\ show "complex_of_real x = 0" by auto qed (use assms in auto) theorem assumes a: "a > 0" and b: "b > (0 :: real)" shows has_integral_Beta_real': "((\u. u powr (b - 1) / (1 + u) powr (a + b)) has_integral Beta a b) {0<..}" and Beta_conv_nn_integral: "Beta a b = (\\<^sup>+u. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) \lborel)" proof - define I where "I = (\\<^sup>+u. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) \lborel)" have "Gamma (a + b) > 0" "Beta a b > 0" using assms by (simp_all add: add_pos_pos Beta_def) from a b have "ennreal (Gamma a * Gamma b) = (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (a - 1) / exp t) \lborel) * (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (b - 1) / exp t) \lborel)" by (subst ennreal_mult') (simp_all add: Gamma_conv_nn_integral_real) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator {0..} t * t powr (a - 1) / exp t) * ennreal (indicator {0..} u * u powr (b - 1) / exp u) \lborel \lborel)" by (simp add: nn_integral_cmult nn_integral_multc) also have "\ = (\\<^sup>+t. indicator {0<..} t * (\\<^sup>+u. indicator {0..} u * t powr (a - 1) * u powr (b - 1) / exp (t + u) \lborel) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"]) (auto simp: indicator_def divide_ennreal ennreal_mult' [symmetric] exp_add mult_ac) also have "\ = (\\<^sup>+t. indicator {0<..} t * (\\<^sup>+u. indicator {0..} u * t powr (a - 1) * u powr (b - 1) / exp (t + u) \(density (distr lborel borel ((*) t)) (\x. ennreal \t\))) \lborel)" by (intro nn_integral_cong mult_indicator_cong, subst lborel_distr_mult' [symmetric]) auto also have "\ = (\\<^sup>+(t::real). indicator {0<..} t * (\\<^sup>+u. indicator {0..} (u * t) * t powr a * (u * t) powr (b - 1) / exp (t + t * u) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: nn_integral_density nn_integral_distr algebra_simps powr_diff simp flip: ennreal_mult) also have "\ = (\\<^sup>+(t::real). \\<^sup>+u. indicator ({0<..}\{0..}) (t, u) * t powr a * (u * t) powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (subst nn_integral_cmult [symmetric], simp, intro nn_integral_cong) (auto simp: indicator_def zero_le_mult_iff algebra_simps) also have "\ = (\\<^sup>+(t::real). \\<^sup>+u. indicator ({0<..}\{0..}) (t, u) * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: powr_add powr_diff indicator_def powr_mult field_simps) also have "\ = (\\<^sup>+(u::real). \\<^sup>+t. indicator ({0<..}\{0..}) (t, u) * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (rule lborel_pair.Fubini') auto also have "\ = (\\<^sup>+(u::real). indicator {0..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"]) (auto simp: indicator_def) also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \(density (distr lborel borel ((*) (1/(1+u)))) (\x. ennreal \1/(1+u)\))) \lborel)" by (intro nn_integral_cong mult_indicator_cong, subst lborel_distr_mult' [symmetric]) auto also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. ennreal (1 / (u + 1)) * ennreal (indicator {0<..} (t / (u + 1)) * (t / (1+u)) powr (a + b - 1) * u powr (b - 1) / exp t) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: nn_integral_distr nn_integral_density add_ac) also have "\ = (\\<^sup>+u. \\<^sup>+t. indicator ({0<..}\{0<..}) (u, t) * 1/(u+1) * (t / (u+1)) powr (a + b - 1) * u powr (b - 1) / exp t \lborel \lborel)" by (subst nn_integral_cmult [symmetric], simp, intro nn_integral_cong) (auto simp: indicator_def field_simps divide_ennreal simp flip: ennreal_mult ennreal_mult') also have "\ = (\\<^sup>+u. \\<^sup>+t. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) * ennreal (indicator {0<..} t * t powr (a + b - 1) / exp t) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def powr_add powr_diff powr_divide powr_minus divide_simps add_ac simp flip: ennreal_mult) also have "\ = I * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) / exp t \lborel)" by (simp add: nn_integral_cmult nn_integral_multc I_def) also have "(\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) / exp t \lborel) = ennreal (Gamma (a + b))" using assms by (subst Gamma_conv_nn_integral_real) (auto intro!: nn_integral_cong_AE[OF AE_I[of _ _ "{0}"]] simp: indicator_def split: if_splits) finally have "ennreal (Gamma a * Gamma b) = I * ennreal (Gamma (a + b))" . hence "ennreal (Gamma a * Gamma b) / ennreal (Gamma (a + b)) = I * ennreal (Gamma (a + b)) / ennreal (Gamma (a + b))" by simp also have "\ = I" using \Gamma (a + b) > 0\ by (intro ennreal_mult_divide_eq) (auto simp: ) also have "ennreal (Gamma a * Gamma b) / ennreal (Gamma (a + b)) = ennreal (Gamma a * Gamma b / Gamma (a + b))" using assms by (intro divide_ennreal) auto also have "\ = ennreal (Beta a b)" by (simp add: Beta_def) finally show *: "ennreal (Beta a b) = I" . define f where "f = (\u. u powr (b - 1) / (1 + u) powr (a + b))" have "((\u. indicator {0<..} u * f u) has_integral Beta a b) UNIV" using * \Beta a b > 0\ by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: f_def measurable_completion nn_integral_completion I_def mult_ac) also have "(\u. indicator {0<..} u * f u) = (\u. if u \ {0<..} then f u else 0)" by (auto simp: fun_eq_iff) also have "(\ has_integral Beta a b) UNIV \ (f has_integral Beta a b) {0<..}" by (rule has_integral_restrict_UNIV) finally show \ by (simp add: f_def) qed lemma has_integral_Beta2: fixes a :: real assumes "a < -1/2" shows "((\x. (1 + x ^ 2) powr a) has_integral Beta (- a - 1 / 2) (1 / 2) / 2) {0<..}" proof - define f where "f = (\u. u powr (-1/2) / (1 + u) powr (-a))" define C where "C = Beta (-a-1/2) (1/2)" have I: "(f has_integral C) {0<..}" using has_integral_Beta_real'[of "-a-1/2" "1/2"] assms by (simp_all add: diff_divide_distrib f_def C_def) define g where "g = (\x. x ^ 2 :: real)" have bij: "bij_betw g {0<..} {0<..}" by (intro bij_betwI[of _ _ _ sqrt]) (auto simp: g_def) have "(f absolutely_integrable_on g ` {0<..} \ integral (g ` {0<..}) f = C)" using I bij by (simp add: bij_betw_def has_integral_iff absolutely_integrable_on_def f_def) also have "?this \ ((\x. \2 * x\ *\<^sub>R f (g x)) absolutely_integrable_on {0<..} \ integral {0<..} (\x. \2 * x\ *\<^sub>R f (g x)) = C)" using bij by (intro has_absolute_integral_change_of_variables_1' [symmetric]) (auto intro!: derivative_eq_intros simp: g_def bij_betw_def) finally have "((\x. \2 * x\ * f (g x)) has_integral C) {0<..}" by (simp add: absolutely_integrable_on_def f_def has_integral_iff) also have "?this \ ((\x::real. 2 * (1 + x\<^sup>2) powr a) has_integral C) {0<..}" by (intro has_integral_cong) (auto simp: f_def g_def powr_def exp_minus ln_realpow field_simps) finally have "((\x::real. 1/2 * (2 * (1 + x\<^sup>2) powr a)) has_integral 1/2 * C) {0<..}" by (intro has_integral_mult_right) thus ?thesis by (simp add: C_def) qed lemma has_integral_Beta3: fixes a b :: real assumes "a < -1/2" and "b > 0" shows "((\x. (b + x ^ 2) powr a) has_integral Beta (-a - 1/2) (1/2) / 2 * b powr (a + 1/2)) {0<..}" proof - define C where "C = Beta (- a - 1 / 2) (1 / 2) / 2" have int: "nn_integral lborel (\x. indicator {0<..} x * (1 + x ^ 2) powr a) = C" using nn_integral_has_integral_lebesgue[OF _ has_integral_Beta2[OF assms(1)]] by (auto simp: C_def) have "nn_integral lborel (\x. indicator {0<..} x * (b + x ^ 2) powr a) = (\\<^sup>+x. ennreal (indicat_real {0<..} (x * sqrt b) * (b + (x * sqrt b)\<^sup>2) powr a * sqrt b) \lborel)" using assms by (subst lborel_distr_mult'[of "sqrt b"]) (auto simp: nn_integral_density nn_integral_distr mult_ac simp flip: ennreal_mult) also have "\ = (\\<^sup>+x. ennreal (indicat_real {0<..} x * (b * (1 + x ^ 2)) powr a * sqrt b) \lborel)" using assms by (intro nn_integral_cong) (auto simp: indicator_def field_simps zero_less_mult_iff) also have "\ = (\\<^sup>+x. ennreal (indicat_real {0<..} x * b powr (a + 1/2) * (1 + x ^ 2) powr a) \lborel)" using assms by (intro nn_integral_cong) (auto simp: indicator_def powr_add powr_half_sqrt powr_mult) also have "\ = b powr (a + 1/2) * (\\<^sup>+x. ennreal (indicat_real {0<..} x * (1 + x ^ 2) powr a) \lborel)" using assms by (subst nn_integral_cmult [symmetric]) (simp_all add: mult_ac flip: ennreal_mult) also have "(\\<^sup>+x. ennreal (indicat_real {0<..} x * (1 + x ^ 2) powr a) \lborel) = C" using int by simp also have "ennreal (b powr (a + 1/2)) * ennreal C = ennreal (C * b powr (a + 1/2))" using assms by (subst ennreal_mult) (auto simp: C_def mult_ac Beta_def) finally have *: "(\\<^sup>+ x. ennreal (indicat_real {0<..} x * (b + x\<^sup>2) powr a) \lborel) = \" . hence "((\x. indicator {0<..} x * (b + x^2) powr a) has_integral C * b powr (a + 1/2)) UNIV" using assms by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: C_def measurable_completion nn_integral_completion Beta_def) also have "(\x. indicator {0<..} x * (b + x^2) powr a) = (\x. if x \ {0<..} then (b + x^2) powr a else 0)" by (auto simp: fun_eq_iff) finally show ?thesis by (subst (asm) has_integral_restrict_UNIV) (auto simp: C_def) qed subsection \Asymptotics of the real $\log\Gamma$ function and its derivatives\ text \ This is the error term that occurs in the expansion of @{term ln_Gamma}. It can be shown to be of order $O(s^{-n})$. \ definition stirling_integral :: "nat \ 'a :: {real_normed_div_algebra, banach} \ 'a" where "stirling_integral n s = lim (\N. integral {0..N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n))" context fixes s :: complex assumes s: "s \ \\<^sub>\\<^sub>0" fixes approx :: "nat \ complex" defines "approx \ (\N. (\n = 1.. \\\ ln_Gamma s\\ (ln_Gamma (of_nat N) - ln (2 * pi / of_nat N) / 2 - of_nat N * ln (of_nat N) + of_nat N) - \ \\\ 0\\ s * (harm (N - 1) - ln (of_nat (N - 1)) - euler_mascheroni) + \ \\\ 0\\ s * (ln (of_nat N + s) - ln (of_nat (N - 1))) - \ \\\ 0\\ (1/2) * (ln (of_nat N + s) - ln (of_nat N)) + \ \\\ 0\\ of_nat N * (ln (of_nat N + s) - ln (of_nat N)) - \ \\\ s\\ (s - 1/2) * ln s - ln (2 * pi) / 2)" begin qualified lemma assumes N: "N > 0" shows integrable_pbernpoly_1: "(\x. of_real (-pbernpoly 1 x) / (of_real x + s)) integrable_on {0..real N}" and integral_pbernpoly_1_aux: "integral {0..real N} (\x. -of_real (pbernpoly 1 x) / (of_real x + s)) = approx N" and has_integral_pbernpoly_1: "((\x. pbernpoly 1 x /(x + s)) has_integral (\mx. -pbernpoly 1 x / (x + s)) has_integral (of_nat n + 1/2 + s) * (ln (of_nat (n + 1) + s) - ln (of_nat n + s)) - 1) {of_nat n..of_nat (n + 1)}" for n proof (rule has_integral_spike) have "((\x. (of_nat n + 1/2 + s) * (1 / (x + s)) - 1) has_integral (of_nat n + 1/2 + s) * (ln (of_real (real (n + 1)) + s) - ln (of_real (real n) + s)) - 1) {of_nat n..of_nat (n + 1)}" using s has_integral_const_real[of 1 "of_nat n" "of_nat (n + 1)"] by (intro has_integral_diff has_integral_mult_right fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: has_field_derivative_iff_has_vector_derivative [symmetric] field_simps complex_nonpos_Reals_iff) thus "((\x. (of_nat n + 1/2 + s) * (1 / (x + s)) - 1) has_integral (of_nat n + 1/2 + s) * (ln (of_nat (n + 1) + s) - ln (of_nat n + s)) - 1) {of_nat n..of_nat (n + 1)}" by simp show "-pbernpoly 1 x / (x + s) = (of_nat n + 1/2 + s) * (1 / (x + s)) - 1" if "x \ {of_nat n..of_nat (n + 1)} - {of_nat (n + 1)}" for x proof - have x: "x \ real n" "x < real (n + 1)" using that by simp_all hence "floor x = int n" by linarith moreover from s x have "complex_of_real x \ -s" by (auto simp add: complex_eq_iff complex_nonpos_Reals_iff simp del: of_nat_Suc) ultimately show "-pbernpoly 1 x / (x + s) = (of_nat n + 1/2 + s) * (1 / (x + s)) - 1" by (auto simp: pbernpoly_def bernpoly_def frac_def divide_simps add_eq_0_iff2) qed qed simp_all hence *: "\I. I\?A \ ((\x. -pbernpoly 1 x / (x + s)) has_integral (Inf I + 1/2 + s) * (ln (Inf I + 1 + s) - ln (Inf I + s)) - 1) I" by (auto simp: add_ac) have "((\x. - pbernpoly 1 x / (x + s)) has_integral (\I\?A. (Inf I + 1 / 2 + s) * (ln (Inf I + 1 + s) - ln (Inf I + s)) - 1)) (\n\{0..x. - pbernpoly 1 x / (x + s)) has_integral ?i) {0..real N}" by (subst has_integral_spike_set_eq) (use Union_atLeastAtMost assms in \auto simp: intro!: empty_imp_negligible\) hence "(\x. - pbernpoly 1 x / (x + s)) integrable_on {0..real N}" and integral: "integral {0..real N} (\x. - pbernpoly 1 x / (x + s)) = ?i" by (simp_all add: has_integral_iff) show "(\x. - pbernpoly 1 x / (x + s)) integrable_on {0..real N}" by fact note has_integral_neg[OF has_integral] also have "-?i = (\xx. of_real (pbernpoly 1 x) / (of_real x + s)) has_integral \) {0..real N}" by simp note integral also have "?i = (\nnnn=1.. = (\n=1..n=1..n=1..n=1.. - (\n = 1..n=1.. 0" for x by (auto simp: complex_nonpos_Reals_iff complex_eq_iff) hence "(\n=1..n=1.. = ln (fact (N - 1)) + (\n=1..n=1..n=1..n=1..n = 1..n = 1..x. -of_real (pbernpoly 1 x) / (of_real x + s)) = \" by simp qed lemma integrable_ln_Gamma_aux: shows "(\x. of_real (pbernpoly n x) / (of_real x + s) ^ n) integrable_on {0..real N}" proof (cases "n = 1") case True with s show ?thesis using integrable_neg[OF integrable_pbernpoly_1[of N]] by (cases "N = 0") (simp_all add: integrable_negligible) next case False from s have "of_real x + s \ 0" if "x \ 0" for x using that by (auto simp: complex_eq_iff add_eq_0_iff2 complex_nonpos_Reals_iff) with False s show ?thesis by (auto intro!: integrable_continuous_real continuous_intros) qed text \ This following proof is based on ``Rudiments of the theory of the gamma function'' by Bruce Berndt~\cite{berndt}. \ lemma tendsto_of_real_0_I: "(f \ 0) G \ ((\x. (of_real (f x))) \ (0 ::'a::real_normed_div_algebra)) G" using tendsto_of_real_iff by force qualified lemma integral_pbernpoly_1: "(\N. integral {0..real N} (\x. pbernpoly 1 x / (x + s))) \ -ln_Gamma s - s + (s - 1 / 2) * ln s + ln (2 * pi) / 2" proof - have neq: "s + of_real x \ 0" if "x \ 0" for x :: real using that s by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) have "(approx \ ln_Gamma s - 0 - 0 + 0 - 0 + s - (s - 1/2) * ln s - ln (2 * pi) / 2) at_top" unfolding approx_def proof (intro tendsto_add tendsto_diff) from s have s': "s \ \\<^sub>\\<^sub>0" by (auto simp: complex_nonpos_Reals_iff elim!: nonpos_Ints_cases) have "(\n. \i=1.. ln_Gamma s + euler_mascheroni * s + ln s" (is "?f \ _") using ln_Gamma_series'_aux[OF s'] unfolding sums_def by (subst filterlim_sequentially_Suc [symmetric], subst (asm) sum.atLeast1_atMost_eq [symmetric]) (simp add: atLeastLessThanSuc_atLeastAtMost) thus "((\n. ?f n - (euler_mascheroni * s + ln s)) \ ln_Gamma s) at_top" by (auto intro: tendsto_eq_intros) next show "(\x. complex_of_real (ln_Gamma (real x) - ln (2 * pi / real x) / 2 - real x * ln (real x) + real x)) \ 0" proof (intro tendsto_of_real_0_I filterlim_compose[OF tendsto_sandwich filterlim_real_sequentially]) show "eventually (\x::real. ln_Gamma x - ln (2 * pi / x) / 2 - x * ln x + x \ 0) at_top" using eventually_ge_at_top[of "1::real"] by eventually_elim (insert ln_Gamma_bounds(1), simp add: algebra_simps) show "eventually (\x::real. ln_Gamma x - ln (2 * pi / x) / 2 - x * ln x + x \ 1 / 12 * inverse x) at_top" using eventually_ge_at_top[of "1::real"] by eventually_elim (insert ln_Gamma_bounds(2), simp add: field_simps) show "((\x::real. 1 / 12 * inverse x) \ 0) at_top" by (intro tendsto_mult_right_zero tendsto_inverse_0_at_top filterlim_ident) qed simp_all next have "(\x. s * of_real (harm (x - 1) - ln (real (x - 1)) - euler_mascheroni)) \ s * of_real (euler_mascheroni - euler_mascheroni)" by (subst filterlim_sequentially_Suc [symmetric], intro tendsto_intros) (insert euler_mascheroni_LIMSEQ, simp_all) also have "?this \ (\x. s * (harm (x - 1) - ln (of_nat (x - 1)) - euler_mascheroni)) \ 0" by (intro filterlim_cong refl eventually_mono[OF eventually_gt_at_top[of "1::nat"]]) (auto simp: Ln_of_nat of_real_harm) finally show "(\x. s * (harm (x - 1) - ln (of_nat (x - 1)) - euler_mascheroni)) \ 0" . next have "((\x. ln (1 + (s + 1) / of_real x)) \ ln (1 + 0)) at_top" (is ?P) by (intro tendsto_intros tendsto_divide_0[OF tendsto_const]) (simp_all add: filterlim_ident filterlim_at_infinity_conv_norm_at_top filterlim_abs_real) also have "ln (of_real (x + 1) + s) - ln (complex_of_real x) = ln (1 + (s + 1) / of_real x)" if "x > 1" for x using that s using Ln_divide_of_real[of x "of_real (x + 1) + s", symmetric] neq[of "x+1"] by (simp add: field_simps Ln_of_real) hence "?P \ ((\x. ln (of_real (x + 1) + s) - ln (of_real x)) \ 0) at_top" by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "1::real"]]) finally have "((\n. ln (of_real (real n + 1) + s) - ln (of_real (real n))) \ 0) at_top" by (rule filterlim_compose[OF _ filterlim_real_sequentially]) hence "((\n. ln (of_nat n + s) - ln (of_nat (n - 1))) \ 0) at_top" by (subst filterlim_sequentially_Suc [symmetric]) (simp add: add_ac) thus "(\x. s * (ln (of_nat x + s) - ln (of_nat (x - 1)))) \ 0" by (rule tendsto_mult_right_zero) next have "((\x. ln (1 + s / of_real x)) \ ln (1 + 0)) at_top" (is ?P) by (intro tendsto_intros tendsto_divide_0[OF tendsto_const]) (simp_all add: filterlim_ident filterlim_at_infinity_conv_norm_at_top filterlim_abs_real) also have "ln (of_real x + s) - ln (of_real x) = ln (1 + s / of_real x)" if "x > 0" for x using Ln_divide_of_real[of x "of_real x + s"] neq[of x] that by (auto simp: field_simps Ln_of_real) hence "?P \ ((\x. ln (of_real x + s) - ln (of_real x)) \ 0) at_top" using s by (intro filterlim_cong refl) (auto intro: eventually_mono [OF eventually_gt_at_top[of "1::real"]]) finally have "(\x. (1/2) * (ln (of_real (real x) + s) - ln (of_real (real x)))) \ 0" by (rule tendsto_mult_right_zero[OF filterlim_compose[OF _ filterlim_real_sequentially]]) thus "(\x. (1/2) * (ln (of_nat x + s) - ln (of_nat x))) \ 0" by simp next have "((\x. x * (ln (1 + s / of_real x))) \ s) at_top" (is ?P) by (rule stirling_limit_aux2) also have "ln (1 + s / of_real x) = ln (of_real x + s) - ln (of_real x)" if "x > 1" for x using that s Ln_divide_of_real [of x "of_real x + s", symmetric] neq[of x] by (auto simp: Ln_of_real field_simps) hence "?P \ ((\x. of_real x * (ln (of_real x + s) - ln (of_real x))) \ s) at_top" by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "1::real"]]) finally have "(\n. of_real (real n) * (ln (of_real (real n) + s) - ln (of_real (real n)))) \ s" by (rule filterlim_compose[OF _ filterlim_real_sequentially]) thus "(\n. of_nat n * (ln (of_nat n + s) - ln (of_nat n))) \ s" by simp qed simp_all also have "?this \ ((\N. integral {0..real N} (\x. -pbernpoly 1 x / (x + s))) \ ln_Gamma s + s - (s - 1/2) * ln s - ln (2 * pi) / 2) at_top" using integral_pbernpoly_1_aux by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) also have "(\N. integral {0..real N} (\x. -pbernpoly 1 x / (x + s))) = (\N. -integral {0..real N} (\x. pbernpoly 1 x / (x + s)))" by (simp add: fun_eq_iff) finally show ?thesis by (simp add: tendsto_minus_cancel_left [symmetric] algebra_simps) qed qualified lemma pbernpoly_integral_conv_pbernpoly_integral_Suc: assumes "n \ 1" shows "integral {0..real N} (\x. pbernpoly n x / (x + s) ^ n) = of_real (pbernpoly (Suc n) (real N)) / (of_nat (Suc n) * (s + of_nat N) ^ n) - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n) + of_nat n / of_nat (Suc n) * integral {0..real N} (\x. of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n)" proof - note [derivative_intros] = has_field_derivative_pbernpoly_Suc' define I where "I = -of_real (pbernpoly (Suc n) (of_nat N)) / (of_nat (Suc n) * (of_nat N + s) ^ n) + of_real (bernoulli (Suc n) / real (Suc n)) / s ^ n + integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)" have "((\x. (-of_nat n * inverse (of_real x + s) ^ Suc n) * (of_real (pbernpoly (Suc n) x) / (of_nat (Suc n)))) has_integral -I) {0..real N}" proof (rule integration_by_parts_interior_strong[OF bounded_bilinear_mult]) fix x :: real assume x: "x \ {0<.. \" proof assume "x \ \" then obtain n where "x = of_int n" by (auto elim!: Ints_cases) with x have x': "x = of_nat (nat n)" by simp from x show False by (auto simp: x') qed hence "((\x. of_real (pbernpoly (Suc n) x / of_nat (Suc n))) has_vector_derivative complex_of_real (pbernpoly n x)) (at x)" by (intro has_vector_derivative_of_real) (auto intro!: derivative_eq_intros) thus "((\x. of_real (pbernpoly (Suc n) x) / of_nat (Suc n)) has_vector_derivative complex_of_real (pbernpoly n x)) (at x)" by simp from x s have "complex_of_real x + s \ 0" by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) thus "((\x. inverse (of_real x + s) ^ n) has_vector_derivative - of_nat n * inverse (of_real x + s) ^ Suc n) (at x)" using x s assms by (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: divide_simps power_add [symmetric] simp del: power_Suc) next have "complex_of_real x + s \ 0" if "x \ 0" for x using that s by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) thus "continuous_on {0..real N} (\x. inverse (of_real x + s) ^ n)" "continuous_on {0..real N} (\x. complex_of_real (pbernpoly (Suc n) x) / of_nat (Suc n))" using assms s by (auto intro!: continuous_intros simp del: of_nat_Suc) next have "((\x. inverse (of_real x + s) ^ n * of_real (pbernpoly n x)) has_integral pbernpoly (Suc n) (of_nat N) / (of_nat (Suc n) * (of_nat N + s) ^ n) - of_real (bernoulli (Suc n) / real (Suc n)) / s ^ n - -I) {0..real N}" using integrable_ln_Gamma_aux[of n N] assms by (auto simp: I_def has_integral_integral divide_simps) thus "((\x. inverse (of_real x + s) ^ n * of_real (pbernpoly n x)) has_integral inverse (of_real (real N) + s) ^ n * (of_real (pbernpoly (Suc n) (real N)) / of_nat (Suc n)) - inverse (of_real 0 + s) ^ n * (of_real (pbernpoly (Suc n) 0) / of_nat (Suc n)) - - I) {0..real N}" by (simp_all add: field_simps) qed simp_all also have "(\x. - of_nat n * inverse (of_real x + s) ^ Suc n * (of_real (pbernpoly (Suc n) x) / of_nat (Suc n))) = (\x. - (of_nat n / of_nat (Suc n) * of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n))" by (simp add: divide_simps fun_eq_iff) finally have "((\x. - (of_nat n / of_nat (Suc n) * of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n)) has_integral - I) {0..real N}" . from has_integral_neg[OF this] show ?thesis by (auto simp add: I_def has_integral_iff algebra_simps integral_mult_right [symmetric] simp del: power_Suc of_nat_Suc ) qed lemma pbernpoly_over_power_tendsto_0: assumes "n > 0" shows "(\x. of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ 0" proof - from s have neq: "s + of_nat n \ 0" for n by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) - from bounded_pbernpoly[of "Suc n"] guess c . note c = this + obtain c where c: "\x. norm (pbernpoly (Suc n) x) \ c" + using bounded_pbernpoly by auto have "eventually (\x. real x + Re s > 0) at_top" by real_asymp hence "eventually (\x. norm (of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ (c / real (Suc n)) / (real x + Re s) ^ n) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim x) have "norm (of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ (c / real (Suc n)) / norm (s + of_nat x) ^ n" (is "_ \ ?rhs") using c[of x] by (auto simp: norm_divide norm_mult norm_power neq field_simps simp del: of_nat_Suc) also have "(real x + Re s) \ cmod (s + of_nat x)" using complex_Re_le_cmod[of "s + of_nat x"] s by (auto simp add: complex_nonpos_Reals_iff) hence "?rhs \ (c / real (Suc n)) / (real x + Re s) ^ n" using s elim c[of 0] neq[of x] by (intro divide_left_mono power_mono mult_pos_pos divide_nonneg_pos zero_less_power) auto finally show ?case . qed moreover have "(\x. (c / real (Suc n)) / (real x + Re s) ^ n) \ 0" using \n > 0\ by real_asymp ultimately show ?thesis by (rule Lim_null_comparison) qed lemma convergent_stirling_integral: assumes "n > 0" shows "convergent (\N. integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n))" (is "convergent (?f n)") proof - have "convergent (?f (Suc n))" for n proof (induction n) case 0 thus ?case using integral_pbernpoly_1 by (auto intro!: convergentI) next case (Suc n) have "convergent (\N. ?f (Suc n) N - of_real (pbernpoly (Suc (Suc n)) (real N)) / (of_nat (Suc (Suc n)) * (s + of_nat N) ^ Suc n) + of_real (bernoulli (Suc (Suc n)) / (real (Suc (Suc n)))) / s ^ Suc n)" (is "convergent ?g") by (intro convergent_add convergent_diff Suc convergent_const convergentI[OF pbernpoly_over_power_tendsto_0]) simp_all also have "?g = (\N. of_nat (Suc n) / of_nat (Suc (Suc n)) * ?f (Suc (Suc n)) N)" using s by (subst pbernpoly_integral_conv_pbernpoly_integral_Suc) (auto simp: fun_eq_iff field_simps simp del: of_nat_Suc power_Suc) also have "convergent \ \ convergent (?f (Suc (Suc n)))" by (intro convergent_mult_const_iff) (simp_all del: of_nat_Suc) finally show ?case . qed from this[of "n - 1"] assms show ?thesis by simp qed lemma stirling_integral_conv_stirling_integral_Suc: assumes "n > 0" shows "stirling_integral n s = of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" proof - have "(\N. of_real (pbernpoly (Suc n) (real N)) / (of_nat (Suc n) * (s + of_nat N) ^ n) - of_real (bernoulli (Suc n)) / (real (Suc n) * s ^ n) + integral {0..real N} (\x. of_nat n / of_nat (Suc n) * (of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n))) \ 0 - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n) + of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s" (is "?f \ _") unfolding stirling_integral_def integral_mult_right using convergent_stirling_integral[of "Suc n"] assms s by (intro tendsto_intros pbernpoly_over_power_tendsto_0) (auto simp: convergent_LIMSEQ_iff simp del: of_nat_Suc) also have "?this \ (\N. integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" using eventually_gt_at_top[of "0::nat"] pbernpoly_integral_conv_pbernpoly_integral_Suc[of n] assms unfolding integral_mult_right by (intro filterlim_cong refl) (auto elim!: eventually_mono simp del: power_Suc) finally show ?thesis unfolding stirling_integral_def[of n] by (rule limI) qed lemma stirling_integral_1_unfold: assumes "m > 0" shows "stirling_integral 1 s = stirling_integral m s / of_nat m - (\k=1..k=1..k = 1.. 0" shows "ln_Gamma s = (s - 1 / 2) * ln s - s + ln (2 * pi) / 2 + (\k=1..k = 1.. 0 \ (\x. integral {0..real x} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ stirling_integral n s" unfolding stirling_integral_def using convergent_stirling_integral[of n] by (simp only: convergent_LIMSEQ_iff) end lemmas has_integral_of_real = has_integral_linear[OF _ bounded_linear_of_real, unfolded o_def] lemmas integral_of_real = integral_linear[OF _ bounded_linear_of_real, unfolded o_def] lemma integrable_ln_Gamma_aux_real: assumes "0 < s" shows "(\x. pbernpoly n x / (x + s) ^ n) integrable_on {0..real N}" proof - have "(\x. complex_of_real (pbernpoly n x / (x + s) ^ n)) integrable_on {0..real N}" using integrable_ln_Gamma_aux[of "of_real s" n N] assms by simp from integrable_linear[OF this bounded_linear_Re] show ?thesis by (simp only: o_def Re_complex_of_real) qed lemma assumes "x > 0" "n > 0" shows stirling_integral_complex_of_real: "stirling_integral n (complex_of_real x) = of_real (stirling_integral n x)" and LIMSEQ_stirling_integral_real: "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ stirling_integral n x" and stirling_integral_real_convergent: "convergent (\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))" proof - have "(\N. integral {0..real N} (\t. of_real (pbernpoly n t / (t + x) ^ n))) \ stirling_integral n (complex_of_real x)" using LIMSEQ_stirling_integral[of "complex_of_real x" n] assms by simp hence "(\N. of_real (integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))) \ stirling_integral n (complex_of_real x)" using integrable_ln_Gamma_aux_real[OF assms(1), of n] by (subst (asm) integral_of_real) simp from tendsto_Re[OF this] have "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ Re (stirling_integral n (complex_of_real x))" by simp thus "convergent (\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))" by (rule convergentI) thus "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ stirling_integral n x" unfolding stirling_integral_def by (simp add: convergent_LIMSEQ_iff) from tendsto_of_real[OF this, where 'a = complex] integrable_ln_Gamma_aux_real[OF assms(1), of n] have "(\xa. integral {0..real xa} (\xa. complex_of_real (pbernpoly n xa) / (complex_of_real xa + x) ^ n)) \ complex_of_real (stirling_integral n x)" by (subst (asm) integral_of_real [symmetric]) simp_all from LIMSEQ_unique[OF this LIMSEQ_stirling_integral[of "complex_of_real x" n]] assms show "stirling_integral n (complex_of_real x) = of_real (stirling_integral n x)" by simp qed lemma ln_Gamma_stirling_real: assumes "x > (0 :: real)" "m > (0::nat)" shows "ln_Gamma x = (x - 1 / 2) * ln x - x + ln (2 * pi) / 2 + (\k = 1..k = 1.. (1::nat)" obtains c where "\s. Re s > 0 \ norm (stirling_integral n s) \ c / Re s ^ (n - 1)" proof - obtain c where c: "norm (pbernpoly n x) \ c" for x by (rule bounded_pbernpoly[of n]) blast have c': "pbernpoly n x \ c" for x using c[of x] by (simp add: abs_real_def split: if_splits) from c[of 0] have c_nonneg: "c \ 0" by simp have "norm (stirling_integral n s) \ c / (real n - 1) / Re s ^ (n - 1)" if s: "Re s > 0" for s proof (rule Lim_norm_ubound[OF _ LIMSEQ_stirling_integral]) have pos: "x + norm s > 0" if "x \ 0" for x using s that by (intro add_nonneg_pos) auto have nz: "of_real x + s \ 0" if "x \ 0" for x using s that by (auto simp: complex_eq_iff) let ?bound = "\N. c / (Re s ^ (n - 1) * (real n - 1)) - c / ((real N + Re s) ^ (n - 1) * (real n - 1))" show "eventually (\N. norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ c / (real n - 1) / Re s ^ (n - 1)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim N) let ?F = "\x. -c / ((x + Re s) ^ (n - 1) * (real n - 1))" from n s have "((\x. c / (x + Re s) ^ n) has_integral (?F (real N) - ?F 0)) {0..real N}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: divide_simps power_diff add_eq_0_iff2 has_field_derivative_iff_has_vector_derivative [symmetric]) also have "?F (real N) - ?F 0 = ?bound N" by simp finally have *: "((\x. c / (x + Re s) ^ n) has_integral ?bound N) {0..real N}" . have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c / (x + Re s) ^ n)" proof (intro integral_norm_bound_integral integrable_ln_Gamma_aux s ballI) fix x assume x: "x \ {0..real N}" have "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ c / norm (of_real x + s) ^ n" unfolding norm_divide norm_power using c by (intro divide_right_mono) simp_all also have "\ \ c / (x + Re s) ^ n" using x c c_nonneg s nz[of x] complex_Re_le_cmod[of "of_real x + s"] by (intro divide_left_mono power_mono mult_pos_pos zero_less_power add_nonneg_pos) auto finally show "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ \" . qed (insert n s * pos nz c, auto simp: complex_nonpos_Reals_iff) also have "\ = ?bound N" using * by (simp add: has_integral_iff) also have "\ \ c / (Re s ^ (n - 1) * (real n - 1))" using c_nonneg elim s n by simp also have "\ = c / (real n - 1) / (Re s ^ (n - 1))" by simp finally show "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ c / (real n - 1) / Re s ^ (n - 1)" . qed qed (insert s n, simp_all add: complex_nonpos_Reals_iff) thus ?thesis by (rule that) qed lemma stirling_integral_bound_aux_integral1: fixes a b c :: real and n :: nat assumes "a \ 0" "b > 0" "c \ 0" "n > 1" "l < a - b" "r > a + b" shows "((\x. c / max b \x - a\ ^ n) has_integral 2*c*(n / (n - 1))/b^(n-1) - c/(n-1) * (1/(a-l)^(n-1) + 1/(r-a)^(n-1))) {l..r}" proof - define x1 x2 where "x1 = a - b" and "x2 = a + b" define F1 where "F1 = (\x::real. c / (a - x) ^ (n - 1) / (n - 1))" define F3 where "F3 = (\x::real. -c / (x - a) ^ (n - 1) / (n - 1))" have deriv: "(F1 has_vector_derivative (c / (a - x) ^ n)) (at x within A)" "(F3 has_vector_derivative (c / (x - a) ^ n)) (at x within A)" if "x \ a" for x :: real and A unfolding F1_def F3_def using assms that by (auto intro!: derivative_eq_intros simp: divide_simps power_diff add_eq_0_iff2 simp flip: has_field_derivative_iff_has_vector_derivative) from assms have "((\x. c / (a - x) ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" by (intro fundamental_theorem_of_calculus deriv) (auto simp: x1_def max_def split: if_splits) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" using assms by (intro has_integral_spike_finite_eq[of "{l}"]) (auto simp: x1_def max_def split: if_splits) finally have I1: "((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" . have "((\x. c / b ^ n) has_integral (x2 - x1) * c / b ^ n) {x1..x2}" using has_integral_const_real[of "c / b ^ n" x1 x2] assms by (simp add: x1_def x2_def) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral ((x2 - x1) * c / b ^ n)) {x1..x2}" by (intro has_integral_spike_finite_eq[of "{x1, x2}"]) (auto simp: x1_def x2_def split: if_splits) finally have I2: "((\x. c / max b \x - a\ ^ n) has_integral ((x2 - x1) * c / b ^ n)) {x1..x2}" . from assms have I3: "((\x. c / (x - a) ^ n) has_integral (F3 r - F3 x2)) {x2..r}" by (intro fundamental_theorem_of_calculus deriv) (auto simp: x2_def min_def split: if_splits) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral (F3 r - F3 x2)) {x2..r}" using assms by (intro has_integral_spike_finite_eq[of "{r}"]) (auto simp: x2_def min_def split: if_splits) finally have I3: "((\x. c / max b \x - a\ ^ n) has_integral (F3 r - F3 x2)) {x2..r}" . have "((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l) + ((x2 - x1) * c / b ^ n) + (F3 r - F3 x2)) {l..r}" using assms by (intro has_integral_combine[OF _ _ has_integral_combine[OF _ _ I1 I2] I3]) (auto simp: x1_def x2_def) also have "(F1 x1 - F1 l) + ((x2 - x1) * c / b ^ n) + (F3 r - F3 x2) = F1 x1 - F1 l + F3 r - F3 x2 + (x2 - x1) * c / b ^ n" by (simp add: algebra_simps) also have "x2 - x1 = 2 * b" using assms by (simp add: x2_def x1_def min_def max_def) also have "2 * b * c / b ^ n = 2 * c / b ^ (n - 1)" using assms by (simp add: power_diff field_simps) also have "F1 x1 - F1 l + F3 r - F3 x2 = c/(n-1) * (2/b^(n-1) - 1/(a-l)^(n-1) - 1/(r-a)^(n-1))" using assms by (simp add: x1_def x2_def F1_def F3_def field_simps) also have "\ + 2 * c / b ^ (n - 1) = 2*c*(1 + 1/(n-1))/b^(n-1) - c/(n-1) * (1/(a-l)^(n-1) + 1/(r-a)^(n-1))" using assms by (simp add: field_simps) also have "1 + 1 / (n - 1) = n / (n - 1)" using assms by (simp add: field_simps) finally show ?thesis . qed lemma stirling_integral_bound_aux_integral2: fixes a b c :: real and n :: nat assumes "a \ 0" "b > 0" "c \ 0" "n > 1" obtains I where "((\x. c / max b \x - a\ ^ n) has_integral I) {l..r}" "I \ 2 * c * (n / (n - 1)) / b ^ (n-1)" proof - define l' where "l' = min l (a - b - 1)" define r' where "r' = max r (a + b + 1)" define A where "A = 2 * c * (n / (n - 1)) / b ^ (n - 1)" define B where "B = c / real (n - 1) * (1 / (a - l') ^ (n - 1) + 1 / (r' - a) ^ (n - 1))" have has_int: "((\x. c / max b \x - a\ ^ n) has_integral (A - B)) {l'..r'}" using assms unfolding A_def B_def by (intro stirling_integral_bound_aux_integral1) (auto simp: l'_def r'_def) have "(\x. c / max b \x - a\ ^ n) integrable_on {l..r}" by (rule integrable_on_subinterval[OF has_integral_integrable[OF has_int]]) (auto simp: l'_def r'_def) then obtain I where has_int': "((\x. c / max b \x - a\ ^ n) has_integral I) {l..r}" by (auto simp: integrable_on_def) from assms have "I \ A - B" by (intro has_integral_subset_le[OF _ has_int' has_int]) (auto simp: l'_def r'_def) also have "\ \ A" using assms by (simp add: B_def l'_def r'_def) finally show ?thesis using that[of I] has_int' unfolding A_def by blast qed lemma stirling_integral_bound_aux': assumes n: "n > (1::nat)" and \: "\ \ {0<..s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral n s) \ c / norm s ^ (n - 1)" proof - obtain c where c: "norm (pbernpoly n x) \ c" for x by (rule bounded_pbernpoly[of n]) blast have c': "pbernpoly n x \ c" for x using c[of x] by (simp add: abs_real_def split: if_splits) from c[of 0] have c_nonneg: "c \ 0" by simp define D where "D = c * Beta (- (real_of_int (- int n) / 2) - 1 / 2) (1 / 2) / 2" define C where "C = max D (2*c*(n/(n-1))/sin \^(n-1))" have *: "norm (stirling_integral n s) \ C / norm s ^ (n - 1)" if s: "s \ complex_cone' \ - {0}" for s :: complex proof (rule Lim_norm_ubound[OF _ LIMSEQ_stirling_integral]) from s \ have Arg: "\Arg s\ \ \" by (auto simp: complex_cone_altdef) have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] \ s by auto from s have [simp]: "s \ 0" by auto show "eventually (\N. norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ C / norm s ^ (n - 1)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim N) show ?case proof (cases "Re s > 0") case True have int: "((\x. c * (x^2 + norm s^2) powr (-n / 2)) has_integral D * (norm s ^ 2) powr (-n / 2 + 1 / 2)) {0<..}" using has_integral_mult_left[OF has_integral_Beta3[of "-n/2" "norm s ^ 2"], of c] assms True unfolding D_def by (simp add: algebra_simps) hence int': "((\x. c * (x^2 + norm s^2) powr (-n / 2)) has_integral D * (norm s ^ 2) powr (-n / 2 + 1 / 2)) {0..}" by (subst has_integral_interior [symmetric]) simp_all hence integrable: "(\x. c * (x^2 + norm s^2) powr (-n / 2)) integrable_on {0..}" by (simp add: has_integral_iff) have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c * (x^2 + norm s^2) powr (-n / 2))" proof (intro integral_norm_bound_integral s ballI integrable_ln_Gamma_aux) have [simp]: "{0<..} - {0::real..} = {}" "{0..} - {0<..} = {0::real}" by auto have "(\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0<..}" using int by (simp add: has_integral_iff) also have "?this \ (\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0..}" by (intro integrable_spike_set_eq) auto finally show "(\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0..real N}" by (rule integrable_on_subinterval) auto next fix x assume x: "x \ {0..real N}" have nz: "complex_of_real x + s \ 0" using True x by (auto simp: complex_eq_iff) have "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ c / norm (of_real x + s) ^ n" unfolding norm_divide norm_power using c by (intro divide_right_mono) simp_all also have "\ \ c / sqrt (x ^ 2 + norm s ^ 2) ^ n" proof (intro divide_left_mono mult_pos_pos zero_less_power power_mono) show "sqrt (x\<^sup>2 + (cmod s)\<^sup>2) \ cmod (complex_of_real x + s)" using x True by (simp add: cmod_def algebra_simps power2_eq_square) qed (use x True c_nonneg assms nz in \auto simp: add_nonneg_pos\) also have "sqrt (x ^ 2 + norm s ^ 2) ^ n = (x ^ 2 + norm s ^ 2) powr (1/2 * n)" by (subst powr_powr [symmetric], subst powr_realpow) (auto simp: powr_half_sqrt add_nonneg_pos) also have "c / \ = c * (x^2 + norm s^2) powr (-n / 2)" by (simp add: powr_minus field_simps) finally show "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ \" . qed fact+ also have "\ \ integral {0..} (\x. c * (x^2 + norm s^2) powr (-n / 2))" using c_nonneg by (intro integral_subset_le integrable integrable_on_subinterval[OF integrable]) auto also have "\ = D * (norm s ^ 2) powr (-n / 2 + 1 / 2)" using int' by (simp add: has_integral_iff) also have "(norm s ^ 2) powr (-n / 2 + 1 / 2) = norm s powr (2 * (-n / 2 + 1 / 2))" by (subst powr_powr [symmetric]) auto also have "\ = norm s powr (-real (n - 1))" using assms by (simp add: of_nat_diff) also have "D * \ = D / norm s ^ (n - 1)" by (auto simp: powr_minus powr_realpow field_simps) also have "\ \ C / norm s ^ (n - 1)" by (intro divide_right_mono) (auto simp: C_def) finally show "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ \" . next case False have "cos \Arg s\ = cos (Arg s)" by (simp add: abs_if) also have "cos (Arg s) = Re (rcis (norm s) (Arg s)) / norm s" by (subst Re_rcis) auto also have "\ = Re s / norm s" by (subst rcis_cmod_Arg) auto also have "\ \ cos (pi / 2)" using False by (auto simp: field_simps) finally have "\Arg s\ \ pi / 2" using Arg \ by (subst (asm) cos_mono_le_eq) auto have "sin \ * norm s = sin (pi - \) * norm s" by simp also have "\ \ sin (pi - \Arg s\) * norm s" using \ Arg \\Arg s\ \ pi / 2\ by (intro mult_right_mono sin_monotone_2pi_le) auto also have "sin \Arg s\ \ 0" using Arg_bounded[of s] by (intro sin_ge_zero) auto hence "sin (pi - \Arg s\) = \sin \Arg s\\" by simp also have "\ = \sin (Arg s)\" by (simp add: abs_if) also have "\ * norm s = \Im (rcis (norm s) (Arg s))\" by (simp add: abs_mult) also have "\ = \Im s\" by (subst rcis_cmod_Arg) auto finally have abs_Im_ge: "\Im s\ \ sin \ * norm s" . have [simp]: "Im s \ 0" "s \ 0" using s \s \ \\<^sub>\\<^sub>0\ False by (auto simp: cmod_def zero_le_mult_iff complex_nonpos_Reals_iff) have "sin \ > 0" using assms by (intro sin_gt_zero) auto obtain I where I: "((\x. c / max \Im s\ \x + Re s\ ^ n) has_integral I) {0..real N}" "I \ 2*c*(n/(n-1)) / \Im s\ ^ (n - 1)" using s c_nonneg assms False stirling_integral_bound_aux_integral2[of "-Re s" "\Im s\" c n 0 "real N"] by auto have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c / max \Im s\ \x + Re s\ ^ n)" proof (intro integral_norm_bound_integral integrable_ln_Gamma_aux s ballI) show "(\x. c / max \Im s\ \x + Re s\ ^ n) integrable_on {0..real N}" using I(1) by (simp add: has_integral_iff) next fix x assume x: "x \ {0..real N}" have nz: "complex_of_real x + s \ 0" by (auto simp: complex_eq_iff) have "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ c / norm (complex_of_real x + s) ^ n" unfolding norm_divide norm_power using c[of x] by (intro divide_right_mono) simp_all also have "\ \ c / max \Im s\ \x + Re s\ ^ n" using c_nonneg nz abs_Re_le_cmod[of "of_real x + s"] abs_Im_le_cmod[of "of_real x + s"] by (intro divide_left_mono power_mono mult_pos_pos zero_less_power) (auto simp: less_max_iff_disj) finally show "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ \" . qed (auto simp: complex_nonpos_Reals_iff) also have "\ \ 2*c*(n/(n-1)) / \Im s\ ^ (n - 1)" using I by (simp add: has_integral_iff) also have "\ \ 2*c*(n/(n-1)) / (sin \ * norm s) ^ (n - 1)" using \sin \ > 0\ s c_nonneg abs_Im_ge by (intro divide_left_mono mult_pos_pos zero_less_power power_mono mult_nonneg_nonneg) auto also have "\ = 2*c*(n/(n-1))/sin \^(n-1) / norm s ^ (n - 1)" by (simp add: field_simps) also have "\ \ C / norm s ^ (n - 1)" by (intro divide_right_mono) (auto simp: C_def) finally show ?thesis . qed qed qed (use that assms complex_cone_inter_nonpos_Reals[of "-\" \] \ in auto) thus ?thesis by (rule that) qed lemma stirling_integral_bound: assumes "n > 0" obtains c where "\s. Re s > 0 \ norm (stirling_integral n s) \ c / Re s ^ n" proof - let ?f = "\s. of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" from stirling_integral_bound_aux[of "Suc n"] assms obtain c where c: "\s. Re s > 0 \ norm (stirling_integral (Suc n) s) \ c / Re s ^ n" by auto define c1 where "c1 = real n / real (Suc n) * c" define c2 where "c2 = \bernoulli (Suc n)\ / real (Suc n)" have c2_nonneg: "c2 \ 0" by (simp add: c2_def) show ?thesis proof (rule that) fix s :: complex assume s: "Re s > 0" hence s': "s \ \\<^sub>\\<^sub>0" by (auto simp: complex_nonpos_Reals_iff) have "stirling_integral n s = ?f s" using s' assms by (rule stirling_integral_conv_stirling_integral_Suc) also have "norm \ \ norm (of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s) + norm (of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n))" by (rule norm_triangle_ineq4) also have "\ = real n / real (Suc n) * norm (stirling_integral (Suc n) s) + c2 / norm s ^ n" (is "_ = ?A + ?B") by (simp add: norm_divide norm_mult norm_power c2_def field_simps del: of_nat_Suc) also have "?A \ real n / real (Suc n) * (c / Re s ^ n)" by (intro mult_left_mono c s) simp_all also have "\ = c1 / Re s ^ n" by (simp add: c1_def) also have "c2 / norm s ^ n \ c2 / Re s ^ n" using s c2_nonneg by (intro divide_left_mono power_mono complex_Re_le_cmod mult_pos_pos zero_less_power) auto also have "c1 / Re s ^ n + c2 / Re s ^ n = (c1 + c2) / Re s ^ n" using s by (simp add: field_simps) finally show "norm (stirling_integral n s) \ (c1 + c2) / Re s ^ n" by - simp_all qed qed lemma stirling_integral_bound': assumes "n > 0" and "\ \ {0<..s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral n s) \ c / norm s ^ n" proof - let ?f = "\s. of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" from stirling_integral_bound_aux'[of "Suc n"] assms obtain c where c: "\s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral (Suc n) s) \ c / norm s ^ n" by auto define c1 where "c1 = real n / real (Suc n) * c" define c2 where "c2 = \bernoulli (Suc n)\ / real (Suc n)" have c2_nonneg: "c2 \ 0" by (simp add: c2_def) show ?thesis proof (rule that) fix s :: complex assume s: "s \ complex_cone' \ - {0}" have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] assms s by auto have "stirling_integral n s = ?f s" using s' assms by (intro stirling_integral_conv_stirling_integral_Suc) auto also have "norm \ \ norm (of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s) + norm (of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n))" by (rule norm_triangle_ineq4) also have "\ = real n / real (Suc n) * norm (stirling_integral (Suc n) s) + c2 / norm s ^ n" (is "_ = ?A + ?B") by (simp add: norm_divide norm_mult norm_power c2_def field_simps del: of_nat_Suc) also have "?A \ real n / real (Suc n) * (c / norm s ^ n)" by (intro mult_left_mono c s) simp_all also have "\ = c1 / norm s ^ n" by (simp add: c1_def) also have "c1 / norm s ^ n + c2 / norm s ^ n = (c1 + c2) / norm s ^ n" using s by (simp add: divide_simps) finally show "norm (stirling_integral n s) \ (c1 + c2) / norm s ^ n" by - simp_all qed qed lemma stirling_integral_holomorphic [holomorphic_intros]: assumes m: "m > 0" and "A \ \\<^sub>\\<^sub>0 = {}" shows "stirling_integral m holomorphic_on A" proof - from assms have [simp]: "z \ \\<^sub>\\<^sub>0" if "z \ A" for z using that by auto let ?f = "\s::complex. of_nat m * ((s - 1 / 2) * Ln s - s + of_real (ln (2 * pi) / 2) + (\k=1.. stirling_integral m holomorphic_on A" using assms by (intro holomorphic_cong refl) (simp_all add: field_simps ln_Gamma_stirling_complex) finally show "stirling_integral m holomorphic_on A" . qed lemma stirling_integral_continuous_on_complex [continuous_intros]: assumes m: "m > 0" and "A \ \\<^sub>\\<^sub>0 = {}" shows "continuous_on A (stirling_integral m :: _ \ complex)" by (intro holomorphic_on_imp_continuous_on stirling_integral_holomorphic assms) lemma has_field_derivative_stirling_integral_complex: fixes x :: complex assumes "x \ \\<^sub>\\<^sub>0" "n > 0" shows "(stirling_integral n has_field_derivative deriv (stirling_integral n) x) (at x)" using assms by (intro holomorphic_derivI[OF stirling_integral_holomorphic, of n "-\\<^sub>\\<^sub>0"]) auto lemma assumes n: "n > 0" and "x > 0" shows deriv_stirling_integral_complex_of_real: "(deriv ^^ j) (stirling_integral n) (complex_of_real x) = complex_of_real ((deriv ^^ j) (stirling_integral n) x)" (is "?lhs x = ?rhs x") and differentiable_stirling_integral_real: "(deriv ^^ j) (stirling_integral n) field_differentiable at x" (is ?thesis2) proof - let ?A = "{s. Re s > 0}" let ?f = "\j x. (deriv ^^ j) (stirling_integral n) (complex_of_real x)" let ?f' = "\j x. complex_of_real ((deriv ^^ j) (stirling_integral n) x)" have [simp]: "open ?A" by (simp add: open_halfspace_Re_gt) have "?lhs x = ?rhs x \ (deriv ^^ j) (stirling_integral n) field_differentiable at x" if "x > 0" for x using that proof (induction j arbitrary: x) case 0 have "((\x. Re (stirling_integral n (of_real x))) has_field_derivative Re (deriv (\x. stirling_integral n x) (of_real x))) (at x)" using 0 n by (auto intro!: derivative_intros has_vector_derivative_real_field field_differentiable_derivI holomorphic_on_imp_differentiable_at[of _ ?A] stirling_integral_holomorphic simp: complex_nonpos_Reals_iff) also have "?this \ (stirling_integral n has_field_derivative Re (deriv (\x. stirling_integral n x) (of_real x))) (at x)" using eventually_nhds_in_open[of "{0<..}" x] 0 n by (intro has_field_derivative_cong_ev refl) (auto elim!: eventually_mono simp: stirling_integral_complex_of_real) finally have "stirling_integral n field_differentiable at x" by (auto simp: field_differentiable_def) with 0 n show ?case by (auto simp: stirling_integral_complex_of_real) next case (Suc j x) note IH = conjunct1[OF Suc.IH] conjunct2[OF Suc.IH] have *: "(deriv ^^ Suc j) (stirling_integral n) (complex_of_real x) = of_real ((deriv ^^ Suc j) (stirling_integral n) x)" if x: "x > 0" for x proof - have "deriv ((deriv ^^ j) (stirling_integral n)) (complex_of_real x) = vector_derivative (\x. (deriv ^^ j) (stirling_integral n) (of_real x)) (at x)" using n x by (intro vector_derivative_of_real_right [symmetric] holomorphic_on_imp_differentiable_at[of _ ?A] holomorphic_higher_deriv stirling_integral_holomorphic) (auto simp: complex_nonpos_Reals_iff) also have "\ = vector_derivative (\x. of_real ((deriv ^^ j) (stirling_integral n) x)) (at x)" using eventually_nhds_in_open[of "{0<..}" x] x by (intro vector_derivative_cong_eq) (auto elim!: eventually_mono simp: IH(1)) also have "\ = of_real (deriv ((deriv ^^ j) (stirling_integral n)) x)" by (intro vector_derivative_of_real_left holomorphic_on_imp_differentiable_at[of _ ?A] field_differentiable_imp_differentiable IH(2) x) finally show ?thesis by simp qed have "((\x. Re ((deriv ^^ Suc j) (stirling_integral n) (of_real x))) has_field_derivative Re (deriv ((deriv ^^ Suc j) (stirling_integral n)) (of_real x))) (at x)" using Suc.prems n by (intro derivative_intros has_vector_derivative_real_field field_differentiable_derivI holomorphic_on_imp_differentiable_at[of _ ?A] stirling_integral_holomorphic holomorphic_higher_deriv) (auto simp: complex_nonpos_Reals_iff) also have "?this \ ((deriv ^^ Suc j) (stirling_integral n) has_field_derivative Re (deriv ((deriv ^^ Suc j) (stirling_integral n)) (of_real x))) (at x)" using eventually_nhds_in_open[of "{0<..}" x] Suc.prems * by (intro has_field_derivative_cong_ev refl) (auto elim!: eventually_mono) finally have "(deriv ^^ Suc j) (stirling_integral n) field_differentiable at x" by (auto simp: field_differentiable_def) with *[OF Suc.prems] show ?case by blast qed from this[OF assms(2)] show "?lhs x = ?rhs x" ?thesis2 by blast+ qed text \ Unfortunately, asymptotic power series cannot, in general, be differentiated. However, since @{term ln_Gamma} is holomorphic on the entire positive real half-space, we can differentiate its asymptotic expansion after all. To do this, we use an ad-hoc version of the more general approach outlined in Erdelyi's ``Asymptotic Expansions'' for holomorphic functions: We bound the value of the $j$-th derivative of the remainder term at some value $x$ by applying Cauchy's integral formula along a circle centred at $x$ with radius $\frac{1}{2} x$. \ lemma deriv_stirling_integral_real_bound: assumes m: "m > 0" shows "(deriv ^^ j) (stirling_integral m) \ O(\x::real. 1 / x ^ (m + j))" proof - - from stirling_integral_bound[OF m] guess c . note c = this + obtain c where c: "\s. 0 < Re s \ cmod (stirling_integral m s) \ c / Re s ^ m" + using stirling_integral_bound[OF m] by auto have "0 \ cmod (stirling_integral m 1)" by simp also have "\ \ c" using c[of 1] by simp finally have c_nonneg: "c \ 0" . define B where "B = c * 2 ^ (m + Suc j)" define B' where "B' = B * fact j / 2" have "eventually (\x::real. norm ((deriv ^^ j) (stirling_integral m) x) \ B' * norm (1 / x ^ (m+ j))) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim case (elim x) have "s \ \\<^sub>\\<^sub>0" if "s \ cball (of_real x) (x/2)" for s :: complex proof - have "x - Re s \ norm (of_real x - s)" using complex_Re_le_cmod[of "of_real x - s"] by simp also from that have "\ \ x/2" by (simp add: dist_complex_def) finally show ?thesis using elim by (auto simp: complex_nonpos_Reals_iff) qed hence "((\u. stirling_integral m u / (u - of_real x) ^ Suc j) has_contour_integral complex_of_real (2 * pi) * \ / fact j * (deriv ^^ j) (stirling_integral m) (of_real x)) (circlepath (of_real x) (x/2))" using m elim by (intro Cauchy_has_contour_integral_higher_derivative_circlepath stirling_integral_continuous_on_complex stirling_integral_holomorphic) auto hence "norm (of_real (2 * pi) * \ / fact j * (deriv ^^ j) (stirling_integral m) (of_real x)) \ B / x ^ (m + Suc j) * (2 * pi * (x / 2))" proof (rule has_contour_integral_bound_circlepath) fix u :: complex assume dist: "norm (u - of_real x) = x / 2" have "Re (of_real x - u) \ norm (of_real x - u)" by (rule complex_Re_le_cmod) also have "\ = x / 2" using dist by (simp add: norm_minus_commute) finally have Re_u: "Re u \ x/2" using elim by simp have "norm (stirling_integral m u / (u - of_real x) ^ Suc j) \ c / Re u ^ m / (x / 2) ^ Suc j" using Re_u elim unfolding norm_divide norm_power dist by (intro divide_right_mono zero_le_power c) simp_all also have "\ \ c / (x/2) ^ m / (x / 2) ^ Suc j" using c_nonneg elim Re_u by (intro divide_right_mono divide_left_mono power_mono) simp_all also have "\ = B / x ^ (m + Suc j)" using elim by (simp add: B_def field_simps power_add) finally show "norm (stirling_integral m u / (u - of_real x) ^ Suc j) \ B / x ^ (m + Suc j)" . qed (insert elim c_nonneg, auto simp: B_def simp del: power_Suc) hence "cmod ((deriv ^^ j) (stirling_integral m) (of_real x)) \ B' / x ^ (j + m)" using elim by (simp add: field_simps norm_divide norm_mult norm_power B'_def) with elim m show ?case by (simp_all add: add_ac deriv_stirling_integral_complex_of_real) qed thus ?thesis by (rule bigoI) qed definition stirling_sum where "stirling_sum j m x = (-1) ^ j * (\k = 1..k\m. (of_real (bernoulli' k) * pochhammer (of_nat (Suc k)) (j - 1) * inverse x ^ (k + j)))" lemma stirling_sum_complex_of_real: "stirling_sum j m (complex_of_real x) = complex_of_real (stirling_sum j m x)" by (simp add: stirling_sum_def pochhammer_of_real [symmetric] del: of_nat_Suc) lemma stirling_sum'_complex_of_real: "stirling_sum' j m (complex_of_real x) = complex_of_real (stirling_sum' j m x)" by (simp add: stirling_sum'_def pochhammer_of_real [symmetric] del: of_nat_Suc) lemma has_field_derivative_stirling_sum_complex [derivative_intros]: "Re x > 0 \ (stirling_sum j m has_field_derivative stirling_sum (Suc j) m x) (at x)" unfolding stirling_sum_def [abs_def] sum_distrib_left by (rule DERIV_sum) (auto intro!: derivative_eq_intros simp del: of_nat_Suc simp: pochhammer_Suc power_diff) lemma has_field_derivative_stirling_sum_real [derivative_intros]: "x > (0::real) \ (stirling_sum j m has_field_derivative stirling_sum (Suc j) m x) (at x)" unfolding stirling_sum_def [abs_def] sum_distrib_left by (rule DERIV_sum) (auto intro!: derivative_eq_intros simp del: of_nat_Suc simp: pochhammer_Suc power_diff) lemma has_field_derivative_stirling_sum'_complex [derivative_intros]: assumes "j > 0" "Re x > 0" shows "(stirling_sum' j m has_field_derivative stirling_sum' (Suc j) m x) (at x)" proof (cases j) case (Suc j') from assms have [simp]: "x \ 0" by auto define c where "c = (\n. (-1) ^ Suc j * complex_of_real (bernoulli' n) * pochhammer (of_nat (Suc n)) j')" define T where "T = (\n x. c n * inverse x ^ (j + n))" define T' where "T' = (\n x. - (of_nat (j + n)) * c n * inverse x ^ (Suc (j + n)))" have "((\x. \k\m. T k x) has_field_derivative (\k\m. T' k x)) (at x)" using assms Suc by (intro DERIV_sum) (auto simp: T_def T'_def intro!: derivative_eq_intros simp: field_simps power_add [symmetric] simp del: of_nat_Suc power_Suc of_nat_add) also have "(\x. (\k\m. T k x)) = stirling_sum' j m" by (simp add: Suc T_def c_def stirling_sum'_def fun_eq_iff add_ac mult.assoc sum_distrib_left) also have "(\k\m. T' k x) = stirling_sum' (Suc j) m x" by (simp add: T'_def c_def Suc stirling_sum'_def sum_distrib_left sum_distrib_right algebra_simps pochhammer_Suc) finally show ?thesis . qed (insert assms, simp_all) lemma has_field_derivative_stirling_sum'_real [derivative_intros]: assumes "j > 0" "x > (0::real)" shows "(stirling_sum' j m has_field_derivative stirling_sum' (Suc j) m x) (at x)" proof (cases j) case (Suc j') from assms have [simp]: "x \ 0" by auto define c where "c = (\n. (-1) ^ Suc j * (bernoulli' n) * pochhammer (of_nat (Suc n)) j')" define T where "T = (\n x. c n * inverse x ^ (j + n))" define T' where "T' = (\n x. - (of_nat (j + n)) * c n * inverse x ^ (Suc (j + n)))" have "((\x. \k\m. T k x) has_field_derivative (\k\m. T' k x)) (at x)" using assms Suc by (intro DERIV_sum) (auto simp: T_def T'_def intro!: derivative_eq_intros simp: field_simps power_add [symmetric] simp del: of_nat_Suc power_Suc of_nat_add) also have "(\x. (\k\m. T k x)) = stirling_sum' j m" by (simp add: Suc T_def c_def stirling_sum'_def fun_eq_iff add_ac mult.assoc sum_distrib_left) also have "(\k\m. T' k x) = stirling_sum' (Suc j) m x" by (simp add: T'_def c_def Suc stirling_sum'_def sum_distrib_left sum_distrib_right algebra_simps pochhammer_Suc) finally show ?thesis . qed (insert assms, simp_all) lemma higher_deriv_stirling_sum_complex: "Re x > 0 \ (deriv ^^ i) (stirling_sum j m) x = stirling_sum (i + j) m x" proof (induction i arbitrary: x) case (Suc i) have "deriv ((deriv ^^ i) (stirling_sum j m)) x = deriv (stirling_sum (i + j) m) x" using eventually_nhds_in_open[of "{x. Re x > 0}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_halfspace_Re_gt Suc.IH) also from Suc.prems have "\ = stirling_sum (Suc (i + j)) m x" by (intro DERIV_imp_deriv has_field_derivative_stirling_sum_complex) finally show ?case by simp qed simp_all definition Polygamma_approx :: "nat \ nat \ 'a \ 'a :: {real_normed_field, ln}" where "Polygamma_approx j m = (deriv ^^ j) (\x. (x - 1 / 2) * ln x - x + of_real (ln (2 * pi)) / 2 + stirling_sum 0 m x)" lemma Polygamma_approx_Suc: "Polygamma_approx (Suc j) m = deriv (Polygamma_approx j m)" by (simp add: Polygamma_approx_def) lemma Polygamma_approx_0: "Polygamma_approx 0 m x = (x - 1/2) * ln x - x + of_real (ln (2*pi)) / 2 + stirling_sum 0 m x" by (simp add: Polygamma_approx_def) lemma Polygamma_approx_1_complex: "Re x > 0 \ Polygamma_approx (Suc 0) m x = ln x - 1 / (2*x) + stirling_sum (Suc 0) m x" unfolding Polygamma_approx_Suc Polygamma_approx_0 by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps) lemma Polygamma_approx_1_real: "x > (0 :: real) \ Polygamma_approx (Suc 0) m x = ln x - 1 / (2*x) + stirling_sum (Suc 0) m x" unfolding Polygamma_approx_Suc Polygamma_approx_0 by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps) lemma stirling_sum_2_conv_stirling_sum'_1: fixes x :: "'a :: {real_div_algebra, field_char_0}" assumes "m > 0" "x \ 0" shows "stirling_sum' 1 m x = 1 / x + 1 / (2 * x^2) + stirling_sum 2 m x" proof - have pochhammer_2: "pochhammer (of_nat k) 2 = of_nat k * of_nat (Suc k)" for k by (simp add: pochhammer_Suc eval_nat_numeral add_ac) have "stirling_sum 2 m x = (\k = Suc 0.. = (\k=0.. = (\k=0.. = (\k\m. of_real (bernoulli' k) * inverse x ^ Suc k)" by (intro sum.cong) auto also have "\ = stirling_sum' 1 m x" by (simp add: stirling_sum'_def) finally show ?thesis by (simp add: add_ac) qed lemma Polygamma_approx_2_real: assumes "x > (0::real)" "m > 0" shows "Polygamma_approx (Suc (Suc 0)) m x = stirling_sum' 1 m x" proof - have "Polygamma_approx (Suc (Suc 0)) m x = deriv (Polygamma_approx (Suc 0) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (\x. ln x - 1 / (2*x) + stirling_sum (Suc 0) m x) x" using eventually_nhds_in_open[of "{0<..}" x] assms by (intro deriv_cong_ev) (auto elim!: eventually_mono simp: Polygamma_approx_1_real) also have "\ = 1 / x + 1 / (2*x^2) + stirling_sum (Suc (Suc 0)) m x" using assms by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps power2_eq_square) also have "\ = stirling_sum' 1 m x" using stirling_sum_2_conv_stirling_sum'_1[of m x] assms by (simp add: eval_nat_numeral) finally show ?thesis . qed lemma Polygamma_approx_2_complex: assumes "Re x > 0" "m > 0" shows "Polygamma_approx (Suc (Suc 0)) m x = stirling_sum' 1 m x" proof - have "Polygamma_approx (Suc (Suc 0)) m x = deriv (Polygamma_approx (Suc 0) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (\x. ln x - 1 / (2*x) + stirling_sum (Suc 0) m x) x" using eventually_nhds_in_open[of "{s. Re s > 0}" x] assms by (intro deriv_cong_ev) (auto simp: open_halfspace_Re_gt elim!: eventually_mono simp: Polygamma_approx_1_complex) also have "\ = 1 / x + 1 / (2*x^2) + stirling_sum (Suc (Suc 0)) m x" using assms by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps power2_eq_square) also have "\ = stirling_sum' 1 m x" using stirling_sum_2_conv_stirling_sum'_1[of m x] assms by (subst stirling_sum_2_conv_stirling_sum'_1) (auto simp: eval_nat_numeral) finally show ?thesis . qed lemma Polygamma_approx_ge_2_real: assumes "x > (0::real)" "m > 0" shows "Polygamma_approx (Suc (Suc j)) m x = stirling_sum' (Suc j) m x" using assms(1) proof (induction j arbitrary: x) case (0 x) with assms show ?case by (simp add: Polygamma_approx_2_real) next case (Suc j x) have "Polygamma_approx (Suc (Suc (Suc j))) m x = deriv (Polygamma_approx (Suc (Suc j)) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (stirling_sum' (Suc j) m) x" using eventually_nhds_in_open[of "{0<..}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH) also have "\ = stirling_sum' (Suc (Suc j)) m x" using Suc.prems by (intro DERIV_imp_deriv derivative_intros) simp_all finally show ?case . qed lemma Polygamma_approx_ge_2_complex: assumes "Re x > 0" "m > 0" shows "Polygamma_approx (Suc (Suc j)) m x = stirling_sum' (Suc j) m x" using assms(1) proof (induction j arbitrary: x) case (0 x) with assms show ?case by (simp add: Polygamma_approx_2_complex) next case (Suc j x) have "Polygamma_approx (Suc (Suc (Suc j))) m x = deriv (Polygamma_approx (Suc (Suc j)) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (stirling_sum' (Suc j) m) x" using eventually_nhds_in_open[of "{x. Re x > 0}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH open_halfspace_Re_gt) also have "\ = stirling_sum' (Suc (Suc j)) m x" using Suc.prems by (intro DERIV_imp_deriv derivative_intros) simp_all finally show ?case . qed lemma Polygamma_approx_complex_of_real: assumes "x > 0" "m > 0" shows "Polygamma_approx j m (complex_of_real x) = of_real (Polygamma_approx j m x)" proof (cases j) case 0 with assms show ?thesis by (simp add: Polygamma_approx_0 Ln_of_real stirling_sum_complex_of_real) next case [simp]: (Suc j') thus ?thesis proof (cases j') case 0 with assms show ?thesis by (simp add: Polygamma_approx_1_complex Polygamma_approx_1_real stirling_sum_complex_of_real Ln_of_real) next case (Suc j'') with assms show ?thesis by (simp add: Polygamma_approx_ge_2_complex Polygamma_approx_ge_2_real stirling_sum'_complex_of_real) qed qed lemma higher_deriv_Polygamma_approx [simp]: "(deriv ^^ j) (Polygamma_approx i m) = Polygamma_approx (j + i) m" by (simp add: Polygamma_approx_def funpow_add) lemma stirling_sum_holomorphic [holomorphic_intros]: "0 \ A \ stirling_sum j m holomorphic_on A" unfolding stirling_sum_def by (intro holomorphic_intros) auto lemma Polygamma_approx_holomorphic [holomorphic_intros]: "Polygamma_approx j m holomorphic_on {s. Re s > 0}" unfolding Polygamma_approx_def by (intro holomorphic_intros) (auto simp: open_halfspace_Re_gt elim!: nonpos_Reals_cases) lemma higher_deriv_lnGamma_stirling: assumes m: "m > 0" shows "(\x::real. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x) \ O(\x. 1 / x ^ (m + j))" proof - have "eventually (\x. \(deriv ^^ j) ln_Gamma x - Polygamma_approx j m x\ = inverse (real m) * \(deriv ^^ j) (stirling_integral m) x\) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim case (elim x) note x = this have "\\<^sub>F y in nhds (complex_of_real x). y \ - \\<^sub>\\<^sub>0" using elim by (intro eventually_nhds_in_open) auto hence "(deriv ^^ j) (\x. ln_Gamma x - Polygamma_approx 0 m x) (complex_of_real x) = (deriv ^^ j) (\x. (-inverse (of_nat m)) * stirling_integral m x) (complex_of_real x)" using x m by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: ln_Gamma_stirling_complex Polygamma_approx_def field_simps open_halfspace_Re_gt stirling_sum_def) also have "\ = - inverse (of_nat m) * (deriv ^^ j) (stirling_integral m) (of_real x)" using x m by (intro higher_deriv_cmult[of _ "-\\<^sub>\\<^sub>0"] stirling_integral_holomorphic) (auto simp: open_halfspace_Re_gt) also have "(deriv ^^ j) (\x. ln_Gamma x - Polygamma_approx 0 m x) (complex_of_real x) = (deriv ^^ j) ln_Gamma (of_real x) - (deriv ^^ j) (Polygamma_approx 0 m) (of_real x)" using x by (intro higher_deriv_diff[of _ "{s. Re s > 0}"]) (auto intro!: holomorphic_intros elim!: nonpos_Reals_cases simp: open_halfspace_Re_gt) also have "(deriv ^^ j) (Polygamma_approx 0 m) (complex_of_real x) = of_real (Polygamma_approx j m x)" using x m by (simp add: Polygamma_approx_complex_of_real) also have "norm (- inverse (of_nat m) * (deriv ^^ j) (stirling_integral m) (complex_of_real x)) = inverse (real m) * \(deriv ^^ j) (stirling_integral m) x\" using x m by (simp add: norm_mult norm_inverse deriv_stirling_integral_complex_of_real) also have "(deriv ^^ j) ln_Gamma (complex_of_real x) = of_real ((deriv ^^ j) ln_Gamma x)" using x by (simp add: higher_deriv_ln_Gamma_complex_of_real) also have "norm (\ - of_real (Polygamma_approx j m x)) = \(deriv ^^ j) ln_Gamma x - Polygamma_approx j m x\" by (simp only: of_real_diff [symmetric] norm_of_real) finally show ?case . qed from bigthetaI_cong[OF this] m have "(\x::real. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x) \ \(\x. (deriv ^^ j) (stirling_integral m) x)" by simp also have "(\x::real. (deriv ^^ j) (stirling_integral m) x) \ O(\x. 1 / x ^ (m + j))" using m by (rule deriv_stirling_integral_real_bound) finally show ?thesis . qed lemma Polygamma_approx_1_real': assumes x: "(x::real) > 0" and m: "m > 0" shows "Polygamma_approx 1 m x = ln x - (\k = Suc 0..m. bernoulli' k * inverse x ^ k / real k)" proof - have "Polygamma_approx 1 m x = ln x - (1 / (2 * x) + (\k=Suc 0..k=Suc 0.. = (\k=0.. = (\k = Suc 0..m. bernoulli' k * inverse x ^ k / real k)" using assms by (subst sum.shift_bounds_Suc_ivl [symmetric]) (simp add: atLeastLessThanSuc_atLeastAtMost) finally show ?thesis . qed theorem assumes m: "m > 0" shows ln_Gamma_real_asymptotics: "(\x. ln_Gamma x - ((x - 1 / 2) * ln x - x + ln (2 * pi) / 2 + (\k = 1.. O(\x. 1 / x ^ m)" (is ?th1) and Digamma_real_asymptotics: "(\x. Digamma x - (ln x - (\k=1..m. bernoulli' k / real k / x ^ k))) \ O(\x. 1 / (x ^ Suc m))" (is ?th2) and Polygamma_real_asymptotics: "j > 0 \ (\x. Polygamma j x - (- 1) ^ Suc j * (\k\m. bernoulli' k * pochhammer (real (Suc k)) (j - 1) / x ^ (k + j))) \ O(\x. 1 / x ^ (m+j+1))" (is "_ \ ?th3") proof - define G :: "nat \ real \ real" where "G = (\m. if m = 0 then ln_Gamma else Polygamma (m - 1))" have *: "(\x. G j x - h x) \ O(\x. 1 / x ^ (m + j))" if "\x::real. x > 0 \ Polygamma_approx j m x = h x" for j h proof - have "(\x. G j x - h x) \ \(\x. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x)" (is "_ \ \(?f)") using that by (intro bigthetaI_cong) (auto intro: eventually_mono[OF eventually_gt_at_top[of "0::real"]] simp del: funpow.simps simp: higher_deriv_ln_Gamma_real G_def) also have "?f \ O(\x::real. 1 / x ^ (m + j))" using m by (rule higher_deriv_lnGamma_stirling) finally show ?thesis . qed note [[simproc del: simplify_landau_sum]] from *[OF Polygamma_approx_0] assms show ?th1 by (simp add: G_def Polygamma_approx_0 stirling_sum_def field_simps) from *[OF Polygamma_approx_1_real'] assms show ?th2 by (simp add: G_def field_simps) assume j: "j > 0" from *[OF Polygamma_approx_ge_2_real, of "j - 1"] assms j show ?th3 by (simp add: G_def stirling_sum'_def power_add power_diff field_simps) qed subsection \Asymptotics of the complex Gamma function\ text \ The \m\-th order remainder of Stirling's formula for $\log\Gamma$ is $O(s^{-m})$ uniformly over any complex cone $\text{Arg}(z) \leq \alpha$, $z\neq 0$ for any angle $\alpha\in(0, \pi)$. This means that there is bounded by $c z^{-m}$ for some constant $c$ for all $z$ in this cone. \ context fixes F and \ assumes \: "\ \ {0<.. principal (complex_cone' \ - {0})" begin lemma stirling_integral_bigo: fixes m :: nat assumes m: "m > 0" shows "stirling_integral m \ O[F](\s. 1 / s ^ m)" proof - obtain c where c: "\s. s \ complex_cone' \ - {0} \ norm (stirling_integral m s) \ c / norm s ^ m" using stirling_integral_bound'[OF \m > 0\ \] by blast have "0 \ norm (stirling_integral m 1 :: complex)" by simp also have "\ \ c" using c[of 1] \ by simp finally have "c \ 0" . have "eventually (\s. s \ complex_cone' \ - {0}) F" unfolding F_def by (auto simp: eventually_principal) hence "eventually (\s. norm (stirling_integral m s) \ c * norm (1 / s ^ m)) F" by eventually_elim (use c in \simp add: norm_divide norm_power\) thus "stirling_integral m \ O[F](\s. 1 / s ^ m)" by (intro bigoI[of _ c]) auto qed end text \ The following is a more explicit statement of this: \ theorem ln_Gamma_complex_asymptotics_explicit: fixes m :: nat and \ :: real assumes "m > 0" and "\ \ {0<.. complex" where "\s::complex. s \ \\<^sub>\\<^sub>0 \ ln_Gamma s = (s - 1/2) * ln s - s + ln (2 * pi) / 2 + (\k=1..s. s \ 0 \ \Arg s\ \ \ \ norm (R s) \ C / norm s ^ m" proof - obtain c where c: "\s. s \ complex_cone' \ - {0} \ norm (stirling_integral m s) \ c / norm s ^ m" using stirling_integral_bound'[OF assms] by blast have "0 \ norm (stirling_integral m 1 :: complex)" by simp also have "\ \ c" using c[of 1] assms by simp finally have "c \ 0" . define R where "R = (\s::complex. stirling_integral m s / of_nat m)" show ?thesis proof (rule that) from ln_Gamma_stirling_complex[of _ m] assms show "\s::complex. s \ \\<^sub>\\<^sub>0 \ ln_Gamma s = (s - 1 / 2) * ln s - s + ln (2 * pi) / 2 + (\k=1..s. s \ 0 \ \Arg s\ \ \ \ cmod (R s) \ c / real m / cmod s ^ m" proof (safe, goal_cases) case (1 s) show ?case using 1 c[of s] assms by (auto simp: complex_cone_altdef abs_le_iff R_def norm_divide field_simps) qed qed qed text \ Lastly, we can also derive the asymptotics of $\Gamma$ itself: \[\Gamma(z) \sim \sqrt{2\pi / z} \left(\frac{z}{e}\right)^z\] uniformly for $|z|\to\infty$ within the cone $\text{Arg}(z) \leq \alpha$ for $\alpha\in(0,\pi)$: \ context fixes F and \ assumes \: "\ \ {0<.. inf at_infinity (principal (complex_cone' \))" begin lemma Gamma_complex_asymp_equiv: "Gamma \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2))" proof - define I :: "complex \ complex" where "I = stirling_integral 1" have "eventually (\s. s \ complex_cone' \) F" by (auto simp: eventually_inf_principal F_def) moreover have "eventually (\s. s \ 0) F" unfolding F_def eventually_inf_principal using eventually_not_equal_at_infinity by eventually_elim auto ultimately have "eventually (\s. Gamma s = sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s)) F" proof eventually_elim case (elim s) from elim have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] \ by auto from elim have [simp]: "s \ 0" by auto from s' have "Gamma s = exp (ln_Gamma s)" unfolding Gamma_complex_altdef using nonpos_Ints_subset_nonpos_Reals by auto also from s' have "ln_Gamma s = (s-1/2) * Ln s - s + complex_of_real (ln (2 * pi) / 2) - I s" by (subst ln_Gamma_stirling_complex[of _ 1]) (simp_all add: exp_add exp_diff I_def) also have "exp \ = exp ((s - 1 / 2) * Ln s) / exp s * exp (complex_of_real (ln (2 * pi) / 2)) / exp (I s)" unfolding exp_diff exp_add by (simp add: exp_diff exp_add) also have "exp ((s - 1 / 2) * Ln s) = s powr (s - 1 / 2)" by (simp add: powr_def) also have "exp (complex_of_real (ln (2 * pi) / 2)) = sqrt (2 * pi)" by (subst exp_of_real) (auto simp: powr_def simp flip: powr_half_sqrt) also have "exp s = exp 1 powr s" by (simp add: powr_def) also have "s powr (s - 1 / 2) / exp 1 powr s = (s powr s / exp 1 powr s) / s powr (1/2)" by (subst powr_diff) auto also have *: "Ln (s / exp 1) = Ln s - 1" using Ln_divide_of_real[of "exp 1" s] by (simp flip: exp_of_real) hence "s powr s / exp 1 powr s = (s / exp 1) powr s" unfolding powr_def by (subst *) (auto simp: exp_diff field_simps) finally show "Gamma s = sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s)" by (simp add: algebra_simps) qed hence "Gamma \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s))" by (rule asymp_equiv_refl_ev) also have "\ \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / 1)" proof (intro asymp_equiv_intros) have "F \ principal (complex_cone' \ - {0})" unfolding le_principal F_def eventually_inf_principal using eventually_not_equal_at_infinity by eventually_elim auto moreover have "I \ O[principal (complex_cone' \ - {0})](\s. 1 / s)" using stirling_integral_bigo[of \ 1] \ unfolding F_def by (simp add: I_def) ultimately have "I \ O[F](\s. 1 / s)" by (rule landau_o.big.filter_mono) also have "(\s. 1 / s) \ o[F](\s. 1)" proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" hence "eventually (\z::complex. norm z \ 1 / c) at_infinity" by (auto simp: eventually_at_infinity) moreover have "eventually (\z::complex. z \ 0) at_infinity" by (rule eventually_not_equal_at_infinity) ultimately show "eventually (\z::complex. norm (1 / z) \ c * norm (1 :: complex)) F" unfolding F_def eventually_inf_principal by eventually_elim (use \c > 0\ in \auto simp: norm_divide field_simps\) qed finally have "I \ o[F](\s. 1)" . from smalloD_tendsto[OF this] have [tendsto_intros]: "(I \ 0) F" by simp show "(\x. exp (I x)) \[F] (\x. 1)" by (rule asymp_equivI' tendsto_eq_intros refl | simp)+ qed finally show ?thesis by simp qed end end diff --git a/thys/Sturm_Sequences/Lib/Misc_Polynomial.thy b/thys/Sturm_Sequences/Lib/Misc_Polynomial.thy --- a/thys/Sturm_Sequences/Lib/Misc_Polynomial.thy +++ b/thys/Sturm_Sequences/Lib/Misc_Polynomial.thy @@ -1,961 +1,961 @@ (* Author: Manuel Eberl *) theory Misc_Polynomial -imports "HOL-Computational_Algebra.Polynomial" "HOL-Computational_Algebra.Polynomial_Factorial" +imports "HOL-Computational_Algebra.Polynomial" "HOL-Computational_Algebra.Polynomial_Factorial" "Pure-ex.Guess" begin subsection \Analysis\ lemma fun_eq_in_ivl: assumes "a \ b" "\x::real. a \ x \ x \ b \ eventually (\\. f \ = f x) (at x)" shows "f a = f b" proof (rule connected_local_const) show "connected {a..b}" "a \ {a..b}" "b \ {a..b}" using \a \ b\ by (auto intro: connected_Icc) show "\aa\{a..b}. eventually (\b. f aa = f b) (at aa within {a..b})" proof fix x assume "x \ {a..b}" with assms(2)[rule_format, of x] show "eventually (\b. f x = f b) (at x within {a..b})" by (auto simp: eventually_at_filter elim: eventually_mono) qed qed subsection \Polynomials\ subsubsection \General simplification lemmas\ lemma pderiv_div: assumes [simp]: "q dvd p" "q \ 0" shows "pderiv (p div q) = (q * pderiv p - p * pderiv q) div (q * q)" "q*q dvd (q * pderiv p - p * pderiv q)" proof- from assms obtain r where "p = q * r" unfolding dvd_def by blast hence "q * pderiv p - p * pderiv q = (q * q) * pderiv r" by (simp add: algebra_simps pderiv_mult) thus "q*q dvd (q * pderiv p - p * pderiv q)" by simp note 0 = pderiv_mult[of q "p div q"] have 1: "q * (p div q) = p" by (metis assms(1) assms(2) dvd_def nonzero_mult_div_cancel_left) have f1: "pderiv (p div q) * (q * q) div (q * q) = pderiv (p div q)" by simp have f2: "pderiv p = q * pderiv (p div q) + p div q * pderiv q" by (metis 0 1) have "p * pderiv q = pderiv q * (q * (p div q))" by (metis 1 mult.commute) then have "p * pderiv q = q * (p div q * pderiv q)" by fastforce then have "q * pderiv p - p * pderiv q = q * (q * pderiv (p div q))" using f2 by (metis add_diff_cancel_right' distrib_left) then show "pderiv (p div q) = (q * pderiv p - p * pderiv q) div (q * q)" using f1 by (metis mult.commute mult.left_commute) qed subsubsection \Divisibility of polynomials\ text \ Two polynomials that are coprime have no common roots. \ lemma coprime_imp_no_common_roots: "\ (poly p x = 0 \ poly q x = 0)" if "coprime p q" for x :: "'a :: field" proof clarify assume "poly p x = 0" "poly q x = 0" then have "[:-x, 1:] dvd p" "[:-x, 1:] dvd q" by (simp_all add: poly_eq_0_iff_dvd) with that have "is_unit [:-x, 1:]" by (rule coprime_common_divisor) then show False by (auto simp add: is_unit_pCons_iff) qed lemma poly_div: assumes "poly q x \ 0" and "(q::'a :: field poly) dvd p" shows "poly (p div q) x = poly p x / poly q x" proof- from assms have [simp]: "q \ 0" by force have "poly q x * poly (p div q) x = poly (q * (p div q)) x" by simp also have "q * (p div q) = p" using assms by (simp add: div_mult_swap) finally show "poly (p div q) x = poly p x / poly q x" using assms by (simp add: field_simps) qed (* TODO: make this less ugly *) lemma poly_div_gcd_squarefree_aux: assumes "pderiv (p::('a::{field_char_0,field_gcd}) poly) \ 0" defines "d \ gcd p (pderiv p)" shows "coprime (p div d) (pderiv (p div d))" and "\x. poly (p div d) x = 0 \ poly p x = 0" proof - obtain r s where "bezout_coefficients p (pderiv p) = (r, s)" by (auto simp add: prod_eq_iff) then have "r * p + s * pderiv p = gcd p (pderiv p)" by (rule bezout_coefficients) then have rs: "d = r * p + s * pderiv p" by (simp add: d_def) define t where "t = p div d" define p' where [simp]: "p' = pderiv p" define d' where [simp]: "d' = pderiv d" define u where "u = p' div d" have A: "p = t * d" and B: "p' = u * d" by (simp_all add: t_def u_def d_def algebra_simps) from poly_squarefree_decomp[OF assms(1) A B[unfolded p'_def] rs] show "\x. poly (p div d) x = 0 \ poly p x = 0" by (auto simp: t_def) from rs have C: "s*t*d' = d * (1 - r*t - s*pderiv t)" by (simp add: A B algebra_simps pderiv_mult) from assms have [simp]: "p \ 0" "d \ 0" "t \ 0" by (force, force, subst (asm) A, force) have "\x. \x dvd t; x dvd (pderiv t)\ \ x dvd 1" proof - fix x assume "x dvd t" "x dvd (pderiv t)" then obtain v w where vw: "t = x*v" "pderiv t = x*w" unfolding dvd_def by blast define x' v' where [simp]: "x' = pderiv x" and [simp]: "v' = pderiv v" from vw have "x*v' + v*x' = x*w" by (simp add: pderiv_mult) hence "v*x' = x*(w - v')" by (simp add: algebra_simps) hence "x dvd v*pderiv x" by simp then obtain y where y: "v*x' = x*y" unfolding dvd_def by force from \t \ 0\ and vw have "x \ 0" by simp have x_pow_n_dvd_d: "\n. x^n dvd d" proof- fix n show "x ^ n dvd d" proof (induction n, simp, rename_tac n, case_tac n) fix n assume "n = (0::nat)" from vw and C have "d = x*(d*r*v + d*s*w + s*v*d')" by (simp add: algebra_simps) with \n = 0\ show "x^Suc n dvd d" by (force intro: dvdI) next fix n n' assume IH: "x^n dvd d" and "n = Suc n'" hence [simp]: "Suc n' = n" "x * x^n' = x^n" by simp_all define c :: "'a poly" where "c = [:of_nat n:]" from pderiv_power_Suc[of x n'] have [simp]: "pderiv (x^n) = c*x^n' * x'" unfolding c_def by simp from IH obtain z where d: "d = x^n * z" unfolding dvd_def by blast define z' where [simp]: "z' = pderiv z" from d \d \ 0\ have "x^n \ 0" "z \ 0" by force+ from C d have "x^n*z = z*r*v*x^Suc n + z*s*c*x^n*(v*x') + s*v*z'*x^Suc n + s*z*(v*x')*x^n + s*z*v'*x^Suc n" by (simp add: algebra_simps vw pderiv_mult) also have "... = x^n*x * (z*r*v + z*s*c*y + s*v*z' + s*z*y + s*z*v')" by (simp only: y, simp add: algebra_simps) finally have "z = x*(z*r*v+z*s*c*y+s*v*z'+s*z*y+s*z*v')" using \x^n \ 0\ by force hence "x dvd z" by (metis dvd_triv_left) with d show "x^Suc n dvd d" by simp qed qed have "degree x = 0" proof (cases "degree x", simp) case (Suc n) hence "x \ 0" by auto with Suc have "degree (x ^ (Suc (degree d))) > degree d" by (subst degree_power_eq, simp_all) moreover from x_pow_n_dvd_d[of "Suc (degree d)"] and \d \ 0\ have "degree (x^Suc (degree d)) \ degree d" by (simp add: dvd_imp_degree_le) ultimately show ?thesis by simp qed then obtain c where [simp]: "x = [:c:]" by (cases x, simp split: if_split_asm) moreover from \x \ 0\ have "c \ 0" by simp ultimately show "x dvd 1" using dvdI[of 1 x "[:inverse c:]"] by simp qed then show "coprime t (pderiv t)" by (rule coprimeI) qed lemma normalize_field: "normalize (x :: 'a :: {field,normalization_semidom}) = (if x = 0 then 0 else 1)" by (auto simp: is_unit_normalize dvd_field_iff) lemma normalize_field_eq_1 [simp]: "x \ 0 \ normalize (x :: 'a :: {field,normalization_semidom}) = 1" by (simp add: normalize_field) lemma unit_factor_field [simp]: "unit_factor (x :: 'a :: {field,normalization_semidom}) = x" by (cases "x = 0") (auto simp: is_unit_unit_factor dvd_field_iff) text \ Dividing a polynomial by its gcd with its derivative yields a squarefree polynomial with the same roots. \ lemma poly_div_gcd_squarefree: assumes "(p :: ('a::{field_char_0,field_gcd}) poly) \ 0" defines "d \ gcd p (pderiv p)" shows "coprime (p div d) (pderiv (p div d))" (is ?A) and "\x. poly (p div d) x = 0 \ poly p x = 0" (is "\x. ?B x") proof- have "?A \ (\x. ?B x)" proof (cases "pderiv p = 0") case False from poly_div_gcd_squarefree_aux[OF this] show ?thesis unfolding d_def by auto next case True then obtain c where [simp]: "p = [:c:]" using pderiv_iszero by blast from assms(1) have "c \ 0" by simp from True have "d = smult (inverse c) p" by (simp add: d_def normalize_poly_def map_poly_pCons field_simps) with \p \ 0\ \c \ 0\ have "p div d = [:c:]" by (simp add: pCons_one) with \c \ 0\ show ?thesis by (simp add: normalize_const_poly is_unit_triv) qed thus ?A and "\x. ?B x" by simp_all qed subsubsection \Sign changes of a polynomial\ text \ If a polynomial has different signs at two points, it has a root inbetween. \ lemma poly_different_sign_imp_root: assumes "a < b" and "sgn (poly p a) \ sgn (poly p (b::real))" shows "\x. a \ x \ x \ b \ poly p x = 0" proof (cases "poly p a = 0 \ poly p b = 0") case True thus ?thesis using assms(1) by (elim disjE, rule_tac exI[of _ a], simp, rule_tac exI[of _ b], simp) next case False hence [simp]: "poly p a \ 0" "poly p b \ 0" by simp_all show ?thesis proof (cases "poly p a < 0") case True hence "sgn (poly p a) = -1" by simp with assms True have "poly p b > 0" by (auto simp: sgn_real_def split: if_split_asm) from poly_IVT_pos[OF \a < b\ True this] guess x .. thus ?thesis by (intro exI[of _ x], simp) next case False hence "poly p a > 0" by (simp add: not_less less_eq_real_def) hence "sgn (poly p a) = 1" by simp with assms False have "poly p b < 0" by (auto simp: sgn_real_def not_less less_eq_real_def split: if_split_asm) from poly_IVT_neg[OF \a < b\ \poly p a > 0\ this] guess x .. thus ?thesis by (intro exI[of _ x], simp) qed qed lemma poly_different_sign_imp_root': assumes "sgn (poly p a) \ sgn (poly p (b::real))" shows "\x. poly p x = 0" using assms by (cases "a < b", auto dest!: poly_different_sign_imp_root simp: less_eq_real_def not_less) lemma no_roots_inbetween_imp_same_sign: assumes "a < b" "\x. a \ x \ x \ b \ poly p x \ (0::real)" shows "sgn (poly p a) = sgn (poly p b)" using poly_different_sign_imp_root assms by auto subsubsection \Limits of polynomials\ lemma poly_neighbourhood_without_roots: assumes "(p :: real poly) \ 0" shows "eventually (\x. poly p x \ 0) (at x\<^sub>0)" proof- { fix \ :: real assume "\ > 0" have fin: "finite {x. \x-x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x = 0}" using poly_roots_finite[OF assms] by simp with \\ > 0\have "\\>0. \\\ \ (\x. \x-x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x \ 0)" proof (induction "card {x. \x-x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x = 0}" arbitrary: \ rule: less_induct) case (less \) let ?A = "{x. \x - x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x = 0}" show ?case proof (cases "card ?A") case 0 hence "?A = {}" using less by auto thus ?thesis using less(2) by (rule_tac exI[of _ \], auto) next case (Suc _) with less(3) have "{x. \x - x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x = 0} \ {}" by force then obtain x where x_props: "\x - x\<^sub>0\ < \" "x \ x\<^sub>0" "poly p x = 0" by blast define \' where "\' = \x - x\<^sub>0\ / 2" have "\' > 0" "\' < \" unfolding \'_def using x_props by simp_all from x_props(1,2) and \\ > 0\ have "x \ {x'. \x' - x\<^sub>0\ < \' \ x' \ x\<^sub>0 \ poly p x' = 0}" (is "_ \ ?B") by (auto simp: \'_def) moreover from x_props have "x \ {x. \x - x\<^sub>0\ < \ \ x \ x\<^sub>0 \ poly p x = 0}" by blast ultimately have "?B \ ?A" by auto hence "card ?B < card ?A" "finite ?B" by (rule psubset_card_mono[OF less(3)], blast intro: finite_subset[OF _ less(3)]) from less(1)[OF this(1) \\' > 0\ this(2)] show ?thesis using \\' < \\ by force qed qed } from this[of "1"] show ?thesis by (auto simp: eventually_at dist_real_def) qed lemma poly_neighbourhood_same_sign: assumes "poly p (x\<^sub>0 :: real) \ 0" shows "eventually (\x. sgn (poly p x) = sgn (poly p x\<^sub>0)) (at x\<^sub>0)" proof - have cont: "isCont (\x. sgn (poly p x)) x\<^sub>0" by (rule isCont_sgn, rule poly_isCont, rule assms) then have "eventually (\x. \sgn (poly p x) - sgn (poly p x\<^sub>0)\ < 1) (at x\<^sub>0)" by (auto simp: isCont_def tendsto_iff dist_real_def) then show ?thesis by (rule eventually_mono) (simp add: sgn_real_def split: if_split_asm) qed lemma poly_lhopital: assumes "poly p (x::real) = 0" "poly q x = 0" "q \ 0" assumes "(\x. poly (pderiv p) x / poly (pderiv q) x) \x\ y" shows "(\x. poly p x / poly q x) \x\ y" using assms proof (rule_tac lhopital) have "isCont (poly p) x" "isCont (poly q) x" by simp_all with assms(1,2) show "poly p \x\ 0" "poly q \x\ 0" by (simp_all add: isCont_def) from \q \ 0\ and \poly q x = 0\ have "pderiv q \ 0" by (auto dest: pderiv_iszero) from poly_neighbourhood_without_roots[OF this] show "eventually (\x. poly (pderiv q) x \ 0) (at x)" . qed (auto intro: poly_DERIV poly_neighbourhood_without_roots) lemma poly_roots_bounds: assumes "p \ 0" obtains l u where "l \ (u :: real)" and "poly p l \ 0" and "poly p u \ 0" and "{x. x > l \ x \ u \ poly p x = 0 } = {x. poly p x = 0}" and "\x. x \ l \ sgn (poly p x) = sgn (poly p l)" and "\x. x \ u \ sgn (poly p x) = sgn (poly p u)" proof from assms have "finite {x. poly p x = 0}" (is "finite ?roots") using poly_roots_finite by fast let ?roots' = "insert 0 ?roots" define l where "l = Min ?roots' - 1" define u where "u = Max ?roots' + 1" from \finite ?roots\ have A: "finite ?roots'" by auto from Min_le[OF this, of 0] and Max_ge[OF this, of 0] show "l \ u" by (simp add: l_def u_def) from Min_le[OF A] have l_props: "\x. x\l \ poly p x \ 0" by (fastforce simp: l_def) from Max_ge[OF A] have u_props: "\x. x\u \ poly p x \ 0" by (fastforce simp: u_def) from l_props u_props show [simp]: "poly p l \ 0" "poly p u \ 0" by auto from l_props have "\x. poly p x = 0 \ x > l" by (metis not_le) moreover from u_props have "\x. poly p x = 0 \ x \ u" by (metis linear) ultimately show "{x. x > l \ x \ u \ poly p x = 0} = ?roots" by auto { fix x assume A: "x < l" "sgn (poly p x) \ sgn (poly p l)" with poly_IVT_pos[OF A(1), of p] poly_IVT_neg[OF A(1), of p] A(2) have False by (auto split: if_split_asm simp: sgn_real_def l_props not_less less_eq_real_def) } thus "\x. x \ l \ sgn (poly p x) = sgn (poly p l)" by (case_tac "x = l", auto simp: less_eq_real_def) { fix x assume A: "x > u" "sgn (poly p x) \ sgn (poly p u)" with u_props poly_IVT_neg[OF A(1), of p] poly_IVT_pos[OF A(1), of p] A(2) have False by (auto split: if_split_asm simp: sgn_real_def l_props not_less less_eq_real_def) } thus "\x. x \ u \ sgn (poly p x) = sgn (poly p u)" by (case_tac "x = u", auto simp: less_eq_real_def) qed definition poly_inf :: "('a::real_normed_vector) poly \ 'a" where "poly_inf p \ sgn (coeff p (degree p))" definition poly_neg_inf :: "('a::real_normed_vector) poly \ 'a" where "poly_neg_inf p \ if even (degree p) then sgn (coeff p (degree p)) else -sgn (coeff p (degree p))" lemma poly_inf_0_iff[simp]: "poly_inf p = 0 \ p = 0" "poly_neg_inf p = 0 \ p = 0" by (auto simp: poly_inf_def poly_neg_inf_def sgn_zero_iff) lemma poly_inf_mult[simp]: fixes p :: "('a::real_normed_field) poly" shows "poly_inf (p*q) = poly_inf p * poly_inf q" "poly_neg_inf (p*q) = poly_neg_inf p * poly_neg_inf q" unfolding poly_inf_def poly_neg_inf_def by ((cases "p = 0 \ q = 0",auto simp: sgn_zero_iff degree_mult_eq[of p q] coeff_mult_degree_sum Real_Vector_Spaces.sgn_mult)[])+ lemma poly_neq_0_at_infinity: assumes "(p :: real poly) \ 0" shows "eventually (\x. poly p x \ 0) at_infinity" proof- from poly_roots_bounds[OF assms] guess l u . note lu_props = this define b where "b = max (-l) u" show ?thesis proof (subst eventually_at_infinity, rule exI[of _ b], clarsimp) fix x assume A: "\x\ \ b" and B: "poly p x = 0" show False proof (cases "x \ 0") case True with A have "x \ u" unfolding b_def by simp with lu_props(3, 6) show False by (metis sgn_zero_iff B) next case False with A have "x \ l" unfolding b_def by simp with lu_props(2, 5) show False by (metis sgn_zero_iff B) qed qed qed lemma poly_limit_aux: fixes p :: "real poly" defines "n \ degree p" shows "((\x. poly p x / x ^ n) \ coeff p n) at_infinity" proof (subst filterlim_cong, rule refl, rule refl) show "eventually (\x. poly p x / x^n = (\i\n. coeff p i / x^(n-i))) at_infinity" proof (rule eventually_mono) show "eventually (\x::real. x \ 0) at_infinity" by (simp add: eventually_at_infinity, rule exI[of _ 1], auto) fix x :: real assume [simp]: "x \ 0" show "poly p x / x ^ n = (\i\n. coeff p i / x ^ (n - i))" by (simp add: n_def sum_divide_distrib power_diff poly_altdef) qed let ?a = "\i. if i = n then coeff p n else 0" have "\i\{..n}. ((\x. coeff p i / x ^ (n - i)) \ ?a i) at_infinity" proof fix i assume "i \ {..n}" hence "i \ n" by simp show "((\x. coeff p i / x ^ (n - i)) \ ?a i) at_infinity" proof (cases "i = n") case True thus ?thesis by (intro tendstoI, subst eventually_at_infinity, intro exI[of _ 1], simp add: dist_real_def) next case False hence "n - i > 0" using \i \ n\ by simp from tendsto_inverse_0 and divide_real_def[of 1] have "((\x. 1 / x :: real) \ 0) at_infinity" by simp from tendsto_power[OF this, of "n - i"] have "((\x::real. 1 / x ^ (n - i)) \ 0) at_infinity" using \n - i > 0\ by (simp add: power_0_left power_one_over) from tendsto_mult_right_zero[OF this, of "coeff p i"] have "((\x. coeff p i / x ^ (n - i)) \ 0) at_infinity" by (simp add: field_simps) thus ?thesis using False by simp qed qed hence "((\x. \i\n. coeff p i / x^(n-i)) \ (\i\n. ?a i)) at_infinity" by (force intro!: tendsto_sum) also have "(\i\n. ?a i) = coeff p n" by (subst sum.delta, simp_all) finally show "((\x. \i\n. coeff p i / x^(n-i)) \ coeff p n) at_infinity" . qed lemma poly_at_top_at_top: fixes p :: "real poly" assumes "degree p \ 1" "coeff p (degree p) > 0" shows "LIM x at_top. poly p x :> at_top" proof- let ?n = "degree p" define f g where "f x = poly p x / x^?n" and "g x = x ^ ?n" for x :: real from poly_limit_aux have "(f \ coeff p (degree p)) at_top" using tendsto_mono at_top_le_at_infinity unfolding f_def by blast moreover from assms have "LIM x at_top. g x :> at_top" by (auto simp add: g_def intro!: filterlim_pow_at_top filterlim_ident) ultimately have "LIM x at_top. f x * g x :> at_top" using filterlim_tendsto_pos_mult_at_top assms by simp also have "eventually (\x. f x * g x = poly p x) at_top" unfolding f_def g_def by (subst eventually_at_top_linorder, rule exI[of _ 1], simp add: poly_altdef field_simps sum_distrib_left power_diff) note filterlim_cong[OF refl refl this] finally show ?thesis . qed lemma poly_at_bot_at_top: fixes p :: "real poly" assumes "degree p \ 1" "coeff p (degree p) < 0" shows "LIM x at_top. poly p x :> at_bot" proof- from poly_at_top_at_top[of "-p"] and assms have "LIM x at_top. -poly p x :> at_top" by simp thus ?thesis by (simp add: filterlim_uminus_at_bot) qed lemma poly_lim_inf: "eventually (\x::real. sgn (poly p x) = poly_inf p) at_top" proof (cases "degree p \ 1") case False hence "degree p = 0" by simp then obtain c where "p = [:c:]" by (cases p, auto split: if_split_asm) thus ?thesis by (simp add: eventually_at_top_linorder poly_inf_def) next case True note deg = this let ?lc = "coeff p (degree p)" from True have "?lc \ 0" by force show ?thesis proof (cases "?lc > 0") case True from poly_at_top_at_top[OF deg this] obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ 1" by (fastforce simp: filterlim_at_top eventually_at_top_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = 1" by force thus ?thesis by (simp only: eventually_at_top_linorder poly_inf_def, intro exI[of _ x\<^sub>0], simp add: True) next case False hence "?lc < 0" using \?lc \ 0\ by linarith from poly_at_bot_at_top[OF deg this] obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ -1" by (fastforce simp: filterlim_at_bot eventually_at_top_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = -1" by force thus ?thesis by (simp only: eventually_at_top_linorder poly_inf_def, intro exI[of _ x\<^sub>0], simp add: \?lc < 0\) qed qed lemma poly_at_top_or_bot_at_bot: fixes p :: "real poly" assumes "degree p \ 1" "coeff p (degree p) > 0" shows "LIM x at_bot. poly p x :> (if even (degree p) then at_top else at_bot)" proof- let ?n = "degree p" define f g where "f x = poly p x / x ^ ?n" and "g x = x ^ ?n" for x :: real from poly_limit_aux have "(f \ coeff p (degree p)) at_bot" using tendsto_mono at_bot_le_at_infinity by (force simp: f_def [abs_def]) moreover from assms have "LIM x at_bot. g x :> (if even (degree p) then at_top else at_bot)" by (auto simp add: g_def split: if_split_asm intro: filterlim_pow_at_bot_even filterlim_pow_at_bot_odd filterlim_ident) ultimately have "LIM x at_bot. f x * g x :> (if even ?n then at_top else at_bot)" by (auto simp: assms intro: filterlim_tendsto_pos_mult_at_top filterlim_tendsto_pos_mult_at_bot) also have "eventually (\x. f x * g x = poly p x) at_bot" unfolding f_def g_def by (subst eventually_at_bot_linorder, rule exI[of _ "-1"], simp add: poly_altdef field_simps sum_distrib_left power_diff) note filterlim_cong[OF refl refl this] finally show ?thesis . qed lemma poly_at_bot_or_top_at_bot: fixes p :: "real poly" assumes "degree p \ 1" "coeff p (degree p) < 0" shows "LIM x at_bot. poly p x :> (if even (degree p) then at_bot else at_top)" proof- from poly_at_top_or_bot_at_bot[of "-p"] and assms have "LIM x at_bot. -poly p x :> (if even (degree p) then at_top else at_bot)" by simp thus ?thesis by (auto simp: filterlim_uminus_at_bot) qed lemma poly_lim_neg_inf: "eventually (\x::real. sgn (poly p x) = poly_neg_inf p) at_bot" proof (cases "degree p \ 1") case False hence "degree p = 0" by simp then obtain c where "p = [:c:]" by (cases p, auto split: if_split_asm) thus ?thesis by (simp add: eventually_at_bot_linorder poly_neg_inf_def) next case True note deg = this let ?lc = "coeff p (degree p)" from True have "?lc \ 0" by force show ?thesis proof (cases "?lc > 0") case True note lc_pos = this show ?thesis proof (cases "even (degree p)") case True from poly_at_top_or_bot_at_bot[OF deg lc_pos] and True obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ 1" by (fastforce simp add: filterlim_at_top filterlim_at_bot eventually_at_bot_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = 1" by force thus ?thesis by (simp add: True eventually_at_bot_linorder poly_neg_inf_def, intro exI[of _ x\<^sub>0], simp add: lc_pos) next case False from poly_at_top_or_bot_at_bot[OF deg lc_pos] and False obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ -1" by (fastforce simp add: filterlim_at_bot eventually_at_bot_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = -1" by force thus ?thesis by (simp add: False eventually_at_bot_linorder poly_neg_inf_def, intro exI[of _ x\<^sub>0], simp add: lc_pos) qed next case False hence lc_neg: "?lc < 0" using \?lc \ 0\ by linarith show ?thesis proof (cases "even (degree p)") case True with poly_at_bot_or_top_at_bot[OF deg lc_neg] obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ -1" by (fastforce simp: filterlim_at_bot eventually_at_bot_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = -1" by force thus ?thesis by (simp only: True eventually_at_bot_linorder poly_neg_inf_def, intro exI[of _ x\<^sub>0], simp add: lc_neg) next case False with poly_at_bot_or_top_at_bot[OF deg lc_neg] obtain x\<^sub>0 where "\x. x \ x\<^sub>0 \ poly p x \ 1" by (fastforce simp: filterlim_at_top eventually_at_bot_linorder less_eq_real_def) hence "\x. x \ x\<^sub>0 \ sgn (poly p x) = 1" by force thus ?thesis by (simp only: False eventually_at_bot_linorder poly_neg_inf_def, intro exI[of _ x\<^sub>0], simp add: lc_neg) qed qed qed subsubsection \Signs of polynomials for sufficiently large values\ lemma polys_inf_sign_thresholds: assumes "finite (ps :: real poly set)" obtains l u where "l \ u" and "\p. \p \ ps; p \ 0\ \ {x. l < x \ x \ u \ poly p x = 0} = {x. poly p x = 0}" and "\p x. \p \ ps; x \ u\ \ sgn (poly p x) = poly_inf p" and "\p x. \p \ ps; x \ l\ \ sgn (poly p x) = poly_neg_inf p" proof goal_cases case prems: 1 have "\l u. l \ u \ (\p x. p \ ps \ x \ u \ sgn (poly p x) = poly_inf p) \ (\p x. p \ ps \ x \ l \ sgn (poly p x) = poly_neg_inf p)" (is "\l u. ?P ps l u") proof (induction rule: finite_subset_induct[OF assms(1), where A = UNIV]) case 1 show ?case by simp next case 2 show ?case by (intro exI[of _ 42], simp) next case prems: (3 p ps) from prems(4) obtain l u where lu_props: "?P ps l u" by blast from poly_lim_inf obtain u' where u'_props: "\x\u'. sgn (poly p x) = poly_inf p" by (force simp add: eventually_at_top_linorder) from poly_lim_neg_inf obtain l' where l'_props: "\x\l'. sgn (poly p x) = poly_neg_inf p" by (force simp add: eventually_at_bot_linorder) show ?case by (rule exI[of _ "min l l'"], rule exI[of _ "max u u'"], insert lu_props l'_props u'_props, auto) qed then obtain l u where lu_props: "l \ u" "\p x. p \ ps \ u \ x \ sgn (poly p x) = poly_inf p" "\p x. p \ ps \ x \ l \ sgn (poly p x) = poly_neg_inf p" by blast moreover { fix p x assume A: "p \ ps" "p \ 0" "poly p x = 0" from A have "l < x" "x < u" by (auto simp: not_le[symmetric] dest: lu_props(2,3)) } note A = this have "\p. p \ ps \ p \ 0 \ {x. l < x \ x \ u \ poly p x = 0} = {x. poly p x = 0}" by (auto dest: A) from prems[OF lu_props(1) this lu_props(2,3)] show thesis . qed subsubsection \Positivity of polynomials\ lemma poly_pos: "(\x::real. poly p x > 0) \ poly_inf p = 1 \ (\x. poly p x \ 0)" proof (intro iffI conjI) assume A: "\x::real. poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. poly p x \ 0" by simp from poly_lim_inf obtain x where "sgn (poly p x) = poly_inf p" by (auto simp: eventually_at_top_linorder) with A show "poly_inf p = 1" by (simp add: sgn_real_def split: if_split_asm) next assume "poly_inf p = 1 \ (\x. poly p x \ 0)" hence A: "poly_inf p = 1" and B: "(\x. poly p x \ 0)" by simp_all from poly_lim_inf obtain x where C: "sgn (poly p x) = poly_inf p" by (auto simp: eventually_at_top_linorder) show "\x. poly p x > 0" proof (rule ccontr) assume "\(\x. poly p x > 0)" then obtain x' where "poly p x' \ 0" by (auto simp: not_less) with A and C have "sgn (poly p x') \ sgn (poly p x)" by (auto simp: sgn_real_def split: if_split_asm) from poly_different_sign_imp_root'[OF this] and B show False by blast qed qed lemma poly_pos_greater: "(\x::real. x > a \ poly p x > 0) \ poly_inf p = 1 \ (\x. x > a \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x::real. x > a \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. x > a \ poly p x \ 0" by auto from poly_lim_inf obtain x\<^sub>0 where "\x\x\<^sub>0. sgn (poly p x) = poly_inf p" by (auto simp: eventually_at_top_linorder) hence "poly_inf p = sgn (poly p (max x\<^sub>0 (a + 1)))" by simp also from A have "... = 1" by force finally show "poly_inf p = 1" . next assume "poly_inf p = 1 \ (\x. x > a \ poly p x \ 0)" hence A: "poly_inf p = 1" and B: "(\x. x > a \ poly p x \ 0)" by simp_all from poly_lim_inf obtain x\<^sub>0 where C: "\x\x\<^sub>0. sgn (poly p x) = poly_inf p" by (auto simp: eventually_at_top_linorder) hence "sgn (poly p (max x\<^sub>0 (a+1))) = poly_inf p" by simp with A have D: "sgn (poly p (max x\<^sub>0 (a+1))) = 1" by simp show "\x. x > a \ poly p x > 0" proof (rule ccontr) assume "\(\x. x > a \ poly p x > 0)" then obtain x' where "x' > a" "poly p x' \ 0" by (auto simp: not_less) with A and D have E: "sgn (poly p x') \ sgn (poly p (max x\<^sub>0(a+1)))" by (auto simp: sgn_real_def split: if_split_asm) show False apply (cases x' "max x\<^sub>0 (a+1)" rule: linorder_cases) using B E \x' > a\ apply (force dest!: poly_different_sign_imp_root[of _ _ p])+ done qed qed lemma poly_pos_geq: "(\x::real. x \ a \ poly p x > 0) \ poly_inf p = 1 \ (\x. x \ a \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x::real. x \ a \ poly p x > 0" hence "\x::real. x > a \ poly p x > 0" by simp also note poly_pos_greater finally have "poly_inf p = 1" "(\x>a. poly p x \ 0)" by simp_all moreover from A have "poly p a > 0" by simp ultimately show "poly_inf p = 1" "\x\a. poly p x \ 0" by (auto simp: less_eq_real_def) next assume "poly_inf p = 1 \ (\x. x \ a \ poly p x \ 0)" hence A: "poly_inf p = 1" and B: "poly p a \ 0" and C: "\x>a. poly p x \ 0" by simp_all from A and C and poly_pos_greater have "\x>a. poly p x > 0" by simp moreover with B C poly_IVT_pos[of a "a+1" p] have "poly p a > 0" by force ultimately show "\x\a. poly p x > 0" by (auto simp: less_eq_real_def) qed lemma poly_pos_less: "(\x::real. x < a \ poly p x > 0) \ poly_neg_inf p = 1 \ (\x. x < a \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x::real. x < a \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. x < a \ poly p x \ 0" by auto from poly_lim_neg_inf obtain x\<^sub>0 where "\x\x\<^sub>0. sgn (poly p x) = poly_neg_inf p" by (auto simp: eventually_at_bot_linorder) hence "poly_neg_inf p = sgn (poly p (min x\<^sub>0 (a - 1)))" by simp also from A have "... = 1" by force finally show "poly_neg_inf p = 1" . next assume "poly_neg_inf p = 1 \ (\x. x < a \ poly p x \ 0)" hence A: "poly_neg_inf p = 1" and B: "(\x. x < a \ poly p x \ 0)" by simp_all from poly_lim_neg_inf obtain x\<^sub>0 where C: "\x\x\<^sub>0. sgn (poly p x) = poly_neg_inf p" by (auto simp: eventually_at_bot_linorder) hence "sgn (poly p (min x\<^sub>0 (a - 1))) = poly_neg_inf p" by simp with A have D: "sgn (poly p (min x\<^sub>0 (a - 1))) = 1" by simp show "\x. x < a \ poly p x > 0" proof (rule ccontr) assume "\(\x. x < a \ poly p x > 0)" then obtain x' where "x' < a" "poly p x' \ 0" by (auto simp: not_less) with A and D have E: "sgn (poly p x') \ sgn (poly p (min x\<^sub>0 (a - 1)))" by (auto simp: sgn_real_def split: if_split_asm) show False apply (cases x' "min x\<^sub>0 (a - 1)" rule: linorder_cases) using B E \x' < a\ apply (auto dest!: poly_different_sign_imp_root[of _ _ p])+ done qed qed lemma poly_pos_leq: "(\x::real. x \ a \ poly p x > 0) \ poly_neg_inf p = 1 \ (\x. x \ a \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x::real. x \ a \ poly p x > 0" hence "\x::real. x < a \ poly p x > 0" by simp also note poly_pos_less finally have "poly_neg_inf p = 1" "(\x 0)" by simp_all moreover from A have "poly p a > 0" by simp ultimately show "poly_neg_inf p = 1" "\x\a. poly p x \ 0" by (auto simp: less_eq_real_def) next assume "poly_neg_inf p = 1 \ (\x. x \ a \ poly p x \ 0)" hence A: "poly_neg_inf p = 1" and B: "poly p a \ 0" and C: "\x 0" by simp_all from A and C and poly_pos_less have "\x 0" by simp moreover with B C poly_IVT_neg[of "a - 1" a p] have "poly p a > 0" by force ultimately show "\x\a. poly p x > 0" by (auto simp: less_eq_real_def) qed lemma poly_pos_between_less_less: "(\x::real. a < x \ x < b \ poly p x > 0) \ (a \ b \ poly p ((a+b)/2) > 0) \ (\x. a < x \ x < b \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x. a < x \ x < b \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. a < x \ x < b \ poly p x \ 0" by auto from A show "a \ b \ poly p ((a+b)/2) > 0" by (cases "a < b", auto) next assume "(b \ a \ 0 < poly p ((a+b)/2)) \ (\x. a x poly p x \ 0)" hence A: "b \ a \ 0 < poly p ((a+b)/2)" and B: "\x. a x poly p x \ 0" by simp_all show "\x. a < x \ x < b \ poly p x > 0" proof (cases "a \ b", simp, clarify, rule_tac ccontr, simp only: not_le not_less) fix x assume "a < b" "a < x" "x < b" "poly p x \ 0" with B have "poly p x < 0" by (simp add: less_eq_real_def) moreover from A and \a < b\ have "poly p ((a+b)/2) > 0" by simp ultimately have "sgn (poly p x) \ sgn (poly p ((a+b)/2))" by simp thus False using B apply (cases x "(a+b)/2" rule: linorder_cases) apply (drule poly_different_sign_imp_root[of _ _ p], assumption, insert \a < b\ \a < x\ \x < b\ , force) [] apply simp apply (drule poly_different_sign_imp_root[of _ _ p], simp, insert \a < b\ \a < x\ \x < b\ , force) done qed qed lemma poly_pos_between_less_leq: "(\x::real. a < x \ x \ b \ poly p x > 0) \ (a \ b \ poly p b > 0) \ (\x. a < x \ x \ b \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x. a < x \ x \ b \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. a < x \ x \ b \ poly p x \ 0" by auto from A show "a \ b \ poly p b > 0" by (cases "a < b", auto) next assume "(b \ a \ 0 < poly p b) \ (\x. a x\b \ poly p x \ 0)" hence A: "b \ a \ 0 < poly p b" and B: "\x. a x\b \ poly p x \ 0" by simp_all show "\x. a < x \ x \ b \ poly p x > 0" proof (cases "a \ b", simp, clarify, rule_tac ccontr, simp only: not_le not_less) fix x assume "a < b" "a < x" "x \ b" "poly p x \ 0" with B have "poly p x < 0" by (simp add: less_eq_real_def) moreover from A and \a < b\ have "poly p b > 0" by simp ultimately have "x < b" using \x \ b\ by (auto simp: less_eq_real_def) from \poly p x < 0\ and \poly p b > 0\ have "sgn (poly p x) \ sgn (poly p b)" by simp from poly_different_sign_imp_root[OF \x < b\ this] and B and \x > a\ show False by auto qed qed lemma poly_pos_between_leq_less: "(\x::real. a \ x \ x < b \ poly p x > 0) \ (a \ b \ poly p a > 0) \ (\x. a \ x \ x < b \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x. a \ x \ x < b \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. a \ x \ x < b \ poly p x \ 0" by auto from A show "a \ b \ poly p a > 0" by (cases "a < b", auto) next assume "(b \ a \ 0 < poly p a) \ (\x. a\x \ x poly p x \ 0)" hence A: "b \ a \ 0 < poly p a" and B: "\x. a\x \ x poly p x \ 0" by simp_all show "\x. a \ x \ x < b \ poly p x > 0" proof (cases "a \ b", simp, clarify, rule_tac ccontr, simp only: not_le not_less) fix x assume "a < b" "a \ x" "x < b" "poly p x \ 0" with B have "poly p x < 0" by (simp add: less_eq_real_def) moreover from A and \a < b\ have "poly p a > 0" by simp ultimately have "x > a" using \x \ a\ by (auto simp: less_eq_real_def) from \poly p x < 0\ and \poly p a > 0\ have "sgn (poly p a) \ sgn (poly p x)" by simp from poly_different_sign_imp_root[OF \x > a\ this] and B and \x < b\ show False by auto qed qed lemma poly_pos_between_leq_leq: "(\x::real. a \ x \ x \ b \ poly p x > 0) \ (a > b \ poly p a > 0) \ (\x. a \ x \ x \ b \ poly p x \ 0)" proof (intro iffI conjI) assume A: "\x. a \ x \ x \ b \ poly p x > 0" have "\x. poly p (x::real) > 0 \ poly p x \ 0" by simp with A show "\x::real. a \ x \ x \ b \ poly p x \ 0" by auto from A show "a > b \ poly p a > 0" by (cases "a \ b", auto) next assume "(b < a \ 0 < poly p a) \ (\x. a\x \ x\b \ poly p x \ 0)" hence A: "b < a \ 0 < poly p a" and B: "\x. a\x \ x\b \ poly p x \ 0" by simp_all show "\x. a \ x \ x \ b \ poly p x > 0" proof (cases "a > b", simp, clarify, rule_tac ccontr, simp only: not_le not_less) fix x assume "a \ b" "a \ x" "x \ b" "poly p x \ 0" with B have "poly p x < 0" by (simp add: less_eq_real_def) moreover from A and \a \ b\ have "poly p a > 0" by simp ultimately have "x > a" using \x \ a\ by (auto simp: less_eq_real_def) from \poly p x < 0\ and \poly p a > 0\ have "sgn (poly p a) \ sgn (poly p x)" by simp from poly_different_sign_imp_root[OF \x > a\ this] and B and \x \ b\ show False by auto qed qed end diff --git a/thys/Sturm_Sequences/ROOT b/thys/Sturm_Sequences/ROOT --- a/thys/Sturm_Sequences/ROOT +++ b/thys/Sturm_Sequences/ROOT @@ -1,20 +1,22 @@ chapter AFP session Sturm_Sequences (AFP) = "HOL-Computational_Algebra" + options [timeout = 600, document_logo = "_", document_variants = "document:outline=/proof,/ML:userguide"] + sessions + "Pure-ex" directories "Lib" "Examples" theories "Lib/Sturm_Library_Document" "Lib/Misc_Polynomial" theories [document = false] "Lib/Sturm_Library" theories Sturm_Theorem Sturm_Method "Examples/Sturm_Ex" document_files "root.tex" "root_userguide.tex" diff --git a/thys/Treaps/Random_List_Permutation.thy b/thys/Treaps/Random_List_Permutation.thy --- a/thys/Treaps/Random_List_Permutation.thy +++ b/thys/Treaps/Random_List_Permutation.thy @@ -1,474 +1,475 @@ (* File: Random_List.thy Authors: Manuel Eberl Some facts about randomly permuted lists and how to obtain them by drawing i.i.d. priorities for every element. *) section \Randomly-permuted lists\ theory Random_List_Permutation imports Probability_Misc Comparison_Sort_Lower_Bound.Linorder_Relations begin subsection \General facts about linear orderings\ text \ We define the set of all linear orderings on a given set and show some properties about it. \ definition linorders_on :: "'a set \ ('a \ 'a) set set" where "linorders_on A = {R. linorder_on A R}" lemma linorders_on_empty [simp]: "linorders_on {} = {{}}" by (auto simp: linorders_on_def linorder_on_def refl_on_def) lemma linorders_finite_nonempty: assumes "finite A" shows "linorders_on A \ {}" proof - from finite_distinct_list[OF assms] obtain xs where "set xs = A" "distinct xs" by blast hence "linorder_on A (linorder_of_list xs)" by auto thus ?thesis by (auto simp: linorders_on_def) qed text \ There is an obvious bijection between permutations of a set (i.\,e.\ lists with all elements from that set without repetition) and linear orderings on it. \ lemma bij_betw_linorders_on: assumes "finite A" shows "bij_betw linorder_of_list (permutations_of_set A) (linorders_on A)" using bij_betw_linorder_of_list[of A] assms unfolding linorders_on_def by simp lemma sorted_wrt_list_of_set_linorder_of_list [simp]: assumes "distinct xs" shows "sorted_wrt_list_of_set (linorder_of_list xs) (set xs) = xs" by (rule sorted_wrt_list_of_set_eqI[of "set xs"]) (insert assms, auto) lemma linorder_of_list_sorted_wrt_list_of_set [simp]: assumes "linorder_on A R" "finite A" shows "linorder_of_list (sorted_wrt_list_of_set R A) = R" proof - from assms(1) have subset: "R \ A \ A" by (auto simp: linorder_on_def refl_on_def) from assms and subset show ?thesis by (auto simp: linorder_of_list_def linorder_sorted_wrt_list_of_set sorted_wrt_linorder_index_le_iff) qed lemma bij_betw_linorders_on': assumes "finite A" shows "bij_betw (\R. sorted_wrt_list_of_set R A) (linorders_on A) (permutations_of_set A)" by (rule bij_betw_byWitness[where f' = linorder_of_list]) (insert assms, auto simp: linorders_on_def permutations_of_set_def linorder_sorted_wrt_list_of_set) lemma finite_linorders_on [intro]: assumes "finite A" shows "finite (linorders_on A)" proof - have "finite (permutations_of_set A)" by simp also have "?this \ finite (linorders_on A)" using assms by (rule bij_betw_finite [OF bij_betw_linorders_on]) finally show ?thesis . qed text \ Next, we look at the ordering defined by a list that is permuted with some permutation function. For this, we first define the composition of a relation with a function. \ definition map_relation :: "'a set \ ('a \ 'b) \ ('b \ 'b) set \ ('a \ 'a) set" where "map_relation A f R = {(x,y)\A\A. (f x, f y) \ R}" lemma index_distinct_eqI: assumes "distinct xs" "i < length xs" "xs ! i = x" shows "index xs x = i" using assms by (induction xs arbitrary: i) (auto simp: nth_Cons split: nat.splits) lemma index_permute_list: assumes "\ permutes {.. set xs" shows "index (permute_list \ xs) x = inv \ (index xs x)" proof - have *: "inv \ permutes {.. permutes {.. xs) = map_relation (set xs) ((!) xs \ inv \ \ index xs) (linorder_of_list xs)" proof - note * = permutes_inv[OF assms(1)] have less: "inv \ i < length xs" if "i < length xs" for i using permutes_in_image[OF *] and that by simp from assms and * show ?thesis by (auto simp: linorder_of_list_def map_relation_def index_nth_id index_permute_list less) qed lemma inj_on_conv_Ex1: "inj_on f A \ (\y\f`A. \!x\A. y = f x)" by (auto simp: inj_on_def) lemma bij_betw_conv_Ex1: "bij_betw f A B \ (\y\B. \!x\A. f x = y) \ B = f ` A" unfolding bij_betw_def inj_on_conv_Ex1 by (auto simp: eq_commute) lemma permutesI: assumes "bij_betw f A A" "\x. x \ A \ f x = x" shows "f permutes A" unfolding permutes_def proof (intro conjI allI impI) fix y from assms have [simp]: "f x \ A \ x \ A" for x by (auto simp: bij_betw_def) show "\!x. f x = y" proof (cases "y \ A") case True also from assms have "A = f ` A" by (auto simp: bij_betw_def) finally obtain x where "x \ A" "y = f x" by auto with assms and \y \ A\ show ?thesis by (intro ex1I[of _ x]) (auto simp: bij_betw_def dest: inj_onD) qed (insert assms, auto) qed (insert assms, auto) text \ We now show the important lemma that any two linear orderings on a finite set can be mapped onto each other by a permutation. \ lemma linorder_permutation_exists: assumes "finite A" "linorder_on A R" "linorder_on A R'" obtains \ where "\ permutes A" "R' = map_relation A \ R" proof - define xs where "xs = sorted_wrt_list_of_set R A" define ys where "ys = sorted_wrt_list_of_set R' A" have xs_ys: "distinct xs" "distinct ys" "set xs = A" "set ys = A" using assms by (simp_all add: linorder_sorted_wrt_list_of_set xs_def ys_def) from xs_ys have "mset ys = mset xs" by (simp add: set_eq_iff_mset_eq_distinct [symmetric]) - from mset_eq_permutation[OF this] guess \ . note \ = this + then obtain \ where \: "\ permutes {.. xs = ys" + by (rule mset_eq_permutation) define \' where "\' = (\x. if x \ A then x else xs ! inv \ (index xs x))" have \': "\' permutes A" proof (rule permutesI) have "bij_betw ((!) xs \ inv \) {.. bij_betw_nth)+ (simp_all add: xs_ys) hence "bij_betw ((!) xs \ inv \ \ index xs) A A" by (rule bij_betw_trans [rotated] bij_betw_index)+ (insert bij_betw_index[of xs A "length xs"], simp_all add: xs_ys atLeast0LessThan) also have "bij_betw ((!) xs \ inv \ \ index xs) A A \ bij_betw \' A A" by (rule bij_betw_cong) (auto simp: \'_def) finally show \ . qed (simp_all add: \'_def) from assms have "R' = linorder_of_list ys" by (simp add: ys_def) also from \ have "ys = permute_list \ xs" by simp also have "linorder_of_list (permute_list \ xs) = map_relation A ((!) xs \ inv \ \ index xs) (linorder_of_list xs)" using \ by (subst linorder_of_list_permute) (simp_all add: xs_ys) also from assms have "linorder_of_list xs = R" by (simp add: xs_def) finally have "R' = map_relation A ((!) xs \ inv \ \ index xs) R" . also have "\ = map_relation A \' R" by (auto simp: map_relation_def \'_def) finally show ?thesis using \' and that[of \'] by simp qed text \ We now define the linear ordering defined by some priority function, i.\,e.\ a function that injectively associates priorities to every element such that elements with lower priority are smaller in the resulting ordering. \ definition linorder_from_keys :: "'a set \ ('a \ 'b :: linorder) \ ('a \ 'a) set" where "linorder_from_keys A f = {(x,y)\A\A. f x \ f y}" lemma linorder_from_keys_permute: assumes "g permutes A" shows "linorder_from_keys A (f \ g) = map_relation A g (linorder_from_keys A f)" using permutes_in_image[OF assms] by (auto simp: map_relation_def linorder_from_keys_def) lemma linorder_on_linorder_from_keys [intro]: assumes "inj_on f A" shows "linorder_on A (linorder_from_keys A f)" using assms by (auto simp: linorder_on_def refl_on_def antisym_def linorder_from_keys_def trans_def total_on_def dest: inj_onD) lemma linorder_from_keys_empty [simp]: "linorder_from_keys {} = (\_. {})" by (simp add: linorder_from_keys_def fun_eq_iff) text \ We now show another important fact, namely that when we draw $n$ values i.\,i.\,d.\ uniformly from a non-trivial real interval, we almost surely get distinct values. \ lemma emeasure_PiM_diagonal: fixes a b :: real assumes "x \ A" "y \ A" "x \ y" assumes "a < b" "finite A" defines "M \ uniform_measure lborel {a..b}" shows "emeasure (PiM A (\_. M)) {h\A \\<^sub>E UNIV. h x = h y} = 0" proof - from assms have M: "prob_space M" unfolding M_def by (intro prob_space_uniform_measure) auto then interpret product_prob_space "\_. M" A unfolding product_prob_space_def product_prob_space_axioms_def product_sigma_finite_def by (auto simp: prob_space_imp_sigma_finite) from M interpret pair_sigma_finite M M by unfold_locales have [measurable]: "{h\extensional {x, y}. h x = h y} \ sets (Pi\<^sub>M {x, y} (\i. lborel))" proof - have "{h\extensional {x,y}. h x = h y} = {h \ space (Pi\<^sub>M {x, y} (\i. lborel)). h x = h y}" by (auto simp: extensional_def space_PiM) also have "... \ sets (Pi\<^sub>M {x, y} (\i. lborel))" by measurable finally show ?thesis . qed have [simp]: "sets (PiM A (\_. M)) = sets (PiM A (\_. lborel))" for A :: "'a set" by (intro sets_PiM_cong refl) (simp_all add: M_def) have sets_M_M: "sets (M \\<^sub>M M) = sets (borel \\<^sub>M borel)" by (intro sets_pair_measure_cong) (auto simp: M_def) have [measurable]: "(\(b, a). if b = a then 1 else 0) \ borel_measurable (M \\<^sub>M M)" unfolding measurable_split_conv by (subst measurable_cong_sets[OF sets_M_M refl]) (auto intro!: measurable_If measurable_const measurable_equality_set) have "{h\A \\<^sub>E UNIV. h x = h y} = (\h. restrict h {x, y}) -` {h\extensional {x, y}. h x = h y} \ space (PiM A (\_. M :: real measure))" by (auto simp: space_PiM PiE_def extensional_def M_def) also have "emeasure (PiM A (\_. M)) \ = emeasure (distr (PiM A (\_. M)) (PiM {x,y} (\_. lborel :: real measure)) (\h. restrict h {x,y})) {h\extensional {x, y}. h x = h y}" proof (rule emeasure_distr [symmetric]) have "(\h. restrict h {x, y}) \ Pi\<^sub>M A (\_. lborel) \\<^sub>M Pi\<^sub>M {x, y} (\_. lborel)" using assms by (intro measurable_restrict_subset) auto also have "\ = Pi\<^sub>M A (\_. M) \\<^sub>M Pi\<^sub>M {x, y} (\_. lborel)" by (intro sets_PiM_cong measurable_cong_sets refl) (simp_all add: M_def) finally show "(\h. restrict h {x, y}) \ \" . next show "{h\extensional {x, y}. h x = h y} \ sets (Pi\<^sub>M {x, y} (\_. lborel))" by simp qed also have "distr (PiM A (\_. M)) (PiM {x,y} (\_. lborel :: real measure)) (\h. restrict h {x,y}) = distr (PiM A (\_. M)) (PiM {x,y} (\_. M)) (\h. restrict h {x,y})" by (intro distr_cong refl sets_PiM_cong) (simp_all add: M_def) also from assms have "\ = Pi\<^sub>M {x, y} (\i. M)" by (intro distr_restrict [symmetric]) auto also have "emeasure \ {h\extensional {x, y}. h x = h y} = nn_integral \ (\h. indicator {h\extensional {x, y}. h x = h y} h)" by (intro nn_integral_indicator [symmetric]) simp_all also have "\ = nn_integral (Pi\<^sub>M {x, y} (\i. M)) (\h. if h x = h y then 1 else 0)" by (intro nn_integral_cong) (auto simp add: indicator_def space_PiM PiE_def) also from \x \ y\ have "\ = (\\<^sup>+ z. (if fst z = snd z then 1 else 0) \(M \\<^sub>M M))" by (intro product_nn_integral_pair) auto also have "\ = (\\<^sup>+ x. (\\<^sup>+y. (if x = y then 1 else 0) \M) \M)" by (subst M.nn_integral_fst [symmetric]) simp_all also have "\ = (\\<^sup>+ x. (\\<^sup>+y. indicator {x} y \M) \M)" by (simp add: indicator_def of_bool_def eq_commute) also have "\ = (\\<^sup>+ x. emeasure M {x} \M)" by (subst nn_integral_indicator) (simp_all add: M_def) also have "\ = (\\<^sup>+ x. 0 \M)" unfolding M_def by (intro nn_integral_cong_AE refl AE_uniform_measureI) auto also have "\ = 0" by simp finally show ?thesis . qed lemma measurable_linorder_from_keys_restrict: assumes fin: "finite A" shows "linorder_from_keys A \ Pi\<^sub>M A (\_. borel :: real measure) \\<^sub>M count_space (Pow (A \ A))" (is "_ : ?M \\<^sub>M _") apply (subst measurable_count_space_eq2) apply (auto simp add: fin linorder_from_keys_def) proof - note fin[simp] fix R assume "R \ A \ A" then have "linorder_from_keys A -` {R} \ space ?M = {f \ space ?M. \x\A. \y\A. (x, y) \ R \ f x \ f y}" by (auto simp add: linorder_from_keys_def set_eq_iff) also have "\ \ sets ?M" by measurable finally show "linorder_from_keys A -` {R} \ space ?M \ sets ?M" . qed lemma measurable_count_space_extend: assumes "f \ measurable M (count_space A)" "A \ B" shows "f \ measurable M (count_space B)" proof - note assms(1) also have "count_space A = restrict_space (count_space B) A" using assms(2) by (subst restrict_count_space) (simp_all add: Int_absorb2) finally show ?thesis by (simp add: measurable_restrict_space2_iff) qed lemma measurable_linorder_from_keys_restrict': assumes fin: "finite A" "A \ B" shows "linorder_from_keys A \ Pi\<^sub>M A (\_. borel :: real measure) \\<^sub>M count_space (Pow (B \ B))" apply (rule measurable_count_space_extend) apply (rule measurable_linorder_from_keys_restrict[OF assms(1)]) using assms by auto context fixes a b :: real and A :: "'a set" and M and B assumes fin: "finite A" and ab: "a < b" and B: "A \ B" defines "M \ distr (PiM A (\_. uniform_measure lborel {a..b})) (count_space (Pow (B \ B))) (linorder_from_keys A)" begin lemma measurable_linorder_from_keys [measurable]: "linorder_from_keys A \ Pi\<^sub>M A (\_. borel :: real measure) \\<^sub>M count_space (Pow (B \ B))" by (rule measurable_linorder_from_keys_restrict') (auto simp: fin B) text \ The ordering defined by randomly-chosen priorities is almost surely linear: \ theorem almost_everywhere_linorder: "AE R in M. linorder_on A R" proof - define N where "N = PiM A (\_. uniform_measure lborel {a..b})" have [simp]: "sets (Pi\<^sub>M A (\_. uniform_measure lborel {a..b})) = sets (PiM A (\_. lborel))" by (intro sets_PiM_cong) simp_all let ?M_A = "(Pi\<^sub>M A (\_. lborel :: real measure))" have meas: "{h \ A \\<^sub>E UNIV. h i = h j} \ sets ?M_A" if [simp]: "i \ A" "j \ A" for i j proof - have "{h \ A \\<^sub>E UNIV. h i = h j} = {h \ space ?M_A. h i = h j}" by (auto simp: space_PiM) also have "... \ sets ?M_A" by measurable finally show ?thesis . qed define X :: "('a \ real) set" where "X = (\x\A. \y\A-{x}. {h\A \\<^sub>E UNIV. h x = h y})" have "AE f in N. inj_on f A" proof (rule AE_I) show "{f \ space N. \ inj_on f A} \ X" by (auto simp: inj_on_def X_def space_PiM N_def) next show "X \ sets N" unfolding X_def N_def using meas by (auto intro!: countable_finite fin sets.countable_UN') next have "emeasure N X \ (\i\A. emeasure N (\y\A - {i}. {h \ A \\<^sub>E UNIV. h i = h y}))" unfolding X_def N_def using fin meas by (intro emeasure_subadditive_finite) (auto simp: disjoint_family_on_def intro!: sets.countable_UN' countable_finite) also have "\ \ (\i\A. \j\A-{i}. emeasure N {h \ A \\<^sub>E UNIV. h i = h j})" unfolding N_def using fin meas by (intro emeasure_subadditive_finite sum_mono) (auto simp: disjoint_family_on_def intro!: sets.countable_UN' countable_finite) also have "\ = (\i\A. \j\A-{i}. 0)" unfolding N_def using fin ab by (intro sum.cong refl emeasure_PiM_diagonal) auto also have "\ = 0" by simp finally show "emeasure N X = 0" by simp qed hence "AE f in N. linorder_on A (linorder_from_keys A f)" by eventually_elim auto thus ?thesis unfolding M_def N_def by (subst AE_distr_iff) auto qed text \ Furthermore, this is equivalent to choosing one of the $|A|!$ linear orderings uniformly at random. \ theorem random_linorder_by_prios: "M = uniform_measure (count_space (Pow (B \ B))) (linorders_on A)" proof - from linorders_finite_nonempty[OF fin] obtain R where R: "linorder_on A R" by (auto simp: linorders_on_def) have *: "emeasure M {R} \ emeasure M {R'}" if "linorder_on A R" "linorder_on A R'" for R R' proof - define N where "N = PiM A (\_. uniform_measure lborel {a..b})" from linorder_permutation_exists[OF fin that] obtain \ where \: "\ permutes A" "R' = map_relation A \ R" by blast have "(\f. f \ \) \ Pi\<^sub>M A (\_. lborel :: real measure) \\<^sub>M Pi\<^sub>M A (\_. lborel :: real measure)" by (auto intro!: measurable_PiM_single' measurable_PiM_component_rev simp: comp_def permutes_in_image[OF \(1)] space_PiM PiE_def extensional_def) also have "\ = N \\<^sub>M Pi\<^sub>M A (\_. lborel)" unfolding N_def by (intro measurable_cong_sets refl sets_PiM_cong) simp_all finally have [measurable]: "(\f. f \ \) \ \" . have [simp]: "measurable N X = measurable (PiM A (\_. lborel)) X" for X :: "('a \ 'a) set measure" unfolding N_def by (intro measurable_cong_sets refl sets_PiM_cong) simp_all have [simp]: "measurable M X = measurable (count_space (Pow (B \ B))) X" for X :: "('a \ 'a) set measure" unfolding M_def by simp have M_eq: "M = distr N (count_space (Pow (B \ B))) (linorder_from_keys A)" by (simp only: M_def N_def) also have "N = distr N (PiM A (\_. lborel)) (\f. f \ \)" unfolding N_def by (rule PiM_uniform_measure_permute [symmetric]) fact+ also have "distr \ (count_space (Pow (B \ B))) (linorder_from_keys A) = distr N (count_space (Pow (B \ B))) (linorder_from_keys A \ (\f. f \ \)) " by (intro distr_distr) simp_all also have "\ = distr N (count_space (Pow (B \ B))) (map_relation A \ \ linorder_from_keys A)" by (intro distr_cong refl) (auto simp: linorder_from_keys_permute[OF \(1)]) also have "\ = distr M (count_space (Pow (B \ B))) (map_relation A \)" unfolding M_eq using B by (intro distr_distr [symmetric]) (auto simp: map_relation_def) finally have M_eq': "distr M (count_space (Pow (B \ B))) (map_relation A \) = M" .. from that have subset': "R \ B \ B" "R' \ B \ B" using B by (auto simp: linorder_on_def refl_on_def) hence "emeasure M {R} \ emeasure M (map_relation A \ -` {R'} \ space M)" using subset' by (intro emeasure_mono) (auto simp: M_def \ ) also have "\ = emeasure (distr M (count_space (Pow (B \ B))) (map_relation A \)) {R'} " by (rule emeasure_distr [symmetric]) (insert subset' B, auto simp: map_relation_def) also note M_eq' finally show ?thesis . qed have same_prob: "emeasure M {R'} = emeasure M {R}" if "linorder_on A R'" for R' using *[of R R'] and *[of R' R] and R and that by simp from ab have "prob_space M" unfolding M_def by (intro prob_space.prob_space_distr prob_space_PiM prob_space_uniform_measure) simp_all hence "1 = emeasure M (Pow (B \ B))" using prob_space.emeasure_space_1[OF \prob_space M\] by (simp add: M_def) also have "(Pow (B \ B)) = linorders_on A \ ((Pow (B \ B))-linorders_on A)" using B by (auto simp: linorders_on_def linorder_on_def refl_on_def) also have "emeasure M \ = emeasure M (linorders_on A) + emeasure M (Pow (B \ B)-linorders_on A)" using B by (subst plus_emeasure) (auto simp: M_def linorders_on_def linorder_on_def refl_on_def) also have "emeasure M (Pow (B \ B)-linorders_on A) = 0" using almost_everywhere_linorder by (subst (asm) AE_iff_measurable) (auto simp: linorders_on_def M_def) also from fin have "emeasure M (linorders_on A) = (\R'\linorders_on A. emeasure M {R'})" using B by (intro emeasure_eq_sum_singleton) (auto simp: M_def linorders_on_def linorder_on_def refl_on_def) also have "\ = (\R'\linorders_on A. emeasure M {R})" by (rule sum.cong) (simp_all add: linorders_on_def same_prob) also from fin have "\ = fact (card A) * emeasure M {R}" by (simp add: linorders_on_def card_finite_linorders) finally have [simp]: "emeasure M {R} = inverse (fact (card A))" by (simp add: inverse_ennreal_unique) show ?thesis proof (rule measure_eqI_countable_AE') show "sets M = Pow (Pow (B \ B))" by (simp add: M_def) next from \finite A\ show "countable (linorders_on A)" by (blast intro: countable_finite) next show "AE R in uniform_measure (count_space (Pow (B \ B))) (linorders_on A). R \ linorders_on A" by (rule AE_uniform_measureI) (insert B, auto simp: linorders_on_def linorder_on_def refl_on_def) next fix R' assume R': "R' \ linorders_on A" have subset: "linorders_on A \ Pow (B \ B)" using B by (auto simp: linorders_on_def linorder_on_def refl_on_def) have "emeasure (uniform_measure (count_space (Pow (B \ B))) (linorders_on A)) {R'} = emeasure (count_space (Pow (B \ B))) (linorders_on A \ {R'}) / emeasure (count_space (Pow (B \ B))) (linorders_on A)" using R' B by (subst emeasure_uniform_measure) (auto simp: linorders_on_def linorder_on_def refl_on_def) also have "\ = 1 / emeasure (count_space (Pow (B \ B))) (linorders_on A)" using R' B by (subst emeasure_count_space) (auto simp: linorders_on_def linorder_on_def refl_on_def) also have "\ = 1 / fact (card A)" using fin finite_linorders_on[of A] by (subst emeasure_count_space [OF subset]) (auto simp: divide_ennreal [symmetric] linorders_on_def card_finite_linorders) also have "\ = emeasure M {R}" by (simp add: field_simps divide_ennreal_def) also have "\ = emeasure M {R'}" using R' by (intro same_prob [symmetric]) (auto simp: linorders_on_def) finally show "emeasure M {R'} = emeasure (uniform_measure (count_space (Pow (B \ B))) (linorders_on A)) {R'}" .. next show "linorders_on A \ Pow (B \ B)" using B by (auto simp: linorders_on_def linorder_on_def refl_on_def) qed (auto simp: M_def linorders_on_def almost_everywhere_linorder) qed end end diff --git a/thys/Twelvefold_Way/Twelvefold_Way_Core.thy b/thys/Twelvefold_Way/Twelvefold_Way_Core.thy --- a/thys/Twelvefold_Way/Twelvefold_Way_Core.thy +++ b/thys/Twelvefold_Way/Twelvefold_Way_Core.thy @@ -1,966 +1,966 @@ (* Author: Lukas Bulwahn *) section \Main Observations on Operations and Permutations\ theory Twelvefold_Way_Core imports Preliminaries begin subsection \Range Multiset\ subsubsection \Existence of a Suitable Finite Function\ lemma obtain_function: assumes "finite A" assumes "size M = card A" shows "\f. image_mset f (mset_set A) = M" using assms proof (induct arbitrary: M rule: finite_induct) case empty from this show ?case by simp next case (insert x A) from insert(1,2,4) have "size M > 0" by (simp add: card_gt_0_iff) from this obtain y where "y \# M" using gr0_implies_Suc size_eq_Suc_imp_elem by blast from insert(1,2,4) this have "size (M - {#y#}) = card A" by (simp add: Diff_insert_absorb card_Diff_singleton_if insertI1 size_Diff_submset) from insert.hyps this obtain f' where "image_mset f' (mset_set A) = M - {#y#}" by blast from this have "image_mset (f'(x := y)) (mset_set (insert x A)) = M" using \finite A\ \x \ A\ \y \# M\ by (simp add: image_mset_fun_upd) from this show ?case by blast qed lemma obtain_function_on_ext_funcset: assumes "finite A" assumes "size M = card A" shows "\f \ A \\<^sub>E set_mset M. image_mset f (mset_set A) = M" proof - obtain f where range_eq_M: "image_mset f (mset_set A) = M" using obtain_function \finite A\ \size M = card A\ by blast let ?f = "\x. if x \ A then f x else undefined" have "?f \ A \\<^sub>E set_mset M" using range_eq_M \finite A\ by auto moreover have "image_mset ?f (mset_set A) = M" using range_eq_M \finite A\ by (auto intro: multiset.map_cong0) ultimately show ?thesis by auto qed subsubsection \Existence of Permutation\ lemma image_mset_eq_implies_bij_betw: fixes f :: "'a1 \ 'b" and f' :: "'a2 \ 'b" assumes "finite A" "finite A'" assumes mset_eq: "image_mset f (mset_set A) = image_mset f' (mset_set A')" obtains bij where "bij_betw bij A A'" and "\x\A. f x = f' (bij x)" proof - from \finite A\ have [simp]: "finite {a \ A. f a = (b::'b)}" for b by auto from \finite A'\ have [simp]: "finite {a \ A'. f' a = (b::'b)}" for b by auto have "f ` A = f' ` A'" proof - have "f ` A = f ` (set_mset (mset_set A))" using \finite A\ by simp also have "\ = f' ` (set_mset (mset_set A'))" by (metis mset_eq multiset.set_map) also have "\ = f' ` A'" using \finite A'\ by simp finally show ?thesis . qed have "\b\(f ` A). \bij. bij_betw bij {a \ A. f a = b} {a \ A'. f' a = b}" proof fix b from mset_eq have "count (image_mset f (mset_set A)) b = count (image_mset f' (mset_set A')) b" by simp from this have "card {a \ A. f a = b} = card {a \ A'. f' a = b}" using \finite A\ \finite A'\ by (simp add: count_image_mset_eq_card_vimage) from this show "\bij. bij_betw bij {a \ A. f a = b} {a \ A'. f' a = b}" by (intro finite_same_card_bij) simp_all qed - hence "\bij. \b\f ` A. bij_betw (bij b) {a \ A. f a = b} {a \ A'. f' a = b}" - by (rule bchoice) - then guess bij .. note bij = this + from bchoice [OF this] + obtain bij where bij: "\b\f ` A. bij_betw (bij b) {a \ A. f a = b} {a \ A'. f' a = b}" + by auto define bij' where "bij' = (\a. bij (f a) a)" have "bij_betw bij' A A'" proof - have "disjoint_family_on (\i. {a \ A'. f' a = i}) (f ` A)" unfolding disjoint_family_on_def by auto moreover have "bij_betw (\a. bij (f a) a) {a \ A. f a = b} {a \ A'. f' a = b}" if b: "b \ f ` A" for b using bij b by (subst bij_betw_cong[where g="bij b"]) auto ultimately have "bij_betw (\a. bij (f a) a) (\b\f ` A. {a \ A. f a = b}) (\b\f ` A. {a \ A'. f' a = b})" by (rule bij_betw_UNION_disjoint) moreover have "(\b\f ` A. {a \ A. f a = b}) = A" by auto moreover have "(\b\f ` A. {a \ A'. f' a = b}) = A'" using \f ` A = f' ` A'\ by auto ultimately show "bij_betw bij' A A'" unfolding bij'_def by (subst bij_betw_cong[where g="(\a. bij (f a) a)"]) auto qed moreover from bij have "\x\A. f x = f' (bij' x)" unfolding bij'_def using bij_betwE by fastforce ultimately show ?thesis by (rule that) qed lemma image_mset_eq_implies_permutes: fixes f :: "'a \ 'b" assumes "finite A" assumes mset_eq: "image_mset f (mset_set A) = image_mset f' (mset_set A)" obtains p where "p permutes A" and "\x\A. f x = f' (p x)" proof - from assms obtain b where "bij_betw b A A" and "\x\A. f x = f' (b x)" using image_mset_eq_implies_bij_betw by blast define p where "p = (\a. if a \ A then b a else a)" have "p permutes A" proof (rule bij_imp_permutes) show "bij_betw p A A" unfolding p_def by (simp add: \bij_betw b A A\ bij_betw_cong) next fix x assume "x \ A" from this show "p x = x" unfolding p_def by simp qed moreover from \\x\A. f x = f' (b x)\ have "\x\A. f x = f' (p x)" unfolding p_def by simp ultimately show ?thesis by (rule that) qed subsection \Domain Partition\ subsubsection \Existence of a Suitable Finite Function\ lemma obtain_function_with_partition: assumes "finite A" "finite B" assumes "partition_on A P" assumes "card P \ card B" shows "\f \ A \\<^sub>E B. (\b. {x \ A. f x = b}) ` B - {{}} = P" proof - obtain g' where "bij_betw g' P (g' ` P)" and "g' ` P \ B" by (meson assms card_le_inj finite_elements inj_on_imp_bij_betw) define f where "\a. f a = (if a \ A then g' (THE X. a \ X \ X \ P) else undefined)" have "f \ A \\<^sub>E B" unfolding f_def using \g' ` P \ B\ assms(3) partition_on_the_part_mem by fastforce moreover have "(\b. {x \ A. f x = b}) ` B - {{}} = P" proof show "(\b. {x \ A. f x = b}) ` B - {{}} \ P" proof fix X assume X:"X \ (\b. {x \ A. f x = b}) ` B - {{}}" from this obtain b where "b \ B" and "X = {x' \ A. f x' = b}" by auto from this X obtain a where "a \ A" and "a \ X" and "f a = b" by blast have "(THE X. a \ X \ X \ P) \ P" using \a \ A\ \partition_on A P\ by (simp add: partition_on_the_part_mem) from \X = {x' \ A. f x' = b}\ have X_eq1: "X = {x' \ A. g' (THE X. x' \ X \ X \ P) = b}" unfolding f_def by auto also have "\ = {x' \ A. (THE X. x' \ X \ X \ P) = inv_into P g' b}" proof - { fix x' assume "x' \ A" have "(THE X. x' \ X \ X \ P) \ P" using \partition_on A P\ \x' \ A\ by (simp add: partition_on_the_part_mem) from X_eq1 \a \ X\ have "g' (THE X. a \ X \ X \ P) = b" unfolding f_def by auto from this \(THE X. a \ X \ X \ P) \ P\ have "b \ g' ` P" by auto have "(g' (THE X. x' \ X \ X \ P) = b) \ ((THE X. x' \ X \ X \ P) = inv_into P g' b)" proof - from \(THE X. x' \ X \ X \ P) \ P\ have "(g' (THE X. x' \ X \ X \ P) = b) \ (inv_into P g' (g' (THE X. x' \ X \ X \ P)) = inv_into P g' b)" using \b \ g' ` P\ by (auto intro: inv_into_injective) moreover have "inv_into P g' (g' (THE X. x' \ X \ X \ P)) = (THE X. x' \ X \ X \ P)" using \bij_betw g' P (g' ` P)\ \(THE X. x' \ X \ X \ P) \ P\ by (simp add: bij_betw_inv_into_left) ultimately show ?thesis by simp qed } from this show ?thesis by auto qed finally have X_eq: "X = {x' \ A. (THE X. x' \ X \ X \ P) = inv_into P g' b}" . moreover have "inv_into P g' b \ P" proof - from X_eq have eq: "inv_into P g' b = (THE X. a \ X \ X \ P)" using \a \ X\ \a \ A\ by auto from this show ?thesis using \(THE X. a \ X \ X \ P) \ P\ by simp qed ultimately have "X = inv_into P g' b" using partition_on_all_in_part_eq_part[OF \partition_on A P\] by blast from this \inv_into P g' b \ P\ show "X \ P" by blast qed next show "P \ (\b. {x \ A. f x = b}) ` B - {{}}" proof fix X assume "X \ P" from assms(3) this have "X \ {}" by (auto elim: partition_onE) moreover have "X \ (\b. {x \ A. f x = b}) ` B" proof show "g' X \ B" using \X \ P\ \g' ` P \ B\ by blast show "X = {x \ A. f x = g' X}" proof show "X \ {x \ A. f x = g' X}" proof fix x assume "x \ X" from this have "x \ A" using \X \ P\ assms(3) by (fastforce elim: partition_onE) have "(THE X. x \ X \ X \ P) = X" using \X \ P\ \x \ X\ assms(3) partition_on_the_part_eq by fastforce from this \x \ A\ have "f x = g' X" unfolding f_def by auto from this \x \ A\ show "x \ {x \ A. f x = g' X}" by auto qed next show "{x \ A. f x = g' X} \ X" proof fix x assume "x \ {x \ A. f x = g' X}" from this have "x \ A" and g_eq: "g' (THE X. x \ X \ X \ P) = g' X" unfolding f_def by auto from \x \ A\ have "(THE X. x \ X \ X \ P) \ P" using assms(3) by (simp add: partition_on_the_part_mem) from this g_eq have "(THE X. x \ X \ X \ P) = X" using \X \ P\ \bij_betw g' P (g' ` P)\ by (metis bij_betw_inv_into_left) from this \x \ A\ assms(3) show "x \ X" using partition_on_in_the_unique_part by fastforce qed qed qed ultimately show "X \ (\b. {x \ A. f x = b}) ` B - {{}}" by auto qed qed ultimately show ?thesis by blast qed subsubsection \Equality under Permutation Application\ lemma permutes_implies_inv_image_on_eq: assumes "p permutes B" shows "(\b. {x \ A. p (f x) = b}) ` B = (\b. {x \ A. f x = b}) ` B" proof - have "\b \ B. \x \ A. p (f x) = b \ f x = inv p b" using \p permutes B\ by (auto simp add: permutes_inverses) from this have "(\b. {x \ A. p (f x) = b}) ` B = (\b. {x \ A. f x = inv p b}) ` B" using image_cong by blast also have "\ = (\b. {x \ A. f x = b}) ` inv p ` B" by (auto simp add: image_comp) also have "\ = (\b. {x \ A. f x = b}) ` B" by (simp add: \p permutes B\ permutes_inv permutes_image) finally show ?thesis . qed subsubsection \Existence of Permutation\ lemma the_elem: assumes "f \ A \\<^sub>E B" "f' \ A \\<^sub>E B" assumes partitions_eq: "(\b. {x \ A. f x = b}) ` B - {{}} = (\b. {x \ A. f' x = b}) ` B - {{}}" assumes "x \ A" shows "the_elem (f ` {xa \ A. f' xa = f' x}) = f x" proof - from \x \ A\ have x: "x \ {x' \ A. f' x' = f' x}" by blast have "f' x \ B" using \x \ A\ \f' \ A \\<^sub>E B\ by blast from this have "{x' \ A. f' x' = f' x} \ (\b. {x \ A. f' x = b}) ` B - {{}}" using \x \ A\ by blast from this have "{x' \ A. f' x' = f' x} \ (\b. {x \ A. f x = b}) ` B - {{}}" using partitions_eq by blast from this obtain b where eq: "{x' \ A. f' x' = f' x} = {x' \ A. f x' = b}" by blast also from x this show "the_elem (f ` {x' \ A. f' x' = f' x}) = f x" by (metis (mono_tags, lifting) empty_iff mem_Collect_eq the_elem_image_unique) qed lemma the_elem_eq: assumes "f \ A \\<^sub>E B" assumes "b \ f ` A" shows "the_elem (f ` {x' \ A. f x' = b}) = b" proof - from \b \ f ` A\ obtain a where "a \ A" and "b = f a" by blast from this show "the_elem (f ` {x' \ A. f x' = b}) = b" using the_elem[OF \f \ A \\<^sub>E B\ \f \ A \\<^sub>E B\] by simp qed lemma partitions_eq_implies: assumes "f \ A \\<^sub>E B" "f' \ A \\<^sub>E B" assumes partitions_eq: "(\b. {x \ A. f x = b}) ` B - {{}} = (\b. {x \ A. f' x = b}) ` B - {{}}" assumes "x \ A" "x' \ A" assumes "f x = f x'" shows "f' x = f' x'" proof - have "f x \ B" and "x \ {a \ A. f a = f x}" and "x' \ {a \ A. f a = f x}" using \f \ A \\<^sub>E B\ \x \ A\ \x' \ A\ \f x = f x'\ by auto moreover have "{a \ A. f a = f x} \ (\b. {x \ A. f x = b}) ` B - {{}}" using \f x \ B\ \x \ {a \ A. f a = f x}\ by auto ultimately obtain b where "x \ {a \ A. f' a = b}" and "x' \ {a \ A. f' a = b}" using partitions_eq by (metis (no_types, lifting) Diff_iff imageE) from this show "f' x = f' x'" by auto qed lemma card_domain_partitions: assumes "f \ A \\<^sub>E B" assumes "finite B" shows "card ((\b. {x \ A. f x = b}) ` B - {{}}) = card (f ` A)" proof - note [simp] = the_elem_eq[OF \f \ A \\<^sub>E B\] have "bij_betw (\X. the_elem (f ` X)) ((\b. {x \ A. f x = b}) ` B - {{}}) (f ` A)" proof (rule bij_betw_imageI) show "inj_on (\X. the_elem (f ` X)) ((\b. {x \ A. f x = b}) ` B - {{}})" proof (rule inj_onI) fix X X' assume X: "X \ (\b. {x \ A. f x = b}) ` B - {{}}" assume X': "X' \ (\b. {x \ A. f x = b}) ` B - {{}}" assume eq: "the_elem (f ` X) = the_elem (f ` X')" from X obtain b where "b \ B" and X_eq: "X = {x \ A. f x = b}" by blast from X this have "b \ f ` A" using Collect_empty_eq Diff_iff image_iff insertCI by auto from X' obtain b' where "b' \ B" and X'_eq: "X' = {x \ A. f x = b'}" by blast from X' this have "b' \ f ` A" using Collect_empty_eq Diff_iff image_iff insertCI by auto from X_eq X'_eq eq \\b. b \ f ` A \ the_elem (f ` {x' \ A. f x' = b}) = b\ \b \ f ` A\ \b' \ f ` A\ have "b = b'" by auto from this show "X = X'" using X_eq X'_eq by simp qed show "(\X. the_elem (f ` X)) ` ((\b. {x \ A. f x = b}) ` B - {{}}) = f ` A" proof show "(\X. the_elem (f ` X)) ` ((\b. {x \ A. f x = b}) ` B - {{}}) \ f ` A" using \\b. b \ f ` A \ the_elem (f ` {x' \ A. f x' = b}) = b\ by auto next show "f ` A \ (\X. the_elem (f ` X)) ` ((\b. {x \ A. f x = b}) ` B - {{}})" proof fix b assume "b \ f ` A" from this have "b = the_elem (f ` {x \ A. f x = b})" using \\b. b \ f ` A \ the_elem (f ` {x' \ A. f x' = b}) = b\ by auto moreover from \b \ f ` A\ have " {x \ A. f x = b} \ (\b. {x \ A. f x = b}) ` B - {{}}" using \f \ A \\<^sub>E B\ by auto ultimately show "b \ (\X. the_elem (f ` X)) ` ((\b. {x \ A. f x = b}) ` B - {{}})" .. qed qed qed from this show ?thesis by (rule bij_betw_same_card) qed lemma partitions_eq_implies_permutes: assumes "f \ A \\<^sub>E B" "f' \ A \\<^sub>E B" assumes "finite B" assumes partitions_eq: "(\b. {x \ A. f x = b}) ` B - {{}} = (\b. {x \ A. f' x = b}) ` B - {{}}" shows "\p. p permutes B \ (\x\A. f x = p (f' x))" proof - have card_eq: "card (f' ` A) = card (f ` A)" using card_domain_partitions[OF \f \ A \\<^sub>E B\ \finite B\] using card_domain_partitions[OF \f' \ A \\<^sub>E B\ \finite B\] using partitions_eq by simp have "f' ` A \ B" "f ` A \ B" using \f \ A \\<^sub>E B\ \f' \ A \\<^sub>E B\ by auto from this card_eq have "card (B - f' ` A) = card (B - f ` A)" using \finite B\ by (auto simp add: card_Diff_subset finite_subset) from this obtain p' where "bij_betw p' (B - f' ` A) (B - f ` A)" using \finite B\ by (metis finite_same_card_bij finite_Diff) from this have "p' ` (B - f' ` A) = (B - f ` A)" by (simp add: bij_betw_imp_surj_on) define p where "\b. p b = (if b \ B then (if b \ f' ` A then the_elem (f ` {x \ A. f' x = b}) else p' b) else b)" have "\x\A. f x = p (f' x)" proof fix x assume "x \ A" from this partitions_eq have "the_elem (f ` {xa \ A. f' xa = f' x}) = f x" using the_elem[OF \f \ A \\<^sub>E B\ \f' \ A \\<^sub>E B\] by auto from this show "f x = p (f' x)" using \x \ A\ p_def \f' \ A \\<^sub>E B\ by auto qed moreover have "p permutes B" proof (rule bij_imp_permutes) let ?invp = "\b. if b \ f ` A then the_elem (f' ` {x \ A. f x = b}) else b" note [simp] = the_elem[OF \f \ A \\<^sub>E B\ \f' \ A \\<^sub>E B\ partitions_eq] show "bij_betw p B B" proof (rule bij_betw_imageI) show "p ` B = B" proof have "(\b. the_elem (f ` {x \ A. f' x = b})) ` (f' ` A) \ B" using \f \ A \\<^sub>E B\ by auto from \p' ` (B - f' ` A) = (B - f ` A)\ this show "p ` B \ B" unfolding p_def \f \ A \\<^sub>E B\ by force next show "B \ p ` B" proof fix b assume "b \ B" show "b \ p ` B" proof (cases "b \ f ` A") assume "b \ f ` A" note \p' ` (B - f' ` A) = (B - f ` A)\ from this \b \ B\ \b \ f ` A\ show ?thesis unfolding p_def by auto next assume "b \ f ` A" from this \\x\A. f x = p (f' x)\ \b \ B\ show ?thesis using \f' \ A \\<^sub>E B\ by auto qed qed qed next show "inj_on p B" proof (rule inj_onI) fix b b' assume "b \ B" "b' \ B" "p b = p b'" have "b \ f' ` A \ b' \ f' ` A" proof - have "b \ f' ` A \ p b \ f ` A" unfolding p_def using \b \ B\ \p' ` (B - f' ` A) = B - f ` A\ by auto also have "p b \ f ` A \ p b' \ f ` A" using \p b = p b'\ by simp also have "p b' \ f ` A \ b' \ f' ` A" unfolding p_def using \b' \ B\ \p' ` (B - f' ` A) = B - f ` A\ by auto finally show ?thesis . qed from this have "(b \ f' ` A \ b' \ f' ` A) \ (b \ f' ` A \ b' \ f' ` A)" by blast from this show "b = b'" proof assume "b \ f' ` A \ b' \ f' ` A" from this obtain a a' where "a \ A" "b = f' a" and "a' \ A" "b' = f' a'" by auto from this \b \ B\ \b' \ B\ have "p b = f a" "p b' = f a'" unfolding p_def by auto from this \p b = p b'\ have "f a = f a'" by simp from this have "f' a = f' a'" using partitions_eq_implies[OF \f \ A \\<^sub>E B\ \f' \ A \\<^sub>E B\ partitions_eq] using \a \ A\ \a' \ A\ by blast from this show "b = b'" using \b' = f' a'\ \b = f' a\ by simp next assume "b \ f' ` A \ b' \ f' ` A" from this \b \ B\ \b' \ B\ have "p b' = p' b'" "p b = p' b" unfolding p_def by auto from this \p b = p b'\ have "p' b = p' b'" by simp moreover have "b \ B - f' ` A" "b' \ B - f' ` A" using \b \ B\ \b' \ B\ \b \ f' ` A \ b' \ f' ` A\ by auto ultimately show "b = b'" using \bij_betw p' _ _\ by (metis bij_betw_inv_into_left) qed qed qed next fix x assume "x \ B" from this show "p x = x" using \f' \ A \\<^sub>E B\ p_def by auto qed ultimately show ?thesis by blast qed subsection \Number Partition of Range\ subsubsection \Existence of a Suitable Finite Function\ lemma obtain_partition: assumes "finite A" assumes "number_partition (card A) N" shows "\P. partition_on A P \ image_mset card (mset_set P) = N" using assms proof (induct N arbitrary: A) case empty from this have "A = {}" unfolding number_partition_def by auto from this have "partition_on A {}" by (simp add: partition_on_empty) moreover have "image_mset card (mset_set {}) = {#}" by simp ultimately show ?case by blast next case (add x N) from add.prems(2) have "0 \# add_mset x N" and "sum_mset (add_mset x N) = card A" unfolding number_partition_def by auto from this have "x \ card A" by auto from this obtain X where "X \ A" and "card X = x" using subset_with_given_card_exists by auto from this have "X \ {}" using \0 \# add_mset x N\ \finite A\ by auto have "sum_mset N = card (A - X)" using \sum_mset (add_mset x N) = card A\ \card X = x\ \X \ A\ by (metis add.commute add.prems(1) add_diff_cancel_right' card_Diff_subset infinite_super sum_mset.add_mset) from this \0 \# add_mset x N\ have "number_partition (card (A - X)) N" unfolding number_partition_def by auto from this obtain P where "partition_on (A - X) P" and eq_N: "image_mset card (mset_set P) = N" using add.hyps \finite A\ by auto from \partition_on (A - X) P\ have "finite P" using \finite A\ finite_elements by blast from \partition_on (A - X) P\ have "X \ P" using \X \ {}\ partition_onD1 by fastforce have "partition_on A (insert X P)" using \partition_on (A - X) P\ \X \ A\ \X \ {}\ by (rule partition_on_insert') moreover have "image_mset card (mset_set (insert X P)) = add_mset x N" using eq_N \card X = x\ \finite P\ \X \ P\ by simp ultimately show ?case by blast qed lemma obtain_extensional_function_from_number_partition: assumes "finite A" "finite B" assumes "number_partition (card A) N" assumes "size N \ card B" shows "\f\A \\<^sub>E B. image_mset (\X. card X) (mset_set (((\b. {x \ A. f x = b})) ` B - {{}})) = N" proof - obtain P where "partition_on A P" and eq_N: "image_mset card (mset_set P) = N" using assms obtain_partition by blast from eq_N[symmetric] \size N \ card B\ have "card P \ card B" by simp from \partition_on A P\ this obtain f where "f \ A \\<^sub>E B" and eq_P: "(\b. {x \ A. f x = b}) ` B - {{}} = P" using obtain_function_with_partition[OF \finite A\ \finite B\] by blast have "image_mset (\X. card X) (mset_set (((\b. {x \ A. f x = b})) ` B - {{}})) = N" using eq_P eq_N by simp from this \f \ A \\<^sub>E B\ show ?thesis by auto qed subsubsection \Equality under Permutation Application\ lemma permutes_implies_multiset_of_partition_cards_eq: assumes "p\<^sub>A permutes A" "p\<^sub>B permutes B" shows "image_mset card (mset_set ((\b. {x \ A. p\<^sub>B (f' (p\<^sub>A x)) = b}) ` B - {{}})) = image_mset card (mset_set ((\b. {x \ A. f' x = b}) ` B - {{}}))" proof - have "inj_on ((`) (inv p\<^sub>A)) ((\b. {x \ A. f' x = b}) ` B - {{}})" by (meson \p\<^sub>A permutes A\ inj_image_eq_iff inj_onI permutes_surj surj_imp_inj_inv) have "image_mset card (mset_set ((\b. {x \ A. p\<^sub>B (f' (p\<^sub>A x)) = b}) ` B - {{}})) = image_mset card (mset_set ((\X. inv p\<^sub>A ` X) ` ((\b. {x \ A. f' x = b}) ` B - {{}})))" proof - have "(\b. {x \ A. p\<^sub>B (f' (p\<^sub>A x)) = b}) ` B - {{}} = (\b. {x \ A. f' (p\<^sub>A x) = b}) ` B - {{}}" using permutes_implies_inv_image_on_eq[OF \p\<^sub>B permutes B\] by metis also have "\ = (\b. inv p\<^sub>A ` {x \ A. f' x = b}) ` B - {{}}" proof - have "{x \ A. f' (p\<^sub>A x) = b} = inv p\<^sub>A ` {x \ A. f' x = b}" for b proof show "{x \ A. f' (p\<^sub>A x) = b} \ inv p\<^sub>A ` {x \ A. f' x = b}" proof fix x assume "x \ {x \ A. f' (p\<^sub>A x) = b}" from this have "x \ A" "f' (p\<^sub>A x) = b" by auto moreover from this \p\<^sub>A permutes A\ have "p\<^sub>A x \ A" by (simp add: permutes_in_image) moreover from \p\<^sub>A permutes A\ have "x = inv p\<^sub>A (p\<^sub>A x)" using permutes_inverses(2) by fastforce ultimately show "x \ inv p\<^sub>A ` {x \ A. f' x = b}" by auto qed next show "inv p\<^sub>A ` {x \ A. f' x = b} \ {x \ A. f' (p\<^sub>A x) = b}" proof fix x assume "x \ inv p\<^sub>A ` {x \ A. f' x = b}" from this obtain x' where x: "x = inv p\<^sub>A x'" "x' \ A" "f' x' = b" by auto from this \p\<^sub>A permutes A\ have "x \ A" by (simp add: permutes_in_image permutes_inv) from \x = inv p\<^sub>A x'\ \f' x' = b\ have "f' (p\<^sub>A x) = b" using \p\<^sub>A permutes A\ permutes_inverses(1) by fastforce from this \x \ A\ show "x \ {x \ A. f' (p\<^sub>A x) = b}" by auto qed qed from this show ?thesis by blast qed also have "\ = (\X. inv p\<^sub>A ` X) ` ((\b. {x \ A. f' x = b}) ` B - {{}})" by auto finally show ?thesis by simp qed also have "\ = image_mset (\X. card (inv p\<^sub>A ` X)) (mset_set ((\b. {x \ A. f' x = b}) ` B - {{}}))" using \inj_on ((`) (inv p\<^sub>A)) ((\b. {x \ A. f' x = b}) ` B - {{}})\ by (simp only: image_mset_mset_set[symmetric] image_mset.compositionality) (meson comp_apply) also have "\ = image_mset card (mset_set ((\b. {x \ A. f' x = b}) ` B - {{}}))" using \p\<^sub>A permutes A\ by (simp add: card_image inj_on_inv_into permutes_surj) finally show ?thesis . qed subsubsection \Existence of Permutation\ lemma partition_implies_permutes: assumes "finite A" assumes "partition_on A P" "partition_on A P'" assumes "image_mset card (mset_set P') = image_mset card (mset_set P)" obtains p where "p permutes A" "P' = (\X. p ` X) ` P" proof - from \partition_on A P\ \partition_on A P'\ have "finite P" "finite P'" using \finite A\ finite_elements by blast+ from this \image_mset card (mset_set P') = image_mset card (mset_set P)\ obtain bij where "bij_betw bij P P'" and "\X\P. card X = card (bij X)" using image_mset_eq_implies_bij_betw by metis have "\X\P. \p'. bij_betw p' X (bij X)" proof fix X assume "X \ P" from this have "X \ A" using \partition_on A P\ partition_onD1 by fastforce from this have "finite X" using \finite A\ rev_finite_subset by blast from \X \ P\ have "bij X \ P'" using \bij_betw bij P P'\ bij_betwE by blast from this have "bij X \ A" using \partition_on A P'\ partition_onD1 by fastforce from this have "finite (bij X)" using \finite A\ rev_finite_subset by blast from \X \ P\ have "card X = card (bij X)" using \\X\P. card X = card (bij X)\ by blast from this show "\p'. bij_betw p' X (bij X)" using \finite (bij X)\ \finite X\ finite_same_card_bij by blast qed from this have "\p'. \X\P. bij_betw (p' X) X (bij X)" by metis from this obtain p' where p': "\X\P. bij_betw (p' X) X (bij X)" .. define p where "\a. p a = (if a \ A then p' (THE X. a \ X \ X \ P) a else a)" have "p permutes A" proof - have "bij_betw p A A" proof - have "disjoint_family_on bij P" proof fix X X' assume XX': "X \ P" "X' \ P" "X \ X'" from this have "bij X \ P'" "bij X' \ P'" using \bij_betw bij P P'\ bij_betwE by blast+ moreover from XX' have "bij X \ bij X'" using \bij_betw bij P P'\ by (metis bij_betw_inv_into_left) ultimately show "bij X \ bij X' = {}" using \partition_on A P'\ by (meson partition_onE) qed moreover have "bij_betw (\a. p' (THE X. a \ X \ X \ P) a) X (bij X)" if "X \ P" for X proof - from \X \ P\ have "bij_betw (p' X) X (bij X)" using \\X\P. bij_betw (p' X) X (bij X)\ by blast moreover from \X \ P\ have "\a\X. (THE X. a \ X \ X \ P) = X" using \partition_on A P\ partition_on_the_part_eq by fastforce ultimately show ?thesis by (auto intro: bij_betw_congI) qed ultimately have "bij_betw (\a. p' (THE X. a \ X \ X \ P) a) (\X\P. X) (\X\P. bij X)" by (rule bij_betw_UNION_disjoint) moreover have "(\X\P. X) = A" "(\X\P'. X) = A" using \partition_on A P\ \partition_on A P'\ partition_onD1 by auto moreover have "(\X\P. bij X) = (\X\P'. X)" using \bij_betw bij P P'\ bij_betw_imp_surj_on by force ultimately have "bij_betw (\a. p' (THE X. a \ X \ X \ P) a) A A" by simp moreover have "\a \ A. p' (THE X. a \ X \ X \ P) a = p a" unfolding p_def by auto ultimately show ?thesis by (rule bij_betw_congI) qed moreover have "p x = x" if "x \ A" for x using \x \ A\ p_def by auto ultimately show ?thesis by (rule bij_imp_permutes) qed moreover have "P' = (\X. p ` X) ` P" proof show "P' \ (\X. p ` X) ` P" proof fix X assume "X \ P'" have in_P: "the_inv_into P bij X \ P" using \X \ P'\ \bij_betw bij P P'\ bij_betwE bij_betw_the_inv_into by blast have eq_X: "bij (the_inv_into P bij X) = X" using \X \ P'\ \bij_betw bij P P'\ by (meson f_the_inv_into_f_bij_betw) have "X = p ` (the_inv_into P bij X)" proof from in_P have "the_inv_into P bij X \ A" using \partition_on A P\ partition_onD1 by fastforce have "(\a. p' (THE X. a \ X \ X \ P) a) ` the_inv_into P bij X = X" proof show "(\a. p' (THE X. a \ X \ X \ P) a) ` the_inv_into P bij X \ X" proof fix x assume "x \ (\a. p' (THE X. a \ X \ X \ P) a) ` the_inv_into P bij X" from this obtain a where a_in: "a \ the_inv_into P bij X" and x_eq: "x = p' (THE X. a \ X \ X \ P) a" by blast have "(THE X. a \ X \ X \ P) = the_inv_into P bij X" using a_in in_P \partition_on A P\ partition_on_the_part_eq by fastforce from this x_eq have x_eq: "x = p' (the_inv_into P bij X) a" by auto from this have "x \ bij (the_inv_into P bij X)" using a_in in_P bij_betwE p' by blast from this eq_X show "x \ X" by blast qed next show "X \ (\a. p' (THE X. a \ X \ X \ P) a) ` the_inv_into P bij X" proof fix x assume "x \ X" let ?X' = "the_inv_into P bij X" define x' where "x' = the_inv_into ?X' (p' ?X') x" from in_P p' eq_X have bij_betw: "bij_betw (p' ?X') ?X' X" by auto from bij_betw \x \ X\ have "x' \ ?X'" unfolding x'_def using bij_betwE bij_betw_the_inv_into by blast from this in_P have "(THE X. x' \ X \ X \ P) = ?X'" using \partition_on A P\ partition_on_the_part_eq by fastforce from this \x \ X\ have "x = p' (THE X. x' \ X \ X \ P) x'" unfolding x'_def using bij_betw f_the_inv_into_f_bij_betw by fastforce from this \x' \ ?X'\ show "x \ (\a. p' (THE X. a \ X \ X \ P) a) ` the_inv_into P bij X" .. qed qed from this \the_inv_into P bij X \ A\ show "X \ p ` the_inv_into P bij X" unfolding p_def by auto next show "p ` the_inv_into P bij X \ X" proof fix x assume "x \ p ` the_inv_into P bij X" from this obtain x' where "x = p x'" and "x' \ the_inv_into P bij X" by auto have "x' \ A" using \x' \ the_inv_into P bij X\ assms(2) in_P partition_onD1 by fastforce have eq: "(THE X. x' \ X \ X \ P) = the_inv_into P bij X" using \x' \ the_inv_into P bij X\ assms(2) in_P partition_on_the_part_eq by fastforce have p': "p' (the_inv_into P bij X) x' \ X" using \x' \ the_inv_into P bij X\ bij_betwE eq_X in_P p' by blast from \x = p x'\ \x' \ A\ eq p' show "x \ X" unfolding p_def by auto qed qed moreover from \X \ P'\ \bij_betw bij P P'\ have "the_inv_into P bij X \ P" using bij_betwE bij_betw_the_inv_into by blast ultimately show "X \ (\X. p ` X) ` P" .. qed next show "(\X. p ` X) ` P \ P'" proof fix X' assume "X' \ (\X. p ` X) ` P" from this obtain X where X'_eq: "X' = p ` X" and "X \ P" .. from \X \ P\ have "X \ A" using assms(2) partition_onD1 by force from \X \ P\ p' have bij: "bij_betw (p' X) X (bij X)" by auto have "p ` X \ P'" proof - from \X \ P\ \bij_betw bij P P'\ have "bij X \ P'" using bij_betwE by blast moreover have "(\a. p' (THE X. a \ X \ X \ P) a) ` X = bij X" proof show "(\a. p' (THE X. a \ X \ X \ P) a) ` X \ bij X" proof fix x' assume "x' \ (\a. p' (THE X. a \ X \ X \ P) a) ` X" from this obtain x where "x \ X" and x'_eq: "x' = p' (THE X. x \ X \ X \ P) x" .. from \X \ P\ \x \ X\ have eq_X: "(THE X. x \ X \ X \ P) = X" using assms(2) partition_on_the_part_eq by fastforce from bij \x \ X\ x'_eq eq_X show "x' \ bij X" using bij_betwE by blast qed next show "bij X \ (\a. p' (THE X. a \ X \ X \ P) a) ` X" proof fix x' assume "x' \ bij X" let ?x = "inv_into X (p' X) x'" from \x' \ bij X\ bij have "?x \ X" by (metis bij_betw_imp_surj_on inv_into_into) from this \X \ P\ have "(THE X. ?x \ X \ X \ P) = X" using assms(2) partition_on_the_part_eq by fastforce from this \x' \ bij X\ bij have "x' = p' (THE X. ?x \ X \ X \ P) ?x" using bij_betw_inv_into_right by fastforce moreover from \x' \ bij X\ bij have "?x \ X" by (metis bij_betw_imp_surj_on inv_into_into) ultimately show "x' \ (\a. p' (THE X. a \ X \ X \ P) a) ` X" .. qed qed ultimately have "(\a. p' (THE X. a \ X \ X \ P) a) ` X \ P'" by simp have "(\a. p' (THE X. a \ X \ X \ P) a) ` X = (\a. if a \ A then p' (THE X. a \ X \ X \ P) a else a) ` X " using \X \ A\ by (auto intro: image_cong) from this show ?thesis using \(\a. p' (THE X. a \ X \ X \ P) a) ` X \ P'\ unfolding p_def by auto qed from this X'_eq show "X' \ P'" by simp qed qed ultimately show thesis using that by blast qed lemma permutes_domain_partition_eq: assumes "f \ A \ B" assumes "p\<^sub>A permutes A" assumes "b \ B" shows "p\<^sub>A ` {x \ A. f x = b} = {x \ A. f (inv p\<^sub>A x) = b}" proof show "p\<^sub>A ` {x \ A. f x = b} \ {x \ A. f (inv p\<^sub>A x) = b}" using \p\<^sub>A permutes A\ permutes_in_image permutes_inverses(2) by fastforce next show "{x \ A. f (inv p\<^sub>A x) = b} \ p\<^sub>A ` {x \ A. f x = b}" proof fix x assume "x \ {x \ A. f (inv p\<^sub>A x) = b}" from this have "x \ A" "f (inv p\<^sub>A x) = b" by auto from \x \ A\ have "x = p\<^sub>A (inv p\<^sub>A x)" using \p\<^sub>A permutes A\ permutes_inverses(1) by fastforce moreover from \f (inv p\<^sub>A x) = b\ \x \ A\ have "inv p\<^sub>A x \ {x \ A. f x = b}" by (simp add: \p\<^sub>A permutes A\ permutes_in_image permutes_inv) ultimately show "x \ p\<^sub>A ` {x \ A. f x = b}" .. qed qed lemma image_domain_partition_eq: assumes "f \ A \\<^sub>E B" assumes "p\<^sub>A permutes A" shows "(\X. p\<^sub>A ` X) ` ((\b. {x \ A. f x = b}) ` B) = (\b. {x \ A. f (inv p\<^sub>A x) = b}) ` B" proof from \f \ A \\<^sub>E B\ have "f \ A \ B" by auto note eq = permutes_domain_partition_eq[OF \f \ A \ B\ \p\<^sub>A permutes A\] show "(\X. p\<^sub>A ` X) ` (\b. {x \ A. f x = b}) ` B \ (\b. {x \ A. f (inv p\<^sub>A x) = b}) ` B" proof fix X assume "X \ (\X. p\<^sub>A ` X) ` (\b. {x \ A. f x = b}) ` B" from this obtain b where "b \ B" and X_eq: "X = p\<^sub>A ` {x \ A. f x = b}" by auto from this eq have "X = {x \ A. f (inv p\<^sub>A x) = b}" by simp from this \b \ B\ show "X \ (\b. {x \ A. f (inv p\<^sub>A x) = b}) ` B" .. qed next from \f \ A \\<^sub>E B\ have "f \ A \ B" by auto note eq = permutes_domain_partition_eq[OF \f \ A \ B\ \p\<^sub>A permutes A\, symmetric] show "(\b. {x \ A. f (inv p\<^sub>A x) = b}) ` B \ (\X. p\<^sub>A ` X) ` (\b. {x \ A. f x = b}) ` B" proof fix X assume "X \ (\b. {x \ A. f (inv p\<^sub>A x) = b}) ` B" from this obtain b where "b \ B" and X_eq: "X = {x \ A. f (inv p\<^sub>A x) = b}" by auto from this eq have "X = p\<^sub>A ` {x \ A. f x = b}" by simp from this \b \ B\ show "X \ (\X. p\<^sub>A ` X) ` (\b. {x \ A. f x = b}) ` B" by auto qed qed lemma multiset_of_partition_cards_eq_implies_permutes: assumes "finite A" "finite B" "f \ A \\<^sub>E B" "f' \ A \\<^sub>E B" assumes eq: "image_mset card (mset_set ((\b. {x \ A. f x = b}) ` B - {{}})) = image_mset card (mset_set ((\b. {x \ A. f' x = b}) ` B - {{}}))" obtains p\<^sub>A p\<^sub>B where "p\<^sub>A permutes A" "p\<^sub>B permutes B" "\x\A. f x = p\<^sub>B (f' (p\<^sub>A x))" proof - have "partition_on A ((\b. {x \ A. f x = b}) ` B - {{}})" using \f \ A \\<^sub>E B\ by (auto intro!: partition_onI) moreover have "partition_on A ((\b. {x \ A. f' x = b}) ` B - {{}})" using \f' \ A \\<^sub>E B\ by (auto intro!: partition_onI) moreover note partition_implies_permutes[OF \finite A\ _ _ eq] ultimately obtain p\<^sub>A where "p\<^sub>A permutes A" and inv_image_eq: "(\b. {x \ A. f x = b}) ` B - {{}} = (`) p\<^sub>A ` ((\b. {x \ A. f' x = b}) ` B - {{}})" by blast from \p\<^sub>A permutes A\ have "inj ((`) p\<^sub>A)" by (meson injI inj_image_eq_iff permutes_inj) have inv_image_eq': "(\b. {x \ A. f x = b}) ` B - {{}} = (\b. {x \ A. f' (inv p\<^sub>A x) = b}) ` B - {{}}" proof - note inv_image_eq also have "(\X. p\<^sub>A ` X) ` ((\b. {x \ A. f' x = b}) ` B - {{}}) = (\b. {x \ A. f' (inv p\<^sub>A x) = b}) ` B - {{}}" using image_domain_partition_eq[OF \f' \ A \\<^sub>E B\ \p\<^sub>A permutes A\] by (simp add: image_set_diff[OF \inj ((`) p\<^sub>A)\]) finally show ?thesis . qed from \p\<^sub>A permutes A\ have "inv p\<^sub>A permutes A" using permutes_inv by blast have "(\x. f' (inv p\<^sub>A x)) \ A \\<^sub>E B" using \f' \ A \\<^sub>E B\ \inv p\<^sub>A permutes A\ permutes_in_image by fastforce from \f \ A \\<^sub>E B\ this \finite B\ obtain p\<^sub>B where "p\<^sub>B permutes B" and eq'': "\x\A. f x = p\<^sub>B (f' (inv p\<^sub>A x))" using partitions_eq_implies_permutes[OF _ _ _ inv_image_eq'] by blast from \inv p\<^sub>A permutes A\ \p\<^sub>B permutes B\ eq'' that show thesis by blast qed subsection \Bijections on Same Domain and Range\ subsubsection \Existence of Domain Permutation\ lemma obtain_domain_permutation_for_two_bijections: assumes "bij_betw f A B" "bij_betw f' A B" obtains p where "p permutes A" and "\a\A. f a = f' (p a)" proof - let ?p = "\a. if a \ A then the_inv_into A f' (f a) else a" have "?p permutes A" proof (rule bij_imp_permutes) show "bij_betw ?p A A" proof (rule bij_betw_imageI) show "inj_on ?p A" proof (rule inj_onI) fix a a' assume "a \ A" "a' \ A" "?p a = ?p a'" from this have "the_inv_into A f' (f a) = the_inv_into A f' (f a')" using \a \ A\ \a' \ A\ by simp from this have "f a = f a'" using \a \ A\ \a' \ A\ assms by (metis bij_betwE f_the_inv_into_f_bij_betw) from this show "a = a'" using \a \ A\ \a' \ A\ assms by (metis bij_betw_inv_into_left) qed next show "?p ` A = A" proof show "?p ` A \ A" proof fix a assume "a \ ?p ` A" from this obtain a' where "a' \ A" and "a = the_inv_into A f' (f a')" by auto from this assms show "a \ A" by (metis bij_betwE bij_betw_imp_inj_on bij_betw_imp_surj_on subset_iff the_inv_into_into) qed next show "A \ ?p ` A" proof fix a assume "a \ A" from this assms have "the_inv_into A f (f' a) \ A" by (meson bij_betwE bij_betw_the_inv_into) moreover from \a \ A\ assms have "a = the_inv_into A f' (f (the_inv_into A f (f' a)))" by (metis bij_betwE bij_betw_imp_inj_on f_the_inv_into_f_bij_betw the_inv_into_f_eq) ultimately show "a \ ?p ` A" by auto qed qed qed next fix a assume "a \ A" from this show "?p a = a" by auto qed moreover have "\a\A. f a = f' (?p a)" using \bij_betw f A B\ \bij_betw f' A B\ using bij_betwE f_the_inv_into_f_bij_betw by fastforce moreover note that ultimately show thesis by auto qed subsubsection \Existence of Range Permutation\ lemma obtain_range_permutation_for_two_bijections: assumes "bij_betw f A B" "bij_betw f' A B" obtains p where "p permutes B" and "\a\A. f a = p (f' a)" proof - let ?p = "\b. if b \ B then f (inv_into A f' b) else b" have "?p permutes B" proof (rule bij_imp_permutes) show "bij_betw ?p B B" proof (rule bij_betw_imageI) show "inj_on ?p B" proof (rule inj_onI) fix b b' assume "b \ B" "b' \ B" "?p b = ?p b'" from this have "f (inv_into A f' b) = f (inv_into A f' b')" using \b \ B\ \b' \ B\ by simp from this have "inv_into A f' b = inv_into A f' b'" using \b \ B\ \b' \ B\ assms by (metis bij_betw_imp_surj_on bij_betw_inv_into_left inv_into_into) from this show "b = b'" using \b \ B\ \b' \ B\ assms(2) by (metis bij_betw_inv_into_right) qed next show "?p ` B = B" proof from assms show "?p ` B \ B" by (auto simp add: bij_betwE bij_betw_def inv_into_into) next show "B \ ?p ` B" proof fix b assume "b \ B" from this assms have "f' (inv_into A f b) \ B" by (metis bij_betwE bij_betw_imp_surj_on inv_into_into) moreover have "b = ?p (f' (inv_into A f b))" using assms \f' (inv_into A f b) \ B\ \b \ B\ by (auto simp add: bij_betw_imp_surj_on bij_betw_inv_into_left bij_betw_inv_into_right inv_into_into) ultimately show "b \ ?p ` B" by auto qed qed qed next fix b assume "b \ B" from this show "?p b = b" by auto qed moreover have "\a\A. f a = ?p (f' a)" using \bij_betw f' A B\ bij_betw_inv_into_left bij_betwE by fastforce moreover note that ultimately show thesis by auto qed end diff --git a/thys/UpDown_Scheme/Triangular_Function.thy b/thys/UpDown_Scheme/Triangular_Function.thy --- a/thys/UpDown_Scheme/Triangular_Function.thy +++ b/thys/UpDown_Scheme/Triangular_Function.thy @@ -1,422 +1,422 @@ section \ Hat Functions \ theory Triangular_Function imports "HOL-Analysis.Equivalence_Lebesgue_Henstock_Integration" Grid begin lemma continuous_on_max[continuous_intros]: fixes f :: "_ \ 'a::linorder_topology" shows "continuous_on S f \ continuous_on S g \ continuous_on S (\x. max (f x) (g x))" by (auto simp: continuous_on_def intro: tendsto_max) definition \ :: "(nat \ int) \ real \ real" where "\ \ (\(l,i) x. max 0 (1 - \ x * 2^(l + 1) - real_of_int i \))" definition \ :: "(nat \ int) list \ (nat \ real) \ real" where "\ p x = (\d (p ! d) (x d))" definition l2_\ where "l2_\ p1 p2 = (\x. \ p1 x * \ p2 x \lborel)" definition l2 where "l2 a b = (\x. \ a x * \ b x \(\\<^sub>M d\{..[measurable]: "\ p \ borel_measurable borel" by (cases p) (simp add: \_def) lemma \_nonneg: "0 \ \ p x" by (simp add: \_def split: prod.split) lemma \_zero_iff: "\ (l,i) x = 0 \ x \ {real_of_int (i - 1) / 2^(l + 1) <..< real_of_int (i + 1) / 2^(l + 1)}" by (auto simp: \_def field_simps split: split_max) lemma \_zero: "x \ {real_of_int (i - 1) / 2^(l + 1) <..< real_of_int (i + 1) / 2^(l + 1)} \ \ (l,i) x = 0" unfolding \_zero_iff by simp lemma \_eq_0: assumes x: "x < 0 \ 1 < x" and i: "0 < i" "i < 2^Suc l" shows "\ (l, i) x = 0" using x proof assume "x < 0" also have "0 \ real_of_int (i - 1) / 2^(l + 1)" using i by (auto simp: field_simps) finally show ?thesis by (auto intro!: \_zero simp: field_simps) next have "real_of_int (i + 1) / 2^(l + 1) \ 1" using i by (subst divide_le_eq_1_pos) (auto simp del: of_int_add power_Suc) also assume "1 < x" finally show ?thesis by (auto intro!: \_zero simp: field_simps) qed lemma ix_lt: "p \ sparsegrid dm lm \ d < dm \ ix p d < 2^(lv p d + 1)" unfolding sparsegrid_def lgrid_def using grid_estimate[of d "start dm" p "{0 ..< dm}"] by auto lemma ix_gt: "p \ sparsegrid dm lm \ d < dm \ 0 < ix p d" unfolding sparsegrid_def lgrid_def using grid_estimate[of d "start dm" p "{0 ..< dm}"] by auto lemma \_eq_0: assumes x: "\d 1 < x d" and p: "p \ sparsegrid dm lm" shows "\ p x = 0" unfolding \_def proof (rule prod_zero) - from x guess d .. + from x obtain d where "d < length p \ (x d < 0 \ 1 < x d)" .. with p[THEN ix_lt, of d] p[THEN ix_gt, of d] p show "\a\{.. (p ! a) (x a) = 0" apply (cases "p!d") apply (intro bexI[of _ d]) apply (auto intro!: \_eq_0 simp: sparsegrid_length ix_def lv_def) done qed simp lemma \_left_support': "x \ {real_of_int (i - 1) / 2^(l + 1) .. real_of_int i / 2^(l + 1)} \ \ (l,i) x = 1 + x * 2^(l + 1) - real_of_int i" by (auto simp: \_def field_simps split: split_max) lemma \_left_support: "x \ {-1 .. 0::real} \ \ (l,i) ((x + real_of_int i) / 2^(l + 1)) = 1 + x" by (auto simp: \_def field_simps split: split_max) lemma \_right_support': "x \ {real_of_int i / 2^(l + 1) .. real_of_int (i + 1) / 2^(l + 1)} \ \ (l,i) x = 1 - x * 2^(l + 1) + real_of_int i" by (auto simp: \_def field_simps split: split_max) lemma \_right_support: "x \ {0 .. 1::real} \ \ (l,i) ((x + real i) / 2^(l + 1)) = 1 - x" by (auto simp: \_def field_simps split: split_max) lemma integrable_\: "integrable lborel (\ p)" proof (induct p) case (Pair l i) have "integrable lborel (\x. indicator {real_of_int (i - 1) / 2^(l + 1) .. real_of_int (i + 1) / 2^(l + 1)} x *\<^sub>R \ (l, i) x)" unfolding \_def by (intro borel_integrable_compact) (auto intro!: continuous_intros) then show ?case by (rule integrable_cong[THEN iffD1, rotated -1]) (auto simp: \_zero_iff) qed lemma integrable_\2: "integrable lborel (\x. \ p x * \ q x)" proof (cases p q rule: prod.exhaust[case_product prod.exhaust]) case (Pair_Pair l i l' i') have "integrable lborel (\x. indicator {real_of_int (i - 1) / 2^(l + 1) .. real_of_int (i + 1) / 2^(l + 1)} x *\<^sub>R (\ (l, i) x * \ (l', i') x))" unfolding \_def by (intro borel_integrable_compact) (auto intro!: continuous_intros) then show ?thesis unfolding Pair_Pair by (rule integrable_cong[THEN iffD1, rotated -1]) (auto simp: \_zero_iff) qed lemma l2_\I_DERIV: assumes n: "\ x. x \ { (real_of_int i' - 1) / 2^(l' + 1) .. real_of_int i' / 2^(l' + 1) } \ DERIV \_n x :> (\ (l', i') x * \ (l, i) x)" (is "\ x. x \ {?a..?b} \ DERIV _ _ :> ?P x") and p: "\ x. x \ { real_of_int i' / 2^(l' + 1) .. (real_of_int i' + 1) / 2^(l' + 1) } \ DERIV \_p x :> (\ (l', i') x * \ (l, i) x)" (is "\ x. x \ {?b..?c} \ _") shows "l2_\ (l', i') (l, i) = (\_n ?b - \_n ?a) + (\_p ?c - \_p ?b)" proof - have "has_bochner_integral lborel (\x. ?P x * indicator {?a..?b} x + ?P x * indicator {?b..?c} x) ((\_n ?b - \_n ?a) + (\_p ?c - \_p ?b))" by (intro has_bochner_integral_add has_bochner_integral_FTC_Icc_nonneg n p) (auto simp: \_nonneg field_simps) then have "has_bochner_integral lborel?P ((\_n ?b - \_n ?a) + (\_p ?c - \_p ?b))" by (rule has_bochner_integral_discrete_difference[where X="{?b}", THEN iffD1, rotated -1]) (auto simp: power_add intro!: \_zero integral_cong split: split_indicator) then show ?thesis by (simp add: has_bochner_integral_iff l2_\_def) qed lemma l2_eq: "length a = length b \ l2 a b = (\d (a!d) (b!d))" unfolding l2_def l2_\_def \_def apply (simp add: prod.distrib[symmetric]) proof (rule product_sigma_finite.product_integral_prod) show "product_sigma_finite (\d. lborel)" .. qed (auto intro: integrable_\2) lemma l2_when_disjoint: assumes "l \ l'" defines "d == l' - l" assumes "(i + 1) * 2^d < i' \ i' < (i - 1) * 2^d" (is "?right \ ?left") shows "l2_\ (l', i') (l, i) = 0" proof - let ?sup = "\l i. {real_of_int (i - 1) / 2^(l + 1) <..< real_of_int (i + 1) / 2^(l + 1)}" have l': "l' = l + d" using assms by simp have *: "\i l. 2 ^ l = real_of_int (2 ^ l::int)" by simp have [arith]: "0 < (2^d::int)" by simp from \?right \ ?left\ \l \ l'\ have empty_support: "?sup l i \ ?sup l' i' = {}" by (auto simp add: min_def max_def divide_simps l' power_add * of_int_mult[symmetric] simp del: of_int_diff of_int_add of_int_mult of_int_power) (simp_all add: field_simps) then have "\x. \ (l', i') x * \ (l, i) x = 0" unfolding \_zero_iff mult_eq_0_iff by blast then show ?thesis by (simp add: l2_\_def del: mult_eq_0_iff vector_space_over_itself.scale_eq_0_iff) qed lemma l2_commutative: "l2_\ p q = l2_\ q p" by (simp add: l2_\_def mult.commute) lemma l2_when_same: "l2_\ (l, i) (l, i) = 1/3 / 2^l" proof (subst l2_\I_DERIV) let ?l = "(2 :: real)^(l + 1)" let ?in = "real_of_int i - 1" let ?ip = "real_of_int i + 1" let ?\ = "\ (l,i)" let ?\2 = "\x. ?\ x * ?\ x" { fix x assume "x \ {?in / ?l .. real_of_int i / ?l}" hence \_eq: "?\ x = ?l * x - ?in" using \_left_support' by auto show "DERIV (\x. x^3 / 3 * ?l^2 + x * ?in^2 - x^2/2 * 2 * ?l * ?in) x :> ?\2 x" by (auto intro!: derivative_eq_intros simp add: power2_eq_square field_simps \_eq) } { fix x assume "x \ {real_of_int i / ?l .. ?ip / ?l}" hence \_eq: "?\ x = ?ip - ?l * x" using \_right_support' by auto show "DERIV (\x. x^3 / 3 * ?l^2 + x * ?ip^2 - x^2/2 * 2 * ?l * ?ip) x :> ?\2 x" by (auto intro!: derivative_eq_intros simp add: power2_eq_square field_simps \_eq) } qed (simp_all add: field_simps power_eq_if[of _ 2] power_eq_if[of _ 3]) lemma l2_when_left_child: assumes "l < l'" and i'_bot: "i' > (i - 1) * 2^(l' - l)" and i'_top: "i' < i * 2^(l' - l)" shows "l2_\ (l', i') (l, i) = (1 + real_of_int i' / 2^(l' - l) - real_of_int i) / 2^(l' + 1)" proof (subst l2_\I_DERIV) let ?l' = "(2 :: real)^(l' + 1)" let ?in' = "real_of_int i' - 1" let ?ip' = "real_of_int i' + 1" let ?l = "(2 :: real)^(l + 1)" let ?i = "real_of_int i - 1" let ?\' = "\ (l',i')" let ?\ = "\ (l, i)" let "?\2 x" = "?\' x * ?\ x" define \_n where "\_n x = x^3 / 3 * ?l' * ?l + x * ?i * ?in' - x^2 / 2 * (?in' * ?l + ?i * ?l')" for x define \_p where "\_p x = x^2 / 2 * (?ip' * ?l + ?i * ?l') - x^3 / 3 * ?l' * ?l - x * ?i * ?ip'" for x have level_diff: "2^(l' - l) = 2^l' / (2^l :: real)" using power_diff[of "2::real" l l'] \l < l'\ by auto { fix x assume x: "x \ {?in' / ?l' .. ?ip' / ?l'}" have "?i * 2^(l' - l) \ ?in'" using i'_bot int_less_real_le by auto hence "?i / ?l \ ?in' / ?l'" using level_diff by (auto simp: field_simps) hence "?i / ?l \ x" using x by auto moreover have "?ip' \ real_of_int i * 2^(l' - l)" using i'_top int_less_real_le by auto hence ip'_le_i: "?ip' / ?l' \ real_of_int i / ?l" using level_diff by (auto simp: field_simps) hence "x \ real_of_int i / ?l" using x by auto ultimately have "?\ x = ?l * x - ?i" using \_left_support' by auto } note \_eq = this { fix x assume x: "x \ {?in' / ?l' .. real_of_int i' / ?l'}" hence \'_eq: "?\' x = ?l' * x - ?in'" using \_left_support' by auto from x have x': "x \ {?in' / ?l' .. ?ip' / ?l'}" by (auto simp add: field_simps) show "DERIV \_n x :> ?\2 x" unfolding \_eq[OF x'] \'_eq \_n_def by (auto intro!: derivative_eq_intros simp add: power2_eq_square algebra_simps) } { fix x assume x: "x \ {real_of_int i' / ?l' .. ?ip' / ?l'}" hence \'_eq: "?\' x = ?ip' - ?l' * x" using \_right_support' by auto from x have x': "x \ {?in' / ?l' .. ?ip' / ?l'}" by (simp add: field_simps) show "DERIV \_p x :> ?\2 x" unfolding \_eq[OF x'] \'_eq \_p_def by (auto intro!: derivative_eq_intros simp add: power2_eq_square algebra_simps) } qed (simp_all add: field_simps power_eq_if[of _ 2] power_eq_if[of _ 3] power_diff[of "2::real", OF _ \l < l'\[THEN less_imp_le]] ) lemma l2_when_right_child: assumes "l < l'" and i'_bot: "i' > i * 2^(l' - l)" and i'_top: "i' < (i + 1) * 2^(l' - l)" shows "l2_\ (l', i') (l, i) = (1 - real_of_int i' / 2^(l' - l) + real_of_int i) / 2^(l' + 1)" proof (subst l2_\I_DERIV) let ?l' = "(2 :: real)^(l' + 1)" let ?in' = "real_of_int i' - 1" let ?ip' = "real_of_int i' + 1" let ?l = "(2 :: real)^(l + 1)" let ?i = "real_of_int i + 1" let ?\' = "\ (l',i')" let ?\ = "\ (l, i)" let ?\2 = "\x. ?\' x * ?\ x" define \_n where "\_n x = x^2 / 2 * (?in' * ?l + ?i * ?l') - x^3 / 3 * ?l' * ?l - x * ?i * ?in'" for x define \_p where "\_p x = x^3 / 3 * ?l' * ?l + x * ?i * ?ip' - x^2 / 2 * (?ip' * ?l + ?i * ?l')" for x have level_diff: "2^(l' - l) = 2^l' / (2^l :: real)" using power_diff[of "2::real" l l'] \l < l'\ by auto { fix x assume x: "x \ {?in' / ?l' .. ?ip' / ?l'}" have "real_of_int i * 2^(l' - l) \ ?in'" using i'_bot int_less_real_le by auto hence "real_of_int i / ?l \ ?in' / ?l'" using level_diff by (auto simp: field_simps) hence "real_of_int i / ?l \ x" using x by auto moreover have "?ip' \ ?i * 2^(l' - l)" using i'_top int_less_real_le by auto hence ip'_le_i: "?ip' / ?l' \ ?i / ?l" using level_diff by (auto simp: field_simps) hence "x \ ?i / ?l" using x by auto ultimately have "?\ x = ?i - ?l * x" using \_right_support' by auto } note \_eq = this { fix x assume x: "x \ {?in' / ?l' .. real_of_int i' / ?l'}" hence \'_eq: "?\' x = ?l' * x - ?in'" using \_left_support' by auto from x have x': "x \ {?in' / ?l' .. ?ip' / ?l'}" by (simp add: field_simps) show "DERIV \_n x :> ?\2 x" unfolding \_n_def \_eq[OF x'] \'_eq by (auto intro!: derivative_eq_intros simp add: simp add: power2_eq_square algebra_simps) } { fix x assume x: "x \ {real_of_int i' / ?l' .. ?ip' / ?l'}" hence \'_eq: "?\' x = ?ip' - ?l' * x" using \_right_support' by auto from x have x': "x \ {?in' / ?l' .. ?ip' / ?l'}" by (auto simp: field_simps) show "DERIV \_p x :> ?\2 x" unfolding \_eq[OF x'] \'_eq \_p_def by (auto intro!: derivative_eq_intros simp add: power2_eq_square algebra_simps) } qed (simp_all add: field_simps power_eq_if[of _ 2] power_eq_if[of _ 3] power_diff[of "2::real", OF _ \l < l'\[THEN less_imp_le]] ) lemma level_shift: "lc > l \ (x :: real) / 2 ^ (lc - Suc l) = x * 2 / 2 ^ (lc - l)" by (auto simp add: power_diff) lemma l2_child: assumes "d < length b" and p_grid: "p \ grid (child b dir d) ds" (is "p \ grid ?child ds") shows "l2_\ (p ! d) (b ! d) = (1 - real_of_int (sgn dir) * (real_of_int (ix p d) / 2^(lv p d - lv b d) - real_of_int (ix b d))) / 2^(lv p d + 1)" proof - have "lv ?child d \ lv p d" using \d < length b\ and p_grid using grid_single_level by auto hence "lv b d < lv p d" using \d < length b\ and p_grid using child_lv by auto let ?i_c = "ix ?child d" and ?l_c = "lv ?child d" let ?i_p = "ix p d" and ?l_p = "lv p d" let ?i_b = "ix b d" and ?l_b = "lv b d" have "(2::int) * 2^(?l_p - ?l_c) = 2^Suc (?l_p - ?l_c)" by auto also have "\ = 2^(Suc ?l_p - ?l_c)" proof - have "Suc (?l_p - ?l_c) = Suc ?l_p - ?l_c" using \lv ?child d \ lv p d\ by auto thus ?thesis by auto qed also have "\ = 2^(?l_p - ?l_b)" using \d < length b\ and \lv b d < lv p d\ by (auto simp add: child_def lv_def) finally have level: "2^(?l_p - ?l_b) = (2::int) * 2^(?l_p - ?l_c)" .. from \d < length b\ and p_grid have range_left: "?i_p > (?i_c - 1) * 2^(?l_p - ?l_c)" and range_right: "?i_p < (?i_c + 1) * 2^(?l_p - ?l_c)" using grid_estimate by auto show ?thesis proof (cases dir) case left with child_ix_left[OF \d < length b\] have "(?i_b - 1) * 2^(?l_p - ?l_b) = (?i_c - 1) * 2^(?l_p - ?l_c)" and "?i_b * 2^(?l_p - ?l_b) = (?i_c + 1) * 2^(?l_p - ?l_c)" using level by auto hence "?i_p > (?i_b - 1) * 2^(?l_p - ?l_b)" and "?i_p < ?i_b * 2^(?l_p - ?l_b)" using range_left and range_right by auto with \?l_b < ?l_p\ have "l2_\ (?l_p, ?i_p) (?l_b, ?i_b) = (1 + real_of_int ?i_p / 2^(?l_p - ?l_b) - real_of_int ?i_b) / 2^(?l_p + 1)" by (rule l2_when_left_child) thus ?thesis using left by (auto simp add: ix_def lv_def) next case right hence "?i_c = 2 * ?i_b + 1" using child_ix_right and \d < length b\ by auto hence "?i_b * 2^(?l_p - ?l_b) = (?i_c - 1) * 2^(?l_p - ?l_c)" and "(?i_b + 1) * 2^(?l_p - ?l_b) = (?i_c + 1) * 2^(?l_p - ?l_c)" using level by auto hence "?i_p > ?i_b * 2^(?l_p - ?l_b)" and "?i_p < (?i_b + 1) * 2^(?l_p - ?l_b)" using range_left and range_right by auto with \?l_b < ?l_p\ have "l2_\ (?l_p, ?i_p) (?l_b, ?i_b) = (1 - real_of_int ?i_p / 2^(?l_p - ?l_b) + real_of_int ?i_b) / 2^(?l_p + 1)" by (rule l2_when_right_child) thus ?thesis using right by (auto simp add: ix_def lv_def) qed qed lemma l2_same: "l2_\ (p!d) (p!d) = 1/3 / 2^(lv p d)" proof - have "l2_\ (p!d) (p!d) = l2_\ (lv p d, ix p d) (lv p d, ix p d)" by (auto simp add: lv_def ix_def) thus ?thesis using l2_when_same by auto qed lemma l2_disjoint: assumes "d < length b" and "p \ grid b {d}" and "p' \ grid b {d}" and "p' \ grid p {d}" and "lv p' d \ lv p d" shows "l2_\ (p' ! d) (p ! d) = 0" proof - have range: "ix p' d > (ix p d + 1) * 2^(lv p' d - lv p d) \ ix p' d < (ix p d - 1) * 2^(lv p' d - lv p d)" proof (rule ccontr) assume "\ ?thesis" hence "ix p' d \ (ix p d + 1) * 2^(lv p' d - lv p d)" and "ix p' d \ (ix p d - 1) * 2^(lv p' d - lv p d)" by auto with \p' \ grid b {d}\ and \p \ grid b {d}\ and \lv p' d \ lv p d\ and \d < length b\ have "p' \ grid p {d}" using grid_part[where p=p and b=b and d=d and p'=p'] by auto with \p' \ grid p {d}\ show False by auto qed have "l2_\ (p' ! d) (p ! d) = l2_\ (lv p' d, ix p' d) (lv p d, ix p d)" by (auto simp add: ix_def lv_def) also have "\ = 0" using range and \lv p' d \ lv p d\ and l2_when_disjoint by auto finally show ?thesis . qed lemma l2_down2: fixes pc pd p assumes "d < length pd" assumes pc_in_grid: "pc \ grid (child pd dir d) {d}" assumes pd_is_child: "pd = child p dir d" (is "pd = ?pd") shows "l2_\ (pc ! d) (pd ! d) / 2 = l2_\ (pc ! d) (p ! d)" proof - have "d < length p" using pd_is_child \d < length pd\ by auto moreover have "pc \ grid ?pd {d}" using pd_is_child and grid_child and pc_in_grid by auto hence "lv p d < lv pc d" using grid_child_level and \d < length pd\ and pd_is_child by auto moreover have "real_of_int (sgn dir) * real_of_int (sgn dir) = 1" by (cases dir, auto) ultimately show ?thesis unfolding l2_child[OF \d < length pd\ pc_in_grid] l2_child[OF \d < length p\ \pc \ grid ?pd {d}\] using child_lv and child_ix and pd_is_child and level_shift by (auto simp add: algebra_simps diff_divide_distrib add_divide_distrib) qed lemma l2_zigzag: assumes "d < length p" and p_child: "p = child p_p dir d" and p'_grid: "p' \ grid (child p (inv dir) d) {d}" and ps_intro: "child p (inv dir) d = child ps dir d" (is "?c_p = ?c_ps") shows "l2_\ (p' ! d) (p_p ! d) = l2_\ (p' ! d) (ps ! d) + l2_\ (p' ! d) (p ! d) / 2" proof - have "length p = length ?c_p" by auto also have "\ = length ?c_ps" using ps_intro by auto finally have "length p = length ps" using ps_intro by auto hence "d < length p_p" using p_child and \d < length p\ by auto moreover from ps_intro have "ps = p[d := (lv p d, ix p d - sgn dir)]" by (rule child_neighbour) hence "lv ps d = lv p d" and "real_of_int (ix ps d) = real_of_int (ix p d) - real_of_int (sgn dir)" using lv_def and ix_def and \length p = length ps\ and \d < length p\ by auto moreover have "d < length ps" and *: "p' \ grid (child ps dir d) {d}" using p'_grid ps_intro \length p = length ps\ \d < length p\ by auto have "p' \ grid p {d}" using p'_grid and grid_child by auto hence p_p_grid: "p' \ grid (child p_p dir d) {d}" using p_child by auto hence "lv p' d > lv p_p d" using grid_child_level and \d < length p_p\ by auto moreover have "real_of_int (sgn dir) * real_of_int (sgn dir) = 1" by (cases dir, auto) ultimately show ?thesis unfolding l2_child[OF \d < length p\ p'_grid] l2_child[OF \d < length ps\ *] l2_child[OF \d < length p_p\ p_p_grid] using child_lv and child_ix and p_child level_shift by (auto simp add: add_divide_distrib algebra_simps diff_divide_distrib) qed end diff --git a/thys/Zeta_Function/ROOT b/thys/Zeta_Function/ROOT --- a/thys/Zeta_Function/ROOT +++ b/thys/Zeta_Function/ROOT @@ -1,23 +1,24 @@ chapter AFP session Zeta_Function (AFP) = Dirichlet_Series + options [timeout = 600] sessions + "Pure-ex" Euler_MacLaurin Bernoulli "HOL-Library" "HOL-Real_Asymp" Dirichlet_Series Winding_Number_Eval theories [document=false] "Euler_MacLaurin.Euler_MacLaurin" "Bernoulli.Bernoulli_Zeta" "HOL-Library.Landau_Symbols" "Dirichlet_Series.Dirichlet_Series_Analysis" theories Zeta_Function Zeta_Laurent_Expansion Hadjicostas_Chapman document_files "root.tex" "root.bib" diff --git a/thys/Zeta_Function/Zeta_Function.thy b/thys/Zeta_Function/Zeta_Function.thy --- a/thys/Zeta_Function/Zeta_Function.thy +++ b/thys/Zeta_Function/Zeta_Function.thy @@ -1,3703 +1,3704 @@ (* File: Zeta_Function.thy Author: Manuel Eberl, TU München *) section \The Hurwitz and Riemann $\zeta$ functions\ theory Zeta_Function imports "Euler_MacLaurin.Euler_MacLaurin" "Bernoulli.Bernoulli_Zeta" "Dirichlet_Series.Dirichlet_Series_Analysis" "Winding_Number_Eval.Winding_Number_Eval" "HOL-Real_Asymp.Real_Asymp" Zeta_Library + "Pure-ex.Guess" begin subsection \Preliminary facts\ (* TODO Move once Landau symbols are in the distribution *) lemma holomorphic_on_extend: assumes "f holomorphic_on S - {\}" "\ \ interior S" "f \ O[at \](\_. 1)" shows "(\g. g holomorphic_on S \ (\z\S - {\}. g z = f z))" by (subst holomorphic_on_extend_bounded) (insert assms, auto elim!: landau_o.bigE) lemma removable_singularities: assumes "finite X" "X \ interior S" "f holomorphic_on (S - X)" assumes "\z. z \ X \ f \ O[at z](\_. 1)" shows "\g. g holomorphic_on S \ (\z\S-X. g z = f z)" using assms proof (induction arbitrary: f rule: finite_induct) case empty thus ?case by auto next case (insert z0 X f) from insert.prems and insert.hyps have z0: "z0 \ interior (S - X)" by (auto simp: interior_diff finite_imp_closed) hence "\g. g holomorphic_on (S - X) \ (\z\S - X - {z0}. g z = f z)" using insert.prems insert.hyps by (intro holomorphic_on_extend) auto then obtain g where g: "g holomorphic_on (S - X)" "\z\S - X - {z0}. g z = f z" by blast have "\h. h holomorphic_on S \ (\z\S - X. h z = g z)" proof (rule insert.IH) fix z0' assume z0': "z0' \ X" hence "eventually (\z. z \ interior S - (X - {z0'}) - {z0}) (nhds z0')" using insert.prems insert.hyps by (intro eventually_nhds_in_open open_Diff finite_imp_closed) auto hence ev: "eventually (\z. z \ S - X - {z0}) (at z0')" unfolding eventually_at_filter by eventually_elim (insert z0' insert.hyps interior_subset[of S], auto) have "g \ \[at z0'](f)" by (intro bigthetaI_cong eventually_mono[OF ev]) (insert g, auto) also have "f \ O[at z0'](\_. 1)" using z0' by (intro insert.prems) auto finally show "g \ \" . qed (insert insert.prems g, auto) then obtain h where "h holomorphic_on S" "\z\S - X. h z = g z" by blast with g have "h holomorphic_on S" "\z\S - insert z0 X. h z = f z" by auto thus ?case by blast qed lemma continuous_imp_bigo_1: assumes "continuous (at x within A) f" shows "f \ O[at x within A](\_. 1)" proof (rule bigoI_tendsto) from assms show "((\x. f x / 1) \ f x) (at x within A)" by (auto simp: continuous_within) qed auto lemma taylor_bigo_linear: assumes "f field_differentiable at x0 within A" shows "(\x. f x - f x0) \ O[at x0 within A](\x. x - x0)" proof - from assms obtain f' where "(f has_field_derivative f') (at x0 within A)" by (auto simp: field_differentiable_def) hence "((\x. (f x - f x0) / (x - x0)) \ f') (at x0 within A)" by (auto simp: has_field_derivative_iff) thus ?thesis by (intro bigoI_tendsto[where c = f']) (auto simp: eventually_at_filter) qed (* END TODO *) (* TODO Move? *) lemma powr_add_minus_powr_asymptotics: fixes a z :: complex shows "((\z. ((1 + z) powr a - 1) / z) \ a) (at 0)" proof (rule Lim_transform_eventually) have "eventually (\z::complex. z \ ball 0 1 - {0}) (at 0)" using eventually_at_ball'[of 1 "0::complex" UNIV] by (simp add: dist_norm) thus "eventually (\z. (\n. (a gchoose (Suc n)) * z ^ n) = ((1 + z) powr a - 1) / z) (at 0)" proof eventually_elim case (elim z) hence "(\n. (a gchoose n) * z ^ n) sums (1 + z) powr a" by (intro gen_binomial_complex) auto hence "(\n. (a gchoose (Suc n)) * z ^ (Suc n)) sums ((1 + z) powr a - 1)" by (subst sums_Suc_iff) simp_all also have "(\n. (a gchoose (Suc n)) * z ^ (Suc n)) = (\n. z * ((a gchoose (Suc n)) * z ^ n))" by (simp add: algebra_simps) finally have "(\n. (a gchoose (Suc n)) * z ^ n) sums (((1 + z) powr a - 1) / z)" by (rule sums_mult_D) (use elim in auto) thus ?case by (simp add: sums_iff) qed next have "conv_radius (\n. a gchoose (n + 1)) = conv_radius (\n. a gchoose n)" using conv_radius_shift[of "\n. a gchoose n" 1] by simp hence "continuous_on (cball 0 (1/2)) (\z. \n. (a gchoose (Suc n)) * (z - 0) ^ n)" using conv_radius_gchoose[of a] by (intro powser_continuous_suminf) (simp_all) hence "isCont (\z. \n. (a gchoose (Suc n)) * z ^ n) 0" by (auto intro: continuous_on_interior) thus "(\z. \n. (a gchoose Suc n) * z ^ n) \0\ a" by (auto simp: isCont_def) qed lemma complex_powr_add_minus_powr_asymptotics: fixes s :: complex assumes a: "a > 0" and s: "Re s < 1" shows "filterlim (\x. of_real (x + a) powr s - of_real x powr s) (nhds 0) at_top" proof (rule Lim_transform_eventually) show "eventually (\x. ((1 + of_real (a / x)) powr s - 1) / of_real (a / x) * of_real x powr (s - 1) * a = of_real (x + a) powr s - of_real x powr s) at_top" (is "eventually (\x. ?f x / ?g x * ?h x * _ = _) _") using eventually_gt_at_top[of a] proof eventually_elim case (elim x) have "?f x / ?g x * ?h x * a = ?f x * (a * ?h x / ?g x)" by simp also have "a * ?h x / ?g x = of_real x powr s" using elim a by (simp add: powr_diff) also have "?f x * \ = of_real (x + a) powr s - of_real x powr s" using a elim by (simp add: algebra_simps powr_times_real [symmetric]) finally show ?case . qed have "filterlim (\x. complex_of_real (a / x)) (nhds (complex_of_real 0)) at_top" by (intro tendsto_of_real real_tendsto_divide_at_top[OF tendsto_const] filterlim_ident) hence "filterlim (\x. complex_of_real (a / x)) (at 0) at_top" using a by (intro filterlim_atI) auto hence "((\x. ?f x / ?g x * ?h x * a) \ s * 0 * a) at_top" using s by (intro tendsto_mult filterlim_compose[OF powr_add_minus_powr_asymptotics] tendsto_const tendsto_neg_powr_complex_of_real filterlim_ident) auto thus "((\x. ?f x / ?g x * ?h x * a) \ 0) at_top" by simp qed (* END TODO *) lemma summable_zeta: assumes "Re s > 1" shows "summable (\n. of_nat (Suc n) powr -s)" proof - have "summable (\n. exp (complex_of_real (ln (real (Suc n))) * - s))" (is "summable ?f") by (subst summable_Suc_iff, rule summable_complex_powr_iff) (use assms in auto) also have "?f = (\n. of_nat (Suc n) powr -s)" by (simp add: powr_def algebra_simps del: of_nat_Suc) finally show ?thesis . qed lemma summable_zeta_real: assumes "x > 1" shows "summable (\n. real (Suc n) powr -x)" proof - have "summable (\n. of_nat (Suc n) powr -complex_of_real x)" using assms by (intro summable_zeta) simp_all also have "(\n. of_nat (Suc n) powr -complex_of_real x) = (\n. of_real (real (Suc n) powr -x))" by (subst powr_Reals_eq) simp_all finally show ?thesis by (subst (asm) summable_complex_of_real) qed lemma summable_hurwitz_zeta: assumes "Re s > 1" "a > 0" shows "summable (\n. (of_nat n + of_real a) powr -s)" proof - have "summable (\n. (of_nat (Suc n) + of_real a) powr -s)" proof (rule summable_comparison_test' [OF summable_zeta_real [OF assms(1)]] ) fix n :: nat have "norm ((of_nat (Suc n) + of_real a) powr -s) = (real (Suc n) + a) powr - Re s" (is "?N = _") using assms by (simp add: norm_powr_real_powr) also have "\ \ real (Suc n) powr -Re s" using assms by (intro powr_mono2') auto finally show "?N \ \" . qed thus ?thesis by (subst (asm) summable_Suc_iff) qed lemma summable_hurwitz_zeta_real: assumes "x > 1" "a > 0" shows "summable (\n. (real n + a) powr -x)" proof - have "summable (\n. (of_nat n + of_real a) powr -complex_of_real x)" using assms by (intro summable_hurwitz_zeta) simp_all also have "(\n. (of_nat n + of_real a) powr -complex_of_real x) = (\n. of_real ((real n + a) powr -x))" using assms by (subst powr_Reals_eq) simp_all finally show ?thesis by (subst (asm) summable_complex_of_real) qed subsection \Definitions\ text \ We use the Euler--MacLaurin summation formula to express $\zeta(s,a) - \frac{a^{1-s}}{s-1}$ as a polynomial plus some remainder term, which is an integral over a function of order $O(-1-2n-\mathfrak{R}(s))$. It is then clear that this integral converges uniformly to an analytic function in $s$ for all $s$ with $\mathfrak{R}(s) > -2n$. \ definition pre_zeta_aux :: "nat \ real \ complex \ complex" where "pre_zeta_aux N a s = a powr - s / 2 + (\i=1..N. (bernoulli (2 * i) / fact (2 * i)) *\<^sub>R (pochhammer s (2*i - 1) * of_real a powr (- s - of_nat (2*i - 1)))) + EM_remainder (Suc (2*N)) (\x. -(pochhammer s (Suc (2*N)) * of_real (x + a) powr (- 1 - 2*N - s))) 0" text \ By iterating the above construction long enough, we can extend this to the entire complex plane. \ definition pre_zeta :: "real \ complex \ complex" where "pre_zeta a s = pre_zeta_aux (nat (1 - \Re s / 2\)) a s" text \ We can then obtain the Hurwitz $\zeta$ function by adding back the pole at 1. Note that it is not necessary to trust that this somewhat complicated definition is, in fact, the correct one, since we will later show that this Hurwitz zeta function fulfils \[\zeta(s,a) = \sum_{n=0}^\infty \frac{1}{(n + a)^s}\] and is analytic on $\mathbb{C}\setminus \{1\}$, which uniquely defines the function due to analytic continuation. It is therefore obvious that any alternative definition that is analytic on $\mathbb{C}\setminus \{1\}$ and satisfies the above equation must be equal to our Hurwitz $\zeta$ function. \ definition hurwitz_zeta :: "real \ complex \ complex" where "hurwitz_zeta a s = (if s = 1 then 0 else pre_zeta a s + of_real a powr (1 - s) / (s - 1))" text \ The Riemann $\zeta$ function is simply the Hurwitz $\zeta$ function with $a = 1$. \ definition zeta :: "complex \ complex" where "zeta = hurwitz_zeta 1" text \ We define the $\zeta$ functions as 0 at their poles. To avoid confusion, these facts are not added as simplification rules by default. \ lemma hurwitz_zeta_1: "hurwitz_zeta c 1 = 0" by (simp add: hurwitz_zeta_def) lemma zeta_1: "zeta 1 = 0" by (simp add: zeta_def hurwitz_zeta_1) lemma zeta_minus_pole_eq: "s \ 1 \ zeta s - 1 / (s - 1) = pre_zeta 1 s" by (simp add: zeta_def hurwitz_zeta_def) context begin private lemma holomorphic_pre_zeta_aux': assumes "a > 0" "bounded U" "open U" "U \ {s. Re s > \}" and \: "\ > - 2 * real n" shows "pre_zeta_aux n a holomorphic_on U" unfolding pre_zeta_aux_def proof (intro holomorphic_intros) define C :: real where "C = max 0 (Sup ((\s. norm (pochhammer s (Suc (2 * n)))) ` closure U))" have "compact (closure U)" using assms by (auto simp: compact_eq_bounded_closed) hence "compact ((\s. norm (pochhammer s (Suc (2 * n)))) ` closure U)" by (rule compact_continuous_image [rotated]) (auto intro!: continuous_intros) hence "bounded ((\s. norm (pochhammer s (Suc (2 * n)))) ` closure U)" by (simp add: compact_eq_bounded_closed) hence C: "cmod (pochhammer s (Suc (2 * n))) \ C" if "s \ U" for s using that closure_subset[of U] unfolding C_def by (intro max.coboundedI2 cSup_upper bounded_imp_bdd_above) (auto simp: image_iff) have C' [simp]: "C \ 0" by (simp add: C_def) let ?g = "\(x::real). C * (x + a) powr (- 1 - 2 * of_nat n - \)" let ?G = "\(x::real). C / (- 2 * of_nat n - \) * (x + a) powr (- 2 * of_nat n - \)" define poch' where "poch' = deriv (\z::complex. pochhammer z (Suc (2 * n)))" have [derivative_intros]: "((\z. pochhammer z (Suc (2 * n))) has_field_derivative poch' z) (at z within A)" for z :: complex and A unfolding poch'_def by (rule holomorphic_derivI [OF holomorphic_pochhammer [of _ UNIV]]) auto have A: "continuous_on A poch'" for A unfolding poch'_def by (rule continuous_on_subset[OF _ subset_UNIV], intro holomorphic_on_imp_continuous_on holomorphic_deriv) (auto intro: holomorphic_pochhammer) note [continuous_intros] = continuous_on_compose2[OF this _ subset_UNIV] define f' where "f' = (\z t. - (poch' z * complex_of_real (t + a) powr (- 1 - 2 * of_nat n - z) - Ln (complex_of_real (t + a)) * complex_of_real (t + a) powr (- 1 - 2 * of_nat n - z) * pochhammer z (Suc (2 * n))))" show "(\z. EM_remainder (Suc (2 * n)) (\x. - (pochhammer z (Suc (2 * n)) * complex_of_real (x + a) powr (- 1 - 2 * of_nat n - z))) 0) holomorphic_on U" unfolding pre_zeta_aux_def proof (rule holomorphic_EM_remainder[of _ ?G ?g _ _ f'], goal_cases) case (1 x) show ?case by (insert 1 \ \a > 0\, rule derivative_eq_intros refl | simp)+ (auto simp: field_simps powr_diff powr_add powr_minus) next case (2 z t x) note [derivative_intros] = has_field_derivative_powr_right [THEN DERIV_chain2] show ?case by (insert 2 \ \a > 0\, (rule derivative_eq_intros refl | (simp add: add_eq_0_iff; fail))+) (simp add: f'_def) next case 3 hence *: "complex_of_real x + complex_of_real a \ \\<^sub>\\<^sub>0" if "x \ 0" for x using nonpos_Reals_of_real_iff[of "x+a", unfolded of_real_add] that \a > 0\ by auto show ?case using \a > 0\ and * unfolding f'_def by (auto simp: case_prod_unfold add_eq_0_iff intro!: continuous_intros) next case (4 b c z e) have "- 2 * real n < \" by (fact \) also from 4 assms have "\ < Re z" by auto finally show ?case using assms 4 by (intro integrable_continuous_real continuous_intros) (auto simp: add_eq_0_iff) next case (5 t x s) thus ?case using \a > 0\ by (intro integrable_EM_remainder') (auto intro!: continuous_intros simp: add_eq_0_iff) next case 6 from \ have "(\y. C / (-2 * real n - \) * (a + y) powr (-2 * real n - \)) \ 0" by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_real_sequentially filterlim_tendsto_add_at_top [OF tendsto_const]) auto thus ?case unfolding convergent_def by (auto simp: add_ac) next case 7 show ?case proof (intro eventually_mono [OF eventually_ge_at_top[of 1]] ballI) fix x :: real and s :: complex assume x: "x \ 1" and s: "s \ U" have "norm (- (pochhammer s (Suc (2 * n)) * of_real (x + a) powr (- 1 - 2 * of_nat n - s))) = norm (pochhammer s (Suc (2 * n))) * (x + a) powr (-1 - 2 * of_nat n - Re s)" (is "?N = _") using 7 \a > 0\ x by (simp add: norm_mult norm_powr_real_powr) also have "\ \ ?g x" using 7 assms x s \a > 0\ by (intro mult_mono C powr_mono) auto finally show "?N \ ?g x" . qed qed (insert assms, auto) qed (insert assms, auto) lemma analytic_pre_zeta_aux: assumes "a > 0" shows "pre_zeta_aux n a analytic_on {s. Re s > - 2 * real n}" unfolding analytic_on_def proof fix s assume s: "s \ {s. Re s > - 2 * real n}" define \ where "\ = (Re s - 2 * real n) / 2" with s have \: "\ > - 2 * real n" by (simp add: \_def field_simps) from s have s': "s \ {s. Re s > \}" by (auto simp: \_def field_simps) have "open {s. Re s > \}" by (rule open_halfspace_Re_gt) with s' obtain \ where "\ > 0" "ball s \ \ {s. Re s > \}" unfolding open_contains_ball by blast with \ have "pre_zeta_aux n a holomorphic_on ball s \" by (intro holomorphic_pre_zeta_aux' [OF assms, of _ \]) auto with \\ > 0\ show "\e>0. pre_zeta_aux n a holomorphic_on ball s e" by blast qed end context fixes s :: complex and N :: nat and \ :: "complex \ complex" and a :: real assumes s: "Re s > 1" and a: "a > 0" defines "\ \ (\s. \n. (of_nat n + of_real a) powr -s)" begin interpretation \: euler_maclaurin_nat' "\x. of_real (x + a) powr (1 - s) / (1 - s)" "\x. of_real (x + a) powr -s" "\n x. (-1) ^ n * pochhammer s n * of_real (x + a) powr -(s + n)" 0 N "\ s" "{}" proof (standard, goal_cases) case 2 show ?case by (simp add: powr_minus field_simps) next case (3 k) have "complex_of_real x + complex_of_real a = 0 \ x = -a" for x by (simp only: of_real_add [symmetric] of_real_eq_0_iff add_eq_0_iff2) with a s show ?case by (intro continuous_intros) (auto simp: add_nonneg_nonneg) next case (4 k x) with a have "0 < x + a" by simp hence *: "complex_of_real x + complex_of_real a \ \\<^sub>\\<^sub>0" using nonpos_Reals_of_real_iff[of "x+a", unfolded of_real_add] by auto have **: "pochhammer z (Suc n) = - pochhammer z n * (-z - of_nat n :: complex)" for z n by (simp add: pochhammer_rec' field_simps) show "((\x. (- 1) ^ k * pochhammer s k * of_real (x + a) powr - (s + of_nat k)) has_vector_derivative (- 1) ^ Suc k * pochhammer s (Suc k) * of_real (x + a) powr - (s + of_nat (Suc k))) (at x)" by (insert 4 *, (rule has_vector_derivative_real_field derivative_eq_intros refl | simp)+) (auto simp: divide_simps powr_add powr_diff powr_minus **) next case 5 with s a show ?case by (auto intro!: continuous_intros simp: minus_equation_iff add_eq_0_iff) next case (6 x) with a have "0 < x + a" by simp hence *: "complex_of_real x + complex_of_real a \ \\<^sub>\\<^sub>0" using nonpos_Reals_of_real_iff[of "x+a", unfolded of_real_add] by auto show ?case unfolding of_real_add by (insert 6 s *, (rule has_vector_derivative_real_field derivative_eq_intros refl | force simp add: minus_equation_iff)+) next case 7 from s a have "(\k. (of_nat k + of_real a) powr -s) sums \ s" unfolding \_def by (intro summable_sums summable_hurwitz_zeta) auto hence 1: "(\b. (\k=0..b. (of_nat k + of_real a) powr -s)) \ \ s" by (simp add: sums_def') { fix z assume "Re z < 0" hence "((\b. (a + real b) powr Re z) \ 0) at_top" by (intro tendsto_neg_powr filterlim_tendsto_add_at_top filterlim_real_sequentially) auto also have "(\b. (a + real b) powr Re z) = (\b. norm ((of_nat b + a) powr z))" using a by (subst norm_powr_real_powr) (auto simp: add_ac) finally have "((\b. (of_nat b + a) powr z) \ 0) at_top" by (subst (asm) tendsto_norm_zero_iff) simp } note * = this have "(\b. (of_nat b + a) powr (1 - s) / (1 - s)) \ 0 / (1 - s)" using s by (intro tendsto_divide tendsto_const *) auto hence 2: "(\b. (of_nat b + a) powr (1 - s) / (1 - s)) \ 0" by simp have "(\b. (\i<2 * N + 1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R ((- 1) ^ i * pochhammer s i * (of_nat b + a) powr -(s + of_nat i)))) \ (\i<2 * N + 1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R ((- 1) ^ i * pochhammer s i * 0))" using s by (intro tendsto_intros *) auto hence 3: "(\b. (\i<2 * N + 1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R ((- 1) ^ i * pochhammer s i * (of_nat b + a) powr -(s + of_nat i)))) \ 0" by simp from tendsto_diff[OF tendsto_diff[OF 1 2] 3] show ?case by simp qed simp_all text \ The pre-$\zeta$ functions agree with the infinite sum that is used to define the $\zeta$ function for $\mathfrak{R}(s)>1$. \ lemma pre_zeta_aux_conv_zeta: "pre_zeta_aux N a s = \ s + a powr (1 - s) / (1 - s)" proof - let ?R = "(\i=1..N. ((bernoulli (2*i) / fact (2*i)) *\<^sub>R pochhammer s (2*i - 1) * of_real a powr (-s - (2*i-1))))" let ?S = "EM_remainder (Suc (2 * N)) (\x. - (pochhammer s (Suc (2*N)) * of_real (x + a) powr (- 1 - 2 * of_nat N - s))) 0" from \.euler_maclaurin_strong_nat'[OF le_refl, simplified] have "of_real a powr -s = a powr (1 - s) / (1 - s) + \ s + a powr -s / 2 + (-?R) - ?S" unfolding sum_negf [symmetric] by (simp add: scaleR_conv_of_real pre_zeta_aux_def mult_ac) thus ?thesis unfolding pre_zeta_aux_def (* TODO: Field_as_Ring causes some problems with field_simps vs. div_mult_self *) by (simp add: field_simps del: div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1) qed end text \ Since all of the partial pre-$\zeta$ functions are analytic and agree in the halfspace with $\mathfrak R(s)>0$, they must agree in their entire domain. \ lemma pre_zeta_aux_eq: assumes "m \ n" "a > 0" "Re s > -2 * real m" shows "pre_zeta_aux m a s = pre_zeta_aux n a s" proof - have "pre_zeta_aux n a s - pre_zeta_aux m a s = 0" proof (rule analytic_continuation[of "\s. pre_zeta_aux n a s - pre_zeta_aux m a s"]) show "(\s. pre_zeta_aux n a s - pre_zeta_aux m a s) holomorphic_on {s. Re s > -2 * real m}" using assms by (intro holomorphic_intros analytic_imp_holomorphic analytic_on_subset[OF analytic_pre_zeta_aux]) auto next fix s assume "s \ {s. Re s > 1}" with \a > 0\ show "pre_zeta_aux n a s - pre_zeta_aux m a s = 0" by (simp add: pre_zeta_aux_conv_zeta) next have "2 \ {s. Re s > 1}" by simp also have "\ = interior \" by (intro interior_open [symmetric] open_halfspace_Re_gt) finally show "2 islimpt {s. Re s > 1}" by (rule interior_limit_point) next show "connected {s. Re s > -2 * real m}" using convex_halfspace_gt[of "-2 * real m" "1::complex"] by (intro convex_connected) auto qed (insert assms, auto simp: open_halfspace_Re_gt) thus ?thesis by simp qed lemma pre_zeta_aux_eq': assumes "a > 0" "Re s > -2 * real m" "Re s > -2 * real n" shows "pre_zeta_aux m a s = pre_zeta_aux n a s" proof (cases m n rule: linorder_cases) case less with assms show ?thesis by (intro pre_zeta_aux_eq) auto next case greater with assms show ?thesis by (subst eq_commute, intro pre_zeta_aux_eq) auto qed auto lemma pre_zeta_aux_eq_pre_zeta: assumes "Re s > -2 * real n" and "a > 0" shows "pre_zeta_aux n a s = pre_zeta a s" unfolding pre_zeta_def proof (intro pre_zeta_aux_eq') from assms show "- 2 * real (nat (1 - \Re s / 2\)) < Re s" by linarith qed (insert assms, simp_all) text \ This means that the idea of iterating that construction infinitely does yield a well-defined entire function. \ lemma analytic_pre_zeta: assumes "a > 0" shows "pre_zeta a analytic_on A" unfolding analytic_on_def proof fix s assume "s \ A" let ?B = "{s'. Re s' > of_int \Re s\ - 1}" have s: "s \ ?B" by simp linarith? moreover have "open ?B" by (rule open_halfspace_Re_gt) ultimately obtain \ where \: "\ > 0" "ball s \ \ ?B" unfolding open_contains_ball by blast define C where "C = ball s \" note analytic = analytic_on_subset[OF analytic_pre_zeta_aux] have "pre_zeta_aux (nat \- Re s\ + 2) a holomorphic_on C" proof (intro analytic_imp_holomorphic analytic subsetI assms, goal_cases) case (1 w) with \ have "w \ ?B" by (auto simp: C_def) thus ?case by (auto simp: ceiling_minus) qed also have "?this \ pre_zeta a holomorphic_on C" proof (intro holomorphic_cong refl pre_zeta_aux_eq_pre_zeta assms) fix w assume "w \ C" with \ have w: "w \ ?B" by (auto simp: C_def) thus " - 2 * real (nat \- Re s\ + 2) < Re w" by (simp add: ceiling_minus) qed finally show "\e>0. pre_zeta a holomorphic_on ball s e" using \\ > 0\ unfolding C_def by blast qed lemma holomorphic_pre_zeta [holomorphic_intros]: "f holomorphic_on A \ a > 0 \ (\z. pre_zeta a (f z)) holomorphic_on A" using holomorphic_on_compose [OF _ analytic_imp_holomorphic [OF analytic_pre_zeta], of f] by (simp add: o_def) corollary continuous_on_pre_zeta: "a > 0 \ continuous_on A (pre_zeta a)" by (intro holomorphic_on_imp_continuous_on holomorphic_intros) auto corollary continuous_on_pre_zeta' [continuous_intros]: "continuous_on A f \ a > 0 \ continuous_on A (\x. pre_zeta a (f x))" using continuous_on_compose2 [OF continuous_on_pre_zeta, of a A f "f ` A"] by (auto simp: image_iff) corollary continuous_pre_zeta [continuous_intros]: "a > 0 \ continuous (at s within A) (pre_zeta a)" by (rule continuous_within_subset[of _ UNIV]) (insert continuous_on_pre_zeta[of a UNIV], auto simp: continuous_on_eq_continuous_at open_Compl) corollary continuous_pre_zeta' [continuous_intros]: "a > 0 \ continuous (at s within A) f \ continuous (at s within A) (\s. pre_zeta a (f s))" using continuous_within_compose3[OF continuous_pre_zeta, of a s A f] by auto text \ It is now obvious that $\zeta$ is holomorphic everywhere except 1, where it has a simple pole with residue 1, which we can simply read off. \ theorem holomorphic_hurwitz_zeta: assumes "a > 0" "1 \ A" shows "hurwitz_zeta a holomorphic_on A" proof - have "(\s. pre_zeta a s + complex_of_real a powr (1 - s) / (s - 1)) holomorphic_on A" using assms by (auto intro!: holomorphic_intros) also from assms have "?this \ ?thesis" by (intro holomorphic_cong) (auto simp: hurwitz_zeta_def) finally show ?thesis . qed corollary holomorphic_hurwitz_zeta' [holomorphic_intros]: assumes "f holomorphic_on A" and "a > 0" and "\z. z \ A \ f z \ 1" shows "(\x. hurwitz_zeta a (f x)) holomorphic_on A" proof - have "hurwitz_zeta a \ f holomorphic_on A" using assms by (intro holomorphic_on_compose_gen[of _ _ _ "f ` A"] holomorphic_hurwitz_zeta assms) auto thus ?thesis by (simp add: o_def) qed theorem holomorphic_zeta: "1 \ A\ zeta holomorphic_on A" unfolding zeta_def by (auto intro: holomorphic_intros) corollary holomorphic_zeta' [holomorphic_intros]: assumes "f holomorphic_on A" and "\z. z \ A \ f z \ 1" shows "(\x. zeta (f x)) holomorphic_on A" using assms unfolding zeta_def by (auto intro: holomorphic_intros) corollary analytic_hurwitz_zeta: assumes "a > 0" "1 \ A" shows "hurwitz_zeta a analytic_on A" proof - from assms(1) have "hurwitz_zeta a holomorphic_on -{1}" by (rule holomorphic_hurwitz_zeta) auto also have "?this \ hurwitz_zeta a analytic_on -{1}" by (intro analytic_on_open [symmetric]) auto finally show ?thesis by (rule analytic_on_subset) (insert assms, auto) qed corollary analytic_zeta: "1 \ A \ zeta analytic_on A" unfolding zeta_def by (rule analytic_hurwitz_zeta) auto corollary continuous_on_hurwitz_zeta: "a > 0 \ 1 \ A \ continuous_on A (hurwitz_zeta a)" by (intro holomorphic_on_imp_continuous_on holomorphic_intros) auto corollary continuous_on_hurwitz_zeta' [continuous_intros]: "continuous_on A f \ a > 0 \ (\x. x \ A \ f x \ 1) \ continuous_on A (\x. hurwitz_zeta a (f x))" using continuous_on_compose2 [OF continuous_on_hurwitz_zeta, of a "f ` A" A f] by (auto simp: image_iff) corollary continuous_on_zeta: "1 \ A \ continuous_on A zeta" by (intro holomorphic_on_imp_continuous_on holomorphic_intros) auto corollary continuous_on_zeta' [continuous_intros]: "continuous_on A f \ (\x. x \ A \ f x \ 1) \ continuous_on A (\x. zeta (f x))" using continuous_on_compose2 [OF continuous_on_zeta, of "f ` A" A f] by (auto simp: image_iff) corollary continuous_hurwitz_zeta [continuous_intros]: "a > 0 \ s \ 1 \ continuous (at s within A) (hurwitz_zeta a)" by (rule continuous_within_subset[of _ UNIV]) (insert continuous_on_hurwitz_zeta[of a "-{1}"], auto simp: continuous_on_eq_continuous_at open_Compl) corollary continuous_hurwitz_zeta' [continuous_intros]: "a > 0 \ f s \ 1 \ continuous (at s within A) f \ continuous (at s within A) (\s. hurwitz_zeta a (f s))" using continuous_within_compose3[OF continuous_hurwitz_zeta, of a f s A] by auto corollary continuous_zeta [continuous_intros]: "s \ 1 \ continuous (at s within A) zeta" unfolding zeta_def by (intro continuous_intros) auto corollary continuous_zeta' [continuous_intros]: "f s \ 1 \ continuous (at s within A) f \ continuous (at s within A) (\s. zeta (f s))" unfolding zeta_def by (intro continuous_intros) auto corollary field_differentiable_at_zeta: assumes "s \ 1" shows "zeta field_differentiable at s" proof - have "zeta holomorphic_on (- {1})" using holomorphic_zeta by force moreover have "open (-{1} :: complex set)" by (intro open_Compl) auto ultimately show ?thesis using assms by (auto simp add: holomorphic_on_open open_halfspace_Re_gt open_Diff field_differentiable_def) qed theorem is_pole_hurwitz_zeta: assumes "a > 0" shows "is_pole (hurwitz_zeta a) 1" proof - from assms have "continuous_on UNIV (pre_zeta a)" by (intro holomorphic_on_imp_continuous_on analytic_imp_holomorphic analytic_pre_zeta) hence "isCont (pre_zeta a) 1" by (auto simp: continuous_on_eq_continuous_at) hence *: "pre_zeta a \1\ pre_zeta a 1" by (simp add: isCont_def) from assms have "isCont (\s. complex_of_real a powr (1 - s)) 1" by (intro isCont_powr_complex) auto with assms have **: "(\s. complex_of_real a powr (1 - s)) \1\ 1" by (simp add: isCont_def) have "(\s::complex. s - 1) \1\ 1 - 1" by (intro tendsto_intros) hence "filterlim (\s::complex. s - 1) (at 0) (at 1)" by (auto simp: filterlim_at eventually_at_filter) hence ***: "filterlim (\s :: complex. a powr (1 - s) / (s - 1)) at_infinity (at 1)" by (intro filterlim_divide_at_infinity [OF **]) auto have "is_pole (\s. pre_zeta a s + complex_of_real a powr (1 - s) / (s - 1)) 1" unfolding is_pole_def hurwitz_zeta_def by (rule tendsto_add_filterlim_at_infinity * ***)+ also have "?this \ ?thesis" unfolding is_pole_def by (intro filterlim_cong refl) (auto simp: eventually_at_filter hurwitz_zeta_def) finally show ?thesis . qed corollary is_pole_zeta: "is_pole zeta 1" unfolding zeta_def by (rule is_pole_hurwitz_zeta) auto theorem zorder_hurwitz_zeta: assumes "a > 0" shows "zorder (hurwitz_zeta a) 1 = -1" proof (rule zorder_eqI[of UNIV]) fix w :: complex assume "w \ UNIV" "w \ 1" thus "hurwitz_zeta a w = (pre_zeta a w * (w - 1) + a powr (1 - w)) * (w - 1) powr (of_int (-1))" apply (subst (1) powr_of_int) by (auto simp add: hurwitz_zeta_def field_simps) qed (insert assms, auto intro!: holomorphic_intros) corollary zorder_zeta: "zorder zeta 1 = - 1" unfolding zeta_def by (rule zorder_hurwitz_zeta) auto theorem residue_hurwitz_zeta: assumes "a > 0" shows "residue (hurwitz_zeta a) 1 = 1" proof - note holo = analytic_imp_holomorphic[OF analytic_pre_zeta] have "residue (hurwitz_zeta a) 1 = residue (\z. pre_zeta a z + a powr (1 - z) / (z - 1)) 1" by (intro residue_cong) (auto simp: eventually_at_filter hurwitz_zeta_def) also have "\ = residue (\z. a powr (1 - z) / (z - 1)) 1" using assms by (subst residue_add [of UNIV]) (auto intro!: holomorphic_intros holo intro: residue_holo[of UNIV, OF _ _ holo]) also have "\ = complex_of_real a powr (1 - 1)" using assms by (intro residue_simple [of UNIV]) (auto intro!: holomorphic_intros) also from assms have "\ = 1" by simp finally show ?thesis . qed corollary residue_zeta: "residue zeta 1 = 1" unfolding zeta_def by (rule residue_hurwitz_zeta) auto lemma zeta_bigo_at_1: "zeta \ O[at 1 within A](\x. 1 / (x - 1))" proof - have "zeta \ \[at 1 within A](\s. pre_zeta 1 s + 1 / (s - 1))" by (intro bigthetaI_cong) (auto simp: eventually_at_filter zeta_def hurwitz_zeta_def) also have "(\s. pre_zeta 1 s + 1 / (s - 1)) \ O[at 1 within A](\s. 1 / (s - 1))" proof (rule sum_in_bigo) have "continuous_on UNIV (pre_zeta 1)" by (intro holomorphic_on_imp_continuous_on holomorphic_intros) auto hence "isCont (pre_zeta 1) 1" by (auto simp: continuous_on_eq_continuous_at) hence "continuous (at 1 within A) (pre_zeta 1)" by (rule continuous_within_subset) auto hence "pre_zeta 1 \ O[at 1 within A](\_. 1)" by (intro continuous_imp_bigo_1) auto also have ev: "eventually (\s. s \ ball 1 1 \ s \ 1 \ s \ A) (at 1 within A)" by (intro eventually_at_ball') auto have "(\_. 1) \ O[at 1 within A](\s. 1 / (s - 1))" by (intro landau_o.bigI[of 1] eventually_mono[OF ev]) (auto simp: eventually_at_filter norm_divide dist_norm norm_minus_commute field_simps) finally show "pre_zeta 1 \ O[at 1 within A](\s. 1 / (s - 1))" . qed simp_all finally show ?thesis . qed theorem assumes "a > 0" "Re s > 1" shows hurwitz_zeta_conv_suminf: "hurwitz_zeta a s = (\n. (of_nat n + of_real a) powr -s)" and sums_hurwitz_zeta: "(\n. (of_nat n + of_real a) powr -s) sums hurwitz_zeta a s" proof - from assms have [simp]: "s \ 1" by auto from assms have "hurwitz_zeta a s = pre_zeta_aux 0 a s + of_real a powr (1 - s) / (s - 1)" by (simp add: hurwitz_zeta_def pre_zeta_def) also from assms have "pre_zeta_aux 0 a s = (\n. (of_nat n + of_real a) powr -s) + of_real a powr (1 - s) / (1 - s)" by (intro pre_zeta_aux_conv_zeta) also have "\ + a powr (1 - s) / (s - 1) = (\n. (of_nat n + of_real a) powr -s) + a powr (1 - s) * (1 / (1 - s) + 1 / (s - 1))" by (simp add: algebra_simps) also have "1 / (1 - s) + 1 / (s - 1) = 0" by (simp add: divide_simps) finally show "hurwitz_zeta a s = (\n. (of_nat n + of_real a) powr -s)" by simp moreover have "(\n. (of_nat n + of_real a) powr -s) sums (\n. (of_nat n + of_real a) powr -s)" by (intro summable_sums summable_hurwitz_zeta assms) ultimately show "(\n. (of_nat n + of_real a) powr -s) sums hurwitz_zeta a s" by simp qed corollary assumes "Re s > 1" shows zeta_conv_suminf: "zeta s = (\n. of_nat (Suc n) powr -s)" and sums_zeta: "(\n. of_nat (Suc n) powr -s) sums zeta s" using hurwitz_zeta_conv_suminf[of 1 s] sums_hurwitz_zeta[of 1 s] assms by (simp_all add: zeta_def add_ac) corollary assumes "n > 1" shows zeta_nat_conv_suminf: "zeta (of_nat n) = (\k. 1 / of_nat (Suc k) ^ n)" and sums_zeta_nat: "(\k. 1 / of_nat (Suc k) ^ n) sums zeta (of_nat n)" proof - have "(\k. of_nat (Suc k) powr -of_nat n) sums zeta (of_nat n)" using assms by (intro sums_zeta) auto also have "(\k. of_nat (Suc k) powr -of_nat n) = (\k. 1 / of_nat (Suc k) ^ n :: complex)" by (simp add: powr_minus divide_simps del: of_nat_Suc) finally show "(\k. 1 / of_nat (Suc k) ^ n) sums zeta (of_nat n)" . thus "zeta (of_nat n) = (\k. 1 / of_nat (Suc k) ^ n)" by (simp add: sums_iff) qed lemma pre_zeta_aux_cnj [simp]: assumes "a > 0" shows "pre_zeta_aux n a (cnj z) = cnj (pre_zeta_aux n a z)" proof - have "cnj (pre_zeta_aux n a z) = of_real a powr -cnj z / 2 + (\x=1..n. (bernoulli (2 * x) / fact (2 * x)) *\<^sub>R a powr (-cnj z - (2*x-1)) * pochhammer (cnj z) (2*x-1)) + EM_remainder (2*n+1) (\x. -(pochhammer (cnj z) (Suc (2 * n)) * cnj (of_real (x + a) powr (-1 - 2 * of_nat n - z)))) 0" (is "_ = _ + ?A + ?B") unfolding pre_zeta_aux_def complex_cnj_add using assms by (subst EM_remainder_cnj [symmetric]) (auto intro!: continuous_intros simp: cnj_powr add_eq_0_iff mult_ac) also have "?B = EM_remainder (2*n+1) (\x. -(pochhammer (cnj z) (Suc (2 * n)) * of_real (x + a) powr (-1 - 2 * of_nat n - cnj z))) 0" using assms by (intro EM_remainder_cong) (auto simp: cnj_powr) also have "of_real a powr -cnj z / 2 + ?A + \ = pre_zeta_aux n a (cnj z)" by (simp add: pre_zeta_aux_def mult_ac) finally show ?thesis .. qed lemma pre_zeta_cnj [simp]: "a > 0 \ pre_zeta a (cnj z) = cnj (pre_zeta a z)" by (simp add: pre_zeta_def) lemma hurwitz_zeta_cnj [simp]: "a > 0 \ hurwitz_zeta a (cnj z) = cnj (hurwitz_zeta a z)" proof - assume "a > 0" moreover have "cnj z = 1 \ z = 1" by (simp add: complex_eq_iff) ultimately show ?thesis by (auto simp: hurwitz_zeta_def cnj_powr) qed lemma zeta_cnj [simp]: "zeta (cnj z) = cnj (zeta z)" by (simp add: zeta_def) corollary hurwitz_zeta_real: "a > 0 \ hurwitz_zeta a (of_real x) \ \" using hurwitz_zeta_cnj [of a "of_real x"] by (simp add: Reals_cnj_iff del: zeta_cnj) corollary zeta_real: "zeta (of_real x) \ \" unfolding zeta_def by (rule hurwitz_zeta_real) auto corollary zeta_real': "z \ \ \ zeta z \ \" by (elim Reals_cases) (auto simp: zeta_real) subsection \Connection to Dirichlet series\ lemma eval_fds_zeta: "Re s > 1 \ eval_fds fds_zeta s = zeta s" using sums_zeta [of s] by (intro eval_fds_eqI) (auto simp: powr_minus divide_simps) theorem euler_product_zeta: assumes "Re s > 1" shows "(\n. \p\n. if prime p then inverse (1 - 1 / of_nat p powr s) else 1) \ zeta s" using euler_product_fds_zeta[of s] assms unfolding nat_power_complex_def by (simp add: eval_fds_zeta) corollary euler_product_zeta': assumes "Re s > 1" shows "(\n. \p | prime p \ p \ n. inverse (1 - 1 / of_nat p powr s)) \ zeta s" proof - note euler_product_zeta [OF assms] also have "(\n. \p\n. if prime p then inverse (1 - 1 / of_nat p powr s) else 1) = (\n. \p | prime p \ p \ n. inverse (1 - 1 / of_nat p powr s))" by (intro ext prod.mono_neutral_cong_right refl) auto finally show ?thesis . qed theorem zeta_Re_gt_1_nonzero: "Re s > 1 \ zeta s \ 0" using eval_fds_zeta_nonzero[of s] by (simp add: eval_fds_zeta) theorem tendsto_zeta_Re_going_to_at_top: "(zeta \ 1) (Re going_to at_top)" proof (rule Lim_transform_eventually) have "eventually (\x::real. x > 1) at_top" by (rule eventually_gt_at_top) hence "eventually (\s. Re s > 1) (Re going_to at_top)" by blast thus "eventually (\z. eval_fds fds_zeta z = zeta z) (Re going_to at_top)" by eventually_elim (simp add: eval_fds_zeta) next have "conv_abscissa (fds_zeta :: complex fds) \ 1" proof (rule conv_abscissa_leI) fix c' assume "ereal c' > 1" thus "\s. s \ 1 = c' \ fds_converges fds_zeta (s::complex)" by (auto intro!: exI[of _ "of_real c'"]) qed hence "(eval_fds fds_zeta \ fds_nth fds_zeta 1) (Re going_to at_top)" by (intro tendsto_eval_fds_Re_going_to_at_top') auto thus "(eval_fds fds_zeta \ 1) (Re going_to at_top)" by simp qed lemma conv_abscissa_zeta [simp]: "conv_abscissa (fds_zeta :: complex fds) = 1" and abs_conv_abscissa_zeta [simp]: "abs_conv_abscissa (fds_zeta :: complex fds) = 1" proof - let ?z = "fds_zeta :: complex fds" have A: "conv_abscissa ?z \ 1" proof (intro conv_abscissa_geI) fix c' assume "ereal c' < 1" hence "\summable (\n. real n powr -c')" by (subst summable_real_powr_iff) auto hence "\summable (\n. of_real (real n powr -c') :: complex)" by (subst summable_of_real_iff) also have "summable (\n. of_real (real n powr -c') :: complex) \ fds_converges fds_zeta (of_real c' :: complex)" unfolding fds_converges_def by (intro summable_cong eventually_mono [OF eventually_gt_at_top[of 0]]) (simp add: fds_nth_zeta powr_Reals_eq powr_minus divide_simps) finally show "\s::complex. s \ 1 = c' \ \fds_converges fds_zeta s" by (intro exI[of _ "of_real c'"]) auto qed have B: "abs_conv_abscissa ?z \ 1" proof (intro abs_conv_abscissa_leI) fix c' assume "1 < ereal c'" thus "\s::complex. s \ 1 = c' \ fds_abs_converges fds_zeta s" by (intro exI[of _ "of_real c'"]) auto qed have "conv_abscissa ?z \ abs_conv_abscissa ?z" by (rule conv_le_abs_conv_abscissa) also note B finally show "conv_abscissa ?z = 1" using A by (intro antisym) note A also have "conv_abscissa ?z \ abs_conv_abscissa ?z" by (rule conv_le_abs_conv_abscissa) finally show "abs_conv_abscissa ?z = 1" using B by (intro antisym) qed theorem deriv_zeta_sums: assumes s: "Re s > 1" shows "(\n. -of_real (ln (real (Suc n))) / of_nat (Suc n) powr s) sums deriv zeta s" proof - from s have "fds_converges (fds_deriv fds_zeta) s" by (intro fds_converges_deriv) simp_all with s have "(\n. -of_real (ln (real (Suc n))) / of_nat (Suc n) powr s) sums deriv (eval_fds fds_zeta) s" unfolding fds_converges_altdef by (simp add: fds_nth_deriv scaleR_conv_of_real eval_fds_deriv eval_fds_zeta) also from s have "eventually (\s. s \ {s. Re s > 1}) (nhds s)" by (intro eventually_nhds_in_open) (auto simp: open_halfspace_Re_gt) hence "eventually (\s. eval_fds fds_zeta s = zeta s) (nhds s)" by eventually_elim (auto simp: eval_fds_zeta) hence "deriv (eval_fds fds_zeta) s = deriv zeta s" by (intro deriv_cong_ev refl) finally show ?thesis . qed theorem inverse_zeta_sums: assumes s: "Re s > 1" shows "(\n. moebius_mu (Suc n) / of_nat (Suc n) powr s) sums inverse (zeta s)" proof - have "fds_converges (fds moebius_mu) s" using assms by (auto intro!: fds_abs_converges_moebius_mu) hence "(\n. moebius_mu (Suc n) / of_nat (Suc n) powr s) sums eval_fds (fds moebius_mu) s" by (simp add: fds_converges_altdef) also have "fds moebius_mu = inverse (fds_zeta :: complex fds)" by (rule fds_moebius_inverse_zeta) also from s have "eval_fds \ s = inverse (zeta s)" by (subst eval_fds_inverse) (auto simp: fds_moebius_inverse_zeta [symmetric] eval_fds_zeta intro!: fds_abs_converges_moebius_mu) finally show ?thesis . qed text \ The following gives an extension of the $\zeta$ functions to the critical strip. \ lemma hurwitz_zeta_critical_strip: fixes s :: complex and a :: real defines "S \ (\n. \i (\n. of_nat n powr (1 - s) / (1 - s))" assumes "Re s > 0" "s \ 1" and "a > 0" shows "(\n. S n - I' n) \ hurwitz_zeta a s" proof - from assms have [simp]: "s \ 1" by auto let ?f = "\x. of_real (x + a) powr -s" let ?fs = "\n x. (-1) ^ n * pochhammer s n * of_real (x + a) powr (-s - of_nat n)" have minus_commute: "-a - b = -b - a" for a b :: complex by (simp add: algebra_simps) define I where "I = (\n. (of_nat n + a) powr (1 - s) / (1 - s))" define R where "R = (\n. EM_remainder' 1 (?fs 1) (real 0) (real n))" define R_lim where "R_lim = EM_remainder 1 (?fs 1) 0" define C where "C = - (a powr -s / 2)" define D where "D = (\n. (1/2) * (of_real (a + real n) powr - s))" define D' where "D' = (\n. of_real (a + real n) powr - s)" define C' where "C' = a powr (1 - s) / (1 - s)" define C'' where "C'' = of_real a powr - s" { fix n :: nat assume n: "n > 0" have "((\x. of_real (x + a) powr -s) has_integral (of_real (real n + a) powr (1-s) / (1 - s) - of_real (0 + a) powr (1 - s) / (1 - s))) {0..real n}" using n assms by (intro fundamental_theorem_of_calculus) (auto intro!: continuous_intros has_vector_derivative_real_field derivative_eq_intros simp: complex_nonpos_Reals_iff) hence I: "((\x. of_real (x + a) powr -s) has_integral (I n - C')) {0..n}" by (auto simp: divide_simps C'_def I_def) have "(\i\{0<..n}. ?f (real i)) - integral {real 0..real n} ?f = (\k<1. (bernoulli' (Suc k) / fact (Suc k)) *\<^sub>R (?fs k (real n) - ?fs k (real 0))) + R n" using n assms unfolding R_def by (intro euler_maclaurin_strong_raw_nat[where Y = "{0}"]) (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_field simp: pochhammer_rec' algebra_simps complex_nonpos_Reals_iff add_eq_0_iff) also have "(\k<1. (bernoulli' (Suc k) / fact (Suc k)) *\<^sub>R (?fs k (real n) - ?fs k (real 0))) = ((n + a) powr - s - a powr - s) / 2" by (simp add: lessThan_nat_numeral scaleR_conv_of_real numeral_2_eq_2 [symmetric]) also have "\ = C + D n" by (simp add: C_def D_def field_simps) also have "integral {real 0..real n} (\x. complex_of_real (x + a) powr - s) = I n - C'" using I by (simp add: has_integral_iff) also have "(\i\{0<..n}. of_real (real i + a) powr - s) = (\i=0..n. of_real (real i + a) powr - s) - of_real a powr -s" using assms by (subst sum.head) auto also have "(\i=0..n. of_real (real i + a) powr - s) = S n + of_real (real n + a) powr -s" unfolding S_def by (subst sum.last_plus) (auto simp: atLeast0LessThan) finally have "C - C' + C'' - D' n + D n + R n + (I n - I' n) = S n - I' n" by (simp add: algebra_simps S_def D'_def C''_def) } hence ev: "eventually (\n. C - C' + C'' - D' n + D n + R n + (I n - I' n) = S n - I' n) at_top" by (intro eventually_mono[OF eventually_gt_at_top[of 0]]) auto have [simp]: "-1 - s = -s - 1" by simp { let ?C = "norm (pochhammer s 1)" have "R \ R_lim" unfolding R_def R_lim_def of_nat_0 proof (subst of_int_0 [symmetric], rule tendsto_EM_remainder) show "eventually (\x. norm (?fs 1 x) \ ?C * (x + a) powr (-Re s - 1)) at_top" using eventually_ge_at_top[of 0] by eventually_elim (insert assms, auto simp: norm_mult norm_powr_real_powr) next fix x assume x: "x \ real_of_int 0" have [simp]: "-numeral n - (x :: real) = -x - numeral n" for x n by (simp add: algebra_simps) show "((\x. ?C / (-Re s) * (x + a) powr (-Re s)) has_real_derivative ?C * (x + a) powr (- Re s - 1)) (at x within {real_of_int 0..})" using assms x by (auto intro!: derivative_eq_intros) next have "(\y. ?C / (- Re s) * (a + real y) powr (- Re s)) \ 0" by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_real_sequentially filterlim_tendsto_add_at_top[OF tendsto_const]) (use assms in auto) thus "convergent (\y. ?C / (- Re s) * (real y + a) powr (- Re s))" by (auto simp: add_ac convergent_def) qed (intro integrable_EM_remainder' continuous_intros, insert assms, auto simp: add_eq_0_iff) } moreover have "(\n. I n - I' n) \ 0" proof - have "(\n. (complex_of_real (real n + a) powr (1 - s) - of_real (real n) powr (1 - s)) / (1 - s)) \ 0 / (1 - s)" using assms(3-5) by (intro filterlim_compose[OF _ filterlim_real_sequentially] tendsto_divide complex_powr_add_minus_powr_asymptotics) auto thus "(\n. I n - I' n) \ 0" by (simp add: I_def I'_def divide_simps) qed ultimately have "(\n. C - C' + C'' - D' n + D n + R n + (I n - I' n)) \ C - C' + C'' - 0 + 0 + R_lim + 0" unfolding D_def D'_def using assms by (intro tendsto_add tendsto_diff tendsto_const tendsto_mult_right_zero tendsto_neg_powr_complex_of_real filterlim_tendsto_add_at_top filterlim_real_sequentially) auto also have "C - C' + C'' - 0 + 0 + R_lim + 0 = (a powr - s / 2) + a powr (1 - s) / (s - 1) + R_lim" by (simp add: C_def C'_def C''_def field_simps) also have "\ = hurwitz_zeta a s" using assms by (simp add: hurwitz_zeta_def pre_zeta_def pre_zeta_aux_def R_lim_def scaleR_conv_of_real) finally have "(\n. C - C' + C'' - D' n + D n + R n + (I n - I' n)) \ hurwitz_zeta a s" . with ev show ?thesis by (blast intro: Lim_transform_eventually) qed lemma zeta_critical_strip: fixes s :: complex and a :: real defines "S \ (\n. \i=1..n. (of_nat i) powr - s)" defines "I \ (\n. of_nat n powr (1 - s) / (1 - s))" assumes s: "Re s > 0" "s \ 1" shows "(\n. S n - I n) \ zeta s" proof - from hurwitz_zeta_critical_strip[OF s zero_less_one] have "(\n. (\i hurwitz_zeta 1 s" by (simp add: add_ac) also have "(\n. (\in. (\i=1..n. of_nat i powr -s))" by (intro ext sum.reindex_bij_witness[of _ "\x. x - 1" Suc]) auto finally show ?thesis by (simp add: zeta_def S_def I_def) qed subsection \The non-vanishing of $\zeta$ for $\mathfrak{R}(s) \geq 1$\ text \ This proof is based on a sketch by Newman~\cite{newman1998analytic}, which was previously formalised in HOL Light by Harrison~\cite{harrison2009pnt}, albeit in a much more concrete and low-level style. Our aim here is to reproduce Newman's proof idea cleanly and on the same high level of abstraction. \ theorem zeta_Re_ge_1_nonzero: fixes s assumes "Re s \ 1" "s \ 1" shows "zeta s \ 0" proof (cases "Re s > 1") case False define a where "a = -Im s" from False assms have s [simp]: "s = 1 - \ * a" and a: "a \ 0" by (auto simp: complex_eq_iff a_def) show ?thesis proof assume "zeta s = 0" hence zero: "zeta (1 - \ * a) = 0" by simp with zeta_cnj[of "1 - \ * a"] have zero': "zeta (1 + \ * a) = 0" by simp \ \We define the function $Q(s) = \zeta(s)^2\zeta(s+ia)\zeta(s-ia)$ and its Dirichlet series. The objective will be to show that this function is entire and its Dirichlet series converges everywhere. Of course, $Q(s)$ has singularities at $1$ and $1\pm ia$, so we need to show they can be removed.\ define Q Q_fds where "Q = (\s. zeta s ^ 2 * zeta (s + \ * a) * zeta (s - \ * a))" and "Q_fds = fds_zeta ^ 2 * fds_shift (\ * a) fds_zeta * fds_shift (-\ * a) fds_zeta" let ?sings = "{1, 1 + \ * a, 1 - \ * a}" \ \We show that @{term Q} is locally bounded everywhere. This is the case because the poles of $\zeta(s)$ cancel with the zeros of $\zeta(s\pm ia)$ and vice versa. This boundedness is then enough to show that @{term Q} has only removable singularities.\ have Q_bigo_1: "Q \ O[at s](\_. 1)" for s proof - have Q_eq: "Q = (\s. (zeta s * zeta (s + \ * a)) * (zeta s * zeta (s - \ * a)))" by (simp add: Q_def power2_eq_square mult_ac) \ \The singularity of $\zeta(s)$ at 1 gets cancelled by the zero of $\zeta(s-ia)$:\ have bigo1: "(\s. zeta s * zeta (s - \ * a)) \ O[at 1](\_. 1)" if "zeta (1 - \ * a) = 0" "a \ 0" for a :: real proof - have "(\s. zeta (s - \ * a) - zeta (1 - \ * a)) \ O[at 1](\s. s - 1)" using that by (intro taylor_bigo_linear holomorphic_on_imp_differentiable_at[of _ "-{1 + \ * a}"] holomorphic_intros) (auto simp: complex_eq_iff) hence "(\s. zeta s * zeta (s - \ * a)) \ O[at 1](\s. 1 / (s - 1) * (s - 1))" using that by (intro landau_o.big.mult zeta_bigo_at_1) simp_all also have "(\s. 1 / (s - 1) * (s - 1)) \ \[at 1](\_. 1)" by (intro bigthetaI_cong) (auto simp: eventually_at_filter) finally show ?thesis . qed \ \The analogous result for $\zeta(s) \zeta(s+ia)$:\ have bigo1': "(\s. zeta s * zeta (s + \ * a)) \ O[at 1](\_. 1)" if "zeta (1 - \ * a) = 0" "a \ 0" for a :: real using bigo1[of "-a"] that zeta_cnj[of "1 - \ * a"] by simp \ \The singularity of $\zeta(s-ia)$ gets cancelled by the zero of $\zeta(s)$:\ have bigo2: "(\s. zeta s * zeta (s - \ * a)) \ O[at (1 + \ * a)](\_. 1)" if "zeta (1 - \ * a) = 0" "a \ 0" for a :: real proof - have "(\s. zeta s * zeta (s - \ * a)) \ O[filtermap (\s. s + \ * a) (at 1)](\_. 1)" using bigo1'[of a] that by (simp add: mult.commute landau_o.big.in_filtermap_iff) also have "filtermap (\s. s + \ * a) (at 1) = at (1 + \ * a)" using filtermap_at_shift[of "-\ * a" 1] by simp finally show ?thesis . qed \ \Again, the analogous result for $\zeta(s) \zeta(s+ia)$:\ have bigo2': "(\s. zeta s * zeta (s + \ * a)) \ O[at (1 - \ * a)](\_. 1)" if "zeta (1 - \ * a) = 0" "a \ 0" for a :: real using bigo2[of "-a"] that zeta_cnj[of "1 - \ * a"] by simp \ \Now the final case distinction to show $Q(s)\in O(1)$ for all $s\in\mathbb{C}$:\ consider "s = 1" | "s = 1 + \ * a" | "s = 1 - \ * a" | "s \ ?sings" by blast thus ?thesis proof cases case 1 thus ?thesis unfolding Q_eq using zero zero' a by (auto intro: bigo1 bigo1' landau_o.big.mult_in_1) next case 2 from a have "isCont (\s. zeta s * zeta (s + \ * a)) (1 + \ * a)" by (auto intro!: continuous_intros) with 2 show ?thesis unfolding Q_eq using zero zero' a by (auto intro: bigo2 landau_o.big.mult_in_1 continuous_imp_bigo_1) next case 3 from a have "isCont (\s. zeta s * zeta (s - \ * a)) (1 - \ * a)" by (auto intro!: continuous_intros) with 3 show ?thesis unfolding Q_eq using zero zero' a by (auto intro: bigo2' landau_o.big.mult_in_1 continuous_imp_bigo_1) qed (auto intro!: continuous_imp_bigo_1 continuous_intros simp: Q_def complex_eq_iff) qed \ \Thus, we can remove the singularities from @{term Q} and extend it to an entire function.\ have "\Q'. Q' holomorphic_on UNIV \ (\z\UNIV - ?sings. Q' z = Q z)" by (intro removable_singularities Q_bigo_1) (auto simp: Q_def complex_eq_iff intro!: holomorphic_intros) then obtain Q' where Q': "Q' holomorphic_on UNIV" "\z. z \ ?sings \ Q' z = Q z" by blast \ \@{term Q'} constitutes an analytic continuation of the Dirichlet series of @{term Q}.\ have eval_Q_fds: "eval_fds Q_fds s = Q' s" if "Re s > 1" for s proof - have "eval_fds Q_fds s = Q s" using that by (simp add: Q_fds_def Q_def eval_fds_mult eval_fds_power fds_abs_converges_mult fds_abs_converges_power eval_fds_zeta) also from that have "\ = Q' s" by (subst Q') auto finally show ?thesis . qed \ \Since $\zeta(s)$ and $\zeta(s\pm ia)$ are completely multiplicative Dirichlet series, the logarithm of their product can be rewritten into the following nice form:\ have ln_Q_fds_eq: "fds_ln 0 Q_fds = fds (\k. of_real (2 * mangoldt k / ln k * (1 + cos (a * ln k))))" proof - note simps = fds_ln_mult[where l' = 0 and l'' = 0] fds_ln_power[where l' = 0] fds_ln_prod[where l' = "\_. 0"] have "fds_ln 0 Q_fds = 2 * fds_ln 0 fds_zeta + fds_shift (\ * a) (fds_ln 0 fds_zeta) + fds_shift (-\ * a) (fds_ln 0 fds_zeta)" by (auto simp: Q_fds_def simps) also have "completely_multiplicative_function (fds_nth (fds_zeta :: complex fds))" by standard auto hence "fds_ln (0 :: complex) fds_zeta = fds (\n. mangoldt n /\<^sub>R ln (real n))" by (subst fds_ln_completely_multiplicative) (auto simp: fds_eq_iff) also have "2 * \ + fds_shift (\ * a) \ + fds_shift (-\ * a) \ = fds (\k. of_real (2 * mangoldt k / ln k * (1 + cos (a * ln k))))" (is "?a = ?b") proof (intro fds_eqI, goal_cases) case (1 n) then consider "n = 1" | "n > 1" by force hence "fds_nth ?a n = mangoldt n / ln (real n) * (2 + (n powr (\ * a) + n powr (-\ * a)))" by cases (auto simp: field_simps scaleR_conv_of_real numeral_fds) also have "n powr (\ * a) + n powr (-\ * a) = 2 * cos (of_real (a * ln n))" using 1 by (subst cos_exp_eq) (simp_all add: powr_def algebra_simps) also have "mangoldt n / ln (real n) * (2 + \) = of_real (2 * mangoldt n / ln n * (1 + cos (a * ln n)))" by (subst cos_of_real) simp_all finally show ?case by (simp add: fds_nth_fds') qed finally show ?thesis . qed \ \It is then obvious that this logarithm series has non-negative real coefficients.\ also have "nonneg_dirichlet_series \" proof (standard, goal_cases) case (1 n) from cos_ge_minus_one[of "a * ln n"] have "1 + cos (a * ln (real n)) \ 0" by linarith thus ?case using 1 by (cases "n = 0") (auto simp: complex_nonneg_Reals_iff fds_nth_fds' mangoldt_nonneg intro!: divide_nonneg_nonneg mult_nonneg_nonneg) qed \ \Therefore, the original series also has non-negative real coefficients.\ finally have nonneg: "nonneg_dirichlet_series Q_fds" by (rule nonneg_dirichlet_series_lnD) (auto simp: Q_fds_def) \ \By the Pringsheim--Landau theorem, a Dirichlet series with non-negative coefficnets that can be analytically continued to the entire complex plane must converge everywhere, i.\,e.\ its abscissa of (absolute) convergence is $-\infty$:\ have abscissa_Q_fds: "abs_conv_abscissa Q_fds \ 1" unfolding Q_fds_def by (auto intro!: abs_conv_abscissa_mult_leI abs_conv_abscissa_power_leI) with nonneg and eval_Q_fds and \Q' holomorphic_on UNIV\ have abscissa: "abs_conv_abscissa Q_fds = -\" by (intro entire_continuation_imp_abs_conv_abscissa_MInfty[where c = 1 and g = Q']) (auto simp: one_ereal_def) \ \This now leads to a contradiction in a very obvious way. If @{term Q_fds} is absolutely convergent, then the subseries corresponding to powers of 2 (\i.e. we delete all summands $a_n / n ^ s$ where $n$ is not a power of 2 from the sum) is also absolutely convergent. We denote this series with $R$.\ define R_fds where "R_fds = fds_primepow_subseries 2 Q_fds" have "conv_abscissa R_fds \ abs_conv_abscissa R_fds" by (rule conv_le_abs_conv_abscissa) also have "abs_conv_abscissa R_fds \ abs_conv_abscissa Q_fds" unfolding R_fds_def by (rule abs_conv_abscissa_restrict) also have "\ = -\" by (simp add: abscissa) finally have abscissa': "conv_abscissa R_fds = -\" by simp \ \Since $\zeta(s)$ and $\zeta(s \pm ia)$ have an Euler product expansion for $\mathfrak{R}(s) > 1$, we have \[R(s) = (1 - 2^{-s})^-2 (1 - 2^{-s+ia})^{-1} (1 - 2^{-s-ia})^{-1}\] there, and since $R$ converges everywhere and the right-hand side is holomorphic for $\mathfrak{R}(s) > 0$, the equation is also valid for all $s$ with $\mathfrak{R}(s) > 0$ by analytic continuation.\ have eval_R: "eval_fds R_fds s = 1 / ((1 - 2 powr -s) ^ 2 * (1 - 2 powr (-s + \ * a)) * (1 - 2 powr (-s - \ * a)))" (is "_ = ?f s") if "Re s > 0" for s proof - show ?thesis proof (rule analytic_continuation_open[where f = "eval_fds R_fds"]) show "?f holomorphic_on {s. Re s > 0}" by (intro holomorphic_intros) (auto simp: powr_def exp_eq_1 Ln_Reals_eq) next fix z assume z: "z \ {s. Re s > 1}" have [simp]: "completely_multiplicative_function (fds_nth fds_zeta)" by standard auto thus "eval_fds R_fds z = ?f z" using z by (simp add: R_fds_def Q_fds_def eval_fds_mult eval_fds_power fds_abs_converges_mult fds_abs_converges_power fds_primepow_subseries_euler_product_cm divide_simps powr_minus powr_diff powr_add fds_abs_summable_zeta) qed (insert that abscissa', auto intro!: exI[of _ 2] convex_connected open_halfspace_Re_gt convex_halfspace_Re_gt holomorphic_intros) qed \ \We now clearly have a contradiction: $R(s)$, being entire, is continuous everywhere, while the function on the right-hand side clearly has a pole at $0$.\ show False proof (rule not_tendsto_and_filterlim_at_infinity) have "((\b. (1-2 powr - b)\<^sup>2 * (1 - 2 powr (-b+\*a)) * (1 - 2 powr (-b-\*a))) \ 0) (at 0 within {s. Re s > 0})" (is "filterlim ?f' _ _") by (intro tendsto_eq_intros) (auto) moreover have "eventually (\s. s \ {s. Re s > 0}) (at 0 within {s. Re s > 0})" by (auto simp: eventually_at_filter) hence "eventually (\s. ?f' s \ 0) (at 0 within {s. Re s > 0})" by eventually_elim (auto simp: powr_def exp_eq_1 Ln_Reals_eq) ultimately have "filterlim ?f' (at 0) (at 0 within {s. Re s > 0})" by (simp add: filterlim_at) hence "filterlim ?f at_infinity (at 0 within {s. Re s > 0})" (is ?lim) by (intro filterlim_divide_at_infinity[OF tendsto_const] tendsto_mult_filterlim_at_infinity) auto also have ev: "eventually (\s. Re s > 0) (at 0 within {s. Re s > 0})" by (auto simp: eventually_at intro!: exI[of _ 1]) have "?lim \ filterlim (eval_fds R_fds) at_infinity (at 0 within {s. Re s > 0})" by (intro filterlim_cong refl eventually_mono[OF ev]) (auto simp: eval_R) finally show \ . next have "continuous (at 0 within {s. Re s > 0}) (eval_fds R_fds)" by (intro continuous_intros) (auto simp: abscissa') thus "((eval_fds R_fds \ eval_fds R_fds 0)) (at 0 within {s. Re s > 0})" by (auto simp: continuous_within) next have "0 \ {s. Re s \ 0}" by simp also have "{s. Re s \ 0} = closure {s. Re s > 0}" using closure_halfspace_gt[of "1::complex" 0] by (simp add: inner_commute) finally have "0 \ \" . thus "at 0 within {s. Re s > 0} \ bot" by (subst at_within_eq_bot_iff) auto qed qed qed (fact zeta_Re_gt_1_nonzero) subsection \Special values of the $\zeta$ functions\ theorem hurwitz_zeta_neg_of_nat: assumes "a > 0" shows "hurwitz_zeta a (-of_nat n) = -bernpoly (Suc n) a / of_nat (Suc n)" proof - have "-of_nat n \ (1::complex)" by (simp add: complex_eq_iff) hence "hurwitz_zeta a (-of_nat n) = pre_zeta a (-of_nat n) + a powr real (Suc n) / (-of_nat (Suc n))" unfolding zeta_def hurwitz_zeta_def using assms by (simp add: powr_of_real [symmetric]) also have "a powr real (Suc n) / (-of_nat (Suc n)) = - (a powr real (Suc n) / of_nat (Suc n))" by (simp add: divide_simps del: of_nat_Suc) also have "a powr real (Suc n) = a ^ Suc n" using assms by (intro powr_realpow) also have "pre_zeta a (-of_nat n) = pre_zeta_aux (Suc n) a (- of_nat n)" using assms by (intro pre_zeta_aux_eq_pre_zeta [symmetric]) auto also have "\ = of_real a powr of_nat n / 2 + (\i = 1..Suc n. (bernoulli (2 * i) / fact (2 * i)) *\<^sub>R (pochhammer (- of_nat n) (2 * i - 1) * of_real a powr (of_nat n - of_nat (2 * i - 1)))) + EM_remainder (Suc (2 * Suc n)) (\x. - (pochhammer (- of_nat n) (2*n+3) * of_real (x + a) powr (- of_nat n - 3))) 0" (is "_ = ?B + sum (\n. ?f (2 * n)) _ + _") unfolding pre_zeta_aux_def by (simp add: add_ac eval_nat_numeral) also have "?B = of_real (a ^ n) / 2" using assms by (subst powr_Reals_eq) (auto simp: powr_realpow) also have "pochhammer (-of_nat n :: complex) (2*n+3) = 0" by (subst pochhammer_eq_0_iff) auto finally have "hurwitz_zeta a (-of_nat n) = - (a ^ Suc n / of_nat (Suc n)) + (a ^ n / 2 + sum (\n. ?f (2 * n)) {1..Suc n})" by simp also have "sum (\n. ?f (2 * n)) {1..Suc n} = sum ?f ((*) 2 ` {1..Suc n})" by (intro sum.reindex_bij_witness[of _ "\i. i div 2" "\i. 2*i"]) auto also have "\ = (\i=2..2*n+2. ?f i)" proof (intro sum.mono_neutral_left ballI, goal_cases) case (3 i) hence "odd i" "i \ 1" by (auto elim!: evenE) thus ?case by (simp add: bernoulli_odd_eq_0) qed auto also have "\ = (\i=2..Suc n. ?f i)" proof (intro sum.mono_neutral_right ballI, goal_cases) case (3 i) hence "pochhammer (-of_nat n :: complex) (i - 1) = 0" by (subst pochhammer_eq_0_iff) auto thus ?case by simp qed auto also have "\ = (\i=Suc 1..Suc n. -of_real (real (Suc n choose i) * bernoulli i * a ^ (Suc n - i)) / of_nat (Suc n))" (is "sum ?lhs _ = sum ?f _") proof (intro sum.cong, goal_cases) case (2 i) hence "of_nat n - of_nat (i - 1) = (of_nat (Suc n - i) :: complex)" by (auto simp: of_nat_diff) also have "of_real a powr \ = of_real (a ^ (Suc n - i))" using 2 assms by (subst powr_nat) auto finally have A: "of_real a powr (of_nat n - of_nat (i - 1)) = \" . have "pochhammer (-of_nat n) (i - 1) = complex_of_real (pochhammer (-real n) (i - 1))" by (simp add: pochhammer_of_real [symmetric]) also have "pochhammer (-real n) (i - 1) = pochhammer (-of_nat (Suc n)) i / (-1 - real n)" using 2 by (subst (2) pochhammer_rec_if) auto also have "-1 - real n = -real (Suc n)" by simp finally have B: "pochhammer (-of_nat n) (i - 1) = -complex_of_real (pochhammer (-real (Suc n)) i / real (Suc n))" by (simp del: of_nat_Suc) have "?lhs i = -complex_of_real (bernoulli i * pochhammer (-real (Suc n)) i / fact i * a ^ (Suc n - i)) / of_nat (Suc n)" by (simp only: A B) (simp add: scaleR_conv_of_real) also have "bernoulli i * pochhammer (-real (Suc n)) i / fact i = (real (Suc n) gchoose i) * bernoulli i" using 2 by (subst gbinomial_pochhammer) (auto simp: minus_one_power_iff bernoulli_odd_eq_0) also have "real (Suc n) gchoose i = Suc n choose i" by (subst binomial_gbinomial) auto finally show ?case by simp qed auto also have "a ^ n / 2 + sum ?f {Suc 1..Suc n} = sum ?f {1..Suc n}" by (subst (2) sum.atLeast_Suc_atMost) (simp_all add: scaleR_conv_of_real del: of_nat_Suc) also have "-(a ^ Suc n / of_nat (Suc n)) + sum ?f {1..Suc n} = sum ?f {0..Suc n}" by (subst (2) sum.atLeast_Suc_atMost) (simp_all add: scaleR_conv_of_real) also have "\ = - bernpoly (Suc n) a / of_nat (Suc n)" unfolding sum_negf sum_divide_distrib [symmetric] by (simp add: bernpoly_def atLeast0AtMost) finally show ?thesis . qed lemma hurwitz_zeta_0 [simp]: "a > 0 \ hurwitz_zeta a 0 = 1 / 2 - a" using hurwitz_zeta_neg_of_nat[of a 0] by (simp add: bernpoly_def) lemma zeta_0 [simp]: "zeta 0 = -1 / 2" by (simp add: zeta_def) theorem zeta_neg_of_nat: "zeta (-of_nat n) = -of_real (bernoulli' (Suc n)) / of_nat (Suc n)" unfolding zeta_def by (simp add: hurwitz_zeta_neg_of_nat bernpoly_1') corollary zeta_trivial_zero: assumes "even n" "n \ 0" shows "zeta (-of_nat n) = 0" using zeta_neg_of_nat[of n] assms by (simp add: bernoulli'_odd_eq_0) theorem zeta_even_nat: "zeta (2 * of_nat n) = of_real ((-1) ^ Suc n * bernoulli (2 * n) * (2 * pi) ^ (2 * n) / (2 * fact (2 * n)))" proof (cases "n = 0") case False hence "(\k. 1 / of_nat (Suc k) ^ (2 * n)) sums zeta (of_nat (2 * n))" by (intro sums_zeta_nat) auto from sums_unique2 [OF this nat_even_power_sums_complex] False show ?thesis by simp qed (insert zeta_neg_of_nat[of 0], simp_all) corollary zeta_even_numeral: "zeta (numeral (Num.Bit0 n)) = of_real ((- 1) ^ Suc (numeral n) * bernoulli (numeral (num.Bit0 n)) * (2 * pi) ^ numeral (num.Bit0 n) / (2 * fact (numeral (num.Bit0 n))))" (is "_ = ?rhs") proof - have *: "(2 * numeral n :: nat) = numeral (Num.Bit0 n)" by (subst numeral.numeral_Bit0, subst mult_2, rule refl) have "numeral (Num.Bit0 n) = (2 * numeral n :: complex)" by (subst numeral.numeral_Bit0, subst mult_2, rule refl) also have "\ = 2 * of_nat (numeral n)" by (simp only: of_nat_numeral of_nat_mult) also have "zeta \ = ?rhs" by (subst zeta_even_nat) (simp only: *) finally show ?thesis . qed corollary zeta_neg_even_numeral [simp]: "zeta (-numeral (Num.Bit0 n)) = 0" proof - have "-numeral (Num.Bit0 n) = (-of_nat (numeral (Num.Bit0 n)) :: complex)" by simp also have "zeta \ = 0" proof (rule zeta_trivial_zero) have "numeral (Num.Bit0 n) = (2 * numeral n :: nat)" by (subst numeral.numeral_Bit0, subst mult_2, rule refl) also have "even \" by (rule dvd_triv_left) finally show "even (numeral (Num.Bit0 n) :: nat)" . qed auto finally show ?thesis . qed corollary zeta_neg_numeral: "zeta (-numeral n) = -of_real (bernoulli' (numeral (Num.inc n)) / numeral (Num.inc n))" proof - have "-numeral n = (- of_nat (numeral n) :: complex)" by simp also have "zeta \ = -of_real (bernoulli' (numeral (Num.inc n)) / numeral (Num.inc n))" by (subst zeta_neg_of_nat) (simp add: numeral_inc) finally show ?thesis . qed corollary zeta_neg1: "zeta (-1) = - 1 / 12" using zeta_neg_of_nat[of 1] by (simp add: eval_bernoulli) corollary zeta_neg3: "zeta (-3) = 1 / 120" by (simp add: zeta_neg_numeral) corollary zeta_neg5: "zeta (-5) = - 1 / 252" by (simp add: zeta_neg_numeral) corollary zeta_2: "zeta 2 = pi ^ 2 / 6" by (simp add: zeta_even_numeral) corollary zeta_4: "zeta 4 = pi ^ 4 / 90" by (simp add: zeta_even_numeral fact_num_eq_if) corollary zeta_6: "zeta 6 = pi ^ 6 / 945" by (simp add: zeta_even_numeral fact_num_eq_if) corollary zeta_8: "zeta 8 = pi ^ 8 / 9450" by (simp add: zeta_even_numeral fact_num_eq_if) subsection \Integral relation between $\Gamma$ and $\zeta$ function\ lemma assumes z: "Re z > 0" and a: "a > 0" shows Gamma_hurwitz_zeta_aux_integral: "Gamma z / (of_nat n + of_real a) powr z = (\s\{0<..}. (s powr (z - 1) / exp ((n+a) * s)) \lebesgue)" and Gamma_hurwitz_zeta_aux_integrable: "set_integrable lebesgue {0<..} (\s. s powr (z - 1) / exp ((n+a) * s))" proof - note integrable = absolutely_integrable_Gamma_integral' [OF z] let ?INT = "set_lebesgue_integral lebesgue {0<..} :: (real \ complex) \ complex" let ?INT' = "set_lebesgue_integral lebesgue {0<..} :: (real \ real) \ real" have meas1: "set_borel_measurable lebesgue {0<..} (\x. of_real x powr (z - 1) / of_real (exp ((n+a) * x)))" unfolding set_borel_measurable_def by (intro measurable_completion, subst measurable_lborel2, intro borel_measurable_continuous_on_indicator) (auto intro!: continuous_intros) show integrable1: "set_integrable lebesgue {0<..} (\s. s powr (z - 1) / exp ((n+a) * s))" using assms by (intro absolutely_integrable_Gamma_integral) auto from assms have pos: "0 < real n + a" by (simp add: add_nonneg_pos) hence "complex_of_real 0 \ of_real (real n + a)" by (simp only: of_real_eq_iff) also have "complex_of_real (real n + a) = of_nat n + of_real a" by simp finally have nz: "\ \ 0" by auto have "((\t. complex_of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0<..}" by (rule Gamma_integral_complex') fact+ hence "Gamma z = ?INT (\t. t powr (z - 1) / of_real (exp t))" using set_lebesgue_integral_eq_integral(2) [OF integrable] by (simp add: has_integral_iff exp_of_real) also have "lebesgue = density (distr lebesgue lebesgue (\t. (real n+a) * t)) (\x. ennreal (real n+a))" using lebesgue_real_scale[of "real n + a"] pos by auto also have "set_lebesgue_integral \ {0<..} (\t. of_real t powr (z - 1) / of_real (exp t)) = set_lebesgue_integral (distr lebesgue lebesgue (\t. (real n + a) * t)) {0<..} (\t. (real n + a) * t powr (z - 1) / exp t)" using integrable pos unfolding set_lebesgue_integral_def by (subst integral_density) (simp_all add: exp_of_real algebra_simps scaleR_conv_of_real set_integrable_def) also have "\ = ?INT (\s. (n + a) * (of_real (n+a) * of_real s) powr (z - 1) / of_real (exp ((n+a) * s)))" unfolding set_lebesgue_integral_def proof (subst integral_distr) show "(*) (real n + a) \ lebesgue \\<^sub>M lebesgue" using lebesgue_measurable_scaling[of "real n + a", where ?'a = real] unfolding real_scaleR_def . next have "(\x. (n+a) * (indicator {0<..} x *\<^sub>R (of_real x powr (z - 1) / of_real (exp x)))) \ lebesgue \\<^sub>M borel" using integrable unfolding set_integrable_def by (intro borel_measurable_times) simp_all thus "(\x. indicator {0<..} x *\<^sub>R (complex_of_real (real n + a) * complex_of_real x powr (z - 1) / exp x)) \ borel_measurable lebesgue" by simp qed (intro Bochner_Integration.integral_cong refl, insert pos, auto simp: indicator_def zero_less_mult_iff) also have "\ = ?INT (\s. ((n+a) powr z) * (s powr (z - 1) / exp ((n+a) * s)))" using pos by (intro set_lebesgue_integral_cong refl allI impI, simp, subst powr_times_real) (auto simp: powr_diff) also have "\ = (n + a) powr z * ?INT (\s. s powr (z - 1) / exp ((n+a) * s))" unfolding set_lebesgue_integral_def by (subst integral_mult_right_zero [symmetric]) simp_all finally show "Gamma z / (of_nat n + of_real a) powr z = ?INT (\s. s powr (z - 1) / exp ((n+a) * s))" using nz by (auto simp add: field_simps) qed lemma assumes x: "x > 0" and "a > 0" shows Gamma_hurwitz_zeta_aux_integral_real: "Gamma x / (real n + a) powr x = set_lebesgue_integral lebesgue {0<..} (\s. s powr (x - 1) / exp ((real n + a) * s))" and Gamma_hurwitz_zeta_aux_integrable_real: "set_integrable lebesgue {0<..} (\s. s powr (x - 1) / exp ((real n + a) * s))" proof - show "set_integrable lebesgue {0<..} (\s. s powr (x - 1) / exp ((real n + a) * s))" using absolutely_integrable_Gamma_integral[of "of_real x" "real n + a"] unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound, goal_cases) case 3 have "set_integrable lebesgue {0<..} (\xa. complex_of_real xa powr (of_real x - 1) / of_real (exp ((n + a) * xa)))" using assms by (intro Gamma_hurwitz_zeta_aux_integrable) auto also have "?this \ integrable lebesgue (\s. complex_of_real (indicator {0<..} s *\<^sub>R (s powr (x - 1) / (exp ((n+a) * s)))))" unfolding set_integrable_def by (intro Bochner_Integration.integrable_cong refl) (auto simp: powr_Reals_eq indicator_def) finally have "set_integrable lebesgue {0<..} (\s. s powr (x - 1) / exp ((n+a) * s))" unfolding set_integrable_def complex_of_real_integrable_eq . thus ?case by (simp add: set_integrable_def) qed (insert assms, auto intro!: AE_I2 simp: indicator_def norm_divide norm_powr_real_powr) from Gamma_hurwitz_zeta_aux_integral[of "of_real x" a n] and assms have "of_real (Gamma x / (real n + a) powr x) = set_lebesgue_integral lebesgue {0<..} (\s. complex_of_real s powr (of_real x - 1) / of_real (exp ((n+a) * s)))" (is "_ = ?I") by (auto simp: Gamma_complex_of_real powr_Reals_eq) also have "?I = lebesgue_integral lebesgue (\s. of_real (indicator {0<..} s *\<^sub>R (s powr (x - 1) / exp ((n+a) * s))))" unfolding set_lebesgue_integral_def using assms by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def powr_Reals_eq) also have "\ = of_real (set_lebesgue_integral lebesgue {0<..} (\s. s powr (x - 1) / exp ((n+a) * s)))" unfolding set_lebesgue_integral_def by (rule Bochner_Integration.integral_complex_of_real) finally show "Gamma x / (real n + a) powr x = set_lebesgue_integral lebesgue {0<..} (\s. s powr (x - 1) / exp ((real n + a) * s))" by (subst (asm) of_real_eq_iff) qed theorem assumes "Re z > 1" and "a > (0::real)" shows Gamma_times_hurwitz_zeta_integral: "Gamma z * hurwitz_zeta a z = (\x\{0<..}. (of_real x powr (z - 1) * of_real (exp (-a*x) / (1 - exp (-x)))) \lebesgue)" and Gamma_times_hurwitz_zeta_integrable: "set_integrable lebesgue {0<..} (\x. of_real x powr (z - 1) * of_real (exp (-a*x) / (1 - exp (-x))))" proof - from assms have z: "Re z > 0" by simp let ?INT = "set_lebesgue_integral lebesgue {0<..} :: (real \ complex) \ complex" let ?INT' = "set_lebesgue_integral lebesgue {0<..} :: (real \ real) \ real" have 1: "complex_set_integrable lebesgue {0<..} (\x. of_real x powr (z - 1) / of_real (exp ((real n + a) * x)))" for n by (rule Gamma_hurwitz_zeta_aux_integrable) (use assms in simp_all) have 2: "summable (\n. norm (indicator {0<..} s *\<^sub>R (of_real s powr (z - 1) / of_real (exp ((n + a) * s)))))" for s proof (cases "s > 0") case True hence "summable (\n. norm (of_real s powr (z - 1)) * exp (-a * s) * exp (-s) ^ n)" using assms by (intro summable_mult summable_geometric) simp_all with True show ?thesis by (simp add: norm_mult norm_divide exp_add exp_diff exp_minus field_simps exp_of_nat_mult [symmetric]) qed simp_all have 3: "summable (\n. \x. norm (indicator {0<..} x *\<^sub>R (complex_of_real x powr (z - 1) / complex_of_real (exp ((n + a) * x)))) \lebesgue)" proof - have "summable (\n. Gamma (Re z) * (real n + a) powr -Re z)" using assms by (intro summable_mult summable_hurwitz_zeta_real) simp_all also have "?this \ summable (\n. ?INT' (\s. norm (of_real s powr (z - 1) / of_real (exp ((n+a) * s)))))" proof (intro summable_cong always_eventually allI, goal_cases) case (1 n) have "Gamma (Re z) * (real n + a) powr -Re z = Gamma (Re z) / (real n + a) powr Re z" by (subst powr_minus) (simp_all add: field_simps) also from assms have "\ = (\x\{0<..}. (x powr (Re z-1) / exp ((n+a) * x)) \lebesgue)" by (subst Gamma_hurwitz_zeta_aux_integral_real) simp_all also have "\ = (\xa\{0<..}. norm (of_real xa powr (z-1) / of_real (exp ((n+a) * xa))) \lebesgue)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def norm_divide norm_powr_real_powr) finally show ?case . qed finally show ?thesis by (simp add: set_lebesgue_integral_def) qed have sum_eq: "(\n. indicator {0<..} s *\<^sub>R (of_real s powr (z - 1) / of_real (exp ((n+a) * s)))) = indicator {0<..} s *\<^sub>R (of_real s powr (z - 1) * of_real (exp (-a * s) / (1 - exp (-s))))" for s proof (cases "s > 0") case True hence "(\n. indicator {0..} s *\<^sub>R (of_real s powr (z - 1) / of_real (exp ((n+a) * s)))) = (\n. of_real s powr (z - 1) * of_real (exp (-a * s)) * of_real (exp (-s)) ^ n)" by (intro suminf_cong) (auto simp: exp_add exp_minus exp_of_nat_mult [symmetric] field_simps of_real_exp) also have "(\n. of_real s powr (z - 1) * of_real (exp (-a * s)) * of_real (exp (-s)) ^ n) = of_real s powr (z - 1) * of_real (exp (-a * s)) * (\n. of_real (exp (-s)) ^ n)" using True by (intro suminf_mult summable_geometric) simp_all also have "(\n. complex_of_real (exp (-s)) ^ n) = 1 / (1 - of_real (exp (-s)))" using True by (intro suminf_geometric) auto also have "of_real s powr (z - 1) * of_real (exp (-a * s)) * \ = of_real s powr (z - 1) * of_real (exp (-a * s) / (1 - exp (-s)))" using \a > 0\ by (auto simp add: divide_simps exp_minus) finally show ?thesis using True by simp qed simp_all show "set_integrable lebesgue {0<..} (\x. of_real x powr (z - 1) * of_real (exp (-a*x) / (1 - exp (-x))))" using 1 unfolding sum_eq [symmetric] set_integrable_def by (intro integrable_suminf[OF _ AE_I2] 2 3) have "(\n. ?INT (\s. s powr (z - 1) / exp ((n+a) * s))) sums lebesgue_integral lebesgue (\s. \n. indicator {0<..} s *\<^sub>R (s powr (z - 1) / exp ((n+a) * s)))" (is "?A sums ?B") using 1 unfolding set_lebesgue_integral_def set_integrable_def by (rule sums_integral[OF _ AE_I2[OF 2] 3]) also have "?A = (\n. Gamma z * (n + a) powr -z)" using assms by (subst Gamma_hurwitz_zeta_aux_integral [symmetric]) (simp_all add: powr_minus divide_simps) also have "?B = ?INT (\s. of_real s powr (z - 1) * of_real (exp (-a * s) / (1 - exp (-s))))" unfolding sum_eq set_lebesgue_integral_def .. finally have "(\n. Gamma z * (of_nat n + of_real a) powr -z) sums ?INT (\x. of_real x powr (z - 1) * of_real (exp (-a * x) / (1 - exp (-x))))" by simp moreover have "(\n. Gamma z * (of_nat n + of_real a) powr -z) sums (Gamma z * hurwitz_zeta a z)" using assms by (intro sums_mult sums_hurwitz_zeta) simp_all ultimately show "Gamma z * hurwitz_zeta a z = (\x\{0<..}. (of_real x powr (z - 1) * of_real (exp (-a * x) / (1 - exp (-x)))) \lebesgue)" by (rule sums_unique2 [symmetric]) qed corollary assumes "Re z > 1" shows Gamma_times_zeta_integral: "Gamma z * zeta z = (\x\{0<..}. (of_real x powr (z - 1) / of_real (exp x - 1)) \lebesgue)" (is ?th1) and Gamma_times_zeta_integrable: "set_integrable lebesgue {0<..} (\x. of_real x powr (z - 1) / of_real (exp x - 1))" (is ?th2) proof - have *: "(\x. indicator {0<..} x *\<^sub>R (of_real x powr (z - 1) * of_real (exp (-x) / (1 - exp (-x))))) = (\x. indicator {0<..} x *\<^sub>R (of_real x powr (z - 1) / of_real (exp x - 1)))" by (intro ext) (simp add: field_simps exp_minus indicator_def) from Gamma_times_hurwitz_zeta_integral [OF assms zero_less_one] and * show ?th1 by (simp add: zeta_def set_lebesgue_integral_def) from Gamma_times_hurwitz_zeta_integrable [OF assms zero_less_one] and * show ?th2 by (simp add: zeta_def set_integrable_def) qed corollary hurwitz_zeta_integral_Gamma_def: assumes "Re z > 1" "a > 0" shows "hurwitz_zeta a z = rGamma z * (\x\{0<..}. (of_real x powr (z - 1) * of_real (exp (-a * x) / (1 - exp (-x)))) \lebesgue)" proof - from assms have "Gamma z \ 0" by (subst Gamma_eq_zero_iff) (auto elim!: nonpos_Ints_cases) with Gamma_times_hurwitz_zeta_integral [OF assms] show ?thesis by (simp add: rGamma_inverse_Gamma field_simps) qed corollary zeta_integral_Gamma_def: assumes "Re z > 1" shows "zeta z = rGamma z * (\x\{0<..}. (of_real x powr (z - 1) / of_real (exp x - 1)) \lebesgue)" proof - from assms have "Gamma z \ 0" by (subst Gamma_eq_zero_iff) (auto elim!: nonpos_Ints_cases) with Gamma_times_zeta_integral [OF assms] show ?thesis by (simp add: rGamma_inverse_Gamma field_simps) qed lemma Gamma_times_zeta_has_integral: assumes "Re z > 1" shows "((\x. x powr (z - 1) / (of_real (exp x) - 1)) has_integral (Gamma z * zeta z)) {0<..}" (is "(?f has_integral _) _") proof - have "(?f has_integral set_lebesgue_integral lebesgue {0<..} ?f) {0<..}" using Gamma_times_zeta_integrable[OF assms] by (intro has_integral_set_lebesgue) auto also have "set_lebesgue_integral lebesgue {0<..} ?f = Gamma z * zeta z" using Gamma_times_zeta_integral[OF assms] by simp finally show ?thesis . qed lemma Gamma_times_zeta_has_integral_real: fixes z :: real assumes "z > 1" shows "((\x. x powr (z - 1) / (exp x - 1)) has_integral (Gamma z * Re (zeta z))) {0<..}" proof - from assms have *: "Re (of_real z) > 1" by simp have "((\x. Re (complex_of_real x powr (complex_of_real z - 1)) / (exp x - 1)) has_integral Gamma z * Re (zeta (complex_of_real z))) {0<..}" using has_integral_linear[OF Gamma_times_zeta_has_integral[OF *] bounded_linear_Re] by (simp add: o_def Gamma_complex_of_real) also have "?this \ ?thesis" using assms by (intro has_integral_cong) (auto simp: powr_Reals_eq) finally show ?thesis . qed lemma Gamma_integral_real': assumes x: "x > (0 :: real)" shows "((\t. t powr (x - 1) / exp t) has_integral Gamma x) {0<..}" using Gamma_integral_real[OF assms] by (subst has_integral_closure [symmetric]) auto subsection \An analytic proof of the infinitude of primes\ text \ We can now also do an analytic proof of the infinitude of primes. \ lemma primes_infinite_analytic: "infinite {p :: nat. prime p}" proof \ \Suppose the set of primes were finite.\ define P :: "nat set" where "P = {p. prime p}" assume fin: "finite P" \ \Then the Euler product form of the $\zeta$ function ranges over a finite set, and since each factor is holomorphic in the positive real half-space, the product is, too.\ define zeta' :: "complex \ complex" where "zeta' = (\s. (\p\P. inverse (1 - 1 / of_nat p powr s)))" have holo: "zeta' holomorphic_on A" if "A \ {s. Re s > 0}" for A proof - { fix p :: nat and s :: complex assume p: "p \ P" and s: "s \ A" from p have p': "real p > 1" by (subst of_nat_1 [symmetric], subst of_nat_less_iff) (simp add: prime_gt_Suc_0_nat P_def) have "norm (of_nat p powr s) = real p powr Re s" by (simp add: norm_powr_real_powr) also have "\ > real p powr 0" using p p' s that by (subst powr_less_cancel_iff) (auto simp: prime_gt_1_nat) finally have "of_nat p powr s \ 1" using p by (auto simp: P_def) } thus ?thesis by (auto simp: zeta'_def P_def intro!: holomorphic_intros) qed \ \Since the Euler product expansion of $\zeta(s)$ is valid for all $s$ with real value at least 1, and both $\zeta(s)$ and the Euler product must be equal in the positive real half-space punctured at 1 by analytic continuation.\ have eq: "zeta s = zeta' s" if "Re s > 0" "s \ 1" for s proof (rule analytic_continuation_open[of "{s. Re s > 1}" "{s. Re s > 0} - {1}" zeta zeta']) fix s assume s: "s \ {s. Re s > 1}" let ?f = "(\n. \p\n. if prime p then inverse (1 - 1 / of_nat p powr s) else 1)" have "eventually (\n. ?f n = zeta' s) sequentially" using eventually_ge_at_top[of "Max P"] proof eventually_elim case (elim n) have "P \ {}" by (auto simp: P_def intro!: exI[of _ 2]) with elim have "P \ {..n}" using fin by auto thus ?case unfolding zeta'_def by (intro prod.mono_neutral_cong_right) (auto simp: P_def) qed moreover from s have "?f \ zeta s" by (intro euler_product_zeta) auto ultimately have "(\_. zeta' s) \ zeta s" by (blast intro: Lim_transform_eventually) thus "zeta s = zeta' s" by (simp add: LIMSEQ_const_iff) qed (auto intro!: exI[of _ 2] open_halfspace_Re_gt connected_open_delete convex_halfspace_Re_gt holomorphic_intros holo that intro: convex_connected) \ \However, since the Euler product is holomorphic on the entire positive real half-space, it cannot have a pole at 1, while $\zeta(s)$ does have a pole at 1. Since they are equal in the punctured neighbourhood of 1, this is a contradiction.\ have ev: "eventually (\s. s \ {s. Re s > 0} - {1}) (at 1)" by (auto simp: eventually_at_filter intro!: open_halfspace_Re_gt eventually_mono[OF eventually_nhds_in_open[of "{s. Re s > 0}"]]) have "\is_pole zeta' 1" by (rule not_is_pole_holomorphic [of "{s. Re s > 0}"]) (auto intro!: holo open_halfspace_Re_gt) also have "is_pole zeta' 1 \ is_pole zeta 1" unfolding is_pole_def by (intro filterlim_cong refl eventually_mono [OF ev] eq [symmetric]) auto finally show False using is_pole_zeta by contradiction qed subsection \The periodic zeta function\ text \ The periodic zeta function $F(s, q)$ (as described e.\,g.\ by Apostol~\cite{apostol1976analytic} is related to the Hurwitz zeta function. It is periodic in \q\ with period 1 and it can be represented by a Dirichlet series that is absolutely convergent for $\mathfrak{R}(s) > 1$. If \q \ \\, it furthermore convergent for $\mathfrak{R}(s) > 0$. It is clear that for integer \q\, we have $F(s, q) = F(s, 0) = \zeta(s)$. Moreover, for non-integer \q\, $F(s, q)$ can be analytically continued to an entire function. \ definition fds_perzeta :: "real \ complex fds" where "fds_perzeta q = fds (\m. exp (2 * pi * \ * m * q))" text \ The definition of the periodic zeta function on the full domain is a bit unwieldy. The precise reasoning for this definition will be given later, and, in any case, it is probably more instructive to look at the derived ``alternative'' definitions later. \ definition perzeta :: "real \ complex \ complex" where "perzeta q' s = (if q' \ \ then zeta s else let q = frac q' in if s = 0 then \ / (2 * pi) * (pre_zeta q 1 - pre_zeta (1 - q) 1 + ln (1 - q) - ln q + pi * \) else if s \ \ then eval_fds (fds_perzeta q) s else complex_of_real (2 * pi) powr (s - 1) * \ * Gamma (1 - s) * (\ powr (-s) * hurwitz_zeta q (1 - s) - \ powr s * hurwitz_zeta (1 - q) (1 - s)))" interpretation fds_perzeta: periodic_fun_simple' fds_perzeta by standard (simp_all add: fds_perzeta_def exp_add ring_distribs exp_eq_polar cis_mult [symmetric] cis_multiple_2pi) interpretation perzeta: periodic_fun_simple' perzeta proof - have [simp]: "n + 1 \ \ \ n \ \" for n :: real by (simp flip: frac_eq_0_iff add: frac.plus_1) show "periodic_fun_simple' perzeta" by standard (auto simp: fun_eq_iff perzeta_def Let_def frac.plus_1) qed lemma perzeta_frac [simp]: "perzeta (frac q) = perzeta q" by (auto simp: perzeta_def fun_eq_iff Let_def) lemma fds_perzeta_frac [simp]: "fds_perzeta (frac q) = fds_perzeta q" using fds_perzeta.plus_of_int[of "frac q" "\q\"] by (simp add: frac_def) lemma abs_conv_abscissa_perzeta: "abs_conv_abscissa (fds_perzeta q) \ 1" proof (rule abs_conv_abscissa_leI) fix c assume c: "ereal c > 1" hence "summable (\n. n powr -c)" by (simp add: summable_real_powr_iff) also have "?this \ fds_abs_converges (fds_perzeta q) (of_real c)" unfolding fds_abs_converges_def by (intro summable_cong eventually_mono[OF eventually_gt_at_top[of 0]]) (auto simp: norm_divide norm_powr_real_powr fds_perzeta_def powr_minus field_simps) finally show "\s. s \ 1 = c \ fds_abs_converges (fds_perzeta q) s" by (intro exI[of _ "of_real c"]) auto qed lemma conv_abscissa_perzeta: "conv_abscissa (fds_perzeta q) \ 1" by (rule order.trans[OF conv_le_abs_conv_abscissa abs_conv_abscissa_perzeta]) lemma fds_perzeta__left_0 [simp]: "fds_perzeta 0 = fds_zeta" by (simp add: fds_perzeta_def fds_zeta_def) lemma perzeta_0_left [simp]: "perzeta 0 s = zeta s" by (simp add: perzeta_def eval_fds_zeta) lemma perzeta_int: "q \ \ \ perzeta q = zeta" by (simp add: perzeta_def fun_eq_iff) lemma fds_perzeta_int: "q \ \ \ fds_perzeta q = fds_zeta" by (auto simp: fds_perzeta.of_int elim!: Ints_cases) lemma sums_fds_perzeta: assumes "Re s > 1" shows "(\m. exp (2 * pi * \ * Suc m * q) / of_nat (Suc m) powr s) sums eval_fds (fds_perzeta q) s" proof - have "conv_abscissa (fds_perzeta q) \ 1" by (rule conv_abscissa_perzeta) also have "\ < ereal (Re s)" using assms by (simp add: one_ereal_def) finally have "fds_converges (fds_perzeta q) s" by (intro fds_converges) auto hence "(\n. fds_nth (fds_perzeta q) (Suc n) / nat_power (Suc n) s) sums eval_fds (fds_perzeta q) s" by (subst sums_Suc_iff) (auto simp: fds_converges_iff) thus ?thesis by (simp add: fds_perzeta_def) qed lemma sum_tendsto_fds_perzeta: assumes "Re s > 1" shows "(\n. \k\{0<..n}. exp (2 * real k * pi * q * \) * of_nat k powr - s) \ eval_fds (fds_perzeta q) s" proof - have "(\m. exp (2 * pi * \ * Suc m * q) / of_nat (Suc m) powr s) sums eval_fds (fds_perzeta q) s" by (intro sums_fds_perzeta assms) hence "(\n. \k) * of_nat (Suc k) powr -s) \ eval_fds (fds_perzeta q) s" (is "filterlim ?f _ _") by (simp add: sums_def powr_minus field_simps) also have "?f = (\n. \k\{0<..n}. exp (2 * real k * pi * q * \) * of_nat k powr -s)" by (intro ext sum.reindex_bij_betw sum.reindex_bij_witness[of _ "\n. n - 1" Suc]) auto finally show ?thesis by simp qed text \ Using the geometric series, it is easy to see that the Dirichlet series for $F(s, q)$ has bounded partial sums for non-integer \q\, so it must converge for any \s\ with $\mathfrak{R}(s) > 0$. \ lemma conv_abscissa_perzeta': assumes "q \ \" shows "conv_abscissa (fds_perzeta q) \ 0" proof (rule conv_abscissa_leI) fix c :: real assume c: "ereal c > 0" have "fds_converges (fds_perzeta q) (of_real c)" proof (rule bounded_partial_sums_imp_fps_converges) define \ where "\ = exp (2 * pi * \ * q)" have [simp]: "norm \ = 1" by (simp add: \_def) define M where "M = 2 / norm (1 - \)" from \q \ \\ have "\ \ 1" by (auto simp: \_def exp_eq_1) hence "M > 0" by (simp add: M_def) show "Bseq (\n. \k\n. fds_nth (fds_perzeta q) k / nat_power k 0)" unfolding Bseq_def proof (rule exI, safe) fix n :: nat show "norm (\k\n. fds_nth (fds_perzeta q) k / nat_power k 0) \ M" proof (cases "n = 0") case False have "(\k\n. fds_nth (fds_perzeta q) k / nat_power k 0) = (\k\{1..1 + (n - 1)}. \ ^ k)" using False by (intro sum.mono_neutral_cong_right) (auto simp: fds_perzeta_def fds_nth_fds exp_of_nat_mult [symmetric] mult_ac \_def) also have "\ = \ * (1 - \ ^ n) / (1 - \)" using \\ \ 1\ and False by (subst sum_gp_offset) simp also have "norm \ \ 1 * (norm (1::complex) + norm (\ ^ n)) / norm (1 - \)" unfolding norm_mult norm_divide by (intro mult_mono divide_right_mono norm_triangle_ineq4) auto also have "\ = M" by (simp add: M_def norm_power) finally show ?thesis . qed (use \M > 0\ in simp_all) qed fact+ qed (insert c, auto) thus "\s. s \ 1 = c \ fds_converges (fds_perzeta q) s" by (intro exI[of _ "of_real c"]) auto qed lemma fds_perzeta_one_half: "fds_perzeta (1 / 2) = fds (\n. (-1) ^ n)" using Complex.DeMoivre[of pi] by (intro fds_eqI) (auto simp: fds_perzeta_def exp_eq_polar mult_ac) lemma perzeta_one_half_1 [simp]: "perzeta (1 / 2) 1 = -ln 2" proof (rule sums_unique2) have *: "(1 / 2 :: real) \ \" using fraction_not_in_ints[of 2 1] by auto have "fds_converges (fds_perzeta (1 / 2)) 1" by (rule fds_converges, rule le_less_trans, rule conv_abscissa_perzeta') (use * in auto) hence "(\n. (-1) ^ Suc n / Suc n) sums eval_fds (fds_perzeta (1 / 2)) 1" unfolding fds_converges_altdef by (simp add: fds_perzeta_one_half) also from * have "eval_fds (fds_perzeta (1 / 2)) 1 = perzeta (1 / 2) 1" by (simp add: perzeta_def) finally show "(\n. -complex_of_real ((-1) ^ n / Suc n)) sums perzeta (1 / 2) 1" by simp hence "(\n. -complex_of_real ((-1) ^ n / Suc n)) sums -of_real (ln 2)" by (intro sums_minus sums_of_real alternating_harmonic_series_sums) thus "(\n. -complex_of_real ((-1) ^ n / Suc n)) sums -(ln 2)" by (simp flip: Ln_of_real) qed subsection \Hurwitz's formula\ text \ We now move on to prove Hurwitz's formula relating the Hurwitz zeta function and the periodic zeta function. We mostly follow Apostol's proof, although we do make some small changes in order to make the proof more amenable to Isabelle's complex analysis library. The big difference is that Apostol integrates along a circle with a slit, where the two sides of the slit lie on different branches of the integrand. This makes sense when looking as the integrand as a Riemann surface, but we do not have a notion of Riemann surfaces in Isabelle. It is therefore much easier to simply cut the circle into an upper and a lower half. In fact, the integral on the lower half can be reduced to the one on the upper half easily by symmetry, so we really only need to handle the integral on the upper half. The integration contour that we will use is therefore a semi-annulus in the upper half of the complex plane, centred around the origin. Now, first of all, we prove the existence of an important improper integral that we will need later. \ (* TODO: Move *) lemma set_integrable_bigo: fixes f g :: "real \ 'a :: {banach, real_normed_field, second_countable_topology}" assumes "f \ O(\x. g x)" and "set_integrable lborel {a..} g" assumes "\b. b \ a \ set_integrable lborel {a.. 0" "\x. x \ x0 \ norm (f x) \ C * norm (g x)" by (fastforce elim!: landau_o.bigE simp: eventually_at_top_linorder) define x0' where "x0' = max a x0" have "set_integrable lborel {a..x. C *\<^sub>R (indicator {x0'..} x *\<^sub>R g x))" unfolding set_integrable_def by (intro integrable_scaleR_right) (simp add: abs_mult norm_mult) next from assms(4) have "set_borel_measurable borel {x0'..} f" by (rule set_borel_measurable_subset) (auto simp: x0'_def) thus "(\x. indicator {x0'..} x *\<^sub>R f x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next show "AE x in lborel. norm (indicator {x0'..} x *\<^sub>R f x) \ norm (C *\<^sub>R (indicator {x0'..} x *\<^sub>R g x))" using C by (intro AE_I2) (auto simp: abs_mult indicator_def x0'_def) qed ultimately have "set_integrable lborel ({a.. {x0'..}) f" by (rule set_integrable_Un) auto also have "{a.. {x0'..} = {a..}" by (auto simp: x0'_def) finally show ?thesis . qed lemma set_integrable_Gamma_hurwitz_aux2_real: fixes s a :: real assumes "r > 0" and "a > 0" shows "set_integrable lborel {r..} (\x. x powr s * (exp (-a * x)) / (1 - exp (-x)))" (is "set_integrable _ _ ?g") proof (rule set_integrable_bigo) have "(\x. exp (-(a/2) * x)) integrable_on {r..}" using assms by (intro integrable_on_exp_minus_to_infinity) auto hence "set_integrable lebesgue {r..} (\x. exp (-(a/2) * x))" by (intro nonnegative_absolutely_integrable) simp_all thus "set_integrable lborel {r..} (\x. exp (-(a/2) * x))" by (simp add: set_integrable_def integrable_completion) next fix y :: real have "set_integrable lborel {r..y} ?g" using assms by (intro borel_integrable_atLeastAtMost') (auto intro!: continuous_intros) thus "set_integrable lborel {r.. O(\x. exp (-(a/2) * x))" by real_asymp qed (auto simp: set_borel_measurable_def) lemma set_integrable_Gamma_hurwitz_aux2: fixes s :: complex and a :: real assumes "r > 0" and "a > 0" shows "set_integrable lborel {r..} (\x. x powr -s * (exp (- a * x)) / (1 - exp (- x)))" (is "set_integrable _ _ ?g") proof - have "set_integrable lborel {r..} (\x. x powr -Re s * (exp (-a * x)) / (1 - exp (-x)))" (is "set_integrable _ _ ?g'") by (rule set_integrable_Gamma_hurwitz_aux2_real) (use assms in auto) also have "?this \ integrable lborel (\x. indicator {r..} x *\<^sub>R ?g' x)" by (simp add: set_integrable_def) also have "(\x. indicator {r..} x *\<^sub>R ?g' x) = (\x. norm (indicator {r..} x *\<^sub>R ?g x))" using assms by (auto simp: indicator_def norm_mult norm_divide norm_powr_real_powr fun_eq_iff exp_of_real exp_minus norm_inverse in_Reals_norm power2_eq_square divide_simps) finally show ?thesis unfolding set_integrable_def by (subst (asm) integrable_norm_iff) auto qed lemma closed_neg_Im_slit: "closed {z. Re z = 0 \ Im z \ 0}" proof - have "closed ({z. Re z = 0} \ {z. Im z \ 0})" by (intro closed_Int closed_halfspace_Re_eq closed_halfspace_Im_le) also have "{z. Re z = 0} \ {z. Im z \ 0} = {z. Re z = 0 \ Im z \ 0}" by blast finally show ?thesis . qed text \ We define tour semi-annulus path. When this path is mirrored into the lower half of the complex plane and subtracted from the original path and the outer radius tends to \\\, this becomes a Hankel contour extending to \-\\. \ definition hankel_semiannulus :: "real \ nat \ real \ complex" where "hankel_semiannulus r N = (let R = (2 * N + 1) * pi in part_circlepath 0 R 0 pi +++ \\Big half circle\ linepath (of_real (-R)) (of_real (-r)) +++ \\Line on the negative real axis\ part_circlepath 0 r pi 0 +++ \\Small half circle\ linepath (of_real r) (of_real R)) \\Line on the positive real axis\" lemma path_hankel_semiannulus [simp, intro]: "path (hankel_semiannulus r R)" and valid_path_hankel_semiannulus [simp, intro]: "valid_path (hankel_semiannulus r R)" and pathfinish_hankel_semiannulus [simp, intro]: "pathfinish (hankel_semiannulus r R) = pathstart (hankel_semiannulus r R)" by (simp_all add: hankel_semiannulus_def Let_def) text \ We set the stage for an application of the Residue Theorem. We define a function \[f(s, z) = z^{-s} \frac{\exp(az)}{1-\exp(-z)}\ ,\] which will be the integrand. However, the principal branch of $z^{-s}$ has a branch cut along the non-positive real axis, which is bad because a part of our integration path also lies on the non-positive real axis. We therefore choose a slightly different branch of $z^{-s}$ by moving the logarithm branch along by $90^{\circ}$ so that the branch cut lies on the non-positive imaginary axis instead. \ context fixes a :: real fixes f :: "complex \ complex \ complex" and g :: "complex \ real \ complex" and h :: "real \ complex \ real \ complex" and Res :: "complex \ nat \ complex" and Ln' :: "complex \ complex" and F :: "real \ complex \ complex" assumes a: "a \ {0<..1}" \\Our custom branch of the logarithm\ defines "Ln' \ (\z. ln (-\ * z) + \ * pi / 2)" \\The integrand\ defines "f \ (\s z. exp (Ln' z * (-s) + of_real a * z) / (1 - exp z))" \\The integrand on the negative real axis\ defines "g \ (\s x. complex_of_real x powr -s * of_real (exp (-a*x)) / of_real (1 - exp (-x)))" \\The integrand on the circular arcs\ defines "h \ (\r s t. r * \ * cis t * exp (a * (r * cis t) - (ln r + \ * t) * s) / (1 - exp (r * cis t)))" \\The interesting part of the residues\ defines "Res \ (\s k. exp (of_real (2 * real k * pi * a) * \) * of_real (2 * real k * pi) powr (-s))" \\The periodic zeta function (at least on $\mathfrak{R}(s) > 1$ half-plane)\ defines "F \ (\q. eval_fds (fds_perzeta q))" begin text \ First, some basic properties of our custom branch of the logarithm: \ lemma Ln'_i: "Ln' \ = \ * pi / 2" by (simp add: Ln'_def) lemma Ln'_of_real_pos: assumes "x > 0" shows "Ln' (of_real x) = of_real (ln x)" proof - have "Ln' (of_real x) = Ln (of_real x * (-\)) + \ * pi / 2" by (simp add: Ln'_def mult_ac) also have "\ = of_real (ln x)" using assms by (subst Ln_times_of_real) (auto simp: Ln_of_real) finally show ?thesis . qed lemma Ln'_of_real_neg: assumes "x < 0" shows "Ln' (of_real x) = of_real (ln (-x)) + \ * pi" proof - have "Ln' (of_real x) = Ln (of_real (-x) * \) + \ * pi / 2" by (simp add: Ln'_def mult_ac) also have "\ = of_real (ln (-x)) + \ * pi" using assms by (subst Ln_times_of_real) (auto simp: Ln_Reals_eq) finally show ?thesis . qed lemma Ln'_times_of_real: "Ln' (of_real x * z) = of_real (ln x) + Ln' z" if "x > 0" "z \ 0" for z x proof - have "Ln' (of_real x * z) = Ln (of_real x * (- \ * z)) + \ * pi / 2" by (simp add: Ln'_def mult_ac) also have "\ = of_real (ln x) + Ln' z" using that by (subst Ln_times_of_real) (auto simp: Ln'_def Ln_of_real) finally show ?thesis . qed lemma Ln'_cis: assumes "t \ {-pi / 2<..3 / 2 * pi}" shows "Ln' (cis t) = \ * t" proof - have "exp (\ * pi / 2) = \" by (simp add: exp_eq_polar) hence "Ln (- (\ * cis t)) = \ * (t - pi / 2)" using assms by (intro Ln_unique) (auto simp: algebra_simps exp_diff cis_conv_exp) thus ?thesis by (simp add: Ln'_def algebra_simps) qed text \ Next, we show that the line and circle integrals are holomorphic using Leibniz's rule: \ lemma contour_integral_part_circlepath_h: assumes r: "r > 0" shows "contour_integral (part_circlepath 0 r 0 pi) (f s) = integral {0..pi} (h r s)" proof - have "contour_integral (part_circlepath 0 r 0 pi) (f s) = integral {0..pi} (\t. f s (r * cis t) * r * \ * cis t)" by (simp add: contour_integral_part_circlepath_eq) also have "integral {0..pi} (\t. f s (r * cis t) * r * \ * cis t) = integral {0..pi} (h r s)" proof (intro integral_cong) fix t assume t: "t \ {0..pi}" have "-(pi / 2) < 0" by simp also have "0 \ t" using t by simp finally have "t \ {-(pi/2)<..3/2*pi}" using t by auto thus "f s (r * cis t) * r * \ * cis t = h r s t" using r by (simp add: f_def Ln'_times_of_real Ln'_cis h_def Ln_Reals_eq) qed finally show ?thesis . qed lemma integral_g_holomorphic: assumes "b > 0" shows "(\s. integral {b..c} (g s)) holomorphic_on A" proof - define g' where "g' = (\s t. - (of_real t powr (-s)) * complex_of_real (ln t) * of_real (exp (- (a * t))) / of_real (1 - exp (- t)))" have "(\s. integral (cbox b c) (g s)) holomorphic_on UNIV" proof (rule leibniz_rule_holomorphic) fix s :: complex and t :: real assume "t \ cbox b c" thus "((\s. g s t) has_field_derivative g' s t) (at s)" using assms by (auto simp: g_def g'_def powr_def Ln_Reals_eq intro!: derivative_eq_intros) next fix s show "g s integrable_on cbox b c" for s unfolding g_def using assms by (intro integrable_continuous continuous_intros) auto next show "continuous_on (UNIV \ cbox b c) (\(s, t). g' s t)" using assms by (auto simp: g'_def case_prod_unfold intro!: continuous_intros) qed auto thus ?thesis by auto qed lemma integral_h_holomorphic: assumes r: "r \ {0<..<2}" shows "(\s. integral {b..c} (h r s)) holomorphic_on A" proof - have no_sing: "exp (r * cis t) \ 1" for t proof define z where "z = r * cis t" assume "exp z = 1" then obtain n where "norm z = 2 * pi * of_int \n\" by (auto simp: exp_eq_1 cmod_def abs_mult) moreover have "norm z = r" using r by (simp add: z_def norm_mult) ultimately have r_eq: "r = 2 * pi * of_int \n\" by simp with r have "n \ 0" by auto moreover from r have "r < 2 * pi" using pi_gt3 by simp with r_eq have "\n\ < 1" by simp ultimately show False by simp qed define h' where "h' = (\s t. exp (a * r * cis t - (ln r + \ * t) * s) * (-ln r - \ * t) * (r * \ * cis t) / (1 - exp (r * cis t)))" have "(\s. integral (cbox b c) (h r s)) holomorphic_on UNIV" proof (rule leibniz_rule_holomorphic) fix s t assume "t \ cbox b c" thus "((\s. h r s t) has_field_derivative h' s t) (at s)" using no_sing r by (auto intro!: derivative_eq_intros simp: h_def h'_def mult_ac Ln_Reals_eq) next fix s show "h r s integrable_on cbox b c" using no_sing unfolding h_def by (auto intro!: integrable_continuous_real continuous_intros) next show "continuous_on (UNIV \ cbox b c) (\(s, t). h' s t)" using no_sing by (auto simp: h'_def case_prod_unfold intro!: continuous_intros) qed auto thus ?thesis by auto qed text \ We now move on to the core result, which uses the Residue Theorem to relate a contour integral along a semi-annulus to a partial sum of the periodic zeta function. \ lemma hurwitz_formula_integral_semiannulus: fixes N :: nat and r :: real and s :: complex defines "R \ real (2 * N + 1) * pi" assumes "r > 0" and "r < 2" shows "exp (-\ * pi * s) * integral {r..R} (\x. x powr (-s) * exp (-a * x) / (1 - exp (-x))) + integral {r..R} (\x. x powr (-s) * exp (a * x) / (1 - exp x)) + contour_integral (part_circlepath 0 R 0 pi) (f s) + contour_integral (part_circlepath 0 r pi 0) (f s) = -2 * pi * \ * exp (- s * of_real pi * \ / 2) * (\k\{0<..N}. Res s k)" (is ?thesis1) and "f s contour_integrable_on hankel_semiannulus r N" proof - have "2 < 1 * pi" using pi_gt3 by simp also have "\ \ R" unfolding R_def by (intro mult_right_mono) auto finally have "R > 2" by (auto simp: R_def) note r_R = \r > 0\ \r < 2\ this \\We integrate along the edge of a semi-annulus in the upper half of the complex plane. It consists of a big semicircle, a small semicircle, and two lines connecting the two circles, one on the positive real axis and one on the negative real axis. The integral along the big circle will vanish as the radius of the circle tends to \\\, whereas the remaining path becomes a Hankel contour, and the integral along that Hankel contour is what we are interested in, since it is connected to the Hurwitz zeta function.\ define big_circle where "big_circle = part_circlepath 0 R 0 pi" define small_circle where "small_circle = part_circlepath 0 r pi 0" define neg_line where "neg_line = linepath (complex_of_real (-R)) (complex_of_real (-r))" define pos_line where "pos_line = linepath (complex_of_real r) (complex_of_real R)" define \ where "\ = hankel_semiannulus r N" have \_altdef: "\ = big_circle +++ neg_line +++ small_circle +++ pos_line" by (simp add: \_def R_def add_ac hankel_semiannulus_def big_circle_def neg_line_def small_circle_def pos_line_def) have [simp]: "path \" "valid_path \" "pathfinish \ = pathstart \" by (simp_all add: \_def) \\The integrand has a branch cut along the non-positive imaginary axis and additional simple poles at $2n\pi i$ for any \n \ \\<^sub>>\<^sub>0\. The radius of the smaller circle will always be less than $2\pi$ and the radius of the bigger circle of the form $(2N+1)\pi$, so we always have precisely the first $N$ poles inside our path.\ define sngs where "sngs = (\n. of_real (2 * pi * real n) * \) ` {0<..}" define sngs' where "sngs' = (\n. of_real (2 * pi * real n) * \) ` {0<..N}" have sngs_subset: "sngs' \ sngs" unfolding sngs_def sngs'_def by (intro image_mono) auto have closed_sngs [intro]: "closed (sngs - sngs')" unfolding sngs_def proof (rule discrete_imp_closed[of "2 * pi"]; safe?) fix m n :: nat assume "dist (of_real (2 * pi * real m) * \) (of_real (2 * pi * real n) * \) < 2 * pi" also have "dist (of_real (2 * pi * real m) * \) (of_real (2 * pi * real n) * \) = norm (of_real (2 * pi * real m) * \ - of_real (2 * pi * real n) * \)" by (simp add: dist_norm) also have "of_real (2 * pi * real m) * \ - of_real (2 * pi * real n) * \ = of_real (2 * pi) * \ * of_int (int m - int n)" by (simp add: algebra_simps) also have "norm \ = 2 * pi * of_int \int m - int n\" unfolding norm_mult norm_of_int by (simp add: norm_mult) finally have "\real m - real n\ < 1" by simp hence "m = n" by linarith thus "of_real (2 * pi * real m) * \ = of_real (2 * pi * real n) * \" by simp qed auto \\We define an area within which the integrand is holomorphic. Choosing this area as big as possible makes things easier later on, so we only remove the branch cut and the poles.\ define S where "S = - {z. Re z = 0 \ Im z \ 0} - (sngs - sngs')" define S' where "S' = - {z. Re z = 0 \ Im z \ 0}" have sngs: "exp z = 1 \ z \ sngs" if "Re z \ 0 \ Im z > 0" for z proof assume "exp z = 1" then obtain n where n: "z = 2 * pi * of_int n * \" unfolding exp_eq_1 by (auto simp: complex_eq_iff mult_ac) moreover from n and pi_gt_zero and that have "n > 0" by (auto simp: zero_less_mult_iff) ultimately have "z = of_real (2 * pi * nat n) * \" and "nat n \ {0<..}" by auto thus "z \ sngs" unfolding sngs_def by blast qed (insert that, auto simp: sngs_def exp_eq_polar) \\We show that the path stays within the well-behaved area.\ have "path_image neg_line = of_real ` {-R..-r}" using r_R by (auto simp: neg_line_def closed_segment_Reals closed_segment_eq_real_ivl) hence "path_image neg_line \ S - sngs'" using r_R sngs_subset by (auto simp: S_def sngs_def complex_eq_iff) have "path_image pos_line = of_real ` {r..R}" using r_R by (auto simp: pos_line_def closed_segment_Reals closed_segment_eq_real_ivl) hence "path_image pos_line \ S - sngs'" using r_R sngs_subset by (auto simp: S_def sngs_def complex_eq_iff) have part_circlepath_in_S: "z \ S - sngs'" if "z \ path_image (part_circlepath 0 r' 0 pi) \ z \ path_image (part_circlepath 0 r' pi 0)" and "r' > 0" "r' \ (\n. 2 * pi * real n) ` {0<..}" for z r' proof - have z: "norm z = r' \ Im z \ 0" using that by (auto simp: path_image_def part_circlepath_def norm_mult Im_exp linepath_def intro!: mult_nonneg_nonneg sin_ge_zero) hence "Re z \ 0 \ Im z > 0" using that by (auto simp: cmod_def) moreover from z and that have "z \ sngs" by (auto simp: sngs_def norm_mult image_iff) ultimately show "z \ S - sngs'" using sngs_subset by (auto simp: S_def) qed { fix n :: nat assume n: "n > 0" have "r < 2 * pi * 1" using pi_gt3 r_R by simp also have "\ \ 2 * pi * real n" using n by (intro mult_left_mono) auto finally have "r < \" . } hence "r \ (\n. 2 * pi * real n) ` {0<..}" using r_R by auto from part_circlepath_in_S[OF _ _ this] r_R have "path_image small_circle \ S - sngs'" by (auto simp: small_circle_def) { fix n :: nat assume n: "n > 0" "2 * pi * real n = real (2 * N + 1) * pi" hence "real (2 * n) = real (2 * N + 1)" unfolding of_nat_mult by simp hence False unfolding of_nat_eq_iff by presburger } hence "R \ (\n. 2 * pi * real n) ` {0<..}" unfolding R_def by force from part_circlepath_in_S[OF _ _ this] r_R have "path_image big_circle \ S - sngs'" by (auto simp: big_circle_def) note path_images = \path_image small_circle \ S - sngs'\ \path_image big_circle \ S - sngs'\ \path_image neg_line \ S - sngs'\ \path_image pos_line \ S - sngs'\ have "path_image \ \ S - sngs'" using path_images by (auto simp: \_altdef path_image_join big_circle_def neg_line_def small_circle_def pos_line_def) \\We need to show that the complex plane is still connected after we removed the branch cut and the poles. We do this by showing that the complex plane with the branch cut removed is starlike and therefore connected. Then we remove the (countably many) poles, which does not break connectedness either.\ have "open S" using closed_neg_Im_slit by (auto simp: S_def) have "starlike (UNIV - {z. Re z = 0 \ Im z \ 0})" (is "starlike ?S'") unfolding starlike_def proof (rule bexI ballI)+ have "1 \ 1 * pi" using pi_gt3 by simp also have "\ < (2 + 2 * real N) * pi" by (intro mult_strict_right_mono) auto finally show *: "\ \ ?S'" by (auto simp: S_def) fix z assume z: "z \ ?S'" have "closed_segment \ z \ {z. Re z = 0 \ Im z \ 0} = {}" proof safe fix s assume s: "s \ closed_segment \ z" "Re s = 0" "Im s \ 0" then obtain t where t: "t \ {0..1}" "s = linepath \ z t" using linepath_image_01 by blast with z s t have z': "Re z = 0" "Im z > 0" by (auto simp: Re_linepath' S_def linepath_0') with s have "Im s \ closed_segment 1 (Im z) \ Im s \ 0" by (subst (asm) closed_segment_same_Re) auto with z' show "s \ {}" by (auto simp: closed_segment_eq_real_ivl split: if_splits) qed thus "closed_segment \ z \ ?S'" by (auto simp: S_def) qed hence "connected ?S'" by (rule starlike_imp_connected) hence "connected S'" by (simp add: Compl_eq_Diff_UNIV S'_def) have "connected S" unfolding S_def by (rule connected_open_diff_countable) (insert \connected S'\, auto simp: sngs_def closed_neg_Im_slit S'_def) \\The integrand is now clearly holomorphic on @{term "S - sngs'"} and we can apply the Residue Theorem.\ have holo: "f s holomorphic_on (S - sngs')" unfolding f_def Ln'_def S_def using sngs by (auto intro!: holomorphic_intros simp: complex_nonpos_Reals_iff) have "contour_integral \ (f s) = of_real (2 * pi) * \ * (\z\sngs'. winding_number \ z * residue (f s) z)" proof (rule Residue_theorem) show "\z. z \ S \ winding_number \ z = 0" proof safe fix z assume "z \ S" hence "Re z = 0 \ Im z \ 0 \ z \ sngs - sngs'" by (auto simp: S_def) thus "winding_number \ z = 0" proof define x where "x = -Im z" assume "Re z = 0 \ Im z \ 0" hence x: "z = -of_real x * \" "x \ 0" unfolding complex_eq_iff by (simp_all add: x_def) obtain B where "\z. norm z \ B \ winding_number \ z = 0" using winding_number_zero_at_infinity[of \] by auto hence "winding_number \ (-of_real (max B 0) * \) = 0" by (auto simp: norm_mult) also have "winding_number \ (-of_real (max B 0) * \) = winding_number \ z" proof (rule winding_number_eq) from x have "closed_segment (-of_real (max B 0) * \) z \ {z. Re z = 0 \ Im z \ 0}" by (auto simp: closed_segment_same_Re closed_segment_eq_real_ivl) with \path_image \ \ S - sngs'\ show "closed_segment (-of_real (max B 0) * \) z \ path_image \ = {}" by (auto simp: S_def) qed auto finally show "winding_number \ z = 0" . next assume z: "z \ sngs - sngs'" show "winding_number \ z = 0" proof (rule winding_number_zero_outside) have "path_image \ = path_image big_circle \ path_image neg_line \ path_image small_circle \ path_image pos_line" unfolding \_altdef small_circle_def big_circle_def pos_line_def neg_line_def by (simp add: path_image_join Un_assoc) also have "\ \ cball 0 ((2 * N + 1) * pi)" using r_R by (auto simp: small_circle_def big_circle_def pos_line_def neg_line_def path_image_join norm_mult R_def path_image_part_circlepath' in_Reals_norm closed_segment_Reals closed_segment_eq_real_ivl) finally show "path_image \ \ \" . qed (insert z, auto simp: sngs_def sngs'_def norm_mult) qed qed qed (insert \path_image \ \ S - sngs'\ \connected S\ \open S\ holo, auto simp: sngs'_def) \\We can use Wenda Li's framework to compute the winding numbers at the poles and show that they are all 1.\ also have "winding_number \ z = 1" if "z \ sngs'" for z proof - have "r < 2 * pi * 1" using pi_gt3 r_R by simp also have "\ \ 2 * pi * real n" if "n > 0" for n using that by (intro mult_left_mono) auto finally have norm_z: "norm z > r" "norm z < R" using that r_R by (auto simp: sngs'_def norm_mult R_def) have "cindex_pathE big_circle z = -1" using r_R that unfolding big_circle_def by (subst cindex_pathE_circlepath_upper(1)) (auto simp: sngs'_def norm_mult R_def) have "cindex_pathE small_circle z = -1" using r_R that norm_z unfolding small_circle_def by (subst cindex_pathE_reversepath', subst reversepath_part_circlepath, subst cindex_pathE_circlepath_upper(2)) (auto simp: sngs'_def norm_mult) have "cindex_pathE neg_line z = 0" "cindex_pathE pos_line z = 0" unfolding neg_line_def pos_line_def using r_R that by (subst cindex_pathE_linepath; force simp: neg_line_def cindex_pathE_linepath closed_segment_Reals closed_segment_eq_real_ivl sngs'_def complex_eq_iff)+ note indices = \cindex_pathE big_circle z = -1\ \cindex_pathE small_circle z = -1\ \cindex_pathE neg_line z = 0\ \cindex_pathE pos_line z = 0\ show ?thesis unfolding \_altdef big_circle_def small_circle_def pos_line_def neg_line_def by eval_winding (insert indices path_images that, auto simp: big_circle_def small_circle_def pos_line_def neg_line_def) qed hence "(\z\sngs'. winding_number \ z * residue (f s) z) = (\z\sngs'. residue (f s) z)" by simp also have "\ = (\k\{0<..N}. residue (f s) (2 * pi * of_nat k * \))" unfolding sngs'_def by (subst sum.reindex) (auto intro!: inj_onI simp: o_def) \\Next, we compute the residues at each pole.\ also have "residue (f s) (2 * pi * of_nat k * \) = -exp (- s * of_real pi * \ / 2) * Res s k" if "k \ {0<..N}" for k unfolding f_def proof (subst residue_simple_pole_deriv) show "open S'" using closed_neg_Im_slit by (auto simp: S'_def) show "connected S'" by fact show "(\z. exp (Ln' z * (-s) + of_real a * z)) holomorphic_on S'" "(\z. 1 - exp z) holomorphic_on S'" by (auto simp: S'_def Ln'_def complex_nonpos_Reals_iff intro!: holomorphic_intros) have "((\z. 1 - exp z) has_field_derivative -exp (2 * pi * k * \)) (at (of_real (2 * pi * real k) * \))" by (auto intro!: derivative_eq_intros) also have "-exp (2 * pi * k * \) = -1" by (simp add: exp_eq_polar) finally show "((\z. 1 - exp z) has_field_derivative -1) (at (of_real (2 * pi * real k) * \))" . have "Im (of_real (2 * pi * real k) * \) > 0" using pi_gt_zero that by auto thus "of_real (2 * pi * real k) * \ \ S'" by (simp add: S'_def) have "exp (\ * pi / 2) = \" by (simp add: exp_eq_polar) hence "exp (Ln' (complex_of_real (2 * pi * real k) * \) * -s + of_real a * (of_real (2 * pi * real k) * \)) / -1 = - exp (2 * k * a * pi * \ - s * pi * \ / 2 - s * ln (2 * k * pi))" (is "?R = _") using that by (subst Ln'_times_of_real) (simp_all add: Ln'_i algebra_simps exp_diff) also have "\ = -exp (- s * of_real pi * \ / 2) * Res s k" using that by (simp add: Res_def exp_diff powr_def exp_minus inverse_eq_divide Ln_Reals_eq mult_ac) finally show "?R = -exp (- s * of_real pi * \ / 2) * Res s k" . qed (insert that, auto simp: S'_def exp_eq_polar) hence "(\k\{0<..N}. residue (f s) (2 * pi * of_nat k * \)) = -exp (- s * of_real pi * \ / 2) * (\k\{0<..N}. Res s k)" by (simp add: sum_distrib_left) \\This gives us the final result:\ finally have "contour_integral \ (f s) = -2 * pi * \ * exp (- s * of_real pi * \ / 2) * (\k\{0<..N}. Res s k)" by simp \\Lastly, we decompose the contour integral into its four constituent integrals because this makes them somewhat nicer to work with later on.\ also show "f s contour_integrable_on \" proof (rule contour_integrable_holomorphic_simple) show "path_image \ \ S - sngs'" by fact have "closed sngs'" by (intro finite_imp_closed) (auto simp: sngs'_def) with \open S\ show "open (S - sngs')" by auto qed (insert holo, auto) hence eq: "contour_integral \ (f s) = contour_integral big_circle (f s) + contour_integral neg_line (f s) + contour_integral small_circle (f s) + contour_integral pos_line (f s)" unfolding \_altdef big_circle_def neg_line_def small_circle_def pos_line_def by simp also have "contour_integral neg_line (f s) = integral {-R..-r} (\x. f s (complex_of_real x))" unfolding neg_line_def using r_R by (subst contour_integral_linepath_Reals_eq) auto also have "\ = exp (- \ * pi * s) * integral {r..R} (\x. exp (-ln x * s) * exp (-a * x) / (1 - exp (-x)))" (is "_ = _ * ?I") unfolding integral_mult_right [symmetric] using r_R by (subst Henstock_Kurzweil_Integration.integral_reflect_real [symmetric], intro integral_cong) (auto simp: f_def exp_of_real Ln'_of_real_neg exp_minus exp_Reals_eq exp_diff exp_add field_simps) also have "?I = integral {r..R} (\x. x powr (-s) * exp (-a * x) / (1 - exp (-x)))" using r_R by (intro integral_cong) (auto simp: powr_def Ln_Reals_eq exp_minus exp_diff field_simps) also have "contour_integral pos_line (f s) = integral {r..R} (\x. f s (complex_of_real x))" unfolding pos_line_def using r_R by (subst contour_integral_linepath_Reals_eq) auto also have "\ = integral {r..R} (\x. x powr (-s) * exp (a * x) / (1 - exp x))" using r_R by (intro integral_cong) (simp add: f_def Ln'_of_real_pos exp_diff exp_minus exp_Reals_eq field_simps powr_def Ln_Reals_eq) finally show ?thesis1 by (simp only: add_ac big_circle_def small_circle_def) qed text \ Next, we need bounds on the integrands of the two semicircles. \ lemma hurwitz_formula_bound1: defines "H \ \z. exp (complex_of_real a * z) / (1 - exp z)" assumes "r > 0" obtains C where "C \ 0" and "\z. z \ (\n::int. ball (2 * n * pi * \) r) \ norm (H z) \ C" proof - define A where "A = cbox (-1 - pi * \) (1 + pi * \) - ball 0 r" { fix z assume "z \ A" have "exp z \ 1" proof assume "exp z = 1" then obtain n :: int where [simp]: "z = 2 * n * pi * \" by (subst (asm) exp_eq_1) (auto simp: complex_eq_iff) from \z \ A\ have "(2 * n) * pi \ (-1) * pi" and "(2 * n) * pi \ 1 * pi" by (auto simp: A_def in_cbox_complex_iff) hence "n = 0" by (subst (asm) (1 2) mult_le_cancel_right) auto with \z \ A\ and \r > 0\ show False by (simp add: A_def) qed } hence "continuous_on A H" by (auto simp: A_def H_def intro!: continuous_intros) moreover have "compact A" by (auto simp: A_def compact_eq_bounded_closed) ultimately have "compact (H ` A)" by (rule compact_continuous_image) hence "bounded (H ` A)" by (rule compact_imp_bounded) then obtain C where bound_inside: "\z. z \ A \ norm (H z) \ C" by (auto simp: bounded_iff) have bound_outside: "norm (H z) \ exp 1 / (exp 1 - 1)" if "\Re z\ > 1" for z proof - have "norm (H z) = exp (a * Re z) / norm (1 - exp z)" by (simp add: H_def norm_divide) also have "\1 - exp (Re z)\ \ norm (1 - exp z)" by (rule order.trans[OF _ norm_triangle_ineq3]) simp hence "exp (a * Re z) / norm (1 - exp z) \ exp (a * Re z) / \1 - exp (Re z)\" using that by (intro divide_left_mono mult_pos_pos) auto also have "\ \ exp 1 / (exp 1 - 1)" proof (cases "Re z > 1") case True hence "exp (a * Re z) / \1 - exp (Re z)\ = exp (a * Re z) / (exp (Re z) - 1)" by simp also have "\ \ exp (Re z) / (exp (Re z) - 1)" using a True by (intro divide_right_mono) auto also have "\ = 1 / (1 - exp (-Re z))" by (simp add: exp_minus field_simps) also have "\ \ 1 / (1 - exp (-1))" using True by (intro divide_left_mono diff_mono) auto also have "\ = exp 1 / (exp 1 - 1)" by (simp add: exp_minus field_simps) finally show ?thesis . next case False with that have "Re z < -1" by simp hence "exp (a * Re z) / \1 - exp (Re z)\ = exp (a * Re z) / (1 - exp (Re z))" by simp also have "\ \ 1 / (1 - exp (Re z))" using a and \Re z < -1\ by (intro divide_right_mono) (auto intro: mult_nonneg_nonpos) also have "\ \ 1 / (1 - exp (-1))" using \Re z < -1\ by (intro divide_left_mono) auto also have "\ = exp 1 / (exp 1 - 1)" by (simp add: exp_minus field_simps) finally show ?thesis . qed finally show ?thesis . qed define D where "D = max C (exp 1 / (exp 1 - 1))" have "D \ 0" by (simp add: D_def max.coboundedI2) have "norm (H z) \ D" if "z \ (\n::int. ball (2 * n * pi * \) r)" for z proof (cases "\Re z\ \ 1") case False with bound_outside[of z] show ?thesis by (simp add: D_def) next case True define n where "n = \Im z / (2 * pi) + 1 / 2\" have "Im (z - 2 * n * pi * \) = frac (Im z / (2 * pi) + 1 / 2) * (2 * pi) - pi" by (simp add: n_def frac_def algebra_simps) also have "\ \ {-pi..)) \ C" using True that by (intro bound_inside) (auto simp: A_def in_cbox_complex_iff dist_norm n_def) also have "exp (2 * pi * n * \) = 1" by (simp add: exp_eq_polar) hence "norm (H (z - 2 * n * pi * \)) = norm (H z)" by (simp add: H_def norm_divide exp_diff mult_ac) also have "C \ D" by (simp add: D_def) finally show ?thesis . qed from \D \ 0\ and this show ?thesis by (rule that) qed lemma hurwitz_formula_bound2: obtains C where "C \ 0" and "\r z. r > 0 \ r < pi \ z \ sphere 0 r \ norm (f s z) \ C * r powr (-Re s - 1)" proof - have "2 * pi > 0" by auto have nz: "1 - exp z \ 0" if "z \ ball 0 (2 * pi) - {0}" for z :: complex proof assume "1 - exp z = 0" then obtain n where "z = 2 * pi * of_int n * \" by (auto simp: exp_eq_1 complex_eq_iff[of z]) moreover have "\real_of_int n\ < 1 \ n = 0" by linarith ultimately show False using that by (auto simp: norm_mult) qed have ev: "eventually (\z::complex. 1 - exp z \ 0) (at 0)" using eventually_at_ball'[OF \2 * pi > 0\] by eventually_elim (use nz in auto) have [simp]: "subdegree (1 - fps_exp (1 :: complex)) = 1" by (intro subdegreeI) auto hence "(\z. exp (a * z) * (if z = 0 then -1 else z / (1 - exp z :: complex))) has_fps_expansion fps_exp a * (fps_X / (fps_const 1 - fps_exp 1))" by (auto intro!: fps_expansion_intros) hence "(\z::complex. exp (a * z) * (if z = 0 then -1 else z / (1 - exp z))) \ O[at 0](\z. 1)" using continuous_imp_bigo_1 has_fps_expansion_imp_continuous by blast also have "?this \ (\z::complex. exp (a * z) * (z / (1 - exp z))) \ O[at 0](\z. 1)" by (intro landau_o.big.in_cong eventually_mono[OF ev]) auto finally have "\g. g holomorphic_on ball 0 (2 * pi) \ (\z\ball 0 (2 * pi) - {0}. g z = exp (of_real a * z) * (z / (1 - exp z)))" using nz by (intro holomorphic_on_extend holomorphic_intros) auto then guess g by (elim exE conjE) note g = this hence "continuous_on (ball 0 (2 * pi)) g" by (auto dest: holomorphic_on_imp_continuous_on) hence "continuous_on (cball 0 pi) g" by (rule continuous_on_subset) (subst cball_subset_ball_iff, use pi_gt_zero in auto) hence "compact (g ` cball 0 pi)" by (intro compact_continuous_image) auto hence "bounded (g ` cball 0 pi)" by (auto simp: compact_imp_bounded) then obtain C where C: "\x\cball 0 pi. norm (g x) \ C" by (auto simp: bounded_iff) { fix r :: real assume r: "r > 0" "r < pi" fix z :: complex assume z: "z \ sphere 0 r" define x where "x = (if Arg z \ -pi / 2 then Arg z + 2 * pi else Arg z)" have "exp (\ * (2 * pi)) = 1" by (simp add: exp_eq_polar) with z have "z = r * exp (\ * x)" using r pi_gt_zero Arg_eq[of z] by (auto simp: x_def exp_add distrib_left) have "x > - pi / 2" "x \ 3 / 2 * pi" using Arg_le_pi[of z] mpi_less_Arg[of z] by (auto simp: x_def) note x = \z = r * exp (\ * x)\ this from x r have z': "z \ cball 0 pi - {0}" using pi_gt3 by (auto simp: norm_mult) also have "cball 0 pi \ ball (0::complex) (2 * pi)" by (subst cball_subset_ball_iff) (use pi_gt_zero in auto) hence "cball 0 pi - {0} \ ball 0 (2 * pi) - {0::complex}" by blast finally have z'': "z \ ball 0 (2 * pi) - {0}" . hence bound: "norm (exp (a * z) * (z / (1 - exp z))) \ C" using C and g and z' by force have "exp z \ 1" using nz z'' by auto with bound z'' have bound': "norm (exp (a * z) / (1 - exp z)) \ C / norm z" by (simp add: norm_divide field_simps norm_mult) have "Ln' z = of_real (ln r) + Ln' (exp (\ * of_real x))" using x r by (simp add: Ln'_times_of_real) also have "exp (\ * pi / 2) = \" by (simp add: exp_eq_polar) hence "Ln' (exp (\ * of_real x)) = Ln (exp (\ * of_real (x - pi / 2))) + \ * pi / 2" by (simp add: algebra_simps Ln'_def exp_diff) also have "\ = \ * x" using x pi_gt3 by (subst Ln_exp) (auto simp: algebra_simps) finally have "norm (exp (-Ln' z * s)) = exp (x * Im s - ln r * Re s)" by simp also { have "x * Im s \ \x * Im s\" by (rule abs_ge_self) also have "\ \ (3/2 * pi) * \Im s\" unfolding abs_mult using x by (intro mult_right_mono) auto finally have "exp (x * Im s - ln r * Re s) \ exp (3 / 2 * pi * \Im s\ - ln r * Re s)" by simp } finally have "norm (exp (-Ln' z * s) * (exp (a * z) / (1 - exp z))) \ exp (3 / 2 * pi * \Im s\ - ln r * Re s) * (C / norm z)" unfolding norm_mult[of "exp t" for t] by (intro mult_mono bound') simp_all also have "norm z = r" using \r > 0\ by (simp add: x norm_mult) also have "exp (3 / 2 * pi * \Im s\ - ln r * Re s) = exp (3 / 2 * pi * \Im s\) * r powr (-Re s)" using r by (simp add: exp_diff powr_def exp_minus inverse_eq_divide) finally have "norm (f s z) \ C * exp (3 / 2 * pi * \Im s\) * r powr (-Re s - 1)" using r by (simp add: f_def exp_diff exp_minus field_simps powr_diff) also have "\ \ max 0 (C * exp (3 / 2 * pi * \Im s\)) * r powr (-Re s - 1)" by (intro mult_right_mono max.coboundedI2) auto finally have "norm (f s z) \ \" . } with that[of "max 0 (C * exp (3 / 2 * pi * \Im s\))"] show ?thesis by auto qed text \ We can now relate the integral along a partial Hankel contour that is cut off at $-\pi$ to $\zeta(1 - s, a) / \Gamma(s)$. \ lemma rGamma_hurwitz_zeta_eq_contour_integral: fixes s :: complex and r :: real assumes "s \ 0" and r: "r \ {1..<2}" and a: "a > 0" defines "err1 \ (\s r. contour_integral (part_circlepath 0 r pi 0) (f s))" defines "err2 \ (\s r. cnj (contour_integral (part_circlepath 0 r pi 0) (f (cnj s))))" shows "2 * \ * pi * rGamma s * hurwitz_zeta a (1 - s) = err2 s r - err1 s r + 2 * \ * sin (pi * s) * (CLBINT x:{r..}. g s x)" (is "?f s = ?g s") proof (rule analytic_continuation_open[where f = ?f]) fix s :: complex assume s: "s \ {s. Re s < 0}" \\We first show that the integrals along the Hankel contour cut off at $-\pi$ all have the same value, no matter what the radius of the circle is (as long as it is small enough). We call this value \C\. This argument could be done by a homotopy argument, but it is easier to simply re-use the above result about the contour integral along the annulus where we fix the radius of the outer circle to $\pi$.\ define C where "C = -contour_integral (part_circlepath 0 pi 0 pi) (f s) + cnj (contour_integral (part_circlepath 0 pi 0 pi) (f (cnj s)))" have integrable: "set_integrable lborel A (g s)" if "A \ sets lborel" "A \ {0<..}" for A proof (rule set_integrable_subset) show "set_integrable lborel {0<..} (g s)" using Gamma_times_hurwitz_zeta_integrable[of "1 - s" a] s a by (simp add: g_def exp_of_real exp_minus integrable_completion set_integrable_def) qed (insert that, auto) { fix r' :: real assume r': "r' \ {0<..<2}" from hurwitz_formula_integral_semiannulus(2)[of r' s 0] and r' have "f s contour_integrable_on part_circlepath 0 r' pi 0" by (auto simp: hankel_semiannulus_def add_ac) } note integrable_circle = this { fix r' :: real assume r': "r' \ {0<..<2}" from hurwitz_formula_integral_semiannulus(2)[of r' "cnj s" 0] and r' have "f (cnj s) contour_integrable_on part_circlepath 0 r' pi 0" by (auto simp: hankel_semiannulus_def add_ac) } note integrable_circle' = this have eq: "-2 * \ * sin (pi * s) * (CLBINT x:{r..pi}. g s x) + (err1 s r - err2 s r) = C" if r: "r \ {0<..<2}" for r :: real proof - have eq1: "integral {r..pi} (\x. cnj (x powr - cnj s) * (exp (- (a * x))) / (1 - (exp (- x)))) = integral {r..pi} (g s)" using r by (intro integral_cong) (auto simp: cnj_powr g_def exp_of_real exp_minus) have eq2: "integral {r..pi} (\x. cnj (x powr - cnj s) * (exp (a * x)) / (1 - (exp x))) = integral {r..pi} (\x. x powr - s * (exp (a * x)) / (1 - (exp x)))" using r by (intro integral_cong) (auto simp: cnj_powr) from hurwitz_formula_integral_semiannulus(1)[of r s 0] hurwitz_formula_integral_semiannulus(1)[of r "cnj s" 0] have "exp (-\*pi * s) * integral {r..real (2*0+1) * pi} (g s) + integral {r..real (2*0+1) * pi} (\x. x powr -s * exp (a * x) / (1 - exp x)) + contour_integral (part_circlepath 0 (real (2 * 0 + 1) * pi) 0 pi) (f s) + contour_integral (part_circlepath 0 r pi 0) (f s) - cnj ( exp (-\*pi * cnj s) * integral {r..real (2*0+1) * pi} (\x. x powr - cnj s * exp (-a*x) / (1 - exp (-x))) + integral {r..real (2*0+1) * pi} (\x. x powr - cnj s * exp (a*x) / (1 - exp x)) + contour_integral (part_circlepath 0 (real (2 * 0 + 1) * pi) 0 pi) (f (cnj s)) + contour_integral (part_circlepath 0 r pi 0) (f (cnj s))) = 0" (is "?lhs = _") unfolding g_def using r by (subst (1 2) hurwitz_formula_integral_semiannulus) auto also have "?lhs = -2 * \ * sin (pi * s) * integral {r..pi} (g s) + err1 s r - err2 s r - C" using eq1 eq2 by (auto simp: integral_cnj exp_cnj err1_def err2_def sin_exp_eq algebra_simps C_def) also have "integral {r..pi} (g s) = (CLBINT x:{r..pi}. g s x)" using r by (intro set_borel_integral_eq_integral(2) [symmetric] integrable) auto finally show "-2 * \ * sin (pi * s) * (CLBINT x:{r..pi}. g s x) + (err1 s r - err2 s r) = C" by (simp add: algebra_simps) qed \\Next, compute the value of @{term C} by letting the radius tend to 0 so that the contribution of the circle vanishes.\ have "((\r. -2 * \ * sin (pi * s) * (CLBINT x:{r..pi}. g s x) + (err1 s r - err2 s r)) \ -2 * \ * sin (pi * s) * (CLBINT x:{0<..pi}. g s x) + 0) (at_right 0)" proof (intro tendsto_intros tendsto_set_lebesgue_integral_at_right integrable) from hurwitz_formula_bound2[of s] guess C1 . note C1 = this from hurwitz_formula_bound2[of "cnj s"] guess C2 . note C2 = this have ev: "eventually (\r::real. r \ {0<..<2}) (at_right 0)" by (intro eventually_at_right_real) auto show "((\r. err1 s r - err2 s r) \ 0) (at_right 0)" proof (rule Lim_null_comparison[OF eventually_mono[OF ev]]) fix r :: real assume r: "r \ {0<..<2}" have "norm (err1 s r - err2 s r) \ norm (err1 s r) + norm (err2 s r)" by (rule norm_triangle_ineq4) also have "norm (err1 s r) \ C1 * r powr (- Re s - 1) * r * \0 - pi\" unfolding err1_def using C1(1) C1(2)[of r] pi_gt3 integrable_circle[of r] path_image_part_circlepath_subset'[of r 0 pi 0] r by (intro contour_integral_bound_part_circlepath) auto also have "\ = C1 * r powr (-Re s) * pi" using r by (simp add: powr_diff field_simps) also have "norm (err2 s r) \ C2 * r powr (- Re s - 1) * r * \0 - pi\" unfolding err2_def complex_mod_cnj using C2(1) C2(2)[of r] r pi_gt3 integrable_circle'[of r] path_image_part_circlepath_subset'[of r 0 pi 0] by (intro contour_integral_bound_part_circlepath) auto also have "\ = C2 * r powr (-Re s) * pi" using r by (simp add: powr_diff field_simps) also have "C1 * r powr (-Re s) * pi + C2 * r powr (-Re s) * pi = (C1 + C2) * pi * r powr (-Re s)" by (simp add: algebra_simps) finally show "norm (err1 s r - err2 s r) \ (C1 + C2) * pi * r powr - Re s" by simp next show "((\x. (C1 + C2) * pi * x powr - Re s) \ 0) (at_right 0)" using s by (auto intro!: tendsto_eq_intros simp: eventually_at exI[of _ 1]) qed qed auto moreover have "eventually (\r::real. r \ {0<..<2}) (at_right 0)" by (intro eventually_at_right_real) auto hence "eventually (\r. -2 * \ * sin (pi * s) * (CLBINT x:{r..pi}. g s x) + (err1 s r - err2 s r) = C) (at_right 0)" by eventually_elim (use eq in auto) hence "((\r. -2 * \ * sin (pi * s) * (CLBINT x:{r..pi}. g s x) + (err1 s r - err2 s r)) \ C) (at_right 0)" by (rule tendsto_eventually) ultimately have [simp]: "C = -2 * \ * sin (pi * s) * (CLBINT x:{0<..pi}. g s x)" using tendsto_unique by force \\We now rearrange everything and obtain the result.\ have "2 * \ * sin (pi * s) * ((CLBINT x:{0<..pi}. g s x) - (CLBINT x:{r..pi}. g s x)) = err2 s r - err1 s r" using eq[of r] r by (simp add: algebra_simps) also have "{0<..pi} = {0<.. {r..pi}" using r pi_gt3 by auto also have "(CLBINT x:\. g s x) - (CLBINT x:{r..pi}. g s x) = (CLBINT x:{0<.. {r..}. g s x) - (CLBINT x:{r..}. g s x)" using r pi_gt3 by (subst set_integral_Un[OF _ integrable integrable]) auto also have "{0<.. {r..} = {0<..}" using r by auto also have "(CLBINT x:{0<..}. g s x) = Gamma (1 - s) * hurwitz_zeta a (1 - s)" using Gamma_times_hurwitz_zeta_integral[of "1 - s" a] s a by (simp add: g_def exp_of_real exp_minus integral_completion set_lebesgue_integral_def) finally have "2 * \ * (sin (pi * s) * Gamma (1 - s)) * hurwitz_zeta a (1 - s) = err2 s r - err1 s r + 2 * \ * sin (pi * s) * (CLBINT x:{r..}. g s x)" by (simp add: algebra_simps) also have "sin (pi * s) * Gamma (1 - s) = pi * rGamma s" proof (cases "s \ \") case False with Gamma_reflection_complex[of s] show ?thesis by (auto simp: divide_simps sin_eq_0 Ints_def rGamma_inverse_Gamma mult_ac split: if_splits) next case True with s have "rGamma s = 0" by (auto simp: rGamma_eq_zero_iff nonpos_Ints_def Ints_def) moreover from True have "sin (pi * s) = 0" by (subst sin_eq_0) (auto elim!: Ints_cases) ultimately show ?thesis by simp qed finally show "2 * \ * pi * rGamma s * hurwitz_zeta a (1 - s) = err2 s r - err1 s r + 2 * \ * sin (pi * s) * (CLBINT x:{r..}. g s x)" by (simp add: mult_ac) next \\By analytic continuation, we lift the result to the case of any non-zero @{term s}.\ show "(\s. 2 * \ * pi * rGamma s * hurwitz_zeta a (1 - s)) holomorphic_on - {0}" using a by (auto intro!: holomorphic_intros) show "(\s. err2 s r - err1 s r + 2 * \ * sin (pi * s) * (CLBINT x:{r..}. g s x)) holomorphic_on -{0}" proof (intro holomorphic_intros) have "(\s. err2 s r) = (\s. - cnj (integral {0..pi} (h r (cnj s))))" using r by (simp add: err2_def contour_integral_part_circlepath_reverse' contour_integral_part_circlepath_h) also have "(\s. - cnj (integral {0..pi} (h r (cnj s)))) = (\s. (integral {0..pi} (\x. h r s (-x))))" using r by (simp add: integral_cnj h_def exp_cnj cis_cnj Ln_Reals_eq) also have "\ = (\s. integral {-pi..0} (h r s))" by (subst Henstock_Kurzweil_Integration.integral_reflect_real [symmetric]) simp finally have "(\s. err2 s r) = \" . moreover have "(\s. integral {-pi..0} (h r s)) holomorphic_on -{0}" using r by (intro integral_h_holomorphic) auto ultimately show "(\s. err2 s r) holomorphic_on -{0}" by simp next have "(\s. - integral {0..pi} (h r s)) holomorphic_on -{0}" using r by (intro holomorphic_intros integral_h_holomorphic) auto also have "(\s. - integral {0..pi} (h r s)) = (\s. err1 s r)" unfolding err1_def using r by (simp add: contour_integral_part_circlepath_reverse' contour_integral_part_circlepath_h) finally show "(\s. err1 s r) holomorphic_on -{0}" . next show "(\s. CLBINT x:{r..}. g s x) holomorphic_on -{0}" proof (rule holomorphic_on_balls_imp_entire') fix R :: real have "eventually (\b. b > r) at_top" by (rule eventually_gt_at_top) hence 1: "eventually (\b. continuous_on (cball 0 R) (\s. CLBINT x:{r..b}. g s x) \ (\s. CLBINT x:{r..b}. g s x) holomorphic_on ball 0 R) at_top" proof eventually_elim case (elim b) have integrable: "set_integrable lborel {r..b} (g s)" for s unfolding g_def using r by (intro borel_integrable_atLeastAtMost' continuous_intros) auto have "(\s. integral {r..b} (g s)) holomorphic_on UNIV" using r by (intro integral_g_holomorphic) auto also have "(\s. integral {r..b} (g s)) = (\s. CLBINT x:{r..b}. g s x)" by (intro ext set_borel_integral_eq_integral(2)[symmetric] integrable) finally have "\ holomorphic_on UNIV" . thus ?case by (auto intro!: holomorphic_on_imp_continuous_on) qed have 2: "uniform_limit (cball 0 R) (\b s. CLBINT x:{r..b}. g s x) (\s. CLBINT x:{r..}. g s x) at_top" proof (rule uniform_limit_set_lebesgue_integral_at_top) fix s :: complex and x :: real assume s: "s \ cball 0 R" and x: "x \ r" have "norm (g s x) = x powr -Re s * exp (-a * x) / (1 - exp (-x))" using x r by (simp add: g_def norm_mult norm_divide in_Reals_norm norm_powr_real_powr) also have "\ \ x powr R * exp (-a * x) / (1 - exp (-x))" using r s x abs_Re_le_cmod[of s] by (intro mult_right_mono divide_right_mono powr_mono) auto finally show "norm (g s x) \ x powr R * exp (- a * x) / (1 - exp (- x))" . next show "set_integrable lborel {r..} (\x. x powr R * exp (-a * x) / (1 - exp (-x)))" using r a by (intro set_integrable_Gamma_hurwitz_aux2_real) auto qed (simp_all add: set_borel_measurable_def g_def) show "(\s. CLBINT x:{r..}. g s x) holomorphic_on ball 0 R" using holomorphic_uniform_limit[OF 1 2] by auto qed qed qed (insert \s \ 0\, auto simp: connected_punctured_universe open_halfspace_Re_lt intro: exI[of _ "-1"]) text \ Finally, we obtain Hurwitz's formula by letting the radius of the outer circle tend to \\\. \ lemma hurwitz_zeta_formula_aux: fixes s :: complex assumes s: "Re s > 1" shows "rGamma s * hurwitz_zeta a (1 - s) = (2 * pi) powr -s * (\ powr (-s) * F a s + \ powr s * F (-a) s)" proof - from s have [simp]: "s \ 0" by auto define r where "r = (1 :: real)" have r: "r \ {0<..<2}" by (simp add: r_def) define R where "R = (\n. real (2 * n + 1) * pi)" define bigc where "bigc = (\n. contour_integral (part_circlepath 0 (R n) 0 pi) (f s) - cnj (contour_integral (part_circlepath 0 (R n) 0 pi) (f (cnj s))))" define smallc where "smallc = contour_integral (part_circlepath 0 r pi 0) (f s) - cnj (contour_integral (part_circlepath 0 r pi 0) (f (cnj s)))" define I where "I = (\n. CLBINT x:{r..R n}. g s x)" define F1 and F2 where "F1 = (\n. exp (-s * pi * \ / 2) * (\k\{0<..n}. exp (2 * real k * pi * a * \) * k powr (-s)))" "F2 = (\n. exp (s * pi * \ / 2) * (\k\{0<..n}. exp (2 * real k * pi * (-a) * \) * k powr (-s)))" have R: "R n \ pi" for n using r by (auto simp: R_def field_simps) have [simp]: "\(pi \ 0)" using pi_gt_zero by linarith have integrable: "set_integrable lborel A (g s)" if "A \ sets lborel" "A \ {r..}" for A proof - have "set_integrable lborel {r..} (g s)" using set_integrable_Gamma_hurwitz_aux2[of r a s] a r by (simp add: g_def exp_of_real exp_minus) thus ?thesis by (rule set_integrable_subset) (use that in auto) qed { fix n :: nat from hurwitz_formula_integral_semiannulus(2)[of "r" s n] and r R[of n] have "f s contour_integrable_on part_circlepath 0 (R n) 0 pi" by (auto simp: hankel_semiannulus_def R_def add_ac) } note integrable_circle = this { fix n :: nat from hurwitz_formula_integral_semiannulus(2)[of "r" "cnj s" n] and r R[of n] have "f (cnj s) contour_integrable_on part_circlepath 0 (R n) 0 pi" by (auto simp: hankel_semiannulus_def R_def add_ac) } note integrable_circle' = this { fix n :: nat have "(exp (-\ * pi * s) * integral {r..R n} (g s) + integral {r..R n} (\x. x powr (-s) * exp (a * x) / (1 - exp x)) + contour_integral (part_circlepath 0 (R n) 0 pi) (f s) + contour_integral (part_circlepath 0 r pi 0) (f s)) - cnj ( exp (-\ * pi * cnj s) * integral {r..R n} (g (cnj s)) + integral {r..R n} (\x. x powr (-cnj s) * exp (a * x) / (1 - exp x)) + contour_integral (part_circlepath 0 (R n) 0 pi) (f (cnj s)) + contour_integral (part_circlepath 0 r pi 0) (f (cnj s))) = -2 * pi * \ * exp (- s * of_real pi * \ / 2) * (\k\{0<..n}. Res s k) - cnj (-2 * pi * \ * exp (- cnj s * of_real pi * \ / 2) * (\k\{0<..n}. Res (cnj s) k))" (is "?lhs = ?rhs") unfolding R_def g_def using r by (subst (1 2) hurwitz_formula_integral_semiannulus) auto also have "?rhs = -2 * pi * \ * (exp (- s * pi * \ / 2) * (\k\{0<..n}. Res s k) + exp (s * pi * \ / 2) * (\k\{0<..n}. cnj (Res (cnj s) k)))" by (simp add: exp_cnj sum.distrib algebra_simps sum_distrib_left sum_distrib_right sum_negf) also have "(\k\{0<..n}. Res s k) = (2 * pi) powr (-s) * (\k\{0<..n}. exp (2 * k * pi * a * \) * k powr (-s))" (is "_ = ?S1") by (simp add: Res_def powr_times_real algebra_simps sum_distrib_left) also have "(\k\{0<..n}. cnj (Res (cnj s) k)) = (2 * pi) powr (-s) * (\k\{0<..n}. exp (-2 * k * pi * a * \) * k powr (-s))" by (simp add: Res_def cnj_powr powr_times_real algebra_simps exp_cnj sum_distrib_left) also have "exp (-s * pi * \ / 2) * ?S1 + exp (s * pi * \ / 2) * \ = (2 * pi) powr (-s) * (exp (-s * pi * \ / 2) * (\k\{0<..n}. exp (2 * k * pi * a * \) * k powr (-s)) + exp (s * pi * \ / 2) * (\k\{0<..n}. exp (-2 * k * pi * a * \) * k powr (-s)))" by (simp add: algebra_simps) also have 1: "integral {r..R n} (g s) = I n" unfolding I_def by (intro set_borel_integral_eq_integral(2) [symmetric] integrable) auto have 2: "cnj (integral {r..R n} (g (cnj s))) = integral {r..R n} (g s)" using r unfolding integral_cnj by (intro integral_cong) (auto simp: g_def cnj_powr) have 3: "integral {r..R n} (\x. exp (x * a) * cnj (x powr - cnj s) / (1 - exp x)) = integral {r..R n} (\x. exp (x * a) * of_real x powr - s / (1 - exp x))" unfolding I_def g_def using r R[of n] by (intro integral_cong; force simp: cnj_powr)+ from 1 2 3 have "?lhs = (exp (-\ * s * pi) - exp (\ * s * pi)) * I n + bigc n + smallc" by (simp add: integral_cnj cnj_powr algebra_simps exp_cnj bigc_def smallc_def g_def) also have "exp (-\ * s * pi) - exp (\ * s * pi) = -2 * \ * sin (s * pi)" by (simp add: sin_exp_eq' algebra_simps) finally have "(- 2 * \ * sin (s * pi) * I n + smallc) + bigc n = -2 * \ * pi * (2 * pi) powr (-s) * (F1 n + F2 n)" by (simp add: F1_F2_def algebra_simps) } note eq = this have "(\n. - 2 * \ * sin (s * pi) * I n + smallc + bigc n) \ (-2 * \ * sin (s * pi)) * (CLBINT x:{r..}. g s x) + smallc + 0" unfolding I_def proof (intro tendsto_intros filterlim_compose[OF tendsto_set_lebesgue_integral_at_top] integrable) show "filterlim R at_top sequentially" unfolding R_def by (intro filterlim_at_top_mult_tendsto_pos[OF tendsto_const] pi_gt_zero filterlim_compose[OF filterlim_real_sequentially] filterlim_subseq) (auto simp: strict_mono_Suc_iff) from hurwitz_formula_bound1[OF pi_gt_zero] guess C . note C = this define D where "D = C * exp (3 / 2 * pi * \Im s\)" from \C \ 0\ have "D \ 0" by (simp add: D_def) show "bigc \ 0" proof (rule Lim_null_comparison[OF always_eventually[OF allI]]) fix n :: nat have bound: "norm (f s' z) \ D * R n powr (-Re s')" if z: "z \ sphere 0 (R n)" "Re s' = Re s" "\Im s'\ = \Im s\" for z s' proof - from z and r R[of n] have [simp]: "z \ 0" by auto have not_in_ball: "z \ ball (2 * m * pi * \) pi" for m :: int proof - have "dist z (2 * m * pi * \) \ \dist z 0 - dist 0 (2 * m * pi * \)\" by (rule abs_dist_diff_le) also have "dist 0 (2 * m * pi * \) = 2 * \m\ * pi" by (simp add: norm_mult) also from z have "dist z 0 = R n" by simp also have "R n - 2 * \m\ * pi = (int (2 * n + 1) - 2 * \m\) * pi" by (simp add: R_def algebra_simps) also have "\\\ = \int (2 * n + 1) - 2 * \m\\ * pi" by (subst abs_mult) simp_all also have "\int (2 * n + 1) - 2 * \m\\ \ 1" by presburger hence "\ * pi \ 1 * pi" by (intro mult_right_mono) auto finally show ?thesis by (simp add: dist_commute) qed have "norm (f s' z) = norm (exp (-Ln' z * s')) * norm (exp (a * z) / (1 - exp z))" by (simp add: f_def exp_diff norm_mult norm_divide mult_ac exp_minus norm_inverse divide_simps del: norm_exp_eq_Re) also have "\ \ norm (exp (-Ln' z * s')) * C" using not_in_ball by (intro mult_left_mono C) auto also have "norm (exp (-Ln' z * s')) = exp (Im s' * (Im (Ln (- (\ * z))) + pi / 2)) / exp (Re s' * ln (R n))" using z r R[of n] pi_gt_zero by (simp add: Ln'_def norm_mult norm_divide exp_add exp_diff exp_minus norm_inverse algebra_simps inverse_eq_divide) also have "\ \ exp (3/2 * pi * \Im s'\) / exp (Re s' * ln (R n))" proof (intro divide_right_mono, subst exp_le_cancel_iff) have "Im s' * (Im (Ln (- (\ * z))) + pi / 2) \ \Im s' * (Im (Ln (- (\ * z))) + pi / 2)\" by (rule abs_ge_self) also have "\ \ \Im s'\ * (pi + pi / 2)" unfolding abs_mult using mpi_less_Im_Ln[of "- (\ * z)"] Im_Ln_le_pi[of "- (\ * z)"] by (intro mult_left_mono order.trans[OF abs_triangle_ineq] add_mono) auto finally show "Im s' * (Im (Ln (- (\ * z))) + pi / 2) \ 3/2 * pi * \Im s'\" by (simp add: algebra_simps) qed auto also have "exp (Re s' * ln (R n)) = R n powr Re s'" using r R[of n] by (auto simp: powr_def) finally show "norm (f s' z) \ D * R n powr (-Re s')" using \C \ 0\ by (simp add: that D_def powr_minus mult_right_mono mult_left_mono field_simps) qed have "norm (bigc n) \ norm (contour_integral (part_circlepath 0 (R n) 0 pi) (f s)) + norm (cnj (contour_integral (part_circlepath 0 (R n) 0 pi) (f (cnj s))))" (is "_ \ norm ?err1 + norm ?err2") unfolding bigc_def by (rule norm_triangle_ineq4) also have "norm ?err1 \ D * R n powr (-Re s) * R n * \pi - 0\" using \D \ 0\ and r R[of n] and pi_gt3 and integrable_circle and path_image_part_circlepath_subset[of 0 pi "R n" 0] and bound[of _ s] by (intro contour_integral_bound_part_circlepath) auto also have "\ = D * pi * R n powr (1 - Re s)" using r R[of n] pi_gt3 by (simp add: powr_diff field_simps powr_minus) also have "norm ?err2 \ D * R n powr (-Re s) * R n * \pi - 0\" unfolding complex_mod_cnj using \D \ 0\ and r R[of n] and pi_gt3 and integrable_circle'[of n] and path_image_part_circlepath_subset[of 0 pi "R n" 0] and bound[of _ "cnj s"] by (intro contour_integral_bound_part_circlepath) auto also have "\ = D * pi * R n powr (1 - Re s)" using r R[of n] pi_gt3 by (simp add: powr_diff field_simps powr_minus) finally show "norm (bigc n) \ 2 * D * pi * R n powr (1 - Re s)" by simp next have "filterlim R at_top at_top" by fact hence "(\x. 2 * D * pi * R x powr (1 - Re s)) \ 2 * D * pi * 0" using s unfolding R_def by (intro tendsto_intros tendsto_neg_powr) auto thus "(\x. 2 * D * pi * R x powr (1 - Re s)) \ 0" by simp qed qed auto also have "(\n. - 2 * \ * sin (s * pi) * I n + smallc + bigc n) = (\n. -2 * \ * pi * (2 * pi) powr -s * (F1 n + F2 n))" by (subst eq) auto finally have "\ \ (-2 * \ * sin (s * pi)) * (CLBINT x:{r..}. g s x) + smallc" by simp moreover have "(\n. -2 * \ * pi * (2 * pi) powr -s * (F1 n + F2 n)) \ -2 * \ * pi * (2 * pi) powr -s * (exp (-s * pi * \ / 2) * F a s + exp (s * pi * \ / 2) * F (-a) s)" unfolding F1_F2_def F_def using s by (intro tendsto_intros sum_tendsto_fds_perzeta) ultimately have "-2 * \ * pi * (2 * pi) powr -s * (exp (-s * pi * \ / 2) * F a s + exp (s * pi * \ / 2) * F (-a) s) = (-2 * \ * sin (s * pi)) * (CLBINT x:{r..}. g s x) + smallc" by (force intro: tendsto_unique) also have "\ = -2 * \ * pi * rGamma s * hurwitz_zeta a (1 - s)" using s r a using rGamma_hurwitz_zeta_eq_contour_integral[of s r] by (simp add: r_def smallc_def algebra_simps) also have "exp (- s * complex_of_real pi * \ / 2) = \ powr (-s)" by (simp add: powr_def field_simps) also have "exp (s * complex_of_real pi * \ / 2) = \ powr s" by (simp add: powr_def field_simps) finally show "rGamma s * hurwitz_zeta a (1 - s) = (2 * pi) powr -s * (\ powr (-s) * F a s + \ powr s * F (-a) s)" by simp qed end text \ We can now use Hurwitz's formula to prove the following nice formula that expresses the periodic zeta function in terms of the Hurwitz zeta function: \[F(s, a) = (2\pi)^{s-1} i \Gamma(1 - s) \left(i^{-s} \zeta(1 - s, a) - i^{s} \zeta(1 - s, 1 - a)\right)\] This holds for all \s\ with \\mathfrak{R}(s) > 0\ as long as \a \ \\. For convenience, we move the \\\ function to the left-hand side in order to avoid having to account for its poles. \ lemma perzeta_conv_hurwitz_zeta_aux: fixes a :: real and s :: complex assumes a: "a \ {0<..<1}" and s: "Re s > 0" shows "rGamma (1 - s) * eval_fds (fds_perzeta a) s = (2 * pi) powr (s - 1) * \ * (\ powr -s * hurwitz_zeta a (1 - s) - \ powr s * hurwitz_zeta (1 - a) (1 - s))" (is "?lhs s = ?rhs s") proof (rule analytic_continuation_open[where f = ?lhs]) show "connected {s. Re s > 0}" by (intro convex_connected convex_halfspace_Re_gt) show "{s. Re s > 1} \ {}" by (auto intro: exI[of _ 2]) show "(\s. rGamma (1 - s) * eval_fds (fds_perzeta a) s) holomorphic_on {s. 0 < Re s}" unfolding perzeta_def using a by (auto intro!: holomorphic_intros le_less_trans[OF conv_abscissa_perzeta'] elim!: Ints_cases) show "?rhs holomorphic_on {s. 0 < Re s}" using assms by (auto intro!: holomorphic_intros) next fix s assume s: "s \ {s. Re s > 1}" have [simp]: "fds_perzeta (1 - a) = fds_perzeta (-a)" using fds_perzeta.plus_of_nat[of "-a" 1] by simp have [simp]: "fds_perzeta (a - 1) = fds_perzeta a" using fds_perzeta.minus_of_nat[of a 1] by simp from s have [simp]: "Gamma s \ 0" by (auto simp: Gamma_eq_zero_iff elim!: nonpos_Ints_cases) have "(2 * pi) powr (-s) * (\ * (\ powr (-s) * (rGamma s * hurwitz_zeta a (1 - s)) - \ powr s * (rGamma s * hurwitz_zeta (1 - a) (1 - s)))) = (2 * pi) powr (-s) * ((\ powr (1 - s) * (rGamma s * hurwitz_zeta a (1 - s)) + \ powr (s - 1) * (rGamma s * hurwitz_zeta (1 - a) (1 - s))))" by (simp add: powr_diff field_simps powr_minus) also have "\ = ((2 * pi) powr (-s)) ^ 2 * ( eval_fds (fds_perzeta a) s * (\ powr s * \ powr (s - 1) + \ powr (-s) * \ powr (1 - s)) + eval_fds (fds_perzeta (-a)) s * (\ powr s * \ powr (1 - s) + \ powr (-s) * \ powr (s - 1)))" using s a by (subst (1 2) hurwitz_zeta_formula_aux) (auto simp: algebra_simps power2_eq_square) also have "(\ powr s * \ powr (1 - s) + \ powr (-s) * \ powr (s - 1)) = exp (\ * complex_of_real pi / 2) + exp (- (\ * complex_of_real pi / 2))" by (simp add: powr_def exp_add [symmetric] field_simps) also have "\ = 0" by (simp add: exp_eq_polar) also have "\ powr s * \ powr (s - 1) = \ powr (2 * s - 1)" by (simp add: powr_def exp_add [symmetric] field_simps) also have "\ powr (-s) * \ powr (1 - s) = \ powr (1 - 2 * s)" by (simp add: powr_def exp_add [symmetric] field_simps) also have "\ powr (2 * s - 1) + \ powr (1 - 2 * s) = 2 * cos ((2 * s - 1) * pi / 2)" by (simp add: powr_def cos_exp_eq algebra_simps minus_divide_left cos_sin_eq) also have "\ = 2 * sin (pi - s * pi)" by (simp add: cos_sin_eq field_simps) also have "\ = 2 * sin (s * pi)" by (simp add: sin_diff) finally have "\ * (rGamma s * \ powr (-s) * hurwitz_zeta a (1 - s) - rGamma s * \ powr s * hurwitz_zeta (1 - a) (1 - s)) = 2 * (2 * pi) powr -s * sin (s * pi) * eval_fds (fds_perzeta a) s" by (simp add: power2_eq_square mult_ac) hence "(2 * pi) powr s / 2 * \ * (\ powr (-s) * hurwitz_zeta a (1 - s) - \ powr s * hurwitz_zeta (1 - a) (1 - s)) = Gamma s * sin (s * pi) * eval_fds (fds_perzeta a) s" by (subst (asm) (2) powr_minus) (simp add: field_simps rGamma_inverse_Gamma) also have "Gamma s * sin (s * pi) = pi * rGamma (1 - s)" using Gamma_reflection_complex[of s] by (auto simp: divide_simps rGamma_inverse_Gamma mult_ac split: if_splits) finally show "?lhs s = ?rhs s" by (simp add: powr_diff) qed (insert s, auto simp: open_halfspace_Re_gt) text \ We can now use the above equation as a defining equation to continue the periodic zeta function $F$ to the entire complex plane except at non-negative integer values for \s\. However, the positive integers are already covered by the original Dirichlet series definition of $F$, so we only need to take care of \s = 0\. We do this by cancelling the pole of \\\ at \0\ with the zero of $i^{-s} \zeta(1 - s, a) - i^s \zeta(1 - s, 1 - a)$. \ lemma assumes "q' \ \" shows holomorphic_perzeta': "perzeta q' holomorphic_on A" and perzeta_altdef2: "Re s > 0 \ perzeta q' s = eval_fds (fds_perzeta q') s" proof - define q where "q = frac q'" from assms have q: "q \ {0<..<1}" by (auto simp: q_def frac_lt_1) hence [simp]: "q \ \" by (auto elim!: Ints_cases) have [simp]: "frac q = q" by (simp add: q_def frac_def) define f where "f = (\s. complex_of_real (2 * pi) powr (s - 1) * \ * Gamma (1 - s) * (\ powr (-s) * hurwitz_zeta q (1 - s) - \ powr s * hurwitz_zeta (1 - q) (1 - s)))" { fix s :: complex assume "1 - s \ \\<^sub>\\<^sub>0" then obtain n where "1 - s = of_int n" "n \ 0" by (auto elim!: nonpos_Ints_cases) hence "s = 1 - of_int n" by (simp add: algebra_simps) also have "\ \ \" using \n \ 0\ by (auto simp: Nats_altdef1 intro: exI[of _ "1 - n"]) finally have "s \ \" . } note * = this hence "f holomorphic_on -\" using q by (auto simp: f_def Nats_altdef2 nonpos_Ints_altdef not_le intro!: holomorphic_intros) also have "?this \ perzeta q holomorphic_on -\" using assms by (intro holomorphic_cong refl) (auto simp: perzeta_def Let_def f_def) finally have holo: "perzeta q holomorphic_on -\" . have f_altdef: "f s = eval_fds (fds_perzeta q) s" if "Re s > 0" and "s \ \" for s using perzeta_conv_hurwitz_zeta_aux[OF q, of s] that * by (auto simp: rGamma_inverse_Gamma Gamma_eq_zero_iff divide_simps f_def perzeta_def split: if_splits) show "perzeta q' s = eval_fds (fds_perzeta q') s" if "Re s > 0" for s using f_altdef[of s] that assms by (auto simp: f_def perzeta_def Let_def q_def) have cont: "isCont (perzeta q) s" if "s \ \" for s proof (cases "s = 0") case False with that obtain n where [simp]: "s = of_nat n" and n: "n > 0" by (auto elim!: Nats_cases) have *: "open ({s. Re s > 0} - (\ - {of_nat n}))" using Nats_subset_Ints by (intro open_Diff closed_subset_Ints open_halfspace_Re_gt) auto have "eventually (\s. s \ {s. Re s > 0} - (\ - {of_nat n})) (nhds (of_nat n))" using \n > 0\ by (intro eventually_nhds_in_open *) auto hence ev: "eventually (\s. eval_fds (fds_perzeta q) s = perzeta q s) (nhds (of_nat n))" proof eventually_elim case (elim s) thus ?case using q f_altdef[of s] by (auto simp: perzeta_def dist_of_nat f_def elim!: Nats_cases Ints_cases) qed have "isCont (eval_fds (fds_perzeta q)) (of_nat n)" using q and \n > 0\ by (intro continuous_eval_fds le_less_trans[OF conv_abscissa_perzeta']) (auto elim!: Ints_cases) also have "?this \ isCont (perzeta q) (of_nat n)" using ev by (intro isCont_cong ev) finally show ?thesis by simp next assume [simp]: "s = 0" define a where "a = Complex (ln q) (-pi / 2)" define b where "b = Complex (ln (1 - q)) (pi / 2)" have "eventually (\s::complex. s \ \) (at 0)" unfolding eventually_at_topological using Nats_subset_Ints by (intro exI[of _ "-(\-{0})"] conjI open_Compl closed_subset_Ints) auto hence ev: "eventually (\s. perzeta q s = (2 * pi) powr (s - 1) * Gamma (1 - s) * \ * (\ powr - s * pre_zeta q (1 - s) -\ powr s * pre_zeta (1 - q) (1 - s) + (exp (b * s) - exp (a * s)) / s)) (at (0::complex))" (is "eventually (\s. _ = ?f s) _") proof eventually_elim case (elim s) have "perzeta q s = (2 * pi) powr (s - 1) * Gamma (1 - s) * \ * (\ powr (-s) * hurwitz_zeta q (1 - s) - \ powr s * hurwitz_zeta (1 - q) (1 - s))" (is "_ = _ * ?T") using elim by (auto simp: perzeta_def powr_diff powr_minus field_simps) also have "?T = \ powr (-s) * pre_zeta q (1 - s) - \ powr s * pre_zeta (1 - q) (1 - s) + (\ powr s * (1 - q) powr s - \ powr (-s) * q powr s) / s" using elim by (auto simp: hurwitz_zeta_def field_simps) also have "\ powr s * (1 - q) powr s = exp (b * s)" using q by (simp add: powr_def exp_add algebra_simps Ln_Reals_eq Complex_eq b_def) also have "\ powr (-s) * q powr s = exp (a * s)" using q by (simp add: powr_def exp_add Ln_Reals_eq exp_diff exp_minus diff_divide_distrib ring_distribs inverse_eq_divide mult_ac Complex_eq a_def) finally show ?case . qed have [simp]: "\(pi \ 0)" using pi_gt_zero by (simp add: not_le) have "(\s::complex. if s = 0 then b - a else (exp (b * s) - exp (a * s)) / s) has_fps_expansion (fps_exp b - fps_exp a) / fps_X" (is "?f' has_fps_expansion _") by (rule fps_expansion_intros)+ (auto intro!: subdegree_geI simp: Ln_Reals_eq a_def b_def) hence "isCont ?f' 0" by (rule has_fps_expansion_imp_continuous) hence "?f' \0\ b - a" by (simp add: isCont_def) also have "?this \ (\s. (exp (b * s) - exp (a * s)) / s) \0\ b - a" by (intro filterlim_cong refl) (auto simp: eventually_at intro: exI[of _ 1]) finally have "?f \0\ of_real (2 * pi) powr (0 - 1) * Gamma (1 - 0) * \ * (\ powr -0 * pre_zeta q (1 - 0) -\ powr 0 * pre_zeta (1 - q) (1 - 0) + (b - a))" (is "filterlim _ (nhds ?c) _") using q by (intro tendsto_intros isContD) (auto simp: complex_nonpos_Reals_iff intro!: continuous_intros) also have "?c = perzeta q 0" using q by (simp add: powr_minus perzeta_def Ln_Reals_eq a_def b_def Complex_eq mult_ac inverse_eq_divide) also have "?f \0\ \ \ perzeta q \0\ \" by (rule sym, intro filterlim_cong refl ev) finally show "isCont (perzeta q) s" by (simp add: isCont_def) qed have "perzeta q field_differentiable at s" for s proof (cases "s \ \") case False with holo have "perzeta q field_differentiable at s within -\" unfolding holomorphic_on_def by blast also have "at s within -\ = at s" using False by (intro at_within_open) auto finally show ?thesis . next case True hence *: "perzeta q holomorphic_on (ball s 1 - {s})" by (intro holomorphic_on_subset[OF holo]) (auto elim!: Nats_cases simp: dist_of_nat) have "perzeta q holomorphic_on ball s 1" using cont True by (intro no_isolated_singularity'[OF _ *]) (auto simp: at_within_open[of _ "ball s 1"] isCont_def) hence "perzeta q field_differentiable at s within ball s 1" unfolding holomorphic_on_def by auto thus ?thesis by (simp add: at_within_open[of _ "ball s 1"]) qed hence "perzeta q holomorphic_on UNIV" by (auto simp: holomorphic_on_def) also have "perzeta q = perzeta q'" by (simp add: q_def) finally show "perzeta q' holomorphic_on A" by auto qed lemma perzeta_altdef1: "Re s > 1 \ perzeta q' s = eval_fds (fds_perzeta q') s" by (cases "q' \ \") (auto simp: perzeta_int eval_fds_zeta fds_perzeta_int perzeta_altdef2) lemma holomorphic_perzeta: "q \ \ \ 1 \ A \ perzeta q holomorphic_on A" by (cases "q \ \") (auto simp: perzeta_int intro: holomorphic_perzeta' holomorphic_zeta) lemma holomorphic_perzeta'' [holomorphic_intros]: assumes "f holomorphic_on A" and "q \ \ \ (\x\A. f x \ 1)" shows "(\x. perzeta q (f x)) holomorphic_on A" proof - have "perzeta q \ f holomorphic_on A" using assms by (intro holomorphic_on_compose holomorphic_perzeta) auto thus ?thesis by (simp add: o_def) qed text \ Using this analytic continuation of the periodic zeta function, Hurwitz's formula now holds (almost) on the entire complex plane. \ theorem hurwitz_zeta_formula: fixes a :: real and s :: complex assumes "a \ {0<..1}" and "s \ 0" and "a \ 1 \ s \ 1" shows "rGamma s * hurwitz_zeta a (1 - s) = (2 * pi) powr - s * (\ powr - s * perzeta a s + \ powr s * perzeta (-a) s)" (is "?f s = ?g s") proof - define A where "A = UNIV - (if a \ \ then {0, 1} else {0 :: complex})" show ?thesis proof (rule analytic_continuation_open[where f = ?f]) show "?f holomorphic_on A" using assms by (auto intro!: holomorphic_intros simp: A_def) show "?g holomorphic_on A" using assms by (auto intro!: holomorphic_intros simp: A_def minus_in_Ints_iff) next fix s assume "s \ {s. Re s > 1}" thus "?f s = ?g s" using hurwitz_zeta_formula_aux[of a s] assms by (simp add: perzeta_altdef1) qed (insert assms, auto simp: open_halfspace_Re_gt A_def elim!: Ints_cases intro: connected_open_delete_finite exI[of _ 2]) qed text \ The equation expressing the periodic zeta function in terms of the Hurwitz zeta function can be extened similarly. \ theorem perzeta_conv_hurwitz_zeta: fixes a :: real and s :: complex assumes "a \ {0<..<1}" and "s \ 0" shows "rGamma (1 - s) * perzeta a s = (2 * pi) powr (s - 1) * \ * (\ powr (-s) * hurwitz_zeta a (1 - s) - \ powr s * hurwitz_zeta (1 - a) (1 - s))" (is "?f s = ?g s") proof (rule analytic_continuation_open[where f = ?f]) show "?f holomorphic_on -{0}" using assms by (auto intro!: holomorphic_intros elim: Ints_cases) show "?g holomorphic_on -{0}" using assms by (auto intro!: holomorphic_intros) next fix s assume "s \ {s. Re s > 1}" thus "?f s = ?g s" using perzeta_conv_hurwitz_zeta_aux[of a s] assms by (simp add: perzeta_altdef1) qed (insert assms, auto simp: open_halfspace_Re_gt connected_punctured_universe intro: exI[of _ 2]) text \ As a simple corollary, we derive the reflection formula for the Riemann zeta function: \ corollary zeta_reflect: fixes s :: complex assumes "s \ 0" "s \ 1" shows "rGamma s * zeta (1 - s) = 2 * (2 * pi) powr -s * cos (s * pi / 2) * zeta s" using hurwitz_zeta_formula[of 1 s] assms by (simp add: zeta_def cos_exp_eq powr_def perzeta_int algebra_simps) corollary zeta_reflect': fixes s :: complex assumes "s \ 0" "s \ 1" shows "rGamma (1 - s) * zeta s = 2 * (2 * pi) powr (s - 1) * sin (s * pi / 2) * zeta (1 - s)" using zeta_reflect[of "1 - s"] assms by (simp add: cos_sin_eq field_simps) text \ It is now easy to see that all the non-trivial zeroes of the Riemann zeta function must lie the critical strip $(0;1)$, and they must be symmetric around the $\mathfrak{R}(z) = \frac{1}{2}$ line. \ corollary zeta_zeroD: assumes "zeta s = 0" "s \ 1" shows "Re s \ {0<..<1} \ (\n::nat. n > 0 \ even n \ s = -real n)" proof (cases "Re s \ 0") case False with zeta_Re_ge_1_nonzero[of s] assms have "Re s < 1" by (cases "Re s < 1") auto with False show ?thesis by simp next case True { assume *: "\n. n > 0 \ even n \ s \ -real n" have "s \ of_int n" for n :: int proof assume [simp]: "s = of_int n" show False proof (cases n "0::int" rule: linorder_cases) assume "n < 0" show False proof (cases "even n") case True hence "nat (-n) > 0" "even (nat (-n))" using \n < 0\ by (auto simp: even_nat_iff) with * have "s \ -real (nat (-n))" . with \n < 0\ and True show False by auto next case False with \n < 0\ have "of_int n = (-of_nat (nat (-n)) :: complex)" by simp also have "zeta \ = -(bernoulli' (Suc (nat (-n)))) / of_nat (Suc (nat (-n)))" using \n < 0\ by (subst zeta_neg_of_nat) (auto) finally have "bernoulli' (Suc (nat (-n))) = 0" using assms by (auto simp del: of_nat_Suc) with False and \n < 0\ show False by (auto simp: bernoulli'_zero_iff even_nat_iff) qed qed (insert assms True, auto) qed hence "rGamma s \ 0" by (auto simp: rGamma_eq_zero_iff nonpos_Ints_def) moreover from assms have [simp]: "s \ 0" by auto ultimately have "zeta (1 - s) = 0" using zeta_reflect[of s] and assms by auto with True zeta_Re_ge_1_nonzero[of "1 - s"] have "Re s > 0" by auto } with True show ?thesis by auto qed lemma zeta_zero_reflect: assumes "Re s \ {0<..<1}" and "zeta s = 0" shows "zeta (1 - s) = 0" proof - from assms have "rGamma s \ 0" by (auto simp: rGamma_eq_zero_iff elim!: nonpos_Ints_cases) moreover from assms have "s \ 0" and "s \ 1" by auto ultimately show ?thesis using zeta_reflect[of s] and assms by auto qed corollary zeta_zero_reflect_iff: assumes "Re s \ {0<..<1}" shows "zeta (1 - s) = 0 \ zeta s = 0" using zeta_zero_reflect[of s] zeta_zero_reflect[of "1 - s"] assms by auto subsection \More functional equations\ lemma perzeta_conv_hurwitz_zeta_multiplication: fixes k :: nat and a :: int and s :: complex assumes "k > 0" "s \ 1" shows "k powr s * perzeta (a / k) s = (\n=1..k. exp (2 * pi * n * a / k * \) * hurwitz_zeta (n / k) s)" (is "?lhs s = ?rhs s") proof (rule analytic_continuation_open[where ?f = ?lhs and ?g = ?rhs]) show "connected (-{1::complex})" by (rule connected_punctured_universe) auto show "{s. Re s > 1} \ {}" by (auto intro!: exI[of _ 2]) next fix s assume s: "s \ {s. Re s > 1}" let ?f = "\n. exp (2 * pi * n * a / k * \)" show "?lhs s = ?rhs s" proof (rule sums_unique2) have "(\m. \n=1..k. ?f n * (of_nat m + of_real (real n / real k)) powr -s) sums (\n=1..k. ?f n * hurwitz_zeta (real n / real k) s)" using assms s by (intro sums_sum sums_mult sums_hurwitz_zeta) auto also have "(\m. \n=1..k. ?f n * (of_nat m + of_real (real n / real k)) powr -s) = (\m. of_nat k powr s * (\n=1..k. ?f n * of_nat (m * k + n) powr -s))" unfolding sum_distrib_left proof (intro ext sum.cong, goal_cases) case (2 m n) hence "m * k + n > 0" by (intro add_nonneg_pos) auto hence "of_nat 0 \ (of_nat (m * k + n) :: complex)" by (simp only: of_nat_eq_iff) also have "of_nat (m * k + n) = of_nat m * of_nat k + (of_nat n :: complex)" by simp finally have nz: "\ \ 0" by auto have "of_nat m + of_real (real n / real k) = (inverse (of_nat k) * of_nat (m * k + n) :: complex)" using assms (* TODO: Field_as_Ring messing up things again *) by (simp add: field_simps del: div_mult_self1 div_mult_self2 div_mult_self3 div_mult_self4) also from nz have "\ powr -s = of_nat k powr s * of_nat (m * k + n) powr -s" by (subst powr_times_real) (auto simp: add_eq_0_iff powr_def exp_minus Ln_inverse) finally show ?case by simp qed auto finally show "\ sums (\n=1..k. ?f n * hurwitz_zeta (real n / real k) s)" . next define g where "g = (\m. exp (2 * pi * \ * m * (real_of_int a / real k)))" have "(\m. g (Suc m) / (Suc m) powr s) sums eval_fds (fds_perzeta (a / k)) s" unfolding g_def using s by (intro sums_fds_perzeta) auto also have "(\m. g (Suc m) / (Suc m) powr s) = (\m. ?f (Suc m) * (Suc m) powr -s)" by (simp add: powr_minus field_simps g_def) also have "eval_fds (fds_perzeta (a / k)) s = perzeta (a / k) s" using s by (simp add: perzeta_altdef1) finally have "(\m. \n=m*k..k > 0\ by (rule sums_group) also have "(\m. \n=m*k..m. \n=1..k. ?f (m * k + n) * of_nat (m * k + n) powr -s)" proof (rule ext, goal_cases) case (1 m) show ?case using assms by (intro ext sum.reindex_bij_witness[of _ "\n. m * k + n - 1" "\n. Suc n - m * k"]) auto qed also have "(\m n. ?f (m * k + n)) = (\m n. ?f n)" proof (intro ext) fix m n :: nat have "?f (m * k + n) / ?f n = exp (2 * pi * m * a * \)" using \k > 0\ by (auto simp: ring_distribs add_divide_distrib exp_add mult_ac) also have "\ = cis (2 * pi * (m * a))" by (simp add: exp_eq_polar mult_ac) also have "\ = 1" by (rule cis_multiple_2pi) auto finally show "?f (m * k + n) = ?f n" by simp qed finally show "(\m. of_nat k powr s * (\n=1..k. ?f n * of_nat (m * k + n) powr -s)) sums (of_nat k powr s * perzeta (a / k) s)" by (rule sums_mult) qed qed (use assms in \auto intro!: holomorphic_intros simp: finite_imp_closed open_halfspace_Re_gt\) lemma perzeta_conv_hurwitz_zeta_multiplication': fixes k :: nat and a :: int and s :: complex assumes "k > 0" "s \ 1" shows "perzeta (a / k) s = k powr -s * (\n=1..k. exp (2 * pi * n * a / k * \) * hurwitz_zeta (n / k) s)" using perzeta_conv_hurwitz_zeta_multiplication[of k s a] assms by (simp add: powr_minus field_simps) lemma zeta_conv_hurwitz_zeta_multiplication: fixes k a :: nat and s :: complex assumes "k > 0" "s \ 1" shows "k powr s * zeta s = (\n=1..k. hurwitz_zeta (n / k) s)" using perzeta_conv_hurwitz_zeta_multiplication[of k s 0] using assms by (simp add: perzeta_int) lemma hurwitz_zeta_one_half_left: assumes "s \ 1" shows "hurwitz_zeta (1 / 2) s = (2 powr s - 1) * zeta s" using zeta_conv_hurwitz_zeta_multiplication[of 2 s] assms by (simp add: eval_nat_numeral zeta_def field_simps) theorem hurwitz_zeta_functional_equation: fixes h k :: nat and s :: complex assumes hk: "k > 0" "h \ {0<..k}" and s: "s \ {0, 1}" defines "a \ real h / real k" shows "rGamma s * hurwitz_zeta a (1 - s) = 2 * (2 * pi * k) powr -s * (\n=1..k. cos (s*pi/2 - 2*pi*n*h/k) * hurwitz_zeta (n / k) s)" proof - from hk have a: "a \ {0<..1}" by (auto simp: a_def) have "rGamma s * hurwitz_zeta a (1 - s) = (2 * pi) powr - s * (\ powr - s * perzeta a s + \ powr s * perzeta (- a) s)" using s a by (intro hurwitz_zeta_formula) auto also have "\ = (2 * pi) powr - s * (\ powr - s * perzeta (of_int (int h) / k) s + \ powr s * perzeta (of_int (-int h) / k) s)" by (simp add: a_def) also have "\ = (2 * pi) powr -s * k powr -s * ((\n=1..k. \ powr -s * cis (2 * pi * n * h / k) * hurwitz_zeta (n / k) s) + (\n=1..k. \ powr s * cis (-2 * pi * n * h / k) * hurwitz_zeta (n / k) s))" (is "_ = _ * (?S1 + ?S2)") using hk a s by (subst (1 2) perzeta_conv_hurwitz_zeta_multiplication') (auto simp: field_simps sum_distrib_left sum_distrib_right exp_eq_polar) also have "(2 * pi) powr -s * k powr -s = (2 * k * pi) powr -s" using hk pi_gt_zero by (simp add: powr_def Ln_times_Reals field_simps exp_add exp_diff exp_minus) also have "?S1 + ?S2 = (\n=1..k. (\ powr -s * cis (2*pi*n*h/k) + \ powr s * cis (-2*pi*n*h/k)) * hurwitz_zeta (n / k) s)" (is "_ = (\n\_. ?c n * _)") by (simp add: algebra_simps sum.distrib) also have "?c = (\n. 2 * cos (s*pi/2 - 2*pi*n*h/k))" proof fix n :: nat have "\ powr -s * cis (2*pi*n*h/k) = exp (-s*pi/2*\ + 2*pi*n*h/k*\)" unfolding exp_add by (simp add: powr_def cis_conv_exp mult_ac) moreover have "\ powr s * cis (-2*pi*n*h/k) = exp (s*pi/2*\ + -2*pi*n*h/k*\)" unfolding exp_add by (simp add: powr_def cis_conv_exp mult_ac) ultimately have "?c n = exp (\ * (s*pi/2 - 2*pi*n*h/k)) + exp (-(\ * (s*pi/2 - 2*pi*n*h/k)))" by (simp add: mult_ac ring_distribs) also have "\ / 2 = cos (s*pi/2 - 2*pi*n*h/k)" by (rule cos_exp_eq [symmetric]) finally show "?c n = 2 * cos (s*pi/2 - 2*pi*n*h/k)" by simp qed also have "(2 * k * pi) powr -s * (\n=1..k. \ n * hurwitz_zeta (n / k) s) = 2 * (2 * pi * k) powr -s * (\n=1..k. cos (s*pi/2 - 2*pi*n*h/k) * hurwitz_zeta (n / k) s)" by (simp add: sum_distrib_left sum_distrib_right mult_ac) finally show ?thesis . qed lemma perzeta_one_half_left: "s \ 1 \ perzeta (1 / 2) s = (2 powr (1-s) - 1) * zeta s" using perzeta_conv_hurwitz_zeta_multiplication'[of 2 s 1] by (simp add: eval_nat_numeral hurwitz_zeta_one_half_left powr_minus field_simps zeta_def powr_diff) lemma perzeta_one_half_left': "perzeta (1 / 2) s = (if s = 1 then -ln 2 else (2 powr (1 - s) - 1) / (s - 1)) * ((s - 1) * pre_zeta 1 s + 1)" by (cases "s = 1") (auto simp: perzeta_one_half_left field_simps zeta_def hurwitz_zeta_def) end