diff --git a/thys/Approximate_Model_Counting/ApproxMCPreliminaries.thy b/thys/Approximate_Model_Counting/ApproxMCPreliminaries.thy --- a/thys/Approximate_Model_Counting/ApproxMCPreliminaries.thy +++ b/thys/Approximate_Model_Counting/ApproxMCPreliminaries.thy @@ -1,342 +1,343 @@ section \Preliminary probability/UHF lemmas\ text \ This section proves some simplified/specialized forms of lemmas that will be used in the algorithm's analysis later. \ theory ApproxMCPreliminaries imports Frequency_Moments.Probability_Ext + Concentration_Inequalities.Bienaymes_Identity Concentration_Inequalities.Paley_Zygmund_Inequality begin lemma card_inter_sum_indicat_real: assumes "finite A" shows "card (A \ B) = sum (indicat_real B) A" by (simp add: assms indicator_def) (* Counting number of finite maps *) lemma card_dom_ran: assumes "finite D" shows "card {w. dom w = D \ ran w \ R} = card R ^ card D" using assms proof (induct rule: finite_induct) case empty have "{w::'a \ 'b option. w = Map.empty \ ran w \ R} = {Map.empty}" by auto then show ?case by auto next case (insert a A) have 1: "inj_on (\(w, r). w(a \ r)) ({w. dom w = A \ ran w \ R} \ R)" unfolding inj_on_def by (smt (verit, del_insts) CollectD SigmaE fun_upd_None_if_notin_dom local.insert(2) map_upd_eqD1 prod.simps(2) restrict_complement_singleton_eq restrict_upd_same) have 21: "(\(w, r). w(a \ r)) ` ({w. dom w = A \ ran w \ R} \ R) \ {w. dom w = insert a A \ ran w \ R}" unfolding image_def using CollectD local.insert(2) by force have "\x. dom x = insert a A \ ran x \ R \ \xa. dom xa = A \ ran xa \ R \ (\y\R. x = xa(a \ y))" by (smt (verit, del_insts) Diff_insert_absorb domD dom_fun_upd fun_upd_triv fun_upd_upd insert.hyps(2) insertCI insert_subset ran_map_upd) then have 22: " {w. dom w = insert a A \ ran w \ R} \ (\(w, r). w(a \ r)) ` ({w. dom w = A \ ran w \ R} \ R)" by (clarsimp simp add: image_def) have "bij_betw (\(w,r). w(a\r)) ({w. dom w = A \ ran w \ R} \ R) {w. dom w = insert a A \ ran w \ R} " unfolding bij_betw_def using 1 21 22 by clarsimp then have "card {w. dom w = insert a A \ ran w \ R} = card ({w. dom w = A \ ran w \ R} \ R)" by (auto simp add: bij_betw_same_card 1 21 22) moreover have "... = card R ^ card A * card R" by(subst card_cartesian_product) (use insert.hyps(3) in auto) ultimately show ?case using insert.hyps by (auto simp add: card_insert_if) qed lemma finite_set_pmf_expectation_sum: fixes f :: "'a \ 'c \ 'b::{banach, second_countable_topology}" assumes "finite (set_pmf A)" shows "measure_pmf.expectation A (\x. sum (f x) T) = (\i\T. measure_pmf.expectation A (\x. f x i))" apply (intro Bochner_Integration.integral_sum integrable_measure_pmf_finite) using assms by auto lemma (in prob_space) k_universal_prob_unif: assumes "k_universal k H D R" assumes "w \ D" "\ \ R" shows "prob {s \ space M. H w s = \} = 1 / card R" proof - have "uniform_on (H w) R" using assms unfolding k_universal_def by auto from uniform_onD[OF this] have "prob {\ \ space M. H w \ \ {\}} = real (card (R \ {\})) / real (card R)" . thus ?thesis using assms by auto qed (* For h: D \ R, k-universal, S \ D. E( | {w \ S. h w = \} | ) = |S| / |R| *) lemma k_universal_expectation_eq: assumes p: "finite (set_pmf p)" assumes ind: "prob_space.k_universal p k H D R" assumes S: "finite S" "S \ D" assumes a: "\ \ R" shows "prob_space.expectation p (\s. real (card (S \ {w. H w s = \}))) = real (card S) / card R" proof - have 1: "prob_space (measure_pmf p)" by (simp add: prob_space_measure_pmf) have 2: "space (measure_pmf p) = UNIV" by auto from prob_space.k_universal_prob_unif[OF 1 ind _ a] have *: "\w. w \ S \ prob_space.prob p {s. H w s = \} = 1 / real (card R)" using S(2) by auto have "measure_pmf.expectation p (\s. real (card (S \ {w. H w s = \}))) = measure_pmf.expectation p (\s. sum (indicat_real {w. H w s = \}) S)" unfolding card_inter_sum_indicat_real[OF S(1)] by presburger moreover have "... = (\i\S. measure_pmf.expectation p (indicat_real {s. H i s = \}))" apply (subst finite_set_pmf_expectation_sum) using assms unfolding indicator_def by auto moreover have " ... = (\i\ S. measure_pmf.prob p {s. H i s = \})" by auto moreover have "... = (\i\S. 1 / card R)" using * by auto ultimately show ?thesis by auto qed lemma (in prob_space) two_universal_indep_var: assumes "k_universal 2 H D R" assumes "w \ D" "w' \ D" "w \ w'" shows "indep_var borel (\x. indicat_real {w. H w x = \} w) borel (\x. indicat_real {w. H w x = \} w')" proof - have Y: "(\z. (of_bool (z = \))::real) \ (count_space UNIV) \\<^sub>M borel" by auto have "k_wise_indep_vars 2 (\_. count_space UNIV) H D" using assms unfolding k_universal_def by auto then have "indep_vars (\_. count_space UNIV) H {w, w'}" unfolding k_wise_indep_vars_def by (metis UNIV_bool assms(2) assms(3) card.empty card.insert card_UNIV_bool card_insert_le empty_iff empty_subsetI finite.emptyI finite.insertI insert_subset order.refl singletonD singleton_insert_inj_eq') from indep_var_from_indep_vars[OF assms(4) this] have "indep_var (count_space UNIV) (H w) (count_space UNIV) (H w')" . from indep_var_compose[OF this Y Y] show ?thesis unfolding indicator_def by (auto simp add: o_def) qed (* For h: D \ R, 2-universal, S \ D. V( | {w \ S. h w = \} | ) \ E( | {w \ S. h w = \} | ) *) lemma two_universal_variance_bound: assumes p: "finite (set_pmf p)" assumes ind: "prob_space.k_universal (measure_pmf p) 2 H D R" assumes S: "finite S" "S \ D" assumes a: "\ \ R" shows "measure_pmf.variance p (\s. real (card (S \ {w. H w s = \}))) \ measure_pmf.expectation p (\s. real (card (S \ {w. H w s = \})))" proof - have vb: "measure_pmf.variance p (\x. (indicat_real {w. H w x = \} i)) \ measure_pmf.expectation p (\x. (indicat_real {w. H w x = \} i))" for i proof - have "measure_pmf.variance p (\x. (indicat_real {w. H w x = \} i)) = measure_pmf.expectation p (\x. (indicat_real {w. H w x = \} i)\<^sup>2) - (measure_pmf.expectation p (\x. indicat_real {w. H w x = \} i))\<^sup>2" apply (subst measure_pmf.variance_eq) by (auto simp add: p integrable_measure_pmf_finite) moreover have "... \ measure_pmf.expectation p (\x. (indicat_real {w. H w x = \} i)\<^sup>2)" by simp moreover have "... = measure_pmf.expectation p (\x. (indicat_real {w. H w x = \} i))" by (metis (mono_tags, lifting) indicator_simps(1) indicator_simps(2) power2_eq_1_iff zero_eq_power2) ultimately show ?thesis by linarith qed have "measure_pmf.variance p (\s. real (card (S \ {w. H w s = \}))) = measure_pmf.variance p (\s. sum (indicat_real {w. H w s = \}) S)" unfolding card_inter_sum_indicat_real[OF S(1)] by presburger then have "... = (\i\S. measure_pmf.variance p (\x. (indicat_real {w. H w x = \} i)))" - apply (subst measure_pmf.var_sum_pairwise_indep) + apply (subst measure_pmf.bienaymes_identity_pairwise_indep) using S prob_space_measure_pmf by (auto intro!: prob_space.two_universal_indep_var[OF _ ind] simp add: p integrable_measure_pmf_finite) moreover have "... \ (\i\S. measure_pmf.expectation p (\x. (indicat_real {w. H w x = \} i)))" by (simp add: sum_mono vb) moreover have "... = measure_pmf.expectation p (\s. sum (indicat_real {w. H w s = \}) S)" apply (subst finite_set_pmf_expectation_sum) using assms by auto ultimately show ?thesis by (simp add: assms(3) card_inter_sum_indicat_real) qed lemma (in prob_space) k_universal_mono: assumes "k' \ k" assumes"k_universal k H D R" shows"k_universal k' H D R" using assms unfolding k_universal_def k_wise_indep_vars_def by auto lemma finite_set_pmf_expectation_add: assumes "finite (set_pmf S)" shows "measure_pmf.expectation S (\x. ((f x)::real) + g x) = measure_pmf.expectation S f + measure_pmf.expectation S g" by (auto intro: Bochner_Integration.integral_add simp add: assms integrable_measure_pmf_finite) lemma finite_set_pmf_expectation_add_const: assumes "finite (set_pmf S)" shows "measure_pmf.expectation S (\x. ((f x)::real) + g) = measure_pmf.expectation S f + g" proof - have "g = measure_pmf.expectation S (\x. g)" by simp thus ?thesis by (simp add: assms finite_set_pmf_expectation_add) qed lemma finite_set_pmf_expectation_diff: assumes "finite (set_pmf S)" shows "measure_pmf.expectation S (\x. ((f x)::real) - g x) = measure_pmf.expectation S f - measure_pmf.expectation S g" by (auto intro: Bochner_Integration.integral_diff simp add: assms integrable_measure_pmf_finite) (* convenient forms of library inequalities *) lemma spec_paley_zygmund_inequality: assumes fin: "finite (set_pmf p)" assumes Zpos: "\z. Z z \ 0" assumes t: "\ \ 1" shows " (measure_pmf.variance p Z + (1-\)^2 * (measure_pmf.expectation p Z)^2) * measure_pmf.prob p {z. Z z > \ * measure_pmf.expectation p Z} \ (1-\)^2 * (measure_pmf.expectation p Z)^2" proof - have "prob_space (measure_pmf p)" by (auto simp add: prob_space_measure_pmf) from prob_space.paley_zygmund_inequality[OF this _ integrable_measure_pmf_finite[OF fin] t] show ?thesis using Zpos by auto qed lemma spec_chebyshev_inequality: assumes fin: "finite (set_pmf p)" assumes pvar: "measure_pmf.variance p Y > 0" assumes k: "k > 0" shows " measure_pmf.prob p {y. (Y y - measure_pmf.expectation p Y)^2 \ k^2 * measure_pmf.variance p Y} \ 1 / k^2" proof - define f where "f x = Y x / sqrt(measure_pmf.variance p Y)" for x have 1: "measure_pmf.random_variable p borel f" by auto have *: "(\x. (f x)\<^sup>2) = (\x. (Y x)\<^sup>2/ measure_pmf.variance p Y)" unfolding f_def by (simp add: power_divide) have 2: "integrable p (\x. (f x)\<^sup>2)" unfolding * by (intro integrable_measure_pmf_finite[OF fin]) from measure_pmf.Chebyshev_inequality[OF 1 2 k] have ineq1:"measure_pmf.prob p {x . k \ \f x - measure_pmf.expectation p f\} \ measure_pmf.expectation p (\x. (f x - measure_pmf.expectation p f)\<^sup>2) / k\<^sup>2" by auto have "(\x. (f x - measure_pmf.expectation p f)\<^sup>2) = (\x. ((Y x - measure_pmf.expectation p Y) / sqrt(measure_pmf.variance p Y))\<^sup>2)" unfolding f_def by (simp add: diff_divide_distrib) moreover have "... = (\x. (Y x - measure_pmf.expectation p Y)^2 / (sqrt(measure_pmf.variance p Y))^2)" by (simp add: power_divide) moreover have "... = (\x. (Y x - measure_pmf.expectation p Y)^2 / measure_pmf.variance p Y)" by simp ultimately have unfold:"(\x. (f x - measure_pmf.expectation p f)\<^sup>2) = (\x. (Y x - measure_pmf.expectation p Y)^2 / measure_pmf.variance p Y)" by auto then have "measure_pmf.expectation p (\x. (f x - measure_pmf.expectation p f)\<^sup>2) / k\<^sup>2 = measure_pmf.expectation p (\x. (Y x - measure_pmf.expectation p Y)^2 / measure_pmf.variance p Y) / k\<^sup>2" by auto moreover have "... = measure_pmf.variance p Y / measure_pmf.variance p Y / k\<^sup>2" by simp moreover have "... = 1 / k\<^sup>2" using pvar by force ultimately have eq1:"measure_pmf.expectation p (\x. (f x - measure_pmf.expectation p f)\<^sup>2) / k\<^sup>2 = 1 / k\<^sup>2" by auto have "(\x. k \ \f x - measure_pmf.expectation p f\) = (\x. k\<^sup>2 \ (f x - measure_pmf.expectation p f)\<^sup>2)" by (metis (no_types, opaque_lifting) abs_of_nonneg k less_le real_le_rsqrt real_sqrt_abs sqrt_ge_absD) moreover have "... = (\x. k\<^sup>2 \ ((Y x - measure_pmf.expectation p Y)^2 / measure_pmf.variance p Y))" by (metis unfold) moreover have "... = (\x. (Y x - measure_pmf.expectation p Y)^2 \ k\<^sup>2 * measure_pmf.variance p Y)" by (simp add: pos_le_divide_eq pvar) ultimately have cond:"(\x. k \ \f x - measure_pmf.expectation p f\) = (\x. (Y x - measure_pmf.expectation p Y)^2 \ k\<^sup>2 * measure_pmf.variance p Y)" by auto show "?thesis" using ineq1 cond eq1 by auto qed end diff --git a/thys/Approximate_Model_Counting/RandomXOR.thy b/thys/Approximate_Model_Counting/RandomXOR.thy --- a/thys/Approximate_Model_Counting/RandomXOR.thy +++ b/thys/Approximate_Model_Counting/RandomXOR.thy @@ -1,1634 +1,1634 @@ section \Random XORs\ text \The goal of this section is to prove that, for a randomly sampled XOR $X$ from a set of variables $V$: \begin{enumerate} \item the probability of an assignment $w$ satisfying $X$ is $\frac{1}{2}$; \item for any distinct assignments $w$, $w'$ the probability of both satisfying $X$ is equal to $\frac{1}{4}$ (2-wise independence); and \item for any distinct assignments $w$, $w'$, $w''$ the probability of all three satisfying $X$ is equal to $\frac{1}{8}$ (3-wise independence). \end{enumerate} \ theory RandomXOR imports ApproxMCPreliminaries - Frequency_Moments.Product_PMF_Ext + Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF Monad_Normalisation.Monad_Normalisation begin text \A random XOR constraint is modeled as a random subset of variables and a randomly chosen RHS bit. \ definition random_xor :: "'a set \ ('a set \ bool) pmf" where "random_xor V = pair_pmf (pmf_of_set (Pow V)) (bernoulli_pmf (1/2))" lemma pmf_of_set_Pow_fin_map: assumes V:"finite V" shows "pmf_of_set (Pow V) = map_pmf (\b. {x \ V. b x = Some True}) (Pi_pmf V def (\_. map_pmf Some (bernoulli_pmf (1 / 2))))" proof - have *: "Pi_pmf V def (\_. map_pmf Some (bernoulli_pmf (1 / 2))) = map_pmf (\f x. if x \ V then f x else def) (Pi_pmf V (Some False) (\_. map_pmf Some (bernoulli_pmf (1 / 2))))" unfolding Pi_pmf_default_swap[OF V] by auto have **: "pmf_of_set (Pow V) = map_pmf (\x. {xa. (xa \ V \ x xa) \ xa \ V}) (Pi_pmf V False (\_. bernoulli_pmf (1 / 2)))" by (smt (verit, ccfv_SIG) Collect_cong pmf_of_set_Pow_conv_bernoulli V pmf.map_cong) show ?thesis unfolding * apply (subst Pi_pmf_map[OF V, of _ "False"]) using ** by (auto simp add: pmf.map_comp o_def) qed (* A random XOR can also be sampled as a map |V| \ bool + a coin flip *) lemma random_xor_from_bits: assumes V:"finite V" shows "random_xor V = pair_pmf (map_pmf (\b. {x \ V. b x = Some True}) (Pi_pmf V def (\_. map_pmf Some (bernoulli_pmf (1/2))))) (bernoulli_pmf (1/2))" unfolding random_xor_def using V pmf_of_set_Pow_fin_map by fastforce (* An assignment is a subset \ of V i.e., those variables assigned to true *) fun satisfies_xor :: "('a set \ bool) \ 'a set \ bool" where "satisfies_xor (x,b) \ = even (card (\ \ x) + of_bool b) " (* (* x_0 \ x_1 \ x_2 = 0 *) value "satisfies_xor ({0,1,2},False) {0::nat}" (* x_0 \ x_1 \ x_2 = 1 *) value "satisfies_xor ({0,1,2},True) {0::nat}" *) lemma satisfies_xor_inter: shows "satisfies_xor ( \ \ x ,b) \ = satisfies_xor (x,b) \" by (auto simp add: Int_commute) lemma prob_bernoulli_bind_pmf: assumes "0 \ p" "p \ 1" assumes "finite E" shows "measure_pmf.prob (bernoulli_pmf p \ x) E = p * (measure_pmf.prob (x True) E) + (1 - p) * (measure_pmf.prob (x False) E) " using assms by (auto simp add: pmf_bind measure_measure_pmf_finite[OF assms(3)] vector_space_over_itself.scale_sum_right comm_monoid_add_class.sum.distrib mult.commute) lemma set_pmf_random_xor: assumes V: "finite V" shows "set_pmf (random_xor V) = (Pow V) \ UNIV" unfolding random_xor_def using assms by (auto simp add: Pow_not_empty Set.basic_monos(7)) lemma pmf_of_set_prod: assumes "P \ {}" "Q \ {}" assumes "finite P" "finite Q" shows "pmf_of_set (P \ Q) = pair_pmf (pmf_of_set P) (pmf_of_set Q)" by (auto intro!: pmf_eqI simp add: indicator_def pmf_pair assms) lemma random_xor_pmf_of_set: assumes V:"finite V" shows "random_xor V = pmf_of_set ((Pow V) \ UNIV)" unfolding random_xor_def apply (subst pmf_of_set_prod) using V bernoulli_pmf_half_conv_pmf_of_set by auto lemma prob_random_xor_with_set_pmf: assumes V: "finite V" shows "prob_space.prob (random_xor V) {c. P c} = prob_space.prob (random_xor V) {c. fst c \ V \ P c}" by (smt (verit, best) PowD assms measure_pmf.measure_pmf_eq mem_Collect_eq mem_Sigma_iff prod.collapse set_pmf_random_xor) lemma prob_set_parity: assumes "measure_pmf.prob M {c. P c} = q" shows "measure_pmf.prob M {c. P c = b} = (if b then q else 1 - q)" proof - { assume b:"b" have ?thesis using assms using b by presburger } moreover { assume b:"\b" have "{c. P c} \ measure_pmf.events M" by auto from measure_pmf.prob_compl[OF this] have "1 - q = measure_pmf.prob M (UNIV - {c. P c})" using assms by auto moreover have " ... = prob_space.prob M {c. P c = b}" by (simp add: b measure_pmf.measure_pmf_eq) ultimately have ?thesis using b by presburger } ultimately show ?thesis by auto qed lemma satisfies_random_xor: assumes V: "finite V" shows "prob_space.prob (random_xor V) {c. satisfies_xor c \} = 1 / 2" proof - have eq: "{(c::'a set \ bool). fst c \ V} = {c. c \ V} \ UNIV" by auto then have "finite {(c::'a set \ bool). fst c \ V}" using assms by auto then have *: " x \ V \ measure_pmf.prob (bernoulli_pmf (1 / 2) \ (\y. return_pmf (x, y))) {c. fst c \ V \ satisfies_xor c \} = 1 / 2" for x apply (subst prob_bernoulli_bind_pmf) by (auto simp add: indicator_def) have "prob_space.prob (random_xor V) {c. satisfies_xor c \} = prob_space.prob (random_xor V) {c. fst c \ V \ satisfies_xor c \}" using prob_random_xor_with_set_pmf[OF V] by auto also have "... = prob_space.expectation (random_xor V) (indicat_real {c. fst c \ V \ satisfies_xor c \})" by auto also have "... = (\a\Pow V. inverse (real (card (Pow V))) * measure_pmf.prob (bernoulli_pmf (1 / 2) \ (\y. return_pmf (a, y))) {c. fst c \ V \ satisfies_xor c \})" unfolding random_xor_def pair_pmf_def apply (subst pmf_expectation_bind_pmf_of_set) using assms by auto also have "... = (\a\Pow V. inverse (real (card (Pow V))) * (1/2))" by (simp add: *) also have "... = inverse (real (card (Pow V))) * (1/2) * (\a\Pow V. 1)" using * by force also have "... = 1/2" by (metis One_nat_def Suc_leI assms card_Pow divide_real_def inverse_inverse_eq inverse_nonzero_iff_nonzero nat_one_le_power nonzero_mult_div_cancel_left not_one_le_zero of_nat_eq_0_iff pos2 real_of_card) finally show ?thesis by presburger qed lemma satisfies_random_xor_parity: assumes V: "finite V" shows "prob_space.prob (random_xor V) {c. satisfies_xor c \ = b} = 1 / 2" using prob_set_parity[OF satisfies_random_xor[OF V]] by auto subsection "Independence properties of random XORs" lemma pmf_of_set_powerset_split: assumes "S \ V" "finite V" shows " map_pmf (\(x,y). x \ y) (pmf_of_set (Pow S \ Pow (V - S))) = pmf_of_set (Pow V)" proof - have spmfS: "set_pmf (pmf_of_set (Pow S)) = Pow S" using assms by (auto intro!: set_pmf_of_set simp add: rev_finite_subset) have spmfVS: "set_pmf (pmf_of_set (Pow (V-S))) = Pow (V-S)" using assms by (auto intro!: set_pmf_of_set) have xsub: "x \ V \ \xa\Pow S. \y\Pow (V - S). x = xa \ y" for x by (metis Diff_subset_conv Pow_iff Un_Diff_Int basic_trans_rules(23) inf_le2 sup_commute) have inj: "inj_on (\(x, y). x \ y) (Pow S \ Pow (V - S))" unfolding inj_on_def by auto then have bij: "bij_betw (\(x, y). x \ y) ((Pow S) \ Pow (V - S)) (Pow V)" unfolding bij_betw_def using assms(1) xsub by (auto simp add: image_def) have "map_pmf (\(x, y). x \ y) (pmf_of_set (Pow S \ Pow (V - S))) = pmf_of_set (Pow V)" apply (subst map_pmf_of_set_inj[OF inj]) subgoal by (auto simp add: image_def) subgoal using bij assms(2) bij_betw_finite by blast apply (intro arg_cong[where f = "pmf_of_set"]) using assms(1) xsub by (auto simp add: image_def) thus ?thesis unfolding spmfS spmfVS by auto qed lemma pmf_of_set_Pow_sing: shows"pmf_of_set (Pow {x}) = bernoulli_pmf (1 / 2) \ (\b. return_pmf (if b then {x} else {}))" apply (intro pmf_eqI) apply (subst pmf_of_set) by (auto simp add: pmf_bind card_Pow indicator_def subset_singleton_iff) lemma pmf_of_set_sing_coin_flip: assumes "finite V" shows "pmf_of_set (Pow {x} \ Pow V) = map_pmf (\(r,c). (if c then {x} else {}, r)) (random_xor V)" proof - have *: "pmf_of_set (Pow {x} \ Pow V) = pair_pmf (pmf_of_set (Pow {x})) (pmf_of_set (Pow V))" apply(intro pmf_of_set_prod) using assms by auto show ?thesis unfolding * apply (intro pmf_eqI) including monad_normalisation by (auto simp add: map_pmf_def pair_pmf_def random_xor_def pmf_of_set_Pow_sing) qed (* TODO: there is a more general version below *) lemma measure_pmf_prob_dependent_product_bound_eq: assumes "countable A" "\i. countable (B i)" assumes "\a. a \ A \ measure_pmf.prob N (B a) = r" shows "measure_pmf.prob (pair_pmf M N) (Sigma A B) = measure_pmf.prob M A * r" proof - have "measure_pmf.prob (pair_pmf M N) (Sigma A B) = (\\<^sub>a(a, b) \ Sigma A B. pmf M a * pmf N b)" by (auto intro!: infsetsum_cong simp add: measure_pmf_conv_infsetsum pmf_pair) also have "... = (\\<^sub>aa\A. \\<^sub>ab\B a. pmf M a * pmf N b)" apply (subst infsetsum_Sigma[OF assms(1-2)]) subgoal by (metis (no_types, lifting) SigmaE abs_summable_on_cong case_prod_conv pmf_abs_summable pmf_pair) by (auto simp add: assms case_prod_unfold) also have "... = (\\<^sub>aa\A. pmf M a * (measure_pmf.prob N (B a)))" by (simp add: infsetsum_cmult_right measure_pmf_conv_infsetsum pmf_abs_summable) also have "... = (\\<^sub>aa\A. pmf M a * r)" using assms(3) by force also have"... = measure_pmf.prob M A * r" by (simp add: infsetsum_cmult_left pmf_abs_summable measure_pmf_conv_infsetsum) finally show ?thesis by linarith qed (* TODO: duplicated in Schwartz_Zippel, but we don't want to pull in the multivariate polynomials library *) lemma measure_pmf_prob_dependent_product_bound_eq': assumes "countable (A \ set_pmf M)" "\i. countable (B i \ set_pmf N)" assumes "\a. a \ A \ set_pmf M \ measure_pmf.prob N (B a \ set_pmf N) = r" shows "measure_pmf.prob (pair_pmf M N) (Sigma A B) = measure_pmf.prob M A * r" proof - have *: "Sigma A B \ (set_pmf M \ set_pmf N) = Sigma (A \ set_pmf M) (\i. B i \ set_pmf N)" by auto have "measure_pmf.prob (pair_pmf M N) (Sigma A B) = measure_pmf.prob (pair_pmf M N) (Sigma (A \ set_pmf M) (\i. B i \ set_pmf N))" by (metis * measure_Int_set_pmf set_pair_pmf) moreover have "... = measure_pmf.prob M (A \ set_pmf M) * r" using measure_pmf_prob_dependent_product_bound_eq[OF assms(1-3)] by auto moreover have "... = measure_pmf.prob M A * r" by (simp add: measure_Int_set_pmf) ultimately show ?thesis by linarith qed lemma single_var_parity_coin_flip: assumes "x \ \" "finite \" assumes "finite a" "x \ a" shows "measure_pmf.prob (pmf_of_set (Pow {x})) {y. even (card ((a \ y) \ \)) = b} = 1/2" proof - have "insert x a \ \ = insert x (a \ \)" using assms by auto then have *: "card (insert x a \ \) = 1 + card (a \ \)" by (simp add: assms(2) assms(4)) have "measure_pmf.prob (pmf_of_set (Pow {x})) {y. even (card ((a \ y) \ \)) = b} = measure_pmf.prob (bernoulli_pmf (1/2)) {odd (card (a \ \)) = b}" unfolding pmf_of_set_Pow_sing map_pmf_def[symmetric] by (auto intro!: measure_prob_cong_0 simp add:image_def * ) moreover have "... = 1/2" by (simp add: measure_pmf_single) ultimately show ?thesis by auto qed (* Fix any event E that does not touch x Then we can sample that event separately *) lemma prob_pmf_of_set_nonempty_parity: assumes V: "finite V" assumes "x \ \" "\ \ V" assumes "\c. c \ E \ c - {x} \ E" shows "prob_space.prob (pmf_of_set (Pow V)) (E \ {c. even (card (c \ \)) = b}) = 1 / 2 * prob_space.prob (pmf_of_set (Pow (V - {x}))) E" proof - have 1: "set_pmf (pmf_of_set (Pow {x})) = Pow {x}" by (simp add: Pow_not_empty) have 2: "set_pmf (pmf_of_set (Pow (V-{x}))) = Pow (V-{x})" by (simp add: Pow_not_empty assms(1)) have 3: "set_pmf (pmf_of_set (Pow {x} \ Pow (V - {x}))) = Pow {x} \ Pow (V - {x})" by (simp add: Pow_not_empty assms(1)) have "{x} \ V" using assms by auto from pmf_of_set_powerset_split[OF this assms(1)] have e: "map_pmf (\(x, y). x \ y) (pmf_of_set (Pow {x} \ Pow (V - {x}))) = pmf_of_set (Pow V)" using 1 2 by auto have "map_pmf (\(x, y). x \ y) (pair_pmf (pmf_of_set (Pow {x})) (pmf_of_set (Pow (V - {x})))) = map_pmf (\(x, y). x \ y) (pair_pmf (pmf_of_set (Pow (V - {x}))) (pmf_of_set (Pow {x})))" apply (subst pair_commute_pmf) by (auto simp add: pmf.map_comp o_def case_prod_unfold Un_commute) then have *: "pmf_of_set (Pow V) = map_pmf (\(x, y). x \ y) (pair_pmf (pmf_of_set (Pow (V - {x}))) (pmf_of_set (Pow {x})))" unfolding e[symmetric] apply (subst pmf_of_set_prod) using V by auto have **: "((\(x, y). x \ y) -` (E \ S)) \ (Pow (V - {x}) \ Pow{x}) = Sigma E (\x. {y. (x \ y) \ S}) \ (Pow (V - {x}) \ Pow{x})" for S proof - have 11: "\a b. a \ b \ E \ a \ b \ S \ a \ V - {x} \ b \ {x} \ a \ E" by (metis Diff_insert_absorb Un_insert_right assms(4) boolean_algebra_cancel.sup0 subset_Diff_insert subset_singleton_iff) have 21: " \a b. a \ E \ a \ b \ S \ a \ V - {x} \ b \ {x} \ a \ b \ E" by (metis Diff_cancel Un_Diff Un_empty_left assms(4) inf_sup_aci(5) subset_singletonD) have 1: " \ab. ab \ ((\(x, y). x \ y) -` (E \ S)) \ (Pow (V - {x}) \ Pow{x}) \ ab \ Sigma E (\x. {y. (x \ y) \ S}) \ (Pow (V - {x}) \ Pow{x})" using 11 by clarsimp also have 2: " \ab. ab \ Sigma E (\x. {y. (x \ y) \ S}) \ (Pow (V - {x}) \ Pow{x}) \ ab \ ((\(x, y). x \ y) -` (E \ S)) \ (Pow (V - {x}) \ Pow{x})" using 21 by clarsimp ultimately show ?thesis apply (intro antisym) by (meson subsetI)+ qed have eR: "\a. a \ E \ set_pmf (pmf_of_set (Pow (V - {x}))) \ measure_pmf.prob (pmf_of_set (Pow {x})) ({y. even (card ((a \ y) \ \)) = b} \ set_pmf (pmf_of_set (Pow {x}))) = 1/2 " apply (subst measure_Int_set_pmf) apply (intro single_var_parity_coin_flip) subgoal using assms by clarsimp subgoal using assms rev_finite_subset by blast subgoal by (metis 2 IntD2 PowD assms(1) finite_Diff rev_finite_subset) using 2 by blast have " prob_space.prob (pmf_of_set (Pow V)) (E \ {c. even (card (c \ \)) = b}) = prob_space.prob (map_pmf (\(x, y). x \ y) (pair_pmf (pmf_of_set (Pow (V - {x}))) (pmf_of_set (Pow {x})))) (E \ {c. even (card (c \ \)) = b})" unfolding * by auto moreover have "... = prob_space.prob (pair_pmf (pmf_of_set (Pow (V - {x}))) (pmf_of_set (Pow {x}))) ((\(x, y). x \ y) -` (E \ {c. even (card (c \ \)) = b}))" by auto moreover have "... = prob_space.prob (pair_pmf (pmf_of_set (Pow (V - {x}))) (pmf_of_set (Pow {x}))) ((\(x, y). x \ y) -` (E \ {c. even (card (c \ \)) = b}) \ (Pow (V - {x}) \ Pow{x}))" by (smt (verit) "1" "2" Int_iff Sigma_cong measure_pmf.measure_pmf_eq set_pair_pmf) moreover have "... = prob_space.prob (pair_pmf (pmf_of_set (Pow (V - {x}))) (pmf_of_set (Pow {x}))) (Sigma E (\x. {y. even (card ((x \ y) \ \)) = b}) \ (Pow (V - {x}) \ Pow{x}))" unfolding ** by auto moreover have "... = prob_space.prob (pair_pmf (pmf_of_set (Pow (V - {x}))) (pmf_of_set (Pow {x}))) (Sigma E (\x. {y. even (card ((x \ y) \ \)) = b}))" by (smt (verit, best) 1 2 Int_iff Sigma_cong measure_pmf.measure_pmf_eq set_pair_pmf) moreover have "... = measure_pmf.prob (pmf_of_set (Pow (V - {x}))) E * (1 / 2)" apply (subst measure_pmf_prob_dependent_product_bound_eq'[OF _ _ eR]) by auto ultimately show ?thesis by auto qed lemma prob_random_xor_split: assumes V: "finite V" shows "prob_space.prob (random_xor V) E = 1 / 2 * prob_space.prob (pmf_of_set (Pow V)) {e. (e,True) \ E} + 1 / 2 * prob_space.prob (pmf_of_set (Pow V)) {e. (e,False) \ E}" proof - have fin: "finite (set_pmf (random_xor V))" by (simp add: V set_pmf_random_xor) have fin2: "finite ((\(x, y). (y, x)) -` set_pmf (random_xor V))" by(auto intro!: finite_vimageI[OF fin] simp add: inj_def) have rw: "{x. (x, b) \ E \ (x, b) \ set_pmf (random_xor V)} = {x. (x,b) \ E} \ set_pmf (pmf_of_set (Pow V))" for b by (auto simp add: V set_pmf_random_xor Pow_not_empty) have "prob_space.prob (random_xor V) E = prob_space.prob (random_xor V) (E \ set_pmf (random_xor V))" by (simp add: measure_Int_set_pmf) moreover have "... = measure_pmf.prob (pair_pmf (bernoulli_pmf (1 / 2)) (pmf_of_set (Pow V))) ((\(x, y). (y, x)) -` (E \ set_pmf (random_xor V)))" unfolding random_xor_def apply (subst pair_commute_pmf) by simp moreover have "... = 1 / 2 * measure_pmf.prob (pmf_of_set (Pow V)) {x. (x, True) \ E} + 1 / 2 * measure_pmf.prob (pmf_of_set (Pow V)) {x. (x, False) \ E}" unfolding pair_pmf_def apply (subst prob_bernoulli_bind_pmf) using fin2 unfolding map_pmf_def[symmetric] measure_map_pmf by (auto simp add: vimage_def rw simp add: measure_Int_set_pmf) ultimately show ?thesis by auto qed lemma prob_random_xor_nonempty_parity: assumes V: "finite V" assumes \: "x \ \" "\ \ V" assumes E: "\c. c \ E \ (fst c - {x},snd c) \ E" shows "prob_space.prob (random_xor V) (E \ {c. satisfies_xor c \ = b}) = 1 / 2 * prob_space.prob (random_xor (V - {x})) E" proof - have *: "{e. (e, b') \ E \ {c. satisfies_xor c \ = b}} = {e. (e, b') \ E} \ {c. even (card (c \ \)) = (b \ b')}" for b' by (auto simp add: Int_commute) have "prob_space.prob (random_xor V) (E \ {c. satisfies_xor c \ = b}) = 1 / 2 * measure_pmf.prob (pmf_of_set (Pow V)) {e. (e, True) \ E \ {c. satisfies_xor c \ = b}} + 1 / 2 * measure_pmf.prob (pmf_of_set (Pow V)) {e. (e, False) \ E \ {c. satisfies_xor c \ = b}}" unfolding prob_random_xor_split[OF V] by auto also have "... = 1 / 2 * measure_pmf.prob (pmf_of_set (Pow V)) ({e. (e, True) \ E} \ {c. even (card (c \ \)) = (b \ True)}) + 1 / 2 * measure_pmf.prob (pmf_of_set (Pow V)) ({e. (e, False) \ E} \ {c. even (card (c \ \)) = (b \ False)})" unfolding * by auto also have "... = 1 / 2 * (1 / 2 * measure_pmf.prob (pmf_of_set (Pow (V - {x}))) {e. (e, True) \ E} + 1 / 2 * measure_pmf.prob (pmf_of_set (Pow (V - {x}))) {e. (e, False) \ E})" apply (subst prob_pmf_of_set_nonempty_parity[OF V \]) subgoal using E by clarsimp apply (subst prob_pmf_of_set_nonempty_parity[OF V \]) using E by auto also have "... = 1 / 2 * measure_pmf.prob (random_xor (V - {x})) E" apply (subst prob_random_xor_split[symmetric]) using V by auto finally show ?thesis by auto qed lemma pair_satisfies_random_xor_parity_1: assumes V:"finite V" assumes x: "x \ \" "x \ \'" assumes \: "\ \ V" "\' \ V" shows "prob_space.prob (random_xor V) {c. satisfies_xor c \ = b \ satisfies_xor c \' = b'} = 1 / 4" proof - have wa: "\ \ (a - {x}) = \ \ a" for a using x by blast have "prob_space.prob (random_xor V) {c. satisfies_xor c \ = b \ satisfies_xor c \' = b'} = prob_space.prob (random_xor V) ({c. satisfies_xor c \ = b} \ {c. satisfies_xor c \' = b'})" by (simp add: Collect_conj_eq) also have "... = 1 / 2 * measure_pmf.prob (random_xor (V - {x})) {c. satisfies_xor c \ = b}" apply (subst prob_random_xor_nonempty_parity[OF V x(2) \(2)]) by (auto simp add: wa) also have "... = 1/4" apply (subst satisfies_random_xor_parity) using V by auto finally show ?thesis by auto qed lemma pair_satisfies_random_xor_parity: assumes V:"finite V" assumes \: "\ \ \'" "\ \ V" "\' \ V" shows "prob_space.prob (random_xor V) {c. satisfies_xor c \ = b \ satisfies_xor c \' = b'} = 1 / 4" proof - obtain x where "x \ \ \ x \ \' \ x \ \' \ x \ \" using \ by blast moreover { assume x: "x \ \" "x \ \'" have ?thesis using pair_satisfies_random_xor_parity_1[OF V x \(2-3)] by blast } moreover { assume x: "x \ \'" "x \ \" then have ?thesis using pair_satisfies_random_xor_parity_1[OF V x \(3) \(2)] by (simp add: Collect_conj_eq Int_commute) } ultimately show ?thesis by auto qed lemma prob_pmf_of_set_nonempty_parity_UNIV: assumes "finite V" assumes "x \ \" "\ \ V" shows "prob_space.prob (pmf_of_set (Pow V)) {c. even (card (c \ \)) = b} = 1 / 2" using prob_pmf_of_set_nonempty_parity[OF assms, of UNIV] by auto lemma prob_Pow_split: assumes "\ \ V" "finite V" shows "prob_space.prob (pmf_of_set (Pow V)) {x. P (\ \ x) \ Q ((V - \) \ x)} = prob_space.prob (pmf_of_set (Pow \)) {x. P x} * prob_space.prob (pmf_of_set (Pow (V - \))) {x. Q x}" proof - have 1: "set_pmf (pmf_of_set (Pow \)) = Pow \" by (meson Pow_not_empty assms(1) assms(2) finite_Pow_iff finite_subset set_pmf_of_set) have 2: "set_pmf (pmf_of_set (Pow (V - \))) = Pow (V - \ )" by (simp add: Pow_not_empty assms(2)) have *: "(pmf_of_set (Pow \ \ Pow (V - \))) = (pair_pmf (pmf_of_set (Pow \)) (pmf_of_set (Pow (V - \))))" unfolding 1 2 apply (subst pmf_of_set_prod) using assms rev_finite_subset by auto have **: "(((\(x, y). x \ y) -` {x. P (\ \ x) \ Q ((V - \) \ x)}) \ ((Pow \) \ (Pow (V - \)))) = ({x. P x} \ Pow \) \ ({x. Q x} \ Pow (V - \))" apply (rule antisym) subgoal apply clarsimp by (smt (verit) Diff_disjoint Int_Un_eq(4) inf.orderE inf_commute inf_sup_distrib1 sup_bot.right_neutral sup_commute) apply (intro subsetI) apply clarsimp by (smt (verit, ccfv_threshold) Diff_Int Diff_disjoint Diff_empty Diff_eq_empty_iff Un_Int_assoc_eq Un_commute sup_bot.left_neutral) have "prob_space.prob (pmf_of_set (Pow V)) {x. P (\ \ x) \ Q ((V - \) \ x)} = prob_space.prob (map_pmf (\(x,y). x \ y) (pmf_of_set (Pow \ \ Pow (V - \)))) {x. P (\ \ x) \ Q ((V - \) \ x)}" apply (subst pmf_of_set_powerset_split[symmetric, OF assms(1-2)]) by auto moreover have "... = measure_pmf.prob (pair_pmf (pmf_of_set (Pow \)) (pmf_of_set (Pow (V - \)))) ((\(x, y). x \ y) -` {x. P (\ \ x) \ Q ((V - \) \ x)})" unfolding measure_map_pmf using * by presburger moreover have "... = measure_pmf.prob (pair_pmf (pmf_of_set (Pow \)) (pmf_of_set (Pow (V - \)))) (((\(x, y). x \ y) -` {x. P (\ \ x) \ Q ((V - \) \ x)}) \ ((Pow \) \ (Pow (V - \))))" using 1 2 by (smt (verit) Int_Collect Int_def Sigma_cong inf_idem measure_pmf.measure_pmf_eq set_pair_pmf) moreover have "... = measure_pmf.prob (pair_pmf (pmf_of_set (Pow \)) (pmf_of_set (Pow (V - \)))) (({x. P x} \ Pow \) \ ({x. Q x} \ Pow (V - \)))" unfolding ** by auto moreover have "... = measure_pmf.prob (pmf_of_set (Pow \)) ({x. P x} \ Pow \) * measure_pmf.prob (pmf_of_set (Pow (V - \))) ({x. Q x} \ Pow (V - \))" apply (intro measure_pmf_prob_product) subgoal by (meson assms(1) assms(2) countable_finite finite_Int finite_Pow_iff rev_finite_subset) by (simp add: assms(2) countable_finite) moreover have "... = measure_pmf.prob (pmf_of_set (Pow \)) {x. P x} * measure_pmf.prob (pmf_of_set (Pow (V - \))) {x. Q x}" by (metis "1" "2" measure_Int_set_pmf) ultimately show ?thesis by auto qed (* Split probability for two disjoint non-empty sets *) lemma disjoint_prob_pmf_of_set_nonempty: assumes \: "x \ \" "\ \ V" assumes \': "x' \ \'" "\' \ V" assumes "\ \ \' = {}" assumes V: "finite V" shows "prob_space.prob (pmf_of_set (Pow V)) {c. even (card (\ \ c)) = b \ even (card (\' \ c)) = b'} = 1 / 4" proof - have "prob_space.prob (pmf_of_set (Pow V)) {c. even (card (\ \ c)) = b \ even (card (\' \ c)) = b'} = prob_space.prob (pmf_of_set (Pow V)) {c. even (card (\ \ c)) = b \ even (card (((V - \) \ c) \ \')) = b'}" by (smt (verit) Collect_cong Diff_Diff_Int Diff_eq_empty_iff Int_Diff Int_commute \'(2) assms(5) inf.orderE) moreover have "... = measure_pmf.prob (pmf_of_set (Pow \)) {x. even (card x) = b} * measure_pmf.prob (pmf_of_set (Pow (V - \))) {x. even (card (x \ \')) = b'}" apply (subst prob_Pow_split) using assms by auto moreover have "... = measure_pmf.prob (pmf_of_set (Pow \)) {x. even (card (x \ \)) = b} * measure_pmf.prob (pmf_of_set (Pow (V - \))) {x. even (card (x \ \')) = b'} " by (smt (verit, best) PowD Pow_not_empty \(2) assms(6) finite_Pow_iff inf.orderE measure_pmf.measure_pmf_eq mem_Collect_eq rev_finite_subset set_pmf_of_set) moreover have "... = 1/4" apply (subst prob_pmf_of_set_nonempty_parity_UNIV[OF _ \(1)]) subgoal using assms rev_finite_subset by blast subgoal by simp apply (subst prob_pmf_of_set_nonempty_parity_UNIV) using assms by auto ultimately show ?thesis by auto qed lemma measure_pmf_prob_product_finite_set_pmf: assumes "finite (set_pmf M)" "finite (set_pmf N)" shows "measure_pmf.prob (pair_pmf M N) (A \ B) = measure_pmf.prob M A * measure_pmf.prob N B" proof - have A: "measure_pmf.prob M A = measure_pmf.prob M (A \ set_pmf M)" by (simp add: measure_Int_set_pmf) have B: "measure_pmf.prob N B = measure_pmf.prob N (B \ set_pmf N)" by (simp add: measure_Int_set_pmf) have "measure_pmf.prob M A * measure_pmf.prob N B = measure_pmf.prob M (A \ set_pmf M) * measure_pmf.prob N (B \ set_pmf N)" using A B by auto moreover have "... = measure_pmf.prob (pair_pmf M N) ((A \ set_pmf M) \ (B \ set_pmf N))" apply (subst measure_pmf_prob_product[symmetric]) by auto moreover have "... = measure_pmf.prob (pair_pmf M N) ((A \ B) \ set_pmf (pair_pmf M N))" by (simp add: Times_Int_Times) moreover have "... = measure_pmf.prob (pair_pmf M N) ((A \ B) )" using measure_Int_set_pmf by blast ultimately show ?thesis by auto qed lemma prob_random_xor_split_space: assumes "\ \ V" "finite V" shows "prob_space.prob (random_xor V) {(x,b). P (\ \ x) b \ Q ((V - \) \ x)} = prob_space.prob (random_xor \) {(x,b). P x b} * prob_space.prob (pmf_of_set (Pow (V - \))) {x. Q x}" proof - have d1: "set_pmf (random_xor \) = Pow \ \ UNIV" by (metis assms(1) assms(2) infinite_super set_pmf_random_xor) have d2: "set_pmf (pmf_of_set (Pow (V - \))) = Pow (V- \)" by (simp add: Pow_not_empty assms(2)) have rhs: "prob_space.prob (random_xor \) {(x,b). P x b} * prob_space.prob (pmf_of_set (Pow (V - \))) {x. Q x} = prob_space.prob (pair_pmf (random_xor \) (pmf_of_set (Pow (V - \)))) ({(x,b). P x b} \ {x. Q x})" apply (subst measure_pmf_prob_product_finite_set_pmf) subgoal by (metis Pow_def assms(1) assms(2) finite_Collect_subsets finite_SigmaI finite_code rev_finite_subset set_pmf_random_xor) using assms by (auto simp add: Pow_not_empty) from pmf_of_set_powerset_split[OF assms] have *: "pmf_of_set (Pow V) = map_pmf (\(x, y). x \ y) (pair_pmf (pmf_of_set (Pow \)) (pmf_of_set (Pow (V - \))))" by (metis Pow_not_empty assms(1) assms(2) finite_Diff finite_Pow_iff pmf_of_set_prod rev_finite_subset) have **: "random_xor V = map_pmf (\((x,b),y). (x \ y,b)) (pair_pmf (random_xor \) (pmf_of_set (Pow (V - \))))" unfolding random_xor_def * including monad_normalisation by (auto simp add: pair_pmf_def map_pmf_def case_prod_unfold) have "prob_space.prob (random_xor V) {(x,b). P (\ \ x) b \ Q ((V - \) \ x)} = measure_pmf.prob (pair_pmf (random_xor \) (pmf_of_set (Pow (V - \)))) {y. P (\ \ (fst (fst y) \ snd y)) (snd (fst y)) \ Q ((V - \) \ (fst (fst y) \ snd y))}" unfolding ** by (auto simp add:case_prod_unfold) moreover have "... = prob_space.prob (pair_pmf (random_xor \) (pmf_of_set (Pow (V - \)))) ({(x,b). P x b} \ {x. Q x})" apply (intro measure_pmf.measure_pmf_eq[where p ="pair_pmf (random_xor \) (pmf_of_set (Pow (V - \)))"]) subgoal by simp apply (clarsimp simp add: d1 d2) by (smt (verit, del_insts) Diff_disjoint Int_Diff boolean_algebra_cancel.sup0 inf.orderE inf_commute inf_sup_distrib1 sup_bot.left_neutral) ultimately show ?thesis using rhs by simp qed lemma three_disjoint_prob_random_xor_nonempty: assumes \: "\ \ {}" "\ \ V" assumes \': "\' \ {}" "\' \ V" assumes I: "I \ V" assumes int: "I \ \ = {}" "I \ \' = {}" "\ \ \' = {}" assumes V: "finite V" shows "prob_space.prob (random_xor V) {c. satisfies_xor c I = b \ even (card (\ \ fst c)) = b' \ even (card (\' \ fst c)) = b''} = 1 / 8" proof - have finI: "finite I" using V I using rev_finite_subset by blast have finVI: "finite (V - I)" using V I using rev_finite_subset by blast obtain x x' where x: "x \ \" "x' \ \'" using \ \' by blast have rw1:"\ \ ((V - I) \ xx) = \ \ xx " for xx by (metis Diff_Int_distrib2 Diff_empty \(2) assms(6) inf.absorb_iff2 inf_assoc inf_left_commute) have rw2:"\' \ ((V - I) \ xx)= \' \ xx" for xx by (metis Diff_Int_distrib2 Diff_empty \'(2) assms(7) inf.absorb_iff2 inf_assoc inf_left_commute) have "prob_space.prob (random_xor V) {c. satisfies_xor c I = b \ even (card ( \ \ fst c )) = b' \ even (card (\' \ fst c )) = b''} = prob_space.prob (random_xor V) {(x,bb). satisfies_xor (I \ x, bb) I = b \ even (card (\ \ ((V - I) \ x))) = b' \ even (card ( \' \ ((V - I) \ x))) = b''}" apply (intro arg_cong[where f = "prob_space.prob (random_xor V)"]) unfolding rw1 rw2 satisfies_xor_inter by (smt (verit) Collect_cong prod.collapse split_conv) moreover have "... = measure_pmf.prob (random_xor I) {c. satisfies_xor c I = b} * measure_pmf.prob (pmf_of_set (Pow (V - I))) {x. even (card (\ \ x)) = b' \ even (card (\' \ x)) = b''}" apply (subst prob_random_xor_split_space[OF I V]) by (metis (no_types, lifting) Collect_cong case_prodE case_prodI2) moreover have "... = 1 / 8" apply (subst satisfies_random_xor_parity[OF finI]) apply (subst disjoint_prob_pmf_of_set_nonempty) using x \ \' int finVI by auto ultimately show ?thesis by auto qed (* Split probability for three disjoint non-empty sets *) lemma three_disjoint_prob_pmf_of_set_nonempty: assumes \: "x \ \" "\ \ V" assumes \': "x' \ \'" "\' \ V" assumes \'': "x'' \ \''" "\'' \ V" assumes int: "\ \ \' = {}" "\' \ \'' = {}" "\'' \ \ = {}" assumes V: "finite V" shows "prob_space.prob (pmf_of_set (Pow V)) {c. even (card (\ \ c)) = b \ even (card (\' \ c)) = b' \ even (card (\'' \ c)) = b''} = 1 / 8" proof - have "prob_space.prob (pmf_of_set (Pow V)) {c. even (card (\ \ c)) = b \ even (card (\' \ c)) = b' \ even (card (\'' \ c)) = b''} = prob_space.prob (pmf_of_set (Pow V)) {c. even (card (\ \ c)) = b \ even (card (((V - \) \ c) \ \')) = b' \ even (card (((V - \) \ c) \ \'')) = b''}" by (smt (verit,best) Collect_cong Diff_Diff_Int Diff_eq_empty_iff Int_Diff Int_commute \' \'' int inf.orderE) moreover have "... = measure_pmf.prob (pmf_of_set (Pow \)) {x. even (card x) = b} * measure_pmf.prob (pmf_of_set (Pow (V - \))) {x. even (card (\' \ x)) = b' \ even (card (\'' \ x)) = b''}" apply (subst prob_Pow_split) using assms by (auto simp add: inf.commute) moreover have "... = measure_pmf.prob (pmf_of_set (Pow \)) {x. even (card (x \ \)) = b} * measure_pmf.prob (pmf_of_set (Pow (V - \))) {x. even (card (\' \ x)) = b'\ even (card (\'' \ x)) = b''} " by (smt (verit, best) PowD Pow_not_empty \ V finite_Pow_iff inf.orderE measure_pmf.measure_pmf_eq mem_Collect_eq rev_finite_subset set_pmf_of_set) moreover have "... = 1/8" apply (subst prob_pmf_of_set_nonempty_parity_UNIV[OF _ \(1)]) subgoal using \(2) V rev_finite_subset by blast subgoal by simp apply (subst disjoint_prob_pmf_of_set_nonempty) using assms by auto ultimately show ?thesis by auto qed lemma four_disjoint_prob_random_xor_nonempty: assumes \: "\ \ {}" "\ \ V" assumes \': "\' \ {}" "\' \ V" assumes \'': "\'' \ {}" "\'' \ V" assumes I: "I \ V" assumes int: "I \ \ = {}" "I \ \' = {}" "I \ \'' = {}" "\ \ \' = {}" "\' \ \'' = {}" "\'' \ \ = {}" assumes V: "finite V" shows "prob_space.prob (random_xor V) {c. satisfies_xor c I = b0 \ even (card (\ \ fst c)) = b \ even (card (\' \ fst c)) = b' \ even (card (\'' \ fst c)) = b''} = 1 / 16" proof - have finI: "finite I" using V I using rev_finite_subset by blast have finVI: "finite (V - I)" using V I using rev_finite_subset by blast obtain x x' x'' where x: "x \ \" "x' \ \'" "x'' \ \''" using \ \' \'' by blast have rw1:"\ \ ((V - I) \ xx) = \ \ xx " for xx by (metis Diff_Int_distrib2 Diff_empty \(2) int(1) inf.absorb_iff2 inf_assoc inf_left_commute) have rw2:"\' \ ((V - I) \ xx)= \' \ xx" for xx by (metis Diff_Int_distrib2 Diff_empty \'(2) int(2) inf.absorb_iff2 inf_assoc inf_left_commute) have rw3:"\'' \ ((V - I) \ xx)= \'' \ xx" for xx by (metis Diff_Int_distrib2 Diff_empty \''(2) int(3) inf.absorb_iff2 inf_assoc inf_left_commute) have "prob_space.prob (random_xor V) {c. satisfies_xor c I = b0 \ even (card (\ \ fst c)) = b \ even (card (\' \ fst c)) = b' \ even (card (\'' \ fst c)) = b''} = prob_space.prob (random_xor V) {(x,bb). satisfies_xor (I \ x, bb) I = b0 \ even (card (\ \ ((V - I) \ x))) = b \ even (card (\' \ ((V - I) \ x))) = b' \ even (card (\'' \ ((V - I) \ x))) = b''}" apply (intro arg_cong[where f = "prob_space.prob (random_xor V)"]) unfolding rw1 rw2 rw3 satisfies_xor_inter by (smt (verit) Collect_cong prod.collapse split_conv) moreover have "... = measure_pmf.prob (random_xor I) {c. satisfies_xor c I = b0} * measure_pmf.prob (pmf_of_set (Pow (V - I))) {x. even (card (\ \ x)) = b \ even (card (\' \ x)) = b' \ even (card (\'' \ x)) = b''}" apply (subst prob_random_xor_split_space[OF I V]) by (metis (no_types, lifting) Collect_cong case_prodE case_prodI2) moreover have "... = 1 / 16" apply(subst satisfies_random_xor_parity[OF finI]) apply (subst three_disjoint_prob_pmf_of_set_nonempty) using x \ \' \'' int finVI by auto ultimately show ?thesis by auto qed lemma three_satisfies_random_xor_parity_1: assumes V:"finite V" assumes \: "\ \ V" "\' \ V" "\'' \ V" assumes x: "x \ \" "x \ \'" "x \ \''" assumes d: "\ \ \'" shows "prob_space.prob (random_xor V) {c. satisfies_xor c \ = b \ satisfies_xor c \' = b' \ satisfies_xor c \'' = b''} = 1 / 8" proof - have wa: "\ \ (a - {x}) = \ \ a" for a using x by blast have wa': "\' \ (a - {x}) = \' \ a" for a using x by blast have "prob_space.prob (random_xor V) {c. satisfies_xor c \ = b \ satisfies_xor c \' = b' \ satisfies_xor c \'' = b''} = prob_space.prob (random_xor V) ({c. satisfies_xor c \ = b \ satisfies_xor c \' = b'} \ {c. satisfies_xor c \'' = b''})" by (simp add: Collect_conj_eq inf_assoc) moreover have "... = 1 / 2 * measure_pmf.prob (random_xor (V - {x})) {c. satisfies_xor c \ = b \ satisfies_xor c \' = b'}" apply (subst prob_random_xor_nonempty_parity[OF V x(3) \(3)]) by (auto simp add: wa wa') moreover have "... = 1/8" apply (subst pair_satisfies_random_xor_parity) using V \ x d by auto ultimately show ?thesis by auto qed lemma split_boolean_eq: shows"(A \ B) = (b \ I) \ (B \ C) = (b' \ I) \ (C \ A) = (b'' \ I) \ I = odd(of_bool b + of_bool b' + of_bool b'') \ (A = True \ B = (b'=b'') \ C = (b = b') \ A = False \ B = (b' \ b'') \ C = (b \ b'))" by auto lemma three_satisfies_random_xor_parity: assumes V:"finite V" assumes \: "\ \ \'" "\ \ \''" "\' \ \''" "\ \ V" "\' \ V" "\'' \ V" shows "prob_space.prob (random_xor V) {c. satisfies_xor c \ = b \ satisfies_xor c \' = b' \ satisfies_xor c \'' = b''} = 1 / 8" proof - have "(\x. x \ \ \ x \ \' \ x \ \'' \ x \ \' \ x \ \ \ x \ \'' \ x \ \'' \ x \ \ \ x \ \') \ \ - (\' \ \'') = {} \ \' - (\ \ \'') = {} \ \'' - (\ \ \') = {}" by blast moreover { assume "(\x. x \ \ \ x \ \' \ x \ \'' \ x \ \' \ x \ \ \ x \ \'' \ x \ \'' \ x \ \ \ x \ \')" then obtain x where " x \ \ \ x \ \' \ x \ \'' \ x \ \' \ x \ \ \ x \ \'' \ x \ \'' \ x \ \ \ x \ \'" by auto moreover { assume x: "x \ \'" "x \ \''" "x \ \" have "measure_pmf.prob (random_xor V) {c. satisfies_xor c \' = b' \ satisfies_xor c \'' = b'' \ satisfies_xor c \ = b} = 1 / 8" apply (intro three_satisfies_random_xor_parity_1[OF V _ _ _ x]) using \ by auto then have ?thesis by (smt (verit, ccfv_SIG) Collect_cong) } moreover { assume x: "x \ \" "x \ \''" "x \ \'" have "measure_pmf.prob (random_xor V) {c. satisfies_xor c \ = b \ satisfies_xor c \'' = b'' \ satisfies_xor c \' = b'} = 1 / 8" apply (intro three_satisfies_random_xor_parity_1[OF V _ _ _ x]) using \ by auto then have ?thesis by (smt (verit, ccfv_SIG) Collect_cong) } moreover { assume x: "x \ \" "x \ \'" "x \ \''" have "measure_pmf.prob (random_xor V) {c. satisfies_xor c \ = b \ satisfies_xor c \' = b' \ satisfies_xor c \'' = b''} = 1 / 8" apply (intro three_satisfies_random_xor_parity_1[OF V _ _ _ x]) using \ by auto then have ?thesis by (smt (verit, ccfv_SIG) Collect_cong) } ultimately have ?thesis by auto } moreover { assume dis: "\ - (\' \ \'') = {} \ \' - (\ \ \'') = {} \ \'' - (\ \ \') = {}" define A where "A = (\ \ \'') - \'" define B where "B = (\ \ \') - \''" define C where "C = (\' \ \'') - \" define I where "I = \ \ \' \ \''" have f: "finite A" "finite B" "finite C" "finite I" unfolding A_def B_def C_def I_def by (meson V \ finite_Diff finite_Int rev_finite_subset)+ have s: "A \ V" "B \ V" "C \ V" "I \ V" unfolding A_def B_def C_def I_def using \ by auto have i: "I \ A = {}" "I \ B = {}" "I \ C = {}" "B \ C = {}" "C \ A = {}" "A \ B = {}" unfolding A_def B_def C_def I_def by blast + have s1: "\ = A \ B \ I" unfolding A_def B_def I_def using dis by auto have sx1: "satisfies_xor (xx,bb) \ = b \ even (card (A \ xx)) = even (card (B \ xx)) = (b \ satisfies_xor(xx,bb) I)" for xx bb unfolding s1 satisfies_xor.simps Int_Un_distrib2 apply (subst card_Un_disjoint) subgoal using f by auto subgoal using f by auto subgoal using A_def B_def I_def by blast apply (subst card_Un_disjoint) subgoal using f by auto subgoal using f by auto subgoal using A_def B_def I_def by blast by auto have s2: "\' = B \ C \ I" unfolding B_def C_def I_def using dis by auto have sx2: "satisfies_xor (xx,bb) \' = b' \ even (card (B \ xx)) = even (card (C \ xx)) = (b' \ satisfies_xor(xx,bb) I)" for xx bb unfolding s2 satisfies_xor.simps Int_Un_distrib2 apply (subst card_Un_disjoint) subgoal using f by auto subgoal using f by auto subgoal using B_def C_def I_def by blast apply (subst card_Un_disjoint) subgoal using f by auto subgoal using f by auto subgoal using B_def C_def I_def by blast by auto have s3: "\'' = A \ C \ I" unfolding A_def C_def I_def using dis by auto have sx3: "satisfies_xor (xx,bb) \'' = b'' \ even (card (C \ xx)) = even (card (A \ xx)) = (b'' \ satisfies_xor(xx,bb) I)" for xx bb unfolding s3 satisfies_xor.simps Int_Un_distrib2 apply (subst card_Un_disjoint) subgoal using f by auto subgoal using f by auto subgoal using A_def B_def C_def I_def by blast apply (subst card_Un_disjoint) subgoal using f by auto subgoal using f by auto subgoal using A_def B_def C_def I_def by blast by auto have "A = {} \ B \ {} \ C \ {} \ A \ {} \ B = {} \ C \ {} \ A \ {} \ B \ {} \ C = {} \ A \ {} \ B \ {} \ C \ {}" by (metis Un_commute \(1) \(2) \(3) s1 s2 s3) moreover { assume as: "A = {}" "B \ {}" "C \ {}" have "satisfies_xor (xx,bb) \ = b \ satisfies_xor (xx,bb) \' = b' \ satisfies_xor (xx,bb) \'' = b'' \ (satisfies_xor (xx, bb) I = odd (of_bool b + of_bool b' + of_bool b'') \ (even (card (B \ xx)) = (b' = b'') \ even (card (C \ xx)) = (b = b')))" for xx bb unfolding sx1 sx2 sx3 apply(subst split_boolean_eq) using as by simp then have *: "{c. satisfies_xor c \ = b \ satisfies_xor c \' = b' \ satisfies_xor c \'' = b''} = {c. satisfies_xor c I = odd (of_bool b + of_bool b' + of_bool b'') \ even (card (B \ fst c)) = (b' = b'') \ even (card (C \ fst c)) = (b = b')}" by (smt (verit,best) Collect_cong prod.collapse) have ?thesis apply (subst *) apply (intro three_disjoint_prob_random_xor_nonempty) using as s i V by auto } moreover { assume as: "A \ {}" "B \ {}" "C = {}" have "satisfies_xor (xx,bb) \'' = b'' \ satisfies_xor (xx,bb) \ = b \ satisfies_xor (xx,bb) \' = b' \ (satisfies_xor (xx, bb) I = odd (of_bool b'' + of_bool b + of_bool b') \ (even (card (A \ xx)) = (b = b') \ even (card (B \ xx)) = (b'' = b)))" for xx bb unfolding sx1 sx2 sx3 apply(subst split_boolean_eq) using as by simp then have *: "{c. satisfies_xor c \ = b \ satisfies_xor c \' = b' \ satisfies_xor c \'' = b''} = {c. satisfies_xor c I = odd (of_bool b'' + of_bool b + of_bool b') \ (even (card (A \ fst c)) = (b = b') \ even (card (B \ fst c)) = (b'' = b))}" by (smt (verit,best) Collect_cong prod.collapse) have ?thesis apply (subst *) apply (intro three_disjoint_prob_random_xor_nonempty) using as s i V by auto } moreover { assume as: "A \ {}" "B = {}" "C \ {}" have "satisfies_xor (xx,bb) \' = b' \ satisfies_xor (xx,bb) \'' = b'' \ satisfies_xor (xx,bb) \ = b \ (satisfies_xor (xx, bb) I = odd (of_bool b' + of_bool b'' + of_bool b) \ (even (card (C \ xx)) = (b'' = b) \ even (card (A \ xx)) = (b' = b'')))" for xx bb unfolding sx1 sx2 sx3 apply(subst split_boolean_eq) using as by simp then have *: "{c. satisfies_xor c \ = b \ satisfies_xor c \' = b' \ satisfies_xor c \'' = b''} = {c. satisfies_xor c I = odd (of_bool b' + of_bool b'' + of_bool b) \ (even (card (C \ fst c)) = (b'' = b) \ even (card (A \ fst c)) = (b' = b''))}" by (smt (verit,best) Collect_cong prod.collapse) have ?thesis apply (subst *) apply (intro three_disjoint_prob_random_xor_nonempty) using as s i V by auto } moreover { assume as: "A \ {}" "B \ {}" "C \ {}" have 1: "satisfies_xor (xx,bb) \ = b \ satisfies_xor (xx,bb) \' = b' \ satisfies_xor (xx,bb) \'' = b'' \ (satisfies_xor (xx, bb) I = odd (of_bool b + of_bool b' + of_bool b'') \ even (card (A \ xx)) = True \ even (card (B \ xx)) = (b' = b'') \ even (card (C \ xx)) = (b = b') \ satisfies_xor (xx, bb) I = odd (of_bool b + of_bool b' + of_bool b'') \ even (card (A \ xx)) = False \ even (card (B \ xx)) = (b' \ b'') \ even (card (C \ xx)) = (b \ b'))" for xx bb unfolding sx1 sx2 sx3 apply(subst split_boolean_eq) by auto have 2: "satisfies_xor c \ = b \ satisfies_xor c \' = b' \ satisfies_xor c \'' = b'' \ (satisfies_xor c I = odd (of_bool b + of_bool b' + of_bool b'') \ even (card (A \ fst c)) = True \ even (card (B \ fst c)) = (b' = b'') \ even (card (C \ fst c)) = (b = b') \ satisfies_xor c I = odd (of_bool b + of_bool b' + of_bool b'') \ even (card (A \ fst c)) = False \ even (card (B \ fst c)) = (b' \ b'') \ even (card (C \ fst c)) = (b \ b'))" for c proof - obtain xx bb where c:"c = (xx,bb)" by fastforce show ?thesis unfolding c apply (subst 1) by auto qed have *: "{c. satisfies_xor c \ = b \ satisfies_xor c \' = b' \ satisfies_xor c \'' = b''} = {c. satisfies_xor c I = odd (of_bool b + of_bool b' + of_bool b'') \ even (card (A \ fst c)) = True \ even (card (B \ fst c)) = (b' = b'') \ even (card (C \ fst c)) = (b = b')} \ {c. satisfies_xor c I = odd (of_bool b + of_bool b' + of_bool b'') \ even (card (A \ fst c)) = False \ even (card (B \ fst c)) = (b' \ b'') \ even (card (C \ fst c)) = (b \ b')}" apply (subst Un_def) apply (intro Collect_cong) apply (subst 2) by simp have **: "1 / 16 + measure_pmf.prob (random_xor V) {c. satisfies_xor c I = odd (of_bool b + of_bool b' + of_bool b'') \ even (card (A \ fst c)) = False \ even (card (B \ fst c)) = (b' \ b'') \ even (card (C \ fst c)) = (b \ b')} = 1 / 8" apply (subst four_disjoint_prob_random_xor_nonempty) using as s i V by auto have ?thesis apply (subst *) apply (subst measure_pmf.finite_measure_Union) subgoal by simp subgoal by simp subgoal by auto apply (subst four_disjoint_prob_random_xor_nonempty) using as s i V ** by auto } ultimately have ?thesis by auto } ultimately show ?thesis by auto qed subsection "Independence for repeated XORs" text \ We can lift the previous result to a list of independent sampled XORs. \ definition random_xors :: "'a set \ nat \ (nat \ 'a set \ bool) pmf" where "random_xors V n = Pi_pmf {..<(n::nat)} None (\_. map_pmf Some (random_xor V))" (* The set of random XORs expressed in various ways*) lemma random_xors_set: assumes V:"finite V" shows "PiE_dflt {.. (\_. map_pmf Some (random_xor V))) = {xors. dom xors = {.. ran xors \ (Pow V) \ UNIV}" (is "?lhs = ?rhs") proof - have "?lhs = {f. dom f = {.. (\x \ {.. Some ` (Pow V \ UNIV))}" unfolding PiE_dflt_def o_def set_map_pmf set_pmf_random_xor[OF V] by force also have "... = ?rhs" apply (rule antisym) subgoal apply clarsimp by (smt (z3) PowD domI image_iff mem_Collect_eq mem_Sigma_iff option.simps(1) ran_def subsetD) apply clarsimp by (smt (verit, ccfv_threshold) domD image_iff lessThan_iff ranI subsetD) finally show ?thesis . qed lemma random_xors_eq: assumes V:"finite V" shows"random_xors V n = pmf_of_set {xors. dom xors = {.. ran xors \ (Pow V) \ UNIV}" proof - have "pmf_of_set {xors. dom xors = {.. ran xors \ (Pow V) \ UNIV} = pmf_of_set (PiE_dflt {.. (\_. map_pmf Some (random_xor V))))" unfolding random_xors_set[OF V] by auto also have "... = Pi_pmf {..x. pmf_of_set ((set_pmf \ (\_. map_pmf Some (random_xor V))) x))" apply (subst Pi_pmf_of_set[symmetric]) by (auto simp add:set_pmf_random_xor[OF V] V) also have "... = random_xors V n" unfolding random_xors_def o_def set_map_pmf apply (subst map_pmf_of_set_inj[symmetric]) subgoal by (auto simp add:set_pmf_random_xor[OF V] V) subgoal by (auto simp add:set_pmf_random_xor[OF V] V) subgoal by (auto simp add:set_pmf_random_xor[OF V] V) by (metis V random_xor_pmf_of_set set_pmf_random_xor) ultimately show ?thesis by auto qed (* The XOR hash function from a mapping nat \ XOR *) definition xor_hash :: "('a \ bool) \ (nat \ ('a set \ bool)) \ (nat \ bool)" where "xor_hash \ xors = (map_option (\xor. satisfies_xor xor {x. \ x = Some True}) \ xors)" lemma finite_map_set_nonempty: assumes "R \ {}" shows " {xors. dom xors = D \ ran xors \ R} \ {}" proof - obtain r where "r \ R" using assms by blast then have "(\x. if x \ D then Some r else None) \ {xors. dom xors = D \ ran xors \ R}" by (auto split:if_splits simp:ran_def) thus ?thesis by auto qed lemma random_xors_set_pmf: assumes V: "finite V" shows " set_pmf (random_xors V n) = {xors. dom xors = {.. ran xors \ (Pow V) \ UNIV}" unfolding random_xors_eq[OF V] apply (intro set_pmf_of_set) subgoal apply (intro finite_map_set_nonempty) by blast apply (intro finite_set_of_finite_maps) by (auto simp add: V) lemma finite_random_xors_set_pmf: assumes V: "finite V" shows " finite (set_pmf (random_xors V n))" unfolding random_xors_set_pmf[OF V] by (auto intro!: finite_set_of_finite_maps simp add: V) lemma map_eq_1: assumes "dom f = dom g" assumes "\x. x \ dom f \ the (f x) = the (g x)" shows "f = g" by (metis assms(1) assms(2) domIff map_le_antisym map_le_def option.expand) lemma xor_hash_eq_iff: assumes "dom \ = {.. x = \ \ (dom x = {.. (\i. i < n \ (\xor. x i = Some xor \ satisfies_xor xor {x. \ x = Some True} = the (\ i)) ))" proof - have 1: "xor_hash \ x = \ \ (dom (xor_hash \ x) = dom \) \ (\i \ dom \. the (xor_hash \ x i) = the (\ i))" using map_eq_1 by fastforce have 2: "dom (xor_hash \ x) = dom x" unfolding xor_hash_def by auto have 3: "\i. i \ dom x \ xor_hash \ x i = Some (satisfies_xor (the (x i)) {x. \ x = Some True})" unfolding xor_hash_def by fastforce show ?thesis unfolding 1 assms 2 using 3 by (smt (verit, best) domD lessThan_iff option.sel) qed lemma xor_hash_eq_PiE_dflt: assumes "dom \ = {.. xors = \} = PiE_dflt {..i. Some ` {xor. satisfies_xor xor {x. \ x = Some True} = the (\ i)})" proof - have *: "\x xa a b. (\ xa < n \ x xa = None) \ x xa = Some (a, b) \ xa < n" by (metis option.distinct(2)) show ?thesis unfolding PiE_dflt_def unfolding xor_hash_eq_iff[OF assms] by (auto intro: * simp del: satisfies_xor.simps) qed lemma prob_random_xors_xor_hash: assumes V: "finite V" assumes \: "dom \ = {.. xors = \} = 1 / 2 ^ n" proof - have "measure_pmf.prob (random_xors V n) {xors. xor_hash \ xors = \} = measure_pmf.prob (Pi_pmf {..<(n::nat)} None (\_. map_pmf Some (random_xor V))) (PiE_dflt {..i. Some ` {xor. satisfies_xor xor {x. \ x = Some True} = the (\ i)}))" unfolding random_xors_def xor_hash_eq_PiE_dflt[OF \] by auto also have "... = (\x x = Some True} = the (\ x)}))" by (subst measure_Pi_pmf_PiE_dflt) (auto simp add: inj_vimage_image_eq) also have "... = (\x PiE_dflt A dflt B' = PiE_dflt A dflt (\b. B b \ B' b)" unfolding PiE_dflt_def by auto lemma random_xors_xor_hash_pair: assumes V: "finite V" assumes \: "dom \ = {..': "dom \' = {..: "dom \ = V" assumes \': "dom \' = V" assumes neq: "\ \ \'" shows " measure_pmf.prob (random_xors V n) {xors. xor_hash \ xors = \ \ xor_hash \' xors = \'} = 1 / 4 ^ n" proof - obtain y where "\ y \ \' y" using neq by blast then have neq:"{x. \ x = Some True} \ {x. \' x = Some True}" by (smt (verit, ccfv_threshold) assms(4) assms(5) domD domIff mem_Collect_eq) have "measure_pmf.prob (random_xors V n) {xors. xor_hash \ xors = \ \ xor_hash \' xors = \'} = measure_pmf.prob (random_xors V n) ({xors. xor_hash \ xors = \} \ {xors. xor_hash \' xors = \'})" by (simp add: Collect_conj_eq) also have "... = measure_pmf.prob (Pi_pmf {..<(n::nat)} None (\_. map_pmf Some (random_xor V))) (PiE_dflt {..i. Some ` {xor. satisfies_xor xor {x. \ x = Some True} = the (\ i)}) \ PiE_dflt {..i. Some ` {xor. satisfies_xor xor {x. \' x = Some True} = the (\' i)}))" unfolding random_xors_def xor_hash_eq_PiE_dflt[OF \] xor_hash_eq_PiE_dflt[OF \'] by auto also have "... = measure_pmf.prob (Pi_pmf {..<(n::nat)} None (\_. map_pmf Some (random_xor V))) (PiE_dflt {..i. Some ` {xor. satisfies_xor xor {x. \ x = Some True} = the (\ i) \ satisfies_xor xor {x. \' x = Some True} = the (\' i)}))" unfolding PiE_dflt_inter apply (subst image_Int[symmetric]) by (auto simp add: Collect_conj_eq) also have "... = (\x x = Some True} = the (\ x) \ satisfies_xor xor {x. \' x = Some True} = the (\' x)}))" by (subst measure_Pi_pmf_PiE_dflt) (auto simp add: inj_vimage_image_eq) also have "... = (\x: "dom \ = {..': "dom \' = {..'': "dom \'' = {..: "dom \ = V" assumes \': "dom \' = V" assumes \': "dom \'' = V" assumes neq: "\ \ \'" "\' \ \''" "\'' \ \" shows " measure_pmf.prob (random_xors V n) {xors. xor_hash \ xors = \ \ xor_hash \' xors = \' \ xor_hash \'' xors = \''} = 1 / 8 ^ n" proof - obtain x y z where "\ x \ \' x" "\' y \ \'' y" "\'' z \ \ z" using neq by blast then have neq: "{x. \ x = Some True} \ {x. \' x = Some True}" "{x. \' x = Some True} \ {x. \'' x = Some True}" "{x. \'' x = Some True} \ {x. \ x = Some True}" by (smt (verit, ccfv_threshold) assms domD domIff mem_Collect_eq)+ have "measure_pmf.prob (random_xors V n) {xors. xor_hash \ xors = \ \ xor_hash \' xors = \' \ xor_hash \'' xors = \''} = measure_pmf.prob (random_xors V n) ({xors. xor_hash \ xors = \} \ {xors. xor_hash \' xors = \'} \ {xors. xor_hash \'' xors = \''})" by (simp add: measure_pmf.measure_pmf_eq) moreover have "... = measure_pmf.prob (Pi_pmf {..<(n::nat)} None (\_. map_pmf Some (random_xor V))) (PiE_dflt {..i. Some ` {xor. satisfies_xor xor {x. \ x = Some True} = the (\ i)}) \ PiE_dflt {..i. Some ` {xor. satisfies_xor xor {x. \' x = Some True} = the (\' i)}) \ PiE_dflt {..i. Some ` {xor. satisfies_xor xor {x. \'' x = Some True} = the (\'' i)}))" unfolding random_xors_def xor_hash_eq_PiE_dflt[OF \] xor_hash_eq_PiE_dflt[OF \'] xor_hash_eq_PiE_dflt[OF \''] by auto moreover have "... = measure_pmf.prob (Pi_pmf {..<(n::nat)} None (\_. map_pmf Some (random_xor V))) (PiE_dflt {..i. Some ` {xor. satisfies_xor xor {x. \ x = Some True} = the (\ i) \ satisfies_xor xor {x. \' x = Some True} = the (\' i) \ satisfies_xor xor {x. \'' x = Some True} = the (\'' i)})) " unfolding PiE_dflt_inter apply (subst image_Int[symmetric]) subgoal by simp apply (intro arg_cong[where f="measure_pmf.prob (Pi_pmf {.._. map_pmf Some (random_xor V)))"]) apply (intro arg_cong[where f="PiE_dflt {..x x = Some True} = the (\ x) \ satisfies_xor xor {x. \' x = Some True} = the (\' x)\ satisfies_xor xor {x. \'' x = Some True} = the (\'' x)}))" apply (subst measure_Pi_pmf_PiE_dflt) by (auto simp add: inj_vimage_image_eq) moreover have "... = (\xPreliminary 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 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 \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"}.\ - -lemmas measure_pmf_cong = measure_pmf_cong -lemmas pmf_mono = pmf_mono -lemmas pmf_add = pmf_add - 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 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 +section \Blind\ +text \Blind section added to preserve section numbers\ + 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 deleted file mode 100644 --- a/thys/Distributed_Distinct_Elements/Pseudorandom_Combinators.thy +++ /dev/null @@ -1,747 +0,0 @@ -section \Combinators for Pseudorandom Objects\ - -theory Pseudorandom_Combinators - imports - Finite_Fields.Card_Irreducible_Polynomials - Universal_Hash_Families.Carter_Wegman_Hash_Family - Distributed_Distinct_Elements_Preliminary - Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF - Expander_Graphs.Expander_Graphs_Strongly_Explicit -begin - -text \Important Note: A more current version of the framework presented here is available at -@{verbatim Universal_Hash_Families.Pseudorandom_Objects}, -@{verbatim Universal_Hash_Families.Pseudorandom_Objects_Hash_Families} and -@{verbatim Expander_Graphs.Pseudorandom_Objects_Expander_Walks}. This section is left here -to prevent possible merge-conflicts.\ - -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\ - -unbundle intro_cong_syntax -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)))) \" - -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"}.\ - -lemmas split_pmf_mod_div' = split_pmf_mod_div' -lemmas pmf_of_set_prod_eq = pmf_of_set_prod_eq -lemmas split_pmf_mod_div = split_pmf_mod_div - -lemma split_pmf_mod: - assumes "a > (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_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,21 +1,20 @@ 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/Finite_Fields/ROOT b/thys/Finite_Fields/ROOT --- a/thys/Finite_Fields/ROOT +++ b/thys/Finite_Fields/ROOT @@ -1,29 +1,30 @@ chapter AFP session Finite_Fields = "HOL-Algebra" + options [timeout = 5400] sessions Digit_Expansions Dirichlet_Series Executable_Randomized_Algorithms Probabilistic_While theories Card_Irreducible_Polynomials Card_Irreducible_Polynomials_Aux Finite_Fields_Factorization_Ext Finite_Fields_Isomorphic Finite_Fields_Preliminary_Results Formal_Polynomial_Derivatives Monic_Polynomial_Factorization Ring_Characteristic Rabin_Irreducibility_Test Rabin_Irreducibility_Test_Code Finite_Fields_More_Bijections + Finite_Fields_More_PMF Finite_Fields_Indexed_Algebra_Code Finite_Fields_Mod_Ring_Code Finite_Fields_Poly_Factor_Ring_Code Finite_Fields_Poly_Ring_Code Find_Irreducible_Poly 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,1315 +1,1315 @@ section \Frequency Moment $0$\label{sec:f0}\ theory Frequency_Moment_0 imports Frequency_Moments_Preliminary_Results Median_Method.Median K_Smallest Universal_Hash_Families.Carter_Wegman_Hash_Family Frequency_Moments Landau_Ext Probability_Ext - Product_PMF_Ext + Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF 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 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); 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) = 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 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. 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" 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 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" 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" 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 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 "... \ 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 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 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 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 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" 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 \} \ (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 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) (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 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)" 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))" by (simp add:\_def) 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" 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) 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) also have "... \ (\ i\ {(u,v) \ {.. {.. 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)))})" 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 * 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 * 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))))" 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) 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} \ 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 {\. 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) 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 "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) 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" 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 (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 \) - 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 \ 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 / \'" 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) \ \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) 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 \))) \ 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) \ 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 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" by (simp add: left_diff_distrib right_diff_distrib diff_divide_distrib) 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) \ 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 \))\" 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 \))\}" 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 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) 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) 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) 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)\} \ 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))" 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))"]) (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 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 + \')))" 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))" 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" 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)" 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 \}" 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)" 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 (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" 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 \" 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' \)" 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 \ 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) 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. \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"]) (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" 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))) \ 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)" 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))) \ 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 (simp add: Transcendental.log_mono aux_2) 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 Transcendental.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\{.. f0_space_usage (n, \, \)" if b: "x \ {..\<^sub>E space" for x proof - have c: "x \ 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) 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) \ 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))) \ 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 (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)) + 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) + 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\{.. 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 \) * (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)) + 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) \ 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' 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 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)))" 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)" 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 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) \ 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] 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)) \ 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)) \ 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] 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))" 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)))" 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"] 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) (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)) \ 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)) \ 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 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))" 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)))" 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 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 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))" 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 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)))" 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]] 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,727 +1,727 @@ 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 + Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF 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) = 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}. (\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 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. (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 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" 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 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, \, \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) next case (snoc a as) 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) = ((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" using oddE by blast have "Collect even \ {..<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) have "p = card {.. {.. ({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)" by (rule Bochner_Integration.integral_cong, simp, simp) 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 "... = ((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) = ((real p - 1) ^ m * (real p + 1) + (- real p - 1) ^ m * (real p - 1)) / (2 * real p)" by simp qed 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 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 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" 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)))" 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 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 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" 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" 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)" 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. 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. 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 * (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 "... = 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 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" 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)" 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" 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)\} \ \.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] 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}" 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" 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} " 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" 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}. 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: "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)" 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) 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 "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) (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: "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))) \ 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)))" 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))) + (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, \, \))" 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, \, \). (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) \ 1/real_of_rat (\_of x) \ \ \ 1/real_of_rat (\_of 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' 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)" 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) note of_nat_int_ceiling [simp del] 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)))\) * (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,1016 +1,1016 @@ section \Frequency Moment $k$\ theory Frequency_Moment_k imports Frequency_Moments Landau_Ext Lp.Lp Median_Method.Median Probability_Ext - Product_PMF_Ext + Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF 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) = 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) 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) = 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 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 (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) = 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) 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)))" by auto have "pmf_of_set {.. (\k. bernoulli_pmf (1 / (real (length xs) + 1)) \ (\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. 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 "\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) 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" 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" 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 \ 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))) \ (\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 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)))" 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)) \ (\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}) (\_. 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 case False hence "{.. {}" by force 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))) \ (\\. 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}) (\_. 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}) (\_. 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))) \ (\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. b * (a^i * b^(k-1-i))) {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))) {0..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) * (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)" 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)" 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)" 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)" 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) * (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" 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 - 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.. length as * (\a\ set as. (\v \ {0..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) \ 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 \ 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) = 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}) (\_. 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 {.. (\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 {.. (\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 {.. (\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 {.. (\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) also have "... = ?rhs" by (simp add:map_pmf_def) finally show ?thesis by simp qed 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 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 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) 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 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 M\<^sub>1_def result_def median_restrict) finally have b: "M = M\<^sub>2 \ return_pmf \ median_rv" by simp 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" 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 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) also have "... = of_rat (F k as * \)^2 * (real s\<^sub>1 / 3)" 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" 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" (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) 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] 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)}" 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)}" 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 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)) \ 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 by (meson lessThan_iff less_imp_le_nat subsetD as_range) qed 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) \ 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) + 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" 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 \ 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 "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 "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. 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) \ 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) \ 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: "(\_. 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)))" 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)))" 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)" 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 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)" 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/Probability_Ext.thy b/thys/Frequency_Moments/Probability_Ext.thy --- a/thys/Frequency_Moments/Probability_Ext.thy +++ b/thys/Frequency_Moments/Probability_Ext.thy @@ -1,82 +1,57 @@ section \Probability Spaces\ text \Some additional results about probability spaces in addition to "HOL-Probability".\ theory Probability_Ext imports "HOL-Probability.Stream_Space" Concentration_Inequalities.Bienaymes_Identity Universal_Hash_Families.Carter_Wegman_Hash_Family Frequency_Moments_Preliminary_Results begin -text \The following aliases are here to prevent possible merge-conflicts. The lemmas have been -moved to @{theory Concentration_Inequalities.Bienaymes_Identity} and/or -@{theory Concentration_Inequalities.Concentration_Inequalities_Preliminary}.\ - -lemmas make_ext = forall_Pi_to_PiE -lemmas PiE_reindex = PiE_reindex - context prob_space begin -lemmas indep_sets_reindex = indep_sets_reindex -lemmas indep_vars_cong_AE = indep_vars_cong_AE -lemmas indep_vars_reindex = indep_vars_reindex -lemmas variance_divide = variance_divide -lemmas covariance_def = covariance_def -lemmas real_prod_integrable = cauchy_schwartz(1) -lemmas covariance_eq = covariance_eq -lemmas covar_integrable = covar_integrable -lemmas sum_square_int = sum_square_int -lemmas var_sum_1 = bienaymes_identity -lemmas covar_self_eq = covar_self_eq -lemmas covar_indep_eq_zero = covar_indep_eq_zero -lemmas var_sum_2 = bienaymes_identity_2 -lemmas var_sum_pairwise_indep = bienaymes_identity_pairwise_indep -lemmas indep_var_from_indep_vars = indep_var_from_indep_vars -lemmas var_sum_pairwise_indep_2 = bienaymes_identity_pairwise_indep_2 -lemmas var_sum_all_indep = bienaymes_identity_full_indep - lemma pmf_mono: assumes "M = measure_pmf p" assumes "\x. x \ P \ x \ set_pmf p \ x \ Q" shows "prob P \ prob Q" proof - have "prob P = prob (P \ (set_pmf p))" by (rule measure_pmf_eq[OF assms(1)], blast) also have "... \ prob Q" using assms by (intro finite_measure.finite_measure_mono, auto) finally show ?thesis by simp qed lemma pmf_add: assumes "M = measure_pmf p" assumes "\x. x \ P \ x \ set_pmf p \ x \ Q \ x \ R" shows "prob P \ prob Q + prob R" proof - have [simp]:"events = UNIV" by (subst assms(1), simp) have "prob P \ prob (Q \ R)" using assms by (intro pmf_mono[OF assms(1)], blast) also have "... \ prob Q + prob R" by (rule measure_subadditive, auto) finally show ?thesis by simp qed lemma pmf_add_2: assumes "M = measure_pmf p" assumes "prob {\. P \} \ r1" assumes "prob {\. Q \} \ r2" shows "prob {\. P \ \ Q \} \ r1 + r2" (is "?lhs \ ?rhs") proof - have "?lhs \ prob {\. P \} + prob {\. Q \}" by (intro pmf_add[OF assms(1)], auto) also have "... \ ?rhs" by (intro add_mono assms(2-3)) finally show ?thesis by simp qed end end diff --git a/thys/Frequency_Moments/Product_PMF_Ext.thy b/thys/Frequency_Moments/Product_PMF_Ext.thy deleted file mode 100644 --- a/thys/Frequency_Moments/Product_PMF_Ext.thy +++ /dev/null @@ -1,30 +0,0 @@ -section \Indexed Products of Probability Mass Functions\ - -theory Product_PMF_Ext - imports - Probability_Ext - Universal_Hash_Families.Universal_Hash_Families_More_Product_PMF -begin - -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}.\ - -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" - -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 - -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,26 @@ chapter AFP session Frequency_Moments = "HOL-Probability" + options [timeout = 1200] sessions Bertrands_Postulate Concentration_Inequalities Equivalence_Relation_Enumeration 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"