diff --git a/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Accuracy_Without_Cutoff.thy b/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Accuracy_Without_Cutoff.thy --- a/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Accuracy_Without_Cutoff.thy +++ b/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Accuracy_Without_Cutoff.thy @@ -1,1007 +1,1008 @@ section \Accuracy without cutoff\label{sec:accuracy_wo_cutoff}\ text \This section verifies that each of the $l$ estimate have the required accuracy with high probability assuming that there was no cut-off, i.e., that $s=0$. Section~\ref{sec:accuracy} will then show that this remains true as long as the cut-off is below @{term "t f"} the subsampling threshold.\ theory Distributed_Distinct_Elements_Accuracy_Without_Cutoff imports + Concentration_Inequalities.Bienaymes_Identity Distributed_Distinct_Elements_Inner_Algorithm Distributed_Distinct_Elements_Balls_and_Bins begin no_notation Polynomials.var ("X\") locale inner_algorithm_fix_A = inner_algorithm + fixes A assumes A_range: "A \ {.. A" begin definition X :: nat where "X = card A" definition q_max where "q_max = nat (\log 2 X\ - b_exp)" definition t :: "(nat \ nat) \ int" where "t f = int (Max (f ` A)) - b_exp + 9" definition s :: "(nat \ nat) \ nat" where "s f = nat (t f)" definition R :: "(nat \ nat) \ nat set" where "R f = {a. a \ A \ f a \ s f}" definition r :: "nat \ (nat \ nat) \ nat" where "r x f = card {a. a \ A \ f a \ x}" definition p where "p = (\(f,g,h). card {j\ {..\<^sub>1 (f,g,h) A 0 j \ s f})" definition Y where "Y = (\(f,g,h). 2 ^ s f * \_inv (p (f,g,h)))" lemma fin_A: "finite A" using A_range finite_nat_iff_bounded by auto lemma X_le_n: "X \ n" proof - have "card A \ card {.. 1" unfolding X_def using fin_A A_nonempty by (simp add: leI) lemma of_bool_square: "(of_bool x)\<^sup>2 = ((of_bool x)::real)" by (cases x, auto) lemma r_eq: "r x f = (\ a \ A.( of_bool( x \ f a) :: real))" unfolding r_def of_bool_def sum.If_cases[OF fin_A] by (simp add: Collect_conj_eq) lemma shows r_exp: "(\\. real (r x \) \ \\<^sub>1) = real X * (of_bool (x \ max (nat \log 2 n\) 1) / 2^x)" and r_var: "measure_pmf.variance \\<^sub>1 (\\. real (r x \)) \ (\\. real (r x \) \ \\<^sub>1)" proof - define V :: "nat \ (nat \ nat) \ real" where "V = (\a f. of_bool (x \ f a))" have V_exp: "(\\. V a \ \\\<^sub>1) = of_bool (x \ max (nat \log 2 n\) 1)/2^x" (is "?L = ?R") if "a \ A" for a proof - have a_le_n: "a < n" using that A_range by auto have "?L = (\\. indicator {f. x \ f a} \ \ \\<^sub>1)" unfolding V_def by (intro integral_cong_AE) auto also have "... = measure (map_pmf (\\. \ a) (sample_pmf \\<^sub>1)) {f. x \ f}" by simp also have "... = measure (\ n_exp) {f. x \ f}" unfolding \\<^sub>1.single[OF a_le_n] by simp also have "... = of_bool (x \ max (nat \log 2 n\) 1)/2^x" unfolding \_prob n_exp_def by simp finally show ?thesis by simp qed have b:"(\\. real (r x \) \ \\<^sub>1) = (\ a \ A. (\\. V a \ \\\<^sub>1))" unfolding r_eq V_def using \\<^sub>1.sample_space by (intro Bochner_Integration.integral_sum) auto also have "... = (\ a \ A. of_bool (x \ max (nat \log 2 n\) 1)/2^x)" using V_exp by (intro sum.cong) auto also have "... = X * (of_bool (x \ max (nat \log 2 n\) 1) / 2^x)" using X_def by simp finally show "(\\. real (r x \) \ \\<^sub>1) = real X * (of_bool (x \ max (nat \log 2 n\) 1)/ 2^x)" by simp have "(\\. (V a \)^2 \ \\<^sub>1) = (\\. V a \ \ \\<^sub>1)" for a unfolding V_def of_bool_square by simp hence a:"measure_pmf.variance \\<^sub>1 (V a) \ measure_pmf.expectation \\<^sub>1 (V a)" for a using \\<^sub>1.sample_space by (subst measure_pmf.variance_eq) auto have "J \ A \ card J = 2 \ prob_space.indep_vars \\<^sub>1 (\_. borel) V J" for J unfolding V_def using A_range finite_subset[OF _ fin_A] by (intro prob_space.indep_vars_compose2[where Y="\i y. of_bool(x \ y)" and M'="\_. discrete"] prob_space.k_wise_indep_vars_subset[OF _ \\<^sub>1.indep]) (auto simp:prob_space_measure_pmf) hence "measure_pmf.variance \\<^sub>1 (\\. real (r x \)) = (\ a \ A. measure_pmf.variance \\<^sub>1 (V a))" unfolding r_eq V_def using \\<^sub>1.sample_space by (intro measure_pmf.bienaymes_identity_pairwise_indep_2 fin_A) (simp_all) also have "... \ (\ a \ A. (\\. V a \ \ \\<^sub>1))" by (intro sum_mono a) also have "... = (\\. real (r x \) \ \\<^sub>1)" unfolding b by simp finally show "measure_pmf.variance \\<^sub>1 (\\. real (r x \)) \ (\\. real (r x \) \ \\<^sub>1)" by simp qed definition E\<^sub>1 where "E\<^sub>1 = (\(f,g,h). 2 powr (-t f) * X \ {b/2^16..b/2})" lemma t_low: "measure \\<^sub>1 {f. of_int (t f) < log 2 (real X) + 1 - b_exp} \ 1/2^7" (is "?L \ ?R") proof (cases "log 2 (real X) \ 8") case True define Z :: "(nat \ nat) \ real" where "Z = r (nat \log 2 (real X) - 8\)" have "log 2 (real X) \ log 2 (real n)" using X_le_n X_ge_1 by (intro log_mono) auto hence "nat \log 2 (real X) - 8\ \ nat \log 2 (real n)\" by (intro nat_mono ceiling_mono) simp hence a:"(nat \log 2 (real X) - 8\ \ max (nat \log 2 (real n)\) 1)" by simp have b:"real (nat (\log 2 (real X)\ - 8)) \ log 2 (real X) - 7" using True by linarith have "2 ^ 7 = real X / (2 powr (log 2 X) * 2 powr (-7))" using X_ge_1 by simp also have "... = real X / (2 powr (log 2 X - 7))" by (subst powr_add[symmetric]) simp also have "... \ real X / (2 powr (real (nat \log 2 (real X) - 8\)))" using b by (intro divide_left_mono powr_mono) auto also have "... = real X / 2 ^ nat \log 2 (real X) - 8\" by (subst powr_realpow) auto finally have "2 ^ 7 \ real X / 2 ^ nat \log 2 (real X) - 8\" by simp hence exp_Z_gt_2_7: "(\\. Z \ \\\<^sub>1) \ 2^7" using a unfolding Z_def r_exp by simp have var_Z_le_exp_Z: "measure_pmf.variance \\<^sub>1 Z \ (\\. Z \ \\\<^sub>1)" unfolding Z_def by (intro r_var) have "?L \ measure \\<^sub>1 {f. of_nat (Max (f ` A)) < log 2 (real X) - 8}" unfolding t_def by (intro pmf_mono) (auto simp add:int_of_nat_def) also have "... \ measure \\<^sub>1 {f \ space \\<^sub>1. (\\. Z \ \\\<^sub>1) \ \Z f - (\\. Z \ \\\<^sub>1) \}" proof (rule pmf_mono) fix f assume "f \ set_pmf (sample_pmf \\<^sub>1)" have fin_f_A: "finite (f ` A)" using fin_A finite_imageI by blast assume " f \ {f. real (Max (f ` A)) < log 2 (real X) - 8}" hence "real (Max (f ` A)) < log 2 (real X) - 8" by auto hence "real (f a) < log 2 (real X) - 8" if "a \ A" for a using Max_ge[OF fin_f_A] imageI[OF that] order_less_le_trans by fastforce hence "of_nat (f a) < \log 2 (real X) - 8\" if "a \ A" for a using that by (subst less_ceiling_iff) auto hence "f a < nat \log 2 (real X) - 8\" if "a \ A" for a using that True by fastforce hence "r (nat \log 2 (real X) - 8\) f = 0" unfolding r_def card_eq_0_iff using not_less by auto hence "Z f = 0" unfolding Z_def by simp thus "f \ {f \ space \\<^sub>1. (\\. Z \ \\\<^sub>1) \ \Z f - (\\. Z \ \\\<^sub>1)\}" by auto qed also have "... \ measure_pmf.variance \\<^sub>1 Z / (\\. Z \ \\\<^sub>1)^2" using exp_Z_gt_2_7 \\<^sub>1.sample_space by (intro measure_pmf.second_moment_method) simp_all also have "... \ (\\. Z \ \\\<^sub>1) / (\\. Z \ \\\<^sub>1)^2" by (intro divide_right_mono var_Z_le_exp_Z) simp also have "... = 1 / (\\. Z \ \\\<^sub>1)" using exp_Z_gt_2_7 by (simp add:power2_eq_square) also have "... \ ?R" using exp_Z_gt_2_7 by (intro divide_left_mono) auto finally show ?thesis by simp next case "False" have "?L \ measure \\<^sub>1 {f. of_nat (Max (f ` A)) < log 2 (real X) - 8}" unfolding t_def by (intro pmf_mono) (auto simp add:int_of_nat_def) also have "... \ measure \\<^sub>1 {}" using False by (intro pmf_mono) simp also have "... = 0" by simp also have "... \ ?R" by simp finally show ?thesis by simp qed lemma t_high: "measure \\<^sub>1 {f. of_int (t f) > log 2 (real X) + 16 - b_exp} \ 1/2^7" (is "?L \ ?R") proof - define Z :: "(nat \ nat) \ real" where "Z = r (nat \log 2 (real X) + 8\)" have Z_nonneg: "Z f \ 0" for f unfolding Z_def r_def by simp have "(\\. Z \ \\\<^sub>1) \ real X / (2 ^ nat \log 2 (real X) + 8\)" unfolding Z_def r_exp by simp also have "... \ real X / (2 powr (real (nat \log 2 (real X) + 8\)))" by (subst powr_realpow) auto also have "... \ real X / (2 powr \log 2 (real X) + 8\)" by (intro divide_left_mono powr_mono) auto also have "... \ real X / (2 powr (log 2 (real X) + 7))" by (intro divide_left_mono powr_mono, linarith) auto also have "... = real X / 2 powr (log 2 (real X)) / 2 powr 7" by (subst powr_add) simp also have "... \ 1/2 powr 7" using X_ge_1 by (subst powr_log_cancel) auto finally have Z_exp: "(\\. Z \ \\\<^sub>1) \ 1/2^7" by simp have "?L \ measure \\<^sub>1 {f. of_nat (Max (f ` A)) > log 2 (real X) + 7}" unfolding t_def by (intro pmf_mono) (auto simp add:int_of_nat_def) also have "... \ measure \\<^sub>1 {f. Z f \ 1}" proof (rule pmf_mono) fix f assume "f \ set_pmf (sample_pmf \\<^sub>1)" assume " f \ {f. real (Max (f ` A)) > log 2 (real X) + 7}" hence "real (Max (f ` A)) > log 2 (real X) + 7" by simp hence "int (Max (f ` A)) \ \log 2 (real X) + 8\" by linarith hence "Max (f ` A) \ nat \log 2 (real X) + 8\" by simp moreover have "f ` A \ {}" "finite (f ` A)" using fin_A finite_imageI A_nonempty by auto ultimately obtain fa where "fa \ f ` A" " fa \ nat \log 2 (real X) + 8\" using Max_in by auto then obtain ae where ae_def: "ae \ A" "nat \log 2 (real X) + 8\ \ f ae" by auto hence "r (nat \log 2 (real X) + 8\) f > 0" unfolding r_def card_gt_0_iff using fin_A by auto hence "Z f \ 1" unfolding Z_def by simp thus "f \ {f. 1 \ Z f}" by simp qed also have "... \ (\\. Z \ \\\<^sub>1) / 1" using Z_nonneg using \\<^sub>1.sample_space by (intro pmf_markov) auto also have "... \ ?R" using Z_exp by simp finally show ?thesis by simp qed lemma e_1: "measure \ {\. \E\<^sub>1 \} \ 1/2^6" proof - have "measure \\<^sub>1 {f. 2 powr (of_int (-t f)) * real X \ {real b/2^16..real b/2}} \ measure \\<^sub>1 {f. 2 powr (of_int (-t f)) * real X < real b/2^16} + measure \\<^sub>1 {f. 2 powr (of_int (-t f)) * real X > real b/2}" by (intro pmf_add) auto also have "... \ measure \\<^sub>1 {f. of_int (t f) > log 2 X + 16 - b_exp} + measure \\<^sub>1 {f. of_int (t f) < log 2 X + 1 - b_exp}" proof (rule add_mono) show "measure \\<^sub>1 {f. 2 powr (of_int (-t f)) * real X < real b/2^16} \ measure \\<^sub>1 {f. of_int (t f) > log 2 X + 16 - b_exp}" proof (rule pmf_mono) fix f assume "f \ {f. 2 powr real_of_int (-t f) * real X < real b / 2 ^ 16}" hence "2 powr real_of_int (-t f) * real X < real b / 2 ^ 16" by simp hence "log 2 (2 powr of_int (-t f) * real X) < log 2 (real b / 2^16)" using b_min X_ge_1 by (intro iffD2[OF log_less_cancel_iff]) auto hence "of_int (-t f) + log 2 (real X) < log 2 (real b / 2^16)" using X_ge_1 by (subst (asm) log_mult) auto also have "... = real b_exp - log 2 (2 powr 16)" unfolding b_def by (subst log_divide) auto also have "... = real b_exp - 16" by (subst log_powr_cancel) auto finally have "of_int (-t f) + log 2 (real X) < real b_exp - 16" by simp thus "f \ {f. of_int (t f) > log 2 (real X) + 16 - b_exp}" by simp qed next show "measure \\<^sub>1 {f. 2 powr of_int (-t f) * real X > real b/2} \ measure \\<^sub>1 {f. of_int (t f) < log 2 X + 1 - b_exp}" proof (rule pmf_mono) fix f assume "f \ {f. 2 powr real_of_int (-t f) * real X > real b / 2}" hence "2 powr real_of_int (-t f) * real X > real b / 2" by simp hence "log 2 (2 powr of_int (-t f) * real X) > log 2 (real b / 2)" using b_min X_ge_1 by (intro iffD2[OF log_less_cancel_iff]) auto hence "of_int (-t f) + log 2 (real X) > log 2 (real b / 2)" using X_ge_1 by (subst (asm) log_mult) auto hence "of_int (-t f) + log 2 (real X) > real b_exp - 1" unfolding b_def by (subst (asm) log_divide) auto hence "of_int (t f) < log 2 (real X) + 1 - b_exp" by simp thus "f \ {f. of_int (t f) < log 2 (real X) + 1 - b_exp}" by simp qed qed also have "... \ 1/2^7 + 1/2^7" by (intro add_mono t_low t_high) also have "... = 1/2^6" by simp finally have "measure \\<^sub>1 {f. 2 powr of_int (-t f) * real X \ {real b/2^16..real b/2}} \ 1/2^6" by simp thus ?thesis unfolding sample_pmf_\ E\<^sub>1_def case_prod_beta by (subst pair_pmf_prob_left) qed definition E\<^sub>2 where "E\<^sub>2 = (\(f,g,h). \card (R f) - X / 2^(s f)\ \ \/3 * X / 2^(s f))" lemma e_2: "measure \ {\. E\<^sub>1 \ \ \E\<^sub>2 \} \ 1/2^6" (is "?L \ ?R") proof - define t\<^sub>m :: int where "t\<^sub>m = \log 2 (real X)\ + 16 - b_exp" have t_m_bound: "t\<^sub>m \ \log 2 (real X)\ - 10" unfolding t\<^sub>m_def using b_exp_ge_26 by simp have "real b / 2^16 = (real X * (1/ X)) * (real b / 2^16)" using X_ge_1 by simp also have "... = (real X * 2 powr (-log 2 X)) * (real b / 2^16)" using X_ge_1 by (subst powr_minus_divide) simp also have "... \ (real X * 2 powr (- \log 2 (real X)\)) * (2 powr b_exp / 2^16)" unfolding b_def using powr_realpow by (intro mult_mono powr_mono) auto also have "... = real X * (2 powr (- \log 2 (real X)\) * 2 powr(real b_exp-16))" by (subst powr_diff) simp also have "... = real X * 2 powr (- \log 2 (real X)\ + (int b_exp - 16))" by (subst powr_add[symmetric]) simp also have "... = real X * 2 powr (-t\<^sub>m)" unfolding t\<^sub>m_def by (simp add:algebra_simps) finally have c:"real b / 2^16 \ real X * 2 powr (-t\<^sub>m)" by simp define T :: "nat set" where "T = {x. (real X / 2^x \ real b / 2^16)}" have "x \ T \ int x \ t\<^sub>m" for x proof - have "x \ T \ 2^ x \ real X * 2^16 / b" using b_min by (simp add: field_simps T_def) also have "... \ log 2 (2^x) \ log 2 (real X * 2^16 / b)" using X_ge_1 b_min by (intro log_le_cancel_iff[symmetric] divide_pos_pos) auto also have "... \ x \ log 2 (real X * 2^16) - log 2 b" using X_ge_1 b_min by (subst log_divide) auto also have "... \ x \ log 2 (real X) + log 2 (2 powr 16) - b_exp" unfolding b_def using X_ge_1 by (subst log_mult) auto also have "... \ x \ \log 2 (real X) + log 2 (2 powr 16) - b_exp\" by linarith also have "... \ x \ \log 2 (real X) + 16 - real_of_int (int b_exp)\" by (subst log_powr_cancel) auto also have "... \ x \ t\<^sub>m" unfolding t\<^sub>m_def by linarith finally show ?thesis by simp qed hence T_eq: "T = {x. int x \ t\<^sub>m}" by auto have "T = {x. int x < t\<^sub>m+1}" unfolding T_eq by simp also have "... = {x. x < nat (t\<^sub>m + 1)}" unfolding zless_nat_eq_int_zless by simp finally have T_eq_2: "T = {x. x < nat (t\<^sub>m + 1)}" by simp have inj_1: "inj_on ((-) (nat t\<^sub>m)) T" unfolding T_eq by (intro inj_onI) simp have fin_T: "finite T" unfolding T_eq_2 by simp have r_exp: "(\\. real (r t \) \\\<^sub>1) = real X / 2^t" if "t \ T" for t proof - have "t \ t\<^sub>m" using that unfolding T_eq by simp also have "... \ \log 2 (real X)\ - 10" using t_m_bound by simp also have "... \ \log 2 (real X)\" by simp also have "... \ \log 2 (real n)\" using X_le_n X_ge_1 by (intro floor_mono log_mono) auto also have "... \ \log 2 (real n)\" by simp finally have "t \ \log 2 (real n)\" by simp hence "t \ max (nat \log 2 (real n)\) 1"by simp thus ?thesis unfolding r_exp by simp qed have r_var: "measure_pmf.variance \\<^sub>1 (\\. real (r t \)) \ real X / 2^t" if "t \ T" for t using r_exp[OF that] r_var by metis have "9 = C\<^sub>4 / \\<^sup>2 * \^2/2^23" using \_gt_0 by (simp add:C\<^sub>4_def) also have "... = 2 powr (log 2 (C\<^sub>4 / \\<^sup>2)) * \^2/2^23" using \_gt_0 C\<^sub>4_def by (subst powr_log_cancel) auto also have "... \ 2 powr b_exp * \^2/2^23" unfolding b_exp_def by (intro divide_right_mono mult_right_mono powr_mono, linarith) auto also have "... = b * \^2/2^23" using powr_realpow unfolding b_def by simp also have "... = (b/2^16) * (\^2/2^7)" by simp also have "... \ (X * 2 powr (-t\<^sub>m)) * (\^2/2^7)" by (intro mult_mono c) auto also have "... = X * (2 powr (-t\<^sub>m) * 2 powr (-7)) * \^2" using powr_realpow by simp also have "... = 2 powr (-t\<^sub>m-7) * (\^2 * X)" by (subst powr_add[symmetric]) (simp ) finally have "9 \ 2 powr (-t\<^sub>m-7) * (\^2 * X)" by simp hence b: "9/ (\^2 * X) \ 2 powr (-t\<^sub>m -7)" using \_gt_0 X_ge_1 by (subst pos_divide_le_eq) auto have a: "measure \\<^sub>1 {f.\real (r t f)-real X/2^t\> \/3 *real X/2^t} \ 2 powr (real t-t\<^sub>m-7)" (is"?L1 \ ?R1") if "t \ T" for t proof - have "?L1 \ \

(f in \\<^sub>1. \real (r t f) - real X / 2^t\ \ \/3 * real X / 2^t)" by (intro pmf_mono) auto also have "... = \

(f in \\<^sub>1. \real (r t f)-(\\. real (r t \) \ \\<^sub>1)\ \ \/3 * real X/2^t)" by (simp add: r_exp[OF that]) also have "... \ measure_pmf.variance \\<^sub>1 (\\. real (r t \)) / (\/3 * real X / 2^t)^2" using X_ge_1 \_gt_0 \\<^sub>1.sample_space by (intro measure_pmf.Chebyshev_inequality divide_pos_pos mult_pos_pos) auto also have "... \ (X / 2^t) / (\/3 * X / 2^t)^2" by (intro divide_right_mono r_var[OF that]) simp also have "... = 2^t*(9/ ( \^2 * X))" by (simp add:power2_eq_square algebra_simps) also have "... \ 2^t*(2 powr (-t\<^sub>m-7))" by (intro mult_left_mono b) simp also have "... = 2 powr t * 2 powr (-t\<^sub>m-7)" by (subst powr_realpow[symmetric]) auto also have "... = ?R1" by (subst powr_add[symmetric]) (simp add:algebra_simps) finally show "?L1 \ ?R1" by simp qed have "\ym + 1). x = nat t\<^sub>m - y" if "x < nat (t\<^sub>m+1)" for x using that by (intro exI[where x="nat t\<^sub>m - x"]) simp hence T_reindex: "(-) (nat t\<^sub>m) ` {x. x < nat (t\<^sub>m + 1)} = {..m + 1)}" by (auto simp add: set_eq_iff image_iff) have "?L \ measure \ {\. (\t \ T. \real (r t (fst \))-real X/2^t\ > \/3 * real X / 2^t)}" proof (rule pmf_mono) fix \ assume "\ \ set_pmf (sample_pmf \)" obtain f g h where \_def: "\ = (f,g,h)" by (metis prod_cases3) assume "\ \ {\. E\<^sub>1 \ \ \ E\<^sub>2 \}" hence a:"2 powr ( -real_of_int (t f)) * real X \ {real b/2^16..real b/2}" and b:"\card (R f) - real X / 2^(s f)\ > \/3 * X / 2^(s f)" unfolding E\<^sub>1_def E\<^sub>2_def by (auto simp add:\_def) have "\card (R f) - X / 2^(s f)\ = 0" if "s f= 0" using that by (simp add:R_def X_def) moreover have "( \/3) * (X / 2^s f) \ 0" using \_gt_0 X_ge_1 by (intro mult_nonneg_nonneg) auto ultimately have "False" if "s f = 0" using b that by simp hence "s f > 0" by auto hence "t f = s f" unfolding s_def by simp hence "2 powr (-real (s f)) * X \ b / 2^16" using a by simp hence "X / 2 powr (real (s f)) \ b / 2^16" by (simp add: divide_powr_uminus mult.commute) hence "real X / 2 ^ (s f) \ b / 2^16" by (subst (asm) powr_realpow, auto) hence "s f \ T" unfolding T_def by simp moreover have "\r (s f) f - X / 2^s f\ > \/3 * X / 2^s f" using R_def r_def b by simp ultimately have "\t \ T. \r t (fst \) - X / 2^t\ > \/3 * X / 2^t" using \_def by (intro bexI[where x="s f"]) simp thus "\ \ {\. (\t \ T. \r t (fst \) - X / 2^t\ > \/3 * X / 2^t)}" by simp qed also have "... = measure \\<^sub>1 {f. (\t \ T. \real (r t f)-real X / 2^t\ > \/3 * real X/2^t)}" unfolding sample_pmf_\ by (intro pair_pmf_prob_left) also have "... = measure \\<^sub>1 (\t \ T. {f. \real (r t f)-real X / 2^t\ > \/3 * real X/2^t})" by (intro measure_pmf_cong) auto also have "... \ (\t \ T. measure \\<^sub>1 {f.\real (r t f)-real X / 2^t\ > \/3 * real X/2^t})" by (intro measure_UNION_le fin_T) (simp) also have "... \ (\t \ T. 2 powr (real t - of_int t\<^sub>m - 7))" by (intro sum_mono a) also have "... = (\t \ T. 2 powr (-int (nat t\<^sub>m-t) - 7))" unfolding T_eq by (intro sum.cong refl arg_cong2[where f="(powr)"]) simp also have "... = (\x \ (\x. nat t\<^sub>m - x) ` T. 2 powr (-real x - 7))" by (subst sum.reindex[OF inj_1]) simp also have "... = (\x \ (\x. nat t\<^sub>m - x) ` T. 2 powr (-7) * 2 powr (-real x))" by (subst powr_add[symmetric]) (simp add:algebra_simps) also have "... = 1/2^7 * (\x \ (\x. nat t\<^sub>m - x) ` T. 2 powr (-real x))" by (subst sum_distrib_left) simp also have "... = 1/2^7 * (\x m+1). 2 powr (-real x))" unfolding T_eq_2 T_reindex by (intro arg_cong2[where f="(*)"] sum.cong) auto also have "... = 1/2^7 * (\x m+1). (2 powr (-1)) powr (real x))" by (subst powr_powr) simp also have "... = 1/2^7 * (\x m+1). (1/2)^x)" using powr_realpow by simp also have "... \ 1/2^7 * 2" by(subst geometric_sum) auto also have "... = 1/2^6" by simp finally show ?thesis by simp qed definition E\<^sub>3 where "E\<^sub>3 = (\(f,g,h). inj_on g (R f))" lemma R_bound: fixes f g h assumes "E\<^sub>1 (f,g,h)" assumes "E\<^sub>2 (f,g,h)" shows "card (R f) \ 2/3 * b" proof - have "real (card (R f)) \ ( \ / 3) * (real X / 2 ^ s f) + real X / 2 ^ s f" using assms(2) unfolding E\<^sub>2_def by simp also have "... \ (1/3) * (real X / 2 ^ s f) + real X / 2 ^ s f" using \_lt_1 by (intro add_mono mult_right_mono) auto also have "... = (4/3) * (real X / 2 powr s f)" using powr_realpow by simp also have "... \ (4/3) * (real X / 2 powr t f)" unfolding s_def by (intro mult_left_mono divide_left_mono powr_mono) auto also have "... = (4/3) * (2 powr (-(of_int (t f))) * real X)" by (subst powr_minus_divide) simp also have "... = (4/3) * (2 powr (- t f) * real X)" by simp also have "... \ (4/3) * (b/2)" using assms(1) unfolding E\<^sub>1_def by (intro mult_left_mono) auto also have "... \ (2/3) * b" by simp finally show ?thesis by simp qed lemma e_3: "measure \ {\. E\<^sub>1 \ \ E\<^sub>2 \ \ \E\<^sub>3 \} \ 1/2^6" (is "?L \ ?R") proof - let ?\ = "(\(z,x,y) f. z < C\<^sub>7*b^2 \ x \ R f \ y \ R f \ x < y)" let ?\ = "(\(z,x,y) g. g x = z \ g y = z)" have \_prob: "measure \\<^sub>2 {g. ?\ \ g} \ (1/real (C\<^sub>7*b^2)^2)" if "?\ \ f" for \ f proof - obtain x y z where \_def: "\ = (z,x,y)" by (metis prod_cases3) have a:"prob_space.k_wise_indep_vars \\<^sub>2 2 (\i. discrete) (\x \. \ x = z) {..\<^sub>2.indep]) (simp_all add:prob_space_measure_pmf) have "u \ R f \ u < n" for u unfolding R_def using A_range by auto hence b: "x < n" "y < n" "card {x, y} = 2" using that \_def by auto have c: "z < C\<^sub>7*b\<^sup>2" using \_def that by simp have "measure \\<^sub>2 {g. ?\ \ g} = measure \\<^sub>2 {g. (\\ \ {x,y}. g \ = z)}" by (simp add:\_def) also have "... = (\\ \ {x,y}. measure \\<^sub>2 {g. g \ = z})" using b by (intro measure_pmf.split_indep_events[OF refl, where I="{x,y}"] prob_space.k_wise_indep_vars_subset[OF _ a]) (simp_all add:prob_space_measure_pmf) also have "... = (\\ \ {x,y}. measure (map_pmf (\\. \ \) (sample_pmf \\<^sub>2)) {g. g = z}) " by (simp add:vimage_def) also have "... = (\\ \ {x,y}. measure [C\<^sub>7 * b\<^sup>2]\<^sub>S {g. g=z})" using b \\<^sub>2.single by (intro prod.cong) fastforce+ also have "... = (\\ \ {x,y}. measure (pmf_of_set {..7 * b\<^sup>2}) {z})" by (subst nat_sample_pmf) simp also have "... = (measure (pmf_of_set {..7 * b\<^sup>2}) {z})^2" using b by simp also have "... \ (1 /(C\<^sub>7*b\<^sup>2))^2" using c by (subst measure_pmf_of_set) auto also have "... = (1 /(C\<^sub>7*b\<^sup>2)^2)" by (simp add:algebra_simps power2_eq_square) finally show ?thesis by simp qed have \_card: "card {\. ?\ \ f} \ (C\<^sub>7*b^2) * (card (R f) * (card (R f)-1)/2)" (is "?TL \ ?TR") and fin_\: "finite {\. ?\ \ f}" (is "?T2") for f proof - have t1: "{\. ?\ \ f} \ {..7*b^2} \ {(x,y) \ R f \ R f. x < y}" by (intro subsetI) auto moreover have "card ({..7*b^2} \ {(x,y) \ R f \ R f. x < y}) = ?TR" using card_ordered_pairs'[where M="R f"] by (simp add: card_cartesian_product) moreover have "finite (R f)" unfolding R_def using fin_A finite_subset by simp hence "finite {(x, y). (x, y) \ R f \ R f \ x < y}" by (intro finite_subset[where B="R f \ R f", OF _ finite_cartesian_product]) auto hence t2: "finite ({..7*b^2} \ {(x,y) \ R f \ R f. x < y})" by (intro finite_cartesian_product) auto ultimately show "?TL \ ?TR" using card_mono of_nat_le_iff by (metis (no_types, lifting)) show ?T2 using finite_subset[OF t1 t2] by simp qed have "?L \ measure \ {(f,g,h). card (R f) \ b \ (\ x y z. ?\ (x,y,z) f \ ?\ (x,y,z) g)}" proof (rule pmf_mono) fix \ assume b:"\ \ set_pmf (sample_pmf \)" obtain f g h where \_def:"\ = (f,g,h)" by (metis prod_cases3) have "(f,g,h) \ sample_set \" using sample_space_alt[OF sample_space_\] b \_def by simp hence c:"g x < C\<^sub>7*b^2" for x using g_range by simp assume a:"\ \ {\. E\<^sub>1 \ \ E\<^sub>2 \ \ \ E\<^sub>3 \}" hence "card (R f) \ 2/3 * b" using R_bound \_def by force moreover have "\a b. a \ R f \ b \ R f \ a \ b \ g a = g b" using a unfolding \_def E\<^sub>3_def inj_on_def by auto hence "\x y. x \ R f \ y \ R f \ x < y \ g x = g y" by (metis not_less_iff_gr_or_eq) hence "\x y z. ?\ (x,y,z) f \ ?\ (x,y,z) g" using c by blast ultimately show "\ \ {(f, g, h). card (R f) \ b \ (\ x y z. ?\ (x,y,z) f \ ?\ (x,y,z) g)}" unfolding \_def by auto qed also have "... = (\f. measure (pair_pmf \\<^sub>2 \\<^sub>3) {g. card (R f) \ b \ (\x y z. ?\ (x,y,z) f \ ?\ (x,y,z) (fst g))} \\\<^sub>1)" unfolding sample_pmf_\ split_pair_pmf by (simp add: case_prod_beta) also have "... = (\f. measure \\<^sub>2 {g. card (R f) \ b \ (\x y z. ?\ (x,y,z) f \ ?\ (x,y,z) g)} \\\<^sub>1)" by (subst pair_pmf_prob_left) simp also have "... \ (\f. 1/real (2*C\<^sub>7) \\\<^sub>1)" proof (rule pmf_exp_mono[OF integrable_sample_pmf[OF \\<^sub>1.sample_space] integrable_sample_pmf[OF \\<^sub>1.sample_space]]) fix f assume "f \ set_pmf (sample_pmf \\<^sub>1)" show "measure \\<^sub>2 {g. card (R f) \ b \ (\x y z. ?\ (x,y,z) f \ ?\ (x,y,z) g)} \ 1 / real (2 * C\<^sub>7)" (is "?L1 \ ?R1") proof (cases "card (R f) \ b") case True have "?L1 \ measure \\<^sub>2 (\ \ \ {\. ?\ \ f}. {g. ?\ \ g})" by (intro pmf_mono) auto also have "... \ (\\ \ {\. ?\ \ f}. measure \\<^sub>2 {g. ?\ \ g})" by (intro measure_UNION_le fin_\) auto also have "... \ (\\ \ {\. ?\ \ f}. (1/real (C\<^sub>7*b^2)^2))" by (intro sum_mono \_prob) auto also have "... = card {\. ?\ \ f} /(C\<^sub>7*b^2)^2" by simp also have "... \ (C\<^sub>7*b^2) * (card (R f) * (card (R f)-1)/2) / (C\<^sub>7*b^2)^2" by (intro \_card divide_right_mono) simp also have "... \ (C\<^sub>7*b^2) * (b * b / 2) / (C\<^sub>7*b^2)^2" unfolding C\<^sub>7_def using True by (intro divide_right_mono Nat.of_nat_mono mult_mono) auto also have "... = 1/(2*C\<^sub>7)" using b_min by (simp add:algebra_simps power2_eq_square) finally show ?thesis by simp next case False then show ?thesis by simp qed qed also have "... \ 1/2^6" unfolding C\<^sub>7_def by simp finally show ?thesis by simp qed definition E\<^sub>4 where "E\<^sub>4 = (\(f,g,h). \p (f,g,h) - \ (card (R f))\ \ \/12 * card (R f))" lemma e_4_h: "9 / sqrt b \ \ / 12" proof - have "108 \ sqrt (C\<^sub>4)" unfolding C\<^sub>4_def by (approximation 5) also have "... \ sqrt( \^2 * real b)" using b_lower_bound \_gt_0 by (intro real_sqrt_le_mono) (simp add: pos_divide_le_eq algebra_simps) also have "... = \ * sqrt b" using \_gt_0 by (simp add:real_sqrt_mult) finally have "108 \ \ * sqrt b" by simp thus ?thesis using b_min by (simp add:pos_divide_le_eq) qed lemma e_4: "measure \ {\. E\<^sub>1 \ \ E\<^sub>2 \ \ E\<^sub>3 \ \ \E\<^sub>4 \} \ 1/2^6" (is "?L \ ?R") proof - have a: "measure \\<^sub>3 {h. E\<^sub>1 (f,g,h) \ E\<^sub>2 (f,g,h) \ E\<^sub>3 (f,g,h) \ \E\<^sub>4 (f,g,h)} \ 1/2^6" (is "?L1 \ ?R1") if "f \ set_pmf (sample_pmf \\<^sub>1)" "g \ set_pmf(sample_pmf \\<^sub>2)" for f g proof (cases "card (R f) \ b \ inj_on g (R f)") case True have g_inj: "inj_on g (R f)" using True by simp have fin_R: "finite (g ` R f)" unfolding R_def using fin_A by (intro finite_imageI) simp interpret B:balls_and_bins_abs "g ` R f" "{.. {..7 * b\<^sup>2}" using g_range_1 that(2) unfolding sample_space_alt[OF \\<^sub>2.sample_space] by auto hence g_ran: "g ` R f \ {..7 * b\<^sup>2}" by auto have "sample_pmf [b]\<^sub>S = pmf_of_set {..\. \ x) (sample_pmf (\ k (C\<^sub>7 * b\<^sup>2) [b]\<^sub>S)) = pmf_of_set {.. g ` R f" for x using g_ran \\<^sub>3.single that by auto moreover have "prob_space.k_wise_indep_vars \\<^sub>3 k (\_. discrete) (\x \. \ x) (g ` R f)" by (intro prob_space.k_wise_indep_subset[OF _ _ \\<^sub>3.indep] g_ran prob_space_measure_pmf) ultimately have lim_balls_and_bins: "B.lim_balls_and_bins k (sample_pmf (\ k (C\<^sub>7 * b\<^sup>2) [b]\<^sub>S))" unfolding B.lim_balls_and_bins_def by auto have card_g_R: "card (g ` R f) = card (R f)" using True card_image by auto hence b_mu: "\ (card (R f)) = B.\" unfolding B.\_def \_def using b_min by (simp add:powr_realpow) have card_g_le_b: "card (g ` R f) \ card {.. measure \\<^sub>3 {h. \B.Y h - B.\\ > 9 * real (card (g ` R f)) / sqrt (card {.. {h. E\<^sub>1 (f,g,h) \ E\<^sub>2 (f,g,h) \ E\<^sub>3 (f,g,h) \ \E\<^sub>4 (f,g,h)}" hence b: "\p (f,g,h) -\ (card (R f))\ > \/12 * card (R f)" unfolding E\<^sub>4_def by simp assume "h \ set_pmf (sample_pmf \\<^sub>3)" hence h_range: "h x < b" for x unfolding sample_space_alt[OF \\<^sub>3.sample_space,symmetric] using h_range_1 by simp have "{j \ {.. \\<^sub>1 (f, g, h) A 0 j} = {j \ {.. max (Max ({int (f a) |a. a \ A \ h (g a) = j} \ {-1})) (- 1)}" unfolding \\<^sub>1_def by simp also have "... = {j \ {.. Max ({int (f a) |a. a \ A \ h (g a) = j} \ {-1})}" using fin_A by (subst max_absorb1) (auto intro: Max_ge) also have "... = {j \ {..a \ R f. h (g a) = j)}" unfolding R_def using fin_A by (subst Max_ge_iff) auto also have "... = {j. \a \ R f. h (g a) = j}" using h_range by auto also have "... = (h \ g) ` (R f)" by (auto simp add:set_eq_iff image_iff) also have "... = h ` (g ` (R f))" by (simp add:image_image) finally have c:"{j \ {.. \\<^sub>1 (f, g, h) A 0 j} = h ` (g ` R f)" by simp have "9 * real (card (g ` R f)) / sqrt (card {.. \/12 * card (R f)" by (intro mult_right_mono e_4_h) simp also have "... < \B.Y h - B.\\" using b c unfolding B.Y_def p_def b_mu by simp finally show "h \ {h. \B.Y h - B.\\ > 9 * real (card (g ` R f)) / sqrt (card {.. 1/2^6" using k_min by (intro B.devitation_bound[OF card_g_le_b lim_balls_and_bins]) auto finally show ?thesis by simp next case False have "?L1 \ measure \\<^sub>3 {}" proof (rule pmf_mono) fix h assume b:"h \ {h. E\<^sub>1 (f, g, h) \ E\<^sub>2 (f, g, h) \ E\<^sub>3 (f, g, h) \ \ E\<^sub>4 (f, g, h)}" hence "card (R f) \ (2/3)*b" by (auto intro!: R_bound[simplified]) hence "card (R f) \ b" by simp moreover have "inj_on g (R f)" using b by (simp add:E\<^sub>3_def) ultimately have "False" using False by simp thus "h \ {}" by simp qed also have "... = 0" by simp finally show ?thesis by simp qed have "?L = (\f. (\g. measure \\<^sub>3 {h. E\<^sub>1 (f,g,h) \ E\<^sub>2 (f,g,h) \ E\<^sub>3 (f,g,h) \ \E\<^sub>4 (f,g,h)} \\\<^sub>2) \\\<^sub>1)" unfolding sample_pmf_\ split_pair_pmf by simp also have "... \ (\f. (\g. 1/2^6 \\\<^sub>2) \\\<^sub>1)" using a \\<^sub>1.sample_space \\<^sub>2.sample_space by (intro integral_mono_AE AE_pmfI) simp_all also have "... = 1/2^6" by simp finally show ?thesis by simp qed lemma \_inverse: "\_inv (\ x) = x" proof - have a:"1-1/b \ 0" using b_min by simp have "\ x = b * (1-(1-1/b) powr x)" unfolding \_def by simp hence "\ x / real b = 1-(1-1/b) powr x" by simp hence "ln (1 - \ x / real b) = ln ((1-1/b) powr x)" by simp also have "... = x * ln (1 - 1/ b)" using a by (intro ln_powr) finally have "ln (1 - \ x / real b) = x * ln (1- 1/ b)" by simp moreover have "ln (1-1/b) < 0" using b_min by (subst ln_less_zero_iff) auto ultimately show ?thesis using \_inv_def by simp qed lemma rho_mono: assumes "x \ y" shows "\ x \ \ y" proof- have "(1 - 1 / real b) powr y \ (1 - 1 / real b) powr x" using b_min by (intro powr_mono_rev assms) auto thus ?thesis unfolding \_def by (intro mult_left_mono) auto qed lemma rho_two_thirds: "\ (2/3 * b) \ 3/5 *b" proof - have "1/3 \ exp ( - 13 / 12::real )" by (approximation 8) also have "... \ exp ( - 1 - 2 / real b )" using b_min by (intro iffD2[OF exp_le_cancel_iff]) (simp add:algebra_simps) also have "... \ exp ( b * (-(1/real b)-2*(1/real b)^2))" using b_min by (simp add:algebra_simps power2_eq_square) also have "... \ exp ( b * ln (1-1/real b))" using b_min by (intro iffD2[OF exp_le_cancel_iff] mult_left_mono ln_one_minus_pos_lower_bound) auto also have "... = exp ( ln ( (1-1/real b) powr b))" using b_min by (subst ln_powr) auto also have "... = (1-1/real b) powr b" using b_min by (subst exp_ln) auto finally have a:"1/3 \ (1-1/real b) powr b" by simp have "2/5 \ (1/3) powr (2/3::real)" by (approximation 5) also have "... \ ((1-1/real b) powr b) powr (2/3)" by (intro powr_mono2 a) auto also have "... = (1-1/real b) powr (2/3 * real b)" by (subst powr_powr) (simp add:algebra_simps) finally have "2/5 \ (1 - 1 / real b) powr (2 / 3 * real b)" by simp hence "1 - (1 - 1 / real b) powr (2 / 3 * real b) \ 3/5" by simp hence "\ (2/3 * b) \ b * (3/5)" unfolding \_def by (intro mult_left_mono) auto thus ?thesis by simp qed definition \_inv' :: "real \ real" where "\_inv' x = -1 / (real b * (1-x / real b) * ln (1 - 1 / real b))" lemma \_inv'_bound: assumes "x \ 0" assumes "x \ 59/90*b" shows "\\_inv' x\ \ 4" proof - have c:"ln (1 - 1 / real b) < 0" using b_min by (subst ln_less_zero_iff) auto hence d:"real b * (1 - x / real b) * ln (1 - 1 / real b) < 0" using b_min assms by (intro Rings.mult_pos_neg) auto have "(1::real) \ 31/30" by simp also have "... \ (31/30) * (b * -(- 1 / real b))" using b_min by simp also have "... \ (31/30) * (b * -ln (1 + (- 1 / real b)))" using b_min by (intro mult_left_mono le_imp_neg_le ln_add_one_self_le_self2) auto also have "... \ 3 * (31/90) * (- b * ln (1 - 1 / real b))" by simp also have "... \ 3 * (1 - x / real b) * (- b * ln (1 - 1 / real b))" using assms b_min pos_divide_le_eq[where c="b"] c by (intro mult_right_mono mult_left_mono mult_nonpos_nonpos) auto also have "... \ 3 * (real b * (1 - x / real b) * (-ln (1 - 1 / real b)))" by (simp add:algebra_simps) finally have "3 * (real b * (1 - x / real b) * (-ln (1 - 1 / real b))) \ 1" by simp hence "3 * (real b * (1 - x / real b) * ln (1 - 1 / real b)) \ -1" by simp hence "\_inv' x \ 3" unfolding \_inv'_def using d by (subst neg_divide_le_eq) auto moreover have "\_inv' x > 0" unfolding \_inv'_def using d by (intro divide_neg_neg) auto ultimately show ?thesis by simp qed lemma \_inv': fixes x :: real assumes "x < b" shows "DERIV \_inv x :> \_inv' x" proof - have "DERIV (ln \ (\x. (1 - x / real b))) x :> 1 / (1-x / real b) * (0 -1/b)" using assms b_min by (intro DERIV_chain DERIV_ln_divide DERIV_cdivide derivative_intros) auto hence "DERIV \_inv x :> (1 / (1-x / real b) * (-1/b)) / ln (1-1/real b)" unfolding comp_def \_inv_def by (intro DERIV_cdivide) auto thus ?thesis by (simp add:\_inv'_def algebra_simps) qed lemma accuracy_without_cutoff: "measure \ {(f,g,h). \Y (f,g,h) - real X\ > \ * X \ s f < q_max} \ 1/2^4" (is "?L \ ?R") proof - have "?L \ measure \ {\. \E\<^sub>1 \ \ \E\<^sub>2 \ \ \E\<^sub>3 \ \ \E\<^sub>4 \}" proof (rule pmf_rev_mono) fix \ assume "\ \ set_pmf (sample_pmf \)" obtain f g h where \_def: "\ = (f,g,h)" by (metis prod_cases3) assume "\ \ {\. \ E\<^sub>1 \ \ \ E\<^sub>2 \ \ \ E\<^sub>3 \ \ \ E\<^sub>4 \}" hence assms: "E\<^sub>1 (f,g,h)" "E\<^sub>2 (f,g,h)" "E\<^sub>3 (f,g,h)" "E\<^sub>4 (f,g,h)" unfolding \_def by auto define I :: "real set" where "I = {0..59/90*b}" have "p (f,g,h) \ \ (card (R f)) + \/12 * card (R f)" using assms(4) E\<^sub>4_def unfolding abs_le_iff by simp also have "... \ \(2/3*b) + 1/12* (2/3*b)" using \_lt_1 R_bound[OF assms(1,2)] by (intro add_mono rho_mono mult_mono) auto also have "... \ 3/5 * b + 1/18*b" by (intro add_mono rho_two_thirds) auto also have "... \ 59/90 * b" by simp finally have "p (f,g,h) \ 59/90 * b" by simp hence p_in_I: "p (f,g,h) \ I" unfolding I_def by simp have "\ (card (R f)) \ \(2/3 * b)" using R_bound[OF assms(1,2)] by (intro rho_mono) auto also have "... \ 3/5 * b" using rho_two_thirds by simp also have "... \ b * 59/90" by simp finally have "\ (card (R f)) \ b * 59/90" by simp moreover have "(1 - 1 / real b) powr (real (card (R f))) \ 1 powr (real (card (R f)))" using b_min by (intro powr_mono2) auto hence "\ (card (R f)) \ 0" unfolding \_def by (intro mult_nonneg_nonneg) auto ultimately have "\ (card (R f)) \ I" unfolding I_def by simp moreover have "interval I" unfolding I_def interval_def by simp moreover have "59 / 90 * b < b" using b_min by simp hence "DERIV \_inv x :> \_inv' x" if "x \ I" for x using that I_def by (intro \_inv') simp ultimately obtain \ :: real where \_def: "\ \ I" "\_inv (p(f,g,h)) - \_inv (\ (card (R f))) = (p (f,g,h) - \(card (R f))) * \_inv' \" using p_in_I MVT_interval by blast have "\\_inv(p (f,g,h)) - card (R f)\ = \\_inv(p (f,g,h)) - \_inv(\(card (R f)))\" by (subst \_inverse) simp also have "... = \(p (f,g,h) - \ (card (R f)))\ * \\_inv' \ \" using \_def(2) abs_mult by simp also have "... \ \p (f,g,h) - \ (card (R f))\ * 4" using \_def(1) I_def by (intro mult_left_mono \_inv'_bound) auto also have "... \ ( \/12 * card (R f)) * 4" using assms(4) E\<^sub>4_def by (intro mult_right_mono) auto also have "... = \/3 * card (R f)" by simp finally have b: "\\_inv(p (f,g,h)) - card (R f)\ \ \/3 * card (R f)" by simp have "\\_inv(p (f,g,h)) - X / 2 ^ (s f)\ \ \\_inv(p (f,g,h)) - card (R f)\ + \card (R f) - X / 2 ^ (s f)\" by simp also have "... \ \/3 * card (R f) + \card (R f) - X / 2 ^ (s f)\" by (intro add_mono b) auto also have "... = \/3 * \X / 2 ^ (s f) + (card (R f) - X / 2 ^ (s f))\ + \card (R f) - X / 2 ^ (s f)\" by simp also have "... \ \/3 * (\X / 2 ^ (s f)\ + \card (R f) - X / 2 ^ (s f)\) + \card (R f) - X / 2 ^ (s f)\" using \_gt_0 by (intro mult_left_mono add_mono abs_triangle_ineq) auto also have "... \ \/3 * \X / 2 ^ (s f)\ + (1+ \/3) * \card (R f) - X / 2 ^ (s f)\" using \_gt_0 \_lt_1 by (simp add:algebra_simps) also have "... \ \/3 * \X / 2 ^ s f\ + (4/3) * ( \ / 3 * real X / 2 ^ s f)" using assms(2) \_gt_0 \_lt_1 unfolding E\<^sub>2_def by (intro add_mono mult_mono) auto also have "... = (7/9) * \ * real X / 2^s f" using X_ge_1 by (subst abs_of_nonneg) auto also have "... \ 1 * \ * real X / 2^s f" using \_gt_0 by (intro mult_mono divide_right_mono) auto also have "... = \ * real X / 2^s f" by simp finally have a:"\\_inv(p (f,g,h)) - X / 2 ^ (s f)\ \ \ * X / 2 ^ (s f)" by simp have "\Y (f, g, h) - real X\ = \2 ^ (s f)\ * \\_inv(p (f,g,h)) - real X / 2 ^ (s f)\" unfolding Y_def by (subst abs_mult[symmetric]) (simp add:algebra_simps powr_add[symmetric]) also have "... \ 2 ^ (s f) * (\ * X / 2 ^ (s f))" by (intro mult_mono a) auto also have "... = \ * X" by (simp add:algebra_simps powr_add[symmetric]) finally have "\Y (f, g, h) - real X\ \ \ * X" by simp moreover have "2 powr (\log 2 (real X)\ - t f) \ 2 powr b_exp" (is "?L1 \ ?R1") proof - have "?L1 \ 2 powr (1 + log 2 (real X)- t f)" by (intro powr_mono, linarith) auto also have "... = 2 powr 1 * 2 powr (log 2 (real X)) * 2 powr (- t f)" unfolding powr_add[symmetric] by simp also have "... = 2 * (2 powr (-t f) * X)" using X_ge_1 by simp also have "... \ 2 * (b/2)" using assms(1) unfolding E\<^sub>1_def by (intro mult_left_mono) auto also have "... = b" by simp also have "... = ?R1" unfolding b_def by (simp add: powr_realpow) finally show ?thesis by simp qed hence "\log 2 (real X)\ - t f \ real b_exp" unfolding not_less[symmetric] using powr_less_mono[where x="2"] by simp hence "s f \ q_max" unfolding s_def q_max_def by (intro nat_mono) auto ultimately show "\ \ {(f, g, h). \ * X < \Y (f, g, h) - real X\ \ s f < q_max}" unfolding \_def by auto qed also have "... \ measure \ {\. \E\<^sub>1 \ \ \E\<^sub>2 \ \ \E\<^sub>3 \} + measure \ {\. E\<^sub>1 \ \ E\<^sub>2 \ \ E\<^sub>3 \ \ \E\<^sub>4 \}" by (intro pmf_add) auto also have "... \ (measure \ {\. \E\<^sub>1 \ \ \E\<^sub>2 \} + measure \ {\. E\<^sub>1 \ \ E\<^sub>2 \ \ \E\<^sub>3 \}) + 1/2^6" by (intro add_mono e_4 pmf_add) auto also have "... \ ((measure \ {\. \E\<^sub>1 \} + measure \ {\. E\<^sub>1 \ \ \E\<^sub>2 \}) + 1/2^6) + 1/2^6" by (intro add_mono e_3 pmf_add) auto also have "... \ ((1/2^6 + 1/2^6) + 1/2^6) + 1/2^6" by (intro add_mono e_2 e_1) auto also have "... = ?R" by simp finally show ?thesis by simp qed end end diff --git a/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Balls_and_Bins.thy b/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Balls_and_Bins.thy --- a/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Balls_and_Bins.thy +++ b/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Balls_and_Bins.thy @@ -1,1926 +1,1926 @@ section \Balls and Bins\ text \The balls and bins model describes the probability space of throwing r balls into b bins. This section derives the expected number of bins hit by at least one ball, as well as the variance in the case that each ball is thrown independently. Further, using an approximation argument it is then possible to derive bounds for the same measures in the case when the balls are being thrown only $k$-wise independently. The proofs follow the reasoning described in \cite[\S A.1]{kane2010} but improve on the constants, as well as constraints.\ theory Distributed_Distinct_Elements_Balls_and_Bins imports Distributed_Distinct_Elements_Preliminary Discrete_Summation.Factorials "HOL-Combinatorics.Stirling" "HOL-Computational_Algebra.Polynomial" "HOL-Decision_Procs.Approximation" begin hide_fact "Henstock_Kurzweil_Integration.integral_sum" hide_fact "Henstock_Kurzweil_Integration.integral_mult_right" hide_fact "Henstock_Kurzweil_Integration.integral_nonneg" hide_fact "Henstock_Kurzweil_Integration.integral_cong" unbundle intro_cong_syntax lemma sum_power_distrib: fixes f :: "'a \ real" assumes "finite R" shows "(\i\R. f i) ^ s = (\xs | set xs \ R \ length xs = s. (\x \ xs. f x))" proof (induction s) case 0 have "{xs. xs = [] \ set xs \ R} = {[]}" by (auto simp add:set_eq_iff) then show ?case by simp next case (Suc s) have a: "(\i\R. (#) i ` {xs. set xs \ R \ length xs = s}) = {xs. set xs \ R \ length xs = Suc s}" by (subst lists_length_Suc_eq) auto have "sum f R ^ Suc s = (sum f R) * (sum f R)^s" by simp also have "... = (sum f R) * (\xs | set xs \ R \ length xs = s. (\x \ xs. f x))" using Suc by simp also have "... = (\i \ R. (\xs | set xs \ R \ length xs = s. (\x \ i#xs. f x)))" by (subst sum_product) simp also have "... = (\i \ R. (\xs \ (\xs. i#xs) ` {xs. set xs \ R \ length xs = s}. (\x \ xs. f x)))" by (subst sum.reindex) (auto) also have "... = (\xs\(\i\R. (#) i ` {xs. set xs \ R \ length xs = s}). (\x \ xs. f x))" by (intro sum.UNION_disjoint[symmetric] assms ballI finite_imageI finite_lists_length_eq) auto also have "... = (\xs| set xs \ R \ length xs = Suc s. (\x \ xs. f x))" by (intro sum.cong a) auto finally show ?case by simp qed lemma sum_telescope_eq: fixes f :: "nat \ 'a :: {comm_ring_1}" shows "(\k\{Suc m..n}. f k - f (k - 1)) = of_bool(m \ n) *(f n - f m)" by (cases "m \ n", subst sum_telescope'', auto) text \An improved version of @{thm [source] diff_power_eq_sum}.\ lemma power_diff_sum: fixes a b :: "'a :: {comm_ring_1,power}" shows "a^k - b^k = (a-b) * (\i = 0.. b" assumes "b \ 0" shows "a^k - b^k \ (a-b) * k * a^(k-1)" proof - have "a^k - b^k = (a-b) * (\i = 0.. (a-b) * (\i = 0.. b" assumes "b \ 0" shows "a^k - b^k \ (a-b) * k * b^(k-1)" proof - have "(a-b) * k * b^(k-1) = (a-b) * (\i=0.. (a-b)* (\i=0.. j \ R. of_bool(f j)) = (of_bool(\j \ R. f j) :: real)" using assms by (induction R rule:finite_induct) auto text \Additional results about falling factorials:\ lemma ffact_nonneg: fixes x :: real assumes "k - 1 \ x" shows "ffact k x \ 0" using assms unfolding prod_ffact[symmetric] by (intro prod_nonneg ballI) simp lemma ffact_pos: fixes x :: real assumes "k - 1 < x" shows "ffact k x > 0" using assms unfolding prod_ffact[symmetric] by (intro prod_pos ballI) simp lemma ffact_mono: fixes x y :: real assumes "k-1 \ x" "x \ y" shows "ffact k x \ ffact k y" using assms unfolding prod_ffact[symmetric] by (intro prod_mono) auto lemma ffact_of_nat_nonneg: fixes x :: "'a :: {comm_ring_1, linordered_nonzero_semiring}" assumes "x \ \" shows "ffact k x \ 0" proof - obtain y where y_def: "x = of_nat y" using assms(1) Nats_cases by auto have "(0::'a) \ of_nat (ffact k y)" by simp also have "... = ffact k x" by (simp add:of_nat_ffact y_def) finally show ?thesis by simp qed lemma ffact_suc_diff: fixes x :: "('a :: comm_ring_1)" shows "ffact k x - ffact k (x-1) = of_nat k * ffact (k-1) (x-1)" (is "?L = ?R") proof (cases "k") case 0 then show ?thesis by simp next case (Suc n) hence "?L = ffact (Suc n) x - ffact (Suc n) (x-1)" by simp also have "... = x * ffact n (x-1) - ((x-1)-of_nat n) * ffact n (x-1)" by (subst (1) ffact_Suc, simp add: ffact_Suc_rev) also have "... = of_nat (Suc n) * ffact n (x-1)" by (simp add:algebra_simps) also have "... = of_nat k * ffact (k-1) (x-1)" using Suc by simp finally show ?thesis by simp qed lemma ffact_bound: "ffact k (n::nat) \ n^k" proof - have "ffact k n = (\i=0.. (\i=0.. :: real assumes "\ \ {0..1}" defines "p \ binomial_pmf n \" shows "(\\. ffact s (real \) \p) = ffact s (real n) * \^s" (is "?L = ?R") proof (cases "s \ n") case True have "?L = (\k\n. (real (n choose k) * \ ^ k * (1 - \) ^ (n - k)) * real (ffact s k))" unfolding p_def using assms by (subst expectation_binomial_pmf') (auto simp add:of_nat_ffact) also have "... = (\k \ {0+s..(n-s)+s}. (real (n choose k) * \ ^ k * (1 - \) ^ (n - k)) * ffact s k)" using True ffact_nat_triv by (intro sum.mono_neutral_cong_right) auto also have "... = (\k=0..n-s. \^s * real (n choose (k+s)) * \^k * (1-\)^(n-(k+s)) *ffact s (k+s))" by (subst sum.atLeastAtMost_shift_bounds, simp add:algebra_simps power_add) also have "... = \^s * (\k\n-s. real (n choose (k+s))*ffact s (k+s)*\^k*(1-\)^((n-s)-k))" using atMost_atLeast0 by (simp add: sum_distrib_left algebra_simps cong:sum.cong) also have "... = \^s * (\k\n-s. real (n choose (k+s))*fact (k+s) / fact k * \^k*(1-\)^((n-s)-k))" using real_of_nat_div[OF fact_dvd[OF le_add1]] by (subst fact_div_fact_ffact_nat[symmetric], auto) also have "... = \^s * (\k\n-s. (fact n / fact (n-s)) * fact (n-s) / (fact ((n-s)-k) * fact k) * \^k*(1-\)^((n-s)-k))" using True by (intro arg_cong2[where f="(*)"] sum.cong) (auto simp add: binomial_fact algebra_simps) also have "... = \^s * (fact n / fact (n - s)) * (\k\n-s. fact (n-s) / (fact ((n-s)-k) * fact k) * \^k*(1-\)^((n-s)-k))" by (simp add:sum_distrib_left algebra_simps) also have "... = \^s * (fact n / fact (n - s)) * (\k\n-s. ((n-s) choose k) * \^k*(1-\)^((n-s)-k))" using True by (intro_cong "[\\<^sub>2(*)]" more: sum.cong) (auto simp add: binomial_fact) also have "... = \^s * real (fact n div fact (n - s)) * (\+(1-\))^(n-s)" using True real_of_nat_div[OF fact_dvd] by (subst binomial_ring, simp) also have "... = \^s * real (ffact s n)" by (subst fact_div_fact_ffact_nat[OF True], simp) also have "... = ?R" by (subst of_nat_ffact, simp) finally show ?thesis by simp next case False have "?L = (\k\n. (real (n choose k) * \ ^ k * (1 - \) ^ (n - k)) * real (ffact s k))" unfolding p_def using assms by (subst expectation_binomial_pmf') (auto simp add:of_nat_ffact) also have "... = (\k\n. (real (n choose k) * \ ^ k * (1 - \) ^ (n - k)) * real 0)" using False by (intro_cong "[\\<^sub>2(*),\\<^sub>1 of_nat]" more: sum.cong ffact_nat_triv) auto also have "... = 0" by simp also have "... = real (ffact s n) * \^s" using False by (subst ffact_nat_triv, auto) also have "... = ?R" by (subst of_nat_ffact, simp) finally show ?thesis by simp qed text \The following describes polynomials of a given maximal degree as a subset of the functions, similar to the subsets @{term "\"} or @{term "\"} as subsets of larger number classes.\ definition Polynomials ("\") where "Polynomials k = {f. \p. f = poly p \ degree p \ k}" lemma Polynomials_mono: assumes "s \ t" shows "\ s \ \ t" using assms unfolding Polynomials_def by auto lemma Polynomials_addI: assumes "f \ \ k" "g \ \ k" shows "(\\. f \ + g \) \ \ k" proof - obtain pf pg where fg_def: "f = poly pf" "degree pf \ k" "g = poly pg" "degree pg \ k" using assms unfolding Polynomials_def by blast hence "degree (pf + pg) \ k" "(\x. f x + g x) = poly (pf + pg)" using degree_add_le by auto thus ?thesis unfolding Polynomials_def by auto qed lemma Polynomials_diffI: fixes f g :: "'a :: comm_ring \ 'a" assumes "f \ \ k" "g \ \ k" shows "(\x. f x - g x) \ \ k" proof - obtain pf pg where fg_def: "f = poly pf" "degree pf \ k" "g = poly pg" "degree pg \ k" using assms unfolding Polynomials_def by blast hence "degree (pf - pg) \ k" "(\x. f x - g x) = poly (pf - pg)" using degree_diff_le by auto thus ?thesis unfolding Polynomials_def by auto qed lemma Polynomials_idI: "(\x. x) \ (\ 1 :: ('a::comm_ring_1 \ 'a) set)" proof - have "(\x. x) = poly [: 0,(1::'a) :]" by (intro ext, auto) also have "... \ \ 1" unfolding Polynomials_def by auto finally show ?thesis by simp qed lemma Polynomials_constI: "(\x. c) \ \ k" proof - have "(\x. c) = poly [: c :]" by (intro ext, simp) also have "... \ \ k" unfolding Polynomials_def by auto finally show ?thesis by simp qed lemma Polynomials_multI: fixes f g :: "'a :: {comm_ring} \ 'a" assumes "f \ \ s" "g \ \ t" shows "(\x. f x * g x) \ \ (s+t)" proof - obtain pf pg where xy_def: "f = poly pf" "degree pf \ s" "g = poly pg" "degree pg \ t" using assms unfolding Polynomials_def by blast have "degree (pf * pg) \ degree pf + degree pg" by (intro degree_mult_le) also have "... \ s + t" using xy_def by (intro add_mono) auto finally have "degree (pf * pg) \ s+t" by simp moreover have "(\x. f x * g x) = poly (pf * pg)" using xy_def by auto ultimately show ?thesis unfolding Polynomials_def by auto qed lemma Polynomials_composeI: fixes f g :: "'a :: {comm_semiring_0, semiring_no_zero_divisors} \ 'a" assumes "f \ \ s" "g \ \ t" shows "(\x. f (g x)) \ \ (s*t)" proof - obtain pf pg where xy_def: "f = poly pf" "degree pf \ s" "g = poly pg" "degree pg \ t" using assms unfolding Polynomials_def by blast have "degree (pf \\<^sub>p pg) = degree pf * degree pg" by (intro degree_pcompose) also have "... \ s * t" using xy_def by (intro mult_mono) auto finally have "degree (pf \\<^sub>p pg) \ s * t" by simp moreover have "(\x. f (g x)) = poly (pf \\<^sub>p pg)" unfolding xy_def by (intro ext poly_pcompose[symmetric]) ultimately show ?thesis unfolding Polynomials_def by auto qed lemma Polynomials_const_left_multI: fixes c :: "'a :: {comm_ring}" assumes "f \ \ k" shows "(\x. c * f x) \ \ k" proof - have "(\x. c * f x) \ \ (0+k)" by (intro Polynomials_multI Polynomials_constI assms) thus ?thesis by simp qed lemma Polynomials_const_right_multI: fixes c :: "'a :: {comm_ring}" assumes "f \ \ k" shows "(\x. f x * c) \ \ k" proof - have "(\x. f x * c) \ \ (k+0)" by (intro Polynomials_multI Polynomials_constI assms) thus ?thesis by simp qed lemma Polynomials_const_divI: fixes c :: "'a :: {field}" assumes "f \ \ k" shows "(\x. f x / c) \ \ k" proof - have "(\x. f x * (1/c)) \ \ (k+0)" by (intro Polynomials_multI Polynomials_constI assms) thus ?thesis by simp qed lemma Polynomials_ffact: "(\x. ffact s (x - y)) \ (\ s :: ('a :: comm_ring_1 \ 'a) set)" proof (induction s arbitrary: y) case 0 then show ?case using Polynomials_constI[where c="1"] by simp next case (Suc s) have "(\(x :: 'a). ffact (Suc s) (x-y)) = (\x. (x-y) * ffact s (x - (y+1)))" by (simp add: ffact_Suc algebra_simps) also have "... \ \ (1+s)" by (intro Polynomials_multI Suc Polynomials_diffI Polynomials_idI Polynomials_constI) finally show ?case by simp qed lemmas Polynomials_intros = Polynomials_const_divI Polynomials_composeI Polynomials_const_left_multI Polynomials_const_right_multI Polynomials_multI Polynomials_addI Polynomials_diffI Polynomials_idI Polynomials_constI Polynomials_ffact definition C\<^sub>2 :: real where "C\<^sub>2 = 7.5" definition C\<^sub>3 :: real where "C\<^sub>3 = 16" text \A locale fixing the sets of balls and bins\ locale balls_and_bins_abs = fixes R :: "'a set" and B :: "'b set" assumes fin_B: "finite B" and B_ne: "B \ {}" assumes fin_R: "finite R" begin text \Independent balls and bins space:\ definition \ where "\ = prod_pmf R (\_. pmf_of_set B)" lemma set_pmf_\: "set_pmf \ = R \\<^sub>E B" unfolding \_def set_prod_pmf[OF fin_R] by (simp add:comp_def set_pmf_of_set[OF B_ne fin_B]) lemma card_B_gt_0: "card B > 0" using B_ne fin_B by auto lemma card_B_ge_1: "card B \ 1" using card_B_gt_0 by simp definition "Z j \ = real (card {i. i \ R \ \ i = (j::'b)})" definition "Y \ = real (card (\ ` R))" definition "\ = real (card B) * (1 - (1-1/real (card B))^card R)" text \Factorial moments for the random variable describing the number of times a bin will be hit:\ lemma fact_moment_balls_and_bins: assumes "J \ B" "J \ {}" shows "(\\. ffact s (\j \ J. Z j \) \\) = ffact s (real (card R)) * (real (card J) / real (card B))^s" (is "?L = ?R") proof - let ?\ = "real (card J) / real (card B)" let ?q = "binomial_pmf (card R) ?\" let ?Y = "(\\. card {r \ R. \ r \ J})" have fin_J: "finite J" using finite_subset assms(1) fin_B by auto have Z_sum_eq: "(\j \ J. Z j \) = real (?Y \)" for \ proof - have "?Y \ = card (\j \ J. {r \ R. \ r= j})" by (intro arg_cong[where f="card"]) auto also have "... = (\i\J. card {r \ R. \ r = i})" using fin_R fin_J by (intro card_UN_disjoint) auto finally have "?Y \ = (\j \ J. card {r \ R. \ r = j})" by simp thus ?thesis unfolding Z_def of_nat_sum[symmetric] by simp qed have card_J: "card J \ card B" using assms(1) fin_B card_mono by auto have \_range: "?\ \ 0" "?\ \ 1" using card_J card_B_gt_0 by auto have "pmf (map_pmf (\\. \ \ J) (pmf_of_set B)) x = pmf (bernoulli_pmf ?\) x" (is "?L1=?R1") for x proof - have "?L1 = real (card (B \ {\. (\ \ J) = x})) / real (card B)" using B_ne fin_B by (simp add:pmf_map measure_pmf_of_set vimage_def) also have "... = (if x then (card J) else (card (B - J))) / real (card B)" using Int_absorb1[OF assms(1)] by (auto simp add:Diff_eq Int_def) also have "... = (if x then (card J) / card B else (real (card B) - card J) / real (card B))" using card_J fin_J assms(1) by (simp add: of_nat_diff card_Diff_subset) also have "... = (if x then ?\ else (1 - ?\))" using card_B_gt_0 by (simp add:divide_simps) also have "... = ?R1" using \_range by auto finally show ?thesis by simp qed hence c:"map_pmf (\\. \ \ J) (pmf_of_set B) = bernoulli_pmf ?\" by (intro pmf_eqI) simp have "map_pmf (\\. \r \ R. \ r \ J) \ = prod_pmf R (\_. (map_pmf (\\. \ \ J) (pmf_of_set B)))" unfolding map_pmf_def \_def restrict_def using fin_R by (subst Pi_pmf_bind[where d'="undefined"]) auto also have "... = prod_pmf R (\_. bernoulli_pmf ?\)" unfolding c by simp finally have b:"map_pmf (\\. \r \ R. \ r \ J) \ = prod_pmf R (\_. bernoulli_pmf ?\)" by simp have "map_pmf ?Y \ = map_pmf ((\\. card {r \ R. \ r}) \ (\\. \r\R. \ r \ J)) \" unfolding comp_def by (intro map_pmf_cong arg_cong[where f="card"]) (auto simp add:comp_def) also have "... = (map_pmf (\\. card {r \ R. \ r}) \ map_pmf (\\. \r \ R. \ r \ J)) \" by (subst map_pmf_compose[symmetric]) auto also have "... = map_pmf (\\. card {r \ R. \ r}) (prod_pmf R (\_. (bernoulli_pmf ?\)))" unfolding comp_def b by simp also have "... = ?q" using \_range by (intro binomial_pmf_altdef'[symmetric] fin_R) auto finally have a:"map_pmf ?Y \ =?q" by simp have "?L = (\\. ffact s (real (?Y \)) \\)" unfolding Z_sum_eq by simp also have "... = (\\. ffact s (real \) \(map_pmf ?Y \))" by simp also have "... = (\\. ffact s (real \) \?q)" unfolding a by simp also have "... = ?R" using \_range by (subst fact_moment_binomial, auto) finally show ?thesis by simp qed text \Expectation and variance for the number of distinct bins that are hit by at least one ball in the fully independent model. The result for the variance is improved by a factor of $4$ w.r.t. the paper.\ lemma shows exp_balls_and_bins: "measure_pmf.expectation \ Y = \" (is "?AL = ?AR") and var_balls_and_bins: "measure_pmf.variance \ Y \ card R * (real (card R) - 1) / card B" (is "?BL \ ?BR") proof - let ?b = "real (card B)" let ?r = "card R" define Z :: "'b \ ('a \ 'b) \ real" where "Z = (\i \. of_bool(i \ \ ` R))" define \ where "\ = (1 - 1 / ?b)^?r" define \ where "\ = (1 - 2 / ?b)^?r" have "card (B \ B \ {x. fst x = snd x}) = card ((\x. (x,x)) ` B)" by (intro arg_cong[where f="card"]) auto also have "... = card B" by (intro card_image, simp add:inj_on_def) finally have d: "card (B \ B \ {x. fst x = snd x}) = card B" by simp hence count_1: "real (card (B \ B \ {x. fst x = snd x})) = card B" by simp have "card B + card (B \ B \ -{x. fst x = snd x}) = card (B \ B \ {x. fst x = snd x}) + card (B \ B \ -{x. fst x = snd x})" by (subst d) simp also have "... = card ((B \ B \ {x. fst x = snd x}) \ (B \ B \ -{x. fst x = snd x}))" using finite_subset[OF _ finite_cartesian_product[OF fin_B fin_B]] by (intro card_Un_disjoint[symmetric]) auto also have "... = card (B \ B)" by (intro arg_cong[where f="card"]) auto also have "... = card B^2" unfolding card_cartesian_product by (simp add:power2_eq_square) finally have "card B + card (B \ B \ -{x. fst x = snd x}) = card B^2" by simp hence count_2: "real (card (B \ B \ -{x. fst x = snd x})) = real (card B)^2 - card B" by (simp add:algebra_simps flip: of_nat_add of_nat_power) hence "finite (set_pmf \)" unfolding set_pmf_\ using fin_R fin_B by (auto intro!:finite_PiE) hence int: "integrable (measure_pmf \) f" for f :: "('a \ 'b) \ real" by (intro integrable_measure_pmf_finite) simp have a:"prob_space.indep_vars (measure_pmf \) (\i. discrete) (\x \. \ x) R" unfolding \_def using indep_vars_Pi_pmf[OF fin_R] by metis have b: "(\\. of_bool (\ ` R \ A) \\) = (real (card (B \ A)) / real (card B))^card R" (is "?L = ?R") for A proof - have "?L = (\\. (\ j \ R. of_bool(\ j \ A)) \\)" by (intro Bochner_Integration.integral_cong ext) (auto simp add: of_bool_prod[OF fin_R]) also have "... = (\ j \ R. (\\. of_bool(\ j \ A) \\))" using fin_R by (intro prob_space.indep_vars_lebesgue_integral[OF prob_space_measure_pmf] int prob_space.indep_vars_compose2[OF prob_space_measure_pmf a]) auto also have "... = (\ j \ R. (\\. of_bool(\ \ A) \(map_pmf (\\. \ j) \)))" by simp also have "... = (\ j \ R. (\\. of_bool(\ \ A) \(pmf_of_set B)))" unfolding \_def by (subst Pi_pmf_component[OF fin_R]) simp also have "... = ((\\\B. of_bool (\ \ A)) / real (card B)) ^ card R" by (simp add: integral_pmf_of_set[OF B_ne fin_B]) also have "... = ?R" unfolding of_bool_def sum.If_cases[OF fin_B] by simp finally show ?thesis by simp qed have Z_exp: "(\\. Z i \ \\) = \" if "i \ B" for i proof - have "real (card (B \ -{i})) = real (card (B - {i}))" by (intro_cong "[\\<^sub>1 card,\\<^sub>1 of_nat]") auto also have "... = real (card B - card {i})" using that by (subst card_Diff_subset) auto also have "... = real (card B) - real (card {i})" using fin_B that by (intro of_nat_diff card_mono) auto finally have c: "real (card (B \ -{i})) = real (card B) - 1" by simp have "(\\. Z i \ \\) = (\\. of_bool(\ ` R \ - {i}) \\)" unfolding Z_def by simp also have "... = (real (card (B \ -{i})) / real (card B))^card R" by (intro b) also have "... = ((real (card B) -1) / real (card B))^card R" by (subst c) simp also have "... = \" unfolding \_def using card_B_gt_0 by (simp add:divide_eq_eq diff_divide_distrib) finally show ?thesis by simp qed have Z_prod_exp: "(\\. Z i \ * Z j \ \\) = (if i = j then \ else \)" if "i \ B" "j \ B" for i j proof - have "real (card (B \ -{i,j})) = real (card (B - {i,j}))" by (intro_cong "[\\<^sub>1 card,\\<^sub>1 of_nat]") auto also have "... = real (card B - card {i,j})" using that by (subst card_Diff_subset) auto also have "... = real (card B) - real (card {i,j})" using fin_B that by (intro of_nat_diff card_mono) auto finally have c: "real (card (B \ -{i,j})) = real (card B) - card {i,j}" by simp have "(\\. Z i \ * Z j \ \\) = (\\. of_bool(\ ` R \ - {i,j}) \\)" unfolding Z_def of_bool_conj[symmetric] by (intro integral_cong ext) auto also have "... = (real (card (B \ -{i,j})) / real (card B))^card R" by (intro b) also have "... = ((real (card B) - card {i,j}) / real (card B))^card R" by (subst c) simp also have "... = (if i = j then \ else \)" unfolding \_def \_def using card_B_gt_0 by (simp add:divide_eq_eq diff_divide_distrib) finally show ?thesis by simp qed have Y_eq: "Y \ = (\i \ B. 1 - Z i \)" if "\ \ set_pmf \" for \ proof - have "set_pmf \ \ Pi R (\_. B)" using set_pmf_\ by (simp add:PiE_def) hence "\ ` R \ B" using that by auto hence "Y \ = card (B \ \ ` R)" unfolding Y_def using Int_absorb1 by metis also have "... = (\ i \ B. of_bool(i \ \ ` R))" unfolding of_bool_def sum.If_cases[OF fin_B] by(simp) also have "... = (\ i \ B. 1 - Z i \)" unfolding Z_def by (intro sum.cong) (auto simp add:of_bool_def) finally show "Y \ = (\i \ B. 1 - Z i \)" by simp qed have Y_sq_eq: "(Y \)\<^sup>2 = (\(i,j) \ B \ B. 1 - Z i \ - Z j \ + Z i \ * Z j \)" if "\ \ set_pmf \" for \ unfolding Y_eq[OF that] power2_eq_square sum_product sum.cartesian_product by (intro sum.cong) (auto simp add:algebra_simps) have "measure_pmf.expectation \ Y = (\\. (\i \ B. 1 - Z i \) \\)" using Y_eq by (intro integral_cong_AE AE_pmfI) auto also have "... = (\i \ B. 1 - (\\. Z i \ \\))" using int by simp also have "... = ?b * (1 - \)" using Z_exp by simp also have "... = ?AR" unfolding \_def \_def by simp finally show "?AL = ?AR" by simp have "measure_pmf.variance \ Y = (\\. Y \^2 \\) - (\\. Y \ \\)^2" using int by (subst measure_pmf.variance_eq) auto also have "... = (\\. (\i \ B \ B. 1 - Z (fst i) \ - Z (snd i) \ + Z (fst i) \ * Z (snd i) \) \\) - (\\. (\i \ B. 1 - Z i \) \\)^2" using Y_eq Y_sq_eq by (intro_cong "[\\<^sub>2(-),\\<^sub>2 power]" more: integral_cong_AE AE_pmfI) (auto simp add:case_prod_beta) also have "... = (\i \ B \ B. (\\. (1 - Z (fst i) \ - Z (snd i) \ + Z (fst i) \ * Z (snd i) \) \\)) - (\i \ B. (\\. (1 - Z i \) \\))^2" by (intro_cong "[\\<^sub>2 (-), \\<^sub>2 power]" more: integral_sum int) also have "... = (\i \ B \ B. (\\. (1 - Z (fst i) \ - Z (snd i) \ + Z (fst i) \ * Z (snd i) \) \\)) - (\i \ B \ B. (\\. (1 - Z (fst i) \) \\) * (\\. (1 - Z (snd i) \) \\))" unfolding power2_eq_square sum_product sum.cartesian_product by (simp add:case_prod_beta) also have "... = (\(i,j) \ B \ B. (\\. (1 - Z i \ - Z j \ + Z i \ * Z j \) \\) - (\\. (1 - Z i \) \\) * (\\. (1 - Z j \) \\))" by (subst sum_subtractf[symmetric], simp add:case_prod_beta) also have "... = (\(i,j) \ B \ B. (\\. Z i \ * Z j \ \\) -(\\. Z i \ \\) * (\ \. Z j \ \\))" using int by (intro sum.cong refl) (simp add:algebra_simps case_prod_beta) also have "... = (\i \ B \ B. (if fst i = snd i then \ - \^2 else \ - \^2))" by (intro sum.cong refl) (simp add:Z_exp Z_prod_exp mem_Times_iff case_prod_beta power2_eq_square) also have "... = ?b * (\ - \\<^sup>2) + (?b^2 - card B) * (\ - \\<^sup>2)" using count_1 count_2 finite_cartesian_product fin_B by (subst sum.If_cases) auto also have "... = ?b^2 * (\ - \\<^sup>2) + ?b * (\ - \)" by (simp add:algebra_simps) also have "... = ?b * ((1-1/?b)^?r - (1-2/?b)^?r) - ?b^2 * (((1-1/?b)^2)^?r - (1-2/?b)^?r)" unfolding \_def \_def by (simp add: power_mult[symmetric] algebra_simps) also have "... \ card R * (real (card R) - 1)/ card B" (is "?L \ ?R") proof (cases "?b \ 2") case True have "?L \ ?b * (((1 - 1 /?b) - (1-2 /?b)) * ?r * (1-1/?b)^(?r - 1)) - ?b^2 * ((((1-1 /?b)^2) - ((1 - 2 /?b))) * ?r * ((1-2/?b))^(?r - 1))" using True by (intro diff_mono mult_left_mono power_diff_est_2 power_diff_est divide_right_mono) (auto simp add:power2_eq_square algebra_simps) also have "... = ?b * ((1/?b) * ?r * (1-1/?b)^(?r-1)) - ?b^2*((1/?b^2)*?r*((1-2/?b))^(?r-1))" by (intro arg_cong2[where f="(-)"] arg_cong2[where f="(*)"] refl) (auto simp add:algebra_simps power2_eq_square) also have "... = ?r * ((1-1/?b)^(?r - 1) - ((1-2/?b))^(?r - 1))" by (simp add:algebra_simps) also have "... \ ?r * (((1-1/?b) - (1-2/?b)) * (?r - 1) * (1-1/?b)^(?r -1 -1))" using True by (intro mult_left_mono power_diff_est) (auto simp add:algebra_simps) also have "... \ ?r * ((1/?b) * (?r - 1) * 1^(?r - 1-1))" using True by (intro mult_left_mono mult_mono power_mono) auto also have "... = ?R" using card_B_gt_0 by auto finally show "?L \ ?R" by simp next case False hence "?b = 1" using card_B_ge_1 by simp thus "?L \ ?R" by (cases "card R =0") auto qed finally show "measure_pmf.variance \ Y \ card R * (real (card R) - 1)/ card B" by simp qed definition "lim_balls_and_bins k p = ( prob_space.k_wise_indep_vars (measure_pmf p) k (\_. discrete) (\x \. \ x) R \ (\x. x \ R \ map_pmf (\\. \ x) p = pmf_of_set B))" lemma indep: assumes "lim_balls_and_bins k p" shows "prob_space.k_wise_indep_vars (measure_pmf p) k (\_. discrete) (\x \. \ x) R" using assms lim_balls_and_bins_def by simp lemma ran: assumes "lim_balls_and_bins k p" "x \ R" shows "map_pmf (\\. \ x) p = pmf_of_set B" using assms lim_balls_and_bins_def by simp lemma Z_integrable: fixes f :: "real \ real" assumes "lim_balls_and_bins k p" shows "integrable p (\\. f (Z i \) )" unfolding Z_def using fin_R card_mono by (intro integrable_pmf_iff_bounded[where C="Max (abs ` f ` real ` {..card R})"]) fastforce+ lemma Z_any_integrable_2: fixes f :: "real \ real" assumes "lim_balls_and_bins k p" shows "integrable p (\\. f (Z i \ + Z j \))" proof - have q:"real (card A) + real (card B) \ real ` {..2 * card R}" if "A \ R" "B \ R" for A B proof - have "card A + card B \ card R + card R" by (intro add_mono card_mono fin_R that) also have "... = 2 * card R" by simp finally show ?thesis by force qed thus ?thesis unfolding Z_def using fin_R card_mono abs_triangle_ineq by (intro integrable_pmf_iff_bounded[where C="Max (abs ` f ` real ` {..2*card R})"] Max_ge finite_imageI imageI) auto qed lemma hit_count_prod_exp: assumes "j1 \ B" "j2 \ B" "s+t \ k" assumes "lim_balls_and_bins k p" defines "L \ {(xs,ys). set xs \ R \ set ys \ R \ (set xs \ set ys = {} \ j1 = j2) \ length xs = s \ length ys = t}" shows "(\\. Z j1 \^s * Z j2 \^t \p) = (\(xs,ys) \ L. (1/real (card B))^(card (set xs \ set ys)))" (is "?L = ?R") proof - define W1 :: "'a \ ('a \ 'b) \ real" where "W1 = (\i \. of_bool (\ i = j1) :: real)" define W2 :: "'a \ ('a \ 'b) \ real" where "W2 = (\i \. of_bool (\ i = j2) :: real)" define \ :: "'a list \ 'a list \ 'a \ 'b" where "\ = (\l x. if x \ set (fst l) then j1 else j2)" have \_check_1: "\ l x = j1" if "x \ set (fst l)" and "l \ L" for x l using that unfolding \_def L_def by auto have \_check_2: "\ l x = j2" if "x \ set (snd l)" and "l \ L" for x l using that unfolding \_def L_def by auto have \_check_3: "\ l x \ B" for x l using assms(1,2) unfolding \_def by simp have Z1_eq: "Z j1 \ = (\i \ R. W1 i \)" for \ using fin_R unfolding Z_def W1_def by (simp add:of_bool_def sum.If_cases Int_def) have Z2_eq: "Z j2 \ = (\i \ R. W2 i \)" for \ using fin_R unfolding Z_def W2_def by (simp add:of_bool_def sum.If_cases Int_def) define \ where "\ = 1 / real (card B)" have a: "(\\. (\x\a. W1 x \) * (\y\b. W2 y \) \p) = 0" (is "?L1 = 0") if "x \ set a \ set b" "j1 \ j2" "length a = s" "length b = t" for x a b proof - have "(\x\a. W1 x \) * (\y\b. W2 y \) = 0" for \ proof - have "W1 x \ = 0 \ W2 x \ = 0" unfolding W1_def W2_def using that by simp hence "(\x\a. W1 x \) = 0 \ (\y\b. W2 y \) = 0" unfolding prod_list_zero_iff using that(1) by auto thus ?thesis by simp qed hence "?L1 = (\\. 0 \p)" by (intro arg_cong2[where f="measure_pmf.expectation"]) auto also have "... = 0" by simp finally show ?thesis by simp qed have b: "prob_space.indep_vars p (\_. discrete) (\i \. \ i) (set (fst x) \ set (snd x))" if "x \ L" for x proof - have "card (set (fst x) \ set (snd x)) \ card (set (fst x)) + card (set (snd x))" by (intro card_Un_le) also have "... \ length (fst x) + length (snd x)" by (intro add_mono card_length) also have "... = s + t" using that L_def by auto also have "... \ k" using assms(3) by simp finally have "card (set (fst x) \ set (snd x)) \ k" by simp moreover have "set (fst x) \ set (snd x) \ R" using that L_def by auto ultimately show ?thesis by (intro prob_space.k_wise_indep_vars_subset[OF prob_space_measure_pmf indep[OF assms(4)]]) auto qed have c: "(\\. of_bool (\ x = z) \p) = \" (is "?L1 = _") if "z \ B" "x \ R" for x z proof - have "?L1 = (\\. indicator {\. \ x = z} \ \p)" unfolding indicator_def by simp also have "... = measure p {\. \ x = z}" by simp also have "... = measure (map_pmf (\\. \ x) p) {z}" by (subst measure_map_pmf) (simp add:vimage_def) also have "... = measure (pmf_of_set B) {z}" using that by (subst ran[OF assms(4)]) auto also have "... = 1/card B" using fin_B that by (subst measure_pmf_of_set) auto also have "... = \" unfolding \_def by simp finally show ?thesis by simp qed have d: "abs x \ 1 \ abs y \ 1 \ abs (x*y) \ 1" for x y :: real by (simp add:abs_mult mult_le_one) have e:"(\x. x \ set xs \ abs x \1) \ abs(prod_list xs) \ 1 " for xs :: "real list" using d by (induction xs, simp, simp) have "?L = (\\. (\j \ R. W1 j \)^s * (\j \ R. W2 j \)^t \p)" unfolding Z1_eq Z2_eq by simp also have "... = (\\. (\xs | set xs \ R \ length xs = s. (\x \ xs. W1 x \)) * (\ys | set ys \ R \ length ys = t. (\y \ ys. W2 y \)) \p)" unfolding sum_power_distrib[OF fin_R] by simp also have "... = (\\. (\l\{xs. set xs \ R \ length xs = s} \ {ys. set ys \ R \ length ys = t}. (\x\fst l. W1 x \) * (\y\snd l. W2 y \)) \p)" by (intro arg_cong[where f="integral\<^sup>L p"]) (simp add: sum_product sum.cartesian_product case_prod_beta) also have "... = (\l\{xs. set xs \ R \ length xs = s} \ {ys. set ys \ R \ length ys = t}. (\\. (\x\fst l. W1 x \) * (\y\snd l. W2 y \) \p))" unfolding W1_def W2_def by (intro integral_sum integrable_pmf_iff_bounded[where C="1"] d e) auto also have "... = (\l\ L. (\\. (\x\fst l. W1 x \) * (\y\snd l. W2 y \) \p))" unfolding L_def using a by (intro sum.mono_neutral_right finite_cartesian_product finite_lists_length_eq fin_R) auto also have "... = (\l\ L. (\\. (\x\fst l. of_bool(\ x = \ l x)) * (\y\snd l. of_bool(\ y = \ l y)) \p))" unfolding W1_def W2_def using \_check_1 \_check_2 by (intro sum.cong arg_cong[where f="integral\<^sup>L p"] ext arg_cong2[where f="(*)"] arg_cong[where f="prod_list"]) auto also have "... = (\l\ L. (\\. (\x\(fst l@snd l). of_bool(\ x = \ l x))\ p))" by simp also have "... = (\l\ L. (\\. (\x \ set (fst l@snd l). of_bool(\ x = \ l x)^count_list (fst l@snd l) x) \ p))" unfolding prod_list_eval by simp also have "... = (\l\ L. (\\. (\x \ set (fst l) \ set (snd l). of_bool(\ x = \ l x)^count_list (fst l@snd l) x) \ p))" by simp also have "... = (\l\ L. (\\. (\x \ set (fst l) \ set (snd l). of_bool(\ x = \ l x)) \ p))" using count_list_gr_1 by (intro sum.cong arg_cong[where f="integral\<^sup>L p"] ext prod.cong) force+ also have "... = (\l\ L. (\x \ set (fst l) \ set (snd l). (\\. of_bool(\ x = \ l x) \ p)))" by (intro sum.cong prob_space.indep_vars_lebesgue_integral[OF prob_space_measure_pmf] integrable_pmf_iff_bounded[where C="1"] prob_space.indep_vars_compose2[OF prob_space_measure_pmf b]) auto also have "... = (\l\ L. (\x \ set (fst l) \ set (snd l). \))" using \_check_3 unfolding L_def by (intro sum.cong prod.cong c) auto also have "... = (\l \ L. \^(card (set (fst l) \ set (snd l))))" by simp also have "... = ?R" unfolding L_def \_def by (simp add:case_prod_beta) finally show ?thesis by simp qed lemma hit_count_prod_pow_eq: assumes "i \ B" "j \ B" assumes "lim_balls_and_bins k p" assumes "lim_balls_and_bins k q" assumes "s+t \ k" shows "(\\. (Z i \)^s * (Z j \)^t \p) = (\\. (Z i \)^s * (Z j \)^t \q)" unfolding hit_count_prod_exp[OF assms(1,2,5,3)] unfolding hit_count_prod_exp[OF assms(1,2,5,4)] by simp lemma hit_count_sum_pow_eq: assumes "i \ B" "j \ B" assumes "lim_balls_and_bins k p" assumes "lim_balls_and_bins k q" assumes "s \ k" shows "(\\. (Z i \ + Z j \)^s \p) = (\\. (Z i \ + Z j \)^s \q)" (is "?L = ?R") proof - have q2: "\Z i x ^ l * Z j x ^ (s - l)\ \ real (card R ^ s)" if "l \ {..s}" for s i j l x proof - have "\Z i x ^ l * Z j x ^ (s - l)\ \ Z i x ^ l * Z j x ^ (s - l)" unfolding Z_def by auto also have "... \ real (card R) ^ l * real (card R) ^ (s-l)" unfolding Z_def by (intro mult_mono power_mono of_nat_mono card_mono fin_R) auto also have "... = real (card R)^s" using that by (subst power_add[symmetric]) simp also have "... = real (card R^s)" by simp finally show ?thesis by simp qed have "?L = (\\. (\l\s. real (s choose l) * (Z i \^l * Z j \^(s-l))) \p)" by (subst binomial_ring) (simp add:algebra_simps) also have "... = (\l\s. (\\. real (s choose l) * (Z i \^l * Z j \^(s-l)) \p))" by (intro integral_sum integrable_mult_right integrable_pmf_iff_bounded[where C="card R^s"] q2) auto also have "... = (\l\s. real (s choose l) * (\\. (Z i \^l * Z j \^(s-l)) \p))" by (intro sum.cong integral_mult_right integrable_pmf_iff_bounded[where C="card R^s"] q2) auto also have "... = (\l\s. real (s choose l) * (\\. (Z i \^l * Z j \^(s-l)) \q))" using assms(5) by (intro_cong "[\\<^sub>2 (*)]" more: sum.cong hit_count_prod_pow_eq[OF assms(1-4)]) auto also have "... = (\l\s. (\\. real (s choose l) * (Z i \^l * Z j \^(s-l)) \q))" by (intro sum.cong integral_mult_right[symmetric] integrable_pmf_iff_bounded[where C="card R^s"] q2) auto also have "... = (\\. (\l\s. real (s choose l) * (Z i \^l * Z j \^(s-l))) \q)" by (intro integral_sum[symmetric] integrable_mult_right integrable_pmf_iff_bounded[where C="card R^s"] q2) auto also have "... = ?R" by (subst binomial_ring) (simp add:algebra_simps) finally show ?thesis by simp qed lemma hit_count_sum_poly_eq: assumes "i \ B" "j \ B" assumes "lim_balls_and_bins k p" assumes "lim_balls_and_bins k q" assumes "f \ \ k" shows "(\\. f (Z i \ + Z j \) \p) = (\\. f (Z i \ + Z j \) \q)" (is "?L = ?R") proof - obtain fp where f_def: "f = poly fp" "degree fp \ k" using assms(5) unfolding Polynomials_def by auto have "?L = (\d\degree fp. (\\. poly.coeff fp d * (Z i \ + Z j \) ^ d \p))" unfolding f_def poly_altdef by (intro integral_sum integrable_mult_right Z_any_integrable_2[OF assms(3)]) also have "... = (\d\degree fp. poly.coeff fp d * (\\. (Z i \ + Z j \) ^ d \p))" by (intro sum.cong integral_mult_right Z_any_integrable_2[OF assms(3)]) simp also have "... = (\d\degree fp. poly.coeff fp d *(\\. (Z i \ + Z j \) ^ d \q))" using f_def by (intro sum.cong arg_cong2[where f="(*)"] hit_count_sum_pow_eq[OF assms(1-4)]) auto also have "... = (\d\degree fp. (\\. poly.coeff fp d * (Z i \ + Z j \) ^ d \q))" by (intro sum.cong) auto also have "... = ?R" unfolding f_def poly_altdef by (intro integral_sum[symmetric] integrable_mult_right Z_any_integrable_2[OF assms(4)]) finally show ?thesis by simp qed lemma hit_count_poly_eq: assumes "b \ B" assumes "lim_balls_and_bins k p" assumes "lim_balls_and_bins k q" assumes "f \ \ k" shows "(\\. f (Z b \) \p) = (\\. f (Z b \) \q)" (is "?L = ?R") proof - have a:"(\a. f (a / 2)) \ \ (k*1)" by (intro Polynomials_composeI[OF assms(4)] Polynomials_intros) have "?L = \\. f ((Z b \ + Z b \)/2) \p" by simp also have "... = \\. f ((Z b \ + Z b \)/2) \q" using a by (intro hit_count_sum_poly_eq[OF assms(1,1,2,3)]) simp also have "... = ?R" by simp finally show ?thesis by simp qed lemma lim_balls_and_bins_from_ind_balls_and_bins: "lim_balls_and_bins k \" proof - have "prob_space.indep_vars (measure_pmf \) (\_. discrete) (\x \. \ x) R" unfolding \_def using indep_vars_Pi_pmf[OF fin_R] by metis hence "prob_space.indep_vars (measure_pmf \) (\_. discrete) (\x \. \ x) J" if "J \ R" for J using prob_space.indep_vars_subset[OF prob_space_measure_pmf _ that] by auto hence a:"prob_space.k_wise_indep_vars (measure_pmf \) k (\_. discrete) (\x \. \ x) R" by (simp add:prob_space.k_wise_indep_vars_def[OF prob_space_measure_pmf]) have b: "map_pmf (\\. \ x) \ = pmf_of_set B" if "x \ R" for x using that unfolding \_def Pi_pmf_component[OF fin_R] by simp show ?thesis using a b fin_R fin_B unfolding lim_balls_and_bins_def by auto qed lemma hit_count_factorial_moments: assumes a:"j \ B" assumes "s \ k" assumes "lim_balls_and_bins k p" shows "(\\. ffact s (Z j \) \p) = ffact s (real (card R)) * (1 / real (card B))^s" (is "?L = ?R") proof - have "(\x. ffact s (x-0::real)) \ \ s" by (intro Polynomials_intros) hence b: "ffact s \ (\ k :: (real \ real) set)" using Polynomials_mono[OF assms(2)] by auto have "?L = (\\. ffact s (Z j \) \\)" by (intro hit_count_poly_eq[OF a assms(3) lim_balls_and_bins_from_ind_balls_and_bins] b) also have "... = (\\. ffact s (\i \ {j}. Z i \) \\)" by simp also have "... = ffact s (real (card R)) * (real (card {j}) / real (card B)) ^ s " using assms(1) by (intro fact_moment_balls_and_bins fin_R fin_B) auto also have "... = ?R" by simp finally show ?thesis by simp qed lemma hit_count_factorial_moments_2: assumes a:"i \ B" "j \ B" assumes "i \ j" "s \ k" "card R \ card B" assumes "lim_balls_and_bins k p" shows "(\\. ffact s (Z i \ + Z j \) \p) \ 2^s" (is "?L \ ?R") proof - have "(\x. ffact s (x-0::real)) \ \ s" by (intro Polynomials_intros) hence b: "ffact s \ (\ k :: (real \ real) set)" using Polynomials_mono[OF assms(4)] by auto have or_distrib: "(a \ b) \ (a \ c) \ a \ (b \ c)" for a b c by auto have "?L = (\\. ffact s (Z i \ + Z j \) \\)" by (intro hit_count_sum_poly_eq[OF a assms(6) lim_balls_and_bins_from_ind_balls_and_bins] b) also have "... = (\\. ffact s ((\t \ {i,j}. Z t \)) \\)" using assms(3) by simp also have "... = ffact s (real (card R)) * (real (card {i,j}) / real (card B)) ^ s " using assms(1,2) by (intro fact_moment_balls_and_bins fin_R fin_B) auto also have "... = real (ffact s (card R)) * (real (card {i,j}) / real (card B)) ^ s " by (simp add:of_nat_ffact) also have "... \ (card R)^s * (real (card {i,j}) / real (card B)) ^ s " by (intro mult_mono of_nat_mono ffact_bound, simp_all) also have "... \ (card B)^s * (real (2) / real (card B)) ^ s " using assms(3) by (intro mult_mono of_nat_mono power_mono assms(5), simp_all) also have "... = ?R" using card_B_gt_0 by (simp add:divide_simps) finally show ?thesis by simp qed lemma balls_and_bins_approx_helper: fixes x :: real assumes "x \ 2" assumes "real k \ 5*x / ln x" shows "k \ 2" and "2^(k+3) / fact k \ (1/exp x)^2" and "2 / fact k \ 1 / (exp 1 * exp x)" proof - have ln_inv: "ln x = - ln (1/ x)" if "x > 0" for x :: real using that by (subst ln_div, auto) have apx: "exp 1 \ (5::real)" "4 * ln 4 \ (2 - 2*exp 1/5)*ln (450::real)" "ln 8 * 2 \ (450::real)" "4 / 5 * 2 * exp 1 + ln (5 / 4) * exp 1 \ (5::real)" "exp 1 \ (2::real)^4" by (approximation 10)+ have "2 \ 5 * (x / (x-1))" using assms(1) by (simp add:divide_simps) also have "... \ 5 * (x / ln x)" using assms(1) by (intro mult_left_mono divide_left_mono ln_le_minus_one mult_pos_pos) auto also have "... \ k" using assms(2) by simp finally show k_ge_2: "k \ 2" by simp have "ln x * (2 * exp 1) = ln (((4/5) * x)*(5/4)) * (2 * exp 1)" by simp also have "... = ln ((4/5) * x) * (2 * exp 1) + ln((5/4))*(2 * exp 1)" using assms(1) by (subst ln_mult, simp_all add:algebra_simps) also have "... < (4/5)* x * (2 * exp 1) + ln (5/4) * (x * exp 1)" using assms(1) by (intro add_less_le_mono mult_strict_right_mono ln_less_self mult_left_mono mult_right_mono) (auto simp add:algebra_simps) also have "... = ((4/5) * 2 * exp 1 + ln(5/4) * exp 1) * x" by (simp add:algebra_simps) also have "... \ 5 * x" using assms(1) apx(4) by (intro mult_right_mono, simp_all) finally have 1: "ln x * (2 * exp 1) \ 5 * x" by simp have "ln 8 \ 3 * x - 5 * x * ln(2*exp 1 /5 * ln x) / ln x" proof (cases "x \ {2..450}") case True then show ?thesis by (approximation 10 splitting: x=10) next case False hence x_ge_450: "x \ 450" using assms(1) by simp have "4 * ln 4 \ (2 - 2*exp 1/5)*ln (450::real)" using apx(2) by (simp) also have "... \ (2 - 2*exp 1/5)* ln x" using x_ge_450 apx(1) by (intro mult_left_mono iffD2[OF ln_le_cancel_iff], simp_all) finally have "(2 - 2*exp 1/5) * ln x \ 4 * ln 4" by simp hence "2*exp 1/5 * ln x + 0 \ 2 * exp 1 / 5 * ln x + ((2 - 2*exp 1/5) * ln x - 4 * ln 4)" by (intro add_mono) auto also have "... = 4 * (1/2) * ln x - 4 * ln 4" by (simp add:algebra_simps) also have "... = 4 * (ln (x powr (1/2)) - ln 4)" using x_ge_450 by (subst ln_powr, auto) also have "... = 4 * (ln (x powr (1/2)/4))" using x_ge_450 by (subst ln_div) auto also have "... < 4 * (x powr (1/2)/4)" using x_ge_450 by (intro mult_strict_left_mono ln_less_self) auto also have "... = x powr (1/2)" by simp finally have "2* exp 1/ 5 * ln x \ x powr (1/2)" by simp hence "ln(2* exp 1/ 5 * ln x) \ ln (x powr (1/2))" using x_ge_450 ln_le_cancel_iff by simp hence 0:"ln(2* exp 1/ 5 * ln x) / ln x \ 1/2" using x_ge_450 by (subst (asm) ln_powr, auto) have "ln 8 \ 3 * x - 5 * x * (1/2)" using x_ge_450 apx(3) by simp also have "... \ 3 * x - 5 * x * (ln(2* exp 1/ 5 * ln x) / ln x)" using x_ge_450 by (intro diff_left_mono mult_left_mono 0) auto finally show ?thesis by simp qed hence "2 * x + ln 8 \ 2 * x + (3 * x - 5 * x * ln(2*exp 1 /5 * ln x) / ln x)" by (intro add_mono, auto) also have "... = 5 * x + 5 * x * ln(5 / (2*exp 1*ln x)) / ln x" using assms(1) by (subst ln_inv) (auto simp add:algebra_simps) also have "... = 5 * x * (ln x + ln(5 / (2*exp 1*ln x))) / ln x" using assms(1) by (simp add:algebra_simps add_divide_distrib) also have "... = 5 * x * (ln (5 * x / (2 * exp 1 * ln x))) / ln x" using assms(1) by (subst ln_mult[symmetric], auto) also have "... = (5 * x / ln x) * ln ((5 * x / ln x) / (2 * exp 1))" by (simp add:algebra_simps) also have "... \ k * ln (k / (2*exp 1))" using assms(1,2) 1 k_ge_2 by (intro mult_mono iffD2[OF ln_le_cancel_iff] divide_right_mono) auto finally have "k * ln (k/(2*exp 1)) \ 2*x + ln 8" by simp hence "k * ln(2*exp 1/k) \ -2*x - ln 8" using k_ge_2 by (subst ln_inv, auto) hence "ln ((2*exp 1/k) powr k) \ ln(exp(-2*x)) - ln 8" using k_ge_2 by (subst ln_powr, auto) also have "... = ln(exp(-2*x)/8)" by (simp add:ln_div) finally have "ln ((2*exp 1/k) powr k) \ ln (exp(-2*x)/8)" by simp hence 1: "(2*exp 1/k) powr k \ exp(-2*x)/8" using k_ge_2 assms(1) by (subst (asm) ln_le_cancel_iff) auto have "2^(k+3)/fact k \ 2^(k+3)/(k / exp 1)^k" using k_ge_2 by (intro divide_left_mono fact_lower_bound_1) auto also have "... = 8 * 2^k * (exp 1 / k)^k" by (simp add:power_add algebra_simps power_divide) also have "... = 8 * (2*exp 1/k) powr k" using k_ge_2 powr_realpow by (simp add:power_mult_distrib[symmetric]) also have "... \ 8 * (exp(-2*x)/8)" by (intro mult_left_mono 1) auto also have "... = exp((-x)*2)" by simp also have "... = exp(-x)^2" by (subst exp_powr[symmetric], simp) also have "... = (1/exp x)^2" by (simp add: exp_minus inverse_eq_divide) finally show 2:"2^(k+3)/fact k \ (1/exp x)^2" by simp have "(2::real)/fact k = (2^(k+3)/fact k)/(2^(k+2))" by (simp add:divide_simps power_add) also have "... \ (1/exp x)^2/(2^(k+2))" by (intro divide_right_mono 2, simp) also have "... \ (1/exp x)^1/(2^(k+2))" using assms(1) by (intro divide_right_mono power_decreasing) auto also have "... \ (1/exp x)^1/(2^4)" using k_ge_2 by (intro divide_left_mono power_increasing) auto also have "... \ (1/exp x)^1/exp(1)" using k_ge_2 apx(5) by (intro divide_left_mono) auto also have "... = 1/(exp 1 * exp x)" by simp finally show "(2::real)/fact k \ 1/(exp 1 * exp x)" by simp qed text \Bounds on the expectation and variance in the k-wise independent case. Here the indepedence assumption is improved by a factor of two compared to the result in the paper.\ lemma assumes "card R \ card B" assumes "\c. lim_balls_and_bins (k+1) (p c)" assumes "\ \ {0<..1/exp(2)}" assumes "k \ 5 * ln (card B / \) / ln (ln (card B / \))" shows exp_approx: "\measure_pmf.expectation (p True) Y - measure_pmf.expectation (p False) Y\ \ \ * real (card R)" (is "?A") and var_approx: "\measure_pmf.variance (p True) Y - measure_pmf.variance (p False) Y\ \ \\<^sup>2" (is "?B") proof - let ?p1 = "p False" let ?p2 = "p True" have "exp (2::real) = 1/ (1/exp 2)" by simp also have "... \ 1/ \" using assms(3) by (intro divide_left_mono) auto also have "... \ real (card B)/ \" using assms(3) card_B_gt_0 by (intro divide_right_mono) auto finally have "exp 2 \ real (card B) / \" by simp hence k_condition_h: "2 \ ln (card B / \)" using assms(3) card_B_gt_0 by (subst ln_ge_iff) auto have k_condition_h_2: "0 < real (card B) / \" using assms(3) card_B_gt_0 by (intro divide_pos_pos) auto note k_condition = balls_and_bins_approx_helper[OF k_condition_h assms(4)] define \ :: "real \ real" where "\ = (\x. min x 1)" define f where "f = (\x. 1 - (-1)^k / real (fact k) * ffact k (x-1))" define g where "g = (\x. \ x - f x)" have \_exp: "\ x = f x + g x" for x unfolding g_def by simp have k_ge_2: "k \ 2" using k_condition(1) by simp define \ where "\ = 1 / real (fact k)" have \_nonneg: "\ \ 0" unfolding \_def by simp have k_le_k_plus_1: "k \ k+1" by simp have "f \ \ k" unfolding f_def by (intro Polynomials_intros) hence f_poly: "f \ \ (k+1)" using Polynomials_mono[OF k_le_k_plus_1] by auto have g_diff: "\g x - g (x-1)\ = ffact (k-1) (x-2) / fact (k-1)" if "x \ k" for x :: real proof - have "x \ 2" using k_ge_2 that by simp hence "\ x = \ (x- 1)" unfolding \_def by simp hence "\g x - g (x-1)\ = \f (x-1) - f x\" unfolding g_def by (simp add:algebra_simps) also have "... = \(-1)^k / real (fact k) * (ffact k (x-2) - ffact k (x-1))\" unfolding f_def by (simp add:algebra_simps) also have "... = 1 / real (fact k) * \ffact k (x-1) - ffact k ((x-1)-1)\" by (simp add:abs_mult) also have "... = 1 / real (fact k) * real k * \ffact (k-1) (x-2)\" by (subst ffact_suc_diff, simp add:abs_mult) also have "... = \ffact (k-1) (x-2)\ / fact (k-1)" using k_ge_2 by (subst fact_reduce) auto also have "... = ffact (k-1) (x-2) / fact (k-1)" unfolding ffact_eq_fact_mult_binomial using that k_ge_2 by (intro arg_cong2[where f="(/)"] abs_of_nonneg ffact_nonneg) auto finally show ?thesis by simp qed have f_approx_\: "f x = \ x" if f_approx_\_1: "x \ real ` {0..k}" for x proof (cases "x = 0") case True hence "f x = 1 - (-1)^k / real (fact k) * (\i = 0..i = 0..i = 0..i \ (\x. x + 1) ` {0..i \ {1..k}. real i))" by (intro arg_cong2[where f="(-)"] arg_cong2[where f="(*)"] prod.cong refl) auto also have "... = 0" unfolding fact_prod by simp also have "... = \ x" using True \_def by simp finally show ?thesis by simp next case False hence a: " x \ 1" using that by auto obtain x' where x'_def: "x' \ {0..k}" "x = real x'" using f_approx_\_1 by auto hence "x' - 1 \ {0.. x" unfolding \_def using a by auto finally show ?thesis by simp qed have q2: "\Z i x ^ l * Z j x ^ (s - l)\ \ real (card R ^ s)" if "l \ {..s}" for s i j l x proof - have "\Z i x ^ l * Z j x ^ (s - l)\ \ Z i x ^ l * Z j x ^ (s - l)" unfolding Z_def by auto also have "... \ real (card R) ^ l * real (card R) ^ (s-l)" unfolding Z_def by (intro mult_mono power_mono of_nat_mono card_mono fin_R) auto also have "... = real (card R)^s" using that by (subst power_add[symmetric]) simp also have "... = real (card R^s)" by simp finally show ?thesis by simp qed have q:"real (card A) + real (card B) \ real ` {..2 * card R}" if "A \ R" "B \ R" for A B proof - have "card A + card B \ card R + card R" by (intro add_mono card_mono fin_R that) also have "... = 2 * card R" by simp finally show ?thesis by force qed have g_eq_0_iff_2: "abs (g x) * y = 0" if "x \ \" "x \ 0" "x \ k" for x y :: real proof - have "\x'. x = real_of_int x' \ x' \ k \ x' \ 0" using that Ints_def by fastforce hence "\x'. x = real x' \ x' \ k" by (metis nat_le_iff of_nat_nat) hence "x \ real ` {0..k}" by auto hence "g x = 0" unfolding g_def using f_approx_\ by simp thus ?thesis by simp qed have g_bound_abs: "\\\. g (f \) \p\ \ (\\. ffact (k+1) (f \) \p) * \" (is "?L \ ?R") if "range f \ real ` {..m}" for m and p :: "('a \ 'b) pmf" and f :: "('a \ 'b) \ real" proof - have f_any_integrable: "integrable p (\\. h (f \))" for h :: "real \ real" using that by (intro integrable_pmf_iff_bounded[where C="Max (abs ` h` real ` {..m})"] Max_ge finite_imageI imageI) auto have f_val: "f \ \ real ` {..m}" for \ using that by auto hence f_nat: "f \ \ \" for \ unfolding Nats_def by auto have f_int: "f \ \ real y + 1" if "f \ > real y" for y \ proof - obtain x where x_def: "f \ = real x" "x \ m" using f_val by auto hence "y < x" using that by simp hence "y + 1 \ x" by simp then show ?thesis using x_def by simp qed have f_nonneg: "f \ \ 0" for \ proof - obtain x where x_def: "f \ = real x" "x \ m" using f_val by auto hence "x \ 0" by simp then show ?thesis using x_def by simp qed have "\(real x \ f \)" if "x > m" for x \ proof - obtain x where x_def: "f \ = real x" "x \ m" using f_val by auto then show ?thesis using x_def that by simp qed hence max_Z1: "measure p {\. real x \ f \} = 0" if "x > m" for x using that by auto have "?L \ (\\. \g (f \)\ \p)" by (intro integral_abs_bound) also have "... = (\y\real ` {..m}. \g y\ * measure p {\. f \ = y})" using that by (intro pmf_exp_of_fin_function) auto also have "... = (\y\{..m}. \g (real y)\ * measure p {\. f \ = real y})" by (subst sum.reindex) (auto simp add:comp_def) also have "... = (\y\{..m}. \g (real y)\ * (measure p ({\. f \ = real y} \ {\. f \ > y}) - measure p {\. f \ > y}))" by (subst measure_Union) auto also have "... = (\y\{..m}. \g (real y)\ * (measure p {\. f \ \ y} - measure p {\. f \ > y}))" by (intro sum.cong arg_cong2[where f="(*)"] arg_cong2[where f="(-)"] arg_cong[where f="measure p"]) auto also have "... = (\y\{..m}. \g (real y)\ * measure p {\. f \ \ y}) - (\y\{..m}. \g (real y)\ * measure p {\. f \ > y})" by (simp add:algebra_simps sum_subtractf) also have "... = (\y\{..m}. \g (real y)\ * measure p {\. f \ \ y}) - (\y\{..m}. \g (real y)\ * measure p {\. f \ \ real (y+1)})" using f_int by (intro sum.cong arg_cong2[where f="(-)"] arg_cong2[where f="(*)"] arg_cong[where f="measure p"]) fastforce+ also have "... = (\y\{..m}. \g (real y) \ * measure p {\. f \ \ real y}) - (\y\Suc ` {..m}. \g (real y - 1)\ * measure p {\. f \ \ real y})" by (subst sum.reindex) (auto simp add:comp_def) also have "... = (\y\{..m}. \g (real y) \ * measure p {\. f \ \ real y}) - (\y\{1..m}. \g (real y - 1)\ * measure p {\. f \ \ real y})" using max_Z1 image_Suc_atMost by (intro arg_cong2[where f="(-)"] sum.mono_neutral_cong) auto also have "... = (\y\{k+1..m}. \g (real y) \ * measure p {\. f \ \ y}) - (\y\{k+1..m}. \g (real y - 1)\ * measure p {\. f \ \ y})" using k_ge_2 by (intro arg_cong2[where f="(-)"] sum.mono_neutral_cong_right ballI g_eq_0_iff_2) auto also have "... = (\y\{k+1..m}. (\g (real y)\ - \g (real y-1)\) * measure p {\. f \ \ y})" by (simp add:algebra_simps sum_subtractf) also have "... \ (\y\{k+1..m}. \g (real y)- g (real y-1)\ * measure p {\. ffact (k+1) (f \) \ ffact (k+1) (real y)})" - using ffact_mono by (intro sum_mono mult_mono measure_pmf.pmf_mono[OF refl]) auto + using ffact_mono by (intro sum_mono mult_mono pmf_mono) auto also have "... = (\y\{k+1..m}. (ffact (k-1) (real y-2) / fact (k-1)) * measure p {\. ffact (k+1) (f \) \ ffact (k+1) (real y)})" by (intro sum.cong, simp_all add: g_diff) also have "... \ (\y\{k+1..m}. (ffact (k-1) (real y-2) / fact (k-1)) * ((\\. ffact (k+1) (f \)\p) / ffact (k+1) (real y)))" using k_ge_2 f_nat by (intro sum_mono mult_left_mono pmf_markov f_any_integrable divide_nonneg_pos ffact_of_nat_nonneg ffact_pos) auto also have "... = (\\. ffact (k+1) (f \) \p) / fact (k-1) * (\y\{k+1..m}. ffact (k-1) (real y - 2) / ffact (Suc (Suc (k-1))) (real y))" using k_ge_2 by (simp add:algebra_simps sum_distrib_left) also have "... = (\\. ffact (k+1) (f \) \p) / fact (k-1) * (\y\{k+1..m}. ffact (k-1) (real y - 2) / (real y * (real y - 1) * ffact (k-1) (real y - 2)))" by (subst ffact_Suc, subst ffact_Suc, simp) also have "... = (\\. ffact (k+1) (f \) \p) / fact (k-1) * (\y\{k+1..m}. 1 / (real y * (real y - 1)))" using order.strict_implies_not_eq[OF ffact_pos] k_ge_2 by (intro arg_cong2[where f="(*)"] sum.cong) auto also have "... = (\\. ffact (k+1) (f \) \p) / fact (k-1) * (\y\{Suc k..m}. 1 / (real y - 1) - 1/(real y))" using k_ge_2 by (intro arg_cong2[where f="(*)"] sum.cong) (auto simp add: divide_simps) also have "... = (\\. ffact (k+1) (f \) \p) / fact (k-1) * (\y\{Suc k..m}. (-1/(real y)) - (-1 / (real (y - 1))))" using k_ge_2 by (intro arg_cong2[where f="(*)"] sum.cong) (auto) also have "... = (\\. ffact (k+1) (f \) \p) / fact (k-1) * (of_bool (k \ m) * (1/real k-1/real m))" by (subst sum_telescope_eq, auto) also have "... \ (\\. ffact (k+1) (f \) \p) / fact (k-1) * (1 / real k)" using k_ge_2 f_nat by (intro mult_left_mono divide_nonneg_nonneg integral_nonneg ffact_of_nat_nonneg) auto also have "... = ?R" using k_ge_2 unfolding \_def by (cases k) (auto simp add:algebra_simps) finally show ?thesis by simp qed have z1_g_bound: "\\\. g (Z i \) \p c\ \ (real (card R) / real (card B)) * \" (is "?L1 \ ?R1") if "i \ B" for i c proof - have "?L1 \ (\\. ffact (k+1) (Z i \) \p c) * \" unfolding Z_def using fin_R by (intro g_bound_abs[where m1="card R"]) (auto intro!:imageI card_mono) also have "... = ffact (k+1) (real (card R)) * (1 / real (card B))^(k+1) * \" using that by (subst hit_count_factorial_moments[OF _ _ assms(2)], simp_all) also have "... = real (ffact (k+1) (card R)) * (1 / real (card B))^(k+1) * \" by (simp add:of_nat_ffact) also have "... \ real (card R^(k+1)) * (1 / real (card B))^(k+1) * \" using \_nonneg by (intro mult_right_mono of_nat_mono ffact_bound, simp_all) also have "... \ (real (card R) / real (card B))^(k+1) * \" by (simp add:divide_simps) also have "... \ (real (card R) / real (card B))^1 * \" using assms(1) card_B_gt_0 \_nonneg by (intro mult_right_mono power_decreasing) auto also have "... = ?R1" by simp finally show ?thesis by simp qed have g_add_bound: "\\\. g (Z i \ + Z j \) \p c\ \ 2^(k+1) * \" (is "?L1 \ ?R1") if ij_in_B: "i \ B" "j \ B" "i \ j" for i j c proof - have "?L1 \ (\\. ffact (k+1) (Z i \ + Z j \) \p c) * \" unfolding Z_def using assms(1) by (intro g_bound_abs[where m1="2*card R"]) (auto intro!:imageI q) also have "... \ 2^(k+1) * \" by (intro \_nonneg mult_right_mono hit_count_factorial_moments_2[OF that(1,2,3) _ assms(1,2)]) auto finally show ?thesis by simp qed have Z_poly_diff: "\(\\. \ (Z i \) \?p1) - (\\. \ (Z i \) \?p2)\ \ 2 * ((real (card R) / card B) * \)" (is "?L \ 2 * ?R") if "i \ B" for i proof - note Z_poly_eq = hit_count_poly_eq[OF that assms(2)[of "True"] assms(2)[of "False"] f_poly] have "?L = \(\\. f (Z i \) \?p1) + (\\. g (Z i \) \?p1) - (\\. f (Z i \) \?p2) - (\\. g (Z i \) \?p2)\" using Z_integrable[OF assms(2)] unfolding \_exp by simp also have "... = \(\\. g (Z i \) \?p1) + (- (\\. g (Z i \) \?p2))\" by (subst Z_poly_eq) auto also have "... \ \(\\. g (Z i \) \?p1)\ + \(\\. g (Z i \) \?p2)\" by simp also have "... \ ?R + ?R" by (intro add_mono z1_g_bound that) also have "... = 2 * ?R" by (simp add:algebra_simps) finally show ?thesis by simp qed have Z_poly_diff_2: "\(\\. \ (Z i \) \?p1) - (\\. \ (Z i \) \?p2)\ \ 2 * \" (is "?L \ ?R") if "i \ B" for i proof - have "?L \ 2 * ((real (card R) / real (card B)) * \)" by (intro Z_poly_diff that) also have "... \ 2 * (1 * \)" using assms fin_B that \_nonneg card_gt_0_iff by (intro mult_mono that iffD2[OF pos_divide_le_eq]) auto also have "... = ?R" by simp finally show ?thesis by simp qed have Z_poly_diff_3: "\(\\. \ (Z i \ + Z j \) \?p2) - (\\. \ (Z i \ + Z j \) \?p1)\ \ 2^(k+2)*\" (is "?L \ ?R") if "i \ B" "j \ B" "i \ j" for i j proof - note Z_poly_eq_2 = hit_count_sum_poly_eq[OF that(1,2) assms(2)[of "True"] assms(2)[of "False"] f_poly] have "?L = \(\\. f (Z i \ + Z j \) \?p2) + (\\. g (Z i \ + Z j \) \?p2) - (\\. f (Z i \ + Z j \) \?p1) - (\\. g (Z i \ + Z j \) \?p1)\" using Z_any_integrable_2[OF assms(2)] unfolding \_exp by simp also have "... = \(\\. g (Z i \ + Z j \) \?p2) + (- (\\. g (Z i \ + Z j \) \?p1))\" by (subst Z_poly_eq_2) auto also have "... \ \(\\. g (Z i \ + Z j \) \?p1)\ + \(\\. g (Z i \ + Z j \) \?p2)\" by simp also have "... \ 2^(k+1)*\ + 2^(k+1)*\" by (intro add_mono g_add_bound that) also have "... = ?R" by (simp add:algebra_simps) finally show ?thesis by simp qed have Y_eq: "Y \ = (\i \ B. \ (Z i \))" if "\ \ set_pmf (p c)" for c \ proof - have "\ ` R \ B" proof (rule image_subsetI) fix x assume a:"x \ R" have "\ x \ set_pmf (map_pmf (\\. \ x) (p c))" using that by (subst set_map_pmf) simp also have "... = set_pmf (pmf_of_set B)" by (intro arg_cong[where f="set_pmf"] assms ran[OF assms(2)] a) also have "... = B" by (intro set_pmf_of_set fin_B B_ne) finally show "\ x \ B" by simp qed hence "(\ ` R) = B \ \ ` R" by auto hence "Y \ = card (B \ \ ` R)" unfolding Y_def by auto also have "... = (\i \ B. of_bool (i \ \ ` R))" unfolding of_bool_def using fin_B by (subst sum.If_cases) auto also have "... = (\i \ B. of_bool (card {r \ R. \ r = i} > 0))" using fin_R by (intro sum.cong arg_cong[where f="of_bool"]) (auto simp add:card_gt_0_iff) also have "... = (\i \ B. \(Z i \))" unfolding \_def Z_def by (intro sum.cong) (auto simp add:of_bool_def) finally show ?thesis by simp qed let ?\2 = "(\x y. \ x + \ y - \ (x+y))" let ?Bd = "{x \ B \ B. fst x \ snd x}" have Y_sq_eq': "Y \^ 2 = (\i \ ?Bd. ?\2 (Z (fst i) \) (Z (snd i) \)) + Y \" (is "?L = ?R") if "\ \ set_pmf (p c)" for c \ proof - have a: "\ (Z x \) = of_bool(card {r \ R. \ r = x} > 0)" for x unfolding \_def Z_def by auto have b: "\ (Z x \ + Z y \) = of_bool( card {r \ R. \ r = x} > 0 \ card {r \ R. \ r = y} > 0)" for x y unfolding \_def Z_def by auto have c: "\ (Z x \) * \(Z y \) = ?\2 (Z x \) (Z y \)" for x y unfolding a b of_bool_def by auto have d: "\ (Z x \) * \(Z x \) = \ (Z x \)" for x unfolding a of_bool_def by auto have "?L = (\i\B \ B. \ (Z (fst i) \) * \ (Z (snd i) \))" unfolding Y_eq[OF that] power2_eq_square sum_product sum.cartesian_product by (simp add:case_prod_beta) also have "... = (\i\?Bd \ {x \ B \ B. fst x = snd x}. \ (Z (fst i) \) * \ (Z (snd i) \))" by (intro sum.cong refl) auto also have "... = (\i\?Bd. \ (Z (fst i) \) * \ (Z (snd i) \)) + (\i\{x \ B \ B. fst x = snd x}. \ (Z (fst i) \) * \ (Z (snd i) \))" using assms fin_B by (intro sum.union_disjoint, auto) also have "... = (\i\?Bd. ?\2 (Z (fst i) \) (Z (snd i) \)) + (\i\{x \ B \ B. fst x = snd x}. \ (Z (fst i) \) * \ (Z (fst i) \))" unfolding c by (intro arg_cong2[where f="(+)"] sum.cong) auto also have "... = (\i\?Bd. ?\2 (Z (fst i) \) (Z (snd i) \)) + (\i\fst ` {x \ B \ B. fst x = snd x}. \ (Z i \) * \ (Z i \))" by (subst sum.reindex, auto simp add:inj_on_def) also have "... = (\i\?Bd. ?\2 (Z (fst i) \) (Z (snd i) \)) + (\i \ B. \ (Z i \))" using d by (intro arg_cong2[where f="(+)"] sum.cong refl d) (auto simp add:image_iff) also have "... = ?R" unfolding Y_eq[OF that] by simp finally show ?thesis by simp qed have "\integral\<^sup>L ?p1 Y - integral\<^sup>L ?p2 Y\ = \(\\. (\i \ B. \(Z i \)) \?p1) - (\\. (\i \ B. \(Z i \)) \?p2)\" by (intro arg_cong[where f="abs"] arg_cong2[where f="(-)"] integral_cong_AE AE_pmfI Y_eq) auto also have "... = \(\i \ B. (\\. \(Z i \) \?p1)) - (\i \ B. (\\. \(Z i \) \?p2))\" by (intro arg_cong[where f="abs"] arg_cong2[where f="(-)"] integral_sum Z_integrable[OF assms(2)]) also have "... = \(\i \ B. (\\. \(Z i \) \?p1) - (\\. \(Z i \) \?p2))\" by (subst sum_subtractf) simp also have "... \ (\i \ B. \(\\. \(Z i \) \?p1) - (\\. \(Z i \) \?p2)\)" by simp also have "... \ (\i \ B. 2 * ((real (card R) / real (card B)) *\))" by (intro sum_mono Z_poly_diff) also have "... \ 2 * real (card R) *\" using \_nonneg by (simp) finally have Y_exp_diff_1: "\integral\<^sup>L ?p1 Y - integral\<^sup>L ?p2 Y\ \ 2 * real (card R) *\" by simp have "\integral\<^sup>L ?p1 Y - integral\<^sup>L ?p2 Y\ \ (2 / fact k) * real (card R)" using Y_exp_diff_1 by (simp add:algebra_simps \_def) also have "... \ 1 / (exp 1 * (real (card B) / \)) * card R" using k_condition(3) k_condition_h_2 by (intro mult_right_mono) auto also have "... = \ / (exp 1 * real (card B)) * card R" by simp also have "... \ \ / (1 * 1) * card R" using assms(3) card_B_gt_0 by (intro mult_right_mono divide_left_mono mult_mono) auto also have "... = \ * card R" by simp finally show ?A by simp have "\integral\<^sup>L ?p1 Y - integral\<^sup>L ?p2 Y\ \ 2 * real (card R) *\" using Y_exp_diff_1 by simp also have "... \ 2 * real (card B) *\" by (intro mult_mono of_nat_mono assms \_nonneg) auto finally have Y_exp_diff_2: "\integral\<^sup>L ?p1 Y - integral\<^sup>L ?p2 Y\ \ 2 *\ * real (card B)" by (simp add:algebra_simps) have int_Y: "integrable (measure_pmf (p c)) Y" for c using fin_R card_image_le unfolding Y_def by (intro integrable_pmf_iff_bounded[where C="card R"]) auto have int_Y_sq: "integrable (measure_pmf (p c)) (\\. Y \^2)" for c using fin_R card_image_le unfolding Y_def by (intro integrable_pmf_iff_bounded[where C="real (card R)^2"]) auto have "\(\\. (\i \ ?Bd. ?\2 (Z (fst i) \) (Z (snd i) \)) \?p1) - (\\. (\i \ ?Bd. ?\2 (Z (fst i) \) (Z (snd i) \)) \?p2)\ \ \(\i \ ?Bd. (\\. \ (Z (fst i) \) \?p1) + (\\. \(Z (snd i) \) \?p1) - (\\. \ (Z (fst i) \ + Z (snd i) \) \?p1) - ( (\\. \(Z (fst i) \) \?p2) + (\\. \ (Z (snd i) \) \?p2) - (\\. \(Z (fst i) \ + Z (snd i) \) \?p2)))\" (is "?R3 \ _ ") using Z_integrable[OF assms(2)] Z_any_integrable_2[OF assms(2)] by (simp add:integral_sum sum_subtractf) also have "... = \(\i \ ?Bd. ((\\. \ (Z (fst i) \) \?p1) - (\\. \(Z (fst i) \) \?p2)) + ((\\. \ (Z (snd i) \) \?p1) - (\\. \(Z (snd i) \) \?p2)) + ((\\. \ (Z (fst i) \ + Z (snd i) \) \?p2) - (\\. \(Z (fst i) \ + Z (snd i) \) \?p1)))\" by (intro arg_cong[where f="abs"] sum.cong) auto also have "... \ (\i \ ?Bd. \ ((\\. \ (Z (fst i) \) \?p1) - (\\. \(Z (fst i) \) \?p2)) + ((\\. \ (Z (snd i) \) \?p1) - (\\. \(Z (snd i) \) \?p2)) + ((\\. \ (Z (fst i) \ + Z (snd i) \) \?p2) - (\\. \(Z (fst i) \ + Z (snd i) \) \?p1))\)" by (intro sum_abs) also have "... \ (\i \ ?Bd. \(\\. \ (Z (fst i) \) \?p1) - (\\. \(Z (fst i) \) \?p2)\ + \(\\. \ (Z (snd i) \) \?p1) - (\\. \(Z (snd i) \) \?p2)\ + \(\\. \ (Z (fst i) \ + Z (snd i) \) \?p2) - (\\. \(Z (fst i) \ + Z (snd i) \) \?p1)\)" by (intro sum_mono) auto also have "... \ (\i \ ?Bd. 2*\ + 2 *\ + 2^(k+2)*\)" by (intro sum_mono add_mono Z_poly_diff_2 Z_poly_diff_3) auto also have "... = (2^(k+2)+4) *\ * real (card ?Bd)" by (simp add:algebra_simps) finally have Y_sq_exp_diff_1:"?R3 \ (2^(k+2)+4) *\ * real (card ?Bd)" by simp have "\(\\. Y \ ^2 \?p1) - (\\. Y \^2 \?p2)\ = \(\\. (\i \ ?Bd. ?\2 (Z (fst i) \) (Z (snd i) \)) + Y \ \?p1) - (\\. (\i \ ?Bd. ?\2 (Z (fst i) \) (Z (snd i) \)) + Y \ \?p2)\" by (intro_cong "[\\<^sub>2 (-), \\<^sub>1 abs]" more: integral_cong_AE AE_pmfI Y_sq_eq') auto also have "... \ \(\\. Y \ \?p1) - (\\. Y \ \?p2)\ + \(\\. (\i \ ?Bd. ?\2 (Z (fst i) \) (Z (snd i) \)) \?p1) - (\\. (\i \ ?Bd. ?\2 (Z (fst i) \) (Z (snd i) \)) \?p2)\" using Z_integrable[OF assms(2)] Z_any_integrable_2[OF assms(2)] int_Y by simp also have "... \ 2 *\ * real (card B) + ?R3" by (intro add_mono Y_exp_diff_2, simp) also have "... \ (2^(k+2)+4) *\ * real (card B) + (2^(k+2)+4) *\ * real (card ?Bd)" using \_nonneg by (intro add_mono Y_sq_exp_diff_1 mult_right_mono) auto also have "... = (2^(k+2)+4) *\ * (real (card B) + real (card ?Bd))" by (simp add:algebra_simps) also have "... = (2^(k+2)+4) * \ * real (card B)^2" using power2_nat_le_imp_le by (simp add:card_distinct_pairs of_nat_diff) finally have Y_sq_exp_diff: "\(\\. Y \ ^2 \?p1) - (\\. Y \^2 \?p2)\ \ (2^(k+2)+4) *\ * real (card B)^2" by simp have Y_exp_rough_bound: "\integral\<^sup>L (p c) Y\ \ card B" (is "?L \ ?R") for c proof - have "?L \ (\\. \Y \\ \(p c))" by (intro integral_abs_bound) also have "... \ (\\. real (card R) \(p c))" unfolding Y_def using card_image_le[OF fin_R] by (intro integral_mono integrable_pmf_iff_bounded[where C="card R"]) auto also have "... = card R" by simp also have "... \ card B" using assms by simp finally show ?thesis by simp qed have "\measure_pmf.variance ?p1 Y - measure_pmf.variance ?p2 Y\ = \(\\. Y \ ^2 \?p1) - (\\. Y \ \ ?p1)^2 - ((\\. Y \ ^2 \?p2) - (\\. Y \ \ ?p2)^2)\" by (intro_cong "[\\<^sub>2 (-), \\<^sub>1 abs]" more: measure_pmf.variance_eq int_Y int_Y_sq) also have "... \ \(\\. Y \^2 \?p1) - (\\. Y \^2 \?p2)\ + \(\\. Y \ \ ?p1)\<^sup>2 - (\\. Y \ \ ?p2)\<^sup>2\" by simp also have "... = \(\\. Y \^2 \?p1) - (\\. Y \^2 \?p2)\ + \(\\. Y \ \ ?p1) - (\\. Y \ \ ?p2)\*\(\\. Y \ \ ?p1) + (\\. Y \ \ ?p2)\" by (simp add:power2_eq_square algebra_simps abs_mult[symmetric]) also have "... \ (2^(k+2)+4) *\ * real (card B)^2 + (2*\ *real (card B)) * (\\\. Y \ \?p1\ + \\\. Y \ \ ?p2\)" using \_nonneg by (intro add_mono mult_mono divide_left_mono Y_sq_exp_diff Y_exp_diff_2) auto also have "... \ (2^(k+2)+4)*\ * real (card B)^2 + (2*\ * real (card B)) * (real (card B) + real (card B))" using \_nonneg by (intro add_mono mult_left_mono Y_exp_rough_bound) auto also have "... = (2^(k+2)+2^3) * \ * real (card B)^2" by (simp add:algebra_simps power2_eq_square) also have "... \ (2^(k+2)+2^(k+2)) * \ * real (card B)^2" using k_ge_2 \_nonneg by (intro mult_right_mono add_mono power_increasing, simp_all) also have "... = (2^(k+3) / fact k) * card B^2" by (simp add:power_add \_def) also have "... \ (1 / (real (card B) / \))^2 * card B^2" using k_condition(2) k_condition_h_2 by (intro mult_right_mono) auto also have "... = \^2" using card_B_gt_0 by (simp add:divide_simps) finally show ?B by simp qed lemma assumes "card R \ card B" assumes "lim_balls_and_bins (k+1) p" assumes "k \ 7.5 * (ln (card B) + 2)" shows exp_approx_2: "\measure_pmf.expectation p Y - \\ \ card R / sqrt (card B)" (is "?AL \ ?AR") and var_approx_2: "measure_pmf.variance p Y \ real (card R)^2 / card B" (is "?BL \ ?BR") proof - define q where "q = (\c. if c then \ else p)" have q_altdef: "q True = \" "q False = p" unfolding q_def by auto have a:"lim_balls_and_bins (k+1) (q c)" for c unfolding q_def using assms lim_balls_and_bins_from_ind_balls_and_bins by auto define \ :: real where "\ = min (sqrt (1/card B)) (1 / exp 2)" have c: "\ \ {0<..1 / exp 2}" using card_B_gt_0 unfolding \_def by auto have b: "5 * ln (card B / \) / ln (ln (card B / \)) \ real k" proof (cases "card B \ exp 4") case True hence "sqrt(1/card B) \ sqrt(1/exp 4)" using card_B_gt_0 by (intro real_sqrt_le_mono divide_left_mono) auto also have "... = (1/exp 2)" by (subst powr_half_sqrt[symmetric]) (auto simp add:powr_divide exp_powr) finally have "sqrt(1/card B) \ (1 / exp 2)" by simp hence \_eq: "\ = sqrt(1 /card B)" unfolding \_def by simp have "exp (6::real) = (exp 4) powr (3/2)" by (simp add:exp_powr) also have "...\ card B powr (3/2)" by (intro powr_mono2 True) auto finally have q4:"exp 6 \ card B powr (3/2)" by simp have "(2::real) \ exp 6" by (approximation 5) hence q1: "2 \ real (card B) powr (3 / 2)" using q4 by argo have "(1::real) < ln(exp 6)" by (approximation 5) also have "... \ ln (card B powr (3 / 2))" using card_B_gt_0 by (intro iffD2[OF ln_le_cancel_iff] q4) auto finally have q2: "1 < ln (card B powr (3 / 2))" by simp have "exp (exp (1::real)) \ exp 6" by (approximation 5) also have "... \ card B powr (3/2)" using q4 by simp finally have "exp (exp 1) \ card B powr (3/2)" by simp hence q3: "1\ln(ln (card B powr (3/2)))" using card_B_gt_0 q1 by (intro iffD2[OF ln_ge_iff] ln_gt_zero, auto) have "5 * ln (card B / \) / ln (ln (card B / \)) = 5 * ln (card B powr (1+1/2)) / ln(ln (card B powr (1+1/2)))" unfolding powr_add by (simp add:real_sqrt_divide powr_half_sqrt[symmetric] \_eq) also have "... \ 5 * ln (card B powr (1+1/2)) / 1" using True q1 q2 q3 by (intro divide_left_mono mult_nonneg_nonneg mult_pos_pos ln_ge_zero ln_gt_zero) auto also have "... = 5 * (1+1/2) * ln(card B)" using card_B_gt_0 by (subst ln_powr) auto also have "... = 7.5 * ln(card B)" by simp also have "... \ k" using assms(3) by simp finally show ?thesis by simp next case False have "(1::real) / exp 2 \ sqrt(1 / exp 4)" by (simp add:real_sqrt_divide powr_half_sqrt[symmetric] exp_powr) also have "...\ sqrt(1 /card B)" using False card_B_gt_0 by (intro real_sqrt_le_mono divide_left_mono mult_pos_pos) auto finally have "1 / exp 2 \ sqrt(1/card B)" by simp hence \_eq: "\ = 1/ exp 2" unfolding \_def by simp have q2:"5 * (ln x + 2) / ln (ln x + 2) \ 7.5 * (ln x + 2)" if "x \ {1..exp 4}" for x:: real using that by (approximation 10 splitting: x=10) have "5 * ln (card B / \) / ln (ln (card B / \)) = 5 * (ln (card B) +2) / ln (ln (card B) + 2)" using card_B_gt_0 unfolding \_eq by (simp add:ln_mult) also have "... \ 7.5 * (ln (card B) + 2)" using False card_B_gt_0 by (intro q2) auto also have "... \ k" using assms(3) by simp finally show ?thesis by simp qed have "?AL = \(\\. Y \ \(q True)) - (\\. Y \ \(q False))\" using exp_balls_and_bins unfolding q_def by simp also have "... \ \ * card R" by (intro exp_approx[OF assms(1) a c b]) also have "... \ sqrt (1 / card B) * card R" unfolding \_def by (intro mult_right_mono) auto also have "... = ?AR" using real_sqrt_divide by simp finally show "?AL \ ?AR" by simp show "?BL \ ?BR" proof (cases "R= {}") case True then show ?thesis unfolding Y_def by simp next case "False" hence "card R > 0" using fin_R by auto hence card_R_ge_1: "real (card R) \ 1" by simp have "?BL \ measure_pmf.variance (q True) Y + \measure_pmf.variance (q True) Y - measure_pmf.variance (q False) Y\" unfolding q_def by auto also have "... \ measure_pmf.variance (q True) Y + \^2" by (intro add_mono var_approx[OF assms(1) a c b]) auto also have "... \ measure_pmf.variance (q True) Y + sqrt(1 / card B)^2" unfolding \_def by (intro add_mono power_mono) auto also have "... \ card R * (real (card R) - 1) / card B + sqrt(1 / card B)^2" unfolding q_altdef by (intro add_mono var_balls_and_bins) auto also have "... = card R * (real (card R) - 1) / card B + 1 / card B" by (auto simp add:power_divide real_sqrt_divide) also have "... \ card R * (real (card R) - 1) / card B + card R / card B" by (intro add_mono divide_right_mono card_R_ge_1) auto also have "... = (card R * (real (card R) - 1) + card R) / card B" by argo also have "... = ?BR" by (simp add:algebra_simps power2_eq_square) finally show "?BL \ ?BR" by simp qed qed lemma devitation_bound: assumes "card R \ card B" assumes "lim_balls_and_bins k p" assumes "real k \ C\<^sub>2 * ln (real (card B)) + C\<^sub>3" shows "measure p {\. \Y \ - \\ > 9 * real (card R) / sqrt (real (card B))} \ 1 / 2^6" (is "?L \ ?R") proof (cases "card R > 0") case True define k' :: nat where "k' = k - 1" have "(1::real) \ 7.5 * 0 + 16" by simp also have "... \ 7.5 * ln (real (card B)) + 16" using card_B_ge_1 by (intro add_mono mult_left_mono ln_ge_zero) auto also have "... \ k" using assms(3) unfolding C\<^sub>2_def C\<^sub>3_def by simp finally have k_ge_1: "k \ 1" by simp have lim: "lim_balls_and_bins (k'+1) p" using k_ge_1 assms(2) unfolding k'_def by simp have k'_min: "real k' \ 7.5 * (ln (real (card B)) + 2)" using k_ge_1 assms(3) unfolding C\<^sub>2_def C\<^sub>3_def k'_def by simp let ?r = "real (card R)" let ?b = "real (card B)" have a: "integrable p (\\. (Y \)\<^sup>2)" unfolding Y_def by (intro integrable_pmf_iff_bounded[where C="real (card R)^2"]) (auto intro!: card_image_le[OF fin_R]) have "?L \ \

(\ in measure_pmf p. \Y \- (\\. Y \ \p)\ \ 8*?r / sqrt ?b)" - proof (rule measure_pmf.pmf_mono[OF refl]) + proof (rule pmf_mono) fix \ assume "\ \ set_pmf p" assume a:"\ \ {\. 9 * real (card R) / sqrt (real (card B)) < \Y \ - \\}" have "8 * ?r / sqrt ?b = 9 * ?r / sqrt ?b - ?r / sqrt ?b" by simp also have "... \ \Y \ - \\ - \ (\\. Y \ \p) - \\" using a by (intro diff_mono exp_approx_2[OF assms(1) lim k'_min]) simp also have "... \ \Y \ - (\\. Y \ \p)\" by simp finally have "8 * ?r / sqrt ?b \ \Y \ - (\\. Y \ \p)\" by simp thus "\ \ {\ \ space (measure_pmf p). 8 * ?r / sqrt ?b \ \Y \ - (\\. Y \ \p)\}" by simp qed also have "... \ measure_pmf.variance p Y / (8*?r / sqrt ?b)^2" using True card_B_gt_0 a by (intro measure_pmf.Chebyshev_inequality) auto also have "... \ (?r^2 / ?b) / (8*?r / sqrt ?b)^2" by (intro divide_right_mono var_approx_2[OF assms(1) lim k'_min]) simp also have "... = 1/2^6" using card_B_gt_0 True by (simp add:power2_eq_square) finally show ?thesis by simp next case False hence "R = {}" "card R = 0" using fin_R by auto thus ?thesis unfolding Y_def \_def by simp qed end unbundle no_intro_cong_syntax end \ No newline at end of file diff --git a/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Preliminary.thy b/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Preliminary.thy --- a/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Preliminary.thy +++ b/thys/Distributed_Distinct_Elements/Distributed_Distinct_Elements_Preliminary.thy @@ -1,349 +1,336 @@ section \Preliminary Results\ text \This section contains various short preliminary results used in the sections below.\ theory Distributed_Distinct_Elements_Preliminary imports Frequency_Moments.Frequency_Moments_Preliminary_Results - Frequency_Moments.Product_PMF_Ext + Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF Median_Method.Median Expander_Graphs.Extra_Congruence_Method Expander_Graphs.Constructive_Chernoff_Bound Frequency_Moments.Landau_Ext Stirling_Formula.Stirling_Formula begin unbundle intro_cong_syntax -text \Simplified versions of measure theoretic results for pmfs:\ - -lemma measure_pmf_cong: - assumes "\x. x \ set_pmf p \ x \ P \ x \ Q" - shows "measure (measure_pmf p) P = measure (measure_pmf p) Q" - using assms - by (intro finite_measure.finite_measure_eq_AE AE_pmfI) auto +text \The following aliases are here to prevent possible merge-conflicts. The lemmas have been moved +to @{theory "Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF"}.\ -lemma pmf_mono: - assumes "\x. x \ set_pmf p \ x \ P \ x \ Q" - shows "measure (measure_pmf p) P \ measure (measure_pmf p) Q" -proof - - have "measure (measure_pmf p) P = measure (measure_pmf p) (P \ (set_pmf p))" - by (intro measure_pmf_cong) auto - also have "... \ measure (measure_pmf p) Q" - using assms by (intro finite_measure.finite_measure_mono) auto - finally show ?thesis by simp -qed +lemmas measure_pmf_cong = measure_pmf_cong +lemmas pmf_mono = pmf_mono lemma pmf_rev_mono: assumes "\x. x \ set_pmf p \ x \ Q \ x \ P" shows "measure p P \ measure p Q" using assms by (intro pmf_mono) blast lemma pmf_exp_mono: fixes f g :: "'a \ real" assumes "integrable (measure_pmf p) f" "integrable (measure_pmf p) g" assumes "\x. x \ set_pmf p \ f x \ g x" shows "integral\<^sup>L (measure_pmf p) f \ integral\<^sup>L (measure_pmf p) g" using assms by (intro integral_mono_AE AE_pmfI) auto lemma pmf_markov: assumes "integrable (measure_pmf p) f" "c > 0" assumes "\x. x \ set_pmf p \ f x \ 0" shows "measure p {\. f \ \ c} \ (\\. f \ \p)/ c" (is "?L \ ?R") proof - have a:"AE \ in (measure_pmf p). 0 \ f \" by (intro AE_pmfI assms(3)) have b:"{} \ measure_pmf.events p" unfolding assms(1) by simp have "?L = \

(\ in (measure_pmf p). f \ \ c)" using assms(1) by simp also have "... \ ?R" by (intro integral_Markov_inequality_measure[OF _ b] assms a) finally show ?thesis by simp qed lemma pmf_add: assumes "\x. x \ P \ x \ set_pmf p \ x \ Q \ x \ R" shows "measure p P \ measure p Q + measure p R" proof - have "measure p P \ measure p (Q \ R)" using assms by (intro pmf_mono) blast also have "... \ measure p Q + measure p R" by (rule measure_subadditive, auto) finally show ?thesis by simp qed lemma pair_pmf_prob_left: "measure_pmf.prob (pair_pmf A B) {\. P (fst \)} = measure_pmf.prob A {\. P \}" (is "?L = ?R") proof - have "?L = measure_pmf.prob (map_pmf fst (pair_pmf A B)) {\. P \}" by (subst measure_map_pmf) simp also have "... = ?R" by (subst map_fst_pair_pmf) simp finally show ?thesis by simp qed lemma pmf_exp_of_fin_function: assumes "finite A" "g ` set_pmf p \ A" shows "(\\. f (g \) \p) = (\y \ A. f y * measure p {\. g \ = y})" (is "?L = ?R") proof - have "?L = integral\<^sup>L (map_pmf g p) f" using integral_map_pmf assms by simp also have "... = (\a\A. f a * pmf (map_pmf g p) a)" using assms by (intro integral_measure_pmf_real) auto also have " ... = (\y \ A. f y * measure p (g -` {y}))" unfolding assms(1) by (intro_cong "[\\<^sub>2 (*)]" more:sum.cong pmf_map) also have "... = ?R" by (intro sum.cong) (auto simp add: vimage_def) finally show ?thesis by simp qed text \Cardinality rules for distinct/ordered pairs of a set without the finiteness constraint - to use in simplification:\ lemma card_distinct_pairs: "card {x \ B \ B. fst x \ snd x} = card B^2 - card B" (is "card ?L = ?R") proof (cases "finite B") case True include intro_cong_syntax have "card ?L = card (B \ B - (\x. (x,x)) ` B)" by (intro arg_cong[where f="card"]) auto also have "... = card (B \ B) - card ((\x. (x,x)) ` B)" by (intro card_Diff_subset finite_imageI True image_subsetI) auto also have "... = ?R" using True by (intro_cong "[\\<^sub>2 (-)]" more: card_image) (auto simp add:power2_eq_square inj_on_def) finally show ?thesis by simp next case False then obtain p where p_in: "p \ B" by fastforce have "False" if "finite ?L" proof - have "(\x. (p,x)) ` (B - {p}) \ ?L" using p_in by (intro image_subsetI) auto hence "finite ((\x. (p,x)) ` (B - {p}))" using finite_subset that by auto hence "finite (B - {p})" by (rule finite_imageD) (simp add:inj_on_def) hence "finite B" by simp thus "False" using False by simp qed hence "infinite ?L" by auto hence "card ?L = 0" by simp also have "... = ?R" using False by simp finally show ?thesis by simp qed lemma card_ordered_pairs': fixes M :: "('a ::linorder) set" shows "card {(x,y) \ M \ M. x < y} = card M * (card M - 1) / 2" proof (cases "finite M") case True show ?thesis using card_ordered_pairs[OF True] by linarith next case False then obtain p where p_in: "p \ M" by fastforce let ?f= "(\x. if x < p then (x,p) else (p,x))" have "False" if "finite {(x,y) \ M \ M. x < y}" (is "finite ?X") proof - have "?f ` (M-{p}) \ ?X" using p_in by (intro image_subsetI) auto hence "finite (?f ` (M-{p}))" using that finite_subset by auto moreover have "inj_on ?f (M-{p})" by (intro inj_onI) (metis Pair_inject) ultimately have "finite (M - {p})" using finite_imageD by blast hence "finite M" using finite_insert[where a="p" and A="M-{p}"] by simp thus "False" using False by simp qed hence "infinite ?X" by auto then show ?thesis using False by simp qed text \The following are versions of the mean value theorem, where the interval endpoints may be reversed.\ lemma MVT_symmetric: assumes "\x. \min a b \ x; x \ max a b\ \ DERIV f x :> f' x" shows "\z::real. min a b \ z \ z \ max a b \ (f b - f a = (b - a) * f' z)" proof - consider (a) "a < b" | (b) "a = b" | (c) "a > b" by argo then show ?thesis proof (cases) case a then obtain z :: real where r: "a < z" "z < b" "f b - f a = (b - a) * f' z" using assms MVT2[where a="a" and b="b" and f="f" and f'="f'"] by auto have "a \ z" "z \ b" using r(1,2) by auto thus ?thesis using a r(3) by auto next case b then show ?thesis by auto next case c then obtain z :: real where r: "b < z" "z < a" "f a - f b = (a - b) * f' z" using assms MVT2[where a="b" and b="a" and f="f" and f'="f'"] by auto have "f b - f a = (b-a) * f' z" using r by argo moreover have "b \ z" "z \ a" using r(1,2) by auto ultimately show ?thesis using c by auto qed qed lemma MVT_interval: fixes I :: "real set" assumes "interval I" "a \ I" "b \ I" assumes "\x. x \ I \ DERIV f x :> f' x" shows "\z. z \ I \ (f b - f a = (b - a) * f' z)" proof - have a:"min a b \ I" using assms(2,3) by (cases "a < b") auto have b:"max a b \ I" using assms(2,3) by (cases "a < b") auto have c:"x \ {min a b..max a b} \ x \ I" for x using interval_def assms(1) a b by auto have "\min a b \ x; x \ max a b\ \ DERIV f x :> f' x" for x using c assms(4) by auto then obtain z where z:"z \ min a b" "z \ max a b" "f b - f a = (b-a) * f' z" using MVT_symmetric by blast have "z \ I" using c z(1,2) by auto thus ?thesis using z(3) by auto qed text \Ln is monotone on the positive numbers and thus commutes with min and max:\ lemma ln_min_swap: "x > (0::real) \ (y > 0) \ ln (min x y) = min (ln x) (ln y)" using ln_less_cancel_iff by fastforce lemma ln_max_swap: "x > (0::real) \ (y > 0) \ ln (max x y) = max (ln x) (ln y)" using ln_le_cancel_iff by fastforce text \Loose lower bounds for the factorial fuction:.\ lemma fact_lower_bound: "sqrt(2*pi*n)*(n/exp(1))^n \ fact n" (is "?L \ ?R") proof (cases "n > 0") case True have "ln ?L = ln (2*pi*n)/2 + n * ln n - n" using True by (simp add: ln_mult ln_sqrt ln_realpow ln_div algebra_simps) also have "... \ ln ?R" by (intro Stirling_Formula.ln_fact_bounds True) finally show ?thesis using iffD1[OF ln_le_cancel_iff] True by simp next case False then show ?thesis by simp qed lemma fact_lower_bound_1: assumes "n > 0" shows "(n/exp 1)^n \ fact n" (is "?L \ ?R") proof - have "2 * pi \ 1" using pi_ge_two by auto moreover have "n \ 1" using assms by simp ultimately have "2 * pi * n \ 1*1" by (intro mult_mono) auto hence a:"2 * pi * n \ 1" by simp have "?L = 1 * ?L" by simp also have "... \ sqrt(2 * pi * n) * ?L" using a by (intro mult_right_mono) auto also have "... \ ?R" using fact_lower_bound by simp finally show ?thesis by simp qed text \Rules to handle O-notation with multiple variables, where some filters may be towards zero:\ lemma real_inv_at_right_0_inf: "\\<^sub>F x in at_right (0::real). c \ 1 / x" proof - have "c \ 1 / x" if b:" x \ {0<..<1 / (max c 1)}" for x proof - have "c * x \ (max c 1) * x" using b by (intro mult_right_mono, linarith, auto) also have "... \ (max c 1) * (1 / (max c 1))" using b by (intro mult_left_mono) auto also have "... \ 1" by (simp add:of_rat_divide) finally have "c * x \ 1" by simp moreover have "0 < x" using b by simp ultimately show ?thesis by (subst pos_le_divide_eq, auto) qed thus ?thesis by (intro eventually_at_rightI[where b="1/(max c 1)"], simp_all) qed lemma bigo_prod_1: assumes "(\x. f x) \ O[F](\x. g x)" "G \ bot" shows "(\x. f (fst x)) \ O[F \\<^sub>F G](\x. g (fst x))" proof - obtain c where a: "\\<^sub>F x in F. norm (f x) \ c * norm (g x)" and c_gt_0: "c > 0" using assms unfolding bigo_def by auto have "\c>0. \\<^sub>F x in F \\<^sub>F G. norm (f (fst x)) \ c * norm (g (fst x))" by (intro exI[where x="c"] conjI c_gt_0 eventually_prod1' a assms(2)) thus ?thesis unfolding bigo_def by simp qed lemma bigo_prod_2: assumes "(\x. f x) \ O[G](\x. g x)" "F \ bot" shows "(\x. f (snd x)) \ O[F \\<^sub>F G](\x. g (snd x))" proof - obtain c where a: "\\<^sub>F x in G. norm (f x) \ c * norm (g x)" and c_gt_0: "c > 0" using assms unfolding bigo_def by auto have "\c>0. \\<^sub>F x in F \\<^sub>F G. norm (f (snd x)) \ c * norm (g (snd x))" by (intro exI[where x="c"] conjI c_gt_0 eventually_prod2' a assms(2)) thus ?thesis unfolding bigo_def by simp qed lemma eventually_inv: fixes P :: "real \ bool" assumes "eventually (\x. P (1/x)) at_top " shows "eventually (\x. P x) (at_right 0)" proof - obtain N where c:"n \ N \ P (1/n)" for n using assms unfolding eventually_at_top_linorder by auto define q where "q = max 1 N" have d: "0 < 1 / q" "q > 0" unfolding q_def by auto have "P x" if "x \ {0<..<1 / q}" for x proof - define n where "n = 1/x" have x_eq: "x = 1 / n" unfolding n_def using that by simp have "N \ q" unfolding q_def by simp also have "... \ n" unfolding n_def using that d by (simp add:divide_simps ac_simps) finally have "N \ n" by simp thus ?thesis unfolding x_eq by (intro c) qed thus ?thesis by (intro eventually_at_rightI[where b="1/q"] d) qed lemma bigo_inv: fixes f g :: "real \ real" assumes "(\x. f (1/x)) \ O(\x. g (1/x))" shows "f \ O[at_right 0](g)" using assms eventually_inv unfolding bigo_def by auto unbundle no_intro_cong_syntax end \ No newline at end of file diff --git a/thys/Distributed_Distinct_Elements/Pseudorandom_Combinators.thy b/thys/Distributed_Distinct_Elements/Pseudorandom_Combinators.thy --- a/thys/Distributed_Distinct_Elements/Pseudorandom_Combinators.thy +++ b/thys/Distributed_Distinct_Elements/Pseudorandom_Combinators.thy @@ -1,787 +1,788 @@ section \Combinators for Pseudo-random Objects\ text \This section introduces a combinator library for pseudo-random objects. Each object can be described as a sample space, a function from an initial segment of the natural numbers that selects a value (or data structure.) Semantically they are multisets with the natural interpretation as a probability space (each element is selected with a probability proportional to its occurence count in the multiset). Operationally the selection procedure describes an algorithm to sample from the space. After general definitions and lemmas basic sample spaces, such as chosing a natural uniformly in an initial segment, a product construction the main pseudo-random objects: hash families and expander graphs are introduced. In both cases the range is itself an arbitrary sample space, such that it is for example possible to construct a pseudo-random object that samples seeds for hash families using an expander walk. The definitions @{term "\"} in Section~\ref{sec:inner_algorithm} and @{term "\"} in Section~\ref{sec:outer_algorithm} are good examples. A nice introduction into such constructions has been published by Goldreich~\cite{goldreich2011}.\ subsection \Definitions and General Lemmas\ theory Pseudorandom_Combinators imports Finite_Fields.Card_Irreducible_Polynomials Universal_Hash_Families.Carter_Wegman_Hash_Family - Frequency_Moments.Product_PMF_Ext Distributed_Distinct_Elements_Preliminary + Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF Expander_Graphs.Expander_Graphs_Strongly_Explicit begin unbundle intro_cong_syntax -hide_const "Quantum.T" -hide_const "Discrete_Topology.discrete" -hide_const "Polynomial.order" +hide_const (open) Quantum.T +hide_const (open) Discrete_Topology.discrete +hide_const (open) Isolated.discrete +hide_const (open) Polynomial.order no_notation Digraph.dominates ("_ \\ _" [100,100] 40) record 'a sample_space = size :: "nat" sample_space_select :: "nat \ 'a" definition sample_pmf where "sample_pmf S = map_pmf (sample_space_select S) (pmf_of_set {.. size S > 0" definition "select S k = (sample_space_select S (if k < size S then k else 0))" definition "sample_set S = select S ` {.. {}" using assms unfolding sample_space_def by auto lemma sample_pmf_alt: assumes "sample_space S" shows "sample_pmf S = map_pmf (select S) (pmf_of_set {.. sample_set S" using assms unfolding sample_space_def select_def sample_set_def by auto declare [[coercion sample_pmf]] lemma integrable_sample_pmf[simp]: fixes f :: "'a \ 'c::{banach, second_countable_topology}" assumes "sample_space S" shows "integrable (measure_pmf (sample_pmf S)) f" proof - have "finite (set_pmf (pmf_of_set {..Basic sample spaces\ text \Sample space for uniformly selecting a natural number less than a given bound:\ definition nat_sample_space :: "nat \ nat sample_space" ("[_]\<^sub>S") where "nat_sample_space n = \ size = n, select = id \" lemma nat_sample_pmf: "sample_pmf ([x]\<^sub>S) = pmf_of_set {.. 0" shows "sample_space [n]\<^sub>S" using assms unfolding sample_space_def nat_sample_space_def by simp text \Sample space for the product of two sample spaces:\ definition prod_sample_space :: "'a sample_space \ 'b sample_space \ ('a \ 'b) sample_space" (infixr "\\<^sub>S" 65) where "prod_sample_space s t = \ size = size s * size t, select = (\i. (select s (i mod (size s)), select t (i div (size s)))) \" lemma split_pmf_mod_div': assumes "a > (0::nat)" assumes "b > 0" shows "map_pmf (\x. (x mod a, x div a)) (pmf_of_set {.. {.. b" using that by simp have "x + a * y < a + a * y" using that by simp also have "... = a * (y+1)" by simp also have "... \ a * b" by (intro mult_left_mono a) auto finally show ?thesis by simp qed hence "bij_betw (\x. (x mod a, x div a)) {.. {.. 0" using assms by simp hence "{.. {}" by blast ultimately show "?thesis" by (intro map_pmf_of_set_bij_betw) auto qed lemma pmf_of_set_prod_eq: assumes "A \ {}" "finite A" assumes "B \ {}" "finite B" shows "pmf_of_set (A \ B) = pair_pmf (pmf_of_set A) (pmf_of_set B)" proof - have "indicat_real (A \ B) (i, j) = indicat_real A i * indicat_real B j" for i j by (cases "i \ A"; cases "j \ B") auto hence "pmf (pmf_of_set (A \ B)) (i,j) = pmf (pair_pmf (pmf_of_set A) (pmf_of_set B)) (i,j)" for i j using assms by (simp add:pmf_pair) thus ?thesis by (intro pmf_eqI) auto qed lemma split_pmf_mod_div: assumes "a > (0::nat)" assumes "b > 0" shows "map_pmf (\x. (x mod a, x div a)) (pmf_of_set {.. (0::nat)" assumes "b > 0" shows "map_pmf (\x. x mod a) (pmf_of_set {..x. x mod a) (pmf_of_set {.. (\x. (x mod a, x div a))) (pmf_of_set {..\<^sub>S T) = pair_pmf (sample_pmf S) (sample_pmf T)" (is "?L = ?R") proof - have size: "size S * size T > 0" using assms sample_space_def by (metis nat_0_less_mult_iff) hence a:"{.. {}" "finite {..i. (select S (i mod size S), select T (i div size S))) (pmf_of_set {..(x,y). (select S x, select T y)) \ (\i. (i mod size S, i div size S))) (pmf_of_set {..(x,y). (select S x, select T y)) (map_pmf (\i. (i mod size S, i div size S)) (pmf_of_set {..(x,y). (select S x, select T y)) (pair_pmf (pmf_of_set {..\<^sub>S T)" using assms unfolding sample_space_def prod_sample_space_def by simp lemma prod_sample_set: assumes "sample_space S" assumes "sample_space T" shows "sample_set (S \\<^sub>S T) = sample_set S \ sample_set T" (is "?L = ?R") using assms by (simp add:sample_space_alt prod_sample_pmf) subsection \Hash Families\ lemma indep_vars_map_pmf: assumes "prob_space.indep_vars (measure_pmf p) (\_. discrete) (\i \. X' i (f \)) I" shows "prob_space.indep_vars (measure_pmf (map_pmf f p)) (\_. discrete) X' I" proof - have "prob_space.indep_vars (measure_pmf p) (\_. discrete) (\i. X' i \ f) I" using assms by (simp add:comp_def) hence "prob_space.indep_vars (distr (measure_pmf p) discrete f) (\_. discrete) X' I" by (intro prob_space.indep_vars_distr prob_space_measure_pmf) auto thus ?thesis using map_pmf_rep_eq by metis qed lemma k_wise_indep_vars_map_pmf: assumes "prob_space.k_wise_indep_vars (measure_pmf p) k (\_. discrete) (\i \. X' i (f \)) I" shows "prob_space.k_wise_indep_vars (measure_pmf (map_pmf f p)) k (\_. discrete) X' I" using assms indep_vars_map_pmf unfolding prob_space.k_wise_indep_vars_def[OF prob_space_measure_pmf] by blast lemma (in prob_space) k_wise_indep_subset: assumes "J \ I" assumes "k_wise_indep_vars k M' X' I" shows "k_wise_indep_vars k M' X' J" using assms unfolding k_wise_indep_vars_def by simp lemma (in prob_space) k_wise_indep_vars_reindex: assumes "inj_on f I" assumes "k_wise_indep_vars k M' X' (f ` I)" shows "k_wise_indep_vars k (M' \ f) (\k \. X' (f k) \) I" proof - have "indep_vars (M' \ f) (\k. X' (f k)) J" if "finite J" "card J \ k" "J \ I" for J proof - have "f ` J \ f ` I" using that by auto moreover have "card (f ` J) \ k" using card_image_le[OF that(1)] that(2) order.trans by auto moreover have "finite (f ` J)" using that by auto ultimately have "indep_vars M' X' (f ` J)" using assms(2) unfolding k_wise_indep_vars_def by simp thus ?thesis using that assms(1) inj_on_subset by (intro indep_vars_reindex) qed thus ?thesis unfolding k_wise_indep_vars_def by simp qed definition GF :: "nat \ int set list set ring" where "GF n = (SOME F. finite_field F \ order F = n)" definition is_prime_power :: "nat \ bool" where "is_prime_power n \ (\p k. Factorial_Ring.prime p \ k > 0 \ n = p^k)" lemma assumes "is_prime_power n" shows GF: "finite_field (GF n)" "order (GF n) = n" proof - obtain p k where p_k: "Factorial_Ring.prime p" "k > 0" "n = p^k" using assms unfolding is_prime_power_def by blast have a:"\(F :: int set list set ring). finite_field F \ order F = n" using existence[OF p_k(2,1)] p_k(3) by simp show "finite_field (GF n)" "order (GF n) = n" unfolding GF_def using someI_ex[OF a] by auto qed lemma is_prime_power: "Factorial_Ring.prime p \ k > 0 \ is_prime_power (p^k)" unfolding is_prime_power_def by auto definition split_prime_power :: "nat \ (nat \ nat)" where "split_prime_power n = (THE (p, k). p^k = n \ Factorial_Ring.prime p \ k > 0)" lemma split_prime_power: assumes "Factorial_Ring.prime p" assumes "k > 0" shows "split_prime_power (p^k) = (p,k)" proof - have "q = p \ l = k" if "q^l = p^k" "Factorial_Ring.prime q" "l > 0" for q l proof - have "q dvd p^k" using that by (metis dvd_power) hence "q dvd p" using prime_dvd_power that by auto moreover have "p dvd q^l" using that assms(2) by (metis dvd_power) hence "p dvd q" using prime_dvd_power that assms by blast ultimately have a:"p = q" by auto hence "l = k" using that prime_power_inj by auto thus ?thesis using a by simp qed thus ?thesis unfolding split_prime_power_def using assms by (intro the_equality) auto qed definition \ :: "nat \ nat \ 'a sample_space \ (nat \ 'a) sample_space" where "\ k d R = ( let (p,n) = split_prime_power (size R); m = (LEAST j. d \ p^j \ j \ n); f = from_nat_into (carrier (GF (p^m))); f' = to_nat_on (carrier (GF (p^m))); g = from_nat_into (bounded_degree_polynomials (GF (p^m)) k) in \ size = p^(m*k), select = (\i x. select R ((f' (ring.hash (GF (p^m)) (f x) (g i))) mod p^n))\)" locale hash_sample_space = fixes k d p n :: nat fixes R :: "'a sample_space" assumes p_prime: "Factorial_Ring.prime p" assumes size_R: "size R = p ^ n" assumes k_gt_0: "k > 0" assumes n_gt_0: "n > 0" begin abbreviation S where "S \ \ k d R" lemma p_n_def: "(p,n) = split_prime_power (size R)" unfolding size_R by (intro split_prime_power[symmetric] n_gt_0 p_prime) definition m where "m = (LEAST j. d \ p^j \ j \ n)" definition f where "f = from_nat_into (carrier (GF (p^m)))" definition f' where "f' = to_nat_on (carrier (GF (p^m)))" lemma n_lt_m: "n \ m" and d_lt_p_m: "d \ p^m" proof - define j :: nat where "j = max n d" have "d \ 2^d" by simp also have "... \ 2^j" unfolding j_def by (intro iffD2[OF power_increasing_iff]) auto also have "... \ p^j" using p_prime prime_ge_2_nat by (intro power_mono) auto finally have "d \ p^j" by simp moreover have "n \ j" unfolding j_def by simp ultimately have "d \ p^m \ m \ n" unfolding m_def by (intro LeastI[where P="\x. d \ p^ x \ x \ n" and k="j"]) auto thus "n \ m" "d \ p^m" by auto qed lemma is_field: "finite_field (GF (p^m))" (is "?A") and field_order: "order (GF(p^m)) = p^m" (is "?B") proof - have "is_prime_power (p^m)" using n_gt_0 n_lt_m by (intro is_prime_power p_prime) auto thus "?A" "?B" using GF by auto qed interpretation cw: carter_wegman_hash_family "GF (p^m)" "k" using finite_field_def is_field finite_field_axioms_def by (intro carter_wegman_hash_familyI k_gt_0) auto lemma field_size: "cw.field_size = p^m" using field_order unfolding Coset.order_def by simp lemma f_bij: "bij_betw f {.. 0" by (metis p_prime gr0I not_prime_0 power_not_zero) lemma p_m_gt_0: "p^m > 0" by (metis p_prime gr0I not_prime_0 power_not_zero) lemma S_eq: "S = \ size = p^(m*k), sample_space_select = (\ i x. select R (f' (cw.hash (f x) (g i)) mod p^n )) \" unfolding \_def by (simp add:p_n_def[symmetric] m_def[symmetric] f_def[symmetric] g_def f'_def Let_def cw.space_def) lemma \_size: "size S > 0" unfolding S_eq using p_m_gt_0 k_gt_0 by simp lemma sample_space: "sample_space S" using \_size unfolding sample_space_def by simp lemma sample_space_R: "sample_space R" using size_R p_n_gt_0 unfolding sample_space_def by auto lemma range: "range (select S i) \ sample_set R" proof - define \ where "\ = select S i" have "\ x \ sample_set R" for x proof - have "\ \ sample_set S" unfolding \_def by (intro select_range sample_space) then obtain j where \_alt: "\ = (\x. select R (f' (cw.hash (f x) (g j)) mod p^n))" "j < p^(m*k)" unfolding sample_set_alt[OF sample_space] unfolding S_eq by auto thus "\ x \ sample_set R" unfolding \_alt by (intro select_range sample_space_R) qed thus ?thesis unfolding \_def by auto qed lemma cw_space: "map_pmf g (pmf_of_set {.. 0" using card_gt_0_iff cw.finite_space cw.non_empty_bounded_degree_polynomials by blast show ?thesis unfolding g_def using card_cw_space card_cw_space_gt_0 bij_betw_from_nat_into_finite[where S="cw.space"] by (intro map_pmf_of_set_bij_betw) auto qed lemma single: assumes "x < d" shows "map_pmf (\\. \ x) (sample_pmf S) = sample_pmf R" (is "?L = ?R") proof - have f_x_carr: "f x \ carrier (GF (p^m))" using assms d_lt_p_m by (intro bij_betw_apply[OF f_bij]) auto have "pmf (map_pmf (cw.hash (f x)) (pmf_of_set cw.space)) i = pmf (pmf_of_set (carrier (GF (p ^ m)))) i" (is "?L1 = ?R1") for i proof - have "?L1 = cw.prob (cw.hash (f x) -` {i})" unfolding cw.M_def by (simp add:pmf_map) also have "... = real (card ({i} \ carrier (GF (p ^ m)))) / real cw.field_size" using cw.prob_range[OF f_x_carr, where A="{i}" ] by (simp add:vimage_def) also have "... = ?R1" by (cases "i \ carrier (GF (p^m))", auto) finally show ?thesis by simp qed hence b: "map_pmf (cw.hash (f x)) (pmf_of_set cw.space) = pmf_of_set (carrier (GF (p^m)))" by (intro pmf_eqI) simp have c: "map_pmf f' (pmf_of_set (carrier (GF (p^m)))) = pmf_of_set {.. m" "p > 0" using n_lt_m p_prime prime_gt_0_nat by auto hence d: "map_pmf (\x. x mod p^n) (pmf_of_set {..\. \ x) \ (sample_space_select S)) (pmf_of_set {..\. sample_space_select S \ x) (pmf_of_set {.. (\x. x mod p^n) \ f' \ (cw.hash (f x)) \ g) (pmf_of_set {.._. discrete) (\i \. \ i) {.._. discrete) cw.hash (f ` {.._. discrete) (\i \. select R (f' (cw.hash i \) mod p^n)) (f ` {.._. discrete) (\i \. select R (f' (cw.hash (f i) \) mod p ^ n)) {.. map_pmf g) (pmf_of_set {.._. discrete) (\i \. \ i) {..i x. ?h (g i) x) (pmf_of_set {.._. discrete) (\i \. \ i) {.. 0" defines m_altdef: "m \ max n (nat \log p d\)" shows "size S = p^(m*k)" proof - have "real d = p powr (log p d)" using assms prime_gt_1_nat[OF p_prime] by (intro powr_log_cancel[symmetric]) auto also have "... \ p powr (nat \log p d\)" using prime_gt_1_nat[OF p_prime] by (intro powr_mono) linarith+ also have "... = p^ (nat \log p d\)" using prime_gt_1_nat[OF p_prime] by (subst powr_realpow) auto also have "... \ p^m" using prime_gt_1_nat[OF p_prime] unfolding m_altdef by (intro power_increasing Nat.of_nat_mono) auto finally have "d \ p ^ m" by simp moreover have "n \ m" unfolding m_altdef by simp moreover have "m \ y" if "d \ p ^ y" "n \ y" for y proof - have "log p d \ log p (p ^ y)" using assms prime_gt_1_nat[OF p_prime] by (intro iffD2[OF log_le_cancel_iff] that(1) Nat.of_nat_mono) auto also have "... = log p (p powr (real y))" using prime_gt_1_nat[OF p_prime] by (subst powr_realpow) auto also have "... = y" using prime_gt_1_nat[OF p_prime] by (intro log_powr_cancel) auto finally have "log p d \ y" by simp hence "nat \log p d\ \ y" by simp thus "m \ y" using that(2) unfolding m_altdef by simp qed ultimately have m_eq: "m = (LEAST j. d \ p ^ j \ n \ j)" by (intro Least_equality[symmetric]) auto show ?thesis unfolding S_eq m_def m_eq by simp qed end text \Sample space with a geometric distribution\ fun count_zeros :: "nat \ nat \ nat" where "count_zeros 0 k = 0" | "count_zeros (Suc n) k = (if odd k then 0 else 1 + count_zeros n (k div 2))" lemma count_zeros_iff: "j \ n \ count_zeros n k \ j \ 2^j dvd k" proof (induction j arbitrary: n k) case 0 then show ?case by simp next case (Suc j) then obtain n' where n_def: "n = Suc n'" using Suc_le_D by presburger show ?case using Suc unfolding n_def by auto qed lemma count_zeros_max: "count_zeros n k \ n" by (induction n arbitrary: k) auto definition \ :: "nat \ nat sample_space" where "\ n = \ size = 2^n, sample_space_select = count_zeros n \" lemma \_sample_space[simp]: "sample_space (\ n)" unfolding sample_space_def \_def by simp lemma \_range: "sample_set (\ n) \ {..n}" using count_zeros_max unfolding sample_set_alt[OF \_sample_space] unfolding \_def by auto lemma \_prob: "measure (sample_pmf (\ n)) {\. \ \ j} = of_bool (j \ n) / 2^j" (is "?L = ?R") proof (cases "j \ n") case True have a:"{..<(2^n)::nat} \ {}" by (simp add: lessThan_empty_iff) have b:"finite {..<(2^n)::nat}" by simp define f :: "nat \ nat" where "f = (\x. x * 2^j)" have d:"inj_on f {..<2^(n-j)}" unfolding f_def by (intro inj_onI) simp have e:"2^j > (0::nat)" by simp have "y \ f ` {..< 2^(n-j)} \ y \ {x. x < 2^n \ 2^j dvd x}" for y :: nat proof - have "y \ f ` {..< 2^(n-j)} \ (\x. x < 2 ^ (n - j) \ y = 2 ^ j * x)" unfolding f_def by auto also have "... \ (\x. 2^j * x < 2^j * 2 ^ (n-j) \ y = 2 ^ j * x)" using e by simp also have "... \ (\x. 2^j * x < 2^n \ y = 2 ^ j * x)" using True by (subst power_add[symmetric]) simp also have "... \ (\x. y < 2^n \ y = x * 2 ^ j)" by (metis Groups.mult_ac(2)) also have "... \ y \ {x. x < 2^n \ 2^j dvd x}" by auto finally show ?thesis by simp qed hence c:"f ` {..< 2^(n-j)} = {x. x < 2^n \ 2^j dvd x}" by auto have "?L = measure (pmf_of_set {..<2^n}) {\. count_zeros n \ \ j}" unfolding sample_pmf_def \_def by simp also have "... = real (card {x::nat. x < 2^n \ 2^j dvd x}) / 2^n" by (simp add: measure_pmf_of_set[OF a b] count_zeros_iff[OF True]) (simp add:lessThan_def Collect_conj_eq) also have "... = real (card (f ` {..<2^(n-j)})) / 2^n" by (simp add:c) also have "... = real (card ({..<(2^(n-j)::nat)})) / 2^n" by (simp add: card_image[OF d]) also have "... = ?R" using True by (simp add:frac_eq_eq power_add[symmetric]) finally show ?thesis by simp next case False have "set_pmf (sample_pmf (\ n)) \ {..n}" unfolding sample_space_alt[OF \_sample_space, symmetric] using \_range by simp hence "?L = measure (sample_pmf (\ n)) {}" using False by (intro measure_pmf_cong) auto also have "... = ?R" using False by simp finally show ?thesis by simp qed lemma \_prob_single: "measure (sample_pmf (\ n)) {j} \ 1 / 2^j" (is "?L \ ?R") proof - have "?L = measure (sample_pmf (\ n)) ({j..}-{j+1..})" by (intro measure_pmf_cong) auto also have "... = measure (sample_pmf (\ n)) {j..} - measure (sample_pmf (\ n)) {j+1..}" by (intro measure_Diff) auto also have "... = measure (sample_pmf (\ n)) {\. \ \ j}-measure (sample_pmf (\ n)) {\. \ \ (j+1)}" by (intro arg_cong2[where f="(-)"] measure_pmf_cong) auto also have "... = of_bool (j \ n) * 1 / 2 ^ j - of_bool (j + 1 \ n) / 2 ^ (j + 1)" unfolding \_prob by simp also have "... \ 1/2^j - 0" by (intro diff_mono) auto also have "... = ?R" by simp finally show ?thesis by simp qed subsection \Expander Walks\ definition \ :: "nat \ real \ 'a sample_space \ (nat \ 'a) sample_space" where "\ l \ S = (let e = see_standard (size S) \ in \ size = see_size e * see_degree e^(l-1), sample_space_select = (\i j. select S (see_sample_walk e (l-1) i ! j)) \) " locale expander_sample_space = fixes l :: nat fixes \ :: real fixes S :: "'a sample_space" assumes l_gt_0: "l > 0" assumes \_gt_0: "\ > 0" assumes sample_space_S: "sample_space S" begin definition e where "e = see_standard (size S) \" lemma size_S_gt_0: "size S > 0" using sample_space_S unfolding sample_space_def by simp lemma \_alt: "(\ l \ S) = \ size = see_size e * see_degree e^(l-1), sample_space_select = (\i j. select S (see_sample_walk e (l-1) i ! j)) \" unfolding \_def e_def[symmetric] by (simp add:Let_def) lemmas see_standard = see_standard[OF size_S_gt_0 \_gt_0] sublocale E: regular_graph "graph_of e" using see_standard(1) unfolding is_expander_def e_def by auto lemma e_deg_gt_0: "see_degree e > 0" unfolding e_def see_standard by simp lemma e_size_gt_0: "see_size e > 0" unfolding e_def see_standard using size_S_gt_0 by simp lemma sample_space: "sample_space (\ l \ S)" unfolding sample_space_def \_alt using e_size_gt_0 e_deg_gt_0 by simp lemma range: "select (\ l \ S) i j \ sample_set S" proof - define \ where "\ = select (\ l \ S) i" have "\ \ sample_set (\ l \ S)" unfolding \_def by (intro select_range sample_space) then obtain k where "\ = sample_space_select (\ l \ S) k" using sample_set_alt[OF sample_space] by auto hence "\ j \ sample_set S" unfolding \_alt using select_range[OF sample_space_S] by simp thus ?thesis unfolding \_def by simp qed lemma sample_set: "sample_set (\ l \ S) \ (UNIV \ sample_set S)" proof (rule subsetI) fix x assume "x \ sample_set (\ l \ S)" then obtain i where "x = select (\ l \ S) i" unfolding sample_set_def by auto thus "x \ UNIV \ sample_set S" using range by auto qed lemma walks: defines "R \ map_pmf (\xs i. select S (xs ! i)) (pmf_of_multiset (walks (graph_of e) l))" shows "sample_pmf (\ l \ S) = R" proof - let ?S = "{.. ?S" using e_size_gt_0 e_deg_gt_0 l_gt_0 by auto hence "?S \ {}" by blast hence "?T = pmf_of_multiset {#see_sample_walk e (l-1) i. i \# mset_set ?S#}" by (subst map_pmf_of_set) simp_all also have "... = pmf_of_multiset (walks' (graph_of e) (l-1))" by (subst see_sample_walk) auto also have "... = pmf_of_multiset (walks (graph_of e) l)" unfolding walks_def using l_gt_0 by (cases l, simp_all) finally have 0:"?T = pmf_of_multiset (walks (graph_of e) l)" by simp have "sample_pmf (\ l \ S) = map_pmf (\xs j. select S (xs ! j)) ?T" unfolding map_pmf_comp sample_pmf_def \_alt by simp also have "... = R" unfolding 0 R_def by simp finally show ?thesis by simp qed lemma uniform_property: assumes "i < l" shows "map_pmf (\w. w i) (\ l \ S) = sample_pmf S" (is "?L = ?R") proof - have "?L = map_pmf (select S) (map_pmf (\xs. (xs ! i)) (pmf_of_multiset (walks (graph_of e) l)))" unfolding walks by (simp add: map_pmf_comp) also have "... = map_pmf (select S) (pmf_of_set (verts (graph_of e)))" unfolding E.uniform_property[OF assms] by simp also have "... = ?R" unfolding sample_pmf_alt[OF sample_space_S] e_def graph_of_def using see_standard by simp finally show ?thesis by simp qed lemma size: "size (\ l \ S) = size S * (16 ^ ((l-1) * nat \ln \ / ln (19 / 20)\))" (is "?L = ?R") proof - have "?L = see_size e * see_degree e ^ (l - 1)" unfolding \_alt by simp also have "... = size S * (16 ^ nat \ln \ / ln (19 / 20)\) ^ (l - 1)" using see_standard unfolding e_def by simp also have "... = size S * (16 ^ ((l-1) * nat \ln \ / ln (19 / 20)\))" unfolding power_mult[symmetric] by (simp add:ac_simps) finally show ?thesis by simp qed end end \ No newline at end of file diff --git a/thys/Distributed_Distinct_Elements/ROOT b/thys/Distributed_Distinct_Elements/ROOT --- a/thys/Distributed_Distinct_Elements/ROOT +++ b/thys/Distributed_Distinct_Elements/ROOT @@ -1,20 +1,21 @@ chapter AFP session Distributed_Distinct_Elements = "Expander_Graphs" + options [timeout = 1200] sessions Discrete_Summation Stirling_Formula + Frequency_Moments theories Distributed_Distinct_Elements_Preliminary Pseudorandom_Combinators Distributed_Distinct_Elements_Balls_and_Bins Distributed_Distinct_Elements_Tail_Bounds Distributed_Distinct_Elements_Inner_Algorithm Distributed_Distinct_Elements_Accuracy_Without_Cutoff Distributed_Distinct_Elements_Cutoff_Level Distributed_Distinct_Elements_Accuracy Distributed_Distinct_Elements_Outer_Algorithm document_files "root.tex" "root.bib" diff --git a/thys/Expander_Graphs/Constructive_Chernoff_Bound.thy b/thys/Expander_Graphs/Constructive_Chernoff_Bound.thy --- a/thys/Expander_Graphs/Constructive_Chernoff_Bound.thy +++ b/thys/Expander_Graphs/Constructive_Chernoff_Bound.thy @@ -1,431 +1,431 @@ subsection \Constructive Chernoff Bound\label{sec:constructive_chernoff_bound}\ text \This section formalizes Theorem~5 by Impagliazzo and Kabanets~\cite{impagliazzo2010}. It is -a general result with which Chernoff-type tail bounds for various kinds of weakly dependent random +a general result with which Chernoff-type tail bounds for various kinds of weakly dependent random variables can be obtained. The results here are general and will be applied in Section~\ref{sec:random_walks} to random walks in expander graphs.\ theory Constructive_Chernoff_Bound - imports - "HOL-Probability.Probability_Measure" - Frequency_Moments.Product_PMF_Ext + imports + "HOL-Probability.Probability_Measure" + Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF Weighted_Arithmetic_Geometric_Mean.Weighted_Arithmetic_Geometric_Mean begin lemma powr_mono_rev: fixes x :: real assumes "a \ b" and "x > 0" "x \ 1" shows "x powr b \ x powr a" proof - have "x powr b = (1/x) powr (-b)" using assms by (simp add: powr_divide powr_minus_divide) also have "... \ (1/x) powr (-a)" using assms by (intro powr_mono) auto also have "... = x powr a" using assms by (simp add: powr_divide powr_minus_divide) finally show ?thesis by simp qed lemma exp_powr: "(exp x) powr y = exp (x*y)" for x :: real unfolding powr_def by simp lemma integrable_pmf_iff_bounded: fixes f :: "'a \ real" assumes "\x. x \ set_pmf p \ abs (f x) \ C" shows "integrable (measure_pmf p) f" proof - obtain x where "x \ set_pmf p" using set_pmf_not_empty by fast hence "C \ 0" using assms(1) by fastforce - hence " (\\<^sup>+ x. ennreal (abs (f x)) \measure_pmf p) \ (\\<^sup>+ x. C \measure_pmf p)" + hence " (\\<^sup>+ x. ennreal (abs (f x)) \measure_pmf p) \ (\\<^sup>+ x. C \measure_pmf p)" using assms ennreal_le_iff by (intro nn_integral_mono_AE AE_pmfI) auto also have "... = C" by simp also have "... < Orderings.top" by simp finally have "(\\<^sup>+ x. ennreal (abs (f x)) \measure_pmf p) < Orderings.top" by simp - thus ?thesis + thus ?thesis by (intro iffD2[OF integrable_iff_bounded]) auto qed -lemma split_pair_pmf: - "measure_pmf.prob (pair_pmf A B) S = integral\<^sup>L A (\a. measure_pmf.prob B {b. (a,b) \ S})" +lemma split_pair_pmf: + "measure_pmf.prob (pair_pmf A B) S = integral\<^sup>L A (\a. measure_pmf.prob B {b. (a,b) \ S})" (is "?L = ?R") proof - have a:"integrable (measure_pmf A) (\x. measure_pmf.prob B {b. (x, b) \ S})" by (intro integrable_pmf_iff_bounded[where C="1"]) simp have "?L = (\\<^sup>+x. indicator S x \(measure_pmf (pair_pmf A B)))" by (simp add: measure_pmf.emeasure_eq_measure) also have "... = (\\<^sup>+x. (\\<^sup>+y. indicator S (x,y) \B) \A)" - by (simp add: nn_integral_pair_pmf') + by (simp add: nn_integral_pair_pmf') also have "... = (\\<^sup>+x. (\\<^sup>+y. indicator {b. (x,b) \ S} y \B) \A)" by (simp add:indicator_def) also have "... = (\\<^sup>+x. (measure_pmf.prob B {b. (x,b) \ S}) \A)" by (simp add: measure_pmf.emeasure_eq_measure) also have "... = ?R" using a by (subst nn_integral_eq_integral) auto finally show ?thesis by simp qed -lemma split_pair_pmf_2: - "measure(pair_pmf A B) S = integral\<^sup>L B (\a. measure_pmf.prob A {b. (b,a) \ S})" +lemma split_pair_pmf_2: + "measure(pair_pmf A B) S = integral\<^sup>L B (\a. measure_pmf.prob A {b. (b,a) \ S})" (is "?L = ?R") proof - have "?L = measure (pair_pmf B A) {\. (snd \, fst \) \ S}" by (subst pair_commute_pmf) (simp add:vimage_def case_prod_beta) also have "... = ?R" unfolding split_pair_pmf by simp finally show ?thesis by simp qed -definition KL_div :: "real \ real \ real" +definition KL_div :: "real \ real \ real" where "KL_div p q = p * ln (p/q) + (1-p) * ln ((1-p)/(1-q))" theorem impagliazzo_kabanets_pmf: - fixes Y :: "nat \ 'a \ bool" + fixes Y :: "nat \ 'a \ bool" fixes p :: "'a pmf" assumes "n > 0" assumes "\i. i \ {.. \ i \ {0..1}" assumes "\S. S \ {.. measure p {\. (\i \ S. Y i \)} \ (\i \ S. \ i)" defines "\_avg \ (\i\ {.. i)/n" assumes "\ \ {\_avg..1}" - assumes "\_avg > 0" - shows "measure p {\. real (card {i \ {..}) \ \ * n} \ exp (-real n * KL_div \ \_avg)" + assumes "\_avg > 0" + shows "measure p {\. real (card {i \ {..}) \ \ * n} \ exp (-real n * KL_div \ \_avg)" (is "?L \ ?R") proof - let ?n = "real n" define q :: real where "q = (if \ = 1 then 1 else (\-\_avg)/(\*(1-\_avg)))" define g where "g \ = card {i. i < n \ \Y i \}" for \ let ?E = "(\\. real (card {i. i < n \ Y i \}) \ \ * n)" let ?\ = "prod_pmf {.._. bernoulli_pmf q)" - have q_range:"q \{0..1}" + have q_range:"q \{0..1}" proof (cases "\ < 1") case True - then show ?thesis + then show ?thesis using assms(5,6) unfolding q_def by (auto intro!:divide_nonneg_pos simp add:algebra_simps) next case False hence "\ = 1" using assms(5) by simp then show ?thesis unfolding q_def by simp qed have abs_pos_le_1I: "abs x \ 1" if "x \ 0" "x \ 1" for x :: real using that by auto - have \_n_nonneg: "\*?n \ 0" - using assms(1,5,6) by simp + have \_n_nonneg: "\*?n \ 0" + using assms(1,5,6) by simp define r where "r = n - nat \\*n\" have 2:"(1-q) ^ r \ (1-q)^ g \" if "?E \" for \ proof - have "g \ = card ({i. i < n} - {i. i < n \ Y i \})" unfolding g_def by (intro arg_cong[where f="\x. card x"]) auto also have "... = card {i. i < n} - card {i. i < n \ Y i \}" by (subst card_Diff_subset, auto) also have "... \ card {i. i < n} - nat \\*n\" using that \_n_nonneg by (intro diff_le_mono2) simp also have "... = r" unfolding r_def by simp finally have "g \ \ r" by simp thus "(1-q) ^ r \ (1-q) ^ (g \)" using q_range by (intro power_decreasing) auto qed have \_gt_0: "\ > 0" using assms(5,6) by simp have q_lt_1: "q < 1" if "\ < 1" proof - have "\_avg < 1" using assms(5) that by simp - hence "(\ - \_avg) / (\ * (1 - \_avg)) < 1" + hence "(\ - \_avg) / (\ * (1 - \_avg)) < 1" using \_gt_0 assms(6) that by (subst pos_divide_less_eq) (auto simp add:algebra_simps) thus "q < 1" unfolding q_def using that by simp qed have 5: "(\_avg * q + (1-q)) / (1-q) powr (1-\) = exp (- KL_div \ \_avg)" (is "?L1 = ?R1") - if "\ < 1" + if "\ < 1" proof - - have \_avg_range: "\_avg \ {0<..<1}" + have \_avg_range: "\_avg \ {0<..<1}" using that assms(5,6) by simp have "?L1 = (1 - (1-\_avg) * q) / (1-q) powr (1-\)" by (simp add:algebra_simps) also have "... = (1 - (\-\_avg) / \ ) / (1-q) powr (1-\)" unfolding q_def using that \_gt_0 \_avg_range by simp also have "... = (\_avg / \) / (1-q) powr (1-\)" using \_gt_0 by (simp add:divide_simps) also have "... = (\_avg / \) * (1/(1-q)) powr (1-\)" using q_lt_1[OF that] by (subst powr_divide, simp_all) also have "... = (\_avg / \) * (1/((\*(1-\_avg)-(\-\_avg))/(\*(1-\_avg)))) powr (1-\)" using \_gt_0 \_avg_range unfolding q_def by (simp add:divide_simps) also have "... = (\_avg / \) * ((\ / \_avg) *((1-\_avg)/(1-\))) powr (1-\)" by (simp add:algebra_simps) also have "... = (\_avg / \) * (\ / \_avg) powr (1-\) *((1-\_avg)/(1-\)) powr (1-\)" using \_gt_0 \_avg_range that by (subst powr_mult, auto) also have "... = (\_avg / \) powr 1 * (\_avg / \) powr -(1-\) *((1-\_avg)/(1-\)) powr (1-\)" using \_gt_0 \_avg_range that unfolding powr_minus_divide by (simp add:powr_divide) also have "... = (\_avg / \) powr \ *((1-\_avg)/(1-\)) powr (1-\)" by (subst powr_add[symmetric]) simp also have "... = exp ( ln ((\_avg / \) powr \ *((1-\_avg)/(1-\)) powr (1-\)))" using \_gt_0 \_avg_range that by (intro exp_ln[symmetric] mult_pos_pos) auto - also have "... = exp ((ln ((\_avg / \) powr \) + ln (((1 - \_avg) / (1 - \)) powr (1-\))))" + also have "... = exp ((ln ((\_avg / \) powr \) + ln (((1 - \_avg) / (1 - \)) powr (1-\))))" using \_gt_0 \_avg_range that by (subst ln_mult) auto - also have "... = exp ((\ * ln (\_avg / \) + (1 - \) * ln ((1 - \_avg) / (1 - \))))" + also have "... = exp ((\ * ln (\_avg / \) + (1 - \) * ln ((1 - \_avg) / (1 - \))))" using \_gt_0 \_avg_range that by (simp add:ln_powr algebra_simps) - also have "... = exp (- (\ * ln (\ / \_avg) + (1 - \) * ln ((1 - \) / (1 - \_avg))))" + also have "... = exp (- (\ * ln (\ / \_avg) + (1 - \) * ln ((1 - \) / (1 - \_avg))))" using \_gt_0 \_avg_range that by (simp add: ln_div algebra_simps) also have "... = ?R1" unfolding KL_div_def by simp finally show ?thesis by simp qed have 3: "(\_avg * q + (1-q)) ^ n / (1-q) ^ r \ exp (- ?n* KL_div \ \_avg)" (is "?L1 \ ?R1") proof (cases "\ < 1") case True - have "\ * real n \ 1 * real n" + have "\ * real n \ 1 * real n" using True by (intro mult_right_mono) auto hence "r = real n - real (nat \\ * real n\)" unfolding r_def by (subst of_nat_diff) auto also have "... = real n - \\ * real n\" using \_n_nonneg by (subst of_nat_nat, auto) also have "... \ ?n - \ * ?n" by (intro diff_mono) auto also have "... = (1-\) *?n" by (simp add:algebra_simps) finally have r_bound: "r \ (1-\)*n" by simp have "?L1 = (\_avg * q + (1-q)) ^ n / (1-q) powr r" using q_lt_1[OF True] assms(1) by (simp add: powr_realpow) also have "... = (\_avg * q + (1-q)) powr n / (1-q) powr r" using q_lt_1[OF True] assms(6) q_range by (subst powr_realpow[symmetric], auto intro!:add_nonneg_pos) also have "... \ (\_avg * q + (1-q)) powr n / (1-q) powr ((1-\)*n)" using q_range q_lt_1[OF True] by (intro divide_left_mono powr_mono_rev r_bound) auto also have "... = (\_avg * q + (1-q)) powr n / ((1-q) powr (1-\)) powr n" unfolding powr_powr by simp also have "... = ((\_avg * q + (1-q)) / (1-q) powr (1-\)) powr n" using assms(6) q_range by (subst powr_divide) auto also have "... = exp (- KL_div \ \_avg) powr real n" unfolding 5[OF True] by simp - also have "... = ?R1" + also have "... = ?R1" unfolding exp_powr by simp finally show ?thesis by simp next case False hence \_eq_1: "\=1" using assms(5) by simp have "?L1 = \_avg ^ n" using \_eq_1 r_def q_def by simp also have "... = exp( - KL_div 1 \_avg) ^ n" - unfolding KL_div_def using assms(6) by (simp add:ln_div) - also have "... = ?R1" + unfolding KL_div_def using assms(6) by (simp add:ln_div) + also have "... = ?R1" using \_eq_1 by (simp add: powr_realpow[symmetric] exp_powr) finally show ?thesis by simp qed - have 4: "(1 - q) ^ r > 0" + have 4: "(1 - q) ^ r > 0" proof (cases "\ < 1") case True then show ?thesis using q_lt_1[OF True] by simp next case False hence "\=1" using assms(5) by simp hence "r=0" unfolding r_def by simp then show ?thesis by simp qed have "(1-q) ^ r * ?L = (\\. indicator {\. ?E \} \ * (1-q) ^ r \p)" by simp also have "... \ (\\. indicator {\. ?E \} \ * (1-q) ^ g \ \p)" - using q_range 2 by (intro integral_mono_AE integrable_pmf_iff_bounded[where C="1"] - abs_pos_le_1I mult_le_one power_le_one AE_pmfI) (simp_all split:split_indicator) + using q_range 2 by (intro integral_mono_AE integrable_pmf_iff_bounded[where C="1"] + abs_pos_le_1I mult_le_one power_le_one AE_pmfI) (simp_all split:split_indicator) also have "... = (\\. indicator {\. ?E \} \ * (\i \ {i. i < n \ \Y i \}. (1-q)) \p)" unfolding g_def using q_range by (intro integral_cong_AE AE_pmfI, simp_all add:powr_realpow) also have "... = (\\. indicator {\. ?E \} \ * measure ?\ ({j. j < n \ \Y j \} \ {False}) \p)" using q_range by (subst prob_prod_pmf') (auto simp add:measure_pmf_single) also have "... = (\\. measure ?\ {\. ?E \ \ (\i\{j. j < n \ \Y j \}. \\ i)} \p)" by (intro integral_cong_AE AE_pmfI, simp_all add:Pi_def split:split_indicator) also have "... = (\\. measure ?\ {\. ?E \ \ (\i\{.. i \ Y i \)} \p)" by (intro integral_cong_AE AE_pmfI measure_eq_AE) auto - also have "... = measure (pair_pmf p ?\) {\.?E (fst \)\(\i \ {.. i \ Y i (fst \))}" + also have "... = measure (pair_pmf p ?\) {\.?E (fst \)\(\i \ {.. i \ Y i (fst \))}" unfolding split_pair_pmf by simp also have "... \ measure (pair_pmf p ?\) {\. (\i \ {j. j < n \ snd \ j}. Y i (fst \))}" - by (intro measure_pmf.pmf_mono[OF refl], auto) + by (intro pmf_mono, auto) also have "... = (\\. measure p {\. \i\{j. j< n \ \ j}. Y i \} \ ?\)" unfolding split_pair_pmf_2 by simp also have "... \ (\a. (\i \ {j. j < n \ a j}. \ i) \ ?\)" - using assms(2) by (intro integral_mono_AE AE_pmfI assms(3) subsetI prod_le_1 prod_nonneg + using assms(2) by (intro integral_mono_AE AE_pmfI assms(3) subsetI prod_le_1 prod_nonneg integrable_pmf_iff_bounded[where C="1"] abs_pos_le_1I) auto also have "... = (\a. (\i \ {.. i^ of_bool(a i)) \ ?\)" - unfolding of_bool_def by (intro integral_cong_AE AE_pmfI) + unfolding of_bool_def by (intro integral_cong_AE AE_pmfI) (auto simp add:if_distrib prod.If_cases Int_def) also have "... = (\ia. (\ i ^ of_bool a) \(bernoulli_pmf q)))" using assms(2) by (intro expectation_prod_Pi_pmf integrable_pmf_iff_bounded[where C="1"]) auto also have "... = (\i i * q + (1-q))" using q_range by simp also have "... = (root (card {..i i * q + (1-q))) ^ (card {.. ((\i i * q + (1-q))/card{..i i * q)/n + (1-q))^n" using assms(1) by (simp add:sum.distrib divide_simps mult.commute) also have "... = (\_avg * q + (1-q))^n" unfolding \_avg_def by (simp add: sum_distrib_right[symmetric]) finally have "(1-q) ^ r * ?L \ (\_avg * q + (1-q)) ^ n" by simp hence "?L \ (\_avg * q + (1-q)) ^ n / (1-q) ^ r" using 4 by (subst pos_le_divide_eq) (auto simp add:algebra_simps) also have "... \ ?R" - by (intro 3) + by (intro 3) finally show ?thesis by simp qed -text \The distribution of a random variable with a countable range is a discrete probability space, +text \The distribution of a random variable with a countable range is a discrete probability space, i.e., induces a PMF. Using this it is possible to generalize the previous result to arbitrary probability spaces.\ lemma (in prob_space) establish_pmf: fixes f :: "'a \ 'b" assumes rv: "random_variable discrete f" assumes "countable (f ` space M)" shows "distr M discrete f \ {M. prob_space M \ sets M = UNIV \ (AE x in M. measure M {x} \ 0)}" proof - define N where "N = {x \ space M.\ prob (f -` {f x} \ space M) \ 0}" define I where "I = {z \ (f ` space M). prob (f -` {z} \ space M) = 0}" - have countable_I: " countable I" + have countable_I: " countable I" unfolding I_def by (intro countable_subset[OF _ assms(2)]) auto - have disj: "disjoint_family_on (\y. f -` {y} \ space M) I" + have disj: "disjoint_family_on (\y. f -` {y} \ space M) I" unfolding disjoint_family_on_def by auto have N_alt_def: "N = (\y \ I. f -` {y} \ space M)" unfolding N_def I_def by (auto simp add:set_eq_iff) have "emeasure M N = \\<^sup>+ y. emeasure M (f -` {y} \ space M) \count_space I" - using rv countable_I unfolding N_alt_def - by (subst emeasure_UN_countable) (auto simp add:disjoint_family_on_def) + using rv countable_I unfolding N_alt_def + by (subst emeasure_UN_countable) (auto simp add:disjoint_family_on_def) also have "... = \\<^sup>+ y. 0 \count_space I" unfolding I_def using emeasure_eq_measure ennreal_0 - by (intro nn_integral_cong) auto + by (intro nn_integral_cong) auto also have "... = 0" by simp finally have 0:"emeasure M N = 0" by simp - have 1:"N \ events" + have 1:"N \ events" unfolding N_alt_def using rv by (intro sets.countable_UN'' countable_I) simp have " AE x in M. prob (f -` {f x} \ space M) \ 0" - using 0 1 by (subst AE_iff_measurable[OF _ N_def[symmetric]]) + using 0 1 by (subst AE_iff_measurable[OF _ N_def[symmetric]]) hence " AE x in M. measure (distr M discrete f) {f x} \ 0" by (subst measure_distr[OF rv], auto) hence "AE x in distr M discrete f. measure (distr M discrete f) {x} \ 0" by (subst AE_distr_iff[OF rv], auto) - thus ?thesis + thus ?thesis using prob_space_distr rv by auto qed lemma singletons_image_eq: "(\x. {x}) ` T \ Pow T" by auto theorem (in prob_space) impagliazzo_kabanets: fixes Y :: "nat \ 'a \ bool" assumes "n > 0" assumes "\i. i \ {.. random_variable discrete (Y i)" assumes "\i. i \ {.. \ i \ {0..1}" assumes "\S. S \ {.. \

(\ in M. (\i \ S. Y i \)) \ (\i \ S. \ i)" - defines "\_avg \ (\i\ {.. i)/n" + defines "\_avg \ (\i\ {.. i)/n" assumes "\ \ {\_avg..1}" "\_avg > 0" - shows "\

(\ in M. real (card {i \ {..}) \ \ * n) \ exp (-real n * KL_div \ \_avg)" + shows "\

(\ in M. real (card {i \ {..}) \ \ * n) \ exp (-real n * KL_div \ \_avg)" (is "?L \ ?R") proof - define f where "f = (\\ i. if i < n then Y i \ else False)" define g where "g = (\\ i. if i < n then \ i else False)" define T where "T = {\. (\i. \ i \ i < n)}" have g_idem: "g \ f = f" unfolding f_def g_def by (simp add:comp_def) - have f_range: " f \ space M \ T" + have f_range: " f \ space M \ T" unfolding T_def f_def by simp have "T = PiE_dflt {.._. UNIV)" unfolding T_def PiE_dflt_def by auto - hence "finite T" + hence "finite T" using finite_PiE_dflt by auto - hence countable_T: "countable T" + hence countable_T: "countable T" by (intro countable_finite) - moreover have "f ` space M \ T" + moreover have "f ` space M \ T" using f_range by auto ultimately have countable_f: "countable (f ` space M)" using countable_subset by auto have "f -` y \ space M \ events" if t:"y \ (\x. {x}) ` T" for y proof - obtain t where "y = {t}" and t_range: "t \ T" using t by auto - hence "f -` y \ space M = {\ \ space M. f \ = t}" + hence "f -` y \ space M = {\ \ space M. f \ = t}" by (auto simp add:vimage_def) also have "... = {\ \ space M. (\i < n. Y i \ = t i)}" using t_range unfolding f_def T_def by auto also have "... = (\i \ {.. \ space M. Y i \ = t i})" using assms(1) by auto also have "... \ events" using assms(1,2) by (intro sets.countable_INT) auto finally show ?thesis by simp qed hence "random_variable (count_space T) f" using sigma_sets_singletons[OF countable_T] singletons_image_eq f_range by (intro measurable_sigma_sets[where \="T" and A=" (\x. {x}) ` T"]) simp_all moreover have "g \ measurable discrete (count_space T)" unfolding g_def T_def by simp ultimately have "random_variable discrete (g \ f)" by simp hence rv:"random_variable discrete f" unfolding g_idem by simp - + define M' :: "(nat \ bool) measure" - where "M' = distr M discrete f" + where "M' = distr M discrete f" define \ where "\ = Abs_pmf M'" have a:"measure_pmf (Abs_pmf M') = M'" unfolding M'_def by (intro Abs_pmf_inverse[OF establish_pmf] rv countable_f) have b:"{i. (i < n \ Y i x) \ i < n} = {i. i < n \ Y i x}" for x by auto have c: "measure \ {\. \i\S. \ i} \ prod \ S" (is "?L1 \ ?R1") if "S \ {.. S \ i < n" for i + have d: "i \ S \ i < n" for i using that by auto have "?L1 = measure M' {\. \i\S. \ i}" unfolding \_def a by simp also have "... = \

(\ in M. (\i \ S. Y i \))" unfolding M'_def using that d by (subst measure_distr[OF rv]) (auto simp add:f_def Int_commute Int_def) also have "... \ ?R1" using that assms(4) by simp finally show ?thesis by simp qed have "?L = measure M' {\. real (card {i. i < n \ \ i}) \ \ * n}" - unfolding M'_def by (subst measure_distr[OF rv]) + unfolding M'_def by (subst measure_distr[OF rv]) (auto simp add:f_def algebra_simps Int_commute Int_def b) also have "... = measure_pmf.prob \ {\. real (card {i \ {.. i}) \ \ * n}" unfolding \_def a by simp also have "... \ ?R" using assms(1,3,6,7) c unfolding \_avg_def by (intro impagliazzo_kabanets_pmf) auto finally show ?thesis by simp qed end \ No newline at end of file diff --git a/thys/Expander_Graphs/Expander_Graphs_Algebra.thy b/thys/Expander_Graphs/Expander_Graphs_Algebra.thy --- a/thys/Expander_Graphs/Expander_Graphs_Algebra.thy +++ b/thys/Expander_Graphs/Expander_Graphs_Algebra.thy @@ -1,958 +1,958 @@ section \Algebra-only Theorems\ text \This section verifies the linear algebraic counter-parts of the graph-theoretic theorems about Random walks. The graph-theoretic results are then derived in Section~\ref{sec:random_walks}.\ theory Expander_Graphs_Algebra - imports + imports "HOL-Library.Monad_Syntax" Expander_Graphs_TTS begin -lemma pythagoras: +lemma pythagoras: fixes v w :: "'a::real_inner" assumes "v \ w = 0" - shows "norm (v+w)^2 = norm v^2 + norm w^2" + shows "norm (v+w)^2 = norm v^2 + norm w^2" using assms by (simp add:power2_norm_eq_inner algebra_simps inner_commute) definition diag :: "('a :: zero)^'n \ 'a^'n^'n" where "diag v = (\ i j. if i = j then (v $ i) else 0)" definition ind_vec :: "'n set \ real^'n" where "ind_vec S = (\ i. of_bool( i \ S))" lemma diag_mult_eq: "diag x ** diag y = diag (x * y)" - unfolding diag_def - by (vector matrix_matrix_mult_def) + unfolding diag_def + by (vector matrix_matrix_mult_def) (auto simp add:if_distrib if_distribR sum.If_cases) lemma diag_vec_mult_eq: "diag x *v y = x * y" - unfolding diag_def matrix_vector_mult_def + unfolding diag_def matrix_vector_mult_def by (simp add:if_distrib if_distribR sum.If_cases times_vec_def) definition matrix_norm_bound :: "real^'n^'m \ real \ bool" where "matrix_norm_bound A l = (\x. norm (A *v x) \ l * norm x)" lemma matrix_norm_boundI: assumes "\x. norm (A *v x) \ l * norm x" shows "matrix_norm_bound A l" using assms unfolding matrix_norm_bound_def by simp lemma matrix_norm_boundD: assumes "matrix_norm_bound A l" shows "norm (A *v x) \ l * norm x" using assms unfolding matrix_norm_bound_def by simp lemma matrix_norm_bound_nonneg: fixes A :: "real^'n^'m" assumes "matrix_norm_bound A l" - shows "l \ 0" + shows "l \ 0" proof - have "0 \ norm (A *v 1)" by simp - also have "... \ l * norm (1::real^'n)" + also have "... \ l * norm (1::real^'n)" using assms(1) unfolding matrix_norm_bound_def by simp finally have "0 \ l * norm (1::real^'n)" by simp moreover have "norm (1::real^'n) > 0" by simp - ultimately show ?thesis + ultimately show ?thesis by (simp add: zero_le_mult_iff) qed -lemma matrix_norm_bound_0: - assumes "matrix_norm_bound A 0" +lemma matrix_norm_bound_0: + assumes "matrix_norm_bound A 0" shows "A = (0::real^'n^'m)" proof (intro iffD2[OF matrix_eq] allI) fix x :: "real^'n" have "norm (A *v x) = 0" using assms unfolding matrix_norm_bound_def by simp thus "A *v x = 0 *v x" by simp qed lemma matrix_norm_bound_diag: fixes x :: "real^'n" assumes "\i. \x $ i\ \ l" shows "matrix_norm_bound (diag x) l" proof (rule matrix_norm_boundI) fix y :: "real^'n" have l_ge_0: "l \ 0" using assms by fastforce have a: "\x $ i * v\ \ \l * v\" for v i using l_ge_0 assms by (simp add:abs_mult mult_right_mono) have "norm (diag x *v y) = sqrt (\i \ UNIV. (x $ i * y $ i)^2)" unfolding matrix_vector_mult_def diag_def norm_vec_def L2_set_def by (auto simp add:if_distrib if_distribR sum.If_cases) also have "... \ sqrt (\i \ UNIV. (l * y $ i)^2)" by (intro real_sqrt_le_mono sum_mono iffD1[OF abs_le_square_iff] a) also have "... = l * norm y" - using l_ge_0 by (simp add:norm_vec_def L2_set_def algebra_simps + using l_ge_0 by (simp add:norm_vec_def L2_set_def algebra_simps sum_distrib_left[symmetric] real_sqrt_mult) finally show "norm (diag x *v y) \ l * norm y" by simp qed -lemma vector_scaleR_matrix_ac_2: "b *\<^sub>R (A::real^'n^'m) *v x = b *\<^sub>R (A *v x)" +lemma vector_scaleR_matrix_ac_2: "b *\<^sub>R (A::real^'n^'m) *v x = b *\<^sub>R (A *v x)" unfolding vector_transpose_matrix[symmetric] transpose_scalar by (intro vector_scaleR_matrix_ac) -lemma matrix_norm_bound_scale: +lemma matrix_norm_bound_scale: assumes "matrix_norm_bound A l" shows "matrix_norm_bound (b *\<^sub>R A) (\b\ * l)" proof (intro matrix_norm_boundI) fix x - have "norm (b *\<^sub>R A *v x) = norm (b *\<^sub>R (A *v x))" + have "norm (b *\<^sub>R A *v x) = norm (b *\<^sub>R (A *v x))" by (metis transpose_scalar vector_scaleR_matrix_ac vector_transpose_matrix) - also have "... = \b\ * norm (A *v x)" + also have "... = \b\ * norm (A *v x)" by simp also have "... \ \b\ * (l * norm x)" using assms matrix_norm_bound_def by (intro mult_left_mono) auto also have "... \ (\b\ * l) * norm x" by simp finally show "norm (b *\<^sub>R A *v x) \ (\b\ * l) * norm x" by simp qed definition nonneg_mat :: "real^'n^'m \ bool" where "nonneg_mat A = (\i j. A $ i $ j \ 0)" lemma nonneg_mat_1: shows "nonneg_mat (mat 1)" unfolding nonneg_mat_def mat_def by auto lemma nonneg_mat_prod: assumes "nonneg_mat A" "nonneg_mat B" shows "nonneg_mat (A ** B)" - using assms unfolding nonneg_mat_def matrix_matrix_mult_def + using assms unfolding nonneg_mat_def matrix_matrix_mult_def by (auto intro:sum_nonneg) lemma nonneg_mat_transpose: "nonneg_mat (transpose A) = nonneg_mat A" - unfolding nonneg_mat_def transpose_def + unfolding nonneg_mat_def transpose_def by auto definition spec_bound :: "real^'n^'n \ real \ bool" where "spec_bound M l = (l \ 0 \ (\v. v \ 1 = 0 \ norm (M *v v) \ l * norm v))" lemma spec_boundD1: assumes "spec_bound M l" - shows "0 \ l" + shows "0 \ l" using assms unfolding spec_bound_def by simp lemma spec_boundD2: assumes "spec_bound M l" assumes "v \ 1 = 0 " - shows "norm (M *v v) \ l * norm v" + shows "norm (M *v v) \ l * norm v" using assms unfolding spec_bound_def by simp lemma spec_bound_mono: assumes "spec_bound M \" "\ \ \" shows "spec_bound M \" proof - have "norm (M *v v) \ \ * norm v" if "inner v 1 = 0" for v proof - - have "norm (M *v v) \ \ * norm v" + have "norm (M *v v) \ \ * norm v" by (intro spec_boundD2[OF assms(1)] that) also have "... \ \ * norm v" by (intro mult_right_mono assms(2)) auto finally show ?thesis by simp qed moreover have "\ \ 0" using assms(2) spec_boundD1[OF assms(1)] by simp - ultimately show ?thesis + ultimately show ?thesis unfolding spec_bound_def by simp qed definition markov :: "real^'n^'n \ bool" where "markov M = (nonneg_mat M \ M *v 1 = 1 \ 1 v* M = 1)" lemma markov_symI: assumes "nonneg_mat A" "transpose A = A" "A *v 1 = 1" shows "markov A" proof - have "1 v* A = transpose A *v 1" unfolding vector_transpose_matrix[symmetric] by simp also have "... = 1" unfolding assms(2,3) by simp finally have "1 v* A = 1" by simp thus ?thesis unfolding markov_def using assms by auto qed lemma markov_apply: assumes "markov M" shows "M *v 1 = 1" "1 v* M = 1" using assms unfolding markov_def by auto lemma markov_transpose: "markov A = markov (transpose A)" unfolding markov_def nonneg_mat_transpose by auto -fun matrix_pow where +fun matrix_pow where "matrix_pow M 0 = mat 1" | "matrix_pow M (Suc n) = M ** (matrix_pow M n)" -lemma markov_orth_inv: +lemma markov_orth_inv: assumes "markov A" shows "inner (A *v x) 1 = inner x 1" proof - have "inner (A *v x) 1 = inner x (1 v* A)" using dot_lmul_matrix inner_commute by metis also have "... = inner x 1" using markov_apply[OF assms(1)] by simp finally show ?thesis by simp qed lemma markov_id: "markov (mat 1)" unfolding markov_def using nonneg_mat_1 by simp lemma markov_mult: assumes "markov A" "markov B" shows "markov (A ** B)" proof - have "nonneg_mat (A ** B)" using assms unfolding markov_def by (intro nonneg_mat_prod) auto - moreover have "(A ** B) *v 1 = 1" + moreover have "(A ** B) *v 1 = 1" using assms unfolding markov_def unfolding matrix_vector_mul_assoc[symmetric] by simp - moreover have "1 v* (A ** B) = 1" + moreover have "1 v* (A ** B) = 1" using assms unfolding markov_def unfolding vector_matrix_mul_assoc[symmetric] by simp ultimately show ?thesis unfolding markov_def by simp qed lemma markov_matrix_pow: assumes "markov A" shows "markov (matrix_pow A k)" using markov_id assms markov_mult by (induction k, auto) -lemma spec_bound_prod: +lemma spec_bound_prod: assumes "markov A" "markov B" assumes "spec_bound A la" "spec_bound B lb" shows "spec_bound (A ** B) (la*lb)" proof - have la_ge_0: "la \ 0" using spec_boundD1[OF assms(3)] by simp have "norm ((A ** B) *v x) \ (la * lb) * norm x" if "inner x 1 = 0" for x proof - have "norm ((A ** B) *v x) = norm (A *v (B *v x))" by (simp add:matrix_vector_mul_assoc) also have "... \ la * norm (B *v x)" by (intro spec_boundD2[OF assms(3)]) (simp add:markov_orth_inv that assms(2)) - also have "... \ la * (lb * norm x)" + also have "... \ la * (lb * norm x)" by (intro spec_boundD2[OF assms(4)] mult_left_mono that la_ge_0) finally show ?thesis by simp qed moreover have "la * lb \ 0" using la_ge_0 spec_boundD1[OF assms(4)] by simp ultimately show ?thesis using spec_bound_def by auto qed -lemma spec_bound_pow: +lemma spec_bound_pow: assumes "markov A" assumes "spec_bound A l" shows "spec_bound (matrix_pow A k) (l^k)" proof (induction k) case 0 then show ?case unfolding spec_bound_def by simp next case (Suc k) have "spec_bound (A ** matrix_pow A k) (l * l ^ k)" by (intro spec_bound_prod assms Suc markov_matrix_pow) thus ?case by simp qed fun intersperse :: "'a \ 'a list \ 'a list" - where + where "intersperse x [] = []" | "intersperse x (y#[]) = y#[]" | "intersperse x (y#z#zs) = y#x#intersperse x (z#zs)" lemma intersperse_snoc: assumes "xs \ []" shows "intersperse z (xs@[y]) = intersperse z xs@[z,y]" using assms proof (induction xs rule:list_nonempty_induct) case (single x) then show ?case by simp next case (cons x xs) then obtain xsh xst where t:"xs = xsh#xst" by (metis neq_Nil_conv) have "intersperse z ((x # xs) @ [y]) = x#z#intersperse z (xs@[y])" unfolding t by simp also have "... = x#z#intersperse z xs@[z,y]" using cons by simp also have "... = intersperse z (x#xs)@[z,y]" unfolding t by simp finally show ?case by simp qed lemma foldl_intersperse: assumes "xs \ []" shows "foldl f a ((intersperse x xs)@[x]) = foldl (\y z. f (f y z) x) a xs" using assms by (induction xs rule:rev_nonempty_induct) (auto simp add:intersperse_snoc) lemma foldl_intersperse_2: shows "foldl f a (intersperse y (x#xs)) = foldl (\x z. f (f x y) z) (f a x) xs" proof (induction xs rule:rev_induct) case Nil then show ?case by simp next case (snoc xst xs) - have "foldl f a (intersperse y ((x # xs) @ [xst])) = foldl (\x. f (f x y)) (f a x) (xs @ [xst])" + have "foldl f a (intersperse y ((x # xs) @ [xst])) = foldl (\x. f (f x y)) (f a x) (xs @ [xst])" by (subst intersperse_snoc, auto simp add:snoc) then show ?case by simp qed context regular_graph_tts begin definition stat :: "real^'n" where "stat = (1 / real CARD('n)) *\<^sub>R 1" definition J :: "('c :: field)^'n^'n" where "J = (\ i j. of_nat 1 / of_nat CARD('n))" lemma inner_1_1: "1 \ (1::real^'n) = CARD('n)" unfolding inner_vec_def by simp definition proj_unit :: "real^'n \ real^'n" where "proj_unit v = (1 \ v) *\<^sub>R stat" -definition proj_rem :: "real^'n \ real^'n" +definition proj_rem :: "real^'n \ real^'n" where "proj_rem v = v - proj_unit v" lemma proj_rem_orth: "1 \ (proj_rem v) = 0" unfolding proj_rem_def proj_unit_def inner_diff_right stat_def by (simp add:inner_1_1) -lemma split_vec: "v = proj_unit v + proj_rem v" +lemma split_vec: "v = proj_unit v + proj_rem v" unfolding proj_rem_def by simp lemma apply_J: "J *v x = proj_unit x" proof (intro iffD2[OF vec_eq_iff] allI) fix i - have "(J *v x) $ i = inner (\ j. 1 / real CARD('n)) x" + have "(J *v x) $ i = inner (\ j. 1 / real CARD('n)) x" unfolding matrix_vector_mul_component J_def by simp also have "... = inner stat x" unfolding stat_def scaleR_vec_def by auto also have "... = (proj_unit x) $ i" unfolding proj_unit_def stat_def by simp finally show "(J *v x) $ i = (proj_unit x) $ i" by simp -qed +qed lemma spec_bound_J: "spec_bound (J :: real^'n^'n) 0" proof - have "norm (J *v v) = 0" if "inner v 1 = 0" for v :: "real^'n" proof - have "inner (proj_unit v + proj_rem v) 1 = 0" using that by (subst (asm) split_vec[of "v"], simp) hence "inner (proj_unit v) 1 = 0" - using proj_rem_orth inner_commute unfolding inner_add_left + using proj_rem_orth inner_commute unfolding inner_add_left by (metis add_cancel_left_right) hence "proj_unit v = 0" unfolding proj_unit_def stat_def by simp hence "J *v v = 0" unfolding apply_J by simp thus ?thesis by simp qed thus ?thesis unfolding spec_bound_def by simp qed lemma matrix_decomposition_lemma_aux: fixes A :: "real^'n^'n" assumes "markov A" shows "spec_bound A l \ matrix_norm_bound (A - (1-l) *\<^sub>R J) l" (is "?L \ ?R") -proof +proof assume a:"?L" - hence l_ge_0: "l \ 0" using spec_boundD1 by auto - show "?R" + hence l_ge_0: "l \ 0" using spec_boundD1 by auto + show "?R" proof (rule matrix_norm_boundI) fix x :: "real^'n" have "(A - (1-l) *\<^sub>R J) *v x = A *v x - (1-l) *\<^sub>R (proj_unit x)" by (simp add:algebra_simps vector_scaleR_matrix_ac_2 apply_J) also have "... = A *v proj_unit x + A *v proj_rem x - (1-l) *\<^sub>R (proj_unit x)" by (subst split_vec[of "x"], simp add:algebra_simps) also have "... = proj_unit x + A *v proj_rem x - (1-l) *\<^sub>R (proj_unit x)" using markov_apply[OF assms(1)] unfolding proj_unit_def stat_def by (simp add:algebra_simps) also have "... = A *v proj_rem x + l *\<^sub>R proj_unit x" (is "_ = ?R1") by (simp add:algebra_simps) finally have d:"(A - (1-l) *\<^sub>R J) *v x = ?R1" by simp - have "inner (l *\<^sub>R proj_unit x) (A *v proj_rem x) = + have "inner (l *\<^sub>R proj_unit x) (A *v proj_rem x) = inner ((l * inner 1 x / real CARD('n)) *\<^sub>R 1 v* A) (proj_rem x)" - by (subst dot_lmul_matrix[symmetric]) (simp add:proj_unit_def stat_def) - also have "... = (l * inner 1 x / real CARD('n)) * inner 1 (proj_rem x)" + by (subst dot_lmul_matrix[symmetric]) (simp add:proj_unit_def stat_def) + also have "... = (l * inner 1 x / real CARD('n)) * inner 1 (proj_rem x)" unfolding scaleR_vector_matrix_assoc markov_apply[OF assms] by simp also have "... = 0" unfolding proj_rem_orth by simp finally have b:"inner (l *\<^sub>R proj_unit x) (A *v proj_rem x) = 0" by simp - have c: "inner (proj_rem x) (proj_unit x) = 0" + have c: "inner (proj_rem x) (proj_unit x) = 0" using proj_rem_orth[of "x"] unfolding proj_unit_def stat_def by (simp add:inner_commute) - have "norm (?R1)^2 = norm (A *v proj_rem x)^2 + norm (l *\<^sub>R proj_unit x)^2" + have "norm (?R1)^2 = norm (A *v proj_rem x)^2 + norm (l *\<^sub>R proj_unit x)^2" using b by (intro pythagoras) (simp add:inner_commute) - also have "... \ (l * norm (proj_rem x))^2 + norm (l *\<^sub>R proj_unit x)^2" + also have "... \ (l * norm (proj_rem x))^2 + norm (l *\<^sub>R proj_unit x)^2" using proj_rem_orth[of "x"] by (intro add_mono power_mono spec_boundD2 a) (auto simp add:inner_commute) also have "... = l^2 * (norm (proj_rem x)^2 + norm (proj_unit x)^2)" by (simp add:algebra_simps) also have "... = l^2 * (norm (proj_rem x + proj_unit x)^2)" using c by (subst pythagoras) auto also have "... = l^2 * norm x^2" by (subst (3) split_vec[of "x"]) (simp add:algebra_simps) also have "... = (l * norm x)^2" by (simp add:algebra_simps) finally have "norm (?R1)^2 \ (l * norm x)^2" by simp hence "norm (?R1) \ l * norm x" using l_ge_0 by (subst (asm) power_mono_iff) auto thus "norm ((A - (1-l) *\<^sub>R J) *v x) \ l * norm x" unfolding d by simp qed -next - assume a:"?R" - have "norm (A *v x) \ l * norm x" if "inner x 1 = 0" for x +next + assume a:"?R" + have "norm (A *v x) \ l * norm x" if "inner x 1 = 0" for x proof - - have "(1 - l) *\<^sub>R J *v x = (1 - l) *\<^sub>R (proj_unit x)" + have "(1 - l) *\<^sub>R J *v x = (1 - l) *\<^sub>R (proj_unit x)" by (simp add:vector_scaleR_matrix_ac_2 apply_J) also have "... = 0" unfolding proj_unit_def using that by (simp add:inner_commute) finally have b: "(1 - l) *\<^sub>R J *v x = 0" by simp have "norm (A *v x) = norm ((A - (1-l) *\<^sub>R J) *v x + ((1-l) *\<^sub>R J) *v x)" by (simp add:algebra_simps) also have "... \ norm ((A - (1-l) *\<^sub>R J) *v x) + norm (((1-l) *\<^sub>R J) *v x)" by (intro norm_triangle_ineq) also have "... \ l * norm x + 0" using a b unfolding matrix_norm_bound_def by (intro add_mono, auto) also have "... = l * norm x" by simp finally show ?thesis by simp qed - moreover have "l \ 0" + moreover have "l \ 0" using a matrix_norm_bound_nonneg by blast - ultimately show "?L" + ultimately show "?L" unfolding spec_bound_def by simp qed lemma matrix_decomposition_lemma: fixes A :: "real^'n^'n" assumes "markov A" - shows "spec_bound A l \ (\E. A = (1-l) *\<^sub>R J + l *\<^sub>R E \ matrix_norm_bound E 1 \ l \ 0)" + shows "spec_bound A l \ (\E. A = (1-l) *\<^sub>R J + l *\<^sub>R E \ matrix_norm_bound E 1 \ l \ 0)" (is "?L \ ?R") proof - - have "?L \ matrix_norm_bound (A - (1-l) *\<^sub>R J) l" + have "?L \ matrix_norm_bound (A - (1-l) *\<^sub>R J) l" using matrix_decomposition_lemma_aux[OF assms] by simp also have "... \ ?R" proof assume a:"matrix_norm_bound (A - (1 - l) *\<^sub>R J) l" hence l_ge_0: "l \ 0" using matrix_norm_bound_nonneg by auto define E where "E = (1/l) *\<^sub>R (A - (1-l) *\<^sub>R J)" - have "A = J" if "l = 0" + have "A = J" if "l = 0" proof - have "matrix_norm_bound (A - J) 0" using a that by simp hence "A - J = 0" using matrix_norm_bound_0 by blast thus "A = J" by simp qed hence "A = (1-l) *\<^sub>R J + l *\<^sub>R E" unfolding E_def by simp - moreover have "matrix_norm_bound E 1" + moreover have "matrix_norm_bound E 1" proof (cases "l = 0") case True hence "E = 0" if "l = 0" unfolding E_def by simp thus "matrix_norm_bound E 1" if "l = 0" using that unfolding matrix_norm_bound_def by auto next case False hence "l > 0" using l_ge_0 by simp moreover have "matrix_norm_bound E (\1 / l\* l)" unfolding E_def by (intro matrix_norm_bound_scale a) ultimately show ?thesis by auto qed ultimately show ?R using l_ge_0 by auto next assume a:?R then obtain E where E_def: "A = (1 - l) *\<^sub>R J + l *\<^sub>R E" "matrix_norm_bound E 1" "l \ 0" by auto - have "matrix_norm_bound (l *\<^sub>R E) (abs l*1)" + have "matrix_norm_bound (l *\<^sub>R E) (abs l*1)" by (intro matrix_norm_bound_scale E_def(2)) - moreover have "l \ 0" using E_def by simp - moreover have " l *\<^sub>R E = (A - (1 - l) *\<^sub>R J)" + moreover have "l \ 0" using E_def by simp + moreover have " l *\<^sub>R E = (A - (1 - l) *\<^sub>R J)" using E_def(1) by simp ultimately show "matrix_norm_bound (A - (1 - l) *\<^sub>R J) l" - by simp + by simp qed finally show ?thesis by simp qed lemma hitting_property_alg: fixes S :: "('n :: finite) set" assumes l_range: "l \ {0..1}" defines "P \ diag (ind_vec S)" defines "\ \ card S / CARD('n)" assumes "\M. M \ set Ms \ spec_bound M l \ markov M" shows "foldl (\x M. P *v (M *v x)) (P *v stat) Ms \ 1 \ (\ + l * (1-\))^(length Ms+1)" proof - define t :: "real^'n" where "t = (\ i. of_bool (i \ S))" define r where "r = foldl (\x M. P *v (M *v x)) (P *v stat) Ms" have P_proj: "P ** P = P" unfolding P_def diag_mult_eq ind_vec_def by (intro arg_cong[where f="diag"]) (vector) have P_1_left: "1 v* P = t" unfolding P_def diag_def ind_vec_def vector_matrix_mult_def t_def by simp have P_1_right: "P *v 1 = t" unfolding P_def diag_def ind_vec_def matrix_vector_mult_def t_def by simp have P_norm :"matrix_norm_bound P 1" unfolding P_def ind_vec_def by (intro matrix_norm_bound_diag) simp - have norm_t: "norm t = sqrt (real (card S))" + have norm_t: "norm t = sqrt (real (card S))" unfolding t_def norm_vec_def L2_set_def of_bool_def by (simp add:sum.If_cases if_distrib if_distribR) - have \_range: "\ \ 0" "\ \ 1" - unfolding \_def by (auto simp add:card_mono) + have \_range: "\ \ 0" "\ \ 1" + unfolding \_def by (auto simp add:card_mono) - define condition :: "real^'n \ nat \ bool" - where "condition = (\x n. norm x \ (\ + l * (1-\))^n * sqrt (card S)/CARD('n) \ P *v x = x)" + define condition :: "real^'n \ nat \ bool" + where "condition = (\x n. norm x \ (\ + l * (1-\))^n * sqrt (card S)/CARD('n) \ P *v x = x)" have a:"condition r (length Ms)" unfolding r_def using assms(4) proof (induction Ms rule:rev_induct) case Nil have "norm (P *v stat) = (1 / real CARD('n)) * norm t" unfolding stat_def matrix_vector_mult_scaleR P_1_right by simp also have "... \ (1 / real CARD('n)) * sqrt (real (card S))" using norm_t by (intro mult_left_mono) auto also have "... = sqrt (card S)/CARD('n)" by simp finally have "norm (P *v stat) \ sqrt (card S)/CARD('n)" by simp moreover have "P *v (P *v stat) = P *v stat" unfolding matrix_vector_mul_assoc P_proj by simp ultimately show ?case unfolding condition_def by simp next case (snoc M xs) hence "spec_bound M l \ markov M" using snoc(2) by simp - then obtain E where E_def: "M = (1-l) *\<^sub>R J + l *\<^sub>R E" "matrix_norm_bound E 1" + then obtain E where E_def: "M = (1-l) *\<^sub>R J + l *\<^sub>R E" "matrix_norm_bound E 1" using iffD1[OF matrix_decomposition_lemma] by auto define y where "y = foldl (\x M. P *v (M *v x)) (P *v stat) xs" have b:"condition y (length xs)" using snoc unfolding y_def by simp hence a:"P *v y = y" using condition_def by simp have "norm (P *v (M *v y)) = norm (P *v ((1-l)*\<^sub>R J *v y) + P *v (l *\<^sub>R E *v y))" by (simp add:E_def algebra_simps) also have "... \ norm (P *v ((1-l)*\<^sub>R J *v y)) + norm (P *v (l *\<^sub>R E *v y)) " by (intro norm_triangle_ineq) also have "... = (1 - l) * norm (P *v (J *v y)) + l * norm (P *v (E *v y))" using l_range by (simp add:vector_scaleR_matrix_ac_2 matrix_vector_mult_scaleR) also have "... = (1-l) * \1 \ (P *v y)/real CARD('n)\ * norm t + l * norm (P *v (E *v y))" - by (subst a[symmetric]) + by (subst a[symmetric]) (simp add:apply_J proj_unit_def stat_def P_1_right matrix_vector_mult_scaleR) also have "... = (1-l) * \t \ y\/real CARD('n) * norm t + l * norm (P *v (E *v y))" by (subst dot_lmul_matrix[symmetric]) (simp add:P_1_left) also have "... \ (1-l) * (norm t * norm y) / real CARD('n) * norm t + l * (1 * norm (E *v y))" using P_norm Cauchy_Schwarz_ineq2 l_range by (intro add_mono mult_right_mono mult_left_mono divide_right_mono matrix_norm_boundD) auto also have "... = (1-l) * \ * norm y + l * norm (E *v y)" unfolding \_def norm_t by simp also have "... \ (1-l) * \ * norm y + l * (1 * norm y)" using \_range l_range by (intro add_mono matrix_norm_boundD mult_left_mono E_def) auto also have "... = (\ + l * (1-\)) * norm y" by (simp add:algebra_simps) also have "... \ (\ + l * (1-\)) * ((\ + l * (1-\))^length xs * sqrt (card S)/CARD('n))" using b \_range l_range unfolding condition_def by (intro mult_left_mono) auto also have "... = (\ + l * (1-\))^(length xs +1) * sqrt (card S)/CARD('n)" by simp finally have "norm (P *v (M *v y)) \ (\ + l * (1-\))^(length xs +1) * sqrt (card S)/CARD('n)" by simp moreover have "P *v (P *v (M *v y)) = P *v (M *v y)" - unfolding matrix_vector_mul_assoc matrix_mul_assoc P_proj + unfolding matrix_vector_mul_assoc matrix_mul_assoc P_proj by simp ultimately have "condition (P *v (M *v y)) (length (xs@[M]))" unfolding condition_def by simp - - then show ?case + + then show ?case unfolding y_def by simp qed have "inner r 1 = inner (P *v r) 1" using a condition_def by simp also have "... = inner (1 v* P) r" unfolding dot_lmul_matrix by (simp add:inner_commute) also have "... = inner t r" unfolding P_1_left by simp also have "... \ norm t * norm r" by (intro norm_cauchy_schwarz) also have "... \ sqrt (card S) * ((\ + l * (1-\))^(length Ms) * sqrt(card S)/CARD('n))" using a unfolding condition_def norm_t by (intro mult_mono) auto also have "... = (\ + 0) * ((\ + l * (1-\))^(length Ms))" by (simp add:\_def) also have "... \ (\ + l * (1-\)) * (\ + l * (1-\))^(length Ms)" using \_range l_range by (intro mult_right_mono zero_le_power add_mono) auto also have "... = (\ + l * (1-\))^(length Ms+1)" by simp - finally show ?thesis + finally show ?thesis unfolding r_def by simp qed lemma upto_append: assumes "i \ j" "j \ k" shows "[i.. (nat list \ nat)" - where "bool_list_split xs = foldl (\(ys,z) x. (if x then (ys@[z],0) else (ys,z+1))) ([],0) xs" + where "bool_list_split xs = foldl (\(ys,z) x. (if x then (ys@[z],0) else (ys,z+1))) ([],0) xs" lemma bool_list_split: assumes "bool_list_split xs = (ys,z)" shows "xs = concat (map (\k. replicate k False@[True]) ys)@replicate z False" using assms proof (induction xs arbitrary: ys z rule:rev_induct) case Nil then show ?case unfolding bool_list_split_def by simp next case (snoc x xs) - obtain u v where uv_def: "bool_list_split xs = (u,v)" + obtain u v where uv_def: "bool_list_split xs = (u,v)" by (metis surj_pair) - show ?case + show ?case proof (cases x) case True have a:"ys = u@[v]" "z = 0" using snoc(2) True uv_def unfolding bool_list_split_def by auto have "xs@[x] = concat (map (\k. replicate k False@[True]) u)@replicate v False@[True]" using snoc(1)[OF uv_def] True by simp also have "... = concat (map (\k. replicate k False@[True]) (u@[v]))@replicate 0 False" by simp also have "... = concat (map (\k. replicate k False@[True]) (ys))@replicate z False" using a by simp finally show ?thesis by simp next case False have a:"ys = u" "z = v+1" using snoc(2) False uv_def unfolding bool_list_split_def by auto have "xs@[x] = concat (map (\k. replicate k False@[True]) u)@replicate (v+1) False" using snoc(1)[OF uv_def] False unfolding replicate_add by simp also have "... = concat (map (\k. replicate k False@[True]) (ys))@replicate z False" using a by simp finally show ?thesis by simp qed qed lemma bool_list_split_count: assumes "bool_list_split xs = (ys,z)" shows "length (filter id xs) = length ys" unfolding bool_list_split[OF assms(1)] by (simp add:filter_concat comp_def) lemma foldl_concat: "foldl f a (concat xss) = foldl (\y xs. foldl f y xs) a xss" by (induction xss rule:rev_induct, auto) lemma hitting_property_alg_2: - fixes S :: "('n :: finite) set" and l :: nat + fixes S :: "('n :: finite) set" and l :: nat fixes M :: "real^'n^'n" assumes \_range: "\ \ {0..1}" assumes "I \ {.. (if i \ I then diag (ind_vec S) else mat 1)" defines "\ \ real (card S) / real (CARD('n))" assumes "spec_bound M \" "markov M" - shows + shows "foldl (\x M. M *v x) stat (intersperse M (map P [0.. 1 \ (\+\*(1-\))^card I" (is "?L \ ?R") proof (cases "I \ {}") case True define xs where "xs = map (\i. i \ I) [0..x. if x then Q else mat 1)" let ?rep = "(\x. replicate x (mat 1))" have P_eq: "P i = P' (i \ I)" for i unfolding P_def P'_def Q_def by simp - have "l > 0" + have "l > 0" using True assms(2) by auto - hence xs_ne: "xs \ []" + hence xs_ne: "xs \ []" unfolding xs_def by simp obtain ys z where ys_z: "bool_list_split xs = (ys,z)" by (metis surj_pair) - have "length ys = length (filter id xs)" + have "length ys = length (filter id xs)" using bool_list_split_count[OF ys_z] by simp also have "... = card (I \ {0.. 0" using True assms(2) by (metis card_gt_0_iff finite_nat_iff_bounded) - then obtain yh yt where ys_split: "ys = yh#yt" + then obtain yh yt where ys_split: "ys = yh#yt" by (metis length_greater_0_conv neq_Nil_conv) have a:"foldl (\x N. M *v (N *v x)) x (?rep z) \ 1 = x \ 1" for x proof (induction z) case 0 then show ?case by simp next case (Suc z) have "foldl (\x N. M *v (N *v x)) x (?rep (z+1)) \ 1 = x \ 1" - unfolding replicate_add using Suc + unfolding replicate_add using Suc by (simp add:markov_orth_inv[OF assms(6)]) then show ?case by simp qed - have "M *v stat = stat" + have "M *v stat = stat" using assms(6) unfolding stat_def matrix_vector_mult_scaleR markov_def by simp hence b: "foldl (\x N. M *v (N *v x)) stat (?rep yh) = stat" by (induction yh, auto) have "foldl (\x N. N *v (M *v x)) a (?rep x) = matrix_pow M x *v a" for x a proof (induction x) case 0 then show ?case by simp next case (Suc x) have "foldl (\x N. N *v (M *v x)) a (?rep (x+1)) = matrix_pow M (x+1) *v a" unfolding replicate_add using Suc by (simp add: matrix_vector_mul_assoc) then show ?case by simp qed hence c: "foldl (\x N. N *v (M *v x)) a (?rep x @ [Q]) = Q *v (matrix_pow M (x+1) *v a)" for x a by (simp add:matrix_vector_mul_assoc matrix_mul_assoc) have d: "spec_bound N \ \ markov N" if t1: "N \ set (map (\x. matrix_pow M (x + 1)) yt)" for N proof - obtain y where N_def: "N = matrix_pow M (y+1)" using t1 by auto hence d1: "spec_bound N (\^(y+1))" unfolding N_def using spec_bound_pow assms(5,6) by blast - have "spec_bound N (\^1)" - using \_range by (intro spec_bound_mono[OF d1] power_decreasing) auto + have "spec_bound N (\^1)" + using \_range by (intro spec_bound_mono[OF d1] power_decreasing) auto moreover have "markov N" unfolding N_def by (intro markov_matrix_pow assms(6)) ultimately show ?thesis by simp qed have "?L = foldl (\x M. M *v x) stat (intersperse M (map P' xs)) \ 1" unfolding P_eq xs_def map_map by (simp add:comp_def) also have "... = foldl (\x M. M *v x) stat (intersperse M (map P' xs)@[M]) \ 1" - by (simp add:markov_orth_inv[OF assms(6)]) + by (simp add:markov_orth_inv[OF assms(6)]) also have "... = foldl (\x N. M *v (N *v x)) stat (map P' xs) \ 1" using xs_ne by (subst foldl_intersperse) auto also have "... = foldl (\x N. M *v (N *v x)) stat ((ys \ (\x. ?rep x @ [Q])) @ ?rep z) \ 1" unfolding bool_list_split[OF ys_z] P'_def List.bind_def by (simp add: comp_def map_concat) also have "... = foldl (\x N. M *v (N *v x)) stat (ys \ (\x. ?rep x @ [Q])) \ 1" by (simp add: a) also have "... = foldl (\x N. M *v (N *v x)) stat (?rep yh @[Q]@(yt \(\x. ?rep x @ [Q]))) \ 1" unfolding ys_split by simp also have "... = foldl (\x N. M *v (N *v x)) stat ([Q]@(yt \(\x. ?rep x @ [Q]))) \ 1" by (simp add:b) also have "... = foldl (\x N. N *v x) stat (intersperse M (Q#(yt \(\x.?rep x@[Q])))@[M])\1" by (subst foldl_intersperse, auto) also have "... = foldl (\x N. N *v x) stat (intersperse M (Q#(yt \(\x.?rep x@[Q])))) \ 1" - by (simp add:markov_orth_inv[OF assms(6)]) - also have "... = foldl (\x N. N *v (M *v x)) (Q *v stat) (yt \(\x.?rep x@[Q])) \ 1" + by (simp add:markov_orth_inv[OF assms(6)]) + also have "... = foldl (\x N. N *v (M *v x)) (Q *v stat) (yt \(\x.?rep x@[Q])) \ 1" by (subst foldl_intersperse_2, simp) also have "... = foldl (\a x. foldl (\x N. N *v (M *v x)) a (?rep x @ [Q])) (Q *v stat) yt \ 1" unfolding List.bind_def foldl_concat foldl_map by simp also have "... = foldl (\a x. Q *v (matrix_pow M (x+1) *v a)) (Q *v stat) yt \ 1" unfolding c by simp also have "... = foldl (\a N. Q *v (N *v a)) (Q *v stat) (map (\x. matrix_pow M (x+1)) yt) \ 1" by (simp add:foldl_map) also have "... \ (\ + \*(1-\))^(length (map (\x. matrix_pow M (x+1)) yt)+1)" unfolding \_def Q_def by (intro hitting_property_alg \_range d) simp - also have "... = (\ + \*(1-\))^(length ys)" + also have "... = (\ + \*(1-\))^(length ys)" unfolding ys_split by simp also have "... = ?R" unfolding len_ys by simp finally show ?thesis by simp next case False hence I_empty: "I = {}" by simp have "?L = stat \ (1 :: real^'n)" proof (cases "l > 0") case True have "?L = foldl (\x M. M *v x) stat ((intersperse M (map P [0.. 1" - by (simp add:markov_orth_inv[OF assms(6)]) + by (simp add:markov_orth_inv[OF assms(6)]) also have "... = foldl (\x N. M *v (N *v x)) stat (map P [0.. 1" using True by (subst foldl_intersperse, auto) also have "... = foldl (\x N. M *v (N *v x)) stat (map (\_. mat 1) [0.. 1" unfolding P_def using I_empty by simp also have "... = foldl (\x _. M *v x) stat [0.. 1" unfolding foldl_map by simp also have "... = stat \ (1 :: real^'n)" by (induction l, auto simp add:markov_orth_inv[OF assms(6)]) finally show ?thesis by simp next case False then show ?thesis by simp qed also have "... = 1" unfolding stat_def by (simp add:inner_vec_def) also have "... \ ?R" unfolding I_empty by simp finally show ?thesis by simp qed lemma uniform_property_alg: fixes x :: "('n :: finite)" and l :: nat assumes "i < l" defines "P j \ (if j = i then diag (ind_vec {x}) else mat 1)" assumes "markov M" shows "foldl (\x M. M *v x) stat (intersperse M (map P [0.. 1 = 1 / CARD('n)" - (is "?L = ?R") + (is "?L = ?R") proof - have a:"l > 0" using assms(1) by simp have 0: "foldl (\x N. M *v (N *v x)) y (xs) \ 1 = y \ 1" if "set xs \ {mat 1}" for xs y using that proof (induction xs rule:rev_induct) case Nil then show ?case by simp next case (snoc x xs) - have "x = mat 1" + have "x = mat 1" using snoc(2) by simp hence "foldl (\x N. M *v (N *v x)) y (xs @ [x]) \ 1 = foldl (\x N. M *v (N *v x)) y xs \ 1" by (simp add:markov_orth_inv[OF assms(3)]) also have "... = y \ 1" using snoc(2) by (intro snoc(1)) auto finally show ?case by simp qed - have M_stat: "M *v stat = stat" + have M_stat: "M *v stat = stat" using assms(3) unfolding stat_def matrix_vector_mult_scaleR markov_def by simp hence 1: "(foldl (\x N. M *v (N *v x)) stat xs) = stat" if "set xs \ {mat 1}" for xs using that by (induction xs, auto) have "?L = foldl (\x M. M *v x) stat ((intersperse M (map P [0.. 1" - by (simp add:markov_orth_inv[OF assms(3)]) + by (simp add:markov_orth_inv[OF assms(3)]) also have "... = foldl (\x N. M *v (N *v x)) stat (map P [0.. 1" using a by (subst foldl_intersperse) auto also have "... = foldl (\x N. M *v (N *v x)) stat (map P ([0.. 1" using assms(1) by (subst upto_append) auto also have "... = foldl (\x N. M *v (N *v x)) stat (map P [0.. 1" unfolding map_append foldl_append P_def by (subst 0) auto also have "... = foldl (\x N. M *v (N *v x)) stat (map P ([0.. 1" by simp also have "... = (M *v (diag (ind_vec {x}) *v stat)) \ 1" unfolding map_append foldl_append P_def by (subst 1) auto - also have "... = (diag (ind_vec {x}) *v stat) \ 1" - by (simp add:markov_orth_inv[OF assms(3)]) - also have "... = ((1/CARD('n)) *\<^sub>R ind_vec {x}) \ 1" - unfolding diag_def ind_vec_def stat_def matrix_vector_mult_def - by (intro arg_cong2[where f="(\)"] refl) + also have "... = (diag (ind_vec {x}) *v stat) \ 1" + by (simp add:markov_orth_inv[OF assms(3)]) + also have "... = ((1/CARD('n)) *\<^sub>R ind_vec {x}) \ 1" + unfolding diag_def ind_vec_def stat_def matrix_vector_mult_def + by (intro arg_cong2[where f="(\)"] refl) (vector of_bool_def sum.If_cases if_distrib if_distribR) also have "... = (1/CARD('n)) * (ind_vec {x} \ 1)" by simp also have "... = (1/CARD('n)) * 1" - unfolding inner_vec_def ind_vec_def of_bool_def + unfolding inner_vec_def ind_vec_def of_bool_def by (intro arg_cong2[where f="(*)"] refl) (simp) finally show ?thesis by simp qed end lemma foldl_matrix_mult_expand: fixes Ms :: "(('r::{semiring_1,comm_monoid_mult})^'a^'a) list" - shows "(foldl (\x M. M *v x) a Ms) $ k = (\x | length x = length Ms+1 \ x! length Ms = k. + shows "(foldl (\x M. M *v x) a Ms) $ k = (\x | length x = length Ms+1 \ x! length Ms = k. (\ i< length Ms. (Ms ! i) $ (x ! (i+1)) $ (x ! i)) * a $ (x ! 0))" proof (induction Ms arbitrary: k rule:rev_induct) case Nil have "length x = Suc 0 \ x = [x!0]" for x :: "'a list" by (cases x, auto) - hence "{x. length x = Suc 0 \ x ! 0 = k} = {[k]}" - by auto + hence "{x. length x = Suc 0 \ x ! 0 = k} = {[k]}" + by auto thus ?case by auto next case (snoc M Ms) let ?l = "length Ms" have 0: "finite {w. length w = Suc (length Ms) \ w ! length Ms = i}" for i :: 'a using finite_lists_length_eq[where A="UNIV::'a set" and n="?l +1"] by simp have "take (?l+1) x @ [x ! (?l+1)] = x" if "length x = ?l+2" for x :: "'a list" proof - have "take (?l+1) x @ [x ! (?l+1)] = take (Suc (?l+1)) x" using that by (intro take_Suc_conv_app_nth[symmetric], simp) - also have "... = x" + also have "... = x" using that by simp finally show ?thesis by simp qed hence 1: "bij_betw (take (?l+1)) {w. length w=?l+2 \ w!(?l+1) =k} {w. length w = ?l+1}" by (intro bij_betwI[where g="\x. x@[k]"]) (auto simp add:nth_append) have "foldl (\x M. M *v x) a (Ms @ [M]) $ k = (\j\UNIV. M$k$j *(foldl (\x M. M *v x) a Ms $ j))" by (simp add:matrix_vector_mult_def) - also have "... = + also have "... = (\j\UNIV. M$k$j * (\w|length w=?l+1\w!?l=j. (\ij\UNIV. (\w|length w=?l+1\w!?l=j. M$k$w!?l * (\iw\ (\j \ UNIV. {w. length w=?l+1 \ w!?l =j}). + also have "... = (\w\ (\j \ UNIV. {w. length w=?l+1 \ w!?l =j}). M$k$w!?l*(\iw | length w=?l+1. M$k$(w!?l)*(\iw \ take (?l+1) ` {w. length w=?l+2 \ w!(?l+1) =k}. + also have "... = (\w \ take (?l+1) ` {w. length w=?l+2 \ w!(?l+1) =k}. M$k$w!?l*(\iw|length w=?l+2\w!(?l+1)=k. M$k$w!?l*(\iw|length w=?l+2\w!(?l+1)=k. + also have "... = (\w|length w=?l+2\w!(?l+1)=k. (Ms@[M])!?l$k$w!?l*(\iw|length w=?l+2\w!(?l+1)=k. (\i<(?l+1). (Ms@[M])!i $ w!(i+1) $ w!i)* a$w!0)" by (intro sum.cong, auto simp add:algebra_simps) - finally have "foldl (\x M. M *v x) a (Ms @ [M]) $ k = + finally have "foldl (\x M. M *v x) a (Ms @ [M]) $ k = (\ w | length w = ?l+2 \ w ! (?l+1) = k. (\i<(?l+1). (Ms@[M])!i $ w!(i+1) $ w!i)* a$w!0)" by simp then show ?case by simp qed lemma foldl_matrix_mult_expand_2: fixes Ms :: "(real^'a^'a) list" - shows "(foldl (\x M. M *v x) a Ms) \ 1 = (\x | length x = length Ms+1. + shows "(foldl (\x M. M *v x) a Ms) \ 1 = (\x | length x = length Ms+1. (\ i< length Ms. (Ms ! i) $ (x ! (i+1)) $ (x ! i)) * a $ (x ! 0))" (is "?L = ?R") proof - let ?l = "length Ms" have "?L = (\j \ UNIV. (foldl (\x M. M *v x) a Ms) $ j)" by (simp add:inner_vec_def) also have "... = (\j\UNIV. \x|length x=?l+1 \ x!?l=j.(\ix \ (\j\ UNIV.{w. length w = length Ms+1 \ w ! length Ms = j}). (\ i< length Ms. (Ms ! i) $ (x ! (i+1)) $ (x ! i)) * a $ (x ! 0))" using finite_lists_length_eq[where A="UNIV::'a set" and n="?l +1"] by (intro sum.UNION_disjoint[symmetric]) auto also have "... = ?R" by (intro sum.cong, auto) finally show ?thesis by simp qed end diff --git a/thys/Expander_Graphs/Expander_Graphs_Cheeger_Inequality.thy b/thys/Expander_Graphs/Expander_Graphs_Cheeger_Inequality.thy --- a/thys/Expander_Graphs/Expander_Graphs_Cheeger_Inequality.thy +++ b/thys/Expander_Graphs/Expander_Graphs_Cheeger_Inequality.thy @@ -1,700 +1,699 @@ section \Cheeger Inequality\label{sec:cheeger_ineq}\ text \The Cheeger inequality relates edge expansion (a combinatorial property) with the second largest eigenvalue.\ theory Expander_Graphs_Cheeger_Inequality - imports - Expander_Graphs_Eigenvalues + imports Expander_Graphs_Eigenvalues begin unbundle intro_cong_syntax hide_const Quantum.T context regular_graph begin lemma edge_expansionD2: assumes "m = card (S \ verts G)" "2*m \ n" shows "\\<^sub>e * m \ real (card (edges_betw S (-S)))" proof - define S' where "S' = S \ verts G" have "\\<^sub>e * m = \\<^sub>e * card S'" using assms(1) S'_def by simp also have "... \ real (card (edges_betw S' (-S')))" using assms unfolding S'_def by (intro edge_expansionD) auto also have "... = real (card (edges_betw S (-S)))" unfolding S'_def edges_betw_def by (intro arg_cong[where f="real"] arg_cong[where f="card"]) auto finally show ?thesis by simp qed lemma edges_betw_sym: "card (edges_betw S T) = card (edges_betw T S)" (is "?L = ?R") proof - have "?L = (\a \ arcs G. of_bool (tail G a \ S \ head G a \ T))" unfolding edges_betw_def of_bool_def by (simp add:sum.If_cases Int_def) also have "... = (\e \# edges G. of_bool (fst e \ S \ snd e \ T))" unfolding sum_unfold_sum_mset edges_def arc_to_ends_def by (simp add:image_mset.compositionality comp_def) also have "... = (\e \# edges G. of_bool (snd e \ S \ fst e \ T))" - by (subst edges_sym[OF sym, symmetric]) + by (subst edges_sym[OF sym, symmetric]) (simp add:image_mset.compositionality comp_def case_prod_beta) also have "... = (\a \ arcs G. of_bool (tail G a \ T \ head G a \ S))" unfolding sum_unfold_sum_mset edges_def arc_to_ends_def by (simp add:image_mset.compositionality comp_def conj.commute) also have "... = ?R" unfolding edges_betw_def of_bool_def by (simp add:sum.If_cases Int_def) finally show ?thesis by simp qed lemma edges_betw_reg: assumes "S \ verts G" shows "card (edges_betw S UNIV) = card S * d" (is "?L = ?R") proof - have "?L = card (\(out_arcs G ` S))" unfolding edges_betw_def out_arcs_def by (intro arg_cong[where f="card"]) auto also have "... = (\i\S. card (out_arcs G i))" using finite_subset[OF assms] unfolding out_arcs_def by (intro card_UN_disjoint) auto also have "... = (\i\S. out_degree G i)" unfolding out_degree_def by simp also have "... = (\i\S. d)" using assms by (intro sum.cong reg) auto also have "... = ?R" by simp finally show ?thesis by simp qed text \The following proof follows Hoory et al.~@{cite \\S 4.5.1\ "hoory2006"}.\ lemma cheeger_aux_2: assumes "n > 1" shows "\\<^sub>e \ d*(1-\\<^sub>2)/2" proof - - have "real (card (edges_betw S (-S))) \ (d * (1 - \\<^sub>2) / 2) * real (card S)" + have "real (card (edges_betw S (-S))) \ (d * (1 - \\<^sub>2) / 2) * real (card S)" if "S \ verts G" "2 * card S \ n" for S proof - - let ?ct = "real (card (verts G - S))" - let ?cs = "real (card S)" + let ?ct = "real (card (verts G - S))" + let ?cs = "real (card S)" have "card (edges_betw S S)+card (edges_betw S (-S))=card(edges_betw S S\edges_betw S (-S))" unfolding edges_betw_def by (intro card_Un_disjoint[symmetric]) auto also have "... = card (edges_betw S UNIV)" unfolding edges_betw_def by (intro arg_cong[where f="card"]) auto also have "... = d * ?cs" using edges_betw_reg[OF that(1)] by simp - finally have "card (edges_betw S S) + card (edges_betw S (-S)) = d * ?cs" by simp + finally have "card (edges_betw S S) + card (edges_betw S (-S)) = d * ?cs" by simp hence 4: "card (edges_betw S S) = d * ?cs - card (edges_betw S (-S))" by simp have "card(edges_betw S(-S))+card(edges_betw(-S)(-S))=card(edges_betw S(-S)\edges_betw(-S)(-S))" unfolding edges_betw_def by (intro card_Un_disjoint[symmetric]) auto also have "... = card (edges_betw UNIV (verts G - S))" unfolding edges_betw_def by (intro arg_cong[where f="card"]) auto also have "... = card (edges_betw (verts G - S) UNIV)" by (intro edges_betw_sym) also have "... = d * ?ct" using edges_betw_reg by auto finally have "card (edges_betw S (-S)) + card (edges_betw (-S) (-S)) = d * ?ct" by simp hence 5: "card (edges_betw (-S) (-S)) = d * ?ct - card (edges_betw S (-S))" by simp - have 6: "card (edges_betw (-S) S) = card (edges_betw S (-S))" + have 6: "card (edges_betw (-S) S) = card (edges_betw S (-S))" by (intro edges_betw_sym) have "?cs + ?ct =real (card (S \ (verts G- S)))" unfolding of_nat_add[symmetric] using finite_subset[OF that(1)] by (intro_cong "[\\<^sub>1 of_nat, \\<^sub>1 card]" more:card_Un_disjoint[symmetric]) auto also have "... = real n" - unfolding n_def using that(1) by (intro_cong "[\\<^sub>1 of_nat, \\<^sub>1 card]") auto + unfolding n_def using that(1) by (intro_cong "[\\<^sub>1 of_nat, \\<^sub>1 card]") auto finally have 7: "?cs + ?ct = n" by simp - define f where + define f where "f x = real (card (verts G - S)) * of_bool (x \ S) - card S * of_bool (x \ S)" for x have "g_inner f (\_. 1) = ?cs * ?ct - real (card (verts G \ {x. x \ S})) * ?cs" unfolding g_inner_def f_def using Int_absorb1[OF that(1)] by (simp add:sum_subtractf) also have "... = ?cs * ?ct - ?ct * ?cs" by (intro_cong "[\\<^sub>2 (-), \\<^sub>2 (*), \\<^sub>1 of_nat, \\<^sub>1 card]") auto also have "... = 0" by simp finally have 11:" g_inner f (\_. 1) = 0" by simp have "g_norm f^2 = (\v\verts G. f v^2)" unfolding g_norm_sq g_inner_def conjugate_real_def by (simp add:power2_eq_square) also have "...=(\v\verts G. ?ct^2*(of_bool (v \ S))\<^sup>2)+(\v\verts G. ?cs^2*(of_bool (v \ S))\<^sup>2)" unfolding f_def power2_diff by (simp add:sum.distrib sum_subtractf power_mult_distrib) also have "... = real (card (verts G \ S))*?ct^2 + real (card (verts G \ {v. v \ S})) * ?cs^2" unfolding of_bool_def by (simp add:if_distrib if_distribR sum.If_cases) also have "... = real(card S)*(real(card(verts G-S)))\<^sup>2 + real(card(verts G-S))*(real(card S))\<^sup>2" using that(1) by (intro_cong "[\\<^sub>2(+), \\<^sub>2 (*), \\<^sub>2 power, \\<^sub>1 of_nat, \\<^sub>1 card]") auto also have "... = real(card S)*real (card (verts G -S))*(?cs + ?ct)" by (simp add:power2_eq_square algebra_simps) also have "... = real(card S)*real (card (verts G -S))*n" unfolding 7 by simp finally have 9:" g_norm f^2 = real(card S)*real (card (verts G -S))*real n" by simp - have "(\a \ arcs G. f (head G a) * f (tail G a)) = - (card (edges_betw S S) * ?ct*?ct) + (card (edges_betw (-S) (-S)) * ?cs*?cs) - + have "(\a \ arcs G. f (head G a) * f (tail G a)) = + (card (edges_betw S S) * ?ct*?ct) + (card (edges_betw (-S) (-S)) * ?cs*?cs) - (card (edges_betw S (-S)) * ?ct*?cs) - (card (edges_betw (-S) S) * ?cs*?ct)" - unfolding f_def by (simp add:of_bool_def algebra_simps Int_def if_distrib if_distribR + unfolding f_def by (simp add:of_bool_def algebra_simps Int_def if_distrib if_distribR edges_betw_def sum.If_cases) also have "... = d*?cs*?ct*(?cs+?ct) - card (edges_betw S (-S))*(?ct*?ct+2*?ct*?cs+?cs*?cs)" unfolding 4 5 6 by (simp add:algebra_simps) also have "... = d*?cs*?ct*n - (?ct+?cs)^2 * card (edges_betw S (-S))" unfolding power2_diff 7 power2_sum by (simp add:ac_simps power2_eq_square) also have "... = d *?cs*?ct*n - n^2 * card (edges_betw S (-S))" using 7 by (simp add:algebra_simps) - finally have 8:"(\a \ arcs G. f(head G a)*f(tail G a))=d*?cs*?ct*n-n^2*card(edges_betw S (-S))" + finally have 8:"(\a \ arcs G. f(head G a)*f(tail G a))=d*?cs*?ct*n-n^2*card(edges_betw S (-S))" by simp have "d*?cs*?ct*n-n^2*card(edges_betw S (-S)) = (\a \ arcs G. f (head G a) * f (tail G a))" unfolding 8 by simp also have "... \ d * (g_inner f (g_step f))" unfolding g_inner_step_eq using d_gt_0 by simp also have "... \ d * (\\<^sub>2 * g_norm f^2)" by (intro mult_left_mono os_expanderD 11) auto also have "... = d * \\<^sub>2 * ?cs*?ct*n" unfolding 9 by simp - finally have "d*?cs*?ct*n-n^2*card(edges_betw S (-S)) \ d * \\<^sub>2 * ?cs*?ct*n" + finally have "d*?cs*?ct*n-n^2*card(edges_betw S (-S)) \ d * \\<^sub>2 * ?cs*?ct*n" by simp hence "n * n * card (edges_betw S (-S)) \ n * (d * ?cs * ?ct * (1-\\<^sub>2))" by (simp add:power2_eq_square algebra_simps) hence 10:"n * card (edges_betw S (-S)) \ d * ?cs * ?ct * ( 1-\\<^sub>2)" - using n_gt_0 by simp + using n_gt_0 by simp have "(d * (1 - \\<^sub>2) / 2) * ?cs = (d * (1-\\<^sub>2) * (1 - 1 / 2)) * ?cs" by simp also have "... \ d * (1-\\<^sub>2) * ((n - ?cs) / n) * ?cs" - using that n_gt_0 \\<^sub>2_le_1 + using that n_gt_0 \\<^sub>2_le_1 by (intro mult_left_mono mult_right_mono mult_nonneg_nonneg) auto also have "... = (d * (1-\\<^sub>2) * ?ct / n) * ?cs" using 7 by simp also have "... = d * ?cs * ?ct * (1-\\<^sub>2) / n" by simp also have "... \ n * card (edges_betw S (-S)) / n" by (intro divide_right_mono 10) auto also have "... = card (edges_betw S (-S))" using n_gt_0 by simp finally show ?thesis by simp qed thus ?thesis by (intro edge_expansionI assms) auto qed end lemma surj_onI: assumes "\x. x \ B \ g x \ A \ f (g x) = x" shows "B \ f ` A" using assms by force lemma find_sorted_bij_1: fixes g :: "'a \ ('b :: linorder)" assumes "finite S" shows "\f. bij_betw f {.. mono_on {.. f)" proof - define h where "h x = from_nat_into S x" for x have h_bij:"bij_betw h {.. h) [0.. h) xs)" unfolding xs_def using sorted_sort_key by simp have "(\i. xs ! i) ` {..i. xs ! i) ` {..i. xs ! i) ` {.. (\i. xs ! i)) {.. f)" - using sorted_nth_mono[OF sorted_xs] l_xs unfolding f_def + have 2: "mono_on {.. f)" + using sorted_nth_mono[OF sorted_xs] l_xs unfolding f_def by (intro mono_onI) simp - thus ?thesis + thus ?thesis using 0 1 2 unfolding bij_betw_def by auto qed lemma find_sorted_bij_2: fixes g :: "'a \ ('b :: linorder)" assumes "finite S" shows "\f. bij_betw f S {.. (\x y. x \ S \ y \ S \ f x < f y \ g x \ g y)" proof - obtain f where f_def: "bij_betw f {.. f)" using find_sorted_bij_1 [OF assms] by auto define h where "h = the_inv_into {.. g y" if "h x < h y" "x \ S" "y \ S" for x y proof - have "h y < card S" "h x < card S" "h x \ h y" using bij_betw_apply[OF bij_h] that by auto hence "g (f (h x)) \ g (f (h y))" using f_def(2) unfolding mono_on_def by simp moreover have "f ` {.. g y" unfolding h_def using that f_the_inv_into_f[OF bij_betw_imp_inj_on[OF f_def(1)]] by auto qed ultimately show ?thesis by auto qed context regular_graph_tts begin text \Normalized Laplacian of the graph\ definition L where "L = mat 1 - A" lemma L_pos_semidefinite: fixes v :: "real ^'n" shows "v \ (L *v v) \ 0" proof - have "0 = v \ v - norm v^2" unfolding power2_norm_eq_inner by simp also have "... \ v \ v - abs (v \ (A *v v))" by (intro diff_mono rayleigh_bound) auto also have "... \ v \ v - v \ (A *v v)" by (intro diff_mono) auto also have "... = v \ (L *v v)" unfolding L_def by (simp add:algebra_simps) finally show ?thesis by simp qed text \The following proof follows Hoory et al.~@{cite \\S 4.5.2\ "hoory2006"}.\ lemma cheeger_aux_1: assumes "n > 1" shows "\\<^sub>e \ d * sqrt (2 * (1-\\<^sub>2))" proof - obtain v where v_def: "v \ 1 = 0" "v \ 0" "A *v v = \\<^sub>2 *s v" using \\<^sub>2_eq_\\<^sub>2 \\<^sub>2_ev[OF assms] by auto have "False" if "2*card {i. (1 *s v) $h i > 0} > n" "2*card {i. ((-1) *s v) $h i > 0} > n" proof - have "2 * n = n + n" by simp also have "... <2 * card {i. (1 *s v) $h i > 0} + 2 * card {i. ((-1) *s v) $h i > 0}" by (intro add_strict_mono that) also have "... = 2 * (card {i. (1 *s v) $h i > 0} + card {i. ((-1) *s v) $h i > 0})" by simp also have "... = 2 * (card ({i. (1 *s v) $h i > 0} \ {i. ((-1) *s v) $h i > 0}))" by (intro arg_cong2[where f="(*)"] card_Un_disjoint[symmetric]) auto also have "... \ 2 * (card (UNIV :: 'n set))" by (intro mult_left_mono card_mono) auto - finally have "2 * n < 2 * n" + finally have "2 * n < 2 * n" unfolding n_def card_n by auto thus ?thesis by simp qed then obtain \ :: real where \_def: "\ = 1 \ \=(-1)" "2* card {i. (\ *s v) $h i > 0 } \ n" unfolding not_le[symmetric] by blast define g where "g = \ *s v" - have g_orth: "g \ 1 = 0" unfolding g_def using v_def(1) + have g_orth: "g \ 1 = 0" unfolding g_def using v_def(1) by (simp add: scalar_mult_eq_scaleR) - have g_nz: "g \ 0" + have g_nz: "g \ 0" unfolding g_def using \_def(1) v_def(2) by auto have g_ev: "A *v g = \\<^sub>2 *s g" unfolding g_def scalar_mult_eq_scaleR matrix_vector_mult_scaleR v_def(3) by auto have g_supp: "2 * card { i. g $h i > 0 } \ n" unfolding g_def using \_def(2) by auto define f where "f = (\ i. max (g $h i) 0)" have "(L *v f) $h i \ (1-\\<^sub>2) * g $h i" (is "?L \ ?R") if "g $h i > 0" for i proof - have "?L = f $h i - (A *v f) $h i" unfolding L_def by (simp add:algebra_simps) also have "... = g $h i - (\j \ UNIV. A $h i $h j * f $h j)" unfolding matrix_vector_mult_def f_def using that by auto also have "... \ g $h i - (\j \ UNIV. A $h i $h j * g $h j)" unfolding f_def A_def by (intro diff_mono sum_mono mult_left_mono) auto - also have "... = g $h i - (A *v g) $h i" + also have "... = g $h i - (A *v g) $h i" unfolding matrix_vector_mult_def by simp also have "... = (1-\\<^sub>2) * g $h i" unfolding g_ev by (simp add:algebra_simps) finally show ?thesis by simp qed - moreover have "f $h i \ 0 \ g $h i > 0 "for i + moreover have "f $h i \ 0 \ g $h i > 0 "for i unfolding f_def by simp ultimately have 0:"(L *v f) $h i \ (1-\\<^sub>2) * g $h i \ f $h i = 0" for i by auto text \Part (i) in Hoory et al. (\S 4.5.2) but the operator L here is normalized.\ have "f \ (L *v f) = (\i\UNIV. (L *v f) $h i * f $h i)" unfolding inner_vec_def by (simp add:ac_simps) also have "... \ (\i\UNIV. ((1-\\<^sub>2) * g $h i) * f $h i)" by (intro sum_mono mult_right_mono' 0) (simp add:f_def) also have "... = (\i\UNIV. (1-\\<^sub>2) * f $h i * f $h i)" unfolding f_def by (intro sum.cong refl) auto also have "... = (1-\\<^sub>2) * (f \ f)" unfolding inner_vec_def by (simp add:sum_distrib_left ac_simps) also have "... = (1 - \\<^sub>2) * norm f^2" by (simp add: power2_norm_eq_inner) finally have h_part_i: "f \ (L *v f) \ (1 - \\<^sub>2) * norm f^2" by simp define f' where "f' x = f $h (enum_verts_inv x)" for x have f'_alt: "f = (\ i. f' (enum_verts i))" unfolding f'_def Rep_inverse by simp define B\<^sub>f where "B\<^sub>f = (\a\arcs G. \f' (tail G a)^2-f' (head G a)^2\)" have "(x + y)^2 \ 2 * (x^2 + y^2)" for x y :: real proof - have "(x + y)^2 = (x^2 + y^2) + 2 * x * y" unfolding power2_sum by simp also have "... \ (x^2 + y^2) + (x^2 + y^2)" by (intro add_mono sum_squares_bound) auto finally show ?thesis by simp qed hence "(\a\arcs G.(f'(tail G a)+ f'(head G a))\<^sup>2)\(\a\arcs G. 2*(f'(tail G a)^2+f'(head G a)^2))" by (intro sum_mono) auto also have "... = 2*((\a\arcs G. f'(tail G a)^2) + (\a\arcs G. f'(head G a)^2))" by (simp add:sum_distrib_left) also have "... = 4 * d * g_norm f'^2" unfolding sum_arcs_tail[where f="\x. f' x^2"] sum_arcs_head[where f="\x. f' x^2"] g_norm_sq g_inner_def by (simp add:power2_eq_square) also have "... = 4 * d * norm f^2" unfolding g_norm_conv f'_alt by simp finally have 1: "(\i\arcs G. (f' (tail G i) + f' (head G i))\<^sup>2) \ 4*d* norm f^2" by simp - have "(\a\arcs G. (f' (tail G a) - f' (head G a))\<^sup>2) = (\a\arcs G. (f' (tail G a))\<^sup>2) + + have "(\a\arcs G. (f' (tail G a) - f' (head G a))\<^sup>2) = (\a\arcs G. (f' (tail G a))\<^sup>2) + (\a\arcs G. (f' (head G a))\<^sup>2) - 2* (\a\arcs G. f' (tail G a) * f' (head G a))" unfolding power2_diff by (simp add:sum_subtractf sum_distrib_left ac_simps) also have "... = 2 * (d * (\v\verts G. (f' v)\<^sup>2) - (\a\arcs G. f' (tail G a) * f' (head G a)))" unfolding sum_arcs_tail[where f="\x. f' x^2"] sum_arcs_head[where f="\x. f' x^2"] by simp also have "... = 2 * (d * g_inner f' f' - d * g_inner f' (g_step f'))" unfolding g_inner_step_eq using d_gt_0 by (intro_cong "[\\<^sub>2 (*), \\<^sub>2 (-)]") (auto simp:power2_eq_square g_inner_def ac_simps) also have "... = 2 * d * (g_inner f' f' -g_inner f' (g_step f'))" - by (simp add:algebra_simps) + by (simp add:algebra_simps) also have "... = 2 * d * (f \ f - f \ (A *v f))" unfolding g_inner_conv g_step_conv f'_alt by simp also have "... = 2 * d * (f \ (L *v f))" unfolding L_def by (simp add:algebra_simps) finally have 2:"(\a\arcs G. (f' (tail G a) - f' (head G a))\<^sup>2) = 2 * d * (f \ (L *v f))" by simp have "B\<^sub>f = (\a\arcs G. \f' (tail G a)+f' (head G a)\*\f' (tail G a)-f' (head G a)\)" unfolding B\<^sub>f_def abs_mult[symmetric] by (simp add:algebra_simps power2_eq_square) - also have "...\ L2_set (\a. f'(tail G a) + f'(head G a)) (arcs G) * + also have "...\ L2_set (\a. f'(tail G a) + f'(head G a)) (arcs G) * L2_set (\a. f' (tail G a) - f'(head G a)) (arcs G)" by (intro L2_set_mult_ineq) also have "... \ sqrt (4*d* norm f^2) * sqrt (2 * d * (f \ (L *v f)))" unfolding L2_set_def 2 - by (intro mult_right_mono iffD2[OF real_sqrt_le_iff] 1 real_sqrt_ge_zero + by (intro mult_right_mono iffD2[OF real_sqrt_le_iff] 1 real_sqrt_ge_zero mult_nonneg_nonneg L_pos_semidefinite) auto also have "... = 2 * sqrt 2 * d * norm f * sqrt (f \ (L *v f))" by (simp add:real_sqrt_mult) - finally have hoory_4_12: "B\<^sub>f \ 2 * sqrt 2 * d * norm f * sqrt (f \ (L *v f))" + finally have hoory_4_12: "B\<^sub>f \ 2 * sqrt 2 * d * norm f * sqrt (f \ (L *v f))" by simp text \The last statement corresponds to Lemma 4.12 in Hoory et al.\ - obtain \ :: "'a \ nat" where \_bij: "bij_betw \ (verts G) {.. :: "'a \ nat" where \_bij: "bij_betw \ (verts G) {.._dec: "\x y. x \ verts G \ y \ verts G \ \ x < \ y \ f' x \ f' y" unfolding n_def using find_sorted_bij_2[where S="verts G" and g="(\x. - f' x)"] by auto define \ where "\ = the_inv_into (verts G) \" have \_bij: "bij_betw \ {.._def by (intro bij_betw_the_inv_into \_bij) have "edges G = {# e \# edges G . \(fst e) \ \(snd e) \ \(fst e) = \(snd e) #}" by simp also have "... = {# e \# edges G . \(fst e) \ \(snd e) #} + {#e\#edges G. \(fst e)=\(snd e)#}" by (simp add:filter_mset_ex_predicates) also have "...={# e\#edges G. \(fst e)<\(snd e)\\(fst e)>\(snd e)#}+{#e\#edges G. fst e=snd e#}" using bij_betw_imp_inj_on[OF \_bij] edge_set by (intro arg_cong2[where f="(+)"] filter_mset_cong refl inj_on_eq_iff[where A="verts G"]) auto - also have "... = {#e \# edges G. \(fst e) < \ (snd e) #} + - {#e \# edges G. \(fst e) > \ (snd e) #} + + also have "... = {#e \# edges G. \(fst e) < \ (snd e) #} + + {#e \# edges G. \(fst e) > \ (snd e) #} + {#e \# edges G. fst e = snd e #}" by (intro arg_cong2[where f="(+)"] filter_mset_ex_predicates[symmetric]) auto - finally have edges_split: "edges G = {#e \# edges G. \(fst e) < \ (snd e) #} + - {#e \# edges G. \(fst e) > \ (snd e) #} + {#e \# edges G. fst e = snd e #}" + finally have edges_split: "edges G = {#e \# edges G. \(fst e) < \ (snd e) #} + + {#e \# edges G. \(fst e) > \ (snd e) #} + {#e \# edges G. fst e = snd e #}" by simp have \_lt_n: "\ x < n" if "x \ verts G" for x using bij_betw_apply[OF \_bij] that by auto have \_\_inv: "\ (\ x) = x" if "x \ verts G" for x unfolding \_def using bij_betw_imp_inj_on[OF \_bij] by (intro the_inv_into_f_f that) auto have \_\_inv: "\ (\ x) = x" if "x < n" for x unfolding \_def using bij_betw_imp_inj_on[OF \_bij] bij_betw_imp_surj_on[OF \_bij] that by (intro f_the_inv_into_f) auto define \ where "\ x = (if x < n then f' (\ x) else 0)" for x - have \_nonneg: "\ k \ 0" for k + have \_nonneg: "\ k \ 0" for k unfolding \_def f'_def f_def by auto have \_antimono: "\ k \ \ l" if " k < l" for k l proof (cases "l \ n") case True hence "\ l = 0" unfolding \_def by simp - then show ?thesis using \_nonneg by simp + then show ?thesis using \_nonneg by simp next case False - hence "\ l = f' (\ l)" + hence "\ l = f' (\ l)" unfolding \_def by simp also have "... \ f' (\ k)" - using \_\_inv False that - by (intro \_dec bij_betw_apply[OF \_bij]) auto + using \_\_inv False that + by (intro \_dec bij_betw_apply[OF \_bij]) auto also have "... = \ k" unfolding \_def using False that by simp finally show ?thesis by simp qed define m :: nat where "m = Min {i. \ i = 0 \ i \ n}" - have "\ n = 0" + have "\ n = 0" unfolding \_def by simp hence "m \ {i. \ i = 0 \ i \ n}" unfolding m_def by (intro Min_in) auto hence m_rel_1: "\ m = 0" and m_le_n: "m \ n" by auto have "\ k > 0" if "k < m" for k proof (rule ccontr) - assume "\(\ k > 0)" + assume "\(\ k > 0)" hence "\ k = 0" by (intro order_antisym \_nonneg) simp hence "k \ {i. \ i = 0 \ i \ n}" using that m_le_n by simp - hence "m \ k" + hence "m \ k" unfolding m_def by (intro Min_le) auto thus "False" using that by simp qed hence m_rel_2: "f' x > 0" if "x \ \ ` {.._def using m_le_n that by auto have "2 * m = 2 * card {.. ` {.. ` {.._bij]] by (intro_cong "[\\<^sub>2 (*)]" more:card_image[symmetric]) auto also have "... \ 2 * card {x \ verts G. f' x > 0}" using m_rel_2 bij_betw_apply[OF \_bij] m_le_n - by (intro mult_left_mono card_mono subsetI) auto + by (intro mult_left_mono card_mono subsetI) auto also have "... = 2 * card (enum_verts_inv ` {x \ verts G. f $h (enum_verts_inv x) > 0})" unfolding f'_def using Abs_inject - by (intro arg_cong2[where f="(*)"] card_image[symmetric] inj_onI) auto + by (intro arg_cong2[where f="(*)"] card_image[symmetric] inj_onI) auto also have "... = 2 * card {x. f $h x > 0}" - using Rep_inverse Rep_range unfolding f'_def by (intro_cong "[\\<^sub>2 (*), \\<^sub>1 card]" - more:subset_antisym image_subsetI surj_onI[where g="enum_verts"]) auto + using Rep_inverse Rep_range unfolding f'_def by (intro_cong "[\\<^sub>2 (*), \\<^sub>1 card]" + more:subset_antisym image_subsetI surj_onI[where g="enum_verts"]) auto also have "... = 2 * card {x. g $h x > 0}" unfolding f_def by (intro_cong "[\\<^sub>2 (*), \\<^sub>1 card]") auto - also have "... \ n" + also have "... \ n" by (intro g_supp) finally have m2_le_n: "2*m \ n" by simp have "\ k \ 0" if "k > m" for k using m_rel_1 \_antimono that by metis hence "\ k \ 0" if "k \ m" for k using m_rel_1 that by (cases "k > m") auto hence \_supp: "\ k = 0" if "k \ m" for k using that by (intro order_antisym \_nonneg) auto have 4: "\ v \ x \ v \ \ ` {..x}" if "v \ verts G" "x < n" for v x proof - - have "\ v \ x \ \ v \ {..x}" + have "\ v \ x \ \ v \ {..x}" by simp also have "... \ \ (\ v) \ \ ` {..x}" using bij_betw_imp_inj_on[OF \_bij] bij_betw_apply[OF \_bij] that by (intro inj_on_image_mem_iff[where B="{.. v \ \ ` {..x}" unfolding \_\_inv[OF that(1)] by simp finally show ?thesis by simp qed have "B\<^sub>f = (\a\arcs G. \f' (tail G a)^2 - f' (head G a)^2\)" unfolding B\<^sub>f_def by simp also have "... = (\e \# edges G. \f' (fst e)^2 - f' (snd e)^2\)" - unfolding edges_def arc_to_ends_def sum_unfold_sum_mset + unfolding edges_def arc_to_ends_def sum_unfold_sum_mset by (simp add:image_mset.compositionality comp_def) - also have "... = + also have "... = (\e\#{#e \# edges G. \ (fst e) < \ (snd e)#}. \(f' (fst e))\<^sup>2 - (f' (snd e))\<^sup>2\) + (\e\#{#e \# edges G. \ (snd e) < \ (fst e)#}. \(f' (fst e))\<^sup>2 - (f' (snd e))\<^sup>2\) + (\e\#{#e \# edges G. fst e = snd e#}. \(f' (fst e))\<^sup>2 - (f' (snd e))\<^sup>2\)" by (subst edges_split) simp also have "... = (\e\#{#e \# edges G. \ (snd e) < \ (fst e)#}. \(f' (fst e))\<^sup>2 - (f' (snd e))\<^sup>2\) + (\e\#{#e \# edges G. \ (snd e) < \ (fst e)#}. \(f' (snd e))\<^sup>2 - (f' (fst e))\<^sup>2\) + (\e\#{#e \# edges G. fst e = snd e#}. \(f' (fst e))\<^sup>2 - (f' (snd e))\<^sup>2\)" - by (subst edges_sym[OF sym, symmetric]) (simp add:image_mset.compositionality + by (subst edges_sym[OF sym, symmetric]) (simp add:image_mset.compositionality comp_def image_mset_filter_mset_swap[symmetric] case_prod_beta) - also have "... = + also have "... = (\e\#{#e \# edges G. \ (snd e) < \ (fst e)#}. \(f' (snd e))\<^sup>2 - (f' (fst e))\<^sup>2\) + (\e\#{#e \# edges G. \ (snd e) < \ (fst e)#}. \(f' (snd e))\<^sup>2 - (f' (fst e))\<^sup>2\) + (\e\#{#e \# edges G. fst e = snd e#}. 0)" by (intro_cong "[\\<^sub>2 (+), \\<^sub>1 sum_mset]" more:image_mset_cong) auto also have "... = 2 * (\e\#{#e\#edges G. \(snd e)<\(fst e)#}. \(f' (snd e))\<^sup>2 - (f' (fst e))\<^sup>2\)" by simp also have "... = 2 *(\a|a\arcs G\\(tail G a)>\(head G a). \f'(head G a)^2 - f'(tail G a)^2\)" - unfolding edges_def arc_to_ends_def sum_unfold_sum_mset + unfolding edges_def arc_to_ends_def sum_unfold_sum_mset by (simp add:image_mset.compositionality comp_def image_mset_filter_mset_swap[symmetric]) also have "... = 2 * (\a|a\arcs G\\(tail G a)>\(head G a). \\(\(head G a))^2 - \(\(tail G a))^2\)" unfolding \_def using \_\_inv \_lt_n by (intro arg_cong2[where f="(*)"] sum.cong refl) auto also have "... = 2 * (\a|a\arcs G\\(tail G a)>\(head G a). \(\(head G a))^2 - \(\(tail G a))^2)" using \_antimono power_mono \_nonneg by (intro arg_cong2[where f="(*)"] sum.cong refl abs_of_nonneg)(auto) also have "... = 2 * (\a|a\arcs G\\(tail G a)>\(head G a). (-(\(\(tail G a))^2)) - (-(\(\(head G a))^2)))" by (simp add:algebra_simps) - also have "... = 2 *(\a|a\arcs G\\(tail G a)>\(head G a). + also have "... = 2 *(\a|a\arcs G\\(tail G a)>\(head G a). (\i=\(head G a)..<\(tail G a). (-(\ (Suc i)^2)) - (-(\ i^2))))" by (intro arg_cong2[where f="(*)"] sum.cong refl sum_Suc_diff'[symmetric]) auto - also have "...=2*(\(a, i)\(SIGMA x:{a \ arcs G. \ (head G a) < \ (tail G a)}. - {\ (head G x)..<\ (tail G x)}). \ i^2 - \ (Suc i)^2)" + also have "...=2*(\(a, i)\(SIGMA x:{a \ arcs G. \ (head G a) < \ (tail G a)}. + {\ (head G x)..<\ (tail G x)}). \ i^2 - \ (Suc i)^2)" by (subst sum.Sigma) auto - also have "...=2*(\p\{(a,i).a \ arcs G\\(head G a)\i\i<\(tail G a)}. \(snd p)^2-\ (snd p+1)^2)" + also have "...=2*(\p\{(a,i).a \ arcs G\\(head G a)\i\i<\(tail G a)}. \(snd p)^2-\ (snd p+1)^2)" by (intro arg_cong2[where f="(*)"] sum.cong refl) (auto simp add:Sigma_def) - also have "...=2*(\p\{(i,a).a \ arcs G\\(head G a) \ i\i < \(tail G a)}. \(fst p)^2-\(fst p+1)^2)" + also have "...=2*(\p\{(i,a).a \ arcs G\\(head G a) \ i\i < \(tail G a)}. \(fst p)^2-\(fst p+1)^2)" by (intro sum.reindex_cong[where l="prod.swap"] arg_cong2[where f="(*)"]) auto also have "...=2* - (\(i, a)\(SIGMA x:{.. arcs G. \ (head G a)\x \ x<\(tail G a)}). \ i^2-\ (i+1)^2)" - using less_trans[OF _ \_lt_n] by (intro sum.cong arg_cong2[where f="(*)"]) auto - also have "...=2*(\ia|a\arcs G \\(head G a) \ i\i < \(tail G a). \ i^2 - \ (i+1)^2))" + (\(i, a)\(SIGMA x:{.. arcs G. \ (head G a)\x \ x<\(tail G a)}). \ i^2-\ (i+1)^2)" + using less_trans[OF _ \_lt_n] by (intro sum.cong arg_cong2[where f="(*)"]) auto + also have "...=2*(\ia|a\arcs G \\(head G a) \ i\i < \(tail G a). \ i^2 - \ (i+1)^2))" by (subst sum.Sigma) auto also have "...=2*(\iarcs G. \(head G a)\i\i<\(tail G a)} * (\ i^2 - \ (i+1)^2))" by simp also have "...=2*(\iarcs G. \(head G a)\i\\(\(tail G a)\i)} * (\ i^2 - \ (i+1)^2))" by (intro_cong "[\\<^sub>2 (*), \\<^sub>1 card, \\<^sub>1 of_nat]" more:sum.cong Collect_cong) auto also have "...=2*(\iarcs G. head G a\\`{..i}\tail G a\\`{..i}} * (\ i^2-\ (i+1)^2))" - using 4 + using 4 by (intro_cong "[\\<^sub>2 (*), \\<^sub>1 card, \\<^sub>1 of_nat, \\<^sub>2 (\)]" more:sum.cong restr_Collect_cong) auto also have "... = 2 * (\i`{..i}) (\`{..i}))) * (\ i^2-\ (i+1)^2))" unfolding edges_betw_def by (auto simp:conj.commute) also have "... = 2 * (\i`{..i}) (-\`{..i}))) * (\ i^2-\ (i+1)^2))" using edges_betw_sym by simp also have "... = 2 * (\i`{..i}) (-\`{..i}))) * (\ i^2-\ (i+1)^2))" using \_supp m_le_n by (intro sum.mono_neutral_right arg_cong2[where f="(*)"]) auto - finally have Bf_eq: + finally have Bf_eq: "B\<^sub>f = 2 * (\i`{..i}) (-\`{..i}))) * (\ i^2-\ (i+1)^2))" by simp - have 3:"card (\ ` {..i} \ verts G) = i + 1" if "i < m" for i + have 3:"card (\ ` {..i} \ verts G) = i + 1" if "i < m" for i proof - have "card (\ ` {..i} \ verts G) = card (\ ` {..i})" - using m_le_n that by (intro arg_cong[where f="card"] Int_absorb2 + using m_le_n that by (intro arg_cong[where f="card"] Int_absorb2 image_subsetI bij_betw_apply[OF \_bij]) auto also have "... = card {..i}" - using m_le_n that by (intro card_image + using m_le_n that by (intro card_image inj_on_subset[OF bij_betw_imp_inj_on[OF \_bij]]) auto also have "... = i+1" by simp finally show ?thesis by simp qed have "2 * \\<^sub>e * norm f^2 = 2 * \\<^sub>e * (g_norm f'^2)" unfolding g_norm_conv f'_alt by simp also have "... \ 2 * \\<^sub>e * (\v\ verts G. f' v^2)" unfolding g_norm_sq g_inner_def by (simp add:power2_eq_square) also have "... = 2 * \\<^sub>e * (\i i)^2)" by (intro arg_cong2[where f="(*)"] refl sum.reindex_bij_betw[symmetric] \_bij) also have "... = 2 * \\<^sub>e * (\i i^2)" unfolding \_def by (intro arg_cong2[where f="(*)"] refl sum.cong) auto also have "... = 2 * \\<^sub>e * (\i i^2)" using \_supp m_le_n by (intro sum.mono_neutral_cong_right arg_cong2[where f="(*)"] refl) auto also have "... \ 2 * \\<^sub>e * ((\i i^2) + (real 0 * \ 0^2 - m * \ m^2))" using \_supp[of "m"] by simp also have "... \ 2 * \\<^sub>e * ((\i i^2) + (\i i^2-(Suc i)*\ (Suc i)^2))" by (subst sum_lessThan_telescope'[symmetric]) simp also have "... \ 2 * (\i\<^sub>e * (i+1)) * (\ i^2-\ (i+1)^2))" by (simp add:sum_distrib_left algebra_simps sum.distrib[symmetric]) also have "... \ 2 * (\i`{..i}) (-\`{..i}))) * (\ i^2-\ (i+1)^2))" using \_nonneg \_antimono power_mono 3 m2_le_n by (intro mult_left_mono sum_mono mult_right_mono edge_expansionD2) auto also have "... = B\<^sub>f" unfolding Bf_eq by simp finally have hoory_4_13: "2 * \\<^sub>e * norm f^2 \ B\<^sub>f" by simp text \Corresponds to Lemma 4.13 in Hoory et al.\ - have f_nz: "f \ 0" + have f_nz: "f \ 0" proof (rule ccontr) - assume f_nz_assms: "\ (f \ 0)" + assume f_nz_assms: "\ (f \ 0)" have "g $h i \ 0" for i proof - - have "g $h i \ max (g $h i) 0" + have "g $h i \ max (g $h i) 0" by simp also have "... = 0" using f_nz_assms unfolding f_def vec_eq_iff by auto finally show ?thesis by simp qed moreover have "(\i \ UNIV. 0-g $h i) = 0" using g_orth unfolding sum_subtractf inner_vec_def by auto ultimately have "\x\UNIV. -(g $h x) = 0" by (intro iffD1[OF sum_nonneg_eq_0_iff]) auto thus "False" using g_nz unfolding vec_eq_iff by simp qed hence norm_f_gt_0: "norm f> 0" by simp have "\\<^sub>e * norm f * norm f \ sqrt 2 * real d * norm f * sqrt (f \ (L *v f))" using order_trans[OF hoory_4_13 hoory_4_12] by (simp add:power2_eq_square) hence "\\<^sub>e \ real d * sqrt 2 * sqrt (f \ (L *v f)) / norm f" using norm_f_gt_0 by (simp add:ac_simps divide_simps) also have "... \ real d * sqrt 2 * sqrt ((1 - \\<^sub>2) * (norm f)\<^sup>2) / norm f" by (intro mult_left_mono divide_right_mono real_sqrt_le_mono h_part_i) auto also have "... = real d * sqrt 2 * sqrt (1- \\<^sub>2)" using f_nz by (simp add:real_sqrt_mult) also have "... = d * sqrt (2 * (1-\\<^sub>2))" by (simp add:real_sqrt_mult[symmetric]) finally show ?thesis by simp qed end context regular_graph -begin +begin -lemmas (in regular_graph) cheeger_aux_1 = +lemmas (in regular_graph) cheeger_aux_1 = regular_graph_tts.cheeger_aux_1[OF eg_tts_1, - internalize_sort "'n :: finite", OF _ regular_graph_axioms, + internalize_sort "'n :: finite", OF _ regular_graph_axioms, unfolded remove_finite_premise, cancel_type_definition, OF verts_non_empty] theorem cheeger_inequality: assumes "n > 1" shows "\\<^sub>e \ {d * (1 - \\<^sub>2) / 2.. d * sqrt (2 * (1 - \\<^sub>2))}" using cheeger_aux_1 cheeger_aux_2 assms by auto unbundle no_intro_cong_syntax end end \ No newline at end of file diff --git a/thys/Expander_Graphs/Expander_Graphs_Definition.thy b/thys/Expander_Graphs/Expander_Graphs_Definition.thy --- a/thys/Expander_Graphs/Expander_Graphs_Definition.thy +++ b/thys/Expander_Graphs/Expander_Graphs_Definition.thy @@ -1,1049 +1,1049 @@ section \Definitions\label{sec:definitions}\ text \This section introduces regular graphs as a sublocale in the graph theory developed by -Lars Noschinski~\cite{Graph_Theory-AFP} and introduces various expansion coefficients.\ +Lars Noschinski~\cite{Graph_Theory-AFP} and introduces various expansion coefficients.\ theory Expander_Graphs_Definition - imports + imports Graph_Theory.Digraph_Isomorphism - "HOL-Analysis.L2_Norm" + "HOL-Analysis.L2_Norm" Extra_Congruence_Method Expander_Graphs_Multiset_Extras Jordan_Normal_Form.Conjugate Interpolation_Polynomials_HOL_Algebra.Interpolation_Polynomial_Cardinalities begin unbundle intro_cong_syntax definition arcs_betw where "arcs_betw G u v = {a. a \ arcs G \ head G a = v \ tail G a = u}" -text \The following is a stronger notion than the notion of symmetry defined in -@{theory "Graph_Theory.Digraph"}, it requires that the number of edges from @{term "v"} to -@{term "w"} must be equal to the number of edges from @{term "w"} to @{term "v"} for any pair of +text \The following is a stronger notion than the notion of symmetry defined in +@{theory "Graph_Theory.Digraph"}, it requires that the number of edges from @{term "v"} to +@{term "w"} must be equal to the number of edges from @{term "w"} to @{term "v"} for any pair of vertices @{term "v"} @{term "w \ verts G"}.\ definition symmetric_multi_graph where "symmetric_multi_graph G = (fin_digraph G \ (\v w. {v, w} \ verts G \ card (arcs_betw G w v) = card (arcs_betw G v w)))" lemma symmetric_multi_graphI: assumes "fin_digraph G" assumes "bij_betw f (arcs G) (arcs G)" assumes "\e. e \ arcs G \ head G (f e) = tail G e \ tail G (f e) = head G e" shows "symmetric_multi_graph G" proof - have "card (arcs_betw G w v) = card (arcs_betw G v w)" (is "?L = ?R") if "v \ verts G" "w \ verts G" for v w proof - have a:"f x \ arcs G" if "x \ arcs G" for x using assms(2) that unfolding bij_betw_def by auto have b:"\y. y \ arcs G \ f y = x" if "x \ arcs G" for x using bij_betw_imp_surj_on[OF assms(2)] that by force - have "inj_on f (arcs G)" + have "inj_on f (arcs G)" using assms(2) unfolding bij_betw_def by simp hence "inj_on f {e \ arcs G. head G e = v \ tail G e = w}" by (rule inj_on_subset, auto) hence "?L = card (f ` {e \ arcs G. head G e = v \ tail G e = w})" unfolding arcs_betw_def by (intro card_image[symmetric]) also have "... = ?R" unfolding arcs_betw_def using a b assms(3) - by (intro arg_cong[where f="card"] order_antisym image_subsetI subsetI) fastforce+ + by (intro arg_cong[where f="card"] order_antisym image_subsetI subsetI) fastforce+ finally show ?thesis by simp qed - thus ?thesis + thus ?thesis using assms(1) unfolding symmetric_multi_graph_def by simp qed lemma symmetric_multi_graphD2: assumes "symmetric_multi_graph G" shows "fin_digraph G" - using assms unfolding symmetric_multi_graph_def by simp + using assms unfolding symmetric_multi_graph_def by simp lemma symmetric_multi_graphD: assumes "symmetric_multi_graph G" - shows "card {e \ arcs G. head G e=v \ tail G e=w} = card {e \ arcs G. head G e=w \ tail G e=v}" + shows "card {e \ arcs G. head G e=v \ tail G e=w} = card {e \ arcs G. head G e=w \ tail G e=v}" (is "card ?L = card ?R") proof (cases "v \ verts G \ w \ verts G") case True - then show ?thesis - using assms unfolding symmetric_multi_graph_def arcs_betw_def by simp + then show ?thesis + using assms unfolding symmetric_multi_graph_def arcs_betw_def by simp next case False interpret fin_digraph G using symmetric_multi_graphD2[OF assms(1)] by simp have 0:"?L = {}" "?R = {}" using False wellformed by auto show ?thesis unfolding 0 by simp qed lemma symmetric_multi_graphD3: assumes "symmetric_multi_graph G" shows - "card {e\arcs G. tail G e=v \ head G e=w}=card {e\arcs G. tail G e=w\head G e=v}" + "card {e\arcs G. tail G e=v \ head G e=w}=card {e\arcs G. tail G e=w\head G e=v}" using symmetric_multi_graphD[OF assms] by (simp add:conj.commute) lemma symmetric_multi_graphD4: assumes "symmetric_multi_graph G" - shows "card (arcs_betw G v w) = card (arcs_betw G w v)" + shows "card (arcs_betw G v w) = card (arcs_betw G w v)" using symmetric_multi_graphD[OF assms] unfolding arcs_betw_def by simp lemma symmetric_degree_eq: assumes "symmetric_multi_graph G" assumes "v \ verts G" shows "out_degree G v = in_degree G v" (is "?L = ?R") proof - - interpret fin_digraph "G" + interpret fin_digraph "G" using assms(1) symmetric_multi_graph_def by auto have "?L = card {e \ arcs G. tail G e = v}" unfolding out_degree_def out_arcs_def by simp also have "... = card (\w \ verts G. {e \ arcs G. head G e = w \ tail G e = v})" by (intro arg_cong[where f="card"]) (auto simp add:set_eq_iff) also have "... = (\w \ verts G. card {e \ arcs G. head G e = w \ tail G e = v})" by (intro card_UN_disjoint) auto also have "... = (\w \ verts G. card {e \ arcs G. head G e = v \ tail G e = w})" by (intro sum.cong refl symmetric_multi_graphD assms) also have "... = card (\w \ verts G. {e \ arcs G. head G e = v \ tail G e = w})" by (intro card_UN_disjoint[symmetric]) auto also have "... = card (in_arcs G v)" - by (intro arg_cong[where f="card"]) (auto simp add:set_eq_iff) - also have "... = ?R" + by (intro arg_cong[where f="card"]) (auto simp add:set_eq_iff) + also have "... = ?R" unfolding in_degree_def by simp finally show ?thesis by simp qed definition edges where "edges G = image_mset (arc_to_ends G) (mset_set (arcs G))" lemma (in fin_digraph) count_edges: "count (edges G) (u,v) = card (arcs_betw G u v)" (is "?L = ?R") proof - have "?L = card {x \ arcs G. arc_to_ends G x = (u, v)}" unfolding edges_def count_mset_exp image_mset_filter_mset_swap[symmetric] by simp also have "... = ?R" unfolding arcs_betw_def arc_to_ends_def by (intro arg_cong[where f="card"]) auto finally show ?thesis by simp qed lemma (in fin_digraph) count_edges_sym: assumes "symmetric_multi_graph G" - shows "count (edges G) (v, w) = count (edges G) (w, v)" + shows "count (edges G) (v, w) = count (edges G) (w, v)" unfolding count_edges using symmetric_multi_graphD4[OF assms] by simp lemma (in fin_digraph) edges_sym: assumes "symmetric_multi_graph G" - shows "{# (y,x). (x,y) \# (edges G) #} = edges G" + shows "{# (y,x). (x,y) \# (edges G) #} = edges G" proof - have "count {#(y, x). (x, y) \# edges G#} x = count (edges G) x" (is "?L = ?R") for x proof - have "?L = count (edges G) (snd x, fst x)" unfolding count_mset_exp by (simp add:image_mset_filter_mset_swap[symmetric] case_prod_beta prod_eq_iff ac_simps) also have "... = count (edges G) (fst x, snd x)" unfolding count_edges_sym[OF assms] by simp also have "... = count (edges G) x" by simp finally show ?thesis by simp qed thus ?thesis by (intro multiset_eqI) simp qed definition "vertices_from G v = {# snd e | e \# edges G. fst e = v #}" definition "vertices_to G v = {# fst e | e \# edges G. snd e = v #}" context fin_digraph begin -lemma edge_set: +lemma edge_set: assumes "x \# edges G" shows "fst x \ verts G" "snd x \ verts G" using assms unfolding edges_def arc_to_ends_def by auto lemma verts_from_alt: "vertices_from G v = image_mset (head G) (mset_set (out_arcs G v))" proof - have "{#x \# mset_set (arcs G). tail G x = v#} = mset_set {a \ arcs G. tail G a = v}" by (intro filter_mset_mset_set) simp thus ?thesis unfolding vertices_from_def out_arcs_def edges_def arc_to_ends_def by (simp add:image_mset.compositionality image_mset_filter_mset_swap[symmetric] comp_def) qed lemma verts_to_alt: "vertices_to G v = image_mset (tail G) (mset_set (in_arcs G v))" proof - have "{#x \# mset_set (arcs G). head G x = v#} = mset_set {a \ arcs G. head G a = v}" by (intro filter_mset_mset_set) simp thus ?thesis unfolding vertices_to_def in_arcs_def edges_def arc_to_ends_def by (simp add:image_mset.compositionality image_mset_filter_mset_swap[symmetric] comp_def) qed lemma set_mset_vertices_from: "set_mset (vertices_from G x) \ verts G" unfolding vertices_from_def using edge_set by auto lemma set_mset_vertices_to: "set_mset (vertices_to G x) \ verts G" unfolding vertices_to_def using edge_set by auto end text \A symmetric multigraph is regular if every node has the same degree. This is the context in which the expansion conditions are introduced.\ locale regular_graph = fin_digraph + assumes sym: "symmetric_multi_graph G" assumes verts_non_empty: "verts G \ {}" assumes arcs_non_empty: "arcs G \ {}" assumes reg': "\v w. v \ verts G \ w \ verts G \ out_degree G v = out_degree G w" begin definition d where "d = out_degree G (SOME v. v \ verts G)" lemmas count_sym = count_edges_sym[OF sym] -lemma reg: +lemma reg: assumes "v \ verts G" shows "out_degree G v = d" "in_degree G v = d" proof - define w where "w = (SOME v. v \ verts G)" have "w \ verts G" unfolding w_def using assms(1) by (rule someI) hence "out_degree G v = out_degree G w" by (intro reg' assms(1)) also have "... = d" unfolding d_def w_def by simp finally show a:"out_degree G v = d" by simp show "in_degree G v = d" using a symmetric_degree_eq[OF sym assms(1)] by simp qed definition n where "n = card (verts G)" lemma n_gt_0: "n > 0" unfolding n_def using verts_non_empty by auto lemma d_gt_0: "d > 0" proof - obtain a where a:"a \ arcs G" using arcs_non_empty by auto hence "a \ in_arcs G (head G a) " unfolding in_arcs_def by simp hence "0 < in_degree G (head G a)" unfolding in_degree_def card_gt_0_iff by blast also have "... = d" using a by (intro reg) simp finally show ?thesis by simp qed -definition g_inner :: "('a \ ('c :: conjugatable_field)) \ ('a \ 'c) \ 'c" +definition g_inner :: "('a \ ('c :: conjugatable_field)) \ ('a \ 'c) \ 'c" where "g_inner f g = (\x \ verts G. (f x) * conjugate (g x))" lemma conjugate_divide[simp]: fixes x y :: "'c :: conjugatable_field" shows "conjugate (x / y) = conjugate x / conjugate y" proof (cases "y = 0") case True then show ?thesis by simp next case False have "conjugate (x/y) * conjugate y = conjugate x" using False by (simp add:conjugate_dist_mul[symmetric]) thus ?thesis by (simp add:divide_simps) qed lemma g_inner_simps: - "g_inner (\x. 0) g = 0" - "g_inner f (\x. 0) = 0" - "g_inner (\x. c * f x) g = c * g_inner f g" - "g_inner f (\x. c * g x) = conjugate c * g_inner f g" - "g_inner (\x. f x - g x) h = g_inner f h - g_inner g h" - "g_inner (\x. f x + g x) h = g_inner f h + g_inner g h" - "g_inner f (\x. g x + h x) = g_inner f g + g_inner f h" - "g_inner f (\x. g x / c) = g_inner f g / conjugate c" - "g_inner (\x. f x / c) g = g_inner f g / c" - unfolding g_inner_def - by (auto simp add:sum.distrib algebra_simps sum_distrib_left sum_subtractf sum_divide_distrib + "g_inner (\x. 0) g = 0" + "g_inner f (\x. 0) = 0" + "g_inner (\x. c * f x) g = c * g_inner f g" + "g_inner f (\x. c * g x) = conjugate c * g_inner f g" + "g_inner (\x. f x - g x) h = g_inner f h - g_inner g h" + "g_inner (\x. f x + g x) h = g_inner f h + g_inner g h" + "g_inner f (\x. g x + h x) = g_inner f g + g_inner f h" + "g_inner f (\x. g x / c) = g_inner f g / conjugate c" + "g_inner (\x. f x / c) g = g_inner f g / c" + unfolding g_inner_def + by (auto simp add:sum.distrib algebra_simps sum_distrib_left sum_subtractf sum_divide_distrib conjugate_dist_mul conjugate_dist_add) definition "g_norm f = sqrt (g_inner f f)" lemma g_norm_eq: "g_norm f = L2_set f (verts G)" unfolding g_norm_def g_inner_def L2_set_def by (intro arg_cong[where f="sqrt"] sum.cong refl) (simp add:power2_eq_square) lemma g_inner_cauchy_schwartz: fixes f g :: "'a \ real" shows "\g_inner f g\ \ g_norm f * g_norm g" proof - - have "\g_inner f g\ \ (\v \ verts G. \f v * g v\)" + have "\g_inner f g\ \ (\v \ verts G. \f v * g v\)" unfolding g_inner_def conjugate_real_def by (intro sum_abs) also have "... \ g_norm f * g_norm g" unfolding g_norm_eq abs_mult by (intro L2_set_mult_ineq) finally show ?thesis by simp qed lemma g_inner_cong: assumes "\x. x \ verts G \ f1 x = f2 x" assumes "\x. x \ verts G \ g1 x = g2 x" shows "g_inner f1 g1 = g_inner f2 g2" unfolding g_inner_def using assms by (intro sum.cong refl) auto lemma g_norm_cong: assumes "\x. x \ verts G \ f x = g x" shows "g_norm f = g_norm g" - unfolding g_norm_def + unfolding g_norm_def by (intro arg_cong[where f="sqrt"] g_inner_cong assms) -lemma g_norm_nonneg: "g_norm f \ 0" +lemma g_norm_nonneg: "g_norm f \ 0" unfolding g_norm_def g_inner_def by (intro real_sqrt_ge_zero sum_nonneg) auto lemma g_norm_sq: - "g_norm f^2 = g_inner f f" + "g_norm f^2 = g_inner f f" using g_norm_nonneg g_norm_def by simp definition g_step :: "('a \ real) \ ('a \ real)" where "g_step f v = (\x \ in_arcs G v. f (tail G x) / real d)" lemma g_step_simps: "g_step (\x. f x + g x) y = g_step f y + g_step g y" "g_step (\x. f x / c) y = g_step f y / c" unfolding g_step_def sum_divide_distrib[symmetric] using finite_in_arcs d_gt_0 by (auto intro:sum.cong simp add:sum.distrib field_simps sum_distrib_left sum_subtractf) lemma g_inner_step_eq: "g_inner f (g_step f) = (\a \ arcs G. f (head G a) * f (tail G a)) / d" (is "?L = ?R") proof - have "?L = (\v\verts G. f v * (\a\in_arcs G v. f (tail G a) / d))" unfolding g_inner_def g_step_def by simp also have "... = (\v\verts G. (\a\in_arcs G v. f v * f (tail G a) / d))" by (subst sum_distrib_left) simp also have "... = (\v\verts G. (\a\in_arcs G v. f (head G a) * f (tail G a) / d))" unfolding in_arcs_def by (intro sum.cong refl) simp also have "... = (\a \ (\ (in_arcs G ` verts G)). f (head G a) * f (tail G a) / d)" - using finite_verts by (intro sum.UNION_disjoint[symmetric] ballI) + using finite_verts by (intro sum.UNION_disjoint[symmetric] ballI) (auto simp add:in_arcs_def) also have "... = (\a \ arcs G. f (head G a) * f (tail G a) / d)" unfolding in_arcs_def using wellformed by (intro sum.cong) auto also have "... = ?R" by (intro sum_divide_distrib[symmetric]) finally show ?thesis by simp qed -definition \_test +definition \_test where "\_test = {f. g_norm f^2 \ 0 \ g_inner f (\_. 1) = 0}" -lemma \_test_ne: +lemma \_test_ne: assumes "n > 1" shows "\_test \ {}" proof - obtain v where v_def: "v \ verts G" using verts_non_empty by auto have "False" if "\w. w \ verts G \ w = v" proof - - have "verts G = {v}" using that v_def + have "verts G = {v}" using that v_def by (intro iffD2[OF set_eq_iff] allI) blast thus False using assms n_def by simp qed - then obtain w where w_def: "w \ verts G" "v \ w" + then obtain w where w_def: "w \ verts G" "v \ w" by auto define f where "f x= (if x = v then 1 else (if x = w then (-1) else (0::real)))" for x have "g_norm f^2 = (\x\verts G. (if x = v then 1 else if x = w then - 1 else 0)\<^sup>2)" - unfolding g_norm_sq g_inner_def conjugate_real_def power2_eq_square[symmetric] + unfolding g_norm_sq g_inner_def conjugate_real_def power2_eq_square[symmetric] by (simp add:f_def) also have "... = (\x \ {v,w}. (if x = v then 1 else if x = w then -1 else 0)\<^sup>2)" using v_def(1) w_def(1) by (intro sum.mono_neutral_cong refl) auto also have "... = (\x \ {v,w}. (if x = v then 1 else - 1)\<^sup>2)" by (intro sum.cong) auto also have "... = 2" using w_def(2) by (simp add:if_distrib if_distribR sum.If_cases) finally have "g_norm f^2 = 2" by simp hence "g_norm f \ 0" by auto moreover have "g_inner f (\_.1) = 0" unfolding g_inner_def f_def using v_def w_def by (simp add:sum.If_cases) ultimately have "f \ \_test" unfolding \_test_def by simp thus ?thesis by auto qed -lemma \_test_empty: +lemma \_test_empty: assumes "n = 1" shows "\_test = {}" proof - obtain v where v_def: "verts G = {v}" using assms card_1_singletonE unfolding n_def by auto have "False" if "f \ \_test" for f proof - have "0 = (g_inner f (\_.1))^2" using that \_test_def by simp also have "... = (f v)^2" unfolding g_inner_def v_def by simp also have "... = g_norm f^2" - unfolding g_norm_sq g_inner_def v_def + unfolding g_norm_sq g_inner_def v_def by (simp add:power2_eq_square) also have "... \ 0" using that \_test_def by simp finally show "False" by simp qed thus ?thesis by auto qed -text \The following are variational definitions for the maxiumum of the spectrum (resp. maximum +text \The following are variational definitions for the maxiumum of the spectrum (resp. maximum modulus of the spectrum) of the stochastic matrix (excluding the Perron eigenvalue $1$). Note that both values can still obtain the value one $1$ (if the multiplicity of the eigenvalue $1$ is larger than $1$ in the stochastic matrix, or in the modulus case if $-1$ is an eigenvalue). The definition relies on the supremum of the Rayleigh-Quotient for vectors orthogonal to the stationary distribution). In Section~\ref{sec:expander_eigenvalues}, the equivalence of this value with the algebraic definition will be shown. The definition here has the advantage that it is (obviously) independent of the matrix representation (ordering of the vertices) used.\ definition \\<^sub>2 :: real where "\\<^sub>2 = (if n > 1 then (SUP f \ \_test. g_inner f (g_step f)/g_inner f f) else 0)" definition \\<^sub>a :: real where "\\<^sub>a = (if n > 1 then (SUP f \ \_test. \g_inner f (g_step f)\/g_inner f f) else 0)" lemma sum_arcs_tail: fixes f :: "'a \ ('c :: semiring_1)" shows "(\a \ arcs G. f (tail G a)) = of_nat d * (\v \ verts G. f v)" (is "?L = ?R") proof - have "?L = (\a\(\(out_arcs G ` verts G)). f (tail G a))" by (intro sum.cong) auto also have "... = (\v \ verts G. (\a \ out_arcs G v. f (tail G a)))" by (intro sum.UNION_disjoint) auto also have "... = (\v \ verts G. of_nat (out_degree G v) * f v)" unfolding out_degree_def by simp also have "... = (\v \ verts G. of_nat d * f v)" by (intro sum.cong arg_cong2[where f="(*)"] arg_cong[where f="of_nat"] reg) auto also have "... = ?R" by (simp add:sum_distrib_left) finally show ?thesis by simp qed lemma sum_arcs_head: fixes f :: "'a \ ('c :: semiring_1)" shows "(\a \ arcs G. f (head G a)) = of_nat d * (\v \ verts G. f v)" (is "?L = ?R") proof - have "?L = (\a\(\(in_arcs G ` verts G)). f (head G a))" by (intro sum.cong) auto also have "... = (\v \ verts G. (\a \ in_arcs G v. f (head G a)))" by (intro sum.UNION_disjoint) auto also have "... = (\v \ verts G. of_nat (in_degree G v) * f v)" unfolding in_degree_def by simp also have "... = (\v \ verts G. of_nat d * f v)" by (intro sum.cong arg_cong2[where f="(*)"] arg_cong[where f="of_nat"] reg) auto also have "... = ?R" by (simp add:sum_distrib_left) finally show ?thesis by simp qed lemma bdd_above_aux: "\\a\arcs G. f(head G a)*f(tail G a)\ \ d* g_norm f^2" (is "?L \ ?R") proof - have "(\a\arcs G. f (head G a)^2) = d * g_norm f^2" unfolding sum_arcs_head[where f="\x. f x^2"] g_norm_sq g_inner_def by (simp add:power2_eq_square) - hence 0:"L2_set (\a. f (head G a)) (arcs G) \ sqrt (d * g_norm f^2)" + hence 0:"L2_set (\a. f (head G a)) (arcs G) \ sqrt (d * g_norm f^2)" using g_norm_nonneg unfolding L2_set_def by simp have "(\a\arcs G. f (tail G a)^2) = d * g_norm f^2" unfolding sum_arcs_tail[where f="\x. f x^2"] sum_distrib_left[symmetric] g_norm_sq g_inner_def by (simp add:power2_eq_square) - hence 1:"L2_set (\a. f (tail G a)) (arcs G) \ sqrt (d * g_norm f^2)" + hence 1:"L2_set (\a. f (tail G a)) (arcs G) \ sqrt (d * g_norm f^2)" unfolding L2_set_def by simp have "?L \ (\a \ arcs G. \f (head G a)\ * \f(tail G a)\)" - unfolding abs_mult[symmetric] by (intro divide_right_mono sum_abs) + unfolding abs_mult[symmetric] by (intro divide_right_mono sum_abs) also have "... \ (L2_set (\a. f (head G a)) (arcs G) * L2_set (\a. f (tail G a)) (arcs G))" - by (intro L2_set_mult_ineq) + by (intro L2_set_mult_ineq) also have "... \ (sqrt (d * g_norm f^2) * sqrt (d * g_norm f^2))" by (intro mult_mono 0 1) auto also have "... = d * g_norm f^2" using d_gt_0 g_norm_nonneg by simp finally show ?thesis by simp qed -lemma bdd_above_aux_2: +lemma bdd_above_aux_2: assumes "f \ \_test" shows "\g_inner f (g_step f)\ / g_inner f f \ 1" proof - - have 0:"g_inner f f > 0" + have 0:"g_inner f f > 0" using assms unfolding \_test_def g_norm_sq[symmetric] by auto have "\g_inner f (g_step f)\ = \\a\arcs G. f (head G a) * f (tail G a)\ / real d" unfolding g_inner_step_eq by simp also have "... \ d * g_norm f^2 / d" by (intro divide_right_mono bdd_above_aux assms) auto also have "... = g_inner f f" using d_gt_0 unfolding g_norm_sq by simp - finally have "\g_inner f (g_step f)\ \ g_inner f f" + finally have "\g_inner f (g_step f)\ \ g_inner f f" by simp thus ?thesis using 0 by simp qed -lemma bdd_above_aux_3: +lemma bdd_above_aux_3: assumes "f \ \_test" shows "g_inner f (g_step f) / g_inner f f \ 1" (is "?L \ ?R") proof - have "?L \ \g_inner f (g_step f)\ / g_inner f f" unfolding g_norm_sq[symmetric] by (intro divide_right_mono) auto also have "... \ 1" using bdd_above_aux_2[OF assms] by simp finally show ?thesis by simp qed lemma bdd_above_\: "bdd_above ((\f. \g_inner f (g_step f)\ / g_inner f f) ` \_test)" using bdd_above_aux_2 by (intro bdd_aboveI[where M="1"]) auto lemma bdd_above_\\<^sub>2: "bdd_above ((\f. g_inner f (g_step f) / g_inner f f) ` \_test)" using bdd_above_aux_3 by (intro bdd_aboveI[where M="1"]) auto lemma \_le_1: "\\<^sub>a \ 1" proof (cases "n > 1") case True - have "(SUP f\\_test. \g_inner f (g_step f)\ / g_inner f f) \ 1" + have "(SUP f\\_test. \g_inner f (g_step f)\ / g_inner f f) \ 1" using bdd_above_aux_2 \_test_ne[OF True] by (intro cSup_least) auto thus "\\<^sub>a \ 1" unfolding \\<^sub>a_def using True by simp next case False thus ?thesis unfolding \\<^sub>a_def by simp qed lemma \\<^sub>2_le_1: "\\<^sub>2 \ 1" proof (cases "n > 1") case True - have "(SUP f\\_test. g_inner f (g_step f) / g_inner f f) \ 1" + have "(SUP f\\_test. g_inner f (g_step f) / g_inner f f) \ 1" using bdd_above_aux_3 \_test_ne[OF True] by (intro cSup_least) auto thus "\\<^sub>2 \ 1" unfolding \\<^sub>2_def using True by simp next case False thus ?thesis unfolding \\<^sub>2_def by simp qed lemma \_ge_0: "\\<^sub>a \ 0" proof (cases "n > 1") case True - obtain f where f_def: "f \ \_test" + obtain f where f_def: "f \ \_test" using \_test_ne[OF True] by auto have "0 \ \g_inner f (g_step f)\ / g_inner f f" unfolding g_norm_sq[symmetric] by (intro divide_nonneg_nonneg) auto also have "... \ (SUP f\\_test. \g_inner f (g_step f)\ / g_inner f f)" using f_def by (intro cSup_upper bdd_above_\) auto finally have "(SUP f\\_test. \g_inner f (g_step f)\ / g_inner f f) \ 0" by simp thus ?thesis unfolding \\<^sub>a_def using True by simp next case False thus ?thesis unfolding \\<^sub>a_def by simp qed lemma os_expanderI: assumes "n > 1" - assumes "\f. g_inner f (\_. 1)=0 \ g_inner f (g_step f) \ C*g_norm f^2" + assumes "\f. g_inner f (\_. 1)=0 \ g_inner f (g_step f) \ C*g_norm f^2" shows "\\<^sub>2 \ C" proof - have "g_inner f (g_step f) / g_inner f f \ C" if "f \ \_test" for f proof - have "g_inner f (g_step f) \ C*g_inner f f" using that \_test_def assms(2) unfolding g_norm_sq by auto moreover have "g_inner f f > 0" using that unfolding \_test_def g_norm_sq[symmetric] by auto - ultimately show ?thesis + ultimately show ?thesis by (simp add:divide_simps) qed hence "(SUP f\\_test. g_inner f (g_step f) / g_inner f f) \ C" using \_test_ne[OF assms(1)] by (intro cSup_least) auto thus ?thesis unfolding \\<^sub>2_def using assms by simp qed lemma os_expanderD: assumes "g_inner f (\_. 1) = 0" shows "g_inner f (g_step f) \ \\<^sub>2 * g_norm f^2" (is "?L \ ?R") proof (cases "g_norm f \ 0") case True have 0:"f \ \_test" unfolding \_test_def using assms True by auto - hence 1:"n > 1" + hence 1:"n > 1" using \_test_empty n_gt_0 by fastforce - have "g_inner f (g_step f)/ g_norm f^2 = g_inner f (g_step f)/g_inner f f" + have "g_inner f (g_step f)/ g_norm f^2 = g_inner f (g_step f)/g_inner f f" unfolding g_norm_sq by simp also have "... \ (SUP f\\_test. g_inner f (g_step f) / g_inner f f)" - by (intro cSup_upper bdd_above_\\<^sub>2 imageI 0) + by (intro cSup_upper bdd_above_\\<^sub>2 imageI 0) also have "... = \\<^sub>2" using 1 unfolding \\<^sub>2_def by simp finally have "g_inner f (g_step f)/ g_norm f^2 \ \\<^sub>2" by simp - thus ?thesis + thus ?thesis using True by (simp add:divide_simps) next case False hence "g_inner f f = 0" unfolding g_norm_sq[symmetric] by simp hence 0:"\v. v \ verts G \ f v = 0" unfolding g_inner_def by (subst (asm) sum_nonneg_eq_0_iff) auto hence "?L = 0" - unfolding g_step_def g_inner_def by simp + unfolding g_step_def g_inner_def by simp also have "... \ \\<^sub>2 * g_norm f^2" using False by simp finally show ?thesis by simp qed lemma expander_intro_1: assumes "C \ 0" - assumes "\f. g_inner f (\_. 1)=0 \ \g_inner f (g_step f)\ \ C*g_norm f^2" + assumes "\f. g_inner f (\_. 1)=0 \ \g_inner f (g_step f)\ \ C*g_norm f^2" shows "\\<^sub>a \ C" proof (cases "n > 1") case True have "\g_inner f (g_step f)\ / g_inner f f \ C" if "f \ \_test" for f proof - have "\g_inner f (g_step f)\ \ C*g_inner f f" using that \_test_def assms(2) unfolding g_norm_sq by auto moreover have "g_inner f f > 0" using that unfolding \_test_def g_norm_sq[symmetric] by auto - ultimately show ?thesis + ultimately show ?thesis by (simp add:divide_simps) qed - hence "(SUP f\\_test. \g_inner f (g_step f)\ / g_inner f f) \ C" + hence "(SUP f\\_test. \g_inner f (g_step f)\ / g_inner f f) \ C" using \_test_ne[OF True] by (intro cSup_least) auto thus ?thesis using True unfolding \\<^sub>a_def by auto next case False then show ?thesis using assms unfolding \\<^sub>a_def by simp qed lemma expander_intro: assumes "C \ 0" - assumes "\f. g_inner f (\_. 1)=0 \ \\a \ arcs G. f(head G a) * f(tail G a)\ \ C*g_norm f^2" + assumes "\f. g_inner f (\_. 1)=0 \ \\a \ arcs G. f(head G a) * f(tail G a)\ \ C*g_norm f^2" shows "\\<^sub>a \ C/d" proof - - have "\g_inner f (g_step f)\ \ C / real d * (g_norm f)\<^sup>2" (is "?L \ ?R") + have "\g_inner f (g_step f)\ \ C / real d * (g_norm f)\<^sup>2" (is "?L \ ?R") if "g_inner f (\_. 1) = 0" for f proof - have "?L = \\a\arcs G. f (head G a) * f (tail G a)\ / real d" unfolding g_inner_step_eq by simp also have "... \ C*g_norm f^2 / real d" by (intro divide_right_mono assms(2)[OF that]) auto also have "... = ?R" by simp finally show ?thesis by simp qed thus ?thesis by (intro expander_intro_1 divide_nonneg_nonneg assms) auto qed lemma expansionD1: assumes "g_inner f (\_. 1) = 0" shows "\g_inner f (g_step f)\ \ \\<^sub>a * g_norm f^2" (is "?L \ ?R") proof (cases "g_norm f \ 0") case True have 0:"f \ \_test" unfolding \_test_def using assms True by auto - hence 1:"n > 1" + hence 1:"n > 1" using \_test_empty n_gt_0 by fastforce - have "\g_inner f (g_step f)\/ g_norm f^2 = \g_inner f (g_step f)\/g_inner f f" + have "\g_inner f (g_step f)\/ g_norm f^2 = \g_inner f (g_step f)\/g_inner f f" unfolding g_norm_sq by simp also have "... \ (SUP f\\_test. \g_inner f (g_step f)\ / g_inner f f)" - by (intro cSup_upper bdd_above_\ imageI 0) + by (intro cSup_upper bdd_above_\ imageI 0) also have "... = \\<^sub>a" using 1 unfolding \\<^sub>a_def by simp finally have "\g_inner f (g_step f)\/ g_norm f^2 \ \\<^sub>a" by simp - thus ?thesis + thus ?thesis using True by (simp add:divide_simps) next case False hence "g_inner f f = 0" unfolding g_norm_sq[symmetric] by simp hence 0:"\v. v \ verts G \ f v = 0" unfolding g_inner_def by (subst (asm) sum_nonneg_eq_0_iff) auto hence "?L = 0" - unfolding g_step_def g_inner_def by simp + unfolding g_step_def g_inner_def by simp also have "... \ \\<^sub>a * g_norm f^2" using False by simp finally show ?thesis by simp qed lemma expansionD: assumes "g_inner f (\_. 1) = 0" shows "\\a \ arcs G. f (head G a) * f (tail G a)\ \ d * \\<^sub>a * g_norm f^2" (is "?L \ ?R") proof - have "?L = \g_inner f (g_step f) * d\" unfolding g_inner_step_eq using d_gt_0 by simp also have "... \ \g_inner f (g_step f)\ * d" by (simp add:abs_mult) also have "... \ (\\<^sub>a * g_norm f^2) * d" by (intro expansionD1 mult_right_mono assms(1)) auto also have "... = ?R" by simp finally show ?thesis by simp qed definition edges_betw where "edges_betw S T = {a \ arcs G. tail G a \ S \ head G a \ T}" text \This parameter is the edge expansion. It is usually denoted by the symbol $h$ or $h(G)$ in text books. Contrary to the previous definitions it doesn't have a spectral theoretic counter part.\ definition \\<^sub>e where "\\<^sub>e = (if n > 1 then (MIN S\{S. S\verts G\2*card S\n\S\{}}. real (card (edges_betw S (-S)))/card S) else 0)" lemma edge_expansionD: assumes "S \ verts G" "2*card S \ n" shows "\\<^sub>e * card S \ real (card (edges_betw S (-S)))" proof (cases "S \ {}") case True - moreover have "finite S" + moreover have "finite S" using finite_subset[OF assms(1)] by simp ultimately have "card S > 0" by auto hence 1: "real (card S) > 0" by simp hence 2: "n > 1" using assms(2) by simp let ?St = "{S. S \ verts G \ 2 * card S \ n \ S \ {}}" have 0: "finite ?St" by (rule finite_subset[where B="Pow (verts G)"]) auto have "\\<^sub>e = (MIN S\?St. real (card (edges_betw S (-S)))/card S)" using 2 unfolding \\<^sub>e_def by simp also have "... \ real (card (edges_betw S (-S))) / card S" using assms True by (intro Min_le finite_imageI imageI) auto finally have "\\<^sub>e \ real (card (edges_betw S (-S))) / card S" by simp thus ?thesis using 1 by (simp add:divide_simps) next case False hence "card S = 0" by simp thus ?thesis by simp qed lemma edge_expansionI: fixes \ :: real assumes "n > 1" - assumes "\S. S \ verts G \ 2*card S \ n \ S \ {} \ card (edges_betw S (-S)) \ \ * card S" + assumes "\S. S \ verts G \ 2*card S \ n \ S \ {} \ card (edges_betw S (-S)) \ \ * card S" shows "\\<^sub>e \ \" proof - define St where "St = {S. S \ verts G \ 2*card S \ n \ S \ {}}" have 0: "finite St" unfolding St_def - by (rule finite_subset[where B="Pow (verts G)"]) auto + by (rule finite_subset[where B="Pow (verts G)"]) auto - obtain v where v_def: "v \ verts G" using verts_non_empty by auto + obtain v where v_def: "v \ verts G" using verts_non_empty by auto - have "{v} \ St" + have "{v} \ St" using assms v_def unfolding St_def n_def by auto hence 1: "St \ {}" by auto - have 2: "\ \ real (card (edges_betw S (- S))) / real (card S)" if "S \ St" for S + have 2: "\ \ real (card (edges_betw S (- S))) / real (card S)" if "S \ St" for S proof - - have "real (card (edges_betw S (- S))) \ \ * card S" + have "real (card (edges_betw S (- S))) \ \ * card S" using assms(2) that unfolding St_def by simp - moreover have "finite S" + moreover have "finite S" using that unfolding St_def by (intro finite_subset[OF _ finite_verts]) auto - hence "card S > 0" + hence "card S > 0" using that unfolding St_def by auto - ultimately show ?thesis + ultimately show ?thesis by (simp add:divide_simps) qed have "\ \ (MIN S\St. real (card (edges_betw S (- S))) / real (card S))" using 0 1 2 by (intro Min.boundedI finite_imageI) auto thus ?thesis - unfolding \\<^sub>e_def St_def[symmetric] using assms by auto + unfolding \\<^sub>e_def St_def[symmetric] using assms by auto qed end lemma regular_graphI: assumes "symmetric_multi_graph G" assumes "verts G \ {}" "d > 0" assumes "\v. v \ verts G \ out_degree G v = d" shows "regular_graph G" proof - obtain v where v_def: "v \ verts G" using assms(2) by auto - have "arcs G \ {}" + have "arcs G \ {}" proof (rule ccontr) assume "\arcs G \ {}" hence "arcs G = {}" by simp hence "out_degree G v = 0" unfolding out_degree_def out_arcs_def by simp hence "d = 0" using v_def assms(4) by simp thus False using assms(3) by simp qed thus ?thesis using assms symmetric_multi_graphD2[OF assms(1)] unfolding regular_graph_def regular_graph_axioms_def by simp qed text \The following theorems verify that a graph isomorphisms preserve symmetry, regularity and all the expansion coefficients.\ lemma (in fin_digraph) symmetric_graph_iso: assumes "digraph_iso G H" assumes "symmetric_multi_graph G" shows "symmetric_multi_graph H" proof - obtain h where hom_iso: "digraph_isomorphism h" and H_alt: "H = app_iso h G" using assms unfolding digraph_iso_def by auto have 0:"fin_digraph H" unfolding H_alt by (intro fin_digraphI_app_iso hom_iso) interpret H:fin_digraph H using 0 by auto have 1:"arcs_betw H (iso_verts h v) (iso_verts h w) = iso_arcs h ` arcs_betw G v w" (is "?L = ?R") if "v \ verts G" "w \ verts G" for v w proof - have "?L = {a \ iso_arcs h ` arcs G. iso_head h a=iso_verts h w \ iso_tail h a=iso_verts h v}" unfolding arcs_betw_def H_alt arcs_app_iso head_app_iso tail_app_iso by simp - also have "... = {a. (\b \ arcs G. a = iso_arcs h b \ iso_verts h (head G b) = iso_verts h w \ + also have "... = {a. (\b \ arcs G. a = iso_arcs h b \ iso_verts h (head G b) = iso_verts h w \ iso_verts h (tail G b) = iso_verts h v)}" using iso_verts_head[OF hom_iso] iso_verts_tail[OF hom_iso] by auto - also have "... = {a. (\b \ arcs G. a = iso_arcs h b \ head G b = w \ tail G b = v)}" + also have "... = {a. (\b \ arcs G. a = iso_arcs h b \ head G b = w \ tail G b = v)}" using that iso_verts_eq_iff[OF hom_iso] by auto also have "... = ?R" unfolding arcs_betw_def by (auto simp add:image_iff set_eq_iff) finally show ?thesis by simp qed - have "card (arcs_betw H w v) = card (arcs_betw H v w)" (is "?L = ?R") + have "card (arcs_betw H w v) = card (arcs_betw H v w)" (is "?L = ?R") if v_range: "v \ verts H" and w_range: "w \ verts H" for v w proof - obtain v' where v': "v = iso_verts h v'" "v' \ verts G" using that v_range verts_app_iso unfolding H_alt by auto obtain w' where w': "w = iso_verts h w'" "w' \ verts G" using that w_range verts_app_iso unfolding H_alt by auto have "?L = card (arcs_betw H (iso_verts h w') (iso_verts h v'))" unfolding v' w' by simp also have "... = card (iso_arcs h ` arcs_betw G w' v')" by (intro arg_cong[where f="card"] 1 v' w') also have "... = card (arcs_betw G w' v')" using iso_arcs_eq_iff[OF hom_iso] unfolding arcs_betw_def by (intro card_image inj_onI) auto also have "... = card (arcs_betw G v' w')" - by (intro symmetric_multi_graphD4 assms(2)) + by (intro symmetric_multi_graphD4 assms(2)) also have "... = card (iso_arcs h ` arcs_betw G v' w')" using iso_arcs_eq_iff[OF hom_iso] unfolding arcs_betw_def by (intro card_image[symmetric] inj_onI) auto also have "... = card (arcs_betw H (iso_verts h v') (iso_verts h w'))" by (intro arg_cong[where f="card"] 1[symmetric] v' w') also have "... = ?R" unfolding v' w' by simp finally show ?thesis by simp qed thus ?thesis using 0 unfolding symmetric_multi_graph_def by auto qed lemma (in regular_graph) assumes "digraph_iso G H" - shows regular_graph_iso: "regular_graph H" + shows regular_graph_iso: "regular_graph H" and regular_graph_iso_size: "regular_graph.n H = n" and regular_graph_iso_degree: "regular_graph.d H = d" and regular_graph_iso_expansion_le: "regular_graph.\\<^sub>a H \ \\<^sub>a" and regular_graph_iso_os_expansion_le: "regular_graph.\\<^sub>2 H \ \\<^sub>2" and regular_graph_iso_edge_expansion_ge: "regular_graph.\\<^sub>e H \ \\<^sub>e" proof - obtain h where hom_iso: "digraph_isomorphism h" and H_alt: "H = app_iso h G" using assms unfolding digraph_iso_def by auto have 0:"symmetric_multi_graph H" by (intro symmetric_graph_iso[OF assms(1)] sym) - have 1:"verts H \ {}" + have 1:"verts H \ {}" unfolding H_alt verts_app_iso using verts_non_empty by simp then obtain h_wit where h_wit: "h_wit \ verts H" by auto have 3:"out_degree H v = d" if v_range: "v \ verts H" for v proof - obtain v' where v': "v = iso_verts h v'" "v' \ verts G" using that v_range verts_app_iso unfolding H_alt by auto have "out_degree H v = out_degree G v'" unfolding v' H_alt by (intro out_degree_app_iso_eq[OF hom_iso] v') also have "... = d" by (intro reg v') finally show ?thesis by simp qed thus 2:"regular_graph H" by (intro regular_graphI[where d="d"] 0 d_gt_0 1) auto interpret H:"regular_graph" H using 2 by auto have "H.n = card (iso_verts h ` verts G)" unfolding H.n_def unfolding H_alt verts_app_iso by simp also have "... = card (verts G)" - by (intro card_image digraph_isomorphism_inj_on_verts hom_iso) + by (intro card_image digraph_isomorphism_inj_on_verts hom_iso) also have "... = n" unfolding n_def by simp finally show n_eq: "H.n = n" by simp have "H.d = out_degree H h_wit" by (intro H.reg[symmetric] h_wit) also have "... = d" by (intro 3 h_wit) finally show 4:"H.d = d" by simp - have "bij_betw (iso_verts h) (verts G) (verts H)" - unfolding H_alt using hom_iso + have "bij_betw (iso_verts h) (verts G) (verts H)" + unfolding H_alt using hom_iso by (simp add: bij_betw_def digraph_isomorphism_inj_on_verts) - hence g_inner_conv: - "H.g_inner f g = g_inner (\x. f (iso_verts h x)) (\x. g (iso_verts h x))" - for f g :: "'c \ real" + hence g_inner_conv: + "H.g_inner f g = g_inner (\x. f (iso_verts h x)) (\x. g (iso_verts h x))" + for f g :: "'c \ real" unfolding g_inner_def H.g_inner_def by (intro sum.reindex_bij_betw[symmetric]) have g_step_conv: "H.g_step f (iso_verts h x) = g_step (\x. f (iso_verts h x)) x" if "x \ verts G" - for f :: "'c \ real" and x + for f :: "'c \ real" and x proof - - have "inj_on (iso_arcs h) (in_arcs G x)" - using inj_on_subset[OF digraph_isomorphism_inj_on_arcs[OF hom_iso]] + have "inj_on (iso_arcs h) (in_arcs G x)" + using inj_on_subset[OF digraph_isomorphism_inj_on_arcs[OF hom_iso]] by (simp add:in_arcs_def) moreover have "in_arcs H (iso_verts h x) = iso_arcs h ` in_arcs G x" - unfolding H_alt by (intro in_arcs_app_iso_eq[OF hom_iso] that) - moreover have "tail H (iso_arcs h a) = iso_verts h (tail G a)" if "a \ in_arcs G x" for a + unfolding H_alt by (intro in_arcs_app_iso_eq[OF hom_iso] that) + moreover have "tail H (iso_arcs h a) = iso_verts h (tail G a)" if "a \ in_arcs G x" for a unfolding H_alt using that by (simp add: hom_iso iso_verts_tail) ultimately show ?thesis - unfolding g_step_def H.g_step_def + unfolding g_step_def H.g_step_def by (intro_cong "[\\<^sub>2(/), \\<^sub>1 f, \\<^sub>1 of_nat]" more: 4 sum.reindex_cong[where l="iso_arcs h"]) qed show "H.\\<^sub>a \ \\<^sub>a" - using expansionD1 by (intro H.expander_intro_1 \_ge_0) + using expansionD1 by (intro H.expander_intro_1 \_ge_0) (simp add:g_inner_conv g_step_conv H.g_norm_sq g_norm_sq cong:g_inner_cong) show "H.\\<^sub>2 \ \\<^sub>2" proof (cases "n > 1") case True hence "H.n > 1" by (simp add:n_eq) thus ?thesis - using os_expanderD by (intro H.os_expanderI) + using os_expanderD by (intro H.os_expanderI) (simp_all add:g_inner_conv g_step_conv H.g_norm_sq g_norm_sq cong:g_inner_cong) - next + next case False thus ?thesis unfolding H.\\<^sub>2_def \\<^sub>2_def by (simp add:n_eq) qed show "H.\\<^sub>e \ \\<^sub>e" proof (cases "n > 1") case True hence n_gt_1: "H.n > 1" by (simp add:n_eq) - have "\\<^sub>e * real (card S) \ real (card (H.edges_betw S (- S)))" - if "S \ verts H" "2 * card S \ H.n" "S \ {}" for S + have "\\<^sub>e * real (card S) \ real (card (H.edges_betw S (- S)))" + if "S \ verts H" "2 * card S \ H.n" "S \ {}" for S proof - define T where "T = iso_verts h -` S \ verts G" have 4:"card T = card S" using that(1) unfolding T_def H_alt verts_app_iso by (intro card_vimage_inj_on digraph_isomorphism_inj_on_verts[OF hom_iso]) auto have "card (H.edges_betw S (-S))=card {a\iso_arcs h`arcs G. iso_tail h a\S\iso_head h a\ -S}" unfolding H.edges_betw_def unfolding H_alt tail_app_iso head_app_iso arcs_app_iso by simp also have "...= card(iso_arcs h` {a \ arcs G. iso_tail h (iso_arcs h a)\S\ iso_head h (iso_arcs h a)\-S})" by (intro arg_cong[where f="card"]) auto also have "... = card {a \ arcs G. iso_tail h (iso_arcs h a)\S\ iso_head h (iso_arcs h a)\-S}" by (intro card_image inj_on_subset[OF digraph_isomorphism_inj_on_arcs[OF hom_iso]]) auto also have "... = card {a \ arcs G. iso_verts h (tail G a) \ S \ iso_verts h (head G a) \ -S}" by (intro restr_Collect_cong arg_cong[where f="card"]) (simp add: iso_verts_tail[OF hom_iso] iso_verts_head[OF hom_iso]) also have "... = card {a \ arcs G. tail G a \ T \ head G a \ -T }" unfolding T_def by (intro_cong "[\\<^sub>1(card),\\<^sub>2 (\)]" more: restr_Collect_cong) auto also have "... = card (edges_betw T (-T))" unfolding edges_betw_def by simp - finally have 5:"card (edges_betw T (-T)) = card (H.edges_betw S (-S))" + finally have 5:"card (edges_betw T (-T)) = card (H.edges_betw S (-S))" by simp have 6: "T \ verts G" unfolding T_def by simp have "\\<^sub>e * real (card S) = \\<^sub>e * real (card T)" unfolding 4 by simp also have "... \ real (card (edges_betw T (-T)))" using that(2) by (intro edge_expansionD 6) (simp add:4 n_eq) also have "... = real (card (H.edges_betw S (-S)))" unfolding 5 by simp finally show ?thesis by simp qed thus ?thesis by (intro H.edge_expansionI n_gt_1) auto next case False thus ?thesis unfolding H.\\<^sub>e_def \\<^sub>e_def by (simp add:n_eq) qed qed lemma (in regular_graph) assumes "digraph_iso G H" shows regular_graph_iso_expansion: "regular_graph.\\<^sub>a H = \\<^sub>a" and regular_graph_iso_os_expansion: "regular_graph.\\<^sub>2 H = \\<^sub>2" and regular_graph_iso_edge_expansion: "regular_graph.\\<^sub>e H = \\<^sub>e" proof - interpret H:"regular_graph" "H" by (intro regular_graph_iso assms) have iso:"digraph_iso H G" using digraph_iso_swap assms wf_digraph_axioms by blast hence "\\<^sub>a \ H.\\<^sub>a" by (intro H.regular_graph_iso_expansion_le) moreover have "H.\\<^sub>a \ \\<^sub>a" using regular_graph_iso_expansion_le[OF assms] by auto ultimately show "H.\\<^sub>a = \\<^sub>a" by auto have "\\<^sub>2 \ H.\\<^sub>2" using iso by (intro H.regular_graph_iso_os_expansion_le) moreover have "H.\\<^sub>2 \ \\<^sub>2" using regular_graph_iso_os_expansion_le[OF assms] by auto ultimately show "H.\\<^sub>2 = \\<^sub>2" by auto have "\\<^sub>e \ H.\\<^sub>e" using iso by (intro H.regular_graph_iso_edge_expansion_ge) moreover have "H.\\<^sub>e \ \\<^sub>e" using regular_graph_iso_edge_expansion_ge[OF assms] by auto ultimately show "H.\\<^sub>e = \\<^sub>e" by auto qed unbundle no_intro_cong_syntax end diff --git a/thys/Expander_Graphs/Expander_Graphs_Eigenvalues.thy b/thys/Expander_Graphs/Expander_Graphs_Eigenvalues.thy --- a/thys/Expander_Graphs/Expander_Graphs_Eigenvalues.thy +++ b/thys/Expander_Graphs/Expander_Graphs_Eigenvalues.thy @@ -1,1469 +1,1469 @@ section \Spectral Theory\label{sec:expander_eigenvalues}\ -text \This section establishes the correspondence of the variationally defined expansion paramters -with the definitions using the spectrum of the stochastic matrix. Additionally stronger +text \This section establishes the correspondence of the variationally defined expansion paramters +with the definitions using the spectrum of the stochastic matrix. Additionally stronger results for the expansion parameters are derived.\ theory Expander_Graphs_Eigenvalues - imports + imports Expander_Graphs_Algebra - Expander_Graphs_TTS + Expander_Graphs_TTS Perron_Frobenius.HMA_Connect Commuting_Hermitian.Commuting_Hermitian begin unbundle intro_cong_syntax hide_const Matrix_Legacy.transpose hide_const Matrix_Legacy.row hide_const Matrix_Legacy.mat hide_const Matrix.mat hide_const Matrix.row hide_fact Matrix_Legacy.row_def hide_fact Matrix_Legacy.mat_def hide_fact Matrix.vec_eq_iff hide_fact Matrix.mat_def hide_fact Matrix.row_def no_notation Matrix.scalar_prod (infix "\" 70) no_notation Ordered_Semiring.max ("Max\") -lemma mult_right_mono': "y \ (0::real) \ x \ z \ y = 0 \ x * y \ z * y" +lemma mult_right_mono': "y \ (0::real) \ x \ z \ y = 0 \ x * y \ z * y" by (metis mult_cancel_right mult_right_mono) lemma poly_prod_zero: fixes x :: "'a :: idom" assumes "poly (\a\#xs. [:- a, 1:]) x = 0" shows "x \# xs" using assms by (induction xs, auto) lemma poly_prod_inj_aux_1: fixes xs ys :: "('a :: idom) multiset" assumes "x \# xs" assumes "(\a\#xs. [:- a, 1:]) = (\a\#ys. [:- a, 1:])" shows "x \# ys" proof - have "poly (\a\#ys. [:- a, 1:]) x = poly (\a\#xs. [:- a, 1:]) x" using assms(2) by simp also have "... = poly (\a\#xs - {#x#} + {#x#}. [:- a, 1:]) x" using assms(1) by simp also have "... = 0" by simp finally have "poly (\a\#ys. [:- a, 1:]) x = 0" by simp thus "x \# ys" using poly_prod_zero by blast qed lemma poly_prod_inj_aux_2: fixes xs ys :: "('a :: idom) multiset" assumes "x \# xs \# ys" assumes "(\a\#xs. [:- a, 1:]) = (\a\#ys. [:- a, 1:])" shows "x \# xs \# ys" proof (cases "x \# xs") case True then show ?thesis using poly_prod_inj_aux_1[OF True assms(2)] by simp next case False hence a:"x \# ys" using assms(1) by simp - then show ?thesis + then show ?thesis using poly_prod_inj_aux_1[OF a assms(2)[symmetric]] by simp qed lemma poly_prod_inj: fixes xs ys :: "('a :: idom) multiset" assumes "(\a\#xs. [:- a, 1:]) = (\a\#ys. [:- a, 1:])" shows "xs = ys" using assms proof (induction "size xs + size ys" arbitrary: xs ys rule:nat_less_induct) case 1 show ?case proof (cases "xs \# ys = {#}") case True then show ?thesis by simp next case False then obtain x where "x \# xs \# ys" by auto hence a:"x \# xs \# ys" by (intro poly_prod_inj_aux_2[OF _ 1(2)]) - have b: "[:- x, 1:] \ 0" + have b: "[:- x, 1:] \ 0" by simp - have c: "size (xs-{#x#}) + size (ys-{#x#}) < size xs + size ys" + have c: "size (xs-{#x#}) + size (ys-{#x#}) < size xs + size ys" using a by (simp add: add_less_le_mono size_Diff1_le size_Diff1_less) have "[:- x, 1:] * (\a\#xs - {#x#}. [:- a, 1:]) = (\a\#xs. [:- a, 1:])" using a by (subst prod_mset.insert[symmetric]) simp also have "... = (\a\#ys. [:- a, 1:])" using 1 by simp also have "... = [:- x, 1:] * (\a\#ys - {#x#}. [:- a, 1:])" using a by (subst prod_mset.insert[symmetric]) simp finally have "[:- x, 1:]*(\a\#xs-{#x#}. [:- a, 1:])=[:-x, 1:]*(\a\#ys-{#x#}. [:- a, 1:])" by simp - hence "(\a\#xs-{#x#}. [:- a, 1:]) = (\a\#ys-{#x#}. [:- a, 1:])" + hence "(\a\#xs-{#x#}. [:- a, 1:]) = (\a\#ys-{#x#}. [:- a, 1:])" using mult_left_cancel[OF b] by simp hence d:"xs - {#x#} = ys - {#x#}" using 1 c by simp have "xs = xs - {#x#} + {#x#}" using a by simp also have "... = ys - {#x#} + {#x#}" unfolding d by simp also have "... = ys" using a by simp finally show ?thesis by simp qed qed -definition eigenvalues :: "('a::comm_ring_1)^'n^'n \ 'a multiset" - where +definition eigenvalues :: "('a::comm_ring_1)^'n^'n \ 'a multiset" + where "eigenvalues A = (SOME as. charpoly A = (\a\#as. [:- a, 1:]) \ size as = CARD ('n))" -lemma char_poly_factorized_hma: +lemma char_poly_factorized_hma: fixes A :: "complex^'n^'n" shows "\as. charpoly A = (\a\as. [:- a, 1:]) \ length as = CARD ('n)" by (transfer_hma rule:char_poly_factorized) lemma eigvals_poly_length: fixes A :: "complex^'n^'n" - shows - "charpoly A = (\a\#eigenvalues A. [:- a, 1:])" (is "?A") + shows + "charpoly A = (\a\#eigenvalues A. [:- a, 1:])" (is "?A") "size (eigenvalues A) = CARD ('n)" (is "?B") proof - define f where "f as = (charpoly A = (\a\#as. [:- a, 1:]) \ size as = CARD('n))" for as obtain as where as_def: "charpoly A = (\a\as. [:- a, 1:])" "length as = CARD('n)" using char_poly_factorized_hma by auto have "charpoly A = (\a\as. [:- a, 1:])" unfolding as_def by simp - also have "... = (\a\#mset as. [:- a, 1:])" + also have "... = (\a\#mset as. [:- a, 1:])" unfolding prod_mset_prod_list[symmetric] mset_map by simp finally have "charpoly A = (\a\#mset as. [:- a, 1:])" by simp moreover have "size (mset as) = CARD('n)" using as_def by simp - ultimately have "f (mset as)" + ultimately have "f (mset as)" unfolding f_def by auto - hence "f (eigenvalues A)" + hence "f (eigenvalues A)" unfolding eigenvalues_def f_def[symmetric] using someI[where x = "mset as" and P="f"] by auto thus ?A ?B unfolding f_def by auto qed lemma similar_matrix_eigvals: fixes A B :: "complex^'n^'n" assumes "similar_matrix A B" shows "eigenvalues A = eigenvalues B" proof - have "(\a\#eigenvalues A. [:- a, 1:]) = (\a\#eigenvalues B. [:- a, 1:])" using similar_matrix_charpoly[OF assms] unfolding eigvals_poly_length(1) by simp thus ?thesis by (intro poly_prod_inj) simp qed definition upper_triangular_hma :: "'a::zero^'n^'n \ bool" where "upper_triangular_hma A \ \i. \ j. (to_nat j < Bij_Nat.to_nat i \ A $h i $h j = 0)" lemma for_all_reindex2: assumes "range f = A" shows "(\x \ A. \y \ A. P x y) \ (\x y. P (f x) (f y))" using assms by auto -lemma upper_triangular_hma: +lemma upper_triangular_hma: fixes A :: "('a::zero)^'n^'n" shows "upper_triangular (from_hma\<^sub>m A) = upper_triangular_hma A" (is "?L = ?R") proof - have "?L \ (\i\{0..j\{0.. A $h from_nat i $h from_nat j = 0)" unfolding upper_triangular_def from_hma\<^sub>m_def by auto also have "... \ (\(i::'n) (j::'n). to_nat j < to_nat i \ A $h from_nat (to_nat i) $h from_nat (to_nat j) = 0)" by (intro for_all_reindex2 range_to_nat[where 'a="'n"]) also have "... \ ?R" unfolding upper_triangular_hma_def by auto finally show ?thesis by simp qed -lemma from_hma_carrier: +lemma from_hma_carrier: fixes A :: "'a^('n::finite)^('m::finite)" shows "from_hma\<^sub>m A \ carrier_mat (CARD ('m)) (CARD ('n))" unfolding from_hma\<^sub>m_def by simp definition diag_mat_hma :: "'a^'n^'n \ 'a multiset" where "diag_mat_hma A = image_mset (\i. A $h i $h i) (mset_set UNIV)" -lemma diag_mat_hma: +lemma diag_mat_hma: fixes A :: "'a^'n^'n" shows "mset (diag_mat (from_hma\<^sub>m A)) = diag_mat_hma A" (is "?L = ?R") -proof - - have "?L = {#from_hma\<^sub>m A $$ (i, i). i \# mset [0..m A $$ (i, i). i \# mset [0..m A $$ (i, i). i \# image_mset to_nat (mset_set (UNIV :: 'n set))#}" using range_to_nat[where 'a="'n"] by (intro arg_cong2[where f="image_mset"] refl) (simp add:image_mset_mset_set[OF inj_to_nat]) also have "... = {#from_hma\<^sub>m A $$ (to_nat i, to_nat i). i \# (mset_set (UNIV :: 'n set))#}" by (simp add:image_mset.compositionality comp_def) also have "... = ?R" unfolding diag_mat_hma_def from_hma\<^sub>m_def using to_nat_less_card[where 'a="'n"] by (intro image_mset_cong) auto finally show ?thesis by simp qed definition adjoint_hma :: "complex^'m^'n \ complex^'n^'m" where "adjoint_hma A = map_matrix cnj (transpose A)" lemma adjoint_hma_eq: "adjoint_hma A $h i $h j = cnj (A $h j $h i)" unfolding adjoint_hma_def map_matrix_def map_vector_def transpose_def by auto lemma adjoint_hma: fixes A :: "complex^('n::finite)^('m::finite)" shows "mat_adjoint (from_hma\<^sub>m A) = from_hma\<^sub>m (adjoint_hma A)" proof - have "mat_adjoint (from_hma\<^sub>m A) $$ (i,j) = from_hma\<^sub>m (adjoint_hma A) $$ (i,j)" if "i < CARD('n)" "j < CARD('m)" for i j - using from_hma_carrier that unfolding mat_adjoint_def from_hma\<^sub>m_def adjoint_hma_def + using from_hma_carrier that unfolding mat_adjoint_def from_hma\<^sub>m_def adjoint_hma_def Matrix.mat_of_rows_def map_matrix_def map_vector_def transpose_def by auto thus ?thesis using from_hma_carrier by (intro eq_matI) auto qed definition cinner where "cinner v w = scalar_product v (map_vector cnj w)" -context - includes lifting_syntax +context + includes lifting_syntax begin -lemma cinner_hma: +lemma cinner_hma: fixes x y :: "complex^'n" shows "cinner x y = (from_hma\<^sub>v x) \c (from_hma\<^sub>v y)" (is "?L = ?R") proof - - have "?L = (\i\UNIV. x $h i * cnj (y $h i))" + have "?L = (\i\UNIV. x $h i * cnj (y $h i))" unfolding cinner_def map_vector_def scalar_product_def by simp also have "... = (\i = 0..v_def by simp finally show ?thesis by simp qed -lemma cinner_hma_transfer[transfer_rule]: +lemma cinner_hma_transfer[transfer_rule]: "(HMA_V ===> HMA_V ===> (=)) (\c) cinner" unfolding HMA_V_def cinner_hma by (auto simp:rel_fun_def) -lemma adjoint_hma_transfer[transfer_rule]: +lemma adjoint_hma_transfer[transfer_rule]: "(HMA_M ===> HMA_M) (mat_adjoint) adjoint_hma" unfolding HMA_M_def rel_fun_def by (auto simp add:adjoint_hma) end lemma adjoint_adjoint_id[simp]: "adjoint_hma (adjoint_hma A ) = A" by (transfer) (simp add:adjoint_adjoint) -lemma adjoint_def_alter_hma: - "cinner (A *v v) w = cinner v (adjoint_hma A *v w)" +lemma adjoint_def_alter_hma: + "cinner (A *v v) w = cinner v (adjoint_hma A *v w)" by (transfer_hma rule:adjoint_def_alter) lemma cinner_0: "cinner 0 0 = 0" by (transfer_hma) lemma cinner_scale_left: "cinner (a *s v) w = a * cinner v w" by transfer_hma lemma cinner_scale_right: "cinner v (a *s w) = cnj a * cinner v w" by transfer (simp add: inner_prod_smult_right) -lemma norm_of_real: +lemma norm_of_real: shows "norm (map_vector complex_of_real v) = norm v" unfolding norm_vec_def map_vector_def by (intro L2_set_cong) auto definition unitary_hma :: "complex^'n^'n \ bool" where "unitary_hma A \ A ** adjoint_hma A = Finite_Cartesian_Product.mat 1" definition unitarily_equiv_hma where "unitarily_equiv_hma A B U \ (unitary_hma U \ similar_matrix_wit A B U (adjoint_hma U))" definition diagonal_mat :: "('a::zero)^('n::finite)^'n \ bool" where "diagonal_mat A \ (\i. \j. i \ j \ A $h i $h j = 0)" -lemma diagonal_mat_ex: +lemma diagonal_mat_ex: assumes "diagonal_mat A" shows "A = diag (\ i. A $h i $h i)" - using assms unfolding diagonal_mat_def diag_def + using assms unfolding diagonal_mat_def diag_def by (intro iffD2[OF vec_eq_iff] allI) auto lemma diag_diagonal_mat[simp]: "diagonal_mat (diag x)" unfolding diag_def diagonal_mat_def by auto lemma diag_imp_upper_tri: "diagonal_mat A \ upper_triangular_hma A" - unfolding diagonal_mat_def upper_triangular_hma_def + unfolding diagonal_mat_def upper_triangular_hma_def by (metis nat_neq_iff) definition unitary_diag where "unitary_diag A b U \ unitarily_equiv_hma A (diag b) U" definition real_diag_decomp_hma where - "real_diag_decomp_hma A d U \ unitary_diag A d U \ + "real_diag_decomp_hma A d U \ unitary_diag A d U \ (\i. d $h i \ Reals)" -definition hermitian_hma :: "complex^'n^'n \ bool" where +definition hermitian_hma :: "complex^'n^'n \ bool" where "hermitian_hma A = (adjoint_hma A = A)" lemma from_hma_one: "from_hma\<^sub>m (mat 1 :: (('a::{one,zero})^'n^'n)) = 1\<^sub>m CARD('n)" unfolding Finite_Cartesian_Product.mat_def from_hma\<^sub>m_def using from_nat_inj by (intro eq_matI) auto -lemma from_hma_mult: +lemma from_hma_mult: fixes A :: "('a :: semiring_1)^'m^'n" fixes B :: "'a^'k^'m::finite" shows "from_hma\<^sub>m A * from_hma\<^sub>m B = from_hma\<^sub>m (A ** B)" using HMA_M_mult unfolding rel_fun_def HMA_M_def by auto lemma hermitian_hma: "hermitian_hma A = hermitian (from_hma\<^sub>m A)" unfolding hermitian_def adjoint_hma hermitian_hma_def by auto lemma unitary_hma: fixes A :: "complex^'n^'n" shows "unitary_hma A = unitary (from_hma\<^sub>m A)" (is "?L = ?R") proof - have "?R \ from_hma\<^sub>m A * mat_adjoint (from_hma\<^sub>m A) = 1\<^sub>m (CARD('n))" using from_hma_carrier unfolding unitary_def inverts_mat_def by simp also have "... \ from_hma\<^sub>m (A ** adjoint_hma A) = from_hma\<^sub>m (mat 1::complex^'n^'n)" unfolding adjoint_hma from_hma_mult from_hma_one by simp also have "... \ A ** adjoint_hma A = Finite_Cartesian_Product.mat 1" unfolding from_hma\<^sub>m_inj by simp also have "... \ ?L" unfolding unitary_hma_def by simp finally show ?thesis by simp qed lemma unitary_hmaD: fixes A :: "complex^'n^'n" assumes "unitary_hma A" shows "adjoint_hma A ** A = mat 1" (is "?A") "A ** adjoint_hma A = mat 1" (is "?B") proof - have "mat_adjoint (from_hma\<^sub>m A) * from_hma\<^sub>m A = 1\<^sub>m CARD('n)" using assms unitary_hma by (intro unitary_simps from_hma_carrier ) auto thus ?A unfolding adjoint_hma from_hma_mult from_hma_one[symmetric] from_hma\<^sub>m_inj by simp show ?B using assms unfolding unitary_hma_def by simp qed lemma unitary_hma_adjoint: assumes "unitary_hma A" shows "unitary_hma (adjoint_hma A)" unfolding unitary_hma_def adjoint_adjoint_id unitary_hmaD[OF assms] by simp lemma unitarily_equiv_hma: fixes A :: "complex^'n^'n" - shows "unitarily_equiv_hma A B U = + shows "unitarily_equiv_hma A B U = unitarily_equiv (from_hma\<^sub>m A) (from_hma\<^sub>m B) (from_hma\<^sub>m U)" (is "?L = ?R") proof - have "?R \ (unitary_hma U \ similar_mat_wit (from_hma\<^sub>m A) (from_hma\<^sub>m B) (from_hma\<^sub>m U) (from_hma\<^sub>m (adjoint_hma U)))" unfolding Spectral_Theory_Complements.unitarily_equiv_def unitary_hma[symmetric] adjoint_hma by simp also have "... \ unitary_hma U \ similar_matrix_wit A B U (adjoint_hma U)" using HMA_similar_mat_wit unfolding rel_fun_def HMA_M_def - by (intro arg_cong2[where f="(\)"] refl) force + by (intro arg_cong2[where f="(\)"] refl) force also have "... \ ?L" unfolding unitarily_equiv_hma_def by auto finally show ?thesis by simp qed -lemma Matrix_diagonal_matD: +lemma Matrix_diagonal_matD: assumes "Matrix.diagonal_mat A" assumes "i j" shows "A $$ (i,j) = 0" using assms unfolding Matrix.diagonal_mat_def by auto lemma diagonal_mat_hma: fixes A :: "('a :: zero)^('n :: finite)^'n" shows "diagonal_mat A = Matrix.diagonal_mat (from_hma\<^sub>m A)" (is "?L = ?R") -proof +proof show "?L \ ?R" - unfolding diagonal_mat_def Matrix.diagonal_mat_def from_hma\<^sub>m_def + unfolding diagonal_mat_def Matrix.diagonal_mat_def from_hma\<^sub>m_def using from_nat_inj by auto next assume a:"?R" have "A $h i $h j = 0" if "i \ j" for i j proof - have "A $h i $h j = (from_hma\<^sub>m A) $$ (to_nat i,to_nat j)" - unfolding from_hma\<^sub>m_def using to_nat_less_card[where 'a="'n"] by simp + unfolding from_hma\<^sub>m_def using to_nat_less_card[where 'a="'n"] by simp also have "... = 0" - using to_nat_less_card[where 'a="'n"] to_nat_inj that + using to_nat_less_card[where 'a="'n"] to_nat_inj that by (intro Matrix_diagonal_matD[OF a]) auto finally show ?thesis by simp qed thus "?L" unfolding diagonal_mat_def by auto qed lemma unitary_diag_hma: fixes A :: "complex^'n^'n" - shows "unitary_diag A d U = + shows "unitary_diag A d U = Spectral_Theory_Complements.unitary_diag (from_hma\<^sub>m A) (from_hma\<^sub>m (diag d)) (from_hma\<^sub>m U)" proof - have "Matrix.diagonal_mat (from_hma\<^sub>m (diag d))" unfolding diagonal_mat_hma[symmetric] by simp thus ?thesis unfolding unitary_diag_def Spectral_Theory_Complements.unitary_diag_def unitarily_equiv_hma by auto qed lemma real_diag_decomp_hma: fixes A :: "complex^'n^'n" - shows "real_diag_decomp_hma A d U = + shows "real_diag_decomp_hma A d U = real_diag_decomp (from_hma\<^sub>m A) (from_hma\<^sub>m (diag d)) (from_hma\<^sub>m U)" proof - have 0:"(\i. d $h i \ \) \ (\i < CARD('n). from_hma\<^sub>m (diag d) $$ (i,i) \ \)" unfolding from_hma\<^sub>m_def diag_def using to_nat_less_card by fastforce show ?thesis unfolding real_diag_decomp_hma_def real_diag_decomp_def unitary_diag_hma 0 by auto qed lemma diagonal_mat_diag_ex_hma: assumes "Matrix.diagonal_mat A" "A \ carrier_mat CARD('n) CARD ('n :: finite)" shows "from_hma\<^sub>m (diag (\ (i::'n). A $$ (to_nat i,to_nat i))) = A" using assms from_nat_inj unfolding from_hma\<^sub>m_def diag_def Matrix.diagonal_mat_def by (intro eq_matI) (auto simp add:to_nat_from_nat_id) theorem commuting_hermitian_family_diag_hma: fixes Af :: "(complex^'n^'n) set" assumes "finite Af" and "Af \ {}" and "\A. A \ Af \ hermitian_hma A" and "\A B. A \ Af \ B\ Af \ A ** B = B ** A" - shows "\ U. \ A\ Af. \B. real_diag_decomp_hma A B U" + shows "\ U. \ A\ Af. \B. real_diag_decomp_hma A B U" proof - have 0:"finite (from_hma\<^sub>m ` Af)" - using assms(1)by (intro finite_imageI) + using assms(1)by (intro finite_imageI) have 1: "from_hma\<^sub>m ` Af \ {}" using assms(2) by simp have 2: "A \ carrier_mat (CARD ('n)) (CARD ('n))" if "A \ from_hma\<^sub>m ` Af" for A using that unfolding from_hma\<^sub>m_def by (auto simp add:image_iff) have 3: "0 < CARD('n)" by simp have 4: "hermitian A" if "A \ from_hma\<^sub>m ` Af" for A using hermitian_hma assms(3) that by auto have 5: "A * B = B * A" if "A \ from_hma\<^sub>m ` Af" "B \ from_hma\<^sub>m ` Af" for A B using that assms(4) by (auto simp add:image_iff from_hma_mult) have "\U. \A\ from_hma\<^sub>m ` Af. \B. real_diag_decomp A B U" using commuting_hermitian_family_diag[OF 0 1 2 3 4 5] by auto then obtain U Bmap where U_def: "\A. A \ from_hma\<^sub>m ` Af \ real_diag_decomp A (Bmap A) U" by metis define U' :: "complex^'n^'n" where "U' = to_hma\<^sub>m U" - define Bmap' :: "complex^'n^'n \ complex^'n" + define Bmap' :: "complex^'n^'n \ complex^'n" where "Bmap' = (\M. (\ i. (Bmap (from_hma\<^sub>m M)) $$ (to_nat i,to_nat i)))" - have "real_diag_decomp_hma A (Bmap' A) U'" if "A \ Af" for A + have "real_diag_decomp_hma A (Bmap' A) U'" if "A \ Af" for A proof - have rdd: "real_diag_decomp (from_hma\<^sub>m A) (Bmap (from_hma\<^sub>m A)) U" using U_def that by simp - have "U \ carrier_mat CARD('n) CARD('n)" "Bmap (from_hma\<^sub>m A) \ carrier_mat CARD('n) CARD('n)" + have "U \ carrier_mat CARD('n) CARD('n)" "Bmap (from_hma\<^sub>m A) \ carrier_mat CARD('n) CARD('n)" "Matrix.diagonal_mat (Bmap (from_hma\<^sub>m A))" using rdd unfolding real_diag_decomp_def Spectral_Theory_Complements.unitary_diag_def Spectral_Theory_Complements.unitarily_equiv_def similar_mat_wit_def by (auto simp add:Let_def) hence "(from_hma\<^sub>m (diag (Bmap' A))) = Bmap (from_hma\<^sub>m A)" "(from_hma\<^sub>m U') = U" unfolding Bmap'_def U'_def by (auto simp add:diagonal_mat_diag_ex_hma) hence "real_diag_decomp (from_hma\<^sub>m A) (from_hma\<^sub>m (diag (Bmap' A))) (from_hma\<^sub>m U')" using rdd by auto thus "?thesis" unfolding real_diag_decomp_hma by simp qed thus ?thesis by (intro exI[where x="U'"]) auto qed -lemma char_poly_upper_triangular: +lemma char_poly_upper_triangular: fixes A :: "complex^'n^'n" assumes "upper_triangular_hma A" shows "charpoly A = (\a \# diag_mat_hma A. [:- a, 1:])" proof - have "charpoly A = char_poly (from_hma\<^sub>m A)" using HMA_char_poly unfolding rel_fun_def HMA_M_def by (auto simp add:eq_commute) also have "... = (\a\diag_mat (from_hma\<^sub>m A). [:- a, 1:])" using assms unfolding upper_triangular_hma[symmetric] by (intro char_poly_upper_triangular[where n="CARD('n)"] from_hma_carrier) auto also have "... = (\a\# mset (diag_mat (from_hma\<^sub>m A)). [:- a, 1:])" unfolding prod_mset_prod_list[symmetric] mset_map by simp also have "... = (\a\# diag_mat_hma A. [:- a, 1:])" unfolding diag_mat_hma by simp finally show "charpoly A = (\a\# diag_mat_hma A. [:- a, 1:])" by simp qed lemma upper_tri_eigvals: fixes A :: "complex^'n^'n" assumes "upper_triangular_hma A" shows "eigenvalues A = diag_mat_hma A" proof - have "(\a\#eigenvalues A. [:- a, 1:]) = charpoly A" unfolding eigvals_poly_length[symmetric] by simp also have "... = (\a\#diag_mat_hma A. [:- a, 1:])" by (intro char_poly_upper_triangular assms) finally have "(\a\#eigenvalues A. [:- a, 1:]) = (\a\#diag_mat_hma A. [:- a, 1:])" by simp thus ?thesis by (intro poly_prod_inj) simp qed lemma cinner_self: fixes v :: "complex^'n" shows "cinner v v = norm v^2" proof - - have 0: "x * cnj x = complex_of_real (x \ x)" for x :: complex + have 0: "x * cnj x = complex_of_real (x \ x)" for x :: complex unfolding inner_complex_def complex_mult_cnj by (simp add:power2_eq_square) thus ?thesis - unfolding cinner_def power2_norm_eq_inner scalar_product_def inner_vec_def - map_vector_def by simp + unfolding cinner_def power2_norm_eq_inner scalar_product_def inner_vec_def + map_vector_def by simp qed lemma unitary_iso: - assumes "unitary_hma U" + assumes "unitary_hma U" shows "norm (U *v v) = norm v" proof - have "norm (U *v v)^2 = cinner (U *v v) (U *v v)" unfolding cinner_self by simp also have "... = cinner v v" unfolding adjoint_def_alter_hma matrix_vector_mul_assoc unitary_hmaD[OF assms] by simp also have "... = norm v^2" unfolding cinner_self by simp finally have "complex_of_real (norm (U *v v)^2) = norm v^2" by simp thus ?thesis by (meson norm_ge_zero of_real_hom.injectivity power2_eq_iff_nonneg) qed lemma (in semiring_hom) mult_mat_vec_hma: "map_vector hom (A *v v) = map_matrix hom A *v map_vector hom v" using mult_mat_vec_hom by transfer auto lemma (in semiring_hom) mat_hom_mult_hma: "map_matrix hom (A ** B) = map_matrix hom A ** map_matrix hom B" using mat_hom_mult by transfer auto -context regular_graph_tts +context regular_graph_tts begin lemma to_nat_less_n: "to_nat (x::'n) < n" - using to_nat_less_card card_n by metis + using to_nat_less_card card_n by metis lemma to_nat_from_nat: "x < n \ to_nat (from_nat x :: 'n) = x" using to_nat_from_nat_id card_n by metis lemma hermitian_A: "hermitian_hma A" - using count_sym unfolding hermitian_hma_def adjoint_hma_def A_def map_matrix_def + using count_sym unfolding hermitian_hma_def adjoint_hma_def A_def map_matrix_def map_vector_def transpose_def by simp lemma nonneg_A: "nonneg_mat A" unfolding nonneg_mat_def A_def by auto lemma g_step_1: assumes "v \ verts G" shows "g_step (\_. 1) v = 1" (is "?L = ?R") proof - have "?L = in_degree G v / d" unfolding g_step_def in_degree_def by simp also have "... = 1" unfolding reg(2)[OF assms] using d_gt_0 by simp finally show ?thesis by simp qed lemma markov: "markov (A :: real^'n^'n)" proof - have "A *v 1 = (1::real ^'n)" (is "?L = ?R") proof - have "A *v 1 = (\ i. g_step (\_. 1) (enum_verts i))" unfolding g_step_conv one_vec_def by simp also have "... = (\ i. 1)" using bij_betw_apply[OF enum_verts] by (subst g_step_1) auto also have "... = 1" unfolding one_vec_def by simp finally show ?thesis by simp qed thus ?thesis by (intro markov_symI nonneg_A symmetric_A) qed lemma nonneg_J: "nonneg_mat J" unfolding nonneg_mat_def J_def by auto lemma J_eigvals: "eigenvalues J = {#1::complex#} + replicate_mset (n - 1) 0" proof - define \ :: "nat \ real" where "\ i = sqrt (i^2+i)" for i :: nat define q :: "nat \ nat \ real" where "q i j = ( if i = 0 then (1/sqrt n) else ( if j < i then ((-1) / \ i) else ( if j = i then (i / \ i) else 0)))" for i j define Q :: "complex^'n^'n" where "Q = (\ i j. complex_of_real (q (to_nat i) (to_nat j)))" - define D :: "complex^'n^'n" where + define D :: "complex^'n^'n" where "D = (\ i j. if to_nat i = 0 \ to_nat j = 0 then 1 else 0)" have 2: "[0..k = 0.. j" "j < n" for i j proof - consider (a) "i = j \ j = 0" | (b) "i = 0 \ i < j" | (c) " 0 < i \ i < j" | (d) "0 < i \ i = j" using 1 by linarith thus ?thesis proof (cases) case a then show ?thesis using n_gt_0 by (simp add:q_def) next case b have "(\k = 0..k\insert j ({0.. {j+1..k=0..k=j+1..k = 0..k\insert i ({0.. {i+1..k=0..k=i+1.. j * i / \ i+ i * ((-1) / \ j * (-1) / \ i)" using c unfolding q_def by simp also have "... = 0" by (simp add:algebra_simps) finally show ?thesis using c by simp next case d have "real i + real i^2 = real (i + i^2)" by simp - also have "... \ real 0" + also have "... \ real 0" unfolding of_nat_eq_iff using d by simp finally have d_1: "real i + real i^2 \ 0" by simp have "(\k = 0..k\insert i ({0.. {i+1..k=0..k=i+1.. i * i / \ i+ i * ((-1) / \ i * (-1) / \ i)" using d that unfolding q_def by simp also have "... = (i^2 + i) / (\ i)^2" by (simp add: power2_eq_square divide_simps) also have "... = 1" - using d_1 unfolding \_def by (simp add:algebra_simps) + using d_1 unfolding \_def by (simp add:algebra_simps) finally show ?thesis using d by simp qed qed have 0:"(\k = 0..k = 0.. j") ( simp_all add:ac_simps cong:sum.cong) also have "... = of_bool (min i j = max i j)" using that by (intro aux0) auto - also have "... = ?R" + also have "... = ?R" by (cases "i \ j") auto finally show ?thesis by simp qed - + have "(\k\UNIV. Q $h j $h k * cnj (Q $h i $h k)) = of_bool (i=j)" (is "?L = ?R") for i j proof - - have "?L = complex_of_real (\k \ (UNIV::'n set). q (to_nat j) (to_nat k) * q (to_nat i) (to_nat k))" - unfolding Q_def + have "?L = complex_of_real (\k \ (UNIV::'n set). q (to_nat j) (to_nat k) * q (to_nat i) (to_nat k))" + unfolding Q_def by (simp add:case_prod_beta scalar_prod_def map_vector_def inner_vec_def row_def inner_complex_def) also have "... = complex_of_real (\k=0.. i j. (if to_nat j = 0 then complex_of_real (1/sqrt n) else 0))" unfolding Q_def D_def by (intro iffD2[OF vec_eq_iff] allI) (auto simp add:adjoint_hma_eq matrix_matrix_mult_def q_def if_distrib if_distribR sum.If_cases) - have "(adjoint_hma Q ** D ** Q) $h i $h j = J $h i $h j" (is "?L1 = ?R1") for i j - proof - + have "(adjoint_hma Q ** D ** Q) $h i $h j = J $h i $h j" (is "?L1 = ?R1") for i j + proof - have "?L1 =1/((sqrt (real n)) * complex_of_real (sqrt (real n)))" unfolding 1 unfolding Q_def using n_gt_0 5 - by (auto simp add:matrix_matrix_mult_def q_def if_distrib if_distribR sum.If_cases) + by (auto simp add:matrix_matrix_mult_def q_def if_distrib if_distribR sum.If_cases) also have "... = 1/sqrt (real n)^2" unfolding of_real_divide of_real_mult power2_eq_square by simp also have "... = J $h i $h j" unfolding J_def card_n using n_gt_0 by simp finally show ?thesis by simp qed hence "adjoint_hma Q ** D ** Q = J" - by (intro iffD2[OF vec_eq_iff] allI) auto + by (intro iffD2[OF vec_eq_iff] allI) auto - hence "similar_matrix_wit J D (adjoint_hma Q) Q" + hence "similar_matrix_wit J D (adjoint_hma Q) Q" unfolding similar_matrix_wit_def unitary_hmaD[OF unit_Q] by auto hence "similar_matrix J D" unfolding similar_matrix_def by auto hence "eigenvalues J = eigenvalues D" by (intro similar_matrix_eigvals) also have "... = diag_mat_hma D" by (intro upper_tri_eigvals diag_imp_upper_tri) (simp add:D_def diagonal_mat_def) also have "... = {# of_bool (to_nat i = 0). i \# mset_set (UNIV :: 'n set)#}" - unfolding diag_mat_hma_def D_def of_bool_def by simp + unfolding diag_mat_hma_def D_def of_bool_def by simp also have "... = {# of_bool (i = 0). i \# mset_set (to_nat ` (UNIV :: 'n set))#}" - unfolding image_mset_mset_set[OF inj_to_nat, symmetric] + unfolding image_mset_mset_set[OF inj_to_nat, symmetric] by (simp add:image_mset.compositionality comp_def) also have "... = mset (map (\i. of_bool(i=0)) [0..i. 0) [1..k\UNIV. A $h k $h i / real CARD('n)) = 1 / real CARD('n)" (is "?L = ?R") for i + have 0: "(\k\UNIV. A $h k $h i / real CARD('n)) = 1 / real CARD('n)" (is "?L = ?R") for i proof - have "?L = (1 v* A) $h i / real CARD('n)" - unfolding vector_matrix_mult_def by (simp add:sum_divide_distrib) + unfolding vector_matrix_mult_def by (simp add:sum_divide_distrib) also have "... = ?R" unfolding markov_apply[OF markov] by simp finally show ?thesis by simp qed - have 1: "(\k\UNIV. A $h i $h k / real CARD('n)) = 1 / real CARD('n)" (is "?L = ?R") for i + have 1: "(\k\UNIV. A $h i $h k / real CARD('n)) = 1 / real CARD('n)" (is "?L = ?R") for i proof - have "?L = (A *v 1) $h i / real CARD('n)" - unfolding matrix_vector_mult_def by (simp add:sum_divide_distrib) + unfolding matrix_vector_mult_def by (simp add:sum_divide_distrib) also have "... = ?R" unfolding markov_apply[OF markov] by simp finally show ?thesis by simp qed show ?thesis unfolding J_def using 0 1 by (intro iffD2[OF vec_eq_iff] allI) (simp add:matrix_matrix_mult_def) qed lemma J_A_comm: "J ** A = A ** (J :: complex^'n^'n)" (is "?L = ?R") proof - have "J ** A = map_matrix complex_of_real (J ** A)" unfolding of_real_hom.mat_hom_mult_hma J_def A_def by (auto simp add:map_matrix_def map_vector_def) also have "... = map_matrix complex_of_real (A ** J)" unfolding J_A_comm_real by simp also have "... = map_matrix complex_of_real A ** map_matrix complex_of_real J" unfolding of_real_hom.mat_hom_mult_hma by simp also have "... = ?R" unfolding A_def J_def by (auto simp add:map_matrix_def map_vector_def) finally show ?thesis by simp qed definition \\<^sub>a :: "'n itself \ real" where "\\<^sub>a _ = (if n > 1 then Max_mset (image_mset cmod (eigenvalues A - {#1#})) else 0)" definition \\<^sub>2 :: "'n itself \ real" where "\\<^sub>2 _ = (if n > 1 then Max_mset {# Re x. x \# (eigenvalues A - {#1#})#} else 0)" lemma J_sym: "hermitian_hma J" unfolding J_def hermitian_hma_def by (intro iffD2[OF vec_eq_iff] allI) (simp add: adjoint_hma_eq) lemma shows evs_real: "set_mset (eigenvalues A::complex multiset) \ \" (is "?R1") and ev_1: "(1::complex) \# eigenvalues A" and \\<^sub>a_ge_0: "\\<^sub>a TYPE ('n) \ 0" - and find_any_ev: - "\\ \# eigenvalues A - {#1#}. \v. cinner v 1 = 0 \ v \ 0 \ A *v v = \ *s v" + and find_any_ev: + "\\ \# eigenvalues A - {#1#}. \v. cinner v 1 = 0 \ v \ 0 \ A *v v = \ *s v" and \\<^sub>a_bound: "\v. cinner v 1 = 0 \ norm (A *v v) \ \\<^sub>a TYPE('n) * norm v" and \\<^sub>2_bound: "\(v::real^'n). v \ 1 = 0 \ v \ (A *v v) \ \\<^sub>2 TYPE ('n) * norm v^2" proof - - have "\ U. \ A\ {J,A}. \B. real_diag_decomp_hma A B U" + have "\ U. \ A\ {J,A}. \B. real_diag_decomp_hma A B U" using J_sym hermitian_A J_A_comm by (intro commuting_hermitian_family_diag_hma) auto - then obtain U Ad Jd + then obtain U Ad Jd where A_decomp: "real_diag_decomp_hma A Ad U" and K_decomp: "real_diag_decomp_hma J Jd U" by auto - have J_sim: "similar_matrix_wit J (diag Jd) U (adjoint_hma U)" and + have J_sim: "similar_matrix_wit J (diag Jd) U (adjoint_hma U)" and unit_U: "unitary_hma U" using K_decomp unfolding real_diag_decomp_hma_def unitary_diag_def unitarily_equiv_hma_def by auto have "diag_mat_hma (diag Jd) = eigenvalues (diag Jd)" by (intro upper_tri_eigvals[symmetric] diag_imp_upper_tri J_sim) auto also have "... = eigenvalues J" using J_sim by (intro similar_matrix_eigvals[symmetric]) (auto simp add:similar_matrix_def) also have "... ={#1::complex#} + replicate_mset (n - 1) 0" unfolding J_eigvals by simp finally have 0:"diag_mat_hma (diag Jd) = {#1::complex#} + replicate_mset (n - 1) 0" by simp hence "1 \# diag_mat_hma (diag Jd)" by simp - then obtain i where i_def:"Jd $h i = 1" + then obtain i where i_def:"Jd $h i = 1" unfolding diag_mat_hma_def diag_def by auto have "{# Jd $h j. j \# mset_set (UNIV - {i}) #} = {#Jd $h j. j \# mset_set UNIV - mset_set {i}#}" unfolding diag_mat_hma_def by (intro arg_cong2[where f="image_mset"] mset_set_Diff) auto also have "... = diag_mat_hma (diag Jd) - {#1#}" unfolding diag_mat_hma_def diag_def by (subst image_mset_Diff) (auto simp add:i_def) also have "... = replicate_mset (n - 1) 0" unfolding 0 by simp finally have "{# Jd $h j. j \# mset_set (UNIV - {i}) #} = replicate_mset (n - 1) 0" by simp - hence "set_mset {# Jd $h j. j \# mset_set (UNIV - {i}) #} \ {0}" + hence "set_mset {# Jd $h j. j \# mset_set (UNIV - {i}) #} \ {0}" by simp hence 1:"Jd $h j = 0" if "j \ i" for j using that by auto define u where "u = adjoint_hma U *v 1" define \ where "\ = u $h i" have "U *v u = (U ** adjoint_hma U) *v 1" unfolding u_def by (simp add:matrix_vector_mul_assoc) also have "... = 1" unfolding unitary_hmaD[OF unit_U] by simp also have "... \ 0" by simp - finally have "U *v u \ 0" by simp + finally have "U *v u \ 0" by simp hence u_nz: "u \ 0" by (cases "u = 0") auto have "diag Jd *v u = adjoint_hma U ** U ** diag Jd ** adjoint_hma U *v 1" unfolding unitary_hmaD[OF unit_U] u_def by (auto simp add:matrix_vector_mul_assoc) also have "... = adjoint_hma U ** (U ** diag Jd ** adjoint_hma U) *v 1" by (simp add:matrix_mul_assoc) also have "... = adjoint_hma U ** J *v 1" using J_sim unfolding similar_matrix_wit_def by simp also have "... = adjoint_hma U *v (map_matrix complex_of_real J *v 1)" by (simp add:map_matrix_def map_vector_def J_def matrix_vector_mul_assoc) also have "... = u" unfolding u_def markov_complex_apply[OF J_markov] by simp finally have u_ev: "diag Jd *v u = u" by simp hence "Jd * u = u" unfolding diag_vec_mult_eq by simp hence "u $h j = 0" if "j \ i" for j using 1 that unfolding times_vec_def vec_eq_iff by auto hence u_alt: "u = axis i \" unfolding \_def axis_def vec_eq_iff by auto hence \_nz: "\ \ 0" using u_nz by (cases "\=0") auto have A_sim: "similar_matrix_wit A (diag Ad) U (adjoint_hma U)" and Ad_real: "\i. Ad $h i \ \" using A_decomp unfolding real_diag_decomp_hma_def unitary_diag_def unitarily_equiv_hma_def by auto have "diag_mat_hma (diag Ad) = eigenvalues (diag Ad)" by (intro upper_tri_eigvals[symmetric] diag_imp_upper_tri A_sim) auto also have "... = eigenvalues A" using A_sim by (intro similar_matrix_eigvals[symmetric]) (auto simp add:similar_matrix_def) - finally have 3:"diag_mat_hma (diag Ad) = eigenvalues A" + finally have 3:"diag_mat_hma (diag Ad) = eigenvalues A" by simp show ?R1 unfolding 3[symmetric] diag_mat_hma_def diag_def using Ad_real by auto have "diag Ad *v u = adjoint_hma U ** U ** diag Ad ** adjoint_hma U *v 1" unfolding unitary_hmaD[OF unit_U] u_def by (auto simp add:matrix_vector_mul_assoc) also have "... = adjoint_hma U ** (U ** diag Ad ** adjoint_hma U) *v 1" by (simp add:matrix_mul_assoc) also have "... = adjoint_hma U ** A *v 1" using A_sim unfolding similar_matrix_wit_def by simp also have "... = adjoint_hma U *v (map_matrix complex_of_real A *v 1)" by (simp add:map_matrix_def map_vector_def A_def matrix_vector_mul_assoc) also have "... = u" unfolding u_def markov_complex_apply[OF markov] by simp finally have u_ev_A: "diag Ad *v u = u" by simp hence "Ad * u = u" unfolding diag_vec_mult_eq by simp hence 5:"Ad $h i = 1" - using \_nz unfolding u_alt times_vec_def vec_eq_iff axis_def by force + using \_nz unfolding u_alt times_vec_def vec_eq_iff axis_def by force thus ev_1: "(1::complex) \# eigenvalues A" unfolding 3[symmetric] diag_mat_hma_def diag_def by auto have "eigenvalues A - {#1#} = diag_mat_hma (diag Ad) - {#1#}" unfolding 3 by simp also have "... = {#Ad $h j. j \# mset_set UNIV#} - {# Ad $h i #}" unfolding 5 diag_mat_hma_def diag_def by simp also have "... = {#Ad $h j. j \# mset_set UNIV - mset_set {i}#}" by (subst image_mset_Diff) auto also have "... = {#Ad $h j. j \# mset_set (UNIV - {i})#}" - by (intro arg_cong2[where f="image_mset"] mset_set_Diff[symmetric]) auto + by (intro arg_cong2[where f="image_mset"] mset_set_Diff[symmetric]) auto finally have 4:"eigenvalues A - {#1#} = {#Ad $h j. j \# mset_set (UNIV - {i})#}" by simp have "cmod (Ad $h k) \ \\<^sub>a TYPE ('n)" if "n > 1" "k \ i" for k unfolding \\<^sub>a_def 4 using that Max_ge by auto moreover have "k = i" if "n = 1" for k - using that to_nat_less_n by simp + using that to_nat_less_n by simp ultimately have norm_Ad: "norm (Ad $h k) \ \\<^sub>a TYPE ('n) \ k = i" for k - using n_gt_0 by (cases "n = 1", auto) + using n_gt_0 by (cases "n = 1", auto) have "Re (Ad $h k) \ \\<^sub>2 TYPE ('n)" if "n > 1" "k \ i" for k unfolding \\<^sub>2_def 4 using that Max_ge by auto moreover have "k = i" if "n = 1" for k - using that to_nat_less_n by simp + using that to_nat_less_n by simp ultimately have Re_Ad: "Re (Ad $h k) \ \\<^sub>2 TYPE ('n) \ k = i" for k - using n_gt_0 by (cases "n = 1", auto) + using n_gt_0 by (cases "n = 1", auto) show \\<^sub>e_ge_0: "\\<^sub>a TYPE ('n) \ 0" proof (cases "n > 1") case True then obtain k where k_def: "k \ i" by (metis (full_types) card_n from_nat_inj n_gt_0 one_neq_zero) - have "0 \ cmod (Ad $h k)" + have "0 \ cmod (Ad $h k)" by simp - also have "... \ \\<^sub>a TYPE ('n)" + also have "... \ \\<^sub>a TYPE ('n)" using norm_Ad k_def by auto finally show ?thesis by auto next case False thus ?thesis unfolding \\<^sub>a_def by simp qed have "\v. cinner v 1 = 0 \ v \ 0 \ A *v v = \ *s v" if \_ran: "\ \# eigenvalues A - {#1#}" for \ proof - obtain j where j_def: "\ = Ad $h j" "j \ i" using \_ran unfolding 4 by auto - define v where "v = U *v axis j 1" + define v where "v = U *v axis j 1" - have "A *v v = A ** U *v axis j 1" + have "A *v v = A ** U *v axis j 1" unfolding v_def by (simp add:matrix_vector_mul_assoc) also have "... = ((U ** diag Ad ** adjoint_hma U) ** U) *v axis j 1" using A_sim unfolding similar_matrix_wit_def by simp also have "... = U ** diag Ad ** (adjoint_hma U ** U) *v axis j 1" by (simp add:matrix_mul_assoc) also have "... = U ** diag Ad *v axis j 1" using unitary_hmaD[OF unit_U] by simp also have "... = U *v (Ad * axis j 1)" by (simp add:matrix_vector_mul_assoc[symmetric] diag_vec_mult_eq) also have "... = U *v (\ *s axis j 1)" - by (intro arg_cong2[where f="(*v)"] iffD2[OF vec_eq_iff]) (auto simp:j_def axis_def) + by (intro arg_cong2[where f="(*v)"] iffD2[OF vec_eq_iff]) (auto simp:j_def axis_def) also have "... = \ *s v" - unfolding v_def by (simp add:vector_scalar_commute) + unfolding v_def by (simp add:vector_scalar_commute) finally have 5:"A *v v = \ *s v" by simp have "cinner v 1 = cinner (axis j 1) (adjoint_hma U *v 1)" unfolding v_def adjoint_def_alter_hma by simp also have "... = cinner (axis j 1) (axis i \)" unfolding u_def[symmetric] u_alt by simp - also have " ... = 0" - using j_def(2) unfolding cinner_def axis_def scalar_product_def map_vector_def + also have " ... = 0" + using j_def(2) unfolding cinner_def axis_def scalar_product_def map_vector_def by (auto simp:if_distrib if_distribR sum.If_cases) finally have 6:"cinner v 1 = 0 " by simp have "cinner v v = cinner (axis j 1) (adjoint_hma U *v (U *v (axis j 1)))" unfolding v_def adjoint_def_alter_hma by simp also have "... = cinner (axis j 1) (axis j 1)" unfolding matrix_vector_mul_assoc unitary_hmaD[OF unit_U] by simp also have "... = 1" - unfolding cinner_def axis_def scalar_product_def map_vector_def + unfolding cinner_def axis_def scalar_product_def map_vector_def by (auto simp:if_distrib if_distribR sum.If_cases) finally have "cinner v v = 1" by simp hence 7:"v \ 0" by (cases "v=0") (auto simp add:cinner_0) show ?thesis by (intro exI[where x="v"] conjI 6 7 5) qed thus "\\ \# eigenvalues A - {#1#}. \v. cinner v 1 = 0 \ v \ 0 \ A *v v = \ *s v" by simp - have "norm (A *v v) \ \\<^sub>a TYPE('n) * norm v" if "cinner v 1 = 0" for v + have "norm (A *v v) \ \\<^sub>a TYPE('n) * norm v" if "cinner v 1 = 0" for v proof - define w where "w= adjoint_hma U *v v" have "w $h i = cinner w (axis i 1)" - unfolding cinner_def axis_def scalar_product_def map_vector_def + unfolding cinner_def axis_def scalar_product_def map_vector_def by (auto simp:if_distrib if_distribR sum.If_cases) also have "... = cinner v (U *v axis i 1)" unfolding w_def adjoint_def_alter_hma by simp also have "... = cinner v ((1 / \) *s (U *v u))" unfolding vector_scalar_commute[symmetric] u_alt using \_nz by (intro_cong "[\\<^sub>2 cinner, \\<^sub>2 (*v)]") (auto simp add:axis_def vec_eq_iff) also have "... = cinner v ((1 / \) *s 1)" unfolding u_def matrix_vector_mul_assoc unitary_hmaD[OF unit_U] by simp also have "... = 0" unfolding cinner_scale_right that by simp finally have w_orth: "w $h i = 0" by simp have "norm (A *v v) = norm (U *v (diag Ad *v w))" using A_sim unfolding matrix_vector_mul_assoc similar_matrix_wit_def w_def by (simp add:matrix_mul_assoc) also have "... = norm (diag Ad *v w)" unfolding unitary_iso[OF unit_U] by simp also have "... = norm (Ad * w)" unfolding diag_vec_mult_eq by simp also have "... = sqrt (\i\UNIV. (cmod (Ad $h i) * cmod (w $h i))\<^sup>2)" unfolding norm_vec_def L2_set_def times_vec_def by (simp add:norm_mult) also have "... \ sqrt (\i\UNIV. ((\\<^sub>a TYPE('n)) * cmod (w $h i))^2)" using w_orth norm_Ad by (intro iffD2[OF real_sqrt_le_iff] sum_mono power_mono mult_right_mono') auto also have "... = \\\<^sub>a TYPE('n)\ * sqrt (\i\UNIV. (cmod (w $h i))\<^sup>2)" by (simp add:power_mult_distrib sum_distrib_left[symmetric] real_sqrt_mult) also have "... = \\\<^sub>a TYPE('n)\ * norm w" unfolding norm_vec_def L2_set_def by simp also have "... = \\<^sub>a TYPE('n) * norm w" using \\<^sub>e_ge_0 by simp also have "... = \\<^sub>a TYPE('n) * norm v" unfolding w_def unitary_iso[OF unitary_hma_adjoint[OF unit_U]] by simp finally show "norm (A *v v) \ \\<^sub>a TYPE('n) * norm v" by simp qed thus "\v. cinner v 1 = 0 \ norm (A *v v) \ \\<^sub>a TYPE('n) * norm v" by auto have "v \ (A *v v) \ \\<^sub>2 TYPE ('n) * norm v^2" if "v \ 1 = 0" for v :: "real^'n" proof - define v' where "v' = map_vector complex_of_real v" define w where "w= adjoint_hma U *v v'" have "w $h i = cinner w (axis i 1)" - unfolding cinner_def axis_def scalar_product_def map_vector_def + unfolding cinner_def axis_def scalar_product_def map_vector_def by (auto simp:if_distrib if_distribR sum.If_cases) also have "... = cinner v' (U *v axis i 1)" unfolding w_def adjoint_def_alter_hma by simp also have "... = cinner v' ((1 / \) *s (U *v u))" unfolding vector_scalar_commute[symmetric] u_alt using \_nz by (intro_cong "[\\<^sub>2 cinner, \\<^sub>2 (*v)]") (auto simp add:axis_def vec_eq_iff) also have "... = cinner v' ((1 / \) *s 1)" unfolding u_def matrix_vector_mul_assoc unitary_hmaD[OF unit_U] by simp also have "... = cnj (1 / \) * cinner v' 1" unfolding cinner_scale_right by simp also have "... = cnj (1 / \) * complex_of_real (v \ 1)" unfolding cinner_def scalar_product_def map_vector_def inner_vec_def v'_def by (intro arg_cong2[where f="(*)"] refl) (simp) also have "... = 0" unfolding that by simp finally have w_orth: "w $h i = 0" by simp - have "complex_of_real (norm v^2) = complex_of_real (v \ v)" + have "complex_of_real (norm v^2) = complex_of_real (v \ v)" by (simp add: power2_norm_eq_inner) also have "... = cinner v' v'" unfolding v'_def cinner_def scalar_product_def inner_vec_def map_vector_def by simp also have "... = norm v'^2" unfolding cinner_self by simp also have "... = norm w^2" unfolding w_def unitary_iso[OF unitary_hma_adjoint[OF unit_U]] by simp also have "... = cinner w w" unfolding cinner_self by simp also have "... = (\j\UNIV. complex_of_real (cmod (w $h j)^2))" - unfolding cinner_def scalar_product_def map_vector_def + unfolding cinner_def scalar_product_def map_vector_def cmod_power2 complex_mult_cnj[symmetric] by simp also have "... = complex_of_real (\j\UNIV. (cmod (w $h j)^2))" by simp finally have "complex_of_real (norm v^2) = complex_of_real (\j\UNIV. (cmod (w $h j)^2))" by simp hence norm_v: "norm v^2 = (\j\UNIV. (cmod (w $h j)^2))" using of_real_hom.injectivity by blast have "complex_of_real (v \ (A *v v)) = cinner v' (map_vector of_real (A *v v))" unfolding v'_def cinner_def scalar_product_def inner_vec_def map_vector_def by simp - also have "... = cinner v' (map_matrix of_real A *v v')" + also have "... = cinner v' (map_matrix of_real A *v v')" unfolding v'_def of_real_hom.mult_mat_vec_hma by simp also have "... = cinner v' (A *v v')" unfolding map_matrix_def map_vector_def A_def by auto also have "... = cinner v' (U ** diag Ad ** adjoint_hma U *v v')" using A_sim unfolding similar_matrix_wit_def by simp also have "... = cinner (adjoint_hma U *v v') (diag Ad ** adjoint_hma U *v v')" unfolding adjoint_def_alter_hma adjoint_adjoint adjoint_adjoint_id by (simp add:matrix_vector_mul_assoc matrix_mul_assoc) also have "... = cinner w (diag Ad *v w)" unfolding w_def by (simp add:matrix_vector_mul_assoc) also have "... = cinner w (Ad * w)" unfolding diag_vec_mult_eq by simp also have "... = (\j\UNIV. cnj (Ad $h j) * cmod (w $h j)^2)" unfolding cinner_def map_vector_def scalar_product_def cmod_power2 complex_mult_cnj[symmetric] by (simp add:algebra_simps) also have "... = (\j\UNIV. Ad $h j * cmod (w $h j)^2)" using Ad_real by (intro sum.cong refl arg_cong2[where f="(*)"] iffD1[OF Reals_cnj_iff]) auto also have "... = (\j\UNIV. complex_of_real (Re (Ad $h j) * cmod (w $h j)^2))" using Ad_real by (intro sum.cong refl) simp also have "... = complex_of_real (\j\ UNIV. Re (Ad $h j) * cmod (w $h j)^2)" by simp finally have "complex_of_real (v\(A *v v)) = of_real(\j\UNIV. Re (Ad $h j) * cmod (w $h j)^2)" by simp - hence "v\(A *v v) = (\j\UNIV. Re (Ad $h j) * cmod (w $h j)^2)" + hence "v\(A *v v) = (\j\UNIV. Re (Ad $h j) * cmod (w $h j)^2)" using of_real_hom.injectivity by blast also have "... \ (\j\UNIV. \\<^sub>2 TYPE ('n) * cmod (w $h j)^2)" using w_orth Re_Ad by (intro sum_mono mult_right_mono') auto also have "... = \\<^sub>2 TYPE ('n) * (\j\UNIV. cmod (w $h j)^2)" by (simp add:sum_distrib_left) also have "... = \\<^sub>2 TYPE ('n) * norm v^2" unfolding norm_v by simp finally show ?thesis by simp qed thus "\(v::real^'n). v \ 1 = 0 \ v \ (A *v v) \ \\<^sub>2 TYPE ('n) * norm v^2" by auto qed lemma find_any_real_ev: assumes "complex_of_real \ \# eigenvalues A - {#1#}" shows "\v. v \ 1 = 0 \ v \ 0 \ A *v v = \ *s v" proof - obtain w where w_def: "cinner w 1 = 0" "w \ 0" "A *v w = \ *s w" using find_any_ev assms by auto have "w = 0" if "map_vector Re (1 *s w) = 0" "map_vector Re (\ *s w) = 0" - using that by (simp add:vec_eq_iff map_vector_def complex_eq_iff) + using that by (simp add:vec_eq_iff map_vector_def complex_eq_iff) then obtain c where c_def: "map_vector Re (c *s w) \ 0" using w_def(2) by blast define u where "u = c *s w" - define v where "v = map_vector Re u" + define v where "v = map_vector Re u" hence "v \ 1 = Re (cinner u 1)" unfolding cinner_def inner_vec_def scalar_product_def map_vector_def by simp also have "... = 0" unfolding u_def cinner_scale_left w_def(1) by simp finally have 1:"v \ 1 = 0" by simp have "A *v v = (\ i. \j\UNIV. A $h i $h j * Re (u $h j))" unfolding matrix_vector_mult_def v_def map_vector_def by simp also have "... = (\ i. \j\UNIV. Re ( of_real (A $h i $h j) * u $h j))" by simp also have "... = (\ i. Re (\j\UNIV. A $h i $h j * u $h j))" unfolding A_def by simp also have "... = map_vector Re (A *v u)" unfolding map_vector_def matrix_vector_mult_def by simp also have "... = map_vector Re (of_real \ *s u)" - unfolding u_def vector_scalar_commute w_def(3) + unfolding u_def vector_scalar_commute w_def(3) by (simp add:ac_simps) also have "... = \ *s v" unfolding v_def by (simp add:vec_eq_iff map_vector_def) finally have 2: "A *v v = \ *s v" by simp - have 3:"v \ 0" + have 3:"v \ 0" unfolding v_def u_def using c_def by simp show ?thesis by (intro exI[where x="v"] conjI 1 2 3) qed lemma size_evs: "size (eigenvalues A - {#1::complex#}) = n-1" proof - have "size (eigenvalues A :: complex multiset) = n" using eigvals_poly_length card_n[symmetric] by auto thus "size (eigenvalues A - {#(1::complex)#}) = n -1" using ev_1 by (simp add: size_Diff_singleton) qed lemma find_\\<^sub>2: assumes "n > 1" shows "\\<^sub>a TYPE('n) \# image_mset cmod (eigenvalues A - {#1::complex#})" proof - have "set_mset (eigenvalues A - {#(1::complex)#}) \ {}" using assms size_evs by auto hence 2: "cmod ` set_mset (eigenvalues A - {#1#}) \ {}" by simp have "\\<^sub>a TYPE('n) \ set_mset (image_mset cmod (eigenvalues A - {#1#}))" unfolding \\<^sub>a_def using assms 2 Max_in by auto thus "\\<^sub>a TYPE('n) \# image_mset cmod (eigenvalues A - {#1#})" by simp qed lemma \\<^sub>2_real_ev: assumes "n > 1" - shows "\v. (\\. abs \=\\<^sub>a TYPE('n) \ v \ 1=0 \ v \ 0 \ A *v v = \ *s v)" + shows "\v. (\\. abs \=\\<^sub>a TYPE('n) \ v \ 1=0 \ v \ 0 \ A *v v = \ *s v)" proof - obtain \ where \_def: "cmod \ = \\<^sub>a TYPE('n)" "\ \# eigenvalues A - {#1#}" using find_\\<^sub>2[OF assms] by auto have "\ \ \" using in_diffD[OF \_def(2)] evs_real by auto - then obtain \ where \_def: "\ = of_real \" + then obtain \ where \_def: "\ = of_real \" using Reals_cases by auto have 0:"complex_of_real \ \# eigenvalues A-{#1#}" using \_def unfolding \_def by auto have 1: "\\\ = \\<^sub>a TYPE('n)" using \_def unfolding \_def by simp show ?thesis using find_any_real_ev[OF 0] 1 by auto qed -lemma \\<^sub>a_real_bound: +lemma \\<^sub>a_real_bound: fixes v :: "real^'n" assumes "v \ 1 = 0" shows "norm (A *v v) \ \\<^sub>a TYPE('n) * norm v" proof - define w where "w = map_vector complex_of_real v" have "cinner w 1 = v \ 1" - unfolding w_def cinner_def map_vector_def scalar_product_def inner_vec_def + unfolding w_def cinner_def map_vector_def scalar_product_def inner_vec_def by simp also have "... = 0" using assms by simp finally have 0: "cinner w 1 = 0" by simp have "norm (A *v v) = norm (map_matrix complex_of_real A *v (map_vector complex_of_real v))" unfolding norm_of_real of_real_hom.mult_mat_vec_hma[symmetric] by simp also have "... = norm (A *v w)" unfolding w_def A_def map_matrix_def map_vector_def by simp also have "... \ \\<^sub>a TYPE('n) * norm w" using \\<^sub>a_bound 0 by auto also have "... = \\<^sub>a TYPE('n) * norm v" unfolding w_def norm_of_real by simp finally show ?thesis by simp qed lemma \\<^sub>e_eq_\: "\\<^sub>a = \\<^sub>a TYPE('n)" proof - - have "\g_inner f (g_step f)\ \ \\<^sub>a TYPE('n) * (g_norm f)\<^sup>2" + have "\g_inner f (g_step f)\ \ \\<^sub>a TYPE('n) * (g_norm f)\<^sup>2" (is "?L \ ?R") if "g_inner f (\_. 1) = 0" for f proof - define v where "v = (\ i. f (enum_verts i))" - have 0: " v \ 1 = 0" + have 0: " v \ 1 = 0" using that unfolding g_inner_conv one_vec_def v_def by auto have "?L = \v \ (A *v v)\" unfolding g_inner_conv g_step_conv v_def by simp also have "... \ (norm v * norm (A *v v))" - by (intro Cauchy_Schwarz_ineq2) + by (intro Cauchy_Schwarz_ineq2) also have "... \ (norm v * (\\<^sub>a TYPE('n) * norm v))" by (intro mult_left_mono \\<^sub>a_real_bound 0) auto also have "... = ?R" unfolding g_norm_conv v_def by (simp add:algebra_simps power2_eq_square) finally show ?thesis by simp qed hence "\\<^sub>a \ \\<^sub>a TYPE('n)" using \\<^sub>a_ge_0 by (intro expander_intro_1) auto - moreover have "\\<^sub>a \ \\<^sub>a TYPE('n)" + moreover have "\\<^sub>a \ \\<^sub>a TYPE('n)" proof (cases "n > 1") case True then obtain v \ where v_def: "abs \ = \\<^sub>a TYPE('n)" "A *v v =\ *s v" "v \ 0" "v \ 1 = 0" using \\<^sub>2_real_ev by auto define f where "f x = v $h enum_verts_inv x" for x have v_alt: "v = (\ i. f (enum_verts i))" unfolding f_def Rep_inverse by simp have "g_inner f (\_. 1) = v \ 1" unfolding g_inner_conv v_alt one_vec_def by simp also have "... = 0" using v_def by simp finally have 2:"g_inner f (\_. 1) = 0" by simp - have "\\<^sub>a TYPE('n) * g_norm f^2 = \\<^sub>a TYPE('n) * norm v^2" + have "\\<^sub>a TYPE('n) * g_norm f^2 = \\<^sub>a TYPE('n) * norm v^2" unfolding g_norm_conv v_alt by simp - also have "... = \\<^sub>a TYPE('n) * \v \ v\" + also have "... = \\<^sub>a TYPE('n) * \v \ v\" by (simp add: power2_norm_eq_inner) - also have "... = \v \ (\ *s v)\" - unfolding v_def(1)[symmetric] scalar_mult_eq_scaleR + also have "... = \v \ (\ *s v)\" + unfolding v_def(1)[symmetric] scalar_mult_eq_scaleR by (simp add:abs_mult) - also have "... = \v \ (A *v v)\" + also have "... = \v \ (A *v v)\" unfolding v_def by simp - also have "... = \g_inner f (g_step f)\" + also have "... = \g_inner f (g_step f)\" unfolding g_inner_conv g_step_conv v_alt by simp - also have "... \ \\<^sub>a * g_norm f^2" - by (intro expansionD1 2) + also have "... \ \\<^sub>a * g_norm f^2" + by (intro expansionD1 2) finally have "\\<^sub>a TYPE('n) * g_norm f^2 \ \\<^sub>a * g_norm f^2" by simp - moreover have "norm v^2 > 0" + moreover have "norm v^2 > 0" using v_def(3) by simp - hence "g_norm f^2 > 0" + hence "g_norm f^2 > 0" unfolding g_norm_conv v_alt by simp - ultimately show ?thesis by simp + ultimately show ?thesis by simp next case False hence "n = 1" using n_gt_0 by simp hence "\\<^sub>a TYPE('n) = 0" unfolding \\<^sub>a_def by simp then show ?thesis using \_ge_0 by simp qed ultimately show ?thesis by simp qed -lemma \\<^sub>2_ev: +lemma \\<^sub>2_ev: assumes "n > 1" shows "\v. v \ 1 = 0 \ v \ 0 \ A *v v = \\<^sub>2 TYPE('n) *s v" proof - have "set_mset (eigenvalues A - {#1::complex#}) \ {}" using size_evs assms by auto hence "Max (Re ` set_mset (eigenvalues A - {#1#})) \ Re ` set_mset (eigenvalues A - {#1#})" by (intro Max_in) auto - hence "\\<^sub>2 TYPE ('n) \ Re ` set_mset (eigenvalues A - {#1#})" + hence "\\<^sub>2 TYPE ('n) \ Re ` set_mset (eigenvalues A - {#1#})" unfolding \\<^sub>2_def using assms by simp then obtain \ where \_def: "\ \ set_mset (eigenvalues A - {#1#})" "\\<^sub>2 TYPE ('n) = Re \" by auto have \_real: "\ \ \" using evs_real in_diffD[OF \_def(1)] by auto have "complex_of_real (\\<^sub>2 TYPE ('n)) = of_real (Re \)" unfolding \_def by simp also have "... = \" using \_real by simp also have "... \# eigenvalues A - {#1#}" using \_def(1) by simp finally have 0:"complex_of_real (\\<^sub>2 TYPE ('n)) \# eigenvalues A - {#1#}" by simp thus ?thesis using find_any_real_ev[OF 0] by auto qed lemma \\<^sub>2_eq_\\<^sub>2: "\\<^sub>2 = \\<^sub>2 TYPE ('n)" proof (cases "n > 1") case True - obtain v where v_def: "v \ 1 = 0" "v \ 0" "A *v v = \\<^sub>2 TYPE('n) *s v" + obtain v where v_def: "v \ 1 = 0" "v \ 0" "A *v v = \\<^sub>2 TYPE('n) *s v" using \\<^sub>2_ev[OF True] by auto define f where "f x = v $h enum_verts_inv x" for x have v_alt: "v = (\ i. f (enum_verts i))" unfolding f_def Rep_inverse by simp have "g_inner f (\_. 1) = v \ 1" unfolding g_inner_conv v_alt one_vec_def by simp also have "... = 0" unfolding v_def(1) by simp finally have f_orth: "g_inner f (\_. 1) = 0" by simp have "\\<^sub>2 TYPE('n) * norm v^2= v \ (\\<^sub>2 TYPE('n) *s v)" unfolding power2_norm_eq_inner by (simp add:algebra_simps scalar_mult_eq_scaleR) also have "... = v \ (A *v v)" unfolding v_def by simp also have "... = g_inner f (g_step f)" unfolding v_alt g_inner_conv g_step_conv by simp also have "... \ \\<^sub>2 * g_norm f^2" - by (intro os_expanderD f_orth) + by (intro os_expanderD f_orth) also have "... = \\<^sub>2 * norm v^2" unfolding v_alt g_norm_conv by simp finally have "\\<^sub>2 TYPE('n) * norm v^2 \ \\<^sub>2 * norm v^2" by simp hence "\\<^sub>2 TYPE('n) \ \\<^sub>2" using v_def(2) by simp moreover have "\\<^sub>2 \ \\<^sub>2 TYPE ('n)" - using \\<^sub>2_bound - by (intro os_expanderI[OF True]) + using \\<^sub>2_bound + by (intro os_expanderI[OF True]) (simp add: g_inner_conv g_step_conv g_norm_conv one_vec_def) ultimately show ?thesis by simp next case False - then show ?thesis + then show ?thesis unfolding \\<^sub>2_def \\<^sub>2_def by simp qed lemma expansionD2: assumes "g_inner f (\_. 1) = 0" shows "g_norm (g_step f) \ \\<^sub>a * g_norm f" (is "?L \ ?R") proof - define v where "v = (\ i. f (enum_verts i))" have "v \ 1 = g_inner f (\_. 1)" unfolding g_inner_conv v_def one_vec_def by simp also have "... = 0" using assms by simp finally have 0:"v \ 1 = 0" by simp have "g_norm (g_step f) = norm (A *v v)" unfolding g_norm_conv g_step_conv v_def by auto also have "... \ \\<^sub>a * norm v" unfolding \\<^sub>e_eq_\ by (intro \\<^sub>a_real_bound 0) also have "... = \\<^sub>a * g_norm f" unfolding g_norm_conv v_def by simp finally show ?thesis by simp qed -lemma rayleigh_bound: +lemma rayleigh_bound: fixes v :: "real^'n" shows "\v \ (A *v v)\ \ norm v^2" proof - define f where "f x = v $h enum_verts_inv x" for x have v_alt: "v = (\ i. f (enum_verts i))" unfolding f_def Rep_inverse by simp have "\v \ (A *v v)\ = \g_inner f (g_step f)\" unfolding v_alt g_inner_conv g_step_conv by simp also have "... = \(\a\arcs G. f (head G a) * f (tail G a))\/d" unfolding g_inner_step_eq by simp also have "... \ (d * (g_norm f)\<^sup>2) / d" by (intro divide_right_mono bdd_above_aux) auto - also have "... = g_norm f^2" + also have "... = g_norm f^2" using d_gt_0 by simp also have "... = norm v^2" unfolding g_norm_conv v_alt by simp finally show ?thesis by simp qed text \The following implies that two-sided expanders are also one-sided expanders.\ lemma \\<^sub>2_range: "\\\<^sub>2\ \ \\<^sub>a" proof (cases "n > 1") case True hence 0:"set_mset (eigenvalues A - {#1::complex#}) \ {}" using size_evs by auto have "\\<^sub>2 TYPE ('n) = Max (Re ` set_mset (eigenvalues A - {#1::complex#}))" unfolding \\<^sub>2_def using True by simp also have "... \ Re ` set_mset (eigenvalues A - {#1::complex#})" using Max_in 0 by simp finally have "\\<^sub>2 TYPE ('n) \ Re ` set_mset (eigenvalues A - {#1::complex#})" by simp then obtain \ where \_def: "\ \ set_mset (eigenvalues A - {#1::complex#})" "\\<^sub>2 TYPE ('n) = Re \" by auto have "\\\<^sub>2\ = \\\<^sub>2 TYPE ('n) \" using \\<^sub>2_eq_\\<^sub>2 by simp - also have "... = \Re \\" + also have "... = \Re \\" using \_def by simp also have "... \ cmod \" using abs_Re_le_cmod by simp also have "... \ Max (cmod ` set_mset (eigenvalues A - {#1#}))" using \_def(1) by (intro Max_ge) auto - also have "... \ \\<^sub>a TYPE('n)" + also have "... \ \\<^sub>a TYPE('n)" unfolding \\<^sub>a_def using True by simp also have "... = \\<^sub>a" using \\<^sub>e_eq_\ by simp finally show ?thesis by simp -next +next case False - thus ?thesis + thus ?thesis unfolding \\<^sub>2_def \\<^sub>a_def by simp qed end -lemmas (in regular_graph) expansionD2 = +lemmas (in regular_graph) expansionD2 = regular_graph_tts.expansionD2[OF eg_tts_1, - internalize_sort "'n :: finite", OF _ regular_graph_axioms, + internalize_sort "'n :: finite", OF _ regular_graph_axioms, unfolded remove_finite_premise, cancel_type_definition, OF verts_non_empty] -lemmas (in regular_graph) \\<^sub>2_range = +lemmas (in regular_graph) \\<^sub>2_range = regular_graph_tts.\\<^sub>2_range[OF eg_tts_1, - internalize_sort "'n :: finite", OF _ regular_graph_axioms, + internalize_sort "'n :: finite", OF _ regular_graph_axioms, unfolded remove_finite_premise, cancel_type_definition, OF verts_non_empty] unbundle no_intro_cong_syntax end \ No newline at end of file diff --git a/thys/Expander_Graphs/Expander_Graphs_MGG.thy b/thys/Expander_Graphs/Expander_Graphs_MGG.thy --- a/thys/Expander_Graphs/Expander_Graphs_MGG.thy +++ b/thys/Expander_Graphs/Expander_Graphs_MGG.thy @@ -1,1260 +1,1260 @@ section \Margulis Gabber Galil Construction\label{sec:margulis}\ text \This section formalizes the Margulis-Gabber-Galil expander graph, which is defined on the product space $\mathbb Z_n \times \mathbb Z_n$. The construction is an adaptation of graph introduced by Margulis~\cite{margulis1973}, for which he gave a non-constructive proof of its spectral gap. Later Gabber and Galil~\cite{gabber1981} adapted the graph and derived an explicit spectral gap, i.e., that the second largest eigenvalue is bounded by $\frac{5}{8} \sqrt{2}$. The -proof was later improved by Jimbo and Marouka~\cite{jimbo1987} using Fourier Analysis. +proof was later improved by Jimbo and Marouka~\cite{jimbo1987} using Fourier Analysis. Hoory et al.~\cite[\S 8]{hoory2006} present a slight simplification of that proof (due to Boppala) which this formalization is based on.\ theory Expander_Graphs_MGG - imports + imports "HOL-Analysis.Complex_Transcendental" "HOL-Decision_Procs.Approximation" Expander_Graphs_Definition begin -datatype ('a, 'b) arc = Arc (arc_tail: 'a) (arc_head: 'a) (arc_label: 'b) +datatype ('a, 'b) arc = Arc (arc_tail: 'a) (arc_head: 'a) (arc_label: 'b) fun mgg_graph_step :: "nat \ (int \ int) \ (nat \ int) \ (int \ int)" - where "mgg_graph_step n (i,j) (l,\) = + where "mgg_graph_step n (i,j) (l,\) = [ ((i+\*(2*j+0)) mod int n, j), (i, (j+\*(2*i+0)) mod int n) , ((i+\*(2*j+1)) mod int n, j), (i, (j+\*(2*i+1)) mod int n) ] ! l" -definition mgg_graph :: "nat \ (int \ int, (int \ int, nat \ int) arc) pre_digraph" where +definition mgg_graph :: "nat \ (int \ int, (int \ int, nat \ int) arc) pre_digraph" where "mgg_graph n = - \ verts = {0.. {0..(t,l). (Arc t (mgg_graph_step n t l) l))`(({0..{0..({..<4}\{-1,1})), + \ verts = {0.. {0..(t,l). (Arc t (mgg_graph_step n t l) l))`(({0..{0..({..<4}\{-1,1})), tail = arc_tail, head = arc_head \" locale margulis_gaber_galil = fixes m :: nat assumes m_gt_0: "m > 0" begin abbreviation G where "G \ mgg_graph m" lemma wf_digraph: "wf_digraph (mgg_graph m)" proof - - have - "tail (mgg_graph m) e \ verts (mgg_graph m)" (is "?A") + have + "tail (mgg_graph m) e \ verts (mgg_graph m)" (is "?A") "head (mgg_graph m) e \ verts (mgg_graph m)" (is "?B") if a:"e \ arcs (mgg_graph m)" for e proof - - obtain t l \ where tl_def: - "t \ {0.. {0.. {..<4}" "\ \ {-1,1}" + obtain t l \ where tl_def: + "t \ {0.. {0.. {..<4}" "\ \ {-1,1}" "e = Arc t (mgg_graph_step m t (l,\)) (l,\)" - using a mgg_graph_def by auto + using a mgg_graph_def by auto thus ?A unfolding mgg_graph_def by auto - have "mgg_graph_step m (fst t, snd t) (l,\) \ {0.. {0..) \ {0.. {0.. {0.. {0.. {0.. {0.. int, nat \ int) arc set" +definition arcs_pos :: "(int \ int, nat \ int) arc set" where "arcs_pos = (\(t,l). (Arc t (mgg_graph_step m t (l,1)) (l, 1)))`(verts G\{..<4})" -definition arcs_neg :: "(int \ int, nat \ int) arc set" +definition arcs_neg :: "(int \ int, nat \ int) arc set" where "arcs_neg = (\(h,l). (Arc (mgg_graph_step m h (l,1)) h (l,-1)))`(verts G\{..<4})" lemma arcs_sym: "arcs G = arcs_pos \ arcs_neg" proof - - have 0: "x \ arcs G" if "x \ arcs_pos" for x - using that unfolding arcs_pos_def mgg_graph_def by auto - have 1: "a \ arcs G" if t:"a \ arcs_neg" for a + have 0: "x \ arcs G" if "x \ arcs_pos" for x + using that unfolding arcs_pos_def mgg_graph_def by auto + have 1: "a \ arcs G" if t:"a \ arcs_neg" for a proof - obtain h l where hl_def: "h \ verts G" "l \ {..<4}" "a = Arc (mgg_graph_step m h (l,1)) h (l,-1)" using t unfolding arcs_neg_def by auto define t where "t = mgg_graph_step m h (l,1)" - have h_ran: "h \ {0.. {0.. {0.. {0.. set [0,1,2,3]" + have l_ran: "l \ set [0,1,2,3]" using hl_def(2) by auto - have "t \ {0.. {0.. {0.. {0.. verts G" + hence t_ran: "t \ verts G" unfolding mgg_graph_def by simp - have "h = mgg_graph_step m t (l,-1)" + have "h = mgg_graph_step m t (l,-1)" using h_ran l_ran unfolding t_def by (cases h, auto simp add:mod_simps) hence "a = Arc t (mgg_graph_step m t (l,-1)) (l,-1)" unfolding t_def hl_def(3) by simp - thus ?thesis + thus ?thesis using t_ran hl_def(2) mgg_graph_def by (simp add:image_iff) qed have "card (arcs_pos \ arcs_neg) = card arcs_pos + card arcs_neg" unfolding arcs_pos_def arcs_neg_def by (intro card_Un_disjoint finite_imageI) auto also have "... = card (verts G\{..<4::nat}) + card (verts G\{..<4::nat})" unfolding arcs_pos_def arcs_neg_def by (intro arg_cong2[where f="(+)"] card_image inj_onI) auto also have "... = card (verts G\{..<4::nat}\{-1,1::int})" by simp also have "... = card ((\(t, l). Arc t (mgg_graph_step m t l) l) ` (verts G \{..<4}\{-1,1}))" by (intro card_image[symmetric] inj_onI) auto also have "... = card (arcs G)" unfolding mgg_graph_def by simp - finally have "card (arcs_pos \ arcs_neg) = card (arcs G)" + finally have "card (arcs_pos \ arcs_neg) = card (arcs G)" by simp hence "arcs_pos \ arcs_neg = arcs G" - using 0 1 by (intro card_subset_eq, auto) + using 0 1 by (intro card_subset_eq, auto) thus ?thesis by simp qed lemma sym: "symmetric_multi_graph (mgg_graph m)" proof - - define f :: "(int \ int, nat \ int) arc \ (int \ int, nat \ int) arc" - where "f a = Arc (arc_head a) (arc_tail a) (apsnd (\x. (-1) * x) (arc_label a))" for a + define f :: "(int \ int, nat \ int) arc \ (int \ int, nat \ int) arc" + where "f a = Arc (arc_head a) (arc_tail a) (apsnd (\x. (-1) * x) (arc_label a))" for a have a: "bij_betw f arcs_pos arcs_neg" by (intro bij_betwI[where g="f"]) (auto simp add:f_def image_iff arcs_pos_def arcs_neg_def) have b: "bij_betw f arcs_neg arcs_pos" by (intro bij_betwI[where g="f"]) (auto simp add:f_def image_iff arcs_pos_def arcs_neg_def) have c:"bij_betw f (arcs_pos \ arcs_neg) (arcs_neg \ arcs_pos)" by (intro bij_betw_combine[OF a b]) (auto simp add:arcs_pos_def arcs_neg_def) hence c:"bij_betw f (arcs G) (arcs G)" unfolding arcs_sym by (subst (2) sup_commute, simp) show ?thesis - by (intro symmetric_multi_graphI[where f="f"] fin_digraph_axioms c) + by (intro symmetric_multi_graphI[where f="f"] fin_digraph_axioms c) (simp add:f_def mgg_graph_def) qed lemma out_deg: assumes "v \ verts G" shows "out_degree G v = 8" proof - have "out_degree (mgg_graph m) v = card (out_arcs (mgg_graph m) v)" unfolding out_degree_def by simp - also have "... = card {e. (\w \ verts (mgg_graph m). \l \ {..<4} \ {-1,1}. - e = Arc w (mgg_graph_step m w l) l \ arc_tail e = v)}" + also have "... = card {e. (\w \ verts (mgg_graph m). \l \ {..<4} \ {-1,1}. + e = Arc w (mgg_graph_step m w l) l \ arc_tail e = v)}" unfolding mgg_graph_def out_arcs_def by (simp add:image_iff) also have "... = card {e. (\l \ {..<4} \ {-1,1}. e = Arc v (mgg_graph_step m v l) l)}" using assms by (intro arg_cong[where f="card"] iffD2[OF set_eq_iff] allI) auto also have "... = card ((\l. Arc v (mgg_graph_step m v l) l) ` ({..<4} \ {-1,1}))" by (intro arg_cong[where f="card"]) (auto simp add:image_iff) also have "... = card ({..<4::nat} \ {-1,1::int})" by (intro card_image inj_onI) simp also have "... = 8" by simp finally show ?thesis by simp qed lemma verts_ne: - "verts G \ {}" + "verts G \ {}" using m_gt_0 unfolding mgg_graph_def by simp sublocale regular_graph "mgg_graph m" using out_deg verts_ne by (intro regular_graphI[where d="8"] sym) auto lemma d_eq_8: "d = 8" proof - obtain v where v_def: "v \ verts G" using verts_ne by auto hence 0:"(SOME v. v \ verts G) \ verts G" by (rule someI[where x="v"]) show ?thesis using out_deg[OF 0] unfolding d_def by simp qed text \We start by introducing Fourier Analysis on the torus $\mathbb Z_n \times \mathbb Z_n$. The following is too specialized for a general AFP entry.\ lemma g_inner_sum_left: assumes "finite I" shows "g_inner (\x. (\i \ I. f i x)) g = (\i\ I. g_inner (f i) g)" using assms by (induction I rule:finite_induct) (auto simp add:g_inner_simps) lemma g_inner_sum_right: assumes "finite I" shows "g_inner f (\x. (\i \ I. g i x)) = (\i\ I. g_inner f (g i))" using assms by (induction I rule:finite_induct) (auto simp add:g_inner_simps) lemma g_inner_reindex: assumes "bij_betw h (verts G) (verts G)" shows "g_inner f g = g_inner (\x. (f (h x))) (\x. (g (h x)))" unfolding g_inner_def by (subst sum.reindex_bij_betw[OF assms,symmetric]) simp - + definition \\<^sub>F :: "real \ complex" where "\\<^sub>F x = cis (2*pi*x/m)" lemma \\<^sub>F_simps: "\\<^sub>F (x + y) = \\<^sub>F x * \\<^sub>F y" "\\<^sub>F (x - y) = \\<^sub>F x * \\<^sub>F (-y)" "cnj (\\<^sub>F x) = \\<^sub>F (-x)" - unfolding \\<^sub>F_def by (auto simp add:algebra_simps diff_divide_distrib + unfolding \\<^sub>F_def by (auto simp add:algebra_simps diff_divide_distrib add_divide_distrib cis_mult cis_divide cis_cnj) lemma \\<^sub>F_cong: fixes x y :: int - assumes "x mod m = y mod m" + assumes "x mod m = y mod m" shows "\\<^sub>F (of_int x) = \\<^sub>F (of_int y)" proof - obtain z :: int where "y = x + m*z" using mod_eqE[OF assms] by auto hence "\\<^sub>F (of_int y) = \\<^sub>F (of_int x + of_int (m*z))" by simp also have "... = \\<^sub>F (of_int x) * \\<^sub>F (of_int (m*z))" by (simp add:\\<^sub>F_simps) also have "... = \\<^sub>F (of_int x) * cis (2 * pi * of_int (z))" - unfolding \\<^sub>F_def using m_gt_0 + unfolding \\<^sub>F_def using m_gt_0 by (intro arg_cong2[where f="(*)"] arg_cong[where f="cis"]) auto also have "... = \\<^sub>F (of_int x) * 1" by (intro arg_cong2[where f="(*)"] cis_multiple_2pi) auto finally show ?thesis by simp qed lemma cis_eq_1_imp: assumes "cis (2 * pi * x) = 1" shows "x \ \" proof - - have "cos (2 * pi * x) = Re (cis (2*pi*x))" + have "cos (2 * pi * x) = Re (cis (2*pi*x))" using cis.simps by simp - also have "... = 1" + also have "... = 1" unfolding assms by simp finally have "cos (2 * pi * x) = 1" by simp then obtain y where "2 * pi * x = of_int y * 2 * pi" using cos_one_2pi_int by auto hence "y = x" by simp thus ?thesis by auto qed lemma \\<^sub>F_eq_1_iff: fixes x :: int shows "\\<^sub>F x = 1 \ x mod m = 0" proof assume "\\<^sub>F (real_of_int x) = 1" hence "cis (2 * pi * real_of_int x / real m) = 1" unfolding \\<^sub>F_def by simp hence "real_of_int x / real m \ \" using cis_eq_1_imp by simp then obtain z :: int where "of_int x / real m = z" using Ints_cases by auto hence "x = z * real m" using m_gt_0 by (simp add: nonzero_divide_eq_eq) hence "x = z * m" using of_int_eq_iff by fastforce thus "x mod m = 0" by simp next assume "x mod m = 0" hence "\\<^sub>F x = \\<^sub>F (of_int 0)" by (intro \\<^sub>F_cong) auto also have "... = 1" unfolding \\<^sub>F_def by simp finally show "\\<^sub>F x= 1" by simp qed definition FT :: "(int \ int \ complex) \ (int \ int \ complex)" where "FT f v = g_inner f (\x. \\<^sub>F (fst x * fst v + snd x * snd v))" lemma FT_altdef: "FT f (u,v) = g_inner f (\x. \\<^sub>F (fst x * u + snd x * v))" unfolding FT_def by (simp add:case_prod_beta) lemma FT_add: "FT (\x. f x + g x) v = FT f v + FT g v" unfolding FT_def by (simp add:g_inner_simps algebra_simps) lemma FT_zero: "FT (\x. 0) v = 0" unfolding FT_def g_inner_def by simp -lemma FT_sum: - assumes "finite I" +lemma FT_sum: + assumes "finite I" shows "FT (\x. (\i \ I. f i x)) v = (\i \ I. FT (f i) v)" using assms by (induction rule: finite_induct, auto simp add:FT_zero FT_add) lemma FT_scale: "FT (\x. c * f x) v = c * FT f v" unfolding FT_def by (simp add: g_inner_simps) lemma FT_cong: assumes "\x. x \ verts G \ f x = g x" shows "FT f = FT g" unfolding FT_def by (intro ext g_inner_cong assms refl) lemma parseval: "g_inner f g = g_inner (FT f) (FT g)/m^2" (is "?L = ?R") proof - define \ :: "(int \ int) \ (int \ int) \ complex" where "\ x y = of_bool (x = y)" for x y have FT_\: "FT (\ v) x = \\<^sub>F (-(fst v *fst x +snd v * snd x))" if "v \ verts G" for v x using that by (simp add:FT_def g_inner_def \_def \\<^sub>F_simps) have 1: "(\x=0..\<^sub>F (z*x)) = m * of_bool(z mod m = 0)" (is "?L1 = ?R1") for z :: int proof (cases "z mod m = 0") case True have "(\x=0..\<^sub>F (z*x)) = (\x=0..\<^sub>F (of_int 0))" using True by (intro sum.cong \\<^sub>F_cong refl) auto also have "... = m * of_bool(z mod m = 0)" unfolding \\<^sub>F_def True by simp finally show ?thesis by simp next case False have "(1-\\<^sub>F z) * ?L1 = (1-\\<^sub>F z) * (\x \ int ` {..\<^sub>F(z*x))" by (intro arg_cong2[where f="(*)"] sum.cong refl) (simp add: image_atLeastZeroLessThan_int) also have "... = (\x\<^sub>F(z*real x) - \\<^sub>F(z*(real (Suc x))))" by (subst sum.reindex, auto simp add:algebra_simps sum_distrib_left \\<^sub>F_simps) also have "... = \\<^sub>F (z * 0) - \\<^sub>F (z * m)" by (subst sum_lessThan_telescope') simp also have "... = \\<^sub>F (of_int 0) - \\<^sub>F (of_int 0)" by (intro arg_cong2[where f="(-)"] \\<^sub>F_cong) auto also have "... = 0" by simp finally have "(1- \\<^sub>F z) * ?L1 = 0" by simp moreover have "\\<^sub>F z \ 1" using \\<^sub>F_eq_1_iff False by simp hence "(1- \\<^sub>F z) \ 0" by simp ultimately have "?L1 = 0" by simp then show ?thesis using False by simp qed have 0:"g_inner (\ v) (\ w) = g_inner (FT (\ v)) (FT (\ w))/m^2" (is "?L1 = ?R1/_") if "v \ verts G" "w \ verts G" for v w proof - have "?R1=g_inner(\x. \\<^sub>F(-(fst v *fst x +snd v * snd x)))(\x. \\<^sub>F(-(fst w *fst x +snd w * snd x)))" using that by (intro g_inner_cong, auto simp add:FT_\) also have "...=(\(x,y)\{0..{0..\<^sub>F((fst w-fst v)*x)*\\<^sub>F((snd w - snd v)* y))" unfolding g_inner_def by (simp add:\\<^sub>F_simps algebra_simps case_prod_beta mgg_graph_def) also have "...=(\x=0..y = 0..\<^sub>F((fst w - fst v)*x)*\\<^sub>F((snd w - snd v) * y))" by (subst sum.cartesian_product[symmetric]) simp also have "...=(\x=0..\<^sub>F((fst w - fst v)*x))*(\y = 0..\<^sub>F((snd w - snd v) * y))" by (subst sum.swap) (simp add:sum_distrib_left sum_distrib_right) - also have "... = of_nat (m * of_bool(fst v mod m = fst w mod m)) * + also have "... = of_nat (m * of_bool(fst v mod m = fst w mod m)) * of_nat (m * of_bool(snd v mod m = snd w mod m))" - using m_gt_0 unfolding 1 - by (intro arg_cong2[where f="(*)"] arg_cong[where f="of_bool"] + using m_gt_0 unfolding 1 + by (intro arg_cong2[where f="(*)"] arg_cong[where f="of_bool"] arg_cong[where f="of_nat"] refl) (auto simp add:algebra_simps cong:mod_diff_cong) also have "... = m^2 * of_bool(v = w)" using that by (auto simp add:prod_eq_iff mgg_graph_def power2_eq_square) - also have "... = m^2 * ?L1" + also have "... = m^2 * ?L1" using that unfolding g_inner_def \_def by simp finally have "?R1 = m^2 * ?L1" by simp thus ?thesis using m_gt_0 by simp qed have "?L = g_inner (\x. (\v \ verts G. (f v) * \ v x)) (\x. (\v \ verts G. (g v) * \ v x))" unfolding \_def by (intro g_inner_cong) auto also have "... = (\v\verts G. (f v) * (\w\verts G. cnj (g w) * g_inner (\ v) (\ w)))" by (simp add:g_inner_simps g_inner_sum_left g_inner_sum_right) also have "... = (\v\verts G. (f v) * (\w\verts G. cnj (g w) * g_inner(FT (\ v)) (FT (\ w))))/m^2" by (simp add:0 sum_divide_distrib sum_distrib_left algebra_simps) also have "...=g_inner(\x.(\v\verts G. (f v)*FT (\ v) x))(\x.(\v\verts G. (g v)*FT (\ v) x))/m\<^sup>2" by (simp add:g_inner_simps g_inner_sum_left g_inner_sum_right) also have "...=g_inner(FT(\x.(\v\verts G.(f v)*\ v x)))(FT(\x.(\v\verts G.(g v)*\ v x)))/m\<^sup>2" by (intro g_inner_cong arg_cong2[where f="(/)"]) (simp_all add: FT_sum FT_scale) also have "... = g_inner (FT f) (FT g)/m^2 " unfolding \_def comp_def by (intro g_inner_cong arg_cong2[where f="(/)"] fun_cong[OF FT_cong]) auto finally show ?thesis by simp qed lemma plancharel: "(\v \ verts G. norm (f v)^2) = (\v \ verts G. norm (FT f v)^2)/m^2" (is "?L = ?R") proof - have "complex_of_real ?L = g_inner f f" by (simp flip:of_real_power add:complex_norm_square g_inner_def algebra_simps) also have "... = g_inner (FT f) (FT f) / m^2" by (subst parseval) simp also have "... = complex_of_real ?R" by (simp flip:of_real_power add:complex_norm_square g_inner_def algebra_simps) simp finally have "complex_of_real ?L = complex_of_real ?R" by simp - thus ?thesis + thus ?thesis using of_real_eq_iff by blast qed lemma FT_swap: "FT (\x. f (snd x, fst x)) (u,v) = FT f (v,u)" proof - have 0:"bij_betw (\(x::int \ int). (snd x, fst x)) (verts G) (verts G)" by (intro bij_betwI[where g="(\(x::int \ int). (snd x, fst x))"]) (auto simp add:mgg_graph_def) show ?thesis unfolding FT_def by (subst g_inner_reindex[OF 0]) (simp add:algebra_simps) qed lemma mod_add_mult_eq: fixes a x y :: int shows "(a + x * (y mod m)) mod m = (a+x*y) mod m" using mod_add_cong mod_mult_right_eq by blast definition periodic where "periodic f = (\x y. f (x,y) = f (x mod int m, y mod int m))" lemma periodicD: assumes "periodic f" shows "f (x,y) = f (x mod m, y mod m)" using assms unfolding periodic_def by simp lemma periodic_comp: assumes "periodic f" shows "periodic (\x. g (f x))" using assms unfolding periodic_def by simp lemma periodic_cong: fixes x y u v :: int assumes "periodic f" - assumes "x mod m = u mod m" "y mod m = v mod m" + assumes "x mod m = u mod m" "y mod m = v mod m" shows "f (x,y) = f (u, v)" using periodicD[OF assms(1)] assms(2,3) by metis lemma periodic_FT: "periodic (FT f)" proof - have "FT f (x,y) = FT f (x mod m,y mod m)" for x y - unfolding FT_altdef by (intro g_inner_cong \\<^sub>F_cong ext) + unfolding FT_altdef by (intro g_inner_cong \\<^sub>F_cong ext) (auto simp add:mod_simps cong:mod_add_cong) - thus ?thesis + thus ?thesis unfolding periodic_def by simp qed lemma FT_sheer_aux: fixes u v c d :: int - assumes "periodic f" - shows "FT (\x. f (fst x,snd x+c*fst x+d)) (u,v) = \\<^sub>F (d* v) * FT f (u-c* v,v)" + assumes "periodic f" + shows "FT (\x. f (fst x,snd x+c*fst x+d)) (u,v) = \\<^sub>F (d* v) * FT f (u-c* v,v)" (is "?L = ?R") proof - define s where "s = (\(x,y). (x, (y - c * x-d) mod m))" define s0 where "s0 = (\(x,y). (x, (y-c*x) mod m))" define s1 where "s1 = (\(x::int,y). (x, (y-d) mod m))" have 0:"bij_betw s0 (verts G) (verts G)" by (intro bij_betwI[where g="\(x,y). (x,(y+c*x) mod m)"]) (auto simp add:mgg_graph_def s0_def Pi_def mod_simps) have 1:"bij_betw s1 (verts G) (verts G)" by (intro bij_betwI[where g="\(x,y). (x,(y+d) mod m)"]) (auto simp add:mgg_graph_def s1_def Pi_def mod_simps) have 2: "s = (s1 \ s0)" by (simp add:s1_def s0_def s_def comp_def mod_simps case_prod_beta ext) have 3:"bij_betw s (verts G) (verts G)" unfolding 2 using bij_betw_trans[OF 0 1] by simp have 4:"(snd (s x) + c * fst x + d) mod int m = snd x mod m" for x unfolding s_def by (simp add:case_prod_beta cong:mod_add_cong) (simp add:algebra_simps) have 5: "fst (s x) = fst x" for x unfolding s_def by (cases x, simp) have "?L = g_inner (\x. f (fst x, snd x + c*fst x+d)) (\x. \\<^sub>F (fst x*u + snd x* v))" unfolding FT_altdef by simp also have "... = g_inner (\x. f (fst x, (snd x + c*fst x+d) mod m)) (\x. \\<^sub>F (fst x*u + snd x* v))" by (intro g_inner_cong periodic_cong[OF assms]) (auto simp add:algebra_simps) also have "... = g_inner (\x. f (fst x, snd x mod m)) (\x. \\<^sub>F (fst x*u+ snd (s x)* v))" by (subst g_inner_reindex[OF 3]) (simp add:4 5) also have "... = g_inner (\x. f (fst x, snd x mod m)) (\x. \\<^sub>F (fst x*u+ ((snd x-c*fst x-d) mod m)* v))" by (simp add:s_def case_prod_beta) also have "... = g_inner f (\x. \\<^sub>F (fst x* (u-c * v) + snd x * v-d * v))" - by (intro g_inner_cong \\<^sub>F_cong) (auto simp add:mgg_graph_def algebra_simps mod_add_mult_eq) + by (intro g_inner_cong \\<^sub>F_cong) (auto simp add:mgg_graph_def algebra_simps mod_add_mult_eq) also have "... = g_inner f (\x. \\<^sub>F (-d* v)*\\<^sub>F (fst x*(u-c* v) + snd x * v))" by (simp add: \\<^sub>F_simps algebra_simps) also have "... = \\<^sub>F (d* v)*g_inner f (\x. \\<^sub>F (fst x*(u-c* v) + snd x * v))" by (simp add:g_inner_simps \\<^sub>F_simps) also have "... = ?R" - unfolding FT_altdef by simp + unfolding FT_altdef by simp finally show ?thesis by simp qed lemma FT_sheer: fixes u v c d :: int - assumes "periodic f" - shows + assumes "periodic f" + shows "FT (\x. f (fst x,snd x+c*fst x+d)) (u,v) = \\<^sub>F (d* v) * FT f (u-c* v,v)" (is "?A") "FT (\x. f (fst x,snd x+c*fst x)) (u,v) = FT f (u-c* v,v)" (is "?B") "FT (\x. f (fst x+c* snd x+d,snd x)) (u,v) = \\<^sub>F (d* u) * FT f (u,v-c*u)" (is "?C") "FT (\x. f (fst x+c* snd x,snd x)) (u,v) = FT f (u,v-c*u)" (is "?D") proof - - have 1: "periodic (\x. f (snd x, fst x))" + have 1: "periodic (\x. f (snd x, fst x))" using assms unfolding periodic_def by simp have 0:"\\<^sub>F 0 = 1" unfolding \\<^sub>F_def by simp show ?A using FT_sheer_aux[OF assms] by simp show ?B using 0 FT_sheer_aux[OF assms, where d="0"] by simp show ?C using FT_sheer_aux[OF 1] by (subst (1 2) FT_swap[symmetric], simp) show ?D using 0 FT_sheer_aux[OF 1, where d="0"] by (subst (1 2) FT_swap[symmetric], simp) qed definition T\<^sub>1 :: "int \ int \ int \ int" where "T\<^sub>1 x = ((fst x + 2 * snd x) mod m, snd x)" definition S\<^sub>1 :: "int \ int \ int \ int" where "S\<^sub>1 x = ((fst x - 2 * snd x) mod m, snd x)" definition T\<^sub>2 :: "int \ int \ int \ int" where "T\<^sub>2 x = (fst x, (snd x + 2 * fst x) mod m)" definition S\<^sub>2 :: "int \ int \ int \ int" where "S\<^sub>2 x = (fst x, (snd x - 2 * fst x) mod m)" -definition \_aux :: "int \ int \ real \ real" +definition \_aux :: "int \ int \ real \ real" where "\_aux x = (\fst x/m-1/2\,\snd x/m-1/2\)" definition compare :: "real \ real \ real \ real \ bool" where "compare x y = (fst x \ fst y \ snd x \ snd y \ x \ y)" -text \The value here is different from the value in the source material. This is because the proof +text \The value here is different from the value in the source material. This is because the proof in Hoory~\cite[\S 8]{hoory2006} only establishes the bound $\frac{73}{80}$ while this formalization establishes the improved bound of $\frac{5}{8} \sqrt 2$.\ -definition \ :: real where "\ = sqrt 2" +definition \ :: real where "\ = sqrt 2" -lemma \_inv: "1/\ = \/2" +lemma \_inv: "1/\ = \/2" unfolding \_def by (simp add: real_div_sqrt) -definition \ :: "int \ int \ int \ int \ real" +definition \ :: "int \ int \ int \ int \ real" where "\ x y = (if compare (\_aux x) (\_aux y) then \ else (if compare (\_aux y) (\_aux x) then (1 / \) else 1))" -lemma \_sym: "\ x y * \ y x = 1" +lemma \_sym: "\ x y * \ y x = 1" unfolding \_def \_def compare_def by (auto simp add:prod_eq_iff) lemma \_nonneg: "\ x y \ 0" unfolding \_def \_def by auto definition \ :: "int \ real" where "\ x = \cos(pi*x/m)\" -definition \' :: "real \ real \ real" +definition \' :: "real \ real \ real" where "\' x y = (if abs (x - 1/2) < abs (y - 1/2) then \ else (if abs (x-1/2) > abs (y-1/2) then (1 / \) else 1))" definition \ :: "real \ real \ real" where "\ x y = \' y (frac(y-2*x))+\' y (frac (y+2*x))" lemma \'_cases: "abs (x-1/2) = abs (y-1/2) \ \' x y = 1" "abs (x-1/2) > abs (y-1/2) \ \' x y = 1/\" "abs (x-1/2) < abs (y-1/2) \ \' x y = \" unfolding \'_def by auto lemma if_cong_direct: assumes "a = b" assumes "c = d'" assumes "e = f" shows "(if a then c else e) = (if b then d' else f)" using assms by (intro if_cong) auto lemma \'_cong: assumes "abs (x-1/2) = abs (u-1/2)" assumes "abs (y-1/2) = abs (v-1/2)" shows "\' x y = \' u v" unfolding \'_def using assms by (intro if_cong_direct refl) auto lemma add_swap_cong: fixes x y u v :: "'a :: ab_semigroup_add" assumes "x = y" "u = v" shows "x + u = v + y" using assms by (simp add:algebra_simps) lemma frac_cong: fixes x y :: real assumes "x - y \ \" shows "frac x = frac y" proof - obtain k where x_eq: "x = y + of_int k" using Ints_cases[OF assms] by (metis add_minus_cancel uminus_add_conv_diff) thus ?thesis unfolding x_eq unfolding frac_def by simp qed lemma frac_expand: fixes x :: real shows "frac x = (if x < (-1) then (x-\x\) else (if x < 0 then (x+1) else (if x < 1 then x else (if x < 2 then (x-1) else (x-\x\)))))" proof - have "real_of_int y = -1 \ y= -1" for y by auto thus ?thesis - unfolding frac_def by (auto simp add:not_less floor_eq_iff) + unfolding frac_def by (auto simp add:not_less floor_eq_iff) qed lemma one_minus_frac: fixes x :: real shows "1 - frac x = (if x \ \ then 1 else frac (-x))" unfolding frac_neg by simp lemma abs_rev_cong: fixes x y :: real assumes "x = - y" shows "abs x = abs y" using assms by simp lemma cos_pi_ge_0: assumes "x \{-1/2.. 1/2}" shows "cos (pi * x) \ 0" proof - have "pi * x \ ((*) pi ` {-1/2..1/2})" by (intro imageI assms) - also have "... = {-pi/2..pi/2}" + also have "... = {-pi/2..pi/2}" by (subst image_mult_atLeastAtMost[OF pi_gt_zero]) simp finally have "pi * x \ {-pi/2..pi/2}" by simp thus ?thesis - by (intro cos_ge_zero) auto + by (intro cos_ge_zero) auto qed text \The following is the first step in establishing Eq. 15 in Hoory et al.~\cite[\S 8]{hoory2006}. Afterwards using various symmetries (diagonal, x-axis, y-axis) the result will follow for the entire square $[0,1] \times [0,1]$.\ lemma fun_bound_real_3: assumes "0 \ x" "x \ y" "y \ 1/2" "(x,y) \ (0,0)" shows "\cos(pi*x)\*\ x y + \cos(pi*y)\*\ y x \ 2.5 * sqrt 2" (is "?L \ ?R") proof - have apx:"4 \ 5 * sqrt (2::real)" "8 * cos (pi / 4) \ 5 * sqrt (2::real)" by (approximation 5)+ - have "cos (pi * x) \ 0" + have "cos (pi * x) \ 0" using assms(1,2,3) by (intro cos_pi_ge_0) simp - moreover have "cos (pi * y) \ 0" + moreover have "cos (pi * y) \ 0" using assms(1,2,3) by (intro cos_pi_ge_0) simp ultimately have 0:"?L = cos(pi*x)*\ x y + cos(pi*y)*\ y x" (is "_ = ?T") by simp consider (a) "x+y < 1/2" | (b) "y = 1/2- x" | (c) "x+y > 1/2" by argo hence "?T \ 2.5 * sqrt 2" (is "?T \ ?R") proof (cases) case a - consider - (1) "x < y" "x > 0" | - (2) "x=0" "y < 1/2" | - (3) "y=x" "x > 0" + consider + (1) "x < y" "x > 0" | + (2) "x=0" "y < 1/2" | + (3) "y=x" "x > 0" using assms(1,2,3,4) a by fastforce - thus ?thesis + thus ?thesis proof (cases) case 1 have "\ x y = \ + 1/\" unfolding \_def using 1 a - by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) + by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) moreover have "\ y x = 1/\ + 1/\" unfolding \_def using 1 a - by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) + by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) ultimately have "?T = cos (pi * x) * (\ + 1/\) + cos (pi * y) * (1/\ + 1/\)" by simp also have "... \ 1 * (\ + 1/\) + 1 * (1/\ + 1/\)" unfolding \_def by (intro add_mono mult_right_mono) auto also have "... = ?R" - unfolding \_def by (simp add:divide_simps) + unfolding \_def by (simp add:divide_simps) finally show ?thesis by simp next case 2 - have y_range: "y \ {0<..<1/2}" + have y_range: "y \ {0<..<1/2}" using assms 2 by simp have "\ 0 y = 1 + 1" unfolding \_def using y_range - by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) - moreover + by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) + moreover have "\x\ * 2 < 1 \ x < 1/2 \ -x < 1/2" for x :: real by auto hence "\ y 0 = 1 / \ + 1/ \" unfolding \_def using y_range by (intro arg_cong2[where f="(+)"] \'_cases) (simp_all add:frac_expand) ultimately have "?T = 2 + cos (pi * y) * (2 / \)" unfolding 2 by simp also have "... \ 2 + 1 * (2 / \)" unfolding \_def by (intro add_mono mult_right_mono) auto - also have "... \ ?R" + also have "... \ ?R" unfolding \_def by (approximation 10) finally show ?thesis by simp next case 3 have "\ x y = 1 + 1/\" unfolding \_def using 3 a - by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) + by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) moreover have "\ y x = 1 + 1/\" unfolding \_def using 3 a - by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) + by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) ultimately have "?T = cos (pi * x) * (2*(1+1/ \))" unfolding 3 by simp also have "... \ 1 * (2*(1+1/ \))" unfolding \_def by (intro mult_right_mono) auto also have "... \ ?R" unfolding \_def by (approximation 10) finally show ?thesis by simp qed next case b have x_range: "x \ {0..1/4}" using assms b by simp then consider (1) "x = 0" | (2) "x = 1/4" | (3) "x \ {0<..<1/4}" by fastforce - thus ?thesis + thus ?thesis proof (cases) case 1 hence y_eq: "y = 1/2" using b by simp - show ?thesis using apx unfolding 1 y_eq \_def by (simp add:\'_def \_def frac_def) + show ?thesis using apx unfolding 1 y_eq \_def by (simp add:\'_def \_def frac_def) next case 2 hence y_eq: "y = 1/4" using b by simp show ?thesis using apx unfolding y_eq 2 \_def by (simp add:\'_def frac_def) next case 3 have "\ x y = \ + 1" unfolding \_def b using 3 - by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) + by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) moreover have "\ y x = 1/\ + 1" unfolding \_def b using 3 - by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) + by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) ultimately have "?T = cos (pi * x) * (\ + 1) + cos (pi * (1 / 2 - x)) * (1/\ + 1)" unfolding b by simp also have "... \ ?R" unfolding \_def using x_range by (approximation 10 splitting: x=10) finally show ?thesis by simp qed next case c - consider - (1) "x < y" "y < 1/2" | - (2) "y=1/2" "x < 1/2" | - (3) "y=x" "x < 1/2" | + consider + (1) "x < y" "y < 1/2" | + (2) "y=1/2" "x < 1/2" | + (3) "y=x" "x < 1/2" | (4) "x=1/2" "y =1/2" using assms(2,3) c by fastforce thus ?thesis proof (cases) case 1 define \ :: real where "\ = arcsin (6 / 10)" have "cos \ = sqrt (1-0.6^2)" unfolding \_def by (intro cos_arcsin) auto also have "... = sqrt ( 0.8^2)" by (intro arg_cong[where f="sqrt"]) (simp add:power2_eq_square) also have "... = 0.8" by simp finally have cos_\: "cos \ = 0.8" by simp have sin_\: "sin \ = 0.6" unfolding \_def by simp have "\ x y = \ + \" unfolding \_def using c 1 by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) moreover have "\ y x = 1/\ + \" unfolding \_def using c 1 by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) ultimately have "?T = cos (pi * x) * (2 * \) + cos (pi * y) * (\ + 1 / \)" by simp also have "... \ cos (pi * (1/2-y)) * (2*\) + cos (pi * y) * (\+1 / \)" unfolding \_def using assms(1,2,3) c by (intro add_mono mult_right_mono order.refl iffD2[OF cos_mono_le_eq]) auto also have "... = (2.5*\)* (sin (pi * y) * 0.8 + cos (pi * y) * 0.6)" unfolding sin_cos_eq \_inv by (simp add:algebra_simps) also have "... = (2.5*\)* sin(pi*y + \)" unfolding sin_add cos_\ sin_\ by (intro arg_cong2[where f="(*)"] arg_cong2[where f="(+)"] refl) also have "... \ (?R) * 1" unfolding \_def by (intro mult_left_mono) auto finally show ?thesis by simp next case 2 have x_range: "x > 0" "x < 1/2" using c 2 by auto have "\ x y = \ + \" unfolding \_def 2 using x_range by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) moreover have "\ y x = 1 + 1" unfolding \_def 2 using x_range by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) ultimately have "?T = cos (pi * x) * (2*\)" unfolding 2 by simp also have "... \ 1 * (2* sqrt 2)" unfolding \_def by (intro mult_right_mono) auto - also have "... \ ?R" + also have "... \ ?R" by (approximation 5) finally show ?thesis by simp next case 3 have x_range: "x \ {1/4..1/2}" using 3 c by simp hence cos_bound: "cos (pi * x) \ 0.71" by (approximation 10) have "\ x y = 1+\" unfolding \_def 3 using 3 c by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) moreover have "\ y x = 1+\" unfolding \_def 3 using 3 c by (intro arg_cong2[where f="(+)"] \'_cases) (auto simp add:frac_expand) ultimately have "?T = 2 * cos (pi * x) * (1+\)" unfolding 3 by simp also have "... \ 2 * 0.71 * (1+sqrt 2)" unfolding \_def by (intro mult_right_mono mult_left_mono cos_bound) auto also have "... \ ?R" by (approximation 6) finally show ?thesis by simp - next + next case 4 show ?thesis unfolding 4 by simp qed qed thus ?thesis using 0 by simp qed text \Extend to square $[0, \frac{1}{2}] \times [0,\frac{1}{2}]$ using symmetry around x=y axis.\ lemma fun_bound_real_2: assumes "x \ {0..1/2}" "y \ {0..1/2}" "(x,y) \ (0,0)" shows "\cos(pi*x)\*\ x y + \cos(pi*y)\*\ y x \ 2.5 * sqrt 2" (is "?L \ ?R") proof (cases "y < x") case True have "?L = \cos(pi*y)\*\ y x + \cos(pi*x)\*\ x y" by simp also have "... \ ?R" using True assms by (intro fun_bound_real_3) auto finally show ?thesis by simp next case False then show ?thesis using assms by (intro fun_bound_real_3) auto qed text \Extend to $x > \frac{1}{2}$ using symmetry around $x=\frac{1}{2}$ axis.\ lemma fun_bound_real_1: assumes "x \ {0..<1}" "y \ {0..1/2}" "(x,y) \ (0,0)" shows "\cos(pi*x)\*\ x y + \cos(pi*y)\*\ y x \ 2.5 * sqrt 2" (is "?L \ ?R") proof (cases "x > 1/2") case True define x' where "x' = 1-x" have "\frac (x - 2 * y) - 1 / 2\ = \frac (1 - x + 2 * y) - 1 / 2\" proof (cases "x - 2 * y \ \") case True - then obtain k where x_eq: "x = 2*y + of_int k" using Ints_cases[OF True] + then obtain k where x_eq: "x = 2*y + of_int k" using Ints_cases[OF True] by (metis add_minus_cancel uminus_add_conv_diff) show ?thesis unfolding x_eq frac_def by simp next case False - hence "1 - x + 2 * y \ \" + hence "1 - x + 2 * y \ \" using Ints_1 Ints_diff by fastforce - thus ?thesis + thus ?thesis by (intro abs_rev_cong) (auto intro:frac_cong simp:one_minus_frac) qed - moreover have "\frac (x + 2 * y) - 1 / 2\ = \frac (1 - x - 2 * y) - 1 / 2\" + moreover have "\frac (x + 2 * y) - 1 / 2\ = \frac (1 - x - 2 * y) - 1 / 2\" proof (cases "x + 2 * y \ \") case True - then obtain k where x_eq: "x = of_int k - 2*y" using Ints_cases[OF True] + then obtain k where x_eq: "x = of_int k - 2*y" using Ints_cases[OF True] by (metis add.right_neutral add_diff_eq cancel_comm_monoid_add_class.diff_cancel) show ?thesis unfolding x_eq frac_def by simp next case False - hence "1 - x - 2 * y \ \" + hence "1 - x - 2 * y \ \" using Ints_1 Ints_diff by fastforce - thus ?thesis + thus ?thesis by (intro abs_rev_cong) (auto intro:frac_cong simp:one_minus_frac) qed - ultimately have "\ y x = \ y x'" - unfolding \_def x'_def by (intro \'_cong add_swap_cong) simp_all + ultimately have "\ y x = \ y x'" + unfolding \_def x'_def by (intro \'_cong add_swap_cong) simp_all moreover have "\ x y = \ x' y" - unfolding \_def x'_def - by (intro \'_cong add_swap_cong refl arg_cong[where f="(\x. abs (x-1/2))"] frac_cong) + unfolding \_def x'_def + by (intro \'_cong add_swap_cong refl arg_cong[where f="(\x. abs (x-1/2))"] frac_cong) (simp_all add:algebra_simps) moreover have "\cos(pi*x)\ = \cos(pi*x')\" unfolding x'_def by (intro abs_rev_cong) (simp add:algebra_simps) ultimately have "?L = \cos(pi*x')\*\ x' y + \cos(pi*y)\*\ y x'" by simp also have "... \ ?R" using assms True by (intro fun_bound_real_2) (auto simp add:x'_def) finally show ?thesis by simp next case False - thus ?thesis using assms fun_bound_real_2 by simp + thus ?thesis using assms fun_bound_real_2 by simp qed text \Extend to $y > \frac{1}{2}$ using symmetry around $y=\frac{1}{2}$ axis.\ lemma fun_bound_real: assumes "x \ {0..<1}" "y \ {0..<1}" "(x,y) \ (0,0)" shows "\cos(pi*x)\*\ x y + \cos(pi*y)\*\ y x \ 2.5 * sqrt 2" (is "?L \ ?R") proof (cases "y > 1/2") case True define y' where "y' = 1-y" have "\frac (y - 2 * x) - 1 / 2\ = \frac (1 - y + 2 * x) - 1 / 2\" proof (cases "y - 2 * x \ \") case True - then obtain k where y_eq: "y = 2*x + of_int k" using Ints_cases[OF True] + then obtain k where y_eq: "y = 2*x + of_int k" using Ints_cases[OF True] by (metis add_minus_cancel uminus_add_conv_diff) show ?thesis unfolding y_eq frac_def by simp next case False - hence "1 - y + 2 * x \ \" + hence "1 - y + 2 * x \ \" using Ints_1 Ints_diff by fastforce - thus ?thesis + thus ?thesis by (intro abs_rev_cong) (auto intro:frac_cong simp:one_minus_frac) qed - moreover have "\frac (y + 2 * x) - 1 / 2\ = \frac (1 - y - 2 * x) - 1 / 2\" + moreover have "\frac (y + 2 * x) - 1 / 2\ = \frac (1 - y - 2 * x) - 1 / 2\" proof (cases "y + 2 * x \ \") case True - then obtain k where y_eq: "y = of_int k - 2*x" using Ints_cases[OF True] + then obtain k where y_eq: "y = of_int k - 2*x" using Ints_cases[OF True] by (metis add.right_neutral add_diff_eq cancel_comm_monoid_add_class.diff_cancel) show ?thesis unfolding y_eq frac_def by simp next case False - hence "1 - y - 2 * x \ \" + hence "1 - y - 2 * x \ \" using Ints_1 Ints_diff by fastforce - thus ?thesis + thus ?thesis by (intro abs_rev_cong) (auto intro:frac_cong simp:one_minus_frac) qed - ultimately have "\ x y = \ x y'" + ultimately have "\ x y = \ x y'" unfolding \_def y'_def by (intro \'_cong add_swap_cong) simp_all moreover have "\ y x = \ y' x" - unfolding \_def y'_def - by (intro \'_cong add_swap_cong refl arg_cong[where f="(\x. abs (x-1/2))"] frac_cong) + unfolding \_def y'_def + by (intro \'_cong add_swap_cong refl arg_cong[where f="(\x. abs (x-1/2))"] frac_cong) (simp_all add:algebra_simps) moreover have "\cos(pi*y)\ = \cos(pi*y')\" unfolding y'_def by (intro abs_rev_cong) (simp add:algebra_simps) ultimately have "?L = \cos(pi*x)\*\ x y' + \cos(pi*y')\*\ y' x" by simp also have "... \ ?R" using assms True by (intro fun_bound_real_1) (auto simp add:y'_def) finally show ?thesis by simp next case False - thus ?thesis using assms fun_bound_real_1 by simp + thus ?thesis using assms fun_bound_real_1 by simp qed -lemma mod_to_frac: +lemma mod_to_frac: fixes x :: int shows "real_of_int (x mod m) = m * frac (x/m)" (is "?L = ?R") proof - obtain y where y_def: "x mod m = x + int m* y" - by (metis mod_eqE mod_mod_trivial) + by (metis mod_eqE mod_mod_trivial) have 0: "x mod int m < m" "x mod int m \ 0" using m_gt_0 by auto have "?L = real m * (of_int (x mod m) / m )" using m_gt_0 by (simp add:algebra_simps) also have "... = real m * frac (of_int (x mod m) / m)" using 0 by (subst iffD2[OF frac_eq]) auto also have "... = real m * frac (x / m + y)" unfolding y_def using m_gt_0 by (simp add:divide_simps mult.commute) - also have "... = ?R" + also have "... = ?R" unfolding frac_def by simp finally show ?thesis by simp qed lemma fun_bound: assumes "v \ verts G" "v \ (0,0)" - shows "\(fst v)*(\ v (S\<^sub>2 v)+\ v (T\<^sub>2 v))+\(snd v)*(\ v (S\<^sub>1 v)+\ v (T\<^sub>1 v)) \ 2.5 * sqrt 2" + shows "\(fst v)*(\ v (S\<^sub>2 v)+\ v (T\<^sub>2 v))+\(snd v)*(\ v (S\<^sub>1 v)+\ v (T\<^sub>1 v)) \ 2.5 * sqrt 2" (is "?L \ ?R") proof - obtain x y where v_def: "v = (x,y)" by (cases v) auto define x' where "x' = x/real m" define y' where "y' = y/real m" have 0:"\ v (S\<^sub>1 v) = \' x' (frac(x'-2*y'))" unfolding \_def \'_def compare_def v_def \_aux_def T\<^sub>1_def S\<^sub>1_def x'_def y'_def using m_gt_0 - by (intro if_cong_direct refl) (auto simp add:case_prod_beta mod_to_frac divide_simps) + by (intro if_cong_direct refl) (auto simp add:case_prod_beta mod_to_frac divide_simps) have 1:"\ v (T\<^sub>1 v) = \' x' (frac(x'+2*y'))" unfolding \_def \'_def compare_def v_def \_aux_def T\<^sub>1_def x'_def y'_def using m_gt_0 - by (intro if_cong_direct refl) (auto simp add:case_prod_beta mod_to_frac divide_simps) + by (intro if_cong_direct refl) (auto simp add:case_prod_beta mod_to_frac divide_simps) have 2:"\ v (S\<^sub>2 v) = \' y' (frac(y'-2*x'))" unfolding \_def \'_def compare_def v_def \_aux_def S\<^sub>2_def x'_def y'_def using m_gt_0 - by (intro if_cong_direct refl) (auto simp add:case_prod_beta mod_to_frac divide_simps) + by (intro if_cong_direct refl) (auto simp add:case_prod_beta mod_to_frac divide_simps) have 3:"\ v (T\<^sub>2 v) = \' y' (frac(y'+2*x'))" unfolding \_def \'_def compare_def v_def \_aux_def T\<^sub>2_def x'_def y'_def using m_gt_0 - by (intro if_cong_direct refl) (auto simp add:case_prod_beta mod_to_frac divide_simps) + by (intro if_cong_direct refl) (auto simp add:case_prod_beta mod_to_frac divide_simps) have 4: "\ (fst v) = \cos(pi*x')\" "\ (snd v) = \cos(pi*y')\" unfolding \_def v_def x'_def y'_def by auto have "x \ {0.. {0.. (0,0)" using assms unfolding v_def mgg_graph_def by auto - hence 5:"x' \ {0..<1}" "y' \ {0..<1}" "(x',y') \ (0,0)" + hence 5:"x' \ {0..<1}" "y' \ {0..<1}" "(x',y') \ (0,0)" unfolding x'_def y'_def by auto have "?L = \cos(pi*x')\*\ x' y' + \cos(pi*y')\*\ y' x'" unfolding 0 1 2 3 4 \_def by simp also have "... \ ?R" - by (intro fun_bound_real 5) + by (intro fun_bound_real 5) finally show ?thesis by simp qed text \Equation 15 in Proof of Theorem 8.8\ lemma hoory_8_8: fixes f :: "int \ int \ real" assumes "\x. f x \ 0" assumes "f (0,0) = 0" assumes "periodic f" shows "g_inner f (\x. f(S\<^sub>2 x)*\ (fst x)+f(S\<^sub>1 x)*\ (snd x))\1.25* sqrt 2*g_norm f^2" (is "?L \ ?R") proof - have 0: "2 * f x * f y \ \ x y * f x^2 + \ y x * f y^2" (is "?L1 \ ?R1") for x y - proof - + proof - have "0 \ ((sqrt (\ x y) * f x) - (sqrt (\ y x) * f y))^2" by simp also have "... = ?R1 - 2 * (sqrt (\ x y) * f x) * (sqrt (\ y x) * f y)" unfolding power2_diff using \_nonneg assms(1) by (intro arg_cong2[where f="(-)"] arg_cong2[where f="(+)"]) (auto simp add: power2_eq_square) also have "... = ?R1 -2 * sqrt (\ x y * \ y x) * f x * f y" unfolding real_sqrt_mult by simp also have "... = ?R1 - ?L1" unfolding \_sym by simp finally have "0 \ ?R1 - ?L1" by simp thus ?thesis by simp qed have [simp]: "fst (S\<^sub>2 x) = fst x" "snd (S\<^sub>1 x) = snd x" for x unfolding S\<^sub>1_def S\<^sub>2_def by auto have S_2_inv [simp]: "T\<^sub>2 (S\<^sub>2 x) = x" if "x \ verts G" for x - using that unfolding T\<^sub>2_def S\<^sub>2_def mgg_graph_def + using that unfolding T\<^sub>2_def S\<^sub>2_def mgg_graph_def by (cases x,simp add:mod_simps) have S_1_inv [simp]: "T\<^sub>1 (S\<^sub>1 x) = x" if "x \ verts G" for x - using that unfolding T\<^sub>1_def S\<^sub>1_def mgg_graph_def + using that unfolding T\<^sub>1_def S\<^sub>1_def mgg_graph_def by (cases x,simp add:mod_simps) have S2_inj: "inj_on S\<^sub>2 (verts G)" using S_2_inv by (intro inj_on_inverseI[where g="T\<^sub>2"]) have S1_inj: "inj_on S\<^sub>1 (verts G)" using S_1_inv by (intro inj_on_inverseI[where g="T\<^sub>1"]) have "S\<^sub>2 ` verts G \ verts G" unfolding mgg_graph_def S\<^sub>2_def by (intro image_subsetI) auto - hence S2_ran: "S\<^sub>2 ` verts G = verts G" + hence S2_ran: "S\<^sub>2 ` verts G = verts G" by (intro card_subset_eq card_image S2_inj) auto have "S\<^sub>1 ` verts G \ verts G" unfolding mgg_graph_def S\<^sub>1_def by (intro image_subsetI) auto hence S1_ran: "S\<^sub>1 ` verts G = verts G" - by (intro card_subset_eq card_image S1_inj) auto + by (intro card_subset_eq card_image S1_inj) auto have 2: "g v * f v^2 \ 2.5 * sqrt 2 * f v^2" if "g v \ 2.5 * sqrt 2 \ v = (0,0)" for v g proof (cases "v=(0,0)") case True then show ?thesis using assms(2) by simp next case False then show ?thesis using that by (intro mult_right_mono) auto qed have "2*?L=(\v\verts G. \(fst v)*(2*f v *f(S\<^sub>2 v)))+(\v\verts G. \(snd v) * (2 * f v * f (S\<^sub>1 v)))" unfolding g_inner_def by (simp add: algebra_simps sum_distrib_left sum.distrib) also have "... \ (\v\verts G. \(fst v)*(\ v (S\<^sub>2 v) * f v^2 + \ (S\<^sub>2 v) v * f(S\<^sub>2 v)^2))+ (\v\verts G. \(snd v)*(\ v (S\<^sub>1 v) * f v^2 + \ (S\<^sub>1 v) v * f(S\<^sub>1 v)^2))" unfolding \_def by (intro add_mono sum_mono mult_left_mono 0) auto - also have "... = + also have "... = (\v\verts G. \(fst v)*\ v (S\<^sub>2 v)*f v^2)+(\v\verts G. \(fst v) * \ (S\<^sub>2 v) v * f(S\<^sub>2 v)^2)+ (\v\verts G. \(snd v)*\ v (S\<^sub>1 v)*f v^2)+(\v\verts G. \(snd v) * \ (S\<^sub>1 v) v * f(S\<^sub>1 v)^2)" by (simp add:sum.distrib algebra_simps) - also have "... = + also have "... = (\v\verts G. \(fst v)*\ v (S\<^sub>2 v)*f v^2)+ (\v\verts G. \(fst (S\<^sub>2 v)) * \ (S\<^sub>2 v) (T\<^sub>2 (S\<^sub>2 v)) * f(S\<^sub>2 v)^2)+ (\v\verts G. \(snd v)*\ v (S\<^sub>1 v)*f v^2)+ (\v\verts G. \(snd (S\<^sub>1 v)) * \ (S\<^sub>1 v) (T\<^sub>1 (S\<^sub>1 v)) * f(S\<^sub>1 v)^2)" by (intro arg_cong2[where f="(+)"] sum.cong refl) simp_all also have "... = (\v\verts G. \(fst v)*\ v (S\<^sub>2 v)*f v^2)+ (\v\S\<^sub>2 ` verts G. \(fst v) * \ v (T\<^sub>2 v) * f v^2)+ (\v\verts G. \(snd v)*\ v (S\<^sub>1 v)*f v^2)+ (\v\S\<^sub>1 ` verts G. \(snd v) * \ v (T\<^sub>1 v) * f v^2)" using S1_inj S2_inj by (simp add:sum.reindex) - also have "... = + also have "... = (\v\verts G. (\(fst v)*(\ v (S\<^sub>2 v)+\ v (T\<^sub>2 v))+\(snd v)*(\ v (S\<^sub>1 v)+\ v (T\<^sub>1 v))) *f v^2)" unfolding S1_ran S2_ran by (simp add:algebra_simps sum.distrib) also have "... \ (\v\verts G. 2.5 * sqrt 2 * f v^2)" using fun_bound by (intro sum_mono 2) auto - also have "... \ 2.5 * sqrt 2 * g_norm f^2" - unfolding g_norm_sq g_inner_def + also have "... \ 2.5 * sqrt 2 * g_norm f^2" + unfolding g_norm_sq g_inner_def by (simp add:algebra_simps power2_eq_square sum_distrib_left) finally have "2 * ?L \ 2.5 * sqrt 2 * g_norm f^2" by simp thus ?thesis by simp qed lemma hoory_8_7: fixes f :: "int\int \ complex" assumes "f (0,0) = 0" assumes "periodic f" shows "norm(g_inner f (\x. f (S\<^sub>2 x) * (1+\\<^sub>F (fst x)) + f (S\<^sub>1 x) * (1+\\<^sub>F (snd x)))) \ (2.5 * sqrt 2) * (\v \ verts G. norm (f v)^2)" (is "?L \ ?R") proof - define g :: "int\int \ real" where "g x = norm (f x)" for x have g_zero: "g (0,0) = 0" using assms(1) unfolding g_def by simp have g_nonneg: "g x \ 0" for x unfolding g_def by simp have g_periodic: "periodic g" unfolding g_def by (intro periodic_comp[OF assms(2)]) have 0: "norm(1+\\<^sub>F x) = 2*\ x" for x :: int proof - have "norm(1+\\<^sub>F x) = norm(\\<^sub>F (-x/2)*(\\<^sub>F 0 + \\<^sub>F x))" unfolding \\<^sub>F_def norm_mult by simp also have "... = norm (\\<^sub>F (0-x/2) + \\<^sub>F (x-x/2))" unfolding \\<^sub>F_simps by (simp add: algebra_simps) also have "... = norm (\\<^sub>F (x/2) + cnj (\\<^sub>F (x/2)))" unfolding \\<^sub>F_simps(3) by (simp add:algebra_simps) also have "... = \2*Re (\\<^sub>F (x/2))\" unfolding complex_add_cnj norm_of_real by simp also have "... = 2*\cos(pi*x/m)\" unfolding \\<^sub>F_def cis.simps by simp also have "... = 2*\ x" unfolding \_def by simp finally show ?thesis by simp qed have "?L\norm(\v\verts G. f v * cnj(f(S\<^sub>2 v)*(1+\\<^sub>F (fst v))+f(S\<^sub>1 v )*(1+\\<^sub>F (snd v))))" unfolding g_inner_def by (simp add:case_prod_beta) also have "...\(\v\verts G. norm(f v * cnj(f (S\<^sub>2 v) *(1+\\<^sub>F (fst v))+f (S\<^sub>1 v)*(1+\\<^sub>F (snd v)))))" by (intro norm_sum) also have "...=(\v\verts G. g v * norm(f (S\<^sub>2 v) *(1+\\<^sub>F (fst v))+f (S\<^sub>1 v)*(1+\\<^sub>F (snd v))))" unfolding norm_mult g_def complex_mod_cnj by simp also have "...\(\v\verts G. g v * (norm (f(S\<^sub>2 v)*(1+\\<^sub>F (fst v)))+norm(f(S\<^sub>1 v)*(1+\\<^sub>F(snd v)))))" by (intro sum_mono norm_triangle_ineq mult_left_mono g_nonneg) also have "...=2*g_inner g (\x. g (S\<^sub>2 x)*\ (fst x)+g(S\<^sub>1 x)*\ (snd x))" - unfolding g_def g_inner_def norm_mult 0 + unfolding g_def g_inner_def norm_mult 0 by (simp add:sum_distrib_left algebra_simps case_prod_beta) also have "... \2*(1.25* sqrt 2*g_norm g^2)" by (intro mult_left_mono hoory_8_8 g_nonneg g_zero g_periodic) auto also have "... = ?R" unfolding g_norm_sq g_def g_inner_def by (simp add:power2_eq_square) finally show ?thesis by simp qed lemma hoory_8_3: assumes "g_inner f (\_. 1) = 0" assumes "periodic f" - shows "\(\(x,y)\verts G. f(x,y)*(f(x+2*y,y)+f(x+2*y+1,y)+f(x,y+2*x)+f(x,y+2*x+1)))\ + shows "\(\(x,y)\verts G. f(x,y)*(f(x+2*y,y)+f(x+2*y+1,y)+f(x,y+2*x)+f(x,y+2*x+1)))\ \ (2.5 * sqrt 2) * g_norm f^2" (is "\?L\ \ ?R") proof - let ?f = "(\x. complex_of_real (f x))" - define Ts :: "(int \ int \ int \ int) list" where + define Ts :: "(int \ int \ int \ int) list" where "Ts = [(\(x,y).(x+2*y,y)),(\(x,y).(x+2*y+1,y)),(\(x,y).(x,y+2*x)),(\(x,y).(x,y+2*x+1))]" - have p: "periodic ?f" + have p: "periodic ?f" by (intro periodic_comp[OF assms(2)]) - have 0: "(\T\Ts. FT (?f\T) v) = FT ?f (S\<^sub>2 v)*(1+\\<^sub>F (fst v))+FT ?f (S\<^sub>1 v)*(1+\\<^sub>F (snd v))" + have 0: "(\T\Ts. FT (?f\T) v) = FT ?f (S\<^sub>2 v)*(1+\\<^sub>F (fst v))+FT ?f (S\<^sub>1 v)*(1+\\<^sub>F (snd v))" (is "?L1 = ?R1") for v :: "int \ int" proof - obtain x y where v_def: "v = (x,y)" by (cases v, auto) have "?L1 = (\T\Ts. FT (?f\T) (x,y))" unfolding v_def by simp also have "... = FT ?f (x,y-2*x)*(1+\\<^sub>F x) + FT ?f (x-2*y,y)*(1+\\<^sub>F y)" unfolding Ts_def by (simp add:FT_sheer[OF p] case_prod_beta comp_def) (simp add:algebra_simps) also have "... = ?R1" unfolding v_def S\<^sub>2_def S\<^sub>1_def by (intro arg_cong2[where f="(+)"] arg_cong2[where f="(*)"] periodic_cong[OF periodic_FT]) auto finally show ?thesis by simp qed have "cmod ((of_nat m)^2) = cmod (of_real (of_nat m^2))" by simp also have "... = abs (of_nat m^2)" by (intro norm_of_real) also have "... = real m^2" by simp finally have 1: "cmod ((of_nat m)\<^sup>2) = (real m)\<^sup>2" by simp have "FT (\x. complex_of_real (f x)) (0, 0) = complex_of_real (g_inner f (\_ . 1))" unfolding FT_def g_inner_def g_inner_def \\<^sub>F_def by simp also have "... = 0" unfolding assms by simp finally have 2: "FT (\x. complex_of_real (f x)) (0, 0) = 0" by simp - have "abs ?L = norm (complex_of_real ?L)" + have "abs ?L = norm (complex_of_real ?L)" unfolding norm_of_real by simp also have "... = norm (\T \ Ts. (g_inner ?f (?f \ T)))" unfolding Ts_def by (simp add:algebra_simps g_inner_def sum.distrib comp_def case_prod_beta) - also have "... = norm (\T \ Ts. (g_inner (FT ?f) (FT (?f \ T)))/m^2)" + also have "... = norm (\T \ Ts. (g_inner (FT ?f) (FT (?f \ T)))/m^2)" by (subst parseval) simp also have "... = norm (g_inner (FT ?f) (\x. (\T \ Ts. (FT (?f \ T) x)))/m^2)" unfolding Ts_def by (simp add:g_inner_simps case_prod_beta add_divide_distrib) also have "...=norm(g_inner(FT ?f)(\x.(FT ?f(S\<^sub>2 x)*(1+\\<^sub>F (fst x))+FT f(S\<^sub>1 x)*(1+\\<^sub>F (snd x)))))/m^2" by (subst 0) (simp add:norm_divide 1) also have "... \ (2.5 * sqrt 2) * (\v \ verts G. norm (FT f v)^2) / m^2" - by (intro divide_right_mono hoory_8_7[where f="FT f"] 2 periodic_FT) auto + by (intro divide_right_mono hoory_8_7[where f="FT f"] 2 periodic_FT) auto also have "... = (2.5 * sqrt 2) * (\v \ verts G. cmod (f v)^2)" by (subst (2) plancharel) simp also have "... = (2.5 * sqrt 2) * (g_inner f f)" unfolding g_inner_def norm_of_real by (simp add: power2_eq_square) - also have "... = ?R" + also have "... = ?R" using g_norm_sq by auto finally show ?thesis by simp qed -text \Inequality stated before Theorem 8.3 in Hoory.\ +text \Inequality stated before Theorem 8.3 in Hoory.\ lemma mgg_numerical_radius_aux: assumes "g_inner f (\_. 1) = 0" shows "\(\a \ arcs G. f (head G a) * f (tail G a))\ \ (5 * sqrt 2) * g_norm f^2" (is "?L \ ?R") proof - define g where "g x = f (fst x mod m, snd x mod m)" for x :: "int \ int" - have 0:"g x = f x" if "x \ verts G" for x + have 0:"g x = f x" if "x \ verts G" for x unfolding g_def using that by (auto simp add:mgg_graph_def mem_Times_iff) have g_mod_simps[simp]: "g (x, y mod m) = g (x, y)" "g (x mod m, y) = g (x, y)" for x y :: int unfolding g_def by auto have periodic_g: "periodic g" unfolding periodic_def by simp have "g_inner g (\_. 1) = g_inner f (\_. 1)" by (intro g_inner_cong 0) auto also have "... = 0" using assms by simp finally have 1:"g_inner g (\_. 1) = 0" by simp have 2:"g_norm g = g_norm f" by (intro g_norm_cong 0) (auto) have "?L = \(\a \ arcs G. g (head G a) * g (tail G a))\" using wellformed by (intro arg_cong[where f="abs"] sum.cong arg_cong2[where f="(*)"] 0[symmetric]) auto also have "...=\(\a\arcs_pos. g(head G a)*g(tail G a))+(\a\arcs_neg. g(head G a)*g(tail G a))\" unfolding arcs_sym arcs_pos_def arcs_neg_def by (intro arg_cong[where f="abs"] sum.union_disjoint) auto also have "... = \2 * (\(v,l)\verts G \ {..<4}. g v * g (mgg_graph_step m v (l, 1)))\" unfolding arcs_pos_def arcs_neg_def by (simp add:inj_on_def sum.reindex case_prod_beta mgg_graph_def algebra_simps) also have "... = 2 * \(\v \ verts G. (\l \ {..<4}. g v * g (mgg_graph_step m v (l, 1))))\" by (subst sum.cartesian_product) (simp add:abs_mult) also have "... = 2*\(\(x,y)\verts G. (\l\[0..<4]. g(x,y)* g (mgg_graph_step m (x,y) (l,1))))\" by (subst interv_sum_list_conv_sum_set_nat) - (auto simp add:atLeast0LessThan case_prod_beta simp del:mgg_graph_step.simps) + (auto simp add:atLeast0LessThan case_prod_beta simp del:mgg_graph_step.simps) also have "... =2*\\(x,y)\verts G. g (x,y)* (g(x+2*y,y)+g(x+2*y+1,y)+g(x,y+2*x)+g(x,y+2*x+1))\" - by (simp add:case_prod_beta numeral_eq_Suc algebra_simps) + by (simp add:case_prod_beta numeral_eq_Suc algebra_simps) also have "... \ 2* ((2.5 * sqrt 2) * g_norm g^2)" by (intro mult_left_mono hoory_8_3 1 periodic_g) auto also have "... \ ?R" unfolding 2 by simp finally show ?thesis by simp qed definition MGG_bound :: real where "MGG_bound = 5 * sqrt 2 / 8" -text \Main result: Theorem 8.2 in Hoory.\ +text \Main result: Theorem 8.2 in Hoory.\ lemma mgg_numerical_radius: "\\<^sub>a \ MGG_bound" proof - have "\\<^sub>a \ (5 * sqrt 2)/real d" by (intro expander_intro mgg_numerical_radius_aux) auto also have "... = MGG_bound" - unfolding MGG_bound_def d_eq_8 by simp + unfolding MGG_bound_def d_eq_8 by simp finally show ?thesis by simp qed end end \ No newline at end of file diff --git a/thys/Expander_Graphs/Expander_Graphs_Multiset_Extras.thy b/thys/Expander_Graphs/Expander_Graphs_Multiset_Extras.thy --- a/thys/Expander_Graphs/Expander_Graphs_Multiset_Extras.thy +++ b/thys/Expander_Graphs/Expander_Graphs_Multiset_Extras.thy @@ -1,195 +1,225 @@ subsection \Multisets\ text \Some preliminary results about multisets.\ theory Expander_Graphs_Multiset_Extras - imports - Frequency_Moments.Frequency_Moments_Preliminary_Results + imports + "HOL-Library.Multiset" Extra_Congruence_Method begin unbundle intro_cong_syntax -lemma sum_mset_conv: +text \This is an induction scheme over the distinct elements of a multisets: +We can represent each multiset as a sum like: +@{text "replicate_mset n\<^sub>1 x\<^sub>1 + replicate_mset n\<^sub>2 x\<^sub>2 + ... + replicate_mset n\<^sub>k x\<^sub>k"} where the +@{term "x\<^sub>i"} are distinct.\ + +lemma disj_induct_mset: + assumes "P {#}" + assumes "\n M x. P M \ \(x \# M) \ n > 0 \ P (M + replicate_mset n x)" + shows "P M" +proof (induction "size M" arbitrary: M rule:nat_less_induct) + case 1 + show ?case + proof (cases "M = {#}") + case True + then show ?thesis using assms by simp + next + case False + then obtain x where x_def: "x \# M" using multiset_nonemptyE by auto + define M1 where "M1 = M - replicate_mset (count M x) x" + then have M_def: "M = M1 + replicate_mset (count M x) x" + by (metis count_le_replicate_mset_subset_eq dual_order.refl subset_mset.diff_add) + have "size M1 < size M" + by (metis M_def x_def count_greater_zero_iff less_add_same_cancel1 size_replicate_mset size_union) + hence "P M1" using 1 by blast + then show "P M" + apply (subst M_def, rule assms(2), simp) + by (simp add:M1_def x_def count_eq_zero_iff[symmetric])+ + qed +qed + +lemma sum_mset_conv: fixes f :: "'a \ 'b::{semiring_1}" shows "sum_mset (image_mset f A) = sum (\x. of_nat (count A x) * f x) (set_mset A)" proof (induction A rule: disj_induct_mset) case 1 then show ?case by simp next case (2 n M x) moreover have "count M x = 0" using 2 by (simp add: count_eq_zero_iff) moreover have "\y. y \ set_mset M \ y \ x" using 2 by blast - ultimately show ?case by (simp add:algebra_simps) + ultimately show ?case by (simp add:algebra_simps) qed -lemma sum_mset_conv_2: +lemma sum_mset_conv_2: fixes f :: "'a \ 'b::{semiring_1}" assumes "set_mset A \ B" "finite B" shows "sum_mset (image_mset f A) = sum (\x. of_nat (count A x) * f x) B" (is "?L = ?R") proof - have "?L = sum (\x. of_nat (count A x) * f x) (set_mset A)" unfolding sum_mset_conv by simp also have "... = ?R" by (intro sum.mono_neutral_left assms) (simp_all add: iffD2[OF count_eq_zero_iff]) finally show ?thesis by simp qed lemma count_mset_exp: "count A x = size (filter_mset (\y. y = x) A)" by (induction A, simp, simp) lemma mset_repl: "mset (replicate k x) = replicate_mset k x" by (induction k, auto) lemma count_image_mset_inj: assumes "inj f" shows "count (image_mset f A) (f x) = count A x" proof (cases "x \ set_mset A") case True - hence "f -` {f x} \ set_mset A = {x}" - using assms by (auto simp add:vimage_def inj_def) + hence "f -` {f x} \ set_mset A = {x}" + using assms by (auto simp add:vimage_def inj_def) then show ?thesis by (simp add:count_image_mset) next case False - hence "f -` {f x} \ set_mset A = {}" - using assms by (auto simp add:vimage_def inj_def) + hence "f -` {f x} \ set_mset A = {}" + using assms by (auto simp add:vimage_def inj_def) thus ?thesis using False by (simp add:count_image_mset count_eq_zero_iff) qed lemma count_image_mset_0_triv: assumes "x \ range f" - shows "count (image_mset f A) x = 0" + shows "count (image_mset f A) x = 0" proof - - have "x \ set_mset (image_mset f A)" + have "x \ set_mset (image_mset f A)" using assms by auto - thus ?thesis + thus ?thesis by (meson count_inI) qed lemma filter_mset_ex_predicates: assumes "\x. \ P x \ \ Q x" shows "filter_mset P M + filter_mset Q M = filter_mset (\x. P x \ Q x) M" using assms by (induction M, auto) -lemma sum_count_2: +lemma sum_count_2: assumes "finite F" shows "sum (count M) F = size (filter_mset (\x. x \ F) M)" using assms proof (induction F rule:finite_induct) case empty then show ?case by simp next case (insert x F) have "sum (count M) (insert x F) = size ({#y \# M. y = x#} + {#x \# M. x \ F#})" using insert(1,2,3) by (simp add:count_mset_exp) also have "... = size ({#y \# M. y = x \ y \ F#})" using insert(2) by (intro arg_cong[where f="size"] filter_mset_ex_predicates) simp also have "... = size (filter_mset (\y. y \ insert x F) M)" by simp finally show ?case by simp qed definition concat_mset :: "('a multiset) multiset \ 'a multiset" where "concat_mset xss = fold_mset (\xs ys. xs + ys) {#} xss" lemma image_concat_mset: "image_mset f (concat_mset xss) = concat_mset (image_mset (image_mset f) xss)" unfolding concat_mset_def by (induction xss, auto) lemma concat_add_mset: "concat_mset (image_mset (\x. f x + g x) xs) = concat_mset (image_mset f xs) + concat_mset (image_mset g xs)" unfolding concat_mset_def by (induction xs) auto lemma concat_add_mset_2: "concat_mset (xs + ys) = concat_mset xs + concat_mset ys" unfolding concat_mset_def by (induction xs, auto) lemma size_concat_mset: "size (concat_mset xss) = sum_mset (image_mset size xss)" unfolding concat_mset_def by (induction xss, auto) lemma filter_concat_mset: "filter_mset P (concat_mset xss) = concat_mset (image_mset (filter_mset P) xss)" unfolding concat_mset_def by (induction xss, auto) lemma count_concat_mset: "count (concat_mset xss) xs = sum_mset (image_mset (\x. count x xs) xss)" unfolding concat_mset_def by (induction xss, auto) lemma set_mset_concat_mset: "set_mset (concat_mset xss) = \ (set_mset ` (set_mset xss))" unfolding concat_mset_def by (induction xss, auto) lemma concat_mset_empty: "concat_mset {#} = {#}" unfolding concat_mset_def by simp lemma concat_mset_single: "concat_mset {#x#} = x" unfolding concat_mset_def by simp lemma concat_disjoint_union_mset: assumes "finite I" assumes "\i. i \ I \ finite (A i)" assumes "\i j. i \ I \ j \ I \ i \ j \ A i \ A j = {}" shows "mset_set (\ (A ` I)) = concat_mset (image_mset (mset_set \ A) (mset_set I))" using assms proof (induction I rule:finite_induct) case empty then show ?case by (simp add:concat_mset_empty) next case (insert x F) have "mset_set (\ (A ` insert x F)) = mset_set (A x \ (\ (A ` F)))" by simp also have "... = mset_set (A x) + mset_set (\ (A ` F))" using insert by (intro mset_set_Union) auto also have "... = mset_set (A x) + concat_mset (image_mset (mset_set \ A) (mset_set F))" using insert by (intro arg_cong2[where f="(+)"] insert(3)) auto also have "... = concat_mset (image_mset (mset_set \ A) ({#x#} + mset_set F))" by (simp add:concat_mset_def) also have "... = concat_mset (image_mset (mset_set \ A) (mset_set (insert x F)))" using insert by (intro_cong "[\\<^sub>1 concat_mset, \\<^sub>2 image_mset]") auto finally show ?case by blast qed lemma size_filter_mset_conv: "size (filter_mset f A) = sum_mset (image_mset (\x. of_bool (f x) :: nat) A)" by (induction A, auto) lemma filter_mset_const: "filter_mset (\_. c) xs = (if c then xs else {#})" by simp -lemma repeat_image_concat_mset: +lemma repeat_image_concat_mset: "repeat_mset n (image_mset f A) = concat_mset (image_mset (\x. replicate_mset n (f x)) A)" unfolding concat_mset_def by (induction A, auto) lemma mset_prod_eq: assumes "finite A" "finite B" - shows + shows "mset_set (A \ B) = concat_mset {# {# (x,y). y \# mset_set B #} .x \# mset_set A #}" using assms(1) proof (induction rule:finite_induct) case empty then show ?case unfolding concat_mset_def by simp next case (insert x F) have "mset_set (insert x F \ B) = mset_set (F \ B \ (\y. (x,y)) ` B)" by (intro arg_cong[where f="mset_set"]) auto also have "... = mset_set (F \ B) + mset_set ((\y. (x,y)) ` B)" using insert(1,2) assms(2) by (intro mset_set_Union finite_cartesian_product) auto also have "... = mset_set (F \ B) + {# (x,y). y \# mset_set B #}" by (intro arg_cong2[where f="(+)"] image_mset_mset_set[symmetric] inj_onI) auto also have "... = concat_mset {#image_mset (Pair x) (mset_set B). x \# {#x#} + (mset_set F)#}" unfolding insert image_mset_union concat_add_mset_2 by (simp add:concat_mset_single) also have "... = concat_mset {#image_mset (Pair x) (mset_set B). x \# mset_set (insert x F)#}" using insert(1,2) by (intro_cong "[\\<^sub>1 concat_mset, \\<^sub>2 image_mset]") auto finally show ?case by simp qed -lemma sum_mset_repeat: +lemma sum_mset_repeat: fixes f :: "'a \ 'b :: {comm_monoid_add,semiring_1}" shows "sum_mset (image_mset f (repeat_mset n A)) = of_nat n * sum_mset (image_mset f A)" by (induction n, auto simp add:sum_mset.distrib algebra_simps) unbundle no_intro_cong_syntax end \ No newline at end of file diff --git a/thys/Expander_Graphs/Expander_Graphs_Power_Construction.thy b/thys/Expander_Graphs/Expander_Graphs_Power_Construction.thy --- a/thys/Expander_Graphs/Expander_Graphs_Power_Construction.thy +++ b/thys/Expander_Graphs/Expander_Graphs_Power_Construction.thy @@ -1,566 +1,566 @@ section \Graph Powers\label{sec:graph_power}\ theory Expander_Graphs_Power_Construction - imports + imports Expander_Graphs_Walks Graph_Theory.Arc_Walk begin unbundle intro_cong_syntax fun is_arc_walk :: "('a, 'b) pre_digraph \ 'a \ 'b list \ bool" - where + where "is_arc_walk G _ [] = True" | "is_arc_walk G y (x#xs) = (is_arc_walk G (head G x) xs \ tail G x = y \ x \ arcs G)" -definition arc_walk_head :: "('a, 'b) pre_digraph \ ('a \ 'b list) \ 'a" +definition arc_walk_head :: "('a, 'b) pre_digraph \ ('a \ 'b list) \ 'a" where "arc_walk_head G x = (if snd x = [] then fst x else head G (last (snd x)))" lemma is_arc_walk_snoc: "is_arc_walk G y (xs@[x]) \ is_arc_walk G y xs \ x \ out_arcs G (arc_walk_head G (y,xs))" by (induction xs arbitrary: y, simp_all add:ac_simps arc_walk_head_def) lemma is_arc_walk_set: - assumes "is_arc_walk G u w" + assumes "is_arc_walk G u w" shows "set w \ arcs G" using assms by (induction w arbitrary: u, auto) lemma (in wf_digraph) awalk_is_arc_walk: assumes "u \ verts G" shows "is_arc_walk G u w \ awalk u w (awlast u w)" using assms unfolding awalk_def by (induction w arbitrary: u, auto) definition arc_walks :: "('a, 'b) pre_digraph \ nat \ ('a \ 'b list) set" where "arc_walks G l = {(u,w). u \ verts G \ is_arc_walk G u w \ length w = l}" lemma arc_walks_len: - assumes "x \ arc_walks G l" + assumes "x \ arc_walks G l" shows "length (snd x) = l" using assms unfolding arc_walks_def by auto lemma (in wf_digraph) awhd_of_arc_walk: assumes "w \ arc_walks G l" shows "awhd (fst w) (snd w) = fst w" - using assms unfolding arc_walks_def awalk_verts_def + using assms unfolding arc_walks_def awalk_verts_def by (cases "snd w", auto) lemma (in wf_digraph) awlast_of_arc_walk: assumes "w \ arc_walks G l" shows "awlast (fst w) (snd w) = arc_walk_head G w" unfolding awalk_verts_conv arc_walk_head_def by simp lemma (in wf_digraph) arc_walk_head_wellformed: assumes "w \ arc_walks G l" shows "arc_walk_head G w \ verts G" proof (cases "snd w = []") case True - then show ?thesis + then show ?thesis using assms unfolding arc_walks_def arc_walk_head_def by auto next case False have 0:"is_arc_walk G (fst w) (snd w)" using assms unfolding arc_walks_def by auto have "last (snd w) \ set (snd w)" using False last_in_set by auto also have "... \ arcs G" by (intro is_arc_walk_set[OF 0]) finally have "last (snd w) \ arcs G" by simp thus ?thesis unfolding arc_walk_head_def using False by simp qed lemma (in wf_digraph) arc_walk_tail_wellformed: assumes "w \ arc_walks G l" shows "fst w \ verts G" using assms unfolding arc_walks_def by auto -lemma (in fin_digraph) arc_walks_fin: +lemma (in fin_digraph) arc_walks_fin: "finite (arc_walks G l)" proof - have 0:"finite (verts G \ {w. set w \ arcs G \ length w = l})" by (intro finite_cartesian_product finite_lists_length_eq) auto show "finite (arc_walks G l)" unfolding arc_walks_def using is_arc_walk_set[where G="G"] by (intro finite_subset[OF _ 0] subsetI) auto qed lemma (in wf_digraph) awalk_verts_unfold: assumes "w \ arc_walks G l" shows "awalk_verts (fst w) (snd w) = fst w#map (head G) (snd w)" (is "?L = ?R") proof - obtain u v where w_def: "w = (u,v)" by fastforce have "awalk u v (awlast u v)" using assms unfolding w_def arc_walks_def by (intro iffD1[OF awalk_is_arc_walk]) auto hence cas: "cas u v (awlast u v)" unfolding awalk_def by simp have 0: "tail G (hd v) = u" if "v \ []" using cas that by (cases v) auto have "?L = awalk_verts u v" unfolding w_def by simp also have "... = (if v = [] then [u] else tail G (hd v) # map (head G) v)" by (intro awalk_verts_conv'[OF cas]) also have "... = u# map (head G) v" using 0 by simp also have "... = ?R" unfolding w_def by simp finally show ?thesis by simp qed -lemma (in fin_digraph) arc_walks_map_walks': +lemma (in fin_digraph) arc_walks_map_walks': "walks' G l = image_mset (case_prod awalk_verts) (mset_set (arc_walks G l))" proof (induction l) case 0 let ?g = "\x. fst x#map (head G) (snd x)" have "walks' G 0 = {#[x]. x \# mset_set (verts G)#}" by simp also have "... = image_mset ?g (image_mset (\x. (x,[])) (mset_set (verts G)))" unfolding image_mset.compositionality by (simp add:comp_def) also have "... = image_mset ?g (mset_set ((\x. (x,[])) ` verts G))" by (intro arg_cong2[where f="image_mset"] image_mset_mset_set inj_onI) auto also have "... = image_mset ?g (mset_set ({(u, w). u \ verts G \ w = []}))" by (intro_cong "[\\<^sub>2 image_mset]") auto also have "... = image_mset ?g (mset_set (arc_walks G 0))" unfolding arc_walks_def by (intro_cong "[\\<^sub>2 image_mset,\\<^sub>1 mset_set]") auto also have "... = image_mset (case_prod awalk_verts) (mset_set (arc_walks G 0))" using arc_walks_fin by (intro image_mset_cong) (simp add:case_prod_beta awalk_verts_unfold) finally show ?case by simp next case (Suc l) let ?f = "\(u,w) a. (u,w@[a])" let ?g = "\x. fst x#map (head G) (snd x)" have "arc_walks G (l+1) = case_prod ?f ` {(x,y). ?f x y \ arc_walks G (l+1)}" using arc_walks_len[where G="G" and l="Suc l", THEN iffD1[OF length_Suc_conv_rev]] by force also have "... = case_prod ?f ` {(x,y). x \ arc_walks G l \ y \ out_arcs G (arc_walk_head G x)}" unfolding arc_walks_def using is_arc_walk_snoc[where G="G"] by (intro_cong "[\\<^sub>2 image]") auto also have "... = (\w \ arc_walks G l. ?f w ` out_arcs G (arc_walk_head G w))" by (auto simp add:image_iff) finally have 0:"arc_walks G (l+1) = (\w \ arc_walks G l. ?f w ` out_arcs G (arc_walk_head G w))" by simp - have "mset_set (arc_walks G (l+1)) = concat_mset (image_mset (mset_set \ + have "mset_set (arc_walks G (l+1)) = concat_mset (image_mset (mset_set \ (\w. ?f w ` out_arcs G (arc_walk_head G w))) (mset_set (arc_walks G l)))" - unfolding 0 by (intro concat_disjoint_union_mset arc_walks_fin finite_imageI) auto - also have "... = concat_mset {# mset_set (?f x ` out_arcs G (arc_walk_head G x)). + unfolding 0 by (intro concat_disjoint_union_mset arc_walks_fin finite_imageI) auto + also have "... = concat_mset {# mset_set (?f x ` out_arcs G (arc_walk_head G x)). x\#mset_set(arc_walks G l)#}" - by (simp add:comp_def case_prod_beta) - also have "... = concat_mset {# {# ?f x y. y \# mset_set (out_arcs G (arc_walk_head G x))#}. + by (simp add:comp_def case_prod_beta) + also have "... = concat_mset {# {# ?f x y. y \# mset_set (out_arcs G (arc_walk_head G x))#}. x \# mset_set (arc_walks G l)#}" by (intro_cong "[\\<^sub>1 concat_mset]" more:image_mset_cong image_mset_mset_set[symmetric] inj_onI) auto - finally have 1:"mset_set (arc_walks G (l+1)) = concat_mset + finally have 1:"mset_set (arc_walks G (l+1)) = concat_mset {# {# ?f x y. y \# mset_set (out_arcs G (arc_walk_head G x))#}. x \# mset_set (arc_walks G l)#}" by simp have "walks' G (l+1)=concat_mset {#{#w @ [z]. z\# vertices_from G (last w)#}. w \# walks' G l#}" by simp also have "... = concat_mset {# - {#awalk_verts (fst x) (snd x) @ [z]. z \# vertices_from G (awlast (fst x) (snd x))#}. + {#awalk_verts (fst x) (snd x) @ [z]. z \# vertices_from G (awlast (fst x) (snd x))#}. x \# mset_set (arc_walks G l)#}" unfolding Suc by (simp add:image_mset.compositionality comp_def case_prod_beta) also have "... = concat_mset {# - {#?g x @ [z]. z \# vertices_from G (awlast (fst x) (snd x))#}. + {#?g x @ [z]. z \# vertices_from G (awlast (fst x) (snd x))#}. x \# mset_set (arc_walks G l)#}" - using arc_walks_fin + using arc_walks_fin by (intro_cong "[\\<^sub>1 concat_mset]" more:image_mset_cong) (auto simp: awalk_verts_unfold) - also have "... = concat_mset {# {#?g x @ [z]. z \# vertices_from G (arc_walk_head G x)#}. + also have "... = concat_mset {# {#?g x @ [z]. z \# vertices_from G (arc_walk_head G x)#}. x \# mset_set (arc_walks G l)#}" using arc_walks_fin awlast_of_arc_walk by (intro_cong "[\\<^sub>1 concat_mset, \\<^sub>2 image_mset]" more: image_mset_cong) auto - also have "... = (concat_mset {# {# ?g (fst x, snd x@[y]). + also have "... = (concat_mset {# {# ?g (fst x, snd x@[y]). y \# mset_set (out_arcs G (arc_walk_head G x))#}. x \# mset_set (arc_walks G l)#})" unfolding verts_from_alt by (simp add:image_mset.compositionality comp_def) - also have "... = image_mset ?g (concat_mset {# {# ?f x y. + also have "... = image_mset ?g (concat_mset {# {# ?f x y. y \# mset_set (out_arcs G (arc_walk_head G x))#}. x \# mset_set (arc_walks G l)#})" unfolding image_concat_mset - by (auto simp add:comp_def case_prod_beta image_mset.compositionality) + by (auto simp add:comp_def case_prod_beta image_mset.compositionality) also have "... = image_mset ?g (mset_set (arc_walks G (l+1)))" unfolding 1 by simp also have "... = image_mset (case_prod awalk_verts) (mset_set (arc_walks G (l+1)))" using arc_walks_fin by (intro image_mset_cong) (simp add:case_prod_beta awalk_verts_unfold) finally show ?case by simp qed -lemma (in fin_digraph) arc_walks_map_walks: +lemma (in fin_digraph) arc_walks_map_walks: "walks G (l+1) = image_mset (case_prod awalk_verts) (mset_set (arc_walks G l))" using arc_walks_map_walks' unfolding walks_def by simp lemma (in wf_digraph) assumes "awalk u a v " "length a = l" "l > 0" shows awalk_ends: "tail G (hd a) = u" "head G (last a) = v" proof - - have 0:"cas u a v" + have 0:"cas u a v" using assms unfolding awalk_def by simp have 1: "a \ []" using assms(2,3) by auto show "tail G (hd a) = u" using 0 unfolding cas_simp[OF 1] by auto show "head G (last a) = v" - using 1 0 by (induction a arbitrary:u rule:list_nonempty_induct) auto + using 1 0 by (induction a arbitrary:u rule:list_nonempty_induct) auto qed definition graph_power :: "('a, 'b) pre_digraph \ nat \ ('a, ('a \ 'b list)) pre_digraph" where "graph_power G l = \ verts = verts G, arcs = arc_walks G l, tail = fst, head = arc_walk_head G \" -lemma (in wf_digraph) graph_power_wf: +lemma (in wf_digraph) graph_power_wf: "wf_digraph (graph_power G l)" proof - have "tail (graph_power G l) a \ verts (graph_power G l)" - "head (graph_power G l) a \ verts (graph_power G l)" + "head (graph_power G l) a \ verts (graph_power G l)" if "a \ arcs (graph_power G l)" for a using that arc_walk_head_wellformed arc_walk_tail_wellformed unfolding graph_power_def by simp_all - thus ?thesis + thus ?thesis unfolding wf_digraph_def by auto qed -lemma (in fin_digraph) graph_power_fin: +lemma (in fin_digraph) graph_power_fin: "fin_digraph (graph_power G l)" proof - - interpret H:wf_digraph "graph_power G l" + interpret H:wf_digraph "graph_power G l" using graph_power_wf by auto - have "finite (arcs (graph_power G l))" + have "finite (arcs (graph_power G l))" using arc_walks_fin unfolding graph_power_def by simp - moreover have "finite (verts (graph_power G l))" + moreover have "finite (verts (graph_power G l))" unfolding graph_power_def by simp ultimately show ?thesis by unfold_locales auto qed lemma (in fin_digraph) graph_power_count_edges: fixes l v w defines "S \ {x. length x=l+1\set x\verts G\hd x=v\last x=w}" shows "count (edges (graph_power G l)) (v,w) = (\x \ S.(\i verts G \ length x = l+1}" by (intro finite_lists_length_eq) auto have fin_S: "finite S" unfolding S_def by (intro finite_subset[OF _ 0]) auto have "?L = size {#x \# mset_set (arc_walks G l). fst x = v \ arc_walk_head G x = w#}" - unfolding graph_power_def edges_def arc_to_ends_def + unfolding graph_power_def edges_def arc_to_ends_def by (simp add:count_mset_exp image_mset_filter_mset_swap[symmetric]) - also have "... = size + also have "... = size {#x \# mset_set (arc_walks G l). awhd (fst x) (snd x) = v \ awlast (fst x) (snd x) = w#}" using awlast_of_arc_walk awhd_of_arc_walk arc_walks_fin by (intro arg_cong[where f="size"] filter_mset_cong refl) simp also have "... = size {#x \# walks G (l+1). hd x = v \ last x = w#}" - unfolding arc_walks_map_walks + unfolding arc_walks_map_walks by (simp add:image_mset_filter_mset_swap[symmetric] case_prod_beta) also have "...=size{#x\# walks G (l+1).x \ S#}" unfolding S_def using set_walks_3 by (intro arg_cong[where f="size"] filter_mset_cong refl) auto also have "...=sum (count (walks G (l+1))) S" - by (intro sum_count_2[symmetric] fin_S) + by (intro sum_count_2[symmetric] fin_S) also have "...=(\x\S.(\i verts (graph_power G l)" "w \ verts (graph_power G l)" - shows "card (arcs_betw (graph_power G l) v w) = card (arcs_betw (graph_power G l) w v)" + shows "card (arcs_betw (graph_power G l) v w) = card (arcs_betw (graph_power G l) w v)" (is "?L = ?R") proof - - interpret H:fin_digraph "graph_power G l" + interpret H:fin_digraph "graph_power G l" using graph_power_fin by auto define S where "S v w = {x. length x=l+1 \ set x \ verts G \ hd x = v \ last x = w}" for v w have 0: "bij_betw rev (S w v) (S v w)" - unfolding S_def by (intro bij_betwI[where g="rev"]) (auto simp add:hd_rev last_rev) + unfolding S_def by (intro bij_betwI[where g="rev"]) (auto simp add:hd_rev last_rev) have 1: "bij_betw ((-) (l - 1)) {..x \ S v w. (\ix \ S w v. (\ix \ S w v. (\ix \ S w v. (\ix \ S w v. (\iv. v \ verts G \ out_degree G v = d" assumes "v \ verts (graph_power G l)" shows "out_degree (graph_power G l) v = d ^ l" (is "?L = ?R") proof - - interpret H:fin_digraph "graph_power G l" + interpret H:fin_digraph "graph_power G l" using graph_power_fin by auto have v_vert: "v \ verts G" using assms unfolding graph_power_def by simp have "?L = size (vertices_from (graph_power G l) v)" unfolding out_degree_def H.verts_from_alt by simp also have "... = size ({# e \# edges (graph_power G l). fst e = v #})" unfolding vertices_from_def by simp also have "... = size {#w \# mset_set (arc_walks G l). fst w = v#}" - unfolding graph_power_def edges_def arc_to_ends_def + unfolding graph_power_def edges_def arc_to_ends_def by (simp add:count_mset_exp image_mset_filter_mset_swap[symmetric]) also have "... = size {#w \# mset_set (arc_walks G l). awhd (fst w) (snd w) = v#}" using awlast_of_arc_walk awhd_of_arc_walk arc_walks_fin by (intro arg_cong[where f="size"] filter_mset_cong refl) simp also have "... = size {#x \# walks' G l. hd x = v #}" - unfolding arc_walks_map_walks' + unfolding arc_walks_map_walks' by (simp add:image_mset_filter_mset_swap[symmetric] case_prod_beta) also have "... = d^l" proof (induction l) case 0 have "size {#x \# walks' G 0. hd x = v#} = card {x. x = v \ x \ verts G}" by (simp add:image_mset_filter_mset_swap[symmetric]) also have "... = card {v}" using v_vert by (intro arg_cong[where f="card"]) auto also have "... = d^0" by simp finally show ?case by simp next case (Suc l) have "size {#x \# walks' G (Suc l). hd x = v#} = (\x\#walks' G l. size {#y \# vertices_from G (last x). hd (x @ [y]) = v#})" - by (simp add:size_concat_mset image_mset_filter_mset_swap[symmetric] + by (simp add:size_concat_mset image_mset_filter_mset_swap[symmetric] filter_concat_mset image_mset.compositionality comp_def) also have "... = (\x\#walks' G l. size {#y \# vertices_from G (last x). hd x = v#})" - using set_walks_2 + using set_walks_2 by (intro_cong "[\\<^sub>1 sum_mset, \\<^sub>1 size]" more:image_mset_cong filter_mset_cong) auto also have "... = (\x\#walks' G l. (if hd x = v then out_degree G (last x) else 0))" unfolding verts_from_alt out_degree_def by (simp add:filter_mset_const if_distribR if_distrib cong:if_cong) also have "... = (\x\#walks' G l. d * of_bool (hd x = v))" using set_walks_2[where l="l"] last_in_set by (intro arg_cong[where f="sum_mset"] image_mset_cong) (auto intro!:reg) - also have "... = d * (\x\#walks' G l. of_bool (hd x = v))" + also have "... = d * (\x\#walks' G l. of_bool (hd x = v))" by (simp add:sum_mset_distrib_left image_mset.compositionality comp_def) also have "... = d * (size {#x \# walks' G l. hd x = v#})" by (simp add:size_filter_mset_conv) also have "... = d * d ^ l" using Suc by simp also have "... = d^Suc l" by simp finally show ?case by simp - qed + qed finally show ?thesis by simp qed lemma (in regular_graph) graph_power_out_degree: assumes "v \ verts (graph_power G l)" shows "out_degree (graph_power G l) v = d ^ l" (is "?L = ?R") by (intro graph_power_out_degree' assms reg) auto -lemma (in regular_graph) graph_power_regular: +lemma (in regular_graph) graph_power_regular: "regular_graph (graph_power G l)" proof - - interpret H:fin_digraph "graph_power G l" + interpret H:fin_digraph "graph_power G l" using graph_power_fin by auto have "verts (graph_power G l) \ {}" using verts_non_empty unfolding graph_power_def by simp moreover have "0 < d^l" using d_gt_0 by simp ultimately show ?thesis using graph_power_out_degree by (intro regular_graphI[where d="d^l"] graph_power_sym sym) qed -lemma (in regular_graph) graph_power_degree: +lemma (in regular_graph) graph_power_degree: "regular_graph.d (graph_power G l) = d^l" (is "?L = ?R") proof - - interpret H:regular_graph "graph_power G l" + interpret H:regular_graph "graph_power G l" using graph_power_regular by auto obtain v where v_set: "v \ verts (graph_power G l)" using H.verts_non_empty by auto - hence "?L = out_degree (graph_power G l) v" + hence "?L = out_degree (graph_power G l) v" using v_set H.reg by auto also have "... =?R" - by (intro graph_power_out_degree[OF v_set]) + by (intro graph_power_out_degree[OF v_set]) finally show ?thesis by simp qed lemma (in regular_graph) graph_power_step: assumes "x \ verts G" shows "regular_graph.g_step (graph_power G l) f x = (g_step^^l) f x" using assms proof (induction l arbitrary: x) case 0 let ?H = "graph_power G 0" interpret H:regular_graph "?H" using graph_power_regular by auto have "regular_graph.g_step (graph_power G 0) f x = H.g_step f x" by simp have "H.g_step f x = (\x\in_arcs ?H x. f (tail ?H x))" unfolding H.g_step_def graph_power_degree by simp also have "... = (\v\{e \ arc_walks G 0. arc_walk_head G e = x}. f (fst v))" unfolding in_arcs_def graph_power_def by (simp add:case_prod_beta) also have "... = (\v\{x}. f v)" unfolding arc_walks_def using 0 by (intro sum.reindex_bij_betw bij_betwI[where g="(\x. (x,[]))"]) (auto simp add:arc_walk_head_def) also have "... = f x" by simp also have "... = (g_step^^0) f x" by simp finally show ?case by simp next case (Suc l) let ?H = "graph_power G l" interpret H:regular_graph "?H" using graph_power_regular by auto let ?HS = "graph_power G (l+1)" interpret HS:regular_graph "?HS" using graph_power_regular by auto let ?bij = "(\(x,(y1,y2)). (y1,y2@[x]))" let ?bijr = "(\(y1,y2). (last y2, (y1,butlast y2)))" define S where "S = {y. fst y \ in_arcs G x \ snd y \ in_arcs ?H (tail G (fst y))}" have "S = {(u,v). u \ arcs G \ head G u = x \ v \ arc_walks G l \ arc_walk_head G v = tail G u}" unfolding S_def graph_power_def in_arcs_def by auto also have "... = {(u,v). (fst v,snd v@[u]) \ arc_walks G (l+1) \ arc_walk_head G (fst v,snd v@[u]) = x}" - unfolding arc_walks_def by (intro iffD2[OF set_eq_iff] allI) + unfolding arc_walks_def by (intro iffD2[OF set_eq_iff] allI) (auto simp add: is_arc_walk_snoc case_prod_beta arc_walk_head_def) also have "... = {(u,v). (fst v,snd v@[u]) \ in_arcs ?HS x}" unfolding in_arcs_def graph_power_def by auto finally have S_alt: "S = {(u,v). (fst v,snd v@[u]) \ in_arcs ?HS x}" by simp have len_in_arcs: "a \ in_arcs ?HS x \ snd a \ []" for a unfolding in_arcs_def graph_power_def arc_walks_def by auto have 0:"bij_betw ?bij S (in_arcs ?HS x)" unfolding S_alt using len_in_arcs by (intro bij_betwI[where g="?bijr"]) auto have "HS.g_step f x = (\y\in_arcs ?HS x. f (tail ?HS y)/ d^(l+1))" unfolding HS.g_step_def graph_power_degree by simp also have "... = (\y\in_arcs ?HS x. f (fst y)/ d^(l+1))" unfolding graph_power_def by simp - also have "... = (\y \ S. f (fst (?bij y))/ d^(l+1))" + also have "... = (\y \ S. f (fst (?bij y))/ d^(l+1))" by (intro sum.reindex_bij_betw[symmetric] 0) - also have "... = (\y \ S. f (fst (snd y))/ d^(l+1))" + also have "... = (\y \ S. f (fst (snd y))/ d^(l+1))" by (intro_cong "[\\<^sub>2 (/),\\<^sub>1 f]" more: sum.cong) (simp add:case_prod_beta) - also have "...=(\y\(\a\in_arcs G x. (Pair a)`in_arcs ?H (tail G a)). f (fst (snd y))/ d^(l+1))" + also have "...=(\y\(\a\in_arcs G x. (Pair a)`in_arcs ?H (tail G a)). f (fst (snd y))/ d^(l+1))" unfolding S_def by (intro sum.cong) auto also have "...=(\a\in_arcs G x. (\y\(Pair a)`in_arcs ?H (tail G a). f (fst (snd y))/ d^(l+1)))" by (intro sum.UNION_disjoint) auto - also have "... = (\a \ in_arcs G x. (\b \ in_arcs ?H (tail G a). f (fst b) / d^(l+1)))" + also have "... = (\a \ in_arcs G x. (\b \ in_arcs ?H (tail G a). f (fst b) / d^(l+1)))" by (intro sum.cong sum.reindex_bij_betw) (auto simp add:bij_betw_def inj_on_def image_iff) - also have "... = (\a \ in_arcs G x. (\b \ in_arcs ?H (tail G a). f (tail ?H b) / d^l)/d)" + also have "... = (\a \ in_arcs G x. (\b \ in_arcs ?H (tail G a). f (tail ?H b) / d^l)/d)" unfolding graph_power_def by (simp add:sum_divide_distrib algebra_simps) - also have "... = (\a \ in_arcs G x. H.g_step f (tail G a)/ d)" + also have "... = (\a \ in_arcs G x. H.g_step f (tail G a)/ d)" unfolding H.g_step_def graph_power_degree by simp also have "... = (\a \ in_arcs G x. (g_step^^l) f (tail G a)/ d)" by (intro sum.cong refl arg_cong2[where f="(/)"] Suc) auto also have "... = g_step ((g_step^^l) f) x" unfolding g_step_def by simp also have "... = (g_step^^(l+1)) f x" by simp finally show ?case by simp qed lemma (in regular_graph) graph_power_expansion: "regular_graph.\\<^sub>a (graph_power G l) \ \\<^sub>a^l" proof - - interpret H:regular_graph "graph_power G l" + interpret H:regular_graph "graph_power G l" using graph_power_regular by auto have "\H.g_inner f (H.g_step f)\ \ \\<^sub>a ^ l * (H.g_norm f)\<^sup>2" (is "?L \ ?R") if "H.g_inner f (\_. 1) = 0" for f proof - have "g_inner f (\_. 1) = H.g_inner f (\_.1)" unfolding g_inner_def H.g_inner_def by (intro sum.cong) (auto simp add:graph_power_def) also have "... = 0" using that by simp finally have 1:"g_inner f (\_. 1) = 0" by simp have 2: "g_inner ((g_step^^l) f) (\_. 1) = 0" for l using g_step_remains_orth 1 by (induction l, auto) have 0: "g_norm ((g_step^^l) f) \ \\<^sub>a ^ l * g_norm f" proof (induction l) case 0 then show ?case by simp next case (Suc l) have "g_norm ((g_step ^^ Suc l) f) = g_norm (g_step ((g_step ^^ l) f))" by simp also have "... \ \\<^sub>a * g_norm (((g_step ^^ l) f))" by (intro expansionD2 2) also have "... \ \\<^sub>a * (\\<^sub>a^l * g_norm f)" by (intro mult_left_mono \_ge_0 Suc) also have "... = \\<^sub>a^(l+1) * g_norm f" by simp finally show ?case by simp qed have "?L = \g_inner f (H.g_step f)\" - unfolding H.g_inner_def g_inner_def + unfolding H.g_inner_def g_inner_def by (intro_cong "[\\<^sub>1 abs]" more:sum.cong) (auto simp add:graph_power_def) also have "... = \g_inner f ((g_step^^l) f)\" by (intro_cong "[\\<^sub>1 abs]" more:g_inner_cong graph_power_step) auto also have "... \ g_norm f * g_norm ((g_step^^l) f)" by (intro g_inner_cauchy_schwartz) also have "... \ g_norm f * (\\<^sub>a ^ l * g_norm f)" - by (intro mult_left_mono 0 g_norm_nonneg) + by (intro mult_left_mono 0 g_norm_nonneg) also have "... = \\<^sub>a ^ l * g_norm f^2" by (simp add:power2_eq_square) - also have "... = ?R" + also have "... = ?R" unfolding g_norm_sq H.g_norm_sq g_inner_def H.g_inner_def by (intro_cong "[\\<^sub>2 (*)]" more:sum.cong) (auto simp add:graph_power_def) finally show ?thesis by simp qed - moreover have " 0 \ \\<^sub>a ^ l" + moreover have " 0 \ \\<^sub>a ^ l" using \_ge_0 by simp ultimately show ?thesis - by (intro H.expander_intro_1) auto + by (intro H.expander_intro_1) auto qed unbundle no_intro_cong_syntax end \ No newline at end of file diff --git a/thys/Expander_Graphs/Expander_Graphs_Strongly_Explicit.thy b/thys/Expander_Graphs/Expander_Graphs_Strongly_Explicit.thy --- a/thys/Expander_Graphs/Expander_Graphs_Strongly_Explicit.thy +++ b/thys/Expander_Graphs/Expander_Graphs_Strongly_Explicit.thy @@ -1,1079 +1,1079 @@ section \Strongly Explicit Expander Graphs\label{sec:see}\ text \In some applications, representing an expander graph using a data structure (for example -as an adjacency lists) would be prohibitive. For such cases strongly explicit expander graphs (SEE) -are relevant. These are expander graphs, which can be represented implicitly using a function that -computes for each vertex its neighbors in space and time logarithmic w.r.t. to the size of the -graph. An application can for example sample a random walk, from a SEE using such a function -efficiently. An example of such a graph is the Margulis construction from +as an adjacency lists) would be prohibitive. For such cases strongly explicit expander graphs (SEE) +are relevant. These are expander graphs, which can be represented implicitly using a function that +computes for each vertex its neighbors in space and time logarithmic w.r.t. to the size of the +graph. An application can for example sample a random walk, from a SEE using such a function +efficiently. An example of such a graph is the Margulis construction from Section~\ref{sec:margulis}. This section presents the latter as a SEE but also shows that two graph -operations that preserve the SEE property, in particular the graph power construction from -Section~\ref{sec:graph_power} and a compression scheme introduced by -Murtagh et al.~\cite[Theorem~20]{murtagh2019}. Combining all of the above it is possible to construct -strongly explicit expander graphs of \emph{every size} and spectral gap, which is formalized in +operations that preserve the SEE property, in particular the graph power construction from +Section~\ref{sec:graph_power} and a compression scheme introduced by +Murtagh et al.~\cite[Theorem~20]{murtagh2019}. Combining all of the above it is possible to construct +strongly explicit expander graphs of \emph{every size} and spectral gap, which is formalized in Subsection~\ref{sec:see_standard}.\ theory Expander_Graphs_Strongly_Explicit imports Expander_Graphs_Power_Construction Expander_Graphs_MGG begin unbundle intro_cong_syntax no_notation Digraph.dominates ("_ \\ _" [100,100] 40) record strongly_explicit_expander = see_size :: nat see_degree :: nat see_step :: "nat \ nat \ nat" definition graph_of :: "strongly_explicit_expander \ (nat, (nat,nat) arc) pre_digraph" - where "graph_of e = - \ verts = {.. verts = {..(v, i). Arc v (see_step e i v) i) ` ({.. {.." definition "is_expander e \\<^sub>a \ regular_graph (graph_of e) \ regular_graph.\\<^sub>a (graph_of e) \ \\<^sub>a" lemma is_expander_mono: assumes "is_expander e a" "a \ b" shows "is_expander e b" using assms unfolding is_expander_def by auto lemma graph_of_finI: assumes "see_step e \ ({.. ({.. {.. verts ?G \ tail ?G a \ verts ?G" if "a \ arcs ?G" for a + have "head ?G a \ verts ?G \ tail ?G a \ verts ?G" if "a \ arcs ?G" for a using assms that unfolding graph_of_def by (auto simp add:Pi_def) hence 0: "wf_digraph ?G" unfolding wf_digraph_def by auto - have 1: "finite (verts ?G)" + have 1: "finite (verts ?G)" unfolding graph_of_def by simp - have 2: "finite (arcs ?G)" + have 2: "finite (arcs ?G)" unfolding graph_of_def by simp show ?thesis using 0 1 2 unfolding fin_digraph_def fin_digraph_axioms_def by auto qed lemma edges_graph_of: "edges(graph_of e)={#(v,see_step e i v). (v,i)\#mset_set ({..{..(v, i). Arc v (see_step e i v) i) ` ({.. {..# mset_set ( {.. {..# mset_set ({.. {..# mset_set ({.. {.. verts (graph_of e)" shows "out_degree (graph_of e) v = see_degree e" (is "?L = ?R") proof - let ?d = "see_degree e" let ?n = "see_size e" - have 0: "v < ?n" + have 0: "v < ?n" using assms unfolding graph_of_def by simp have "?L = card {a. (\x\{..y\{.. arc_tail a = v}" unfolding out_degree_def out_arcs_def graph_of_def by (simp add:image_iff) also have "... = card {a. (\y\{..y. Arc v (see_step e y v) y) ` {..v \ verts ?G. {x. fst x = v \ is_arc_walk ?G v (snd x) \ length (snd x) = n})" unfolding arc_walks_def by (intro arg_cong[where f="card"]) auto also have "... = (\v \ verts ?G. card {x. fst x=v\is_arc_walk ?G v (snd x)\length (snd x) = n})" using is_arc_walk_set[where G="?G"] by (intro card_UN_disjoint ballI finite_cartesian_product subsetI finite_lists_length_eq finite_subset[where B="verts ?G \ {x. set x \ arcs ?G \ length x = n}"]) force+ also have "... = (\v \ verts ?G. out_degree (graph_power ?G n) v)" - unfolding out_degree_def graph_power_def out_arcs_def arc_walks_def + unfolding out_degree_def graph_power_def out_arcs_def arc_walks_def by (intro sum.cong arg_cong[where f="card"]) auto also have "... = (\v \ verts ?G. see_degree e^n)" by (intro sum.cong graph_power_out_degree' out_degree_see refl) (simp_all add: graph_power_def) also have "... = ?R" by (simp add:graph_of_def) finally show ?thesis by simp qed lemma regular_graph_degree_eq_see_degree: assumes "regular_graph (graph_of e)" shows "regular_graph.d (graph_of e) = see_degree e" (is "?L = ?R") proof - interpret regular_graph "graph_of e" using assms(1) by simp obtain v where v_set: "v \ verts (graph_of e)" using verts_non_empty by auto - hence "?L = out_degree (graph_of e) v" + hence "?L = out_degree (graph_of e) v" using v_set reg by auto also have "... = see_degree e" by (intro out_degree_see v_set) finally show ?thesis by simp qed text \The following introduces the compression scheme, described in \cite[Theorem 20]{murtagh2019}.\ fun see_compress :: "nat \ strongly_explicit_expander \ strongly_explicit_expander" - where "see_compress m e = + where "see_compress m e = \ see_size = m, see_degree = see_degree e * 2 - , see_step = (\k v. - if k < see_degree e - then (see_step e k v) mod m + , see_step = (\k v. + if k < see_degree e + then (see_step e k v) mod m else (if v+m < see_size e then (see_step e (k-see_degree e) (v+m)) mod m else v)) \" lemma edges_of_compress: fixes e m assumes "2*m \ see_size e" "m \ see_size e" defines "A \ {# (x mod m, y mod m). (x,y) \# edges (graph_of e)#}" defines "B \ repeat_mset (see_degree e) {# (x,x). x \# (mset_set {see_size e - m.. v \ v < ?n \ v - m = v mod m" for v using assms by (simp add: le_mod_geq) let ?M = "mset_set ({..{..<2*?d})" define M1 where "M1 = mset_set ({.. {.. {?d..<2*?d})" define M3 where "M3 = mset_set ({?n-m.. {?d..<2*?d})" have "M2 = mset_set ((\(x,y). (x-m,y+?d)) ` ({m.. {..(x,y). (x-m,y+?d)) (mset_set ({m.. {..(x,y). (x-m,y+?d)) (mset_set ({m.. {..{.. {..{?d..<2*?d} \ {?n-m..{?d..<2*?d})" using assms(1,2) by (intro arg_cong[where f="mset_set"]) auto also have "... = mset_set ({..{.. {..{?d..<2*?d}) + M3" unfolding M3_def by (intro mset_set_Union) auto also have "... = M1 + M2 + M3" unfolding M1_def M2_def by (intro arg_cong2[where f="(+)"] mset_set_Union) auto finally have 0:"mset_set ({.. {..<2*?d}) = M1 + M2 + M3" by simp have 1:"{#(v,?c i v). (v,i)\#M1#}={#(v mod m,?s i v mod m). (v,i)\#mset_set ({..{..#M2#}={#(fst x-m,?c(snd x+?d)(fst x-m)).x\#mset_set({m..{..#mset_set ({m..{..#mset_set ({m..{..#M2#}={#(v mod m,?s i v mod m). (v,i)\#mset_set ({m..{..#M3#} = {#(v,v). (v,i) \# mset_set ({?n-m.. {?d..<2*?d})#}" unfolding M3_def by (intro image_mset_cong) auto also have "... = concat_mset {#{#(x, x). xa \# mset_set {?d..<2 * ?d}#}. x \# mset_set {?n - m..# mset_set {?n - m..#M3#}=B" by simp have "A = {#(fst x mod m, ?s (snd x) (fst x) mod m). x \# mset_set ({.. {..#mset_set({..{..#mset_set({..{..# ?M #}" unfolding edges_graph_of by (simp add:ac_simps) also have "... = {#(v,?c i v). (v,i)\#M1#}+{#(v,?c i v). (v,i)\#M2#}+{#(v,?c i v). (v,i)\#M3#}" unfolding 0 image_mset_union by simp also have "...={#(v mod m,?s i v mod m). (v,i)\#mset_set({..{..{m..{..\<^sub>2 (+), \\<^sub>2 image_mset]" more: mset_set_Union[symmetric]) auto also have "...={#(v mod m,?s i v mod m). (v,i)\#mset_set({..{..\<^sub>2 (+), \\<^sub>2 image_mset, \\<^sub>1 mset_set]") auto also have "... = A + B" unfolding 4 by simp finally show ?thesis by simp qed lemma see_compress_sym: assumes "2*m \ see_size e" "m \ see_size e" assumes "symmetric_multi_graph (graph_of e)" shows "symmetric_multi_graph (graph_of (see_compress m e))" proof - let ?c = "see_compress m e" let ?d = "see_degree e" let ?G = "graph_of e" let ?H = "graph_of (see_compress m e)" interpret G:"fin_digraph" "?G" - by (intro symmetric_multi_graphD2[OF assms(3)]) + by (intro symmetric_multi_graphD2[OF assms(3)]) interpret H:"fin_digraph" "?H" by (intro graph_of_finI) simp have deg_compres: "see_degree ?c = 2 * see_degree e" by simp - have 1: "card (arcs_betw ?H v w) = card (arcs_betw ?H w v)" (is "?L = ?R") - if "v \ verts ?H" "w \ verts ?H" for v w + have 1: "card (arcs_betw ?H v w) = card (arcs_betw ?H w v)" (is "?L = ?R") + if "v \ verts ?H" "w \ verts ?H" for v w proof - define b where "b =count {#(x, x). x \# mset_set {see_size e - m..# mset_set {see_size e - m..# edges (graph_of e)#} (v, w) + ?d * b" unfolding edges_of_compress[OF assms(1,2)] b_def by simp also have "... = count {#(snd e mod m, fst e mod m). e \# edges (graph_of e)#} (v, w) + ?d * b" by (subst G.edges_sym[OF assms(3),symmetric]) (simp add:image_mset.compositionality comp_def case_prod_beta) also have "... = count {#(x mod m, y mod m). (x,y) \# edges (graph_of e)#} (w, v) + ?d * b" - unfolding count_mset_exp + unfolding count_mset_exp by (simp add:image_mset_filter_mset_swap[symmetric] ac_simps case_prod_beta) also have "... = count (edges ?H) (w,v)" unfolding edges_of_compress[OF assms(1,2)] b_alt_def by simp - also have "... = ?R" + also have "... = ?R" unfolding H.count_edges by simp finally show ?thesis by simp qed show ?thesis using 1 H.fin_digraph_axioms unfolding symmetric_multi_graph_def by auto qed lemma see_compress: assumes "is_expander e \\<^sub>a" assumes "2*m \ see_size e" "m \ see_size e" shows "is_expander (see_compress m e) (\\<^sub>a/2 + 1/2)" proof - let ?H = "graph_of (see_compress m e)" let ?G = "graph_of e" let ?d = "see_degree e" let ?n = "see_size e" - interpret G:regular_graph "graph_of e" + interpret G:regular_graph "graph_of e" using assms(1) is_expander_def by simp have d_eq: "?d = G.d" using regular_graph_degree_eq_see_degree[OF G.regular_graph_axioms] by simp have n_eq: "G.n = ?n" unfolding G.n_def by (simp add:graph_of_def) have n_gt_1: "?n > 0" using G.n_gt_0 n_eq by auto have "symmetric_multi_graph (graph_of (see_compress m e))" by (intro see_compress_sym assms(2,3) G.sym) - moreover have "see_size e > 0" + moreover have "see_size e > 0" using G.verts_non_empty unfolding graph_of_def by auto hence "m > 0" using assms(2) by simp - hence "verts (graph_of (see_compress m e)) \ {}" + hence "verts (graph_of (see_compress m e)) \ {}" unfolding graph_of_def by auto - moreover have 1:"0 < see_degree e" + moreover have 1:"0 < see_degree e" using d_eq G.d_gt_0 by auto hence "0 < see_degree (see_compress m e)" by simp ultimately have 0:"regular_graph ?H" by (intro regular_graphI[where d="see_degree (see_compress m e)"] out_degree_see) auto interpret H:regular_graph ?H using 0 by auto - have "\\a\arcs ?H. f (head ?H a) * f (tail ?H a)\ \ (real G.d * G.\\<^sub>a + G.d) * (H.g_norm f)\<^sup>2" + have "\\a\arcs ?H. f (head ?H a) * f (tail ?H a)\ \ (real G.d * G.\\<^sub>a + G.d) * (H.g_norm f)\<^sup>2" (is "?L \ ?R") if "H.g_inner f (\_. 1) = 0" for f proof - define f' where "f' x = f (x mod m)" for x let ?L1 = "G.g_norm f'^2 + \\x=?n-m.." let ?L2 = "G.g_inner f' (\_.1)^2/ G.n + \\x=?n-m.." have "?L1 = (\x\x=?n-m.." unfolding G.g_norm_sq G.g_inner_def f'_def by (simp add:graph_of_def power2_eq_square) - also have "... = (\x\{0.. {m..x=?n-m..x\{0.. {m..x=?n-m..\<^sub>2 (+)]" more:sum.cong abs_of_nonneg sum_nonneg) auto also have "...=(\x=0..x=m..x=?n-m..\<^sub>2 (+)]" more:sum.union_disjoint) auto also have "... = (\x=0..x=0..x=?n-m..\<^sub>2 (+)]" more: sum.reindex_bij_betw bij_betwI[where g="(\x. x+m)"]) (auto simp add:le_mod_geq) also have "... = (\x=0..x=0..x=?n-m..x=0..x=0..x=?n-m..x=0..x\{0..{?n-m..xxx\{..{m..x=?n-m..\<^sub>2 (+), \\<^sub>2 (/), \\<^sub>2 (power)]" more: sum.cong abs_of_nonneg sum_nonneg) (auto simp add:graph_of_def) also have "...=((\xx=m..x=?n-m..\<^sub>2 (+), \\<^sub>2 (/), \\<^sub>2 (power)]" more:sum.union_disjoint) auto also have "...=((\xx=0..x=?n-m..\<^sub>2 (+), \\<^sub>2 (/), \\<^sub>2 (power)]" + using assms(2,3) by (intro_cong "[\\<^sub>2 (+), \\<^sub>2 (/), \\<^sub>2 (power)]" more:sum.reindex_bij_betw bij_betwI[where g="(\x. x+m)"]) (auto simp add:le_mod_geq) also have "...=(H.g_inner f (\_. 1) +(\xx=?n-m..\<^sub>2 (+), \\<^sub>2 (/), \\<^sub>2 (power)]" more: sum.cong) (auto simp:graph_of_def) + by (intro_cong "[\\<^sub>2 (+), \\<^sub>2 (/), \\<^sub>2 (power)]" more: sum.cong) (auto simp:graph_of_def) also have "...=(\xx=?n-m.. (\xf x\ * \1\)^2/G.n + (\x=?n-m.. (L2_set f {.._. 1) {..x=?n-m..x x=?n-m..\<^sub>2 (+), \\<^sub>2 (/),\\<^sub>2 (*)]" more:real_sqrt_pow2 sum_nonneg) auto also have "... = (\x x=?n-m.. (\x x=?n-m..x\{..{?n-m.. H.g_norm f^2" by simp have "?L = \\(u, v)\#edges ?H. f v * f u\" - unfolding edges_def arc_to_ends_def sum_unfold_sum_mset + unfolding edges_def arc_to_ends_def sum_unfold_sum_mset by (simp add:image_mset.compositionality comp_def del:see_compress.simps) also have "...=\(\x \# edges ?G.f(snd x mod m)*f(fst x mod m))+(\x=?n-m.." - unfolding edges_of_compress[OF assms(2,3)] sum_unfold_sum_mset - by (simp add:image_mset.compositionality sum_mset_repeat comp_def + unfolding edges_of_compress[OF assms(2,3)] sum_unfold_sum_mset + by (simp add:image_mset.compositionality sum_mset_repeat comp_def case_prod_beta power2_eq_square del:see_compress.simps) also have "...=\(\(u,v) \# edges ?G.f(u mod m)*f(v mod m))+(\x=?n-m.." by (intro_cong "[\\<^sub>1 abs, \\<^sub>2 (+), \\<^sub>1 sum_mset]" more:image_mset_cong) (simp_all add:case_prod_beta) also have "... \ \\(u,v) \# edges ?G.f(u mod m)*f(v mod m)\+\\x=?n-m.. " by (intro abs_triangle_ineq) also have "... = ?d * (\\(u,v) \# edges ?G.f(v mod m)*f(u mod m)\/G.d+\\x=?n-m..)" - unfolding d_eq using G.d_gt_0 + unfolding d_eq using G.d_gt_0 by (simp add:divide_simps ac_simps sum_distrib_left[symmetric] abs_mult) - also have "... = ?d * (\G.g_inner f' (G.g_step f')\ + \\x=?n-m..)" + also have "... = ?d * (\G.g_inner f' (G.g_step f')\ + \\x=?n-m..)" unfolding G.g_inner_step_eq sum_unfold_sum_mset edges_def arc_to_ends_def f'_def by (simp add:image_mset.compositionality comp_def del:see_compress.simps) also have "...\ ?d * ((G.\\<^sub>a * G.g_norm f'^2 + (1-G.\\<^sub>a)*G.g_inner f' (\_.1)^2/ G.n) + \\x=?n-m..)" by (intro add_mono G.expansionD3 mult_left_mono) auto also have "... = ?d * (G.\\<^sub>a * ?L1 + (1 - G.\\<^sub>a) * ?L2)" by (simp add:algebra_simps) also have "... \ ?d * (G.\\<^sub>a * (2 * H.g_norm f^2) + (1-G.\\<^sub>a) * H.g_norm f^2)" unfolding 2 using G.\_ge_0 G.\_le_1 by (intro mult_left_mono add_mono 3) auto also have "... = ?R" unfolding d_eq[symmetric] by (simp add:algebra_simps) finally show ?thesis by simp qed hence "H.\\<^sub>a \ (G.d*G.\\<^sub>a+G.d)/H.d" using G.d_gt_0 G.\_ge_0 by (intro H.expander_intro) (auto simp del:see_compress.simps) also have "... = (see_degree e * G.\\<^sub>a + see_degree e) / (2* see_degree e)" unfolding d_eq[symmetric] regular_graph_degree_eq_see_degree[OF H.regular_graph_axioms] by simp also have "... = G.\\<^sub>a/2 + 1/2" using 1 by (simp add:field_simps) also have "... \ \\<^sub>a/2 + 1/2" using assms(1) unfolding is_expander_def by simp finally have "H.\\<^sub>a \ \\<^sub>a/2 + 1/2" by simp thus ?thesis unfolding is_expander_def using 0 by simp qed text \The graph power of a strongly explicit expander graph is itself a strongly explicit expander graph.\ fun to_digits :: "nat \ nat \ nat \ nat list" - where + where "to_digits _ 0 _ = []" | "to_digits b (Suc l) k = (k mod b)# to_digits b l (k div b)" fun from_digits :: "nat \ nat list \ nat" where "from_digits b [] = 0" | "from_digits b (x#xs) = x + b * from_digits b xs" lemma to_from_digits: assumes "length xs = n" "set xs \ {.. {.. 0") case True have "from_digits b xs \ b^length xs - 1" - using assms(2) + using assms(2) proof (induction xs) case Nil then show ?case by simp next case (Cons a xs) have "from_digits b (a # xs) = a + b * from_digits b xs" by simp also have "... \ (b-1) + b * from_digits b xs" using Cons by (intro add_mono) auto also have "... \ (b-1) + b * (b^length xs-1)" using Cons(2) by (intro add_mono mult_left_mono Cons(1)) auto also have "... = b^length (a#xs) - 1" using True by (simp add:algebra_simps) - finally show "from_digits b (a # xs) \ b^length (a#xs) - 1" by simp + finally show "from_digits b (a # xs) \ b^length (a#xs) - 1" by simp qed also have "... < b^n" using True assms(1) by simp finally show ?thesis by simp -next +next case False hence "b = 0" by simp hence "xs = []" using assms(2) by simp thus ?thesis using assms(1) by simp qed lemma from_digits_inj: "inj_on (from_digits b) {xs. set xs \ {.. length xs = n}" by (intro inj_on_inverseI[where g="to_digits b n"] to_from_digits) auto - + fun see_power :: "nat \ strongly_explicit_expander \ strongly_explicit_expander" where "see_power l e = \ see_size = see_size e, see_degree = see_degree e^l , see_step = (\k v. foldl (\y x. see_step e x y) v (to_digits (see_degree e) l k)) \" lemma graph_power_iso_see_power: assumes "fin_digraph (graph_of e)" shows "digraph_iso (graph_power (graph_of e) n) (graph_of (see_power n e))" proof - let ?G = "graph_of e" let ?P = "graph_power (graph_of e) n" let ?H = "graph_of (see_power n e)" let ?d = "see_degree e" let ?n = "see_size e" interpret fin_digraph "(graph_of e)" using assms by auto interpret P:fin_digraph ?P by (intro graph_power_fin) - define \ where + define \ where "\ = (\(u,v). Arc u (arc_walk_head ?G (u, v)) (from_digits ?d (map arc_label v)))" - define iso where "iso = - \ iso_verts = id, iso_arcs = \, iso_head = arc_head, iso_tail = arc_tail \" + define iso where "iso = + \ iso_verts = id, iso_arcs = \, iso_head = arc_head, iso_tail = arc_tail \" - have "xs = ys" if "length xs = length ys" "map arc_label xs = map arc_label ys" + have "xs = ys" if "length xs = length ys" "map arc_label xs = map arc_label ys" "is_arc_walk ?G u xs \ is_arc_walk ?G u ys \ u \ verts ?G" for xs ys u using that proof (induction xs ys arbitrary: u rule:list_induct2) case Nil then show ?case by simp next case (Cons x xs y ys) have "arc_label x = arc_label y" "u \ verts ?G" "x \ out_arcs ?G u" "y \ out_arcs ?G u" using Cons by auto - hence a:"x = y" + hence a:"x = y" unfolding graph_of_def by auto moreover have "head ?G y \ verts ?G" using Cons by auto - ultimately have "xs = ys" + ultimately have "xs = ys" using Cons(3,4) by (intro Cons(2)[of "head ?G y"]) auto thus ?case using a by auto qed hence 5:"inj_on (\(u,v). (u, map arc_label v)) (arc_walks ?G n)" unfolding arc_walks_def by (intro inj_onI) auto - have 3:"set (map arc_label (snd xs)) \ {.. {.. arc_walks ?G n" for xs proof - show "length (snd xs) = n" using subsetD[OF is_arc_walk_set[where G="?G"]] that unfolding arc_walks_def by auto have "set (snd xs) \ arcs ?G" using subsetD[OF is_arc_walk_set[where G="?G"]] that unfolding arc_walks_def by auto thus "set (map arc_label (snd xs)) \ {..(u,v). (u, from_digits ?d (map arc_label v))) (arc_walks ?G n)" using inj_onD[OF 5] inj_onD[OF from_digits_inj] by (intro inj_onI) auto hence "inj_on \ (arc_walks ?G n)" unfolding inj_on_def \_def by auto - hence "inj_on (iso_arcs iso) (arcs (graph_power (graph_of e) n))" + hence "inj_on (iso_arcs iso) (arcs (graph_power (graph_of e) n))" unfolding iso_def graph_power_def by simp - moreover have "inj_on (iso_verts iso) (verts (graph_power (graph_of e) n))" + moreover have "inj_on (iso_verts iso) (verts (graph_power (graph_of e) n))" unfolding iso_def by simp - moreover have + moreover have "iso_verts iso (tail ?P a) = iso_tail iso (iso_arcs iso a)" - "iso_verts iso (head ?P a) = iso_head iso (iso_arcs iso a)" if "a \ arcs ?P" for a + "iso_verts iso (head ?P a) = iso_head iso (iso_arcs iso a)" if "a \ arcs ?P" for a unfolding \_def iso_def graph_power_def by (simp_all add:case_prod_beta) ultimately have 0:"P.digraph_isomorphism iso" - unfolding P.digraph_isomorphism_def by (intro conjI ballI P.wf_digraph_axioms) auto + unfolding P.digraph_isomorphism_def by (intro conjI ballI P.wf_digraph_axioms) auto have "card((\(u, v).(u,from_digits ?d (map arc_label v)))`arc_walks ?G n)=card(arc_walks ?G n)" by (intro card_image 7) also have "... = ?d^n * ?n" - by (intro card_arc_walks_see fin_digraph_axioms) + by (intro card_arc_walks_see fin_digraph_axioms) finally have "card((\(u, v).(u,from_digits ?d (map arc_label v)))`arc_walks ?G n) = ?d^n * ?n" by simp moreover have "fst v \ {.. arc_walks ?G n" for v using that unfolding arc_walks_def graph_of_def by auto moreover have "from_digits ?d (map arc_label (snd v)) < ?d ^ n" if "v \ arc_walks ?G n" for v using 3[OF that] by (intro from_digits_range) auto - ultimately have 2: - "{..{..(u,v). (u, from_digits ?d (map arc_label v))) ` arc_walks ?G n" + ultimately have 2: + "{..{..(u,v). (u, from_digits ?d (map arc_label v))) ` arc_walks ?G n" by (intro card_subset_eq[symmetric]) auto - have "foldl (\y x. see_step e x y) u (map arc_label w) = arc_walk_head ?G (u,w)" - if "is_arc_walk ?G u w" "u \ verts ?G" for u w + have "foldl (\y x. see_step e x y) u (map arc_label w) = arc_walk_head ?G (u,w)" + if "is_arc_walk ?G u w" "u \ verts ?G" for u w using that proof (induction w rule:rev_induct) case Nil then show ?case by (simp add:arc_walk_head_def) next case (snoc x xs) - hence "x \ arcs ?G" by (simp add:is_arc_walk_snoc) - hence "see_step e (arc_label x) (tail ?G x) = (head ?G x)" + hence "x \ arcs ?G" by (simp add:is_arc_walk_snoc) + hence "see_step e (arc_label x) (tail ?G x) = (head ?G x)" unfolding graph_of_def by (auto simp add:image_iff) also have "... = arc_walk_head (graph_of e) (u, xs @ [x])" unfolding arc_walk_head_def by simp finally have "see_step e (arc_label x) (tail ?G x) = arc_walk_head (graph_of e) (u, xs @ [x])" by simp thus ?case using snoc by (simp add:is_arc_walk_snoc) qed - hence 4: "foldl (\y x. see_step e x y) (fst x) (map arc_label (snd x)) = arc_walk_head ?G x" - if "x \ arc_walks (graph_of e) n" for x + hence 4: "foldl (\y x. see_step e x y) (fst x) (map arc_label (snd x)) = arc_walk_head ?G x" + if "x \ arc_walks (graph_of e) n" for x using that unfolding arc_walks_def by (simp add:case_prod_beta) have "arcs ?H = (\(v, i). Arc v (see_step (see_power n e) i v) i) ` ({..{..(v,w). Arc v (see_step (see_power n e) (from_digits ?d (map arc_label w)) v) + unfolding graph_of_def by simp + also have "... = (\(v,w). Arc v (see_step (see_power n e) (from_digits ?d (map arc_label w)) v) (from_digits ?d (map arc_label w))) ` arc_walks ?G n" unfolding 2 image_image by (simp del:see_power.simps add: case_prod_beta comp_def) - also have "... = (\(v,w). Arc v (foldl (\y x. see_step e x y) v (map arc_label w)) + also have "... = (\(v,w). Arc v (foldl (\y x. see_step e x y) v (map arc_label w)) (from_digits ?d (map arc_label w))) ` arc_walks ?G n" using 3 by (intro image_cong refl) (simp add:case_prod_beta to_from_digits) also have "... = \ ` arc_walks ?G n" - unfolding \_def using 4 by (simp add:case_prod_beta) + unfolding \_def using 4 by (simp add:case_prod_beta) also have "... = iso_arcs iso ` arcs ?P" unfolding iso_def graph_power_def by simp finally have "arcs ?H = iso_arcs iso ` arcs ?P" by simp moreover have "verts ?H = iso_verts iso ` verts ?P" unfolding iso_def graph_of_def graph_power_def by simp moreover have "tail ?H = iso_tail iso" unfolding iso_def graph_of_def by simp moreover have "head ?H = iso_head iso" unfolding iso_def graph_of_def by simp ultimately have 1:"?H = app_iso iso ?P" - unfolding app_iso_def + unfolding app_iso_def by (intro pre_digraph.equality) (simp_all del:see_power.simps) - show ?thesis + show ?thesis using 0 1 unfolding digraph_iso_def by auto qed lemma see_power: assumes "is_expander e \\<^sub>a" shows "is_expander (see_power n e) (\\<^sub>a^n)" proof - - interpret G: "regular_graph" "graph_of e" + interpret G: "regular_graph" "graph_of e" using assms unfolding is_expander_def by auto interpret H:"regular_graph" "graph_power (graph_of e) n" by (intro G.graph_power_regular) have 0:"digraph_iso (graph_power (graph_of e) n) (graph_of (see_power n e))" - by (intro graph_power_iso_see_power) auto + by (intro graph_power_iso_see_power) auto have "regular_graph.\\<^sub>a (graph_of (see_power n e)) = H.\\<^sub>a" using H.regular_graph_iso_expansion[OF 0] by auto also have "... \ G.\\<^sub>a^n" by (intro G.graph_power_expansion) also have "... \ \\<^sub>a^n" using assms(1) unfolding is_expander_def by (intro power_mono G.\_ge_0) auto finally have "regular_graph.\\<^sub>a (graph_of (see_power n e)) \ \\<^sub>a^n" by simp moreover have "regular_graph (graph_of (see_power n e))" using H.regular_graph_iso[OF 0] by auto ultimately show ?thesis unfolding is_expander_def by auto qed text \The Margulis Construction from Section~\ref{sec:margulis} is a strongly explicit expander graph.\ definition mgg_vert :: "nat \ nat \ (int \ int)" where "mgg_vert n x = (x mod n, x div n)" definition mgg_vert_inv :: "nat \ (int \ int) \ nat" where "mgg_vert_inv n x = nat (fst x) + nat (snd x) * n" lemma mgg_vert_inv: assumes "n > 0" "x \ {0..{0.. (nat \ int)" where "mgg_arc k = (k mod 4, if k \ 4 then (-1) else 1)" definition mgg_arc_inv :: "(nat \ int) \ nat" where "mgg_arc_inv x = (nat (fst x) + 4 * of_bool (snd x < 0))" lemma mgg_arc_inv: assumes "x \ {..<4}\{-1,1}" shows "mgg_arc (mgg_arc_inv x) = x" using assms unfolding mgg_arc_def mgg_arc_inv_def by auto definition see_mgg :: "nat \ strongly_explicit_expander" where - "see_mgg n = \ see_size = n^2, see_degree = 8, + "see_mgg n = \ see_size = n^2, see_degree = 8, see_step = (\i v. mgg_vert_inv n (mgg_graph_step n (mgg_vert n v) (mgg_arc i))) \" lemma mgg_graph_iso: assumes "n > 0" shows "digraph_iso (mgg_graph n) (graph_of (see_mgg n))" proof - let ?v = "mgg_vert n" let ?vi = "mgg_vert_inv n" let ?a = "mgg_arc" let ?ai = "mgg_arc_inv" let ?G = "graph_of (see_mgg n)" let ?s = "mgg_graph_step n" define \ where "\ a = Arc (?vi (arc_tail a)) (?vi (arc_head a)) (?ai (arc_label a))" for a - define iso where "iso = - \ iso_verts = mgg_vert_inv n, iso_arcs = \, iso_head = arc_head, iso_tail = arc_tail \" + define iso where "iso = + \ iso_verts = mgg_vert_inv n, iso_arcs = \, iso_head = arc_head, iso_tail = arc_tail \" interpret M: margulis_gaber_galil n using assms by unfold_locales have inj_vi: "inj_on ?vi (verts M.G)" unfolding mgg_graph_def mgg_vert_inv_def by (intro inj_on_inverseI[where g="mgg_vert n"]) (auto simp:mgg_vert_def) have "card (?vi ` verts M.G) = card (verts M.G)" by (intro card_image inj_vi) moreover have "card (verts M.G) = n\<^sup>2" unfolding mgg_graph_def by (auto simp:power2_eq_square) - moreover have "mgg_vert_inv n x \ {..2}" if "x \ verts M.G" for x + moreover have "mgg_vert_inv n x \ {..2}" if "x \ verts M.G" for x proof - have "mgg_vert_inv n x = nat (fst x) + nat (snd x) * n" unfolding mgg_vert_inv_def by simp also have "... \ (n-1) + (n-1) * n" using that unfolding mgg_graph_def by (intro add_mono mult_right_mono) auto also have "... = n * n - 1" using assms by (simp add:algebra_simps) also have "... < n^2" using assms by (simp add: power2_eq_square) finally have "mgg_vert_inv n x < n^2" by simp thus ?thesis by simp qed ultimately have 0:"{.. {-1,1})" unfolding mgg_arc_inv_def by (intro inj_onI) auto have "card (?ai ` ({..<4} \ {- 1, 1})) = card ({..<4::nat} \ {-1,1::int})" by (intro card_image inj_ai) hence 1:"{..<8} = ?ai ` ({..<4} \ {-1,1})" by (intro card_subset_eq[symmetric] image_subsetI) (auto simp add:mgg_arc_inv_def) have "arcs ?G = (\(v, i). Arc v (?vi (?s (?v v) (?a i))) i) ` ({..2} \ {..<8})" - by (simp add:see_mgg_def graph_of_def) - also have "... = (\(v, i). Arc (?vi v) (?vi (?s (?v (?vi v)) (?a (?ai i)))) (?ai i)) ` + by (simp add:see_mgg_def graph_of_def) + also have "... = (\(v, i). Arc (?vi v) (?vi (?s (?v (?vi v)) (?a (?ai i)))) (?ai i)) ` (verts M.G \ ({..<4} \ {-1,1}))" unfolding 0 1 mgg_arc_inv by (auto simp add:image_iff) also have "... = (\(v, i). Arc (?vi v) (?vi (?s v i)) (?ai i)) ` (verts M.G \ ({..<4} \ {-1,1}))" using mgg_vert_inv[OF assms] mgg_arc_inv unfolding mgg_graph_def by (intro image_cong) auto - also have "... = (\ \ (\(t, l). Arc t (?s t l) l)) ` (verts M.G \ ({..<4} \ {-1,1}))" + also have "... = (\ \ (\(t, l). Arc t (?s t l) l)) ` (verts M.G \ ({..<4} \ {-1,1}))" unfolding \_def by (intro image_cong refl) ( simp add:comp_def case_prod_beta ) - also have "... = \ ` arcs M.G" + also have "... = \ ` arcs M.G" unfolding mgg_graph_def by (simp add:image_image) also have "... = iso_arcs iso ` arcs (mgg_graph n)" unfolding iso_def by simp - finally have "arcs (graph_of (see_mgg n)) = iso_arcs iso ` arcs (mgg_graph n)" + finally have "arcs (graph_of (see_mgg n)) = iso_arcs iso ` arcs (mgg_graph n)" by simp moreover have "verts ?G = iso_verts iso ` verts (mgg_graph n)" - unfolding iso_def graph_of_def see_mgg_def using 0 by simp + unfolding iso_def graph_of_def see_mgg_def using 0 by simp moreover have "tail ?G = iso_tail iso" unfolding iso_def graph_of_def by simp moreover have "head ?G = iso_head iso" unfolding iso_def graph_of_def by simp - ultimately have 0:"?G = app_iso iso (mgg_graph n)" + ultimately have 0:"?G = app_iso iso (mgg_graph n)" unfolding app_iso_def by (intro pre_digraph.equality) simp_all have "inj_on \ (arcs M.G)" proof (rule inj_onI) fix x y assume assms': "x \ arcs M.G" "y \ arcs M.G" "\ x = \ y" have "?vi (head M.G x) = ?vi (head M.G y)" using assms'(3) unfolding \_def mgg_graph_def by auto hence "head M.G x = head M.G y" using assms'(1,2) by (intro inj_onD[OF inj_vi]) auto hence "arc_head x = arc_head y" unfolding mgg_graph_def by simp moreover have "?vi (tail M.G x) = ?vi (tail M.G y)" using assms'(3) unfolding \_def mgg_graph_def by auto hence "tail M.G x = tail M.G y" using assms'(1,2) by (intro inj_onD[OF inj_vi]) auto hence "arc_tail x = arc_tail y" unfolding mgg_graph_def by simp moreover have "?ai (arc_label x) = ?ai (arc_label y)" using assms'(3) unfolding \_def by auto hence "arc_label x = arc_label y" using assms'(1,2) unfolding mgg_graph_def by (intro inj_onD[OF inj_ai]) (auto simp del:mgg_graph_step.simps) ultimately show "x = y" by (intro arc.expand) auto qed - hence "inj_on (iso_arcs iso) (arcs M.G)" + hence "inj_on (iso_arcs iso) (arcs M.G)" unfolding iso_def by simp - moreover have "inj_on (iso_verts iso) (verts M.G)" + moreover have "inj_on (iso_verts iso) (verts M.G)" using inj_vi unfolding iso_def by simp - moreover have - "iso_verts iso (tail M.G a) = iso_tail iso (iso_arcs iso a)" + moreover have + "iso_verts iso (tail M.G a) = iso_tail iso (iso_arcs iso a)" "iso_verts iso (head M.G a) = iso_head iso (iso_arcs iso a)" if "a \ arcs M.G" for a unfolding iso_def \_def mgg_graph_def by auto ultimately have 1:"M.digraph_isomorphism iso" unfolding M.digraph_isomorphism_def by (intro conjI ballI M.wf_digraph_axioms) auto show ?thesis unfolding digraph_iso_def using 0 1 by auto qed lemma see_mgg: assumes "n > 0" shows "is_expander (see_mgg n) (5* sqrt 2 / 8)" proof - - interpret G: "margulis_gaber_galil" "n" + interpret G: "margulis_gaber_galil" "n" using assms by unfold_locales auto note 0 = mgg_graph_iso[OF assms] have "regular_graph.\\<^sub>a (graph_of (see_mgg n)) = G.\\<^sub>a" using G.regular_graph_iso_expansion[OF 0] by auto also have "... \ (5* sqrt 2 / 8)" using G.mgg_numerical_radius unfolding G.MGG_bound_def by simp finally have "regular_graph.\\<^sub>a (graph_of (see_mgg n)) \ (5* sqrt 2 / 8)" by simp moreover have "regular_graph (graph_of (see_mgg n))" using G.regular_graph_iso[OF 0] by auto ultimately show ?thesis unfolding is_expander_def by auto qed -text \Using all of the above it is possible to construct strongly explicit expanders of every +text \Using all of the above it is possible to construct strongly explicit expanders of every size and spectral gap with asymptotically optimal degree.\ -definition see_standard_aux +definition see_standard_aux where "see_standard_aux n = see_compress n (see_mgg (nat \sqrt n\))" lemma see_standard_aux: assumes "n > 0" - shows + shows "is_expander (see_standard_aux n) ((8+5 * sqrt 2) / 16)" (is "?A") "see_degree (see_standard_aux n) = 16" (is "?B") "see_size (see_standard_aux n) = n" (is "?C") proof - have 2:"sqrt (real n) > -1" by (rule less_le_trans[where y="0"]) auto - have 0:"real n \ of_int \sqrt (real n)\^2" + have 0:"real n \ of_int \sqrt (real n)\^2" by (simp add:sqrt_le_D) - consider (a) "n = 1" | (b) "n \ 2 \ n \ 4" | (c) "n \ 5 \ n \ 9" | (d) "n \ 10" + consider (a) "n = 1" | (b) "n \ 2 \ n \ 4" | (c) "n \ 5 \ n \ 9" | (d) "n \ 10" using assms by linarith - hence 1:"of_int \sqrt (real n)\^2 \ 2 * real n" + hence 1:"of_int \sqrt (real n)\^2 \ 2 * real n" proof (cases) case a then show ?thesis by simp next case b hence "real_of_int \sqrt (real n)\^2 \ of_int \sqrt (real 4)\^2" using 2 by (intro power_mono iffD2[OF of_int_le_iff] ceiling_mono iffD2[OF real_sqrt_le_iff]) auto also have "... = 2 * real 2" by simp also have "... \ 2 * real n" using b by (intro mult_left_mono) auto finally show ?thesis by simp next case c hence "real_of_int \sqrt (real n)\^2 \ of_int \sqrt (real 9)\^2" using 2 by (intro power_mono iffD2[OF of_int_le_iff] ceiling_mono iffD2[OF real_sqrt_le_iff]) auto also have "... = 9" by simp also have "... \ 2 * real 5" by simp also have "... \ 2 * real n" using c by (intro mult_left_mono) auto finally show ?thesis by simp next case d have "real_of_int \sqrt (real n)\^2 \ (sqrt (real n)+1)^2" using 2 by (intro power_mono) auto also have "... = real n + sqrt (4 * real n + 0) + 1" using real_sqrt_pow2 by (simp add:power2_eq_square algebra_simps real_sqrt_mult) also have "... \ real n + sqrt (4 * real n + (real n * (real n - 6) + 1)) + 1" using d by (intro add_mono iffD2[OF real_sqrt_le_iff]) auto also have "... = real n + sqrt ((real n-1)^2) + 1" by (intro_cong "[\\<^sub>2 (+), \\<^sub>1 sqrt]") (auto simp add:power2_eq_square algebra_simps) also have "... = 2 * real n" using d by simp finally show ?thesis by simp qed have "real (nat \sqrt (real n)\^2) = of_int \sqrt (real n)\^2" unfolding of_nat_power using 2 by (simp add:not_less) - also have "... \ {real n..2 * real n}" + also have "... \ {real n..2 * real n}" using 0 1 by auto also have "... = {real n..real (2*n)}" by simp finally have "real (nat \sqrt (real n)\^2) \ {real n..real (2*n)}" by simp hence "nat \sqrt (real n)\^2 \ {n..2*n}" by (simp del:of_nat_mult) hence "see_size (see_mgg (nat \sqrt (real n)\)) \ {n..2*n}" by (simp add:see_mgg_def) moreover have "sqrt (real n) > 0" using assms by simp hence "0 < nat \sqrt (real n)\" by simp ultimately have "is_expander (see_standard_aux n) ((5* sqrt 2 / 8)/2 + 1/2)" unfolding see_standard_aux_def by (intro see_compress see_mgg) auto - thus ?A + thus ?A by (auto simp add:field_simps) - show ?B + show ?B unfolding see_standard_aux_def by (simp add:see_mgg_def) - show ?C + show ?C unfolding see_standard_aux_def by simp qed definition see_standard_power where "see_standard_power x = (if x \ (0::real) then 0 else nat \ln x / ln 0.95\)" lemma see_standard_power: assumes "\\<^sub>a > 0" shows "0.95^(see_standard_power \\<^sub>a) \ \\<^sub>a" (is "?L \ ?R") proof (cases "\\<^sub>a \ 1") case True - hence "0 \ ln \\<^sub>a / ln 0.95" + hence "0 \ ln \\<^sub>a / ln 0.95" using assms by (intro divide_nonpos_neg) auto - hence 1:"0 \ \ln \\<^sub>a / ln 0.95\" + hence 1:"0 \ \ln \\<^sub>a / ln 0.95\" by simp have "?L = 0.95^nat \ln \\<^sub>a / ln 0.95\" using assms unfolding see_standard_power_def by simp also have "... = 0.95 powr (of_nat (nat (\ln \\<^sub>a / ln 0.95\)))" - by (subst powr_realpow) auto + by (subst powr_realpow) auto also have "... = 0.95 powr \ln \\<^sub>a / ln 0.95\" using 1 by (subst of_nat_nat) auto also have "... \ 0.95 powr (ln \\<^sub>a / ln 0.95)" by (intro powr_mono_rev) auto also have "... = ?R" using assms unfolding powr_def by simp finally show ?thesis by simp next case False hence "ln \\<^sub>a / ln 0.95 \ 0" by (subst neg_divide_le_eq) auto hence "see_standard_power \\<^sub>a = 0" unfolding see_standard_power_def by simp then show ?thesis using False by simp qed -lemma see_standard_power_eval[code]: +lemma see_standard_power_eval[code]: "see_standard_power x = (if x \ 0 \ x \ 1 then 0 else (1+see_standard_power (x/0.95)))" proof (cases "x \ 0 \ x \ 1") case True have "ln x / ln (19 / 20) \ 0" if "x > 0" proof - have "x \ 1" using that True by auto thus ?thesis by (intro divide_nonneg_neg) auto qed then show ?thesis using True unfolding see_standard_power_def by simp next case False hence x_range: "x > 0" "x < 1" by auto - have "ln (x / 0.95) < ln (1/0.95)" + have "ln (x / 0.95) < ln (1/0.95)" using x_range by (intro iffD2[OF ln_less_cancel_iff]) auto - also have "... = - ln 0.95" + also have "... = - ln 0.95" by (subst ln_div) auto finally have "ln (x / 0.95) < - ln 0.95" by simp hence 0: "-1 < ln (x / 0.95) / ln 0.95" by (subst neg_less_divide_eq) auto have "see_standard_power x = nat \ln x / ln 0.95\" - using x_range unfolding see_standard_power_def by simp + using x_range unfolding see_standard_power_def by simp also have "... = nat \ln (x/0.95) / ln 0.95 + 1\" by (subst ln_div[OF x_range(1)]) (simp_all add:field_simps ) also have "... = nat (\ln (x/0.95) / ln 0.95\+1)" by (intro arg_cong[where f="nat"]) simp also have "... = 1 + nat \ln (x/0.95) / ln 0.95\" using 0 by (subst nat_add_distrib) auto also have "... = (if x \ 0 \ 1 \ x then 0 else 1 + see_standard_power (x/0.95))" unfolding see_standard_power_def using x_range by auto finally show ?thesis by simp qed definition see_standard :: "nat \ real \ strongly_explicit_expander" where "see_standard n \\<^sub>a = see_power (see_standard_power \\<^sub>a) (see_standard_aux n)" theorem see_standard: assumes "n > 0" "\\<^sub>a > 0" shows "is_expander (see_standard n \\<^sub>a) \\<^sub>a" and "see_size (see_standard n \\<^sub>a) = n" and "see_degree (see_standard n \\<^sub>a) = 16 ^ (nat \ln \\<^sub>a / ln 0.95\)" (is "?C") proof - have 0:"is_expander (see_standard_aux n) 0.95" by (intro see_standard_aux(1)[OF assms(1)] is_expander_mono[where a="(8+5 * sqrt 2) / 16"]) (approximation 10) show "is_expander (see_standard n \\<^sub>a) \\<^sub>a" - unfolding see_standard_def + unfolding see_standard_def by (intro see_power 0 is_expander_mono[where a="0.95^(see_standard_power \\<^sub>a)"] see_standard_power assms(2)) show "see_size (see_standard n \\<^sub>a) = n" unfolding see_standard_def using see_standard_aux[OF assms(1)] by simp have "see_degree (see_standard n \\<^sub>a) = 16 ^ (see_standard_power \\<^sub>a)" - unfolding see_standard_def using see_standard_aux[OF assms(1)] by simp + unfolding see_standard_def using see_standard_aux[OF assms(1)] by simp also have "... = 16 ^ (nat \ln \\<^sub>a / ln 0.95\)" unfolding see_standard_power_def using assms(2) by simp finally show ?C by simp qed fun see_sample_walk :: "strongly_explicit_expander \ nat \ nat \ nat list" - where + where "see_sample_walk e 0 x = [x]" | - "see_sample_walk e (Suc l) x = (let w = see_sample_walk e l (x div (see_degree e)) in + "see_sample_walk e (Suc l) x = (let w = see_sample_walk e l (x div (see_degree e)) in w@[see_step e (x mod (see_degree e)) (last w)])" theorem see_sample_walk: fixes e l assumes "fin_digraph (graph_of e)" - defines "r \ see_size e * see_degree e ^l" + defines "r \ see_size e * see_degree e ^l" shows "{# see_sample_walk e l k. k \# mset_set {.. {j * ?d..<(j + 1) * ?d} = {}" if "i \ j" for i j using that index_div_eq by blast - have 2:"vertices_from ?G x = {# see_step e i x. i \# mset_set {..# mset_set {.. verts ?G" for x proof - - have "x < ?n" + have "x < ?n" using that unfolding graph_of_def by simp hence 1:"out_arcs ?G x = (\i. Arc x (see_step e i x) i) ` {..# mset_set (out_arcs ?G x) #}" unfolding verts_from_alt by (simp add:graph_of_def) - also have "... = {# arc_head a. a \# {# Arc x (see_step e i x) i. i \# mset_set {..# {# Arc x (see_step e i x) i. i \# mset_set {..ww < r. card {w * ?d..<(w + 1) *?d})" using 1 by (intro card_UN_disjoint) auto also have "... = r * ?d" by simp finally have "card (\w ?d * r" if "z < r" for z - proof - + moreover have "?d + z * ?d \ ?d * r" if "z < r" for z + proof - have "?d + z * ?d = ?d * (z + 1)" by simp - also have "... \ ?d * r" + also have "... \ ?d * r" using that by (intro mult_left_mono) auto finally show ?thesis by simp qed ultimately have 0: "(\w# mset_set {..# mset_set {..# mset_set (\w (\w. {w * ?d..<(w + 1) * ?d})) (mset_set {..#mset_set {w*?d..<(w+1)*?d}#}. w\#mset_set {..#mset_set ((+)(w*?d)`{..#mset_set {..\<^sub>1 concat_mset, \\<^sub>2 image_mset, \\<^sub>1 mset_set]" more:ext) (simp add: atLeast0LessThan[symmetric]) - also have "... = concat_mset + also have "... = concat_mset {#{#?w (l+1) i. i\#image_mset ((+) (w*?d)) (mset_set {..#mset_set {..\<^sub>1 concat_mset, \\<^sub>2 image_mset]" more:image_mset_cong + by (intro_cong "[\\<^sub>1 concat_mset, \\<^sub>2 image_mset]" more:image_mset_cong image_mset_mset_set[symmetric] inj_onI) auto also have "... = concat_mset {#{#?w (l+1) (w*?d+i).i\#mset_set {..#mset_set {..#mset_set {..#mset_set {..\<^sub>1 concat_mset]" more:image_mset_cong) (simp add:Let_def) also have "... = concat_mset {#{#w@[see_step e i (last w)].i\#mset_set {..#walks' ?G l#}" unfolding r_def Suc[symmetric] image_mset.compositionality comp_def by simp - also have "... = concat_mset + also have "... = concat_mset {#{#w@[x].x\#{# see_step e i (last w). i\#mset_set {..# walks' ?G l#}" unfolding image_mset.compositionality comp_def by simp also have "... = concat_mset {#{#w@[x].x\#vertices_from ?G (last w)#}. w \# walks' ?G l#}" using last_in_set set_walks_2(1,2) by (intro_cong "[\\<^sub>1 concat_mset, \\<^sub>2 image_mset]" more:image_mset_cong 2[symmetric]) blast also have "... = walks' (graph_of e) (l+1)" by (simp add:image_mset.compositionality comp_def) - finally show ?case by simp + finally show ?case by simp qed unbundle no_intro_cong_syntax end diff --git a/thys/Expander_Graphs/Expander_Graphs_TTS.thy b/thys/Expander_Graphs/Expander_Graphs_TTS.thy --- a/thys/Expander_Graphs/Expander_Graphs_TTS.thy +++ b/thys/Expander_Graphs/Expander_Graphs_TTS.thy @@ -1,141 +1,141 @@ section \Setup for Types to Sets\label{sec:tts}\ theory Expander_Graphs_TTS - imports - Expander_Graphs_Definition + imports + Expander_Graphs_Definition "HOL-Analysis.Cartesian_Space" "HOL-Types_To_Sets.Types_To_Sets" begin text \This section sets up a sublocale with the assumption that there is a finite type with the same -cardinality as the vertex set of a regular graph. This allows defining the adjacency matrix for +cardinality as the vertex set of a regular graph. This allows defining the adjacency matrix for the graph using type-based linear algebra. Theorems shown in the sublocale that do not refer to the local type are then lifted to the @{locale regular_graph} locale using the Types-To-Sets mechanism.\ locale regular_graph_tts = regular_graph + fixes n_itself :: "('n :: finite) itself" assumes td: "\(f :: ('n \ 'a)) g. type_definition f g (verts G)" begin -definition td_components :: "('n \ 'a) \ ('a \ 'n)" +definition td_components :: "('n \ 'a) \ ('a \ 'n)" where "td_components = (SOME q. type_definition (fst q) (snd q) (verts G))" definition enum_verts where "enum_verts = fst td_components" definition enum_verts_inv where "enum_verts_inv = snd td_components" sublocale type_definition "enum_verts" "enum_verts_inv" "verts G" proof - have 0:"\q. type_definition ((fst q)::('n \ 'a)) (snd q) (verts G)" using td by simp show "type_definition enum_verts enum_verts_inv (verts G)" unfolding td_components_def enum_verts_def enum_verts_inv_def using someI_ex[OF 0] by simp qed lemma enum_verts: "bij_betw enum_verts UNIV (verts G)" unfolding bij_betw_def by (simp add: Rep_inject Rep_range inj_on_def) text \The stochastic matrix associated to the graph.\ definition A :: "('c::field)^'n^'n" where "A = (\ i j. of_nat (count (edges G) (enum_verts j,enum_verts i))/of_nat d)" lemma card_n: "CARD('n) = n" unfolding n_def card by simp lemma symmetric_A: "transpose A = A" proof - have "A $ i $ j = A $ j $ i" for i j unfolding A_def count_edges arcs_betw_def using symmetric_multi_graphD[OF sym] by auto thus ?thesis unfolding transpose_def by (intro iffD2[OF vec_eq_iff] allI) auto qed -lemma g_step_conv: +lemma g_step_conv: "(\ i. g_step f (enum_verts i)) = A *v (\ i. f (enum_verts i))" proof - have "g_step f (enum_verts i) = (\j\UNIV. A $ i $ j * f (enum_verts j))" (is "?L = ?R") for i proof - have "?L = (\x\in_arcs G (enum_verts i). f (tail G x) / d)" unfolding g_step_def by simp also have "... = (\x\#vertices_to G (enum_verts i). f x/d)" unfolding verts_to_alt sum_unfold_sum_mset by (simp add:image_mset.compositionality comp_def) also have "... = (\j\verts G. (count (vertices_to G (enum_verts i)) j) * (f j / real d))" by (intro sum_mset_conv_2 set_mset_vertices_to) auto also have "... = (\j\verts G. (count (edges G) (j,enum_verts i)) * (f j / real d))" - unfolding vertices_to_def count_mset_exp + unfolding vertices_to_def count_mset_exp by (intro sum.cong arg_cong[where f="real"] arg_cong2[where f="(*)"]) (auto simp add:filter_filter_mset image_mset_filter_mset_swap[symmetric] prod_eq_iff ac_simps) also have "...=(\j\UNIV.(count(edges G)(enum_verts j,enum_verts i))*(f(enum_verts j)/real d))" by (intro sum.reindex_bij_betw[symmetric] enum_verts) also have "... = ?R" unfolding A_def by simp finally show ?thesis by simp qed thus ?thesis unfolding matrix_vector_mult_def by (intro iffD2[OF vec_eq_iff] allI) simp qed -lemma g_inner_conv: +lemma g_inner_conv: "g_inner f g = (\ i. f (enum_verts i)) \ (\ i. g (enum_verts i))" unfolding inner_vec_def g_inner_def vec_lambda_beta inner_real_def conjugate_real_def by (intro sum.reindex_bij_betw[symmetric] enum_verts) -lemma g_norm_conv: +lemma g_norm_conv: "g_norm f = norm (\ i. f (enum_verts i))" proof - have "g_norm f^2 = norm (\ i. f (enum_verts i))^2" unfolding g_norm_sq power2_norm_eq_inner g_inner_conv by simp thus ?thesis using g_norm_nonneg norm_ge_zero by simp qed end lemma eg_tts_1: assumes "regular_graph G" assumes "\(f::('n::finite) \ 'a) g. type_definition f g (verts G)" shows "regular_graph_tts TYPE('n) G" - using assms + using assms unfolding regular_graph_tts_def regular_graph_tts_axioms_def by auto -context regular_graph +context regular_graph begin lemma remove_finite_premise_aux: assumes "\(Rep :: 'n \ 'a) Abs. type_definition Rep Abs (verts G)" shows "class.finite TYPE('n)" proof - obtain Rep :: "'n \ 'a" and Abs where d:"type_definition Rep Abs (verts G)" using assms by auto interpret type_definition Rep Abs "verts G" using d by simp - - have "finite (verts G)" by simp + + have "finite (verts G)" by simp thus ?thesis unfolding class.finite_def univ by auto qed -lemma remove_finite_premise: - "(class.finite TYPE('n) \ \(Rep :: 'n \ 'a) Abs. type_definition Rep Abs (verts G) \ PROP Q) - \ (\(Rep :: 'n \ 'a) Abs. type_definition Rep Abs (verts G) \ PROP Q)" +lemma remove_finite_premise: + "(class.finite TYPE('n) \ \(Rep :: 'n \ 'a) Abs. type_definition Rep Abs (verts G) \ PROP Q) + \ (\(Rep :: 'n \ 'a) Abs. type_definition Rep Abs (verts G) \ PROP Q)" (is "?L \ ?R") proof (intro Pure.equal_intr_rule) assume e:"\(Rep :: 'n \ 'a) Abs. type_definition Rep Abs (verts G)" and l:"PROP ?L" - hence f: "class.finite TYPE('n)" + hence f: "class.finite TYPE('n)" using remove_finite_premise_aux[OF e] by simp show "PROP ?R" using l[OF f] by auto next assume "\(Rep :: 'n \ 'a) Abs. type_definition Rep Abs (verts G)" and l:"PROP ?R" show "PROP ?L" using l by auto qed end end \ No newline at end of file diff --git a/thys/Expander_Graphs/Expander_Graphs_Walks.thy b/thys/Expander_Graphs/Expander_Graphs_Walks.thy --- a/thys/Expander_Graphs/Expander_Graphs_Walks.thy +++ b/thys/Expander_Graphs/Expander_Graphs_Walks.thy @@ -1,580 +1,580 @@ section \Random Walks\label{sec:random_walks}\ theory Expander_Graphs_Walks imports Expander_Graphs_Algebra Expander_Graphs_Eigenvalues Expander_Graphs_TTS Constructive_Chernoff_Bound -begin +begin unbundle intro_cong_syntax no_notation Matrix.vec_index (infixl "$" 100) hide_const Matrix.vec_index hide_const Matrix.vec no_notation Matrix.scalar_prod (infix "\" 70) fun walks' :: "('a,'b) pre_digraph \ nat \ ('a list) multiset" - where + where "walks' G 0 = image_mset (\x. [x]) (mset_set (verts G))" | - "walks' G (Suc n) = - concat_mset {#{#w @[z].z\# vertices_from G (last w)#}. w \# walks' G n#}" + "walks' G (Suc n) = + concat_mset {#{#w @[z].z\# vertices_from G (last w)#}. w \# walks' G n#}" definition "walks G l = (case l of 0 \ {#[]#} | Suc pl \ walks' G pl)" lemma Union_image_mono: "(\x. x \ A \ f x \ g x) \ \ (f ` A) \ \ (g ` A)" by auto context fin_digraph begin lemma count_walks': assumes "set xs \ verts G" assumes "length xs = l+1" shows "count (walks' G l) xs = (\i \ {.. []" using assms(2) by auto have "count (walks' G (length xs-1)) xs = (\i verts G" by simp - hence "count {#[x]. x \# mset_set (verts G)#} [x] = 1" - by (subst count_image_mset_inj, auto simp add:inj_def) - then show ?case by simp + hence "count {#[x]. x \# mset_set (verts G)#} [x] = 1" + by (subst count_image_mset_inj, auto simp add:inj_def) + then show ?case by simp next case (snoc x xs) have set_xs: "set xs \ verts G" using snoc by simp - define l where "l = length xs - 1" + define l where "l = length xs - 1" have l_xs: "length xs = l + 1" unfolding l_def using snoc by simp have "count (walks' G (length (xs @ [x]) - 1)) (xs @ [x]) = (\ys\#walks' G l. count {#ys @ [z]. z \# vertices_from G (last ys)#} (xs @ [x]))" by (simp add:l_xs count_concat_mset image_mset.compositionality comp_def) - also have "... = (\ys\#walks' G l. + also have "... = (\ys\#walks' G l. (if ys = xs then count {#xs @ [z]. z \# vertices_from G (last xs)#} (xs @ [x]) else 0))" - by (intro arg_cong[where f="sum_mset"] image_mset_cong) (auto intro!: count_image_mset_0_triv) + by (intro arg_cong[where f="sum_mset"] image_mset_cong) (auto intro!: count_image_mset_0_triv) also have "... = (\ys\#walks' G l.(if ys=xs then count (vertices_from G (last xs)) x else 0))" by (subst count_image_mset_inj, auto simp add:inj_def) also have "... = count (walks' G l) xs * count (vertices_from G (last xs)) x" by (subst sum_mset_delta, simp) also have "... = count (walks' G l) xs * count (edges G) (last xs, x)" - unfolding vertices_from_def count_mset_exp image_mset_filter_mset_swap[symmetric] + unfolding vertices_from_def count_mset_exp image_mset_filter_mset_swap[symmetric] filter_filter_mset by (simp add:prod_eq_iff) also have "... = count (walks' G l) xs * count (edges G) ((xs@[x])!l, (xs@[x])!(l+1))" - using snoc(1) unfolding l_def nth_append last_conv_nth[OF snoc(1)] by simp + using snoc(1) unfolding l_def nth_append last_conv_nth[OF snoc(1)] by simp also have "... = (\ii verts G" assumes "length xs = l" "l > 0" shows "count (walks G l) xs = (\i \ {.. {xs. set xs \ verts G \ length xs = (l+1)}" proof (induction l) case 0 - then show ?case by auto + then show ?case by auto next case (Suc l) have "set_mset (walks' G (Suc l)) = (\x\set_mset (walks' G l). (\z. x @ [z]) ` set_mset (vertices_from G (last x)))" by (simp add:set_mset_concat_mset) - also have "... \ (\x\{xs. set xs \ verts G \ length xs = l + 1}. + also have "... \ (\x\{xs. set xs \ verts G \ length xs = l + 1}. (\z. x @ [z]) ` set_mset (vertices_from G (last x)))" by (intro Union_mono image_mono Suc) also have "... \ (\x\{xs. set xs \ verts G \ length xs = l + 1}. (\z. x @ [z]) ` verts G)" by (intro Union_image_mono image_mono set_mset_vertices_from) also have "... \ {xs. set xs \ verts G \ length xs = (Suc l + 1)}" by (intro subsetI) auto finally show ?case by simp qed lemma set_walks: "set_mset (walks G l) \ {xs. set xs \ verts G \ length xs = l}" unfolding walks_def using set_walks' by (cases l, auto) lemma set_walks_2: assumes "xs \# walks' G l" shows "set xs \ verts G" "xs \ []" proof - have a:"xs \ set_mset (walks' G l)" using assms by simp thus "set xs \ verts G" using set_walks' by auto have "length xs \ 0" using set_walks' a by fastforce thus "xs \ []" by simp qed lemma set_walks_3: assumes "xs \# walks G l" shows "set xs \ verts G" "length xs = l" using set_walks assms by auto end lemma measure_pmf_of_multiset: assumes "A \ {#}" - shows "measure (pmf_of_multiset A) S = real (size (filter_mset (\x. x \ S) A)) / size A" + shows "measure (pmf_of_multiset A) S = real (size (filter_mset (\x. x \ S) A)) / size A" (is "?L = ?R") proof - have "sum (count A) (S \ set_mset A) = size (filter_mset (\x. x \ S \ set_mset A) A)" by (intro sum_count_2) simp also have "... = size (filter_mset (\x. x \ S) A)" by (intro arg_cong[where f="size"] filter_mset_cong) auto - finally have a: "sum (count A) (S \ set_mset A) = size (filter_mset (\x. x \ S) A)" + finally have a: "sum (count A) (S \ set_mset A) = size (filter_mset (\x. x \ S) A)" by simp have "?L = measure (pmf_of_multiset A) (S \ set_mset A)" using assms by (intro measure_eq_AE AE_pmfI) auto also have "... = sum (pmf (pmf_of_multiset A)) (S \ set_mset A)" by (intro measure_measure_pmf_finite) simp also have "... = (\x \ S \ set_mset A. count A x / size A)" using assms by (intro sum.cong, auto) also have "... = (\x \ S \ set_mset A. count A x) / size A" by (simp add:sum_divide_distrib) - also have "... = ?R" + also have "... = ?R" using a by simp finally show ?thesis by simp qed lemma pmf_of_multiset_image_mset: assumes "A \ {#}" shows "pmf_of_multiset (image_mset f A) = map_pmf f (pmf_of_multiset A)" - using assms by (intro pmf_eqI) (simp add:pmf_map measure_pmf_of_multiset count_mset_exp + using assms by (intro pmf_eqI) (simp add:pmf_map measure_pmf_of_multiset count_mset_exp image_mset_filter_mset_swap[symmetric]) -context regular_graph +context regular_graph begin lemma size_walks': "size (walks' G l) = card (verts G) * d^l" proof (induction l) case 0 then show ?case by simp next case (Suc l) have a:"out_degree G (last x) = d" if "x \# walks' G l" for x proof - - have "last x \ verts G" + have "last x \ verts G" using set_walks_2 that by fastforce thus ?thesis using reg by simp qed have "size (walks' G (Suc l)) = (\x\#walks' G l. out_degree G (last x))" by (simp add:size_concat_mset image_mset.compositionality comp_def verts_from_alt out_degree_def) also have "... = (\x\#walks' G l. d)" by (intro arg_cong[where f="sum_mset"] image_mset_cong a) simp also have "... = size (walks' G l) * d" by simp also have "... = card (verts G) * d^(Suc l)" using Suc by simp finally show ?case by simp qed lemma size_walks: "size (walks G l) = (if l > 0 then n * d^(l-1) else 1)" using size_walks' unfolding walks_def n_def by (cases l, auto) -lemma walks_nonempty: +lemma walks_nonempty: "walks G l \ {#}" proof - have "size (walks G l) > 0" - unfolding size_walks using d_gt_0 n_gt_0 by auto + unfolding size_walks using d_gt_0 n_gt_0 by auto thus "walks G l \ {#}" by auto qed end context regular_graph_tts begin lemma g_step_remains_orth: assumes "g_inner f (\_. 1) = 0" shows "g_inner (g_step f) (\_. 1) = 0" (is "?L = ?R") proof - have "?L = (A *v (\ i. f (enum_verts i))) \ 1" unfolding g_inner_conv g_step_conv one_vec_def by simp also have "... = (\ i. f (enum_verts i)) \ 1" by (intro markov_orth_inv markov) also have "... = g_inner f (\_. 1)" unfolding g_inner_conv one_vec_def by simp also have "... = 0" using assms by simp finally show ?thesis by simp qed lemma spec_bound: "spec_bound A \\<^sub>a" proof - have "norm (A *v v) \ \\<^sub>a * norm v" if "v \ 1 = (0::real)" for v::"real^'n" unfolding \\<^sub>e_eq_\ by (intro \\<^sub>a_real_bound that) thus ?thesis unfolding spec_bound_def using \_ge_0 by auto qed -text \A spectral expansion rule that does not require orthogonality of the vector for the stationary +text \A spectral expansion rule that does not require orthogonality of the vector for the stationary distribution:\ lemma expansionD3: "\g_inner f (g_step f)\ \ \\<^sub>a * g_norm f^2 + (1-\\<^sub>a) * g_inner f (\_. 1)^2 / n" (is "?L \ ?R") proof - define v where "v = (\ i. f (enum_verts i))" define v1 :: "real^ 'n" where "v1 = ((v \ 1) / n) *\<^sub>R 1" define v2 :: "real^ 'n" where "v2 = v - v1" have v_eq: "v = v1 + v2" unfolding v2_def by simp - have 0: "A *v v1 = v1" - unfolding v1_def using markov_apply[OF markov] + have 0: "A *v v1 = v1" + unfolding v1_def using markov_apply[OF markov] by (simp add:algebra_simps) - have 1: "v1 v* A = v1" - unfolding v1_def using markov_apply[OF markov] + have 1: "v1 v* A = v1" + unfolding v1_def using markov_apply[OF markov] by (simp add:algebra_simps scaleR_vector_matrix_assoc) have "v2 \ 1 = v \ 1 - v1 \ 1" unfolding v2_def by (simp add:algebra_simps) also have "... = v \ 1 - v \ 1 * real CARD('n) / real n" unfolding v1_def by (simp add:inner_1_1) also have "... = 0 " using verts_non_empty unfolding card n_def by simp finally have 4:"v2 \ 1 = 0" by simp hence 2: "v1 \ v2 = 0" unfolding v1_def by (simp add:inner_commute) define f2 where "f2 i = v2 $ (enum_verts_inv i)" for i have f2_def: "v2 = (\ i. f2 (enum_verts i))" unfolding f2_def Rep_inverse by simp have 6: "g_inner f2 (\_. 1) = 0" unfolding g_inner_conv f2_def[symmetric] one_vec_def[symmetric] 4 by simp have "\v2 \ (A *v v2)\ = \g_inner f2 (g_step f2)\" unfolding f2_def g_inner_conv g_step_conv by simp also have "... \ \\<^sub>a * (g_norm f2)\<^sup>2" - by (intro expansionD1 6) + by (intro expansionD1 6) also have "... = \\<^sub>a * (norm v2)^2" unfolding g_norm_conv f2_def by simp finally have 5:"\v2 \ (A *v v2)\ \ \\<^sub>a * (norm v2)\<^sup>2" by simp have 3: "norm (1 :: real^'n)^2 = n" - unfolding power2_norm_eq_inner inner_1_1 card n_def by presburger + unfolding power2_norm_eq_inner inner_1_1 card n_def by presburger have "?L = \v \ (A *v v)\" unfolding g_inner_conv g_step_conv v_def by simp also have "... = \v1 \ (A *v v1) + v2 \ (A *v v1) + v1 \ (A *v v2) + v2 \ (A *v v2)\" unfolding v_eq by (simp add:algebra_simps) also have "... = \v1 \ v1 + v2 \ v1 + v1 \ v2 + v2 \ (A *v v2)\" unfolding dot_lmul_matrix[where x="v1",symmetric] 0 1 by simp also have "... = \v1 \ v1 + v2 \ (A *v v2)\" using 2 by (simp add:inner_commute) also have "... \ \norm v1^2\ + \v2 \ (A *v v2)\" unfolding power2_norm_eq_inner by (intro abs_triangle_ineq) also have "... \ norm v1^2 + \\<^sub>a * norm v2^2" by (intro add_mono 5) auto also have "... = \\<^sub>a * (norm v1^2 + norm v2^2) + (1 - \\<^sub>a) * norm v1^2" by (simp add:algebra_simps) also have "... = \\<^sub>a * norm v^2 + (1 - \\<^sub>a) * norm v1^2" unfolding v_eq pythagoras[OF 2] by simp also have "... = \\<^sub>a * norm v^2 + ((1 - \\<^sub>a)) * ((v \ 1)^2*n)/n^2" unfolding v1_def by (simp add:power_divide power_mult_distrib 3) also have "... = \\<^sub>a * norm v^2 + ((1 - \\<^sub>a)/n) * (v \ 1)^2" by (simp add:power2_eq_square) also have "... = ?R" unfolding g_norm_conv g_inner_conv v_def one_vec_def by (simp add:field_simps) finally show ?thesis by simp qed definition ind_mat where "ind_mat S = diag (ind_vec (enum_verts -` S))" lemma walk_distr: "measure (pmf_of_multiset (walks G l)) {\. (\i ! i \ S i)} = - foldl (\x M. M *v x) stat (intersperse A (map (\i. ind_mat (S i)) [0..1" + foldl (\x M. M *v x) stat (intersperse A (map (\i. ind_mat (S i)) [0..1" (is "?L = ?R") proof (cases "l > 0") case True let ?n = "real n" let ?d = "real d" let ?W = "{(w::'a list). set w \ verts G \ length w = l}" let ?V = "{(w::'n list). length w = l}" have a: "set_mset (walks G l) \ ?W" using set_walks by auto have b: "finite ?W" by (intro finite_lists_length_eq) auto define lp where "lp = l - 1" define xs where "xs = map (\i. ind_mat (S i)) [0.. []" unfolding xs_def using True by simp then obtain xh xt where xh_xt: "xh#xt=xs" by (cases xs, auto) have "length xs = l" unfolding xs_def by simp - hence len_xt: "length xt = lp" + hence len_xt: "length xt = lp" using True unfolding xh_xt[symmetric] lp_def by simp - have "xh = xs ! 0" + have "xh = xs ! 0" unfolding xh_xt[symmetric] by simp also have "... = ind_mat (S 0)" using True unfolding xs_def by simp - finally have xh_eq: "xh = ind_mat (S 0)" + finally have xh_eq: "xh = ind_mat (S 0)" by simp have inj_map_enum_verts: "inj_on (map enum_verts) ?V" using bij_betw_imp_inj_on[OF enum_verts] inj_on_subset by (intro inj_on_mapI) auto have "card ?W = card (verts G)^l" by (intro card_lists_length_eq) simp also have "... = card {w. set w \ (UNIV :: 'n set) \ length w = l}" unfolding card[symmetric] by (intro card_lists_length_eq[symmetric]) simp also have "... = card ?V" by (intro arg_cong[where f="card"]) auto also have "... = card (map enum_verts ` ?V)" by (intro card_image[symmetric] inj_map_enum_verts) finally have "card ?W = card (map enum_verts ` ?V)" by simp hence "map enum_verts ` ?V = ?W" using bij_betw_apply[OF enum_verts] by (intro card_subset_eq b image_subsetI) auto hence bij_map_enum_verts: "bij_betw (map enum_verts) ?V ?W" using inj_map_enum_verts unfolding bij_betw_def by auto have "?L = size {# w \# walks G l. \i S i #} / (?n * ?d^(l-1))" using True unfolding size_walks measure_pmf_of_multiset[OF walks_nonempty] by simp also have "... = (\w\?W. real (count (walks G l) w) * of_bool (\i S i))/(?n*?d^(l-1))" unfolding size_filter_mset_conv sum_mset_conv_2[OF a b] by simp - also have "... = (\w\?W. (\iw\?W. (\ii S i)))/(?n*?d^(l-1))" using True by (intro sum.cong arg_cong2[where f="(/)"]) (auto simp add: count_walks) - also have "... = + also have "... = (\w\?W. (\ii S i)))/?n" using True unfolding prod_dividef by (simp add:sum_divide_distrib algebra_simps) - also have "... = - (\w\?V. (\iw\?V. (\ii S i)))/?n" by (intro sum.reindex_bij_betw[symmetric] arg_cong2[where f="(/)"] refl bij_map_enum_verts) - also have "... = + also have "... = (\w\?V. (\ii S i)))/?n" unfolding A_def lp_def using True by simp - also have "... = (\w\?V. (\iw\?V. (\ii\insert 0 (Suc ` {.. S i)))/?n" using lessThan_Suc_eq_insert_0 by (intro sum.cong arg_cong2[where f="(/)"] arg_cong2[where f="(*)"] prod.cong) auto - also have "... = (\w\?V. (\iS(i+1))* A$ w!(i+1) $ w!i) + also have "... = (\w\?V. (\iS(i+1))* A$ w!(i+1) $ w!i) * of_bool(enum_verts(w!0)\S 0))/?n" by (simp add:prod.reindex algebra_simps prod.distrib) - also have "... = + also have "... = (\w\?V. (\iS 0))/?n" unfolding diag_def ind_vec_def matrix_matrix_mult_def ind_mat_def - by (intro sum.cong arg_cong2[where f="(/)"] arg_cong2[where f="(*)"] prod.cong refl) + by (intro sum.cong arg_cong2[where f="(/)"] arg_cong2[where f="(*)"] prod.cong refl) (simp add:if_distrib if_distribR sum.If_cases) - also have "... = + also have "... = (\w\?V. (\iS 0))/?n" unfolding xs_def lp_def True - by (intro sum.cong arg_cong2[where f="(/)"] arg_cong2[where f="(*)"] prod.cong refl) auto - also have "... = + by (intro sum.cong arg_cong2[where f="(/)"] arg_cong2[where f="(*)"] prod.cong refl) auto + also have "... = (\w\?V. (\iS 0))/?n" unfolding xh_xt[symmetric] by auto also have "... = (\w\?V. (\iw\?V. (\ix M. M *v x) (xh *v stat) (map (\x. x ** A) xt) \ 1" + also have "... = foldl (\x M. M *v x) (xh *v stat) (map (\x. x ** A) xt) \ 1" using True unfolding foldl_matrix_mult_expand_2 by (simp add:len_xt lp_def) - also have "... = foldl (\x M. M *v (A *v x)) (xh *v stat) xt \ 1" + also have "... = foldl (\x M. M *v (A *v x)) (xh *v stat) xt \ 1" by (simp add: matrix_vector_mul_assoc foldl_map) - also have "... = foldl (\x M. M *v x) stat (intersperse A (xh#xt)) \ 1" + also have "... = foldl (\x M. M *v x) stat (intersperse A (xh#xt)) \ 1" by (subst foldl_intersperse_2, simp) also have "... = ?R" unfolding xh_xt xs_def by simp finally show ?thesis by simp next case False hence "l = 0" by simp thus ?thesis unfolding stat_def by (simp add: inner_1_1) qed lemma hitting_property: - assumes "S \ verts G" + assumes "S \ verts G" assumes "I \ {.. \ real (card S) / card (verts G)" - shows "measure (pmf_of_multiset (walks G l)) {w. set (nths w I) \ S} \ (\+\\<^sub>a*(1-\))^card I" + shows "measure (pmf_of_multiset (walks G l)) {w. set (nths w I) \ S} \ (\+\\<^sub>a*(1-\))^card I" (is "?L \ ?R") proof - - define T where "T = (\i. if i \ I then S else UNIV)" + define T where "T = (\i. if i \ I then S else UNIV)" have 0: "ind_mat UNIV = mat 1" unfolding ind_mat_def diag_def ind_vec_def Finite_Cartesian_Product.mat_def by vector have \_range: "\\<^sub>a \ {0..1}" using \_ge_0 \_le_1 by simp - have "S \ range enum_verts" + have "S \ range enum_verts" using assms(1) enum_verts unfolding bij_betw_def by simp - moreover have "inj enum_verts" + moreover have "inj enum_verts" using bij_betw_imp_inj_on[OF enum_verts] by simp ultimately have \_alt: "\ = real (card (enum_verts -` S)) / CARD ('n)" unfolding \_def card by (subst card_vimage_inj) auto have "?L = measure (pmf_of_multiset (walks G l)) {w. \i T i}" using walks_nonempty set_walks_3 unfolding T_def set_nths by (intro measure_eq_AE AE_pmfI) auto - also have "... = foldl (\x M. M *v x) stat + also have "... = foldl (\x M. M *v x) stat (intersperse A (map (\i. (if i \ I then ind_mat S else mat 1)) [0.. 1" unfolding walk_distr T_def by (simp add:if_distrib if_distribR 0 cong:if_cong) also have "... \ ?R" unfolding \_alt ind_mat_def by (intro hitting_property_alg_2[OF \_range assms(2) spec_bound markov]) finally show ?thesis by simp qed lemma uniform_property: assumes "i < l" "x \ verts G" - shows "measure (pmf_of_multiset (walks G l)) {w. w ! i = x} = 1/real (card (verts G))" + shows "measure (pmf_of_multiset (walks G l)) {w. w ! i = x} = 1/real (card (verts G))" (is "?L = ?R") proof - - obtain xi where xi_def: "enum_verts xi = x" + obtain xi where xi_def: "enum_verts xi = x" using assms(2) bij_betw_imp_surj_on[OF enum_verts] by force define T where "T = (\j. if j = i then {x} else UNIV)" have "diag (ind_vec UNIV) = mat 1" unfolding diag_def ind_vec_def Finite_Cartesian_Product.mat_def by vector moreover have "enum_verts -` {x} = {xi}" using bij_betw_imp_inj_on[OF enum_verts] unfolding vimage_def xi_def[symmetric] by (auto simp add:inj_on_def) ultimately have 0: "ind_mat (T j) = (if j = i then diag (ind_vec {xi}) else mat 1)" for j unfolding T_def ind_mat_def by (cases "j = i", auto) have "?L = measure (pmf_of_multiset (walks G l)) {w. \j < l. w ! j \ T j}" unfolding T_def using assms(1) by simp also have "... = foldl (\x M. M *v x) stat (intersperse A (map (\j. ind_mat (T j)) [0.. 1" unfolding walk_distr by simp also have "... = 1/CARD('n)" unfolding 0 uniform_property_alg[OF assms(1) markov] by simp - also have "... = ?R" + also have "... = ?R" unfolding card by simp finally show ?thesis by simp qed end context regular_graph begin -lemmas expansionD3 = +lemmas expansionD3 = regular_graph_tts.expansionD3[OF eg_tts_1, - internalize_sort "'n :: finite", OF _ regular_graph_axioms, + internalize_sort "'n :: finite", OF _ regular_graph_axioms, unfolded remove_finite_premise, cancel_type_definition, OF verts_non_empty] -lemmas g_step_remains_orth = +lemmas g_step_remains_orth = regular_graph_tts.g_step_remains_orth[OF eg_tts_1, - internalize_sort "'n :: finite", OF _ regular_graph_axioms, + internalize_sort "'n :: finite", OF _ regular_graph_axioms, unfolded remove_finite_premise, cancel_type_definition, OF verts_non_empty] -lemmas hitting_property = +lemmas hitting_property = regular_graph_tts.hitting_property[OF eg_tts_1, - internalize_sort "'n :: finite", OF _ regular_graph_axioms, + internalize_sort "'n :: finite", OF _ regular_graph_axioms, unfolded remove_finite_premise, cancel_type_definition, OF verts_non_empty] -lemmas uniform_property_2 = +lemmas uniform_property_2 = regular_graph_tts.uniform_property[OF eg_tts_1, - internalize_sort "'n :: finite", OF _ regular_graph_axioms, + internalize_sort "'n :: finite", OF _ regular_graph_axioms, unfolded remove_finite_premise, cancel_type_definition, OF verts_non_empty] theorem uniform_property: assumes "i < l" shows "map_pmf (\w. w ! i) (pmf_of_multiset (walks G l)) = pmf_of_set (verts G)" (is "?L = ?R") proof (rule pmf_eqI) fix x :: "'a" have a:"measure (pmf_of_multiset (walks G l)) {w. w ! i = x} = 0" (is "?L1 = ?R1") if "x \ verts G" proof - have "?L1 \ measure (pmf_of_multiset (walks G l)) {w. set w \ verts G \ x \ set w}" using walks_nonempty set_walks_3 assms(1) - by (intro measure_pmf.pmf_mono[OF refl]) auto + by (intro pmf_mono) auto also have "... \ measure (pmf_of_multiset (walks G l)) {}" - using that by (intro measure_pmf.pmf_mono[OF refl]) auto + using that by (intro pmf_mono) auto also have "... = 0" by simp finally have "?L1 \ 0" by simp thus ?thesis using measure_le_0_iff by blast qed have "pmf ?L x = measure (pmf_of_multiset (walks G l)) {w. w ! i = x}" unfolding pmf_map by (simp add:vimage_def) also have "... = indicator (verts G) x/real (card (verts G))" using uniform_property_2[OF assms(1)] a by (cases "x \ verts G", auto) also have "... = pmf ?R x" - using verts_non_empty by (intro pmf_of_set[symmetric]) auto + using verts_non_empty by (intro pmf_of_set[symmetric]) auto finally show "pmf ?L x = pmf ?R x" by simp qed lemma uniform_property_gen: fixes S :: "'a set" assumes "S \ verts G" "i < l" defines "\ \ real (card S) / card (verts G)" shows "measure (pmf_of_multiset (walks G l)) {w. w ! i \ S} = \" (is "?L = ?R") proof - have "?L = measure (map_pmf (\w. w ! i) (pmf_of_multiset (walks G l))) S" unfolding measure_map_pmf by (simp add:vimage_def) also have "... = measure (pmf_of_set (verts G)) S" unfolding uniform_property[OF assms(2)] by simp also have "... = ?R" - using verts_non_empty Int_absorb1[OF assms(1)] - unfolding \_def by (subst measure_pmf_of_set) auto + using verts_non_empty Int_absorb1[OF assms(1)] + unfolding \_def by (subst measure_pmf_of_set) auto finally show ?thesis by simp qed theorem kl_chernoff_property: assumes "l > 0" assumes "S \ verts G" defines "\ \ real (card S) / card (verts G)" assumes "\ \ 1" "\ + \\<^sub>a * (1-\) \ {0<..\}" - shows "measure (pmf_of_multiset (walks G l)) {w. real (card {i \ {.. S}) \ \*l} + shows "measure (pmf_of_multiset (walks G l)) {w. real (card {i \ {.. S}) \ \*l} \ exp (- real l * KL_div \ (\+\\<^sub>a*(1-\)))" (is "?L \ ?R") proof - let ?\ = "(\i+\\<^sub>a*(1-\))/l" have a: "measure (pmf_of_multiset (walks G l)) {w. \i\T. w ! i \ S} \ (\ + \\<^sub>a*(1-\)) ^ card T" - (is "?L1 \ ?R1") if "T \ {.. ?R1") if "T \ {.. S}" unfolding set_nths setcompr_eq_image using that set_walks_3 walks_nonempty by (intro measure_eq_AE AE_pmfI) (auto simp add:image_subset_iff) also have "... \ ?R1" unfolding \_def by (intro hitting_property[OF assms(2) that]) finally show ?thesis by simp qed - + have "?L \ exp ( - real l * KL_div \ ?\)" using assms(1,4,5) a by (intro impagliazzo_kabanets_pmf) simp_all also have "... = ?R" by simp finally show ?thesis by simp qed end unbundle no_intro_cong_syntax end diff --git a/thys/Expander_Graphs/Extra_Congruence_Method.thy b/thys/Expander_Graphs/Extra_Congruence_Method.thy --- a/thys/Expander_Graphs/Extra_Congruence_Method.thy +++ b/thys/Expander_Graphs/Extra_Congruence_Method.thy @@ -1,57 +1,57 @@ subsection \Congruence Method\ text \The following is a method for proving equalities of large terms by checking the equivalence of subterms. It is possible to precisely control which operators to split by.\ theory Extra_Congruence_Method - imports - Main + imports + Main "HOL-Eisbach.Eisbach" begin datatype cong_tag_type = CongTag -definition cong_tag_1 :: "('a \ 'b) \ cong_tag_type" +definition cong_tag_1 :: "('a \ 'b) \ cong_tag_type" where "cong_tag_1 x = CongTag" definition cong_tag_2 :: "('a \ 'b \ 'c) \ cong_tag_type" where "cong_tag_2 x = CongTag" -definition cong_tag_3 :: "('a \ 'b \ 'c \ 'd) \ cong_tag_type" +definition cong_tag_3 :: "('a \ 'b \ 'c \ 'd) \ cong_tag_type" where "cong_tag_3 x = CongTag" lemma arg_cong3: assumes "x1 = x2" "y1 = y2" "z1 = z2" shows "f x1 y1 z1 = f x2 y2 z2" using assms by auto -method intro_cong for A :: "cong_tag_type list" uses more = - (match (A) in - "cong_tag_1 f#h" (multi) for f :: "'a \ 'b" and h +method intro_cong for A :: "cong_tag_type list" uses more = + (match (A) in + "cong_tag_1 f#h" (multi) for f :: "'a \ 'b" and h \ \intro_cong h more:more arg_cong[where f="f"]\ - \ "cong_tag_2 f#h" (multi) for f :: "'a \ 'b \ 'c" and h + \ "cong_tag_2 f#h" (multi) for f :: "'a \ 'b \ 'c" and h \ \intro_cong h more:more arg_cong2[where f="f"]\ - \ "cong_tag_3 f#h" (multi) for f :: "'a \ 'b \ 'c \ 'd" and h + \ "cong_tag_3 f#h" (multi) for f :: "'a \ 'b \ 'c \ 'd" and h \ \intro_cong h more:more arg_cong3[where f="f"]\ \ _ \ \intro more refl\) bundle intro_cong_syntax begin notation cong_tag_1 ("\\<^sub>1") notation cong_tag_2 ("\\<^sub>2") notation cong_tag_3 ("\\<^sub>3") end bundle no_intro_cong_syntax begin no_notation cong_tag_1 ("\\<^sub>1") no_notation cong_tag_2 ("\\<^sub>2") no_notation cong_tag_3 ("\\<^sub>3") end lemma restr_Collect_cong: assumes "\x. x \ A \ P x = Q x" shows "{x \ A. P x} = {x \ A. Q x}" using assms by auto end diff --git a/thys/Expander_Graphs/ROOT b/thys/Expander_Graphs/ROOT --- a/thys/Expander_Graphs/ROOT +++ b/thys/Expander_Graphs/ROOT @@ -1,25 +1,26 @@ chapter AFP -session Expander_Graphs = Frequency_Moments + +session Expander_Graphs = "HOL-Probability" + options [timeout = 1200] sessions Graph_Theory Perron_Frobenius Commuting_Hermitian Weighted_Arithmetic_Geometric_Mean + Universal_Hash_Families theories Constructive_Chernoff_Bound Expander_Graphs_Algebra Expander_Graphs_Cheeger_Inequality Expander_Graphs_Definition Expander_Graphs_Eigenvalues Expander_Graphs_MGG Expander_Graphs_Multiset_Extras Expander_Graphs_Power_Construction Expander_Graphs_Strongly_Explicit Expander_Graphs_TTS Expander_Graphs_Walks Extra_Congruence_Method document_files "root.tex" "root.bib" diff --git a/thys/Frequency_Moments/Frequency_Moment_0.thy b/thys/Frequency_Moments/Frequency_Moment_0.thy --- a/thys/Frequency_Moments/Frequency_Moment_0.thy +++ b/thys/Frequency_Moments/Frequency_Moment_0.thy @@ -1,1314 +1,1315 @@ section \Frequency Moment $0$\label{sec:f0}\ theory Frequency_Moment_0 imports Frequency_Moments_Preliminary_Results - Median_Method.Median - K_Smallest + Median_Method.Median + K_Smallest Universal_Hash_Families.Carter_Wegman_Hash_Family Frequency_Moments - Landau_Ext + Landau_Ext + Probability_Ext Product_PMF_Ext Universal_Hash_Families.Universal_Hash_Families_More_Finite_Fields begin text \This section contains a formalization of a new algorithm for the zero-th frequency moment inspired by ideas described in \<^cite>\"baryossef2002"\. -It is a KMV-type ($k$-minimum value) algorithm with a rounding method and matches the space complexity +It is a KMV-type ($k$-minimum value) algorithm with a rounding method and matches the space complexity of the best algorithm described in \<^cite>\"baryossef2002"\. In addition to the Isabelle proof here, there is also an informal hand-written proof in Appendix~\ref{sec:f0_proof}.\ type_synonym f0_state = "nat \ nat \ nat \ nat \ (nat \ nat list) \ (nat \ float set)" definition hash where "hash p = ring.hash (mod_ring p)" fun f0_init :: "rat \ rat \ nat \ f0_state pmf" where "f0_init \ \ n = do { let s = nat \-18 * ln (real_of_rat \)\; let t = nat \80 / (real_of_rat \)\<^sup>2\; let p = prime_above (max n 19); - let r = nat (4 * \log 2 (1 / real_of_rat \)\ + 23); + let r = nat (4 * \log 2 (1 / real_of_rat \)\ + 23); h \ prod_pmf {.._. pmf_of_set (bounded_degree_polynomials (mod_ring p) 2)); return_pmf (s, t, p, r, h, (\_ \ {0.. f0_state \ f0_state pmf" where - "f0_update x (s, t, p, r, h, sketch) = + "f0_update x (s, t, p, r, h, sketch) = return_pmf (s, t, p, r, h, \i \ {.. rat pmf" where "f0_result (s, t, p, r, h, sketch) = return_pmf (median s (\i \ {.. rat \ rat) \ real" where "f0_space_usage (n, \, \) = ( - let s = nat \-18 * ln (real_of_rat \)\ in + let s = nat \-18 * ln (real_of_rat \)\ in let r = nat (4 * \log 2 (1 / real_of_rat \)\ + 23) in let t = nat \80 / (real_of_rat \)\<^sup>2 \ in 6 + 2 * log 2 (real s + 1) + 2 * log 2 (real t + 1) + 2 * log 2 (real n + 21) + 2 * log 2 (real r + 1) + real s * (5 + 2 * log 2 (21 + real n) + real t * (13 + 4 * r + 2 * log 2 (log 2 (real n + 13)))))" definition encode_f0_state :: "f0_state \ bool list option" where - "encode_f0_state = - N\<^sub>e \\<^sub>e (\s. + "encode_f0_state = + N\<^sub>e \\<^sub>e (\s. N\<^sub>e \\<^sub>e ( - N\<^sub>e \\<^sub>e (\p. - N\<^sub>e \\<^sub>e ( + N\<^sub>e \\<^sub>e (\p. + N\<^sub>e \\<^sub>e ( ([0..\<^sub>e (P\<^sub>e p 2)) \\<^sub>e ([0..\<^sub>e (S\<^sub>e F\<^sub>e))))))" lemma "inj_on encode_f0_state (dom encode_f0_state)" proof - - have "is_encoding encode_f0_state" + have "is_encoding encode_f0_state" unfolding encode_f0_state_def by (intro dependent_encoding exp_golomb_encoding poly_encoding fun_encoding set_encoding float_encoding) thus ?thesis by (rule encoding_imp_inj) qed context fixes \ \ :: rat fixes n :: nat fixes as :: "nat list" fixes result assumes \_range: "\ \ {0<..<1}" assumes \_range: "\ \ {0<..<1}" assumes as_range: "set as \ {.. fold (\a state. state \ f0_update a) as (f0_init \ \ n) \ f0_result" -begin +begin private definition t where "t = nat \80 / (real_of_rat \)\<^sup>2\" private lemma t_gt_0: "t > 0" using \_range by (simp add:t_def) private definition s where "s = nat \-(18 * ln (real_of_rat \))\" private lemma s_gt_0: "s > 0" using \_range by (simp add:s_def) private definition p where "p = prime_above (max n 19)" private lemma p_prime:"Factorial_Ring.prime p" using p_def prime_above_prime by presburger private lemma p_ge_18: "p \ 18" proof - - have "p \ 19" + have "p \ 19" by (metis p_def prime_above_lower_bound max.bounded_iff) thus ?thesis by simp qed private lemma p_gt_0: "p > 0" using p_ge_18 by simp private lemma p_gt_1: "p > 1" using p_ge_18 by simp private lemma n_le_p: "n \ p" proof - have "n \ max n 19" by simp also have "... \ p" unfolding p_def by (rule prime_above_lower_bound) finally show ?thesis by simp qed private lemma p_le_n: "p \ 2*n + 40" proof - have "p \ 2 * (max n 19) + 2" by (subst p_def, rule prime_above_upper_bound) also have "... \ 2 * n + 40" by (cases "n \ 19", auto) finally show ?thesis by simp qed -private lemma as_lt_p: "\x. x \ set as \ x < p" +private lemma as_lt_p: "\x. x \ set as \ x < p" using as_range atLeastLessThan_iff by (intro order_less_le_trans[OF _ n_le_p]) blast private lemma as_subset_p: "set as \ {..log 2 (1 / real_of_rat \)\ + 23)" private lemma r_bound: "4 * log 2 (1 / real_of_rat \) + 23 \ r" proof - - have "0 \ log 2 (1 / real_of_rat \)" using \_range by simp + have "0 \ log 2 (1 / real_of_rat \)" using \_range by simp hence "0 \ \log 2 (1 / real_of_rat \)\" by simp hence "0 \ 4 * \log 2 (1 / real_of_rat \)\ + 23" by (intro add_nonneg_nonneg mult_nonneg_nonneg, auto) thus ?thesis by (simp add:r_def) qed private lemma r_ge_23: "r \ 23" proof - have "(23::real) = 0 + 23" by simp - also have "... \ 4 * log 2 (1 / real_of_rat \) + 23" - using \_range by (intro add_mono mult_nonneg_nonneg, auto) + also have "... \ 4 * log 2 (1 / real_of_rat \) + 23" + using \_range by (intro add_mono mult_nonneg_nonneg, auto) also have "... \ r" using r_bound by simp finally show "23 \ r" by simp qed private lemma two_pow_r_le_1: "0 < 1 - 2 powr - real r" proof - have a: "2 powr (0::real) = 1" by simp - show ?thesis using r_ge_23 + show ?thesis using r_ge_23 by (simp, subst a[symmetric], intro powr_less_mono, auto) qed interpretation carter_wegman_hash_family "mod_ring p" 2 rewrites "ring.hash (mod_ring p) = Frequency_Moment_0.hash p" using carter_wegman_hash_familyI[OF mod_ring_is_field mod_ring_finite] using hash_def p_prime by auto private definition tr_hash where "tr_hash x \ = truncate_down r (hash x \)" private definition sketch_rv where "sketch_rv \ = least t ((\x. float_of (tr_hash x \)) ` set as)" -private definition estimate +private definition estimate where "estimate S = (if card S < t then of_nat (card S) else of_nat t * of_nat p / rat_of_float (Max S))" private definition sketch_rv' where "sketch_rv' \ = least t ((\x. tr_hash x \) ` set as)" private definition estimate' where "estimate' S = (if card S < t then real (card S) else real t * real p / Max S)" private definition \\<^sub>0 where "\\<^sub>0 = prod_pmf {.._. pmf_of_set space)" private lemma f0_alg_sketch: defines "sketch \ fold (\a state. state \ f0_update a) as (f0_init \ \ n)" - shows "sketch = map_pmf (\x. (s,t,p,r, x, \i \ {..\<^sub>0" - unfolding sketch_rv_def + shows "sketch = map_pmf (\x. (s,t,p,r, x, \i \ {..\<^sub>0" + unfolding sketch_rv_def proof (subst sketch_def, induction as rule:rev_induct) case Nil then show ?case by (simp add:s_def p_def[symmetric] map_pmf_def t_def r_def Let_def least_def restrict_def space_def \\<^sub>0_def) next case (snoc x xs) let ?sketch = "\\ xs. least t ((\a. float_of (tr_hash a \)) ` set xs)" have "fold (\a state. state \ f0_update a) (xs @ [x]) (f0_init \ \ n) = (map_pmf (\\. (s, t, p, r, \, \i \ {.. i) xs)) \\<^sub>0) \ f0_update x" by (simp add: restrict_def snoc del:f0_init.simps) also have "... = \\<^sub>0 \ (\\. f0_update x (s, t, p, r, \, \i\{.. i) xs)) " by (simp add:map_pmf_def bind_assoc_pmf bind_return_pmf del:f0_update.simps) also have "... = map_pmf (\\. (s, t, p, r, \, \i\{.. i) (xs@[x]))) \\<^sub>0" by (simp add:least_insert map_pmf_def tr_hash_def cong:restrict_cong) finally show ?case by blast qed private lemma card_nat_in_ball: fixes x :: nat fixes q :: real assumes "q \ 0" defines "A \ {k. abs (real x - real k) \ q \ k \ x}" shows "real (card A) \ 2 * q" and "finite A" proof - have a: "of_nat x \ {\real x-q\..\real x+q\}" - using assms + using assms by (simp add: ceiling_le_iff) have "card A = card (int ` A)" by (rule card_image[symmetric], simp) also have "... \ card ({\real x-q\..\real x+q\} - {of_nat x})" by (intro card_mono image_subsetI, simp_all add:A_def abs_le_iff, linarith) also have "... = card {\real x-q\..\real x+q\} - 1" by (rule card_Diff_singleton, rule a) also have "... = int (card {\real x-q\..\real x+q\}) - int 1" by (intro of_nat_diff) (metis a card_0_eq empty_iff finite_atLeastAtMost_int less_one linorder_not_le) also have "... \ \q+real x\+1 -\real x-q\ - 1" using assms by (simp, linarith) also have "... \ 2*q" by linarith finally show "card A \ 2 * q" by simp have "A \ {..x + nat \q\}" by (rule subsetI, simp add:A_def abs_le_iff, linarith) thus "finite A" by (rule finite_subset, simp) qed private lemma prob_degree_lt_1: - "prob {\. degree \ < 1} \ 1/real p" + "prob {\. degree \ < 1} \ 1/real p" proof - have "space \ {\. length \ \ Suc 0} = bounded_degree_polynomials (mod_ring p) 1" by (auto simp:set_eq_iff bounded_degree_polynomials_def space_def) moreover have "field_size = p" by (simp add:mod_ring_def) hence "real (card (bounded_degree_polynomials (mod_ring p) (Suc 0))) / real (card space) = 1 / real p" by (simp add:space_def bounded_degree_polynomials_card power2_eq_square) ultimately show ?thesis by (simp add:M_def measure_pmf_of_set) qed private lemma collision_prob: assumes "c \ 1" - shows "prob {\. \x \ set as. \y \ set as. x \ y \ tr_hash x \ \ c \ tr_hash x \ = tr_hash y \} \ + shows "prob {\. \x \ set as. \y \ set as. x \ y \ tr_hash x \ \ c \ tr_hash x \ = tr_hash y \} \ (5/2) * (real (card (set as)))\<^sup>2 * c\<^sup>2 * 2 powr -(real r) / (real p)\<^sup>2 + 1/real p" (is "prob {\. ?l \} \ ?r1 + ?r2") proof - define \ :: real where "\ = 9/8" - have rho_c_ge_0: "\ * c \ 0" unfolding \_def using assms by simp + have rho_c_ge_0: "\ * c \ 0" unfolding \_def using assms by simp have c_ge_0: "c\0" using assms by simp - + have "degree \ \ 1 \ \ \ space \ degree \ = 1" for \ - by (simp add:bounded_degree_polynomials_def space_def) + by (simp add:bounded_degree_polynomials_def space_def) (metis One_nat_def Suc_1 le_less_Suc_eq less_imp_diff_less list.size(3) pos2) - hence a: "\\ x y. x < p \ y < p \ x \ y \ degree \ \ 1 \ \ \ space \ hash x \ \ hash y \" - using inj_onD[OF inj_if_degree_1] mod_ring_carr by blast + hence a: "\\ x y. x < p \ y < p \ x \ y \ degree \ \ 1 \ \ \ space \ hash x \ \ hash y \" + using inj_onD[OF inj_if_degree_1] mod_ring_carr by blast have b: "prob {\. degree \ \ 1 \ tr_hash x \ \ c \ tr_hash x \ = tr_hash y \} \ 5 * c\<^sup>2 * 2 powr (-real r) /(real p)\<^sup>2" if b_assms: "x \ set as" "y \ set as" "x < y" for x y proof - have c: "real u \ \ * c \ \real u - real v\ \ \ * c * 2 powr (-real r)" if c_assms:"truncate_down r (real u) \ c" "truncate_down r (real u) = truncate_down r (real v)" for u v proof - - have "9 * 2 powr - real r \ 9 * 2 powr (- real 23)" + have "9 * 2 powr - real r \ 9 * 2 powr (- real 23)" using r_ge_23 by (intro mult_left_mono powr_mono, auto) also have "... \ 1" by simp finally have "9 * 2 powr - real r \ 1" by simp - hence "1 \ \ * (1 - 2 powr (- real r))" + hence "1 \ \ * (1 - 2 powr (- real r))" by (simp add:\_def) - hence d: "(c*1) / (1 - 2 powr (-real r)) \ c * \" + hence d: "(c*1) / (1 - 2 powr (-real r)) \ c * \" using assms two_pow_r_le_1 by (simp add: pos_divide_le_eq) - have "\x. truncate_down r (real x) \ c \ real x * (1 - 2 powr - real r) \ c * 1" + have "\x. truncate_down r (real x) \ c \ real x * (1 - 2 powr - real r) \ c * 1" using truncate_down_pos[OF of_nat_0_le_iff] order_trans by (simp, blast) hence "\x. truncate_down r (real x) \ c \ real x \ c * \" - using two_pow_r_le_1 by (intro order_trans[OF _ d], simp add: pos_le_divide_eq) + using two_pow_r_le_1 by (intro order_trans[OF _ d], simp add: pos_le_divide_eq) - hence e: "real u \ c * \" "real v \ c * \" + hence e: "real u \ c * \" "real v \ c * \" using c_assms by auto have " \real u - real v\ \ (max \real u\ \real v\) * 2 powr (-real r)" using c_assms by (intro truncate_down_eq, simp) also have "... \ (c * \) * 2 powr (-real r)" using e by (intro mult_right_mono, auto) finally have "\real u - real v\ \ \ * c * 2 powr (-real r)" by (simp add:algebra_simps) thus ?thesis using e by (simp add:algebra_simps) qed have "prob {\. degree \ \ 1 \ tr_hash x \ \ c \ tr_hash x \ = tr_hash y \} \ prob (\ i \ {(u,v) \ {.. {.. v \ truncate_down r u \ c \ truncate_down r u = truncate_down r v}. {\. hash x \ = fst i \ hash y \ = snd i})" - using a by (intro pmf_mono[OF M_def], simp add:tr_hash_def) - (metis hash_range mod_ring_carr b_assms as_subset_p lessThan_iff nat_neq_iff subset_eq) + using a by (intro pmf_mono[OF M_def], simp add:tr_hash_def) + (metis hash_range mod_ring_carr b_assms as_subset_p lessThan_iff nat_neq_iff subset_eq) also have "... \ (\ i\ {(u,v) \ {.. {.. v \ - truncate_down r u \ c \ truncate_down r u = truncate_down r v}. + truncate_down r u \ c \ truncate_down r u = truncate_down r v}. prob {\. hash x \ = fst i \ hash y \ = snd i})" by (intro measure_UNION_le finite_cartesian_product finite_subset[where B="{0.. {0.. (\ i\ {(u,v) \ {.. {.. v \ - truncate_down r u \ c \ truncate_down r u = truncate_down r v}. - prob {\. (\u \ {x,y}. hash u \ = (if u = x then (fst i) else (snd i)))})" + truncate_down r u \ c \ truncate_down r u = truncate_down r v}. + prob {\. (\u \ {x,y}. hash u \ = (if u = x then (fst i) else (snd i)))})" by (intro sum_mono pmf_mono[OF M_def]) force also have "... \ (\ i\ {(u,v) \ {.. {.. v \ truncate_down r u \ c \ truncate_down r u = truncate_down r v}. 1/(real p)\<^sup>2)" using assms as_subset_p b_assms by (intro sum_mono, subst hash_prob) (auto simp add: mod_ring_def power2_eq_square) - also have "... = 1/(real p)\<^sup>2 * + also have "... = 1/(real p)\<^sup>2 * card {(u,v) \ {0.. {0.. v \ truncate_down r u \ c \ truncate_down r u = truncate_down r v}" by simp - also have "... \ 1/(real p)\<^sup>2 * + also have "... \ 1/(real p)\<^sup>2 * card {(u,v) \ {.. {.. v \ real u \ \ * c \ abs (real u - real v) \ \ * c * 2 powr (-real r)}" using c by (intro mult_mono of_nat_mono card_mono finite_cartesian_product finite_subset[where B="{..{.. 1/(real p)\<^sup>2 * card (\u' \ {u. u < p \ real u \ \ * c}. {(u::nat,v::nat). u = u' \ abs (real u - real v) \ \ * c * 2 powr (-real r) \ v < p \ v \ u'})" by (intro mult_left_mono of_nat_mono card_mono finite_cartesian_product finite_subset[where B="{..{.. 1/(real p)\<^sup>2 * (\ u' \ {u. u < p \ real u \ \ * c}. card {(u,v). u = u' \ abs (real u - real v) \ \ * c * 2 powr (-real r) \ v < p \ v \ u'})" by (intro mult_left_mono of_nat_mono card_UN_le, auto) also have "... = 1/(real p)\<^sup>2 * (\ u' \ {u. u < p \ real u \ \ * c}. card ((\x. (u' ,x)) ` {v. abs (real u' - real v) \ \ * c * 2 powr (-real r) \ v < p \ v \ u'}))" by (intro arg_cong2[where f="(*)"] arg_cong[where f="real"] sum.cong arg_cong[where f="card"]) (auto simp add:set_eq_iff) also have "... \ 1/(real p)\<^sup>2 * (\ u' \ {u. u < p \ real u \ \ * c}. card {v. abs (real u' - real v) \ \ * c * 2 powr (-real r) \ v < p \ v \ u'})" by (intro mult_left_mono of_nat_mono sum_mono card_image_le, auto) also have "... \ 1/(real p)\<^sup>2 * (\ u' \ {u. u < p \ real u \ \ * c}. card {v. abs (real u' - real v) \ \ * c * 2 powr (-real r) \ v \ u'})" by (intro mult_left_mono sum_mono of_nat_mono card_mono card_nat_in_ball subsetI) auto also have "... \ 1/(real p)\<^sup>2 * (\ u' \ {u. u < p \ real u \ \ * c}. real (card {v. abs (real u' - real v) \ \ * c * 2 powr (-real r) \ v \ u'}))" by simp also have "... \ 1/(real p)\<^sup>2 * (\ u' \ {u. u < p \ real u \ \ * c}. 2 * (\ * c * 2 powr (-real r)))" by (intro mult_left_mono sum_mono card_nat_in_ball(1), auto) also have "... = 1/(real p)\<^sup>2 * (real (card {u. u < p \ real u \ \ * c}) * (2 * (\ * c * 2 powr (-real r))))" by simp also have "... \ 1/(real p)\<^sup>2 * (real (card {u. u \ nat (\\ * c \)}) * (2 * (\ * c * 2 powr (-real r))))" using rho_c_ge_0 le_nat_floor by (intro mult_left_mono mult_right_mono of_nat_mono card_mono subsetI) auto also have "... \ 1/(real p)\<^sup>2 * ((1+\ * c) * (2 * (\ * c * 2 powr (-real r))))" using rho_c_ge_0 by (intro mult_left_mono mult_right_mono, auto) - also have "... \ 1/(real p)\<^sup>2 * (((1+\) * c) * (2 * (\ * c * 2 powr (-real r))))" + also have "... \ 1/(real p)\<^sup>2 * (((1+\) * c) * (2 * (\ * c * 2 powr (-real r))))" using assms by (intro mult_mono, auto simp add:distrib_left distrib_right \_def) also have "... = (\ * (2 + \ * 2)) * c\<^sup>2 * 2 powr (-real r) /(real p)\<^sup>2" - by (simp add:ac_simps power2_eq_square) + by (simp add:ac_simps power2_eq_square) also have "... \ 5 * c\<^sup>2 * 2 powr (-real r) /(real p)\<^sup>2" by (intro divide_right_mono mult_right_mono) (auto simp add:\_def) finally show ?thesis by simp qed - have "prob {\. ?l \ \ degree \ \ 1} \ + have "prob {\. ?l \ \ degree \ \ 1} \ prob (\ i \ {(x,y) \ (set as) \ (set as). x < y}. {\. degree \ \ 1 \ tr_hash (fst i) \ \ c \ tr_hash (fst i) \ = tr_hash (snd i) \})" by (rule pmf_mono[OF M_def], simp, metis linorder_neqE_nat) - also have "... \ (\ i \ {(x,y) \ (set as) \ (set as). x < y}. prob + also have "... \ (\ i \ {(x,y) \ (set as) \ (set as). x < y}. prob {\. degree \ \ 1 \ tr_hash (fst i) \ \ c \ tr_hash (fst i) \ = tr_hash (snd i) \})" unfolding M_def by (intro measure_UNION_le finite_cartesian_product finite_subset[where B="(set as) \ (set as)"]) auto also have "... \ (\ i \ {(x,y) \ (set as) \ (set as). x < y}. 5 * c\<^sup>2 * 2 powr (-real r) /(real p)\<^sup>2)" using b by (intro sum_mono, simp add:case_prod_beta) also have "... = ((5/2) * c\<^sup>2 * 2 powr (-real r) /(real p)\<^sup>2) * (2 * card {(x,y) \ (set as) \ (set as). x < y})" by simp also have "... = ((5/2) * c\<^sup>2 * 2 powr (-real r) /(real p)\<^sup>2) * (card (set as) * (card (set as) - 1))" - by (subst card_ordered_pairs, auto) + by (subst card_ordered_pairs, auto) also have "... \ ((5/2) * c\<^sup>2 * 2 powr (-real r) /(real p)\<^sup>2) * (real (card (set as)))\<^sup>2" by (intro mult_left_mono) (auto simp add:power2_eq_square mult_left_mono) also have "... = (5/2) * (real (card (set as)))\<^sup>2 * c\<^sup>2 * 2 powr (-real r) /(real p)\<^sup>2" by (simp add:algebra_simps) finally have f:"prob {\. ?l \ \ degree \ \ 1} \ ?r1" by simp have "prob {\. ?l \} \ prob {\. ?l \ \ degree \ \ 1} + prob {\. degree \ < 1}" by (rule pmf_add[OF M_def], auto) also have "... \ ?r1 + ?r2" by (intro add_mono f prob_degree_lt_1) finally show ?thesis by simp qed private lemma of_bool_square: "(of_bool x)\<^sup>2 = ((of_bool x)::real)" by (cases x, auto) private definition Q where "Q y \ = card {x \ set as. int (hash x \) < y}" private definition m where "m = card (set as)" private lemma assumes "a \ 0" assumes "a \ int p" shows exp_Q: "expectation (\\. real (Q a \)) = real m * (of_int a) / p" and var_Q: "variance (\\. real (Q a \)) \ real m * (of_int a) / p" proof - have exp_single: "expectation (\\. of_bool (int (hash x \) < a)) = real_of_int a /real p" if a:"x \ set as" for x proof - have x_le_p: "x < p" using a as_lt_p by simp have "expectation (\\. of_bool (int (hash x \) < a)) = expectation (indicat_real {\. int (Frequency_Moment_0.hash p x \) < a})" by (intro arg_cong2[where f="integral\<^sup>L"] ext, simp_all) also have "... = prob {\. hash x \ \ {k. int k < a}}" by (simp add:M_def) also have "... = card ({k. int k < a} \ {..\. of_bool (int (hash x \) < a)) = real_of_int a /real p" by simp qed have "expectation(\\. real (Q a \)) = expectation (\\. (\x \ set as. of_bool (int (hash x \) < a)))" by (simp add:Q_def Int_def) also have "... = (\x \ set as. expectation (\\. of_bool (int (hash x \) < a)))" by (rule Bochner_Integration.integral_sum, simp) also have "... = (\ x \ set as. a /real p)" by (rule sum.cong, simp, subst exp_single, simp, simp) also have "... = real m * real_of_int a / real p" by (simp add:m_def) finally show "expectation (\\. real (Q a \)) = real m * real_of_int a / p" by simp have indep: "J \ set as \ card J = 2 \ indep_vars (\_. borel) (\i x. of_bool (int (hash i x) < a)) J" for J using as_subset_p mod_ring_carr by (intro indep_vars_compose2[where Y="\i x. of_bool (int x < a)" and M'="\_. discrete"] k_wise_indep_vars_subset[OF k_wise_indep] finite_subset[OF _ finite_set]) auto have rv: "\x. x \ set as \ random_variable borel (\\. of_bool (int (hash x \) < a))" by (simp add:M_def) have "variance (\\. real (Q a \)) = variance (\\. (\x \ set as. of_bool (int (hash x \) < a)))" by (simp add:Q_def Int_def) also have "... = (\x \ set as. variance (\\. of_bool (int (hash x \) < a)))" by (intro bienaymes_identity_pairwise_indep_2 indep rv) auto also have "... \ (\ x \ set as. a / real p)" by (rule sum_mono, simp add: variance_eq of_bool_square, simp add: exp_single) also have "... = real m * real_of_int a /real p" by (simp add:m_def) finally show "variance (\\. real (Q a \)) \ real m * real_of_int a / p" by simp qed private lemma t_bound: "t \ 81 / (real_of_rat \)\<^sup>2" proof - have "t \ 80 / (real_of_rat \)\<^sup>2 + 1" using t_def t_gt_0 by linarith also have "... \ 80 / (real_of_rat \)\<^sup>2 + 1 / (real_of_rat \)\<^sup>2" using \_range by (intro add_mono, simp, simp add:power_le_one) also have "... = 81 / (real_of_rat \)\<^sup>2" by simp finally show ?thesis by simp qed private lemma t_r_bound: "18 * 40 * (real t)\<^sup>2 * 2 powr (-real r) \ 1" proof - have "720 * (real t)\<^sup>2 * 2 powr (-real r) \ 720 * (81 / (real_of_rat \)\<^sup>2)\<^sup>2 * 2 powr (-4 * log 2 (1 / real_of_rat \) - 23)" using r_bound t_bound by (intro mult_left_mono mult_mono power_mono powr_mono, auto) also have "... \ 720 * (81 / (real_of_rat \)\<^sup>2)\<^sup>2 * (2 powr (-4 * log 2 (1 / real_of_rat \)) * 2 powr (-23))" using \_range by (intro mult_left_mono mult_mono power_mono add_mono) (simp_all add:power_le_one powr_diff) also have "... = 720 * (81\<^sup>2 / (real_of_rat \)^4) * (2 powr (log 2 ((real_of_rat \)^4)) * 2 powr (-23))" using \_range by (intro arg_cong2[where f="(*)"]) (simp_all add:power2_eq_square power4_eq_xxxx log_divide log_powr[symmetric]) also have "... = 720 * 81\<^sup>2 * 2 powr (-23)" using \_range by simp also have "... \ 1" by simp finally show ?thesis by simp qed private lemma m_eq_F_0: "real m = of_rat (F 0 as)" by (simp add:m_def F_def) private lemma estimate'_bounds: "prob {\. of_rat \ * real_of_rat (F 0 as) < \estimate' (sketch_rv' \) - of_rat (F 0 as)\} \ 1/3" proof (cases "card (set as) \ t") case True define \' where "\' = 3 * real_of_rat \ / 4" define u where "u = \real t * p / (m * (1+\'))\" define v where "v = \real t * p / (m * (1-\'))\" - define has_no_collision where + define has_no_collision where "has_no_collision = (\\. \x\ set as. \y \ set as. (tr_hash x \ = tr_hash y \ \ x = y) \ tr_hash x \ > v)" have "2 powr (-real r) \ 2 powr (-(4 * log 2 (1 / real_of_rat \) + 23))" using r_bound by (intro powr_mono, linarith, simp) also have "... = 2 powr (-4 * log 2 (1 /real_of_rat \) -23)" by (rule arg_cong2[where f="(powr)"], auto simp add:algebra_simps) also have "... \ 2 powr ( -1 * log 2 (1 /real_of_rat \) -4)" using \_range by (intro powr_mono diff_mono, auto) also have "... = 2 powr ( -1 * log 2 (1 /real_of_rat \)) / 16" by (simp add: powr_diff) also have "... = real_of_rat \ / 16" using \_range by (simp add:log_divide) also have "... < real_of_rat \ / 8" using \_range by (subst pos_divide_less_eq, auto) finally have r_le_\: "2 powr (-real r) < real_of_rat \ / 8" by simp have \'_gt_0: "\' > 0" using \_range by (simp add:\'_def) have "\' < 3/4" using \_range by (simp add:\'_def)+ also have "... < 1" by simp finally have \'_lt_1: "\' < 1" by simp have "t \ 81 / (real_of_rat \)\<^sup>2" using t_bound by simp also have "... = (81*9/16) / (\')\<^sup>2" by (simp add:\'_def power2_eq_square) also have "... \ 46 / \'\<^sup>2" by (intro divide_right_mono, simp, simp) finally have t_le_\': "t \ 46/ \'\<^sup>2" by simp have "80 \ (real_of_rat \)\<^sup>2 * (80 / (real_of_rat \)\<^sup>2)" using \_range by simp also have "... \ (real_of_rat \)\<^sup>2 * t" by (intro mult_left_mono, simp add:t_def of_nat_ceiling, simp) finally have "80 \ (real_of_rat \)\<^sup>2 * t" by simp hence t_ge_\': "45 \ t * \' * \'" by (simp add:\'_def power2_eq_square) have "m \ card {.. p" using n_le_p by simp finally have m_le_p: "m \ p" by simp hence t_le_m: "t \ card (set as)" using True by simp have m_ge_0: "real m > 0" using m_def True t_gt_0 by simp have "v \ real t * real p / (real m * (1 - \'))" by (simp add:v_def) also have "... \ real t * real p / (real m * (1/4))" using \'_lt_1 m_ge_0 \_range by (intro divide_left_mono mult_left_mono mult_nonneg_nonneg mult_pos_pos, simp_all add:\'_def) finally have v_ubound: "v \ 4 * real t * real p / real m" by (simp add:algebra_simps) have a_ge_1: "u \ 1" using \'_gt_0 p_gt_0 m_ge_0 t_gt_0 - by (auto intro!:mult_pos_pos divide_pos_pos simp add:u_def) + by (auto intro!:mult_pos_pos divide_pos_pos simp add:u_def) hence a_ge_0: "u \ 0" by simp have "real m * (1 - \') < real m" using \'_gt_0 m_ge_0 by simp also have "... \ 1 * real p" using m_le_p by simp also have "... \ real t * real p" using t_gt_0 by (intro mult_right_mono, auto) finally have " real m * (1 - \') < real t * real p" by simp hence v_gt_0: "v > 0" using mult_pos_pos m_ge_0 \'_lt_1 by (simp add:v_def) hence v_ge_1: "real_of_int v \ 1" by linarith have "real t \ real m" using True m_def by linarith also have "... < (1 + \') * real m" using \'_gt_0 m_ge_0 by force finally have a_le_p_aux: "real t < (1 + \') * real m" by simp have "u \ real t * real p / (real m * (1 + \'))+1" by (simp add:u_def) - also have "... < real p + 1" + also have "... < real p + 1" using m_ge_0 \'_gt_0 a_le_p_aux a_le_p_aux p_gt_0 - by (simp add: pos_divide_less_eq ac_simps) - finally have "u \ real p" + by (simp add: pos_divide_less_eq ac_simps) + finally have "u \ real p" by (metis int_less_real_le not_less of_int_le_iff of_int_of_nat_eq) hence u_le_p: "u \ int p" by linarith - have "prob {\. Q u \ \ t} \ prob {\ \ Sigma_Algebra.space M. abs (real (Q u \) - + have "prob {\. Q u \ \ t} \ prob {\ \ Sigma_Algebra.space M. abs (real (Q u \) - expectation (\\. real (Q u \))) \ 3 * sqrt (m * real_of_int u / p)}" proof (rule pmf_mono[OF M_def]) fix \ assume "\ \ {\. t \ Q u \}" hence t_le: "t \ Q u \" by simp have "real m * real_of_int u / real p \ real m * (real t * real p / (real m * (1 + \'))+1) / real p" using m_ge_0 p_gt_0 by (intro divide_right_mono mult_left_mono, simp_all add: u_def) also have "... = real m * real t * real p / (real m * (1+\') * real p) + real m / real p" by (simp add:distrib_left add_divide_distrib) also have "... = real t / (1+\') + real m / real p" using p_gt_0 m_ge_0 by simp also have "... \ real t / (1+\') + 1" using m_le_p p_gt_0 by (intro add_mono, auto) finally have "real m * real_of_int u / real p \ real t / (1 + \') + 1" by simp - hence "3 * sqrt (real m * of_int u / real p) + real m * of_int u / real p \ + hence "3 * sqrt (real m * of_int u / real p) + real m * of_int u / real p \ 3 * sqrt (t / (1+\')+1)+(t/(1+\')+1)" by (intro add_mono mult_left_mono real_sqrt_le_mono, auto) also have "... \ 3 * sqrt (real t+1) + ((t * (1 - \' / (1+\'))) + 1)" using \'_gt_0 t_gt_0 by (intro add_mono mult_left_mono real_sqrt_le_mono) (simp_all add: pos_divide_le_eq left_diff_distrib) also have "... = 3 * sqrt (real t+1) + (t - \' * t / (1+\')) + 1" by (simp add:algebra_simps) also have "... \ 3 * sqrt (46 / \'\<^sup>2 + 1 / \'\<^sup>2) + (t - \' * t/2) + 1 / \'" using \'_gt_0 t_gt_0 \'_lt_1 add_pos_pos t_le_\' by (intro add_mono mult_left_mono real_sqrt_le_mono add_mono) (simp_all add: power_le_one pos_le_divide_eq) - also have "... \ (21 / \' + (t - 45 / (2*\'))) + 1 / \'" + also have "... \ (21 / \' + (t - 45 / (2*\'))) + 1 / \'" using \'_gt_0 t_ge_\' by (intro add_mono) (simp_all add:real_sqrt_divide divide_le_cancel real_le_lsqrt pos_divide_le_eq ac_simps) also have "... \ t" using \'_gt_0 by simp also have "... \ Q u \" using t_le by simp finally have "3 * sqrt (real m * of_int u / real p) + real m * of_int u / real p \ Q u \" by simp hence " 3 * sqrt (real m * real_of_int u / real p) \ \real (Q u \) - expectation (\\. real (Q u \))\" using a_ge_0 u_le_p True by (simp add:exp_Q abs_ge_iff) - thus "\ \ {\ \ Sigma_Algebra.space M. 3 * sqrt (real m * real_of_int u / real p) \ + thus "\ \ {\ \ Sigma_Algebra.space M. 3 * sqrt (real m * real_of_int u / real p) \ \real (Q u \) - expectation (\\. real (Q u \))\}" by (simp add: M_def) qed also have "... \ variance (\\. real (Q u \)) / (3 * sqrt (real m * of_int u / real p))\<^sup>2" - using a_ge_1 p_gt_0 m_ge_0 - by (intro Chebyshev_inequality, simp add:M_def, auto) + using a_ge_1 p_gt_0 m_ge_0 + by (intro Chebyshev_inequality, simp add:M_def, auto) also have "... \ (real m * real_of_int u / real p) / (3 * sqrt (real m * of_int u / real p))\<^sup>2" using a_ge_0 u_le_p by (intro divide_right_mono var_Q, auto) also have "... \ 1/9" using a_ge_0 by simp finally have case_1: "prob {\. Q u \ \ t} \ 1/9" by simp have case_2: "prob {\. Q v \ < t} \ 1/9" proof (cases "v \ p") case True - have "prob {\. Q v \ < t} \ prob {\ \ Sigma_Algebra.space M. abs (real (Q v \) - expectation (\\. real (Q v \))) + have "prob {\. Q v \ < t} \ prob {\ \ Sigma_Algebra.space M. abs (real (Q v \) - expectation (\\. real (Q v \))) \ 3 * sqrt (m * real_of_int v / p)}" proof (rule pmf_mono[OF M_def]) fix \ assume "\ \ set_pmf (pmf_of_set space)" have "(real t + 3 * sqrt (real t / (1 - \') )) * (1 - \') = real t - \' * t + 3 * ((1-\') * sqrt( real t / (1-\') ))" by (simp add:algebra_simps) also have "... = real t - \' * t + 3 * sqrt ( (1-\')\<^sup>2 * (real t / (1-\')))" using \'_lt_1 by (subst real_sqrt_mult, simp) also have "... = real t - \' * t + 3 * sqrt ( real t * (1- \'))" by (simp add:power2_eq_square distrib_left) also have "... \ real t - 45/ \' + 3 * sqrt ( real t )" using \'_gt_0 t_ge_\' \'_lt_1 by (intro add_mono mult_left_mono real_sqrt_le_mono) (simp_all add:pos_divide_le_eq ac_simps left_diff_distrib power_le_one) also have "... \ real t - 45/ \' + 3 * sqrt ( 46 / \'\<^sup>2)" using t_le_\' \'_lt_1 \'_gt_0 by (intro add_mono mult_left_mono real_sqrt_le_mono, simp_all add:pos_divide_le_eq power_le_one) also have "... = real t + (3 * sqrt(46) - 45)/ \'" using \'_gt_0 by (simp add:real_sqrt_divide diff_divide_distrib) also have "... \ t" using \'_gt_0 by (simp add:pos_divide_le_eq real_le_lsqrt) finally have aux: "(real t + 3 * sqrt (real t / (1 - \'))) * (1 - \') \ real t " by simp assume "\ \ {\. Q v \ < t}" hence "Q v \ < t" by simp - hence "real (Q v \) + 3 * sqrt (real m * real_of_int v / real p) + hence "real (Q v \) + 3 * sqrt (real m * real_of_int v / real p) \ real t - 1 + 3 * sqrt (real m * real_of_int v / real p)" using m_le_p p_gt_0 by (intro add_mono, auto simp add: algebra_simps add_divide_distrib) also have "... \ (real t-1) + 3 * sqrt (real m * (real t * real p / (real m * (1- \'))) / real p)" by (intro add_mono mult_left_mono real_sqrt_le_mono divide_right_mono) (auto simp add:v_def) also have "... \ real t + 3 * sqrt(real t / (1-\')) - 1" using m_ge_0 p_gt_0 by simp - also have "... \ real t / (1-\')-1" - using \'_lt_1 aux by (simp add: pos_le_divide_eq) + also have "... \ real t / (1-\')-1" + using \'_lt_1 aux by (simp add: pos_le_divide_eq) also have "... \ real m * (real t * real p / (real m * (1-\'))) / real p - 1" using p_gt_0 m_ge_0 by simp also have "... \ real m * (real t * real p / (real m * (1-\'))) / real p - real m / real p" using m_le_p p_gt_0 by (intro diff_mono, auto) - also have "... = real m * (real t * real p / (real m * (1-\'))-1) / real p" + also have "... = real m * (real t * real p / (real m * (1-\'))-1) / real p" by (simp add: left_diff_distrib right_diff_distrib diff_divide_distrib) - also have "... \ real m * real_of_int v / real p" + also have "... \ real m * real_of_int v / real p" by (intro divide_right_mono mult_left_mono, simp_all add:v_def) - finally have "real (Q v \) + 3 * sqrt (real m * real_of_int v / real p) + finally have "real (Q v \) + 3 * sqrt (real m * real_of_int v / real p) \ real m * real_of_int v / real p" by simp - hence " 3 * sqrt (real m * real_of_int v / real p) \ \real (Q v \) -expectation (\\. real (Q v \))\" + hence " 3 * sqrt (real m * real_of_int v / real p) \ \real (Q v \) -expectation (\\. real (Q v \))\" using v_gt_0 True by (simp add: exp_Q abs_ge_iff) - thus "\ \ {\\ Sigma_Algebra.space M. 3 * sqrt (real m * real_of_int v / real p) \ - \real (Q v \) - expectation (\\. real (Q v \))\}" + thus "\ \ {\\ Sigma_Algebra.space M. 3 * sqrt (real m * real_of_int v / real p) \ + \real (Q v \) - expectation (\\. real (Q v \))\}" by (simp add:M_def) qed - also have "... \ variance (\\. real (Q v \)) / (3 * sqrt (real m * real_of_int v / real p))\<^sup>2" - using v_gt_0 p_gt_0 m_ge_0 + also have "... \ variance (\\. real (Q v \)) / (3 * sqrt (real m * real_of_int v / real p))\<^sup>2" + using v_gt_0 p_gt_0 m_ge_0 by (intro Chebyshev_inequality, simp add:M_def, auto) also have "... \ (real m * real_of_int v / real p) / (3 * sqrt (real m * real_of_int v / real p))\<^sup>2" using v_gt_0 True by (intro divide_right_mono var_Q, auto) also have "... = 1/9" using p_gt_0 v_gt_0 m_ge_0 by (simp add:power2_eq_square) finally show ?thesis by simp next case False have "prob {\. Q v \ < t} \ prob {\. False}" proof (rule pmf_mono[OF M_def]) fix \ assume a:"\ \ {\. Q v \ < t}" assume "\ \ set_pmf (pmf_of_set space)" - hence b:"\x. x < p \ hash x \ < p" - using hash_range mod_ring_carr by (simp add:M_def measure_pmf_inverse) + hence b:"\x. x < p \ hash x \ < p" + using hash_range mod_ring_carr by (simp add:M_def measure_pmf_inverse) have "t \ card (set as)" using True by simp also have "... \ Q v \" - unfolding Q_def using b False as_lt_p by (intro card_mono subsetI, simp, force) + unfolding Q_def using b False as_lt_p by (intro card_mono subsetI, simp, force) also have "... < t" using a by simp finally have "False" by auto thus "\ \ {\. False}" by simp qed also have "... = 0" by auto finally show ?thesis by simp qed have "prob {\. \has_no_collision \} \ prob {\. \x \ set as. \y \ set as. x \ y \ tr_hash x \ \ real_of_int v \ tr_hash x \ = tr_hash y \}" - by (rule pmf_mono[OF M_def]) (simp add:has_no_collision_def M_def, force) + by (rule pmf_mono[OF M_def]) (simp add:has_no_collision_def M_def, force) also have "... \ (5/2) * (real (card (set as)))\<^sup>2 * (real_of_int v)\<^sup>2 * 2 powr - real r / (real p)\<^sup>2 + 1 / real p" using collision_prob v_ge_1 by blast also have "... \ (5/2) * (real m)\<^sup>2 * (real_of_int v)\<^sup>2 * 2 powr - real r / (real p)\<^sup>2 + 1 / real p" by (intro divide_right_mono add_mono mult_right_mono mult_mono power_mono, simp_all add:m_def) also have "... \ (5/2) * (real m)\<^sup>2 * (4 * real t * real p / real m)\<^sup>2 * (2 powr - real r) / (real p)\<^sup>2 + 1 / real p" using v_def v_ge_1 v_ubound by (intro add_mono divide_right_mono mult_right_mono mult_left_mono, auto) also have "... = 40 * (real t)\<^sup>2 * (2 powr -real r) + 1 / real p" using p_gt_0 m_ge_0 t_gt_0 by (simp add:algebra_simps power2_eq_square) also have "... \ 1/18 + 1/18" using t_r_bound p_ge_18 by (intro add_mono, simp_all add: pos_le_divide_eq) also have "... = 1/9" by simp finally have case_3: "prob {\. \has_no_collision \} \ 1/9" by simp - have "prob {\. real_of_rat \ * of_rat (F 0 as) < \estimate' (sketch_rv' \) - of_rat (F 0 as)\} \ + have "prob {\. real_of_rat \ * of_rat (F 0 as) < \estimate' (sketch_rv' \) - of_rat (F 0 as)\} \ prob {\. Q u \ \ t \ Q v \ < t \ \(has_no_collision \)}" proof (rule pmf_mono[OF M_def], rule ccontr) fix \ assume "\ \ set_pmf (pmf_of_set space)" assume "\ \ {\. real_of_rat \ * real_of_rat (F 0 as) < \estimate' (sketch_rv' \) - real_of_rat (F 0 as)\}" hence est: "real_of_rat \ * real_of_rat (F 0 as) < \estimate' (sketch_rv' \) - real_of_rat (F 0 as)\" by simp assume "\ \ {\. t \ Q u \ \ Q v \ < t \ \ has_no_collision \}" hence "\( t \ Q u \ \ Q v \ < t \ \ has_no_collision \)" by simp hence lb: "Q u \ < t" and ub: "Q v \ \ t" and no_col: "has_no_collision \" by simp+ define y where "y = nth_mset (t-1) {#int (hash x \). x \# mset_set (set as)#}" define y' where "y' = nth_mset (t-1) {#tr_hash x \. x \# mset_set (set as)#}" have rank_t_lb: "u \ y" unfolding y_def using True t_gt_0 lb by (intro nth_mset_bound_left, simp_all add:count_less_def swap_filter_image Q_def) - + have rank_t_ub: "y \ v - 1" unfolding y_def using True t_gt_0 ub by (intro nth_mset_bound_right, simp_all add:Q_def swap_filter_image count_le_def) have y_ge_0: "real_of_int y \ 0" using rank_t_lb a_ge_0 by linarith - have "mono (\x. truncate_down r (real_of_int x))" + have "mono (\x. truncate_down r (real_of_int x))" by (metis truncate_down_mono mono_def of_int_le_iff) hence y'_eq: "y' = truncate_down r y" unfolding y_def y'_def using True t_gt_0 - by (subst nth_mset_commute_mono[where f="(\x. truncate_down r (of_int x))"]) + by (subst nth_mset_commute_mono[where f="(\x. truncate_down r (of_int x))"]) (simp_all add: multiset.map_comp comp_def tr_hash_def) have "real_of_int u * (1 - 2 powr -real r) \ real_of_int y * (1 - 2 powr (-real r))" using rank_t_lb of_int_le_iff two_pow_r_le_1 by (intro mult_right_mono, auto) also have "... \ y'" using y'_eq truncate_down_pos[OF y_ge_0] by simp finally have rank_t_lb': "u * (1 - 2 powr -real r) \ y'" by simp have "y' \ real_of_int y" by (subst y'_eq, rule truncate_down_le, simp) also have "... \ real_of_int (v-1)" using rank_t_ub of_int_le_iff by blast finally have rank_t_ub': "y' \ v-1" by simp have "0 < u * (1-2 powr -real r)" using a_ge_1 two_pow_r_le_1 by (intro mult_pos_pos, auto) hence y'_pos: "y' > 0" using rank_t_lb' by linarith have no_col': "\x. x \ y' \ count {#tr_hash x \. x \# mset_set (set as)#} x \ 1" - using rank_t_ub' no_col + using rank_t_ub' no_col by (simp add:vimage_def card_le_Suc0_iff_eq count_image_mset has_no_collision_def) force have h_1: "Max (sketch_rv' \) = y'" using True t_gt_0 no_col' by (simp add:sketch_rv'_def y'_def nth_mset_max) have "card (sketch_rv' \) = card (least ((t-1)+1) (set_mset {#tr_hash x \. x \# mset_set (set as)#}))" using t_gt_0 by (simp add:sketch_rv'_def) also have "... = (t-1) +1" using True t_gt_0 no_col' by (intro nth_mset_max(2), simp_all add:y'_def) also have "... = t" using t_gt_0 by simp finally have "card (sketch_rv' \) = t" by simp hence h_3: "estimate' (sketch_rv' \) = real t * real p / y'" using h_1 by (simp add:estimate'_def) - have "(real t) * real p \ (1 + \') * real m * ((real t) * real p / (real m * (1 + \')))" + have "(real t) * real p \ (1 + \') * real m * ((real t) * real p / (real m * (1 + \')))" using \'_lt_1 m_def True t_gt_0 \'_gt_0 by auto also have "... \ (1+\') * m * u" using \'_gt_0 by (intro mult_left_mono, simp_all add:u_def) also have "... < ((1 + real_of_rat \)*(1-real_of_rat \/8)) * m * u" using True m_def t_gt_0 a_ge_1 \_range by (intro mult_strict_right_mono, auto simp add:\'_def right_diff_distrib) also have "... \ ((1 + real_of_rat \)*(1-2 powr (-r))) * m * u" using r_le_\ \_range a_ge_0 by (intro mult_right_mono mult_left_mono, auto) - also have "... = (1 + real_of_rat \) * m * (u * (1-2 powr -real r))" + also have "... = (1 + real_of_rat \) * m * (u * (1-2 powr -real r))" by simp also have "... \ (1 + real_of_rat \) * m * y'" using \_range by (intro mult_left_mono rank_t_lb', simp) finally have "real t * real p < (1 + real_of_rat \) * m * y'" by simp hence f_1: "estimate' (sketch_rv' \) < (1 + real_of_rat \) * m" using y'_pos by (simp add: h_3 pos_divide_less_eq) - have "(1 - real_of_rat \) * m * y' \ (1 - real_of_rat \) * m * v" + have "(1 - real_of_rat \) * m * y' \ (1 - real_of_rat \) * m * v" using \_range rank_t_ub' y'_pos by (intro mult_mono rank_t_ub', simp_all) also have "... = (1-real_of_rat \) * (real m * v)" by simp - also have "... < (1-\') * (real m * v)" + also have "... < (1-\') * (real m * v)" using \_range m_ge_0 v_ge_1 by (intro mult_strict_right_mono mult_pos_pos, simp_all add:\'_def) also have "... \ (1-\') * (real m * (real t * real p / (real m * (1-\'))))" using \'_gt_0 \'_lt_1 by (intro mult_left_mono, auto simp add:v_def) also have "... = real t * real p" using \'_gt_0 \'_lt_1 t_gt_0 p_gt_0 m_ge_0 by auto finally have "(1 - real_of_rat \) * m * y' < real t * real p" by simp hence f_2: "estimate' (sketch_rv' \) > (1 - real_of_rat \) * m" using y'_pos by (simp add: h_3 pos_less_divide_eq) have "abs (estimate' (sketch_rv' \) - real_of_rat (F 0 as)) < real_of_rat \ * (real_of_rat (F 0 as))" using f_1 f_2 by (simp add:abs_less_iff algebra_simps m_eq_F_0) thus "False" using est by linarith qed also have "... \ 1/9 + (1/9 + 1/9)" by (intro pmf_add_2[OF M_def] case_1 case_2 case_3) also have "... = 1/3" by simp finally show ?thesis by simp next case False have "prob {\. real_of_rat \ * of_rat (F 0 as) < \estimate' (sketch_rv' \) - of_rat (F 0 as)\} \ - prob {\. \x \ set as. \y \ set as. x \ y \ tr_hash x \ \ real p \ tr_hash x \ = tr_hash y \}" + prob {\. \x \ set as. \y \ set as. x \ y \ tr_hash x \ \ real p \ tr_hash x \ = tr_hash y \}" proof (rule pmf_mono[OF M_def]) fix \ assume a:"\ \ {\. real_of_rat \ * real_of_rat (F 0 as) < \estimate' (sketch_rv' \) - real_of_rat (F 0 as)\}" - assume b:"\ \ set_pmf (pmf_of_set space)" + assume b:"\ \ set_pmf (pmf_of_set space)" have c: "card (set as) < t" using False by auto hence "card ((\x. tr_hash x \) ` set as) < t" using card_image_le order_le_less_trans by blast hence d:"card (sketch_rv' \) = card ((\x. tr_hash x \) ` (set as))" by (simp add:sketch_rv'_def card_least) have "card (sketch_rv' \) < t" by (metis List.finite_set c d card_image_le order_le_less_trans) hence "estimate' (sketch_rv' \) = card (sketch_rv' \)" by (simp add:estimate'_def) hence "card (sketch_rv' \) \ real_of_rat (F 0 as)" - using a \_range by simp + using a \_range by simp (metis abs_zero cancel_comm_monoid_add_class.diff_cancel of_nat_less_0_iff pos_prod_lt zero_less_of_rat_iff) hence "card (sketch_rv' \) \ card (set as)" using m_def m_eq_F_0 by linarith hence "\inj_on (\x. tr_hash x \) (set as)" using card_image d by auto moreover have "tr_hash x \ \ real p" if a:"x \ set as" for x proof - - have "hash x \ < p" + have "hash x \ < p" using hash_range as_lt_p a b by (simp add:mod_ring_carr M_def) thus "tr_hash x \ \ real p" using truncate_down_le by (simp add:tr_hash_def) qed ultimately show "\ \ {\. \x \ set as. \y \ set as. x \ y \ tr_hash x \ \ real p \ tr_hash x \ = tr_hash y \}" by (simp add:inj_on_def, blast) qed also have "... \ (5/2) * (real (card (set as)))\<^sup>2 * (real p)\<^sup>2 * 2 powr - real r / (real p)\<^sup>2 + 1 / real p" using p_gt_0 by (intro collision_prob, auto) also have "... = (5/2) * (real (card (set as)))\<^sup>2 * 2 powr (- real r) + 1 / real p" using p_gt_0 by (simp add:ac_simps power2_eq_square) also have "... \ (5/2) * (real t)\<^sup>2 * 2 powr (-real r) + 1 / real p" using False by (intro add_mono mult_right_mono mult_left_mono power_mono, auto) also have "... \ 1/6 + 1/6" using t_r_bound p_ge_18 by (intro add_mono, simp_all) also have "... \ 1/3" by simp finally show ?thesis by simp qed private lemma median_bounds: "\

(\ in measure_pmf \\<^sub>0. \median s (\i. estimate (sketch_rv (\ i))) - F 0 as\ \ \ * F 0 as) \ 1 - real_of_rat \" proof - have "strict_mono_on A real_of_float" for A by (meson less_float.rep_eq strict_mono_onI) - hence real_g_2: "\\. sketch_rv' \ = real_of_float ` sketch_rv \" + hence real_g_2: "\\. sketch_rv' \ = real_of_float ` sketch_rv \" by (simp add: sketch_rv'_def sketch_rv_def tr_hash_def least_mono_commute image_comp) moreover have "inj_on real_of_float A" for A using real_of_float_inject by (simp add:inj_on_def) - ultimately have card_eq: "\\. card (sketch_rv \) = card (sketch_rv' \)" + ultimately have card_eq: "\\. card (sketch_rv \) = card (sketch_rv' \)" using real_g_2 by (auto intro!: card_image[symmetric]) - have "Max (sketch_rv' \) = real_of_float (Max (sketch_rv \))" if a:"card (sketch_rv' \) \ t" for \ + have "Max (sketch_rv' \) = real_of_float (Max (sketch_rv \))" if a:"card (sketch_rv' \) \ t" for \ proof - have "mono real_of_float" using less_eq_float.rep_eq mono_def by blast moreover have "finite (sketch_rv \)" by (simp add:sketch_rv_def least_def) moreover have " sketch_rv \ \ {}" - using card_eq[symmetric] card_gt_0_iff t_gt_0 a by (simp, force) + using card_eq[symmetric] card_gt_0_iff t_gt_0 a by (simp, force) ultimately show ?thesis by (subst mono_Max_commute[where f="real_of_float"], simp_all add:real_g_2) qed hence real_g: "\\. estimate' (sketch_rv' \) = real_of_rat (estimate (sketch_rv \))" by (simp add:estimate_def estimate'_def card_eq of_rat_divide of_rat_mult of_rat_add real_of_rat_of_float) have indep: "prob_space.indep_vars (measure_pmf \\<^sub>0) (\_. borel) (\i \. estimate' (sketch_rv' (\ i))) {0..\<^sub>0_def by (rule indep_vars_restrict_intro', auto simp add:restrict_dfl_def lessThan_atLeast0) moreover have "- (18 * ln (real_of_rat \)) \ real s" using of_nat_ceiling by (simp add:s_def) blast moreover have "i < s \ measure \\<^sub>0 {\. of_rat \ * of_rat (F 0 as) < \estimate' (sketch_rv' (\ i)) - of_rat (F 0 as)\} \ 1/3" for i using estimate'_bounds unfolding \\<^sub>0_def M_def by (subst prob_prod_pmf_slice, simp_all) - + ultimately have "1-real_of_rat \ \ \

(\ in measure_pmf \\<^sub>0. \median s (\i. estimate' (sketch_rv' (\ i))) - real_of_rat (F 0 as)\ \ real_of_rat \ * real_of_rat (F 0 as))" using \_range prob_space_measure_pmf by (intro prob_space.median_bound_2) auto - also have "... = \

(\ in measure_pmf \\<^sub>0. + also have "... = \

(\ in measure_pmf \\<^sub>0. \median s (\i. estimate (sketch_rv (\ i))) - F 0 as\ \ \ * F 0 as)" - using s_gt_0 median_rat[symmetric] real_g by (intro arg_cong2[where f="measure"]) + using s_gt_0 median_rat[symmetric] real_g by (intro arg_cong2[where f="measure"]) (simp_all add:of_rat_diff[symmetric] of_rat_mult[symmetric] of_rat_less_eq) finally show "\

(\ in measure_pmf \\<^sub>0. \median s (\i. estimate (sketch_rv (\ i))) - F 0 as\ \ \ * F 0 as) \ 1-real_of_rat \" by blast qed lemma f0_alg_correct': "\

(\ in measure_pmf result. \\ - F 0 as\ \ \ * F 0 as) \ 1 - of_rat \" proof - have f0_result_elim: "\x. f0_result (s, t, p, r, x, \i\{..i. estimate (sketch_rv (x i))))" by (simp add:estimate_def, rule median_cong, simp) - + have "result = map_pmf (\x. (s, t, p, r, x, \i\{..\<^sub>0 \ f0_result" by (subst result_def, subst f0_alg_sketch, simp) also have "... = \\<^sub>0 \ (\x. return_pmf (s, t, p, r, x, \i\{.. f0_result" by (simp add:t_def p_def r_def s_def map_pmf_def) also have "... = \\<^sub>0 \ (\x. return_pmf (median s (\i. estimate (sketch_rv (x i)))))" by (subst bind_assoc_pmf, subst bind_return_pmf, subst f0_result_elim) simp finally have a:"result = \\<^sub>0 \ (\x. return_pmf (median s (\i. estimate (sketch_rv (x i)))))" by simp show ?thesis using median_bounds by (simp add: a map_pmf_def[symmetric]) qed private lemma f_subset: assumes "g ` A \ h ` B" shows "(\x. f (g x)) ` A \ (\x. f (h x)) ` B" using assms by auto lemma f0_exact_space_usage': defines "\ \ fold (\a state. state \ f0_update a) as (f0_init \ \ n)" shows "AE \ in \. bit_count (encode_f0_state \) \ f0_space_usage (n, \, \)" proof - - - have log_2_4: "log 2 4 = 2" + + have log_2_4: "log 2 4 = 2" by (metis log2_of_power_eq mult_2 numeral_Bit0 of_nat_numeral power2_eq_square) - have a: "bit_count (F\<^sub>e (float_of (truncate_down r y))) \ + have a: "bit_count (F\<^sub>e (float_of (truncate_down r y))) \ ereal (12 + 4 * real r + 2 * log 2 (log 2 (n+13)))" if a_1:"y \ {.. 1") case True - have aux_1: " 0 < 2 + log 2 (real y)" + have aux_1: " 0 < 2 + log 2 (real y)" using True by (intro add_pos_nonneg, auto) have aux_2: "0 < 2 + log 2 (real p)" using p_gt_1 by (intro add_pos_nonneg, auto) - have "bit_count (F\<^sub>e (float_of (truncate_down r y))) \ + have "bit_count (F\<^sub>e (float_of (truncate_down r y))) \ ereal (10 + 4 * real r + 2 * log 2 (2 + \log 2 \real y\\))" by (rule truncate_float_bit_count) also have "... = ereal (10 + 4 * real r + 2 * log 2 (2 + (log 2 (real y))))" using True by simp also have "... \ ereal (10 + 4 * real r + 2 * log 2 (2 + log 2 p))" using aux_1 aux_2 True p_gt_0 a_1 by simp also have "... \ ereal (10 + 4 * real r + 2 * log 2 (log 2 4 + log 2 (2 * n + 40)))" using log_2_4 p_le_n p_gt_0 by (intro ereal_mono add_mono mult_left_mono log_mono of_nat_mono add_pos_nonneg, auto) also have "... = ereal (10 + 4 * real r + 2 * log 2 (log 2 (8 * n + 160)))" by (simp add:log_mult[symmetric]) also have "... \ ereal (10 + 4 * real r + 2 * log 2 (log 2 ((n+13) powr 2)))" by (intro ereal_mono add_mono mult_left_mono log_mono of_nat_mono add_pos_nonneg) (auto simp add:power2_eq_square algebra_simps) also have "... = ereal (10 + 4 * real r + 2 * log 2 (log 2 4 * log 2 (n + 13)))" by (subst log_powr, simp_all add:log_2_4) also have "... = ereal (12 + 4 * real r + 2 * log 2 (log 2 (n + 13)))" by (subst log_mult, simp_all add:log_2_4) finally show ?thesis by simp next case False hence "y = 0" using a_1 by simp then show ?thesis by (simp add:float_bit_count_zero) qed - have "bit_count (encode_f0_state (s, t, p, r, x, \i\{.. + have "bit_count (encode_f0_state (s, t, p, r, x, \i\{.. f0_space_usage (n, \, \)" if b: "x \ {..\<^sub>E space" for x proof - - have c: "x \ extensional {.. extensional {.. (\k. float_of (truncate_down r k)) ` {.. (\xa. float_of (truncate_down r (hash xa (x y)))) ` set as" - using least_subset by (auto simp add:sketch_rv_def tr_hash_def) + using least_subset by (auto simp add:sketch_rv_def tr_hash_def) also have "... \ (\k. float_of (truncate_down r (real k))) ` {..y. y < s \ finite (sketch_rv (x y))" unfolding sketch_rv_def by (rule finite_subset[OF least_subset], simp) moreover have card_sketch: "\y. y < s \ card (sketch_rv (x y)) \ t " by (simp add:sketch_rv_def card_least) - moreover have "\y z. y < s \ z \ sketch_rv (x y) \ + moreover have "\y z. y < s \ z \ sketch_rv (x y) \ bit_count (F\<^sub>e z) \ ereal (12 + 4 * real r + 2 * log 2 (log 2 (real n + 13)))" using a d by auto - ultimately have e: "\y. y < s \ bit_count (S\<^sub>e F\<^sub>e (sketch_rv (x y))) + ultimately have e: "\y. y < s \ bit_count (S\<^sub>e F\<^sub>e (sketch_rv (x y))) \ ereal (real t) * (ereal (12 + 4 * real r + 2 * log 2 (log 2 (real (n + 13)))) + 1) + 1" using float_encoding by (intro set_bit_count_est, auto) have f: "\y. y < s \ bit_count (P\<^sub>e p 2 (x y)) \ ereal (real 2 * (log 2 (real p) + 1))" using p_gt_1 b by (intro bounded_degree_polynomial_bit_count) (simp_all add:space_def PiE_def Pi_def) have "bit_count (encode_f0_state (s, t, p, r, x, \i\{..e s) + bit_count (N\<^sub>e t) + bit_count (N\<^sub>e p) + bit_count (N\<^sub>e r) + bit_count (([0..\<^sub>e P\<^sub>e p 2) x) + bit_count (([0..\<^sub>e S\<^sub>e F\<^sub>e) (\i\{.. ereal (2* log 2 (real s + 1) + 1) + ereal (2* log 2 (real t + 1) + 1) + ereal (2* log 2 (real p + 1) + 1) + ereal (2 * log 2 (real r + 1) + 1) - + (ereal (real s) * (ereal (real 2 * (log 2 (real p) + 1)))) - + (ereal (real s) * ((ereal (real t) * + + (ereal (real s) * (ereal (real 2 * (log 2 (real p) + 1)))) + + (ereal (real s) * ((ereal (real t) * (ereal (12 + 4 * real r + 2 * log 2 (log 2 (real (n + 13)))) + 1) + 1)))" using c e f by (intro add_mono exp_golomb_bit_count fun_bit_count_est[where xs="[0.. ereal ( 4 + 2 * log 2 (real s + 1) + 2 * log 2 (real t + 1) + - 2 * log 2 (2 * (21 + real n)) + 2 * log 2 (real r + 1) + real s * (3 + 2 * log 2 (2 * (21 + real n)) + + also have "... \ ereal ( 4 + 2 * log 2 (real s + 1) + 2 * log 2 (real t + 1) + + 2 * log 2 (2 * (21 + real n)) + 2 * log 2 (real r + 1) + real s * (3 + 2 * log 2 (2 * (21 + real n)) + real t * (13 + (4 * real r + 2 * log 2 (log 2 (real n + 13))))))" using p_le_n p_gt_0 by (intro ereal_mono add_mono mult_left_mono, auto) - also have "... = ereal (6 + 2 * log 2 (real s + 1) + 2 * log 2 (real t + 1) + - 2 * log 2 (21 + real n) + 2 * log 2 (real r + 1) + real s * (5 + 2 * log 2 (21 + real n) + + also have "... = ereal (6 + 2 * log 2 (real s + 1) + 2 * log 2 (real t + 1) + + 2 * log 2 (21 + real n) + 2 * log 2 (real r + 1) + real s * (5 + 2 * log 2 (21 + real n) + real t * (13 + (4 * real r + 2 * log 2 (log 2 (real n + 13))))))" by (subst (1 2) log_mult, auto) also have "... \ f0_space_usage (n, \, \)" by (simp add:s_def[symmetric] r_def[symmetric] t_def[symmetric] Let_def) (simp add:algebra_simps) - finally show "bit_count (encode_f0_state (s, t, p, r, x, \i\{.. + finally show "bit_count (encode_f0_state (s, t, p, r, x, \i\{.. f0_space_usage (n, \, \)" by simp qed hence "\x. x \ set_pmf \\<^sub>0 \ bit_count (encode_f0_state (s, t, p, r, x, \i\{.. ereal (f0_space_usage (n, \, \))" by (simp add:\\<^sub>0_def set_prod_pmf del:f0_space_usage.simps) hence "\y. y \ set_pmf \ \ bit_count (encode_f0_state y) \ ereal (f0_space_usage (n, \, \))" by (simp add: \_def f0_alg_sketch del:f0_space_usage.simps f0_init.simps) (metis (no_types, lifting) image_iff pmf.set_map) thus ?thesis by (simp add: AE_measure_pmf_iff del:f0_space_usage.simps) qed end text \Main results of this section:\ theorem f0_alg_correct: assumes "\ \ {0<..<1}" assumes "\ \ {0<..<1}" assumes "set as \ {.. \ fold (\a state. state \ f0_update a) as (f0_init \ \ n) \ f0_result" shows "\

(\ in measure_pmf \. \\ - F 0 as\ \ \ * F 0 as) \ 1 - of_rat \" using f0_alg_correct'[OF assms(1-3)] unfolding \_def by blast theorem f0_exact_space_usage: assumes "\ \ {0<..<1}" assumes "\ \ {0<..<1}" assumes "set as \ {.. \ fold (\a state. state \ f0_update a) as (f0_init \ \ n)" shows "AE \ in \. bit_count (encode_f0_state \) \ f0_space_usage (n, \, \)" using f0_exact_space_usage'[OF assms(1-3)] unfolding \_def by blast theorem f0_asymptotic_space_complexity: - "f0_space_usage \ O[at_top \\<^sub>F at_right 0 \\<^sub>F at_right 0](\(n, \, \). ln (1 / of_rat \) * + "f0_space_usage \ O[at_top \\<^sub>F at_right 0 \\<^sub>F at_right 0](\(n, \, \). ln (1 / of_rat \) * (ln (real n) + 1 / (of_rat \)\<^sup>2 * (ln (ln (real n)) + ln (1 / of_rat \))))" (is "_ \ O[?F](?rhs)") proof - define n_of :: "nat \ rat \ rat \ nat" where "n_of = (\(n, \, \). n)" define \_of :: "nat \ rat \ rat \ rat" where "\_of = (\(n, \, \). \)" define \_of :: "nat \ rat \ rat \ rat" where "\_of = (\(n, \, \). \)" define t_of where "t_of = (\x. nat \80 / (real_of_rat (\_of x))\<^sup>2\)" define s_of where "s_of = (\x. nat \-(18 * ln (real_of_rat (\_of x)))\)" define r_of where "r_of = (\x. nat (4 * \log 2 (1 / real_of_rat (\_of x))\ + 23))" - define g where "g = (\x. ln (1 / of_rat (\_of x)) * (ln (real (n_of x)) + + define g where "g = (\x. ln (1 / of_rat (\_of x)) * (ln (real (n_of x)) + 1 / (of_rat (\_of x))\<^sup>2 * (ln (ln (real (n_of x))) + ln (1 / of_rat (\_of x)))))" - have evt: "(\x. - 0 < real_of_rat (\_of x) \ 0 < real_of_rat (\_of x) \ + have evt: "(\x. + 0 < real_of_rat (\_of x) \ 0 < real_of_rat (\_of x) \ 1/real_of_rat (\_of x) \ \ \ 1/real_of_rat (\_of x) \ \ \ real (n_of x) \ n \ P x) \ eventually P ?F" (is "(\x. ?prem x \ _) \ _") for \ \ n P apply (rule eventually_mono[where P="?prem" and Q="P"]) apply (simp add:\_of_def case_prod_beta' \_of_def n_of_def) - apply (intro eventually_conj eventually_prod1' eventually_prod2' + apply (intro eventually_conj eventually_prod1' eventually_prod2' sequentially_inf eventually_at_right_less inv_at_right_0_inf) by (auto simp add:prod_filter_eq_bot) have exp_pos: "exp k \ real x \ x > 0" for k x - using exp_gt_zero gr0I by force + using exp_gt_zero gr0I by force have exp_gt_1: "exp 1 \ (1::real)" by simp have 1: "(\_. 1) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" by (auto intro!:landau_o.big_mono evt[where \="exp 1"] iffD2[OF ln_ge_iff] simp add:abs_ge_iff) - have 2: "(\_. 1) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" + have 2: "(\_. 1) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" by (auto intro!:landau_o.big_mono evt[where \="exp 1"] iffD2[OF ln_ge_iff] simp add:abs_ge_iff) have 3: " (\x. 1) \ O[?F](\x. ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x)))" using exp_pos by (intro landau_sum_2 2 evt[where n="exp 1" and \="1"] ln_ge_zero iffD2[OF ln_ge_iff], auto) - have 4: "(\_. 1) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" + have 4: "(\_. 1) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" using one_le_power by (intro landau_o.big_mono evt[where \="1"], auto simp add:power_one_over[symmetric]) have "(\x. 80 * (1 / (real_of_rat (\_of x))\<^sup>2)) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" by (subst landau_o.big.cmult_in_iff, auto) hence 5: "(\x. real (t_of x)) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" - unfolding t_of_def + unfolding t_of_def by (intro landau_real_nat landau_ceil 4, auto) have "(\x. ln (real_of_rat (\_of x))) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" by (intro landau_o.big_mono evt[where \="1"], auto simp add:ln_div) hence 6: "(\x. real (s_of x)) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" unfolding s_of_def by (intro landau_nat_ceil 1, simp) have 7: " (\x. 1) \ O[?F](\x. ln (real (n_of x)))" using exp_pos by (auto intro!: landau_o.big_mono evt[where n="exp 1"] iffD2[OF ln_ge_iff] simp: abs_ge_iff) - have 8:" (\_. 1) \ + have 8:" (\_. 1) \ O[?F](\x. ln (real (n_of x)) + 1 / (real_of_rat (\_of x))\<^sup>2 * (ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x))))" using order_trans[OF exp_gt_1] exp_pos - by (intro landau_sum_1 7 evt[where n="exp 1" and \="1"] ln_ge_zero iffD2[OF ln_ge_iff] + by (intro landau_sum_1 7 evt[where n="exp 1" and \="1"] ln_ge_zero iffD2[OF ln_ge_iff] mult_nonneg_nonneg add_nonneg_nonneg) auto have "(\x. ln (real (s_of x) + 1)) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" by (intro landau_ln_3 sum_in_bigo 6 1, simp) hence 9: "(\x. log 2 (real (s_of x) + 1)) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1 8, auto simp:log_def) have 10: "(\x. 1) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1 8 1) - have "(\x. ln (real (t_of x) + 1)) \ + have "(\x. ln (real (t_of x) + 1)) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2 * (ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x))))" using 5 by (intro landau_o.big_mult_1 3 landau_ln_3 sum_in_bigo 4, simp_all) - hence " (\x. log 2 (real (t_of x) + 1)) \ + hence " (\x. log 2 (real (t_of x) + 1)) \ O[?F](\x. ln (real (n_of x)) + 1 / (real_of_rat (\_of x))\<^sup>2 * (ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x))))" using order_trans[OF exp_gt_1] exp_pos - by (intro landau_sum_2 evt[where n="exp 1" and \="1"] ln_ge_zero iffD2[OF ln_ge_iff] + by (intro landau_sum_2 evt[where n="exp 1" and \="1"] ln_ge_zero iffD2[OF ln_ge_iff] mult_nonneg_nonneg add_nonneg_nonneg) (auto simp add:log_def) hence 11: "(\x. log 2 (real (t_of x) + 1)) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1' 1, auto) - have " (\x. 1) \ O[?F](\x. real (n_of x))" + have " (\x. 1) \ O[?F](\x. real (n_of x))" by (intro landau_o.big_mono evt[where n="1"], auto) - hence "(\x. ln (real (n_of x) + 21)) \ O[?F](\x. ln (real (n_of x)))" + hence "(\x. ln (real (n_of x) + 21)) \ O[?F](\x. ln (real (n_of x)))" by (intro landau_ln_2[where a="2"] evt[where n="2"] sum_in_bigo, auto) hence 12: "(\x. log 2 (real (n_of x) + 21)) \ O[?F](g)" unfolding g_def using exp_pos order_trans[OF exp_gt_1] - by (intro landau_o.big_mult_1' 1 landau_sum_1 evt[where n="exp 1" and \="1"] + by (intro landau_o.big_mult_1' 1 landau_sum_1 evt[where n="exp 1" and \="1"] ln_ge_zero iffD2[OF ln_ge_iff] mult_nonneg_nonneg add_nonneg_nonneg) (auto simp add:log_def) - have "(\x. ln (1 / real_of_rat (\_of x))) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" - by (intro landau_ln_3 evt[where \="1"] landau_o.big_mono) + have "(\x. ln (1 / real_of_rat (\_of x))) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" + by (intro landau_ln_3 evt[where \="1"] landau_o.big_mono) (auto simp add:power_one_over[symmetric] self_le_power) hence " (\x. real (nat (4*\log 2 (1 / real_of_rat (\_of x))\+23))) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" using 4 by (auto intro!: landau_real_nat sum_in_bigo landau_ceil simp:log_def) hence " (\x. ln (real (r_of x) + 1)) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" unfolding r_of_def by (intro landau_ln_3 sum_in_bigo 4, auto) - hence " (\x. log 2 (real (r_of x) + 1)) \ + hence " (\x. log 2 (real (r_of x) + 1)) \ O[?F](\x. (1 / (real_of_rat (\_of x))\<^sup>2) * (ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x))))" by (intro landau_o.big_mult_1 3, simp add:log_def) - hence " (\x. log 2 (real (r_of x) + 1)) \ + hence " (\x. log 2 (real (r_of x) + 1)) \ O[?F](\x. ln (real (n_of x)) + 1 / (real_of_rat (\_of x))\<^sup>2 * (ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x))))" using exp_pos order_trans[OF exp_gt_1] - by (intro landau_sum_2 evt[where n="exp 1" and \="1"] ln_ge_zero + by (intro landau_sum_2 evt[where n="exp 1" and \="1"] ln_ge_zero iffD2[OF ln_ge_iff] add_nonneg_nonneg mult_nonneg_nonneg) (auto) hence 13: "(\x. log 2 (real (r_of x) + 1)) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1' 1, auto) - have 14: "(\x. 1) \ O[?F](\x. real (n_of x))" + have 14: "(\x. 1) \ O[?F](\x. real (n_of x))" by (intro landau_o.big_mono evt[where n="1"], auto) - have "(\x. ln (real (n_of x) + 13)) \ O[?F](\x. ln (real (n_of x)))" + have "(\x. ln (real (n_of x) + 13)) \ O[?F](\x. ln (real (n_of x)))" using 14 by (intro landau_ln_2[where a="2"] evt[where n="2"] sum_in_bigo, auto) hence "(\x. ln (log 2 (real (n_of x) + 13))) \ O[?F](\x. ln (ln (real (n_of x))))" using exp_pos by (intro landau_ln_2[where a="2"] iffD2[OF ln_ge_iff] evt[where n="exp 2"]) (auto simp add:log_def) hence "(\x. log 2 (log 2 (real (n_of x) + 13))) \ O[?F](\x. ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x)))" using exp_pos by (intro landau_sum_1 evt[where n="exp 1" and \="1"] ln_ge_zero iffD2[OF ln_ge_iff]) (auto simp add:log_def) moreover have "(\x. real (r_of x)) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" unfolding r_of_def using 2 by (auto intro!: landau_real_nat sum_in_bigo landau_ceil simp:log_def) hence "(\x. real (r_of x)) \ O[?F](\x. ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x)))" - using exp_pos + using exp_pos by (intro landau_sum_2 evt[where n="exp 1" and \="1"] ln_ge_zero iffD2[OF ln_ge_iff], auto) ultimately have 15:" (\x. real (t_of x) * (13 + 4 * real (r_of x) + 2 * log 2 (log 2 (real (n_of x) + 13)))) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2 * (ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x))))" - using 5 3 + using 5 3 by (intro landau_o.mult sum_in_bigo, auto) have "(\x. 5 + 2 * log 2 (21 + real (n_of x)) + real (t_of x) * (13 + 4 * real (r_of x) + 2 * log 2 (log 2 (real (n_of x) + 13)))) \ O[?F](\x. ln (real (n_of x)) + 1 / (real_of_rat (\_of x))\<^sup>2 * (ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x))))" proof - - have "\\<^sub>F x in ?F. 0 \ ln (real (n_of x))" + have "\\<^sub>F x in ?F. 0 \ ln (real (n_of x))" by (intro evt[where n="1"] ln_ge_zero, auto) moreover have "\\<^sub>F x in ?F. 0 \ 1 / (real_of_rat (\_of x))\<^sup>2 * (ln (ln (real (n_of x))) + ln (1 / real_of_rat (\_of x)))" using exp_pos - by (intro evt[where n="exp 1" and \="1"] mult_nonneg_nonneg add_nonneg_nonneg + by (intro evt[where n="exp 1" and \="1"] mult_nonneg_nonneg add_nonneg_nonneg ln_ge_zero iffD2[OF ln_ge_iff]) auto - moreover have " (\x. ln (21 + real (n_of x))) \ O[?F](\x. ln (real (n_of x)))" + moreover have " (\x. ln (21 + real (n_of x))) \ O[?F](\x. ln (real (n_of x)))" using 14 by (intro landau_ln_2[where a="2"] sum_in_bigo evt[where n="2"], auto) hence "(\x. 5 + 2 * log 2 (21 + real (n_of x))) \ O[?F](\x. ln (real (n_of x)))" using 7 by (intro sum_in_bigo, auto simp add:log_def) ultimately show ?thesis using 15 by (rule landau_sum) qed hence 16: "(\x. real (s_of x) * (5 + 2 * log 2 (21 + real (n_of x)) + real (t_of x) * (13 + 4 * real (r_of x) + 2 * log 2 (log 2 (real (n_of x) + 13))))) \ O[?F](g)" unfolding g_def by (intro landau_o.mult 6, auto) have "f0_space_usage = (\x. f0_space_usage (n_of x, \_of x, \_of x))" by (simp add:case_prod_beta' n_of_def \_of_def \_of_def) also have "... \ O[?F](g)" using 9 10 11 12 13 16 - by (simp add:fun_cong[OF s_of_def[symmetric]] fun_cong[OF t_of_def[symmetric]] + by (simp add:fun_cong[OF s_of_def[symmetric]] fun_cong[OF t_of_def[symmetric]] fun_cong[OF r_of_def[symmetric]] Let_def) (intro sum_in_bigo, auto) also have "... = O[?F](?rhs)" by (simp add:case_prod_beta' g_def n_of_def \_of_def \_of_def) finally show ?thesis by simp qed end diff --git a/thys/Frequency_Moments/Frequency_Moment_2.thy b/thys/Frequency_Moments/Frequency_Moment_2.thy --- a/thys/Frequency_Moments/Frequency_Moment_2.thy +++ b/thys/Frequency_Moments/Frequency_Moment_2.thy @@ -1,721 +1,725 @@ section \Frequency Moment $2$\ theory Frequency_Moment_2 imports Universal_Hash_Families.Carter_Wegman_Hash_Family Universal_Hash_Families.Universal_Hash_Families_More_Finite_Fields Equivalence_Relation_Enumeration.Equivalence_Relation_Enumeration Landau_Ext Median_Method.Median + Probability_Ext Product_PMF_Ext - Frequency_Moments + Frequency_Moments begin +hide_const (open) Discrete_Topology.discrete +hide_const (open) Isolated.discrete + text \This section contains a formalization of the algorithm for the second frequency moment. It is based on the algorithm described in \<^cite>\\\textsection 2.2\ in "alon1999"\. The only difference is that the algorithm is adapted to work with prime field of odd order, which greatly reduces the implementation complexity.\ fun f2_hash where "f2_hash p h k = (if even (ring.hash (mod_ring p) k h) then int p - 1 else - int p - 1)" type_synonym f2_state = "nat \ nat \ nat \ (nat \ nat \ nat list) \ (nat \ nat \ int)" fun f2_init :: "rat \ rat \ nat \ f2_state pmf" where "f2_init \ \ n = do { let s\<^sub>1 = nat \6 / \\<^sup>2\; let s\<^sub>2 = nat \-(18 * ln (real_of_rat \))\; let p = prime_above (max n 3); h \ prod_pmf ({..1} \ {..2}) (\_. pmf_of_set (bounded_degree_polynomials (mod_ring p) 4)); return_pmf (s\<^sub>1, s\<^sub>2, p, h, (\_ \ {..1} \ {..2}. (0 :: int))) }" fun f2_update :: "nat \ f2_state \ f2_state pmf" where - "f2_update x (s\<^sub>1, s\<^sub>2, p, h, sketch) = + "f2_update x (s\<^sub>1, s\<^sub>2, p, h, sketch) = return_pmf (s\<^sub>1, s\<^sub>2, p, h, \i \ {..1} \ {..2}. f2_hash p (h i) x + sketch i)" fun f2_result :: "f2_state \ rat pmf" where - "f2_result (s\<^sub>1, s\<^sub>2, p, h, sketch) = - return_pmf (median s\<^sub>2 (\i\<^sub>2 \ {..2}. + "f2_result (s\<^sub>1, s\<^sub>2, p, h, sketch) = + return_pmf (median s\<^sub>2 (\i\<^sub>2 \ {..2}. (\i\<^sub>1\{..1} . (rat_of_int (sketch (i\<^sub>1, i\<^sub>2)))\<^sup>2) / (((rat_of_nat p)\<^sup>2-1) * rat_of_nat s\<^sub>1)))" fun f2_space_usage :: "(nat \ nat \ rat \ rat) \ real" where "f2_space_usage (n, m, \, \) = ( let s\<^sub>1 = nat \6 / \\<^sup>2 \ in - let s\<^sub>2 = nat \-(18 * ln (real_of_rat \))\ in + let s\<^sub>2 = nat \-(18 * ln (real_of_rat \))\ in 3 + 2 * log 2 (s\<^sub>1 + 1) + 2 * log 2 (s\<^sub>2 + 1) + 2 * log 2 (9 + 2 * real n) + s\<^sub>1 * s\<^sub>2 * (5 + 4*log 2 (8 + 2 * real n) + 2 * log 2 (real m * (18 + 4 * real n) + 1 )))" definition encode_f2_state :: "f2_state \ bool list option" where - "encode_f2_state = - N\<^sub>e \\<^sub>e (\s\<^sub>1. - N\<^sub>e \\<^sub>e (\s\<^sub>2. - N\<^sub>e \\<^sub>e (\p. + "encode_f2_state = + N\<^sub>e \\<^sub>e (\s\<^sub>1. + N\<^sub>e \\<^sub>e (\s\<^sub>2. + N\<^sub>e \\<^sub>e (\p. (List.product [0..1] [0..2] \\<^sub>e P\<^sub>e p 4) \\<^sub>e (List.product [0..1] [0..2] \\<^sub>e I\<^sub>e))))" lemma "inj_on encode_f2_state (dom encode_f2_state)" proof - have " is_encoding encode_f2_state" unfolding encode_f2_state_def by (intro dependent_encoding exp_golomb_encoding fun_encoding list_encoding int_encoding poly_encoding) - + thus ?thesis by (rule encoding_imp_inj) qed context fixes \ \ :: rat fixes n :: nat fixes as :: "nat list" fixes result assumes \_range: "\ \ {0<..<1}" assumes \_range: "\ > 0" assumes as_range: "set as \ {.. fold (\a state. state \ f2_update a) as (f2_init \ \ n) \ f2_result" -begin +begin private definition s\<^sub>1 where "s\<^sub>1 = nat \6 / \\<^sup>2\" lemma s1_gt_0: "s\<^sub>1 > 0" using \_range by (simp add:s\<^sub>1_def) private definition s\<^sub>2 where "s\<^sub>2 = nat \-(18* ln (real_of_rat \))\" lemma s2_gt_0: "s\<^sub>2 > 0" using \_range by (simp add:s\<^sub>2_def) private definition p where "p = prime_above (max n 3)" - -lemma p_prime: "Factorial_Ring.prime p" + +lemma p_prime: "Factorial_Ring.prime p" unfolding p_def using prime_above_prime by blast lemma p_ge_3: "p \ 3" unfolding p_def by (meson max.boundedE prime_above_lower_bound) lemma p_gt_0: "p > 0" using p_ge_3 by linarith lemma p_gt_1: "p > 1" using p_ge_3 by simp lemma p_ge_n: "p \ n" unfolding p_def by (meson max.boundedE prime_above_lower_bound ) interpretation carter_wegman_hash_family "mod_ring p" 4 using carter_wegman_hash_familyI[OF mod_ring_is_field mod_ring_finite] using p_prime by auto definition sketch where "sketch = fold (\a state. state \ f2_update a) as (f2_init \ \ n)" -private definition \ where"\ = prod_pmf ({..1} \ {..2}) (\_. pmf_of_set space)" -private definition \\<^sub>p where"\\<^sub>p = measure_pmf \" +private definition \ where"\ = prod_pmf ({..1} \ {..2}) (\_. pmf_of_set space)" +private definition \\<^sub>p where"\\<^sub>p = measure_pmf \" private definition sketch_rv where "sketch_rv \ = of_int (sum_list (map (f2_hash p \) as))^2" private definition mean_rv where "mean_rv \ = (\i\<^sub>2. (\i\<^sub>1 = 0..1. sketch_rv (\ (i\<^sub>1, i\<^sub>2))) / (((of_nat p)\<^sup>2 - 1) * of_nat s\<^sub>1))" private definition result_rv where "result_rv \ = median s\<^sub>2 (\i\<^sub>2\{..2}. mean_rv \ i\<^sub>2)" lemma mean_rv_alg_sketch: "sketch = \ \ (\\. return_pmf (s\<^sub>1, s\<^sub>2, p, \, \i \ {..1} \ {..2}. sum_list (map (f2_hash p (\ i)) as)))" proof - have "sketch = fold (\a state. state \ f2_update a) as (f2_init \ \ n)" by (simp add:sketch_def) - also have "... = \ \ (\\. return_pmf (s\<^sub>1, s\<^sub>2, p, \, + also have "... = \ \ (\\. return_pmf (s\<^sub>1, s\<^sub>2, p, \, \i \ {..1} \ {..2}. sum_list (map (f2_hash p (\ i)) as)))" proof (induction as rule:rev_induct) case Nil - then show ?case - by (simp add:s\<^sub>1_def s\<^sub>2_def space_def p_def[symmetric] \_def restrict_def Let_def) + then show ?case + by (simp add:s\<^sub>1_def s\<^sub>2_def space_def p_def[symmetric] \_def restrict_def Let_def) next case (snoc a as) - have "fold (\a state. state \ f2_update a) (as @ [a]) (f2_init \ \ n) = \ \ + have "fold (\a state. state \ f2_update a) (as @ [a]) (f2_init \ \ n) = \ \ (\\. return_pmf (s\<^sub>1, s\<^sub>2, p, \, \s \ {..1} \ {..2}. (\x \ as. f2_hash p (\ s) x)) \ f2_update a)" using snoc by (simp add: bind_assoc_pmf restrict_def del:f2_hash.simps f2_init.simps) also have "... = \ \ (\\. return_pmf (s\<^sub>1, s\<^sub>2, p, \, \i \ {..1} \ {..2}. (\x \ as@[a]. f2_hash p (\ i) x)))" by (subst bind_return_pmf) (simp add: add.commute del:f2_hash.simps cong:restrict_cong) finally show ?case by blast qed finally show ?thesis by auto qed lemma distr: "result = map_pmf result_rv \" proof - have "result = sketch \ f2_result" by (simp add:result_def sketch_def) also have "... = \ \ (\x. f2_result (s\<^sub>1, s\<^sub>2, p, x, \i\{..1} \ {..2}. sum_list (map (f2_hash p (x i)) as)))" by (simp add: mean_rv_alg_sketch bind_assoc_pmf bind_return_pmf) also have "... = map_pmf result_rv \" by (simp add:map_pmf_def result_rv_def mean_rv_def sketch_rv_def lessThan_atLeast0 cong:restrict_cong) finally show ?thesis by simp qed private lemma f2_hash_pow_exp: assumes "k < p" shows - "expectation (\\. real_of_int (f2_hash p \ k) ^m) = + "expectation (\\. real_of_int (f2_hash p \ k) ^m) = ((real p - 1) ^ m * (real p + 1) + (- real p - 1) ^ m * (real p - 1)) / (2 * real p)" proof - have "odd p" using p_prime p_ge_3 prime_odd_nat assms by simp - then obtain t where t_def: "p=2*t+1" + then obtain t where t_def: "p=2*t+1" using oddE by blast - have "Collect even \ {..<2 * t + 1} \ (*) 2 ` {.. {..<2 * t + 1} \ (*) 2 ` {.. Collect even \ {..<2 * t + 1}" by (rule image_subsetI, simp) ultimately have "card ({k. even k} \ {..x. 2*x) ` {.. {.. {... hash k \ \ Collect even} = (real p + 1)/(2*real p)" - using assms by (subst prob_range, auto simp:frac_eq_eq p_gt_0 mod_ring_def) + using assms by (subst prob_range, auto simp:frac_eq_eq p_gt_0 mod_ring_def) have "p = card {.. {.. ({k. even k} \ {.. {.. ({k. even k} \ {.. {.. {.. {.. {.. {.. {... hash k \ \ Collect odd} = (real p - 1)/(2*real p)" using assms by (subst prob_range, auto simp add: frac_eq_eq mod_ring_def) have "expectation (\x. real_of_int (f2_hash p x k) ^ m) = - expectation (\\. indicator {\. even (hash k \)} \ * (real p - 1)^m + - indicator {\. odd (hash k \)} \ * (-real p - 1)^m)" + expectation (\\. indicator {\. even (hash k \)} \ * (real p - 1)^m + + indicator {\. odd (hash k \)} \ * (-real p - 1)^m)" by (rule Bochner_Integration.integral_cong, simp, simp) - also have "... = - prob {\. hash k \ \ Collect even} * (real p - 1) ^ m + + also have "... = + prob {\. hash k \ \ Collect even} * (real p - 1) ^ m + prob {\. hash k \ \ Collect odd} * (-real p - 1) ^ m " by (simp, simp add:M_def) also have "... = (real p + 1) * (real p - 1) ^ m / (2 * real p) + (real p - 1) * (- real p - 1) ^ m / (2 * real p)" by (subst prob_even, subst prob_odd, simp) - also have "... = + also have "... = ((real p - 1) ^ m * (real p + 1) + (- real p - 1) ^ m * (real p - 1)) / (2 * real p)" by (simp add:add_divide_distrib ac_simps) - finally show "expectation (\x. real_of_int (f2_hash p x k) ^ m) = + finally show "expectation (\x. real_of_int (f2_hash p x k) ^ m) = ((real p - 1) ^ m * (real p + 1) + (- real p - 1) ^ m * (real p - 1)) / (2 * real p)" by simp qed -lemma +lemma shows var_sketch_rv:"variance sketch_rv \ 2*(real_of_rat (F 2 as)^2) * ((real p)\<^sup>2-1)\<^sup>2" (is "?A") and exp_sketch_rv:"expectation sketch_rv = real_of_rat (F 2 as) * ((real p)\<^sup>2-1)" (is "?B") proof - define h where "h = (\\ x. real_of_int (f2_hash p \ x))" define c where "c = (\x. real (count_list as x))" define r where "r = (\(m::nat). ((real p - 1) ^ m * (real p + 1) + (- real p - 1) ^ m * (real p - 1)) / (2 * real p))" - define h_prod where "h_prod = (\as \. prod_list (map (h \) as))" + define h_prod where "h_prod = (\as \. prod_list (map (h \) as))" define exp_h_prod :: "nat list \ real" where "exp_h_prod = (\as. (\i \ set as. r (count_list as i)))" have f_eq: "sketch_rv = (\\. (\x \ set as. c x * h \ x)^2)" by (rule ext, simp add:sketch_rv_def c_def h_def sum_list_eval del:f2_hash.simps) have r_one: "r (Suc 0) = 0" by (simp add:r_def algebra_simps) have r_two: "r 2 = (real p^2-1)" - using p_gt_0 unfolding r_def power2_eq_square + using p_gt_0 unfolding r_def power2_eq_square by (simp add:nonzero_divide_eq_eq, simp add:algebra_simps) have"(real p)^2 \ 2^2" by (rule power_mono, use p_gt_1 in linarith, simp) hence p_square_ge_4: "(real p)\<^sup>2 \ 4" by simp - have "r 4 = (real p)^4+2*(real p)\<^sup>2 - 3" + have "r 4 = (real p)^4+2*(real p)\<^sup>2 - 3" using p_gt_0 unfolding r_def by (subst nonzero_divide_eq_eq, auto simp:power4_eq_xxxx power2_eq_square algebra_simps) also have "... \ (real p)^4+2*(real p)\<^sup>2 + 3" by simp also have "... \ 3 * r 2 * r 2" using p_square_ge_4 by (simp add:r_two power4_eq_xxxx power2_eq_square algebra_simps mult_left_mono) finally have r_four_est: "r 4 \ 3 * r 2 * r 2" by simp - have exp_h_prod_elim: "exp_h_prod = (\as. prod_list (map (r \ count_list as) (remdups as)))" + have exp_h_prod_elim: "exp_h_prod = (\as. prod_list (map (r \ count_list as) (remdups as)))" by (simp add:exp_h_prod_def prod.set_conv_list[symmetric]) have exp_h_prod: "\x. set x \ set as \ length x \ 4 \ expectation (h_prod x) = exp_h_prod x" proof - - fix x + fix x assume "set x \ set as" hence x_sub_p: "set x \ {..k. k \ set x \ k < p" by auto assume "length x \ 4" hence card_x: "card (set x) \ 4" using card_length dual_order.trans by blast have "set x \ carrier (mod_ring p) " using x_sub_p by (simp add:mod_ring_def) hence h_indep: "indep_vars (\_. borel) (\i \. h \ i ^ count_list x i) (set x)" using k_wise_indep_vars_subset[OF k_wise_indep] card_x as_range h_def by (auto intro:indep_vars_compose2[where X="hash" and M'=" (\_. discrete)"]) have "expectation (h_prod x) = expectation (\\. \ i \ set x. h \ i^(count_list x i))" by (simp add:h_prod_def prod_list_eval) also have "... = (\i \ set x. expectation (\\. h \ i^(count_list x i)))" by (simp add: indep_vars_lebesgue_integral[OF _ h_indep]) also have "... = (\i \ set x. r (count_list x i))" - using f2_hash_pow_exp x_le_p + using f2_hash_pow_exp x_le_p by (simp add:h_def r_def M_def[symmetric] del:f2_hash.simps) also have "... = exp_h_prod x" by (simp add:exp_h_prod_def) finally show "expectation (h_prod x) = exp_h_prod x" by simp qed - have "\x y. kernel_of x = kernel_of y \ exp_h_prod x = exp_h_prod y" + have "\x y. kernel_of x = kernel_of y \ exp_h_prod x = exp_h_prod y" proof - fix x y :: "nat list" assume a:"kernel_of x = kernel_of y" then obtain f where b:"bij_betw f (set x) (set y)" and c:"\z. z \ set x \ count_list x z = count_list y (f z)" using kernel_of_eq_imp_bij by blast have "exp_h_prod x = prod ( (\i. r(count_list y i)) \ f) (set x)" by (simp add:exp_h_prod_def c) also have "... = (\i \ f ` (set x). r(count_list y i))" by (metis b bij_betw_def prod.reindex) also have "... = exp_h_prod y" unfolding exp_h_prod_def by (rule prod.cong, metis b bij_betw_def) simp finally show "exp_h_prod x = exp_h_prod y" by simp qed - hence exp_h_prod_cong: "\p x. of_bool (kernel_of x = kernel_of p) * exp_h_prod p = - of_bool (kernel_of x = kernel_of p) * exp_h_prod x" + hence exp_h_prod_cong: "\p x. of_bool (kernel_of x = kernel_of p) * exp_h_prod p = + of_bool (kernel_of x = kernel_of p) * exp_h_prod x" by (metis (full_types) of_bool_eq_0_iff vector_space_over_itself.scale_zero_left) have c:"(\p\enum_rgfs n. of_bool (kernel_of xs = kernel_of p) * r) = r" if a:"length xs = n" for xs :: "nat list" and n and r :: real proof - have "(\p\enum_rgfs n. of_bool (kernel_of xs = kernel_of p) * 1) = (1::real)" - using equiv_rels_2[OF a[symmetric]] by (simp add:equiv_rels_def comp_def) - thus "(\p\enum_rgfs n. of_bool (kernel_of xs = kernel_of p) * r) = (r::real)" + using equiv_rels_2[OF a[symmetric]] by (simp add:equiv_rels_def comp_def) + thus "(\p\enum_rgfs n. of_bool (kernel_of xs = kernel_of p) * r) = (r::real)" by (simp add:sum_list_mult_const) qed have "expectation sketch_rv = (\i\set as. (\j\set as. c i * c j * expectation (h_prod [i,j])))" by (simp add:f_eq h_prod_def power2_eq_square sum_distrib_left sum_distrib_right Bochner_Integration.integral_sum algebra_simps) also have "... = (\i\set as. (\j\set as. c i * c j * exp_h_prod [i,j]))" by (simp add:exp_h_prod) - also have "... = (\i \ set as. (\j \ set as. + also have "... = (\i \ set as. (\j \ set as. c i * c j * (sum_list (map (\p. of_bool (kernel_of [i,j] = kernel_of p) * exp_h_prod p) (enum_rgfs 2)))))" by (subst exp_h_prod_cong, simp add:c) also have "... = (\i\set as. c i * c i * r 2)" by (simp add: numeral_eq_Suc kernel_of_eq All_less_Suc exp_h_prod_elim r_one distrib_left sum.distrib sum_collapse) also have "... = real_of_rat (F 2 as) * ((real p)^2-1)" by (simp add: sum_distrib_right[symmetric] c_def F_def power2_eq_square of_rat_sum of_rat_mult r_two) finally show b:?B by simp have "expectation (\x. (sketch_rv x)\<^sup>2) = (\i1 \ set as. (\i2 \ set as. (\i3 \ set as. (\i4 \ set as. c i1 * c i2 * c i3 * c i4 * expectation (h_prod [i1, i2, i3, i4])))))" by (simp add:f_eq h_prod_def power4_eq_xxxx sum_distrib_left sum_distrib_right Bochner_Integration.integral_sum algebra_simps) - also have "... = (\i1 \ set as. (\i2 \ set as. (\i3 \ set as. (\i4 \ set as. + also have "... = (\i1 \ set as. (\i2 \ set as. (\i3 \ set as. (\i4 \ set as. c i1 * c i2 * c i3 * c i4 * exp_h_prod [i1,i2,i3,i4]))))" by (simp add:exp_h_prod) - also have "... = (\i1 \ set as. (\i2 \ set as. (\i3 \ set as. (\i4 \ set as. - c i1 * c i2 * c i3 * c i4 * + also have "... = (\i1 \ set as. (\i2 \ set as. (\i3 \ set as. (\i4 \ set as. + c i1 * c i2 * c i3 * c i4 * (sum_list (map (\p. of_bool (kernel_of [i1,i2,i3,i4] = kernel_of p) * exp_h_prod p) (enum_rgfs 4)))))))" by (subst exp_h_prod_cong, simp add:c) - also have "... = + also have "... = 3 * (\i \ set as. (\j \ set as. c i^2 * c j^2 * r 2 * r 2)) + ((\ i \ set as. c i^4 * r 4) - 3 * (\ i \ set as. c i ^ 4 * r 2 * r 2))" apply (simp add: numeral_eq_Suc exp_h_prod_elim r_one) (* large intermediate terms *) apply (simp add: kernel_of_eq All_less_Suc numeral_eq_Suc distrib_left sum.distrib sum_collapse neq_commute of_bool_not_iff) apply (simp add: algebra_simps sum_subtractf sum_collapse) apply (simp add: sum_distrib_left algebra_simps) done also have "... = 3 * (\i \ set as. c i^2 * r 2)^2 + (\ i \ set as. c i ^ 4 * (r 4 - 3 * r 2 * r 2))" by (simp add:power2_eq_square sum_distrib_left algebra_simps sum_subtractf) also have "... = 3 * (\i \ set as. c i^2)^2 * (r 2)^2 + (\i \ set as. c i ^ 4 * (r 4 - 3 * r 2 * r 2))" by (simp add:power_mult_distrib sum_distrib_right[symmetric]) also have "... \ 3 * (\i \ set as. c i^2)^2 * (r 2)^2 + (\i \ set as. c i ^ 4 * 0)" - using r_four_est + using r_four_est by (auto intro!: sum_nonpos simp add:mult_nonneg_nonpos) - also have "... = 3 * (real_of_rat (F 2 as)^2) * ((real p)\<^sup>2-1)\<^sup>2" + also have "... = 3 * (real_of_rat (F 2 as)^2) * ((real p)\<^sup>2-1)\<^sup>2" by (simp add:c_def r_two F_def of_rat_sum of_rat_power) finally have "expectation (\x. (sketch_rv x)\<^sup>2) \ 3 * (real_of_rat (F 2 as)^2) * ((real p)\<^sup>2-1)\<^sup>2" by simp - + thus "variance sketch_rv \ 2*(real_of_rat (F 2 as)^2) * ((real p)\<^sup>2-1)\<^sup>2" by (simp add: variance_eq, simp add:power_mult_distrib b) qed lemma space_omega_1 [simp]: "Sigma_Algebra.space \\<^sub>p = UNIV" by (simp add:\\<^sub>p_def) interpretation \: prob_space "\\<^sub>p" by (simp add:\\<^sub>p_def prob_space_measure_pmf) lemma integrable_\: fixes f :: "((nat \ nat) \ (nat list)) \ real" shows "integrable \\<^sub>p f" unfolding \\<^sub>p_def \_def by (rule integrable_measure_pmf_finite, auto intro:finite_PiE simp:set_prod_pmf) lemma sketch_rv_exp: assumes "i\<^sub>2 < s\<^sub>2" assumes "i\<^sub>1 \ {0..1}" shows "\.expectation (\\. sketch_rv (\ (i\<^sub>1, i\<^sub>2))) = real_of_rat (F 2 as) * ((real p)\<^sup>2 - 1)" proof - have "\.expectation (\\. (sketch_rv (\ (i\<^sub>1, i\<^sub>2))) :: real) = expectation sketch_rv" using integrable_\ integrable_M assms unfolding \_def \\<^sub>p_def M_def by (subst expectation_Pi_pmf_slice, auto) also have "... = (real_of_rat (F 2 as)) * ((real p)\<^sup>2 - 1)" using exp_sketch_rv by simp finally show ?thesis by simp qed lemma sketch_rv_var: assumes "i\<^sub>2 < s\<^sub>2" assumes "i\<^sub>1 \ {0..1}" shows "\.variance (\\. sketch_rv (\ (i\<^sub>1, i\<^sub>2))) \ 2 * (real_of_rat (F 2 as))\<^sup>2 * ((real p)\<^sup>2 - 1)\<^sup>2" proof - have "\.variance (\\. (sketch_rv (\ (i\<^sub>1, i\<^sub>2)) :: real)) = variance sketch_rv" using integrable_\ integrable_M assms unfolding \_def \\<^sub>p_def M_def by (subst variance_prod_pmf_slice, auto) also have "... \ 2 * (real_of_rat (F 2 as))\<^sup>2 * ((real p)\<^sup>2 - 1)\<^sup>2" using var_sketch_rv by simp finally show ?thesis by simp qed lemma mean_rv_exp: assumes "i < s\<^sub>2" shows "\.expectation (\\. mean_rv \ i) = real_of_rat (F 2 as)" proof - have a:"(real p)\<^sup>2 > 1" using p_gt_1 by simp have "\.expectation (\\. mean_rv \ i) = (\i\<^sub>1 = 0..1. \.expectation (\\. sketch_rv (\ (i\<^sub>1, i)))) / (((real p)\<^sup>2 - 1) * real s\<^sub>1)" using assms integrable_\ by (simp add:mean_rv_def) - also have "... = (\i\<^sub>1 = 0..1. real_of_rat (F 2 as) * ((real p)\<^sup>2 - 1)) / (((real p)\<^sup>2 - 1) * real s\<^sub>1)" + also have "... = (\i\<^sub>1 = 0..1. real_of_rat (F 2 as) * ((real p)\<^sup>2 - 1)) / (((real p)\<^sup>2 - 1) * real s\<^sub>1)" using sketch_rv_exp[OF assms] by simp also have "... = real_of_rat (F 2 as)" using s1_gt_0 a by simp finally show ?thesis by simp qed lemma mean_rv_var: assumes "i < s\<^sub>2" shows "\.variance (\\. mean_rv \ i) \ (real_of_rat (\ * F 2 as))\<^sup>2 / 3" proof - have a: "\.indep_vars (\_. borel) (\i\<^sub>1 x. sketch_rv (x (i\<^sub>1, i))) {0..1}" using assms unfolding \\<^sub>p_def \_def by (intro indep_vars_restrict_intro'[where f="fst"]) (auto simp add: restrict_dfl_def case_prod_beta lessThan_atLeast0) - have p_sq_ne_1: "(real p)^2 \ 1" + have p_sq_ne_1: "(real p)^2 \ 1" by (metis p_gt_1 less_numeral_extra(4) of_nat_power one_less_power pos2 semiring_char_0_class.of_nat_eq_1_iff) have s1_bound: " 6 / (real_of_rat \)\<^sup>2 \ real s\<^sub>1" unfolding s\<^sub>1_def by (metis (mono_tags, opaque_lifting) of_rat_ceiling of_rat_divide of_rat_numeral_eq of_rat_power real_nat_ceiling_ge) have "\.variance (\\. mean_rv \ i) = \.variance (\\. \i\<^sub>1 = 0..1. sketch_rv (\ (i\<^sub>1, i))) / (((real p)\<^sup>2 - 1) * real s\<^sub>1)\<^sup>2" unfolding mean_rv_def by (subst \.variance_divide[OF integrable_\], simp) also have "... = (\i\<^sub>1 = 0..1. \.variance (\\. sketch_rv (\ (i\<^sub>1, i)))) / (((real p)\<^sup>2 - 1) * real s\<^sub>1)\<^sup>2" by (subst \.bienaymes_identity_full_indep[OF _ _ integrable_\ a]) (auto simp: \_def \\<^sub>p_def) also have "... \ (\i\<^sub>1 = 0..1. 2*(real_of_rat (F 2 as)^2) * ((real p)\<^sup>2-1)\<^sup>2) / (((real p)\<^sup>2 - 1) * real s\<^sub>1)\<^sup>2" by (rule divide_right_mono, rule sum_mono[OF sketch_rv_var[OF assms]], auto) also have "... = 2 * (real_of_rat (F 2 as)^2) / real s\<^sub>1" using p_sq_ne_1 s1_gt_0 by (subst frac_eq_eq, auto simp:power2_eq_square) also have "... \ 2 * (real_of_rat (F 2 as)^2) / (6 / (real_of_rat \)\<^sup>2)" using s1_gt_0 \_range by (intro divide_left_mono mult_pos_pos s1_bound) auto also have "... = (real_of_rat (\ * F 2 as))\<^sup>2 / 3" by (simp add:of_rat_mult algebra_simps) finally show ?thesis by simp qed lemma mean_rv_bounds: assumes "i < s\<^sub>2" shows "\.prob {\. real_of_rat \ * real_of_rat (F 2 as) < \mean_rv \ i - real_of_rat (F 2 as)\} \ 1/3" proof (cases "as = []") case True then show ?thesis using assms by (subst mean_rv_def, subst sketch_rv_def, simp add:F_def) next case False hence "F 2 as > 0" using F_gr_0 by auto hence a: "0 < real_of_rat (\ * F 2 as)" using \_range by simp have [simp]: "(\\. mean_rv \ i) \ borel_measurable \\<^sub>p" by (simp add:\_def \\<^sub>p_def) - have "\.prob {\. real_of_rat \ * real_of_rat (F 2 as) < \mean_rv \ i - real_of_rat (F 2 as)\} \ + have "\.prob {\. real_of_rat \ * real_of_rat (F 2 as) < \mean_rv \ i - real_of_rat (F 2 as)\} \ \.prob {\. real_of_rat (\ * F 2 as) \ \mean_rv \ i - real_of_rat (F 2 as)\}" by (rule \.pmf_mono[OF \\<^sub>p_def], simp add:of_rat_mult) also have "... \ \.variance (\\. mean_rv \ i) / (real_of_rat (\ * F 2 as))\<^sup>2" - using \.Chebyshev_inequality[where a="real_of_rat (\ * F 2 as)" and f="\\. mean_rv \ i",simplified] + using \.Chebyshev_inequality[where a="real_of_rat (\ * F 2 as)" and f="\\. mean_rv \ i",simplified] a prob_space_measure_pmf[where p="\"] mean_rv_exp[OF assms] integrable_\ by simp also have "... \ ((real_of_rat (\ * F 2 as))\<^sup>2/3) / (real_of_rat (\ * F 2 as))\<^sup>2" by (rule divide_right_mono, rule mean_rv_var[OF assms], simp) also have "... = 1/3" using a by force finally show ?thesis by blast qed lemma f2_alg_correct': "\

(\ in measure_pmf result. \\ - F 2 as\ \ \ * F 2 as) \ 1-of_rat \" proof - - have a: "\.indep_vars (\_. borel) (\i \. mean_rv \ i) {0..2}" + have a: "\.indep_vars (\_. borel) (\i \. mean_rv \ i) {0..2}" using s1_gt_0 unfolding \\<^sub>p_def \_def by (intro indep_vars_restrict_intro'[where f="snd"]) (auto simp: \\<^sub>p_def \_def mean_rv_def restrict_dfl_def) - have b: "- 18 * ln (real_of_rat \) \ real s\<^sub>2" + have b: "- 18 * ln (real_of_rat \) \ real s\<^sub>2" unfolding s\<^sub>2_def using of_nat_ceiling by auto have "1 - of_rat \ \ \.prob {\. \median s\<^sub>2 (mean_rv \) - real_of_rat (F 2 as) \ \ of_rat \ * of_rat (F 2 as)}" using \_range \.median_bound_2[OF _ a b, where \="real_of_rat \ * real_of_rat (F 2 as)" and \="real_of_rat (F 2 as)"] mean_rv_bounds by simp also have "... = \.prob {\. \real_of_rat (result_rv \) - of_rat (F 2 as) \ \ of_rat \ * of_rat (F 2 as)}" by (simp add:result_rv_def median_restrict lessThan_atLeast0 median_rat[OF s2_gt_0] mean_rv_def sketch_rv_def of_rat_divide of_rat_sum of_rat_mult of_rat_diff of_rat_power) - also have "... = \.prob {\. \result_rv \ - F 2 as\ \ \ * F 2 as} " + also have "... = \.prob {\. \result_rv \ - F 2 as\ \ \ * F 2 as} " by (simp add:of_rat_less_eq of_rat_mult[symmetric] of_rat_diff[symmetric] set_eq_iff) finally have "\.prob {y. \result_rv y - F 2 as\ \ \ * F 2 as} \ 1-of_rat \ " by simp thus ?thesis by (simp add: distr \\<^sub>p_def) qed lemma f2_exact_space_usage': "AE \ in sketch . bit_count (encode_f2_state \) \ f2_space_usage (n, length as, \, \)" proof - have "p \ 2 * max n 3 + 2" by (subst p_def, rule prime_above_upper_bound) also have "... \ 2 * n + 8" by (cases "n \ 2", simp_all) - finally have p_bound: "p \ 2 * n + 8" + finally have p_bound: "p \ 2 * n + 8" by simp have "bit_count (N\<^sub>e p) \ ereal (2 * log 2 (real p + 1) + 1)" by (rule exp_golomb_bit_count) also have "... \ ereal (2 * log 2 (2 * real n + 9) + 1)" using p_bound by simp finally have p_bit_count: "bit_count (N\<^sub>e p) \ ereal (2 * log 2 (2 * real n + 9) + 1)" by simp - have a: "bit_count (encode_f2_state (s\<^sub>1, s\<^sub>2, p, y, \i\{..1} \ {..2}. + have a: "bit_count (encode_f2_state (s\<^sub>1, s\<^sub>2, p, y, \i\{..1} \ {..2}. sum_list (map (f2_hash p (y i)) as))) \ ereal (f2_space_usage (n, length as, \, \))" if a:"y\{..1} \ {..2} \\<^sub>E bounded_degree_polynomials (mod_ring p) 4" for y proof - have "y \ extensional ({..1} \ {..2})" using a PiE_iff by blast hence y_ext: "y \ extensional (set (List.product [0..1] [0..2]))" by (simp add:lessThan_atLeast0) have h_bit_count_aux: "bit_count (P\<^sub>e p 4 (y x)) \ ereal (4 + 4 * log 2 (8 + 2 * real n))" if b:"x \ set (List.product [0..1] [0..2])" for x proof - have "y x \ bounded_degree_polynomials (mod_ring p) 4" using b a by force hence "bit_count (P\<^sub>e p 4 (y x)) \ ereal (real 4 * (log 2 (real p) + 1))" by (rule bounded_degree_polynomial_bit_count[OF p_gt_1] ) also have "... \ ereal (real 4 * (log 2 (8 + 2 * real n) + 1) )" using p_gt_0 p_bound by simp also have "... \ ereal (4 + 4 * log 2 (8 + 2 * real n))" by simp finally show ?thesis by blast qed - have h_bit_count: + have h_bit_count: "bit_count ((List.product [0..1] [0..2] \\<^sub>e P\<^sub>e p 4) y) \ ereal (real s\<^sub>1 * real s\<^sub>2 * (4 + 4 * log 2 (8 + 2 * real n)))" using fun_bit_count_est[where e="P\<^sub>e p 4", OF y_ext h_bit_count_aux] by simp have sketch_bit_count_aux: "bit_count (I\<^sub>e (sum_list (map (f2_hash p (y x)) as))) \ ereal (1 + 2 * log 2 (real (length as) * (18 + 4 * real n) + 1))" (is "?lhs \ ?rhs") if " x \ {0..1} \ {0..2}" for x proof - - have "\sum_list (map (f2_hash p (y x)) as)\ \ sum_list (map (abs \ (f2_hash p (y x))) as)" + have "\sum_list (map (f2_hash p (y x)) as)\ \ sum_list (map (abs \ (f2_hash p (y x))) as)" by (subst map_map[symmetric]) (rule sum_list_abs) also have "... \ sum_list (map (\_. (int p+1)) as)" - by (rule sum_list_mono) (simp add:p_gt_0) + by (rule sum_list_mono) (simp add:p_gt_0) also have "... = int (length as) * (int p+1)" by (simp add: sum_list_triv) also have "... \ int (length as) * (9+2*(int n))" using p_bound by (intro mult_mono, auto) finally have "\sum_list (map (f2_hash p (y x)) as)\ \ int (length as) * (9 + 2 * int n)" by simp hence "?lhs \ ereal (2 * log 2 (real_of_int (2* (int (length as) * (9 + 2 * int n)) + 1)) + 1)" by (rule int_bit_count_est) also have "... = ?rhs" by (simp add:algebra_simps) finally show "?thesis" by simp qed - have + have "bit_count ((List.product [0..1] [0..2] \\<^sub>e I\<^sub>e) (\i\{..1} \ {..2}. sum_list (map (f2_hash p (y i)) as))) \ ereal (real (length (List.product [0..1] [0..2]))) * (ereal (1 + 2 * log 2 (real (length as) * (18 + 4 * real n) + 1)))" - by (intro fun_bit_count_est) + by (intro fun_bit_count_est) (simp_all add:extensional_def lessThan_atLeast0 sketch_bit_count_aux del:f2_hash.simps) also have "... = ereal (real s\<^sub>1 * real s\<^sub>2 * (1 + 2 * log 2 (real (length as) * (18 + 4 * real n) + 1)))" by simp - finally have sketch_bit_count: + finally have sketch_bit_count: "bit_count ((List.product [0..1] [0..2] \\<^sub>e I\<^sub>e) (\i\{..1} \ {..2}. sum_list (map (f2_hash p (y i)) as))) \ ereal (real s\<^sub>1 * real s\<^sub>2 * (1 + 2 * log 2 (real (length as) * (18 + 4 * real n) + 1)))" by simp - have "bit_count (encode_f2_state (s\<^sub>1, s\<^sub>2, p, y, \i\{..1} \ {..2}. sum_list (map (f2_hash p (y i)) as))) \ + have "bit_count (encode_f2_state (s\<^sub>1, s\<^sub>2, p, y, \i\{..1} \ {..2}. sum_list (map (f2_hash p (y i)) as))) \ bit_count (N\<^sub>e s\<^sub>1) + bit_count (N\<^sub>e s\<^sub>2) +bit_count (N\<^sub>e p) + bit_count ((List.product [0..1] [0..2] \\<^sub>e P\<^sub>e p 4) y) + - bit_count ((List.product [0..1] [0..2] \\<^sub>e I\<^sub>e) (\i\{..1} \ {..2}. sum_list (map (f2_hash p (y i)) as)))" + bit_count ((List.product [0..1] [0..2] \\<^sub>e I\<^sub>e) (\i\{..1} \ {..2}. sum_list (map (f2_hash p (y i)) as)))" by (simp add:Let_def s\<^sub>1_def s\<^sub>2_def encode_f2_state_def dependent_bit_count add.assoc) - also have "... \ ereal (2 * log 2 (real s\<^sub>1 + 1) + 1) + ereal (2 * log 2 (real s\<^sub>2 + 1) + 1) + ereal (2 * log 2 (2 * real n + 9) + 1) + - (ereal (real s\<^sub>1 * real s\<^sub>2) * (4 + 4 * log 2 (8 + 2 * real n))) + + also have "... \ ereal (2 * log 2 (real s\<^sub>1 + 1) + 1) + ereal (2 * log 2 (real s\<^sub>2 + 1) + 1) + ereal (2 * log 2 (2 * real n + 9) + 1) + + (ereal (real s\<^sub>1 * real s\<^sub>2) * (4 + 4 * log 2 (8 + 2 * real n))) + (ereal (real s\<^sub>1 * real s\<^sub>2) * (1 + 2 * log 2 (real (length as) * (18 + 4 * real n) + 1) ))" by (intro add_mono exp_golomb_bit_count p_bit_count, auto intro: h_bit_count sketch_bit_count) also have "... = ereal (f2_space_usage (n, length as, \, \))" by (simp add:distrib_left add.commute s\<^sub>1_def[symmetric] s\<^sub>2_def[symmetric] Let_def) - finally show "bit_count (encode_f2_state (s\<^sub>1, s\<^sub>2, p, y, \i\{..1} \ {..2}. sum_list (map (f2_hash p (y i)) as))) \ - ereal (f2_space_usage (n, length as, \, \))" + finally show "bit_count (encode_f2_state (s\<^sub>1, s\<^sub>2, p, y, \i\{..1} \ {..2}. sum_list (map (f2_hash p (y i)) as))) \ + ereal (f2_space_usage (n, length as, \, \))" by simp qed have "set_pmf \ = {..1} \ {..2} \\<^sub>E bounded_degree_polynomials (mod_ring p) 4" by (simp add: \_def set_prod_pmf) (simp add: space_def) thus ?thesis by (simp add:mean_rv_alg_sketch AE_measure_pmf_iff del:f2_space_usage.simps, metis a) qed end text \Main results of this section:\ theorem f2_alg_correct: assumes "\ \ {0<..<1}" assumes "\ > 0" assumes "set as \ {.. \ fold (\a state. state \ f2_update a) as (f2_init \ \ n) \ f2_result" shows "\

(\ in measure_pmf \. \\ - F 2 as\ \ \ * F 2 as) \ 1-of_rat \" using f2_alg_correct'[OF assms(1,2,3)] \_def by auto theorem f2_exact_space_usage: assumes "\ \ {0<..<1}" assumes "\ > 0" assumes "set as \ {.. fold (\a state. state \ f2_update a) as (f2_init \ \ n)" shows "AE \ in M. bit_count (encode_f2_state \) \ f2_space_usage (n, length as, \, \)" using f2_exact_space_usage'[OF assms(1,2,3)] by (subst (asm) sketch_def[OF assms(1,2,3)], subst M_def, simp) theorem f2_asymptotic_space_complexity: - "f2_space_usage \ O[at_top \\<^sub>F at_top \\<^sub>F at_right 0 \\<^sub>F at_right 0](\ (n, m, \, \). + "f2_space_usage \ O[at_top \\<^sub>F at_top \\<^sub>F at_right 0 \\<^sub>F at_right 0](\ (n, m, \, \). (ln (1 / of_rat \)) / (of_rat \)\<^sup>2 * (ln (real n) + ln (real m)))" (is "_ \ O[?F](?rhs)") proof - define n_of :: "nat \ nat \ rat \ rat \ nat" where "n_of = (\(n, m, \, \). n)" define m_of :: "nat \ nat \ rat \ rat \ nat" where "m_of = (\(n, m, \, \). m)" define \_of :: "nat \ nat \ rat \ rat \ rat" where "\_of = (\(n, m, \, \). \)" define \_of :: "nat \ nat \ rat \ rat \ rat" where "\_of = (\(n, m, \, \). \)" define g where "g = (\x. (1/ (of_rat (\_of x))\<^sup>2) * (ln (1 / of_rat (\_of x))) * (ln (real (n_of x)) + ln (real (m_of x))))" - have evt: "(\x. - 0 < real_of_rat (\_of x) \ 0 < real_of_rat (\_of x) \ + have evt: "(\x. + 0 < real_of_rat (\_of x) \ 0 < real_of_rat (\_of x) \ 1/real_of_rat (\_of x) \ \ \ 1/real_of_rat (\_of x) \ \ \ - real (n_of x) \ n \ real (m_of x) \ m\ P x) + real (n_of x) \ n \ real (m_of x) \ m\ P x) \ eventually P ?F" (is "(\x. ?prem x \ _) \ _") for \ \ n m P apply (rule eventually_mono[where P="?prem" and Q="P"]) apply (simp add:\_of_def case_prod_beta' \_of_def n_of_def m_of_def) - apply (intro eventually_conj eventually_prod1' eventually_prod2' + apply (intro eventually_conj eventually_prod1' eventually_prod2' sequentially_inf eventually_at_right_less inv_at_right_0_inf) by (auto simp add:prod_filter_eq_bot) have unit_1: "(\_. 1) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" using one_le_power by (intro landau_o.big_mono evt[where \="1"], auto simp add:power_one_over[symmetric]) have unit_2: "(\_. 1) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" by (intro landau_o.big_mono evt[where \="exp 1"]) (auto intro!:iffD2[OF ln_ge_iff] simp add:abs_ge_iff) have unit_3: "(\_. 1) \ O[?F](\x. real (n_of x))" by (intro landau_o.big_mono evt, auto) have unit_4: "(\_. 1) \ O[?F](\x. real (m_of x))" by (intro landau_o.big_mono evt, auto) have unit_5: "(\_. 1) \ O[?F](\x. ln (real (n_of x)))" by (auto intro!: landau_o.big_mono evt[where n="exp 1"]) (metis abs_ge_self linorder_not_le ln_ge_iff not_exp_le_zero order.trans) have unit_6: "(\_. 1) \ O[?F](\x. ln (real (n_of x)) + ln (real (m_of x)))" by (intro landau_sum_1 evt unit_5 iffD2[OF ln_ge_iff], auto) have unit_7: "(\_. 1) \ O[?F](\x. 1 / real_of_rat (\_of x))" by (intro landau_o.big_mono evt[where \="1"], auto) - - have unit_8: "(\_. 1) \ O[?F](g)" + + have unit_8: "(\_. 1) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1 unit_1 unit_2 unit_6) have unit_9: "(\_. 1) \ O[?F](\x. real (n_of x) * real (m_of x))" by (intro landau_o.big_mult_1 unit_3 unit_4) have " (\x. 6 * (1 / (real_of_rat (\_of x))\<^sup>2)) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" by (subst landau_o.big.cmult_in_iff, simp_all) hence l1: "(\x. real (nat \6 / (\_of x)\<^sup>2\)) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" by (intro landau_real_nat landau_rat_ceil[OF unit_1]) (simp_all add:of_rat_divide of_rat_power) have "(\x. - ( ln (real_of_rat (\_of x)))) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" by (intro landau_o.big_mono evt) (subst ln_div, auto) hence l2: "(\x. real (nat \- (18 * ln (real_of_rat (\_of x)))\)) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" by (intro landau_real_nat landau_ceil[OF unit_2], simp) have l3_aux: " (\x. real (m_of x) * (18 + 4 * real (n_of x)) + 1) \ O[?F](\x. real (n_of x) * real (m_of x))" by (rule sum_in_bigo[OF _unit_9], subst mult.commute) (intro landau_o.mult sum_in_bigo, auto simp:unit_3) have "(\x. ln (real (m_of x) * (18 + 4 * real (n_of x)) + 1)) \ O[?F](\x. ln (real (n_of x) * real (m_of x)))" apply (rule landau_ln_2[where a="2"], simp, simp) apply (rule evt[where m="2" and n="1"]) apply (metis dual_order.trans mult_left_mono mult_of_nat_commute of_nat_0_le_iff verit_prod_simplify(1)) using l3_aux by simp also have "(\x. ln (real (n_of x) * real (m_of x))) \ O[?F](\x. ln (real (n_of x)) + ln(real (m_of x)))" by (intro landau_o.big_mono evt[where m="1" and n="1"], auto simp add:ln_mult) finally have l3: "(\x. ln (real (m_of x) * (18 + 4 * real (n_of x)) + 1)) \ O[?F](\x. ln (real (n_of x)) + ln (real (m_of x)))" using landau_o.big_trans by simp have l4: "(\x. ln (8 + 2 * real (n_of x))) \ O[?F](\x. ln (real (n_of x)) + ln (real (m_of x)))" by (intro landau_sum_1 evt[where n="2"] landau_ln_2[where a="2"] iffD2[OF ln_ge_iff]) (auto intro!: sum_in_bigo simp add:unit_3) have l5: "(\x. ln (9 + 2 * real (n_of x))) \ O[?F](\x. ln (real (n_of x)) + ln (real (m_of x)))" by (intro landau_sum_1 evt[where n="2"] landau_ln_2[where a="2"] iffD2[OF ln_ge_iff]) (auto intro!: sum_in_bigo simp add:unit_3) have l6: "(\x. ln (real (nat \6 / (\_of x)\<^sup>2\) + 1)) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1 landau_ln_3 sum_in_bigo unit_6 unit_2 l1 unit_1, simp) have l7: "(\x. ln (9 + 2 * real (n_of x))) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1' unit_1 unit_2 l5) have l8: "(\x. ln (real (nat \- (18 * ln (real_of_rat (\_of x)))\) + 1) ) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1 unit_6 landau_o.big_mult_1' unit_1 landau_ln_3 sum_in_bigo l2 unit_2) simp have l9: "(\x. 5 + 4 * ln (8 + 2 * real (n_of x)) / ln 2 + 2 * ln (real (m_of x) * (18 + 4 * real (n_of x)) + 1) / ln 2) \ O[?F](\x. ln (real (n_of x)) + ln (real (m_of x)))" by (intro sum_in_bigo, auto simp: l3 l4 unit_6) - have l10: "(\x. real (nat \6 / (\_of x)\<^sup>2\) * real (nat \- (18 * ln (real_of_rat (\_of x)))\) * + have l10: "(\x. real (nat \6 / (\_of x)\<^sup>2\) * real (nat \- (18 * ln (real_of_rat (\_of x)))\) * (5 + 4 * ln (8 + 2 * real (n_of x)) / ln 2 + 2 * ln(real (m_of x) * (18 + 4 * real (n_of x)) + 1) / ln 2)) \ O[?F](g)" unfolding g_def by (intro landau_o.mult, auto simp: l1 l2 l9) have "f2_space_usage = (\x. f2_space_usage (n_of x, m_of x, \_of x, \_of x))" by (simp add:case_prod_beta' n_of_def \_of_def \_of_def m_of_def) also have "... \ O[?F](g)" by (auto intro!:sum_in_bigo simp:Let_def log_def l6 l7 l8 l10 unit_8) also have "... = O[?F](?rhs)" by (simp add:case_prod_beta' g_def n_of_def \_of_def \_of_def m_of_def) finally show ?thesis by simp qed end diff --git a/thys/Frequency_Moments/Frequency_Moment_k.thy b/thys/Frequency_Moments/Frequency_Moment_k.thy --- a/thys/Frequency_Moments/Frequency_Moment_k.thy +++ b/thys/Frequency_Moments/Frequency_Moment_k.thy @@ -1,1015 +1,1016 @@ section \Frequency Moment $k$\ theory Frequency_Moment_k - imports + imports Frequency_Moments - Landau_Ext + Landau_Ext Lp.Lp Median_Method.Median - Product_PMF_Ext + Probability_Ext + Product_PMF_Ext begin text \This section contains a formalization of the algorithm for the $k$-th frequency moment. It is based on the algorithm described in \<^cite>\\\textsection 2.1\ in "alon1999"\.\ type_synonym fk_state = "nat \ nat \ nat \ nat \ (nat \ nat \ (nat \ nat))" fun fk_init :: "nat \ rat \ rat \ nat \ fk_state pmf" where "fk_init k \ \ n = do { let s\<^sub>1 = nat \3 * real k * n powr (1-1/real k) / (real_of_rat \)\<^sup>2\; let s\<^sub>2 = nat \-18 * ln (real_of_rat \)\; return_pmf (s\<^sub>1, s\<^sub>2, k, 0, (\_ \ {0..1} \ {0..2}. (0,0))) }" fun fk_update :: "nat \ fk_state \ fk_state pmf" where - "fk_update a (s\<^sub>1, s\<^sub>2, k, m, r) = + "fk_update a (s\<^sub>1, s\<^sub>2, k, m, r) = do { coins \ prod_pmf ({0..1} \ {0..2}) (\_. bernoulli_pmf (1/(real m+1))); - return_pmf (s\<^sub>1, s\<^sub>2, k, m+1, \i \ {0..1} \ {0..2}. - if coins i then - (a,0) + return_pmf (s\<^sub>1, s\<^sub>2, k, m+1, \i \ {0..1} \ {0..2}. + if coins i then + (a,0) else ( let (x,l) = r i in (x, l + of_bool (x=a)) ) ) }" fun fk_result :: "fk_state \ rat pmf" where - "fk_result (s\<^sub>1, s\<^sub>2, k, m, r) = + "fk_result (s\<^sub>1, s\<^sub>2, k, m, r) = return_pmf (median s\<^sub>2 (\i\<^sub>2 \ {0..2}. (\i\<^sub>1\{0..1}. rat_of_nat (let t = snd (r (i\<^sub>1, i\<^sub>2)) + 1 in m * (t^k - (t - 1)^k))) / (rat_of_nat s\<^sub>1)) )" lemma bernoulli_pmf_1: "bernoulli_pmf 1 = return_pmf True" by (rule pmf_eqI, simp add:indicator_def) fun fk_space_usage :: "(nat \ nat \ nat \ rat \ rat) \ real" where "fk_space_usage (k, n, m, \, \) = ( let s\<^sub>1 = nat \3*real k* (real n) powr (1-1/ real k) / (real_of_rat \)\<^sup>2 \ in - let s\<^sub>2 = nat \-(18 * ln (real_of_rat \))\ in + let s\<^sub>2 = nat \-(18 * ln (real_of_rat \))\ in 4 + 2 * log 2 (s\<^sub>1 + 1) + 2 * log 2 (s\<^sub>2 + 1) + 2 * log 2 (real k + 1) + 2 * log 2 (real m + 1) + s\<^sub>1 * s\<^sub>2 * (2 + 2 * log 2 (real n+1) + 2 * log 2 (real m+1)))" definition encode_fk_state :: "fk_state \ bool list option" where - "encode_fk_state = - N\<^sub>e \\<^sub>e (\s\<^sub>1. - N\<^sub>e \\<^sub>e (\s\<^sub>2. - N\<^sub>e \\<^sub>e - N\<^sub>e \\<^sub>e + "encode_fk_state = + N\<^sub>e \\<^sub>e (\s\<^sub>1. + N\<^sub>e \\<^sub>e (\s\<^sub>2. + N\<^sub>e \\<^sub>e + N\<^sub>e \\<^sub>e (List.product [0..1] [0..2] \\<^sub>e (N\<^sub>e \\<^sub>e N\<^sub>e))))" lemma "inj_on encode_fk_state (dom encode_fk_state)" proof - have "is_encoding encode_fk_state" by (simp add:encode_fk_state_def) (intro dependent_encoding exp_golomb_encoding fun_encoding) thus ?thesis by (rule encoding_imp_inj) qed text \This is an intermediate non-parallel form @{term "fk_update"} used only in the correctness proof.\ fun fk_update_2 :: "'a \ (nat \ 'a \ nat) \ (nat \ 'a \ nat) pmf" where - "fk_update_2 a (m,x,l) = + "fk_update_2 a (m,x,l) = do { coin \ bernoulli_pmf (1/(real m+1)); return_pmf (m+1,if coin then (a,0) else (x, l + of_bool (x=a))) }" definition sketch where "sketch as i = (as ! i, count_list (drop (i+1) as) (as ! i))" lemma fk_update_2_distr: assumes "as \ []" shows "fold (\x s. s \ fk_update_2 x) as (return_pmf (0,0,0)) = pmf_of_set {.. (\k. return_pmf (length as, sketch as k))" using assms proof (induction as rule:rev_nonempty_induct) case (single x) - show ?case using single - by (simp add:bind_return_pmf pmf_of_set_singleton bernoulli_pmf_1 lessThan_def sketch_def) + show ?case using single + by (simp add:bind_return_pmf pmf_of_set_singleton bernoulli_pmf_1 lessThan_def sketch_def) next case (snoc x xs) let ?h = "(\xs k. count_list (drop (Suc k) xs) (xs ! k))" let ?q = "(\xs k. (length xs, sketch xs k))" have non_empty: " {.. {}" " {.. {}" using snoc by auto - have fk_update_2_eta:"fk_update_2 x = (\a. fk_update_2 x (fst a, fst (snd a), snd (snd a)))" + have fk_update_2_eta:"fk_update_2 x = (\a. fk_update_2 x (fst a, fst (snd a), snd (snd a)))" by auto have "pmf_of_set {.. (\k. bernoulli_pmf (1 / (real (length xs) + 1)) \ - (\coin. return_pmf (if coin then length xs else k))) = + (\coin. return_pmf (if coin then length xs else k))) = bernoulli_pmf (1 / (real (length xs) + 1)) \ (\y. pmf_of_set {.. (\k. return_pmf (if y then length xs else k)))" by (subst bind_commute_pmf, simp) also have "... = pmf_of_set {.. (\k. bernoulli_pmf (1 / (real (length xs) + 1)) \ (\coin. return_pmf (if coin then length xs else k))) = pmf_of_set {..x s. (s \ fk_update_2 x)) (xs@[x]) (return_pmf (0,0,0)) = (pmf_of_set {.. (\k. return_pmf (length xs, sketch xs k))) \ fk_update_2 x" using snoc by (simp add:case_prod_beta') - also have "... = (pmf_of_set {.. (\k. return_pmf (length xs, sketch xs k))) \ - (\(m,a,l). bernoulli_pmf (1 / (real m + 1)) \ (\coin. + also have "... = (pmf_of_set {.. (\k. return_pmf (length xs, sketch xs k))) \ + (\(m,a,l). bernoulli_pmf (1 / (real m + 1)) \ (\coin. return_pmf (m + 1, if coin then (x, 0) else (a, (l + of_bool (a = x))))))" by (subst fk_update_2_eta, subst fk_update_2.simps, simp add:case_prod_beta') also have "... = pmf_of_set {.. (\k. bernoulli_pmf (1 / (real (length xs) + 1)) \ (\coin. return_pmf (length xs + 1, if coin then (x, 0) else (xs ! k, ?h xs k + of_bool (xs ! k = x)))))" by (subst bind_assoc_pmf, simp add: bind_return_pmf sketch_def) also have "... = pmf_of_set {.. (\k. bernoulli_pmf (1 / (real (length xs) + 1)) \ (\coin. return_pmf (if coin then length xs else k) \ (\k'. return_pmf (?q (xs@[x]) k'))))" using non_empty by (intro bind_pmf_cong, auto simp add:bind_return_pmf nth_append count_list_append sketch_def) also have "... = pmf_of_set {.. (\k. bernoulli_pmf (1 / (real (length xs) + 1)) \ (\coin. return_pmf (if coin then length xs else k))) \ (\k'. return_pmf (?q (xs@[x]) k'))" by (subst bind_assoc_pmf, subst bind_assoc_pmf, simp) also have "... = pmf_of_set {.. (\k'. return_pmf (?q (xs@[x]) k'))" by (subst b, simp) finally show ?case by simp qed context fixes \ \ :: rat fixes n k :: nat fixes as assumes k_ge_1: "k \ 1" assumes \_range: "\ \ {0<..<1}" assumes \_range: "\ > 0" assumes as_range: "set as \ {..1 where "s\<^sub>1 = nat \3 * real k * (real n) powr (1-1/real k) / (real_of_rat \)\<^sup>2\" definition s\<^sub>2 where "s\<^sub>2 = nat \-(18 * ln (real_of_rat \))\" definition "M\<^sub>1 = {(u, v). v < count_list as u}" definition "\\<^sub>1 = measure_pmf (pmf_of_set M\<^sub>1)" definition "M\<^sub>2 = prod_pmf ({0..1} \ {0..2}) (\_. pmf_of_set M\<^sub>1)" definition "\\<^sub>2 = measure_pmf M\<^sub>2" interpretation prob_space "\\<^sub>1" unfolding \\<^sub>1_def by (simp add:prob_space_measure_pmf) interpretation \\<^sub>2:prob_space "\\<^sub>2" unfolding \\<^sub>2_def by (simp add:prob_space_measure_pmf) lemma split_space: "(\a\M\<^sub>1. f (snd a)) = (\u \ set as. (\v \{0..u. {u} \ {v. v < count_list as u})" - have a: "inj_on snd (A x)" for x - by (simp add:A_def inj_on_def) + have a: "inj_on snd (A x)" for x + by (simp add:A_def inj_on_def) have "\u v. u < count_list as v \ v \ set as" by (subst count_list_gr_1, force) hence "M\<^sub>1 = \ (A ` set as)" by (auto simp add:set_eq_iff A_def M\<^sub>1_def) hence "(\a\M\<^sub>1. f (snd a)) = sum (f \ snd) (\ (A ` set as))" by (intro sum.cong, auto) also have "... = sum (\x. sum (f \ snd) (A x)) (set as)" - by (rule sum.UNION_disjoint, simp, simp add:A_def, simp add:A_def, blast) + by (rule sum.UNION_disjoint, simp, simp add:A_def, simp add:A_def, blast) also have "... = sum (\x. sum f (snd ` A x)) (set as)" by (intro sum.cong, auto simp add:sum.reindex[OF a]) also have "... = (\u \ set as. (\v \{0.. []" - shows fin_space: "finite M\<^sub>1" + shows fin_space: "finite M\<^sub>1" and non_empty_space: "M\<^sub>1 \ {}" and card_space: "card M\<^sub>1 = length as" proof - have "M\<^sub>1 \ set as \ {k. k < length as}" proof (rule subsetI) fix x assume a:"x \ M\<^sub>1" have "fst x \ set as" using a by (simp add:case_prod_beta count_list_gr_1 M\<^sub>1_def) moreover have "snd x < length as" using a count_le_length order_less_le_trans by (simp add:case_prod_beta M\<^sub>1_def) fast ultimately show "x \ set as \ {k. k < length as}" by (simp add:mem_Times_iff) qed thus fin_space: "finite M\<^sub>1" using finite_subset by blast - have "(as ! 0, 0) \ M\<^sub>1" + have "(as ! 0, 0) \ M\<^sub>1" using assms(1) unfolding M\<^sub>1_def by (simp, metis count_list_gr_1 gr0I length_greater_0_conv not_one_le_zero nth_mem) thus "M\<^sub>1 \ {}" by blast show "card M\<^sub>1 = length as" using fin_space split_space[where f="\_. (1::nat)"] by (simp add:sum_count_set[where X="set as" and xs="as", simplified]) qed lemma assumes "as \ []" shows integrable_1: "integrable \\<^sub>1 (f :: _ \ real)" and integrable_2: "integrable \\<^sub>2 (g :: _ \ real)" proof - have fin_omega: "finite (set_pmf (pmf_of_set M\<^sub>1))" using fin_space[OF assms] non_empty_space[OF assms] by auto thus "integrable \\<^sub>1 f" unfolding \\<^sub>1_def by (rule integrable_measure_pmf_finite) have "finite (set_pmf M\<^sub>2)" unfolding M\<^sub>2_def using fin_omega by (subst set_prod_pmf) (auto intro:finite_PiE) thus "integrable \\<^sub>2 g" unfolding \\<^sub>2_def by (intro integrable_measure_pmf_finite) qed lemma sketch_distr: assumes "as \ []" shows "pmf_of_set {.. (\k. return_pmf (sketch as k)) = pmf_of_set M\<^sub>1" proof - - have "x < y \ y < length as \ + have "x < y \ y < length as \ count_list (drop (y+1) as) (as ! y) < count_list (drop (x+1) as) (as ! y)" for x y by (intro count_list_lt_suffix suffix_drop_drop, simp_all) (metis Suc_diff_Suc diff_Suc_Suc diff_add_inverse lessI less_natE) hence a1: "inj_on (sketch as) {k. k < length as}" unfolding sketch_def by (intro inj_onI) (metis Pair_inject mem_Collect_eq nat_neq_iff) have "x < length as \ count_list (drop (x+1) as) (as ! x) < count_list as (as ! x)" for x by (rule count_list_lt_suffix, auto simp add:suffix_drop) hence "sketch as ` {k. k < length as} \ M\<^sub>1" by (intro image_subsetI, simp add:sketch_def M\<^sub>1_def) moreover have "card M\<^sub>1 \ card (sketch as ` {k. k < length as})" by (simp add: card_space[OF assms(1)] card_image[OF a1]) ultimately have "sketch as ` {k. k < length as} = M\<^sub>1" using fin_space[OF assms(1)] by (intro card_seteq, simp_all) hence "bij_betw (sketch as) {k. k < length as} M\<^sub>1" using a1 by (simp add:bij_betw_def) hence "map_pmf (sketch as) (pmf_of_set {k. k < length as}) = pmf_of_set M\<^sub>1" using assms by (intro map_pmf_of_set_bij_betw, auto) thus ?thesis by (simp add: sketch_def map_pmf_def lessThan_def) qed lemma fk_update_distr: - "fold (\x s. s \ fk_update x) as (fk_init k \ \ n) = - prod_pmf ({0..1} \ {0..2}) (\_. fold (\x s. s \ fk_update_2 x) as (return_pmf (0,0,0))) + "fold (\x s. s \ fk_update x) as (fk_init k \ \ n) = + prod_pmf ({0..1} \ {0..2}) (\_. fold (\x s. s \ fk_update_2 x) as (return_pmf (0,0,0))) \ (\x. return_pmf (s\<^sub>1,s\<^sub>2,k, length as, \i\{0..1}\{0..2}. snd (x i)))" proof (induction as rule:rev_induct) case Nil - then show ?case + then show ?case by (auto simp:Let_def s\<^sub>1_def[symmetric] s\<^sub>2_def[symmetric] bind_return_pmf) next case (snoc x xs) - have fk_update_2_eta:"fk_update_2 x = (\a. fk_update_2 x (fst a, fst (snd a), snd (snd a)))" + have fk_update_2_eta:"fk_update_2 x = (\a. fk_update_2 x (fst a, fst (snd a), snd (snd a)))" by auto have a: "fk_update x (s\<^sub>1, s\<^sub>2, k, length xs, \i\{0..1} \ {0..2}. snd (f i)) = - prod_pmf ({0..1} \ {0..2}) (\i. fk_update_2 x (f i)) \ + prod_pmf ({0..1} \ {0..2}) (\i. fk_update_2 x (f i)) \ (\a. return_pmf (s\<^sub>1,s\<^sub>2, k, Suc (length xs), \i\{0..1} \ {0..2}. snd (a i)))" - if b: "f \ set_pmf (prod_pmf ({0..1} \ {0..2}) + if b: "f \ set_pmf (prod_pmf ({0..1} \ {0..2}) (\_. fold (\a s. s \ fk_update_2 a) xs (return_pmf (0, 0, 0))))" for f proof - have c:"fst (f i) = length xs" if d:"i \ {0..1} \{0..2}" for i proof (cases "xs = []") case True then show ?thesis using b d by (simp add: set_Pi_pmf) - next + next case False hence "{.. {}" by force - thus ?thesis using b d + thus ?thesis using b d by (simp add:set_Pi_pmf fk_update_2_distr[OF False] PiE_dflt_def) force qed show ?thesis apply (subst fk_update_2_eta, subst fk_update_2.simps, simp) apply (simp add: Pi_pmf_bind_return[where d'="undefined"] bind_assoc_pmf) apply (rule bind_pmf_cong, simp add:c cong:Pi_pmf_cong) by (auto simp add:bind_return_pmf case_prod_beta) qed - - have "fold (\x s. s \ fk_update x) (xs @ [x]) (fk_init k \ \ n) = - prod_pmf ({0..1} \ {0..2}) (\_. fold (\x s. s \ fk_update_2 x) xs (return_pmf (0,0,0))) + + have "fold (\x s. s \ fk_update x) (xs @ [x]) (fk_init k \ \ n) = + prod_pmf ({0..1} \ {0..2}) (\_. fold (\x s. s \ fk_update_2 x) xs (return_pmf (0,0,0))) \ (\\. return_pmf (s\<^sub>1,s\<^sub>2,k, length xs, \i\{0..1}\{0..2}. snd (\ i)) \ fk_update x)" using snoc by (simp add:restrict_def bind_assoc_pmf del:fk_init.simps) - also have "... = prod_pmf ({0..1} \ {0..2}) + also have "... = prod_pmf ({0..1} \ {0..2}) (\_. fold (\a s. s \ fk_update_2 a) xs (return_pmf (0, 0, 0))) \ (\f. prod_pmf ({0..1} \ {0..2}) (\i. fk_update_2 x (f i)) \ (\a. return_pmf (s\<^sub>1, s\<^sub>2, k, Suc (length xs), \i\{0..1} \ {0..2}. snd (a i))))" using a by (intro bind_pmf_cong, simp_all add:bind_return_pmf del:fk_update.simps) - also have "... = prod_pmf ({0..1} \ {0..2}) + also have "... = prod_pmf ({0..1} \ {0..2}) (\_. fold (\a s. s \ fk_update_2 a) xs (return_pmf (0, 0, 0))) \ (\f. prod_pmf ({0..1} \ {0..2}) (\i. fk_update_2 x (f i))) \ (\a. return_pmf (s\<^sub>1, s\<^sub>2, k, Suc (length xs), \i\{0..1} \ {0..2}. snd (a i)))" by (simp add:bind_assoc_pmf) - also have "... = (prod_pmf ({0..1} \ {0..2}) - (\_. fold (\a s. s \ fk_update_2 a) (xs@[x]) (return_pmf (0,0,0))) + also have "... = (prod_pmf ({0..1} \ {0..2}) + (\_. fold (\a s. s \ fk_update_2 a) (xs@[x]) (return_pmf (0,0,0))) \ (\a. return_pmf (s\<^sub>1,s\<^sub>2,k, length (xs@[x]), \i\{0..1}\{0..2}. snd (a i))))" by (simp, subst Pi_pmf_bind, auto) finally show ?case by blast qed lemma power_diff_sum: fixes a b :: "'a :: {comm_ring_1,power}" assumes "k > 0" shows "a^k - b^k = (a-b) * (\i = 0.. insert m {Suc m..i. a * (a^i * b^(k-1-i))) {0..i. a * (a^i * b^(k-1-i))) {0..i. b * (a^i * b^(k-1-i))) {0..i. (a^i * b^(k-i))) \ (\i. i+1)) {0..i. (a^i * b^(k-i))) \ (\i. i+1)) {0..i. (a^i * (b^(1+(k-1-i))))) {0..i. (a^i * b^(k-i))) \ (\i. i+1)) {0..i. (a^i * b^(k-i))) \ (\i. i+1)) {0..i. (a^i * b^(k-i))) {0..i. (a^i * b^(k-i))) (insert k {1..i. (a^i * b^(k-i))) (insert k {1..i. (a^i * b^(k-i))) (insert 0 {Suc 0.. 0" assumes "(a :: real) \ b" assumes "b \ 0" shows "a^k -b^k \ (a-b) * k * a^(k-1)" proof - have " \i. i < k \ a ^ i * b ^ (k - 1 - i) \ a ^ i * a ^ (k-1-i)" using assms by (intro mult_left_mono power_mono) auto also have "\i. i < k \ a ^ i * a ^ (k - 1 - i) = a ^ (k - Suc 0)" using assms(1) by (subst power_add[symmetric], simp) finally have a: "\i. i < k \ a ^ i * b ^ (k - 1 - i) \ a ^ (k - Suc 0)" by blast have "a^k - b^k = (a-b) * (\i = 0.. (a-b) * (\i = 0..Specialization of the Hoelder inquality for sums.\ lemma Holder_inequality_sum: assumes "p > (0::real)" "q > 0" "1/p + 1/q = 1" assumes "finite A" shows "\\x\A. f x * g x\ \ (\x\A. \f x\ powr p) powr (1/p) * (\x\A. \g x\ powr q) powr (1/q)" proof - - have "\LINT x|count_space A. f x * g x\ \ - (LINT x|count_space A. \f x\ powr p) powr (1 / p) * + have "\LINT x|count_space A. f x * g x\ \ + (LINT x|count_space A. \f x\ powr p) powr (1 / p) * (LINT x|count_space A. \g x\ powr q) powr (1 / q)" using assms integrable_count_space by (intro Lp.Holder_inequality, auto) thus ?thesis using assms by (simp add: lebesgue_integral_count_space_finite[symmetric]) qed lemma real_count_list_pos: assumes "x \ set as" shows "real (count_list as x) > 0" using count_list_gr_1 assms by force lemma fk_estimate: assumes "as \ []" shows "length as * of_rat (F (2*k-1) as) \ n powr (1 - 1 / real k) * (of_rat (F k as))^2" (is "?lhs \ ?rhs") proof (cases "k \ 2") case True - define M where "M = Max (count_list as ` set as)" + define M where "M = Max (count_list as ` set as)" have "M \ count_list as ` set as" unfolding M_def using assms by (intro Max_in, auto) then obtain m where m_in: "m \ set as" and m_def: "M = count_list as m" by blast have a: "real M > 0" using m_in count_list_gr_1 by (simp add:m_def, force) have b: "2*k-1 = (k-1) + k" by simp - have " 0 < real (count_list as m)" + have " 0 < real (count_list as m)" using m_in count_list_gr_1 by force hence "M powr k = real (count_list as m) ^ k" by (simp add: powr_realpow m_def) also have "... \ (\x\set as. real (count_list as x) ^ k)" using m_in by (intro member_le_sum, simp_all) also have "... \ real_of_rat (F k as)" by (simp add:F_def of_rat_sum of_rat_power) finally have d: "M powr k \ real_of_rat (F k as)" by simp - have e: "0 \ real_of_rat (F k as)" + have e: "0 \ real_of_rat (F k as)" using F_gr_0[OF assms(1)] by (simp add: order_le_less) have "real (k - 1) / real k + 1 = real (k - 1) / real k + real k / real k" using assms True by simp also have "... = real (2 * k - 1) / real k" using b by (subst add_divide_distrib[symmetric], force) finally have f: "real (k - 1) / real k + 1 = real (2 * k - 1) / real k" by blast - have "real_of_rat (F (2*k-1) as) = - (\x\set as. real (count_list as x) ^ (k - 1) * real (count_list as x) ^ k)" + have "real_of_rat (F (2*k-1) as) = + (\x\set as. real (count_list as x) ^ (k - 1) * real (count_list as x) ^ k)" using b by (simp add:F_def of_rat_sum sum_distrib_left of_rat_mult power_add of_rat_power) also have "... \ (\x\set as. real M ^ (k - 1) * real (count_list as x) ^ k)" by (intro sum_mono mult_right_mono power_mono of_nat_mono) (auto simp:M_def) also have "... = M powr (k-1) * of_rat (F k as)" using a by (simp add:sum_distrib_left F_def of_rat_mult of_rat_sum of_rat_power powr_realpow) also have "... = (M powr k) powr (real (k - 1) / real k) * of_rat (F k as) powr 1" using e by (simp add:powr_powr) also have "... \ (real_of_rat (F k as)) powr ((k-1)/k) * (real_of_rat (F k as) powr 1)" using d by (intro mult_right_mono powr_mono2, auto) also have "... = (real_of_rat (F k as)) powr ((2*k-1) / k)" by (subst powr_add[symmetric], subst f, simp) finally have a: "real_of_rat (F (2*k-1) as) \ (real_of_rat (F k as)) powr ((2*k-1) / k)" by blast have g: "card (set as) \ n" using card_mono[OF _ as_range] by simp have "length as = abs (sum (\x. real (count_list as x)) (set as))" by (subst of_nat_sum[symmetric], simp add: sum_count_set) - also have "... \ card (set as) powr ((k-Suc 0)/k) * + also have "... \ card (set as) powr ((k-Suc 0)/k) * (sum (\x. \real (count_list as x)\ powr k) (set as)) powr (1/k)" using assms True by (intro Holder_inequality_sum[where p="k/(k-1)" and q="k" and f="\_.1", simplified]) (auto simp add:algebra_simps add_divide_distrib[symmetric]) also have "... = (card (set as)) powr ((k-1) / real k) * of_rat (F k as) powr (1/ k)" using real_count_list_pos by (simp add:F_def of_rat_sum of_rat_power powr_realpow) also have "... = (card (set as)) powr (1 - 1 / real k) * of_rat (F k as) powr (1/ k)" using k_ge_1 by (subst of_nat_diff[OF k_ge_1], subst diff_divide_distrib, simp) also have "... \ n powr (1 - 1 / real k) * of_rat (F k as) powr (1/ k)" using k_ge_1 g by (intro mult_right_mono powr_mono2, auto) finally have h: "length as \ n powr (1 - 1 / real k) * of_rat (F k as) powr (1/real k)" by blast have i:"1 / real k + real (2 * k - 1) / real k = real 2" using True by (subst add_divide_distrib[symmetric], simp_all add:of_nat_diff) have "?lhs \ n powr (1 - 1/k) * of_rat (F k as) powr (1/k) * (of_rat (F k as)) powr ((2*k-1) / k)" using a h F_ge_0 by (intro mult_mono mult_nonneg_nonneg, auto) also have "... = ?rhs" using i F_gr_0[OF assms] by (simp add:powr_add[symmetric] powr_realpow[symmetric]) finally show ?thesis by blast next case False have "n = 0 \ False" using as_range assms by auto - hence "n > 0" + hence "n > 0" by auto moreover have "k = 1" using assms k_ge_1 False by linarith moreover have "length as = real_of_rat (F (Suc 0) as)" by (simp add:F_def sum_count_set of_nat_sum[symmetric] del:of_nat_sum) ultimately show ?thesis by (simp add:power2_eq_square) qed definition result where "result a = of_nat (length as) * of_nat (Suc (snd a) ^ k - snd a ^ k)" lemma result_exp_1: assumes "as \ []" shows "expectation result = real_of_rat (F k as)" proof - have "expectation result = (\a\M\<^sub>1. result a * pmf (pmf_of_set M\<^sub>1) a)" unfolding \\<^sub>1_def using non_empty_space assms fin_space by (subst integral_measure_pmf_real) auto also have "... = (\a\M\<^sub>1. result a / real (length as))" using non_empty_space assms fin_space card_space by simp also have "... = (\a\M\<^sub>1. real (Suc (snd a) ^ k - snd a ^ k))" using assms by (simp add:result_def) also have "... = (\u\set as. \v = 0..u\set as. real (count_list as u)^k)" using k_ge_1 by (subst sum_Suc_diff') (auto simp add:zero_power) also have "... = of_rat (F k as)" by (simp add:F_def of_rat_sum of_rat_power) finally show ?thesis by simp qed lemma result_var_1: assumes "as \ []" shows "variance result \ (of_rat (F k as))\<^sup>2 * k * n powr (1 - 1 / real k)" -proof - +proof - have k_gt_0: "k > 0" using k_ge_1 by linarith have c:"real (Suc v ^ k) - real (v ^ k) \ k * real (count_list as a) ^ (k - Suc 0)" if c_1: "v < count_list as a" for a v proof - have "real (Suc v ^ k) - real (v ^ k) \ (real (v+1) - real v) * k * (1 + real v) ^ (k - Suc 0)" using k_gt_0 power_diff_est[where a="Suc v" and b="v"] by simp moreover have "(real (v+1) - real v) = 1" by auto ultimately have "real (Suc v ^ k) - real (v ^ k) \ k * (1 + real v) ^ (k - Suc 0)" by auto also have "... \ k * real (count_list as a) ^ (k- Suc 0)" using c_1 by (intro mult_left_mono power_mono, auto) finally show ?thesis by blast qed - + have "length as * (\a\ M\<^sub>1. (real (Suc (snd a) ^ k - (snd a) ^ k))\<^sup>2) = - length as * (\a\ set as. (\v \ {0..a\ set as. (\v \ {0.. length as * (\a\ set as. (\v \ {0.. length as * (\a\ set as. (\v \ {0..a\ set as. real (count_list as a) ^ (k-1) * + also have "... = length as * k * (\a\ set as. real (count_list as a) ^ (k-1) * (\v \ {0..a\ set as. real (count_list as a ^ (2*k-1)))" using assms k_ge_1 by (subst sum_Suc_diff', auto simp: zero_power[OF k_gt_0] mult_2 power_add[symmetric]) also have "... = k * (length as * of_rat (F (2*k-1) as))" by (simp add:sum_distrib_left[symmetric] F_def of_rat_sum of_rat_power) also have "... \ k * (of_rat (F k as)^2 * n powr (1 - 1 / real k))" using fk_estimate[OF assms] by (intro mult_left_mono) (auto simp: mult.commute) - finally have b: "real (length as) * (\a\ M\<^sub>1. (real (Suc (snd a) ^ k - (snd a) ^ k))\<^sup>2) \ + finally have b: "real (length as) * (\a\ M\<^sub>1. (real (Suc (snd a) ^ k - (snd a) ^ k))\<^sup>2) \ k * ((of_rat (F k as))\<^sup>2 * n powr (1 - 1 / real k))" by blast have "expectation (\\. (result \ :: real)^2) - (expectation result)^2 \ expectation (\\. result \^2)" by simp also have "... = (\a\M\<^sub>1. (length as * real (Suc (snd a) ^ k - snd a ^ k))\<^sup>2 * pmf (pmf_of_set M\<^sub>1) a)" using fin_space non_empty_space assms unfolding \\<^sub>1_def result_def by (subst integral_measure_pmf_real[where A="M\<^sub>1"], auto) also have "... = (\a\M\<^sub>1. length as * (real (Suc (snd a) ^ k - snd a ^ k))\<^sup>2)" using assms non_empty_space fin_space by (subst pmf_of_set) (simp_all add:card_space power_mult_distrib power2_eq_square ac_simps) also have "... \ k * ((of_rat (F k as))\<^sup>2 * n powr (1 - 1 / real k))" using b by (simp add:sum_distrib_left[symmetric]) also have "... = of_rat (F k as)^2 * k * n powr (1 - 1 / real k)" by (simp add:ac_simps) - finally have "expectation (\\. result \^2) - (expectation result)^2 \ + finally have "expectation (\\. result \^2) - (expectation result)^2 \ of_rat (F k as)^2 * k * n powr (1 - 1 / real k)" by blast thus ?thesis using integrable_1[OF assms] by (simp add:variance_eq) qed theorem fk_alg_sketch: assumes "as \ []" - shows "fold (\a state. state \ fk_update a) as (fk_init k \ \ n) = + shows "fold (\a state. state \ fk_update a) as (fk_init k \ \ n) = map_pmf (\x. (s\<^sub>1,s\<^sub>2,k,length as, x)) M\<^sub>2" (is "?lhs = ?rhs") proof - - have "?lhs = prod_pmf ({0..1} \ {0..2}) + have "?lhs = prod_pmf ({0..1} \ {0..2}) (\_. fold (\x s. s \ fk_update_2 x) as (return_pmf (0, 0, 0))) \ (\x. return_pmf (s\<^sub>1, s\<^sub>2, k, length as, \i\{0..1} \ {0..2}. snd (x i)))" by (subst fk_update_distr, simp) - also have "... = prod_pmf ({0..1} \ {0..2}) (\_. pmf_of_set {.. + also have "... = prod_pmf ({0..1} \ {0..2}) (\_. pmf_of_set {.. (\k. return_pmf (length as, sketch as k))) \ (\x. return_pmf (s\<^sub>1, s\<^sub>2, k, length as, \i\{0..1} \ {0..2}. snd (x i)))" by (subst fk_update_2_distr[OF assms], simp) - also have "... = prod_pmf ({0..1} \ {0..2}) (\_. pmf_of_set {.. + also have "... = prod_pmf ({0..1} \ {0..2}) (\_. pmf_of_set {.. (\k. return_pmf (sketch as k)) \ (\s. return_pmf (length as, s))) \ (\x. return_pmf (s\<^sub>1, s\<^sub>2, k, length as, \i\{0..1} \ {0..2}. snd (x i)))" by (subst bind_assoc_pmf, subst bind_return_pmf, simp) - also have "... = prod_pmf ({0..1} \ {0..2}) (\_. pmf_of_set {.. + also have "... = prod_pmf ({0..1} \ {0..2}) (\_. pmf_of_set {.. (\k. return_pmf (sketch as k))) \ (\x. return_pmf (\i \ {0..1} \ {0..2}. (length as, x i))) \ (\x. return_pmf (s\<^sub>1, s\<^sub>2, k, length as, \i\{0..1} \ {0..2}. snd (x i)))" by (subst Pi_pmf_bind_return[where d'="undefined"], simp, simp add:restrict_def) - also have "... = prod_pmf ({0..1} \ {0..2}) (\_. pmf_of_set {.. + also have "... = prod_pmf ({0..1} \ {0..2}) (\_. pmf_of_set {.. (\k. return_pmf (sketch as k))) \ (\x. return_pmf (s\<^sub>1, s\<^sub>2, k, length as, restrict x ({0..1} \ {0..2})))" by (subst bind_assoc_pmf, simp add:bind_return_pmf cong:restrict_cong) also have "... = M\<^sub>2 \ (\x. return_pmf (s\<^sub>1, s\<^sub>2, k, length as, restrict x ({0..1} \ {0..2})))" by (subst sketch_distr[OF assms], simp add:M\<^sub>2_def) also have "... = M\<^sub>2 \ (\x. return_pmf (s\<^sub>1, s\<^sub>2, k, length as, x))" - by (rule bind_pmf_cong, auto simp add:PiE_dflt_def M\<^sub>2_def set_Pi_pmf) + by (rule bind_pmf_cong, auto simp add:PiE_dflt_def M\<^sub>2_def set_Pi_pmf) also have "... = ?rhs" by (simp add:map_pmf_def) finally show ?thesis by simp qed -definition mean_rv +definition mean_rv where "mean_rv \ i\<^sub>2 = (\i\<^sub>1 = 0..1. result (\ (i\<^sub>1, i\<^sub>2))) / of_nat s\<^sub>1" -definition median_rv +definition median_rv where "median_rv \ = median s\<^sub>2 (\i\<^sub>2. mean_rv \ i\<^sub>2)" lemma fk_alg_correct': defines "M \ fold (\a state. state \ fk_update a) as (fk_init k \ \ n) \ fk_result" shows "\

(\ in measure_pmf M. \\ - F k as\ \ \ * F k as) \ 1 - of_rat \" proof (cases "as = []") case True - have a: "nat \- (18 * ln (real_of_rat \))\ > 0" using \_range by simp - show ?thesis using True \_range + have a: "nat \- (18 * ln (real_of_rat \))\ > 0" using \_range by simp + show ?thesis using True \_range by (simp add:F_def M_def bind_return_pmf median_const[OF a] Let_def) next case False have "set as \ {}" using assms False by blast hence n_nonzero: "n > 0" using as_range by fastforce have fk_nonzero: "F k as > 0" using F_gr_0[OF False] by simp have s1_nonzero: "s\<^sub>1 > 0" using \_range k_ge_1 n_nonzero by (simp add:s\<^sub>1_def) have s2_nonzero: "s\<^sub>2 > 0" - using \_range by (simp add:s\<^sub>2_def) + using \_range by (simp add:s\<^sub>2_def) have real_of_rat_mean_rv: "\x i. mean_rv x = (\i. real_of_rat (mean_rv x i))" by (rule ext, simp add:of_rat_divide of_rat_sum of_rat_mult result_def mean_rv_def) have real_of_rat_median_rv: "\x. median_rv x = real_of_rat (median_rv x)" unfolding median_rv_def using s2_nonzero by (subst real_of_rat_mean_rv, simp add: median_rat median_restrict) have space_\\<^sub>2: "space \\<^sub>2 = UNIV" by (simp add:\\<^sub>2_def) - have fk_result_eta: "fk_result = (\(x,y,z,u,v). fk_result (x,y,z,u,v))" - by auto + have fk_result_eta: "fk_result = (\(x,y,z,u,v). fk_result (x,y,z,u,v))" + by auto - have a:"fold (\x state. state \ fk_update x) as (fk_init k \ \ n) = + have a:"fold (\x state. state \ fk_update x) as (fk_init k \ \ n) = map_pmf (\x. (s\<^sub>1,s\<^sub>2,k,length as, x)) M\<^sub>2" by (subst fk_alg_sketch[OF False]) (simp add:s\<^sub>1_def[symmetric] s\<^sub>2_def[symmetric]) have "M = map_pmf (\x. (s\<^sub>1, s\<^sub>2, k, length as, x)) M\<^sub>2 \ fk_result" by (subst M_def, subst a, simp) also have "... = M\<^sub>2 \ return_pmf \ median_rv" by (subst fk_result_eta) - (auto simp add:map_pmf_def bind_assoc_pmf bind_return_pmf median_rv_def mean_rv_def comp_def + (auto simp add:map_pmf_def bind_assoc_pmf bind_return_pmf median_rv_def mean_rv_def comp_def M\<^sub>1_def result_def median_restrict) finally have b: "M = M\<^sub>2 \ return_pmf \ median_rv" by simp - have result_exp: + have result_exp: "i\<^sub>1 < s\<^sub>1 \ i\<^sub>2 < s\<^sub>2 \ \\<^sub>2.expectation (\x. result (x (i\<^sub>1, i\<^sub>2))) = real_of_rat (F k as)" for i\<^sub>1 i\<^sub>2 unfolding \\<^sub>2_def M\<^sub>2_def using integrable_1[OF False] result_exp_1[OF False] by (subst expectation_Pi_pmf_slice, auto simp:\\<^sub>1_def) - have result_var: "\\<^sub>2.variance (\\. result (\ (i\<^sub>1, i\<^sub>2))) \ of_rat (\ * F k as)^2 * real s\<^sub>1 / 3" + have result_var: "\\<^sub>2.variance (\\. result (\ (i\<^sub>1, i\<^sub>2))) \ of_rat (\ * F k as)^2 * real s\<^sub>1 / 3" if result_var_assms: "i\<^sub>1 < s\<^sub>1" "i\<^sub>2 < s\<^sub>2" for i\<^sub>1 i\<^sub>2 proof - have "3 * real k * n powr (1 - 1 / real k) = (of_rat \)\<^sup>2 * (3 * real k * n powr (1 - 1 / real k) / (of_rat \)\<^sup>2)" using \_range by simp also have "... \ (real_of_rat \)\<^sup>2 * (real s\<^sub>1)" unfolding s\<^sub>1_def by (intro mult_mono of_nat_ceiling, simp_all) finally have f2_var_2: "3 * real k * n powr (1 - 1 / real k) \ (of_rat \)\<^sup>2 * (real s\<^sub>1)" by blast have "\\<^sub>2.variance (\\. result (\ (i\<^sub>1, i\<^sub>2)) :: real) = variance result" using result_var_assms integrable_1[OF False] - unfolding \\<^sub>2_def M\<^sub>2_def \\<^sub>1_def + unfolding \\<^sub>2_def M\<^sub>2_def \\<^sub>1_def by (subst variance_prod_pmf_slice, auto) also have "... \ of_rat (F k as)^2 * real k * n powr (1 - 1 / real k)" using assms False result_var_1 \\<^sub>1_def by simp also have "... = of_rat (F k as)^2 * (real k * n powr (1 - 1 / real k))" by (simp add:ac_simps) also have "... \ of_rat (F k as)^2 * (of_rat \^2 * (real s\<^sub>1 / 3))" - using f2_var_2 by (intro mult_left_mono, auto) + using f2_var_2 by (intro mult_left_mono, auto) also have "... = of_rat (F k as * \)^2 * (real s\<^sub>1 / 3)" - by (simp add: of_rat_mult power_mult_distrib) + by (simp add: of_rat_mult power_mult_distrib) also have "... = of_rat (\ * F k as)^2 * real s\<^sub>1 / 3" by (simp add:ac_simps) finally show ?thesis by simp qed have mean_rv_exp: "\\<^sub>2.expectation (\\. mean_rv \ i) = real_of_rat (F k as)" if mean_rv_exp_assms: "i < s\<^sub>2" for i proof - have "\\<^sub>2.expectation (\\. mean_rv \ i) = \\<^sub>2.expectation (\\. \n = 0..1. result (\ (n, i)) / real s\<^sub>1)" by (simp add:mean_rv_def sum_divide_distrib) also have "... = (\n = 0..1. \\<^sub>2.expectation (\\. result (\ (n, i))) / real s\<^sub>1)" using integrable_2[OF False] by (subst Bochner_Integration.integral_sum, auto) also have "... = of_rat (F k as)" using s1_nonzero mean_rv_exp_assms by (simp add:result_exp) finally show ?thesis by simp qed - have mean_rv_var: "\\<^sub>2.variance (\\. mean_rv \ i) \ real_of_rat (\ * F k as)^2/3" + have mean_rv_var: "\\<^sub>2.variance (\\. mean_rv \ i) \ real_of_rat (\ * F k as)^2/3" if mean_rv_var_assms: "i < s\<^sub>2" for i proof - have a:"\\<^sub>2.indep_vars (\_. borel) (\n x. result (x (n, i)) / real s\<^sub>1) {0..1}" unfolding \\<^sub>2_def M\<^sub>2_def using mean_rv_var_assms by (intro indep_vars_restrict_intro'[where f="fst"], simp, simp add:restrict_dfl_def, simp, simp) have "\\<^sub>2.variance (\\. mean_rv \ i) = \\<^sub>2.variance (\\. \j = 0..1. result (\ (j, i)) / real s\<^sub>1)" by (simp add:mean_rv_def sum_divide_distrib) also have "... = (\j = 0..1. \\<^sub>2.variance (\\. result (\ (j, i)) / real s\<^sub>1))" using a integrable_2[OF False] by (subst \\<^sub>2.bienaymes_identity_full_indep, auto simp add:\\<^sub>2_def) also have "... = (\j = 0..1. \\<^sub>2.variance (\\. result (\ (j, i))) / real s\<^sub>1^2)" using integrable_2[OF False] by (subst \\<^sub>2.variance_divide, auto) also have "... \ (\j = 0..1. ((real_of_rat (\ * F k as))\<^sup>2 * real s\<^sub>1 / 3) / (real s\<^sub>1^2))" using result_var[OF _ mean_rv_var_assms] by (intro sum_mono divide_right_mono, auto) also have "... = real_of_rat (\ * F k as)^2/3" using s1_nonzero by (simp add:algebra_simps power2_eq_square) finally show ?thesis by simp qed - have "\\<^sub>2.prob {y. of_rat (\ * F k as) < \mean_rv y i - real_of_rat (F k as)\} \ 1/3" + have "\\<^sub>2.prob {y. of_rat (\ * F k as) < \mean_rv y i - real_of_rat (F k as)\} \ 1/3" (is "?lhs \ _") if c_assms: "i < s\<^sub>2" for i proof - define a where "a = real_of_rat (\ * F k as)" have c: "0 < a" unfolding a_def using assms \_range fk_nonzero by (metis zero_less_of_rat_iff mult_pos_pos) have "?lhs \ \\<^sub>2.prob {y \ space \\<^sub>2. a \ \mean_rv y i - \\<^sub>2.expectation (\\. mean_rv \ i)\}" - by (intro \\<^sub>2.pmf_mono[OF \\<^sub>2_def], simp add:a_def mean_rv_exp[OF c_assms] space_\\<^sub>2) + by (intro \\<^sub>2.pmf_mono[OF \\<^sub>2_def], simp add:a_def mean_rv_exp[OF c_assms] space_\\<^sub>2) also have "... \ \\<^sub>2.variance (\\. mean_rv \ i)/a^2" by (intro \\<^sub>2.Chebyshev_inequality integrable_2 c False) (simp add:\\<^sub>2_def) also have "... \ 1/3" using c - using mean_rv_var[OF c_assms] + using mean_rv_var[OF c_assms] by (simp add:algebra_simps, simp add:a_def) finally show ?thesis by blast qed moreover have "\\<^sub>2.indep_vars (\_. borel) (\i \. mean_rv \ i) {0..2}" using s1_nonzero unfolding \\<^sub>2_def M\<^sub>2_def by (intro indep_vars_restrict_intro'[where f="snd"] finite_cartesian_product) (simp_all add:mean_rv_def restrict_dfl_def space_\\<^sub>2) moreover have " - (18 * ln (real_of_rat \)) \ real s\<^sub>2" by (simp add:s\<^sub>2_def, linarith) - ultimately have "1 - of_rat \ \ - \\<^sub>2.prob {y \ space \\<^sub>2. \median s\<^sub>2 (mean_rv y) - real_of_rat (F k as)\ \ of_rat (\ * F k as)}" + ultimately have "1 - of_rat \ \ + \\<^sub>2.prob {y \ space \\<^sub>2. \median s\<^sub>2 (mean_rv y) - real_of_rat (F k as)\ \ of_rat (\ * F k as)}" using \_range by (intro \\<^sub>2.median_bound_2, simp_all add:space_\\<^sub>2) - also have "... = \\<^sub>2.prob {y. \median_rv y - real_of_rat (F k as)\ \ real_of_rat (\ * F k as)}" + also have "... = \\<^sub>2.prob {y. \median_rv y - real_of_rat (F k as)\ \ real_of_rat (\ * F k as)}" by (simp add:median_rv_def space_\\<^sub>2) also have "... = \\<^sub>2.prob {y. \median_rv y - F k as\ \ \ * F k as}" by (simp add:real_of_rat_median_rv of_rat_less_eq flip: of_rat_diff) also have "... = \

(\ in measure_pmf M. \\ - F k as\ \ \ * F k as)" by (simp add: b comp_def map_pmf_def[symmetric] \\<^sub>2_def) finally show ?thesis by simp qed lemma fk_exact_space_usage': defines "M \ fold (\a state. state \ fk_update a) as (fk_init k \ \ n)" shows "AE \ in M. bit_count (encode_fk_state \) \ fk_space_usage (k, n, length as, \, \)" (is "AE \ in M. (_ \ ?rhs)") proof - define H where "H = (if as = [] then return_pmf (\i\ {0..1}\{0..2}. (0,0)) else M\<^sub>2)" have a:"M = map_pmf (\x.(s\<^sub>1,s\<^sub>2,k,length as, x)) H" proof (cases "as \ []") case True - then show ?thesis + then show ?thesis unfolding M_def fk_alg_sketch[OF True] H_def by (simp add:M\<^sub>2_def) next case False then show ?thesis by (simp add:H_def M_def s\<^sub>1_def[symmetric] Let_def s\<^sub>2_def[symmetric] map_pmf_def bind_return_pmf) qed have "bit_count (encode_fk_state (s\<^sub>1, s\<^sub>2, k, length as, y)) \ ?rhs" if b:"y \ set_pmf H" for y proof - have b0:" as \ [] \ y \ {0..1} \ {0..2} \\<^sub>E M\<^sub>1" using b non_empty_space fin_space by (simp add:H_def M\<^sub>2_def set_prod_pmf) - have "bit_count ((N\<^sub>e \\<^sub>e N\<^sub>e) (y x)) \ + have "bit_count ((N\<^sub>e \\<^sub>e N\<^sub>e) (y x)) \ ereal (2 * log 2 (real n + 1) + 1) + ereal (2 * log 2 (real (length as) + 1) + 1)" (is "_ \ ?rhs1") if b1_assms: "x \ {0..1}\{0..2}" for x proof - have "fst (y x) \ n" proof (cases "as = []") case True then show ?thesis using b b1_assms by (simp add:H_def) next case False hence "1 \ count_list as (fst (y x))" using b0 b1_assms by (simp add:PiE_iff case_prod_beta M\<^sub>1_def, fastforce) hence "fst (y x) \ set as" using count_list_gr_1 by metis - then show ?thesis + then show ?thesis by (meson lessThan_iff less_imp_le_nat subsetD as_range) qed - moreover have "snd (y x) \ length as" + moreover have "snd (y x) \ length as" proof (cases "as = []") case True then show ?thesis using b b1_assms by (simp add:H_def) next case False hence "(y x) \ M\<^sub>1" using b0 b1_assms by auto hence "snd (y x) \ count_list as (fst (y x))" by (simp add:M\<^sub>1_def case_prod_beta) then show ?thesis using count_le_length by (metis order_trans) qed ultimately have "bit_count (N\<^sub>e (fst (y x))) + bit_count (N\<^sub>e (snd (y x))) \ ?rhs1" using exp_golomb_bit_count_est by (intro add_mono, auto) thus ?thesis by (subst dependent_bit_count_2, simp) qed moreover have "y \ extensional ({0..1} \ {0..2})" using b0 b PiE_iff by (cases "as = []", auto simp:H_def PiE_iff) - ultimately have "bit_count ((List.product [0..1] [0..2] \\<^sub>e N\<^sub>e \\<^sub>e N\<^sub>e) y) \ + ultimately have "bit_count ((List.product [0..1] [0..2] \\<^sub>e N\<^sub>e \\<^sub>e N\<^sub>e) y) \ ereal (real s\<^sub>1 * real s\<^sub>2) * (ereal (2 * log 2 (real n + 1) + 1) + ereal (2 * log 2 (real (length as) + 1) + 1))" by (intro fun_bit_count_est[where xs="(List.product [0..1] [0..2])", simplified], auto) - hence "bit_count (encode_fk_state (s\<^sub>1, s\<^sub>2, k, length as, y)) \ - ereal (2 * log 2 (real s\<^sub>1 + 1) + 1) + - (ereal (2 * log 2 (real s\<^sub>2 + 1) + 1) + - (ereal (2 * log 2 (real k + 1) + 1) + - (ereal (2 * log 2 (real (length as) + 1) + 1) + - (ereal (real s\<^sub>1 * real s\<^sub>2) * (ereal (2 * log 2 (real n+1) + 1) + + hence "bit_count (encode_fk_state (s\<^sub>1, s\<^sub>2, k, length as, y)) \ + ereal (2 * log 2 (real s\<^sub>1 + 1) + 1) + + (ereal (2 * log 2 (real s\<^sub>2 + 1) + 1) + + (ereal (2 * log 2 (real k + 1) + 1) + + (ereal (2 * log 2 (real (length as) + 1) + 1) + + (ereal (real s\<^sub>1 * real s\<^sub>2) * (ereal (2 * log 2 (real n+1) + 1) + ereal (2 * log 2 (real (length as)+1) + 1))))))" unfolding encode_fk_state_def dependent_bit_count by (intro add_mono exp_golomb_bit_count, auto) - also have "... \ ?rhs" + also have "... \ ?rhs" by (simp add: s\<^sub>1_def[symmetric] s\<^sub>2_def[symmetric] Let_def) (simp add:ac_simps) finally show "bit_count (encode_fk_state (s\<^sub>1, s\<^sub>2, k, length as, y)) \ ?rhs" by blast qed thus ?thesis by (simp add: a AE_measure_pmf_iff del:fk_space_usage.simps) qed end text \Main results of this section:\ theorem fk_alg_correct: assumes "k \ 1" assumes "\ \ {0<..<1}" assumes "\ > 0" assumes "set as \ {.. fold (\a state. state \ fk_update a) as (fk_init k \ \ n) \ fk_result" shows "\

(\ in measure_pmf M. \\ - F k as\ \ \ * F k as) \ 1 - of_rat \" unfolding M_def using fk_alg_correct'[OF assms(1-4)] by blast theorem fk_exact_space_usage: assumes "k \ 1" assumes "\ \ {0<..<1}" assumes "\ > 0" assumes "set as \ {.. fold (\a state. state \ fk_update a) as (fk_init k \ \ n)" shows "AE \ in M. bit_count (encode_fk_state \) \ fk_space_usage (k, n, length as, \, \)" unfolding M_def using fk_exact_space_usage'[OF assms(1-4)] by blast theorem fk_asymptotic_space_complexity: - "fk_space_usage \ + "fk_space_usage \ O[at_top \\<^sub>F at_top \\<^sub>F at_top \\<^sub>F at_right (0::rat) \\<^sub>F at_right (0::rat)](\ (k, n, m, \, \). real k * real n powr (1-1/ real k) / (of_rat \)\<^sup>2 * (ln (1 / of_rat \)) * (ln (real n) + ln (real m)))" (is "_ \ O[?F](?rhs)") proof - define k_of :: "nat \ nat \ nat \ rat \ rat \ nat" where "k_of = (\(k, n, m, \, \). k)" define n_of :: "nat \ nat \ nat \ rat \ rat \ nat" where "n_of = (\(k, n, m, \, \). n)" define m_of :: "nat \ nat \ nat \ rat \ rat \ nat" where "m_of = (\(k, n, m, \, \). m)" define \_of :: "nat \ nat \ nat \ rat \ rat \ rat" where "\_of = (\(k, n, m, \, \). \)" define \_of :: "nat \ nat \ nat \ rat \ rat \ rat" where "\_of = (\(k, n, m, \, \). \)" - define g1 where + define g1 where "g1 = (\x. real (k_of x)*(real (n_of x)) powr (1-1/ real (k_of x)) * (1 / of_rat (\_of x)^2))" - define g where + define g where "g = (\x. g1 x * (ln (1 / of_rat (\_of x))) * (ln (real (n_of x)) + ln (real (m_of x))))" - define s1_of where "s1_of = (\x. + define s1_of where "s1_of = (\x. nat \3 * real (k_of x) * real (n_of x) powr (1 - 1 / real (k_of x)) / (real_of_rat (\_of x))\<^sup>2\)" define s2_of where "s2_of = (\x. nat \- (18 * ln (real_of_rat (\_of x)))\)" - have evt: "(\x. - 0 < real_of_rat (\_of x) \ 0 < real_of_rat (\_of x) \ + have evt: "(\x. + 0 < real_of_rat (\_of x) \ 0 < real_of_rat (\_of x) \ 1/real_of_rat (\_of x) \ \ \ 1/real_of_rat (\_of x) \ \ \ - real (n_of x) \ n \ real (k_of x) \ k \ real (m_of x) \ m\ P x) + real (n_of x) \ n \ real (k_of x) \ k \ real (m_of x) \ m\ P x) \ eventually P ?F" (is "(\x. ?prem x \ _) \ _") for \ \ n k m P apply (rule eventually_mono[where P="?prem" and Q="P"]) apply (simp add:\_of_def case_prod_beta' \_of_def n_of_def k_of_def m_of_def) apply (intro eventually_conj eventually_prod1' eventually_prod2' sequentially_inf eventually_at_right_less inv_at_right_0_inf) by (auto simp add:prod_filter_eq_bot) - have 1: + have 1: "(\_. 1) \ O[?F](\x. real (n_of x))" "(\_. 1) \ O[?F](\x. real (m_of x))" "(\_. 1) \ O[?F](\x. real (k_of x))" by (intro landau_o.big_mono eventually_mono[OF evt], auto)+ have "(\x. ln (real (m_of x) + 1)) \ O[?F](\x. ln (real (m_of x)))" by (intro landau_ln_2[where a="2"] evt[where m="2"] sum_in_bigo 1, auto) hence 2: " (\x. log 2 (real (m_of x) + 1)) \ O[?F](\x. ln (real (n_of x)) + ln (real (m_of x)))" by (intro landau_sum_2 eventually_mono[OF evt[where n="1" and m="1"]]) (auto simp add:log_def) - have 3: "(\_. 1) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" + have 3: "(\_. 1) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" using order_less_le_trans[OF exp_gt_zero] ln_ge_iff by (intro landau_o.big_mono evt[where \="exp 1"]) (simp add: abs_ge_iff, blast) have 4: "(\_. 1) \ O[?F](\x. 1 / (real_of_rat (\_of x))\<^sup>2)" using one_le_power by (intro landau_o.big_mono evt[where \="1"]) (simp add:power_one_over[symmetric], blast) have "(\x. 1) \ O[?F](\x. ln (real (n_of x)))" using order_less_le_trans[OF exp_gt_zero] ln_ge_iff by (intro landau_o.big_mono evt[where n="exp 1"]) (simp add: abs_ge_iff, blast) hence 5: "(\x. 1) \ O[?F](\x. ln (real (n_of x)) + ln (real (m_of x)))" by (intro landau_sum_1 evt[where n="1" and m="1"], auto) - have "(\x. -ln(of_rat (\_of x))) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" + have "(\x. -ln(of_rat (\_of x))) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" by (intro landau_o.big_mono evt) (auto simp add:ln_div) hence 6: "(\x. real (s2_of x)) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" unfolding s2_of_def by (intro landau_nat_ceil 3, simp) have 7: "(\_. 1) \ O[?F](\x. real (n_of x) powr (1 - 1 / real (k_of x)))" by (intro landau_o.big_mono evt[where n="1" and k="1"]) (auto simp add: ge_one_powr_ge_zero) have 8: "(\_. 1) \ O[?F](g1)" unfolding g1_def by (intro landau_o.big_mult_1 1 7 4) have "(\x. 3 * (real (k_of x) * (n_of x) powr (1 - 1 / real (k_of x)) / (of_rat (\_of x))\<^sup>2)) \ O[?F](g1)" by (subst landau_o.big.cmult_in_iff, simp, simp add:g1_def) hence 9: "(\x. real (s1_of x)) \ O[?F](g1)" unfolding s1_of_def by (intro landau_nat_ceil 8, auto simp:ac_simps) - have 10: "(\_. 1) \ O[?F](g)" + have 10: "(\_. 1) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1 8 3 5) - + have "(\x. real (s1_of x)) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1 5 3 9) hence "(\x. ln (real (s1_of x) + 1)) \ O[?F](g)" using 10 by (intro landau_ln_3 sum_in_bigo, auto) hence 11: "(\x. log 2 (real (s1_of x) + 1)) \ O[?F](g)" by (simp add:log_def) have 12: " (\x. ln (real (s2_of x) + 1)) \ O[?F](\x. ln (1 / real_of_rat (\_of x)))" using evt[where \="2"] 6 3 by (intro landau_ln_3 sum_in_bigo, auto) - have 13: "(\x. log 2 (real (s2_of x) + 1)) \ O[?F](g)" - unfolding g_def + have 13: "(\x. log 2 (real (s2_of x) + 1)) \ O[?F](g)" + unfolding g_def by (rule landau_o.big_mult_1, rule landau_o.big_mult_1', auto simp add: 8 5 12 log_def) have "(\x. real (k_of x)) \ O[?F](g1)" unfolding g1_def using 7 4 by (intro landau_o.big_mult_1, simp_all) hence "(\x. log 2 (real (k_of x) + 1)) \ O[?F](g1)" by (simp add:log_def) (intro landau_ln_3 sum_in_bigo 8, auto) hence 14: "(\x. log 2 (real (k_of x) + 1)) \ O[?F](g)" unfolding g_def by (intro landau_o.big_mult_1 3 5) have 15: "(\x. log 2 (real (m_of x) + 1)) \ O[?F](g)" unfolding g_def using 2 8 3 by (intro landau_o.big_mult_1', simp_all) have "(\x. ln (real (n_of x) + 1)) \ O[?F](\x. ln (real (n_of x)))" by (intro landau_ln_2[where a="2"] eventually_mono[OF evt[where n="2"]] sum_in_bigo 1, auto) hence " (\x. log 2 (real (n_of x) + 1)) \ O[?F](\x. ln (real (n_of x)) + ln (real (m_of x)))" by (intro landau_sum_1 evt[where n="1" and m="1"]) (auto simp add:log_def) hence 16: "(\x. real (s1_of x) * real (s2_of x) * - (2 + 2 * log 2 (real (n_of x) + 1) + 2 * log 2 (real (m_of x) + 1))) \ O[?F](g)" + (2 + 2 * log 2 (real (n_of x) + 1) + 2 * log 2 (real (m_of x) + 1))) \ O[?F](g)" unfolding g_def using 9 6 5 2 by (intro landau_o.mult sum_in_bigo, auto) have "fk_space_usage = (\x. fk_space_usage (k_of x, n_of x, m_of x, \_of x, \_of x))" by (simp add:case_prod_beta' k_of_def n_of_def \_of_def \_of_def m_of_def) also have "... \ O[?F](g)" using 10 11 13 14 15 16 by (simp add:fun_cong[OF s1_of_def[symmetric]] fun_cong[OF s2_of_def[symmetric]] Let_def) (intro sum_in_bigo, auto) also have "... = O[?F](?rhs)" by (simp add:case_prod_beta' g1_def g_def n_of_def \_of_def \_of_def m_of_def k_of_def) finally show ?thesis by simp qed end diff --git a/thys/Frequency_Moments/Frequency_Moments.thy b/thys/Frequency_Moments/Frequency_Moments.thy --- a/thys/Frequency_Moments/Frequency_Moments.thy +++ b/thys/Frequency_Moments/Frequency_Moments.thy @@ -1,115 +1,115 @@ section "Frequency Moments" theory Frequency_Moments - imports + imports Frequency_Moments_Preliminary_Results Universal_Hash_Families.Universal_Hash_Families_More_Finite_Fields Interpolation_Polynomials_HOL_Algebra.Interpolation_Polynomial_Cardinalities begin -text \This section contains a definition of the frequency moments of a stream and a few general results about +text \This section contains a definition of the frequency moments of a stream and a few general results about frequency moments..\ definition F where "F k xs = (\ x \ set xs. (rat_of_nat (count_list xs x)^k))" lemma F_ge_0: "F k as \ 0" unfolding F_def by (rule sum_nonneg, simp) lemma F_gr_0: assumes "as \ []" shows "F k as > 0" proof - have "rat_of_nat 1 \ rat_of_nat (card (set as))" - using assms card_0_eq[where A="set as"] + using assms card_0_eq[where A="set as"] by (intro of_nat_mono) (metis List.finite_set One_nat_def Suc_leI neq0_conv set_empty) also have "... = (\x\set as. 1)" by simp also have "... \ (\x\set as. rat_of_nat (count_list as x) ^ k)" by (intro sum_mono one_le_power) (metis count_list_gr_1 of_nat_1 of_nat_le_iff) also have "... \ F k as" by (simp add:F_def) finally show ?thesis by simp qed definition P\<^sub>e :: "nat \ nat \ nat list \ bool list option" where "P\<^sub>e p n f = (if p > 1 \ f \ bounded_degree_polynomials (mod_ring p) n then ([0..\<^sub>e Nb\<^sub>e p) (\i \ {..e p n)" proof (cases "p > 1") case True interpret cring "mod_ring p" using mod_ring_is_cring True by blast have a:"inj_on (\x. (\i \ {.. bounded_degree_polynomials (mod_ring p) n" assume c:"y \ bounded_degree_polynomials (mod_ring p) n" assume d:"restrict (coeff x) {.. n" by linarith have "coeff x i = \\<^bsub>mod_ring p\<^esub>" using b e by (subst coeff_length, auto simp:bounded_degree_polynomials_length) also have "... = coeff y i" using c e by (subst coeff_length, auto simp:bounded_degree_polynomials_length) finally show ?thesis by simp qed then show "x = y" - using b c univ_poly_carrier - by (subst coeff_iff_polynomial_cond) (auto simp:bounded_degree_polynomials_length) + using b c univ_poly_carrier + by (subst coeff_iff_polynomial_cond) (auto simp:bounded_degree_polynomials_length) qed have "is_encoding (\f. P\<^sub>e p n f)" unfolding P\<^sub>e_def using a True - by (intro encoding_compose[where f="([0..\<^sub>e Nb\<^sub>e p)"] fun_encoding bounded_nat_encoding) + by (intro encoding_compose[where f="([0..\<^sub>e Nb\<^sub>e p)"] fun_encoding bounded_nat_encoding) auto thus ?thesis by simp next case False hence "is_encoding (\f. P\<^sub>e p n f)" unfolding P\<^sub>e_def using encoding_triv by simp then show ?thesis by simp qed lemma bounded_degree_polynomial_bit_count: assumes "p > 1" assumes "x \ bounded_degree_polynomials (mod_ring p) n" shows "bit_count (P\<^sub>e p n x) \ ereal (real n * (log 2 p + 1))" proof - interpret cring "mod_ring p" using mod_ring_is_cring assms by blast have a: "x \ carrier (poly_ring (mod_ring p))" using assms(2) by (simp add:bounded_degree_polynomials_def) have "real_of_int \log 2 (p-1)\+1 \ log 2 (p-1) + 1" - using floor_eq_iff by (intro add_mono, auto) + using floor_eq_iff by (intro add_mono, auto) also have "... \ log 2 p + 1" using assms by (intro add_mono, auto) finally have b: "\log 2 (p-1)\+1 \ log 2 p + 1" by simp have "bit_count (P\<^sub>e p n x) = (\ k \ [0..e p (coeff x k)))" - using assms restrict_extensional + using assms restrict_extensional by (auto intro!:arg_cong[where f="sum_list"] simp add:P\<^sub>e_def fun_bit_count lessThan_atLeast0) also have "... = (\ k \ [0..log 2 (p-1)\+1)" + also have "... = n * real_of_int (\log 2 (p-1)\+1)" using assms(1) by (simp add:floorlog_def) - also have "... \ ereal (real n * (log 2 p + 1))" + also have "... \ ereal (real n * (log 2 p + 1))" by (subst ereal_less_eq, intro mult_left_mono b, auto) finally show ?thesis by simp qed end \ No newline at end of file diff --git a/thys/Frequency_Moments/Frequency_Moments_Preliminary_Results.thy b/thys/Frequency_Moments/Frequency_Moments_Preliminary_Results.thy --- a/thys/Frequency_Moments/Frequency_Moments_Preliminary_Results.thy +++ b/thys/Frequency_Moments/Frequency_Moments_Preliminary_Results.thy @@ -1,491 +1,464 @@ section \Preliminary Results\ theory Frequency_Moments_Preliminary_Results - imports - "HOL.Transcendental" + imports + "HOL.Transcendental" "HOL-Computational_Algebra.Primes" "HOL-Library.Extended_Real" "HOL-Library.Multiset" "HOL-Library.Sublist" Prefix_Free_Code_Combinators.Prefix_Free_Code_Combinators Bertrands_Postulate.Bertrand + Expander_Graphs.Expander_Graphs_Multiset_Extras begin text \This section contains various preliminary results.\ lemma card_ordered_pairs: - fixes M :: "('a ::linorder) set" + fixes M :: "('a ::linorder) set" assumes "finite M" shows "2 * card {(x,y) \ M \ M. x < y} = card M * (card M - 1)" proof - have a: "finite (M \ M)" using assms by simp have inj_swap: "inj (\x. (snd x, fst x))" by (rule inj_onI, simp add: prod_eq_iff) have "2 * card {(x,y) \ M \ M. x < y} = card {(x,y) \ M \ M. x < y} + card ((\x. (snd x, fst x))`{(x,y) \ M \ M. x < y})" by (simp add: card_image[OF inj_on_subset[OF inj_swap]]) also have "... = card {(x,y) \ M \ M. x < y} + card {(x,y) \ M \ M. y < x}" by (auto intro: arg_cong[where f="card"] simp add:set_eq_iff image_iff) also have "... = card ({(x,y) \ M \ M. x < y} \ {(x,y) \ M \ M. y < x})" by (intro card_Un_disjoint[symmetric] a finite_subset[where B="M \ M"] subsetI) auto also have "... = card ((M \ M) - {(x,y) \ M \ M. x = y})" - by (auto intro: arg_cong[where f="card"] simp add:set_eq_iff) + by (auto intro: arg_cong[where f="card"] simp add:set_eq_iff) also have "... = card (M \ M) - card {(x,y) \ M \ M. x = y}" by (intro card_Diff_subset a finite_subset[where B="M \ M"] subsetI) auto also have "... = card M ^ 2 - card ((\x. (x,x)) ` M)" using assms by (intro arg_cong2[where f="(-)"] arg_cong[where f="card"]) (auto simp:power2_eq_square set_eq_iff image_iff) also have "... = card M ^ 2 - card M" by (intro arg_cong2[where f="(-)"] card_image inj_onI, auto) also have "... = card M * (card M - 1)" by (cases "card M \ 0", auto simp:power2_eq_square algebra_simps) finally show ?thesis by simp qed lemma ereal_mono: "x \ y \ ereal x \ ereal y" by simp lemma log_mono: "a > 1 \ x \ y \ 0 < x \ log a x \ log a y" by (subst log_le_cancel_iff, auto) lemma abs_ge_iff: "((x::real) \ abs y) = (x \ y \ x \ -y)" by linarith lemma count_list_gr_1: "(x \ set xs) = (count_list xs x \ 1)" by (induction xs, simp, simp) lemma count_list_append: "count_list (xs@ys) v = count_list xs v + count_list ys v" by (induction xs, simp, simp) lemma count_list_lt_suffix: assumes "suffix a b" assumes "x \ {b ! i| i. i < length b - length a}" shows "count_list a x < count_list b x" proof - - have "length a \ length b" using assms(1) + have "length a \ length b" using assms(1) by (simp add: suffix_length_le) hence "x \ set (nths b {i. i < length b - length a})" - using assms diff_commute by (auto simp add:set_nths) + using assms diff_commute by (auto simp add:set_nths) hence a:"x \ set (take (length b - length a) b)" by (subst (asm) lessThan_def[symmetric], simp) have "b = (take (length b - length a) b)@drop (length b - length a) b" by simp also have "... = (take (length b - length a) b)@a" - using assms(1) suffix_take by auto + using assms(1) suffix_take by auto finally have b:"b = (take (length b - length a) b)@a" by simp have "count_list a x < 1 + count_list a x" by simp also have "... \ count_list (take (length b - length a) b) x + count_list a x" using a count_list_gr_1 - by (intro add_mono, fast, simp) + by (intro add_mono, fast, simp) also have "... = count_list b x" using b count_list_append by metis finally show ?thesis by simp qed lemma suffix_drop_drop: assumes "x \ y" shows "suffix (drop x a) (drop y a)" proof - have "drop y a = take (x - y) (drop y a)@drop (x- y) (drop y a)" by (subst append_take_drop_id, simp) also have " ... = take (x-y) (drop y a)@drop x a" using assms by simp finally have "drop y a = take (x-y) (drop y a)@drop x a" by simp - thus ?thesis - by (auto simp add:suffix_def) + thus ?thesis + by (auto simp add:suffix_def) qed lemma count_list_card: "count_list xs x = card {k. k < length xs \ xs ! k = x}" proof - have "count_list xs x = length (filter ((=) x) xs)" by (induction xs, simp, simp) also have "... = card {k. k < length xs \ xs ! k = x}" by (subst length_filter_conv_card, metis) finally show ?thesis by simp qed lemma card_gr_1_iff: assumes "finite S" "x \ S" "y \ S" "x \ y" shows "card S > 1" using assms card_le_Suc0_iff_eq leI by auto lemma count_list_ge_2_iff: assumes "y < z" assumes "z < length xs" assumes "xs ! y = xs ! z" shows "count_list xs (xs ! y) > 1" proof - have " 1 < card {k. k < length xs \ xs ! k = xs ! y}" using assms by (intro card_gr_1_iff[where x="y" and y="z"], auto) thus ?thesis by (simp add: count_list_card) qed text \Results about multisets and sorting\ -text \This is a induction scheme over the distinct elements of a multisets: -We can represent each multiset as a sum like: -@{text "replicate_mset n\<^sub>1 x\<^sub>1 + replicate_mset n\<^sub>2 x\<^sub>2 + ... + replicate_mset n\<^sub>k x\<^sub>k"} where the -@{term "x\<^sub>i"} are distinct.\ +lemmas disj_induct_mset = disj_induct_mset -lemma disj_induct_mset: - assumes "P {#}" - assumes "\n M x. P M \ \(x \# M) \ n > 0 \ P (M + replicate_mset n x)" - shows "P M" -proof (induction "size M" arbitrary: M rule:nat_less_induct) - case 1 - show ?case - proof (cases "M = {#}") - case True - then show ?thesis using assms by simp - next - case False - then obtain x where x_def: "x \# M" using multiset_nonemptyE by auto - define M1 where "M1 = M - replicate_mset (count M x) x" - then have M_def: "M = M1 + replicate_mset (count M x) x" - by (metis count_le_replicate_mset_subset_eq dual_order.refl subset_mset.diff_add) - have "size M1 < size M" - by (metis M_def x_def count_greater_zero_iff less_add_same_cancel1 size_replicate_mset size_union) - hence "P M1" using 1 by blast - then show "P M" - apply (subst M_def, rule assms(2), simp) - by (simp add:M1_def x_def count_eq_zero_iff[symmetric])+ - qed -qed - -lemma prod_mset_conv: +lemma prod_mset_conv: fixes f :: "'a \ 'b::{comm_monoid_mult}" shows "prod_mset (image_mset f A) = prod (\x. f x^(count A x)) (set_mset A)" proof (induction A rule: disj_induct_mset) case 1 then show ?case by simp next case (2 n M x) moreover have "count M x = 0" using 2 by (simp add: count_eq_zero_iff) moreover have "\y. y \ set_mset M \ y \ x" using 2 by blast - ultimately show ?case by (simp add:algebra_simps) + ultimately show ?case by (simp add:algebra_simps) qed text \There is a version @{thm [source] sum_list_map_eq_sum_count} but it doesn't work if the function maps into the reals.\ lemma sum_list_eval: fixes f :: "'a \ 'b::{ring,semiring_1}" shows "sum_list (map f xs) = (\x \ set xs. of_nat (count_list xs x) * f x)" proof - define M where "M = mset xs" have "sum_mset (image_mset f M) = (\x \ set_mset M. of_nat (count M x) * f x)" proof (induction "M" rule:disj_induct_mset) case 1 then show ?case by simp next case (2 n M x) have a:"\y. y \ set_mset M \ y \ x" using 2(2) by blast show ?case using 2 by (simp add:a count_eq_zero_iff[symmetric]) qed - moreover have "\x. count_list xs x = count (mset xs) x" + moreover have "\x. count_list xs x = count (mset xs) x" by (induction xs, simp, simp) ultimately show ?thesis by (simp add:M_def sum_mset_sum_list[symmetric]) qed lemma prod_list_eval: fixes f :: "'a \ 'b::{ring,semiring_1,comm_monoid_mult}" shows "prod_list (map f xs) = (\x \ set xs. (f x)^(count_list xs x))" proof - define M where "M = mset xs" have "prod_mset (image_mset f M) = (\x \ set_mset M. f x ^ (count M x))" proof (induction "M" rule:disj_induct_mset) case 1 then show ?case by simp next case (2 n M x) have a:"\y. y \ set_mset M \ y \ x" using 2(2) by blast - have b:"count M x = 0" using 2 by (subst count_eq_zero_iff) blast + have b:"count M x = 0" using 2 by (subst count_eq_zero_iff) blast show ?case using 2 by (simp add:a b mult.commute) qed - moreover have "\x. count_list xs x = count (mset xs) x" + moreover have "\x. count_list xs x = count (mset xs) x" by (induction xs, simp, simp) ultimately show ?thesis by (simp add:M_def prod_mset_prod_list[symmetric]) qed lemma sorted_sorted_list_of_multiset: "sorted (sorted_list_of_multiset M)" - by (induction M, auto simp:sorted_insort) + by (induction M, auto simp:sorted_insort) lemma count_mset: "count (mset xs) a = count_list xs a" by (induction xs, auto) lemma swap_filter_image: "filter_mset g (image_mset f A) = image_mset f (filter_mset (g \ f) A)" by (induction A, auto) lemma list_eq_iff: assumes "mset xs = mset ys" assumes "sorted xs" assumes "sorted ys" - shows "xs = ys" + shows "xs = ys" using assms properties_for_sort by blast lemma sorted_list_of_multiset_image_commute: assumes "mono f" shows "sorted_list_of_multiset (image_mset f M) = map f (sorted_list_of_multiset M)" proof - - have "sorted (sorted_list_of_multiset (image_mset f M))" + have "sorted (sorted_list_of_multiset (image_mset f M))" by (simp add:sorted_sorted_list_of_multiset) moreover have " sorted_wrt (\x y. f x \ f y) (sorted_list_of_multiset M)" - by (rule sorted_wrt_mono_rel[where P="\x y. x \ y"]) + by (rule sorted_wrt_mono_rel[where P="\x y. x \ y"]) (auto intro: monoD[OF assms] sorted_sorted_list_of_multiset) hence "sorted (map f (sorted_list_of_multiset M))" by (subst sorted_wrt_map) ultimately show ?thesis by (intro list_eq_iff, auto) qed text \Results about rounding and floating point numbers\ lemma round_down_ge: "x \ round_down prec x + 2 powr (-prec)" using round_down_correct by (simp, meson diff_diff_eq diff_eq_diff_less_eq) lemma truncate_down_ge: "x \ truncate_down prec x + abs x * 2 powr (-prec)" proof (cases "abs x > 0") case True have "x \ round_down (int prec - \log 2 \x\\) x + 2 powr (-real_of_int(int prec - \log 2 \x\\))" by (rule round_down_ge) also have "... \ truncate_down prec x + 2 powr ( \log 2 \x\\) * 2 powr (-real prec)" by (rule add_mono, simp_all add:powr_add[symmetric] truncate_down_def) also have "... \ truncate_down prec x + \x\ * 2 powr (-real prec)" using True by (intro add_mono mult_right_mono, simp_all add:le_log_iff[symmetric]) finally show ?thesis by simp next case False then show ?thesis by simp qed lemma truncate_down_pos: assumes "x \ 0" shows "x * (1 - 2 powr (-prec)) \ truncate_down prec x" by (simp add:right_diff_distrib diff_le_eq) (metis truncate_down_ge assms abs_of_nonneg) lemma truncate_down_eq: assumes "truncate_down r x = truncate_down r y" shows "abs (x-y) \ max (abs x) (abs y) * 2 powr (-real r)" -proof - +proof - have "x - y \ truncate_down r x + abs x * 2 powr (-real r) - y" by (rule diff_right_mono, rule truncate_down_ge) also have "... \ y + abs x * 2 powr (-real r) - y" using truncate_down_le by (intro diff_right_mono add_mono, subst assms(1), simp_all) also have "... \ abs x * 2 powr (-real r)" by simp also have "... \ max (abs x) (abs y) * 2 powr (-real r)" by simp finally have a:"x - y \ max (abs x) (abs y) * 2 powr (-real r)" by simp have "y - x \ truncate_down r y + abs y * 2 powr (-real r) - x" by (rule diff_right_mono, rule truncate_down_ge) also have "... \ x + abs y * 2 powr (-real r) - x" using truncate_down_le by (intro diff_right_mono add_mono, subst assms(1)[symmetric], auto) also have "... \ abs y * 2 powr (-real r)" by simp also have "... \ max (abs x) (abs y) * 2 powr (-real r)" by simp finally have b:"y - x \ max (abs x) (abs y) * 2 powr (-real r)" by simp show ?thesis using abs_le_iff a b by linarith qed -definition rat_of_float :: "float \ rat" where - "rat_of_float f = of_int (mantissa f) * - (if exponent f \ 0 then 2 ^ (nat (exponent f)) else 1 / 2 ^ (nat (-exponent f)))" +definition rat_of_float :: "float \ rat" where + "rat_of_float f = of_int (mantissa f) * + (if exponent f \ 0 then 2 ^ (nat (exponent f)) else 1 / 2 ^ (nat (-exponent f)))" lemma real_of_rat_of_float: "real_of_rat (rat_of_float x) = real_of_float x" proof - have "real_of_rat (rat_of_float x) = mantissa x * (2 powr (exponent x))" by (simp add:rat_of_float_def of_rat_mult of_rat_divide of_rat_power powr_realpow[symmetric] powr_minus_divide) also have "... = real_of_float x" using mantissa_exponent by simp - finally show ?thesis by simp + finally show ?thesis by simp qed lemma log_est: "log 2 (real n + 1) \ n" proof - have "1 + real n = real (n + 1)" by simp also have "... \ real (2 ^ n)" by (intro of_nat_mono suc_n_le_2_pow_n) also have "... = 2 powr (real n)" by (simp add:powr_realpow) finally have "1 + real n \ 2 powr (real n)" by simp thus ?thesis by (simp add: Transcendental.log_le_iff) qed lemma truncate_mantissa_bound: "abs (\x * 2 powr (real r - real_of_int \log 2 \x\\)\) \ 2 ^ (r+1)" (is "?lhs \ _") proof - define q where "q = \x * 2 powr (real r - real_of_int (\log 2 \x\\))\" have "abs q \ 2 ^ (r + 1)" if a:"x > 0" proof - have "abs q = q" using a by (intro abs_of_nonneg, simp add:q_def) also have "... \ x * 2 powr (real r - real_of_int \log 2 \x\\)" unfolding q_def using of_int_floor_le by blast also have "... = x * 2 powr real_of_int (int r - \log 2 \x\\)" by auto also have "... = 2 powr (log 2 x + real_of_int (int r - \log 2 \x\\))" using a by (simp add:powr_add) also have "... \ 2 powr (real r + 1)" - using a by (intro powr_mono, linarith+) + using a by (intro powr_mono, linarith+) also have "... = 2 ^ (r+1)" by (subst powr_realpow[symmetric], simp_all add:add.commute) - finally show "abs q \ 2 ^ (r+1)" + finally show "abs q \ 2 ^ (r+1)" by (metis of_int_le_iff of_int_numeral of_int_power) qed - + moreover have "abs q \ (2 ^ (r + 1))" if a: "x < 0" proof - have "-(2 ^ (r+1) + 1) = -(2 powr (real r + 1)+1)" by (subst powr_realpow[symmetric], simp_all add: add.commute) also have "... < -(2 powr (log 2 (- x) + (r - \log 2 \x\\)) + 1)" using a by (simp, linarith) also have "... = x * 2 powr (r - \log 2 \x\\) - 1" using a by (simp add:powr_add) also have "... \ q" by (simp add:q_def) also have "... = - abs q" using a by (subst abs_of_neg, simp_all add: mult_pos_neg2 q_def) finally have "-(2 ^ (r+1)+1) < - abs q" using of_int_less_iff by fastforce hence "-(2 ^ (r+1)) \ - abs q" by linarith thus "abs q \ 2^(r+1)" by linarith qed moreover have "x = 0 \ abs q \ 2^(r+1)" by (simp add:q_def) ultimately have "abs q \ 2^(r+1)" by fastforce thus ?thesis using q_def by blast qed lemma truncate_float_bit_count: - "bit_count (F\<^sub>e (float_of (truncate_down r x))) \ 10 + 4 * real r + 2*log 2 (2 + \log 2 \x\\)" + "bit_count (F\<^sub>e (float_of (truncate_down r x))) \ 10 + 4 * real r + 2*log 2 (2 + \log 2 \x\\)" (is "?lhs \ ?rhs") proof - define m where "m = \x * 2 powr (real r - real_of_int \log 2 \x\\)\" define e where "e = \log 2 \x\\ - int r" have a: "(real_of_int \log 2 \x\\ - real r) = e" by (simp add:e_def) have "abs m + 2 \ 2 ^ (r + 1) + 2^1" using truncate_mantissa_bound by (intro add_mono, simp_all add:m_def) also have "... \ 2 ^ (r+2)" by simp finally have b:"abs m + 2 \ 2 ^ (r+2)" by simp - hence "real_of_int (\m\ + 2) \ real_of_int (4 * 2 ^ r)" + hence "real_of_int (\m\ + 2) \ real_of_int (4 * 2 ^ r)" by (subst of_int_le_iff, simp) - hence "\real_of_int m\ + 2 \ 4 * 2 ^ r" + hence "\real_of_int m\ + 2 \ 4 * 2 ^ r" by simp hence c:"log 2 (real_of_int (\m\ + 2)) \ r+2" by (simp add: Transcendental.log_le_iff powr_add powr_realpow) have "real_of_int (abs e + 1) \ real_of_int \\log 2 \x\\\ + real_of_int r + 1" by (simp add:e_def) also have "... \ 1 + abs (log 2 (abs x)) + real_of_int r + 1" by (simp add:abs_le_iff, linarith) also have "... \ (real_of_int r+ 1) * (2 + abs (log 2 (abs x)))" by (simp add:distrib_left distrib_right) finally have d:"real_of_int (abs e + 1) \ (real_of_int r+ 1) * (2 + abs (log 2 (abs x)))" by simp have "log 2 (real_of_int (abs e + 1)) \ log 2 (real_of_int r + 1) + log 2 (2 + abs (log 2 (abs x)))" using d by (simp add: log_mult[symmetric]) also have "... \ r + log 2 (2 + abs (log 2 (abs x)))" using log_est by (intro add_mono, simp_all add:add.commute) finally have e: "log 2 (real_of_int (abs e + 1)) \ r + log 2 (2 + abs (log 2 (abs x)))" by simp have "?lhs = bit_count (F\<^sub>e (float_of (real_of_int m * 2 powr real_of_int e)))" by (simp add:truncate_down_def round_down_def m_def[symmetric] a) also have "... \ ereal (6 + (2 * log 2 (real_of_int (\m\ + 2)) + 2 * log 2 (real_of_int (\e\ + 1))))" using float_bit_count_2 by simp also have "... \ ereal (6 + (2 * real (r+2) + 2 * (r + log 2 (2 + abs (log 2 (abs x))))))" using c e - by (subst ereal_less_eq, intro add_mono mult_left_mono, linarith+) + by (subst ereal_less_eq, intro add_mono mult_left_mono, linarith+) also have "... = ?rhs" by simp finally show ?thesis by simp qed -definition prime_above :: "nat \ nat" +definition prime_above :: "nat \ nat" where "prime_above n = (SOME x. x \ {n..(2*n+2)} \ prime x)" text \The term @{term"prime_above n"} returns a prime between @{term "n::nat"} and @{term "2*(n::nat)+2"}. Because of Bertrand's postulate there always is such a value. In a refinement of the algorithms, it may make sense to replace this with an algorithm, that finds such a prime exactly or approximately. The definition is intentionally inexact, to allow refinement with various algorithms, without modifying the high-level mathematical correctness proof.\ lemma ex_subset: assumes "\x \ A. P x" assumes "A \ B" shows "\x \ B. P x" using assms by auto lemma shows prime_above_prime: "prime (prime_above n)" and prime_above_range: "prime_above n \ {n..(2*n+2)}" proof - define r where "r = (\x. x \ {n..(2*n+2)} \ prime x)" have "\x. r x" proof (cases "n>2") case True hence "n-1 > 1" by simp hence "\x \ {(n-1)<..<(2*(n-1))}. prime x" using bertrand by simp moreover have "{n - 1<..<2 * (n - 1)} \ {n..2 * n + 2}" - by (intro subsetI, auto) + by (intro subsetI, auto) ultimately have "\x \ {n..(2*n+2)}. prime x" by (rule ex_subset) then show ?thesis by (simp add:r_def Bex_def) next case False - hence "2 \ {n..(2*n+2)}" + hence "2 \ {n..(2*n+2)}" by simp - moreover have "prime (2::nat)" + moreover have "prime (2::nat)" using two_is_prime_nat by blast ultimately have "r 2" using r_def by simp then show ?thesis by (rule exI) qed moreover have "prime_above n = (SOME x. r x)" by (simp add:prime_above_def r_def) ultimately have a:"r (prime_above n)" using someI_ex by metis show "prime (prime_above n)" using a unfolding r_def by blast show "prime_above n \ {n..(2*n+2)}" using a unfolding r_def by blast qed lemma prime_above_min: "prime_above n \ 2" - using prime_above_prime + using prime_above_prime by (simp add: prime_ge_2_nat) lemma prime_above_lower_bound: "prime_above n \ n" using prime_above_range by simp lemma prime_above_upper_bound: "prime_above n \ 2*n+2" using prime_above_range by simp end diff --git a/thys/Frequency_Moments/K_Smallest.thy b/thys/Frequency_Moments/K_Smallest.thy --- a/thys/Frequency_Moments/K_Smallest.thy +++ b/thys/Frequency_Moments/K_Smallest.thy @@ -1,385 +1,385 @@ section \Ranks, $k$ smallest element and elements\ theory K_Smallest - imports + imports Frequency_Moments_Preliminary_Results Interpolation_Polynomials_HOL_Algebra.Interpolation_Polynomial_Cardinalities begin text \This section contains definitions and results for the selection of the $k$ smallest elements, the $k$-th smallest element, rank of an element in an ordered set.\ -definition rank_of :: "'a :: linorder \ 'a set \ nat" where "rank_of x S = card {y \ S. y < x}" +definition rank_of :: "'a :: linorder \ 'a set \ nat" where "rank_of x S = card {y \ S. y < x}" text \The function @{term "rank_of"} returns the rank of an element within a set.\ lemma rank_mono: assumes "finite S" shows "x \ y \ rank_of x S \ rank_of y S" unfolding rank_of_def using assms by (intro card_mono, auto) lemma rank_mono_2: assumes "finite S" shows "S' \ S \ rank_of x S' \ rank_of x S" unfolding rank_of_def using assms by (intro card_mono, auto) lemma rank_mono_commute: assumes "finite S" assumes "S \ T" assumes "strict_mono_on T f" assumes "x \ T" shows "rank_of x S = rank_of (f x) (f ` S)" proof - have a: "inj_on f T" by (metis assms(3) strict_mono_on_imp_inj_on) have "rank_of (f x) (f ` S) = card (f ` {y \ S. f y < f x})" unfolding rank_of_def by (intro arg_cong[where f="card"], auto) also have "... = card (f ` {y \ S. y < x})" using assms by (intro arg_cong[where f="card"] arg_cong[where f="(`) f"]) (meson in_mono linorder_not_le strict_mono_onD strict_mono_on_leD set_eq_iff) also have "... = card {y \ S. y < x}" using assms by (intro card_image inj_on_subset[OF a], blast) also have "... = rank_of x S" by (simp add:rank_of_def) finally show ?thesis by simp qed definition least where "least k S = {y \ S. rank_of y S < k}" text \The function @{term "least"} returns the k smallest elements of a finite set.\ -lemma rank_strict_mono: +lemma rank_strict_mono: assumes "finite S" shows "strict_mono_on S (\x. rank_of x S)" proof - have "\x y. x \ S \ y \ S \ x < y \ rank_of x S < rank_of y S" - unfolding rank_of_def using assms + unfolding rank_of_def using assms by (intro psubset_card_mono, auto) thus ?thesis by (simp add:rank_of_def strict_mono_on_def) qed lemma rank_of_image: assumes "finite S" shows "(\x. rank_of x S) ` S = {0..x. x \ S \ card {y \ S. y < x} < card S" by (rule psubset_card_mono, metis assms, blast) thus "(\x. rank_of x S) ` S \ {0..x. rank_of x S) S" - by (metis strict_mono_on_imp_inj_on rank_strict_mono assms) + by (metis strict_mono_on_imp_inj_on rank_strict_mono assms) thus "card {0.. card ((\x. rank_of x S) ` S)" by (simp add:card_image) qed lemma card_least: assumes "finite S" shows "card (least k S) = min k (card S)" proof (cases "card S < k") case True - have "\t. rank_of t S \ card S" - unfolding rank_of_def using assms + have "\t. rank_of t S \ card S" + unfolding rank_of_def using assms by (intro card_mono, auto) - hence "\t. rank_of t S < k" + hence "\t. rank_of t S < k" by (metis True not_less_iff_gr_or_eq order_less_le_trans) hence "least k S = S" by (simp add:least_def) then show ?thesis using True by simp next case False hence a:"card S \ k" using leI by blast hence "card ((\x. rank_of x S) -` {0.. S) = card {0.. S" by (simp add:least_def) lemma least_mono_commute: assumes "finite S" assumes "strict_mono_on S f" shows "f ` least k S = least k (f ` S)" proof - - have a:"inj_on f S" + have a:"inj_on f S" using strict_mono_on_imp_inj_on[OF assms(2)] by simp have "card (least k (f ` S)) = min k (card (f ` S))" by (subst card_least, auto simp add:assms) also have "... = min k (card S)" by (subst card_image, metis a, auto) also have "... = card (least k S)" by (subst card_least, auto simp add:assms) also have "... = card (f ` least k S)" by (subst card_image[OF inj_on_subset[OF a]], simp_all add:least_def) finally have b: "card (least k (f ` S)) \ card (f ` least k S)" by simp have c: "f ` least k S \least k (f ` S)" - using assms by (intro image_subsetI) + using assms by (intro image_subsetI) (simp add:least_def rank_mono_commute[symmetric, where T="S"]) show ?thesis using b c assms by (intro card_seteq, simp_all add:least_def) qed lemma least_eq_iff: assumes "finite B" assumes "A \ B" assumes "\x. x \ B \ rank_of x B < k \ x \ A" shows "least k A = least k B" proof - have "least k B \ least k A" using assms rank_mono_2[OF assms(1,2)] order_le_less_trans - by (simp add:least_def, blast) + by (simp add:least_def, blast) moreover have "card (least k B) \ card (least k A)" using assms finite_subset[OF assms(2,1)] card_mono[OF assms(1,2)] by (simp add: card_least min_le_iff_disj) - moreover have "finite (least k A)" + moreover have "finite (least k A)" using finite_subset least_subset assms(1,2) by metis ultimately show ?thesis by (intro card_seteq[symmetric], simp_all) qed -lemma least_insert: +lemma least_insert: assumes "finite S" shows "least k (insert x (least k S)) = least k (insert x S)" (is "?lhs = ?rhs") proof (rule least_eq_iff) show "finite (insert x S)" using assms(1) by simp show "insert x (least k S) \ insert x S" using least_subset by blast show "y \ insert x (least k S)" if a: "y \ insert x S" and b: "rank_of y (insert x S) < k" for y proof - have "rank_of y S \ rank_of y (insert x S)" using assms by (intro rank_mono_2, auto) also have "... < k" using b by simp finally have "rank_of y S < k" by simp - hence "y = x \ (y \ S \ rank_of y S < k)" + hence "y = x \ (y \ S \ rank_of y S < k)" using a by simp thus ?thesis by (simp add:least_def) qed qed definition count_le where "count_le x M = size {#y \# M. y \ x#}" definition count_less where "count_less x M = size {#y \# M. y < x#}" definition nth_mset :: "nat \ ('a :: linorder) multiset \ 'a" where "nth_mset k M = sorted_list_of_multiset M ! k" lemma nth_mset_bound_left: assumes "k < size M" assumes "count_less x M \ k" shows "x \ nth_mset k M" proof (rule ccontr) define xs where "xs = sorted_list_of_multiset M" have s_xs: "sorted xs" by (simp add:xs_def sorted_sorted_list_of_multiset) have l_xs: "k < length xs" - using assms(1) by (simp add:xs_def size_mset[symmetric]) + using assms(1) by (simp add:xs_def size_mset[symmetric]) have M_xs: "M = mset xs" by (simp add:xs_def) hence a:"\i. i \ k \ xs ! i \ xs ! k" using s_xs l_xs sorted_iff_nth_mono by blast assume "\(x \ nth_mset k M)" hence "x > nth_mset k M" by simp hence b:"x > xs ! k" by (simp add:nth_mset_def xs_def[symmetric]) have "k < card {0..k}" by simp also have "... \ card {i. i < length xs \ xs ! i < x}" - using a b l_xs order_le_less_trans + using a b l_xs order_le_less_trans by (intro card_mono subsetI) auto also have "... = length (filter (\y. y < x) xs)" by (subst length_filter_conv_card, simp) also have "... = size (mset (filter (\y. y < x) xs))" by (subst size_mset, simp) also have "... = count_less x M" by (simp add:count_less_def M_xs) also have "... \ k" using assms by simp finally show "False" by simp qed lemma nth_mset_bound_left_excl: assumes "k < size M" assumes "count_le x M \ k" shows "x < nth_mset k M" proof (rule ccontr) define xs where "xs = sorted_list_of_multiset M" have s_xs: "sorted xs" by (simp add:xs_def sorted_sorted_list_of_multiset) - have l_xs: "k < length xs" - using assms(1) by (simp add:xs_def size_mset[symmetric]) + have l_xs: "k < length xs" + using assms(1) by (simp add:xs_def size_mset[symmetric]) have M_xs: "M = mset xs" by (simp add:xs_def) hence a:"\i. i \ k \ xs ! i \ xs ! k" using s_xs l_xs sorted_iff_nth_mono by blast assume "\(x < nth_mset k M)" hence "x \ nth_mset k M" by simp hence b:"x \ xs ! k" by (simp add:nth_mset_def xs_def[symmetric]) have "k+1 \ card {0..k}" by simp also have "... \ card {i. i < length xs \ xs ! i \ xs ! k}" using a b l_xs order_le_less_trans by (intro card_mono subsetI, auto) also have "... \ card {i. i < length xs \ xs ! i \ x}" using b by (intro card_mono subsetI, auto) also have "... = length (filter (\y. y \ x) xs)" by (subst length_filter_conv_card, simp) also have "... = size (mset (filter (\y. y \ x) xs))" by (subst size_mset, simp) also have "... = count_le x M" by (simp add:count_le_def M_xs) also have "... \ k" using assms by simp finally show "False" by simp qed lemma nth_mset_bound_right: assumes "k < size M" assumes "count_le x M > k" shows "nth_mset k M \ x" proof (rule ccontr) define xs where "xs = sorted_list_of_multiset M" have s_xs: "sorted xs" by (simp add:xs_def sorted_sorted_list_of_multiset) - have l_xs: "k < length xs" - using assms(1) by (simp add:xs_def size_mset[symmetric]) + have l_xs: "k < length xs" + using assms(1) by (simp add:xs_def size_mset[symmetric]) have M_xs: "M = mset xs" by (simp add:xs_def) assume "\(nth_mset k M \ x)" hence "x < nth_mset k M" by simp - hence "x < xs ! k" + hence "x < xs ! k" by (simp add:nth_mset_def xs_def[symmetric]) hence a:"\i. i < length xs \ xs ! i \ x \ i < k" using s_xs l_xs sorted_iff_nth_mono leI by fastforce have "count_le x M = size (mset (filter (\y. y \ x) xs))" by (simp add:count_le_def M_xs) also have "... = length (filter (\y. y \ x) xs)" by (subst size_mset, simp) also have "... = card {i. i < length xs \ xs ! i \ x}" by (subst length_filter_conv_card, simp) also have "... \ card {i. i < k}" using a by (intro card_mono subsetI, auto) also have "... = k" by simp finally have "count_le x M \ k" by simp thus "False" using assms by simp qed lemma nth_mset_commute_mono: assumes "mono f" assumes "k < size M" shows "f (nth_mset k M) = nth_mset k (image_mset f M)" proof - have a:"k < length (sorted_list_of_multiset M)" by (metis assms(2) mset_sorted_list_of_multiset size_mset) show ?thesis using a by (simp add:nth_mset_def sorted_list_of_multiset_image_commute[OF assms(1)]) -qed +qed -lemma nth_mset_max: +lemma nth_mset_max: assumes "size A > k" assumes "\x. x \ nth_mset k A \ count A x \ 1" shows "nth_mset k A = Max (least (k+1) (set_mset A))" and "card (least (k+1) (set_mset A)) = k+1" proof - define xs where "xs = sorted_list_of_multiset A" have k_bound: "k < length xs" unfolding xs_def - by (metis size_mset mset_sorted_list_of_multiset assms(1)) + by (metis size_mset mset_sorted_list_of_multiset assms(1)) have A_def: "A = mset xs" by (simp add:xs_def) have s_xs: "sorted xs" by (simp add:xs_def sorted_sorted_list_of_multiset) have "\x. x \ xs ! k \ count A x \ Suc 0" using assms(2) by (simp add:xs_def[symmetric] nth_mset_def) - hence no_col: "\x. x \ xs ! k \ count_list xs x \ 1" - by (simp add:A_def count_mset) + hence no_col: "\x. x \ xs ! k \ count_list xs x \ 1" + by (simp add:A_def count_mset) have inj_xs: "inj_on (\k. xs ! k) {0..k}" by (rule inj_onI, simp) (metis (full_types) count_list_ge_2_iff k_bound no_col le_neq_implies_less linorder_not_le order_le_less_trans s_xs sorted_iff_nth_mono) have "\y. y < length xs \ rank_of (xs ! y) (set xs) < k+1 \ y < k+1" proof (rule ccontr) fix y assume b:"y < length xs" assume "\y < k +1" hence a:"k + 1 \ y" by simp have d:"Suc k < length xs" using a b by simp - have "k+1 = card ((!) xs ` {0..k})" + have "k+1 = card ((!) xs ` {0..k})" by (subst card_image[OF inj_xs], simp) also have "... \ rank_of (xs ! (k+1)) (set xs)" unfolding rank_of_def using k_bound - by (intro card_mono image_subsetI conjI, simp_all) (metis count_list_ge_2_iff no_col not_le le_imp_less_Suc s_xs + by (intro card_mono image_subsetI conjI, simp_all) (metis count_list_ge_2_iff no_col not_le le_imp_less_Suc s_xs sorted_iff_nth_mono d order_less_le) also have "... \ rank_of (xs ! y) (set xs)" unfolding rank_of_def by (intro card_mono subsetI, simp_all) (metis Suc_eq_plus1 a b s_xs order_less_le_trans sorted_iff_nth_mono) also assume "... < k+1" finally show "False" by force qed moreover have "rank_of (xs ! y) (set xs) < k+1" if a:"y < k + 1" for y proof - have "rank_of (xs ! y) (set xs) \ card ((\k. xs ! k) ` {k. k < length xs \ xs ! k < xs ! y})" unfolding rank_of_def by (intro card_mono subsetI, simp) (metis (no_types, lifting) imageI in_set_conv_nth mem_Collect_eq) also have "... \ card {k. k < length xs \ xs ! k < xs ! y}" by (rule card_image_le, simp) also have "... \ card {k. k < y}" by (intro card_mono subsetI, simp_all add:not_less) (metis sorted_iff_nth_mono s_xs linorder_not_less) also have "... = y" by simp also have "... < k + 1" using a by simp finally show "rank_of (xs ! y) (set xs) < k+1" by simp qed ultimately have rank_conv: "\y. y < length xs \ rank_of (xs ! y) (set xs) < k+1 \ y < k+1" by blast have "y \ xs ! k" if a:"y \ least (k+1) (set xs)" for y proof - have "y \ set xs" using a least_subset by blast then obtain i where i_bound: "i < length xs" and y_def: "y = xs ! i" using in_set_conv_nth by metis hence "rank_of (xs ! i) (set xs) < k+1" using a y_def i_bound by (simp add: least_def) hence "i < k+1" using rank_conv i_bound by blast hence "i \ k" by linarith hence "xs ! i \ xs ! k" using s_xs i_bound k_bound sorted_nth_mono by blast thus "y \ xs ! k" using y_def by simp qed moreover have "xs ! k \ least (k+1) (set xs)" using k_bound rank_conv by (simp add:least_def) ultimately have "Max (least (k+1) (set xs)) = xs ! k" by (intro Max_eqI finite_subset[OF least_subset], auto) - hence "nth_mset k A = Max (K_Smallest.least (Suc k) (set xs))" + hence "nth_mset k A = Max (K_Smallest.least (Suc k) (set xs))" by (simp add:nth_mset_def xs_def[symmetric]) also have "... = Max (least (k+1) (set_mset A))" by (simp add:A_def) finally show "nth_mset k A = Max (least (k+1) (set_mset A))" by simp - have "k + 1 = card ((\i. xs ! i) ` {0..k})" - by (subst card_image[OF inj_xs], simp) + have "k + 1 = card ((\i. xs ! i) ` {0..k})" + by (subst card_image[OF inj_xs], simp) also have "... \ card (least (k+1) (set xs))" using rank_conv k_bound by (intro card_mono image_subsetI finite_subset[OF least_subset], simp_all add:least_def) finally have "card (least (k+1) (set xs)) \ k+1" by simp moreover have "card (least (k+1) (set xs)) \ k+1" by (subst card_least, simp, simp) ultimately have "card (least (k+1) (set xs)) = k+1" by simp thus "card (least (k+1) (set_mset A)) = k+1" by (simp add:A_def) qed end diff --git a/thys/Frequency_Moments/Landau_Ext.thy b/thys/Frequency_Moments/Landau_Ext.thy --- a/thys/Frequency_Moments/Landau_Ext.thy +++ b/thys/Frequency_Moments/Landau_Ext.thy @@ -1,243 +1,243 @@ section \Landau Symbols\ theory Landau_Ext - imports + imports "HOL-Library.Landau_Symbols" "HOL.Topological_Spaces" begin text \This section contains results about Landau Symbols in addition to "HOL-Library.Landau".\ lemma landau_sum: - assumes "eventually (\x. g1 x \ (0::real)) F" - assumes "eventually (\x. g2 x \ 0) F" + assumes "eventually (\x. g1 x \ (0::real)) F" + assumes "eventually (\x. g2 x \ 0) F" assumes "f1 \ O[F](g1)" assumes "f2 \ O[F](g2)" shows "(\x. f1 x + f2 x) \ O[F](\x. g1 x + g2 x)" proof - obtain c1 where a1: "c1 > 0" and b1: "eventually (\x. abs (f1 x) \ c1 * abs (g1 x)) F" using assms(3) by (simp add:bigo_def, blast) obtain c2 where a2: "c2 > 0" and b2: "eventually (\x. abs (f2 x) \ c2 * abs (g2 x)) F" using assms(4) by (simp add:bigo_def, blast) have "eventually (\x. abs (f1 x + f2 x) \ (max c1 c2) * abs (g1 x + g2 x)) F" proof (rule eventually_mono[OF eventually_conj[OF b1 eventually_conj[OF b2 eventually_conj[OF assms(1,2)]]]]) fix x assume a: "\f1 x\ \ c1 * \g1 x\ \ \f2 x\ \ c2 * \g2 x\ \ 0 \ g1 x \ 0 \ g2 x" have "\f1 x + f2 x\ \ \f1 x \ + \f2 x\" using abs_triangle_ineq by blast also have "... \ c1 * \g1 x\ + c2 * \g2 x\" using a add_mono by blast - also have "... \ max c1 c2 * \g1 x\ + max c1 c2 * \g2 x\" + also have "... \ max c1 c2 * \g1 x\ + max c1 c2 * \g2 x\" by (intro add_mono mult_right_mono) auto also have "... = max c1 c2 * (\g1 x\ + \g2 x\)" by (simp add:algebra_simps) also have "... \ max c1 c2 * (\g1 x + g2 x\)" using a a1 a2 by (intro mult_left_mono) auto finally show "\f1 x + f2 x\ \ max c1 c2 * \g1 x + g2 x\" by (simp add:algebra_simps) qed hence " 0 < max c1 c2 \ (\\<^sub>F x in F. \f1 x + f2 x\ \ max c1 c2 * \g1 x + g2 x\)" using a1 a2 by linarith thus ?thesis - by (simp add: bigo_def, blast) + by (simp add: bigo_def, blast) qed lemma landau_sum_1: - assumes "eventually (\x. g1 x \ (0::real)) F" - assumes "eventually (\x. g2 x \ 0) F" + assumes "eventually (\x. g1 x \ (0::real)) F" + assumes "eventually (\x. g2 x \ 0) F" assumes "f \ O[F](g1)" shows "f \ O[F](\x. g1 x + g2 x)" proof - have "f = (\x. f x + 0)" by simp also have "... \ O[F](\x. g1 x + g2 x)" using assms zero_in_bigo by (intro landau_sum) finally show ?thesis by simp qed lemma landau_sum_2: - assumes "eventually (\x. g1 x \ (0::real)) F" - assumes "eventually (\x. g2 x \ 0) F" + assumes "eventually (\x. g1 x \ (0::real)) F" + assumes "eventually (\x. g2 x \ 0) F" assumes "f \ O[F](g2)" shows "f \ O[F](\x. g1 x + g2 x)" proof - have "f = (\x. 0 + f x)" by simp also have "... \ O[F](\x. g1 x + g2 x)" using assms zero_in_bigo by (intro landau_sum) finally show ?thesis by simp qed lemma landau_ln_3: - assumes "eventually (\x. (1::real) \ f x) F" - assumes "f \ O[F](g)" - shows "(\x. ln (f x)) \ O[F](g)" + assumes "eventually (\x. (1::real) \ f x) F" + assumes "f \ O[F](g)" + shows "(\x. ln (f x)) \ O[F](g)" proof - have "1 \ x \ \ln x\ \ \x\" for x :: real using ln_bound by auto hence "(\x. ln (f x)) \ O[F](f)" by (intro landau_o.big_mono eventually_mono[OF assms(1)]) simp thus ?thesis using assms(2) landau_o.big_trans by blast qed lemma landau_ln_2: assumes "a > (1::real)" - assumes "eventually (\x. 1 \ f x) F" - assumes "eventually (\x. a \ g x) F" - assumes "f \ O[F](g)" - shows "(\x. ln (f x)) \ O[F](\x. ln (g x))" + assumes "eventually (\x. 1 \ f x) F" + assumes "eventually (\x. a \ g x) F" + assumes "f \ O[F](g)" + shows "(\x. ln (f x)) \ O[F](\x. ln (g x))" proof - obtain c where a: "c > 0" and b: "eventually (\x. abs (f x) \ c * abs (g x)) F" using assms(4) by (simp add:bigo_def, blast) define d where "d = 1 + (max 0 (ln c)) / ln a" have d:"eventually (\x. abs (ln (f x)) \ d * abs (ln (g x))) F" proof (rule eventually_mono[OF eventually_conj[OF b eventually_conj[OF assms(3,2)]]]) fix x - assume c:"\f x\ \ c * \g x\ \ a \ g x \ 1 \ f x" + assume c:"\f x\ \ c * \g x\ \ a \ g x \ 1 \ f x" have "abs (ln (f x)) = ln (f x)" by (subst abs_of_nonneg, rule ln_ge_zero, metis c, simp) also have "... \ ln (c * abs (g x))" - using c assms(1) mult_pos_pos[OF a] by auto + using c assms(1) mult_pos_pos[OF a] by auto also have "... \ ln c + ln (abs (g x))" using c assms(1) by (simp add: ln_mult[OF a]) also have "... \ (d-1)*ln a + ln (g x)" using assms(1) c by (intro add_mono iffD2[OF ln_le_cancel_iff], simp_all add:d_def) also have "... \ (d-1)* ln (g x) + ln (g x)" using assms(1) c by (intro add_mono mult_left_mono iffD2[OF ln_le_cancel_iff], simp_all add:d_def) also have "... = d * ln (g x)" by (simp add:algebra_simps) also have "... = d * abs (ln (g x))" using c assms(1) by auto finally show "abs (ln (f x)) \ d * abs (ln (g x))" by simp qed hence "\\<^sub>F x in F. \ln (f x)\ \ d * \ln (g x)\" by simp moreover have "0 < d" unfolding d_def using assms(1) by (intro add_pos_nonneg divide_nonneg_pos, auto) ultimately show ?thesis by (auto simp:bigo_def) qed lemma landau_real_nat: fixes f :: "'a \ int" assumes "(\x. of_int (f x)) \ O[F](g)" shows "(\x. real (nat (f x))) \ O[F](g)" proof - obtain c where a: "c > 0" and b: "eventually (\x. abs (of_int (f x)) \ c * abs (g x)) F" using assms(1) by (simp add:bigo_def, blast) have "\\<^sub>F x in F. real (nat (f x)) \ c * \g x\" by (rule eventually_mono[OF b], simp) thus ?thesis using a by (auto simp:bigo_def) qed lemma landau_ceil: assumes "(\_. 1) \ O[F'](g)" assumes "f \ O[F'](g)" shows "(\x. real_of_int \f x\) \ O[F'](g)" proof - have "(\x. real_of_int \f x\) \ O[F'](\x. 1 + abs (f x))" by (intro landau_o.big_mono always_eventually allI, simp, linarith) also have "(\x. 1 + abs(f x)) \ O[F'](g)" using assms(2) by (intro sum_in_bigo assms(1), auto) finally show ?thesis by simp qed lemma landau_rat_ceil: assumes "(\_. 1) \ O[F'](g)" assumes "(\x. real_of_rat (f x)) \ O[F'](g)" shows "(\x. real_of_int \f x\) \ O[F'](g)" proof - have a:"\real_of_int \x\\ \ 1 + real_of_rat \x\" for x :: rat proof (cases "x \ 0") case True - then show ?thesis + then show ?thesis by (simp, metis add.commute of_int_ceiling_le_add_one of_rat_ceiling) next case False have "real_of_rat x - 1 \ real_of_rat x" by simp - also have "... \ real_of_int \x\" + also have "... \ real_of_int \x\" by (metis ceiling_correct of_rat_ceiling) finally have " real_of_rat (x)-1 \ real_of_int \x\" by simp hence "- real_of_int \x\ \ 1 + real_of_rat (- x)" by (simp add: of_rat_minus) - then show ?thesis using False by simp + then show ?thesis using False by simp qed - have "(\x. real_of_int \f x\) \ O[F'](\x. 1 + abs (real_of_rat (f x)))" + have "(\x. real_of_int \f x\) \ O[F'](\x. 1 + abs (real_of_rat (f x)))" using a by (intro landau_o.big_mono always_eventually allI, simp) also have "(\x. 1 + abs (real_of_rat (f x))) \ O[F'](g)" using assms by (intro sum_in_bigo assms(1), subst landau_o.big.abs_in_iff, simp) finally show ?thesis by simp qed lemma landau_nat_ceil: assumes "(\_. 1) \ O[F'](g)" assumes "f \ O[F'](g)" shows "(\x. real (nat \f x\)) \ O[F'](g)" using assms by (intro landau_real_nat landau_ceil, auto) lemma eventually_prod1': assumes "B \ bot" assumes " (\\<^sub>F x in A. P x)" shows "(\\<^sub>F x in A \\<^sub>F B. P (fst x))" proof - have "(\\<^sub>F x in A \\<^sub>F B. P (fst x)) = (\\<^sub>F (x,y) in A \\<^sub>F B. P x)" by (simp add:case_prod_beta') also have "... = (\\<^sub>F x in A. P x)" by (subst eventually_prod1[OF assms(1)], simp) finally show ?thesis using assms(2) by simp qed lemma eventually_prod2': assumes "A \ bot" assumes " (\\<^sub>F x in B. P x)" shows "(\\<^sub>F x in A \\<^sub>F B. P (snd x))" proof - have "(\\<^sub>F x in A \\<^sub>F B. P (snd x)) = (\\<^sub>F (x,y) in A \\<^sub>F B. P y)" by (simp add:case_prod_beta') also have "... = (\\<^sub>F x in B. P x)" by (subst eventually_prod2[OF assms(1)], simp) finally show ?thesis using assms(2) by simp qed lemma sequentially_inf: "\\<^sub>F x in sequentially. n \ real x" by (meson eventually_at_top_linorder nat_ceiling_le_eq) instantiation rat :: linorder_topology begin definition open_rat :: "rat set \ bool" where "open_rat = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" instance by standard (rule open_rat_def) end lemma inv_at_right_0_inf: "\\<^sub>F x in at_right 0. c \ 1 / real_of_rat x" proof - have a:" c \ 1 / real_of_rat x" if b:" x \ {0<..<1 / rat_of_int (max \c\ 1)}" for x proof - have "c * real_of_rat x \ real_of_int (max \c\ 1) * real_of_rat x" using b by (intro mult_right_mono, linarith, auto) also have "... < real_of_int (max \c\ 1) * real_of_rat (1/rat_of_int (max \c\ 1) )" using b by (intro mult_strict_left_mono iffD2[OF of_rat_less], auto) also have "... \ 1" by (simp add:of_rat_divide) finally have "c * real_of_rat x \ 1" by simp moreover have "0 < real_of_rat x" using b by simp ultimately show ?thesis by (subst pos_le_divide_eq, auto) qed show ?thesis using a by (intro eventually_at_rightI[where b="1/rat_of_int (max \c\ 1)"], simp_all) qed end \ No newline at end of file diff --git a/thys/Frequency_Moments/Product_PMF_Ext.thy b/thys/Frequency_Moments/Product_PMF_Ext.thy --- a/thys/Frequency_Moments/Product_PMF_Ext.thy +++ b/thys/Frequency_Moments/Product_PMF_Ext.thy @@ -1,210 +1,30 @@ section \Indexed Products of Probability Mass Functions\ theory Product_PMF_Ext - imports Main Probability_Ext "HOL-Probability.Product_PMF" Universal_Hash_Families.Universal_Hash_Families_More_Independent_Families + imports + Probability_Ext + Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF begin -hide_const "Isolated.discrete" - -text \This section introduces a restricted version of @{term "Pi_pmf"} where the default value is @{term "undefined"} -and contains some additional results about that case in addition to @{theory "HOL-Probability.Product_PMF"}\ - -abbreviation prod_pmf where "prod_pmf I M \ Pi_pmf I undefined M" - -lemma pmf_prod_pmf: - assumes "finite I" - shows "pmf (prod_pmf I M) x = (if x \ extensional I then \i \ I. (pmf (M i)) (x i) else 0)" - by (simp add: pmf_Pi[OF assms(1)] extensional_def) - -lemma PiE_defaut_undefined_eq: "PiE_dflt I undefined M = PiE I M" - by (simp add:PiE_dflt_def PiE_def extensional_def Pi_def set_eq_iff) blast - -lemma set_prod_pmf: - assumes "finite I" - shows "set_pmf (prod_pmf I M) = PiE I (set_pmf \ M)" - by (simp add:set_Pi_pmf[OF assms] PiE_defaut_undefined_eq) - -text \A more general version of @{thm [source] "measure_Pi_pmf_Pi"}.\ -lemma prob_prod_pmf': - assumes "finite I" - assumes "J \ I" - shows "measure (measure_pmf (Pi_pmf I d M)) (Pi J A) = (\ i \ J. measure (M i) (A i))" -proof - - have a:"Pi J A = Pi I (\i. if i \ J then A i else UNIV)" - using assms by (simp add:Pi_def set_eq_iff, blast) - show ?thesis - using assms - by (simp add:if_distrib a measure_Pi_pmf_Pi[OF assms(1)] prod.If_cases[OF assms(1)] Int_absorb1) -qed - -lemma prob_prod_pmf_slice: - assumes "finite I" - assumes "i \ I" - shows "measure (measure_pmf (prod_pmf I M)) {\. P (\ i)} = measure (M i) {\. P \}" - using prob_prod_pmf'[OF assms(1), where J="{i}" and M="M" and A="\_. Collect P"] - by (simp add:assms Pi_def) - - -definition restrict_dfl where "restrict_dfl f A d = (\x. if x \ A then f x else d)" - -lemma pi_pmf_decompose: - assumes "finite I" - shows "Pi_pmf I d M = map_pmf (\\. restrict_dfl (\i. \ (f i) i) I d) (Pi_pmf (f ` I) (\_. d) (\j. Pi_pmf (f -` {j} \ I) d M))" -proof - - have fin_F_I:"finite (f ` I)" using assms by blast - - have "finite I \ ?thesis" - using fin_F_I - proof (induction "f ` I" arbitrary: I rule:finite_induct) - case empty - then show ?case by (simp add:restrict_dfl_def) - next - case (insert x F) - have a: "(f -` {x} \ I) \ (f -` F \ I) = I" - using insert(4) by blast - have b: "F = f ` (f -` F \ I) " using insert(2,4) - by (auto simp add:set_eq_iff image_def vimage_def) - have c: "finite (f -` F \ I)" using insert by blast - have d: "\j. j \ F \ (f -` {j} \ (f -` F \ I)) = (f -` {j} \ I)" - using insert(4) by blast - - have " Pi_pmf I d M = Pi_pmf ((f -` {x} \ I) \ (f -` F \ I)) d M" - by (simp add:a) - also have "... = map_pmf (\(g, h) i. if i \ f -` {x} \ I then g i else h i) - (pair_pmf (Pi_pmf (f -` {x} \ I) d M) (Pi_pmf (f -` F \ I) d M))" - using insert by (subst Pi_pmf_union) auto - also have "... = map_pmf (\(g,h) i. if f i = x \ i \ I then g i else if f i \ F \ i \ I then h (f i) i else d) - (pair_pmf (Pi_pmf (f -` {x} \ I) d M) (Pi_pmf F (\_. d) (\j. Pi_pmf (f -` {j} \ (f -` F \ I)) d M)))" - by (simp add:insert(3)[OF b c] map_pmf_comp case_prod_beta' apsnd_def map_prod_def - pair_map_pmf2 b[symmetric] restrict_dfl_def) (metis fst_conv snd_conv) - also have "... = map_pmf (\(g,h) i. if i \ I then (h(x := g)) (f i) i else d) - (pair_pmf (Pi_pmf (f -` {x} \ I) d M) (Pi_pmf F (\_. d) (\j. Pi_pmf (f -` {j} \ I) d M)))" - using insert(4) d - by (intro arg_cong2[where f="map_pmf"] ext) (auto simp add:case_prod_beta' cong:Pi_pmf_cong) - also have "... = map_pmf (\\ i. if i \ I then \ (f i) i else d) (Pi_pmf (insert x F) (\_. d) (\j. Pi_pmf (f -` {j} \ I) d M))" - by (simp add:Pi_pmf_insert[OF insert(1,2)] map_pmf_comp case_prod_beta') - finally show ?case by (simp add:insert(4) restrict_dfl_def) - qed - thus ?thesis using assms by blast -qed - -lemma restrict_dfl_iter: "restrict_dfl (restrict_dfl f I d) J d = restrict_dfl f (I \ J) d" - by (rule ext, simp add:restrict_dfl_def) +text \The following aliases are here to prevent possible merge-conflicts. The lemmas have been +moved to @{theory Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF}.\ -lemma indep_vars_restrict': - assumes "finite I" - shows "prob_space.indep_vars (Pi_pmf I d M) (\_. discrete) (\i \. restrict_dfl \ (f -` {i} \ I) d) (f ` I)" -proof - - let ?Q = "(Pi_pmf (f ` I) (\_. d) (\i. Pi_pmf (I \ f -` {i}) d M))" - have a:"prob_space.indep_vars ?Q (\_. discrete) (\i \. \ i) (f ` I)" - using assms by (intro indep_vars_Pi_pmf, blast) - have b: "AE x in measure_pmf ?Q. \i\f ` I. x i = restrict_dfl (\i. x (f i) i) (I \ f -` {i}) d" - using assms - by (auto simp add:PiE_dflt_def restrict_dfl_def AE_measure_pmf_iff set_Pi_pmf comp_def Int_commute) - have "prob_space.indep_vars ?Q (\_. discrete) (\i x. restrict_dfl (\i. x (f i) i) (I \ f -` {i}) d) (f ` I)" - by (rule prob_space.indep_vars_cong_AE[OF prob_space_measure_pmf b a], simp) - thus ?thesis - using prob_space_measure_pmf - by (auto intro!:prob_space.indep_vars_distr simp:pi_pmf_decompose[OF assms, where f="f"] - map_pmf_rep_eq comp_def restrict_dfl_iter Int_commute) -qed - -lemma indep_vars_restrict_intro': - assumes "finite I" - assumes "\i \. i \ J \ X' i \ = X' i (restrict_dfl \ (f -` {i} \ I) d)" - assumes "J = f ` I" - assumes "\\ i. i \ J \ X' i \ \ space (M' i)" - shows "prob_space.indep_vars (measure_pmf (Pi_pmf I d p)) M' (\i \. X' i \) J" -proof - - define M where "M \ measure_pmf (Pi_pmf I d p)" - interpret prob_space "M" - using M_def prob_space_measure_pmf by blast - have "indep_vars (\_. discrete) (\i x. restrict_dfl x (f -` {i} \ I) d) (f ` I)" - unfolding M_def by (rule indep_vars_restrict'[OF assms(1)]) - hence "indep_vars M' (\i \. X' i (restrict_dfl \ ( f -` {i} \ I) d)) (f ` I)" - using assms(4) - by (intro indep_vars_compose2[where Y="X'" and N="M'" and M'="\_. discrete"]) (auto simp:assms(3)) - hence "indep_vars M' (\i \. X' i \) (f ` I)" - using assms(2)[symmetric] - by (simp add:assms(3) cong:indep_vars_cong) - thus ?thesis - unfolding M_def using assms(3) by simp -qed - -lemma - fixes f :: "'b \ ('c :: {second_countable_topology,banach,real_normed_field})" - assumes "finite I" - assumes "i \ I" - assumes "integrable (measure_pmf (M i)) f" - shows integrable_Pi_pmf_slice: "integrable (Pi_pmf I d M) (\x. f (x i))" - and expectation_Pi_pmf_slice: "integral\<^sup>L (Pi_pmf I d M) (\x. f (x i)) = integral\<^sup>L (M i) f" -proof - - have a:"distr (Pi_pmf I d M) (M i) (\\. \ i) = distr (Pi_pmf I d M) discrete (\\. \ i)" - by (rule distr_cong, auto) - - have b: "measure_pmf.random_variable (M i) borel f" - using assms(3) by simp - - have c:"integrable (distr (Pi_pmf I d M) (M i) (\\. \ i)) f" - using assms(1,2,3) - by (subst a, subst map_pmf_rep_eq[symmetric], subst Pi_pmf_component, auto) +abbreviation prod_pmf where "prod_pmf \ Universal_Hash_Families_More_Product_PMF.prod_pmf" +abbreviation restrict_dfl where "restrict_dfl \ Universal_Hash_Families_More_Product_PMF.restrict_dfl" - show "integrable (Pi_pmf I d M) (\x. f (x i))" - by (rule integrable_distr[where f="f" and M'="M i"]) (auto intro: c) - - have "integral\<^sup>L (Pi_pmf I d M) (\x. f (x i)) = integral\<^sup>L (distr (Pi_pmf I d M) (M i) (\\. \ i)) f" - using b by (intro integral_distr[symmetric], auto) - also have "... = integral\<^sup>L (map_pmf (\\. \ i) (Pi_pmf I d M)) f" - by (subst a, subst map_pmf_rep_eq[symmetric], simp) - also have "... = integral\<^sup>L (M i) f" - using assms(1,2) by (simp add: Pi_pmf_component) - finally show "integral\<^sup>L (Pi_pmf I d M) (\x. f (x i)) = integral\<^sup>L (M i) f" by simp -qed - -text \This is an improved version of @{thm [source] "expectation_prod_Pi_pmf"}. -It works for general normed fields instead of non-negative real functions .\ +lemmas pmf_prod_pmf = pmf_prod_pmf +lemmas PiE_defaut_undefined_eq = PiE_defaut_undefined_eq +lemmas set_prod_pmf = set_prod_pmf +lemmas prob_prod_pmf' = prob_prod_pmf' +lemmas prob_prod_pmf_slice = prob_prod_pmf_slice +lemmas pi_pmf_decompose = pi_pmf_decompose +lemmas restrict_dfl_iter = restrict_dfl_iter +lemmas indep_vars_restrict' = indep_vars_restrict' +lemmas indep_vars_restrict_intro' = indep_vars_restrict_intro' +lemmas integrable_Pi_pmf_slice = integrable_Pi_pmf_slice +lemmas expectation_Pi_pmf_slice = expectation_Pi_pmf_slice +lemmas expectation_prod_Pi_pmf = expectation_prod_Pi_pmf +lemmas variance_prod_pmf_slice = variance_prod_pmf_slice +lemmas Pi_pmf_bind_return = Pi_pmf_bind_return -lemma expectation_prod_Pi_pmf: - fixes f :: "'a \ 'b \ ('c :: {second_countable_topology,banach,real_normed_field})" - assumes "finite I" - assumes "\i. i \ I \ integrable (measure_pmf (M i)) (f i)" - shows "integral\<^sup>L (Pi_pmf I d M) (\x. (\i \ I. f i (x i))) = (\ i \ I. integral\<^sup>L (M i) (f i))" -proof - - have a: "prob_space.indep_vars (measure_pmf (Pi_pmf I d M)) (\_. borel) (\i \. f i (\ i)) I" - by (intro prob_space.indep_vars_compose2[where Y="f" and M'="\_. discrete"] - prob_space_measure_pmf indep_vars_Pi_pmf assms(1)) auto - have "integral\<^sup>L (Pi_pmf I d M) (\x. (\i \ I. f i (x i))) = (\ i \ I. integral\<^sup>L (Pi_pmf I d M) (\x. f i (x i)))" - by (intro prob_space.indep_vars_lebesgue_integral prob_space_measure_pmf assms(1,2) - a integrable_Pi_pmf_slice) auto - also have "... = (\ i \ I. integral\<^sup>L (M i) (f i))" - by (intro prod.cong expectation_Pi_pmf_slice assms(1,2)) auto - finally show ?thesis by simp -qed - -lemma variance_prod_pmf_slice: - fixes f :: "'a \ real" - assumes "i \ I" "finite I" - assumes "integrable (measure_pmf (M i)) (\\. f \^2)" - shows "prob_space.variance (Pi_pmf I d M) (\\. f (\ i)) = prob_space.variance (M i) f" -proof - - have a:"integrable (measure_pmf (M i)) f" - using assms(3) measure_pmf.square_integrable_imp_integrable by auto - have b:" integrable (measure_pmf (Pi_pmf I d M)) (\x. (f (x i))\<^sup>2)" - by (rule integrable_Pi_pmf_slice[OF assms(2) assms(1)], metis assms(3)) - have c:" integrable (measure_pmf (Pi_pmf I d M)) (\x. (f (x i)))" - by (rule integrable_Pi_pmf_slice[OF assms(2) assms(1)], metis a) - - have "measure_pmf.expectation (Pi_pmf I d M) (\x. (f (x i))\<^sup>2) - (measure_pmf.expectation (Pi_pmf I d M) (\x. f (x i)))\<^sup>2 = - measure_pmf.expectation (M i) (\x. (f x)\<^sup>2) - (measure_pmf.expectation (M i) f)\<^sup>2" - using assms a b c by ((subst expectation_Pi_pmf_slice[OF assms(2,1)])?, simp)+ - - thus ?thesis - using assms a b c by (simp add: measure_pmf.variance_eq) -qed - -lemma Pi_pmf_bind_return: - assumes "finite I" - shows "Pi_pmf I d (\i. M i \ (\x. return_pmf (f i x))) = Pi_pmf I d' M \ (\x. return_pmf (\i. if i \ I then f i (x i) else d))" - using assms by (simp add: Pi_pmf_bind[where d'="d'"]) - -end +end \ No newline at end of file diff --git a/thys/Frequency_Moments/ROOT b/thys/Frequency_Moments/ROOT --- a/thys/Frequency_Moments/ROOT +++ b/thys/Frequency_Moments/ROOT @@ -1,27 +1,27 @@ chapter AFP session Frequency_Moments = "HOL-Probability" + options [timeout = 1200] sessions Bertrands_Postulate Concentration_Inequalities Equivalence_Relation_Enumeration - "HOL-Algebra" Interpolation_Polynomials_HOL_Algebra Lp Prefix_Free_Code_Combinators Median_Method Universal_Hash_Families + Expander_Graphs theories Frequency_Moments_Preliminary_Results Frequency_Moments Frequency_Moment_0 Frequency_Moment_2 Frequency_Moment_k Landau_Ext K_Smallest Probability_Ext Product_PMF_Ext document_files "root.tex" "root.bib" diff --git a/thys/Universal_Hash_Families/ROOT b/thys/Universal_Hash_Families/ROOT --- a/thys/Universal_Hash_Families/ROOT +++ b/thys/Universal_Hash_Families/ROOT @@ -1,16 +1,18 @@ chapter AFP session Universal_Hash_Families = "HOL-Probability" + options [timeout = 600] sessions "HOL-Algebra" - "Finite_Fields" - "Interpolation_Polynomials_HOL_Algebra" + Concentration_Inequalities + Finite_Fields + Interpolation_Polynomials_HOL_Algebra theories Universal_Hash_Families Universal_Hash_Families_More_Independent_Families Carter_Wegman_Hash_Family Universal_Hash_Families_More_Finite_Fields + Universal_Hash_Families_More_Product_PMF document_files "root.tex" "root.bib" diff --git a/thys/Universal_Hash_Families/Universal_Hash_Families_More_Product_PMF.thy b/thys/Universal_Hash_Families/Universal_Hash_Families_More_Product_PMF.thy new file mode 100644 --- /dev/null +++ b/thys/Universal_Hash_Families/Universal_Hash_Families_More_Product_PMF.thy @@ -0,0 +1,229 @@ +section \Indexed Products of Probability Mass Functions\ + +theory Universal_Hash_Families_More_Product_PMF + imports + "HOL-Probability.Product_PMF" + Concentration_Inequalities.Concentration_Inequalities_Preliminary + Universal_Hash_Families_More_Independent_Families +begin + +hide_const (open) Isolated.discrete + +text \This section introduces a restricted version of @{term "Pi_pmf"} where the default value is @{term "undefined"} +and contains some additional results about that case in addition to @{theory "HOL-Probability.Product_PMF"}\ + +abbreviation prod_pmf where "prod_pmf I M \ Pi_pmf I undefined M" + +lemma measure_pmf_cong: + assumes "\x. x \ set_pmf p \ x \ P \ x \ Q" + shows "measure (measure_pmf p) P = measure (measure_pmf p) Q" + using assms + by (intro finite_measure.finite_measure_eq_AE AE_pmfI) auto + +lemma pmf_mono: + assumes "\x. x \ set_pmf p \ x \ P \ x \ Q" + shows "measure (measure_pmf p) P \ measure (measure_pmf p) Q" +proof - + have "measure (measure_pmf p) P = measure (measure_pmf p) (P \ (set_pmf p))" + by (intro measure_pmf_cong) auto + also have "... \ measure (measure_pmf p) Q" + using assms by (intro finite_measure.finite_measure_mono) auto + finally show ?thesis by simp +qed + +lemma pmf_prod_pmf: + assumes "finite I" + shows "pmf (prod_pmf I M) x = (if x \ extensional I then \i \ I. (pmf (M i)) (x i) else 0)" + by (simp add: pmf_Pi[OF assms(1)] extensional_def) + +lemma PiE_defaut_undefined_eq: "PiE_dflt I undefined M = PiE I M" + by (simp add:PiE_dflt_def PiE_def extensional_def Pi_def set_eq_iff) blast + +lemma set_prod_pmf: + assumes "finite I" + shows "set_pmf (prod_pmf I M) = PiE I (set_pmf \ M)" + by (simp add:set_Pi_pmf[OF assms] PiE_defaut_undefined_eq) + +text \A more general version of @{thm [source] "measure_Pi_pmf_Pi"}.\ +lemma prob_prod_pmf': + assumes "finite I" + assumes "J \ I" + shows "measure (measure_pmf (Pi_pmf I d M)) (Pi J A) = (\ i \ J. measure (M i) (A i))" +proof - + have a:"Pi J A = Pi I (\i. if i \ J then A i else UNIV)" + using assms by (simp add:Pi_def set_eq_iff, blast) + show ?thesis + using assms + by (simp add:if_distrib a measure_Pi_pmf_Pi[OF assms(1)] prod.If_cases[OF assms(1)] Int_absorb1) +qed + +lemma prob_prod_pmf_slice: + assumes "finite I" + assumes "i \ I" + shows "measure (measure_pmf (prod_pmf I M)) {\. P (\ i)} = measure (M i) {\. P \}" + using prob_prod_pmf'[OF assms(1), where J="{i}" and M="M" and A="\_. Collect P"] + by (simp add:assms Pi_def) + +definition restrict_dfl where "restrict_dfl f A d = (\x. if x \ A then f x else d)" + +lemma pi_pmf_decompose: + assumes "finite I" + shows "Pi_pmf I d M = map_pmf (\\. restrict_dfl (\i. \ (f i) i) I d) (Pi_pmf (f ` I) (\_. d) (\j. Pi_pmf (f -` {j} \ I) d M))" +proof - + have fin_F_I:"finite (f ` I)" using assms by blast + + have "finite I \ ?thesis" + using fin_F_I + proof (induction "f ` I" arbitrary: I rule:finite_induct) + case empty + then show ?case by (simp add:restrict_dfl_def) + next + case (insert x F) + have a: "(f -` {x} \ I) \ (f -` F \ I) = I" + using insert(4) by blast + have b: "F = f ` (f -` F \ I) " using insert(2,4) + by (auto simp add:set_eq_iff image_def vimage_def) + have c: "finite (f -` F \ I)" using insert by blast + have d: "\j. j \ F \ (f -` {j} \ (f -` F \ I)) = (f -` {j} \ I)" + using insert(4) by blast + + have " Pi_pmf I d M = Pi_pmf ((f -` {x} \ I) \ (f -` F \ I)) d M" + by (simp add:a) + also have "... = map_pmf (\(g, h) i. if i \ f -` {x} \ I then g i else h i) + (pair_pmf (Pi_pmf (f -` {x} \ I) d M) (Pi_pmf (f -` F \ I) d M))" + using insert by (subst Pi_pmf_union) auto + also have "... = map_pmf (\(g,h) i. if f i = x \ i \ I then g i else if f i \ F \ i \ I then h (f i) i else d) + (pair_pmf (Pi_pmf (f -` {x} \ I) d M) (Pi_pmf F (\_. d) (\j. Pi_pmf (f -` {j} \ (f -` F \ I)) d M)))" + by (simp add:insert(3)[OF b c] map_pmf_comp case_prod_beta' apsnd_def map_prod_def + pair_map_pmf2 b[symmetric] restrict_dfl_def) (metis fst_conv snd_conv) + also have "... = map_pmf (\(g,h) i. if i \ I then (h(x := g)) (f i) i else d) + (pair_pmf (Pi_pmf (f -` {x} \ I) d M) (Pi_pmf F (\_. d) (\j. Pi_pmf (f -` {j} \ I) d M)))" + using insert(4) d + by (intro arg_cong2[where f="map_pmf"] ext) (auto simp add:case_prod_beta' cong:Pi_pmf_cong) + also have "... = map_pmf (\\ i. if i \ I then \ (f i) i else d) (Pi_pmf (insert x F) (\_. d) (\j. Pi_pmf (f -` {j} \ I) d M))" + by (simp add:Pi_pmf_insert[OF insert(1,2)] map_pmf_comp case_prod_beta') + finally show ?case by (simp add:insert(4) restrict_dfl_def) + qed + thus ?thesis using assms by blast +qed + +lemma restrict_dfl_iter: "restrict_dfl (restrict_dfl f I d) J d = restrict_dfl f (I \ J) d" + by (rule ext, simp add:restrict_dfl_def) + +lemma indep_vars_restrict': + assumes "finite I" + shows "prob_space.indep_vars (Pi_pmf I d M) (\_. discrete) (\i \. restrict_dfl \ (f -` {i} \ I) d) (f ` I)" +proof - + let ?Q = "(Pi_pmf (f ` I) (\_. d) (\i. Pi_pmf (I \ f -` {i}) d M))" + have a:"prob_space.indep_vars ?Q (\_. discrete) (\i \. \ i) (f ` I)" + using assms by (intro indep_vars_Pi_pmf, blast) + have b: "AE x in measure_pmf ?Q. \i\f ` I. x i = restrict_dfl (\i. x (f i) i) (I \ f -` {i}) d" + using assms + by (auto simp add:PiE_dflt_def restrict_dfl_def AE_measure_pmf_iff set_Pi_pmf comp_def Int_commute) + have "prob_space.indep_vars ?Q (\_. discrete) (\i x. restrict_dfl (\i. x (f i) i) (I \ f -` {i}) d) (f ` I)" + by (rule prob_space.indep_vars_cong_AE[OF prob_space_measure_pmf b a], simp) + thus ?thesis + using prob_space_measure_pmf + by (auto intro!:prob_space.indep_vars_distr simp:pi_pmf_decompose[OF assms, where f="f"] + map_pmf_rep_eq comp_def restrict_dfl_iter Int_commute) +qed + +lemma indep_vars_restrict_intro': + assumes "finite I" + assumes "\i \. i \ J \ X' i \ = X' i (restrict_dfl \ (f -` {i} \ I) d)" + assumes "J = f ` I" + assumes "\\ i. i \ J \ X' i \ \ space (M' i)" + shows "prob_space.indep_vars (measure_pmf (Pi_pmf I d p)) M' (\i \. X' i \) J" +proof - + define M where "M \ measure_pmf (Pi_pmf I d p)" + interpret prob_space "M" + using M_def prob_space_measure_pmf by blast + have "indep_vars (\_. discrete) (\i x. restrict_dfl x (f -` {i} \ I) d) (f ` I)" + unfolding M_def by (rule indep_vars_restrict'[OF assms(1)]) + hence "indep_vars M' (\i \. X' i (restrict_dfl \ ( f -` {i} \ I) d)) (f ` I)" + using assms(4) + by (intro indep_vars_compose2[where Y="X'" and N="M'" and M'="\_. discrete"]) (auto simp:assms(3)) + hence "indep_vars M' (\i \. X' i \) (f ` I)" + using assms(2)[symmetric] + by (simp add:assms(3) cong:indep_vars_cong) + thus ?thesis + unfolding M_def using assms(3) by simp +qed + +lemma + fixes f :: "'b \ ('c :: {second_countable_topology,banach,real_normed_field})" + assumes "finite I" + assumes "i \ I" + assumes "integrable (measure_pmf (M i)) f" + shows integrable_Pi_pmf_slice: "integrable (Pi_pmf I d M) (\x. f (x i))" + and expectation_Pi_pmf_slice: "integral\<^sup>L (Pi_pmf I d M) (\x. f (x i)) = integral\<^sup>L (M i) f" +proof - + have a:"distr (Pi_pmf I d M) (M i) (\\. \ i) = distr (Pi_pmf I d M) discrete (\\. \ i)" + by (rule distr_cong, auto) + + have b: "measure_pmf.random_variable (M i) borel f" + using assms(3) by simp + + have c:"integrable (distr (Pi_pmf I d M) (M i) (\\. \ i)) f" + using assms(1,2,3) + by (subst a, subst map_pmf_rep_eq[symmetric], subst Pi_pmf_component, auto) + + show "integrable (Pi_pmf I d M) (\x. f (x i))" + by (rule integrable_distr[where f="f" and M'="M i"]) (auto intro: c) + + have "integral\<^sup>L (Pi_pmf I d M) (\x. f (x i)) = integral\<^sup>L (distr (Pi_pmf I d M) (M i) (\\. \ i)) f" + using b by (intro integral_distr[symmetric], auto) + also have "... = integral\<^sup>L (map_pmf (\\. \ i) (Pi_pmf I d M)) f" + by (subst a, subst map_pmf_rep_eq[symmetric], simp) + also have "... = integral\<^sup>L (M i) f" + using assms(1,2) by (simp add: Pi_pmf_component) + finally show "integral\<^sup>L (Pi_pmf I d M) (\x. f (x i)) = integral\<^sup>L (M i) f" by simp +qed + +text \This is an improved version of @{thm [source] "expectation_prod_Pi_pmf"}. +It works for general normed fields instead of non-negative real functions .\ + +lemma expectation_prod_Pi_pmf: + fixes f :: "'a \ 'b \ ('c :: {second_countable_topology,banach,real_normed_field})" + assumes "finite I" + assumes "\i. i \ I \ integrable (measure_pmf (M i)) (f i)" + shows "integral\<^sup>L (Pi_pmf I d M) (\x. (\i \ I. f i (x i))) = (\ i \ I. integral\<^sup>L (M i) (f i))" +proof - + have a: "prob_space.indep_vars (measure_pmf (Pi_pmf I d M)) (\_. borel) (\i \. f i (\ i)) I" + by (intro prob_space.indep_vars_compose2[where Y="f" and M'="\_. discrete"] + prob_space_measure_pmf indep_vars_Pi_pmf assms(1)) auto + have "integral\<^sup>L (Pi_pmf I d M) (\x. (\i \ I. f i (x i))) = (\ i \ I. integral\<^sup>L (Pi_pmf I d M) (\x. f i (x i)))" + by (intro prob_space.indep_vars_lebesgue_integral prob_space_measure_pmf assms(1,2) + a integrable_Pi_pmf_slice) auto + also have "... = (\ i \ I. integral\<^sup>L (M i) (f i))" + by (intro prod.cong expectation_Pi_pmf_slice assms(1,2)) auto + finally show ?thesis by simp +qed + +lemma variance_prod_pmf_slice: + fixes f :: "'a \ real" + assumes "i \ I" "finite I" + assumes "integrable (measure_pmf (M i)) (\\. f \^2)" + shows "prob_space.variance (Pi_pmf I d M) (\\. f (\ i)) = prob_space.variance (M i) f" +proof - + have a:"integrable (measure_pmf (M i)) f" + using assms(3) measure_pmf.square_integrable_imp_integrable by auto + have b:" integrable (measure_pmf (Pi_pmf I d M)) (\x. (f (x i))\<^sup>2)" + by (rule integrable_Pi_pmf_slice[OF assms(2) assms(1)], metis assms(3)) + have c:" integrable (measure_pmf (Pi_pmf I d M)) (\x. (f (x i)))" + by (rule integrable_Pi_pmf_slice[OF assms(2) assms(1)], metis a) + + have "measure_pmf.expectation (Pi_pmf I d M) (\x. (f (x i))\<^sup>2) - (measure_pmf.expectation (Pi_pmf I d M) (\x. f (x i)))\<^sup>2 = + measure_pmf.expectation (M i) (\x. (f x)\<^sup>2) - (measure_pmf.expectation (M i) f)\<^sup>2" + using assms a b c by ((subst expectation_Pi_pmf_slice[OF assms(2,1)])?, simp)+ + + thus ?thesis + using assms a b c by (simp add: measure_pmf.variance_eq) +qed + +lemma Pi_pmf_bind_return: + assumes "finite I" + shows "Pi_pmf I d (\i. M i \ (\x. return_pmf (f i x))) = Pi_pmf I d' M \ (\x. return_pmf (\i. if i \ I then f i (x i) else d))" + using assms by (simp add: Pi_pmf_bind[where d'="d'"]) + +end