diff --git a/thys/Density_Compiler/PDF_Transformations.thy b/thys/Density_Compiler/PDF_Transformations.thy --- a/thys/Density_Compiler/PDF_Transformations.thy +++ b/thys/Density_Compiler/PDF_Transformations.thy @@ -1,478 +1,478 @@ (* Theory: PDF_Transformations.thy Author: Manuel Eberl Provides lemmas for transformations of measure spaces with a density. *) section \Measure Space Transformations\ theory PDF_Transformations imports Density_Predicates begin lemma not_top_le_1_ennreal[simp]: "\ top \ (1::ennreal)" by (simp add: top_unique) lemma range_int: "range int = {n. n \ 0}" proof (intro equalityI subsetI) fix n :: int assume "n \ {n. n \ 0}" hence "n = int (nat n)" by simp thus "n \ range int" by blast qed auto lemma range_exp: "range (exp :: real \ real) = {x. x > 0}" proof (intro equalityI subsetI) fix x :: real assume "x \ {x. x > 0}" hence "x = exp (ln x)" by simp thus "x \ range exp" by blast qed auto lemma Int_stable_Icc: "Int_stable (range (\(a, b). {a .. b::real}))" by (auto simp: Int_stable_def) lemma distr_mult_real: assumes "c \ 0" "has_density M lborel (f :: real \ ennreal)" shows "has_density (distr M borel ((*) c)) lborel (\x. f (x / c) * inverse (abs c))" (is "has_density ?M' _ ?f'") proof from assms(2) have "M = density lborel f" by (rule has_densityD) also from assms have Mf[measurable]: "f \ borel_measurable borel" by (auto dest: has_densityD) hence "distr (density lborel f) borel ((*) c) = density lborel ?f'" (is "?M1 = ?M2") proof (intro measure_eqI) fix X assume X[measurable]: "X \ sets (distr (density lborel f) borel ((*) c))" with assms have "emeasure ?M1 X = \\<^sup>+x. f x * indicator X (c * x) \lborel" by (subst emeasure_distr, simp, simp, subst emeasure_density) (auto dest: has_densityD intro!: measurable_sets_borel nn_integral_cong split: split_indicator) also from assms(1) and X have "... = \\<^sup>+x. ?f' x * indicator X x \lborel" apply (subst lborel_distr_mult'[of "inverse c"]) apply simp apply (subst nn_integral_density) apply (simp_all add: nn_integral_distr field_simps) done also from X have "... = emeasure ?M2 X" by (subst emeasure_density) auto finally show "emeasure ?M1 X = emeasure ?M2 X" . qed simp finally show "distr M borel ((*) c) = density lborel ?f'" . qed (insert assms, auto dest: has_densityD) lemma distr_uminus_real: assumes "has_density M lborel (f :: real \ ennreal)" shows "has_density (distr M borel uminus) lborel (\x. f (- x))" proof- from assms have "has_density (distr M borel ((*) (- 1))) lborel (\x. f (x / -1) * ennreal (inverse (abs (-1))))" by (intro distr_mult_real) simp_all also have "(*) (-1) = (uminus :: real \ real)" by (intro ext) simp also have "(\x. f (x / -1) * ennreal (inverse (abs (-1)))) = (\x. f (-x))" by (intro ext) (simp add: one_ennreal_def[symmetric]) finally show ?thesis . qed lemma distr_plus_real: assumes "has_density M lborel (f :: real \ ennreal)" shows "has_density (distr M borel ((+) c)) lborel (\x. f (x - c))" proof from assms have "M = density lborel f" by (rule has_densityD) also from assms have Mf[measurable]: "f \ borel_measurable borel" by (auto dest: has_densityD) hence "distr (density lborel f) borel ((+) c) = density lborel (\x. f (x - c))" (is "?M1 = ?M2") proof (intro measure_eqI) fix X assume X: "X \ sets (distr (density lborel f) borel ((+) c))" with assms have "emeasure ?M1 X = \\<^sup>+x. f x * indicator X (c + x) \lborel" by (subst emeasure_distr, simp, simp, subst emeasure_density) (auto dest: has_densityD intro!: measurable_sets_borel nn_integral_cong split: split_indicator) also from X have "... = \\<^sup>+x. f (x - c) * indicator X x \lborel" by (subst lborel_distr_plus[where c = "-c", symmetric], subst nn_integral_distr) auto also from X have "... = emeasure ?M2 X" by (subst emeasure_density) (auto simp: emeasure_density intro!: measurable_compose[OF borel_measurable_diff Mf]) finally show "emeasure ?M1 X = emeasure ?M2 X" . qed simp finally show "distr M borel ((+) c) = density lborel (\x. f (x - c))" . qed (insert assms, auto dest: has_densityD) lemma count_space_uminus: "count_space UNIV = distr (count_space UNIV) (count_space UNIV) (uminus :: ('a :: ring \ _))" proof (rule distr_bij_count_space[symmetric]) show "bij (uminus :: 'a \ 'a)" by (auto intro!: o_bij[where g=uminus]) qed lemma count_space_plus: "count_space UNIV = distr (count_space UNIV) (count_space UNIV) ((+) (c :: ('a :: ring)))" by (rule distr_bij_count_space [symmetric]) simp lemma distr_uminus_ring_count_space: assumes "has_density M (count_space UNIV) (f :: _ :: ring \ ennreal)" shows "has_density (distr M (count_space UNIV) uminus) (count_space UNIV) (\x. f (- x))" proof from assms have "M = density (count_space UNIV) f" by (rule has_densityD) also have "distr (density (count_space UNIV) f) (count_space UNIV) uminus = density (count_space UNIV)(\x. f (- x))" (is "?M1 = ?M2") proof (intro measure_eqI) fix X assume X: "X \ sets (distr (density (count_space UNIV) f) (count_space UNIV) uminus)" with assms have "emeasure ?M1 X = \\<^sup>+x. f x * indicator X (-x) \count_space UNIV" by (subst emeasure_distr, simp, simp, subst emeasure_density) (auto dest: has_densityD intro!: measurable_sets_borel nn_integral_cong split: split_indicator) also from X have "... = emeasure ?M2 X" by (subst count_space_uminus) (simp_all add: nn_integral_distr emeasure_density) finally show "emeasure ?M1 X = emeasure ?M2 X" . qed simp finally show "distr M (count_space UNIV) uminus = density (count_space UNIV) (\x. f (- x))" . qed (insert assms, auto dest: has_densityD) lemma distr_plus_ring_count_space: assumes "has_density M (count_space UNIV) (f :: _ :: ring \ ennreal)" shows "has_density (distr M (count_space UNIV) ((+) c)) (count_space UNIV) (\x. f (x - c))" proof from assms have "M = density (count_space UNIV) f" by (rule has_densityD) also have "distr (density (count_space UNIV) f) (count_space UNIV) ((+) c) = density (count_space UNIV)(\x. f (x - c))" (is "?M1 = ?M2") proof (intro measure_eqI) fix X assume X: "X \ sets (distr (density (count_space UNIV) f) (count_space UNIV) ((+) c))" with assms have "emeasure ?M1 X = \\<^sup>+x. f x * indicator X (c + x) \count_space UNIV" by (subst emeasure_distr, simp, simp, subst emeasure_density) (auto dest: has_densityD intro!: measurable_sets_borel nn_integral_cong split: split_indicator) also from X have "... = emeasure ?M2 X" by (subst count_space_plus[of "-c"]) (simp_all add: nn_integral_distr emeasure_density) finally show "emeasure ?M1 X = emeasure ?M2 X" . qed simp finally show "distr M (count_space UNIV) ((+) c) = density (count_space UNIV) (\x. f (x - c))" . qed (insert assms, auto dest: has_densityD) lemma subprob_density_distr_real_eq: assumes dens: "has_subprob_density M lborel f" assumes Mh: "h \ borel_measurable borel" assumes Mg: "g \ borel_measurable borel" assumes measure_eq: "\a b. a \ b \ emeasure (distr (density lborel f) lborel h) {a..b} = emeasure (density lborel g) {a..b}" shows "has_subprob_density (distr M borel (h :: real \ real)) lborel g" proof (rule has_subprob_densityI) from dens have sets_M: "sets M = sets borel" by (auto dest: has_subprob_densityD) have meas_M[simp]: "measurable M = measurable borel" by (intro ext, subst measurable_cong_sets[OF sets_M refl]) auto from Mh and dens show subprob_space: "subprob_space (distr M borel h)" by (intro subprob_space.subprob_space_distr) (auto dest: has_subprob_densityD) show "distr M borel h = density lborel g" proof (rule measure_eqI_generator_eq[OF Int_stable_Icc, of UNIV]) { fix x :: real obtain n :: nat where "n > abs x" using reals_Archimedean2 by auto hence "\n::nat. x \ {-real n..real n}" by (intro exI[of _ n]) auto } thus "(\i::nat. {-real i..real i}) = UNIV" by blast next fix i :: nat from subprob_space have "emeasure (distr M borel h) {-real i..real i} \ 1" by (intro subprob_space.subprob_emeasure_le_1) (auto dest: has_subprob_densityD) thus "emeasure (distr M borel h) {- real i..real i} \ \" by auto next fix X :: "real set" assume "X \ range (\(a,b). {a..b})" then obtain a b where "X = {a..b}" by auto with dens have "emeasure (distr M lborel h) X = emeasure (density lborel g) X" by (cases "a \ b") (auto simp: measure_eq dest: has_subprob_densityD) also have "distr M lborel h = distr M borel h" by (rule distr_cong) auto finally show "emeasure (distr M borel h) X = emeasure (density lborel g) X" . qed (auto simp: borel_eq_atLeastAtMost) qed (insert assms, auto) lemma subprob_density_distr_real_exp: assumes dens: "has_subprob_density M lborel f" shows "has_subprob_density (distr M borel exp) lborel (\x. if x > 0 then f (ln x) * ennreal (inverse x) else 0)" (is "has_subprob_density _ _ ?g") proof (rule subprob_density_distr_real_eq[OF dens]) from dens have [measurable]: "f \ borel_measurable borel" by (auto dest: has_subprob_densityD) have Mf: "(\x. f (ln x) * ennreal (inverse x)) \ borel_measurable borel" by simp fix a b :: real assume "a \ b" let ?A = "\i. {inverse (Suc i) :: real ..}" let ?M1 = "distr (density lborel f) lborel exp" and ?M2 = "density lborel ?g" { fix x :: real assume "\i. x < inverse (Suc i)" hence "x \ 0" by (intro tendsto_lowerbound[OF LIMSEQ_inverse_real_of_nat]) (auto intro!: always_eventually less_imp_le) } hence decomp: "{a..b} = {x\{a..b}. x \ 0} \ (\i. ?A i \ {a..b})" (is "_ = ?C \ ?D") by (auto simp: not_le) have inv_le: "\x i. x \ inverse (real (Suc i)) \ \(x \ 0)" by (subst not_le, erule dual_order.strict_trans1, simp) hence "emeasure ?M1 {a..b} = emeasure ?M1 ?C + emeasure ?M1 ?D" by (subst decomp, intro plus_emeasure[symmetric]) auto also have "emeasure ?M1 ?C = 0" by (subst emeasure_distr) auto also have "0 = emeasure ?M2 ?C" by (subst emeasure_density, simp, simp, rule sym, subst nn_integral_0_iff) auto also have "emeasure ?M1 (\i. ?A i \ {a..b}) = (SUP i. emeasure ?M1 (?A i \ {a..b}))" by (rule SUP_emeasure_incseq[symmetric]) (auto simp: incseq_def max_def not_le dest: order.strict_trans1) also have "\i. emeasure ?M1 (?A i \ {a..b}) = emeasure ?M2 (?A i \ {a..b})" proof (case_tac "inverse (Suc i) \ b") fix i assume True: "inverse (Suc i) \ b" let ?a = "inverse (Suc i)" from \a \ b\ have A: "?A i \ {a..b} = {max ?a a..b}" (is "?E = ?F") by auto hence "emeasure ?M1 ?E = emeasure ?M1 ?F" by simp - also have "strict_mono_on ln {max (inverse (real (Suc i))) a..b}" + also have "strict_mono_on {max (inverse (real (Suc i))) a..b} ln" by (rule strict_mono_onI, subst ln_less_cancel_iff) (auto dest: inv_le simp del: of_nat_Suc) with \a \ b\ True dens have "emeasure ?M1 ?F = emeasure (density lborel (\x. f (ln x) * inverse x)) ?F" by (intro emeasure_density_distr_interval) (auto simp: Mf not_less not_le range_exp dest: has_subprob_densityD dest!: inv_le intro!: DERIV_ln continuous_on_inverse continuous_on_id simp del: of_nat_Suc) also note A[symmetric] also have "emeasure (density lborel (\x. f (ln x) * inverse x)) ?E = emeasure ?M2 ?E" by (subst (1 2) emeasure_density) (auto intro!: nn_integral_cong split: split_indicator dest!: inv_le simp del: of_nat_Suc) finally show "emeasure ?M1 (?A i \ {a..b}) = emeasure ?M2 (?A i \ {a..b})" . qed simp hence "(SUP i. emeasure ?M1 (?A i \ {a..b})) = (SUP i. emeasure ?M2 (?A i \ {a..b}))" by simp also have "... = emeasure ?M2 (\i. ?A i \ {a..b})" by (rule SUP_emeasure_incseq) (auto simp: incseq_def max_def not_le dest: order.strict_trans1) also have "emeasure ?M2 ?C + emeasure ?M2 ?D = emeasure ?M2 (?C \ ?D)" by (rule plus_emeasure) (auto dest: inv_le simp del: of_nat_Suc) also note decomp[symmetric] finally show "emeasure ?M1 {a..b} = emeasure ?M2 {a..b}" . qed (insert dens, auto dest!: has_subprob_densityD(1)) lemma subprob_density_distr_real_inverse_aux: assumes dens: "has_subprob_density M lborel f" shows "has_subprob_density (distr M borel (\x. - inverse x)) lborel (\x. f (-inverse x) * ennreal (inverse (x * x)))" (is "has_subprob_density _ _ ?g") proof (rule subprob_density_distr_real_eq[OF dens]) from dens have Mf[measurable]: "f \ borel_measurable borel" by (auto dest: has_subprob_densityD) show Mg: "?g \ borel_measurable borel" by measurable have surj[simp]: "surj (\x. - inverse x :: real)" by (intro surjI[of _ "\x. - inverse x"]) (simp add: field_simps) fix a b :: real assume "a \ b" let ?A1 = "\i. {..-inverse (Suc i) :: real}" and ?A2 = "\i. {inverse (Suc i) :: real ..}" let ?C = "if 0 \ {a..b} then {0} else {}" let ?M1 = "distr (density lborel f) lborel (\x. - inverse x)" and ?M2 = "density lborel ?g" have inv_le: "\x i. x \ inverse (real (Suc i)) \ \(x \ 0)" by (subst not_le, erule dual_order.strict_trans1, simp) have "\x. x > 0 \ \i. x \ inverse (Suc i)" proof (rule ccontr) fix x :: real assume "x > 0" "\(\i. x \ inverse (Suc i))" hence "x \ 0" by (intro tendsto_lowerbound[OF LIMSEQ_inverse_real_of_nat]) (auto intro!: always_eventually less_imp_le simp: not_le) with \x > 0\ show False by simp qed hence A: "(\i. ?A2 i) = {0<..}" by (auto dest: inv_le simp del: of_nat_Suc) moreover have "\x. x < 0 \ \i. x \ -inverse (Suc i)" proof (rule ccontr) fix x :: real assume "x < 0" "\(\i. x \ -inverse (Suc i))" hence "x \ 0" by (intro tendsto_upperbound, simp) (auto intro!: always_eventually less_imp_le LIMSEQ_inverse_real_of_nat_add_minus simp: not_le) with \x < 0\ show False by simp qed hence B: "(\i. ?A1 i) = {..<0}" by (auto simp: le_minus_iff[of _ "inverse x" for x] dest!: inv_le simp del: of_nat_Suc) ultimately have C: "UNIV = (\i. ?A1 i) \ (\i. ?A2 i) \ {0}" by (subst A, subst B) force have UN_Int_distrib: "\f A. (\i. f i) \ A = (\i. f i \ A)" by blast have decomp: "{a..b} = (\i. ?A1 i \ {a..b}) \ (\i. ?A2 i \ {a..b}) \ ?C" (is "_ = ?D \ ?E \ _") by (subst Int_UNIV_left[symmetric], simp only: C Int_Un_distrib2 UN_Int_distrib) (simp split: if_split) have "emeasure ?M1 {a..b} = emeasure ?M1 ?D + emeasure ?M1 ?E + emeasure ?M1 ?C" apply (subst decomp) apply (subst plus_emeasure[symmetric], simp, simp, simp) apply (subst plus_emeasure[symmetric]) apply (auto dest!: inv_le simp: not_le le_minus_iff[of _ "inverse x" for x] simp del: of_nat_Suc) done also have "(\x. - inverse x) -` {0 :: real} = {0}" by (auto simp: field_simps) hence "emeasure ?M1 ?C = 0" by (subst emeasure_distr) (auto split: if_split simp: emeasure_density Mf) also have "emeasure ?M2 {0} = 0" by (simp add: emeasure_density) hence "0 = emeasure ?M2 ?C" by (rule_tac sym, rule_tac order.antisym, rule_tac order.trans, rule_tac emeasure_mono[of _ "{0}"]) simp_all also have "emeasure ?M1 (\i. ?A1 i \ {a..b}) = (SUP i. emeasure ?M1 (?A1 i \ {a..b}))" by (rule SUP_emeasure_incseq[symmetric]) (auto simp: incseq_def max_def not_le dest: order.strict_trans1) also have "\i. emeasure ?M1 (?A1 i \ {a..b}) = emeasure ?M2 (?A1 i \ {a..b})" proof (case_tac "-inverse (Suc i) \ a") fix i assume True: "-inverse (Suc i) \ a" let ?a = "-inverse (Suc i)" from \a \ b\ have A: "?A1 i \ {a..b} = {a..min ?a b}" (is "?F = ?G") by auto hence "emeasure ?M1 ?F = emeasure ?M1 ?G" by simp - also have "strict_mono_on (\x. -inverse x) {a..min ?a b}" + also have "strict_mono_on {a..min ?a b} (\x. -inverse x)" by (rule strict_mono_onI) (auto simp: le_minus_iff[of _ "inverse x" for x] dest!: inv_le simp del: of_nat_Suc) with \a \ b\ True dens have "emeasure ?M1 ?G = emeasure ?M2 ?G" by (intro emeasure_density_distr_interval) (auto simp: Mf not_less dest: has_subprob_densityD inv_le intro!: derivative_eq_intros continuous_on_mult continuous_on_inverse continuous_on_id) also note A[symmetric] finally show "emeasure ?M1 (?A1 i \ {a..b}) = emeasure ?M2 (?A1 i \ {a..b})" . qed simp hence "(SUP i. emeasure ?M1 (?A1 i \ {a..b})) = (SUP i. emeasure ?M2 (?A1 i \ {a..b}))" by simp also have "... = emeasure ?M2 (\i. ?A1 i \ {a..b})" by (rule SUP_emeasure_incseq) (auto simp: incseq_def max_def not_le dest: order.strict_trans1) also have "emeasure ?M1 (\i. ?A2 i \ {a..b}) = (SUP i. emeasure ?M1 (?A2 i \ {a..b}))" by (rule SUP_emeasure_incseq[symmetric]) (auto simp: incseq_def max_def not_le dest: order.strict_trans1) also have "\i. emeasure ?M1 (?A2 i \ {a..b}) = emeasure ?M2 (?A2 i \ {a..b})" proof (case_tac "inverse (Suc i) \ b") fix i assume True: "inverse (Suc i) \ b" let ?a = "inverse (Suc i)" from \a \ b\ have A: "?A2 i \ {a..b} = {max ?a a..b}" (is "?F = ?G") by auto hence "emeasure ?M1 ?F = emeasure ?M1 ?G" by simp - also have "strict_mono_on (\x. -inverse x) {max ?a a..b}" + also have "strict_mono_on {max ?a a..b} (\x. -inverse x)" by (rule strict_mono_onI) (auto dest!: inv_le simp: not_le simp del: of_nat_Suc) with \a \ b\ True dens have "emeasure ?M1 ?G = emeasure ?M2 ?G" by (intro emeasure_density_distr_interval) (auto simp: Mf not_less dest: has_subprob_densityD inv_le intro!: derivative_eq_intros continuous_on_mult continuous_on_inverse continuous_on_id) also note A[symmetric] finally show "emeasure ?M1 (?A2 i \ {a..b}) = emeasure ?M2 (?A2 i \ {a..b})" . qed simp hence "(SUP i. emeasure ?M1 (?A2 i \ {a..b})) = (SUP i. emeasure ?M2 (?A2 i \ {a..b}))" by simp also have "... = emeasure ?M2 (\i. ?A2 i \ {a..b})" by (rule SUP_emeasure_incseq) (auto simp: incseq_def max_def not_le dest: order.strict_trans1) also have "emeasure ?M2 ?D + emeasure ?M2 ?E + emeasure ?M2 ?C = emeasure ?M2 {a..b}" apply (subst (4) decomp) apply (subst plus_emeasure, simp, simp) apply (auto dest!: inv_le simp: not_le le_minus_iff[of _ "inverse x" for x] simp del: of_nat_Suc) apply (subst plus_emeasure) apply (auto dest!: inv_le simp: not_le le_minus_iff[of _ "inverse x" for x]) done finally show "emeasure ?M1 {a..b} = emeasure ?M2 {a..b}" . qed simp lemma subprob_density_distr_real_inverse: assumes dens: "has_subprob_density M lborel f" shows "has_subprob_density (distr M borel inverse) lborel (\x. f (inverse x) * ennreal (inverse (x * x)))" proof (unfold has_subprob_density_def, intro conjI) let ?g' = "(\x. f (-inverse x) * ennreal (inverse (x * x)))" have prob: "has_subprob_density (distr M borel (\x. -inverse x)) lborel ?g'" by (rule subprob_density_distr_real_inverse_aux[OF assms]) from assms have sets_M: "sets M = sets borel" by (auto dest: has_subprob_densityD) have [simp]: "measurable M = measurable borel" by (intro ext, subst measurable_cong_sets[OF sets_M refl]) auto from prob have dens: "has_density (distr M lborel (\x. -inverse x)) lborel (\x. f (-inverse x) * ennreal (inverse (x * x)))" unfolding has_subprob_density_def by (simp cong: distr_cong) from distr_uminus_real[OF this] show "has_density (distr M borel inverse) lborel (\x. f (inverse x) * ennreal (inverse (x * x)))" by (simp add: distr_distr o_def cong: distr_cong) show "subprob_space (distr M borel inverse)" by (intro subprob_space.subprob_space_distr has_subprob_densityD[OF assms]) simp_all qed lemma distr_convolution_real: assumes "has_density M lborel (f :: (real \ real) \ ennreal)" shows "has_density (distr M borel (case_prod (+))) lborel (\z. \\<^sup>+x. f (x, z - x) \lborel)" (is "has_density ?M' _ ?f'") proof from has_densityD[OF assms] have Mf[measurable]: "f \ borel_measurable borel" by simp show Mf': "(\z. \\<^sup>+ x. f (x, z - x) \lborel) \ borel_measurable lborel" by measurable from assms have sets_M: "sets M = sets borel" by (auto dest: has_densityD) hence [simp]: "space M = UNIV" by (subst sets_eq_imp_space_eq[OF sets_M]) simp from sets_M have [simp]: "measurable M = measurable borel" by (intro ext measurable_cong_sets) simp_all have M_add: "case_prod (+) \ borel_measurable (borel :: (real \ real) measure)" by (simp add: borel_prod[symmetric]) show "distr M borel (case_prod (+)) = density lborel ?f'" proof (rule measure_eqI) fix X :: "real set" assume X[measurable]: "X \ sets (distr M borel (case_prod (+)))" hence "emeasure (distr M borel (case_prod (+))) X = emeasure M ((\(x, y). x + y) -` X)" by (simp_all add: M_add emeasure_distr) also from X have "... = \\<^sup>+z. f z * indicator ((\(x, y). x + y) -` X) z \(lborel \\<^sub>M lborel)" by (simp add: emeasure_density has_densityD[OF assms] measurable_sets_borel[OF M_add] lborel_prod) also have "... = \\<^sup>+x. \\<^sup>+y. f (x, y) * indicator ((\(x, y). x + y) -` X) (x,y) \lborel \lborel" apply (rule lborel.nn_integral_fst[symmetric]) apply measurable apply (simp_all add: borel_prod) done also have "... = \\<^sup>+x. \\<^sup>+y. f (x, y) * indicator ((\(x, y). x + y) -` X) (x,y) \distr lborel borel ((+) (-x)) \lborel" by (rule nn_integral_cong, subst lborel_distr_plus) simp also have "... = \\<^sup>+x. \\<^sup>+z. f (x, z-x) * indicator ((\(x, y). x + y) -` X) (x, z-x) \lborel \lborel" apply (rule nn_integral_cong) apply (subst nn_integral_distr) apply simp_all apply measurable apply (subst space_count_space) apply auto done also have "... = \\<^sup>+x. \\<^sup>+z. f (x, z-x) * indicator X z \lborel \lborel" by (intro nn_integral_cong) (simp split: split_indicator) also have "... = \\<^sup>+z. \\<^sup>+x. f (x, z-x) * indicator X z \lborel \lborel" using X by (subst lborel_pair.Fubini') (simp_all add: pair_sigma_finite_def) also have "... = \\<^sup>+z. (\\<^sup>+x. f (x, z-x) \lborel) * indicator X z \lborel" by (rule nn_integral_cong) (simp split: split_indicator) also have "... = emeasure (density lborel ?f') X" using X by (simp add: emeasure_density) finally show "emeasure (distr M borel (case_prod (+))) X = emeasure (density lborel ?f') X" . qed (insert assms, auto dest: has_densityD) qed simp_all lemma distr_convolution_ring_count_space: assumes C: "countable (UNIV :: 'a set)" assumes "has_density M (count_space UNIV) (f :: (('a :: ring) \ 'a) \ ennreal)" shows "has_density (distr M (count_space UNIV) (case_prod (+))) (count_space UNIV) (\z. \\<^sup>+x. f (x, z - x) \count_space UNIV)" (is "has_density ?M' _ ?f'") proof let ?CS = "count_space UNIV :: 'a measure" and ?CSP = "count_space UNIV :: ('a \ 'a) measure" show Mf': "(\z. \\<^sup>+ x. f (x, z - x) \count_space UNIV) \ borel_measurable ?CS" by simp from assms have sets_M: "sets M = UNIV" and [simp]: "space M = UNIV" by (auto dest: has_densityD) from assms have [simp]: "measurable M = measurable (count_space UNIV)" by (intro ext measurable_cong_sets) (simp_all add: sets_M) interpret sigma_finite_measure ?CS by (rule sigma_finite_measure_count_space_countable[OF C]) show "distr M ?CS (case_prod (+)) = density ?CS ?f'" proof (rule measure_eqI) fix X :: "'a set" assume X: "X \ sets (distr M ?CS (case_prod (+)))" hence "emeasure (distr M ?CS (case_prod (+))) X = emeasure M ((\(x, y). x + y) -` X)" by (simp_all add: emeasure_distr) also from X have "... = \\<^sup>+z. f z * indicator ((\(x, y). x + y) -` X) z \(?CS \\<^sub>M ?CS)" by (simp add: emeasure_density has_densityD[OF assms(2)] sets_M pair_measure_countable C) also have "... = \\<^sup>+x. \\<^sup>+y. f (x, y) * indicator ((\(x, y). x + y) -` X) (x,y) \?CS \?CS" by (rule nn_integral_fst[symmetric]) (simp add: pair_measure_countable C) also have "... = \\<^sup>+x. \\<^sup>+y. f (x, y) * indicator ((\(x, y). x + y) -` X) (x,y) \distr ?CS ?CS ((+) (-x)) \?CS" by (rule nn_integral_cong, subst count_space_plus) simp also have "... = \\<^sup>+x. \\<^sup>+z. f (x, z-x) * indicator ((\(x, y). x + y) -` X) (x, z-x) \?CS \?CS" by (rule nn_integral_cong) (simp_all add: nn_integral_distr) also have "... = \\<^sup>+x. \\<^sup>+z. f (x, z-x) * indicator X z \?CS \?CS" by (intro nn_integral_cong) (simp split: split_indicator) also have "... = \\<^sup>+z. \\<^sup>+x. f (x, z-x) * indicator X z \?CS \?CS" using X by (subst pair_sigma_finite.Fubini') (simp_all add: pair_sigma_finite_def sigma_finite_measure_count_space_countable C pair_measure_countable) also have "... = \\<^sup>+z. (\\<^sup>+x. f (x, z-x) \?CS) * indicator X z \?CS" by (rule nn_integral_cong) (simp split: split_indicator) also have "... = emeasure (density ?CS ?f') X" using X by (simp add: emeasure_density) finally show "emeasure (distr M ?CS (case_prod (+))) X = emeasure (density ?CS ?f') X" . qed (insert assms, auto dest: has_densityD) qed simp_all end diff --git a/thys/Frequency_Moments/Frequency_Moment_0.thy b/thys/Frequency_Moments/Frequency_Moment_0.thy --- a/thys/Frequency_Moments/Frequency_Moment_0.thy +++ b/thys/Frequency_Moments/Frequency_Moment_0.thy @@ -1,1314 +1,1314 @@ 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 Product_PMF_Ext Universal_Hash_Families.Field begin text \This section contains a formalization of a new algorithm for the zero-th frequency moment inspired by ideas described in \cite{baryossed2002}. 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 var_sum_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 real_of_float A" for A by (meson less_float.rep_eq strict_mono_onI) + 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 (intro ereal_mono add_mono mult_left_mono log_mono of_nat_mono add_pos_nonneg, auto) also have "... = ereal (10 + 4 * real r + 2 * log 2 (log 2 (8 * n + 160)))" by (simp add:log_mult[symmetric]) also have "... \ ereal (10 + 4 * real r + 2 * log 2 (log 2 ((n+13) powr 2)))" by (intro ereal_mono add_mono mult_left_mono log_mono of_nat_mono add_pos_nonneg) (auto simp add:power2_eq_square algebra_simps) also have "... = ereal (10 + 4 * real r + 2 * log 2 (log 2 4 * log 2 (n + 13)))" by (subst log_powr, simp_all add:log_2_4) also have "... = ereal (12 + 4 * real r + 2 * log 2 (log 2 (n + 13)))" by (subst log_mult, simp_all add:log_2_4) finally show ?thesis by simp next case False hence "y = 0" using a_1 by simp then show ?thesis by (simp add:float_bit_count_zero) qed have "bit_count (encode_f0_state (s, t, p, r, x, \i\{.. 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/K_Smallest.thy b/thys/Frequency_Moments/K_Smallest.thy --- a/thys/Frequency_Moments/K_Smallest.thy +++ b/thys/Frequency_Moments/K_Smallest.thy @@ -1,385 +1,385 @@ section \Ranks, $k$ smallest element and elements\ theory K_Smallest imports Frequency_Moments_Preliminary_Results Interpolation_Polynomials_HOL_Algebra.Interpolation_Polynomial_Cardinalities begin text \This section contains definitions and results for the selection of the $k$ smallest elements, the $k$-th smallest element, rank of an element in an ordered set.\ definition rank_of :: "'a :: linorder \ 'a set \ nat" where "rank_of x S = card {y \ S. y < x}" text \The function @{term "rank_of"} returns the rank of an element within a set.\ lemma rank_mono: assumes "finite S" shows "x \ y \ rank_of x S \ rank_of y S" unfolding rank_of_def using assms by (intro card_mono, auto) lemma rank_mono_2: assumes "finite S" shows "S' \ S \ rank_of x S' \ rank_of x S" unfolding rank_of_def using assms by (intro card_mono, auto) lemma rank_mono_commute: assumes "finite S" assumes "S \ T" - assumes "strict_mono_on f T" + assumes "strict_mono_on T f" assumes "x \ T" shows "rank_of x S = rank_of (f x) (f ` S)" proof - have a: "inj_on f T" by (metis assms(3) strict_mono_on_imp_inj_on) have "rank_of (f x) (f ` S) = card (f ` {y \ S. f y < f x})" unfolding rank_of_def by (intro arg_cong[where f="card"], auto) also have "... = card (f ` {y \ S. y < x})" using assms by (intro arg_cong[where f="card"] arg_cong[where f="(`) f"]) (meson in_mono linorder_not_le strict_mono_onD strict_mono_on_leD set_eq_iff) also have "... = card {y \ S. y < x}" using assms by (intro card_image inj_on_subset[OF a], blast) also have "... = rank_of x S" by (simp add:rank_of_def) finally show ?thesis by simp qed definition least where "least k S = {y \ S. rank_of y S < k}" text \The function @{term "least"} returns the k smallest elements of a finite set.\ lemma rank_strict_mono: assumes "finite S" - shows "strict_mono_on (\x. rank_of x S) S" + shows "strict_mono_on S (\x. rank_of x S)" proof - have "\x y. x \ S \ y \ S \ x < y \ rank_of x S < rank_of y S" unfolding rank_of_def using assms by (intro psubset_card_mono, auto) thus ?thesis by (simp add:rank_of_def strict_mono_on_def) qed lemma rank_of_image: assumes "finite S" shows "(\x. rank_of x S) ` S = {0..x. x \ S \ card {y \ S. y < x} < card S" by (rule psubset_card_mono, metis assms, blast) thus "(\x. rank_of x S) ` S \ {0..x. rank_of x S) S" by (metis strict_mono_on_imp_inj_on rank_strict_mono assms) thus "card {0.. card ((\x. rank_of x S) ` S)" by (simp add:card_image) qed lemma card_least: assumes "finite S" shows "card (least k S) = min k (card S)" proof (cases "card S < k") case True have "\t. rank_of t S \ card S" unfolding rank_of_def using assms by (intro card_mono, auto) hence "\t. rank_of t S < k" by (metis True not_less_iff_gr_or_eq order_less_le_trans) hence "least k S = S" by (simp add:least_def) then show ?thesis using True by simp next case False hence a:"card S \ k" using leI by blast hence "card ((\x. rank_of x S) -` {0.. S) = card {0.. S" by (simp add:least_def) lemma least_mono_commute: assumes "finite S" - assumes "strict_mono_on f S" + assumes "strict_mono_on S f" shows "f ` least k S = least k (f ` S)" proof - have a:"inj_on f S" using strict_mono_on_imp_inj_on[OF assms(2)] by simp have "card (least k (f ` S)) = min k (card (f ` S))" by (subst card_least, auto simp add:assms) also have "... = min k (card S)" by (subst card_image, metis a, auto) also have "... = card (least k S)" by (subst card_least, auto simp add:assms) also have "... = card (f ` least k S)" by (subst card_image[OF inj_on_subset[OF a]], simp_all add:least_def) finally have b: "card (least k (f ` S)) \ card (f ` least k S)" by simp have c: "f ` least k S \least k (f ` S)" using assms by (intro image_subsetI) (simp add:least_def rank_mono_commute[symmetric, where T="S"]) show ?thesis using b c assms by (intro card_seteq, simp_all add:least_def) qed lemma least_eq_iff: assumes "finite B" assumes "A \ B" assumes "\x. x \ B \ rank_of x B < k \ x \ A" shows "least k A = least k B" proof - have "least k B \ least k A" using assms rank_mono_2[OF assms(1,2)] order_le_less_trans by (simp add:least_def, blast) moreover have "card (least k B) \ card (least k A)" using assms finite_subset[OF assms(2,1)] card_mono[OF assms(1,2)] by (simp add: card_least min_le_iff_disj) moreover have "finite (least k A)" using finite_subset least_subset assms(1,2) by metis ultimately show ?thesis by (intro card_seteq[symmetric], simp_all) qed lemma least_insert: assumes "finite S" shows "least k (insert x (least k S)) = least k (insert x S)" (is "?lhs = ?rhs") proof (rule least_eq_iff) show "finite (insert x S)" using assms(1) by simp show "insert x (least k S) \ insert x S" using least_subset by blast show "y \ insert x (least k S)" if a: "y \ insert x S" and b: "rank_of y (insert x S) < k" for y proof - have "rank_of y S \ rank_of y (insert x S)" using assms by (intro rank_mono_2, auto) also have "... < k" using b by simp finally have "rank_of y S < k" by simp hence "y = x \ (y \ S \ rank_of y S < k)" using a by simp thus ?thesis by (simp add:least_def) qed qed definition count_le where "count_le x M = size {#y \# M. y \ x#}" definition count_less where "count_less x M = size {#y \# M. y < x#}" definition nth_mset :: "nat \ ('a :: linorder) multiset \ 'a" where "nth_mset k M = sorted_list_of_multiset M ! k" lemma nth_mset_bound_left: assumes "k < size M" assumes "count_less x M \ k" shows "x \ nth_mset k M" proof (rule ccontr) define xs where "xs = sorted_list_of_multiset M" have s_xs: "sorted xs" by (simp add:xs_def sorted_sorted_list_of_multiset) have l_xs: "k < length xs" using assms(1) by (simp add:xs_def size_mset[symmetric]) have M_xs: "M = mset xs" by (simp add:xs_def) hence a:"\i. i \ k \ xs ! i \ xs ! k" using s_xs l_xs sorted_iff_nth_mono by blast assume "\(x \ nth_mset k M)" hence "x > nth_mset k M" by simp hence b:"x > xs ! k" by (simp add:nth_mset_def xs_def[symmetric]) have "k < card {0..k}" by simp also have "... \ card {i. i < length xs \ xs ! i < x}" using a b l_xs order_le_less_trans by (intro card_mono subsetI) auto also have "... = length (filter (\y. y < x) xs)" by (subst length_filter_conv_card, simp) also have "... = size (mset (filter (\y. y < x) xs))" by (subst size_mset, simp) also have "... = count_less x M" by (simp add:count_less_def M_xs) also have "... \ k" using assms by simp finally show "False" by simp qed lemma nth_mset_bound_left_excl: assumes "k < size M" assumes "count_le x M \ k" shows "x < nth_mset k M" proof (rule ccontr) define xs where "xs = sorted_list_of_multiset M" have s_xs: "sorted xs" by (simp add:xs_def sorted_sorted_list_of_multiset) have l_xs: "k < length xs" using assms(1) by (simp add:xs_def size_mset[symmetric]) have M_xs: "M = mset xs" by (simp add:xs_def) hence a:"\i. i \ k \ xs ! i \ xs ! k" using s_xs l_xs sorted_iff_nth_mono by blast assume "\(x < nth_mset k M)" hence "x \ nth_mset k M" by simp hence b:"x \ xs ! k" by (simp add:nth_mset_def xs_def[symmetric]) have "k+1 \ card {0..k}" by simp also have "... \ card {i. i < length xs \ xs ! i \ xs ! k}" using a b l_xs order_le_less_trans by (intro card_mono subsetI, auto) also have "... \ card {i. i < length xs \ xs ! i \ x}" using b by (intro card_mono subsetI, auto) also have "... = length (filter (\y. y \ x) xs)" by (subst length_filter_conv_card, simp) also have "... = size (mset (filter (\y. y \ x) xs))" by (subst size_mset, simp) also have "... = count_le x M" by (simp add:count_le_def M_xs) also have "... \ k" using assms by simp finally show "False" by simp qed lemma nth_mset_bound_right: assumes "k < size M" assumes "count_le x M > k" shows "nth_mset k M \ x" proof (rule ccontr) define xs where "xs = sorted_list_of_multiset M" have s_xs: "sorted xs" by (simp add:xs_def sorted_sorted_list_of_multiset) have l_xs: "k < length xs" using assms(1) by (simp add:xs_def size_mset[symmetric]) have M_xs: "M = mset xs" by (simp add:xs_def) assume "\(nth_mset k M \ x)" hence "x < nth_mset k M" by simp hence "x < xs ! k" by (simp add:nth_mset_def xs_def[symmetric]) hence a:"\i. i < length xs \ xs ! i \ x \ i < k" using s_xs l_xs sorted_iff_nth_mono leI by fastforce have "count_le x M = size (mset (filter (\y. y \ x) xs))" by (simp add:count_le_def M_xs) also have "... = length (filter (\y. y \ x) xs)" by (subst size_mset, simp) also have "... = card {i. i < length xs \ xs ! i \ x}" by (subst length_filter_conv_card, simp) also have "... \ card {i. i < k}" using a by (intro card_mono subsetI, auto) also have "... = k" by simp finally have "count_le x M \ k" by simp thus "False" using assms by simp qed lemma nth_mset_commute_mono: assumes "mono f" assumes "k < size M" shows "f (nth_mset k M) = nth_mset k (image_mset f M)" proof - have a:"k < length (sorted_list_of_multiset M)" by (metis assms(2) mset_sorted_list_of_multiset size_mset) show ?thesis using a by (simp add:nth_mset_def sorted_list_of_multiset_image_commute[OF assms(1)]) qed lemma nth_mset_max: assumes "size A > k" assumes "\x. x \ nth_mset k A \ count A x \ 1" shows "nth_mset k A = Max (least (k+1) (set_mset A))" and "card (least (k+1) (set_mset A)) = k+1" proof - define xs where "xs = sorted_list_of_multiset A" have k_bound: "k < length xs" unfolding xs_def by (metis size_mset mset_sorted_list_of_multiset assms(1)) have A_def: "A = mset xs" by (simp add:xs_def) have s_xs: "sorted xs" by (simp add:xs_def sorted_sorted_list_of_multiset) have "\x. x \ xs ! k \ count A x \ Suc 0" using assms(2) by (simp add:xs_def[symmetric] nth_mset_def) hence no_col: "\x. x \ xs ! k \ count_list xs x \ 1" by (simp add:A_def count_mset) have inj_xs: "inj_on (\k. xs ! k) {0..k}" by (rule inj_onI, simp) (metis (full_types) count_list_ge_2_iff k_bound no_col le_neq_implies_less linorder_not_le order_le_less_trans s_xs sorted_iff_nth_mono) have "\y. y < length xs \ rank_of (xs ! y) (set xs) < k+1 \ y < k+1" proof (rule ccontr) fix y assume b:"y < length xs" assume "\y < k +1" hence a:"k + 1 \ y" by simp have d:"Suc k < length xs" using a b by simp have "k+1 = card ((!) xs ` {0..k})" by (subst card_image[OF inj_xs], simp) also have "... \ rank_of (xs ! (k+1)) (set xs)" unfolding rank_of_def using k_bound by (intro card_mono image_subsetI conjI, simp_all) (metis count_list_ge_2_iff no_col not_le le_imp_less_Suc s_xs sorted_iff_nth_mono d order_less_le) also have "... \ rank_of (xs ! y) (set xs)" unfolding rank_of_def by (intro card_mono subsetI, simp_all) (metis Suc_eq_plus1 a b s_xs order_less_le_trans sorted_iff_nth_mono) also assume "... < k+1" finally show "False" by force qed moreover have "rank_of (xs ! y) (set xs) < k+1" if a:"y < k + 1" for y proof - have "rank_of (xs ! y) (set xs) \ card ((\k. xs ! k) ` {k. k < length xs \ xs ! k < xs ! y})" unfolding rank_of_def by (intro card_mono subsetI, simp) (metis (no_types, lifting) imageI in_set_conv_nth mem_Collect_eq) also have "... \ card {k. k < length xs \ xs ! k < xs ! y}" by (rule card_image_le, simp) also have "... \ card {k. k < y}" by (intro card_mono subsetI, simp_all add:not_less) (metis sorted_iff_nth_mono s_xs linorder_not_less) also have "... = y" by simp also have "... < k + 1" using a by simp finally show "rank_of (xs ! y) (set xs) < k+1" by simp qed ultimately have rank_conv: "\y. y < length xs \ rank_of (xs ! y) (set xs) < k+1 \ y < k+1" by blast have "y \ xs ! k" if a:"y \ least (k+1) (set xs)" for y proof - have "y \ set xs" using a least_subset by blast then obtain i where i_bound: "i < length xs" and y_def: "y = xs ! i" using in_set_conv_nth by metis hence "rank_of (xs ! i) (set xs) < k+1" using a y_def i_bound by (simp add: least_def) hence "i < k+1" using rank_conv i_bound by blast hence "i \ k" by linarith hence "xs ! i \ xs ! k" using s_xs i_bound k_bound sorted_nth_mono by blast thus "y \ xs ! k" using y_def by simp qed moreover have "xs ! k \ least (k+1) (set xs)" using k_bound rank_conv by (simp add:least_def) ultimately have "Max (least (k+1) (set xs)) = xs ! k" by (intro Max_eqI finite_subset[OF least_subset], auto) hence "nth_mset k A = Max (K_Smallest.least (Suc k) (set xs))" by (simp add:nth_mset_def xs_def[symmetric]) also have "... = Max (least (k+1) (set_mset A))" by (simp add:A_def) finally show "nth_mset k A = Max (least (k+1) (set_mset A))" by simp have "k + 1 = card ((\i. xs ! i) ` {0..k})" by (subst card_image[OF inj_xs], simp) also have "... \ card (least (k+1) (set xs))" using rank_conv k_bound by (intro card_mono image_subsetI finite_subset[OF least_subset], simp_all add:least_def) finally have "card (least (k+1) (set xs)) \ k+1" by simp moreover have "card (least (k+1) (set xs)) \ k+1" by (subst card_least, simp, simp) ultimately have "card (least (k+1) (set xs)) = k+1" by simp thus "card (least (k+1) (set_mset A)) = k+1" by (simp add:A_def) qed end diff --git a/thys/Ordinal_Partitions/Erdos_Milner.thy b/thys/Ordinal_Partitions/Erdos_Milner.thy --- a/thys/Ordinal_Partitions/Erdos_Milner.thy +++ b/thys/Ordinal_Partitions/Erdos_Milner.thy @@ -1,1241 +1,1241 @@ theory Erdos_Milner imports Partitions begin subsection \Erdős-Milner theorem\ text \P. Erdős and E. C. Milner, A Theorem in the Partition Calculus. Canadian Math. Bull. 15:4 (1972), 501-505. Corrigendum, Canadian Math. Bull. 17:2 (1974), 305.\ text \The paper defines strong types as satisfying the criteria below. It remarks that ordinals satisfy those criteria. Here is a (too complicated) proof.\ proposition strong_ordertype_eq: assumes D: "D \ elts \" and "Ord \" obtains L where "\(List.set L) = D" "\X. X \ List.set L \ indecomposable (tp X)" and "\M. \M \ D; \X. X \ List.set L \ tp (M \ X) \ tp X\ \ tp M = tp D" proof - define \ where "\ \ inv_into D (ordermap D VWF)" then have bij_\: "bij_betw \ (elts (tp D)) D" using D bij_betw_inv_into down ordermap_bij by blast have \_cancel_left: "\d. d \ D \ \ (ordermap D VWF d) = d" by (metis D \_def bij_betw_inv_into_left down ordermap_bij total_on_VWF wf_VWF) have \_cancel_right: "\\. \ \ elts (tp D) \ ordermap D VWF (\ \) = \" by (metis \_def f_inv_into_f ordermap_surj subsetD) have "small D" "D \ ON" using assms down elts_subset_ON [of \] by auto then have \_less_iff: "\\ \. \\ \ elts (tp D); \ \ elts (tp D)\ \ \ \ < \ \ \ \ < \" by (metis \_def inv_ordermap_VWF_mono_iff inv_ordermap_mono_eq less_V_def) obtain \s where "List.set \s \ ON" and "\_dec \s" and tpD_eq: "tp D = \_sum \s" using Ord_ordertype \_nf_exists by blast \ \Cantor normal form of the ordertype\ have ord [simp]: "Ord (\s ! k)" "Ord (\_sum (take k \s))" if "k < length \s" for k using that \list.set \s \ ON\ by (auto simp: dual_order.trans set_take_subset elim: ON_imp_Ord) define E where "E \ \k. lift (\_sum (take k \s)) (\\(\s!k))" define L where "L \ map (\k. \ ` (elts (E k))) [0..s]" have lengthL [simp]: "length L = length \s" by (simp add: L_def) have in_elts_E_less: "elts (E k') \ elts (E k)" if "k's" for k k' \ \The ordinals have been partitioned into disjoint intervals\ proof - have ord\: "Ord (\ \ \s ! k')" using that by auto from that id_take_nth_drop [of k' "take k \s"] obtain l where "take k \s = take k' \s @ (\s!k') # l" by (simp add: min_def) then show ?thesis using that unfolding E_def lift_def less_sets_def by (auto dest!: OrdmemD [OF ord\] intro: less_le_trans) qed have elts_E: "elts (E k) \ elts (\_sum \s)" if "k < length \s" for k proof - have "\ \ (\s!k) \ \_sum (drop k \s)" by (metis that Cons_nth_drop_Suc \_sum_Cons add_le_cancel_left0) then have "(+) (\_sum (take k \s)) ` elts (\ \ (\s!k)) \ elts (\_sum (take k \s) + \_sum (drop k \s))" by blast also have "\ = elts (\_sum \s)" using \_sum_take_drop by auto finally show ?thesis by (simp add: lift_def E_def) qed have \_sum_in_tpD: "\_sum (take k \s) + \ \ elts (tp D)" if "\ \ elts (\ \ \s!k)" "k < length \s" for \ k using elts_E lift_def that tpD_eq by (auto simp: E_def) have Ord_\: "Ord (\ (\_sum (take k \s) + \))" if "\ \ elts (\ \ \s!k)" "k < length \s" for \ k using \_sum_in_tpD \D \ ON\ bij_\ bij_betw_imp_surj_on that by fastforce define \ where "\ \ \k. ((\y. odiff y (\_sum (take k \s))) \ ordermap D VWF)" \ \mapping the segments of @{term D} into some power of @{term \}\ have bij_\: "bij_betw (\ k) (\ ` elts (E k)) (elts (\ \ (\s!k)))" if "k < length \s" for k using that by (auto simp: bij_betw_def \_def E_def inj_on_def lift_def image_comp \_sum_in_tpD \_cancel_right that) have \_iff: "\ k x < \ k y \ (x,y) \ VWF" if "x \ \ ` elts (E k)" "y \ \ ` elts (E k)" "k < length \s" for k x y using that by (auto simp: \_def E_def lift_def \_sum_in_tpD \_cancel_right Ord_\ \_less_iff) have tp_E_eq [simp]: "tp (\ ` elts (E k)) = \\(\s!k)" if k: "k < length \s" for k by (metis Ord_\ Ord_oexp \_iff bij_\ ord(1) ordertype_VWF_eq_iff replacement small_elts that) have tp_L_eq [simp]: "tp (L!k) = \\(\s!k)" if "k < length \s" for k by (simp add: L_def that) have UL_eq_D: "\ (list.set L) = D" proof (rule antisym) show "\ (list.set L) \ D" by (force simp: L_def tpD_eq bij_betw_apply [OF bij_\] dest: elts_E) show "D \ \ (list.set L)" proof fix \ assume "\ \ D" then have "ordermap D VWF \ \ elts (\_sum \s)" by (metis \small D\ ordermap_in_ordertype tpD_eq) then show "\ \ \ (list.set L)" using \\ \ D\ \_cancel_left in_elts_\_sum by (fastforce simp: L_def E_def image_iff lift_def) qed qed show thesis proof show "indecomposable (tp X)" if "X \ list.set L" for X using that by (auto simp: L_def indecomposable_\_power) next fix M assume "M \ D" and *: "\X. X \ list.set L \ tp X \ tp (M \ X)" show "tp M = tp D" proof (rule antisym) show "tp M \ tp D" by (simp add: \M \ D\ \small D\ ordertype_VWF_mono) define \ where "\ \ \X. inv_into (M \ X) (ordermap (M \ X) VWF)" \ \The bijection from an @{term \}-power into the appropriate segment of @{term M}\ have bij_\: "bij_betw (\ X) (elts (tp (M \ X))) (M \ X)" for X unfolding \_def by (meson \M \ D\ \small D\ bij_betw_inv_into inf_le1 ordermap_bij smaller_than_small total_on_VWF wf_VWF) have Ord_\: "Ord (\ X \)" if "\ \ elts (tp (M \ X))" for \ X using \D \ ON\ \M \ D\ bij_betw_apply [OF bij_\] that by blast have \_cancel_right: "\\ X. \ \ elts (tp (M \ X)) \ ordermap (M \ X) VWF (\ X \) = \" by (metis \_def f_inv_into_f ordermap_surj subsetD) have smM: "small (M \ X)" for X by (meson \M \ D\ \small D\ inf_le1 subset_iff_less_eq_V) have "\k < length \s. \ \ elts (E k)" if \: "\ \ elts (tp D)" for \ proof - obtain X where "X \ set L" and X: "\ \ \ X" by (metis UL_eq_D \ Union_iff \_def in_mono inv_into_into ordermap_surj) then have "\k < length \s. \ \ elts (E k) \ X = \ ` elts (E k)" apply (clarsimp simp: L_def) by (metis \ \_cancel_right elts_E in_mono tpD_eq) then show ?thesis by blast qed then obtain K where K: "\\. \ \ elts (tp D) \ K \ < length \s \ \ \ elts (E (K \))" by metis \ \The index from @{term "tp D"} to the appropriate segment number\ define \ where "\ \ \d. \ (L ! K (ordermap D VWF d)) (\ (K (ordermap D VWF d)) d)" show "tp D \ tp M" proof (rule ordertype_inc_le) show "small D" "small M" using \M \ D\ \small D\ subset_iff_less_eq_V by auto next fix d' d assume xy: "d' \ D" "d \ D" and "(d',d) \ VWF" then have "d' < d" using ON_imp_Ord \D \ ON\ by auto let ?\' = "ordermap D VWF d'" let ?\ = "ordermap D VWF d" have len': "K ?\' < length \s" and elts': "?\' \ elts (E (K ?\'))" and len: "K ?\ < length \s" and elts: "?\ \ elts (E (K ?\))" using K \d' \ D\ \d \ D\ by (auto simp: \small D\ ordermap_in_ordertype) have Ord_\L: "Ord (\ (L!k) (\ k d))" if "d \ \ ` elts (E k)" "k < length \s" for k d by (metis (mono_tags) "*" Ord_\ bij_\ bij_betw_apply lengthL nth_mem that tp_L_eq vsubsetD) have "?\' < ?\" by (simp add: \(d', d) \ VWF\ \small D\ ordermap_mono_less xy) then have "K ?\' \ K ?\" using elts' elts in_elts_E_less by (meson leI len' less_asym less_sets_def) then consider (less) "K ?\' < K ?\" | (equal) "K ?\' = K ?\" by linarith then have "\ (L ! K ?\') (\ (K ?\') d') < \ (L ! K ?\) (\ (K ?\) d)" proof cases case less obtain \: "\ (L ! K ?\') (\ (K ?\') d') \ M \ L ! K ?\'" "\ (L ! K ?\) (\ (K ?\) d) \ M \ L ! K ?\" using elts' elts len' len * [THEN vsubsetD] by (metis lengthL \_cancel_left bij_\ bij_\ bij_betw_imp_surj_on imageI nth_mem tp_L_eq xy) then have "ordermap D VWF (\ (L ! K ?\') (\ (K ?\') d')) \ elts (E (K ?\'))" "ordermap D VWF (\ (L ! K ?\) (\ (K ?\) d)) \ elts (E (K ?\))" using len' len elts_E tpD_eq by (fastforce simp: L_def \_cancel_right)+ then have "ordermap D VWF (\ (L ! K ?\') (\ (K ?\') d')) < ordermap D VWF (\ (L ! K ?\) (\ (K ?\) d))" using in_elts_E_less len less by (meson less_sets_def) moreover have "\ (L ! K ?\') (\ (K ?\') d') \ D" "\ (L ! K ?\) (\ (K ?\) d) \ D" using \M \ D\ \ by auto ultimately show ?thesis by (metis \small D\ \_cancel_left \_less_iff ordermap_in_ordertype) next case equal have \_less: "\X \ \. \\ < \; \ \ elts (tp (M \ X)); \ \ elts (tp (M \ X))\ \ \ X \ < \ X \" by (metis \D \ ON\ \M \ D\ \_def dual_order.trans inv_ordermap_VWF_strict_mono_iff le_infI1 smM) have "\ (K ?\) d' < \ (K ?\) d" by (metis equal \(d', d) \ VWF\ \_cancel_left \_iff elts elts' imageI len xy) then show ?thesis unfolding equal by (metis * [THEN vsubsetD] lengthL \_cancel_left \_less bij_\ bij_betw_imp_surj_on elts elts' image_eqI len local.equal nth_mem tp_L_eq xy) qed moreover have "Ord (\ (L ! K ?\') (\ (K ?\') d'))" "Ord (\ (L ! K ?\) (\ (K ?\) d))" using Ord_\L \_cancel_left elts len elts' len' xy by fastforce+ ultimately show "(\ d', \ d) \ VWF" by (simp add: \_def) next show "\ ` D \ M" proof (clarsimp simp: \_def) fix d assume "d \ D" let ?\ = "ordermap D VWF d" have len: "K ?\ < length \s" and elts: "?\ \ elts (E (K ?\))" using K \d \ D\ by (auto simp: \small D\ ordermap_in_ordertype) have "\ (K ?\) d \ elts (tp (L! (K ?\)))" using bij_\ [OF len] \d \ D\ by (metis \_cancel_left bij_betw_apply elts imageI len tp_L_eq) then show "\ (L ! K (ordermap D VWF d)) (\ (K (ordermap D VWF d)) d) \ M" by (metis "*" lengthL Int_iff bij_\ bij_betw_apply len nth_mem vsubsetD) qed qed auto qed qed (simp add: UL_eq_D) qed text \The ``remark'' of Erdős and E. C. Milner, Canad. Math. Bull. Vol. 17 (2), 1974\ proposition indecomposable_imp_Ex_less_sets: assumes indec: "indecomposable \" and "\ \ \" and A: "tp A = \" "small A" "A \ ON" and "x \ A" and A1: "tp A1 = \" "A1 \ A" obtains A2 where "tp A2 = \" "A2 \ A1" "{x} \ A2" proof - have "Ord \" using indec indecomposable_imp_Ord by blast have "Limit \" by (meson \_gt1 assms indec indecomposable_imp_Limit less_le_trans) define \ where "\ \ inv_into A (ordermap A VWF)" then have bij_\: "bij_betw \ (elts \) A" using A bij_betw_inv_into down ordermap_bij by blast have bij_om: "bij_betw (ordermap A VWF) A (elts \)" using A down ordermap_bij by blast define \ where "\ \ ordermap A VWF x" have \: "\ \ elts \" unfolding \_def using A \x \ A\ down by auto then have "Ord \" using Ord_in_Ord \Ord \\ by blast define B where "B \ \ ` (elts (succ \))" have B: "{y \ A. ordermap A VWF y \ ordermap A VWF x} \ B" apply (clarsimp simp add: B_def \_def image_iff \_def) by (metis Ord_linear Ord_ordermap OrdmemD bij_betw_inv_into_left bij_om leD) show thesis proof have "small A1" by (meson \small A\ A1 smaller_than_small) then have "tp (A1 - B) \ tp A1" by (simp add: ordertype_VWF_mono) moreover have "tp (A1 - B) \ \" proof - have "\ (\ \ tp B)" unfolding B_def proof (subst ordertype_VWF_inc_eq) show "elts (succ \) \ ON" by (auto simp: \_def ordertype_VWF_inc_eq intro: Ord_in_Ord) have sub: "elts (succ \) \ elts \" using Ord_trans \ \Ord \\ by auto then show "\ ` elts (succ \) \ ON" using \A \ ON\ bij_\ bij_betw_imp_surj_on by blast have "succ \ \ elts \" using \ Limit_def \Limit \\ by blast with A sub show "\ u < \ v" if "u \ elts (succ \)" and "v \ elts (succ \)" and "u < v" for u v by (metis \_def inv_ordermap_VWF_strict_mono_iff subsetD that) show "\ \ \ tp (elts (succ \))" by (metis Limit_def Ord_succ \ \Limit \\ \Ord \\ mem_not_refl ordertype_eq_Ord vsubsetD) qed auto then show ?thesis using indecomposable_ordertype_ge [OF indec, of A1 B] \small A1\ A1 by (auto simp: B_def) qed ultimately show "tp (A1 - B) = \" using A1 by blast have "{x} \ (A - B)" using assms B apply (clarsimp simp: less_sets_def \_def subset_iff) by (metis Ord_not_le VWF_iff_Ord_less less_V_def order_refl ordermap_mono_less trans_VWF wf_VWF) with \A1 \ A\ show "{x} \ (A1 - B)" by auto qed auto qed text \the main theorem, from which they derive the headline result\ theorem Erdos_Milner_aux: assumes part: "partn_lst_VWF \ [k, \] 2" and indec: "indecomposable \" and "k > 1" "Ord \" and \: "\ \ elts \1" shows "partn_lst_VWF (\*\) [ord_of_nat (2*k), min \ (\*\)] 2" proof (cases "\\1 \ \=0") case True have "Ord \" using indec indecomposable_def by blast show ?thesis proof (cases "\=0") case True then show ?thesis by (simp add: partn_lst_triv0 [where i=1]) next case False then consider (0) "\=0" | (1) "\=1" by (metis Ord_0 OrdmemD True \Ord \\ mem_0_Ord one_V_def order.antisym succ_le_iff) then show ?thesis proof cases case 0 with part show ?thesis by (force simp add: partn_lst_def nsets_empty_iff less_Suc_eq) next case 1 then obtain "Ord \" by (meson ON_imp_Ord Ord_\1 True \ elts_subset_ON) then obtain i where "i < Suc (Suc 0)" "[ord_of_nat k, \] ! i \ \" using partn_lst_VWF_nontriv [OF part] 1 by auto then have "\ \ 1" using \\=1\ \k > 1\ by (fastforce simp: less_Suc_eq) then have "min \ (\*\) \ 1" by (metis Ord_1 Ord_\ Ord_linear_le Ord_mult \Ord \\ min_def order_trans) then show ?thesis using False by (auto simp: True \Ord \\ \\\0\ \\=1\ intro!: partn_lst_triv1 [where i=1]) qed qed next case False then have "\ \ \" by (meson Ord_1 Ord_not_less indec indecomposable_def indecomposable_imp_Limit omega_le_Limit) have "0 \ elts \" "\ \ 0" using False Ord_\1 Ord_in_Ord \ mem_0_Ord by blast+ show ?thesis unfolding partn_lst_def proof clarsimp fix f assume "f \ [elts (\*\)]\<^bsup>2\<^esup> \ {.. [elts (\*\)]\<^bsup>2\<^esup> \ {..<2::nat}" by (simp add: eval_nat_numeral) obtain ord [iff]: "Ord \" "Ord \" "Ord (\*\)" using Ord_\1 Ord_in_Ord \ indec indecomposable_imp_Ord Ord_mult by blast have *: False if i [rule_format]: "\H. tp H = ord_of_nat (2*k) \ H \ elts (\*\) \ \ f ` [H]\<^bsup>2\<^esup> \ {0}" and ii [rule_format]: "\H. tp H = \ \ H \ elts (\*\) \ \ f ` [H]\<^bsup>2\<^esup> \ {1}" and iii [rule_format]: "\H. tp H = (\*\) \ H \ elts (\*\) \ \ f ` [H]\<^bsup>2\<^esup> \ {1}" proof - have Ak0: "\X \ [A]\<^bsup>k\<^esup>. f ` [X]\<^bsup>2\<^esup> \ {0}" \\remark (8) about @{term"A \ S"}\ if A_\\: "A \ elts (\*\)" and ot: "tp A \ \" for A proof - let ?g = "inv_into A (ordermap A VWF)" have "small A" using down that by auto then have inj_g: "inj_on ?g (elts \)" by (meson inj_on_inv_into less_eq_V_def ordermap_surj ot subset_trans) have om_A_less: "\x y. \x \ A; y \ A; x < y\ \ ordermap A VWF x < ordermap A VWF y" using ord by (meson A_\\ Ord_in_Ord VWF_iff_Ord_less \small A\ in_mono ordermap_mono_less trans_VWF wf_VWF) have \_sub: "elts \ \ ordermap A VWF ` A" by (metis \small A\ elts_of_set less_eq_V_def ordertype_def ot replacement) have g: "?g \ elts \ \ elts (\*\)" by (meson A_\\ Pi_I' \_sub inv_into_into subset_eq) then have fg: "f \ (\X. ?g ` X) \ [elts \]\<^bsup>2\<^esup> \ {..<2}" by (rule nsets_compose_image_funcset [OF f _ inj_g]) obtain i H where "i < 2" "H \ elts \" and ot_eq: "tp H = [k,\]!i" "(f \ (\X. ?g ` X)) ` (nsets H 2) \ {i}" using ii partn_lst_E [OF part fg] by (auto simp: eval_nat_numeral) then consider (0) "i=0" | (1) "i=1" by linarith then show ?thesis proof cases case 0 then have "f ` [inv_into A (ordermap A VWF) ` H]\<^bsup>2\<^esup> \ {0}" using ot_eq \H \ elts \\ \_sub by (auto simp: nsets_def [of _ k] inj_on_inv_into elim!: nset_image_obtains) moreover have "finite H \ card H = k" using 0 ot_eq \H \ elts \\ down by (simp add: finite_ordertype_eq_card) then have "inv_into A (ordermap A VWF) ` H \ [A]\<^bsup>k\<^esup>" using \H \ elts \\ \_sub by (auto simp: nsets_def [of _ k] card_image inj_on_inv_into inv_into_into) ultimately show ?thesis by blast next case 1 have gH: "?g ` H \ elts (\*\)" by (metis A_\\ \_sub \H \ elts \\ image_subsetI inv_into_into subset_eq) have g_less: "?g x < ?g y" if "x < y" "x \ elts \" "y \ elts \" for x y using Pi_mem [OF g] ord that by (meson A_\\ Ord_in_Ord Ord_not_le \small A\ dual_order.trans elts_subset_ON inv_ordermap_VWF_mono_le ot vsubsetD) have [simp]: "tp (?g ` H) = tp H" by (meson \H \ elts \\ ord down dual_order.trans elts_subset_ON gH g_less ordertype_VWF_inc_eq subsetD) show ?thesis using ii [of "?g ` H"] subset_inj_on [OF inj_g \H \ elts \\] ot_eq 1 by (auto simp: gH elim!: nset_image_obtains) qed qed define K where "K \ \i x. {y \ elts (\*\). x \ y \ f{x,y} = i}" have small_K: "small (K i x)" for i x by (simp add: K_def) define KI where "KI \ \i X. (\x\X. K i x)" have KI_disj_self: "X \ KI i X = {}" for i X by (auto simp: KI_def K_def) define M where "M \ \D \ x. {\::V. \ \ D \ tp (K 1 x \ \ \) \ \}" have M_sub_D: "M D \ x \ D" for D \ x by (auto simp: M_def) have small_M [simp]: "small (M D \ x)" if "small D" for D \ x by (simp add: M_def that) have 9: "tp {x \ A. tp (M D \ x) \ tp D} \ \" (is "ordertype ?AD _ \ \") if inD: "indecomposable (tp D)" and D: "D \ elts \" and A: "A \ elts (\*\)" and tpA: "tp A = \" and \: "\ \ D \ {X. X \ elts (\*\) \ tp X = \}" for D A \ \\remark (9), assuming an indecomposable order type\ proof (rule ccontr) define A' where "A' \ {x \ A. \ tp (M D \ x) \ tp D}" have small [iff]: "small A" "small D" using A D down by (auto simp: M_def) have small_\: "small (\ \)" if "\ \ D" for \ using that \ by (auto simp: Pi_iff subset_iff_less_eq_V) assume not_\_le: "\ \ \ tp {x \ A. tp (M D \ x) \ tp D}" moreover obtain "small A" "small A'" "A' \ A" and A'_sub: "A' \ elts (\*\)" using A'_def A down by auto moreover have "A' = A - ?AD" by (force simp: A'_def) ultimately have A'_ge: "tp A' \ \" by (metis (no_types, lifting) dual_order.refl indec indecomposable_ordertype_eq mem_Collect_eq subsetI tpA) obtain X where "X \ A'" "finite X" "card X = k" and fX0: "f ` [X]\<^bsup>2\<^esup> \ {0}" using Ak0 [OF A'_sub A'_ge] by (auto simp: nsets_def [of _ k]) then have \: "\ tp (M D \ x) \ tp D" if "x \ X" for x using that by (auto simp: A'_def) obtain x where "x \ X" using \card X = k\ \k>1\ by fastforce have "\ D \ (\ x\X. M D \ x)" proof assume not: "D \ (\x\X. M D \ x)" have "\X\M D \ ` X. tp D \ tp X" proof (rule indecomposable_ordertype_finite_ge [OF inD]) show "M D \ ` X \ {}" using A'_def A'_ge not not_\_le by auto show "small (\ (M D \ ` X))" using \finite X\ by (simp add: finite_imp_small) qed (use \finite X\ not in auto) then show False by (simp add: \) qed then obtain \ where "\ \ D" and \: "\ \ (\ x\X. M D \ x)" by blast define \ where "\ \ {KI 0 X \ \ \, \x\X. K 1 x \ \ \, X \ \ \}" have \\: "X \ elts (\*\)" "\ \ \ elts (\*\)" using A'_sub \X \ A'\ \ \\ \ D\ by auto then have "KI 0 X \ (\x\X. K 1 x) \ X = elts (\*\)" using \x \ X\ f by (force simp: K_def KI_def Pi_iff less_2_cases_iff) with \\ have \\_\: "finite \" "\ \ \ \\" by (auto simp: \_def) then have "\ tp (K 1 x \ \ \) \ \" if "x \ X" for x using that \\ \ D\ \ \k > 1\ \card X = k\ by (fastforce simp: M_def) moreover have sm_K1: "small (\x\X. K 1 x \ \ \)" by (meson Finite_V Int_lower2 \\ \ D\ \finite X\ small_\ small_UN smaller_than_small) ultimately have not1: "\ tp (\x\X. K 1 x \ \ \) \ \" using \finite X\ \x \ X\ indecomposable_ordertype_finite_ge [OF indec, of "(\x. K 1 x \ \ \) ` X"] by blast moreover have "\ tp (X \ \ \) \ \" using \finite X\ \\ \ \\ by (meson finite_Int mem_not_refl ordertype_VWF_\ vsubsetD) moreover have "\ \ tp (\ \)" using \ \\ \ D\ small_\ by fastforce+ moreover have "small (\ \)" using \\ \ D\ small_\ by (fastforce simp: \_def intro: smaller_than_small sm_K1) ultimately have K0\_ge: "tp (KI 0 X \ \ \) \ \" using indecomposable_ordertype_finite_ge [OF indec \\_\] by (auto simp: \_def) have \\: "\ \ \ elts (\*\)" "tp (\ \) = \" using \\ \ D\ \ by blast+ then obtain Y where Ysub: "Y \ KI 0 X \ \ \" and "finite Y" "card Y = k" and fY0: "f ` [Y]\<^bsup>2\<^esup> \ {0}" using Ak0 [OF _ K0\_ge] by (auto simp: nsets_def [of _ k]) have \: "X \ Y = {}" using Ysub KI_disj_self by blast then have "card (X \ Y) = 2*k" by (simp add: \card X = k\ \card Y = k\ \finite X\ \finite Y\ card_Un_disjoint) moreover have "X \ Y \ elts (\*\)" using A'_sub \X \ A'\ \\(1) \Y \ KI 0 X \ \ \\ by auto moreover have "f ` [X \ Y]\<^bsup>2\<^esup> \ {0}" using fX0 fY0 Ysub by (auto simp: \ nsets_disjoint_2 image_Un image_UN KI_def K_def) ultimately show False using i \finite X\ \finite Y\ ordertype_VWF_finite_nat by auto qed have IX: "tp {x \ A. tp (M D \ x) \ tp D} \ \" if D: "D \ elts \" and A: "A \ elts (\*\)" and tpA: "tp A = \" and \: "\ \ D \ {X. X \ elts (\*\) \ tp X = \}" for D A \ \\remark (9) for any order type\ proof - obtain L where UL: "\(List.set L) \ D" and indL: "\X. X \ List.set L \ indecomposable (tp X)" and eqL: "\M. \M \ D; \X. X \ List.set L \ tp (M \ X) \ tp X\ \ tp M = tp D" using ord by (metis strong_ordertype_eq D order_refl) obtain A'' where A'': "A'' \ A" "tp A'' \ \" and "\x X. \x \ A''; X \ List.set L\ \ tp (M D \ x \ X) \ tp X" using UL indL proof (induction L arbitrary: thesis) case (Cons X L) then obtain A'' where A'': "A'' \ A" "tp A'' \ \" and "X \ D" and ge_X: "\x X. \x \ A''; X \ List.set L\ \ tp (M D \ x \ X) \ tp X" by auto then have tp_A'': "tp A'' = \" by (metis A antisym down ordertype_VWF_mono tpA) have ge_\: "tp {x \ A''. tp (M X \ x) \ tp X} \ \" by (rule 9) (use A A'' tp_A'' Cons.prems \D \ elts \\ \X \ D\ \ in auto) let ?A = "{x \ A''. tp (M D \ x \ X) \ tp X}" have M_eq: "M D \ x \ X = M X \ x" if "x \ A''" for x using that \X \ D\ by (auto simp: M_def) show thesis proof (rule Cons.prems) show "\ \ tp ?A" using ge_\ by (simp add: M_eq cong: conj_cong) show "tp Y \ tp (M D \ x \ Y)" if "x \ ?A" "Y \ list.set (X # L)" for x Y using that ge_X by force qed (use A'' in auto) qed (use tpA in auto) then have tp_M_ge: "tp (M D \ x) \ tp D" if "x \ A''" for x using eqL that by (auto simp: M_def) have "\ \ tp A''" by (simp add: A'') also have "\ \ tp {x \ A''. tp (M D \ x) \ tp D}" by (metis (mono_tags, lifting) tp_M_ge eq_iff mem_Collect_eq subsetI) also have "\ \ tp {x \ A. tp D \ tp (M D \ x)}" by (rule ordertype_mono) (use A'' A down in auto) finally show ?thesis . qed have IX': "tp {x \ A'. tp (K 1 x \ A) \ \} \ \" if A: "A \ elts (\*\)" "tp A = \" and A': "A' \ elts (\*\)" "tp A' = \" for A A' proof - have \: "\ \ tp (K 1 t \ A)" if "1 \ tp {\. \ = 0 \ \ \ tp (K 1 t \ A)}" for t using that by (metis Collect_empty_eq less_eq_V_0_iff ordertype_empty zero_neq_one) have "tp {x \ A'. 1 \ tp {\. \ = 0 \ \ \ tp (K 1 x \ A)}} \ tp {x \ A'. \ \ tp (K 1 x \ A)}" by (rule ordertype_mono) (use "\" A' in \auto simp: down\) then show ?thesis using IX [of "{0}" A' "\x. A"] that \0 \ elts \\ by (force simp: M_def) qed - have 10: "\x0 \ A. \g \ elts \ \ elts \. strict_mono_on g (elts \) \ (\\ \ F. g \ = \) + have 10: "\x0 \ A. \g \ elts \ \ elts \. strict_mono_on (elts \) g \ (\\ \ F. g \ = \) \ (\\ \ elts \. tp (K 1 x0 \ \ (g \)) \ \)" if F: "finite F" "F \ elts \" and A: "A \ elts (\*\)" "tp A = \" and \: "\ \ elts \ \ {X. X \ elts (\*\) \ tp X = \}" for F A \ proof - define p where "p \ card F" have "\ \ F" using that by auto - then obtain \ :: "nat \ V" where bij\: "bij_betw \ {..p} (insert \ F)" and mon\: "strict_mono_on \ {..p}" + then obtain \ :: "nat \ V" where bij\: "bij_betw \ {..p} (insert \ F)" and mon\: "strict_mono_on {..p} \" using ZFC_Cardinals.ex_bij_betw_strict_mono_card [of "insert \ F"] elts_subset_ON \Ord \\ F by (simp add: p_def lessThan_Suc_atMost) blast have less_\_I: "\ k < \ l" if "k < l" "l \ p" for k l using mon\ that by (auto simp: strict_mono_on_def) then have less_\_D: "k < l" if "\ k < \ l" "k \ p" for k l by (metis less_asym linorder_neqE_nat that) have Ord_\: "Ord (\ k)" if "k \ p" for k by (metis (no_types, lifting) ON_imp_Ord atMost_iff insert_subset mem_Collect_eq order_trans \F \ elts \\ bij\ bij_betwE elts_subset_ON \Ord \\ that) have le_\0 [simp]: "\j. j \ p \ \ 0 \ \ j" by (metis eq_refl leI le_0_eq less_\_I less_imp_le) have le_\: "\ i \ \ (j - Suc 0)" if "i < j" "j \ p" for i j proof (cases i) case 0 then show ?thesis using le_\0 that by auto next case (Suc i') then show ?thesis by (metis (no_types, opaque_lifting) Suc_pred le_less less_Suc_eq less_Suc_eq_0_disj less_\_I not_less_eq that) qed have [simp]: "\ p = \" proof - obtain k where k: "\ k = \" "k \ p" by (meson atMost_iff bij\ bij_betw_iff_bijections insertI1) then have "k = p \ k < p" by linarith then show ?thesis using bij\ ord k that(2) by (metis OrdmemD atMost_iff bij_betw_iff_bijections insert_iff leD less_\_D order_refl subsetD) qed have F_imp_Ex: "\k < p. \ = \ k" if "\ \ F" for \ proof - obtain k where k: "k \ p" "\ = \ k" by (metis \\ \ F\ atMost_iff bij\ bij_betw_def imageE insert_iff) then show ?thesis using \\ \ F\ \\ p = \\ le_imp_less_or_eq that by blast qed have F_imp_ge: "\ \ \ 0" if "\ \ F" for \ using F_imp_Ex [OF that] by (metis dual_order.order_iff_strict le0 less_\_I) define D where "D \ \k. (if k=0 then {..<\ 0} else {\ (k-1)<..<\ k}) \ elts \" have D\: "D k \ elts \" for k by (auto simp: D_def) then have small_D [simp]: "small (D k)" for k by (meson down) have M_Int_D: "M (elts \) \ x \ D k = M (D k) \ x" if "k \ p" for x k using D\ by (auto simp: M_def) have \_le_if_D: "\ k \ \" if "\ \ D (Suc k)" for \ k using that by (simp add: D_def order.order_iff_strict) have mono_D: "D i \ D j" if "i < j" "j \ p" for i j proof (cases j) case (Suc j') with that show ?thesis apply (simp add: less_sets_def D_def Ball_def) by (metis One_nat_def diff_Suc_1 le_\ less_le_trans less_trans) qed (use that in auto) then have disjnt_DD: "disjnt (D i) (D j)" if "i \ j" "i \ p" "j \ p" for i j by (meson disjnt_sym less_linear less_sets_imp_disjnt that) have UN_D_eq: "(\l \ k. D l) = {..<\ k} \ (elts \ - F)" if "k \ p" for k using that proof (induction k) case 0 then show ?case by (auto simp: D_def F_imp_ge leD) next case (Suc k) have "D (Suc k) \ {..<\ k} \ (elts \ - F) = {..<\ (Suc k)} \ (elts \ - F)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using Suc.prems by (auto simp: D_def if_split_mem2 intro: less_\_I less_trans dest!: less_\_D F_imp_Ex) have "\x. \x < \ (Suc k); x \ elts \; x \ F; \ k \ x\ \ \ k < x" using Suc.prems \F \ elts \\ bij\ le_imp_less_or_eq by (fastforce simp: bij_betw_iff_bijections) then show "?rhs \ ?lhs" using Suc.prems by (auto simp: D_def Ord_not_less Ord_in_Ord [OF \Ord \\] Ord_\ if_split_mem2) qed then show ?case using Suc by (simp add: atMost_Suc) qed have \_decomp: "elts \ = F \ (\k \ p. D k)" using \F \ elts \\ OrdmemD [OF \Ord \\] by (auto simp: UN_D_eq) define \idx where "\idx \ \\. @k. \ \ D k \ k \ p" have \idx: "\ \ D (\idx \) \ \idx \ \ p" if "\ \ elts \ - F" for \ using that by (force simp: \idx_def \_decomp intro: someI_ex del: conjI) have any_imp_\idx: "k = \idx \" if "\ \ D k" "k \ p" for k \ proof (rule ccontr) assume non: "k \ \idx \" have "\ \ F" using that UN_D_eq by auto then show False using disjnt_DD [OF non] by (metis D\ Diff_iff \idx disjnt_iff subsetD that) qed have "\A'. A' \ A \ tp A' = \ \ (\x \ A'. F \ M (elts \) \ x)" using F proof induction case (insert \ F) then obtain A' where "A' \ A" and A': "A' \ elts (\*\)" "tp A' = \" and FN: "\x. x \ A' \ F \ M (elts \) \ x" using A(1) by auto define A'' where "A'' \ {x \ A'. \ \ tp (K 1 x \ \ \)}" have "\ \ elts \" "F \ elts \" using insert by auto note ordertype_eq_Ord [OF \Ord \\, simp] show ?case proof (intro exI conjI) show "A'' \ A" using \A' \ A\ by (auto simp: A''_def) show "tp A'' = \" proof (rule antisym) show "tp A'' \ \" using \A'' \ A\ down ordertype_VWF_mono A by blast have "\ \ \ elts (\*\)" "tp (\ \) = \" using \ \\ \ elts \\ by auto then show "\ \ tp A''" using IX' [OF _ _ A'] by (simp add: A''_def) qed show "\x\A''. insert \ F \ M (elts \) \ x" using A''_def FN M_def \\ \ elts \\ by blast qed qed (use A in auto) then obtain A' where A': "A' \ A" "tp A' = \" and FN: "\x. x \ A' \ F \ M (elts \) \ x" by metis have False - if *: "\x0 g. \x0 \ A; g \ elts \ \ elts \; strict_mono_on g (elts \)\ + if *: "\x0 g. \x0 \ A; g \ elts \ \ elts \; strict_mono_on (elts \) g\ \ (\\\F. g \ \ \) \ (\\\elts \. tp (K 1 x0 \ \ (g \)) < \)" proof - { fix x \ \construction of the monotone map @{term g} mentioned above\ assume "x \ A'" with A' have "x \ A" by blast have "\k. k \ p \ tp (M (D k) \ x) < tp (D k)" (is "?P") proof (rule ccontr) assume "\ ?P" then have le: "tp (D k) \ tp (M (D k) \ x)" if "k \ p" for k by (meson Ord_linear2 Ord_ordertype that) - have "\f\D k \ M (D k) \ x. inj_on f (D k) \ (strict_mono_on f (D k))" + have "\f\D k \ M (D k) \ x. inj_on f (D k) \ (strict_mono_on (D k) f)" if "k \ p" for k using le [OF that] that VWF_iff_Ord_less apply (clarsimp simp: ordertype_le_ordertype strict_mono_on_def) by (metis (full_types) D\ M_sub_D Ord_in_Ord PiE VWF_iff_Ord_less ord(2) subsetD) then obtain h where fun_h: "\k. k \ p \ h k \ D k \ M (D k) \ x" and inj_h: "\k. k \ p \ inj_on (h k) (D k)" - and mono_h: "\k x y. k \ p \ strict_mono_on (h k) (D k)" + and mono_h: "\k x y. k \ p \ strict_mono_on (D k) (h k)" by metis then have fun_hD: "\k. k \ p \ h k \ D k \ D k" by (auto simp: M_def) have h_increasing: "\ \ h k \" if "k \ p" and \: "\ \ D k" for k \ proof (rule Ord_mono_imp_increasing) show "h k \ D k \ D k" by (simp add: fun_hD that(1)) show "D k \ ON" using D\ elts_subset_ON ord(2) by blast qed (auto simp: that mono_h) define g where "g \ \\. if \ \ F then \ else h (\idx \) \" have [simp]: "g \ = \" if "\ \ F" for \ using that by (auto simp: g_def) have fun_g: "g \ elts \ \ elts \" proof (rule Pi_I) fix x assume "x \ elts \" then have "x \ D (\idx x)" "\idx x \ p" if "x \ F" using that by (auto simp: \idx) then show "g x \ elts \" using fun_h D\ M_sub_D \x \ elts \\ by (simp add: g_def) blast qed have h_in_D: "h (\idx \) \ \ D (\idx \)" if "\ \ F" "\ \ elts \" for \ using \idx fun_hD that by fastforce have 1: "\ k < h (\idx \) \" if "k < p" and \: "\ \ F" "\ \ elts \" and "\ k < \" for k \ using that h_in_D [OF \] \idx by (fastforce simp: D_def dest: h_increasing split: if_split_asm) moreover have 2: "h (\idx \) \ < \ k" if \: "\ \ F" "\ \ elts \" and "k < p" "\ < \ k" for \ k proof - have "\idx \ \ k" proof (rule ccontr) assume "\ \idx \ \ k" then have "k < \idx \" by linarith then show False using \_le_if_D \idx by (metis Diff_iff Suc_pred le0 leD le_\ le_less_trans \ \\ < \ k\) qed then show ?thesis using that h_in_D [OF \] by (smt (verit, best) Int_lower1 UN_D_eq UN_I atMost_iff lessThan_iff less_imp_le subset_eq) qed moreover have "h (\idx \) \ < h (\idx \) \" if \: "\ \ F" "\ \ elts \" and \: "\ \ F" "\ \ elts \" and "\ < \" for \ \ proof - have le: "\idx \ \ \idx \" if "\ (\idx \ - Suc 0) < h (\idx \) \" "h (\idx \) \ < \ (\idx \)" by (metis 2 that Diff_iff \idx \ \ \\ < \\ dual_order.strict_implies_order dual_order.strict_trans1 h_increasing leI le_\ less_asym) have "h 0 \ < h 0 \" if "\idx \ = 0" "\idx \ = 0" using that mono_h unfolding strict_mono_on_def by (metis Diff_iff \idx \ \ \\ < \\) moreover have "h 0 \ < h (\idx \) \" if "0 < \idx \" "h 0 \ < \ 0" and "\ (\idx \ - Suc 0) < h (\idx \) \" by (meson DiffI \idx \ le_\ le_less_trans less_le_not_le that) moreover have "\idx \ \ 0" if "0 < \idx \" "h 0 \ < \ 0" "\ (\idx \ - Suc 0) < h (\idx \) \" using le le_0_eq that by fastforce moreover have "h (\idx \) \ < h (\idx \) \" if "\ (\idx \ - Suc 0) < h (\idx \) \" "h (\idx \) \ < \ (\idx \)" "h (\idx \) \ < \ (\idx \)" "\ (\idx \ - Suc 0) < h (\idx \) \" using mono_h unfolding strict_mono_on_def by (metis le Diff_iff \idx \ \ \\ < \\ le_\ le_less le_less_trans that) ultimately show ?thesis using h_in_D [OF \] h_in_D [OF \] by (simp add: D_def split: if_split_asm) qed - ultimately have sm_g: "strict_mono_on g (elts \)" + ultimately have sm_g: "strict_mono_on (elts \) g" by (auto simp: g_def strict_mono_on_def dest!: F_imp_Ex) have False if "\ \ elts \" and \: "tp (K 1 x \ \ (g \)) < \" for \ proof - have "F \ M (elts \) \ x" by (meson FN \x \ A'\) then have False if "tp (K (Suc 0) x \ \ \) < \" "\ \ F" using that by (auto simp: M_def) moreover have False if "tp (K (Suc 0) x \ \ (g \)) < \" "\ \ D k" "k \ p" "\ \ F" for k proof - have "h (\idx \) \ \ M (D (\idx \)) \ x" using fun_h \idx \\ \ elts \\ \\ \ F\ by auto then show False using that by (simp add: M_def g_def leD) qed ultimately show False using \\ \ elts \\ \ by (force simp: \_decomp) qed then show False using * [OF \x \ A\ fun_g sm_g] by auto qed then have "\l. l \ p \ tp (M (elts \) \ x \ D l) < tp (D l)" using M_Int_D by auto } then obtain l where lp: "\x. x \ A'\ l x \ p" and lless: "\x. x \ A'\ tp (M (elts \) \ x \ D (l x)) < tp (D (l x))" by metis obtain A'' L where "A'' \ A'" and A'': "A'' \ elts (\*\)" "tp A'' = \" and lL: "\x. x \ A'' \ l x = L" proof - have eq: "A' = (\i\p. {x \ A'. l x = i})" using lp by auto have "\X\(\i. {x \ A'. l x = i}) ` {..p}. \ \ tp X" proof (rule indecomposable_ordertype_finite_ge [OF indec]) show "small (\i\p. {x \ A'. l x = i})" by (metis A'(1) A(1) eq down smaller_than_small) qed (use A' eq in auto) then show thesis proof fix A'' assume A'': "A'' \ (\i. {x \ A'. l x = i}) ` {..p}" and "\ \ tp A''" then obtain L where L: "\x. x \ A'' \ l x = L" by auto have "A'' \ A'" using A'' by force then have "tp A'' \ tp A'" by (meson A' A down order_trans ordertype_VWF_mono) with \\ \ tp A''\ have "tp A'' = \" using A'(2) by auto moreover have "A'' \ elts (\*\)" using A' A \A'' \ A'\ by auto ultimately show thesis using L that [OF \A'' \ A'\] by blast qed qed have \D: "\ \ D L \ {X. X \ elts (\*\) \ tp X = \}" using \ D\ by blast have \: "\ \ tp {x \ A''. tp (D L) \ tp (M (D L) \ x)}" using IX [OF D\ A'' \D] by simp have "M (elts \) \ x \ D L = M (D L) \ x" for x using D\ by (auto simp: M_def) then have "tp (M (D L) \ x) < tp (D L)" if "x \ A''" for x using lless that \A'' \ A'\ lL by force then have [simp]: "{x \ A''. tp (D L) \ tp (M (D L) \ x)} = {}" using leD by blast show False using \ \\ \ \\ by simp qed then show ?thesis by (meson Ord_linear2 Ord_ordertype \Ord \\) qed let ?U = "UNIV :: nat set" define \ where "\ \ fst \ from_nat_into (elts \ \ ?U)" define q where "q \ to_nat_on (elts \ \ ?U)" have co_\U: "countable (elts \ \ ?U)" by (simp add: \ less_\1_imp_countable) moreover have "elts \ \ ?U \ {}" using \0 \ elts \\ by blast ultimately have "range (from_nat_into (elts \ \ ?U)) = (elts \ \ ?U)" by (metis range_from_nat_into) then have \_in_\ [simp]: "\ i \ elts \" for i by (metis SigmaE \_def comp_apply fst_conv range_eqI) then have Ord_\ [simp]: "Ord (\ i)" for i using Ord_in_Ord by blast have inf_\U: "infinite (elts \ \ ?U)" using \0 \ elts \\ finite_cartesian_productD2 by auto have 11 [simp]: "\ (q (\,n)) = \" if "\ \ elts \" for \ n by (simp add: \_def q_def that co_\U) have range_\ [simp]: "range \ = elts \" by (auto simp: image_iff) (metis 11) have [simp]: "KI i {} = UNIV" "KI i (insert a X) = K i a \ KI i X" for i a X by (auto simp: KI_def) define \ where "\ \ \n::nat. \\ x. (\\ \ elts \. \ \ \ elts (\*\) \ tp (\ \) = \) \ x ` {.. elts (\*\) \ (\\ \ elts \. \ \) \ KI 1 (x ` {.. strict_mono_sets (elts \) \" - define \ where "\ \ \n::nat. \g \ \' xn. g \ elts \ \ elts \ \ strict_mono_on g (elts \) + define \ where "\ \ \n::nat. \g \ \' xn. g \ elts \ \ elts \ \ strict_mono_on (elts \) g \ (\i\n. g (\ i) = \ i) \ (\\ \ elts \. \' \ \ K 1 xn \ \ (g \)) \ {xn} \ (\' (\ n)) \ xn \ \ (\ n)" let ?\0 = "\\. plus (\ * \) ` elts \" have base: "\ 0 ?\0 x" for x by (auto simp: \_def add_mult_less add_mult_less_add_mult ordertype_image_plus strict_mono_sets_def less_sets_def) have step: "Ex (\(g,\',xn). \ n g \ \' xn \ \ (Suc n) \' (x(n:=xn)))" if "\ n \ x" for n \ x proof - have \: "\\. \ \ elts \ \ \ \ \ elts (\*\) \ tp (\ \) = \" and x: "x ` {.. elts (\*\)" and sub: "\ (\ ` elts \) \ KI (Suc 0) (x ` {..) \" and \\: "\ ` {..n} \ elts \" and \sub: "\ (\ n) \ elts (\*\)" and \fun: "\ \ elts \ \ {X. X \ elts (\*\) \ tp X = \}" using that by (auto simp: \_def) then obtain xn g where xn: "xn \ \ (\ n)" and g: "g \ elts \ \ elts \" - and sm_g: "strict_mono_on g (elts \)" and g_\: "\\ \ \`{..n}. g \ = \" + and sm_g: "strict_mono_on (elts \) g" and g_\: "\\ \ \`{..n}. g \ = \" and g_\: "\\ \ elts \. \ \ tp (K 1 xn \ \ (g \))" using 10 [OF _ \\ \sub _ \fun] by (auto simp: \) have tp1: "tp (K 1 xn \ \ (g \)) = \" if "\ \ elts \" for \ proof (rule antisym) have "tp (K 1 xn \ \ (g \)) \ tp (\ (g \))" by (metis Int_lower2 PiE \ down g ordertype_VWF_mono that) also have "\ = \" using \ g that by force finally show "tp (K 1 xn \ \ (g \)) \ \" . qed (use that g_\ in auto) have tp2: "tp (\ (\ n)) = \" by (auto simp: \) obtain "small (\ (\ n))" "\ (\ n) \ ON" by (meson \sub ord down elts_subset_ON subset_trans) then obtain A2 where A2: "tp A2 = \" "A2 \ K 1 xn \ \ (\ n)" "{xn} \ A2" using indecomposable_imp_Ex_less_sets [OF indec \\ \ \\ tp2] by (metis \_in_\ atMost_iff image_eqI inf_le2 le_refl xn tp1 g_\) then have A2_sub: "A2 \ \ (\ n)" by simp let ?\ = "\\. if \ = \ n then A2 else K 1 xn \ \ (g \)" have [simp]: "({.. {x. x \ n}) = ({.. (\x\elts \ \ {\. \ \ \ n}. \ (g x)) \ KI (Suc 0) (x ` {.. KI (Suc 0) (x ` {.. elts (\*\)" "xn \ elts (\*\)" using \sub sub A2 xn by fastforce+ moreover have "strict_mono_sets (elts \) ?\" using sm sm_g g g_\ A2_sub unfolding strict_mono_sets_def strict_mono_on_def less_sets_def Pi_iff subset_iff Ball_def Bex_def image_iff by (simp (no_asm_use) add: if_split_mem2) (smt order_refl) ultimately have "\ (Suc n) ?\ (x(n := xn))" using tp1 x A2 by (auto simp: \_def K_def) with A2 show ?thesis by (rule_tac x="(g,?\,xn)" in exI) (simp add: \_def g sm_g g_\ xn) qed define G where "G \ \n \ x. @(g,\',x'). \xn. \ n g \ \' xn \ x' = (x(n:=xn)) \ \ (Suc n) \' x'" have G\: "(\(g,\',x'). \ (Suc n) \' x') (G n \ x)" and G\: "(\(g,\',x'). \ n g \ \' (x' n)) (G n \ x)" if "\ n \ x" for n \ x using step [OF that] by (force simp: G_def dest: some_eq_imp)+ define H where "H \ rec_nat (id,?\0,undefined) (\n (g0,\,x0). G n \ x0)" have "(\(g,\,x). \ n \ x) (H n)" for n unfolding H_def by (induction n) (use G\ base in fastforce)+ then have H_imp_\: "\ n \ x" if "H n = (g,\,x)" for g \ x n by (metis case_prodD that) then have H_imp_\: "(\(g,\',x'). let (g0,\,x) = H n in \ n g \ \' (x' n)) (H (Suc n))" for n using G\ by (fastforce simp: H_def split: prod.split) define g where "g \ \n. (\(g,\,x). g) (H (Suc n))" - have g: "g n \ elts \ \ elts \" and sm_g: "strict_mono_on (g n) (elts \)" + have g: "g n \ elts \ \ elts \" and sm_g: "strict_mono_on (elts \) (g n)" and 13: "\i. i\n \ g n (\ i) = \ i" for n using H_imp_\ [of n] by (auto simp: g_def \_def) define \ where "\ \ \n. (\(g,\,x). \) (H n)" define x where "x \ \n. (\(g,\,x). x n) (H (Suc n))" have 14: "\ (Suc n) \ \ K 1 (x n) \ \ n (g n \)" if "\ \ elts \" for \ n using H_imp_\ [of n] that by (force simp: \_def \_def x_def g_def) then have x14: "\ (Suc n) \ \ \ n (g n \)" if "\ \ elts \" for \ n using that by blast have 15: "x n \ \ n (\ n)" and 16: "{x n} \ (\ (Suc n) (\ n))" for n using H_imp_\ [of n] by (force simp: \_def \_def x_def)+ have \_\\: "\ n \ \ elts (\*\)" if "\ \ elts \" for \ n using H_imp_\ [of n] that by (auto simp: \_def \_def split: prod.split) have 12: "strict_mono_sets (elts \) (\ n)" for n using H_imp_\ [of n] that by (auto simp: \_def \_def split: prod.split) have tp_\: "tp (\ n \) = \" if "\ \ elts \" for \ n using H_imp_\ [of n] that by (auto simp: \_def \_def split: prod.split) let ?Z = "range x" have S_dec: "\ (\ (m+k) ` elts \) \ \ (\ m ` elts \)" for k m by (induction k) (use 14 g in \fastforce+\) have "x n \ K 1 (x m)" if "m (\\ \ elts \. \ n \)" by (meson "15" UN_I \_in_\) also have "\ \ (\\ \ elts \. \ (Suc m) \)" using S_dec [of "Suc m"] less_iff_Suc_add that by auto also have "\ \ K 1 (x m)" using 14 by auto finally show ?thesis . qed then have "f{x m, x n} = 1" if "m2\<^esup> \ {1}" by (clarsimp simp: nsets_2_eq) (metis insert_commute less_linear) moreover have Z_sub: "?Z \ elts (\*\)" using "15" \_\\ \_in_\ by blast moreover have "tp ?Z \ \ * \" proof - define \ where "\ \ \i j x. wfrec (measure (\k. j-k)) (\\ k. if k (Suc k)) else x) i" have \: "\ i j x = (if i (Suc i) j x) else x)" for i j x by (simp add: \_def wfrec cut_apply) have 17: "\ k j (\ i) = \ i" if "i \ k" for i j k using wf_measure [of "\k. j-k"] that by (induction k rule: wf_induct_rule) (simp add: "13" \ le_imp_less_Suc) have \_in_\: "\ i j \ \ elts \" if "\ \ elts \" for i j \ using wf_measure [of "\k. j-k"] that proof (induction i rule: wf_induct_rule) case (less i) with g show ?case by (force simp: \ [of i]) qed then have \_fun: "\ i j \ elts \ \ elts \" for i j by simp - have sm_\: "strict_mono_on (\ i j) (elts \)" for i j + have sm_\: "strict_mono_on (elts \) (\ i j)" for i j using wf_measure [of "\k. j-k"] proof (induction i rule: wf_induct_rule) case (less i) with sm_g show ?case by (auto simp: \ [of i] strict_mono_on_def \_in_\) qed have *: "\ j (\ j) \ \ i (\ i j (\ j))" if "i < j" for i j using wf_measure [of "\k. j-k"] that proof (induction i rule: wf_induct_rule) case (less i) then have "j - Suc i < j - i" by (metis (no_types) Suc_diff_Suc lessI) with less \_in_\ show ?case by (simp add: \ [of i]) (metis 17 Suc_lessI \_in_\ order_refl order_trans x14) qed have le_iff: "\ i j (\ j) \ \ i \ \ j \ \ i" for i j using sm_\ unfolding strict_mono_on_def by (metis "17" Ord_in_Ord Ord_linear2 \_in_\ leD le_refl less_V_def \Ord \\) then have less_iff: "\ i j (\ j) < \ i \ \ j < \ i" for i j by (metis (no_types, lifting) "17" \_in_\ less_V_def order_refl sm_\ strict_mono_on_def) have eq_iff: "\ i j (\ j) = \ i \ \ j = \ i" for i j by (metis eq_refl le_iff less_iff less_le) have \_if_ne: "\ m < \ n" if mn: "\ m (\ m) \ \ n (\ n)" "m \ n" for m n proof - have xmn: "x m < x n" using "15" less_setsD that(1) by blast have Ord\: "Ord (\ n m (\ m))" using Ord_in_Ord \_in_\ \_in_\ ord(2) by presburger have "\ \ m (\ m) \ \ n (\ n)" if "\ n = \ m" using "*" "15" eq_iff that unfolding less_sets_def by (metis in_mono less_irrefl not_less_iff_gr_or_eq) moreover have "\ n (\ n) \ \ m (\ m n (\ n)) \ \ m (\ m) \ \ n (\ n m (\ m))" using * mn by (meson antisym_conv3) then have False if "\ n < \ m" using strict_mono_setsD [OF 12] 15 xmn \_in_\ \_in_\ that by (smt (verit, best) Ord\ Ord_\ Ord_linear2 leD le_iff less_asym less_iff less_setsD subset_iff) ultimately show "\ m < \ n" by (meson that(1) Ord_\ Ord_linear_lt) qed have 18: "\ m (\ m) \ \ n (\ n) \ \ m < \ n" for m n proof (cases n m rule: linorder_cases) case less show ?thesis proof (intro iffI) assume "\ m < \ n" then have "\ n (\ n m (\ m)) \ \ n (\ n)" by (metis "12" \_in_\ \_in_\ eq_iff le_iff less_V_def strict_mono_sets_def) then show "\ m (\ m) \ \ n (\ n)" by (meson "*" less less_sets_weaken1) qed (use \_if_ne less in blast) next case equal with 15 show ?thesis by auto next case greater show ?thesis proof (intro iffI) assume "\ m < \ n" then have "\ m (\ m) \ (\ m (\ m n (\ n)))" by (meson 12 Ord_in_Ord Ord_linear2 \_in_\ \_in_\ le_iff leD ord(2) strict_mono_sets_def) then show "\ m (\ m) \ \ n (\ n)" by (meson "*" greater less_sets_weaken2) qed (use \_if_ne greater in blast) qed have \_increasing_\: "\ n (\ n) \ \ m (\ m)" if "m \ n" "\ m = \ n" for m n by (metis "*" "17" dual_order.order_iff_strict that) moreover have INF: "infinite {n. n \ m \ \ m = \ n}" for m proof - have "infinite (range (\n. q (\ m, n)))" unfolding q_def using to_nat_on_infinite [OF co_\U inf_\U] finite_image_iff by (simp add: finite_image_iff inj_on_def) moreover have "(range (\n. q (\ m, n))) \ {n. \ m = \ n}" using 11 [of "\ m"] by auto ultimately have "infinite {n. \ m = \ n}" using finite_subset by auto then have "infinite ({n. \ m = \ n} - {.. n" "\ p = \ n" "\ m = \ n" "n < p" with 16 [of n] show "x n < x p" by (metis "*" "15" "17" Suc_lessI insert_absorb insert_subset le_SucI less_sets_singleton1) qed then have inj_x: "inj_on x (?eqv m)" for m using strict_mono_on_imp_inj_on by blast define ZA where "ZA \ \m. ?Z \ \ m (\ m)" have small_ZA [simp]: "small (ZA m)" for m by (metis ZA_def inf_le1 small_image_nat smaller_than_small) have 19: "tp (ZA m) \ \" for m proof - have "x ` {n. m \ n \ \ m = \ n} \ ZA m" unfolding ZA_def using "15" \_increasing_\ by blast then have "infinite (ZA m)" using INF [of m] finite_image_iff [OF inj_x] by (meson finite_subset) then show ?thesis by (simp add: ordertype_infinite_ge_\) qed - have "\f \ elts \ \ ZA m. strict_mono_on f (elts \)" for m + have "\f \ elts \ \ ZA m. strict_mono_on (elts \) f" for m proof - obtain Z where "Z \ ZA m" "tp Z = \" by (meson 19 Ord_\ le_ordertype_obtains_subset small_ZA) moreover have "ZA m \ ON" using Ord_in_Ord \_\\ \_in_\ unfolding ZA_def by blast ultimately show ?thesis by (metis strict_mono_on_ordertype Pi_mono small_ZA smaller_than_small subset_iff) qed then obtain \ where \: "\m. \ m \ elts \ \ ZA m" - and sm_\: "\m. strict_mono_on (\ m) (elts \)" + and sm_\: "\m. strict_mono_on (elts \) (\ m)" by metis have "Ex(\(m,\). \ \ elts \ \ \ = \ * \ + ord_of_nat m)" if "\ \ elts (\ * \)" for \ using that by (auto simp: mult [of \ \] lift_def elts_\) then obtain split where split: "\\. \ \ elts (\ * \) \ (\(m,\). \ \ elts \ \ \ = \ * \ + ord_of_nat m)(split \)" by meson have split_eq [simp]: "split (\ * \ + ord_of_nat m) = (m,\)" if "\ \ elts \" for \ m proof - have [simp]: "\ * \ + ord_of_nat m = \ * \ + ord_of_nat n \ \ = \ \ n = m" if "\ \ elts \" for \ n by (metis Ord_\ that Ord_mem_iff_less_TC mult_cancellation_lemma ord_of_nat_\ ord_of_nat_inject) show ?thesis using split [of "\*\ + m"] that by (auto simp: mult [of \ \] lift_def cong: conj_cong) qed define \ where "\ \ \\. (\(m,\). \ (q(\,0)) m)(split \)" have \_Pi: "\ \ elts (\ * \) \ (\m. ZA m)" using \ by (fastforce simp: \_def mult [of \ \] lift_def elts_\) moreover have "(\m. ZA m) \ ON" unfolding ZA_def using \_\\ \_in_\ elts_subset_ON by blast ultimately have Ord_\_Pi: "\ \ elts (\ * \) \ ON" by fastforce show "tp ?Z \ \ * \" proof - have \: "(\m. ZA m) = ?Z" using "15" by (force simp: ZA_def) moreover have "tp (elts (\ * \)) \ tp (\m. ZA m)" proof (rule ordertype_inc_le) show "\ ` elts (\ * \) \ (\m. ZA m)" using \_Pi by blast next fix u v assume x: "u \ elts (\ * \)" and y: "v \ elts (\ * \)" and "(u,v) \ VWF" then have "u Ord_in_Ord Ord_mult VWF_iff_Ord_less ord(2)) moreover obtain m \ n \ where ueq: "u = \ * \ + ord_of_nat m" and \: "\ \ elts \" and veq: "v = \ * \ + ord_of_nat n" and \: "\ \ elts \" using x y by (auto simp: mult [of \ \] lift_def elts_\) ultimately have "\ \ \" by (meson Ord_\ Ord_in_Ord Ord_linear2 \Ord \\ add_mult_less_add_mult less_asym ord_of_nat_\) consider (eq) "\ = \" | (lt) "\ < \" using \\ \ \\ le_neq_trans by blast then have "\ u < \ v" proof cases case eq then have "m < n" using ueq veq \u by simp then have "\ (q (\, 0)) m < \ (q (\, 0)) n" using sm_\ strict_mono_onD by blast then show ?thesis using eq ueq veq \ \m < n\ by (simp add: \_def) next case lt have "\ (q(\,0)) m \ \ (q(\,0)) (\(q(\,0)))" "\ (q (\,0)) n \ \ (q(\,0)) (\(q(\,0)))" using \ unfolding ZA_def by blast+ then show ?thesis using lt ueq veq \ \ 18 [of "q(\,0)" "q(\,0)"] by (simp add: \_def less_sets_def) qed then show "(\ u, \ v) \ VWF" using \_Pi by (metis Ord_\_Pi PiE VWF_iff_Ord_less x y mem_Collect_eq) qed (use \ in auto) ultimately show ?thesis by simp qed qed then obtain Z where "Z \ ?Z" "tp Z = \ * \" by (meson Ord_\ Ord_mult ord Z_sub down le_ordertype_obtains_subset) ultimately show False using iii [of Z] by (meson dual_order.trans image_mono nsets_mono) qed have False if 0: "\H. tp H = ord_of_nat (2*k) \ H \ elts (\*\) \ \ f ` [H]\<^bsup>2\<^esup> \ {0}" and 1: "\H. tp H = min \ (\ * \) \ H \ elts (\*\) \ \ f ` [H]\<^bsup>2\<^esup> \ {1}" proof (cases "\*\ \ \") case True then have \: "\H'\H. tp H' = \ * \" if "tp H = \" "small H" for H by (metis Ord_\ Ord_\1 Ord_in_Ord Ord_mult \ le_ordertype_obtains_subset that) have [simp]: "min \ (\*\) = \*\" by (simp add: min_absorb2 that True) then show ?thesis using * [OF 0] 1 True by simp (meson \ down image_mono nsets_mono subset_trans) next case False then have \: "\H'\H. tp H' = \" if "tp H = \ * \" "small H" for H by (metis Ord_linear_le Ord_ordertype \Ord \\ le_ordertype_obtains_subset that) then have "\ \ \*\" by (meson Ord_\ Ord_\1 Ord_in_Ord Ord_linear_le Ord_mult \ \Ord \\ False) then have [simp]: "min \ (\*\) = \" by (simp add: min_absorb1) then show ?thesis using * [OF 0] 1 False by simp (meson \ down image_mono nsets_mono subset_trans) qed then show "\iH\elts (\*\). tp H = [ord_of_nat (2*k), min \ (\*\)] ! i \ f ` [H]\<^bsup>2\<^esup> \ {i}" by force qed qed theorem Erdos_Milner: assumes \: "\ \ elts \1" shows "partn_lst_VWF (\\(1 + \ * n)) [ord_of_nat (2^n), \\(1+\)] 2" proof (induction n) case 0 then show ?case using partn_lst_VWF_degenerate [of 1 2] by simp next case (Suc n) have "Ord \" using Ord_\1 Ord_in_Ord assms by blast have "1+\ \ \+1" by (simp add: \Ord \\ one_V_def plus_Ord_le) then have [simp]: "min (\ \ (1 + \)) (\ * \ \ \) = \ \ (1+\)" by (simp add: \Ord \\ oexp_add min_def) have ind: "indecomposable (\ \ (1 + \ * ord_of_nat n))" by (simp add: \Ord \\ indecomposable_\_power) show ?case proof (cases "n = 0") case True then show ?thesis using partn_lst_VWF_\_2 \Ord \\ one_V_def by auto next case False then have "Suc 0 < 2 ^ n" using less_2_cases not_less_eq by fastforce then have "partn_lst_VWF (\ \ (1 + \ * n) * \ \ \) [ord_of_nat (2 * 2 ^ n), \ \ (1 + \)] 2" using Erdos_Milner_aux [OF Suc ind, where \ = "\\\"] \Ord \\ \ by (auto simp: countable_oexp) then show ?thesis using \Ord \\ by (simp add: mult_succ mult.assoc oexp_add) qed qed corollary remark_3: "partn_lst_VWF (\\(Suc(4*k))) [4, \\(Suc(2*k))] 2" using Erdos_Milner [of "2*k" 2] apply (simp flip: ord_of_nat_mult ord_of_nat.simps) by (simp add: one_V_def) text \Theorem 3.2 of Jean A. Larson, ibid.\ corollary Theorem_3_2: fixes k n::nat shows "partn_lst_VWF (\\(n*k)) [\\n, ord_of_nat k] 2" proof (cases "n=0 \ k=0") case True then show ?thesis by (auto intro: partn_lst_triv0 [where i=1] partn_lst_triv1 [where i=0]) next case False then have "n > 0" "k > 0" by auto from \k > 0\ less_exp [of \k - 1\] have \k \ 2 ^ (k - 1)\ by (cases k) (simp_all add: less_eq_Suc_le) have PV: "partn_lst_VWF (\ \ (1 + ord_of_nat (n-1) * ord_of_nat (k-1))) [ord_of_nat (2 ^ (k-1)), \ \ (1 + ord_of_nat (n-1))] 2" using Erdos_Milner [of "ord_of_nat (n-1)" "k-1"] Ord_\1 Ord_mem_iff_lt less_imp_le by blast have "k+n \ Suc (k * n)" using False not0_implies_Suc by fastforce then have "1 + (n - 1) * (k - 1) \ n*k" using False by (auto simp: algebra_simps) then have "(1 + ord_of_nat (n - 1) * ord_of_nat (k - 1)) \ ord_of_nat(n*k)" by (metis (mono_tags, lifting) One_nat_def one_V_def ord_of_nat.simps ord_of_nat_add ord_of_nat_mono_iff ord_of_nat_mult) then have x: "\ \ (1 + ord_of_nat (n - 1) * ord_of_nat (k - 1)) \ \\(n*k)" by (simp add: oexp_mono_le) then have "partn_lst_VWF (\\(n*k)) [\ \ (1 + ord_of_nat (n-1)), ord_of_nat (2 ^ (k-1))] 2" by (metis PV partn_lst_two_swap Partitions.partn_lst_greater_resource less_eq_V_def) then have "partn_lst_VWF (\\(n*k)) [\ \ n, ord_of_nat (2 ^ (k-1))] 2" using ord_of_minus_1 [OF \n > 0\] by (simp add: one_V_def) then show ?thesis using \k \ 2 ^ (k - 1)\ by (auto elim!: partn_lst_less simp add: less_Suc_eq) qed end diff --git a/thys/Ordinal_Partitions/Library_Additions.thy b/thys/Ordinal_Partitions/Library_Additions.thy --- a/thys/Ordinal_Partitions/Library_Additions.thy +++ b/thys/Ordinal_Partitions/Library_Additions.thy @@ -1,338 +1,338 @@ section \Library additions\ theory Library_Additions imports "ZFC_in_HOL.Ordinal_Exp" "HOL-Library.Ramsey" "Nash_Williams.Nash_Williams" begin lemma finite_enumerate_Diff_singleton: fixes S :: "'a::wellorder set" assumes "finite S" and i: "i < card S" "enumerate S i < x" shows "enumerate (S - {x}) i = enumerate S i" using i proof (induction i) case 0 have "(LEAST i. i \ S \ i\x) = (LEAST i. i \ S)" proof (rule Least_equality) have "\t. t \ S \ t\x" using 0 \finite S\ finite_enumerate_in_set by blast then show "(LEAST i. i \ S) \ S \ (LEAST i. i \ S) \ x" by (metis "0.prems"(2) LeastI enumerate_0 not_less_Least) qed (simp add: Least_le) then show ?case by (auto simp: enumerate_0) next case (Suc i) then have x: "enumerate S i < x" by (meson enumerate_step finite_enumerate_step less_trans) have cardSx: "Suc i < card (S - {x})" and "i < card S" using Suc \finite S\ card_Diff_singleton_if[of S] finite_enumerate_Ex by fastforce+ have "(LEAST s. s \ S \ s\x \ enumerate (S - {x}) i < s) = (LEAST s. s \ S \ enumerate S i < s)" (is "_ = ?r") proof (intro Least_equality conjI) show "?r \ S" by (metis (lifting) LeastI Suc.prems(1) assms(1) finite_enumerate_in_set finite_enumerate_step) show "?r \ x" using Suc.prems not_less_Least [of _ "\t. t \ S \ enumerate S i < t"] \finite S\ finite_enumerate_in_set finite_enumerate_step by blast show "enumerate (S - {x}) i < ?r" by (metis (full_types) Suc.IH Suc.prems(1) \i < card S\ enumerate_Suc'' enumerate_step finite_enumerate_Suc'' finite_enumerate_step x) show "\y. y \ S \ y \ x \ enumerate (S - {x}) i < y \ ?r \ y" by (simp add: Least_le Suc.IH \i < card S\ x) qed then show ?case using Suc assms by (simp add: finite_enumerate_Suc'' cardSx) qed lemma hd_lex: "\hd ms < hd ns; length ms = length ns; ns \ []\ \ (ms, ns) \ lex less_than" by (metis hd_Cons_tl length_0_conv less_than_iff lexord_cons_cons lexord_lex) lemma sorted_hd_le: assumes "sorted xs" "x \ list.set xs" shows "hd xs \ x" using assms by (induction xs) (auto simp: less_imp_le) lemma sorted_le_last: assumes "sorted xs" "x \ list.set xs" shows "x \ last xs" using assms by (induction xs) (auto simp: less_imp_le) lemma hd_list_of: assumes "finite A" "A \ {}" shows "hd (sorted_list_of_set A) = Min A" proof (rule antisym) have "Min A \ A" by (simp add: assms) then show "hd (sorted_list_of_set A) \ Min A" by (simp add: sorted_hd_le \finite A\) next show "Min A \ hd (sorted_list_of_set A)" by (metis Min_le assms hd_in_set set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) qed lemma sorted_hd_le_last: assumes "sorted xs" "xs \ []" shows "hd xs \ last xs" using assms by (simp add: sorted_hd_le) lemma sorted_list_of_set_set_of [simp]: "strict_sorted l \ sorted_list_of_set (list.set l) = l" by (simp add: strict_sorted_equal) lemma range_strict_mono_ext: fixes f::"nat \ 'a::linorder" assumes eq: "range f = range g" and sm: "strict_mono f" "strict_mono g" shows "f = g" proof fix n show "f n = g n" proof (induction n rule: less_induct) case (less n) obtain x y where xy: "f n = g y" "f x = g n" by (metis eq imageE rangeI) then have "n = y" by (metis (no_types) less.IH neq_iff sm strict_mono_less xy) then show ?case using xy by auto qed qed subsection \Other material\ definition strict_mono_sets :: "['a::order set, 'a::order \ 'b::order set] \ bool" where "strict_mono_sets A f \ \x\A. \y\A. x < y \ less_sets (f x) (f y)" lemma strict_mono_setsD: assumes "strict_mono_sets A f" "x < y" "x \ A" "y \ A" shows "less_sets (f x) (f y)" using assms by (auto simp: strict_mono_sets_def) -lemma strict_mono_on_o: "\strict_mono_on r A; strict_mono_on s B; s ` B \ A\ \ strict_mono_on (r \ s) B" +lemma strict_mono_on_o: "\strict_mono_on A r; strict_mono_on B s; s ` B \ A\ \ strict_mono_on B (r \ s)" by (auto simp: image_subset_iff strict_mono_on_def) lemma strict_mono_sets_imp_disjoint: fixes A :: "'a::linorder set" assumes "strict_mono_sets A f" shows "pairwise (\x y. disjnt (f x) (f y)) A" using assms unfolding strict_mono_sets_def pairwise_def by (meson antisym_conv3 disjnt_sym less_sets_imp_disjnt) lemma strict_mono_sets_subset: assumes "strict_mono_sets B f" "A \ B" shows "strict_mono_sets A f" using assms by (auto simp: strict_mono_sets_def) lemma strict_mono_less_sets_Min: assumes "strict_mono_sets I f" "finite I" "I \ {}" shows "less_sets (f (Min I)) (\ (f ` (I - {Min I})))" using assms by (simp add: strict_mono_sets_def less_sets_UN2 dual_order.strict_iff_order) lemma pair_less_iff1 [simp]: "((x,y), (x,z)) \ pair_less \ y" "\\{}" "\A. A \ \ \ infinite A" and "\A B. \A \ \; B \ \\ \ A \ B \ \" shows "infinite (\\)" by (simp add: assms finite_Inf_in) lemma atLeast_less_sets: "\less_sets A {x}; B \ {x..}\ \ less_sets A B" by (force simp: less_sets_def subset_iff) subsection \The list-of function\ lemma sorted_list_of_set_insert_remove_cons: assumes "finite A" "less_sets {a} A" shows "sorted_list_of_set (insert a A) = a # sorted_list_of_set A" proof - have "strict_sorted (a # sorted_list_of_set A)" using assms less_setsD by auto moreover have "list.set (a # sorted_list_of_set A) = insert a A" using assms by force moreover have "length (a # sorted_list_of_set A) = card (insert a A)" using assms card_insert_if less_setsD by fastforce ultimately show ?thesis by (metis \finite A\ finite_insert sorted_list_of_set_unique) qed lemma sorted_list_of_set_Un: assumes AB: "less_sets A B" and fin: "finite A" "finite B" shows "sorted_list_of_set (A \ B) = sorted_list_of_set A @ sorted_list_of_set B" proof - have "strict_sorted (sorted_list_of_set A @ sorted_list_of_set B)" using AB unfolding less_sets_def by (metis fin set_sorted_list_of_set sorted_wrt_append strict_sorted_list_of_set) moreover have "card A + card B = card (A \ B)" using less_sets_imp_disjnt [OF AB] by (simp add: assms card_Un_disjoint disjnt_def) ultimately show ?thesis by (simp add: assms strict_sorted_equal) qed lemma sorted_list_of_set_UN_lessThan: fixes k::nat assumes sm: "strict_mono_sets {..i. i < k \ finite (A i)" shows "sorted_list_of_set (\i A) (sorted_list_of_set {.. (A ` {.. (A ` {.. (A ` {.. A k)" by (simp add: Un_commute lessThan_Suc) also have "\ = sorted_list_of_set (\ (A ` {.. = concat (map (sorted_list_of_set \ A) (sorted_list_of_set {.. = concat (map (sorted_list_of_set \ A) (sorted_list_of_set {..i. i \ k \ finite (A i)" shows "sorted_list_of_set (\i\k. A i) = concat (map (sorted_list_of_set \ A) (sorted_list_of_set {..k}))" by (metis assms lessThan_Suc_atMost less_Suc_eq_le sorted_list_of_set_UN_lessThan) subsection \Monotonic enumeration of a countably infinite set\ abbreviation "enum \ enumerate" text \Could be generalised to infinite countable sets of any type\ lemma nat_infinite_iff: fixes N :: "nat set" shows "infinite N \ (\f::nat\nat. N = range f \ strict_mono f)" proof safe assume "infinite N" then show "\f. N = range (f::nat \ nat) \ strict_mono f" by (metis bij_betw_imp_surj_on bij_enumerate enumerate_mono strict_mono_def) next fix f :: "nat \ nat" assume "strict_mono f" and "N = range f" and "finite (range f)" then show False using range_inj_infinite strict_mono_imp_inj_on by blast qed lemma enum_works: fixes N :: "nat set" assumes "infinite N" shows "N = range (enum N) \ strict_mono (enum N)" by (metis assms bij_betw_imp_surj_on bij_enumerate enumerate_mono strict_monoI) lemma range_enum: "range (enum N) = N" and strict_mono_enum: "strict_mono (enum N)" if "infinite N" for N :: "nat set" using enum_works [OF that] by auto lemma enum_0_eq_Inf: fixes N :: "nat set" assumes "infinite N" shows "enum N 0 = Inf N" proof - have "enum N 0 \ N" using assms range_enum by auto moreover have "\x. x \ N \ enum N 0 \ x" by (metis (mono_tags, opaque_lifting) assms imageE le0 less_mono_imp_le_mono range_enum strict_monoD strict_mono_enum) ultimately show ?thesis by (metis cInf_eq_minimum) qed lemma enum_works_finite: fixes N :: "nat set" assumes "finite N" - shows "N = enum N ` {.. strict_mono_on (enum N) {.. strict_mono_on {.. N" "finite N" obtains i where "i < card N" "x = enum N i" by (metis \x \ N\ \finite N\ enum_works_finite imageE lessThan_iff) lemma enum_0_eq_Inf_finite: fixes N :: "nat set" assumes "finite N" "N \ {}" shows "enum N 0 = Inf N" proof - have "enum N 0 \ N" by (metis Nat.neq0_conv assms empty_is_image enum_works_finite image_eqI lessThan_empty_iff lessThan_iff) moreover have "enum N 0 \ x" if "x \ N" for x proof - obtain i where "i < card N" "x = enum N i" by (metis \x \ N\ \finite N\ enum_obtain_index_finite) with assms show ?thesis by (metis Nat.neq0_conv finite_enumerate_mono less_or_eq_imp_le) qed ultimately show ?thesis by (metis cInf_eq_minimum) qed lemma greaterThan_less_enum: fixes N :: "nat set" assumes "N \ {x<..}" "infinite N" shows "x < enum N i" using assms range_enum by fastforce lemma atLeast_le_enum: fixes N :: "nat set" assumes "N \ {x..}" "infinite N" shows "x \ enum N i" using assms range_enum by fastforce lemma less_sets_empty1 [simp]: "less_sets {} A" and less_sets_empty2 [simp]: "less_sets A {}" by (simp_all add: less_sets_def) lemma less_sets_singleton1 [simp]: "less_sets {a} A \ (\x\A. a < x)" and less_sets_singleton2 [simp]: "less_sets A {a} \ (\x\A. x < a)" by (simp_all add: less_sets_def) lemma less_sets_atMost [simp]: "less_sets {..a} A \ (\x\A. a < x)" and less_sets_alLeast [simp]: "less_sets A {a..} \ (\x\A. x < a)" by (auto simp: less_sets_def) lemma less_sets_imp_strict_mono_sets: assumes "\i. less_sets (A i) (A (Suc i))" "\i. i>0 \ A i \ {}" shows "strict_mono_sets UNIV A" proof (clarsimp simp: strict_mono_sets_def) fix i j::nat assume "i < j" then show "less_sets (A i) (A j)" proof (induction "j-i" arbitrary: i j) case (Suc x) then show ?case by (metis Suc_diff_Suc Suc_inject Suc_mono assms less_Suc_eq less_sets_trans zero_less_Suc) qed auto qed lemma less_sets_Suc_Max: assumes "finite A" shows "less_sets A {Suc (Max A)..}" proof (cases "A = {}") case False then show ?thesis by (simp add: assms less_Suc_eq_le) qed auto lemma infinite_nat_greaterThan: fixes m::nat assumes "infinite N" shows "infinite (N \ {m<..})" proof - have "N \ -{m<..} \ (N \ {m<..})" by blast moreover have "finite (-{m<..})" by simp ultimately show ?thesis using assms finite_subset by blast qed end diff --git a/thys/Ordinal_Partitions/Omega_Omega.thy b/thys/Ordinal_Partitions/Omega_Omega.thy --- a/thys/Ordinal_Partitions/Omega_Omega.thy +++ b/thys/Ordinal_Partitions/Omega_Omega.thy @@ -1,4278 +1,4278 @@ section \An ordinal partition theorem by Jean A. Larson\ text \Jean A. Larson, A short proof of a partition theorem for the ordinal $\omega^\omega$. \emph{Annals of Mathematical Logic}, 6:129--145, 1973.\ theory Omega_Omega imports "HOL-Library.Product_Lexorder" Erdos_Milner begin abbreviation "list_of \ sorted_list_of_set" subsection \Cantor normal form for ordinals below @{term "\\\"}\ text \Unlike @{term Cantor_sum}, there is no list of ordinal exponents, which are instead taken as consecutive. We obtain an order-isomorphism between @{term "\\\"} and increasing lists of natural numbers (ordered lexicographically).\ fun omega_sum_aux where Nil: "omega_sum_aux 0 _ = 0" | Suc: "omega_sum_aux (Suc n) [] = 0" | Cons: "omega_sum_aux (Suc n) (m#ms) = (\\n) * (ord_of_nat m) + omega_sum_aux n ms" abbreviation omega_sum where "omega_sum ms \ omega_sum_aux (length ms) ms" text \A normal expansion has no leading zeroes\ inductive normal:: "nat list \ bool" where normal_Nil[iff]: "normal []" | normal_Cons: "m > 0 \ normal (m#ms)" inductive_simps normal_Cons_iff [simp]: "normal (m#ms)" lemma omega_sum_0_iff [simp]: "normal ns \ omega_sum ns = 0 \ ns = []" by (induction ns rule: normal.induct) auto lemma Ord_omega_sum_aux [simp]: "Ord (omega_sum_aux k ms)" by (induction rule: omega_sum_aux.induct) auto lemma Ord_omega_sum: "Ord (omega_sum ms)" by simp lemma omega_sum_less_\\ [intro]: "omega_sum ms < \\\" proof (induction ms) case (Cons m ms) have "\ \ (length ms) * ord_of_nat m \ elts (\ \ Suc (length ms))" using Ord_mem_iff_lt by auto then have "\\(length ms) * ord_of_nat m \ elts (\\\)" using Ord_ord_of_nat oexp_mono_le omega_nonzero ord_of_nat_le_omega by blast with Cons show ?case by (auto simp: mult_succ OrdmemD oexp_less indecomposableD indecomposable_\_power) qed (auto simp: zero_less_Limit) lemma omega_sum_aux_less: "omega_sum_aux k ms < \ \ k" proof (induction rule: omega_sum_aux.induct) case (3 n m ms) have " \\n * ord_of_nat m + \\n < \\n * \" by (metis Ord_ord_of_nat \_power_succ_gtr mult_succ oexp_succ ord_of_nat.simps(2)) with 3 show ?case using dual_order.strict_trans by force qed auto lemma omega_sum_less: "omega_sum ms < \ \ (length ms)" by (rule omega_sum_aux_less) lemma omega_sum_ge: "m \ 0 \ \ \ (length ms) \ omega_sum (m#ms)" apply clarsimp by (metis Ord_ord_of_nat add_le_cancel_left0 le_mult Nat.neq0_conv ord_of_eq_0_iff vsubsetD) lemma omega_sum_length_less: assumes "normal ns" "length ms < length ns" shows "omega_sum ms < omega_sum ns" using assms proof (induction rule: normal.induct) case (normal_Cons n ns') have "\ \ length ms \ \ \ length ns'" using normal_Cons oexp_mono_le by auto then show ?case by (metis gr_implies_not_zero less_le_trans normal_Cons.hyps omega_sum_aux_less omega_sum_ge) qed auto lemma omega_sum_length_leD: assumes "omega_sum ms \ omega_sum ns" "normal ms" shows "length ms \ length ns" by (meson assms leD leI omega_sum_length_less) lemma omega_sum_less_eqlen_iff_cases [simp]: assumes "length ms = length ns" shows "omega_sum (m#ms) < omega_sum (n#ns) \ m m=n \ omega_sum ms < omega_sum ns" (is "?lhs = ?rhs") proof assume L: ?lhs have "\ Suc n < Suc m" using omega_sum_less [of ms] omega_sum_less [of ns] L assms mult_nat_less_add_less by fastforce with L assms show ?rhs by auto qed (auto simp: mult_nat_less_add_less omega_sum_aux_less assms) lemma omega_sum_lex_less_iff_cases: "((length ms, omega_sum (m#ms)), (length ns, omega_sum (n#ns))) \ less_than <*lex*> VWF \ length ms < length ns \ length ms = length ns \ m m=n \ ((length ms, omega_sum ms), (length ns, omega_sum ns)) \ less_than <*lex*> VWF" using omega_sum_less_eqlen_iff_cases by force lemma omega_sum_less_iff_cases: assumes "m > 0" "n > 0" shows "omega_sum (m#ms) < omega_sum (n#ns) \ length ms < length ns \ length ms = length ns \ m length ms = length ns \ m=n \ omega_sum ms < omega_sum ns" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis Suc_less_eq \m>0\ length_Cons less_asym nat_neq_iff normal_Cons omega_sum_length_less omega_sum_less_eqlen_iff_cases) next assume ?rhs then show ?lhs by (metis (full_types) Suc_less_eq \n>0\ length_Cons normal_Cons omega_sum_length_less omega_sum_less_eqlen_iff_cases) qed lemma omega_sum_less_iff: "((length ms, omega_sum ms), (length ns, omega_sum ns)) \ less_than <*lex*> VWF \ (ms,ns) \ lenlex less_than" proof (induction ms arbitrary: ns) case (Cons m ms) then show ?case proof (induction ns) case (Cons n ns') show ?case using Cons.prems Cons_lenlex_iff omega_sum_less_eqlen_iff_cases by fastforce qed auto qed auto lemma eq_omega_sum_less_iff: assumes "length ms = length ns" shows "(omega_sum ms, omega_sum ns) \ VWF \ (ms,ns) \ lenlex less_than" by (metis assms in_lex_prod less_not_refl less_than_iff omega_sum_less_iff) lemma eq_omega_sum_eq_iff: assumes "length ms = length ns" shows "omega_sum ms = omega_sum ns \ ms=ns" proof assume "omega_sum ms = omega_sum ns" then have "(omega_sum ms, omega_sum ns) \ VWF" "(omega_sum ns, omega_sum ms) \ VWF" by auto then obtain "(ms,ns) \ lenlex less_than" "(ns,ms) \ lenlex less_than" using assms eq_omega_sum_less_iff by metis moreover have "total (lenlex less_than)" by (simp add: total_lenlex total_less_than) ultimately show "ms=ns" by (meson UNIV_I total_on_def) qed auto lemma inj_omega_sum: "inj_on omega_sum {l. length l = n}" unfolding inj_on_def using eq_omega_sum_eq_iff by fastforce lemma Ex_omega_sum: "\ \ elts (\\n) \ \ns. \ = omega_sum ns \ length ns = n" proof (induction n arbitrary: \) case 0 then show ?case by (rule_tac x="[]" in exI) auto next case (Suc n) then obtain k::nat where k: "\ \ elts (\ \ n * k)" and kmin: "\k'. k' \ \ elts (\ \ n * k')" by (metis Ord_ord_of_nat elts_mult_\E oexp_succ ord_of_nat.simps(2)) show ?case proof (cases k) case (Suc k') then obtain \ where \: "\ = (\ \ n * k') + \" by (metis lessI mult_succ ord_of_nat.simps(2) k kmin mem_plus_V_E) then have \in: "\ \ elts (\ \ n)" using Suc k mult_succ by auto then obtain ns where ns: "\ = omega_sum ns" and len: "length ns = n" using Suc.IH by auto moreover have "omega_sum ns < \\n" using OrdmemD ns \in by auto ultimately show ?thesis by (rule_tac x="k'#ns" in exI) (simp add: \) qed (use k in auto) qed lemma omega_sum_drop [simp]: "omega_sum (dropWhile (\n. n=0) ns) = omega_sum ns" by (induction ns) auto lemma normal_drop [simp]: "normal (dropWhile (\n. n=0) ns)" by (induction ns) auto lemma omega_sum_\\: assumes "\ \ elts (\\\)" obtains ns where "\ = omega_sum ns" "normal ns" proof - obtain ms where "\ = omega_sum ms" using assms Ex_omega_sum by (auto simp: oexp_Limit elts_\) show thesis proof show "\ = omega_sum (dropWhile (\n. n=0) ms)" by (simp add: \\ = omega_sum ms\) show "normal (dropWhile (\n. n=0) ms)" by auto qed qed definition Cantor_\\ :: "V \ nat list" where "Cantor_\\ \ \x. SOME ns. x = omega_sum ns \ normal ns" lemma assumes "\ \ elts (\\\)" shows Cantor_\\: "omega_sum (Cantor_\\ \) = \" and normal_Cantor_\\: "normal (Cantor_\\ \)" by (metis (mono_tags, lifting) Cantor_\\_def assms omega_sum_\\ someI)+ subsection \Larson's set $W(n)$\ definition WW :: "nat list set" where "WW \ {l. strict_sorted l}" fun into_WW :: "nat \ nat list \ nat list" where "into_WW k [] = []" | "into_WW k (n#ns) = (k+n) # into_WW (Suc (k+n)) ns" fun from_WW :: "nat \ nat list \ nat list" where "from_WW k [] = []" | "from_WW k (n#ns) = (n - k) # from_WW (Suc n) ns" lemma from_into_WW [simp]: "from_WW k (into_WW k ns) = ns" by (induction ns arbitrary: k) auto lemma inj_into_WW: "inj (into_WW k)" by (metis from_into_WW injI) lemma into_from_WW_aux: "\strict_sorted ns; \n\list.set ns. k \ n\ \ into_WW k (from_WW k ns) = ns" by (induction ns arbitrary: k) (auto simp: Suc_leI) lemma into_from_WW [simp]: "strict_sorted ns \ into_WW 0 (from_WW 0 ns) = ns" by (simp add: into_from_WW_aux) lemma into_WW_imp_ge: "y \ List.set (into_WW x ns) \ x \ y" by (induction ns arbitrary: x) fastforce+ lemma strict_sorted_into_WW: "strict_sorted (into_WW x ns)" by (induction ns arbitrary: x) (auto simp: dest: into_WW_imp_ge) lemma length_into_WW: "length (into_WW x ns) = length ns" by (induction ns arbitrary: x) auto lemma WW_eq_range_into: "WW = range (into_WW 0)" proof - have "\ns. strict_sorted ns \ ns \ range (into_WW 0)" by (metis into_from_WW rangeI) then show ?thesis by (auto simp: WW_def strict_sorted_into_WW) qed lemma into_WW_lenlex_iff: "(into_WW k ms, into_WW k ns) \ lenlex less_than \ (ms, ns) \ lenlex less_than" proof (induction ms arbitrary: ns k) case Nil then show ?case by simp (metis length_0_conv length_into_WW) next case (Cons m ms) then show ?case by (induction ns) (auto simp: Cons_lenlex_iff length_into_WW) qed lemma wf_llt [simp]: "wf (lenlex less_than)" and trans_llt [simp]: "trans (lenlex less_than)" by blast+ lemma total_llt [simp]: "total_on A (lenlex less_than)" by (meson UNIV_I total_lenlex total_less_than total_on_def) lemma omega_sum_1_less: assumes "(ms,ns) \ lenlex less_than" shows "omega_sum (1#ms) < omega_sum (1#ns)" proof - have "omega_sum (1#ms) < omega_sum (1#ns)" if "length ms < length ns" using omega_sum_less_iff_cases that zero_less_one by blast then show ?thesis using assms by (auto simp: mult_succ simp flip: omega_sum_less_iff) qed lemma ordertype_WW_1: "ordertype WW (lenlex less_than) \ ordertype UNIV (lenlex less_than)" by (rule ordertype_mono) auto lemma ordertype_WW_2: "ordertype UNIV (lenlex less_than) \ \\\" proof (rule ordertype_inc_le_Ord) show "range (\ms. omega_sum (1#ms)) \ elts (\\\)" by (meson Ord_\ Ord_mem_iff_lt Ord_oexp Ord_omega_sum image_subset_iff omega_sum_less_\\) qed (use omega_sum_1_less in auto) lemma ordertype_WW_3: "\\\ \ ordertype WW (lenlex less_than)" proof - define \ where "\ \ into_WW 0 \ Cantor_\\" have \\: "\\\ = tp (elts (\\\))" by simp also have "\ \ ordertype WW (lenlex less_than)" proof (rule ordertype_inc_le) fix \ \ assume \: "\ \ elts (\\\)" and \: "\ \ elts (\\\)" and "(\, \) \ VWF" then obtain *: "Ord \" "Ord \" "\<\" by (metis Ord_in_Ord Ord_ordertype VWF_iff_Ord_less \\) then have "length (Cantor_\\ \) \ length (Cantor_\\ \)" using \ \ by (simp add: Cantor_\\ normal_Cantor_\\ omega_sum_length_leD) with \ \ * have "(Cantor_\\ \, Cantor_\\ \) \ lenlex less_than" by (auto simp: Cantor_\\ simp flip: omega_sum_less_iff) then show "(\ \, \ \) \ lenlex less_than" by (simp add: \_def into_WW_lenlex_iff) qed (auto simp: \_def WW_def strict_sorted_into_WW) finally show "\\\ \ ordertype WW (lenlex less_than)" . qed lemma ordertype_WW: "ordertype WW (lenlex less_than) = \\\" and ordertype_UNIV_\\: "ordertype UNIV (lenlex less_than) = \\\" using ordertype_WW_1 ordertype_WW_2 ordertype_WW_3 by auto lemma ordertype_\\: fixes F :: "nat \ nat list set" assumes "\j::nat. ordertype (F j) (lenlex less_than) = \\j" shows "ordertype (\j. F j) (lenlex less_than) = \\\" proof (rule antisym) show "ordertype (\ (range F)) (lenlex less_than) \ \ \ \" by (metis ordertype_UNIV_\\ ordertype_mono small top_greatest trans_llt wf_llt) have "\n. \ \ ord_of_nat n \ ordertype (\ (range F)) (lenlex less_than)" by (metis TC_small Union_upper assms ordertype_mono rangeI trans_llt wf_llt) then show "\ \ \ \ ordertype (\ (range F)) (lenlex less_than)" by (auto simp: oexp_\_Limit ZFC_in_HOL.SUP_le_iff elts_\) qed definition WW_seg :: "nat \ nat list set" where "WW_seg n \ {l \ WW. length l = n}" lemma WW_seg_subset_WW: "WW_seg n \ WW" by (auto simp: WW_seg_def) lemma WW_eq_UN_WW_seg: "WW = (\ n. WW_seg n)" by (auto simp: WW_seg_def) lemma ordertype_list_seg: "ordertype {l. length l = n} (lenlex less_than) = \\n" proof - have "bij_betw omega_sum {l. length l = n} (elts (\\n))" unfolding WW_seg_def bij_betw_def by (auto simp: inj_omega_sum Ord_mem_iff_lt omega_sum_less dest: Ex_omega_sum) then show ?thesis by (force simp: ordertype_eq_iff simp flip: eq_omega_sum_less_iff) qed lemma ordertype_WW_seg: "ordertype (WW_seg n) (lenlex less_than) = \\n" (is "ordertype ?W ?R = \\n") proof - have "ordertype {l. length l = n} ?R = ordertype ?W ?R" proof (subst ordertype_eq_ordertype) show "\f. bij_betw f {l. length l = n} ?W \ (\x\{l. length l = n}. \y\{l. length l = n}. ((f x, f y) \ lenlex less_than) = ((x, y) \ lenlex less_than))" proof (intro exI conjI) have "inj_on (into_WW 0) {l. length l = n}" by (metis from_into_WW inj_onI) then show "bij_betw (into_WW 0) {l. length l = n} ?W" by (auto simp: bij_betw_def WW_seg_def WW_eq_range_into length_into_WW) qed (simp add: into_WW_lenlex_iff) qed auto then show ?thesis using ordertype_list_seg by auto qed subsection \Definitions required for the lemmas\ subsubsection \Larson's "$<$"-relation on ordered lists\ instantiation list :: (ord)ord begin definition "xs < ys \ xs \ [] \ ys \ [] \ last xs < hd ys" for xs ys :: "'a list" definition "xs \ ys \ xs < ys \ xs = ys" for xs ys :: "'a list" instance by standard end lemma less_Nil [simp]: "xs < []" "[] < xs" by (auto simp: less_list_def) lemma less_sets_imp_list_less: assumes "list.set xs \ list.set ys" shows "xs < ys" by (metis assms last_in_set less_list_def less_sets_def list.set_sel(1)) lemma less_sets_imp_sorted_list_of_set: assumes "A \ B" "finite A" "finite B" shows "list_of A < list_of B" by (simp add: assms less_sets_imp_list_less) lemma sorted_list_of_set_imp_less_sets: assumes "xs < ys" "sorted xs" "sorted ys" shows "list.set xs \ list.set ys" using assms sorted_hd_le sorted_le_last by (force simp: less_list_def less_sets_def intro: order.trans) lemma less_list_iff_less_sets: assumes "sorted xs" "sorted ys" shows "xs < ys \ list.set xs \ list.set ys" using assms sorted_hd_le sorted_le_last by (force simp: less_list_def less_sets_def intro: order.trans) lemma strict_sorted_append_iff: "strict_sorted (xs @ ys) \ xs < ys \ strict_sorted xs \ strict_sorted ys" (is "?lhs = ?rhs") by (metis less_list_iff_less_sets less_setsD sorted_wrt_append strict_sorted_imp_less_sets strict_sorted_imp_sorted) lemma singleton_less_list_iff: "sorted xs \ [n] < xs \ {..n} \ list.set xs = {}" apply (simp add: less_list_def disjoint_iff) by (metis empty_iff less_le_trans list.set(1) list.set_sel(1) not_le sorted_hd_le) lemma less_hd_imp_less: "xs < [hd ys] \ xs < ys" by (simp add: less_list_def) lemma strict_sorted_concat_I: assumes "\x. x \ list.set xs \ strict_sorted x" "\n. Suc n < length xs \ xs!n < xs!Suc n" "xs \ lists (- {[]})" shows "strict_sorted (concat xs)" using assms proof (induction xs) case (Cons x xs) then have "x < concat xs" apply (simp add: less_list_def) by (metis Compl_iff hd_concat insertI1 length_greater_0_conv length_pos_if_in_set list.sel(1) lists.cases nth_Cons_0) with Cons show ?case by (force simp: strict_sorted_append_iff) qed auto subsection \Nash Williams for lists\ subsubsection \Thin sets of lists\ inductive initial_segment :: "'a list \ 'a list \ bool" where "initial_segment xs (xs@ys)" definition thin :: "'a list set \ bool" where "thin A \ \ (\x y. x \ A \ y \ A \ x \ y \ initial_segment x y)" lemma initial_segment_ne: assumes "initial_segment xs ys" "xs \ []" shows "ys \ [] \ hd ys = hd xs" using assms by (auto elim!: initial_segment.cases) lemma take_initial_segment: assumes "initial_segment xs ys" "k \ length xs" shows "take k xs = take k ys" by (metis append_eq_conv_conj assms initial_segment.cases min_def take_take) lemma initial_segment_length_eq: assumes "initial_segment xs ys" "length xs = length ys" shows "xs = ys" using assms initial_segment.cases by fastforce lemma initial_segment_Nil [simp]: "initial_segment [] ys" by (simp add: initial_segment.simps) lemma initial_segment_Cons [simp]: "initial_segment (x#xs) (y#ys) \ x=y \ initial_segment xs ys" by (metis append_Cons initial_segment.simps list.inject) lemma init_segment_iff_initial_segment: assumes "strict_sorted xs" "strict_sorted ys" shows "init_segment (list.set xs) (list.set ys) \ initial_segment xs ys" (is "?lhs = ?rhs") proof assume ?lhs then obtain S' where S': "list.set ys = list.set xs \ S'" "list.set xs \ S'" by (auto simp: init_segment_def) then have "finite S'" by (metis List.finite_set finite_Un) have "ys = xs @ list_of S'" using S' \strict_sorted xs\ proof (induction xs) case Nil with \strict_sorted ys\ show ?case by auto next case (Cons a xs) with \finite S'\ have "ys = a # xs @ list_of S'" by (metis List.finite_set append_Cons assms(2) sorted_list_of_set_Un sorted_list_of_set_set_of) then show ?case by (auto simp: Cons) qed then show ?rhs using initial_segment.intros by blast next assume ?rhs then show ?lhs proof cases case (1 ys) with assms(2) show ?thesis by (metis init_segment_def set_append sorted_list_of_set_imp_less_sets strict_sorted_append_iff strict_sorted_imp_sorted) qed qed theorem Nash_Williams_WW: fixes h :: "nat list \ nat" assumes "infinite M" and h: "h ` {l \ A. List.set l \ M} \ {..<2}" and "thin A" "A \ WW" obtains i N where "i < 2" "infinite N" "N \ M" "h ` {l \ A. List.set l \ N} \ {i}" proof - define AM where "AM \ {l \ A. List.set l \ M}" have "thin_set (list.set ` A)" using \thin A\ \A \ WW\ unfolding thin_def thin_set_def WW_def by (auto simp: subset_iff init_segment_iff_initial_segment) then have "thin_set (list.set ` AM)" by (simp add: AM_def image_subset_iff thin_set_def) then have "Ramsey (list.set ` AM) 2" using Nash_Williams_2 by metis moreover have "(h \ list_of) \ list.set ` AM \ {..<2}" unfolding AM_def proof clarsimp fix l assume "l \ A" "list.set l \ M" then have "strict_sorted l" using WW_def \A \ WW\ by blast then show "h (list_of (list.set l)) < 2" using h \l \ A\ \list.set l \ M\ by auto qed ultimately obtain N i where N: "N \ M" "infinite N" "i<2" and "list.set ` AM \ Pow N \ (h \ list_of) -` {i}" unfolding Ramsey_eq by (metis \infinite M\) then have N_disjoint: "(h \ list_of) -` {1-i} \ (list.set ` AM) \ Pow N = {}" unfolding subset_vimage_iff less_2_cases_iff by force have "h ` {l \ A. list.set l \ N} \ {i}" proof clarify fix l assume "l \ A" and "list.set l \ N" then have "h l < 2" using h \N \ M\ by force with \i<2\ have "h l \ Suc 0 - i \ h l = i" by (auto simp: eval_nat_numeral less_Suc_eq) moreover have "strict_sorted l" using \A \ WW\ \l \ A\ unfolding WW_def by blast moreover have "h (list_of (list.set l)) = 1 - i \ \ (list.set l \ N)" using N_disjoint \N \ M\ \l \ A\ by (auto simp: AM_def) ultimately show "h l = i" using N \N \ M\ \l \ A\ \list.set l \ N\ by (auto simp: vimage_def set_eq_iff AM_def WW_def subset_iff) qed then show thesis using that \i<2\ N by auto qed subsection \Specialised functions on lists\ lemma mem_lists_non_Nil: "xss \ lists (- {[]}) \ (\x \ list.set xss. x \ [])" by auto fun acc_lengths :: "nat \ 'a list list \ nat list" where "acc_lengths acc [] = []" | "acc_lengths acc (l#ls) = (acc + length l) # acc_lengths (acc + length l) ls" lemma length_acc_lengths [simp]: "length (acc_lengths acc ls) = length ls" by (induction ls arbitrary: acc) auto lemma acc_lengths_eq_Nil_iff [simp]: "acc_lengths acc ls = [] \ ls = []" by (metis length_0_conv length_acc_lengths) lemma set_acc_lengths: assumes "ls \ lists (- {[]})" shows "list.set (acc_lengths acc ls) \ {acc<..}" using assms by (induction ls rule: acc_lengths.induct) fastforce+ text \Useful because @{text acc_lengths.simps} will sometimes be deleted from the simpset.\ lemma hd_acc_lengths [simp]: "hd (acc_lengths acc (l#ls)) = acc + length l" by simp lemma last_acc_lengths [simp]: "ls \ [] \ last (acc_lengths acc ls) = acc + sum_list (map length ls)" by (induction acc ls rule: acc_lengths.induct) auto lemma nth_acc_lengths [simp]: "\ls \ []; k < length ls\ \ acc_lengths acc ls ! k = acc + sum_list (map length (take (Suc k) ls))" by (induction acc ls arbitrary: k rule: acc_lengths.induct) (fastforce simp: less_Suc_eq nth_Cons')+ lemma acc_lengths_plus: "acc_lengths (m+n) as = map ((+)m) (acc_lengths n as)" by (induction n as arbitrary: m rule: acc_lengths.induct) (auto simp: add.assoc) lemma acc_lengths_shift: "NO_MATCH 0 acc \ acc_lengths acc as = map ((+)acc) (acc_lengths 0 as)" by (metis acc_lengths_plus add.comm_neutral) lemma length_concat_acc_lengths: "ls \ [] \ k + length (concat ls) \ list.set (acc_lengths k ls)" by (metis acc_lengths_eq_Nil_iff last_acc_lengths last_in_set length_concat) lemma strict_sorted_acc_lengths: assumes "ls \ lists (- {[]})" shows "strict_sorted (acc_lengths acc ls)" using assms proof (induction ls rule: acc_lengths.induct) case (2 acc l ls) then have "strict_sorted (acc_lengths (acc + length l) ls)" by auto then show ?case using set_acc_lengths "2.prems" by auto qed auto lemma acc_lengths_append: "acc_lengths acc (xs @ ys) = acc_lengths acc xs @ acc_lengths (acc + sum_list (map length xs)) ys" by (induction acc xs rule: acc_lengths.induct) (auto simp: add.assoc) lemma length_concat_ge: assumes "as \ lists (- {[]})" shows "length (concat as) \ length as" using assms proof (induction as) case (Cons a as) then have "length a \ Suc 0" "\l. l \ list.set as \ length l \ Suc 0" by (auto simp: Suc_leI) then show ?case using Cons.IH by force qed auto fun interact :: "'a list list \ 'a list list \ 'a list" where "interact [] ys = concat ys" | "interact xs [] = concat xs" | "interact (x#xs) (y#ys) = x @ y @ interact xs ys" lemma (in monoid_add) length_interact: "length (interact xs ys) = sum_list (map length xs) + sum_list (map length ys)" by (induction rule: interact.induct) (auto simp: length_concat) lemma length_interact_ge: assumes "xs \ lists (- {[]})" "ys \ lists (- {[]})" shows "length (interact xs ys) \ length xs + length ys" by (metis add_mono assms length_concat length_concat_ge length_interact) lemma set_interact [simp]: shows "list.set (interact xs ys) = list.set (concat xs) \ list.set (concat ys)" by (induction rule: interact.induct) auto lemma interact_eq_Nil_iff [simp]: assumes "xs \ lists (- {[]})" "ys \ lists (- {[]})" shows "interact xs ys = [] \ xs=[] \ ys=[]" using length_interact_ge [OF assms] by fastforce lemma interact_sing [simp]: "interact [x] ys = x @ concat ys" by (metis (no_types) concat.simps(2) interact.simps neq_Nil_conv) lemma hd_interact: "\xs \ []; hd xs \ []\ \ hd (interact xs ys) = hd (hd xs)" by (smt (verit, best) hd_append2 hd_concat interact.elims list.sel(1)) lemma acc_lengths_concat_injective: assumes "concat as' = concat as" "acc_lengths n as' = acc_lengths n as" shows "as' = as" using assms proof (induction as arbitrary: n as') case Nil then show ?case by (metis acc_lengths_eq_Nil_iff) next case (Cons a as) then obtain a' bs where "as' = a'#bs" by (metis Suc_length_conv length_acc_lengths) with Cons show ?case by simp qed lemma acc_lengths_interact_injective: assumes "interact as' bs' = interact as bs" "acc_lengths a as' = acc_lengths a as" "acc_lengths b bs' = acc_lengths b bs" shows "as' = as \ bs' = bs" using assms proof (induction as bs arbitrary: a b as' bs' rule: interact.induct) case (1 cs) then show ?case by (metis acc_lengths_concat_injective acc_lengths_eq_Nil_iff interact.simps(1)) next case (2 c cs) then show ?case by (metis acc_lengths_concat_injective acc_lengths_eq_Nil_iff interact.simps(2) list.exhaust) next case (3 x xs y ys) then obtain a' us b' vs where "as' = a'#us" "bs' = b'#vs" by (metis length_Suc_conv length_acc_lengths) with 3 show ?case by auto qed lemma strict_sorted_interact_I: assumes "length ys \ length xs" "length xs \ Suc (length ys)" "\x. x \ list.set xs \ strict_sorted x" "\y. y \ list.set ys \ strict_sorted y" "\n. n < length ys \ xs!n < ys!n" "\n. Suc n < length xs \ ys!n < xs!Suc n" assumes "xs \ lists (- {[]})" "ys \ lists (- {[]})" shows "strict_sorted (interact xs ys)" using assms proof (induction rule: interact.induct) case (3 x xs y ys) then have "x < y" by force moreover have "strict_sorted (interact xs ys)" using 3 by simp (metis Suc_less_eq nth_Cons_Suc) moreover have "y < interact xs ys" using 3 apply (simp add: less_list_def) by (metis hd_interact le_zero_eq length_greater_0_conv list.sel(1) list.set_sel(1) list.size(3) lists.simps mem_lists_non_Nil nth_Cons_0) ultimately show ?case using 3 by (simp add: strict_sorted_append_iff less_list_def) qed auto subsection \Forms and interactions\ subsubsection \Forms\ inductive Form_Body :: "[nat, nat, nat list, nat list, nat list] \ bool" where "Form_Body ka kb xs ys zs" if "length xs < length ys" "xs = concat (a#as)" "ys = concat (b#bs)" "a#as \ lists (- {[]})" "b#bs \ lists (- {[]})" "length (a#as) = ka" "length (b#bs) = kb" "c = acc_lengths 0 (a#as)" "d = acc_lengths 0 (b#bs)" "zs = concat [c, a, d, b] @ interact as bs" "strict_sorted zs" inductive Form :: "[nat, nat list set] \ bool" where "Form 0 {xs,ys}" if "length xs = length ys" "xs \ ys" | "Form (2*k-1) {xs,ys}" if "Form_Body k k xs ys zs" "k > 0" | "Form (2*k) {xs,ys}" if "Form_Body (Suc k) k xs ys zs" "k > 0" inductive_cases Form_0_cases_raw: "Form 0 u" lemma Form_elim_upair: assumes "Form l U" obtains xs ys where "xs \ ys" "U = {xs,ys}" "length xs \ length ys" using assms by (smt (verit, best) Form.simps Form_Body.cases less_or_eq_imp_le nat_neq_iff) lemma assumes "Form_Body ka kb xs ys zs" shows Form_Body_WW: "zs \ WW" and Form_Body_nonempty: "length zs > 0" and Form_Body_length: "length xs < length ys" using Form_Body.cases [OF assms] by (fastforce simp: WW_def)+ lemma form_cases: fixes l::nat obtains (zero) "l = 0" | (nz) ka kb where "l = ka+kb - 1" "0 < kb" "kb \ ka" "ka \ Suc kb" proof - have "l = 0 \ (\ka kb. l = ka+kb - 1 \ 0 < kb \ kb \ ka \ ka \ Suc kb)" by presburger then show thesis using nz zero by blast qed subsubsection \Interactions\ lemma interact: assumes "Form l U" "l>0" obtains ka kb xs ys zs where "l = ka+kb - 1" "U = {xs,ys}" "Form_Body ka kb xs ys zs" "0 < kb" "kb \ ka" "ka \ Suc kb" using assms unfolding Form.simps by (smt (verit, best) add_Suc diff_Suc_1 lessI mult_2 nat_less_le order_refl) definition inter_scheme :: "nat \ nat list set \ nat list" where "inter_scheme l U \ SOME zs. \k xs ys. U = {xs,ys} \ (l = 2*k-1 \ Form_Body k k xs ys zs \ l = 2*k \ Form_Body (Suc k) k xs ys zs)" lemma inter_scheme: assumes "Form l U" "l>0" obtains ka kb xs ys where "l = ka+kb - 1" "U = {xs,ys}" "Form_Body ka kb xs ys (inter_scheme l U)" "0 < kb" "kb \ ka" "ka \ Suc kb" using interact [OF \Form l U\] proof cases case (2 ka kb xs ys zs) then have \

: "\ka kb zs. \ Form_Body ka kb ys xs zs" using Form_Body_length less_asym' by blast have "Form_Body ka kb xs ys (inter_scheme l U)" proof (cases "ka = kb") case True with 2 have l: "\k. l \ k * 2" by presburger have [simp]: "\k. kb + kb - Suc 0 = k * 2 - Suc 0 \ k=kb" by auto show ?thesis unfolding inter_scheme_def using 2 l True by (auto simp: \
\l > 0\ Set.doubleton_eq_iff conj_disj_distribR ex_disj_distrib algebra_simps some_eq_ex) next case False with 2 have l: "\k. l \ k * 2 - Suc 0" and [simp]: "ka = Suc kb" by presburger+ have [simp]: "\k. kb + kb = k * 2 \ k=kb" by auto show ?thesis unfolding inter_scheme_def using 2 l False by (auto simp: \
\l > 0\ Set.doubleton_eq_iff conj_disj_distribR ex_disj_distrib algebra_simps some_eq_ex) qed then show ?thesis by (simp add: 2 that) qed (use \l > 0\ in auto) lemma inter_scheme_strict_sorted: assumes "Form l U" "l>0" shows "strict_sorted (inter_scheme l U)" using Form_Body.simps assms inter_scheme by fastforce lemma inter_scheme_simple: assumes "Form l U" "l>0" shows "inter_scheme l U \ WW \ length (inter_scheme l U) > 0" using inter_scheme [OF assms] by (meson Form_Body_WW Form_Body_nonempty) subsubsection \Injectivity of interactions\ proposition inter_scheme_injective: assumes "Form l U" "Form l U'" "l > 0" and eq: "inter_scheme l U' = inter_scheme l U" shows "U' = U" proof - obtain ka kb xs ys where l: "l = ka+kb - 1" and U: "U = {xs,ys}" and FB: "Form_Body ka kb xs ys (inter_scheme l U)" and kb: "0 < kb" "kb \ ka" "ka \ Suc kb" using assms inter_scheme by blast then obtain a as b bs c d where xs: "xs = concat (a#as)" and ys: "ys = concat (b#bs)" and len: "length (a#as) = ka" "length (b#bs) = kb" and c: "c = acc_lengths 0 (a#as)" and d: "d = acc_lengths 0 (b#bs)" and Ueq: "inter_scheme l U = concat [c, a, d, b] @ interact as bs" by (auto simp: Form_Body.simps) obtain ka' kb' xs' ys' where l': "l = ka'+kb' - 1" and U': "U' = {xs',ys'}" and FB': "Form_Body ka' kb' xs' ys' (inter_scheme l U')" and kb': "0 < kb'" "kb' \ ka'" "ka' \ Suc kb'" using assms inter_scheme by blast then obtain a' as' b' bs' c' d' where xs': "xs' = concat (a'#as')" and ys': "ys' = concat (b'#bs')" and len': "length (a'#as') = ka'" "length (b'#bs') = kb'" and c': "c' = acc_lengths 0 (a'#as')" and d': "d' = acc_lengths 0 (b'#bs')" and Ueq': "inter_scheme l U' = concat [c', a', d', b'] @ interact as' bs'" using Form_Body.simps by auto have [simp]: "ka' = ka \ kb' = kb" using \l > 0\ l l' kb kb' le_SucE le_antisym mult_2 by linarith have [simp]: "length c = length c'" "length d = length d'" using c c' d d' len' len by auto have c_off: "c' = c" "a' @ d' @ b' @ interact as' bs' = a @ d @ b @ interact as bs" using eq by (auto simp: Ueq Ueq') then have len_a: "length a' = length a" by (metis acc_lengths.simps(2) add.left_neutral c c' nth_Cons_0) with c_off have \
: "a' = a" "d' = d" "b' @ interact as' bs' = b @ interact as bs" by auto then have "length (interact as' bs') = length (interact as bs)" by (metis acc_lengths.simps(2) add_left_cancel append_eq_append_conv d d' list.inject) with \
have "b' = b" "interact as' bs' = interact as bs" by auto moreover have "acc_lengths 0 as' = acc_lengths 0 as" using \a' = a\ \c' = c\ by (simp add: c' c acc_lengths_shift) moreover have "acc_lengths 0 bs' = acc_lengths 0 bs" using \b' = b\ \d' = d\ by (simp add: d' d acc_lengths_shift) ultimately have "as' = as \ bs' = bs" using acc_lengths_interact_injective by blast then show ?thesis by (simp add: \a' = a\ U U' \b' = b\ xs xs' ys ys') qed lemma strict_sorted_interact_imp_concat: "strict_sorted (interact as bs) \ strict_sorted (concat as) \ strict_sorted (concat bs)" proof (induction as bs rule: interact.induct) case (3 x xs y ys) have "x < concat xs" using "3.prems" by (smt (verit, del_insts) Un_iff hd_in_set interact.simps(3) last_in_set less_list_def set_append set_interact sorted_wrt_append) moreover have "y < concat ys" using 3 sorted_wrt_append strict_sorted_append_iff by fastforce ultimately show ?case using 3 by (auto simp add: strict_sorted_append_iff) qed auto lemma strict_sorted_interact_hd: "\strict_sorted (interact cs ds); cs \ []; ds \ []; hd cs \ []; hd ds \ []\ \ hd (hd cs) < hd (hd ds)" by (metis append_is_Nil_conv hd_append2 hd_in_set interact.simps(3) list.exhaust_sel sorted_wrt_append) text \the lengths of the two lists can differ by one\ proposition interaction_scheme_unique_aux: assumes "concat as = concat as'" and ys': "concat bs = concat bs'" and "as \ lists (- {[]})" "bs \ lists (- {[]})" and "strict_sorted (interact as bs)" and "length bs \ length as" "length as \ Suc (length bs)" and "as' \ lists (- {[]})" "bs' \ lists (- {[]})" and "strict_sorted (interact as' bs')" and "length bs' \ length as'" "length as' \ Suc (length bs')" and "length as = length as'" "length bs = length bs'" shows "as = as' \ bs = bs'" using assms proof (induction "length as" arbitrary: as bs as' bs') case 0 then show ?case by auto next case (Suc k) show ?case proof (cases k) case 0 then obtain a a' where aa': "as = [a]" "as' = [a']" by (metis Suc.hyps(2) \length as = length as'\ Suc_length_conv length_0_conv) show ?thesis proof show "as = as'" using aa' \concat as = concat as'\ by force with Suc 0 show " bs = bs'" by (metis Suc_leI append_Nil2 concat.simps impossible_Cons le_antisym length_greater_0_conv list.exhaust) qed next case (Suc k') then obtain a cs b ds where eq: "as = a#cs" "bs = b#ds" using Suc.prems by (metis Suc.hyps(2) le0 list.exhaust list.size(3) not_less_eq_eq) have "length as' \ 0" using Suc.hyps(2) \length as = length as'\ by force then obtain a' cs' b' ds' where eq': "as' = a'#cs'" "bs' = b'#ds'" by (metis \length bs = length bs'\ eq(2) length_0_conv list.exhaust) obtain k: "k = length cs" "k \ Suc (length ds)" using eq \Suc k = length as\ \length bs \ length as\ \length as \ Suc (length bs)\ by auto case (Suc k') obtain [simp]: "b \ []" "b' \ []" "a \ []" "a' \ []" using Suc.prems by (simp add: eq eq') then have "hd b' = hd b" using Suc.prems(2) by (metis concat.simps(2) eq'(2) eq(2) hd_append2) have ss_ab: "strict_sorted (concat as)" "strict_sorted (concat bs)" using strict_sorted_interact_imp_concat Suc.prems(5) by blast+ have sw_ab: "strict_sorted (a @ b @ interact cs ds)" by (metis Suc.prems(5) eq interact.simps(3)) then obtain "a < b" "strict_sorted a" "strict_sorted b" by (metis append_assoc strict_sorted_append_iff) have b_cs: "strict_sorted (concat (b # cs))" by (metis append.simps(1) concat.simps(2) interact.simps(3) strict_sorted_interact_imp_concat sw_ab) then have "b < concat cs" using strict_sorted_append_iff by auto have "strict_sorted (a @ concat cs)" using eq(1) ss_ab(1) by force have "list.set a = list.set (concat as) \ {..< hd b}" proof - have "x \ list.set a" if "x < hd b" and "l \ list.set cs" and "x \ list.set l" for x l using b_cs sorted_hd_le strict_sorted_imp_sorted that by fastforce then show ?thesis using \b \ []\ sw_ab by (force simp: strict_sorted_append_iff sorted_wrt_append eq) qed moreover have ss_ab': "strict_sorted (concat as')" "strict_sorted (concat bs')" using strict_sorted_interact_imp_concat Suc.prems(10) by blast+ have sw_ab': "strict_sorted (a' @ b' @ interact cs' ds')" by (metis Suc.prems(10) eq' interact.simps(3)) then obtain "a' < b'" "strict_sorted a'" "strict_sorted b'" by (metis append_assoc strict_sorted_append_iff) have b_cs': "strict_sorted (concat (b' # cs'))" by (metis (no_types) Suc.prems(10) append_Nil eq' interact.simps(3) strict_sorted_append_iff strict_sorted_interact_imp_concat) then have "b' < concat cs'" by (simp add: strict_sorted_append_iff) then have "hd b' \ list.set (concat cs')" by (metis Un_iff \b' \ []\ list.set_sel(1) not_less_iff_gr_or_eq set_interact sorted_wrt_append sw_ab') have "strict_sorted (a' @ concat cs')" using eq'(1) ss_ab'(1) by force then have b_cs': "strict_sorted (b' @ concat cs')" using \b' < concat cs'\ eq'(2) ss_ab'(2) strict_sorted_append_iff by auto have "list.set a' = list.set (concat as') \ {..< hd b'}" proof - have "x \ list.set a'" if "x < hd b'" and "l \ list.set cs'" and "x \ list.set l" for x l using b_cs' sorted_hd_le strict_sorted_imp_sorted that by fastforce then show ?thesis using \b' \ []\ sw_ab' by (force simp: strict_sorted_append_iff sorted_wrt_append eq') qed ultimately have "a=a'" by (simp add: Suc.prems(1) \hd b' = hd b\ \strict_sorted a'\ \strict_sorted a\ strict_sorted_equal) moreover have ccat_cs_cs': "concat cs = concat cs'" using Suc.prems(1) \a = a'\ eq'(1) eq(1) by fastforce have "b=b'" proof (cases "ds = [] \ ds' = []") case True then show ?thesis using \length bs = length bs'\ Suc.prems(2) eq'(2) eq(2) by auto next case False then have "ds \ []" "ds' \ []" "sorted (concat ds)" "sorted (concat ds')" using eq(2) ss_ab(2) eq'(2) ss_ab'(2) strict_sorted_append_iff strict_sorted_imp_sorted by auto have "strict_sorted b" "strict_sorted b'" using b_cs b_cs' sorted_wrt_append by auto moreover have "cs \ []" using k local.Suc by auto then obtain "hd cs \ []" "hd ds \ []" using Suc.prems(3) Suc.prems(4) eq list.set_sel(1) by (simp add: \ds \ []\ mem_lists_non_Nil) then have "concat cs \ []" using \cs \ []\ hd_in_set by auto have "hd (concat cs) < hd (concat ds)" using strict_sorted_interact_hd by (metis \cs \ []\ \ds \ []\ \hd cs \ []\ \hd ds \ []\ hd_concat sorted_wrt_append sw_ab) have "list.set b = list.set (concat bs) \ {..< hd (concat cs)}" proof - have 1: "x \ list.set b" if "x < hd (concat cs)" and "l \ list.set ds" and "x \ list.set l" for x l using \hd (concat cs) < hd (concat ds)\ \sorted (concat ds)\ sorted_hd_le that by fastforce have 2: "l < hd (concat cs)" if "l \ list.set b" for l by (meson \b < concat cs\ \b \ []\ \concat cs \ []\ \strict_sorted b\ le_less_trans less_list_def sorted_le_last strict_sorted_imp_sorted that) show ?thesis using 1 2 by (auto simp: strict_sorted_append_iff sorted_wrt_append eq) qed moreover have "cs' \ []" using k local.Suc \concat cs \ []\ ccat_cs_cs' by auto then obtain "hd cs' \ []" "hd ds' \ []" using Suc.prems(8,9) \ds' \ []\ eq'(1) eq'(2) list.set_sel(1) by auto then have "concat cs' \ []" using \cs' \ []\ hd_in_set by auto have "hd (concat cs') < hd (concat ds')" using strict_sorted_interact_hd by (metis \cs' \ []\ \ds' \ []\ \hd cs' \ []\ \hd ds' \ []\ hd_concat sorted_wrt_append sw_ab') have "list.set b' = list.set (concat bs') \ {..< hd (concat cs')}" proof - have 1: "x \ list.set b'" if "x < hd (concat cs')" and "l \ list.set ds'" and "x \ list.set l" for x l using \hd (concat cs') < hd (concat ds')\ \sorted (concat ds')\ sorted_hd_le that by fastforce have 2: "l < hd (concat cs')" if "l \ list.set b'" for l by (metis \concat cs' \ []\ b_cs' list.set_sel(1) sorted_wrt_append that) show ?thesis using 1 2 by (auto simp: strict_sorted_append_iff sorted_wrt_append eq') qed ultimately show "b = b'" by (simp add: Suc.prems(2) ccat_cs_cs' strict_sorted_equal) qed moreover have "cs = cs' \ ds = ds'" proof (rule Suc.hyps) show "k = length cs" using eq Suc.hyps(2) by auto[1] show "concat ds = concat ds'" using Suc.prems(2) \b = b'\ eq'(2) eq(2) by auto show "strict_sorted (interact cs ds)" using eq Suc.prems(5) strict_sorted_append_iff by auto show "length ds \ length cs" "length cs \ Suc (length ds)" using eq Suc.hyps(2) Suc.prems(6) k by auto show "strict_sorted (interact cs' ds')" using eq' Suc.prems(10) strict_sorted_append_iff by auto show "length cs = length cs'" using Suc.hyps(2) Suc.prems(13) eq'(1) k(1) by force qed (use ccat_cs_cs' eq eq' Suc.prems in auto) ultimately show ?thesis by (simp add: \a = a'\ \b = b'\ eq eq') qed qed proposition Form_Body_unique: assumes "Form_Body ka kb xs ys zs" "Form_Body ka kb xs ys zs'" and "kb \ ka" "ka \ Suc kb" shows "zs' = zs" proof - obtain a as b bs c d where xs: "xs = concat (a#as)" and ys: "ys = concat (b#bs)" and ne: "a#as \ lists (- {[]})" "b#bs \ lists (- {[]})" and len: "length (a#as) = ka" "length (b#bs) = kb" and c: "c = acc_lengths 0 (a#as)" and d: "d = acc_lengths 0 (b#bs)" and Ueq: "zs = concat [c, a, d, b] @ interact as bs" and ss_zs: "strict_sorted zs" using Form_Body.cases [OF assms(1)] by (metis (no_types)) obtain a' as' b' bs' c' d' where xs': "xs = concat (a'#as')" and ys': "ys = concat (b'#bs')" and ne': "a'#as' \ lists (- {[]})" "b'#bs' \ lists (- {[]})" and len': "length (a'#as') = ka" "length (b'#bs') = kb" and c': "c' = acc_lengths 0 (a'#as')" and d': "d' = acc_lengths 0 (b'#bs')" and Ueq': "zs' = concat [c', a', d', b'] @ interact as' bs'" and ss_zs': "strict_sorted zs'" using Form_Body.cases [OF assms(2)] by (metis (no_types)) have [simp]: "length c = length c'" "length d = length d'" using c c' d d' len' len by auto note acc_lengths.simps [simp del] have "a < b" using ss_zs by (auto simp: Ueq strict_sorted_append_iff less_list_def c d) have "a' < b'" using ss_zs' by (auto simp: Ueq' strict_sorted_append_iff less_list_def c' d') have "a#as = a'#as' \ b#bs = b'#bs'" proof (rule interaction_scheme_unique_aux) show "strict_sorted (interact (a # as) (b # bs))" using ss_zs \a < b\ by (auto simp: Ueq strict_sorted_append_iff less_list_def d) show "strict_sorted (interact (a' # as') (b' # bs'))" using ss_zs' \a' < b'\ by (auto simp: Ueq' strict_sorted_append_iff less_list_def d') show "length (b # bs) \ length (a # as)" "length (b' # bs') \ length (a' # as')" using \kb \ ka\ len len' by auto show "length (a # as) \ Suc (length (b # bs))" using \ka \ Suc kb\ len by linarith then show "length (a' # as') \ Suc (length (b' # bs'))" using len len' by fastforce qed (use len len' xs xs' ys ys' ne ne' in fastforce)+ then show ?thesis using Ueq Ueq' c c' d d' by blast qed lemma Form_Body_imp_inter_scheme: assumes FB: "Form_Body ka kb xs ys zs" and "0 < kb" "kb \ ka" "ka \ Suc kb" shows "zs = inter_scheme ((ka+kb) - Suc 0) {xs,ys}" proof - have "length xs < length ys" by (meson Form_Body_length assms(1)) have [simp]: "a + a = b + b \ a=b" "a + a - Suc 0 = b + b - Suc 0 \ a=b" for a b::nat by auto show ?thesis proof (cases "ka = kb") case True show ?thesis unfolding inter_scheme_def apply (rule some_equality [symmetric], metis One_nat_def True FB mult_2) subgoal for zs' using assms \length xs < length ys\ by (auto simp: True mult_2 Set.doubleton_eq_iff Form_Body_unique dest: Form_Body_length, presburger) done next case False then have eq: "ka = Suc kb" using assms by linarith show ?thesis unfolding inter_scheme_def apply (rule some_equality [symmetric], use assms False mult_2 one_is_add eq in fastforce) subgoal for zs' using assms \length xs < length ys\ by (auto simp: eq mult_2 Set.doubleton_eq_iff Form_Body_unique dest: Form_Body_length, presburger) done qed qed subsection \For Lemma 3.8 AND PROBABLY 3.7\ definition grab :: "nat set \ nat \ nat set \ nat set" where "grab N n \ (N \ enumerate N ` {.. {enumerate N n..})" lemma grab_0 [simp]: "grab N 0 = ({}, N)" by (fastforce simp: grab_def enumerate_0 Least_le) lemma less_sets_grab: "infinite N \ fst (grab N n) \ snd (grab N n)" by (auto simp: grab_def less_sets_def intro: enumerate_mono less_le_trans) lemma finite_grab [iff]: "finite (fst (grab N n))" by (simp add: grab_def) lemma card_grab [simp]: assumes "infinite N" shows "card (fst (grab N n)) = n" proof - have "N \ enumerate N ` {.. N" using grab_def range_enum by fastforce lemma snd_grab_subset: "snd (grab N n) \ N" by (auto simp: grab_def) lemma grab_Un_eq: assumes "infinite N" shows "fst (grab N n) \ snd (grab N n) = N" proof show "N \ fst (grab N n) \ snd (grab N n)" unfolding grab_def using assms enumerate_Ex le_less_linear strict_mono_enum strict_mono_less by fastforce qed (simp add: grab_def) lemma finite_grab_iff [simp]: "finite (snd (grab N n)) \ finite N" by (metis finite_grab grab_Un_eq infinite_Un infinite_super snd_grab_subset) lemma grab_eqD: "\grab N n = (A,M); infinite N\ \ A \ M \ finite A \ card A = n \ infinite M \ A \ N \ M \ N" using card_grab grab_def less_sets_grab finite_grab_iff by auto lemma less_sets_fst_grab: "A \ N \ A \ fst (grab N n)" by (simp add: fst_grab_subset less_sets_weaken2) text\Possibly redundant, given @{term grab}\ definition nxt where "nxt \ \N. \n::nat. N \ {n<..}" lemma infinite_nxtN: "infinite N \ infinite (nxt N n)" by (simp add: infinite_nat_greaterThan nxt_def) lemma nxt_subset: "nxt N n \ N" unfolding nxt_def by blast lemma nxt_subset_greaterThan: "m \ n \ nxt N n \ {m<..}" by (auto simp: nxt_def) lemma nxt_subset_atLeast: "m \ n \ nxt N n \ {m..}" by (auto simp: nxt_def) lemma enum_nxt_ge: "infinite N \ a \ enum (nxt N a) n" by (simp add: atLeast_le_enum infinite_nxtN nxt_subset_atLeast) lemma inj_enum_nxt: "infinite N \ inj_on (enum (nxt N a)) A" by (simp add: infinite_nxtN strict_mono_enum strict_mono_imp_inj_on) subsection \Larson's Lemma 3.11\ text \Again from Jean A. Larson, A short proof of a partition theorem for the ordinal $\omega^\omega$. \emph{Annals of Mathematical Logic}, 6:129--145, 1973.\ lemma lemma_3_11: assumes "l > 0" shows "thin (inter_scheme l ` {U. Form l U})" using form_cases [of l] proof cases case zero then show ?thesis using assms by auto next case (nz ka kb) note acc_lengths.simps [simp del] show ?thesis unfolding thin_def proof clarify fix U U' assume ne: "inter_scheme l U \ inter_scheme l U'" and init: "initial_segment (inter_scheme l U) (inter_scheme l U')" assume "Form l U" then obtain kp kq xs ys where "l = kp+kq - 1" "U = {xs,ys}" and U: "Form_Body kp kq xs ys (inter_scheme l U)" and "0 < kq" "kq \ kp" "kp \ Suc kq" using assms inter_scheme by blast then have "kp = ka \ kq = kb" using nz by linarith then obtain a as b bs c d where len: "length (a#as) = ka" "length (b#bs) = kb" and c: "c = acc_lengths 0 (a#as)" and d: "d = acc_lengths 0 (b#bs)" and Ueq: "inter_scheme l U = concat [c, a, d, b] @ interact as bs" using U by (auto simp: Form_Body.simps) assume "Form l U'" then obtain kp' kq' xs' ys' where "l = kp'+kq' - 1" "U' = {xs',ys'}" and U': "Form_Body kp' kq' xs' ys' (inter_scheme l U')" and "0 < kq'" "kq' \ kp'" "kp' \ Suc kq'" using assms inter_scheme by blast then have "kp' = ka \ kq' = kb" using nz by linarith then obtain a' as' b' bs' c' d' where len': "length (a'#as') = ka" "length (b'#bs') = kb" and c': "c' = acc_lengths 0 (a'#as')" and d': "d' = acc_lengths 0 (b'#bs')" and Ueq': "inter_scheme l U' = concat [c', a', d', b'] @ interact as' bs'" using U' by (auto simp: Form_Body.simps) have [simp]: "length bs' = length bs" "length as' = length as" using len len' by auto have "inter_scheme l U \ []" "inter_scheme l U' \ []" using Form_Body_nonempty U U' by auto define u1 where "u1 \ hd (inter_scheme l U)" have u1_eq': "u1 = hd (inter_scheme l U')" using \inter_scheme l U \ []\ init u1_def initial_segment_ne by fastforce have au1: "u1 = length a" by (simp add: u1_def Ueq c) have au1': "u1 = length a'" by (simp add: u1_eq' Ueq' c') have len_eqk: "length c' = ka" "length d' = kb" "length c' = ka" "length d' = kb" using c d len c' d' len' by auto have take: "take (ka + u1 + kb) (c @ a @ d @ l) = c @ a @ d" "take (ka + u1 + kb) (c' @ a' @ d' @ l) = c' @ a' @ d'" for l using c d c' d' len by (simp_all flip: au1 au1') have leU: "ka + u1 + kb \ length (inter_scheme l U)" using c d len by (simp add: au1 Ueq) then have "take (ka + u1 + kb) (inter_scheme l U) = take (ka + u1 + kb) (inter_scheme l U')" using take_initial_segment init by blast then have \
: "c @ a @ d = c' @ a' @ d'" by (metis Ueq Ueq' append.assoc concat.simps(2) take) have "length (inter_scheme l U) = ka + (c @ a @ d)!(ka-1) + kb + last d" by (simp add: Ueq c d length_interact nth_append flip: len) moreover have "length (inter_scheme l U') = ka + (c' @ a' @ d')!(ka-1) + kb + last d'" by (simp add: Ueq' c' d' length_interact nth_append flip: len') moreover have "last d = last d'" using "\
" c d d' len'(1) len_eqk(1) by auto ultimately have "length (inter_scheme l U) = length (inter_scheme l U')" by (simp add: \
) then show False using init initial_segment_length_eq ne by blast qed qed subsection \Larson's Lemma 3.6\ proposition lemma_3_6: fixes g :: "nat list set \ nat" assumes g: "g \ [WW]\<^bsup>2\<^esup> \ {0,1}" obtains N j where "infinite N" and "\k u. \k > 0; u \ [WW]\<^bsup>2\<^esup>; Form k u; [enum N k] < inter_scheme k u; List.set (inter_scheme k u) \ N\ \ g u = j k" proof - define \ where "\ \ \m::nat. \M. infinite M \ m < Inf M" define \ where "\ \ \l m n::nat. \M N j. n > m \ N \ M \ n \ M \ (\U. Form l U \ U \ WW \ [n] < inter_scheme l U \ list.set (inter_scheme l U) \ N \ g U = j)" have *: "\n N j. \ n N \ \ l m n M N j" if "l > 0" "\ m M" for l m::nat and M :: "nat set" proof - define FF where "FF \ {U \ [WW]\<^bsup>2\<^esup>. Form l U}" define h where "h \ \zs. g (inv_into FF (inter_scheme l) zs)" have "thin (inter_scheme l ` FF)" using \l > 0\ lemma_3_11 by (simp add: thin_def FF_def) moreover have "inter_scheme l ` FF \ WW" using inter_scheme_simple \0 < l\ FF_def by blast moreover have "h ` {xs \ inter_scheme l ` FF. List.set xs \ M} \ {..<2}" using g inv_into_into[of concl: "FF" "inter_scheme l"] by (force simp: h_def FF_def Pi_iff) ultimately obtain j N where "j < 2" "infinite N" "N \ M" and hj: "h ` {xs \ inter_scheme l ` FF. List.set xs \ N} \ {j}" using \\ m M\ unfolding \_def by (blast intro: Nash_Williams_WW [of M]) let ?n = "Inf N" have "?n > m" using \\ m M\ \infinite N\ unfolding \_def Inf_nat_def infinite_nat_iff_unbounded by (metis LeastI_ex \N \ M\ le_less_trans not_less not_less_Least subsetD) have "g U = j" if "Form l U" "U \ WW" "[?n] < inter_scheme l U" "list.set (inter_scheme l U) \ N - {?n}" for U proof - obtain xs ys where xys: "xs \ ys" "U = {xs,ys}" using Form_elim_upair \Form l U\ by blast moreover have "inj_on (inter_scheme l) FF" using \0 < l\ inj_on_def inter_scheme_injective FF_def by blast moreover have "g (inv_into FF (inter_scheme l) (inter_scheme l U)) = j" using hj that xys subset_Diff_insert by (fastforce simp: h_def FF_def image_iff) ultimately show ?thesis using that FF_def by auto qed moreover have "?n < Inf (N - {?n})" by (metis Diff_iff Inf_nat_def Inf_nat_def1 \infinite N\ finite.emptyI infinite_remove linorder_neqE_nat not_less_Least singletonI) moreover have "?n \ M" by (metis Inf_nat_def1 \N \ M\ \infinite N\ finite.emptyI subsetD) ultimately have "\ ?n (N - {?n}) \ \ l m ?n M (N - {?n}) j" using \\ m M\ \infinite N\ \N \ M\ \?n > m\ by (auto simp: \_def \_def) then show ?thesis by blast qed have base: "\ 0 {0<..}" unfolding \_def by (metis infinite_Ioi Inf_nat_def1 greaterThan_iff greaterThan_non_empty) have step: "Ex (\(n,N,j). \ n N \ \ l m n M N j)" if "\ m M" "l > 0" for m M l using * [of l m M] that by (auto simp: \_def) define G where "G \ \l m M. @(n,N,j). \ n N \ \ (Suc l) m n M N j" have G\: "(\(n,N,j). \ n N) (G l m M)" and G\: "(\(n,N,j). \ (Suc l) m n M N j) (G l m M)" if "\ m M" for l m M using step [OF that, of "Suc l"] by (force simp: G_def dest: some_eq_imp)+ have G_increasing: "(\(n,N,j). n > m \ N \ M \ n \ M) (G l m M)" if "\ m M" for l m M using G\ [OF that, of l] that by (simp add: \_def split: prod.split_asm) define H where "H \ rec_nat (0,{0<..},0) (\l (m,M,j). G l m M)" have H_simps: "H 0 = (0,{0<..},0)" "\l. H (Suc l) = (case H l of (m,M,j) \ G l m M)" by (simp_all add: H_def) have H\: "(\(n,N,j). \ n N) (H l)" for l by (induction l) (use base G\ in \auto simp: H_simps split: prod.split_asm\) define \ where "\ \ (\l. case H l of (n,M,j) \ n)" have H_inc: "\ l \ l" for l proof (induction l) case (Suc l) then show ?case using H\ [of l] G_increasing [of "\ l"] apply (clarsimp simp: H_simps \_def split: prod.split) by (metis (no_types, lifting) case_prodD leD le_less_trans not_less_eq_eq) qed auto let ?N = "range \" define j where "j \ \l. case H l of (n,M,j) \ j" have H_increasing_Suc: "(case H k of (n, N, j') \ N) \ (case H (Suc k) of (n, N, j') \ insert n N)" for k using H\ [of k] by (force simp: H_simps split: prod.split dest: G_increasing [where l=k]) have H_increasing_superset: "(case H k of (n, N, j') \ N) \ (case H (n+k) of (n, N, j') \ N)" for k n proof (induction n) case (Suc n) then show ?case using H_increasing_Suc [of "n+k"] by (auto split: prod.split_asm) qed auto then have H_increasing_less: "(case H k of (n, N, j') \ N) \ (case H l of (n, N, j') \ insert n N)" if "k k < \ (Suc k)" for k using H\ [of k] unfolding \_def by (auto simp: H_simps split: prod.split dest: G_increasing [where l=k]) then have strict_mono_\: "strict_mono \" by (simp add: strict_mono_Suc_iff) then have enum_N: "enum ?N = \" by (metis enum_works nat_infinite_iff range_strict_mono_ext) have **: "?N \ {n<..} \ N'" if H: "H k = (n, N', j)" for n N' k j proof clarify fix l assume "n < \ l" then have False if "l \ k" using that strict_monoD [OF strict_mono_\, of l k ] H by (force simp: \_def) then have "k < l" using not_less by blast then obtain M j where Mj: "H l = (\ l,M,j)" unfolding \_def by (metis (mono_tags, lifting) case_prod_conv old.prod.exhaust) then show "\ l \ N'" using that H_increasing_less [OF \k] Mj by auto qed show thesis proof show "infinite (?N::nat set)" using H_inc infinite_nat_iff_unbounded_le by auto next fix l U assume "0 < l" and U: "U \ [WW]\<^bsup>2\<^esup>" and interU: "[enum ?N l] < inter_scheme l U" "Form l U" and sub: "list.set (inter_scheme l U) \ ?N" obtain k where k: "l = Suc k" using \0 < l\ gr0_conv_Suc by blast have "g U = v" if "H k = (m, M, j0)" and "G k m M = (n, N', v)" for m M j0 n N' v proof - have n: "\ (Suc k) = n" using that by (simp add: \_def H_simps) have "{..enum (range \) l} \ list.set (inter_scheme l U) = {}" using inter_scheme_strict_sorted \0 < l\ interU singleton_less_list_iff strict_sorted_iff by blast then have "list.set (inter_scheme (Suc k) U) \ N'" using that sub ** [of "Suc k" n N' v] Suc_le_eq not_less_eq_eq by (fastforce simp: k n enum_N H_simps) then show ?thesis using that interU U G\ [of m M k] H\ [of k] by (auto simp: \_def k enum_N H_simps n nsets_def) qed with U show "g U = j l" by (auto simp: k j_def H_simps split: prod.split) qed qed subsection \Larson's Lemma 3.7\ subsubsection \Preliminaries\ text \Analogous to @{thm [source] ordered_nsets_2_eq}, but without type classes\ lemma total_order_nsets_2_eq: assumes tot: "total_on A r" and irr: "irrefl r" shows "nsets A 2 = {{x,y} | x y. x \ A \ y \ A \ (x,y) \ r}" (is "_ = ?rhs") proof show "nsets A 2 \ ?rhs" using tot unfolding numeral_nat total_on_def nsets_def by (fastforce simp: card_Suc_eq Set.doubleton_eq_iff not_less) show "?rhs \ nsets A 2" using irr unfolding numeral_nat by (force simp: nsets_def card_Suc_eq irrefl_def) qed lemma lenlex_nsets_2_eq: "nsets A 2 = {{x,y} | x y. x \ A \ y \ A \ (x,y) \ lenlex less_than}" using total_order_nsets_2_eq by (simp add: total_order_nsets_2_eq irrefl_def) lemma sum_sorted_list_of_set_map: "finite I \ sum_list (map f (list_of I)) = sum f I" proof (induction "card I" arbitrary: I) case (Suc n I) then have [simp]: "I \ {}" by auto have "sum_list (map f (list_of (I - {Min I}))) = sum f (I - {Min I})" using Suc by auto then show ?case using Suc.prems sum.remove [of I "Min I" f] by (simp add: sorted_list_of_set_nonempty Suc) qed auto lemma sorted_list_of_set_UN_eq_concat: assumes I: "strict_mono_sets I f" "finite I" and fin: "\i. finite (f i)" shows "list_of (\i \ I. f i) = concat (map (list_of \ f) (list_of I))" using I proof (induction "card I" arbitrary: I) case (Suc n I) then have "I \ {}" and Iexp: "I = insert (Min I) (I - {Min I})" using Min_in Suc.hyps(2) Suc.prems(2) by fastforce+ have IH: "list_of (\ (f ` (I - {Min I}))) = concat (map (list_of \ f) (list_of (I - {Min I})))" using Suc unfolding strict_mono_sets_def by (metis DiffE Iexp card_Diff_singleton diff_Suc_1 finite_Diff insertI1) have "list_of (\ (f ` I)) = list_of (\ (f ` (insert (Min I) (I - {Min I}))))" using Iexp by auto also have "\ = list_of (f (Min I) \ \ (f ` (I - {Min I})))" by (metis Union_image_insert) also have "\ = list_of (f (Min I)) @ list_of (\ (f ` (I - {Min I})))" proof (rule sorted_list_of_set_Un) show "f (Min I) \ \ (f ` (I - {Min I}))" using Suc.prems \I \ {}\ strict_mono_less_sets_Min by blast show "finite (\ (f ` (I - {Min I})))" by (simp add: \finite I\ fin) qed (use fin in auto) also have "\ = list_of (f (Min I)) @ concat (map (list_of \ f) (list_of (I - {Min I})))" using IH by metis also have "\ = concat (map (list_of \ f) (list_of I))" by (simp add: Suc.prems(2) \I \ {}\ sorted_list_of_set_nonempty) finally show ?case . qed auto subsubsection \Lemma 3.7 of Jean A. Larson, ibid.\ proposition lemma_3_7: assumes "infinite N" "l > 0" obtains M where "M \ [WW]\<^bsup>m\<^esup>" "\U. U \ [M]\<^bsup>2\<^esup> \ Form l U \ List.set (inter_scheme l U) \ N" proof (cases "m < 2") case True obtain w where w: "w \ WW" using WW_def strict_sorted_into_WW by auto define M where "M \ if m=0 then {} else {w}" have M: "M \ [WW]\<^bsup>m\<^esup>" using True by (auto simp: M_def nsets_def w) have [simp]: "[M]\<^bsup>2\<^esup> = {}" using True by (auto simp: M_def nsets_def w dest: subset_singletonD) show ?thesis using M that by fastforce next case False then have "m \ 2" by auto have nonz: "(enum N \ Suc) i > 0" for i using assms(1) le_enumerate less_le_trans by fastforce note infinite_nxt_N = infinite_nxtN [OF \infinite N\, iff] note \infinite N\ [ iff] have [simp]: "{n<.. \k D. enum (nxt N (enum (nxt N (Max D)) (Inf D - 1))) ` {.. \k n. (DF_Suc k ^^ n) ((enum N \ Suc) ` {.. Suc) ` {.. Suc) {..infinite N\ DF_simps DF_Suc_def card_image infinite_nxtN strict_mono_enum strict_mono_imp_inj_on) qed have DF_ne: "DF k i \ {}" for i k by (metis card_DF card_lessThan lessThan_empty_iff nat.simps(3)) have finite_DF: "finite (DF k i)" for i k by (induction i) (auto simp: DF_simps DF_Suc_def) have DF_Suc: "DF k i \ DF k (Suc i)" for i k unfolding less_sets_def by (force simp: finite_DF DF_simps DF_Suc_def intro!: greaterThan_less_enum nxt_subset_greaterThan atLeast_le_enum nxt_subset_atLeast infinite_nxtN [OF \infinite N\]) have DF_DF: "DF k i \ DF k j" if "i N" for i k proof (induction i) case 0 then show ?case using \infinite N\ range_enum by (auto simp: DF_simps) next case (Suc i) then show ?case unfolding DF_simps DF_Suc_def image_subset_iff by (metis IntE \infinite N\ enumerate_in_set infinite_nxtN nxt_def) qed - have sm_enum_DF: "strict_mono_on (enum (DF k i)) {..k}" for k i + have sm_enum_DF: "strict_mono_on {..k} (enum (DF k i))" for k i by (metis card_DF enum_works_finite finite_DF lessThan_Suc_atMost) define AF where "AF \ \k i. enum (nxt N (Max (DF k i))) ` {.. {}" for i k by (auto simp: AF_def lessThan_empty_iff DF_gt0) have finite_AF [simp]: "finite (AF k i)" for i k by (simp add: AF_def) have card_AF: "card (AF k i) = \ (DF k i)" for k i by (simp add: AF_def card_image inj_enum_nxt) have DF_AF: "DF k i \ AF k i" for i k unfolding less_sets_def AF_def by (simp add: finite_DF greaterThan_less_enum nxt_subset_greaterThan) have E: "\x \ y; infinite M\ \ enum M x < enum (nxt N (enum M y)) z" for x y z M by (metis infinite_nxt_N dual_order.eq_iff enumerate_mono greaterThan_less_enum nat_less_le nxt_subset_greaterThan) have AF_DF_Suc: "AF k i \ DF k (Suc i)" for i k by (auto simp: DF_simps DF_Suc_def less_sets_def AF_def E) have AF_DF: "AF k p \ DF k q" if "p AF k (Suc i)" for i k using AF_DF_Suc DF_AF DF_ne less_sets_trans by blast then have sm_AF: "strict_mono_sets UNIV (AF k)" for k by (simp add: AF_ne less_sets_imp_strict_mono_sets) define del where "del \ \k i j. enum (DF k i) j - enum (DF k i) (j - 1)" define QF where "QF k \ wfrec pair_less (\f (j,i). if j=0 then AF k i else let r = (if i=0 then f (j-1,m-1) else f (j,i-1)) in enum (nxt N (Suc (Max r))) ` {..< del k (if j=k then m - Suc i else i) j})" for k note cut_apply [simp] have finite_QF [simp]: "finite (QF k p)" for p k using wf_pair_less proof (induction p rule: wf_induct_rule) case (less p) then show ?case by (simp add: def_wfrec [OF QF_def, of k p] split: prod.split) qed have del_gt_0: "\j < Suc k; 0 < j\ \ 0 < del k i j" for i j k by (simp add: card_DF del_def finite_DF) have QF_ne [simp]: "QF k (j,i) \ {}" if j: "j < Suc k" for j i k using wf_pair_less j proof (induction "(j,i)" rule: wf_induct_rule) case less then show ?case by (auto simp: def_wfrec [OF QF_def, of k "(j,i)"] AF_ne lessThan_empty_iff del_gt_0) qed have QF_0 [simp]: "QF k (0,i) = AF k i" for i k by (simp add: def_wfrec [OF QF_def]) have QF_Suc: "QF k (Suc j,0) = enum (nxt N (Suc (Max (QF k (j, m - 1))))) ` {..< del k (if Suc j = k then m - 1 else 0) (Suc j)}" for j k apply (simp add: def_wfrec [OF QF_def, of k "(Suc j,0)"] One_nat_def) apply (simp add: pair_less_def cut_def) done have QF_Suc_Suc: "QF k (Suc j, Suc i) = enum (nxt N (Suc (Max (QF k (Suc j, i))))) ` {..< del k (if Suc j = k then m - Suc(Suc i) else Suc i) (Suc j)}" for i j k by (simp add: def_wfrec [OF QF_def, of k "(Suc j,Suc i)"]) have less_QF1: "QF k (j, m - 1) \ QF k (Suc j,0)" for j k by (auto simp: def_wfrec [OF QF_def, of k "(Suc j,0)"] pair_lessI1 enum_nxt_ge intro!: less_sets_weaken2 [OF less_sets_Suc_Max]) have less_QF2: "QF k (j,i) \ QF k (j, Suc i)" for j i k by (auto simp: def_wfrec [OF QF_def, of k "(j, Suc i)"] pair_lessI2 enum_nxt_ge intro: less_sets_weaken2 [OF less_sets_Suc_Max] strict_mono_setsD [OF sm_AF]) have less_QF_same: "QF k (j,i') \ QF k (j,i)" if "i' < i" "j \ k" for i' i j k proof (rule strict_mono_setsD [OF less_sets_imp_strict_mono_sets]) show "QF k (j, i) \ QF k (j, Suc i)" for i by (simp add: less_QF2) show "QF k (j, i) \ {}" if "0 < i" for i using that by (simp add: \j \ k\ le_imp_less_Suc) qed (use that in auto) have less_QF_step: "QF k (j-1, i') \ QF k (j,i)" if "0 < j" "j \ k" "i' < m" for j i' i k proof - have less_QF1': "QF k (j - 1, m-1) \ QF k (j,0)" if "j > 0" for j by (metis less_QF1 that Suc_pred One_nat_def) have "QF k (j-1, i') \ QF k (j,0)" proof (cases "i' = m - 1") case True then show ?thesis using less_QF1' \0 < j\ by blast next case False show ?thesis using False that less_sets_trans [OF less_QF_same less_QF1' QF_ne] by auto qed then show ?thesis by (metis QF_ne less_QF_same less_Suc_eq_le less_sets_trans \j \ k\ zero_less_iff_neq_zero) qed have less_QF: "QF k (j',i') \ QF k (j,i)" if j: "j' < j" "j \ k" and i: "i' < m" "i < m" for j' j i' i k using j proof (induction "j-j'" arbitrary: j) case (Suc d) show ?case proof (cases "j' < j - 1") case True then have "QF k (j', i') \ QF k (j - 1, i)" using Suc.hyps Suc.prems(2) by force then show ?thesis by (rule less_sets_trans [OF _ less_QF_step QF_ne]) (use Suc i in auto) next case False then have "j' = j - 1" using \j' < j\ by linarith then show ?thesis using Suc.hyps \j \ k\ less_QF_step i by auto qed qed auto have sm_QF: "strict_mono_sets ({..k} \ {.. {..k} \ {.. {..k} \ {..: "p = (j',i')" "q = (j,i)" "i' < m" "i < m" "j' \ k" "j \ k" using surj_pair [of p] surj_pair [of q] by blast with \p < q\ have "j' < j \ j' = j \ i' < i" by auto then show "QF k p \ QF k q" using \
less_QF less_QF_same by presburger qed then have sm_QF1: "strict_mono_sets {..j. QF k (j,i))" if "i ka" "ka \ k" for ka k i proof - have "{.. {..k}" by (metis lessThan_Suc_atMost lessThan_subset_iff \Suc k \ ka\) then show ?thesis by (simp add: less_QF strict_mono_sets_def subset_iff that) qed have disjoint_QF: "i'=i \ j'=j" if "\ disjnt (QF k (j', i')) (QF k (j,i))" "j' \ k" "j \ k" "i' < m" "i < m" for i' i j' j k using that strict_mono_sets_imp_disjoint [OF sm_QF] by (force simp: pairwise_def) have card_QF: "card (QF k (j,i)) = (if j=0 then \ (DF k i) else del k (if j = k then m - Suc i else i) j)" for i k j proof (cases j) case 0 then show ?thesis by (simp add: AF_def card_image inj_enum_nxt) next case (Suc j') show ?thesis by (cases i; simp add: Suc One_nat_def QF_Suc QF_Suc_Suc card_image inj_enum_nxt) qed have AF_non_Nil: "list_of (AF k i) \ []" for k i by (simp add: AF_ne) have QF_non_Nil: "list_of (QF k (j,i)) \ []" if "j < Suc k" for i j k by (simp add: that) have AF_subset_N: "AF k i \ N" for i k unfolding AF_def image_subset_iff using nxt_subset enumerate_in_set infinite_nxtN \infinite N\ by blast have QF_subset_N: "QF k (j,i) \ N" for i j k proof (induction j) case (Suc j) show ?case by (cases i) (use nxt_subset enumerate_in_set in \(force simp: QF_Suc QF_Suc_Suc)+\) qed (use AF_subset_N in auto) obtain ka k where "k>0" and kka: "k \ ka" "ka \ Suc k" "l = ((ka+k) - 1)" by (metis One_nat_def assms(2) diff_add_inverse form_cases le0 le_refl) then have "ka > 0" using dual_order.strict_trans1 by blast have ka_k_or_Suc: "ka = k \ ka = Suc k" using kka by linarith have lessThan_k: "{..0" for k::nat using that by auto then have sorted_list_of_set_k: "list_of {..0" for k::nat using sorted_list_of_set_insert_remove_cons [of concl: 0 "{0<.. \j i. if j = k then QF k (j, m - Suc i) else QF k (j,i)" have RF_subset_N: "RF j i \ N" if "i0 < k\ by auto have disjoint_RF: "i'=i \ j'=j" if "\ disjnt (RF j' i') (RF j i)" "j' \ k" "j \ k" "i' < m" "i < m" for i' i j' j using disjoint_QF that by (auto simp: RF_def split: if_split_asm dest: disjoint_QF) have sum_card_RF [simp]: "(\j\n. card (RF j i)) = enum (DF k i) n" if "n \ k" "i < m" for i n using that proof (induction n) case 0 then show ?case using DF_ne [of k i] finite_DF [of k i] \k>0\ by (simp add: RF_def AF_def card_image inj_enum_nxt enum_0_eq_Inf_finite) next case (Suc n) then have "enum (DF k i) 0 \ enum (DF k i) n \ enum (DF k i) n \ enum (DF k i) (Suc n)" using sm_enum_DF [of k i] by (metis Suc_leD card_DF dual_order.eq_iff finite_DF finite_enumerate_mono le_imp_less_Suc less_imp_le_nat not_gr_zero) with Suc show ?case by (auto simp: RF_def card_QF del_def) qed have DF_in_N: "enum (DF k i) j \ N" if "j \ k" for i j by (metis DF_N card_DF finite_DF finite_enumerate_in_set le_imp_less_Suc subsetD that) have Inf_DF_N: "\(DF k p) \ N" for k p using DF_N DF_ne Inf_nat_def1 by blast have RF_in_N: "(\j\n. card (RF j i)) \ N" if "n \ k" "i < m" for i n by (auto simp: DF_in_N that) have "ka - 1 \ k" using kka(2) by linarith then have sum_card_RF' [simp]: "(\j0 < ka\ lessThan_Suc_atMost that) have enum_DF_le_iff [simp]: "enum (DF k i) j \ enum (DF k i') j \ i \ i'" (is "?lhs = _") if "j \ k" for i' i j k proof show "i \ i'" if ?lhs proof - have "enum (DF k i) j \ DF k i" by (simp add: card_DF finite_enumerate_in_set finite_DF le_imp_less_Suc \j \ k\) moreover have "enum (DF k i') j \ DF k i'" by (simp add: \j \ k\ card_DF finite_enumerate_in_set finite_DF le_imp_less_Suc that) ultimately have "enum (DF k i') j < enum (DF k i) j" if "i' < i" using sm_DF [of k] by (meson UNIV_I less_sets_def strict_mono_setsD that) then show ?thesis using not_less that by blast qed show ?lhs if "i \ i'" using sm_DF [of k] that \j \ k\ card_DF finite_enumerate_in_set finite_DF le_eq_less_or_eq by (force simp: strict_mono_sets_def less_sets_def finite_enumerate_in_set) qed then have enum_DF_eq_iff[simp]: "enum (DF k i) j = enum (DF k i') j \ i = i'" if "j \ k" for i' i j k by (metis le_antisym order_refl that) have enum_DF_less_iff [simp]: "enum (DF k i) j < enum (DF k i') j \ i < i'" if "j \ k" for i' i j k by (meson enum_DF_le_iff not_less that) have card_AF_sum: "card (AF k i) + (\j\{0<..k > 0\ \k \ ka\ \ka \ Suc k\ by (simp add: lessThan_k RF_0 flip: sum_card_RF') have sorted_list_of_set_iff [simp]: "list_of {0<.. k = 1" if "k>0" for k::nat proof - have "list_of {0<.. {0<.. \ k = 1" using \k > 0\ atLeastSucLessThan_greaterThanLessThan by fastforce finally show ?thesis . qed show thesis \\proof of main result\ proof have inj: "inj_on (\i. list_of (\jjjjj RF 0 x" using AF_ne QF_0 \0 < k\ Inf_nat_def1 \k \ ka\ by (force simp: RF_def) with eq \ka > 0\ obtain j' where "j' < ka" "n \ RF j' y" by blast then show ?thesis using disjoint_QF [of k 0 x j'] n \x < m\ \y < m\ \ka \ Suc k\ \0 < k\ by (force simp: RF_def disjnt_iff simp del: QF_0 split: if_split_asm) qed qed define M where "M \ (\i. list_of (\jk \ ka\ card_image inj) moreover have "M \ WW" by (force simp: M_def WW_def) ultimately show "M \ [WW]\<^bsup>m\<^esup>" by (simp add: nsets_def) have sm_RF: "strict_mono_sets {..j. RF j i)" if "i []" if "j < Suc k" for i j using that by (simp add: RF_def) have less_RF_same: "RF j i' \ RF j i" if "i' < i" "j < k" for i' i j using that by (simp add: less_QF_same RF_def) have less_RF_same_k: "RF k i' \ RF k i" \\reversed version for @{term k}\ if "i < i'" "i' < m" for i' i using that by (simp add: less_QF_same RF_def) show "Form l U \ list.set (inter_scheme l U) \ N" if "U \ [M]\<^bsup>2\<^esup>" for U proof - from that obtain x y where "U = {x,y}" "x \ M" "y \ M" and xy: "(x,y) \ lenlex less_than" by (auto simp: lenlex_nsets_2_eq) let ?R = "\p. list_of \ (\j. RF j p)" obtain p q where x: "x = list_of (\jjx \ M\ \y \ M\ by (auto simp: M_def) then have pq: "pk \ ka\ \ka \ Suc k\ lexl_not_refl [OF irrefl_less_than] by (auto simp: lenlex_def sm_RF sorted_list_of_set_UN_lessThan length_concat sum_sorted_list_of_set_map) moreover have xc: "x = concat (map (?R p) (list_of {..k \ ka\ \ka \ Suc k\ \p < m\ sm_RF) have yc: "y = concat (map (?R q) (list_of {..k \ ka\ \ka \ Suc k\ \q < m\ sm_RF) have enum_DF_AF: "enum (DF k p) (ka - 1) < hd (list_of (AF k p))" for p proof (rule less_setsD [OF DF_AF]) show "enum (DF k p) (ka - 1) \ DF k p" using \ka \ Suc k\ card_DF finite_DF by (auto simp: finite_enumerate_in_set) show "hd (list_of (AF k p)) \ AF k p" using AF_non_Nil finite_AF hd_in_set set_sorted_list_of_set by blast qed have less_RF_RF: "RF n p \ RF n q" if "n < k" for n using that \p by (simp add: less_RF_same) have less_RF_Suc: "RF n q \ RF (Suc n) q" if "n < k" for n using \q < m\ that by (auto simp: RF_def less_QF) have less_RF_k: "RF k q \ RF k p" using \q < m\ less_RF_same_k \p by blast have less_RF_k_ka: "RF (k-1) p \ RF (ka - 1) q" using ka_k_or_Suc less_RF_RF by (metis One_nat_def RF_def \0 < k\ \ka - 1 \ k\ \p < m\ diff_Suc_1 diff_Suc_less less_QF_step) have Inf_DF_eq_enum: "\ (DF k i) = enum (DF k i) 0" for k i by (simp add: Inf_nat_def enumerate_0) have Inf_DF_less: "\ (DF k i') < \ (DF k i)" if "i'x. x \ AF k i \ \ (DF k i') < x" if "i'\i" for i' i k using less_setsD [OF DF_AF] DF_ne that by (metis Inf_DF_less Inf_nat_def1 dual_order.order_iff_strict dual_order.strict_trans) show ?thesis \\The general case requires @{term\k>1\}, necessitating a painful special case\ proof (cases "k=1") case True with kka consider "ka=1" | "ka=2" by linarith then show ?thesis proof cases case 1 define zs where "zs = card (AF 1 p) # list_of (AF 1 p) @ card (AF 1 q) # list_of (AF 1 q)" have zs: "Form_Body ka k x y zs" proof (intro that exI conjI Form_Body.intros) show "x = concat ([list_of (AF k p)])" "y = concat ([list_of (AF k q)])" by (simp_all add: x y 1 lessThan_Suc RF_0) have "AF k p \ insert (\ (DF k q)) (AF k q)" by (metis AF_DF DF_ne Inf_nat_def1 RF_0 \0 < k\ insert_iff less_RF_RF less_sets_def pq(1)) then have "strict_sorted (list_of (AF k p) @ \ (DF k q) # list_of (AF k q))" by (auto simp: strict_sorted_append_iff intro: less_sets_imp_list_less AF_Inf_DF_less) moreover have "\x. x \ AF k q \ \ (DF k p) < x" by (meson AF_Inf_DF_less less_imp_le_nat \p < q\) moreover have "\x. x \ AF 1 p \ \ (DF 1 p) < x" by (meson DF_AF DF_ne Inf_nat_def1 less_setsD) ultimately show "strict_sorted zs" using \p < q\ True Inf_DF_less DF_AF DF_ne by (auto simp: zs_def less_sets_def card_AF AF_Inf_DF_less) qed (auto simp: \k=1\ \ka=1\ zs_def AF_ne \length x < length y\) have zs_N: "list.set zs \ N" using AF_subset_N by (auto simp: zs_def card_AF Inf_DF_N \k=1\) show ?thesis proof have "l = 1" using kka \k=1\ \ka=1\ by auto have "Form (2*1-1) {x,y}" using "1" Form.intros(2) True zs by fastforce then show "Form l U" by (simp add: \U = {x,y}\ \l = 1\ One_nat_def) show "list.set (inter_scheme l U) \ N" using kka zs zs_N \k=1\ Form_Body_imp_inter_scheme by (fastforce simp: \U = {x,y}\) qed next case 2 note True [simp] note 2 [simp] have [simp]: "{0<..<2} = {1::nat}" by auto have enum_DF1_eq: "enum (DF 1 i) 1 = card (AF 1 i) + card (RF 1 i)" if "i < m" for i using card_AF_sum that by (simp add: One_nat_def) have card_RF: "card (RF 1 i) = enum (DF 1 i) 1 - enum (DF 1 i) 0" if "i < m" for i using that by (auto simp: RF_def card_QF del_def) have list_of_AF_RF: "list_of (AF 1 q \ RF 1 q) = list_of (AF 1 q) @ list_of (RF 1 q)" by (metis One_nat_def RF_0 True \0 < k\ finite_RF less_RF_Suc sorted_list_of_set_Un) define zs where "zs = card (AF 1 p) # (card (AF 1 p) + card (RF 1 p)) # list_of (AF 1 p) @ (card (AF 1 q) + card (RF 1 q)) # list_of (AF 1 q) @ list_of (RF 1 q) @ list_of (RF 1 p)" have zs: "Form_Body ka k x y zs" proof (intro that exI conjI Form_Body.intros) have "x = list_of (RF 0 p \ RF 1 p)" by (simp add: x eval_nat_numeral lessThan_Suc RF_0 Un_commute One_nat_def) also have "\ = list_of (RF 0 p) @ list_of (RF 1 p)" using RF_def True \p < m\ less_QF_step by (metis QF_0 RF_0 diff_self_eq_0 finite_RF le_refl sorted_list_of_set_Un zero_less_one) finally show "x = concat ([list_of (AF 1 p),list_of (RF 1 p)])" by (simp add: RF_0) show "y = concat [list_of (RF 1 q \ AF 1 q)]" by (simp add: y eval_nat_numeral lessThan_Suc RF_0 One_nat_def) show zs: "zs = concat [[card (AF 1 p), card (AF 1 p) + card (RF 1 p)], list_of (AF 1 p), [card (AF 1 q) + card (RF 1 q)], list_of (RF 1 q \ AF 1 q)] @ interact [list_of (RF 1 p)] []" using list_of_AF_RF by (simp add: zs_def Un_commute) show "strict_sorted zs" proof (simp add: \p \q \p zs_def strict_sorted_append_iff, intro conjI strip) show "0 < card (RF 1 p)" using \p by (simp add: card_RF card_DF finite_DF) show "card (AF 1 p) < card (AF 1 q) + card (RF 1 q)" using \p \q by (simp add: Inf_DF_less card_AF trans_less_add1) show "card (AF 1 p) < x" if "x \ AF 1 p \ (AF 1 q \ (RF 1 q \ RF 1 p))" for x using that apply (simp add: card_AF) by (metis AF_ne DF_AF DF_ne less_RF_RF less_RF_Suc less_RF_k Inf_nat_def1 One_nat_def RF_0 RF_non_Nil True finite_RF lessI less_setsD less_sets_trans sorted_list_of_set_eq_Nil_iff) show "card (AF 1 p) + card (RF 1 p) < card (AF 1 q) + card (RF 1 q)" using \p < q\ \p < m\ \q < m\ by (metis enum_DF1_eq enum_DF_less_iff le_refl) show "card (AF 1 p) + card (RF 1 p) < x" if "x \ AF 1 p \ (AF 1 q \ (RF 1 q \ RF 1 p))" for x using that \p < m\ apply (simp flip: enum_DF1_eq) by (metis AF_ne DF_AF less_RF_RF less_RF_Suc less_RF_k One_nat_def RF_0 RF_non_Nil Suc_mono True \0 < k\ card_DF finite_enumerate_in_set finite_DF less_setsD less_sets_trans sorted_list_of_set_empty) have "list_of (AF 1 p) < list_of {enum (DF 1 q) 1}" proof (rule less_sets_imp_sorted_list_of_set) show "AF 1 p \ {enum (DF 1 q) 1}" by (metis AF_DF card_DF empty_subsetI finite_DF finite_enumerate_in_set insert_subset less_Suc_eq less_sets_weaken2 pq(1)) qed auto then show "list_of (AF 1 p) < (card (AF 1 q) + card (RF 1 q)) # list_of (AF 1 q) @ list_of (RF 1 q) @ list_of (RF 1 p)" using \q < m\ by (simp add: less_list_def enum_DF1_eq) show "card (AF 1 q) + card (RF 1 q) < x" if "x \ AF 1 q \ (RF 1 q \ RF 1 p)" for x using that \q < m\ apply (simp flip: enum_DF1_eq) by (metis AF_ne DF_AF less_RF_Suc less_RF_k One_nat_def RF_0 RF_non_Nil True card_DF finite_enumerate_in_set finite_DF finite_RF lessI less_setsD less_sets_trans sorted_list_of_set_eq_Nil_iff) have "list_of (AF 1 q) < list_of (RF 1 q)" proof (rule less_sets_imp_sorted_list_of_set) show "AF 1 q \ RF 1 q" by (metis less_RF_Suc One_nat_def RF_0 True \0 < k\) qed auto then show "list_of (AF 1 q) < list_of (RF 1 q) @ list_of (RF 1 p)" using RF_non_Nil by (auto simp: less_list_def) show "list_of (RF 1 q) < list_of (RF 1 p)" proof (rule less_sets_imp_sorted_list_of_set) show "RF 1 q \ RF 1 p" by (metis less_RF_k True) qed auto qed show "[list_of (AF 1 p), list_of (RF 1 p)] \ lists (- {[]})" using RF_non_Nil \0 < k\ by (auto simp: zs_def AF_ne) show "[card (AF 1 q) + card (RF 1 q)] = acc_lengths 0 [list_of (RF 1 q \ AF 1 q)]" using list_of_AF_RF by (auto simp: zs_def AF_ne sup_commute) qed (auto simp: zs_def AF_ne \length x < length y\) have zs_N: "list.set zs \ N" using \p < m\ \q < m\ DF_in_N enum_DF1_eq [symmetric] by (auto simp: zs_def card_AF AF_subset_N RF_subset_N Inf_DF_N) show ?thesis proof have "Form (2*1) {x,y}" by (metis "2" Form.simps Suc_1 True zero_less_one zs) with kka show "Form l U" by (simp add: \U = {x,y}\) show "list.set (inter_scheme l U) \ N" using kka zs zs_N \k=1\ Form_Body_imp_inter_scheme by (fastforce simp: \U = {x, y}\) qed qed next case False then have "k \ 2" "ka \ 2" using kka \k>0\ by auto then have k_minus_1 [simp]: "Suc (k - Suc (Suc 0)) = k - Suc 0" by auto have [simp]: "Suc (k - 2) = k-1" using \k \ 2\ by linarith define PP where "PP \ map (?R p) (list_of {0<.. map (?R q) (list_of {0<.. RF (ka-1) q)])" let ?INT = "interact PP QQ" \\No separate sets A and B as in the text, but instead we treat both cases as once\ have [simp]: "length PP = ka - 1" by (simp add: PP_def) have [simp]: "length QQ = k-1" using \k \ 2\ by (simp add: QQ_def) have PP_n: "PP ! n = list_of (RF (Suc n) p)" if "n < ka-1" for n using that kka by (auto simp: PP_def nth_sorted_list_of_set_greaterThanLessThan) have QQ_n: "QQ ! n = (if n < k-2 then list_of (RF (Suc n) q) else list_of (RF (k-1) q \ RF (ka - 1) q))" if "n < k-1" for n using that kka by (auto simp: QQ_def nth_append nth_sorted_list_of_set_greaterThanLessThan) have QQ_n_same: "QQ ! n = list_of (RF (Suc n) q)" if "n < k-1" "k=ka" for n using that kka Suc_diff_Suc by (fastforce simp: One_nat_def QQ_def nth_append nth_sorted_list_of_set_greaterThanLessThan) have split_nat_interval: "{0<.. 2" for n::nat using that by auto have split_list_interval: "list_of{0<.. 2" for n::nat proof (intro sorted_list_of_set_unique [THEN iffD1] conjI) have "list_of {0<..n \ 2\ in auto) have list_of_RF_Un: "list_of (RF (k-1) q \ RF k q) = list_of (RF (k-1) q) @ list_of (RF k q)" by (metis Suc_diff_1 \0 < k\ finite_RF lessI less_RF_Suc sorted_list_of_set_Un) have card_AF_sum_QQ: "card (AF k q) + sum_list (map length QQ) = (\j RF k q = {}" using less_RF_Suc [of "k-1"] \k > 0\ by (auto simp: less_sets_def) then have "card (RF (k-1) q \ RF k q) = card (RF (k-1) q) + card (RF k q)" by (simp add: card_Un_disjoint) then show ?thesis using \k\2\ \q < m\ apply (simp add: QQ_def True flip: RF_0) apply (simp add: lessThan_k split_nat_interval sum_sorted_list_of_set_map) done next case False with kka have "ka=k" by linarith with \k\2\ show ?thesis by (simp add: QQ_def lessThan_k split_nat_interval sum_sorted_list_of_set_map flip: RF_0) qed define LENS where "LENS \ \i. acc_lengths 0 (list_of (AF k i) # map (?R i) (list_of {0<.. N" if "i < m" for i proof - have eq: "(list_of (AF k i) # map (?R i) (list_of {0<..0 < ka\ sorted_list_of_set_k by auto let ?f = "rec_nat [card (AF k i)] (\n r. r @ [(\j\Suc n. card (RF j i))])" have f: "acc_lengths 0 (map (?R i) (list_of {..v})) = ?f v" for v by (induction v) (auto simp: RF_0 acc_lengths_append sum_sorted_list_of_set_map) have 3: "list.set (?f v) \ N" if "v \ k" for v using that proof (induction v) case 0 have "card (AF k i) \ N" by (metis DF_N DF_ne Inf_nat_def1 card_AF subsetD) with 0 show ?case by simp next case (Suc v) then have "enum (DF k i) (Suc v) \ N" by (metis DF_N card_DF finite_enumerate_in_set finite_DF in_mono le_imp_less_Suc) with Suc \i < m\ show ?case by (simp del: sum.atMost_Suc) qed show ?thesis unfolding LENS_def by (metis 3 Suc_pred' \0 < ka\ \ka - 1 \ k\ eq f lessThan_Suc_atMost) qed define LENS_QQ where "LENS_QQ \ acc_lengths 0 (list_of (AF k q) # QQ)" have LENS_QQ_subset: "list.set LENS_QQ \ list.set (LENS q)" proof (cases "ka = Suc k") case True with \k \ 2\ show ?thesis unfolding QQ_def LENS_QQ_def LENS_def by (auto simp: list_of_RF_Un split_list_interval acc_lengths_append) next case False then have "ka=k" using kka by linarith with \k \ 2\ show ?thesis by (simp add: QQ_def LENS_QQ_def LENS_def split_list_interval) qed have ss_INT: "strict_sorted ?INT" proof (rule strict_sorted_interact_I) fix n assume "n < length QQ" then have n: "n < k-1" by simp have "n = k - 2" if "\ n < k - 2" using n that by linarith moreover have "list_of (RF (Suc (k - 2)) p) < list_of (RF (k-1) q \ RF (ka - 1) q)" by (auto simp: less_sets_imp_sorted_list_of_set less_sets_Un2 less_RF_RF less_RF_k_ka \0 < k\) ultimately show "PP ! n < QQ ! n" using \k \ ka\ n by (auto simp: PP_n QQ_n less_sets_imp_sorted_list_of_set less_RF_RF) next fix n have V: "\Suc n < ka - 1\ \ list_of (RF (Suc n) q) < list_of (RF (Suc (Suc n)) p)" for n by (smt RF_def Suc_leI \ka - 1 \ k\ \q < m\ diff_Suc_1 finite_RF less_QF_step less_le_trans less_sets_imp_sorted_list_of_set nat_neq_iff zero_less_Suc) have "RF (k - 1) q \ RF k p" by (metis One_nat_def RF_non_Nil Suc_pred \0 < k\ finite_RF lessI less_RF_Suc less_RF_k less_sets_trans sorted_list_of_set_eq_Nil_iff) with kka have "RF (k-1) q \ RF (ka - 1) q \ RF k p" by (metis less_RF_k One_nat_def less_sets_Un1 antisym_conv2 diff_Suc_1 le_less_Suc_eq) then have VI: "list_of (RF (k-1) q \ RF (ka - 1) q) < list_of (RF k p)" by (rule less_sets_imp_sorted_list_of_set) auto assume "Suc n < length PP" with \ka \ Suc k\ VI show "QQ ! n < PP ! Suc n" apply (clarsimp simp: PP_n QQ_n V) by (metis One_nat_def Suc_1 Suc_lessI add.right_neutral add_Suc_right diff_Suc_Suc ka_k_or_Suc less_diff_conv) next show "PP \ lists (- {[]})" using RF_non_Nil kka by (clarsimp simp: PP_def) (metis RF_non_Nil less_le_trans) show "QQ \ lists (- {[]})" using RF_non_Nil kka by (clarsimp simp: QQ_def) (metis RF_non_Nil Suc_pred \0 < k\ less_SucI One_nat_def) qed (use kka PP_def QQ_def in auto) then have ss_QQ: "strict_sorted (concat QQ)" using strict_sorted_interact_imp_concat by blast obtain zs where zs: "Form_Body ka k x y zs" and zs_N: "list.set zs \ N" proof (intro that exI conjI Form_Body.intros [OF \length x < length y\]) show "x = concat (list_of (AF k p) # PP)" using \ka > 0\ by (simp add: PP_def RF_0 xc sorted_list_of_set_k) let ?YR = "(map (list_of \ (\j. RF j q)) (list_of {0<..ka - 1 \ k\ less_le_trans) then show "list_of (RF (list_of {0<.. lists (- {[]})" using RF_non_Nil \ka \ Suc k\ by (auto simp: mem_lists_non_Nil) qed auto show "list.set (concat ?YR) = list.set (concat QQ)" using ka_k_or_Suc proof assume "ka = k" then show "list.set (concat (map (list_of \ (\j. RF j q)) (list_of {0<..k\2\ by simp (simp add: split_nat_interval QQ_def) next assume "ka = Suc k" then show "list.set (concat (map (list_of \ (\j. RF j q)) (list_of {0<..k\2\ by simp (auto simp: QQ_def split_nat_interval) qed qed then show "y = concat (list_of (AF k q) # QQ)" using \ka > 0\ by (simp add: RF_0 yc sorted_list_of_set_k) show "list_of (AF k p) # PP \ lists (- {[]})" "list_of (AF k q) # QQ \ lists (- {[]})" using RF_non_Nil kka by (auto simp: AF_ne PP_def QQ_def eq_commute [of "[]"]) show "list.set ((LENS p @ list_of (AF k p) @ LENS_QQ @ list_of (AF k q) @ ?INT)) \ N" using AF_subset_N RF_subset_N LENS_subset_N \p < m\ \q < m\ LENS_QQ_subset by (auto simp: subset_iff PP_def QQ_def) show "length (list_of (AF k p) # PP) = ka" "length (list_of (AF k q) # QQ) = k" using \0 < ka\ \0 < k\ by auto show "LENS p = acc_lengths 0 (list_of (AF k p) # PP)" by (auto simp: LENS_def PP_def) show "strict_sorted (LENS p @ list_of (AF k p) @ LENS_QQ @ list_of (AF k q) @ ?INT)" unfolding strict_sorted_append_iff proof (intro conjI ss_INT) show "LENS p < list_of (AF k p) @ LENS_QQ @ list_of (AF k q) @ ?INT" using AF_non_Nil [of k p] \k \ ka\ \ka \ Suc k\ \p < m\ card_AF_sum enum_DF_AF by (simp add: enum_DF_AF less_list_def card_AF_sum LENS_def sum_sorted_list_of_set_map del: acc_lengths.simps) show "strict_sorted (LENS p)" unfolding LENS_def by (rule strict_sorted_acc_lengths) (use RF_non_Nil AF_non_Nil kka in \auto simp: in_lists_conv_set\) show "strict_sorted LENS_QQ" unfolding LENS_QQ_def QQ_def by (rule strict_sorted_acc_lengths) (use RF_non_Nil AF_non_Nil kka in \auto simp: in_lists_conv_set\) have last_AF_DF: "last (list_of (AF k p)) < \ (DF k q)" using AF_DF [OF \p < q\, of k] AF_non_Nil [of k p] DF_ne [of k q] by (metis Inf_nat_def1 finite_AF last_in_set less_sets_def set_sorted_list_of_set) then show "list_of (AF k p) < LENS_QQ @ list_of (AF k q) @ ?INT" by (simp add: less_list_def card_AF LENS_QQ_def) show "LENS_QQ < list_of (AF k q) @ ?INT" using AF_non_Nil [of k q] \q < m\ card_AF_sum enum_DF_AF card_AF_sum_QQ by (auto simp: less_list_def AF_ne hd_append card_AF_sum LENS_QQ_def) show "list_of (AF k q) < ?INT" proof - have "AF k q \ RF 1 p" using \0 < k\ \p < m\ \q < m\ by (simp add: RF_def less_QF flip: QF_0) then have "last (list_of (AF k q)) < hd (list_of (RF 1 p))" proof (rule less_setsD) show "last (list_of (AF k q)) \ AF k q" using AF_non_Nil finite_AF last_in_set set_sorted_list_of_set by blast show "hd (list_of (RF 1 p)) \ RF 1 p" by (metis One_nat_def RF_non_Nil \0 < k\ finite_RF hd_in_set not_less_eq set_sorted_list_of_set) qed with \k > 0\ \ka \ 2\ RF_non_Nil show ?thesis by (simp add: One_nat_def hd_interact less_list_def sorted_list_of_set_greaterThanLessThan PP_def QQ_def) qed qed auto qed (auto simp: LENS_QQ_def) show ?thesis proof (cases "ka = k") case True then have "l = 2*k-1" by (simp add: kka(3) mult_2) then show ?thesis by (metis One_nat_def Form.intros(2) Form_Body_imp_inter_scheme True \0 < k\ \U = {x, y}\ kka zs zs_N) next case False then have "l = 2*k" using kka by linarith then show ?thesis by (metis One_nat_def False Form.intros(3) Form_Body_imp_inter_scheme \0 < k\ \U = {x, y}\ antisym kka le_SucE zs zs_N) qed qed qed qed qed subsection \Larson's Lemma 3.8\ subsubsection \Primitives needed for the inductive construction of @{term b}\ definition IJ where "IJ \ \k. Sigma {..k} (\j::nat. {.. IJ k \ (\j i. u = (j,i) \ i j\k)" by (auto simp: IJ_def) lemma finite_IJ: "finite (IJ k)" by (auto simp: IJ_def) fun prev where "prev 0 0 = None" | "prev (Suc 0) 0 = None" | "prev (Suc j) 0 = Some (j, j - Suc 0)" | "prev j (Suc i) = Some (j,i)" lemma prev_eq_None_iff: "prev j i = None \ j \ Suc 0 \ i = 0" by (auto simp: le_Suc_eq elim: prev.elims) lemma prev_pair_less: "prev j i = Some ji' \ (ji', (j,i)) \ pair_less" by (auto simp: pair_lessI1 elim: prev.elims) lemma prev_Some_less: "\prev j i = Some (j',i'); i \ j\ \ i' < j'" by (auto elim: prev.elims) lemma prev_maximal: "\prev j i = Some (j',i'); (ji'', (j,i)) \ pair_less; ji'' \ IJ k\ \ (ji'', (j',i')) \ pair_less \ ji'' = (j',i')" by (force simp: IJ_def pair_less_def elim: prev.elims) lemma pair_less_prev: assumes "(u, (j,i)) \ pair_less" "u \ IJ k" shows "prev j i = Some u \ (\x. (u, x) \ pair_less \ prev j i = Some x)" proof (cases "prev j i") case None then show ?thesis using assms by (force simp: prev_eq_None_iff pair_less_def IJ_def split: prod.split) next case (Some a) then show ?thesis by (metis assms prev_maximal prod.exhaust_sel) qed subsubsection \Special primitives for the ordertype proof\ definition USigma :: "'a set set \ ('a set \ 'a set) \ 'a set set" where "USigma \ B \ \X\\. \y \ B X. {insert y X}" definition usplit where "usplit f A \ f (A - {Max A}) (Max A)" lemma USigma_empty [simp]: "USigma {} B = {}" by (auto simp: USigma_def) lemma USigma_iff: assumes "\I j. I \ \ \ I \ J I \ finite I" shows "x \ USigma \ J \ usplit (\I j. I \ \ \ j \ J I \ x = insert j I) x" proof - have [simp]: "\I j. \I \ \; j \ J I\ \ Max (insert j I) = j" by (meson Max_insert2 assms less_imp_le less_sets_def) show ?thesis proof - have \
: "j \ I" if "I \ \" "j \ J I" for I j using that by (metis assms less_irrefl less_sets_def) have "\I\\. \j\J I. x = insert j I" if "x - {Max x} \ \" and "Max x \ J (x - {Max x})" "x \ {}" using that by (metis Max_in assms infinite_remove insert_Diff) then show ?thesis by (auto simp: USigma_def usplit_def \
) qed qed proposition ordertype_append_image_IJ: assumes lenB [simp]: "\i j. i \ \ \ j \ J i \ length (B j) = c" and AB: "\i j. i \ \ \ j \ J i \ A i < B j" and IJ: "\i. i \ \ \ i \ J i \ finite i" and \: "\i. i \ \ \ ordertype (B ` J i) (lenlex less_than) = \" and A: "inj_on A \" shows "ordertype (usplit (\i j. A i @ B j) ` USigma \ J) (lenlex less_than) = \ * ordertype (A ` \) (lenlex less_than)" (is "ordertype ?AB ?R = _ * ?\") proof (cases "\ = {}") case False have "Ord \" using \ False wf_Ord_ordertype by fastforce show ?thesis proof (subst ordertype_eq_iff) define split where "split \ \l::nat list. (take (length l - c) l, (drop (length l - c) l))" have oB: "ordermap (B ` J i) ?R (B j) \ \" if \i \ \\ \j \ J i\ for i j using \ less_TC_iff that by fastforce then show "Ord (\ * ?\)" by (intro \Ord \\ wf_Ord_ordertype Ord_mult; simp) define f where "f \ \u. let (x,y) = split u in let i = inv_into \ A x in \ * ordermap (A`\) ?R x + ordermap (B`J i) ?R y" have inv_into_IA [simp]: "inv_into \ A (A i) = i" if "i \ \" for i by (simp add: A that) show "\f. bij_betw f ?AB (elts (\ * ?\)) \ (\x\?AB. \y\?AB. (f x < f y) = ((x, y) \ ?R))" unfolding bij_betw_def proof (intro exI conjI strip) show "inj_on f ?AB" proof (clarsimp simp: f_def inj_on_def split_def USigma_iff IJ usplit_def) fix x y assume \
: "\ * ordermap (A ` \) ?R (A (x - {Max x})) + ordermap (B ` J (x - {Max x})) ?R (B (Max x)) = \ * ordermap (A ` \) ?R (A (y - {Max y})) + ordermap (B ` J (y - {Max y})) ?R (B (Max y))" and x: "x - {Max x} \ \" and y: "y - {Max y} \ \" and mx: "Max x \ J (x - {Max x})" and "x = insert (Max x) x" and my: "Max y \ J (y - {Max y})" have "ordermap (A`\) ?R (A (x - {Max x})) = ordermap (A`\) ?R (A (y - {Max y}))" and B_eq: "ordermap (B ` J (x - {Max x})) ?R (B (Max x)) = ordermap (B ` J (y - {Max y})) ?R (B (Max y))" using mult_cancellation_lemma [OF \
] oB mx my x y by blast+ then have "A (x - {Max x}) = A (y - {Max y})" using x y by auto then have "x - {Max x} = y - {Max y}" by (metis x y inv_into_IA) then show "A (x - {Max x}) = A (y - {Max y}) \ B (Max x) = B (Max y)" using B_eq mx my by auto qed show "f ` ?AB = elts (\ * ?\)" proof show "f ` ?AB \ elts (\ * ?\)" using \Ord \\ apply (clarsimp simp add: f_def split_def USigma_iff IJ usplit_def) by (metis TC_small \ add_mult_less image_eqI ordermap_in_ordertype trans_llt wf_Ord_ordertype wf_llt) show "elts (\ * ?\) \ f ` ?AB" proof (clarsimp simp: f_def split_def image_iff USigma_iff IJ usplit_def Bex_def elim!: elts_multE split: prod.split) fix \ \ assume \: "\ \ elts \" and \: "\ \ elts ?\" have "\ \ ordermap (A ` \) (lenlex less_than) ` A ` \" by (meson \ ordermap_surj subset_iff) then obtain i where "i \ \" and yv: "\ = ordermap (A`\) ?R (A i)" by blast have "\ \ ordermap (B ` J i) (lenlex less_than) ` B ` J i" by (metis (no_types) \ \ \i \ \\ in_mono ordermap_surj) then obtain j where "j \ J i" and xu: "\ = ordermap (B`J i) ?R (B j)" by blast then have mji: "Max (insert j i) = j" by (meson IJ Max_insert2 \i \ \\ less_imp_le less_sets_def) have [simp]: "i - {j} = i" using IJ \i \ \\ \j \ J i\ less_setsD by fastforce show "\l. (\K. K - {Max K} \ \ \ Max K \ J (K - {Max K}) \ K = insert (Max K) K \ l = A (K - {Max K}) @ B (Max K)) \ \ * \ + \ = \ * ordermap (A ` \) ?R (take (length l - c) l) + ordermap (B ` J (inv_into \ A (take (length l - c) l))) ?R (drop (length l - c) l)" proof (intro conjI exI) let ?ji = "insert j i" show "A i @ B j = A (?ji - {Max ?ji}) @ B (Max ?ji)" by (auto simp: mji) qed (use \i \ \\ \j \ J i\ mji xu yv in auto) qed qed next fix p q assume "p \ ?AB" and "q \ ?AB" then obtain x y where peq: "p = A (x - {Max x}) @ B (Max x)" and qeq: "q = A (y - {Max y}) @ B (Max y)" and x: "x - {Max x} \ \" and y: "y - {Max y} \ \" and mx: "Max x \ J (x - {Max x})" and my: "Max y \ J (y - {Max y})" by (auto simp: USigma_iff IJ usplit_def) let ?mx = "x - {Max x}" let ?my = "y - {Max y}" show "(f p < f q) \ ((p, q) \ ?R)" proof assume "f p < f q" then consider "ordermap (A`\) ?R (A (x - {Max x})) < ordermap (A`\) ?R (A (y - {Max y}))" | "ordermap (A`\) ?R (A (x - {Max x})) = ordermap (A`\) ?R (A (y - {Max y}))" "ordermap (B`J (x - {Max x})) ?R (B (Max x)) < ordermap (B`J (y - {Max y})) ?R (B (Max y))" using x y mx my by (auto dest: mult_cancellation_less simp: f_def split_def peq qeq oB) then have "(A ?mx @ B (Max x), A ?my @ B (Max y)) \ ?R" proof cases case 1 then have "(A ?mx, A ?my) \ ?R" using x y by (force simp: Ord_mem_iff_lt intro: converse_ordermap_mono) then show ?thesis using x y mx my lenB lenlex_append1 by blast next case 2 then have "A ?mx = A ?my" using \?my \ \\ \?mx \ \\ by auto then have eq: "?mx = ?my" by (metis \?my \ \\ \?mx \ \\ inv_into_IA) then have "(B (Max x), B (Max y)) \ ?R" using mx my 2 by (force simp: Ord_mem_iff_lt intro: converse_ordermap_mono) with 2 show ?thesis by (simp add: eq irrefl_less_than) qed then show "(p,q) \ ?R" by (simp add: peq qeq f_def split_def sorted_list_of_set_Un AB) next assume pqR: "(p,q) \ ?R" then have \
: "(A ?mx @ B (Max x), A ?my @ B (Max y)) \ ?R" using peq qeq by blast then consider "(A ?mx, A ?my) \ ?R" | "A ?mx = A ?my \ (B (Max x), B (Max y)) \ ?R" proof (cases "(A ?mx, A ?my) \ ?R") case False have False if "(A ?my, A ?mx) \ ?R" by (metis \?my \ \\ \?mx \ \\ "\
" \(Max y) \ J ?my\ \(Max x) \ J ?mx\ lenB lenlex_append1 omega_sum_1_less order.asym that) then have "A ?mx = A ?my" by (meson False UNIV_I total_llt total_on_def) then show ?thesis using "\
" irrefl_less_than that(2) by auto qed (use that in blast) then have "\ * ordermap (A`\) ?R (A ?mx) + ordermap (B`J ?mx) ?R (B (Max x)) < \ * ordermap (A`\) ?R (A ?my) + ordermap (B`J ?my) ?R (B (Max y))" proof cases case 1 show ?thesis proof (rule add_mult_less_add_mult) show "ordermap (A`\) (lenlex less_than) (A ?mx) < ordermap (A`\) (lenlex less_than) (A ?my)" by (simp add: "1" \?my \ \\ \?mx \ \\ ordermap_mono_less) show "Ord (ordertype (A`\) ?R)" using wf_Ord_ordertype by blast show "ordermap (B ` J ?mx) ?R (B (Max x)) \ elts \" using Ord_less_TC_mem \Ord \\ \?mx \ \\ \(Max x) \ J ?mx\ oB by blast show "ordermap (B ` J ?my) ?R (B (Max y)) \ elts \" using Ord_less_TC_mem \Ord \\ \?my \ \\ \(Max y) \ J ?my\ oB by blast qed (use \?my \ \\ \?mx \ \\ \Ord \\ in auto) next case 2 with \?mx \ \\ show ?thesis using \(Max y) \ J ?my\ \(Max x) \ J ?mx\ ordermap_mono_less by (metis (no_types, opaque_lifting) Kirby.add_less_cancel_left TC_small image_iff inv_into_IA trans_llt wf_llt y) qed then show "f p < f q" using \?my \ \\ \?mx \ \\ \(Max y) \ J ?my\ \(Max x) \ J ?mx\ by (auto simp: peq qeq f_def split_def AB) qed qed qed auto qed auto subsubsection \The final part of 3.8, where two sequences are merged\ inductive merge :: "[nat list list,nat list list,nat list list,nat list list] \ bool" where NullNull: "merge [] [] [] []" | Null: "as \ [] \ merge as [] [concat as] []" | App: "\as1 \ []; bs1 \ []; concat as1 < concat bs1; concat bs1 < concat as2; merge as2 bs2 as bs\ \ merge (as1@as2) (bs1@bs2) (concat as1 # as) (concat bs1 # bs)" inductive_simps Null1 [simp]: "merge [] bs us vs" inductive_simps Null2 [simp]: "merge as [] us vs" lemma merge_single: "\concat as < concat bs; concat as \ []; concat bs \ []\ \ merge as bs [concat as] [concat bs]" using merge.App [of as bs "[]" "[]"] by (fastforce simp: less_list_def) lemma merge_length1_nonempty: assumes "merge as bs us vs" "as \ lists (- {[]})" shows "us \ lists (- {[]})" using assms by induction (auto simp: mem_lists_non_Nil) lemma merge_length2_nonempty: assumes "merge as bs us vs" "bs \ lists (- {[]})" shows "vs \ lists (- {[]})" using assms by induction (auto simp: mem_lists_non_Nil) lemma merge_length1_gt_0: assumes "merge as bs us vs" "as \ []" shows "length us > 0" using assms by induction auto lemma merge_length_le: assumes "merge as bs us vs" shows "length vs \ length us" using assms by induction auto lemma merge_length_le_Suc: assumes "merge as bs us vs" shows "length us \ Suc (length vs)" using assms by induction auto lemma merge_length_less2: assumes "merge as bs us vs" shows "length vs \ length as" using assms proof induction case (App as1 bs1 as2 bs2 as bs) then show ?case using length_greater_0_conv [of as1] by (simp, presburger) qed auto lemma merge_preserves: assumes "merge as bs us vs" shows "concat as = concat us \ concat bs = concat vs" using assms by induction auto lemma merge_interact: assumes "merge as bs us vs" "strict_sorted (concat as)" "strict_sorted (concat bs)" "bs \ lists (- {[]})" shows "strict_sorted (interact us vs)" using assms proof induction case (App as1 bs1 as2 bs2 as bs) then have bs: "concat bs1 < concat bs" "concat bs1 < concat as" and xx: "concat bs1 \ []" using merge_preserves strict_sorted_append_iff by fastforce+ then have "concat bs1 < interact as bs" unfolding less_list_def using App bs by (metis (no_types, lifting) Un_iff concat_append hd_in_set last_in_set merge_preserves set_interact sorted_wrt_append strict_sorted_append_iff) with App show ?case apply (simp add: strict_sorted_append_iff del: concat_eq_Nil_conv) by (metis hd_append2 less_list_def xx) qed auto lemma acc_lengths_merge1: assumes "merge as bs us vs" shows "list.set (acc_lengths k us) \ list.set (acc_lengths k as)" using assms proof (induction arbitrary: k) case (App as1 bs1 as2 bs2 as bs) then show ?case apply (simp add: acc_lengths_append strict_sorted_append_iff length_concat_acc_lengths) by (simp add: le_supI2 length_concat) qed (auto simp: length_concat_acc_lengths) lemma acc_lengths_merge2: assumes "merge as bs us vs" shows "list.set (acc_lengths k vs) \ list.set (acc_lengths k bs)" using assms proof (induction arbitrary: k) case (App as1 bs1 as2 bs2 as bs) then show ?case apply (simp add: acc_lengths_append strict_sorted_append_iff length_concat_acc_lengths) by (simp add: le_supI2 length_concat) qed (auto simp: length_concat_acc_lengths) lemma length_hd_le_concat: assumes "as \ []" shows "length (hd as) \ length (concat as)" by (metis (no_types) add.commute assms concat.simps(2) le_add2 length_append list.exhaust_sel) lemma length_hd_merge2: assumes "merge as bs us vs" shows "length (hd bs) \ length (hd vs)" using assms by induction (auto simp: length_hd_le_concat) lemma merge_less_sets_hd: assumes "merge as bs us vs" "strict_sorted (concat as)" "strict_sorted (concat bs)" "bs \ lists (- {[]})" shows "list.set (hd us) \ list.set (concat vs)" using assms proof induction case (App as1 bs1 as2 bs2 as bs) then have \
: "list.set (concat bs1) \ list.set (concat bs2)" by (force simp: dest: strict_sorted_imp_less_sets) have *: "list.set (concat as1) \ list.set (concat bs1)" using App by (metis concat_append strict_sorted_append_iff strict_sorted_imp_less_sets) then have "list.set (concat as1) \ list.set (concat bs)" using App \
less_sets_trans merge_preserves by (metis List.set_empty append_in_lists_conv le_zero_eq length_0_conv length_concat_ge) with * App.hyps show ?case by (fastforce simp: less_sets_UN1 less_sets_UN2 less_sets_Un2) qed auto lemma set_takeWhile: assumes "strict_sorted (concat as)" "as \ lists (- {[]})" shows "list.set (takeWhile (\x. x < y) as) = {x \ list.set as. x < y}" using assms proof (induction as) case (Cons a as) have "a < y" if a: "a < concat as" "strict_sorted a" "strict_sorted (concat as)" "x < y" "x \ []" "x \ list.set as" for x proof - have "last x \ list.set (concat as)" using set_concat that(5) that(6) by fastforce then have "last a < hd (concat as)" using Cons.prems that by (auto simp: less_list_def) also have "\ \ hd y" if "y \ []" using that a by (meson \last x \ list.set (concat as)\ dual_order.strict_trans less_list_def not_le sorted_hd_le strict_sorted_imp_sorted) finally show ?thesis by (simp add: less_list_def) qed then show ?case using Cons by (auto simp: strict_sorted_append_iff) qed auto proposition merge_exists: assumes "strict_sorted (concat as)" "strict_sorted (concat bs)" "as \ lists (- {[]})" "bs \ lists (- {[]})" "hd as < hd bs" "as \ []" "bs \ []" and disj: "\a b. \a \ list.set as; b \ list.set bs\ \ a bus vs. merge as bs us vs" using assms proof (induction "length as + length bs" arbitrary: as bs rule: less_induct) case (less as bs) obtain as1 as2 bs1 bs2 where A: "as1 \ []" "bs1 \ []" "concat as1 < concat bs1" "concat bs1 < concat as2" and B: "as = as1@as2" "bs = bs1@bs2" and C: "bs2 = [] \ (as2 \ [] \ hd as2 < hd bs2)" proof define as1 where "as1 \ takeWhile (\x. x < hd bs) as" define as2 where "as2 \ dropWhile (\x. x < hd bs) as" define bs1 where "bs1 \ if as2=[] then bs else takeWhile (\x. x < hd as2) bs" define bs2 where "bs2 \ if as2=[] then [] else dropWhile (\x. x < hd as2) bs" have as1: "as1 = takeWhile (\x. last x < hd (hd bs)) as" using less.prems by (auto simp: as1_def less_list_def cong: takeWhile_cong) have as2: "as2 = dropWhile (\x. last x < hd (hd bs)) as" using less.prems by (auto simp: as2_def less_list_def cong: dropWhile_cong) have hd_as2: "as2 \ [] \ \ hd as2 < hd bs" using as2_def hd_dropWhile by metis have hd_bs2: "bs2 \ [] \ \ hd bs2 < hd as2" using bs2_def hd_dropWhile by metis show "as1 \ []" by (simp add: as1_def less.prems takeWhile_eq_Nil_iff) show "bs1 \ []" by (metis as2 bs1_def hd_as2 hd_in_set less.prems(7) less.prems(8) set_dropWhileD takeWhile_eq_Nil_iff) show "bs2 = [] \ (as2 \ [] \ hd as2 < hd bs2)" by (metis as2_def bs2_def hd_bs2 less.prems(8) list.set_sel(1) set_dropWhileD) have AB: "list.set A \ list.set B" if "A \ list.set as1" "B \ list.set bs" for A B proof - have "A \ list.set as" using that by (metis as1 set_takeWhileD) then have "sorted A" by (metis concat.simps(2) concat_append less.prems(1) sorted_append split_list_last strict_sorted_imp_sorted) moreover have "sorted (hd bs)" by (metis concat.simps(2) hd_Cons_tl less.prems(2) less.prems(7) strict_sorted_append_iff strict_sorted_imp_sorted) ultimately show ?thesis using that less.prems apply (clarsimp simp add: as1_def set_takeWhile less_list_iff_less_sets less_sets_def) by (metis (full_types) UN_I hd_concat less_le_trans list.set_sel(1) set_concat sorted_hd_le strict_sorted_imp_sorted) qed show "as = as1@as2" by (simp add: as1_def as2_def) show "bs = bs1@bs2" by (simp add: bs1_def bs2_def) have "list.set (concat as1) \ list.set (concat bs1)" using AB set_takeWhileD by (fastforce simp: as1_def bs1_def less_sets_UN1 less_sets_UN2) then show "concat as1 < concat bs1" by (rule less_sets_imp_list_less) have "list.set (concat bs1) \ list.set (concat as2)" if "as2 \ []" proof (clarsimp simp add: bs1_def less_sets_UN1 less_sets_UN2 set_takeWhile less.prems) fix A B assume "A \ list.set as2" "B \ list.set bs" "B < hd as2" with that show "list.set B \ list.set A" using hd_as2 less.prems(1,2) apply (clarsimp simp add: less_sets_def less_list_def) apply (auto simp: as2_def) apply (simp flip: as2_def) by (smt (verit, ccfv_SIG) UN_I \as = as1 @ as2\ concat.simps(2) concat_append hd_concat in_set_conv_decomp_first le_less_trans less_le_trans set_concat sorted_append sorted_hd_le sorted_le_last strict_sorted_imp_sorted that) qed then show "concat bs1 < concat as2" by (simp add: bs1_def less_sets_imp_list_less) qed obtain cs ds where "merge as2 bs2 cs ds" proof (cases "as2 = [] \ bs2 = []") case True then show thesis using that C NullNull Null by metis next have \: "length as2 + length bs2 < length as + length bs" by (simp add: A B) case False moreover have "strict_sorted (concat as2)" "strict_sorted (concat bs2)" "as2 \ lists (- {[]})" "bs2 \ lists (- {[]})" "\a b. \a \ list.set as2; b \ list.set bs2\ \ a < b \ b < a" using B less.prems strict_sorted_append_iff by auto ultimately show ?thesis using C less.hyps [OF \] False that by force qed then obtain cs where "merge (as1 @ as2) (bs1 @ bs2) (concat as1 # cs) (concat bs1 # ds)" using A merge.App by blast then show ?case using B by blast qed subsubsection \Actual proof of Larson's Lemma 3.8\ proposition lemma_3_8: assumes "infinite N" obtains X where "X \ WW" "ordertype X (lenlex less_than) = \\\" "\u. u \ [X]\<^bsup>2\<^esup> \ \l. Form l u \ (l > 0 \ [enum N l] < inter_scheme l u \ List.set (inter_scheme l u) \ N)" proof - let ?LL = "lenlex less_than" define bf where "bf \ \M q. wfrec pair_less (\f (j,i). let R = (case prev j i of None \ M | Some u \ snd (f u)) in grab R (q j i))" have bf_rec: "bf M q (j,i) = (let R = (case prev j i of None \ M | Some u \ snd (bf M q u)) in grab R (q j i))" for M q j i by (subst (1) bf_def) (simp add: Let_def wfrec bf_def cut_apply prev_pair_less cong: conj_cong split: option.split) have "infinite (snd (bf M q u)) = infinite M \ fst (bf M q u) \ M \ snd (bf M q u) \ M" for M q u using wf_pair_less proof (induction u rule: wf_induct_rule) case (less u) then show ?case proof (cases u) case (Pair j i) with less.IH prev_pair_less show ?thesis apply (simp add: bf_rec [of M q j i] split: option.split) using fst_grab_subset snd_grab_subset by blast qed qed then have infinite_bf [simp]: "infinite (snd (bf M q u)) = infinite M" and bf_subset: "fst (bf M q u) \ M \ snd (bf M q u) \ M" for M q u by auto have bf_less_sets: "fst (bf M q ij) \ snd (bf M q ij)" if "infinite M" for M q ij using wf_pair_less proof (induction ij rule: wf_induct_rule) case (less u) then show ?case proof (cases u) case (Pair j i) with less_sets_grab show ?thesis by (simp add: bf_rec [of M q j i] less.IH prev_pair_less that split: option.split) qed qed have card_fst_bf: "finite (fst (bf M q (j,i))) \ card (fst (bf M q (j,i))) = q j i" if "infinite M" for M q j i by (simp add: that bf_rec [of M q j i] split: option.split) have bf_cong: "bf M q u = bf M q' u" if "snd u \ fst u" and eq: "\y x. \x\y; y\fst u\ \ q' y x = q y x" for M q q' u using wf_pair_less that proof (induction u rule: wf_induct_rule) case (less u) show ?case proof (cases u) case (Pair j i) with less.prems show ?thesis proof (clarsimp simp add: bf_rec [of M _ j i] split: option.split) fix j' i' assume *: "prev j i = Some (j',i')" then have **: "((j', i'), u) \ pair_less" by (simp add: Pair prev_pair_less) moreover have "i' < j'" using Pair less.prems by (simp add: prev_Some_less [OF *]) moreover have "\x y. \x \ y; y \ j'\ \ q' y x = q y x" using ** less.prems by (auto simp: pair_less_def Pair) ultimately show "grab (snd (bf M q (j',i'))) (q j i) = grab (snd (bf M q' (j',i'))) (q j i)" using less.IH by auto qed qed qed define ediff where "ediff \ \D:: nat \ nat set. \j i. enum (D j) (Suc i) - enum (D j) i" define F where "F \ \l (dl,a0::nat set,b0::nat \ nat \ nat set,M). let (d,Md) = grab (nxt M (enum N (Suc (2 * Suc l)))) (Suc l) in let (a,Ma) = grab Md (Min d) in let Gb = bf Ma (ediff (dl(l := d))) in let dl' = dl(l := d) in (dl', a, fst \ Gb, snd (Gb(l, l-1)))" define DF where "DF \ rec_nat (\i\{..<0}. {}, {}, \p. {}, N) F" have DF_simps: "DF 0 = (\i\{..<0}. {}, {}, \p. {}, N)" "DF (Suc l) = F l (DF l)" for l by (auto simp: DF_def) note cut_apply [simp] have inf [rule_format]: "\dl al bl L. DF l = (dl,al,bl,L) \ infinite L" for l by (induction l) (auto simp: DF_simps F_def Let_def grab_eqD infinite_nxtN assms split: prod.split) define \ where "\ \ \(dl, a, b, M). \l::nat. dl l \ a \ card a > 0 \ (\j\l. card (dl j) = Suc j) \ a \ \(range b) \ range b \ Collect finite \ a \ N \ \(range b) \ N \ infinite M \ b(l,l-1) \ M \ M \ N" have \_DF: "\ (DF (Suc l)) l" for l proof (induction l) case 0 show ?case using assms apply (clarsimp simp add: bf_rec F_def DF_simps \_def split: prod.split) apply (drule grab_eqD, blast dest: grab_eqD infinite_nxtN)+ apply (auto simp: less_sets_UN2 less_sets_grab card_fst_bf elim!: less_sets_weaken2) apply (metis card_1_singleton_iff Min_singleton greaterThan_iff insertI1 le0 nxt_subset_greaterThan subsetD) using nxt_subset snd_grab_subset bf_subset by blast+ next case (Suc l) then show ?case using assms unfolding Let_def DF_simps(2)[of "Suc l"] F_def \_def apply (clarsimp simp add: bf_rec DF_simps split: prod.split) apply (drule grab_eqD, metis grab_eqD infinite_nxtN)+ apply (safe, simp_all add: less_sets_UN2 less_sets_grab card_fst_bf card_Suc_eq_finite) apply (meson less_sets_weaken2) apply (metis Min_in gr0I greaterThan_iff insert_not_empty le_inf_iff less_asym nxt_def subsetD) apply (meson bf_subset less_sets_weaken2) apply (meson nxt_subset subset_eq) apply (meson bf_subset nxt_subset subset_eq) using bf_rec infinite_bf apply force using bf_less_sets bf_rec apply force by (metis bf_rec bf_subset nxt_subset subsetD) qed define d where "d \ \k. let (dk,ak,bk,M) = DF(Suc k) in dk k" define a where "a \ \k. let (dk,ak,bk,M) = DF(Suc k) in ak" define b where "b \ \k. let (dk,ak,bk,M) = DF(Suc k) in bk" define M where "M \ \k. let (dk,ak,bk,M) = DF k in M" have infinite_M [simp]: "infinite (M k)" for k by (auto simp: M_def inf split: prod.split) have M_Suc_subset: "M (Suc k) \ M k" for k apply (clarsimp simp add: Let_def M_def F_def DF_simps split: prod.split) apply (drule grab_eqD, blast dest: infinite_nxtN local.inf)+ using bf_subset nxt_subset by blast have Inf_M_Suc_ge: "Inf (M k) \ Inf (M (Suc k))" for k by (simp add: M_Suc_subset cInf_superset_mono infinite_imp_nonempty) have Inf_M_telescoping: "{Inf (M k)..} \ {Inf (M k')..}" if k': "k'\k" for k k' using that Inf_nat_def1 infinite_M unfolding Inf_nat_def atLeast_subset_iff by (metis M_Suc_subset finite.emptyI le_less_linear lift_Suc_antimono_le not_less_Least subsetD) have d_eq: "d k = fst (grab (nxt (M k) (enum N (Suc (2 * Suc k)))) (Suc k))" for k by (simp add: d_def M_def Let_def DF_simps F_def split: prod.split) then have finite_d [simp]: "finite (d k)" for k by simp then have d_ne [simp]: "d k \ {}" for k by (metis card.empty card_grab d_eq infinite_M infinite_nxtN nat.distinct(1)) have a_eq: "\M. a k = fst (grab M (Min (d k))) \ infinite M" for k apply (simp add: a_def d_def M_def Let_def DF_simps F_def split: prod.split) by (metis fst_conv grab_eqD infinite_nxtN local.inf) then have card_a: "card (a k) = Inf (d k)" for k by (metis cInf_eq_Min card_grab d_ne finite_d) have d_eq_dl: "d k = dl k" if "(dl,a,b,P) = DF l" "k < l" for k l dl a b P using that by (induction l arbitrary: dl a b P) (simp_all add: d_def DF_simps F_def Let_def split: prod.split_asm prod.split) have card_d [simp]: "card (d k) = Suc k" for k by (auto simp: d_eq infinite_nxtN) have d_ne [simp]: "d j \ {}" and a_ne [simp]: "a j \ {}" and finite_d [simp]: "finite (d j)" and finite_a [simp]: "finite (a j)" for j using \_DF [of "j"] by (auto simp: \_def a_def d_def card_gt_0_iff split: prod.split_asm) have da: "d k \ a k" for k using \_DF [of "k"] by (simp add: \_def a_def d_def split: prod.split_asm) have ab_same: "a k \ \(range(b k))" for k using \_DF [of "k"] by (simp add: \_def a_def b_def M_def split: prod.split_asm) have snd_bf_subset: "snd (bf M r (j,i)) \ snd (bf M r (j',i'))" if ji: "((j',i'), (j,i)) \ pair_less" "(j',i') \ IJ k" for M r k j i j' i' using wf_pair_less ji proof (induction rule: wf_induct_rule [where a= "(j,i)"]) case (less u) show ?case proof (cases u) case (Pair j i) then consider "prev j i = Some (j', i')" | x where "((j', i'), x) \ pair_less" "prev j i = Some x" using less.prems pair_less_prev by blast then show ?thesis proof cases case 2 with less.IH show ?thesis unfolding bf_rec Pair by (metis in_mono option.simps(5) prev_pair_less snd_grab_subset subsetI that(2)) qed (simp add: Pair bf_rec snd_grab_subset) qed qed have less_bf: "fst (bf M r (j',i')) \ fst (bf M r (j,i))" if ji: "((j',i'), (j,i)) \ pair_less" "(j',i') \ IJ k" and "infinite M" for M r k j i j' i' proof - consider "prev j i = Some (j', i')" | j'' i'' where "((j', i'), (j'',i'')) \ pair_less" "prev j i = Some (j'',i'')" by (metis pair_less_prev ji prod.exhaust_sel) then show ?thesis proof cases case 1 then show ?thesis using bf_less_sets bf_rec less_sets_fst_grab \infinite M\ by force next case 2 then have "fst (bf M r (j',i')) \ snd (bf M r (j'',i''))" by (meson bf_less_sets snd_bf_subset less_sets_weaken2 that) with 2 show ?thesis using bf_rec bf_subset less_sets_fst_grab \infinite M\ by auto qed qed have aM: "a k \ M (Suc k)" for k apply (clarsimp simp add: a_def M_def DF_simps F_def Let_def split: prod.split) by (meson bf_subset grab_eqD infinite_nxtN less_sets_weaken2 local.inf) then have "a k \ a (Suc k)" for k by (metis IntE card_d card.empty d_eq da fst_grab_subset less_sets_trans less_sets_weaken2 nat.distinct(1) nxt_def subsetI) then have aa: "a j \ a k" if "j b k (j,i)" if "k'\k" for k k' j i by (metis a_ne ab_same le_less less_sets_UN2 less_sets_trans rangeI that) have db: "d j \ b k (j,i)" if "j\k" for k j i by (meson a_ne ab da less_sets_trans that) have bMkk: "b k (k,k-1) \ M (Suc k)" for k using \_DF [of k] by (simp add: \_def b_def d_def M_def split: prod.split_asm) have b: "\P \ M k. infinite P \ (\j i. i\j \ j\k \ b k (j,i) = fst (bf P (ediff d) (j,i)))" for k proof (clarsimp simp: b_def DF_simps F_def Let_def split: prod.split) fix a a' d' dl bb P M' M'' assume gr: "grab M'' (Min d') = (a', M')" "grab (nxt P (enum N (Suc (Suc (Suc (2 * k)))))) (Suc k) = (d', M'')" and DF: "DF k = (dl, a, bb, P)" have deq: "d j = (if j = k then d' else dl j)" if "j\k" for j proof (cases "j < k") case True then show ?thesis by (metis DF d_eq_dl less_not_refl) next case False then show ?thesis using that DF gr by (auto simp: d_def DF_simps F_def Let_def split: prod.split) qed have "M' \ P" by (metis gr in_mono nxt_subset snd_conv snd_grab_subset subsetI) also have "P \ M k" using DF by (simp add: M_def) finally have "M' \ M k" . moreover have "infinite M'" using DF by (metis (mono_tags) finite_grab_iff gr infinite_nxtN local.inf snd_conv) moreover have "ediff (dl(k := d')) j i = ediff d j i" if "j\k" for j i by (simp add: deq that ediff_def) then have "bf M' (ediff (dl(k := d'))) (j,i) = bf M' (ediff d) (j,i)" if "i \ j" "j\k" for j i using bf_cong that by fastforce ultimately show "\P\M k. infinite P \ (\j i. i \ j \ j \ k \ fst (bf M' (ediff (dl(k := d'))) (j,i)) = fst (bf P (ediff d) (j,i)))" by auto qed have card_b: "card (b k (j,i)) = enum (d j) (Suc i) - enum (d j) i" if "j\k" for k j i \\there's a short proof of this from the previous result but it would need @{term"i\j"}\ proof (clarsimp simp: b_def DF_simps F_def Let_def split: prod.split) fix dl and a a' d':: "nat set" and bb M M' M'' assume gr: "grab M'' (Min d') = (a', M')" "grab (nxt M (enum N (Suc (Suc (Suc (2 * k)))))) (Suc k) = (d',M'')" and DF: "DF k = (dl, a, bb, M)" have "d j = (if j = k then d' else dl j)" proof (cases "j < k") case True then show ?thesis by (metis DF d_eq_dl less_not_refl) next case False then show ?thesis using that DF gr by (auto simp: d_def DF_simps F_def Let_def split: prod.split) qed then show "card (fst (bf M' (ediff (dl(k := d'))) (j,i))) = enum (d j) (Suc i) - enum (d j) i" using DF gr card_fst_bf grab_eqD infinite_nxtN local.inf ediff_def by auto qed have card_b_pos: "card (b k (j,i)) > 0" if "i < j" "j\k" for k j i by (simp add: card_b that finite_enumerate_step) have b_ne [simp]: "b k (j,i) \ {}" if "i < j" "j\k" for k j i using card_b_pos [OF that] less_imp_neq by fastforce+ have card_b_finite [simp]: "finite (b k u)" for k u using \_DF [of k] by (fastforce simp: \_def b_def) have bM: "b k (j,i) \ M (Suc k)" if "ik" for i j k proof - obtain M' where "M' \ M k" "infinite M'" and bk: "\j i. i\j \ j\k \ b k (j,i) = fst (bf M' (ediff d) (j,i))" using b by (metis (no_types, lifting)) show ?thesis proof (cases "j=k \ i = k-1") case False show ?thesis proof (rule less_sets_trans [OF _ bMkk]) show "b k (j,i) \ b k (k, k-1)" using that \infinite M'\ False by (force simp: bk pair_less_def IJ_def intro: less_bf) show "b k (k, k-1) \ {}" using b_ne that by auto qed qed (use bMkk in auto) qed have b_InfM: "\ (range (b k)) \ {\(M k)..}" for k proof (clarsimp simp add: \_def b_def M_def DF_simps F_def Let_def split: prod.split) fix r dl :: "nat \ nat set" and a b and d' a' M'' M' P and x j' i' :: nat assume gr: "grab M'' (Min d') = (a', M')" "grab (nxt P (enum N (Suc (Suc (Suc (2 * k)))))) (Suc k) = (d', M'')" and DF: "DF k = (dl, a, b, P)" and x: "x \ fst (bf M' (ediff (dl(k := d'))) (j', i'))" have "infinite P" using DF local.inf by blast then have "M' \ P" by (meson gr grab_eqD infinite_nxtN nxt_subset order.trans) with bf_subset show "\ P \ x" using Inf_nat_def x le_less_linear not_less_Least by fastforce qed have b_Inf_M_Suc: "b k (j,i) \ {Inf(M (Suc k))}" if "ik" for k j i using bMkk [of k] that by (metis Inf_nat_def1 bM finite.emptyI infinite_M less_setsD less_sets_singleton2) have bb_same: "b k (j',i') \ b k (j,i)" if "((j',i'), (j,i)) \ pair_less" "(j',i') \ IJ k" for k j i j' i' using that unfolding b_def DF_simps F_def Let_def by (auto simp: less_bf grab_eqD infinite_nxtN local.inf split: prod.split) have bb: "b k' (j',i') \ b k (j,i)" if j: "i' < j'" "j'\k'" and k: "k' {Inf(M (Suc k'))}" using Suc_lessD b_Inf_M_Suc nat_less_le j by blast show "b k (j,i) \ {Inf(M (Suc k'))..}" by (meson Inf_M_telescoping Suc_leI UnionI b_InfM rangeI subset_eq k) qed have M_subset_N: "M k \ N" for k proof (cases k) case (Suc k') with \_DF [of k'] show ?thesis by (auto simp: M_def Let_def \_def split: prod.split) qed (auto simp: M_def DF_simps) have a_subset_N: "a k \ N" for k using \_DF [of k] by (simp add: a_def \_def split: prod.split prod.split_asm) have d_subset_N: "d k \ N" for k using M_subset_N [of k] d_eq fst_grab_subset nxt_subset by blast have b_subset_N: "b k (j,i) \ N" for k j i using \_DF [of k] by (force simp: b_def \_def) define \:: "[nat,nat] \ nat set set" where "\ \ \j0 j. nsets {j0<..} j" have \_finite: "finite K" and \_card: "card K = j" if "K \ \ j0 j" for K j0 j using that by (auto simp add: \_def nsets_def) have \_enum: "j0 < enum K i" if "K \ \ j0 j" "i < card K" for K j0 j i using that by (auto simp: \_def nsets_def finite_enumerate_in_set subset_eq) have \_0 [simp]: "\ k 0 = {{}}" for k by (auto simp: \_def) have \_Suc: "\ j0 (Suc j) = USigma (\ j0 j) (\K. {Max (insert j0 K)<..})" (is "?lhs = ?rhs") for j j0 proof show "\ j0 (Suc j) \ USigma (\ j0 j) (\K. {Max (insert j0 K)<..})" unfolding \_def nsets_def USigma_def proof clarsimp fix K assume K: "K \ {j0<..}" "finite K" "card K = Suc j" then have "Max K \ K" by (metis Max_in card_0_eq nat.distinct(1)) then obtain i where "Max (insert j0 (K - {Max K})) < i" "K = insert i (K - {Max K})" using K by (simp add: subset_iff) (metis DiffE Max.coboundedI insertCI insert_Diff le_neq_implies_less) then show "\L\{j0<..}. finite L \ card L = j \ (\i\{Max (insert j0 L)<..}. K = insert i L)" using K by (metis \Max K \ K\ card_Diff_singleton_if diff_Suc_1 finite_Diff greaterThan_iff insert_subset) qed show "?rhs \ \ j0 (Suc j)" by (force simp: \_def nsets_def USigma_def) qed define BB where "BB \ \j0 j K. list_of (a j0 \ (\i \j. BB j j ` \ j j" have less_list_of: "BB j i K < list_of (b l (j,i))" if K: "K \ \ j i" "\j\K. j < l" and "i \ j" "j \ l" for j i K l unfolding BB_def proof (rule less_sets_imp_sorted_list_of_set) have "\i. i < card K \ b (enum K i) (j,i) \ b l (j, card K)" using that by (metis \_card \_enum \_finite bb finite_enumerate_in_set nat_less_le less_le_trans) then show "a j \ (\i b l (j,i)" using that unfolding \_def nsets_def by (auto simp: less_sets_Un1 less_sets_UN1 ab finite_enumerate_in_set subset_eq) qed auto have BB_Suc: "BB j0 (Suc j) K = usplit (\L k. BB j0 j L @ list_of (b k (j0, j))) K" if j: "j \ j0" and K: "K \ \ j0 (Suc j)" for j0 j K \\towards the ordertype proof\ proof - have Kj: "K \ {j0<..}" and [simp]: "finite K" and cardK: "card K = Suc j" using K by (auto simp: \_def nsets_def) have KMK: "K - {Max K} \ \ j0 j" using that by (simp add: \_Suc USigma_iff \_finite less_sets_def usplit_def) have "j0 < Max K" by (metis Kj Max_in cardK card_gt_0_iff greaterThan_iff subsetD zero_less_Suc) have MaxK: "Max K = enum K j" proof (rule Max_eqI) fix k assume "k \ K" with K cardK show "k \ enum K j" by (metis \finite K\ finite_enumerate_Ex finite_enumerate_mono_iff leI lessI not_less_eq) qed (auto simp: cardK finite_enumerate_in_set) have ene: "i enum (K - {enum K j}) i = enum K i" for i using finite_enumerate_Diff_singleton [OF \finite K\] by (simp add: cardK) have "BB j0 (Suc j) K = list_of ((a j0 \ (\x b (enum K j) (j0, j))" by (simp add: BB_def lessThan_Suc Un_ac) also have "\ = list_of ((a j0 \ (\i b (enum K j) (j0, j)" if "i enum K i" using that K by (metis \_enum cardK less_SucI less_imp_le_nat) show "enum K i < enum K j" by (simp add: cardK finite_enumerate_mono that) qed moreover have "a j0 \ b (enum K j) (j0, j)" using MaxK \j0 < Max K\ ab by auto ultimately show "a j0 \ (\x b (enum K j) (j0, j)" by (simp add: less_sets_Un1 less_sets_UN1) qed (auto simp: finite_UnI) also have "\ = BB j0 j (K - {Max K}) @ list_of (b (Max K) (j0, j))" by (simp add: BB_def MaxK ene) also have "\ = usplit (\L k. BB j0 j L @ list_of (b k (j0, j))) K" by (simp add: usplit_def) finally show ?thesis . qed have enum_d_0: "enum (d j) 0 = Inf (d j)" for j using enum_0_eq_Inf_finite by auto have Inf_b_less: "\(b k' (j',i')) < \(b k (j,i))" if j: "i' < j'" "i < j" "j'\k'" "j\k" and k: "k' (b k (k, k-1)) \ k-1" for k proof (induction k) case (Suc k) show ?case proof (cases "k=0") case False then have "\ (b k (k, k - 1)) < \ (b (Suc k) (Suc k, k))" using Inf_b_less by auto with Suc show ?thesis by simp qed auto qed auto have b_ge: "\ (b k (j,i)) \ k-1" if "k \ j" "j > i" for k j i proof - have "\ Suc (\ (b k (j, i))) < k" by (metis (no_types) Inf_b_less Suc_leI b_ge_k diff_Suc_1 lessI not_less that) then show ?thesis by simp qed have hd_b: "hd (list_of (b k (j,i))) = \ (b k (j,i))" if "i < j" "j \ k" for k j i using that by (simp add: hd_list_of cInf_eq_Min) have b_disjoint_less: "b (enum K i) (j0, i) \ b (enum K i') (j0, i') = {}" if K: "K \ {j0<..}" "finite K" "card K \ j0" "i' < j" "i < i'" "j \ j0" for i i' j j0 K proof (intro bb less_sets_imp_disjnt [unfolded disjnt_def]) show "i < j0" using that by linarith then show "j0 \ enum K i" by (meson K finite_enumerate_in_set greaterThan_iff less_imp_le_nat less_le_trans subsetD) show "enum K i < enum K i'" using K \j \ j0\ that by auto qed have b_disjoint: "b (enum K i) (j0, i) \ b (enum K i') (j0, i') = {}" if K: "K \ {j0<..}" "finite K" "card K \ j0" "i < j" "i' < j" "i \ i'" "j \ j0" for i i' j j0 K using that b_disjoint_less inf_commute neq_iff by metis have ot\: "ordertype ((\k. list_of (b k (j,i))) ` {Max (insert j K)<..}) ?LL = \" (is "?lhs = _") if K: "K \ \ j i" "j > i" for j i K proof - have Sucj: "Suc (Max (insert j K)) \ j" using \_finite that(1) le_Suc_eq by auto let ?N = "{Inf(b k (j,i))| k. Max (insert j K) < k}" have infN: "infinite ?N" proof (clarsimp simp add: infinite_nat_iff_unbounded_le) fix m show "\n\m. \k. n = \ (b k (j,i)) \ Max (insert j K) < k" using b_ge \j > i\ Sucj by (metis (no_types, lifting) diff_Suc_1 le_SucI le_trans less_Suc_eq_le nat_le_linear) qed have [simp]: "Max (insert j K) < k \ j < k \ (\a\K. a < k)" for k using that by (auto simp: \_finite) have "?lhs = ordertype ?N less_than" proof (intro ordertype_eqI strip) have "list_of (b k (j,i)) = list_of (b k' (j,i))" if "j \ k" "j \ k'" "hd (list_of (b k (j,i))) = hd (list_of (b k' (j,i)))" for k k' by (metis Inf_b_less \i < j\ hd_b nat_less_le not_le that) moreover have "\k' j' i'. hd (list_of (b k (j,i))) = \ (b k' (j', i')) \ i' < j' \ j' \ k'" if "j \ k" for k using that \i < j\ hd_b less_imp_le_nat by blast moreover have "\k'. hd (list_of (b k (j,i))) = \ (b k' (j,i)) \ j < k' \ (\a\K. a < k')" if "j < k" "\a\K. a < k" for k using that K hd_b less_imp_le_nat by blast moreover have "\ (b k (j,i)) \ hd ` (\k. list_of (b k (j,i))) ` {Max (insert j K)<..}" if "j < k" "\a\K. a < k" for k using that K by (auto simp: hd_b image_iff) ultimately show "bij_betw hd ((\k. list_of (b k (j,i))) ` {Max (insert j K)<..}) {\ (b k (j,i)) |k. Max (insert j K) < k}" by (auto simp: bij_betw_def inj_on_def) next fix ms ns assume "ms \ (\k. list_of (b k (j,i))) ` {Max (insert j K)<..}" and "ns \ (\k. list_of (b k (j,i))) ` {Max (insert j K)<..}" with that obtain k k' where ms: "ms = list_of (b k (j,i))" and ns: "ns = list_of (b k' (j,i))" and "j < k" "j < k'" and lt_k: "\a\K. a < k" and lt_k': "\a\K. a < k'" by (auto simp: \_finite) then have len_eq [simp]: "length ns = length ms" by (simp add: card_b) have nz: "length ns \ 0" using b_ne \i < j\ \j < k'\ ns by auto show "(hd ms, hd ns) \ less_than \ (ms, ns) \ ?LL" proof assume "(hd ms, hd ns) \ less_than" then show "(ms, ns) \ ?LL" using that nz by (fastforce simp: lenlex_def \_finite card_b intro: hd_lex) next assume \
: "(ms, ns) \ ?LL" then have "(list_of (b k' (j,i)), list_of (b k (j,i))) \ ?LL" using less_asym ms ns omega_sum_1_less by blast then show "(hd ms, hd ns) \ less_than" using \j < k\ \j < k'\ Inf_b_less [of i j i j] ms ns by (metis Cons_lenlex_iff \
len_eq b_ne card_b_finite diff_Suc_1 hd_Cons_tl hd_b length_Cons less_or_eq_imp_le less_than_iff linorder_neqE_nat sorted_list_of_set_eq_Nil_iff that(2)) qed qed auto also have "\ = \" using infN ordertype_nat_\ by blast finally show ?thesis . qed have ot\j: "ordertype (BB j0 j ` \ j0 j) ?LL = \\j" if "j \ j0" for j j0 using that proof (induction j) \\a difficult proof, but no hints in Larson's text\ case 0 then show ?case by (auto simp: XX_def) next case (Suc j) then have ih: "ordertype (BB j0 j ` \ j0 j) ?LL = \ \ j" by simp have "j \ j0" by (simp add: Suc.prems Suc_leD) have inj_BB: "inj_on (BB j0 j) ([{j0<..}]\<^bsup>j\<^esup>)" proof (clarsimp simp: inj_on_def BB_def nsets_def sorted_list_of_set_Un less_sets_UN2) fix X Y assume X: "X \ {j0<..}" and Y: "Y \ {j0<..}" and "finite X" "finite Y" and jeq: "j = card X" and "card Y = card X" and eq: "list_of (a j0 \ (\i (\in. \n < card X\ \ j0 \ enum X n" using X \finite X\ finite_enumerate_in_set less_imp_le_nat by blast have enumY: "\n. \n < card X\ \ j0 \ enum Y n" using subsetD [OF Y] by (metis \card Y = card X\ \finite Y\ finite_enumerate_in_set greaterThan_iff less_imp_le_nat) have smX: "strict_mono_sets {..i. b (enum X i) (j0, i))" and smY: "strict_mono_sets {..i. b (enum Y i) (j0, i))" using Suc.prems \card Y = card X\ \finite X\ \finite Y\ bb enumX enumY jeq by (auto simp: strict_mono_sets_def) have len_eq: "length ms = length ns" if "(ms, ns) \ list.set (zip (map (list_of \ (\i. b (enum X i) (j0,i))) (list_of {.. (\i. b (enum Y i) (j0,i))) (list_of {.. card X" for ms ns n using that by (induction n rule: nat.induct) (auto simp: card_b enumX enumY) have "concat (map (list_of \ (\i. b (enum X i) (j0, i))) (list_of {.. (\i. b (enum Y i) (j0, i))) (list_of {.. (\i. b (enum X i) (j0, i))) (list_of {.. (\i. b (enum Y i) (j0, i))) (list_of {.. (b (enum X i) (j0,i))" "Inf (b (enum Y i) (j0,i)) \ (b (enum Y i) (j0,i))" "i < j0" using Inf_nat_def1 Suc.prems b_ne enumX enumY jeq that by auto ultimately show ?thesis by (metis Inf_b_less enumX enumY leI nat_less_le that) qed then show "X = Y" by (simp add: \card Y = card X\ \finite X\ \finite Y\ finite_enum_ext) qed have BB_Suc': "BB j0 (Suc j) X = usplit (\L k. BB j0 j L @ list_of (b k (j0, j))) X" if "X \ USigma (\ j0 j) (\K. {Max (insert j0 K)<..})" for X using that by (simp add: USigma_iff \_finite less_sets_def usplit_def \_Suc BB_Suc \j \ j0\) have "ordertype (BB j0 (Suc j) ` \ j0 (Suc j)) ?LL = ordertype (usplit (\L k. BB j0 j L @ list_of (b k (j0, j))) ` USigma (\ j0 j) (\K. {Max (insert j0 K)<..})) ?LL" by (simp add: BB_Suc' \_Suc) also have "\ = \ * ordertype (BB j0 j ` \ j0 j) ?LL" proof (intro ordertype_append_image_IJ) fix L k assume "L \ \ j0 j" and "k \ {Max (insert j0 L)<..}" then have "j0 < k" and L: "\a. a \ L \ a < k" by (simp_all add: \_finite) then show "BB j0 j L < list_of (b k (j0, j))" by (simp add: \L \ \ j0 j\ \j \ j0\ \_finite less_list_of) next show "inj_on (BB j0 j) (\ j0 j)" by (simp add: \_def inj_BB) next fix L assume L: "L \ \ j0 j" then show "L \ {Max (insert j0 L)<..} \ finite L" by (simp add: \_finite less_sets_def) show "ordertype ((\i. list_of (b i (j0, j))) ` {Max (insert j0 L)<..}) ?LL = \" using L Suc.prems Suc_le_lessD ot\ by blast qed (auto simp: \_finite card_b) also have "\ = \ \ ord_of_nat (Suc j)" by (simp add: oexp_mult_commute ih) finally show ?case . qed define seqs where "seqs \ \j0 j K. list_of (a j0) # (map (list_of \ (\i. b (enum K i) (j0,i))) (list_of {.. lists (- {[]})" if K: "K \ \ j0 j" and "j \ j0" for K j j0 proof - have j0: "\i. i < card K \ j0 \ enum K i" and le_j0: "card K \ j0" using finite_enumerate_in_set that unfolding \_def nsets_def by fastforce+ show "BB j0 j K = concat (seqs j0 j K)" using that unfolding BB_def \_def nsets_def seqs_def by (fastforce simp: j0 ab bb less_sets_UN2 sorted_list_of_set_Un strict_mono_sets_def sorted_list_of_set_UN_lessThan) have "b (enum K i) (j0, i) \ {}" if "i < card K" for i using j0 le_j0 less_le_trans that by simp moreover have "card K = j" using K \_card by blast ultimately show "seqs j0 j K \ lists (- {[]})" by (clarsimp simp: seqs_def) (metis card_b_finite sorted_list_of_set_eq_Nil_iff) qed have BB_decomp: "\cs. BB j0 j K = concat cs \ cs \ lists (- {[]})" if K: "K \ \ j0 j" and "j \ j0" for K j j0 using BB_eq_concat_seqs seqs_ne K that(2) by blast have a_subset_M: "a k \ M k" for k apply (clarsimp simp: a_def M_def DF_simps F_def Let_def split: prod.split_asm) by (metis (no_types) fst_conv fst_grab_subset nxt_subset snd_conv snd_grab_subset subsetD) have ba_Suc: "b k (j,i) \ a (Suc k)" if "i < j" "j \ k" for i j k by (meson a_subset_M bM less_sets_weaken2 nat_less_le that) have ba: "b k (j,i) \ a r" if "i < j" "j \ k" "k < r" for i j k r by (metis Suc_lessI a_ne aa ba_Suc less_sets_trans that) have disjnt_ba: "disjnt (b k (j,i)) (a r)" if "i < j" "j \ k" for i j k r by (meson ab ba disjnt_sym less_sets_imp_disjnt not_le that) have bb_disjnt: "disjnt (b k (j,i)) (b l (r,q))" if "q < r" "i < j" "j \ k" "r \ l" "j < r" for i j q r k l proof (cases "k=l") case True with that show ?thesis by (force simp: pair_less_def IJ_def intro: bb_same less_sets_imp_disjnt) next case False with that show ?thesis by (metis bb less_sets_imp_disjnt disjnt_sym nat_neq_iff) qed have sum_card_b: "(\i {j0<..}" "finite K" "card K \ j0" and "j \ j0" for j0 j K using \j \ j0\ proof (induction j) case 0 then show ?case by auto next case (Suc j) then have "j < card K" using that(3) by linarith have dis: "disjnt (b (enum K j) (j0, j)) (\ij < card K\ by (force simp: finite_enumerate_in_set) have "(\ii = card (b (enum K j) (j0, j)) + enum (d j0) j - enum (d j0) 0" using \Suc j \ j0\ by (simp add: Suc.IH split: nat_diff_split) also have "\ = enum (d j0) (Suc j) - enum (d j0) 0" using j0_less Suc.prems card_b less_or_eq_imp_le by force finally show ?case . qed have card_UN_b: "card (\i {j0<..}" "finite K" "card K \ j0" and "j \ j0" for j0 j K using that by (simp add: card_UN_disjoint sum_card_b b_disjoint) have len_BB: "length (BB j j K) = enum (d j) j" if K: "K \ \ j j" and "j \ j" for j K proof - have dis_ab: "\i. i < j \ disjnt (a j) (b (enum K i) (j,i))" using K \_card \_enum ab less_sets_imp_disjnt nat_less_le by blast show ?thesis using K unfolding BB_def \_def nsets_def by (simp add: card_UN_b card_Un_disjnt dis_ab card_a cInf_le_finite finite_enumerate_in_set enum_0_eq_Inf_finite) qed have "d k \ d (Suc k)" for k by (metis aM a_ne d_eq da less_sets_fst_grab less_sets_trans less_sets_weaken2 nxt_subset) then have dd: "d k' \ d k" if "k' < k" for k' k by (meson UNIV_I d_ne less_sets_imp_strict_mono_sets strict_mono_sets_def that) show thesis proof show "(\ (range XX)) \ WW" by (auto simp: XX_def BB_def WW_def) show "ordertype (\ (range XX)) (?LL) = \ \ \" using ot\j by (simp add: XX_def ordertype_\\) next fix U assume U: "U \ [\ (range XX)]\<^bsup>2\<^esup>" then obtain x y where Ueq: "U = {x,y}" and len_xy: "length x \ length y" by (auto simp: lenlex_nsets_2_eq lenlex_length) show "\l. Form l U \ (0 < l \ [enum N l] < inter_scheme l U \ list.set (inter_scheme l U) \ N)" proof (cases "length x = length y") case True then show ?thesis using Form.intros(1) U Ueq by fastforce next case False then have xy: "length x < length y" using len_xy by auto obtain j r K L where K: "K \ \ j j" and xeq: "x = BB j j K" and ne: "BB j j K \ BB r r L" and L: "L \ \ r r" and yeq: "y = BB r r L" using U by (auto simp: Ueq XX_def) then have "length x = enum (d j) j" "length y = enum (d r) r" by (auto simp: len_BB) then have "j < r" using xy dd by (metis card_d finite_enumerate_in_set finite_d lessI less_asym less_setsD linorder_neqE_nat) then have aj_ar: "a j \ a r" using aa by auto have Ksub: "K \ {j<..}" and "finite K" "card K \ j" using K by (auto simp: \_def nsets_def) have Lsub: "L \ {r<..}" and "finite L" "card L \ r" using L by (auto simp: \_def nsets_def) have enumK: "enum K i > j" if "i < j" for i using K \_card \_enum that by blast have enumL: "enum L i > r" if "i < r" for i using L \_card \_enum that by blast have "list.set (acc_lengths w (seqs j0 j K)) \ (+) w ` d j0" if K: "K \ {j0<..}" "finite K" "card K \ j0" and "j \ j0" for j0 j K w using \j \ j0\ proof (induction j arbitrary: w) case 0 then show ?case by (simp add: seqs_def Inf_nat_def1 card_a) next case (Suc j) let ?db = "\ (d j0) + ((\i \ (d j0)" using Suc.prems card_d by (simp add: cInf_le_finite finite_enumerate_in_set) then have "?db = enum (d j0) (Suc j)" using Suc.prems that by (simp add: cInf_le_finite finite_enumerate_in_set sum_card_b card_b enum_d_0 \j0 < enum K j\ less_or_eq_imp_le) then have "?db \ d j0" using Suc.prems finite_enumerate_in_set by (auto simp: finite_enumerate_in_set) moreover have "list.set (acc_lengths w (seqs j0 j K)) \ (+) w ` d j0" by (simp add: Suc Suc_leD) then have "list.set (acc_lengths (w + \ (d j0)) (map (list_of \ (\i. b (enum K i) (j0,i))) (list_of {.. (+) w ` d j0" by (simp add: seqs_def card_a subset_insertI) ultimately show ?case by (simp add: seqs_def acc_lengths_append image_iff Inf_nat_def1 sum_sorted_list_of_set_map card_a) qed then have acc_lengths_subset_d: "list.set (acc_lengths 0 (seqs j0 j K)) \ d j0" if K: "K \ {j0<..}" "finite K" "card K \ j0" and "j \ j0" for j0 j K by (metis image_add_0 that) have "strict_sorted x" "strict_sorted y" by (auto simp: xeq yeq BB_def) have disjnt_xy: "disjnt (list.set x) (list.set y)" proof - have "disjnt (a j) (a r)" using \j < r\ aa less_sets_imp_disjnt by blast moreover have "disjnt (b (enum K i) (j,i)) (a r)" if "i < j" for i by (simp add: disjnt_ba enumK less_imp_le_nat that) moreover have "disjnt (a j) (b (enum L q) (r,q))" if "q < r" for q by (meson disjnt_ba disjnt_sym enumL less_imp_le_nat that) moreover have "disjnt (b (enum K i) (j,i)) (b (enum L q) (r,q))" if "i < j" "q < r" for i q by (meson \j < r\ bb_disjnt enumK enumL less_imp_le that) ultimately show ?thesis by (simp add: xeq yeq BB_def) qed have "\us vs. merge (seqs j j K) (seqs r r L) us vs" proof (rule merge_exists) show "strict_sorted (concat (seqs j j K))" using BB_eq_concat_seqs K \strict_sorted x\ xeq by auto show "strict_sorted (concat (seqs r r L))" using BB_eq_concat_seqs L \strict_sorted y\ yeq by auto show "seqs j j K \ lists (- {[]})" "seqs r r L \ lists (- {[]})" by (auto simp: K L seqs_ne) show "hd (seqs j j K) < hd (seqs r r L)" by (simp add: aj_ar less_sets_imp_list_less seqs_def) show "seqs j j K \ []" "seqs r r L \ []" using seqs_def by blast+ have less_bb: "b (enum K i) (j,i) \ b (enum L p) (r, p)" if "\ b (enum L p) (r, p) \ b (enum K i) (j,i)" and "i < j" "p < r" for i p by (metis IJ_iff \j < r\ bb bb_same enumK enumL less_imp_le_nat linorder_neqE_nat pair_lessI1 that) show "u < v \ v < u" if "u \ list.set (seqs j j K)" and "v \ list.set (seqs r r L)" for u v using that enumK enumL unfolding seqs_def apply (auto simp: seqs_def aj_ar intro!: less_bb less_sets_imp_list_less) apply (meson ab ba less_imp_le_nat not_le)+ done qed then obtain uus vvs where merge: "merge (seqs j j K) (seqs r r L) uus vvs" by metis then have "uus \ []" using merge_length1_gt_0 by (auto simp: seqs_def) then obtain u1 us where us: "u1#us = uus" by (metis neq_Nil_conv) define ku where "ku \ length (u1#us)" define ps where "ps \ acc_lengths 0 (u1#us)" have us_ne: "u1#us \ lists (- {[]})" using merge_length1_nonempty seqs_ne us merge us K by auto have xu_eq: "x = concat (u1#us)" using BB_eq_concat_seqs K merge merge_preserves us xeq by auto then have "strict_sorted u1" using \strict_sorted x\ strict_sorted_append_iff by auto have u_sub: "list.set ps \ list.set (acc_lengths 0 (seqs j j K))" using acc_lengths_merge1 merge ps_def us by blast have "vvs \ []" using merge BB_eq_concat_seqs L merge_preserves xy yeq by auto then obtain v1 vs where vs: "v1#vs = vvs" by (metis neq_Nil_conv) define kv where "kv \ length (v1#vs)" define qs where "qs \ acc_lengths 0 (v1#vs)" have vs_ne: "v1#vs \ lists (- {[]})" using L merge merge_length2_nonempty seqs_ne vs by auto have yv_eq: "y = concat (v1#vs)" using BB_eq_concat_seqs L merge merge_preserves vs yeq by auto then have "strict_sorted v1" using \strict_sorted y\ strict_sorted_append_iff by auto have v_sub: "list.set qs \ list.set (acc_lengths 0 (seqs r r L))" using acc_lengths_merge2 merge qs_def vs by blast have ss_concat_jj: "strict_sorted (concat (seqs j j K))" using BB_eq_concat_seqs K \strict_sorted x\ xeq by auto then obtain k: "0 < kv" "kv \ ku" "ku \ Suc kv" "kv \ Suc j" using us vs merge_length_le merge_length_le_Suc merge_length_less2 merge unfolding ku_def kv_def by fastforce define zs where "zs \ concat [ps,u1,qs,v1] @ interact us vs" have ss: "strict_sorted zs" proof - have ssp: "strict_sorted ps" unfolding ps_def by (meson strict_sorted_acc_lengths us_ne) have ssq: "strict_sorted qs" unfolding qs_def by (meson strict_sorted_acc_lengths vs_ne) have "d j \ list.set x" using da [of j] db [of j] K \_card \_enum nat_less_le by (auto simp: xeq BB_def less_sets_Un2 less_sets_UN2) then have ac_x: "acc_lengths 0 (seqs j j K) < x" by (meson Ksub \finite K\ \j \ card K\ acc_lengths_subset_d le_refl less_sets_imp_list_less less_sets_weaken1) then have "ps < x" by (meson Ksub \d j \ list.set x\ \finite K\ \j \ card K\ acc_lengths_subset_d le_refl less_sets_imp_list_less less_sets_weaken1 u_sub) then have "ps < u1" by (metis Nil_is_append_conv concat.simps(2) hd_append2 less_list_def xu_eq) have "d r \ list.set y" using da [of r] db [of r] L \_card \_enum nat_less_le by (auto simp: yeq BB_def less_sets_Un2 less_sets_UN2) then have "acc_lengths 0 (seqs r r L) < y" by (meson Lsub \finite L\ \r \ card L\ acc_lengths_subset_d le_refl less_sets_imp_list_less less_sets_weaken1) then have "qs < y" by (metis L Lsub \_card \d r \ list.set y\ \finite L\ acc_lengths_subset_d less_sets_imp_list_less less_sets_weaken1 order_refl v_sub) then have "qs < v1" by (metis concat.simps(2) gr_implies_not0 hd_append2 less_list_def list.size(3) xy yv_eq) have carda_v1: "card (a r) \ length v1" using length_hd_merge2 [OF merge] unfolding vs [symmetric] by (simp add: seqs_def) have ab_enumK: "\i. i < j \ a j \ b (enum K i) (j,i)" by (meson ab enumK le_trans less_imp_le_nat) have ab_enumL: "\q. q < r \ a j \ b (enum L q) (r,q)" by (meson \j < r\ ab enumL le_trans less_imp_le_nat) then have ay: "a j \ list.set y" by (auto simp: yeq BB_def less_sets_Un2 less_sets_UN2 aj_ar) have disjnt_hd_last_K_y: "disjnt {hd l..last l} (list.set y)" if l: "l \ list.set (seqs j j K)" for l proof (clarsimp simp add: yeq BB_def disjnt_iff Ball_def, intro conjI strip) fix u assume u: "u \ last l" and "hd l \ u" with l consider "u \ last (list_of (a j))" "hd (list_of (a j)) \ u" | i where "i last (list_of (b (enum K i) (j,i)))" "hd (list_of (b (enum K i) (j,i))) \ u" by (force simp: seqs_def) note l_cases = this then show "u \ a r" proof cases case 1 then show ?thesis by (metis a_ne aj_ar finite_a last_in_set leD less_setsD set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) next case 2 then show ?thesis by (metis enumK ab ba Inf_nat_def1 b_ne card_b_finite hd_b last_in_set less_asym less_setsD not_le set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) qed fix q assume "q < r" show "u \ b (enum L q) (r,q)" using l_cases proof cases case 1 then show ?thesis by (metis \q < r\ a_ne ab_enumL finite_a last_in_set leD less_setsD set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) next case 2 show ?thesis proof (cases "enum K i = enum L q") case True then show ?thesis using 2 bb_same [of concl: "enum L q" j i r q] \j < r\ u by (metis IJ_iff b_ne card_b_finite enumK last_in_set leD less_imp_le_nat less_setsD pair_lessI1 set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) next case False with 2 bb enumK enumL show ?thesis unfolding less_sets_def by (metis \q < r\ b_ne card_b_finite last_in_set leD less_imp_le_nat list.set_sel(1) nat_neq_iff set_sorted_list_of_set sorted_list_of_set_eq_Nil_iff) qed qed qed have u1_y: "list.set u1 \ list.set y" using vs yv_eq L \strict_sorted y\ merge merge_less_sets_hd merge_preserves seqs_ne ss_concat_jj us by fastforce have u1_subset_seqs: "list.set u1 \ list.set (concat (seqs j j K))" using merge_preserves [OF merge] us by auto have "b k (j,i) \ d (Suc k)" if "j\k" "i d k'" if "j\k" "i d (Suc k)" for k by (metis aM d_eq less_sets_fst_grab less_sets_weaken2 nxt_subset) then have ad: "a k \ d k'" if "k list.set u1" for n proof - obtain l where l: "l \ list.set (seqs j j K)" and n: "n \ list.set l" using n u1_subset_seqs by auto then consider "l = list_of (a j)" | i where "l = list_of (b (enum K i) (j,i))" "i < j" by (force simp: seqs_def) then show ?thesis proof cases case 1 then show ?thesis by (metis Inf_nat_def1 \j < r\ ad d_ne finite_a less_setsD n set_sorted_list_of_set) next case 2 then have "hd (list_of (b (enum K i) (j,i))) = Min (b (enum K i) (j,i))" by (meson b_ne card_b_finite enumK hd_list_of less_imp_le_nat) also have "\ \ n" using 2 n by (simp add: less_list_def disjnt_iff less_sets_def) also have f8: "n < hd y" using less_setsD that u1_y by (metis gr_implies_not0 list.set_sel(1) list.size(3) xy) finally have "l < y" using 2 disjnt_hd_last_K_y [OF l] by (simp add: disjnt_iff) (metis leI less_imp_le_nat less_list_def list.set_sel(1)) moreover have "last (list_of (b (enum K i) (j,i))) < hd (list_of (a r))" using \l < y\ L n by (auto simp: 2yeq BB_eq_concat_seqs seqs_def less_list_def) then have "enum K i < r" by (metis "2"(1) a_ne ab card_b_finite empty_iff finite.emptyI finite_a last_in_set leI less_asym less_setsD list.set_sel(1) n set_sorted_list_of_set) moreover have "j \ enum K i" by (simp add: "2"(2) enumK less_imp_le_nat) ultimately show ?thesis using 2 n bd [of j "enum K i" i r] Inf_nat_def1 less_setsD by force qed qed then have "last u1 < Inf (d r)" using \uus \ []\ us_ne by auto also have "\ \ length v1" using card_a carda_v1 by auto finally have "last u1 < length v1" . then have "u1 < qs" by (simp add: qs_def less_list_def) have "strict_sorted (interact (u1#us) (v1#vs))" using L \strict_sorted x\ \strict_sorted y\ merge merge_interact merge_preserves seqs_ne us vs xu_eq yv_eq by auto then have "strict_sorted (interact us vs)" "v1 < interact us vs" by (auto simp: strict_sorted_append_iff) moreover have "ps < u1 @ qs @ v1 @ interact us vs" using \ps < u1\ us_ne unfolding less_list_def by auto moreover have "u1 < qs @ v1 @ interact us vs" by (metis \u1 < qs\ \vvs \ []\ acc_lengths_eq_Nil_iff hd_append less_list_def qs_def vs) moreover have "qs < v1 @ interact us vs" using \qs < v1\ us_ne \last u1 < length v1\ vs_ne by (auto simp: less_list_def) ultimately show ?thesis by (simp add: zs_def strict_sorted_append_iff ssp ssq \strict_sorted u1\ \strict_sorted v1\) qed have ps_subset_d: "list.set ps \ d j" using K Ksub \_card \finite K\ acc_lengths_subset_d u_sub by blast have ps_less_u1: "ps < u1" proof - have "hd u1 = hd x" using us_ne by (auto simp: xu_eq) then have "hd u1 \ a j" by (simp add: xeq BB_eq_concat_seqs K seqs_def hd_append hd_list_of) then have "list.set ps \ {hd u1}" by (metis da ps_subset_d less_sets_def singletonD subset_iff) then show ?thesis by (metis less_hd_imp_less list.set(2) empty_set less_sets_imp_list_less) qed have qs_subset_d: "list.set qs \ d r" using L Lsub \_card \finite L\ acc_lengths_subset_d v_sub by blast have qs_less_v1: "qs < v1" proof - have "hd v1 = hd y" using vs_ne by (auto simp: yv_eq) then have "hd v1 \ a r" by (simp add: yeq BB_eq_concat_seqs L seqs_def hd_append hd_list_of) then have "list.set qs \ {hd v1}" by (metis da qs_subset_d less_sets_def singletonD subset_iff) then show ?thesis by (metis less_hd_imp_less list.set(2) empty_set less_sets_imp_list_less) qed have FB: "Form_Body ku kv x y zs" unfolding Form_Body.simps using ku_def kv_def ps_def qs_def ss us_ne vs_ne xu_eq xy yv_eq zs_def by blast then have "zs = (inter_scheme ((ku+kv) - Suc 0) {x,y})" by (simp add: Form_Body_imp_inter_scheme k) obtain l where "l \ 2 * (Suc j)" and l: "Form l U" and zs_eq_interact: "zs = inter_scheme l {x,y}" proof show "ku+kv-1 \ 2 * (Suc j)" using k by auto show "Form (ku+kv-1) U" proof (cases "ku=kv") case True then show ?thesis using FB Form.simps Ueq \0 < kv\ by (auto simp: mult_2) next case False then have "ku = Suc kv" using k by auto then show ?thesis using FB Form.simps Ueq \0 < kv\ by auto qed show "zs = inter_scheme (ku + kv - 1) {x, y}" using Form_Body_imp_inter_scheme by (simp add: FB k) qed then have "enum N l \ enum N (Suc (2 * Suc j))" by (simp add: assms less_imp_le_nat) also have "\ < Min (d j)" by (smt (verit, best) Min_gr_iff d_eq d_ne finite_d fst_grab_subset greaterThan_iff in_mono le_inf_iff nxt_def) finally have ls: "{enum N l} \ d j" by simp have "l > 0" by (metis l False Form_0_cases_raw Set.doubleton_eq_iff Ueq gr0I) show ?thesis unfolding Ueq proof (intro exI conjI impI) have zs_subset: "list.set zs \ list.set (acc_lengths 0 (seqs j j K)) \ list.set (acc_lengths 0 (seqs r r L)) \ list.set x \ list.set y" using u_sub v_sub by (auto simp: zs_def xu_eq yv_eq) also have "\ \ N" proof (simp, intro conjI) show "list.set (acc_lengths 0 (seqs j j K)) \ N" using d_subset_N Ksub \finite K\ \j \ card K\ acc_lengths_subset_d by blast show "list.set (acc_lengths 0 (seqs r r L)) \ N" using d_subset_N Lsub \finite L\ \r \ card L\ acc_lengths_subset_d by blast show "list.set x \ N" "list.set y \ N" by (simp_all add: xeq yeq BB_def a_subset_N UN_least b_subset_N) qed finally show "list.set (inter_scheme l {x, y}) \ N" using zs_eq_interact by blast have "[enum N l] < ps" using ps_subset_d ls by (metis empty_set less_sets_imp_list_less less_sets_weaken2 list.simps(15)) then show "[enum N l] < inter_scheme l {x, y}" by (simp add: zs_def less_list_def ps_def flip: zs_eq_interact) qed (use Ueq l in blast) qed qed qed subsection \The main partition theorem for @{term "\\\"}\ definition iso_ll where "iso_ll A B \ iso (lenlex less_than \ (A\A)) (lenlex less_than \ (B\B))" corollary ordertype_eq_ordertype_iso_ll: assumes "Field (Restr (lenlex less_than) A) = A" "Field (Restr (lenlex less_than) B) = B" shows "(ordertype A (lenlex less_than) = ordertype B (lenlex less_than)) \ (\f. iso_ll A B f)" proof - have "total_on A (lenlex less_than) \ total_on B (lenlex less_than)" by (meson UNIV_I total_lenlex total_on_def total_on_less_than) then show ?thesis by (simp add: assms wf_lenlex lenlex_transI iso_ll_def ordertype_eq_ordertype_iso_Restr) qed theorem partition_\\_aux: assumes "\ \ elts \" shows "partn_lst (lenlex less_than) WW [\\\,\] 2" (is "partn_lst ?R WW [\\\,\] 2") proof (cases "\ \ 1") case True then show ?thesis using strict_sorted_into_WW unfolding WW_def by (auto intro!: partn_lst_triv1[where i=1]) next case False obtain m where m: "\ = ord_of_nat m" using assms elts_\ by auto then have "m>1" using False by auto show ?thesis unfolding partn_lst_def proof clarsimp fix f assume f: "f \ [WW]\<^bsup>2\<^esup> \ {..: "?P0 \ ?P1" proof (rule disjCI) assume not1: "\ ?P1" have "\W'. ordertype W' ?R = \\n \ f ` [W']\<^bsup>2\<^esup> \ {0} \ W' \ WW_seg (n*m)" for n::nat proof - have fnm: "f \ [WW_seg (n*m)]\<^bsup>2\<^esup> \ {..\n, ord_of_nat m] 2" using ordertype_WW_seg [of "n*m"] by (simp add: partn_lst_VWF_imp_partn_lst [OF Theorem_3_2]) show ?thesis using partn_lst_E [OF * fnm, simplified] by (metis One_nat_def WW_seg_subset_WW less_2_cases m not1 nth_Cons_0 nth_Cons_Suc numeral_2_eq_2 subset_trans) qed then obtain W':: "nat \ nat list set" where otW': "\n. ordertype (W' n) ?R = \\n" and f_W': "\n. f ` [W' n]\<^bsup>2\<^esup> \ {0}" and seg_W': "\n. W' n \ WW_seg (n*m)" by metis define WW' where "WW' \ (\n. W' n)" have "WW' \ WW" using seg_W' WW_seg_subset_WW by (force simp: WW'_def) with f have f': "f \ [WW']\<^bsup>2\<^esup> \ {..\\" proof (rule antisym) have "ordertype WW' ?R \ ordertype WW ?R" by (simp add: \WW' \ WW\ lenlex_transI ordertype_mono wf_lenlex) with ordertype_WW show "ordertype WW' ?R \ \ \ \" by simp have "\ \ n \ ordertype (\ (range W')) ?R" for n::nat using oexp_Limit ordertype_\\ otW' by auto then show "\ \ \ \ ordertype WW' ?R" by (auto simp: elts_\ oexp_Limit ZFC_in_HOL.SUP_le_iff WW'_def) qed have FR_WW: "Field (Restr (lenlex less_than) WW) = WW" by (simp add: Limit_omega_oexp Limit_ordertype_imp_Field_Restr ordertype_WW) have FR_WW': "Field (Restr (lenlex less_than) WW') = WW'" by (simp add: Limit_omega_oexp Limit_ordertype_imp_Field_Restr ot') have FR_W: "Field (Restr (lenlex less_than) (WW_seg n)) = WW_seg n" if "n>0" for n by (simp add: Limit_omega_oexp ordertype_WW_seg that Limit_ordertype_imp_Field_Restr) have FR_W': "Field (Restr (lenlex less_than) (W' n)) = W' n" if "n>0" for n by (simp add: Limit_omega_oexp otW' that Limit_ordertype_imp_Field_Restr) have "\h. iso_ll (WW_seg n) (W' n) h" if "n>0" for n proof (subst ordertype_eq_ordertype_iso_ll [symmetric]) show "ordertype (WW_seg n) (lenlex less_than) = ordertype (W' n) (lenlex less_than)" by (simp add: ordertype_WW_seg otW') qed (auto simp: FR_W FR_W' that) then obtain h_seg where h_seg: "\n. n > 0 \ iso_ll (WW_seg n) (W' n) (h_seg n)" by metis define h where "h \ \l. if l=[] then [] else h_seg (length l) l" have bij_h_seg: "\n. n > 0 \ bij_betw (h_seg n) (WW_seg n) (W' n)" using h_seg by (simp add: iso_ll_def iso_iff2 FR_W FR_W') have len_h_seg: "length (h_seg (length l) l) = length l * m" if "length l > 0" "l \ WW" for l using bij_betwE [OF bij_h_seg] seg_W' that by (simp add: WW_seg_def subset_iff) have hlen: "length (h x) = length (h y) \ length x = length y" if "x \ WW" "y \ WW" for x y using that \1 < m\ h_def len_h_seg by force have h: "iso_ll WW WW' h" unfolding iso_ll_def iso_iff2 FR_WW FR_WW' proof (intro conjI strip) have W'_ne: "W' n \ {}" for n using otW' [of n] by auto then have "[] \ WW'" using seg_W' [of 0] by (auto simp: WW'_def WW_seg_def) let ?g = "\l. if l=[] then l else inv_into (WW_seg (length l div m)) (h_seg (length l div m)) l" have h_seg_iff: "\n a b. \a \ WW_seg n; b \ WW_seg n; n>0\ \ (a, b) \ lenlex less_than \ (h_seg n a, h_seg n b) \ lenlex less_than \ h_seg n a \ W' n \ h_seg n b \ W' n" using h_seg by (auto simp: iso_ll_def iso_iff2 FR_W FR_W') show "bij_betw h WW WW'" unfolding bij_betw_iff_bijections proof (intro exI conjI ballI) fix l assume "l \ WW" then have l: "l \ WW_seg (length l)" by (simp add: WW_seg_def) have "h l \ W' (length l)" proof (cases "l=[]") case True with seg_W' [of 0] W'_ne show ?thesis by (auto simp: WW_seg_def h_def) next case False then show ?thesis using bij_betwE bij_h_seg h_def l by fastforce qed show "h l \ WW'" using WW'_def \h l \ W' (length l)\ by blast show "?g (h l) = l" proof (cases "l=[]") case False then have "length l > 0" by auto then have "h_seg (length l) l \ []" using \1 < m\ \l \ WW\ len_h_seg by fastforce moreover have "bij_betw (h_seg (length l)) (WW_seg (length l)) (W' (length l))" using \0 < length l\ bij_h_seg by presburger ultimately show ?thesis using \l \ WW\ bij_betw_inv_into_left h_def l len_h_seg by fastforce qed (auto simp: h_def) next fix l assume "l \ WW'" then have l: "l \ W' (length l div m)" using WW_seg_def \1 < m\ seg_W' by (fastforce simp: WW'_def) show "?g l \ WW" proof (cases "l=[]") case False then have "l \ W' 0" using WW_seg_def seg_W' by fastforce with l have "inv_into (WW_seg (length l div m)) (h_seg (length l div m)) l \ WW_seg (length l div m)" by (metis Nat.neq0_conv bij_betwE bij_betw_inv_into bij_h_seg) then show ?thesis using False WW_seg_subset_WW by auto qed (auto simp: WW_def) show "h (?g l) = l" proof (cases "l=[]") case False then have "0 < length l div m" using WW_seg_def l seg_W' by fastforce then have "inv_into (WW_seg (length l div m)) (h_seg (length l div m)) l \ WW_seg (length l div m)" by (metis bij_betw_imp_surj_on bij_h_seg inv_into_into l) then show ?thesis using bij_h_seg [of "length l div m"] WW_seg_def \0 < length l div m\ bij_betw_inv_into_right l by (fastforce simp: h_def) qed (auto simp: h_def) qed fix a b assume "a \ WW" "b \ WW" show "(a, b) \ Restr (lenlex less_than) WW \ (h a, h b) \ Restr (lenlex less_than) WW'" (is "?lhs = ?rhs") proof assume L: ?lhs then consider "length a < length b" | "length a = length b" "(a, b) \ lex less_than" by (auto simp: lenlex_conv) then show ?rhs proof cases case 1 then have "length (h a) < length (h b)" using \1 < m\ \a \ WW\ \b \ WW\ h_def len_h_seg by auto then have "(h a, h b) \ lenlex less_than" by (auto simp: lenlex_conv) then show ?thesis using \a \ WW\ \b \ WW\ \bij_betw h WW WW'\ bij_betwE by fastforce next case 2 then have ab: "a \ WW_seg (length a)" "b \ WW_seg (length a)" using \a \ WW\ \b \ WW\ by (auto simp: WW_seg_def) have "length (h a) = length (h b)" using 2 \a \ WW\ \b \ WW\ h_def len_h_seg by force moreover have "(a, b) \ lenlex less_than" using L by blast then have "(h_seg (length a) a, h_seg (length a) b) \ lenlex less_than" using 2 ab h_seg_iff by blast ultimately show ?thesis using 2 \a \ WW\ \b \ WW\ \bij_betw h WW WW'\ bij_betwE h_def by fastforce qed next assume R: ?rhs then have R': "(h a, h b) \ lenlex less_than" by blast then consider "length a < length b" | "length a = length b" "(h a, h b) \ lex less_than" using \a \ WW\ \b \ WW\ \m > 1\ by (auto simp: lenlex_conv h_def len_h_seg split: if_split_asm) then show ?lhs proof cases case 1 then show ?thesis using omega_sum_less_iff \a \ WW\ \b \ WW\ by auto next case 2 then have ab: "a \ WW_seg (length a)" "b \ WW_seg (length a)" using \a \ WW\ \b \ WW\ by (auto simp: WW_seg_def) then have "(a, b) \ lenlex less_than" using bij_betwE [OF bij_h_seg] \a \ WW\ \b \ WW\ R' 2 by (simp add: h_def h_seg_iff split: if_split_asm) then show ?thesis using \a \ WW\ \b \ WW\ by blast qed qed qed let ?fh = "f \ image h" have "bij_betw h WW WW'" using h unfolding iso_ll_def iso_iff2 by (fastforce simp: FR_WW FR_WW') moreover have "{.. [WW]\<^bsup>2\<^esup> \ {0,1}" unfolding Pi_iff using bij_betwE f' bij_betw_nsets by (metis PiE comp_apply) have "f{x,y} = 0" if "x \ WW'" "y \ WW'" "length x = length y" "x \ y" for x y proof - obtain p q where "x \ W' p" and "y \ W' q" using WW'_def \x \ WW'\ \y \ WW'\ by blast then obtain n where "{x,y} \ [W' n]\<^bsup>2\<^esup>" using seg_W' \1 < m\ \length x = length y\ \x \ y\ by (auto simp: WW'_def WW_seg_def subset_iff) then show "f{x,y} = 0" using f_W' by blast qed then have fh_eq_0_eqlen: "?fh{x,y} = 0" if "x \ WW" "y \ WW" "length x = length y" "x\y" for x y using \bij_betw h WW WW'\ that hlen by (simp add: bij_betw_iff_bijections) metis have m_f_0: "\x\[M]\<^bsup>2\<^esup>. f x = 0" if "M \ WW" "card M = m" for M proof - have "finite M" using False m that by auto with not1 [simplified, rule_format, of M] f show ?thesis using that \1 < m\ apply (simp add: Pi_iff image_subset_iff finite_ordertype_eq_card m) by (metis less_2_cases nsets_mono numeral_2_eq_2 subset_iff) qed have m_fh_0: "\x\[M]\<^bsup>2\<^esup>. ?fh x = 0" if "M \ WW" "card M = m" for M proof - have "h ` M \ WW" using \WW' \ WW\ \bij_betw h WW WW'\ bij_betwE that(1) by fastforce moreover have "card (h ` M) = m" by (metis \bij_betw h WW WW'\ bij_betw_def bij_betw_subset card_image that) ultimately have "\x \ [h ` M]\<^bsup>2\<^esup>. f x = 0" by (metis m_f_0) then obtain Y where Y: "f (h ` Y) = 0" "Y \ M" and "finite (h ` Y)" "card (h ` Y) = 2" by (auto simp: nsets_def subset_image_iff) then have "card Y = 2" using \bij_betw h WW WW'\ \M \ WW\ by (metis bij_betw_def card_image inj_on_subset) with Y card.infinite[of Y] show ?thesis by (auto simp: nsets_def) qed obtain N j where "infinite N" and N: "\k u. \k > 0; u \ [WW]\<^bsup>2\<^esup>; Form k u; [enum N k] < inter_scheme k u; List.set (inter_scheme k u) \ N\ \ ?fh u = j k" using lemma_3_6 [OF fh] by blast have infN': "infinite (enum N ` {k<..})" for k by (simp add: \infinite N\ enum_works finite_image_iff infinite_Ioi strict_mono_imp_inj_on) have j_0: "j k = 0" if "k>0" for k proof - obtain M where M: "M \ [WW]\<^bsup>m\<^esup>" and MF: "\u. u \ [M]\<^bsup>2\<^esup> \ Form k u" and Mi: "\u. u \ [M]\<^bsup>2\<^esup> \ List.set (inter_scheme k u) \ enum N ` {k<..}" using lemma_3_7 [OF infN' \k > 0\] by metis obtain u where u: "u \ [M]\<^bsup>2\<^esup>" "?fh u = 0" using m_fh_0 [of M] M [unfolded nsets_def] by force moreover have \
: "Form k u" "List.set (inter_scheme k u) \ enum N ` {k<..}" by (simp_all add: MF Mi \u \ [M]\<^bsup>2\<^esup>\) then have "hd (inter_scheme k u) \ enum N ` {k<..}" using hd_in_set inter_scheme_simple that by blast then have "[enum N k] < inter_scheme k u" using strict_mono_enum [OF \infinite N\] by (auto simp: less_list_def strict_mono_def) moreover have "u \ [WW]\<^bsup>2\<^esup>" using M u by (auto simp: nsets_def) moreover have "enum N ` {k<..} \ N" using \infinite N\ range_enum by auto ultimately show ?thesis using N \
that by auto qed obtain X where "X \ WW" and otX: "ordertype X (lenlex less_than) = \\\" and X: "\u. u \ [X]\<^bsup>2\<^esup> \ \l. Form l u \ (l > 0 \ [enum N l] < inter_scheme l u \ List.set (inter_scheme l u) \ N)" using lemma_3_8 [OF \infinite N\] ot' by blast have 0: "?fh ` [X]\<^bsup>2\<^esup> \ {0}" proof clarsimp fix u assume u: "u \ [X]\<^bsup>2\<^esup>" obtain l where "Form l u" and l: "l > 0 \ [enum N l] < inter_scheme l u \ List.set (inter_scheme l u) \ N" using u X by blast have "?fh u = 0" proof (cases "l = 0") case True then show ?thesis by (metis Form_0_cases_raw \Form l u\ \X \ WW\ doubleton_in_nsets_2 fh_eq_0_eqlen subset_iff u) next case False then obtain "[enum N l] < inter_scheme l u" "List.set (inter_scheme l u) \ N" "j l = 0" using Nat.neq0_conv j_0 l by blast with False show ?thesis using \X \ WW\ N inter_scheme \Form l u\ doubleton_in_nsets_2 u by (auto simp: nsets_def) qed then show "f (h ` u) = 0" by auto qed show ?P0 proof (intro exI conjI) show "h ` X \ WW" using \WW' \ WW\ \X \ WW\ \bij_betw h WW WW'\ bij_betw_imp_surj_on by fastforce show "ordertype (h ` X) (lenlex less_than) = \ \ \" proof (subst ordertype_inc_eq) show "(h x, h y) \ lenlex less_than" if "x \ X" "y \ X" "(x, y) \ lenlex less_than" for x y using that h \X \ WW\ by (auto simp: FR_WW FR_WW' iso_iff2 iso_ll_def) qed (use otX in auto) show "f ` [h ` X]\<^bsup>2\<^esup> \ {0}" proof (clarsimp simp: image_subset_iff nsets_def) fix Y assume "Y \ h ` X" and "finite Y" and "card Y = 2" have "inv_into WW h ` Y \ X" using \X \ WW\ \Y \ h ` X\ \bij_betw h WW WW'\ bij_betw_inv_into_LEFT by blast moreover have "finite (inv_into WW h ` Y)" using \finite Y\ by blast moreover have "card (inv_into WW h ` Y) = 2" by (metis \X \ WW\ \Y \ h ` X\ \card Y = 2\ card_image inj_on_inv_into subset_image_iff subset_trans) ultimately have "f (h ` inv_into WW h ` Y) = 0" using 0 by (auto simp: image_subset_iff nsets_def) then show "f Y = 0" by (metis \X \ WW\ \Y \ h ` X\ image_inv_into_cancel image_mono order_trans) qed qed qed then show "\iH\WW. ordertype H ?R = [\\\, \] ! i \ f ` [H]\<^bsup>2\<^esup> \ {i}" by (metis One_nat_def lessI nth_Cons_0 nth_Cons_Suc zero_less_Suc) qed qed text \Theorem 3.1 of Jean A. Larson, ibid.\ theorem partition_\\: "\ \ elts \ \ partn_lst_VWF (\\\) [\\\,\] 2" using partn_lst_imp_partn_lst_VWF_eq [OF partition_\\_aux] ordertype_WW by auto end diff --git a/thys/Poincare_Bendixson/ODE_Misc.thy b/thys/Poincare_Bendixson/ODE_Misc.thy --- a/thys/Poincare_Bendixson/ODE_Misc.thy +++ b/thys/Poincare_Bendixson/ODE_Misc.thy @@ -1,1111 +1,1111 @@ section \Additions to the ODE Library\ theory ODE_Misc imports Ordinary_Differential_Equations.ODE_Analysis Analysis_Misc begin lemma local_lipschitz_compact_bicomposeE: assumes ll: "local_lipschitz T X f" assumes cf: "\x. x \ X \ continuous_on I (\t. f t x)" assumes cI: "compact I" assumes "I \ T" assumes cv: "continuous_on I v" assumes cw: "continuous_on I w" assumes v: "v ` I \ X" assumes w: "w ` I \ X" obtains L where "L > 0" "\x. x \ I \ dist (f x (v x)) (f x (w x)) \ L * dist (v x) (w x)" proof - from v w have "v ` I \ w ` I \ X" by auto with ll \I \ T\ have llI:"local_lipschitz I (v ` I \ w ` I) f" by (rule local_lipschitz_subset) have cvwI: "compact (v ` I \ w ` I)" by (auto intro!: compact_continuous_image cv cw cI) from local_lipschitz_compact_implies_lipschitz[OF llI cvwI \compact I\ cf] obtain L where L: "\t. t \ I \ L-lipschitz_on (v ` I \ w ` I) (f t)" using v w by blast define L' where "L' = max L 1" with L have "L' > 0" "\x. x \ I \ dist (f x (v x)) (f x (w x)) \ L' * dist (v x) (w x)" apply (auto simp: lipschitz_on_def L'_def) by (smt Un_iff image_eqI mult_right_mono zero_le_dist) then show ?thesis .. qed subsection \Comparison Principle\ lemma comparison_principle_le: fixes f::"real \ real \ real" and \ \::"real \ real" assumes ll: "local_lipschitz X Y f" assumes cf: "\x. x \ Y \ continuous_on {a..b} (\t. f t x)" assumes abX: "{a .. b} \ X" assumes \': "\x. x \ {a .. b} \ (\ has_real_derivative \' x) (at x)" assumes \': "\x. x \ {a .. b} \ (\ has_real_derivative \' x) (at x)" assumes \_in: "\ ` {a..b} \ Y" assumes \_in: "\ ` {a..b} \ Y" assumes init: "\ a \ \ a" assumes defect: "\x. x \ {a .. b} \ \' x - f x (\ x) \ \' x - f x (\ x)" shows "\x \ {a .. b}. \ x \ \ x" (is "?th1") (* "(\x \ {a .. b}. \ x < \ x) \ (\c\{a..b}. (\x\{a..c}. \ x \ \ x) \ (\x\{c<..b}. \ x < \ x))" (is "?th2") *) unfolding atomize_conj apply (cases "a \ b") defer subgoal by simp proof - assume "a \ b" note \_cont = has_real_derivative_imp_continuous_on[OF \'] note \_cont = has_real_derivative_imp_continuous_on[OF \'] from local_lipschitz_compact_bicomposeE[OF ll cf compact_Icc abX \_cont \_cont \_in \_in] obtain L where L: "L > 0" "\x. x \ {a..b} \ dist (f x (\ x)) (f x (\ x)) \ L * dist (\ x) (\ x)" by blast define w where "w x = \ x - \ x" for x have w'[derivative_intros]: "\x. x \ {a .. b} \ (w has_real_derivative \' x - \' x) (at x)" using \' \' by (auto simp: has_vderiv_on_def w_def[abs_def] intro!: derivative_eq_intros) note w_cont[continuous_intros] = has_real_derivative_imp_continuous_on[OF w', THEN continuous_on_compose2] have "w d \ 0" if "d \ {a .. b}" for d proof (rule ccontr, unfold not_le) assume "w d < 0" let ?N = "(w -` {..0} \ {a .. d})" from \w d < 0\ that have "d \ ?N" by auto then have "?N \ {}" by auto have "closed ?N" unfolding compact_eq_bounded_closed using that by (intro conjI closed_vimage_Int) (auto intro!: continuous_intros) let ?N' = "{a0 \ {a .. d}. w ` {a0 .. d} \ {..0}}" from \w d < 0\ that have "d \ ?N'" by simp then have "?N' \ {}" by auto have "compact ?N'" unfolding compact_eq_bounded_closed proof have "?N' \ {a .. d}" using that by auto then show "bounded ?N'" by (rule bounded_subset[rotated]) simp have "w u \ 0" if "(\n. x n \ ?N')" "x \ l" "l \ u" "u \ d" for x l u proof cases assume "l = u" have "\n. x n \ ?N" using that(1) by force from closed_sequentially[OF \closed ?N\ this \x \ l\] show ?thesis by (auto simp: \l = u\) next assume "l \ u" with that have "l < u" by auto from order_tendstoD(2)[OF \x \ l\ \l < u\] obtain n where "x n < u" by (auto dest: eventually_happens) with that show ?thesis using \l < u\ by (auto dest!: spec[where x=n] simp: image_subset_iff) qed then show "closed ?N'" unfolding closed_sequential_limits by (auto simp: Lim_bounded Lim_bounded2) qed from compact_attains_inf[OF \compact ?N'\ \?N' \ {}\] obtain a0 where a0: "a \ a0" "a0 \ d" "w ` {a0..d} \ {..0}" and a0_least: "\x. a \ x \ x \ d \ w ` {x..d} \ {..0} \ a0 \ x" by auto have a0d: "{a0 .. d} \ {a .. b}" using that a0 by auto have L_w_bound: "L * w x \ \' x - \' x" if "x \ {a0 .. d}" for x proof - from set_mp[OF a0d that] have "x \ {a .. b}" . from defect[OF this] have "\' x - \' x \ dist (f x (\ x)) (f x (\ x))" by (simp add: dist_real_def) also have "\ \ L * dist (\ x) (\ x)" using \x \ {a .. b}\ by (rule L) also have "\ \ -L * w x" using \0 < L\ a0 that by (force simp add: dist_real_def abs_real_def w_def algebra_split_simps ) finally show ?thesis by simp qed - have mono: "mono_on (\x. w x * exp(-L*x)) {a0..d}" + have mono: "mono_on {a0..d} (\x. w x * exp(-L*x))" apply (rule mono_onI) apply (rule DERIV_nonneg_imp_nondecreasing, assumption) using a0d by (auto intro!: exI[where x="(\' x - \' x) * exp (- (L * x)) - exp (- (L * x)) * L * w x" for x] derivative_eq_intros L_w_bound simp: ) then have "w a0 * exp (-L * a0) \ w d * exp (-L * d)" by (rule mono_onD) (use that a0 in auto) also have "\ < 0" using \w d < 0\ by (simp add: algebra_split_simps) finally have "w a0 * exp (- L * a0) < 0" . then have "w a0 < 0" by (simp add: algebra_split_simps) have "a0 \ a" proof (rule ccontr, unfold not_le) assume "a < a0" have "continuous_on {a.. a0} w" by (rule continuous_intros, assumption) (use a0 a0d in auto) from continuous_on_Icc_at_leftD[OF this \a < a0\] have "(w \ w a0) (at_left a0)" . from order_tendstoD(2)[OF this \w a0 < 0\] have "\\<^sub>F x in at_left a0. w x < 0" . moreover have "\\<^sub>F x in at_left a0. a < x" by (rule order_tendstoD) (auto intro!: \a < a0\) ultimately have "\\<^sub>F x in at_left a0. a < x \ w x < 0" by eventually_elim auto then obtain a1' where "a1'y. y > a1' \ y < a0 \ a < y \ w y < 0" unfolding eventually_at_left_field by auto define a1 where "a1 = (a1' + a0)/2" have "a1 < a0" using \a1' < a0\ by (auto simp: a1_def) have "a \ a1" using \a < a0\ a1_neg by (force simp: a1_def) moreover have "a1 \ d" using \a1' < a0\ a0(2) by (auto simp: a1_def) moreover have "w ` {a1..a0} \ {..0}" using \w a0 < 0\ a1_neg a0(3) by (auto simp: a1_def) smt moreover have "w ` {a0..d} \ {..0}" using a0 by auto ultimately have "a0 \ a1" apply (intro a0_least) apply assumption apply assumption by (smt atLeastAtMost_iff image_subset_iff) with \a1 show False by simp qed then have "a0 = a" using \a \ a0\ by simp with \w a0 < 0\ have "w a < 0" by simp with init show False by (auto simp: w_def) qed then show ?thesis by (auto simp: w_def) qed lemma local_lipschitz_mult: shows "local_lipschitz (UNIV::real set) (UNIV::real set) (*)" apply (auto intro!: c1_implies_local_lipschitz[where f'="\p. blinfun_mult_left (fst p)"]) apply (simp add: has_derivative_mult_right mult_commute_abs) by (auto intro!: continuous_intros) lemma comparison_principle_le_linear: fixes \ :: "real \ real" assumes "continuous_on {a..b} g" assumes "(\t. t \ {a..b} \ (\ has_real_derivative \' t) (at t))" assumes "\ a \ 0" assumes "(\t. t \ {a..b} \ \' t \ g t *\<^sub>R \ t)" shows "\t\{a..b}. \ t \ 0" proof - have *: "\x. continuous_on {a..b} (\t. g t * x)" using assms(1) continuous_on_mult_right by blast then have "local_lipschitz (g`{a..b}) UNIV (*)" using local_lipschitz_subset[OF local_lipschitz_mult] by blast from local_lipschitz_compose1[OF this assms(1)] have "local_lipschitz {a..b} UNIV (\t. (*) (g t))" . from comparison_principle_le[OF this _ _ assms(2) _ _ _ assms(3), of b "\t.0"] * assms(4) show ?thesis by auto qed subsection \Locally Lipschitz ODEs\ context ll_on_open_it begin lemma flow_lipschitzE: assumes "{a .. b} \ existence_ivl t0 x" obtains L where "L-lipschitz_on {a .. b} (flow t0 x)" proof - have f': "(flow t0 x has_derivative (\i. i *\<^sub>R f t (flow t0 x t))) (at t within {a .. b})" if "t \ {a .. b}" for t using flow_has_derivative[of t x] assms that by (auto simp: has_derivative_at_withinI) have "compact ((\t. f t (flow t0 x t)) ` {a .. b})" using assms apply (auto intro!: compact_continuous_image continuous_intros) using local.existence_ivl_empty2 apply fastforce apply (meson atLeastAtMost_iff general.existence_ivl_subset in_mono) by (simp add: general.flow_in_domain subset_iff) then obtain C where "t \ {a .. b} \ norm (f t (flow t0 x t)) \ C" for t by (fastforce dest!: compact_imp_bounded simp: bounded_iff intro: that) then have "t \ {a..b} \ onorm (\i. i *\<^sub>R f t (flow t0 x t)) \ max 0 C" for t apply (subst onorm_scaleR_left) apply (auto simp: onorm_id max_def) by (metis diff_0_right diff_mono diff_self norm_ge_zero) from bounded_derivative_imp_lipschitz[OF f' _ this] have "(max 0 C)-lipschitz_on {a..b} (flow t0 x)" by auto then show ?thesis .. qed lemma flow_undefined0: "t \ existence_ivl t0 x \ flow t0 x t = 0" unfolding flow_def by auto lemma csols_undefined: "x \ X \ csols t0 x = {}" apply (auto simp: csols_def) using general.existence_ivl_empty2 general.existence_ivl_maximal_segment apply blast done lemmas existence_ivl_undefined = existence_ivl_empty2 end subsection \Reverse flow as Sublocale\ lemma range_preflect_0[simp]: "range (preflect 0) = UNIV" by (auto simp: preflect_def) lemma range_uminus[simp]: "range uminus = (UNIV::'a::ab_group_add set)" by auto context auto_ll_on_open begin sublocale rev: auto_ll_on_open "-f" rewrites "-(-f) = f" apply unfold_locales using auto_local_lipschitz auto_open_domain unfolding fun_Compl_def local_lipschitz_minus by auto lemma existence_ivl_eq_rev0: "existence_ivl0 y = uminus ` rev.existence_ivl0 y" for y by (auto simp: existence_ivl_eq_rev rev.existence_ivl0_def preflect_def) lemma rev_existence_ivl_eq0: "rev.existence_ivl0 y = uminus ` existence_ivl0 y" for y using uminus_uminus_image[of "rev.existence_ivl0 y"] by (simp add: existence_ivl_eq_rev0) lemma flow_eq_rev0: "flow0 y t = rev.flow0 y (-t)" for y t apply (cases "t \ existence_ivl0 y") subgoal apply (subst flow_eq_rev(2), assumption) apply (subst rev.flow0_def) by (simp add: preflect_def) subgoal apply (frule flow_undefined0) by (auto simp: existence_ivl_eq_rev0 rev.flow_undefined0) done lemma rev_eq_flow: "rev.flow0 y t = flow0 y (-t)" for y t apply (subst flow_eq_rev0) using uminus_uminus_image[of "rev.existence_ivl0 y"] apply - apply (subst (asm) existence_ivl_eq_rev0[symmetric]) by auto lemma rev_flow_image_eq: "rev.flow0 x ` S = flow0 x ` (uminus ` S)" unfolding rev_eq_flow[abs_def] by force lemma flow_image_eq_rev: "flow0 x ` S = rev.flow0 x ` (uminus ` S)" unfolding rev_eq_flow[abs_def] by force end context c1_on_open begin sublocale rev: c1_on_open "-f" "-f'" rewrites "-(-f) = f" and "-(-f') = f'" by (rule c1_on_open_rev) auto end context c1_on_open_euclidean begin sublocale rev: c1_on_open_euclidean "-f" "-f'" rewrites "-(-f) = f" and "-(-f') = f'" by unfold_locales auto end subsection \Autonomous LL ODE : Existence Interval and trapping on the interval\ lemma bdd_above_is_intervalI: "bdd_above I" if "is_interval I" "a \ b" "a \ I" "b \ I" for I::"real set" by (meson bdd_above_def is_interval_1 le_cases that) lemma bdd_below_is_intervalI: "bdd_below I" if "is_interval I" "a \ b" "a \ I" "b \ I" for I::"real set" by (meson bdd_below_def is_interval_1 le_cases that) context auto_ll_on_open begin lemma open_existence_ivl0: assumes x : "x \ X" shows "\a b. a < 0 \ 0 < b \ {a..b} \ existence_ivl0 x" proof - have a1:"0 \ existence_ivl0 x" by (simp add: x) have a2: "open (existence_ivl0 x)" by (simp add: x) from a1 a2 obtain d where "d > 0" "ball 0 d \ existence_ivl0 x" using openE by blast have "{-d/2..d/2} \ ball 0 d" using \0 < d\ dist_norm mem_ball by auto thus ?thesis by (smt \0 < d\ \ball 0 d \ existence_ivl0 x\ divide_minus_left half_gt_zero order_trans) qed lemma open_existence_ivl': assumes x : "x \ X" obtains a where "a > 0" "{-a..a} \ existence_ivl0 x" proof - from open_existence_ivl0[OF assms(1)] obtain a b where ab: "a < 0" "0 < b" "{a..b} \ existence_ivl0 x" by auto then have "min (-a) b > 0" by linarith have "{-min (-a) b .. min(-a) b} \ {a..b}" by auto thus ?thesis using ab(3) that[OF \min (-a) b > 0\] by blast qed lemma open_existence_ivl_on_compact: assumes C: "C \ X" and "compact C" "C \ {}" obtains a where "a > 0" "\x. x \ C \ {-a..a} \ existence_ivl0 x" proof - from existence_ivl_cballs have "\x\C. \e>0. \t>0. \y\cball x e. cball 0 t\existence_ivl0 y" by (metis (full_types) C Int_absorb1 Int_iff UNIV_I) then obtain d' t' where *: "\x\C. 0 < d' x \ t' x > 0 \ (\y\cball x (d' x). cball 0 (t' x) \ existence_ivl0 y)" by metis with compactE_image[OF \compact C\, of C "\x. ball x (d' x)"] obtain C' where "C' \ C" and [simp]: "finite C'" and C_subset: "C \ (\c\C'. ball c (d' c))" by force from C_subset \C \ {}\ have [simp]: "C' \ {}" by auto define d where "d = Min (d' ` C')" define t where "t = Min (t' ` C')" have "t > 0" using * \C' \ C\ by (auto simp: t_def) moreover have "{-t .. t} \ existence_ivl0 x" if "x \ C" for x proof - from C_subset that \C' \ C\ obtain c where c: "c \ C'" "x \ ball c (d' c)" "c \ C" by force then have "{-t .. t} \ cball 0 (t' c)" by (auto simp: abs_real_def t_def minus_le_iff) also from c have "cball 0 (t' c) \ existence_ivl0 x" using *[rule_format, OF \c \ C\] by auto finally show ?thesis . qed ultimately show ?thesis .. qed definition "trapped_forward x K \ (flow0 x ` (existence_ivl0 x \ {0..}) \ K)" \ \TODO: use this for backwards trapped, invariant, and all assumptions\ definition "trapped_backward x K \ (flow0 x ` (existence_ivl0 x \ {..0}) \ K)" definition "trapped x K \ trapped_forward x K \ trapped_backward x K" lemma trapped_iff_on_existence_ivl0: "trapped x K \ (flow0 x ` (existence_ivl0 x) \ K)" unfolding trapped_def trapped_forward_def trapped_backward_def apply (auto) by (metis IntI atLeast_iff atMost_iff image_subset_iff less_eq_real_def linorder_not_less) end context auto_ll_on_open begin lemma infinite_rev_existence_ivl0_rewrites: "{0..} \ rev.existence_ivl0 x \ {..0} \ existence_ivl0 x" "{..0} \ rev.existence_ivl0 x \ {0..} \ existence_ivl0 x" apply (auto simp add: rev.rev_existence_ivl_eq0 subset_iff) using neg_le_0_iff_le apply fastforce using neg_0_le_iff_le by fastforce lemma trapped_backward_iff_rev_trapped_forward: "trapped_backward x K \ rev.trapped_forward x K" unfolding trapped_backward_def rev.trapped_forward_def by (auto simp add: rev_flow_image_eq existence_ivl_eq_rev0 image_subset_iff) text \If solution is trapped in a compact set at some time on its existence interval then it is trapped forever\ lemma trapped_sol_right: \ \TODO: when building on afp-devel (??? outdated): \<^url>\https://bitbucket.org/isa-afp/afp-devel/commits/0c3edf9248d5389197f248c723b625c419e4d3eb\\ assumes "compact K" "K \ X" assumes "x \ X" "trapped_forward x K" shows "{0..} \ existence_ivl0 x" proof (rule ccontr) assume "\ {0..} \ existence_ivl0 x" from this obtain t where "0 \ t" "t \ existence_ivl0 x" by blast then have bdd: "bdd_above (existence_ivl0 x)" by (auto intro!: bdd_above_is_intervalI \x \ X\) from flow_leaves_compact_ivl_right [OF UNIV_I \x \ X\ bdd UNIV_I assms(1-2)] show False by (metis assms(4) trapped_forward_def IntI atLeast_iff image_subset_iff) qed lemma trapped_sol_right_gen: assumes "compact K" "K \ X" assumes "t \ existence_ivl0 x" "trapped_forward (flow0 x t) K" shows "{t..} \ existence_ivl0 x" proof - have "x \ X" using assms(3) local.existence_ivl_empty_iff by fastforce have xtk: "flow0 x t \ X" by (simp add: assms(3) local.flow_in_domain) from trapped_sol_right[OF assms(1-2) xtk assms(4)] have "{0..} \ existence_ivl0 (flow0 x t)" . thus "{t..} \ existence_ivl0 x" using existence_ivl_trans[OF assms(3)] by (metis add.commute atLeast_iff diff_add_cancel le_add_same_cancel1 subset_iff) qed lemma trapped_sol_left: \ \TODO: when building on afp-devel: \<^url>\https://bitbucket.org/isa-afp/afp-devel/commits/0c3edf9248d5389197f248c723b625c419e4d3eb\\ assumes "compact K" "K \ X" assumes "x \ X" "trapped_backward x K" shows "{..0} \ existence_ivl0 x" proof (rule ccontr) assume "\ {..0} \ existence_ivl0 x" from this obtain t where "t \ 0" "t \ existence_ivl0 x" by blast then have bdd: "bdd_below (existence_ivl0 x)" by (auto intro!: bdd_below_is_intervalI \x \ X\) from flow_leaves_compact_ivl_left [OF UNIV_I \x \ X\ bdd UNIV_I assms(1-2)] show False by (metis IntI assms(4) atMost_iff auto_ll_on_open.trapped_backward_def auto_ll_on_open_axioms image_subset_iff) qed lemma trapped_sol_left_gen: assumes "compact K" "K \ X" assumes "t \ existence_ivl0 x" "trapped_backward (flow0 x t) K" shows "{..t} \ existence_ivl0 x" proof - have "x \ X" using assms(3) local.existence_ivl_empty_iff by fastforce have xtk: "flow0 x t \ X" by (simp add: assms(3) local.flow_in_domain) from trapped_sol_left[OF assms(1-2) xtk assms(4)] have "{..0} \ existence_ivl0 (flow0 x t)" . thus "{..t} \ existence_ivl0 x" using existence_ivl_trans[OF assms(3)] by (metis add.commute add_le_same_cancel1 atMost_iff diff_add_cancel subset_eq) qed lemma trapped_sol: assumes "compact K" "K \ X" assumes "x \ X" "trapped x K" shows "existence_ivl0 x = UNIV" by (metis (mono_tags, lifting) assms existence_ivl_zero image_subset_iff interval local.existence_ivl_initial_time_iff local.existence_ivl_subset local.subset_mem_compact_implies_subset_existence_interval order_refl subset_antisym trapped_iff_on_existence_ivl0) (* Follows from rectification *) lemma regular_locally_noteq:\ \TODO: should be true in \ll_on_open_it\\ assumes "x \ X" "f x \ 0" shows "eventually (\t. flow0 x t \ x) (at 0)" proof - have nf:"norm (f x) > 0" by (simp add: assms(2)) (* By continuity of solutions and f probably *) obtain a where a: "a>0" "{-a--a} \ existence_ivl0 x" "0 \ {-a--a}" "\t. t \ {-a--a} \ norm(f (flow0 x t) - f (flow0 x 0)) \ norm(f x)/2" proof - from open_existence_ivl'[OF assms(1)] obtain a1 where a1: "a1 > 0" "{-a1..a1} \ existence_ivl0 x" . have "continuous (at 0) (\t. norm(f (flow0 x t) - f (flow0 x 0) ))" apply (auto intro!: continuous_intros) by (simp add: assms(1) local.f_flow_continuous) then obtain a2 where "a2>0" "\t. norm t < a2 \ norm (f (flow0 x t) - f (flow0 x 0)) < norm(f x)/2" unfolding continuous_at_real_range by (metis abs_norm_cancel cancel_comm_monoid_add_class.diff_cancel diff_zero half_gt_zero nf norm_zero) then have t: "\t. t \ {-a2<-- norm(f (flow0 x t) - f (flow0 x 0)) \ norm(f x)/2" by (smt open_segment_bound(2) open_segment_bound1 real_norm_def) define a where "a = min a1 (a2/2)" have t1:"a > 0" unfolding a_def using \a1 > 0\ \a2 > 0\ by auto then have t3:"0 \{-a--a}" using closed_segment_eq_real_ivl by auto have "{-a--a} \ {-a1..a1}" unfolding a_def using \a1 > 0\ \a2 > 0\ using ODE_Auxiliarities.closed_segment_eq_real_ivl by auto then have t2:"{-a--a} \ existence_ivl0 x" using a1 by auto have "{-a--a} \ {-a2<--a1 > 0\ \a2 > 0\ by (smt Diff_iff closed_segment_eq_real_ivl atLeastAtMost_iff empty_iff half_gt_zero insert_iff pos_half_less segment(1) subset_eq) then have t4:"\t. t \ {-a--a} \ norm(f (flow0 x t) - f (flow0 x 0)) \ norm(f x)/2" using t by auto show ?thesis using t1 t2 t3 t4 that by auto qed have "\t. t \ {-a--a} \ (flow0 x has_vector_derivative f (flow0 x t)) (at t within {-a--a})" apply (rule has_vector_derivative_at_within) using a(2) by (auto intro!:flow_has_vector_derivative) from vector_differentiable_bound_linearization[OF this _ a(4)] have nb:"\c d. {c--d} \ {-a--a} \ norm (flow0 x d - flow0 x c - (d - c) *\<^sub>R f (flow0 x 0)) \ norm (d - c) * (norm (f x) / 2)" using a(3) by blast have "\t. dist t 0 < a \ t \ 0 \ flow0 x t \ x" proof (rule ccontr) fix t assume "dist t 0 < a" "t \ 0" "\ flow0 x t \ x" then have tx:"flow0 x t = x" by auto have "t \ {-a--a}" using closed_segment_eq_real_ivl \dist t 0 < a\ by auto have "t > 0 \ t < 0" using \t \ 0\ by linarith moreover { assume "t > 0" then have "{0--t} \ {-a--a}" by (simp add: \t \ {-a--a}\ a(3) subset_closed_segment) from nb[OF this] have "norm (flow0 x t - x - t *\<^sub>R f x) \ norm t * (norm (f x) / 2)" by (simp add: assms(1)) then have "norm (t *\<^sub>R f x) \ norm t * (norm (f x) / 2)" using tx by auto then have False using nf using \0 < t\ by auto } moreover { assume "t < 0" then have "{t--0} \ {-a--a}" by (simp add: \t \ {-a--a}\ a(3) subset_closed_segment) from nb[OF this] have "norm (x - flow0 x t + t *\<^sub>R f x) \ norm t * (norm (f x) / 2)" by (simp add: assms(1)) then have "norm (t *\<^sub>R f x) \ norm t * (norm (f x) / 2)" using tx by auto then have False using nf using \t < 0\ by auto } ultimately show False by blast qed thus ?thesis unfolding eventually_at using a(1) by blast qed lemma compact_max_time_flow_in_closed: assumes "closed M" and t_ex: "t \ existence_ivl0 x" shows "compact {s \ {0..t}. flow0 x ` {0..s} \ M}" (is "compact ?C") unfolding compact_eq_bounded_closed proof have "bounded {0 .. t}" by auto then show "bounded ?C" by (rule bounded_subset) auto show "closed ?C" unfolding closed_def proof (rule topological_space_class.openI, clarsimp) \ \TODO: there must be a more abstract argument for this, e.g., with @{thm continuous_on_closed_vimageI} and then reasoning about the connected component around 0?\ fix s assume notM: "s \ t \ 0 \ s \ \ flow0 x ` {0..s} \ M" consider "0 \ s" "s \ t" "flow0 x s \ M" | "0 \ s" "s \ t" "flow0 x s \ M" | "s < 0" | "s > t" by arith then show "\T. open T \ s \ T \ T \ - {s. 0 \ s \ s \ t \ flow0 x ` {0..s} \ M}" proof cases assume s: "0 \ s" "s \ t" and sM: "flow0 x s \ M" have "isCont (flow0 x) s" using s ivl_subset_existence_ivl[OF t_ex] by (auto intro!: flow_continuous) from this[unfolded continuous_at_open, rule_format, of "-M"] sM \closed M\ obtain S where "open S" "s \ S" "(\x'\S. flow0 x x' \ - M)" by auto then show ?thesis by (force intro!: exI[where x=S]) next assume s: "0 \ s" "s \ t" and sM: "flow0 x s \ M" from this notM obtain s0 where s0: "0 \ s0" "s0 < s" "flow0 x s0 \ M" by force from order_tendstoD(1)[OF tendsto_ident_at \s0 < s\, of UNIV, unfolded eventually_at_topological] obtain S where "open S" "s \ S" "\x. x \ S \ x \ s \ s0 < x" by auto then show ?thesis using s0 by (auto simp: intro!: exI[where x=S]) (smt atLeastAtMost_iff image_subset_iff) qed (force intro: exI[where x="{t<..}"] exI[where x="{..<0}"])+ qed qed lemma flow_in_closed_max_timeE: assumes "closed M" "t \ existence_ivl0 x" "0 \ t" "x \ M" obtains T where "0 \ T" "T \ t" "flow0 x ` {0..T} \ M" "\s'. 0 \ s' \ s' \ t \ flow0 x ` {0..s'} \ M \ s' \ T" proof - let ?C = "{s \ {0..t}. flow0 x ` {0..s} \ M}" have "?C \ {}" using assms using local.mem_existence_ivl_iv_defined by (auto intro!: exI[where x=0]) from compact_max_time_flow_in_closed[OF assms(1,2)] have "compact ?C" . from compact_attains_sup[OF this \?C \ {}\] obtain s where s: "0 \ s" "s \ t" "flow0 x ` {0..s} \ M" and s_max: "\s'. 0 \ s' \ s' \ t \ flow0 x ` {0..s'} \ M \ s' \ s" by auto then show ?thesis .. qed lemma flow_leaves_closed_at_frontierE: assumes "closed M" and t_ex: "t \ existence_ivl0 x" and "0 \ t" "x \ M" "flow0 x t \ M" obtains s where "0 \ s" "s < t" "flow0 x ` {0..s} \ M" "flow0 x s \ frontier M" "\\<^sub>F s' in at_right s. flow0 x s' \ M" proof - from flow_in_closed_max_timeE[OF assms(1-4)] assms(5) obtain s where s: "0 \ s" "s < t" "flow0 x ` {0..s} \ M" and s_max: "\s'. 0 \ s' \ s' \ t \ flow0 x ` {0..s'} \ M \ s' \ s" by (smt atLeastAtMost_iff image_subset_iff) note s moreover have "flow0 x s \ interior M" proof assume interior: "flow0 x s \ interior M" have "s \ existence_ivl0 x" using ivl_subset_existence_ivl[OF \t \ _\] s by auto from flow_continuous[OF this, THEN isContD, THEN topological_tendstoD, OF open_interior interior] have "\\<^sub>F s' in at s. flow0 x s' \ interior M" by auto then have "\\<^sub>F s' in at_right s. flow0 x s' \ interior M" by (auto simp: eventually_at_split) moreover have "\\<^sub>F s' in at_right s. s' < t" using tendsto_ident_at \s < t\ by (rule order_tendstoD) ultimately have "\\<^sub>F s' in at_right s. flow0 x s' \ M \ s' < t" by eventually_elim (use interior_subset[of M] in auto) then obtain s' where s': "s < s'" "s' < t" "\y. y > s \ y \ s' \ flow0 x y \ M" by (auto simp: eventually_at_right_field_le) have s'_ivl: "flow0 x ` {0..s'} \ M" proof safe fix s'' assume "s'' \ {0 .. s'}" then show "flow0 x s'' \ M" using s interior_subset[of M] s' by (cases "s'' \ s") auto qed with s_max[of s'] \s' < t\ \0 \ s\ \s < s'\ show False by auto qed then have "flow0 x s \ frontier M" using s closure_subset[of M] by (force simp: frontier_def) moreover have "compact (flow0 x -` M \ {s..t})" (is "compact ?C") unfolding compact_eq_bounded_closed proof have "bounded {s .. t}" by simp then show "bounded ?C" by (rule bounded_subset) auto show "closed ?C" using \closed M\ assms mem_existence_ivl_iv_defined(2)[OF t_ex] ivl_subset_existence_ivl[OF t_ex] \0 \ s\ by (intro closed_vimage_Int) (auto intro!: continuous_intros) qed have "\\<^sub>F s' in at_right s. flow0 x s' \ M" apply (rule ccontr) unfolding not_frequently proof - assume "\\<^sub>F s' in at_right s. \ flow0 x s' \ M" moreover have "\\<^sub>F s' in at_right s. s' < t" using tendsto_ident_at \s < t\ by (rule order_tendstoD) ultimately have "\\<^sub>F s' in at_right s. flow0 x s' \ M \ s' < t" by eventually_elim auto then obtain s' where s': "s < s'" "\y. y > s \ y < s' \ flow0 x y \ M" "\y. y > s \ y < s' \ y < t" by (auto simp: eventually_at_right_field) define s'' where "s'' = (s + s') / 2" have "0 \ s''" "s'' \ t" "s < s''" "s'' < s'" using s s' by (auto simp del: divide_le_eq_numeral1 le_divide_eq_numeral1 simp: s''_def) fastforce then have "flow0 x ` {0..s''} \ M" using s s' apply auto subgoal for u by (cases "u\s") auto done from s_max[OF \0 \ s''\ \s''\ t\ this] \s'' > s\ show False by simp qed ultimately show ?thesis .. qed subsection \Connectedness\ lemma fcontX: shows "continuous_on X f" using auto_local_lipschitz local_lipschitz_continuous_on by blast lemma fcontx: assumes "x \ X" shows "continuous (at x) f" proof - have "open X" by simp from continuous_on_eq_continuous_at[OF this] show ?thesis using fcontX assms(1) by blast qed lemma continuous_at_imp_cball: assumes "continuous (at x) g" assumes "g x > (0::real)" obtains r where "r > 0" "\y \ cball x r. g y > 0" proof - from assms(1) obtain d where "d>0" "g ` (ball x d) \ ball (g x) ((g x)/2)" by (meson assms(2) continuous_at_ball half_gt_zero) then have "\y \ cball x (d/2). g y > 0" by (smt assms(2) dist_norm image_subset_iff mem_ball mem_cball pos_half_less real_norm_def) thus ?thesis using \0 < d\ that half_gt_zero by blast qed text \ \flow0\ is \path_connected\ \ lemma flow0_path_connected_time: assumes "ts \ existence_ivl0 x" "path_connected ts" shows "path_connected (flow0 x ` ts)" proof - have "continuous_on ts (flow0 x)" by (meson assms continuous_on_sequentially flow_continuous_on subsetD) from path_connected_continuous_image[OF this assms(2)] show ?thesis . qed lemma flow0_path_connected: assumes "path_connected D" "path_connected ts" "\x. x \ D \ ts \ existence_ivl0 x" shows "path_connected ( (\(x, y). flow0 x y) ` (D \ ts))" proof - have "D \ ts \ Sigma X existence_ivl0" using assms(3) subset_iff by fastforce then have a1:"continuous_on (D \ ts) (\(x, y). flow0 x y)" using flow_continuous_on_state_space continuous_on_subset by blast have a2 : "path_connected (D \ ts)" using path_connected_Times assms by auto from path_connected_continuous_image[OF a1 a2] show ?thesis . qed end subsection \Return Time and Implicit Function Theorem\ context c1_on_open_euclidean begin lemma flow_implicit_function: \ \TODO: generalization of @{thm returns_to_implicit_function}!\ fixes s::"'a::euclidean_space \ real" and S::"'a set" assumes t: "t \ existence_ivl0 x" and x: "x \ X" and st: "s (flow0 x t) = 0" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" assumes DsC: "isCont Ds (flow0 x t)" assumes nz: "Ds (flow0 x t) (f (flow0 x t)) \ 0" obtains u e where "s (flow0 x (u x)) = 0" "u x = t" "(\y. y \ cball x e \ s (flow0 y (u y)) = 0)" "continuous_on (cball x e) u" "(\t. (t, u t)) ` cball x e \ Sigma X existence_ivl0" "0 < e" "(u has_derivative (- blinfun_scaleR_left (inverse (blinfun_apply (Ds (flow0 x t)) (f (flow0 x t)))) o\<^sub>L (Ds (flow0 x t) o\<^sub>L flowderiv x t) o\<^sub>L embed1_blinfun)) (at x)" proof - note [derivative_intros] = has_derivative_compose[OF _ Ds] have cont_s: "continuous_on UNIV s" by (rule has_derivative_continuous_on[OF Ds]) note cls[simp, intro] = closed_levelset[OF cont_s] then have xt1: "(x, t) \ Sigma X existence_ivl0" by (auto simp: t x) have D: "(\x. x \ Sigma X existence_ivl0 \ ((\(x, t). s (flow0 x t)) has_derivative blinfun_apply (Ds (flow0 (fst x) (snd x)) o\<^sub>L (flowderiv (fst x) (snd x)))) (at x))" by (auto intro!: derivative_eq_intros) have C: "isCont (\x. Ds (flow0 (fst x) (snd x)) o\<^sub>L flowderiv (fst x) (snd x)) (x, t)" using flowderiv_continuous_on[unfolded continuous_on_eq_continuous_within, rule_format, OF xt1] using at_within_open[OF xt1 open_state_space] by (auto intro!: continuous_intros tendsto_eq_intros x t isCont_tendsto_compose[OF DsC, unfolded poincare_map_def] simp: split_beta' isCont_def) have Z: "(case (x, t) of (x, t) \ s (flow0 x t)) = 0" by (auto simp: st) have I1: "blinfun_scaleR_left (inverse (Ds (flow0 x t)(f (flow0 x t)))) o\<^sub>L ((Ds (flow0 (fst (x, t)) (snd (x, t))) o\<^sub>L flowderiv (fst (x, t)) (snd (x, t))) o\<^sub>L embed2_blinfun) = 1\<^sub>L" using nz by (auto intro!: blinfun_eqI simp: flowderiv_def blinfun.bilinear_simps inverse_eq_divide poincare_map_def) have I2: "((Ds (flow0 (fst (x, t)) (snd (x, t))) o\<^sub>L flowderiv (fst (x, t)) (snd (x, t))) o\<^sub>L embed2_blinfun) o\<^sub>L blinfun_scaleR_left (inverse (Ds (flow0 x t)(f (flow0 x t)))) = 1\<^sub>L" using nz by (auto intro!: blinfun_eqI simp: flowderiv_def blinfun.bilinear_simps inverse_eq_divide poincare_map_def) show ?thesis apply (rule implicit_function_theorem[where f="\(x, t). s (flow0 x t)" and S="Sigma X existence_ivl0", OF D xt1 open_state_space order_refl C Z I1 I2]) apply blast unfolding split_beta' fst_conv snd_conv poincare_map_def[symmetric] .. qed lemma flow_implicit_function_at: fixes s::"'a::euclidean_space \ real" and S::"'a set" assumes x: "x \ X" and st: "s x = 0" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" assumes DsC: "isCont Ds x" assumes nz: "Ds x (f x) \ 0" assumes pos: "e > 0" obtains u d where "0 < d" "u x = 0" "\y. y \ cball x d \ s (flow0 y (u y)) = 0" "\y. y \ cball x d \ \u y\ < e" "\y. y \ cball x d \ u y \ existence_ivl0 y" "continuous_on (cball x d) u" "(u has_derivative -Ds x /\<^sub>R (Ds x) (f x)) (at x)" proof - have x0: "flow0 x 0 = x" by (simp add: x) from flow_implicit_function[OF existence_ivl_zero[OF x] x, unfolded x0, of s, OF st Ds DsC nz] obtain u d0 where s0: "s (flow0 x (u x)) = 0" and u0: "u x = 0" and u: "\y. y \ cball x d0 \ s (flow0 y (u y)) = 0" and uc: "continuous_on (cball x d0) u" and uex: "(\t. (t, u t)) ` cball x d0 \ Sigma X existence_ivl0" and d0: "0 < d0" and u': "(u has_derivative blinfun_apply (- blinfun_scaleR_left (inverse (blinfun_apply (Ds x) (f x))) o\<^sub>L (Ds x o\<^sub>L flowderiv x 0) o\<^sub>L embed1_blinfun)) (at x)" by blast have "at x within cball x d0 = at x" by (rule at_within_interior) (auto simp: \0 < d0\) then have "(u \ 0) (at x)" using uc d0 by (auto simp: continuous_on_def u0 dest!: bspec[where x=x]) from tendstoD[OF this \0 < e\] pos u0 obtain d1 where d1: "0 < d1" "\xa. dist xa x \ d1 \ \u xa\ < e" unfolding eventually_at_le by force define d where "d = min d0 d1" have "0 < d" by (auto simp: d_def d0 d1) moreover note u0 moreover have "\y. y \ cball x d \ s (flow0 y (u y)) = 0" by (auto intro!: u simp: d_def) moreover have "\y. y \ cball x d \ \u y\ < e" using d1 by (auto simp: d_def dist_commute) moreover have "\y. y \ cball x d \ u y \ existence_ivl0 y" using uex by (force simp: d_def) moreover have "continuous_on (cball x d) u" using uc by (rule continuous_on_subset) (auto simp: d_def) moreover have "(u has_derivative -Ds x /\<^sub>R (Ds x) (f x)) (at x)" using u' by (rule has_derivative_subst) (auto intro!: ext simp: x x0 flowderiv_def blinfun.bilinear_simps) ultimately show ?thesis .. qed lemma returns_to_implicit_function_gen: \ \TODO: generalizes proof of @{thm returns_to_implicit_function}!\ fixes s::"'a::euclidean_space \ real" assumes rt: "returns_to {x \ S. s x = 0} x" (is "returns_to ?P x") assumes cS: "closed S" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" "isCont Ds (poincare_map ?P x)" "Ds (poincare_map ?P x) (f (poincare_map ?P x)) \ 0" obtains u e where "s (flow0 x (u x)) = 0" "u x = return_time ?P x" "(\y. y \ cball x e \ s (flow0 y (u y)) = 0)" "continuous_on (cball x e) u" "(\t. (t, u t)) ` cball x e \ Sigma X existence_ivl0" "0 < e" "(u has_derivative (- blinfun_scaleR_left (inverse (blinfun_apply (Ds (poincare_map ?P x)) (f (poincare_map ?P x)))) o\<^sub>L (Ds (poincare_map ?P x) o\<^sub>L flowderiv x (return_time ?P x)) o\<^sub>L embed1_blinfun)) (at x)" proof - note [derivative_intros] = has_derivative_compose[OF _ Ds(1)] have cont_s: "continuous_on UNIV s" by (rule has_derivative_continuous_on[OF Ds(1)]) note cls[simp, intro] = closed_levelset[OF cont_s] let ?t1 = "return_time ?P x" have cls[simp, intro]: "closed {x \ S. s x = 0}" by (rule closed_levelset_within) (auto intro!: cS continuous_on_subset[OF cont_s]) have *: "poincare_map ?P x = flow0 x (return_time {x \ S. s x = 0} x)" by (simp add: poincare_map_def) have "return_time {x \ S. s x = 0} x \ existence_ivl0 x" "x \ X" "s (poincare_map ?P x) = 0" using poincare_map_returns rt by (auto intro!: return_time_exivl rt) note E = flow_implicit_function[of "return_time ?P x" x s Ds, OF this[unfolded *] Ds[unfolded *], folded *] show ?thesis by (rule E) rule qed text \c.f. Perko Section 3.7 Lemma 2 part 1.\ lemma flow_transversal_surface_finite_intersections: fixes s::"'a \ 'b::real_normed_vector" and Ds::"'a \ 'a \\<^sub>L 'b" assumes "closed S" assumes "\x. (s has_derivative (Ds x)) (at x)" assumes "\x. x \ S \ s x = 0 \ Ds x (f x) \ 0" assumes "a \ b" "{a .. b} \ existence_ivl0 x" shows "finite {t\{a..b}. flow0 x t \ {x \ S. s x = 0}}" \ \TODO: define notion of (compact/closed)-(continuous/differentiable/C1)-surface?\ proof cases note Ds = \\x. (s has_derivative (Ds x)) (at x)\ note transversal = \\x. x \ S \ s x = 0 \ Ds x (f x) \ 0\ assume "a < b" show ?thesis proof (rule ccontr) let ?S = "{x \ S. s x = 0}" let ?T = "{t\{a..b}. flow0 x t \ {x \ S. s x = 0}}" define \ where "\ = flow0 x" have [THEN continuous_on_compose2, continuous_intros]: "continuous_on S s" by (auto simp: intro!: has_derivative_continuous_on Ds intro: has_derivative_at_withinI) assume "infinite ?T" from compact_sequentialE[OF compact_Icc[of a b] this] obtain t tl where t: "t n \ ?T" "flow0 x (t n) \ ?S" "t n \ {a .. b}" "t n \ tl" and tl: "t \ tl" "tl \ {a..b}" for n by force have tl_ex: "tl \ existence_ivl0 x" using \{a .. b} \ existence_ivl0 x\ \tl \ {a .. b}\ by auto have "closed ?S" by (auto intro!: closed_levelset_within \closed S\ continuous_intros) moreover have "\n. flow0 x (t n) \ ?S" using t by auto moreover have flow_t: "(\n. flow0 x (t n)) \ flow0 x tl" by (auto intro!: tendsto_eq_intros tl_ex tl) ultimately have "flow0 x tl \ ?S" by (rule closed_sequentially) let ?qt = "\t. (flow0 x t - flow0 x tl) /\<^sub>R (t - tl)" from flow_has_vector_derivative[OF tl_ex, THEN has_vector_derivative_withinD] have qt_tendsto: "?qt \tl\ f (flow0 x tl)" . let ?q = "\n. ?qt (t n)" have "filterlim t (at tl) sequentially" using tl(1) by (rule filterlim_atI) (simp add: t) with qt_tendsto have "?q \ f (flow0 x tl)" by (rule filterlim_compose) then have "((\n. Ds (flow0 x tl) (?q n))) \ Ds (flow0 x tl) (f (flow0 x tl))" by (auto intro!: tendsto_intros) moreover from flow_lipschitzE[OF \{a .. b} \ existence_ivl0 x\] obtain L' where L': "L'-lipschitz_on {a..b} (flow0 x)" . define L where "L = L' + 1" from lipschitz_on_le[OF L', of L] lipschitz_on_nonneg[OF L'] have L: "L-lipschitz_on {a .. b} (flow0 x)" and "L > 0" by (auto simp: L_def) from flow_lipschitzE[OF \{a .. b} \ existence_ivl0 x\] obtain L' where "L'-lipschitz_on {a..b} (flow0 x)" . \ \TODO: is this reasoning (below) with this Lipschitz constant really necessary?\ have s[simp]: "s (flow0 x (t n)) = 0""s (flow0 x tl) = 0" for n using t \flow0 x tl \ ?S\ by auto from Ds(1)[of "flow0 x tl", unfolded has_derivative_within] have "(\y. (1 / norm (y - flow0 x tl)) *\<^sub>R (s y - (s (flow0 x tl) + blinfun_apply (Ds (flow0 x tl)) (y - flow0 x tl)))) \flow0 x tl\ 0" by auto then have "((\y. (1 / norm (y - flow0 x tl)) *\<^sub>R (s y - (s (flow0 x tl) + blinfun_apply (Ds (flow0 x tl)) (y - flow0 x tl)))) \ 0) (nhds (flow0 x tl))" by (rule tendsto_nhds_continuousI) simp from filterlim_compose[OF this flow_t] have "(\xa. (blinfun_apply (Ds (flow0 x tl)) (flow0 x (t xa) - flow0 x tl)) /\<^sub>R norm (flow0 x (t xa) - flow0 x tl)) \ 0" using t by (auto simp: inverse_eq_divide tendsto_minus_cancel_right) from tendsto_mult[OF tendsto_const[of "L"] tendsto_norm[OF this, simplified, simplified divide_inverse_commute[symmetric]]]\ \TODO: uuugly\ have Ds0: "(\xa. norm (blinfun_apply (Ds (flow0 x tl)) (flow0 x (t xa) - flow0 x tl)) / (norm (flow0 x (t xa) - flow0 x tl)/(L))) \ 0" by (auto simp: ac_simps) from _ Ds0 have "((\n. Ds (flow0 x tl) (?q n)) \ 0)" apply (rule Lim_null_comparison) apply (rule eventuallyI) unfolding norm_scaleR blinfun.scaleR_right abs_inverse divide_inverse_commute[symmetric] subgoal for n apply (cases "flow0 x (t n) = flow0 x tl") subgoal by (simp add: blinfun.bilinear_simps) subgoal apply (rule divide_left_mono) using lipschitz_onD[OF L, of "t n" tl] \0 < L\ t(3) tl(2) by (auto simp: algebra_split_simps zero_less_divide_iff dist_norm pos_divide_le_eq intro!: add_pos_nonneg) done done ultimately have "Ds (flow0 x tl) (f (flow0 x tl)) = 0" by (rule LIMSEQ_unique) moreover have "Ds (flow0 x tl) (f (flow0 x tl)) \ 0" by (rule transversal) (use \flow0 x tl \ ?S\ in auto) ultimately show False by auto qed qed (use assms in auto) lemma uniform_limit_flow0_state:\ \TODO: is that something more general?\ assumes "compact C" assumes "C \ X" shows "uniform_limit C (\s x. flow0 x s) (\x. flow0 x 0) (at 0)" proof (cases "C = {}") case True then show ?thesis by auto next case False show ?thesis proof (rule uniform_limitI) fix e::real assume "0 < e" { fix x assume "x \ C" with assms have "x \ X" by auto from existence_ivl_cballs[OF UNIV_I \x \ X\] obtain t L u where "\y. y \ cball x u \ cball 0 t \ existence_ivl0 y" "\s y. y \ cball x u \ s \ cball 0 t \ flow0 y s \ cball y u" "L-lipschitz_on (cball 0 t\cball x u) (\(t, x). flow0 x t)" "\y. y \ cball x u \ cball y u \ X" "0 < t" "0 < u" by metis then have "\L. \u>0. \t>0. L-lipschitz_on (cball 0 t\cball x u) (\(t, x). flow0 x t)" by blast } then have "\x\C. \L. \u>0. \t>0. L-lipschitz_on (cball 0 t\cball x u) (\(t, x). flow0 x t)" .. then obtain L d' u' where L: "\x. x \ C \ (L x)-lipschitz_on (cball 0 (d' x)\cball x (u' x)) (\(t, x). flow0 x t)" and d': "\x. x \ C \ d' x > 0" and u': "\x. x \ C \ u' x > 0" by metis have "C \ (\c\C. ball c (u' c))" using u' by auto from compactE_image[OF \compact C\ _ this] obtain C' where "C' \ C" and [simp]: "finite C'" and C'_cover: "C \ (\c\C'. ball c (u' c))" by auto from C'_cover obtain c' where c': "x \ C \ x \ ball (c' x) (u' (c' x))" "x \ C \ c' x \ C'" for x by (auto simp: subset_iff) metis have "\\<^sub>F s in at 0. \x\ball c (u' c). dist (flow0 x s) (flow0 x 0) < e" if "c \ C'" for c proof - have cC: "c \ C" using c' \c \ C'\ d' \C' \ C\ by auto have *: "dist (flow0 x s) (flow0 x 0) \ L c * \s\" if "x\ball c (u' c)" "s \ cball 0 (d' c)" for x s proof - from L[OF cC, THEN lipschitz_onD, of "(0, x)" "(s, x)"] d'[OF cC] that show ?thesis by (auto simp: dist_prod_def dist_commute) qed have "\\<^sub>F s in at 0. abs s < d' c" by (rule order_tendstoD tendsto_intros)+ (use d' cC in auto) moreover have "\\<^sub>F s in at 0. L c * \s\ < e" by (rule order_tendstoD tendsto_intros)+ (use \0 < e\ in auto) ultimately show ?thesis apply eventually_elim apply (use * in auto) by smt qed then have "\\<^sub>F s in at 0. \c\C'. \x\ball c (u' c). dist (flow0 x s) (flow0 x 0) < e" by (simp add: eventually_ball_finite_distrib) then show "\\<^sub>F s in at 0. \x\C. dist (flow0 x s) (flow0 x 0) < e" apply eventually_elim apply auto subgoal for s x apply (drule bspec[where x="c' x"]) apply (simp add: c'(2)) apply (drule bspec) prefer 2 apply assumption apply auto using c'(1) by auto done qed qed end subsection \Fixpoints\ context auto_ll_on_open begin lemma fixpoint_sol: assumes "x \ X" "f x = 0" shows "existence_ivl0 x = UNIV" "flow0 x t = x" proof - have sol: "((\t::real. x) solves_ode (\_. f)) UNIV X" apply (rule solves_odeI) by(auto simp add: assms intro!: derivative_intros) from maximal_existence_flow[OF sol] have "UNIV \ existence_ivl0 x" "flow0 x t = x" by auto thus "existence_ivl0 x = UNIV" "flow0 x t = x" by auto qed end end \ No newline at end of file diff --git a/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy b/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy --- a/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy +++ b/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy @@ -1,447 +1,447 @@ (* File: Partial_Zeta_Bounds.thy Author: Manuel Eberl, TU München Asymptotic bounds on sums over k^(-s) for k=1..n, for fixed s and variable n. *) section \Bounds on partial sums of the $\zeta$ function\ theory Partial_Zeta_Bounds imports Euler_MacLaurin.Euler_MacLaurin_Landau Zeta_Function.Zeta_Function Prime_Number_Theorem.Prime_Number_Theorem_Library Prime_Distribution_Elementary_Library begin (* TODO: This does not necessarily require the full complex zeta function. *) text \ We employ Euler--MacLaurin's summation formula to obtain asymptotic estimates for the partial sums of the Riemann $\zeta(s)$ function for fixed real $a$, i.\,e.\ the function \[f(n) = \sum_{k=1}^n k^{-s}\ .\] We distinguish various cases. The case $s = 1$ is simply the Harmonic numbers and is treated apart from the others. \ lemma harm_asymp_equiv: "sum_upto (\n. 1 / n) \[at_top] ln" proof - have "sum_upto (\n. n powr -1) \[at_top] (\x. ln (\x\ + 1))" proof (rule asymp_equiv_sandwich) have "eventually (\x. sum_upto (\n. n powr -1) x \ {ln (\x\ + 1)..ln \x\ + 1}) at_top" using eventually_ge_at_top[of 1] proof eventually_elim case (elim x) have "sum_upto (\n. real n powr -1) x = harm (nat \x\)" unfolding sum_upto_altdef harm_def by (intro sum.cong) (auto simp: field_simps powr_minus) also have "\ \ {ln (\x\ + 1)..ln \x\ + 1}" using elim harm_le[of "nat \x\"] ln_le_harm[of "nat \x\"] by (auto simp: le_nat_iff le_floor_iff) finally show ?case by simp qed thus "eventually (\x. sum_upto (\n. n powr -1) x \ ln (\x\ + 1)) at_top" "eventually (\x. sum_upto (\n. n powr -1) x \ ln \x\ + 1) at_top" by (eventually_elim; simp)+ qed real_asymp+ also have "\ \[at_top] ln" by real_asymp finally show ?thesis by (simp add: powr_minus field_simps) qed lemma fixes s :: real assumes s: "s > 0" "s \ 1" shows zeta_partial_sum_bigo_pos: "(\n. (\k=1..n. real k powr -s) - real n powr (1 - s) / (1 - s) - Re (zeta s)) \ O(\x. real x powr -s)" and zeta_partial_sum_bigo_pos': "(\n. \k=1..n. real k powr -s) =o (\n. real n powr (1 - s) / (1 - s) + Re (zeta s)) +o O(\x. real x powr -s)" proof - define F where "F = (\x. x powr (1 - s) / (1 - s))" define f where "f = (\x. x powr -s)" define f' where "f' = (\x. -s * x powr (-s-1))" define z where "z = Re (zeta s)" interpret euler_maclaurin_nat' F f "(!) [f, f']" 1 0 z "{}" proof have "(\b. (\k=1..b. real k powr -s) - real b powr (1 - s) / (1 - s) - real b powr -s / 2) \ Re (zeta s) - 0" proof (intro tendsto_diff) let ?g = "\b. (\i\<^sub>F b in at_top. Re (?g b) = (\k=1..b. real k powr -s) - real b powr (1 - s) / (1 - s)" using eventually_ge_at_top[of 1] proof eventually_elim case (elim b) have "(\k=1..b. real k powr -s) = (\kn. n - 1"]) auto also have "\ - real b powr (1 - s) / (1 - s) = Re (?g b)" by (auto simp: powr_Reals_eq add_ac) finally show ?case .. qed moreover have "(\b. Re (?g b)) \ Re (zeta s)" using hurwitz_zeta_critical_strip[of "of_real s" 1] s by (intro tendsto_intros) (simp add: zeta_def) ultimately show "(\b. (\k=1..b. real k powr -s) - real b powr (1 - s) / (1 - s)) \ Re (zeta s)" by (blast intro: Lim_transform_eventually) qed (use s in real_asymp) thus "(\b. (\k = 1..b. f (real k)) - F (real b) - (\i<2 * 0 + 1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R ([f, f'] ! i) (real b))) \ z" by (simp add: f_def F_def z_def) qed (use \s \ 1\ in \auto intro!: derivative_eq_intros continuous_intros simp flip: has_real_derivative_iff_has_vector_derivative simp: F_def f_def f'_def nth_Cons split: nat.splits\) { fix n :: nat assume n: "n \ 1" have "(\k=1..n. real k powr -s) = n powr (1 - s) / (1 - s) + z + 1/2 * n powr -s - EM_remainder 1 f' (int n)" using euler_maclaurin_strong_nat'[of n] n by (simp add: F_def f_def) } note * = this have "(\n. (\k=1..n. real k powr -s) - n powr (1 - s) / (1 - s) - z) \ \(\n. 1/2 * n powr -s - EM_remainder 1 f' (int n))" using * by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\n. 1/2 * n powr -s - EM_remainder 1 f' (int n)) \ O(\n. n powr -s)" proof (intro sum_in_bigo) have "(\x. norm (EM_remainder 1 f' (int x))) \ O(\x. real x powr - s)" proof (rule EM_remainder_strong_bigo_nat[where a = 1 and Y = "{}"]) fix x :: real assume "x \ 1" show "norm (f' x) \ s * x powr (-s-1)" using s by (simp add: f'_def) next from s show "((\x. x powr -s) \ 0) at_top" by real_asymp qed (auto simp: f'_def intro!: continuous_intros derivative_eq_intros) thus "(\x. EM_remainder 1 f' (int x)) \ O(\x. real x powr -s)" by simp qed real_asymp+ finally show "(\n. (\k=1..n. real k powr -s) - real n powr (1 - s) / (1 - s) - z) \ O(\x. real x powr -s)" . thus"(\n. \k=1..n. real k powr -s) =o (\n. real n powr (1 - s) / (1 - s) + z) +o O(\x. real x powr -s)" by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def algebra_simps) qed lemma zeta_tail_bigo: fixes s :: real assumes s: "s > 1" shows "(\n. Re (hurwitz_zeta (real n + 1) s)) \ O(\x. real x powr (1 - s))" proof - have [simp]: "complex_of_real (Re (zeta s)) = zeta s" using zeta_real[of s] by (auto elim!: Reals_cases) from s have s': "s > 0" "s \ 1" by auto have "(\n. -Re (hurwitz_zeta (real n + 1) s) - real n powr (1 - s) / (1 - s) + real n powr (1 - s) / (1 - s)) \ O(\x. real x powr (1 - s))" proof (rule sum_in_bigo) have "(\n. -Re (hurwitz_zeta (real n + 1) s) - real n powr (1 - s) / (1 - s)) = (\n. (\k=1..n. real k powr -s) - real n powr (1 - s) / (1 - s) - Re (zeta s))" (is "?lhs = ?rhs") proof fix n :: nat have "hurwitz_zeta (1 + real n) s = zeta s - (\kauto simp: zeta_def powr_Reals_eq\) also have "(\kk=1..n. real k powr -s)" by (rule sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto finally show "?lhs n = ?rhs n" by (simp add: add_ac) qed also have "\ \ O(\x. real x powr (-s))" by (rule zeta_partial_sum_bigo_pos) (use s in auto) also have "(\x. real x powr (-s)) \ O(\x. real x powr (1-s))" by real_asymp finally show "(\n. -Re (hurwitz_zeta (real n + 1) s) - real n powr (1 - s) / (1 - s)) \ \" . qed (use s in real_asymp) thus ?thesis by simp qed lemma zeta_tail_bigo': fixes s :: real assumes s: "s > 1" shows "(\n. Re (hurwitz_zeta (real n) s)) \ O(\x. real x powr (1 - s))" proof - have "(\n. Re (hurwitz_zeta (real n) s)) \ \(\n. Re (hurwitz_zeta (real (n - 1) + 1) s))" by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]]) (auto simp: of_nat_diff) also have "(\n. Re (hurwitz_zeta (real (n - 1) + 1) s)) \ O(\x. real (x - 1) powr (1 - s))" by (rule landau_o.big.compose[OF zeta_tail_bigo[OF assms]]) real_asymp also have "(\x. real (x - 1) powr (1 - s)) \ O(\x. real x powr (1 - s))" by real_asymp finally show ?thesis . qed lemma fixes s :: real assumes s: "s > 0" shows zeta_partial_sum_bigo_neg: "(\n. (\i=1..n. real i powr s) - n powr (1 + s) / (1 + s)) \ O(\n. n powr s)" and zeta_partial_sum_bigo_neg': "(\n. (\i=1..n. real i powr s)) =o (\n. n powr (1 + s) / (1 + s)) +o O(\n. n powr s)" proof - define F where "F = (\x. x powr (1 + s) / (1 + s))" define f where "f = (\x. x powr s)" define f' where "f' = (\x. s * x powr (s - 1))" have "(\i=1..n. f (real i)) - F n = 1 / 2 - F 1 + f n / 2 + EM_remainder' 1 f' 1 (real n)" if n: "n \ 1" for n proof - have "(\i\{1<..n}. f (real i)) - integral {real 1..real n} f = (\k<1. (bernoulli' (Suc k) / fact (Suc k)) *\<^sub>R (([f, f'] ! k) (real n) - ([f, f'] ! k) (real 1))) + EM_remainder' 1 ([f, f'] ! 1) (real 1) (real n)" by (rule euler_maclaurin_strong_raw_nat[where Y = "{}"]) (use \s > 0\ \n \ 1\ in \auto intro!: derivative_eq_intros continuous_intros simp flip: has_real_derivative_iff_has_vector_derivative simp: F_def f_def f'_def nth_Cons split: nat.splits\) also have "(\i\{1<..n}. f (real i)) = (\i\insert 1 {1<..n}. f (real i)) - f 1" using n by (subst sum.insert) auto also from n have "insert 1 {1<..n} = {1..n}" by auto finally have "(\i=1..n. f (real i)) - F n = f 1 + (integral {1..real n} f - F n) + (f (real n) - f 1) / 2 + EM_remainder' 1 f' 1 (real n)" by simp hence "(\i=1..n. f (real i)) - F n = 1 / 2 + (integral {1..real n} f - F n) + f (real n) / 2 + EM_remainder' 1 f' 1 (real n)" using s by (simp add: f_def diff_divide_distrib) also have "(f has_integral (F (real n) - F 1)) {1..real n}" using assms n by (intro fundamental_theorem_of_calculus) (auto simp flip: has_real_derivative_iff_has_vector_derivative simp: F_def f_def intro!: derivative_eq_intros continuous_intros) hence "integral {1..real n} f - F n = - F 1" by (simp add: has_integral_iff) also have "1 / 2 + (-F 1) + f (real n) / 2 = 1 / 2 - F 1 + f n / 2" by simp finally show ?thesis . qed hence "(\n. (\i=1..n. f (real i)) - F n) \ \(\n. 1 / 2 - F 1 + f n / 2 + EM_remainder' 1 f' 1 (real n))" by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]]) also have "(\n. 1 / 2 - F 1 + f n / 2 + EM_remainder' 1 f' 1 (real n)) \ O(\n. real n powr s)" unfolding F_def f_def proof (intro sum_in_bigo) have "(\x. integral {1..real x} (\t. pbernpoly 1 t *\<^sub>R f' t)) \ O(\n. 1 / s * real n powr s)" proof (intro landau_o.big.compose[OF integral_bigo]) have "(\x. pbernpoly 1 x * f' x) \ O(\x. 1 * x powr (s - 1))" by (intro landau_o.big.mult pbernpoly_bigo) (auto simp: f'_def) thus "(\x. pbernpoly 1 x *\<^sub>R f' x) \ O(\x. x powr (s - 1))" by simp from s show "filterlim (\a. 1 / s * a powr s) at_top at_top" by real_asymp next fix a' x :: real assume "a' \ 1" "a' \ x" thus "(\a. pbernpoly 1 a *\<^sub>R f' a) integrable_on {a'..x}" by (intro integrable_EM_remainder') (auto intro!: continuous_intros simp: f'_def) qed (use s in \auto intro!: filterlim_real_sequentially continuous_intros derivative_eq_intros\) thus "(\x. EM_remainder' 1 f' 1 (real x)) \ O(\n. real n powr s)" using \s > 0\ by (simp add: EM_remainder'_def) qed (use \s > 0\ in real_asymp)+ finally show "(\n. (\i=1..n. real i powr s) - n powr (1 + s) / (1 + s)) \ O(\n. n powr s)" by (simp add: f_def F_def) thus "(\n. (\i=1..n. real i powr s)) =o (\n. n powr (1 + s) / (1 + s)) +o O(\n. n powr s)" by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def algebra_simps) qed lemma zeta_partial_sum_le_pos: assumes "s > 0" "s \ 1" defines "z \ Re (zeta (complex_of_real s))" shows "\c>0. \x\1. \sum_upto (\n. n powr -s) x - (x powr (1-s) / (1-s) + z)\ \ c * x powr -s" proof (rule sum_upto_asymptotics_lift_nat_real) show "(\n. (\k = 1..n. real k powr - s) - (real n powr (1 - s) / (1 - s) + z)) \ O(\n. real n powr - s)" using zeta_partial_sum_bigo_pos[OF assms(1,2)] unfolding z_def by (simp add: algebra_simps) from assms have "s < 1 \ s > 1" by linarith thus "(\n. real n powr (1 - s) / (1 - s) + z - (real (Suc n) powr (1 - s) / (1 - s) + z)) \ O(\n. real n powr - s)" by standard (use \s > 0\ in \real_asymp+\) show "(\n. real n powr - s) \ O(\n. real (Suc n) powr - s)" by real_asymp - show "mono_on (\a. a powr - s) {1..} \ mono_on (\x. - (x powr - s)) {1..}" + show "mono_on {1..} (\a. a powr - s) \ mono_on {1..} (\x. - (x powr - s))" using assms by (intro disjI2) (auto intro!: mono_onI powr_mono2') from assms have "s < 1 \ s > 1" by linarith - hence "mono_on (\a. a powr (1 - s) / (1 - s) + z) {1..}" + hence "mono_on {1..} (\a. a powr (1 - s) / (1 - s) + z)" proof assume "s < 1" thus ?thesis using \s > 0\ by (intro mono_onI powr_mono2 divide_right_mono add_right_mono) auto next assume "s > 1" thus ?thesis by (intro mono_onI le_imp_neg_le add_right_mono divide_right_mono_neg powr_mono2') auto qed - thus "mono_on (\a. a powr (1 - s) / (1 - s) + z) {1..} \ - mono_on (\x. - (x powr (1 - s) / (1 - s) + z)) {1..}" by blast + thus "mono_on {1..} (\a. a powr (1 - s) / (1 - s) + z) \ + mono_on {1..} (\x. - (x powr (1 - s) / (1 - s) + z))" by blast qed auto lemma zeta_partial_sum_le_pos': assumes "s > 0" "s \ 1" defines "z \ Re (zeta (complex_of_real s))" shows "\c>0. \x\1. \sum_upto (\n. n powr -s) x - x powr (1-s) / (1-s)\ \ c" proof - have "\c>0. \x\1. \sum_upto (\n. n powr -s) x - x powr (1-s) / (1-s)\ \ c * 1" proof (rule sum_upto_asymptotics_lift_nat_real) have "(\n. (\k = 1..n. real k powr - s) - (real n powr (1 - s) / (1 - s) + z)) \ O(\n. real n powr - s)" using zeta_partial_sum_bigo_pos[OF assms(1,2)] unfolding z_def by (simp add: algebra_simps) also have "(\n. real n powr -s) \ O(\n. 1)" using assms by real_asymp finally have "(\n. (\k = 1..n. real k powr - s) - real n powr (1 - s) / (1 - s) - z) \ O(\n. 1)" by (simp add: algebra_simps) hence "(\n. (\k = 1..n. real k powr - s) - real n powr (1 - s) / (1 - s) - z + z) \ O(\n. 1)" by (rule sum_in_bigo) auto thus "(\n. (\k = 1..n. real k powr - s) - (real n powr (1 - s) / (1 - s))) \ O(\n. 1)" by simp from assms have "s < 1 \ s > 1" by linarith thus "(\n. real n powr (1 - s) / (1 - s) - (real (Suc n) powr (1 - s) / (1 - s))) \ O(\n. 1)" by standard (use \s > 0\ in \real_asymp+\) - show "mono_on (\a. 1) {1..} \ mono_on (\x::real. -1 :: real) {1..}" + show "mono_on {1..} (\a. 1) \ mono_on {1..} (\x::real. -1 :: real)" using assms by (intro disjI2) (auto intro!: mono_onI powr_mono2') from assms have "s < 1 \ s > 1" by linarith - hence "mono_on (\a. a powr (1 - s) / (1 - s)) {1..}" + hence "mono_on {1..} (\a. a powr (1 - s) / (1 - s))" proof assume "s < 1" thus ?thesis using \s > 0\ by (intro mono_onI powr_mono2 divide_right_mono add_right_mono) auto next assume "s > 1" thus ?thesis by (intro mono_onI le_imp_neg_le add_right_mono divide_right_mono_neg powr_mono2') auto qed - thus "mono_on (\a. a powr (1 - s) / (1 - s)) {1..} \ - mono_on (\x. - (x powr (1 - s) / (1 - s))) {1..}" by blast + thus "mono_on {1..} (\a. a powr (1 - s) / (1 - s)) \ + mono_on {1..} (\x. - (x powr (1 - s) / (1 - s)))" by blast qed auto thus ?thesis by simp qed lemma zeta_partial_sum_le_pos'': assumes "s > 0" "s \ 1" shows "\c>0. \x\1. \sum_upto (\n. n powr -s) x\ \ c * x powr max 0 (1 - s)" proof - from zeta_partial_sum_le_pos'[OF assms] obtain c where c: "c > 0" "\x. x \ 1 \ \sum_upto (\x. real x powr - s) x - x powr (1 - s) / (1 - s)\ \ c" by auto { fix x :: real assume x: "x \ 1" have "\sum_upto (\x. real x powr - s) x\ \ \x powr (1 - s) / (1 - s)\ + c" using c(1) c(2)[OF x] x by linarith also have "\x powr (1 - s) / (1 - s)\ = x powr (1 - s) / \1 - s\" using assms by simp also have "\ \ x powr max 0 (1 - s) / \1 - s\" using x by (intro divide_right_mono powr_mono) auto also have "c = c * x powr 0" using x by simp also have "c * x powr 0 \ c * x powr max 0 (1 - s)" using c(1) x by (intro mult_left_mono powr_mono) auto also have "x powr max 0 (1 - s) / \1 - s\ + c * x powr max 0 (1 - s) = (1 / \1 - s\ + c) * x powr max 0 (1 - s)" by (simp add: algebra_simps) finally have "\sum_upto (\x. real x powr - s) x\ \ (1 / \1 - s\ + c) * x powr max 0 (1 - s)" by simp } moreover have "(1 / \1 - s\ + c) > 0" using c assms by (intro add_pos_pos divide_pos_pos) auto ultimately show ?thesis by blast qed lemma zeta_partial_sum_le_pos_bigo: assumes "s > 0" "s \ 1" shows "(\x. sum_upto (\n. n powr -s) x) \ O(\x. x powr max 0 (1 - s))" proof - from zeta_partial_sum_le_pos''[OF assms] obtain c where "\x\1. \sum_upto (\n. n powr -s) x\ \ c * x powr max 0 (1 - s)" by auto thus ?thesis by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto qed lemma zeta_partial_sum_01_asymp_equiv: assumes "s \ {0<..<1}" shows "sum_upto (\n. n powr -s) \[at_top] (\x. x powr (1 - s) / (1 - s))" proof - from zeta_partial_sum_le_pos'[of s] assms obtain c where c: "c > 0" "\x\1. \sum_upto (\x. real x powr -s) x - x powr (1 - s) / (1 - s)\ \ c" by auto hence "(\x. sum_upto (\x. real x powr -s) x - x powr (1 - s) / (1 - s)) \ O(\_. 1)" by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\_. 1) \ o(\x. x powr (1 - s) / (1 - s))" using assms by real_asymp finally show ?thesis by (rule smallo_imp_asymp_equiv) qed lemma zeta_partial_sum_gt_1_asymp_equiv: fixes s :: real assumes "s > 1" defines "\ \ Re (zeta s)" shows "sum_upto (\n. n powr -s) \[at_top] (\x. \)" proof - have [simp]: "\ \ 0" using assms zeta_Re_gt_1_nonzero[of s] zeta_real[of s] by (auto elim!: Reals_cases) from zeta_partial_sum_le_pos[of s] assms obtain c where c: "c > 0" "\x\1. \sum_upto (\x. real x powr -s) x - (x powr (1 - s) / (1 - s) + \)\ \ c * x powr -s" by (auto simp: \_def) hence "(\x. sum_upto (\x. real x powr -s) x - \ - x powr (1 - s) / (1 - s)) \ O(\x. x powr -s)" by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\x. x powr -s) \ o(\_. 1)" using \s > 1\ by real_asymp finally have "(\x. sum_upto (\x. real x powr -s) x - \ - x powr (1 - s) / (1 - s) + x powr (1 - s) / (1 - s)) \ o(\_. 1)" by (rule sum_in_smallo) (use \s > 1\ in real_asymp) thus ?thesis by (simp add: smallo_imp_asymp_equiv) qed lemma zeta_partial_sum_pos_bigtheta: assumes "s > 0" "s \ 1" shows "sum_upto (\n. n powr -s) \ \(\x. x powr max 0 (1 - s))" proof (cases "s > 1") case False thus ?thesis using asymp_equiv_imp_bigtheta[OF zeta_partial_sum_01_asymp_equiv[of s]] assms by (simp add: max_def) next case True have [simp]: "Re (zeta s) \ 0" using True zeta_Re_gt_1_nonzero[of s] zeta_real[of s] by (auto elim!: Reals_cases) show ?thesis using True asymp_equiv_imp_bigtheta[OF zeta_partial_sum_gt_1_asymp_equiv[of s]] by (simp add: max_def) qed lemma zeta_partial_sum_le_neg: assumes "s > 0" shows "\c>0. \x\1. \sum_upto (\n. n powr s) x - x powr (1 + s) / (1 + s)\ \ c * x powr s" proof (rule sum_upto_asymptotics_lift_nat_real) show "(\n. (\k = 1..n. real k powr s) - (real n powr (1 + s) / (1 + s))) \ O(\n. real n powr s)" using zeta_partial_sum_bigo_neg[OF assms(1)] by (simp add: algebra_simps) show "(\n. real n powr (1 + s) / (1 + s) - (real (Suc n) powr (1 + s) / (1 + s))) \ O(\n. real n powr s)" using assms by real_asymp show "(\n. real n powr s) \ O(\n. real (Suc n) powr s)" by real_asymp - show "mono_on (\a. a powr s) {1..} \ mono_on (\x. - (x powr s)) {1..}" + show "mono_on {1..} (\a. a powr s) \ mono_on {1..} (\x. - (x powr s))" using assms by (intro disjI1) (auto intro!: mono_onI powr_mono2) - show "mono_on (\a. a powr (1 + s) / (1 + s)) {1..} \ - mono_on (\x. - (x powr (1 + s) / (1 + s))) {1..}" + show "mono_on {1..} (\a. a powr (1 + s) / (1 + s)) \ + mono_on {1..} (\x. - (x powr (1 + s) / (1 + s)))" using assms by (intro disjI1 divide_right_mono powr_mono2 mono_onI) auto qed auto lemma zeta_partial_sum_neg_asymp_equiv: assumes "s > 0" shows "sum_upto (\n. n powr s) \[at_top] (\x. x powr (1 + s) / (1 + s))" proof - from zeta_partial_sum_le_neg[of s] assms obtain c where c: "c > 0" "\x\1. \sum_upto (\x. real x powr s) x - x powr (1 + s) / (1 + s)\ \ c * x powr s" by auto hence "(\x. sum_upto (\x. real x powr s) x - x powr (1 + s) / (1 + s)) \ O(\x. x powr s)" by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\x. x powr s) \ o(\x. x powr (1 + s) / (1 + s))" using assms by real_asymp finally show ?thesis by (rule smallo_imp_asymp_equiv) qed end \ No newline at end of file diff --git a/thys/Prime_Distribution_Elementary/Prime_Distribution_Elementary_Library.thy b/thys/Prime_Distribution_Elementary/Prime_Distribution_Elementary_Library.thy --- a/thys/Prime_Distribution_Elementary/Prime_Distribution_Elementary_Library.thy +++ b/thys/Prime_Distribution_Elementary/Prime_Distribution_Elementary_Library.thy @@ -1,594 +1,594 @@ (* File: Prime_Distribution_Elementary_Library.thy Author: Manuel Eberl, TU München Various auxiliary material, much of which should probably be moved somewhere else eventually. *) section \Auxiliary material\ theory Prime_Distribution_Elementary_Library imports Zeta_Function.Zeta_Function Prime_Number_Theorem.Prime_Counting_Functions Stirling_Formula.Stirling_Formula begin lemma divisor_count_pos [intro]: "n > 0 \ divisor_count n > 0" by (auto simp: divisor_count_def intro!: Nat.gr0I) lemma divisor_count_eq_0_iff [simp]: "divisor_count n = 0 \ n = 0" by (cases "n = 0") auto lemma divisor_count_pos_iff [simp]: "divisor_count n > 0 \ n > 0" by (cases "n = 0") auto lemma smallest_prime_beyond_eval: "prime n \ smallest_prime_beyond n = n" "\prime n \ smallest_prime_beyond n = smallest_prime_beyond (Suc n)" proof - assume "prime n" thus "smallest_prime_beyond n = n" by (rule smallest_prime_beyond_eq) auto next assume "\prime n" show "smallest_prime_beyond n = smallest_prime_beyond (Suc n)" proof (rule antisym) show "smallest_prime_beyond n \ smallest_prime_beyond (Suc n)" by (rule smallest_prime_beyond_smallest) (auto intro: order.trans[OF _ smallest_prime_beyond_le]) next have "smallest_prime_beyond n \ n" using prime_smallest_prime_beyond[of n] \\prime n\ by metis hence "smallest_prime_beyond n > n" using smallest_prime_beyond_le[of n] by linarith thus "smallest_prime_beyond n \ smallest_prime_beyond (Suc n)" by (intro smallest_prime_beyond_smallest) auto qed qed lemma nth_prime_numeral: "nth_prime (numeral n) = smallest_prime_beyond (Suc (nth_prime (pred_numeral n)))" by (subst nth_prime_Suc[symmetric]) auto lemmas nth_prime_eval = smallest_prime_beyond_eval nth_prime_Suc nth_prime_numeral lemma nth_prime_1 [simp]: "nth_prime (Suc 0) = 3" by (simp add: nth_prime_eval) lemma nth_prime_2 [simp]: "nth_prime 2 = 5" by (simp add: nth_prime_eval) lemma nth_prime_3 [simp]: "nth_prime 3 = 7" by (simp add: nth_prime_eval) (* TODO: copied from afp-devel; delete in AFP 2019 *) lemma strict_mono_sequence_partition: assumes "strict_mono (f :: nat \ 'a :: {linorder, no_top})" assumes "x \ f 0" assumes "filterlim f at_top at_top" shows "\k. x \ {f k.. x)" { obtain n where "x \ f n" using assms by (auto simp: filterlim_at_top eventually_at_top_linorder) also have "f n < f (Suc n)" using assms by (auto simp: strict_mono_Suc_iff) finally have "\n. f (Suc n) > x" by auto } from LeastI_ex[OF this] have "x < f (Suc k)" by (simp add: k_def) moreover have "f k \ x" proof (cases k) case (Suc k') have "k \ k'" if "f (Suc k') > x" using that unfolding k_def by (rule Least_le) with Suc show "f k \ x" by (cases "f k \ x") (auto simp: not_le) qed (use assms in auto) ultimately show ?thesis by auto qed lemma nth_prime_partition: assumes "x \ 2" shows "\k. x \ {nth_prime k.. 2" shows "\k. x \ {real (nth_prime k).. nth_prime k" "n < nth_prime (Suc k)" shows "\prime n" using assms by (metis Suc_leI not_le nth_prime_Suc smallest_prime_beyond_smallest) lemma nth_prime_partition'': includes prime_counting_notation assumes "x \ (2 :: real)" shows "x \ {real (nth_prime (nat \\ x\ - 1))..\ x\))}" proof - obtain n where n: "x \ {nth_prime n.. (nth_prime n) = \ x" unfolding \_def using between_nth_primes_imp_nonprime n by (intro prime_sum_upto_eqI) (auto simp: le_nat_iff le_floor_iff) hence "real n = \ x - 1" by simp hence n_eq: "n = nat \\ x\ - 1" "Suc n = nat \\ x\" by linarith+ with n show ?thesis by simp qed (* END TODO *) (* TODO: Move *) lemma asymp_equivD_strong: assumes "f \[F] g" "eventually (\x. f x \ 0 \ g x \ 0) F" shows "((\x. f x / g x) \ 1) F" proof - from assms(1) have "((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F" by (rule asymp_equivD) also have "?this \ ?thesis" by (intro filterlim_cong eventually_mono[OF assms(2)]) auto finally show ?thesis . qed lemma hurwitz_zeta_shift: fixes s :: complex assumes "a > 0" and "s \ 1" shows "hurwitz_zeta (a + real n) s = hurwitz_zeta a s - (\k {s. Re s > 1}" have "(\k. (a + of_nat (k + n)) powr -s) sums hurwitz_zeta (a + real n) s" using sums_hurwitz_zeta[of "a + real n" s] s assms by (simp add: add_ac) moreover have "(\k. (a + of_nat k) powr -s) sums hurwitz_zeta a s" using sums_hurwitz_zeta[of a s] s assms by (simp add: add_ac) hence "(\k. (a + of_nat (k + n)) powr -s) sums (hurwitz_zeta a s - (\kkauto intro!: holomorphic_intros open_halfspace_Re_gt exI[of _ 2]\) lemma pbernpoly_bigo: "pbernpoly n \ O(\_. 1)" proof - from bounded_pbernpoly[of n] obtain c where "\x. norm (pbernpoly n x) \ c" by auto thus ?thesis by (intro bigoI[of _ c]) auto qed lemma harm_le: "n \ 1 \ harm n \ ln n + 1" using euler_mascheroni_sequence_decreasing[of 1 n] by (simp add: harm_expand) lemma sum_upto_1 [simp]: "sum_upto f 1 = f 1" proof - have "{0<..Suc 0} = {1}" by auto thus ?thesis by (simp add: sum_upto_altdef) qed (* TODO: replace original *) lemma sum_upto_cong' [cong]: "(\n. n > 0 \ real n \ x \ f n = f' n) \ x = x' \ sum_upto f x = sum_upto f' x'" unfolding sum_upto_def by (intro sum.cong) auto lemma finite_primes_le: "finite {p. prime p \ real p \ x}" by (rule finite_subset[of _ "{..nat \x\}"]) (auto simp: le_nat_iff le_floor_iff) lemma frequently_filtermap: "frequently P (filtermap f F) = frequently (\n. P (f n)) F" by (auto simp: frequently_def eventually_filtermap) lemma frequently_mono_filter: "frequently P F \ F \ F' \ frequently P F'" using filter_leD[of F F' "\x. \P x"] by (auto simp: frequently_def) lemma \_at_top: "filterlim primes_pi at_top at_top" unfolding filterlim_at_top proof safe fix C :: real define x0 where "x0 = real (nth_prime (nat \max 0 C\))" show "eventually (\x. primes_pi x \ C) at_top" using eventually_ge_at_top proof eventually_elim fix x assume "x \ x0" have "C \ real (nat \max 0 C\ + 1)" by linarith also have "real (nat \max 0 C\ + 1) = primes_pi x0" unfolding x0_def by simp also have "\ \ primes_pi x" by (rule \_mono) fact finally show "primes_pi x \ C" . qed qed lemma sum_upto_ln_stirling_weak_bigo: "(\x. sum_upto ln x - x * ln x + x) \ O(ln)" proof - let ?f = "\x. x * ln x - x + ln (2 * pi * x) / 2" have "ln (fact n) - (n * ln n - n + ln (2 * pi * n) / 2) \ {0..1/(12*n)}" if "n > 0" for n :: nat using ln_fact_bounds[OF that] by (auto simp: algebra_simps) hence "(\n. ln (fact n) - ?f n) \ O(\n. 1 / real n)" by (intro bigoI[of _ "1/12"] eventually_mono[OF eventually_gt_at_top[of 0]]) auto hence "(\x. ln (fact (nat \x\)) - ?f (nat \x\)) \ O(\x. 1 / real (nat \x\))" by (rule landau_o.big.compose) (intro filterlim_compose[OF filterlim_nat_sequentially] filterlim_floor_sequentially) also have "(\x. 1 / real (nat \x\)) \ O(\x::real. ln x)" by real_asymp finally have "(\x. ln (fact (nat \x\)) - ?f (nat \x\) + (?f (nat \x\) - ?f x)) \ O(\x. ln x)" by (rule sum_in_bigo) real_asymp hence "(\x. ln (fact (nat \x\)) - ?f x) \ O(\x. ln x)" by (simp add: algebra_simps) hence "(\x. ln (fact (nat \x\)) - ?f x + ln (2 * pi * x) / 2) \ O(\x. ln x)" by (rule sum_in_bigo) real_asymp thus ?thesis by (simp add: sum_upto_ln_conv_ln_fact algebra_simps) qed subsection \Various facts about Dirichlet series\ lemma fds_mangoldt': "fds mangoldt = fds_zeta * fds_deriv (fds moebius_mu)" proof - have "fds mangoldt = (fds moebius_mu * fds (\n. of_real (ln (real n)) :: 'a))" (is "_ = ?f") by (subst fds_mangoldt) auto also have "\ = fds_zeta * fds_deriv (fds moebius_mu)" proof (intro fds_eqI) fix n :: nat assume n: "n > 0" have "fds_nth ?f n = (\d | d dvd n. moebius_mu d * of_real (ln (real (n div d))))" by (auto simp: fds_eq_iff fds_nth_mult dirichlet_prod_def) also have "\ = (\d | d dvd n. moebius_mu d * of_real (ln (real n / real d)))" by (intro sum.cong) (auto elim!: dvdE simp: ln_mult split: if_splits) also have "\ = (\d | d dvd n. moebius_mu d * of_real (ln n - ln d))" using n by (intro sum.cong refl) (subst ln_div, auto elim!: dvdE) also have "\ = of_real (ln n) * (\d | d dvd n. moebius_mu d) - (\d | d dvd n. of_real (ln d) * moebius_mu d)" by (simp add: sum_subtractf sum_distrib_left sum_distrib_right algebra_simps) also have "of_real (ln n) * (\d | d dvd n. moebius_mu d) = 0" by (subst sum_moebius_mu_divisors') auto finally show "fds_nth ?f n = fds_nth (fds_zeta * fds_deriv (fds moebius_mu) :: 'a fds) n" by (simp add: fds_nth_mult dirichlet_prod_altdef1 fds_nth_deriv sum_negf scaleR_conv_of_real) qed finally show ?thesis . qed lemma sum_upto_divisor_sum1: "sum_upto (\n. \d | d dvd n. f d :: real) x = sum_upto (\n. f n * floor (x / n)) x" proof - have "sum_upto (\n. \d | d dvd n. f d :: real) x = sum_upto (\n. f n * real (nat (floor (x / n)))) x" using sum_upto_dirichlet_prod[of f "\_. 1" x] by (simp add: dirichlet_prod_def sum_upto_altdef) also have "\ = sum_upto (\n. f n * floor (x / n)) x" unfolding sum_upto_def by (intro sum.cong) auto finally show ?thesis . qed lemma sum_upto_divisor_sum2: "sum_upto (\n. \d | d dvd n. f d :: real) x = sum_upto (\n. sum_upto f (x / n)) x" using sum_upto_dirichlet_prod[of "\_. 1" f x] by (simp add: dirichlet_prod_altdef1) lemma sum_upto_moebius_times_floor_linear: "sum_upto (\n. moebius_mu n * \x / real n\) x = (if x \ 1 then 1 else 0)" proof - have "real_of_int (sum_upto (\n. moebius_mu n * \x / real n\) x) = sum_upto (\n. moebius_mu n * of_int \x / real n\) x" by (simp add: sum_upto_def) also have "\ = sum_upto (\n. \d | d dvd n. moebius_mu d :: real) x" using sum_upto_divisor_sum1[of moebius_mu x] by auto also have "\ = sum_upto (\n. if n = 1 then 1 else 0) x" by (intro sum_upto_cong sum_moebius_mu_divisors' refl) also have "\ = real_of_int (if x \ 1 then 1 else 0)" by (auto simp: sum_upto_def) finally show ?thesis unfolding of_int_eq_iff . qed lemma ln_fact_conv_sum_mangoldt: "sum_upto (\n. mangoldt n * \x / real n\) x = ln (fact (nat \x\))" proof - have "sum_upto (\n. mangoldt n * of_int \x / real n\) x = sum_upto (\n. \d | d dvd n. mangoldt d :: real) x" using sum_upto_divisor_sum1[of mangoldt x] by auto also have "\ = sum_upto (\n. of_real (ln (real n))) x" by (intro sum_upto_cong mangoldt_sum refl) auto also have "\ = (\n\{0<..nat \x\}. ln n)" by (simp add: sum_upto_altdef) also have "\ = ln (\{0<..nat \x\})" unfolding of_nat_prod by (subst ln_prod) auto also have "{0<..nat \x\} = {1..nat \x\}" by auto also have "\\ = fact (nat \x\)" by (simp add: fact_prod) finally show ?thesis by simp qed subsection \Facts about prime-counting functions\ lemma abs_\ [simp]: "\primes_pi x\ = primes_pi x" by (subst abs_of_nonneg) auto lemma \_less_self: includes prime_counting_notation assumes "x > 0" shows "\ x < x" proof - have "\ x \ (\n\{1<..nat \x\}. 1)" unfolding \_def prime_sum_upto_altdef2 by (intro sum_mono2) (auto dest: prime_gt_1_nat) also have "\ = real (nat \x\ - 1)" using assms by simp also have "\ < x" using assms by linarith finally show ?thesis . qed lemma \_le_self': includes prime_counting_notation assumes "x \ 1" shows "\ x \ x - 1" proof - have "\ x \ (\n\{1<..nat \x\}. 1)" unfolding \_def prime_sum_upto_altdef2 by (intro sum_mono2) (auto dest: prime_gt_1_nat) also have "\ = real (nat \x\ - 1)" using assms by simp also have "\ \ x - 1" using assms by linarith finally show ?thesis . qed lemma \_le_self: includes prime_counting_notation assumes "x \ 0" shows "\ x \ x" using \_less_self[of x] assms by (cases "x = 0") auto subsection \Strengthening `Big-O' bounds\ text \ The following two statements are crucial: They allow us to strengthen a `Big-O' statement for $n\to\infty$ or $x\to\infty$ to a bound for \<^emph>\all\ $n\geq n_0$ or all $x\geq x_0$ under some mild conditions. This allows us to use all the machinery of asymptotics in Isabelle and still get a bound that is applicable over the full domain of the function in the end. This is important because Newman often shows that $f(x) \in O(g(x))$ and then writes \[\sum_{n\leq x} f(\frac{x}{n}) = \sum_{n\leq x} O(g(\frac{x}{n}))\] which is not easy to justify otherwise. \ lemma natfun_bigoE: fixes f :: "nat \ _" assumes bigo: "f \ O(g)" and nz: "\n. n \ n0 \ g n \ 0" obtains c where "c > 0" "\n. n \ n0 \ norm (f n) \ c * norm (g n)" proof - from bigo obtain c where c: "c > 0" "eventually (\n. norm (f n) \ c * norm (g n)) at_top" by (auto elim: landau_o.bigE) then obtain n0' where n0': "\n. n \ n0' \ norm (f n) \ c * norm (g n)" by (auto simp: eventually_at_top_linorder) define c' where "c' = Max ((\n. norm (f n) / norm (g n)) ` (insert n0 {n0.. max 1 (max c c') * norm (g n)" if "n \ n0" for n proof (cases "n \ n0'") case False with that have "norm (f n) / norm (g n) \ c'" unfolding c'_def by (intro Max.coboundedI) auto also have "\ \ max 1 (max c c')" by simp finally show ?thesis using nz[of n] that by (simp add: field_simps) next case True hence "norm (f n) \ c * norm (g n)" by (rule n0') also have "\ \ max 1 (max c c') * norm (g n)" by (intro mult_right_mono) auto finally show ?thesis . qed with that[of "max 1 (max c c')"] show ?thesis by auto qed lemma bigoE_bounded_real_fun: fixes f g :: "real \ real" assumes "f \ O(g)" assumes "\x. x \ x0 \ \g x\ \ cg" "cg > 0" assumes "\b. b \ x0 \ bounded (f ` {x0..b})" shows "\c>0. \x\x0. \f x\ \ c * \g x\" proof - from assms(1) obtain c where c: "c > 0" "eventually (\x. \f x\ \ c * \g x\) at_top" by (elim landau_o.bigE) auto then obtain b where b: "\x. x \ b \ \f x\ \ c * \g x\" by (auto simp: eventually_at_top_linorder) have "bounded (f ` {x0..max x0 b})" by (intro assms) auto then obtain C where C: "\x. x \ {x0..max x0 b} \ \f x\ \ C" unfolding bounded_iff by fastforce define c' where "c' = max c (C / cg)" have "\f x\ \ c' * \g x\" if "x \ x0" for x proof (cases "x \ b") case False then have "\f x\ \ C" using C that by auto with False have "\f x\ / \g x\ \ C / cg" by (meson abs_ge_zero assms frac_le landau_omega.R_trans that) also have "\ \ c'" by (simp add: c'_def) finally show "\f x\ \ c' * \g x\" using that False assms(2)[of x] assms(3) by (auto simp add: divide_simps split: if_splits) next case True hence "\f x\ \ c * \g x\" by (intro b) auto also have "\ \ c' * \g x\" by (intro mult_right_mono) (auto simp: c'_def) finally show ?thesis . qed moreover from c(1) have "c' > 0" by (auto simp: c'_def) ultimately show ?thesis by blast qed lemma sum_upto_asymptotics_lift_nat_real_aux: fixes f :: "nat \ real" and g :: "real \ real" assumes bigo: "(\n. (\k=1..n. f k) - g (real n)) \ O(\n. h (real n))" assumes g_bigo_self: "(\n. g (real n) - g (real (Suc n))) \ O(\n. h (real n))" assumes h_bigo_self: "(\n. h (real n)) \ O(\n. h (real (Suc n)))" assumes h_pos: "\x. x \ 1 \ h x > 0" - assumes mono_g: "mono_on g {1..} \ mono_on (\x. - g x) {1..}" - assumes mono_h: "mono_on h {1..} \ mono_on (\x. - h x) {1..}" + assumes mono_g: "mono_on {1..} g \ mono_on {1..} (\x. - g x)" + assumes mono_h: "mono_on {1..} h \ mono_on {1..} (\x. - h x)" shows "\c>0. \x\1. sum_upto f x - g x \ c * h x" proof - have h_nz: "h (real n) \ 0" if "n \ 1" for n using h_pos[of n] that by simp from natfun_bigoE[OF bigo h_nz] obtain c1 where c1: "c1 > 0" "\n. n \ 1 \ norm ((\k=1..n. f k) - g (real n)) \ c1 * norm (h (real n))" by auto from natfun_bigoE[OF g_bigo_self h_nz] obtain c2 where c2: "c2 > 0" "\n. n \ 1 \ norm (g (real n) - g (real (Suc n))) \ c2 * norm (h (real n))" by auto from natfun_bigoE[OF h_bigo_self h_nz] obtain c3 where c3: "c3 > 0" "\n. n \ 1 \ norm (h (real n)) \ c3 * norm (h (real (Suc n)))" by auto { fix x :: real assume x: "x \ 1" define n where "n = nat \x\" from x have n: "n \ 1" unfolding n_def by linarith have "(\k = 1..n. f k) - g x \ (c1 + c2) * h (real n)" using mono_g proof - assume mono: "mono_on (\x. -g x) {1..}" + assume mono: "mono_on {1..} (\x. -g x)" from x have "x \ real (Suc n)" unfolding n_def by linarith hence "(\k=1..n. f k) - g x \ (\k=1..n. f k) - g n + (g n - g (Suc n))" using mono_onD[OF mono, of x "real (Suc n)"] x by auto also have "\ \ norm ((\k=1..n. f k) - g n) + norm (g n - g (Suc n))" by simp also have "\ \ c1 * norm (h n) + c2 * norm (h n)" using n by (intro add_mono c1 c2) auto also have "\ = (c1 + c2) * h n" using h_pos[of "real n"] n by (simp add: algebra_simps) finally show ?thesis . next - assume mono: "mono_on g {1..}" + assume mono: "mono_on {1..} g" have "(\k=1..n. f k) - g x \ (\k=1..n. f k) - g n" using x by (intro diff_mono mono_onD[OF mono]) (auto simp: n_def) also have "\ \ c1 * h (real n)" using c1(2)[of n] n h_pos[of n] by simp also have "\ \ (c1 + c2) * h (real n)" using c2 h_pos[of n] n by (intro mult_right_mono) auto finally show ?thesis . qed also have "(c1 + c2) * h (real n) \ (c1 + c2) * (1 + c3) * h x" using mono_h proof - assume mono: "mono_on (\x. -h x) {1..}" + assume mono: "mono_on {1..} (\x. -h x)" have "(c1 + c2) * h (real n) \ (c1 + c2) * (c3 * h (real (Suc n)))" using c3(2)[of n] n h_pos[of n] h_pos[of "Suc n"] c1(1) c2(1) by (intro mult_left_mono) (auto) also have "\ = (c1 + c2) * c3 * h (real (Suc n))" by (simp add: mult_ac) also have "\ \ (c1 + c2) * (1 + c3) * h (real (Suc n))" using c1(1) c2(1) c3(1) h_pos[of "Suc n"] by (intro mult_left_mono mult_right_mono) auto also from x have "x \ real (Suc n)" unfolding n_def by linarith hence "(c1 + c2) * (1 + c3) * h (real (Suc n)) \ (c1 + c2) * (1 + c3) * h x" using c1(1) c2(1) c3(1) mono_onD[OF mono, of x "real (Suc n)"] x by (intro mult_left_mono) (auto simp: n_def) finally show "(c1 + c2) * h (real n) \ (c1 + c2) * (1 + c3) * h x" . next - assume mono: "mono_on h {1..}" + assume mono: "mono_on {1..} h" have "(c1 + c2) * h (real n) = 1 * ((c1 + c2) * h (real n))" by simp also have "\ \ (1 + c3) * ((c1 + c2) * h (real n))" using c1(1) c2(1) c3(1) h_pos[of n] x n by (intro mult_right_mono) auto also have "\ = (1 + c3) * (c1 + c2) * h (real n)" by (simp add: mult_ac) also have "\ \ (1 + c3) * (c1 + c2) * h x" using x c1(1) c2(1) c3(1) h_pos[of n] n by (intro mult_left_mono mono_onD[OF mono]) (auto simp: n_def) finally show "(c1 + c2) * h (real n) \ (c1 + c2) * (1 + c3) * h x" by (simp add: mult_ac) qed also have "(\k = 1..n. f k) = sum_upto f x" unfolding sum_upto_altdef n_def by (intro sum.cong) auto finally have "sum_upto f x - g x \ (c1 + c2) * (1 + c3) * h x" . } moreover have "(c1 + c2) * (1 + c3) > 0" using c1(1) c2(1) c3(1) by (intro mult_pos_pos add_pos_pos) auto ultimately show ?thesis by blast qed lemma sum_upto_asymptotics_lift_nat_real: fixes f :: "nat \ real" and g :: "real \ real" assumes bigo: "(\n. (\k=1..n. f k) - g (real n)) \ O(\n. h (real n))" assumes g_bigo_self: "(\n. g (real n) - g (real (Suc n))) \ O(\n. h (real n))" assumes h_bigo_self: "(\n. h (real n)) \ O(\n. h (real (Suc n)))" assumes h_pos: "\x. x \ 1 \ h x > 0" - assumes mono_g: "mono_on g {1..} \ mono_on (\x. - g x) {1..}" - assumes mono_h: "mono_on h {1..} \ mono_on (\x. - h x) {1..}" + assumes mono_g: "mono_on {1..} g \ mono_on {1..} (\x. - g x)" + assumes mono_h: "mono_on {1..} h \ mono_on {1..} (\x. - h x)" shows "\c>0. \x\1. \sum_upto f x - g x\ \ c * h x" proof - have "\c>0. \x\1. sum_upto f x - g x \ c * h x" by (intro sum_upto_asymptotics_lift_nat_real_aux assms) then obtain c1 where c1: "c1 > 0" "\x. x \ 1 \ sum_upto f x - g x \ c1 * h x" by auto have "(\n. -(g (real n) - g (real (Suc n)))) \ O(\n. h (real n))" by (subst landau_o.big.uminus_in_iff) fact also have "(\n. -(g (real n) - g (real (Suc n)))) = (\n. g (real (Suc n)) - g (real n))" by simp finally have "(\n. g (real (Suc n)) - g (real n)) \ O(\n. h (real n))" . moreover { have "(\n. -((\k=1..n. f k) - g (real n))) \ O(\n. h (real n))" by (subst landau_o.big.uminus_in_iff) fact also have "(\n. -((\k=1..n. f k) - g (real n))) = (\n. (\k=1..n. -f k) + g (real n))" by (simp add: sum_negf) finally have "(\n. (\k=1..n. - f k) + g (real n)) \ O(\n. h (real n))" . } ultimately have "\c>0. \x\1. sum_upto (\n. -f n) x - (-g x) \ c * h x" using mono_g by (intro sum_upto_asymptotics_lift_nat_real_aux assms) (simp_all add: disj_commute) then obtain c2 where c2: "c2 > 0" "\x. x \ 1 \ sum_upto (\n. - f n) x + g x \ c2 * h x" by auto { fix x :: real assume x: "x \ 1" have "sum_upto f x - g x \ max c1 c2 * h x" using h_pos[of x] x by (intro order.trans[OF c1(2)] mult_right_mono) auto moreover have "sum_upto (\n. -f n) x + g x \ max c1 c2 * h x" using h_pos[of x] x by (intro order.trans[OF c2(2)] mult_right_mono) auto hence "-(sum_upto f x - g x) \ max c1 c2 * h x" by (simp add: sum_upto_def sum_negf) ultimately have "\sum_upto f x - g x\ \ max c1 c2 * h x" by linarith } moreover from c1(1) c2(1) have "max c1 c2 > 0" by simp ultimately show ?thesis by blast qed lemma (in factorial_semiring) primepow_divisors_induct [case_names zero unit factor]: assumes "P 0" "\x. is_unit x \ P x" "\p k x. prime p \ k > 0 \ \p dvd x \ P x \ P (p ^ k * x)" shows "P x" proof - have "finite (prime_factors x)" by simp thus ?thesis proof (induction "prime_factors x" arbitrary: x rule: finite_induct) case empty hence "prime_factors x = {}" by metis hence "prime_factorization x = {#}" by simp thus ?case using assms(1,2) by (auto simp: prime_factorization_empty_iff) next case (insert p A x) define k where "k = multiplicity p x" have "k > 0" using insert.hyps by (auto simp: prime_factors_multiplicity k_def) have p: "p \ prime_factors x" using insert.hyps by auto from p have "x \ 0" "\is_unit p" by (auto simp: in_prime_factors_iff) from multiplicity_decompose'[OF this] obtain y where y: "x = p ^ k * y" "\p dvd y" by (auto simp: k_def) have "prime_factorization x = replicate_mset k p + prime_factorization y" using p \k > 0\ y unfolding y by (subst prime_factorization_mult) (auto simp: prime_factorization_prime_power in_prime_factors_iff) moreover from y p have "p \ prime_factors y" by (auto simp: in_prime_factors_iff) ultimately have "prime_factors y = prime_factors x - {p}" by auto also have "\ = A" using insert.hyps by auto finally have "P y" using insert by auto thus "P x" unfolding y using y \k > 0\ p by (intro assms(3)) (auto simp: in_prime_factors_iff) qed qed end \ No newline at end of file diff --git a/thys/Prime_Distribution_Elementary/Primorial.thy b/thys/Prime_Distribution_Elementary/Primorial.thy --- a/thys/Prime_Distribution_Elementary/Primorial.thy +++ b/thys/Prime_Distribution_Elementary/Primorial.thy @@ -1,353 +1,353 @@ (* File: Primorial.thy Author: Manuel Eberl, TU München The primorial function x#, i.\,e. the product of all primes no greater than x. *) section \The Primorial function\ theory Primorial imports Prime_Distribution_Elementary_Library Primes_Omega begin subsection \Definition and basic properties\ definition primorial :: "real \ nat" where "primorial x = \{p. prime p \ real p \ x}" lemma primorial_mono: "x \ y \ primorial x \ primorial y" unfolding primorial_def by (intro dvd_imp_le prod_dvd_prod_subset) (auto intro!: prod_pos finite_primes_le dest: prime_gt_0_nat) lemma prime_factorization_primorial: "prime_factorization (primorial x) = mset_set {p. prime p \ real p \ x}" proof (intro multiset_eqI) fix p :: nat note fin = finite_primes_le[of x] show "count (prime_factorization (primorial x)) p = count (mset_set {p. prime p \ real p \ x}) p" proof (cases "prime p") case True hence "count (prime_factorization (primorial x)) p = sum (multiplicity p) {p. prime p \ real p \ x}" unfolding primorial_def count_prime_factorization using fin by (subst prime_elem_multiplicity_prod_distrib) auto also from True have "\ = sum (\_. 1) (if p \ x then {p} else {})" using fin by (intro sum.mono_neutral_cong_right) (auto simp: prime_multiplicity_other split: if_splits) also have "\ = count (mset_set {p. prime p \ real p \ x}) p" using True fin by auto finally show ?thesis . qed auto qed lemma prime_factors_primorial [simp]: "prime_factors (primorial x) = {p. prime p \ real p \ x}" unfolding prime_factorization_primorial using finite_primes_le[of x] by simp lemma primorial_pos [simp, intro]: "primorial x > 0" unfolding primorial_def by (intro prod_pos) (auto dest: prime_gt_0_nat) lemma primorial_neq_zero [simp]: "primorial x \ 0" by auto lemma of_nat_primes_omega_primorial [simp]: "real (primes_omega (primorial x)) = primes_pi x" by (simp add: primes_omega_def primes_pi_def prime_sum_upto_def) lemma primes_omega_primorial: "primes_omega (primorial x) = nat \primes_pi x\" by (simp add: primes_omega_def primes_pi_def prime_sum_upto_def) lemma prime_dvd_primorial_iff: "prime p \ p dvd primorial x \ p \ x" using finite_primes_le[of x] by (auto simp: primorial_def prime_dvd_prod_iff dest: primes_dvd_imp_eq) lemma squarefree_primorial [intro]: "squarefree (primorial x)" unfolding primorial_def by (intro squarefree_prod_coprime) (auto simp: squarefree_prime intro: primes_coprime) lemma primorial_ge: "primorial x \ 2 powr primes_pi x" proof - have "2 powr primes_pi x = real (\p | prime p \ real p \ x. 2)" by (simp add: primes_pi_def prime_sum_upto_def powr_realpow) also have "(\p | prime p \ real p \ x. 2) \ (\p | prime p \ real p \ x. p)" by (intro prod_mono) (auto dest: prime_gt_1_nat) also have "\ = primorial x" by (simp add: primorial_def) finally show ?thesis by simp qed lemma primorial_at_top: "filterlim primorial at_top at_top" proof - have "filterlim (\x. real (primorial x)) at_top at_top" proof (rule filterlim_at_top_mono) show "eventually (\x. primorial x \ 2 powr primes_pi x) at_top" by (intro always_eventually primorial_ge allI) have "filterlim (\x. exp (ln 2 * primes_pi x)) at_top at_top" by (intro filterlim_compose[OF exp_at_top] filterlim_tendsto_pos_mult_at_top[OF tendsto_const] \_at_top) auto thus "filterlim (\x. 2 powr primes_pi x) at_top at_top" by (simp add: powr_def mult_ac) qed thus ?thesis unfolding filterlim_sequentially_iff_filterlim_real [symmetric] . qed lemma totient_primorial: "real (totient (primorial x)) = real (primorial x) * (\p | prime p \ real p \ x. 1 - 1 / real p)" for x proof - have "real (totient (primorial x)) = primorial x * (\p\prime_factors (primorial x). 1 - 1 / p)" by (rule totient_formula2) thus ?thesis by simp qed lemma ln_primorial: "ln (primorial x) = primes_theta x" proof - have "finite {p. prime p \ real p \ x}" by (rule finite_subset[of _ "{..nat \x\}"]) (auto simp: le_nat_iff le_floor_iff) thus ?thesis unfolding of_nat_prod primorial_def by (subst ln_prod) (auto dest: prime_gt_0_nat simp: \_def prime_sum_upto_def) qed lemma divisor_count_primorial: "divisor_count (primorial x) = 2 powr primes_pi x" proof - have "divisor_count (primorial x) = (\p | prime p \ real p \ x. divisor_count p)" unfolding primorial_def by (subst divisor_count.prod_coprime) (auto simp: primes_coprime) also have "\ = (\p | prime p \ real p \ x. 2)" by (intro prod.cong divisor_count.prime) auto also have "\ = 2 powr primes_pi x" by (simp add: primes_pi_def prime_sum_upto_def powr_realpow) finally show ?thesis . qed subsection \An alternative view on primorials\ text \ The following function is an alternative representation of primorials; instead of taking the product of all primes up to a given real bound \x\, it takes the product of the first \k\ primes. This is sometimes more convenient. \ definition primorial' :: "nat \ nat" where "primorial' n = (\k 0" unfolding primorial'_def by (auto intro: prime_gt_0_nat) lemma primorial'_neq_0 [simp]: "primorial' n \ 0" by auto lemma strict_mono_primorial': "strict_mono primorial'" unfolding strict_mono_Suc_iff proof fix n :: nat have "primorial' n * 1 < primorial' n * nth_prime n" using prime_gt_1_nat[OF prime_nth_prime[of n]] by (intro mult_strict_left_mono) auto thus "primorial' n < primorial' (Suc n)" by (simp add: primorial'_Suc) qed lemma prime_factorization_primorial': "prime_factorization (primorial' k) = mset_set (nth_prime ` {..i = (\i = (\p\nth_prime ` {.. = mset_set (nth_prime ` {..iprimes_pi x\)" unfolding primorial_def primorial'_def proof (rule prod.reindex_bij_witness[of _ nth_prime "\p. nat \primes_pi (real p)\ - 1"]) fix p assume p: "p \ {p. prime p \ real p \ x}" have [simp]: "primes_pi 2 = 1" by (auto simp: eval_\) have "primes_pi p \ 1" using p \_mono[of 2 "real p"] by (auto dest!: prime_gt_1_nat) with p show "nth_prime (nat \primes_pi p\ - 1) = p" using \_pos[of "real p"] by (intro nth_prime_eqI'') (auto simp: le_nat_iff le_floor_iff primes_pi_def prime_sum_upto_def of_nat_diff) from p have "nat \primes_pi (real p)\ \ nat \primes_pi x\" by (intro nat_mono floor_mono \_mono) auto hence "nat \primes_pi (real p)\ - 1 < nat \primes_pi x\" using \primes_pi p \ 1\ by linarith thus "nat \primes_pi (real p)\ - 1 \ {..primes_pi x\}" by simp show "nth_prime (nat \primes_pi (real p)\ - 1) = p" using p \primes_pi p \ 1\ by (intro nth_prime_eqI'') (auto simp: le_nat_iff primes_pi_def prime_sum_upto_def) next fix k assume k: "k \ {..primes_pi x\}" thus *: "nat \primes_pi (real (nth_prime k))\ - 1 = k" by auto from k have "\(x < 2)" by (intro notI) auto hence "x \ 2" by simp have "real (nth_prime k) \ real (nth_prime (nat \primes_pi x\ - 1))" using k by simp also have "\ \ x" using nth_prime_partition''[of x] \x \ 2\ by auto finally show "nth_prime k \ {p. prime p \ real p \ x}" by auto qed lemma primorial'_conv_primorial: assumes "n > 0" shows "primorial' n = primorial (nth_prime (n - 1))" proof - have "primorial (nth_prime (n - 1)) = (\kMaximal compositeness of primorials\ text \ Primorials are maximally composite, i.\,e.\ any number with \k\ distinct prime factors is as least as big as the primorial with \k\ distinct prime factors, and and number less than a primorial has strictly less prime factors. \ lemma nth_prime_le_prime_sequence: fixes p :: "nat \ nat" - assumes "strict_mono_on p {..k. k < n \ prime (p k)" and "k < n" + assumes "strict_mono_on {..k. k < n \ prime (p k)" and "k < n" shows "nth_prime k \ p k" using assms(3) proof (induction k) case 0 hence "prime (p 0)" by (intro assms) hence "p 0 \ 2" by (auto dest: prime_gt_1_nat) thus ?case by simp next case (Suc k) have IH: "Suc (nth_prime k) \ Suc (p k)" using Suc by simp have "nth_prime (Suc k) = smallest_prime_beyond (Suc (nth_prime k))" by (simp add: nth_prime_Suc) also { have "Suc (nth_prime k) \ Suc (p k)" using Suc by simp also have "\ \ smallest_prime_beyond (Suc (p k))" by (rule smallest_prime_beyond_le) finally have "smallest_prime_beyond (Suc (nth_prime k)) \ smallest_prime_beyond (Suc (p k))" by (rule smallest_prime_beyond_smallest[OF prime_smallest_prime_beyond]) } also have "p k < p (Suc k)" using Suc by (intro strict_mono_onD[OF assms(1)]) auto hence "smallest_prime_beyond (Suc (p k)) \ p (Suc k)" using Suc.prems by (intro smallest_prime_beyond_smallest assms) auto finally show ?case . qed theorem primorial'_primes_omega_le: fixes n :: nat assumes n: "n > 0" shows "primorial' (primes_omega n) \ n" proof (cases "n = 1") case True thus ?thesis by simp next case False with assms have "n > 1" by simp define m where "m = primes_omega n" define P where "P = {p. prime p \ real p \ primes_pi n}" define ps where "ps = sorted_list_of_set (prime_factors n)" have set_ps: "set ps = prime_factors n" by (simp add: ps_def) have [simp]: "length ps = m" by (simp add: ps_def m_def primes_omega_def) have "sorted ps" "distinct ps" by (simp_all add: ps_def) - hence mono: "strict_mono_on (\k. ps ! k) {..k. ps ! k)" by (intro strict_mono_onI le_neq_trans) (auto simp: sorted_nth_mono distinct_conv_nth) from \n > 1\ have "m > 0" by (auto simp: m_def prime_factorization_empty_iff intro!: Nat.gr0I) have "primorial' m = (\km > 0\ by (simp add: of_nat_diff primorial'_def m_def) also have "(\k (\k {.. prime_factors n" using set_ps by (auto simp: set_conv_nth) with i set_ps have "nth_prime i \ ps ! i" by (intro nth_prime_le_prime_sequence[where n = m] mono) (auto simp: set_conv_nth) also have "\ \ ps ! i ^ multiplicity (ps ! i) n" using p by (intro self_le_power) (auto simp: prime_factors_multiplicity dest: prime_gt_1_nat) finally show "nth_prime i \ \" . qed auto also have "\ = (\p\(\i. ps ! i) ` {..distinct ps\ by (subst prod.reindex) (auto intro!: inj_onI simp: distinct_conv_nth) also have "(\i. ps ! i) ` {..p\prime_factors n. p ^ multiplicity p n) = n" using \n > 1\ by (intro prime_factorization_nat [symmetric]) auto finally show "primorial' m \ n" . qed lemma primes_omega_less_primes_omega_primorial: fixes n :: nat assumes n: "n > 0" and "n < primorial x" shows "primes_omega n < primes_omega (primorial x)" proof (cases "n > 1") case False have [simp]: "primes_pi 2 = 1" by (simp add: eval_\) from False assms have [simp]: "n = 1" by auto from assms have "\(x < 2)" by (intro notI) (auto simp: primorial_conv_primorial') thus ?thesis using assms \_mono[of 2 x] by auto next case True define m where "m = primes_omega n" have le: "primorial' m \ n" using primorial'_primes_omega_le[of n] \n > 1\ by (simp add: m_def primes_omega_def) also have "\ < primorial x" by fact also have "\ = primorial' (nat \primes_pi x\)" by (simp add: primorial_conv_primorial') finally have "m < nat \primes_pi x\" using strict_mono_less[OF strict_mono_primorial'] by simp hence "m < primes_pi x" by linarith also have "\ = primes_omega (primorial x)" by simp finally show ?thesis unfolding m_def of_nat_less_iff . qed lemma primes_omega_le_primes_omega_primorial: fixes n :: nat assumes "n \ primorial x" shows "primes_omega n \ primes_omega (primorial x)" proof - consider "n = 0" | "n > 0" "n = primorial x" | "n > 0" "n \ primorial x" by force thus ?thesis by cases (use primes_omega_less_primes_omega_primorial[of n x] assms in auto) qed end \ No newline at end of file diff --git a/thys/Random_BSTs/Random_BSTs.thy b/thys/Random_BSTs/Random_BSTs.thy --- a/thys/Random_BSTs/Random_BSTs.thy +++ b/thys/Random_BSTs/Random_BSTs.thy @@ -1,818 +1,818 @@ (* File: Random_BSTs.thy Author: Manuel Eberl Expected shape of random Binary Search Trees *) section \Expected shape of random Binary Search Trees\ theory Random_BSTs imports Complex_Main "HOL-Probability.Random_Permutations" "HOL-Data_Structures.Tree_Set" Quick_Sort_Cost.Quick_Sort_Average_Case begin (* TODO: Hide this in the proper place *) hide_const (open) Tree_Set.insert subsection \Auxiliary lemmas\ (* TODO: Move? *) lemma linorder_on_linorder_class [intro]: "linorder_on UNIV {(x, y). x \ (y :: 'a :: linorder)}" by (auto simp: linorder_on_def refl_on_def antisym_def trans_def total_on_def) lemma Nil_in_permutations_of_set_iff [simp]: "[] \ permutations_of_set A \ A = {}" by (auto simp: permutations_of_set_def) lemma max_power_distrib_right: fixes a :: "'a :: linordered_semidom" shows "a > 1 \ max (a ^ b) (a ^ c) = a ^ max b c" by (auto simp: max_def) lemma set_tree_empty_iff [simp]: "set_tree t = {} \ t = Leaf" by (cases t) auto lemma card_set_tree_bst: "bst t \ card (set_tree t) = size t" proof (induction t) case (Node l x r) have "set_tree \l, x, r\ = insert x (set_tree l \ set_tree r)" by simp also from Node.prems have "card \ = Suc (card (set_tree l \ set_tree r))" by (intro card_insert_disjoint) auto also from Node have "card (set_tree l \ set_tree r) = size l + size r" by (subst card_Un_disjoint) force+ finally show ?case by simp qed simp_all lemma pair_pmf_cong: "p = p' \ q = q' \ pair_pmf p q = pair_pmf p' q'" by simp lemma expectation_add_pair_pmf: fixes f :: "'a \ 'c::{banach, second_countable_topology}" assumes "finite (set_pmf p)" and "finite (set_pmf q)" shows "measure_pmf.expectation (pair_pmf p q) (\(x,y). f x + g y) = measure_pmf.expectation p f + measure_pmf.expectation q g" proof - have "measure_pmf.expectation (pair_pmf p q) (\(x,y). f x + g y) = measure_pmf.expectation (pair_pmf p q) (\z. f (fst z) + g (snd z))" by (simp add: case_prod_unfold) also have "\ = measure_pmf.expectation (pair_pmf p q) (\z. f (fst z)) + measure_pmf.expectation (pair_pmf p q) (\z. g (snd z))" by (intro Bochner_Integration.integral_add integrable_measure_pmf_finite) (auto intro: assms) also have "measure_pmf.expectation (pair_pmf p q) (\z. f (fst z)) = measure_pmf.expectation (map_pmf fst (pair_pmf p q)) f" by simp also have "map_pmf fst (pair_pmf p q) = p" by (rule map_fst_pair_pmf) also have "measure_pmf.expectation (pair_pmf p q) (\z. g (snd z)) = measure_pmf.expectation (map_pmf snd (pair_pmf p q)) g" by simp also have "map_pmf snd (pair_pmf p q) = q" by (rule map_snd_pair_pmf) finally show ?thesis . qed subsection \Creating a BST from a list\ text \ The following recursive function creates a binary search tree from a given list of elements by inserting them into an initially empty BST from left to right. We will prove that this is the case later, but the recursive definition has the advantage of giving us a useful induction rule, so we chose that definition and prove the alternative definitions later. This recursion, which already almost looks like QuickSort, will be key in analysing the shape distributions of random BSTs. \ fun bst_of_list :: "'a :: linorder list \ 'a tree" where "bst_of_list [] = Leaf" | "bst_of_list (x # xs) = Node (bst_of_list [y \ xs. y < x]) x (bst_of_list [y \ xs. y > x])" lemma bst_of_list_eq_Leaf_iff [simp]: "bst_of_list xs = Leaf \ xs = []" by (induction xs) auto lemma bst_of_list_snoc [simp]: "bst_of_list (xs @ [y]) = Tree_Set.insert y (bst_of_list xs)" by (induction xs rule: bst_of_list.induct) auto lemma bst_of_list_append: "bst_of_list (xs @ ys) = fold Tree_Set.insert ys (bst_of_list xs)" proof (induction ys arbitrary: xs) case (Cons y ys) have "bst_of_list (xs @ (y # ys)) = bst_of_list ((xs @ [y]) @ ys)" by simp also have "\ = fold Tree_Set.insert ys (bst_of_list (xs @ [y]))" by (rule Cons.IH) finally show ?case by simp qed simp_all text \ The following now shows that the recursive function indeed corresponds to the notion of inserting the elements from the list from left to right. \ lemma bst_of_list_altdef: "bst_of_list xs = fold Tree_Set.insert xs Leaf" using bst_of_list_append[of "[]" xs] by simp lemma size_bst_insert: "x \ set_tree t \ size (Tree_Set.insert x t) = Suc (size t)" by (induction t) auto lemma set_bst_insert [simp]: "set_tree (Tree_Set.insert x t) = insert x (set_tree t)" by (induction t) auto lemma set_bst_of_list [simp]: "set_tree (bst_of_list xs) = set xs" by (induction xs rule: rev_induct) simp_all lemma size_bst_of_list_distinct [simp]: assumes "distinct xs" shows "size (bst_of_list xs) = length xs" using assms by (induction xs rule: rev_induct) (auto simp: size_bst_insert) lemma strict_mono_on_imp_less_iff: - assumes "strict_mono_on f A" "x \ A" "y \ A" + assumes "strict_mono_on A f" "x \ A" "y \ A" shows "f x < (f y :: 'b :: linorder) \ x < (y :: 'a :: linorder)" using assms by (cases x y rule: linorder_cases; force simp: strict_mono_on_def)+ lemma bst_of_list_map: fixes f :: "'a :: linorder \ 'b :: linorder" - assumes "strict_mono_on f A" "set xs \ A" + assumes "strict_mono_on A f" "set xs \ A" shows "bst_of_list (map f xs) = map_tree f (bst_of_list xs)" using assms proof (induction xs rule: bst_of_list.induct) case (2 x xs) have "[xa\xs . f xa < f x] = [xa\xs . xa < x]" and "[xa\xs . f xa > f x] = [xa\xs . xa > x]" using "2.prems" by (auto simp: strict_mono_on_imp_less_iff intro!: filter_cong) with 2 show ?case by (auto simp: filter_map o_def) qed auto subsection \Random BSTs\ text \ Analogously to the previous section, we can now view the concept of a random BST (i.\,e.\ a BST obtained by inserting a given set of elements in random order) in two different ways. We again start with the recursive variant: \ function random_bst :: "'a :: linorder set \ 'a tree pmf" where "random_bst A = (if \finite A \ A = {} then return_pmf Leaf else do { x \ pmf_of_set A; l \ random_bst {y \ A. y < x}; r \ random_bst {y \ A. y > x}; return_pmf (Node l x r) })" by auto termination by (relation finite_psubset) auto declare random_bst.simps [simp del] lemma random_bst_empty [simp]: "random_bst {} = return_pmf Leaf" by (simp add: random_bst.simps) lemma set_pmf_random_permutation [simp]: "finite A \ set_pmf (pmf_of_set (permutations_of_set A)) = {xs. distinct xs \ set xs = A}" by (subst set_pmf_of_set) (auto dest: permutations_of_setD) text \ The alternative characterisation is the more intuitive one where we simply pick a random permutation of the set elements uniformly at random and insert them into an empty tree from left to right: \ lemma random_bst_altdef: assumes "finite A" shows "random_bst A = map_pmf bst_of_list (pmf_of_set (permutations_of_set A))" using assms proof (induction A rule: finite_psubset_induct) case (psubset A) define L R where "L = (\x. {y\A. y < x})" and "R = (\x. {y\A. y > x})" { fix x assume x: "x \ A" hence *: "L x \ A" "R x \ A" by (auto simp: L_def R_def) note this [THEN psubset.IH] } note IH = this show ?case proof (cases "A = {}") case False note A = \finite A\ \A \ {}\ have "random_bst A = do { x \ pmf_of_set A; (l, r) \ pair_pmf (random_bst (L x)) (random_bst (R x)); return_pmf (Node l x r) }" using A unfolding pair_pmf_def L_def R_def by (subst random_bst.simps) (simp add: bind_return_pmf bind_assoc_pmf) also have "\ = do { x \ pmf_of_set A; (l, r) \ pair_pmf (map_pmf bst_of_list (pmf_of_set (permutations_of_set (L x)))) (map_pmf bst_of_list (pmf_of_set (permutations_of_set (R x)))); return_pmf (Node l x r) }" using A by (intro bind_pmf_cong refl) (simp_all add: IH) also have "\ = do { x \ pmf_of_set A; (ls, rs) \ pair_pmf (pmf_of_set (permutations_of_set (L x))) (pmf_of_set (permutations_of_set (R x))); return_pmf (Node (bst_of_list ls) x (bst_of_list rs)) }" unfolding map_pair [symmetric] by (simp add: map_pmf_def case_prod_unfold bind_return_pmf bind_assoc_pmf) also have "L = (\x. {y \ A - {x}. y \ x})" by (auto simp: L_def) also have "R = (\x. {y \ A - {x}. \y \ x})" by (auto simp: R_def) also have "do { x \ pmf_of_set A; (ls, rs) \ pair_pmf (pmf_of_set (permutations_of_set {y \ A - {x}. y \ x})) (pmf_of_set (permutations_of_set {y \ A - {x}. \y \ x})); return_pmf (Node (bst_of_list ls) x (bst_of_list rs)) } = do { x \ pmf_of_set A; (ls, rs) \ map_pmf (partition (\y. y \ x)) (pmf_of_set (permutations_of_set (A - {x}))); return_pmf (Node (bst_of_list ls) x (bst_of_list rs)) }" using \finite A\ by (intro bind_pmf_cong refl partition_random_permutations [symmetric]) auto also have "\ = do { x \ pmf_of_set A; (ls, rs) \ map_pmf (\xs. ([y\xs. y < x], [y\xs. y > x])) (pmf_of_set (permutations_of_set (A - {x}))); return_pmf (Node (bst_of_list ls) x (bst_of_list rs)) }" using A by (intro bind_pmf_cong refl map_pmf_cong) (auto intro!: filter_cong dest: permutations_of_setD simp: order.strict_iff_order) also have "\ = map_pmf bst_of_list (pmf_of_set (permutations_of_set A))" using A by (subst random_permutation_of_set[of A]) (auto simp: map_pmf_def bind_return_pmf o_def bind_assoc_pmf not_le) finally show ?thesis . qed (simp_all add: pmf_of_set_singleton) qed lemma finite_set_random_bst [simp, intro]: "finite A \ finite (set_pmf (random_bst A))" by (simp add: random_bst_altdef) lemma random_bst_code [code]: "random_bst (set xs) = map_pmf bst_of_list (pmf_of_set (permutations_of_set (set xs)))" by (rule random_bst_altdef) simp_all lemma random_bst_singleton [simp]: "random_bst {x} = return_pmf (Node Leaf x Leaf)" by (simp add: random_bst_altdef pmf_of_set_singleton) lemma size_random_bst: assumes "t \ set_pmf (random_bst A)" "finite A" shows "size t = card A" proof - from assms obtain xs where "distinct xs" "A = set xs" "t = bst_of_list xs" by (auto simp: random_bst_altdef dest: permutations_of_setD) thus ?thesis using \finite A\ by (simp add: distinct_card) qed lemma random_bst_image: - assumes "finite A" "strict_mono_on f A" + assumes "finite A" "strict_mono_on A f" shows "random_bst (f ` A) = map_pmf (map_tree f) (random_bst A)" proof - from assms(2) have inj: "inj_on f A" by (rule strict_mono_on_imp_inj_on) with assms have "inj_on (map f) (permutations_of_set A)" by (intro inj_on_mapI) auto with assms inj have "random_bst (f ` A) = map_pmf (\x. bst_of_list (map f x)) (pmf_of_set (permutations_of_set A))" by (simp add: random_bst_altdef permutations_of_set_image_inj map_pmf_of_set_inj [symmetric] pmf.map_comp o_def) also have "\ = map_pmf (map_tree f) (random_bst A)" unfolding random_bst_altdef[OF \finite A\] pmf.map_comp o_def using assms - by (intro map_pmf_cong refl bst_of_list_map[of f A]) (auto dest: permutations_of_setD) + by (intro map_pmf_cong refl bst_of_list_map[of A f]) (auto dest: permutations_of_setD) finally show ?thesis . qed text \ We can also re-phrase the non-recursive definition using the @{const fold_random_permutation} combinator from the HOL-Probability library, which folds over a given set in random order. \ lemma random_bst_altdef': assumes "finite A" shows "random_bst A = fold_random_permutation Tree_Set.insert Leaf A" proof - have "random_bst A = map_pmf bst_of_list (pmf_of_set (permutations_of_set A))" using assms by (simp add: random_bst_altdef) also have "\ = map_pmf (\xs. fold Tree_Set.insert xs Leaf) (pmf_of_set (permutations_of_set A))" using assms by (intro map_pmf_cong refl) (auto simp: bst_of_list_altdef) also from assms have "\ = fold_random_permutation Tree_Set.insert Leaf A" by (simp add: fold_random_permutation_fold) finally show ?thesis . qed subsection \Expected height\ text \ For the purposes of the analysis of the expected height, we define the following notion of `expected height', which is essentially two to the power of the height (as defined by Cormen \textit{et al.}) with a special treatment for the empty tree, which has exponential height 0. Note that the height defined by Cormen \textit{et al.}\ differs from the @{const height} function here in Isabelle in that for them, the height of the empty tree is undefined and the height of a singleton tree is 0 etc., whereas in Isabelle, the height of the empty tree is 0 and the height of a singleton tree is 1. \ definition eheight :: "'a tree \ nat" where "eheight t = (if t = Leaf then 0 else 2 ^ (height t - 1))" lemma eheight_Leaf [simp]: "eheight Leaf = 0" by (simp add: eheight_def) lemma eheight_Node_singleton [simp]: "eheight (Node Leaf x Leaf) = 1" by (simp add: eheight_def) lemma eheight_Node: "l \ Leaf \ r \ Leaf \ eheight (Node l x r) = 2 * max (eheight l) (eheight r)" by (cases l; cases r) (simp_all add: eheight_def max_power_distrib_right) fun eheight_rbst :: "nat \ nat pmf" where "eheight_rbst 0 = return_pmf 0" | "eheight_rbst (Suc 0) = return_pmf 1" | "eheight_rbst (Suc n) = do { k \ pmf_of_set {..n}; h1 \ eheight_rbst k; h2 \ eheight_rbst (n - k); return_pmf (2 * max h1 h2)}" definition eheight_exp :: "nat \ real" where "eheight_exp n = measure_pmf.expectation (eheight_rbst n) real" lemma eheight_rbst_reduce: assumes "n > 1" shows "eheight_rbst n = do {k \ pmf_of_set {.. eheight_rbst k; h2 \ eheight_rbst (n - k - 1); return_pmf (2 * max h1 h2)}" using assms by (cases n rule: eheight_rbst.cases) (simp_all add: lessThan_Suc_atMost) lemma Leaf_in_set_random_bst_iff: assumes "finite A" shows "Leaf \ set_pmf (random_bst A) \ A = {}" proof assume "Leaf \ set_pmf (random_bst A)" from size_random_bst[OF this] and assms show "A = {}" by auto qed auto lemma eheight_rbst: assumes "finite A" shows "eheight_rbst (card A) = map_pmf eheight (random_bst A)" using assms proof (induction A rule: finite_psubset_induct) case (psubset A) define rank where "rank = linorder_rank {(x,y). x \ y} A" from \finite A\ have "A = {} \ is_singleton A \ card A > 1" by (auto simp: not_less le_Suc_eq is_singleton_altdef) then consider "A = {}" | "is_singleton A" | "card A > 1" by blast thus ?case proof cases case 3 hence nonempty: "A \ {}" by auto from 3 have "\is_singleton A" by (auto simp: is_singleton_def) hence exists_other: "\y\A. y \ x" for x using \A \ {}\ by (force simp: is_singleton_def) hence "map_pmf eheight (random_bst A) = do { x \ pmf_of_set A; l \ random_bst {y \ A. y < x}; r \ random_bst {y \ A. y > x}; return_pmf (eheight (Node l x r)) }" using \finite A\ by (subst random_bst.simps) (auto simp: map_bind_pmf) also have "\ = do { x \ pmf_of_set A; l \ random_bst {y \ A. y < x}; r \ random_bst {y \ A. y > x}; return_pmf (2 * max (eheight l) (eheight r)) }" using 3 \finite A\ exists_other by (intro bind_pmf_cong refl, subst eheight_Node) (force simp: Leaf_in_set_random_bst_iff not_less nonempty eheight_Node)+ also have "\ = do { x \ pmf_of_set A; h1 \ map_pmf eheight (random_bst {y \ A. y < x}); h2 \ map_pmf eheight (random_bst {y \ A. y > x}); return_pmf (2 * max h1 h2) }" by (simp add: bind_map_pmf) also have "\ = do { x \ pmf_of_set A; h1 \ eheight_rbst (card {y \ A. y < x}); h2 \ eheight_rbst (card {y \ A. y > x}); return_pmf (2 * max h1 h2) }" using \A \ {}\ \finite A\ by (intro bind_pmf_cong psubset.IH [symmetric] refl) auto also have "\ = do { k \ map_pmf rank (pmf_of_set A); h1 \ eheight_rbst k; h2 \ eheight_rbst (card A - k - 1); return_pmf (2 * max h1 h2) }" unfolding bind_map_pmf proof (intro bind_pmf_cong refl, goal_cases) case (1 x) have "rank x = card {y\A-{x}. y \ x}" by (simp add: rank_def linorder_rank_def) also have "{y\A-{x}. y \ x} = {y\A. y < x}" by auto finally show ?case by simp next case (2 x) have "A - {x} = {y\A-{x}. y \ x} \ {y\A. y > x}" by auto also have "card \ = rank x + card {y\A. y > x}" using \finite A\ by (subst card_Un_disjoint) (auto simp: rank_def linorder_rank_def) finally have "card {y\A. y > x} = card A - rank x - 1" using 2 \finite A\ \A \ {}\ by simp thus ?case by simp qed also have "map_pmf rank (pmf_of_set A) = pmf_of_set {..A \ {}\ \finite A\ unfolding rank_def by (intro map_pmf_of_set_bij_betw bij_betw_linorder_rank[of UNIV]) auto also have "do { k \ pmf_of_set {.. eheight_rbst k; h2 \ eheight_rbst (card A - k - 1); return_pmf (2 * max h1 h2) } = eheight_rbst (card A)" by (rule eheight_rbst_reduce [symmetric]) fact+ finally show ?thesis .. qed (auto simp: is_singleton_def) qed lemma finite_pmf_set_eheight_rbst [simp, intro]: "finite (set_pmf (eheight_rbst n))" proof - have "eheight_rbst n = map_pmf eheight (random_bst {..)" by simp finally show ?thesis . qed lemma eheight_exp_0 [simp]: "eheight_exp 0 = 0" by (simp add: eheight_exp_def) lemma eheight_exp_1 [simp]: "eheight_exp (Suc 0) = 1" by (simp add: eheight_exp_def lessThan_Suc) lemma eheight_exp_reduce_bound: assumes "n > 1" shows "eheight_exp n \ 4 / n * (\k = 1 / real n * (\k(h1,h2). 2 * max h1 h2) (?p k)) real)" (is "_ = _ * ?S") unfolding pair_pmf_def map_bind_pmf by (subst eheight_rbst_reduce [OF assms], subst pmf_expectation_bind_pmf_of_set) (insert assms, auto simp: sum_divide_distrib divide_simps) also have "?S = (\kx. 2 * x) (map_pmf ?f (?p k))) real)" by (simp only: pmf.map_comp o_def case_prod_unfold) also have "\ = 2 * (\kk(h1,h2). max (real h1) (real h2)))" by (simp add: case_prod_unfold) also have "\ \ (\k(h1,h2). real h1 + real h2))" unfolding integral_map_pmf case_prod_unfold by (intro sum_mono Bochner_Integration.integral_mono integrable_measure_pmf_finite) auto also have "\ = (\kkkkk. n - Suc k" "\k. n - Suc k"]) auto also have "1 / real n * (2 * (\ + \)) = 4 / real n * \" by simp finally show ?thesis using assms by (simp_all add: mult_left_mono divide_right_mono) qed text \ We now define the following upper bound on the expected exponential height due to Cormen\ \textit{et\ al.}~\cite{cormen}: \ lemma eheight_exp_bound: "eheight_exp n \ real ((n + 3) choose 3) / 4" proof (induction n rule: less_induct) case (less n) consider "n = 0" | "n = 1" | "n > 1" by force thus ?case proof cases case 3 hence "eheight_exp n \ 4 / n * (\kk (\k = real (\kkk\n - 1. ((k + 3) choose 3))" using \n > 1\ by (intro sum.cong) auto also have "\ = ((n + 3) choose 4)" using choose_rising_sum(1)[of 3 "n - 1"] and \n > 1\ by (simp add: add_ac Suc3_eq_add_3) also have "4 / real n * (\ / 4) = real ((n + 3) choose 3) / 4" using \n > 1\ by (cases n) (simp_all add: binomial_fact fact_numeral divide_simps) finally show ?thesis using \n > 1\ by (simp add: mult_left_mono divide_right_mono) qed (auto simp: eval_nat_numeral) qed text \ We then show that this is indeed an upper bound on the expected exponential height by induction over the set of elements. This proof mostly follows that by Cormen\ \textit{et al.}~\cite{cormen}, and partially an answer on the Computer Science Stack Exchange~\cite{sofl}. \ text \ Since the function $\uplambda x.\ 2 ^ x$ is convex, we can then easily derive a bound on the actual height using Jensen's inequality: \ definition height_exp_approx :: "nat \ real" where "height_exp_approx n = log 2 (real ((n + 3) choose 3) / 4) + 1" theorem height_expectation_bound: assumes "finite A" "A \ {}" shows "measure_pmf.expectation (random_bst A) height \ height_exp_approx (card A)" proof - have "convex_on UNIV ((powr) 2)" by (intro convex_on_realI[where f' = "\x. ln 2 * 2 powr x"]) (auto intro!: derivative_eq_intros DERIV_powr simp: powr_def [abs_def]) hence "2 powr measure_pmf.expectation (random_bst A) (\t. real (height t - 1)) \ measure_pmf.expectation (random_bst A) (\t. 2 powr real (height t - 1))" using assms by (intro measure_pmf.jensens_inequality[where I = UNIV]) (auto intro!: integrable_measure_pmf_finite) also have "(\t. 2 powr real (height t - 1)) = (\t. 2 ^ (height t - 1))" by (simp add: powr_realpow) also have "measure_pmf.expectation (random_bst A) (\t. 2 ^ (height t - 1)) = measure_pmf.expectation (random_bst A) (\t. real (eheight t))" using assms by (intro integral_cong_AE) (auto simp: AE_measure_pmf_iff random_bst_altdef eheight_def) also have "\ = measure_pmf.expectation (map_pmf eheight (random_bst A)) real" by simp also have "map_pmf eheight (random_bst A) = eheight_rbst (card A)" by (rule eheight_rbst [symmetric]) fact+ also have "measure_pmf.expectation \ real = eheight_exp (card A)" by (simp add: eheight_exp_def) also have "\ \ real ((card A + 3) choose 3) / 4" by (rule eheight_exp_bound) also have "measure_pmf.expectation (random_bst A) (\t. real (height t - 1)) = measure_pmf.expectation (random_bst A) (\t. real (height t) - 1)" proof (intro integral_cong_AE AE_pmfI, goal_cases) case (3 t) with \A \ {}\ and assms show ?case by (subst of_nat_diff) (auto simp: Suc_le_eq random_bst_altdef) qed auto finally have "2 powr measure_pmf.expectation (random_bst A) (\t. real (height t) - 1) \ real ((card A + 3) choose 3) / 4" . hence "log 2 (2 powr measure_pmf.expectation (random_bst A) (\t. real (height t) - 1)) \ log 2 (real ((card A + 3) choose 3) / 4)" (is "?lhs \ ?rhs") by (subst log_le_cancel_iff) auto also have "?lhs = measure_pmf.expectation (random_bst A) (\t. real (height t) - 1)" by simp also have "\ = measure_pmf.expectation (random_bst A) (\t. real (height t)) - 1" using assms by (subst Bochner_Integration.integral_diff) (auto intro!: integrable_measure_pmf_finite) finally show ?thesis by (simp add: height_exp_approx_def) qed text \ This upper bound is asymptotically equivalent to $c \ln n$ with $c = \frac{3}{\ln 2} \approx 4.328$. This is actually a relatively tight upper bound, since the exact asymptotics of the expected height of a random BST is $c \ln n$ with $c \approx 4.311$.~\cite{reed} However, the proof of these precise asymptotics is very intricate and we will therefore be content with the upper bound. In particular, we can now show that the expected height is $O(\log n)$. \ lemma ln_sum_bigo_ln: "(\x::real. ln (x + c)) \ O(ln)" proof (rule bigoI_tendsto) from eventually_gt_at_top[of "1::real"] show "eventually (\x::real. ln x \ 0) at_top" by eventually_elim simp_all next show "((\x. ln (x + c) / ln x) \ 1) at_top" proof (rule lhospital_at_top_at_top) show "eventually (\x. ((\x. ln (x + c)) has_real_derivative inverse (x + c)) (at x)) at_top" using eventually_gt_at_top[of "-c"] by eventually_elim (auto intro!: derivative_eq_intros simp: field_simps) show "eventually (\x. ((\x. ln x) has_real_derivative inverse x) (at x)) at_top" using eventually_gt_at_top[of 0] by eventually_elim (auto intro!: derivative_eq_intros simp: field_simps) show "((\x. inverse (x + c) / inverse x) \ 1) at_top" proof (rule Lim_transform_eventually) show "eventually (\x. inverse (1 + c / x) = inverse (x + c) / inverse x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "-c"] by eventually_elim (simp add: field_simps) have "((\x. inverse (1 + c / x)) \ inverse (1 + 0)) at_top" by (intro tendsto_inverse tendsto_add tendsto_const real_tendsto_divide_at_top[OF tendsto_const] filterlim_ident) simp_all thus "((\x. inverse (1 + c / x)) \ 1) at_top" by simp qed qed (auto simp: ln_at_top eventually_at_top_not_equal) qed corollary height_expectation_bigo: "height_exp_approx \ O(ln)" proof - let ?T = "\x::real. log 2 (x + 1) + log 2 (x + 2) + log 2 (x + 3) + (1 - log 2 24)" have "eventually (\n. height_exp_approx n = log 2 (real n + 1) + log 2 (real n + 2) + log 2 (real n + 3) + (1 - log 2 24)) at_top" (is "eventually (\n. _ = ?T n) at_top") using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim n) have "height_exp_approx n = log 2 (real (n + 3 choose 3) / 4) + 1" by (simp add: height_exp_approx_def log_divide) also have "real ((n + 3) choose 3) = real (n + 3) gchoose 3" by (simp add: binomial_gbinomial) also have "\ / 4 = (real n + 1) * (real n + 2) * (real n + 3) / 24" by (simp add: gbinomial_pochhammer' numeral_3_eq_3 pochhammer_Suc add_ac) also have "log 2 \ = log 2 (real n + 1) + log 2 (real n + 2) + log 2 (real n + 3) - log 2 24" by (simp add: log_divide log_mult) finally show ?case by simp qed hence "height_exp_approx \ \(?T)" by (rule bigthetaI_cong) also have *: "(\x. ln (x + c) / ln 2) \ O(ln)" for c :: real by (subst landau_o.big.cdiv_in_iff') (auto intro!: ln_sum_bigo_ln) have "?T \ O(\n. ln (real n))" unfolding log_def by (intro bigo_real_nat_transfer sum_in_bigo ln_sum_bigo_ln *) simp_all finally show ?thesis . qed subsection \Lookup costs\ text \ The following function describes the cost incurred when looking up a specific element in a specific BST. The cost corresponds to the number of edges traversed in the lookup. \ primrec lookup_cost :: "'a :: linorder \ 'a tree \ nat" where "lookup_cost x Leaf = 0" | "lookup_cost x (Node l y r) = (if x = y then 0 else if x < y then Suc (lookup_cost x l) else Suc (lookup_cost x r))" text \ Some of the literature defines these costs as 1 in the case that the current node is the correct one, i.\,e.\ their costs are our costs plus 1. These alternative costs are exactly the number of comparisons performed in the lookup. Our cost function has the advantage of precisely summing up to the internal path length and therefore gives us slightly nicer results, and since the difference is only a ${}+1$ in the end, this variant seemed more reasonable. \ text \ It can be shown with a simple induction that The sum of all lookup costs in a tree is the internal path length of the tree. \ theorem sum_lookup_costs: fixes t :: "'a :: linorder tree" assumes "bst t" shows "(\x\set_tree t. lookup_cost x t) = ipl t" using assms proof (induction t) case (Node l x r) from Node.prems have disj: "x \ set_tree l" "x \ set_tree r" "set_tree l \ set_tree r = {}" by force+ have "set_tree (Node l x r) = insert x (set_tree l \ set_tree r)" by simp also have "(\y\\. lookup_cost y (Node l x r)) = lookup_cost x \l, x, r\ + (\y\set_tree l. lookup_cost y \l, x, r\) + (\y\set_tree r. lookup_cost y \l, x, r\)" using disj by (simp add: sum.union_disjoint) also have "(\y\set_tree l. lookup_cost y \l, x, r\) = (\y\set_tree l. 1 + lookup_cost y l)" using disj and Node by (intro sum.cong refl) auto also have "\ = size l + ipl l" using Node by (subst sum.distrib) (simp_all add: card_set_tree_bst) also have "(\y\set_tree r. lookup_cost y \l, x, r\) = (\y\set_tree r. 1 + lookup_cost y r)" using disj and Node by (intro sum.cong refl) auto also have "\ = size r + ipl r" using Node by (subst sum.distrib) (simp_all add: card_set_tree_bst) finally show ?case by simp qed simp_all text \ This allows us to easily show that the expected cost of looking up a random element in a fixed tree is the internal path length divided by the number of elements. \ theorem expected_lookup_cost: assumes "bst t" "t \ Leaf" shows "measure_pmf.expectation (pmf_of_set (set_tree t)) (\x. lookup_cost x t) = ipl t / size t" using assms by (subst integral_pmf_of_set) (simp_all add: sum_lookup_costs of_nat_sum [symmetric] card_set_tree_bst) text \ Therefore, we will now turn to analysing the internal path length of a random BST. This then clearly related to the expected lookup costs of a random element in a random BST by the above result. \ subsection \Average Path Length\ text \ The internal path length satisfies the recursive equation @{thm ipl.simps(2)[of l x r]}. This is quite similar to the number of comparisons performed by QuickSort, and indeed, we can reduce the internal path length of a random BST to the number of comparisons performed by QuickSort on a randomly-ordered list relatively easily: \ theorem map_pmf_random_bst_eq_rqs_cost: assumes "finite A" shows "map_pmf ipl (random_bst A) = rqs_cost (card A)" using assms proof (induction A rule: finite_psubset_induct) case (psubset A) show ?case proof (cases "A = {}") case False note A = \finite A\ \A \ {}\ define n where "n = card A - 1" define rank :: "'a \ nat" where "rank = linorder_rank {(x,y). x \ y} A" from A have card: "card A = Suc n" by (cases "card A") (auto simp: n_def) from A have "map_pmf ipl (random_bst A) = do { x \ pmf_of_set A; (l,r) \ pair_pmf (random_bst {y \ A. y < x}) (random_bst {y \ A. y > x}); return_pmf (ipl (Node l x r)) }" by (subst random_bst.simps) (simp_all add: pair_pmf_def card map_pmf_def bind_assoc_pmf bind_return_pmf) also have "\ = do { x \ pmf_of_set A; (l,r) \ pair_pmf (random_bst {y \ A. y < x}) (random_bst {y \ A. y > x}); return_pmf (n + ipl l + ipl r) }" proof (intro bind_pmf_cong refl, clarify, goal_cases) case (1 x l r) from 1 and A have "n = card (A - {x})" by (simp add: n_def) also have "A - {x} = {y\A. y < x} \ {y\A. y > x}" by auto also have "card \ = card {y\A. y < x} + card {y\A. y > x}" using \finite A\ by (intro card_Un_disjoint) auto also from 1 and A have "card {y\A. y < x} = size l" by (auto dest: size_random_bst) also from 1 and A have "card {y\A. y > x} = size r" by (auto dest: size_random_bst) finally show ?case by simp qed also have "\ = do { x \ pmf_of_set A; (l,r) \ pair_pmf (map_pmf ipl (random_bst {y \ A. y < x})) (map_pmf ipl (random_bst {y \ A. y > x})); return_pmf (n + l + r) }" by (simp add: map_pair [symmetric] case_prod_unfold bind_map_pmf) also have "\ = do { i \ map_pmf rank (pmf_of_set A); (l,r) \ pair_pmf (rqs_cost i) (rqs_cost (n - i)); return_pmf (n + l + r) }" (is "_ = bind_pmf _ ?f") unfolding bind_map_pmf proof (intro bind_pmf_cong refl pair_pmf_cong, goal_cases) case (1 x) have "map_pmf ipl (random_bst {y \ A. y < x}) = rqs_cost (card {y \ A. y < x})" using 1 and A by (intro psubset.IH) auto also have "{y \ A. y < x} = {y \ A - {x}. y \ x}" by auto hence "card {y \ A. y < x} = rank x" by (simp add: rank_def linorder_rank_def) finally show ?case . next case (2 x) have "map_pmf ipl (random_bst {y \ A. y > x}) = rqs_cost (card {y \ A. y > x})" using 2 and A by (intro psubset.IH) auto also have "{y \ A. y > x} = A - {x} - {y \ A - {x}. y \ x}" by auto hence "card {y \ A. y > x} = card \" by (simp only:) also from 2 and A have "\ = n - rank x" by (subst card_Diff_subset) (auto simp: rank_def linorder_rank_def n_def) finally show ?case . qed also from A have "map_pmf rank (pmf_of_set A) = pmf_of_set {.. \ ?f = rqs_cost (card A)" by (simp add: pair_pmf_def bind_assoc_pmf bind_return_pmf card) finally show ?thesis . qed simp_all qed text \ In particular, this means that the expected values are the same: \ corollary expected_ipl_random_bst_eq: assumes "finite A" shows "measure_pmf.expectation (random_bst A) ipl = rqs_cost_exp (card A)" proof - have "measure_pmf.expectation (random_bst A) ipl = measure_pmf.expectation (map_pmf ipl (random_bst A)) real" by simp also from assms have "map_pmf ipl (random_bst A) = rqs_cost (card A)" by (rule map_pmf_random_bst_eq_rqs_cost) also have "measure_pmf.expectation \ real = rqs_cost_exp (card A)" by (rule expectation_rqs_cost) finally show ?thesis . qed text \ Therefore, the results about the expected number of comparisons of QuickSort carry over to the expected internal path length: \ corollary expected_ipl_random_bst_eq': assumes "finite A" shows "measure_pmf.expectation (random_bst A) ipl = 2 * real (card A + 1) * harm (card A) - 4 * real (card A)" by (simp add: expected_ipl_random_bst_eq rqs_cost_exp_eq assms) end diff --git a/thys/Simplicial_complexes_and_boolean_functions/Bij_betw_simplicial_complex_bool_func.thy b/thys/Simplicial_complexes_and_boolean_functions/Bij_betw_simplicial_complex_bool_func.thy --- a/thys/Simplicial_complexes_and_boolean_functions/Bij_betw_simplicial_complex_bool_func.thy +++ b/thys/Simplicial_complexes_and_boolean_functions/Bij_betw_simplicial_complex_bool_func.thy @@ -1,247 +1,247 @@ theory Bij_betw_simplicial_complex_bool_func imports Simplicial_complex begin section\Bijection between simplicial complexes and monotone Boolean functions\ context simplicial_complex begin lemma ceros_of_boolean_input_in_set: assumes s: "\ \ simplices" shows "ceros_of_boolean_input (vec n (\i. i \ \)) = \" unfolding ceros_of_boolean_input_def using s unfolding simplices_def by auto lemma assumes sigma: "\ \ simplices" and nempty: "\ \ {}" shows "Max \ < n" proof - have "Max \ \ \" using linorder_class.Max_in [OF finite_simplex [OF sigma] nempty] . thus ?thesis using sigma unfolding simplices_def by auto qed definition bool_vec_from_simplice :: "nat set => (bool vec)" where "bool_vec_from_simplice \ = vec n (\i. i \ \)" lemma [simp]: assumes "\ \ simplices" shows "ceros_of_boolean_input (bool_vec_from_simplice \) = \" unfolding bool_vec_from_simplice_def unfolding ceros_of_boolean_input_def unfolding dim_vec using assms unfolding simplices_def by auto lemma [simp]: assumes n: "dim_vec f = n" shows "bool_vec_from_simplice (ceros_of_boolean_input f) = f" unfolding bool_vec_from_simplice_def unfolding ceros_of_boolean_input_def using n by auto lemma "bool_vec_from_simplice {0} = vec n (\i. i \ {0})" unfolding bool_vec_from_simplice_def by auto lemma "bool_vec_from_simplice {1,2} = vec n (\i. i \ {1,2})" unfolding bool_vec_from_simplice_def by auto lemma simplicial_complex_implies_true: assumes "\ \ simplicial_complex_induced_by_monotone_boolean_function n f" shows "f (bool_vec_from_simplice \)" unfolding bool_vec_from_simplice_def using assms unfolding simplicial_complex_induced_by_monotone_boolean_function_def unfolding ceros_of_boolean_input_def apply auto by (smt (verit, best) dim_vec eq_vecI index_vec) definition bool_vec_set_from_simplice_set :: "nat set set => (bool vec) set" where "bool_vec_set_from_simplice_set K = {\. (dim_vec \ = n) \ (\k\K. \ = bool_vec_from_simplice k)}" definition boolean_function_from_simplicial_complex :: "nat set set => (bool vec => bool)" where "boolean_function_from_simplicial_complex K = (\x. x \ (bool_vec_set_from_simplice_set K))" lemma "Collect (boolean_function_from_simplicial_complex K) = (bool_vec_set_from_simplice_set K)" unfolding boolean_function_from_simplicial_complex_def by simp text\The Boolean function induced by a simplicial complex is monotone. This result is proven in Scoville as part of the proof of Proposition 6.16. The opposite direction has been proven as @{thm monotone_bool_fun_induces_simplicial_complex}.\ lemma simplicial_complex_induces_monotone_bool_fun: assumes mon: "simplicial_complex (K::nat set set)" shows "boolean_functions.monotone_bool_fun n (boolean_function_from_simplicial_complex K)" proof (unfold boolean_functions.monotone_bool_fun_def) - show "mono_on (boolean_function_from_simplicial_complex K) (carrier_vec n)" + show "mono_on (carrier_vec n) (boolean_function_from_simplicial_complex K)" proof (intro mono_onI) fix r and s::"bool vec" assume r_le_s: "r \ s" show "boolean_function_from_simplicial_complex K r \ boolean_function_from_simplicial_complex K s" proof (cases "boolean_function_from_simplicial_complex K r") case False then show ?thesis by simp next case True have ce: "ceros_of_boolean_input s \ ceros_of_boolean_input r" using monotone_ceros_of_boolean_input [OF r_le_s] . from True obtain k where r_def: "r = vec n (\i. i \ k)" and k: "k \ K" unfolding boolean_function_from_simplicial_complex_def unfolding bool_vec_set_from_simplice_set_def unfolding bool_vec_from_simplice_def by auto have r_in_K: "ceros_of_boolean_input r \ K" using k mon unfolding r_def unfolding ceros_of_boolean_input_def unfolding dim_vec using simplicial_complex_def [of K] by fastforce have "boolean_function_from_simplicial_complex K s" proof (unfold boolean_function_from_simplicial_complex_def bool_vec_set_from_simplice_set_def, rule, rule conjI) show "dim_vec s = n" by (metis less_eq_vec_def dim_vec r_def r_le_s) show "\k\K. s = bool_vec_from_simplice k" proof (rule bexI [of _ "ceros_of_boolean_input s"], unfold bool_vec_from_simplice_def) show "ceros_of_boolean_input s \ K" using simplicial_complex_monotone [OF mon r_in_K ce] . show "s = vec n (\i. i \ ceros_of_boolean_input s)" using ce unfolding ceros_of_boolean_input_def using r_le_s unfolding less_eq_vec_def unfolding r_def unfolding dim_vec by auto qed qed thus ?thesis by simp qed qed qed lemma shows "(simplicial_complex_induced_by_monotone_boolean_function n) \ boolean_functions.monotone_bool_fun_set n \ (simplicial_complex_set::nat set set set)" proof fix x::"bool vec \ bool" assume x: "x \ boolean_functions.monotone_bool_fun_set n" show "simplicial_complex_induced_by_monotone_boolean_function n x \ simplicial_complex_set" using monotone_bool_fun_induces_simplicial_complex [of x] x unfolding boolean_functions.monotone_bool_fun_set_def unfolding boolean_functions.monotone_bool_fun_def simplicial_complex_set_def by auto qed lemma shows "boolean_function_from_simplicial_complex \ (simplicial_complex_set::nat set set set) \ boolean_functions.monotone_bool_fun_set n" proof fix x::"nat set set" assume x: "x \ simplicial_complex_set" show "boolean_function_from_simplicial_complex x \ boolean_functions.monotone_bool_fun_set n" using simplicial_complex_induces_monotone_bool_fun [of x] unfolding boolean_functions.monotone_bool_fun_set_def unfolding boolean_functions.monotone_bool_fun_def using x unfolding simplicial_complex_set_def unfolding mem_Collect_eq by fast qed text\Given a Boolean function @{term f}, if we build its associated simplicial complex and then the associated Boolean function, we obtain @{term f}. The result holds for every Boolean function @{term f} (the premise on @{term f} being monotone can be omitted).\ lemma boolean_function_from_simplicial_complex_simplicial_complex_induced_by_monotone_boolean_function: fixes f :: "bool vec \ bool" assumes dim: "v \ carrier_vec n" shows "boolean_function_from_simplicial_complex (simplicial_complex_induced_by_monotone_boolean_function n f) v = f v" proof (intro iffI) assume xb: "f v" show bf: "boolean_function_from_simplicial_complex (simplicial_complex_induced_by_monotone_boolean_function n f) v" proof - have "f v \ v = bool_vec_from_simplice (ceros_of_boolean_input v)" unfolding ceros_of_boolean_input_def unfolding bool_vec_from_simplice_def using xb dim unfolding carrier_vec_def by auto then show ?thesis unfolding simplicial_complex_induced_by_monotone_boolean_function_def unfolding boolean_function_from_simplicial_complex_def unfolding bool_vec_set_from_simplice_set_def unfolding mem_Collect_eq using dim unfolding carrier_vec_def by blast qed next assume "boolean_function_from_simplicial_complex (simplicial_complex_induced_by_monotone_boolean_function n f) v" then show "f v" unfolding simplicial_complex_induced_by_monotone_boolean_function_def unfolding boolean_function_from_simplicial_complex_def unfolding bool_vec_set_from_simplice_set_def unfolding mem_Collect_eq using \boolean_function_from_simplicial_complex (simplicial_complex_induced_by_monotone_boolean_function n f) v\ boolean_function_from_simplicial_complex_def simplicial_complex.bool_vec_set_from_simplice_set_def simplicial_complex_implies_true by fastforce qed text\Given a simplicial complex @{term K}, if we build its associated Boolean function, and then the associated simplicial complex, we obtain @{term K}.\ lemma simplicial_complex_induced_by_monotone_boolean_function_boolean_function_from_simplicial_complex: fixes K :: "nat set set" assumes K: "simplicial_complex K" shows "simplicial_complex_induced_by_monotone_boolean_function n (boolean_function_from_simplicial_complex K) = K" proof (intro equalityI) show "simplicial_complex_induced_by_monotone_boolean_function n (boolean_function_from_simplicial_complex K) \ K" proof fix x :: "nat set" assume x: "x \ simplicial_complex_induced_by_monotone_boolean_function n (boolean_function_from_simplicial_complex K)" show "x \ K" using x unfolding boolean_function_from_simplicial_complex_def unfolding simplicial_complex_induced_by_monotone_boolean_function_def unfolding bool_vec_from_simplice_def bool_vec_set_from_simplice_set_def using K unfolding simplicial_complex_def simplices_def by auto (metis assms bool_vec_from_simplice_def ceros_of_boolean_input_in_set simplicial_complex_def) qed next show "K \ simplicial_complex_induced_by_monotone_boolean_function n (boolean_function_from_simplicial_complex K)" proof fix x :: "nat set" assume "x \ K" hence x: "x \ simplices" using K unfolding simplicial_complex_def by simp have bvs: "ceros_of_boolean_input (bool_vec_from_simplice x) = x" unfolding one_bool_def unfolding bool_vec_from_simplice_def using ceros_of_boolean_input_in_set [OF x] . show "x \ simplicial_complex_induced_by_monotone_boolean_function n (boolean_function_from_simplicial_complex K)" unfolding boolean_function_from_simplicial_complex_def unfolding simplicial_complex_induced_by_monotone_boolean_function_def unfolding bool_vec_from_simplice_def bool_vec_set_from_simplice_set_def using x bool_vec_from_simplice_def bvs by (metis (mono_tags, lifting) \x \ K\ dim_vec mem_Collect_eq) qed qed end end \ No newline at end of file diff --git a/thys/Simplicial_complexes_and_boolean_functions/Boolean_functions.thy b/thys/Simplicial_complexes_and_boolean_functions/Boolean_functions.thy --- a/thys/Simplicial_complexes_and_boolean_functions/Boolean_functions.thy +++ b/thys/Simplicial_complexes_and_boolean_functions/Boolean_functions.thy @@ -1,119 +1,119 @@ theory Boolean_functions imports Main "Jordan_Normal_Form.Matrix" begin section\Boolean functions\ text\Definition of monotonicity\ text\We consider (monotone) Boolean functions over vectors of length $n$, so that we can later prove that those are isomorphic to simplicial complexes of dimension $n$ (in $n$ vertexes).\ locale boolean_functions = fixes n::"nat" begin definition bool_fun_dim_n :: "(bool vec => bool) set" where "bool_fun_dim_n = {f. f \ carrier_vec n \ (UNIV::bool set)}" definition monotone_bool_fun :: "(bool vec => bool) => bool" - where "monotone_bool_fun f \ (mono_on f (carrier_vec n))" + where "monotone_bool_fun \ (mono_on (carrier_vec n))" definition monotone_bool_fun_set :: "(bool vec => bool) set" where "monotone_bool_fun_set = (Collect monotone_bool_fun)" text\Some examples of Boolean functions\ definition bool_fun_top :: "bool vec => bool" where "bool_fun_top f = True" definition bool_fun_bot :: "bool vec => bool" where "bool_fun_bot f = False" end section\Threshold function\ definition count_true :: "bool vec => nat" where "count_true v = sum (\i. if vec_index v i then 1 else 0::nat) {0..i. False)) 2 = False" by simp lemma "vec_index (vec (5::nat) (\i. True)) 3 = True" by simp lemma "count_true (vec (1::nat) (\i. True)) = 1" unfolding count_true_def by simp lemma "count_true (vec (2::nat) (\i. True)) = 2" unfolding count_true_def by simp lemma "count_true (vec (5::nat) (\i. True)) = 5" unfolding count_true_def by simp text\The threshold function is a Boolean function which also satisfies the condition of being \emph{evasive}. We follow the definition by Scoville~\cite[Problem 6.5]{SC19}.\ definition bool_fun_threshold :: "nat => (bool vec => bool)" where "bool_fun_threshold i = (\v. if i \ count_true v then True else False)" context boolean_functions begin -lemma "mono_on bool_fun_top UNIV" +lemma "mono_on UNIV bool_fun_top" by (simp add: bool_fun_top_def mono_onI monotone_bool_fun_def) lemma "monotone_bool_fun bool_fun_top" by (simp add: bool_fun_top_def mono_onI monotone_bool_fun_def) -lemma "mono_on bool_fun_bot UNIV" +lemma "mono_on UNIV bool_fun_bot" by (simp add: bool_fun_bot_def mono_onI monotone_bool_fun_def) lemma "monotone_bool_fun bool_fun_bot" by (simp add: bool_fun_bot_def mono_onI monotone_bool_fun_def) lemma monotone_count_true: assumes ulev: "(u::bool vec) \ v" shows "count_true u \ count_true v" unfolding count_true_def using Groups_Big.ordered_comm_monoid_add_class.sum_mono [of "{0..i. if vec_index u i then 1 else 0)" "(\i. if vec_index v i then 1 else 0)"] using ulev unfolding Matrix.less_eq_vec_def by fastforce text\The threshold function is monotone.\ lemma monotone_threshold: assumes ulev: "(u::bool vec) \ v" shows "bool_fun_threshold n u \ bool_fun_threshold n v" unfolding bool_fun_threshold_def using monotone_count_true [OF ulev] by simp lemma assumes "(u::bool vec) \ v" and "n < dim_vec u" shows "bool_fun_threshold n u \ bool_fun_threshold n v" using monotone_threshold [OF assms(1)] . -lemma "mono_on (bool_fun_threshold n) UNIV" +lemma "mono_on UNIV (bool_fun_threshold n)" by (meson mono_onI monotone_bool_fun_def monotone_threshold) lemma "monotone_bool_fun (bool_fun_threshold n)" unfolding monotone_bool_fun_def by (meson boolean_functions.monotone_threshold mono_onI) end end diff --git a/thys/Simplicial_complexes_and_boolean_functions/Simplicial_complex.thy b/thys/Simplicial_complexes_and_boolean_functions/Simplicial_complex.thy --- a/thys/Simplicial_complexes_and_boolean_functions/Simplicial_complex.thy +++ b/thys/Simplicial_complexes_and_boolean_functions/Simplicial_complex.thy @@ -1,793 +1,793 @@ theory Simplicial_complex imports Boolean_functions begin section\Simplicial Complexes\ lemma Pow_singleton: "Pow {a} = {{},{a}}" by auto lemma Pow_pair: "Pow {a,b} = {{},{a},{b},{a,b}}" by auto locale simplicial_complex = fixes n::"nat" begin text\A simplex (in $n$ vertexes) is any set of vertexes, including the empty set.\ definition simplices :: "nat set set" where "simplices = Pow {0.. simplices" unfolding simplices_def by simp lemma "{0.. simplices" unfolding simplices_def by simp lemma finite_simplex: assumes "\ \ simplices" shows "finite \" by (metis Pow_iff assms finite_atLeastLessThan finite_subset simplices_def) text\A simplicial complex (in $n$ vertexes) is a collection of sets of vertexes such that every subset of a set of vertexes also belongs to the simplicial complex.\ definition simplicial_complex :: "nat set set => bool" where "simplicial_complex K \ (\\\K. (\ \ simplices) \ (Pow \) \ K)" lemma finite_simplicial_complex: assumes "simplicial_complex K" shows "finite K" by (metis assms finite_Pow_iff finite_atLeastLessThan rev_finite_subset simplices_def simplicial_complex_def subsetI) lemma finite_simplices: assumes "simplicial_complex K" and "v \ K" shows "finite v" using assms finite_simplex simplicial_complex.simplicial_complex_def by blast definition simplicial_complex_set :: "nat set set set" where "simplicial_complex_set = (Collect simplicial_complex)" lemma simplicial_complex_empty_set: fixes K::"nat set set" assumes k: "simplicial_complex K" shows "K = {} \ {} \ K" using k unfolding simplicial_complex_def Pow_def by auto lemma simplicial_complex_monotone: fixes K::"nat set set" assumes k: "simplicial_complex K" and s: "s \ K" and rs: "r \ s" shows "r \ K" using k rs s unfolding simplicial_complex_def Pow_def by auto text\One example of simplicial complex with four simplices.\ lemma assumes three: "(3::nat) < n" shows "simplicial_complex {{},{0},{1},{2},{3}}" apply (simp_all add: Pow_singleton simplicial_complex_def simplices_def) using Suc_lessD three by presburger lemma "\ simplicial_complex {{0,1},{1}}" by (simp add: Pow_pair simplicial_complex_def) text\Another example of simplicial complex with five simplices.\ lemma assumes three: "(3::nat) < n" shows "simplicial_complex {{},{0},{1},{2},{3},{0,1}}" apply (simp add: Pow_pair Pow_singleton simplicial_complex_def simplices_def) using Suc_lessD three by presburger text\Another example of simplicial complex with ten simplices.\ lemma assumes three: "(3::nat) < n" shows "simplicial_complex {{2,3},{1,3},{1,2},{0,3},{0,2},{3},{2},{1},{0},{}}" apply (simp add: Pow_pair Pow_singleton simplicial_complex_def simplices_def) using Suc_lessD three by presburger end section\Simplicial complex induced by a monotone Boolean function\ text\In this section we introduce the definition of the simplicial complex induced by a monotone Boolean function, following the definition in Scoville~\cite[Def. 6.9]{SC19}.\ text\First we introduce the set of tuples for which a Boolean function is @{term False}.\ definition ceros_of_boolean_input :: "bool vec => nat set" where "ceros_of_boolean_input v = {x. x < dim_vec v \ vec_index v x = False}" lemma ceros_of_boolean_input_l_dim: assumes a: "a \ ceros_of_boolean_input v" shows "a < dim_vec v" using a unfolding ceros_of_boolean_input_def by simp lemma "ceros_of_boolean_input v = {x. x < dim_vec v \ \ vec_index v x}" unfolding ceros_of_boolean_input_def by simp lemma ceros_of_boolean_input_complementary: shows "ceros_of_boolean_input v = {x. x < dim_vec v} - {x. vec_index v x}" unfolding ceros_of_boolean_input_def by auto (*lemma ceros_in_UNIV: "ceros_of_boolean_input f \ (UNIV::nat set)" using subset_UNIV .*) lemma monotone_ceros_of_boolean_input: fixes r and s::"bool vec" assumes r_le_s: "r \ s" shows "ceros_of_boolean_input s \ ceros_of_boolean_input r" proof (intro subsetI, unfold ceros_of_boolean_input_def, intro CollectI, rule conjI) fix x assume "x \ {x. x < dim_vec s \ vec_index s x = False}" hence xl: "x < dim_vec s" and nr: "vec_index s x = False" by simp_all show "vec_index r x = False" using r_le_s nr xl unfolding less_eq_vec_def by auto show "x < dim_vec r" using r_le_s xl unfolding less_eq_vec_def by auto qed text\We introduce here instantiations of the typ\bool\ type for the type classes class\zero\ and class\one\ that will simplify notation at some points:\ instantiation bool :: "{zero,one}" begin definition zero_bool_def: "0 == False" definition one_bool_def: "1 == True" instance proof qed end text\Definition of the simplicial complex induced by a Boolean function \f\ in dimension \n\.\ definition simplicial_complex_induced_by_monotone_boolean_function :: "nat => (bool vec => bool) => nat set set" where "simplicial_complex_induced_by_monotone_boolean_function n f = {y. \x. dim_vec x = n \ f x \ ceros_of_boolean_input x = y}" text\The simplicial complex induced by a Boolean function is a subset of the powerset of the set of vertexes.\ lemma simplicial_complex_induced_by_monotone_boolean_function_subset: "simplicial_complex_induced_by_monotone_boolean_function n (v::bool vec => bool) \ Pow (({0..n}::nat set))" using ceros_of_boolean_input_def simplicial_complex_induced_by_monotone_boolean_function_def by force corollary "simplicial_complex_induced_by_monotone_boolean_function n (v::bool vec => bool) \ Pow ((UNIV::nat set))" by simp text\The simplicial complex induced by a monotone Boolean function is a simplicial complex. This result is proven in Scoville as part of the proof of Proposition 6.16~\cite[Prop. 6.16]{SC19}.\ context simplicial_complex begin lemma monotone_bool_fun_induces_simplicial_complex: assumes mon: "boolean_functions.monotone_bool_fun n f" shows "simplicial_complex (simplicial_complex_induced_by_monotone_boolean_function n f)" unfolding simplicial_complex_def proof (rule, unfold simplicial_complex_induced_by_monotone_boolean_function_def, safe) fix \ :: "nat set" and x :: "bool vec" assume fx: "f x" and dim_vec_x: "n = dim_vec x" show "ceros_of_boolean_input x \ simplicial_complex.simplices (dim_vec x)" using ceros_of_boolean_input_def dim_vec_x simplices_def by force next fix \ :: "nat set" and x :: "bool vec" and \ :: "nat set" assume fx: "f x" and dim_vec_x: "n = dim_vec x" and tau_def: "\ \ ceros_of_boolean_input x" show "\xb. dim_vec xb = dim_vec x \ f xb \ ceros_of_boolean_input xb = \" proof (rule exI [of _ "vec n (\i. if i \ \ then False else True)"], intro conjI) show "dim_vec (vec n (\i. if i \ \ then False else True)) = dim_vec x" unfolding dim_vec using dim_vec_x . - from mon have mono: "mono_on f (carrier_vec n)" + from mon have mono: "mono_on (carrier_vec n) f" unfolding boolean_functions.monotone_bool_fun_def . show "f (vec n (\i. if i \ \ then False else True))" proof - have "f x \ f (vec n (\i. if i \ \ then False else True))" proof (rule mono_onD [OF mono]) show "x \ carrier_vec n" using dim_vec_x by simp show "vec n (\i. if i \ \ then False else True) \ carrier_vec n" by simp show "x \ vec n (\i. if i \ \ then False else True)" using tau_def dim_vec_x unfolding ceros_of_boolean_input_def using less_eq_vec_def by fastforce qed thus ?thesis using fx by simp qed show "ceros_of_boolean_input (vec n (\i. if i \ \ then False else True)) = \" using \\ \ ceros_of_boolean_input x\ ceros_of_boolean_input_def dim_vec_x by auto qed qed end text\Example 6.10 in Scoville, the threshold function for $2$ in dimension $4$ (with vertexes $0$,$1$,$2$,$3$)\ definition bool_fun_threshold_2_3 :: "bool vec => bool" where "bool_fun_threshold_2_3 = (\v. if 2 \ count_true v then True else False)" lemma set_list_four: shows "{0..<4} = set [0,1,2,3::nat]" by auto lemma comp_fun_commute_lambda: "comp_fun_commute_on UNIV ((+) \ (\i. if vec 4 f $ i then 1 else (0::nat)))" unfolding comp_fun_commute_on_def by auto lemma "bool_fun_threshold_2_3 (vec 4 (\i. if i = 0 \ i = 1 then True else False)) = True" unfolding bool_fun_threshold_2_3_def unfolding count_true_def unfolding dim_vec unfolding sum.eq_fold using index_vec [of _ 4] apply auto unfolding set_list_four unfolding comp_fun_commute_on.fold_set_fold_remdups [OF comp_fun_commute_lambda, simplified] by simp lemma "0 \ ceros_of_boolean_input (vec 4 (\i. if i = 0 \ i = 1 then True else False))" and "1 \ ceros_of_boolean_input (vec 4 (\i. if i = 0 \ i = 1 then True else False))" and "2 \ ceros_of_boolean_input (vec 4 (\i. if i = 0 \ i = 1 then True else False))" and "3 \ ceros_of_boolean_input (vec 4 (\i. if i = 0 \ i = 1 then True else False))" and "{2,3} \ ceros_of_boolean_input (vec 4 (\i. if i = 0 \ i = 1 then True else False))" unfolding ceros_of_boolean_input_def by simp_all lemma "bool_fun_threshold_2_3 (vec 4 (\i. if i = 3 then True else False)) = False" unfolding bool_fun_threshold_2_3_def unfolding count_true_def unfolding dim_vec unfolding sum.eq_fold using index_vec [of _ 4] apply auto unfolding set_list_four unfolding comp_fun_commute_on.fold_set_fold_remdups [OF comp_fun_commute_lambda, simplified] by simp lemma "bool_fun_threshold_2_3 (vec 4 (\i. if i = 0 then False else True))" unfolding bool_fun_threshold_2_3_def unfolding count_true_def unfolding dim_vec unfolding sum.eq_fold using index_vec [of _ 4] apply auto unfolding set_list_four unfolding comp_fun_commute_on.fold_set_fold_remdups [OF comp_fun_commute_lambda, simplified] by simp section\The simplicial complex induced by the threshold function\ lemma empty_set_in_simplicial_complex_induced: "{} \ simplicial_complex_induced_by_monotone_boolean_function 4 bool_fun_threshold_2_3" unfolding simplicial_complex_induced_by_monotone_boolean_function_def unfolding bool_fun_threshold_2_3_def apply rule apply (rule exI [of _ "vec 4 (\x. True)"]) unfolding count_true_def ceros_of_boolean_input_def by auto lemma singleton_in_simplicial_complex_induced: assumes x: "x < 4" shows "{x} \ simplicial_complex_induced_by_monotone_boolean_function 4 bool_fun_threshold_2_3" (is "?A \ simplicial_complex_induced_by_monotone_boolean_function 4 bool_fun_threshold_2_3") proof (unfold simplicial_complex_induced_by_monotone_boolean_function_def, rule, rule exI [of _ "vec 4 (\i. if i \ ?A then False else True)"], intro conjI) show "dim_vec (vec 4 (\i. if i \ {x} then False else True)) = 4" by simp show "bool_fun_threshold_2_3 (vec 4 (\i. if i \ ?A then False else True))" unfolding bool_fun_threshold_2_3_def unfolding count_true_def unfolding dim_vec unfolding sum.eq_fold using index_vec [of _ 4] apply auto unfolding set_list_four unfolding comp_fun_commute_on.fold_set_fold_remdups [OF comp_fun_commute_lambda, simplified] by simp show "ceros_of_boolean_input (vec 4 (\i. if i \ ?A then False else True)) = ?A" unfolding ceros_of_boolean_input_def using x by auto qed lemma pair_in_simplicial_complex_induced: assumes x: "x < 4" and y: "y < 4" shows "{x,y} \ simplicial_complex_induced_by_monotone_boolean_function 4 bool_fun_threshold_2_3" (is "?A \ simplicial_complex_induced_by_monotone_boolean_function 4 bool_fun_threshold_2_3") proof (unfold simplicial_complex_induced_by_monotone_boolean_function_def, rule, rule exI [of _ "vec 4 (\i. if i \ ?A then False else True)"], intro conjI) show "dim_vec (vec 4 (\i. if i \ {x, y} then False else True)) = 4" by simp show "bool_fun_threshold_2_3 (vec 4 (\i. if i \ ?A then False else True))" unfolding bool_fun_threshold_2_3_def unfolding count_true_def unfolding dim_vec unfolding sum.eq_fold using index_vec [of _ 4] apply auto unfolding set_list_four unfolding comp_fun_commute_on.fold_set_fold_remdups [OF comp_fun_commute_lambda, simplified] by simp show "ceros_of_boolean_input (vec 4 (\i. if i \ ?A then False else True)) = ?A" unfolding ceros_of_boolean_input_def using x y by auto qed lemma finite_False: "finite {x. x < dim_vec a \ vec_index (a::bool vec) x = False}" by auto lemma finite_True: "finite {x. x < dim_vec a \ vec_index (a::bool vec) x = True}" by auto lemma UNIV_disjoint: "{x. x < dim_vec a \ vec_index (a::bool vec) x = True} \ {x. x < dim_vec a \ vec_index (a::bool vec) x = False} = {}" by auto lemma UNIV_union: "{x. x < dim_vec a \ vec_index (a::bool vec) x = True} \ {x. x < dim_vec a \ vec_index (a::bool vec) x = False} = {x. x < dim_vec a}" by auto lemma card_UNIV_union: "card {x. x < dim_vec a \ vec_index (a::bool vec) x = True} + card {x. x < dim_vec a \ vec_index (a::bool vec) x = False} = card {x. x < dim_vec a}" (is "card ?true + card ?false = _") proof - have "card ?true + card ?false = card (?true \ ?false) + card (?true \ ?false)" using card_Un_Int [OF finite_True [of a] finite_False [of a]] . also have "... = card {x. x < dim_vec a}" unfolding UNIV_union UNIV_disjoint by simp finally show ?thesis by simp qed lemma card_complementary: "card (ceros_of_boolean_input v) + card {x. x < (dim_vec v) \ (vec_index v x = True)} = (dim_vec v)" unfolding ceros_of_boolean_input_def using card_UNIV_union [of v] by simp corollary card_ceros_of_boolean_input: shows "card (ceros_of_boolean_input a) \ dim_vec a" using card_complementary [of a] by simp lemma vec_fun: assumes "v \ carrier_vec n" shows "\f. v = vec n f" using assms unfolding carrier_vec_def by fastforce corollary assumes "dim_vec v = n" shows "\f. v = vec n f" using carrier_vecI [OF assms] unfolding carrier_vec_def by fastforce lemma vec_l_eq: assumes "i < n" shows "vec (Suc n) f $ i = vec n f $ i" by (simp add: assms less_SucI) lemma card_boolean_function: assumes d: "v \ carrier_vec n" shows "card {x. x < n \ v $ x = True} = (\i = 0..mx. x \ carrier_vec m \ card {xa. xa < m \ x $ xa = True} = (\i = 0.. carrier_vec n" show "card {x. x < n \ v $ x = True} = (\i = 0.. carrier_vec n" obtain f :: "nat => bool" where v_f: "v = vec n f" using vec_fun [OF v] by auto have "card {x. x < m \ (vec m f) $ x = True} = (\i = 0.. vec (Suc m) f $ x = True} = ({x. x < m \ vec (Suc m) f $ x = True} \ {x. x = m \ (vec (Suc m) f) $ x = True})" by auto have two: "disjnt {x. x < m \ vec (Suc m) f $ x = True} {x. x = m \ (vec (Suc m) f) $ x = True}" using disjnt_iff by blast have "card {x. x < Suc m \ vec (Suc m) f $ x = True} = card {x. x < m \ (vec (Suc m) f) $ x = True} + card {x. x = m \ (vec (Suc m) f) $ x = True}" unfolding one by (rule card_Un_disjnt [OF _ _ two], simp_all) also have "... = card {x. x < m \ (vec m f) $ x = True} + 1" proof - have one: "{x. x < m \ vec (Suc m) f $ x = True} = {x. x < m \ vec m f $ x = True}" using vec_l_eq [of _ m] by auto have eq: "{x. x = m \ vec (Suc m) f $ x = True} = {m}" using True by auto hence two: "card {x. x = m \ vec (Suc m) f $ x = True} = 1" by simp show ?thesis using one two by simp qed finally have lhs: "card {x. x < Suc m \ vec (Suc m) f $ x = True} = card {x. x < m \ vec m f $ x = True} + 1" . have "(\i = 0..i = 0..i = 0..i = 0..i = 0.. vec (Suc m) f $ x = True} = (\i = 0.. vec (Suc m) f $ x = True} = ({x. x < m \ vec (Suc m) f $ x = True} \ {x. x = m \ (vec (Suc m) f) $ x = True})" by auto have two: "disjnt {x. x < m \ vec (Suc m) f $ x = True} {x. x = m \ (vec (Suc m) f) $ x = True}" using disjnt_iff by blast have "card {x. x < Suc m \ vec (Suc m) f $ x = True} = card {x. x < m \ (vec (Suc m) f) $ x = True} + card {x. x = m \ (vec (Suc m) f) $ x = True}" unfolding one by (rule card_Un_disjnt [OF _ _ two], simp_all) also have "... = card {x. x < m \ (vec m f) $ x = True} + 0" proof - have one: "{x. x < m \ vec (Suc m) f $ x = True} = {x. x < m \ vec m f $ x = True}" using vec_l_eq [of _ m] by auto have eq: "{x. x = m \ vec (Suc m) f $ x = True} = {}" using False by auto hence two: "card {x. x = m \ vec (Suc m) f $ x = True} = 0" by simp show ?thesis using one two by simp qed finally have lhs: "card {x. x < Suc m \ vec (Suc m) f $ x = True} = card {x. x < m \ vec m f $ x = True} + 0" . have "(\i = 0..i = 0..i = 0..i = 0..i = 0.. vec (Suc m) f $ x = True} = (\i = 0..We calculate the carrier set of the @{const ceros_of_boolean_input} function for dimensions $2$, $3$ and $4$.\ text\Vectors of dimension $2$.\ lemma dim_vec_2_cases: assumes dx: "dim_vec x = 2" shows "(x $ 0 = x $ 1 = True) \ (x $ 0 = False \ x $ 1 = True) \ (x $ 0 = True \ x $ 1 = False) \ (x $ 0 = x $ 1 = False)" by auto lemma tt_2: assumes dx: "dim_vec x = 2" and be: "x $ 0 = True \ x $ 1 = True" shows "ceros_of_boolean_input x = {}" using dx be unfolding ceros_of_boolean_input_def using less_2_cases by auto lemma tf_2: assumes dx: "dim_vec x = 2" and be: "x $ 0 = True \ x $ 1 = False" shows "ceros_of_boolean_input x = {1}" using dx be unfolding ceros_of_boolean_input_def using less_2_cases by auto lemma ft_2: assumes dx: "dim_vec x = 2" and be: "x $ 0 = False \ x $ 1 = True" shows "ceros_of_boolean_input x = {0}" using dx be unfolding ceros_of_boolean_input_def using less_2_cases by auto lemma ff_2: assumes dx: "dim_vec x = 2" and be: "x $ 0 = False \ x $ 1 = False" shows "ceros_of_boolean_input x = {0,1}" using dx be unfolding ceros_of_boolean_input_def using less_2_cases by auto lemma assumes dx: "dim_vec x = 2" shows "ceros_of_boolean_input x \ {{},{0},{1},{0,1}}" using dim_vec_2_cases [OF ] using tt_2 [OF dx] tf_2 [OF dx] ft_2 [OF dx] ff_2 [OF dx] by (metis insertCI) text\Vectors of dimension $3$.\ lemma less_3_cases: assumes n: "n < 3" shows "n = 0 \ n = 1 \ n = (2::nat)" using n by linarith lemma dim_vec_3_cases: assumes dx: "dim_vec x = 3" shows "(x $ 0 = x $ 1 = x $ 2 = False) \ (x $ 0 = x $ 1 = False \ x $ 2 = True) \ (x $ 0 = x $ 2 = False \ x $ 1 = True) \ (x $ 0 = False \ x $ 1 = x $ 2 = True) \ (x $ 0 = True \ x $ 1 = x $ 2 = False) \ (x $ 0 = x $ 2 = True \ x $ 1 = False) \ (x $ 0 = x $ 1 = True \ x $ 2 = False) \ (x $ 0 = x $ 1 = x $ 2 = True)" by auto lemma fff_3: assumes dx: "dim_vec x = 3" and be: "x $ 0 = False \ x $ 1 = False \ x $ 2 = False" shows "ceros_of_boolean_input x = {0,1,2}" using dx be unfolding ceros_of_boolean_input_def using less_3_cases by auto lemma fft_3: assumes dx: "dim_vec x = 3" and be: "x $ 0 = False \ x $ 1 = False \ x $ 2 = True" shows "ceros_of_boolean_input x = {0,1}" using dx be unfolding ceros_of_boolean_input_def using less_3_cases by auto lemma ftf_3: assumes dx: "dim_vec x = 3" and be: "x $ 0 = False \ x $ 1 = True \ x $ 2 = False" shows "ceros_of_boolean_input x = {0,2}" using dx be unfolding ceros_of_boolean_input_def using less_3_cases by fastforce lemma ftt_3: assumes dx: "dim_vec x = 3" and be: "x $ 0 = False \ x $ 1 = True \ x $ 2 = True" shows "ceros_of_boolean_input x = {0}" using dx be unfolding ceros_of_boolean_input_def using less_3_cases by auto lemma tff_3: assumes dx: "dim_vec x = 3" and be: "x $ 0 = True \ x $ 1 = False \ x $ 2 = False" shows "ceros_of_boolean_input x = {1,2}" using dx be unfolding ceros_of_boolean_input_def using less_3_cases by auto lemma tft_3: assumes dx: "dim_vec x = 3" and be: "x $ 0 = True \ x $ 1 = False \ x $ 2 = True" shows "ceros_of_boolean_input x = {1}" using dx be unfolding ceros_of_boolean_input_def using less_3_cases by auto lemma ttf_3: assumes dx: "dim_vec x = 3" and be: "x $ 0 = True \ x $ 1 = True \ x $ 2 = False" shows "ceros_of_boolean_input x = {2}" using dx be unfolding ceros_of_boolean_input_def using less_3_cases by fastforce lemma ttt_3: assumes dx: "dim_vec x = 3" and be: "x $ 0 = True \ x $ 1 = True \ x $ 2 = True" shows "ceros_of_boolean_input x = {}" using dx be unfolding ceros_of_boolean_input_def using less_3_cases by auto lemma assumes dx: "dim_vec x = 3" shows "ceros_of_boolean_input x \ {{},{0},{1},{2},{0,1},{0,2},{1,2},{0,1,2}}" using dim_vec_3_cases [OF ] using fff_3 [OF dx] fft_3 [OF dx] ftf_3 [OF dx] ftt_3 [OF dx] using tff_3 [OF dx] tft_3 [OF dx] ttf_3 [OF dx] ttt_3 [OF dx] by (smt (z3) insertCI) text\Vectors of dimension $4$.\ lemma less_4_cases: assumes n: "n < 4" shows "n = 0 \ n = 1 \ n = 2 \ n = (3::nat)" using n by linarith lemma dim_vec_4_cases: assumes dx: "dim_vec x = 4" shows "(x $ 0 = x $ 1 = x $ 2 = x $ 3 = False) \ (x $ 0 = x $ 1 = x $ 2 = False \ x $ 3 = True) \ (x $ 0 = x $ 1 = x $ 3 = False \ x $ 2 = True) \ (x $ 0 = x $ 1 = False \ x $ 2 = x $ 3 = True) \ (x $ 0 = x $ 2 = x $ 3 = False \ x $ 1 = True) \ (x $ 0 = x $ 2 = False \ x $ 1 = x $ 3 = True) \ (x $ 0 = x $ 3 = False \ x $ 1 = x $ 2 = True) \ (x $ 0 = False \ x $ 1 = x $ 2 = x $ 3 = True) \ (x $ 0 = True \ x $ 1 = x $ 2 = x $ 3 = False) \ (x $ 0 = x $ 3 = True \ x $ 1 = x $ 2 = False) \ (x $ 0 = x $ 2 = True \ x $ 1 = x $ 3 = False) \ (x $ 0 = x $ 2 = x $ 3 = True \ x $ 1 = False) \ (x $ 0 = x $ 1 = True \ x $ 2 = x $ 3 = False) \ (x $ 0 = x $ 1 = x $ 3 = True \ x $ 2 = False) \ (x $ 0 = x $ 1 = x $ 2 = True \ x $ 3 = False) \ (x $ 0 = x $ 1 = x $ 2 = x $ 3 = True)" by blast lemma ffff_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = False \ x $ 1 = False \ x $ 2 = False \ x $ 3 = False" shows "ceros_of_boolean_input x = {0,1,2,3}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma ffft_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = False \ x $ 1 = False \ x $ 2 = False \ x $ 3 = True" shows "ceros_of_boolean_input x = {0,1,2}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma fftf_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = False \ x $ 1 = False \ x $ 2 = True \ x $ 3 = False" shows "ceros_of_boolean_input x = {0,1,3}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma fftt_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = False \ x $ 1 = False \ x $ 2 = True \ x $ 3 = True" shows "ceros_of_boolean_input x = {0,1}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma ftff_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = False \ x $ 1 = True \ x $ 2 = False \ x $ 3 = False" shows "ceros_of_boolean_input x = {0,2,3}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma ftft_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = False \ x $ 1 = True \ x $ 2 = False \ x $ 3 = True" shows "ceros_of_boolean_input x = {0,2}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma fttf_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = False \ x $ 1 = True \ x $ 2 = True \ x $ 3 = False" shows "ceros_of_boolean_input x = {0,3}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma fttt_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = False \ x $ 1 = True \ x $ 2 = True \ x $ 3 = True" shows "ceros_of_boolean_input x = {0}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma tfff_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = True \ x $ 1 = False \ x $ 2 = False \ x $ 3 = False" shows "ceros_of_boolean_input x = {1,2,3}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma tfft_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = True \ x $ 1 = False \ x $ 2 = False \ x $ 3 = True" shows "ceros_of_boolean_input x = {1,2}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma tftf_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = True \ x $ 1 = False \ x $ 2 = True \ x $ 3 = False" shows "ceros_of_boolean_input x = {1,3}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma tftt_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = True \ x $ 1 = False \ x $ 2 = True \ x $ 3 = True" shows "ceros_of_boolean_input x = {1}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma ttff_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = True \ x $ 1 = True \ x $ 2 = False \ x $ 3 = False" shows "ceros_of_boolean_input x = {2,3}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma ttft_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = True \ x $ 1 = True \ x $ 2 = False \ x $ 3 = True" shows "ceros_of_boolean_input x = {2}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma tttf_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = True \ x $ 1 = True \ x $ 2 = True \ x $ 3 = False" shows "ceros_of_boolean_input x = {3}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma tttt_4: assumes dx: "dim_vec x = 4" and be: "x $ 0 = True \ x $ 1 = True \ x $ 2 = True \ x $ 3 = True" shows "ceros_of_boolean_input x = {}" using dx be unfolding ceros_of_boolean_input_def using less_4_cases by auto lemma ceros_of_boolean_input_set: assumes dx: "dim_vec x = 4" shows "ceros_of_boolean_input x \ {{},{0},{1},{2},{3},{0,1},{0,2},{0,3},{1,2},{1,3},{2,3}, {0,1,2},{0,1,3},{0,2,3},{1,2,3},{0,1,2,3}}" using dim_vec_4_cases [OF ] using ffff_4 [OF dx] ffft_4 [OF dx] fftf_4 [OF dx] fftt_4 [OF dx] using ftff_4 [OF dx] ftft_4 [OF dx] fttf_4 [OF dx] fttt_4 [OF dx] using tfff_4 [OF dx] tfft_4 [OF dx] tftf_4 [OF dx] tftt_4 [OF dx] using ttff_4 [OF dx] ttft_4 [OF dx] tttf_4 [OF dx] tttt_4 [OF dx] by (smt (z3) insertCI) context simplicial_complex begin text\The simplicial complex induced by the monotone Boolean function @{const bool_fun_threshold_2_3} has the following explicit expression.\ lemma simplicial_complex_induced_by_monotone_boolean_function_4_bool_fun_threshold_2_3: shows "{{},{0},{1},{2},{3},{0,1},{0,2},{0,3},{1,2},{1,3},{2,3}} = simplicial_complex_induced_by_monotone_boolean_function 4 bool_fun_threshold_2_3" (is "{{},?a,?b,?c,?d,?e,?f,?g,?h,?i,?j} = _") proof (rule) show "{{},?a,?b,?c,?d,?e,?f,?g,?h,?i,?j} \ simplicial_complex_induced_by_monotone_boolean_function 4 bool_fun_threshold_2_3" by (simp add: empty_set_in_simplicial_complex_induced singleton_in_simplicial_complex_induced pair_in_simplicial_complex_induced)+ show "simplicial_complex_induced_by_monotone_boolean_function 4 bool_fun_threshold_2_3 \ {{},?a,?b,?c,?d,?e,?f,?g,?h,?i,?j}" unfolding simplicial_complex_induced_by_monotone_boolean_function_def unfolding bool_fun_threshold_2_3_def proof fix y::"nat set" assume y: "y \ {y. \x. dim_vec x = 4 \ (if 2 \ count_true x then True else False) \ ceros_of_boolean_input x = y}" then obtain x::"bool vec" where ct_ge_2: "(if 2 \ count_true x then True else False)" and cx: "ceros_of_boolean_input x = y" and dx: "dim_vec x = 4" by auto have "count_true x + card (ceros_of_boolean_input x) = dim_vec x" using card_ceros_count_UNIV [of x] by simp hence "card (ceros_of_boolean_input x) \ 2" using ct_ge_2 using card_boolean_function using dx by presburger hence card_le: "card y \ 2" using cx by simp have "y \ {{},?a,?b,?c,?d,?e,?f,?g,?h,?i,?j}" proof (rule ccontr) assume "y \ {{},?a,?b,?c,?d,?e,?f,?g,?h,?i,?j}" then have y_nin: "y \ set [{},?a,?b,?c,?d,?e,?f,?g,?h,?i,?j]" by simp have "y \ set [{0,1,2},{0,1,3},{0,2,3},{1,2,3},{0,1,2,3}]" using ceros_of_boolean_input_set [OF dx] y_nin unfolding cx by simp hence "card y \ 3" by auto thus False using card_le by simp qed then show "y \ {{},?a,?b,?c,?d,?e,?f,?g,?h,?i,?j}" by simp qed qed end end diff --git a/thys/Skip_Lists/Misc.thy b/thys/Skip_Lists/Misc.thy --- a/thys/Skip_Lists/Misc.thy +++ b/thys/Skip_Lists/Misc.thy @@ -1,192 +1,192 @@ (* File: Misc.thy Authors: Max W. Haslbeck, Manuel Eberl *) section \Auxiliary material\ theory Misc imports "HOL-Analysis.Analysis" begin text \Based on @{term sorted_list_of_set} and @{term the_inv_into} we construct a bijection between a finite set A of type 'a::linorder and a set of natural numbers @{term "{..< card A}"}\ lemma bij_betw_mono_on_the_inv_into: fixes A::"'a::linorder set" and B::"'b::linorder set" - assumes b: "bij_betw f A B" and m: "mono_on f A" - shows "mono_on (the_inv_into A f) B" + assumes b: "bij_betw f A B" and m: "mono_on A f" + shows "mono_on B (the_inv_into A f)" proof (rule ccontr) - assume "\ mono_on (the_inv_into A f) B" + assume "\ mono_on B (the_inv_into A f)" then have "\r s. r \ B \ s \ B \ r \ s \ \ the_inv_into A f s \ the_inv_into A f r" unfolding mono_on_def by blast then obtain r s where rs: "r \ B" "s \ B" "r \ s" "the_inv_into A f s < the_inv_into A f r" by fastforce have f: "f (the_inv_into A f b) = b" if "b \ B" for b using that assms f_the_inv_into_f_bij_betw by metis have "the_inv_into A f s \ A" "the_inv_into A f r \ A" using rs assms by (auto simp add: bij_betw_def the_inv_into_into) then have "f (the_inv_into A f s) \ f (the_inv_into A f r)" using rs by (intro mono_onD[OF m]) (auto) then have "r = s" using rs f by simp then show False using rs by auto qed lemma rev_removeAll_removeAll_rev: "rev (removeAll x xs) = removeAll x (rev xs)" by (simp add: removeAll_filter_not_eq rev_filter) lemma sorted_list_of_set_Min_Cons: assumes "finite A" "A \ {}" shows "sorted_list_of_set A = Min A # sorted_list_of_set (A - {Min A})" proof - have *: "A = insert (Min A) A" using assms Min_in by (auto) then have "sorted_list_of_set A = insort (Min A) (sorted_list_of_set (A - {Min A}))" using assms by (subst *, intro sorted_list_of_set_insert_remove) auto also have "\ = Min A # sorted_list_of_set (A - {Min A})" using assms by (intro insort_is_Cons) (auto) finally show ?thesis by simp qed lemma sorted_list_of_set_filter: assumes "finite A" shows "sorted_list_of_set ({x\A. P x}) = filter P (sorted_list_of_set A)" using assms proof (induction "sorted_list_of_set A" arbitrary: A) case (Cons x xs) have x: "x \ A" using Cons sorted_list_of_set list.set_intros(1) by metis have "sorted_list_of_set A = Min A # sorted_list_of_set (A - {Min A})" using Cons by (intro sorted_list_of_set_Min_Cons) auto then have 1: "x = Min A" "xs = sorted_list_of_set (A - {x})" using Cons by auto { assume Px: "P x" have 2: "sorted_list_of_set {x \ A. P x} = Min {x \ A. P x} # sorted_list_of_set ({x \ A. P x} - {Min {x \ A. P x}})" using Px Cons 1 sorted_list_of_set_eq_Nil_iff by (intro sorted_list_of_set_Min_Cons) fastforce+ also have 3: "Min {x \ A. P x} = x" using Cons 1 Px x by (auto intro!: Min_eqI) also have 4: "{x \ A. P x} - {x} = {y \ A - {x}. P y}" by blast also have 5: "sorted_list_of_set {y \ A - {x}. P y} = filter P (sorted_list_of_set (A - {x}))" using 1 Cons by (intro Cons) (auto) also have "\ = filter P xs" using 1 by simp also have "filter P (sorted_list_of_set A) = x # filter P xs" using Px by (simp flip: \x # xs = sorted_list_of_set A\) finally have ?case by auto } moreover { assume Px: "\ P x" then have "{x \ A. P x} = {y \ A - {x}. P y}" by blast also have "sorted_list_of_set \ = filter P (sorted_list_of_set (A - {x}))" using 1 Cons by (intro Cons) auto also have "filter P (sorted_list_of_set (A - {x})) = filter P (sorted_list_of_set A)" using 1 Px by (simp flip: \x # xs = sorted_list_of_set A\) finally have ?case by simp } ultimately show ?case by blast qed (use sorted_list_of_set_eq_Nil_iff in fastforce) lemma sorted_list_of_set_Max_snoc: assumes "finite A" "A \ {}" shows "sorted_list_of_set A = sorted_list_of_set (A - {Max A}) @ [Max A]" proof - have *: "A = insert (Max A) A" using assms Max_in by (auto) then have "sorted_list_of_set A = insort (Max A) (sorted_list_of_set (A - {Max A}))" using assms by (subst *, intro sorted_list_of_set_insert_remove) auto also have "\ = sorted_list_of_set (A - {Max A}) @ [Max A]" using assms by (intro sorted_insort_is_snoc) (auto) finally show ?thesis by simp qed lemma sorted_list_of_set_image: - assumes "mono_on g A" "inj_on g A" + assumes "mono_on A g" "inj_on g A" shows "(sorted_list_of_set (g ` A)) = map g (sorted_list_of_set A)" proof (cases "finite A") case True then show ?thesis using assms proof (induction "sorted_list_of_set A" arbitrary: A) case Nil then show ?case using sorted_list_of_set_eq_Nil_iff by fastforce next case (Cons x xs A) have not_empty_A: "A \ {}" using Cons sorted_list_of_set_eq_Nil_iff by auto have *: "Min (g ` A) = g (Min A)" proof - have "g (Min A) \ g a" if "a \ A" for a - using that Cons Min_in Min_le not_empty_A by (auto intro!: mono_onD[of g]) + using that Cons Min_in Min_le not_empty_A by (auto intro!: mono_onD[of _ g]) then show ?thesis using Cons not_empty_A by (intro Min_eqI) auto qed have "g ` A \ {}" "finite (g ` A)" using Cons by auto then have "(sorted_list_of_set (g ` A)) = Min (g ` A) # sorted_list_of_set ((g ` A) - {Min (g ` A)})" by (auto simp add: sorted_list_of_set_Min_Cons) also have "(g ` A) - {Min (g ` A)} = g ` (A - {Min A})" using Cons Min_in not_empty_A * by (subst inj_on_image_set_diff[of _ A]) auto also have "sorted_list_of_set (g ` (A - {Min A})) = map g (sorted_list_of_set (A - {Min A}))" - using not_empty_A Cons mono_on_subset[of _ A "A - {Min A}"] inj_on_subset[of _ A "A - {Min A}"] + using not_empty_A Cons mono_on_subset[of A _ "A - {Min A}"] inj_on_subset[of _ A "A - {Min A}"] by (intro Cons) (auto simp add: sorted_list_of_set_Min_Cons) finally show ?case using Cons not_empty_A * by (auto simp add: sorted_list_of_set_Min_Cons) qed next case False then show ?thesis using assms by (simp add: finite_image_iff) qed lemma sorted_list_of_set_length: "length (sorted_list_of_set A) = card A" using distinct_card sorted_list_of_set[of A] by (cases "finite A") fastforce+ lemma sorted_list_of_set_bij_betw: assumes "finite A" shows "bij_betw (\n. sorted_list_of_set A ! n) {..n. xs ! n) {..n. xs ! n)" using assms by (intro mono_onI sorted_nth_mono) (auto simp add: distinct_card) lemma sorted_list_of_set_mono_on: - "finite A \ mono_on (\n. sorted_list_of_set A ! n) {.. mono_on {..n. sorted_list_of_set A ! n)" by (rule nth_mono_on) (auto) definition bij_mono_map_set_to_nat :: "'a::linorder set \ 'a \ nat" where "bij_mono_map_set_to_nat A = (\x. if x \ A then the_inv_into {..Randomized Skip Lists\ theory Skip_List imports Geometric_PMF Misc "Monad_Normalisation.Monad_Normalisation" begin subsection \Preliminaries\ lemma bind_pmf_if': "(do {c \ C; ab \ (if c then A else B); D ab}::'a pmf) = do {c \ C; (if c then (A \ D) else (B \ D))}" by (metis (mono_tags, lifting)) abbreviation (input) Max\<^sub>0 where "Max\<^sub>0 \ (\A. Max (A \ {0}))" subsection \Definition of a Randomised Skip List\ text \ Given a set A we assign a geometric random variable (counting the number of failed Bernoulli trials before the first success) to every element in A. That means an arbitrary element of A is on level n with probability $(1-p)^{n}p$. We define he height of the skip list as the maximum assigned level. So a skip list with only one level has height 0 but the calculation of the expected height is cleaner this way. \ locale random_skip_list = fixes p::real begin definition q where "q = 1 - p" definition SL :: "('a::linorder) set \ ('a \ nat) pmf" where "SL A = Pi_pmf A 0 (\_. geometric_pmf p)" definition SL\<^sub>N :: "nat \ (nat \ nat) pmf" where "SL\<^sub>N n = SL {..Height of Skip List\ definition H where "H A = map_pmf (\f. Max\<^sub>0 (f ` A)) (SL A)" definition H\<^sub>N :: "nat \ nat pmf" where "H\<^sub>N n = H {.. The height of a skip list is independent of the values in a set A. For simplicity we can therefore work on the skip list over the set @{term "{..< card A}"} \ lemma assumes "finite A" shows "H A = H\<^sub>N (card A)" proof - define f' where "f' = (\x. if x \ A then the_inv_into {..0 ((f \ f') ` A) = Max\<^sub>0 (f ` {.. nat" using bij_betw_imp_surj_on bij_f' image_comp by metis have "H A = map_pmf (\f. Max\<^sub>0 (f ` A)) (map_pmf (\g. g \ f') (SL\<^sub>N (card A)))" using assms bij_f' unfolding H_def SL_def SL\<^sub>N_def by (subst Pi_pmf_bij_betw[of _ f' "{.. = H\<^sub>N (card A)" unfolding H\<^sub>N_def H_def SL\<^sub>N_def using * by (auto intro!: bind_pmf_cong simp add: map_pmf_def) finally show ?thesis by simp qed text \ The cumulative distribution function (CDF) of the height is the CDF of the geometric PMF to the power of n \ lemma prob_Max_IID_geometric_atMost: assumes "p \ {0..1}" shows "measure_pmf.prob (H\<^sub>N n) {..i} = (measure_pmf.prob (geometric_pmf p) {..i}) ^ n" (is "?lhs = ?rhs") proof - note SL_def[simp] SL\<^sub>N_def[simp] H_def[simp] H\<^sub>N_def[simp] have "{f. Max\<^sub>0 (f ` {.. i} = {.. {..i}" by auto then have "?lhs = measure_pmf.prob (SL\<^sub>N n) ({.. {..i})" by (simp add: vimage_def) also have "\ = measure_pmf.prob (SL\<^sub>N n) (PiE_dflt {.._. {..i}))" by (intro measure_prob_cong_0) (auto simp add: PiE_dflt_def pmf_Pi split: if_splits) also have "\ = measure_pmf.prob (geometric_pmf p) {..i} ^ n" using assms by (auto simp add: measure_Pi_pmf_PiE_dflt) finally show ?thesis by simp qed lemma prob_Max_IID_geometric_greaterThan: assumes "p \ {0<..1}" shows "measure_pmf.prob (H\<^sub>N n) {i<..} = 1 - (1 - q ^ (i + 1)) ^ n" proof - have "UNIV - {..i} = {i<..}" by auto then have "measure_pmf.prob (H\<^sub>N n) {i<..} = measure_pmf.prob (H\<^sub>N n) (space (measure_pmf (H\<^sub>N n)) - {..i})" by (auto) also have "\ = 1 - (measure_pmf.prob (geometric_pmf p) {..i}) ^ n" using assms by (subst measure_pmf.prob_compl) (auto simp add: prob_Max_IID_geometric_atMost) also have "\ = 1 - (1 - q ^ (i + 1)) ^ n" using assms unfolding q_def by (subst geometric_pmf_prob_atMost) auto finally show ?thesis by simp qed end (* context includes monad_normalisation *) end (* locale skip_list *) text \ An alternative definition of the expected value of a non-negative random variable \footnote{\url{https://en.wikipedia.org/w/index.php?title=Expected\_value&oldid=881384346\#Formula\_for\_non-negative\_random\_variables}} \ lemma expectation_prob_atLeast: assumes "(\i. measure_pmf.prob N {i..}) abs_summable_on {1..}" shows "measure_pmf.expectation N real = infsetsum (\i. measure_pmf.prob N {i..}) {1..}" "integrable N real" proof - have "(\(x, y). pmf N y) abs_summable_on Sigma {Suc 0..} atLeast" using assms by (auto simp add: measure_pmf_conv_infsetsum abs_summable_on_Sigma_iff) then have summable: "(\(x, y). pmf N x) abs_summable_on Sigma {Suc 0..} (atLeastAtMost (Suc 0))" by (subst abs_summable_on_reindex_bij_betw[of "\(x,y). (y,x)", symmetric]) (auto intro!: bij_betw_imageI simp add: inj_on_def case_prod_beta) have "measure_pmf.expectation N real = (\\<^sub>ax. pmf N x *\<^sub>R real x)" by (auto simp add: infsetsum_def integral_density measure_pmf_eq_density) also have "\ = (\\<^sub>ax \ ({0} \ {Suc 0..}). pmf N x *\<^sub>R real x)" by (auto intro!: infsetsum_cong) also have "\ = (\\<^sub>ax\{Suc 0..}. pmf N x * real x)" proof - have "(\x. pmf N x *\<^sub>R real x) abs_summable_on {0} \ {Suc 0..}" using summable by (subst (asm) abs_summable_on_Sigma_iff) (auto simp add: mult.commute) then show ?thesis by (subst infsetsum_Un_Int) auto qed also have "\ = (\\<^sub>a(x, y)\Sigma {Suc 0..} (atLeastAtMost (Suc 0)). pmf N x)" using summable by (subst infsetsum_Sigma) (auto simp add: mult.commute) also have "\ = (\\<^sub>ax\Sigma {Suc 0..} atLeast. pmf N (snd x))" by (subst infsetsum_reindex_bij_betw[of "\(x,y). (y,x)", symmetric]) (auto intro!: bij_betw_imageI simp add: inj_on_def case_prod_beta) also have "\ = infsetsum (\i. measure_pmf.prob N {i..}) {1..}" using assms by (subst infsetsum_Sigma) (auto simp add: measure_pmf_conv_infsetsum abs_summable_on_Sigma_iff infsetsum_Sigma') finally show "measure_pmf.expectation N real = infsetsum (\i. measure_pmf.prob N {i..}) {1..}" by simp have "(\x. pmf N x *\<^sub>R real x) abs_summable_on {0} \ {Suc 0..}" using summable by (subst (asm) abs_summable_on_Sigma_iff) (auto simp add: mult.commute) then have "(\x. pmf N x *\<^sub>R real x) abs_summable_on UNIV" by (simp add: atLeast_Suc) then have "integrable (count_space UNIV) (\x. pmf N x *\<^sub>R real x)" by (subst abs_summable_on_def[symmetric]) blast then show "integrable N real" by (subst measure_pmf_eq_density, subst integrable_density) auto qed text \ The expected height of a skip list has no closed-form expression but we can approximate it. We start by showing how we can calculate an infinite sum over the natural numbers with an integral over the positive reals and the floor function. \ lemma infsetsum_set_nn_integral_reals: assumes "f abs_summable_on UNIV" "\n. f n \ 0" shows "infsetsum f UNIV = set_nn_integral lborel {0::real..} (\x. f (nat (floor x)))" proof - have "x < 1 + (floor x)"for x::real by linarith then have "\n. real n \ x \ x < 1 + real n" if "x \ 0" for x using that of_nat_floor by (intro exI[of _ "nat (floor x)"]) auto then have "{0..} = (\n. {real n..\<^sup>+x\{0::real..}. ennreal (f (nat \x\))\lborel = (\n. \\<^sup>+x\{real n..<1 + real n}. ennreal (f (nat \x\))\lborel)" by (auto simp add: disjoint_family_on_def nn_integral_disjoint_family) also have "\ = (\n. \\<^sup>+x\{real n..<1 + real n}. ennreal (f n)\lborel)" by(subst suminf_cong, rule nn_integral_cong_AE) (auto intro!: eventuallyI simp add: indicator_def floor_eq4) also have "\ = (\n. ennreal (f n))" by (auto intro!: suminf_cong simp add: nn_integral_cmult) also have "\ = infsetsum f {0..}" using assms suminf_ennreal2 abs_summable_on_nat_iff' summable_norm_cancel by (auto simp add: infsetsum_nat) finally show ?thesis by simp qed lemma nn_integral_nats_reals: shows "(\\<^sup>+ i. ennreal (f i) \count_space UNIV) = \\<^sup>+x\{0::real..}. ennreal (f (nat \x\))\lborel" proof - have "x < 1 + (floor x)"for x::real by linarith then have "\n. real n \ x \ x < 1 + real n" if "x \ 0" for x using that of_nat_floor by (intro exI[of _ "nat (floor x)"]) auto then have "{0..} = (\n. {real n..\<^sup>+x\{0::real..}. f (nat \x\)\lborel = (\n. \\<^sup>+x\{real n..<1 + real n}. ennreal (f (nat \x\))\lborel)" by (auto simp add: disjoint_family_on_def nn_integral_disjoint_family) also have "\ = (\n. \\<^sup>+x\{real n..<1 + real n}. ennreal (f n)\lborel)" by(subst suminf_cong,rule nn_integral_cong_AE) (auto intro!: eventuallyI simp add: indicator_def floor_eq4) also have "\ = (\n. ennreal (f n))" by (auto intro!: suminf_cong simp add: nn_integral_cmult) also have "\ = (\\<^sup>+ i. ennreal (f i) \count_space UNIV)" by (simp add: nn_integral_count_space_nat) finally show ?thesis by simp qed lemma nn_integral_floor_less_eq: assumes "\x y. x \ y \ f y \ f x" shows "\\<^sup>+x\{0::real..}. ennreal (f x)\lborel \ \\<^sup>+x\{0::real..}. ennreal (f (nat \x\))\lborel" using assms by (auto simp add: indicator_def intro!: nn_integral_mono ennreal_leI) lemma nn_integral_finite_imp_abs_sumable_on: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "nn_integral (count_space A) (\x. norm (f x)) < \" shows "f abs_summable_on A" using assms unfolding abs_summable_on_def integrable_iff_bounded by auto lemma nn_integral_finite_imp_abs_sumable_on': assumes "nn_integral (count_space A) (\x. ennreal (f x)) < \" "\x. f x \ 0" shows "f abs_summable_on A" using assms unfolding abs_summable_on_def integrable_iff_bounded by auto text \ We now show that $\int_0^\infty 1 - (1 - q^x) ^ n\;dx = \frac{- H_n}{\ln q}$ if $0 < q < 1$. \ lemma harm_integral_x_raised_n: "set_integrable lborel {0::real..1} (\x. (\i\{..i\{..x. (\i\{..i\{..i\{..x. (\i\{..auto simp add: einterval_def\) moreover have "set_integrable lborel (einterval 0 1) (\x. (x ^ n))" proof - have "set_integrable lborel {0::real..1} (\x. (x ^ n))" by (rule borel_integrable_atLeastAtMost') (auto intro!: borel_integrable_atLeastAtMost' continuous_intros) then show ?thesis by (rule set_integrable_subset) (auto simp add: einterval_def) qed ultimately show ?thesis by (auto intro!: borel_integrable_atLeastAtMost' simp add: interval_lebesgue_integrable_def) qed also have "(LBINT x=0..1. x ^ n) = 1 / (1 + real n)" proof - have "(LBINT x=0..1. x ^ n) = LBINT x. x ^ n * indicator {0..1} x " proof - have "AE x in lborel. x ^ n * indicator {0..1} x = indicator (einterval 0 1) x * x ^ n" by(rule eventually_mono[OF eventually_conj[OF AE_lborel_singleton[of 1] AE_lborel_singleton[of 0]]]) (auto simp add: indicator_def einterval_def) then show ?thesis using integral_cong_AE unfolding interval_lebesgue_integral_def set_lebesgue_integral_def by (auto intro!: integral_cong_AE) qed then show ?thesis by (auto simp add: integral_power) qed finally show ?case using Suc by (auto simp add: harm_def inverse_eq_divide) qed (auto simp add: harm_def) qed lemma harm_integral_0_1_fraction: "set_integrable lborel {0::real..1} (\x. (1 - x ^ n) / (1 - x))" "(LBINT x = 0..1. ((1 - x ^ n) / (1 - x))) = harm n" proof - show "set_integrable lborel {0::real..1} (\x. (1 - x ^ n) / (1 - x))" proof - have "AE x\{0::real..1} in lborel. (1 - x ^ n) / (1 - x) = sum ((^) x) {..{0::real<..<1} in lborel. (1 - x ^ n) / (1 - x) = sum ((^) x) {.. {0<..<1}" shows "set_integrable lborel (einterval 0 \) (\x. (1 - (1 - q powr x) ^ n))" "(LBINT x=0..\. 1 - (1 - q powr x) ^ n) = - harm n / ln q" proof - have [simp]: "q powr (log q (1-x)) = 1 - x" if "x \ {0<..<1}" for x using that assms by (subst powr_log_cancel) auto have 1: "((ereal \ (\x. log q (1 - x)) \ real_of_ereal) \ 0) (at_right 0)" using assms unfolding zero_ereal_def ereal_tendsto_simps by (auto intro!: tendsto_eq_intros) have 2: "((ereal \ (\x. log q (1-x)) \ real_of_ereal) \ \) (at_left 1)" proof - have "filterlim ((-) 1) (at_right 0) (at_left (1::real))" by (intro filterlim_at_withinI eventually_at_leftI[of 0]) (auto intro!: tendsto_eq_intros) then have "LIM x at_left 1. - inverse (ln q) * - ln (1 - x) :> at_top" using assms by (intro filterlim_tendsto_pos_mult_at_top [OF tendsto_const]) (auto simp: filterlim_uminus_at_top intro!: filterlim_compose[OF ln_at_0]) then show ?thesis unfolding one_ereal_def ereal_tendsto_simps log_def by (simp add: field_simps) qed have 3: "set_integrable lborel (einterval 0 1) (\x. (1 - (1 - q powr (log q (1 - x))) ^ n) * (- 1 / (ln q * (1 - x))))" proof - have "set_integrable lborel (einterval 0 1) (\x. - (1 / ln q) * ((1 - x ^ n) / (1 - x)))" by(intro set_integrable_mult_right) (auto intro!: harm_integral_0_1_fraction intro: set_integrable_subset simp add: einterval_def) then show ?thesis by(subst set_integrable_cong_AE[where g="\x. - (1 / ln q) * ((1 - x ^ n) / (1 - x))"]) (auto intro!: eventuallyI simp add: einterval_def) qed have 4: "LBINT x=0..1. - ((1 - (1 - q powr log q (1 - x)) ^ n) / (ln q * (1 - x))) = - (harm n / ln q)" (is "?lhs = ?rhs") proof - have "?lhs = LBINT x=0..1. ((1 - x ^ n) / (1 - x)) * (- 1 / ln q)" using assms by (intro interval_integral_cong_AE) (auto intro!: eventuallyI simp add: max_def einterval_def field_simps) also have "\ = harm n * (-1 / ln q)" using harm_integral_0_1_fraction by (subst interval_lebesgue_integral_mult_left) auto finally show ?thesis by auto qed note sub = interval_integral_substitution_nonneg [where f = "(\x. (1 - (1 - q powr x) ^ n))" and g="(\x. log q (1-x))" and g'="(\x. - 1 / (ln q * (1 - x)))" and a = 0 and b = 1] show "set_integrable lborel (einterval 0 \) (\x. (1 - (1 - q powr x) ^ n))" using assms 1 2 3 4 by (intro sub) (auto intro!: derivative_eq_intros mult_nonneg_nonpos2 tendsto_intros power_le_one) show "(LBINT x=0..\. 1 - (1 - q powr x) ^ n) = - harm n / ln q" using assms 1 2 3 4 by (subst sub) (auto intro!: derivative_eq_intros mult_nonneg_nonpos2 tendsto_intros power_le_one) qed lemma one_minus_one_minus_q_x_n_nn_integral: fixes q::real assumes "q \ {0<..<1}" shows "set_nn_integral lborel {0..} (\x. (1 - (1 - q powr x) ^ n)) = LBINT x=0..\. 1 - (1 - q powr x) ^ n" proof - have "set_nn_integral lborel {0..} (\x. (1 - (1 - q powr x) ^ n)) = nn_integral lborel (\x. indicator (einterval 0 \) x * (1 - (1 - q powr x) ^ n))" using assms by (intro nn_integral_cong_AE eventually_mono[OF AE_lborel_singleton[of 0]]) (auto simp add: indicator_def einterval_def) also have "\ = ennreal (LBINT x. indicator (einterval 0 \) x * (1 - (1 - q powr x) ^ n))" using one_minus_one_minus_q_x_n_integral assms by(intro nn_integral_eq_integral) (auto simp add: indicator_def einterval_def set_integrable_def intro!: eventuallyI power_le_one powr_le1) finally show ?thesis by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def) qed text \ We can now derive bounds for the expected height. \ context random_skip_list begin definition EH\<^sub>N where "EH\<^sub>N n = measure_pmf.expectation (H\<^sub>N n) real" lemma EH\<^sub>N_bounds': fixes n::nat assumes "p \ {0<..<1}" "0 < n" shows "- harm n / ln q - 1 \ EH\<^sub>N n" "EH\<^sub>N n \ - harm n / ln q" "integrable (H\<^sub>N n) real" proof - define f where "f = (\x. 1 - (1 - q ^ x) ^ n)" define f' where "f' = (\x. 1 - (1 - q powr x) ^ n)" have q: "q \ {0<..<1}" unfolding q_def using assms by auto have f_descending: "f y \ f x" if "x \ y" for x y unfolding f_def using that q by (auto intro!: power_mono simp add: power_decreasing power_le_one_iff) have f'_descending: "f' y \ f' x" if "x \ y" "0 \ x" for x y unfolding f'_def using that q by (auto intro!: power_mono simp add: ln_powr powr_def mult_nonneg_nonpos) have [simp]: "harm n / ln q <= 0" using harm_nonneg ln_ge_zero_imp_ge_one q by (intro divide_nonneg_neg) auto have f_nn_integral_harm: "- harm n / ln q \ \\<^sup>+ x. (f x) \count_space UNIV" "(\\<^sup>+ i. f (i + 1) \count_space UNIV) \ - harm n / ln q" proof - have "(\\<^sup>+ i. f (i + 1) \count_space UNIV) = (\\<^sup>+x\{0::real..}. (f (nat \x\ + 1))\lborel)" using nn_integral_nats_reals by auto also have "\ = \\<^sup>+x\{0::real..}. ennreal (f' (nat \x\ + 1))\lborel" proof - have "0 \ x \ (1 - q * q ^ nat \x\) ^ n = (1 - q powr (1 + real_of_int \x\)) ^ n" for x::real using q by (subst powr_realpow [symmetric]) (auto simp: powr_add) then show ?thesis unfolding f_def f'_def using q by (auto intro!: nn_integral_cong ennreal_cong simp add: powr_real_of_int indicator_def) qed also have "\ \ set_nn_integral lborel {0..} f'" proof - have "x \ 1 + real_of_int \x\" for x by linarith then show ?thesis by (auto simp add: indicator_def intro!: f'_descending nn_integral_mono ennreal_leI) qed also have harm_integral_f': "\ = - harm n / ln q" unfolding f'_def using q by (auto intro!: ennreal_cong simp add: one_minus_one_minus_q_x_n_nn_integral one_minus_one_minus_q_x_n_integral) finally show "(\\<^sup>+ i. f (i + 1) \count_space UNIV) \ - harm n / ln q" by simp note harm_integral_f'[symmetric] also have "set_nn_integral lborel {0..} f' \ \\<^sup>+x\{0::real..}. f' (nat \x\)\lborel" using assms f'_descending by (auto simp add: indicator_def intro!: nn_integral_mono ennreal_leI) also have "\ = \\<^sup>+x\{0::real..}. f (nat \x\)\lborel" unfolding f_def f'_def using q by (auto intro!: nn_integral_cong ennreal_cong simp add: powr_real_of_int indicator_def) also have "\ = (\\<^sup>+ x. f x \count_space UNIV)" using nn_integral_nats_reals by auto finally show "- harm n / ln q \ \\<^sup>+ x. f x \count_space UNIV" by simp qed then have f1_abs_summable_on: "(\i. f (i + 1)) abs_summable_on UNIV" unfolding f_def using q by (intro nn_integral_finite_imp_abs_sumable_on') (auto simp add: f_def le_less_trans intro!: power_le_one mult_le_one) then have f_abs_summable_on: "f abs_summable_on {1..}" using Suc_le_lessD greaterThan_0 by (subst abs_summable_on_reindex_bij_betw[symmetric, where g="\x. x + 1" and A="UNIV"]) auto also have "(f abs_summable_on {1..}) = ((\x. measure_pmf.prob (H\<^sub>N n) {x..}) abs_summable_on {1..})" proof - have "((\x. measure_pmf.prob (H\<^sub>N n) {x..}) abs_summable_on {1..}) = ((\x. measure_pmf.prob (H\<^sub>N n) {x - 1<..}) abs_summable_on {1..})" by (auto intro!: measure_prob_cong_0 abs_summable_on_cong) also have "\ = (f abs_summable_on {1..})" using assms by (intro abs_summable_on_cong) (auto simp add: f_def prob_Max_IID_geometric_greaterThan) finally show ?thesis by simp qed finally have EH\<^sub>N_sum: "EH\<^sub>N n = (\\<^sub>ai\{1..}. measure_pmf.prob (H\<^sub>N n) {i..})" "integrable (measure_pmf (H\<^sub>N n)) real" unfolding EH\<^sub>N_def using expectation_prob_atLeast by auto then show "integrable (measure_pmf (H\<^sub>N n)) real" by simp have EH\<^sub>N_sum': "EH\<^sub>N n = infsetsum f {1..}" proof - have "EH\<^sub>N n = (\\<^sub>ak\{1..}. measure_pmf.prob (H\<^sub>N n) {k - 1<..})" unfolding EH\<^sub>N_sum by (auto intro!: measure_prob_cong_0 infsetsum_cong) also have "\ = infsetsum f {1..}" using assms by (intro infsetsum_cong) (auto simp add: f_def prob_Max_IID_geometric_greaterThan) finally show ?thesis by simp qed also have "\ = (\\<^sub>ak. f (k + 1))" using Suc_le_lessD greaterThan_0 by (subst infsetsum_reindex_bij_betw[symmetric, where g="\x. x + 1" and A="UNIV"]) auto also have "ennreal \ = (\\<^sup>+x\{0::real..}. f (nat \x\ + 1)\lborel)" using f1_abs_summable_on q by (intro infsetsum_set_nn_integral_reals) (auto simp add: f_def mult_le_one power_le_one) also have "\ = (\\<^sup>+ i. f (i + 1) \count_space UNIV)" using nn_integral_nats_reals by auto also have "\ \ - harm n / ln q" using f_nn_integral_harm by auto finally show "EH\<^sub>N n \ - harm n / ln q" by (subst (asm) ennreal_le_iff) (auto) have "EH\<^sub>N n + 1 = (\\<^sub>ax\{Suc 0..}. f x) + (\\<^sub>ax\{0}. f x)" using assms by (subst EH\<^sub>N_sum') (auto simp add: f_def) also have "\ = infsetsum f UNIV" using f_abs_summable_on by (subst infsetsum_Un_disjoint[symmetric]) (auto intro!: infsetsum_cong) also have "\ = (\\<^sup>+x\{0::real..}. f (nat \x\)\lborel)" proof - have "f abs_summable_on ({0} \ {1..})" using f_abs_summable_on by (intro abs_summable_on_union) (auto) also have "{0::nat} \ {1..} = UNIV" by auto finally show ?thesis using q by (intro infsetsum_set_nn_integral_reals) (auto simp add: f_def mult_le_one power_le_one) qed also have "\ = (\\<^sup>+ x. f x \count_space UNIV)" using nn_integral_nats_reals by auto also have "... \ - harm n / ln q" using f_nn_integral_harm by auto finally have "- harm n / ln q \ EH\<^sub>N n + 1" by (subst (asm) ennreal_le_iff) (auto simp add: EH\<^sub>N_def) then show "- harm n / ln q - 1 \ EH\<^sub>N n" by simp qed theorem EH\<^sub>N_bounds: fixes n::nat assumes "p \ {0<..<1}" shows "- harm n / ln q - 1 \ EH\<^sub>N n" "EH\<^sub>N n \ - harm n / ln q" "integrable (H\<^sub>N n) real" proof - show "- harm n / ln q - 1 \ EH\<^sub>N n" using assms EH\<^sub>N_bounds' by (cases "n = 0") (auto simp add: EH\<^sub>N_def H\<^sub>N_def H_def SL_def harm_expand) show "EH\<^sub>N n \ - harm n / ln q" using assms EH\<^sub>N_bounds' by (cases "n = 0") (auto simp add: EH\<^sub>N_def H\<^sub>N_def H_def SL_def harm_expand) show "integrable (H\<^sub>N n) real" using assms EH\<^sub>N_bounds' by (cases "n = 0") (auto simp add: H\<^sub>N_def H_def SL_def intro!: integrable_measure_pmf_finite) qed end (* context random_skip_list *) subsection \Expected Length of Search Path\ text \ Let @{term "A::'a::linorder set"} and @{term "f::'a \ nat"} where f is an abstract description of a skip list (assign each value its maximum level). steps A f s u l starts on the rightmost element on level s in the skip lists. If possible it moves up, if not it moves to the left. For every step up it adds cost u and for every step to the left it adds cost l. steps A f 0 1 1 therefore walks from the bottom right corner of a skip list to the top left corner of a skip list and counts all steps. \ \ \NOTE: You could also define steps with lsteps and then prove that the following recursive definition holds\ function steps :: "'a :: linorder set \ ('a \ nat) \ nat \ nat \ nat \ nat" where "steps A f l up left = (if A = {} \ infinite A then 0 else (let m = Max A in (if f m < l then steps (A - {m}) f l up left else (if f m > l then up + steps A f (l + 1) up left else left + steps (A - {m}) f l up left))))" by pat_completeness auto termination proof (relation "(\(A,f,l,a,b). card A) <*mlex*> (\(A,f,l,a,b). Max (f ` A) - l) <*mlex*> {}", goal_cases) case 1 then show ?case by(intro wf_mlex wf_empty) next case 2 then show ?case by (intro mlex_less) (auto simp: card_gt_0_iff) next case (3 A f l a b x) then have "Max (f ` A) - Suc l < Max (f ` A) - l" by (meson Max_gr_iff Max_in diff_less_mono2 finite_imageI imageI image_is_empty lessI) with 3 have "((A, f, l + 1, a, b), A, f, l, a, b) \ (\(A, f, l, a, b). Max (f ` A) - l) <*mlex*> {}" by (intro mlex_less) (auto) with 3 show ?case apply - apply(rule mlex_leq) by auto next case 4 then show ?case by (intro mlex_less) (auto simp: card_gt_0_iff) qed declare steps.simps[simp del] text \ lsteps is similar to steps but is using lists instead of sets. This makes the proofs where we use induction easier. \ function lsteps :: "'a list \ ('a \ nat) \ nat \ nat \ nat \ nat" where "lsteps [] f l up left = 0" | "lsteps (x#xs) f l up left = (if f x < l then lsteps xs f l up left else (if f x > l then up + lsteps (x#xs) f (l + 1) up left else left + lsteps xs f l up left))" by pat_completeness auto termination proof (relation "(\(xs,f,l,a,b). length xs) <*mlex*> (\(xs,f,l,a,b). Max (f ` set xs) - l) <*mlex*> {}", goal_cases) case 1 then show ?case by(intro wf_mlex wf_empty) next case 2 then show ?case by (auto intro: mlex_less simp: card_gt_0_iff) next case (3 n f l a b) show ?case by (rule mlex_leq) (use 3 in \auto intro: mlex_less mlex_leq intro!: diff_less_mono2 simp add: Max_gr_iff\) next case 4 then show ?case by (intro mlex_less) (auto simp: card_gt_0_iff) qed declare lsteps.simps(2)[simp del] lemma steps_empty [simp]: "steps {} f l up left = 0" by (simp add: steps.simps) lemma steps_lsteps: "steps A f l u v = lsteps (rev (sorted_list_of_set A)) f l u v" proof (cases "finite A \ A \ {}") case True then show ?thesis proof(induction "(rev (sorted_list_of_set A))" f l u v arbitrary: A rule: lsteps.induct) case (2 y ys f l u v A) then have y_ys: "y = Max A" "ys = rev (sorted_list_of_set (A - {y}))" by (auto simp add: sorted_list_of_set_Max_snoc) consider (a) "l < f y" | (b) "f y < l" | (c) "f y = l" by fastforce then have "steps A f l u v = lsteps (y#ys) f l u v" proof cases case a then show ?thesis by (subst steps.simps, subst lsteps.simps) (use y_ys 2 in auto) next case b then show ?thesis using y_ys 2(1) by (cases "ys = []") (auto simp add: steps.simps lsteps.simps) next case c then have "steps (A - {Max A}) f l u v = lsteps (rev (sorted_list_of_set (A - {Max A}))) f l u v" by (cases "A = {Max A}") (use y_ys 2 in \auto intro!: 2(3) simp add: steps.simps\) then show ?thesis by (subst steps.simps, subst lsteps.simps) (use y_ys 2 in auto) qed then show ?case using 2 by simp qed (auto simp add: steps.simps) qed (auto simp add: steps.simps) lemma lsteps_comp_map: "lsteps zs (f \ g) l u v = lsteps (map g zs) f l u v" by (induction zs "f \ g" l u v rule: lsteps.induct) (auto simp add: lsteps.simps) lemma steps_image: - assumes "finite A" "mono_on g A" "inj_on g A" + assumes "finite A" "mono_on A g" "inj_on g A" shows "steps A (f \ g) l u v = steps (g ` A) f l u v" proof - have "(sorted_list_of_set (g ` A)) = map g (sorted_list_of_set A)" using sorted_list_of_set_image assms by auto also have "rev \ = map g (rev (sorted_list_of_set A))" using rev_map by auto finally show ?thesis by (simp add: steps_lsteps lsteps_comp_map) qed lemma lsteps_cong: assumes "ys = xs" "\x. x \ set xs \ f x = g x" "l = l'" shows "lsteps xs f l u v = lsteps ys g l' u v" using assms proof (induction xs f l u v arbitrary: ys l' rule: lsteps.induct) case (2 x xs f l up left) then show ?case by (subst \ys = x # xs\, subst lsteps.simps, subst (2) lsteps.simps) auto qed (auto) lemma steps_cong: assumes "A = B" "\x. x \ A \ f x = g x" "l = l'" shows "steps A f l u v = steps B g l' u v" using assms by (cases "A = {} \ infinite A") (auto simp add: steps_lsteps steps.simps intro!: lsteps_cong) lemma lsteps_f_add': shows "lsteps xs f l u v = lsteps xs (\x. f x + m) (l + m) u v" by (induction xs f l u v rule: lsteps.induct) (auto simp add: lsteps.simps) lemma steps_f_add': shows "steps A f l u v = steps A (\x. f x + m) (l + m) u v" by (cases "A = {} \ infinite A") (auto simp add: steps_lsteps steps.simps intro!: lsteps_f_add') lemma lsteps_smaller_set: assumes "m \ l" shows "lsteps xs f l u v = lsteps [x \ xs. m \ f x] f l u v" using assms by (induction xs f l u v rule: lsteps.induct) (auto simp add: lsteps.simps) lemma steps_smaller_set: assumes "finite A" "m \ l" shows "steps A f l u v = steps {x\A. f x \ m} f l u v" using assms by(cases "A = {} \ infinite A") (auto simp add: steps_lsteps steps.simps rev_filter sorted_list_of_set_filter intro!: lsteps_smaller_set) lemma lsteps_level_greater_fun_image: assumes "\x. x \ set xs \ f x < l" shows "lsteps xs f l u v = 0" using assms by (induction xs f l u v rule: lsteps.induct) (auto simp add: lsteps.simps) lemma lsteps_smaller_card_Max_fun': assumes "\x \ set xs. l \ f x" shows "lsteps xs f l u v + l * u \ v * length xs + u * Max ((f ` (set xs)) \ {0})" using assms proof (induction xs f l u v rule: lsteps.induct) case (1 f l up left) then show ?case by (simp) next case (2 x xs f l up left) consider "l = f x" "\y\set xs. l \ f y" | "f x = l" "\ (\y\set xs. l \ f y)" | "f x < l" | "l < f x" by fastforce then show ?case proof cases assume a: "l = f x" "\y\set xs. l \ f y" have "lsteps (x # xs) f l up left + l * up = lsteps xs f l up left + f x * up + left" using a by (auto simp add: lsteps.simps) also have "lsteps xs f l up left + f x * up \ left * length xs + up * Max (f ` set xs \ {0})" using a 2 by blast also have "up * Max (f ` set xs \ {0}) \ up * Max (insert (f x) (f ` set xs))" by simp finally show ?case by auto next assume a: "f x = l" "\ (\y\set xs. l \ f y)" have "lsteps (x # xs) f l up left + l * up = lsteps xs f l up left + f x * up + left" using a by (auto simp add: lsteps.simps) also have "lsteps xs f l up left = 0" using a by (subst lsteps_level_greater_fun_image) auto also have "f x * up \ up * Max (insert (f x) (f ` set xs))" by simp finally show ?case by simp next assume a: "f x < l" then have "lsteps (x # xs) f l up left = lsteps xs f l up left" by (auto simp add: lsteps.simps) also have "\ + l * up \ left * length (x # xs) + up * Max (insert 0 (f ` set xs))" using a 2 by auto also have "Max (insert 0 (f ` set xs)) \ Max (f ` set (x # xs) \ {0})" by simp finally show ?case by simp next assume "f x > l" then show ?case using 2 by (subst lsteps.simps) auto qed qed lemma steps_smaller_card_Max_fun': assumes "finite A" "\x\A. l \ f x" shows "steps A f l up left + l * up \ left * card A + up * Max\<^sub>0 (f ` A)" proof - let ?xs = "rev (sorted_list_of_set A)" have "steps A f l up left = lsteps (rev (sorted_list_of_set A)) f l up left" using steps_lsteps by blast also have "\ + l * up \ left * length ?xs + up * Max (f ` set ?xs \ {0})" using assms by (intro lsteps_smaller_card_Max_fun') auto also have "left * length ?xs = left * card A" using assms sorted_list_of_set_length by (auto) also have "set ?xs = A" using assms by (auto) finally show ?thesis by simp qed lemma lsteps_height: assumes "\x \ set xs. l \ f x" shows "lsteps xs f l up 0 + up * l = up * Max\<^sub>0 (f ` (set xs))" using assms proof (induction xs f l up "0::nat" rule: lsteps.induct) case (2 x xs f l up) consider "l = f x" "\y\set xs. l \ f y" | "f x = l" "\ (\y\set xs. l \ f y)" | "f x < l" | "l < f x" by fastforce then show ?case proof cases assume 0: "l = f x" "\y\set xs. l \ f y" then have 1: "set xs \ {}" using 2 by auto then have "\xa\set xs. f x \ f xa" using 0 2 by force then have "f x \ Max (f ` set xs)" using 0 2 by (subst Max_ge_iff) auto then have "max (f x) (Max (f ` set xs)) = (Max (f ` set xs))" using 0 2 by (auto intro!: simp add: max_def) then show ?case using 0 1 2 by (subst lsteps.simps) (auto) next assume 0: "f x = l" "\ (\y\set xs. l \ f y)" then have "Max (insert l (f ` set xs)) = l" by (intro Max_eqI) (auto) moreover have "lsteps xs f l up 0 = 0" using 0 by (subst lsteps_level_greater_fun_image) auto ultimately show ?case using 0 by (subst lsteps.simps) auto next assume 0: "f x < l" then have 1: "set xs \ {}" using 2 by auto then have "\xa\set xs. f x \ f xa" using 0 2 by force then have " f x \ Max (f ` set xs)" using 0 2 by (subst Max_ge_iff) auto then have "max (f x) (Max (f ` set xs)) = Max (f ` set xs)" using 0 2 by (auto intro!: simp add: max_def) then show ?case using 0 1 2 by (subst lsteps.simps) (auto) next assume "f x > l" then show ?case using 2 by (subst lsteps.simps) auto qed qed (simp) lemma steps_height: assumes "finite A" shows "steps A f 0 up 0 = up * Max\<^sub>0 (f ` A)" proof - have "steps A f 0 up 0 = lsteps (rev (sorted_list_of_set A)) f 0 up 0 + up * 0" by (subst steps_lsteps) simp also have "\ = up * Max (f ` A \ {0})" if "A \ {}" using assms that by (subst lsteps_height) auto finally show ?thesis using assms by (cases "A = {}") (auto) qed context random_skip_list begin text \ We can now define the pmf describing the length of the search path in a skip list. Like the height it only depends on the number of elements in the skip list's underlying set. \ definition R where "R A u l = map_pmf (\f. steps A f 0 u l) (SL A)" definition R\<^sub>N :: "nat \ nat \ nat \ nat pmf" where "R\<^sub>N n u l = R {..N_alt_def: "R\<^sub>N n u l = map_pmf (\f. steps {..N n)" unfolding SL\<^sub>N_def R\<^sub>N_def R_def by simp context includes monad_normalisation begin lemma R_R\<^sub>N: assumes "finite A" "p \ {0..1}" shows "R A u l = R\<^sub>N (card A) u l" proof - let ?steps = "\A f. steps A f 0 u l" let ?f' = "bij_mono_map_set_to_nat A" have "R A u l = SL A \ (\f. return_pmf (?steps A f))" unfolding R_def map_pmf_def by simp also have "\ = SL\<^sub>N (card A) \ (\f. return_pmf (?steps A (f \ ?f')))" proof - have "?f' x \ {.. A" for x using that unfolding bij_mono_map_set_to_nat_def by (auto) then show ?thesis using assms bij_mono_map_set_to_nat unfolding SL_def SL\<^sub>N_def by (subst Pi_pmf_bij_betw[of _ ?f' "{.. = SL\<^sub>N (card A) \ (\f. return_pmf (?steps {..N_def R_def SL\<^sub>N_def SL_def by (simp add: map_pmf_def) qed text \ @{const R\<^sub>N} fulfills a recurrence relation. If we move up or to the left the ``remaining'' length of the search path is again a slightly different probability distribution over the length. \ lemma R\<^sub>N_recurrence: assumes "0 < n" "p \ {0<..1}" shows "R\<^sub>N n u l = do { b \ bernoulli_pmf p; if b then \ \leftwards\ map_pmf (\n. n + l) (R\<^sub>N (n - 1) u l) else do { \ \upwards\ m \ binomial_pmf (n - 1) (1 - p); map_pmf (\n. n + u) (R\<^sub>N (m + 1) u l) } }" proof - define B where "B = (\b. insert (n-1) {x \ {.. b x})" have "R\<^sub>N n u l = map_pmf (\f. steps {..N n)" by (auto simp add: R\<^sub>N_def R_def SL\<^sub>N_def) also have "\ = map_pmf (\f. steps {..(y, f). f(n-1 := y)) (pair_pmf (geometric_pmf p) (SL\<^sub>N (n - 1))))" proof - have "{.._. geometric_pmf p)) = map_pmf (\(y, f). f(n - 1 := y)) (pair_pmf (geometric_pmf p) (Pi_pmf {.._. geometric_pmf p)))" using assms by (subst Pi_pmf_insert[of "{.._. geometric_pmf p", symmetric]) (auto) then show ?thesis by (simp add: SL\<^sub>N_def SL_def) qed also have "\ = do { g \ geometric_pmf p; f \ SL\<^sub>N (n - 1); return_pmf (steps {.. = do { b \ bernoulli_pmf p; g \ if b then return_pmf 0 else map_pmf Suc (geometric_pmf p); f \ SL\<^sub>N (n - 1); return_pmf (steps {.. = do { b \ bernoulli_pmf p; if b then do { g \ return_pmf 0; f \ SL\<^sub>N (n - 1); return_pmf (steps {.. map_pmf Suc (geometric_pmf p); f \ SL\<^sub>N (n - 1); return_pmf (steps {.. return_pmf 0; f \ SL\<^sub>N (n - 1); return_pmf (steps {.. SL\<^sub>N (n - 1); return_pmf (steps {.. = map_pmf (\n. n + l) (map_pmf (\f. steps {..N (n - 1)))" proof - have I: "{.. = map_pmf (\n. n + l) (R\<^sub>N (n - 1) u l)" unfolding R\<^sub>N_def R_def SL\<^sub>N_def by simp also have "map_pmf Suc (geometric_pmf p) \ (\g. SL\<^sub>N (n - 1) \ (\f. return_pmf (steps {.._. bernoulli_pmf p) \ (\b. map_pmf Suc (geometric_pmf p) \ (\g. Pi_pmf {x \ {.. b x} 0 (\_. map_pmf Suc (geometric_pmf p)) \ (\f. return_pmf (steps {..N_def SL_def by (subst Pi_pmf_geometric_filter) (auto) also have "\ = do { b \ Pi_pmf {.._. bernoulli_pmf p); f \ Pi_pmf (insert (n-1) {x \ {.. b x}) 0 (\_. map_pmf Suc (geometric_pmf p)); return_pmf (steps {.. = do { b \ Pi_pmf {.._. bernoulli_pmf p); f \ Pi_pmf (B b) 1 (\_. map_pmf Suc (geometric_pmf p)); return_pmf (steps {..x. if x \ (B b) then f x else 0) 0 u l)}" by (subst Pi_pmf_default_swap[symmetric, of _ _ _ 1]) (auto simp add: map_pmf_def B_def) also have "\ = do { b \ Pi_pmf {.._. bernoulli_pmf p); f \ SL (B b); return_pmf (steps {..x. if x \ (B b) then Suc (f x) else 0) 0 u l)}" proof - have *: "(Suc \ f) x = Suc (f x)" for x and f::"nat \ nat" by simp have "(\f. return_pmf (steps {..x. if x \ B b then (Suc \ f) x else 0) 0 u l)) = (\f. return_pmf (steps {..x. if x \ B b then Suc (f x) else 0) 0 u l))" for b by (subst *) (simp) then show ?thesis by (subst Pi_pmf_map[of _ _ 0]) (auto simp add: map_pmf_def B_def SL_def) qed also have "\ = do { b \ Pi_pmf {.._. bernoulli_pmf p); r \ R (B b) u l; return_pmf (u + r)}" proof - have "steps {..x. if x \ B b then Suc (f x) else 0) 0 u l = u + steps (B b) f 0 u l" for f b proof - have "Max {..x. if x \ B b then Suc (f x) else 0) 0 u l = u + (steps {..x. if x \ (B b) then Suc (f x) else 0) 1 u l)" unfolding B_def using assms by (subst steps.simps) (auto simp add: Let_def) also have "steps {..x. if x \ (B b) then Suc (f x) else 0) 1 u l = steps (B b) (\x. if x \ (B b) then Suc (f x) else 0) 1 u l" proof - have "{x \ {.. (if x \ B b then Suc (f x) else 0)} = B b" using assms unfolding B_def by force then show ?thesis by (subst steps_smaller_set[of _ 1]) auto qed also have "\ = steps (B b) (\x. f x + 1) 1 u l" by (rule steps_cong) (auto) also have "\ = steps (B b) f 0 u l" by (subst (2) steps_f_add'[of _ _ _ _ _ 1]) simp finally show ?thesis by auto qed then show ?thesis by (simp add: R_def map_pmf_def) qed also have "\ = do { b \ Pi_pmf {.._. bernoulli_pmf (1 - p)); let m = 1 + card {x. x < n - 1 \ b x}; r \ R {.. b x}) = (Suc (card {x. x < n - 1 \ b x}))" for b using assms by (auto simp add: card_insert_if) have "Pi_pmf {.._. bernoulli_pmf p) = Pi_pmf {.._. map_pmf Not (bernoulli_pmf (1 - p)))" using assms by (subst bernoulli_pmf_Not) auto also have "\ = map_pmf ((\) Not) (Pi_pmf {.._. bernoulli_pmf (1 - p)))" using assms by (subst Pi_pmf_map[of _ _ False]) auto finally show ?thesis unfolding B_def using assms * by (subst R_R\<^sub>N) (auto simp add: R_R\<^sub>N map_pmf_def) qed also have "\ = binomial_pmf (n - 1) (1 - p) \ (\m. map_pmf (\n. n + u) (R\<^sub>N (m + 1) u l))" using assms by (subst binomial_pmf_altdef'[where A = "{..N_def R_def SL_def map_pmf_def ac_simps) finally show ?thesis by simp qed end (* context includes monad_normalisation *) text \ The expected height and length of search path defined as non-negative integral. It's easier to prove the recurrence relation of the expected length of the search path using non-negative integrals. \ definition NH\<^sub>N where "NH\<^sub>N n = nn_integral (H\<^sub>N n) real" definition NR\<^sub>N where "NR\<^sub>N n u l = nn_integral (R\<^sub>N n u l) real" lemma NH\<^sub>N_EH\<^sub>N: assumes "p \ {0<..<1}" shows "NH\<^sub>N n = EH\<^sub>N n" using assms EH\<^sub>N_bounds unfolding EH\<^sub>N_def NH\<^sub>N_def by (subst nn_integral_eq_integral) (auto) lemma R\<^sub>N_0 [simp]: "R\<^sub>N 0 u l = return_pmf 0" unfolding R\<^sub>N_def R_def SL_def by (auto simp add: steps.simps) lemma NR\<^sub>N_bounds: fixes u l::nat shows "NR\<^sub>N n u l \ l * n + u * NH\<^sub>N n" proof - have "NR\<^sub>N n u l = \\<^sup>+ x. x \measure_pmf (R\<^sub>N n u l)" unfolding NR\<^sub>N_def R\<^sub>N_alt_def by (simp add: ennreal_of_nat_eq_real_of_nat) also have "\ \ \\<^sup>+ x. x \(measure_pmf (map_pmf (\f. l * n + u * Max\<^sub>0 (f ` {..N n)))" using of_nat_mono[OF steps_smaller_card_Max_fun'[of "{..N_alt_def by (cases "n = 0") (auto intro!: nn_integral_mono) also have "\ = l * n + u * NH\<^sub>N n" unfolding NH\<^sub>N_def H\<^sub>N_def H_def SL\<^sub>N_def by (auto simp add: nn_integral_add nn_integral_cmult ennreal_of_nat_eq_real_of_nat ennreal_mult) finally show "NR\<^sub>N n u l \ l * n + u * NH\<^sub>N n" by simp qed lemma NR\<^sub>N_recurrence: assumes "0 < n" "p \ {0<..<1}" shows "NR\<^sub>N n u l = (p * (l + NR\<^sub>N (n - 1) u l) + q * (u + (\kN (k + 1) u l * (pmf (binomial_pmf (n - 1) q) k)))) / (1 - (q ^ n))" proof - define B where "B = (\n k. pmf (binomial_pmf n q) k)" have q: "q \ {0<..<1}" using assms unfolding q_def by auto then have "q ^ n < 1" using assms power_Suc_less_one by (induction n) (auto) then have qn: "q ^ n \ {0<..<1}" using assms q by (auto) have "NR\<^sub>N n u l = p * (l + NR\<^sub>N (n - 1) u l) + q * (u + \\<^sup>+ k. NR\<^sub>N (k + 1) u l \measure_pmf (binomial_pmf (n - 1) q))" using assms unfolding NR\<^sub>N_def by(subst R\<^sub>N_recurrence) (auto simp add: field_simps nn_integral_add q_def ennreal_of_nat_eq_real_of_nat) also have "(\\<^sup>+ m. NR\<^sub>N (m + 1) u l \measure_pmf (binomial_pmf (n - 1) q)) = (\k\n - 1. NR\<^sub>N (k + 1) u l * B (n - 1) k)" using assms unfolding B_def q_def by (auto simp add: nn_integral_measure_pmf_finite) also have "\ = (\k\{.. {n - 1}. NR\<^sub>N (k + 1) u l * B (n - 1) k)" by (rule sum.cong) (auto) also have "\ = (\kN (k + 1) u l * B (n - 1) k) + NR\<^sub>N n u l * q ^ (n - 1)" unfolding B_def q_def using assms by (subst sum.union_disjoint) (auto) finally have "NR\<^sub>N n u l = p * (l + NR\<^sub>N (n - 1) u l) + q * ((\kN (k + 1) u l * B (n - 1) k) + u) + NR\<^sub>N n u l * (q ^ (n - 1)) * q" using assms by (auto simp add: field_simps numerals) also have "NR\<^sub>N n u l * (q ^ (n - 1)) * q = (q ^ n) * NR\<^sub>N n u l" using q power_minus_mult[of _ q] assms by (subst mult_ac, subst ennreal_mult[symmetric], auto simp add: mult_ac) finally have 1: "NR\<^sub>N n u l = p * (l + NR\<^sub>N (n - 1) u l) + q * (u + (\kN (k + 1) u l * (B (n - 1) k))) + (q ^ n) * NR\<^sub>N n u l " by (simp add: add_ac) have "x - z = y" if "x = y + z" "z \ \" for x y z::ennreal using that by (subst that) (auto) have "NR\<^sub>N n u l \ l * n + u * NH\<^sub>N n" using NR\<^sub>N_bounds by (auto simp add: ennreal_of_nat_eq_real_of_nat) also have "NH\<^sub>N n = EH\<^sub>N n" using assms NH\<^sub>N_EH\<^sub>N by auto also have "(l * n) + u * ennreal (EH\<^sub>N n) < \" by (simp add: ennreal_mult_less_top of_nat_less_top) finally have 3: "NR\<^sub>N n u l \ \" by simp have 2: "x = y / (1 - a)" if "x = y + a * x" and t: "x \ \" "a \ {0<..<1}" for x y::ennreal and a::real proof - have "y = x - a * x" using t by (subst that) (auto simp add: ennreal_mult_eq_top_iff) also have "\ = x * (ennreal 1 - ennreal a)" using that by (auto simp add: mult_ac ennreal_right_diff_distrib) also have "ennreal 1 - ennreal a = ennreal (1 - a)" using that by (subst ennreal_minus) (auto) also have "x * (1 - a) / (1 - a) = x" using that ennreal_minus_eq_0 not_less by (subst mult_divide_eq_ennreal) auto finally show ?thesis by simp qed have "NR\<^sub>N n u l = (p * (l + NR\<^sub>N (n - 1) u l) + q * (u + (\kN (k + 1) u l * (B (n - 1) k)))) / (1 - (q ^ n))" using 1 3 assms qn by (intro 2) auto then show ?thesis unfolding B_def by simp qed lemma NR\<^sub>n_NH\<^sub>N: "NR\<^sub>N n u 0 = u * NH\<^sub>N n" proof - have "NR\<^sub>N n u 0 = \\<^sup>+ f. steps {..measure_pmf (SL\<^sub>N n)" unfolding NR\<^sub>N_def R\<^sub>N_alt_def by (auto simp add: ennreal_of_nat_eq_real_of_nat) also have "\ = \\<^sup>+ f. of_nat u * of_nat (Max\<^sub>0 (f ` {..measure_pmf (SL\<^sub>N n)" by (intro nn_integral_cong) (auto simp add: steps_height) also have "\ = u * NH\<^sub>N n" by (auto simp add: NH\<^sub>N_def H\<^sub>N_def H_def SL\<^sub>N_def ennreal_of_nat_eq_real_of_nat nn_integral_cmult) finally show ?thesis by simp qed lemma NR\<^sub>N_recurrence': assumes "0 < n" "p \ {0<..<1}" shows "NR\<^sub>N n u l = (p * l + p * NR\<^sub>N (n - 1) u l + q * u + q * (\kN (k + 1) u l * (pmf (binomial_pmf (n - 1) q) k))) / (1 - (q ^ n))" unfolding NR\<^sub>N_recurrence[OF assms] by (auto simp add: field_simps ennreal_of_nat_eq_real_of_nat ennreal_mult' ennreal_mult'') lemma NR\<^sub>N_l_0: assumes "0 < n" "p \ {0<..<1}" shows "NR\<^sub>N n u 0 = (p * NR\<^sub>N (n - 1) u 0 + q * (u + (\kN (k + 1) u 0 * (pmf (binomial_pmf (n - 1) q) k)))) / (1 - (q ^ n))" unfolding NR\<^sub>N_recurrence[OF assms] by (simp) lemma NR\<^sub>N_u_0: assumes "0 < n" "p \ {0<..<1}" shows "NR\<^sub>N n 0 l = (p * (l + NR\<^sub>N (n - 1) 0 l) + q * (\kN (k + 1) 0 l * (pmf (binomial_pmf (n - 1) q) k))) / (1 - (q ^ n))" unfolding NR\<^sub>N_recurrence[OF assms] by (simp) lemma NR\<^sub>N_0[simp]: "NR\<^sub>N 0 u l = 0" unfolding NR\<^sub>N_def R\<^sub>N_def R_def by (auto) lemma NR\<^sub>N_1: assumes "p \ {0<..<1}" shows "NR\<^sub>N 1 u l = (u * q + l * p) / p" proof - have "NR\<^sub>N 1 u l = (ennreal p * of_nat l + ennreal q * of_nat u) / ennreal (1 - q)" using assms by (subst NR\<^sub>N_recurrence) auto also have "(ennreal p * of_nat l + ennreal q * of_nat u) = (u * q + l * p)" using assms q_def by (subst ennreal_plus) (auto simp add: field_simps ennreal_mult' ennreal_of_nat_eq_real_of_nat) also have "\ / ennreal (1 - q) = ennreal ((u * q + l * p) / (1 - q))" using q_def assms by (intro divide_ennreal) auto finally show ?thesis unfolding q_def by simp qed lemma NR\<^sub>N_NR\<^sub>N_l_0: assumes n: "0 < n" and p: "p \ {0<..<1}" and "u \ 1" shows "NR\<^sub>N n u 0 = (u * q / (u * q + l * p)) * NR\<^sub>N n u l" using n proof (induction n rule: less_induct) case (less i) have 1: "0 < u * q" unfolding q_def using assms by simp moreover have "0 \ l * p" using assms by auto ultimately have 2: "0 < u * q + l * p" by arith define c where "c = ennreal (u * q / (u * q + l * p))" have [simp]: "c / c = 1" proof - have "u * q / (u * q + l * p) \ 0" using assms q_def 2 by auto then show ?thesis unfolding c_def using p q_def by (auto intro!: ennreal_divide_self) qed show ?case proof (cases "i = 1") case True have "c * NR\<^sub>N i u l = c * ((u * q + l * p) / p)" unfolding c_def True by (subst NR\<^sub>N_1[OF p]) auto also have "\ = ennreal ((u * q / (u * q + l * p)) * ((u * q + l * p) / p))" unfolding c_def using assms q_def by (subst ennreal_mult'') auto also have "(u * q / (u * q + l * p)) * ((u * q + l * p) / p) = u * q / p" proof - have I: "(a / b) * (b / c) = a / c" if "0 < b" for a b c::"real" using that by (auto) show ?thesis using 2 q_def by (intro I) auto qed also have "\ = NR\<^sub>N i u 0" unfolding True c_def by (subst NR\<^sub>N_1[OF p]) (auto) finally show ?thesis unfolding c_def using True by simp next case False then have i: "i > 1" using less by auto define c where "c = ennreal (u * q / (u * q + l * p))" define B where "B = (\kN (k + 1) u l * ennreal (pmf (binomial_pmf (i - 1) q) k))" have "NR\<^sub>N i u 0 = (p * NR\<^sub>N (i - 1) u 0 + q * (u + (\kN (k + 1) u 0 * (pmf (binomial_pmf (i - 1) q) k)))) / (1 - (q ^ i))" using less assms by (subst NR\<^sub>N_l_0) auto also have "q * (u + (\kN (k + 1) u 0 * (pmf (binomial_pmf (i - 1) q) k))) = q * u + q * (\kN (k + 1) u 0 * (pmf (binomial_pmf (i - 1) q) k))" using assms q_def by (auto simp add: field_simps ennreal_of_nat_eq_real_of_nat ennreal_mult) also have "NR\<^sub>N (i - 1) u 0 = c * NR\<^sub>N (i - 1) u l" unfolding c_def using less i by (intro less) (auto) also have "(\kN (k + 1) u 0 * ennreal (pmf (binomial_pmf (i - 1) q) k)) = (\kN (k + 1) u l * ennreal (pmf (binomial_pmf (i - 1) q) k))" by (auto intro!: sum.cong simp add: less c_def) also have "\ = c * B" unfolding B_def by (subst sum_distrib_left) (auto intro!: sum.cong mult_ac) also have "q * (c * B) = c * (q * B)" by (simp add: mult_ac) also have "ennreal (q * real u) = q * u * ((u * q + l * p) / (u * q + l * p))" using assms 2 by (auto simp add: field_simps q_def) also have "\ = c * (real u * q + real l * p)" unfolding c_def using 2 by (subst ennreal_mult''[symmetric]) (auto simp add: mult_ac) also have "c * ennreal (real u * q + real l * p) + c * (ennreal q * B) = c * (ennreal (real u * q + real l * p) + (ennreal q * B))" by (auto simp add: field_simps) also have "ennreal p * (c * NR\<^sub>N (i - 1) u l) = c * (ennreal p * NR\<^sub>N (i - 1) u l)" by (simp add: mult_ac) also have "(c * (ennreal p * NR\<^sub>N (i - 1) u l) + c * (ennreal (u * q + l * p) + ennreal q * B)) = c * ((ennreal p * NR\<^sub>N (i - 1) u l) + (ennreal (u * q + l * p) + ennreal q * B))" by (auto simp add: field_simps) also have " c * (ennreal p * NR\<^sub>N (i - 1) u l + (ennreal (u * q + l * p) + ennreal q * B)) / ennreal (1 - q ^ i) = c * ((ennreal p * NR\<^sub>N (i - 1) u l + (ennreal (u * q + l * p) + ennreal q * B)) / ennreal (1 - q ^ i))" by (auto simp add: ennreal_times_divide) also have "(ennreal p * NR\<^sub>N (i - 1) u l + (ennreal (real u * q + real l * p) + ennreal q * B)) / ennreal (1 - q ^ i) = NR\<^sub>N i u l" apply(subst (2) NR\<^sub>N_recurrence') using i assms q_def by (auto simp add: field_simps B_def ennreal_of_nat_eq_real_of_nat ennreal_mult' ennreal_mult'') finally show ?thesis unfolding c_def by simp qed qed text \ Assigning 1 as the cost for going up and/or left, we can now show the relation between the expected length of the reverse search path and the expected height. \ definition EL\<^sub>N where "EL\<^sub>N n = measure_pmf.expectation (R\<^sub>N n 1 1) real" theorem EH\<^sub>N_EL\<^sub>s\<^sub>p: assumes "p \ {0<..<1}" shows "1 / q * EH\<^sub>N n = EL\<^sub>N n" proof - have 1: "ennreal (1 / y * x) = r" if "ennreal x = y * r" "x \ 0" "y > 0" for x y::real and r::ennreal proof - have "ennreal ((1 / y) * x) = ennreal (1 / y) * ennreal x" using that apply(subst ennreal_mult'') by auto also note that(1) also have "ennreal (1 / y) * (ennreal y * r) = ennreal ((1 / y) * y) * r" using that by (subst ennreal_mult'') (auto simp add: mult_ac) also have "(1 / y) * y = 1" using that by (auto) finally show ?thesis by auto qed have "EH\<^sub>N n = NH\<^sub>N n" using NH\<^sub>N_EH\<^sub>N assms by auto also have "NH\<^sub>N n = NR\<^sub>N n 1 0" using NR\<^sub>n_NH\<^sub>N by auto also have "NR\<^sub>N n 1 0 = q * NR\<^sub>N n 1 1" if "n > 0" using NR\<^sub>N_NR\<^sub>N_l_0[of _ 1 1] that assms q_def by force finally have "ennreal (EH\<^sub>N n) = q * NR\<^sub>N n 1 1" if "n > 0" using that by blast then have "1 / q * EH\<^sub>N n = NR\<^sub>N n 1 1" if "n > 0" using that assms q_def by (intro 1) (auto simp add: EH\<^sub>N_def H\<^sub>N_def H_def) moreover have "1 / q * EH\<^sub>N n = NR\<^sub>N n 1 1" if "n = 0" unfolding that by (auto simp add: EH\<^sub>N_def H\<^sub>N_def H_def) ultimately have 2: "ennreal (1 / q * EH\<^sub>N n) = NR\<^sub>N n 1 1" by blast also have "NR\<^sub>N n 1 1 = EL\<^sub>N n" using 2 assms EH\<^sub>N_bounds unfolding EL\<^sub>N_def NR\<^sub>N_def by(subst nn_integral_eq_integral) (auto intro!: integrableI_nn_integral_finite[where x="EH\<^sub>N n / q"]) finally show ?thesis using assms q_def ennreal_inj unfolding EL\<^sub>N_def EH\<^sub>N_def H\<^sub>N_def H_def SL_def by (auto) qed end (* context random_skip_list *) thm random_skip_list.EH\<^sub>N_EL\<^sub>s\<^sub>p[unfolded random_skip_list.q_def] random_skip_list.EH\<^sub>N_bounds'[unfolded random_skip_list.q_def] end diff --git a/thys/Smith_Normal_Form/Cauchy_Binet.thy b/thys/Smith_Normal_Form/Cauchy_Binet.thy --- a/thys/Smith_Normal_Form/Cauchy_Binet.thy +++ b/thys/Smith_Normal_Form/Cauchy_Binet.thy @@ -1,1423 +1,1423 @@ (* Author: Jose Divasón Email: jose.divason@unirioja.es *) section \The Cauchy--Binet formula\ theory Cauchy_Binet imports Diagonal_To_Smith SNF_Missing_Lemmas begin subsection \Previous missing results about @{text "pick"} and @{text "insert"}\ lemma pick_insert: assumes a_notin_I: "a \ I" and i2: "i < card I" and a_def: "pick (insert a I) a' = a" (*Alternative: index (insort a (sorted_list_of_set I)) a = a'*) and ia': "i < a'" (*Case 1*) and a'_card: "a' < card I + 1" shows "pick (insert a I) i = pick I i" proof - have finI: "finite I" using i2 using card.infinite by force have "pick (insert a I) i = sorted_list_of_set (insert a I) ! i" proof (rule sorted_list_of_set_eq_pick[symmetric]) have "finite (insert a I)" using card.infinite i2 by force thus "i < length (sorted_list_of_set (insert a I))" by (metis a_notin_I card_insert_disjoint distinct_card finite_insert i2 less_Suc_eq sorted_list_of_set(1) sorted_list_of_set(3)) qed also have "... = insort a (sorted_list_of_set I) ! i" by (simp add: a_notin_I finI) also have "... = (sorted_list_of_set I) ! i" proof (rule insort_nth[OF]) show "sorted (sorted_list_of_set I)" by auto show "a \ set (sorted_list_of_set I)" using a_notin_I by (metis card.infinite i2 not_less_zero set_sorted_list_of_set) have "index (sorted_list_of_set (insert a I)) a = a'" using pick_index a_def using a'_card a_notin_I finI by auto hence "index (insort a (sorted_list_of_set I)) a = a'" by (simp add: a_notin_I finI) thus "i < index (insort a (sorted_list_of_set I)) a" using ia' by auto show "sorted_list_of_set I \ []" using finI i2 by fastforce qed also have "... = pick I i" proof (rule sorted_list_of_set_eq_pick) have "finite I" using card.infinite i2 by fastforce thus "i < length (sorted_list_of_set I)" by (metis distinct_card distinct_sorted_list_of_set i2 set_sorted_list_of_set) qed finally show ?thesis . qed lemma pick_insert2: assumes a_notin_I: "a \ I" and i2: "i < card I" and a_def: "pick (insert a I) a' = a" (*Alternative: index (sorted_list_of_set (insert a I)) a = a'*) and ia': "i \ a'" (*Case 2*) and a'_card: "a' < card I + 1" shows "pick (insert a I) i < pick I i" proof (cases "i = 0") case True then show ?thesis by (auto, metis (mono_tags, lifting) DL_Missing_Sublist.pick.simps(1) Least_le a_def a_notin_I dual_order.order_iff_strict i2 ia' insertCI le_zero_eq not_less_Least pick_in_set_le) next case False hence i0: "i = Suc (i - 1)" using a'_card ia' by auto have finI: "finite I" using i2 card.infinite by force have index_a'1: "index (sorted_list_of_set (insert a I)) a = a'" using pick_index using a'_card a_def a_notin_I finI by auto hence index_a': "index (insort a (sorted_list_of_set I)) a = a'" by (simp add: a_notin_I finI) have i1_length: "i - 1 < length (sorted_list_of_set I)" using i2 by (metis distinct_card distinct_sorted_list_of_set finI less_imp_diff_less set_sorted_list_of_set) have 1: "pick (insert a I) i = sorted_list_of_set (insert a I) ! i" proof (rule sorted_list_of_set_eq_pick[symmetric]) have "finite (insert a I)" using card.infinite i2 by force thus "i < length (sorted_list_of_set (insert a I))" by (metis a_notin_I card_insert_disjoint distinct_card finite_insert i2 less_Suc_eq sorted_list_of_set(1) sorted_list_of_set(3)) qed also have 2: "... = insort a (sorted_list_of_set I) ! i" by (simp add: a_notin_I finI) also have "... = insort a (sorted_list_of_set I) ! Suc (i-1)" using i0 by auto also have "... < pick I i" proof (cases "i = a'") case True have "(sorted_list_of_set I) ! i > a" by (smt "1" Suc_less_eq True a_def a_notin_I distinct_card distinct_sorted_list_of_set finI i2 ia' index_a' insort_nth2 length_insort lessI list.size(3) nat_less_le not_less_zero pick_in_set_le set_sorted_list_of_set sorted_list_of_set(2) sorted_list_of_set_insert sorted_list_of_set_eq_pick sorted_wrt_nth_less) moreover have "a = insort a (sorted_list_of_set I) ! i" using True 1 2 a_def by auto ultimately show ?thesis using 1 2 by (metis distinct_card finI i0 i2 set_sorted_list_of_set sorted_list_of_set(3) sorted_list_of_set_eq_pick) next case False have "insort a (sorted_list_of_set I) ! Suc (i-1) = (sorted_list_of_set I) ! (i-1)" by (rule insort_nth2, insert i1_length False ia' index_a', auto simp add: a_notin_I finI) also have "... = pick I (i-1)" by (rule sorted_list_of_set_eq_pick[OF i1_length]) also have "... < pick I i" using i0 i2 pick_mono_le by auto finally show ?thesis . qed finally show ?thesis . qed lemma pick_insert3: assumes a_notin_I: "a \ I" and i2: "i < card I" and a_def: "pick (insert a I) a' = a" (*Alternative: index (sorted_list_of_set (insert a I)) a = a'.*) and ia': "i \ a'" (*Case 2*) and a'_card: "a' < card I + 1" shows "pick (insert a I) (Suc i) = pick I i" proof (cases "i = 0") case True have a_LEAST: "a < (LEAST aa. aa\I)" using True a_def a_notin_I i2 ia' pick_insert2 by fastforce have Least_rw: "(LEAST aa. aa = a \ aa \ I) = a" by (rule Least_equality, insert a_notin_I, auto, metis a_LEAST le_less_trans nat_le_linear not_less_Least) let ?P = "\aa. (aa = a \ aa \ I) \ (LEAST aa. aa = a \ aa \ I) < aa" let ?Q = "\aa. aa \ I" have "?P = ?Q" unfolding Least_rw fun_eq_iff by (auto, metis a_LEAST le_less_trans not_le not_less_Least) thus ?thesis using True by auto next case False have finI: "finite I" using i2 card.infinite by force have index_a'1: "index (sorted_list_of_set (insert a I)) a = a'" using pick_index using a'_card a_def a_notin_I finI by auto hence index_a': "index (insort a (sorted_list_of_set I)) a = a'" by (simp add: a_notin_I finI) have i1_length: "i < length (sorted_list_of_set I)" using i2 by (metis distinct_card distinct_sorted_list_of_set finI set_sorted_list_of_set) have 1: "pick (insert a I) (Suc i) = sorted_list_of_set (insert a I) ! (Suc i)" proof (rule sorted_list_of_set_eq_pick[symmetric]) have "finite (insert a I)" using card.infinite i2 by force thus "Suc i < length (sorted_list_of_set (insert a I))" by (metis Suc_mono a_notin_I card_insert_disjoint distinct_card distinct_sorted_list_of_set finI i2 set_sorted_list_of_set) qed also have 2: "... = insort a (sorted_list_of_set I) ! Suc i" by (simp add: a_notin_I finI) also have "... = pick I i" proof (cases "i = a'") case True show ?thesis by (metis True a_notin_I finI i1_length index_a' insort_nth2 le_refl list.size(3) not_less0 set_sorted_list_of_set sorted_list_of_set(2) sorted_list_of_set_eq_pick) next case False have "insort a (sorted_list_of_set I) ! Suc i = (sorted_list_of_set I) ! i" by (rule insort_nth2, insert i1_length False ia' index_a', auto simp add: a_notin_I finI) also have "... = pick I i" by (rule sorted_list_of_set_eq_pick[OF i1_length]) finally show ?thesis . qed finally show ?thesis . qed lemma pick_insert_index: assumes Ik: "card I = k" and a_notin_I: "a \ I" and ik: "i < k" and a_def: "pick (insert a I) a' = a" and a'k: "a' < card I + 1" shows "pick (insert a I) (insert_index a' i) = pick I i" proof (cases "i I" by (simp add: Ik ik pick_in_set_le) show "pick (insert a I) i < pick I i" by (rule pick_insert2[OF a_notin_I _ a_def _ a'k], insert False, auto simp add: Ik ik) fix y assume y: "y \ I \ pick (insert a I) i < y" let ?xs = "sorted_list_of_set (insert a I)" have "y \ set ?xs" using y by (metis fin_aI insertI2 set_sorted_list_of_set y) from this obtain j where xs_j_y: "?xs ! j = y" and j: "j < length ?xs" using in_set_conv_nth by metis have ij: "i pick (insert a I) j" by (metis Ik Suc_lessI card.infinite distinct_card distinct_sorted_list_of_set eq_iff finite_insert ij ik j less_imp_le_nat not_less_zero pick_mono_le set_sorted_list_of_set) also have "... = ?xs ! j" by (rule sorted_list_of_set_eq_pick[symmetric, OF j]) also have "... = y" by (rule xs_j_y) finally show "pick I i \ y" . qed finally show ?thesis unfolding insert_index_def using False by auto qed subsection\Start of the proof\ definition "strict_from_inj n f = (\i. if i\{0.. nat" - assumes "inj_on f {0.. {0.. {0.. f ` {0.. f ` {0.. strict_from_inj n f ` {0..)|f \. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i)} \ {\. \ permutes {0.. carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = det (mat\<^sub>r n n (\i. finsum_vec TYPE('a::comm_ring_1) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) {0..T \ carrier_mat m n" using A by auto have BT: "B\<^sup>T \ carrier_mat n m" using B by auto let ?f = "(\i. finsum_vec TYPE('a) n (\k. B\<^sup>T $$ (i, k) \\<^sub>v Matrix.row A\<^sup>T k) {0..k\{0..T $$ (i, k) \\<^sub>v row A\<^sup>T k) $ j)" by (rule index_finsum_vec[OF _ j_n], auto simp add: A) also have "... = (\k\{0..\<^sub>v col A k) $ j)" proof (rule sum.cong, auto) fix x assume x: "xT x = col A x" by (rule row_transpose, insert A x, auto) have B_rw: "B\<^sup>T $$ (i,x) = B $$ (x, i)" by (rule index_transpose_mat, insert x i B, auto) have "(B\<^sup>T $$ (i, x) \\<^sub>v Matrix.row A\<^sup>T x) $v j = B\<^sup>T $$ (i, x) * Matrix.row A\<^sup>T x $v j" by (rule index_smult_vec, insert A j_n, auto) also have "... = B $$ (x, i) * col A x $v j" unfolding row_rw B_rw by simp also have "... = (B $$ (x, i) \\<^sub>v col A x) $v j" by (rule index_smult_vec[symmetric], insert A j_n, auto) finally show " (B\<^sup>T $$ (i, x) \\<^sub>v Matrix.row A\<^sup>T x) $v j = (B $$ (x, i) \\<^sub>v col A x) $v j" . qed also have "... = ?g i $v j" by (rule index_finsum_vec[symmetric, OF _ j_n], auto simp add: A) also have "... = ?rhs $$ (i, j)" by (rule index_mat[symmetric], insert i j, auto) finally show "?lhs $$ (i, j) = ?rhs $$ (i, j)" . qed have "det (A*B) = det (B\<^sup>T*A\<^sup>T)" using det_transpose by (metis A B Matrix.transpose_mult mult_carrier_mat) also have "... = det (mat\<^sub>r n n (\i. finsum_vec TYPE('a) n (\k. B\<^sup>T $$ (i, k) \\<^sub>v Matrix.row A\<^sup>T k) {0..r n n (\i. finsum_vec TYPE('a) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) {0.. carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = (\f | (\i\{0.. {0.. (\i. i \ {0.. f i = i). (\i = 0..r n n (\i. col A (f i))))" proof - let ?V="{0..r n n (\i. B $$ (f i, i) \\<^sub>v col A (f i))) = (prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i)))" if f: "(\i\{0.. {0.. (\i. i \ {0.. f i = i)" for f by (rule det_rows_mul, insert A col_dim, auto) have "det (A*B) = det (mat\<^sub>r n n (\i. finsum_vec TYPE('a::comm_ring_1) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) ?U))" by (rule det_mul_finsum_alt[OF A B]) also have "... = sum ?g ?F" by (rule det_linear_rows_sum[OF fm], auto simp add: A) also have "... = (\f\?F. prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i))))" using det_rw by auto finally show ?thesis . qed lemma det_cols_mul': assumes A: "A \ carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = (\f | (\i\{0.. {0.. (\i. i \ {0.. f i = i). (\i = 0..r n n (\i. row B (f i))))" proof - let ?F="{f. (\i\{0.. {0.. (\i. i \ {0.. f i = i)}" have t: "A * B = (B\<^sup>T*A\<^sup>T)\<^sup>T" using transpose_mult[OF A B] transpose_transpose by metis have "det (B\<^sup>T*A\<^sup>T) = (\f\?F. (\i = 0..T $$ (f i, i)) * det (mat\<^sub>r n n (\i. col B\<^sup>T (f i))))" by (rule det_cols_mul, auto simp add: A B) also have "... = (\f \?F. (\i = 0..r n n (\i. row B (f i))))" proof (rule sum.cong, rule refl) fix f assume f: "f \ ?F" have "(\i = 0..T $$ (f i, i)) = (\i = 0.. {0..T $$ (f x, x) = A $$ (x, f x)" by (rule index_transpose_mat(1), insert f A x, auto) qed moreover have "det (mat\<^sub>r n n (\i. col B\<^sup>T (f i))) = det (mat\<^sub>r n n (\i. row B (f i)))" proof - have row_eq_colT: "row B (f i) $v j = col B\<^sup>T (f i) $v j" if i: "i < n" and j: "j < n" for i j proof - have fi_m: "f i < m" using f i by auto have "col B\<^sup>T (f i) $v j = B\<^sup>T $$(j, f i)" by (rule index_col, insert B fi_m j, auto) also have "... = B $$ (f i, j)" using B fi_m j by auto also have "... = row B (f i) $v j" by (rule index_row[symmetric], insert B fi_m j, auto) finally show ?thesis .. qed show ?thesis by (rule arg_cong[of _ _ det], rule eq_matI, insert row_eq_colT, auto) qed ultimately show "(\i = 0..T $$ (f i, i)) * det (mat\<^sub>r n n (\i. col B\<^sup>T (f i))) = (\i = 0..r n n (\i. row B (f i)))" by simp qed finally show ?thesis by (metis (no_types, lifting) A B det_transpose transpose_mult mult_carrier_mat) qed (*We need a more general version of this lemma*) lemma assumes F: "F= {f. f \ {0.. {0.. (\i. i \ {0.. f i = i)}" and p: " \ permutes {0..f\F. (\i = 0.. i))) = (\f\F. (\i = 0.. \ = g \ \" have "f x = g x" for x proof (cases "x \ {0.. F" show "xa \ \ \ F" unfolding o_def F using F PiE p xa by (auto, smt F atLeastLessThan_iff mem_Collect_eq p permutes_def xa) show "\x\F. xa = x \ \" proof (rule bexI[of _ "xa \ Hilbert_Choice.inv \"]) show "xa = xa \ Hilbert_Choice.inv \ \ \" using p by auto show "xa \ Hilbert_Choice.inv \ \ F" unfolding o_def F using F PiE p xa by (auto, smt atLeastLessThan_iff permutes_def permutes_less(3)) qed qed have prod_rw: "(\i = 0..i = 0.. i), \ i))" if "f\F" for f using prod.permute[OF p] by auto let ?g = "\f. (\i = 0.. i))" have "(\f\F. (\i = 0..f\F. (\i = 0.. i), \ i)))" using prod_rw by auto also have "... = (\f\(?h`F). \i = 0.. i))" using sum.reindex[OF inj_on_F, of ?g] unfolding hF by auto also have "... = (\f\F. \i = 0.. i))" unfolding hF by auto finally show ?thesis .. qed lemma detAB_Znm_aux: assumes F: "F= {f. f \ {0.. {0.. (\i. i \ {0.. f i = i)}" shows"(\\ | \ permutes {0..f\F. prod (\i. B $$ (f i, i)) {0.. * (\i = 0.. i, f i))))) = (\\ | \ permutes {0..f\F. (\i = 0.. i)) * (signof \ * (\i = 0..\ | \ permutes {0..f\F. prod (\i. B $$ (f i, i)) {0.. * (\i = 0.. i, f i))))) = (\\ | \ permutes {0..f\F. signof \ * (\i = 0.. i, f i)))" by (smt mult.left_commute prod.cong prod.distrib sum.cong) also have "... = (\\ | \ permutes {0..f\F. signof (Hilbert_Choice.inv \) * (\i = 0.. i, f i)))" by (rule sum_permutations_inverse) also have "... = (\\ | \ permutes {0..f\F. signof (Hilbert_Choice.inv \) * (\i = 0.. i), (\ i)) * A $$ (Hilbert_Choice.inv \ (\ i), f (\ i))))" proof (rule sum.cong) fix x assume x: "x \ {\. \ permutes {0..i = 0..i = 0.. F" for f using prod.permute[OF p] by auto then show "(\f\F. signof ?inv_x * (\i = 0..f\F. signof ?inv_x * (\i = 0..\ | \ permutes {0..f\F. signof \ * (\i = 0.. i), (\ i)) * A $$ (i, f (\ i))))" by (rule sum.cong, auto, rule sum.cong, auto) (metis (no_types, lifting) finite_atLeastLessThan signof_inv) also have "... = (\\ | \ permutes {0..f\F. signof \ * (\i = 0.. i)) * A $$ (i, f i)))" proof (rule sum.cong) fix \ assume p: "\ \ {\. \ permutes {0.. permutes {0.. ?inv_pi = g \ ?inv_pi" have "f x = g x" for x proof (cases "x \ {0.. F" show "xa \ ?inv_pi \ F" unfolding o_def F using F PiE p xa by (auto, smt atLeastLessThan_iff permutes_def permutes_less(3)) show "\x\F. xa = x \ ?inv_pi" proof (rule bexI[of _ "xa \ \"]) show "xa = xa \ \ \ Hilbert_Choice.inv \ " using p by auto show "xa \ \ \ F" unfolding o_def F using F PiE p xa by (auto, smt atLeastLessThan_iff permutes_def permutes_less(3)) qed qed let ?g = "\f. signof \ * (\i = 0.. i), \ i) * A $$ (i, f (\ i)))" show "(\f\F. signof \ * (\i = 0.. i), \ i) * A $$ (i, f (\ i)))) = (\f\F. signof \ * (\i = 0.. i) * A $$ (i, f i)))" using sum.reindex[OF inj_on_F, of "?g"] p unfolding hF unfolding o_def by auto qed (simp) also have "... = (\\ | \ permutes {0..f\F. (\i = 0.. i)) * (signof \ * (\i = 0.. carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = (\(f, \)\Z n m. signof \ * (\i = 0.. i)))" proof - let ?V="{0.. {0.. {0.. (\i. i \ {0.. f i = i)}" by auto have det_rw: "det (mat\<^sub>r n n (\i. B $$ (f i, i) \\<^sub>v col A (f i))) = (prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i)))" if f: "(\i\{0.. {0.. (\i. i \ {0.. f i = i)" for f by (rule det_rows_mul, insert A col_dim, auto) have det_rw2: "det (mat\<^sub>r n n (\i. col A (f i))) = (\\ | \ permutes {0.. * (\i = 0.. i, f i)))" if f: "f \ ?F" for f proof (unfold Determinant.det_def, auto, rule sum.cong, auto) fix x assume x: "x permutes {0..i = 0..i = 0.. {0..i = 0..i = 0..r n n (\i. finsum_vec TYPE('a::comm_ring_1) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) {0..f\?F. prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i))))" using det_rw by auto also have "... = (\f\?F. prod (\i. B $$ (f i, i)) {0..\ | \ permutes {0.. * (\i = 0.. i, f (i)))))" by (rule sum.cong, auto simp add: det_rw2) also have "... = (\f\?F. \\ | \ permutes {0..i. B $$ (f i, i)) {0.. * (\i = 0.. i, f (i)))))" by (simp add: mult_hom.hom_sum) also have "... = (\\ | \ permutes {0..f\?F.prod (\i. B $$ (f i, i)) {0.. * (\i = 0.. i, f i))))" by (rule VS_Connect.class_semiring.finsum_finsum_swap, insert finite_permutations finite_bounded_functions[OF fin_m fin_n], auto) thm detAB_Znm_aux also have "... = (\\ | \ permutes {0..f\?F. (\i = 0.. i)) * (signof \ * (\i = 0..f\?F.\\ | \ permutes {0..i = 0.. i)) * (signof \ * (\i = 0..f\?F.\\ | \ permutes {0.. * (\i = 0.. i)))" unfolding prod.distrib by (rule sum.cong, auto, rule sum.cong, auto) also have "... = sum (\(f,\). (signof \) * (prod (\i. A$$(i,f i) * B $$ (f i, \ i)) {0.. carrier_mat n m" and B: "B \ carrier_mat m n" begin private definition "Z_inj = ({f. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ \ inj_on f {0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) - \ strict_mono_on f {0.. {\. \ permutes {0.. strict_mono_on {0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) - \ \ strict_mono_on f {0.. {\. \ permutes {0.. \ strict_mono_on {0.. {\. \ permutes {0.. = (signof \) * (prod (\i. A$$(i,f i) * B $$ (f i, \ i)) {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ strict_mono_on f {0.. (\i. i \ {0.. f i = i) \ strict_mono_on {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ \ inj_on f {0.. {0.. {0.. (\i. i \ {0.. f i = i)}" text\The Cauchy--Binet formula is proven in \url{https://core.ac.uk/download/pdf/82475020.pdf} In that work, they define @{text "\ \ inv \ \ \"}. I had problems following this proof in Isabelle, since I was demanded to show that such permutations commute, which is false. It is a notation problem of the @{text "\"} operator, the author means @{text "\ \ \ \ inv \"} using the Isabelle notation (i.e., @{text "\ x = \ ((inv \) x)"}). \ lemma step_weight: fixes \ \ defines "\ \ \ \ Hilbert_Choice.inv \" assumes f_inj: "f \ F_inj" and gF: "g \ F" and pi: "\ permutes {0.. permutes {0..x \ {0.. x)" shows "weight f \ = (signof \) * (\i = 0.. i))) * (signof \) * (\i = 0.. i))" proof - let ?A = "(\i = 0.. i))) " let ?B = "(\i = 0.. i))" have sigma: "\ permutes {0.._def by (rule permutes_compose[OF permutes_inv[OF phi] pi]) have A_rw: "?A = (\i = 0..i = 0.. i), \ (\ i)))" by (rule prod.permute[unfolded o_def, OF phi]) also have "... = (\i = 0.. i))" using fg_phi unfolding \_def unfolding o_def unfolding permutes_inverses(2)[OF phi] by auto finally have B_rw: "?B = (\i = 0.. i))" . have "(signof \) * ?A * (signof \) * ?B = (signof \) * (signof \) * ?A * ?B" by auto also have "... = signof (\ \ \) * ?A * ?B" unfolding signof_compose[OF phi sigma] by simp also have "... = signof \ * ?A * ?B" by (metis (no_types, lifting) \_def mult.commute o_inv_o_cancel permutes_inj phi sigma signof_compose) also have "... = signof \ * (\i = 0..i = 0.. i))" using A_rw B_rw by auto also have "... = signof \ * (\i = 0.. i))" by auto also have "... = weight f \" unfolding weight_def by simp finally show ?thesis .. qed lemma Z_good_fun_alt_sum: fixes g defines "Z_good_fun \ {f. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. F_inj" shows "(\f\Z_good_fun. P f)= (\\\{\. \ permutes {0.. \))" proof - let ?f = "\\. g \ \" let ?P = "{\. \ permutes {0.. i < n" hence "xa i = i" unfolding permutes_def by auto thus "g (xa i) = i" using g i_ge_n unfolding F_inj_def by auto next fix xa assume "xa permutes {0.. xa) {0.. xb assume "\ permutes {0.. (\x. g (\ x)) ` {0.. {0.. {0..i. \ i < n \ x i = i" and inj_on_x: "inj_on x {0.. = "\i. if i x i = g j) else i" show "x \ (\) g ` {\. \ permutes {0..], rule conjI) have "?\ i = i" if i: "i \ {0..!j. ?\ j = i" for i proof (cases "i x a = g j) = i" proof (rule theI2) show "i < n \ x a = g i" using xa_gi True by auto fix xa assume "xa < n \ x a = g xa" thus "xa = i" by (metis (mono_tags, lifting) F_inj_def True atLeast0LessThan g inj_onD lessThan_iff mem_Collect_eq xa_gi) thus "xa = i" . qed thus ta: "?\ a = i" using a by auto fix j assume tj: "?\ j = i" show "j = a" proof (cases "j x j = g ja) = i" using tj True by auto have "?P (THE ja. ?P ja)" proof (rule theI) show "b < n \ x j = g b" using xj_gb b by auto fix xa assume "xa < n \ x j = g xa" thus "xa = b" by (metis (mono_tags, lifting) F_inj_def b atLeast0LessThan g inj_onD lessThan_iff mem_Collect_eq xj_gb) qed hence "x j = g i" unfolding the_ji by auto hence "x j = x a" using xa_gi by auto then show ?thesis using inj_on_x a True unfolding inj_on_def by auto next case False then show ?thesis using tj True by auto qed qed next case False note i_ge_n = False show ?thesis proof (rule ex1I[of _ i]) show "?\ i = i" using False by simp fix j assume tj: "?\ j = i" show "j = i" proof (cases "j x j = g ja) < n" proof (rule theI2) show "a < n \ x j = g a" using xj_ga a by auto fix xa assume a1: "xa < n \ x j = g xa" thus "xa = a" using F_inj_def a atLeast0LessThan g inj_on_eq_iff xj_ga by fastforce show "xa < n" by (simp add: a1) qed then show ?thesis using tj i_ge_n by auto next case False then show ?thesis using tj by auto qed qed qed ultimately show "?\ permutes {0.. ?\" proof - have "x xa = g (THE j. j < n \ x xa = g j)" if xa: "xa < n" for xa proof - obtain c where c: "c < n" and xxa_gc: "x xa = g c" by (metis (mono_tags, opaque_lifting) atLeast0LessThan imageE imageI lessThan_iff xa xg) show ?thesis proof (rule theI2) show c1: "c < n \ x xa = g c" using c xxa_gc by auto fix xb assume c2: "xb < n \ x xa = g xb" thus "xb = c" by (metis (mono_tags, lifting) F_inj_def c1 atLeast0LessThan g inj_onD lessThan_iff mem_Collect_eq) show "x xa = g xb" using c1 c2 by simp qed qed moreover have "x xa = g xa" if xa: "\ xa < n" for xa using g x1 x2 xa unfolding F_inj_def by simp ultimately show ?thesis unfolding o_def fun_eq_iff by auto qed qed qed have inj: "inj_on ?f ?P" proof (rule inj_onI) fix x y assume x: "x \ ?P" and y: "y \ ?P" and gx_gy: "g \ x = g \ y" have "x i = y i" for i proof (cases "i {0.. {0..f\Z_good_fun. P f) = (\f\?f`?P. P f)" using fP by simp also have "... = sum (P \ (\) g) {\. \ permutes {0..\ | \ permutes {0.. \))" by auto finally show ?thesis . qed lemma F_injI: assumes "f \ {0.. {0..i. i \ {0.. f i = i)" and "inj_on f {0.. F_inj" using assms unfolding F_inj_def by simp lemma F_inj_composition_permutation: assumes phi: "\ permutes {0.. F_inj" shows "g \ \ \ F_inj" proof (rule F_injI) show "g \ \ \ {0.. {0..i. i \ {0.. (g \ \) i = i" using g phi unfolding permutes_def F_inj_def by simp show "inj_on (g \ \) {0.. F_strict" shows "f \ F_inj" using f strict_mono_on_imp_inj_on unfolding F_strict_def F_inj_def by auto lemma one_step: assumes g1: "g \ F_strict" shows "det (submatrix A UNIV (g`{0..(x, y) \ Z_good g. weight x y)" (is "?lhs = ?rhs") proof - define Z_good_fun where "Z_good_fun = {f. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. F_inj" by (rule F_strict_imp_F_inj[OF g1]) have detA: "(\\\{\. \ permutes {0.. * (\i = 0.. i)))) = det (submatrix A UNIV (g`{0.. j \ g ` {0.. g ` {0.. j \ g ` {0.. carrier_mat n n" unfolding submatrix_def card_J using A by auto have "det (submatrix A UNIV (g`{0..p | p permutes {0..i = 0..\\{\. \ permutes {0.. * (\i = 0.. i))))" proof (rule sum.cong) fix x assume x: "x \ {\. \ permutes {0..i = 0..i = 0.. {0.. (g ` {0..i = 0..i = 0..\ \ ?Perm. signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i))) = (\\ \ ?Perm. signof (\) * (\i = 0.. i)))" if phi: "\ permutes {0.. proof - let ?h="\\. \ \ ?inv \" let ?g = "\\. signof (\) * (\i = 0.. i))" have "?h`?Perm = ?Perm" proof - have "\ \ ?inv \ permutes {0.. permutes {0.. using permutes_compose permutes_inv phi that by blast moreover have "x \ (\\. \ \ ?inv \) ` ?Perm" if "x permutes {0.. \ permutes {0.. \ \ ?inv \" using phi by auto ultimately show ?thesis unfolding image_def by auto qed ultimately show ?thesis by auto qed hence "(\\ \ ?Perm. ?g \) = (\\ \ ?h`?Perm. ?g \)" by simp also have "... = sum (?g \ ?h) ?Perm" proof (rule sum.reindex) show "inj_on (\\. \ \ ?inv \) {\. \ permutes {0..\ \ ?Perm. signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i)))" unfolding o_def by auto finally show ?thesis by simp qed have detB: "det (submatrix B (g`{0..\ \ ?Perm. signof \ * (\i = 0.. i)))" proof - have "{i. i < dim_row B \ i \ g ` {0.. g ` {0.. j \ g ` {0.. carrier_mat n n" unfolding submatrix_def using card_I B by auto have "det (submatrix B (g`{0..p \ ?Perm. signof p * (\i=0..\ \ ?Perm. signof \ * (\i = 0.. i)))" proof (rule sum.cong, rule refl) fix x assume x: "x \ {\. \ permutes {0..i=0..i=0.. {0.. (g ` {0..i = 0..i = 0..f\Z_good_fun. \\\?Perm. weight f \)" unfolding Z_good_def sum.cartesian_product Z_good_fun_def by blast also have "... = (\\\{\. \ permutes {0.. \))" unfolding Z_good_fun_def by (rule Z_good_fun_alt_sum[OF g]) also have "... = (\\\{\. \ permutes {0..\\{\. \ permutes {0.. * (\i = 0.. i))) * signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i)))" proof (rule sum.cong, simp, rule sum.cong, simp) fix \ \ assume phi: "\ \ ?Perm" and pi: "\ \ ?Perm" show "weight (g \ \) \ = signof \ * (\i = 0.. i))) * signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i))" proof (rule step_weight) show "g \ \ \ F_inj" by (rule F_inj_composition_permutation[OF _ g], insert phi, auto) show "g \ F" using g unfolding F_def F_inj_def by simp qed (insert phi pi, auto) qed also have "... = (\\\{\. \ permutes {0.. * (\i = 0.. i))) * (\\ | \ permutes {0.. \ ?inv \) * (\i = 0.. \ ?inv \) i))))" by (metis (mono_tags, lifting) Groups.mult_ac(1) semiring_0_class.sum_distrib_left sum.cong) also have "... = (\\ \ ?Perm. signof \ * (\i = 0.. i))) * (\\ \ ?Perm. signof \ * (\i = 0.. i))))" using detB_rw by auto also have "... = (\\ \ ?Perm. signof \ * (\i = 0.. i)))) * (\\ \ ?Perm. signof \ * (\i = 0.. i)))" by (simp add: semiring_0_class.sum_distrib_right) also have "... = ?lhs" unfolding detA detB .. finally show ?thesis .. qed lemma gather_by_strictness: "sum (\g. sum (\(f,\). weight f \) (Z_good g)) F_strict = sum (\g. det (submatrix A UNIV (g`{0.. F_strict" show "(\(x, y)\Z_good f. weight x y) = det (submatrix A UNIV (f ` {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite ?A" using rev_finite_subset by blast show "finite {\. \ permutes {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite ?A" using rev_finite_subset by blast show "finite {\. \ permutes {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite ?A" using rev_finite_subset by blast show "finite {\. \ permutes {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite F_inj" unfolding F_inj_def using rev_finite_subset by blast qed lemma finite_F_strict[simp]: "finite F_strict" proof - have finN: "finite {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite F_strict" unfolding F_strict_def using rev_finite_subset by blast qed lemma nth_strict_mono: fixes f::"nat \ nat" assumes strictf: "strict_mono f" and i: "i ?I. a < f i} = i" using i proof (induct i) case 0 then show ?case by (auto simp add: strict_mono_less strictf) next case (Suc i) have i: "i < n" using Suc.prems by auto let ?J'="{a \ f ` {0.. f i" and 2: "f xa < f (Suc i)" show "f xa < f i" using 1 2 not_less_less_Suc_eq strict_mono_less strictf by fastforce next fix xa assume "f xa < f i" thus "f xa < f (Suc i)" using less_SucI strict_mono_less strictf by blast next show "f i \ f ` {0..card (f ` {0.. i) also have "... = pick ?I (card {a \ ?I. a < f i})" unfolding card_eq by simp also have "... = f i" by (rule pick_card_in_set, simp add: i) finally show ?thesis .. qed lemma nth_strict_mono_on: fixes f::"nat \ nat" - assumes strictf: "strict_mono_on f {0.. ?I. a < f i} = i" using i proof (induct i) case 0 then show ?case by (auto, metis (no_types, lifting) atLeast0LessThan lessThan_iff less_Suc_eq not_less0 not_less_eq strict_mono_on_def strictf) next case (Suc i) have i: "i < n" using Suc.prems by auto let ?J'="{a \ f ` {0.. f i" and 2: "f xa < f (Suc i)" and 3: "xa < n" show "f xa < f i" by (metis (full_types) 1 2 3 antisym_conv3 atLeast0LessThan i lessThan_iff less_SucE order.asym strict_mono_onD strictf) next fix xa assume "f xa < f i" and "xa < n" thus "f xa < f (Suc i)" using less_SucI strictf by (metis (no_types, lifting) Suc.prems atLeast0LessThan lessI lessThan_iff less_trans strict_mono_onD) next show "f i \ f ` {0..card (f ` {0.. i) also have "... = pick ?I (card {a \ ?I. a < f i})" unfolding card_eq by simp also have "... = f i" by (rule pick_card_in_set, simp add: i) finally show ?thesis .. qed lemma strict_fun_eq: assumes f: "f \ F_strict" and g: "g \ F_strict" and fg: "f`{0.. F_inj" shows "strict_from_inj n f \ F" proof - { fix x assume x: "x < n" have inj_on: "inj_on f {0.. a \ f ` {0.. a \ f ` {0.. F_strict" if xa: "xa \ F_inj" for xa proof - - have "strict_mono_on (strict_from_inj n xa) {0.. F_inj" shows "strict_from_inj n f ` {0.. a \ f ` {0.. a \ f ` {0.. f ` {0..card (f ` {0.. xa) finally show "strict_from_inj n f xa \ f ` {0.. strict_from_inj n f ` {0.. F_strict" shows "Z_good g = {x \ F_inj. strict_from_inj n x = g} \ {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. F_inj. strict_from_inj n x = g}" proof (auto) fix f assume f: "f \ Z_good_fun" thus f_inj: "f \ F_inj" unfolding F_inj_def Z_good_fun_def by auto show "strict_from_inj n f = g" proof (rule strict_fun_eq[OF _ g]) show "strict_from_inj n f ` {0.. F_strict" using F_strict_def f_inj strict_from_inj_F_strict by blast qed next fix f assume f_inj: "f \ F_inj" and g_strict_f: "g = strict_from_inj n f" have "f xa \ g ` {0.. f ` {0.. Z_good_fun" using f_inj g_strict_f unfolding Z_good_fun_def F_inj_def by auto qed thus ?thesis unfolding Z_good_fun_def Z_good_def by simp qed lemma weight_0: "(\(f, \) \ Z_not_inj. weight f \) = 0" proof - let ?F="{f. (\i\{0.. {0.. (\i. i \ {0.. f i = i)}" let ?Perm = "{\. \ permutes {0..(f, \)\Z_not_inj. weight f \) = (\f \ F_not_inj. (\i = 0..r n n (\i. row B (f i))))" proof - have dim_row_rw: "dim_row (mat\<^sub>r n n (\i. col A (f i))) = n" for f by auto have dim_row_rw2: "dim_row (mat\<^sub>r n n (\i. Matrix.row B (f i))) = n" for f by auto have prod_rw: "(\i = 0.. i)) = (\i = 0.. i)" if f: "f \ F_not_inj" and pi: "\ \ ?Perm" for f \ proof (rule prod.cong, rule refl) fix x assume x: "x \ {0.. x < dim_col B" using x pi B by auto ultimately show "B $$ (f x, \ x) = Matrix.row B (f x) $v \ x" by (rule index_row[symmetric]) qed have sum_rw: "(\\ | \ permutes {0.. * (\i = 0.. i))) = det (mat\<^sub>r n n (\i. row B (f i)))" if f: "f \ F_not_inj" for f unfolding Determinant.det_def using dim_row_rw2 prod_rw f by auto have "(\(f, \)\Z_not_inj. weight f \) = (\f\F_not_inj.\\ \ ?Perm. weight f \)" unfolding Z_not_inj_def unfolding sum.cartesian_product unfolding F_not_inj_def by simp also have "... = (\f\F_not_inj. \\ | \ permutes {0.. * (\i = 0.. i)))" unfolding weight_def by simp also have "... = (\f\F_not_inj. (\i = 0..\ | \ permutes {0.. * (\i = 0.. i))))" by (rule sum.cong, rule refl, auto) (metis (no_types, lifting) mult.left_commute mult_hom.hom_sum sum.cong) also have "... = (\f \ F_not_inj. (\i = 0..r n n (\i. row B (f i))))" using sum_rw by auto finally show ?thesis by auto qed also have "... = 0" by (rule sum.neutral, insert det_not_inj_on[of _ n B], auto simp add: F_not_inj_def) finally show ?thesis . qed subsection \Final theorem\ lemma Cauchy_Binet1: shows "det (A*B) = sum (\f. det (submatrix A UNIV (f`{0..(f, \) \ Z_not_inj. weight f \) = 0" by (rule weight_0) let ?f = "strict_from_inj n" have sum_rw: "sum g F_inj = (\y \ F_strict. sum g {x \ F_inj. ?f x = y})" for g by (rule sum.group[symmetric], insert strict_from_inj_F_strict, auto) have Z_Union: "Z_inj \ Z_not_inj = Z n m" unfolding Z_def Z_not_inj_def Z_inj_def by auto have Z_Inter: "Z_inj \ Z_not_inj = {}" unfolding Z_def Z_not_inj_def Z_inj_def by auto have "det (A*B) = (\(f, \)\Z n m. weight f \)" using detAB_Znm[OF A B] unfolding weight_def by auto also have "... = (\(f, \)\Z_inj. weight f \) + (\(f, \)\Z_not_inj. weight f \)" by (metis Z_Inter Z_Union finite_Un finite_Znm sum.union_disjoint) also have "... = (\(f, \)\Z_inj. weight f \)" using sum0 by force also have "... = (\f \ F_inj. \\\{\. \ permutes {0..)" unfolding Z_inj_def unfolding F_inj_def sum.cartesian_product .. also have "... = (\y\F_strict. \f\{x \ F_inj. strict_from_inj n x = y}. sum (weight f) {\. \ permutes {0..y\F_strict. \(f,\)\({x \ F_inj. strict_from_inj n x = y} \ {\. \ permutes {0..)" unfolding F_inj_def sum.cartesian_product .. also have "... = sum (\g. sum (\(f,\). weight f \) (Z_good g)) F_strict" using Z_good_alt by auto also have "... = ?rhs" unfolding gather_by_strictness by simp finally show ?thesis . qed lemma Cauchy_Binet: "det (A*B) = (\I\{I. I\{0.. card I=n}. det (submatrix A UNIV I) * det (submatrix B I UNIV))" proof - let ?f="(\I. (\i. if i ?setI" and J: "J \ ?setI" and fI_fJ: "?f I = ?f J" have "x \ J" if x: "x \ I" for x by (metis (mono_tags) fI_fJ I J distinct_card in_set_conv_nth mem_Collect_eq sorted_list_of_set(1) sorted_list_of_set(3) subset_eq_atLeast0_lessThan_finite x) moreover have "x \ I" if x: "x \ J" for x by (metis (mono_tags) fI_fJ I J distinct_card in_set_conv_nth mem_Collect_eq sorted_list_of_set(1) sorted_list_of_set(3) subset_eq_atLeast0_lessThan_finite x) ultimately show "I = J" by auto qed have rw: "?f I ` {0.. ?setI" for I proof - have "sorted_list_of_set I ! xa \ I" if "xa < n" for xa by (metis (mono_tags, lifting) I distinct_card distinct_sorted_list_of_set mem_Collect_eq nth_mem set_sorted_list_of_set subset_eq_atLeast0_lessThan_finite that) moreover have "\xa\{0..I" for x by (metis (full_types) x I atLeast0LessThan distinct_card in_set_conv_nth mem_Collect_eq lessThan_iff sorted_list_of_set(1) sorted_list_of_set(3) subset_eq_atLeast0_lessThan_finite) ultimately show ?thesis unfolding image_def by auto qed have f_setI: "?f` ?setI = F_strict" proof - have "sorted_list_of_set I ! xa < m" if I: "I \ {0..xa < card I\ atLeast0LessThan distinct_card finite_atLeastLessThan lessThan_iff pick_in_set_le rev_finite_subset sorted_list_of_set(1) sorted_list_of_set(3) sorted_list_of_set_eq_pick subsetCE) - moreover have "strict_mono_on (\i. if i < card I then sorted_list_of_set I ! i else i) {0..i. if i < card I then sorted_list_of_set I ! i else i)" if "I \ {0..I \ {0.. atLeastLessThan_iff distinct_card finite_atLeastLessThan pick_mono_le rev_finite_subset sorted_list_of_set(1) sorted_list_of_set(3) sorted_list_of_set_eq_pick strict_mono_on_def) moreover have "x \ ?f ` {I. I \ {0.. card I = n}" if x1: "x \ {0.. {0..i. \ i < n \ x i = i" - and s: "strict_mono_on x {0..i. if i < n then sorted_list_of_set (x ` {0..f. det (submatrix A UNIV (f ` {0.. ?f) {I. I \ {0.. card I = n}" unfolding Cauchy_Binet1 f_setI[symmetric] by (rule sum.reindex[OF inj_on]) also have "... = (\I\{I. I\{0.. card I=n}.det(submatrix A UNIV I)*det(submatrix B I UNIV))" by (rule sum.cong, insert rw, auto) finally show ?thesis . qed end end diff --git a/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy b/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy --- a/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy +++ b/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy @@ -1,1168 +1,1168 @@ (* Author: Jose Divasón Email: jose.divason@unirioja.es *) section \Missing results\ theory SNF_Missing_Lemmas imports Hermite.Hermite Mod_Type_Connect Jordan_Normal_Form.DL_Rank_Submatrix "List-Index.List_Index" begin text \This theory presents some missing lemmas that are required for the Smith normal form development. Some of them could be added to different AFP entries, such as the Jordan Normal Form AFP entry by Ren\'e Thiemann and Akihisa Yamada. However, not all the lemmas can be added directly, since some imports are required.\ hide_const (open) C hide_const (open) measure subsection \Miscellaneous lemmas\ lemma sum_two_rw: "(\i = 0..<2. f i) = (\i \ {0,1::nat}. f i)" by (rule sum.cong, auto) lemma sum_common_left: fixes f::"'a \ 'b::comm_ring_1" assumes "finite A" shows "sum (\i. c * f i) A = c * sum f A" by (simp add: mult_hom.hom_sum) lemma prod3_intro: assumes "fst A = a" and "fst (snd A) = b" and "snd (snd A) = c" shows "A = (a,b,c)" using assms by auto subsection \Transfer rules for the HMA\_Connect file of the Perron-Frobenius development\ hide_const (open) HMA_M HMA_I to_hma\<^sub>m from_hma\<^sub>m hide_fact (open) from_hma\<^sub>m_def from_hma_to_hma\<^sub>m HMA_M_def HMA_I_def dim_row_transfer_rule dim_col_transfer_rule context includes lifting_syntax begin lemma HMA_invertible_matrix[transfer_rule]: "((HMA_Connect.HMA_M :: _ \ 'a :: comm_ring_1 ^ 'n ^ 'n \ _) ===> (=)) invertible_mat invertible" proof (intro rel_funI, goal_cases) case (1 x y) note rel_xy[transfer_rule] = "1" have eq_dim: "dim_col x = dim_row x" using HMA_Connect.dim_col_transfer_rule HMA_Connect.dim_row_transfer_rule rel_xy by fastforce moreover have "\A'. y ** A' = Finite_Cartesian_Product.mat 1 \ A' ** y = Finite_Cartesian_Product.mat 1" if xB: "x * B = 1\<^sub>m (dim_row x)" and Bx: "B * x = 1\<^sub>m (dim_row B)" for B proof - let ?A' = "HMA_Connect.to_hma\<^sub>m B:: 'a :: comm_ring_1 ^ 'n ^ 'n" have rel_BA[transfer_rule]: "HMA_M B ?A'" by (metis (no_types, lifting) Bx HMA_M_def eq_dim carrier_mat_triv dim_col_mat(1) from_hma\<^sub>m_def from_hma_to_hma\<^sub>m index_mult_mat(3) index_one_mat(3) rel_xy xB) have [simp]: "dim_row B = CARD('n)" using dim_row_transfer_rule rel_BA by blast have [simp]: "dim_row x = CARD('n)" using dim_row_transfer_rule rel_xy by blast have "y ** ?A' = Finite_Cartesian_Product.mat 1" using xB by (transfer, simp) moreover have "?A' ** y = Finite_Cartesian_Product.mat 1" using Bx by (transfer, simp) ultimately show ?thesis by blast qed moreover have "\B. x * B = 1\<^sub>m (dim_row x) \ B * x = 1\<^sub>m (dim_row B)" if yA: "y ** A' = Finite_Cartesian_Product.mat 1" and Ay: "A' ** y = Finite_Cartesian_Product.mat 1" for A' proof - let ?B = "(from_hma\<^sub>m A')" have [simp]: "dim_row x = CARD('n)" using dim_row_transfer_rule rel_xy by blast have [transfer_rule]: "HMA_M ?B A'" by (simp add: HMA_M_def) hence [simp]: "dim_row ?B = CARD('n)" using dim_row_transfer_rule by auto have "x * ?B = 1\<^sub>m (dim_row x)" using yA by (transfer', auto) moreover have "?B * x = 1\<^sub>m (dim_row ?B)" using Ay by (transfer', auto) ultimately show ?thesis by auto qed ultimately show ?case unfolding invertible_mat_def invertible_def inverts_mat_def by auto qed end subsection \Lemmas obtained from HOL Analysis using local type definitions\ thm Cartesian_Space.invertible_mult (*In HOL Analysis*) thm invertible_iff_is_unit (*In HOL Analysis*) thm det_non_zero_imp_unit (*In JNF, but only for fields*) thm mat_mult_left_right_inverse (*In JNF, but only for fields*) lemma invertible_mat_zero: assumes A: "A \ carrier_mat 0 0" shows "invertible_mat A" using A unfolding invertible_mat_def inverts_mat_def one_mat_def times_mat_def scalar_prod_def Matrix.row_def col_def carrier_mat_def by (auto, metis (no_types, lifting) cong_mat not_less_zero) lemma invertible_mult_JNF: fixes A::"'a::comm_ring_1 mat" assumes A: "A\carrier_mat n n" and B: "B\carrier_mat n n" and inv_A: "invertible_mat A" and inv_B: "invertible_mat B" shows "invertible_mat (A*B)" proof (cases "n = 0") case True then show ?thesis using assms by (simp add: invertible_mat_zero) next case False then show ?thesis using invertible_mult[where ?'a="'a::comm_ring_1", where ?'b="'n::finite", where ?'c="'n::finite", where ?'d="'n::finite", untransferred, cancel_card_constraint, OF assms] by auto qed lemma invertible_iff_is_unit_JNF: assumes A: "A \ carrier_mat n n" shows "invertible_mat A \ (Determinant.det A) dvd 1" proof (cases "n=0") case True then show ?thesis using det_dim_zero invertible_mat_zero A by auto next case False then show ?thesis using invertible_iff_is_unit[untransferred, cancel_card_constraint] A by auto qed subsection \Lemmas about matrices, submatrices and determinants\ (*This is a generalization of thm mat_mult_left_right_inverse*) thm mat_mult_left_right_inverse lemma mat_mult_left_right_inverse: fixes A :: "'a::comm_ring_1 mat" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and AB: "A * B = 1\<^sub>m n" shows "B * A = 1\<^sub>m n" proof - have "Determinant.det (A * B) = Determinant.det (1\<^sub>m n)" using AB by auto hence "Determinant.det A * Determinant.det B = 1" using Determinant.det_mult[OF A B] det_one by auto hence det_A: "(Determinant.det A) dvd 1" and det_B: "(Determinant.det B) dvd 1" using dvd_triv_left dvd_triv_right by metis+ hence inv_A: "invertible_mat A" and inv_B: "invertible_mat B" using A B invertible_iff_is_unit_JNF by blast+ obtain B' where inv_BB': "inverts_mat B B'" and inv_B'B: "inverts_mat B' B" using inv_B unfolding invertible_mat_def by auto have B'_carrier: "B' \ carrier_mat n n" by (metis B inv_B'B inv_BB' carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mult_mat(3) index_one_mat(3) inverts_mat_def) have "B * A * B = B" using A AB B by auto hence "B * A * (B * B') = B * B'" by (smt A AB B B'_carrier assoc_mult_mat carrier_matD(1) inv_BB' inverts_mat_def one_carrier_mat) thus ?thesis by (metis A B carrier_matD(1) carrier_matD(2) index_mult_mat(3) inv_BB' inverts_mat_def right_mult_one_mat') qed context comm_ring_1 begin lemma col_submatrix_UNIV: assumes "j < card {i. i < dim_col A \ i \ J}" shows "col (submatrix A UNIV J) j = col A (pick J j)" proof (rule eq_vecI) show dim_eq:"dim_vec (col (submatrix A UNIV J) j) = dim_vec (col A (pick J j))" by (simp add: dim_submatrix(1)) fix i assume "i < dim_vec (col A (pick J j))" show "col (submatrix A UNIV J) j $v i = col A (pick J j) $v i" by (smt Collect_cong assms col_def dim_col dim_eq dim_submatrix(1) eq_vecI index_vec pick_UNIV submatrix_index) qed lemma submatrix_split2: "submatrix A I J = submatrix (submatrix A I UNIV) UNIV J" (is "?lhs = ?rhs") proof (rule eq_matI) show dr: "dim_row ?lhs = dim_row ?rhs" by (simp add: dim_submatrix(1)) show dc: "dim_col ?lhs = dim_col ?rhs" by (simp add: dim_submatrix(2)) fix i j assume i: "i < dim_row ?rhs" and j: "j < dim_col ?rhs" have "?rhs $$ (i, j) = (submatrix A I UNIV) $$ (pick UNIV i, pick J j)" proof (rule submatrix_index) show "i < card {i. i < dim_row (submatrix A I UNIV) \ i \ UNIV}" by (metis (full_types) dim_submatrix(1) i) show "j < card {j. j < dim_col (submatrix A I UNIV) \ j \ J}" by (metis (full_types) dim_submatrix(2) j) qed also have "... = A $$ (pick I (pick UNIV i), pick UNIV (pick J j))" proof (rule submatrix_index) show "pick UNIV i < card {i. i < dim_row A \ i \ I}" by (metis (full_types) dr dim_submatrix(1) i pick_UNIV) show "pick J j < card {j. j < dim_col A \ j \ UNIV}" by (metis (full_types) dim_submatrix(2) j pick_le) qed also have "... = ?lhs $$ (i,j)" proof (unfold pick_UNIV, rule submatrix_index[symmetric]) show "i < card {i. i < dim_row A \ i \ I}" by (metis (full_types) dim_submatrix(1) dr i) show "j < card {j. j < dim_col A \ j \ J}" by (metis (full_types) dim_submatrix(2) dc j) qed finally show "?lhs $$ (i, j) = ?rhs $$ (i, j)" .. qed lemma submatrix_mult: "submatrix (A*B) I J = submatrix A I UNIV * submatrix B UNIV J" (is "?lhs = ?rhs") proof (rule eq_matI) show "dim_row ?lhs = dim_row ?rhs" unfolding submatrix_def by auto show "dim_col ?lhs = dim_col ?rhs" unfolding submatrix_def by auto fix i j assume i: "i < dim_row ?rhs" and j: "j < dim_col ?rhs" have i1: "i < card {i. i < dim_row (A * B) \ i \ I}" by (metis (full_types) dim_submatrix(1) i index_mult_mat(2)) have j1: "j < card {j. j < dim_col (A * B) \ j \ J}" by (metis dim_submatrix(2) index_mult_mat(3) j) have pi: "pick I i < dim_row A" using i1 pick_le by auto have pj: "pick J j < dim_col B" using j1 pick_le by auto have row_rw: "Matrix.row (submatrix A I UNIV) i = Matrix.row A (pick I i)" using i1 row_submatrix_UNIV by auto have col_rw: "col (submatrix B UNIV J) j = col B (pick J j)" using j1 col_submatrix_UNIV by auto have "?lhs $$ (i,j) = (A*B) $$ (pick I i, pick J j)" by (rule submatrix_index[OF i1 j1]) also have "... = Matrix.row A (pick I i) \ col B (pick J j)" by (rule index_mult_mat(1)[OF pi pj]) also have "... = Matrix.row (submatrix A I UNIV) i \ col (submatrix B UNIV J) j" using row_rw col_rw by simp also have "... = (?rhs) $$ (i,j)" by (rule index_mult_mat[symmetric], insert i j, auto) finally show "?lhs $$ (i, j) = ?rhs $$ (i, j)" . qed lemma det_singleton: assumes A: "A \ carrier_mat 1 1" shows "det A = A $$ (0,0)" using A unfolding carrier_mat_def Determinant.det_def by auto lemma submatrix_singleton_index: assumes A: "A \ carrier_mat n m" and an: "a < n" and bm: "b < m" shows "submatrix A {a} {b} $$ (0,0) = A $$ (a,b)" proof - have a: "{i. i = a \ i < dim_row A} = {a}" using an A unfolding carrier_mat_def by auto have b: "{i. i = b \ i < dim_col A} = {b}" using bm A unfolding carrier_mat_def by auto have "submatrix A {a} {b} $$ (0,0) = A $$ (pick {a} 0,pick {b} 0)" by (rule submatrix_index, insert a b, auto) moreover have "pick {a} 0 = a" by (auto, metis (full_types) LeastI) moreover have "pick {b} 0 = b" by (auto, metis (full_types) LeastI) ultimately show ?thesis by simp qed end lemma det_not_inj_on: assumes not_inj_on: "\ inj_on f {0..r n n (\i. Matrix.row B (f i))) = 0" proof - obtain i j where i: "ij" using not_inj_on unfolding inj_on_def by auto show ?thesis proof (rule det_identical_rows[OF _ ij i j]) let ?B="(mat\<^sub>r n n (\i. row B (f i)))" show "row ?B i = row ?B j" proof (rule eq_vecI, auto) fix ia assume ia: "ia < n" have "row ?B i $ ia = ?B $$ (i, ia)" by (rule index_row(1), insert i ia, auto) also have "... = ?B $$ (j, ia)" by (simp add: fi_fj i ia j) also have "... = row ?B j $ ia" by (rule index_row(1)[symmetric], insert j ia, auto) finally show "row ?B i $ ia = row (mat\<^sub>r n n (\i. row B (f i))) j $ ia" by simp qed show "mat\<^sub>r n n (\i. Matrix.row B (f i)) \ carrier_mat n n" by auto qed qed lemma mat_row_transpose: "(mat\<^sub>r nr nc f)\<^sup>T = mat nc nr (\(i,j). vec_index (f j) i)" by (rule eq_matI, auto) lemma obtain_inverse_matrix: assumes A: "A \ carrier_mat n n" and i: "invertible_mat A" obtains B where "inverts_mat A B" and "inverts_mat B A" and "B \ carrier_mat n n" proof - have "(\B. inverts_mat A B \ inverts_mat B A)" using i unfolding invertible_mat_def by auto from this obtain B where AB: "inverts_mat A B" and BA: "inverts_mat B A" by auto moreover have "B \ carrier_mat n n" using A AB BA unfolding carrier_mat_def inverts_mat_def by (auto, metis index_mult_mat(3) index_one_mat(3))+ ultimately show ?thesis using that by blast qed lemma invertible_mat_smult_mat: fixes A :: "'a::comm_ring_1 mat" assumes inv_A: "invertible_mat A" and k: "k dvd 1" shows "invertible_mat (k \\<^sub>m A)" proof - obtain n where A: "A \ carrier_mat n n" using inv_A unfolding invertible_mat_def by auto have det_dvd_1: "Determinant.det A dvd 1" using inv_A invertible_iff_is_unit_JNF[OF A] by auto have "Determinant.det (k \\<^sub>m A) = k ^ dim_col A * Determinant.det A" by simp also have "... dvd 1" by (rule unit_prod, insert k det_dvd_1 dvd_power_same, force+) finally show ?thesis using invertible_iff_is_unit_JNF by (metis A smult_carrier_mat) qed lemma invertible_mat_one[simp]: "invertible_mat (1\<^sub>m n)" unfolding invertible_mat_def using inverts_mat_def by fastforce lemma four_block_mat_dim0: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n 0" and C: "C \ carrier_mat 0 n" and D: "D \ carrier_mat 0 0" shows "four_block_mat A B C D = A" unfolding four_block_mat_def using assms by auto lemma det_four_block_mat_lower_right_id: assumes A: "A \ carrier_mat m m" and B: "B = 0\<^sub>m m (n-m)" and C: "C = 0\<^sub>m (n-m) m" and D: "D = 1\<^sub>m (n-m)" and "n>m" shows "Determinant.det (four_block_mat A B C D) = Determinant.det A" using assms proof (induct n arbitrary: A B C D) case 0 then show ?case by auto next case (Suc n) let ?block = "(four_block_mat A B C D)" let ?B = "Matrix.mat m (n-m) (\(i,j). 0)" let ?C = "Matrix.mat (n-m) m (\(i,j). 0)" let ?D = "1\<^sub>m (n-m)" have mat_eq: "(mat_delete ?block n n) = four_block_mat A ?B ?C ?D" (is "?lhs = ?rhs") proof (rule eq_matI) fix i j assume i: "i < dim_row (four_block_mat A ?B ?C ?D)" and j: "j < dim_col (four_block_mat A ?B ?C ?D)" let ?f = " (if i < dim_row A then if j < dim_col A then A $$ (i, j) else B $$ (i, j - dim_col A) else if j < dim_col A then C $$ (i - dim_row A, j) else D $$ (i - dim_row A, j - dim_col A))" let ?g = "(if i < dim_row A then if j < dim_col A then A $$ (i, j) else ?B $$ (i, j - dim_col A) else if j < dim_col A then ?C $$ (i - dim_row A, j) else ?D $$ (i - dim_row A, j - dim_col A))" have "(mat_delete ?block n n) $$ (i,j) = ?block $$ (i,j)" using i j Suc.prems unfolding mat_delete_def by auto also have "... = ?f" by (rule index_mat_four_block, insert Suc.prems i j, auto) also have "... = ?g" using i j Suc.prems by auto also have "... = four_block_mat A ?B ?C ?D $$ (i,j)" by (rule index_mat_four_block[symmetric], insert Suc.prems i j, auto) finally show "?lhs $$ (i,j) = ?rhs $$ (i,j)" . qed (insert Suc.prems, auto) have nn_1: "?block $$ (n, n) = 1" using Suc.prems by auto have rw0: "(\i {..ii carrier_mat 1 n" and B: "B \ carrier_mat m n" and m0: "m \ 0" and r: "Matrix.row A 0 = Matrix.row B 0" shows "Matrix.row (A * V) 0 = Matrix.row (B * V) 0" proof (rule eq_vecI) show "dim_vec (Matrix.row (A * V) 0) = dim_vec (Matrix.row (B * V) 0)" using A B r by auto fix i assume i: "i < dim_vec (Matrix.row (B * V) 0)" have "Matrix.row (A * V) 0 $v i = (A * V) $$ (0,i)" by (rule index_row, insert i A, auto) also have "... = Matrix.row A 0 \ col V i" by (rule index_mult_mat, insert A i, auto) also have "... = Matrix.row B 0 \ col V i" using r by auto also have "... = (B * V) $$ (0,i)" by (rule index_mult_mat[symmetric], insert m0 B i, auto) also have "... = Matrix.row (B * V) 0 $v i" by (rule index_row[symmetric], insert i B m0, auto) finally show "Matrix.row (A * V) 0 $v i = Matrix.row (B * V) 0 $v i" . qed lemma smult_mat_mat_one_element: assumes A: "A \ carrier_mat 1 1" and B: "B \ carrier_mat 1 n" shows "A * B = A $$ (0,0) \\<^sub>m B" proof (rule eq_matI) fix i j assume i: "i < dim_row (A $$ (0, 0) \\<^sub>m B)" and j: "j < dim_col (A $$ (0, 0) \\<^sub>m B)" have i0: "i = 0" using A B i by auto have "(A * B) $$ (i, j) = Matrix.row A i \ col B j" by (rule index_mult_mat, insert i j A B, auto) also have "... = Matrix.row A i $v 0 * col B j $v 0" unfolding scalar_prod_def using B by auto also have "... = A$$(i,i) * B$$(i,j)" using A i i0 j by auto also have "... = (A $$ (i, i) \\<^sub>m B) $$ (i, j)" unfolding i by (rule index_smult_mat[symmetric], insert i j B, auto) finally show "(A * B) $$ (i, j) = (A $$ (0, 0) \\<^sub>m B) $$ (i, j)" using i0 by simp qed (insert A B, auto) lemma determinant_one_element: assumes A: "A \ carrier_mat 1 1" shows "Determinant.det A = A $$ (0,0)" proof - have "Determinant.det A = prod_list (diag_mat A)" by (rule det_upper_triangular[OF _ A], insert A, unfold upper_triangular_def, auto) also have "... = A $$ (0,0)" using A unfolding diag_mat_def by auto finally show ?thesis . qed lemma invertible_mat_transpose: assumes inv_A: "invertible_mat (A::'a::comm_ring_1 mat)" shows "invertible_mat A\<^sup>T" proof - obtain n where A: "A \ carrier_mat n n" using inv_A unfolding invertible_mat_def square_mat.simps by auto hence At: "A\<^sup>T \ carrier_mat n n" by simp have "Determinant.det A\<^sup>T = Determinant.det A" by (metis Determinant.det_def Determinant.det_transpose carrier_matI index_transpose_mat(2) index_transpose_mat(3)) also have "... dvd 1" using invertible_iff_is_unit_JNF[OF A] inv_A by simp finally show ?thesis using invertible_iff_is_unit_JNF[OF At] by auto qed lemma dvd_elements_mult_matrix_left: assumes A: "(A::'a::comm_ring_1 mat) \ carrier_mat m n" and P: "P \ carrier_mat m m" and x: "(\i j. i j x dvd A$$(i,j))" shows "(\i j. i j x dvd (P*A)$$(i,j))" proof - have "x dvd (P * A) $$ (i, j)" if i: "i < m" and j: "j < n" for i j proof - have "(P * A) $$ (i, j) = (\ia = 0..ia = 0.. carrier_mat m n" and Q: "Q \ carrier_mat n n" and x: "(\i j. i j x dvd A$$(i,j))" shows "(\i j. i j x dvd (A*Q)$$(i,j))" proof - have "x dvd (A*Q) $$ (i, j)" if i: "i < m" and j: "j < n" for i j proof - have "(A*Q) $$ (i, j) = (\ia = 0..ia = 0.. carrier_mat m n" and P: "P \ carrier_mat m m" and Q: "Q \ carrier_mat n n" and x: "(\i j. i j x dvd A$$(i,j))" shows "(\i j. i j x dvd (P*A*Q)$$(i,j))" using dvd_elements_mult_matrix_left[OF A P x] by (meson P A Q dvd_elements_mult_matrix_right mult_carrier_mat) definition append_cols :: "'a :: zero mat \ 'a mat \ 'a mat" (infixr "@\<^sub>c" 65)where "A @\<^sub>c B = four_block_mat A B (0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 (dim_col B))" lemma append_cols_carrier[simp,intro]: "A \ carrier_mat n a \ B \ carrier_mat n b \ (A @\<^sub>c B) \ carrier_mat n (a+b)" unfolding append_cols_def by auto lemma append_cols_mult_left: assumes A: "A \ carrier_mat n a" and B: "B \ carrier_mat n b" and P: "P \ carrier_mat n n" shows "P * (A @\<^sub>c B) = (P*A) @\<^sub>c (P*B)" proof - let ?P = "four_block_mat P (0\<^sub>m n 0) (0\<^sub>m 0 n) (0\<^sub>m 0 0)" have "P = ?P" by (rule eq_matI, auto) hence "P * (A @\<^sub>c B) = ?P * (A @\<^sub>c B)" by simp also have "?P * (A @\<^sub>c B) = four_block_mat (P * A + 0\<^sub>m n 0 * 0\<^sub>m 0 (dim_col A)) (P * B + 0\<^sub>m n 0 * 0\<^sub>m 0 (dim_col B)) (0\<^sub>m 0 n * A + 0\<^sub>m 0 0 * 0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 n * B + 0\<^sub>m 0 0 * 0\<^sub>m 0 (dim_col B))" unfolding append_cols_def by (rule mult_four_block_mat, insert A B P, auto) also have "... = four_block_mat (P * A) (P * B) (0\<^sub>m 0 (dim_col (P*A))) (0\<^sub>m 0 (dim_col (P*B)))" by (rule cong_four_block_mat, insert P, auto) also have "... = (P*A) @\<^sub>c (P*B)" unfolding append_cols_def by auto finally show ?thesis . qed lemma append_cols_mult_right_id: assumes A: "(A::'a::semiring_1 mat) \ carrier_mat n 1" and B: "B \ carrier_mat n (m-1)" and C: "C = four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) D" and D: "D \ carrier_mat (m-1) (m-1)" shows "(A @\<^sub>c B) * C = A @\<^sub>c (B * D)" proof - let ?C = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) D" have "(A @\<^sub>c B) * C = (A @\<^sub>c B) * ?C" unfolding C by auto also have "... = four_block_mat A B (0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 (dim_col B)) * ?C" unfolding append_cols_def by auto also have "... = four_block_mat (A * 1\<^sub>m 1 + B * 0\<^sub>m (m - 1) 1) (A * 0\<^sub>m 1 (m - 1) + B * D) (0\<^sub>m 0 (dim_col A) * 1\<^sub>m 1 + 0\<^sub>m 0 (dim_col B) * 0\<^sub>m (m - 1) 1) (0\<^sub>m 0 (dim_col A) * 0\<^sub>m 1 (m - 1) + 0\<^sub>m 0 (dim_col B) * D)" by (rule mult_four_block_mat, insert assms, auto) also have "... = four_block_mat A (B * D) (0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 (dim_col (B*D)))" by (rule cong_four_block_mat, insert assms, auto) also have "... = A @\<^sub>c (B * D)" unfolding append_cols_def by auto finally show ?thesis . qed lemma append_cols_mult_right_id2: assumes A: "(A::'a::semiring_1 mat) \ carrier_mat n a" and B: "B \ carrier_mat n b" and C: "C = four_block_mat D (0\<^sub>m a b) (0\<^sub>m b a) (1\<^sub>m b)" and D: "D \ carrier_mat a a" shows "(A @\<^sub>c B) * C = (A * D) @\<^sub>c B" proof - let ?C = "four_block_mat D (0\<^sub>m a b) (0\<^sub>m b a) (1\<^sub>m b)" have "(A @\<^sub>c B) * C = (A @\<^sub>c B) * ?C" unfolding C by auto also have "... = four_block_mat A B (0\<^sub>m 0 a) (0\<^sub>m 0 b) * ?C" unfolding append_cols_def using A B by auto also have "... = four_block_mat (A * D + B * 0\<^sub>m b a) (A * 0\<^sub>m a b + B * 1\<^sub>m b) (0\<^sub>m 0 a * D + 0\<^sub>m 0 b * 0\<^sub>m b a) (0\<^sub>m 0 a * 0\<^sub>m a b + 0\<^sub>m 0 b * 1\<^sub>m b)" by (rule mult_four_block_mat, insert A B C D, auto) also have "... = four_block_mat (A * D) B (0\<^sub>m 0 (dim_col (A*D))) (0\<^sub>m 0 (dim_col B))" by (rule cong_four_block_mat, insert assms, auto) also have "... = (A * D) @\<^sub>c B" unfolding append_cols_def by auto finally show ?thesis . qed lemma append_cols_nth: assumes A: "A \ carrier_mat n a" and B: "B \ carrier_mat n b" and i: "ic B) $$ (i, j) = (if j < dim_col A then A $$(i,j) else B$$(i,j-a))" (is "?lhs = ?rhs") proof - let ?C = "(0\<^sub>m 0 (dim_col A))" let ?D = "(0\<^sub>m 0 (dim_col B))" have i2: "i < dim_row A + dim_row ?D" using i A by auto have j2: "j < dim_col A + dim_col (0\<^sub>m 0 (dim_col B))" using j B A by auto have "(A @\<^sub>c B) $$ (i, j) = four_block_mat A B ?C ?D $$ (i, j)" unfolding append_cols_def by auto also have "... = (if i < dim_row A then if j < dim_col A then A $$ (i, j) else B $$ (i, j - dim_col A) else if j < dim_col A then ?C $$ (i - dim_row A, j) else 0\<^sub>m 0 (dim_col B) $$ (i - dim_row A, j - dim_col A))" by (rule index_mat_four_block(1)[OF i2 j2]) also have "... = ?rhs" using i A by auto finally show ?thesis . qed lemma append_cols_split: assumes d: "dim_col A > 0" shows "A = mat_of_cols (dim_row A) [col A 0] @\<^sub>c mat_of_cols (dim_row A) (map (col A) [1..c ?A2") proof (rule eq_matI) fix i j assume i: "i < dim_row (?A1 @\<^sub>c ?A2)" and j: "j < dim_col (?A1 @\<^sub>c ?A2)" have "(?A1 @\<^sub>c ?A2) $$ (i, j) = (if j < dim_col ?A1 then ?A1 $$(i,j) else ?A2$$(i,j-(dim_col ?A1)))" by (rule append_cols_nth, insert i j, auto simp add: append_cols_def) also have "... = A $$ (i,j)" proof (cases "j< dim_col ?A1") case True then show ?thesis by (metis One_nat_def Suc_eq_plus1 add.right_neutral append_cols_def col_def i index_mat_four_block(2) index_vec index_zero_mat(2) less_one list.size(3) list.size(4) mat_of_cols_Cons_index_0 mat_of_cols_carrier(2) mat_of_cols_carrier(3)) next case False then show ?thesis by (metis (no_types, lifting) Suc_eq_plus1 Suc_less_eq Suc_pred add_diff_cancel_right' append_cols_def diff_zero i index_col index_mat_four_block(2) index_mat_four_block(3) index_zero_mat(2) index_zero_mat(3) j length_map length_upt linordered_semidom_class.add_diff_inverse list.size(3) list.size(4) mat_of_cols_carrier(2) mat_of_cols_carrier(3) mat_of_cols_index nth_map_upt plus_1_eq_Suc upt_0) qed finally show "A $$ (i, j) = (?A1 @\<^sub>c ?A2) $$ (i, j)" .. qed (auto simp add: append_cols_def d) lemma append_rows_nth: assumes A: "A \ carrier_mat a n" and B: "B \ carrier_mat b n" and i: "ir B) $$ (i, j) = (if i < dim_row A then A $$(i,j) else B$$(i-a,j))" (is "?lhs = ?rhs") proof - let ?C = "(0\<^sub>m (dim_row A) 0)" let ?D = "(0\<^sub>m (dim_row B) 0)" have i2: "i < dim_row A + dim_row ?D" using i j A B by auto have j2: "j < dim_col A + dim_col ?D" using i j A B by auto have "(A @\<^sub>r B) $$ (i, j) = four_block_mat A ?C B ?D $$ (i, j)" unfolding append_rows_def by auto also have "... = (if i < dim_row A then if j < dim_col A then A $$ (i, j) else ?C $$ (i, j - dim_col A) else if j < dim_col A then B $$ (i - dim_row A, j) else ?D $$ (i - dim_row A, j - dim_col A))" by (rule index_mat_four_block(1)[OF i2 j2]) also have "... = ?rhs" using i A j B by auto finally show ?thesis . qed lemma append_rows_split: assumes k: "k\dim_row A" shows "A = (mat_of_rows (dim_col A) [Matrix.row A i. i \ [0..r (mat_of_rows (dim_col A) [Matrix.row A i. i \ [k..r ?A2") proof (rule eq_matI) have "(?A1 @\<^sub>r ?A2) \ carrier_mat (k + (dim_row A-k)) (dim_col A)" by (rule carrier_append_rows, insert k, auto) hence A1_A2: "(?A1 @\<^sub>r ?A2) \ carrier_mat (dim_row A) (dim_col A)" using k by simp thus "dim_row A = dim_row (?A1 @\<^sub>r ?A2)" and "dim_col A = dim_col (?A1 @\<^sub>r ?A2)" by auto fix i j assume i: "i < dim_row (?A1 @\<^sub>r ?A2)" and j: "j < dim_col (?A1 @\<^sub>r ?A2)" have "(?A1 @\<^sub>r ?A2) $$ (i, j) = (if i < dim_row ?A1 then ?A1 $$(i,j) else ?A2$$(i-(dim_row ?A1),j))" by (rule append_rows_nth, insert k i j, auto simp add: append_rows_def) also have "... = A $$ (i,j)" proof (cases "ir ?A2) $$ (i,j)" by simp qed lemma transpose_mat_append_rows: assumes A: "A \ carrier_mat a n" and B: "B \ carrier_mat b n" shows "(A @\<^sub>r B)\<^sup>T = A\<^sup>T @\<^sub>c B\<^sup>T" by (smt append_cols_def append_rows_def A B carrier_matD(1) index_transpose_mat(3) transpose_four_block_mat zero_carrier_mat zero_transpose_mat) lemma transpose_mat_append_cols: assumes A: "A \ carrier_mat n a" and B: "B \ carrier_mat n b" shows "(A @\<^sub>c B)\<^sup>T = A\<^sup>T @\<^sub>r B\<^sup>T" by (metis Matrix.transpose_transpose A B carrier_matD(1) carrier_mat_triv index_transpose_mat(3) transpose_mat_append_rows) lemma append_rows_mult_right: assumes A: "(A::'a::comm_semiring_1 mat) \ carrier_mat a n" and B: "B \ carrier_mat b n" and Q: "Q\ carrier_mat n n" shows "(A @\<^sub>r B) * Q = (A * Q) @\<^sub>r (B*Q)" proof - have "transpose_mat ((A @\<^sub>r B) * Q) = Q\<^sup>T * (A @\<^sub>r B)\<^sup>T" by (rule transpose_mult, insert A B Q, auto) also have "... = Q\<^sup>T * (A\<^sup>T @\<^sub>c B\<^sup>T)" using transpose_mat_append_rows assms by metis also have "... = Q\<^sup>T * A\<^sup>T @\<^sub>c Q\<^sup>T * B\<^sup>T" using append_cols_mult_left assms by (metis transpose_carrier_mat) also have "transpose_mat ... = (A * Q) @\<^sub>r (B*Q)" by (smt A B Matrix.transpose_mult Matrix.transpose_transpose append_cols_def append_rows_def Q carrier_mat_triv index_mult_mat(2) index_transpose_mat(2) transpose_four_block_mat zero_carrier_mat zero_transpose_mat) finally show ?thesis by simp qed lemma append_rows_mult_left_id: assumes A: "(A::'a::comm_semiring_1 mat) \ carrier_mat 1 n" and B: "B \ carrier_mat (m-1) n" and C: "C = four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) D" and D: "D \ carrier_mat (m-1) (m-1)" shows "C * (A @\<^sub>r B) = A @\<^sub>r (D * B)" proof - have "transpose_mat (C * (A @\<^sub>r B)) = (A @\<^sub>r B)\<^sup>T * C\<^sup>T" by (metis (no_types, lifting) B C D Matrix.transpose_mult append_rows_def A carrier_matD carrier_mat_triv index_mat_four_block(2,3) index_zero_mat(2) one_carrier_mat) also have "... = (A\<^sup>T @\<^sub>c B\<^sup>T) * C\<^sup>T" using transpose_mat_append_rows[OF A B] by auto also have "... = A\<^sup>T @\<^sub>c (B\<^sup>T * D\<^sup>T)" by (rule append_cols_mult_right_id, insert A B C D, auto) also have "transpose_mat ... = A @\<^sub>r (D * B)" by (smt B D Matrix.transpose_mult Matrix.transpose_transpose append_cols_def append_rows_def A carrier_matD(2) carrier_mat_triv index_mult_mat(3) index_transpose_mat(3) transpose_four_block_mat zero_carrier_mat zero_transpose_mat) finally show ?thesis by auto qed lemma append_rows_mult_left_id2: assumes A: "(A::'a::comm_semiring_1 mat) \ carrier_mat a n" and B: "B \ carrier_mat b n" and C: "C = four_block_mat D (0\<^sub>m a b) (0\<^sub>m b a) (1\<^sub>m b)" and D: "D \ carrier_mat a a" shows "C * (A @\<^sub>r B) = (D * A) @\<^sub>r B" proof - have "(C * (A @\<^sub>r B))\<^sup>T = (A @\<^sub>r B)\<^sup>T * C\<^sup>T" by (rule transpose_mult, insert assms, auto) also have "... = (A\<^sup>T @\<^sub>c B\<^sup>T) * C\<^sup>T" by (metis A B transpose_mat_append_rows) also have "... = (A\<^sup>T * D\<^sup>T @\<^sub>c B\<^sup>T)" by (rule append_cols_mult_right_id2, insert assms, auto) also have "...\<^sup>T = (D * A) @\<^sub>r B" by (metis A B D transpose_mult transpose_transpose mult_carrier_mat transpose_mat_append_rows) finally show ?thesis by simp qed lemma four_block_mat_preserves_column: assumes A: "(A::'a::semiring_1 mat) \ carrier_mat n m" and B: "B = four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) C" and C: "C \ carrier_mat (m-1) (m-1)" and i: "ic ?A2" by (rule append_cols_split[of A, unfolded n2], insert m A, auto) hence "A * B = (?A1 @\<^sub>c ?A2) * B" by simp also have "... = ?A1 @\<^sub>c (?A2 * C)" by (rule append_cols_mult_right_id[OF _ _ B C], insert A, auto) also have "... $$ (i,0) = ?A1 $$ (i,0)" using append_cols_nth by (simp add: append_cols_def i) also have "... = A $$ (i,0)" by (metis A i carrier_matD(1) col_def index_vec mat_of_cols_Cons_index_0) finally show ?thesis . qed definition "lower_triangular A = (\i j. i < j \ i < dim_row A \ j < dim_col A \ A $$ (i,j) = 0)" lemma lower_triangular_index: assumes "lower_triangular A" "i carrier_mat n n" shows "A * (k \\<^sub>m (1\<^sub>m n)) = (k \\<^sub>m (1\<^sub>m n)) * A" proof - have "(\ia = 0..ia = 0..ia \ ({0..ia \ ({0..ia \ ({0..ia \ ({0.. carrier_mat 2 2" shows "Determinant.det A = A$$(0,0) * A $$ (1,1) - A$$(0,1)*A$$(1,0)" proof - let ?A = "(Mod_Type_Connect.to_hma\<^sub>m A)::'a^2^2" have [transfer_rule]: "Mod_Type_Connect.HMA_M A ?A" unfolding Mod_Type_Connect.HMA_M_def using from_hma_to_hma\<^sub>m A by auto have [transfer_rule]: "Mod_Type_Connect.HMA_I 0 0" unfolding Mod_Type_Connect.HMA_I_def by (simp add: to_nat_0) have [transfer_rule]: "Mod_Type_Connect.HMA_I 1 1" unfolding Mod_Type_Connect.HMA_I_def by (simp add: to_nat_1) have "Determinant.det A = Determinants.det ?A" by (transfer, simp) also have "... = ?A $h 1 $h 1 * ?A $h 2 $h 2 - ?A $h 1 $h 2 * ?A $h 2 $h 1" unfolding det_2 by simp also have "... = ?A $h 0 $h 0 * ?A $h 1 $h 1 - ?A $h 0 $h 1 * ?A $h 1 $h 0" by (smt Groups.mult_ac(2) exhaust_2 semiring_norm(160)) also have "... = A$$(0,0) * A $$ (1,1) - A$$(0,1)*A$$(1,0)" unfolding index_hma_def[symmetric] by (transfer, auto) finally show ?thesis . qed lemma mat_diag_smult: "mat_diag n (\ x. (k::'a::comm_ring_1)) = (k \\<^sub>m 1\<^sub>m n)" proof - have "mat_diag n (\ x. k) = mat_diag n (\ x. k * 1)" by auto also have "... = mat_diag n (\ x. k) * mat_diag n (\ x. 1)" using mat_diag_diag by (simp add: mat_diag_def) also have "... = mat_diag n (\ x. k) * (1\<^sub>m n)" by auto thm mat_diag_mult_left also have "... = Matrix.mat n n (\(i, j). k * (1\<^sub>m n) $$ (i, j))" by (rule mat_diag_mult_left, auto) also have "... = (k \\<^sub>m 1\<^sub>m n)" unfolding smult_mat_def by auto finally show ?thesis . qed lemma invertible_mat_four_block_mat_lower_right: assumes A: "(A::'a::comm_ring_1 mat) \ carrier_mat n n" and inv_A: "invertible_mat A" shows "invertible_mat (four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 n) (0\<^sub>m n 1) A)" proof - let ?I = "(four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 n) (0\<^sub>m n 1) A)" have "Determinant.det ?I = Determinant.det (1\<^sub>m 1) * Determinant.det A" by (rule det_four_block_mat_lower_left_zero_col, insert assms, auto) also have "... = Determinant.det A" by auto finally have "Determinant.det ?I = Determinant.det A" . thus ?thesis by (metis (no_types, lifting) assms carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mat_four_block(2) index_mat_four_block(3) index_one_mat(2) index_one_mat(3) invertible_iff_is_unit_JNF) qed lemma invertible_mat_four_block_mat_lower_right_id: assumes A: "(A::'a::comm_ring_1 mat) \ carrier_mat m m" and B: "B = 0\<^sub>m m (n-m)" and C: "C = 0\<^sub>m (n-m) m" and D: "D = 1\<^sub>m (n-m)" and "n>m" and inv_A: "invertible_mat A" shows "invertible_mat (four_block_mat A B C D)" proof - have "Determinant.det (four_block_mat A B C D) = Determinant.det A" by (rule det_four_block_mat_lower_right_id, insert assms, auto) thus ?thesis using inv_A by (metis (no_types, lifting) assms(1) assms(4) carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mat_four_block(2) index_mat_four_block(3) index_one_mat(2) index_one_mat(3) invertible_iff_is_unit_JNF) qed lemma split_block4_decreases_dim_row: assumes E: "(A,B,C,D) = split_block E 1 1" and E1: "dim_row E > 1" and E2: "dim_col E > 1" shows "dim_row D < dim_row E" proof - have "D \ carrier_mat (1 + (dim_row E - 2)) (1 + (dim_col E - 2))" by (rule split_block(4)[OF E[symmetric]], insert E1 E2, auto) hence "D \ carrier_mat (dim_row E - 1) (dim_col E - 1)" using E1 E2 by auto thus ?thesis using E1 by auto qed lemma inv_P'PAQQ': assumes A: "A \ carrier_mat n n" and P: "P \ carrier_mat n n" and inv_P: "inverts_mat P' P" and inv_Q: "inverts_mat Q Q'" and Q: "Q \ carrier_mat n n" and P': "P' \ carrier_mat n n" and Q': "Q' \ carrier_mat n n" shows "(P'*(P*A*Q)*Q') = A" proof - have "(P'*(P*A*Q)*Q') = (P'*(P*A*Q*Q'))" by (smt P P' Q Q' assoc_mult_mat carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mult_mat(2) index_mult_mat(3)) also have "... = ((P'*P)*A*(Q*Q'))" by (smt A P P' Q Q' assoc_mult_mat carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mult_mat(3) inv_Q inverts_mat_def right_mult_one_mat') finally show ?thesis by (metis P' Q A inv_P inv_Q carrier_matD(1) inverts_mat_def left_mult_one_mat right_mult_one_mat) qed lemma assumes "U \ carrier_mat 2 2" and "V \ carrier_mat 2 2" and "A = U * V" shows mat_mult2_00: "A $$ (0,0) = U $$ (0,0)*V $$ (0,0) + U $$ (0,1)*V $$ (1,0)" and mat_mult2_01: "A $$ (0,1) = U $$ (0,0)*V $$ (0,1) + U $$ (0,1)*V $$ (1,1)" and mat_mult2_10: "A $$ (1,0) = U $$ (1,0)*V $$ (0,0) + U $$ (1,1)*V $$ (1,0)" and mat_mult2_11: "A $$ (1,1) = U $$ (1,0)*V $$ (0,1) + U $$ (1,1)*V $$ (1,1)" using assms unfolding times_mat_def Matrix.row_def col_def scalar_prod_def using sum_two_rw by auto subsection\Lemmas about @{text "sorted lists"}, @{text "insort"} and @{text "pick"}\ lemma sorted_distinct_imp_sorted_wrt: assumes "sorted xs" and "distinct xs" shows "sorted_wrt (<) xs" using assms by (induct xs, insert le_neq_trans, auto) lemma sorted_map_strict: - assumes "strict_mono_on g {0.. g ` {0..x\set (map g [0.. g n" using sg unfolding strict_mono_on_def by (simp add: less_imp_le) qed finally show ?case . qed lemma sorted_nth_strict_mono: "sorted xs \ distinct xs \i < j \ j < length xs \ xs!i < xs!j" by (simp add: less_le nth_eq_iff_index_eq sorted_iff_nth_mono_less) lemma sorted_list_of_set_0_LEAST: assumes finI: "finite I" and I: "I \ {}" shows "sorted_list_of_set I ! 0 = (LEAST n. n\I)" proof (rule Least_equality[symmetric]) show "sorted_list_of_set I ! 0 \ I" by (metis I Max_in finI gr_zeroI in_set_conv_nth not_less_zero set_sorted_list_of_set) fix y assume "y \ I" thus "sorted_list_of_set I ! 0 \ y" by (metis eq_iff finI in_set_conv_nth neq0_conv sorted_iff_nth_mono_less sorted_list_of_set(1) sorted_sorted_list_of_set) qed lemma sorted_list_of_set_eq_pick: assumes i: "i < length (sorted_list_of_set I)" shows "sorted_list_of_set I ! i = pick I i" proof - have finI: "finite I" proof (rule ccontr) assume "infinite I" hence "length (sorted_list_of_set I) = 0" by auto thus False using i by simp qed show ?thesis using i proof (induct i) case 0 have I: "I \ {}" using "0.prems" sorted_list_of_set_empty by blast show ?case unfolding pick.simps by (rule sorted_list_of_set_0_LEAST[OF finI I]) next case (Suc i) note x_less = Suc.prems show ?case proof (unfold pick.simps, rule Least_equality[symmetric], rule conjI) show 1: "pick I i < sorted_list_of_set I ! Suc i" by (metis Suc.hyps Suc.prems Suc_lessD distinct_sorted_list_of_set find_first_unique lessI nat_less_le sorted_sorted_list_of_set sorted_wrt_nth_less) show "sorted_list_of_set I ! Suc i \ I" using Suc.prems finI nth_mem set_sorted_list_of_set by blast have rw: "sorted_list_of_set I ! i = pick I i" using Suc.hyps Suc_lessD x_less by blast have sorted_less: "sorted_list_of_set I ! i < sorted_list_of_set I ! Suc i" by (simp add: 1 rw) fix y assume y: "y \ I \ pick I i < y" show "sorted_list_of_set I ! Suc i \ y" by (smt antisym_conv finI in_set_conv_nth less_Suc_eq less_Suc_eq_le nat_neq_iff rw sorted_iff_nth_mono_less sorted_list_of_set(1) sorted_sorted_list_of_set x_less y) qed qed qed text\$b$ is the position where we add, $a$ the element to be added and $i$ the position that is checked\ lemma insort_nth': assumes "\j set xs" and "i < length xs + 1" and "i < b" and "xs \ []" and "b < length xs" shows "insort a xs ! i = xs ! i" using assms proof (induct xs arbitrary: a b i) case Nil then show ?case by auto next case (Cons x xs) note less = Cons.prems(1) note sorted = Cons.prems(2) note a_notin = Cons.prems(3) note i_length = Cons.prems(4) note i_b = Cons.prems(5) note b_length = Cons.prems(7) show ?case proof (cases "a \ x") case True have "insort a (x # xs) ! i = (a # x # xs) ! i" using True by simp also have "... = (x # xs) ! i" using Cons.prems(1) Cons.prems(5) True by force finally show ?thesis . next case False note x_less_a = False have "insort a (x # xs) ! i = (x # insort a xs) ! i" using False by simp also have "... = (x # xs) ! i" proof (cases "i = 0") case True then show ?thesis by auto next case False have "(x # insort a xs) ! i = (insort a xs) ! (i-1)" by (simp add: False nth_Cons') also have "... = xs ! (i-1)" proof (rule Cons.hyps) show "sorted xs" using sorted by simp show "a \ set xs" using a_notin by simp show "i - 1 < length xs + 1" using i_length False by auto show "xs \ []" using i_b b_length by force show "i - 1 < b - 1" by (simp add: False diff_less_mono i_b leI) show "b - 1 < length xs" using b_length i_b by auto show "\j set xs" and "i < index (insort a xs) a" and "xs \ []" shows "insort a xs ! i = xs ! i" using assms proof (induct xs arbitrary: a i) case Nil then show ?case by auto next case (Cons x xs) note sorted = Cons.prems(1) note a_notin = Cons.prems(2) note i_index = Cons.prems(3) show ?case proof (cases "a \ x") case True have "insort a (x # xs) ! i = (a # x # xs) ! i" using True by simp also have "... = (x # xs) ! i" using Cons.prems(1) Cons.prems(3) True by force finally show ?thesis . next case False note x_less_a = False show ?thesis proof (cases "xs = []") case True have "x \ a" using False by auto then show ?thesis using True i_index False by auto next case False note xs_not_empty = False have "insort a (x # xs) ! i = (x # insort a xs) ! i" using x_less_a by simp also have "... = (x # xs) ! i" proof (cases "i = 0") case True then show ?thesis by auto next case False note i0 = False have "(x # insort a xs) ! i = (insort a xs) ! (i-1)" by (simp add: False nth_Cons') also have "... = xs ! (i-1)" proof (rule Cons.hyps[OF _ _ _ xs_not_empty]) show "sorted xs" using sorted by simp show "a \ set xs" using a_notin by simp have "index (insort a (x # xs)) a = index ((x # insort a xs)) a" using x_less_a by auto also have "... = index (insort a xs) a + 1" unfolding index_Cons using x_less_a by simp finally show "i - 1 < index (insort a xs) a" using False i_index by linarith qed also have "... = (x # xs) ! i" by (simp add: False nth_Cons') finally show ?thesis . qed finally show ?thesis . qed qed qed lemma insort_nth2: assumes "sorted xs" and "a \ set xs" and "i < length xs" and "i \ index (insort a xs) a" and "xs \ []" shows "insort a xs ! (Suc i) = xs ! i" using assms proof (induct xs arbitrary: a i) case Nil then show ?case by auto next case (Cons x xs) note sorted = Cons.prems(1) note a_notin = Cons.prems(2) note i_length = Cons.prems(3) note index_i = Cons.prems(4) show ?case proof (cases "a \ x") case True have "insort a (x # xs) ! (Suc i) = (a # x # xs) ! (Suc i)" using True by simp also have "... = (x # xs) ! i" using Cons.prems(1) Cons.prems(5) True by force finally show ?thesis . next case False note x_less_a = False have "insort a (x # xs) ! (Suc i) = (x # insort a xs) ! (Suc i)" using False by simp also have "... = (x # xs) ! i" proof (cases "i = 0") case True then show ?thesis using index_i linear x_less_a by fastforce next case False note i0 = False show ?thesis proof - have Suc_i: "Suc (i - 1) = i" using i0 by auto have "(x # insort a xs) ! (Suc i) = (insort a xs) ! i" by (simp add: nth_Cons') also have "... = (insort a xs) ! Suc (i - 1)" using Suc_i by simp also have "... = xs ! (i - 1)" proof (rule Cons.hyps) show "sorted xs" using sorted by simp show "a \ set xs" using a_notin by simp show "i - 1 < length xs" using i_length using Suc_i by auto thus "xs \ []" by auto have "index (insort a (x # xs)) a = index ((x # insort a xs)) a" using x_less_a by simp also have "... = index (insort a xs) a + 1" unfolding index_Cons using x_less_a by simp finally show "index (insort a xs) a \ i - 1" using index_i i0 by auto qed also have "... = (x # xs) ! i" using Suc_i by auto finally show ?thesis . qed qed finally show ?thesis . qed qed lemma pick_index: assumes a: "a \ I" and a'_card: "a' < card I" shows "(pick I a' = a) = (index (sorted_list_of_set I) a = a')" proof - have finI: "finite I" using a'_card card.infinite by force have length_I: "length (sorted_list_of_set I) = card I" by (metis a'_card card.infinite distinct_card distinct_sorted_list_of_set not_less_zero set_sorted_list_of_set) let ?i = "index (sorted_list_of_set I) a" have "(sorted_list_of_set I) ! a' = pick I a'" by (rule sorted_list_of_set_eq_pick, auto simp add: finI a'_card length_I) moreover have "(sorted_list_of_set I) ! ?i = a" by (rule nth_index, simp add: a finI) ultimately show ?thesis by (metis a'_card distinct_sorted_list_of_set index_nth_id length_I) qed end diff --git a/thys/Smith_Normal_Form/SNF_Uniqueness.thy b/thys/Smith_Normal_Form/SNF_Uniqueness.thy --- a/thys/Smith_Normal_Form/SNF_Uniqueness.thy +++ b/thys/Smith_Normal_Form/SNF_Uniqueness.thy @@ -1,1180 +1,1180 @@ (* Author: Jose Divasón Email: jose.divason@unirioja.es *) section \Uniqueness of the Smith normal form\ theory SNF_Uniqueness imports Cauchy_Binet Smith_Normal_Form_JNF Admits_SNF_From_Diagonal_Iff_Bezout_Ring begin lemma dvd_associated1: fixes a::"'a::comm_ring_1" assumes "\u. u dvd 1 \ a = u*b" shows "a dvd b \ b dvd a" using assms by auto text \This is a key lemma. It demands the type class to be an integral domain. This means that the uniqueness result will be obtained for GCD domains, instead of rings.\ lemma dvd_associated2: fixes a::"'a::idom" assumes ab: "a dvd b" and ba: "b dvd a" and a: "a\0" shows "\u. u dvd 1 \ a = u*b" proof - obtain k where a_kb: "a = k*b" using ab unfolding dvd_def by (metis Groups.mult_ac(2) ba dvdE) obtain q where b_qa: "b = q*a" using ba unfolding dvd_def by (metis Groups.mult_ac(2) ab dvdE) have 1: "a = k*q*a" using a_kb b_qa by auto hence "k*q = 1" using a by simp thus ?thesis using 1 by (metis a_kb dvd_triv_left) qed corollary dvd_associated: fixes a::"'a::idom" assumes "a\0" shows "(a dvd b \ b dvd a) = (\u. u dvd 1 \ a = u*b)" using assms dvd_associated1 dvd_associated2 by metis lemma exists_inj_ge_index: assumes S: "S \ {0..f. inj_on f {0.. f`{0.. (\i\{0.. f i)" proof - have "\h. bij_betw h {0..i\{0.. ?f i" proof fix i assume i: "i \ {0.. ?xs ! i" proof (rule sorted_wrt_less_idx, rule sorted_distinct_imp_sorted_wrt) show "sorted ?xs" using sorted_sorted_list_of_set by blast show "distinct ?xs" using distinct_sorted_list_of_set by blast show "i < length ?xs" by (metis S Sk atLeast0LessThan distinct_card distinct_sorted_list_of_set gk_S i lessThan_iff set_sorted_list_of_set subset_eq_atLeast0_lessThan_finite) qed ultimately show "i \ ?f i" by auto qed show ?thesis using 1 2 3 by auto qed subsection \More specific results about submatrices\ lemma diagonal_imp_submatrix0: assumes dA: "diagonal_mat A" and A_carrier: "A\ carrier_mat n m" and Ik: "card I = k" and Jk: "card J = k" and r: "\row_index \ I. row_index < n" (*I \ {0..col_index \ J. col_index < m" and a: "a submatrix A I J $$ (a,b) = A $$(pick I a, pick I a)" proof (cases "submatrix A I J $$ (a, b) = 0") case True then show ?thesis by auto next case False note not0 = False have aux: "submatrix A I J $$ (a, b) = A $$(pick I a, pick J b)" proof (rule submatrix_index) have "card {i. i < dim_row A \ i \ I} = k" by (smt A_carrier Ik carrier_matD(1) equalityI mem_Collect_eq r subsetI) moreover have "card {i. i < dim_col A \ i \ J} = k" by (metis (no_types, lifting) A_carrier Jk c carrier_matD(2) carrier_mat_def equalityI mem_Collect_eq subsetI) ultimately show " a < card {i. i < dim_row A \ i \ I}" and "b < card {i. i < dim_col A \ i \ J}" using a b by auto qed thus ?thesis proof (cases "pick I a = pick J b") case True then show ?thesis using aux by auto next case False then show ?thesis by (metis aux A_carrier Ik Jk a b c carrier_matD dA diagonal_mat_def pick_in_set_le r) qed qed lemma diagonal_imp_submatrix_element_not0: assumes dA: "diagonal_mat A" and A_carrier: "A \ carrier_mat n m" and Ik: "card I = k" and Jk: "card J = k" and I: "I \ {0.. {0..i. i submatrix A I J $$ (i, b) \ 0" shows "\!i. i submatrix A I J $$ (i, b) \ 0" proof - have I_eq: "I = {i. i < dim_row A \ i \ I}" using I A_carrier unfolding carrier_mat_def by auto have J_eq: "J = {i. i < dim_col A \ i \ J}" using J A_carrier unfolding carrier_mat_def by auto obtain a where sub_ab: "submatrix A I J $$ (a, b) \ 0" and ak: "a < k" using ex_not0 by auto moreover have "i = a" if sub_ib: "submatrix A I J $$ (i, b) \ 0" and ik: "i < k" for i proof - have 1: "pick I i < dim_row A" using I_eq Ik ik pick_in_set_le by auto have 2: "pick J b < dim_col A" using J_eq Jk b pick_le by auto have 3: "pick I a < dim_row A" using I_eq Ik calculation(2) pick_le by auto have "submatrix A I J $$ (i, b) = A $$ (pick I i, pick J b)" by (rule submatrix_index, insert I_eq Ik ik J_eq Jk b, auto) hence pick_Ii_Jb: "pick I i = pick J b" using dA sub_ib 1 2 unfolding diagonal_mat_def by auto have "submatrix A I J $$ (a, b) = A $$ (pick I a, pick J b)" by (rule submatrix_index, insert I_eq Ik ak J_eq Jk b, auto) hence pick_Ia_Jb: "pick I a = pick J b" using dA sub_ab 3 2 unfolding diagonal_mat_def by auto have pick_Ia_Ii: "pick I a = pick I i" using pick_Ii_Jb pick_Ia_Jb by simp thus ?thesis by (metis Ik ak ik nat_neq_iff pick_mono_le) qed ultimately show ?thesis by auto qed lemma submatrix_index_exists: assumes A_carrier: "A\ carrier_mat n m" and Ik: "card I = k" and Jk: "card J = k" and a: "a \ I" and b: "b \ J" and k: "k > 0" and I: "I \ {0.. {0..a' b'. a' < k \ b' < k \ submatrix A I J $$ (a',b') = A $$ (a,b) \ a = pick I a' \ b = pick J b'" proof - let ?xs = "sorted_list_of_set I" let ?ys = "sorted_list_of_set J" have finI: "finite I" and finJ: "finite J" using k Ik Jk card_ge_0_finite by metis+ have set_xs: "set ?xs = I" by (rule set_sorted_list_of_set[OF finI]) have set_ys: "set ?ys = J" by (rule set_sorted_list_of_set[OF finJ]) have a_in_xs: "a \ set ?xs" and b_in_ys: "b \ set ?ys" using set_xs a set_ys b by auto have length_xs: "length ?xs = k" by (metis Ik distinct_card set_xs sorted_list_of_set(3)) have length_ys: "length ?ys = k" by (metis Jk distinct_card set_ys sorted_list_of_set(3)) obtain a' where a': "?xs ! a' = a" and a'_length: "a' < length ?xs" by (meson a_in_xs in_set_conv_nth) obtain b' where b': "?ys ! b' = b" and b'_length: "b' < length ?ys" by (meson b_in_ys in_set_conv_nth) have pick_a: "a = pick I a'" using a' a'_length finI sorted_list_of_set_eq_pick by auto have pick_b: "b = pick J b'" using b' b'_length finJ sorted_list_of_set_eq_pick by auto have I_rw: "I = {i. i < dim_row A \ i \ I}" and J_rw: "J = {i. i < dim_col A \ i \ J}" using I A_carrier J by auto have a'k: "a' < k" using a'_length length_xs by auto moreover have b'k: "b' carrier_mat n m" and Ik: "card I = k" and Jk: "card J = k" and I: "I \ {0.. {0.. I" and b_notin_J: "b \ J" and a'k: "a' < Suc k" and b'k: "b' < Suc k" and a_def: "pick (insert a I) a' = a" and b_def: "pick (insert b J) b' = b" shows "mat_delete (submatrix A (insert a I) (insert b J)) a' b' = submatrix A I J" (is "?lhs = ?rhs") proof (rule eq_matI) have I_eq: "I = {i. i < dim_row A \ i \ I}" using I A_carrier unfolding carrier_mat_def by auto have J_eq: "J = {i. i < dim_col A \ i \ J}" using J A_carrier unfolding carrier_mat_def by auto have insert_I_eq: "insert a I = {i. i < dim_row A \ i \ insert a I}" using I A_carrier a k unfolding carrier_mat_def by auto have card_Suc_k: "card {i. i < dim_row A \ i \ insert a I} = Suc k" using insert_I_eq Ik a_notin_I by (metis I card_insert_disjoint finite_atLeastLessThan finite_subset) have insert_J_eq: "insert b J = {i. i < dim_col A \ i \ insert b J}" using J A_carrier b k unfolding carrier_mat_def by auto have card_Suc_k': "card {i. i < dim_col A \ i \ insert b J} = Suc k" using insert_J_eq Jk b_notin_J by (metis J card_insert_disjoint finite_atLeastLessThan finite_subset) show "dim_row ?lhs = dim_row ?rhs" unfolding mat_delete_dim unfolding dim_submatrix using card_Suc_k I_eq Ik by auto show "dim_col ?lhs = dim_col ?rhs" unfolding mat_delete_dim unfolding dim_submatrix using card_Suc_k' J_eq Jk by auto fix i j assume i: "i < dim_row (submatrix A I J)" and j: "j < dim_col (submatrix A I J)" have ik: "i < k" by (metis I_eq Ik dim_submatrix(1) i) have jk: "j < k" by (metis J_eq Jk dim_submatrix(2) j) show "?lhs $$ (i, j) = ?rhs $$ (i, j)" proof - have index_eq1: "pick (insert a I) (insert_index a' i) = pick I i" by (rule pick_insert_index[OF Ik a_notin_I ik a_def], simp add: Ik a'k) have index_eq2: "pick (insert b J) (insert_index b' j) = pick J j" by (rule pick_insert_index[OF Jk b_notin_J jk b_def], simp add: Jk b'k) have "?lhs $$ (i,j) = (submatrix A (insert a I) (insert b J)) $$ (insert_index a' i, insert_index b' j)" proof (rule mat_delete_index[symmetric, OF _ a'k b'k ik jk]) show "submatrix A (insert a I) (insert b J) \ carrier_mat (Suc k) (Suc k)" by (metis card_Suc_k card_Suc_k' carrier_matI dim_submatrix(1) dim_submatrix(2)) qed also have "... = A $$ (pick (insert a I) (insert_index a' i), pick (insert b J) (insert_index b' j))" proof (rule submatrix_index) show "insert_index a' i < card {i. i < dim_row A \ i \ insert a I}" using card_Suc_k ik insert_index_def by auto show "insert_index b' j < card {j. j < dim_col A \ j \ insert b J}" using card_Suc_k' insert_index_def jk by auto qed also have "... = A $$ (pick I i, pick J j)" unfolding index_eq1 index_eq2 by auto also have "... = submatrix A I J $$ (i,j)" by (rule submatrix_index[symmetric], insert ik I_eq Ik Jk J_eq jk, auto) finally show ?thesis . qed qed subsection \On the minors of a diagonal matrix\ lemma det_minors_diagonal: assumes dA: "diagonal_mat A" and A_carrier: "A \ carrier_mat n m" and Ik: "card I = k" and Jk: "card J = k" and r: "I \ {0.. {0..0" shows "det (submatrix A I J) = 0 \ (\xs. (det (submatrix A I J) = prod_list xs \ det (submatrix A I J) = - prod_list xs) \ set xs \ {A$$(i,i)|i. i A$$(i,i)\ 0} \ length xs = k)" using Ik Jk r c k proof (induct k arbitrary: I J) case 0 then show ?case by auto next case (Suc k) note cardI = Suc.prems(1) note cardJ = Suc.prems(2) note I = Suc.prems(3) note J = Suc.prems(4) have *: "{i. i < dim_row A \ i \ I} = I" using I Ik A_carrier carrier_mat_def by auto have **: "{j. j < dim_col A \ j \ J} = J" using J Jk A_carrier carrier_mat_def by auto show ?case proof (cases "k = 0") case True note k0 = True from this obtain a where aI: "I = {a}" using True cardI card_1_singletonE by auto from this obtain b where bJ: "J = {b}" using True cardJ card_1_singletonE by auto have an: "a carrier_mat 1 1" unfolding carrier_mat_def submatrix_def using * ** aI bJ by auto have 1: "det (submatrix A {a} {b}) = (submatrix A {a} {b}) $$ (0,0)" by (rule det_singleton[OF sub_carrier]) have 2: "... = A $$ (a,b)" by (rule submatrix_singleton_index[OF A_carrier an bm]) show ?thesis proof (cases "A $$ (a,b) \ 0") let ?xs = "[submatrix A {a} {b} $$ (0,0)]" case True hence "a = b" using dA A_carrier an bm unfolding diagonal_mat_def carrier_mat_def by auto hence "set ?xs \ {A $$ (i, i) |i. i < min n m \ A $$ (i, i) \ 0}" using 2 True an bm by auto moreover have "det (submatrix A {a} {b}) = prod_list ?xs" using 1 by auto moreover have "length ?xs = Suc k" using k0 by auto ultimately show ?thesis using an bm unfolding aI bJ by blast next case False then show ?thesis using 1 2 aI bJ by auto qed next case False hence k0: "0 < k" by simp have k: "k < min n m" by (metis I J cardI cardJ le_imp_less_Suc less_Suc_eq_le min.commute min_def not_less subset_eq_atLeast0_lessThan_card) have subIJ_carrier: "(submatrix A I J) \ carrier_mat (Suc k) (Suc k)" unfolding carrier_mat_def using * ** cardI cardJ unfolding submatrix_def by auto obtain b' where b'k: "b' < Suc k" by auto let ?f="\i. submatrix A I J $$ (i, b') * cofactor (submatrix A I J) i b'" have det_rw: "det (submatrix A I J) = (\ia' 0") case True obtain a' where sub_IJ_0: "submatrix A I J $$ (a',b') \ 0" and a'k: "a' < Suc k" and unique: "\j. j submatrix A I J $$ (j,b') \ 0 \ j = a'" using diagonal_imp_submatrix_element_not0[OF dA A_carrier cardI cardJ I J b'k True] by auto have "submatrix A I J $$ (a', b') = A $$ (pick I a', pick J b')" by (rule submatrix_index, auto simp add: "*" a'k cardI "**" b'k cardJ) from this obtain a b where an: "a < n" and bm: "b < m" and sub_index: "submatrix A I J $$ (a', b') = A $$ (a, b)" and pick_a: "pick I a' = a" and pick_b: "pick J b' = b" using * ** A_carrier a'k b'k cardI cardJ pick_le by fastforce obtain I' where aI': "I = insert a I'" and a_notin: "a \ I'" by (metis Set.set_insert a'k cardI pick_a pick_in_set_le) obtain J' where bJ': "J = insert b J'" and b_notin: "b \ J'" by (metis Set.set_insert b'k cardJ pick_b pick_in_set_le) have Suc_k0: "0 < Suc k" by simp have aI: "a \ I" using aI' by auto have bJ: "b \ J" using bJ' by auto have cardI': "card I' = k" by (metis aI' a_notin cardI card.infinite card_insert_disjoint finite_insert nat.inject nat.simps(3)) have cardJ': "card J' = k" by (metis bJ' b_notin cardJ card.infinite card_insert_disjoint finite_insert nat.inject nat.simps(3)) have I': "I' \ {0.. {0.. (\xs. (det (submatrix A I' J') = prod_list xs \ det (submatrix A I' J') = - prod_list xs) \ set xs \ {A $$ (i, i) |i. i < min n m \ A $$ (i, i) \ 0} \ length xs = k)" proof (rule Suc.hyps[OF cardI' cardJ' _ _ k0]) show "I' \ {0.. {0..i det (submatrix A I' J') = - prod_list xs" and xs: "set xs \ {A $$ (i, i) |i. i < min n m \ A $$ (i, i) \ 0}" and length_xs: "length xs = k" using det_sub_I'J' by blast let ?ys = "A$$(a,b) # xs" have length_ys: "length ?ys = Suc k" using length_xs by auto have a_eq_b: "a=b" using A_carrier an bm sub_IJ_0 sub_index dA unfolding diagonal_mat_def by auto have A_aa_in: "A$$(a,a) \ {A $$ (i, i) |i. i < min n m \ A $$ (i, i) \ 0}" using a_eq_b an bm sub_IJ_0 sub_index by auto have ys: "set ?ys \ {A $$ (i, i) |i. i < min n m \ A $$ (i, i) \ 0}" using xs A_aa_in a_eq_b by auto show ?thesis proof (cases "even (a'+b')") case True have det_submatrix_IJ: "det (submatrix A I J) = A $$ (a, b) * det (submatrix A I' J')" using det_submatrix_IJ True by auto show ?thesis proof (cases "det (submatrix A I' J') = prod_list xs") case True have "det (submatrix A I J) = prod_list ?ys" using det_submatrix_IJ unfolding True by auto then show ?thesis using ys length_ys by blast next case False hence "det (submatrix A I' J') = - prod_list xs" using prod_list_xs by simp hence "det (submatrix A I J) = - prod_list ?ys" using det_submatrix_IJ by auto then show ?thesis using ys length_ys by blast qed next case False have det_submatrix_IJ: "det (submatrix A I J) = A $$ (a, b) * - det (submatrix A I' J')" using det_submatrix_IJ False by auto show ?thesis proof (cases "det (submatrix A I' J') = prod_list xs") case True have "det (submatrix A I J) = - prod_list ?ys" using det_submatrix_IJ unfolding True by auto then show ?thesis using ys length_ys by blast next case False hence "det (submatrix A I' J') = - prod_list xs" using prod_list_xs by simp hence "det (submatrix A I J) = prod_list ?ys" using det_submatrix_IJ by auto then show ?thesis using ys length_ys by blast qed qed qed next case False have "sum ?f {0.. {0.. J \ {0.. card I = k \ card J = k}" lemma Gcd_minors_dvd: fixes A::"'a::{semiring_Gcd,comm_ring_1} mat" assumes PAQ_B: "P * A * Q = B" and P: "P \ carrier_mat m m" and A: "A \ carrier_mat m n" and Q: "Q \ carrier_mat n n" and I: "I \ {0.. {0.. carrier_mat k n" proof - have "I = {i. i < dim_row P \ i \ I}" using P I A by auto hence "card {i. i < dim_row P \ i \ I} = k" using Ik by auto thus ?thesis using A unfolding submatrix_def by auto qed have subQ: "submatrix Q UNIV J \ carrier_mat n k" proof - have J_eq: "J = {j. j < dim_col Q \ j \ J}" using Q J A by auto hence "card {j. j < dim_col Q \ j \ J} = k" using Jk by auto moreover have "card {i. i < dim_row Q \ i \ UNIV} = n" using Q by auto ultimately show ?thesis unfolding submatrix_def by auto qed have sub_sub_PA: "(submatrix ?subPA UNIV I') = submatrix (P * A) I I'" for I' using submatrix_split2[symmetric] by auto have det_subPA_rw: "det (submatrix (P * A) I I') = (\J' | J' \ {0.. card J' = k. det ((submatrix P I J')) * det (submatrix A J' I'))" if I'1: "I' \ {0..C | C \ {0.. card C = k. det (submatrix (submatrix P I UNIV) UNIV C) * det (submatrix (submatrix A UNIV I') C UNIV))" proof (rule Cauchy_Binet) have "I = {i. i < dim_row P \ i \ I}" using P I A by auto thus "submatrix P I UNIV \ carrier_mat k m" using Ik P unfolding submatrix_def by auto have "I' = {j. j < dim_col A \ j \ I'}" using I'1 A by auto thus "submatrix A UNIV I' \ carrier_mat m k" using I'2 A unfolding submatrix_def by auto qed also have "... = (\J' | J' \ {0.. card J' = k. det (submatrix P I J') * det (submatrix A J' I'))" unfolding submatrix_split2[symmetric] submatrix_split[symmetric] by simp finally show ?thesis . qed have "det (submatrix B I J) = det (submatrix (P*A*Q) I J)" using PAQ_B by simp also have "... = det (?subPA * ?subQ)" unfolding submatrix_mult by auto also have "... = (\I' | I' \ {0.. card I' = k. det (submatrix ?subPA UNIV I') * det (submatrix ?subQ I' UNIV))" by (rule Cauchy_Binet[OF subPA subQ]) also have "... = (\I' | I' \ {0.. card I' = k. det (submatrix (P * A) I I') * det (submatrix Q I' J))" using submatrix_split[symmetric, of Q] submatrix_split2[symmetric, of "P*A"] by presburger also have "... = (\I' | I' \ {0.. card I' = k. \J' | J' \ {0.. card J' = k. det (submatrix P I J') * det (submatrix A J' I') * det (submatrix Q I' J))" using det_subPA_rw by (simp add: semiring_0_class.sum_distrib_right) finally have det_rw: "det (submatrix B I J) = (\I' | I' \ {0.. card I' = k. \J' | J' \ {0.. card J' = k. det (submatrix P I J') * det (submatrix A J' I') * det (submatrix Q I' J))" . show ?thesis proof (unfold det_rw, (rule dvd_sum)+) fix I' J' assume I': "I' \ {I'. I' \ {0.. card I' = k}" and J': "J' \ {J'. J' \ {0.. card J' = k}" have "Gcd (minors A k) dvd det (submatrix A J' I')" by (rule Gcd_dvd, unfold minors_def, insert A I' J', auto) then show "Gcd (minors A k) dvd det (submatrix P I J') * det (submatrix A J' I') * det (submatrix Q I' J)" by auto qed qed (*The conclusion could be simplified since we have S = I.*) lemma det_minors_diagonal2: assumes dA: "diagonal_mat A" and A_carrier: "A \ carrier_mat n m" and Ik: "card I = k" and Jk: "card J = k" and r: "I \ {0.. {0..0" shows "det (submatrix A I J) = 0 \ (\S. S \ {0.. card S = k \ S=I \ (det (submatrix A I J) = (\i\S. A $$ (i,i)) \ det (submatrix A I J) = - (\i\S. A $$ (i,i))))" using Ik Jk r c k proof (induct k arbitrary: I J) case 0 then show ?case by auto next case (Suc k) note cardI = Suc.prems(1) note cardJ = Suc.prems(2) note I = Suc.prems(3) note J = Suc.prems(4) have *: "{i. i < dim_row A \ i \ I} = I" using I Ik A_carrier carrier_mat_def by auto have **: "{j. j < dim_col A \ j \ J} = J" using J Jk A_carrier carrier_mat_def by auto show ?case proof (cases "k = 0") case True note k0 = True from this obtain a where aI: "I = {a}" using True cardI card_1_singletonE by auto from this obtain b where bJ: "J = {b}" using True cardJ card_1_singletonE by auto have an: "a carrier_mat 1 1" unfolding carrier_mat_def submatrix_def using * ** aI bJ by auto have 1: "det (submatrix A {a} {b}) = (submatrix A {a} {b}) $$ (0,0)" by (rule det_singleton[OF sub_carrier]) have 2: "... = A $$ (a,b)" by (rule submatrix_singleton_index[OF A_carrier an bm]) show ?thesis proof (cases "A $$ (a,b) \ 0") let ?S="{a}" case True hence ab: "a = b" using dA A_carrier an bm unfolding diagonal_mat_def carrier_mat_def by auto hence "?S \ {0..i\?S. A $$ (i, i))" using 1 2 ab by auto moreover have "card ?S = Suc k" using k0 by auto ultimately show ?thesis using an bm unfolding aI bJ by blast next case False then show ?thesis using 1 2 aI bJ by auto qed next case False hence k0: "0 < k" by simp have k: "k < min n m" by (metis I J cardI cardJ le_imp_less_Suc less_Suc_eq_le min.commute min_def not_less subset_eq_atLeast0_lessThan_card) have subIJ_carrier: "(submatrix A I J) \ carrier_mat (Suc k) (Suc k)" unfolding carrier_mat_def using * ** cardI cardJ unfolding submatrix_def by auto obtain b' where b'k: "b' < Suc k" by auto let ?f="\i. submatrix A I J $$ (i, b') * cofactor (submatrix A I J) i b'" have det_rw: "det (submatrix A I J) = (\ia' 0") case True obtain a' where sub_IJ_0: "submatrix A I J $$ (a',b') \ 0" and a'k: "a' < Suc k" and unique: "\j. j submatrix A I J $$ (j,b') \ 0 \ j = a'" using diagonal_imp_submatrix_element_not0[OF dA A_carrier cardI cardJ I J b'k True] by auto have "submatrix A I J $$ (a', b') = A $$ (pick I a', pick J b')" by (rule submatrix_index, auto simp add: "*" a'k cardI "**" b'k cardJ) from this obtain a b where an: "a < n" and bm: "b < m" and sub_index: "submatrix A I J $$ (a', b') = A $$ (a, b)" and pick_a: "pick I a' = a" and pick_b: "pick J b' = b" using * ** A_carrier a'k b'k cardI cardJ pick_le by fastforce obtain I' where aI': "I = insert a I'" and a_notin: "a \ I'" by (metis Set.set_insert a'k cardI pick_a pick_in_set_le) obtain J' where bJ': "J = insert b J'" and b_notin: "b \ J'" by (metis Set.set_insert b'k cardJ pick_b pick_in_set_le) have Suc_k0: "0 < Suc k" by simp have aI: "a \ I" using aI' by auto have bJ: "b \ J" using bJ' by auto have cardI': "card I' = k" by (metis aI' a_notin cardI card.infinite card_insert_disjoint finite_insert nat.inject nat.simps(3)) have cardJ': "card J' = k" by (metis bJ' b_notin cardJ card.infinite card_insert_disjoint finite_insert nat.inject nat.simps(3)) have I': "I' \ {0.. {0.. (\S\{0.. S=I' \ (det (submatrix A I' J') = (\i\S. A $$ (i, i)) \ det (submatrix A I' J') = - (\i\S. A $$ (i, i))))" proof (rule Suc.hyps[OF cardI' cardJ' _ _ k0]) show "I' \ {0.. {0..ii\xs. A $$ (i, i)) \ det (submatrix A I' J') = - (\i\xs. A $$ (i, i))" and xs: "xs\{0.. xs" by (simp add: xs_I' a_notin) have length_ys: "card ?ys = Suc k" using length_xs a_notin_xs by (simp add: card_ge_0_finite k0) have a_eq_b: "a=b" using A_carrier an bm sub_IJ_0 sub_index dA unfolding diagonal_mat_def by auto have A_aa_in: "A$$(a,a) \ {A $$ (i, i) |i. i < min n m \ A $$ (i, i) \ 0}" using a_eq_b an bm sub_IJ_0 sub_index by auto show ?thesis proof (cases "even (a'+b')") case True have det_submatrix_IJ: "det (submatrix A I J) = A $$ (a, b) * det (submatrix A I' J')" using det_submatrix_IJ True by auto show ?thesis proof (cases "det (submatrix A I' J') = (\i\xs. A $$ (i, i))") case True have "det (submatrix A I J) = (\i\?ys. A $$ (i, i))" using det_submatrix_IJ unfolding True a_eq_b by (metis (no_types, lifting) a_notin_xs a_eq_b card_ge_0_finite k0 length_xs prod.insert) then show ?thesis using length_ys using a_eq_b an bm xs xs_I' by (simp add: aI') next case False hence "det (submatrix A I' J') = - (\i\xs. A $$ (i, i))" using prod_list_xs by simp hence "det (submatrix A I J) = -(\i\?ys. A $$ (i, i))" using det_submatrix_IJ a_eq_b by (metis (no_types, lifting) a_notin_xs card_ge_0_finite k0 length_xs mult_minus_right prod.insert) then show ?thesis using length_ys using a_eq_b an bm xs aI' xs_I' by force qed next case False have det_submatrix_IJ: "det (submatrix A I J) = A $$ (a, b) * - det (submatrix A I' J')" using det_submatrix_IJ False by auto show ?thesis proof (cases "det (submatrix A I' J') = (\i\xs. A $$ (i, i))") case True have "det (submatrix A I J) = - (\i\?ys. A $$ (i, i))" using det_submatrix_IJ unfolding True by (metis (no_types, lifting) a_eq_b a_notin_xs card_ge_0_finite k0 length_xs mult_minus_right prod.insert) then show ?thesis using length_ys using a_eq_b an bm xs aI' xs_I' by force next case False hence "det (submatrix A I' J') = - (\i\xs. A $$ (i, i))" using prod_list_xs by simp hence "det (submatrix A I J) = (\i\?ys. A $$ (i, i))" using det_submatrix_IJ by (metis (mono_tags, lifting) a_eq_b a_notin_xs card_ge_0_finite equation_minus_iff k0 length_xs prod.insert) then show ?thesis using length_ys using a_eq_b an bm xs aI' xs_I' by force qed qed qed next case False have "sum ?f {0..Relating minors and GCD\ lemma diagonal_dvd_Gcd_minors: fixes A::"'a::{semiring_Gcd,comm_ring_1} mat" assumes A: "A \ carrier_mat n m" and SNF_A: "Smith_normal_form_mat A" shows "(\i=0.. minors A k" show "(\i = 0.. {0.. {0.. {0..i\S. A $$ (i,i)) \ det (submatrix A I J) = -(\i\S. A $$ (i,i))" using det_minors_diagonal2[OF diag_A A Ik Jk _ _ k] I J A False b by auto obtain f where inj_f: "inj_on f {0..i\{0.. f i)" using exists_inj_ge_index[OF S Sk] by blast have "(\i = 0..i\{0..i\f`{0..i\S. A $$ (i, i))" using fk_S by auto finally have *: "(\i = 0..i\S. A $$ (i, i))" . show "(\i = 0.. carrier_mat n m" and SNF_A: "Smith_normal_form_mat A" and k: "k \ min n m" shows "Gcd (minors A k) dvd (\i=0..i = 0..(i, j). A $$ (i, j))" proof (rule eq_matI, auto) have "I = {i. i < dim_row A \ i \ I}" unfolding I_def using A k by auto hence ck: "card {i. i < dim_row A \ i \ I} = k" unfolding I_def using card_atLeastLessThan by presburger have "I = {i. i < dim_col A \ i \ I}" unfolding I_def using A k by auto hence ck2: "card {j. j < dim_col A \ j \ I} = k" unfolding I_def using card_atLeastLessThan by presburger show dr: "dim_row (submatrix A I I) = k" using ck unfolding submatrix_def by auto show dc: "dim_col (submatrix A I I) = k" using ck2 unfolding submatrix_def by auto fix i j assume i: "i < k" and j: "j < k" have p1: "pick I i = i" proof - have "{0.. I. a < i}" using I_def i by auto hence i_eq: "i = card {a \ I. a < i}" by (metis card_atLeastLessThan diff_zero) have "pick I i = pick I (card {a \ I. a < i})" using i_eq by simp also have "... = i" by (rule pick_card_in_set, insert i I_def, simp) finally show ?thesis . qed have p2: "pick I j = j" proof - have "{0.. I. a < j}" using I_def j by auto hence j_eq: "j = card {a \ I. a < j}" by (metis card_atLeastLessThan diff_zero) have "pick I j = pick I (card {a \ I. a < j})" using j_eq by simp also have "... = j" by (rule pick_card_in_set, insert j I_def, simp) finally show ?thesis . qed have "submatrix A I I $$ (i, j) = A $$ (pick I i, pick I j)" proof (rule submatrix_index) show "i < card {i. i < dim_row A \ i \ I}" by (metis dim_submatrix(1) dr i) show "j < card {j. j < dim_col A \ j \ I}" by (metis dim_submatrix(2) dc j) qed also have "... = A $$ (i,j)" using p1 p2 by simp finally show "submatrix A I I $$ (i, j) = A $$ (i, j)" . qed hence "det (submatrix A I I) = det (mat k k (\(i, j). A $$ (i, j)))" by simp also have "... = prod_list (diag_mat (mat k k (\(i, j). A $$ (i, j))))" proof (rule det_upper_triangular) show "mat k k (\(i, j). A $$ (i, j)) \ carrier_mat k k" by auto show "upper_triangular (Matrix.mat k k (\(i, j). A $$ (i, j)))" using SNF_A A k unfolding Smith_normal_form_mat_def isDiagonal_mat_def by auto qed also have "... = (\i = 0.. {0.. {0..i = 0.. minors A k" unfolding minors_def by auto qed lemma Gcd_minors_A_dvd_Gcd_minors_PAQ: fixes A::"'a::{semiring_Gcd,comm_ring_1} mat" assumes A: "A \ carrier_mat m n" and P: "P \ carrier_mat m m" and Q: "Q \ carrier_mat n n" shows "Gcd (minors A k) dvd Gcd (minors (P*A*Q) k)" proof (rule Gcd_greatest) let ?B="(P * A * Q)" fix b assume "b \ minors ?B k" from this obtain I J where b: "b = det (submatrix ?B I J)" and I: "I \ {0.. {0.. carrier_mat m n" and P: "P \ carrier_mat m m" and Q: "Q \ carrier_mat n n" and inv_P: "invertible_mat P" and inv_Q: "invertible_mat Q" shows "Gcd (minors (P*A*Q) k) dvd Gcd (minors A k)" proof (rule Gcd_greatest) let ?B = "P * A * Q" fix b assume "b \ minors A k" from this obtain I J where b: "b = det (submatrix A I J)" and I: "I \ {0.. {0.. carrier_mat m m" using PP' P'P unfolding inverts_mat_def by (metis P carrier_matD(1) carrier_matD(2) carrier_matI index_mult_mat(3) index_one_mat(3)) have Q': "Q' \ carrier_mat n n" using QQ' Q'Q unfolding inverts_mat_def by (metis Q carrier_matD(1) carrier_matD(2) carrier_matI index_mult_mat(3) index_one_mat(3)) have rw: "P' *?B *Q' = A" proof - have f1: "P' * P = 1\<^sub>m m" by (metis (no_types) P' P'P carrier_matD(1) inverts_mat_def) have *: "P' * P * A = P' * (P * A)" by (meson A P P' assoc_mult_mat) have " P' * (P * A * Q) * Q' = P' * P * A * Q * Q'" by (smt A P P' Q assoc_mult_mat mult_carrier_mat) also have "... = P' * P * (A * Q * Q')" using A P P' Q Q' f1 * by auto also have "... = A * Q * Q'" using P'P A P' unfolding inverts_mat_def by auto also have "... = A" using QQ' A Q' Q unfolding inverts_mat_def by auto finally show ?thesis . qed have "Gcd (minors ?B k) dvd det (submatrix (P'*?B*Q') I J)" by (rule Gcd_minors_dvd[OF _ P' _ Q' _ _ Ik Jk], insert P A Q I J, auto) also have "... = det (submatrix A I J)" using rw by simp finally show "Gcd (minors ?B k) dvd b" using b by simp qed lemma Gcd_minors_dvd_diag_PAQ: fixes P A Q::"'a::{semiring_Gcd,comm_ring_1} mat" assumes A: "A \ carrier_mat m n" and P: "P \ carrier_mat m m" and Q: "Q \ carrier_mat n n" and SNF: "Smith_normal_form_mat (P*A*Q)" and k: "k\min m n" shows "Gcd (minors A k) dvd (\i=0..i=0.. carrier_mat m n" and P: "P \ carrier_mat m m" and Q: "Q \ carrier_mat n n" and inv_P: "invertible_mat P" and inv_Q: "invertible_mat Q" and SNF: "Smith_normal_form_mat (P*A*Q)" shows "(\i=0..i=0.. carrier_mat m n" and SNF: "Smith_normal_form_mat A" and prod_0: "(\j=0..Final theorem\ lemma Smith_normal_form_uniqueness_aux: fixes P A Q::"'a::{idom,semiring_Gcd} mat" assumes A: "A \ carrier_mat m n" (*PAQ = B with B in SNF and P,Q invertible matrices*) and P: "P \ carrier_mat m m" and Q: "Q \ carrier_mat n n" and inv_P: "invertible_mat P" and inv_Q: "invertible_mat Q" and PAQ_B: "P*A*Q = B" and SNF: "Smith_normal_form_mat B" (*P'AQ' = B' with B' in SNF and P',Q' invertible matrices*) and P': "P' \ carrier_mat m m" and Q': "Q' \ carrier_mat n n" and inv_P': "invertible_mat P'" and inv_Q': "invertible_mat Q'" and P'AQ'_B': "P'*A*Q' = B'" and SNF_B': "Smith_normal_form_mat B'" and k: "ki\k. B$$(i,i) dvd B'$$(i,i) \ B'$$(i,i) dvd B$$(i,i)" proof (rule allI, rule impI) fix i assume ik: "i \ k" show " B $$ (i, i) dvd B' $$ (i, i) \ B' $$ (i, i) dvd B $$ (i, i)" proof - let ?\Bi = "(\i=0..B'i = "(\i=0..B'i dvd Gcd (minors A i)" by (unfold P'AQ'_B'[symmetric], rule diag_PAQ_dvd_Gcd_minors[OF A P' Q' inv_P' inv_Q'], insert P'AQ'_B' SNF_B' ik k, auto ) also have "... dvd ?\Bi" by (unfold PAQ_B[symmetric], rule Gcd_minors_dvd_diag_PAQ[OF A P Q], insert PAQ_B SNF ik k, auto) finally have B'_i_dvd_B_i: "?\B'i dvd ?\Bi" . have "?\Bi dvd Gcd (minors A i)" by (unfold PAQ_B[symmetric], rule diag_PAQ_dvd_Gcd_minors[OF A P Q inv_P inv_Q], insert PAQ_B SNF ik k, auto ) also have "... dvd ?\B'i" by (unfold P'AQ'_B'[symmetric], rule Gcd_minors_dvd_diag_PAQ[OF A P' Q'], insert P'AQ'_B' SNF_B' ik k, auto) finally have B_i_dvd_B'_i: "?\Bi dvd ?\B'i" . let ?\B_Suc = "(\i=0..B'_Suc = "(\i=0..B'_Suc dvd Gcd (minors A (Suc i))" by (unfold P'AQ'_B'[symmetric], rule diag_PAQ_dvd_Gcd_minors[OF A P' Q' inv_P' inv_Q'], insert P'AQ'_B' SNF_B' ik k, auto ) also have "... dvd ?\B_Suc" by (unfold PAQ_B[symmetric], rule Gcd_minors_dvd_diag_PAQ[OF A P Q], insert PAQ_B SNF ik k, auto) finally have 3: "?\B'_Suc dvd ?\B_Suc" . have "?\B_Suc dvd Gcd (minors A (Suc i))" by (unfold PAQ_B[symmetric], rule diag_PAQ_dvd_Gcd_minors[OF A P Q inv_P inv_Q], insert PAQ_B SNF ik k, auto ) also have "... dvd ?\B'_Suc" by (unfold P'AQ'_B'[symmetric], rule Gcd_minors_dvd_diag_PAQ[OF A P' Q'], insert P'AQ'_B' SNF_B' ik k, auto) finally have 4: "?\B_Suc dvd ?\B'_Suc" . show ?thesis proof (cases "?\B_Suc = 0") case True have True2: "?\B'_Suc = 0" using 4 True by fastforce have "B$$(i,i) = 0" by (rule Smith_prod_zero_imp_last_zero[OF _ SNF True], insert ik k PAQ_B P Q, auto) moreover have "B'$$(i,i) = 0" by (rule Smith_prod_zero_imp_last_zero[OF _ SNF_B' True2], insert ik k P'AQ'_B' P' Q', auto) ultimately show ?thesis by auto next case False have "\u. u dvd 1 \ ?\B'i = u * ?\Bi" by (rule dvd_associated2[OF B'_i_dvd_B_i B_i_dvd_B'_i], insert False B'_i_dvd_B_i, force) from this obtain u where eq1: "(\i=0..i=0..u. u dvd 1 \ ?\B_Suc = u * ?\B'_Suc" by (rule dvd_associated2[OF 4 3 False]) from this obtain w where eq2: "(\i=0..i=0..i=0..i=0..i=0..i=0..i=0..u. is_unit u \ B $$ (i, i) = u * B' $$ (i, i)" by auto thus ?thesis using dvd_associated2 by force qed qed qed lemma Smith_normal_form_uniqueness: fixes P A Q::"'a::{idom,semiring_Gcd} mat" assumes A: "A \ carrier_mat m n" (*PAQ = B with B in SNF and P,Q invertible matrices*) and P: "P \ carrier_mat m m" and Q: "Q \ carrier_mat n n" and inv_P: "invertible_mat P" and inv_Q: "invertible_mat Q" and PAQ_B: "P*A*Q = B" and SNF: "Smith_normal_form_mat B" (*P'AQ' = B' with B' in SNF and P',Q' invertible matrices*) and P': "P' \ carrier_mat m m" and Q': "Q' \ carrier_mat n n" and inv_P': "invertible_mat P'" and inv_Q': "invertible_mat Q'" and P'AQ'_B': "P'*A*Q' = B'" and SNF_B': "Smith_normal_form_mat B'" and i: "i < min m n" shows "\u. u dvd 1 \ B $$ (i,i) = u * B' $$ (i,i)" proof (cases "B $$ (i,i) = 0") case True let ?\B_Suc = "(\i=0..B'_Suc = "(\i=0..B_Suc dvd Gcd (minors A (Suc i))" by (unfold PAQ_B[symmetric], rule diag_PAQ_dvd_Gcd_minors[OF A P Q inv_P inv_Q], insert PAQ_B SNF i, auto) also have "... dvd ?\B'_Suc" by (unfold P'AQ'_B'[symmetric], rule Gcd_minors_dvd_diag_PAQ[OF A P' Q'], insert P'AQ'_B' SNF_B' i, auto) finally have 4: "?\B_Suc dvd ?\B'_Suc" . have prod0: "?\B_Suc=0" using True by auto have True2: "?\B'_Suc = 0" using 4 by (metis dvd_0_left_iff prod0) have "B'$$(i,i) = 0" by (rule Smith_prod_zero_imp_last_zero[OF _ SNF_B' True2], insert i P'AQ'_B' P' Q', auto) thus ?thesis using True by auto next case False have "\a\i. B$$(a,a) dvd B'$$(a,a) \ B'$$(a,a) dvd B$$(a,a)" by (rule Smith_normal_form_uniqueness_aux[OF assms]) hence "B$$(i,i) dvd B'$$(i,i) \ B'$$(i,i) dvd B$$(i,i)" using i by auto thus ?thesis using dvd_associated2 False by blast qed text \The final theorem, moved to HOL Analysis\ lemma Smith_normal_form_uniqueness_HOL_Analysis: fixes A::"'a::{idom,semiring_Gcd}^'m::mod_type^'n::mod_type" and P P'::"'a^'n::mod_type^'n::mod_type" and Q Q'::"'a^'m::mod_type^'m::mod_type" assumes (*PAQ = B with B in SNF and P,Q invertible matrices*) inv_P: "invertible P" and inv_Q: "invertible Q" and PAQ_B: "P**A**Q = B" and SNF: "Smith_normal_form B" (*P'AQ' = B' with B' in SNF and P',Q' invertible matrices*) and inv_P': "invertible P'" and inv_Q': "invertible Q'" and P'AQ'_B': "P'**A**Q' = B'" and SNF_B': "Smith_normal_form B'" and i: "i < min (nrows A) (ncols A)" shows "\u. u dvd 1 \ B $h Mod_Type.from_nat i $h Mod_Type.from_nat i = u * B' $h Mod_Type.from_nat i $h Mod_Type.from_nat i" proof - let ?P = "Mod_Type_Connect.from_hma\<^sub>m P" let ?A = "Mod_Type_Connect.from_hma\<^sub>m A" let ?Q = "Mod_Type_Connect.from_hma\<^sub>m Q" let ?B = "Mod_Type_Connect.from_hma\<^sub>m B" let ?P' = "Mod_Type_Connect.from_hma\<^sub>m P'" let ?Q' = "Mod_Type_Connect.from_hma\<^sub>m Q'" let ?B' = "Mod_Type_Connect.from_hma\<^sub>m B'" let ?i = "(Mod_Type.from_nat i)::'n" let ?i' = "(Mod_Type.from_nat i)::'m" have [transfer_rule]: "Mod_Type_Connect.HMA_M ?P P" by (simp add: Mod_Type_Connect.HMA_M_def) have [transfer_rule]: "Mod_Type_Connect.HMA_M ?A A" by (simp add: Mod_Type_Connect.HMA_M_def) have [transfer_rule]: "Mod_Type_Connect.HMA_M ?Q Q" by (simp add: Mod_Type_Connect.HMA_M_def) have [transfer_rule]: "Mod_Type_Connect.HMA_M ?B B" by (simp add: Mod_Type_Connect.HMA_M_def) have [transfer_rule]: "Mod_Type_Connect.HMA_M ?P' P'" by (simp add: Mod_Type_Connect.HMA_M_def) have [transfer_rule]: "Mod_Type_Connect.HMA_M ?Q' Q'" by (simp add: Mod_Type_Connect.HMA_M_def) have [transfer_rule]: "Mod_Type_Connect.HMA_M ?B' B'" by (simp add: Mod_Type_Connect.HMA_M_def) have [transfer_rule]: "Mod_Type_Connect.HMA_I i ?i" by (metis Mod_Type_Connect.HMA_I_def i min.strict_boundedE mod_type_class.to_nat_from_nat_id nrows_def) have [transfer_rule]: "Mod_Type_Connect.HMA_I i ?i'" by (metis Mod_Type_Connect.HMA_I_def i min.strict_boundedE mod_type_class.to_nat_from_nat_id ncols_def) have i2: "i < min CARD('m) CARD('n)" using i unfolding nrows_def ncols_def by auto have "\u. u dvd 1 \ ?B $$(i,i) = u * ?B' $$ (i,i)" proof (rule Smith_normal_form_uniqueness[of _ "CARD('n)" "CARD('m)"]) show "?P*?A*?Q=?B" using PAQ_B by (transfer', auto) show "Smith_normal_form_mat ?B" using SNF by (transfer', auto) show "?P'*?A*?Q'=?B'" using P'AQ'_B' by (transfer', auto) show "Smith_normal_form_mat ?B'" using SNF_B' by (transfer', auto) show "invertible_mat ?P" using inv_P by (transfer, auto) show "invertible_mat ?P'" using inv_P' by (transfer, auto) show "invertible_mat ?Q" using inv_Q by (transfer, auto) show "invertible_mat ?Q'" using inv_Q' by (transfer, auto) qed (insert i2, auto) hence "\u. u dvd 1 \ (index_hma B ?i ?i') = u * (index_hma B' ?i ?i')" by (transfer', rule) thus ?thesis unfolding index_hma_def by simp qed subsection \Uniqueness fixing a complete set of non-associates\ definition "Smith_normal_form_wrt A \ = ( (\a b. Mod_Type.to_nat a = Mod_Type.to_nat b \ Mod_Type.to_nat a + 1 < nrows A \ Mod_Type.to_nat b + 1 < ncols A \ A $h a $h b dvd A $h (a+1) $h (b+1)) \ isDiagonal A \ Complete_set_non_associates \ \ (\a b. Mod_Type.to_nat a = Mod_Type.to_nat b \ Mod_Type.to_nat a < min (nrows A) (ncols A) \ Mod_Type.to_nat b < min (nrows A) (ncols A) \ A $h a $h b \ \) )" lemma Smith_normal_form_wrt_uniqueness_HOL_Analysis: fixes A::"'a::{idom,semiring_Gcd}^'m::mod_type^'n::mod_type" and P P'::"'a^'n::mod_type^'n::mod_type" and Q Q'::"'a^'m::mod_type^'m::mod_type" assumes (*PAQ = S with S in SNF and P,Q invertible matrices*) P: "invertible P" and Q: "invertible Q" and PAQ_S: "P**A**Q = S" and SNF: "Smith_normal_form_wrt S \" (*P'AQ' = S' with S' in SNF and P',Q' invertible matrices*) and P': "invertible P'" and Q': "invertible Q'" and P'AQ'_S': "P'**A**Q' = S'" and SNF_S': "Smith_normal_form_wrt S' \" shows "S = S'" proof - have "S $h i $h j = S' $h i $h j" for i j proof (cases "Mod_Type.to_nat i \ Mod_Type.to_nat j") case True then show ?thesis using SNF SNF_S' unfolding Smith_normal_form_wrt_def isDiagonal_def by auto next case False let ?i = "Mod_Type.to_nat i" let ?j = "Mod_Type.to_nat j" have complete_set: "Complete_set_non_associates \" using SNF_S' unfolding Smith_normal_form_wrt_def by simp have ij: "?i = ?j" using False by auto show ?thesis proof (rule ccontr) assume d: "S $h i $h j \ S' $h i $h j" have n: "normalize (S $h i $h j) \ normalize (S' $h i $h j)" proof (rule in_Ass_not_associated[OF complete_set _ _ d]) show "S $h i $h j \ \" using SNF unfolding Smith_normal_form_wrt_def by (metis False min_less_iff_conj mod_type_class.to_nat_less_card ncols_def nrows_def) show "S' $h i $h j \ \" using SNF_S' unfolding Smith_normal_form_wrt_def by (metis False min_less_iff_conj mod_type_class.to_nat_less_card ncols_def nrows_def) qed have "\u. u dvd 1 \ S $h i $h j = u * S' $h i $h j" proof - have "\u. u dvd 1 \ S $h Mod_Type.from_nat ?i $h Mod_Type.from_nat ?i = u * S' $h Mod_Type.from_nat ?i $h Mod_Type.from_nat ?i" proof (rule Smith_normal_form_uniqueness_HOL_Analysis[OF P Q PAQ_S _ P' Q' P'AQ'_S' _]) show "Smith_normal_form S" and "Smith_normal_form S'" using SNF SNF_S' Smith_normal_form_def Smith_normal_form_wrt_def by blast+ show "?i < min (nrows A) (ncols A)" by (metis ij min_less_iff_conj mod_type_class.to_nat_less_card ncols_def nrows_def) qed thus ?thesis using False by auto qed from this obtain u where "is_unit u" and "S $h i $h j = u * S' $h i $h j" by auto thus False using n by (simp add: normalize_1_iff normalize_mult) qed qed thus ?thesis by vector qed end \ No newline at end of file diff --git a/thys/Youngs_Inequality/Youngs.thy b/thys/Youngs_Inequality/Youngs.thy --- a/thys/Youngs_Inequality/Youngs.thy +++ b/thys/Youngs_Inequality/Youngs.thy @@ -1,870 +1,870 @@ section \Young's Inequality for Increasing Functions\ text \From the following paper: Cunningham, F., and Nathaniel Grossman. “On Young’s Inequality.” The American Mathematical Monthly 78, no. 7 (1971): 781–83. \url{https://doi.org/10.2307/2318018}\ theory Youngs imports "HOL-Analysis.Analysis" begin subsection \Library Extras: already added to the repository\ text \In fact, strict inequality is required only at a single point within the box.\ lemma integral_less: fixes f :: "'n::euclidean_space \ real" assumes cont: "continuous_on (cbox a b) f" "continuous_on (cbox a b) g" and ne: "box a b \ {}" and fg: "\x. x \ box a b \ f x < g x" shows "integral (cbox a b) f < integral (cbox a b) g" proof - obtain int: "f integrable_on (cbox a b)" "g integrable_on (cbox a b)" using cont integrable_continuous by blast then have "integral (cbox a b) f \ integral (cbox a b) g" by (metis fg integrable_on_open_interval integral_le integral_open_interval less_eq_real_def) moreover have "integral (cbox a b) f \ integral (cbox a b) g" proof (rule ccontr) assume "\ integral (cbox a b) f \ integral (cbox a b) g" then have 0: "((\x. g x - f x) has_integral 0) (cbox a b)" by (metis (full_types) cancel_comm_monoid_add_class.diff_cancel has_integral_diff int integrable_integral) have cgf: "continuous_on (cbox a b) (\x. g x - f x)" using cont continuous_on_diff by blast show False using has_integral_0_cbox_imp_0 [OF cgf _ 0] ne box_subset_cbox fg by fastforce qed ultimately show ?thesis by linarith qed lemma integral_less_real: fixes f :: "real \ real" assumes "continuous_on {a..b} f" "continuous_on {a..b} g" and "{a<.. {}" and "\x. x \ {a<.. f x < g x" shows "integral {a..b} f < integral {a..b} g" by (metis assms box_real integral_less) lemma has_integral_UN: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "finite I" and int: "\i. i \ I \ (f has_integral (g i)) (\ i)" and neg: "pairwise (\i i'. negligible (\ i \ \ i')) I" shows "(f has_integral (sum g I)) (\i\I. \ i)" proof - let ?\ = "((\(a,b). \ a \ \ b) ` {(a,b). a \ I \ b \ I-{a}})" have "((\x. if x \ (\i\I. \ i) then f x else 0) has_integral sum g I) UNIV" proof (rule has_integral_spike) show "negligible (\?\)" proof (rule negligible_Union) have "finite (I \ I)" by (simp add: \finite I\) moreover have "{(a,b). a \ I \ b \ I-{a}} \ I \ I" by auto ultimately show "finite ?\" by (simp add: finite_subset) show "\t. t \ ?\ \ negligible t" using neg unfolding pairwise_def by auto qed next show "(if x \ (\i\I. \ i) then f x else 0) = (\i\I. if x \ \ i then f x else 0)" if "x \ UNIV - (\?\)" for x proof clarsimp fix i assume i: "i \ I" "x \ \ i" then have "\j\I. x \ \ j \ j = i" using that by blast with i show "f x = (\i\I. if x \ \ i then f x else 0)" by (simp add: sum.delta[OF \finite I\]) qed next show "((\x. (\i\I. if x \ \ i then f x else 0)) has_integral sum g I) UNIV" using int by (simp add: has_integral_restrict_UNIV has_integral_sum[OF \finite I\]) qed then show ?thesis using has_integral_restrict_UNIV by blast qed lemma integrable_mono_on_nonneg: fixes f :: "real \ real" - assumes mon: "mono_on f {a..b}" and 0: "\x. 0 \ f x" + assumes mon: "mono_on {a..b} f" and 0: "\x. 0 \ f x" shows "integrable (lebesgue_on {a..b}) f" proof - have "space lborel = space lebesgue" "sets borel \ sets lebesgue" by force+ then have fborel: "f \ borel_measurable (lebesgue_on {a..b})" by (metis mon borel_measurable_mono_on_fnc borel_measurable_subalgebra mono_restrict_space space_lborel space_restrict_space) then obtain g where g: "incseq g" and simple: "\i. simple_function (lebesgue_on {a..b}) (g i)" and bdd: " (\x. bdd_above (range (\i. g i x)))" and nonneg: "\i x. 0 \ g i x" and fsup: "f = (SUP i. g i)" by (metis borel_measurable_implies_simple_function_sequence_real 0) have "f ` {a..b} \ {f a..f b}" using assms by (auto simp: mono_on_def) have g_le_f: "g i x \ f x" for i x proof - have "bdd_above ((\h. h x) ` range g)" using bdd cSUP_lessD linorder_not_less by fastforce then show ?thesis by (metis SUP_apply UNIV_I bdd cSUP_upper fsup) qed then have gfb: "g i x \ f b" if "x \ {a..b}" for i x by (smt (verit, best) mon atLeastAtMost_iff mono_on_def that) have g_le: "g i x \ g j x" if "i\j" for i j x using g by (simp add: incseq_def le_funD that) show "integrable (lebesgue_on {a..b}) ( f)" proof (rule integrable_dominated_convergence) show "f \ borel_measurable (lebesgue_on {a..b})" using fborel by blast have "\x. (\i. g i x) \ (SUP h \ range g. h x)" proof (rule order_tendstoI) show "\\<^sub>F i in sequentially. y < g i x" if "y < (SUP h\range g. h x)" for x y proof - from that obtain h where h: "h \ range g" "y < h x" using g_le_f by (subst (asm)less_cSUP_iff) fastforce+ then show ?thesis by (smt (verit, ccfv_SIG) eventually_sequentially g_le imageE) qed show "\\<^sub>F i in sequentially. g i x < y" if "(SUP h\range g. h x) < y" for x y by (smt (verit, best) that Sup_apply g_le_f always_eventually fsup image_cong) qed then show "AE x in lebesgue_on {a..b}. (\i. g i x) \ f x" by (simp add: fsup) fix i show "g i \ borel_measurable (lebesgue_on {a..b})" using borel_measurable_simple_function simple by blast show "AE x in lebesgue_on {a..b}. norm (g i x) \ f b" by (simp add: gfb nonneg Measure_Space.AE_I' [of "{}"]) qed auto qed lemma integrable_mono_on: fixes f :: "real \ real" - assumes "mono_on f {a..b}" + assumes "mono_on {a..b} f" shows "integrable (lebesgue_on {a..b}) f" proof - define f' where "f' \ \x. if x \ {a..b} then f x - f a else 0" - have "mono_on f' {a..b}" + have "mono_on {a..b} f'" by (smt (verit, best) assms f'_def mono_on_def) moreover have 0: "\x. 0 \ f' x" by (smt (verit, best) assms atLeastAtMost_iff f'_def mono_on_def) ultimately have "integrable (lebesgue_on {a..b}) f'" using integrable_mono_on_nonneg by presburger then have "integrable (lebesgue_on {a..b}) (\x. f' x + f a)" by force moreover have "space lborel = space lebesgue" "sets borel \ sets lebesgue" by force+ then have fborel: "f \ borel_measurable (lebesgue_on {a..b})" using borel_measurable_mono_on_fnc [OF assms] by (metis borel_measurable_subalgebra mono_restrict_space space_lborel space_restrict_space) ultimately show ?thesis by (rule integrable_cong_AE_imp) (auto simp: f'_def) qed lemma integrable_on_mono_on: fixes f :: "real \ real" - assumes "mono_on f {a..b}" + assumes "mono_on {a..b} f" shows "f integrable_on {a..b}" by (simp add: assms integrable_mono_on integrable_on_lebesgue_on) lemma strict_mono_image_endpoints: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" - assumes "strict_mono_on f {a..b}" and f: "continuous_on {a..b} f" and "a \ b" + assumes "strict_mono_on {a..b} f" and f: "continuous_on {a..b} f" and "a \ b" shows "f ` {a..b} = {f a..f b}" proof show "f ` {a..b} \ {f a..f b}" using assms(1) strict_mono_on_leD by fastforce show "{f a..f b} \ f ` {a..b}" using assms IVT'[OF _ _ _ f] by (force simp: Bex_def) qed subsection \Toward Young's inequality\ text \Generalisations of the type of @{term f} are not obvious\ lemma strict_mono_continuous_invD: fixes f :: "real \ real" - assumes sm: "strict_mono_on f {a..}" and contf: "continuous_on {a..} f" + assumes sm: "strict_mono_on {a..} f" and contf: "continuous_on {a..} f" and fim: "f ` {a..} = {f a..}" and g: "\x. x \ a \ g (f x) = x" shows "continuous_on {f a..} g" proof (clarsimp simp add: continuous_on_eq_continuous_within) fix y assume "f a \ y" then obtain u where u: "y+1 = f u" "u \ a" by (smt (verit, best) atLeast_iff fim imageE) have "continuous_on {f a..y+1} g" proof - - obtain "continuous_on {a..u} f" "strict_mono_on f {a..u}" + obtain "continuous_on {a..u} f" "strict_mono_on {a..u} f" using contf sm continuous_on_subset by (force simp add: strict_mono_on_def) moreover have "continuous_on (f ` {a..u}) g" using assms continuous_on_subset by (intro continuous_on_inv) fastforce+ ultimately show ?thesis - using strict_mono_image_endpoints [of f] + using strict_mono_image_endpoints [of _ _ f] by (simp add: strict_mono_image_endpoints u) qed then have *: "continuous (at y within {f a..y+1}) g" by (simp add: \f a \ y\ continuous_on_imp_continuous_within) show "continuous (at y within {f a..}) g" proof (clarsimp simp add: continuous_within_topological Ball_def) fix B assume "open B" and "g y \ B" with * obtain A where A: "open A" "y \ A" and "\x. f a \ x \ x \ y+1 \ x \ A \ g x \ B" by (force simp: continuous_within_topological) then have "\x\f a. x \ A \ ball y 1 \ g x \ B" by (smt (verit, ccfv_threshold) IntE dist_norm mem_ball real_norm_def) then show "\A. open A \ y \ A \ (\x\f a. x \ A \ g x \ B)" by (metis Elementary_Metric_Spaces.open_ball Int_iff A centre_in_ball open_Int zero_less_one) qed qed subsection \Regular divisions\ text \Our lack of the Riemann integral forces us to construct explicitly the step functions mentioned in the text.\ definition "segment \ \n k. {real k / real n..(1 + k) / real n}" lemma segment_nonempty: "segment n k \ {}" by (auto simp: segment_def divide_simps) lemma segment_Suc: "segment n ` {.. (segment n ` {.. \n. segment n ` {.. (segments n) = (if n=0 then {} else {0..1})" by (simp add: segments_def Union_segment_image) definition "regular_division \ \a b n. (image ((+) a \ (*) (b-a))) ` (segments n)" lemma translate_scale_01: assumes "a \ b" shows "(\x. a + (b - a) * x) ` {0..1} = {a..b::real}" using closed_segment_real_eq [of a b] assms closed_segment_eq_real_ivl by auto lemma finite_regular_division [simp]: "finite (regular_division a b n)" by (simp add: regular_division_def segments_def) lemma card_regular_division [simp]: assumes "a (*) (b - a))) (segments n)" proof fix x y assume "((+) a \ (*) (b - a)) ` x = ((+) a \ (*) (b - a)) ` y" then have "(+) (-a) ` ((+) a \ (*) (b - a)) ` x = (+) (-a) ` ((+) a \ (*) (b - a)) ` y" by simp then have "((*) (b - a)) ` x = ((*) (b - a)) ` y" by (simp add: image_comp) then have "(*) (inverse(b - a)) ` (*) (b - a) ` x = (*) (inverse(b - a)) ` (*) (b - a) ` y" by simp then show "x = y" using assms by (simp add: image_comp mult_ac) qed then show ?thesis by (metis card_image card_segments regular_division_def) qed lemma Union_regular_division: assumes "a \ b" shows "\(regular_division a b n) = (if n=0 then {} else {a..b})" using assms by (auto simp: regular_division_def Union_segments translate_scale_01 simp flip: image_Union) lemma regular_division_eqI: assumes K: "K = {a + (b-a)*(real k / n) .. a + (b-a)*((1 + real k) / n)}" and "a regular_division a b n" unfolding regular_division_def segments_def image_comp proof have "K = (\x. (b-a) * x + a) ` {real k / real n..(1 + real k) / real n}" using K \a by (simp add: image_affinity_atLeastAtMost divide_simps) then show "K = ((`) ((+) a \ (*) (b - a)) \ segment n) k" by (simp add: segment_def add.commute) qed (use assms in auto) lemma regular_divisionE: assumes "K \ regular_division a b n" "ax. a + (b - a) * x) = (\x. a + x) \ (\x. (b - a) * x)" by (simp add: o_def) obtain k where "kx. a+x) \ (\x. (b-a) * x)) ` {k/n .. (1 + real k) / n}" using assms by (auto simp: regular_division_def segments_def segment_def) with that \a show ?thesis unfolding image_comp [symmetric] by auto qed lemma regular_division_division_of: assumes "a < b" "n>0" shows "(regular_division a b n) division_of {a..b}" proof (rule division_ofI) show "finite (regular_division a b n)" by (simp add: regular_division_def segments_def) show \
: "\ (regular_division a b n) = {a..b}" using Union_regular_division assms by simp fix K assume K: "K \ regular_division a b n" then obtain k where Keq: "K = {a + (b-a)*(k/n) .. a + (b-a)*((1 + real k) / n)}" using \a regular_divisionE by meson show "K \ {a..b}" using K Union_regular_division \n>0\ by (metis Union_upper \
) show "K \ {}" using K by (auto simp: regular_division_def segment_nonempty segments_def) show "\a b. K = cbox a b" by (metis K \a box_real(2) regular_divisionE) fix K' assume K': "K' \ regular_division a b n" and "K \ K'" then obtain k' where Keq': "K' = {a + (b-a)*(k'/n) .. a + (b-a)*((1 + real k') / n)}" using K \a regular_divisionE by meson consider "1 + real k \ k'" | "1 + real k' \ k" using Keq Keq' \K \ K'\ by force then show "interior K \ interior K' = {}" proof cases case 1 then show ?thesis by (simp add: Keq Keq' min_def max_def divide_right_mono assms) next case 2 then have "interior K' \ interior K = {}" by (simp add: Keq Keq' min_def max_def divide_right_mono assms) then show ?thesis by (simp add: inf_commute) qed qed subsection \Special cases of Young's inequality\ lemma weighted_nesting_sum: fixes g :: "nat \ 'a::comm_ring_1" shows "(\ki real" - assumes sm: "strict_mono_on f {0..}" and cont: "continuous_on {0..} f" and a: "a\0" + assumes sm: "strict_mono_on {0..} f" and cont: "continuous_on {0..} f" and a: "a\0" and f: "f 0 = 0" "f a = b" and g: "\x. \0 \ x; x \ a\ \ g (f x) = x" shows "a*b = integral {0..a} f + integral {0..b} g" proof (cases "a=0") case False with \a \ 0\ have "a > 0" by linarith then have "b \ 0" by (smt (verit, best) atLeast_iff f sm strict_mono_onD) - have sm_0a: "strict_mono_on f {0..a}" + have sm_0a: "strict_mono_on {0..a} f" by (metis atLeastAtMost_iff atLeast_iff sm strict_mono_on_def) have cont_0a: "continuous_on {0..a} f" using cont continuous_on_subset by fastforce with sm_0a have "continuous_on {0..b} g" by (metis a atLeastAtMost_iff compact_Icc continuous_on_inv f g strict_mono_image_endpoints) then have intgb_g: "g integrable_on {0..b}" using integrable_continuous_interval by blast have intgb_f: "f integrable_on {0..a}" using cont_0a integrable_continuous_real by blast have f_iff [simp]: "f x < f y \ x < y" "f x \ f y \ x \ y" if "x \ 0" "y \ 0" for x y using that by (smt (verit, best) atLeast_iff sm strict_mono_onD)+ have fim: "f ` {0..a} = {0..b}" by (simp add: \a \ 0\ cont_0a strict_mono_image_endpoints strict_mono_on_def f) have "uniformly_continuous_on {0..a} f" using compact_uniformly_continuous cont_0a by blast then obtain del where del_gt0: "\e. e>0 \ del e > 0" and del: "\e x x'. \\x'-x\ < del e; e>0; x \ {0..a}; x' \ {0..a}\ \ \f x' - f x\ < e" unfolding uniformly_continuous_on_def dist_real_def by metis have *: "\a * b - integral {0..a} f - integral {0..b} g\ < 2*\" if "\ > 0" for \ proof - define \ where "\ = min a (del (\/a)) / 2" have "\ > 0" "\ \ a" using \a > 0\ \\ > 0\ del_gt0 by (auto simp: \_def) define n where "n \ nat\a / \\" define a_seg where "a_seg \ \u::real. u * a/n" have "n > 0" using \a > 0\ \\ > 0\ \\ \ a\ by (simp add: n_def) have a_seg_ge_0 [simp]: "a_seg x \ 0 \ x \ 0" and a_seg_le_a [simp]: "a_seg x \ a \ x \ n" for x using \n > 0\ \a > 0\ by (auto simp: a_seg_def zero_le_mult_iff divide_simps) have a_seg_le_iff [simp]: "a_seg x \ a_seg y \ x \ y" and a_seg_less_iff [simp]: "a_seg x < a_seg y \ x < y" for x y using \n > 0\ \a > 0\ by (auto simp: a_seg_def zero_le_mult_iff divide_simps) have "strict_mono a_seg" by (simp add: strict_mono_def) have a_seg_eq_a_iff: "a_seg x = a \ x=n" for x using \0 < n\ \a > 0\ by (simp add: a_seg_def nonzero_divide_eq_eq) have fa_eq_b: "f (a_seg n) = b" using a_seg_eq_a_iff f by fastforce have "a/d < real_of_int \a * 2 / min a d\" if "d>0" for d by (smt (verit) \0 < \\ \\ \ a\ add_divide_distrib divide_less_eq_1_pos floor_eq_iff that) then have an_less_del: "a/n < del (\/a)" using \a > 0\ \\ > 0\ del_gt0 by (simp add: n_def \_def field_simps) define lower where "lower \ \x. a_seg\(real n * x) / a\" define f1 where "f1 \ f \ lower" have f1_lower: "f1 x \ f x" if "0 \ x" "x \ a" for x proof - have "lower x \ x" using \n > 0\ floor_divide_lower [OF \a > 0\] by (auto simp: lower_def a_seg_def field_simps) moreover have "lower x \ 0" unfolding lower_def using \n > 0\ \a \ 0\ \0 \ x\ by force ultimately show ?thesis using sm strict_mono_on_leD by (fastforce simp add: f1_def) qed define upper where "upper \ \x. a_seg\real n * x / a\" define f2 where "f2 \ f \ upper" have f2_upper: "f2 x \ f x" if "0 \ x" "x \ a" for x proof - have "x \ upper x" using \n > 0\ ceiling_divide_upper [OF \a > 0\] by (simp add: upper_def a_seg_def field_simps) then show ?thesis using sm strict_mono_on_leD \0 \ x\ by (force simp: f2_def) qed let ?\ = "regular_division 0 a n" have div: "?\ division_of {0..a}" using \a > 0\ \n > 0\ regular_division_division_of zero_less_nat_eq by presburger have int_f1_D: "(f1 has_integral f(Inf K) * (a/n)) K" and int_f2_D: "(f2 has_integral f(Sup K) * (a/n)) K" and less: "\f(Sup K) - f(Inf K)\ < \/a" if "K\?\" for K proof - from regular_divisionE [OF that] \a > 0\ obtain k where "k a_seg k" define v where "v \ a_seg (Suc k)" have "u < v" "0 \ u" "0 \ v" "u \ a" "v \ a" and Kuv: "K = {u..v}" using \n > 0\ \k < n\ \a > 0\ by (auto simp: k u_def v_def divide_simps) have InfK: "Inf K = u" and SupK: "Sup K = v" using Kuv \u < v\ apply force using \n > 0\ \a > 0\ by (auto simp: divide_right_mono k u_def v_def) have f1: "f1 x = f (Inf K)" if "x \ K - {v}" for x proof - have "x \ {u..real_of_int n * x / a\ = int k" using \n > 0\ \a > 0\ by (simp add: field_simps u_def v_def a_seg_def floor_eq_iff) then show ?thesis by (simp add: InfK f1_def lower_def a_seg_def mult.commute u_def) qed have "((\x. f (Inf K)) has_integral (f (Inf K) * (a/n))) K" using has_integral_const_real [of "f (Inf K)" u v] \n > 0\ \a > 0\ by (simp add: Kuv field_simps a_seg_def u_def v_def) then show "(f1 has_integral (f (Inf K) * (a/n))) K" using has_integral_spike_finite_eq [of "{v}" K "\x. f (Inf K)" f1] f1 by simp have f2: "f2 x = f (Sup K)" if "x \ K - {u}" for x proof - have "x \ {u<..v}" using that Kuv greaterThanAtMost_eq_atLeastAtMost_diff by blast then have "\x * real_of_int n / a\ = 1 + int k" using \n > 0\ \a > 0\ by (simp add: field_simps u_def v_def a_seg_def ceiling_eq_iff) then show ?thesis by (simp add: mult.commute f2_def upper_def a_seg_def SupK v_def) qed have "((\x. f (Sup K)) has_integral (f (Sup K) * (a/n))) K" using \n > 0\ \a > 0\ has_integral_const_real [of "f (Sup K)" u v] by (simp add: Kuv field_simps u_def v_def a_seg_def) then show "(f2 has_integral (f (Sup K) * (a/n))) K" using has_integral_spike_finite_eq [of "{u}" K "\x. f (Sup K)" f2] f2 by simp have "\v - u\ < del (\/a)" using \n > 0\ \a > 0\ by (simp add: v_def u_def a_seg_def field_simps an_less_del) then have "\f v - f u\ < \/a" using \\ > 0\ \a > 0\ \0 \ u\ \u \ a\ \0 \ v\ \v \ a\ by (intro del) auto then show "\f(Sup K) - f(Inf K)\ < \/a" using InfK SupK by blast qed have int_21_D: "((\x. f2 x - f1 x) has_integral (f(Sup K) - f(Inf K)) * (a/n)) K" if "K\?\" for K using that has_integral_diff [OF int_f2_D int_f1_D] by (simp add: algebra_simps) have D_ne: "?\ \ {}" by (metis \0 < a\ \n > 0\ card_gt_0_iff card_regular_division) have f12: "((\x. f2 x - f1 x) has_integral (\K\?\. (f(Sup K) - f(Inf K)) * (a/n))) {0..a}" by (intro div int_21_D has_integral_combine_division) moreover have "(\K\?\. (f(Sup K) - f(Inf K)) * (a/n)) < \" proof - have "(\K\?\. (f(Sup K) - f(Inf K)) * (a/n)) \ (\K\?\. \f(Sup K) - f(Inf K)\ * (a/n))" using \n > 0\ \a > 0\ by (smt (verit) divide_pos_pos of_nat_0_less_iff sum_mono zero_le_mult_iff) also have "\ < (\K\?\. \/n)" using \n > 0\ \a > 0\ less by (intro sum_strict_mono finite_regular_division D_ne) (simp add: field_simps) also have "\ = \" using \n > 0\ \a > 0\ by simp finally show ?thesis . qed ultimately have f2_near_f1: "integral {0..a} (\x. f2 x - f1 x) < \" by (simp add: integral_unique) define yidx where "yidx \ \y. LEAST k. y < f (a_seg (Suc k))" have fa_yidx_le: "f (a_seg (yidx y)) \ y" and yidx_gt: "y < f (a_seg (Suc (yidx y)))" if "y \ {0..b}" for y proof - obtain x where x: "f x = y" "x \ {0..a}" using Topological_Spaces.IVT' [OF _ _ _ cont_0a] assms by (metis \y \ {0..b}\ atLeastAtMost_iff) define k where "k \ nat \x/a * n\" have x_lims: "a_seg k \ x" "x < a_seg (Suc k)" using \n > 0\ \0 < a\ floor_divide_lower floor_divide_upper [of a "x*n"] x by (auto simp: k_def a_seg_def field_simps) with that x obtain f_lims: "f (a_seg k) \ y" "y < f (a_seg (Suc k))" using strict_mono_onD [OF sm] by force then have "a_seg (yidx y) \ a_seg k" by (simp add: Least_le \strict_mono a_seg\ strict_mono_less_eq yidx_def) then have "f (a_seg (yidx y)) \ f (a_seg k)" using strict_mono_onD [OF sm] by simp then show "f (a_seg (yidx y)) \ y" using f_lims by linarith show "y < f (a_seg (Suc (yidx y)))" by (metis LeastI f_lims(2) yidx_def) qed have yidx_equality: "yidx y = k" if "y \ {0..b}" "y \ {f (a_seg k).. k" unfolding yidx_def by (metis atLeastLessThan_iff that(2) Least_le) have "(a_seg (real k)) < a_seg (1 + real (yidx y))" using yidx_gt [OF that(1)] that(2) strict_mono_onD [OF sm] order_le_less_trans by fastforce then have "real k < 1 + real (yidx y)" by (simp add: \strict_mono a_seg\ strict_mono_less) then show "k \ yidx y" by simp qed have "yidx b = n" proof - have "a < (1 + real n) * a / real n" using \0 < n\ \0 < a\ by (simp add: divide_simps) then have "b < f (a_seg (1 + real n))" using f \a \ 0\ a_seg_def sm strict_mono_onD by fastforce then show ?thesis using \0 \ b\ by (auto simp: f a_seg_def yidx_equality) qed moreover have yidx_less_n: "yidx y < n" if "y < b" for y by (metis \0 < n\ fa_eq_b gr0_conv_Suc less_Suc_eq_le that Least_le yidx_def) ultimately have yidx_le_n: "yidx y \ n" if "y \ b" for y by (metis dual_order.order_iff_strict that) have zero_to_b_eq: "{0..b} = (\k ?rhs" proof fix y assume y: "y \ {0..b}" have fn: "f (a_seg n) = b" using a_seg_eq_a_iff \f a = b\ by fastforce show "y \ ?rhs" proof (cases "y=b") case True with fn \n>0\ show ?thesis by (rule_tac a="n-1" in UN_I) auto next case False with y show ?thesis apply (simp add: subset_iff Bex_def) by (metis atLeastAtMost_iff of_nat_Suc order_le_less yidx_gt fa_yidx_le yidx_less_n) qed qed show "?rhs \ ?lhs" apply clarsimp by (smt (verit, best) a_seg_ge_0 a_seg_le_a f f_iff(2) nat_less_real_le of_nat_0_le_iff) qed define g1 where "g1 \ \y. if y=b then a else a_seg (Suc (yidx y))" define g2 where "g2 \ \y. if y=0 then 0 else a_seg (yidx y)" have g1: "g1 y \ {0..a}" if "y \ {0..b}" for y using that \a > 0\ yidx_less_n [of y] by (auto simp: g1_def a_seg_def divide_simps) have g2: "g2 y \ {0..a}" if "y \ {0..b}" for y using that \a > 0\ yidx_le_n [of y] by (simp add: g2_def a_seg_def divide_simps) have g2_le_g: "g2 y \ g y" if "y \ {0..b}" for y proof - have "f (g2 y) \ y" using \f 0 = 0\ g2_def that fa_yidx_le by presburger then have "f (g2 y) \ f (g y)" using that g by (smt (verit, best) atLeastAtMost_iff fim image_iff) then show ?thesis by (smt (verit, best) atLeastAtMost_iff fim g g2 imageE sm_0a strict_mono_onD that) qed have g_le_g1: "g y \ g1 y" if "y \ {0..b}" for y proof - have "y \ f (g1 y)" by (smt (verit, best) \f a = b\ g1_def that yidx_gt) then have "f (g y) \ f (g1 y)" using that g by (smt (verit, best) atLeastAtMost_iff fim image_iff) then show ?thesis by (smt (verit, ccfv_threshold) atLeastAtMost_iff f_iff(1) g1 that) qed define DN where "DN \ \K. nat \Inf K * real n / a\" have [simp]: "DN {a * real k / n..a * (1 + real k) / n} = k" for k using \n > 0\ \a > 0\ by (simp add: DN_def divide_simps) have DN: "bij_betw DN ?\ {.. regular_division 0 a n" with \a > 0\ obtain k where k: "K = {a * (real k / n) .. a * (1 + real k) / n}" by (force elim: regular_divisionE) assume "K' \ regular_division 0 a n" with \a > 0\ obtain k' where k': "K' = {a * (real k' / n) .. a * (1 + real k') / n}" by (force elim: regular_divisionE) assume "DN K = DN K'" then show "K = K'" by (simp add: k k') qed have "\K\regular_division 0 a n. k = nat \Inf K * real n / a\" if "k < n" for k using \n > 0\ \a > 0\ that by (force simp: divide_simps intro: regular_division_eqI [OF refl]) with \a>0\ show "DN ` regular_division 0 a n = {..k ?\" for K using that \a>0\ by (auto simp: DN_def field_simps a_seg_def elim: regular_divisionE) then have "(\K\?\. f(Inf K) * (a/n)) = (\kK\?\. f(Inf K) * (a/n))) {0..a}" by (intro div int_f1_D has_integral_combine_division) ultimately show ?thesis by (metis sum_distrib_right) qed text \The claim @{term "(f2 has_integral (\k have int_g1_D: "(g1 has_integral a_seg (Suc k) * (f (a_seg (Suc k)) - f (a_seg k))) {f(a_seg k)..f(a_seg (Suc k))}" and int_g2_D: "(g2 has_integral a_seg k * (f (a_seg (Suc k)) - f (a_seg k))) {f(a_seg k)..f(a_seg (Suc k))}" if "k < n" for k proof - define u where "u \ f (a_seg k)" define v where "v \ f (a_seg (Suc k))" obtain "u < v" "0 \ u" "0 \ v" unfolding u_def v_def assms by (smt (verit, best) a_seg_ge_0 a_seg_le_iff f(1) f_iff(1) of_nat_0_le_iff of_nat_Suc) have "u \ b" "v \ b" using \k < n\ \a \ 0\ by (simp_all add: u_def v_def flip: \f a = b\) have yidx_eq: "yidx x = k" if "x \ {u..0 \ u\ \v \ b\ that u_def v_def yidx_equality by auto have "g1 x = a_seg (Suc k)" if "x \ {u..v \ b\ by (simp add: g1_def yidx_eq) moreover have "((\x. a_seg (Suc k)) has_integral (a_seg (Suc k) * (v-u))) {u..v}" using has_integral_const_real \u < v\ by (metis content_real_if less_eq_real_def mult.commute real_scaleR_def) ultimately show "(g1 has_integral (a_seg (Suc k) * (v-u))) {u..v}" using has_integral_spike_finite_eq [of "{v}" "{u..v}" "\x. a_seg (Suc k)" g1] by simp have g2: "g2 x = a_seg k" if "x \ {u<..0 \ u\ by (simp add: g2_def yidx_eq) moreover have "((\x. a_seg k) has_integral (a_seg k * (v-u))) {u..v}" using has_integral_const_real \u < v\ by (metis content_real_if less_eq_real_def mult.commute real_scaleR_def) ultimately show "(g2 has_integral (a_seg k * (v-u))) {u..v}" using has_integral_spike_finite_eq [of "{u,v}" "{u..v}" "\x. a_seg k" g2] by simp qed have int_g1: "(g1 has_integral (\kkkk = (n * f (a_seg n) - (\k = a * b - (\kn > 0\ by (simp add: fa_eq_b field_simps) finally have int_g1': "(g1 has_integral a * b - (\kThe claim @{term "(g2 has_integral a * b - (\k have a_seg_diff: "a_seg (Suc k) - a_seg k = a/n" for k by (simp add: a_seg_def field_split_simps) have f_a_seg_diff: "\f (a_seg (Suc k)) - f (a_seg k)\ < \/a" if "ka > 0\ a_seg_diff an_less_del \\ > 0\ by (intro del) auto have "((\x. g1 x - g2 x) has_integral (\kk" proof - have "(\k (\kf (a_seg (Suc k)) - f (a_seg k)\ * (a/n))" by simp also have "\ < (\k/a) * (a/n))" proof (rule sum_strict_mono) fix k assume "k \ {..n > 0\ \a > 0\ divide_strict_right_mono f_a_seg_diff pos_less_divide_eq show "\f (a_seg (Suc k)) - f (a_seg k)\ * (a/n) < \/a * (a/n)" by fastforce qed (use \n > 0\ in auto) also have "\ = \" using \n > 0\ \a > 0\ by simp finally show ?thesis . qed ultimately have g2_near_g1: "integral {0..b} (\x. g1 x - g2 x) < \" by (simp add: integral_unique) have ab1: "integral {0..a} f1 + integral {0..b} g1 = a*b" using int_f1 int_g1' by (simp add: integral_unique) have "integral {0..a} (\x. f x - f1 x) \ integral {0..a} (\x. f2 x - f1 x)" proof (rule integral_le) show "(\x. f x - f1 x) integrable_on {0..a}" "(\x. f2 x - f1 x) integrable_on {0..a}" using Henstock_Kurzweil_Integration.integrable_diff int_f1 intgb_f f12 by blast+ qed (auto simp: f2_upper) with f2_near_f1 have "integral {0..a} (\x. f x - f1 x) < \" by simp moreover have "integral {0..a} f1 \ integral {0..a} f" by (intro integral_le has_integral_integral intgb_f has_integral_integrable [OF int_f1]) (simp add: f1_lower) ultimately have f_error: "\integral {0..a} f - integral {0..a} f1\ < \" using Henstock_Kurzweil_Integration.integral_diff int_f1 intgb_f by fastforce have "integral {0..b} (\x. g1 x - g x) \ integral {0..b} (\x. g1 x - g2 x)" proof (rule integral_le) show "(\x. g1 x - g x) integrable_on {0..b}" "(\x. g1 x - g2 x) integrable_on {0..b}" using Henstock_Kurzweil_Integration.integrable_diff int_g1 int_g2 intgb_g by blast+ qed (auto simp: g2_le_g) with g2_near_g1 have "integral {0..b} (\x. g1 x - g x) < \" by simp moreover have "integral {0..b} g \ integral {0..b} g1" by (intro integral_le has_integral_integral intgb_g has_integral_integrable [OF int_g1]) (simp add: g_le_g1) ultimately have g_error: "\integral {0..b} g1 - integral {0..b} g\ < \" using integral_diff int_g1 intgb_g by fastforce show ?thesis using f_error g_error ab1 by linarith qed show ?thesis using * [of "\a * b - integral {0..a} f - integral {0..b} g\ / 2"] by fastforce qed (use assms in force) corollary Youngs_strict: fixes f :: "real \ real" - assumes sm: "strict_mono_on f {0..}" and cont: "continuous_on {0..} f" and "a>0" "b\0" + assumes sm: "strict_mono_on {0..} f" and cont: "continuous_on {0..} f" and "a>0" "b\0" and f: "f 0 = 0" "f a \ b" and fim: "f ` {0..} = {0..}" and g: "\x. 0 \ x \ g (f x) = x" shows "a*b < integral {0..a} f + integral {0..b} g" proof - have f_iff [simp]: "f x < f y \ x < y" "f x \ f y \ x \ y" if "x \ 0" "y \ 0" for x y using that by (smt (verit, best) atLeast_iff sm strict_mono_onD)+ let ?b' = "f a" have "?b' \ 0" by (smt (verit, best) \0 < a\ atLeast_iff f sm strict_mono_onD) - then have sm_gx: "strict_mono_on g {0..}" + then have sm_gx: "strict_mono_on {0..} g" unfolding strict_mono_on_def by (smt (verit, best) atLeast_iff f_iff(1) f_inv_into_f fim g inv_into_into) show ?thesis proof (cases "?b' < b") case True have gt_a: "a < g y" if "y \ {?b'<..b}" for y proof - have "a = g ?b'" using \a > 0\ g by force also have "\ < g y" using \0 \ ?b'\ sm_gx strict_mono_onD that by fastforce finally show ?thesis . qed have "continuous_on {0..} g" by (metis cont f(1) fim g sm strict_mono_continuous_invD) then have contg: "continuous_on {?b'..b} g" by (meson Icc_subset_Ici_iff \0 \ f a\ continuous_on_subset) - have "mono_on g {0..}" + have "mono_on {0..} g" by (simp add: sm_gx strict_mono_on_imp_mono_on) then have int_g0b: "g integrable_on {0..b}" by (simp add: integrable_on_mono_on mono_on_subset) then have int_gb'b: "g integrable_on {?b'..b}" by (simp add: \0 \ ?b'\ integrable_on_subinterval) have "a * (b - ?b') = integral {?b'..b} (\y. a)" using True by force also have "\ < integral {?b'..b} g" using contg True gt_a by (intro integral_less_real) auto finally have *: "a * (b - ?b') < integral {?b'..b} g" . have "a*b = a * ?b' + a * (b - ?b')" by (simp add: algebra_simps) also have "\ = integral {0..a} f + integral {0..?b'} g + a * (b - ?b')" using Youngs_exact \a > 0\ cont \f 0 = 0\ g sm by force also have "\ < integral {0..a} f + integral {0..?b'} g + integral {?b'..b} g" by (simp add: *) also have "\ = integral {0..a} f + integral {0..b} g" by (smt (verit) Henstock_Kurzweil_Integration.integral_combine True \0 \ ?b'\ int_g0b) finally show ?thesis . next case False with f have "b < ?b'" by force obtain a' where "f a' = b" "a' \ 0" using fim \b \ 0\ by force then have "a' < a" using \b < f a\ \a > 0\ by force have gt_b: "b < f x" if "x \ {a'<..a}" for x using \0 \ a'\ \f a' = b\ that by fastforce have int_f0a: "f integrable_on {0..a}" by (simp add: integrable_on_mono_on mono_on_def) then have int_fa'a: "f integrable_on {a'..a}" by (simp add: \0 \ a'\ integrable_on_subinterval) have cont_f': "continuous_on {a'..a} f" by (meson Icc_subset_Ici_iff \0 \ a'\ cont continuous_on_subset) have "b * (a - a') = integral {a'..a} (\x. b)" using \a' < a\ by simp also have "\ < integral {a'..a} f" using cont_f' \a' < a\ gt_b by (intro integral_less_real) auto finally have *: "b * (a - a') < integral {a'..a} f" . have "a*b = a' * b + b * (a - a')" by (simp add: algebra_simps) also have "\ = integral {0..a'} f + integral {0..b} g + b * (a - a')" by (simp add: Youngs_exact \0 \ a'\ \f a' = b\ cont f g sm) also have "\ < integral {0..a'} f + integral {0..b} g + integral {a'..a} f" by (simp add: *) also have "\ = integral {0..a} f + integral {0..b} g" by (smt (verit) Henstock_Kurzweil_Integration.integral_combine \0 \ a'\ \a' < a\ int_f0a) finally show ?thesis . qed qed corollary Youngs_inequality: fixes f :: "real \ real" - assumes sm: "strict_mono_on f {0..}" and cont: "continuous_on {0..} f" and "a\0" "b\0" + assumes sm: "strict_mono_on {0..} f" and cont: "continuous_on {0..} f" and "a\0" "b\0" and f: "f 0 = 0" and fim: "f ` {0..} = {0..}" and g: "\x. 0 \ x \ g (f x) = x" shows "a*b \ integral {0..a} f + integral {0..b} g" proof (cases "a=0") case True have "g x \ 0" if "x \ 0" for x by (metis atLeast_iff fim g imageE that) then have "0 \ integral {0..b} g" by (metis Henstock_Kurzweil_Integration.integral_nonneg atLeastAtMost_iff not_integrable_integral order_refl) then show ?thesis by (simp add: True) next case False then show ?thesis by (smt (verit) assms Youngs_exact Youngs_strict) qed end diff --git a/thys/ZFC_in_HOL/ZFC_Cardinals.thy b/thys/ZFC_in_HOL/ZFC_Cardinals.thy --- a/thys/ZFC_in_HOL/ZFC_Cardinals.thy +++ b/thys/ZFC_in_HOL/ZFC_Cardinals.thy @@ -1,2708 +1,2708 @@ section \Cartesian products, Disjoint Sums, Ranks, Cardinals\ theory ZFC_Cardinals imports ZFC_in_HOL begin declare [[coercion_enabled]] declare [[coercion "ord_of_nat :: nat \ V"]] subsection \Ordered Pairs\ lemma singleton_eq_iff [iff]: "set {a} = set {b} \ a=b" by simp lemma doubleton_eq_iff: "set {a,b} = set {c,d} \ (a=c \ b=d) \ (a=d \ b=c)" by (simp add: Set.doubleton_eq_iff) definition vpair :: "V \ V \ V" where "vpair a b = set {set {a},set {a,b}}" definition vfst :: "V \ V" where "vfst p \ THE x. \y. p = vpair x y" definition vsnd :: "V \ V" where "vsnd p \ THE y. \x. p = vpair x y" definition vsplit :: "[[V, V] \ 'a, V] \ 'a::{}" \ \for pattern-matching\ where "vsplit c \ \p. c (vfst p) (vsnd p)" nonterminal Vs syntax (ASCII) "_Tuple" :: "[V, Vs] \ V" ("<(_,/ _)>") "_hpattern" :: "[pttrn, patterns] \ pttrn" ("<_,/ _>") syntax "" :: "V \ Vs" ("_") "_Enum" :: "[V, Vs] \ Vs" ("_,/ _") "_Tuple" :: "[V, Vs] \ V" ("\(_,/ _)\") "_hpattern" :: "[pttrn, patterns] \ pttrn" ("\_,/ _\") translations "" \ ">" "" \ "CONST vpair x y" "" \ ">" "\. b" \ "CONST vsplit(\x . b)" "\. b" \ "CONST vsplit(\x y. b)" lemma vpair_def': "vpair a b = set {set {a,a},set {a,b}}" by (simp add: vpair_def) lemma vpair_iff [simp]: "vpair a b = vpair a' b' \ a=a' \ b=b'" unfolding vpair_def' doubleton_eq_iff by auto lemmas vpair_inject = vpair_iff [THEN iffD1, THEN conjE, elim!] lemma vfst_conv [simp]: "vfst \a,b\ = a" by (simp add: vfst_def) lemma vsnd_conv [simp]: "vsnd \a,b\ = b" by (simp add: vsnd_def) lemma vsplit [simp]: "vsplit c \a,b\ = c a b" by (simp add: vsplit_def) lemma vpair_neq_fst: "\a,b\ \ a" by (metis elts_of_set insertI1 mem_not_sym small_upair vpair_def') lemma vpair_neq_snd: "\a,b\ \ b" by (metis elts_of_set insertI1 mem_not_sym small_upair subsetD subset_insertI vpair_def') lemma vpair_nonzero [simp]: "\x,y\ \ 0" by (metis elts_0 elts_of_set empty_not_insert small_upair vpair_def) lemma zero_notin_vpair: "0 \ elts \x,y\" by (auto simp: vpair_def) lemma inj_on_vpair [simp]: "inj_on (\(x, y). \x, y\) A" by (auto simp: inj_on_def) subsection \Generalized Cartesian product\ definition VSigma :: "V \ (V \ V) \ V" where "VSigma A B \ set(\x \ elts A. \y \ elts (B x). {\x,y\})" abbreviation vtimes where "vtimes A B \ VSigma A (\x. B)" definition pairs :: "V \ (V * V)set" where "pairs r \ {(x,y). \x,y\ \ elts r} " lemma pairs_iff_elts: "(x,y) \ pairs z \ \x,y\ \ elts z" by (simp add: pairs_def) lemma VSigma_iff [simp]: "\a,b\ \ elts (VSigma A B) \ a \ elts A \ b \ elts (B a)" by (auto simp: VSigma_def UNION_singleton_eq_range) lemma VSigmaI [intro!]: "\ a \ elts A; b \ elts (B a)\ \ \a,b\ \ elts (VSigma A B)" by simp lemmas VSigmaD1 = VSigma_iff [THEN iffD1, THEN conjunct1] lemmas VSigmaD2 = VSigma_iff [THEN iffD1, THEN conjunct2] text \The general elimination rule\ lemma VSigmaE [elim!]: assumes "c \ elts (VSigma A B)" obtains x y where "x \ elts A" "y \ elts (B x)" "c=\x,y\" using assms by (auto simp: VSigma_def split: if_split_asm) lemma VSigmaE2 [elim!]: assumes "\a,b\ \ elts (VSigma A B)" obtains "a \ elts A" and "b \ elts (B a)" using assms by auto lemma VSigma_empty1 [simp]: "VSigma 0 B = 0" by auto lemma times_iff [simp]: "\a,b\ \ elts (vtimes A B) \ a \ elts A \ b \ elts B" by simp lemma timesI [intro!]: "\a \ elts A; b \ elts B\ \ \a,b\ \ elts (vtimes A B)" by simp lemma times_empty2 [simp]: "vtimes A 0 = 0" using elts_0 by blast lemma times_empty_iff: "VSigma A B = 0 \ A=0 \ (\x \ elts A. B x = 0)" by (metis VSigmaE VSigmaI elts_0 empty_iff trad_foundation) lemma elts_VSigma: "elts (VSigma A B) = (\(x,y). vpair x y) ` Sigma (elts A) (\x. elts (B x))" by auto lemma small_Sigma [simp]: assumes A: "small A" and B: "\x. x \ A \ small (B x)" shows "small (Sigma A B)" proof - obtain a where "elts a \ A" by (meson assms small_eqpoll) then obtain f where f: "bij_betw f (elts a) A" using eqpoll_def by blast have "\y. elts y \ B x" if "x \ A" for x using B small_eqpoll that by blast then obtain g where g: "\x. x \ A \ elts (g x) \ B x" by metis with f have "elts (VSigma a (g \ f)) \ Sigma A B" by (simp add: elts_VSigma Sigma_eqpoll_cong bij_betwE) then show ?thesis using small_eqpoll by blast qed lemma small_Times [simp]: assumes "small A" "small B" shows "small (A \ B)" by (simp add: assms) lemma small_Times_iff: "small (A \ B) \ small A \ small B \ A={} \ B={}" (is "_ = ?rhs") proof assume *: "small (A \ B)" { have "small A \ small B" if "x \ A" "y \ B" for x y proof - have "A \ fst ` (A \ B)" "B \ snd ` (A \ B)" using that by auto with that show ?thesis by (metis * replacement smaller_than_small) qed } then show ?rhs by (metis equals0I) next assume ?rhs then show "small (A \ B)" by auto qed subsection \Disjoint Sum\ definition vsum :: "V \ V \ V" (infixl "\" 65) where "A \ B \ (VSigma (set {0}) (\x. A)) \ (VSigma (set {1}) (\x. B))" definition Inl :: "V\V" where "Inl a \ \0,a\" definition Inr :: "V\V" where "Inr b \ \1,b\" lemmas sum_defs = vsum_def Inl_def Inr_def lemma Inl_nonzero [simp]:"Inl x \ 0" by (metis Inl_def vpair_nonzero) lemma Inr_nonzero [simp]:"Inr x \ 0" by (metis Inr_def vpair_nonzero) subsubsection\Equivalences for the injections and an elimination rule\ lemma Inl_in_sum_iff [iff]: "Inl a \ elts (A \ B) \ a \ elts A" by (auto simp: sum_defs) lemma Inr_in_sum_iff [iff]: "Inr b \ elts (A \ B) \ b \ elts B" by (auto simp: sum_defs) lemma sumE [elim!]: assumes u: "u \ elts (A \ B)" obtains x where "x \ elts A" "u=Inl x" | y where "y \ elts B" "u=Inr y" using u by (auto simp: sum_defs) subsubsection \Injection and freeness equivalences, for rewriting\ lemma Inl_iff [iff]: "Inl a=Inl b \ a=b" by (simp add: sum_defs) lemma Inr_iff [iff]: "Inr a=Inr b \ a=b" by (simp add: sum_defs) lemma inj_on_Inl [simp]: "inj_on Inl A" by (simp add: inj_on_def) lemma inj_on_Inr [simp]: "inj_on Inr A" by (simp add: inj_on_def) lemma Inl_Inr_iff [iff]: "Inl a=Inr b \ False" by (simp add: sum_defs) lemma Inr_Inl_iff [iff]: "Inr b=Inl a \ False" by (simp add: sum_defs) lemma sum_empty [simp]: "0 \ 0 = 0" by auto lemma elts_vsum: "elts (a \ b) = Inl ` (elts a) \ Inr ` (elts b)" by auto lemma sum_iff: "u \ elts (A \ B) \ (\x. x \ elts A \ u=Inl x) \ (\y. y \ elts B \ u=Inr y)" by blast lemma sum_subset_iff: "A\B \ C\D \ A\C \ B\D" by (auto simp: less_eq_V_def) lemma sum_equal_iff: fixes A :: V shows "A\B = C\D \ A=C \ B=D" by (simp add: eq_iff sum_subset_iff) definition is_sum :: "V \ bool" where "is_sum z = (\x. z = Inl x \ z = Inr x)" definition sum_case :: "(V \ 'a) \ (V \ 'a) \ V \ 'a" where "sum_case f g a \ THE z. (\x. a = Inl x \ z = f x) \ (\y. a = Inr y \ z = g y) \ (\ is_sum a \ z = undefined)" lemma sum_case_Inl [simp]: "sum_case f g (Inl x) = f x" by (simp add: sum_case_def is_sum_def) lemma sum_case_Inr [simp]: "sum_case f g (Inr y) = g y" by (simp add: sum_case_def is_sum_def) lemma sum_case_non [simp]: "\ is_sum a \ sum_case f g a = undefined" by (simp add: sum_case_def is_sum_def) lemma is_sum_cases: "(\x. z = Inl x \ z = Inr x) \ \ is_sum z" by (auto simp: is_sum_def) lemma sum_case_split: "P (sum_case f g a) \ (\x. a = Inl x \ P(f x)) \ (\y. a = Inr y \ P(g y)) \ (\ is_sum a \ P undefined)" by (cases "is_sum a") (auto simp: is_sum_def) lemma sum_case_split_asm: "P (sum_case f g a) \ \ ((\x. a = Inl x \ \ P(f x)) \ (\y. a = Inr y \ \ P(g y)) \ (\ is_sum a \ \ P undefined))" by (auto simp: sum_case_split) subsubsection \Applications of disjoint sums and pairs: general union theorems for small sets\ lemma small_Un: assumes X: "small X" and Y: "small Y" shows "small (X \ Y)" proof - obtain x y where "elts x \ X" "elts y \ Y" by (meson assms small_eqpoll) then have "X \ Y \ Inl ` (elts x) \ Inr ` (elts y)" by (metis (mono_tags, lifting) Inr_Inl_iff Un_lepoll_mono disjnt_iff eqpoll_imp_lepoll eqpoll_sym f_inv_into_f inj_on_Inl inj_on_Inr inj_on_image_lepoll_2) then show ?thesis by (metis lepoll_iff replacement small_elts small_sup_iff smaller_than_small) qed lemma small_UN [simp,intro]: assumes A: "small A" and B: "\x. x \ A \ small (B x)" shows "small (\x\A. B x)" proof - obtain a where "elts a \ A" by (meson assms small_eqpoll) then obtain f where f: "bij_betw f (elts a) A" using eqpoll_def by blast have "\y. elts y \ B x" if "x \ A" for x using B small_eqpoll that by blast then obtain g where g: "\x. x \ A \ elts (g x) \ B x" by metis have sm: "small (Sigma (elts a) (elts \ g \ f))" by simp have "(\x\A. B x) \ Sigma A B" by (metis image_lepoll snd_image_Sigma) also have "... \ Sigma (elts a) (elts \ g \ f)" by (smt (verit) Sigma_eqpoll_cong bij_betw_iff_bijections comp_apply eqpoll_imp_lepoll eqpoll_sym f g) finally show ?thesis using lepoll_small sm by blast qed lemma small_Union [simp,intro]: assumes "\ \ Collect small" "small \" shows "small (\ \)" using small_UN [of \ "\x. x"] assms by (simp add: subset_iff) subsection\Generalised function space and lambda\ definition VLambda :: "V \ (V \ V) \ V" where "VLambda A b \ set ((\x. \x,b x\) ` elts A)" definition app :: "[V,V] \ V" where "app f x \ THE y. \x,y\ \ elts f" lemma beta [simp]: assumes "x \ elts A" shows "app (VLambda A b) x = b x" using assms by (auto simp: VLambda_def app_def) definition VPi :: "V \ (V \ V) \ V" where "VPi A B \ set {f \ elts (VPow(VSigma A B)). elts A \ Domain (pairs f) \ single_valued (pairs f)}" lemma VPi_I: assumes "\x. x \ elts A \ b x \ elts (B x)" shows "VLambda A b \ elts (VPi A B)" proof (clarsimp simp: VPi_def, intro conjI impI) show "VLambda A b \ VSigma A B" by (auto simp: assms VLambda_def split: if_split_asm) show "elts A \ Domain (pairs (VLambda A b))" by (force simp: VLambda_def pairs_iff_elts) show "single_valued (pairs (VLambda A b))" by (auto simp: VLambda_def single_valued_def pairs_iff_elts) show "small {f. f \ VSigma A B \ elts A \ Domain (pairs f) \ single_valued (pairs f)}" by (metis (mono_tags, lifting) down VPow_iff mem_Collect_eq subsetI) qed lemma apply_pair: assumes f: "f \ elts (VPi A B)" and x: "x \ elts A" shows "\x, app f x\ \ elts f" proof - have "x \ Domain (pairs f)" by (metis (no_types, lifting) VPi_def assms elts_of_set empty_iff mem_Collect_eq subsetD) then obtain y where y: "\x,y\ \ elts f" using pairs_iff_elts by auto show ?thesis unfolding app_def proof (rule theI) show "\x, y\ \ elts f" by (rule y) show "z = y" if "\x, z\ \ elts f" for z using f unfolding VPi_def by (metis (mono_tags, lifting) that elts_of_set empty_iff mem_Collect_eq pairs_iff_elts single_valued_def y) qed qed lemma VPi_D: assumes f: "f \ elts (VPi A B)" and x: "x \ elts A" shows "app f x \ elts (B x)" proof - have "f \ VSigma A B" by (metis (no_types, lifting) VPi_def elts_of_set empty_iff f VPow_iff mem_Collect_eq) then show ?thesis using apply_pair [OF assms] by blast qed lemma VPi_memberD: assumes f: "f \ elts (VPi A B)" and p: "p \ elts f" obtains x where "x \ elts A" "p = \x, app f x\" proof - have "f \ VSigma A B" by (metis (no_types, lifting) VPi_def elts_of_set empty_iff f VPow_iff mem_Collect_eq) then obtain x y where "p = \x,y\" "x \ elts A" using p by blast then have "y = app f x" by (metis (no_types, lifting) VPi_def apply_pair elts_of_set equals0D f mem_Collect_eq p pairs_iff_elts single_valuedD) then show thesis using \p = \x, y\\ \x \ elts A\ that by blast qed lemma fun_ext: assumes "f \ elts (VPi A B)" "g \ elts (VPi A B)" "\x. x \ elts A \ app f x = app g x" shows "f = g" by (metis VPi_memberD V_equalityI apply_pair assms) lemma eta[simp]: assumes "f \ elts (VPi A B)" shows "VLambda A ((app)f) = f" proof (rule fun_ext [OF _ assms]) show "VLambda A (app f) \ elts (VPi A B)" using VPi_D VPi_I assms by auto qed auto lemma fst_pairs_VLambda: "fst ` pairs (VLambda A f) = elts A" by (force simp: VLambda_def pairs_def) lemma snd_pairs_VLambda: "snd ` pairs (VLambda A f) = f ` elts A" by (force simp: VLambda_def pairs_def) lemma VLambda_eq_D1: "VLambda A f = VLambda B g \ A = B" by (metis ZFC_in_HOL.ext fst_pairs_VLambda) lemma VLambda_eq_D2: "\VLambda A f = VLambda A g; x \ elts A\ \ f x = g x" by (metis beta) subsection\Transitive closure of a set\ definition TC :: "V\V" where "TC \ transrec (\f x. x \ \ (f ` elts x))" lemma TC: "TC a = a \ \ (TC ` elts a)" by (metis (no_types, lifting) SUP_cong TC_def restrict_apply' transrec) lemma TC_0 [simp]: "TC 0 = 0" by (metis TC ZFC_in_HOL.Sup_empty elts_0 image_is_empty sup_V_0_left) lemma arg_subset_TC: "a \ TC a" by (metis (no_types) TC sup_ge1) lemma Transset_TC: "Transset(TC a)" proof (induction a rule: eps_induct) case (step x) have 1: "v \ elts (TC x)" if "v \ elts u" "u \ elts x" for u v using that unfolding TC [of x] using arg_subset_TC by fastforce have 2: "v \ elts (TC x)" if "v \ elts u" "\x\elts x. u \ elts (TC x)" for u v using that step unfolding TC [of x] Transset_def by auto show ?case unfolding Transset_def by (subst TC) (force intro: 1 2) qed lemma TC_least: "\Transset x; a\x\ \ TC a \ x" proof (induction a rule: eps_induct) case (step y) show ?case proof (cases "y=0") case True then show ?thesis by auto next case False have "\ (TC ` elts y) \ x" proof (rule cSup_least) show "TC ` elts y \ {}" using False by auto show "z \ x" if "z \ TC ` elts y" for z using that by (metis Transset_def image_iff step.IH step.prems vsubsetD) qed then show ?thesis by (simp add: step TC [of y]) qed qed definition less_TC (infix "\" 50) where "x \ y \ x \ elts (TC y)" definition le_TC (infix "\" 50) where "x \ y \ x \ y \ x=y" lemma less_TC_imp_not_le: "x \ a \ \ a \ x" proof (induction a arbitrary: x rule: eps_induct) case (step a) then show ?case unfolding TC[of a] less_TC_def using Transset_TC Transset_def by force qed lemma non_TC_less_0 [iff]: "\ (x \ 0)" using less_TC_imp_not_le by blast lemma less_TC_iff: "x \ y \ (\z \ elts y. x \ z)" by (auto simp: less_TC_def le_TC_def TC [of y]) lemma nonzero_less_TC: "x \ 0 \ 0 \ x" by (metis eps_induct le_TC_def less_TC_iff trad_foundation) lemma less_irrefl_TC [simp]: "\ x \ x" using less_TC_imp_not_le by blast lemma less_asym_TC: "\x \ y; y \ x\ \ False" by (metis TC_least Transset_TC Transset_def antisym_conv less_TC_def less_TC_imp_not_le order_refl) lemma le_antisym_TC: "\x \ y; y \ x\ \ x = y" using less_asym_TC by blast lemma less_imp_le_TC [iff]: "x \ y \ x \ y" by (simp add: le_TC_def) lemma le_TC_refl [iff]: "x \ x" by (simp add: le_TC_def) lemma less_TC_trans [trans]: "\x \ y; y \ z\ \ x \ z" by (meson TC_least Transset_TC Transset_def less_TC_def less_eq_V_def subsetD) lemma less_le_TC_trans [trans]: "\x \ y; y \ z\ \ x \ z" using le_TC_def less_TC_trans by blast lemma le_less_TC_trans [trans]: "\x \ y; y \ z\ \ x \ z" using le_TC_def less_TC_trans by blast lemma le_TC_trans [trans]: "\x \ y; y \ z\ \ x \ z" using le_TC_def le_less_TC_trans by blast lemma TC_sup_distrib: "TC (x \ y) = TC x \ TC y" by (simp add: Sup_Un_distrib TC [of "x \ y"] TC [of x] TC [of y] image_Un sup.assoc sup_left_commute) lemma TC_Sup_distrib: assumes "small X" shows "TC (\X) = \(TC ` X)" proof - have "\X \ \ (TC ` X)" using arg_subset_TC by fastforce moreover have "\ (\x\X. TC ` elts x) \ \ (TC ` X)" using assms by clarsimp (meson TC_least Transset_TC Transset_def arg_subset_TC replacement vsubsetD) ultimately have "\ X \ \ (\x\X. TC ` elts x) \ \ (TC ` X)" by simp moreover have "\ (TC ` X) \ \ X \ \ (\x\X. TC ` elts x)" proof (clarsimp simp add: Sup_le_iff assms) show "\x\X. y \ elts x" if "x \ X" "y \ elts (TC x)" "\x\X. \u\elts x. y \ elts (TC u)" for x y using that by (auto simp: TC [of x]) qed ultimately show ?thesis using Sup_Un_distrib TC [of "\X"] image_Union assms by (simp add: image_Union inf_sup_aci(5) sup.absorb_iff2) qed lemma TC': "TC x = x \ TC (\ (elts x))" by (simp add: TC [of x] TC_Sup_distrib) lemma TC_eq_0_iff [simp]: "TC x = 0 \ x=0" using arg_subset_TC by fastforce text\A distinctive induction principle\ lemma TC_induct_down_lemma: assumes ab: "a \ b" and base: "b \ d" and step: "\y z. \y \ b; y \ elts d; z \ elts y\ \ z \ elts d" shows "a \ elts d" proof - have "Transset (TC b \ d)" using Transset_TC unfolding Transset_def by (metis inf.bounded_iff less_TC_def less_eq_V_def local.step subsetI vsubsetD) moreover have "b \ TC b \ d" by (simp add: arg_subset_TC base) ultimately show ?thesis using TC_least [THEN vsubsetD] ab unfolding less_TC_def by (meson TC_least le_inf_iff vsubsetD) qed lemma TC_induct_down [consumes 1, case_names base step small]: assumes "a \ b" and "\y. y \ elts b \ P y" and "\y z. \y \ b; P y; z \ elts y\ \ P z" and "small (Collect P)" shows "P a" using TC_induct_down_lemma [of a b "set (Collect P)"] assms by (metis elts_of_set mem_Collect_eq vsubsetI) subsection\Rank of a set\ definition rank :: "V\V" where "rank a \ transrec (\f x. set (\y\elts x. elts (succ(f y)))) a" lemma rank: "rank a = set(\y \ elts a. elts (succ(rank y)))" by (subst rank_def [THEN def_transrec], simp) lemma rank_Sup: "rank a = \((\y. succ(rank y)) ` elts a)" by (metis elts_Sup image_image rank replacement set_of_elts small_elts) lemma Ord_rank [simp]: "Ord(rank a)" proof (induction a rule: eps_induct) case (step x) then show ?case unfolding rank_Sup [of x] by (metis (mono_tags, lifting) Ord_Sup Ord_succ imageE) qed lemma rank_of_Ord: "Ord i \ rank i = i" by (induction rule: Ord_induct) (metis (no_types, lifting) Ord_equality SUP_cong rank_Sup) lemma Ord_iff_rank: "Ord x \ rank x = x" using Ord_rank [of x] rank_of_Ord by fastforce lemma rank_lt: "a \ elts b \ rank a < rank b" by (metis Ord_linear2 Ord_rank ZFC_in_HOL.SUP_le_iff rank_Sup replacement small_elts succ_le_iff order.irrefl) lemma rank_0 [simp]: "rank 0 = 0" using transrec Ord_0 rank_def rank_of_Ord by presburger lemma rank_succ [simp]: "rank(succ x) = succ(rank x)" proof (rule order_antisym) show "rank (succ x) \ succ (rank x)" by (metis (no_types, lifting) Sup_insert elts_of_set elts_succ image_insert rank small_UN small_elts subset_insertI sup.orderE vsubsetI) show "succ (rank x) \ rank (succ x)" by (metis (mono_tags, lifting) ZFC_in_HOL.Sup_upper elts_succ image_insert insertI1 rank_Sup replacement small_elts) qed lemma rank_mono: "a \ b \ rank a \ rank b" using rank [of a] rank [of b] small_UN by force lemma VsetI: "rank b \ i \ b \ elts (Vset i)" proof (induction i arbitrary: b rule: eps_induct) case (step x) then consider "rank b \ elts x" | "(\y\elts x. rank b \ elts (TC y))" using le_TC_def less_TC_def less_TC_iff by fastforce then have "\y\elts x. b \ Vset y" proof cases case 1 then have "b \ Vset (rank b)" unfolding less_eq_V_def subset_iff by (meson Ord_mem_iff_lt Ord_rank le_TC_refl less_TC_iff rank_lt step.IH) then show ?thesis using "1" by blast next case 2 then show ?thesis using step.IH unfolding less_eq_V_def subset_iff less_TC_def by (meson Ord_mem_iff_lt Ord_rank Transset_TC Transset_def rank_lt vsubsetD) qed then show ?case by (simp add: Vset [of x]) qed lemma Ord_VsetI: "\Ord i; rank b < i\ \ b \ elts (Vset i)" by (meson Ord_mem_iff_lt Ord_rank VsetI arg_subset_TC less_TC_def vsubsetD) lemma arg_le_Vset_rank: "a \ Vset(rank a)" by (simp add: Ord_VsetI rank_lt vsubsetI) lemma two_in_Vset: obtains \ where "x \ elts (Vset \)" "y \ elts (Vset \)" by (metis Ord_rank Ord_VsetI elts_of_set insert_iff rank_lt small_elts small_insert_iff) lemma rank_eq_0_iff [simp]: "rank x = 0 \ x=0" using arg_le_Vset_rank by fastforce lemma small_ranks_imp_small: assumes "small (rank ` A)" shows "small A" proof - define i where "i \ set (\(elts ` (rank ` A)))" have "Ord i" unfolding i_def using Ord_Union Ord_rank assms imageE by blast have *: "Vset (rank x) \ (Vset i)" if "x \ A" for x unfolding i_def by (metis Ord_rank Sup_V_def ZFC_in_HOL.Sup_upper Vfrom_mono assms imageI le_less that) have "A \ elts (VPow (Vset i))" by (meson "*" VPow_iff arg_le_Vset_rank order.trans subsetI) then show ?thesis using down by blast qed lemma rank_Union: "rank(\ A) = \ (rank ` A)" proof (rule order_antisym) have "elts (\y\elts (\ A). succ (rank y)) \ elts (\ (rank ` A))" by clarsimp (meson Ord_mem_iff_lt Ord_rank less_V_def rank_lt vsubsetD) then show "rank (\ A) \ \ (rank ` A)" by (metis less_eq_V_def rank_Sup) show "\ (rank ` A) \ rank (\ A)" proof (cases "small A") case True then show ?thesis by (simp add: ZFC_in_HOL.SUP_le_iff ZFC_in_HOL.Sup_upper rank_mono) next case False then have "\ small (rank ` A)" using small_ranks_imp_small by blast then show ?thesis by blast qed qed lemma small_bounded_rank: "small {x. rank x \ elts a}" proof - have "{x. rank x \ elts a} \ {x. rank x \ a}" using less_TC_iff by auto also have "\ \ elts (Vset a)" using VsetI by blast finally show ?thesis using down by simp qed lemma small_bounded_rank_le: "small {x. rank x \ a}" using small_bounded_rank [of "VPow a"] VPow_iff [of _ a] by simp lemma TC_rank_lt: "a \ b \ rank a < rank b" proof (induction rule: TC_induct_down) case (base y) then show ?case by (simp add: rank_lt) next case (step y z) then show ?case using less_trans rank_lt by blast next case small show ?case using smaller_than_small [OF small_bounded_rank_le [of "rank b"]] by (simp add: Collect_mono less_V_def) qed lemma TC_rank_mem: "x \ y \ rank x \ elts (rank y)" by (simp add: Ord_mem_iff_lt TC_rank_lt) lemma wf_TC_less: "wf {(x,y). x \ y}" proof (rule wf_subset [OF wf_inv_image [OF foundation, of rank]]) show "{(x, y). x \ y} \ inv_image {(x, y). x \ elts y} rank" by (auto simp: TC_rank_mem inv_image_def) qed lemma less_TC_minimal: assumes "P a" obtains x where "P x" "x \ a" "\y. y \ x \ \ P y" using wfE_min' [OF wf_TC_less, of "{x. P x \ x \ a}"] by simp (metis le_TC_def less_le_TC_trans assms) lemma Vfrom_rank_eq: "Vfrom A (rank(x)) = Vfrom A x" proof (rule order_antisym) show "Vfrom A (rank x) \ Vfrom A x" proof (induction x rule: eps_induct) case (step x) have "(\j\elts (rank x). VPow (Vfrom A j)) \ (\j\elts x. VPow (Vfrom A j))" apply (rule Sup_least) apply (clarsimp simp add: rank [of x]) by (meson Ord_in_Ord Ord_rank OrdmemD Vfrom_mono order.trans less_imp_le order.refl step) then show ?case by (simp add: Vfrom [of _ x] Vfrom [of _ "rank(x)"] sup.coboundedI2) qed show "Vfrom A x \ Vfrom A (rank x)" proof (induction x rule: eps_induct) case (step x) have "(\j\elts x. VPow (Vfrom A j)) \ (\j\elts (rank x). VPow (Vfrom A j))" using step.IH TC_rank_mem less_TC_iff by force then show ?case by (simp add: Vfrom [of _ x] Vfrom [of _ "rank(x)"] sup.coboundedI2) qed qed lemma Vfrom_succ: "Vfrom A (succ(i)) = A \ VPow(Vfrom A i)" by (metis Ord_rank Vfrom_rank_eq Vfrom_succ_Ord rank_succ) lemma Vset_succ_TC: assumes "x \ elts (Vset (ZFC_in_HOL.succ k))" "u \ x" shows "u \ elts (Vset k)" using assms using TC_least Transset_Vfrom Vfrom_succ less_TC_def by auto subsection\Cardinal Numbers\ text\We extend the membership relation to a wellordering\ definition VWO :: "(V \ V) set" where "VWO \ @r. {(x,y). x \ elts y} \ r \ Well_order r \ Field r = UNIV" lemma VWO: "{(x,y). x \ elts y} \ VWO \ Well_order VWO \ Field VWO = UNIV" unfolding VWO_def by (metis (mono_tags, lifting) VWO_def foundation someI_ex total_well_order_extension) lemma wf_VWO: "wf(VWO - Id)" using VWO well_order_on_def by blast lemma wf_Ord_less: "wf {(x, y). Ord y \ x < y}" by (metis (no_types, lifting) Ord_mem_iff_lt eps_induct wfPUNIVI wfP_def) lemma refl_VWO: "refl VWO" using VWO order_on_defs by fastforce lemma trans_VWO: "trans VWO" using VWO by (simp add: VWO wo_rel.TRANS wo_rel_def) lemma antisym_VWO: "antisym VWO" using VWO by (simp add: VWO wo_rel.ANTISYM wo_rel_def) lemma total_VWO: "total VWO" using VWO by (metis wo_rel.TOTAL wo_rel.intro) lemma total_VWOId: "total (VWO-Id)" by (simp add: total_VWO) lemma Linear_order_VWO: "Linear_order VWO" using VWO well_order_on_def by blast lemma wo_rel_VWO: "wo_rel VWO" using VWO wo_rel_def by blast subsubsection \Transitive Closure and VWO\ lemma mem_imp_VWO: "x \ elts y \ (x,y) \ VWO" using VWO by blast lemma less_TC_imp_VWO: "x \ y \ (x,y) \ VWO" unfolding less_TC_def proof (induction y arbitrary: x rule: eps_induct) case (step y' u) then consider "u \ elts y'" | v where "v \ elts y'" "u \ elts (TC v)" by (auto simp: TC [of y']) then show ?case proof cases case 2 then show ?thesis by (meson mem_imp_VWO step.IH transD trans_VWO) qed (use mem_imp_VWO in blast) qed lemma le_TC_imp_VWO: "x \ y \ (x,y) \ VWO" by (metis Diff_iff Linear_order_VWO Linear_order_in_diff_Id UNIV_I VWO le_TC_def less_TC_imp_VWO) lemma le_TC_0_iff [simp]: "x \ 0 \ x = 0" by (simp add: le_TC_def) lemma less_TC_succ: " x \ succ \ \ x \ \ \ x = \" by (metis elts_succ insert_iff le_TC_def less_TC_iff) lemma le_TC_succ: "x \ succ \ \ x \ \ \ x = succ \" by (simp add: le_TC_def less_TC_succ) lemma Transset_TC_eq [simp]: "Transset x \ TC x = x" by (simp add: TC_least arg_subset_TC eq_iff) lemma Ord_TC_less_iff: "\Ord \; Ord \\ \ \ \ \ \ \ < \" by (metis Ord_def Ord_mem_iff_lt Transset_TC_eq less_TC_def) lemma Ord_mem_iff_less_TC: "Ord l \ k \ elts l \ k \ l" by (simp add: Ord_def less_TC_def) lemma le_TC_Ord: "\\ \ \; Ord \\ \ Ord \" by (metis Ord_def Ord_in_Ord Transset_TC_eq le_TC_def less_TC_def) lemma Ord_less_TC_mem: assumes "Ord \" "\ \ \" shows "\ \ elts \" using Ord_def assms less_TC_def by auto lemma VWO_TC_le: "\Ord \; Ord \; (\, \) \ VWO\ \ \ \ \" proof (induct \ arbitrary: \ rule: Ord_induct) case (step \) then show ?case by (metis DiffI IdD Linear_order_VWO Linear_order_in_diff_Id Ord_linear Ord_mem_iff_less_TC VWO iso_tuple_UNIV_I le_TC_def mem_imp_VWO) qed lemma VWO_iff_Ord_le [simp]: "\Ord \; Ord \\ \ (\, \) \ VWO \ \ \ \" by (metis VWO_TC_le Ord_TC_less_iff le_TC_def le_TC_imp_VWO le_less) lemma zero_TC_le [iff]: "0 \ y" using le_TC_def nonzero_less_TC by auto lemma succ_le_TC_iff: "Ord j \ succ i \ j \ i \ j" by (metis Ord_in_Ord Ord_linear Ord_mem_iff_less_TC Ord_succ le_TC_def less_TC_succ less_asym_TC) lemma VWO_0_iff [simp]: "(x,0) \ VWO \ x=0" proof show "x = 0" if "(x, 0) \ VWO" using zero_TC_le [of x] le_TC_imp_VWO that by (metis DiffI Linear_order_VWO Linear_order_in_diff_Id UNIV_I VWO pair_in_Id_conv) qed auto lemma VWO_antisym: assumes "(x,y) \ VWO" "(y,x) \ VWO" shows "x=y" by (metis Diff_iff IdD Linear_order_VWO Linear_order_in_diff_Id UNIV_I VWO assms) subsubsection \Relation VWF\ definition VWF where "VWF \ VWO - Id" lemma wf_VWF [iff]: "wf VWF" by (simp add: VWF_def wf_VWO) lemma trans_VWF [iff]: "trans VWF" by (simp add: VWF_def antisym_VWO trans_VWO trans_diff_Id) lemma asym_VWF [iff]: "asym VWF" by (metis VWF_def asym.intros irrefl_diff_Id wf_VWF wf_not_sym) lemma total_VWF [iff]: "total VWF" using VWF_def total_VWOId by auto lemma total_on_VWF [iff]: "total_on A VWF" by (meson UNIV_I total_VWF total_on_def) lemma VWF_asym: assumes "(x,y) \ VWF" "(y,x) \ VWF" shows False using VWF_def assms wf_VWO wf_not_sym by fastforce lemma VWF_non_refl [iff]: "(x,x) \ VWF" by simp lemma VWF_iff_Ord_less [simp]: "\Ord \; Ord \\ \ (\,\) \ VWF \ \ < \" by (simp add: VWF_def less_V_def) lemma mem_imp_VWF: "x \ elts y \ (x,y) \ VWF" using VWF_def mem_imp_VWO by fastforce subsection\Order types\ definition ordermap :: "'a set \ ('a \ 'a) set \ 'a \ V" where "ordermap A r \ wfrec r (\f x. set (f ` {y \ A. (y,x) \ r}))" definition ordertype :: "'a set \ ('a \ 'a) set \ V" where "ordertype A r \ set (ordermap A r ` A)" lemma ordermap_type: "small A \ ordermap A r \ A \ elts (ordertype A r)" by (simp add: ordertype_def) lemma ordermap_in_ordertype [intro]: "\a \ A; small A\ \ ordermap A r a \ elts (ordertype A r)" by (simp add: ordertype_def) lemma ordermap: "wf r \ ordermap A r a = set (ordermap A r ` {y \ A. (y,a) \ r})" unfolding ordermap_def by (auto simp: wfrec_fixpoint adm_wf_def) lemma wf_Ord_ordermap [iff]: assumes "wf r" "trans r" shows "Ord (ordermap A r x)" using \wf r\ proof (induction x rule: wf_induct_rule) case (less u) have "Transset (set (ordermap A r ` {y \ A. (y, u) \ r}))" proof (clarsimp simp add: Transset_def) show "x \ ordermap A r ` {y \ A. (y, u) \ r}" if "small (ordermap A r ` {y \ A. (y, u) \ r})" and x: "x \ elts (ordermap A r y)" and "y \ A" "(y, u) \ r" for x y proof - have "ordermap A r y = ZFC_in_HOL.set (ordermap A r ` {a \ A. (a, y) \ r})" using ordermap assms(1) by force then have "x \ ordermap A r ` {z \ A. (z, y) \ r}" by (metis (no_types, lifting) elts_of_set empty_iff x) then have "\v. v \ A \ (v, u) \ r \ x = ordermap A r v" using that transD [OF \trans r\] by blast then show ?thesis by blast qed qed moreover have "Ord x" if "x \ elts (set (ordermap A r ` {y \ A. (y, u) \ r}))" for x using that less by (auto simp: split: if_split_asm) ultimately show ?case by (metis (full_types) Ord_def ordermap assms(1)) qed lemma wf_Ord_ordertype: assumes "wf r" "trans r" shows "Ord(ordertype A r)" proof - have "y \ set (ordermap A r ` A)" if "y = ordermap A r x" "x \ A" "small (ordermap A r ` A)" for x y using that by (auto simp: less_eq_V_def ordermap [OF \wf r\, of A x]) moreover have "z \ y" if "y \ ordermap A r ` A" "z \ elts y" for y z by (metis wf_Ord_ordermap OrdmemD assms imageE order.strict_implies_order that) ultimately show ?thesis unfolding ordertype_def Ord_def Transset_def by simp qed lemma Ord_ordertype [simp]: "Ord(ordertype A VWF)" using wf_Ord_ordertype by blast lemma Ord_ordermap [simp]: "Ord (ordermap A VWF x)" by blast lemma ordertype_singleton [simp]: assumes "wf r" shows "ordertype {x} r = 1" proof - have \: "{y. y = x \ (y, x) \ r} = {}" using assms by auto show ?thesis by (auto simp add: ordertype_def assms \ ordermap [where a=x]) qed subsubsection\@{term ordermap} preserves the orderings in both directions\ lemma ordermap_mono: assumes wx: "(w, x) \ r" and "wf r" "w \ A" "small A" shows "ordermap A r w \ elts (ordermap A r x)" proof - have "small {a \ A. (a, x) \ r} \ w \ A \ (w, x) \ r" by (simp add: assms) then show ?thesis using assms ordermap [of r A] by (metis (no_types, lifting) elts_of_set image_eqI mem_Collect_eq replacement) qed lemma converse_ordermap_mono: assumes "ordermap A r y \ elts (ordermap A r x)" "wf r" "total_on A r" "x \ A" "y \ A" "small A" shows "(y, x) \ r" proof (cases "x = y") case True then show ?thesis using assms(1) mem_not_refl by blast next case False then consider "(x,y) \ r" | "(y,x) \ r" using \total_on A r\ assms by (meson UNIV_I total_on_def) then show ?thesis by (meson ordermap_mono assms mem_not_sym) qed lemma converse_ordermap_mono_iff: assumes "wf r" "total_on A r" "x \ A" "y \ A" "small A" shows "ordermap A r y \ elts (ordermap A r x) \ (y, x) \ r" by (metis assms converse_ordermap_mono ordermap_mono) lemma ordermap_surj: "elts (ordertype A r) \ ordermap A r ` A" unfolding ordertype_def by simp lemma ordermap_bij: assumes "wf r" "total_on A r" "small A" shows "bij_betw (ordermap A r) A (elts (ordertype A r))" unfolding bij_betw_def proof (intro conjI) show "inj_on (ordermap A r) A" unfolding inj_on_def by (metis assms mem_not_refl ordermap_mono total_on_def) show "ordermap A r ` A = elts (ordertype A r)" by (metis ordertype_def \small A\ elts_of_set replacement) qed lemma ordermap_eq_iff [simp]: "\x \ A; y \ A; wf r; total_on A r; small A\ \ ordermap A r x = ordermap A r y \ x = y" by (metis bij_betw_iff_bijections ordermap_bij) lemma inv_into_ordermap: "\ \ elts (ordertype A r) \ inv_into A (ordermap A r) \ \ A" by (meson in_mono inv_into_into ordermap_surj) lemma ordertype_nat_imp_finite: assumes "ordertype A r = ord_of_nat m" "small A" "wf r" "total_on A r" shows "finite A" proof - have "A \ elts m" using eqpoll_def assms ordermap_bij by fastforce then show ?thesis using eqpoll_finite_iff finite_Ord_omega by blast qed lemma wf_ordertype_eqpoll: assumes "wf r" "total_on A r" "small A" shows "elts (ordertype A r) \ A" using assms eqpoll_def eqpoll_sym ordermap_bij by blast lemma ordertype_eqpoll: assumes "small A" shows "elts (ordertype A VWF) \ A" using assms wf_ordertype_eqpoll total_VWF wf_VWF by (simp add: wf_ordertype_eqpoll total_on_def) subsection \More advanced @{term ordertype} and @{term ordermap} results\ lemma ordermap_VWF_0 [simp]: "ordermap A VWF 0 = 0" by (simp add: ordermap wf_VWO VWF_def) lemma ordertype_empty [simp]: "ordertype {} r = 0" by (simp add: ordertype_def) lemma ordertype_eq_0_iff [simp]: "\small X; wf r\ \ ordertype X r = 0 \ X = {}" by (metis ordertype_def elts_of_set replacement image_is_empty zero_V_def) lemma ordermap_mono_less: assumes "(w, x) \ r" and "wf r" "trans r" and "w \ A" "x \ A" and "small A" shows "ordermap A r w < ordermap A r x" by (simp add: OrdmemD assms ordermap_mono) lemma ordermap_mono_le: assumes "(w, x) \ r \ w=x" and "wf r" "trans r" and "w \ A" "x \ A" and "small A" shows "ordermap A r w \ ordermap A r x" by (metis assms dual_order.strict_implies_order eq_refl ordermap_mono_less) lemma converse_ordermap_le_mono: assumes "ordermap A r y \ ordermap A r x" "wf r" "total r" "x \ A" "small A" shows "(y, x) \ r \ y=x" by (meson UNIV_I assms mem_not_refl ordermap_mono total_on_def vsubsetD) lemma ordertype_mono: assumes "X \ Y" and r: "wf r" "trans r" and "small Y" shows "ordertype X r \ ordertype Y r" proof - have "small X" using assms smaller_than_small by fastforce have *: "ordermap X r x \ ordermap Y r x" for x using \wf r\ proof (induction x rule: wf_induct_rule) case (less x) have "ordermap X r z < ordermap Y r x" if "z \ X" and zx: "(z,x) \ r" for z using less [OF zx] assms by (meson Ord_linear2 OrdmemD wf_Ord_ordermap ordermap_mono in_mono leD that(1) vsubsetD zx) then show ?case by (auto simp add: ordermap [of _ X x] \small X\ Ord_mem_iff_lt set_image_le_iff less_eq_V_def r) qed show ?thesis proof - have "ordermap Y r ` Y = elts (ordertype Y r)" by (metis ordertype_def \small Y\ elts_of_set replacement) then have "ordertype Y r \ ordermap X r ` X" using "*" \X \ Y\ by fastforce then show ?thesis by (metis Ord_linear2 Ord_mem_iff_lt ordertype_def wf_Ord_ordertype \small X\ elts_of_set replacement r) qed qed corollary ordertype_VWF_mono: assumes "X \ Y" "small Y" shows "ordertype X VWF \ ordertype Y VWF" using assms by (simp add: ordertype_mono) lemma ordertype_UNION_ge: assumes "A \ \" "wf r" "trans r" "\ \ Collect small" "small \" shows "ordertype A r \ ordertype (\\) r" by (rule ordertype_mono) (use assms in auto) lemma inv_ordermap_mono_less: assumes "(inv_into M (ordermap M r) \, inv_into M (ordermap M r) \) \ r" and "small M" and \: "\ \ elts (ordertype M r)" and \: "\ \ elts (ordertype M r)" and "wf r" "trans r" shows "\ < \" proof - have "\ = ordermap M r (inv_into M (ordermap M r) \)" by (metis \ f_inv_into_f ordermap_surj subset_eq) also have "\ < ordermap M r (inv_into M (ordermap M r) \)" by (meson \ \ assms in_mono inv_into_into ordermap_mono_less ordermap_surj) also have "\ = \" by (meson \ f_inv_into_f in_mono ordermap_surj) finally show ?thesis . qed lemma inv_ordermap_mono_eq: assumes "inv_into M (ordermap M r) \ = inv_into M (ordermap M r) \" and "\ \ elts (ordertype M r)" "\ \ elts (ordertype M r)" shows "\ = \" by (metis assms f_inv_into_f ordermap_surj subsetD) lemma inv_ordermap_VWF_mono_le: assumes "inv_into M (ordermap M VWF) \ \ inv_into M (ordermap M VWF) \" and "M \ ON" "small M" and \: "\ \ elts (ordertype M VWF)" and \: "\ \ elts (ordertype M VWF)" shows "\ \ \" proof - have "\ = ordermap M VWF (inv_into M (ordermap M VWF) \)" by (metis \ f_inv_into_f ordermap_surj subset_eq) also have "\ \ ordermap M VWF (inv_into M (ordermap M VWF) \)" by (metis ON_imp_Ord VWF_iff_Ord_less assms dual_order.strict_implies_order elts_of_set eq_refl inv_into_into order.not_eq_order_implies_strict ordermap_mono_less ordertype_def replacement trans_VWF wf_VWF) also have "\ = \" by (meson \ f_inv_into_f in_mono ordermap_surj) finally show ?thesis . qed lemma inv_ordermap_VWF_mono_iff: assumes "M \ ON" "small M" and "\ \ elts (ordertype M VWF)" and "\ \ elts (ordertype M VWF)" shows "inv_into M (ordermap M VWF) \ \ inv_into M (ordermap M VWF) \ \ \ \ \" by (metis ON_imp_Ord Ord_linear_le assms dual_order.eq_iff inv_into_ordermap inv_ordermap_VWF_mono_le) lemma inv_ordermap_VWF_strict_mono_iff: assumes "M \ ON" "small M" and "\ \ elts (ordertype M VWF)" and "\ \ elts (ordertype M VWF)" shows "inv_into M (ordermap M VWF) \ < inv_into M (ordermap M VWF) \ \ \ < \" by (simp add: assms inv_ordermap_VWF_mono_iff less_le_not_le) lemma strict_mono_on_ordertype: assumes "M \ ON" "small M" - obtains f where "f \ elts (ordertype M VWF) \ M" "strict_mono_on f (elts (ordertype M VWF))" + obtains f where "f \ elts (ordertype M VWF) \ M" "strict_mono_on (elts (ordertype M VWF)) f" proof show "inv_into M (ordermap M VWF) \ elts (ordertype M VWF) \ M" by (meson Pi_I' in_mono inv_into_into ordermap_surj) - show "strict_mono_on (inv_into M (ordermap M VWF)) (elts (ordertype M VWF))" + show "strict_mono_on (elts (ordertype M VWF)) (inv_into M (ordermap M VWF))" proof (clarsimp simp: strict_mono_on_def) fix x y assume "x \ elts (ordertype M VWF)" "y \ elts (ordertype M VWF)" "x < y" then show "inv_into M (ordermap M VWF) x < inv_into M (ordermap M VWF) y" using assms by (meson ON_imp_Ord Ord_linear2 inv_into_into inv_ordermap_VWF_mono_le leD ordermap_surj subsetD) qed qed lemma ordermap_inc_eq: assumes "x \ A" "small A" and \: "\x y. \x\A; y\A; (x,y) \ r\ \ (\ x, \ y) \ s" and r: "wf r" "total_on A r" and "wf s" shows "ordermap (\ ` A) s (\ x) = ordermap A r x" using \wf r\ \x \ A\ proof (induction x rule: wf_induct_rule) case (less x) then have 1: "{y \ A. (y, x) \ r} = A \ {y. (y, x) \ r}" using r by auto have 2: "{y \ \ ` A. (y, \ x) \ s} = \ ` A \ {y. (y, \ x) \ s}" by auto have inv\: "\x y. \x\A; y\A; (\ x, \ y) \ s\ \ (x, y) \ r" by (metis \ \wf s\ \total_on A r\ total_on_def wf_not_sym) have eq: "f ` (\ ` A \ {y. (y, \ x) \ s}) = (f \ \) ` (A \ {y. (y, x) \ r})" for f :: "'b \ V" using less by (auto simp: image_subset_iff inv\ \) show ?case using less by (simp add: ordermap [OF \wf r\, of _ x] ordermap [OF \wf s\, of _ "\ x"] 1 2 eq) qed lemma ordertype_inc_eq: assumes "small A" and \: "\x y. \x\A; y\A; (x,y) \ r\ \ (\ x, \ y) \ s" and r: "wf r" "total_on A r" and "wf s" shows "ordertype (\ ` A) s = ordertype A r" proof - have "ordermap (\ ` A) s (\ x) = ordermap A r x" if "x \ A" for x using assms that by (auto simp: ordermap_inc_eq) then show ?thesis unfolding ordertype_def by (metis (no_types, lifting) image_cong image_image) qed lemma ordertype_inc_le: assumes "small A" "small B" and \: "\x y. \x\A; y\A; (x,y) \ r\ \ (\ x, \ y) \ s" and r: "wf r" "total_on A r" and "wf s" "trans s" and "\ ` A \ B" shows "ordertype A r \ ordertype B s" by (metis assms ordertype_inc_eq ordertype_mono) corollary ordertype_VWF_inc_eq: assumes "A \ ON" "\ ` A \ ON" "small A" and "\x y. \x\A; y\A; x \ \ x < \ y" shows "ordertype (\ ` A) VWF = ordertype A VWF" proof (rule ordertype_inc_eq) show "(\ x, \ y) \ VWF" if "x \ A" "y \ A" "(x, y) \ VWF" for x y using that ON_imp_Ord assms by auto show "total_on A VWF" by (meson UNIV_I total_VWF total_on_def) qed (use assms in auto) lemma ordertype_image_ordermap: assumes "small A" "X \ A" "wf r" "trans r" "total_on X r" shows "ordertype (ordermap A r ` X) VWF = ordertype X r" proof (rule ordertype_inc_eq) show "small X" by (meson assms smaller_than_small) show "(ordermap A r x, ordermap A r y) \ VWF" if "x \ X" "y \ X" "(x, y) \ r" for x y by (meson that wf_Ord_ordermap VWF_iff_Ord_less assms ordermap_mono_less subsetD) qed (use assms in auto) lemma ordertype_map_image: assumes "B \ A" "small A" shows "ordertype (ordermap A VWF ` A - ordermap A VWF ` B) VWF = ordertype (A - B) VWF" proof - have "ordermap A VWF ` A - ordermap A VWF ` B = ordermap A VWF ` (A - B)" using assms by auto then have "ordertype (ordermap A VWF ` A - ordermap A VWF ` B) VWF = ordertype (ordermap A VWF ` (A - B)) VWF" by simp also have "\ = ordertype (A - B) VWF" using \small A\ ordertype_image_ordermap by fastforce finally show ?thesis . qed proposition ordertype_le_ordertype: assumes r: "wf r" "total_on A r" and "small A" assumes s: "wf s" "total_on B s" "trans s" and "small B" shows "ordertype A r \ ordertype B s \ (\f \ A \ B. inj_on f A \ (\x \ A. \y \ A. ((x,y) \ r \ (f x, f y) \ s)))" (is "?lhs = ?rhs") proof assume L: ?lhs define f where "f \ inv_into B (ordermap B s) \ ordermap A r" show ?rhs proof (intro bexI conjI ballI impI) have AB: "elts (ordertype A r) \ ordermap B s ` B" by (metis L assms(7) ordertype_def replacement set_of_elts small_elts subset_iff_less_eq_V) have bijA: "bij_betw (ordermap A r) A (elts (ordertype A r))" using ordermap_bij \small A\ r by blast have "inv_into B (ordermap B s) (ordermap A r i) \ B" if "i \ A" for i by (meson L \small A\ inv_into_into ordermap_in_ordertype ordermap_surj subsetD that vsubsetD) then show "f \ A \ B" by (auto simp: Pi_iff f_def) show "inj_on f A" proof (clarsimp simp add: f_def inj_on_def) fix x y assume "x \ A" "y \ A" and "inv_into B (ordermap B s) (ordermap A r x) = inv_into B (ordermap B s) (ordermap A r y)" then have "ordermap A r x = ordermap A r y" by (meson AB \small A\ inv_into_injective ordermap_in_ordertype subsetD) then show "x = y" by (metis \x \ A\ \y \ A\ bijA bij_betw_inv_into_left) qed next fix x y assume "x \ A" "y \ A" and "(x, y) \ r" have \: "ordermap A r y \ ordermap B s ` B" by (meson L \y \ A\ \small A\ in_mono ordermap_in_ordertype ordermap_surj vsubsetD) moreover have \: "\x. inv_into B (ordermap B s) (ordermap A r x) = f x" by (simp add: f_def) then have *: "ordermap B s (f y) = ordermap A r y" using \ by (metis f_inv_into_f) moreover have "ordermap A r x \ ordermap B s ` B" by (meson L \x \ A\ \small A\ in_mono ordermap_in_ordertype ordermap_surj vsubsetD) moreover have "ordermap A r x < ordermap A r y" using * r s by (metis (no_types) wf_Ord_ordermap OrdmemD \(x, y) \ r\ \x \ A\ \small A\ ordermap_mono) ultimately show "(f x, f y) \ s" using \ s by (metis assms(7) f_inv_into_f inv_into_into less_asym ordermap_mono_less total_on_def) qed next assume R: ?rhs then obtain f where f: "f\A \ B" "inj_on f A" "\x\A. \y\A. (x, y) \ r \ (f x, f y) \ s" by blast show ?lhs by (rule ordertype_inc_le [where \=f]) (use f assms in auto) qed lemma iso_imp_ordertype_eq_ordertype: assumes iso: "iso r r' f" and "wf r" and "Total r" and sm: "small (Field r)" shows "ordertype (Field r) r = ordertype (Field r') r'" by (metis (no_types, lifting) iso_forward iso_wf assms iso_Field ordertype_inc_eq sm) lemma ordertype_infinite_ge_\: assumes "infinite A" "small A" shows "ordertype A VWF \ \" proof - have "inj_on (ordermap A VWF) A" by (meson ordermap_bij \small A\ bij_betw_def total_on_VWF wf_VWF) then have "infinite (ordermap A VWF ` A)" using \infinite A\ finite_image_iff by blast then show ?thesis using Ord_ordertype \small A\ infinite_Ord_omega by (auto simp: ordertype_def) qed lemma ordertype_eqI: assumes "wf r" "total_on A r" "small A" "wf s" "bij_betw f A B" "(\x \ A. \y \ A. (f x, f y) \ s \ (x,y) \ r)" shows "ordertype A r = ordertype B s" by (metis assms bij_betw_imp_surj_on ordertype_inc_eq) lemma ordermap_eq_self: assumes "Ord \" and x: "x \ elts \" shows "ordermap (elts \) VWF x = x" using Ord_in_Ord [OF assms] x proof (induction x rule: Ord_induct) case (step x) have 1: "{y \ elts \. (y, x) \ VWF} = elts x" (is "?A = _") proof show "?A \ elts x" using \Ord \\ by clarify (meson Ord_in_Ord Ord_mem_iff_lt VWF_iff_Ord_less step.hyps) show "elts x \ ?A" using \Ord \\ by clarify (meson Ord_in_Ord Ord_trans OrdmemD VWF_iff_Ord_less step.prems) qed show ?case using step by (simp add: ordermap [OF wf_VWF, of _ x] 1 Ord_trans [of _ _ \] step.prems \Ord \\ cong: image_cong) qed lemma ordertype_eq_Ord [simp]: assumes "Ord \" shows "ordertype (elts \) VWF = \" using assms ordermap_eq_self [OF assms] by (simp add: ordertype_def) proposition ordertype_eq_iff: assumes \: "Ord \" and r: "wf r" and "small A" "total_on A r" "trans r" shows "ordertype A r = \ \ (\f. bij_betw f A (elts \) \ (\x \ A. \y \ A. f x < f y \ (x,y) \ r))" (is "?lhs = ?rhs") proof safe assume eq: "\ = ordertype A r" show "\f. bij_betw f A (elts (ordertype A r)) \ (\x\A. \y\A. f x < f y \ ((x, y) \ r))" proof (intro exI conjI ballI) show "bij_betw (ordermap A r) A (elts (ordertype A r))" by (simp add: assms ordermap_bij) then show "ordermap A r x < ordermap A r y \ (x, y) \ r" if "x \ A" "y \ A" for x y using that assms by (metis order.asym ordermap_mono_less total_on_def) qed next fix f assume f: "bij_betw f A (elts \)" "\x\A. \y\A. f x < f y \ (x, y) \ r" have "ordertype A r = ordertype (elts \) VWF" proof (rule ordertype_eqI) show "\x\A. \y\A. ((f x, f y) \ VWF) = ((x, y) \ r)" by (meson Ord_in_Ord VWF_iff_Ord_less \ bij_betwE f) qed (use assms f in auto) then show ?lhs by (simp add: \) qed corollary ordertype_VWF_eq_iff: assumes "Ord \" "small A" shows "ordertype A VWF = \ \ (\f. bij_betw f A (elts \) \ (\x \ A. \y \ A. f x < f y \ (x,y) \ VWF))" by (metis UNIV_I assms ordertype_eq_iff total_VWF total_on_def trans_VWF wf_VWF) lemma ordertype_le_Ord: assumes "Ord \" "X \ elts \" shows "ordertype X VWF \ \" by (metis assms ordertype_VWF_mono ordertype_eq_Ord small_elts) lemma ordertype_inc_le_Ord: assumes "small A" "Ord \" and \: "\x y. \x\A; y\A; (x,y) \ r\ \ \ x < \ y" and "wf r" "total_on A r" and sub: "\ ` A \ elts \" shows "ordertype A r \ \" proof - have "\x y. \x\A; y\A; (x,y) \ r\ \ (\ x, \ y) \ VWF" by (meson Ord_in_Ord VWF_iff_Ord_less \ \Ord \\ sub image_subset_iff) with assms show ?thesis by (metis ordertype_inc_eq ordertype_le_Ord wf_VWF) qed lemma le_ordertype_obtains_subset: assumes \: "\ \ \" "ordertype H VWF = \" and "small H" "Ord \" obtains G where "G \ H" "ordertype G VWF = \" proof (intro exI conjI that) let ?f = "ordermap H VWF" show \: "inv_into H ?f ` elts \ \ H" unfolding image_subset_iff by (metis \ inv_into_into ordermap_surj subsetD vsubsetD) have "\f. bij_betw f (inv_into H ?f ` elts \) (elts \) \ (\x\inv_into H ?f ` elts \. \y\inv_into H ?f ` elts \. (f x < f y) = ((x, y) \ VWF))" proof (intro exI conjI ballI iffI) show "bij_betw ?f (inv_into H ?f ` elts \) (elts \)" using ordermap_bij [OF wf_VWF total_on_VWF \small H\] \ by (metis bij_betw_inv_into_RIGHT bij_betw_subset less_eq_V_def \) next fix x y assume x: "x \ inv_into H ?f ` elts \" and y: "y \ inv_into H ?f ` elts \" show "?f x < ?f y" if "(x,y) \ VWF" using that \ \small H\ in_mono ordermap_mono_less x y by fastforce show "(x,y) \ VWF" if "?f x < ?f y" using that \ \small H\ in_mono ordermap_mono_less [OF _ wf_VWF trans_VWF] x y by (metis UNIV_I less_imp_not_less total_VWF total_on_def) qed then show "ordertype (inv_into H ?f ` elts \) VWF = \" by (subst ordertype_eq_iff) (use assms in auto) qed lemma ordertype_infinite_\: assumes "A \ elts \" "infinite A" shows "ordertype A VWF = \" proof (rule antisym) show "ordertype A VWF \ \" by (simp add: assms ordertype_le_Ord) show "\ \ ordertype A VWF" using assms down ordertype_infinite_ge_\ by auto qed text \For infinite sets of natural numbers\ lemma ordertype_nat_\: assumes "infinite N" shows "ordertype N less_than = \" proof - have "small N" by (meson inj_on_def ord_of_nat_inject small_def small_iff_range small_image_nat_V) have "ordertype (ord_of_nat ` N) VWF = \" by (force simp: assms finite_image_iff inj_on_def intro: ordertype_infinite_\) moreover have "ordertype (ord_of_nat ` N) VWF = ordertype N less_than" by (auto intro: ordertype_inc_eq \small N\) ultimately show ?thesis by simp qed proposition ordertype_eq_ordertype: assumes r: "wf r" "total_on A r" "trans r" and "small A" assumes s: "wf s" "total_on B s" "trans s" and "small B" shows "ordertype A r = ordertype B s \ (\f. bij_betw f A B \ (\x \ A. \y \ A. (f x, f y) \ s \ (x,y) \ r))" (is "?lhs = ?rhs") proof assume L: ?lhs define \ where "\ = ordertype A r" have A: "bij_betw (ordermap A r) A (ordermap A r ` A)" by (meson ordermap_bij assms(4) bij_betw_def r) have B: "bij_betw (ordermap B s) B (ordermap B s ` B)" by (meson ordermap_bij assms(8) bij_betw_def s) define f where "f \ inv_into B (ordermap B s) \ ordermap A r" show ?rhs proof (intro exI conjI) have bijA: "bij_betw (ordermap A r) A (elts \)" unfolding \_def using ordermap_bij \small A\ r by blast moreover have bijB: "bij_betw (ordermap B s) B (elts \)" by (simp add: L \_def ordermap_bij \small B\ s) ultimately show bij: "bij_betw f A B" unfolding f_def using bij_betw_comp_iff bij_betw_inv_into by blast have invB: "\\. \ \ elts \ \ ordermap B s (inv_into B (ordermap B s) \) = \" by (meson bijB bij_betw_inv_into_right) have ordermap_A_\: "\a. a \ A \ ordermap A r a \ elts \" using bijA bij_betwE by auto have f_in_B: "\a. a \ A \ f a \ B" using bij bij_betwE by fastforce show "\x\A. \y\A. (f x, f y) \ s \ (x, y) \ r" proof (intro iffI ballI) fix x y assume "x \ A" "y \ A" and ins: "(f x, f y) \ s" then have "ordermap A r x < ordermap A r y" unfolding o_def by (metis (mono_tags, lifting) f_def \small B\ comp_apply f_in_B invB ordermap_A_\ ordermap_mono_less s(1) s(3)) then show "(x, y) \ r" by (metis \x \ A\ \y \ A\ \small A\ order.asym ordermap_mono_less r total_on_def) next fix x y assume "x \ A" "y \ A" and "(x, y) \ r" then have "ordermap A r x < ordermap A r y" by (simp add: \small A\ ordermap_mono_less r) then have "(f y, f x) \ s" by (metis (mono_tags, lifting) \x \ A\ \y \ A\ \small B\ comp_apply f_def f_in_B invB order.asym ordermap_A_\ ordermap_mono_less s(1) s(3)) moreover have "f y \ f x" by (metis \(x, y) \ r\ \x \ A\ \y \ A\ bij bij_betw_inv_into_left r(1) wf_not_sym) ultimately show "(f x, f y) \ s" by (meson \x \ A\ \y \ A\ f_in_B s(2) total_on_def) qed qed next assume ?rhs then show ?lhs using assms ordertype_eqI by blast qed corollary ordertype_eq_ordertype_iso: assumes r: "wf r" "total_on A r" "trans r" and "small A" and FA: "Field r = A" assumes s: "wf s" "total_on B s" "trans s" and "small B" and FB: "Field s = B" shows "ordertype A r = ordertype B s \ (\f. iso r s f)" (is "?lhs = ?rhs") proof assume L: ?lhs then obtain f where "bij_betw f A B" "\x \ A. \y \ A. (f x, f y) \ s \ (x,y) \ r" using assms ordertype_eq_ordertype by blast then show ?rhs using FA FB iso_iff2 by blast next assume ?rhs then show ?lhs using FA FB \small A\ iso_imp_ordertype_eq_ordertype r by blast qed lemma Limit_ordertype_imp_Field_Restr: assumes Lim: "Limit (ordertype A r)" and r: "wf r" "total_on A r" and "small A" shows "Field (Restr r A) = A" proof - have "\y\A. (x,y) \ r" if "x \ A" for x proof - let ?oy = "succ (ordermap A r x)" have \
: "?oy \ elts (ordertype A r)" by (simp add: Lim \small A\ ordermap_in_ordertype succ_in_Limit_iff that) then have A: "inv_into A (ordermap A r) ?oy \ A" by (simp add: inv_into_ordermap) moreover have "(x, inv_into A (ordermap A r) ?oy) \ r" proof - have "ordermap A r x \ elts (ordermap A r (inv_into A (ordermap A r) ?oy))" by (metis "\
" elts_succ f_inv_into_f insert_iff ordermap_surj subsetD) then show ?thesis by (metis \small A\ A converse_ordermap_mono r that) qed ultimately show ?thesis .. qed then have "A \ Field (Restr r A)" by (auto simp: Field_def) then show ?thesis by (simp add: Field_Restr_subset subset_antisym) qed lemma ordertype_Field_Restr: assumes "wf r" "total_on A r" "trans r" "small A" "Field (Restr r A) = A" shows "ordertype (Field (Restr r A)) (Restr r A) = ordertype A r" using assms by (force simp: ordertype_eq_ordertype wf_Restr total_on_def trans_Restr) proposition ordertype_eq_ordertype_iso_Restr: assumes r: "wf r" "total_on A r" "trans r" and "small A" and FA: "Field (Restr r A) = A" assumes s: "wf s" "total_on B s" "trans s" and "small B" and FB: "Field (Restr s B) = B" shows "ordertype A r = ordertype B s \ (\f. iso (Restr r A) (Restr s B) f)" (is "?lhs = ?rhs") proof assume L: ?lhs then obtain f where "bij_betw f A B" "\x \ A. \y \ A. (f x, f y) \ s \ (x,y) \ r" using assms ordertype_eq_ordertype by blast then show ?rhs using FA FB bij_betwE unfolding iso_iff2 by fastforce next assume ?rhs moreover have "ordertype (Field (Restr r A)) (Restr r A) = ordertype A r" using FA \small A\ ordertype_Field_Restr r by blast moreover have "ordertype (Field (Restr s B)) (Restr s B) = ordertype B s" using FB \small B\ ordertype_Field_Restr s by blast ultimately show ?lhs using iso_imp_ordertype_eq_ordertype FA FB \small A\ r by (fastforce intro: total_on_imp_Total_Restr trans_Restr wf_Int1) qed lemma ordermap_insert: assumes "Ord \" and y: "Ord y" "y \ \" and U: "U \ elts \" shows "ordermap (insert \ U) VWF y = ordermap U VWF y" using y proof (induction rule: Ord_induct) case (step y) then have 1: "{u \ U. (u, y) \ VWF} = elts y \ U" apply (simp add: set_eq_iff) by (meson Ord_in_Ord Ord_mem_iff_lt VWF_iff_Ord_less assms subsetD) have 2: "{u \ insert \ U. (u, y) \ VWF} = elts y \ U" apply (simp add: set_eq_iff) by (meson Ord_in_Ord Ord_mem_iff_lt VWF_iff_Ord_less assms leD step.hyps step.prems subsetD) show ?case using step apply (simp only: ordermap [OF wf_VWF, of _ y] 1 2) by (meson Int_lower1 Ord_is_Transset Sup.SUP_cong Transset_def assms(1) in_mono vsubsetD) qed lemma ordertype_insert: assumes "Ord \" and U: "U \ elts \" shows "ordertype (insert \ U) VWF = succ (ordertype U VWF)" proof - have \: "{y \ insert \ U. (y, \) \ VWF} = U" "{y \ U. (y, \) \ VWF} = U" using Ord_in_Ord OrdmemD assms by auto have eq: "\x. x \ U \ ordermap (insert \ U) VWF x = ordermap U VWF x" by (meson Ord_in_Ord Ord_is_Transset Transset_def U assms(1) in_mono ordermap_insert) have "ordertype (insert \ U) VWF = ZFC_in_HOL.set (insert (ordermap U VWF \) (ordermap U VWF ` U))" by (simp add: ordertype_def ordermap_insert assms eq) also have "\ = succ (ZFC_in_HOL.set (ordermap U VWF ` U))" using "\" U by (simp add: ordermap [OF wf_VWF, of _ \] down succ_def vinsert_def) also have "\ = succ (ordertype U VWF)" by (simp add: ordertype_def) finally show ?thesis . qed lemma finite_ordertype_le_card: assumes "finite A" "wf r" "trans r" shows "ordertype A r \ ord_of_nat (card A)" proof - have "Ord (ordertype A r)" by (simp add: wf_Ord_ordertype assms) moreover have "ordermap A r ` A = elts (ordertype A r)" by (simp add: ordertype_def finite_imp_small \finite A\) moreover have "card (ordermap A r ` A) \ card A" using \finite A\ card_image_le by blast ultimately show ?thesis by (metis Ord_linear_le Ord_ord_of_nat \finite A\ card_ord_of_nat card_seteq finite_imageI less_eq_V_def) qed lemma ordertype_VWF_\: assumes "finite A" shows "ordertype A VWF \ elts \" proof - have "finite (ordermap A VWF ` A)" using assms by blast then have "ordertype A VWF < \" by (meson Ord_\ OrdmemD trans_VWF wf_VWF assms finite_ordertype_le_card le_less_trans ord_of_nat_\) then show ?thesis by (simp add: Ord_mem_iff_lt) qed lemma ordertype_VWF_finite_nat: assumes "finite A" shows "ordertype A VWF = ord_of_nat (card A)" by (metis finite_imp_small ordermap_bij total_on_VWF wf_VWF \_def assms bij_betw_same_card card_ord_of_nat elts_of_set f_inv_into_f inf ordertype_VWF_\) lemma finite_ordertype_eq_card: assumes "small A" "wf r" "trans r" "total_on A r" shows "ordertype A r = ord_of_nat m \ finite A \ card A = m" using ordermap_bij [OF \wf r\] proof - have *: "bij_betw (ordermap A r) A (elts (ordertype A r))" by (simp add: assms ordermap_bij) moreover have "card (ordermap A r ` A) = card A" by (meson bij_betw_def * card_image) ultimately show ?thesis using assms bij_betw_finite bij_betw_imp_surj_on finite_Ord_omega ordertype_VWF_finite_nat wf_Ord_ordertype by fastforce qed lemma ex_bij_betw_strict_mono_card: assumes "finite M" "M \ ON" - obtains h where "bij_betw h {..finite M\ ordermap_bij ordertype_VWF_finite_nat by fastforce let ?h = "(inv_into M (ordermap M VWF)) \ ord_of_nat" show thesis proof show bijh: "bij_betw ?h {.. elts (ordertype M VWF)" "n \ elts (ordertype M VWF)" using \m < n\ \n < card M\ \finite M\ ordertype_VWF_finite_nat by auto have ord: "Ord (?h m)" "Ord (?h n)" using bijh assms(2) bij_betwE that by fastforce+ moreover assume "\ ?h m < ?h n" ultimately consider "?h m = ?h n" | "?h m > ?h n" using Ord_linear_lt by blast then show False proof cases case 1 then have "m = n" by (metis inv_ordermap_mono_eq mn comp_apply ord_of_nat_inject) with \m < n\ show False by blast next case 2 then have "ord_of_nat n \ ord_of_nat m" by (metis Finite_V mn assms comp_def inv_ordermap_VWF_mono_le less_imp_le) then show ?thesis using leD \m < n\ by blast qed qed with assms show ?thesis by (auto simp: strict_mono_on_def) qed qed qed lemma ordertype_finite_less_than [simp]: assumes "finite A" shows "ordertype A less_than = card A" proof - let ?M = "ord_of_nat ` A" obtain M: "finite ?M" "?M \ ON" using Ord_ord_of_nat assms by blast have "ordertype A less_than = ordertype ?M VWF" by (rule ordertype_inc_eq [symmetric]) (use assms finite_imp_small total_on_def in \force+\) also have "\ = card A" proof (subst ordertype_eq_iff) let ?M = "ord_of_nat ` A" - obtain h where bijh: "bij_betw h {.. ord_of_nat \ inv_into {..f. bij_betw f ?M (elts (card A)) \ (\x\?M. \y\?M. f x < f y \ ((x, y) \ VWF))" proof (intro exI conjI ballI) have "bij_betw (ord_of_nat \ inv_into {.. ?M" "y \ ?M" then obtain m n where "x = ord_of_nat m" "y = ord_of_nat n" by auto have "(f x < f y) \ ((h \ inv_into {.. inv_into {.. = (x < y)" using bijh by (simp add: bij_betw_inv_into_right xy) also have "\ \ ((x, y) \ VWF)" using M(2) ON_imp_Ord xy by auto finally show "(f x < f y) \ ((x, y) \ VWF)" . qed qed auto finally show ?thesis . qed subsection \Cardinality of an arbitrary HOL set\ definition gcard :: "'a set \ V" where "gcard X \ if small X then (LEAST i. Ord i \ elts i \ X) else 0" subsection\Cardinality of a set\ definition vcard :: "V\V" where "vcard a \ (LEAST i. Ord i \ elts i \ elts a)" lemma gcard_eq_vcard [simp]: "gcard (elts x) = vcard x" by (simp add: vcard_def gcard_def) definition Card:: "V\bool" where "Card i \ i = vcard i" abbreviation CARD where "CARD \ Collect Card" lemma cardinal_cong: "elts x \ elts y \ vcard x = vcard y" unfolding vcard_def by (meson eqpoll_sym eqpoll_trans) lemma gcardinal_cong: assumes "X \ Y" shows "gcard X = gcard Y" proof - have "(LEAST i. Ord i \ elts i \ X) = (LEAST i. Ord i \ elts i \ Y)" by (meson eqpoll_sym eqpoll_trans assms) then show ?thesis unfolding gcard_def by (meson eqpoll_sym small_eqcong assms) qed lemma vcard_set_image: "inj_on f (elts x) \ vcard (set (f ` elts x)) = vcard x" by (simp add: cardinal_cong) lemma gcard_image: "inj_on f X \ gcard (f ` X) = gcard X" by (simp add: gcardinal_cong) lemma Card_cardinal_eq: "Card \ \ vcard \ = \" by (simp add: Card_def) lemma Card_is_Ord: assumes "Card \" shows "Ord \" proof - obtain \ where "Ord \" "elts \ \ elts \" using Ord_ordertype ordertype_eqpoll by blast then have "Ord (LEAST i. Ord i \ elts i \ elts \)" by (metis Ord_Least) then show ?thesis using Card_def vcard_def assms by auto qed lemma cardinal_eqpoll: "elts (vcard a) \ elts a" unfolding vcard_def using ordertype_eqpoll [of "elts a"] Ord_LeastI by (meson Ord_ordertype small_elts) lemma inj_into_vcard: obtains f where "f \ elts A \ elts (vcard A)" "inj_on f (elts A)" using cardinal_eqpoll [of A] inj_on_the_inv_into the_inv_into_onto by (fastforce simp: Pi_iff bij_betw_def eqpoll_def) lemma cardinal_idem [simp]: "vcard (vcard a) = vcard a" using cardinal_cong cardinal_eqpoll by blast lemma subset_smaller_vcard: assumes "\ \ vcard x" "Card \" obtains y where "y \ x" "vcard y = \" proof - obtain \ where \: "bij_betw \ (elts (vcard x)) (elts x)" using cardinal_eqpoll eqpoll_def by blast show thesis proof let ?y = "ZFC_in_HOL.set (\ ` elts \)" show "?y \ x" by (meson \ assms bij_betwE set_image_le_iff small_elts vsubsetD) show "vcard ?y = \" by (metis vcard_set_image Card_def assms bij_betw_def bij_betw_subset \ less_eq_V_def) qed qed text\every natural number is a (finite) cardinal\ lemma nat_into_Card: assumes "\ \ elts \" shows "Card(\)" proof (unfold Card_def vcard_def, rule sym) obtain n where n: "\ = ord_of_nat n" by (metis \_def assms elts_of_set imageE inf) have "Ord(\)" using assms by auto moreover { fix \ assume "\ < \" "Ord \" "elts \ \ elts \" with n have "elts \ \ {..Ord \\ \\ < \\ \Ord(\)\ by (metis \elts \ \ elts \\ card_seteq eqpoll_finite_iff eqpoll_iff_card finite_lessThan less_eq_V_def less_le_not_le order_refl) } ultimately show "(LEAST i. Ord i \ elts i \ elts \) = \" by (metis (no_types, lifting) Least_equality Ord_linear_le eqpoll_refl less_le_not_le) qed lemma Card_ord_of_nat [simp]: "Card (ord_of_nat n)" by (simp add: \_def nat_into_Card) lemma Card_0 [iff]: "Card 0" by (simp add: nat_into_Card) lemma CardI: "\Ord i; \j. \j < i; Ord j\ \ \ elts j \ elts i\ \ Card i" unfolding Card_def vcard_def by (metis Ord_Least Ord_linear_lt cardinal_eqpoll eqpoll_refl not_less_Ord_Least vcard_def) lemma vcard_0 [simp]: "vcard 0 = 0" using Card_0 Card_def by auto lemma Ord_cardinal [simp,intro!]: "Ord(vcard a)" unfolding vcard_def by (metis Card_def Card_is_Ord cardinal_cong cardinal_eqpoll vcard_def) lemma gcard_big_0: "\ small X \ gcard X = 0" by (simp add: gcard_def) lemma gcard_eq_card: assumes "finite X" shows "gcard X = ord_of_nat (card X)" proof - have "\y. Ord y \ elts y \ X \ ord_of_nat (card X) \ y" by (metis assms eqpoll_finite_iff eqpoll_iff_card order_le_less ordertype_VWF_finite_nat ordertype_eq_Ord) then have "(LEAST i. Ord i \ elts i \ X) = ord_of_nat (card X)" by (simp add: assms eqpoll_iff_card finite_Ord_omega Least_equality) with assms show ?thesis by (simp add: finite_imp_small gcard_def) qed lemma gcard_empty_0 [simp]: "gcard {} = 0" by (simp add: gcard_eq_card) lemma gcard_single_1 [simp]: "gcard {x} = 1" by (simp add: gcard_eq_card one_V_def) lemma Card_gcard [iff]: "Card (gcard X)" by (metis Card_0 Card_def cardinal_idem gcard_big_0 gcardinal_cong small_eqpoll gcard_eq_vcard) lemma gcard_eqpoll: "small X \ elts (gcard X) \ X" by (metis cardinal_eqpoll eqpoll_trans gcard_eq_vcard gcardinal_cong small_eqpoll) lemma lepoll_imp_gcard_le: assumes "A \ B" "small B" shows "gcard A \ gcard B" proof - have "elts (gcard A) \ A" "elts (gcard B) \ B" by (meson assms gcard_eqpoll lepoll_small)+ with \A \ B\ show ?thesis by (metis Ord_cardinal Ord_linear2 eqpoll_sym gcard_eq_vcard gcardinal_cong lepoll_antisym lepoll_trans2 less_V_def less_eq_V_def subset_imp_lepoll) qed lemma gcard_image_le: assumes "small A" shows "gcard (f ` A) \ gcard A" using assms image_lepoll lepoll_imp_gcard_le by blast lemma subset_imp_gcard_le: assumes "A \ B" "small B" shows "gcard A \ gcard B" by (simp add: assms lepoll_imp_gcard_le subset_imp_lepoll) lemma gcard_le_lepoll: "\gcard A \ \; small A\ \ A \ elts \" by (meson eqpoll_sym gcard_eqpoll lepoll_trans1 less_eq_V_def subset_imp_lepoll) subsection\Cardinality of a set\ text\The cardinals are the initial ordinals.\ lemma Card_iff_initial: "Card \ \ Ord \ \ (\\. Ord \ \ \ < \ \ ~ elts \ \ elts \)" by (metis CardI Card_def Card_is_Ord not_less_Ord_Least vcard_def) lemma Card_\ [iff]: "Card \" by (meson CardI Ord_\ eqpoll_finite_iff infinite_Ord_omega infinite_\ leD) lemma lt_Card_imp_lesspoll: "\i < a; Card a; Ord i\ \ elts i \ elts a" by (meson Card_iff_initial less_eq_V_def less_imp_le lesspoll_def subset_imp_lepoll) lemma lepoll_imp_Card_le: assumes "elts a \ elts b" shows "vcard a \ vcard b" using assms lepoll_imp_gcard_le by fastforce lemma lepoll_cardinal_le: "\elts A \ elts i; Ord i\ \ vcard A \ i" by (metis Ord_Least Ord_linear2 dual_order.trans eqpoll_refl lepoll_imp_Card_le not_less_Ord_Least vcard_def) lemma cardinal_le_lepoll: "vcard A \ \ \ elts A \ elts \" by (meson cardinal_eqpoll eqpoll_sym lepoll_trans1 less_eq_V_def subset_imp_lepoll) lemma lesspoll_imp_Card_less: assumes "elts a \ elts b" shows "vcard a < vcard b" by (metis assms cardinal_eqpoll eqpoll_sym eqpoll_trans lepoll_imp_Card_le less_V_def lesspoll_def) lemma Card_Union [simp,intro]: assumes A: "\x. x \ A \ Card(x)" shows "Card(\A)" proof (rule CardI) show "Ord(\A)" using A by (simp add: Card_is_Ord Ord_Sup) next fix j assume j: "j < \A" "Ord j" hence "\c\A. j < c \ Card(c)" using A by (meson Card_is_Ord Ord_linear2 ZFC_in_HOL.Sup_least leD) then obtain c where c: "c\A" "j < c" "Card(c)" by blast hence jls: "elts j \ elts c" using j(2) lt_Card_imp_lesspoll by blast { assume eqp: "elts j \ elts (\A)" have "elts c \ elts (\A)" using c by (metis Card_def Sup_V_def ZFC_in_HOL.Sup_upper cardinal_le_lepoll j(1) not_less_0) also have "... \ elts j" by (rule eqpoll_sym [OF eqp]) also have "... \ elts c" by (rule jls) finally have "elts c \ elts c" . hence False by auto } thus "\ elts j \ elts (\A)" by blast qed lemma Card_UN: "(\x. x \ A \ Card(K x)) ==> Card(Sup (K ` A))" by blast subsection\Transfinite recursion for definitions based on the three cases of ordinals\ definition transrec3 :: "[V, [V,V]\V, [V,V\V]\V, V] \ V" where "transrec3 a b c \ transrec (\r x. if x=0 then a else if Limit x then c x (\y \ elts x. r y) else b(pred x) (r (pred x)))" lemma transrec3_0 [simp]: "transrec3 a b c 0 = a" by (simp add: transrec transrec3_def) lemma transrec3_succ [simp]: "transrec3 a b c (succ i) = b i (transrec3 a b c i)" by (simp add: transrec transrec3_def) lemma transrec3_Limit [simp]: "Limit i \ transrec3 a b c i = c i (\j \ elts i. transrec3 a b c j)" unfolding transrec3_def by (subst transrec) auto subsection \Cardinal Addition\ definition cadd :: "[V,V]\V" (infixl \\\ 65) where "\ \ \ \ vcard (\ \ \)" subsubsection\Cardinal addition is commutative\ lemma vsum_commute_eqpoll: "elts (a\b) \ elts (b\a)" proof - have "bij_betw (\z \ elts (a\b). sum_case Inr Inl z) (elts (a\b)) (elts (b\a))" unfolding bij_betw_def proof (intro conjI inj_onI) show "restrict (sum_case Inr Inl) (elts (a \ b)) ` elts (a \ b) = elts (b \ a)" apply auto apply (metis (no_types) imageI sum_case_Inr sum_iff) by (metis Inl_in_sum_iff imageI sum_case_Inl) qed auto then show ?thesis using eqpoll_def by blast qed lemma cadd_commute: "i \ j = j \ i" by (simp add: cadd_def cardinal_cong vsum_commute_eqpoll) subsubsection\Cardinal addition is associative\ lemma sum_assoc_bij: "bij_betw (\z \ elts ((a\b)\c). sum_case(sum_case Inl (\y. Inr(Inl y))) (\y. Inr(Inr y)) z) (elts ((a\b)\c)) (elts (a\(b\c)))" by (rule_tac f' = "sum_case (\x. Inl (Inl x)) (sum_case (\x. Inl (Inr x)) Inr)" in bij_betw_byWitness) auto lemma sum_assoc_eqpoll: "elts ((a\b)\c) \ elts (a\(b\c))" unfolding eqpoll_def by (metis sum_assoc_bij) lemma elts_vcard_vsum_eqpoll: "elts (vcard (i \ j)) \ Inl ` elts i \ Inr ` elts j" proof - have "elts (i \ j) \ Inl ` elts i \ Inr ` elts j" by (simp add: elts_vsum) then show ?thesis using cardinal_eqpoll eqpoll_trans by blast qed lemma cadd_assoc: "(i \ j) \ k = i \ (j \ k)" proof (unfold cadd_def, rule cardinal_cong) have "elts (vcard(i \ j) \ k) \ elts ((i \ j) \ k)" by (auto simp: disjnt_def elts_vsum elts_vcard_vsum_eqpoll intro: Un_eqpoll_cong) also have "\ \ elts (i \ (j \ k))" by (rule sum_assoc_eqpoll) also have "\ \ elts (i \ vcard(j \ k))" by (auto simp: disjnt_def elts_vsum elts_vcard_vsum_eqpoll [THEN eqpoll_sym] intro: Un_eqpoll_cong) finally show "elts (vcard (i \ j) \ k) \ elts (i \ vcard (j \ k))" . qed lemma cadd_left_commute: "j \ (i \ k) = i \ (j \ k)" using cadd_assoc cadd_commute by force lemmas cadd_ac = cadd_assoc cadd_commute cadd_left_commute text\0 is the identity for addition\ lemma vsum_0_eqpoll: "elts (0\a) \ elts a" by (simp add: elts_vsum) lemma cadd_0 [simp]: "Card \ \ 0 \ \ = \" by (metis Card_def cadd_def cardinal_cong vsum_0_eqpoll) lemma cadd_0_right [simp]: "Card \ \ \ \ 0 = \" by (simp add: cadd_commute) lemma vsum_lepoll_self: "elts a \ elts (a\b)" unfolding elts_vsum by (meson Inl_iff Un_upper1 inj_onI lepoll_def) lemma cadd_le_self: assumes \: "Card \" shows "\ \ \ \ a" proof (unfold cadd_def) have "\ \ vcard \" using Card_def \ by auto also have "\ \ vcard (\ \ a)" by (simp add: lepoll_imp_Card_le vsum_lepoll_self) finally show "\ \ vcard (\ \ a)" . qed text\Monotonicity of addition\ lemma cadd_le_mono: "\\' \ \; \' \ \\ \ \' \ \' \ \ \ \" unfolding cadd_def by (metis (no_types) lepoll_imp_Card_le less_eq_V_def subset_imp_lepoll sum_subset_iff) subsection\Cardinal multiplication\ definition cmult :: "[V,V]\V" (infixl \\\ 70) where "\ \ \ \ vcard (VSigma \ (\z. \))" subsubsection\Cardinal multiplication is commutative\ lemma prod_bij: "\bij_betw f A C; bij_betw g B D\ \ bij_betw (\(x, y). (f x, g y)) (A \ B) (C \ D)" apply (rule bij_betw_byWitness [where f' = "\(x,y). (inv_into A f x, inv_into B g y)"]) apply (auto simp: bij_betw_inv_into_left bij_betw_inv_into_right bij_betwE) using bij_betwE bij_betw_inv_into apply blast+ done lemma cmult_commute: "i \ j = j \ i" proof - have "(\(x, y). \x, y\) ` (elts i \ elts j) \ (\(x, y). \x, y\) ` (elts j \ elts i)" by (simp add: times_commute_eqpoll) then show ?thesis unfolding cmult_def using cardinal_cong elts_VSigma by auto qed subsubsection\Cardinal multiplication is associative\ lemma elts_vcard_VSigma_eqpoll: "elts (vcard (vtimes i j)) \ elts i \ elts j" proof - have "elts (vtimes i j) \ elts i \ elts j" by (simp add: elts_VSigma) then show ?thesis using cardinal_eqpoll eqpoll_trans by blast qed lemma elts_cmult: "elts (\' \ \) \ elts \' \ elts \" by (simp add: cmult_def elts_vcard_VSigma_eqpoll) lemma cmult_assoc: "(i \ j) \ k = i \ (j \ k)" unfolding cmult_def proof (rule cardinal_cong) have "elts (vcard (vtimes i j)) \ elts k \ (elts i \ elts j) \ elts k" by (blast intro: times_eqpoll_cong elts_vcard_VSigma_eqpoll cardinal_eqpoll) also have "\ \ elts i \ (elts j \ elts k)" by (rule times_assoc_eqpoll) also have "\ \ elts i \ elts (vcard (vtimes j k))" by (blast intro: times_eqpoll_cong elts_vcard_VSigma_eqpoll cardinal_eqpoll eqpoll_sym) finally show "elts (VSigma (vcard (vtimes i j)) (\z. k)) \ elts (VSigma i (\z. vcard (vtimes j k)))" by (simp add: elts_VSigma) qed subsubsection\Cardinal multiplication distributes over addition\ lemma cadd_cmult_distrib: "(i \ j) \ k = (i \ k) \ (j \ k)" unfolding cadd_def cmult_def proof (rule cardinal_cong) have "elts (vtimes (vcard (i \ j)) k) \ elts (vcard (vsum i j)) \ elts k" using cardinal_eqpoll elts_vcard_VSigma_eqpoll eqpoll_sym eqpoll_trans by blast also have "\ \ (Inl ` elts i \ Inr ` elts j) \ elts k" using elts_vcard_vsum_eqpoll times_eqpoll_cong by blast also have "\ \ (Inl ` elts i) \ elts k \ (Inr ` elts j) \ elts k" by (simp add: Sigma_Un_distrib1) also have "\ \ elts (vtimes i k \ vtimes j k)" unfolding Plus_def by (auto simp: elts_vsum elts_VSigma disjnt_iff intro!: Un_eqpoll_cong times_eqpoll_cong) also have "\ \ elts (vcard (vtimes i k \ vtimes j k))" by (simp add: cardinal_eqpoll eqpoll_sym) also have "\ \ elts (vcard (vtimes i k) \ vcard (vtimes j k))" by (metis cadd_assoc cadd_def cardinal_cong cardinal_eqpoll vsum_0_eqpoll vsum_commute_eqpoll) finally show "elts (VSigma (vcard (i \ j)) (\z. k)) \ elts (vcard (vtimes i k) \ vcard (vtimes j k))" . qed text\Multiplication by 0 yields 0\ lemma cmult_0 [simp]: "0 \ i = 0" using Card_0 Card_def cmult_def by auto text\1 is the identity for multiplication\ lemma cmult_1 [simp]: assumes "Card \" shows "1 \ \ = \" proof - have "elts (vtimes (set {0}) \) \ elts \" by (auto simp: elts_VSigma intro!: times_singleton_eqpoll) then show ?thesis by (metis Card_def assms cardinal_cong cmult_def elts_1 set_of_elts) qed subsection\Some inequalities for multiplication\ lemma cmult_square_le: assumes "Card \" shows "\ \ \ \ \" proof - have "elts \ \ elts (\ \ \)" using times_square_lepoll [of "elts \"] cmult_def elts_vcard_VSigma_eqpoll eqpoll_sym lepoll_trans2 by fastforce then show ?thesis using Card_def assms cmult_def lepoll_cardinal_le by fastforce qed text\Multiplication by a non-empty set\ lemma cmult_le_self: assumes "Card \" "\ \ 0" shows "\ \ \ \ \" proof - have "\ = vcard \" using Card_def \Card \\ by blast also have "\ \ vcard (vtimes \ \)" apply (rule lepoll_imp_Card_le) apply (simp add: elts_VSigma) by (metis ZFC_in_HOL.ext \\ \ 0\ elts_0 lepoll_times1) also have "\ = \ \ \" by (simp add: cmult_def) finally show ?thesis . qed text\Monotonicity of multiplication\ lemma cmult_le_mono: "\\' \ \; \' \ \\ \ \' \ \' \ \ \ \" unfolding cmult_def by (auto simp: elts_VSigma intro!: lepoll_imp_Card_le times_lepoll_mono subset_imp_lepoll) lemma vcard_Sup_le_cmult: assumes "small U" and \: "\x. x \ U \ vcard x \ \" shows "vcard (\U) \ vcard (set U) \ \" proof - have "\f. f \ elts x \ elts \ \ inj_on f (elts x)" if "x \ U" for x using \ [OF that] by (metis cardinal_le_lepoll image_subset_iff_funcset lepoll_def) then obtain \ where \: "\x. x \ U \ (\ x) \ elts x \ elts \ \ inj_on (\ x) (elts x)" by metis define u where "u \ \y. @x. x \ U \ y \ elts x" have u: "u y \ U \ y \ elts (u y)" if "y \ \(elts ` U)" for y unfolding u_def by (metis (mono_tags, lifting)that someI2_ex UN_iff) define \ where "\ \ \y. (u y, \ (u y) y)" have U: "elts (vcard (set U)) \ U" by (metis \small U\ cardinal_eqpoll elts_of_set) have "elts (\U) = \(elts ` U)" using \small U\ by blast also have "\ \ U \ elts \" unfolding lepoll_def proof (intro exI conjI) show "inj_on \ (\ (elts ` U))" using \ u by (smt (verit) \_def inj_on_def prod.inject) show "\ ` \ (elts ` U) \ U \ elts \" using \ u by (auto simp: \_def) qed also have "\ \ elts (vcard (set U) \ \)" using U elts_cmult eqpoll_sym eqpoll_trans times_eqpoll_cong by blast finally have "elts (\ U) \ elts (vcard (set U) \ \)" . then show ?thesis by (simp add: cmult_def lepoll_cardinal_le) qed subsection\The finite cardinals\ lemma succ_lepoll_succD: "elts (succ(m)) \ elts (succ(n)) \ elts m \ elts n" by (simp add: insert_lepoll_insertD) text\Congruence law for @{text succ} under equipollence\ lemma succ_eqpoll_cong: "elts a \ elts b \ elts (succ(a)) \ elts (succ(b))" by (simp add: succ_def insert_eqpoll_cong) lemma sum_succ_eqpoll: "elts (succ a \ b) \ elts (succ(a\b))" unfolding eqpoll_def proof (rule exI) let ?f = "\z. if z=Inl a then a\b else z" let ?g = "\z. if z=a\b then Inl a else z" show "bij_betw ?f (elts (succ a \ b)) (elts (succ (a \ b)))" apply (rule bij_betw_byWitness [where f' = ?g], auto) apply (metis Inl_in_sum_iff mem_not_refl) by (metis Inr_in_sum_iff mem_not_refl) qed lemma cadd_succ: "succ m \ n = vcard (succ(m \ n))" proof (unfold cadd_def) have [intro]: "elts (m \ n) \ elts (vcard (m \ n))" using cardinal_eqpoll eqpoll_sym by blast have "vcard (succ m \ n) = vcard (succ(m \ n))" by (rule sum_succ_eqpoll [THEN cardinal_cong]) also have "\ = vcard (succ(vcard (m \ n)))" by (blast intro: succ_eqpoll_cong cardinal_cong) finally show "vcard (succ m \ n) = vcard (succ(vcard (m \ n)))" . qed lemma nat_cadd_eq_add: "ord_of_nat m \ ord_of_nat n = ord_of_nat (m + n)" proof (induct m) case (Suc m) thus ?case by (metis Card_def Card_ord_of_nat add_Suc cadd_succ ord_of_nat.simps(2)) qed auto lemma vcard_disjoint_sup: assumes "x \ y = 0" shows "vcard (x \ y) = vcard x \ vcard y" proof - have "elts (x \ y) \ elts (x \ y)" unfolding eqpoll_def proof (rule exI) let ?f = "\z. if z \ elts x then Inl z else Inr z" let ?g = "sum_case id id" show "bij_betw ?f (elts (x \ y)) (elts (x \ y))" by (rule bij_betw_byWitness [where f' = ?g]) (use assms V_disjoint_iff in auto) qed then show ?thesis by (metis cadd_commute cadd_def cardinal_cong cardinal_idem vsum_0_eqpoll cadd_assoc) qed lemma vcard_sup: "vcard (x \ y) \ vcard x \ vcard y" proof - have "elts (x \ y) \ elts (x \ y)" unfolding lepoll_def proof (intro exI conjI) let ?f = "\z. if z \ elts x then Inl z else Inr z" show "inj_on ?f (elts (x \ y))" by (simp add: inj_on_def) show "?f ` elts (x \ y) \ elts (x \ y)" by force qed then show ?thesis using cadd_ac by (metis cadd_def cardinal_cong cardinal_idem lepoll_imp_Card_le vsum_0_eqpoll) qed subsection\Infinite cardinals\ definition InfCard :: "V\bool" where "InfCard \ \ Card \ \ \ \ \" lemma InfCard_iff: "InfCard \ \ Card \ \ infinite (elts \)" proof (cases "\ \ \") case True then show ?thesis using inj_ord_of_nat lepoll_def less_eq_V_def by (auto simp: InfCard_def \_def infinite_le_lepoll) next case False then show ?thesis using Card_iff_initial InfCard_def infinite_Ord_omega by blast qed lemma InfCard_ge_ord_of_nat: assumes "InfCard \" shows "ord_of_nat n \ \" using InfCard_def assms ord_of_nat_le_omega by blast lemma InfCard_not_0[iff]: "\ InfCard 0" by (simp add: InfCard_iff) definition csucc :: "V\V" where "csucc \ \ LEAST \'. Ord \' \ (Card \' \ \ < \')" lemma less_vcard_VPow: "vcard A < vcard (VPow A)" proof (rule lesspoll_imp_Card_less) show "elts A \ elts (VPow A)" by (simp add: elts_VPow down inj_on_def lesspoll_Pow_self) qed lemma greater_Card: assumes "Card \" shows "\ < vcard (VPow \)" proof - have "\ = vcard \" using Card_def assms by blast also have "\ < vcard (VPow \)" proof (rule lesspoll_imp_Card_less) show "elts \ \ elts (VPow \)" by (simp add: elts_VPow down inj_on_def lesspoll_Pow_self) qed finally show ?thesis . qed lemma assumes "Card \" shows Card_csucc [simp]: "Card (csucc \)" and less_csucc [simp]: "\ < csucc \" proof - have "Card (csucc \) \ \ < csucc \" unfolding csucc_def proof (rule Ord_LeastI2) show "Card (vcard (VPow \)) \ \ < (vcard (VPow \))" using Card_def assms greater_Card by auto qed auto then show "Card (csucc \)" "\ < csucc \" by auto qed lemma le_csucc: assumes "Card \" shows "\ \ csucc \" by (simp add: assms less_csucc less_imp_le) lemma csucc_le: "\Card \; \ \ elts \\ \ csucc \ \ \" unfolding csucc_def by (simp add: Card_is_Ord Ord_Least_le OrdmemD) lemma finite_csucc: "a \ elts \ \ csucc a = succ a" unfolding csucc_def proof (rule Least_equality) show "Ord (ZFC_in_HOL.succ a) \ Card (ZFC_in_HOL.succ a) \ a < ZFC_in_HOL.succ a" if "a \ elts \" using that by (auto simp: less_V_def less_eq_V_def nat_into_Card) show "ZFC_in_HOL.succ a \ y" if "a \ elts \" and "Ord y \ Card y \ a < y" for y :: V using that using Ord_mem_iff_lt dual_order.strict_implies_order by fastforce qed lemma Finite_imp_cardinal_cons [simp]: assumes FA: "finite A" and a: "a \ A" shows "vcard (set (insert a A)) = csucc(vcard (set A))" proof - show ?thesis unfolding csucc_def proof (rule Least_equality [THEN sym]) have "small A" by (simp add: FA Finite_V) then have "\ elts (set A) \ elts (set (insert a A))" using FA a eqpoll_imp_lepoll eqpoll_sym finite_insert_lepoll by fastforce then show "Ord (vcard (set (insert a A))) \ Card (vcard (set (insert a A))) \ vcard (set A) < vcard (set (insert a A))" by (simp add: Card_def lesspoll_imp_Card_less lesspoll_def subset_imp_lepoll subset_insertI) show "vcard (set (insert a A)) \ i" if "Ord i \ Card i \ vcard (set A) < i" for i proof - have "elts (vcard (set A)) \ A" by (metis FA finite_imp_small cardinal_eqpoll elts_of_set) then have less: "A \ elts i" using eq_lesspoll_trans eqpoll_sym lt_Card_imp_lesspoll that by blast show ?thesis using that less by (auto simp: less_imp_insert_lepoll lepoll_cardinal_le) qed qed qed lemma vcard_finite_set: "finite A \ vcard (set A) = ord_of_nat (card A)" by (induction A rule: finite_induct) (auto simp: set_empty \_def finite_csucc) lemma lt_csucc_iff: assumes "Ord \" "Card \" shows "\ < csucc \ \ vcard \ \ \" proof show "vcard \ \ \" if "\ < csucc \" proof - have "vcard \ \ csucc \" by (meson \Ord \\ dual_order.trans lepoll_cardinal_le lepoll_refl less_le_not_le that) then show ?thesis by (metis (no_types) Card_def Card_iff_initial Ord_linear2 Ord_mem_iff_lt assms cardinal_eqpoll cardinal_idem csucc_le eq_iff eqpoll_sym that) qed show "\ < csucc \" if "vcard \ \ \" proof - have "\ csucc \ \ \" using that by (metis Card_csucc Card_def assms(2) le_less_trans lepoll_imp_Card_le less_csucc less_eq_V_def less_le_not_le subset_imp_lepoll) then show ?thesis by (meson Card_csucc Card_is_Ord Ord_linear2 assms) qed qed lemma Card_lt_csucc_iff: "\Card \'; Card \\ \ (\' < csucc \) = (\' \ \)" by (simp add: lt_csucc_iff Card_cardinal_eq Card_is_Ord) lemma csucc_lt_csucc_iff: "\Card \'; Card \\ \ (csucc \' < csucc \) = (\' < \)" by (metis Card_csucc Card_is_Ord Card_lt_csucc_iff Ord_not_less) lemma csucc_le_csucc_iff: "\Card \'; Card \\ \ (csucc \' \ csucc \) = (\' \ \)" by (metis Card_csucc Card_is_Ord Card_lt_csucc_iff Ord_not_less) lemma csucc_0 [simp]: "csucc 0 = 1" by (simp add: finite_csucc one_V_def) lemma Card_Un [simp,intro]: assumes "Card x" "Card y" shows "Card(x \ y)" by (metis Card_is_Ord Ord_linear_le assms sup.absorb2 sup.orderE) lemma InfCard_csucc: "InfCard \ \ InfCard (csucc \)" using InfCard_def le_csucc by auto text\Kunen's Lemma 10.11\ lemma InfCard_is_Limit: assumes "InfCard \" shows "Limit \" proof (rule non_succ_LimitI) show "\ \ 0" using InfCard_def assms mem_not_refl by blast show "Ord \" using Card_is_Ord InfCard_def assms by blast show "ZFC_in_HOL.succ y \ \" for y proof assume "succ y = \" then have "Card (succ y)" using InfCard_def assms by auto moreover have "\ \ y" by (metis InfCard_iff Ord_in_Ord \Ord \\ \ZFC_in_HOL.succ y = \\ assms elts_succ finite_insert infinite_Ord_omega insertI1) moreover have "elts y \ elts (succ y)" using InfCard_iff \ZFC_in_HOL.succ y = \\ assms eqpoll_sym infinite_insert_eqpoll by fastforce ultimately show False by (metis Card_iff_initial Ord_in_Ord OrdmemD elts_succ insertI1) qed qed subsection\Toward's Kunen's Corollary 10.13 (1)\ text\Kunen's Theorem 10.12\ lemma InfCard_csquare_eq: assumes "InfCard(\)" shows "\ \ \ = \" using infinite_times_eqpoll_self [of "elts \"] assms unfolding InfCard_iff Card_def by (metis cardinal_cong cardinal_eqpoll cmult_def elts_vcard_VSigma_eqpoll eqpoll_trans) lemma InfCard_le_cmult_eq: assumes "InfCard \" "\ \ \" "\ \ 0" shows "\ \ \ = \" proof (rule order_antisym) have "\ \ \ \ \ \ \" by (simp add: assms(2) cmult_le_mono) also have "\ \ \" by (simp add: InfCard_csquare_eq assms(1)) finally show "\ \ \ \ \" . show "\ \ \ \ \" using InfCard_def assms(1) assms(3) cmult_le_self by auto qed text\Kunen's Corollary 10.13 (1), for cardinal multiplication\ lemma InfCard_cmult_eq: "\InfCard \; InfCard \\ \ \ \ \ = \ \ \" by (metis Card_is_Ord InfCard_def InfCard_le_cmult_eq Ord_linear_le cmult_commute inf_sup_aci(5) mem_not_refl sup.orderE sup_V_0_right zero_in_omega) lemma cmult_succ: "succ(m) \ n = n \ (m \ n)" unfolding cmult_def cadd_def proof (rule cardinal_cong) have "elts (vtimes (ZFC_in_HOL.succ m) n) \ elts n <+> elts m \ elts n" by (simp add: elts_VSigma prod_insert_eqpoll) also have "\ \ elts (n \ vcard (vtimes m n))" unfolding elts_VSigma elts_vsum Plus_def proof (rule Un_eqpoll_cong) show "(Sum_Type.Inr ` (elts m \ elts n)::(V + V \ V) set) \ Inr ` elts (vcard (vtimes m n))" by (simp add: elts_vcard_VSigma_eqpoll eqpoll_sym) qed (auto simp: disjnt_def) finally show "elts (vtimes (ZFC_in_HOL.succ m) n) \ elts (n \ vcard (vtimes m n))" . qed lemma cmult_2: assumes "Card n" shows "ord_of_nat 2 \ n = n \ n" proof - have "ord_of_nat 2 = succ (succ 0)" by force then show ?thesis by (simp add: cmult_succ assms) qed lemma InfCard_cdouble_eq: assumes "InfCard \" shows "\ \ \ = \" proof - have "\ \ \ = \ \ ord_of_nat 2" using InfCard_def assms cmult_2 cmult_commute by auto also have "\ = \" by (simp add: InfCard_le_cmult_eq InfCard_ge_ord_of_nat assms) finally show ?thesis . qed text\Corollary 10.13 (1), for cardinal addition\ lemma InfCard_le_cadd_eq: "\InfCard \; \ \ \\ \ \ \ \ = \" by (metis InfCard_cdouble_eq InfCard_def antisym cadd_le_mono cadd_le_self) lemma InfCard_cadd_eq: "\InfCard \; InfCard \\ \ \ \ \ = \ \ \" by (metis Card_iff_initial InfCard_def InfCard_le_cadd_eq Ord_linear_le cadd_commute sup.absorb2 sup.orderE) lemma csucc_le_Card_iff: "\Card \'; Card \\ \ csucc \' \ \ \ \' < \" by (metis Card_csucc Card_is_Ord Card_lt_csucc_iff Ord_not_le) lemma cadd_InfCard_le: assumes "\ \ \" "\ \ \" "InfCard \" shows "\ \ \ \ \" using assms by (metis InfCard_cdouble_eq cadd_le_mono) lemma cmult_InfCard_le: assumes "\ \ \" "\ \ \" "InfCard \" shows "\ \ \ \ \" using assms by (metis InfCard_csquare_eq cmult_le_mono) subsection \The Aleph-seqence\ text \This is the well-known transfinite enumeration of the cardinal numbers.\ definition Aleph :: "V \ V" (\\_\ [90] 90) where "Aleph \ transrec (\f x. \ \ \((\y. csucc(f y)) ` elts x))" lemma Aleph: "Aleph \ = \ \ (\y\elts \. csucc (Aleph y))" by (simp add: Aleph_def transrec[of _ \]) lemma InfCard_Aleph [simp, intro]: "InfCard(Aleph x)" proof (induction x rule: eps_induct) case (step x) then show ?case by (simp add: Aleph [of x] InfCard_def Card_UN step) qed lemma Card_Aleph [simp, intro]: "Card(Aleph x)" using InfCard_def by auto lemma Aleph_0 [simp]: "\0 = \" by (simp add: Aleph [of 0]) lemma mem_Aleph_succ: "\\ \ elts (Aleph (succ \))" apply (simp add: Aleph [of "succ \"]) by (meson InfCard_Aleph Card_csucc Card_is_Ord InfCard_def Ord_mem_iff_lt less_csucc) lemma Aleph_lt_succD [simp]: "\\ < Aleph (succ \)" by (simp add: InfCard_is_Limit Limit_is_Ord OrdmemD mem_Aleph_succ) lemma Aleph_succ [simp]: "Aleph (succ x) = csucc (Aleph x)" proof (rule antisym) show "Aleph (ZFC_in_HOL.succ x) \ csucc (Aleph x)" apply (simp add: Aleph [of "succ x"]) by (metis Aleph InfCard_Aleph InfCard_def Sup_V_insert le_csucc le_sup_iff order_refl replacement small_elts) show "csucc (Aleph x) \ Aleph (ZFC_in_HOL.succ x)" by (force simp add: Aleph [of "succ x"]) qed lemma csucc_Aleph_le_Aleph: "\ \ elts \ \ csucc (\\) \ \\" by (metis Aleph ZFC_in_HOL.SUP_le_iff replacement small_elts sup_ge2) lemma Aleph_in_Aleph: "\ \ elts \ \ \\ \ elts (\\)" using csucc_Aleph_le_Aleph mem_Aleph_succ by auto lemma Aleph_Limit: assumes "Limit \" shows "Aleph \ = \ (Aleph ` elts \)" proof - have [simp]: "\ \ 0" using assms by auto show ?thesis proof (rule antisym) show "Aleph \ \ \ (Aleph ` elts \)" apply (simp add: Aleph [of \]) by (metis (mono_tags, lifting) Aleph_0 Aleph_succ Limit_def ZFC_in_HOL.SUP_le_iff ZFC_in_HOL.Sup_upper assms imageI replacement small_elts) show "\ (Aleph ` elts \) \ Aleph \" apply (simp add: cSup_le_iff) by (meson InfCard_Aleph InfCard_def csucc_Aleph_le_Aleph le_csucc order_trans) qed qed lemma Aleph_increasing: assumes ab: "\ < \" "Ord \" "Ord \" shows "\\ < \\" by (meson Aleph_in_Aleph InfCard_Aleph Card_iff_initial InfCard_def Ord_mem_iff_lt assms) lemma countable_iff_le_Aleph0: "countable (elts A) \ vcard A \ \0" proof show "vcard A \ \0" if "countable (elts A)" proof (cases "finite (elts A)") case True then show ?thesis using vcard_finite_set by fastforce next case False then have "elts \ \ elts A" using countableE_infinite [OF that] by (simp add: eqpoll_def \_def) (meson bij_betw_def bij_betw_inv bij_betw_trans inj_ord_of_nat) then show ?thesis using Card_\ Card_def cardinal_cong vcard_def by auto qed show "countable (elts A)" if "vcard A \ Aleph 0" proof - have "elts A \ elts \" using cardinal_le_lepoll [OF that] by simp then show ?thesis by (simp add: countable_iff_lepoll \_def inj_ord_of_nat) qed qed lemma Aleph_csquare_eq [simp]: "\\ \ \\ = \\" using InfCard_csquare_eq by auto lemma vcard_Aleph [simp]: "vcard (\\) = \\" using Card_def InfCard_Aleph InfCard_def by auto lemma omega_le_Aleph [simp]: "\ \ \\" using InfCard_def by auto lemma finite_iff_less_Aleph0: "finite (elts x) \ vcard x < \" proof show "finite (elts x) \ vcard x < \" by (metis Card_\ Card_def finite_lesspoll_infinite infinite_\ lesspoll_imp_Card_less) show "vcard x < \ \ finite (elts x)" by (meson Ord_cardinal cardinal_eqpoll eqpoll_finite_iff infinite_Ord_omega less_le_not_le) qed lemma countable_iff_vcard_less1: "countable (elts x) \ vcard x < \1" by (simp add: countable_iff_le_Aleph0 lt_csucc_iff one_V_def) lemma countable_infinite_vcard: "countable (elts x) \ infinite (elts x) \ vcard x = \0" by (metis Aleph_0 countable_iff_le_Aleph0 dual_order.refl finite_iff_less_Aleph0 less_V_def) subsection \The ordinal @{term "\1"}\ abbreviation "\1 \ Aleph 1" lemma Ord_\1 [simp]: "Ord \1" by (metis Ord_cardinal vcard_Aleph) lemma omega_\1 [iff]: "\ \ elts \1" by (metis Aleph_0 mem_Aleph_succ one_V_def) lemma ord_of_nat_\1 [iff]: "ord_of_nat n \ elts \1" using Ord_\1 Ord_trans by blast lemma countable_iff_less_\1: assumes "Ord \" shows "countable (elts \) \ \ < \1" by (simp add: assms countable_iff_le_Aleph0 lt_csucc_iff one_V_def) lemma less_\1_imp_countable: assumes "\ \ elts \1" shows "countable (elts \)" using Ord_\1 Ord_in_Ord OrdmemD assms countable_iff_less_\1 by blast lemma \1_gt0 [simp]: "\1 > 0" using Ord_\1 Ord_trans OrdmemD by blast lemma \1_gt1 [simp]: "\1 > 1" using Ord_\1 OrdmemD \_gt1 less_trans by blast lemma Limit_\1 [simp]: "Limit \1" by (simp add: InfCard_def InfCard_is_Limit le_csucc one_V_def) end diff --git a/thys/ZFC_in_HOL/ZFC_in_HOL.thy b/thys/ZFC_in_HOL/ZFC_in_HOL.thy --- a/thys/ZFC_in_HOL/ZFC_in_HOL.thy +++ b/thys/ZFC_in_HOL/ZFC_in_HOL.thy @@ -1,1299 +1,1299 @@ section \The ZF Axioms, Ordinals and Transfinite Recursion\ theory ZFC_in_HOL imports ZFC_Library begin subsection\Syntax and axioms\ hide_const (open) list.set Sum subset unbundle lattice_syntax typedecl V text\Presentation refined by Dmitriy Traytel\ axiomatization elts :: "V \ V set" where ext [intro?]: "elts x = elts y \ x=y" and down_raw: "Y \ elts x \ Y \ range elts" and Union_raw: "X \ range elts \ Union (elts ` X) \ range elts" and Pow_raw: "X \ range elts \ inv elts ` Pow X \ range elts" and replacement_raw: "X \ range elts \ f ` X \ range elts" and inf_raw: "range (g :: nat \ V) \ range elts" and foundation: "wf {(x,y). x \ elts y}" lemma mem_not_refl [simp]: "i \ elts i" using wf_not_refl [OF foundation] by force lemma mem_not_sym: "\ (x \ elts y \ y \ elts x)" using wf_not_sym [OF foundation] by force text \A set is small if it can be injected into the extension of a V-set.\ definition small :: "'a set \ bool" where "small X \ \V_of :: 'a \ V. inj_on V_of X \ V_of ` X \ range elts" lemma small_empty [iff]: "small {}" by (simp add: small_def down_raw) lemma small_iff_range: "small X \ X \ range elts" apply (simp add: small_def) by (metis inj_on_id2 replacement_raw the_inv_into_onto) lemma small_eqpoll: "small A \ (\x. elts x \ A)" unfolding small_def by (metis UNIV_I bij_betw_def eqpoll_def eqpoll_sym imageE image_eqI) text\Small classes can be mapped to sets.\ definition set :: "V set \ V" where "set X \ (if small X then inv elts X else inv elts {})" lemma set_of_elts [simp]: "set (elts x) = x" by (force simp add: ext set_def f_inv_into_f small_def) lemma elts_of_set [simp]: "elts (set X) = (if small X then X else {})" by (simp add: ZFC_in_HOL.set_def down_raw f_inv_into_f small_iff_range) lemma down: "Y \ elts x \ small Y" by (simp add: down_raw small_iff_range) lemma Union [intro]: "small X \ small (Union (elts ` X))" by (simp add: Union_raw small_iff_range) lemma Pow: "small X \ small (set ` Pow X)" unfolding small_iff_range using Pow_raw set_def down by force declare replacement_raw [intro,simp] lemma replacement [intro,simp]: assumes "small X" shows "small (f ` X)" proof - let ?A = "inv_into X f ` (f ` X)" have AX: "?A \ X" by (simp add: image_subsetI inv_into_into) have inj: "inj_on f ?A" by (simp add: f_inv_into_f inj_on_def) have injo: "inj_on (inv_into X f) (f ` X)" using inj_on_inv_into by blast have "\V_of. inj_on V_of (f ` X) \ V_of ` f ` X \ range elts" if "inj_on V_of X" and "V_of ` X = elts x" for V_of :: "'a \ V" and x proof (intro exI conjI) show "inj_on (V_of \ inv_into X f) (f ` X)" by (meson \inv_into X f ` f ` X \ X\ comp_inj_on inj_on_subset injo that) have "(\x. V_of (inv_into X f (f x))) ` X = elts (set (V_of ` ?A))" by (metis AX down elts_of_set image_image image_mono that(2)) then show "(V_of \ inv_into X f) ` f ` X \ range elts" by (metis image_comp image_image rangeI) qed then show ?thesis using assms by (auto simp: small_def) qed lemma small_image_iff [simp]: "inj_on f A \ small (f ` A) \ small A" by (metis replacement the_inv_into_onto) text \A little bootstrapping is needed to characterise @{term small} for sets of arbitrary type.\ lemma inf: "small (range (g :: nat \ V))" by (simp add: inf_raw small_iff_range) lemma small_image_nat_V [simp]: "small (g ` N)" for g :: "nat \ V" by (metis (mono_tags, opaque_lifting) down elts_of_set image_iff inf rangeI subsetI) lemma Finite_V: fixes X :: "V set" assumes "finite X" shows "small X" using ex_bij_betw_nat_finite [OF assms] unfolding bij_betw_def by (metis small_image_nat_V) lemma small_insert_V: fixes X :: "V set" assumes "small X" shows "small (insert a X)" proof (cases "finite X") case True then show ?thesis by (simp add: Finite_V) next case False show ?thesis using infinite_imp_bij_betw2 [OF False] by (metis replacement Un_insert_right assms bij_betw_imp_surj_on sup_bot.right_neutral) qed lemma small_UN_V [simp,intro]: fixes B :: "'a \ V set" assumes X: "small X" and B: "\x. x \ X \ small (B x)" shows "small (\x\X. B x)" proof - have "(\ (elts ` (\x. ZFC_in_HOL.set (B x)) ` X)) = (\ (B ` X))" using B by force then show ?thesis using Union [OF replacement [OF X, of "\x. ZFC_in_HOL.set (B x)"]] by simp qed definition vinsert where "vinsert x y \ set (insert x (elts y))" lemma elts_vinsert [simp]: "elts (vinsert x y) = insert x (elts y)" using down small_insert_V vinsert_def by auto definition succ where "succ x \ vinsert x x" lemma elts_succ [simp]: "elts (succ x) = insert x (elts x)" by (simp add: succ_def) lemma finite_imp_small: assumes "finite X" shows "small X" using assms proof induction case empty then show ?case by simp next case (insert a X) then obtain V_of u where u: "inj_on V_of X" "V_of ` X = elts u" by (meson small_def image_iff) show ?case unfolding small_def proof (intro exI conjI) show "inj_on (V_of(a:=u)) (insert a X)" using u apply (clarsimp simp add: inj_on_def) by (metis image_eqI mem_not_refl) have "(V_of(a:=u)) ` insert a X = elts (vinsert u u)" using insert.hyps(2) u(2) by auto then show "(V_of(a:=u)) ` insert a X \ range elts" by (blast intro: elim: ) qed qed lemma small_insert: assumes "small X" shows "small (insert a X)" proof (cases "finite X") case True then show ?thesis by (simp add: finite_imp_small) next case False show ?thesis using infinite_imp_bij_betw2 [OF False] by (metis replacement Un_insert_right assms bij_betw_imp_surj_on sup_bot.right_neutral) qed lemma smaller_than_small: assumes "small A" "B \ A" shows "small B" using assms by (metis down elts_of_set image_mono small_def small_iff_range subset_inj_on) lemma small_insert_iff [iff]: "small (insert a X) \ small X" by (meson small_insert smaller_than_small subset_insertI) lemma small_iff: "small X \ (\x. X = elts x)" by (metis down elts_of_set subset_refl) lemma small_elts [iff]: "small (elts x)" by (auto simp: small_iff) lemma small_diff [iff]: "small (elts a - X)" by (meson Diff_subset down) lemma small_set [simp]: "small (list.set xs)" by (simp add: ZFC_in_HOL.finite_imp_small) lemma small_upair: "small {x,y}" by simp lemma small_Un_elts: "small (elts x \ elts y)" using Union [OF small_upair] by auto lemma small_eqcong: "\small X; X \ Y\ \ small Y" by (metis bij_betw_imp_surj_on eqpoll_def replacement) lemma lepoll_small: "\small Y; X \ Y\ \ small X" by (meson lepoll_iff replacement smaller_than_small) lemma big_UNIV [simp]: "\ small (UNIV::V set)" (is "\ small ?U") proof assume "small ?U" then have "small A" for A :: "V set" by (metis (full_types) UNIV_I down small_iff subsetI) then have "range elts = UNIV" by (meson small_iff surj_def) then show False by (metis Cantors_paradox Pow_UNIV) qed lemma inj_on_set: "inj_on set (Collect small)" by (metis elts_of_set inj_onI mem_Collect_eq) lemma set_injective [simp]: "\small X; small Y\ \ set X = set Y \ X=Y" by (metis elts_of_set) subsection\Type classes and other basic setup\ instantiation V :: zero begin definition zero_V where "0 \ set {}" instance .. end lemma elts_0 [simp]: "elts 0 = {}" by (simp add: zero_V_def) lemma set_empty [simp]: "set {} = 0" by (simp add: zero_V_def) instantiation V :: one begin definition one_V where "1 \ succ 0" instance .. end lemma elts_1 [simp]: "elts 1 = {0}" by (simp add: one_V_def) lemma insert_neq_0 [simp]: "set (insert a X) = 0 \ \ small X" unfolding zero_V_def by (metis elts_of_set empty_not_insert set_of_elts small_insert_iff) lemma elts_eq_empty_iff [simp]: "elts x = {} \ x=0" by (auto simp: ZFC_in_HOL.ext) instantiation V :: distrib_lattice begin definition inf_V where "inf_V x y \ set (elts x \ elts y)" definition sup_V where "sup_V x y \ set (elts x \ elts y)" definition less_eq_V where "less_eq_V x y \ elts x \ elts y" definition less_V where "less_V x y \ less_eq x y \ x \ (y::V)" instance proof show "(x < y) = (x \ y \ \ y \ x)" for x :: V and y :: V using ext less_V_def less_eq_V_def by auto show "x \ x" for x :: V by (simp add: less_eq_V_def) show "x \ z" if "x \ y" "y \ z" for x y z :: V using that by (auto simp: less_eq_V_def) show "x = y" if "x \ y" "y \ x" for x y :: V using that by (simp add: ext less_eq_V_def) show "inf x y \ x" for x y :: V by (metis down elts_of_set inf_V_def inf_sup_ord(1) less_eq_V_def) show "inf x y \ y" for x y :: V by (metis Int_lower2 down elts_of_set inf_V_def less_eq_V_def) show "x \ inf y z" if "x \ y" "x \ z" for x y z :: V proof - have "small (elts y \ elts z)" by (meson down inf.cobounded1) then show ?thesis using elts_of_set inf_V_def less_eq_V_def that by auto qed show "x \ x \ y" "y \ x \ y" for x y :: V by (simp_all add: less_eq_V_def small_Un_elts sup_V_def) show "sup y z \ x" if "y \ x" "z \ x" for x y z :: V using less_eq_V_def sup_V_def that by auto show "sup x (inf y z) = inf (x \ y) (sup x z)" for x y z :: V proof - have "small (elts y \ elts z)" by (meson down inf.cobounded2) then show ?thesis by (simp add: Un_Int_distrib inf_V_def small_Un_elts sup_V_def) qed qed end lemma V_equalityI [intro]: "(\x. x \ elts a \ x \ elts b) \ a = b" by (meson dual_order.antisym less_eq_V_def subsetI) lemma vsubsetI [intro!]: "(\x. x \ elts a \ x \ elts b) \ a \ b" by (simp add: less_eq_V_def subsetI) lemma vsubsetD [elim, intro?]: "a \ b \ c \ elts a \ c \ elts b" using less_eq_V_def by auto lemma rev_vsubsetD: "c \ elts a \ a \ b \ c \ elts b" \ \The same, with reversed premises for use with @{method erule} -- cf. @{thm rev_mp}.\ by (rule vsubsetD) lemma vsubsetCE [elim,no_atp]: "a \ b \ (c \ elts a \ P) \ (c \ elts b \ P) \ P" \ \Classical elimination rule.\ using vsubsetD by blast lemma set_image_le_iff: "small A \ set (f ` A) \ B \ (\x\A. f x \ elts B)" by auto lemma eq0_iff: "x = 0 \ (\y. y \ elts x)" by auto lemma less_eq_V_0_iff [simp]: "x \ 0 \ x = 0" for x::V by auto lemma subset_iff_less_eq_V: assumes "small B" shows "A \ B \ set A \ set B \ small A" using assms down small_iff by auto lemma small_Collect [simp]: "small A \ small {x \ A. P x}" by (simp add: smaller_than_small) lemma small_Union_iff: "small (\(elts ` X)) \ small X" proof show "small X" if "small (\ (elts ` X))" proof - have "X \ set ` Pow (\ (elts ` X))" by fastforce then show ?thesis using Pow subset_iff_less_eq_V that by auto qed qed auto lemma not_less_0 [iff]: fixes x::V shows "\ x < 0" by (simp add: less_eq_V_def less_le_not_le) lemma le_0 [iff]: fixes x::V shows "0 \ x" by auto lemma min_0L [simp]: "min 0 n = 0" for n :: V by (simp add: min_absorb1) lemma min_0R [simp]: "min n 0 = 0" for n :: V by (simp add: min_absorb2) lemma neq0_conv: "\n::V. n \ 0 \ 0 < n" by (simp add: less_V_def) definition VPow :: "V \ V" where "VPow x \ set (set ` Pow (elts x))" lemma VPow_iff [iff]: "y \ elts (VPow x) \ y \ x" using down Pow apply (auto simp: VPow_def less_eq_V_def) using less_eq_V_def apply fastforce done lemma VPow_le_VPow_iff [simp]: "VPow a \ VPow b \ a \ b" by auto lemma elts_VPow: "elts (VPow x) = set ` Pow (elts x)" by (auto simp: VPow_def Pow) lemma small_sup_iff [simp]: "small (X \ Y) \ small X \ small Y" for X::"V set" by (metis down elts_of_set small_Un_elts sup_ge1 sup_ge2) lemma elts_sup_iff [simp]: "elts (x \ y) = elts x \ elts y" by (simp add: sup_V_def) lemma trad_foundation: assumes z: "z \ 0" shows "\w. w \ elts z \ w \ z = 0" using foundation assms by (simp add: wf_eq_minimal) (metis Int_emptyI equals0I inf_V_def set_of_elts zero_V_def) instantiation "V" :: Sup begin definition Sup_V where "Sup_V X \ if small X then set (Union (elts ` X)) else 0" instance .. end instantiation "V" :: Inf begin definition Inf_V where "Inf_V X \ if X = {} then 0 else set (Inter (elts ` X))" instance .. end lemma V_disjoint_iff: "x \ y = 0 \ elts x \ elts y = {}" by (metis down elts_of_set inf_V_def inf_le1 zero_V_def) text\I've no idea why @{term bdd_above} is treated differently from @{term bdd_below}, but anyway\ lemma bdd_above_iff_small [simp]: "bdd_above X = small X" for X::"V set" proof show "small X" if "bdd_above X" proof - obtain a where "\x\X. x \ a" using that \bdd_above X\ bdd_above_def by blast then show "small X" by (meson VPow_iff \\x\X. x \ a\ down subsetI) qed show "bdd_above X" if "small X" proof - have "\x\X. x \ \ X" by (simp add: SUP_upper Sup_V_def Union less_eq_V_def that) then show ?thesis by (meson bdd_above_def) qed qed instantiation "V" :: conditionally_complete_lattice begin definition bdd_below_V where "bdd_below_V X \ X \ {}" instance proof show "\ X \ x" if "x \ X" "bdd_below X" for x :: V and X :: "V set" using that by (auto simp: bdd_below_V_def Inf_V_def split: if_split_asm) show "z \ \ X" if "X \ {}" "\x. x \ X \ z \ x" for X :: "V set" and z :: V using that apply (clarsimp simp add: bdd_below_V_def Inf_V_def less_eq_V_def split: if_split_asm) by (meson INT_subset_iff down eq_refl equals0I) show "x \ \ X" if "x \ X" and "bdd_above X" for x :: V and X :: "V set" using that Sup_V_def by auto show "\ X \ (z::V)" if "X \ {}" "\x. x \ X \ x \ z" for X :: "V set" and z :: V using that by (simp add: SUP_least Sup_V_def less_eq_V_def) qed end lemma Sup_upper: "\x \ A; small A\ \ x \ \A" for A::"V set" by (auto simp: Sup_V_def SUP_upper Union less_eq_V_def) lemma Sup_least: fixes z::V shows "(\x. x \ A \ x \ z) \ \A \ z" by (auto simp: Sup_V_def SUP_least less_eq_V_def) lemma Sup_empty [simp]: "\{} = (0::V)" using Sup_V_def by auto lemma elts_Sup [simp]: "small X \ elts (\ X) = \(elts ` X)" by (auto simp: Sup_V_def) lemma sup_V_0_left [simp]: "0 \ a = a" and sup_V_0_right [simp]: "a \ 0 = a" for a::V by auto lemma Sup_V_insert: fixes x::V assumes "small A" shows "\(insert x A) = x \ \A" by (simp add: assms cSup_insert_If) lemma Sup_Un_distrib: "\small A; small B\ \ \(A \ B) = \A \ \B" for A::"V set" by auto lemma SUP_sup_distrib: fixes f :: "V \ V" shows "small A \ (\x\A. f x \ g x) = \ (f ` A) \ \ (g ` A)" by (force simp:) lemma SUP_const [simp]: "(\y \ A. a) = (if A = {} then (0::V) else a)" by simp lemma cSUP_subset_mono: fixes f :: "'a \ V set" and g :: "'a \ V set" shows "\A \ B; \x. x \ A \ f x \ g x\ \ \ (f ` A) \ \ (g ` B)" by (simp add: SUP_subset_mono) lemma mem_Sup_iff [iff]: "x \ elts (\X) \ x \ \ (elts ` X) \ small X" using Sup_V_def by auto lemma cSUP_UNION: fixes B :: "V \ V set" and f :: "V \ V" assumes ne: "small A" and bdd_UN: "small (\x\A. f ` B x)" shows "\(f ` (\x\A. B x)) = \((\x. \(f ` B x)) ` A)" proof - have bdd: "\x. x \ A \ small (f ` B x)" using bdd_UN subset_iff_less_eq_V by (meson SUP_upper smaller_than_small) then have bdd2: "small ((\x. \(f ` B x)) ` A)" using ne(1) by blast have "\(f ` (\x\A. B x)) \ \((\x. \(f ` B x)) ` A)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper2 simp: bdd2 bdd) moreover have "\((\x. \(f ` B x)) ` A) \ \(f ` (\x\A. B x))" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper simp: image_UN bdd_UN) ultimately show ?thesis by (rule order_antisym) qed lemma Sup_subset_mono: "small B \ A \ B \ Sup A \ Sup B" for A::"V set" by auto lemma Sup_le_iff: "small A \ Sup A \ a \ (\x\A. x \ a)" for A::"V set" by auto lemma SUP_le_iff: "small (f ` A) \ \(f ` A) \ u \ (\x\A. f x \ u)" for f :: "V \ V" by blast lemma Sup_eq_0_iff [simp]: "\A = 0 \ A \ {0} \ \ small A" for A :: "V set" using Sup_upper by fastforce lemma Sup_Union_commute: fixes f :: "V \ V set" assumes "small A" "\x. x\A \ small (f x)" shows "\ (\x\A. f x) = (\x\A. \ (f x))" using assms by (force simp: subset_iff_less_eq_V intro!: antisym) lemma Sup_eq_Sup: fixes B :: "V set" assumes "B \ A" "small A" and *: "\x. x \ A \ \y \ B. x \ y" shows "Sup A = Sup B" proof - have "small B" using assms subset_iff_less_eq_V by auto moreover have "\y\B. u \ elts y" if "x \ A" "u \ elts x" for u x using that "*" by blast moreover have "\x\A. v \ elts x" if "y \ B" "v \ elts y" for v y using that \B \ A\ by blast ultimately show ?thesis using assms by auto qed subsection\Successor function\ lemma vinsert_not_empty [simp]: "vinsert a A \ 0" and empty_not_vinsert [simp]: "0 \ vinsert a A" by (auto simp: vinsert_def) lemma succ_not_0 [simp]: "succ n \ 0" and zero_not_succ [simp]: "0 \ succ n" by (auto simp: succ_def) instantiation V :: zero_neq_one begin instance by intro_classes (metis elts_0 elts_succ empty_iff insert_iff one_V_def set_of_elts) end instantiation V :: zero_less_one begin instance by intro_classes (simp add: less_V_def) end lemma succ_ne_self [simp]: "i \ succ i" by (metis elts_succ insertI1 mem_not_refl) lemma succ_notin_self: "succ i \ elts i" using elts_succ mem_not_refl by blast lemma le_succE: "succ i \ succ j \ i \ j" using less_eq_V_def mem_not_sym by auto lemma succ_inject_iff [iff]: "succ i = succ j \ i = j" by (simp add: dual_order.antisym le_succE) lemma inj_succ: "inj succ" by (simp add: inj_def) lemma succ_neq_zero: "succ x \ 0" by (metis elts_0 elts_succ insert_not_empty) definition pred where "pred i \ THE j. i = succ j" lemma pred_succ [simp]: "pred (succ i) = i" by (simp add: pred_def) subsection \Ordinals\ definition Transset where "Transset x \ \y \ elts x. y \ x" definition Ord where "Ord x \ Transset x \ (\y \ elts x. Transset y)" abbreviation ON where "ON \ Collect Ord" subsubsection \Transitive sets\ lemma Transset_0 [iff]: "Transset 0" by (auto simp: Transset_def) lemma Transset_succ [intro]: assumes "Transset x" shows "Transset (succ x)" using assms by (auto simp: Transset_def succ_def less_eq_V_def) lemma Transset_Sup: assumes "\x. x \ X \ Transset x" shows "Transset (\X)" proof (cases "small X") case True with assms show ?thesis by (simp add: Transset_def) (meson Sup_upper assms dual_order.trans) qed (simp add: Sup_V_def) lemma Transset_sup: assumes "Transset x" "Transset y" shows "Transset (x \ y)" using Transset_def assms by fastforce lemma Transset_inf: "\Transset i; Transset j\ \ Transset (i \ j)" by (simp add: Transset_def rev_vsubsetD) lemma Transset_VPow: "Transset(i) \ Transset(VPow(i))" by (auto simp: Transset_def) lemma Transset_Inf: "(\i. i \ A \ Transset i) \ Transset (\ A)" by (force simp: Transset_def Inf_V_def) lemma Transset_SUP: "(\x. x \ A \ Transset (B x)) \ Transset (\ (B ` A))" by (metis Transset_Sup imageE) lemma Transset_INT: "(\x. x \ A \ Transset (B x)) \ Transset (\ (B ` A))" by (metis Transset_Inf imageE) subsubsection \Zero, successor, sups\ lemma Ord_0 [iff]: "Ord 0" by (auto simp: Ord_def) lemma Ord_succ [intro]: assumes "Ord x" shows "Ord (succ x)" using assms by (auto simp: Ord_def) lemma Ord_Sup: assumes "\x. x \ X \ Ord x" shows "Ord (\X)" proof (cases "small X") case True with assms show ?thesis by (auto simp: Ord_def Transset_Sup) qed (simp add: Sup_V_def) lemma Ord_Union: assumes "\x. x \ X \ Ord x" "small X" shows "Ord (set (\ (elts ` X)))" by (metis Ord_Sup Sup_V_def assms) lemma Ord_sup: assumes "Ord x" "Ord y" shows "Ord (x \ y)" using assms proof (clarsimp simp: Ord_def) show "Transset (x \ y) \ (\y\elts x \ elts y. Transset y)" if "Transset x" "Transset y" "\y\elts x. Transset y" "\y\elts y. Transset y" using Ord_def Transset_sup assms by auto qed lemma big_ON [simp]: "\ small ON" proof assume "small ON" then have "set ON \ ON" by (metis Ord_Union Ord_succ Sup_upper elts_Sup elts_succ insertI1 mem_Collect_eq mem_not_refl set_of_elts vsubsetD) then show False by (metis \small ON\ elts_of_set mem_not_refl) qed lemma Ord_1 [iff]: "Ord 1" using Ord_succ one_V_def succ_def vinsert_def by fastforce lemma OrdmemD: "Ord k \ j \ elts k \ j < k" using Ord_def Transset_def less_V_def by auto lemma Ord_trans: "\ i \ elts j; j \ elts k; Ord k \ \ i \ elts k" using Ord_def Transset_def by blast lemma mem_0_Ord: assumes k: "Ord k" and knz: "k \ 0" shows "0 \ elts k" by (metis Ord_def Transset_def inf.orderE k knz trad_foundation) lemma Ord_in_Ord: "\ Ord k; m \ elts k \ \ Ord m" using Ord_def Ord_trans by blast lemma OrdI: "\Transset i; \x. x \ elts i \ Transset x\ \ Ord i" by (simp add: Ord_def) lemma Ord_is_Transset: "Ord i \ Transset i" by (simp add: Ord_def) lemma Ord_contains_Transset: "\Ord i; j \ elts i\ \ Transset j" using Ord_def by blast lemma ON_imp_Ord: assumes "H \ ON" "x \ H" shows "Ord x" using assms by blast lemma elts_subset_ON: "Ord \ \ elts \ \ ON" using Ord_in_Ord by blast lemma Transset_pred [simp]: "Transset x \ \(elts (succ x)) = x" by (fastforce simp: Transset_def) lemma Ord_pred [simp]: "Ord \ \ \ (insert \ (elts \)) = \" using Ord_def Transset_pred by auto subsubsection \Induction, Linearity, etc.\ lemma Ord_induct [consumes 1, case_names step]: assumes k: "Ord k" and step: "\x.\ Ord x; \y. y \ elts x \ P y \ \ P x" shows "P k" using foundation k proof (induction k rule: wf_induct_rule) case (less x) then show ?case using Ord_in_Ord local.step by auto qed text \Comparability of ordinals\ lemma Ord_linear: "Ord k \ Ord l \ k \ elts l \ k=l \ l \ elts k" proof (induct k arbitrary: l rule: Ord_induct) case (step k) note step_k = step show ?case using \Ord l\ proof (induct l rule: Ord_induct) case (step l) thus ?case using step_k by (metis Ord_trans V_equalityI) qed qed text \The trichotomy law for ordinals\ lemma Ord_linear_lt: assumes "Ord k" "Ord l" obtains (lt) "k < l" | (eq) "k=l" | (gt) "l < k" using Ord_linear OrdmemD assms by blast lemma Ord_linear2: assumes "Ord k" "Ord l" obtains (lt) "k < l" | (ge) "l \ k" by (metis Ord_linear_lt eq_refl assms order.strict_implies_order) lemma Ord_linear_le: assumes "Ord k" "Ord l" obtains (le) "k \ l" | (ge) "l \ k" by (meson Ord_linear2 le_less assms) lemma union_less_iff [simp]: "\Ord i; Ord j\ \ i \ j < k \ i j Ord l \ k \ elts l \ k < l" by (metis Ord_linear OrdmemD less_le_not_le) lemma Ord_Collect_lt: "Ord \ \ {\. Ord \ \ \ < \} = elts \" by (auto simp flip: Ord_mem_iff_lt elim: Ord_in_Ord OrdmemD) lemma Ord_not_less: "\Ord x; Ord y\ \ \ x < y \ y \ x" by (metis (no_types) Ord_linear2 leD) lemma Ord_not_le: "\Ord x; Ord y\ \ \ x \ y \ y < x" by (metis (no_types) Ord_linear2 leD) lemma le_succ_iff: "Ord i \ Ord j \ succ i \ succ j \ i \ j" by (metis Ord_linear_le Ord_succ le_succE order_antisym) lemma succ_le_iff: "Ord i \ Ord j \ succ i \ j \ i < j" using Ord_mem_iff_lt dual_order.strict_implies_order less_eq_V_def by fastforce lemma succ_in_Sup_Ord: assumes eq: "succ \ = \A" and "small A" "A \ ON" "Ord \" shows "succ \ \ A" proof - have "\ \A \ \" using eq \Ord \\ succ_le_iff by fastforce then show ?thesis using assms by (metis Ord_linear2 Sup_least Sup_upper eq_iff mem_Collect_eq subsetD succ_le_iff) qed lemma in_succ_iff: "Ord i \ j \ elts (ZFC_in_HOL.succ i) \ Ord j \ j \ i" by (metis Ord_in_Ord Ord_mem_iff_lt Ord_not_le Ord_succ succ_le_iff) lemma zero_in_succ [simp,intro]: "Ord i \ 0 \ elts (succ i)" using mem_0_Ord by auto lemma less_succ_self: "x < succ x" by (simp add: less_eq_V_def order_neq_le_trans subset_insertI) lemma Ord_finite_Sup: "\finite A; A \ ON; A \ {}\ \ \A \ A" proof (induction A rule: finite_induct) case (insert x A) then have *: "small A" "A \ ON" "Ord x" by (auto simp add: ZFC_in_HOL.finite_imp_small insert.hyps) show ?case proof (cases "A = {}") case False then have "\A \ A" using insert by blast then have "\A \ x" if "x \ \A \ A" using * by (metis ON_imp_Ord Ord_linear_le sup.absorb2 that) then show ?thesis by (fastforce simp: \small A\ Sup_V_insert) qed auto qed auto subsubsection \The natural numbers\ primrec ord_of_nat :: "nat \ V" where "ord_of_nat 0 = 0" | "ord_of_nat (Suc n) = succ (ord_of_nat n)" lemma ord_of_nat_eq_initial: "ord_of_nat n = set (ord_of_nat ` {.. elts (ord_of_nat n) \ (\m i = \ (succ ` elts i)" by (force intro: Ord_trans) lemma Ord_ord_of_nat [simp]: "Ord (ord_of_nat k)" by (induct k, auto) lemma ord_of_nat_equality: "ord_of_nat n = \ ((succ \ ord_of_nat) ` {.. :: V where "\ \ set (range ord_of_nat)" lemma elts_\: "elts \ = {\. \n. \ = ord_of_nat n}" by (auto simp: \_def image_iff) lemma nat_into_Ord [simp]: "n \ elts \ \ Ord n" by (metis Ord_ord_of_nat \_def elts_of_set image_iff inf) lemma Sup_\: "\(elts \) = \" unfolding \_def by force lemma Ord_\ [iff]: "Ord \" by (metis Ord_Sup Sup_\ nat_into_Ord) lemma zero_in_omega [iff]: "0 \ elts \" by (metis \_def elts_of_set inf ord_of_nat.simps(1) rangeI) lemma succ_in_omega [simp]: "n \ elts \ \ succ n \ elts \" by (metis \_def elts_of_set image_iff small_image_nat_V ord_of_nat.simps(2) rangeI) lemma ord_of_eq_0: "ord_of_nat j = 0 \ j = 0" by (induct j) (auto simp: succ_neq_zero) lemma ord_of_nat_le_omega: "ord_of_nat n \ \" by (metis Sup_\ ZFC_in_HOL.Sup_upper \_def elts_of_set inf rangeI) lemma ord_of_eq_0_iff [simp]: "ord_of_nat n = 0 \ n=0" by (auto simp: ord_of_eq_0) lemma ord_of_nat_inject [iff]: "ord_of_nat i = ord_of_nat j \ i=j" proof (induct i arbitrary: j) case 0 show ?case using ord_of_eq_0 by auto next case (Suc i) then show ?case by auto (metis elts_0 elts_succ insert_not_empty not0_implies_Suc ord_of_nat.simps succ_inject_iff) qed corollary inj_ord_of_nat: "inj ord_of_nat" by (simp add: linorder_injI) corollary countable: assumes "countable X" shows "small X" proof - have "X \ range (from_nat_into X)" by (simp add: assms subset_range_from_nat_into) then show ?thesis by (meson inf_raw inj_ord_of_nat replacement small_def smaller_than_small) qed corollary infinite_\: "infinite (elts \)" using range_inj_infinite [of ord_of_nat] by (simp add: \_def inj_ord_of_nat) corollary ord_of_nat_mono_iff [iff]: "ord_of_nat i \ ord_of_nat j \ i \ j" by (metis Ord_def Ord_ord_of_nat Transset_def eq_iff mem_ord_of_nat_iff not_less ord_of_nat_inject) corollary ord_of_nat_strict_mono_iff [iff]: "ord_of_nat i < ord_of_nat j \ i < j" by (simp add: less_le_not_le) lemma small_image_nat [simp]: fixes N :: "nat set" shows "small (g ` N)" by (simp add: countable) lemma finite_Ord_omega: "\ \ elts \ \ finite (elts \)" proof (clarsimp simp add: \_def) show "finite (elts (ord_of_nat n))" if "\ = ord_of_nat n" for n using that by (simp add: ord_of_nat_eq_initial [of n]) qed lemma infinite_Ord_omega: "Ord \ \ infinite (elts \) \ \ \ \" by (meson Ord_\ Ord_linear2 Ord_mem_iff_lt finite_Ord_omega) lemma ord_of_minus_1: "n > 0 \ ord_of_nat n = succ (ord_of_nat (n - 1))" by (metis Suc_diff_1 ord_of_nat.simps(2)) lemma card_ord_of_nat [simp]: "card (elts (ord_of_nat m)) = m" by (induction m) (auto simp: \_def finite_Ord_omega) lemma ord_of_nat_\ [iff]:"ord_of_nat n \ elts \" by (simp add: \_def) lemma succ_\_iff [iff]: "succ n \ elts \ \ n \ elts \" by (metis Ord_\ OrdmemD elts_vinsert insert_iff less_V_def succ_def succ_in_omega vsubsetD) lemma \_gt0 [simp]: "\ > 0" by (simp add: OrdmemD) lemma \_gt1 [simp]: "\ > 1" by (simp add: OrdmemD one_V_def) subsubsection\Limit ordinals\ definition Limit :: "V\bool" where "Limit i \ Ord i \ 0 \ elts i \ (\y. y \ elts i \ succ y \ elts i)" lemma zero_not_Limit [iff]: "\ Limit 0" by (simp add: Limit_def) lemma not_succ_Limit [simp]: "\ Limit(succ i)" by (metis Limit_def Ord_mem_iff_lt elts_succ insertI1 less_irrefl) lemma Limit_is_Ord: "Limit \ \ Ord \" by (simp add: Limit_def) lemma succ_in_Limit_iff: "Limit \ \ succ \ \ elts \ \ \ \ elts \" by (metis Limit_def OrdmemD elts_succ insertI1 less_V_def vsubsetD) lemma Limit_eq_Sup_self [simp]: "Limit i \ Sup (elts i) = i" apply (rule order_antisym) apply (simp add: Limit_def Ord_def Transset_def Sup_least) by (metis Limit_def Ord_equality Sup_V_def SUP_le_iff Sup_upper small_elts) lemma zero_less_Limit: "Limit \ \ 0 < \" by (simp add: Limit_def OrdmemD) lemma non_Limit_ord_of_nat [iff]: "\ Limit (ord_of_nat m)" by (metis Limit_def mem_ord_of_nat_iff not_succ_Limit ord_of_eq_0_iff ord_of_minus_1) lemma Limit_omega [iff]: "Limit \" by (simp add: Limit_def) lemma omega_nonzero [simp]: "\ \ 0" using Limit_omega by fastforce lemma Ord_cases_lemma: assumes "Ord k" shows "k = 0 \ (\j. k = succ j) \ Limit k" proof (cases "Limit k") case False have "succ j \ elts k" if "\j. k \ succ j" "j \ elts k" for j by (metis Ord_in_Ord Ord_linear Ord_succ assms elts_succ insertE mem_not_sym that) with assms show ?thesis by (auto simp: Limit_def mem_0_Ord) qed auto lemma Ord_cases [cases type: V, case_names 0 succ limit]: assumes "Ord k" obtains "k = 0" | l where "Ord l" "succ l = k" | "Limit k" by (metis assms Ord_cases_lemma Ord_in_Ord elts_succ insertI1) lemma non_succ_LimitI: assumes "i\0" "Ord(i)" "\y. succ(y) \ i" shows "Limit(i)" using Ord_cases_lemma assms by blast lemma Ord_induct3 [consumes 1, case_names 0 succ Limit, induct type: V]: assumes \: "Ord \" and P: "P 0" "\\. \Ord \; P \\ \ P (succ \)" "\\. \Limit \; \\. \ \ elts \ \ P \\ \ P (\\ \ elts \. \)" shows "P \" using \ proof (induction \ rule: Ord_induct) case (step \) then show ?case by (metis Limit_eq_Sup_self Ord_cases P elts_succ image_ident insertI1) qed subsubsection\Properties of LEAST for ordinals\ lemma assumes "Ord k" "P k" shows Ord_LeastI: "P (LEAST i. Ord i \ P i)" and Ord_Least_le: "(LEAST i. Ord i \ P i) \ k" proof - have "P (LEAST i. Ord i \ P i) \ (LEAST i. Ord i \ P i) \ k" using assms proof (induct k rule: Ord_induct) case (step x) then have "P x" by simp show ?case proof (rule classical) assume assm: "\ (P (LEAST a. Ord a \ P a) \ (LEAST a. Ord a \ P a) \ x)" have "\y. Ord y \ P y \ x \ y" proof (rule classical) fix y assume y: "Ord y \ P y" "\ x \ y" with step obtain "P (LEAST a. Ord a \ P a)" and le: "(LEAST a. Ord a \ P a) \ y" by (meson Ord_linear2 Ord_mem_iff_lt) with assm have "x < (LEAST a. Ord a \ P a)" by (meson Ord_linear_le y order.trans \Ord x\) then show "x \ y" using le by auto qed then have Least: "(LEAST a. Ord a \ P a) = x" by (simp add: Least_equality \Ord x\ step.prems) with \P x\ show ?thesis by simp qed qed then show "P (LEAST i. Ord i \ P i)" and "(LEAST i. Ord i \ P i) \ k" by auto qed lemma Ord_Least: assumes "Ord k" "P k" shows "Ord (LEAST i. Ord i \ P i)" proof - have "Ord (LEAST i. Ord i \ (Ord i \ P i))" using Ord_LeastI [where P = "\i. Ord i \ P i"] assms by blast then show ?thesis by simp qed \ \The following 3 lemmas are due to Brian Huffman\ lemma Ord_LeastI_ex: "\i. Ord i \ P i \ P (LEAST i. Ord i \ P i)" using Ord_LeastI by blast lemma Ord_LeastI2: "\Ord a; P a; \x. \Ord x; P x\ \ Q x\ \ Q (LEAST i. Ord i \ P i)" by (blast intro: Ord_LeastI Ord_Least) lemma Ord_LeastI2_ex: "\a. Ord a \ P a \ (\x. \Ord x; P x\ \ Q x) \ Q (LEAST i. Ord i \ P i)" by (blast intro: Ord_LeastI_ex Ord_Least) lemma Ord_LeastI2_wellorder: assumes "Ord a" "P a" and "\a. \ P a; \b. Ord b \ P b \ a \ b \ \ Q a" shows "Q (LEAST i. Ord i \ P i)" proof (rule LeastI2_order) show "Ord (LEAST i. Ord i \ P i) \ P (LEAST i. Ord i \ P i)" using Ord_Least Ord_LeastI assms by auto next fix y assume "Ord y \ P y" thus "(LEAST i. Ord i \ P i) \ y" by (simp add: Ord_Least_le) next fix x assume "Ord x \ P x" "\y. Ord y \ P y \ x \ y" thus "Q x" by (simp add: assms(3)) qed lemma Ord_LeastI2_wellorder_ex: assumes "\x. Ord x \ P x" and "\a. \ P a; \b. Ord b \ P b \ a \ b \ \ Q a" shows "Q (LEAST i. Ord i \ P i)" using assms by clarify (blast intro!: Ord_LeastI2_wellorder) lemma not_less_Ord_Least: "\k < (LEAST x. Ord x \ P x); Ord k\ \ \ P k" using Ord_Least_le less_le_not_le by auto lemma exists_Ord_Least_iff: "(\\. Ord \ \ P \) \ (\\. Ord \ \ P \ \ (\\ < \. Ord \ \ \ P \))" (is "?lhs \ ?rhs") proof assume ?rhs thus ?lhs by blast next assume H: ?lhs then obtain \ where \: "Ord \" "P \" by blast let ?x = "LEAST \. Ord \ \ P \" have "Ord ?x" by (metis Ord_Least \) moreover { fix \ assume m: "\ < ?x" "Ord \" from not_less_Ord_Least[OF m] have "\ P \" . } ultimately show ?rhs using Ord_LeastI_ex[OF H] by blast qed lemma Ord_mono_imp_increasing: assumes fun_hD: "h \ D \ D" - and mono_h: "strict_mono_on h D" + and mono_h: "strict_mono_on D h" and "D \ ON" and \: "\ \ D" shows "\ \ h \" proof (rule ccontr) assume non: "\ \ \ h \" define \ where "\ \ LEAST \. Ord \ \ \ \ \ h \ \ \ \ D" have "Ord \" using \ \D \ ON\ by blast then have \: "\ \ \ h \ \ \ \ D" unfolding \_def by (rule Ord_LeastI) (simp add: \ non) have "Ord (h \)" using assms by auto then have "Ord (h (h \))" by (meson ON_imp_Ord \ assms funcset_mem) have "Ord \" using \ \D \ ON\ by blast then have "h \ < \" by (metis ON_imp_Ord Ord_linear2 PiE \ \D \ ON\ fun_hD) then have "\ h \ \ h (h \)" using \ fun_hD mono_h by (force simp: strict_mono_on_def) moreover have *: "h \ \ D" using \ fun_hD by auto moreover have "Ord (h \)" using \D \ ON\ * by blast ultimately have "\ \ h \" by (simp add: \_def Ord_Least_le) then show False using \ by blast qed lemma le_Sup_iff: assumes "A \ ON" "Ord x" "small A" shows "x \ \A \ (\y \ ON. y (\a\A. y < a))" proof (intro iffI ballI impI) show "\a\A. y < a" if "x \ \ A" "y \ ON" "y < x" for y proof - have "\ \ A \ y" "Ord y" using that by auto then show ?thesis by (metis Ord_linear2 Sup_least \A \ ON\ mem_Collect_eq subset_eq) qed show "x \ \ A" if "\y\ON. y < x \ (\a\A. y < a)" using that assms by (metis Ord_Sup Ord_linear_le Sup_upper less_le_not_le mem_Collect_eq subsetD) qed lemma le_SUP_iff: "\f ` A \ ON; Ord x; small A\ \ x \ \(f ` A) \ (\y \ ON. y (\i\A. y < f i))" by (simp add: le_Sup_iff) subsection\Transfinite Recursion and the V-levels\ definition transrec :: "((V \ 'a) \ V \ 'a) \ V \ 'a" where "transrec H a \ wfrec {(x,y). x \ elts y} H a" lemma transrec: "transrec H a = H (\x \ elts a. transrec H x) a" proof - have "(cut (wfrec {(x, y). x \ elts y} H) {(x, y). x \ elts y} a) = (\x\elts a. wfrec {(x, y). x \ elts y} H x)" by (force simp: cut_def) then show ?thesis unfolding transrec_def by (simp add: foundation wfrec) qed text\Avoids explosions in proofs; resolve it with a meta-level definition\ lemma def_transrec: "\\x. f x \ transrec H x\ \ f a = H(\x \ elts a. f x) a" by (metis restrict_ext transrec) lemma eps_induct [case_names step]: assumes "\x. (\y. y \ elts x \ P y) \ P x" shows "P a" using wf_induct [OF foundation] assms by auto definition Vfrom :: "[V,V] \ V" where "Vfrom a \ transrec (\f x. a \ \((\y. VPow(f y)) ` elts x))" abbreviation Vset :: "V \ V" where "Vset \ Vfrom 0" lemma Vfrom: "Vfrom a i = a \ \((\j. VPow(Vfrom a j)) ` elts i)" apply (subst Vfrom_def) apply (subst transrec) using Vfrom_def by auto lemma Vfrom_0 [simp]: "Vfrom a 0 = a" by (subst Vfrom) auto lemma Vset: "Vset i = \((\j. VPow(Vset j)) ` elts i)" by (subst Vfrom) auto lemma Vfrom_mono1: assumes "a \ b" shows "Vfrom a i \ Vfrom b i" proof (induction i rule: eps_induct) case (step i) then have "a \ (\j\elts i. VPow (Vfrom a j)) \ b \ (\j\elts i. VPow (Vfrom b j))" by (intro sup_mono cSUP_subset_mono \a \ b\) auto then show ?case by (metis Vfrom) qed lemma Vfrom_mono2: "Vfrom a i \ Vfrom a (i \ j)" proof (induction arbitrary: j rule: eps_induct) case (step i) then have "a \ (\j\elts i. VPow (Vfrom a j)) \ a \ (\j\elts (i \ j). VPow (Vfrom a j))" by (intro sup_mono cSUP_subset_mono order_refl) auto then show ?case by (metis Vfrom) qed lemma Vfrom_mono: "\Ord i; a\b; i\j\ \ Vfrom a i \ Vfrom b j" by (metis (no_types) Vfrom_mono1 Vfrom_mono2 dual_order.trans sup.absorb_iff2) lemma Transset_Vfrom: "Transset(A) \ Transset(Vfrom A i)" proof (induction i rule: eps_induct) case (step i) then show ?case by (metis Transset_SUP Transset_VPow Transset_sup Vfrom) qed lemma Transset_Vset [simp]: "Transset(Vset i)" by (simp add: Transset_Vfrom) lemma Vfrom_sup: "Vfrom a (i \ j) = Vfrom a i \ Vfrom a j" proof (rule order_antisym) show "Vfrom a (i \ j) \ Vfrom a i \ Vfrom a j" by (simp add: Vfrom [of a "i \ j"] Vfrom [of a i] Vfrom [of a j] Sup_Un_distrib image_Un sup.assoc sup.left_commute) show "Vfrom a i \ Vfrom a j \ Vfrom a (i \ j)" by (metis Vfrom_mono2 le_supI sup_commute) qed lemma Vfrom_succ_Ord: assumes "Ord i" shows "Vfrom a (succ i) = a \ VPow(Vfrom a i)" proof (cases "i = 0") case True then show ?thesis by (simp add: Vfrom [of _ "succ 0"]) next case False have *: "(\x\elts i. VPow (Vfrom a x)) \ VPow (Vfrom a i)" proof (rule cSup_least) show "(\x. VPow (Vfrom a x)) ` elts i \ {}" using False by auto show "x \ VPow (Vfrom a i)" if "x \ (\x. VPow (Vfrom a x)) ` elts i" for x using that by clarsimp (meson Ord_in_Ord Ord_linear_le Vfrom_mono assms mem_not_refl order_refl vsubsetD) qed show ?thesis proof (rule Vfrom [THEN trans]) show "a \ (\j\elts (succ i). VPow (Vfrom a j)) = a \ VPow (Vfrom a i)" using assms by (intro sup_mono order_antisym) (auto simp: Sup_V_insert *) qed qed lemma Vset_succ: "Ord i \ Vset(succ(i)) = VPow(Vset(i))" by (simp add: Vfrom_succ_Ord) lemma Vfrom_Sup: assumes "X \ {}" "small X" shows "Vfrom a (Sup X) = (\y\X. Vfrom a y)" proof (rule order_antisym) have "Vfrom a (\ X) = a \ (\j\elts (\ X). VPow (Vfrom a j))" by (metis Vfrom) also have "\ \ \ (Vfrom a ` X)" proof - have "a \ \ (Vfrom a ` X)" by (metis Vfrom all_not_in_conv assms bdd_above_iff_small cSUP_upper2 replacement sup_ge1) moreover have "(\j\elts (\ X). VPow (Vfrom a j)) \ \ (Vfrom a ` X)" proof - have "VPow (Vfrom a x) \ \ (Vfrom a ` X)" if "y \ X" "x \ elts y" for x y proof - have "VPow (Vfrom a x) \ Vfrom a y" by (metis Vfrom bdd_above_iff_small cSUP_upper2 le_supI2 order_refl replacement small_elts that(2)) also have "\ \ \ (Vfrom a ` X)" using assms that by (force intro: cSUP_upper) finally show ?thesis . qed then show ?thesis by (simp add: SUP_le_iff \small X\) qed ultimately show ?thesis by auto qed finally show "Vfrom a (\ X) \ \ (Vfrom a ` X)" . have "\x. x \ X \ a \ (\j\elts x. VPow (Vfrom a j)) \ a \ (\j\elts (\ X). VPow (Vfrom a j))" using cSUP_subset_mono \small X\ by auto then show "\ (Vfrom a ` X) \ Vfrom a (\ X)" by (metis Vfrom assms(1) cSUP_least) qed lemma Limit_Vfrom_eq: "Limit(i) \ Vfrom a i = (\y \ elts i. Vfrom a y)" by (metis Limit_def Limit_eq_Sup_self Vfrom_Sup ex_in_conv small_elts) end