diff --git a/thys/Ordinal_Partitions/Erdos_Milner.thy b/thys/Ordinal_Partitions/Erdos_Milner.thy new file mode 100644 --- /dev/null +++ b/thys/Ordinal_Partitions/Erdos_Milner.thy @@ -0,0 +1,1349 @@ +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_raw ordermap_bij small_iff_range 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)\ \ \ \ < \ \ \ \ < \" + using ordermap_mono_less [of _ _ VWF D] bij_betw_apply [OF bij_\] VWF_iff_Ord_less \_cancel_right trans_VWF wf_VWF + by (metis ON_imp_Ord Ord_linear2 less_V_def order.asym) + 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 [simp]: "length L = length \s" + by (simp add: L_def) + have in_elts_E_less: "x' < x" + if "x' \ elts (E k')" "x \ elts (E k)" "k's" for k k' x' x + \ \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 by (auto simp: E_def lift_def add.assoc 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 + using ordertype_eq_iff \_iff bij_\ that + by (meson Ord_\ Ord_oexp ord(1) ordertype_VWF_eq_iff replacement small_elts) + 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 (metis (no_types) \M \ D\ \small D\ bij_betw_inv_into inf_le1 ordermap_bij subset_iff_less_eq_V 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) + then have \_less: "\X \ \. \\ < \; \ \ elts (tp (M \ X)); \ \ elts (tp (M \ X))\ + \ \ X \ < \ X \" + using ordermap_mono_less bij_betw_apply [OF bij_\] VWF_iff_Ord_less \_cancel_right trans_VWF wf_VWF + by (metis Ord_\ Ord_linear_lt less_imp_not_less ordermap_mono_less) + 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 + define \' where "\' \ ordermap D VWF d'" + define \ where "\ \ 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: \'_def \_def \small D\ ordermap_in_ordertype) + have **: "\X w. \X \ list.set L; w \ elts (tp X)\ \ w \ elts (tp (M \ X))" + using "*" by blast + have Ord_\L: "Ord (\ (L!k) (\ k d))" if "d \ \ ` elts (E k)" "k < length \s" for k d + by (metis "**" Ord_\ \length L = length \s\ bij_\ bij_betw_imp_surj_on imageI nth_mem that tp_L_eq) + have "\' < \" + by (metis \'_def \_def \d' < d\ \small D\ \_cancel_left \_less_iff ordermap_in_ordertype xy) + then have "K \' \ K \" + using elts' elts by (meson in_elts_E_less leI len' less_asym) + 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 + unfolding \'_def \_def + by (metis "**" \length L = length \s\ \_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 \'_def \_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 blast + 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 + show ?thesis + unfolding equal + proof (rule \_less) + show "\ (K \) d' < \ (K \) d" + by (metis equal \'_def \_def \(d', d) \ VWF\ \_cancel_left \_iff elts elts' imageI len xy) + have "\ (K \) d' \ elts (tp (L ! K \))" "\ (K \) d \ elts (tp (L ! K \))" + using equal \_cancel_left \'_def elts' len' \_def elts len xy + by (force intro: bij_betw_apply [OF bij_\])+ + then show "\ (K \) d' \ elts (tp (M \ L ! K \))" "\ (K \) d \ elts (tp (M \ L ! K \))" + by (simp_all add: "**" len) + qed + qed + moreover have "Ord (\ (L ! K \') (\ (K \') d'))" + using Ord_\L \'_def \_cancel_left elts' len' xy(1) by fastforce + moreover have "Ord (\ (L ! K \) (\ (K \) d))" + using Ord_\L \_def \_cancel_left elts len xy by fastforce + ultimately show "(\ d', \ d) \ VWF" + by (simp add: \_def \'_def \_def) + next + show "\ ` D \ M" + proof (clarsimp simp: \_def) + fix d + assume "d \ D" + define \ where "\ \ ordermap D VWF d" + have len: "K \ < length \s" and elts: "\ \ elts (E (K \))" + using K \d \ D\ by (auto simp: \_def \small D\ ordermap_in_ordertype) + have "\ (K \) d \ elts (tp (L! (K \)))" + using bij_\ [OF len] \d \ D\ + apply (simp add: L_def len) + by (metis \_def \_cancel_left bij_betw_imp_surj_on elts imageI) + then have "\ (L! (K \)) (\ (K \) d) \ M \ (L! (K \))" + using smM bij_betw_imp_surj_on [OF ordermap_bij] \length L = length \s\ + unfolding \_def + by (metis (no_types) "*" inv_into_into len nth_mem vsubsetD total_on_VWF wf_VWF) + then show "\ (L ! K (ordermap D VWF d)) (\ (K (ordermap D VWF d)) d) \ M" + using \_def by blast + 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 "\ > 1" "Ord \" and A: "tp A = \" "A \ elts (\*\)" + and "x \ A" and A1: "tp A1 = \" "A1 \ A" + obtains A2 where "tp A2 = \" "A2 \ A1" "less_sets {x} A2" +proof - + have "Ord \" + using indec indecomposable_imp_Ord by blast + have "Limit \" + by (simp add: assms indecomposable_imp_Limit) + have "small A" + by (meson A small_elts smaller_than_small) + 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 + have "A \ ON" + by (meson Ord_mult \Ord \\ \Ord \\ A dual_order.trans elts_subset_ON) + define B where "B \ \ ` (elts (succ \))" + show thesis + proof + have "small A1" + by (meson \small A\ A1 smaller_than_small) + then have "tp (A1 - B) \ tp A1" + unfolding B_def by (auto intro!: ordertype_VWF_mono del: vsubsetI) + 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 "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 + show "\ u < \ v" + if "u \ elts (succ \)" and "v \ elts (succ \)" and "u < v" for u v + proof - + have "succ \ \ elts \" + using \ Limit_def \Limit \\ by blast + then have "u \ elts \" "v \ elts \" + using Ord_trans \Ord \\ that by blast+ + then show ?thesis + using that \Limit \\ \small A\ A bij_betwE [OF bij_\] + by (metis ON_imp_Ord Ord_linear2 \A \ ON\ \_def inv_ordermap_VWF_mono_le leD) + qed + show "\ \ \ tp (elts (succ \))" + proof (subst ordertype_eq_Ord) + show "\ \ \ succ \" + by (meson \ \Limit \\ less_eq_V_def mem_not_refl subsetD succ_in_Limit_iff) + qed (use \Ord \\ in blast) + 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 + show "less_sets {x} (A1 - B)" + proof (clarsimp simp: less_sets_def B_def simp del: elts_succ) + fix y + assume "y \ A1" and y: "y \ \ ` elts (succ \)" + obtain "Ord x" "Ord y" + using \A \ ON\ \x \ A\ \y \ A1\ A1 by auto + have "y \ \ ` elts (succ \)" if "y \ elts (succ x)" + proof - + have "ordermap A VWF y \ elts (ZFC_in_HOL.succ (ordermap A VWF x))" + using A1 + by (metis insert_iff ordermap_mono subset_iff that wf_VWF OrdmemD VWF_iff_Ord_less \Ord x\ \Ord y\ \small A\ \y \ A1\ elts_succ) + then show ?thesis + using that A1 unfolding \_def + by (metis \y \ A1\ \_def bij_betw_inv_into_left bij_om imageI subsetD) + qed + then show "x < y" + by (meson Ord_linear2 Ord_mem_iff_lt Ord_succ \Ord x\ \Ord y\ y succ_le_iff) + qed + qed auto +qed + + +text \the main theorem, from which they derive the headline result\ +theorem Erdos_Milner_aux: + fixes k::nat and \::V + assumes part: "partn_lst_VWF \ [ord_of_nat 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 Ord_\1 Ord_in_Ord \ by blast + show ?thesis + proof (cases "\=0") + case True + moreover have "min \ 0 = 0" + by (simp add: min_def) + ultimately show ?thesis + by (simp add: partn_lst_triv0 [where i=1]) + next + case False + then have "\=1" + using True by blast + then obtain i where "i < Suc (Suc 0)" "[ord_of_nat k, \] ! i \ \" + using partn_lst_VWF_nontriv [OF part] 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) + moreover have "elts \ \ {}" + using False by auto + ultimately show ?thesis + by (auto simp: True \Ord \\ \\\0\ \\=1\ intro!: partn_lst_triv1 [where i=1]) + qed +next + case False + then have "\ \ 1" "\ \ 0" + by auto + then have "0 \ elts \" + using Ord_\1 Ord_in_Ord \ mem_0_Ord by blast + show ?thesis + proof (cases "\=0") + case True + have \: "[ord_of_nat (2 * k), 0] ! 1 = 0" + by simp + show ?thesis + using True assms + by (force simp: partn_lst_def nsets_empty_iff simp flip: numeral_2_eq_2 dest!: less_2_cases intro: \) + next + case False + then have "\ \ \" + using indec \\ \ 1\ + by (metis Ord_\ indecomposable_is_\_power le_oexp oexp_0_right) + then have "\ > 1" + using \_gt1 dual_order.strict_trans1 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 Aless: "\x y. \x \ A; y \ A; x < y\ \ (x,y) \ VWF" + by (meson Ord_in_Ord VWF_iff_Ord_less \Ord(\*\)\ subsetD that(1)) + then have om_A_less: "\x y. \x \ A; y \ A; x < y\ \ ordermap A VWF x < ordermap A VWF y" + by (auto simp: \small A\ ordermap_mono_less) + have \_sub: "elts \ \ ordermap A VWF ` A" + by (metis \small A\ elts_of_set less_eq_V_def ordertype_def ot replacement) + have g_less: "?g x < ?g y" if "x < y" "x \ elts \" "y \ elts \" for x y + proof - + have "?g x \ A" "?g y \ A" + using that by (meson \_sub inv_into_into subsetD)+ + moreover have "x \ ordermap A VWF ` A" "y \ ordermap A VWF ` A" + using \_sub that by blast+ + moreover have "A \ ON" + using A_\\ elts_subset_ON \Ord(\*\)\ by blast + ultimately show ?thesis + by (metis ON_imp_Ord Ord_linear_lt f_inv_into_f less_not_sym om_A_less \x < y\) + qed + have "?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 [simp]: "tp (?g ` H) = tp H" + proof (rule ordertype_VWF_inc_eq) + show "?g ` H \ ON" + using elts_subset_ON gH ord(3) by auto + show "?g x < inv_into A (ordermap A VWF) y" if "x \ H" "y \ H" "x < y" for x y + using that \H \ elts \\ g_less by blast + qed (use \H \ elts \\ elts_subset_ON ord down in auto) + show ?thesis + using ii [of "?g ` H"] ot_eq 1 + apply (auto simp: gH elim!: nset_image_obtains) + apply (meson \H \ elts \\ inj_g bij_betw_def inj_on_subset) + done + 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 [simp]: "tp {0} = 1" + using ordertype_eq_Ord by fastforce + 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 "t \ A'" "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 (simp add: M_def cong: conj_cong) + qed + + have 10: "\x0 \ A. \g \ elts \ \ elts \. strict_mono_on g (elts \) \ (\\ \ 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}" + 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, hide_lams) 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 have "k \ p" + using that F by auto + with k show ?thesis + using le_neq_implies_less 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 split: if_split_asm) + have "disjnt (D i) (D j)" if "i < j" "j \ p" for i j + proof (cases j) + case (Suc j') + then show ?thesis + using that + apply (auto simp: disjnt_def D_def) + using not_less_eq by (blast intro: less_\_D less_trans Suc_leD)+ + qed (use that in auto) + then have disjnt_DD: "disjnt (D i) (D j)" if "i \ j" "i \ p" "j \ p" for i j + using disjnt_sym nat_neq_iff that by auto + 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) + have "tp A'' \ \" + using \A'' \ A\ down ordertype_VWF_mono A by blast + moreover have "\ \ \ elts (\*\)" "tp (\ \) = \" + using \ \\ \ elts \\ by auto + then have "\ \ tp A''" + using IX' [OF _ _ A'] by (simp add: A''_def) + ultimately show "tp A'' = \" + by (rule antisym) + have "\ \ M (elts \) \ x" "F \ M (elts \) \ x" + if "x \ A''" for x + proof - + show "F \ M (elts \) \ x" + using A''_def FN that by blast + show "\ \ M (elts \) \ x" + using \\ \ elts \\ that by (simp add: M_def A''_def) + qed + then show "\x\A''. insert \ F \ M (elts \) \ x" + 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 \)\ + \ (\\\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))" + 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)" + 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 \] + apply (simp add: D_def split: if_split_asm) + apply (metis (no_types) dual_order.order_iff_strict le0 less_\_I less_trans) + by (metis (no_types) dual_order.order_iff_strict less_\_I less_trans) + 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 \)" + by (auto simp: g_def strict_mono_on_def dest!: F_imp_Ex) + show False + using * [OF \x \ A\ fun_g sm_g] + proof safe + fix \ + assume "\ \ elts \" and \: "tp (K 1 x \ \ (g \)) < \" + have FM: "F \ M (elts \) \ x" + by (meson FN \x \ A'\) + have False if "tp (K (Suc 0) x \ \ \) < \" "\ \ F" + using that FM 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 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 "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 \: "{x \ A''. tp (D L) \ tp (M (D L) \ x)} = {}" + using leD by blast + have "\ \ tp {x \ A''. tp (D L) \ tp (M (D L) \ x)}" + using IX [OF D\ A'' \D] by simp + then show False + using \\ \ 0\ by (simp add: \) + 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 \) \ (\i\n. g (\ i) = \ i) + \ (\\ \ elts \. \' \ \ K 1 xn \ \ (g \)) + \ less_sets {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 ` {..) \" + using that by (auto simp: \_def) + have \\: "\ ` {..n} \ elts \" and \sub: "\ (\ n) \ elts (\ * \)" + by (auto simp: \) + have \fun: "\ \ elts \ \ {X. X \ elts (\ * \) \ tp X = \}" + by (simp add: \) + 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 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 \))" + proof (rule ordertype_VWF_mono) + show "small (\ (g \))" + by (metis PiE \ down g that) + qed auto + 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 A2 where A2: "tp A2 = \" "A2 \ K 1 xn \ \ (\ n)" "less_sets {xn} A2" + using indecomposable_imp_Ex_less_sets [OF indec \\ > 1\ \Ord \\ tp2] + by (metis \sub \_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 (\ * \)" + using \sub sub A2 by fastforce+ + moreover have "xn \ elts (\ * \)" + using \sub xn by blast + 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 H_Suc: "H (Suc n) = (case H n of (g0, xa, xb) \ G n xa xb)" for n + by (simp add: H_def) + have "(\(g,\,x). \ n \ x) (H n)" for n + proof (induction n) + case 0 show ?case + by (simp add: H_def base) + next + case (Suc n) then show ?case + using G\ by (fastforce simp: H_Suc) + qed + 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_Suc 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 \)" + 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: "less_sets {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 + 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: "\ 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: "\ 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: "\ i j (\ j) = \ i \ \ j = \ i" for i j + by (metis eq_refl le less less_le) + have 18: "less_sets (\ m (\ m)) (\ n (\ n)) \ \ m < \ n" for m n + proof (cases n m rule: linorder_cases) + case less + show ?thesis + proof (intro iffI) + assume "less_sets (\ m (\ m)) (\ n (\ n))" + moreover + have "\ less_sets (\ m (\ m)) (\ n (\ n))" if "\ n = \ m" + by (metis "*" "15" eq less less_V_def less_sets_def less_sets_weaken2 that) + moreover + have "\ less_sets (\ m (\ m)) (\ n (\ n))" if "\ n < \ m" + using that 12 15 * [OF less] + apply (clarsimp simp: less_sets_def strict_mono_sets_def) + by (metis Ord_in_Ord Ord_linear2 \_in_\ \_in_\ \Ord \\ le leD less_asym subsetD) + ultimately show "\ m < \ n" + by (meson Ord_in_Ord Ord_linear_lt \_in_\ \Ord \\) + next + assume "\ m < \ n" + then have "less_sets (\ n (\ n m (\ m))) (\ n (\ n))" + by (metis "12" \_in_\ \_in_\ eq le less_V_def strict_mono_sets_def) + then show "less_sets (\ m (\ m)) (\ n (\ n))" + by (meson *[OF less] less_sets_weaken1) + qed + next + case equal + with 15 show ?thesis by auto + next + case greater + show ?thesis + proof (intro iffI) + assume "less_sets (\ m (\ m)) (\ n (\ n))" + moreover + have "\ less_sets (\ m (\ m)) (\ n (\ n))" if "\ n = \ m" + by (metis "*" "15" disjnt_iff eq greater in_mono less_sets_imp_disjnt that) + moreover + have "\ less_sets (\ m (\ m)) (\ n (\ n))" if "\ n < \ m" + using that 12 15 * [OF greater] + apply (clarsimp simp: less_sets_def strict_mono_sets_def) + by (meson \_in_\ \_in_\ in_mono less less_asym) + ultimately show "\ m < \ n" + by (meson Ord_\ Ord_linear_lt) + next + assume "\ m < \ n" + then have "less_sets (\ m (\ m)) (\ m (\ m n (\ n)))" + by (meson 12 Ord_in_Ord Ord_linear2 \_in_\ \_in_\ le leD ord(2) strict_mono_sets_def) + then show "less_sets (\ m (\ m)) (\ n (\ n))" + by (meson "*" greater less_sets_weaken2) + qed + 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 (simp add: less_sets_def) (metis "*" "15" "17" Suc_lessI le_SucI subsetD) + 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 + 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 \)" + 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 +qed + + +theorem Erdos_Milner: + assumes \: "\ \ elts \1" + shows "partn_lst_VWF (\\(1 + \ * ord_of_nat 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] simp add:) +next + case False + then have "n > 0" "k > 0" + by auto + 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 (Suc(k-1) * Suc(n-1))" + by simp + also have "\ \ Suc (k * n)" + using False by auto + finally 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)) [ord_of_nat (2 ^ (k-1)), \ \ (1 + ord_of_nat (n-1))] 2" + using Partitions.partn_lst_greater_resource PV x by blast + then have "partn_lst_VWF (\\(n*k)) [\ \ (1 + ord_of_nat (n-1)), ord_of_nat (2 ^ (k-1))] 2" + using partn_lst_two_swap by blast + moreover have "(1 + ord_of_nat (n-1)) = ord_of_nat n" + using ord_of_minus_1 [OF \n > 0\] + by (simp add: one_V_def) + ultimately have "partn_lst_VWF (\\(n*k)) [\ \ n, ord_of_nat (2 ^ (k-1))] 2" + by simp + then show ?thesis + using power_gt_expt [of 2 "k-1"] + by (force simp: less_Suc_eq intro: partn_lst_less) +qed + +end diff --git a/thys/Ordinal_Partitions/Library_Additions.thy b/thys/Ordinal_Partitions/Library_Additions.thy new file mode 100644 --- /dev/null +++ b/thys/Ordinal_Partitions/Library_Additions.thy @@ -0,0 +1,1177 @@ +section \Library additions\ + +theory Library_Additions + imports "ZFC_in_HOL.Ordinal_Exp" "HOL-Library.Ramsey" "Nash_Williams.Nash_Williams" + +begin + +subsection \Already in the development version\ + +declare \_gt0 [simp] + +lemma irrefl_less_than: "irrefl less_than" + using irrefl_def by blast + +lemma total_on_less_than [simp]: "total_on A less_than" + using total_on_def by force+ + +lemma takeWhile_eq_Nil_iff: "takeWhile P xs = [] \ xs = [] \ \P (hd xs)" +by (cases xs) auto + +lemma lenlex_append1: + assumes len: "(us,xs) \ lenlex R" and eq: "length vs = length ys" + shows "(us @ vs, xs @ ys) \ lenlex R" + using len +proof (induction us) + case Nil + then show ?case + by (simp add: lenlex_def eq) +next + case (Cons u us) + with lex_append_rightI show ?case + by (fastforce simp add: lenlex_def eq) +qed + +lemma lenlex_append2 [simp]: + assumes "irrefl R" + shows "(us @ xs, us @ ys) \ lenlex R \ (xs, ys) \ lenlex R" +proof (induction us) + case Nil + then show ?case + by (simp add: lenlex_def) +next + case (Cons u us) + with assms show ?case + by (auto simp: lenlex_def irrefl_def) +qed + +lemma hd_concat: "\xs \ []; hd xs \ []\ \ hd (concat xs) = hd (hd xs)" + by (metis concat.simps(2) hd_Cons_tl hd_append2) + +lemma sorted_list_of_set_lessThan_Suc [simp]: + "sorted_list_of_set {..finite S; n < card S\ \ enumerate S n \ S" +proof (induction n arbitrary: S) + case 0 + then show ?case + by (metis all_not_in_conv card_empty enumerate.simps(1) not_less0 wellorder_Least_lemma(1)) +next + case (Suc n) + show ?case + using Suc.prems Suc.IH [of "S - {LEAST n. n \ S}"] + apply (simp add: enumerate.simps) + by (metis Diff_empty Diff_insert0 Suc_lessD card.remove less_Suc_eq) +qed + +lemma finite_enumerate_step: "\finite S; Suc n < card S\ \ enumerate S n < enumerate S (Suc n)" +proof (induction n arbitrary: S) + case 0 + then have "enumerate S 0 \ enumerate S (Suc 0)" + by (simp add: Least_le enumerate.simps(1) finite_enumerate_in_set) + moreover have "enumerate (S - {enumerate S 0}) 0 \ S - {enumerate S 0}" + by (metis 0 Suc_lessD Suc_less_eq card_Suc_Diff1 enumerate_in_set finite_enumerate_in_set) + then have "enumerate S 0 \ enumerate (S - {enumerate S 0}) 0" + by auto + ultimately show ?case + by (simp add: enumerate_Suc') +next + case (Suc n) + then show ?case + by (simp add: enumerate_Suc' finite_enumerate_in_set) +qed + +lemma finite_enumerate_mono: "\m < n; finite S; n < card S\ \ enumerate S m < enumerate S n" + by (induct m n rule: less_Suc_induct) (auto intro: finite_enumerate_step) + +lemma finite_enumerate_Suc'': + fixes S :: "'a::wellorder set" + assumes "finite S" "Suc n < card S" + shows "enumerate S (Suc n) = (LEAST s. s \ S \ enumerate S n < s)" + using assms +proof (induction n arbitrary: S) + case 0 + then have "\s \ S. enumerate S 0 \ s" + by (auto simp: enumerate.simps intro: Least_le) + then show ?case + unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"] + by (metis Diff_iff dual_order.strict_iff_order singletonD singletonI) +next + case (Suc n S) + then have "Suc n < card (S - {enumerate S 0})" + using Suc.prems(2) finite_enumerate_in_set by force + then show ?case + apply (subst (1 2) enumerate_Suc') + apply (simp add: Suc) + apply (intro arg_cong[where f = Least] HOL.ext) + using finite_enumerate_mono[OF zero_less_Suc \finite S\, of n] Suc.prems + by (auto simp flip: enumerate_Suc') +qed + +lemma finite_enumerate_initial_segment: + fixes S :: "'a::wellorder set" + assumes "finite S" and n: "n < card (S \ {.. {.. S \ n < s) = (LEAST n. n \ S)" + proof (rule Least_equality) + have "\t. t \ S \ t < s" + by (metis "0" card_gt_0_iff disjoint_iff_not_equal lessThan_iff) + then show "(LEAST n. n \ S) \ S \ (LEAST n. n \ S) < s" + by (meson LeastI Least_le le_less_trans) + qed (simp add: Least_le) + then show ?case + by (auto simp: enumerate_0) +next + case (Suc n) + then have less_card: "Suc n < card S" + by (meson assms(1) card_mono inf_sup_ord(1) leD le_less_linear order.trans) + obtain T where T: "T \ {s \ S. enumerate S n < s}" + by (metis Infinite_Set.enumerate_step enumerate_in_set finite_enumerate_in_set finite_enumerate_step less_card mem_Collect_eq) + have "(LEAST x. x \ S \ x < s \ enumerate S n < x) = (LEAST x. x \ S \ enumerate S n < x)" + (is "_ = ?r") + proof (intro Least_equality conjI) + show "?r \ S" + by (metis (mono_tags, lifting) LeastI mem_Collect_eq T) + have "\ s \ ?r" + using not_less_Least [of _ "\x. x \ S \ enumerate S n < x"] Suc assms + by (metis (mono_tags, lifting) Int_Collect Suc_lessD finite_Int finite_enumerate_in_set finite_enumerate_step lessThan_def less_le_trans) + then show "?r < s" + by auto + show "enumerate S n < ?r" + by (metis (no_types, lifting) LeastI mem_Collect_eq T) + qed (auto simp: Least_le) + then show ?case + using Suc assms by (simp add: finite_enumerate_Suc'' less_card) +qed + +lemma finite_enumerate_Ex: + fixes S :: "'a::wellorder set" + assumes S: "finite S" + and s: "s \ S" + shows "\ny\S. y < s") + case True + let ?T = "S \ {..finite S\]) + from True have y: "\x. Max ?T < x \ (\s'\S. s' < s \ s' < x)" + by (subst Max_less_iff) (auto simp: \finite ?T\) + then have y_in: "Max ?T \ {s'\S. s' < s}" + using Max_in \finite ?T\ by fastforce + with less.IH[of "Max ?T" ?T] obtain n where n: "enumerate ?T n = Max ?T" "n < card ?T" + using \finite ?T\ by blast + then have "Suc n < card S" + using TS less_trans_Suc by blast + with S n have "enumerate S (Suc n) = s" + by (subst finite_enumerate_Suc'') (auto simp: y finite_enumerate_initial_segment less finite_enumerate_Suc'' intro!: Least_equality) + then show ?thesis + using \Suc n < card S\ by blast + next + case False + then have "\t\S. s \ t" by auto + moreover have "0 < card S" + using card_0_eq less.prems by blast + ultimately show ?thesis + using \s \ S\ + by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) + qed +qed + +lemma finite_bij_enumerate: + fixes S :: "'a::wellorder set" + assumes S: "finite S" + shows "bij_betw (enumerate S) {..n m. \n \ m; n < card S; m < card S\ \ enumerate S n \ enumerate S m" + using finite_enumerate_mono[OF _ \finite S\] by (auto simp: neq_iff) + then have "inj_on (enumerate S) {..s \ S. \ifinite S\ + ultimately show ?thesis + unfolding bij_betw_def by (auto intro: finite_enumerate_in_set) +qed + +lemma length_sorted_list_of_set [simp]: "length (sorted_list_of_set A) = card A" +proof (cases "finite A") + case True + then show ?thesis + by(metis distinct_card distinct_sorted_list_of_set set_sorted_list_of_set) +qed auto + +lemmas sorted_list_of_set = set_sorted_list_of_set sorted_sorted_list_of_set distinct_sorted_list_of_set + +lemma strict_sorted_equal: + assumes "strict_sorted xs" + and "strict_sorted ys" + and "set ys = set xs" + shows "ys = xs" + using assms +proof (induction xs arbitrary: ys) + case (Cons x xs) + show ?case + proof (cases ys) + case Nil + then show ?thesis + using Cons.prems by auto + next + case (Cons y ys') + then have "xs = ys'" + by (metis Cons.prems list.inject sorted_distinct_set_unique strict_sorted_iff) + moreover have "x = y" + using Cons.prems \xs = ys'\ local.Cons by fastforce + ultimately show ?thesis + using local.Cons by blast + qed +qed auto + +lemma sorted_list_of_set_inject: + assumes "sorted_list_of_set A = sorted_list_of_set B" "finite A" "finite B" + shows "A = B" + using assms set_sorted_list_of_set by fastforce + +lemma sorted_list_of_set_unique: + assumes "finite A" + shows "strict_sorted l \ set l = A \ length l = card A \ sorted_list_of_set A = l" + using assms strict_sorted_equal by force + +lemma iso_iff2: "iso r r' f \ + bij_betw f (Field r) (Field r') \ + (\a \ Field r. \b \ Field r. (a, b) \ r \ (f a, f b) \ r')" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + then have "bij_betw f (Field r) (Field r')" and emb: "embed r r' f" + by (auto simp: bij_betw_def iso_def) + then obtain g where g: "\x. x \ Field r \ g (f x) = x" + by (auto simp: bij_betw_iff_bijections) + moreover + have "(a, b) \ r" if "a \ Field r" "b \ Field r" "(f a, f b) \ r'" for a b + using that emb g g [OF FieldI1] \\yes it's weird\ + by (force simp: embed_def under_def bij_betw_iff_bijections) + ultimately show ?rhs + using L by (auto simp: compat_def iso_def dest: embed_compat) +next + assume R: ?rhs + then show ?lhs + apply (clarsimp simp add: iso_def embed_def under_def bij_betw_iff_bijections) + apply (rule_tac x=g in exI) + apply (fastforce simp add: intro: FieldI1)+ + done +qed + +lemma sorted_list_of_set_atMost_Suc [simp]: + "sorted_list_of_set {..Suc k} = sorted_list_of_set {..k} @ [Suc k]" + using lessThan_Suc_atMost sorted_list_of_set_lessThan_Suc by fastforce + +lemma sorted_list_of_set_greaterThanLessThan: + assumes "Suc i < j" + shows "sorted_list_of_set {i<.. j" + shows "sorted_list_of_set {i<..j} = Suc i # sorted_list_of_set {Suc i<..j}" + using sorted_list_of_set_greaterThanLessThan [of i "Suc j"] + by (metis assms greaterThanAtMost_def greaterThanLessThan_eq le_imp_less_Suc lessThan_Suc_atMost) + +lemma nth_sorted_list_of_set_greaterThanLessThan: + "n < j - Suc i \ sorted_list_of_set {i<.. sorted_list_of_set {i<..j} ! n = Suc (i+n)" + using nth_sorted_list_of_set_greaterThanLessThan [of n "Suc j" i] + by (simp add: greaterThanAtMost_def greaterThanLessThan_eq lessThan_Suc_atMost) + +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 elts_multE: + assumes "z \ elts (x * y)" + obtains u v where "u \ elts x" "v \ elts y" "z = x*v + u" + using mult [of x y] lift_def assms by auto + +lemma Ord_add_mult_iff: + assumes "\ \ elts \" "\' \ elts \" "Ord \" "Ord \'" "Ord \" + shows "\ * \ + \ \ elts (\ * \' + \') \ \ \ elts \' \ \ = \' \ \ \ elts \'" (is "?lhs \ ?rhs") +proof + assume L: ?lhs + show ?rhs + proof (cases "\ \ elts \'") + case False + with assms have "\ = \'" + by (meson L Ord_linear Ord_mult Ord_trans add_mult_less not_add_mem_right) + then show ?thesis + using L less_V_def by auto + qed auto +next + assume R: ?rhs + then show ?lhs + proof + assume "\ \ elts \'" + then obtain \ where "\' = \+\" + by (metis OrdmemD assms(3) assms(4) le_Ord_diff less_V_def) + show ?lhs + using assms + by (meson \\ \ elts \'\ add_le_cancel_left0 add_mult_less vsubsetD) + next + assume "\ = \' \ \ \ elts \'" + then show ?lhs + using less_V_def by auto + qed +qed + +lemma small_Times [simp]: + assumes "small A" "small B" + shows "small (A \ B)" +proof - + obtain f a g b where "inj_on f A" "inj_on g B" and f: "f ` A = elts a" and g: "g ` B = elts b" + using assms by (auto simp: small_def) + define h where "h \ \(x,y). \f x, g y\" + show ?thesis + unfolding small_def + proof (intro exI conjI) + show "inj_on h (A \ B)" + using \inj_on f A\ \inj_on g B\ by (simp add: h_def inj_on_def) + have "h ` (A \ B) = elts (vtimes a b)" + using f g by (fastforce simp: h_def image_iff split: prod.split) + then show "h ` (A \ B) \ range elts" + by blast + qed +qed + +lemma ordertype_Times: + assumes "small A" "small B" and r: "wf r" "trans r" "total_on A r" and s: "wf s" "trans s" "total_on B s" + shows "ordertype (A\B) (r <*lex*> s) = ordertype B s * ordertype A r" (is "_ = ?\ * ?\") +proof (subst ordertype_eq_iff) + show "Ord (?\ * ?\)" + by (intro wf_Ord_ordertype Ord_mult r s; simp) + define f where "f \ \(x,y). ?\ * ordermap A r x + (ordermap B s y)" + show "\f. bij_betw f (A \ B) (elts (?\ * ?\)) \ (\x\A \ B. \y\A \ B. (f x < f y) = ((x, y) \ (r <*lex*> s)))" + unfolding bij_betw_def + proof (intro exI conjI strip) + show "inj_on f (A \ B)" + proof (clarsimp simp: f_def inj_on_def) + fix x y x' y' + assume "x \ A" "y \ B" "x' \ A" "y' \ B" + and eq: "?\ * ordermap A r x + ordermap B s y = ?\ * ordermap A r x' + ordermap B s y'" + have "ordermap A r x = ordermap A r x' \ + ordermap B s y = ordermap B s y'" + proof (rule mult_cancellation_lemma [OF eq]) + show "ordermap B s y \ ?\" + using ordermap_in_ordertype [OF \y \ B\, of s] less_TC_iff \small B\ by blast + show "ordermap B s y' \ ?\" + using ordermap_in_ordertype [OF \y' \ B\, of s] less_TC_iff \small B\ by blast + qed + then show "x = x' \ y = y'" + using \x \ A\ \x' \ A\ \y \ B\ \y' \ B\ r s \small A\ \small B\ by auto + qed + show "f ` (A \ B) = elts (?\ * ?\)" (is "?lhs = ?rhs") + proof + show "f ` (A \ B) \ elts (?\ * ?\)" + apply (auto simp: f_def add_mult_less ordermap_in_ordertype wf_Ord_ordertype r s) + by (simp add: add_mult_less assms ordermap_in_ordertype wf_Ord_ordertype) + show "elts (?\ * ?\) \ f ` (A \ B)" + proof (clarsimp simp: f_def image_iff elim!: elts_multE split: prod.split) + fix u v + assume u: "u \ elts (?\)" and v: "v \ elts ?\" + have "inv_into B (ordermap B s) u \ B" + by (simp add: inv_into_ordermap u) + moreover have "inv_into A (ordermap A r) v \ A" + by (simp add: inv_into_ordermap v) + ultimately show "\x\A. \y\B. ?\ * v + u = ?\ * ordermap A r x + ordermap B s y" + by (metis \small A\ \small B\ bij_betw_inv_into_right ordermap_bij r(1) r(3) s(1) s(3) u v) + qed + qed + next + fix p q + assume "p \ A \ B" and "q \ A \ B" + then obtain u v x y where \
: "p = (u,v)" "u \ A" "v \ B" "q = (x,y)" "x \ A" "y \ B" + by blast + show "((f p) < f q) = ((p, q) \ (r <*lex*> s))" + proof + assume "f p < f q" + with \
assms have "(u, x) \ r \ u=x \ (v, y) \ s" + apply (simp add: f_def) + by (metis Ord_add Ord_add_mult_iff Ord_mem_iff_lt Ord_mult Ord_ordermap converse_ordermap_mono + ordermap_eq_iff ordermap_in_ordertype wf_Ord_ordertype) + then show "(p,q) \ (r <*lex*> s)" + by (simp add: \
) + next + assume "(p,q) \ (r <*lex*> s)" + then have "(u, x) \ r \ u = x \ (v, y) \ s" + by (simp add: \
) + then show "f p < f q" + proof + assume ux: "(u, x) \ r" + have oo: "\x. Ord (ordermap A r x)" "\y. Ord (ordermap B s y)" + by (simp_all add: r s) + show "f p < f q" + proof (clarsimp simp: f_def split: prod.split) + fix a b a' b' + assume "p = (a, b)" and "q = (a', b')" + then have "?\ * ordermap A r a + ordermap B s b < ?\ * ordermap A r a'" + using ux assms \
+ by (metis Ord_mult Ord_ordermap OrdmemD Pair_inject add_mult_less ordermap_in_ordertype ordermap_mono wf_Ord_ordertype) + also have "\ \ ?\ * ordermap A r a' + ordermap B s b'" + by simp + finally show "?\ * ordermap A r a + ordermap B s b < ?\ * ordermap A r a' + ordermap B s b'" . + qed + next + assume "u = x \ (v, y) \ s" + then show "f p < f q" + using \
assms by (fastforce simp: f_def split: prod.split intro: ordermap_mono_less) + qed + qed + qed +qed (use assms small_Times in auto) + +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" + using total_on_def by (fastforce intro!: ordertype_inc_eq \small N\) + ultimately show ?thesis + by simp +qed +lemma infinite_infinite_partition: + assumes "infinite A" + obtains C :: "nat \ 'a set" + where "pairwise (\i j. disjnt (C i) (C j)) UNIV" "(\i. C i) \ A" "\i. infinite (C i)" +proof - + obtain f :: "nat\'a" where "range f \ A" "inj f" + using assms infinite_countable_subset by blast + let ?C = "\i. range (\j. f (prod_encode (i,j)))" + show thesis + proof + show "pairwise (\i j. disjnt (?C i) (?C j)) UNIV" + by (auto simp: pairwise_def disjnt_def inj_on_eq_iff [OF \inj f\] inj_on_eq_iff [OF inj_prod_encode, of _ UNIV]) + show "(\i. ?C i) \ A" + using \range f \ A\ by blast + have "infinite (range (\j. f (prod_encode (i, j))))" for i + by (rule range_inj_infinite) (meson Pair_inject \inj f\ inj_def prod_encode_eq) + then show "\i. infinite (?C i)" + using that by auto + qed +qed + +text \This is already installed in the development AFP entry\ +lemma mult_cancellation_half: + assumes "a*x + r \ a*y + s" "r \ a" "s \ a" + shows "x \ y" +proof - + have "x \ y" if "Ord \" "x \ elts (Vset \)" "y \ elts (Vset \)" for \ + using that assms + proof (induction \ arbitrary: x y r s rule: Ord_induct3) + case 0 + then show ?case + by auto + next + case (succ k) + show ?case + proof + fix u + assume u: "u \ elts x" + have u_k: "u \ elts (Vset k)" + using Vset_succ succ.hyps succ.prems(1) u by auto + obtain r' where "r' \ elts a" "r \ r'" + using less_TC_iff succ.prems(4) by blast + have "a*u + r' \ elts (lift (a*u) a)" + by (simp add: \r' \ elts a\ lift_def) + also have "\ \ elts (a*x)" + using u by (force simp: mult [of _ x]) + also have "\ \ elts (a*y + s)" + using plus_eq_lift succ.prems(3) by auto + also have "\ = elts (a*y) \ elts (lift (a*y) s)" + by (simp add: plus_eq_lift) + finally have "a * u + r' \ elts (a * y) \ elts (lift (a * y) s)" . + then show "u \ elts y" + proof + assume *: "a * u + r' \ elts (a * y)" + show "u \ elts y" + proof - + obtain v e where v: "v \ elts y" "e \ elts a" "a * u + r' = a * v + e" + using * by (auto simp: mult [of _ y] lift_def) + then have v_k: "v \ elts (Vset k)" + using Vset_succ_TC less_TC_iff succ.prems(2) by blast + then show ?thesis + by (metis \r' \ elts a\ antisym le_TC_refl less_TC_iff order_refl succ.IH u_k v) + qed + next + assume "a * u + r' \ elts (lift (a * y) s)" + then obtain t where "t \ elts s" and t: "a * u + r' = a * y + t" + using lift_def by auto + have noteq: "a*y \ a*u" + proof + assume "a*y = a*u" + then have "lift (a*y) a = lift (a*u) a" + by metis + also have "\ \ a*x" + unfolding mult [of _ x] using \u \ elts x\ by (auto intro: cSUP_upper) + also have "\ \ a*y \ lift (a*y) s" + using \elts (a * x) \ elts (a * y + s)\ plus_eq_lift by auto + finally have "lift (a*y) a \ a*y \ lift (a*y) s" . + then have "lift (a*y) a \ lift (a*y) s" + using add_le_cancel_left less_TC_imp_not_le plus_eq_lift \s \ a\ by auto + then have "a \ s" + by (simp add: le_iff_sup lift_eq_lift lift_sup_distrib) + then show False + using \s \ a\ less_TC_imp_not_le by auto + qed + consider "a * u \ a * y" | "a * y \ a * u" + using t comparable vle_comparable_def by blast + then have "False" + proof cases + case 1 + then obtain c where "a*y = a*u + c" + by (metis vle_def) + then have "c+t = r'" + by (metis add_right_cancel add.assoc t) + then have "c \ a" + using \r' \ elts a\ less_TC_iff vle2 vle_def by force + moreover have "c \ 0" + using \a * y = a * u + c\ noteq by auto + ultimately show ?thesis + using \a * y = a * u + c\ mult_eq_imp_0 by blast + next + case 2 + then obtain c where "a*u = a*y + c" + by (metis vle_def) + then have "c+r' = t" + by (metis add_right_cancel add.assoc t) + then have "c \ a" + by (metis \t \ elts s\ less_TC_iff less_TC_trans \s \ a\ vle2 vle_def) + moreover have "c \ 0" + using \a * u = a * y + c\ noteq by auto + ultimately show ?thesis + using \a * u = a * y + c\ mult_eq_imp_0 by blast + qed + then show "u \ elts y" .. + qed + qed + next + case (Limit k) + obtain i j where k: "i \ elts k" "j \ elts k" + and x: "x \ elts (Vset i)" and y: "y \ elts (Vset j)" + using that Limit by (auto simp: Limit_Vfrom_eq) + show ?case + proof (rule Limit.IH [of "i \ j"]) + show "i \ j \ elts k" + by (meson k x y Limit.hyps Limit_def Ord_in_Ord Ord_mem_iff_lt Ord_sup union_less_iff) + show "x \ elts (Vset (i \ j))" "y \ elts (Vset (i \ j))" + using x y by (auto simp: Vfrom_sup) + show "a * x + r \ a * y + s" + by (simp add: Limit.prems) + qed (auto simp: Limit.prems) + qed + then show ?thesis + by (metis two_in_Vset Ord_rank Ord_VsetI rank_lt) +qed + +corollary mult_cancellation_less: + assumes lt: "a*x + r < a*y + s" and "r \ a" "s \ a" + obtains "x < y" | "x = y" "r < s" +proof - + have "x \ y" + by (meson assms dual_order.strict_implies_order mult_cancellation_half) + then consider "x < y" | "x = y" + using less_V_def by blast + with lt that show ?thesis by blast +qed + +subsection \For HOL\ + +lemma enumerate_mono_iff [simp]: + "infinite S \ enumerate S m < enumerate S n \ m < n" + by (metis enumerate_mono less_asym less_linear) + +lemma finite_enumerate_mono_iff [simp]: + "\finite S; m < card S; n < card S\ \ enumerate S m < enumerate S n \ m < n" + by (metis finite_enumerate_mono less_asym less_linear) + +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 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 lexl_not_refl [simp]: "irrefl r \ (x,x) \ lex r" + by (meson irrefl_def lex_take_index) + +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 finite_enum_subset: + assumes "\i. i < card X \ enumerate X i = enumerate Y i" and "finite X" "finite Y" "card X \ card Y" + shows "X \ Y" + by (metis assms finite_enumerate_Ex finite_enumerate_in_set less_le_trans subsetI) + +lemma finite_enum_ext: + assumes "\i. i < card X \ enumerate X i = enumerate Y i" and "finite X" "finite Y" "card X = card Y" + shows "X = Y" + by (intro antisym finite_enum_subset) (auto simp: assms) + +thm card_Un_disjoint +lemma card_Un_disjnt: "\finite A; finite B; disjnt A B\ \ card (A \ B) = card A + card B" + by (simp add: card_Un_disjoint disjnt_def) + + +lemma sorted_list_of_set_nonempty: + assumes "finite I" "I \ {}" + shows "sorted_list_of_set I = Min I # sorted_list_of_set (I - {Min I})" + using assms by (auto simp: less_le simp flip: sorted_list_of_set_unique intro: Min_in) +lemma strict_sorted_imp_sorted: "strict_sorted xs \ sorted xs" + by (auto simp: strict_sorted_iff) + +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 finite_Inf_in: + fixes A :: "'a::complete_lattice set" + assumes "finite A" "A\{}" and inf: "\x y. \x \ A; y \ A\ \ inf x y \ A" + shows "Inf A \ A" +proof - + have "Inf B \ A" if "B \ A" "B\{}" for B + using finite_subset [OF \B \ A\ \finite A\] that + by (induction B) (use inf in \force+\) + then show ?thesis + by (simp add: assms) +qed + +lemma finite_Sup_in: + fixes A :: "'a::complete_lattice set" + assumes "finite A" "A\{}" and sup: "\x y. \x \ A; y \ A\ \ sup x y \ A" + shows "Sup A \ A" +proof - + have "Sup B \ A" if "B \ A" "B\{}" for B + using finite_subset [OF \B \ A\ \finite A\] that + by (induction B) (use sup in \force+\) + then show ?thesis + by (simp add: assms) +qed + +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 + +(*METIS NOT ALLOWED!*) + +lemma iso_trans: + assumes "trans r" "iso r r' f" shows "trans r'" + using assms unfolding trans_def iso_iff2 bij_betw_iff_bijections + by (metis (full_types) FieldI1 FieldI2) + +lemma iso_Total: + assumes "Total r" "iso r r' f" shows "Total r'" + using assms unfolding total_on_def iso_iff2 bij_betw_iff_bijections by metis + +lemma iso_wf: + assumes "wf r" "iso r r' f" shows "wf r'" +proof - + have bij: "bij_betw f (Field r) (Field r')" + and iff: "(\a \ Field r. \b \ Field r. (a, b) \ r \ (f a, f b) \ r')" + using assms by (auto simp: iso_iff2) + show ?thesis + proof (rule wfI_min) + fix x::'b and Q + assume "x \ Q" + let ?Q = "inv_into (Field r) f ` Q" + obtain z where "z \ ?Q" "\x y. \(y, z) \ r; x \ Q\ \ y \ inv_into (Field r) f x" + by (metis \x \ Q\ \wf r\ image_eqI wfE_min) + with bij show "\z\Q. \y. (y, z) \ r' \ y \ Q" + by (metis (no_types, lifting) FieldI2 bij_betw_imp_surj_on f_inv_into_f iff inv_into_into) + qed +qed + + +subsection \Also in the Nash-Williams development\ +text \FIXME: these contain duplicates and need consolidation\ + +lemma less_setsD: "\less_sets A B; a \ A; b \ B\ \ a < b" + by (auto simp: less_sets_def) + +lemma less_sets_irrefl [simp]: "less_sets A A \ A = {}" + by (auto simp: less_sets_def) + +lemma less_sets_trans: "\less_sets A B; less_sets B C; B \ {}\ \ less_sets A C" + unfolding less_sets_def using less_trans by blast + +lemma less_sets_weaken1: "\less_sets A' B; A \ A'\ \ less_sets A B" + by (auto simp: less_sets_def) + +lemma less_sets_weaken2: "\less_sets A B'; B \ B'\ \ less_sets A B" + by (auto simp: less_sets_def) + +lemma less_sets_imp_disjnt: "less_sets A B \ disjnt A B" + by (auto simp: less_sets_def disjnt_def) + +lemma less_sets_Un1: "less_sets (A \ A') B \ less_sets A B \ less_sets A' B" + by (auto simp: less_sets_def) + +lemma less_sets_Un2: "less_sets A (B \ B') \ less_sets A B \ less_sets A B'" + by (auto simp: less_sets_def) + +lemma less_sets_UN1: "less_sets (\\) B \ (\A\\. less_sets A B)" + by (auto simp: less_sets_def) + +lemma less_sets_UN2: "less_sets A (\ \) \ (\B\\. less_sets A B)" + by (auto simp: less_sets_def) + +lemma strict_sorted_imp_less_sets: + "strict_sorted (as @ bs) \ less_sets (list.set as) (list.set bs)" + by (simp add: less_sets_def sorted_wrt_append strict_sorted_sorted_wrt) + +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" + 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\ + +(*NOT THE FOLLOWING*) + +lemma sorted_list_of_set_insert: + 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 assms(1) 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 strict_sorted_sorted_wrt) + 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 \Ramsey\ + +lemma nsets_Pi_contra: "A' \ A \ Pi ([A]\<^bsup>n\<^esup>) B \ Pi ([A']\<^bsup>n\<^esup>) B" + by (auto simp: nsets_def) + +subsection \Misc additions to the ZF libraries\ + + +lemma oexp_\_Limit: "Limit \ \ \\\ = (SUP \ \ elts \. \\\)" + by (simp add: oexp_Limit) + + +lemma oexp_mult_commute: + fixes j::nat + shows "Ord \ \ (\ \ j) * \ = \ * (\ \ j)" + by (metis Ord_1 Ord_ord_of_nat oexp_1_right oexp_add oexp_succ one_V_def ord_of_nat_\ succ_0_plus_eq) + +lemma iso_imp_ordertype_eq_ordertype: + assumes iso: "iso r r' f" + and "wf r" + and "Total r" + and "trans r" + and sm: "small (Field r)" + shows "ordertype (Field r) r = ordertype (Field r') r'" +proof (subst ordertype_eq_ordertype) + show "small (Field r')" + by (metis iso sm iso_Field replacement) + show "\f. bij_betw f (Field r) (Field r') \ (\x\Field r. \y\Field r. ((f x, f y) \ r') = ((x, y) \ r))" + using assms(1) iso_iff2 by blast +qed (use assms iso_wf iso_Total iso_trans in auto) + +proposition 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 total_on_imp_Total_Restr: "total_on A r \ Total (Restr r A)" + by (auto simp: Field_def total_on_def) + +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 + +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, hide_lams) 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) {.. 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 new file mode 100644 --- /dev/null +++ b/thys/Ordinal_Partitions/Omega_Omega.thy @@ -0,0 +1,4670 @@ +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\ +primrec normal:: "nat list \ bool" where + normal_0: "normal [] = True" +| normal_Suc: "normal (m#ms) = (m > 0)" + +lemma omega_sum_0_iff [simp]: "normal ns \ omega_sum ns = 0 \ ns = []" + by (induction ns) 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 Nil + then show ?case + by (auto simp: zero_less_Limit) +next + 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 + +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 "length ms < length ns" "normal ns" + shows "omega_sum ms < omega_sum ns" +proof (cases ns) + case Nil + then show ?thesis + using assms by auto +next + case (Cons n ns') + have "\ \ length ms \ \ \ length ns'" + using assms local.Cons by (simp add: oexp_mono_le) + then have "\ omega_sum (n#ns') \ omega_sum ms" + using omega_sum_ge [of n ns'] omega_sum_less [of ms] \normal ns\ local.Cons by auto + then show ?thesis + by (metis Ord_linear2 Ord_omega_sum local.Cons) +qed + +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 + then have "m\n" + by auto + with L assms show ?rhs + by auto +next + assume ?rhs + then show ?lhs + by (auto simp: mult_nat_less_add_less omega_sum_aux_less assms) +qed + +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 L: ?lhs + then have "length ms \ length ns" + using omega_sum_length_leD [OF less_imp_le [OF L]] by (simp add: \m > 0\) + moreover have "m\n" if "length ms = length ns" + using L omega_sum_less_eqlen_iff_cases that by auto + ultimately show ?rhs + using L by auto +next + assume ?rhs + moreover + have "omega_sum (m # ms) < omega_sum (n # ns)" + if "length ms < length ns" + using that by (metis Suc_mono \n > 0\ length_Cons normal_Suc omega_sum_length_less) + ultimately show ?lhs + using omega_sum_less_eqlen_iff_cases by force +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 Nil + then show ?case + by auto +next + case (Cons m ms) + then show ?case + proof (induction ns) + case (Cons n ns') + show ?case + using omega_sum_lex_less_iff_cases [of ms m ns' n] Cons.prems + by (simp add: Cons_lenlex_iff lenlex_length order.not_eq_order_implies_strict nat_less_le) + qed auto +qed + +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. @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)" + by (metis (mono_tags, hide_lams) WW_def equalityI image_subset_iff into_from_WW mem_Collect_eq rangeI strict_sorted_into_WW subset_iff) + +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)" + by blast + +lemma 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) + next + show "\ ` elts (\\\) \ WW" + by (auto simp: \_def WW_def strict_sorted_into_WW) + qed auto + 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 "less_sets (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 "less_sets 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 "less_sets (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 \ less_sets (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 sorted_trans: + assumes "xs < ys" "ys < zs" "sorted ys" "ys \ []" shows "xs < zs" + using assms unfolding less_list_def + by (metis dual_order.strict_trans last_in_set leD neqE sorted_hd_le) + +lemma strict_sorted_imp_append_less: + assumes "strict_sorted (xs @ ys)" + shows "xs < ys" + using assms by (simp add: less_list_def sorted_wrt_append strict_sorted_sorted_wrt) + +lemma strict_sorted_append_iff: + "strict_sorted (xs @ ys) \ xs < ys \ strict_sorted xs \ strict_sorted ys" (is "?lhs = ?rhs") +proof + assume ?lhs then show ?rhs + by (auto simp: sorted_wrt_append strict_sorted_sorted_wrt less_list_def) +next + assume R: ?rhs + then have "\x y. \x \ list.set xs; y \ list.set ys\ \ x < y" + using less_setsD sorted_list_of_set_imp_less_sets strict_sorted_imp_sorted by blast + with R show ?lhs + by (auto simp: sorted_wrt_append strict_sorted_sorted_wrt) +qed + +lemma singleton_less_list_iff: "sorted xs \ [n] < xs \ {..n} \ list.set xs = {}" + apply (simp add: less_list_def set_eq_iff) + by (metis empty_iff less_le_trans list.set(1) list.set_sel(1) not_le sorted_hd_le) + +lemma less_last_iff: "xs@[x] < ys \ [x] < ys" + by (simp add: less_list_def) + +lemma less_Cons_iff: "NO_MATCH [] ys \ xs < y#ys \ xs < [y]" + by (simp add: less_list_def) + +lemma less_hd_imp_less: "xs < [hd ys] \ xs < ys" + by (simp add: less_list_def) + +lemma last_less_imp_less: "[last xs] < 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" + assumes "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 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'" "less_sets (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 \finite S'\ 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 + using sorted_list_of_set_imp_less_sets strict_sorted_imp_sorted + by (auto simp: init_segment_def strict_sorted_append_iff) + 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 "\j. \j<2; i\j\ \ (h \ list_of) -` {j} \ (list.set ` AM) \ Pow N = {}" + unfolding Ramsey_def by (metis \infinite M\) + then have N_disjoint: "(h \ list_of) -` {1-i} \ (list.set ` AM) \ Pow N = {}" + by (metis One_nat_def diff_less_Suc not_less_eq numeral_2_eq_2 zero_less_diff) + 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+ + +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)" + using "2" 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) + + +declare acc_lengths.simps [simp del] + +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 mem_lists_non_Nil 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 (metis concat.simps(2) hd_append2 interact.simps(2) interact.simps(3) list.exhaust 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 add: acc_lengths.simps) +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 have "as' = []" + by (metis acc_lengths_eq_Nil_iff) + with 1 show ?case + using acc_lengths_concat_injective by auto +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 simp: acc_lengths.simps) +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" + proof (clarsimp simp add: less_list_def) + assume "y \ []" and ne: "interact xs ys \ []" + then show "last y < hd (interact xs ys)" + using 3 + apply simp + by (metis dual_order.strict_trans1 hd_interact length_greater_0_conv less_list_def list.sel(1) lists.simps mem_lists_non_Nil nth_Cons' nth_mem) + qed + 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 (elim Form.cases Form_Body.cases; metis dual_order.order_iff_strict less_not_refl) + + +lemma Form_Body_WW: + assumes "Form_Body ka kb xs ys zs" + shows "zs \ WW" + by (rule Form_Body.cases [OF assms]) (auto simp: WW_def) + +lemma Form_Body_nonempty: + assumes "Form_Body ka kb xs ys zs" + shows "length zs > 0" + by (rule Form_Body.cases [OF assms]) auto + +lemma Form_Body_length: + assumes "Form_Body ka kb xs ys zs" + shows "length xs < length ys" + using Form_Body.cases assms by blast + +lemma form_cases: + fixes l::nat + obtains (zero) "l = 0" | (odd) k where "l = 2*k-1" "k > 0" | (even) k where "l = 2*k" "k > 0" +proof - + have "l = 0 \ (\k. l = 2*k-1 \ k > 0) \ (\k. l = 2*k \ k > 0)" + by presburger + then show thesis + using even odd zero by blast +qed + +lemma odd_eq_iffs: "2 * k - Suc 0 = 2 * k' - Suc 0 \ k = k' \ k=0 \ k'=0" + "2 * k - Suc 0 = 2 * k' \ k=0 \ k'=0" + "2 * k = 2 * k' - Suc 0 \ k=0 \ k'=0" + by presburger+ + +subsubsection \Interactions\ + +lemma interact: + assumes "Form l U" "l>0" + obtains k xs ys zs where "l = 2*k-1" "U = {xs,ys}" "Form_Body k k xs ys zs" "k > 0" + | k xs ys zs where "l = 2*k" "U = {xs,ys}" "Form_Body (Suc k) k xs ys zs" "k > 0" + by (rule Form.cases [OF \Form l U\]) (use \l>0\ in \force+\) + + +definition inter_scheme :: "nat \ nat list set \ nat list" + where "inter_scheme l U \ @zs. \k xs ys. l > 0 + \ (l = 2*k-1 \ U = {xs,ys} \ Form_Body k k xs ys zs + \ l = 2*k \ U = {xs,ys} \ Form_Body (Suc k) k xs ys zs)" + + +lemma inter_scheme: + assumes "Form l U" "l>0" + obtains (odd) k xs ys where "l = 2*k-1" "U = {xs,ys}" "Form_Body k k xs ys (inter_scheme l U)" + | (even) k xs ys where "l = 2*k" "U = {xs,ys}" "Form_Body (Suc k) k xs ys (inter_scheme l U)" + using form_cases [of l] +proof cases + case zero + with \l > 0\ show ?thesis + by auto +next + case (odd k) + show ?thesis + using interact [OF assms] + proof cases + case (1 k' xs ys zs) + show ?thesis + proof (rule that(1)) + have "\ Form_Body k k ys xs zs" for zs + using 1 Form_Body_length less_asym' by blast + then show "Form_Body k k xs ys (inter_scheme l U)" + using odd 1 + by (force simp: Set.doubleton_eq_iff odd_eq_iffs inter_scheme_def conj_disj_distribR ex_disj_distrib some_eq_ex) + qed (use 1 odd in auto) + qed (use odd in presburger) +next + case (even k) + show ?thesis + using interact [OF assms] + proof cases + case (2 k' xs ys zs) + show ?thesis + proof (rule that(2)) + have "\ Form_Body (Suc k) k ys xs zs" for zs + using 2 Form_Body_length less_asym' by blast + then show "Form_Body (Suc k) k xs ys (inter_scheme l U)" + using even 2 + by (force simp: Set.doubleton_eq_iff odd_eq_iffs inter_scheme_def conj_disj_distribR ex_disj_distrib some_eq_ex) + qed (use 2 even in auto) + qed (use even in presburger) +qed + + +lemma inter_scheme_simple: + assumes "Form k U" "k>0" + shows "inter_scheme k U \ WW \ length (inter_scheme k U) > 0" + using inter_scheme by (meson Form_Body_WW Form_Body_nonempty assms) + +lemma inter_scheme_strict_sorted: + assumes "Form k U" "k>0" + shows "strict_sorted (inter_scheme k U)" + using inter_scheme_simple [OF assms] by (auto simp: WW_def) + +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" + using inter_scheme [OF \Form l U\ \l>0\] +proof cases + case (1 k xs ys) + 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) = k" "length (b#bs) = k" + 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 + note one = 1 + show ?thesis + using inter_scheme [OF \Form l U'\ \l>0\] + proof cases + case (1 k' xs' ys') + 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') = k'" "length (b'#bs') = k'" + 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]: "k' = k" + using 1 one \l>0\ by (simp add: odd_eq_iffs) + 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.simps 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.simps acc_lengths_shift) + ultimately have "as' = as \ bs' = bs" + using acc_lengths_interact_injective by blast + with one 1 show ?thesis + by (simp add: xs ys xs' ys' \a' = a\ \b' = b\) + next + case (2 k' xs' ys') + then show ?thesis + using 1 \l>0\ by (simp add: odd_eq_iffs) + qed +next + case (2 k xs ys) + 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) = Suc k" "length (b#bs) = k" + 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 + note two = 2 + show ?thesis + using inter_scheme [OF \Form l U'\ \l>0\] + proof cases + case (1 k' xs' ys') + then show ?thesis + using 2 \l>0\ by (simp add: odd_eq_iffs) + next + case (2 k' xs' ys') + 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') = Suc k'" "length (b'#bs') = k'" + 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]: "k' = k" + using 2 two \l>0\ by (simp add: odd_eq_iffs) + 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.simps 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.simps acc_lengths_shift) + ultimately have "as' = as \ bs' = bs" + using acc_lengths_interact_injective by blast + with two 2 show ?thesis + by (simp add: xs ys xs' ys' \a' = a\ \b' = b\) + qed +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) + show ?case + proof (cases x) + case Nil + show ?thesis + proof (cases y) + case Nil + then show ?thesis + using "3" strict_sorted_append_iff by (auto simp: \x = []\) + next + case (Cons a list) + with Nil 3 show ?thesis + apply (simp add: strict_sorted_append_iff) + by (metis (no_types, lifting) Un_iff set_interact sorted_wrt_append strict_sorted_append_iff strict_sorted_sorted_wrt) + qed + next + case (Cons a list) + have \
: "sorted_wrt (<) ((a # list) @ y @ interact xs ys)" + by (metis (no_types) "3.prems" interact.simps(3) local.Cons strict_sorted_sorted_wrt) + then have "list = [] \ concat xs = [] \ last list < hd (concat xs)" + by (metis (full_types) Un_iff hd_in_set last_ConsR last_in_set list.simps(3) set_append set_interact sorted_wrt_append) + then have "list < concat xs" + using less_list_def by blast + have "list < y" + by (metis \
append.assoc last.simps less_list_def list.distinct(1) strict_sorted_append_iff strict_sorted_sorted_wrt) + note Cons1 = Cons + show ?thesis + proof (cases y) + case Nil + then show ?thesis + using 3 by (simp add: sorted_wrt_append strict_sorted_sorted_wrt) + next + case (Cons a' list') + have "strict_sorted (list' @ concat ys)" + apply (simp add: strict_sorted_sorted_wrt) + by (metis "3.IH" "\
" Un_iff append_Cons local.Cons set_interact sorted_wrt_append strict_sorted.simps(2) strict_sorted_sorted_wrt) + moreover have "y < concat ys" + by (metis "\
" Un_iff hd_in_set last_in_set less_list_def set_interact sorted_wrt_append) + ultimately show ?thesis + using 3 \list < concat xs\ + by (auto simp: Cons1 strict_sorted_append_iff) + qed + qed +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 Nil_is_append_conv hd_append2 hd_in_set interact.simps(3) list.exhaust_sel sorted_wrt_append strict_sorted_sorted_wrt) + + +text \the lengths of the two lists can differ by one\ +proposition interaction_scheme_unique_aux: + assumes eq: "concat as = concat as'" and ys': "concat bs = concat bs'" + and ne: "as \ lists (- {[]})" "bs \ lists (- {[]})" + and ss_zs: "strict_sorted (interact as bs)" + and "length bs \ length as" "length as \ Suc (length bs)" + and ne': "as' \ lists (- {[]})" "bs' \ lists (- {[]})" + and ss_zs': "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 have "length as = Suc 0" + using Suc.hyps(2) by auto + then obtain a a' where "as = [a]" "as' = [a']" + by (metis \length as = length as'\ length_0_conv length_Suc_conv) + with 0 show ?thesis + using Suc.prems + apply (simp add: le_Suc_eq) + by (metis concat.simps length_0_conv length_Suc_conv self_append_conv) + 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) Suc.prems(1) Suc.prems(3) interact_eq_Nil_iff by auto + then obtain a' cs' b' ds' where eq': "as' = a'#cs'" "bs' = b'#ds'" + by (metis Suc.prems(14) 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 "a < b" + by (metis eq Suc.prems(5) append.assoc interact.simps(3) strict_sorted_append_iff) + have sw_ab: "sorted_wrt (<) (a @ b @ interact cs ds)" + by (metis Suc.prems(5) eq interact.simps(3) strict_sorted_sorted_wrt) + 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) + have "b < concat cs" + using eq \strict_sorted (interact as bs)\ + apply (simp add: strict_sorted_append_iff) + by (metis Un_iff sw_ab last_in_set less_list_def list.set_sel(1) set_interact sorted_wrt_append) + have "strict_sorted (a @ concat cs)" + using eq(1) ss_ab(1) by force + then have b_cs: "strict_sorted (b @ concat cs)" + by (metis \b < concat cs\ strict_sorted_append_iff strict_sorted_sorted_wrt sw_ab) + 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 "a' < b'" + by (metis eq' Suc.prems(10) append.assoc interact.simps(3) strict_sorted_append_iff) + have sw_ab': "sorted_wrt (<) (a' @ b' @ interact cs' ds')" + by (metis Suc.prems(10) eq' interact.simps(3) strict_sorted_sorted_wrt) + 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) + have "b' < concat cs'" + using eq' \strict_sorted (interact as' bs')\ + apply (simp add: strict_sorted_append_iff) + by (metis Un_iff last_in_set less_list_def list.set_sel(1) 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 (metis Suc.prems(1) \hd b' = hd b\ sorted_wrt_append strict_sorted_equal strict_sorted_sorted_wrt sw_ab sw_ab') + 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 Suc.prems(14) Suc.prems(2) eq'(2) eq(2) by auto + next + case False + then have "ds \ []" "ds' \ []" + by auto + have "strict_sorted b" + by (metis Suc.prems(2) concat.simps(2) eq(2) ss_ab'(2) strict_sorted_append_iff) + 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 strict_sorted_append_iff strict_sorted_sorted_wrt sw_ab) + then have "list.set b = list.set (concat bs) \ {..< hd (concat cs)}" + using ss_ab + apply (auto simp: strict_sorted_append_iff eq) + apply (metis \b < concat cs\ \concat cs \ []\ hd_in_set sorted_wrt_append strict_sorted_append_iff strict_sorted_sorted_wrt) + by (metis strict_sorted_iff UN_I dual_order.strict_trans2 order.asym set_concat sorted_hd_le) + 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 strict_sorted_append_iff strict_sorted_sorted_wrt sw_ab') + then have "list.set b' = list.set (concat bs') \ {..< hd (concat cs')}" + using ss_ab' + apply (auto simp: strict_sorted_append_iff eq') + apply (meson strict_sorted_iff \b' < concat cs'\ \b' \ []\ \concat cs' \ []\ dual_order.strict_trans2 less_list_def sorted_le_last) + by (metis strict_sorted_iff UN_I dual_order.strict_trans2 order.asym set_concat sorted_hd_le) + + ultimately show "b = b'" + by (metis Suc.prems(2) ccat_cs_cs' strict_sorted_append_iff strict_sorted_equal strict_sorted_sorted_wrt sw_ab') + 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 + have "a < b" + using ss_zs apply (simp add: Ueq strict_sorted_append_iff) + by (metis strict_sorted_iff append.assoc d length_0_conv length_acc_lengths list.distinct(1) strict_sorted_append_iff sorted_trans) + have "a' < b'" + using ss_zs' apply (simp add: Ueq' strict_sorted_append_iff) + by (metis strict_sorted_iff append.assoc d' length_0_conv length_acc_lengths list.distinct(1) strict_sorted_append_iff sorted_trans) + have "a#as = a'#as' \ b#bs = b'#bs'" + proof (rule interaction_scheme_unique_aux) + show "concat (a # as) = concat (a' # as')" + using xs xs' by blast + show "concat (b # bs) = concat (b' # bs')" + using ys ys' by blast + show "a # as \ lists (- {[]})" "b # bs \ lists (- {[]})" + using ne by auto + show "strict_sorted (interact (a # as) (b # bs))" + using ss_zs \a < b\ apply (simp add: Ueq strict_sorted_append_iff) + by (metis strict_sorted_iff append.assoc append.left_neutral strict_sorted_append_iff sorted_trans) + 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 + show "a' # as' \ lists (- {[]})" "b' # bs' \ lists (- {[]})" + using ne' by auto + show "strict_sorted (interact (a' # as') (b' # bs'))" + using ss_zs' \a' < b'\ apply (simp add: Ueq' strict_sorted_append_iff) + by (metis strict_sorted_iff append.assoc append.left_neutral strict_sorted_append_iff sorted_trans) + show "length (a # as) = length (a' # as')" + using len'(1) len(1) by blast + show "length (b # bs) = length (b' # bs')" + using len'(2) len(2) by blast + qed + then show ?thesis + using Ueq Ueq' c c' d d' by blast +qed + + +lemma Form_Body_imp_inter_scheme: + assumes "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]) + using assms True mult_2 not_gr0 one_is_add apply fastforce + using assms \length xs < length ys\ + apply (auto simp: True mult_2 Set.doubleton_eq_iff Form_Body_unique dest: Form_Body_length) + by presburger + next + case False + then have eq: "ka = Suc kb" + using assms by linarith + show ?thesis + unfolding inter_scheme_def + apply (rule some_equality [symmetric]) + using assms False mult_2 one_is_add eq apply fastforce + using assms \length xs < length ys\ + apply (auto simp: eq mult_2 Set.doubleton_eq_iff Form_Body_unique dest: Form_Body_length) + by presburger + 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 add: grab_def enumerate_0 Least_le) + +lemma less_sets_grab: + "infinite N \ less_sets (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\ + \ less_sets 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: "less_sets A N \ less_sets 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 (odd k) + 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 xs ys where "U = {xs,ys}" + and U: "Form_Body k k xs ys (inter_scheme l U)" + using inter_scheme [OF \Form l U\ \l > 0\] + by (metis One_nat_def less_numeral_extra(3) odd odd_eq_iffs(1) odd_eq_iffs(3)) + 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) = k" "length (b#bs) = k" + 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.cases by metis + assume "Form l U'" + then obtain xs' ys' where "U' = {xs',ys'}" + and U': "Form_Body k k xs' ys' (inter_scheme l U')" + using inter_scheme [OF \Form l U'\ \l > 0\] + by (metis One_nat_def Suc_pred assms double_not_eq_Suc_double odd(1) odd_eq_iffs(1) odd_eq_iffs(3)) + 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') = k" "length (b'#bs') = k" + 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.cases by metis + have [simp]: "length bs = length as" "length bs' = 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 acc_lengths.simps) + have au1': "u1 = length a'" + by (simp add: u1_eq' Ueq' c' acc_lengths.simps) + have len_eqk: "length c = k" "length d = k" "length c' = k" "length d' = k" + using c d len c' d' len' by auto + have take: "take (k + u1 + k) (c @ a @ d @ l) = c @ a @ d" + "take (k + u1 + k) (c' @ a' @ d' @ l) = c' @ a' @ d'" for l + by (simp_all add: len_eqk flip: au1 au1') + have leU: "k + u1 + k \ length (inter_scheme l U)" + by (simp add: len_eqk au1 Ueq) + then have "take (k + u1 + k) (inter_scheme l U) = take (k + u1 + k) (inter_scheme l U')" + using take_initial_segment init by blast + then have \
: "c @ a @ d = c' @ a' @ d'" + by (metis append.assoc Ueq Ueq' concat.simps take) + have "length (inter_scheme l U) = k + (c @ a @ d)!(k-1) + k + last d" + by (simp add: Ueq c d length_interact nth_append len_eqk flip: len) + moreover + have "length (inter_scheme l U') = k + (c' @ a' @ d')!(k-1) + k + last d'" + by (simp add: Ueq' c' d' length_interact nth_append len_eqk flip: len') + moreover have "last d = last d'" + using "\
" au1 au1' len_eqk 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 +next + case (even k) + 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 xs ys where "U = {xs,ys}" + and U: "Form_Body (Suc k) k xs ys (inter_scheme l U)" + using inter_scheme [OF \Form l U\ \l > 0\] + by (metis One_nat_def Suc_1 Suc_mult_cancel1 even less_numeral_extra(3) odd_eq_iffs(3)) + 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) = Suc k" "length (b#bs) = k" + 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.cases by metis + assume "Form l U'" + then obtain xs' ys' where "U' = {xs',ys'}" + and U': "Form_Body (Suc k) k xs' ys' (inter_scheme l U')" + using inter_scheme [OF \Form l U'\ \l > 0\] + by (metis One_nat_def Suc_1 Suc_mult_cancel1 even not_gr0 odd_eq_iffs(3)) + 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') = Suc k" "length (b'#bs') = k" + 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.cases by metis + have [simp]: "length as = Suc (length bs)" "length as' = Suc (length bs')" + 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 acc_lengths.simps) + have au1': "u1 = length a'" + by (simp add: u1_eq' Ueq' c' acc_lengths.simps) + have len_eqk: "length c = Suc k" "length d = k" "length c' = Suc k" "length d' = k" + using c d len c' d' len' by auto + have take: "take (Suc k + u1 + k) (c @ a @ d @ l) = c @ a @ d" + "take (Suc k + u1 + k) (c' @ a' @ d' @ l) = c' @ a' @ d'" for l + by (simp_all add: len_eqk flip: au1 au1') + have leU: "Suc k + u1 + k \ length (inter_scheme l U)" + by (simp add: len_eqk au1 Ueq) + then have "take (Suc k + u1 + k) (inter_scheme l U) = take (Suc k + u1 + k) (inter_scheme l U')" + using take_initial_segment init by blast + then have \
: "c @ a @ d = c' @ a' @ d'" + by (metis append.assoc Ueq Ueq' concat.simps take) + have "length (inter_scheme l U) = Suc k + (c @ a @ d)!k + k + last d" + by (simp add: Ueq c d length_interact nth_append len_eqk flip: len) + moreover + have "length (inter_scheme l U') = Suc k + (c' @ a' @ d')!k + k + last d'" + by (simp add: Ueq' c' d' length_interact nth_append len_eqk flip: len') + moreover have "last d = last d'" + using "\
" au1 au1' len_eqk 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 + assumes g: "g \ [WW]\<^bsup>2\<^esup> \ {..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)" + { fix l m::nat and M :: "nat set" + assume "l > 0" "\ m M" + let ?A = "inter_scheme l ` {U \ [WW]\<^bsup>2\<^esup>. Form l U}" + define h where "h \ \zs. g (inv_into {U \ [WW]\<^bsup>2\<^esup>. Form l U} (inter_scheme l) zs)" + have "thin ?A" + using \l > 0\ lemma_3_11 by (simp add: thin_def) + moreover + have "?A \ WW" + using inter_scheme_simple \0 < l\ by blast + moreover + have "h ` {l \ ?A. List.set l \ M} \ {..<2}" + using g inv_into_into[of concl: "{U \ [WW]\<^bsup>2\<^esup>. Form l U}" "inter_scheme l"] + by (force simp: h_def Pi_iff) + ultimately + obtain j N where "j < 2" "infinite N" "N \ M" and hj: "h ` {l \ ?A. List.set l \ N} \ {j}" + using \\ m M\ unfolding \_def by (blast intro: Nash_Williams_WW [of M]) + define n where "n \ Inf N" + have "n > m" + using \\ m M\ \infinite N\ unfolding n_def \_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) {U \ [WW]\<^bsup>2\<^esup>. Form l U}" + by (metis (mono_tags, lifting) inter_scheme_injective \0 < l\ inj_onI mem_Collect_eq) + moreover have "g (inv_into {U \ [WW]\<^bsup>2\<^esup>. Form l U} (inter_scheme l) (inter_scheme l U)) = j" + using hj that xys + apply (simp add: h_def image_subset_iff image_iff) + by (metis (no_types, lifting) Diff_subset doubleton_in_nsets_2 dual_order.trans) + ultimately show ?thesis + using that by auto + qed + moreover have "n < Inf (N - {n})" + unfolding n_def + 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 n_def 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 have "\n N j. \ n N \ \ l m n M N j" + by blast + } note * = this + 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 + proof (induction l) + case 0 + with base show ?case + by (auto simp: H_simps) + next + case (Suc l) + with G\ show ?case + by (force simp: H_simps split: prod.split prod.split_asm) + qed + define \ where "\ \ (\l. case H l of (n,M,j) \ n)" + have H_inc: "\ l \ l" for l + proof (induction l) + case 0 + then show ?case + by auto + next + case (Suc l) + then show ?case + using H\ G_increasing [of "\ l"] + apply (auto simp: H_simps \_def split: prod.split prod.split_asm) + by (metis (mono_tags, lifting) Suc_leI Suc_le_mono case_prod_conv dual_order.trans) + qed + 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 arbitrary:) + 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 "U \ WW" + using U by (auto simp: nsets_def) + moreover + 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 add: k n enum_N H_simps) + then show ?thesis + using that interU \U \ WW\ G\ [of m M k] H\ [of k] + by (auto simp: \_def k enum_N H_simps n) + qed + ultimately 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" + unfolding numeral_nat + apply (clarsimp simp add: nsets_def card_Suc_eq Set.doubleton_eq_iff not_less) + by (metis tot total_on_def) + 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 0 + then show ?case + by auto +next + case (Suc n I) + then have [simp]: "I \ {}" + by auto + moreover have "sum_list (map f (list_of (I - {Min I}))) = sum f (I - {Min I})" + using Suc by auto + ultimately show ?case + using Suc.prems sum.remove [of I "Min I" f] + by (simp add: sorted_list_of_set_nonempty Suc) +qed + + +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 0 + then have "I={}" by auto + then show ?case by auto +next + 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 + by (metis DiffE Min_in \I \ {}\ card_Diff_singleton diff_Suc_1 finite_Diff strict_mono_sets_def) + 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 "less_sets (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 + +subsubsection \Lemma 3.7 of Jean A. Larson, ibid.\ + +text \Possibly should be redone using grab\ +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_nxtN [OF \infinite N\, iff] + + have [simp]: "{n<.. \k. rec_nat ((enum N \ Suc) ` {..r D. enum (nxt N (enum (nxt N (Max D)) (Inf D - Suc 0))) ` {.. Suc) ` {.. Suc) {..infinite N\ DF_simps 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 DF_N: "DF k i \ N \ {0<..}" for i k + proof (induction i) + case 0 + then show ?case + using \infinite N\ range_enum nonz by (auto simp: DF_simps) + next + case (Suc i) + then show ?case + unfolding DF_simps image_subset_iff + using infinite_nxtN [OF \infinite N\] + by (metis Int_iff enumerate_in_set greaterThan_iff not_gr0 not_less0 nxt_def) + qed + then have DF_gt0: "0 < Inf (DF k i)" for i k + using DF_ne Inf_nat_def1 by blast + have finite_DF: "finite (DF k i)" for i k + by (induction i) (auto simp: DF_simps) + + have sm_enum_DF: "strict_mono_on (enum (DF k i)) {..k}" for k i + by (metis card_DF enum_works_finite finite_DF lessThan_Suc_atMost) + + have DF_Suc: "less_sets (DF k i) (DF k (Suc i))" for i k + unfolding less_sets_def + by (force simp: finite_DF DF_simps + intro!: greaterThan_less_enum nxt_subset_greaterThan atLeast_le_enum nxt_subset_atLeast infinite_nxtN [OF \infinite N\]) + have DF_DF: "less_sets (DF k i) (DF k j)" if "i \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 \infinite N\ card_image inj_enum_nxt) + + have DF_AF: "less_sets (DF k i) (AF k i)" for i k + unfolding less_sets_def AF_def + by (simp add: finite_DF \infinite N\ greaterThan_less_enum nxt_subset_greaterThan) + + have AF_DF_Suc: "less_sets (AF k i) (DF k (Suc i))" for i k + apply (clarsimp simp: DF_simps less_sets_def AF_def) + using strict_monoD [OF strict_mono_enum] + by (metis DF_gt0 Suc_pred assms(1) dual_order.order_iff_strict greaterThan_less_enum + infinite_nxtN linorder_neqE_nat not_less_eq nxt_subset_greaterThan) + + have AF_DF: "less_sets (AF k p) (DF k q)" if "p \k i j. enum (DF k i) j - enum (DF k i) (j - Suc 0)" + + 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 - Suc 0))))) ` + {..< 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)"]) + 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: "less_sets (QF k (j, m - Suc 0)) (QF k (Suc j,0))" for j k + by (auto simp: def_wfrec [OF QF_def, of k "(Suc j,0)"] pair_lessI1 \infinite N\ enum_nxt_ge + intro!: less_sets_weaken2 [OF less_sets_Suc_Max]) + + have less_QF2: "less_sets (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 \infinite N\ enum_nxt_ge + intro: less_sets_weaken2 [OF less_sets_Suc_Max] strict_mono_setsD [OF sm_AF]) + + have less_QF_same: "less_sets (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 [of "\i. QF k (j,i)"]]) + show "less_sets (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: "less_sets (QF k (j - Suc 0, i')) (QF k (j,i))" + if "0 < j" "j \ k" "i' < m" for j i' i k + proof - + have less_QF1': "less_sets (QF k (j - Suc 0, m - Suc 0)) (QF k (j,0))" if "j > 0" for j + by (metis less_QF1 that Suc_pred) + have \
: "less_sets (QF k (j - Suc 0, i')) (QF k (j,0))" + proof (cases "i' = m - Suc 0") + 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: "less_sets (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 0 + then show ?case + by auto + next + case (Suc d) + then have eq: "d = (j - Suc 0) - j'" + by linarith + show ?case + proof (cases "j' < j - Suc 0") + case True + then have "less_sets (QF k (j', i')) (QF k (j - Suc 0, i))" + using Suc eq by auto + 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 - Suc 0" + using \j' < j\ by linarith + then show ?thesis + using Suc.hyps \j \ k\ less_QF_step i by auto + qed + qed + + 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 "less_sets (QF k p) (QF k q)" + proof (elim conjE disjE) + assume "j' < j" + show "less_sets (QF k p) (QF k q)" + by (simp add: \
\j' < j\ less_QF that) + qed (use \
in \simp add: that less_QF_same\) + 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 \infinite N\ inj_enum_nxt) + next + case (Suc j') + show ?thesis + by (cases i; simp add: Suc QF_Suc QF_Suc_Suc card_image \infinite N\ 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 0 + with AF_subset_N show ?case + by auto + next + case (Suc j) + show ?case + by (cases i) (use nxt_subset enumerate_in_set in \(force simp: QF_Suc QF_Suc_Suc)+\) + qed + + obtain ka k where "k>0" and kka: "k \ ka" "ka \ Suc k" "l = ((ka+k) - Suc 0)" + proof - + consider (odd) k where "l = 2*k - Suc 0" "k > 0" | (even) k where "l = 2*k" "k > 0" + by (metis One_nat_def \l > 0\ form_cases not_gr0) + then show thesis + proof cases + case odd + then show ?thesis + by (metis lessI less_or_eq_imp_le mult_2 that) + next + case even + then show ?thesis + by (metis add_Suc diff_Suc_Suc lessI less_or_eq_imp_le minus_nat.diff_0 mult_2 that) + qed + qed + 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 [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 \infinite N\ 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] + apply (simp add: strict_mono_on_def) + by (metis Suc_leD dual_order.order_iff_strict le0) + 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 + using DF_N [of k i] card_DF finite_enumerate_in_set finite_DF that + by (metis inf.boundedE le_imp_less_Suc subsetD) + 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 - Suc 0 \ 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 = Suc 0" if "k>0" for k::nat + proof - + have "list_of {0<.. {0<.. \ k = Suc 0" + 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: "less_sets (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: "less_sets (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 - Suc 0) < hd (list_of (AF k p))" for p + proof (rule less_setsD [OF DF_AF]) + show "enum (DF k p) (ka - Suc 0) \ 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: "less_sets (RF n p) (RF n q)" if "n < k" for n + using that \p by (simp add: less_RF_same) + have less_RF_Suc: "less_sets (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: "less_sets (RF k q) (RF k p)" + using \q < m\ less_RF_same_k \p by blast + have less_RF_k_ka: "less_sets (RF (k - Suc 0) p) (RF (ka - Suc 0) q)" + using ka_k_or_Suc less_RF_RF + by (metis One_nat_def RF_def \0 < k\ \ka - Suc 0 \ 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 + 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 + 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 [OF \length x < length y\]) + 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 "less_sets (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\) + ultimately show "strict_sorted zs" + using \p < q\ True Inf_DF_less DF_AF DF_ne + apply (auto simp: zs_def less_sets_def card_AF AF_Inf_DF_less) + by (meson Inf_nat_def1) + qed (auto simp: \k=1\ \ka=1\ acc_lengths.simps zs_def AF_ne) + 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\) + show "list.set (inter_scheme l U) \ N" + using kka zs zs_N \k=1\ Form_Body_imp_inter_scheme by (fastforce simp add: \U = {x,y}\) + qed + next + case 2 + note True [simp] note 2 [simp] + have [simp]: "{0<..<2} = {Suc 0}" + by auto + have enum_DF1_eq: "enum (DF (Suc 0) i) (Suc 0) = card (AF (Suc 0) i) + card (RF (Suc 0) i)" + if "i < m" for i + using card_AF_sum that by simp + have card_RF: "card (RF (Suc 0) i) = enum (DF (Suc 0) i) (Suc 0) - enum (DF (Suc 0) 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 (Suc 0) q \ RF (Suc 0) q) = list_of (AF (Suc 0) q) @ list_of (RF (Suc 0) q)" + using RF_def \q < m\ less_QF_step by (fastforce intro!: 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 [OF \length x < length y\]) + have "x = list_of (RF 0 p \ RF (Suc 0) p)" + by (simp add: x eval_nat_numeral lessThan_Suc RF_0 Un_commute) + also have "\ = list_of (RF 0 p) @ list_of (RF (Suc 0) p)" + using RF_def True \p < m\ less_QF_step + by (subst sorted_list_of_set_Un) (fastforce+) + 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) + show "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 (Suc 0) p)" + using \p by (simp add: card_RF card_DF finite_DF) + show "card (AF (Suc 0) p) < card (AF (Suc 0) q) + card (RF (Suc 0) q)" + using \p \q by (simp add: Inf_DF_less card_AF trans_less_add1) + show "card (AF (Suc 0) p) < x" + if "x \ AF (Suc 0) p \ (AF (Suc 0) q \ (RF (Suc 0) q \ RF (Suc 0) 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 (Suc 0) p) + card (RF (Suc 0) p) < card (AF (Suc 0) q) + card (RF (Suc 0) q)" + using \p < q\ \p < m\ \q < m\ by (metis enum_DF1_eq enum_DF_less_iff le_refl) + show "card (AF (Suc 0) p) + card (RF (Suc 0) p) < x" + if "x \ AF (Suc 0) p \ (AF (Suc 0) q \ (RF (Suc 0) q \ RF (Suc 0) p))" for x + using that \p < m\ + apply (simp add: 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 (Suc 0) p) < list_of {enum (DF (Suc 0) q) (Suc 0)}" + proof (rule less_sets_imp_sorted_list_of_set) + show "less_sets (AF (Suc 0) p) {enum (DF (Suc 0) q) (Suc 0)}" + by (metis (no_types, lifting) AF_DF DF_ne Inf_nat_def1 \q < m\ card_AF enum_DF1_eq less_setsD less_sets_singleton2 pq(1) trans_less_add1) + qed auto + then show "list_of (AF (Suc 0) p) < (card (AF (Suc 0) q) + card (RF (Suc 0) q)) # list_of (AF (Suc 0) q) @ list_of (RF (Suc 0) q) @ list_of (RF (Suc 0) p)" + using \q < m\ by (simp add: less_list_def enum_DF1_eq) + show "card (AF (Suc 0) q) + card (RF (Suc 0) q) < x" + if "x \ AF (Suc 0) q \ (RF (Suc 0) q \ RF (Suc 0) 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 (Suc 0) q) < list_of (RF (Suc 0) q)" + proof (rule less_sets_imp_sorted_list_of_set) + show "less_sets (AF (Suc 0) q) (RF (Suc 0) q)" + by (metis less_RF_Suc One_nat_def RF_0 True \0 < k\) + qed auto + then show "list_of (AF (Suc 0) q) < list_of (RF (Suc 0) q) @ list_of (RF (Suc 0) p)" + using RF_non_Nil by (auto simp: less_list_def) + show "list_of (RF (Suc 0) q) < list_of (RF (Suc 0) p)" + proof (rule less_sets_imp_sorted_list_of_set) + show "less_sets (RF (Suc 0) q) (RF (Suc 0) p)" + by (metis less_RF_k One_nat_def 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: acc_lengths.simps 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: acc_lengths.simps zs_def AF_ne sup_commute) + qed (auto simp: acc_lengths.simps zs_def AF_ne) + 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 add: \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 + + 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 - Suc (Suc 0) then list_of (RF (Suc n) q) + else list_of (RF (k - Suc 0) q \ RF (ka - Suc 0) 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 - Suc 0" "k=ka" for n + using that kka Suc_diff_Suc + by (fastforce simp add: QQ_def nth_append nth_sorted_list_of_set_greaterThanLessThan) + + have split_nat_interval: "{0<.. 2" for n + using that by auto + have split_list_interval: "list_of{0<.. 2" for n + 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 - Suc 0) q \ RF k q) = list_of (RF (k - Suc 0) q) @ list_of (RF k q)" + by (metis less_RF_Suc Suc_pred \0 < k\ diff_Suc_less finite_RF 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 - Suc 0"] \k > 0\ by (auto simp: less_sets_def) + then have "card (RF (k - Suc 0) q \ RF k q) = card (RF (k - Suc 0) 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.simps 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 Int_subset_iff 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 Int_subset_iff 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 - Suc 0 \ 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.simps 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 - Suc (Suc 0)" if "\ n < k - Suc (Suc 0)" + using n that by linarith + with \p n show "PP ! n < QQ ! n" + using \0 < k\ \k \ ka\ \ka \ Suc k\ \p n + by (auto simp: PP_n QQ_n less_RF_same less_sets_imp_sorted_list_of_set less_sets_Un2 less_RF_RF less_RF_k_ka) + next + fix n + have V: "\Suc n < ka - Suc 0\ \ list_of (RF (Suc n) q) < list_of (RF (Suc (Suc n)) p)" for n + by (smt One_nat_def RF_def Suc_leI \ka - Suc 0 \ 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 "less_sets (RF (k - Suc 0) q) (RF k p)" + by (metis less_RF_Suc less_RF_k RF_non_Nil Suc_pred \0 < k\ finite_RF lessI less_sets_trans sorted_list_of_set_eq_Nil_iff) + with kka have "less_sets (RF (k - Suc 0) q \ RF (ka - Suc 0) 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 - Suc 0) q \ RF (ka - Suc 0) 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 (auto simp: PP_n QQ_n less_RF_same less_sets_imp_sorted_list_of_set less_sets_Un1 less_sets_Un2 less_RF_RF less_RF_k_ka V) + by (metis One_nat_def Suc_less_eq Suc_pred \0 < k\ diff_Suc_1 k_minus_1 ka_k_or_Suc less_SucE) + 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) + 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 - Suc 0 \ k\ add.left_neutral finite_RF less_le_trans less_sets_imp_sorted_list_of_set nth_sorted_list_of_set_greaterThanLessThan) + next + show "?YR \ 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) + 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 "less_sets (AF k q) (RF (Suc 0) 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 (Suc 0) 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 (Suc 0) p)) \ RF (Suc 0) p" + by (metis RF_non_Nil Suc_mono \0 < k\ finite_RF hd_in_set set_sorted_list_of_set) + qed + with \k > 0\ \ka \ 2\ RF_non_Nil show ?thesis + by (simp add: 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 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 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 + show ?thesis + proof (cases u) + case (Pair j' i') + then show ?thesis + using assms None by (simp add: prev_eq_None_iff pair_less_def IJ_def) + qed +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 \ \ \ less_sets 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 "I - {j} \ \" if "I \ \" "j \ J I" for I j + using that by (metis Diff_empty Diff_insert0 assms less_irrefl less_sets_def) + moreover have "j \ J (I - {j})" if "I \ \" "j \ J I" for I j + using that by (metis Diff_empty Diff_insert0 assms less_irrefl less_setsD) + moreover 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) + ultimately show ?thesis + by (auto simp: USigma_def usplit_def) + qed +qed + + +lemma 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 \ \ \ less_sets 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 "\ = {}") +next + 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 Ord_mem_iff_less_TC TC_small add_mult_less image_eqI oB 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, hide_lams) 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\ + +definition cconcat where [simp]: "cconcat x l \ if x=[] then l else concat x#l" + +inductive merge :: "[nat list list,nat list list,nat list list,nat list list] \ bool" + where Null: "merge as [] (cconcat as []) []" + | App: "\as1 \ []; bs1 \ []; + concat as1 < concat bs1; concat bs1 < concat as2; merge as2 bs2 as bs\ + \ merge (as1@as2) (bs1@bs2) (cconcat as1 as) (cconcat 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 add: 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 + by simp (metis One_nat_def Suc_eq_plus1 Suc_leI add.commute add_mono length_greater_0_conv) +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 "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" + using App + apply (simp add: less_list_def del: concat_eq_Nil_conv) + by (metis (full_types) Un_iff \concat bs1 < concat as\ \concat bs1 < concat bs\ last_in_set list.set_sel(1) set_interact sorted_wrt_append strict_sorted_append_iff strict_sorted_interact_imp_concat strict_sorted_sorted_wrt xx) + 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 acc_lengths.simps strict_sorted_append_iff length_concat_acc_lengths) + by (simp add: le_supI2 length_concat) +qed (auto simp: acc_lengths.simps 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 acc_lengths.simps strict_sorted_append_iff length_concat_acc_lengths) + by (simp add: le_supI2 length_concat) +qed (auto simp: acc_lengths.simps 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 "less_sets (list.set (hd us)) (list.set (concat vs))" + using assms +proof induction + case (App as1 bs1 as2 bs2 as bs) + then have \
: "less_sets (list.set (concat bs1)) (list.set (concat bs2))" + by (force simp: dest: strict_sorted_imp_less_sets)+ + have *: "less_sets (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 "less_sets (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_concat_ge length_greater_0_conv list.size(3) mem_lists_non_Nil) + with * App.hyps show ?case + by (fastforce simp add: 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: "less_sets (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 + apply (clarsimp simp add: as1_def less.prems set_takeWhile less_list_iff_less_sets less_sets_def) + by (smt UN_I dual_order.strict_trans2 hd_concat less.prems(2) less.prems(4) less.prems(7) list.set_sel(1) mem_lists_non_Nil not_le 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 "less_sets (list.set (concat as1)) (list.set (concat bs1))" + using AB set_takeWhileD by (fastforce simp add: as1_def bs1_def less_sets_UN1 less_sets_UN2) + then show "concat as1 < concat bs1" + by (rule less_sets_imp_list_less) + have "less_sets (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 "less_sets (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 (metis UN_I \as = as1 @ as2\ concat.simps(2) concat_append dual_order.strict_trans2 hd_concat in_set_conv_decomp_last not_le set_concat sorted_hd_le sorted_le_last sorted_sorted_wrt sorted_wrt_append 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 "bs2 = []") + case True + show ?thesis + proof + show "merge as2 bs2 (cconcat as2 []) (cconcat bs2 [])" + by (simp add: True) + qed + 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) (cconcat as1 cs) (cconcat bs1 ds)" + using A merge.App by blast + then show ?case + using B by blast +qed + +subsubsection \Actual proof of lemma 3.8\ + +text \Lemma 3.8 of Jean A. Larson, ibid.\ +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_bf [simp]: "infinite (snd (bf M q u)) = infinite 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 + by (auto simp: bf_rec [of M q j i] split: option.split) + qed + qed + + have bf_less_sets: "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 bf_subset: "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) + show ?case + proof (cases u) + case (Pair j i) + then show ?thesis + apply (simp add: bf_rec [of M q j i] that split: option.split) + using fst_grab_subset less.IH prev_pair_less snd_grab_subset by blast + 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 :: nat \ nat \ nat set, M::nat set). \l::nat. + less_sets (dl l) a \ finite a \ dl l \ {} \ a \ {} \ + (\j\l. card (dl j) = Suc j) \ less_sets a (\(range b)) \ range b \ Collect finite \ + a \ N \ \(range b) \ N \ infinite M \ less_sets (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 Min_in card_eq_0_iff greaterThan_iff le_inf_iff less_nat_zero_code n_not_Suc_n nxt_def 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) + apply (meson less_sets_weaken2) + apply (metis (no_types, hide_lams) IntE Min_in card_empty greaterThan_iff leD not_less_eq_eq nxt_def subsetD zero_less_Suc) + 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" for k k' + using that + by (induction "k-k'")(auto simp: Inf_M_Suc_ge M_Suc_subset cInf_superset_mono infinite_imp_nonempty lift_Suc_antimono_le) + + 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 split: prod.split_asm) + + have da: "less_sets (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: "less_sets (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 1 + then show ?thesis + by (simp add: Pair bf_rec snd_grab_subset) + next + case 2 + then have "snd (bf M r x) \ snd (bf M r (j', i'))" + by (simp add: Pair less.IH prev_pair_less that(2)) + moreover have "snd (bf M r u) \ snd (bf M r x)" + by (simp add: 2 Pair bf_rec snd_grab_subset) + ultimately show ?thesis + by auto + qed + qed + qed + + have less_bf: "less_sets (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 infinite_bf less_sets_fst_grab \infinite M\ by auto + next + case 2 + then have "less_sets (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 infinite_bf less_sets_fst_grab \infinite M\ by auto + qed + qed + + have aM: "less_sets (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 "less_sets (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: "less_sets (a j) (a k)" if "jk" for k k' j i + by (metis a_ne ab_same le_less less_sets_UN2 less_sets_trans rangeI that) + have db: "less_sets (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: "less_sets (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 add: \_def b_def) + + have bM: "less_sets (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 "less_sets (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 :: "nat set" + 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::nat)" + using Inf_nat_def x le_less_linear not_less_Least by fastforce + qed + + have b_Inf_M_Suc: "less_sets (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: "less_sets (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: "less_sets (b k' (j',i')) (b k (j,i))" + if j: "i' < j'" "j'\k'" and k: "k' {\(M k)..}" + by (rule order_trans [OF _ b_InfM]) auto + also have "\ \ {Inf(M (Suc k'))..}" + using Inf_M_telescoping k by auto + finally show "b k (j,i) \ {Inf(M (Suc 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: "K \ \ j0 j \ finite K" for K j0 j + by (simp add: \_def nsets_def) + have \_card: "K \ \ j0 j \ card K = j" for K j0 j + by (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 obtain i where "Max (insert j0 (K - {Max K})) < i" "K = insert i (K - {Max K})" + apply (simp add: subset_iff) + by (metis Diff_iff Max.coboundedI Max_in card_0_eq insert_Diff insert_iff le_neq_implies_less nat.distinct(1)) + then show "\L\{j0<..}. finite L \ card L = j \ + (\i\{Max (insert j0 L)<..}. K = insert i L)" + using K + by (metis Max_in card_Diff_singleton_if card_gt_0_iff diff_Suc_1 finite_Diff greaterThan_iff insert_subset zero_less_Suc) + 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 \ less_sets (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 "less_sets (a j \ (\i_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) + show "enum K j \ K" + by (simp add: cardK finite_enumerate_in_set) + show "k \ enum K j" if "k \ K" for k + using that K + by (metis \finite K\ cardK enum_obtain_index_finite finite_enumerate_mono leI less_Suc_eq less_asym) + qed auto + 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 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 "less_sets (a j0) (b (enum K j) (j0, j))" + using MaxK \j0 < Max K\ ab by auto + ultimately show "less_sets (a j0 \ (\x = 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" if "k>0" for k + using that + proof (induction k) + case (Suc k) + show ?case + proof (cases "k=0") + case False + have "\ (b k (k, k - Suc 0)) < \ (b (Suc k) (Suc k, k))" + using False Inf_b_less by auto + with False Suc show ?thesis + by simp + qed auto + qed auto + + have b_ge: "\ (b k (j,i)) \ k-1" if k: "k>0" "k \ j" and "j > i" for k j i + using k + proof (induction k) + case (Suc k) + show ?case + proof (cases "j \ k") + case True + have "\ (b k (j,i)) < \ (b (Suc k) (j,i))" + using \j > i\ Suc True by (force intro: Inf_b_less) + then show ?thesis + using Suc.IH True by linarith + next + case False + then have "j = Suc k" + using Suc.prems(2) by linarith + with \i < j\ have "i < Suc k" + by fastforce + moreover have "\ \ (b (Suc k) (j,i)) < \ (b (Suc k) (j,i))" + by fastforce + ultimately have "\ Suc (\ (b (Suc k) (j,i))) < Suc k" + by (metis Inf_b_less \j = Suc k\ b_ge_k diff_Suc_1 leD le_refl lessI zero_less_Suc) + then show ?thesis + by simp + qed + qed auto + + 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 [of _ j i] \j > i\ Sucj + by (metis (no_types, lifting) diff_Suc_1 le_SucI le_trans less_Suc_eq_le nat_le_linear zero_less_Suc) + 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 add: 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) + 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 subset_iff sorted_list_of_set_Un less_sets_UN2) + fix X Y + assume X [rule_format]: "\t. t \ X \ j0 < t" + and Y [rule_format]: "\t. t \ Y \ j0 < t" + and "finite X" + and jeq: "j = card X" + and "finite Y" + 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" + by (simp add: Y \card Y = card X\ \finite Y\ finite_enumerate_in_set 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 (rule 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 "less_sets L {Max (insert j0 L)<..} \ finite L" + by (metis \_finite atLeast_Suc_greaterThan finite_insert less_sets_Suc_Max less_sets_weaken1 subset_insertI) + 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 (metis ih One_nat_def Ord_\ Ord_ord_of_nat oexp_1_right oexp_add one_V_def ord_of_nat.simps(1) ord_of_nat.simps(2) ord_of_nat_add plus_1_eq_Suc) + 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: "less_sets (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(1) that(2)) + have ba: "less_sets (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 + proof (cases "k < r") + case True + then show ?thesis + by (simp add: ba less_sets_imp_disjnt that) + next + case False + then show ?thesis + proof - + have "less_sets (a r) (b k (j,i))" + by (metis False a_ne aa ab_same less_linear less_sets_UN2 less_sets_trans rangeI) + then show ?thesis + using disjnt_sym less_sets_imp_disjnt by blast + qed + qed + + 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) + have dis: "disjnt (b (enum K j) (j0, j)) (\ii < j\ b_disjoint_less disjnt_def disjnt_sym less_Suc_eq that) + qed + have j0_less: "j0 < enum K j" + by (meson Suc.prems Suc_le_lessD finite_enumerate_in_set greaterThan_iff less_le_trans subsetD K) + 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 + apply (simp add: card_b split: nat_diff_split) + by (metis Suc.prems card_d finite_d finite_enumerate_step le_imp_less_Suc less_asym) + 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 "less_sets (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: "less_sets (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: "less_sets (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 acc_lengths.simps Inf_nat_def1 card_a) + next + case (Suc j) + let ?db = "\ (d j0) + ((\i 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 acc_lengths.simps card_a subset_insertI) + ultimately show ?case + by (simp add: seqs_def acc_lengths.simps 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: "less_sets (b (enum K i) (j,i)) (b (enum L p) (r, p))" + if neg: "\ less_sets (b (enum L p) (r, p)) (b (enum K i) (j,i))" and "i < j" "p < r" + for i p + proof (cases "enum K i" "enum L p" rule: linorder_cases) + case less + then show ?thesis + by (simp add: bb enumK less_imp_le_nat \i < j\) + next + case equal + then show ?thesis + using \j < r\ enumK \i < j\ by (force simp: IJ_iff pair_less_def intro: bb_same) + next + case greater + then show ?thesis + using bb enumL less_imp_le_nat neg \p < r\ by blast + qed + 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 + 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 "less_sets (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 dual_order.refl less_sets_imp_list_less less_sets_weaken1) + then have "ps < u1" + by (metis K Ksub UnI1 \_card \finite K\ \j \ card K\ \less_sets (d j) (list.set x)\ acc_lengths_subset_d concat.simps(2) empty_iff empty_set hd_append2 less_list_def less_sets_imp_list_less less_sets_weaken1 list.set_sel(1) set_append u_sub xu_eq) + + have "less_sets (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 dual_order.refl less_sets_imp_list_less less_sets_weaken1) + then have "qs < v1" + by (metis L Lsub UnI1 \_card \finite L\ \r \ card L\ \less_sets (d r) (list.set y)\ acc_lengths_subset_d concat.simps(2) empty_iff empty_set hd_append2 less_list_def less_sets_imp_list_less less_sets_weaken1 list.set_sel(1) set_append v_sub 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 \ less_sets (a j) (b (enum K i) (j,i))" + by (meson ab enumK le_trans less_imp_le_nat) + + have ab_enumL: "\q. q < r \ less_sets (a j) (b (enum L q) (r,q))" + by (meson \j < r\ ab enumL le_trans less_imp_le_nat) + then have ay: "less_sets (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\ + apply (simp add: IJ_def pair_less_def less_sets_def) + by (metis enumK b_ne card_b_finite last_in_set leD less_imp_le_nat 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: "less_sets (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 "less_sets (b k (j,i)) (d (Suc k))" if "j\k" "ik" "i 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 "Min (b (enum K i) (j,i)) \ n" + using 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] u1_y + apply (simp add: less_list_def disjnt_iff) + by (metis card_b_finite hd_list_of leI less_imp_le_nat 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 acc_lengths.simps 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 "less_sets (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 "less_sets (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 (metis Min_in card_0_eq card_d d_eq finite_d fst_grab_subset greaterThan_iff in_mono le_inf_iff nxt_def old.nat.distinct(2)) + finally have ls: "less_sets {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 (no_types, hide_lams) One_nat_def Suc_1 WW_seg_subset_WW order.trans less_2_cases m not1 nth_Cons' nth_Cons_Suc) + 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 + by (metis TC_small UNIV_I UN_I otW' lenlex_transI ordertype_mono subsetI trans_less_than wf_lenlex wf_less_than) + 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 + with \1 < m\ show ?thesis + apply (simp add: h_def len_h_seg \l \ WW\) + by (meson \0 < length l\ bij_betw_inv_into_left bij_h_seg l) + 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 add: 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 have "(a, b) \ lenlex less_than" + using omega_sum_less_iff by auto + then show ?thesis + by (simp add: \a \ WW\ \b \ WW\) + 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') + then have fh: "?fh \ [WW]\<^bsup>2\<^esup> \ {.. 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 "f (h ` Y) = 0" "finite Y" "card Y = 2" "Y \ M" + apply (simp add: nsets_def subset_image_iff) + by (metis \M \ WW\ \bij_betw h WW WW'\ bij_betw_def card_image card_infinite inj_on_subset zero_neq_numeral) + then 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>\) + 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 + moreover + have "[enum N k] < inter_scheme k u" + using inter_scheme [of k u] strict_mono_enum [OF \infinite N\] \
+ apply (auto simp: less_list_def subset_image_iff subset_eq Bex_def image_iff) + by (metis hd_in_set strict_mono_def) + 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 False + then have "l = 0" + by blast + 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 True + 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 True 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/Ordinal_Partitions/Partitions.thy b/thys/Ordinal_Partitions/Partitions.thy new file mode 100644 --- /dev/null +++ b/thys/Ordinal_Partitions/Partitions.thy @@ -0,0 +1,938 @@ +section \Ordinal Partitions\ + +text \Material 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. +Also from ``Partition Relations'' by A. Hajnal and J. A. Larson, +in \emph{Handbook of Set Theory}, edited by Matthew Foreman and Akihiro Kanamori +(Springer, 2010).\ + +theory Partitions + imports Library_Additions "ZFC_in_HOL.ZFC_Typeclasses" "ZFC_in_HOL.Cantor_NF" + +begin + +abbreviation tp :: "V set \ V" + where "tp A \ ordertype A VWF" + +subsection \Ordinal Partitions: Definitions\ + +definition partn_lst :: "[('a \ 'a) set, 'a set, V list, nat] \ bool" + where "partn_lst r B \ n \ \f \ nsets B n \ {..}. + \i < length \. \H. H \ B \ ordertype H r = (\!i) \ f ` (nsets H n) \ {i}" + +abbreviation partn_lst_VWF :: "V \ V list \ nat \ bool" + where "partn_lst_VWF \ \ partn_lst VWF (elts \)" + +lemma partn_lst_E: + assumes "partn_lst r B \ n" "f \ nsets B n \ {.." + obtains i H where "i < l" "H \ B" + "ordertype H r = \!i" "f ` (nsets H n) \ {i}" + using assms by (auto simp: partn_lst_def) + +lemma partn_lst_VWF_nontriv: + assumes "partn_lst_VWF \ \ n" "l = length \" "Ord \" "l > 0" + obtains i where "i < l" "\!i \ \" +proof - + have "{.. {}" + by (simp add: \l > 0\ lessThan_empty_iff) + then obtain f where "f \ nsets (elts \) n \ {.. elts \" and eq: "tp H = \!i" + using assms by (metis partn_lst_E) + then have "\!i \ \" + by (metis \H \ elts \\ \Ord \\ eq ordertype_le_Ord) + then show thesis + using \i < l\ that by auto +qed + +lemma partn_lst_triv0: + assumes "\!i = 0" "i < length \" "n \ 0" + shows "partn_lst r B \ n" + by (metis partn_lst_def assms bot_least image_empty nsets_empty_iff ordertype_empty) + +lemma partn_lst_triv1: + assumes "\!i \ 1" "i < length \" "n > 1" "B \ {}" "wf r" + shows "partn_lst r B \ n" + unfolding partn_lst_def +proof clarsimp + obtain \ where "\ \ B" "\ \ []" + using assms mem_0_Ord by fastforce + have 01: "\!i = 0 \ \!i = 1" + using assms by (fastforce simp: one_V_def) + fix f + assume f: "f \ [B]\<^bsup>n\<^esup> \ {..}" + with assms have "ordertype {\} r = 1 \ f ` [{\}]\<^bsup>n\<^esup> \ {i}" + "ordertype {} r = 0 \ f ` [{}]\<^bsup>n\<^esup> \ {i}" + by (auto simp: one_V_def ordertype_insert nsets_eq_empty) + with assms 01 show "\i. \H\B. ordertype H r = \ ! i \ f ` [H]\<^bsup>n\<^esup> \ {i}" + using \\ \ B\ by auto +qed + +lemma partn_lst_two_swap: + assumes "partn_lst r B [x,y] n" shows "partn_lst r B [y,x] n" +proof - + { fix f :: "'a set \ nat" + assume f: "f \ [B]\<^bsup>n\<^esup> \ {..<2}" + then have f': "(\i. 1 - i) \ f \ [B]\<^bsup>n\<^esup> \ {..<2}" + by (auto simp: Pi_def) + obtain i H where "i<2" "H \ B" "ordertype H r = ([x,y]!i)" "((\i. 1 - i) \ f) ` (nsets H n) \ {i}" + by (auto intro: partn_lst_E [OF assms f']) + moreover have "f x = Suc 0" if "Suc 0 \ f x" "x\[H]\<^bsup>n\<^esup>" for x + using f that \H \ B\ nsets_mono by (fastforce simp: Pi_iff) + ultimately have "ordertype H r = [y,x] ! (1-i) \ f ` [H]\<^bsup>n\<^esup> \ {1-i}" + by (force simp: eval_nat_numeral less_Suc_eq) + then have "\i H. i<2 \ H\B \ ordertype H r = [y,x] ! i \ f ` [H]\<^bsup>n\<^esup> \ {i}" + by (metis Suc_1 \H \ B\ diff_less_Suc) } + then show ?thesis + by (auto simp: partn_lst_def eval_nat_numeral) +qed + +lemma partn_lst_greater_resource: + assumes M: "partn_lst r B \ n" and "B \ C" + shows "partn_lst r C \ n" +proof (clarsimp simp: partn_lst_def) + fix f + assume "f \ [C]\<^bsup>n\<^esup> \ {..}" + then have "f \ [B]\<^bsup>n\<^esup> \ {..}" + by (metis \B \ C\ part_fn_def part_fn_subset) + then obtain i H where "i < length \" + and "H \ B" "ordertype H r = (\!i)" + and "f ` nsets H n \ {i}" + using M partn_lst_def by metis + then show "\i. \H\C. ordertype H r = \ ! i \ f ` [H]\<^bsup>n\<^esup> \ {i}" + using \B \ C\ by blast +qed + + +lemma partn_lst_less: + assumes M: "partn_lst r B \ n" and eq: "length \' = length \" and "List.set \' \ ON" + and le: "\i. i < length \ \ \'!i \ \!i " + and r: "wf r" "trans r" "total_on B r" and "small B" + shows "partn_lst r B \' n" +proof (clarsimp simp: partn_lst_def) + fix f + assume "f \ [B]\<^bsup>n\<^esup> \ {..'}" + then obtain i H where "i < length \" + and "H \ B" "small H" and H: "ordertype H r = (\!i)" + and fi: "f ` nsets H n \ {i}" + using assms by (auto simp: partn_lst_def smaller_than_small) + then have bij: "bij_betw (ordermap H r) H (elts (\!i))" + using ordermap_bij [of r H] + by (smt assms(8) in_mono r(1) r(3) smaller_than_small total_on_def) + define H' where "H' = inv_into H (ordermap H r) ` (elts (\'!i))" + have "H' \ H" + using bij \i < length \\ bij_betw_imp_surj_on le + by (force simp: H'_def image_subset_iff intro: inv_into_into) + moreover have ot: "ordertype H' r = (\'!i)" + proof (subst ordertype_eq_iff) + show "Ord (\' ! i)" + using assms by (simp add: \i < length \\ subset_eq) + show "small H'" + by (simp add: H'_def) + show "\f. bij_betw f H' (elts (\' ! i)) \ (\x\H'. \y\H'. (f x < f y) = ((x, y) \ r))" + proof (intro exI conjI ballI) + show "bij_betw (ordermap H r) H' (elts (\' ! i))" + using \H' \ H\ + by (metis H'_def \i < length \\ bij bij_betw_inv_into_RIGHT bij_betw_subset le less_eq_V_def) + show "(ordermap H r x < ordermap H r y) = ((x, y) \ r)" + if "x \ H'" "y \ H'" for x y + proof (intro iffI ordermap_mono_less) + assume "ordermap H r x < ordermap H r y" + then show "(x, y) \ r" + by (metis \H \ B\ assms(8) calculation in_mono leD ordermap_mono_le r smaller_than_small that total_on_def) + qed (use assms that \H' \ H\ \small H\ in auto) + qed + show "total_on H' r" + using r by (meson \H \ B\ \H' \ H\ subsetD total_on_def) + qed (use r in auto) + ultimately show "\i'. \H\B. ordertype H r = \' ! i \ f ` [H]\<^bsup>n\<^esup> \ {i}" + using \H \ B\ \i < length \\ fi assms + by (metis image_mono nsets_mono subset_trans) +qed + + +text \Holds because no $n$-sets exist!\ +lemma partn_lst_VWF_degenerate: + assumes "k < n" + shows "partn_lst_VWF \ (ord_of_nat k # \s) n" +proof (clarsimp simp: partn_lst_def) + fix f :: "V set \ nat" + have "[elts (ord_of_nat k)]\<^bsup>n\<^esup> = {}" + by (simp add: nsets_eq_empty assms finite_Ord_omega) + then have "f ` [elts (ord_of_nat k)]\<^bsup>n\<^esup> \ {0}" + by auto + then show "\i < Suc (length \s). \H\elts \. tp H = (ord_of_nat k # \s) ! i \ f ` [H]\<^bsup>n\<^esup> \ {i}" + using assms ordertype_eq_Ord [of "ord_of_nat k"] elts_ord_of_nat less_Suc_eq_0_disj + by fastforce +qed + +lemma partn_lst_VWF_\_2: + assumes "Ord \" + shows "partn_lst_VWF (\ \ (1+\)) [2, \ \ (1+\)] 2" (is "partn_lst_VWF ?\ _ _") +proof (clarsimp simp: partn_lst_def) + fix f + assume f: "f \ [elts ?\]\<^bsup>2\<^esup> \ {..iH\elts ?\. tp H = [2, ?\] ! i \ f ` [H]\<^bsup>2\<^esup> \ {i}" + proof (cases "\x \ elts ?\. \y \ elts ?\. x \ y \ f{x,y} = 0") + case True + then obtain x y where "x \ elts ?\" "y \ elts ?\" "x \ y" "f {x, y} = 0" + by auto + then have "{x,y} \ elts ?\" "tp {x,y} = 2" + "f ` [{x, y}]\<^bsup>2\<^esup> \ {0}" + by auto (simp add: eval_nat_numeral ordertype_VWF_finite_nat) + with \x \ y\ show ?thesis + by (metis nth_Cons_0 zero_less_Suc) + next + case False + with f have "\x\elts ?\. \y\elts ?\. x \ y \ f {x, y} = 1" + unfolding Pi_iff using lessThan_Suc by force + then have "tp (elts ?\) = ?\" "f ` [elts ?\]\<^bsup>2\<^esup> \ {Suc 0}" + by (auto simp: assms nsets_2_eq) + then show ?thesis + by (metis lessI nth_Cons_0 nth_Cons_Suc subsetI) + qed +qed + + +subsection \Relating partition properties on @{term VWF} to the general case\ + +text \Two very similar proofs here!\ + +lemma partn_lst_imp_partn_lst_VWF_eq: + assumes part: "partn_lst r U \ n" and \: "ordertype U r = \" "small U" + and r: "wf r" "trans r" "total_on U r" + shows "partn_lst_VWF \ \ n" + unfolding partn_lst_def +proof clarsimp + fix f + assume f: "f \ [elts \]\<^bsup>n\<^esup> \ {..}" + define cv where "cv \ \X. ordermap U r ` X" + have bij: "bij_betw (ordermap U r) U (elts \)" + using ordermap_bij [of "r" U] assms by blast + then have bij_cv: "bij_betw cv ([U]\<^bsup>n\<^esup>) ([elts \]\<^bsup>n\<^esup>)" + using bij_betw_nsets cv_def by blast + then have func: "f \ cv \ [U]\<^bsup>n\<^esup> \ {..}" and "inj_on (ordermap U r) U" + using bij bij_betw_def bij_betw_apply f by fastforce+ + then have cv_part: "\\x\[X]\<^bsup>n\<^esup>. f (cv x) = i; X \ U; a \ [cv X]\<^bsup>n\<^esup>\ \ f a = i" for a X i n + by (force simp: cv_def nsets_def subset_image_iff inj_on_subset finite_image_iff card_image) + have ot_eq [simp]: "tp (cv X) = ordertype X r" if "X \ U" for X + unfolding cv_def + proof (rule ordertype_inc_eq) + fix u v + assume "u \ X" "v \ X" and "(u,v) \ r" + with that have "ordermap U r u < ordermap U r v" + by (simp add: assms ordermap_mono_less subset_eq) + then show "(ordermap U r u, ordermap U r v) \ VWF" + by (simp add: r) + next + show "total_on X r" + using that r by (auto simp: total_on_def) + show "small X" + by (meson \small U\ smaller_than_small that) + qed (use assms in auto) + obtain X i where "X \ U" and X: "ordertype X r = \!i" "(f \ cv) ` [X]\<^bsup>n\<^esup> \ {i}" + and "i < length \" + using part func by (auto simp: partn_lst_def) + show "\i < length \. \H\elts \. tp H = \!i \ f ` [H]\<^bsup>n\<^esup> \ {i}" + proof (intro exI conjI) + show "i < length \" + by (simp add: \i < length \\) + show "cv X \ elts \" + using \X \ U\ bij bij_betw_imp_surj_on cv_def by blast + show "tp (cv X) = \ ! i" + by (simp add: X(1) \X \ U\) + show "f ` [cv X]\<^bsup>n\<^esup> \ {i}" + using X \X \ U\ cv_part unfolding image_subset_iff cv_def + by (metis comp_apply insertCI singletonD) + qed +qed + +lemma partn_lst_imp_partn_lst_VWF: + assumes part: "partn_lst r U \ n" and \: "ordertype U r \ \" "small U" + and r: "wf r" "trans r" "total_on U r" + shows "partn_lst_VWF \ \ n" + by (metis assms less_eq_V_def partn_lst_imp_partn_lst_VWF_eq partn_lst_greater_resource) + +lemma partn_lst_VWF_imp_partn_lst_eq: + assumes part: "partn_lst_VWF \ \ n" and \: "ordertype U r = \" "small U" + and r: "wf r" "trans r" "total_on U r" + shows "partn_lst r U \ n" + unfolding partn_lst_def +proof clarsimp + fix f + assume f: "f \ [U]\<^bsup>n\<^esup> \ {..}" + define cv where "cv \ \X. inv_into U (ordermap U r) ` X" + have bij: "bij_betw (ordermap U r) U (elts \)" + using ordermap_bij [of "r" U] assms by blast + then have bij_cv: "bij_betw cv ([elts \]\<^bsup>n\<^esup>) ([U]\<^bsup>n\<^esup>)" + using bij_betw_nsets bij_betw_inv_into unfolding cv_def by blast + then have func: "f \ cv \ [elts \]\<^bsup>n\<^esup> \ {..}" + using bij_betw_apply f by fastforce + have "inj_on (ordermap U r) U" + using bij bij_betw_def by blast + then have cv_part: "\\x\[X]\<^bsup>n\<^esup>. f (cv x) = i; X \ elts \; a \ [cv X]\<^bsup>n\<^esup>\ \ f a = i" for a X i n + apply ( simp add: cv_def nsets_def subset_image_iff inj_on_subset finite_image_iff card_image) + by (metis bij bij_betw_def card_image inj_on_finite inj_on_inv_into subset_eq) + have ot_eq [simp]: "ordertype (cv X) r = tp X" if "X \ elts \" for X + unfolding cv_def + proof (rule ordertype_inc_eq) + show "small X" + using down that by auto + show "(inv_into U (ordermap U r) x, inv_into U (ordermap U r) y) \ r" + if "x \ X" "y \ X" and "(x,y) \ VWF" for x y + proof - + have xy: "x \ ordermap U r ` U" "y \ ordermap U r ` U" + using \X \ elts \\ \x \ X\ \y \ X\ bij bij_betw_imp_surj_on by blast+ + then have eq: "ordermap U r (inv_into U (ordermap U r) x) = x" "ordermap U r (inv_into U (ordermap U r) y) = y" + by (meson f_inv_into_f)+ + then have "y \ elts x" + by (metis (no_types) VWF_non_refl mem_imp_VWF that(3) trans_VWF trans_def) + then show ?thesis + by (metis (no_types) VWF_non_refl xy eq assms(3) inv_into_into ordermap_mono r(1) r(3) that(3) total_on_def) + qed + qed (use r in auto) + obtain X i where "X \ elts \" and X: "tp X = \!i" "(f \ cv) ` [X]\<^bsup>n\<^esup> \ {i}" + and "i < length \" + using part func by (auto simp: partn_lst_def) + show "\i < length \. \H\U. ordertype H r = \!i \ f ` [H]\<^bsup>n\<^esup> \ {i}" + proof (intro exI conjI) + show "i < length \" + by (simp add: \i < length \\) + show "cv X \ U" + using \X \ elts \\ bij bij_betw_imp_surj_on bij_betw_inv_into cv_def by blast + show "ordertype (cv X) r = \ ! i" + by (simp add: X(1) \X \ elts \\) + show "f ` [cv X]\<^bsup>n\<^esup> \ {i}" + using X \X \ elts \\ cv_part unfolding image_subset_iff cv_def + by (metis comp_apply insertCI singletonD) + qed +qed + +corollary partn_lst_VWF_imp_partn_lst: + assumes "partn_lst_VWF \ \ n" and \: "ordertype U r \ \" "small U" + "wf r" "trans r" "total_on U r" + shows "partn_lst r U \ n" + by (metis assms less_eq_V_def partn_lst_VWF_imp_partn_lst_eq partn_lst_greater_resource) + +subsection \Simple consequences of the definitions\ + +text \A restatement of the infinite Ramsey theorem using partition notation\ +lemma Ramsey_partn: "partn_lst_VWF \ [\,\] 2" +proof (clarsimp simp: partn_lst_def) + fix f + assume "f \ [elts \]\<^bsup>2\<^esup> \ {..x\elts \. \y\elts \. x \ y \ f {x, y} < 2" + by (auto simp: nsets_def eval_nat_numeral) + obtain H i where H: "H \ elts \" and "infinite H" + and t: "i < Suc (Suc 0)" + and teq: "\x\H. \y\H. x \ y \ f {x, y} = i" + using Ramsey2 [OF infinite_\ *] by (auto simp: eval_nat_numeral) + then have "tp H = [\, \] ! i" + using less_2_cases eval_nat_numeral ordertype_infinite_\ by force + moreover have "f ` {N. N \ H \ finite N \ card N = 2} \ {i}" + by (force simp: teq card_2_iff) + ultimately have "f ` [H]\<^bsup>2\<^esup> \ {i}" + by (metis (no_types) nsets_def numeral_2_eq_2) + then show "\iH\elts \. tp H = [\,\] ! i \ f ` [H]\<^bsup>2\<^esup> \ {i}" + using H \tp H = [\, \] ! i\ t by blast +qed + +text \This is the counterexample sketched in Hajnal and Larson, section 9.1.\ +proposition omega_basic_counterexample: + assumes "Ord \" + shows "\ partn_lst_VWF \ [succ (vcard \), \] 2" +proof - + obtain \ where fun\: "\ \ elts \ \ elts (vcard \)" and inj\: "inj_on \ (elts \)" + using inj_into_vcard by auto + have Ord\: "Ord (\ x)" if "x \ elts \" for x + using Ord_in_Ord fun\ that by fastforce + define f where "f A \ @i::nat. \x y. A = {x,y} \ x < y \ (\ x < \ y \ i=0 \ \ y < \ x \ i=1)" for A + have f_Pi: "f \ [elts \]\<^bsup>2\<^esup> \ {.. [elts \]\<^bsup>2\<^esup>" + then obtain x y where xy: "x \ elts \" "y \ elts \" "x < y" and A: "A = {x,y}" + apply (clarsimp simp: nsets_2_eq) + by (metis Ord_in_Ord Ord_linear_lt assms insert_commute) + consider "\ x < \ y" | "\ y < \ x" + by (metis Ord\ Ord_linear_lt inj\ inj_onD less_imp_not_eq2 xy) + then show "f A \ {..x < y\ A exE_some [OF _ f_def]) + qed + have fiff: "\ x < \ y \ i=0 \ \ y < \ x \ i=1" + if f: "f {x,y} = i" and xy: "x \ elts \" "y \ elts \" "x x < \ y" | "\ y < \ x" + using xy by (metis Ord\ Ord_linear_lt inj\ inj_onD less_V_def) + then show ?thesis + proof cases + case 1 + then have "f{x,y} = 0" + using \x by (force simp: f_def Set.doubleton_eq_iff) + then show ?thesis + using "1" f by auto + next + case 2 + then have "f{x,y} = 1" + using \x by (force simp: f_def Set.doubleton_eq_iff) + then show ?thesis + using "2" f by auto + qed + qed + have False + if eq: "tp H = succ (vcard \)" and H: "H \ elts \" + and 0: "\A. A \ [H]\<^bsup>2\<^esup> \ f A = 0" for H + proof - + have [simp]: "small H" + using H down by auto + have OH: "Ord x" if "x \ H" for x + using H Ord_in_Ord \Ord \\ that by blast + have \: "\ x < \ y" if "x\H" "y\H" "x ` H \ elts (vcard \)" + using H fun\ by auto + have "tp H = tp (\ ` H)" + proof (rule ordertype_VWF_inc_eq [symmetric]) + show "\ ` H \ ON" + using H Ord\ by blast + qed (auto simp: \ OH subsetI) + also have "\ \ vcard \" + by (simp add: H sub_vcard assms ordertype_le_Ord) + finally show False + by (simp add: eq succ_le_iff) + qed + moreover have False + if eq: "tp H = \" and H: "H \ elts \" + and 1: "\A. A \ [H]\<^bsup>2\<^esup> \ f A = Suc 0" for H + proof - + have [simp]: "small H" + using H down by auto + define \ where "\ \ inv_into H (ordermap H VWF) \ ord_of_nat" + have bij: "bij_betw (ordermap H VWF) H (elts \)" + by (metis ordermap_bij \small H\ eq total_on_VWF wf_VWF) + then have "bij_betw (inv_into H (ordermap H VWF)) (elts \) H" + by (simp add: bij_betw_inv_into) + then have \: "bij_betw \ UNIV H" + unfolding \_def + by (metis \_def bij_betw_comp_iff2 bij_betw_def elts_of_set inf inj_ord_of_nat order_refl) + have Ord\: "Ord (\ k)" for k + by (meson H Ord_in_Ord UNIV_I \ assms bij_betw_apply subsetD) + obtain k where k: "((\ \ \)(Suc k), (\ \ \) k) \ VWF" + using wf_VWF wf_iff_no_infinite_down_chain by blast + have \: "\ y < \ x" if "x\H" "y\H" "x (Suc k) \ \ k" + proof - + have "(\ (Suc k), \ k) \ VWF \ \ (Suc k) = \ k" + using that Ord\ Ord_mem_iff_lt by auto + then have "ordermap H VWF (\ (Suc k)) \ ordermap H VWF (\ k)" + by (metis \ \small H\ bij_betw_imp_surj_on ordermap_mono_le rangeI trans_VWF wf_VWF) + moreover have "ordermap H VWF (\ (Suc k)) = succ (ord_of_nat k)" + unfolding \_def using bij bij_betw_inv_into_right by force + moreover have "ordermap H VWF (\ k) = ord_of_nat k" + apply (simp add: \_def) + by (meson bij bij_betw_inv_into_right ord_of_nat_\) + ultimately have "succ (ord_of_nat k) \ ord_of_nat k" + by simp + then show False + by (simp add: less_eq_V_def) + qed + then have "\ k < \ (Suc k)" + by (metis Ord\ Ord_linear_lt dual_order.strict_implies_order eq_refl) + then have "(\ \ \)(Suc k) < (\ \ \)k" + using \ \ bij_betw_apply by force + then show False + using k + apply (simp add: subset_iff) + by (metis H Ord\ UNIV_I VWF_iff_Ord_less \ bij_betw_imp_surj_on image_subset_iff) + qed + ultimately show ?thesis + apply (simp add: partn_lst_def image_subset_iff) + by (metis f_Pi less_2_cases nth_Cons_0 nth_Cons_Suc numeral_2_eq_2) +qed + +subsection \Specker's theorem\ + +definition form_split :: "[nat,nat,nat,nat,nat] \ bool" where + "form_split a b c d i \ a \ c \ (i=0 \ a b c + i=1 \ a c b + i=2 \ a c d + i=3 \ a=c \ b\d)" + +definition form :: "[(nat*nat)set, nat] \ bool" where + "form u i \ \a b c d. u = {(a,b),(c,d)} \ form_split a b c d i" + +definition scheme :: "[(nat*nat)set] \ nat set" where + "scheme u \ fst ` u \ snd ` u" + +definition UU :: "(nat*nat) set" + where "UU \ {(a,b). a < b}" + +lemma ordertype_UNIV_\2: "ordertype UNIV pair_less = \\2" + using ordertype_Times [of concl: UNIV UNIV less_than less_than] + by (simp add: total_less_than pair_less_def ordertype_nat_\ numeral_2_eq_2) + +lemma ordertype_UU_ge_\2: "ordertype UNIV pair_less \ ordertype UU pair_less" +proof (rule ordertype_inc_le) + define \ where "\ \ \(m,n). (m, Suc (m+n))" + show "(\ (x::nat \ nat), \ y) \ pair_less" if "(x, y) \ pair_less" for x y + using that by (auto simp: \_def pair_less_def split: prod.split) + show "range \ \ UU" + by (auto simp: \_def UU_def) +qed auto + +lemma ordertype_UU_\2: "ordertype UU pair_less = \\2" + by (metis eq_iff ordertype_UNIV_\2 ordertype_UU_ge_\2 ordertype_mono small top_greatest trans_pair_less wf_pair_less) + + +text \Lemma 2.3 of 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_2_3: + fixes f :: "(nat \ nat) set \ nat" + assumes "f \ [UU]\<^bsup>2\<^esup> \ {..k u. \k < 4; u \ [UU]\<^bsup>2\<^esup>; form u k; scheme u \ N\ \ f u = js!k" +proof - + have f_less2: "f {p,q} < Suc (Suc 0)" if "p \ q" "p \ UU" "q \ UU" for p q + proof - + have "{p,q} \ [UU]\<^bsup>2\<^esup>" + using that by (simp add: nsets_def) + then show ?thesis + using assms by (simp add: Pi_iff) + qed + define f0 where "f0 \ (\A::nat set. THE x. \a b c d. A = {a,b,c,d} \ a b c x = f {(a,b),(c,d)})" + have f0: "f0 {a,b,c,d} = f {(a,b),(c,d)}" if "a [X]\<^bsup>4\<^esup>" + using that by (auto simp: nsets_def) + then obtain a b c d where "X = {a,b,c,d} \ a b cN t. infinite N \ t < Suc (Suc 0) + \ (\X. X \ N \ finite X \ card X = 4 \ f0 X = t)" + using Ramsey [of UNIV 4 f0 2] by (simp add: eval_nat_numeral) + then obtain N0 j0 where "infinite N0" and j0: "j0 < Suc (Suc 0)" and N0: "\A. A \ [N0]\<^bsup>4\<^esup> \ f0 A = j0" + by (auto simp: nsets_def) + + define f1 where "f1 \ (\A::nat set. THE x. \a b c d. A = {a,b,c,d} \ a b c x = f {(a,c),(b,d)})" + have f1: "f1 {a,b,c,d} = f {(a,c),(b,d)}" if "a [X]\<^bsup>4\<^esup>" + using that by (auto simp: nsets_def) + then obtain a b c d where "X = {a,b,c,d} \ a b cN t. N \ N0 \ infinite N \ t < Suc (Suc 0) + \ (\X. X \ N \ finite X \ card X = 4 \ f1 X = t)" + using \infinite N0\ Ramsey [of N0 4 f1 2] by (simp add: eval_nat_numeral) + then obtain N1 j1 where "N1 \ N0" "infinite N1" and j1: "j1 < Suc (Suc 0)" and N1: "\A. A \ [N1]\<^bsup>4\<^esup> \ f1 A = j1" + by (auto simp: nsets_def) + + define f2 where "f2 \ (\A::nat set. THE x. \a b c d. A = {a,b,c,d} \ a b c x = f {(a,d),(b,c)})" + have f2: "f2 {a,b,c,d} = f {(a,d),(b,c)}" if "a [X]\<^bsup>4\<^esup>" + using that by (auto simp: nsets_def) + then obtain a b c d where "X = {a,b,c,d} \ a b cN t. N \ N1 \ infinite N \ t < Suc (Suc 0) + \ (\X. X \ N \ finite X \ card X = 4 \ f2 X = t)" + using \infinite N1\ Ramsey [of N1 4 f2 2] by (simp add: eval_nat_numeral) + then obtain N2 j2 where "N2 \ N1" "infinite N2" and j2: "j2 < Suc (Suc 0)" and N2: "\A. A \ [N2]\<^bsup>4\<^esup> \ f2 A = j2" + by (auto simp: nsets_def) + + define f3 where "f3 \ (\A::nat set. THE x. \a b c. A = {a,b,c} \ a b x = f {(a,b),(a,c)})" + have f3: "f3 {a,b,c} = f {(a,b),(a,c)}" if "a [X]\<^bsup>3\<^esup>" + using that by (auto simp: nsets_def) + then obtain a b c where "X = {a,b,c} \ a bN t. N \ N2 \ infinite N \ t < Suc (Suc 0) + \ (\X. X \ N \ finite X \ card X = 3 \ f3 X = t)" + using \infinite N2\ Ramsey [of N2 3 f3 2] by (simp add: eval_nat_numeral) + then obtain N3 j3 where "N3 \ N2" "infinite N3" and j3: "j3 < Suc (Suc 0)" and N3: "\A. A \ [N3]\<^bsup>3\<^esup> \ f3 A = j3" + by (auto simp: nsets_def) + + show thesis + proof + fix k u + assume "k < 4" + and u: "form u k" "scheme u \ N3" + and UU: "u \ [UU]\<^bsup>2\<^esup>" + then consider (0) "k=0" | (1) "k=1" | (2) "k=2" | (3) "k=3" + by linarith + then show "f u = [j0,j1,j2,j3] ! k" + proof cases + case 0 + have "N3 \ N0" + using \N1 \ N0\ \N2 \ N1\ \N3 \ N2\ by auto + then show ?thesis + using u 0 + apply (auto simp: form_def form_split_def scheme_def simp flip: f0) + apply (force simp: nsets_def intro: N0) + done + next + case 1 + have "N3 \ N1" + using \N2 \ N1\ \N3 \ N2\ by auto + then show ?thesis + using u 1 + apply (auto simp: form_def form_split_def scheme_def simp flip: f1) + apply (force simp: nsets_def intro: N1) + done + next + case 2 + then show ?thesis + using u \N3 \ N2\ + apply (auto simp: form_def form_split_def scheme_def nsets_def simp flip: f2) + apply (force simp: nsets_def intro: N2) + done + next + case 3 + { fix a b d + assume "{(a, b), (a, d)} \ [UU]\<^bsup>2\<^esup>" + and *: "a \ N3" "b \ N3" "d \ N3" "b \ d" + then have "ainfinite N3\) +qed + + +text \Lemma 2.4 of Jean A. Larson, ibid.\ +lemma lemma_2_4: + assumes "infinite N" "k < 4" + obtains M where "M \ [UU]\<^bsup>m\<^esup>" "\u. u \ [M]\<^bsup>2\<^esup> \ form u k" "\u. u \ [M]\<^bsup>2\<^esup> \ scheme u \ N" +proof - + obtain f:: "nat \ nat" where "bij_betw f UNIV N" "strict_mono f" + using assms by (meson bij_enumerate enumerate_mono strict_monoI) + then have iff[simp]: "f x = f y \ x=y" "f x < f y \ x N" for x + using bij_betw_apply [OF \bij_betw f UNIV N\] by blast + define M0 where "M0 = (\i. (f(2*i), f(Suc(2*i)))) ` {..i. (f i, f(m+i))) ` {..i. (f i, f(2*m-i))) ` {..i. (f 0, f (Suc i))) ` {..i. (f (2 * i), f (Suc (2 * i)))) {.. [UU]\<^bsup>m\<^esup>" + by (simp add: M0_def nsets_def card_image UU_def image_subset_iff) + next + fix u + assume u: "(u::(nat \ nat) set) \ [M0]\<^bsup>2\<^esup>" + then obtain x y where "u = {x,y}" "x \ y" "x \ M0" "y \ M0" + by (auto simp: nsets_2_eq) + then obtain i j where "i f (2 * j)" + by (simp add: \i less_imp_le_nat) + ultimately show "form u k" + apply (simp add: 0 form_def form_split_def nsets_def) + apply (rule_tac x="f (2 * i)" in exI) + apply (rule_tac x="f (Suc (2 * i))" in exI) + apply (rule_tac x="f (2 * j)" in exI) + apply (rule_tac x="f (Suc (2 * j))" in exI) + apply auto + done + show "scheme u \ N" + using ueq by (auto simp: scheme_def) + qed + next + case 1 + show ?thesis + proof + have "inj_on (\i. (f i, f(m+i))) {.. [UU]\<^bsup>m\<^esup>" + by (simp add: M1_def nsets_def card_image UU_def image_subset_iff) + next + fix u + assume u: "(u::(nat \ nat) set) \ [M1]\<^bsup>2\<^esup>" + then obtain x y where "u = {x,y}" "x \ y" "x \ M1" "y \ M1" + by (auto simp: nsets_2_eq) + then obtain i j where "i N" + using ueq by (auto simp: scheme_def) + qed + next + case 2 + show ?thesis + proof + have "inj_on (\i. (f i, f(2*m-i))) {.. [UU]\<^bsup>m\<^esup>" + by (auto simp: M2_def nsets_def card_image UU_def image_subset_iff) + next + fix u + assume u: "(u::(nat \ nat) set) \ [M2]\<^bsup>2\<^esup>" + then obtain x y where "u = {x,y}" "x \ y" "x \ M2" "y \ M2" + by (auto simp: nsets_2_eq) + then obtain i j where "i N" + using ueq by (auto simp: scheme_def) + qed + next + case 3 + show ?thesis + proof + have "inj_on (\i. (f 0, f (Suc i))) {.. [UU]\<^bsup>m\<^esup>" + by (auto simp: M3_def nsets_def card_image UU_def image_subset_iff) + next + fix u + assume u: "(u::(nat \ nat) set) \ [M3]\<^bsup>2\<^esup>" + then obtain x y where "u = {x,y}" "x \ y" "x \ M3" "y \ M3" + by (auto simp: nsets_2_eq) + then obtain i j where "i N" + using ueq by (auto simp: scheme_def) + qed + qed +qed + + +text \Lemma 2.5 of Jean A. Larson, ibid.\ +lemma lemma_2_5: + assumes "infinite N" + obtains X where "X \ UU" "ordertype X pair_less = \\2" + "\u. u \ [X]\<^bsup>2\<^esup> \ (\k<4. form u k) \ scheme u \ N" +proof - + obtain C + where dis: "pairwise (\i j. disjnt (C i) (C j)) UNIV" + and N: "(\i. C i) \ N" and infC: "\i::nat. infinite (C i)" + using assms infinite_infinite_partition by blast + then have "\\::nat \ nat. inj \ \ range \ = C i \ strict_mono \" for i + by (metis bij_betw_imp_inj_on bij_betw_imp_surj_on bij_enumerate enumerate_mono infC strict_mono_def) + then obtain \:: "[nat,nat] \ nat" + where \: "\i. inj (\ i) \ range (\ i) = C i \ strict_mono (\ i)" + by metis + then have \_in_C [simp]: "\ i j \ C i' \ i'=i" for i i' j + using dis by (fastforce simp: pairwise_def disjnt_def) + have less_iff [simp]: "\ i j' < \ i j \ j' < j" for i j' j + by (simp add: \ strict_mono_less) + let ?a = "\ 0" + define X where "X \ {(?a i, b) | i b. ?a i < b \ b \ C (Suc i)}" + show thesis + proof + show "X \ UU" + by (auto simp: X_def UU_def) + show "ordertype X pair_less = \\2" + proof (rule antisym) + have "ordertype X pair_less \ ordertype UU pair_less" + by (simp add: \X \ UU\ ordertype_mono) + then show "ordertype X pair_less \ \\2" + using ordertype_UU_\2 by auto + define \ where "\ \ \(i,j::nat). (?a i, \ (Suc i) (?a j))" + have "\i j. i < j \ \ 0 i < \ (Suc i) (\ 0 j)" + by (meson \ le_less_trans less_iff strict_mono_imp_increasing) + then have subX: "\ ` UU \ X" + by (auto simp: UU_def \_def X_def) + then have "ordertype (\ ` UU) pair_less \ ordertype X pair_less" + by (simp add: ordertype_mono) + moreover have "ordertype (\ ` UU) pair_less = ordertype UU pair_less" + proof (rule ordertype_inc_eq) + show "(\ x, \ y) \ pair_less" + if "x \ UU" "y \ UU" and "(x, y) \ pair_less" for x y + using that by (auto simp: UU_def \_def pair_less_def) + qed auto + ultimately show "\\2 \ ordertype X pair_less" + using ordertype_UU_\2 by simp + qed + next + fix U + assume "U \ [X]\<^bsup>2\<^esup>" + then obtain a b c d where Ueq: "U = {(a,b),(c,d)}" and ne: "(a,b) \ (c,d)" and inX: "(a,b) \ X" "(c,d) \ X" and "a \ c" + apply (auto simp: nsets_def subset_iff eval_nat_numeral card_Suc_eq Set.doubleton_eq_iff) + apply (metis nat_le_linear)+ + done + show "(\k<4. form U k) \ scheme U \ N" + proof + show "scheme U \ N" + using inX N \ by (fastforce simp: scheme_def Ueq X_def) + next + consider "a < c" | "a = c \ b \ d" + using \a \ c\ ne nat_less_le by blast + then show "\k<4. form U k" + proof cases + case 1 + have *: "a < b" "b \ c" "c < d" + using inX by (auto simp: X_def) + moreover have "\a < c; c < b; \ d < b\ \ b < d" + using inX apply (clarsimp simp: X_def not_less) + by (metis \ \_in_C imageE nat.inject nat_less_le) + ultimately consider (k0) "a b c c b c da \ c\ by blast + then show ?thesis by force + next + case k1 + then have "form U 1" + unfolding form_def form_split_def using Ueq \a \ c\ by blast + then show ?thesis by force + next + case k2 + then have "form U 2" + unfolding form_def form_split_def using Ueq \a \ c\ by blast + then show ?thesis by force + qed + next + case 2 + then have "form_split a b c d 3" + by (auto simp: form_split_def) + then show ?thesis + using Ueq form_def leI by force + qed + qed + qed +qed + +text \Theorem 2.1 of Jean A. Larson, ibid.\ +lemma Specker_aux: + assumes "\ \ elts \" + shows "partn_lst pair_less UU [\\2,\] 2" + unfolding partn_lst_def +proof clarsimp + fix f + assume f: "f \ [UU]\<^bsup>2\<^esup> \ {..: "?P0 \ ?P1" + proof (rule disjCI) + assume "\ ?P1" + then have not1: "\M. \M \ UU; ordertype M pair_less = \\ \ \x\[M]\<^bsup>2\<^esup>. f x \ Suc 0" + by auto + obtain m where m: "\ = ord_of_nat m" + using assms elts_\ by auto + then have f_eq_0: "M \ [UU]\<^bsup>m\<^esup> \ \x\[M]\<^bsup>2\<^esup>. f x = 0" for M + using not1 [of M] finite_ordertype_eq_card [of M pair_less m] f + apply (clarsimp simp: nsets_def eval_nat_numeral Pi_def) + by (meson less_Suc0 not_less_less_Suc_eq subset_trans) + obtain N js where "infinite N" and N: "\k u. \k < 4; u \ [UU]\<^bsup>2\<^esup>; form u k; scheme u \ N\ \ f u = js!k" + using f lemma_2_3 by blast + obtain M0 where M0: "M0 \ [UU]\<^bsup>m\<^esup>" "\u. u \ [M0]\<^bsup>2\<^esup> \ form u 0" "\u. u \ [M0]\<^bsup>2\<^esup> \ scheme u \ N" + by (rule lemma_2_4 [OF \infinite N\]) auto + obtain M1 where M1: "M1 \ [UU]\<^bsup>m\<^esup>" "\u. u \ [M1]\<^bsup>2\<^esup> \ form u 1" "\u. u \ [M1]\<^bsup>2\<^esup> \ scheme u \ N" + by (rule lemma_2_4 [OF \infinite N\]) auto + obtain M2 where M2: "M2 \ [UU]\<^bsup>m\<^esup>" "\u. u \ [M2]\<^bsup>2\<^esup> \ form u 2" "\u. u \ [M2]\<^bsup>2\<^esup> \ scheme u \ N" + by (rule lemma_2_4 [OF \infinite N\]) auto + obtain M3 where M3: "M3 \ [UU]\<^bsup>m\<^esup>" "\u. u \ [M3]\<^bsup>2\<^esup> \ form u 3" "\u. u \ [M3]\<^bsup>2\<^esup> \ scheme u \ N" + by (rule lemma_2_4 [OF \infinite N\]) auto + have "js!0 = 0" + using N [of 0 ] M0 f_eq_0 [of M0] by (force simp: nsets_def eval_nat_numeral) + moreover have "js!1 = 0" + using N [of 1] M1 f_eq_0 [of M1] by (force simp: nsets_def eval_nat_numeral) + moreover have "js!2 = 0" + using N [of 2 ] M2 f_eq_0 [of M2] by (force simp: nsets_def eval_nat_numeral) + moreover have "js!3 = 0" + using N [of 3 ] M3 f_eq_0 [of M3] by (force simp: nsets_def eval_nat_numeral) + ultimately have js0: "js!k = 0" if "k < 4" for k + using that by (auto simp: eval_nat_numeral less_Suc_eq) + obtain X where "X \ UU" and otX: "ordertype X pair_less = \\2" + and X: "\u. u \ [X]\<^bsup>2\<^esup> \ (\k<4. form u k) \ scheme u \ N" + using \infinite N\ lemma_2_5 by auto + moreover have "f ` [X]\<^bsup>2\<^esup> \ {0}" + proof (clarsimp simp: image_subset_iff) + fix u + assume u: "u \ [X]\<^bsup>2\<^esup>" + then have u_UU2: "u \ [UU]\<^bsup>2\<^esup>" + using \X \ UU\ nsets_mono by blast + show "f u = 0" + using X u N [OF _ u_UU2] js0 by auto + qed + ultimately show "\X \ UU. ordertype X pair_less = \\2 \ f ` [X]\<^bsup>2\<^esup> \ {0}" + by blast + qed + then show "\iH\UU. ordertype H pair_less = [\\2, \] ! i \ f ` [H]\<^bsup>2\<^esup> \ {i}" + proof + show "?P0 \ ?thesis" + by (metis nth_Cons_0 numeral_2_eq_2 pos2) + show "?P1 \ ?thesis" + by (metis One_nat_def lessI nth_Cons_0 nth_Cons_Suc) + qed +qed + +theorem Specker: "\ \ elts \ \ partn_lst_VWF (\\2) [\\2,\] 2" + using partn_lst_imp_partn_lst_VWF_eq [OF Specker_aux] ordertype_UU_\2 wf_pair_less by blast + +end diff --git a/thys/Ordinal_Partitions/ROOT b/thys/Ordinal_Partitions/ROOT new file mode 100644 --- /dev/null +++ b/thys/Ordinal_Partitions/ROOT @@ -0,0 +1,12 @@ +chapter AFP + +session Ordinal_Partitions (AFP) = HOL + + options [timeout = 600] + sessions + "HOL-Library" "ZFC_in_HOL" "Nash_Williams" + theories + Omega_Omega + document_files + "root.tex" + "root.bib" + diff --git a/thys/Ordinal_Partitions/document/root.bib b/thys/Ordinal_Partitions/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Ordinal_Partitions/document/root.bib @@ -0,0 +1,51 @@ +%% This BibTeX bibliography file was created using BibDesk. +%% http://bibdesk.sourceforge.net/ + + +%% Created for Larry Paulson at 2020-08-03 15:23:09 +0100 + + +%% Saved with string encoding Unicode (UTF-8) + + + +@article{erdos-theorem-partition-corr, + Author = {Paul Erd{\"o}s and E. C. Milner}, + Date-Added = {2020-08-03 15:20:59 +0100}, + Date-Modified = {2020-08-03 15:23:08 +0100}, + Doi = {10.4153/CMB-1974-062-6}, + Journal = {Canadian Mathematical Bulletin}, + Month = jun, + Number = {2}, + Pages = {305}, + Title = {A Theorem in the Partition Calculus Corrigendum}, + Volume = {17}, + Year = {1974}, + Bdsk-Url-1 = {https://doi.org/10.4153/CMB-1972-088-1}} + +@article{erdos-theorem-partition, + Author = {Paul Erd{\"o}s and E. C. Milner}, + Date-Added = {2020-08-02 14:51:07 +0100}, + Date-Modified = {2020-08-02 14:53:46 +0100}, + Doi = {10.4153/CMB-1972-088-1}, + Journal = {Canadian Mathematical Bulletin}, + Month = dec, + Number = {4}, + Pages = {501-505}, + Title = {A Theorem in the Partition Calculus}, + Volume = {15}, + Year = {1972}} + +@article{larson-short-proof, + Author = {Jean A. Larson}, + Date-Added = {2020-08-02 14:49:21 +0100}, + Date-Modified = {2020-08-02 14:49:21 +0100}, + Doi = {10.1016/0003-4843(73)90006-5}, + Journal = {Annals of Mathematical Logic}, + Month = dec, + Number = {2}, + Pages = {129-145}, + Title = {A Short Proof of a Partition Theorem for the Ordinal $\omega^\omega$}, + Volume = {6}, + Year = {1973}, + Bdsk-Url-1 = {https://doi.org/10.1016/0003-4843(73)90006-5}} diff --git a/thys/Ordinal_Partitions/document/root.tex b/thys/Ordinal_Partitions/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Ordinal_Partitions/document/root.tex @@ -0,0 +1,40 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} +\usepackage{amssymb} +\usepackage{stmaryrd} + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{A Partition Theorem for the Ordinal $\omega^\omega$} +\author{Lawrence C. Paulson} +\maketitle + +\begin{abstract} +The theory of partition relations concerns generalisations of Ramsey's theorem. +For any ordinal $\alpha$, write $\alpha \to (\alpha, m)^2$ if for each function~$f$ from unordered pairs of elements of~$\alpha$ into $\{0,1\}$, either there is a subset $X\subseteq \alpha$ order-isomorphic to $\alpha$ such that $f\{x,y\}=0$ for all $\{x,y\}\subseteq X$, or there is an $m$ element set $Y\subseteq \alpha$ such that $f\{x,y\}=1$ for all $\{x,y\}\subseteq Y$. (In both cases, with $\{x,y\}$ we require $x\not=y$.) +In particular, the infinite Ramsey theorem can be written in this notation as $\omega \to (\omega, \omega)^2$, or if we restrict~$m$ to the positive integers as above, then $\omega \to (\omega, m)^2$ for all~$m$ \cite{larson-short-proof}. + +This entry formalises Larson's proof of $\omega^\omega \to (\omega^\omega, m)^2$ along with a similar proof of a result due to Specker: $\omega^2 \to (\omega^2, m)^2$. Also proved is a necessary result by Erd{\H o}s and Milner~\cite{erdos-theorem-partition,erdos-theorem-partition-corr}: $\omega^{1+\alpha\cdot n} \to (\omega^{1+\alpha}, 2^n)^2$. + +These examples demonstrate the use of Isabelle/HOL to formalise advanced results that combine ZF set theory with basic concepts like lists and natural numbers. +\end{abstract} + +\tableofcontents + +% include generated text of all theories +\input{session} + +\section{Acknowledgements} +The author was supported by the ERC Advanced Grant ALEXANDRIA (Project 742178) funded by the European Research Council. Many thanks to Mirna D\v{z}amonja (who suggested the project) and Angeliki Koutsoukou-Argyraki for assistance at tricky moments. + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,549 +1,550 @@ ADS_Functor AODV Attack_Trees Auto2_HOL Auto2_Imperative_HOL AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BNF_Operations Banach_Steinhaus Bell_Numbers_Spivey Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BNF_CC Bondy Boolean_Expression_Checkers Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CYK CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean ClockSynchInst Closest_Pair_Points CofGroups Coinductive Coinductive_Languages Collections Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Completeness Complete_Non_Orders Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Consensus_Refined Constructive_Cryptography Constructor_Funs Containers CoreC++ Core_DOM Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties DFS_Framework DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Deep_Learning Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series Discrete_Summation DiscretePricing DiskPaxos DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Example-Submission Factored_Transition_System_Bounding Farkas FFT FLP FOL-Fitting FOL_Harrison FOL_Seq_Calc1 Falling_Factorial_Sum FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite_Automata_HF First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_SSA Formula_Derivatives Fourier Free-Boolean-Algebra Free-Groups FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Heard_Of Hello_World HereditarilyFinite Hermite Hidden_Markov_Models Higher_Order_Terms Hoare_Time HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL IEEE_Floating_Point IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Inductive_Confidentiality InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interval_Arithmetic_Word32 Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Isabelle_C Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_to_DRA LTL_to_GBA LTL_Master_Theorem LTL_Normal_Form Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Lambda_Free_EPO Launchbury Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MFMC_Countable MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Of_Medians_Selection Menger Mersenne_Primes MFODL_Monitor_Optimized MFOTL_Monitor MiniML Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_Assembly_Kit_Security Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multirelations Multi_Party_Computation Myhill-Nerode Name_Carrying_Type_Inference Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Octonions Open_Induction OpSets Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal +Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PCF PLM Pell POPLmark-deBruijn PSemigroupsConvolution Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Perfect-Number-Thm Perron_Frobenius Pi_Calculus Pi_Transcendental Planarity_Certificates Polynomial_Factorization Polynomial_Interpolation Polynomials Poincare_Bendixson Poincare_Disc Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Projective_Geometry Program-Conflict-Analysis Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem QHLProver QR_Decomposition Quantales Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Randomised_BSTs Random_Graph_Subgraph_Threshold Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Regular-Sets Regular_Algebras Relation_Algebra Relational-Incorrectness-Logic Relational_Paths Rep_Fin_Groups Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Root_Balanced_Tree Routing Roy_Floyd_Warshall SATSolverVerification SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares SuperCalc Surprise_Paradox Symmetric_Polynomials Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Timed_Automata Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God Universal_Turing_Machine UPF UPF_Firewall UpDown_Scheme UTP Valuation VectorSpace VeriComp Verified-Prover VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves VolpanoSmith WHATandWHERE_Security WebAssembly Weight_Balanced_Trees Well_Quasi_Orders Winding_Number_Eval WOOT_Strong_Eventual_Consistency Word_Lib WorkerWrapper XML Zeta_Function Zeta_3_Irrational ZFC_in_HOL pGCL