diff --git a/thys/Fishers_Inequality/Design_Extras.thy b/thys/Fishers_Inequality/Design_Extras.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Design_Extras.thy @@ -0,0 +1,620 @@ +(* Title: Design Extras.thy + Author: Chelsea Edmonds +*) + +section \ Micellaneous Design Extras \ + +text \Extension's to the author's previous entry on Design Theory \ + +theory Design_Extras imports Set_Multiset_Extras Design_Theory.BIBD +begin + +subsection \Extensions to existing Locales and Properties \ + +text \Extend lemmas on intersection number\ +lemma inter_num_max_bound: + assumes "finite b1" "finite b2" + shows "b1 |\| b2 \ card b1" "b1 |\| b2 \ card b2" + by(simp_all add: assms intersection_number_def card_mono) + +lemma inter_eq_blocks_eq_card: "card b1 = card b2 \ finite b1 \ finite b2 \ b1 |\| b2 = card b1 + \ b1 = b2" + using equal_card_inter_fin_eq_sets intersection_number_def by (metis) + +lemma inter_num_of_eq_blocks: "b1 = b2 \ b1 |\| b2 = card b1" + by (simp add: intersection_number_def) + +lemma intersect_num_same_eq_size[simp]: "bl |\| bl = card bl" + by (simp add: intersection_number_def) + +lemma index_lt_rep_general: "x \ ps \ B index ps \ B rep x" + by (simp add: points_index_def point_replication_number_def) + (metis filter_filter_mset_cond_simp size_filter_mset_lesseq subset_iff) + +context incidence_system +begin + +lemma block_size_alt: + assumes "bl \# \" + shows "card bl = card {x \ \ . x \ bl}" +proof - + have "\ x. x \ bl \ x \ \" using wellformed assms by auto + thus ?thesis + by (metis (no_types, lifting) Collect_cong Collect_mem_eq) +qed + +lemma complement_image: "\\<^sup>C = image_mset block_complement \" + by (simp add: complement_blocks_def) + +lemma point_in_block_rep_min_iff: + assumes "x \ \" + shows "\ bl . bl \# \ \ x \ bl \ (\ rep x > 0)" + using rep_number_g0_exists + by (metis block_complement_elem_iff block_complement_inv wellformed) + +lemma points_inter_num_rep: + assumes "b1 \# \" and "b2 \# \ - {#b1#}" + shows "card {v \ \ . v \ b1 \ v \ b2} = b1 |\| b2" +proof - + have "\ x. x \ b1 \ b2 \ x \ \" using wellformed assms by auto + then have "{v \ \ . v \ (b1 \ b2)} = (b1 \ b2)" + by blast + then have "card {v \ \ . v \ b1 \ v \ b2} = card (b1 \ b2)" + by simp + thus ?thesis using assms intersection_number_def by metis +qed + +text \Extensions on design operation lemmas \ +lemma del_block_b: + "bl \# \ \ size (del_block bl) = \ - 1" + "bl \# \ \ size (del_block bl) = \" + by (simp_all add: del_block_def size_Diff_singleton) + +lemma del_block_points_index: + assumes "ps \ \" + assumes "card ps = 2" + assumes "bl \# \" + shows "ps \ bl \ points_index (del_block bl) ps = points_index \ ps - 1" + "\ (ps \ bl) \ points_index (del_block bl) ps = points_index \ ps" +proof - + assume "ps \ bl" + then show "points_index (del_block bl) ps = points_index \ ps - 1" + using point_index_diff del_block_def + by (metis assms(3) insert_DiffM2 points_index_singleton) +next + assume "\ ps \ bl" + then show "del_block bl index ps = \ index ps" + using point_index_diff del_block_def + by (metis add_block_def add_block_index_not_in assms(3) insert_DiffM2) +qed + +end + +text \Extensions to properties of design sub types \ + +context finite_incidence_system +begin + +lemma complete_block_size_eq_points: "bl \# \ \ card bl = \ \ bl = \" + using wellformed by (simp add: card_subset_eq finite_sets) + +lemma complete_block_all_subsets: "bl \# \ \ card bl = \ \ ps \ \ \ ps \ bl" + using complete_block_size_eq_points by auto + +lemma del_block_complete_points_index: "ps \ \ \ card ps = 2 \ bl \# \ \ card bl = \ \ + points_index (del_block bl) ps = points_index \ ps - 1" + using complete_block_size_eq_points del_block_points_index(1) by blast + +end + +context design +begin + +lemma block_num_rep_bound: "\ \ (\ x \ \. \ rep x)" +proof - + have exists: "\ bl. bl \# \ \ (\ x \ \ . bl \# {#b \# \. x \ b#})" using wellformed + using blocks_nempty by fastforce + then have bss: "\ \# \\<^sub># (image_mset (\ v. {#b \# \. v \ b#}) (mset_set \))" + proof (intro mset_subset_eqI) + fix bl + show "count \ bl \ count (\v\#mset_set \. filter_mset ((\) v) \) bl" + proof (cases "bl \# \") + case True + then obtain x where xin: "x \ \" and blin: "bl \# filter_mset ((\) x) \" using exists by auto + then have eq: "count \ bl = count (filter_mset ((\) x) \) bl" by simp + have "(\v\#mset_set \. filter_mset ((\) v) \) = (filter_mset ((\) x) \) + + (\v\#(mset_set \) - {#x#}. filter_mset ((\) v) \)" + using xin by (simp add: finite_sets mset_set.remove) + then have "count (\v\#mset_set \. filter_mset ((\) v) \) bl = count (filter_mset ((\) x) \) bl + + count (\v\#(mset_set \) - {#x#}. filter_mset ((\) v) \) bl" + by simp + then show ?thesis using eq by linarith + next + case False + then show ?thesis by (metis count_eq_zero_iff le0) + qed + qed + have "(\ x \ \. \ rep x) = (\ x \ \. size ({#b \# \. x \ b#}))" + by (simp add: point_replication_number_def) + also have "... = (\ x \# (mset_set \). size ({#b \# \. x \ b#}))" + by (simp add: sum_unfold_sum_mset) + also have "... = (\ x \# (image_mset (\ v. {#b \# \. v \ b#}) (mset_set \)) . size x)" + by auto + finally have "(\ x \ \. \ rep x) = size (\\<^sub># (image_mset (\ v. {#b \# \. v \ b#}) (mset_set \)))" + using size_big_union_sum by metis + then show ?thesis using bss + by (simp add: size_mset_mono) +qed + +end + +context proper_design +begin + +lemma del_block_proper: + assumes "\ > 1" + shows "proper_design \ (del_block bl)" +proof - + interpret d: design \ "(del_block bl)" + using delete_block_design by simp + have "d.\ > 0" using del_block_b assms + by (metis b_positive zero_less_diff) + then show ?thesis by(unfold_locales) (auto) +qed + +end + +context simple_design +begin + +lemma inter_num_lt_block_size_strict: + assumes "bl1 \# \" + assumes "bl2 \# \" + assumes "bl1 \ bl2" + assumes "card bl1 = card bl2" + shows "bl1 |\| bl2 < card bl1" "bl1 |\| bl2 < card bl2" +proof - + have lt: "bl1 |\| bl2 \ card bl1" using finite_blocks + by (simp add: \bl1 \# \\ \bl2 \# \\ inter_num_max_bound(1)) + have ne: "bl1 |\| bl2 \ card bl1" + proof (rule ccontr, simp) + assume "bl1 |\| bl2 = card bl1" + then have "bl1 = bl2" using assms(4) inter_eq_blocks_eq_card assms(1) assms(2) finite_blocks + by blast + then show False using assms(3) by simp + qed + then show "bl1 |\| bl2 < card bl1" using lt by simp + have "bl1 |\| bl2 \ card bl2" using ne by (simp add: assms(4)) + then show "bl1 |\| bl2 < card bl2" using lt assms(4) by simp +qed + +lemma block_mset_distinct: "distinct_mset \" using simple + by (simp add: distinct_mset_def) + +end + +context constant_rep_design +begin + +lemma index_lt_const_rep: + assumes "ps \ \" + assumes "ps \ {}" + shows "\ index ps \ \" +proof - + obtain x where xin: "x \ ps" using assms by auto + then have "\ rep x = \" + by (meson assms(1) in_mono rep_number_alt_def_all) + thus ?thesis using index_lt_rep_general xin by auto +qed + +end + +context t_wise_balance +begin + +lemma obtain_t_subset_with_point: + assumes "x \ \" + obtains ps where "ps \ \" and "card ps = \" and "x \ ps" +proof (cases "\ = 1") + case True + have "{x} \ \" "card {x} = 1" "x \ {x}" + using assms by simp_all + then show ?thesis + using True that by blast +next + case False + have "\ - 1 \ card (\ - {x})" + by (simp add: assms diff_le_mono finite_sets t_lt_order) + then obtain ps' where psss: "ps' \ (\ - {x})" and psc: "card ps' = \ - 1" + by (meson obtain_subset_with_card_n) + then have xs: "(insert x ps') \ \" + using assms by blast + have xnotin: "x \ ps'" using psss + by blast + then have "card (insert x ps') = Suc (card ps')" + by (meson \insert x ps' \ \\ finite_insert card_insert_disjoint finite_sets finite_subset) + then have "card (insert x ps') = card ps' + 1" + by presburger + then have xc: "card (insert x ps') = \" using psc + using add.commute add_diff_inverse t_non_zero by linarith + have "x \ (insert x ps')" by simp + then show ?thesis using xs xc that by blast +qed + +lemma const_index_lt_rep: + assumes "x \ \" + shows "\\<^sub>t \ \ rep x" +proof - + obtain ps where psin: "ps \ \" and "card ps = \" and xin: "x \ ps" + using assms t_lt_order obtain_t_subset_with_point by auto + then have "\ index ps = \\<^sub>t " using balanced by simp + thus ?thesis using index_lt_rep_general xin + by (meson) +qed + +end + +context pairwise_balance +begin + +lemma index_zero_iff: "\ = 0 \ (\ bl \# \ . card bl = 1)" +proof (auto) + fix bl assume l0: "\ = 0" assume blin: "bl \# \" + have "card bl = 1" + proof (rule ccontr) + assume "card bl \ 1" + then have "card bl \ 2" using block_size_gt_0 + by (metis Suc_1 Suc_leI blin less_one nat_neq_iff) + then obtain ps where psss: "ps \ bl" and pscard: "card ps = 2" + by (meson obtain_subset_with_card_n) + then have psin: "\ index ps \ 1" + using blin points_index_count_min by auto + have "ps \ \" using wellformed psss blin by auto + then show False using balanced l0 psin pscard by auto + qed + thus "card bl = (Suc 0)" by simp +next + assume a: "\bl\#\. card bl = Suc 0" + obtain ps where psss: "ps \ \" and ps2: "card ps = 2" + by (meson obtain_t_subset_points) + then have "\ bl. bl \# \ \ (card ps > card bl)" using a + by simp + then have cond: "\ bl. bl \# \ \ \( ps \ bl)" + by (metis card_mono finite_blocks le_antisym less_imp_le_nat less_not_refl3) + have "\ index ps = size {# bl \# \ . ps \ bl #}" by (simp add:points_index_def) + then have "\ index ps = size {#}" using cond + by (metis points_index_0_iff size_empty) + thus "\ = 0" using psss ps2 balanced by simp +qed + +lemma count_complete_lt_balance: "count \ \ \ \" +proof (rule ccontr) + assume a: "\ count \ \ \ \" + then have assm: "count \ \ > \" + by simp + then have gt: "size {# bl \# \ . bl = \#} > \" + by (simp add: count_size_set_repr) + obtain ps where psss: "ps \ \" and pscard: "card ps = 2" using t_lt_order + by (meson obtain_t_subset_points) + then have "{# bl \# \ . bl = \#} \# {# bl \# \ . ps \ bl #}" + by (metis a balanced le_refl points_index_count_min) + then have "size {# bl \# \ . bl = \#} \ \ index ps " + using points_index_def[of \ ps] size_mset_mono by simp + thus False using pscard psss balanced gt by auto +qed + +lemma eq_index_rep_imp_complete: + assumes "\ = \ rep x" + assumes "x \ \" + assumes "bl \# \" + assumes "x \ bl" + shows "card bl = \" +proof - + have "\ y. y \ \ \ y \ x \ card {x, y} = 2 \ {x, y} \ \" using assms by simp + then have size_eq: "\ y. y \ \ \ y \ x \ size {# b \# \ . {x, y} \ b#} = size {# b \# \ . x \ b#}" + using point_replication_number_def balanced points_index_def assms by metis + have "\ y b. y \ \ \ y \ x \ b \# \ \ {x, y} \ b \ x \ b" by simp + then have "\ y. y \ \ \ y \ x \ {# b \# \ . {x, y} \ b#} \# {# b \# \ . x \ b#}" + using multiset_filter_mono2 assms by auto + then have eq_sets: "\ y. y \ \ \ y \ x \ {# b \# \ . {x, y} \ b#} = {# b \# \ . x \ b#}" + using size_eq by (smt (z3) Diff_eq_empty_iff_mset cancel_comm_monoid_add_class.diff_cancel + size_Diff_submset size_empty size_eq_0_iff_empty subset_mset.antisym) + have "bl \# {# b \# \ . x \ b#}" using assms by simp + then have "\ y. y \ \ \ y \ x \ {x, y} \ bl" using eq_sets + by (metis (no_types, lifting) Multiset.set_mset_filter mem_Collect_eq) + then have "\ y. y \ \ \ y \ bl" using assms by blast + then have "bl = \" using wellformed assms(3) by blast + thus ?thesis by simp +qed + +lemma incomplete_index_strict_lt_rep: + assumes "\ bl. bl \# \ \ incomplete_block bl" + assumes "x \ \" + assumes "\ > 0" + shows "\ < \ rep x" +proof (rule ccontr) + assume "\ (\ < \ rep x)" + then have a: "\ \ \ rep x" + by simp + then have "\ = \ rep x" using const_index_lt_rep + using assms(2) le_antisym by blast + then obtain bl where xin: "x \ bl" and blin: "bl \# \" + by (metis assms(3) rep_number_g0_exists) + thus False using assms eq_index_rep_imp_complete incomplete_alt_size + using \\ = \ rep x\ nat_less_le by blast +qed + +text \Construct new PBD's from existing PBD's \ + +lemma remove_complete_block_pbd: + assumes "\ \ 2" + assumes "bl \# \" + assumes "card bl = \" + shows "pairwise_balance \ (del_block bl) (\ - 1)" +proof - + interpret pd: proper_design \ "(del_block bl)" using assms(1) del_block_proper by simp + show ?thesis using t_lt_order assms del_block_complete_points_index + by (unfold_locales) (simp_all) +qed + +lemma remove_complete_block_pbd_alt: + assumes "\ \ 2" + assumes "bl \# \" + assumes "bl = \" + shows "pairwise_balance \ (del_block bl) (\ - 1)" + using remove_complete_block_pbd assms by blast + +lemma b_gt_index:"\ \ \" +proof (rule ccontr) + assume blt: "\ \ \ \" + obtain ps where "card ps = 2" and "ps \ \" using t_lt_order + by (meson obtain_t_subset_points) + then have "size {#bl \# \. ps \ bl#} = \" using balanced by (simp add: points_index_def) + thus False using blt by auto +qed + +lemma remove_complete_blocks_set_pbd: + assumes "x < \" + assumes "size A = x" + assumes "A \# \" + assumes "\ a. a \# A \ a = \" + shows "pairwise_balance \ (\ - A) (\ - x)" +using assms proof (induct "x" arbitrary: A) + case 0 + then have beq: "\ - A = \" by simp + have "pairwise_balance \ \ \" by (unfold_locales) + then show ?case using beq by simp +next + case (Suc x) + then have "size A > 0" by simp + let ?A' = "A - {#\#}" + have ss: "?A' \# \" + using Suc.prems(3) by (metis diff_subset_eq_self subset_mset.le_less_trans) + have sx: "size ?A' = x" + by (metis Suc.prems(2) Suc.prems(4) Suc_inject size_Suc_Diff1 size_eq_Suc_imp_elem) + have xlt: "x < \" + by (simp add: Suc.prems(1) Suc_lessD) + have av: "\ a. a \# ?A' \ a = \" using Suc.prems(4) + by (meson in_remove1_mset_neq) + then interpret pbd: pairwise_balance \ "(\ - ?A')" "(\ - x)" using Suc.hyps sx ss xlt by simp + have "Suc x < \" using Suc.prems(3) + by (metis Suc.prems(2) mset_subset_size) + then have "\ - x \ 2" + by linarith + then have bgt: "size (\ - ?A') \ 2" using ss size_Diff_submset + by (metis subset_msetE sx) + have ar: "add_mset \ (remove1_mset \ A) = A" using Suc.prems(2) Suc.prems(4) + by (metis insert_DiffM size_eq_Suc_imp_elem) + then have db: "pbd.del_block \ = \ - A" by(simp add: pbd.del_block_def) + then have "\ - ?A' = \ - A + {#\#}" using Suc.prems(2) Suc.prems(4) + by (metis (no_types, lifting) Suc.prems(3) ar add_diff_cancel_left' add_mset_add_single add_right_cancel + pbd.del_block_def remove_1_mset_id_iff_notin ss subset_mset.lessE trivial_add_mset_remove_iff) + then have "\ \# (\ - ?A')" by simp + then have "pairwise_balance \ (\ - A) (\ - (Suc x))" using db bgt diff_Suc_eq_diff_pred + diff_commute pbd.remove_complete_block_pbd_alt by presburger + then show ?case by simp +qed + +lemma remove_all_complete_blocks_pbd: + assumes "count \ \ < \" + shows "pairwise_balance \ (removeAll_mset \ \) (\ - (count \ \))" (is "pairwise_balance \ ?B ?\") +proof - + let ?A = "replicate_mset (count \ \) \" + let ?x = "size ?A" + have blt: "count \ \ \ \" using b_gt_index assms + by linarith + have xeq: "?x = count \ \" by simp + have av: "\ a. a \# ?A \ a = \" + by (metis in_replicate_mset) + have "?A \# \" + by (meson count_le_replicate_mset_subset_eq le_eq_less_or_eq) + then have "?A \# \" using blt + by (metis subset_mset.nless_le xeq) + thus ?thesis using assms av xeq remove_complete_blocks_set_pbd + by presburger +qed + +end + +context bibd +begin +lemma symmetric_bibdIII: "\ = \ \ symmetric_bibd \ \ \ \" + using necessary_condition_one symmetric_condition_1 by (unfold_locales) (simp) +end + +subsection \ New Design Locales \ +text \ We establish a number of new locales and link them to the existing locale hierarchy +in order to reason in contexts requiring specific combinations of contexts \ + +text \Regular t-wise balance \ +locale regular_t_wise_balance = t_wise_balance + constant_rep_design +begin + +lemma reg_index_lt_rep: + shows "\\<^sub>t \ \" +proof - + obtain ps where psin: "ps \ \" and pst: "card ps = \" + by (metis obtain_t_subset_points) + then have ne: "ps \ {}" using t_non_zero by auto + then have "\ index ps = \\<^sub>t" using balanced pst psin by simp + thus ?thesis using index_lt_const_rep + using ne psin by auto +qed + +end + +locale regular_pairwise_balance = regular_t_wise_balance \ \ 2 \ \ + pairwise_balance \ \ \ + for \ and \ and \ and \ + +text \ Const Intersect Design \ +text \ This is the dual of a balanced design, and used extensively in the remaining formalisation \ + +locale const_intersect_design = proper_design + + fixes \ :: nat + assumes const_intersect: "b1 \# \ \ b2 \# (\ - {#b1#}) \ b1 |\| b2 = \" + +sublocale symmetric_bibd \ const_intersect_design \ \ \ + by (unfold_locales) (simp) + +context const_intersect_design +begin + +lemma inter_num_le_block_size: + assumes "bl \# \" + assumes "\ \ 2" + shows "\ \ card bl" +proof (rule ccontr) + assume a: "\ (\ \ card bl)" + obtain bl' where blin: "bl' \# \ - {#bl#}" + using assms by (metis add_mset_add_single diff_add_inverse2 diff_is_0_eq' multiset_nonemptyE + nat_1_add_1 remove1_mset_eqE size_single zero_neq_one) + then have const: "bl |\| bl' = \" using const_intersect assms by auto + thus False using inter_num_max_bound(1) finite_blocks + by (metis a blin assms(1) finite_blocks in_diffD) +qed + +lemma const_inter_multiplicity_one: + assumes "bl \# \" + assumes "\ < card bl" + shows "multiplicity bl = 1" +proof (rule ccontr) + assume "multiplicity bl \ 1" + then have "multiplicity bl > 1" using assms + by (simp add: le_neq_implies_less) + then obtain bl2 where "bl = bl2" and "bl2 \# \ - {#bl#}" + by (metis count_single in_diff_count) + then have "bl |\| bl2 = card bl" + using inter_num_of_eq_blocks by blast + thus False using assms const_intersect + by (simp add: \bl2 \# remove1_mset bl \\) +qed + +lemma mult_blocks_const_inter: + assumes "bl \# \" + assumes "multiplicity bl > 1" + assumes "\ \ 2" + shows "\ = card bl" +proof (rule ccontr) + assume "\ \ card bl" + then have "\ < card bl" using inter_num_le_block_size assms + using nat_less_le by blast + then have "multiplicity bl = 1" using const_inter_multiplicity_one assms by simp + thus False using assms(2) by simp +qed + +lemma simple_const_inter_block_size: "(\ bl. bl \# \ \ \ < card bl) \ simple_design \ \" + using const_inter_multiplicity_one by (unfold_locales) (simp) + +lemma simple_const_inter_iff: + assumes "\ \ 2" + shows "size {#bl \# \ . card bl = \ #} \ 1 \ simple_design \ \" +proof (intro iffI) + assume a: "size {#bl \# \. card bl = \#} \ 1" + show "simple_design \ \" + proof (unfold_locales) + fix bl assume blin: "bl \# \" + show "multiplicity bl = 1" + proof (cases "card bl = \") + case True + then have m: "multiplicity bl = size {#b \# \ . b = bl#}" + by (simp add: count_size_set_repr) + then have "{#b \# \ . b = bl#} \# {#bl \# \. card bl = \#}" using True + by (simp add: mset_subset_eqI) + then have "size {#b \# \ . b = bl#} \ size {#bl \# \. card bl = \#}" + by (simp add: size_mset_mono) + then show ?thesis using a blin + by (metis count_eq_zero_iff le_neq_implies_less le_trans less_one m) + next + case False + then have "\ < card bl" using assms + by (simp add: blin inter_num_le_block_size le_neq_implies_less) + then show ?thesis using const_inter_multiplicity_one + by (simp add: blin) + qed + qed +next + assume simp: "simple_design \ \" + then have mult: "\ bl. bl \# \ \ multiplicity bl = 1" + using simple_design.axioms(2) simple_incidence_system.simple_alt_def_all by blast + show "size {#bl \# \ . card bl = \ #} \ 1" + proof (rule ccontr) + assume "\ size {#bl \# \. card bl = \#} \ 1" + then have "size {#bl \# \ . card bl = \ #} > 1" by simp + then obtain bl1 bl2 where blin: "bl1 \# \" and bl2in: "bl2 \# \ - {#bl1#}" and + card1: "card bl1 = \" and card2: "card bl2 = \" + using obtain_two_items_mset_filter by blast + then have "bl1 |\| bl2 = \" using const_intersect by simp + then have "bl1 = bl2" + by (metis blin bl2in card1 card2 finite_blocks in_diffD inter_eq_blocks_eq_card) + then have "multiplicity bl1 > 1" + using \bl2 \# remove1_mset bl1 \\ count_eq_zero_iff by force + thus False using mult blin by simp + qed +qed + +lemma empty_inter_implies_rep_one: + assumes "\ = 0" + assumes "x \ \" + shows "\ rep x \ 1" +proof (rule ccontr) + assume a: "\ \ rep x \ 1" + then have gt1: "\ rep x > 1" by simp + then obtain bl1 where blin1: "bl1 \# \" and xin1: "x \ bl1" + by (metis gr_implies_not0 linorder_neqE_nat rep_number_g0_exists) + then have "(\ - {#bl1#}) rep x > 0" using gt1 point_rep_number_split point_rep_singleton_val + by (metis a add_0 eq_imp_le neq0_conv remove1_mset_eqE) + then obtain bl2 where blin2: "bl2 \# (\ - {#bl1#})" and xin2: "x \ bl2" + by (metis rep_number_g0_exists) + then have "x \ (bl1 \ bl2)" using xin1 by simp + then have "bl1 |\| bl2 \ 0" + by (metis blin1 empty_iff finite_blocks intersection_number_empty_iff) + thus False using const_intersect assms blin1 blin2 by simp +qed + +lemma empty_inter_implies_b_lt_v: + assumes "\ = 0" + shows "\ \ \" +proof - + have le1: "\ x. x \ \ \ \ rep x \ 1" using empty_inter_implies_rep_one assms by simp + have disj: "{v \ \ . \ rep v = 0} \ {v \ \ . \ (\ rep v = 0)} = {}" by auto + have eqv: "\ = ({v \ \ . \ rep v = 0} \ {v \ \ . \ (\ rep v = 0)})" by auto + have "\ \ (\ x \ \ . \ rep x)" using block_num_rep_bound by simp + also have 1: "... \ (\ x \ ({v \ \ . \ rep v = 0} \ {v \ \ . \ (\ rep v = 0)}) . \ rep x)" + using eqv by simp + also have "... \ (\ x \ ({v \ \ . \ rep v = 0}) . \ rep x) + (\ x \ ({v \ \ . \ (\ rep v = 0)}) . \ rep x)" + using sum.union_disjoint finite_sets eqv disj + by (metis (no_types, lifting) 1 finite_Un) + also have "... \ (\ x \ ({v \ \ . \ (\ rep v = 0)}) . \ rep x)" by simp + also have "... \ (\ x \ ({v \ \ . \ (\ rep v = 0)}) . 1)" using le1 + by (metis (mono_tags, lifting) mem_Collect_eq sum_mono) + also have "... \ card {v \ \ . \ (\ rep v = 0)}" by simp + also have "... \ card \" using finite_sets + using card_mono eqv by blast + finally show ?thesis by simp +qed + +end + +locale simple_const_intersect_design = const_intersect_design + simple_design + +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/Dual_Systems.thy b/thys/Fishers_Inequality/Dual_Systems.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Dual_Systems.thy @@ -0,0 +1,601 @@ +(* Title: Dual_Systems.thy + Author: Chelsea Edmonds +*) + +section \ Dual Systems \ +text \The concept of a dual incidence system \cite{colbournHandbookCombinatorialDesigns2007} + is an important property in design theory. It enables us to reason on the existence of several +different types of design constructs through dual properties \cite{stinsonCombinatorialDesignsConstructions2004}\ + +theory Dual_Systems imports Incidence_Matrices +begin + +subsection \Dual Blocks \ +text \A dual design of $(\mathcal{V}, \mathcal{B})$, is the design where each block in $\mathcal{B}$ +represents a point $x$, and a block in a dual design is a set of blocks which $x$ is in from the original design. +It is important to note that if a block repeats in $\mathcal{B}$, each instance of the block is a distinct point. +As such the definition below uses each block's list index as its identifier. The list of points would simply be the +indices $0..<$length $Bs$ \ + +definition dual_blocks :: "'a set \ 'a set list \ nat set multiset" where +"dual_blocks \ \s \ {# {y . y < length \s \ x \ \s ! y} . x \# (mset_set \)#}" + +lemma dual_blocks_wf: "b \# dual_blocks V Bs \ b \ {0..s*") where +"dual_blocks_ordered \ map (\ x . {y . y < length \s \ x \ \s ! y}) \s" + +lemma dual_blocks_ordered_eq: "dual_blocks \ \s= mset (\s*)" + by (auto simp add: distinct dual_blocks_def dual_blocks_ordered_def mset_set_set) + +lemma dual_blocks_len: "length \s* = length \s" + by (simp add: dual_blocks_ordered_def) + +text \A dual system is an incidence system \ +sublocale dual_sys: finite_incidence_system "{0..s}" "dual_blocks \ \s" + using dual_blocks_wf by(unfold_locales) (auto) + +lemma dual_is_ordered_inc_sys: "ordered_incidence_system [0..s] \s*" + using inc_sys_orderedI dual_blocks_ordered_eq + by (metis atLeastLessThan_upt distinct_upt dual_sys.incidence_system_axioms) + +interpretation ordered_dual_sys: ordered_incidence_system "[0..s]" "\s*" + using dual_is_ordered_inc_sys by simp + +subsection \Basic Dual Properties\ +lemma ord_dual_blocks_b: "ordered_dual_sys.\ = \" + using dual_blocks_len by (simp add: points_list_length) + +lemma dual_blocks_b: "dual_sys.\ = \" + using points_list_length + by (simp add: dual_blocks_len dual_blocks_ordered_eq) + +lemma dual_blocks_v: "dual_sys.\ = \" + by fastforce + +lemma ord_dual_blocks_v: "ordered_dual_sys.\ = \" + by fastforce + +lemma dual_point_block: "i < \ \ \s* ! i = {y. y < length \s \ (\s ! i) \ \s ! y}" + by (simp add: dual_blocks_ordered_def points_list_length) + +lemma dual_incidence_iff: "i < \ \ j < \ \ \s ! j = bl \ \s ! i = x \ (x \ bl \ j \ \s* ! i)" + using dual_point_block by (intro iffI)(simp_all) + +lemma dual_incidence_iff2: "i < \ \ j < \ \ (\s ! i \ \s ! j \ j \ \s* ! i)" + using dual_incidence_iff by simp + +lemma dual_blocks_point_exists: "bl \# dual_blocks \ \s \ + \ x. x \ \ \ bl = {y . y < length \s \ x \ \s ! y}" + by (auto simp add: dual_blocks_def) + +lemma dual_blocks_ne_index_ne: "j1 < length \s* \ j2 < length \s* \ \s* ! j1 \ \s* ! j2 \ j1 \ j2" + by auto + +lemma dual_blocks_list_index_img: "image_mset (\x . \s* ! x) (mset_set {0..s*}) = mset \s*" + using lessThan_atLeast0 ordered_dual_sys.blocks_list_length ordered_dual_sys.blocks_mset_image + by presburger + +lemma dual_blocks_elem_iff: + assumes "j < \" + shows "x \ (\s* ! j) \ \s ! j \ \s ! x \ x < \" +proof (intro iffI conjI) + show "x \ \s* ! j \ \s ! j \ \s ! x" + using assms ordered_incidence_system.dual_point_block ordered_incidence_system_axioms + by fastforce + show "x \ \s* ! j \ x < \" + using assms dual_blocks_ordered_def dual_point_block by fastforce + show "\s ! j \ \s ! x \ x < \ \ x \ \s* ! j" + by (metis (full_types) assms blocks_list_length dual_incidence_iff) +qed + +text \The incidence matrix of the dual of a design is just the transpose \ +lemma dual_incidence_mat_eq_trans: "ordered_dual_sys.N = N\<^sup>T" +proof (intro eq_matI) + show dimr: "dim_row ordered_dual_sys.N = dim_row N\<^sup>T" using dual_blocks_v by (simp) + show dimc: "dim_col ordered_dual_sys.N = dim_col N\<^sup>T" using ord_dual_blocks_b by (simp) + show "\i j. i < dim_row N\<^sup>T \ j < dim_col N\<^sup>T \ ordered_dual_sys.N $$ (i, j) = N\<^sup>T $$ (i, j)" + proof - + fix i j assume ilt: "i < dim_row N\<^sup>T" assume jlt: "j < dim_col N\<^sup>T" + then have ilt2: "i < length \s"using dimr + using blocks_list_length ord_dual_blocks_v ilt ordered_dual_sys.dim_row_is_v by linarith + then have ilt3: "i < \" by simp + have jlt2: "j < \" using jlt + using dim_row_is_v by fastforce + have "ordered_dual_sys.N $$ (i, j) = (if ([0..s] ! i) \ (\s* ! j) then 1 else 0)" + using dimr dual_blocks_len ilt jlt inc_matrix_elems_one_zero + by (metis inc_mat_dim_row inc_matrix_point_in_block_iff index_transpose_mat(3) ) + then have "ordered_dual_sys.N $$ (i, j) = (if \s ! j \ \s ! i then 1 else 0)" + using ilt3 jlt2 dual_incidence_iff2 by simp + thus "ordered_dual_sys.N $$ (i, j) = N\<^sup>T $$ (i, j)" + using ilt3 jlt2 dim_row_is_v dim_col_is_b N_trans_index_val by simp + qed +qed + +lemma dual_incidence_mat_eq_trans_rev: "(ordered_dual_sys.N)\<^sup>T = N" + using dual_incidence_mat_eq_trans by simp + +subsection \Incidence System Dual Properties\ +text \Many common design properties have a dual in the dual design which enables extensive reasoning +Using incidence matrices and the transpose property these are easy to prove. We leave examples of +counting proofs (commented out), to demonstrate how incidence matrices can significantly simplify +reasoning \ + +lemma dual_blocks_nempty: + assumes "(\ x . x \ \ \ \ rep x > 0)" + assumes "bl \# dual_blocks \ \s" + shows "bl \ {}" +proof - + have "bl \# {# {y . y < length \s \ x \ \s ! y} . x \# (mset_set \)#}" + using assms dual_blocks_def by metis + then obtain x where "x \# (mset_set \)" and blval: "bl = {y . y < length \s \ x \ \s ! y}" + by blast + then obtain bl' where "bl' \# \" and xin: "x \ bl'" using assms(1) + using point_in_block_rep_min_iff by auto + then obtain y where "y < length \s" and "\s ! y = bl'" + using valid_blocks_index_cons by auto + then have "y \ bl" + by (simp add: xin blval) + thus ?thesis by blast +qed + +lemma dual_blocks_size_is_rep: "j < length \s* \ card (\s* ! j) = \ rep (\s ! j)" + using dual_incidence_mat_eq_trans trans_mat_rep_block_size_sym(2) + by (metis dual_blocks_len dual_is_ordered_inc_sys inc_mat_dim_row mat_rep_num_N_row + ordered_incidence_system.mat_block_size_N_col points_list_length size_mset) + +(* Old Counting proof +proof - + have 1: "card (\s* ! j) = card {y . y < length \s \ (\s ! j) \ \s ! y}" + using assms dual_blocks_len dual_point_block points_list_length by force + also have 2: "... = card {y \ {0..s} . (\s ! j) \ \s ! y}" by simp + also have "... = size (mset_set {y \ {0..s} . (\s ! j) \ \s ! y})" by simp + also have "... = size {# y \# mset_set {0..< length \s} . (\s ! j) \ \s ! y #}" + using filter_mset_mset_set by simp + finally have "card (\s* ! j) = size {# bl \# \ . (\s ! j) \ bl #}" + by (metis 1 2 filter_size_blocks_eq_card_indexes lessThan_atLeast0 size_mset) + thus ?thesis by (simp add: point_replication_number_def) +qed +*) + +lemma dual_blocks_size_is_rep_obtain: + assumes "bl \# dual_blocks \ \s" + obtains x where "x \ \" and "card bl = \ rep x" +proof - + obtain j where jlt1: "j < length \s*" and bleq: "\s* ! j = bl" + by (metis assms dual_blocks_ordered_eq in_mset_conv_nth) + then have jlt: "j < \" + by (simp add: dual_blocks_len points_list_length) + let ?x = "\s ! j" + have xin: "?x \ \" using jlt + by (simp add: valid_points_index) + have "card bl = \ rep ?x" using dual_blocks_size_is_rep jlt1 bleq by auto + thus ?thesis using xin that by auto +qed + +lemma dual_blocks_rep_is_size: + assumes "i < length \s" + shows "(mset \s*) rep i = card (\s ! i)" +proof - + have "[0..s] ! i = i" using assms by simp + then have "(mset \s*) rep i = mat_rep_num ordered_dual_sys.N i" + using ordered_dual_sys.mat_rep_num_N_row assms length_upt minus_nat.diff_0 + ordered_dual_sys.points_list_length by presburger + also have "... = mat_block_size (ordered_dual_sys.N)\<^sup>T i" using dual_incidence_mat_eq_trans + trans_mat_rep_block_size_sym(2) by (metis assms inc_mat_dim_col index_transpose_mat(2)) + finally show ?thesis using dual_incidence_mat_eq_trans_rev + by (metis assms blocks_list_length mat_block_size_N_col) +qed + +(* Counting Proof +proof - + have "(mset \s* ) rep i = size {# bl \# (mset \s* ) . i \ bl #}" + by (simp add: point_replication_number_def) + also have 1: "... = size {# bl \# {# {y . y < length \s \ x \ \s ! y} . x \# (mset_set \)#} . i \ bl #}" + using dual_blocks_ordered_eq dual_blocks_def by metis + also have "... = size (filter_mset (\ bl . i \ bl) + (image_mset (\ x . {y . y < length \s \ x \ \s ! y}) (mset_set \)))" by simp + finally have "(mset \s* ) rep i = size (image_mset (\ x . {y . y < length \s \ x \ \s ! y}) + (filter_mset (\ bl . i \ {y . y < length \s \ bl \ \s ! y}) (mset_set \)))" + using filter_mset_image_mset by (metis 1 ordered_dual_sys.point_rep_number_alt_def) + then have "(mset \s* ) rep i = size (filter_mset (\ bl . i \ {y . y < length \s \ bl \ \s ! y}) + (mset_set \))" + by fastforce + then have "(mset \s* ) rep i = size (filter_mset (\ bl . bl \ \s ! i) (mset_set \))" + using assms by simp + then have "(mset \s* ) rep i = card {x \ \ . x \ (\s ! i)}" by simp + thus ?thesis using assms block_size_alt by auto +qed +*) + +lemma dual_blocks_inter_index: + assumes "j1 < length \s*" "j2 < length \s*" + shows "(\s* ! j1) |\| (\s* ! j2) = points_index \ {\s ! j1, \s ! j2}" +proof - + have assms2: "j1 < \" "j2 < \" using assms + by (simp_all add: dual_blocks_len points_list_length) + have "(\s* ! j1) |\| (\s* ! j2) = mat_inter_num (ordered_dual_sys.N) j1 j2" + by (simp add: assms(1) assms(2) ordered_dual_sys.mat_inter_num_conv) + also have "... = mat_point_index N {j1, j2}" using dual_incidence_mat_eq_trans_rev trans_mat_point_index_inter_sym(2) + by (metis assms inc_mat_dim_col) + finally show ?thesis using assms2 incidence_mat_two_index + by presburger +qed +(* Counting Proof + have fin: "finite {0..s}" + by auto + have j1lt: "j1 < \" using assms + using dual_blocks_len points_list_length by auto + have j2lt: "j2 < \" using assms dual_blocks_len points_list_length by auto + have iff: "\ x. (x \(\s* ! j1) \ x \ (\s* ! j2)) \ (\s ! j1 \ \s ! x \ x < \ \ \s ! j2 \ \s ! x)" + by (auto simp add: dual_blocks_elem_iff j1lt j2lt) + have pi: "points_index \ {\s ! j1, \s ! j2} = size {# bl \# \ . \s !j1 \ bl \ \s ! j2 \ bl#}" + by (auto simp add: points_index_def) + have "(\s* ! j1) |\| (\s* ! j2) = card ({x . x s \ x \ (\s* ! j1) \ x \ (\s* ! j2)})" + apply (auto simp add: intersection_number_def) + by (smt (verit) Collect_cong Int_Collect blocks_list_length dual_blocks_elem_iff inf.idem inf_set_def j2lt mem_Collect_eq) + then have "(\s* ! j1) |\| (\s* ! j2) = card ({x . x s \ \s ! j1 \ \s ! x \ \s ! j2 \ \s ! x})" using iff + size_mset by (smt (verit, best) Collect_cong) + then have "(\s* ! j1) |\| (\s* ! j2) = size (mset_set {x \ {0..s}. \s ! j1 \ \s ! x \ \s ! j2 \ \s ! x})" by simp + then have "(\s* ! j1) |\| (\s* ! j2) = size ({#x \# mset_set {0..s}. \s ! j1 \ \s ! x \ \s ! j2 \ \s ! x#})" using fin by simp + then have "(\s* ! j1) |\| (\s* ! j2) = size (filter_mset (\ x . \s ! j1 \ \s ! x \ \s ! j2 \ \s ! x) (mset_set {0..s}))" by simp + then have "(\s* ! j1) |\| (\s* ! j2) = size (image_mset (\ i. \s ! i) (filter_mset (\ x . \s ! j1 \ \s ! x \ \s ! j2 \ \s ! x) (mset_set {0..s})))" + by simp + then have "(\s* ! j1) |\| (\s* ! j2) = size (filter_mset (\ x . \s ! j1 \ x \ \s ! j2 \ x) (image_mset (\ i. \s ! i) (mset_set {0..s})))" + by (simp add: filter_mset_image_mset) + then have "(\s* ! j1) |\| (\s* ! j2) = size {# bl \# \ . \s !j1 \ bl \ \s ! j2 \ bl#}" + by (metis blocks_list_length blocks_mset_image lessThan_atLeast0) + thus ?thesis using pi by simp +qed +*) + +lemma dual_blocks_points_index_inter: + assumes "i1 < \" "i2 < \" + shows "(mset \s*) index {i1, i2} = (\s ! i1) |\| (\s ! i2)" +proof - + have "(mset \s*) index {i1, i2} = mat_point_index (ordered_dual_sys.N) {i1, i2}" + using assms(1) assms(2) blocks_list_length ord_dual_blocks_v ordered_dual_sys.dim_row_is_v + ordered_dual_sys.incidence_mat_two_index ordered_dual_sys.mat_ord_inc_sys_point by presburger + also have "... = mat_inter_num N i1 i2" using dual_incidence_mat_eq_trans trans_mat_point_index_inter_sym(1) + by (metis assms(1) assms(2) dual_incidence_mat_eq_trans_rev ord_dual_blocks_v ordered_dual_sys.dim_row_is_v) + finally show ?thesis using mat_inter_num_conv + using assms(1) assms(2) by auto +qed + +(* Counting Proof +proof - + have "\ j . j \# mset_set {0..s*} \ j < \" + by (metis atLeastLessThan_iff atLeastLessThan_upt dual_blocks_len mset_upt points_list_length set_mset_mset) + then have iff: "\ j i. j \# mset_set {0..s*} \ i < \ \ i \ (\s* ! j) \ (\s ! j) \ (\s ! i)" + using assms dual_incidence_iff2 by simp + then have iff2: "\ j . j \# mset_set {0..s*} \ i1 \ (\s* ! j) \ i2 \ (\s* ! j) \ (\s ! j) \ (\s ! i1) \ (\s ! j) \ (\s ! i2)" + using assms by auto + have ss2: "(\s ! i2) \ \" using wellformed assms by auto + then have ss: "{x . x \ (\s ! i1) \ x \ (\s ! i2)} \ \" + by auto + then have inter: "(\s ! i1) |\| (\s ! i2) = card {x \ \. x \ (\s ! i1) \ x \ (\s ! i2)}" + using intersection_number_def by (metis Collect_conj_eq Collect_mem_eq Int_absorb1) + have inj: "inj_on (\ j. \s ! j) {j \ {0..s} . (\s ! j) \ (\s ! i1) \ (\s ! j) \ (\s ! i2)}" + by (simp add: inj_on_nth distinct) + have init: "(mset \s* ) index {i1, i2} = size {#bl \# (mset \s* ) . i1 \ bl \ i2 \ bl#}" + by (simp add: points_index_def) + then have "size {#bl \# (mset \s* ) . i1 \ bl \ i2 \ bl#} = size {#j \# mset_set {0..s*} . i1 \ (\s* ! j) \ i2 \ (\s* ! j)#}" + proof - + have "size {#j \# mset_set {0..s*} . i1 \ (\s* ! j) \ i2 \ (\s* ! j)#} + = size (filter_mset (\ j. i1 \ (\s* ! j) \ i2 \ (\s* ! j)) (mset_set {0..s*})) " by simp + also have s1: "... = size (image_mset (\x . \s* ! x) (filter_mset (\ j. i1 \ (\s* ! j) \ i2 \ (\s* ! j)) (mset_set {0..s*})))" by fastforce + also have s2: "... = size (filter_mset (\ j. i1 \ j \ i2 \ j) (image_mset (\x . \s* ! x) (mset_set {0..s*})))" + by (simp add: filter_mset_image_mset) + finally have "size {#j \# mset_set {0..s*} . i1 \ (\s* ! j) \ i2 \ (\s* ! j)#} = size (filter_mset (\ j. i1 \ j \ i2 \ j) (mset \s* ))" + using dual_blocks_list_index_img s2 s1 by presburger + thus ?thesis by simp + qed + then have "size {#bl \# (mset \s* ) . i1 \ bl \ i2 \ bl#} = size {#j \# mset_set {0..s*} . (\s ! j) \ (\s ! i1) \ (\s ! j) \ (\s ! i2)#}" using iff2 + by (smt (verit, ccfv_SIG) filter_mset_cong) + then have "size {#bl \# (mset \s* ) . i1 \ bl \ i2 \ bl#} = + card ({j \ {0..s*} . (\s ! j) \ (\s ! i1) \ (\s ! j) \ (\s ! i2)})" by simp + then have "size {#bl \# (mset \s* ) . i1 \ bl \ i2 \ bl#} = + card ({j \ {0..s} . (\s ! j) \ (\s ! i1) \ (\s ! j) \ (\s ! i2)})" using dual_blocks_len by presburger + then have "size {#bl \# (mset \s* ) . i1 \ bl \ i2 \ bl#} = + card (image (\ j. \s ! j) {j \ {0..s} . (\s ! j) \ (\s ! i1) \ (\s ! j) \ (\s ! i2)})" + using inj card_image[of "(\ j. \s ! j)" "{j \ {0..s} . (\s ! j) \ (\s ! i1) \ (\s ! j) \ (\s ! i2)}"] by simp + then have "size {#bl \# (mset \s* ) . i1 \ bl \ i2 \ bl#} = + card {j \ image (\ j. \s ! j) {0..s}. j \ (\s ! i1) \ j \ (\s ! i2)}" + using Compr_image_eq[of "(\ j. \s ! j)" "{0..s}" "(\ j . j \ (\s ! i1) \ j \ (\s ! i2))"] by simp + then have "size {#bl \# (mset \s* ) . i1 \ bl \ i2 \ bl#} = + card {j \ \. j \ (\s ! i1) \ j \ (\s ! i2)}" + using lessThan_atLeast0 points_list_length points_set_index_img by presburger + thus ?thesis using init inter by simp +qed*) +end + +subsection \Dual Properties for Design sub types \ +context ordered_design +begin + +lemma dual_is_design: + assumes "(\ x . x \ \ \ \ rep x > 0)" \ \ Required to ensure no blocks are empty \ + shows "design {0..s} (dual_blocks \ \s)" + using dual_blocks_nempty assms by (unfold_locales) (simp) +end + +context ordered_proper_design +begin + +lemma dual_sys_b_non_zero: "dual_sys.\ \ 0" + using v_non_zero dual_blocks_b by auto + +lemma dual_is_proper_design: + assumes "(\ x . x \ \ \ \ rep x > 0)" \ \ Required to ensure no blocks are empty \ + shows "proper_design {0..s} (dual_blocks \ \s)" + using dual_blocks_nempty dual_sys_b_non_zero assms by (unfold_locales) (simp_all) + +end + +context ordered_block_design +begin + +lemma dual_blocks_const_rep: "i \ {0..s} \ (mset \s*) rep i = \" + using dual_blocks_rep_is_size uniform by (metis atLeastLessThan_iff nth_mem_mset) + +lemma dual_blocks_constant_rep_design: + assumes "(\ x . x \ \ \ \ rep x > 0)" + shows "constant_rep_design {0..s} (dual_blocks \ \s) \" +proof - + interpret des: proper_design "{0..s}" "(dual_blocks \ \s)" + using dual_is_proper_design assms by simp + show ?thesis using dual_blocks_const_rep dual_blocks_ordered_eq by (unfold_locales) (simp) +qed + + +end + +context ordered_constant_rep +begin + +lemma dual_blocks_const_size: "j < length \s* \ card (\s* ! j) = \" + using dual_blocks_rep_is_size dual_blocks_len dual_blocks_size_is_rep by fastforce + +lemma dual_is_block_design: "block_design {0..s} (dual_blocks \ \s) \" +proof - + have "\ > 0" by (simp add: r_gzero) + then have "(\ x . x \ \ \ \ rep x > 0)" using rep_number by simp + then interpret pdes: proper_design "{0..s}" "(dual_blocks \ \s)" + using dual_is_proper_design by simp + have "\ bl. bl \# dual_blocks \ \s \ card bl = \" + using dual_blocks_const_size + by (metis dual_blocks_ordered_eq in_set_conv_nth set_mset_mset) + thus ?thesis by (unfold_locales) (simp) +qed + +end + +context ordered_pairwise_balance +begin + +lemma dual_blocks_const_intersect: + assumes "j1 < length \s*" "j2 < length \s*" + assumes "j1 \ j2" + shows "(\s* ! j1) |\| (\s* ! j2) = \" +proof - + have "\s ! j1 \ \s ! j2" using assms(3) + using assms(1) assms(2) distinct dual_blocks_len nth_eq_iff_index_eq by auto + then have c: "card {\s ! j1, \s ! j2} = 2" + using card_2_iff by blast + have ss: "{\s ! j1, \s ! j2} \ \" using assms points_list_length + using dual_blocks_len by auto + have "(\s* ! j1) |\| (\s* ! j2) = points_index \ {\s ! j1, \s ! j2}" + using dual_blocks_inter_index assms by simp + thus ?thesis using ss c balanced + by blast +qed + +lemma dual_is_const_intersect_des: + assumes "\ > 0" + shows "const_intersect_design {0..<(length \s)} (dual_blocks \ \s) \" +proof - + have "(\ x . x \ \ \ \ rep x \ \)" using const_index_lt_rep by simp + then have "(\ x . x \ \ \ \ rep x > 0)" using assms + by (metis gr_zeroI le_zero_eq) + then interpret pd: proper_design "{0..<(length \s)}" "(dual_blocks \ \s)" + using dual_is_proper_design by (simp) + show ?thesis proof (unfold_locales) + fix b1 b2 + assume b1in: "b1 \# dual_blocks \ \s" + assume b2in: "b2 \# remove1_mset b1 (dual_blocks \ \s)" + obtain j1 where b1eq: "b1 = \s* ! j1" and j1lt: "j1 < length \s*" using b1in + by (metis dual_blocks_ordered_eq in_set_conv_nth set_mset_mset) + obtain j2 where b2eq: "b2 = \s* ! j2" and j2lt: "j2 < length \s*" and "j1 \ j2" + using b2in index_remove1_mset_ne + by (metis (mono_tags) b1eq dual_blocks_ordered_eq j1lt nth_mem set_mset_mset) + then show "b1 |\| b2 = \" + using dual_blocks_const_intersect b1eq b2eq j1lt j2lt by simp + qed +qed + + +lemma dual_is_simp_const_inter_des: + assumes "\ > 0" + assumes "\ bl. bl \# \ \ incomplete_block bl" + shows "simple_const_intersect_design {0..<(length \s)} (dual_blocks \ \s) \" +proof - + interpret d: const_intersect_design "{0..<(length \s)}" "(dual_blocks \ \s)" "\" + using assms dual_is_const_intersect_des by simp + \ \ Show that m < block size for all blocks \ + have "\ x. x \ \ \ \ < \ rep x" using assms incomplete_index_strict_lt_rep + by blast + then have "\ bl. bl \# (dual_blocks \ \s) \ \ < card bl" + by (metis dual_blocks_size_is_rep_obtain) + then interpret s: simple_design "{0..<(length \s)}" "(dual_blocks \ \s)" + using d.simple_const_inter_block_size by simp + show ?thesis by (unfold_locales) +qed +end + +context ordered_const_intersect_design +begin + +lemma dual_is_balanced: + assumes "ps \ {0..s}" + assumes "card ps = 2" + shows "(dual_blocks \ \s) index ps = \" +proof - + obtain i1 i2 where psin: "ps = {i1, i2}" and neq: "i1 \ i2" using assms + by (meson card_2_iff) + then have lt: "i1 < \" using assms + by (metis atLeastLessThan_iff blocks_list_length insert_subset) + have lt2: "i2 < \" using assms psin + by (metis atLeastLessThan_iff blocks_list_length insert_subset) + then have inter: "(dual_blocks \ \s) index ps = (\s ! i1) |\| (\s ! i2)" using dual_blocks_points_index_inter neq lt + using dual_blocks_ordered_eq psin by presburger + have inb1: "(\s ! i1) \# \" + using lt by auto + have inb2: "(\s ! i2) \# (\ - {#(\s ! i1)#})" using lt2 neq blocks_index_ne_belong + by (metis blocks_list_length lt) + thus ?thesis using const_intersect inb1 inb2 inter by blast +qed + +lemma dual_is_pbd: + assumes "(\ x . x \ \ \ \ rep x > 0)" + assumes "\ \ 2" + shows "pairwise_balance {0..<(length \s)} (dual_blocks \ \s) \" +proof - + interpret pd: proper_design "{0..<(length \s)}" "(dual_blocks \ \s)" + using dual_is_proper_design + by (simp add: assms) + show ?thesis proof (unfold_locales) + show "(1 ::nat) \ 2" by simp + then show "2 \ dual_sys.\" using assms(2) + by fastforce + show "\ps. ps \ {0..s} \ card ps = 2 \ dual_blocks \ \s index ps = \" + using dual_is_balanced by simp + qed +qed + +end + +context ordered_sym_bibd +begin + +lemma dual_is_balanced: + assumes "ps \ {0..s}" + assumes "card ps = 2" + shows "(dual_blocks \ \s) index ps = \" +proof - + obtain i1 i2 where psin: "ps = {i1, i2}" and neq: "i1 \ i2" + using assms by (meson card_2_iff) + then have lt: "i1 < \" using assms + by (metis atLeastLessThan_iff blocks_list_length insert_subset) + have lt2: "i2 < \" using assms psin + by (metis atLeastLessThan_iff blocks_list_length insert_subset) + then have inter: "(dual_blocks \ \s) index ps = (\s ! i1) |\| (\s ! i2)" + using dual_blocks_points_index_inter neq lt dual_blocks_ordered_eq psin by presburger + have inb1: "(\s ! i1) \# \" + using lt by auto + have inb2: "(\s ! i2) \# (\ - {#(\s ! i1)#})" using lt2 neq blocks_index_simp_unique + by (metis blocks_list_length in_remove1_mset_neq lt valid_blocks_index) + thus ?thesis using sym_block_intersections_index inb1 inter by blast +qed + +lemma dual_bibd: "bibd {0..<(length \s)} (dual_blocks \ \s) \ \" +proof - + interpret block: block_design "{0..<(length \s)}" "(dual_blocks \ \s)" \ + using dual_is_block_design by simp + show ?thesis proof (unfold_locales) + show "\ < dual_sys.\" + using dual_blocks_v incomplete symmetric_condition_1 symmetric_condition_2 by presburger + show "(1 ::nat) \ 2" by simp + have "\ \ 2" + by (simp add: t_lt_order) + then have "\ \ 2" using local.symmetric by auto + then show "2 \ dual_sys.\" by simp + show "\ps. ps \ {0..s} \ card ps = 2 \ dual_blocks \ \s index ps = \" + using dual_is_balanced by simp + show "2 \ \" using r_ge_two by blast + qed +qed + +text \The dual of a BIBD must by symmetric \ + +lemma dual_bibd_symmetric: "symmetric_bibd {0..<(length \s)} (dual_blocks \ \s) \ \" +proof - + interpret bibd: bibd "{0..<(length \s)}" "(dual_blocks \ \s)" \ \ + using dual_bibd by simp + show ?thesis using dual_blocks_b local.symmetric by (unfold_locales) (simp) +qed + +end + +subsection \Generalise Dual Concept \ +text \The above formalisation relies on one translation of a dual design. However, any design +with an ordering of points and blocks such that the matrix is the transpose of the original is +a dual. The definition below encapsulates this concept. Additionally, we prove an isomorphism +exists between the generated dual from @{term "dual_blocks"} and any design satisfying the is dual +definition\ + +context ordered_incidence_system +begin + +definition is_dual:: "'b list \ 'b set list \ bool" where +"is_dual Vs' Bs' \ ordered_incidence_system Vs' Bs' \ (inc_mat_of Vs' Bs' = N\<^sup>T)" + +lemma is_dualI: + assumes "ordered_incidence_system Vs' Bs'" + assumes "(inc_mat_of Vs' Bs' = N\<^sup>T)" + shows "is_dual Vs' Bs'" + by (auto simp add: is_dual_def assms) + +lemma is_dualD1: + assumes "is_dual Vs' Bs'" + shows "(inc_mat_of Vs' Bs' = N\<^sup>T)" + using is_dual_def assms + by auto + +lemma is_dualD2: + assumes "is_dual Vs' Bs'" + shows "ordered_incidence_system Vs' Bs'" + using is_dual_def assms + by auto + +lemma generated_is_dual: "is_dual [0..<(length \s)] \s*" +proof - + interpret osys: ordered_incidence_system "[0..<(length \s)]" "\s*" using dual_is_ordered_inc_sys by simp + show ?thesis using is_dual_def + by (simp add: is_dual_def dual_incidence_mat_eq_trans osys.ordered_incidence_system_axioms) +qed + +lemma is_dual_isomorphism_generated: + assumes "is_dual Vs' Bs'" + shows "\ \. incidence_system_isomorphism (set Vs') (mset Bs') ({0..<(length \s)}) (dual_blocks \ \s) \" +proof - + interpret os2: ordered_incidence_system "([0..<(length \s)])" "(\s*)" + by (simp add: dual_is_ordered_inc_sys) + interpret os1: ordered_incidence_system Vs' Bs' using assms + by (simp add: is_dualD2) + interpret tos: two_ordered_sys Vs' Bs' "([0..<(length \s)])" "(\s*)" + using assms ordered_incidence_system_axioms two_ordered_sys.intro + by (simp add: is_dualD2 two_ordered_sys.intro dual_is_ordered_inc_sys) + have os2V: "os2.\ = {0..<(length \s)}" + by auto + have os2B: "os2.\ = dual_blocks \ \s" + by (simp add: dual_blocks_ordered_eq) + have "os1.N = inc_mat_of Vs' Bs'" by simp + then have "os2.N = os1.N" + using assms is_dualD1 dual_incidence_mat_eq_trans by fastforce + thus ?thesis using tos.equal_inc_mat_isomorphism_ex os2V os2B by auto +qed + +interpretation ordered_dual_sys: ordered_incidence_system "[0..s]" "\s*" + using dual_is_ordered_inc_sys by simp + +text \Original system is dual of the dual \ +lemma is_dual_rev: "ordered_dual_sys.is_dual \s \s" + by (simp add: dual_incidence_mat_eq_trans_rev ordered_dual_sys.is_dualI ordered_incidence_system_axioms) + +end + +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/Fishers_Inequality.thy b/thys/Fishers_Inequality/Fishers_Inequality.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Fishers_Inequality.thy @@ -0,0 +1,338 @@ +(* Title: Fishers_Inequality.thy + Author: Chelsea Edmonds +*) + +section \Fisher's Inequality\ + +text \This theory presents the proof of Fisher's Inequality \cite{fisherExaminationDifferentPossible1940a} + on BIBD's (i.e. uniform Fisher's) and the generalised nonuniform Fisher's Inequality \ +theory Fishers_Inequality imports Rank_Argument_General Linear_Bound_Argument +begin + +subsection \ Uniform Fisher's Inequality \ +context ordered_bibd +begin + +text \Row/Column transformation steps \ + +text\Following design theory lecture notes from MATH3301 at the University of Queensland + \cite{HerkeLectureNotes2016}, a simple transformation to produce an upper triangular matrix using elementary +row operations is to (1) Subtract the first row from each other row, and (2) add all columns to the first column\ + +lemma transform_N_step1_vals: + defines mdef: "M \ (N * N\<^sup>T)" + assumes "i < dim_row M" + assumes "j < dim_col M" + shows "i = 0 \ j = 0 \ (add_row_to_multiple (-1) [1..)" \ \ top left elem \ + and "i \ 0 \ j = 0 \ (add_row_to_multiple (-1) [1..) - (int \)" \ \ first column ex. 1 \ + and "i = 0 \ j \ 0 \ (add_row_to_multiple (-1) [1..)" \ \ first row ex. 1 \ + and "i \ 0 \ j \ 0 \ i = j \ (add_row_to_multiple (-1) [1..) - (int \)" \ \ diagonal ex. 1 \ + and "i \ 0 \ j \ 0 \ i \ j \ (add_row_to_multiple (-1) [1.. \ everything else \ + using transpose_N_mult_diag v_non_zero assms +proof (simp) + show "i = 0 \ j \ 0 \ (add_row_to_multiple (-1) [1..)" + using transpose_N_mult_off_diag v_non_zero assms transpose_N_mult_dim(2) by force +next + assume a: "j = 0" "i\0" + then have ail: "((-1) * M $$(0, j)) = -(int \)" + using transpose_N_mult_diag v_non_zero mdef by auto + then have ijne: "j \ i" using a by simp + then have aij: "M $$ (i, j) = (int \)" using assms(2) mdef transpose_N_mult_off_diag a v_non_zero + by (metis transpose_N_mult_dim(1)) + then have "add_row_to_multiple (-1) [1..) + (int \)" + using ail add_first_row_to_multiple_index(2) assms(2) assms(3) a by (metis mult_minus1) + then show "(add_row_to_multiple (-1) [1..) - (int \)" + by linarith +next + assume a: "i \ 0" "j \ 0" + have ail: "((-1) * M $$(0, j)) = -(int \)" + using assms transpose_N_mult_off_diag a v_non_zero transpose_N_mult_dim(1) by auto + then have "i = j \ M $$ (i, j) = (int \)" + using assms transpose_N_mult_diag a v_non_zero by (metis transpose_N_mult_dim(1)) + then show "i = j \ (add_row_to_multiple (-1) [1..) - (int \)" + using ail add_first_row_to_multiple_index(2) assms a by (metis uminus_add_conv_diff) + then have "i \ j \ M $$ (i, j) = (int \)" using assms transpose_N_mult_off_diag a v_non_zero + by (metis transpose_N_mult_dim(1) transpose_N_mult_dim(2)) + then show "i \ j \ (add_row_to_multiple (-1) [1.. (add_row_to_multiple (-1) [1..T)] 0 (N * N\<^sup>T))" + assumes "i < dim_row (M)" + assumes "j < dim_col (M)" + shows "i = 0 \ j = 0 \ add_multiple_cols 1 0 [1..) + (int \) * (\ - 1)" \ \ top left element \ + and "i = 0 \ j \ 0 \ add_multiple_cols 1 0 [1..)" \ \ top row \ + and "i \ 0 \ i = j \ add_multiple_cols 1 0 [1..) - (int \)" \ \ Diagonal \ + and "i \ 0 \ i \ j \ add_multiple_cols 1 0 [1.. \Everything else\ +proof - + show "i = 0 \ j \ 0 \ add_multiple_cols 1 0 [1..)" + using add_all_cols_to_first assms transform_N_step1_vals(3) by simp + show "i \ 0 \ i = j \ add_multiple_cols 1 0 [1..) - (int \)" + using add_all_cols_to_first assms transform_N_step1_vals(4) by simp +next + assume a: "i = 0" "j =0" + then have size: "card {1.. - 1" using assms by simp + have val: "\ l . l \ {1.. M $$ (i, l) = (int \)" + using mdef transform_N_step1_vals(3) by (simp add: a(1)) + have "add_multiple_cols 1 0 [1..l\{1..l\{1..)) + M$$(i,0)" using val by simp + also have "... = (\ - 1) * (int \) + M$$(i,0)" using size by (metis sum_constant) + finally show "add_multiple_cols 1 0 [1..) + (int \) * (\ - 1)" + using transform_N_step1_vals(1) a(1) a(2) assms(1) assms(2) by simp +next + assume a: "i \ 0" "i \ j" + then show "add_multiple_cols 1 0 [1.. 0") + case True + then show ?thesis using add_all_cols_to_first assms a transform_N_step1_vals(5) by simp + next + case False + then have iin: "i \ {1.. l . l \ {1.. l T) \ l \ 0" using assms by simp + then have val: "\ l . l \ {1.. M $$ (i, l) = 0" + using assms(3) transform_N_step1_vals(5) a False assms(1) + by (metis DiffE iin index_mult_mat(2) index_mult_mat(3) index_transpose_mat(3) insertI1) + have val2: "M $$ (i, i) = (int \) - (int \)" using mdef transform_N_step1_vals(4) a False + assms(1) transpose_N_mult_dim(1) transpose_N_mult_dim(2) + by (metis cond iin) + have val3: "M$$ (i, 0) = (int \) - (int \)" + using assms(3) transform_N_step1_vals(2) a False assms(1) assms(2) + by (metis add_row_to_multiple_dim(1) transpose_N_mult_dim(2) v_non_zero) + have "add_multiple_cols 1 0 [1..l\{1..l\{1..Transformed matrix is upper triangular \ +lemma transform_upper_triangular: + defines mdef: "M \ (add_row_to_multiple (-1) [1..T)] 0 (N * N\<^sup>T))" + shows "upper_triangular (add_multiple_cols 1 0 [1..Find the determinant of the $NN^T$ matrix using transformed matrix values\ +lemma determinant_inc_mat_square: "det (N * N\<^sup>T) = (\ + \ * (\ - 1))* (\ - \)^(\ - 1)" +proof - +\ \ Show the matrix is now lower triangular, therefore the det is the product of the sum of diagonal \ + have cm: "(N * N\<^sup>T) \ carrier_mat \ \" + using transpose_N_mult_dim(1) transpose_N_mult_dim(2) by blast + define C where "C \(add_row_to_multiple (-1) [1..T)] 0 (N * N\<^sup>T))" + have "0 \ set [1..T)]" by simp + then have detbc: "det (N * N\<^sup>T) = det C" + using C_def add_row_to_multiple_det v_non_zero by (metis cm) + define D where "D \ add_multiple_cols 1 0 [1..) + (int \) * (\ - 1))" using transform_N_step2_vals(1) + by (simp add: C_def D_def v_non_zero) + have ine0: "\ i. i \ {1.. i \ 0" by simp + have "\ i. i \ {1.. i < dim_row (N * N\<^sup>T)" using D_def C_def by simp + then have diagnon0: "\ i. i \ {1..<\} \ D $$ (i, i) = (int \) - (int \)" + using transform_N_step2_vals(3) ine0 D_def C_def by simp (* Slow *) + have alll: "\ l. l \ set [1.. l < \" using C_def by simp + have cmc: "C \ carrier_mat \ \" using cm C_def + by (simp add: add_row_to_multiple_carrier) + have dimgt2: "dim_row D \ 2" + using t_lt_order D_def C_def by (simp) + then have fstterm: "0 \ { 0 ..< dim_row D}" by (simp add: points_list_length) + have "0 \ set [1..T) = det D" using add_multiple_cols_det alll cmc D_def C_def + by (metis detbc) + also have "... = prod_list (diag_mat D)" using det_upper_triangular + using transform_upper_triangular D_def C_def by fastforce + also have "... = (\ i = 0 ..< dim_row D. D $$ (i,i))" using prod_list_diag_prod by blast + also have "... = (\ i = 0 ..< \. D $$ (i,i))" by (simp add: D_def C_def) + finally have "det (N * N\<^sup>T) = D $$ (0, 0) * (\ i = 1 ..< \. D $$ (i,i))" + using dimgt2 by (simp add: prod.atLeast_Suc_lessThan v_non_zero) + then have "det (N * N\<^sup>T) = (\ + \ * (\ - 1)) * ((int \) - (int \))^(\ - 1)" + using d00 diagnon0 by simp + then have "det (N * N\<^sup>T) = (\ + \ * (\ - 1)) * ( \ - \)^(\ - 1)" + using index_lt_replication + by (metis (no_types, lifting) less_imp_le_nat of_nat_diff of_nat_mult of_nat_power) + then show ?thesis by blast +qed + +text \Fisher's Inequality using the rank argument. +Note that to use the rank argument we must first map N to a real matrix. It is useful to explicitly +include the parameters which should be used in the application of the @{thm [source] "rank_argument_det"} lemma \ +theorem Fishers_Inequality_BIBD: "\ \ \" +proof (intro rank_argument_det[of "map_mat real_of_int N" "\" "\"], simp_all) + show "N \ carrier_mat \ (length \s)" using blocks_list_length N_carrier_mat by simp + let ?B = "map_mat (real_of_int) (N * N\<^sup>T)" + have b_split: "?B = map_mat (real_of_int) N * (map_mat (real_of_int) N)\<^sup>T" + using semiring_hom.mat_hom_mult of_int_hom.semiring_hom_axioms transpose_carrier_mat map_mat_transpose + by (metis (no_types, lifting) N_carrier_mat) + have db: "det ?B = (\ + \ * (\ - 1))* (\ - \)^(\ - 1)" + using determinant_inc_mat_square by simp + have lhn0: "(\ + \ * (\ - 1)) > 0" + using r_gzero by blast + have "(\ - \) > 0" + using index_lt_replication zero_less_diff by blast + then have det_not_0: "det ?B \ 0" using lhn0 db + by (metis gr_implies_not0 mult_is_0 of_nat_eq_0_iff power_not_zero) + thus "det (of_int_hom.mat_hom N * (of_int_hom.mat_hom N)\<^sup>T) \ (0:: real)" using b_split by simp +qed + +end + +subsection \Generalised Fisher's Inequality \ + +context simp_ordered_const_intersect_design +begin + +text \Lemma to reason on sum coefficients \ +lemma sum_split_coeffs_0: + fixes c :: "nat \ real" + assumes "\ \ 2" + assumes "\ > 0" + assumes "j' < \" + assumes "0 = (\ j \ {0..<\} . (c j)^2 * ((card (\s ! j))- (int \))) + + \ * ((\ j \ {0..<\} . c j)^2)" + shows "c j' = 0" +proof (rule ccontr) + assume cine0: "c j' \ 0" + have innerge: "\ j . j < \ \ (c j)^2 * (card (\s ! j) - (int \)) \ 0" + using inter_num_le_block_size assms(1) by simp + then have lhsge: "(\ j \ {0..<\} . (c j)^2 * ((card (\s ! j))- (int \))) \ 0" + using sum_bounded_below[of "{0..<\}" "0" "\ i. (c i)^2 * ((card (\s ! i))- (int \))"] by simp + have "\ * ((\ j \ {0..<\} . c j)^2) \ 0" by simp + then have rhs0: "\ * ((\ j \ {0..<\} . c j)^2) = 0" + using assms(2) assms(4) lhsge by linarith + then have "(\ j \ {0..<\} . (c j)^2 * ((card (\s ! j))- (int \))) = 0" + using assms by linarith + then have lhs0inner: "\ j . j < \ \ (c j)^2 * (card (\s ! j) - (int \)) = 0" + using innerge sum_nonneg_eq_0_iff[of "{0..<\}" "\ j . (c j)^2 * (card (\s ! j) - (int \))"] + by simp + thus False proof (cases "card (\s ! j') = \") + case True + then have cj0: "\ j. j \ {0..<\} - {j'} \ (c j) = 0" + using lhs0inner const_intersect_block_size_diff assms True by auto + then have "(\ i \ {0..<\} . c i) \ 0" + using sum.remove[of "{0..<\}" "j'" "\ i. c i"] assms(3) cine0 cj0 by simp + then show ?thesis using rhs0 assms by simp + next + case False + then have ne: "(card (\s ! j') - (int \)) \ 0" + by linarith + then have "(c j')^2 * (card (\s ! j') - (int \)) \ 0" using cine0 + by auto + then show ?thesis using lhs0inner assms(3) by auto + qed +qed + +text \The general non-uniform version of fisher's inequality is also known as the "Block town problem". +In this case we are working in a constant intersect design, hence the inequality is the opposite +way around compared to the BIBD version. The theorem below is the more traditional set theoretic +approach. This proof is based off one by Jukna \cite{juknaExtremalCombinatorics2011} \ +theorem general_fishers_inequality: "\ \ \" +proof (cases "\ = 0 \ \ = 1") + case True + then show ?thesis using empty_inter_implies_b_lt_v v_non_zero by linarith +next + case False + then have mge: "\ > 0" by simp + then have bge: "\ \ 2" using b_positive False blocks_list_length by linarith + define NR :: "real mat" where "NR \ lift_01_mat N" + show ?thesis + proof (intro lin_bound_argument2[of NR]) + show "distinct (cols NR)" using lift_01_distinct_cols_N NR_def by simp + show nrcm: "NR \ carrier_mat \ \" using NR_def N_carrier_mat_01_lift by simp + have scalar_prod_real1: "\ i. i <\ \ ((col NR i) \ (col NR i)) = card (\s ! i)" + using scalar_prod_block_size_lift_01 NR_def by auto + have scalar_prod_real2: "\ i j. i <\ \ j <\ \ i \ j \ ((col NR i) \ (col NR j)) = \" + using scalar_prod_inter_num_lift_01 NR_def indexed_const_intersect by auto + show "\f. vec \ (\i. \j = 0..<\. f (col NR j) * (col NR j) $ i) = 0\<^sub>v \ \ \v\set (cols NR). f v = 0" + proof (intro ballI) + fix f v + assume eq0: "vec \ (\i. \j = 0..<\. f (col NR j) * col NR j $ i) = 0\<^sub>v \" + assume vin: "v \ set (cols NR)" + define c where "c \ (\ j. f (col NR j))" + obtain j' where v_def: "col NR j' = v" and jvlt: "j' < dim_col NR" + using vin by (metis cols_length cols_nth index_less_size_conv nth_index) + have dim_col: "\j. j \ {0..< \} \ dim_vec (col NR j) = \" using nrcm by auto + \ \ Summation reasoning to obtain conclusion on coefficients \ + have "0 = (vec \ (\i. \j = 0..<\. c j * (col NR j) $ i)) \ (vec \ (\i. \j = 0..<\. c j * (col NR j) $ i))" + using vec_prod_zero eq0 c_def by simp + also have "... = (\ j1 \ {0..<\} . c j1 * c j1 * ((col NR j1) \ (col NR j1))) + (\ j1 \ {0..<\} . + (\ j2 \ ({0..< \} - {j1}) . c j1 * c j2 * ((col NR j1) \ (col NR j2))))" + using scalar_prod_double_sum_fn_vec[of \ "col NR" \ c] dim_col by simp + also have "... = (\ j1 \ {0..<\} . (c j1) * (c j1) * (card (\s ! j1))) + (\ j1 \ {0..<\} . + (\ j2 \ ({0..< \} - {j1}) . c j1 * c j2 * ((col NR j1) \ (col NR j2))))" + using scalar_prod_real1 by simp + also have "... = (\ j1 \ {0..<\} . (c j1)^2 * (card (\s ! j1))) + (\ j1 \ {0..<\} . + (\ j2 \ ({0..< \} - {j1}) . c j1 * c j2 * ((col NR j1) \ (col NR j2))))" + by (metis power2_eq_square) + also have "... = (\ j1 \ {0..<\} . (c j1)^2 * (card (\s ! j1))) + (\ j1 \ {0..<\} . + (\ j2 \ ({0..< \} - {j1}) . c j1 * c j2 * \))" using scalar_prod_real2 by auto + also have "... = (\ j1 \ {0..<\} . (c j1)^2 * (card (\s ! j1))) + + \ * (\ j1 \ {0..<\} . (\ j2 \ ({0..< \} - {j1}) . c j1 * c j2))" + using double_sum_mult_hom[of "\" "\ i j . c i * c j" "\ i.{0..<\} - {i}" "{0..<\}"] + by (metis (no_types, lifting) mult_of_nat_commute sum.cong) + also have "... = (\ j \ {0..<\} . (c j)^2 * (card (\s ! j))) + + \ * ((\ j \ {0..<\} . c j)^2 - (\ j \ {0..<\} . c j * c j))" + using double_sum_split_square_diff by auto + also have "... = (\ j \ {0..<\} . (c j)^2 * (card (\s ! j))) + (-\) * (\ j \ {0..<\} . (c j)^2) + + \ * ((\ j \ {0..<\} . c j)^2)" by (simp add: algebra_simps power2_eq_square) + also have "... = (\ j \ {0..<\} . (c j)^2 * (card (\s ! j))) + (\ j \ {0..<\} . (-\)* (c j)^2) + + \ * ((\ j \ {0..<\} . c j)^2)" by (simp add: sum_distrib_left) + also have "... = (\ j \ {0..<\} . (c j)^2 * (card (\s ! j))+ (-\)* (c j)^2) + + \ * ((\ j \ {0..<\} . c j)^2)" by (metis (no_types) sum.distrib) + finally have sum_rep: "0 = (\ j \ {0..<\} . (c j)^2 * ((card (\s ! j))- (int \))) + + \ * ((\ j \ {0..<\} . c j)^2)" by (simp add: algebra_simps) + thus "f v = 0" using sum_split_coeffs_0[of "j'" "c"] mge bge jvlt nrcm c_def v_def by simp + qed + qed +qed + +end + +text \Using the dual design concept, it is easy to translate the set theoretic general definition +of Fisher's inequality to a more traditional design theoretic version on pairwise balanced designs. +Two versions of this are given using different trivial (but crucial) conditions on design properties\ +context ordered_pairwise_balance +begin + +corollary general_nonuniform_fishers: \ \only valid on incomplete designs \ + assumes "\ > 0" + assumes "\ bl. bl \# \ \ incomplete_block bl" + \ \ i.e. not a super trivial design with only complete blocks \ + shows "\ \ \" +proof - + have "mset (\s*) = dual_blocks \ \s" using dual_blocks_ordered_eq by simp + then interpret des: simple_const_intersect_design "set [0..<(length \s)]" "mset (\s*)" \ + using assms dual_is_simp_const_inter_des by simp + interpret odes: simp_ordered_const_intersect_design "[0..s]" "\s*" \ + using distinct_upt des.wellformed by (unfold_locales) (blast) + have "length (\s*) \ length [0..s]" using odes.general_fishers_inequality + using odes.blocks_list_length odes.points_list_length by presburger + then have "\ \ length \s" + by (simp add: dual_blocks_len points_list_length) + then show ?thesis by auto +qed + +corollary general_nonuniform_fishers_comp: + assumes "\ > 0" + assumes "count \ \ < \" \ \ i.e. not a super trivial design with only complete blocks and single blocks \ + shows "\ \ \" +proof - + define B where "B = (removeAll_mset \ \)" + have b_smaller: "size B \ \" using B_def removeAll_size_lt by simp + then have b_incomp: "\ bl. bl \# B \ card bl < \" + using wellformed B_def by (simp add: psubsetI psubset_card_mono) + have index_gt: "(\ - (count \ \)) > 0" using assms by simp + interpret pbd: pairwise_balance \ B "(\ - (count \ \))" + using remove_all_complete_blocks_pbd B_def assms(2) by blast + obtain Bs where m: "mset Bs = B" + using ex_mset by blast + interpret opbd: ordered_pairwise_balance \s Bs "(\ - (count \ \))" + by (intro pbd.ordered_pbdI) (simp_all add: m distinct) + have "\ \ (size B)" using b_incomp opbd.general_nonuniform_fishers + using index_gt m by blast + then show ?thesis using b_smaller m by auto +qed + +end +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/Fishers_Inequality_Root.thy b/thys/Fishers_Inequality/Fishers_Inequality_Root.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Fishers_Inequality_Root.thy @@ -0,0 +1,21 @@ +(* Title: Fishers_Inequality_Root.thy + Author: Chelsea Edmonds +*) + +theory Fishers_Inequality_Root +imports + Set_Multiset_Extras + Matrix_Vector_Extras + Design_Extras + + Incidence_Matrices + Dual_Systems + Rank_Argument_General + Linear_Bound_Argument + + Fishers_Inequality + Vector_Matrix_Mod + Fishers_Inequality_Variations +begin + +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/Fishers_Inequality_Variations.thy b/thys/Fishers_Inequality/Fishers_Inequality_Variations.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Fishers_Inequality_Variations.thy @@ -0,0 +1,211 @@ +(* Title: Fishers_Inequality_Variations.thy + Author: Chelsea Edmonds +*) + +section \Variations on Fisher's Inequality \ + +theory Fishers_Inequality_Variations imports Dual_Systems Rank_Argument_General +Vector_Matrix_Mod Linear_Bound_Argument +begin + +subsection \ Matrix mod properties \ +text \This is reasoning on properties specific to incidence matrices under @{term "mat_mod"}. Ultimately, +this definition was not used in the final proof but it is left as is in case of future use \ + +context mat_mod +begin + +lemma mat_mod_proper_iff: "proper_inc_mat (mat_mod N) \ proper_inc_mat N" + by (simp add: proper_inc_mat_def) + +lemma mat_mod_rep_num_eq: "i < dim_row N \ elements_mat N \ {0.. + mat_rep_num (mat_mod N) i = mat_rep_num N i" + by (simp add: mat_mod_count_row_eq mat_rep_num_def) + +lemma mat_point_index_eq: "elements_mat N \ {0.. + mat_point_index (mat_mod N) I = mat_point_index N I" + by (simp add: mat_mod_eq_cond) + +lemma mod_mat_inter_num_eq: "elements_mat N \ {0.. + mat_inter_num (mat_mod N) j1 j2 = mat_inter_num N j1 j2" + by (simp add: mat_mod_eq_cond) + +lemma mod_mat_block_size: "elements_mat N \ {0.. mat_block_size (mat_mod N) j = mat_block_size N j" + by (simp add: mat_mod_eq_cond) + +lemma mat_mod_non_empty_col_iff: "elements_mat M \ {0.. + non_empty_col (mat_mod M) j \ non_empty_col M j" + using mat_mod_eq_cond by auto +end + +context mat_mod_type +begin + +lemma mat_rep_num_MM_Rel: + assumes "MM_Rel A B" + assumes "i < dim_row A" + shows "mat_rep_num (mat_mod A) i = mat_rep_num B i" + unfolding mat_rep_num_def using vec_count_MV_Rel_direct assms mat_mod_vec_mod_row row_map_mat + by (metis MM_Rel_def MV_Rel_def index_map_mat(2) mat_mod_dim(1) to_int_mod_ring_hom.hom_one) + + +lemma mat_block_size_MM_Rel: + assumes "MM_Rel A B" + assumes " j < dim_col A" + shows "mat_block_size (mat_mod A) j = mat_block_size B j" + unfolding mat_block_size_def using vec_count_MV_Rel_direct assms MM_Rel_MV_Rel_col + by (metis mat_mod_vec_mod_col to_int_mod_ring_hom.hom_one) + +lemma mat_inter_num_MM_Rel: + assumes "MM_Rel A B" + assumes "j1 < dim_col A" "j2 < dim_col B" + shows "mat_inter_num (mat_mod A) j1 j2 = mat_inter_num B j1 j2" + unfolding mat_inter_num_def using assms index_map_mat mat_mod_dim(2) + by (smt (z3) Collect_cong MM_Rel_def to_int_mod_ring_hom.hom_1 to_int_mod_ring_hom.hom_one) + + +text \ Lift 01 and mat mod equivalence on 0-1 matrices \ + +lemma of_int_mod_ring_lift_01_eq: + assumes "zero_one_matrix N" + shows "map_mat (of_int_mod_ring) N = (lift_01_mat) N" + apply (auto simp add: mat_eq_iff[of "map_mat (of_int_mod_ring) N" "lift_01_mat N"]) + using assms zero_one_matrix.M_not_one_simp by fastforce + +lemma to_int_mod_ring_lift_01_eq: + assumes "zero_one_matrix N" + shows "to_int_mat N = (lift_01_mat) N" + apply (auto simp add: mat_eq_iff[of "to_int_mat N" "lift_01_mat N"]) + using assms using zero_one_matrix.M_not_zero_simp by fastforce + +end + +subsection \The Odd-town Problem\ +text \ The odd-town problem \cite{abaiLINEARALGEBRAMETHODS1988} is perhaps one of the most common +introductory problems for applying the linear algebra bound method to a combinatorial problem. +In mathematical literature, it is considered simpler than Fisher's Inequality, however presents some +interesting challenges to formalisation. Most significantly, it considers the incidence matrix to have +elements of types integers mod 2. \ + +text \Initially, define a locale to represent the odd town context (a town with v people, and b groups) +which must each be of odd size, but have an even intersection number with any other group \ +locale odd_town = ordered_design + + assumes odd_groups: "bl \# \ \ odd (card bl)" + and even_inters: "bl1 \# \ \ bl2 \# (\ - {#bl1#}) \ even (bl1 |\| bl2)" +begin + +lemma odd_town_no_repeat_clubs: "distinct_mset \" +proof (rule ccontr) + assume "\ distinct_mset \" + then obtain a where ain: "a \# \" and countne: "count \ a \ 1" + by (auto simp add: distinct_mset_def) + then have "count \ a > 1" + using nat_less_le by auto + then have ain2: "a \# (\ - {#a#})" + by (simp add: in_diff_count) + then have "odd (a |\| a)" using odd_groups ain by simp + thus False using even_inters ain ain2 + by blast +qed + +lemma odd_blocks_mat_block_size: "j < dim_col N \ odd (mat_block_size N j)" + using mat_block_size_conv odd_groups + by (metis dim_col_is_b valid_blocks_index) + +lemma odd_block_size_mod_2: + assumes "CARD('b::prime_card) = 2" + assumes "j < \" + shows "of_nat (card (\s ! j)) = (1 :: 'b mod_ring)" +proof - + have cb2: "CARD('b) = 2" using assms by simp + then have "odd (card (\s ! j))" using \j < \\ odd_groups by auto + then show "of_nat (card (\s ! j)) = (1 :: 'b mod_ring)" + by(transfer' fixing: j \s, simp add: cb2) presburger +qed + +lemma valid_indices_block_min: "j1 < dim_col N \ j2 < dim_col N \ j1 \ j2 \ \ \ 2" + by simp + +lemma even_inter_mat_intersections: "j1 < dim_col N \ j2 < dim_col N \ j1 \ j2 + \ even (mat_inter_num N j1 j2)" + using even_inters mat_inter_num_conv valid_indices_block_min + by (metis dim_col_is_b obtains_two_diff_block_indexes) + +lemma even_inter_mod_2: + assumes "CARD('b::prime_card) = 2" + assumes "i < \" and jlt: "j < \" and ne: "i \ j" + shows "of_nat ((\s ! i) |\| (\s ! j)) = (0 :: 'b mod_ring)" +proof - + have cb2: "CARD('b) = 2" using assms by simp + have "even ((\s ! i) |\| (\s ! j))" using even_inters assms + using blocks_index_ne_belong blocks_list_length valid_blocks_index by presburger + then show "of_nat ((\s ! i) |\| (\s ! j)) = (0 :: 'b mod_ring)" + by (transfer' fixing: i j \s, simp add: cb2) +qed + +end + +text \The odd town locale must be simple by definition \ +sublocale odd_town \ ordered_simple_design + using odd_town_no_repeat_clubs by (unfold_locales) (meson distinct_mset_def) + +context odd_town +begin + +text \The upper bound lemma (i.e. variation on Fisher's) for the odd town property using the linear +bound argument. Notice it follows exactly the same pattern as the generalised version, however +it's sum manipulation argument is significantly simpler (in line with the mathematical proofs) \ +lemma upper_bound_clubs: + assumes "CARD('b::prime_card) = 2" + shows "\ \ \" +proof - + have cb2: "CARD('b) = 2" using assms by simp + then interpret mmt: mat_mod_type 2 "TYPE('b::prime_card)" + using assms by (unfold_locales) (simp_all) + define N2 :: "'b mod_ring mat" where "N2 \ lift_01_mat N" + show ?thesis proof (intro lin_bound_argument2[of "N2"]) + show "distinct (cols (N2))" using lift_01_distinct_cols_N N2_def by simp + show n2cm:"N2 \ carrier_mat \ \" using N2_def N_carrier_mat_01_lift by simp + have scalar_prod_odd: "\ i. i <\ \ ((col N2 i) \ (col N2 i)) = 1" + using scalar_prod_block_size_lift_01 N2_def odd_block_size_mod_2 assms by (metis cb2) + have scalar_prod_even: "\ i j. i <\ \ j <\ \ i \ j \ ((col N2 i) \ (col N2 j)) = 0" + using even_inter_mod_2 scalar_prod_inter_num_lift_01 N2_def assms by metis + show "\f. vec \ (\i. \j = 0..<\. f (col N2 j) * (col N2 j) $ i) = 0\<^sub>v \ \ \v\set (cols N2). f v = 0" + proof (auto) + fix f v + assume eq0: "vec \ (\i. \j = 0..s. f (col N2 j) * (col N2 j) $ i) = 0\<^sub>v \" + assume vin: "v \ set (cols N2)" + define c where "c \ (\ j. f (col N2 j))" + have inner: "\ j l. v $ l * (c j * (col N2 j) $ l) = c j * v $ l * (col N2 j) $ l" + using mult.commute by auto + obtain j' where v_def: "col N2 j' = v" and jvlt: "j' < dim_col N2" + using vin by (metis cols_length cols_nth index_less_size_conv nth_index) + then have jvltb: "j' < \" using n2cm by simp + then have even0: "\ j. j \ {0..<\} - {j'} \ c j * (v \ (col N2 j)) = 0" + using scalar_prod_even v_def by fastforce + have vinc: "v \ carrier_vec \" using n2cm set_cols_carrier vin by blast + then have "0 = v \ vec \ (\i. \j = 0..<\. c j * (col N2 j) $ i)" + using eq0 c_def by auto + also have "... = (\ l =0.. j = 0.. l =0.. j = 0.. j \ {0.. (c j \\<^sub>v (col N2 j)))" + using sum.swap scalar_prod_def[of v] by simp + also have "... = v \ (c j' \\<^sub>v v) + (\ j \ {0.. (c j \\<^sub>v (col N2 j)))" + using jvlt sum.remove[of "{0.. j. v \ (c j \\<^sub>v (col N2 j))"] v_def by simp + also have "... = v \ (c j' \\<^sub>v v) + (\ j \ {0..<\} - {j'}. c j * (v \ (col N2 j)))" + using n2cm scalar_prod_smult_distrib col_dim v_def by force + also have "... = v \ (c j' \\<^sub>v v)" + using even0 by (simp add: sum.neutral) + also have "... = (c j') * (v \ v)" + using scalar_prod_smult_distrib by (simp add: v_def) + finally have "0 = (c j')" using v_def jvlt n2cm scalar_prod_odd by fastforce + then show "f v = 0" using c_def v_def by simp + qed + qed +qed + +end + +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/Incidence_Matrices.thy b/thys/Fishers_Inequality/Incidence_Matrices.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Incidence_Matrices.thy @@ -0,0 +1,2084 @@ +(* Author: Chelsea Edmonds +Theory: Incidence_Matrices.thy +*) + +section \ Incidence Vectors and Matrices \ +text \Incidence Matrices are an important representation for any incidence set system. The majority +of basic definitions and properties proved in this theory are based on Stinson \cite{stinsonCombinatorialDesignsConstructions2004} +and Colbourn \cite{colbournHandbookCombinatorialDesigns2007}.\ + +theory Incidence_Matrices imports "Design_Extras" Matrix_Vector_Extras "List-Index.List_Index" + "Design_Theory.Design_Isomorphisms" +begin + +subsection \Incidence Vectors \ +text \A function which takes an ordered list of points, and a block, +returning a 0-1 vector $v$ where there is a 1 in the ith position if point i is in that block \ + +definition inc_vec_of :: "'a list \ 'a set \ ('b :: {ring_1}) vec" where +"inc_vec_of Vs bl \ vec (length Vs) (\ i . if (Vs ! i) \ bl then 1 else 0)" + +lemma inc_vec_one_zero_elems: "set\<^sub>v (inc_vec_of Vs bl) \ {0, 1}" + by (auto simp add: vec_set_def inc_vec_of_def) + +lemma finite_inc_vec_elems: "finite (set\<^sub>v (inc_vec_of Vs bl))" + using finite_subset inc_vec_one_zero_elems by blast + +lemma inc_vec_elems_max_two: "card (set\<^sub>v (inc_vec_of Vs bl)) \ 2" + using card_mono inc_vec_one_zero_elems finite.insertI card_0_eq card_2_iff + by (smt (verit) insert_absorb2 linorder_le_cases linordered_nonzero_semiring_class.zero_le_one + obtain_subset_with_card_n one_add_one subset_singletonD trans_le_add1) + +lemma inc_vec_dim: "dim_vec (inc_vec_of Vs bl) = length Vs" + by (simp add: inc_vec_of_def) + +lemma inc_vec_index: "i < length Vs \ inc_vec_of Vs bl $ i = (if (Vs ! i) \ bl then 1 else 0)" + by (simp add: inc_vec_of_def) + +lemma inc_vec_index_one_iff: "i < length Vs \ inc_vec_of Vs bl $ i = 1 \ Vs ! i \ bl" + by (auto simp add: inc_vec_of_def ) + +lemma inc_vec_index_zero_iff: "i < length Vs \ inc_vec_of Vs bl $ i = 0 \ Vs ! i \ bl" + by (auto simp add: inc_vec_of_def) + +lemma inc_vec_of_bij_betw: + assumes "inj_on f (set Vs)" + assumes "bl \ (set Vs)" + shows "inc_vec_of Vs bl = inc_vec_of (map f Vs) (f ` bl)" +proof (intro eq_vecI, simp_all add: inc_vec_dim) + fix i assume "i < length Vs" + then have "Vs ! i \ bl \ (map f Vs) ! i \ (f ` bl)" + by (metis assms(1) assms(2) inj_on_image_mem_iff nth_map nth_mem) + then show "inc_vec_of Vs bl $ i = inc_vec_of (map f Vs) (f ` bl) $ i" + using inc_vec_index by (metis \i < length Vs\ length_map) +qed + +subsection \ Incidence Matrices \ + +text \ A function which takes a list of points, and list of sets of points, and returns +a $v \times b$ 0-1 matrix $M$, where $v$ is the number of points, and $b$ the number of sets, such +that there is a 1 in the i, j position if and only if point i is in block j. The matrix has +type @{typ "('b :: ring_1) mat"} to allow for operations commonly used on matrices \cite{stinsonCombinatorialDesignsConstructions2004}\ + +definition inc_mat_of :: "'a list \ 'a set list \ ('b :: {ring_1}) mat" where +"inc_mat_of Vs Bs \ mat (length Vs) (length Bs) (\ (i,j) . if (Vs ! i) \ (Bs ! j) then 1 else 0)" + +text \ Basic lemmas on the @{term "inc_mat_of"} matrix result (elements/dimensions/indexing)\ + +lemma inc_mat_one_zero_elems: "elements_mat (inc_mat_of Vs Bs) \ {0, 1}" + by (auto simp add: inc_mat_of_def elements_mat_def) + +lemma fin_incidence_mat_elems: "finite (elements_mat (inc_mat_of Vs Bs))" + using finite_subset inc_mat_one_zero_elems by auto + +lemma inc_matrix_elems_max_two: "card (elements_mat (inc_mat_of Vs Bs)) \ 2" + using inc_mat_one_zero_elems order_trans card_2_iff + by (smt (verit, del_insts) antisym bot.extremum card.empty insert_commute insert_subsetI + is_singletonI is_singleton_altdef linorder_le_cases not_one_le_zero one_le_numeral subset_insert) + +lemma inc_mat_of_index [simp]: "i < dim_row (inc_mat_of Vs Bs) \ j < dim_col (inc_mat_of Vs Bs) \ + inc_mat_of Vs Bs $$ (i, j) = (if (Vs ! i) \ (Bs ! j) then 1 else 0)" + by (simp add: inc_mat_of_def) + +lemma inc_mat_dim_row: "dim_row (inc_mat_of Vs Bs) = length Vs" + by (simp add: inc_mat_of_def) + +lemma inc_mat_dim_vec_row: "dim_vec (row (inc_mat_of Vs Bs) i) = length Bs" + by (simp add: inc_mat_of_def) + +lemma inc_mat_dim_col: "dim_col (inc_mat_of Vs Bs) = length Bs" + by (simp add: inc_mat_of_def) + +lemma inc_mat_dim_vec_col: "dim_vec (col (inc_mat_of Vs Bs) i) = length Vs" + by (simp add: inc_mat_of_def) + +lemma inc_matrix_point_in_block_one: "i < length Vs \ j < length Bs \ Vs ! i \ Bs ! j + \ (inc_mat_of Vs Bs) $$ (i, j) = 1" + by (simp add: inc_mat_of_def) + +lemma inc_matrix_point_not_in_block_zero: "i < length Vs \ j < length Bs \ Vs ! i \ Bs ! j \ + (inc_mat_of Vs Bs) $$ (i, j) = 0" + by(simp add: inc_mat_of_def) + +lemma inc_matrix_point_in_block: "i < length Vs \ j < length Bs \ (inc_mat_of Vs Bs) $$ (i, j) = 1 + \ Vs ! i \ Bs ! j" + using inc_matrix_point_not_in_block_zero by (metis zero_neq_one) + +lemma inc_matrix_point_not_in_block: "i < length Vs \ j < length Bs \ + (inc_mat_of Vs Bs) $$ (i, j) = 0 \ Vs ! i \ Bs ! j" + using inc_matrix_point_in_block_one by (metis zero_neq_one) + +lemma inc_matrix_point_not_in_block_iff: "i < length Vs \ j < length Bs \ + (inc_mat_of Vs Bs) $$ (i, j) = 0 \ Vs ! i \ Bs ! j" + using inc_matrix_point_not_in_block inc_matrix_point_not_in_block_zero by blast + +lemma inc_matrix_point_in_block_iff: "i < length Vs \ j < length Bs \ + (inc_mat_of Vs Bs) $$ (i, j) = 1 \ Vs ! i \ Bs ! j" + using inc_matrix_point_in_block inc_matrix_point_in_block_one by blast + +lemma inc_matrix_subset_implies_one: + assumes "I \ {..< length Vs}" + assumes "j < length Bs" + assumes "(!) Vs ` I \ Bs ! j" + assumes "i \ I" + shows "(inc_mat_of Vs Bs) $$ (i, j) = 1" +proof - + have iin: "Vs ! i \ Bs ! j" using assms(3) assms(4) by auto + have "i < length Vs" using assms(1) assms(4) by auto + thus ?thesis using iin inc_matrix_point_in_block_iff assms(2) by blast +qed + +lemma inc_matrix_one_implies_membership: "I \ {..< length Vs} \ j < length Bs \ + (\ i. i\I \ (inc_mat_of Vs Bs) $$ (i, j) = 1) \ i \ I \ Vs ! i \ Bs ! j" + using inc_matrix_point_in_block subset_iff by blast + +lemma inc_matrix_elems_one_zero: "i < length Vs \ j < length Bs \ + (inc_mat_of Vs Bs) $$ (i, j) = 0 \ (inc_mat_of Vs Bs) $$ (i, j) = 1" + using inc_matrix_point_in_block_one inc_matrix_point_not_in_block_zero by blast + +text \Reasoning on Rows/Columns of the incidence matrix \ + +lemma inc_mat_col_def: "j < length Bs \ i < length Vs \ + (col (inc_mat_of Vs Bs) j) $ i = (if (Vs ! i \ Bs ! j) then 1 else 0)" + by (simp add: inc_mat_of_def) + +lemma inc_mat_col_list_map_elem: "j < length Bs \ i < length Vs \ + col (inc_mat_of Vs Bs) j $ i = map_vec (\ x . if (x \ (Bs ! j)) then 1 else 0) (vec_of_list Vs) $ i" + by (simp add: inc_mat_of_def index_vec_of_list) + +lemma inc_mat_col_list_map: "j < length Bs \ + col (inc_mat_of Vs Bs) j = map_vec (\ x . if (x \ (Bs ! j)) then 1 else 0) (vec_of_list Vs)" + by (intro eq_vecI) + (simp_all add: inc_mat_dim_row inc_mat_dim_col inc_mat_col_list_map_elem index_vec_of_list) + +lemma inc_mat_row_def: "j < length Bs \ i < length Vs \ + (row (inc_mat_of Vs Bs) i) $ j = (if (Vs ! i \ Bs ! j) then 1 else 0)" + by (simp add: inc_mat_of_def) + +lemma inc_mat_row_list_map_elem: "j < length Bs \ i < length Vs \ + row (inc_mat_of Vs Bs) i $ j = map_vec (\ bl . if ((Vs ! i) \ bl) then 1 else 0) (vec_of_list Bs) $ j" + by (simp add: inc_mat_of_def vec_of_list_index) + +lemma inc_mat_row_list_map: "i < length Vs \ + row (inc_mat_of Vs Bs) i = map_vec (\ bl . if ((Vs ! i) \ bl) then 1 else 0) (vec_of_list Bs)" + by (intro eq_vecI) + (simp_all add: inc_mat_dim_row inc_mat_dim_col inc_mat_row_list_map_elem index_vec_of_list) + +text \ Connecting @{term "inc_vec_of"} and @{term "inc_mat_of"} \ + +lemma inc_mat_col_inc_vec: "j < length Bs \ col (inc_mat_of Vs Bs) j = inc_vec_of Vs (Bs ! j)" + by (auto simp add: inc_mat_of_def inc_vec_of_def) + +lemma inc_mat_of_cols_inc_vecs: "cols (inc_mat_of Vs Bs) = map (\ j . inc_vec_of Vs j) Bs" +proof (intro nth_equalityI) + have l1: "length (cols (inc_mat_of Vs Bs)) = length Bs" + using inc_mat_dim_col by simp + have l2: "length (map (\ j . inc_vec_of Vs j) Bs) = length Bs" + using length_map by simp + then show "length (cols (inc_mat_of Vs Bs)) = length (map (inc_vec_of Vs) Bs)" + using l1 l2 by simp + show "\ i. i < length (cols (inc_mat_of Vs Bs)) \ + (cols (inc_mat_of Vs Bs) ! i) = (map (\ j . inc_vec_of Vs j) Bs) ! i" + using inc_mat_col_inc_vec l1 by (metis cols_nth inc_mat_dim_col nth_map) +qed + +lemma inc_mat_of_bij_betw: + assumes "inj_on f (set Vs)" + assumes "\ bl . bl \ (set Bs) \ bl \ (set Vs)" + shows "inc_mat_of Vs Bs = inc_mat_of (map f Vs) (map ((`) f) Bs)" +proof (intro eq_matI, simp_all add: inc_mat_dim_row inc_mat_dim_col, intro impI) + fix i j assume ilt: "i < length Vs" and jlt: " j < length Bs" and "Vs ! i \ Bs ! j" + then show "f (Vs ! i) \ f ` Bs ! j" + by (meson assms(1) assms(2) ilt inj_on_image_mem_iff jlt nth_mem) +qed + +text \Definitions for the incidence matrix representation of common incidence system properties \ + +definition non_empty_col :: "('a :: {zero_neq_one}) mat \ nat \ bool" where +"non_empty_col M j \ \ k. k \ 0 \ k \$ col M j" + +definition proper_inc_mat :: "('a :: {zero_neq_one}) mat \ bool" where +"proper_inc_mat M \ (dim_row M > 0 \ dim_col M > 0)" + +text \Matrix version of the representation number property @{term "point_replication_number"}\ +definition mat_rep_num :: "('a :: {zero_neq_one}) mat \ nat \ nat" where +"mat_rep_num M i \ count_vec (row M i) 1" + +text \Matrix version of the points index property @{term "points_index"}\ +definition mat_point_index :: "('a :: {zero_neq_one}) mat \ nat set \ nat" where +"mat_point_index M I \ card {j . j < dim_col M \ (\ i \ I. M $$ (i, j) = 1)}" + +definition mat_inter_num :: "('a :: {zero_neq_one}) mat \ nat \ nat \ nat" where +"mat_inter_num M j1 j2 \ card {i . i < dim_row M \ M $$ (i, j1) = 1 \ M $$ (i, j2) = 1}" + +text \Matrix version of the block size property\ +definition mat_block_size :: "('a :: {zero_neq_one}) mat \ nat \ nat" where +"mat_block_size M j \ count_vec (col M j) 1" + +lemma non_empty_col_obtains: + assumes "non_empty_col M j" + obtains i where "i < dim_row M" and "(col M j) $ i \ 0" +proof - + have d: "dim_vec (col M j) = dim_row M" by simp + from assms obtain k where "k \ 0" and "k \$ col M j" + by (auto simp add: non_empty_col_def) + thus ?thesis using vec_contains_obtains_index d + by (metis that) +qed + +lemma non_empty_col_alt_def: + assumes "j < dim_col M" + shows "non_empty_col M j \ (\ i. i < dim_row M \ M $$ (i, j) \ 0)" +proof (intro iffI) + show "non_empty_col M j \ \i 0" + by(metis assms index_col non_empty_col_obtains) +next + assume "\i 0" + then obtain i where ilt: " i < dim_row M" and ne: "M $$ (i, j) \ 0" by blast + then have ilt2: " i < dim_vec (col M j)" by simp + then have "(col M j) $ i \ 0" using ne by (simp add: assms) + then obtain k where "(col M j) $ i = k" and "k \ 0" + by simp + then show "non_empty_col M j " using non_empty_col_def + by (metis ilt2 vec_setI) +qed + +lemma proper_inc_mat_map: "proper_inc_mat M \ proper_inc_mat (map_mat f M)" + by (simp add: proper_inc_mat_def) + +lemma mat_point_index_alt: "mat_point_index M I = card {j \ {0.. i \ I . M $$(i, j) = 1)}" + by (simp add: mat_point_index_def) + +lemma mat_block_size_sum_alt: + fixes M :: "'a :: {ring_1} mat" + shows "elements_mat M \ {0, 1} \ j < dim_col M \ of_nat (mat_block_size M j) = sum_vec (col M j)" + unfolding mat_block_size_def using count_vec_sum_ones_alt col_elems_subset_mat subset_trans + by metis + +lemma mat_rep_num_sum_alt: + fixes M :: "'a :: {ring_1} mat" + shows "elements_mat M \ {0, 1} \ i < dim_row M \ of_nat (mat_rep_num M i) = sum_vec (row M i)" + using count_vec_sum_ones_alt + by (metis mat_rep_num_def row_elems_subset_mat subset_trans) + +lemma mat_point_index_two_alt: + assumes "i1 < dim_row M" + assumes "i2 < dim_row M" + shows "mat_point_index M {i1, i2} = card {j . j < dim_col M \ M $$(i1, j) = 1 \ M $$ (i2, j) = 1}" +proof - + let ?I = "{i1, i2}" + have ss: "{i1, i2} \ {.. j . j < dim_col M \ (\ i \ ?I . M $$(i, j) = 1) \ M $$(i1, j) = 1 \ M $$ (i2, j) = 1" + by auto + have "?I \ {..< dim_row M}" using assms(1) assms(2) by fastforce + thus ?thesis using filter ss unfolding mat_point_index_def + by meson +qed + +text \ Transpose symmetries \ + +lemma trans_mat_rep_block_size_sym: "j < dim_col M \ mat_block_size M j = mat_rep_num M\<^sup>T j" + "i < dim_row M \ mat_rep_num M i = mat_block_size M\<^sup>T i" + unfolding mat_block_size_def mat_rep_num_def by simp_all + +lemma trans_mat_point_index_inter_sym: + "i1 < dim_row M \ i2 < dim_row M \ mat_point_index M {i1, i2} = mat_inter_num M\<^sup>T i1 i2" + "j1 < dim_col M \ j2 < dim_col M \ mat_inter_num M j1 j2 = mat_point_index M\<^sup>T {j1, j2}" + apply (simp_all add: mat_inter_num_def mat_point_index_two_alt) + apply (metis (no_types, lifting) index_transpose_mat(1)) + by (metis (no_types, lifting) index_transpose_mat(1)) + +subsection \0-1 Matrices \ +text \Incidence matrices contain only two elements: 0 and 1. We define a locale which provides +a context to work in for matrices satisfying this condition for any @{typ "'b :: zero_neq_one"} type.\ +locale zero_one_matrix = + fixes matrix :: "'b :: {zero_neq_one} mat" ("M") + assumes elems01: "elements_mat M \ {0, 1}" +begin + +text \ Row and Column Properties of the Matrix \ + +lemma row_elems_ss01:"i < dim_row M \ vec_set (row M i) \ {0, 1}" + using row_elems_subset_mat elems01 by blast + +lemma col_elems_ss01: + assumes "j < dim_col M" + shows "vec_set (col M j) \ {0, 1}" +proof - + have "vec_set (col M j) \ elements_mat M" using assms + by (simp add: col_elems_subset_mat assms) + thus ?thesis using elems01 by blast +qed + +lemma col_nth_0_or_1_iff: + assumes "j < dim_col M" + assumes "i < dim_row M" + shows "col M j $ i = 0 \ col M j $ i \ 1" +proof (intro iffI) + have dv: "i < dim_vec (col M j)" using assms by simp + have sv: "set\<^sub>v (col M j) \ {0, 1}" using col_elems_ss01 assms by simp + then show "col M j $ i = 0 \ col M j $ i \ 1" using dv by simp + show "col M j $ i \ 1 \ col M j $ i = 0" using dv sv + by (meson insertE singletonD subset_eq vec_setI) +qed + +lemma row_nth_0_or_1_iff: + assumes "j < dim_col M" + assumes "i < dim_row M" + shows "row M i $ j = 0 \ row M i $ j \ 1" +proof (intro iffI) + have dv: "j < dim_vec (row M i)" using assms by simp + have sv: "set\<^sub>v (row M i) \ {0, 1}" using row_elems_ss01 assms by simp + then show "row M i $ j = 0 \ row M i $ j \ 1" by simp + show "row M i $ j \ 1 \ row M i $ j = 0" using dv sv + by (meson insertE singletonD subset_eq vec_setI) +qed + +lemma transpose_entries: "elements_mat (M\<^sup>T) \ {0, 1}" + using elems01 transpose_mat_elems by metis + +lemma M_not_zero_simp: "j < dim_col M \ i < dim_row M \ M $$ (i, j) \ 0 \ M $$ (i, j) = 1" + using elems01 by auto + +lemma M_not_one_simp: "j < dim_col M \ i < dim_row M \ M $$ (i, j) \ 1 \ M $$ (i, j) = 0" + using elems01 by auto + +text \Definition for mapping a column to a block \ +definition map_col_to_block :: "'a :: {zero_neq_one} vec \ nat set" where +"map_col_to_block c \ { i \ {.. c$ i = 1}" + by (simp add: map_col_to_block_def) + +lemma map_col_to_block_elem: "i < dim_vec c \ i \ map_col_to_block c \ c $ i = 1" + by (simp add: map_col_to_block_alt) + +lemma in_map_col_valid_index: "i \ map_col_to_block c \ i < dim_vec c" + by (simp add: map_col_to_block_alt) + +lemma map_col_to_block_size: "j < dim_col M \ card (map_col_to_block (col M j)) = mat_block_size M j" + unfolding mat_block_size_def map_col_to_block_alt using count_vec_alt[of "col M j" "1"] Collect_cong + by (metis (no_types, lifting)) + +lemma in_map_col_valid_index_M: "j < dim_col M \ i \ map_col_to_block (col M j) \ i < dim_row M" + using in_map_col_valid_index by (metis dim_col) + +lemma map_col_to_block_elem_not: "c \ set (cols M) \ i < dim_vec c \ i \ map_col_to_block c \ c $ i = 0" + apply (auto simp add: map_col_to_block_alt) + using elems01 by (metis col_nth_0_or_1_iff dim_col obtain_col_index) + +lemma obtain_block_index_map_block_set: + assumes "bl \# {# map_col_to_block c . c \# mset (cols M)#}" + obtains j where "j < dim_col M" and "bl = map_col_to_block (col M j)" +proof - + obtain c where bleq: "bl = map_col_to_block c" and "c \# mset (cols M)" + using assms by blast + then have "c \ set (cols M)" by simp + thus ?thesis using bleq obtain_col_index + by (metis that) +qed + +lemma mat_ord_inc_sys_point[simp]: "x < dim_row M \ [0..<(dim_row M)] ! x = x" + by simp + +lemma mat_ord_inc_sys_block[simp]: "j < dim_col M \ + (map (map_col_to_block) (cols M)) ! j = map_col_to_block (col M j)" + by auto + +lemma ordered_to_mset_col_blocks: + "{# map_col_to_block c . c \# mset (cols M)#} = mset (map (map_col_to_block) (cols M))" + by simp + +text \ Lemmas on incidence matrix properties \ +lemma non_empty_col_01: + assumes "j < dim_col M" + shows "non_empty_col M j \ 1 \$ col M j" +proof (intro iffI) + assume "non_empty_col M j" + then obtain k where kn0: "k \ 0" and kin: "k \$ col M j" using non_empty_col_def + by blast + then have "k \ elements_mat M" using vec_contains_col_elements_mat assms + by metis + then have "k = 1" using kn0 + using elems01 by blast + thus "1 \$ col M j" using kin by simp +next + assume "1 \$ col M j" + then show "non_empty_col M j" using non_empty_col_def + by (metis zero_neq_one) +qed + +lemma mat_rep_num_alt: + assumes "i < dim_row M" + shows "mat_rep_num M i = card {j . j < dim_col M \ M $$ (i, j) = 1}" +proof (simp add: mat_rep_num_def) + have eq: "\ j. (j < dim_col M \ M $$ (i, j) = 1) = (row M i $ j = 1 \ j < dim_vec (row M i))" + using assms by auto + have "count_vec (row M i) 1 = card {j. (row M i) $ j = 1 \ j < dim_vec (row M i)}" + using count_vec_alt[of "row M i" "1"] by simp + thus "count_vec (row M i) 1 = card {j. j < dim_col M \ M $$ (i, j) = 1}" + using eq Collect_cong by simp +qed + +lemma mat_rep_num_alt_col: "i < dim_row M \ mat_rep_num M i = size {#c \# (mset (cols M)) . c $ i = 1#}" + using mat_rep_num_alt index_to_col_card_size_prop[of i M] by auto + +text \ A zero one matrix is an incidence system \ + +lemma map_col_to_block_wf: "\c. c \ set (cols M) \ map_col_to_block c \ {0.. 1 \$ (col M j) \ map_col_to_block (col M j) \ {}" + unfolding map_col_to_block_def using vec_setE by force + +interpretation incidence_sys: incidence_system "{0..# mset (cols M)#}" + using map_col_to_block_wf by (unfold_locales) auto + +interpretation fin_incidence_sys: finite_incidence_system "{0..# mset (cols M)#}" + by (unfold_locales) (simp) + +lemma block_nempty_implies_all_zeros: "j < dim_col M \ map_col_to_block (col M j) = {} \ + i < dim_row M \ col M j $ i = 0" + by (metis col_nth_0_or_1_iff dim_col one_implies_block_nempty vec_setI) + +lemma block_nempty_implies_no_one: "j < dim_col M \ map_col_to_block (col M j) = {} \ \ (1 \$ (col M j))" + using one_implies_block_nempty by blast + +lemma mat_is_design: + assumes "\j. j < dim_col M\ 1 \$ (col M j)" + shows "design {0..# mset (cols M)#}" +proof (unfold_locales) + fix bl + assume "bl \# {# map_col_to_block c . c \# mset (cols M)#}" + then obtain j where "j < dim_col M" and map: "bl = map_col_to_block (col M j)" + using obtain_block_index_map_block_set by auto + thus "bl \ {}" using assms one_implies_block_nempty + by simp +qed + +lemma mat_is_proper_design: + assumes "\j. j < dim_col M \ 1 \$ (col M j)" + assumes "dim_col M > 0" + shows "proper_design {0..# mset (cols M)#}" +proof - + interpret des: design "{0..# mset (cols M)#}" + using mat_is_design assms by simp + show ?thesis proof (unfold_locales) + have "length (cols M) \ 0" using assms(2) by auto + then have "size {# map_col_to_block c . c \# mset (cols M)#} \ 0" by auto + then show "incidence_sys.\ \ 0" by simp + qed +qed + +text \ Show the 01 injective function preserves system properties \ + +lemma inj_on_01_hom_index: + assumes "inj_on_01_hom f" + assumes "i < dim_row M" "j < dim_col M" + shows "M $$ (i, j) = 1 \ (map_mat f M) $$ (i, j) = 1" + and "M $$ (i, j) = 0 \ (map_mat f M) $$ (i, j) = 0" +proof - + interpret hom: inj_on_01_hom f using assms by simp + show "M $$ (i, j) = 1 \ (map_mat f M) $$ (i, j) = 1" + using assms col_nth_0_or_1_iff + by (simp add: hom.inj_1_iff) + show "M $$ (i, j) = 0 \ (map_mat f M) $$ (i, j) = 0" + using assms col_nth_0_or_1_iff + by (simp add: hom.inj_0_iff) +qed + +lemma preserve_non_empty: + assumes "inj_on_01_hom f" + assumes "j < dim_col M" + shows "non_empty_col M j \ non_empty_col (map_mat f M) j" +proof(simp add: non_empty_col_def, intro iffI) + interpret hom: inj_on_01_hom f using assms(1) by simp + assume "\k. k \ 0 \ k \$ col M j" + then obtain k where kneq: "k \ 0" and kin: "k \$ col M j" by blast + then have "f k \$ col (map_mat f M) j" using vec_contains_img + by (metis assms(2) col_map_mat) + then have "f k \ 0" using assms(1) kneq kin assms(2) col_elems_ss01 hom.inj_0_iff by blast + thus "\k. k \ 0 \ k \$ col (map_mat f M) j" + using \f k \$ col (map_mat f M) j\ by blast +next + interpret hom: inj_on_01_hom f using assms(1) by simp + assume "\k. k \ 0 \ k \$ col (map_mat f M) j" + then obtain k where kneq: "k \ 0" and kin: "k \$ col (map_mat f M) j" by blast + then have "k \$ map_vec f (col M j)" using assms(2) col_map_mat by simp + then have "k \ f ` set\<^sub>v (col M j)" + by (smt (verit) image_eqI index_map_vec(1) index_map_vec(2) vec_setE vec_setI) + then obtain k' where keq: "k = f k'" and kin2: "k' \ set\<^sub>v (col M j)" + by blast + then have "k' \ 0" using assms(1) kneq hom.inj_0_iff by blast + thus "\k. k \ 0 \ k \$ col M j" using kin2 by auto +qed + +lemma preserve_mat_rep_num: + assumes "inj_on_01_hom f" + assumes "i < dim_row M" + shows "mat_rep_num M i = mat_rep_num (map_mat f M) i" + unfolding mat_rep_num_def using injective_lim.lim_inj_hom_count_vec inj_on_01_hom_def row_map_mat + by (metis assms(1) assms(2) inj_on_01_hom.inj_1_iff insert_iff row_elems_ss01) + +lemma preserve_mat_block_size: + assumes "inj_on_01_hom f" + assumes "j < dim_col M" + shows "mat_block_size M j = mat_block_size (map_mat f M) j" + unfolding mat_block_size_def using injective_lim.lim_inj_hom_count_vec inj_on_01_hom_def col_map_mat + by (metis assms(1) assms(2) inj_on_01_hom.inj_1_iff insert_iff col_elems_ss01) + + +lemma preserve_mat_point_index: + assumes "inj_on_01_hom f" + assumes "\ i. i \ I \ i < dim_row M" + shows "mat_point_index M I = mat_point_index (map_mat f M) I" +proof - + have "\ i j. i \ I \ j < dim_col M \ M $$ (i, j) = 1 \ + j < dim_col (map_mat f M) \ (map_mat f M) $$ (i, j) = 1" + using assms(2) inj_on_01_hom_index(1) assms(1) by (metis index_map_mat(3)) + thus ?thesis unfolding mat_point_index_def + by (metis (no_types, opaque_lifting) index_map_mat(3)) +qed + +lemma preserve_mat_inter_num: + assumes "inj_on_01_hom f" + assumes "j1 < dim_col M" "j2 < dim_col M" + shows "mat_inter_num M j1 j2 = mat_inter_num (map_mat f M) j1 j2" + unfolding mat_inter_num_def using assms + by (metis (no_types, opaque_lifting) index_map_mat(2) inj_on_01_hom_index(1)) + +lemma lift_mat_01_index_iff: + "i < dim_row M \ j < dim_col M \ (lift_01_mat M) $$ (i, j) = 0 \ M $$ (i, j) = 0" + "i < dim_row M \ j < dim_col M \ (lift_01_mat M) $$ (i, j) = 1 \ M $$ (i, j) = 1" + by (simp) (metis col_nth_0_or_1_iff index_col lift_01_mat_simp(3) of_zero_neq_one_def zero_neq_one) + +lemma lift_mat_elems: "elements_mat (lift_01_mat M) \ {0, 1}" +proof - + have "elements_mat (lift_01_mat M) = of_zero_neq_one ` (elements_mat M)" + by (simp add: lift_01_mat_def map_mat_elements) + then have "elements_mat (lift_01_mat M) \ of_zero_neq_one ` {0, 1}" using elems01 + by fastforce + thus ?thesis + by simp +qed + +lemma lift_mat_is_0_1: "zero_one_matrix (lift_01_mat M)" + using lift_mat_elems by (unfold_locales) + +lemma lift_01_mat_distinct_cols: "distinct (cols M) \ distinct (cols (lift_01_mat M))" + using of_injective_lim.mat_cols_hom_lim_distinct_iff lift_01_mat_def + by (metis elems01 map_vec_mat_cols) + +end + +text \Some properties must be further restricted to matrices having a @{typ "'a :: ring_1"} type \ +locale zero_one_matrix_ring_1 = zero_one_matrix M for M :: "'b :: {ring_1} mat" +begin + +lemma map_col_block_eq: + assumes "c \ set(cols M)" + shows "inc_vec_of [0..i. i < dim_vec c \ c $ i \ 1 \ c $ i = 0" + using assms map_col_to_block_elem map_col_to_block_elem_not by auto + show "dim_vec (inc_vec_of [0..i j. i < dim_row M \ j < dim_col M \ i \ map_col_to_block (col M j) \ M $$ (i, j) = 1" + by (simp add: map_col_to_block_elem) + show "\i j. i < dim_row M \ j < dim_col M \ i \ map_col_to_block (col M j) \ M $$ (i, j) = 0" + by (metis col_nth_0_or_1_iff dim_col index_col map_col_to_block_elem) +qed + +lemma M_index_square_itself: "j < dim_col M \ i < dim_row M \ (M $$ (i, j))^2 = M $$ (i, j)" + using M_not_zero_simp by (cases "M $$ (i, j) = 0")(simp_all, metis power_one) + +lemma M_col_index_square_itself: "j < dim_col M \ i < dim_row M \ ((col M j) $ i)^2 = (col M j) $ i" + using index_col M_index_square_itself by auto + + +text \ Scalar Prod Alternative definitions for matrix properties \ + +lemma scalar_prod_inc_vec_block_size_mat: + assumes "j < dim_col M" + shows "(col M j) \ (col M j) = of_nat (mat_block_size M j)" +proof - + have "(col M j) \ (col M j) = (\ i \ {0.. i \ {0.. i \ {0.. (col M j2) = of_nat (mat_inter_num M j1 j2)" +proof - + have split: "{0.. {0.. (M $$ (i, j2) = 1) } \ + {i \ {0.. M $$ (i, j2) = 0}" using assms M_not_zero_simp by auto + have inter: "{i \ {0.. (M $$ (i, j2) = 1) } \ + {i \ {0.. M $$ (i, j2) = 0} = {}" by auto + have "(col M j1) \ (col M j2) = (\ i \ {0.. i \ {0.. i \ {i \ {0.. (M $$ (i, j2) = 1) } . M $$ (i, j1) * M $$ (i, j2)) + + (\ i \ {i \ {0.. M $$ (i, j2) = 0} . M $$ (i, j1) * M $$ (i, j2))" + using split inter sum.union_disjoint[of " {i \ {0.. (M $$ (i, j2) = 1)}" + "{i \ {0.. M $$ (i, j2) = 0}" "\ i . M $$ (i, j1) * M $$ (i, j2)"] + by (metis (no_types, lifting) finite_Un finite_atLeastLessThan) + also have "... = (\ i \ {i \ {0.. (M $$ (i, j2) = 1) } . 1) + + (\ i \ {i \ {0.. M $$ (i, j2) = 0} . 0)" + using sum.cong mem_Collect_eq by (smt (z3) mult.right_neutral mult_not_zero) + finally have "(col M j1) \ (col M j2) = + of_nat (card {i . i < dim_row M \ (M $$ (i, j1) = 1) \ (M $$ (i, j2) = 1)})" + by simp + then show ?thesis using mat_inter_num_def[of M j1 j2] by simp +qed + +end + +text \ Any matrix generated by @{term "inc_mat_of"} is a 0-1 matrix.\ +lemma inc_mat_of_01_mat: "zero_one_matrix_ring_1 (inc_mat_of Vs Bs)" + by (unfold_locales) (simp add: inc_mat_one_zero_elems) + +subsection \Ordered Incidence Systems \ +text \We impose an arbitrary ordering on the point set and block collection to enable +matrix reasoning. Note that this is also common in computer algebra representations of designs \ + +locale ordered_incidence_system = + fixes \s :: "'a list" and \s :: "'a set list" + assumes wf_list: "b \# (mset \s) \ b \ set \s" + assumes distinct: "distinct \s" + +text \An ordered incidence system, as it is defined on lists, can only represent finite incidence systems \ +sublocale ordered_incidence_system \ finite_incidence_system "set \s" "mset \s" + by(unfold_locales) (auto simp add: wf_list) + +lemma ordered_incidence_sysI: + assumes "finite_incidence_system \ \" + assumes "\s \ permutations_of_set \" and "\s \ permutations_of_multiset \" + shows "ordered_incidence_system \s \s" +proof - + have veq: "\ = set \s" using assms permutations_of_setD(1) by auto + have beq: "\ = mset \s" using assms permutations_of_multisetD by auto + interpret fisys: finite_incidence_system "set \s" "mset \s" using assms(1) veq beq by simp + show ?thesis proof (unfold_locales) + show "\b. b \# mset \s \ b \ set \s" using fisys.wellformed + by simp + show "distinct \s" using assms permutations_of_setD(2) by auto + qed +qed + +lemma ordered_incidence_sysII: + assumes "finite_incidence_system \ \" and "set \s = \" and "distinct \s" and "mset \s = \" + shows "ordered_incidence_system \s \s" +proof - + interpret fisys: finite_incidence_system "set \s" "mset \s" using assms by simp + show ?thesis using fisys.wellformed assms by (unfold_locales) (simp_all) +qed + +context ordered_incidence_system +begin +text \For ease of notation, establish the same notation as for incidence systems \ + +abbreviation "\ \ set \s" +abbreviation "\ \ mset \s" + +text \Basic properties on ordered lists \ +lemma points_indexing: "\s \ permutations_of_set \" + by (simp add: permutations_of_set_def distinct) + +lemma blocks_indexing: "\s \ permutations_of_multiset \" + by (simp add: permutations_of_multiset_def) + +lemma points_list_empty_iff: "\s = [] \ \ = {}" + using finite_sets points_indexing + by (simp add: elem_permutation_of_set_empty_iff) + +lemma points_indexing_inj: "\ i \ I . i < length \s \ inj_on ((!) \s) I" + by (simp add: distinct inj_on_nth) + +lemma blocks_list_empty_iff: "\s = [] \ \ = {#}" + using blocks_indexing by (simp) + +lemma blocks_list_nempty: "proper_design \ \ \ \s \ []" + using mset.simps(1) proper_design.design_blocks_nempty by blast + +lemma points_list_nempty: "proper_design \ \ \ \s \ []" + using proper_design.design_points_nempty points_list_empty_iff by blast + +lemma points_list_length: "length \s = \" + using points_indexing + by (simp add: length_finite_permutations_of_set) + +lemma blocks_list_length: "length \s = \" + using blocks_indexing length_finite_permutations_of_multiset by blast + +lemma valid_points_index: "i < \ \ \s ! i \ \" + using points_list_length by simp + +lemma valid_points_index_cons: "x \ \ \ \ i. \s ! i = x \ i < \" + using points_list_length by (auto simp add: in_set_conv_nth) + +lemma valid_points_index_obtains: + assumes "x \ \" + obtains i where "\s ! i = x \ i < \" + using valid_points_index_cons assms by auto + +lemma valid_blocks_index: "j < \ \ \s ! j \# \" + using blocks_list_length by (metis nth_mem_mset) + +lemma valid_blocks_index_cons: "bl \# \ \ \ j . \s ! j = bl \ j < \" + by (auto simp add: in_set_conv_nth) + +lemma valid_blocks_index_obtains: + assumes "bl \# \" + obtains j where "\s ! j = bl \ j < \" + using assms valid_blocks_index_cons by auto + +lemma block_points_valid_point_index: + assumes "bl \# \" "x \ bl" + obtains i where "i < length \s \ \s ! i = x" + using wellformed valid_points_index_obtains assms + by (metis points_list_length wf_invalid_point) + +lemma points_set_index_img: "\ = image(\ i . (\s ! i)) ({..<\})" + using valid_points_index_cons valid_points_index by auto + +lemma blocks_mset_image: "\ = image_mset (\ i . (\s ! i)) (mset_set {..<\})" + by (simp add: mset_list_by_index) + +lemma incidence_cond_indexed[simp]: "i < \ \ j < \ \ incident (\s ! i) (\s ! j) \ (\s ! i) \ (\s ! j)" + using incidence_alt_def valid_points_index valid_blocks_index by simp + +lemma bij_betw_points_index: "bij_betw (\ i. \s ! i) {0..<\} \" +proof (simp add: bij_betw_def, intro conjI) + show "inj_on ((!) \s) {0..<\}" + by (simp add: points_indexing_inj points_list_length) + show "(!) \s ` {0..<\} = \" + proof (intro subset_antisym subsetI) + fix x assume "x \ (!) \s ` {0..<\}" + then obtain i where "x = \s ! i" and "i < \" by auto + then show "x \ \" + by (simp add: valid_points_index) + next + fix x assume "x \ \" + then obtain i where "\s ! i = x" and "i <\" + using valid_points_index_cons by auto + then show "x \ (!) \s ` {0..<\}" by auto + qed +qed + +text \Some lemmas on cardinality due to different set descriptor filters \ +lemma card_filter_point_indices: "card {i \ {0..<\}. P (\s ! i)} = card {v \ \ . P v }" +proof - + have "{v \ \ . P v }= (\ i. \s ! i) ` {i \ {0..<\}. P (\s ! i)}" + by (metis Compr_image_eq lessThan_atLeast0 points_set_index_img) + thus ?thesis using inj_on_nth points_list_length + by (metis (no_types, lifting) card_image distinct lessThan_atLeast0 lessThan_iff mem_Collect_eq) +qed + +lemma card_block_points_filter: + assumes "j < \" + shows "card (\s ! j) = card {i \ {0..<\} . (\s ! i) \ (\s ! j)}" +proof - + obtain bl where "bl \# \" and blis: "bl = \s ! j" + using assms by auto + then have cbl: "card bl = card {v \ \ . v \ bl}" using block_size_alt by simp + have "\ = (\ i. \s ! i) ` {0..<\}" using bij_betw_points_index + using lessThan_atLeast0 points_set_index_img by presburger + then have "Set.filter (\ v . v \ bl) \ = Set.filter (\ v . v \ bl) ((\ i. \s ! i) ` {0..<\})" + by presburger + have "card {i \ {0..<\} . (\s ! i) \ bl} = card {v \ \ . v \ bl}" + using card_filter_point_indices by simp + thus ?thesis using cbl blis by simp +qed + +lemma obtains_two_diff_block_indexes: + assumes "j1 < \" + assumes "j2 < \" + assumes "j1 \ j2" + assumes "\ \ 2" + obtains bl1 bl2 where "bl1 \# \" and "\s ! j1 = bl1" and "bl2 \# \ - {#bl1#}" and "\s ! j2 = bl2" +proof - + have j1lt: "min j1 (length \s) = j1" using assms by auto + obtain bl1 where bl1in: "bl1 \# \" and bl1eq: "\s ! j1 = bl1" + using assms(1) valid_blocks_index by blast + then have split: "\s = take j1 \s @ \s!j1 # drop (Suc j1) \s" + using assms id_take_nth_drop by auto + then have lj1: "length (take j1 \s) = j1" using j1lt by (simp add: length_take[of j1 \s]) + have "\ = mset (take j1 \s @ \s!j1 # drop (Suc j1) \s)" using split assms(1) by presburger + then have bsplit: "\ = mset (take j1 \s) + {#bl1#} + mset (drop (Suc j1) \s)" by (simp add: bl1eq) + then have btake: "\ - {#bl1#} = mset (take j1 \s) + mset (drop (Suc j1) \s)" by simp + thus ?thesis proof (cases "j2 < j1") + case True + then have "j2 < length (take j1 \s)" using lj1 by simp + then obtain bl2 where bl2eq: "bl2 = (take j1 \s) ! j2" by auto + then have bl2eq2: "bl2 = \s ! j2" + by (simp add: True) + then have "bl2 \# \ - {#bl1#}" using btake + by (metis bl2eq \j2 < length (take j1 \s)\ nth_mem_mset union_iff) + then show ?thesis using bl2eq2 bl1in bl1eq that by auto + next + case False + then have j2gt: "j2 \ Suc j1" using assms by simp + then obtain i where ieq: "i = j2 - Suc j1" + by simp + then have j2eq: "j2 = (Suc j1) + i" using j2gt by presburger + have "length (drop (Suc j1) \s) = \ - (Suc j1)" using blocks_list_length by auto + then have "i < length (drop (Suc j1) \s)" using ieq assms blocks_list_length + using diff_less_mono j2gt by presburger + then obtain bl2 where bl2eq: "bl2 = (drop (Suc j1) \s) ! i" by auto + then have bl2in: "bl2 \# \ - {#bl1#}" using btake nth_mem_mset union_iff + by (metis \i < length (drop (Suc j1) \s)\) + then have "bl2 = \s ! j2" using bl2eq nth_drop blocks_list_length assms j2eq + by (metis Suc_leI) + then show ?thesis using bl1in bl1eq bl2in that by auto + qed +qed + +lemma filter_size_blocks_eq_card_indexes: "size {# b \# \ . P b #} = card {j \ {..<(\)}. P (\s ! j)}" +proof - + have "\ = image_mset (\ j . \s ! j) (mset_set {..<(\)})" + using blocks_mset_image by simp + then have helper: "{# b \# \ . P b #} = image_mset (\ j . \s ! j) {# j \# (mset_set {..< \}). P (\s ! j) #} " + by (simp add: filter_mset_image_mset) + have "card {j \ {..<\}. P (\s ! j)} = size {# j \# (mset_set {..< \}). P (\s ! j) #}" + using card_size_filter_eq [of "{..<\}"] by simp + thus ?thesis using helper by simp +qed + +lemma blocks_index_ne_belong: + assumes "i1 < length \s" + assumes "i2 < length \s" + assumes "i1 \ i2" + shows "\s ! i2 \# \ - {#(\s ! i1)#}" +proof (cases "\s ! i1 = \s ! i2") + case True + then have "count (mset \s) (\s ! i1) \ 2" using count_min_2_indices assms by fastforce + then have "count ((mset \s) - {#(\s ! i1)#}) (\s ! i1) \ 1" + by (metis Nat.le_diff_conv2 add_leD2 count_diff count_single nat_1_add_1) + then show ?thesis + by (metis True count_inI not_one_le_zero) +next + case False + have "\s ! i2 \# \" using assms + by simp + then show ?thesis using False + by (metis in_remove1_mset_neq) +qed + +lemma inter_num_points_filter_def: + assumes "j1 < \" "j2 < \" "j1 \ j2" + shows "card {x \ {0..<\} . ((\s ! x) \ (\s ! j1) \ (\s ! x) \ (\s ! j2)) } = (\s ! j1) |\| (\s ! j2)" +proof - + have inter: "\ v. v \ \ \ v \ (\s ! j1) \ v \ (\s ! j2) \ v \ (\s ! j1) \ (\s ! j2)" + by simp + obtain bl1 bl2 where bl1in: "bl1 \# \" and bl1eq: "\s ! j1 = bl1" and bl2in: "bl2 \# \ - {#bl1#}" + and bl2eq: "\s ! j2 = bl2" + using assms obtains_two_diff_block_indexes + by (metis blocks_index_ne_belong size_mset valid_blocks_index) + have "card {x \ {0..<\} . (\s ! x) \ (\s ! j1) \ (\s ! x) \ (\s ! j2) } = + card {v \ \ . v \ (\s ! j1) \ v \ (\s ! j2) }" + using card_filter_point_indices by simp + also have "... = card {v \ \ . v \ bl1 \ v \ bl2 }" using bl1eq bl2eq by simp + finally show ?thesis using points_inter_num_rep bl1in bl2in + by (simp add: bl1eq bl2eq) +qed + +text \Define an incidence matrix for this ordering of an incidence system \ + +abbreviation N :: "int mat" where +"N \ inc_mat_of \s \s" + +sublocale zero_one_matrix_ring_1 "N" + using inc_mat_of_01_mat . + +lemma N_alt_def_dim: "N = mat \ \ (\ (i,j) . if (incident (\s ! i) (\s ! j)) then 1 else 0) " + using incidence_cond_indexed inc_mat_of_def + by (intro eq_matI) (simp_all add: inc_mat_dim_row inc_mat_dim_col inc_matrix_point_in_block_one + inc_matrix_point_not_in_block_zero points_list_length) + +text \Matrix Dimension related lemmas \ + +lemma N_carrier_mat: "N \ carrier_mat \ \" + by (simp add: N_alt_def_dim) + +lemma dim_row_is_v[simp]: "dim_row N = \" + by (simp add: N_alt_def_dim) + +lemma dim_col_is_b[simp]: "dim_col N = \" + by (simp add: N_alt_def_dim) + +lemma dim_vec_row_N: "dim_vec (row N i) = \" + by (simp add: N_alt_def_dim) + +lemma dim_vec_col_N: "dim_vec (col N i) = \" by simp + +lemma dim_vec_N_col: + assumes "j < \" + shows "dim_vec (cols N ! j) = \" +proof - + have "cols N ! j = col N j" using assms dim_col_is_b by simp + then have "dim_vec (cols N ! j) = dim_vec (col N j)" by simp + thus ?thesis using dim_col assms by (simp) +qed + +lemma N_carrier_mat_01_lift: "lift_01_mat N \ carrier_mat \ \" + by auto + +text \Transpose properties \ + +lemma transpose_N_mult_dim: "dim_row (N * N\<^sup>T) = \" "dim_col (N * N\<^sup>T) = \" + by (simp_all) + +lemma N_trans_index_val: "i < dim_col N \ j < dim_row N \ + N\<^sup>T $$ (i, j) = (if (\s ! j) \ (\s ! i) then 1 else 0)" + by (simp add: inc_mat_of_def) + +text \Matrix element and index related lemmas \ +lemma mat_row_elems: "i < \ \ vec_set (row N i) \ {0, 1}" + using points_list_length + by (simp add: row_elems_ss01) + +lemma mat_col_elems: "j < \ \ vec_set (col N j) \ {0, 1}" + using blocks_list_length by (metis col_elems_ss01 dim_col_is_b) + +lemma matrix_elems_one_zero: "i < \ \ j < \ \ N $$ (i, j) = 0 \ N $$ (i, j) = 1" + by (metis blocks_list_length inc_matrix_elems_one_zero points_list_length) + +lemma matrix_point_in_block_one: "i < \ \ j < \ \ (\s ! i)\ (\s ! j) \N $$ (i, j) = 1" + by (metis inc_matrix_point_in_block_one points_list_length blocks_list_length ) + +lemma matrix_point_not_in_block_zero: "i < \ \ j < \ \ \s ! i \ \s ! j \ N $$ (i, j) = 0" + by(metis inc_matrix_point_not_in_block_zero points_list_length blocks_list_length) + +lemma matrix_point_in_block: "i < \ \ j < \ \ N $$ (i, j) = 1 \ \s ! i \ \s ! j" + by (metis blocks_list_length points_list_length inc_matrix_point_in_block) + +lemma matrix_point_not_in_block: "i < \ \ j < \ \ N $$ (i, j) = 0 \ \s ! i \ \s ! j" + by (metis blocks_list_length points_list_length inc_matrix_point_not_in_block) + +lemma matrix_point_not_in_block_iff: "i < \ \ j < \ \ N $$ (i, j) = 0 \ \s ! i \ \s ! j" + by (metis blocks_list_length points_list_length inc_matrix_point_not_in_block_iff) + +lemma matrix_point_in_block_iff: "i < \ \ j < \ \ N $$ (i, j) = 1 \ \s ! i \ \s ! j" + by (metis blocks_list_length points_list_length inc_matrix_point_in_block_iff) + +lemma matrix_subset_implies_one: "I \ {..< \} \ j < \ \ (!) \s ` I \ \s ! j \ i \ I \ + N $$ (i, j) = 1" + by (metis blocks_list_length points_list_length inc_matrix_subset_implies_one) + +lemma matrix_one_implies_membership: +"I \ {..< \} \ j < size \ \ \i\I. N $$ (i, j) = 1 \ i \ I \ \s ! i \ \s ! j" + by (simp add: matrix_point_in_block_iff subset_iff) + +text \Incidence Vector's of Incidence Matrix columns \ + +lemma col_inc_vec_of: "j < length \s \ inc_vec_of \s (\s ! j) = col N j" + by (simp add: inc_mat_col_inc_vec) + +lemma inc_vec_eq_iff_blocks: + assumes "bl \# \" + assumes "bl' \# \" + shows "inc_vec_of \s bl = inc_vec_of \s bl' \ bl = bl'" +proof (intro iffI eq_vecI, simp_all add: inc_vec_dim assms) + define v1 :: "'c :: {ring_1} vec" where "v1 = inc_vec_of \s bl" + define v2 :: "'c :: {ring_1} vec" where "v2 = inc_vec_of \s bl'" + assume a: "v1 = v2" + then have "dim_vec v1 = dim_vec v2" + by (simp add: inc_vec_dim) + then have "\ i. i < dim_vec v1 \ v1 $ i = v2 $ i" using a by simp + then have "\ i. i < length \s \ v1 $ i = v2 $ i" by (simp add: v1_def inc_vec_dim) + then have "\ i. i < length \s \ (\s ! i) \ bl \ (\s ! i) \ bl'" + using inc_vec_index_one_iff v1_def v2_def by metis + then have "\ x. x \ \ \ x \ bl \ x \ bl'" + using points_list_length valid_points_index_cons by auto + then show "bl = bl'" using wellformed assms + by (meson subset_antisym subset_eq) +qed + +text \Incidence matrix column properties\ + +lemma N_col_def: "j < \ \ i < \ \ (col N j) $ i = (if (\s ! i \ \s ! j) then 1 else 0)" + by (metis inc_mat_col_def points_list_length blocks_list_length) + +lemma N_col_def_indiv: "j < \ \ i < \ \ \s ! i \ \s ! j \ (col N j) $ i = 1" + "j < \ \ i < \ \ \s ! i \ \s ! j \ (col N j) $ i = 0" + by(simp_all add: inc_matrix_point_in_block_one inc_matrix_point_not_in_block_zero points_list_length) + +lemma N_col_list_map_elem: "j < \ \ i < \ \ + col N j $ i = map_vec (\ x . if (x \ (\s ! j)) then 1 else 0) (vec_of_list \s) $ i" + by (metis inc_mat_col_list_map_elem points_list_length blocks_list_length) + +lemma N_col_list_map: "j < \ \ col N j = map_vec (\ x . if (x \ (\s ! j)) then 1 else 0) (vec_of_list \s)" + by (metis inc_mat_col_list_map blocks_list_length) + +lemma N_col_mset_point_set_img: "j < \ \ + vec_mset (col N j) = image_mset (\ x. if (x \ (\s ! j)) then 1 else 0) (mset_set \)" + using vec_mset_img_map N_col_list_map points_indexing + by (metis (no_types, lifting) finite_sets permutations_of_multisetD permutations_of_set_altdef) + +lemma matrix_col_to_block: + assumes "j < \" + shows "\s ! j = (\ k . \s ! k) ` {i \ {..< \} . (col N j) $ i = 1}" +proof (intro subset_antisym subsetI) + fix x assume assm1: "x \ \s ! j" + then have "x \ \" using wellformed assms valid_blocks_index by blast + then obtain i where vs: "\s ! i = x" and "i < \" + using valid_points_index_cons by auto + then have inset: "i \ {..< \}" + by fastforce + then have "col N j $ i = 1" using assm1 N_col_def assms vs + using \i < \\ by presburger + then have "i \ {i. i \ {..< \} \ col N j $ i = 1}" + using inset by blast + then show "x \ (!) \s ` {i. i \ {..<\} \ col N j $ i = 1}" using vs by blast +next + fix x assume assm2: "x \ ((\ k . \s ! k) ` {i \ {..< \} . col N j $ i = 1})" + then obtain k where "x = \s !k" and inner: "k \{i \ {..< \} . col N j $ i = 1}" + by blast + then have ilt: "k < \" by auto + then have "N $$ (k, j) = 1" using inner + by (metis (mono_tags) N_col_def assms matrix_point_in_block_iff matrix_point_not_in_block_zero mem_Collect_eq) + then show "x \ \s ! j" using ilt + using \x = \s ! k\ assms matrix_point_in_block_iff by blast +qed + +lemma matrix_col_to_block_v2: "j < \ \ \s ! j = (\ k . \s ! k) ` map_col_to_block (col N j)" + using matrix_col_to_block map_col_to_block_def by fastforce + +lemma matrix_col_in_blocks: "j < \ \ (!) \s ` map_col_to_block (col N j) \# \" + using matrix_col_to_block_v2 by (metis (no_types, lifting) valid_blocks_index) + +lemma inc_matrix_col_block: + assumes "c \ set (cols N)" + shows "(\ x. \s ! x) ` (map_col_to_block c) \# \" +proof - + obtain j where "c = col N j" and "j < \" using assms cols_length cols_nth in_mset_conv_nth + ordered_incidence_system_axioms set_mset_mset by (metis dim_col_is_b) + thus ?thesis + using matrix_col_in_blocks by blast +qed + +text \ Incidence Matrix Row Definitions \ +lemma N_row_def: "j < \ \ i < \ \ (row N i) $ j = (if (\s ! i \ \s ! j) then 1 else 0)" + by (metis inc_mat_row_def points_list_length blocks_list_length) + +lemma N_row_list_map_elem: "j < \ \ i < \ \ + row N i $ j = map_vec (\ bl . if ((\s ! i) \ bl) then 1 else 0) (vec_of_list \s) $ j" + by (metis inc_mat_row_list_map_elem points_list_length blocks_list_length) + +lemma N_row_list_map: "i < \ \ + row N i = map_vec (\ bl . if ((\s ! i) \ bl) then 1 else 0) (vec_of_list \s)" + by (simp add: inc_mat_row_list_map points_list_length blocks_list_length) + +lemma N_row_mset_blocks_img: "i < \ \ + vec_mset (row N i) = image_mset (\ x . if ((\s ! i) \ x) then 1 else 0) \" + using vec_mset_img_map N_row_list_map by metis + +text \Alternate Block representations \ + +lemma block_mat_cond_rep: + assumes "j < length \s" + shows "(\s ! j) = {\s ! i | i. i < length \s \ N $$ (i, j) = 1}" +proof - + have cond: "\ i. i < length \s \ N $$ (i, j) = 1 \i \ {..< \} \ (col N j) $ i = 1" + using assms points_list_length by auto + have "(\s ! j) = (\ k . \s ! k) ` {i \ {..< \} . (col N j) $ i = 1}" + using matrix_col_to_block assms by simp + also have "... = {\s ! i | i. i \ {..< \} \ (col N j) $ i = 1}" by auto + finally show "(\s ! j) = {\s ! i | i. i < length \s \ N $$ (i, j) = 1}" + using Collect_cong cond by auto +qed + +lemma block_mat_cond_rep': "j < length \s \ (\s ! j) = ((!) \s) ` {i . i < length \s \ N $$ (i, j) = 1}" + by (simp add: block_mat_cond_rep setcompr_eq_image) + +lemma block_mat_cond_rev: + assumes "j < length \s" + shows "{i . i < length \s \ N $$ (i, j) = 1} = ((List_Index.index) \s) ` (\s ! j)" +proof (intro Set.set_eqI iffI) + fix i assume a1: "i \ {i. i < length \s \ N $$ (i, j) = 1}" + then have ilt1: "i < length \s" and Ni1: "N $$ (i, j) = 1" by auto + then obtain x where "\s ! i = x" and "x \ (\s ! j)" + using assms inc_matrix_point_in_block by blast + then have "List_Index.index \s x = i" using distinct index_nth_id ilt1 by auto + then show "i \ List_Index.index \s ` \s ! j" by (metis \x \ \s ! j\ imageI) +next + fix i assume a2: "i \ List_Index.index \s ` \s ! j" + then obtain x where ieq: "i = List_Index.index \s x" and xin: "x \ \s !j" + by blast + then have ilt: "i < length \s" + by (smt (z3) assms index_first index_le_size nat_less_le nth_mem_mset points_list_length + valid_points_index_cons wf_invalid_point) + then have "N $$ (i, j) = 1" using xin inc_matrix_point_in_block_one + by (metis ieq assms index_conv_size_if_notin less_irrefl_nat nth_index) + then show "i \ {i. i < length \s \ N $$ (i, j) = 1}" using ilt by simp +qed + +text \Incidence Matrix incidence system properties \ +lemma incomplete_block_col: + assumes "j < \" + assumes "incomplete_block (\s ! j)" + shows "0 \$ (col N j)" +proof - + obtain x where "x \ \" and "x \ (\s ! j)" + by (metis Diff_iff assms(2) incomplete_block_proper_subset psubset_imp_ex_mem) + then obtain i where "\s ! i = x" and "i< \" + using valid_points_index_cons by blast + then have "N $$ (i, j) = 0" + using \x \ \s ! j\ assms(1) matrix_point_not_in_block_zero by blast + then have "col N j $ i = 0" + using N_col_def \\s ! i = x\ \i < \\ \x \ \s ! j\ assms(1) by fastforce + thus ?thesis using vec_setI + by (smt (z3) \i < \\ dim_col dim_row_is_v) +qed + +lemma mat_rep_num_N_row: + assumes "i < \" + shows "mat_rep_num N i = \ rep (\s ! i)" +proof - + have "count (image_mset (\ x . if ((\s ! i) \ x) then 1 else (0 :: int )) \) 1 = + size (filter_mset (\ x . (\s ! i) \ x) \)" + using count_mset_split_image_filter[of "\" "1" "\ x . (0 :: int)" "\ x . (\s ! i) \ x"] by simp + then have "count (image_mset (\ x . if ((\s ! i) \ x) then 1 else (0 :: int )) \) 1 + = \ rep (\s ! i)" by (simp add: point_rep_number_alt_def) + thus ?thesis using N_row_mset_blocks_img assms + by (simp add: mat_rep_num_def) +qed + +lemma point_rep_mat_row_sum: "i < \ \ sum_vec (row N i) = \ rep (\s ! i)" + using count_vec_sum_ones_alt mat_rep_num_N_row mat_row_elems mat_rep_num_def by metis + +lemma mat_block_size_N_col: + assumes "j < \" + shows "mat_block_size N j = card (\s ! j)" +proof - + have val_b: "\s ! j \# \" using assms valid_blocks_index by auto + have "\ x. x \# mset_set \ \ (\x . (0 :: int)) x \ 1" using zero_neq_one by simp + then have "count (image_mset (\ x. if (x \ (\s ! j)) then 1 else (0 :: int)) (mset_set \)) 1 = + size (filter_mset (\ x . x \ (\s ! j)) (mset_set \))" + using count_mset_split_image_filter [of "mset_set \" "1" "(\ x . (0 :: int))" "\ x . x \ \s ! j"] + by simp + then have "count (image_mset (\ x. if (x \ (\s ! j)) then 1 else (0 :: int)) (mset_set \)) 1 = card (\s ! j)" + using val_b block_size_alt by (simp add: finite_sets) + thus ?thesis using N_col_mset_point_set_img assms mat_block_size_def by metis +qed + +lemma block_size_mat_rep_sum: "j < \ \ sum_vec (col N j) = mat_block_size N j" + using count_vec_sum_ones_alt mat_block_size_N_col mat_block_size_def by (metis mat_col_elems) + +lemma mat_point_index_rep: + assumes "I \ {..<\}" + shows "mat_point_index N I = \ index ((\ i. \s ! i) ` I)" +proof - + have "\ i . i \ I \ \s ! i \ \" using assms valid_points_index by auto + then have eqP: "\ j. j < dim_col N \ ((\ i. \s ! i) ` I) \ (\s ! j) \ (\ i \ I . N $$ (i, j) = 1)" + proof (intro iffI subsetI, simp_all) + show "\j i. j < length \s \ (\i. i \ I \ \s ! i \ \) \ (!) \s ` I \ \s ! j \ + \i\I. N $$ (i, j) = 1" + using matrix_subset_implies_one assms by simp + have "\x. x\ (!) \s ` I \ \ i \ I. \s ! i = x" + by auto + then show "\j x. j < length \s \ \i\I. N $$ (i, j) = 1 \ x \ (!) \s ` I + \ (\i. i \ I \ \s ! i \ \) \ x \ \s ! j" + using assms matrix_one_implies_membership by (metis blocks_list_length) + qed + have "card {j . j < dim_col N \ (\ i \ I . N $$(i, j) = 1)} = + card {j . j < dim_col N \ ((\ i . \s ! i) ` I) \ \s ! j}" + using eqP by (metis (mono_tags, lifting)) + also have "... = size {# b \# \ . ((\ i . \s ! i) ` I) \ b #}" + using filter_size_blocks_eq_card_indexes by auto + also have "... = points_index \ ((\ i . \s ! i) ` I)" + by (simp add: points_index_def) + finally have "card {j . j < dim_col N \ (\ i \ I . N $$(i, j) = 1)} = \ index ((\ i . \s ! i) ` I)" + by blast + thus ?thesis unfolding mat_point_index_def by simp +qed + +lemma incidence_mat_two_index: "i1 < \ \ i2 < \ \ + mat_point_index N {i1, i2} = \ index {\s ! i1, \s ! i2}" + using mat_point_index_two_alt[of i1 N i2 ] mat_point_index_rep[of "{i1, i2}"] dim_row_is_v + by (metis (no_types, lifting) empty_subsetI image_empty image_insert insert_subset lessThan_iff) + +lemma ones_incidence_mat_block_size: + assumes "j < \" + shows "((u\<^sub>v \) \<^sub>v* N) $ j = mat_block_size N j" +proof - + have "dim_vec ((u\<^sub>v \) \<^sub>v* N) = \" by (simp) + then have "((u\<^sub>v \) \<^sub>v* N) $ j = (u\<^sub>v \) \ col N j" using assms by simp + also have "... = (\ i \ {0 ..< \}. (u\<^sub>v \) $ i * (col N j) $ i)" + by (simp add: scalar_prod_def) + also have "... = sum_vec (col N j)" using dim_row_is_v by (simp add: sum_vec_def) + finally show ?thesis using block_size_mat_rep_sum assms by simp +qed + +lemma mat_block_size_conv: "j < dim_col N \ card (\s ! j) = mat_block_size N j" + by (simp add: mat_block_size_N_col) + +lemma mat_inter_num_conv: + assumes "j1 < dim_col N" "j2 < dim_col N" + shows "(\s ! j1) |\| (\s ! j2) = mat_inter_num N j1 j2" +proof - + have eq_sets: "\ P. (\ i . \s ! i) ` {i \ {0..<\}. P (\s ! i)} = {x \ \ . P x}" + by (metis Compr_image_eq lessThan_atLeast0 points_set_index_img) + have bin: "\s ! j1 \# \" "\s ! j2 \# \" using assms dim_col_is_b by simp_all + have "(\s ! j1) |\| (\s ! j2) = card ((\s ! j1) \ (\s ! j2))" + by (simp add: intersection_number_def) + also have "... = card {x . x \ (\s ! j1) \ x \ (\s ! j2)}" + by (simp add: Int_def) + also have "... = card {x \ \. x \ (\s ! j1) \ x \ (\s ! j2)}" using wellformed bin + by (meson wf_invalid_point) + also have "... = card ((\ i . \s ! i) ` {i \ {0..<\}. (\s ! i) \ (\s ! j1) \ (\s ! i) \ (\s ! j2)})" + using eq_sets[of "\ x. x \ (\s ! j1) \ x \ (\s ! j2)"] by simp + also have "... = card ({i \ {0..<\}. (\s ! i) \ (\s ! j1) \ (\s ! i) \ (\s ! j2)})" + using points_indexing_inj card_image + by (metis (no_types, lifting) lessThan_atLeast0 lessThan_iff mem_Collect_eq points_list_length) + also have "... = card ({i . i < \ \ (\s ! i) \ (\s ! j1) \ (\s ! i) \ (\s ! j2)})" by auto + also have "... = card ({i . i < \ \ N $$ (i, j1) = 1 \ N $$ (i, j2) = 1})" using assms + by (metis (no_types, opaque_lifting) inc_mat_dim_col inc_matrix_point_in_block_iff points_list_length) + finally have "(\s ! j1) |\| (\s ! j2) = card {i . i < dim_row N \ N $$ (i, j1) = 1 \ N $$ (i, j2) = 1}" + using dim_row_is_v by presburger + thus ?thesis using assms by (simp add: mat_inter_num_def) +qed + +lemma non_empty_col_map_conv: + assumes "j < dim_col N" + shows "non_empty_col N j \ \s ! j \ {}" +proof (intro iffI) + assume "non_empty_col N j" + then obtain i where ilt: "i < dim_row N" and "(col N j) $ i \ 0" + using non_empty_col_obtains assms by blast + then have "(col N j) $ i = 1" + using assms + by (metis N_col_def_indiv(1) N_col_def_indiv(2) dim_col_is_b dim_row_is_v) + then have "\s ! i \ \s ! j" + by (smt (verit, best) assms ilt inc_mat_col_def dim_col_is_b inc_mat_dim_col inc_mat_dim_row) + thus "\s ! j \ {}" by blast +next + assume a: "\s ! j \ {}" + have "\s ! j \# \" using assms dim_col_is_b by simp + then obtain x where "x \ \s ! j" and "x \ \" using wellformed a by auto + then obtain i where "\s ! i \ \s ! j" and "i < dim_row N" using dim_row_is_v + using valid_points_index_cons by auto + then have "N $$ (i, j) = 1" + using assms by (meson inc_mat_of_index) + then show "non_empty_col N j" using non_empty_col_alt_def + using \i < dim_row N\ assms by fastforce +qed + +lemma scalar_prod_inc_vec_inter_num: + assumes "j1 < \" "j2 < \" + shows "(col N j1) \ (col N j2) = (\s ! j1) |\| (\s ! j2)" + using scalar_prod_inc_vec_mat_inter_num assms N_carrier_mat + by (simp add: mat_inter_num_conv) + +lemma scalar_prod_block_size_lift_01: + assumes "i < \" + shows "((col (lift_01_mat N) i) \ (col (lift_01_mat N) i)) = (of_nat (card (\s ! i)) :: ('b :: {ring_1}))" +proof - + interpret z1: zero_one_matrix_ring_1 "(lift_01_mat N)" + by (intro_locales) (simp add: lift_mat_is_0_1) + show ?thesis using assms z1.scalar_prod_inc_vec_block_size_mat preserve_mat_block_size + mat_block_size_N_col lift_01_mat_def + by (metis inc_mat_dim_col lift_01_mat_simp(2) of_inj_on_01_hom.inj_on_01_hom_axioms size_mset) +qed + +lemma scalar_prod_inter_num_lift_01: + assumes "j1 < \" "j2 < \" + shows "((col (lift_01_mat N) j1) \ (col (lift_01_mat N) j2)) = (of_nat ((\s ! j1) |\| (\s ! j2)) :: ('b :: {ring_1}))" +proof - + interpret z1: zero_one_matrix_ring_1 "(lift_01_mat N)" + by (intro_locales) (simp add: lift_mat_is_0_1) + show ?thesis using assms z1.scalar_prod_inc_vec_mat_inter_num preserve_mat_inter_num + mat_inter_num_conv lift_01_mat_def blocks_list_length inc_mat_dim_col + by (metis lift_01_mat_simp(2) of_inj_on_01_hom.inj_on_01_hom_axioms) +qed + +text \ The System complement's incidence matrix flips 0's and 1's \ + +lemma map_block_complement_entry: "j < \ \ (map block_complement \s) ! j = block_complement (\s ! j)" + using blocks_list_length by (metis nth_map) + +lemma complement_mat_entries: + assumes "i < \" and "j < \" + shows "(\s ! i \ \s ! j) \ (\s ! i \ (map block_complement \s) ! j)" + using assms block_complement_def map_block_complement_entry valid_points_index by simp + +lemma length_blocks_complement: "length (map block_complement \s) = \" + by auto + +lemma ordered_complement: "ordered_incidence_system \s (map block_complement \s)" +proof - + interpret inc: finite_incidence_system \ "complement_blocks" + by (simp add: complement_finite) + have "map inc.block_complement \s \ permutations_of_multiset complement_blocks" + using complement_image by (simp add: permutations_of_multiset_def) + then show ?thesis using ordered_incidence_sysI[of "\" "complement_blocks" "\s" "(map block_complement \s)"] + by (simp add: inc.finite_incidence_system_axioms points_indexing) +qed + +interpretation ordered_comp: ordered_incidence_system "\s" "(map block_complement \s)" + using ordered_complement by simp + +lemma complement_mat_entries_val: + assumes "i < \" and "j < \" + shows "ordered_comp.N $$ (i, j) = (if \s ! i \ \s ! j then 0 else 1)" +proof - + have cond: "(\s ! i \ \s ! j) \ (\s ! i \ (map block_complement \s) ! j)" + using complement_mat_entries assms by simp + then have "ordered_comp.N $$ (i, j) = (if (\s ! i \ (map block_complement \s) ! j) then 1 else 0)" + using assms ordered_comp.matrix_point_in_block_one ordered_comp.matrix_point_not_in_block_iff + by force + then show ?thesis using cond by simp +qed + +lemma ordered_complement_mat: "ordered_comp.N = mat \ \ (\ (i,j) . if (\s ! i) \ (\s ! j) then 0 else 1)" + using complement_mat_entries_val by (intro eq_matI, simp_all) + +lemma ordered_complement_mat_map: "ordered_comp.N = map_mat (\x. if x = 1 then 0 else 1) N" + apply (intro eq_matI, simp_all) + using ordered_incidence_system.matrix_point_in_block_iff ordered_incidence_system_axioms + complement_mat_entries_val by (metis blocks_list_length) + + +end + +text \Establishing connection between incidence system and ordered incidence system locale \ + +lemma (in incidence_system) alt_ordering_sysI: "Vs \ permutations_of_set \ \ Bs \ permutations_of_multiset \ \ + ordered_incidence_system Vs Bs" + by (unfold_locales) (simp_all add: permutations_of_multisetD permutations_of_setD wellformed) + +lemma (in finite_incidence_system) exists_ordering_sysI: "\ Vs Bs . Vs \ permutations_of_set \ \ + Bs \ permutations_of_multiset \ \ ordered_incidence_system Vs Bs" +proof - + obtain Vs where "Vs \ permutations_of_set \" + by (meson all_not_in_conv finite_sets permutations_of_set_empty_iff) + obtain Bs where "Bs \ permutations_of_multiset \" + by (meson all_not_in_conv permutations_of_multiset_not_empty) + then show ?thesis using alt_ordering_sysI \Vs \ permutations_of_set \\ by blast +qed + +lemma inc_sys_orderedI: + assumes "incidence_system V B" and "distinct Vs" and "set Vs = V" and "mset Bs = B" + shows "ordered_incidence_system Vs Bs" +proof - + interpret inc: incidence_system V B using assms by simp + show ?thesis proof (unfold_locales) + show "\b. b \# mset Bs \ b \ set Vs" using inc.wellformed assms by simp + show "distinct Vs" using assms(2)permutations_of_setD(2) by auto + qed +qed + +text \Generalise the idea of an incidence matrix to an unordered context \ + +definition is_incidence_matrix :: "'c :: {ring_1} mat \ 'a set \ 'a set multiset \ bool" where +"is_incidence_matrix N V B \ + (\ Vs Bs . (Vs \ permutations_of_set V \ Bs \ permutations_of_multiset B \ N = (inc_mat_of Vs Bs)))" + +lemma (in incidence_system) is_incidence_mat_alt: "is_incidence_matrix N \ \ \ + (\ Vs Bs. (set Vs = \ \ mset Bs = \ \ ordered_incidence_system Vs Bs \ N = (inc_mat_of Vs Bs)))" +proof (intro iffI, simp add: is_incidence_matrix_def) + assume "\Vs. Vs \ permutations_of_set \ \ (\Bs. Bs \ permutations_of_multiset \ \ N = inc_mat_of Vs Bs)" + then obtain Vs Bs where "Vs \ permutations_of_set \ \ Bs \ permutations_of_multiset \ \ N = inc_mat_of Vs Bs" + by auto + then show "\Vs. set Vs = \ \ (\Bs. mset Bs = \ \ ordered_incidence_system Vs Bs \ N = inc_mat_of Vs Bs)" + using incidence_system.alt_ordering_sysI incidence_system_axioms permutations_of_multisetD permutations_of_setD(1) + by blast +next + assume "\Vs Bs. set Vs = \ \ mset Bs = \ \ ordered_incidence_system Vs Bs \ N = inc_mat_of Vs Bs" + then obtain Vs Bs where s: "set Vs = \" and ms: "mset Bs = \" and "ordered_incidence_system Vs Bs" + and n: "N = inc_mat_of Vs Bs" by auto + then interpret ois: ordered_incidence_system Vs Bs by simp + have vs: "Vs \ permutations_of_set \" + using ois.points_indexing s by blast + have "Bs \ permutations_of_multiset \" using ois.blocks_indexing ms by blast + then show "is_incidence_matrix N \ \ " using n vs + using is_incidence_matrix_def by blast +qed + +lemma (in ordered_incidence_system) is_incidence_mat_true: "is_incidence_matrix N \ \ = True" + using blocks_indexing is_incidence_matrix_def points_indexing by blast + +subsection\Incidence Matrices on Design Subtypes \ + +locale ordered_design = ordered_incidence_system \s \s + design "set \s" "mset \s" + for \s and \s +begin + +lemma incidence_mat_non_empty_blocks: + assumes "j < \" + shows "1 \$ (col N j)" +proof - + obtain bl where isbl: "\s ! j = bl" by simp + then have "bl \# \" + using assms valid_blocks_index by auto + then obtain x where inbl: "x \ bl" + using blocks_nempty by blast + then obtain i where isx: "\s ! i = x" and vali: "i < \" + using \bl \# \\ valid_points_index_cons wf_invalid_point by blast + then have "N $$ (i, j) = 1" + using \\s ! j = bl\ \x \ bl\ assms matrix_point_in_block_one by blast + thus ?thesis using vec_setI + by (smt (verit, ccfv_SIG) N_col_def isx vali isbl inbl assms dim_vec_col_N of_nat_less_imp_less) +qed + +lemma all_cols_non_empty: "j < dim_col N \ non_empty_col N j" + using blocks_nempty non_empty_col_map_conv dim_col_is_b by simp +end + +locale ordered_simple_design = ordered_design \s \s + simple_design "(set \s)" "mset \s" for \s \s +begin + +lemma block_list_distinct: "distinct \s" + using block_mset_distinct by auto + +lemma distinct_cols_N: "distinct (cols N)" +proof - + have "inj_on (\ bl . inc_vec_of \s bl) (set \s)" using inc_vec_eq_iff_blocks + by (simp add: inc_vec_eq_iff_blocks inj_on_def) + then show ?thesis using distinct_map inc_mat_of_cols_inc_vecs block_list_distinct + by (simp add: distinct_map inc_mat_of_cols_inc_vecs ) +qed + +lemma simp_blocks_length_card: "length \s = card (set \s)" + using design_support_def simple_block_size_eq_card by fastforce + +lemma blocks_index_inj_on: "inj_on (\ i . \s ! i) {0..s}" + by (auto simp add: inj_on_def) (metis simp_blocks_length_card card_distinct nth_eq_iff_index_eq) + +lemma x_in_block_set_img: assumes "x \ set \s" shows "x \ (!) \s ` {0..s}" +proof - + obtain i where "\s ! i = x" and "i < length \s" using assms + by (meson in_set_conv_nth) + thus ?thesis by auto +qed + +lemma blocks_index_simp_bij_betw: "bij_betw (\ i . \s ! i) {0..s} (set \s)" + using blocks_index_inj_on x_in_block_set_img by (auto simp add: bij_betw_def) + +lemma blocks_index_simp_unique: "i1 < length \s \ i2 < length \s \ i1 \ i2 \ \s ! i1 \ \s ! i2" + using block_list_distinct nth_eq_iff_index_eq by blast + +lemma lift_01_distinct_cols_N: "distinct (cols (lift_01_mat N))" + using lift_01_mat_distinct_cols distinct_cols_N by simp + +end + +locale ordered_proper_design = ordered_design \s \s + proper_design "set \s" "mset \s" + for \s and \s +begin + +lemma mat_is_proper: "proper_inc_mat N" + using design_blocks_nempty v_non_zero + by (auto simp add: proper_inc_mat_def) + +end + +locale ordered_constant_rep = ordered_proper_design \s \s + constant_rep_design "set \s" "mset \s" \ + for \s and \s and \ + +begin + +lemma incidence_mat_rep_num: "i < \ \ mat_rep_num N i = \" + using mat_rep_num_N_row rep_number valid_points_index by simp + +lemma incidence_mat_rep_num_sum: "i < \ \ sum_vec (row N i) = \" + using incidence_mat_rep_num mat_rep_num_N_row + by (simp add: point_rep_mat_row_sum) + +lemma transpose_N_mult_diag: + assumes "i = j" and "i < \" and "j < \" + shows "(N * N\<^sup>T) $$ (i, j) = \" +proof - + have unsq: "\ k . k < \ \ (N $$ (i, k))^2 = N $$ (i, k)" + using assms(2) matrix_elems_one_zero by fastforce + then have "(N * N\<^sup>T) $$ (i, j) = (\k \{0..<\} . N $$ (i, k) * N $$ (j, k))" + using assms(2) assms(3) transpose_mat_mult_entries[of "i" "N" "j"] by (simp) + also have "... = (\k \{0..<\} . (N $$ (i, k))^2)" using assms(1) + by (simp add: power2_eq_square) + also have "... = (\k \{0..<\} . N $$ (i, k))" + by (meson atLeastLessThan_iff sum.cong unsq) + also have "... = (\k \{0..<\} . (row N i) $ k)" + using assms(2) dim_col_is_b dim_row_is_v by auto + finally have "(N * N\<^sup>T) $$ (i, j) = sum_vec (row N i)" + by (simp add: sum_vec_def) + thus ?thesis using incidence_mat_rep_num_sum + using assms(2) by presburger +qed + +end + +locale ordered_block_design = ordered_proper_design \s \s + block_design "set \s" "mset \s" \ + for \s and \s and \ + +begin + +(* Every col has k ones *) +lemma incidence_mat_block_size: "j < \ \ mat_block_size N j = \" + using mat_block_size_N_col uniform valid_blocks_index by fastforce + +lemma incidence_mat_block_size_sum: "j < \ \ sum_vec (col N j) = \" + using incidence_mat_block_size block_size_mat_rep_sum by presburger + +lemma ones_mult_incidence_mat_k_index: "j < \ \ ((u\<^sub>v \) \<^sub>v* N) $ j = \" + using ones_incidence_mat_block_size uniform incidence_mat_block_size by blast + +lemma ones_mult_incidence_mat_k: "((u\<^sub>v \) \<^sub>v* N) = \ \\<^sub>v (u\<^sub>v \)" + using ones_mult_incidence_mat_k_index dim_col_is_b by (intro eq_vecI) (simp_all) + +end + +locale ordered_incomplete_design = ordered_block_design \s \s \ + incomplete_design \ \ \ + for \s and \s and \ + +begin + +lemma incidence_mat_incomplete: "j < \ \ 0 \$ (col N j)" + using valid_blocks_index incomplete_block_col incomplete_imp_incomp_block by blast + +end + +locale ordered_t_wise_balance = ordered_proper_design \s \s + t_wise_balance "set \s" "mset \s" \ \\<^sub>t + for \s and \s and \ and \\<^sub>t + +begin + +lemma incidence_mat_des_index: + assumes "I \ {0..<\}" + assumes "card I = \" + shows "mat_point_index N I = \\<^sub>t" +proof - + have card: "card ((!) \s ` I) = \" using assms points_indexing_inj + by (metis (mono_tags, lifting) card_image ex_nat_less_eq not_le points_list_length subset_iff) + have "((!) \s ` I) \ \" using assms + by (metis atLeastLessThan_iff image_subset_iff subsetD valid_points_index) + then have "\ index ((!) \s ` I) = \\<^sub>t" using balanced assms(2) card by simp + thus ?thesis using mat_point_index_rep assms(1) lessThan_atLeast0 by presburger +qed + +end + +locale ordered_pairwise_balance = ordered_t_wise_balance \s \s 2 \ + pairwise_balance "set \s" "mset \s" \ + for \s and \s and \ +begin + +lemma incidence_mat_des_two_index: + assumes "i1 < \" + assumes "i2 < \" + assumes "i1 \ i2" + shows "mat_point_index N {i1, i2} = \" + using incidence_mat_des_index incidence_mat_two_index +proof - + have "\s ! i1 \ \s ! i2" using assms(3) + by (simp add: assms(1) assms(2) distinct nth_eq_iff_index_eq points_list_length) + then have pair: "card {\s ! i1, \s ! i2} = 2" using card_2_iff by blast + have "{\s ! i1, \s ! i2} \ \" using assms + by (simp add: valid_points_index) + then have "\ index {\s ! i1, \s ! i2} = \" using pair + using balanced by blast + thus ?thesis using incidence_mat_two_index assms by simp +qed + +lemma transpose_N_mult_off_diag: + assumes "i \ j" and "i < \" and "j < \" + shows "(N * N\<^sup>T) $$ (i, j) = \" +proof - + have rev: "\ k. k \ {0..<\} \ \ (N $$ (i, k) = 1 \ N $$ (j, k) = 1) \ N $$ (i, k) = 0 \ N $$ (j, k) = 0" + using assms matrix_elems_one_zero by auto + then have split: "{0..<\} = {k \ {0..<\}. N $$ (i, k) = 1 \ N $$ (j, k) = 1} \ + {k \ {0..<\}. N $$ (i, k) = 0 \ N $$ (j, k) = 0}" + by blast + have zero: "\ k . k \ {0..<\} \ N $$ (i, k) = 0 \ N $$ (j, k) = 0 \ N $$ (i, k) * N$$ (j, k) = 0" + by simp + have djnt: "{k \ {0..<\}. N $$ (i, k) = 1 \ N $$ (j, k) = 1} \ + {k \ {0..<\}. N $$ (i, k) = 0 \ N $$ (j, k) = 0} = {}" using rev by auto + have fin1: "finite {k \ {0..<\}. N $$ (i, k) = 1 \ N $$ (j, k) = 1}" by simp + have fin2: "finite {k \ {0..<\}. N $$ (i, k) = 0 \ N $$ (j, k) = 0}" by simp + have "(N * N\<^sup>T) $$ (i, j) = (\k \{0..<\} . N $$ (i, k) * N $$ (j, k))" + using assms(2) assms(3) transpose_mat_mult_entries[of "i" "N" "j"] by (simp) + also have "... = (\k \({k' \ {0..<\}. N $$ (i, k') = 1 \ N $$ (j, k') = 1} \ + {k' \ {0..<\}. N $$ (i, k') = 0 \ N $$ (j, k') = 0}) . N $$ (i, k) * N $$ (j, k))" + using split by metis + also have "... = (\k \{k' \ {0..<\}. N $$ (i, k') = 1 \ N $$ (j, k') = 1} . N $$ (i, k) * N $$ (j, k)) + + (\k \{k' \ {0..<\}. N $$ (i, k') = 0 \ N $$ (j, k') = 0} . N $$ (i, k) * N $$ (j, k))" + using fin1 fin2 djnt sum.union_disjoint by blast + also have "... = card {k' \ {0..<\}. N $$ (i, k') = 1 \ N $$ (j, k') = 1}" + by (simp add: zero) + also have "... = mat_point_index N {i, j}" + using assms mat_point_index_two_alt[of i N j] by simp + finally show ?thesis using incidence_mat_des_two_index assms by simp +qed + +end + +context pairwise_balance +begin + +lemma ordered_pbdI: + assumes "\ = mset \s" and "\ = set \s" and "distinct \s" + shows "ordered_pairwise_balance \s \s \" +proof - + interpret ois: ordered_incidence_system \s \s + using ordered_incidence_sysII assms finite_incidence_system_axioms by blast + show ?thesis using b_non_zero blocks_nempty assms t_lt_order balanced + by (unfold_locales)(simp_all) +qed +end + +locale ordered_regular_pairwise_balance = ordered_pairwise_balance "\s" "\s" \ + + regular_pairwise_balance "set \s" "mset \s" \ \ for \s and \s and \ and \ + +sublocale ordered_regular_pairwise_balance \ ordered_constant_rep + by unfold_locales + +context ordered_regular_pairwise_balance +begin + +text \ Stinson's Theorem 1.15. Stinson \cite{stinsonCombinatorialDesignsConstructions2004} +gives an iff condition for incidence matrices of regular pairwise +balanced designs. The other direction is proven in the @{term "zero_one_matrix"} context \ +lemma rpbd_incidence_matrix_cond: "N * (N\<^sup>T) = \ \\<^sub>m (J\<^sub>m \) + (\ - \) \\<^sub>m (1\<^sub>m \)" +proof (intro eq_matI) + fix i j + assume ilt: "i < dim_row (int \ \\<^sub>m J\<^sub>m \ + int (\ - \) \\<^sub>m 1\<^sub>m \)" + and jlt: "j < dim_col (int \ \\<^sub>m J\<^sub>m \ + int (\ - \) \\<^sub>m 1\<^sub>m \)" + then have "(int \ \\<^sub>m J\<^sub>m \ + int (\ - \) \\<^sub>m 1\<^sub>m \) $$ (i, j) = + (int \ \\<^sub>m J\<^sub>m \) $$(i, j) + (int (\ - \) \\<^sub>m 1\<^sub>m \) $$ (i, j)" + by simp + then have split: "(int \ \\<^sub>m J\<^sub>m \ + int (\ - \) \\<^sub>m 1\<^sub>m \) $$ (i, j) = + (int \ \\<^sub>m J\<^sub>m \) $$(i, j) + (\ - \) * ((1\<^sub>m \) $$ (i, j))" + using ilt jlt by simp + have lhs: "(int \ \\<^sub>m J\<^sub>m \) $$(i, j) = \" using ilt jlt by simp + show "(N * N\<^sup>T) $$ (i, j) = (int \ \\<^sub>m J\<^sub>m \ + int (\ - \) \\<^sub>m 1\<^sub>m \) $$ (i, j)" + proof (cases "i = j") + case True + then have rhs: "(int (\ - \) \\<^sub>m 1\<^sub>m \) $$ (i, j) = (\ - \)" using ilt by fastforce + have "(int \ \\<^sub>m J\<^sub>m \ + int (\ - \) \\<^sub>m 1\<^sub>m \) $$ (i, j) = \ + (\ - \)" + using True jlt by auto + then have "(int \ \\<^sub>m J\<^sub>m \ + int (\ - \) \\<^sub>m 1\<^sub>m \) $$ (i, j) = \" + using reg_index_lt_rep by (simp add: nat_diff_split) + then show ?thesis using lhs split rhs True transpose_N_mult_diag ilt jlt by simp + next + case False + then have "(1\<^sub>m \) $$ (i, j) = 0" using ilt jlt by simp + then have "(\ - \) * ((1\<^sub>m \) $$ (i, j)) = 0" using ilt jlt + by (simp add: \1\<^sub>m \ $$ (i, j) = 0\) + then show ?thesis using lhs transpose_N_mult_off_diag ilt jlt False by simp + qed +next + show "dim_row (N * N\<^sup>T) = dim_row (int \ \\<^sub>m J\<^sub>m \ + int (\ - \) \\<^sub>m 1\<^sub>m \)" + using transpose_N_mult_dim(1) by auto +next + show "dim_col (N * N\<^sup>T) = dim_col (int \ \\<^sub>m J\<^sub>m \ + int (\ - \) \\<^sub>m 1\<^sub>m \)" + using transpose_N_mult_dim(1) by auto +qed +end + +locale ordered_bibd = ordered_proper_design \s \s + bibd "set \s" "mset \s" \ \ + for \s and \s and \ and \ + +sublocale ordered_bibd \ ordered_incomplete_design + by unfold_locales + +sublocale ordered_bibd \ ordered_constant_rep \s \s \ + by unfold_locales + +sublocale ordered_bibd \ ordered_pairwise_balance + by unfold_locales + +locale ordered_sym_bibd = ordered_bibd \s \s \ \ + symmetric_bibd "set \s" "mset \s" \ \ + for \s and \s and \ and \ + + +sublocale ordered_sym_bibd \ ordered_simple_design + by (unfold_locales) + +locale ordered_const_intersect_design = ordered_proper_design \s \s + const_intersect_design "set \s" "mset \s" \ + for \s \s \ + + +locale simp_ordered_const_intersect_design = ordered_const_intersect_design + ordered_simple_design +begin + +lemma max_one_block_size_inter: + assumes "\ \ 2" + assumes "bl \# \" + assumes "card bl = \" + assumes "bl2 \# \ - {#bl#}" + shows "\ < card bl2" +proof - + have sd: "simple_design \ \" + by (simp add: simple_design_axioms) + have bl2in: "bl2 \# \" using assms(4) + by (meson in_diffD) + have blin: "bl \# {#b \# \ . card b = \#}" using assms(3) assms(2) by simp + then have slt: "size {#b \# \ . card b = \#} = 1" using simple_const_inter_iff sd assms(1) + by (metis count_empty count_eq_zero_iff less_one nat_less_le size_eq_0_iff_empty) + then have "size {#b \# (\ - {#bl#}) . card b = \#} = 0" using blin + by (smt (verit) add_mset_eq_singleton_iff count_eq_zero_iff count_filter_mset + filter_mset_add_mset insert_DiffM size_1_singleton_mset size_eq_0_iff_empty) + then have ne: "card bl2 \ \" using assms(4) + by (metis (mono_tags, lifting) filter_mset_empty_conv size_eq_0_iff_empty) + thus ?thesis using inter_num_le_block_size assms bl2in nat_less_le by presburger +qed + +lemma block_size_inter_num_cases: + assumes "bl \# \" + assumes "\ \ 2" + shows "\ < card bl \ (card bl = \ \ (\ bl' \# (\ - {#bl#}) . \ < card bl'))" +proof (cases "card bl = \") + case True + have "(\ bl'. bl' \# (\ - {#bl#}) \ \ < card bl')" + using max_one_block_size_inter True assms by simp + then show ?thesis using True by simp +next + case False + then have "\ < card bl" using assms inter_num_le_block_size nat_less_le by presburger + then show ?thesis by simp +qed + +lemma indexed_const_intersect: + assumes "j1 < \" + assumes "j2 < \" + assumes "j1 \ j2" + shows "(\s ! j1) |\| (\s ! j2) = \" +proof - + obtain bl1 bl2 where "bl1 \# \" and "\s ! j1 = bl1" and "bl2 \# \ - {#bl1#}" and "\s ! j2 = bl2" + using obtains_two_diff_block_indexes assms by fastforce + thus ?thesis by (simp add: const_intersect) +qed + +lemma const_intersect_block_size_diff: + assumes "j' < \" and "j < \" and "j \ j'" and "card (\s ! j') = \" and "\ \ 2" + shows "card (\s ! j) - \ > 0" +proof - + obtain bl1 bl2 where "bl1 \# \" and "\s ! j' = bl1" and "bl2 \# \ - {#bl1#}" and "\s ! j = bl2" + using assms(1) assms(2) assms(3) obtains_two_diff_block_indexes by fastforce + then have "\ < card (bl2)" + using max_one_block_size_inter assms(4) assms(5) by blast + thus ?thesis + by (simp add: \\s ! j = bl2\) +qed + +lemma scalar_prod_inc_vec_const_inter: + assumes "j1 < \" "j2 < \" "j1 \ j2" + shows "(col N j1) \ (col N j2) = \" + using scalar_prod_inc_vec_inter_num indexed_const_intersect assms by simp + +end + +subsection \ Zero One Matrix Incidence System Existence \ +text \We prove 0-1 matrices with certain properties imply the existence of an incidence system +with particular properties. This leads to Stinson's theorem in the other direction \cite{stinsonCombinatorialDesignsConstructions2004} \ + +context zero_one_matrix +begin + +lemma mat_is_ordered_incidence_sys: "ordered_incidence_system [0..<(dim_row M)] (map (map_col_to_block) (cols M))" + apply (unfold_locales, simp_all) + using map_col_to_block_wf atLeastLessThan_upt by blast + +interpretation mat_ord_inc_sys: ordered_incidence_system "[0..<(dim_row M)]" "(map (map_col_to_block) (cols M))" + by (simp add: mat_is_ordered_incidence_sys) + +lemma mat_ord_inc_sys_N: "mat_ord_inc_sys.N = lift_01_mat M" + by (intro eq_matI, simp_all add: inc_mat_of_def map_col_to_block_elem) + (metis lift_01_mat_simp(3) lift_mat_01_index_iff(2) of_zero_neq_one_def) + +lemma map_col_to_block_mat_rep_num: + assumes "x # mset (cols M)#} rep x) = mat_rep_num M x" +proof - + have "mat_rep_num M x = mat_rep_num (lift_01_mat M) x" + using preserve_mat_rep_num mat_ord_inc_sys_N + by (metis assms lift_01_mat_def of_inj_on_01_hom.inj_on_01_hom_axioms) + then have "mat_rep_num M x = (mat_rep_num mat_ord_inc_sys.N x)" using mat_ord_inc_sys_N by (simp) + then have "mat_rep_num M x = mset (map (map_col_to_block) (cols M)) rep x" + using assms atLeastLessThan_upt card_atLeastLessThan mat_ord_inc_sys.mat_rep_num_N_row + mat_ord_inc_sys_point minus_nat.diff_0 by presburger + thus ?thesis using ordered_to_mset_col_blocks + by presburger +qed + +end + +context zero_one_matrix_ring_1 +begin + +lemma transpose_cond_index_vals: + assumes "M * (M\<^sup>T) = \ \\<^sub>m (J\<^sub>m (dim_row M)) + (r - \) \\<^sub>m (1\<^sub>m (dim_row M))" + assumes "i < dim_row (M * (M\<^sup>T))" + assumes "j < dim_col (M * (M\<^sup>T))" + shows "i = j \ (M * (M\<^sup>T)) $$ (i, j) = r" "i \ j \ (M * (M\<^sup>T)) $$ (i, j) = \" + using assms by auto + +end + +locale zero_one_matrix_int = zero_one_matrix_ring_1 M for M :: "int mat" +begin + +text \Some useful conditions on the transpose product for matrix system properties \ +lemma transpose_cond_diag_r: + assumes "i < dim_row (M * (M\<^sup>T))" + assumes "\ j. i = j \ (M * (M\<^sup>T)) $$ (i, j) = r" + shows "mat_rep_num M i = r" +proof - + have eqr: "(M * M\<^sup>T) $$ (i, i) = r" using assms(2) + by simp + have unsq: "\ k . k < dim_col M \ (M $$ (i, k))^2 = M $$ (i, k)" + using assms elems01 by fastforce + have "sum_vec (row M i) = (\k \{0..<(dim_col M)} . (row M i) $ k)" + using assms by (simp add: sum_vec_def) + also have "... = (\k \{0..<(dim_col M)} . M $$ (i, k))" + using assms by auto + also have "... = (\k \{0..<(dim_col M)} . M $$ (i, k)^2)" + using atLeastLessThan_iff sum.cong unsq by simp + also have "... = (\k \{0..<(dim_col M)} . M $$ (i, k) * M $$ (i, k))" + using assms by (simp add: power2_eq_square) + also have "... = (M * M\<^sup>T) $$ (i, i)" + using assms transpose_mat_mult_entries[of "i" "M" "i"] by simp + finally have "sum_vec (row M i) = r" using eqr by simp + thus ?thesis using mat_rep_num_sum_alt + by (metis assms(1) elems01 index_mult_mat(2) of_nat_eq_iff) +qed + + +lemma transpose_cond_non_diag: + assumes "i1 < dim_row (M * (M\<^sup>T))" + assumes "i2 < dim_row (M * (M\<^sup>T))" + assumes "i1 \ i2" + assumes "\ j i. j \ i \ i < dim_row (M * (M\<^sup>T)) \ j < dim_row (M * (M\<^sup>T)) \ (M * (M\<^sup>T)) $$ (i, j) = \" + shows "\ = mat_point_index M {i1, i2}" +proof - + have ilt: "i1 < dim_row M" "i2 < dim_row M" + using assms(1) assms (2) by auto + have rev: "\ k. k \ {0.. + \ (M $$ (i1, k) = 1 \ M $$ (i2, k) = 1) \ M $$ (i1, k) = 0 \ M $$ (i2, k) = 0" + using assms elems01 by fastforce + then have split: "{0.. {0.. M $$ (i2, k) = 1} \ + {k \ {0.. M $$ (i2, k) = 0}" + by blast + have zero: "\ k . k \ {0.. M $$ (i1, k) = 0 \ M $$ (i2, k) = 0 \ M $$ (i1, k) * M$$ (i2, k) = 0" + by simp + have djnt: "{k \ {0.. M $$ (i2, k) = 1} \ + {k \ {0.. M $$ (i2, k) = 0} = {}" + using rev by auto + have fin1: "finite {k \ {0.. M $$ (i2, k) = 1}" by simp + have fin2: "finite {k \ {0.. M $$ (i2, k) = 0}" by simp + have "mat_point_index M {i1, i2} = card {k' \ {0..M $$ (i2, k') = 1}" + using mat_point_index_two_alt ilt assms(3) by auto + then have "mat_point_index M {i1, i2} = + (\k \{k' \ {0.. M $$ (i2, k') = 1} . M $$ (i1, k) * M $$ (i2, k)) + + (\k \{k' \ {0.. M $$ (i2, k') = 0} . M $$ (i1, k) * M $$ (i2, k))" + by (simp add: zero) (* Odd behaviour if I use also have here *) + also have "... = (\k \({k' \ {0.. M $$ (i2, k') = 1} \ + {k' \ {0.. M $$ (i2, k') = 0}) . M $$ (i1, k) * M $$ (i2, k))" + using fin1 fin2 djnt sum.union_disjoint by (metis (no_types, lifting)) + also have "... = (\k \{0..T)) $$ (i1, i2)" + using assms(1) assms(2) transpose_mat_mult_entries[of "i1" "M" "i2"] by simp + thus ?thesis using assms by presburger +qed + +lemma trans_cond_implies_map_rep_num: + assumes "M * (M\<^sup>T) = \ \\<^sub>m (J\<^sub>m (dim_row M)) + (r - \) \\<^sub>m (1\<^sub>m (dim_row M))" + assumes "x < dim_row M" + shows "(image_mset map_col_to_block (mset (cols M))) rep x = r" +proof - + interpret ois: ordered_incidence_system "[0.. rep x = sum_vec (row M x)" using ois.point_rep_mat_row_sum + by (simp add: assms(2) inc_mat_of_map_rev) + then have "\ j. x = j \ (M * (M\<^sup>T)) $$ (x, j) = r" using assms(1) transpose_cond_index_vals + by (metis assms(2) index_mult_mat(2) index_mult_mat(3) index_transpose_mat(3)) + thus ?thesis using eq transpose_cond_diag_r assms(2) index_mult_mat(2) + by (metis map_col_to_block_mat_rep_num) +qed + +lemma trans_cond_implies_map_index: + assumes "M * (M\<^sup>T) = \ \\<^sub>m (J\<^sub>m (dim_row M)) + (r - \) \\<^sub>m (1\<^sub>m (dim_row M))" + assumes "ps \ {0.." +proof - + interpret ois: ordered_incidence_system "[0.. i2" + using assms(2) assms(3) card_2_iff insert_subset by (metis atLeastLessThan_iff) + have cond: "\ j i. j \ i \ i < dim_row (M * (M\<^sup>T)) \ j < dim_row (M * (M\<^sup>T)) \ (M * (M\<^sup>T)) $$ (i, j) = \" + using assms(1) by simp + then have "(image_mset map_col_to_block (mset (cols M))) index ps = mat_point_index M ps" + using ois.incidence_mat_two_index psis i1in i2in by (simp add: neqi inc_mat_of_map_rev) + thus ?thesis using cond transpose_cond_non_diag[of i1 i2 \] i1in i2in index_mult_mat(2)[of "M" "M\<^sup>T"] + neqi of_nat_eq_iff psis by simp +qed + +text \ Stinson Theorem 1.15 existence direction \ +lemma rpbd_exists: + assumes "dim_row M \ 2" \ \Min two points\ + assumes "dim_col M \ 1" \ \Min one block\ + assumes "\ j. j < dim_col M \ 1 \$ col M j" \ \no empty blocks \ + assumes "M * (M\<^sup>T) = \ \\<^sub>m (J\<^sub>m (dim_row M)) + (r - \) \\<^sub>m (1\<^sub>m (dim_row M))" + shows "ordered_regular_pairwise_balance [0.. r" +proof - + interpret ois: ordered_incidence_system "[0..v (dim_row M)) \<^sub>v* M) = k \\<^sub>v (u\<^sub>v (dim_col M))" + assumes "j < dim_col M" + shows "mat_block_size M j = k" +proof - + have "mat_block_size M j = sum_vec (col M j)" using assms(2) + by (simp add: elems01 mat_block_size_sum_alt) + also have "... = ((u\<^sub>v (dim_row M)) \<^sub>v* M) $ j" using assms(2) + by (simp add: sum_vec_def scalar_prod_def) + finally show ?thesis using assms(1) assms(2) by (simp) +qed + +lemma vec_k_impl_uniform_block_size: + assumes "((u\<^sub>v (dim_row M)) \<^sub>v* M) = k \\<^sub>v (u\<^sub>v (dim_col M))" + assumes "bl \# (image_mset map_col_to_block (mset (cols M)))" + shows "card bl = k" +proof - + obtain j where jlt: "j < dim_col M" and bleq: "bl = map_col_to_block (col M j)" + using assms(2) obtain_block_index_map_block_set by blast + then have "card (map_col_to_block (col M j)) = mat_block_size M j" + by (simp add: map_col_to_block_size) + thus ?thesis using vec_k_uniform_mat_block_size assms(1) bleq jlt by blast +qed + +lemma bibd_exists: + assumes "dim_col M \ 1" \ \Min one block\ + assumes "\ j. j < dim_col M \ 1 \$ col M j" \ \no empty blocks \ + assumes "M * (M\<^sup>T) = \ \\<^sub>m (J\<^sub>m (dim_row M)) + (r - \) \\<^sub>m (1\<^sub>m (dim_row M))" + assumes "((u\<^sub>v (dim_row M)) \<^sub>v* M) = k \\<^sub>v (u\<^sub>v (dim_col M))" + assumes "(r ::nat) \ 0" + assumes "k \ 2" "k < dim_row M" + shows "ordered_bibd [0.." +proof - + interpret ipbd: ordered_regular_pairwise_balance "[0.. r + using rpbd_exists assms by simp + show ?thesis using vec_k_impl_uniform_block_size by (unfold_locales, simp_all add: assms) +qed + +end + +subsection \Isomorphisms and Incidence Matrices \ +text \If two incidence systems have the same incidence matrix, they are isomorphic. Similarly +if two incidence systems are isomorphic there exists an ordering such that they have the same +incidence matrix \ +locale two_ordered_sys = D1: ordered_incidence_system \s \s + D2: ordered_incidence_system \s' \s' + for "\s" and "\s" and "\s'" and "\s'" + +begin + +lemma equal_inc_mat_isomorphism: + assumes "D1.N = D2.N" + shows "incidence_system_isomorphism D1.\ D1.\ D2.\ D2.\ (\ x . \s' ! (List_Index.index \s x))" +proof (unfold_locales) + show "bij_betw (\x. \s' ! List_Index.index \s x) D1.\ D2.\" + proof - + have comp: "(\x. \s' ! List_Index.index \s x) = (\ i. \s' ! i) \ (\ y . List_Index.index \s y)" + by (simp add: comp_def) + have leq: "length \s = length \s'" + using assms D1.dim_row_is_v D1.points_list_length D2.dim_row_is_v D2.points_list_length by force + have bij1: "bij_betw (\ i. \s' !i) {..s} (set \s') " using leq + by (simp add: bij_betw_nth D2.distinct) + have "bij_betw (List_Index.index \s) (set \s) {..s}" using D1.distinct + by (simp add: bij_betw_index lessThan_atLeast0) + thus ?thesis using bij_betw_trans comp bij1 by simp + qed +next + have len: "length (map ((`) (\x. \s' ! List_Index.index \s x)) \s) = length \s'" + using length_map assms D1.dim_col_is_b by force + have mat_eq: "\ i j . D1.N $$ (i, j) = D2.N $$ (i, j)" using assms + by simp + have vslen: "length \s = length \s'" using assms + using D1.dim_row_is_v D1.points_list_length D2.dim_row_is_v D2.points_list_length by force + have "\ j. j < length \s' \ (map ((`) (\x. \s' ! List_Index.index \s x)) \s) ! j = \s' ! j" + proof - + fix j assume a: "j < length \s'" + then have "(map ((`) (\x. \s' ! List_Index.index \s x)) \s) ! j = (\x. \s' ! List_Index.index \s x) ` (\s ! j)" + by (metis D1.blocks_list_length D1.dim_col_is_b D2.blocks_list_length D2.dim_col_is_b assms nth_map) + also have "... = (\ i . \s' ! i) ` ((\ x. List_Index.index \s x) ` (\s ! j))" + by blast + also have "... = ((\ i . \s' ! i) ` {i . i < length \s \ D1.N $$ (i, j) = 1})" + using D1.block_mat_cond_rev a assms + by (metis (no_types, lifting) D1.blocks_list_length D1.dim_col_is_b D2.blocks_list_length D2.dim_col_is_b) + also have "... = ((\ i . \s' ! i) ` {i . i < length \s' \ D2.N $$ (i, j) = 1})" + using vslen mat_eq by simp + finally have "(map ((`) (\x. \s' ! List_Index.index \s x)) \s) ! j = (\s' ! j)" + using D2.block_mat_cond_rep' a by presburger + then show "(map ((`) (\x. \s' ! List_Index.index \s x)) \s) ! j = (\s' ! j)" by simp + qed + then have "map ((`) (\x. \s' ! List_Index.index \s x)) \s = \s'" + using len nth_equalityI[of "(map ((`) (\x. \s' ! List_Index.index \s x)) \s)" "\s'"] by simp + then show "image_mset ((`) (\x. \s' ! List_Index.index \s x)) D1.\ = D2.\" + using mset_map by auto +qed + +lemma equal_inc_mat_isomorphism_ex: "D1.N = D2.N \ \ \ . incidence_system_isomorphism D1.\ D1.\ D2.\ D2.\ \" + using equal_inc_mat_isomorphism by auto + +lemma equal_inc_mat_isomorphism_obtain: + assumes "D1.N = D2.N" + obtains \ where "incidence_system_isomorphism D1.\ D1.\ D2.\ D2.\ \" + using equal_inc_mat_isomorphism assms by auto + +end + +context incidence_system_isomorphism +begin + +lemma exists_eq_inc_mats: + assumes "finite \" "finite \'" + obtains N where "is_incidence_matrix N \ \" and "is_incidence_matrix N \' \'" +proof - + obtain Vs where vsis: "Vs \ permutations_of_set \" using assms + by (meson all_not_in_conv permutations_of_set_empty_iff) + obtain Bs where bsis: "Bs \ permutations_of_multiset \" + by (meson all_not_in_conv permutations_of_multiset_not_empty) + have inj: "inj_on \ \" using bij + by (simp add: bij_betw_imp_inj_on) + then have mapvs: "map \ Vs \ permutations_of_set \'" using permutations_of_set_image_inj + using \Vs \ permutations_of_set \\ iso_points_map by blast + have "permutations_of_multiset (image_mset ((`)\) \) = map ((`) \) ` permutations_of_multiset \" + using block_img permutations_of_multiset_image by blast + then have mapbs: "map ((`) \) Bs \ permutations_of_multiset \'" using bsis block_img by blast + define N :: "'c :: {ring_1} mat" where "N \ inc_mat_of Vs Bs" + have "is_incidence_matrix N \ \" + using N_def bsis is_incidence_matrix_def vsis by blast + have "\ bl . bl \ (set Bs) \ bl \ (set Vs)" + by (meson bsis in_multiset_in_set ordered_incidence_system.wf_list source.alt_ordering_sysI vsis) + then have "N = inc_mat_of (map \ Vs) (map ((`) \) Bs)" + using inc_mat_of_bij_betw inj + by (metis N_def permutations_of_setD(1) vsis) + then have "is_incidence_matrix N \' \'" + using mapbs mapvs is_incidence_matrix_def by blast + thus ?thesis + using \is_incidence_matrix N \ \\ that by auto +qed + +end + +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/Linear_Bound_Argument.thy b/thys/Fishers_Inequality/Linear_Bound_Argument.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Linear_Bound_Argument.thy @@ -0,0 +1,161 @@ +(* Title: Linear_Bound_Argument.thy + Author: Chelsea Edmonds +*) +section \Linear Bound Argument - General \ +text \Lemmas to enable general reasoning using the linear bound argument for combinatorial proofs. +Jukna \cite{juknaExtremalCombinatorics2011} presents a good overview of the mathematical background +this theory is based on and applications \ +theory Linear_Bound_Argument imports Incidence_Matrices Jordan_Normal_Form.DL_Rank +Jordan_Normal_Form.Ring_Hom_Matrix +begin + +subsection \Vec Space Extensions\ +text \Simple extensions to the existing vector space locale on linear independence \ +context vec_space +begin +lemma lin_indpt_set_card_lt_dim: + fixes A :: "'a vec set" + assumes "A \ carrier_vec n" + assumes "lin_indpt A" + shows "card A \ dim" + using assms(1) assms(2) fin_dim li_le_dim(2) by blast + +lemma lin_indpt_dim_col_lt_dim: + fixes A :: "'a mat" + assumes "A \ carrier_mat n nc" + assumes "distinct (cols A)" + assumes "lin_indpt (set (cols A))" + shows "nc \ dim" +proof - + have b: "card (set (cols A)) = dim_col A" using cols_length assms(2) + by (simp add: distinct_card) + have "set (cols A) \ carrier_vec n" using assms(1) cols_dim by blast + thus ?thesis using lin_indpt_set_card_lt_dim assms b by auto +qed + +lemma lin_comb_imp_lin_dep_fin: + fixes A :: "'a vec set" + assumes "finite A" + assumes "A \ carrier_vec n" + assumes "lincomb c A = 0\<^sub>v n" + assumes "\ a. a \ A \ c a \ 0" + shows "lin_dep A" + unfolding lin_dep_def using assms lincomb_as_lincomb_list_distinct sumlist_nth by auto + +text \While a trivial definition, this enables us to directly reference the definition outside +of a locale context, as @{term "lin_indpt"} is an inherited definition \ +definition lin_indpt_vs:: "'a vec set \ bool" where +"lin_indpt_vs A \ lin_indpt A" + +lemma lin_comb_sum_lin_indpt: + fixes A :: "'a vec list" + assumes "set (A) \ carrier_vec n" + assumes "distinct A" + assumes "\ f. lincomb_list (\i. f (A ! i)) A = 0\<^sub>v n \ \v\ (set A). f v = 0" + shows "lin_indpt (set A)" + by (rule finite_lin_indpt2, auto simp add: assms lincomb_as_lincomb_list_distinct) + +lemma lin_comb_mat_lin_indpt: + fixes A :: "'a vec list" + assumes "set (A) \ carrier_vec n" + assumes "distinct A" + assumes "\ f. mat_of_cols n A *\<^sub>v vec (length A) (\i. f (A ! i)) = 0\<^sub>v n \ \v\ (set A). f v = 0" + shows "lin_indpt (set A)" +proof (rule lin_comb_sum_lin_indpt, auto simp add: assms) + fix f v + have "\ v. v \ set A \ dim_vec v = n" + using assms by auto + then show "lincomb_list (\i. f (A ! i)) A = 0\<^sub>v n \ v \ set A \ f v = 0" + using lincomb_list_as_mat_mult[of A "(\i. f (A ! i))"] assms(3)[of f] by auto +qed + +lemma lin_comb_mat_lin_indpt_vs: + fixes A :: "'a vec list" + assumes "set (A) \ carrier_vec n" + assumes "distinct A" + assumes "\ f. mat_of_cols n A *\<^sub>v vec (length A) (\i. f (A ! i)) = 0\<^sub>v n \ \v\ (set A). f v = 0" + shows "lin_indpt_vs (set A)" + using lin_comb_mat_lin_indpt lin_indpt_vs_def assms by auto + + +end + +subsection \Linear Bound Argument Lemmas\ + +text \Three general representations of the linear bound argument, requiring a direct fact of +linear independence onthe rows of the vector space over either a set A of vectors, list xs of vectors +or a Matrix' columns \ +lemma lin_bound_arg_general_set: + fixes A ::"('a :: {field})vec set" + assumes "A \ carrier_vec nr" + assumes "vec_space.lin_indpt_vs nr A" + shows "card A \ nr" + using vec_space.lin_indpt_set_card_lt_dim[of "A" "nr"] vec_space.lin_indpt_vs_def[of nr A] + vec_space.dim_is_n assms by metis + +lemma lin_bound_arg_general_list: + fixes xs ::"('a :: {field})vec list" + assumes "distinct xs" + assumes "(set xs) \ carrier_vec nr" + assumes "vec_space.lin_indpt_vs nr (set xs)" + shows "length (xs) \ nr" + using lin_bound_arg_general_set[of "set xs" nr] distinct_card assms + by force + +lemma lin_bound_arg_general: + fixes A ::"('a :: {field}) mat" + assumes "distinct (cols A)" + assumes "A \ carrier_mat nr nc" + assumes "vec_space.lin_indpt_vs nr (set (cols A))" + shows "nc \ nr" +proof - + have ss: "set (cols A) \ carrier_vec nr" using assms cols_dim by blast + have "length (cols A) = nc" + using assms(2) cols_length by blast + thus ?thesis using lin_bound_arg_general_list[of "cols A" "nr"] ss assms by simp +qed + +text\The linear bound argument lemma on a matrix requiring the lower level assumption on a linear +combination. This removes the need to directly refer to any aspect of the linear algebra libraries \ +lemma lin_bound_argument: + fixes A :: "('a :: {field}) mat" + assumes "distinct (cols A)" + assumes "A \ carrier_mat nr nc" + assumes "\ f. A *\<^sub>v vec nc (\i. f (col A i)) = 0\<^sub>v nr \ \v\ (set (cols A)). f v = 0" + shows "nc \ nr" +proof (intro lin_bound_arg_general[of "A" nr nc] vec_space.lin_comb_mat_lin_indpt_vs, simp_all add: assms) + show "set (cols A) \ carrier_vec nr" using assms cols_dim by blast +next + have mA: "mat_of_cols nr (cols A) = A" using mat_of_cols_def assms by auto + have "\ f. vec (dim_col A) (\i. f (cols A ! i)) = vec nc (\i. f (col A i))" + proof (intro eq_vecI, simp_all add: assms) + show "\f i. i < nc \ vec (dim_col A) (\i. f (cols A ! i)) $ i = f (col A i)" + using assms(2) by force + show "dim_col A = nc " using assms by simp + qed + then show "\f. mat_of_cols nr (cols A) *\<^sub>v vec (dim_col A) (\i. f (cols A ! i)) = 0\<^sub>v nr \ + \v\set (cols A). f v = 0" + using mA assms(3) by metis +qed + +text \A further extension to present the linear combination directly as a sum. This manipulation from +vector product to a summation was found to commonly be repeated in proofs applying this rule \ +lemma lin_bound_argument2: + fixes A :: "('a :: {field}) mat" + assumes "distinct (cols A)" + assumes "A \ carrier_mat nr nc" + assumes "\ f. vec nr (\i. \ j \ {0..v nr \ + \v\ (set (cols A)). f v = 0" + shows "nc \ nr" +proof (intro lin_bound_argument[of A nr nc], simp add: assms, simp add: assms) + fix f + have "A *\<^sub>v vec nc (\i. f (col A i)) = + vec (dim_row A) (\i. \ j \ {0..i. \ j \ {0..v vec nc (\i. f (col A i)) = 0\<^sub>v nr \ \v\set (cols A). f v = 0" + using assms(3)[of f] assms(2) by fastforce +qed + +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/Matrix_Vector_Extras.thy b/thys/Fishers_Inequality/Matrix_Vector_Extras.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Matrix_Vector_Extras.thy @@ -0,0 +1,775 @@ +(* Title: Matrix_Vector_Extras.thy + Author: Chelsea Edmonds +*) +section \ Matrix and Vector Additions \ + +theory Matrix_Vector_Extras imports Set_Multiset_Extras Jordan_Normal_Form.Matrix +Design_Theory.Multisets_Extras "Groebner_Bases.Macaulay_Matrix" "Polynomial_Factorization.Missing_List" +begin + + +subsection \Vector Extras\ +text \For ease of use, a number of additions to the existing vector library as initially developed +in the JNF AFP Entry, are given below\ + +text \We define the concept of summing up elements of a vector \ + +definition (in comm_monoid_add) sum_vec :: "'a vec \ 'a" where +"sum_vec v \ sum (\ i . v $ i) {0.. i . v $ i) {0.. i . (vCons a v) $ Suc i) {0..< dim_vec v}" + by force + also have "... = sum (\ i . (vCons a v) $ i) {Suc 0..< (Suc (dim_vec v))}" + by (metis sum.shift_bounds_Suc_ivl) + finally have sum: "sum_vec v = sum (\ i . (vCons a v) $ i) {Suc 0..< dim_vec (vCons a v)}" by simp + have "sum_vec (vCons a v) = sum (\ i . (vCons a v) $ i){0..< dim_vec (vCons a v)}" + by (simp add: sum_vec_def) + then have "sum_vec (vCons a v) = (vCons a v) $ 0 + sum (\ i . (vCons a v) $ i){Suc 0..< dim_vec (vCons a v)}" + by (metis dim_vec_vCons sum.atLeast_Suc_lessThan zero_less_Suc) + thus ?thesis using sum 0 by simp +qed + +lemma sum_vec_list: "sum_list (list_of_vec v) = sum_vec v" + by (induct v)(simp_all add: sum_vec_vCons) + +lemma sum_vec_mset: "sum_vec v = (\ x \# (mset (list_of_vec v)) . x)" + by (simp add: sum_vec_list) + +lemma dim_vec_vCons_ne_0: "dim_vec (vCons a v) > 0" + by (cases v) simp_all + +lemma sum_vec_vCons_lt: + assumes "\ i. i < dim_vec (vCons a v) \ (vCons a v) $ i \ (n ::int)" + assumes "sum_vec v \ m" + shows "sum_vec (vCons a v) \ n + m" +proof - + have split: "sum_vec (vCons a v) = a + sum_vec v" by (simp add: sum_vec_vCons) + have a: "(vCons a v) $ 0 = a" by simp + then have "0 < dim_vec (vCons a v)" using dim_vec_vCons_ne_0 by simp + then have "a \ n" using assms by (metis a) + thus ?thesis using split assms + by (simp add: add_mono) +qed + +lemma sum_vec_one_zero: + assumes "\ i. i < dim_vec (v :: int vec) \ v $ i \ (1 ::int)" + shows "sum_vec v \ dim_vec v" + using assms +proof (induct v) + case vNil + then show ?case by simp +next + case (vCons a v) + then have "(\ i. i < dim_vec v \ v $ i \ (1 ::int))" + using vCons.prems by force + then have lt: "sum_vec v \ int (dim_vec v)" by (simp add: vCons.hyps) + then show ?case using sum_vec_vCons_lt lt vCons.prems by simp +qed + +text \Definition to convert a vector to a multiset \ + +definition vec_mset:: "'a vec \ 'a multiset" where +"vec_mset v \ image_mset (vec_index v) (mset_set {.. i \ {.. x \# vec_mset v" + by (auto simp add: vec_mset_def) + +lemma mset_vec_same_size: "dim_vec v = size (vec_mset v)" + by (simp add: vec_mset_def) + +lemma mset_vec_eq_mset_list: "vec_mset v = mset (list_of_vec v)" + by (auto simp add: vec_mset_def) + (metis list_of_vec_map mset_map mset_set_upto_eq_mset_upto) + +lemma vec_mset_img_map: "image_mset f (mset (xs)) = vec_mset (map_vec f (vec_of_list xs))" + by (metis list_vec mset_map mset_vec_eq_mset_list vec_of_list_map) + +lemma vec_mset_vNil: "vec_mset vNil = {#}" + by (simp add: vec_mset_def) + +lemma vec_mset_vCons: "vec_mset (vCons x v) = add_mset x (vec_mset v)" +proof - + have "vec_mset (vCons x v) = mset (list_of_vec (vCons x v))" + by (simp add: mset_vec_eq_mset_list) + then have "mset (list_of_vec (vCons x v)) = add_mset x (mset (list_of_vec v))" + by simp + thus ?thesis + by (metis mset_vec_eq_mset_list) +qed + +lemma vec_mset_set: "vec_set v = set_mset (vec_mset v)" + by (simp add: mset_vec_eq_mset_list set_list_of_vec) + +lemma vCons_set_contains_in: "a \ set\<^sub>v v \ set\<^sub>v (vCons a v) = set\<^sub>v v" + by (metis remdups_mset_singleton_sum set_mset_remdups_mset vec_mset_set vec_mset_vCons) + +lemma vCons_set_contains_add: "a \ set\<^sub>v v \ set\<^sub>v (vCons a v) = set\<^sub>v v \ {a}" + using vec_mset_set vec_mset_vCons + by (metis Un_insert_right set_mset_add_mset_insert sup_bot_right) + +lemma setv_vec_mset_not_in_iff: "a \ set\<^sub>v v \ a \# vec_mset v" + by (simp add: vec_mset_set) + +text \Abbreviation for counting occurrences of an element in a vector \ + +abbreviation "count_vec v a \ count (vec_mset v) a" + +lemma vec_count_lt_dim: "count_vec v a \ dim_vec v" + by (metis mset_vec_same_size order_refl set_count_size_min) + +lemma count_vec_empty: "dim_vec v = 0 \ count_vec v a = 0" + by (simp add: mset_vec_same_size) + +lemma count_vec_vNil: "count_vec vNil a = 0" + by (simp add: vec_mset_def) + +lemma count_vec_vCons: "count_vec (vCons aa v) a = (if (aa = a) then count_vec v a + 1 else count_vec v a)" + by(simp add: vec_mset_vCons) + +lemma elem_exists_count_min: "\ i \{.. count_vec v x \ 1" + by (simp add: vec_elem_exists_mset) + +lemma count_vec_count_mset: "vec_mset v = image_mset f A \ count_vec v a = count (image_mset f A) a" + by (simp) + +lemma count_vec_alt_list: "count_vec v a = length (filter (\y. a = y) (list_of_vec v))" + by (simp add: mset_vec_eq_mset_list) (metis count_mset) + +lemma count_vec_alt: "count_vec v x = card { i. v $ i = x \ i< dim_vec v}" +proof - + have "count_vec v x = count (image_mset (($) v) (mset_set {..# (image_mset (($) v) (mset_set {..# (mset_set {.. {.. i. i < dim_vec v \ v $ i = 1 \ v $ i = 0" + shows "of_nat (count_vec v 1) = sum_vec v" + using assms +proof (induct v) + case vNil + then show ?case + by (simp add: vec_mset_vNil) + next + case (vCons a v) + then have lim: "dim_vec (vCons a v) \ 1" + by simp + have "(\ i. i < dim_vec v \ v $ i = 1 \ v$ i = 0)" + using vCons.prems by force + then have hyp: "of_nat (count_vec v 1) = sum_vec v" + using vCons.hyps by blast + have "sum (($) (vCons a v)) {0.. i. i < dim_vec v \ v $ i = 1 \ v $ i = 0" + shows "count_vec v 1 + count_vec v 0 = dim_vec v" +proof - + have ss: "vec_set v \ {0, 1}" using assms by (auto simp add: vec_set_def) + have "dim_vec v = size (vec_mset v)" + by (simp add: mset_vec_same_size) + have "size (vec_mset v) = (\ x \ (vec_set v) . count (vec_mset v) x)" + by (simp add: vec_mset_set size_multiset_overloaded_eq) + also have "... = (\ x \ {0, 1} . count (vec_mset v) x)" + using size_count_mset_ss ss + by (metis calculation finite.emptyI finite.insertI vec_mset_set) + finally have "size (vec_mset v) = count (vec_mset v) 0 + count (vec_mset v) 1" by simp + thus ?thesis + by (simp add: \dim_vec v = size (vec_mset v)\) +qed + +lemma count_vec_sum_zeros: + fixes v :: "'a :: {ring_1} vec" + assumes "\ i. i < dim_vec v \ v $ i = 1 \ v $ i = 0" + shows "of_nat (count_vec v 0) = of_nat (dim_vec v) - sum_vec v" + using count_vec_two_elems assms count_vec_sum_ones + by (metis add_diff_cancel_left' of_nat_add) + +lemma count_vec_sum_ones_alt: + fixes v :: "'a :: {ring_1} vec" + assumes "vec_set v \ {0, 1}" + shows "of_nat (count_vec v 1) = sum_vec v" +proof - + have "\ i. i < dim_vec v \ v $ i = 1 \ v $ i = 0" using assms + by (meson insertE singletonD subsetD vec_setI) + thus ?thesis using count_vec_sum_ones + by blast +qed + +lemma setv_not_in_count0_iff: "a \ set\<^sub>v v \ count_vec v a = 0" + using setv_vec_mset_not_in_iff + by (metis count_eq_zero_iff) + +lemma sum_count_vec: + assumes "finite (set\<^sub>v v)" + shows "(\ i \ set\<^sub>v v. count_vec v i) = dim_vec v" +using assms proof (induct "v") + case vNil + then show ?case + by (simp add: count_vec_empty) +next + case (vCons a v) + then show ?case proof (cases "a \ set\<^sub>v v") + case True + have cv: "\ x. x \(set\<^sub>v v) - {a} \ count_vec (vCons a v) x = count_vec v x" + using count_vec_vCons by (metis DiffD2 singletonI) + then have "sum (count_vec (vCons a v)) (set\<^sub>v (vCons a v)) = sum (count_vec (vCons a v)) (set\<^sub>v v)" + using vCons_set_contains_in True by metis + also have "... = count_vec (vCons a v) a + sum (count_vec (vCons a v)) ((set\<^sub>v v) - {a})" + using sum.remove True vCons.prems(1) by (metis vCons_set_contains_in) + also have "... = count_vec v a + 1 + sum (count_vec v) ((set\<^sub>v v) - {a})" + using cv count_vec_vCons by (metis sum.cong) + also have "... = 1 + sum (count_vec v) ((set\<^sub>v v))" + using sum.remove add.commute vCons.prems vCons_set_contains_in True + by (metis (no_types, opaque_lifting) ab_semigroup_add_class.add_ac(1)) + also have "... = 1 + dim_vec v" using vCons.hyps + by (metis True vCons.prems vCons_set_contains_in) + finally show ?thesis by simp + next + case False + then have cv: "\ x. x \(set\<^sub>v v) \ count_vec (vCons a v) x = count_vec v x" + using count_vec_vCons by (metis) + have f: "finite (set\<^sub>v v)" + using vCons.prems False vCons_set_contains_add by (metis Un_infinite) + have "sum (count_vec (vCons a v)) (set\<^sub>v (vCons a v)) = + count_vec (vCons a v) a + sum (count_vec (vCons a v)) (set\<^sub>v v)" + using False vCons_set_contains_add + by (metis Un_insert_right finite_Un sum.insert sup_bot_right vCons.prems ) + also have "... = count_vec v a + 1 + sum (count_vec v) ((set\<^sub>v v) )" + using cv count_vec_vCons by (metis sum.cong) + also have "... = 1 + sum (count_vec v) ((set\<^sub>v v))" + using False setv_not_in_count0_iff by (metis add_0) + finally show ?thesis using vCons.hyps f by simp + qed +qed + +lemma sum_setv_subset_eq: + assumes "finite A" + assumes "set\<^sub>v v \ A" + shows "(\ i \ set\<^sub>v v. count_vec v i) = (\ i \ A. count_vec v i)" +proof - + have ni: "\ x. x \ set\<^sub>v v \ count_vec v x = 0" + by (simp add: setv_not_in_count0_iff) + have "(\ i \ A. count_vec v i) = (\ i \ A - (set\<^sub>v v). count_vec v i) + (\ i \ (set\<^sub>v v). count_vec v i)" + using sum.subset_diff assms by auto + then show ?thesis using ni + by simp +qed + +lemma sum_count_vec_subset: "finite A \ set\<^sub>v v \ A \ (\ i \ A. count_vec v i) = dim_vec v" + using sum_setv_subset_eq sum_count_vec finite_subset by metis + +text \An abbreviation for checking if an element is in a vector \ + +abbreviation vec_contains :: "'a \ 'a vec \ bool" (infix "\$" 50)where +"a \$ v \ a \ set\<^sub>v v" + +lemma vec_set_mset_contains_iff: "a \$ v \ a \# vec_mset v" + by (simp add: vec_mset_def vec_set_def) + +lemma vec_contains_count_gt1_iff: "a \$ v \ count_vec v a \ 1" + by (simp add: vec_set_mset_contains_iff) + +lemma vec_contains_obtains_index: + assumes "a \$ v" + obtains i where "i < dim_vec v" and "v $ i = a" + by (metis assms vec_setE) + +lemma vec_count_eq_list_count: "count (mset xs) a = count_vec (vec_of_list xs) a" + by (simp add: list_vec mset_vec_eq_mset_list) + +lemma vec_contains_col_elements_mat: + assumes "j < dim_col M" + assumes "a \$ col M j" + shows "a \ elements_mat M" +proof - + have "dim_vec (col M j) = dim_row M" by simp + then obtain i where ilt: "i < dim_row M" and "(col M j) $ i = a" + using vec_setE by (metis assms(2)) + then have "M $$ (i, j) = a" + by (simp add: assms(1)) + thus ?thesis using assms(1) ilt + by blast +qed + +lemma vec_contains_row_elements_mat: + assumes "i < dim_row M" + assumes "a \$ row M i" + shows "a \ elements_mat M" +proof - + have "dim_vec (row M i) = dim_col M" by simp + then obtain j where jlt: "j < dim_col M" and "(row M i) $ j = a" using vec_setE + by (metis assms(2)) + then have "M $$ (i, j) = a" + by (simp add: assms(1)) + thus ?thesis using assms(1) jlt + by blast +qed + +lemma vec_contains_img: "a \$ v \ f a \$ (map_vec f v)" + by (metis index_map_vec(1) index_map_vec(2) vec_contains_obtains_index vec_setI) + +text \ The existing vector library contains the identity and zero vectors, but no definition +of a vector where all elements are 1, as defined below \ + +definition all_ones_vec :: "nat \ 'a :: {zero,one} vec" ("u\<^sub>v") where + "u\<^sub>v n \ vec n (\ i. 1)" + +lemma dim_vec_all_ones[simp]: "dim_vec (u\<^sub>v n) = n" + by (simp add: all_ones_vec_def) + +lemma all_ones_index [simp]: "i < n \ u\<^sub>v n $ i = 1" + by (simp add: all_ones_vec_def) + +lemma dim_vec_mult_vec_mat [simp]: "dim_vec (v \<^sub>v* A) = dim_col A" + unfolding mult_vec_mat_def by simp + +lemma all_ones_vec_smult[simp]: "i < n \ ((k :: ('a :: {one, zero, monoid_mult})) \\<^sub>v (u\<^sub>v n)) $ i = k" + by (simp add: smult_vec_def) + +text \Extra lemmas on existing vector operations \ + +lemma smult_scalar_prod_sum: + fixes x :: "'a :: {comm_ring_1}" + assumes "vx \ carrier_vec n" + assumes "vy \ carrier_vec n" + shows "(\ i \ {0..\<^sub>v vx) $ i) * ((y \\<^sub>v vy) $ i)) = x * y * (vx \ vy)" +proof - + have "\ i . i < n \ ((x \\<^sub>v vx) $ i) * ((y \\<^sub>v vy) $ i) = x * y * (vx $ i) * (vy $ i)" + using assms by simp + then have "(\ i \ {0..\<^sub>v vx) $ i) * ((y \\<^sub>v vy) $ i)) = + (\ i \ {0.. i \ {0.. i. (vx $ i) * (vy $ i))" "{0.. i \ {0..\<^sub>v vx) $ i) * ((y \\<^sub>v vy) $ i)) = x * y * (vx \ vy)" + using scalar_prod_def assms by (metis carrier_vecD) + thus ?thesis by simp +qed + +lemma scalar_prod_double_sum_fn_vec: + fixes c :: "nat \ ('a :: {comm_semiring_0})" + fixes f :: "nat \ 'a vec" + assumes "\ j . j < k \ dim_vec (f j) = n" + shows "(vec n (\i. \j = 0.. (vec n (\i. \j = 0.. j1 \ {0.. (f j1))) + + (\ j1 \ {0.. j2 \ ({0..< k} - {j1}) . c j1 * c j2 * ((f j1) \ (f j2))))" +proof - + have sum_simp: "\ j1 j2. (\l \ {0..l \ {0..l \ {0..l \ {0..l \ {0..l \ {0.. l. (f j1) $ l * (f j2) $ l" "{0..i. \j = 0.. (vec n (\i. \j = 0.. l = 0..j1 = 0..j2 = 0.. l \ {0.. j1 \ {0.. j2 \ {0..< k}. c j1 * (f j1) $ l * (c j2 * (f j2) $ l))))" + by (metis (no_types) sum_product) + also have "... = (\ j1 \ {0.. j2 \ {0..l \ {0.. l j1 j2 .(c j1 * (f j1) $ l * (c j2 * (f j2) $ l))" "{0.. j1 \ {0.. j2 \ {0..l \ {0.. j1 \ {0.. j2 \ {0.. (f j2))))" + unfolding scalar_prod_def using dim_col assms by simp + finally show ?thesis + using double_sum_split_case by fastforce +qed + +lemma vec_prod_zero: "(0\<^sub>v n) \ (0\<^sub>v n) = 0" + by simp + +lemma map_vec_compose: "map_vec f (map_vec g v) = map_vec (f \ g) v" + by auto + +subsection \Matrix Extras\ + +text \As with vectors, the all ones mat definition defines the concept of a matrix where all +elements are 1 \ + +definition all_ones_mat :: "nat \ 'a :: {zero,one} mat" ("J\<^sub>m") where + "J\<^sub>m n \ mat n n (\ (i,j). 1)" + +lemma all_ones_mat_index[simp]: "i < dim_row (J\<^sub>m n) \ j < dim_col (J\<^sub>m n) \ J\<^sub>m n $$ (i, j)= 1" + by (simp add: all_ones_mat_def) + +lemma all_ones_mat_dim_row[simp]: "dim_row (J\<^sub>m n) = n" + by (simp add: all_ones_mat_def) + +lemma all_ones_mat_dim_col[simp]: "dim_col (J\<^sub>m n) = n" + by (simp add: all_ones_mat_def) + +text \Basic lemmas on existing matrix operations \ +lemma index_mult_vec_mat[simp]: "j < dim_col A \ (v \<^sub>v* A) $ j = v \ col A j" + by (auto simp: mult_vec_mat_def) + +lemma transpose_mat_mult_entries: "i < dim_row A \ j < dim_row A \ + (A * A\<^sup>T) $$ (i, j) = (\k\ {0..<(dim_col A)}. (A $$ (i, k)) * (A $$ (j, k)))" + by (simp add: times_mat_def scalar_prod_def) + +lemma transpose_mat_elems: "elements_mat A = elements_mat A\<^sup>T" + apply (auto simp add: transpose_mat_def) + apply (metis elements_matD elements_matI index_transpose_mat(1) mat_carrier transpose_mat_def) + by fastforce + +lemma row_elems_subset_mat: "i < dim_row N \ vec_set (row N i) \ elements_mat N" + by (auto simp add: vec_set_def elements_matI) + +lemma col_elems_subset_mat: "i < dim_col N \ vec_set (col N i) \ elements_mat N" + by (auto simp add: vec_set_def elements_matI) + +lemma obtain_row_index: + assumes "r \ set (rows M)" + obtains i where "row M i = r" and "i < dim_row M" + by (metis assms in_set_conv_nth length_rows nth_rows) + +lemma row_prop_cond: "(\ i. i < dim_row M \ P (row M i)) \ r \ set (rows M) \ P r" + using obtain_row_index by metis + +lemma obtain_col_index: + assumes "c \ set (cols M)" + obtains j where "col M j = c" and "j < dim_col M" + by (metis assms cols_length cols_nth obtain_set_list_item) + +lemma col_prop_cond: "(\ j. j < dim_col M \ P (col M j)) \ c \ set (cols M) \ P c" + using obtain_col_index by metis + +text \ Lemmas on the @{term "map_mat"} definition \ + +lemma row_map_mat[simp]: + assumes "i < dim_row A" shows "row (map_mat f A) i = map_vec f (row A i)" + unfolding map_mat_def map_vec_def using assms by auto + +lemma map_vec_mat_rows: "map (map_vec f) (rows M) = rows ((map_mat f) M)" + by (simp add: map_nth_eq_conv) + +lemma map_vec_mat_cols: "map (map_vec f) (cols M) = cols ((map_mat f) M)" + by (simp add: map_nth_eq_conv) + +lemma map_mat_compose: "map_mat f (map_mat g A) = map_mat (f \ g) A" + by (auto) + +lemma map_mat_elements: "elements_mat (map_mat f A) = f ` (elements_mat A)" +proof (auto) + fix x assume "x \ elements_mat (map_mat f A)" + then obtain i j where "i < dim_row (map_mat f A)" and "j < dim_col (map_mat f A)" and "(map_mat f A) $$ (i, j) = x" + by auto + then show "x \ f ` elements_mat A" by auto +next + fix xa assume "xa \ elements_mat A" + then obtain i j where "i < dim_row A" and "j < dim_col A" and "A $$ (i, j) = xa" by auto + then show "f xa \ elements_mat (map_mat f A)" by auto +qed + +text \ Reasoning on sets and multisets of matrix elements \ +lemma set_cols_carrier: "A \ carrier_mat m n \ v \ set (cols A) \ v \ carrier_vec m" + by (auto simp: cols_def) + +lemma mset_cols_index_map: "image_mset (\ j. col M j) (mset_set {0..< dim_col M}) = mset (cols M)" + by (simp add: cols_def) + +lemma mset_rows_index_map: "image_mset (\ i. row M i) (mset_set {0..< dim_row M}) = mset (rows M)" + by (simp add: rows_def) + +lemma index_to_col_card_size_prop: + assumes "i < dim_row M" + assumes "\ j. j < dim_col M \ P j \ Q (col M j)" + shows "card {j . j < dim_col M \ P j} = size {#c \# (mset (cols M)) . Q c #}" +proof - + have "card {j . j < dim_col M \ P j} = size (mset_set({j \ {0.. {0.. j. col M j) {# j \# mset_set {0..< dim_col M} . Q (col M j) #})" + by simp + also have "... = size ({# c \# (image_mset (\ j. col M j) (mset_set {0..< dim_col M})) . Q c #})" + using image_mset_filter_swap[of "(\ j. col M j)" "Q" "(mset_set {0..< dim_col M})"] by simp + finally have "card {j . j < dim_col M \ P j} = size ({# c \# (mset (cols M)) . Q c #})" + using mset_cols_index_map by metis + thus ?thesis by simp +qed + +lemma index_to_row_card_size_prop: + assumes "j < dim_col M" + assumes "\ i. i < dim_row M \ P i \ Q (row M i)" + shows "card {i . i < dim_row M \ P i} = size {#r \# (mset (rows M)) . Q r #}" +proof - + have "card {i . i < dim_row M \ P i} = size (mset_set({i \ {0.. {0.. i. row M i) {# i \# mset_set {0..< dim_row M} . Q (row M i) #})" + by simp + also have "... = size ({# r \# (image_mset (\ i. row M i) (mset_set {0..< dim_row M})) . Q r #})" + using image_mset_filter_swap[of "(\ j. row M j)" "Q" "(mset_set {0..< dim_row M})"] by simp + finally have "card {j . j < dim_row M \ P j} = size ({# c \# (mset (rows M)) . Q c #})" + using mset_rows_index_map by metis + thus ?thesis by simp +qed + +lemma setv_row_subset_mat_elems: + assumes "v \ set (rows M)" + shows "set\<^sub>v v \ elements_mat M" +proof (intro subsetI) + fix x assume "x \$ v" + then obtain i where "v = row M i" and "i < dim_row M" + by (metis assms obtain_row_index) + then show "x \ elements_mat M" + by (metis \x \$ v\ vec_contains_row_elements_mat) +qed + +lemma setv_col_subset_mat_elems: + assumes "v \ set (cols M)" + shows "set\<^sub>v v \ elements_mat M" +proof (intro subsetI) + fix x assume "x \$ v" + then obtain i where "v = col M i" and "i < dim_col M" + by (metis assms obtain_col_index) + then show "x \ elements_mat M" + by (metis \x \$ v\ vec_contains_col_elements_mat) +qed + +subsection \ Vector and Matrix Homomorphism \ + +text \We extend on the existing lemmas on homomorphism mappings as applied to vectors and matrices \ + +context semiring_hom +begin + +lemma vec_hom_smult2: + assumes "dim_vec v2 \ dim_vec v1" + shows "hom (v1 \ v2) = vec\<^sub>h v1 \ vec\<^sub>h v2" + unfolding scalar_prod_def using index_map_vec assms by (auto simp add: hom_distribs) + +end + + +lemma map_vec_vCons: "vCons (f a) (map_vec f v) = map_vec f (vCons a v)" + by (intro eq_vecI, simp_all add: vec_index_vCons) + +context inj_zero_hom +begin + +lemma vec_hom_zero_iff[simp]: "(map_vec hom x = 0\<^sub>v n) = (x = 0\<^sub>v n)" +proof - + { + fix i + assume i: "i < n" "dim_vec x = n" + hence "map_vec hom x $ i = 0 \ x $ i = 0" + using index_map_vec(1)[of i x] by simp + } note main = this + show ?thesis unfolding vec_eq_iff by (simp, insert main, auto) +qed + +lemma mat_hom_inj: "map_mat hom A = map_mat hom B \ A = B" + unfolding mat_eq_iff by auto + +lemma vec_hom_inj: "map_vec hom v = map_vec hom w \ v = w" + unfolding vec_eq_iff by auto + +lemma vec_hom_set_distinct_iff: + fixes xs :: "'a vec list" + shows "distinct xs \ distinct (map (map_vec hom) xs)" + using vec_hom_inj by (induct xs) (auto) + +lemma vec_hom_mset: "image_mset hom (vec_mset v) = vec_mset (map_vec hom v)" + apply (induction v) + apply (metis mset.simps(1) vec_mset_img_map vec_mset_vNil vec_of_list_Nil) + by (metis mset_vec_eq_mset_list vec_list vec_mset_img_map) + +lemma vec_hom_set: "hom ` set\<^sub>v v = set\<^sub>v (map_vec hom v)" +proof (induct v) + case vNil + then show ?case by (metis image_mset_empty set_image_mset vec_hom_zero_iff vec_mset_set vec_mset_vNil zero_vec_zero) +next + case (vCons a v) + have "hom ` set\<^sub>v (vCons a v) = hom ` ({a} \ set\<^sub>v v)" + by (metis Un_commute insert_absorb insert_is_Un vCons_set_contains_add vCons_set_contains_in) + also have "... = {hom a} \ (hom ` (set\<^sub>v v))" by simp + also have "... = {hom a} \ (set\<^sub>v (map_vec hom v))" using vCons.hyps by simp + also have "... = set\<^sub>v (vCons (hom a) (map_vec hom v))" + by (metis Un_commute insert_absorb insert_is_Un vCons_set_contains_add vCons_set_contains_in) + finally show ?case using map_vec_vCons + by metis +qed + +end + +subsection \ Zero One injections and homomorphisms \ + +text \Define a locale to encapsulate when a function is injective on a certain set (i.e. not +a universal homomorphism for the type\ +locale injective_lim = + fixes A :: "'a set" + fixes f :: "'a \ 'b" assumes injectivity_lim: "\x y. x \ A \ y \ A \ f x = f y \ x = y" +begin + lemma eq_iff[simp]: "x \ A \ y \ A \ f x = f y \ x = y" using injectivity_lim by auto +lemma inj_on_f: "inj_on f A" by (auto intro: inj_onI) + +end + +sublocale injective \ injective_lim Univ + by(unfold_locales) simp + +context injective_lim +begin + +lemma mat_hom_inj_lim: + assumes "elements_mat M \ A" and "elements_mat N \ A" + shows "map_mat f M = map_mat f N \ M = N" + unfolding mat_eq_iff apply auto + using assms injectivity_lim by blast + +lemma vec_hom_inj_lim: assumes "set\<^sub>v v \ A" "set\<^sub>v w \ A" + shows "map_vec f v = map_vec f w \ v = w" + unfolding vec_eq_iff apply (auto) + using vec_setI in_mono assms injectivity_lim by metis + +lemma lim_inj_hom_count_vec: + assumes "set\<^sub>v v \ A" + assumes "x \ A" + shows "count_vec v x = count_vec (map_vec f v) (f x)" +using assms proof (induct v) + case vNil + have "(map_vec f vNil) = vNil" by auto + then show ?case + by (smt (verit) count_vec_vNil) +next + case (vCons a v) + have 1: "map_vec f (vCons a v) = vCons (f a) (map_vec f v)" + by (simp add: map_vec_vCons) + then show ?case proof (cases "a = x") + case True + have "count_vec (vCons a v) x = count_vec v x + 1" + by (simp add: True count_vec_vCons) + then show ?thesis using Un_subset_iff 1 count_vec_vCons vCons.hyps vCons.prems(1) + vCons.prems(2) vCons_set_contains_add vCons_set_contains_in + by metis + next + case False + then have "count_vec (vCons a v) x = count_vec v x" + by (simp add: count_vec_vCons) + then show ?thesis using "1" Un_empty_right Un_insert_right count_vec_vCons insert_absorb insert_subset + vCons.hyps vCons.prems(1) vCons.prems(2) vCons_set_contains_add vCons_set_contains_in + by (metis (no_types, lifting) injectivity_lim) + qed +qed + +lemma vec_hom_lim_set_distinct_iff: + fixes xs :: "'a vec list" + assumes "\ v . v \ set (xs) \ set\<^sub>v v \ A" + shows "distinct xs \ distinct (map (map_vec f) xs)" + using assms vec_hom_inj_lim by (induct xs, simp_all) (metis (no_types, lifting) image_iff) + +lemma mat_rows_hom_lim_distinct_iff: + assumes "elements_mat M \ A" + shows "distinct (rows M) \ distinct (map (map_vec f) (rows M))" + apply (intro vec_hom_lim_set_distinct_iff) + using setv_row_subset_mat_elems assms by blast + +lemma mat_cols_hom_lim_distinct_iff: + assumes "elements_mat M \ A" + shows "distinct (cols M) \ distinct (map (map_vec f) (cols M))" + apply (intro vec_hom_lim_set_distinct_iff) + using setv_col_subset_mat_elems assms by blast + +end + +locale inj_on_01_hom = zero_hom + one_hom + injective_lim "{0, 1}" hom +begin + +lemma inj_0_iff: "x \ {0, 1} \ hom x = 0 \ x = 0" + by (metis hom_zero insertI1 local.eq_iff) + +lemma inj_1_iff: "x \ {0, 1} \ hom x = 1 \ x = 1" + using inj_0_iff by fastforce + +end + +context zero_neq_one +begin + +definition of_zero_neq_one :: "'b :: {zero_neq_one} \ 'a" where +"of_zero_neq_one x \ if (x = 0) then 0 else 1" + +lemma of_zero_neq_one_1 [simp]: "of_zero_neq_one 1 = 1" + by (simp add: of_zero_neq_one_def) + +lemma of_zero_neq_one_0 [simp]: "of_zero_neq_one 0 = 0" + by (simp add: of_zero_neq_one_def) + +lemma of_zero_neq_one_0_iff[iff]: "of_zero_neq_one x = 0 \ x = 0" + by (simp add: of_zero_neq_one_def) + +lemma of_zero_neq_one_lim_eq: "x \ {0, 1} \ y \ {0, 1} \ of_zero_neq_one x = of_zero_neq_one y \ x = y" + by (auto simp add: of_zero_neq_one_def) + + +end + +interpretation of_zero_hom: zero_hom_0 of_zero_neq_one + by(unfold_locales) (simp_all) + +interpretation of_injective_lim: injective_lim "{0, 1}" of_zero_neq_one + by (unfold_locales)(simp_all add: of_zero_neq_one_lim_eq) + +interpretation of_inj_on_01_hom: inj_on_01_hom of_zero_neq_one + by (unfold_locales)(simp_all add: of_zero_neq_one_lim_eq) + +text \We want the ability to transform any 0-1 vector or matrix to another @{typ "'c :: zero_neq_one"} type \ +definition lift_01_vec :: "'b :: {zero_neq_one} vec \ 'c :: {zero_neq_one} vec" where +"lift_01_vec v \ map_vec of_zero_neq_one v" + +lemma lift_01_vec_simp[simp]: "dim_vec (lift_01_vec v) = dim_vec v" +"i < dim_vec v \ (lift_01_vec v) $ i = of_zero_neq_one (v $ i)" + by (simp_all add: lift_01_vec_def) + +lemma lift_01_vec_count: + assumes "set\<^sub>v v \ {0, 1}" + assumes "x \ {0, 1}" + shows "count_vec v x = count_vec (lift_01_vec v) (of_zero_neq_one x)" + using of_injective_lim.lim_inj_hom_count_vec + by (metis assms(1) assms(2) lift_01_vec_def) + +definition lift_01_mat :: "'b :: {zero_neq_one} mat \ 'c :: {zero_neq_one} mat" where +"lift_01_mat M \ map_mat of_zero_neq_one M" + +lemma lift_01_mat_simp[simp]: "dim_row (lift_01_mat M) = dim_row M" + "dim_col (lift_01_mat M) = dim_col M" + "i < dim_row M \ j < dim_col M \ (lift_01_mat M) $$ (i, j) = of_zero_neq_one (M $$ (i, j))" + by (simp_all add: lift_01_mat_def) + +lemma lift_01_mat_carrier: "lift_01_mat M \ carrier_mat (dim_row M) (dim_col M)" + using lift_01_mat_def by auto + +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/ROOT b/thys/Fishers_Inequality/ROOT new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/ROOT @@ -0,0 +1,16 @@ +chapter AFP + +session Fishers_Inequality (AFP) = "Berlekamp_Zassenhaus" + + options [timeout = 600] + sessions + Design_Theory + "HOL-Combinatorics" + Groebner_Bases + Polynomial_Factorization + "List-Index" + BenOr_Kozen_Reif + theories + Fishers_Inequality_Root + document_files + "root.bib" + "root.tex" diff --git a/thys/Fishers_Inequality/Rank_Argument_General.thy b/thys/Fishers_Inequality/Rank_Argument_General.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Rank_Argument_General.thy @@ -0,0 +1,443 @@ +(* Title: Rank_Argument_General.thy + Author: Chelsea Edmonds +*) + +section \Rank Argument - General \ +text \General lemmas to enable reasoning using the rank argument. This is described by Godsil +\cite{godsilToolsLinearAlgebra} and Bukh \cite{bukhAlgebraicMethodsCombinatoricsa}, both of whom +present it as a foundational technique \ +theory Rank_Argument_General imports Dual_Systems Jordan_Normal_Form.Determinant +Jordan_Normal_Form.DL_Rank Jordan_Normal_Form.Ring_Hom_Matrix BenOr_Kozen_Reif.More_Matrix +begin + +subsection \Row/Column Operations \ +text \Extensions to the existing elementary operations are made to enable reasoning on multiple +operations at once, similar to mathematical literature\ + +lemma index_mat_addrow_basic [simp]: + "i < dim_row A \ j < dim_col A \ addrow a k l A $$ (i,j) = (if k = i then + ( a * (A $$ (l,j)) + (A $$ (i,j))) else A $$ (i,j))" + "i < dim_row A \ j < dim_col A \ addrow a i l A $$ (i,j) = (a * (A $$ (l,j)) + (A $$ (i,j)))" + "i < dim_row A \ j < dim_col A \ k \ i \ addrow a k l A $$ (i,j) = A $$(i,j)" + "dim_row (addrow a k l A) = dim_row A" "dim_col (addrow a k l A) = dim_col A" + unfolding mat_addrow_def by auto + +text\Function to add a column to multiple other columns \ +fun add_col_to_multiple :: "'a :: semiring_1 \ nat list \ nat \ 'a mat \ 'a mat" where +"add_col_to_multiple a [] l A = A" | +"add_col_to_multiple a (k # ks) l A = (addcol a k l (add_col_to_multiple a ks l A))" + +text \Function to add a row to multiple other rows \ +fun add_row_to_multiple :: "'a :: semiring_1 \ nat list \ nat \ 'a mat \ 'a mat" where +"add_row_to_multiple a [] l A = A" | +"add_row_to_multiple a (k # ks) l A = (addrow a k l (add_row_to_multiple a ks l A))" + +text \Function to add multiple rows to a single row \ +fun add_multiple_rows :: "'a :: semiring_1 \ nat \ nat list \ 'a mat \ 'a mat" where +"add_multiple_rows a k [] A = A" | +"add_multiple_rows a k (l # ls) A = (addrow a k l (add_multiple_rows a k ls A))" + +text \Function to add multiple columns to a single col \ +fun add_multiple_cols :: "'a :: semiring_1 \ nat \ nat list \ 'a mat \ 'a mat" where +"add_multiple_cols a k [] A = A" | +"add_multiple_cols a k (l # ls) A = (addcol a k l (add_multiple_cols a k ls A))" + +text \Basic lemmas on dimension and indexing of resulting matrix from above functions \ +lemma add_multiple_rows_dim [simp]: +"dim_row (add_multiple_rows a k ls A) = dim_row A" +"dim_col (add_multiple_rows a k ls A) = dim_col A" + by (induct ls) simp_all + +lemma add_multiple_rows_index_unchanged [simp]: + "i < dim_row A \ j < dim_col A \ k \ i \ add_multiple_rows a k ls A $$ (i,j) = A $$(i,j)" + by (induct ls) (simp_all) + +lemma add_multiple_rows_index_eq: + assumes "i < dim_row A" and "j < dim_col A" and "i \ set ls" and "\ l . l \ set ls \ l < dim_row A" + shows "add_multiple_rows a i ls A $$ (i,j) = (\l\ls. a * A $$(l,j)) + A$$(i,j)" + using assms proof (induct ls) + case Nil + then show ?case by simp +next + case (Cons aa ls) + then have ne: "i \ aa" + by auto + have lt: "aa < dim_row A" using assms(1) + by (simp add: Cons.prems(4)) + have "(add_multiple_rows a i (aa # ls) A) $$ (i, j) = + (addrow a i aa (add_multiple_rows a i ls A)) $$ (i, j)" + by simp + also have "... = a * add_multiple_rows a i ls A $$ (aa, j) + (add_multiple_rows a i ls A) $$ (i, j)" + using assms(1) assms(2) index_mat_addrow_basic(2)[of "i" "(add_multiple_rows a i ls A)" "j" "a" "aa"] + by simp + also have "... = a * A $$(aa, j) + (add_multiple_rows a i ls A) $$ (i, j)" + using lt ne by (simp add: assms(2)) + also have "... = a * A $$(aa, j) + (\l\ls. a * A $$ (l, j)) + A $$ (i, j)" + using Cons.hyps assms(1) assms(2) Cons.prems(3) Cons.prems(4) + by (metis (mono_tags, lifting) ab_semigroup_add_class.add_ac(1) list.set_intros(2)) + finally show "(add_multiple_rows a i (aa # ls) A) $$ (i, j) = + (\l\(aa #ls). a * A $$ (l, j)) + A $$ (i, j)" + by simp +qed + +lemma add_multiple_rows_index_eq_bounds: + assumes "i < dim_row A" and "j < dim_col A" and "i < low \ i \ up" and "up \ dim_row A" + shows "add_multiple_rows a i [low..l=low.. set [low.. l . l \ set [low.. l < dim_row A" using assms(4) by auto + thus ?thesis using add_multiple_rows_index_eq[of i A j "[low.. l. a * A $$(l,j)" low up] notin assms(1) assms(2) by simp +qed + +lemma add_multiple_cols_dim [simp]: + "dim_row (add_multiple_cols a k ls A) = dim_row A" + "dim_col (add_multiple_cols a k ls A) = dim_col A" + by (induct ls) simp_all + +lemma add_multiple_cols_index_unchanged [simp]: + "i < dim_row A \ j < dim_col A \ k \ j \ add_multiple_cols a k ls A $$ (i,j) = A $$(i,j)" + by (induct ls) (simp_all) + +lemma add_multiple_cols_index_eq: + assumes "i < dim_row A" and "j < dim_col A" and "j \ set ls" and "\ l . l \ set ls \ l < dim_col A" + shows "add_multiple_cols a j ls A $$ (i,j) = (\l\ls. a * A $$(i,l)) + A$$(i,j)" + using assms +proof (induct ls) + case Nil + then show ?case by simp +next + case (Cons aa ls) + then have ne: "j \ aa" + by auto + have lt: "aa < dim_col A" using assms + by (simp add: Cons.prems(4)) + have "(add_multiple_cols a j (aa # ls) A) $$ (i, j) = (addcol a j aa (add_multiple_cols a j ls A)) $$ (i, j)" + by simp + also have "... = a * add_multiple_cols a j ls A $$ (i, aa) + (add_multiple_cols a j ls A) $$ (i, j)" + using assms index_mat_addcol by simp + also have "... = a * A $$(i, aa) + (add_multiple_cols a j ls A) $$ (i, j)" + using lt ne by (simp add: assms(1)) + also have "... = a * A $$(i, aa) + (\l\ls. a * A $$ (i, l)) + A $$ (i, j)" + using Cons.hyps assms(1) assms(2) Cons.prems(3) Cons.prems(4) + by (metis (mono_tags, lifting) ab_semigroup_add_class.add_ac(1) list.set_intros(2)) + finally show ?case by simp +qed + +lemma add_multiple_cols_index_eq_bounds: + assumes "i < dim_row A" and "j < dim_col A" and "j < low \ j \ up" and "up \ dim_col A" + shows "add_multiple_cols a j [low..l=low.. set [low.. l . l \ set [low.. l < dim_col A" using assms(4) by auto + thus ?thesis using add_multiple_cols_index_eq[of i A j "[low.. l. a * A $$(i,l)" low up] notin assms(1) assms(2) by simp +qed + +lemma add_row_to_multiple_dim [simp]: + "dim_row (add_row_to_multiple a ks l A) = dim_row A" + "dim_col (add_row_to_multiple a ks l A) = dim_col A" + by (induct ks) simp_all + +lemma add_row_to_multiple_index_unchanged [simp]: + "i < dim_row A \ j < dim_col A \ i \ set ks \ add_row_to_multiple a ks l A $$ (i,j) = A $$(i,j)" + by (induct ks) simp_all + +lemma add_row_to_multiple_index_unchanged_bound: + "i < dim_row A \ j < dim_col A \ i < low \ i \ up \ + add_row_to_multiple a [low.. set ks" and "distinct ks" and "l \ set ks" + and "l < dim_row A" + shows "add_row_to_multiple a ks l A $$ (i,j) = (a * A$$(l, j)) + A$$(i,j)" + using assms +proof (induct ks) + case Nil + then show ?case by simp +next + case (Cons aa ls) + then have lnotin: "l \ set ls" using assms by simp + then show ?case + proof (cases "i = aa") + case True + then have inotin: "i \ set ls" using assms + using Cons.prems(4) by fastforce + have "add_row_to_multiple a (aa # ls) l A $$ (i, j) = + (addrow a aa l (add_row_to_multiple a ls l A)) $$ (i, j)" by simp + also have "... = (a * ((add_row_to_multiple a ls l A) $$ (l,j)) + + ((add_row_to_multiple a ls l A) $$ (i,j)))" + using True assms(1) assms(2) by auto + also have "... = a* A $$ (l, j) + ((add_row_to_multiple a ls l A) $$ (i,j))" + using assms lnotin by simp + finally have "add_row_to_multiple a (aa # ls) l A $$ (i, j) = a* A $$ (l,j) + A $$ (i, j)" + using inotin assms by simp + then show ?thesis by simp + next + case False + then have iin: "i \ set ls" using assms + by (meson Cons.prems(3) set_ConsD) + have "add_row_to_multiple a (aa # ls) l A $$ (i, j) = (addrow a aa l (add_row_to_multiple a ls l A)) $$ (i, j)" + by simp + also have "... = ((add_row_to_multiple a ls l A) $$ (i,j))" + using False assms by auto + finally have "add_row_to_multiple a (aa # ls) l A $$ (i, j) = a * A $$ (l, j) + A $$ (i, j)" + using Cons.hyps by (metis Cons.prems(4) assms(1) assms(2) assms(6) distinct.simps(2) iin lnotin) + then show ?thesis by simp + qed +qed + +lemma add_row_to_multiple_index_change_bounds: + assumes "i < dim_row A" and "j < dim_col A" and "i \ low" and "i < up" and "l < low \ l \ up" + and "l < dim_row A" + shows "add_row_to_multiple a [low.. set [low.. set [low.. j < dim_col A \ j \ set ks \ add_col_to_multiple a ks l A $$ (i,j) = A $$(i,j)" + by (induct ks) simp_all + +lemma add_col_to_multiple_index_unchanged_bound: + "i < dim_row A \ j < dim_col A \ j < low \ j \ up \ + add_col_to_multiple a [low.. set ks" and "distinct ks" and "l \ set ks" + and "l < dim_col A" + shows "add_col_to_multiple a ks l A $$ (i,j) = (a * A$$(i, l)) + A$$(i,j)" + using assms +proof (induct ks) + case Nil + then show ?case by simp +next + case (Cons aa ls) + then have lnotin: "l \ set ls" using assms by simp + then show ?case + proof (cases "j = aa") + case True + then have inotin: "j \ set ls" using assms + using Cons.prems(4) by fastforce + have "add_col_to_multiple a (aa # ls) l A $$ (i, j) = + (addcol a aa l (add_col_to_multiple a ls l A)) $$ (i, j)" by simp + also have "... = (a * ((add_col_to_multiple a ls l A) $$ (i,l)) + + ((add_col_to_multiple a ls l A) $$ (i,j)))" + using True assms(1) assms(2) by auto + also have "... = a* A $$ (i, l) + ((add_col_to_multiple a ls l A) $$ (i,j))" + using assms lnotin by simp + finally have "add_col_to_multiple a (aa # ls) l A $$ (i, j) = a* A $$ (i,l) + A $$ (i, j)" + using inotin assms by simp + then show ?thesis by simp + next + case False + then have iin: "j \ set ls" using assms + by (meson Cons.prems(3) set_ConsD) + have "add_col_to_multiple a (aa # ls) l A $$ (i, j) = + (addcol a aa l (add_col_to_multiple a ls l A)) $$ (i, j)" by simp + also have "... = ((add_col_to_multiple a ls l A) $$ (i,j))" + using False assms by auto + finally have "add_col_to_multiple a (aa # ls) l A $$ (i, j) = a * A $$ (i, l) + A $$ (i, j)" + using Cons.hyps by (metis Cons.prems(4) assms(1) assms(2) assms(6) distinct.simps(2) iin lnotin) + then show ?thesis by simp + qed +qed + +lemma add_col_to_multiple_index_change_bounds: + assumes "i < dim_row A" and "j < dim_col A" and "j \ low" and "j < up" and "l < low \ l \ up" + and "l < dim_col A" + shows "add_col_to_multiple a [low.. set [low.. set [low.. Operations specifically on 1st row/column \ + +lemma add_first_row_to_multiple_index: + assumes "i < dim_row M" and "j < dim_col M" + shows "i = 0 \ (add_row_to_multiple a [1.. 0 \ (add_row_to_multiple a [1.. 0 \ add_multiple_cols 1 0 [1.. add_multiple_cols 1 0 [1..l = 1..Lemmas on the determinant of the matrix under extended row/column operations \ + +lemma add_row_to_multiple_carrier: + "A \ carrier_mat n n \ add_row_to_multiple a ks l A \ carrier_mat n n" + by (metis add_row_to_multiple_dim(1) add_row_to_multiple_dim(2) carrier_matD(1) carrier_matD(2) carrier_matI) + +lemma add_col_to_multiple_carrier: + "A \ carrier_mat n n \ add_col_to_multiple a ks l A \ carrier_mat n n" + by (metis add_col_to_multiple_dim carrier_matD(1) carrier_matD(2) carrier_matI) + +lemma add_multiple_rows_carrier: + "A \ carrier_mat n n \ add_multiple_rows a k ls A \ carrier_mat n n" + by (metis add_multiple_rows_dim carrier_matD(1) carrier_matD(2) carrier_matI) + +lemma add_multiple_cols_carrier: + "A \ carrier_mat n n \ add_multiple_cols a k ls A \ carrier_mat n n" + by (metis add_multiple_cols_dim carrier_matD(1) carrier_matD(2) carrier_matI) + +lemma add_row_to_multiple_det: + assumes "l \ set ks" and "l < n" and "A \ carrier_mat n n" + shows "det (add_row_to_multiple a ks l A) = det A" + using assms +proof (induct ks) + case Nil + then show ?case by simp +next + case (Cons aa ks) + have ne: "aa \ l" + using Cons.prems(1) by auto + have "det (add_row_to_multiple a (aa # ks) l A) = det (addrow a aa l (add_row_to_multiple a ks l A))" + by simp + also have "... = det (add_row_to_multiple a ks l A)" + by (meson det_addrow add_row_to_multiple_carrier ne assms) + finally have "det (add_row_to_multiple a (aa # ks) l A) = det A" + using Cons.hyps assms by (metis Cons.prems(1) list.set_intros(2)) + then show ?case by simp +qed + +lemma add_col_to_multiple_det: + assumes "l \ set ks" and "l < n" and "A \ carrier_mat n n" + shows "det (add_col_to_multiple a ks l A) = det A" + using assms +proof (induct ks) + case Nil + then show ?case by simp +next + case (Cons aa ks) + have ne: "aa \ l" + using Cons.prems(1) by auto + have "det (add_col_to_multiple a (aa # ks) l A) = det (addcol a aa l (add_col_to_multiple a ks l A))" + by simp + also have "... = det (add_col_to_multiple a ks l A)" + by (meson det_addcol add_col_to_multiple_carrier ne assms) + finally have "det (add_col_to_multiple a (aa # ks) l A) = det A" + using Cons.hyps assms by (metis Cons.prems(1) list.set_intros(2)) + then show ?case by simp +qed + +lemma add_multiple_cols_det: + assumes "k \ set ls" and "\l. l \ set ls \ l < n" and "A \ carrier_mat n n" + shows "det (add_multiple_cols a k ls A) = det A" + using assms +proof (induct ls) + case Nil + then show ?case by simp +next + case (Cons aa ls) + have ne: "aa \ k" + using Cons.prems(1) by auto + have "det (add_multiple_cols a k (aa # ls) A) = det (addcol a k aa (add_multiple_cols a k ls A))" + by simp + also have "... = det (add_multiple_cols a k ls A)" + using det_addcol add_multiple_cols_carrier ne assms by (metis Cons.prems(2) list.set_intros(1)) + finally have "det (add_multiple_cols a k (aa # ls) A) = det A" + using Cons.hyps assms by (metis Cons.prems(1) Cons.prems(2) list.set_intros(2)) + then show ?case by simp +qed + +lemma add_multiple_rows_det: + assumes "k \ set ls" and "\l. l \ set ls \ l < n" and "A \ carrier_mat n n" + shows "det (add_multiple_rows a k ls A) = det A" + using assms +proof (induct ls) + case Nil + then show ?case by simp +next + case (Cons aa ls) + have ne: "aa \ k" + using Cons.prems(1) by auto + have "det (add_multiple_rows a k (aa # ls) A) = det (addrow a k aa (add_multiple_rows a k ls A))" + by simp + also have "... = det (add_multiple_rows a k ls A)" + using det_addrow add_multiple_rows_carrier ne assms by (metis Cons.prems(2) list.set_intros(1)) + finally have "det (add_multiple_rows a k (aa # ls) A) = det A" + using Cons.hyps assms by (metis Cons.prems(1) Cons.prems(2) list.set_intros(2)) + then show ?case by simp +qed + +subsection \Rank and Linear Independence\ + +abbreviation "rank v M \ vec_space.rank v M" + +text \Basic lemma: the rank of the multiplication of two matrices will be less than the minimum +of the individual ranks of those matrices. This directly follows from an existing lemmas in the +linear algebra library which show independently that the resulting matrices rank is less than either +the right or left matrix rank in the product \ +lemma rank_mat_mult_lt_min_rank_factor: + fixes A :: "'a::{conjugatable_ordered_field} mat" + assumes "A \ carrier_mat n m" + assumes "B \ carrier_mat m nc" + shows "rank n (A * B) \ min (rank n A) (rank m B)" +proof - + have 1: "rank n (A * B) \ (rank n A)" + using assms(1) assms(2) vec_space.rank_mat_mul_right by blast + have "rank n (A* B) \ rank m B" + by (meson assms(1) assms(2) rank_mat_mul_left) + thus ?thesis using 1 by simp +qed + +text \Rank Argument 1: Given two a $x \times y$ matrix $M$ where $MM^T$ has rank x, $x \le y$\ +lemma rank_argument: + fixes M :: "('c :: {conjugatable_ordered_field}) mat" + assumes "M \ carrier_mat x y" + assumes "vec_space.rank x (M* M\<^sup>T) = x" + shows "x \ y" +proof - + let ?B = "(M * M\<^sup>T)" + have Mt_car: "M\<^sup>T \ carrier_mat y x" using assms by simp + have b_car: "?B \ carrier_mat x x" + using transpose_carrier_mat assms by simp + then have "rank x ?B \ min (rank x M) (rank y M\<^sup>T)" + using rank_mat_mult_lt_min_rank_factor Mt_car b_car assms(1) by blast + thus ?thesis using le_trans vec_space.rank_le_nc + by (metis assms(1) assms(2) min.bounded_iff) +qed + + +text \Generalise the rank argument to use the determinant. If the determinant of the matrix +is non-zero, than it's rank must be equal to $x$. This removes the need for someone to use +facts on rank in their proofs. \ +lemma rank_argument_det: + fixes M :: "('c :: {conjugatable_ordered_field}) mat" + assumes "M \ carrier_mat x y" + assumes "det (M* M\<^sup>T) \ 0" + shows "x \ y" +proof - + let ?B = "(M * M\<^sup>T)" + have Mt_car: "M\<^sup>T \ carrier_mat y x" using assms by simp + have b_car: "?B \ carrier_mat x x" + using transpose_carrier_mat assms by simp + then have b_rank: "vec_space.rank x ?B = x" + using vec_space.low_rank_det_zero assms(2) by blast + then have "rank x ?B \ min (rank x M) (rank y M\<^sup>T)" + using rank_mat_mult_lt_min_rank_factor Mt_car b_car assms(1) by blast + thus ?thesis using le_trans vec_space.rank_le_nc + by (metis assms(1) b_rank min.bounded_iff) +qed + +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/Set_Multiset_Extras.thy b/thys/Fishers_Inequality/Set_Multiset_Extras.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Set_Multiset_Extras.thy @@ -0,0 +1,283 @@ +(* Title: Set_Multiset_Extras.thy + Author: Chelsea Edmonds +*) + +section \Micellaneous Multset/Set Extras \ + +theory Set_Multiset_Extras imports Design_Theory.Multisets_Extras "HOL-Combinatorics.Multiset_Permutations" +begin + +subsection \ Set extras \ +text \ Minor set extras on cardinality and filtering \ +lemma equal_card_inter_fin_eq_sets: "finite A \ finite B \ card A = card B \ + card (A \ B) = card A \ A = B" + by (metis Int_lower1 Int_lower2 card_subset_eq) + +lemma insert_filter_set_true: "P x \ {a \ (insert x A) . P a} = insert x {a \ A . P a}" + by auto + +lemma insert_filter_set_false: "\ P x \ {a \ (insert x A) . P a} = {a \ A . P a}" + by auto + + +subsection \Multiset Extras \ +text \ Minor multiset extras on size and element reasoning \ + +lemma obtain_two_items_mset: + assumes "size A > 1" + obtains x y where "x \# A" and "y \# A - {#x#}" +proof - + obtain x where "x \# A" + by (metis assms gr_implies_not_zero multiset_nonemptyE size_empty) + have "size (A - {#x#}) > 0" + by (metis \x \# A\ assms insert_DiffM less_irrefl_nat nonempty_has_size size_single) + then obtain bl2 where "bl2 \# A - {#x#}" + by (metis less_not_refl multiset_nonemptyE size_empty) + thus ?thesis + using \x \# A\ that by blast +qed + +lemma obtain_two_items_mset_filter: + assumes "size {#a \# A . P a #} > 1" + obtains x y where "x \# A" and "y \# A - {#x#}" and "P x" and "P y" +proof - + obtain x y where xin: "x \# {#a \# A . P a #}" and yin: "y \# {#a \# A . P a #} - {#x#}" + using obtain_two_items_mset assms by blast + then have xdets: "x \# A" "P x" by auto + then have yin2: "y \# {#a \# (A - {#x#}) . P a #}" using yin + by force + then have "y \# (A - {#x#})" "P y" + by (metis multiset_partition union_iff) (meson yin2 filter_mset_eq_conv) + thus ?thesis using xdets that by blast +qed + +lemma size_count_mset_ss: + assumes "finite B" + assumes "(set_mset A) \ B" + shows "size A = (\ x \ B . count A x)" +proof - + obtain C where cdef: "B - (set_mset A) = C" using assms + by simp + have fin: "finite (set_mset A)" using assms by auto + have un: "C \ (set_mset A) = B" + using Diff_partition \B - set_mset A = C\ assms by blast + have disj: "C \ (set_mset A) = {}" + using \B - set_mset A = C\ by auto + have zero: "\ x . x\ C \ count A x = 0" + by (meson count_eq_zero_iff disj disjoint_iff_not_equal) + then have "(\ x \ B . count A x) = (\ x \ (C \ set_mset A) . count A x)" using un by simp + also have "... = (\ x \ C . count A x) + (\ x \ (set_mset A) . count A x) " + using disj fin assms cdef sum.subset_diff by (metis un) + also have "... = (\ x \ (set_mset A) . count A x)" using zero by auto + finally have "(\ x \ B . count A x) = size A" + by (simp add: size_multiset_overloaded_eq) + thus ?thesis by simp +qed + +lemma mset_list_by_index: "mset (xs) = image_mset (\ i . (xs ! i)) (mset_set {.. x. x \#A \ a \ g x" + shows "count (image_mset (\x. if P x then a else g x) A ) a = size (filter_mset P A)" + using image_mset_If image_mset_filter_swap size_image_mset + by (smt (verit) assms count_size_set_repr filter_mset_cong) + +lemma count_mset_split_image_filter_not: + assumes "\ x. x \#A \ b \ f x" + shows "count (image_mset (\x. if P x then f x else b) A ) b = size (filter_mset (\ x. \ P x) A)" + using image_mset_If image_mset_filter_swap size_image_mset + by (smt (verit) assms count_size_set_repr filter_mset_cong) + +lemma removeAll_size_lt: "size (removeAll_mset C M) \ size M" + by (simp add: size_mset_mono) + +lemma mset_image_eq_filter_eq: "A = image_mset f B \ + filter_mset P A = (image_mset f (filter_mset (\ x. P (f x)) B))" + by (simp add: filter_mset_image_mset) + +subsection \Permutation on Sets and Multisets \ + +lemma elem_permutation_of_set_empty_iff: "finite A \ xs \ permutations_of_set A \ + xs = [] \ A = {}" + using permutations_of_setD(1) by fastforce + +lemma elem_permutation_of_mset_empty_iff: "xs \ permutations_of_multiset A \ xs = [] \ A = {#}" + using permutations_of_multisetD by fastforce + +subsection \ Lists \ +text \Further lemmas on the relationship between lists and multisets \ + +lemma count_distinct_mset_list_index: "i1 < length xs \ i2 < length xs \ i1 \ i2 \ + distinct_mset (mset xs) \ xs ! i1 \ xs ! i2" + by (simp add: nth_eq_iff_index_eq) + +lemma index_remove1_mset_ne: + assumes "x \# (mset xs)" + assumes "y \# remove1_mset x (mset xs)" + assumes "xs ! j1 = x" + assumes "j1 < length xs" + obtains j2 where "xs ! j2 = y" and "j2 < length xs" and "j1 \ j2" +proof (cases "x = y") + case True + then have "count (mset xs) x \ 2" + using assms(2) count_eq_zero_iff by fastforce + then have crm: "count (remove1_mset x (mset xs)) x \ 1" + by (metis True assms(2) count_greater_eq_one_iff) + obtain ys zs where xseq: "xs = ys @ (x # zs)" and yseq: "ys = take j1 xs" and zseq: "zs = drop (Suc j1) xs" + using assms(1) assms(3) id_take_nth_drop in_mset_conv_nth assms(4) by blast + have "mset xs = mset ys + mset (x # zs)" + by (simp add: xseq) + then have "remove1_mset x (mset xs) = mset (ys) + mset (zs)" + using assms by simp + then have "y \# (mset ys + mset zs)" using crm + using True \remove1_mset x (mset xs) = mset ys + mset zs\ assms(2) by presburger + then have yinor: "y \# mset ys \ y \# mset zs" by simp + then show ?thesis proof (cases "y \# mset ys") + case True + then obtain j2 where yeq: "ys ! j2 = y" and j2lt: "j2 < length ys" + by (meson in_mset_conv_nth) + then have 1: "xs ! j2 = y" using yseq by simp + have "j2 < j1" using yseq j2lt by simp + then show ?thesis using that 1 j2lt xseq by simp + next + case False + then have "y \# mset zs" using yinor by simp + then obtain j2 where zsy: "zs ! j2 = y" and j2lt: "j2 < length zs" + by (meson in_mset_conv_nth) + then have 1: "xs ! ((Suc j1) + j2) = y" using zseq zsy assms(4) by simp + have "length xs = (Suc j1) + length zs" using zseq xseq + by (metis Suc_diff_Suc add_Suc_shift add_diff_inverse_nat assms(4) length_drop less_imp_not_less) + then have 2: "(Suc j1) + j2 < length xs" using j2lt by simp + then show ?thesis using 1 that by simp + qed +next + case False + then show ?thesis + by (metis that assms(2) assms(3) in_diffD in_mset_conv_nth) +qed + +lemma count_list_mset: "count_list xs x = count (mset xs) x" +proof (induct xs) + case Nil + then show ?case by simp +next + case (Cons a xs) + then show ?case proof (cases "a = x") + case True + have mset_add_split: "count (mset (a # xs)) x = count (add_mset a (mset xs)) x" + by simp + then have "count (mset (a # xs)) x = count (mset xs) x + 1" + by (metis True Suc_eq_plus1 count_add_mset) + then show ?thesis using True Cons.hyps by simp + next + case False + then show ?thesis using Cons.hyps by simp + qed +qed + +lemma count_min_2_indices_lt: + assumes "i1 < i2" + assumes "xs ! i1 = x" + assumes "xs ! i2 = x" + assumes "i1 < length xs" "i2 < length xs" + shows "count (mset xs) x \ 2" +proof - + obtain xs1 xs2 where xse: "xs = xs1 @ xs2" and xs1: "xs1 = take i2 xs" and xs2: "xs2 = drop i2 xs" + by simp + have "i1 < length xs1" using assms length_take + by (simp add: assms(4) \xs1 = take i2 xs\) + then have xs1in: "xs ! i1 \# mset xs1" + by (simp add: nth_append xse) + have "i2 \ length xs1" using assms length_take xs1 by simp + then have xs2in: "xs ! i2 \# mset xs2" using xse nth_append + by (metis (no_types, lifting) assms(5) in_mset_conv_nth leD leI take_all_iff take_append) + have "count (mset xs) x = count ((mset xs1) + (mset xs2)) x" + by (simp add: xse) + then have "count (mset xs) x = count (mset xs1) x + count (mset xs2) x" by simp + thus ?thesis using xs1in xs2in + by (metis add_mono assms(2) assms(3) count_greater_eq_one_iff nat_1_add_1) +qed + +lemma count_min_2_indices: "i1 \ i2 \ xs ! i1 = x \ xs ! i2 = x \ i1 < length xs \ + i2 < length xs \ count (mset xs) x \ 2" + apply (cases "i1 < i2", simp add: count_min_2_indices_lt) + by (metis count_min_2_indices_lt linorder_cases) + +lemma obtain_set_list_item: + assumes "x \ set xs" + obtains i where "i < length xs" and "xs ! i = x" + by (meson assms in_set_conv_nth) + +subsection \Summation Rules\ + +text \ Some lemmas to make it simpler to work with double and triple summations \ +context comm_monoid_add +begin + +lemma sum_reorder_triple: "(\ l \ A . (\ i \ B . (\ j \ C . g l i j))) = + (\ i \ B . (\ j \ C . (\ l \ A . g l i j)))" +proof - + have "(\ l \ A . (\ i \ B . (\ j \ C . g l i j))) = (\i \ B . (\ l \ A . (\ j \ C . g l i j)))" + using sum.swap[of "(\ l i . (\ j \ C . g l i j))" "B" "A"] by simp + also have "... = (\i \ B . (\ j \ C . (\l \ A . g l i j)))" using sum.swap by metis + finally show ?thesis by simp +qed + +lemma double_sum_mult_hom: + fixes k :: "'b :: {comm_ring_1}" + shows "(\ i \ A . (\ j \ g i . k * (f i j))) = k* (\ i \ A . (\ j \ g i . f i j))" + by (metis (mono_tags, lifting) comm_monoid_add_class.sum.cong sum_distrib_left) + +lemma double_sum_split_case: + assumes "finite A" + shows "(\ i \ A . (\ j \ A . f i j)) = (\ i \ A . (f i i)) + (\ i \ A . (\ j \ (A - {i}) . f i j))" +proof - + have "\ i. i \ A \ (\ j \ A . f i j) = f i i + (\ j \ (A - {i}) . f i j)" + using sum.remove assms by metis + then show ?thesis by (simp add: sum.distrib) +qed + +lemma double_sum_split_case2: "(\ i \ A . (\ j \ A . g i j)) = + (\ i \ A . (g i i)) + (\ i \ A . (\ j \ {a \ A . a \ i} . g i j)) " +proof - + have "\ i. A = {a \ A . a = i} \ {a \ A . a \ i}" by auto + then have part: "\ i. i \ A \ A = {i} \ {a \ A . a \ i}" by auto + have empt:"\ i. {i} \ {a \ A . a \ i} = {}" + by simp + then have "(\ i \ A . (\ j \ A . g i j)) = + (\ i \ A . ((\ j \ {i} . g i j) + (\ j \ {a \ A . a \ i} . g i j)))" using part + by (smt (verit) finite_Un local.sum.cong local.sum.infinite local.sum.union_disjoint) + also have "... = (\ i \ A . ((\ j \ {i} . g i j))) + (\ i \ A . (\ j \ {a \ A . a \ i} . g i j))" + by (simp add: local.sum.distrib) + finally show ?thesis by simp +qed + +end + +context comm_ring_1 +begin + +lemma double_sum_split_case_square: + assumes "finite A" + shows "(\ i \ A . f i )^2 = (\ i \ A . (f i * f i)) + (\ i \ A . (\ j \ (A - {i}) . f i * f j))" +proof - + have "(\ i \ A . f i )^2 = (\ i \ A . f i) * (\ i \ A . f i)" + using power2_eq_square by blast + then have "(\ i \ A . f i) * (\ i \ A . f i) = (\ i \ A . f i) * (\ j \ A . f j)" by simp + also have 1: "... = (\ i \ A . f i * (\ j \ A . f j))" using sum_distrib_right by simp + also have 2: "... = (\ i \ A . (\ j \ A . f i * f j))" using sum_distrib_left by metis + finally have "(\ i \ A . f i) * (\ i \ A . f i) = + (\ i \ A . (f i * f i)) + (\ i \ A . (\ j \ (A - {i}) . f i * f j))" + using assms double_sum_split_case[of "A" "\ i j . f i * f j"] 1 2 by presburger + then show ?thesis + using power2_eq_square by presburger +qed + +lemma double_sum_split_square_diff: "finite {0.. + (\ i \ {0.. j \ ({0..< x} - {i}) . c i * c j)) = + (\ i \ {0.. i \ {0.. i. c i"] by fastforce + +end +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/Vector_Matrix_Mod.thy b/thys/Fishers_Inequality/Vector_Matrix_Mod.thy new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/Vector_Matrix_Mod.thy @@ -0,0 +1,533 @@ +(* Title: Vector_Matrix_Mod.thy + Author: Chelsea Edmonds +*) + +section \ Matrices/Vectors mod x \ +text \ This section formalises operations and properties mod some integer x on integer matrices and +vectors. Much of this file was no longer needed for fishers once the generic idea of lifting a 0-1 matrix +was introduced, however it is left as an example for future work on matrices mod n, beyond 0-1 matrices \ +theory Vector_Matrix_Mod imports Matrix_Vector_Extras +Berlekamp_Zassenhaus.Finite_Field Berlekamp_Zassenhaus.More_Missing_Multiset +begin + +text \Simple abbreviations for main mapping functions \ +abbreviation to_int_mat :: "'a :: finite mod_ring mat \ int mat" where + "to_int_mat \ map_mat to_int_mod_ring" + +abbreviation to_int_vec :: "'a :: finite mod_ring vec \ int vec" where +"to_int_vec \ map_vec to_int_mod_ring" + +interpretation of_int_mod_ring_hom_sr: semiring_hom of_int_mod_ring +proof (unfold_locales) + show "\x y. of_int_mod_ring (x + y) = of_int_mod_ring x + of_int_mod_ring y" + by (transfer,presburger) + show "of_int_mod_ring 1 = 1" by (metis of_int_hom.hom_one of_int_of_int_mod_ring) + show "\x y. of_int_mod_ring (x * y) = of_int_mod_ring x * of_int_mod_ring y" + by (transfer, simp add: mod_mult_eq) +qed + +text \NOTE: The context directly below is copied from Matrix Vector Extras, as for some reason +they can't be used locally if not? Ideally remove in future if possible \ + +context inj_zero_hom +begin + +lemma vec_hom_zero_iff[simp]: "(map_vec hom x = 0\<^sub>v n) = (x = 0\<^sub>v n)" +proof - + { + fix i + assume i: "i < n" "dim_vec x = n" + hence "map_vec hom x $ i = 0 \ x $ i = 0" + using index_map_vec(1)[of i x] by simp + } note main = this + show ?thesis unfolding vec_eq_iff by (simp, insert main, auto) +qed + +lemma mat_hom_inj: "map_mat hom A = map_mat hom B \ A = B" + unfolding mat_eq_iff by auto + +lemma vec_hom_inj: "map_vec hom v = map_vec hom w \ v = w" + unfolding vec_eq_iff by auto + +lemma vec_hom_set_distinct_iff: + fixes xs :: "'a vec list" + shows "distinct xs \ distinct (map (map_vec hom) xs)" + using vec_hom_inj by (induct xs) (auto) +end + +subsection \ Basic Mod Context \ + +locale mat_mod = fixes m :: int +assumes non_triv_m: "m > 1" +begin + +text \First define the mod operations on vectors \ +definition vec_mod :: "int vec \ int vec" where +"vec_mod v \ map_vec (\ x . x mod m) v" + +(* Parse tree ambiguity is caused by bad definitions in the MPoly theory files. Issue raised +on Isabelle Mailing List *) + +lemma vec_mod_dim[simp]: "dim_vec (vec_mod v) = dim_vec v" + using vec_mod_def by simp + +lemma vec_mod_index[simp]: "i < dim_vec v \ (vec_mod v) $ i = (v $ i) mod m" + using vec_mod_def by simp + +lemma vec_mod_index_same[simp]: "i < dim_vec v \ v $ i < m \ v $ i \ 0 \ (vec_mod v) $ i = v $ i" + by simp + +lemma vec_setI2: "i < dim_vec v \ v $ i \ set\<^sub>v v" + by (simp add: vec_setI) + +lemma vec_mod_eq: "set\<^sub>v v \ {0.. vec_mod v = v" + apply (intro eq_vecI, simp_all) + using vec_setI2 vec_mod_index_same by (metis atLeastLessThan_iff subset_iff zmod_trivial_iff) + +lemma vec_mod_set_vec_same:"set\<^sub>v v \ {0.. set\<^sub>v (vec_mod v) = set\<^sub>v v" + using vec_mod_eq by auto + +text \Define the mod operation on matrices \ + +definition mat_mod :: "int mat \ int mat" where +"mat_mod M \ map_mat (\ x. x mod m) M" + +lemma mat_mod_dim[simp]: "dim_row (mat_mod M) = dim_row M" "dim_col (mat_mod M) = dim_col M" + using mat_mod_def by simp_all + +lemma mat_mod_index [simp]: "i < dim_row M \ j < dim_col M \ (mat_mod M) $$ (i, j) = (M $$ (i, j)) mod m" + by(simp add: mat_mod_def) + +lemma mat_mod_index_same[simp]: "i < dim_row M \ j < dim_col M \ M $$ (i, j) < m \ + M $$ (i, j) \ 0 \ mat_mod M $$ (i, j) = M $$ (i, j)" + by (simp add: mat_mod_def) + +lemma elements_matI2: "i < dim_row A \ j < dim_col A \ A $$ (i, j) \ elements_mat A" + by auto + +lemma mat_mod_elements_in: + assumes "x \ elements_mat M" + shows "x mod m \ elements_mat (mat_mod M)" +proof - + obtain i j where "M $$ (i, j) = x" and ilt: "i < dim_row M" and jlt: "j < dim_col M" using assms by auto + then have "mat_mod M $$ (i, j) = x mod m" by simp + thus ?thesis using ilt jlt + by (metis elements_matI2 mat_mod_dim(1) mat_mod_dim(2)) +qed + +lemma mat_mod_elements_map: + assumes "x \ elements_mat M" + shows "elements_mat (mat_mod M) = (\ x. x mod m) ` (elements_mat M)" +proof (auto simp add: mat_mod_elements_in) + fix x assume "x \ elements_mat (local.mat_mod M)" + then obtain i j where "(mat_mod M) $$ (i, j) = x" and "i < dim_row (mat_mod M)" and "j < dim_col (mat_mod M)" by auto + then show "x \ (\x. x mod m) ` elements_mat M" + by auto +qed + +lemma mat_mod_eq_cond: + assumes "elements_mat M \ {0.. {0.. {0.. elements_mat (mat_mod M) = elements_mat M" + using mat_mod_eq_cond by auto + +lemma mat_mod_vec_mod_row: "i < dim_row A \ row (mat_mod A) i = vec_mod (row A i)" + unfolding mat_mod_def vec_mod_def by (simp) + +lemma mat_mod_vec_mod_col: "j < dim_col A \ col (mat_mod A) j = vec_mod (col A j)" + unfolding mat_mod_def vec_mod_def by (simp) + +lemma count_vec_mod_eq: "set\<^sub>v v \ {0.. count_vec v x = count_vec (vec_mod v) x" + using vec_mod_eq by (simp) + +lemma elems_mat_setv_row_0m: "i < dim_row M \ elements_mat M \ {0.. set\<^sub>v (row M i) \ {0.. elements_mat M \ {0.. set\<^sub>v (col M j) \ {0.. elements_mat M \ {0.. + count_vec (row (mat_mod M) i) x = count_vec (row M i) x" + using count_vec_mod_eq mat_mod_vec_mod_row elems_mat_setv_row_0m by simp + +lemma mat_mod_count_col_eq: "j < dim_col M \ elements_mat M \ {0.. + count_vec (col (mat_mod M) j) x = count_vec (col M j) x" + using count_vec_mod_eq mat_mod_vec_mod_col elems_mat_setv_col_0m by simp + +lemma mod_mat_one: "mat_mod (1\<^sub>m n) = (1\<^sub>m n)" + by (intro eq_matI, simp_all add: mat_mod_def non_triv_m) + +lemma mod_mat_zero: "mat_mod (0\<^sub>m nr nc) = (0\<^sub>m nr nc)" + by (intro eq_matI, simp_all add: mat_mod_def non_triv_m) + +lemma vec_mod_unit: "vec_mod (unit_vec n i) = (unit_vec n i)" + by (intro eq_vecI, simp_all add: unit_vec_def vec_mod_def non_triv_m) + +lemma vec_mod_zero: "vec_mod (0\<^sub>v n) = (0\<^sub>v n)" + by (intro eq_vecI, simp_all add: non_triv_m) + +lemma mat_mod_cond_iff: "elements_mat M \ {0.. P M \ P (mat_mod M)" + by (simp add: mat_mod_eq_cond) + +end + +subsection \Mod Type \ +text \ The below locale takes lemmas from the Poly Mod Finite Field theory in the Berlekamp Zassenhaus +AFP entry, however has removed any excess material on polynomials mod, and only included the general +factors. Ideally, this could be used as the base locale for both in the future \ + +locale mod_type = + fixes m :: int and ty :: "'a :: nontriv itself" + assumes m: "m = CARD('a)" +begin + +lemma m1: "m > 1" using nontriv[where 'a = 'a] by (auto simp:m) + +definition M :: "int \ int" where "M x = x mod m" + +lemma M_0[simp]: "M 0 = 0" + by (auto simp add: M_def) + +lemma M_M[simp]: "M (M x) = M x" + by (auto simp add: M_def) + +lemma M_plus[simp]: "M (M x + y) = M (x + y)" "M (x + M y) = M (x + y)" + by (auto simp add: M_def mod_simps) + +lemma M_minus[simp]: "M (M x - y) = M (x - y)" "M (x - M y) = M (x - y)" + by (auto simp add: M_def mod_simps) + +lemma M_times[simp]: "M (M x * y) = M (x * y)" "M (x * M y) = M (x * y)" + by (auto simp add: M_def mod_simps) + +lemma M_1[simp]: "M 1 = 1" unfolding M_def + using m1 by auto + +lemma M_sum: "M (sum (\ x. M (f x)) A) = M (sum f A)" +proof (induct A rule: infinite_finite_induct) + case (insert x A) + from insert(1-2) have "M (\x\insert x A. M (f x)) = M (f x + M ((\x\A. M (f x))))" by simp + also have "M ((\x\A. M (f x))) = M ((\x\A. f x))" using insert by simp + finally show ?case using insert by simp +qed auto + +definition inv_M :: "int \ int" where + "inv_M x = (if x + x \ m then x else x - m)" + +lemma M_inv_M_id[simp]: "M (inv_M x) = M x" + unfolding inv_M_def M_def by simp + +definition M_Rel :: "int \ 'a mod_ring \ bool" + where "M_Rel x x' \ (M x = to_int_mod_ring x')" + +lemma to_int_mod_ring_plus: "to_int_mod_ring ((x :: 'a mod_ring) + y) = M (to_int_mod_ring x + to_int_mod_ring y)" + unfolding M_def using m by (transfer, auto) + +lemma to_int_mod_ring_times: "to_int_mod_ring ((x :: 'a mod_ring) * y) = M (to_int_mod_ring x * to_int_mod_ring y)" + unfolding M_def using m by (transfer, auto) + +lemma eq_M_Rel[transfer_rule]: "(M_Rel ===> M_Rel ===> (=)) (\ x y. M x = M y) (=)" + unfolding M_Rel_def rel_fun_def by auto + +lemma one_M_Rel[transfer_rule]: "M_Rel 1 1" + unfolding M_Rel_def M_def + unfolding m by auto + +lemma zero_M_Rel[transfer_rule]: "M_Rel 0 0" + unfolding M_Rel_def M_def + unfolding m by auto + +lemma M_to_int_mod_ring: "M (to_int_mod_ring (x :: 'a mod_ring)) = to_int_mod_ring x" + unfolding M_def unfolding m by (transfer, auto) + + +lemma right_total_M_Rel[transfer_rule]: "right_total M_Rel" + unfolding right_total_def M_Rel_def using M_to_int_mod_ring by blast + +lemma left_total_M_Rel[transfer_rule]: "left_total M_Rel" + unfolding left_total_def M_Rel_def[abs_def] +proof + fix x + show "\ x' :: 'a mod_ring. M x = to_int_mod_ring x'" unfolding M_def unfolding m + by (rule exI[of _ "of_int x"], transfer, simp) +qed + +lemma bi_total_M_Rel[transfer_rule]: "bi_total M_Rel" + using right_total_M_Rel left_total_M_Rel by (metis bi_totalI) + +lemma to_int_mod_ring_of_int_M: "to_int_mod_ring (of_int x :: 'a mod_ring) = M x" unfolding M_def + unfolding m by transfer auto + +lemma UNIV_M_Rel[transfer_rule]: "rel_set M_Rel {0.. Mat mod type \ +text \ Define a context to work on matrices and vectors of type @{typ "'a mod_ring"} \ + +locale mat_mod_type = mat_mod + mod_type +begin + +lemma to_int_mod_ring_plus: "to_int_mod_ring ((x :: 'a mod_ring) + y) = (to_int_mod_ring x + to_int_mod_ring y) mod m" + using m by (transfer, auto) + +lemma to_int_mod_ring_times: "to_int_mod_ring ((x :: 'a mod_ring) * y) = (to_int_mod_ring x * to_int_mod_ring y) mod m" + using m by (transfer, auto) + +text \ Set up transfer relation for matrices and vectors \ +definition MM_Rel :: "int mat \ 'a mod_ring mat \ bool" + where "MM_Rel f f' \ (mat_mod f = to_int_mat f')" + +definition MV_Rel :: "int vec \ 'a mod_ring vec \ bool" + where "MV_Rel v v' \ (vec_mod v = to_int_vec v')" + +lemma to_int_mat_index[simp]: "i < dim_row N \ j < dim_col N \ (to_int_mat N $$ (i, j)) = to_int_mod_ring (N $$ (i, j))" + by simp + +lemma to_int_vec_index[simp]: "i < dim_vec v \ (to_int_vec v $i) = to_int_mod_ring (v $i)" + by simp + +lemma eq_dim_row_MM_Rel[transfer_rule]: "(MM_Rel ===> (=)) dim_row dim_row " + by (metis (mono_tags) MM_Rel_def index_map_mat(2) mat_mod_dim(1) rel_funI) + +lemma lt_dim_row_MM_Rel[transfer_rule]: "(MM_Rel ===> (=) ===> (=)) (\ M i. i < dim_row M) (\ M i. i < dim_row M)" + using eq_dim_row_MM_Rel unfolding MM_Rel_def rel_fun_def by auto + +lemma eq_dim_col_MM_Rel[transfer_rule]: "(MM_Rel ===> (=)) dim_col dim_col " + unfolding MM_Rel_def rel_fun_def + by (metis index_map_mat(3) mat_mod_dim(2)) + +lemma lt_dim_col_MM_Rel[transfer_rule]: "(MM_Rel ===> (=) ===> (=)) (\ M j. j < dim_col M) (\ M j. j < dim_col M)" + using eq_dim_col_MM_Rel unfolding MM_Rel_def rel_fun_def by auto + +lemma eq_dim_vec_MV_Rel[transfer_rule]: "(MV_Rel ===> (=)) dim_vec dim_vec" + unfolding MV_Rel_def rel_fun_def using index_map_vec(2) vec_mod_dim by metis + +lemma lt_dim_vec_MV_Rel[transfer_rule]: "(MV_Rel ===> (=) ===> (=)) (\ v j. j < dim_vec v) (\ v j. j < dim_vec v)" + unfolding MV_Rel_def rel_fun_def using index_map_vec(2) vec_mod_dim by metis + +lemma eq_MM_Rel[transfer_rule]: "(MM_Rel ===> MM_Rel ===> (=)) (\ f f' . mat_mod f = mat_mod f') (=) " + unfolding MM_Rel_def rel_fun_def using to_int_mod_ring_hom.mat_hom_inj by (auto) + +lemma eq_MV_Rel[transfer_rule]: "(MV_Rel ===> MV_Rel ===> (=)) (\ v v' . vec_mod v = vec_mod v') (=) " + unfolding MV_Rel_def rel_fun_def using to_int_mod_ring_hom.vec_hom_inj by auto + + +lemma index_MV_Rel[transfer_rule]: "(MV_Rel ===> (=) ===> M_Rel) + (\ v i. if i < dim_vec v then v $ i else 0) (\ v i. if i < dim_vec v then v $ i else 0)" + using lt_dim_vec_MV_Rel unfolding MV_Rel_def M_Rel_def M_def rel_fun_def + by (simp, metis to_int_vec_index vec_mod_index) + +lemma index_MM_Rel[transfer_rule]: "(MM_Rel ===> (=) ===> (=) ===> M_Rel) + (\ M i j. if (i < dim_row M \ j < dim_col M) then M $$ (i, j) else 0) + (\ M i j. if (i < dim_row M \ j < dim_col M) then M $$ (i, j) else 0)" + using lt_dim_row_MM_Rel lt_dim_col_MM_Rel unfolding M_Rel_def M_def rel_fun_def + by (simp, metis mat_mod_index to_int_mat_index MM_Rel_def) + +lemma index_MM_Rel_explicit: + assumes "MM_Rel N N'" + assumes "i < dim_row N" "j < dim_col N" + shows "(N $$ (i, j)) mod m = to_int_mod_ring (N' $$ (i, j))" +proof - + have eq: "(to_int_mat N') $$ (i, j) = to_int_mod_ring (N' $$ (i, j))" + by (metis MM_Rel_def assms(1) assms(2) assms(3) index_map_mat mat_mod.mat_mod_dim mat_mod_axioms) + have "mat_mod N = to_int_mat N'" using assms by (simp add: MM_Rel_def) + then have "(mat_mod N) $$ (i, j) = (to_int_mat N') $$ (i, j)" + by simp + thus ?thesis using mat_mod_index eq + using assms(2) assms(3) by auto +qed + +lemma one_MV_Rel[transfer_rule]: "MV_Rel (unit_vec n i) (unit_vec n i)" + unfolding MV_Rel_def vec_mod_unit non_triv_m unit_vec_def + by (intro eq_vecI, simp_all add: non_triv_m) + +lemma one_MM_Rel[transfer_rule]: "MM_Rel (1\<^sub>m n) (1\<^sub>m n)" + unfolding MM_Rel_def mod_mat_one + by (intro eq_matI, simp_all) + +lemma zero_MM_Rel[transfer_rule]: "MM_Rel (0\<^sub>m nr nc) (0\<^sub>m nr nc)" + unfolding MM_Rel_def + by (intro eq_matI, simp_all) + +lemma zero_MV_Rel[transfer_rule]: "MV_Rel (0\<^sub>v n) (0\<^sub>v n)" + unfolding MV_Rel_def by (intro eq_vecI, simp_all) + +lemma right_unique_MV_Rel[transfer_rule]: "right_unique MV_Rel" + unfolding right_unique_def MV_Rel_def + using to_int_mod_ring_hom.vec_hom_inj by auto + +lemma right_unique_MM_Rel[transfer_rule]: "right_unique MM_Rel" + unfolding right_unique_def MM_Rel_def + using to_int_mod_ring_hom.mat_hom_inj by auto + +lemma mod_to_int_mod_ring: "(to_int_mod_ring (x :: 'a mod_ring)) mod m = to_int_mod_ring x" + unfolding m by (transfer, auto) + +lemma mat_mod_to_int_mat: "mat_mod (to_int_mat (N :: 'a mod_ring mat)) = to_int_mat N" + using mod_to_int_mod_ring by (intro eq_matI, simp_all) + +lemma vec_mod_to_int_vec: "vec_mod (to_int_vec (v :: 'a mod_ring vec)) = to_int_vec v" + using mod_to_int_mod_ring by (intro eq_vecI, simp_all) + +lemma right_total_MM_Rel[transfer_rule]: "right_total MM_Rel" + unfolding right_total_def MM_Rel_def +proof + fix M :: "'a mod_ring mat" + show "\x. mat_mod x = to_int_mat M" + by (intro exI[of _ "to_int_mat M"], simp add: mat_mod_to_int_mat) +qed + +lemma right_total_MV_Rel[transfer_rule]: "right_total MV_Rel" + unfolding right_total_def MV_Rel_def +proof + fix v :: "'a mod_ring vec" + show "\x. vec_mod x = to_int_vec v" + by (intro exI[of _ "to_int_vec v"], simp add: vec_mod_to_int_vec) +qed + +lemma to_int_mod_ring_of_int_mod: "to_int_mod_ring (of_int x :: 'a mod_ring) = x mod m" + unfolding m by transfer auto + +lemma vec_mod_v_representative: "vec_mod v = to_int_vec (map_vec of_int v :: 'a mod_ring vec)" + unfolding mat_mod_def by (auto simp: to_int_mod_ring_of_int_mod) + +lemma mat_mod_N_representative: "mat_mod N = to_int_mat (map_mat of_int N :: 'a mod_ring mat)" + unfolding mat_mod_def by (auto simp: to_int_mod_ring_of_int_mod) + +lemma left_total_MV_Rel[transfer_rule]: "left_total MV_Rel" + unfolding left_total_def MV_Rel_def[abs_def] using vec_mod_v_representative by blast + +lemma left_total_MM_Rel[transfer_rule]: "left_total MM_Rel" + unfolding left_total_def MM_Rel_def[abs_def] using mat_mod_N_representative by blast + +lemma bi_total_MV_Rel[transfer_rule]: "bi_total MV_Rel" + using right_total_MV_Rel left_total_MV_Rel by (metis bi_totalI) + +lemma bi_total_MM_Rel[transfer_rule]: "bi_total MM_Rel" + using right_total_MM_Rel left_total_MM_Rel by (metis bi_totalI) + +lemma domain_MV_rel[transfer_domain_rule]: "Domainp MV_Rel = (\ f. True)" +proof + fix v :: "int vec" + show "Domainp MV_Rel v = True" unfolding MV_Rel_def[abs_def] Domainp.simps + by (auto simp: vec_mod_v_representative) +qed + +lemma domain_MM_rel[transfer_domain_rule]: "Domainp MM_Rel = (\ f. True)" +proof + fix N :: "int mat" + show "Domainp MM_Rel N = True" unfolding MM_Rel_def[abs_def] Domainp.simps + by (auto simp: mat_mod_N_representative) +qed + +lemma mem_MV_Rel[transfer_rule]: + "(MV_Rel ===> rel_set MV_Rel ===> (=)) (\ x Y. \y \ Y. vec_mod x = vec_mod y) (\)" +proof (intro rel_funI iffI) + fix x y X Y assume xy: "MV_Rel x y" and XY: "rel_set MV_Rel X Y" + { assume "\x' \ X. vec_mod x = vec_mod x'" + then obtain x' where x'X: "x' \ X" and xx': "vec_mod x = vec_mod x'" by auto + with xy have x'y: "MV_Rel x' y" by (auto simp: MV_Rel_def) + from rel_setD1[OF XY x'X] obtain y' where "MV_Rel x' y'" and "y' \ Y" by auto + with x'y + show "y \ Y" using to_int_mod_ring_hom.vec_hom_inj by (auto simp: MV_Rel_def) + } + assume "y \ Y" + from rel_setD2[OF XY this] obtain x' where x'X: "x' \ X" and x'y: "MV_Rel x' y" by auto + from xy x'y have "vec_mod x = vec_mod x'" by (auto simp: MV_Rel_def) + with x'X show "\x' \ X. vec_mod x = vec_mod x'" by auto +qed + +lemma mem_MM_Rel[transfer_rule]: + "(MM_Rel ===> rel_set MM_Rel ===> (=)) (\ x Y. \y \ Y. mat_mod x = mat_mod y) (\)" +proof (intro rel_funI iffI) + fix x y X Y assume xy: "MM_Rel x y" and XY: "rel_set MM_Rel X Y" + { assume "\x' \ X. mat_mod x = mat_mod x'" + then obtain x' where x'X: "x' \ X" and xx': "mat_mod x = mat_mod x'" by auto + with xy have x'y: "MM_Rel x' y" by (auto simp: MM_Rel_def) + from rel_setD1[OF XY x'X] obtain y' where "MM_Rel x' y'" and "y' \ Y" by auto + with x'y + show "y \ Y" using to_int_mod_ring_hom.mat_hom_inj by (auto simp: MM_Rel_def) + } + assume "y \ Y" + from rel_setD2[OF XY this] obtain x' where x'X: "x' \ X" and x'y: "MM_Rel x' y" by auto + from xy x'y have "mat_mod x = mat_mod x'" by (auto simp: MM_Rel_def) + with x'X show "\x' \ X. mat_mod x = mat_mod x'" by auto +qed + +lemma conversep_MM_Rel_OO_MM_Rel [simp]: "MM_Rel\\ OO MM_Rel = (=)" + using mat_mod_to_int_mat apply (intro ext, auto simp: OO_def MM_Rel_def) + using to_int_mod_ring_hom.mat_hom_inj by auto + +lemma MM_Rel_OO_conversep_MM_Rel [simp]: "MM_Rel OO MM_Rel\\ = (\ M M' . mat_mod M = mat_mod M')" + by (intro ext, auto simp: OO_def MM_Rel_def mat_mod_N_representative) + +lemma conversep_MM_Rel_OO_eq_m [simp]: "MM_Rel\\ OO (\ M M' . mat_mod M = mat_mod M') = MM_Rel\\" + by (intro ext, auto simp: OO_def MM_Rel_def) + +lemma eq_m_OO_MM_Rel [simp]: "(\ M M' . mat_mod M = mat_mod M') OO MM_Rel = MM_Rel" + by (intro ext, auto simp: OO_def MM_Rel_def) + +lemma eq_mset_MM_Rel [transfer_rule]: + "(rel_mset MM_Rel ===> rel_mset MM_Rel ===> (=)) (rel_mset (\ M M' . mat_mod M = mat_mod M')) (=)" +proof (intro rel_funI iffI) + fix A B X Y + assume AX: "rel_mset MM_Rel A X" and BY: "rel_mset MM_Rel B Y" + { + assume AB: "rel_mset (\ M M' . mat_mod M = mat_mod M') A B" + from AX have "rel_mset MM_Rel\\ X A" by (simp add: multiset.rel_flip) + note rel_mset_OO[OF this AB] + note rel_mset_OO[OF this BY] + then show "X = Y" by (simp add: multiset.rel_eq) + } + assume "X = Y" + with BY have "rel_mset MM_Rel\\ X B" by (simp add: multiset.rel_flip) + from rel_mset_OO[OF AX this] + show "rel_mset (\ M M' . mat_mod M = mat_mod M') A B" by simp +qed + +lemma vec_mset_MV_Rel[transfer_rule]: + "(MV_Rel ===> (=)) (\ v. vec_mset (vec_mod v)) (\ v. image_mset (to_int_mod_ring) (vec_mset v))" + unfolding MV_Rel_def rel_fun_def +proof (intro allI impI subset_antisym subsetI) + fix x :: "int vec" fix y :: "'a mod_ring vec" + assume assm: "vec_mod x = to_int_vec y" + have "image_mset to_int_mod_ring (vec_mset y) = vec_mset (to_int_vec y)" + using inj_zero_hom.vec_hom_mset to_int_mod_ring_hom.inj_zero_hom_axioms by auto + then show " vec_mset (vec_mod x) = image_mset to_int_mod_ring (vec_mset y)" using assm by simp +qed + +lemma vec_count_MV_Rel_direct: + assumes "MV_Rel v1 v2" + shows "count_vec v2 i = count_vec (vec_mod v1) (to_int_mod_ring i)" +proof- + have eq_vecs: "to_int_vec v2 = vec_mod v1" using assms unfolding MV_Rel_def by simp + have "count_vec v2 i = count (vec_mset v2) i" by simp + also have 1: "... = count (image_mset to_int_mod_ring (vec_mset v2)) (to_int_mod_ring i)" + using count_image_mset_inj by (metis to_int_mod_ring_hom.inj_f) + also have 2: "... = count (vec_mset (vec_mod v1)) (to_int_mod_ring i)" using assms + by (simp add: eq_vecs inj_zero_hom.vec_hom_mset to_int_mod_ring_hom.inj_zero_hom_axioms) + finally show "count_vec v2 i = count_vec (vec_mod v1) (to_int_mod_ring i)" + by (simp add: 1 2 ) +qed + +lemma MM_Rel_MV_Rel_row: "MM_Rel A B \ i < dim_row A \ MV_Rel (row A i) (row B i)" + unfolding MM_Rel_def MV_Rel_def + by (metis index_map_mat(2) mat_mod_dim(1) mat_mod_vec_mod_row row_map_mat) + +lemma MM_Rel_MV_Rel_col: "MM_Rel A B \ j < dim_col A \ MV_Rel (col A j) (col B j)" + unfolding MM_Rel_def MV_Rel_def + using index_map_mat(3) mat_mod_dim(2) mat_mod_vec_mod_col col_map_mat by (metis) + +end +end \ No newline at end of file diff --git a/thys/Fishers_Inequality/document/root.bib b/thys/Fishers_Inequality/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/document/root.bib @@ -0,0 +1,103 @@ +@book{cameronCombinatoricsTopicsTechniques1994, + abstract = {Cambridge Core - Algorithmics, Complexity, Computer Algebra, Computational Geometry - Combinatorics - by Peter J. Cameron}, + address = {{Cambridge}}, + author = {Cameron, Peter J.}, + file = {C\:\\Users\\cledm\\Zotero\\storage\\V9GWRW39\\951A2163C96B61B09140F054E021A9FE.html}, + isbn = {978-0-521-45133-8 978-0-511-80388-8 978-0-521-45761-3}, + language = {en}, + month = oct, + note = {\url{/core/books/combinatorics/951A2163C96B61B09140F054E021A9FE}}, + publisher = {{Cambridge University Press}}, + shorttitle = {Combinatorics}, + title = {Combinatorics: {{Topics}}, {{Techniques}}, {{Algorithms}}}, + year = {1994}} + +@book{colbournHandbookCombinatorialDesigns2007, + author = {Colbourn, C. J and Dinitz, Jeffrey H.}, + date-modified = {2021-03-22 15:57:21 +0000}, + edition = {2nd}, + file = {C\:\\Users\\cledm\\Zotero\\storage\\ALJQH3YJ\\C. J Colbourn_Dinitz_2007_Handbook of combinatorial designs - edited by Charles J.pdf}, + isbn = {978-1-58488-506-1}, + keywords = {Combinatorial designs and configurations}, + language = {eng}, + lccn = {QA166.25 .H363 2007}, + publisher = {{Chapman \& Hall/CRC}}, + title = {Handbook of Combinatorial Designs / Edited by {{Charles J}}. {{Colbourn}}, {{Jeffrey H}}. {{Dinitz}}.}, + year = {2007}} + +@book{stinsonCombinatorialDesignsConstructions2004, + abstract = {Created to teach students many of the most important techniques used for constructing combinatorial designs, this is an ideal textbook for advanced undergraduate and graduate courses in combinatorial design theory. The text features clear explanations of basic designs, such as Steiner and Kirkman triple systems, mutual orthogonal Latin squares, finite projective and affine planes, and Steiner quadruple systems. In these settings, the student will master various construction techniques, both classic and modern, and will be well-prepared to construct a vast array of combinatorial designs. Design theory offers a progressive approach to the subject, with carefully ordered results. It begins with simple constructions that gradually increase in complexity. Each design has a construction that contains new ideas or that reinforces and builds upon similar ideas previously introduced. A new text/reference covering all apsects of modern combinatorial design theory. Graduates and professionals in computer science, applied mathematics, combinatorics, and applied statistics will find the book an essential resource.}, + annotation = {https://www.springer.com/gp/book/9780387954875}, + author = {Stinson, Douglas}, + date-modified = {2021-03-22 16:11:11 +0000}, + file = {C\:\\Users\\cledm\\Zotero\\storage\\WN767Z98\\Stinson_2004_Combinatorial Designs.pdf;C\:\\Users\\cledm\\Zotero\\storage\\FYAWHUHN\\9780387954875.html}, + isbn = {978-0-387-95487-5}, + language = {en}, + publisher = {Springer}, + shorttitle = {Combinatorial {{Designs}}}, + title = {Combinatorial {{Designs}}: {{Constructions}} and {{Analysis}}}, + year = {2004}} + +@misc{HerkeLectureNotes2016, + author = {Sara Herke}, + title = {MATH3301 Lecture Notes in Combinatorial Design Theory}, + month = {July}, + year = {2016}, + publisher = {University of Queensland} + } + +@incollection{godsilToolsLinearAlgebra, + title = {Tools from {{Linear Algebra}}}, + booktitle = {Handbook of {{Combinatorics}}}, + author = {Godsil, C. D.}, + editor = {Graham RL, Gr{\"o}tschel M, Lov{\'a}sz L}, + volume = {2}, + publisher = {{Elsevier}}, + address = {{Amsterdam}}, + langid = {english}, + file = {C\:\\Users\\cledm\\Zotero\\storage\\J9WN7TIC\\Godsil - Chapter 31 Tools from Linear Algebra.pdf} +} + +@misc{bukhAlgebraicMethodsCombinatoricsa, + title = {Lecture Notes in Algebraic {{Methods}} in {{Combinatorics}}: {{Rank}} Argument}, + author = {Bukh, Boris}, + year = {2014}, + publisher ={Carnegie Mellon University}, + url = {http://www.borisbukh.org/AlgMethods14/rank_notes.pdf}, + langid = {english}, + file = {C\:\\Users\\cledm\\Zotero\\storage\\9FWFUJ85\\Bukh - Algebraic Methods in Combinatorics Rank argument.pdf} +} +@book{juknaExtremalCombinatorics2011, + title = {Extremal {{Combinatorics}}}, + author = {Jukna, Stasys}, + year = {2011}, + series = {Texts in {{Theoretical Computer Science}}. {{An EATCS Series}}}, + publisher = {{Springer Berlin Heidelberg}}, + address = {{Berlin, Heidelberg}}, + isbn = {978-3-642-17363-9 978-3-642-17364-6}, + langid = {english}, + file = {C\:\\Users\\cledm\\Zotero\\storage\\BV8FFARY\\Jukna - 2011 - Extremal Combinatorics.pdf} +} + +@article{fisherExaminationDifferentPossible1940a, + title = {An {{Examination}} of the {{Different Possible Solutions}} of a {{Problem}} in {{Incomplete Blocks}}}, + author = {Fisher, R. A.}, + year = {1940}, + journal = {Annals of Eugenics}, + volume = {10}, + number = {1}, + pages = {52--75}, + issn = {2050-1439}, + langid = {english}, + annotation = {\_eprint: https://onlinelibrary.wiley.com/doi/pdf/10.1111/j.1469-1809.1940.tb02237.x}, + file = {C\:\\Users\\cledm\\Zotero\\storage\\V3PM5T83\\Fisher_1940_An Examination of the Different Possible Solutions of a Problem in Incomplete.pdf;C\:\\Users\\cledm\\Zotero\\storage\\TTXKNKWU\\j.1469-1809.1940.tb02237.html} +} + +@book{babaiLINEARALGEBRAMETHODS1988, + title = {Linear Algebra Methods in Combinatorics}, + author = {Babai, L\'{a}szl\'{o} and Frankl, P\'{e}ter}, + year = {2020}, + edition = {2.1}, + langid = {english}, + file = {C\:\\Users\\cledm\\Zotero\\storage\\WXLVYMME\\Babai and Frankl - 1988 - LINEAR ALGEBRA METHODS IN COMBINATORICS.pdf} +} \ No newline at end of file diff --git a/thys/Fishers_Inequality/document/root.tex b/thys/Fishers_Inequality/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Fishers_Inequality/document/root.tex @@ -0,0 +1,30 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} + +% 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{Fisher's Inequality: Linear Algebraic Proof Techniques for Combinatorics} +\author{Chelsea Edmonds and Lawrence C. Paulson} +\maketitle + +\begin{abstract} + Linear algebraic techniques are powerful, yet often underrated tools in combinatorial proofs. This formalisation provides a library including matrix representations of incidence set systems, general formal proof techniques for the rank argument and linear bound argument, and finally a formalisation of a number of variations of the well-known Fisher's inequality. We build on our prior work formalising combinatorial design theory using a locale-centric approach, including extensions such as constant intersect designs and dual incidence systems. In addition to Fisher's inequality, we also formalise proofs on other incidence system properties using the incidence matrix representation, such as design existence, dual system relationships and incidence system isomorphisms. This formalisation is presented in the paper "Formalising Fisher's Inequality: Formal Linear Algebraic Techniques in Combinatorics", accepted to ITP 2022. +\end{abstract} + +\tableofcontents + +% include generated text of all theories +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,676 +1,677 @@ ADS_Functor AI_Planning_Languages_Semantics AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Ackermanns_not_PR Actuarial_Mathematics 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 Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Blue_Eyes 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 CSP_RefTK CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions 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 CoCon CoSMeDis CoSMed CofGroups Coinductive Coinductive_Languages Collections Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Cotangent_PFD_Formula Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties Cubic_Quartic_Equations DFS_Framework DOM_Components DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Dedekind_Real Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos Dominance_CHK DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Eval_FO Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 FOL_Seq_Calc2 FOL_Seq_Calc3 Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates +Fishers_Inequality Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier FO_Theory_Rewriting Free-Boolean-Algebra Free-Groups Frequency_Moments Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Gale_Shapley 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 Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Hahn_Jordan_Decomposition Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL Hyperdual IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Independence_CH Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpolation_Polynomials_HOL_Algebra Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Irrationals_From_THEBOOK IsaGeoCoq Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knights_Tour Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent 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 Logging_Independent_Anonymity Lowe_Ontological_Argument Lower_Semicontinuous Lp LP_Duality Lucas_Theorem MDP-Algorithms MDP-Rewards MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multiset_Ordering_NPC Multi_Party_Computation Multirelations 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 OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PAL PCF PLM POPLmark-deBruijn PSemigroupsConvolution Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Prefix_Free_Code_Combinators 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 Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic QHLProver QR_Decomposition Quantales Quasi_Borel_Spaces Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Real_Power Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Registers Regression_Test_Selection Regular-Sets Regular_Algebras Regular_Tree_Relations Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups ResiduatedTransitionSystem Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Roth_Arithmetic_Progressions Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Saturation_Framework_Extensions Schutz_Spacetime Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem SpecCheck 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 Sunflowers SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Universal_Hash_Families Universal_Turing_Machine UpDown_Scheme Valuation Van_Emde_Boas_Trees Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith VYDRA_MDL WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Weighted_Path_Order Well_Quasi_Orders Wetzels_Problem Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML Youngs_Inequality ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL