diff --git a/thys/Affine_Arithmetic/Affine_Form.thy b/thys/Affine_Arithmetic/Affine_Form.thy --- a/thys/Affine_Arithmetic/Affine_Form.thy +++ b/thys/Affine_Arithmetic/Affine_Form.thy @@ -1,1989 +1,1989 @@ section \Affine Form\ theory Affine_Form imports "HOL-Analysis.Multivariate_Analysis" - "HOL-Library.List_Permutation" + "HOL-Combinatorics.List_Permutation" Affine_Arithmetic_Auxiliarities Executable_Euclidean_Space begin subsection \Auxiliary developments\ lemma sum_list_mono: fixes xs ys::"'a::ordered_ab_group_add list" shows "length xs = length ys \ (\x y. (x, y) \ set (zip xs ys) \ x \ y) \ sum_list xs \ sum_list ys" by (induct xs ys rule: list_induct2) (auto simp: algebra_simps intro: add_mono) lemma fixes xs::"'a::ordered_comm_monoid_add list" shows sum_list_nonneg: "(\x. x \ set xs \ x \ 0) \ sum_list xs \ 0" by (induct xs) (auto intro!: add_nonneg_nonneg) lemma map_filter: "map f (filter (\x. P (f x)) xs) = filter P (map f xs)" by (induct xs) simp_all lemma map_of_zip_upto2_length_eq_nth: assumes "distinct B" assumes "i < length B" shows "(map_of (zip B [0.. (i, a) \ set xs \ (i, b) \ set xs \ a = b" by (metis (lifting) map_of_is_SomeI option.inject) lemma length_filter_snd_zip: "length ys = length xs \ length (filter (p \ snd) (zip ys xs)) = length (filter p xs)" by (induct ys xs rule: list_induct2) (auto ) lemma filter_snd_nth: "length ys = length xs \ n < length (filter p xs) \ snd (filter (p \ snd) (zip ys xs) ! n) = filter p xs ! n" by (induct ys xs arbitrary: n rule: list_induct2) (auto simp: o_def nth_Cons split: nat.split) lemma distinct_map_snd_fst_eqD: "distinct (map snd xs) \ (i, a) \ set xs \ (j, a) \ set xs \ i = j" by (metis Pair_inject inj_on_contraD snd_conv distinct_map) lemma map_of_mapk_inj_on_SomeI: "inj_on f (fst ` (set t)) \ map_of t k = Some x \ map_of (map (case_prod (\k. Pair (f k))) t) (f k) = Some x" by (induct t) (auto simp add: inj_on_def dest!: map_of_SomeD split: if_split_asm) lemma map_abs_nonneg[simp]: fixes xs::"'a::ordered_ab_group_add_abs list" shows "list_all (\x. x \ 0) xs \ map abs xs = xs" by (induct xs) auto lemma the_inv_into_image_eq: "inj_on f A \ Y \ f ` A \ the_inv_into A f ` Y = f -` Y \ A" using f_the_inv_into_f the_inv_into_f_f[where f = f and A = A] by force lemma image_fst_zip: "length ys = length xs \ fst ` set (zip ys xs) = set ys" by (metis dom_map_of_conv_image_fst dom_map_of_zip) lemma inj_on_fst_set_zip_distinct[simp]: "distinct xs \ length xs = length ys \ inj_on fst (set (zip xs ys))" by (force simp add: in_set_zip distinct_conv_nth intro!: inj_onI) lemma mem_greaterThanLessThan_absI: fixes x::real assumes "abs x < 1" shows "x \ {-1 <..< 1}" using assms by (auto simp: abs_real_def split: if_split_asm) lemma minus_one_less_divideI: "b > 0 \ -b < a \ -1 < a / (b::real)" by (auto simp: field_simps) lemma divide_less_oneI: "b > 0 \ b > a \ a / (b::real) < 1" by (auto simp: field_simps) lemma closed_segment_real: fixes a b::real shows "closed_segment a b = (if a \ b then {a .. b} else {b .. a})" (is "_ = ?if") proof safe fix x assume "x \ closed_segment a b" from segment_bound[OF this] show "x \ ?if" by (auto simp: abs_real_def split: if_split_asm) next fix x assume "x \ ?if" thus "x \ closed_segment a b" by (auto simp: closed_segment_def intro!: exI[where x="(x - a)/(b - a)"] simp: divide_simps algebra_simps) qed subsection \Partial Deviations\ typedef (overloaded) 'a pdevs = "{x::nat \ 'a::zero. finite {i. x i \ 0}}" \ \TODO: unify with polynomials\ morphisms pdevs_apply Abs_pdev by (auto intro!: exI[where x="\x. 0"]) setup_lifting type_definition_pdevs lemma pdevs_eqI: "(\i. pdevs_apply x i = pdevs_apply y i) \ x = y" by transfer auto definition pdevs_val :: "(nat \ real) \ 'a::real_normed_vector pdevs \ 'a" where "pdevs_val e x = (\i. e i *\<^sub>R pdevs_apply x i)" definition valuate:: "((nat \ real) \ 'a) \ 'a set" where "valuate x = x ` (UNIV \ {-1 .. 1})" lemma valuate_ex: "x \ valuate f \ (\e. (\i. e i \ {-1 .. 1}) \ x = f e)" unfolding valuate_def by (auto simp add: valuate_def Pi_iff) blast instantiation pdevs :: (equal) equal begin definition equal_pdevs::"'a pdevs \ 'a pdevs \ bool" where "equal_pdevs a b \ a = b" instance proof qed (simp add: equal_pdevs_def) end subsection \Affine Forms\ text \The data structure of affine forms represents particular sets, zonotopes\ type_synonym 'a aform = "'a \ 'a pdevs" subsection \Evaluation, Range, Joint Range\ definition aform_val :: "(nat \ real) \ 'a::real_normed_vector aform \ 'a" where "aform_val e X = fst X + pdevs_val e (snd X)" definition Affine :: "'a::real_normed_vector aform \ 'a set" where "Affine X = valuate (\e. aform_val e X)" definition Joints :: "'a::real_normed_vector aform list \ 'a list set" where "Joints XS = valuate (\e. map (aform_val e) XS)" lemma Joints_nthE: assumes "zs \ Joints ZS" obtains e where "\i. i < length zs \ zs ! i = aform_val e (ZS ! i)" "\i. e i \ {-1..1}" using assms by atomize_elim (auto simp: Joints_def Pi_iff valuate_ex) lemma Joints_mapE: assumes "ys \ Joints YS" obtains e where "ys = map (\x. aform_val e x) YS" "\i. e i \ {-1 .. 1}" using assms by (force simp: Joints_def valuate_def) lemma zipped_subset_mapped_Elem: assumes "xs = map (aform_val e) XS" assumes e: "\i. e i \ {-1 .. 1}" assumes [simp]: "length xs = length XS" assumes [simp]: "length ys = length YS" assumes "set (zip ys YS) \ set (zip xs XS)" shows "ys = map (aform_val e) YS" proof - from assms have ys: "\i. i < length xs \ xs ! i = aform_val e (XS ! i)" by auto from assms have set_eq: "{(ys ! i, YS ! i) |i. i < length ys \ i < length YS} \ {(xs ! i, XS ! i) |i. i < length xs \ i < length XS}" using assms(2) by (auto simp: set_zip) hence "\ij YS ! i = XS ! j" by auto then obtain j where j: "\i. i < length YS \ ys ! i = xs ! (j i)" "\i. i < length YS \ YS ! i = XS ! (j i)" "\i. i < length YS \ j i < length XS" by metis show ?thesis using assms by (auto simp: Joints_def j ys intro!: exI[where x=e] nth_equalityI) qed lemma Joints_set_zip_subset: assumes "xs \ Joints XS" assumes "length xs = length XS" assumes "length ys = length YS" assumes "set (zip ys YS) \ set (zip xs XS)" shows "ys \ Joints YS" proof - from Joints_mapE assms obtain e where ys: "xs = map (\x. aform_val e x) XS" and e: "\i. e i \ {-1 .. 1}" by blast show "ys \ Joints YS" using e zipped_subset_mapped_Elem[OF ys e assms(2-4)] by (auto simp: Joints_def valuate_def intro!: exI[where x=e]) qed lemma Joints_set_zip: assumes "ys \ Joints YS" assumes "length xs = length XS" assumes "length YS = length XS" assumes sets_eq: "set (zip xs XS) = set (zip ys YS)" shows "xs \ Joints XS" proof - from assms have "length ys = length YS" by (auto simp: Joints_def valuate_def) from assms(1) this assms(2) show ?thesis by (rule Joints_set_zip_subset) (simp add: assms) qed definition Joints2 :: "'a::real_normed_vector aform list \'b::real_normed_vector aform \ ('a list \ 'b) set" where "Joints2 XS Y = valuate (\e. (map (aform_val e) XS, aform_val e Y))" lemma Joints2E: assumes "zs_y \ Joints2 ZS Y" obtains e where "\i. i < length (fst zs_y) \ (fst zs_y) ! i = aform_val e (ZS ! i)" "snd (zs_y) = aform_val e Y" "\i. e i \ {-1..1}" using assms by atomize_elim (auto simp: Joints2_def Pi_iff valuate_ex) lemma nth_in_AffineI: assumes "xs \ Joints XS" assumes "i < length XS" shows "xs ! i \ Affine (XS ! i)" using assms by (force simp: Affine_def Joints_def valuate_def) lemma Cons_nth_in_Joints1: assumes "xs \ Joints XS" assumes "i < length XS" shows "((xs ! i) # xs) \ Joints ((XS ! i) # XS)" using assms by (force simp: Joints_def valuate_def) lemma Cons_nth_in_Joints2: assumes "xs \ Joints XS" assumes "i < length XS" assumes "j < length XS" shows "((xs ! i) #(xs ! j) # xs) \ Joints ((XS ! i)#(XS ! j) # XS)" using assms by (force simp: Joints_def valuate_def) lemma Joints_swap: "x#y#xs\Joints (X#Y#XS) \ y#x#xs \ Joints (Y#X#XS)" by (force simp: Joints_def valuate_def) lemma Joints_swap_Cons_append: "length xs = length XS \ x#ys@xs\Joints (X#YS@XS) \ ys@x#xs \ Joints (YS@X#XS)" by (auto simp: Joints_def valuate_def) lemma Joints_ConsD: "x#xs\Joints (X#XS) \ xs \ Joints XS" by (force simp: Joints_def valuate_def) lemma Joints_appendD1: "ys@xs\Joints (YS@XS) \ length xs = length XS \ xs \ Joints XS" by (force simp: Joints_def valuate_def) lemma Joints_appendD2: "ys@xs\Joints (YS@XS) \ length ys = length YS \ ys \ Joints YS" by (force simp: Joints_def valuate_def) lemma Joints_imp_length_eq: "xs \ Joints XS \ length xs = length XS" by (auto simp: Joints_def valuate_def) lemma Joints_rotate[simp]: "xs@[x] \ Joints (XS @[X]) \ x#xs \ Joints (X#XS)" by (auto simp: Joints_def valuate_def) subsection \Domain\ definition "pdevs_domain x = {i. pdevs_apply x i \ 0}" lemma finite_pdevs_domain[intro, simp]: "finite (pdevs_domain x)" unfolding pdevs_domain_def by transfer lemma in_pdevs_domain[simp]: "i \ pdevs_domain x \ pdevs_apply x i \ 0" by (auto simp: pdevs_domain_def) subsection \Least Fresh Index\ definition degree::"'a::real_vector pdevs \ nat" where "degree x = (LEAST i. \j\i. pdevs_apply x j = 0)" lemma degree[rule_format, intro, simp]: shows "\j\degree x. pdevs_apply x j = 0" unfolding degree_def proof (rule LeastI_ex) have "\j. j > Max (pdevs_domain x) \ j \ (pdevs_domain x)" by (metis Max_less_iff all_not_in_conv less_irrefl_nat finite_pdevs_domain) then show "\xa. \j\xa. pdevs_apply x j = 0" by (auto intro!: exI[where x="Max (pdevs_domain x) + 1"]) qed lemma degree_le: assumes d: "\j \ d. pdevs_apply x j = 0" shows "degree x \ d" unfolding degree_def by (rule Least_le) (rule d) lemma degree_gt: "pdevs_apply x j \ 0 \ degree x > j" by auto lemma pdevs_val_pdevs_domain: "pdevs_val e X = (\i\pdevs_domain X. e i *\<^sub>R pdevs_apply X i)" by (auto simp: pdevs_val_def intro!: suminf_finite) lemma pdevs_val_sum_le: "degree X \ d \ pdevs_val e X = (\i < d. e i *\<^sub>R pdevs_apply X i)" by (force intro!: degree_gt sum.mono_neutral_cong_left simp: pdevs_val_pdevs_domain) lemmas pdevs_val_sum = pdevs_val_sum_le[OF order_refl] lemma pdevs_val_zero[simp]: "pdevs_val (\_. 0) x = 0" by (auto simp: pdevs_val_sum) lemma degree_eqI: assumes "pdevs_apply x d \ 0" assumes "\j. j > d \ pdevs_apply x j = 0" shows "degree x = Suc d" unfolding eq_iff by (auto intro!: degree_gt degree_le assms simp: Suc_le_eq) lemma finite_degree_nonzero[intro, simp]: "finite {i. pdevs_apply x i \ 0}" by transfer (auto simp: vimage_def Collect_neg_eq) lemma degree_eq_Suc_max: "degree x = (if (\i. pdevs_apply x i = 0) then 0 else Suc (Max {i. pdevs_apply x i \ 0}))" proof - { assume "\i. pdevs_apply x i = 0" hence ?thesis by auto (metis degree_le le_0_eq) } moreover { fix i assume "pdevs_apply x i \ 0" hence ?thesis using Max_in[OF finite_degree_nonzero, of x] by (auto intro!: degree_eqI) (metis Max.coboundedI[OF finite_degree_nonzero] in_pdevs_domain le_eq_less_or_eq less_asym pdevs_domain_def) } ultimately show ?thesis by blast qed lemma pdevs_val_degree_cong: assumes "b = d" assumes "\i. i < degree b \ a i = c i" shows "pdevs_val a b = pdevs_val c d" using assms by (auto simp: pdevs_val_sum) abbreviation degree_aform::"'a::real_vector aform \ nat" where "degree_aform X \ degree (snd X)" lemma degree_cong: "(\i. (pdevs_apply x i = 0) = (pdevs_apply y i = 0)) \ degree x = degree y" unfolding degree_def by auto lemma Least_True_nat[intro, simp]: "(LEAST i::nat. True) = 0" by (metis (lifting) One_nat_def less_one not_less_Least not_less_eq) lemma sorted_list_of_pdevs_domain_eq: "sorted_list_of_set (pdevs_domain X) = filter ((\) 0 o pdevs_apply X) [0..x. x", simplified]) subsection \Total Deviation\ definition tdev::"'a::ordered_euclidean_space pdevs \ 'a" where "tdev x = (\ipdevs_apply x i\)" lemma abs_pdevs_val_le_tdev: "e \ UNIV \ {-1 .. 1} \ \pdevs_val e x\ \ tdev x" by (force simp: pdevs_val_sum tdev_def abs_scaleR Pi_iff intro!: order_trans[OF sum_abs] sum_mono scaleR_left_le_one_le intro: abs_leI) subsection \Binary Pointwise Operations\ definition binop_pdevs_raw::"('a::zero \ 'b::zero \ 'c::zero) \ (nat \ 'a) \ (nat \ 'b) \ nat \ 'c" where "binop_pdevs_raw f x y i = (if x i = 0 \ y i = 0 then 0 else f (x i) (y i))" lemma nonzeros_binop_pdevs_subset: "{i. binop_pdevs_raw f x y i \ 0} \ {i. x i \ 0} \ {i. y i \ 0}" by (auto simp: binop_pdevs_raw_def) lift_definition binop_pdevs:: "('a \ 'b \ 'c) \ 'a::zero pdevs \ 'b::zero pdevs \ 'c::zero pdevs" is binop_pdevs_raw using nonzeros_binop_pdevs_subset by (rule finite_subset) auto lemma pdevs_apply_binop_pdevs[simp]: "pdevs_apply (binop_pdevs f x y) i = (if pdevs_apply x i = 0 \ pdevs_apply y i = 0 then 0 else f (pdevs_apply x i) (pdevs_apply y i))" by transfer (auto simp: binop_pdevs_raw_def) subsection \Addition\ definition add_pdevs::"'a::real_vector pdevs \ 'a pdevs \ 'a pdevs" where "add_pdevs = binop_pdevs (+)" lemma pdevs_apply_add_pdevs[simp]: "pdevs_apply (add_pdevs X Y) n = pdevs_apply X n + pdevs_apply Y n" by (auto simp: add_pdevs_def) lemma pdevs_val_add_pdevs[simp]: fixes x y::"'a::euclidean_space" shows "pdevs_val e (add_pdevs X Y) = pdevs_val e X + pdevs_val e Y" proof - let ?sum = "\m X. \i < m. e i *\<^sub>R pdevs_apply X i" let ?m = "max (degree X) (degree Y)" have "pdevs_val e X + pdevs_val e Y = ?sum (degree X) X + ?sum (degree Y) Y" by (simp add: pdevs_val_sum) also have "?sum (degree X) X = ?sum ?m X" by (rule sum.mono_neutral_cong_left) auto also have "?sum (degree Y) Y = ?sum ?m Y" by (rule sum.mono_neutral_cong_left) auto also have "?sum ?m X + ?sum ?m Y = (\i < ?m. e i *\<^sub>R (pdevs_apply X i + pdevs_apply Y i))" by (simp add: scaleR_right_distrib sum.distrib) also have "\ = (\i. e i *\<^sub>R (pdevs_apply X i + pdevs_apply Y i))" by (rule suminf_finite[symmetric]) auto also have "\ = pdevs_val e (add_pdevs X Y)" by (simp add: pdevs_val_def) finally show "pdevs_val e (add_pdevs X Y) = pdevs_val e X + pdevs_val e Y" by simp qed subsection \Total Deviation\ lemma tdev_eq_zero_iff: fixes X::"real pdevs" shows "tdev X = 0 \ (\e. pdevs_val e X = 0)" by (force simp add: pdevs_val_sum tdev_def sum_nonneg_eq_0_iff dest!: spec[where x="\i. if pdevs_apply X i \ 0 then 1 else -1"] split: if_split_asm) lemma tdev_nonneg[intro, simp]: "tdev X \ 0" by (auto simp: tdev_def) lemma tdev_nonpos_iff[simp]: "tdev X \ 0 \ tdev X = 0" by (auto simp: order.antisym) subsection \Unary Operations\ definition unop_pdevs_raw:: "('a::zero \ 'b::zero) \ (nat \ 'a) \ nat \ 'b" where "unop_pdevs_raw f x i = (if x i = 0 then 0 else f (x i))" lemma nonzeros_unop_pdevs_subset: "{i. unop_pdevs_raw f x i \ 0} \ {i. x i \ 0}" by (auto simp: unop_pdevs_raw_def) lift_definition unop_pdevs:: "('a \ 'b) \ 'a::zero pdevs \ 'b::zero pdevs" is unop_pdevs_raw using nonzeros_unop_pdevs_subset by (rule finite_subset) auto lemma pdevs_apply_unop_pdevs[simp]: "pdevs_apply (unop_pdevs f x) i = (if pdevs_apply x i = 0 then 0 else f (pdevs_apply x i))" by transfer (auto simp: unop_pdevs_raw_def) lemma pdevs_domain_unop_linear: assumes "linear f" shows "pdevs_domain (unop_pdevs f x) \ pdevs_domain x" proof - interpret f: linear f by fact show ?thesis by (auto simp: f.zero) qed lemma pdevs_val_unop_linear: assumes "linear f" shows "pdevs_val e (unop_pdevs f x) = f (pdevs_val e x)" proof - interpret f: linear f by fact have *: "\i. (if pdevs_apply x i = 0 then 0 else f (pdevs_apply x i)) = f (pdevs_apply x i)" by (auto simp: f.zero) have "pdevs_val e (unop_pdevs f x) = (\i\pdevs_domain (unop_pdevs f x). e i *\<^sub>R f (pdevs_apply x i))" by (auto simp add: pdevs_val_pdevs_domain *) also have "\ = (\xa\pdevs_domain x. e xa *\<^sub>R f (pdevs_apply x xa))" by (auto intro!: sum.mono_neutral_cong_left) also have "\ = f (pdevs_val e x)" by (auto simp add: pdevs_val_pdevs_domain f.sum f.scaleR) finally show ?thesis . qed subsection \Pointwise Scaling of Partial Deviations\ definition scaleR_pdevs::"real \ 'a::real_vector pdevs \ 'a pdevs" where "scaleR_pdevs r x = unop_pdevs ((*\<^sub>R) r) x" lemma pdevs_apply_scaleR_pdevs[simp]: "pdevs_apply (scaleR_pdevs x Y) n = x *\<^sub>R pdevs_apply Y n" by (auto simp: scaleR_pdevs_def) lemma degree_scaleR_pdevs[simp]: "degree (scaleR_pdevs r x) = (if r = 0 then 0 else degree x)" unfolding degree_def by auto lemma pdevs_val_scaleR_pdevs[simp]: fixes x::real and Y::"'a::real_normed_vector pdevs" shows "pdevs_val e (scaleR_pdevs x Y) = x *\<^sub>R pdevs_val e Y" by (auto simp: pdevs_val_sum scaleR_sum_right ac_simps) subsection \Partial Deviations Scale Pointwise\ definition pdevs_scaleR::"real pdevs \ 'a::real_vector \ 'a pdevs" where "pdevs_scaleR r x = unop_pdevs (\r. r *\<^sub>R x) r" lemma pdevs_apply_pdevs_scaleR[simp]: "pdevs_apply (pdevs_scaleR X y) n = pdevs_apply X n *\<^sub>R y" by (auto simp: pdevs_scaleR_def) lemma degree_pdevs_scaleR[simp]: "degree (pdevs_scaleR r x) = (if x = 0 then 0 else degree r)" unfolding degree_def by auto lemma pdevs_val_pdevs_scaleR[simp]: fixes X::"real pdevs" and y::"'a::real_normed_vector" shows "pdevs_val e (pdevs_scaleR X y) = pdevs_val e X *\<^sub>R y" by (auto simp: pdevs_val_sum scaleR_sum_left) subsection \Pointwise Unary Minus\ definition uminus_pdevs::"'a::real_vector pdevs \ 'a pdevs" where "uminus_pdevs = unop_pdevs uminus" lemma pdevs_apply_uminus_pdevs[simp]: "pdevs_apply (uminus_pdevs x) = - pdevs_apply x" by (auto simp: uminus_pdevs_def) lemma degree_uminus_pdevs[simp]: "degree (uminus_pdevs x) = degree x" by (rule degree_cong) simp lemma pdevs_val_uminus_pdevs[simp]: "pdevs_val e (uminus_pdevs x) = - pdevs_val e x" unfolding pdevs_val_sum by (auto simp: sum_negf) definition "uminus_aform X = (- fst X, uminus_pdevs (snd X))" lemma fst_uminus_aform[simp]: "fst (uminus_aform Y) = - fst Y" by (simp add: uminus_aform_def) lemma aform_val_uminus_aform[simp]: "aform_val e (uminus_aform X) = - aform_val e X" by (auto simp: uminus_aform_def aform_val_def) subsection \Constant\ lift_definition zero_pdevs::"'a::zero pdevs" is "\_. 0" by simp lemma pdevs_apply_zero_pdevs[simp]: "pdevs_apply zero_pdevs i = 0" by transfer simp lemma pdevs_val_zero_pdevs[simp]: "pdevs_val e zero_pdevs = 0" by (auto simp: pdevs_val_def) definition "num_aform f = (f, zero_pdevs)" subsection \Inner Product\ definition pdevs_inner::"'a::euclidean_space pdevs \ 'a \ real pdevs" where "pdevs_inner x b = unop_pdevs (\x. x \ b) x" lemma pdevs_apply_pdevs_inner[simp]: "pdevs_apply (pdevs_inner p a) i = pdevs_apply p i \ a" by (simp add: pdevs_inner_def) lemma pdevs_val_pdevs_inner[simp]: "pdevs_val e (pdevs_inner p a) = pdevs_val e p \ a" by (auto simp add: inner_sum_left pdevs_val_pdevs_domain intro!: sum.mono_neutral_cong_left) definition inner_aform::"'a::euclidean_space aform \ 'a \ real aform" where "inner_aform X b = (fst X \ b, pdevs_inner (snd X) b)" subsection \Inner Product Pair\ definition inner2::"'a::euclidean_space \ 'a \ 'a \ real*real" where "inner2 x n l = (x \ n, x \ l)" definition pdevs_inner2::"'a::euclidean_space pdevs \ 'a \ 'a \ (real*real) pdevs" where "pdevs_inner2 X n l = unop_pdevs (\x. inner2 x n l) X" lemma pdevs_apply_pdevs_inner2[simp]: "pdevs_apply (pdevs_inner2 p a b) i = (pdevs_apply p i \ a, pdevs_apply p i \ b)" by (simp add: pdevs_inner2_def inner2_def zero_prod_def) definition inner2_aform::"'a::euclidean_space aform \ 'a \ 'a \ (real*real) aform" where "inner2_aform X a b = (inner2 (fst X) a b, pdevs_inner2 (snd X) a b)" lemma linear_inner2[intro, simp]: "linear (\x. inner2 x n i)" by unfold_locales (auto simp: inner2_def algebra_simps) lemma aform_val_inner2_aform[simp]: "aform_val e (inner2_aform Z n i) = inner2 (aform_val e Z) n i" proof - have "aform_val e (inner2_aform Z n i) = inner2 (fst Z) n i + inner2 (pdevs_val e (snd Z)) n i" by (auto simp: aform_val_def inner2_aform_def pdevs_inner2_def pdevs_val_unop_linear) also have "\ = inner2 (aform_val e Z) n i" by (simp add: inner2_def algebra_simps aform_val_def) finally show ?thesis . qed subsection \Update\ lemma pdevs_val_upd[simp]: "pdevs_val (e(n := e')) X = pdevs_val e X - e n * pdevs_apply X n + e' * pdevs_apply X n" unfolding pdevs_val_def by (subst suminf_finite[OF finite.insertI[OF finite_degree_nonzero], of n X], auto simp: pdevs_val_def sum.insert_remove)+ lemma nonzeros_fun_upd: "{i. (f(n := a)) i \ 0} \ {i. f i \ 0} \ {n}" by (auto split: if_split_asm) lift_definition pdev_upd::"'a::real_vector pdevs \ nat \ 'a \ 'a pdevs" is "\x n a. x(n:=a)" by (rule finite_subset[OF nonzeros_fun_upd]) simp lemma pdevs_apply_pdev_upd[simp]: "pdevs_apply (pdev_upd X n x) = (pdevs_apply X)(n:=x)" by transfer simp lemma pdevs_val_pdev_upd[simp]: "pdevs_val e (pdev_upd X n x) = pdevs_val e X + e n *\<^sub>R x - e n *\<^sub>R pdevs_apply X n" unfolding pdevs_val_def by (subst suminf_finite[OF finite.insertI[OF finite_degree_nonzero], of n X], auto simp: pdevs_val_def sum.insert_remove)+ lemma degree_pdev_upd: assumes "x = 0 \ pdevs_apply X n = 0" shows "degree (pdev_upd X n x) = degree X" using assms by (auto intro!: degree_cong split: if_split_asm) lemma degree_pdev_upd_le: assumes "degree X \ n" shows "degree (pdev_upd X n x) \ Suc n" using assms by (auto intro!: degree_le) subsection \Inf/Sup\ definition "Inf_aform X = fst X - tdev (snd X)" definition "Sup_aform X = fst X + tdev (snd X)" lemma Inf_aform: assumes "e \ UNIV \ {-1 .. 1}" shows "Inf_aform X \ aform_val e X" using order_trans[OF abs_ge_minus_self abs_pdevs_val_le_tdev[OF assms]] by (auto simp: Inf_aform_def aform_val_def minus_le_iff) lemma Sup_aform: assumes "e \ UNIV \ {-1 .. 1}" shows "aform_val e X \ Sup_aform X" using order_trans[OF abs_ge_self abs_pdevs_val_le_tdev[OF assms]] by (auto simp: Sup_aform_def aform_val_def) subsection \Minkowski Sum\ definition msum_pdevs_raw::"nat\(nat \ 'a::real_vector)\(nat \ 'a)\(nat\'a)" where "msum_pdevs_raw n x y i = (if i < n then x i else y (i - n))" lemma nonzeros_msum_pdevs_raw: "{i. msum_pdevs_raw n f g i \ 0} = ({0.. {i. f i \ 0}) \ (+) n ` ({i. g i \ 0})" by (force simp: msum_pdevs_raw_def not_less split: if_split_asm) lift_definition msum_pdevs::"nat\'a::real_vector pdevs\'a pdevs\'a pdevs" is msum_pdevs_raw unfolding nonzeros_msum_pdevs_raw by simp lemma pdevs_apply_msum_pdevs: "pdevs_apply (msum_pdevs n f g) i = (if i < n then pdevs_apply f i else pdevs_apply g (i - n))" by transfer (auto simp: msum_pdevs_raw_def) lemma degree_least_nonzero: assumes "degree f \ 0" shows "pdevs_apply f (degree f - 1) \ 0" proof assume H: "pdevs_apply f (degree f - 1) = 0" { fix j assume "j\degree f - 1" with H have "pdevs_apply f j = 0" by (cases "degree f - 1 = j") auto } from degree_le[rule_format, OF this] have "degree f \ degree f - 1" by blast with assms show False by simp qed lemma degree_leI: assumes "(\i. pdevs_apply y i = 0 \ pdevs_apply x i = 0)" shows "degree x \ degree y" proof cases assume "degree x \ 0" from degree_least_nonzero[OF this] have "pdevs_apply y (degree x - 1) \ 0" by (auto simp: assms split: if_split_asm) from degree_gt[OF this] show ?thesis by simp qed simp lemma degree_msum_pdevs_ge1: shows "degree f \ n \ degree f \ degree (msum_pdevs n f g)" by (rule degree_leI) (auto simp: pdevs_apply_msum_pdevs split: if_split_asm) lemma degree_msum_pdevs_ge2: assumes "degree f \ n" shows "degree g \ degree (msum_pdevs n f g) - n" proof cases assume "degree g \ 0" hence "pdevs_apply g (degree g - 1) \ 0" by (rule degree_least_nonzero) hence "pdevs_apply (msum_pdevs n f g) (n + degree g - 1) \ 0" using assms by (auto simp: pdevs_apply_msum_pdevs) from degree_gt[OF this] show ?thesis by simp qed simp lemma degree_msum_pdevs_le: shows "degree (msum_pdevs n f g) \ n + degree g" by (auto intro!: degree_le simp: pdevs_apply_msum_pdevs) lemma sum_msum_pdevs_cases: assumes "degree f \ n" assumes [simp]: "\i. e i 0 = 0" shows "(\i i i i\{.. {i. i < n}. e i (pdevs_apply f i)) + (\i\{.. - {i. i < n}. e i (pdevs_apply g (i - n)))" (is "_ = ?sum_f + ?sum_g") by (simp add: sum.If_cases if_distrib) also have "?sum_f = (\i = 0..i\{0 + n.. = (\i = 0.. = (\i = 0.. n \ tdev (msum_pdevs n f g) = tdev f + tdev g" by (auto simp: tdev_def pdevs_apply_msum_pdevs intro!: sum_msum_pdevs_cases) lemma pdevs_val_msum_pdevs: "degree f \ n \ pdevs_val e (msum_pdevs n f g) = pdevs_val e f + pdevs_val (\i. e (i + n)) g" by (auto simp: pdevs_val_sum pdevs_apply_msum_pdevs intro!: sum_msum_pdevs_cases) definition msum_aform::"nat \ 'a::real_vector aform \ 'a aform \ 'a aform" where "msum_aform n f g = (fst f + fst g, msum_pdevs n (snd f) (snd g))" lemma fst_msum_aform[simp]: "fst (msum_aform n f g) = fst f + fst g" by (simp add: msum_aform_def) lemma snd_msum_aform[simp]: "snd (msum_aform n f g) = msum_pdevs n (snd f) (snd g)" by (simp add: msum_aform_def) lemma finite_nonzero_summable: "finite {i. f i \ 0} \ summable f" by (auto intro!: sums_summable sums_finite) lemma aform_val_msum_aform: assumes "degree_aform f \ n" shows "aform_val e (msum_aform n f g) = aform_val e f + aform_val (\i. e (i + n)) g" using assms by (auto simp: pdevs_val_msum_pdevs aform_val_def) lemma Inf_aform_msum_aform: "degree_aform X \ n \ Inf_aform (msum_aform n X Y) = Inf_aform X + Inf_aform Y" by (simp add: Inf_aform_def tdev_msum_pdevs) lemma Sup_aform_msum_aform: "degree_aform X \ n \ Sup_aform (msum_aform n X Y) = Sup_aform X + Sup_aform Y" by (simp add: Sup_aform_def tdev_msum_pdevs) definition "independent_from d Y = msum_aform d (0, zero_pdevs) Y" definition "independent_aform X Y = independent_from (degree_aform X) Y" lemma degree_zero_pdevs[simp]: "degree zero_pdevs = 0" by (metis degree_least_nonzero pdevs_apply_zero_pdevs) lemma independent_aform_Joints: assumes "x \ Affine X" assumes "y \ Affine Y" shows "[x, y] \ Joints [X, independent_aform X Y]" using assms unfolding Affine_def valuate_def Joints_def apply safe subgoal premises prems for e ea using prems by (intro image_eqI[where x="\i. if i < degree_aform X then e i else ea (i - degree_aform X)"]) (auto simp: aform_val_def pdevs_val_msum_pdevs Pi_iff independent_aform_def independent_from_def intro!: pdevs_val_degree_cong) done lemma msum_aform_Joints: assumes "d \ degree_aform X" assumes "\X. X \ set XS \ d \ degree_aform X" assumes "(x#xs) \ Joints (X#XS)" assumes "y \ Affine Y" shows "((x + y)#x#xs) \ Joints (msum_aform d X Y#X#XS)" using assms unfolding Joints_def valuate_def Affine_def proof (safe, goal_cases) case (1 e ea a b zs) then show ?case by (intro image_eqI[where x = "\i. if i < d then e i else ea (i - d)"]) (force simp: aform_val_def pdevs_val_msum_pdevs intro!: intro!: pdevs_val_degree_cong)+ qed lemma Joints_msum_aform: assumes "d \ degree_aform X" assumes "\Y. Y \ set YS \ d \ degree_aform Y" shows "Joints (msum_aform d X Y#YS) = {((x + y)#ys) |x y ys. y \ Affine Y \ x#ys \ Joints (X#YS)}" unfolding Affine_def valuate_def Joints_def proof (safe, goal_cases) case (1 x e) thus ?case using assms by (intro exI[where x = "aform_val e X"] exI[where x = "aform_val ((\i. e (i + d))) Y"]) (auto simp add: aform_val_def pdevs_val_msum_pdevs) next case (2 x xa y ys e ea) thus ?case using assms by (intro image_eqI[where x="\i. if i < d then ea i else e (i - d)"]) (force simp: aform_val_def pdevs_val_msum_pdevs Pi_iff intro!: pdevs_val_degree_cong)+ qed lemma Joints_singleton_image: "Joints [x] = (\x. [x]) ` Affine x" by (auto simp: Joints_def Affine_def valuate_def) lemma Collect_extract_image: "{g (f x y) |x y. P x y} = g ` {f x y |x y. P x y}" by auto lemma inj_Cons: "inj (\x. x#xs)" by (auto intro!: injI) lemma Joints_Nil[simp]: "Joints [] = {[]}" by (force simp: Joints_def valuate_def) lemma msum_pdevs_zero_ident[simp]: "msum_pdevs 0 zero_pdevs x = x" by transfer (auto simp: msum_pdevs_raw_def) lemma msum_aform_zero_ident[simp]: "msum_aform 0 (0, zero_pdevs) x = x" by (simp add: msum_aform_def) lemma mem_Joints_singleton: "(x \ Joints [X]) = (\y. x = [y] \ y \ Affine X)" by (auto simp: Affine_def valuate_def Joints_def) lemma singleton_mem_Joints[simp]: "[x] \ Joints [X] \ x \ Affine X" by (auto simp: mem_Joints_singleton) lemma msum_aform_Joints_without_first: assumes "d \ degree_aform X" assumes "\X. X \ set XS \ d \ degree_aform X" assumes "(x#xs) \ Joints (X#XS)" assumes "y \ Affine Y" assumes "z = x + y" shows "z#xs \ Joints (msum_aform d X Y#XS)" unfolding \z = x + y\ using msum_aform_Joints[OF assms(1-4)] by (force simp: Joints_def valuate_def) lemma Affine_msum_aform: assumes "d \ degree_aform X" shows "Affine (msum_aform d X Y) = {x + y |x y. x \ Affine X \ y \ Affine Y}" using Joints_msum_aform[OF assms, of Nil Y, simplified, unfolded mem_Joints_singleton] by (auto simp add: Joints_singleton_image Collect_extract_image[where g="\x. [x]"] inj_image_eq_iff[OF inj_Cons] ) lemma Affine_zero_pdevs[simp]: "Affine (0, zero_pdevs) = {0}" by (force simp: Affine_def valuate_def aform_val_def) lemma Affine_independent_aform: "Affine (independent_aform X Y) = Affine Y" by (auto simp: independent_aform_def independent_from_def Affine_msum_aform) lemma abs_diff_eq1: fixes l u::"'a::ordered_euclidean_space" shows "l \ u \ \u - l\ = u - l" by (metis abs_of_nonneg diff_add_cancel le_add_same_cancel2) lemma compact_sum: fixes f :: "'a \ 'b::topological_space \ 'c::real_normed_vector" assumes "finite I" assumes "\i. i \ I \ compact (S i)" assumes "\i. i \ I \ continuous_on (S i) (f i)" assumes "I \ J" shows "compact {\i\I. f i (x i) | x. x \ Pi J S}" using assms proof (induct I) case empty thus ?case proof (cases "\x. x \ Pi J S") case False hence *: "{\i\{}. f i (x i) |x. x \ Pi J S} = {}" by (auto simp: Pi_iff) show ?thesis unfolding * by simp qed auto next case (insert a I) hence "{\i\insert a I. f i (xa i) |xa. xa \ Pi J S} = {x + y |x y. x \ f a ` S a \ y \ {\i\I. f i (x i) |x. x \ Pi J S}}" proof safe fix s x assume "s \ S a" "x \ Pi J S" thus "\xa. f a s + (\i\I. f i (x i)) = (\i\insert a I. f i (xa i)) \ xa \ Pi J S" using insert by (auto intro!: exI[where x="x(a:=s)"] sum.cong) qed force also have "compact \" using insert by (intro compact_sums) (auto intro!: compact_continuous_image) finally show ?case . qed lemma compact_Affine: fixes X::"'a::ordered_euclidean_space aform" shows "compact (Affine X)" proof - have "Affine X = {x + y|x y. x \ {fst X} \ y \ {(\i \ {0..R pdevs_apply (snd X) i) | e. e \ UNIV \ {-1 .. 1}}}" by (auto simp: Affine_def valuate_def aform_val_def pdevs_val_sum atLeast0LessThan) also have "compact \" by (rule compact_sums) (auto intro!: compact_sum continuous_intros) finally show ?thesis . qed lemma Joints2_JointsI: "(xs, x) \ Joints2 XS X \ x#xs \ Joints (X#XS)" by (auto simp: Joints_def Joints2_def valuate_def) subsection \Splitting\ definition "split_aform X i = (let xi = pdevs_apply (snd X) i /\<^sub>R 2 in ((fst X - xi, pdev_upd (snd X) i xi), (fst X + xi, pdev_upd (snd X) i xi)))" lemma split_aformE: assumes "e \ UNIV \ {-1 .. 1}" assumes "x = aform_val e X" obtains err where "x = aform_val (e(i:=err)) (fst (split_aform X i))" "err \ {-1 .. 1}" | err where "x = aform_val (e(i:=err)) (snd (split_aform X i))" "err \ {-1 .. 1}" proof (atomize_elim) let ?thesis = "(\err. x = aform_val (e(i := err)) (fst (split_aform X i)) \ err \ {- 1..1}) \ (\err. x = aform_val (e(i := err)) (snd (split_aform X i)) \ err \ {- 1..1})" { assume "pdevs_apply (snd X) i = 0" hence "X = fst (split_aform X i)" by (auto simp: split_aform_def intro!: prod_eqI pdevs_eqI) with assms have ?thesis by (auto intro!: exI[where x="e i"]) } moreover { assume "pdevs_apply (snd X) i \ 0" hence [simp]: "degree_aform X > i" by (rule degree_gt) note assms(2) also have "aform_val e X = fst X + (\iR pdevs_apply (snd X) i)" by (simp add: aform_val_def pdevs_val_sum) also have rewr: "{.. {i}" by auto have "(\iR pdevs_apply (snd X) i) = (\i \ {0..R pdevs_apply (snd X) i) + e i *\<^sub>R pdevs_apply (snd X) i" by (subst rewr, subst sum.union_disjoint) auto finally have "x = fst X + \" . hence "x = aform_val (e(i:=2 * e i - 1)) (snd (split_aform X i))" "x = aform_val (e(i:=2 * e i + 1)) (fst (split_aform X i))" by (auto simp: aform_val_def split_aform_def Let_def pdevs_val_sum atLeast0LessThan Diff_eq degree_pdev_upd if_distrib sum.If_cases field_simps scaleR_left_distrib[symmetric]) moreover have "2 * e i - 1 \ {-1 .. 1} \ 2 * e i + 1 \ {-1 .. 1}" using assms by (auto simp: not_le Pi_iff dest!: spec[where x=i]) ultimately have ?thesis by blast } ultimately show ?thesis by blast qed lemma pdevs_val_add: "pdevs_val (\i. e i + f i) xs = pdevs_val e xs + pdevs_val f xs" by (auto simp: pdevs_val_pdevs_domain algebra_simps sum.distrib) lemma pdevs_val_minus: "pdevs_val (\i. e i - f i) xs = pdevs_val e xs - pdevs_val f xs" by (auto simp: pdevs_val_pdevs_domain algebra_simps sum_subtractf) lemma pdevs_val_cmul: "pdevs_val (\i. u * e i) xs = u *\<^sub>R pdevs_val e xs" by (auto simp: pdevs_val_pdevs_domain scaleR_sum_right) lemma atLeastAtMost_absI: "- a \ a \ \x::real\ \ \a\ \ x \ atLeastAtMost (- a) a" by auto lemma divide_atLeastAtMost_1_absI: "\x::real\ \ \a\ \ x/a \ {-1 .. 1}" by (intro atLeastAtMost_absI) (auto simp: divide_le_eq_1) lemma convex_scaleR_aux: "u + v = 1 \ u *\<^sub>R x + v *\<^sub>R x = (x::'a::real_vector)" by (metis scaleR_add_left scaleR_one) lemma convex_mult_aux: "u + v = 1 \ u * x + v * x = (x::real)" using convex_scaleR_aux[of u v x] by simp lemma convex_Affine: "convex (Affine X)" proof (rule convexI) fix x y::'a and u v::real assume "x \ Affine X" "y \ Affine X" and convex: "0 \ u" "0 \ v" "u + v = 1" then obtain e f where x: "x = aform_val e X" "e \ UNIV \ {-1 .. 1}" and y: "y = aform_val f X" "f \ UNIV \ {-1 .. 1}" by (auto simp: Affine_def valuate_def) let ?conv = "\i. u * e i + v * f i" { fix i have "\?conv i\ \ u * \e i\ + v * \f i\" using convex by (intro order_trans[OF abs_triangle_ineq]) (simp add: abs_mult) also have "\ \ 1" using convex x y by (intro convex_bound_le) (auto simp: Pi_iff abs_real_def) finally have "?conv i \ 1" "-1 \ ?conv i" by (auto simp: abs_real_def split: if_split_asm) } thus "u *\<^sub>R x + v *\<^sub>R y \ Affine X" using convex x y by (auto simp: Affine_def valuate_def aform_val_def pdevs_val_add pdevs_val_cmul algebra_simps convex_scaleR_aux intro!: image_eqI[where x="?conv"]) qed lemma segment_in_aform_val: assumes "e \ UNIV \ {-1 .. 1}" assumes "f \ UNIV \ {-1 .. 1}" shows "closed_segment (aform_val e X) (aform_val f X) \ Affine X" proof - have "aform_val e X \ Affine X" "aform_val f X \ Affine X" using assms by (auto simp: Affine_def valuate_def) with convex_Affine[of X, simplified convex_contains_segment] show ?thesis by simp qed subsection \From List of Generators\ lift_definition pdevs_of_list::"'a::zero list \ 'a pdevs" is "\xs i. if i < length xs then xs ! i else 0" by auto lemma pdevs_apply_pdevs_of_list: "pdevs_apply (pdevs_of_list xs) i = (if i < length xs then xs ! i else 0)" by transfer simp lemma pdevs_apply_pdevs_of_list_Nil[simp]: "pdevs_apply (pdevs_of_list []) i = 0" by transfer auto lemma pdevs_apply_pdevs_of_list_Cons: "pdevs_apply (pdevs_of_list (x # xs)) i = (if i = 0 then x else pdevs_apply (pdevs_of_list xs) (i - 1))" by transfer auto lemma pdevs_domain_pdevs_of_list_Cons[simp]: "pdevs_domain (pdevs_of_list (x # xs)) = (if x = 0 then {} else {0}) \ (+) 1 ` pdevs_domain (pdevs_of_list xs)" by (force simp: pdevs_apply_pdevs_of_list_Cons split: if_split_asm) lemma pdevs_val_pdevs_of_list_eq[simp]: "pdevs_val e (pdevs_of_list (x # xs)) = e 0 *\<^sub>R x + pdevs_val (e o (+) 1) (pdevs_of_list xs)" proof - have "pdevs_val e (pdevs_of_list (x # xs)) = (\i\pdevs_domain (pdevs_of_list (x # xs)) \ {0}. e i *\<^sub>R x) + (\i\pdevs_domain (pdevs_of_list (x # xs)) \ - {0}. e i *\<^sub>R pdevs_apply (pdevs_of_list xs) (i - Suc 0))" (is "_ = ?l + ?r") by (simp add: pdevs_val_pdevs_domain if_distrib sum.If_cases pdevs_apply_pdevs_of_list_Cons) also have "?r = (\i\pdevs_domain (pdevs_of_list xs). e (Suc i) *\<^sub>R pdevs_apply (pdevs_of_list xs) i)" by (rule sum.reindex_cong[of "\i. i + 1"]) auto also have "\ = pdevs_val (e o (+) 1) (pdevs_of_list xs)" by (simp add: pdevs_val_pdevs_domain ) also have "?l = (\i\{0}. e i *\<^sub>R x)" by (rule sum.mono_neutral_cong_left) auto also have "\ = e 0 *\<^sub>R x" by simp finally show ?thesis . qed lemma less_degree_pdevs_of_list_imp_less_length: assumes "i < degree (pdevs_of_list xs)" shows "i < length xs" proof - from assms have "pdevs_apply (pdevs_of_list xs) (degree (pdevs_of_list xs) - 1) \ 0" by (metis degree_least_nonzero less_nat_zero_code) hence "degree (pdevs_of_list xs) - 1 < length xs" by (simp add: pdevs_apply_pdevs_of_list split: if_split_asm) with assms show ?thesis by simp qed lemma tdev_pdevs_of_list[simp]: "tdev (pdevs_of_list xs) = sum_list (map abs xs)" by (auto simp: tdev_def pdevs_apply_pdevs_of_list sum_list_sum_nth less_degree_pdevs_of_list_imp_less_length intro!: sum.mono_neutral_cong_left degree_gt) lemma pdevs_of_list_Nil[simp]: "pdevs_of_list [] = zero_pdevs" by (auto intro!: pdevs_eqI) lemma pdevs_val_inj_sumI: fixes K::"'a set" and g::"'a \ nat" assumes "finite K" assumes "inj_on g K" assumes "pdevs_domain x \ g ` K" assumes "\i. i \ K \ g i \ pdevs_domain x \ f i = 0" assumes "\i. i \ K \ g i \ pdevs_domain x \ f i = e (g i) *\<^sub>R pdevs_apply x (g i)" shows "pdevs_val e x = (\i\K. f i)" proof - have [simp]: "inj_on (the_inv_into K g) (pdevs_domain x)" using assms by (auto simp: intro!: subset_inj_on[OF inj_on_the_inv_into]) { fix y assume y: "y \ pdevs_domain x" have g_inv: "g (the_inv_into K g y) = y" by (meson assms(2) assms(3) y f_the_inv_into_f subset_eq) have inv_in: "the_inv_into K g y \ K" by (meson assms(2) assms(3) y subset_iff in_pdevs_domain the_inv_into_into) have inv3: "the_inv_into (pdevs_domain x) (the_inv_into K g) (the_inv_into K g y) = g (the_inv_into K g y)" using assms y by (subst the_inv_into_f_f) (auto simp: f_the_inv_into_f[OF assms(2)]) note g_inv inv_in inv3 } note this[simp] have "pdevs_val e x = (\i\pdevs_domain x. e i *\<^sub>R pdevs_apply x i)" by (simp add: pdevs_val_pdevs_domain) also have "\ = (\i \ the_inv_into K g ` pdevs_domain x. e (g i) *\<^sub>R pdevs_apply x (g i))" by (rule sum.reindex_cong[OF inj_on_the_inv_into]) auto also have "\ = (\i\K. f i)" using assms by (intro sum.mono_neutral_cong_left) (auto simp: the_inv_into_image_eq) finally show ?thesis . qed lemma pdevs_domain_pdevs_of_list_le: "pdevs_domain (pdevs_of_list xs) \ {0..(i,x)\zip [0..R x)" by (auto simp: sum_list_distinct_conv_sum_set in_set_zip image_fst_zip pdevs_apply_pdevs_of_list distinct_zipI1 intro!: pdevs_val_inj_sumI[of _ fst] split: if_split_asm) lemma scaleR_sum_list: fixes xs::"'a::real_vector list" shows "a *\<^sub>R sum_list xs = sum_list (map (scaleR a) xs)" by (induct xs) (auto simp: algebra_simps) lemma pdevs_val_const_pdevs_of_list: "pdevs_val (\_. c) (pdevs_of_list xs) = c *\<^sub>R sum_list xs" unfolding pdevs_val_zip split_beta' scaleR_sum_list by (rule arg_cong) (auto intro!: nth_equalityI) lemma pdevs_val_partition: assumes "e \ UNIV \ I" obtains f g where "pdevs_val e (pdevs_of_list xs) = pdevs_val f (pdevs_of_list (filter p xs)) + pdevs_val g (pdevs_of_list (filter (Not o p) xs))" "f \ UNIV \ I" "g \ UNIV \ I" proof - obtain i where i: "i \ I" by (metis assms funcset_mem iso_tuple_UNIV_I) let ?zip = "zip [0.. snd) ?zip" let ?f = "(\n. if n < degree (pdevs_of_list (filter p xs)) then e (map fst (fst part) ! n) else i)" let ?g = "(\n. if n < degree (pdevs_of_list (filter (Not \ p) xs)) then e (map fst (snd part) ! n) else i)" show ?thesis proof have "pdevs_val e (pdevs_of_list xs) = (\(i,x)\?zip. e i *\<^sub>R x)" by (rule pdevs_val_zip) also have "\ = (\(i, x)\set ?zip. e i *\<^sub>R x)" by (simp add: sum_list_distinct_conv_sum_set distinct_zipI1) also have [simp]: "set (fst part) \ set (snd part) = {}" by (auto simp: part_def) from partition_set[of "p o snd" ?zip "fst part" "snd part"] have "set ?zip = set (fst part) \ set (snd part)" by (auto simp: part_def) also have "(\a\set (fst part) \ set (snd part). case a of (i, x) \ e i *\<^sub>R x) = (\(i, x)\set (fst part). e i *\<^sub>R x) + (\(i, x)\set (snd part). e i *\<^sub>R x)" by (auto simp: split_beta sum_Un) also have "(\(i, x)\set (fst part). e i *\<^sub>R x) = (\(i, x)\(fst part). e i *\<^sub>R x)" by (simp add: sum_list_distinct_conv_sum_set distinct_zipI1 part_def) also have "\ = (\i e i *\<^sub>R x)" by (subst sum_list_sum_nth) (simp add: split_beta' atLeast0LessThan) also have "\ = pdevs_val (\n. e (map fst (fst part) ! n)) (pdevs_of_list (map snd (fst part)))" by (force simp: pdevs_val_zip sum_list_distinct_conv_sum_set distinct_zipI1 split_beta' in_set_zip intro!: sum.reindex_cong[where l=fst] image_eqI[where x = "(x, map snd (fst part) ! x)" for x]) also have "(\(i, x)\set (snd part). e i *\<^sub>R x) = (\(i, x)\(snd part). e i *\<^sub>R x)" by (simp add: sum_list_distinct_conv_sum_set distinct_zipI1 part_def) also have "\ = (\i e i *\<^sub>R x)" by (subst sum_list_sum_nth) (simp add: split_beta' atLeast0LessThan) also have "\ = pdevs_val (\n. e (map fst (snd part) ! n)) (pdevs_of_list (map snd (snd part)))" by (force simp: pdevs_val_zip sum_list_distinct_conv_sum_set distinct_zipI1 split_beta' in_set_zip intro!: sum.reindex_cong[where l=fst] image_eqI[where x = "(x, map snd (snd part) ! x)" for x]) also have "pdevs_val (\n. e (map fst (fst part) ! n)) (pdevs_of_list (map snd (fst part))) = pdevs_val (\n. if n < degree (pdevs_of_list (map snd (fst part))) then e (map fst (fst part) ! n) else i) (pdevs_of_list (map snd (fst part)))" by (rule pdevs_val_degree_cong) simp_all also have "pdevs_val (\n. e (map fst (snd part) ! n)) (pdevs_of_list (map snd (snd part))) = pdevs_val (\n. if n < degree (pdevs_of_list (map snd (snd part))) then e (map fst (snd part) ! n) else i) (pdevs_of_list (map snd (snd part)))" by (rule pdevs_val_degree_cong) simp_all also have "map snd (snd part) = filter (Not o p) xs" by (simp add: part_def filter_map[symmetric] o_assoc) also have "map snd (fst part) = filter p xs" by (simp add: part_def filter_map[symmetric]) finally show "pdevs_val e (pdevs_of_list xs) = pdevs_val ?f (pdevs_of_list (filter p xs)) + pdevs_val ?g (pdevs_of_list (filter (Not \ p) xs))" . show "?f \ UNIV \ I" "?g \ UNIV \ I" using assms \i\I\ by (auto simp: Pi_iff) qed qed lemma pdevs_apply_pdevs_of_list_append: "pdevs_apply (pdevs_of_list (xs @ zs)) i = (if i < length xs then pdevs_apply (pdevs_of_list xs) i else pdevs_apply (pdevs_of_list zs) (i - length xs))" by (auto simp: pdevs_apply_pdevs_of_list nth_append) lemma degree_pdevs_of_list_le_length[intro, simp]: "degree (pdevs_of_list xs) \ length xs" by (metis less_irrefl_nat le_less_linear less_degree_pdevs_of_list_imp_less_length) lemma degree_pdevs_of_list_append: "degree (pdevs_of_list (xs @ ys)) \ length xs + degree (pdevs_of_list ys)" by (rule degree_le) (auto simp: pdevs_apply_pdevs_of_list_append) lemma pdevs_val_pdevs_of_list_append: assumes "f \ UNIV \ I" assumes "g \ UNIV \ I" obtains e where "pdevs_val f (pdevs_of_list xs) + pdevs_val g (pdevs_of_list ys) = pdevs_val e (pdevs_of_list (xs @ ys))" "e \ UNIV \ I" proof let ?e = "(\i. if i < length xs then f i else g (i - length xs))" have f: "pdevs_val f (pdevs_of_list xs) = (\i\{..R pdevs_apply (pdevs_of_list (xs @ ys)) i)" by (auto simp: pdevs_val_sum degree_gt pdevs_apply_pdevs_of_list_append intro: sum.mono_neutral_cong_left) have g: "pdevs_val g (pdevs_of_list ys) = (\i=length xs ..R pdevs_apply (pdevs_of_list (xs @ ys)) i)" (is "_ = ?sg") by (auto simp: pdevs_val_sum pdevs_apply_pdevs_of_list_append intro!: inj_onI image_eqI[where x="length xs + x" for x] sum.reindex_cong[where l="\i. i - length xs"]) show "pdevs_val f (pdevs_of_list xs) + pdevs_val g (pdevs_of_list ys) = pdevs_val ?e (pdevs_of_list (xs @ ys))" unfolding f g by (subst sum.union_disjoint[symmetric]) (force simp: pdevs_val_sum ivl_disj_un degree_pdevs_of_list_append intro!: sum.mono_neutral_cong_right split: if_split_asm)+ show "?e \ UNIV \ I" using assms by (auto simp: Pi_iff) qed lemma sum_general_mono: fixes f::"'a\('b::ordered_ab_group_add)" assumes [simp,intro]: "finite s" "finite t" assumes f: "\x. x \ s - t \ f x \ 0" assumes g: "\x. x \ t - s \ g x \ 0" assumes fg: "\x. x \ s \ t \ f x \ g x" shows "(\x \ s. f x) \ (\x \ t. g x)" proof - have "s = (s - t) \ (s \ t)" and [intro, simp]: "(s - t) \ (s \ t) = {}" by auto hence "(\x \ s. f x) = (\x \ s - t \ s \ t. f x)" using assms by simp also have "\ = (\x \ s - t. f x) + (\x \ s \ t. f x)" by (simp add: sum_Un) also have "(\x \ s - t. f x) \ 0" by (auto intro!: sum_nonpos f) also have "0 \ (\x \ t - s. g x)" by (auto intro!: sum_nonneg g) also have "(\x \ s \ t. f x) \ (\x \ s \ t. g x)" by (auto intro!: sum_mono fg) also have [intro, simp]: "(t - s) \ (s \ t) = {}" by auto hence "sum g (t - s) + sum g (s \ t) = sum g ((t - s) \ (s \ t))" by (simp add: sum_Un) also have "\ = sum g t" by (auto intro!: sum.cong) finally show ?thesis by simp qed lemma pdevs_val_perm_ex: assumes "xs <~~> ys" assumes mem: "e \ UNIV \ I" shows "\e'. e' \ UNIV \ I \ pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list ys)" using assms proof (induct arbitrary: e) case Nil thus ?case by auto next case (Cons xs ys z) hence "(e \ (+) (Suc 0)) \ UNIV \ I" by auto from Cons(2)[OF this] obtain e' where "e' \ UNIV \ I" "pdevs_val (e \ (+) (Suc 0)) (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list ys)" by metis thus ?case using Cons by (auto intro!: exI[where x="\x. if x = 0 then e 0 else e' (x - 1)"] simp: o_def Pi_iff) next case (trans xs ys zs) thus ?case by metis next case (swap y x l) thus ?case by (auto intro!: exI[where x="\i. if i = 0 then e 1 else if i = 1 then e 0 else e i"] simp: o_def Pi_iff) qed lemma pdevs_val_perm: assumes "xs <~~> ys" assumes mem: "e \ UNIV \ I" obtains e' where "e' \ UNIV \ I" "pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list ys)" using assms by (metis pdevs_val_perm_ex) lemma set_distinct_permI: "set xs = set ys \ distinct xs \ distinct ys \ xs <~~> ys" by (metis eq_set_perm_remdups remdups_id_iff_distinct) lemmas pdevs_val_permute = pdevs_val_perm[OF set_distinct_permI] lemma partition_permI: "filter p xs @ filter (Not o p) xs <~~> xs" proof (induct xs) case (Cons x xs) have swap_app_Cons: "filter p xs @ x # [a\xs . \ p a] <~~> x # filter p xs @ [a\xs . \ p a]" by (metis perm_sym perm_append_Cons) also have "\ <~~> x#xs" using Cons by auto finally (trans) show ?case using Cons by simp qed simp lemma pdevs_val_eqI: assumes "\i. i \ pdevs_domain y \ i \ pdevs_domain x \ e i *\<^sub>R pdevs_apply x i = f i *\<^sub>R pdevs_apply y i" assumes "\i. i \ pdevs_domain y \ i \ pdevs_domain x \ f i *\<^sub>R pdevs_apply y i = 0" assumes "\i. i \ pdevs_domain x \ i \ pdevs_domain y \ e i *\<^sub>R pdevs_apply x i = 0" shows "pdevs_val e x = pdevs_val f y" using assms by (force simp: pdevs_val_pdevs_domain intro!: sum.reindex_bij_witness_not_neutral[where i=id and j = id and S'="pdevs_domain x - pdevs_domain y" and T'="pdevs_domain y - pdevs_domain x"]) definition filter_pdevs_raw::"(nat \ 'a \ bool) \ (nat \ 'a::real_vector) \ (nat \ 'a)" where "filter_pdevs_raw I X = (\i. if I i (X i) then X i else 0)" lemma filter_pdevs_raw_nonzeros: "{i. filter_pdevs_raw s f i \ 0} = {i. f i \ 0} \ {x. s x (f x)}" by (auto simp: filter_pdevs_raw_def) lift_definition filter_pdevs::"(nat \ 'a \ bool) \ 'a::real_vector pdevs \ 'a pdevs" is filter_pdevs_raw by (simp add: filter_pdevs_raw_nonzeros) lemma pdevs_apply_filter_pdevs[simp]: "pdevs_apply (filter_pdevs I x) i = (if I i (pdevs_apply x i) then pdevs_apply x i else 0)" by transfer (auto simp: filter_pdevs_raw_def) lemma degree_filter_pdevs_le: "degree (filter_pdevs I x) \ degree x" by (rule degree_leI) (simp split: if_split_asm) lemma pdevs_val_filter_pdevs: "pdevs_val e (filter_pdevs I x) = (\i \ {.. {i. I i (pdevs_apply x i)}. e i *\<^sub>R pdevs_apply x i)" by (auto simp: pdevs_val_sum if_distrib sum.inter_restrict degree_filter_pdevs_le degree_gt intro!: sum.mono_neutral_cong_left split: if_split_asm) lemma pdevs_val_filter_pdevs_dom: "pdevs_val e (filter_pdevs I x) = (\i \ pdevs_domain x \ {i. I i (pdevs_apply x i)}. e i *\<^sub>R pdevs_apply x i)" by (auto simp: pdevs_val_pdevs_domain if_distrib sum.inter_restrict degree_filter_pdevs_le degree_gt intro!: sum.mono_neutral_cong_left split: if_split_asm) lemma pdevs_val_filter_pdevs_eval: "pdevs_val e (filter_pdevs p x) = pdevs_val (\i. if p i (pdevs_apply x i) then e i else 0) x" by (auto split: if_split_asm intro!: pdevs_val_eqI) definition "pdevs_applys X i = map (\x. pdevs_apply x i) X" definition "pdevs_vals e X = map (pdevs_val e) X" definition "aform_vals e X = map (aform_val e) X" definition "filter_pdevs_list I X = map (filter_pdevs (\i _. I i (pdevs_applys X i))) X" lemma pdevs_applys_filter_pdevs_list[simp]: "pdevs_applys (filter_pdevs_list I X) i = (if I i (pdevs_applys X i) then pdevs_applys X i else map (\_. 0) X)" by (auto simp: filter_pdevs_list_def o_def pdevs_applys_def) definition "degrees X = Max (insert 0 (degree ` set X))" abbreviation "degree_aforms X \ degrees (map snd X)" lemma degrees_leI: assumes "\x. x \ set X \ degree x \ K" shows "degrees X \ K" using assms by (auto simp: degrees_def intro!: Max.boundedI) lemma degrees_leD: assumes "degrees X \ K" shows "\x. x \ set X \ degree x \ K" using assms by (auto simp: degrees_def intro!: Max.boundedI) lemma degree_filter_pdevs_list_le: "degrees (filter_pdevs_list I x) \ degrees x" by (rule degrees_leI) (auto simp: filter_pdevs_list_def intro!: degree_le dest!: degrees_leD) definition "dense_list_of_pdevs x = map (\i. pdevs_apply x i) [0..(reverse) ordered coefficients as list\ definition "list_of_pdevs x = map (\i. (i, pdevs_apply x i)) (rev (sorted_list_of_set (pdevs_domain x)))" lemma list_of_pdevs_zero_pdevs[simp]: "list_of_pdevs zero_pdevs = []" by (auto simp: list_of_pdevs_def) lemma sum_list_list_of_pdevs: "sum_list (map snd (list_of_pdevs x)) = sum_list (dense_list_of_pdevs x)" by (auto intro!: sum.mono_neutral_cong_left simp add: degree_gt sum_list_distinct_conv_sum_set dense_list_of_pdevs_def list_of_pdevs_def) lemma sum_list_filter_dense_list_of_pdevs[symmetric]: "sum_list (map snd (filter (p o snd) (list_of_pdevs x))) = sum_list (filter p (dense_list_of_pdevs x))" by (auto intro!: sum.mono_neutral_cong_left simp add: degree_gt sum_list_distinct_conv_sum_set dense_list_of_pdevs_def list_of_pdevs_def o_def filter_map) lemma pdevs_of_list_dense_list_of_pdevs: "pdevs_of_list (dense_list_of_pdevs x) = x" by (auto simp: pdevs_apply_pdevs_of_list dense_list_of_pdevs_def pdevs_eqI) lemma pdevs_val_sum_list: "pdevs_val (\_. c) X = c *\<^sub>R sum_list (map snd (list_of_pdevs X))" by (auto simp: pdevs_val_sum sum_list_list_of_pdevs pdevs_val_const_pdevs_of_list[symmetric] pdevs_of_list_dense_list_of_pdevs) lemma list_of_pdevs_all_nonzero: "list_all (\x. x \ 0) (map snd (list_of_pdevs xs))" by (auto simp: list_of_pdevs_def list_all_iff) lemma list_of_pdevs_nonzero: "x \ set (map snd (list_of_pdevs xs)) \ x \ 0" by (auto simp: list_of_pdevs_def) lemma pdevs_of_list_scaleR_0[simp]: fixes xs::"'a::real_vector list" shows "pdevs_of_list (map ((*\<^sub>R) 0) xs) = zero_pdevs" by (auto simp: pdevs_apply_pdevs_of_list intro!: pdevs_eqI) lemma degree_pdevs_of_list_scaleR: "degree (pdevs_of_list (map ((*\<^sub>R) c) xs)) = (if c \ 0 then degree (pdevs_of_list xs) else 0)" by (auto simp: pdevs_apply_pdevs_of_list intro!: degree_cong) lemma list_of_pdevs_eq: "rev (list_of_pdevs X) = (filter ((\) 0 o snd) (map (\i. (i, pdevs_apply X i)) [0..i. if i < d then 1 else 0) (pdevs_of_list xs)" proof - have "sum_list (take d xs) = 1 *\<^sub>R sum_list (take d xs)" by simp also note pdevs_val_const_pdevs_of_list[symmetric] also have "pdevs_val (\_. 1) (pdevs_of_list (take d xs)) = pdevs_val (\i. if i < d then 1 else 0) (pdevs_of_list xs)" by (auto simp: pdevs_apply_pdevs_of_list split: if_split_asm intro!: pdevs_val_eqI) finally show ?thesis . qed lemma zero_in_range_pdevs_apply[intro, simp]: fixes X::"'a::real_vector pdevs" shows "0 \ range (pdevs_apply X)" by (metis degree_gt less_irrefl rangeI) lemma dense_list_in_range: "x \ set (dense_list_of_pdevs X) \ x \ range (pdevs_apply X)" by (auto simp: dense_list_of_pdevs_def) lemma not_in_dense_list_zeroD: assumes "pdevs_apply X i \ set (dense_list_of_pdevs X)" shows "pdevs_apply X i = 0" proof (rule ccontr) assume "pdevs_apply X i \ 0" hence "i < degree X" by (rule degree_gt) thus False using assms by (auto simp: dense_list_of_pdevs_def) qed lemma list_all_list_of_pdevsI: assumes "\i. i \ pdevs_domain X \ P (pdevs_apply X i)" shows "list_all (\x. P x) (map snd (list_of_pdevs X))" using assms by (auto simp: list_all_iff list_of_pdevs_def) lemma pdevs_of_list_map_scaleR: "pdevs_of_list (map (scaleR r) xs) = scaleR_pdevs r (pdevs_of_list xs)" by (auto intro!: pdevs_eqI simp: pdevs_apply_pdevs_of_list) lemma map_permI: assumes "xs <~~> ys" shows "map f xs <~~> map f ys" using assms by induct auto lemma rev_perm: "rev xs <~~> ys \ xs <~~> ys" by (metis perm.trans perm_rev rev_rev_ident) lemma list_of_pdevs_perm_filter_nonzero: "map snd (list_of_pdevs X) <~~> (filter ((\) 0) (dense_list_of_pdevs X))" proof - have zip_map: "zip [0..i. (i, pdevs_apply X i)) [0.. filter ((\) 0 o snd) (zip [0.. map snd (filter ((\) 0 \ snd) (zip [0..) 0 \ snd) (zip [0..) 0) (dense_list_of_pdevs X)" using map_filter[of snd "(\) 0" "(zip [0.. UNIV \ I" assumes "0 \ I" obtains e' where "pdevs_val e (pdevs_of_list (filter p xs)) = pdevs_val e' (pdevs_of_list xs)" "e' \ UNIV \ I" unfolding pdevs_val_filter_pdevs_eval proof - have "(\_::nat. 0) \ UNIV \ I" using assms by simp have "pdevs_val e (pdevs_of_list (filter p xs)) = pdevs_val e (pdevs_of_list (filter p xs)) + pdevs_val (\_. 0) (pdevs_of_list (filter (Not o p) xs))" by (simp add: pdevs_val_sum) also from pdevs_val_pdevs_of_list_append[OF \e \ _\ \(\_. 0) \ _\] obtain e' where "e' \ UNIV \ I" "\ = pdevs_val e' (pdevs_of_list (filter p xs @ filter (Not o p) xs))" by metis note this(2) also from pdevs_val_perm[OF partition_permI \e' \ _\] obtain e'' where "\ = pdevs_val e'' (pdevs_of_list xs)" "e'' \ UNIV \ I" by metis note this(1) finally show ?thesis using \e'' \ _\ .. qed lemma pdevs_val_of_list_of_pdevs: assumes "e \ UNIV \ I" assumes "0 \ I" obtains e' where "pdevs_val e (pdevs_of_list (map snd (list_of_pdevs X))) = pdevs_val e' X" "e' \ UNIV \ I" proof - obtain e' where "e' \ UNIV \ I" and "pdevs_val e (pdevs_of_list (map snd (list_of_pdevs X))) = pdevs_val e' (pdevs_of_list (filter ((\) 0) (dense_list_of_pdevs X)))" by (rule pdevs_val_perm[OF list_of_pdevs_perm_filter_nonzero assms(1)]) note this(2) also from pdevs_val_filter[OF \e' \ _\ \0 \ I\, of "(\) 0" "dense_list_of_pdevs X"] obtain e'' where "e'' \ UNIV \ I" and "\ = pdevs_val e'' (pdevs_of_list (dense_list_of_pdevs X))" by metis note this(2) also have "\ = pdevs_val e'' X" by (simp add: pdevs_of_list_dense_list_of_pdevs) finally show ?thesis using \e'' \ UNIV \ I\ .. qed lemma pdevs_val_of_list_of_pdevs2: assumes "e \ UNIV \ I" obtains e' where "pdevs_val e X = pdevs_val e' (pdevs_of_list (map snd (list_of_pdevs X)))" "e' \ UNIV \ I" proof - from list_of_pdevs_perm_filter_nonzero[of X] have perm: "(filter ((\) 0) (dense_list_of_pdevs X)) <~~> map snd (list_of_pdevs X)" by (simp add: perm_sym) have "pdevs_val e X = pdevs_val e (pdevs_of_list (dense_list_of_pdevs X))" by (simp add: pdevs_of_list_dense_list_of_pdevs) also from pdevs_val_partition[OF \e \ _\, of "dense_list_of_pdevs X" "(\) 0"] obtain f g where "f \ UNIV \ I" "g \ UNIV \ I" "\ = pdevs_val f (pdevs_of_list (filter ((\) 0) (dense_list_of_pdevs X))) + pdevs_val g (pdevs_of_list (filter (Not \ (\) 0) (dense_list_of_pdevs X)))" (is "_ = ?f + ?g") by metis note this(3) also have "pdevs_of_list [x\dense_list_of_pdevs X . x = 0] = zero_pdevs" by (auto intro!: pdevs_eqI simp: pdevs_apply_pdevs_of_list dest!: nth_mem) hence "?g = 0" by (auto simp: o_def ) also obtain e' where "e' \ UNIV \ I" and "?f = pdevs_val e' (pdevs_of_list (map snd (list_of_pdevs X)))" by (rule pdevs_val_perm[OF perm \f \ _\]) note this(2) finally show ?thesis using \e' \ UNIV \ I\ by (auto intro!: that) qed lemma dense_list_of_pdevs_scaleR: "r \ 0 \ map ((*\<^sub>R) r) (dense_list_of_pdevs x) = dense_list_of_pdevs (scaleR_pdevs r x)" by (auto simp: dense_list_of_pdevs_def) lemma degree_pdevs_of_list_eq: "(\x. x \ set xs \ x \ 0) \ degree (pdevs_of_list xs) = length xs" by (cases xs) (auto simp add: pdevs_apply_pdevs_of_list nth_Cons intro!: degree_eqI split: nat.split) lemma dense_list_of_pdevs_pdevs_of_list: "(\x. x \ set xs \ x \ 0) \ dense_list_of_pdevs (pdevs_of_list xs) = xs" by (auto simp: dense_list_of_pdevs_def degree_pdevs_of_list_eq pdevs_apply_pdevs_of_list intro!: nth_equalityI) lemma pdevs_of_list_sum: assumes "distinct xs" assumes "e \ UNIV \ I" obtains f where "f \ UNIV \ I" "pdevs_val e (pdevs_of_list xs) = (\P\set xs. f P *\<^sub>R P)" proof - define f where "f X = e (the (map_of (zip xs [0.. UNIV \ I" by (auto simp: f_def) moreover have "pdevs_val e (pdevs_of_list xs) = (\P\set xs. f P *\<^sub>R P)" by (auto simp add: pdevs_val_zip f_def assms sum_list_distinct_conv_sum_set[symmetric] in_set_zip map_of_zip_upto2_length_eq_nth intro!: sum_list_nth_eqI) ultimately show ?thesis .. qed lemma pdevs_domain_eq_pdevs_of_list: assumes nz: "\x. x \ set (xs) \ x \ 0" shows "pdevs_domain (pdevs_of_list xs) = {0..x. x \ set xs \ x \ 0" shows "length (list_of_pdevs (pdevs_of_list xs)) = length xs" using nz by (auto simp: list_of_pdevs_def pdevs_domain_eq_pdevs_of_list) lemma nth_list_of_pdevs_pdevs_of_list: assumes nz: "\x. x \ set xs \ x \ 0" assumes l: "n < length xs" shows "list_of_pdevs (pdevs_of_list xs) ! n = ((length xs - Suc n), xs ! (length xs - Suc n))" using nz l by (auto simp: list_of_pdevs_def pdevs_domain_eq_pdevs_of_list rev_nth pdevs_apply_pdevs_of_list) lemma list_of_pdevs_pdevs_of_list_eq: "(\x. x \ set xs \ x \ 0) \ list_of_pdevs (pdevs_of_list xs) = zip (rev [0..x. x \ set xs \ x \ 0" shows "sum_list (filter p (map snd (list_of_pdevs (pdevs_of_list xs)))) = sum_list (filter p xs)" using assms by (auto simp: list_of_pdevs_pdevs_of_list_eq rev_filter[symmetric]) lemma sum_list_partition: fixes xs::"'a::comm_monoid_add list" shows "sum_list (filter p xs) + sum_list (filter (Not o p) xs) = sum_list xs" by (induct xs) (auto simp: ac_simps) subsection \2d zonotopes\ definition "prod_of_pdevs x y = binop_pdevs Pair x y" lemma apply_pdevs_prod_of_pdevs[simp]: "pdevs_apply (prod_of_pdevs x y) i = (pdevs_apply x i, pdevs_apply y i)" unfolding prod_of_pdevs_def by (simp add: zero_prod_def) lemma pdevs_domain_prod_of_pdevs[simp]: "pdevs_domain (prod_of_pdevs x y) = pdevs_domain x \ pdevs_domain y" by (auto simp: zero_prod_def) lemma pdevs_val_prod_of_pdevs[simp]: "pdevs_val e (prod_of_pdevs x y) = (pdevs_val e x, pdevs_val e y)" proof - have "pdevs_val e x = (\i\pdevs_domain x \ pdevs_domain y. e i *\<^sub>R pdevs_apply x i)" (is "_ = ?x") unfolding pdevs_val_pdevs_domain by (rule sum.mono_neutral_cong_left) auto moreover have "pdevs_val e y = (\i\pdevs_domain x \ pdevs_domain y. e i *\<^sub>R pdevs_apply y i)" (is "_ = ?y") unfolding pdevs_val_pdevs_domain by (rule sum.mono_neutral_cong_left) auto ultimately have "(pdevs_val e x, pdevs_val e y) = (?x, ?y)" by auto also have "\ = pdevs_val e (prod_of_pdevs x y)" by (simp add: sum_prod pdevs_val_pdevs_domain) finally show ?thesis by simp qed definition prod_of_aforms (infixr "\\<^sub>a" 80) where "prod_of_aforms x y = ((fst x, fst y), prod_of_pdevs (snd x) (snd y))" subsection \Intervals\ definition One_pdevs_raw::"nat \ 'a::executable_euclidean_space" where "One_pdevs_raw i = (if i < length (Basis_list::'a list) then Basis_list ! i else 0)" lemma zeros_One_pdevs_raw: "One_pdevs_raw -` {0::'a::executable_euclidean_space} = {length (Basis_list::'a list)..}" by (auto simp: One_pdevs_raw_def nonzero_Basis split: if_split_asm dest!: nth_mem) lemma nonzeros_One_pdevs_raw: "{i. One_pdevs_raw i \ (0::'a::executable_euclidean_space)} = - {length (Basis_list::'a list)..}" using zeros_One_pdevs_raw by blast lift_definition One_pdevs::"'a::executable_euclidean_space pdevs" is One_pdevs_raw by (auto simp: nonzeros_One_pdevs_raw) lemma pdevs_apply_One_pdevs[simp]: "pdevs_apply One_pdevs i = (if i < length (Basis_list::'a::executable_euclidean_space list) then Basis_list ! i else 0::'a)" by transfer (simp add: One_pdevs_raw_def) lemma Max_Collect_less_nat: "Max {i::nat. i < k} = (if k = 0 then Max {} else k - 1)" by (auto intro!: Max_eqI) lemma degree_One_pdevs[simp]: "degree (One_pdevs::'a pdevs) = length (Basis_list::'a::executable_euclidean_space list)" by (auto simp: degree_eq_Suc_max Basis_list_nth_nonzero Max_Collect_less_nat intro!: Max_eqI DIM_positive) definition inner_scaleR_pdevs::"'a::euclidean_space \ 'a pdevs \ 'a pdevs" where "inner_scaleR_pdevs b x = unop_pdevs (\x. (b \ x) *\<^sub>R x) x" lemma pdevs_apply_inner_scaleR_pdevs[simp]: "pdevs_apply (inner_scaleR_pdevs a x) i = (a \ (pdevs_apply x i)) *\<^sub>R (pdevs_apply x i)" by (simp add: inner_scaleR_pdevs_def) lemma degree_inner_scaleR_pdevs_le: "degree (inner_scaleR_pdevs (l::'a::executable_euclidean_space) One_pdevs) \ degree (One_pdevs::'a pdevs)" by (rule degree_leI) (auto simp: inner_scaleR_pdevs_def One_pdevs_raw_def) definition "pdevs_of_ivl l u = scaleR_pdevs (1/2) (inner_scaleR_pdevs (u - l) One_pdevs)" lemma degree_pdevs_of_ivl_le: "degree (pdevs_of_ivl l u::'a::executable_euclidean_space pdevs) \ DIM('a)" using degree_inner_scaleR_pdevs_le by (simp add: pdevs_of_ivl_def) lemma pdevs_apply_pdevs_of_ivl: defines "B \ Basis_list::'a::executable_euclidean_space list" shows "pdevs_apply (pdevs_of_ivl l u) i = (if i < length B then ((u - l)\(B!i)/2)*\<^sub>R(B!i) else 0)" by (auto simp: pdevs_of_ivl_def B_def) lemma deg_length_less_imp[simp]: "k < degree (pdevs_of_ivl l u::'a::executable_euclidean_space pdevs) \ k < length (Basis_list::'a list)" by (metis (no_types, hide_lams) degree_One_pdevs degree_inner_scaleR_pdevs_le degree_scaleR_pdevs dual_order.strict_trans length_Basis_list_pos nat_neq_iff not_le pdevs_of_ivl_def) lemma tdev_pdevs_of_ivl: "tdev (pdevs_of_ivl l u) = \u - l\ /\<^sub>R 2" proof - have "tdev (pdevs_of_ivl l u) = (\i pdevs_apply (pdevs_of_ivl l u) i\)" by (auto simp: tdev_def) also have "\ = (\i = 0..pdevs_apply (pdevs_of_ivl l u) i\)" using degree_pdevs_of_ivl_le[of l u] by (intro sum.mono_neutral_cong_left) auto also have "\ = (\i = 0..((u - l) \ Basis_list ! i / 2) *\<^sub>R Basis_list ! i\)" by (auto simp: pdevs_apply_pdevs_of_ivl) also have "\ = (\b \ Basis_list. \((u - l) \ b / 2) *\<^sub>R b\)" by (auto simp: sum_list_sum_nth) also have "\ = (\b\Basis. \((u - l) \ b / 2) *\<^sub>R b\)" by (auto simp: sum_list_distinct_conv_sum_set) also have "\ = \u - l\ /\<^sub>R 2" by (subst euclidean_representation[symmetric, of "\u - l\ /\<^sub>R 2"]) (simp add: abs_inner abs_scaleR) finally show ?thesis . qed definition "aform_of_ivl l u = ((l + u)/\<^sub>R2, pdevs_of_ivl l u)" definition "aform_of_point x = aform_of_ivl x x" lemma Elem_affine_of_ivl_le: assumes "e \ UNIV \ {-1 .. 1}" assumes "l \ u" shows "l \ aform_val e (aform_of_ivl l u)" proof - have "l = (1 / 2) *\<^sub>R l + (1 / 2) *\<^sub>R l" by (simp add: scaleR_left_distrib[symmetric]) also have "\ = (l + u)/\<^sub>R2 - tdev (pdevs_of_ivl l u)" by (auto simp: assms tdev_pdevs_of_ivl algebra_simps) also have "\ \ aform_val e (aform_of_ivl l u)" using abs_pdevs_val_le_tdev[OF assms(1), of "pdevs_of_ivl l u"] by (auto simp: aform_val_def aform_of_ivl_def minus_le_iff dest!: abs_le_D2) finally show ?thesis . qed lemma Elem_affine_of_ivl_ge: assumes "e \ UNIV \ {-1 .. 1}" assumes "l \ u" shows "aform_val e (aform_of_ivl l u) \ u" proof - have "aform_val e (aform_of_ivl l u) \ (l + u)/\<^sub>R2 + tdev (pdevs_of_ivl l u)" using abs_pdevs_val_le_tdev[OF assms(1), of "pdevs_of_ivl l u"] by (auto simp: aform_val_def aform_of_ivl_def minus_le_iff dest!: abs_le_D1) also have "\ = (1 / 2) *\<^sub>R u + (1 / 2) *\<^sub>R u" by (auto simp: assms tdev_pdevs_of_ivl algebra_simps) also have "\ = u" by (simp add: scaleR_left_distrib[symmetric]) finally show ?thesis . qed lemma map_of_zip_upto_length_eq_nth: assumes "i < length B" assumes "d = length B" shows "(map_of (zip [0.. {l .. u}" obtains e where "e \ UNIV \ {-1 .. 1}" "k = aform_val e (aform_of_ivl l u)" proof atomize_elim define e where [abs_def]: "e i = (let b = if i R 2) \ b) / (((u - l) /\<^sub>R 2) \ b))" for i let ?B = "Basis_list::'a list" have "k = (1 / 2) *\<^sub>R (l + u) + (\b \ Basis. (if (u - l) \ b = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ b)) *\<^sub>R b)" (is "_ = _ + ?dots") using assms by (force simp add: algebra_simps eucl_le[where 'a='a] intro!: euclidean_eqI[where 'a='a]) also have "?dots = (\b \ Basis. (if (u - l) \ b = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ b) *\<^sub>R b))" by (auto intro!: sum.cong) also have "\ = (\b \ ?B. (if (u - l) \ b = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ b) *\<^sub>R b))" by (auto simp: sum_list_distinct_conv_sum_set) also have "\ = (\i = 0.. ?B ! i = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ ?B ! i) *\<^sub>R ?B ! i))" by (auto simp: sum_list_sum_nth) also have "\ = (\i = 0.. Basis_list ! i = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ Basis_list ! i) *\<^sub>R Basis_list ! i))" using degree_inner_scaleR_pdevs_le[of "u - l"] by (intro sum.mono_neutral_cong_right) (auto dest!: degree) also have "(1 / 2) *\<^sub>R (l + u) + (\i = 0.. Basis_list ! i = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ Basis_list ! i) *\<^sub>R Basis_list ! i)) = aform_val e (aform_of_ivl l u)" using degree_inner_scaleR_pdevs_le[of "u - l"] by (auto simp: aform_val_def aform_of_ivl_def pdevs_of_ivl_def map_of_zip_upto_length_eq_nth e_def Let_def pdevs_val_sum intro!: sum.cong) finally have "k = aform_val e (aform_of_ivl l u)" . moreover { fix k l u::real assume *: "l \ k" "k \ u" let ?m = "l / 2 + u / 2" have "\k - ?m\ \ \if k \ ?m then ?m - l else u - ?m\" using * by auto also have "\ \ \u / 2 - l / 2\" by (auto simp: abs_real_def) finally have "\k - (l / 2 + u / 2)\ \ \u / 2 - l/2\" . } note midpoint_abs = this have "e \ UNIV \ {- 1..1}" using assms unfolding e_def Let_def by (intro Pi_I divide_atLeastAtMost_1_absI) (auto simp: map_of_zip_upto_length_eq_nth eucl_le[where 'a='a] divide_le_eq_1 not_less inner_Basis algebra_simps intro!: midpoint_abs dest!: nth_mem) ultimately show "\e. e \ UNIV \ {- 1..1} \ k = aform_val e (aform_of_ivl l u)" by blast qed lemma Inf_aform_aform_of_ivl: assumes "l \ u" shows "Inf_aform (aform_of_ivl l u) = l" using assms by (auto simp: Inf_aform_def aform_of_ivl_def tdev_pdevs_of_ivl abs_diff_eq1 algebra_simps) (metis field_sum_of_halves scaleR_add_left scaleR_one) lemma Sup_aform_aform_of_ivl: assumes "l \ u" shows "Sup_aform (aform_of_ivl l u) = u" using assms by (auto simp: Sup_aform_def aform_of_ivl_def tdev_pdevs_of_ivl abs_diff_eq1 algebra_simps) (metis field_sum_of_halves scaleR_add_left scaleR_one) lemma Affine_aform_of_ivl: "a \ b \ Affine (aform_of_ivl a b) = {a .. b}" by (force simp: Affine_def valuate_def intro!: Elem_affine_of_ivl_ge Elem_affine_of_ivl_le elim!: in_ivl_affine_of_ivlE) end diff --git a/thys/Affine_Arithmetic/ROOT b/thys/Affine_Arithmetic/ROOT --- a/thys/Affine_Arithmetic/ROOT +++ b/thys/Affine_Arithmetic/ROOT @@ -1,14 +1,15 @@ chapter AFP session Affine_Arithmetic (AFP) = "HOL-Analysis" + options [timeout = 2400] sessions "HOL-Decision_Procs" + "HOL-Combinatorics" "List-Index" Deriving Show theories Affine_Arithmetic document_files "root.tex" "root.bib" diff --git a/thys/Bell_Numbers_Spivey/Bell_Numbers.thy b/thys/Bell_Numbers_Spivey/Bell_Numbers.thy --- a/thys/Bell_Numbers_Spivey/Bell_Numbers.thy +++ b/thys/Bell_Numbers_Spivey/Bell_Numbers.thy @@ -1,534 +1,534 @@ (* Author: Lukas Bulwahn *) section \Bell Numbers and Spivey's Generalized Recurrence\ theory Bell_Numbers imports "HOL-Library.FuncSet" "HOL-Library.Monad_Syntax" - "HOL-Library.Stirling" + "HOL-Combinatorics.Stirling" Card_Partitions.Injectivity_Solver Card_Partitions.Card_Partitions begin subsection \Preliminaries\ subsubsection \Additions to FuncSet\ (* this is clearly to be added to FuncSet *) lemma extensional_funcset_ext: assumes "f \ A \\<^sub>E B" "g \ A \\<^sub>E B" assumes "\x. x \ A \ f x = g x" shows "f = g" using assms by (metis PiE_iff extensionalityI) subsubsection \Additions for Injectivity Proofs\ lemma inj_on_impl_inj_on_image: assumes "inj_on f A" assumes "\x. x \ X \ x \ A" shows "inj_on ((`) f) X" using assms by (meson inj_onI inj_on_image_eq_iff) lemma injectivity_union: assumes "A \ B = C \ D" assumes "P A" "P C" assumes "Q B" "Q D" "\S T. P S \ Q T \ S \ T = {}" shows "A = C \ B = D" using assms Int_Un_distrib Int_commute inf_sup_absorb by blast+ lemma injectivity_image: assumes "f ` A = g ` A" assumes "\x\A. invert (f x) = x \ invert (g x) = x" shows "\x\A. f x = g x" using assms by (metis (no_types, lifting) image_iff) lemma injectivity_image_union: assumes "(\X. X \ F X) ` P = (\X. X \ G X) ` P'" assumes "\X \ P. X \ A" "\X \ P'. X \ A" assumes "\X \ P. \y\F X. y \ A" "\X \ P'. \y\G X. y \ A" shows "P = P'" proof show "P \ P'" proof fix X assume "X \ P" from assms(1) this obtain X' where "X' \ P'" and "X \ F X = X' \ G X'" by (metis imageE image_eqI) moreover from assms(2,4) \X \ P\ have X: "(X \ F X) \ A = X" by auto moreover from assms(3,5) \X' \ P'\ have X': "(X' \ G X') \ A = X'" by auto ultimately have "X = X'" by simp from this \X' \ P'\ show "X \ P'" by auto qed next show "P' \ P" proof fix X' assume "X' \ P'" from assms(1) this obtain X where "X \ P" and "X \ F X = X' \ G X'" by (metis imageE image_eqI) moreover from assms(2,4) \X \ P\ have X: "(X \ F X) \ A = X" by auto moreover from assms(3,5) \X' \ P'\ have X': "(X' \ G X') \ A = X'" by auto ultimately have "X = X'" by simp from this \X \ P\ show "X' \ P" by auto qed qed subsection \Definition of Bell Numbers\ definition Bell :: "nat \ nat" where "Bell n = card {P. partition_on {0..finite A\ obtain f where bij: "bij_betw f {0..x \ {P. partition_on {0.. Pow {0..finite A\ show ?thesis unfolding Bell_def by (subst bij_betw_iff_card[symmetric]) (auto intro: finitely_many_partition_on) qed lemma Bell_0: "Bell 0 = 1" by (auto simp add: Bell_def partition_on_empty) subsection \Construction of the Partitions\ definition construct_partition_on :: "'a set \ 'a set \ 'a set set set" where "construct_partition_on B C = do { k \ {0..card B}; j \ {0..card C}; P \ {P. partition_on C P \ card P = j}; B' \ {B'. B' \ B \ card B' = k}; Q \ {Q. partition_on B' Q}; f \ (B - B') \\<^sub>E P; P' \ {(\X. X \ {x \ B - B'. f x = X}) ` P}; {P' \ Q} }" lemma construct_partition_on: assumes "finite B" "finite C" assumes "B \ C = {}" shows "construct_partition_on B C = {P. partition_on (B \ C) P}" proof (rule set_eqI') fix Q' assume "Q' \ construct_partition_on B C" from this obtain j k P P' Q B' f where "j \ card C" and "k \ card B" and P: "partition_on C P \ card P = j" and B': "B' \ B \ card B' = k" and Q: "partition_on B' Q" and f: "f \ B - B' \\<^sub>E P" and P': "P' = (\X. X \ {x \ B - B'. f x = X}) ` P" and Q': "Q' = P' \ Q" unfolding construct_partition_on_def by auto from P f have "partition_on (B - B' \ C) P'" unfolding P' using \B \ C = {}\ by (intro partition_on_insert_elements) auto from this Q have "partition_on ((B - B' \ C) \ B') Q'" unfolding Q' using B' \B \ C = {}\ by (auto intro: partition_on_union) from this have "partition_on (B \ C) Q'" using B' by (metis Diff_partition sup.assoc sup.commute) from this show "Q' \ {P. partition_on (B \ C) P}" by auto next fix Q' assume Q': "Q' \ {Q'. partition_on (B \ C) Q'}" from Q' have "{} \ Q'" by (auto elim!: partition_onE) obtain Q where Q: "Q = ((\X. if X \ B then X else {}) ` Q') - {{}}" by blast obtain P' where P': "P' = ((\X. if X \ B then {} else X) ` Q') - {{}}" by blast from P' Q \{} \ Q'\ have Q'_prop: "Q' = P' \ Q" by auto have P'_nosubset: "\X \ P'. \ X \ B" unfolding P' by auto moreover have "\X \ P'. X \ B \ C" using Q' P' by (auto elim: partition_onE) ultimately have P'_witness: "\X \ P'. \x. x \ X \ C" using \B \ C = {}\ by fastforce obtain B' where B': "B' = \Q" by blast have Q_prop: "partition_on B' Q" using B' Q' Q'_prop partition_on_split2 mem_Collect_eq by blast have "\P' = B - B' \ C" proof have "\Q' = B \ C" "\X\Q'. \X'\Q'. X \ X' \ X \ X' = {}" using Q' unfolding partition_on_def disjoint_def by auto from this show "\P' \ B - B' \ C" unfolding P' B' Q by auto blast next show "B - B' \ C \ \P'" proof fix x assume "x \ B - B' \ C" from this obtain X where X: "x \ X" "X \ Q'" using Q' by (metis Diff_iff Un_iff mem_Collect_eq partition_on_partition_on_unique) have "\X \ Q'. X \ B \ X \ B'" unfolding B' Q by auto from this X \x \ B - B' \ C\ have "\ X \ B" using \B \ C = {}\ by auto from this \X \ Q'\ have "X \ P'" using P' by auto from this \x \ X\ show "x \ \P'" by auto qed qed from this have partition_on_P': "partition_on (B - B' \ C) P'" using partition_on_split1 Q'_prop Q' mem_Collect_eq by fastforce obtain P where P: "P = (\X. X \ C) ` P'" by blast from P partition_on_P' P'_witness have "partition_on C P" using partition_on_intersect_on_elements by auto obtain f where f: "f = (\x. if x \ B - B' then (THE X. x \ X \ X \ P') \ C else undefined)" by blast have P'_prop: "P' = (\X. X \ {x \ B - B'. f x = X}) ` P" proof { fix X assume "X \ P'" have X_subset: "X \ (B - B') \ C" using partition_on_P' \X \ P'\ by (auto elim: partition_onE) have "X = X \ C \ {x \ B - B'. f x = X \ C}" proof { fix x assume "x \ X" from this X_subset have "x \ (B - B') \ C" by auto from this have "x \ X \ C \ {xa \ B - B'. f xa = X \ C}" proof assume "x \ C" from this \x \ X\ show ?thesis by simp next assume "x \ B - B'" from partition_on_P' \x \ X\ \X \ P'\ have "(THE X. x \ X \ X \ P') = X" by (simp add: partition_on_the_part_eq) from \x \ B - B'\ this show ?thesis unfolding f by auto qed } from this show "X \ X \ C \ {x \ B - B'. f x = X \ C}" by auto next show "X \ C \ {xa \ B - B'. f xa = X \ C} \ X" proof fix x assume "x \ X \ C \ {x \ B - B'. f x = X \ C}" from this show "x \ X" proof assume "x \ X \ C" from this show ?thesis by simp next assume x_in: "x \ {x \ B - B'. f x = X \ C}" from this have ex1: "\!X. x \ X \ X \ P'" using partition_on_P' by (auto intro!: partition_on_partition_on_unique) from x_in X_subset have eq: "(THE X. x \ X \ X \ P') \ C = X \ C" unfolding f by auto from P'_nosubset \X \ P'\ have "\ X \ B" by simp from this have "X \ C \ {}" using X_subset assms(3) by blast from this obtain y where y: "y \ X \ C" by auto from this eq have y_in: "y \ (THE X. x \ X \ X \ P') \ C" by simp from y y_in have "y \ X" "y \ (THE X. x \ X \ X \ P')" by auto moreover from y have "\!X. y \ X \ X \ P'" using partition_on_P' by (simp add: partition_on_partition_on_unique) moreover have "(THE X. x \ X \ X \ P') \ P'" using ex1 by (rule the1I2) auto ultimately have "(THE X. x \ X \ X \ P') = X" using \X \ P'\ by auto from this ex1 show ?thesis by (auto intro: the1I2) qed qed qed from \X \ P'\ this have "X \ (\X. X \ {x \ B - B'. f x = X}) ` P" unfolding P by simp } from this show "P' \ (\X. X \ {x \ B - B'. f x = X}) ` P" .. next { fix x assume x_in_image: "x \ (\X. X \ {x \ B - B'. f x = X}) ` P" { fix X assume "X \ P'" have "{x \ B - B'. f x = X \ C} = {x \ B - B'. x \ X}" proof - { fix x assume "x \ B - B'" from this have ex1: "\!X. x \ X \ X \ P'" using partition_on_P' by (auto intro!: partition_on_partition_on_unique) from this have in_p: "(THE X. x \ X \ X \ P') \ P'" and x_in: "x \ (THE X. x \ X \ X \ P')" by (metis (mono_tags, lifting) theI)+ have "f x = X \ C \ (THE X. x \ X \ X \ P') \ C = X \ C" using \x \ B - B'\ unfolding f by auto also have "... \ (THE X. x \ X \ X \ P') = X" proof assume "(THE X. x \ X \ X \ P') = X" from this show "(THE X. x \ X \ X \ P') \ C = X \ C" by auto next assume "(THE X. x \ X \ X \ P') \ C = X \ C" have "(THE X. x \ X \ X \ P') \ X \ {}" using P'_witness \(THE X. x \ X \ X \ P') \ C = X \ C\ \X \ P'\ by fastforce from this show "(THE X. x \ X \ X \ P') = X" using partition_on_P'[unfolded partition_on_def disjoint_def] in_p \X \ P'\ by metis qed also have "... \ x \ X" using ex1 \X \ P'\ x_in by (auto; metis (no_types, lifting) the_equality) finally have "f x = X \ C \ x \ X" . } from this show ?thesis by auto qed moreover have "X \ B - B' \ C" using partition_on_P' \X \ P'\ by (blast elim: partition_onE) ultimately have "X \ C \ {x \ B. x \ B' \ f x = X \ C} = X" by auto } from this x_in_image have "x \ P'" unfolding P by auto } from this show "(\X. X \ {x \ B - B'. f x = X}) ` P \ P'" .. qed from partition_on_P' have f_prop: "f \ (B - B') \\<^sub>E P" unfolding f P by (auto simp add: partition_on_the_part_mem) from Q B' have "B' \ B" by auto obtain k where k: "k = card B'" by blast from \finite B\ \B' \ B\ k have k_prop: "k \ {0..card B}" by (simp add: card_mono) obtain j where j: "j = card P" by blast from j \partition_on C P\ have j_prop: "j \ {0..card C}" by (simp add: assms(2) partition_on_le_set_elements) from \partition_on C P\ j have P_prop: "partition_on C P \ card P = j" by auto from k \B' \ B\ have B'_prop: "B' \ B \ card B' = k" by auto show "Q' \ construct_partition_on B C" using j_prop k_prop P_prop B'_prop Q_prop P'_prop f_prop Q'_prop unfolding construct_partition_on_def by (auto simp del: atLeastAtMost_iff) blast qed subsection \Injectivity of the Set Construction\ lemma injectivity: assumes "B \ C = {}" assumes P: "(partition_on C P \ card P = j) \ (partition_on C P' \ card P' = j')" assumes B': "(B' \ B \ card B' = k) \ (B'' \ B \ card B'' = k')" assumes Q: "partition_on B' Q \ partition_on B'' Q'" assumes f: "f \ B - B' \\<^sub>E P \ g \ B - B'' \\<^sub>E P'" assumes P': "P'' = (\X. X \ {x \ B - B'. f x = X}) ` P \ P''' = (\X. X \ {x \ B - B''. g x = X}) ` P'" assumes eq_result: "P'' \ Q = P''' \ Q'" shows "f = g" and "Q = Q'" and "B' = B''" and "P = P'" and "j = j'" and "k = k'" proof - have P_nonempty_sets: "\X\P. \c\C. c \ X" "\X\P'. \c\C. c \ X" using P by (force elim: partition_onE)+ have 1: "\X\P''. \c\C. c \ X" "\X\P'''. \c\C. c \ X" using P' P_nonempty_sets by fastforce+ have 2: "\X\Q. \x\X. x \ C" "\X\Q'. \x\X. x \ C" using \B \ C = {}\ Q B' by (auto elim: partition_onE) from eq_result have "P'' = P'''" and "Q = Q'" by (auto dest: injectivity_union[OF _ 1 2]) from this Q show "Q = Q'" and "B' = B''" by (auto intro!: partition_on_eq_implies_eq_carrier) have subset_C: "\X\P. X \ C" "\X\P'. X \ C" using P by (auto elim: partition_onE) have eq_image: "(\X. X \ {x \ B - B'. f x = X}) ` P = (\X. X \ {x \ B - B''. g x = X}) ` P'" using P' \P'' = P'''\ by auto from this \B \ C = {}\ show "P = P'" by (auto dest: injectivity_image_union[OF _ subset_C]) have eq2: "(\X. X \ {x \ B - B'. f x = X}) ` P = (\X. X \ {x \ B - B'. g x = X}) ` P" using \P = P'\ \B' = B''\ eq_image by simp from P have P_props: "\X \ P. X \ C" "\X \ P. X \ {}" by (auto elim: partition_onE) have invert: "\X\P. (X \ {x \ B - B'. f x = X}) \ C = X \ (X \ {x \ B - B'. g x = X}) \ C = X" using \B \ C = {}\ P_props by auto have eq3: "\X \ P. (X \ {x \ B - B'. f x = X}) = (X \ {x \ B - B'. g x = X})" using injectivity_image[OF eq2 invert] by blast have eq4: "\X \ P. {x \ B - B'. f x = X} = {x \ B - B'. g x = X}" proof fix X assume "X \ P" from this P have "X \ C" by (auto elim: partition_onE) have disjoint: "X \ {x \ B - B'. f x = X} = {}" "X \ {x \ B - B'. g x = X} = {}" using \B \ C = {}\ \X \ C\ by auto from eq3 \X \ P\ have "X \ {x \ B - B'. f x = X} = X \ {x \ B - B'. g x = X}" by auto from this disjoint show "{x \ B - B'. f x = X} = {x \ B - B'. g x = X}" by (auto intro: injectivity_union) qed from eq4 f have eq5: "\b\B - B'. f b = g b" by blast from eq5 f \B' = B''\ \P = P'\ show eq6: "f = g" by (auto intro: extensional_funcset_ext) from P \P = P'\ show "j = j'" by simp from B' \B' = B''\ show "k = k'" by simp qed subsection \The Generalized Bell Recurrence Relation\ theorem Bell_eq: "Bell (n + m) = (\k\n. \j\m. j ^ (n - k) * Stirling m j * (n choose k) * Bell k)" proof - define A where "A = {0.. C" "B \ C = {}" "finite B" "card B = n" "finite C" "card C = m" unfolding A_def B_def C_def by auto have step1: "Bell (n + m) = card {P. partition_on A P}" unfolding Bell_def A_def .. from \A = B \ C\ \B \ C = {}\ \finite B\ \finite C\ have step2: "card {P. partition_on A P} = card (construct_partition_on B C)" by (simp add: construct_partition_on) note injectivity = injectivity[OF \B \ C = {}\] let ?expr = "do { k \ {0..card B}; j \ {0..card C}; P \ {P. partition_on C P \ card P = j}; B' \ {B'. B' \ B \ card B' = k}; Q \ {Q. partition_on B' Q}; f \ (B - B') \\<^sub>E P; P' \ {(\X. X \ {x \ B - B'. f x = X}) ` P}; {P' \ Q} }" let "?S \ ?comp" = ?expr { fix k assume k: "k \ {..card B}" let ?expr = "?comp k" let "?S \ ?comp" = ?expr { fix j assume "j \ {.. card C}" let ?expr = "?comp j" let "?S \ ?comp" = ?expr from \finite C\ have "finite ?S" by (intro finite_Collect_conjI disjI1 finitely_many_partition_on) { fix P assume P: "P \ {P. partition_on C P \ card P = j}" from this have "partition_on C P" by simp let ?expr = "?comp P" let "?S \ ?comp" = ?expr have "finite P" using P \finite C\ by (auto intro: finite_elements) from \finite B\ have "finite ?S" by (auto simp add: finite_subset) moreover { fix B' assume B': "B' \ {B'. B' \ B \ card B' = k}" from this have "B' \ B" by simp let ?expr = "?comp B'" let "?S \ ?comp" = ?expr from \finite B\ have "finite B'" using B' by (auto simp add: finite_subset) from \finite B'\ have "finite {Q. partition_on B' Q}" by (rule finitely_many_partition_on) moreover { fix Q assume Q: "Q \ {Q. partition_on B' Q}" let ?expr = "?comp Q" let "?S \ ?comp" = ?expr { fix f assume "f \ B - B' \\<^sub>E P" let ?expr = "?comp f" let "?S \ ?comp" = ?expr have "disjoint_family_on ?comp ?S" by (auto intro: disjoint_family_onI) from this have "card ?expr = 1" by (simp add: card_bind_constant) moreover have "finite ?expr" by (simp add: finite_bind) ultimately have "finite ?expr \ card ?expr = 1" by blast } moreover have "finite ?S" using \finite B\ \finite P\ by (auto intro: finite_PiE) moreover have "disjoint_family_on ?comp ?S" using P B' Q by (injectivity_solver rule: local.injectivity(1)) moreover have "card ?S = j ^ (n - k)" proof - have "card (B - B') = n - k" using B' \finite B'\ \card B = n\ by (subst card_Diff_subset) auto from this show ?thesis using \finite B\ P by (subst card_PiE) (simp add: prod_constant)+ qed ultimately have "card ?expr = j ^ (n - k)" by (simp add: card_bind_constant) moreover have "finite ?expr" using \finite ?S\ \finite {P. partition_on C P \ card P = j}\ by (auto intro!: finite_bind) ultimately have "finite ?expr \ card ?expr = j ^ (n - k)" by blast } note inner = this moreover have "card ?S = Bell k" using B' \finite B'\ by (auto simp add: Bell_altdef[symmetric]) moreover have "disjoint_family_on ?comp ?S" using P B' by (injectivity_solver rule: local.injectivity(2)) ultimately have "card ?expr = j ^ (n - k) * Bell k" by (subst card_bind_constant) auto moreover have "finite ?expr" using inner \finite ?S\ by (auto intro: finite_bind) ultimately have "finite ?expr \ card ?expr = j ^ (n - k) * Bell k" by blast } note inner = this moreover have "card ?S = n choose k" using \card B = n\ \finite B\ by (simp add: n_subsets) moreover have "disjoint_family_on ?comp ?S" using P by (injectivity_solver rule: local.injectivity(3)) ultimately have "card ?expr = j ^ (n - k) * (n choose k) * Bell k" by (subst card_bind_constant) auto moreover have "finite ?expr" using inner \finite ?S\ by (auto intro: finite_bind) ultimately have "finite ?expr \ card ?expr = j ^ (n - k) * (n choose k) * Bell k" by blast } note inner = this moreover note \finite ?S\ moreover have "card ?S = Stirling m j" using \finite C\ \card C = m\ by (simp add: card_partition_on) moreover have "disjoint_family_on ?comp ?S" by (injectivity_solver rule: local.injectivity(4)) ultimately have "card ?expr = j ^ (n - k) * Stirling m j * (n choose k) * Bell k" by (subst card_bind_constant) auto moreover have "finite ?expr" using inner \finite ?S\ by (auto intro: finite_bind) ultimately have "finite ?expr \ card ?expr = j ^ (n - k) * Stirling m j * (n choose k) * Bell k" by blast } note inner = this moreover have "finite ?S" by simp moreover have "disjoint_family_on ?comp ?S" by (injectivity_solver rule: local.injectivity(5)) ultimately have "card ?expr = (\j\m. j ^ (n - k) * Stirling m j * (n choose k) * Bell k)" (is "_ = ?formula") using \card C = m\ by (subst card_bind) (auto intro: sum.cong) moreover have "finite ?expr" using inner \finite ?S\ by (auto intro: finite_bind) ultimately have "finite ?expr \ card ?expr = ?formula" by blast } moreover have "finite ?S" by simp moreover have "disjoint_family_on ?comp ?S" by (injectivity_solver rule: local.injectivity(6)) ultimately have step3: "card (construct_partition_on B C) = (\k\n. \j\m. j ^ (n - k) * Stirling m j * (n choose k) * Bell k)" unfolding construct_partition_on_def using \card B = n\ by (subst card_bind) (auto intro: sum.cong) from step1 step2 step3 show ?thesis by auto qed subsection \Corollaries of the Generalized Bell Recurrence\ corollary Bell_Stirling_eq: "Bell m = (\j\m. Stirling m j)" proof - have "Bell m = Bell (0 + m)" by simp also have "... = (\j\m. Stirling m j)" unfolding Bell_eq[of 0] by (simp add: Bell_0) finally show ?thesis . qed corollary Bell_recursive_eq: "Bell (n + 1) = (\k\n. (n choose k) * Bell k)" unfolding Bell_eq[of _ 1] by simp end diff --git a/thys/Berlekamp_Zassenhaus/Missing_Multiset2.thy b/thys/Berlekamp_Zassenhaus/Missing_Multiset2.thy --- a/thys/Berlekamp_Zassenhaus/Missing_Multiset2.thy +++ b/thys/Berlekamp_Zassenhaus/Missing_Multiset2.thy @@ -1,270 +1,270 @@ theory Missing_Multiset2 - imports "HOL-Library.Permutations" + imports "HOL-Combinatorics.Permutations" Containers.Containers_Auxiliary (* only for a lemma *) begin subsubsection \Missing muiltiset\ lemma id_imp_bij: assumes id: "\x. f (f x) = x" shows "bij f" proof (intro bijI injI surjI[of f, OF id]) fix x y assume "f x = f y" then have "f (f x) = f (f y)" by auto with id show "x = y" by auto qed lemma rel_mset_Zero_iff[simp]: shows "rel_mset rel {#} Y \ Y = {#}" and "rel_mset rel X {#} \ X = {#}" using rel_mset_Zero rel_mset_size by (fastforce, fastforce) definition "is_mset_set X \ \x \# X. count X x = 1" lemma is_mset_setD[dest]: "is_mset_set X \ x \# X \ count X x = 1" unfolding is_mset_set_def by auto lemma is_mset_setI[intro]: assumes "\x. x \# X \ count X x = 1" shows "is_mset_set X" using assms unfolding is_mset_set_def by auto lemma is_mset_set[simp]: "is_mset_set (mset_set X)" unfolding is_mset_set_def by (meson count_mset_set(1) count_mset_set(2) count_mset_set(3) not_in_iff) lemma is_mset_set_add[simp]: "is_mset_set (X + {#x#}) \ is_mset_set X \ x \# X" (is "?L \ ?R") proof(intro iffI conjI) assume L: ?L with count_eq_zero_iff count_single show "is_mset_set X" unfolding is_mset_set_def by (metis (no_types, hide_lams) add_mset_add_single count_add_mset nat.inject set_mset_add_mset_insert union_single_eq_member) show "x \# X" proof assume "x \# X" then have "count (X + {#x#}) x > 1" by auto with L show False by (auto simp: is_mset_set_def) qed next assume R: ?R show ?L proof fix x' assume x': "x' \# X + {#x#}" show "count (X + {#x#}) x' = 1" proof(cases "x' \# X") case True with R have "count X x' = 1" by auto moreover from True R have "count {#x#} x' = 0" by auto ultimately show ?thesis by auto next case False then have "count X x' = 0" by (simp add: not_in_iff) with R x' show ?thesis by auto qed qed qed lemma mset_set_id[simp]: assumes "is_mset_set X" shows "mset_set (set_mset X) = X" using assms unfolding is_mset_set_def by (metis count_eq_zero_iff count_mset_set(1) count_mset_set(3) finite_set_mset multiset_eqI) lemma count_image_mset: shows "count (image_mset f X) y = (\x | x \# X \ y = f x. count X x)" proof(induct X) case empty show ?case by auto next case (add x X) define X' where "X' \ X + {#x#}" have "(\z | z \# X' \ y = f z. count (X + {#x#}) z) = (\z | z \# X' \ y = f z. count X z) + (\z | z \# X' \ y = f z. count {#x#} z)" unfolding plus_multiset.rep_eq sum.distrib.. also have split: "{z. z \# X' \ y = f z} = {z. z \# X' \ y = f z \ z \ x} \ {z. z \# X' \ y = f z \ z = x}" by blast then have "(\z | z \# X' \ y = f z. count {#x#} z) = (\z | z \# X' \ y = f z \ z = x. count {#x#} z)" unfolding split by (subst sum.union_disjoint, auto) also have "... = (if y = f x then 1 else 0)" using card_eq_Suc_0_ex1 by (auto simp: X'_def) also have "(\z | z \# X' \ y = f z. count X z) = (\z | z \# X \ y = f z. count X z)" proof(cases "x \# X") case True then have "z \# X' \ z \# X" for z by (auto simp: X'_def) then show ?thesis by auto next case False have split: "{z. z \# X' \ y = f z} = {z. z \# X \ y = f z} \ {z. z = x \ y = f z}" by (auto simp: X'_def) also have "sum (count X) ... = (\z | z \# X \ y = f z. count X z) + (\z | z = x \ y = f z. count X z)" by (subst sum.union_disjoint, auto simp: False) also with False have "\z. z = x \ y = f z \ count X z = 0" by (meson count_inI) with sum.neutral_const have "(\z | z = x \ y = f z. count X z) = 0" by auto finally show ?thesis by auto qed also have "... = count (image_mset f X) y" using add by auto finally show ?case by (simp add: X'_def) qed lemma is_mset_set_image: assumes "inj_on f (set_mset X)" and "is_mset_set X" shows "is_mset_set (image_mset f X)" proof (cases X) case empty then show ?thesis by auto next case (add x X) define X' where "X' \ add_mset x X" with assms add have inj:"inj_on f (set_mset X')" and X': "is_mset_set X'" by auto show ?thesis proof(unfold add, intro is_mset_setI, fold X'_def) fix y assume "y \# image_mset f X'" then have "y \ f ` set_mset X'" by auto with inj have "\!x'. x' \# X' \ y = f x'" by (meson imageE inj_onD) then obtain x' where x': "{x'. x' \# X' \ y = f x'} = {x'}" by auto then have "count (image_mset f X') y = count X' x'" unfolding count_image_mset by auto also from X' x' have "... = 1" by auto finally show "count (image_mset f X') y = 1". qed qed (* a variant for "right" *) lemma ex_mset_zip_right: assumes "length xs = length ys" "mset ys' = mset ys" shows "\xs'. length ys' = length xs' \ mset (zip xs' ys') = mset (zip xs ys)" using assms proof (induct xs ys arbitrary: ys' rule: list_induct2) case Nil thus ?case by auto next case (Cons x xs y ys ys') obtain j where j_len: "j < length ys'" and nth_j: "ys' ! j = y" by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD) define ysa where "ysa = take j ys' @ drop (Suc j) ys'" have "mset ys' = {#y#} + mset ysa" unfolding ysa_def using j_len nth_j by (metis Cons_nth_drop_Suc union_mset_add_mset_right add_mset_remove_trivial add_diff_cancel_left' append_take_drop_id mset.simps(2) mset_append) hence ms_y: "mset ysa = mset ys" by (simp add: Cons.prems) then obtain xsa where len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)" using Cons.hyps(2) by blast define xs' where "xs' = take j xsa @ x # drop j xsa" have ys': "ys' = take j ysa @ y # drop j ysa" using ms_y j_len nth_j Cons.prems ysa_def by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons length_drop size_mset) have j_len': "j \ length ysa" using j_len ys' ysa_def by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less) have "length ys' = length xs'" unfolding xs'_def using Cons.prems len_a ms_y by (metis add_Suc_right append_take_drop_id length_Cons length_append mset_eq_length) moreover have "mset (zip xs' ys') = mset (zip (x # xs) (y # ys))" unfolding ys' xs'_def apply (rule HOL.trans[OF mset_zip_take_Cons_drop_twice]) using j_len' by (auto simp: len_a ms_a) ultimately show ?case by blast qed lemma list_all2_reorder_right_invariance: assumes rel: "list_all2 R xs ys" and ms_y: "mset ys' = mset ys" shows "\xs'. list_all2 R xs' ys' \ mset xs' = mset xs" proof - have len: "length xs = length ys" using rel list_all2_conv_all_nth by auto obtain xs' where len': "length xs' = length ys'" and ms_xy: "mset (zip xs' ys') = mset (zip xs ys)" using len ms_y by (metis ex_mset_zip_right) have "list_all2 R xs' ys'" using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: mset_eq_setD) moreover have "mset xs' = mset xs" using len len' ms_xy map_fst_zip mset_map by metis ultimately show ?thesis by blast qed lemma rel_mset_via_perm: "rel_mset rel (mset xs) (mset ys) \ (\zs. mset xs = mset zs \ list_all2 rel zs ys)" proof (unfold rel_mset_def, intro iffI, goal_cases) case 1 then obtain zs ws where zs: "mset zs = mset xs" and ws: "mset ws = mset ys" and zsws: "list_all2 rel zs ws" by auto note list_all2_reorder_right_invariance[OF zsws ws[symmetric], unfolded zs] then show ?case by (auto dest: sym) next case 2 from this show ?case by force qed lemma rel_mset_free: assumes rel: "rel_mset rel X Y" and xs: "mset xs = X" shows "\ys. mset ys = Y \ list_all2 rel xs ys" proof- from rel[unfolded rel_mset_def] obtain xs' ys' where xs': "mset xs' = X" and ys': "mset ys' = Y" and xsys': "list_all2 rel xs' ys'" by auto from xs' xs have "mset xs = mset xs'" by auto from mset_eq_permutation[OF this] obtain f where perm: "f permutes {..i. i < length xs \ xs ! i = xs' ! f i" by auto note [simp] = list_all2_lengthD[OF xsys',symmetric] note [simp] = atLeast0LessThan[symmetric] note bij = permutes_bij[OF perm] define ys where "ys \ map (nth ys' \ f) [0..Y1 Y2. Y = Y1 + Y2 \ rel_mset rel X1 Y1 \ rel_mset rel X2 Y2" proof- obtain xs1 where xs1: "mset xs1 = X1" using ex_mset by auto obtain xs2 where xs2: "mset xs2 = X2" using ex_mset by auto from xs1 xs2 have "mset (xs1 @ xs2) = X1 + X2" by auto from rel_mset_free[OF rel this] obtain ys where ys: "mset ys = Y" "list_all2 rel (xs1 @ xs2) ys" by auto then obtain ys1 ys2 where ys12: "ys = ys1 @ ys2" and xs1ys1: "list_all2 rel xs1 ys1" and xs2ys2: "list_all2 rel xs2 ys2" using list_all2_append1 by blast from ys12 ys have "Y = mset ys1 + mset ys2" by auto moreover from xs1 xs1ys1 have "rel_mset rel X1 (mset ys1)" unfolding rel_mset_def by auto moreover from xs2 xs2ys2 have "rel_mset rel X2 (mset ys2)" unfolding rel_mset_def by auto ultimately show ?thesis by (subst exI[of _ "mset ys1"], subst exI[of _ "mset ys2"],auto) qed lemma rel_mset_OO: assumes AB: "rel_mset R A B" and BC: "rel_mset S B C" shows "rel_mset (R OO S) A C" proof- from AB obtain as bs where A_as: "A = mset as" and B_bs: "B = mset bs" and as_bs: "list_all2 R as bs" by (auto simp: rel_mset_def) from rel_mset_free[OF BC] B_bs obtain cs where C_cs: "C = mset cs" and bs_cs: "list_all2 S bs cs" by auto from list_all2_trans[OF _ as_bs bs_cs, of "R OO S"] A_as C_cs show ?thesis by (auto simp: rel_mset_def) qed end diff --git a/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy b/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy --- a/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy +++ b/thys/Berlekamp_Zassenhaus/Unique_Factorization.thy @@ -1,971 +1,971 @@ theory Unique_Factorization imports Polynomial_Interpolation.Ring_Hom_Poly Polynomial_Factorization.Polynomial_Divisibility - "HOL-Library.Permutations" + "HOL-Combinatorics.Permutations" "HOL-Computational_Algebra.Euclidean_Algorithm" Containers.Containers_Auxiliary (* only for a lemma *) Missing_Multiset2 "HOL-Algebra.Divisibility" begin hide_const(open) Divisibility.prime Divisibility.irreducible hide_fact(open) Divisibility.irreducible_def Divisibility.irreducibleI Divisibility.irreducibleD Divisibility.irreducibleE hide_const (open) Rings.coprime lemma irreducible_uminus [simp]: fixes a::"'a::idom" shows "irreducible (-a) \ irreducible a" using irreducible_mult_unit_left[of "-1::'a"] by auto context comm_monoid_mult begin definition coprime :: "'a \ 'a \ bool" where coprime_def': "coprime p q \ \r. r dvd p \ r dvd q \ r dvd 1" lemma coprimeI: assumes "\r. r dvd p \ r dvd q \ r dvd 1" shows "coprime p q" using assms by (auto simp: coprime_def') lemma coprimeE: assumes "coprime p q" and "(\r. r dvd p \ r dvd q \ r dvd 1) \ thesis" shows thesis using assms by (auto simp: coprime_def') lemma coprime_commute [ac_simps]: "coprime p q \ coprime q p" by (auto simp add: coprime_def') lemma not_coprime_iff_common_factor: "\ coprime p q \ (\r. r dvd p \ r dvd q \ \ r dvd 1)" by (auto simp add: coprime_def') end lemma (in algebraic_semidom) coprime_iff_coprime [simp, code]: "coprime = Rings.coprime" by (simp add: fun_eq_iff coprime_def coprime_def') lemma (in comm_semiring_1) coprime_0 [simp]: "coprime p 0 \ p dvd 1" "coprime 0 p \ p dvd 1" by (auto intro: coprimeI elim: coprimeE dest: dvd_trans) (**** until here ****) (* TODO: move or...? *) lemma dvd_rewrites: "dvd.dvd ((*)) = (dvd)" by (unfold dvd.dvd_def dvd_def, rule) subsection \Interfacing UFD properties\ hide_const (open) Divisibility.irreducible context comm_monoid_mult_isom begin lemma coprime_hom[simp]: "coprime (hom x) y' \ coprime x (Hilbert_Choice.inv hom y')" proof- show ?thesis by (unfold coprime_def', fold ball_UNIV, subst surj[symmetric], simp) qed lemma coprime_inv_hom[simp]: "coprime (Hilbert_Choice.inv hom x') y \ coprime x' (hom y)" proof- interpret inv: comm_monoid_mult_isom "Hilbert_Choice.inv hom".. show ?thesis by simp qed end subsubsection \Original part\ lemma dvd_dvd_imp_smult: fixes p q :: "'a :: idom poly" assumes pq: "p dvd q" and qp: "q dvd p" shows "\c. p = smult c q" proof (cases "p = 0") case True then show ?thesis by auto next case False from qp obtain r where r: "p = q * r" by (elim dvdE, auto) with False qp have r0: "r \ 0" and q0: "q \ 0" by auto with divides_degree[OF pq] divides_degree[OF qp] False have "degree p = degree q" by auto with r degree_mult_eq[OF q0 r0] have "degree r = 0" by auto from degree_0_id[OF this] obtain c where "r = [:c:]" by metis from r[unfolded this] show ?thesis by auto qed lemma dvd_const: assumes pq: "(p::'a::semidom poly) dvd q" and q0: "q \ 0" and degq: "degree q = 0" shows "degree p = 0" proof- from dvdE[OF pq] obtain r where *: "q = p * r". with q0 have "p \ 0" "r \ 0" by auto from degree_mult_eq[OF this] degq * show "degree p = 0" by auto qed context Rings.dvd begin abbreviation ddvd (infix "ddvd" 40) where "x ddvd y \ x dvd y \ y dvd x" lemma ddvd_sym[sym]: "x ddvd y \ y ddvd x" by auto end context comm_monoid_mult begin lemma ddvd_trans[trans]: "x ddvd y \ y ddvd z \ x ddvd z" using dvd_trans by auto lemma ddvd_transp: "transp (ddvd)" by (intro transpI, fact ddvd_trans) end context comm_semiring_1 begin definition mset_factors where "mset_factors F p \ F \ {#} \ (\f. f \# F \ irreducible f) \ p = prod_mset F" lemma mset_factorsI[intro!]: assumes "\f. f \# F \ irreducible f" and "F \ {#}" and "prod_mset F = p" shows "mset_factors F p" unfolding mset_factors_def using assms by auto lemma mset_factorsD: assumes "mset_factors F p" shows "f \# F \ irreducible f" and "F \ {#}" and "prod_mset F = p" using assms[unfolded mset_factors_def] by auto lemma mset_factorsE[elim]: assumes "mset_factors F p" and "(\f. f \# F \ irreducible f) \ F \ {#} \ prod_mset F = p \ thesis" shows thesis using assms[unfolded mset_factors_def] by auto lemma mset_factors_imp_not_is_unit: assumes "mset_factors F p" shows "\ p dvd 1" proof(cases F) case empty with assms show ?thesis by auto next case (add f F) with assms have "\ f dvd 1" "p = f * prod_mset F" by (auto intro!: irreducible_not_unit) then show ?thesis by auto qed definition primitive_poly where "primitive_poly f \ \d. (\i. d dvd coeff f i) \ d dvd 1" end lemma(in semidom) mset_factors_imp_nonzero: assumes "mset_factors F p" shows "p \ 0" proof assume "p = 0" moreover from assms have "prod_mset F = p" by auto ultimately obtain f where "f \# F" "f = 0" by auto with assms show False by auto qed class ufd = idom + assumes mset_factors_exist: "\x. x \ 0 \ \ x dvd 1 \ \F. mset_factors F x" and mset_factors_unique: "\x F G. mset_factors F x \ mset_factors G x \ rel_mset (ddvd) F G" subsubsection \Connecting to HOL/Divisibility\ context comm_semiring_1 begin abbreviation "mk_monoid \ \carrier = UNIV - {0}, mult = (*), one = 1\" lemma carrier_0[simp]: "x \ carrier mk_monoid \ x \ 0" by auto lemmas mk_monoid_simps = carrier_0 monoid.simps abbreviation irred where "irred \ Divisibility.irreducible mk_monoid" abbreviation factor where "factor \ Divisibility.factor mk_monoid" abbreviation factors where "factors \ Divisibility.factors mk_monoid" abbreviation properfactor where "properfactor \ Divisibility.properfactor mk_monoid" lemma factors: "factors fs y \ prod_list fs = y \ Ball (set fs) irred" proof - have "prod_list fs = foldr (*) fs 1" by (induct fs, auto) thus ?thesis unfolding factors_def by auto qed lemma factor: "factor x y \ (\z. z \ 0 \ x * z = y)" unfolding factor_def by auto lemma properfactor_nz: shows "(y :: 'a) \ 0 \ properfactor x y \ x dvd y \ \ y dvd x" by (auto simp: properfactor_def factor_def dvd_def) lemma mem_Units[simp]: "y \ Units mk_monoid \ y dvd 1" unfolding dvd_def Units_def by (auto simp: ac_simps) end context idom begin lemma irred_0[simp]: "irred (0::'a)" by (unfold Divisibility.irreducible_def, auto simp: factor properfactor_def) lemma factor_idom[simp]: "factor (x::'a) y \ (if y = 0 then x = 0 else x dvd y)" by (cases "y = 0"; auto intro: exI[of _ 1] elim: dvdE simp: factor) lemma associated_connect[simp]: "(\\<^bsub>mk_monoid\<^esub>) = (ddvd)" by (intro ext, unfold associated_def, auto) lemma essentially_equal_connect[simp]: "essentially_equal mk_monoid fs gs \ rel_mset (ddvd) (mset fs) (mset gs)" by (auto simp: essentially_equal_def rel_mset_via_perm perm_iff_eq_mset) lemma irred_idom_nz: assumes x0: "(x::'a) \ 0" shows "irred x \ irreducible x" using x0 by (auto simp: irreducible_altdef Divisibility.irreducible_def properfactor_nz) lemma dvd_dvd_imp_unit_mult: assumes xy: "x dvd y" and yx: "y dvd x" shows "\z. z dvd 1 \ y = x * z" proof(cases "x = 0") case True with xy show ?thesis by (auto intro: exI[of _ 1]) next case x0: False from xy obtain z where z: "y = x * z" by (elim dvdE, auto) from yx obtain w where w: "x = y * w" by (elim dvdE, auto) from z w have "x * (z * w) = x" by (auto simp: ac_simps) then have "z * w = 1" using x0 by auto with z show ?thesis by (auto intro: exI[of _ z]) qed lemma irred_inner_nz: assumes x0: "x \ 0" shows "(\b. b dvd x \ \ x dvd b \ b dvd 1) \ (\a b. x = a * b \ a dvd 1 \ b dvd 1)" (is "?l \ ?r") proof (intro iffI allI impI) assume l: ?l fix a b assume xab: "x = a * b" then have ax: "a dvd x" and bx: "b dvd x" by auto { assume a1: "\ a dvd 1" with l ax have xa: "x dvd a" by auto from dvd_dvd_imp_unit_mult[OF ax xa] obtain z where z1: "z dvd 1" and xaz: "x = a * z" by auto from xab x0 have "a \ 0" by auto with xab xaz have "b = z" by auto with z1 have "b dvd 1" by auto } then show "a dvd 1 \ b dvd 1" by auto next assume r: ?r fix b assume bx: "b dvd x" and xb: "\ x dvd b" then obtain a where xab: "x = a * b" by (elim dvdE, auto simp: ac_simps) with r consider "a dvd 1" | "b dvd 1" by auto then show "b dvd 1" proof(cases) case 2 then show ?thesis by auto next case 1 then obtain c where ac1: "a * c = 1" by (elim dvdE, auto) from xab have "x * c = b * (a * c)" by (auto simp: ac_simps) with ac1 have "x * c = b" by auto then have "x dvd b" by auto with xb show ?thesis by auto qed qed lemma irred_idom[simp]: "irred x \ x = 0 \ irreducible x" by (cases "x = 0"; simp add: irred_idom_nz irred_inner_nz irreducible_def) lemma assumes "x \ 0" and "factors fs x" and "f \ set fs" shows "f \ 0" using assms by (auto simp: factors) lemma factors_as_mset_factors: assumes x0: "x \ 0" and x1: "x \ 1" shows "factors fs x \ mset_factors (mset fs) x" using assms by (auto simp: factors prod_mset_prod_list) end context ufd begin interpretation comm_monoid_cancel: comm_monoid_cancel "mk_monoid::'a monoid" apply (unfold_locales) apply simp_all using mult_left_cancel apply (auto simp: ac_simps) done lemma factors_exist: assumes "a \ 0" and "\ a dvd 1" shows "\fs. set fs \ UNIV - {0} \ factors fs a" proof- from mset_factors_exist[OF assms] obtain F where "mset_factors F a" by auto also from ex_mset obtain fs where "F = mset fs" by metis finally have fs: "mset_factors (mset fs) a". then have "factors fs a" using assms by (subst factors_as_mset_factors, auto) moreover have "set fs \ UNIV - {0}" using fs by (auto elim!: mset_factorsE) ultimately show ?thesis by auto qed lemma factors_unique: assumes fs: "factors fs a" and gs: "factors gs a" and a0: "a \ 0" and a1: "\ a dvd 1" shows "rel_mset (ddvd) (mset fs) (mset gs)" proof- from a1 have "a \ 1" by auto with a0 fs gs have "mset_factors (mset fs) a" "mset_factors (mset gs) a" by (unfold factors_as_mset_factors) from mset_factors_unique[OF this] show ?thesis. qed lemma factorial_monoid: "factorial_monoid (mk_monoid :: 'a monoid)" by (unfold_locales; auto simp add: factors_exist factors_unique) end lemma (in idom) factorial_monoid_imp_ufd: assumes "factorial_monoid (mk_monoid :: 'a monoid)" shows "class.ufd ((*) :: 'a \ _) 1 (+) 0 (-) uminus" proof (unfold_locales) interpret factorial_monoid "mk_monoid :: 'a monoid" by (fact assms) { fix x assume x: "x \ 0" "\ x dvd 1" note * = factors_exist[simplified, OF this] with x show "\F. mset_factors F x" by (subst(asm) factors_as_mset_factors, auto) } fix x F G assume FG: "mset_factors F x" "mset_factors G x" with mset_factors_imp_not_is_unit have x1: "\ x dvd 1" by auto from FG(1) have x0: "x \ 0" by (rule mset_factors_imp_nonzero) obtain fs gs where fsgs: "F = mset fs" "G = mset gs" using ex_mset by metis note FG = FG[unfolded this] then have 0: "0 \ set fs" "0 \ set gs" by (auto elim!: mset_factorsE) from x1 have "x \ 1" by auto note FG[folded factors_as_mset_factors[OF x0 this]] from factors_unique[OF this, simplified, OF x0 x1, folded fsgs] 0 show "rel_mset (ddvd) F G" by auto qed subsection \Preservation of Irreducibility\ locale comm_semiring_1_hom = comm_monoid_mult_hom hom + zero_hom hom for hom :: "'a :: comm_semiring_1 \ 'b :: comm_semiring_1" locale irreducibility_hom = comm_semiring_1_hom + assumes irreducible_imp_irreducible_hom: "irreducible a \ irreducible (hom a)" begin lemma hom_mset_factors: assumes F: "mset_factors F p" shows "mset_factors (image_mset hom F) (hom p)" proof (unfold mset_factors_def, intro conjI allI impI) from F show "hom p = prod_mset (image_mset hom F)" "image_mset hom F \ {#}" by (auto simp: hom_distribs) fix f' assume "f' \# image_mset hom F" then obtain f where f: "f \# F" and f'f: "f' = hom f" by auto with F irreducible_imp_irreducible_hom show "irreducible f'" unfolding f'f by auto qed end locale unit_preserving_hom = comm_semiring_1_hom + assumes is_unit_hom_if: "\x. hom x dvd 1 \ x dvd 1" begin lemma is_unit_hom_iff[simp]: "hom x dvd 1 \ x dvd 1" using is_unit_hom_if hom_dvd by force lemma irreducible_hom_imp_irreducible: assumes irr: "irreducible (hom a)" shows "irreducible a" proof (intro irreducibleI) from irr show "a \ 0" by auto from irr show "\ a dvd 1" by (auto dest: irreducible_not_unit) fix b c assume "a = b * c" then have "hom a = hom b * hom c" by (simp add: hom_distribs) with irr have "hom b dvd 1 \ hom c dvd 1" by (auto dest: irreducibleD) then show "b dvd 1 \ c dvd 1" by simp qed end locale factor_preserving_hom = unit_preserving_hom + irreducibility_hom begin lemma irreducible_hom[simp]: "irreducible (hom a) \ irreducible a" using irreducible_hom_imp_irreducible irreducible_imp_irreducible_hom by metis end lemma factor_preserving_hom_comp: assumes f: "factor_preserving_hom f" and g: "factor_preserving_hom g" shows "factor_preserving_hom (f o g)" proof- interpret f: factor_preserving_hom f by (rule f) interpret g: factor_preserving_hom g by (rule g) show ?thesis by (unfold_locales, auto simp: hom_distribs) qed context comm_semiring_isom begin sublocale unit_preserving_hom by (unfold_locales, auto) sublocale factor_preserving_hom proof (standard) fix a :: 'a assume "irreducible a" note a = this[unfolded irreducible_def] show "irreducible (hom a)" proof (rule ccontr) assume "\ irreducible (hom a)" from this[unfolded Factorial_Ring.irreducible_def,simplified] a obtain hb hc where eq: "hom a = hb * hc" and nu: "\ hb dvd 1" "\ hc dvd 1" by auto from bij obtain b where hb: "hb = hom b" by (elim bij_pointE) from bij obtain c where hc: "hc = hom c" by (elim bij_pointE) from eq[unfolded hb hc, folded hom_mult] have "a = b * c" by auto with nu hb hc have "a = b * c" "\ b dvd 1" "\ c dvd 1" by auto with a show False by auto qed qed end subsubsection\Back to divisibility\ lemma(in comm_semiring_1) mset_factors_mult: assumes F: "mset_factors F a" and G: "mset_factors G b" shows "mset_factors (F+G) (a*b)" proof(intro mset_factorsI) fix f assume "f \# F + G" then consider "f \# F" | "f \# G" by auto then show "irreducible f" by(cases, insert F G, auto) qed (insert F G, auto) lemma(in ufd) dvd_imp_subset_factors: assumes ab: "a dvd b" and F: "mset_factors F a" and G: "mset_factors G b" shows "\G'. G' \# G \ rel_mset (ddvd) F G'" proof- from F G have a0: "a \ 0" and b0: "b \ 0" by (simp_all add: mset_factors_imp_nonzero) from ab obtain c where c: "b = a * c" by (elim dvdE, auto) with b0 have c0: "c \ 0" by auto show ?thesis proof(cases "c dvd 1") case True show ?thesis proof(cases F) case empty with F show ?thesis by auto next case (add f F') with F have a: "f * prod_mset F' = a" and F': "\f. f \# F' \ irreducible f" and irrf: "irreducible f" by auto from irrf have f0: "f \ 0" and f1: "\f dvd 1" by (auto dest: irreducible_not_unit) from a c have "(f * c) * prod_mset F' = b" by (auto simp: ac_simps) moreover { have "irreducible (f * c)" using True irrf by (subst irreducible_mult_unit_right) with F' irrf have "\f'. f' \# F' + {#f * c#} \ irreducible f'" by auto } ultimately have "mset_factors (F' + {#f * c#}) b" by (intro mset_factorsI, auto) from mset_factors_unique[OF this G] have F'G: "rel_mset (ddvd) (F' + {#f * c#}) G". from True add have FF': "rel_mset (ddvd) F (F' + {#f * c#})" by (auto simp add: multiset.rel_refl intro!: rel_mset_Plus) have "rel_mset (ddvd) F G" apply(rule transpD[OF multiset.rel_transp[OF transpI] FF' F'G]) using ddvd_trans. then show ?thesis by auto qed next case False from mset_factors_exist[OF c0 this] obtain H where H: "mset_factors H c" by auto from c mset_factors_mult[OF F H] have "mset_factors (F + H) b" by auto note mset_factors_unique[OF this G] from rel_mset_split[OF this] obtain G1 G2 where "G = G1 + G2" "rel_mset (ddvd) F G1" "rel_mset (ddvd) H G2" by auto then show ?thesis by (intro exI[of _ "G1"], auto) qed qed lemma(in idom) irreducible_factor_singleton: assumes a: "irreducible a" shows "mset_factors F a \ F = {#a#}" proof(cases F) case empty with mset_factorsD show ?thesis by auto next case (add f F') show ?thesis proof assume F: "mset_factors F a" from add mset_factorsD[OF F] have *: "a = f * prod_mset F'" by auto then have fa: "f dvd a" by auto from * a have f0: "f \ 0" by auto from add have "f \# F" by auto with F have f: "irreducible f" by auto from add have "F' \# F" by auto then have unitemp: "prod_mset F' dvd 1 \ F' = {#}" proof(induct F') case empty then show ?case by auto next case (add f F') from add have "f \# F" by (simp add: mset_subset_eq_insertD) with F irreducible_not_unit have "\ f dvd 1" by auto then have "\ (prod_mset F' * f) dvd 1" by simp with add show ?case by auto qed show "F = {#a#}" proof(cases "a dvd f") case True then obtain r where "f = a * r" by (elim dvdE, auto) with * have "f = (r * prod_mset F') * f" by (auto simp: ac_simps) with f0 have "r * prod_mset F' = 1" by auto then have "prod_mset F' dvd 1" by (metis dvd_triv_right) with unitemp * add show ?thesis by auto next case False with fa a f show ?thesis by (auto simp: irreducible_altdef) qed qed (insert a, auto) qed lemma(in ufd) irreducible_dvd_imp_factor: assumes ab: "a dvd b" and a: "irreducible a" and G: "mset_factors G b" shows "\g \# G. a ddvd g" proof- from a have "mset_factors {#a#} a" by auto from dvd_imp_subset_factors[OF ab this G] obtain G' where G'G: "G' \# G" and rel: "rel_mset (ddvd) {#a#} G'" by auto with rel_mset_size size_1_singleton_mset size_single obtain g where gG': "G' = {#g#}" by fastforce from rel[unfolded this rel_mset_def] have "a ddvd g" by auto with gG' G'G show ?thesis by auto qed lemma(in idom) prod_mset_remove_units: "prod_mset F ddvd prod_mset {# f \# F. \f dvd 1 #}" proof(induct F) case (add f F) then show ?case by (cases "f = 0", auto) qed auto lemma(in comm_semiring_1) mset_factors_imp_dvd: assumes "mset_factors F x" and "f \# F" shows "f dvd x" using assms by (simp add: dvd_prod_mset mset_factors_def) lemma(in ufd) prime_elem_iff_irreducible[iff]: "prime_elem x \ irreducible x" proof (intro iffI, fact prime_elem_imp_irreducible, rule prime_elemI) assume r: "irreducible x" then show x0: "x \ 0" and x1: "\ x dvd 1" by (auto dest: irreducible_not_unit) from irreducible_factor_singleton[OF r] have *: "mset_factors {#x#} x" by auto fix a b assume "x dvd a * b" then obtain c where abxc: "a * b = x * c" by (elim dvdE, auto) show "x dvd a \ x dvd b" proof(cases "c = 0 \ a = 0 \ b = 0") case True with abxc show ?thesis by auto next case False then have a0: "a \ 0" and b0: "b \ 0" and c0: "c \ 0" by auto from x0 c0 have xc0: "x * c \ 0" by auto from x1 have xc1: "\ x * c dvd 1" by auto show ?thesis proof (cases "a dvd 1 \ b dvd 1") case False then have a1: "\ a dvd 1" and b1: "\ b dvd 1" by auto from mset_factors_exist[OF a0 a1] obtain F where Fa: "mset_factors F a" by auto then have F0: "F \ {#}" by auto from mset_factors_exist[OF b0 b1] obtain G where Gb: "mset_factors G b" by auto then have G0: "G \ {#}" by auto from mset_factors_mult[OF Fa Gb] have FGxc: "mset_factors (F + G) (x * c)" by (simp add: abxc) show ?thesis proof (cases "c dvd 1") case True from r irreducible_mult_unit_right[OF this] have "irreducible (x*c)" by simp note irreducible_factor_singleton[OF this] FGxc with F0 G0 have False by (cases F; cases G; auto) then show ?thesis by auto next case False from mset_factors_exist[OF c0 this] obtain H where "mset_factors H c" by auto with * have xHxc: "mset_factors (add_mset x H) (x * c)" by force note rel = mset_factors_unique[OF this FGxc] obtain hs where "mset hs = H" using ex_mset by auto then have "mset (x#hs) = add_mset x H" by auto from rel_mset_free[OF rel this] obtain jjs where jjsGH: "mset jjs = F + G" and rel: "list_all2 (ddvd) (x # hs) jjs" by auto then obtain j js where jjs: "jjs = j # js" by (cases jjs, auto) with rel have xj: "x ddvd j" by auto from jjs jjsGH have j: "j \ set_mset (F + G)" by (intro union_single_eq_member, auto) from j consider "j \# F" | "j \# G" by auto then show ?thesis proof(cases) case 1 with Fa have "j dvd a" by (auto intro: mset_factors_imp_dvd) with xj dvd_trans have "x dvd a" by auto then show ?thesis by auto next case 2 with Gb have "j dvd b" by (auto intro: mset_factors_imp_dvd) with xj dvd_trans have "x dvd b" by auto then show ?thesis by auto qed qed next case True then consider "a dvd 1" | "b dvd 1" by auto then show ?thesis proof(cases) case 1 then obtain d where ad: "a * d = 1" by (elim dvdE, auto) from abxc have "x * (c * d) = a * b * d" by (auto simp: ac_simps) also have "... = a * d * b" by (auto simp: ac_simps) finally have "x dvd b" by (intro dvdI, auto simp: ad) then show ?thesis by auto next case 2 then obtain d where bd: "b * d = 1" by (elim dvdE, auto) from abxc have "x * (c * d) = a * b * d" by (auto simp: ac_simps) also have "... = (b * d) * a" by (auto simp: ac_simps) finally have "x dvd a" by (intro dvdI, auto simp:bd) then show ?thesis by auto qed qed qed qed subsection\Results for GCDs etc.\ lemma prod_list_remove1: "(x :: 'b :: comm_monoid_mult) \ set xs \ prod_list (remove1 x xs) * x = prod_list xs" by (induct xs, auto simp: ac_simps) (* Isabelle 2015-style and generalized gcd-class without normalization and factors *) class comm_monoid_gcd = gcd + comm_semiring_1 + assumes gcd_dvd1[iff]: "gcd a b dvd a" and gcd_dvd2[iff]: "gcd a b dvd b" and gcd_greatest: "c dvd a \ c dvd b \ c dvd gcd a b" begin lemma gcd_0_0[simp]: "gcd 0 0 = 0" using gcd_greatest[OF dvd_0_right dvd_0_right, of 0] by auto lemma gcd_zero_iff[simp]: "gcd a b = 0 \ a = 0 \ b = 0" proof assume "gcd a b = 0" from gcd_dvd1[of a b, unfolded this] gcd_dvd2[of a b, unfolded this] show "a = 0 \ b = 0" by auto qed auto lemma gcd_zero_iff'[simp]: "0 = gcd a b \ a = 0 \ b = 0" using gcd_zero_iff by metis lemma dvd_gcd_0_iff[simp]: shows "x dvd gcd 0 a \ x dvd a" (is ?g1) and "x dvd gcd a 0 \ x dvd a" (is ?g2) proof- have "a dvd gcd a 0" "a dvd gcd 0 a" by (auto intro: gcd_greatest) with dvd_refl show ?g1 ?g2 by (auto dest: dvd_trans) qed lemma gcd_dvd_1[simp]: "gcd a b dvd 1 \ coprime a b" using dvd_trans[OF gcd_greatest[of _ a b], of _ 1] by (cases "a = 0 \ b = 0") (auto intro!: coprimeI elim: coprimeE) lemma dvd_imp_gcd_dvd_gcd: "b dvd c \ gcd a b dvd gcd a c" by (meson gcd_dvd1 gcd_dvd2 gcd_greatest dvd_trans) definition listgcd :: "'a list \ 'a" where "listgcd xs = foldr gcd xs 0" lemma listgcd_simps[simp]: "listgcd [] = 0" "listgcd (x # xs) = gcd x (listgcd xs)" by (auto simp: listgcd_def) lemma listgcd: "x \ set xs \ listgcd xs dvd x" proof (induct xs) case (Cons y ys) show ?case proof (cases "x = y") case False with Cons have dvd: "listgcd ys dvd x" by auto thus ?thesis unfolding listgcd_simps using dvd_trans by blast next case True thus ?thesis unfolding listgcd_simps using dvd_trans by blast qed qed simp lemma listgcd_greatest: "(\ x. x \ set xs \ y dvd x) \ y dvd listgcd xs" by (induct xs arbitrary:y, auto intro: gcd_greatest) end context Rings.dvd begin definition "is_gcd x a b \ x dvd a \ x dvd b \ (\y. y dvd a \ y dvd b \ y dvd x)" definition "some_gcd a b \ SOME x. is_gcd x a b" lemma is_gcdI[intro!]: assumes "x dvd a" "x dvd b" "\y. y dvd a \ y dvd b \ y dvd x" shows "is_gcd x a b" by (insert assms, auto simp: is_gcd_def) lemma is_gcdE[elim!]: assumes "is_gcd x a b" and "x dvd a \ x dvd b \ (\y. y dvd a \ y dvd b \ y dvd x) \ thesis" shows thesis by (insert assms, auto simp: is_gcd_def) lemma is_gcd_some_gcdI: assumes "\x. is_gcd x a b" shows "is_gcd (some_gcd a b) a b" by (unfold some_gcd_def, rule someI_ex[OF assms]) end context comm_semiring_1 begin lemma some_gcd_0[intro!]: "is_gcd (some_gcd a 0) a 0" "is_gcd (some_gcd 0 b) 0 b" by (auto intro!: is_gcd_some_gcdI intro: exI[of _ a] exI[of _ b]) lemma some_gcd_0_dvd[intro!]: "some_gcd a 0 dvd a" "some_gcd 0 b dvd b" using some_gcd_0 by auto lemma dvd_some_gcd_0[intro!]: "a dvd some_gcd a 0" "b dvd some_gcd 0 b" using some_gcd_0[of a] some_gcd_0[of b] by auto end context idom begin lemma is_gcd_connect: assumes "a \ 0" "b \ 0" shows "isgcd mk_monoid x a b \ is_gcd x a b" using assms by (force simp: isgcd_def) lemma some_gcd_connect: assumes "a \ 0" and "b \ 0" shows "somegcd mk_monoid a b = some_gcd a b" using assms by (auto intro!: arg_cong[of _ _ Eps] simp: is_gcd_connect some_gcd_def somegcd_def) end context comm_monoid_gcd begin lemma is_gcd_gcd: "is_gcd (gcd a b) a b" using gcd_greatest by auto lemma is_gcd_some_gcd: "is_gcd (some_gcd a b) a b" by (insert is_gcd_gcd, auto intro!: is_gcd_some_gcdI) lemma gcd_dvd_some_gcd: "gcd a b dvd some_gcd a b" using is_gcd_some_gcd by auto lemma some_gcd_dvd_gcd: "some_gcd a b dvd gcd a b" using is_gcd_some_gcd by (auto intro: gcd_greatest) lemma some_gcd_ddvd_gcd: "some_gcd a b ddvd gcd a b" by (auto intro: gcd_dvd_some_gcd some_gcd_dvd_gcd) lemma some_gcd_dvd: "some_gcd a b dvd d \ gcd a b dvd d" "d dvd some_gcd a b \ d dvd gcd a b" using some_gcd_ddvd_gcd[of a b] by (auto dest:dvd_trans) end class idom_gcd = comm_monoid_gcd + idom begin interpretation raw: comm_monoid_cancel "mk_monoid :: 'a monoid" by (unfold_locales, auto intro: mult_commute mult_assoc) interpretation raw: gcd_condition_monoid "mk_monoid :: 'a monoid" by (unfold_locales, auto simp: is_gcd_connect intro!: exI[of _ "gcd _ _"] dest: gcd_greatest) lemma gcd_mult_ddvd: "d * gcd a b ddvd gcd (d * a) (d * b)" proof (cases "d = 0") case True then show ?thesis by auto next case d0: False show ?thesis proof (cases "a = 0 \ b = 0") case False note some_gcd_ddvd_gcd[of a b] with d0 have "d * gcd a b ddvd d * some_gcd a b" by auto also have "d * some_gcd a b ddvd some_gcd (d * a) (d * b)" using False d0 raw.gcd_mult by (simp add: some_gcd_connect) also note some_gcd_ddvd_gcd finally show ?thesis. next case True with d0 show ?thesis apply (elim disjE) apply (rule ddvd_trans[of _ "d * b"]; force) apply (rule ddvd_trans[of _ "d * a"]; force) done qed qed lemma gcd_greatest_mult: assumes cad: "c dvd a * d" and cbd: "c dvd b * d" shows "c dvd gcd a b * d" proof- from gcd_greatest[OF assms] have c: "c dvd gcd (d * a) (d * b)" by (auto simp: ac_simps) note gcd_mult_ddvd[of d a b] then have "gcd (d * a) (d * b) dvd gcd a b * d" by (auto simp: ac_simps) from dvd_trans[OF c this] show ?thesis . qed lemma listgcd_greatest_mult: "(\ x :: 'a. x \ set xs \ y dvd x * z) \ y dvd listgcd xs * z" proof (induct xs) case (Cons x xs) from Cons have "y dvd x * z" "y dvd listgcd xs * z" by auto thus ?case unfolding listgcd_simps by (rule gcd_greatest_mult) qed (simp) lemma dvd_factor_mult_gcd: assumes dvd: "k dvd p * q" "k dvd p * r" and q0: "q \ 0" and r0: "r \ 0" shows "k dvd p * gcd q r" proof - from dvd gcd_greatest[of k "p * q" "p * r"] have "k dvd gcd (p * q) (p * r)" by simp also from gcd_mult_ddvd[of p q r] have "... dvd (p * gcd q r)" by auto finally show ?thesis . qed lemma coprime_mult_cross_dvd: assumes coprime: "coprime p q" and eq: "p' * p = q' * q" shows "p dvd q'" (is ?g1) and "q dvd p'" (is ?g2) proof (atomize(full), cases "p = 0 \ q = 0") case True then show "?g1 \ ?g2" proof assume p0: "p = 0" with coprime have "q dvd 1" by auto with eq p0 show ?thesis by auto next assume q0: "q = 0" with coprime have "p dvd 1" by auto with eq q0 show ?thesis by auto qed next case False { fix p q r p' q' :: 'a assume cop: "coprime p q" and eq: "p' * p = q' * q" and p: "p \ 0" and q: "q \ 0" and r: "r dvd p" "r dvd q" let ?gcd = "gcd q p" from eq have "p' * p dvd q' * q" by auto hence d1: "p dvd q' * q" by (rule dvd_mult_right) have d2: "p dvd q' * p" by auto from dvd_factor_mult_gcd[OF d1 d2 q p] have 1: "p dvd q' * ?gcd" . from q p have 2: "?gcd dvd q" by auto from q p have 3: "?gcd dvd p" by auto from cop[unfolded coprime_def', rule_format, OF 3 2] have "?gcd dvd 1" . from 1 dvd_mult_unit_iff[OF this] have "p dvd q'" by auto } note main = this from main[OF coprime eq,of 1] False coprime coprime_commute main[OF _ eq[symmetric], of 1] show "?g1 \ ?g2" by auto qed end subclass (in ring_gcd) idom_gcd by (unfold_locales, auto) lemma coprime_rewrites: "comm_monoid_mult.coprime ((*)) 1 = coprime" apply (intro ext) apply (subst comm_monoid_mult.coprime_def') apply (unfold_locales) apply (unfold dvd_rewrites) apply (fold coprime_def') .. (* TODO: incorporate into the default class hierarchy *) locale gcd_condition = fixes ty :: "'a :: idom itself" assumes gcd_exists: "\a b :: 'a. \x. is_gcd x a b" begin sublocale idom_gcd "(*)" "1 :: 'a" "(+)" 0 "(-)" uminus some_gcd rewrites "dvd.dvd ((*)) = (dvd)" and "comm_monoid_mult.coprime ((*) ) 1 = Unique_Factorization.coprime" proof- have "is_gcd (some_gcd a b) a b" for a b :: 'a by (intro is_gcd_some_gcdI gcd_exists) from this[unfolded is_gcd_def] show "class.idom_gcd (*) (1 :: 'a) (+) 0 (-) uminus some_gcd" by (unfold_locales, auto simp: dvd_rewrites) qed (simp_all add: dvd_rewrites coprime_rewrites) end instance semiring_gcd \ comm_monoid_gcd by (intro_classes, auto) lemma listgcd_connect: "listgcd = gcd_list" proof (intro ext) fix xs :: "'a list" show "listgcd xs = gcd_list xs" by(induct xs, auto) qed interpretation some_gcd: gcd_condition "TYPE('a::ufd)" proof(unfold_locales, intro exI) interpret factorial_monoid "mk_monoid :: 'a monoid" by (fact factorial_monoid) note d = dvd.dvd_def some_gcd_def carrier_0 fix a b :: 'a show "is_gcd (some_gcd a b) a b" proof (cases "a = 0 \ b = 0") case True thus ?thesis using some_gcd_0 by auto next case False with gcdof_exists[of a b] show ?thesis by (auto intro!: is_gcd_some_gcdI simp add: is_gcd_connect some_gcd_connect) qed qed lemma some_gcd_listgcd_dvd_listgcd: "some_gcd.listgcd xs dvd listgcd xs" by (induct xs, auto simp:some_gcd_dvd intro:dvd_imp_gcd_dvd_gcd) lemma listgcd_dvd_some_gcd_listgcd: "listgcd xs dvd some_gcd.listgcd xs" by (induct xs, auto simp:some_gcd_dvd intro:dvd_imp_gcd_dvd_gcd) context factorial_ring_gcd begin text \Do not declare the following as subclass, to avoid conflict in \field \ gcd_condition\ vs. \factorial_ring_gcd \ gcd_condition\. \ sublocale as_ufd: ufd proof(unfold_locales, goal_cases) case (1 x) from prime_factorization_exists[OF \x \ 0\] obtain F where f: "\f. f \# F \ prime_elem f" and Fx: "normalize (prod_mset F) = normalize x" by auto from associatedE2[OF Fx] obtain u where u: "is_unit u" "x = u * prod_mset F" by blast from \\ is_unit x\ Fx have "F \ {#}" by auto then obtain g G where F: "F = add_mset g G" by (cases F, auto) then have "g \# F" by auto with f[OF this]prime_elem_iff_irreducible irreducible_mult_unit_left[OF unit_factor_is_unit[OF \x \ 0\]] have g: "irreducible (u * g)" using u(1) by (subst irreducible_mult_unit_left) simp_all show ?case proof (intro exI conjI mset_factorsI) show "prod_mset (add_mset (u * g) G) = x" using \x \ 0\ by (simp add: F ac_simps u) fix f assume "f \# add_mset (u * g) G" with f[unfolded F] g prime_elem_iff_irreducible show "irreducible f" by auto qed auto next case (2 x F G) note transpD[OF multiset.rel_transp[OF ddvd_transp],trans] obtain fs where F: "F = mset fs" by (metis ex_mset) have "list_all2 (ddvd) fs (map normalize fs)" by (intro list_all2_all_nthI, auto) then have FH: "rel_mset (ddvd) F (image_mset normalize F)" by (unfold rel_mset_def F, force) also have FG: "image_mset normalize F = image_mset normalize G" proof (intro prime_factorization_unique'') from 2 have xF: "x = prod_mset F" and xG: "x = prod_mset G" by auto from xF have "normalize x = normalize (prod_mset (image_mset normalize F))" by (simp add: normalize_prod_mset_normalize) with xG have nFG: "\ = normalize (prod_mset (image_mset normalize G))" by (simp_all add: normalize_prod_mset_normalize) then show "normalize (\i\#image_mset normalize F. i) = normalize (\i\#image_mset normalize G. i)" by auto next from 2 prime_elem_iff_irreducible have "f \# F \ prime_elem f" "g \# G \ prime_elem g" for f g by (auto intro: prime_elemI) then show " Multiset.Ball (image_mset normalize F) prime" "Multiset.Ball (image_mset normalize G) prime" by auto qed also obtain gs where G: "G = mset gs" by (metis ex_mset) have "list_all2 ((ddvd)\\) gs (map normalize gs)" by (intro list_all2_all_nthI, auto) then have "rel_mset (ddvd) (image_mset normalize G) G" by (subst multiset.rel_flip[symmetric], unfold rel_mset_def G, force) finally show ?case. qed end instance int :: ufd by (intro class.ufd.of_class.intro as_ufd.ufd_axioms) instance int :: idom_gcd by (intro_classes, auto) instance field \ ufd by (intro_classes, auto simp: dvd_field_iff) end diff --git a/thys/Bernoulli/Bernoulli_FPS.thy b/thys/Bernoulli/Bernoulli_FPS.thy --- a/thys/Bernoulli/Bernoulli_FPS.thy +++ b/thys/Bernoulli/Bernoulli_FPS.thy @@ -1,1370 +1,1370 @@ (* File: Bernoulli_FPS.thy Author: Manuel Eberl Connection of Bernoulli numbers to formal power series; proof B_n = 0 for odd n > 1; Akiyama-Tanigawa algorithm. *) section \Connection of Bernoulli numbers to formal power series\ theory Bernoulli_FPS imports Bernoulli "HOL-Computational_Algebra.Computational_Algebra" + "HOL-Combinatorics.Stirling" "HOL-Number_Theory.Number_Theory" - "HOL-Library.Stirling" begin subsection \Preliminaries\ context factorial_semiring begin lemma multiplicity_prime_prime: "prime p \ prime q \ multiplicity p q = (if p = q then 1 else 0)" by (simp add: prime_multiplicity_other) lemma prime_prod_dvdI: fixes f :: "'b \ 'a" assumes "finite A" assumes "\x. x \ A \ prime (f x)" assumes "\x. x \ A \ f x dvd y" assumes "inj_on f A" shows "prod f A dvd y" proof (cases "y = 0") case False have nz: "f x \ 0" if "x \ A" for x using assms(2)[of x] that by auto have "prod f A \ 0" using assms nz by (subst prod_zero_iff) auto thus ?thesis proof (rule multiplicity_le_imp_dvd) fix p :: 'a assume "prime p" show "multiplicity p (prod f A) \ multiplicity p y" proof (cases "p dvd prod f A") case True then obtain x where x: "x \ A" and "p dvd f x" using \prime p\ assms by (subst (asm) prime_dvd_prod_iff) auto have "multiplicity p (prod f A) = (\x\A. multiplicity p (f x))" using assms \prime p\ nz by (intro prime_elem_multiplicity_prod_distrib) auto also have "\ = (\x\{x}. 1 :: nat)" using assms \prime p\ \p dvd f x\ primes_dvd_imp_eq x by (intro Groups_Big.sum.mono_neutral_cong_right) (auto simp: multiplicity_prime_prime inj_on_def) finally have "multiplicity p (prod f A) = 1" by simp also have "1 \ multiplicity p y" using assms nz \prime p\ \y \ 0\ x \p dvd f x\ by (intro multiplicity_geI) force+ finally show ?thesis . qed (auto simp: not_dvd_imp_multiplicity_0) qed qed auto end (* TODO: Move? *) context semiring_gcd begin lemma gcd_add_dvd_right1: "a dvd b \ gcd a (b + c) = gcd a c" by (elim dvdE) (simp add: gcd_add_mult mult.commute[of a]) lemma gcd_add_dvd_right2: "a dvd c \ gcd a (b + c) = gcd a b" using gcd_add_dvd_right1[of a c b] by (simp add: add_ac) lemma gcd_add_dvd_left1: "a dvd b \ gcd (b + c) a = gcd c a" using gcd_add_dvd_right1[of a b c] by (simp add: gcd.commute) lemma gcd_add_dvd_left2: "a dvd c \ gcd (b + c) a = gcd b a" using gcd_add_dvd_right2[of a c b] by (simp add: gcd.commute) end context ring_gcd begin lemma gcd_diff_dvd_right1: "a dvd b \ gcd a (b - c) = gcd a c" using gcd_add_dvd_right1[of a b "-c"] by simp lemma gcd_diff_dvd_right2: "a dvd c \ gcd a (b - c) = gcd a b" using gcd_add_dvd_right2[of a "-c" b] by simp lemma gcd_diff_dvd_left1: "a dvd b \ gcd (b - c) a = gcd c a" using gcd_add_dvd_left1[of a b "-c"] by simp lemma gcd_diff_dvd_left2: "a dvd c \ gcd (b - c) a = gcd b a" using gcd_add_dvd_left2[of a "-c" b] by simp end lemma cong_int: "[a = b] (mod m) \ [int a = int b] (mod m)" by (simp add: cong_int_iff) lemma Rats_int_div_natE: assumes "(x :: 'a :: field_char_0) \ \" obtains m :: int and n :: nat where "n > 0" and "x = of_int m / of_nat n" and "coprime m n" proof - from assms obtain r where [simp]: "x = of_rat r" by (auto simp: Rats_def) obtain a b where [simp]: "r = Rat.Fract a b" and ab: "b > 0" "coprime a b" by (cases r) from ab show ?thesis by (intro that[of "nat b" a]) (auto simp: of_rat_rat) qed lemma sum_in_Ints: "(\x. x \ A \ f x \ \) \ sum f A \ \" by (induction A rule: infinite_finite_induct) auto lemma Ints_real_of_nat_divide: "b dvd a \ real a / real b \ \" by auto lemma product_dvd_fact: assumes "a > 1" "b > 1" "a = b \ a > 2" shows "(a * b) dvd fact (a * b - 1)" proof (cases "a = b") case False have "a * 1 < a * b" and "1 * b < a * b" using assms by (intro mult_strict_left_mono mult_strict_right_mono; simp)+ hence ineqs: "a \ a * b - 1" "b \ a * b - 1" by linarith+ from False have "a * b = \{a,b}" by simp also have "\ dvd \{1..a * b - 1}" using assms ineqs by (intro prod_dvd_prod_subset) auto finally show ?thesis by (simp add: fact_prod) next case [simp]: True from assms have "a > 2" by auto hence "a * 2 < a * b" using assms by (intro mult_strict_left_mono; simp) hence *: "2 * a \ a * b - 1" by linarith have "a * a dvd (2 * a) * a" by simp also have "\ = \{2*a, a}" using assms by auto also have "\ dvd \{1..a * b - 1}" using assms * by (intro prod_dvd_prod_subset) auto finally show ?thesis by (simp add: fact_prod) qed lemma composite_imp_factors_nat: assumes "m > 1" "\prime (m::nat)" shows "\n k. m = n * k \ 1 < n \ n < m \ 1 < k \ k < m" proof - from assms have "\irreducible m" by (simp flip: prime_elem_iff_irreducible ) then obtain a where a: "a dvd m" "\m dvd a" "a \ 1" using assms by (auto simp: irreducible_altdef) then obtain b where [simp]: "m = a * b" by auto from a assms have "a \ 0" "b \ 0" "b \ 1" by (auto intro!: Nat.gr0I) with a have "a > 1" "b > 1" by linarith+ moreover from this and a have "a < m" "b < m" by auto ultimately show ?thesis using \m = a * b\ by blast qed text \ This lemma describes what the numerator and denominator of a finite subseries of the harmonic series are when it is written as a single fraction. \ lemma sum_inverses_conv_fraction: fixes f :: "'a \ 'b :: field" assumes "\x. x \ A \ f x \ 0" "finite A" shows "(\x\A. 1 / f x) = (\x\A. \y\A-{x}. f y) / (\x\A. f x)" proof - have "(\x\A. (\y\A. f y) / f x) = (\x\A. \y\A-{x}. f y)" using prod.remove[of A _ f] assms by (intro sum.cong refl) (auto simp: field_simps) thus ?thesis using assms by (simp add: field_simps sum_distrib_right sum_distrib_left) qed text \ If all terms in the subseries are primes, this fraction is automatically on lowest terms. \ lemma sum_prime_inverses_fraction_coprime: fixes f :: "'a \ nat" assumes "finite A" and primes: "\x. x \ A \ prime (f x)" and inj: "inj_on f A" defines "a \ (\x\A. \y\A-{x}. f y)" shows "coprime a (\x\A. f x)" proof (intro prod_coprime_right) fix x assume x: "x \ A" have "a = (\y\A-{x}. f y) + (\y\A-{x}. \z\A-{y}. f z)" unfolding a_def using \finite A\ and x by (rule sum.remove) also have "gcd \ (f x) = gcd (\y\A-{x}. f y) (f x)" using \finite A\ and x by (intro gcd_add_dvd_left2 dvd_sum dvd_prodI) auto also from x primes inj have "coprime (\y\A-{x}. f y) (f x)" by (intro prod_coprime_left) (auto intro!: primes_coprime simp: inj_on_def) hence "gcd (\y\A-{x}. f y) (f x) = 1" by simp finally show "coprime a (f x)" by (simp only: coprime_iff_gcd_eq_1) qed (* END TODO *) text \ In the following, we will prove the correctness of the Akiyama--Tanigawa algorithm~\cite{kaneko2000}, which is a simple algorithm for computing Bernoulli numbers that was discovered by Akiyama and Tanigawa~\cite{aki_tani1999} essentially as a by-product of their studies of the Euler--Zagier multiple zeta function. The algorithm is based on a number triangle (similar to Pascal's triangle) in which the Bernoulli numbers are the leftmost diagonal. While the algorithm itself is quite simple, proving it correct is not entirely trivial. We will use generating functions and Stirling numbers, mostly following the presentation by Kaneko~\cite{kaneko2000}. \ text \ The following operator is a variant of the @{term fps_XD} operator where the multiplication is not with @{term fps_X}, but with an arbitrary formal power series. It is not quite clear if this operator has a less ad-hoc meaning than the fashion in which we use it; it is, however, very useful for proving the relationship between Stirling numbers and Bernoulli numbers. \ context includes fps_notation begin definition fps_XD' where "fps_XD' a = (\b. a * fps_deriv b)" lemma fps_XD'_0 [simp]: "fps_XD' a 0 = 0" by (simp add: fps_XD'_def) lemma fps_XD'_1 [simp]: "fps_XD' a 1 = 0" by (simp add: fps_XD'_def) lemma fps_XD'_fps_const [simp]: "fps_XD' a (fps_const b) = 0" by (simp add: fps_XD'_def) lemma fps_XD'_fps_of_nat [simp]: "fps_XD' a (of_nat b) = 0" by (simp add: fps_XD'_def) lemma fps_XD'_fps_of_int [simp]: "fps_XD' a (of_int b) = 0" by (simp add: fps_XD'_def) lemma fps_XD'_fps_numeral [simp]: "fps_XD' a (numeral b) = 0" by (simp add: fps_XD'_def) lemma fps_XD'_add [simp]: "fps_XD' a (b + c :: 'a :: comm_ring_1 fps) = fps_XD' a b + fps_XD' a c" by (simp add: fps_XD'_def algebra_simps) lemma fps_XD'_minus [simp]: "fps_XD' a (b - c :: 'a :: comm_ring_1 fps) = fps_XD' a b - fps_XD' a c" by (simp add: fps_XD'_def algebra_simps) lemma fps_XD'_prod: "fps_XD' a (b * c :: 'a :: comm_ring_1 fps) = fps_XD' a b * c + b * fps_XD' a c" by (simp add: fps_XD'_def algebra_simps) lemma fps_XD'_power: "fps_XD' a (b ^ n :: 'a :: idom fps) = of_nat n * b ^ (n - 1) * fps_XD' a b" proof (cases "n = 0") case False have "b * fps_XD' a (b ^ n) = of_nat n * b ^ n * fps_XD' a b" by (induction n) (simp_all add: fps_XD'_prod algebra_simps) also have "\ = b * (of_nat n * b ^ (n - 1) * fps_XD' a b)" by (cases n) (simp_all add: algebra_simps) finally show ?thesis using False by (subst (asm) mult_cancel_left) (auto simp: power_0_left) qed simp_all lemma fps_XD'_power_Suc: "fps_XD' a (b ^ Suc n :: 'a :: idom fps) = of_nat (Suc n) * b ^ n * fps_XD' a b" by (subst fps_XD'_power) simp_all lemma fps_XD'_sum: "fps_XD' a (sum f A) = sum (\x. fps_XD' (a :: 'a :: comm_ring_1 fps) (f x)) A" by (induction A rule: infinite_finite_induct) simp_all lemma fps_XD'_funpow_affine: fixes G H :: "real fps" assumes [simp]: "fps_deriv G = 1" defines "S \ \n i. fps_const (real (Stirling n i))" shows "(fps_XD' G ^^ n) H = (\m\n. S n m * G ^ m * (fps_deriv ^^ m) H)" proof (induction n arbitrary: H) case 0 thus ?case by (simp add: S_def) next case (Suc n H) have "(\m\Suc n. S (Suc n) m * G ^ m * (fps_deriv ^^ m) H) = (\i\n. of_nat (Suc i) * S n (Suc i) * G ^ Suc i * (fps_deriv ^^ Suc i) H) + (\i\n. S n i * G ^ Suc i * (fps_deriv ^^ Suc i) H)" (is "_ = sum (\i. ?f (Suc i)) \ + ?S2") by (subst sum.atMost_Suc_shift) (simp_all add: sum.distrib algebra_simps fps_of_nat S_def fps_const_add [symmetric] fps_const_mult [symmetric] del: fps_const_add fps_const_mult) also have "sum (\i. ?f (Suc i)) {..n} = sum (\i. ?f (Suc i)) {.. = ?f 0 + \" by simp also have "\ = sum ?f {..n}" by (subst sum.atMost_shift [symmetric]) simp_all also have "\ + ?S2 = (\x\n. fps_XD' G (S n x * G ^ x * (fps_deriv ^^ x) H))" unfolding sum.distrib [symmetric] proof (rule sum.cong, goal_cases) case (2 i) thus ?case unfolding fps_XD'_prod fps_XD'_power by (cases i) (auto simp: fps_XD'_prod fps_XD'_power_Suc algebra_simps of_nat_diff S_def fps_XD'_def) qed simp_all also have "\ = (fps_XD' G ^^ Suc n) H" by (simp add: Suc.IH fps_XD'_sum) finally show ?case .. qed subsection \Generating function of Stirling numbers\ lemma Stirling_n_0: "Stirling n 0 = (if n = 0 then 1 else 0)" by (cases n) simp_all text \ The generating function of Stirling numbers w.\,r.\,t.\ their first argument: \[\sum_{n=0}^\infty \genfrac{\{}{\}}{0pt}{}{n}{m} \frac{x^n}{n!} = \frac{(e^x - 1)^m}{m!}\] \ definition Stirling_fps :: "nat \ real fps" where "Stirling_fps m = fps_const (1 / fact m) * (fps_exp 1 - 1) ^ m" theorem sum_Stirling_binomial: "Stirling (Suc n) (Suc m) = (\i = 0..n. Stirling i m * (n choose i))" proof - have "real (Stirling (Suc n) (Suc m)) = real (\i = 0..n. Stirling i m * (n choose i))" proof (induction n arbitrary: m) case (Suc n m) have "real (\i = 0..Suc n. Stirling i m * (Suc n choose i)) = real (\i = 0..n. Stirling (Suc i) m * (Suc n choose Suc i)) + real (Stirling 0 m)" by (subst sum.atLeast0_atMost_Suc_shift) simp_all also have "real (\i = 0..n. Stirling (Suc i) m * (Suc n choose Suc i)) = real (\i = 0..n. (n choose i) * Stirling (Suc i) m) + real (\i = 0..n. (n choose Suc i) * Stirling (Suc i) m)" by (simp add: algebra_simps sum.distrib) also have "(\i = 0..n. (n choose Suc i) * Stirling (Suc i) m) = (\i = Suc 0..Suc n. (n choose i) * Stirling i m)" by (subst sum.shift_bounds_cl_Suc_ivl) simp_all also have "\ = (\i = Suc 0..n. (n choose i) * Stirling i m)" by (intro sum.mono_neutral_right) auto also have "\ = real (\i = 0..n. Stirling i m * (n choose i)) - real (Stirling 0 m)" by (simp add: sum.atLeast_Suc_atMost mult_ac) also have "real (\i = 0..n. Stirling i m * (n choose i)) = real (Stirling (Suc n) (Suc m))" by (rule Suc.IH [symmetric]) also have "real (\i = 0..n. (n choose i) * Stirling (Suc i) m) = real m * real (Stirling (Suc n) (Suc m)) + real (Stirling (Suc n) m)" by (cases m; (simp only: Suc.IH, simp add: algebra_simps sum.distrib sum_distrib_left sum_distrib_right)) also have "\ + (real (Stirling (Suc n) (Suc m)) - real (Stirling 0 m)) + real (Stirling 0 m) = real (Suc m * Stirling (Suc n) (Suc m) + Stirling (Suc n) m)" by (simp add: algebra_simps del: Stirling.simps) also have "Suc m * Stirling (Suc n) (Suc m) + Stirling (Suc n) m = Stirling (Suc (Suc n)) (Suc m)" by (rule Stirling.simps(4) [symmetric]) finally show ?case .. qed simp_all thus ?thesis by (subst (asm) of_nat_eq_iff) qed lemma Stirling_fps_aux: "(fps_exp 1 - 1) ^ m $ n * fact n = fact m * real (Stirling n m)" proof (induction m arbitrary: n) case 0 thus ?case by (simp add: Stirling_n_0) next case (Suc m n) show ?case proof (cases n) case 0 thus ?thesis by simp next case (Suc n') hence "(fps_exp 1 - 1 :: real fps) ^ Suc m $ n * fact n = fps_deriv ((fps_exp 1 - 1) ^ Suc m) $ n' * fact n'" by (simp_all add: algebra_simps del: power_Suc) also have "fps_deriv ((fps_exp 1 - 1 :: real fps) ^ Suc m) = fps_const (real (Suc m)) * ((fps_exp 1 - 1) ^ m * fps_exp 1)" by (subst fps_deriv_power) simp_all also have "\ $ n' * fact n' = real (Suc m) * ((\i = 0..n'. (fps_exp 1 - 1) ^ m $ i / fact (n' - i)) * fact n')" unfolding fps_mult_left_const_nth by (simp add: fps_mult_nth Suc.IH sum_distrib_right del: of_nat_Suc) also have "(\i = 0..n'. (fps_exp 1 - 1 :: real fps) ^ m $ i / fact (n' - i)) * fact n' = (\i = 0..n'. (fps_exp 1 - 1) ^ m $ i * fact n' / fact (n' - i))" by (subst sum_distrib_right, rule sum.cong) (simp_all add: divide_simps) also have "\ = (\i = 0..n'. (fps_exp 1 - 1) ^ m $ i * fact i * (n' choose i))" by (intro sum.cong refl) (simp_all add: binomial_fact) also have "\ = (\i = 0..n'. fact m * real (Stirling i m) * real (n' choose i))" by (simp only: Suc.IH) also have "real (Suc m) * \ = fact (Suc m) * (\i = 0..n'. real (Stirling i m) * real (n' choose i))" (is "_ = _ * ?S") by (simp add: sum_distrib_left sum_distrib_right mult_ac del: of_nat_Suc) also have "?S = Stirling (Suc n') (Suc m)" by (subst sum_Stirling_binomial) simp also have "Suc n' = n" by (simp add: Suc) finally show ?thesis . qed qed lemma Stirling_fps_nth: "Stirling_fps m $ n = Stirling n m / fact n" unfolding Stirling_fps_def using Stirling_fps_aux[of m n] by (simp add: field_simps) theorem Stirling_fps_altdef: "Stirling_fps m = Abs_fps (\n. Stirling n m / fact n)" by (simp add: fps_eq_iff Stirling_fps_nth) theorem Stirling_closed_form: "real (Stirling n k) = (\j\k. (-1)^(k - j) * real (k choose j) * real j ^ n) / fact k" proof - have "(fps_exp 1 - 1 :: real fps) = (fps_exp 1 + (-1))" by simp also have "\ ^ k = (\j\k. of_nat (k choose j) * fps_exp 1 ^ j * (- 1) ^ (k - j))" unfolding binomial_ring .. also have "\ = (\j\k. fps_const ((-1) ^ (k - j) * real (k choose j)) * fps_exp (real j))" by (simp add: fps_const_mult [symmetric] fps_const_power [symmetric] fps_const_neg [symmetric] mult_ac fps_of_nat fps_exp_power_mult del: fps_const_mult fps_const_power fps_const_neg) also have "\ $ n = (\j\k. (- 1) ^ (k - j) * real (k choose j) * real j ^ n) / fact n" by (simp add: fps_sum_nth sum_divide_distrib) also have "\ * fact n = (\j\k. (- 1) ^ (k - j) * real (k choose j) * real j ^ n)" by simp also note Stirling_fps_aux[of k n] finally show ?thesis by (simp add: atLeast0AtMost field_simps) qed subsection \Generating function of Bernoulli numbers\ text \ We will show that the negative and positive Bernoulli numbers are the coefficients of the exponential generating function $\frac{x}{e^x - 1}$ (resp. $\frac{x}{1-e^{-x}}$), i.\,e. \[\sum_{n=0}^\infty B_n^{-} \frac{x^n}{n!} = \frac{x}{e^x - 1}\] \[\sum_{n=0}^\infty B_n^{+} \frac{x^n}{n!} = \frac{x}{1 - e^{-1}}\] \ definition bernoulli_fps :: "'a :: real_normed_field fps" where "bernoulli_fps = fps_X / (fps_exp 1 - 1)" definition bernoulli'_fps :: "'a :: real_normed_field fps" where "bernoulli'_fps = fps_X / (1 - (fps_exp (-1)))" lemma bernoulli_fps_altdef: "bernoulli_fps = Abs_fps (\n. of_real (bernoulli n) / fact n :: 'a)" and bernoulli_fps_aux: "bernoulli_fps * (fps_exp 1 - 1 :: 'a :: real_normed_field fps) = fps_X" proof - have *: "Abs_fps (\n. of_real (bernoulli n) / fact n :: 'a) * (fps_exp 1 - 1) = fps_X" proof (rule fps_ext) fix n have "(Abs_fps (\n. of_real (bernoulli n) / fact n :: 'a) * (fps_exp 1 - 1)) $ n = (\i = 0..n. of_real (bernoulli i) * (1 / fact (n - i) - (if n = i then 1 else 0)) / fact i)" by (auto simp: fps_mult_nth divide_simps split: if_splits intro!: sum.cong) also have "\ = (\i = 0..n. of_real (bernoulli i) / (fact i * fact (n - i)) - (if n = i then of_real (bernoulli i) / fact i else 0))" by (intro sum.cong) (simp_all add: field_simps) also have "\ = (\i = 0..n. of_real (bernoulli i) / (fact i * fact (n - i))) - of_real (bernoulli n) / fact n" unfolding sum_subtractf by (subst sum.delta') simp_all also have "\ = (\i = (\iii = of_real (\i / fact n = fps_X $ n" by (subst sum_binomial_times_bernoulli') simp_all finally show "(Abs_fps (\n. of_real (bernoulli n) / fact n :: 'a) * (fps_exp 1 - 1)) $ n = fps_X $ n" . qed moreover show "bernoulli_fps = Abs_fps (\n. of_real (bernoulli n) / fact n :: 'a)" unfolding bernoulli_fps_def by (subst * [symmetric]) simp_all ultimately show "bernoulli_fps * (fps_exp 1 - 1 :: 'a fps) = fps_X" by simp qed theorem fps_nth_bernoulli_fps [simp]: "fps_nth bernoulli_fps n = of_real (bernoulli n) / fact n" by (simp add: bernoulli_fps_altdef) lemma bernoulli'_fps_aux: "(fps_exp 1 - 1) * Abs_fps (\n. of_real (bernoulli' n) / fact n :: 'a) = fps_exp 1 * fps_X" and bernoulli'_fps_aux': "(1 - fps_exp (-1)) * Abs_fps (\n. of_real (bernoulli' n) / fact n :: 'a) = fps_X" and bernoulli'_fps_altdef: "bernoulli'_fps = Abs_fps (\n. of_real (bernoulli' n) / fact n :: 'a :: real_normed_field)" proof - have "Abs_fps (\n. of_real (bernoulli' n) / fact n :: 'a) = bernoulli_fps + fps_X" by (simp add: fps_eq_iff bernoulli'_def) also have "(fps_exp 1 - 1) * \ = fps_exp 1 * fps_X" using bernoulli_fps_aux by (simp add: algebra_simps) finally show "(fps_exp 1 - 1) * Abs_fps (\n. of_real (bernoulli' n) / fact n :: 'a) = fps_exp 1 * fps_X" . also have "(fps_exp 1 - 1) = fps_exp 1 * (1 - fps_exp (-1 :: 'a))" by (simp add: algebra_simps fps_exp_add_mult [symmetric]) also note mult.assoc finally show *: "(1 - fps_exp (-1)) * Abs_fps (\n. of_real (bernoulli' n) / fact n :: 'a) = fps_X" by (subst (asm) mult_left_cancel) simp_all show "bernoulli'_fps = Abs_fps (\n. of_real (bernoulli' n) / fact n :: 'a)" unfolding bernoulli'_fps_def by (subst * [symmetric]) simp_all qed theorem fps_nth_bernoulli'_fps [simp]: "fps_nth bernoulli'_fps n = of_real (bernoulli' n) / fact n" by (simp add: bernoulli'_fps_altdef) lemma bernoulli_fps_conv_bernoulli'_fps: "bernoulli_fps = bernoulli'_fps - fps_X" by (simp add: fps_eq_iff bernoulli'_def) lemma bernoulli'_fps_conv_bernoulli_fps: "bernoulli'_fps = bernoulli_fps + fps_X" by (simp add: fps_eq_iff bernoulli'_def) theorem bernoulli_odd_eq_0: assumes "n \ 1" and "odd n" shows "bernoulli n = 0" proof - from bernoulli_fps_aux have "2 * bernoulli_fps * (fps_exp 1 - 1) = 2 * fps_X" by simp hence "(2 * bernoulli_fps + fps_X) * (fps_exp 1 - 1) = fps_X * (fps_exp 1 + 1)" by (simp add: algebra_simps) also have "fps_exp 1 - 1 = fps_exp (1/2) * (fps_exp (1/2) - fps_exp (-1/2 :: real))" by (simp add: algebra_simps fps_exp_add_mult [symmetric]) also have "fps_exp 1 + 1 = fps_exp (1/2) * (fps_exp (1/2) + fps_exp (-1/2 :: real))" by (simp add: algebra_simps fps_exp_add_mult [symmetric]) finally have "fps_exp (1/2) * ((2 * bernoulli_fps + fps_X) * (fps_exp (1/2) - fps_exp (- 1/2))) = fps_exp (1/2) * (fps_X * (fps_exp (1/2) + fps_exp (-1/2 :: real)))" by (simp add: algebra_simps) hence *: "(2 * bernoulli_fps + fps_X) * (fps_exp (1/2) - fps_exp (- 1/2)) = fps_X * (fps_exp (1/2) + fps_exp (-1/2 :: real))" (is "?lhs = ?rhs") by (subst (asm) mult_cancel_left) simp_all have "fps_compose ?lhs (-fps_X) = fps_compose ?rhs (-fps_X)" by (simp only: *) also have "fps_compose ?lhs (-fps_X) = (-2 * (bernoulli_fps oo - fps_X) + fps_X) * (fps_exp ((1/2)) - fps_exp (-1/2))" by (simp add: fps_compose_mult_distrib fps_compose_add_distrib fps_compose_sub_distrib algebra_simps) also have "fps_compose ?rhs (-fps_X) = -?rhs" by (simp add: fps_compose_mult_distrib fps_compose_add_distrib fps_compose_sub_distrib) also note * [symmetric] also have "- ((2 * bernoulli_fps + fps_X) * (fps_exp (1/2) - fps_exp (-1/2))) = ((-2 * bernoulli_fps - fps_X) * (fps_exp (1/2) - fps_exp (-1/2)))" by (simp add: algebra_simps) finally have "2 * (bernoulli_fps oo - fps_X) = 2 * (bernoulli_fps + fps_X :: real fps)" by (subst (asm) mult_cancel_right) (simp add: algebra_simps) hence **: "bernoulli_fps oo -fps_X = (bernoulli_fps + fps_X :: real fps)" by (subst (asm) mult_cancel_left) simp from assms have "(bernoulli_fps oo -fps_X) $ n = bernoulli n / fact n" by (subst **) simp also have "-fps_X = fps_const (-1 :: real) * fps_X" by (simp only: fps_const_neg [symmetric] fps_const_1_eq_1) simp also from assms have "(bernoulli_fps oo \) $ n = - bernoulli n / fact n" by (subst fps_compose_linear) simp finally show ?thesis by simp qed lemma bernoulli'_odd_eq_0: "n \ 1 \ odd n \ bernoulli' n = 0" by (simp add: bernoulli'_def bernoulli_odd_eq_0) text \ The following simplification rule takes care of rewriting @{term "bernoulli n"} to $0$ for any odd numeric constant greater than $1$: \ lemma bernoulli_odd_numeral_eq_0 [simp]: "bernoulli (numeral (Num.Bit1 n)) = 0" by (rule bernoulli_odd_eq_0[OF _ odd_numeral]) auto lemma bernoulli'_odd_numeral_eq_0 [simp]: "bernoulli' (numeral (Num.Bit1 n)) = 0" by (simp add: bernoulli'_def) text \ The following explicit formula for Bernoulli numbers can also derived reasonably easily using the generating functions of Stirling numbers and Bernoulli numbers. The proof follows an answer by Marko Riedel on the Mathematics StackExchange~\cite{riedel_mathse_2014}. \ theorem bernoulli_altdef: "bernoulli n = (\m\n. \k\m. (-1)^k * real (m choose k) * real k^n / real (Suc m))" proof - have "(\m\n. \k\m. (-1)^k * real (m choose k) * real k^n / real (Suc m)) = (\m\n. (\k\m. (-1)^k * real (m choose k) * real k^n) / real (Suc m))" by (subst sum_divide_distrib) simp_all also have "\ = fact n * (\m\n. (- 1) ^ m / real (Suc m) * (fps_exp 1 - 1) ^ m $ n)" proof (subst sum_distrib_left, intro sum.cong refl) fix m assume m: "m \ {..n}" have "(\k\m. (-1)^k * real (m choose k) * real k^n) = (-1)^m * (\k\m. (-1)^(m - k) * real (m choose k) * real k^n)" by (subst sum_distrib_left, intro sum.cong refl) (auto simp: minus_one_power_iff) also have "\ = (-1) ^ m * (real (Stirling n m) * fact m)" by (subst Stirling_closed_form) simp_all also have "real (Stirling n m) = Stirling_fps m $ n * fact n" by (subst Stirling_fps_nth) simp_all also have "\ * fact m = (fps_exp 1 - 1) ^ m $ n * fact n" by (simp add: Stirling_fps_def) finally show "(\k\m. (-1)^k * real (m choose k) * real k^n) / real (Suc m) = fact n * ((- 1) ^ m / real (Suc m) * (fps_exp 1 - 1) ^ m $ n)" by simp qed also have "(\m\n. (- 1) ^ m / real (Suc m) * (fps_exp 1 - 1) ^ m $ n) = fps_compose (Abs_fps (\m. (-1) ^ m / real (Suc m))) (fps_exp 1 - 1) $ n" by (simp add: fps_compose_def atLeast0AtMost fps_sum_nth) also have "fps_ln 1 = fps_X * Abs_fps (\m. (-1) ^ m / real (Suc m))" unfolding fps_ln_def by (auto simp: fps_eq_iff) hence "Abs_fps (\m. (-1) ^ m / real (Suc m)) = fps_ln 1 / fps_X" by (metis fps_X_neq_zero nonzero_mult_div_cancel_left) also have "fps_compose \ (fps_exp 1 - 1) = fps_compose (fps_ln 1) (fps_exp 1 - 1) / (fps_exp 1 - 1)" by (subst fps_compose_divide_distrib) auto also have "fps_compose (fps_ln 1) (fps_exp 1 - 1 :: real fps) = fps_X" by (simp add: fps_ln_fps_exp_inv fps_inv_fps_exp_compose) also have "(fps_X / (fps_exp 1 - 1)) = bernoulli_fps" by (simp add: bernoulli_fps_def) also have "fact n * \ $ n = bernoulli n" by simp finally show ?thesis .. qed corollary%important bernoulli_conv_Stirling: "bernoulli n = (\k\n. (-1) ^ k * fact k / real (k + 1) * Stirling n k)" proof - have "(\k\n. (-1) ^ k * fact k / (k + 1) * Stirling n k) = (\k\n. \i\k. (-1) ^ i * (k choose i) * i ^ n / real (k + 1))" proof (intro sum.cong, goal_cases) case (2 k) have "(-1) ^ k * fact k / (k + 1) * Stirling n k = (\j\k. (-1) ^ k * (-1) ^ (k - j) * (k choose j) * j ^ n / (k + 1))" by (simp add: Stirling_closed_form sum_distrib_left sum_divide_distrib mult_ac) also have "\ = (\j\k. (-1) ^ j * (k choose j) * j ^ n / (k + 1))" by (intro sum.cong) (auto simp: uminus_power_if split: if_splits) finally show ?case . qed auto also have "\ = bernoulli n" by (simp add: bernoulli_altdef) finally show ?thesis .. qed subsection \Von Staudt--Clausen Theorem\ lemma vonStaudt_Clausen_lemma: assumes "n > 0" and "prime p" shows "[(\m 0" "m < p" for m proof - from True obtain q where "2 * n = (p - 1) * q" by blast hence "[m ^ (2 * n) = (m ^ (p - 1)) ^ q] (mod p)" by (simp add: power_mult) also have "[(m ^ (p - 1)) ^ q = 1 ^ q] (mod p)" using assms \m > 0\ \m < p\ by (intro cong_pow fermat_theorem) auto finally show ?thesis by simp qed have "(\mm\{0<..n > 0\ by (intro sum.mono_neutral_right) auto also have "[\ = (\m\{0<..m\{0<..m\insert 0 {0<..m\p-1. (-1)^m * ((p - 1) choose m)) = 0" using prime_gt_1_nat[of p] assms by (subst choose_alternating_sum) auto finally show ?thesis using True by simp next case False define n' where "n' = (2 * n) mod (p - 1)" from assms False have "n' > 0" by (auto simp: n'_def dvd_eq_mod_eq_0) from False have "p \ 2" by auto with assms have "odd p" using prime_prime_factor two_is_prime_nat by blast have cong_pow_2n: "[m ^ (2*n) = m ^ n'] (mod p)" if "m > 0" "m < p" for m proof - from assms and that have "coprime p m" by (intro prime_imp_coprime) auto have "[2 * n = n'] (mod (p - 1))" by (simp add: n'_def) moreover have "ord p m dvd (p - 1)" using order_divides_totient[of p m] \coprime p m\ assms by (auto simp: totient_prime) ultimately have "[2 * n = n'] (mod ord p m)" by (rule cong_dvd_modulus_nat) thus ?thesis using \coprime p m\ by (subst order_divides_expdiff) auto qed have "(\mm\{0<..n > 0\ by (intro sum.mono_neutral_right) auto also have "[\ = (\m\{0<..m\{0<..m\p-1. (-1)^m * ((p - 1) choose m) * m ^ n')" using \n' > 0\ by (intro sum.mono_neutral_left) auto also have "\ = (\m\p-1. (-1)^(p - Suc m) * ((p - 1) choose m) * m ^ n')" using \n' > 0\ assms \odd p\ by (intro sum.cong) (auto simp: uminus_power_if) also have "\ = 0" proof - have "of_int (\m\p-1. (-1)^(p - Suc m) * ((p - 1) choose m) * m ^ n') = real (Stirling n' (p - 1)) * fact (p - 1)" by (simp add: Stirling_closed_form) also have "n' < p - 1" using assms prime_gt_1_nat[of p] by (auto simp: n'_def) hence "Stirling n' (p - 1) = 0" by simp finally show ?thesis by linarith qed finally show ?thesis using False by simp qed text \ The Von Staudt--Clausen theorem states that for \n > 0\, \[B_{2n} + \sum\limits_{p - 1\mid 2n} \frac{1}{p}\] is an integer. \ theorem vonStaudt_Clausen: assumes "n > 0" shows "bernoulli (2 * n) + (\p | prime p \ (p - 1) dvd (2 * n). 1 / real p) \ \" (is "_ + ?P \ \") proof - define P :: "nat \ real" where "P = (\m. if prime (m + 1) \ m dvd (2 * n) then 1 / (m + 1) else 0)" define P' :: "nat \ int" where "P' = (\m. if prime (m + 1) \ m dvd (2 * n) then 1 else 0)" have "?P = (\p | prime (p + 1) \ p dvd (2 * n). 1 / real (p + 1))" by (rule sum.reindex_bij_witness[of _ "\p. p + 1" "\p. p - 1"]) (use prime_gt_0_nat in auto) also have "\ = (\m\2*n. P m)" using \n > 0\ by (intro sum.mono_neutral_cong_left) (auto simp: P_def dest!: dvd_imp_le) finally have "bernoulli (2 * n) + ?P = (\m\2*n. (-1)^m * (of_int (fact m * Stirling (2*n) m) / (m + 1)) + P m)" by (simp add: sum.distrib bernoulli_conv_Stirling sum_divide_distrib algebra_simps) also have "\ = (\m\2*n. of_int ((-1)^m * fact m * Stirling (2*n) m + P' m) / (m + 1))" by (intro sum.cong) (auto simp: P'_def P_def field_simps) also have "\ \ \" proof (rule sum_in_Ints, goal_cases) case (1 m) have "m = 0 \ m = 3 \ prime (m + 1) \ (\prime (m + 1) \ m > 3)" by (cases "m = 1"; cases "m = 2") (auto simp flip: numeral_2_eq_2) then consider "m = 0" | "m = 3" | "prime (m + 1)" | "\prime (m + 1)" "m > 3" by blast thus ?case proof cases assume "m = 0" thus ?case by auto next assume [simp]: "m = 3" have "real_of_int (fact m * Stirling (2 * n) m) = real_of_int (9 ^ n + 3 - 3 * 4 ^ n)" using \n > 0\ by (auto simp: P'_def fact_numeral Stirling_closed_form power_mult atMost_nat_numeral binomial_fact zero_power) hence "int (fact m * Stirling (2 * n) m) = 9 ^ n + 3 - 3 * 4 ^ n" by linarith also have "[\ = 1 ^ n + (-1) - 3 * 0 ^ n] (mod 4)" by (intro cong_add cong_diff cong_mult cong_pow) (auto simp: cong_def) finally have dvd: "4 dvd int (fact m * Stirling (2 * n) m)" using \n > 0\ by (simp add: cong_0_iff zero_power) have "real_of_int ((- 1) ^ m * fact m * Stirling (2 * n) m + P' m) / (m + 1) = -(real_of_int (int (fact m * Stirling (2 * n) m)) / real_of_int 4)" using \n > 0\ by (auto simp: P'_def) also have "\ \ \" by (intro Ints_minus of_int_divide_in_Ints dvd) finally show ?case . next assume composite: "\prime (m + 1)" and "m > 3" obtain a b where ab: "a * b = m + 1" "a > 1" "b > 1" using \m > 3\ composite composite_imp_factors_nat[of "m + 1"] by auto have "a = b \ a > 2" proof assume "a = b" hence "a ^ 2 > 2 ^ 2" using \m > 3\ and ab by (auto simp: power2_eq_square) thus "a > 2" using power_less_imp_less_base by blast qed hence dvd: "(m + 1) dvd fact m" using product_dvd_fact[of a b] ab by auto have "real_of_int ((- 1) ^ m * fact m * Stirling (2 * n) m + P' m) / real (m + 1) = real_of_int ((- 1) ^ m * Stirling (2 * n) m) * (real (fact m) / (m + 1))" using composite by (auto simp: P'_def) also have "\ \ \" by (intro Ints_mult Ints_real_of_nat_divide dvd) auto finally show ?case . next assume prime: "prime (m + 1)" have "real_of_int ((-1) ^ m * fact m * int (Stirling (2 * n) m)) = (\j\m. (-1) ^ m * (-1) ^ (m - j) * (m choose j) * real_of_int j ^ (2 * n))" by (simp add: Stirling_closed_form sum_divide_distrib sum_distrib_left mult_ac) also have "\ = real_of_int (\j\m. (-1) ^ j * (m choose j) * j ^ (2 * n))" unfolding of_int_sum by (intro sum.cong) (auto simp: uminus_power_if) finally have "(-1) ^ m * fact m * int (Stirling (2 * n) m) = (\j\m. (-1) ^ j * (m choose j) * j ^ (2 * n))" by linarith also have "\ = (\j = (if m dvd 2 * n then - 1 else 0)] (mod (m + 1))" using vonStaudt_Clausen_lemma[of n "m + 1"] prime \n > 0\ by simp also have "(if m dvd 2 * n then - 1 else 0) = - P' m" using prime by (simp add: P'_def) finally have "int (m + 1) dvd ((- 1) ^ m * fact m * int (Stirling (2 * n) m) + P' m)" by (simp add: cong_iff_dvd_diff) hence "real_of_int ((-1)^m * fact m * int (Stirling (2*n) m) + P' m) / of_int (int (m+1)) \ \" by (intro of_int_divide_in_Ints) thus ?case by simp qed qed finally show ?thesis . qed subsection \Denominators of Bernoulli numbers\ text \ A consequence of the Von Staudt--Clausen theorem is that the denominator of $B_{2n}$ for $n > 0$ is precisely the product of all prime numbers \p\ such that \p - 1\ divides $2n$. Since the denominator is obvious in all other cases, this fully characterises the denominator of Bernoulli numbers. \ definition bernoulli_denom :: "nat \ nat" where "bernoulli_denom n = (if n = 1 then 2 else if n = 0 \ odd n then 1 else \{p. prime p \ (p - 1) dvd n})" definition bernoulli_num :: "nat \ int" where "bernoulli_num n = \bernoulli n * bernoulli_denom n\" lemma finite_bernoulli_denom_set: "n > (0 :: nat) \ finite {p. prime p \ (p - 1) dvd n}" by (rule finite_subset[of _ "{..2*n+1}"]) (auto dest!: dvd_imp_le) lemma bernoulli_denom_0 [simp]: "bernoulli_denom 0 = 1" and bernoulli_denom_1 [simp]: "bernoulli_denom 1 = 2" and bernoulli_denom_Suc_0 [simp]: "bernoulli_denom (Suc 0) = 2" and bernoulli_denom_odd [simp]: "n \ 1 \ odd n \ bernoulli_denom n = 1" and bernoulli_denom_even: "n > 0 \ even n \ bernoulli_denom n = \{p. prime p \ (p - 1) dvd n}" by (auto simp: bernoulli_denom_def) lemma bernoulli_denom_pos: "bernoulli_denom n > 0" by (auto simp: bernoulli_denom_def intro!: prod_pos) lemma bernoulli_denom_nonzero [simp]: "bernoulli_denom n \ 0" using bernoulli_denom_pos[of n] by simp lemma bernoulli_denom_code [code]: "bernoulli_denom n = (if n = 1 then 2 else if n = 0 \ odd n then 1 else prod_list (filter (\p. (p - 1) dvd n) (primes_upto (n + 1))))" (is "_ = ?rhs") proof (cases "even n \ n > 0") case True hence "?rhs = prod_list (filter (\p. (p - 1) dvd n) (primes_upto (n + 1)))" by auto also have "\ = \(set (filter (\p. (p - 1) dvd n) (primes_upto (n + 1))))" by (subst prod.distinct_set_conv_list) auto also have "(set (filter (\p. (p - 1) dvd n) (primes_upto (n + 1)))) = {p\{..n+1}. prime p \ (p - 1) dvd n}" by (auto simp: set_primes_upto) also have "\ = {p. prime p \ (p - 1) dvd n}" using True by (auto dest: dvd_imp_le) also have "\\ = bernoulli_denom n" using True by (simp add: bernoulli_denom_even) finally show ?thesis .. qed auto corollary%important bernoulli_denom_correct: obtains a :: int where "coprime a (bernoulli_denom m)" "bernoulli m = of_int a / of_nat (bernoulli_denom m)" proof - consider "m = 0" | "m = 1" | "odd m" "m \ 1" | "even m" "m > 0" by auto thus ?thesis proof cases assume "m = 0" thus ?thesis by (intro that[of 1]) (auto simp: bernoulli_denom_def) next assume "m = 1" thus ?thesis by (intro that[of "-1"]) (auto simp: bernoulli_denom_def) next assume "odd m" "m \ 1" thus ?thesis by (intro that[of 0]) (auto simp: bernoulli_denom_def bernoulli_odd_eq_0) next assume "even m" "m > 0" define n where "n = m div 2" have [simp]: "m = 2 * n" and n: "n > 0" using \even m\ \m > 0\ by (auto simp: n_def intro!: Nat.gr0I) obtain a b where ab: "bernoulli (2 * n) = a / b" "coprime a (int b)" "b > 0" using Rats_int_div_natE[OF bernoulli_in_Rats] by metis define P where "P = {p. prime p \ (p - 1) dvd (2 * n)}" have "finite P" unfolding P_def using n by (intro finite_bernoulli_denom_set) auto from vonStaudt_Clausen[of n] obtain k where k: "bernoulli (2 * n) + (\p\P. 1/p) = of_int k" using \n > 0\ by (auto simp: P_def Ints_def) define c where "c = (\p\P. \(P-{p}))" from \finite P\ have "(\p\P. 1 / p) = c / \P" by (subst sum_inverses_conv_fraction) (auto simp: P_def prime_gt_0_nat c_def) moreover have P_nz: "prod real P > 0" using prime_gt_0_nat by (auto simp: P_def intro!: prod_pos) ultimately have eq: "bernoulli (2 * n) = (k * \P - c) / \P" using ab P_nz by (simp add: field_simps k [symmetric]) have "gcd (k * \P - int c) (\P) = gcd (int c) (\P)" by (simp add: gcd_diff_dvd_left1) also have "\ = int (gcd c (\P))" by (simp flip: gcd_int_int_eq) also have "coprime c (\P)" unfolding c_def using \finite P\ by (intro sum_prime_inverses_fraction_coprime) (auto simp: P_def) hence "gcd c (\P) = 1" by simp finally have coprime: "coprime (k * \P - int c) (\P)" by (simp only: coprime_iff_gcd_eq_1) have eq': "\P = bernoulli_denom (2 * n)" using n by (simp add: bernoulli_denom_def P_def) show ?thesis by (rule that[of "k * \P - int c"]) (use eq eq' coprime in simp_all) qed qed lemma bernoulli_conv_num_denom: "bernoulli n = bernoulli_num n / bernoulli_denom n" (is ?th1) and coprime_bernoulli_num_denom: "coprime (bernoulli_num n) (bernoulli_denom n)" (is ?th2) proof - obtain a :: int where a: "coprime a (bernoulli_denom n)" "bernoulli n = a / bernoulli_denom n" using bernoulli_denom_correct[of n] by blast thus ?th1 by (simp add: bernoulli_num_def) with a show ?th2 by auto qed text \ Two obvious consequences from this are that the denominators of all odd Bernoulli numbers except for the first one are squarefree and multiples of 6: \ lemma six_divides_bernoulli_denom: assumes "even n" "n > 0" shows "6 dvd bernoulli_denom n" proof - from assms have "\{2, 3} dvd \{p. prime p \ (p - 1) dvd n}" by (intro prod_dvd_prod_subset finite_bernoulli_denom_set) auto with assms show ?thesis by (simp add: bernoulli_denom_even) qed lemma squarefree_bernoulli_denom: "squarefree (bernoulli_denom n)" by (auto intro!: squarefree_prod_coprime primes_coprime simp: bernoulli_denom_def squarefree_prime) text \ Furthermore, the denominator of $B_n$ divides $2(2^n - 1)$. This also gives us an upper bound on the denominators. \ lemma bernoulli_denom_dvd: "bernoulli_denom n dvd (2 * (2 ^ n - 1))" proof (cases "even n \ n > 0") case True hence "bernoulli_denom n = \{p. prime p \ (p - 1) dvd n}" by (auto simp: bernoulli_denom_def) also have "\ dvd (2 * (2 ^ n - 1))" proof (rule prime_prod_dvdI; clarify?) from True show "finite {p. prime p \ (p - 1) dvd n}" by (intro finite_bernoulli_denom_set) auto next fix p assume p: "prime p" "(p - 1) dvd n" show "p dvd (2 * (2 ^ n - 1))" proof (cases "p = 2") case False with p have "p > 2" using prime_gt_1_nat[of p] by force have "[2 ^ n - 1 = 1 - 1] (mod p)" using p \p > 2\ prime_odd_nat by (intro cong_diff_nat Carmichael_divides) (auto simp: Carmichael_prime) hence "p dvd (2 ^ n - 1)" by (simp add: cong_0_iff) thus ?thesis by simp qed auto qed auto finally show ?thesis . qed (auto simp: bernoulli_denom_def) corollary bernoulli_bound: assumes "n > 0" shows "bernoulli_denom n \ 2 * (2 ^ n - 1)" proof - from assms have "2 ^ n > (1 :: nat)" by (intro one_less_power) auto thus ?thesis by (intro dvd_imp_le[OF bernoulli_denom_dvd]) auto qed text \ It can also be shown fairly easily from the von Staudt--Clausen theorem that if \p\ is prime and \2p + 1\ is not, then $B_{2p} \equiv \frac{1}{6}\ (\text{mod}\ 1)$ or, equivalently, the denominator of $B_{2p}$ is 6 and the numerator is of the form $6k+1$. This is the case e.\,g.\ for any primes of the form $3k+1$ or $5k+2$. \ lemma bernoulli_denom_prime_nonprime: assumes "prime p" "\prime (2 * p + 1)" shows "bernoulli (2 * p) - 1 / 6 \ \" "[bernoulli_num (2 * p) = 1] (mod 6)" "bernoulli_denom (2 * p) = 6" proof - from assms have "p > 0" using prime_gt_0_nat by auto define P where "P = {q. prime q \ (q - 1) dvd (2 * p)}" have P_eq: "P = {2, 3}" proof (intro equalityI subsetI) fix q assume "q \ P" hence q: "prime q" "(q - 1) dvd (2 * p)" by (simp_all add: P_def) have "q - 1 \ {1, 2, p, 2 * p}" proof - obtain b c where bc: "b dvd 2" "c dvd p" "q - 1 = b * c" using division_decomp[OF q(2)] by auto from bc have "b \ {1, 2}" and "c \ {1, p}" using prime_nat_iff two_is_prime_nat \prime p\ by blast+ with bc show ?thesis by auto qed hence "q \ {2, 3, p + 1, 2 * p + 1}" using prime_gt_0_nat[OF \prime q\] by force moreover have "q \ p + 1" proof assume [simp]: "q = p + 1" have "even q \ even p" by auto with \prime q\ and \prime p\ have "p = 2" using prime_odd_nat[of p] prime_odd_nat[of q] prime_gt_1_nat[of p] prime_gt_1_nat[of q] by force with assms show False by (simp add: cong_def) qed ultimately show "q \ {2, 3}" using assms \prime q\ by auto qed (auto simp: P_def) show [simp]: "bernoulli_denom (2 * p) = 6" using \p > 0\ P_eq by (subst bernoulli_denom_even) (auto simp: P_def) have "bernoulli (2 * p) + 5 / 6 \ \" using \p > 0\ P_eq vonStaudt_Clausen[of p] by (auto simp: P_def) hence "bernoulli (2 * p) + 5 / 6 - 1 \ \" by (intro Ints_diff) auto thus "bernoulli (2 * p) - 1 / 6 \ \" by simp then obtain a where "of_int a = bernoulli (2 * p) - 1 / 6" by (elim Ints_cases) auto hence "real_of_int a = real_of_int (bernoulli_num (2 * p) - 1) / 6" by (auto simp: bernoulli_conv_num_denom) hence "bernoulli_num (2 * p) - 1 = 6 * a" by simp thus "[bernoulli_num (2 * p) = 1] (mod 6)" by (auto simp: cong_iff_dvd_diff) qed subsection \Akiyama--Tanigawa algorithm\ text \ First, we define the Akiyama--Tanigawa number triangle as shown by Kaneko~\cite{kaneko2000}. We define this generically, parametrised by the first row. This makes the proofs a little bit more modular. \ fun gen_akiyama_tanigawa :: "(nat \ real) \ nat \ nat \ real" where "gen_akiyama_tanigawa f 0 m = f m" | "gen_akiyama_tanigawa f (Suc n) m = real (Suc m) * (gen_akiyama_tanigawa f n m - gen_akiyama_tanigawa f n (Suc m))" lemma gen_akiyama_tanigawa_0 [simp]: "gen_akiyama_tanigawa f 0 = f" by (simp add: fun_eq_iff) text \ The ``regular'' Akiyama--Tanigawa triangle is the one that is used for reading off Bernoulli numbers: \ definition akiyama_tanigawa where "akiyama_tanigawa = gen_akiyama_tanigawa (\n. 1 / real (Suc n))" context begin private definition AT_fps :: "(nat \ real) \ nat \ real fps" where "AT_fps f n = (fps_X - 1) * Abs_fps (gen_akiyama_tanigawa f n)" private lemma AT_fps_Suc: "AT_fps f (Suc n) = (fps_X - 1) * fps_deriv (AT_fps f n)" proof (rule fps_ext) fix m :: nat show "AT_fps f (Suc n) $ m = ((fps_X - 1) * fps_deriv (AT_fps f n)) $ m" by (cases m) (simp_all add: AT_fps_def fps_deriv_def algebra_simps) qed private lemma AT_fps_altdef: "AT_fps f n = (\m\n. fps_const (real (Stirling n m)) * (fps_X - 1)^m * (fps_deriv ^^ m) (AT_fps f 0))" proof - have "AT_fps f n = (fps_XD' (fps_X - 1) ^^ n) (AT_fps f 0)" by (induction n) (simp_all add: AT_fps_Suc fps_XD'_def) also have "\ = (\m\n. fps_const (real (Stirling n m)) * (fps_X - 1) ^ m * (fps_deriv ^^ m) (AT_fps f 0))" by (rule fps_XD'_funpow_affine) simp_all finally show ?thesis . qed private lemma AT_fps_0_nth: "AT_fps f 0 $ n = (if n = 0 then -f 0 else f (n - 1) - f n)" by (simp add: AT_fps_def algebra_simps) text \ The following fact corresponds to Proposition 1 in Kaneko's proof: \ lemma gen_akiyama_tanigawa_n_0: "gen_akiyama_tanigawa f n 0 = (\k\n. (- 1) ^ k * fact k * real (Stirling (Suc n) (Suc k)) * f k)" proof (cases "n = 0") case False note [simp del] = gen_akiyama_tanigawa.simps have "gen_akiyama_tanigawa f n 0 = -(AT_fps f n $ 0)" by (simp add: AT_fps_def) also have "AT_fps f n $ 0 = (\k\n. real (Stirling n k) * (- 1) ^ k * (fact k * AT_fps f 0 $ k))" by (subst AT_fps_altdef) (simp add: fps_sum_nth fps_nth_power_0 fps_0th_higher_deriv) also have "\ = (\k\n. real (Stirling n k) * (- 1) ^ k * (fact k * (f (k - 1) - f k)))" using False by (intro sum.cong refl) (auto simp: Stirling_n_0 AT_fps_0_nth) also have "\ = (\k\n. fact k * (real (Stirling n k) * (- 1) ^ k) * f (k - 1)) - (\k\n. fact k * (real (Stirling n k) * (- 1) ^ k) * f k)" (is "_ = sum ?f _ - ?S2") by (simp add: sum_subtractf algebra_simps) also from False have "sum ?f {..n} = sum ?f {0<..n}" by (intro sum.mono_neutral_right) (auto simp: Stirling_n_0) also have "\ = sum ?f {0<..Suc n}" by (intro sum.mono_neutral_left) auto also have "{0<..Suc n} = {Suc 0..Suc n}" by auto also have "sum ?f \ = sum (\n. ?f (Suc n)) {0..n}" by (subst sum.atLeast_Suc_atMost_Suc_shift) simp_all also have "{0..n} = {..n}" by auto also have "sum (\n. ?f (Suc n)) \ - ?S2 = (\k\n. -((-1)^k * fact k * real (Stirling (Suc n) (Suc k)) * f k))" by (subst sum_subtractf [symmetric], intro sum.cong) (simp_all add: algebra_simps) also have "-\ = (\k\n. ((-1)^k * fact k * real (Stirling (Suc n) (Suc k)) * f k))" by (simp add: sum_negf) finally show ?thesis . qed simp_all text \ The following lemma states that for $A(x) := \sum_{k=0}^\infty a_{0,k} x^k$, we have \[\sum_{n=0}^\infty a_{n,0}\frac{x^n}{n!} = e^x A(1 - e^x)\] which correspond's to Kaneko's remark at the end of Section 2. This seems to be easier to formalise than his actual proof of his Theorem 1, since his proof contains an infinite sum of formal power series, and it was unclear to us how to capture this formally. \ lemma gen_akiyama_tanigawa_fps: "Abs_fps (\n. gen_akiyama_tanigawa f n 0 / fact n) = fps_exp 1 * fps_compose (Abs_fps f) (1 - fps_exp 1)" proof (rule fps_ext) fix n :: nat have "(fps_const (fact n) * (fps_compose (Abs_fps (\n. gen_akiyama_tanigawa f 0 n)) (1 - fps_exp 1) * fps_exp 1)) $ n = (\m\n. \k\m. (1 - fps_exp 1) ^ k $ m * fact n / fact (n - m) * f k)" unfolding fps_mult_left_const_nth by (simp add: fps_times_def fps_compose_def gen_akiyama_tanigawa_n_0 sum_Stirling_binomial field_simps sum_distrib_left sum_distrib_right atLeast0AtMost del: Stirling.simps of_nat_Suc) also have "\ = (\m\n. \k\m. (-1)^k * fact k * real (Stirling m k) * real (n choose m) * f k)" proof (intro sum.cong refl, goal_cases) case (1 m k) have "(1 - fps_exp 1 :: real fps) ^ k = (-fps_exp 1 + 1 :: real fps) ^ k" by simp also have "\ = (\i\k. of_nat (k choose i) * (-1) ^ i * fps_exp (real i))" by (subst binomial_ring) (simp add: atLeast0AtMost power_minus' fps_exp_power_mult mult.assoc) also have "\ = (\i\k. fps_const (real (k choose i) * (-1) ^ i) * fps_exp (real i))" by (simp add: fps_const_mult [symmetric] fps_of_nat fps_const_power [symmetric] fps_const_neg [symmetric] del: fps_const_mult fps_const_power fps_const_neg) also have "\ $ m = (\i\k. real (k choose i) * (- 1) ^ i * real i ^ m) / fact m" (is "_ = ?S / _") by (simp add: fps_sum_nth sum_divide_distrib [symmetric]) also have "?S = (-1) ^ k * (\i\k. (-1) ^ (k - i) * real (k choose i) * real i ^ m)" by (subst sum_distrib_left, intro sum.cong refl) (auto simp: minus_one_power_iff) also have "(\i\k. (-1) ^ (k - i) * real (k choose i) * real i ^ m) = real (Stirling m k) * fact k" by (subst Stirling_closed_form) (simp_all add: field_simps) finally have *: "(1 - fps_exp 1 :: real fps) ^ k $ m * fact n / fact (n - m) = (- 1) ^ k * fact k * real (Stirling m k) * real (n choose m)" using 1 by (simp add: binomial_fact del: of_nat_Suc) show ?case using 1 by (subst *) simp qed also have "\ = (\m\n. \k\n. (- 1) ^ k * fact k * real (Stirling m k) * real (n choose m) * f k)" by (rule sum.cong[OF refl], rule sum.mono_neutral_left) auto also have "\ = (\k\n. \m\n. (- 1) ^ k * fact k * real (Stirling m k) * real (n choose m) * f k)" by (rule sum.swap) also have "\ = gen_akiyama_tanigawa f n 0" by (simp add: gen_akiyama_tanigawa_n_0 sum_Stirling_binomial sum_distrib_left sum_distrib_right mult.assoc atLeast0AtMost del: Stirling.simps) finally show "Abs_fps (\n. gen_akiyama_tanigawa f n 0 / fact n) $ n = (fps_exp 1 * (Abs_fps f oo 1 - fps_exp 1)) $ n" by (subst (asm) fps_mult_left_const_nth) (simp add: field_simps del: of_nat_Suc) qed text \ As Kaneko notes in his afore-mentioned remark, if we let $a_{0,k} = \frac{1}{k+1}$, we obtain \[A(z) = \sum_{k=0}^\infty \frac{x^k}{k+1} = -\frac{\ln (1 - x)}{x}\] and therefore \[\sum_{n=0}^\infty a_{n,0} \frac{x^n}{n!} = \frac{x e^x}{e^x - 1} = \frac{x}{1 - e^{-x}},\] which immediately gives us the connection to the positive Bernoulli numbers. \ theorem bernoulli'_conv_akiyama_tanigawa: "bernoulli' n = akiyama_tanigawa n 0" proof - define f where "f = (\n. 1 / real (Suc n))" note gen_akiyama_tanigawa_fps[of f] also { have "fps_ln 1 = fps_X * Abs_fps (\n. (-1)^n / real (Suc n))" by (intro fps_ext) (simp del: of_nat_Suc add: fps_ln_def) hence "fps_ln 1 / fps_X = Abs_fps (\n. (-1)^n / real (Suc n))" by (metis fps_X_neq_zero nonzero_mult_div_cancel_left) also have "fps_compose \ (-fps_X) = Abs_fps f" by (simp add: fps_compose_uminus' fps_eq_iff f_def) finally have "Abs_fps f = fps_compose (fps_ln 1 / fps_X) (-fps_X)" .. also have "fps_ln 1 / fps_X oo - fps_X oo 1 - fps_exp (1::real) = fps_ln 1 / fps_X oo fps_exp 1 - 1" by (subst fps_compose_assoc [symmetric]) (simp_all add: fps_compose_uminus) also have "\ = (fps_ln 1 oo fps_exp 1 - 1) / (fps_exp 1 - 1)" by (subst fps_compose_divide_distrib) auto also have "\ = fps_X / (fps_exp 1 - 1)" by (simp add: fps_ln_fps_exp_inv fps_inv_fps_exp_compose) finally have "Abs_fps f oo 1 - fps_exp 1 = fps_X / (fps_exp 1 - 1)" . } also have "fps_exp (1::real) - 1 = (1 - fps_exp (-1)) * fps_exp 1" by (simp add: algebra_simps fps_exp_add_mult [symmetric]) also have "fps_exp 1 * (fps_X / \) = bernoulli'_fps" unfolding bernoulli'_fps_def by (subst dvd_div_mult2_eq) (auto simp: fps_dvd_iff intro!: subdegree_leI) finally have "Abs_fps (\n. gen_akiyama_tanigawa f n 0 / fact n) = bernoulli'_fps" . thus ?thesis by (simp add: fps_eq_iff akiyama_tanigawa_def f_def) qed theorem bernoulli_conv_akiyama_tanigawa: "bernoulli n = akiyama_tanigawa n 0 - (if n = 1 then 1 else 0)" using bernoulli'_conv_akiyama_tanigawa[of n] by (auto simp: bernoulli_conv_bernoulli') end end subsection \Efficient code\ text \ We can now compute parts of the Akiyama--Tanigawa (and thereby Bernoulli numbers) with reasonable efficiency but iterating the recurrence row by row. We essentially start with some finite prefix of the zeroth row, say of length $n$, and then apply the recurrence one to get a prefix of the first row of length $n - 1$ etc. \ fun akiyama_tanigawa_step_aux :: "nat \ real list \ real list" where "akiyama_tanigawa_step_aux m (x # y # xs) = real m * (x - y) # akiyama_tanigawa_step_aux (Suc m) (y # xs)" | "akiyama_tanigawa_step_aux m xs = []" lemma length_akiyama_tanigawa_step_aux [simp]: "length (akiyama_tanigawa_step_aux m xs) = length xs - 1" by (induction m xs rule: akiyama_tanigawa_step_aux.induct) simp_all lemma akiyama_tanigawa_step_aux_eq_Nil_iff [simp]: "akiyama_tanigawa_step_aux m xs = [] \ length xs < 2" by (subst length_0_conv [symmetric]) auto lemma nth_akiyama_tanigawa_step_aux: "n < length xs - 1 \ akiyama_tanigawa_step_aux m xs ! n = real (m + n) * (xs ! n - xs ! Suc n)" proof (induction m xs arbitrary: n rule: akiyama_tanigawa_step_aux.induct) case (1 m x y xs n) thus ?case by (cases n) auto qed auto definition gen_akiyama_tanigawa_row where "gen_akiyama_tanigawa_row f n l u = map (gen_akiyama_tanigawa f n) [l.. l \ u" by (auto simp add: gen_akiyama_tanigawa_row_def) lemma nth_gen_akiyama_tanigawa_row: "i < u - l \ gen_akiyama_tanigawa_row f n l u ! i = gen_akiyama_tanigawa f n (i + l)" by (simp add: gen_akiyama_tanigawa_row_def add_ac) lemma gen_akiyama_tanigawa_row_0 [code]: "gen_akiyama_tanigawa_row f 0 l u = map f [l.. l \ u" by (auto simp add: akiyama_tanigawa_row_def) lemma nth_akiyama_tanigawa_row: "i < u - l \ akiyama_tanigawa_row n l u ! i = akiyama_tanigawa n (i + l)" by (simp add: akiyama_tanigawa_row_def add_ac) lemma akiyama_tanigawa_row_0 [code]: "akiyama_tanigawa_row 0 l u = map (\n. inverse (real (Suc n))) [l.. n = 1 \ odd n") case False thus ?thesis by (auto simp add: bernoulli_conv_akiyama_tanigawa) qed (auto simp: bernoulli_odd_eq_0) lemma bernoulli'_code [code]: "bernoulli' n = (if n = 0 then 1 else if n = 1 then 1/2 else if odd n then 0 else akiyama_tanigawa n 0)" by (simp add: bernoulli'_def bernoulli_code) text \ Evaluation with the simplifier is much slower than by reflection, but can still be done with much better efficiency than before: \ lemmas eval_bernoulli = akiyama_tanigawa_code akiyama_tanigawa_row_numeral numeral_2_eq_2 [symmetric] akiyama_tanigawa_row_Suc upt_conv_Cons akiyama_tanigawa_row_0 bernoulli_code[of "numeral n" for n] lemmas eval_bernoulli' = eval_bernoulli bernoulli'_code[of "numeral n" for n] lemmas eval_bernpoly = bernpoly_def atMost_nat_numeral power_eq_if binomial_fact fact_numeral eval_bernoulli (* This should only take a few seconds *) lemma bernoulli_upto_20 [simp]: "bernoulli 2 = 1 / 6" "bernoulli 4 = -(1 / 30)" "bernoulli 6 = 1 / 42" "bernoulli 8 = - (1 / 30)" "bernoulli 10 = 5 / 66" "bernoulli 12 = - (691 / 2730)" "bernoulli 14 = 7 / 6" "bernoulli 16 = -(3617 / 510)" "bernoulli 18 = 43867 / 798" "bernoulli 20 = -(174611 / 330)" by (simp_all add: eval_bernoulli) lemma bernoulli'_upto_20 [simp]: "bernoulli' 2 = 1 / 6" "bernoulli' 4 = -(1 / 30)" "bernoulli' 6 = 1 / 42" "bernoulli' 8 = - (1 / 30)" "bernoulli' 10 = 5 / 66" "bernoulli' 12 = - (691 / 2730)" "bernoulli' 14 = 7 / 6" "bernoulli' 16 = -(3617 / 510)" "bernoulli' 18 = 43867 / 798" "bernoulli' 20 = -(174611 / 330)" by (simp_all add: bernoulli'_def) end diff --git a/thys/Bernoulli/ROOT b/thys/Bernoulli/ROOT --- a/thys/Bernoulli/ROOT +++ b/thys/Bernoulli/ROOT @@ -1,16 +1,17 @@ chapter AFP session Bernoulli (AFP) = "HOL-Complex_Analysis" + options [timeout = 600] sessions "HOL-Analysis" "HOL-Number_Theory" + "HOL-Combinatorics" "HOL-Computational_Algebra" theories Bernoulli Periodic_Bernpoly Bernoulli_FPS Bernoulli_Zeta document_files "root.tex" "root.bib" diff --git a/thys/CakeML/ROOT b/thys/CakeML/ROOT --- a/thys/CakeML/ROOT +++ b/thys/CakeML/ROOT @@ -1,68 +1,69 @@ chapter AFP session LEM (AFP) in "generated" = HOL + options [timeout = 300] sessions Word_Lib theories Lem_pervasives Lem_pervasives_extra Lem_list_extra Lem_set_extra Lem_string Lem_string_extra session CakeML (AFP) = LEM + options [timeout = 3600] sessions + "HOL-Combinatorics" Coinductive IEEE_Floating_Point Show directories "doc" "generated/CakeML" "Tests" theories "doc/Doc_Generated" "generated/CakeML/Ast" "generated/CakeML/AstAuxiliary" "generated/CakeML/BigSmallInvariants" "generated/CakeML/BigStep" "generated/CakeML/Evaluate" "generated/CakeML/Lib" "generated/CakeML/LibAuxiliary" "generated/CakeML/Namespace" "generated/CakeML/NamespaceAuxiliary" "generated/CakeML/PrimTypes" "generated/CakeML/SemanticPrimitives" "generated/CakeML/SemanticPrimitivesAuxiliary" "generated/CakeML/SimpleIO" "generated/CakeML/Tokens" "generated/CakeML/TypeSystem" "generated/CakeML/TypeSystemAuxiliary" "doc/Doc_Proofs" Semantic_Extras Evaluate_Termination Evaluate_Clock Evaluate_Single Big_Step_Determ Big_Step_Fun_Equiv Big_Step_Total Big_Step_Unclocked Big_Step_Unclocked_Single Big_Step_Clocked Matching CakeML_Code CakeML_Quickcheck CakeML_Compiler theories [condition = "ISABELLE_CAKEML_HOME,ISABELLE_CC", document = false] "Tests/Compiler_Test" theories [condition = "ISABELLE_GHC", document = false] "Tests/Code_Test_Haskell" document_files "root.tex" diff --git a/thys/CakeML/generated/LemExtraDefs.thy b/thys/CakeML/generated/LemExtraDefs.thy --- a/thys/CakeML/generated/LemExtraDefs.thy +++ b/thys/CakeML/generated/LemExtraDefs.thy @@ -1,1257 +1,1257 @@ (*========================================================================*) (* Lem *) (* *) (* Dominic Mulligan, University of Cambridge *) (* Francesco Zappa Nardelli, INRIA Paris-Rocquencourt *) (* Gabriel Kerneis, University of Cambridge *) (* Kathy Gray, University of Cambridge *) (* Peter Boehm, University of Cambridge (while working on Lem) *) (* Peter Sewell, University of Cambridge *) (* Scott Owens, University of Kent *) (* Thomas Tuerk, University of Cambridge *) (* Brian Campbell, University of Edinburgh *) (* Shaked Flur, University of Cambridge *) (* Thomas Bauereiss, University of Cambridge *) (* Stephen Kell, University of Cambridge *) (* Thomas Williams *) (* Lars Hupel *) (* Basile Clement *) (* *) (* The Lem sources are copyright 2010-2018 *) (* by the authors above and Institut National de Recherche en *) (* Informatique et en Automatique (INRIA). *) (* *) (* All files except ocaml-lib/pmap.{ml,mli} and ocaml-libpset.{ml,mli} *) (* are distributed under the license below. The former are distributed *) (* under the LGPLv2, as in the LICENSE file. *) (* *) (* *) (* Redistribution and use in source and binary forms, with or without *) (* modification, are permitted provided that the following conditions *) (* are met: *) (* 1. Redistributions of source code must retain the above copyright *) (* notice, this list of conditions and the following disclaimer. *) (* 2. Redistributions in binary form must reproduce the above copyright *) (* notice, this list of conditions and the following disclaimer in the *) (* documentation and/or other materials provided with the distribution. *) (* 3. The names of the authors may not be used to endorse or promote *) (* products derived from this software without specific prior written *) (* permission. *) (* *) (* THIS SOFTWARE IS PROVIDED BY THE AUTHORS ``AS IS'' AND ANY EXPRESS *) (* OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED *) (* WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE *) (* ARE DISCLAIMED. IN NO EVENT SHALL THE AUTHORS BE LIABLE FOR ANY *) (* DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL *) (* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE *) (* GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS *) (* INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER *) (* IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR *) (* OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, EVEN *) (* IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *) (*========================================================================*) chapter \Auxiliary Definitions needed by Lem\ theory "LemExtraDefs" imports Main + "HOL-Combinatorics.List_Permutation" "HOL-Library.While_Combinator" - "HOL-Library.List_Permutation" begin hide_const (open) sign subsection \General\ consts failwith :: " 'a \ 'b" subsection \Lists\ fun index :: " 'a list \ nat \ 'a option " where "index [] n = None" | "index (x # xs) 0 = Some x" | "index (x # xs) (Suc n) = index xs n" lemma index_eq_some [simp]: "index l n = Some x \ (n < length l \ (x = l ! n))" proof (induct l arbitrary: n x) case Nil thus ?case by simp next case (Cons e es n x) note ind_hyp = this show ?case proof (cases n) case 0 thus ?thesis by auto next case (Suc n') with ind_hyp show ?thesis by simp qed qed lemma index_eq_none [simp]: "index l n = None \ length l \ n" by (rule iffD1[OF Not_eq_iff]) auto lemma index_simps [simp]: "length l \ n \ index l n = None" "n < length l \ index l n = Some (l ! n)" by (simp_all) fun find_indices :: "('a \ bool) \ 'a list \ nat list" where "find_indices P [] = []" | "find_indices P (x # xs) = (if P x then 0 # (map Suc (find_indices P xs)) else (map Suc (find_indices P xs)))" lemma length_find_indices : "length (find_indices P l) \ length l" by (induct l) auto lemma sorted_map_suc : "sorted l \ sorted (map Suc l)" by (induct l) (simp_all) lemma sorted_find_indices : "sorted (find_indices P xs)" proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) from sorted_map_suc[OF this] show ?case by (simp) qed lemma find_indices_set [simp] : "set (find_indices P l) = {i. i < length l \ P (l ! i)}" proof (intro set_eqI) fix i show "i \ set (find_indices P l) \ (i \ {i. i < length l \ P (l ! i)})" proof (induct l arbitrary: i) case Nil thus ?case by simp next case (Cons x l' i) note ind_hyp = this show ?case proof (cases i) case 0 thus ?thesis by auto next case (Suc i') with ind_hyp[of i'] show ?thesis by auto qed qed qed definition find_index where "find_index P xs = (case find_indices P xs of [] \ None | i # _ \ Some i)" lemma find_index_eq_some [simp] : "(find_index P xs = Some ii) \ (ii < length xs \ P (xs ! ii) \ (\i' < ii. \(P (xs ! i'))))" (is "?lhs = ?rhs") proof (cases "find_indices P xs") case Nil with find_indices_set[of P xs] show ?thesis unfolding find_index_def by auto next case (Cons i il) note find_indices_eq = this from sorted_find_indices[of P xs] find_indices_eq have "sorted (i # il)" by simp hence i_leq: "\i'. i' \ set (i # il) \ i \ i'" by auto from find_indices_set[of P xs, unfolded find_indices_eq] have set_i_il_eq:"\i'. i' \ set (i # il) = (i' < length xs \ P (xs ! i'))" by simp have lhs_eq: "find_index P xs = Some i" unfolding find_index_def find_indices_eq by simp show ?thesis proof (intro iffI) assume ?lhs with lhs_eq have ii_eq[simp]: "ii = i" by simp from set_i_il_eq[of i] i_leq[unfolded set_i_il_eq] show ?rhs by auto (metis leD less_trans) next assume ?rhs with set_i_il_eq[of ii] have "ii \ set (i # il) \ (ii \ i)" by (metis leI length_pos_if_in_set nth_Cons_0 nth_mem set_i_il_eq) with i_leq [of ii] have "i = ii" by simp thus ?lhs unfolding lhs_eq by simp qed qed lemma find_index_eq_none [simp] : "(find_index P xs = None) \ (\x \ set xs. \(P x))" (is "?lhs = ?rhs") proof (rule iffD1[OF Not_eq_iff], intro iffI) assume "\ ?lhs" then obtain i where "find_index P xs = Some i" by auto hence "i < length xs \ P (xs ! i)" by simp thus "\ ?rhs" by auto next let ?p = "(\i. i < length xs \ P(xs ! i))" assume "\ ?rhs" then obtain i where "?p i" by (metis in_set_conv_nth) from LeastI [of ?p, OF \?p i\] have "?p (Least ?p)" . hence "find_index P xs = Some (Least ?p)" by (subst find_index_eq_some) (metis (mono_tags) less_trans not_less_Least) thus "\ ?lhs" by blast qed definition genlist where "genlist f n = map f (upt 0 n)" lemma genlist_length [simp] : "length (genlist f n) = n" unfolding genlist_def by simp lemma genlist_simps [simp]: "genlist f 0 = []" "genlist f (Suc n) = genlist f n @ [f n]" unfolding genlist_def by auto definition split_at where "split_at n l = (take n l, drop n l)" fun delete_first :: "('a \ bool) \ 'a list \ ('a list) option " where "delete_first P [] = None" | "delete_first P (x # xs) = (if (P x) then Some xs else map_option (\xs'. x # xs') (delete_first P xs))" declare delete_first.simps [simp del] lemma delete_first_simps [simp] : "delete_first P [] = None" "P x \ delete_first P (x # xs) = Some xs" "\(P x) \ delete_first P (x # xs) = map_option (\xs'. x # xs') (delete_first P xs)" unfolding delete_first.simps by auto lemmas delete_first_unroll = delete_first.simps(2) lemma delete_first_eq_none [simp] : "delete_first P l = None \ (\x \ set l. \ (P x))" by (induct l) (auto simp add: delete_first_unroll) lemma delete_first_eq_some : "delete_first P l = (Some l') \ (\l1 x l2. P x \ (\x \ set l1. \(P x)) \ (l = l1 @ (x # l2)) \ (l' = l1 @ l2))" (is "?lhs l l' = (\l1 x l2. ?rhs_body l1 x l2 l l')") proof (induct l arbitrary: l') case Nil thus ?case by simp next case (Cons e l l') note ind_hyp = this show ?case proof (cases "P e") case True show ?thesis proof (rule iffI) assume "?lhs (e # l) l'" with \P e\ have "l = l'" by simp with \P e\ have "?rhs_body [] e l' (e # l) l'" by simp thus "\l1 x l2. ?rhs_body l1 x l2 (e # l) l'" by blast next assume "\l1 x l2. ?rhs_body l1 x l2 (e # l) l'" then obtain l1 x l2 where body_ok: "?rhs_body l1 x l2 (e # l) l'" by blast from body_ok \P e\ have l1_eq[simp]: "l = l'" by (cases l1) (simp_all) with \P e\ show "?lhs (e # l) l'" by simp qed next case False define rhs_pred where "rhs_pred \ \l1 x l2 l l'. ?rhs_body l1 x l2 l l'" have rhs_fold: "\l1 x l2 l l'. ?rhs_body l1 x l2 l l' = rhs_pred l1 x l2 l l'" unfolding rhs_pred_def by simp have "(\z l1 x l2. rhs_pred l1 x l2 l z \ e # z = l') = (\l1 x l2. rhs_pred l1 x l2 (e # l) l')" proof (intro iffI) assume "\z l1 x l2. rhs_pred l1 x l2 l z \ e # z = l'" then obtain z l1 x l2 where "rhs_pred l1 x l2 l z" and l'_eq: "l' = e # z" by auto with \\(P e)\ have "rhs_pred (e # l1) x l2 (e # l) l'" unfolding rhs_pred_def by simp thus "\l1 x l2. rhs_pred l1 x l2 (e # l) l'" by blast next assume "\l1 x l2. rhs_pred l1 x l2 (e # l) l'" then obtain l1 x l2 where "rhs_pred l1 x l2 (e # l) l'" by blast with \\ (P e)\ obtain l1' where l1_eq[simp]: "l1 = e # l1'" unfolding rhs_pred_def by (cases l1) (auto) with \rhs_pred l1 x l2 (e # l) l'\ have "rhs_pred l1' x l2 l (l1' @ l2) \ e # (l1' @ l2) = l'" unfolding rhs_pred_def by (simp) thus "\z l1 x l2. rhs_pred l1 x l2 l z \ e # z = l'" by blast qed with \\ P e\ show ?thesis unfolding rhs_fold by (simp add: ind_hyp[unfolded rhs_fold]) qed qed lemma perm_eval [code] : "perm [] l \ l = []" (is ?g1) "perm (x # xs) l \ (case delete_first (\e. e = x) l of None => False | Some l' => perm xs l')" (is ?g2) proof - show ?g1 by auto next show ?g2 proof (cases "delete_first (\e. e = x) l") case None note del_eq = this hence "x \ set l" by auto with perm_set_eq [of "x # xs" l] have "\ perm (x # xs) l" by auto thus ?thesis unfolding del_eq by simp next case (Some l') note del_eq = this from del_eq[unfolded delete_first_eq_some] obtain l1 l2 where l_eq: "l = l1 @ [x] @ l2" and l'_eq: "l' = l1 @ l2" by auto have "(x # xs <~~> l1 @ x # l2) = (xs <~~> l1 @ l2)" proof - from perm_append_swap [of l1 "[x]"] perm_append2 [of "l1 @ [x]" "x # l1" l2] have "l1 @ x # l2 <~~> x # (l1 @ l2)" by simp hence "x # xs <~~> l1 @ x # l2 \ x # xs <~~> x # (l1 @ l2)" by (metis perm.trans perm_sym) thus ?thesis by simp qed with del_eq l_eq l'_eq show ?thesis by simp qed qed fun sorted_by :: "('a \ 'a \ bool)\ 'a list \ bool " where "sorted_by cmp [] = True" | "sorted_by cmp [_] = True" | "sorted_by cmp (x1 # x2 # xs) = ((cmp x1 x2) \ sorted_by cmp (x2 # xs))" lemma sorted_by_lesseq [simp] : "sorted_by ((\) :: ('a::{linorder}) => 'a => bool) = sorted" proof (rule ext) fix l :: "'a list" show "sorted_by (\) l = sorted l" proof (induct l) case Nil thus ?case by simp next case (Cons x xs) thus ?case by (cases xs) (simp_all del: sorted.simps(2) add: sorted2_simps) qed qed lemma sorted_by_cons_imp : "sorted_by cmp (x # xs) \ sorted_by cmp xs" by (cases xs) simp_all lemma sorted_by_cons_trans : assumes trans_cmp: "transp cmp" shows "sorted_by cmp (x # xs) = ((\x' \ set xs . cmp x x') \ sorted_by cmp xs)" proof (induct xs arbitrary: x) case Nil thus ?case by simp next case (Cons x2 xs x1) note ind_hyp = this from trans_cmp show ?case by (auto simp add: ind_hyp transp_def) qed fun insert_sort_insert_by :: "('a \ 'a \ bool)\ 'a \ 'a list \ 'a list " where "insert_sort_insert_by cmp e ([]) = ( [e])" | "insert_sort_insert_by cmp e (x # xs) = ( if cmp e x then (e # (x # xs)) else x # (insert_sort_insert_by cmp e xs))" lemma insert_sort_insert_by_length [simp] : "length (insert_sort_insert_by cmp e l) = Suc (length l)" by (induct l) auto lemma insert_sort_insert_by_set [simp] : "set (insert_sort_insert_by cmp e l) = insert e (set l)" by (induct l) auto lemma insert_sort_insert_by_perm : "(insert_sort_insert_by cmp e l) <~~> (e # l)" proof (induct l) case Nil thus ?case by simp next case (Cons e2 l') note ind_hyp = this have "e2 # e # l' <~~> e # e2 # l'" by (rule perm.swap) hence "e2 # insert_sort_insert_by cmp e l' <~~> e # e2 # l'" using ind_hyp by (metis cons_perm_eq perm.trans) thus ?case by simp qed lemma insert_sort_insert_by_sorted_by : assumes cmp_cases: "\y. y \ set l \ \ (cmp e y) \ cmp y e" assumes cmp_trans: "transp cmp" shows "sorted_by cmp l \ sorted_by cmp (insert_sort_insert_by cmp e l)" using cmp_cases proof (induct l) case Nil thus ?case by simp next case (Cons x1 l') note ind_hyp = Cons(1) note sorted_x1_l' = Cons(2) note cmp_cases = Cons(3) show ?case proof (cases l') case Nil with cmp_cases show ?thesis by simp next case (Cons x2 l'') note l'_eq = this from l'_eq sorted_x1_l' have "cmp x1 x2" "sorted_by cmp l'" by simp_all show ?thesis proof (cases "cmp e x1") case True with \cmp x1 x2\ \sorted_by cmp l'\ have "sorted_by cmp (x1 # l')" unfolding l'_eq by (simp) with \cmp e x1\ show ?thesis by simp next case False with cmp_cases have "cmp x1 e" by simp have "\x'. x' \ set l' \ cmp x1 x'" proof - fix x' assume "x' \ set l'" hence "x' = x2 \ cmp x2 x'" using \sorted_by cmp l'\ l'_eq sorted_by_cons_trans [OF cmp_trans, of x2 l''] by auto with transpD[OF cmp_trans, of x1 x2 x'] \cmp x1 x2\ show "cmp x1 x'" by blast qed hence "sorted_by cmp (x1 # insert_sort_insert_by cmp e l')" using ind_hyp [OF \sorted_by cmp l'\] \cmp x1 e\ cmp_cases unfolding sorted_by_cons_trans[OF cmp_trans] by simp with \\(cmp e x1)\ show ?thesis by simp qed qed qed fun insert_sort_by :: "('a \ 'a \ bool) \ 'a list \ 'a list " where "insert_sort_by cmp [] = []" | "insert_sort_by cmp (x # xs) = insert_sort_insert_by cmp x (insert_sort_by cmp xs)" lemma insert_sort_by_perm : "(insert_sort_by cmp l) <~~> l" proof (induct l) case Nil thus ?case by simp next case (Cons x l) thus ?case by simp (metis cons_perm_eq insert_sort_insert_by_perm perm.trans) qed lemma insert_sort_by_length [simp]: "length (insert_sort_by cmp l) = length l" by (induct l) auto lemma insert_sort_by_set [simp]: "set (insert_sort_by cmp l) = set l" by (induct l) auto definition sort_by where "sort_by = insert_sort_by" lemma sort_by_simps [simp]: "length (sort_by cmp l) = length l" "set (sort_by cmp l) = set l" unfolding sort_by_def by simp_all lemma sort_by_perm : "sort_by cmp l <~~> l" unfolding sort_by_def by (simp add: insert_sort_by_perm) subsection \Maps\ definition map_image :: "('v \ 'w) \ ('k, 'v) map \ ('k, 'w) map" where "map_image f m = (\k. map_option f (m k))" definition map_domain_image :: "('k \ 'v \ 'w) \ ('k, 'v) map \ ('k, 'w) map" where "map_domain_image f m = (\k. map_option (f k) (m k))" lemma map_image_simps [simp]: "(map_image f m) k = None \ m k = None" "(map_image f m) k = Some x \ (\x'. (m k = Some x') \ (x = f x'))" "(map_image f Map.empty) = Map.empty" "(map_image f (m (k \ v)) = (map_image f m) (k \ f v))" unfolding map_image_def by auto lemma map_image_dom_ran [simp]: "dom (map_image f m) = dom m" "ran (map_image f m) = f ` (ran m)" unfolding dom_def ran_def by auto definition map_to_set :: "('k, 'v) map \ ('k * 'v) set" where "map_to_set m = { (k, v) . m k = Some v }" lemma map_to_set_simps [simp] : "map_to_set Map.empty = {}" (is ?g1) "map_to_set (m ((k::'k) \ (v::'v))) = (insert (k, v) (map_to_set (m |` (- {k}))))" (is ?g2) proof - show ?g1 unfolding map_to_set_def by simp next show ?g2 proof (rule set_eqI) fix kv :: "('k * 'v)" obtain k' v' where kv_eq[simp]: "kv = (k', v')" by (rule prod.exhaust) show "(kv \ map_to_set (m(k \ v))) = (kv \ insert (k, v) (map_to_set (m |` (- {k}))))" by (auto simp add: map_to_set_def) qed qed subsection \Sets\ definition "set_choose s \ (SOME x. (x \ s))" definition without_trans_edges :: "('a \ 'a) set \ ('a \ 'a) set" where "without_trans_edges S \ let ts = trancl S in { (x, y) \ S. \z \ snd ` S. x \ z \ y \ z \ \ ((x, z) \ ts \ (z, y) \ ts)}" definition unbounded_lfp :: "'a set \ ('a set \ 'a set) \ 'a set" where "unbounded_lfp S f \ while (\x. x \ S) f S" definition unbounded_gfp :: "'a set \ ('a set \ 'a set) \ 'a set" where "unbounded_gfp S f \ while (\x. S \ x) f S" lemma set_choose_thm[simp]: "s \ {} \ (set_choose s) \ s" unfolding set_choose_def by (rule someI_ex) auto lemma set_choose_sing [simp]: "set_choose {x} = x" unfolding set_choose_def by auto lemma set_choose_code [code]: "set_choose (set [x]) = x" by auto lemma set_choose_in [intro] : assumes "s \ {}" shows "set_choose s \ s" proof - from \s \ {}\ obtain x where "x \ s" by auto thus ?thesis unfolding set_choose_def by (rule someI) qed definition set_case where "set_case s c_empty c_sing c_else = (if (s = {}) then c_empty else (if (card s = 1) then c_sing (set_choose s) else c_else))" lemma set_case_simps [simp] : "set_case {} c_empty c_sing c_else = c_empty" "set_case {x} c_empty c_sing c_else = c_sing x" "card s > 1 \ set_case s c_empty c_sing c_else = c_else" "\(finite s) \ set_case s c_empty c_sing c_else = c_else" unfolding set_case_def by auto lemma set_case_simp_insert2 [simp] : assumes x12_neq: "x1 \ x2" shows "set_case (insert x1 (insert x2 xs)) c_empty c_sing c_else = c_else" proof (cases "finite xs") case False thus ?thesis by (simp) next case True note fin_xs = this have "card {x1,x2} \ card (insert x1 (insert x2 xs))" by (rule card_mono) (auto simp add: fin_xs) with x12_neq have "1 < card (insert x1 (insert x2 xs))" by simp thus ?thesis by auto qed lemma set_case_code [code] : "set_case (set []) c_empty c_sing c_else = c_empty" "set_case (set [x]) c_empty c_sing c_else = c_sing x" "set_case (set (x1 # x2 # xs)) c_empty c_sing c_else = (if (x1 = x2) then set_case (set (x2 # xs)) c_empty c_sing c_else else c_else)" by auto definition set_lfp:: "'a set \ ('a set \ 'a set) \ 'a set" where "set_lfp s f = lfp (\s'. f s' \ s)" lemma set_lfp_tail_rec_def : assumes mono_f: "mono f" shows "set_lfp s f = (if (f s) \ s then s else (set_lfp (s \ f s) f))" (is "?ls = ?rs") proof (cases "f s \ s") case True note fs_sub_s = this from fs_sub_s have "\{u. f u \ s \ u} = s" by auto hence "?ls = s" unfolding set_lfp_def lfp_def . with fs_sub_s show "?ls = ?rs" by simp next case False note not_fs_sub_s = this from mono_f have mono_f': "mono (\s'. f s' \ s)" unfolding mono_def by auto have "\{u. f u \ s \ u} = \{u. f u \ (s \ f s) \ u}" (is "\?S1 = \?S2") proof have "?S2 \ ?S1" by auto thus "\?S1 \ \?S2" by (rule Inf_superset_mono) next { fix e assume "e \ \?S2" hence S2_prop: "\s'. f s' \ s' \ s \ s' \ f s \ s' \ e \ s'" by simp { fix s' assume "f s' \ s'" "s \ s'" from mono_f \s \ s'\ have "f s \ f s'" unfolding mono_def by simp with \f s' \ s'\ have "f s \ s'" by simp with \f s' \ s'\ \s \ s'\ S2_prop have "e \ s'" by simp } hence "e \ \?S1" by simp } thus "\?S2 \ \?S1" by auto qed hence "?ls = (set_lfp (s \ f s) f)" unfolding set_lfp_def lfp_def . with not_fs_sub_s show "?ls = ?rs" by simp qed lemma set_lfp_simps [simp] : "mono f \ f s \ s \ set_lfp s f = s" "mono f \ \(f s \ s) \ set_lfp s f = (set_lfp (s \ f s) f)" by (metis set_lfp_tail_rec_def)+ fun insert_in_list_at_arbitrary_pos where "insert_in_list_at_arbitrary_pos x [] = {[x]}" | "insert_in_list_at_arbitrary_pos x (y # ys) = insert (x # y # ys) ((\l. y # l) ` (insert_in_list_at_arbitrary_pos x ys))" lemma insert_in_list_at_arbitrary_pos_thm : "xl \ insert_in_list_at_arbitrary_pos x l \ (\l1 l2. l = l1 @ l2 \ xl = l1 @ [x] @ l2)" proof (induct l arbitrary: xl) case Nil thus ?case by simp next case (Cons y l xyl) note ind_hyp = this show ?case proof (rule iffI) assume xyl_in: "xyl \ insert_in_list_at_arbitrary_pos x (y # l)" show "\l1 l2. y # l = l1 @ l2 \ xyl = l1 @ [x] @ l2" proof (cases "xyl = x # y # l") case True hence "y # l = [] @ (y # l) \ xyl = [] @ [x] @ (y # l)" by simp thus ?thesis by blast next case False with xyl_in have "xyl \ (#) y ` insert_in_list_at_arbitrary_pos x l" by simp with ind_hyp obtain l1 l2 where "l = l1 @ l2 \ xyl = y # l1 @ x # l2" by (auto simp add: image_def Bex_def) hence "y # l = (y # l1) @ l2 \ xyl = (y # l1) @ [x] @ l2" by simp thus ?thesis by blast qed next assume "\l1 l2. y # l = l1 @ l2 \ xyl = l1 @ [x] @ l2" then obtain l1 l2 where yl_eq: "y # l = l1 @ l2" and xyl_eq: "xyl = l1 @ [x] @ l2" by blast show "xyl \ insert_in_list_at_arbitrary_pos x (y # l)" proof (cases l1) case Nil with yl_eq xyl_eq have "xyl = x # y # l" by simp thus ?thesis by simp next case (Cons y' l1') with yl_eq have l1_eq: "l1 = y # l1'" and l_eq: "l = l1' @ l2" by simp_all have "\l1'' l2''. l = l1'' @ l2'' \ l1' @ [x] @ l2 = l1'' @ [x] @ l2''" apply (rule_tac exI[where x = l1']) apply (rule_tac exI [where x = l2]) apply (simp add: l_eq) done hence "(l1' @ [x] @ l2) \ insert_in_list_at_arbitrary_pos x l" unfolding ind_hyp by blast hence "\l'. l' \ insert_in_list_at_arbitrary_pos x l \ l1 @ x # l2 = y # l'" by (rule_tac exI [where x = "l1' @ [x] @ l2"]) (simp add: l1_eq) thus ?thesis by (simp add: image_def Bex_def xyl_eq) qed qed qed definition list_of_set_set :: "'a set \ ('a list) set" where "list_of_set_set s = { l . (set l = s) \ distinct l }" lemma list_of_set_set_empty [simp]: "list_of_set_set {} = {[]}" unfolding list_of_set_set_def by auto lemma list_of_set_set_insert [simp] : "list_of_set_set (insert x s) = \ ((insert_in_list_at_arbitrary_pos x) ` (list_of_set_set (s - {x})))" (is "?lhs = ?rhs") proof (intro set_eqI) fix l have "(set l = insert x s \ distinct l) \ (\l1 l2. set (l1 @ l2) = s - {x} \ distinct (l1 @ l2) \ l = l1 @ x # l2)" proof (intro iffI) assume "set l = insert x s \ distinct l" hence set_l_eq: "set l = insert x s" and "distinct l" by simp_all from \set l = insert x s\ have "x \ set l" by simp then obtain l1 l2 where l_eq: "l = l1 @ x # l2" unfolding in_set_conv_decomp by blast from \distinct l\ l_eq have "distinct (l1 @ l2)" and x_nin: "x \ set (l1 @ l2)" by auto from x_nin set_l_eq[unfolded l_eq] have set_l12_eq: "set (l1 @ l2) = s - {x}" by auto from \distinct (l1 @ l2)\ l_eq set_l12_eq show "\l1 l2. set (l1 @ l2) = s - {x} \ distinct (l1 @ l2) \ l = l1 @ x # l2" by blast next assume "\l1 l2. set (l1 @ l2) = s - {x} \ distinct (l1 @ l2) \ l = l1 @ x # l2" then obtain l1 l2 where "set (l1 @ l2) = s - {x}" "distinct (l1 @ l2)" "l = l1 @ x # l2" by blast thus "set l = insert x s \ distinct l" by auto qed thus "l \ list_of_set_set (insert x s) \ l \ (\ ((insert_in_list_at_arbitrary_pos x) ` (list_of_set_set (s - {x}))))" unfolding list_of_set_set_def by (simp add: insert_in_list_at_arbitrary_pos_thm ex_simps[symmetric] del: ex_simps) qed lemma list_of_set_set_code [code]: "list_of_set_set (set []) = {[]}" "list_of_set_set (set (x # xs)) = \ ((insert_in_list_at_arbitrary_pos x) ` (list_of_set_set ((set xs) - {x})))" by simp_all lemma list_of_set_set_is_empty : "list_of_set_set s = {} \ \ (finite s)" proof - have "finite s \ (\l. set l = s \ distinct l)" proof (rule iffI) assume "\l. set l = s \ distinct l" then obtain l where "s = set l" by blast thus "finite s" by simp next assume "finite s" thus "\l. set l = s \ distinct l" proof (induct s) case empty show ?case by auto next case (insert e s) note e_nin_s = insert(2) from insert(3) obtain l where set_l: "set l = s" and dist_l: "distinct l" by blast from set_l have set_el: "set (e # l) = insert e s" by auto from dist_l set_l e_nin_s have dist_el: "distinct (e # l)" by simp from set_el dist_el show ?case by blast qed qed thus ?thesis unfolding list_of_set_set_def by simp qed definition list_of_set :: "'a set \ 'a list" where "list_of_set s = set_choose (list_of_set_set s)" lemma list_of_set [simp] : assumes fin_s: "finite s" shows "set (list_of_set s) = s" "distinct (list_of_set s)" proof - from fin_s list_of_set_set_is_empty[of s] have "\ (list_of_set_set s = {})" by simp hence "list_of_set s \ list_of_set_set s" unfolding list_of_set_def by (rule set_choose_thm) thus "set (list_of_set s) = s" "distinct (list_of_set s)" unfolding list_of_set_set_def by simp_all qed lemma list_of_set_in: "finite s \ list_of_set s \ list_of_set_set s" unfolding list_of_set_def by (metis list_of_set_set_is_empty set_choose_thm) definition ordered_list_of_set where "ordered_list_of_set cmp s = set_choose (sort_by cmp ` list_of_set_set s)" subsection \sum\ find_consts "'a list => ('a list * _)" fun sum_partition :: "('a + 'b) list \ 'a list * 'b list" where "sum_partition [] = ([], [])" | "sum_partition ((Inl l) # lrs) = (let (ll, rl) = sum_partition lrs in (l # ll, rl))" | "sum_partition ((Inr r) # lrs) = (let (ll, rl) = sum_partition lrs in (ll, r # rl))" lemma sum_partition_length : "List.length lrs = List.length (fst (sum_partition lrs)) + List.length (snd (sum_partition lrs))" proof (induct lrs) case Nil thus ?case by simp next case (Cons lr lrs) thus ?case by (cases lr) (auto split: prod.split) qed subsection \sorting\ subsection \Strings\ declare String.literal.explode_inverse [simp] subsection \num to string conversions\ definition nat_list_to_string :: "nat list \ string" where "nat_list_to_string nl = map char_of nl" definition is_digit where "is_digit (n::nat) = (n < 10)" lemma is_digit_simps[simp] : "n < 10 \ is_digit n" "\(n < 10) \ \(is_digit n)" unfolding is_digit_def by simp_all lemma is_digit_expand : "is_digit n \ (n = 0) \ (n = 1) \ (n = 2) \ (n = 3) \ (n = 4) \ (n = 5) \ (n = 6) \ (n = 7) \ (n = 8) \ (n = 9)" unfolding is_digit_def by auto lemmas is_digitE = is_digit_expand[THEN iffD1,elim_format] lemmas is_digitI = is_digit_expand[THEN iffD2,rule_format] definition is_digit_char where "is_digit_char c \ (c = CHR ''0'') \ (c = CHR ''5'') \ (c = CHR ''1'') \ (c = CHR ''6'') \ (c = CHR ''2'') \ (c = CHR ''7'') \ (c = CHR ''3'') \ (c = CHR ''8'') \ (c = CHR ''4'') \ (c = CHR ''9'')" lemma is_digit_char_simps[simp] : "is_digit_char (CHR ''0'')" "is_digit_char (CHR ''1'')" "is_digit_char (CHR ''2'')" "is_digit_char (CHR ''3'')" "is_digit_char (CHR ''4'')" "is_digit_char (CHR ''5'')" "is_digit_char (CHR ''6'')" "is_digit_char (CHR ''7'')" "is_digit_char (CHR ''8'')" "is_digit_char (CHR ''9'')" unfolding is_digit_char_def by simp_all lemmas is_digit_charE = is_digit_char_def[THEN iffD1,elim_format] lemmas is_digit_charI = is_digit_char_def[THEN iffD2,rule_format] definition digit_to_char :: "nat \ char" where "digit_to_char n = ( if n = 0 then CHR ''0'' else if n = 1 then CHR ''1'' else if n = 2 then CHR ''2'' else if n = 3 then CHR ''3'' else if n = 4 then CHR ''4'' else if n = 5 then CHR ''5'' else if n = 6 then CHR ''6'' else if n = 7 then CHR ''7'' else if n = 8 then CHR ''8'' else if n = 9 then CHR ''9'' else CHR ''X'')" lemma digit_to_char_simps [simp]: "digit_to_char 0 = CHR ''0''" "digit_to_char (Suc 0) = CHR ''1''" "digit_to_char 2 = CHR ''2''" "digit_to_char 3 = CHR ''3''" "digit_to_char 4 = CHR ''4''" "digit_to_char 5 = CHR ''5''" "digit_to_char 6 = CHR ''6''" "digit_to_char 7 = CHR ''7''" "digit_to_char 8 = CHR ''8''" "digit_to_char 9 = CHR ''9''" "n > 9 \ digit_to_char n = CHR ''X''" unfolding digit_to_char_def by simp_all definition char_to_digit :: "char \ nat" where "char_to_digit c = ( if c = CHR ''0'' then 0 else if c = CHR ''1'' then 1 else if c = CHR ''2'' then 2 else if c = CHR ''3'' then 3 else if c = CHR ''4'' then 4 else if c = CHR ''5'' then 5 else if c = CHR ''6'' then 6 else if c = CHR ''7'' then 7 else if c = CHR ''8'' then 8 else if c = CHR ''9'' then 9 else 10)" lemma char_to_digit_simps [simp]: "char_to_digit (CHR ''0'') = 0" "char_to_digit (CHR ''1'') = 1" "char_to_digit (CHR ''2'') = 2" "char_to_digit (CHR ''3'') = 3" "char_to_digit (CHR ''4'') = 4" "char_to_digit (CHR ''5'') = 5" "char_to_digit (CHR ''6'') = 6" "char_to_digit (CHR ''7'') = 7" "char_to_digit (CHR ''8'') = 8" "char_to_digit (CHR ''9'') = 9" unfolding char_to_digit_def by simp_all lemma diget_to_char_inv[simp]: assumes is_digit: "is_digit n" shows "char_to_digit (digit_to_char n) = n" using is_digit unfolding is_digit_expand by auto lemma char_to_diget_inv[simp]: assumes is_digit: "is_digit_char c" shows "digit_to_char (char_to_digit c) = c" using is_digit unfolding char_to_digit_def is_digit_char_def by auto lemma char_to_digit_div_mod [simp]: assumes is_digit: "is_digit_char c" shows "char_to_digit c < 10" using is_digit unfolding char_to_digit_def is_digit_char_def by auto lemma is_digit_char_intro[simp]: "is_digit (char_to_digit c) = is_digit_char c" unfolding char_to_digit_def is_digit_char_def is_digit_expand by auto lemma is_digit_intro[simp]: "is_digit_char (digit_to_char n) = is_digit n" unfolding digit_to_char_def is_digit_char_def is_digit_expand by auto lemma digit_to_char_11: "digit_to_char n1 = digit_to_char n2 \ (is_digit n1 = is_digit n2) \ (is_digit n1 \ (n1 = n2))" by (metis diget_to_char_inv is_digit_intro) lemma char_to_digit_11: "char_to_digit c1 = char_to_digit c2 \ (is_digit_char c1 = is_digit_char c2) \ (is_digit_char c1 \ (c1 = c2))" by (metis char_to_diget_inv is_digit_char_intro) function nat_to_string :: "nat \ string" where "nat_to_string n = (if (is_digit n) then [digit_to_char n] else nat_to_string (n div 10) @ [digit_to_char (n mod 10)])" by auto termination by (relation "measure id") (auto simp add: is_digit_def) definition int_to_string :: "int \ string" where "int_to_string i \ if i < 0 then ''-'' @ nat_to_string (nat (abs i)) else nat_to_string (nat i)" lemma nat_to_string_simps[simp]: "is_digit n \ nat_to_string n = [digit_to_char n]" "\(is_digit n) \ nat_to_string n = nat_to_string (n div 10) @ [digit_to_char (n mod 10)]" by simp_all declare nat_to_string.simps[simp del] lemma nat_to_string_neq_nil[simp]: "nat_to_string n \ []" by (cases "is_digit n") simp_all lemmas nat_to_string_neq_nil2[simp] = nat_to_string_neq_nil[symmetric] lemma nat_to_string_char_to_digit [simp]: "is_digit_char c \ nat_to_string (char_to_digit c) = [c]" by auto lemma nat_to_string_11[simp] : "(nat_to_string n1 = nat_to_string n2) \ n1 = n2" proof (rule iffI) assume "n1 = n2" thus "nat_to_string n1 = nat_to_string n2" by simp next assume "nat_to_string n1 = nat_to_string n2" thus "n1 = n2" proof (induct n2 arbitrary: n1 rule: less_induct) case (less n2') note ind_hyp = this(1) note n2s_eq = less(2) have is_dig_eq: "is_digit n2' = is_digit n1" using n2s_eq apply (cases "is_digit n2'") apply (case_tac [!] "is_digit n1") apply (simp_all) done show ?case proof (cases "is_digit n2'") case True with n2s_eq is_dig_eq show ?thesis by simp (metis digit_to_char_11) next case False with is_dig_eq have not_digs : "\ (is_digit n1)" "\ (is_digit n2')" by simp_all from not_digs(2) have "n2' div 10 < n2'" unfolding is_digit_def by auto note ind_hyp' = ind_hyp [OF this, of "n1 div 10"] from not_digs n2s_eq ind_hyp' digit_to_char_11[of "n1 mod 10" "n2' mod 10"] have "(n1 mod 10) = (n2' mod 10)" "n1 div 10 = n2' div 10" by simp_all thus "n1 = n2'" by (metis div_mult_mod_eq) qed qed qed definition "is_nat_string s \ (\c\set s. is_digit_char c)" definition "is_strong_nat_string s \ is_nat_string s \ (s \ []) \ (hd s = CHR ''0'' \ length s = 1)" lemma is_nat_string_simps[simp]: "is_nat_string []" "is_nat_string (c # s) \ is_digit_char c \ is_nat_string s" unfolding is_nat_string_def by simp_all lemma is_strong_nat_string_simps[simp]: "\(is_strong_nat_string [])" "is_strong_nat_string (c # s) \ is_digit_char c \ is_nat_string s \ (c = CHR ''0'' \ s = [])" unfolding is_strong_nat_string_def by simp_all fun string_to_nat_aux :: "nat \ string \ nat" where "string_to_nat_aux n [] = n" | "string_to_nat_aux n (d#ds) = string_to_nat_aux (n*10 + char_to_digit d) ds" definition string_to_nat :: "string \ nat option" where "string_to_nat s \ (if is_nat_string s then Some (string_to_nat_aux 0 s) else None)" definition string_to_nat' :: "string \ nat" where "string_to_nat' s \ the (string_to_nat s)" lemma string_to_nat_aux_inv : assumes "is_nat_string s" assumes "n > 0 \ is_strong_nat_string s" shows "nat_to_string (string_to_nat_aux n s) = (if n = 0 then '''' else nat_to_string n) @ s" using assms proof (induct s arbitrary: n) case Nil thus ?case by (simp add: is_strong_nat_string_def) next case (Cons c s n) from Cons(2) have "is_digit_char c" "is_nat_string s" by simp_all note cs_ok = Cons(3) let ?m = "n*10 + char_to_digit c" note ind_hyp = Cons(1)[OF \is_nat_string s\, of ?m] from \is_digit_char c\ have m_div: "?m div 10 = n" and m_mod: "?m mod 10 = char_to_digit c" unfolding is_digit_char_def by auto show ?case proof (cases "?m = 0") case True with \is_digit_char c\ have "n = 0" "c = CHR ''0''" unfolding is_digit_char_def by auto moreover with cs_ok have "s = []" by simp ultimately show ?thesis by simp next case False note m_neq_0 = this with ind_hyp have ind_hyp': "nat_to_string (string_to_nat_aux ?m s) = (nat_to_string ?m) @ s" by auto hence "nat_to_string (string_to_nat_aux n (c # s)) = (nat_to_string ?m) @ s" by simp with \is_digit_char c\ m_div show ?thesis by simp qed qed lemma string_to_nat_inv : assumes strong_nat_s: "is_strong_nat_string s" assumes s2n_s: "string_to_nat s = Some n" shows "nat_to_string n = s" proof - from strong_nat_s have nat_s: "is_nat_string s" unfolding is_strong_nat_string_def by simp with s2n_s have n_eq: "n = string_to_nat_aux 0 s" unfolding string_to_nat_def by simp from string_to_nat_aux_inv[of s 0, folded n_eq] nat_s strong_nat_s show ?thesis by simp qed lemma nat_to_string_induct [case_names "digit" "non_digit"]: assumes digit: "\d. is_digit d \ P d" assumes not_digit: "\n. \(is_digit n) \ P (n div 10) \ P (n mod 10) \ P n" shows "P n" proof (induct n rule: less_induct) case (less n) note ind_hyp = this show ?case proof (cases "is_digit n") case True with digit show ?thesis by simp next case False note not_dig = this hence "n div 10 < n" "n mod 10 < n" unfolding is_digit_def by auto with not_dig ind_hyp not_digit show ?thesis by simp qed qed lemma nat_to_string___is_nat_string [simp]: "is_nat_string (nat_to_string n)" unfolding is_nat_string_def proof (induct n rule: nat_to_string_induct) case (digit d) thus ?case by simp next case (non_digit n) thus ?case by simp qed lemma nat_to_string___eq_0 [simp]: "(nat_to_string n = (CHR ''0'')#s) \ (n = 0 \ s = [])" unfolding is_nat_string_def proof (induct n arbitrary: s rule: nat_to_string_induct) case (digit d) thus ?case unfolding is_digit_expand by auto next case (non_digit n) obtain c s' where ns_eq: "nat_to_string (n div 10) = c # s'" by (cases "nat_to_string (n div 10)") auto from non_digit(1) have "n div 10 \ 0" unfolding is_digit_def by auto with non_digit(2) ns_eq have c_neq: "c \ CHR ''0''" by auto from \\ (is_digit n)\ c_neq ns_eq show ?case by auto qed lemma nat_to_string___is_strong_nat_string: "is_strong_nat_string (nat_to_string n)" proof (cases "is_digit n") case True thus ?thesis by simp next case False note not_digit = this obtain c s' where ns_eq: "nat_to_string n = c # s'" by (cases "nat_to_string n") auto from not_digit have "0 < n" unfolding is_digit_def by simp with ns_eq have c_neq: "c \ CHR ''0''" by auto hence "hd (nat_to_string n) \ CHR ''0''" unfolding ns_eq by simp thus ?thesis unfolding is_strong_nat_string_def by simp qed lemma nat_to_string_inv : "string_to_nat (nat_to_string n) = Some n" by (metis nat_to_string_11 nat_to_string___is_nat_string nat_to_string___is_strong_nat_string string_to_nat_def string_to_nat_inv) definition The_opt where "The_opt p = (if (\!x. p x) then Some (The p) else None)" lemma The_opt_eq_some [simp] : "((The_opt p) = (Some x)) \ ((p x) \ ((\ y. p y \ (x = y))))" (is "?lhs = ?rhs") proof (cases "\!x. p x") case True note exists_unique = this then obtain x where p_x: "p x" by auto from the1_equality[of p x] exists_unique p_x have the_opt_eq: "The_opt p = Some x" unfolding The_opt_def by simp from exists_unique the_opt_eq p_x show ?thesis by auto next case False note not_unique = this hence "The_opt p = None" unfolding The_opt_def by simp with not_unique show ?thesis by auto qed lemma The_opt_eq_none [simp] : "((The_opt p) = None) \ \(\!x. p x)" unfolding The_opt_def by auto end diff --git a/thys/CakeML/generated/Lem_sorting.thy b/thys/CakeML/generated/Lem_sorting.thy --- a/thys/CakeML/generated/Lem_sorting.thy +++ b/thys/CakeML/generated/Lem_sorting.thy @@ -1,110 +1,110 @@ chapter \Generated by Lem from \sorting.lem\.\ theory "Lem_sorting" imports Main "Lem_bool" "Lem_basic_classes" "Lem_maybe" "Lem_list" "Lem_num" "Lem" - "HOL-Library.List_Permutation" + "HOL-Combinatorics.List_Permutation" begin \ \\open import Bool Basic_classes Maybe List Num\\ \ \\open import {isabelle} `HOL-Library.Permutation`\\ \ \\open import {coq} `Coq.Lists.List`\\ \ \\open import {hol} `sortingTheory` `permLib`\\ \ \\open import {isabelle} `$LIB_DIR/Lem`\\ \ \\ ------------------------- \\ \ \\ permutations \\ \ \\ ------------------------- \\ \ \\val isPermutation : forall 'a. Eq 'a => list 'a -> list 'a -> bool\\ \ \\val isPermutationBy : forall 'a. ('a -> 'a -> bool) -> list 'a -> list 'a -> bool\\ fun isPermutationBy :: "('a \ 'a \ bool)\ 'a list \ 'a list \ bool " where " isPermutationBy eq ([]) l2 = ( (l2 = []))" |" isPermutationBy eq (x # xs) l2 = ( ( (case delete_first (eq x) l2 of None => False | Some ys => isPermutationBy eq xs ys ) ))" \ \\ ------------------------- \\ \ \\ isSorted \\ \ \\ ------------------------- \\ \ \\ isSortedBy R l checks, whether the list l is sorted by ordering R. R should represent an order, i.e. it should be transitive. Different backends defined "isSorted" slightly differently. However, the definitions coincide for transitive R. Therefore there is the following restriction: WARNING: Use isSorted and isSortedBy only with transitive relations! \\ \ \\val isSorted : forall 'a. Ord 'a => list 'a -> bool\\ \ \\val isSortedBy : forall 'a. ('a -> 'a -> bool) -> list 'a -> bool\\ \ \\ DPM: rejigged the definition with a nested match to get past Coq's termination checker. \\ \ \\let rec isSortedBy cmp l= match l with | [] -> true | x1 :: xs -> match xs with | [] -> true | x2 :: _ -> (cmp x1 x2 && isSortedBy cmp xs) end end\\ \ \\ ----------------------- \\ \ \\ insertion sort \\ \ \\ ----------------------- \\ \ \\val insert : forall 'a. Ord 'a => 'a -> list 'a -> list 'a\\ \ \\val insertBy : forall 'a. ('a -> 'a -> bool) -> 'a -> list 'a -> list 'a\\ \ \\val insertSort: forall 'a. Ord 'a => list 'a -> list 'a\\ \ \\val insertSortBy: forall 'a. ('a -> 'a -> bool) -> list 'a -> list 'a\\ \ \\let rec insertBy cmp e l= match l with | [] -> [e] | x :: xs -> if cmp x e then x :: (insertBy cmp e xs) else (e :: x :: xs) end\\ \ \\let insertSortBy cmp l= List.foldl (fun l e -> insertBy cmp e l) [] l\\ \ \\ ----------------------- \\ \ \\ general sorting \\ \ \\ ----------------------- \\ \ \\val sort: forall 'a. Ord 'a => list 'a -> list 'a\\ \ \\val sortBy: forall 'a. ('a -> 'a -> bool) -> list 'a -> list 'a\\ \ \\val sortByOrd: forall 'a. ('a -> 'a -> ordering) -> list 'a -> list 'a\\ \ \\val predicate_of_ord : forall 'a. ('a -> 'a -> ordering) -> 'a -> 'a -> bool\\ definition predicate_of_ord :: "('a \ 'a \ ordering)\ 'a \ 'a \ bool " where " predicate_of_ord f x y = ( (case f x y of LT => True | EQ => True | GT => False ))" end diff --git a/thys/Card_Partitions/Card_Partitions.thy b/thys/Card_Partitions/Card_Partitions.thy --- a/thys/Card_Partitions/Card_Partitions.thy +++ b/thys/Card_Partitions/Card_Partitions.thy @@ -1,353 +1,353 @@ (* Author: Lukas Bulwahn *) section \Cardinality of Set Partitions\ theory Card_Partitions imports - "HOL-Library.Stirling" + "HOL-Combinatorics.Stirling" Set_Partition Injectivity_Solver begin lemma set_partition_on_insert_with_fixed_card_eq: assumes "finite A" assumes "a \ A" shows "{P. partition_on (insert a A) P \ card P = Suc k} = (do { P <- {P. partition_on A P \ card P = Suc k}; p <- P; {insert (insert a p) (P - {p})} }) \ (do { P <- {P. partition_on A P \ card P = k}; {insert {a} P} })" (is "?S = ?T") proof show "?S \ ?T" proof fix P assume "P \ {P. partition_on (insert a A) P \ card P = Suc k}" from this have "partition_on (insert a A) P" and "card P = Suc k" by auto show "P \ ?T" proof cases assume "{a} \ P" have "partition_on A (P - {{a}})" using \{a} \ P\ \partition_on (insert a A) P\[THEN partition_on_Diff, of "{{a}}"] \a \ A\ by auto moreover from \{a} \ P\ \card P = Suc k\ have "card (P - {{a}}) = k" by (subst card_Diff_singleton) (auto intro: card_ge_0_finite) moreover from \{a} \ P\ have "P = insert {a} (P - {{a}})" by auto ultimately have "P \ {P. partition_on A P \ card P = k} \ (\P. {insert {a} P})" by auto from this show ?thesis by auto next assume "{a} \ P" let ?p' = "(THE X. a \ X \ X \ P)" let ?p = "(THE X. a \ X \ X \ P) - {a}" let ?P' = "insert ?p (P - {?p'})" from \partition_on (insert a A) P\ have "a \ (THE X. a \ X \ X \ P)" using partition_on_in_the_unique_part by fastforce from \partition_on (insert a A) P\ have "(THE X. a \ X \ X \ P) \ P" using partition_on_the_part_mem by fastforce from this \partition_on (insert a A) P\ have "(THE X. a \ X \ X \ P) - {a} \ P" using partition_subset_imp_notin \a \ (THE X. a \ X \ X \ P)\ by blast have "(THE X. a \ X \ X \ P) \ {a}" using \(THE X. a \ X \ X \ P) \ P\ \{a} \ P\ by auto from \partition_on (insert a A) P\ have "(THE X. a \ X \ X \ P) \ insert a A" using \(THE X. a \ X \ X \ P) \ P\ partition_onD1 by fastforce note facts_on_the_part_of = \a \ (THE X. a \ X \ X \ P)\ \(THE X. a \ X \ X \ P) \ P\ \(THE X. a \ X \ X \ P) - {a} \ P\ \(THE X. a \ X \ X \ P) \ {a}\ \(THE X. a \ X \ X \ P) \ insert a A\ from \partition_on (insert a A) P\ \finite A\ have "finite P" by (meson finite.insertI finite_elements) from \partition_on (insert a A) P\ \a \ A\ have "partition_on (A - ?p) (P - {?p'})" using facts_on_the_part_of by (auto intro: partition_on_remove_singleton) from this have "partition_on A ?P'" using facts_on_the_part_of by (auto intro: partition_on_insert simp add: disjnt_iff) moreover have "card ?P' = Suc k" proof - from \card P = Suc k\ have "card (P - {THE X. a \ X \ X \ P}) = k" using \finite P\ \(THE X. a \ X \ X \ P) \ P\ by simp from this show ?thesis using \finite P\ \(THE X. a \ X \ X \ P) - {a} \ P\ by (simp add: card_insert_if) qed moreover have "?p \ ?P'" by auto moreover have "P = insert (insert a ?p) (?P' - {?p})" using facts_on_the_part_of by (auto simp add: insert_absorb) ultimately have "P \ {P. partition_on A P \ card P = Suc k} \ (\P. P \ (\p. {insert (insert a p) (P - {p})}))" by auto then show ?thesis by auto qed qed next show "?T \ ?S" proof fix P assume "P \ ?T" (is "_ \ ?subexpr1 \ ?subexpr2") from this show "P \ ?S" proof assume "P \ ?subexpr1" from this obtain p P' where "P = insert (insert a p) (P' - {p})" and "partition_on A P'" and "card P' = Suc k" and "p \ P'" by auto from \p \ P'\ \partition_on A P'\ have "partition_on (A - p) (P' - {p})" by (simp add: partition_on_remove_singleton) from \partition_on A P'\ \finite A\ have "finite P" using \P = _\ finite_elements by auto from \partition_on A P'\ \a \ A\ have "insert a p \ P' - {p}" using partition_onD1 by fastforce from \P = _\ this \card P' = Suc k\ \finite P\ \p \ P'\ have "card P = Suc k" by auto moreover have "partition_on (insert a A) P" using \partition_on (A - p) (P' - {p})\ \a \ A\ \p \ P'\ \partition_on A P'\ \P = _\ by (auto intro!: partition_on_insert dest: partition_onD1 simp add: disjnt_iff) ultimately show "P \ ?S" by auto next assume "P \ ?subexpr2" from this obtain P' where "P = insert {a} P'" and "partition_on A P'" and "card P' = k" by auto from \partition_on A P'\ \finite A\ have "finite P" using \P = insert {a} P'\ finite_elements by auto from \partition_on A P'\ \a \ A\ have "{a} \ P'" using partition_onD1 by fastforce from \P = insert {a} P'\ \card P' = k\ this \finite P\ have "card P = Suc k" by auto moreover from \partition_on A P'\ \a \ A\ have "partition_on (insert a A) P" using \P = insert {a} P'\ by (simp add: partition_on_insert_singleton) ultimately show "P \ ?S" by auto qed qed qed lemma injectivity_subexpr1: assumes "a \ A" assumes "X \ P \ X' \ P'" assumes "insert (insert a X) (P - {X}) = insert (insert a X') (P' - {X'})" assumes "(partition_on A P \ card P = Suc k') \ (partition_on A P' \ card P' = Suc k')" shows "P = P'" and "X = X'" proof - from assms(1, 2, 4) have "a \ X" "a \ X'" using partition_onD1 by auto from assms(1, 4) have "insert a X \ P" "insert a X' \ P'" using partition_onD1 by auto from assms(1, 3, 4) have "insert a X = insert a X'" by (metis Diff_iff insertE insertI1 mem_simps(9) partition_onD1) from this \a \ X'\ \a \ X\ show "X = X'" by (meson insert_ident) from assms(2, 3) show "P = P'" using \insert a X = insert a X'\ \insert a X \ P\ \insert a X' \ P'\ by (metis insert_Diff insert_absorb insert_commute insert_ident) qed lemma injectivity_subexpr2: assumes "a \ A" assumes "insert {a} P = insert {a} P'" assumes "(partition_on A P \ card P = k') \ partition_on A P' \ card P' = k'" shows "P = P'" proof - from assms(1, 3) have "{a} \ P" and "{a} \ P'" using partition_onD1 by auto from \{a} \ P\ have "P = insert {a} P - {{a}}" by simp also from \insert {a} P = insert {a} P'\ have "\ = insert {a} P' - {{a}}" by simp also from \{a} \ P'\ have "\ = P'" by simp finally show ?thesis . qed theorem card_partition_on: assumes "finite A" shows "card {P. partition_on A P \ card P = k} = Stirling (card A) k" using assms proof (induct A arbitrary: k) case empty have eq: "{P. P = {} \ card P = 0} = {{}}" by auto show ?case by (cases k) (auto simp add: partition_on_empty eq) next case (insert a A) from this show ?case proof (cases k) case 0 from insert(1) have empty: "{P. partition_on (insert a A) P \ card P = 0} = {}" unfolding partition_on_def by (auto simp add: card_eq_0_iff finite_UnionD) from 0 insert show ?thesis by (auto simp add: empty) next case (Suc k') let ?subexpr1 = "do { P <- {P. partition_on A P \ card P = Suc k'}; p <- P; {insert (insert a p) (P - {p})} }" let ?subexpr2 = "do { P <- {P. partition_on A P \ card P = k'}; {insert {a} P} }" let ?expr = "?subexpr1 \ ?subexpr2" have "card {P. partition_on (insert a A) P \ card P = k} = card ?expr" using \finite A\ \a \ A\ \k = Suc k'\ by (simp add: set_partition_on_insert_with_fixed_card_eq) also have "card ?expr = Stirling (card A) k' + Stirling (card A) (Suc k') * Suc k'" proof - have "finite ?subexpr1 \ card ?subexpr1 = Stirling (card A) (Suc k') * Suc k'" proof - from \finite A\ have "finite {P. partition_on A P \ card P = Suc k'}" by (simp add: finitely_many_partition_on) moreover have "\X\{P. partition_on A P \ card P = Suc k'}. finite (X \ (\p. {insert (insert a p) (X - {p})}))" using finite_elements \finite A\ finite_bind by (metis (no_types, lifting) finite.emptyI finite_insert mem_Collect_eq) moreover have "disjoint_family_on (\P. P \ (\p. {insert (insert a p) (P - {p})})) {P. partition_on A P \ card P = Suc k'}" by (injectivity_solver rule: injectivity_subexpr1(1)[OF \a \ A\]) moreover have "card (P \ (\p. {insert (insert a p) (P - {p})})) = Suc k'" if "P \ {P. partition_on A P \ card P = Suc k'}" for P proof - from that \finite A\ have "finite P" using finite_elements by blast moreover have "inj_on (\p. insert (insert a p) (P - {p})) P" using that injectivity_subexpr1(2)[OF \a \ A\] by (simp add: inj_onI) moreover from that have "card P = Suc k'" by simp ultimately show ?thesis by (simp add: card_bind_singleton) qed ultimately have "card ?subexpr1 = card {P. partition_on A P \ card P = Suc k'} * Suc k'" by (subst card_bind_constant) simp+ from this have "card ?subexpr1 = Stirling (card A) (Suc k') * Suc k'" using insert.hyps(3) by simp moreover have "finite ?subexpr1" using \finite {P. partition_on A P \ card P = Suc k'}\ \\X\{P. partition_on A P \ card P = Suc k'}. finite (X \ (\p. {insert (insert a p) (X - {p})}))\ by (auto intro: finite_bind) ultimately show ?thesis by blast qed moreover have "finite ?subexpr2 \ card ?subexpr2 = Stirling (card A) k'" proof - from \finite A\ have "finite {P. partition_on A P \ card P = k'}" by (simp add: finitely_many_partition_on) moreover have " inj_on (insert {a}) {P. partition_on A P \ card P = k'}" using injectivity_subexpr2[OF \a \ A\] by (simp add: inj_on_def) ultimately have "card ?subexpr2 = card {P. partition_on A P \ card P = k'}" by (simp add: card_bind_singleton) also have "\ = Stirling (card A) k'" using insert.hyps(3) . finally have "card ?subexpr2 = Stirling (card A) k'" . moreover have "finite ?subexpr2" by (simp add: \finite {P. partition_on A P \ card P = k'}\ finite_bind) ultimately show ?thesis by blast qed moreover have "?subexpr1 \ ?subexpr2 = {}" proof - have "\P\?subexpr1. {a} \ P" using insert.hyps(2) by (force elim!: partition_onE) moreover have "\P\?subexpr2. {a} \ P" by auto ultimately show "?subexpr1 \ ?subexpr2 = {}" by blast qed ultimately show ?thesis by (simp add: card_Un_disjoint) qed also have "\ = Stirling (card (insert a A)) k" using insert(1, 2) \k = Suc k'\ by simp finally show ?thesis . qed qed theorem card_partition_on_at_most_size: assumes "finite A" shows "card {P. partition_on A P \ card P \ k} = (\j\k. Stirling (card A) j)" proof - have "card {P. partition_on A P \ card P \ k} = card (\j\k. {P. partition_on A P \ card P = j})" by (rule arg_cong[where f="card"]) auto also have "\ = (\j\k. card {P. partition_on A P \ card P = j})" by (subst card_UN_disjoint) (auto simp add: \finite A\ finitely_many_partition_on) also have "(\j\k. card {P. partition_on A P \ card P = j}) = (\j\k. Stirling (card A) j)" using \finite A\ by (simp add: card_partition_on) finally show ?thesis . qed theorem partition_on_size1: assumes "finite A" shows "{P. partition_on A P \ (\X\P. card X = 1)} = {(\a. {a}) ` A}" proof show "{P. partition_on A P \ (\X\P. card X = 1)} \ {(\a. {a}) ` A}" proof fix P assume P: "P \ {P. partition_on A P \ (\X\P. card X = 1)}" have "P = (\a. {a}) ` A" proof show "P \ (\a. {a}) ` A" proof fix X assume "X \ P" from P this obtain x where "X = {x}" by (auto simp add: card_Suc_eq) from this \X \ P\ have "x \ A" using P unfolding partition_on_def by blast from this \X = {x}\ show "X \(\a. {a}) ` A" by auto qed next show "(\a. {a}) ` A \ P" proof fix X assume "X \ (\a. {a}) ` A" from this obtain x where X: "X = {x}" "x \ A" by auto have "\P = A" using P unfolding partition_on_def by blast from this \x \ A\ obtain X' where "x \ X'" and "X' \ P" using UnionE by blast from \X' \ P\ have "card X' = 1" using P unfolding partition_on_def by auto from this \x \ X'\ have "X' = {x}" using card_1_singletonE by blast from this X(1) \X' \ P\ show "X \ P" by auto qed qed from this show "P \ {(\a. {a}) ` A}" by auto qed next show "{(\a. {a}) ` A} \ {P. partition_on A P \ (\X\P. card X = 1)}" proof fix P assume "P \ {(\a. {a}) ` A}" from this have P: "P = (\a. {a}) ` A" by auto from this have "partition_on A P" by (auto intro: partition_onI) from P this show "P \ {P. partition_on A P \ (\X\P. card X = 1)}" by auto qed qed theorem card_partition_on_size1: assumes "finite A" shows "card {P. partition_on A P \ (\X\P. card X = 1)} = 1" using assms partition_on_size1 by fastforce lemma card_partition_on_size1_eq_1: assumes "finite A" assumes "card A \ k" shows "card {P. partition_on A P \ card P \ k \ (\X\P. card X = 1)} = 1" proof - { fix P assume "partition_on A P" "\X\P. card X = 1" from this have "P \ {P. partition_on A P \ (\X\P. card X = 1)}" by simp from this have "P \ {(\a. {a}) ` A}" using partition_on_size1 \finite A\ by auto from this have "P = (\a. {a}) ` A" by auto moreover from this have "card P = card A" by (auto intro: card_image) } from this have "{P. partition_on A P \ card P \ k \ (\X\P. card X = 1)} = {P. partition_on A P \ (\X\P. card X = 1)}" using \card A \ k\ by auto from this show ?thesis using \finite A\ by (simp only: card_partition_on_size1) qed lemma card_partition_on_size1_eq_0: assumes "finite A" assumes "k < card A" shows "card {P. partition_on A P \ card P \ k \ (\X\P. card X = 1)} = 0" proof - { fix P assume "partition_on A P" "\X\P. card X = 1" from this have "P \ {P. partition_on A P \ (\X\P. card X = 1)}" by simp from this have "P \ {(\a. {a}) ` A}" using partition_on_size1 \finite A\ by auto from this have "P = (\a. {a}) ` A" by auto from this have "card P = card A" by (auto intro: card_image) } from this assms(2) have "{P. partition_on A P \ card P \ k \ (\X\P. card X = 1)} = {}" using Collect_empty_eq leD by fastforce from this show ?thesis by (simp only: card.empty) qed end diff --git a/thys/Card_Partitions/ROOT b/thys/Card_Partitions/ROOT --- a/thys/Card_Partitions/ROOT +++ b/thys/Card_Partitions/ROOT @@ -1,12 +1,13 @@ chapter AFP session Card_Partitions (AFP) = HOL + options [timeout = 600] sessions + "HOL-Combinatorics" "HOL-Eisbach" Discrete_Summation theories Card_Partitions document_files "root.bib" "root.tex" diff --git a/thys/Comparison_Sort_Lower_Bound/Linorder_Relations.thy b/thys/Comparison_Sort_Lower_Bound/Linorder_Relations.thy --- a/thys/Comparison_Sort_Lower_Bound/Linorder_Relations.thy +++ b/thys/Comparison_Sort_Lower_Bound/Linorder_Relations.thy @@ -1,545 +1,545 @@ (* File: Linorder_Relations.thy Author: Manuel Eberl Linear orderings represented as relations (i.e. set of pairs). Also contains material connecting such orderings to lists, and insertion sort w.r.t. a given ordering relation. *) section \Linear orderings as relations\ theory Linorder_Relations imports Complex_Main - "HOL-Library.Multiset_Permutations" + "HOL-Combinatorics.Multiset_Permutations" "List-Index.List_Index" begin subsection \Auxiliary facts\ (* TODO: Move *) lemma distinct_count_atmost_1': "distinct xs = (\a. count (mset xs) a \ 1)" proof - { fix x have "count (mset xs) x = (if x \ set xs then 1 else 0) \ count (mset xs) x \ 1" using count_eq_zero_iff[of "mset xs" x] by (cases "count (mset xs) x") (auto simp del: count_mset_0_iff) } thus ?thesis unfolding distinct_count_atmost_1 by blast qed lemma distinct_mset_mono: assumes "distinct ys" "mset xs \# mset ys" shows "distinct xs" unfolding distinct_count_atmost_1' proof fix x from assms(2) have "count (mset xs) x \ count (mset ys) x" by (rule mset_subset_eq_count) also from assms(1) have "\ \ 1" unfolding distinct_count_atmost_1' .. finally show "count (mset xs) x \ 1" . qed lemma mset_eq_imp_distinct_iff: assumes "mset xs = mset ys" shows "distinct xs \ distinct ys" using assms by (simp add: distinct_count_atmost_1') lemma total_on_subset: "total_on B R \ A \ B \ total_on A R" by (auto simp: total_on_def) subsection \Sortedness w.r.t. a relation\ inductive sorted_wrt :: "('a \ 'a) set \ 'a list \ bool" for R where "sorted_wrt R []" | "sorted_wrt R xs \ (\y. y \ set xs \ (x,y) \ R) \ sorted_wrt R (x # xs)" lemma sorted_wrt_Nil [simp]: "sorted_wrt R []" by (rule sorted_wrt.intros) lemma sorted_wrt_Cons: "sorted_wrt R (x # xs) \ (\y\set xs. (x,y) \ R) \ sorted_wrt R xs" by (auto intro: sorted_wrt.intros elim: sorted_wrt.cases) lemma sorted_wrt_singleton [simp]: "sorted_wrt R [x]" by (intro sorted_wrt.intros) simp_all lemma sorted_wrt_many: assumes "trans R" shows "sorted_wrt R (x # y # xs) \ (x,y) \ R \ sorted_wrt R (y # xs)" by (force intro: sorted_wrt.intros transD[OF assms] elim: sorted_wrt.cases) lemma sorted_wrt_imp_le_last: assumes "sorted_wrt R xs" "xs \ []" "x \ set xs" "x \ last xs" shows "(x, last xs) \ R" using assms by induction auto lemma sorted_wrt_append: assumes "sorted_wrt R xs" "sorted_wrt R ys" "\x y. x \ set xs \ y \ set ys \ (x,y) \ R" "trans R" shows "sorted_wrt R (xs @ ys)" using assms by (induction xs) (auto simp: sorted_wrt_Cons) lemma sorted_wrt_snoc: assumes "sorted_wrt R xs" "(last xs, y) \ R" "trans R" shows "sorted_wrt R (xs @ [y])" using assms(1,2) proof induction case (2 xs x) show ?case proof (cases "xs = []") case False with 2 have "(z,y) \ R" if "z \ set xs" for z using that by (cases "z = last xs") (auto intro: assms transD[OF assms(3), OF sorted_wrt_imp_le_last[OF 2(1)]]) from False have *: "last xs \ set xs" by simp moreover from 2 False have "(x,y) \ R" by (intro transD[OF assms(3) 2(2)[OF *]]) simp ultimately show ?thesis using 2 False by (auto intro!: sorted_wrt.intros) qed (insert 2, auto intro: sorted_wrt.intros) qed simp_all lemma sorted_wrt_conv_nth: "sorted_wrt R xs \ (\i j. i < j \ j < length xs \ (xs!i, xs!j) \ R)" by (induction xs) (auto simp: sorted_wrt_Cons nth_Cons set_conv_nth split: nat.splits) subsection \Linear orderings\ definition linorder_on :: "'a set \ ('a \ 'a) set \ bool" where "linorder_on A R \ refl_on A R \ antisym R \ trans R \ total_on A R" lemma linorder_on_cases: assumes "linorder_on A R" "x \ A" "y \ A" shows "x = y \ ((x, y) \ R \ (y, x) \ R) \ ((y, x) \ R \ (x, y) \ R)" using assms by (auto simp: linorder_on_def refl_on_def total_on_def antisym_def) lemma sorted_wrt_linorder_imp_index_le: assumes "linorder_on A R" "set xs \ A" "sorted_wrt R xs" "x \ set xs" "y \ set xs" "(x,y) \ R" shows "index xs x \ index xs y" proof - define i j where "i = index xs x" and "j = index xs y" { assume "j < i" moreover from assms have "i < length xs" by (simp add: i_def) ultimately have "(xs!j,xs!i) \ R" using assms by (auto simp: sorted_wrt_conv_nth) with assms have "x = y" by (auto simp: linorder_on_def antisym_def i_def j_def) } hence "i \ j \ x = y" by linarith thus ?thesis by (auto simp: i_def j_def) qed lemma sorted_wrt_linorder_index_le_imp: assumes "linorder_on A R" "set xs \ A" "sorted_wrt R xs" "x \ set xs" "y \ set xs" "index xs x \ index xs y" shows "(x,y) \ R" proof (cases "x = y") case False define i j where "i = index xs x" and "j = index xs y" from False and assms have "i \ j" by (simp add: i_def j_def) with \index xs x \ index xs y\ have "i < j" by (simp add: i_def j_def) moreover from assms have "j < length xs" by (simp add: j_def) ultimately have "(xs ! i, xs ! j) \ R" using assms(3) by (auto simp: sorted_wrt_conv_nth) with assms show ?thesis by (simp_all add: i_def j_def) qed (insert assms, auto simp: linorder_on_def refl_on_def) lemma sorted_wrt_linorder_index_le_iff: assumes "linorder_on A R" "set xs \ A" "sorted_wrt R xs" "x \ set xs" "y \ set xs" shows "index xs x \ index xs y \ (x,y) \ R" using sorted_wrt_linorder_index_le_imp[OF assms] sorted_wrt_linorder_imp_index_le[OF assms] by blast lemma sorted_wrt_linorder_index_less_iff: assumes "linorder_on A R" "set xs \ A" "sorted_wrt R xs" "x \ set xs" "y \ set xs" shows "index xs x < index xs y \ (y,x) \ R" by (subst sorted_wrt_linorder_index_le_iff[OF assms(1-3) assms(5,4), symmetric]) auto lemma sorted_wrt_distinct_linorder_nth: assumes "linorder_on A R" "set xs \ A" "sorted_wrt R xs" "distinct xs" "i < length xs" "j < length xs" shows "(xs!i, xs!j) \ R \ i \ j" proof (cases i j rule: linorder_cases) case less with assms show ?thesis by (simp add: sorted_wrt_conv_nth) next case equal from assms have "xs ! i \ set xs" "xs ! j \ set xs" by (auto simp: set_conv_nth) with assms(2) have "xs ! i \ A" "xs ! j \ A" by blast+ with \linorder_on A R\ and equal show ?thesis by (simp add: linorder_on_def refl_on_def) next case greater with assms have "(xs!j, xs!i) \ R" by (auto simp add: sorted_wrt_conv_nth) moreover from assms and greater have "xs ! i \ xs ! j" by (simp add: nth_eq_iff_index_eq) ultimately show ?thesis using \linorder_on A R\ greater by (auto simp: linorder_on_def antisym_def) qed subsection \Converting a list into a linear ordering\ definition linorder_of_list :: "'a list \ ('a \ 'a) set" where "linorder_of_list xs = {(a,b). a \ set xs \ b \ set xs \ index xs a \ index xs b}" lemma linorder_linorder_of_list [intro, simp]: assumes "distinct xs" shows "linorder_on (set xs) (linorder_of_list xs)" unfolding linorder_on_def using assms by (auto simp: refl_on_def antisym_def trans_def total_on_def linorder_of_list_def) lemma sorted_wrt_linorder_of_list [intro, simp]: "distinct xs \ sorted_wrt (linorder_of_list xs) xs" by (auto simp: sorted_wrt_conv_nth linorder_of_list_def index_nth_id) subsection \Insertion sort\ primrec insert_wrt :: "('a \ 'a) set \ 'a \ 'a list \ 'a list" where "insert_wrt R x [] = [x]" | "insert_wrt R x (y # ys) = (if (x, y) \ R then x # y # ys else y # insert_wrt R x ys)" lemma set_insert_wrt [simp]: "set (insert_wrt R x xs) = insert x (set xs)" by (induction xs) auto lemma mset_insert_wrt [simp]: "mset (insert_wrt R x xs) = add_mset x (mset xs)" by (induction xs) auto lemma length_insert_wrt [simp]: "length (insert_wrt R x xs) = Suc (length xs)" by (induction xs) simp_all definition insort_wrt :: "('a \ 'a) set \ 'a list \ 'a list" where "insort_wrt R xs = foldr (insert_wrt R) xs []" lemma set_insort_wrt [simp]: "set (insort_wrt R xs) = set xs" by (induction xs) (simp_all add: insort_wrt_def) lemma mset_insort_wrt [simp]: "mset (insort_wrt R xs) = mset xs" by (induction xs) (simp_all add: insort_wrt_def) lemma length_insort_wrt [simp]: "length (insort_wrt R xs) = length xs" by (induction xs) (simp_all add: insort_wrt_def) lemma sorted_wrt_insert_wrt [intro]: "linorder_on A R \ set (x # xs) \ A \ sorted_wrt R xs \ sorted_wrt R (insert_wrt R x xs)" proof (induction xs) case (Cons y ys) from Cons.prems have "(x,y) \ R \ (y,x) \ R" by (cases "x = y") (auto simp: linorder_on_def refl_on_def total_on_def) with Cons show ?case by (auto simp: sorted_wrt_Cons intro: transD simp: linorder_on_def) qed auto lemma sorted_wrt_insort [intro]: assumes "linorder_on A R" "set xs \ A" shows "sorted_wrt R (insort_wrt R xs)" proof - from assms have "set (insort_wrt R xs) = set xs \ sorted_wrt R (insort_wrt R xs)" by (induction xs) (auto simp: insort_wrt_def intro!: sorted_wrt_insert_wrt) thus ?thesis .. qed lemma distinct_insort_wrt [simp]: "distinct (insort_wrt R xs) \ distinct xs" by (simp add: distinct_count_atmost_1) lemma sorted_wrt_linorder_unique: assumes "linorder_on A R" "mset xs = mset ys" "sorted_wrt R xs" "sorted_wrt R ys" shows "xs = ys" proof - from \mset xs = mset ys\ have "length xs = length ys" by (rule mset_eq_length) from this and assms(2-) show ?thesis proof (induction xs ys rule: list_induct2) case (Cons x xs y ys) have "set (x # xs) = set_mset (mset (x # xs))" by simp also have "mset (x # xs) = mset (y # ys)" by fact also have "set_mset \ = set (y # ys)" by simp finally have eq: "set (x # xs) = set (y # ys)" . have "x = y" proof (rule ccontr) assume "x \ y" with eq have "x \ set ys" "y \ set xs" by auto with Cons.prems and assms(1) and eq have "(x, y) \ R" "(y, x) \ R" by (auto simp: sorted_wrt_Cons) with assms(1) have "x = y" by (auto simp: linorder_on_def antisym_def) with \x \ y\ show False by contradiction qed with Cons show ?case by (auto simp: sorted_wrt_Cons) qed auto qed subsection \Obtaining a sorted list of a given set\ definition sorted_wrt_list_of_set where "sorted_wrt_list_of_set R A = (if finite A then (THE xs. set xs = A \ distinct xs \ sorted_wrt R xs) else [])" lemma mset_remdups: "mset (remdups xs) = mset_set (set xs)" proof (induction xs) case (Cons x xs) thus ?case by (cases "x \ set xs") (auto simp: insert_absorb) qed auto lemma sorted_wrt_list_set: assumes "linorder_on A R" "set xs \ A" shows "sorted_wrt_list_of_set R (set xs) = insort_wrt R (remdups xs)" proof - have "sorted_wrt_list_of_set R (set xs) = (THE xsa. set xsa = set xs \ distinct xsa \ sorted_wrt R xsa)" by (simp add: sorted_wrt_list_of_set_def) also have "\ = insort_wrt R (remdups xs)" proof (rule the_equality) fix xsa assume xsa: "set xsa = set xs \ distinct xsa \ sorted_wrt R xsa" from xsa have "mset xsa = mset_set (set xsa)" by (subst mset_set_set) simp_all also from xsa have "set xsa = set xs" by simp also have "mset_set \ = mset (remdups xs)" by (simp add: mset_remdups) finally show "xsa = insort_wrt R (remdups xs)" using xsa assms by (intro sorted_wrt_linorder_unique[OF assms(1)]) (auto intro!: sorted_wrt_insort) qed (insert assms, auto intro!: sorted_wrt_insort) finally show ?thesis . qed lemma linorder_sorted_wrt_exists: assumes "linorder_on A R" "finite B" "B \ A" shows "\xs. set xs = B \ distinct xs \ sorted_wrt R xs" proof - from \finite B\ obtain xs where "set xs = B" "distinct xs" using finite_distinct_list by blast hence "set (insort_wrt R xs) = B" "distinct (insort_wrt R xs)" by simp_all moreover have "sorted_wrt R (insort_wrt R xs)" using assms \set xs = B\ by (intro sorted_wrt_insort[OF assms(1)]) auto ultimately show ?thesis by blast qed lemma linorder_sorted_wrt_list_of_set: assumes "linorder_on A R" "finite B" "B \ A" shows "set (sorted_wrt_list_of_set R B) = B" "distinct (sorted_wrt_list_of_set R B)" "sorted_wrt R (sorted_wrt_list_of_set R B)" proof - have "\!xs. set xs = B \ distinct xs \ sorted_wrt R xs" proof (rule ex_ex1I) show "\xs. set xs = B \ distinct xs \ sorted_wrt R xs" by (rule linorder_sorted_wrt_exists assms)+ next fix xs ys assume "set xs = B \ distinct xs \ sorted_wrt R xs" "set ys = B \ distinct ys \ sorted_wrt R ys" thus "xs = ys" by (intro sorted_wrt_linorder_unique[OF assms(1)]) (auto simp: set_eq_iff_mset_eq_distinct) qed from theI'[OF this] show "set (sorted_wrt_list_of_set R B) = B" "distinct (sorted_wrt_list_of_set R B)" "sorted_wrt R (sorted_wrt_list_of_set R B)" by (simp_all add: sorted_wrt_list_of_set_def \finite B\) qed lemma sorted_wrt_list_of_set_eqI: assumes "linorder_on B R" "A \ B" "set xs = A" "distinct xs" "sorted_wrt R xs" shows "sorted_wrt_list_of_set R A = xs" proof (rule sorted_wrt_linorder_unique) show "linorder_on B R" by fact let ?ys = "sorted_wrt_list_of_set R A" have fin [simp]: "finite A" by (simp_all add: assms(3) [symmetric]) have *: "distinct ?ys" "set ?ys = A" "sorted_wrt R ?ys" by (rule linorder_sorted_wrt_list_of_set[OF assms(1)] fin assms)+ from assms * show "mset ?ys = mset xs" by (subst set_eq_iff_mset_eq_distinct [symmetric]) simp_all show "sorted_wrt R ?ys" by fact qed fact+ subsection \Rank of an element in an ordering\ text \ The `rank' of an element in a set w.r.t. an ordering is how many smaller elements exist. This is particularly useful in linear orders, where there exists a unique $n$-th element for every $n$. \ definition linorder_rank where "linorder_rank R A x = card {y\A-{x}. (y,x) \ R}" lemma linorder_rank_le: assumes "finite A" shows "linorder_rank R A x \ card A" unfolding linorder_rank_def using assms by (rule card_mono) auto lemma linorder_rank_less: assumes "finite A" "x \ A" shows "linorder_rank R A x < card A" proof - have "linorder_rank R A x \ card (A - {x})" unfolding linorder_rank_def using assms by (intro card_mono) auto also from assms have "\ < card A" by (intro psubset_card_mono) auto finally show ?thesis . qed lemma linorder_rank_union: assumes "finite A" "finite B" "A \ B = {}" shows "linorder_rank R (A \ B) x = linorder_rank R A x + linorder_rank R B x" proof - have "linorder_rank R (A \ B) x = card {y\(A\B)-{x}. (y,x) \ R}" by (simp add: linorder_rank_def) also have "{y\(A\B)-{x}. (y,x) \ R} = {y\A-{x}. (y,x) \ R} \ {y\B-{x}. (y,x) \ R}" by blast also have "card \ = linorder_rank R A x + linorder_rank R B x" unfolding linorder_rank_def using assms by (intro card_Un_disjoint) auto finally show ?thesis . qed lemma linorder_rank_empty [simp]: "linorder_rank R {} x = 0" by (simp add: linorder_rank_def) lemma linorder_rank_singleton: "linorder_rank R {y} x = (if x \ y \ (y,x) \ R then 1 else 0)" proof - have "linorder_rank R {y} x = card {z\{y}-{x}. (z,x) \ R}" by (simp add: linorder_rank_def) also have "{z\{y}-{x}. (z,x) \ R} = (if x \ y \ (y,x) \ R then {y} else {})" by auto also have "card \ = (if x \ y \ (y,x) \ R then 1 else 0)" by simp finally show ?thesis . qed lemma linorder_rank_insert: assumes "finite A" "y \ A" shows "linorder_rank R (insert y A) x = (if x \ y \ (y,x) \ R then 1 else 0) + linorder_rank R A x" using linorder_rank_union[of "{y}" A R x] assms by (auto simp: linorder_rank_singleton) lemma linorder_rank_mono: assumes "linorder_on B R" "finite A" "A \ B" "(x, y) \ R" shows "linorder_rank R A x \ linorder_rank R A y" unfolding linorder_rank_def proof (rule card_mono) from assms have trans: "trans R" and antisym: "antisym R" by (simp_all add: linorder_on_def) from assms antisym show "{y \ A - {x}. (y, x) \ R} \ {ya \ A - {y}. (ya, y) \ R}" by (auto intro: transD[OF trans] simp: antisym_def) qed (insert assms, simp_all) lemma linorder_rank_strict_mono: assumes "linorder_on B R" "finite A" "A \ B" "y \ A" "(y, x) \ R" "x \ y" shows "linorder_rank R A y < linorder_rank R A x" proof - from assms(1) have trans: "trans R" by (simp add: linorder_on_def) from assms have *: "(x, y) \ R" by (auto simp: linorder_on_def antisym_def) from this and \(y,x) \ R\ have "{z\A-{y}. (z, y) \ R} \ {z\A-{x}. (z,x) \ R}" by (auto intro: transD[OF trans]) moreover from * and assms have "y \ {z\A-{y}. (z, y) \ R}" "y \ {z\A-{x}. (z, x) \ R}" by auto ultimately have "{z\A-{y}. (z, y) \ R} \ {z\A-{x}. (z,x) \ R}" by blast thus ?thesis using assms unfolding linorder_rank_def by (intro psubset_card_mono) auto qed lemma linorder_rank_le_iff: assumes "linorder_on B R" "finite A" "A \ B" "x \ A" "y \ A" shows "linorder_rank R A x \ linorder_rank R A y \ (x, y) \ R" proof (cases "x = y") case True with assms show ?thesis by (auto simp: linorder_on_def refl_on_def) next case False from assms(1) have trans: "trans R" by (simp_all add: linorder_on_def) from assms have "x \ B" "y \ B" by auto with \linorder_on B R\ and False have "((x,y) \ R \ (y,x) \ R) \ ((y,x) \ R \ (x,y) \ R)" by (fastforce simp: linorder_on_def antisym_def total_on_def) thus ?thesis proof assume "(x,y) \ R \ (y,x) \ R" with assms show ?thesis by (auto intro!: linorder_rank_mono) next assume *: "(y,x) \ R \ (x,y) \ R" with linorder_rank_strict_mono[OF assms(1-3), of y x] assms False show ?thesis by auto qed qed lemma linorder_rank_eq_iff: assumes "linorder_on B R" "finite A" "A \ B" "x \ A" "y \ A" shows "linorder_rank R A x = linorder_rank R A y \ x = y" proof assume "linorder_rank R A x = linorder_rank R A y" with linorder_rank_le_iff[OF assms(1-5)] linorder_rank_le_iff[OF assms(1-3) assms(5,4)] have "(x, y) \ R" "(y, x) \ R" by simp_all with assms show "x = y" by (auto simp: linorder_on_def antisym_def) qed simp_all lemma linorder_rank_set_sorted_wrt: assumes "linorder_on B R" "set xs \ B" "sorted_wrt R xs" "x \ set xs" "distinct xs" shows "linorder_rank R (set xs) x = index xs x" proof - define j where "j = index xs x" from assms have j: "j < length xs" by (simp add: j_def) have *: "x = y \ ((x, y) \ R \ (y, x) \ R) \ ((y, x) \ R \ (x, y) \ R)" if "y \ set xs" for y using linorder_on_cases[OF assms(1), of x y] assms that by auto from assms have "{y\set xs-{x}. (y, x) \ R} = {y\set xs-{x}. index xs y < index xs x}" by (auto simp: sorted_wrt_linorder_index_less_iff[OF assms(1-3)] dest: *) also have "\ = {y\set xs. index xs y < j}" by (auto simp: j_def) also have "\ = (\i. xs ! i) ` {i. i < j}" proof safe fix y assume "y \ set xs" "index xs y < j" moreover from this and j have "y = xs ! index xs y" by simp ultimately show "y \ (!) xs ` {i. i < j}" by blast qed (insert assms j, auto simp: index_nth_id) also from assms and j have "card \ = card {i. i < j}" by (intro card_image) (auto simp: inj_on_def nth_eq_iff_index_eq) also have "\ = j" by simp finally show ?thesis by (simp only: j_def linorder_rank_def) qed lemma bij_betw_linorder_rank: assumes "linorder_on B R" "finite A" "A \ B" shows "bij_betw (linorder_rank R A) A {..distinct xs\ have len_xs: "length xs = card A" by (subst \set xs = A\ [symmetric]) (auto simp: distinct_card) have rank: "linorder_rank R (set xs) x = index xs x" if "x \ A" for x using linorder_rank_set_sorted_wrt[OF assms(1), of xs x] assms that xs by simp_all from xs len_xs show ?thesis by (intro bij_betw_byWitness[where f' = "\i. xs ! i"]) (auto simp: rank index_nth_id intro!: nth_mem) qed subsection \The bijection between linear orderings and lists\ theorem bij_betw_linorder_of_list: assumes "finite A" shows "bij_betw linorder_of_list (permutations_of_set A) {R. linorder_on A R}" proof (intro bij_betw_byWitness[where f' = "\R. sorted_wrt_list_of_set R A"] ballI subsetI, goal_cases) case (1 xs) thus ?case by (intro sorted_wrt_list_of_set_eqI) (auto simp: permutations_of_set_def) next case (2 R) hence R: "linorder_on A R" by simp from R have in_R: "x \ A" "y \ A" if "(x,y) \ R" for x y using that by (auto simp: linorder_on_def refl_on_def) let ?xs = "sorted_wrt_list_of_set R A" have xs: "distinct ?xs" "set ?xs = A" "sorted_wrt R ?xs" by (rule linorder_sorted_wrt_list_of_set[OF R] assms order.refl)+ thus ?case using sorted_wrt_linorder_index_le_iff[OF R, of ?xs] by (auto simp: linorder_of_list_def dest: in_R) next case (4 xs) then obtain R where R: "linorder_on A R" and xs [simp]: "xs = sorted_wrt_list_of_set R A" by auto let ?xs = "sorted_wrt_list_of_set R A" have xs: "distinct ?xs" "set ?xs = A" "sorted_wrt R ?xs" by (rule linorder_sorted_wrt_list_of_set[OF R] assms order.refl)+ thus ?case by auto qed (auto simp: permutations_of_set_def) corollary card_finite_linorders: assumes "finite A" shows "card {R. linorder_on A R} = fact (card A)" proof - have "card {R. linorder_on A R} = card (permutations_of_set A)" by (rule sym, rule bij_betw_same_card [OF bij_betw_linorder_of_list[OF assms]]) also from assms have "\ = fact (card A)" by (rule card_permutations_of_set) finally show ?thesis . qed end diff --git a/thys/Comparison_Sort_Lower_Bound/ROOT b/thys/Comparison_Sort_Lower_Bound/ROOT --- a/thys/Comparison_Sort_Lower_Bound/ROOT +++ b/thys/Comparison_Sort_Lower_Bound/ROOT @@ -1,13 +1,14 @@ chapter AFP session Comparison_Sort_Lower_Bound (AFP) = Stirling_Formula + options [timeout = 300] sessions + "HOL-Combinatorics" "List-Index" Landau_Symbols theories Linorder_Relations Comparison_Sort_Lower_Bound document_files "root.tex" "root.bib" diff --git a/thys/Derangements/Derangements.thy b/thys/Derangements/Derangements.thy --- a/thys/Derangements/Derangements.thy +++ b/thys/Derangements/Derangements.thy @@ -1,518 +1,518 @@ (* Author: Lukas Bulwahn *) section \Derangements\ theory Derangements imports Complex_Main - "HOL-Library.Permutations" + "HOL-Combinatorics.Permutations" begin subsection \Preliminaries\ subsubsection \Additions to @{theory HOL.Finite_Set} Theory\ lemma card_product_dependent: assumes "finite S" "\x \ S. finite (T x)" shows "card {(x, y). x \ S \ y \ T x} = (\x \ S. card (T x))" using card_SigmaI[OF assms, symmetric] by (auto intro!: arg_cong[where f=card] simp add: Sigma_def) -subsubsection \Additions to @{theory "HOL-Library.Permutations"} Theory\ +subsubsection \Additions to @{theory "HOL-Combinatorics.Permutations"} Theory\ lemma permutes_imp_bij': assumes "p permutes S" shows "bij p" using assms by (meson bij_def permutes_inj permutes_surj) lemma permutesE: assumes "p permutes S" obtains "bij p" "\x. x \ S \ p x = x" using assms by (simp add: permutes_def permutes_imp_bij') lemma bij_imp_permutes': assumes "bij p" "\x. x \ A \ p x = x" shows "p permutes A" using assms bij_imp_permutes permutes_superset by force lemma permutes_swap: assumes "p permutes S" shows "Fun.swap x y p permutes (insert x (insert y S))" proof - from assms have "p permutes (insert x (insert y S))" by (meson permutes_subset subset_insertI) moreover have "Fun.swap x y id permutes (insert x (insert y S))" by (simp add: permutes_swap_id) ultimately show "Fun.swap x y p permutes (insert x (insert y S))" by (metis comp_id comp_swap permutes_compose) qed lemma bij_extends: "bij p \ p x = x \ bij (p(x := y, inv p y := x))" unfolding bij_def proof (rule conjI; erule conjE) assume a: "inj p" "p x = x" show "inj (p(x := y, inv p y := x))" proof (intro injI) fix z z' assume "(p(x := y, inv p y := x)) z = (p(x := y, inv p y := x)) z'" from this a show "z = z'" by (auto split: if_split_asm simp add: inv_f_eq inj_eq) qed next assume a: "inj p" "surj p" "p x = x" { fix x' from a have "(p(x := y, inv p y := x)) (((inv p)(y := x, x := inv p y)) x') = x'" by (auto split: if_split_asm) (metis surj_f_inv_f)+ } from this show "surj (p(x := y, inv p y := x))" by (metis surjI) qed lemma permutes_add_one: assumes "p permutes S" "x \ S" "y \ S" shows "p(x := y, inv p y := x) permutes (insert x S)" proof (rule bij_imp_permutes') from assms show "bij (p(x := y, inv p y := x))" by (meson bij_extends permutes_def permutes_imp_bij') from assms show "\z. z \ insert x S \ (p(x := y, inv p y := x)) z = z" by (metis fun_upd_apply insertCI permutes_def permutes_inverses(1)) qed lemma permutations_skip_one: assumes "p permutes S" "x : S" shows "p(x := x, inv p x := p x) permutes (S - {x})" proof (rule bij_imp_permutes') from assms show "\z. z \ S - {x} \ (p(x := x, inv p x := p x)) z = z" by (auto elim: permutesE simp add: bij_inv_eq_iff) (simp add: assms(1) permutes_in_image permutes_inv) from assms have "inj (p(x := x, inv p x := p x))" by (intro injI) (auto split: if_split_asm; metis permutes_inverses(2))+ moreover have "surj (p(x := x, inv p x := p x))" using assms UNIV_I bij_betw_swap_iff permutes_inj permutes_surj surj_f_inv_f by (metis (no_types, hide_lams) Fun.swap_def bij_betw_def) ultimately show "bij (p(x := x, inv p x := p x))" by (rule bijI) qed lemma permutes_drop_cycle_size_two: assumes "p permutes S" "p (p x) = x" shows "Fun.swap x (p x) p permutes (S - {x, p x})" using assms by (auto intro!: bij_imp_permutes' elim!: permutesE) (metis swap_apply(1,3)) subsection \Fixpoint-Free Permutations\ definition derangements :: "nat set \ (nat \ nat) set" where "derangements S = {p. p permutes S \ (\x \ S. p x \ x)}" lemma derangementsI: assumes "p permutes S" "\x. x \ S \ p x \ x" shows "p \ derangements S" using assms unfolding derangements_def by auto lemma derangementsE: assumes "d : derangements S" obtains "d permutes S" "\x\S. d x \ x" using assms unfolding derangements_def by auto subsection \Properties of Derangements\ lemma derangements_inv: assumes d: "d \ derangements S" shows "inv d \ derangements S" using assms by (auto intro!: derangementsI elim!: derangementsE simp add: permutes_inv permutes_inv_eq) lemma derangements_in_image: assumes "d \ derangements A" "x \ A" shows "d x \ A" using assms by (auto elim: derangementsE simp add: permutes_in_image) lemma derangements_in_image_strong: assumes "d \ derangements A" "x \ A" shows "d x \ A - {x}" using assms by (auto elim: derangementsE simp add: permutes_in_image) lemma derangements_inverse_in_image: assumes "d \ derangements A" "x \ A" shows "inv d x \ A" using assms by (auto intro: derangements_in_image derangements_inv) lemma derangements_fixpoint: assumes "d \ derangements A" "x \ A" shows "d x = x" using assms by (auto elim!: derangementsE simp add: permutes_def) lemma derangements_no_fixpoint: assumes "d \ derangements A" "x \ A" shows "d x \ x" using assms by (auto elim: derangementsE) lemma finite_derangements: assumes "finite A" shows "finite (derangements A)" using assms unfolding derangements_def by (auto simp add: finite_permutations) subsection \Construction of Derangements\ lemma derangements_empty[simp]: "derangements {} = {id}" unfolding derangements_def by auto lemma derangements_singleton[simp]: "derangements {x} = {}" unfolding derangements_def by auto lemma derangements_swap: assumes "d \ derangements S" "x \ S" "y \ S" "x \ y" shows "Fun.swap x y d \ derangements (insert x (insert y S))" proof (rule derangementsI) from assms show "Fun.swap x y d permutes (insert x (insert y S))" by (auto intro: permutes_swap elim: derangementsE) from assms have s: "d x = x" "d y = y" by (auto intro: derangements_fixpoint) { fix x' assume "x' : insert x (insert y S)" from s assms \x \ y\ this show "Fun.swap x y d x' \ x'" by (cases "x' = x"; cases "x' = y") (auto dest: derangements_no_fixpoint) } qed lemma derangements_skip_one: assumes d: "d \ derangements S" and "x \ S" "d (d x) \ x" shows "d(x := x, inv d x := d x) \ derangements (S - {x})" proof - from d have bij: "bij d" by (auto elim: derangementsE simp add: permutes_imp_bij') from d \x : S\ have that: "d x : S - {x}" by (auto dest: derangements_in_image derangements_no_fixpoint) from d \d (d x) \ x\ bij have "\x'\S - {x}. (d(x := x, inv d x := d x)) x' \ x'" by (auto elim!: derangementsE simp add: bij_inv_eq_iff) from d \x : S\ this show derangements: "d(x:=x, inv d x:= d x) : derangements (S - {x})" by (meson derangementsE derangementsI permutations_skip_one) qed lemma derangements_add_one: assumes "d \ derangements S" "x \ S" "y \ S" shows "d(x := y, inv d y := x) \ derangements (insert x S)" proof (rule derangementsI) from assms show "d(x := y, inv d y := x) permutes (insert x S)" by (auto intro: permutes_add_one elim: derangementsE) next fix z assume "z : insert x S" from assms this derangements_inverse_in_image[OF assms(1), of y] show "(d(x := y, inv d y := x)) z \ z" by (auto elim: derangementsE) qed lemma derangements_drop_minimal_cycle: assumes "d \ derangements S" "d (d x) = x" shows "Fun.swap x (d x) d \ derangements (S - {x, d x})" proof (rule derangementsI) from assms show "Fun.swap x (d x) d permutes (S - {x, d x})" by (meson derangementsE permutes_drop_cycle_size_two) next fix y assume "y \ S - {x, d x}" from assms this show "Fun.swap x (d x) d y \ y" by (auto elim: derangementsE) qed subsection \Cardinality of Derangements\ subsubsection \Recursive Characterization\ fun count_derangements :: "nat \ nat" where "count_derangements 0 = 1" | "count_derangements (Suc 0) = 0" | "count_derangements (Suc (Suc n)) = (n + 1) * (count_derangements (Suc n) + count_derangements n)" lemma card_derangements: assumes "finite S" "card S = n" shows "card (derangements S) = count_derangements n" using assms proof (induct n arbitrary: S rule: count_derangements.induct) case 1 from this show ?case by auto next case 2 from this derangements_singleton finite_derangements show ?case using Finite_Set.card_0_eq card_eq_SucD count_derangements.simps(2) by fastforce next case (3 n) from 3(4) obtain x where "x \ S" using card_eq_SucD insertI1 by auto let ?D1 = "(\(y, d). Fun.swap x y d) ` {(y, d). y \ S & y \ x & d : derangements (S - {x, y})}" let ?D2 = "(\(y, f). f(x:=y, inv f y := x)) ` ((S - {x}) \ derangements (S - {x}))" from \x \ S\ have subset1: "?D1 \ derangements S" proof (auto) fix y d assume "y \ S" "y \ x" assume d: "d \ derangements (S - {x, y})" from \x : S\ \y : S\ have S: "S = insert x (insert y (S - {x, y}))" by auto from d \x : S\ \y : S\ \y \ x\ show "Fun.swap x y d \ derangements S" by (subst S) (auto intro!: derangements_swap) qed have subset2: "?D2 \ derangements S" proof (rule subsetI, erule imageE, simp split: prod.split_asm, (erule conjE)+) fix d y assume "d : derangements (S - {x})" "y : S" "y \ x" from this have "d(x := y, inv d y := x) \ derangements (insert x (S - {x}))" by (intro derangements_add_one) auto from this \x : S\ show "d(x := y, inv d y := x) \ derangements S" using insert_Diff by fastforce qed have split: "derangements S = ?D1 \ ?D2" proof from subset1 subset2 show "?D1 \ ?D2 \ derangements S" by simp next show "derangements S \ ?D1 \ ?D2" proof fix d assume d: "d : derangements S" show "d : ?D1 \ ?D2" proof (cases "d (d x) = x") case True from \x : S\ d have "d x \ S" "d x \ x" by (auto simp add: derangements_in_image derangements_no_fixpoint) from d True have "Fun.swap x (d x) d \ derangements (S - {x, d x})" by (rule derangements_drop_minimal_cycle) from \d x \ S\ \d x \ x\ this have "d : ?D1" by (auto intro: image_eqI[where x = "(d x, Fun.swap x (d x) d)"]) from this show ?thesis by auto next case False from d have bij: "bij d" by (auto elim: derangementsE simp add: permutes_imp_bij') from d \x : S\ have that: "d x : S - {x}" by (intro derangements_in_image_strong) from d \x : S\ False have derangements: "d(x:=x, inv d x:= d x) : derangements (S - {x})" by (auto intro: derangements_skip_one) from this have "bij (d(x := x, inv d x:= d x))" by (metis derangementsE permutes_imp_bij')+ from this have a: "inv (d(x := x, inv d x := d x)) (d x) = inv d x" by (metis bij_inv_eq_iff fun_upd_same) from bij have x: "d (inv d x) = x" by (meson bij_inv_eq_iff) from d derangements_inv[of d] \x : S\ have "inv d x \ x" "d x \ x" by (auto dest: derangements_no_fixpoint) from this a x have d_eq: "d = d(inv d x := d x, x := d x, inv (d(x := x, inv d x := d x)) (d x) := x)" by auto from derangements that have "(d x, d(x:=x, inv d x:=d x)) : ((S - {x}) \ derangements (S - {x}))" by auto from d_eq this have "d : ?D2" by (auto intro: image_eqI[where x = "(d x, d(x:=x, inv d x:=d x))"]) from this show ?thesis by auto qed qed qed have no_intersect: "?D1 \ ?D2 = {}" proof - have that: "\d. d \ ?D1 \ d (d x) = x" using Diff_iff Diff_insert2 derangements_fixpoint insertI1 swap_apply(2) by fastforce have "\d. d \ ?D2 \ d (d x) \ x" proof - fix d assume a: "d \ ?D2" from a obtain y d' where d: "d = d'(x := y, inv d' y := x)" "d' \ derangements (S - {x})" "y \ S - {x}" by auto from d(2) have inv: "inv d' \ derangements (S - {x})" by (rule derangements_inv) from d have inv_x: "inv d' y \ x" by (auto dest: derangements_inverse_in_image) from inv have inv_y: "inv d' y \ y" using d(3) derangements_no_fixpoint by blast from d inv_x have 1: "d x = y" by auto from d inv_y have 2: "d y = d' y" by auto from d(2, 3) have 3: "d' y \ S - {x}" by (auto dest: derangements_in_image) from 1 2 3 show "d (d x) \ x" by auto qed from this that show ?thesis by blast qed have inj: "inj_on (\(y, f). Fun.swap x y f) {(y, f). y \ S & y \ x & f : derangements (S - {x, y})}" unfolding inj_on_def by (clarify; metis DiffD2 derangements_fixpoint insertI1 insert_commute swap_apply(1) swap_nilpotent) have eq: "{(y, f). y \ S & y \ x & f : derangements (S - {x, y})} = {(y, f). y \ S - {x} & f : derangements (S - {x, y})}" by simp have eq': "{(y, f). y \ S & y \ x & f : derangements (S - {x, y})} = Sigma (S - {x}) (%y. derangements (S - {x, y}))" unfolding Sigma_def by auto have card: "\ y. y \ S - {x} \ card (derangements (S - {x, y})) = count_derangements n" proof - fix y assume "y \ S - {x}" from 3(3, 4) \x \ S\ this have "card (S - {x, y}) = n" by (metis Diff_insert2 card_Diff_singleton diff_Suc_1 finite_Diff) from 3(3) 3(2)[OF _ this] show "card (derangements (S - {x, y})) = count_derangements n" by auto qed from 3(3, 4) \x : S\ have card2: "card (S - {x}) = Suc n" by (simp add: card.insert_remove insert_absorb) from inj have "card ?D1 = card {(y, f). y \ S - {x} \ f \ derangements (S - {x, y})}" by (simp add: card_image) also from 3(3) have "... = (\y\S - {x}. card (derangements (S - {x, y})))" by (subst card_product_dependent) (simp add: finite_derangements)+ finally have card_1: "card ?D1 =(Suc n) * count_derangements n" using card card2 by auto have inj_2: "inj_on (\(y, f). f(x := y, inv f y := x)) ((S - {x}) \ derangements (S - {x}))" proof - { fix d d' y y' assume 1: "d \ derangements (S - {x})" "d' \ derangements (S - {x})" assume 2: "y \ S" "y \ x" "y' \ S" "y' \ x" assume eq: "d(x := y, inv d y := x) = d'(x := y', inv d' y' := x)" from 1 2 eq \x \ S\ have y: "y = y'" by (metis Diff_insert_absorb derangements_in_image derangements_inv fun_upd_same fun_upd_twist insert_iff mk_disjoint_insert) have "d = d'" proof fix z from 1 have d_x: "d x = d' x" by (auto dest!: derangements_fixpoint) from 1 have bij: "bij d" "bij d'" by (metis derangementsE permutes_imp_bij')+ from this have d_d: "d (inv d y) = y" "d' (inv d' y') = y'" by (simp add: bij_is_surj surj_f_inv_f)+ from 1 2 bij have neq: "d (inv d' y') \ x \ d' (inv d y) \ x" by (metis Diff_iff bij_inv_eq_iff derangements_fixpoint singletonI) from eq have "(d(x := y, inv d y := x)) z = (d'(x := y', inv d' y' := x)) z" by auto from y d_x d_d neq this show "d z = d' z" by (auto split: if_split_asm) qed from y this have "y = y' & d = d'" by auto } from this show ?thesis unfolding inj_on_def by auto qed from 3(3) 3(1)[OF _ card2] have card3: "card (derangements (S - {x})) = count_derangements (Suc n)" by auto from 3(3) inj_2 have card_2: "card ?D2 = (Suc n) * count_derangements (Suc n)" by (simp only: card_image) (auto simp add: card_cartesian_product card3 card2) from \finite S\have finite1: "finite ?D1" unfolding eq' by (auto intro: finite_derangements) from \finite S\ have finite2: "finite ?D2" by (auto intro: finite_cartesian_product finite_derangements) show ?case by (simp add: split card_Un_disjoint finite1 finite2 no_intersect card_1 card_2 field_simps) qed subsubsection \Closed-Form Characterization\ lemma count_derangements: "real (count_derangements n) = fact n * (\k \ {0..n}. (-1) ^ k / fact k)" proof (induct n rule: count_derangements.induct) case (3 n) let ?f = "\n. fact n * (\k = 0..n. (- 1) ^ k / fact k)" have "real (count_derangements (Suc (Suc n))) = (n + 1) * (count_derangements (n + 1) + count_derangements n)" unfolding count_derangements.simps by simp also have "... = real (n + 1) * (real (count_derangements (n + 1)) + real (count_derangements n))" by (simp only: of_nat_mult of_nat_add) also have "... = (n + 1) * (?f (n + 1) + ?f n)" unfolding 3(2) 3(1)[unfolded Suc_eq_plus1] .. also have "(n + 1) * (?f (n + 1) + ?f n) = ?f (n + 2)" proof - define f where "f n = ((fact n) :: real) * (\k = 0..n. (- 1) ^ k / fact k)" for n have f_eq: "\n. f (n + 1) = (n + 1) * f n + (- 1) ^ (n + 1)" proof - fix n have "?f (n + 1) = (n + 1) * fact n * (\k = 0..n. (- 1) ^ k / fact k) + fact (n + 1) * ((- 1) ^ (n + 1) / fact (n + 1))" by (simp add: field_simps) also have "... = (n + 1) * ?f n + (- 1) ^ (n + 1)" by (simp del: One_nat_def) finally show "?thesis n" unfolding f_def by simp qed from this have f_eq': "\ n. (n + 1) * f n = f (n + 1) + (- 1) ^ n" by auto from f_eq'[of n] have "(n + 1) * (f (n + 1) + f n) = ((n + 1) * f (n + 1)) + f (n + 1) + (- 1) ^ n" by (simp only: distrib_left of_nat_add of_nat_1) also have "... = (n + 2) * f (n + 1) + (- 1) ^ (n + 2)" by (simp del: One_nat_def add_2_eq_Suc' add: algebra_simps) simp also from f_eq[of "n + 1"] have "... = f (n + 2)" by simp finally show ?thesis unfolding f_def by simp qed finally show ?case by simp qed (auto) subsubsection \Approximation of Cardinality\ lemma two_power_fact_le_fact: assumes "n \ 1" shows "2^k * fact n \ (fact (n + k) :: 'a :: {semiring_char_0,linordered_semidom})" proof (induction k) case (Suc k) have "2 ^ Suc k * fact n = 2 * (2 ^ k * fact n)" by (simp add: algebra_simps) also note Suc.IH also from assms have "of_nat 1 + of_nat 1 \ of_nat n + (of_nat (Suc k) :: 'a)" by (intro add_mono) (unfold of_nat_le_iff, simp_all) hence "2 * (fact (n + k) :: 'a) \ of_nat (n + Suc k) * fact (n + k)" by (intro mult_right_mono) (simp_all add: add_ac) also have "\ = fact (n + Suc k)" by simp finally show ?case by - (simp add: mult_left_mono) qed simp_all lemma exp1_approx: assumes "n > 0" shows "exp (1::real) - (\k {0..2 / fact n}" proof (unfold atLeastAtMost_iff, safe) have "(\k. 1^k / fact k) sums exp (1::real)" using exp_converges[of "1::real"] by (simp add: field_simps) from sums_split_initial_segment[OF this, of n] have sums: "(\k. 1 / fact (n+k)) sums (exp 1 - (\kk 0" by (intro sums_le[OF _ sums_zero sums]) auto show "(exp 1 - (\k 2 / fact n" proof (rule sums_le) from assms have "(\k. (1 / fact n) * (1 / 2)^k :: real) sums ((1 / fact n) * (1 / (1 - 1 / 2)))" by (intro sums_mult geometric_sums) simp_all also have "\ = 2 / fact n" by simp finally show "(\k. 1 / fact n * (1 / 2) ^ k) sums (2 / fact n :: real)" . next fix k show "(1 / fact (n+k)) \ (1 / fact n) * (1 / 2 :: real)^k" for k using two_power_fact_le_fact[of n k] assms by (auto simp: field_simps) qed fact+ qed lemma exp1_bounds: "exp 1 \ {8 / 3..11 / 4 :: real}" using exp1_approx[of 4] by (simp add: eval_nat_numeral) lemma count_derangements_approximation: assumes "n \ 0" shows "abs(real (count_derangements n) - fact n / exp 1) < 1 / 2" proof (cases "n \ 5") case False from assms this have n: "n = 1 \ n = 2 \ n = 3 \ n = 4" by auto from exp1_bounds have 1: "abs(real (count_derangements 1) - fact 1 / exp 1) < 1 / 2" by simp from exp1_bounds have 2: "abs(real (count_derangements 2) - fact 2 / exp 1) < 1 / 2" by (auto simp add: eval_nat_numeral abs_real_def) from exp1_bounds have 3: "abs(real (count_derangements 3) - fact 3 / exp 1) < 1 / 2" by (auto simp add: eval_nat_numeral abs_if field_simps) from exp1_bounds have 4: "abs(real (count_derangements 4) - fact 4 / exp 1) < 1 / 2" by (auto simp: abs_if field_simps eval_nat_numeral) from 1 2 3 4 n show ?thesis by auto next case True from Maclaurin_exp_le[of "- 1" "n + 1"] obtain t where t: "abs (t :: real) \ abs (- 1)" and exp: "exp (- 1) = (\mk = 0..n. (- 1) ^ k / fact k) = exp (- 1) - exp t / fact (n + 1) * (- 1) ^ (n + 1)" by (simp add: atLeast0AtMost ivl_disj_un(2)[symmetric]) have fact_plus1: "fact (n + 1) = (n + 1) * fact n" by simp have eq: "\real (count_derangements n) - fact n / exp 1\ = exp t / (n + 1)" by (simp del: One_nat_def add: count_derangements sum_eq_exp exp_minus inverse_eq_divide algebra_simps abs_mult) (simp add: fact_plus1) from t have "exp t \ exp 1" by simp also from exp1_bounds have "\ < 3" by simp finally show ?thesis using \n \ 5\ by (simp add: eq) qed theorem derangements_formula: assumes "n \ 0" "finite S" "card S = n" shows "int (card (derangements S)) = round (fact n / exp 1 :: real)" using count_derangements_approximation[of n] assms by (intro round_unique' [symmetric]) (auto simp: card_derangements abs_minus_commute) theorem derangements_formula': assumes "n \ 0" "finite S" "card S = n" shows "card (derangements S) = nat (round (fact n / exp 1 :: real))" using derangements_formula[OF assms] by simp end diff --git a/thys/Derangements/ROOT b/thys/Derangements/ROOT --- a/thys/Derangements/ROOT +++ b/thys/Derangements/ROOT @@ -1,11 +1,12 @@ chapter AFP session Derangements (AFP) = HOL + options [timeout = 600] sessions "HOL-Library" + "HOL-Combinatorics" theories Derangements document_files "root.tex" "root.bib" diff --git a/thys/Discrete_Summation/Factorials.thy b/thys/Discrete_Summation/Factorials.thy --- a/thys/Discrete_Summation/Factorials.thy +++ b/thys/Discrete_Summation/Factorials.thy @@ -1,322 +1,322 @@ (* Authors: Amine Chaieb & Florian Haftmann, TU Muenchen with contributions by Lukas Bulwahn *) section \Falling factorials\ theory Factorials - imports Complex_Main "HOL-Library.Stirling" + imports Complex_Main "HOL-Combinatorics.Stirling" begin lemma pochhammer_0 [simp]: \ \TODO move\ "pochhammer 0 n = (0::nat)" if "n > 0" using that by (simp add: pochhammer_prod) definition ffact :: "nat \ 'a::comm_semiring_1_cancel \ 'a" where "ffact n a = pochhammer (a + 1 - of_nat n) n" lemma ffact_0 [simp]: "ffact 0 = (\x. 1)" by (simp add: fun_eq_iff ffact_def) lemma ffact_Suc: "ffact (Suc n) a = a * ffact n (a - 1)" for a :: "'a :: comm_ring_1" by (simp add: ffact_def pochhammer_prod prod.atLeast0_lessThan_Suc algebra_simps) (* TODO: what's the right class here? *) lemma ffact_Suc_rev: "ffact (Suc n) m = (m - of_nat n) * ffact n m" for m :: "'a :: {comm_semiring_1_cancel, ab_group_add}" unfolding ffact_def pochhammer_rec by (simp add: diff_add_eq) lemma ffact_nat_triv: "ffact n m = 0" if "m < n" using that by (simp add: ffact_def) lemma ffact_Suc_nat: "ffact (Suc n) m = m * ffact n (m - 1)" for m :: nat proof (cases "n \ m") case True then show ?thesis by (simp add: ffact_def pochhammer_prod algebra_simps prod.atLeast0_lessThan_Suc) next case False then have "m < n" by simp then show ?thesis by (simp add: ffact_nat_triv) qed lemma ffact_Suc_rev_nat: "ffact (Suc n) m = (m - n) * ffact n m" proof (cases "n \ m") case True then show ?thesis by (simp add: ffact_def pochhammer_rec Suc_diff_le) next case False then have "m < n" by simp then show ?thesis by (simp add: ffact_nat_triv) qed lemma fact_div_fact_ffact: "fact n div fact m = ffact (n - m) n" if "m \ n" proof - from that have "fact n = ffact (n - m) n * fact m" by (simp add: ffact_def pochhammer_product pochhammer_fact) moreover have "fact m dvd (fact n :: nat)" using that by (rule fact_dvd) ultimately show ?thesis by simp qed lemma fact_div_fact_ffact_nat: "fact n div fact (n - k) = ffact k n" if "k \ n" using that by (simp add: fact_div_fact_ffact) lemma ffact_fact: "ffact n (of_nat n) = (of_nat (fact n) :: 'a :: comm_ring_1)" by (induct n) (simp_all add: algebra_simps ffact_Suc) lemma ffact_add_diff_assoc: "(a - of_nat n) * ffact n a + of_nat n * ffact n a = a * ffact n a" for a :: "'a :: comm_ring_1" by (simp add: algebra_simps) lemma mult_ffact: "a * ffact n a = ffact (Suc n) a + of_nat n * ffact n a" for a :: "'a :: comm_ring_1" proof - have "ffact (Suc n) a + of_nat n * (ffact n a) = (a - of_nat n) * (ffact n a) + of_nat n * (ffact n a)" using ffact_Suc_rev [of n] by auto also have "\ = a * ffact n a" using ffact_add_diff_assoc by (simp add: algebra_simps) finally show ?thesis by simp qed (* TODO: what's the right class here? *) lemma prod_ffact: fixes m :: "'a :: {ord, ring_1, comm_monoid_mult, comm_semiring_1_cancel}" shows "(\i = 0..j. j - 1) {1..n}" by (force intro: inj_on_diff_nat) moreover have "{0..j. j - 1) ` {1..n}" proof - have "i \ (\j. j - 1) ` {1..n}" if "i \ {0.. {1..n}" for i using that by (simp add: of_nat_diff) ultimately have "(\i = 0..i = 1..n. m + 1 - of_nat n + of_nat (n - i))" by (rule prod.reindex_cong) from this show ?thesis unfolding ffact_def by (simp only: pochhammer_prod_rev) qed lemma prod_ffact_nat: fixes m :: nat shows "(\i = 0.. m" have "inj_on (\j. j - 1) {1..n}" by (force intro: inj_on_diff_nat) moreover have "{0..j. j - 1) ` {1..n}" proof - have "i \ (\j. j - 1) ` {1..n}" if "i \ {0..i = 0..i = 1..n. (m + 1) - i)" by (auto intro: prod.reindex_cong[where l="\i. i - 1"]) from this \n \ m\ show ?thesis unfolding ffact_def by (simp add: pochhammer_prod_rev) next assume "\ n \ m" from this show ?thesis by (auto simp add: ffact_nat_triv) qed (* TODO: what's the right class here? *) lemma prod_rev_ffact: fixes m :: "'a :: {ord, ring_1, comm_monoid_mult, comm_semiring_1_cancel}" shows "(\i = 1..n. m - of_nat n + of_nat i) = ffact n m" proof - have "inj_on (\i. i + 1) {0..i. i + 1) ` {0..i = 1..n. m - of_nat n + of_nat i) = (\i = 0.. m" shows "(\i = 1..n. m - n + i) = ffact n m" proof - have "inj_on (\i. i + 1) {0..i. i + 1) ` {0..n \ m\ by auto ultimately have "(\i = 1..n. m - n + i) = (\i = 0.. m" shows "\{m - n + 1..m} = ffact n m" proof - have "inj_on (\i. m - n + i) {1::nat..n}" by (auto intro: inj_onI) moreover have "{m - n + 1..m} = (\i. m - n + i) ` {1::nat..n}" proof - have "i \ (\i. m + i - n) ` {Suc 0..n}" if "i \ {m - n + 1..m}" for i using that \n \ m\ by (auto intro!: image_eqI[where x="i - (m - n)"]) with \n \ m\ show ?thesis by auto qed moreover have "m - n + i = m - n + i" for i .. ultimately have "\{m - n + (1::nat)..m} = (\i = 1..n. m - n + i)" by (rule prod.reindex_cong) from this show ?thesis using \n \ m\ by (simp only: prod_rev_ffact_nat) qed lemma ffact_eq_fact_mult_binomial: "ffact k n = fact k * (n choose k)" proof cases assume "k \ n" have "ffact k n = fact n div fact (n - k)" using \k \ n\ by (simp add: fact_div_fact_ffact_nat) also have "\ = fact k * (n choose k)" using \k \ n\ by (simp add: binomial_fact_lemma[symmetric]) finally show ?thesis . next assume "\ k \ n" from this ffact_nat_triv show ?thesis by force qed lemma of_nat_ffact: "of_nat (ffact n m) = ffact n (of_nat m :: 'a :: comm_ring_1)" proof (induct n arbitrary: m) case 0 then show ?case by simp next case (Suc n) show ?case proof (cases m) case 0 then show ?thesis by (simp add: ffact_Suc_nat ffact_Suc) next case (Suc m) with Suc.hyps show ?thesis by (simp add: algebra_simps ffact_Suc_nat ffact_Suc) qed qed lemma of_int_ffact: "of_int (ffact n k) = ffact n (of_int k :: 'a :: comm_ring_1)" proof (induct n arbitrary: k) case 0 then show ?case by simp next case (Suc n k) then have "of_int (ffact n (k - 1)) = ffact n (of_int (k - 1) :: 'a)" . then show ?case by (simp add: ffact_Suc_nat ffact_Suc) qed lemma ffact_minus: fixes x :: "'a :: comm_ring_1" shows "ffact n (- x) = (- 1) ^ n * pochhammer x n" proof - have "ffact n (- x) = pochhammer (- x + 1 - of_nat n) n" unfolding ffact_def .. also have "\ = pochhammer (- x - of_nat n + 1) n" by (simp add: diff_add_eq) also have "\ = (- 1) ^ n * pochhammer (- (- x)) n" by (rule pochhammer_minus') also have "\ = (- 1) ^ n * pochhammer x n" by simp finally show ?thesis . qed text \Conversion of natural potences into falling factorials and back\ lemma monomial_ffact: "a ^ n = (\k = 0..n. of_nat (Stirling n k) * ffact k a)" for a :: "'a :: comm_ring_1" proof (induct n) case 0 then show ?case by simp next case (Suc n) then have "a ^ Suc n = a * (\k = 0..n. of_nat (Stirling n k) * ffact k a)" by simp also have "\ = (\k = 0..n. of_nat (Stirling n k) * (a * ffact k a))" by (simp add: sum_distrib_left algebra_simps) also have "\ = (\k = 0..n. of_nat (Stirling n k) * ffact (Suc k) a) + (\k = 0..n. of_nat (Stirling n k) * (of_nat k * ffact k a))" by (simp add: sum.distrib algebra_simps mult_ffact) also have "\ = (\k = 0.. Suc n. of_nat (Stirling n k) * ffact (Suc k) a) + (\k = 0..Suc n. of_nat ((Suc k) * (Stirling n (Suc k))) * (ffact (Suc k) a))" proof - have "(\k = 0..n. of_nat (Stirling n k) * (of_nat k * ffact k a)) = (\k = 0..n+2. of_nat (Stirling n k) * (of_nat k * ffact k a))" by simp also have "\ = (\k = Suc 0 .. Suc (Suc n). of_nat (Stirling n k) * (of_nat k * ffact k a)) " by (simp only: sum.atLeast_Suc_atMost [of 0 "n + 2"]) simp also have "\ = (\k = 0 .. Suc n. of_nat (Stirling n (Suc k)) * (of_nat (Suc k) * ffact (Suc k) a))" by (simp only: image_Suc_atLeastAtMost sum.shift_bounds_cl_Suc_ivl) also have "\ = (\k = 0 .. Suc n. of_nat ((Suc k) * Stirling n (Suc k)) * ffact (Suc k) a)" by (simp only: of_nat_mult algebra_simps) finally have "(\k = 0..n. of_nat (Stirling n k) * (of_nat k * ffact k a)) = (\k = 0..Suc n. of_nat (Suc k * Stirling n (Suc k)) * ffact (Suc k) a)" by simp then show ?thesis by simp qed also have "\ = (\k = 0..n. of_nat (Stirling (Suc n) (Suc k)) * ffact (Suc k) a)" by (simp add: algebra_simps sum.distrib) also have "\ = (\k = Suc 0..Suc n. of_nat (Stirling (Suc n) k) * ffact k a)" by (simp only: image_Suc_atLeastAtMost sum.shift_bounds_cl_Suc_ivl) also have "\ = (\k = 0..Suc n. of_nat (Stirling (Suc n) k) * ffact k a)" by (simp only: sum.atLeast_Suc_atMost [of "0" "Suc n"]) simp finally show ?case by simp qed lemma ffact_monomial: "ffact n a = (\k = 0..n. (- 1) ^ (n - k) * of_nat (stirling n k) * a ^ k)" for a :: "'a :: comm_ring_1" proof (induct n) case 0 show ?case by simp next case (Suc n) then have "ffact (Suc n) a = (a - of_nat n) * (\k = 0..n. (- 1) ^ (n - k) * of_nat (stirling n k) * a ^ k)" by (simp add: ffact_Suc_rev) also have "\ = (\k = 0..n. (- 1) ^ (n - k) * of_nat (stirling n k) * a ^ (Suc k)) + (\k = 0..n. (- 1) * (- 1) ^ (n - k) * of_nat (n * (stirling n k)) * a ^ k)" by (simp only: diff_conv_add_uminus distrib_right) (simp add: sum_distrib_left field_simps) also have "\ = (\k = 0..n. (- 1) ^ (Suc n - Suc k) * of_nat (stirling n k) * a ^ Suc k) + (\k = 0..n. (- 1) ^ (Suc n - Suc k) * of_nat (n * stirling n (Suc k)) * a ^ Suc k)" proof - have "(\k = 0..n. (- 1) * (- 1) ^ (n - k) * of_nat (n * stirling n k) * a ^ k) = (\k = 0..n. (- 1) ^ (Suc n - k) * of_nat (n * stirling n k) * a ^ k)" by (simp add: Suc_diff_le) also have "\ = (\k = Suc 0..Suc n. (- 1) ^ (Suc n - k) * of_nat (n * stirling n k) * a ^ k)" by (simp add: sum.atLeast_Suc_atMost) (cases n; simp) also have "\ = (\k = 0..n. (- 1) ^ (Suc n - Suc k) * of_nat (n * stirling n (Suc k)) * a ^ Suc k)" by (simp only: sum.shift_bounds_cl_Suc_ivl) finally show ?thesis by simp qed also have "\ = (\k = 0..n. (- 1) ^ (Suc n - Suc k) * of_nat (n * stirling n (Suc k) + stirling n k) * a ^ Suc k)" by (simp add: sum.distrib algebra_simps) also have "\ = (\k = 0..n. (- 1) ^ (Suc n - Suc k) * of_nat (stirling (Suc n) (Suc k)) * a ^ Suc k)" by (simp only: stirling.simps) also have "\ = (\k = Suc 0..Suc n. (- 1) ^ (Suc n - k) * of_nat (stirling (Suc n) k) * a ^ k)" by (simp only: sum.shift_bounds_cl_Suc_ivl) also have "\ = (\k = 0..Suc n. (- 1) ^ (Suc n - k) * of_nat (stirling (Suc n) k) * a ^ k)" by (simp add: sum.atLeast_Suc_atMost) finally show ?case . qed end diff --git a/thys/Discrete_Summation/ROOT b/thys/Discrete_Summation/ROOT --- a/thys/Discrete_Summation/ROOT +++ b/thys/Discrete_Summation/ROOT @@ -1,10 +1,11 @@ chapter AFP session Discrete_Summation (AFP) = HOL + options [timeout = 600] sessions "HOL-Library" + "HOL-Combinatorics" theories Examples document_files "root.tex" diff --git a/thys/Graph_Theory/Bidirected_Digraph.thy b/thys/Graph_Theory/Bidirected_Digraph.thy --- a/thys/Graph_Theory/Bidirected_Digraph.thy +++ b/thys/Graph_Theory/Bidirected_Digraph.thy @@ -1,110 +1,110 @@ theory Bidirected_Digraph imports Digraph - "HOL-Library.Permutations" + "HOL-Combinatorics.Permutations" begin section \Bidirected Graphs\ locale bidirected_digraph = wf_digraph G for G + fixes arev :: "'b \ 'b" assumes arev_dom: "\a. a \ arcs G \ arev a \ a" assumes arev_arev_raw: "\a. a \ arcs G \ arev (arev a) = a" assumes tail_arev[simp]: "\a. a \ arcs G \ tail G (arev a) = head G a" lemma (in wf_digraph) bidirected_digraphI: assumes arev_eq: "\a. a \ arcs G \ arev a = a" assumes arev_neq: "\a. a \ arcs G \ arev a \ a" assumes arev_arev_raw: "\a. a \ arcs G \ arev (arev a) = a" assumes tail_arev: "\a. a \ arcs G \ tail G (arev a) = head G a" shows "bidirected_digraph G arev" using assms by unfold_locales (auto simp: permutes_def) context bidirected_digraph begin lemma bidirected_digraph[intro!]: "bidirected_digraph G arev" by unfold_locales lemma arev_arev[simp]: "arev (arev a) = a" using arev_dom by (cases "a \ arcs G") (auto simp: arev_arev_raw) lemma arev_o_arev[simp]: "arev o arev = id" by (simp add: fun_eq_iff) lemma arev_eq: "a \ arcs G \ arev a = a" by (simp add: arev_dom) lemma arev_neq: "a \ arcs G \ arev a \ a" by (simp add: arev_dom) lemma arev_in_arcs[simp]: "a \ arcs G \ arev a \ arcs G" by (metis arev_arev arev_dom) lemma head_arev[simp]: assumes "a \ arcs G" shows "head G (arev a) = tail G a" proof - from assms have "head G (arev a) = tail G (arev (arev a)) " by (simp only: tail_arev arev_in_arcs) then show ?thesis by simp qed lemma ate_arev[simp]: assumes "a \ arcs G" shows "arc_to_ends G (arev a) = prod.swap (arc_to_ends G a)" using assms by (auto simp: arc_to_ends_def) lemma bij_arev: "bij arev" using arev_arev by (metis bij_betw_imageI inj_on_inverseI surjI) lemma arev_permutes_arcs: "arev permutes arcs G" using arev_dom bij_arev by (auto simp: permutes_def bij_iff) lemma arev_eq_iff: "\x y. arev x = arev y \ x = y" by (metis arev_arev) lemma in_arcs_eq: "in_arcs G w = arev ` out_arcs G w" by auto (metis arev_arev arev_in_arcs image_eqI in_out_arcs_conv tail_arev) lemma inj_on_arev[intro!]: "inj_on arev S" by (metis arev_arev inj_on_inverseI) lemma even_card_loops: "even (card (in_arcs G w \ out_arcs G w))" (is "even (card ?S)") proof - { assume "\finite ?S" then have ?thesis by simp } moreover { assume A:"finite ?S" have "card ?S = card (\{{a,arev a} | a. a \ ?S})" (is "_ = card (\ ?T)") by (rule arg_cong[where f=card]) (auto intro!: exI[where x="{x, arev x}" for x]) also have "\= sum card ?T" proof (rule card_Union_disjoint) show "\A. A\{{a, arev a} |a. a \ ?S} \ finite A" by auto show "pairwise disjnt {{a, arev a} |a. a \ in_arcs G w \ out_arcs G w}" unfolding pairwise_def disjnt_def by safe (simp_all add: arev_eq_iff) qed also have "\ = sum (\a. 2) ?T" by (intro sum.cong) (auto simp: card_insert_if dest: arev_neq) also have "\ = 2 * card ?T" by simp finally have ?thesis by simp } ultimately show ?thesis by blast qed end (*XXX*) sublocale bidirected_digraph \ sym_digraph proof (unfold_locales, unfold symmetric_def, intro symI) fix u v assume "u \\<^bsub>G\<^esub> v" then obtain a where "a \ arcs G" "arc_to_ends G a = (u,v)" by (auto simp: arcs_ends_def) then have "arev a \ arcs G" "arc_to_ends G (arev a) = (v,u)" by (auto simp: arc_to_ends_def) then show "v \\<^bsub>G\<^esub> u" by (auto simp: arcs_ends_def intro: rev_image_eqI) qed end diff --git a/thys/Graph_Theory/Funpow.thy b/thys/Graph_Theory/Funpow.thy --- a/thys/Graph_Theory/Funpow.thy +++ b/thys/Graph_Theory/Funpow.thy @@ -1,1051 +1,1051 @@ theory Funpow imports "HOL-Library.FuncSet" - "HOL-Library.Permutations" + "HOL-Combinatorics.Permutations" begin section \Auxiliary Lemmas about @{term "(^^)"}\ lemma funpow_simp_l: "f ((f ^^ n) x) = (f ^^ Suc n) x" by (metis comp_apply funpow.simps(2)) lemma funpow_add_app: "(f ^^ n) ((f ^^ m) x) = (f ^^ (n + m)) x" by (metis comp_apply funpow_add) lemma funpow_mod_eq: assumes "(f ^^ n) x = x" "0 < n" shows "(f ^^ (m mod n)) x = (f ^^ m) x" proof (induct m rule: less_induct) case (less m) { assume "m < n" then have ?case by simp } moreover { assume "m = n" then have ?case by (simp add: \_ = x\)} moreover { assume "n < m" then have "m - n < m" "0 < m - n" using \0 < n\ by arith+ have "(f ^^ (m mod n)) x = (f ^^ ((m - n) mod n)) x" using \0 < m - n\ by (simp add: mod_geq) also have "\ = (f ^^ (m - n)) x" using \m - n < m\ by (rule less) also have "\ = (f ^^ (m - n)) ((f ^^ n) x)" by (simp add: assms) also have "\ = (f ^^ m) x" using \0 < m - n\ by (simp add: funpow_add_app) finally have ?case . } ultimately show ?case by (metis linorder_neqE_nat) qed lemma id_funpow_id: assumes "f x = x" shows "(f ^^ n) x = x" using assms by (induct n) auto lemma inv_id_abs[simp]: "inv (\a. a) = id" unfolding id_def[symmetric] by simp lemma inj_funpow: fixes f :: "'a \ 'a" assumes "inj f" shows "inj (f ^^ n)" proof (induct n) case 0 then show ?case by (auto simp: id_def[symmetric]) next case (Suc n) with assms show ?case unfolding funpow.simps by (rule inj_compose) qed lemma funpow_inj_finite: assumes "inj p" "finite {(p ^^ n) x |n. True}" shows "\n>0. (p ^^ n) x = x" proof - have "\finite {0::nat..}" by simp moreover have "{(p ^^ n) x |n. True} = (\n. (p ^^ n) x) ` {0..}" by auto with assms have "finite \" by simp ultimately have "\n\{0..}. \ finite {m \ {0..}. (p ^^ m) x = (p ^^ n) x}" by (rule pigeonhole_infinite) then obtain n where "\finite {m. (p ^^ m) x = (p ^^ n) x}" by auto then have "\finite ({m. (p ^^ m) x = (p ^^ n) x} - {n})" by auto then have "({m. (p ^^ m) x = (p ^^ n) x} - {n}) \ {}" by (metis finite.emptyI) then obtain m where m: "(p ^^ m) x = (p ^^ n) x" "m \ n" by auto { fix m n assume "(p ^^ n) x = (p ^^ m) x" "m < n" have "(p ^^ (n - m)) x = inv (p ^^ m) ((p ^^ m) ((p ^^ (n - m)) x))" using \inj p\ by (simp add: inv_f_f inj_funpow) also have "((p ^^ m) ((p ^^ (n - m)) x)) = (p ^^ n) x" using \m < n\ by (simp add: funpow_add_app) also have "inv (p ^^ m) \ = x" using \inj p\ by (simp add: \(p ^^ n) x = _\ inj_funpow) finally have "(p ^^ (n - m)) x = x" "0 < n - m" using \m < n\ by auto } note general = this show ?thesis proof (cases m n rule: linorder_cases) case less then show ?thesis using general m by metis next case equal then show ?thesis using m by metis next case greater then show ?thesis using general m by metis qed qed lemma permutes_in_funpow_image: assumes "f permutes S" "x \ S" shows "(f ^^ n) x \ S" using assms by (induct n) (auto simp: permutes_in_image) (* XXX move*) lemma permutation_self: assumes "permutation p" shows "\n>0. (p ^^ n) x = x" proof cases assume "p x = x" then show ?thesis by auto next assume "p x \ x" from assms have "inj p" by (intro permutation_bijective bij_is_inj) { fix n from \p x \ x\ have "(p ^^ Suc n) x \ (p ^^ n) x" proof (induct n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have "p (p x) \ p x" proof (rule notI) assume "p (p x) = p x" then show False using \p x \ x\ \inj p\ by (simp add: inj_eq) qed have "(p ^^ Suc (Suc n)) x = (p ^^ Suc n) (p x)" by (metis funpow_simp_l funpow_swap1) also have "\ \ (p ^^ n) (p x)" by (rule Suc) fact also have "(p ^^ n) (p x) = (p ^^ Suc n) x" by (metis funpow_simp_l funpow_swap1) finally show ?case by simp qed } then have "{(p ^^ n) x | n. True} \ {x. p x \ x}" by auto then have "finite {(p ^^ n) x | n. True}" using permutation_finite_support[OF assms] by (rule finite_subset) with \inj p\ show ?thesis by (rule funpow_inj_finite) qed (* XXX move *) lemma (in -) funpow_invs: assumes "m \ n" and inv: "\x. f (g x) = x" shows "(f ^^ m) ((g ^^ n) x) = (g ^^ (n - m)) x" using \m \ n\ proof (induction m) case (Suc m) moreover then have "n - m = Suc (n - Suc m)" by auto ultimately show ?case by (auto simp: inv) qed simp section \Function-power distance between values\ (* xxx move *) definition funpow_dist :: "('a \ 'a) \ 'a \ 'a \ nat" where "funpow_dist f x y \ LEAST n. (f ^^ n) x = y" abbreviation funpow_dist1 :: "('a \ 'a) \ 'a \ 'a \ nat" where "funpow_dist1 f x y \ Suc (funpow_dist f (f x) y)" lemma funpow_dist_0: assumes "x = y" shows "funpow_dist f x y = 0" using assms unfolding funpow_dist_def by (intro Least_eq_0) simp lemma funpow_dist_least: assumes "n < funpow_dist f x y" shows "(f ^^ n) x \ y" proof (rule notI) assume "(f ^^ n) x = y" then have "funpow_dist f x y \ n" unfolding funpow_dist_def by (rule Least_le) with assms show False by linarith qed lemma funpow_dist1_least: assumes "0 < n" "n < funpow_dist1 f x y" shows "(f ^^ n) x \ y" proof (rule notI) assume "(f ^^ n) x = y" then have "(f ^^ (n - 1)) (f x) = y" using \0 < n\ by (cases n) (simp_all add: funpow_swap1) then have "funpow_dist f (f x) y \ n - 1" unfolding funpow_dist_def by (rule Least_le) with assms show False by simp qed section \Cyclic Permutations\ inductive_set orbit :: "('a \ 'a) \ 'a \ 'a set" for f x where base: "f x \ orbit f x" | step: "y \ orbit f x \ f y \ orbit f x" definition cyclic_on :: "('a \ 'a) \ 'a set \ bool" where "cyclic_on f S \ (\s\S. S = orbit f s)" lemma orbit_altdef: "orbit f x = {(f ^^ n) x | n. 0 < n}" (is "?L = ?R") proof (intro set_eqI iffI) fix y assume "y \ ?L" then show "y \ ?R" by (induct rule: orbit.induct) (auto simp: exI[where x=1] exI[where x="Suc n" for n]) next fix y assume "y \ ?R" then obtain n where "y = (f ^^ n) x" "0 < n" by blast then show "y \ ?L" proof (induction n arbitrary: y) case (Suc n) then show ?case by (cases "n = 0") (auto intro: orbit.intros) qed simp qed lemma orbit_trans: assumes "s \ orbit f t" "t \ orbit f u" shows "s \ orbit f u" using assms by induct (auto intro: orbit.intros) lemma orbit_subset: assumes "s \ orbit f (f t)" shows "s \ orbit f t" using assms by (induct) (auto intro: orbit.intros) lemma orbit_sim_step: assumes "s \ orbit f t" shows "f s \ orbit f (f t)" using assms by induct (auto intro: orbit.intros) lemma orbit_step: assumes "y \ orbit f x" "f x \ y" shows "y \ orbit f (f x)" using assms proof induction case (step y) then show ?case by (cases "x = y") (auto intro: orbit.intros) qed simp lemma self_in_orbit_trans: assumes "s \ orbit f s" "t \ orbit f s" shows "t \ orbit f t" using assms(2,1) by induct (auto intro: orbit_sim_step) lemma orbit_swap: assumes "s \ orbit f s" "t \ orbit f s" shows "s \ orbit f t" using assms(2,1) proof induction case base then show ?case by (cases "f s = s") (auto intro: orbit_step) next case (step x) then show ?case by (cases "f x = s") (auto intro: orbit_step) qed lemma permutation_self_in_orbit: assumes "permutation f" shows "s \ orbit f s" unfolding orbit_altdef using permutation_self[OF assms, of s] by simp metis lemma orbit_altdef_self_in: assumes "s \ orbit f s" shows "orbit f s = {(f ^^ n) s | n. True}" proof (intro set_eqI iffI) fix x assume "x \ {(f ^^ n) s | n. True}" then obtain n where "x = (f ^^ n) s" by auto then show "x \ orbit f s" using assms by (cases "n = 0") (auto simp: orbit_altdef) qed (auto simp: orbit_altdef) lemma orbit_altdef_permutation: assumes "permutation f" shows "orbit f s = {(f ^^ n) s | n. True}" using assms by (intro orbit_altdef_self_in permutation_self_in_orbit) lemma orbit_altdef_bounded: assumes "(f ^^ n) s = s" "0 < n" shows "orbit f s = {(f ^^ m) s| m. m < n}" proof - from assms have "s \ orbit f s" unfolding orbit_altdef by auto metis then have "orbit f s = {(f ^^ m) s|m. True}" by (rule orbit_altdef_self_in) also have "\ = {(f ^^ m) s| m. m < n}" using assms by (auto simp: funpow_mod_eq intro: exI[where x="m mod n" for m]) finally show ?thesis . qed lemma funpow_in_orbit: assumes "s \ orbit f t" shows "(f ^^ n) s \ orbit f t" using assms by (induct n) (auto intro: orbit.intros) lemma finite_orbit: assumes "s \ orbit f s" shows "finite (orbit f s)" proof - from assms obtain n where n: "0 < n" "(f ^^n) s = s" by (auto simp: orbit_altdef) then show ?thesis by (auto simp: orbit_altdef_bounded) qed lemma self_in_orbit_step: assumes "s \ orbit f s" shows "orbit f (f s) = orbit f s" proof (intro set_eqI iffI) fix t assume "t \ orbit f s" then show "t \ orbit f (f s)" using assms by (auto intro: orbit_step orbit_sim_step) qed (auto intro: orbit_subset) lemma permutation_orbit_step: assumes "permutation f" shows "orbit f (f s) = orbit f s" using assms by (intro self_in_orbit_step permutation_self_in_orbit) lemma orbit_nonempty: "orbit f s \ {}" using orbit.base by fastforce lemma orbit_inv_eq: assumes "permutation f" shows "orbit (inv f) x = orbit f x" (is "?L = ?R") proof - { fix g y assume A: "permutation g" "y \ orbit (inv g) x" have "y \ orbit g x" proof - have inv_g: "\y. x = g y \ inv g x = y" "\y. inv g (g y) = y" by (metis A(1) bij_inv_eq_iff permutation_bijective)+ { fix y assume "y \ orbit g x" then have "inv g y \ orbit g x" by (cases) (simp_all add: inv_g A(1) permutation_self_in_orbit) } note inv_g_in_orb = this from A(2) show ?thesis by induct (simp_all add: inv_g_in_orb A permutation_self_in_orbit) qed } note orb_inv_ss = this have "inv (inv f) = f" by (simp add: assms inv_inv_eq permutation_bijective) then show ?thesis using orb_inv_ss[OF assms] orb_inv_ss[OF permutation_inverse[OF assms]] by auto qed lemma cyclic_on_alldef: "cyclic_on f S \ S \ {} \ (\s\S. S = orbit f s)" unfolding cyclic_on_def by (auto intro: orbit.step orbit_swap orbit_trans) lemma cyclic_on_funpow_in: assumes "cyclic_on f S" "s \ S" shows "(f^^n) s \ S" using assms unfolding cyclic_on_def by (auto intro: funpow_in_orbit) lemma finite_cyclic_on: assumes "cyclic_on f S" shows "finite S" using assms by (auto simp: cyclic_on_def finite_orbit) lemma cyclic_on_singleI: assumes "s \ S" "S = orbit f s" shows "cyclic_on f S" using assms unfolding cyclic_on_def by blast lemma inj_on_funpow_least: assumes "(f ^^ n) s = s" "\m. \m < n; 0 < m\ \ (f ^^ m) s \ s" shows "inj_on (\k. (f^^k) s) {0.. l" "(f ^^ k) s = (f ^^ l) s" define k' l' where "k' = min k l" and "l' = max k l" with A have A': "k' < l'" "(f ^^ k') s = (f ^^ l') s" "l' < n" by (auto simp: min_def max_def) have "s = (f ^^ ((n - l') + l')) s" using assms \l' < n\ by simp also have "\ = (f ^^ (n - l')) ((f ^^ l') s)" by (simp add: funpow_add) also have "(f ^^ l') s = (f ^^ k') s" by (simp add: A') also have "(f ^^ (n - l')) \ = (f ^^ (n - l' + k')) s" by (simp add: funpow_add) finally have "(f ^^ (n - l' + k')) s = s" by simp moreover have "n - l' + k' < n" "0 < n - l' + k'"using A' by linarith+ ultimately have False using assms(2) by auto } then show ?thesis by (intro inj_onI) auto qed lemma cyclic_on_inI: assumes "cyclic_on f S" "s \ S" shows "f s \ S" using assms by (auto simp: cyclic_on_def intro: orbit.intros) lemma bij_betw_funpow: assumes "bij_betw f S S" shows "bij_betw (f ^^ n) S S" proof (induct n) case 0 then show ?case by (auto simp: id_def[symmetric]) next case (Suc n) then show ?case unfolding funpow.simps using assms by (rule bij_betw_trans) qed (*XXX rename move*) lemma orbit_FOO: assumes self:"a \ orbit g a" and eq: "\x. x \ orbit g a \ g' (f x) = f (g x)" shows "f ` orbit g a = orbit g' (f a)" (is "?L = ?R") proof (intro set_eqI iffI) fix x assume "x \ ?L" then obtain x0 where "x0 \ orbit g a" "x = f x0" by auto then show "x \ ?R" proof (induct arbitrary: x) case base then show ?case by (auto simp: self orbit.base eq[symmetric]) next case step then show ?case by cases (auto simp: eq[symmetric] orbit.intros) qed next fix x assume "x \ ?R" then show "x \ ?L" proof (induct arbitrary: ) case base then show ?case by (auto simp: self orbit.base eq) next case step then show ?case by cases (auto simp: eq orbit.intros) qed qed lemma cyclic_on_FOO: assumes "cyclic_on f S" assumes "\x. x \ S \ g (h x) = h (f x)" shows "cyclic_on g (h ` S)" using assms by (auto simp: cyclic_on_def) (meson orbit_FOO) lemma cyclic_on_f_in: assumes "f permutes S" "cyclic_on f A" "f x \ A" shows "x \ A" proof - from assms have fx_in_orb: "f x \ orbit f (f x)" by (auto simp: cyclic_on_alldef) from assms have "A = orbit f (f x)" by (auto simp: cyclic_on_alldef) moreover then have "\ = orbit f x" using \f x \ A\ by (auto intro: orbit_step orbit_subset) ultimately show ?thesis by (metis (no_types) orbit.simps permutes_inverses(2)[OF assms(1)]) qed lemma permutes_not_in: assumes "f permutes S" "x \ S" shows "f x = x" using assms by (auto simp: permutes_def) lemma orbit_cong0: assumes "x \ A" "f \ A \ A" "\y. y \ A \ f y = g y" shows "orbit f x = orbit g x" proof - { fix n have "(f ^^ n) x = (g ^^ n) x \ (f ^^ n) x \ A" by (induct n rule: nat.induct) (insert assms, auto) } then show ?thesis by (auto simp: orbit_altdef) qed lemma orbit_cong: assumes self_in: "t \ orbit f t" and eq: "\s. s \ orbit f t \ g s = f s" shows "orbit g t = orbit f t" using assms(1) _ assms(2) by (rule orbit_cong0) (auto simp: orbit.step eq) lemma cyclic_cong: assumes "\s. s \ S \ f s = g s" shows "cyclic_on f S = cyclic_on g S" proof - have "(\s\S. orbit f s = orbit g s) \ cyclic_on f S = cyclic_on g S" by (metis cyclic_on_alldef cyclic_on_def) then show ?thesis by (metis assms orbit_cong cyclic_on_def) qed lemma permutes_comp_preserves_cyclic1: assumes "g permutes B" "cyclic_on f C" assumes "A \ B = {}" "C \ A" shows "cyclic_on (f o g) C" proof - have *: "\c. c \ C \ f (g c) = f c" using assms by (subst permutes_not_in[where f=g]) auto with assms(2) show ?thesis by (simp cong: cyclic_cong) qed lemma permutes_comp_preserves_cyclic2: assumes "f permutes A" "cyclic_on g C" assumes "A \ B = {}" "C \ B" shows "cyclic_on (f o g) C" proof - obtain c where c: "c \ C" "C = orbit g c" "c \ orbit g c" using \cyclic_on g C\ by (auto simp: cyclic_on_def) then have "\c. c \ C \ f (g c) = g c" using assms c by (subst permutes_not_in[where f=f]) (auto intro: orbit.intros) with assms(2) show ?thesis by (simp cong: cyclic_cong) qed (*XXX merge with previous section?*) subsection \Orbits\ lemma permutes_orbit_subset: assumes "f permutes S" "x \ S" shows "orbit f x \ S" proof fix y assume "y \ orbit f x" then show "y \ S" by induct (auto simp: permutes_in_image assms) qed lemma cyclic_on_orbit': assumes "permutation f" shows "cyclic_on f (orbit f x)" unfolding cyclic_on_alldef using orbit_nonempty[of f x] by (auto intro: assms orbit_swap orbit_trans permutation_self_in_orbit) (* XXX remove? *) lemma cyclic_on_orbit: assumes "f permutes S" "finite S" shows "cyclic_on f (orbit f x)" using assms by (intro cyclic_on_orbit') (auto simp: permutation_permutes) lemma orbit_cyclic_eq3: assumes "cyclic_on f S" "y \ S" shows "orbit f y = S" using assms unfolding cyclic_on_alldef by simp (*XXX move*) lemma orbit_eq_singleton_iff: "orbit f x = {x} \ f x = x" (is "?L \ ?R") proof assume A: ?R { fix y assume "y \ orbit f x" then have "y = x" by induct (auto simp: A) } then show ?L by (metis orbit_nonempty singletonI subsetI subset_singletonD) next assume A: ?L then have "\y. y \ orbit f x \ f x = y" by - (erule orbit.cases, simp_all) then show ?R using A by blast qed (* XXX move *) lemma eq_on_cyclic_on_iff1: assumes "cyclic_on f S" "x \ S" obtains "f x \ S" "f x = x \ card S = 1" proof from assms show "f x \ S" by (auto simp: cyclic_on_def intro: orbit.intros) from assms have "S = orbit f x" by (auto simp: cyclic_on_alldef) then have "f x = x \ S = {x}" by (metis orbit_eq_singleton_iff) then show "f x = x \ card S = 1" using \x \ S\ by (auto simp: card_Suc_eq) qed subsection \Decomposition of Arbitrary Permutations\ definition perm_restrict :: "('a \ 'a) \ 'a set \ ('a \ 'a)" where "perm_restrict f S x \ if x \ S then f x else x" lemma perm_restrict_comp: assumes "A \ B = {}" "cyclic_on f B" shows "perm_restrict f A o perm_restrict f B = perm_restrict f (A \ B)" proof - have "\x. x \ B \ f x \ B" using \cyclic_on f B\ by (rule cyclic_on_inI) with assms show ?thesis by (auto simp: perm_restrict_def fun_eq_iff) qed lemma perm_restrict_simps: "x \ S \ perm_restrict f S x = f x" "x \ S \ perm_restrict f S x = x" by (auto simp: perm_restrict_def) lemma perm_restrict_perm_restrict: "perm_restrict (perm_restrict f A) B = perm_restrict f (A \ B)" by (auto simp: perm_restrict_def) lemma perm_restrict_union: assumes "perm_restrict f A permutes A" "perm_restrict f B permutes B" "A \ B = {}" shows "perm_restrict f A o perm_restrict f B = perm_restrict f (A \ B)" using assms by (auto simp: fun_eq_iff perm_restrict_def permutes_def) (metis Diff_iff Diff_triv) lemma perm_restrict_id[simp]: assumes "f permutes S" shows "perm_restrict f S = f" using assms by (auto simp: permutes_def perm_restrict_def) lemma cyclic_on_perm_restrict: "cyclic_on (perm_restrict f S) S \ cyclic_on f S" by (simp add: perm_restrict_def cong: cyclic_cong) lemma perm_restrict_diff_cyclic: assumes "f permutes S" "cyclic_on f A" shows "perm_restrict f (S - A) permutes (S - A)" proof - { fix y have "\x. perm_restrict f (S - A) x = y" proof cases assume A: "y \ S - A" with \f permutes S\ obtain x where "f x = y" "x \ S" unfolding permutes_def by auto metis moreover with A have "x \ A" by (metis Diff_iff assms(2) cyclic_on_inI) ultimately have "perm_restrict f (S - A) x = y" by (simp add: perm_restrict_simps) then show ?thesis .. next assume "y \ S - A" then have "perm_restrict f (S - A) y = y" by (simp add: perm_restrict_simps) then show ?thesis .. qed } note X = this { fix x y assume "perm_restrict f (S - A) x = perm_restrict f (S - A) y" with assms have "x = y" by (auto simp: perm_restrict_def permutes_def split: if_splits intro: cyclic_on_f_in) } note Y = this show ?thesis by (auto simp: permutes_def perm_restrict_simps X intro: Y) qed lemma orbit_eqI: "y = f x \ y \ orbit f x" "z = f y \y \ orbit f x \z \ orbit f x" by (metis orbit.base) (metis orbit.step) lemma permutes_decompose: assumes "f permutes S" "finite S" shows "\C. (\c \ C. cyclic_on f c) \ \C = S \ (\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = {})" using assms(2,1) proof (induction arbitrary: f rule: finite_psubset_induct) case (psubset S) show ?case proof (cases "S = {}") case True then show ?thesis by (intro exI[where x="{}"]) auto next case False then obtain s where "s \ S" by auto with \f permutes S\ have "orbit f s \ S" by (rule permutes_orbit_subset) have cyclic_orbit: "cyclic_on f (orbit f s)" using \f permutes S\ \finite S\ by (rule cyclic_on_orbit) let ?f' = "perm_restrict f (S - orbit f s)" have "f s \ S" using \f permutes S\ \s \ S\ by (auto simp: permutes_in_image) then have "S - orbit f s \ S" using orbit.base[of f s] \s \ S\ by blast moreover have "?f' permutes (S - orbit f s)" using \f permutes S\ cyclic_orbit by (rule perm_restrict_diff_cyclic) ultimately obtain C where C: "\c. c \ C \ cyclic_on ?f' c" "\C = S - orbit f s" "\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = {}" using psubset.IH by metis { fix c assume "c \ C" then have *: "\x. x \ c \ perm_restrict f (S - orbit f s) x = f x" using C(2) \f permutes S\ by (auto simp add: perm_restrict_def) then have "cyclic_on f c" using C(1)[OF \c \ C\] by (simp cong: cyclic_cong add: *) } note in_C_cyclic = this have Un_ins: "\(insert (orbit f s) C) = S" using \\C = _\ \orbit f s \ S\ by blast have Disj_ins: "(\c1 \ insert (orbit f s) C. \c2 \ insert (orbit f s) C. c1 \ c2 \ c1 \ c2 = {})" using C by auto show ?thesis by (intro conjI Un_ins Disj_ins exI[where x="insert (orbit f s) C"]) (auto simp: cyclic_orbit in_C_cyclic) qed qed subsection \Funpow + Orbit\ lemma funpow_dist_prop: "y \ orbit f x \ (f ^^ funpow_dist f x y) x = y" unfolding funpow_dist_def by (rule LeastI_ex) (auto simp: orbit_altdef) lemma funpow_dist_0_eq: assumes "y \ orbit f x" shows "funpow_dist f x y = 0 \ x = y" using assms by (auto simp: funpow_dist_0 dest: funpow_dist_prop) lemma funpow_dist_step: assumes "x \ y" "y \ orbit f x" shows "funpow_dist f x y = Suc (funpow_dist f (f x) y)" proof - from \y \ _\ obtain n where "(f ^^ n) x = y" by (auto simp: orbit_altdef) with \x \ y\ obtain n' where [simp]: "n = Suc n'" by (cases n) auto show ?thesis unfolding funpow_dist_def proof (rule Least_Suc2) show "(f ^^ n) x = y" by fact then show "(f ^^ n') (f x) = y" by (simp add: funpow_swap1) show "(f ^^ 0) x \ y" using \x \ y\ by simp show "\k. ((f ^^ Suc k) x = y) = ((f ^^ k) (f x) = y)" by (simp add: funpow_swap1) qed qed lemma funpow_dist1_prop: assumes "y \ orbit f x" shows "(f ^^ funpow_dist1 f x y) x = y" by (metis assms funpow.simps(1) funpow_dist_0 funpow_dist_prop funpow_simp_l funpow_swap1 id_apply orbit_step) (*XXX simplify? *) lemma funpow_neq_less_funpow_dist: assumes "y \ orbit f x" "m \ funpow_dist f x y" "n \ funpow_dist f x y" "m \ n" shows "(f ^^ m) x \ (f ^^ n) x" proof (rule notI) assume A: "(f ^^ m) x = (f ^^ n) x" define m' n' where "m' = min m n" and "n' = max m n" with A assms have A': "m' < n'" "(f ^^ m') x = (f ^^ n') x" "n' \ funpow_dist f x y" by (auto simp: min_def max_def) have "y = (f ^^ funpow_dist f x y) x" using \y \ _\ by (simp only: funpow_dist_prop) also have "\ = (f ^^ ((funpow_dist f x y - n') + n')) x" using \n' \ _\ by simp also have "\ = (f ^^ ((funpow_dist f x y - n') + m')) x" by (simp add: funpow_add \(f ^^ m') x = _\) also have "(f ^^ ((funpow_dist f x y - n') + m')) x \ y" using A' by (intro funpow_dist_least) linarith finally show "False" by simp qed (* XXX reduce to funpow_neq_less_funpow_dist? *) lemma funpow_neq_less_funpow_dist1: assumes "y \ orbit f x" "m < funpow_dist1 f x y" "n < funpow_dist1 f x y" "m \ n" shows "(f ^^ m) x \ (f ^^ n) x" proof (rule notI) assume A: "(f ^^ m) x = (f ^^ n) x" define m' n' where "m' = min m n" and "n' = max m n" with A assms have A': "m' < n'" "(f ^^ m') x = (f ^^ n') x" "n' < funpow_dist1 f x y" by (auto simp: min_def max_def) have "y = (f ^^ funpow_dist1 f x y) x" using \y \ _\ by (simp only: funpow_dist1_prop) also have "\ = (f ^^ ((funpow_dist1 f x y - n') + n')) x" using \n' < _\ by simp also have "\ = (f ^^ ((funpow_dist1 f x y - n') + m')) x" by (simp add: funpow_add \(f ^^ m') x = _\) also have "(f ^^ ((funpow_dist1 f x y - n') + m')) x \ y" using A' by (intro funpow_dist1_least) linarith+ finally show "False" by simp qed lemma inj_on_funpow_dist: assumes "y \ orbit f x" shows "inj_on (\n. (f ^^ n) x) {0..funpow_dist f x y}" using funpow_neq_less_funpow_dist[OF assms] by (intro inj_onI) auto lemma inj_on_funpow_dist1: assumes "y \ orbit f x" shows "inj_on (\n. (f ^^ n) x) {0.. orbit f x" shows "orbit f x = (\n. (f ^^ n) x) ` {0.. orbit f x" by (auto simp: orbit_altdef) then show ?thesis by (rule funpow_dist1_prop) qed lemma funpow_dist1_dist: assumes "funpow_dist1 f x y < funpow_dist1 f x z" assumes "{y,z} \ orbit f x" shows "funpow_dist1 f x z = funpow_dist1 f x y + funpow_dist1 f y z" (is "?L = ?R") proof - have x_z: "(f ^^ funpow_dist1 f x z) x = z" using assms by (blast intro: funpow_dist1_prop) have x_y: "(f ^^ funpow_dist1 f x y) x = y" using assms by (blast intro: funpow_dist1_prop) have "(f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) y = (f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) ((f ^^ funpow_dist1 f x y) x)" using x_y by simp also have "\ = z" using assms x_z by (simp del: funpow.simps add: funpow_add_app) finally have y_z_diff: "(f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) y = z" . then have "(f ^^ funpow_dist1 f y z) y = z" using assms by (intro funpow_dist1_prop1) auto then have "(f ^^ funpow_dist1 f y z) ((f ^^ funpow_dist1 f x y) x) = z" using x_y by simp then have "(f ^^ (funpow_dist1 f y z + funpow_dist1 f x y)) x = z" by (simp del: funpow.simps add: funpow_add_app) show ?thesis proof (rule antisym) from y_z_diff have "(f ^^ funpow_dist1 f y z) y = z" using assms by (intro funpow_dist1_prop1) auto then have "(f ^^ funpow_dist1 f y z) ((f ^^ funpow_dist1 f x y) x) = z" using x_y by simp then have "(f ^^ (funpow_dist1 f y z + funpow_dist1 f x y)) x = z" by (simp del: funpow.simps add: funpow_add_app) then have "funpow_dist1 f x z \ funpow_dist1 f y z + funpow_dist1 f x y" using funpow_dist1_least not_less by fastforce then show "?L \ ?R" by presburger next have "funpow_dist1 f y z \ funpow_dist1 f x z - funpow_dist1 f x y" using y_z_diff assms(1) by (metis not_less zero_less_diff funpow_dist1_least) then show "?R \ ?L" by linarith qed qed lemma funpow_dist1_le_self: assumes "(f ^^ m) x = x" "0 < m" "y \ orbit f x" shows "funpow_dist1 f x y \ m" proof (cases "x = y") case True with assms show ?thesis by (auto dest!: funpow_dist1_least) next case False have "(f ^^ funpow_dist1 f x y) x = (f ^^ (funpow_dist1 f x y mod m)) x" using assms by (simp add: funpow_mod_eq) with False \y \ orbit f x\ have "funpow_dist1 f x y \ funpow_dist1 f x y mod m" by auto (metis funpow_dist_least funpow_dist_prop funpow_dist_step funpow_simp_l not_less) with \m > 0\ show ?thesis by (auto intro: order_trans) qed subsection \Permutation Domains\ definition has_dom :: "('a \ 'a) \ 'a set \ bool" where "has_dom f S \ \s. s \ S \ f s = s" lemma permutes_conv_has_dom: "f permutes S \ bij f \ has_dom f S" by (auto simp: permutes_def has_dom_def bij_iff) section \Segments\ inductive_set segment :: "('a \ 'a) \ 'a \ 'a \ 'a set" for f a b where base: "f a \ b \ f a \ segment f a b" | step: "x \ segment f a b \ f x \ b \ f x \ segment f a b" lemma segment_step_2D: assumes "x \ segment f a (f b)" shows "x \ segment f a b \ x = b" using assms by induct (auto intro: segment.intros) lemma not_in_segment2D: assumes "x \ segment f a b" shows "x \ b" using assms by induct auto lemma segment_altdef: assumes "b \ orbit f a" shows "segment f a b = (\n. (f ^^ n) a) ` {1.. ?L" have "f a \b \ b \ orbit f (f a)" using assms by (simp add: orbit_step) then have *: "f a \ b \ 0 < funpow_dist f (f a) b" using assms using gr0I funpow_dist_0_eq[OF \_ \ b \ orbit f (f a)\] by (simp add: orbit.intros) from \x \ ?L\ show "x \ ?R" proof induct case base then show ?case by (intro image_eqI[where x=1]) (auto simp: *) next case step then show ?case using assms funpow_dist1_prop less_antisym by (fastforce intro!: image_eqI[where x="Suc n" for n]) qed next fix x assume "x \ ?R" then obtain n where "(f ^^ n ) a = x" "0 < n" "n < funpow_dist1 f a b" by auto then show "x \ ?L" proof (induct n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have "(f ^^ Suc n) a \ b" using Suc by (meson funpow_dist1_least) with Suc show ?case by (cases "n = 0") (auto intro: segment.intros) qed qed (*XXX move up*) lemma segmentD_orbit: assumes "x \ segment f y z" shows "x \ orbit f y" using assms by induct (auto intro: orbit.intros) lemma segment1_empty: "segment f x (f x) = {}" by (auto simp: segment_altdef orbit.base funpow_dist_0) lemma segment_subset: assumes "y \ segment f x z" assumes "w \ segment f x y" shows "w \ segment f x z" using assms by (induct arbitrary: w) (auto simp: segment1_empty intro: segment.intros dest: segment_step_2D elim: segment.cases) (* XXX move up*) lemma not_in_segment1: assumes "y \ orbit f x" shows "x \ segment f x y" proof assume "x \ segment f x y" then obtain n where n: "0 < n" "n < funpow_dist1 f x y" "(f ^^ n) x = x" using assms by (auto simp: segment_altdef Suc_le_eq) then have neq_y: "(f ^^ (funpow_dist1 f x y - n)) x \ y" by (simp add: funpow_dist1_least) have "(f ^^ (funpow_dist1 f x y - n)) x = (f ^^ (funpow_dist1 f x y - n)) ((f ^^ n) x)" using n by (simp add: funpow_add_app) also have "\ = (f ^^ funpow_dist1 f x y) x" using \n < _\ by (simp add: funpow_add_app) also have "\ = y" using assms by (rule funpow_dist1_prop) finally show False using neq_y by contradiction qed lemma not_in_segment2: "y \ segment f x y" using not_in_segment2D by metis (*XXX move*) lemma in_segmentE: assumes "y \ segment f x z" "z \ orbit f x" obtains "(f ^^ funpow_dist1 f x y) x = y" "funpow_dist1 f x y < funpow_dist1 f x z" proof from assms show "(f ^^ funpow_dist1 f x y) x = y" by (intro segmentD_orbit funpow_dist1_prop) moreover obtain n where "(f ^^ n) x = y" "0 < n" "n < funpow_dist1 f x z" using assms by (auto simp: segment_altdef) moreover then have "funpow_dist1 f x y \ n" by (meson funpow_dist1_least not_less) ultimately show "funpow_dist1 f x y < funpow_dist1 f x z" by linarith qed (*XXX move*) lemma cyclic_split_segment: assumes S: "cyclic_on f S" "a \ S" "b \ S" and "a \ b" shows "S = {a,b} \ segment f a b \ segment f b a" (is "?L = ?R") proof (intro set_eqI iffI) fix c assume "c \ ?L" with S have "c \ orbit f a" unfolding cyclic_on_alldef by auto then show "c \ ?R" by induct (auto intro: segment.intros) next fix c assume "c \ ?R" moreover have "segment f a b \ orbit f a" "segment f b a \ orbit f b" by (auto dest: segmentD_orbit) ultimately show "c \ ?L" using S by (auto simp: cyclic_on_alldef) qed (*XXX move*) lemma segment_split: assumes y_in_seg: "y \ segment f x z" shows "segment f x z = segment f x y \ {y} \ segment f y z" (is "?L = ?R") proof (intro set_eqI iffI) fix w assume "w \ ?L" then show "w \ ?R" by induct (auto intro: segment.intros) next fix w assume "w \ ?R" moreover { assume "w \ segment f x y" then have "w \ segment f x z" using segment_subset[OF y_in_seg] by auto } moreover { assume "w \ segment f y z" then have "w \ segment f x z" using y_in_seg by induct (auto intro: segment.intros) } ultimately show "w \ ?L" using y_in_seg by (auto intro: segment.intros) qed lemma in_segmentD_inv: assumes "x \ segment f a b" "x \ f a" assumes "inj f" shows "inv f x \ segment f a b" using assms by (auto elim: segment.cases) lemma in_orbit_invI: assumes "b \ orbit f a" assumes "inj f" shows "a \ orbit (inv f) b" using assms(1) apply induct apply (simp add: assms(2) orbit_eqI(1)) by (metis assms(2) inv_f_f orbit.base orbit_trans) lemma segment_step_2: assumes A: "x \ segment f a b" "b \ a" and "inj f" shows "x \ segment f a (f b)" using A by induct (auto intro: segment.intros dest: not_in_segment2D injD[OF \inj f\]) lemma inv_end_in_segment: assumes "b \ orbit f a" "f a \ b" "bij f" shows "inv f b \ segment f a b" using assms(1,2) proof induct case base then show ?case by simp next case (step x) moreover from \bij f\ have "inj f" by (rule bij_is_inj) moreover then have "x \ f x \ f a = x \ x \ segment f a (f x)" by (meson segment.simps) moreover have "x \ f x" using step \inj f\ by (metis in_orbit_invI inv_f_eq not_in_segment1 segment.base) then have "inv f x \ segment f a (f x) \ x \ segment f a (f x)" using \bij f\ \inj f\ by (auto dest: segment.step simp: surj_f_inv_f bij_is_surj) then have "inv f x \ segment f a x \ x \ segment f a (f x)" using \f a \ f x\ \inj f\ by (auto dest: segment_step_2 injD) ultimately show ?case by (cases "f a = x") simp_all qed lemma segment_overlapping: assumes "x \ orbit f a" "x \ orbit f b" "bij f" shows "segment f a x \ segment f b x \ segment f b x \ segment f a x" using assms(1,2) proof induction case base then show ?case by (simp add: segment1_empty) next case (step x) from \bij f\ have "inj f" by (simp add: bij_is_inj) have *: "\f x y. y \ segment f x (f x) \ False" by (simp add: segment1_empty) { fix y z assume A: "y \ segment f b (f x)" "y \ segment f a (f x)" "z \ segment f a (f x)" from \x \ orbit f a\ \f x \ orbit f b\ \y \ segment f b (f x)\ have "x \ orbit f b" by (metis * inv_end_in_segment[OF _ _ \bij f\] inv_f_eq[OF \inj f\] segmentD_orbit) moreover with \x \ orbit f a\ step.IH have "segment f a (f x) \ segment f b (f x) \ segment f b (f x) \ segment f a (f x)" apply auto apply (metis * inv_end_in_segment[OF _ _ \bij f\] inv_f_eq[OF \inj f\] segment_step_2D segment_subset step.prems subsetCE) by (metis (no_types, lifting) \inj f\ * inv_end_in_segment[OF _ _ \bij f\] inv_f_eq orbit_eqI(2) segment_step_2D segment_subset subsetCE) ultimately have "segment f a (f x) \ segment f b (f x)" using A by auto } note C = this then show ?case by auto qed lemma segment_disj: assumes "a \ b" "bij f" shows "segment f a b \ segment f b a = {}" proof (rule ccontr) assume "\?thesis" then obtain x where x: "x \ segment f a b" "x \ segment f b a" by blast then have "segment f a b = segment f a x \ {x} \ segment f x b" "segment f b a = segment f b x \ {x} \ segment f x a" by (auto dest: segment_split) then have o: "x \ orbit f a" "x \ orbit f b" by (auto dest: segmentD_orbit) note * = segment_overlapping[OF o \bij f\] have "inj f" using \bij f\ by (simp add: bij_is_inj) have "segment f a x = segment f b x" proof (intro set_eqI iffI) fix y assume A: "y \ segment f b x" then have "y \ segment f a x \ f a \ segment f b a" using * x(2) by (auto intro: segment.base segment_subset) then show "y \ segment f a x" using \inj f\ A by (metis (no_types) not_in_segment2 segment_step_2) next fix y assume A: "y \ segment f a x " then have "y \ segment f b x \ f b \ segment f a b" using * x(1) by (auto intro: segment.base segment_subset) then show "y \ segment f b x" using \inj f\ A by (metis (no_types) not_in_segment2 segment_step_2) qed moreover have "segment f a x \ segment f b x" by (metis assms bij_is_inj not_in_segment2 segment.base segment_step_2 segment_subset x(1)) ultimately show False by contradiction qed lemma segment_x_x_eq: assumes "permutation f" shows "segment f x x = orbit f x - {x}" (is "?L = ?R") proof (intro set_eqI iffI) fix y assume "y \ ?L" then show "y \ ?R" by (auto dest: segmentD_orbit simp: not_in_segment2) next fix y assume "y \ ?R" then have "y \ orbit f x" "y \ x" by auto then show "y \ ?L" by induct (auto intro: segment.intros) qed section \Lists of Powers\ definition iterate :: "nat \ nat \ ('a \ 'a ) \ 'a \ 'a list" where "iterate m n f x = map (\n. (f^^n) x) [m..k. (f ^^ k) x) ` {m.. m \ n" by (auto simp: iterate_def) lemma iterate_length[simp]: "length (iterate m n f x) = n - m" by (auto simp: iterate_def) lemma iterate_nth[simp]: assumes "k < n - m" shows "iterate m n f x ! k = (f^^(m+k)) x" using assms by (induct k arbitrary: m) (auto simp: iterate_def) lemma iterate_applied: "iterate n m f (f x) = iterate (Suc n) (Suc m) f x" by (induct m arbitrary: n) (auto simp: iterate_def funpow_swap1) end diff --git a/thys/Graph_Theory/ROOT b/thys/Graph_Theory/ROOT --- a/thys/Graph_Theory/ROOT +++ b/thys/Graph_Theory/ROOT @@ -1,9 +1,11 @@ chapter AFP session "Graph_Theory" (AFP) = "HOL-Library" + options [timeout = 600] + sessions + "HOL-Combinatorics" theories Graph_Theory document_files "root.tex" "root.bib" diff --git a/thys/Groebner_Macaulay/Hilbert_Function.thy b/thys/Groebner_Macaulay/Hilbert_Function.thy --- a/thys/Groebner_Macaulay/Hilbert_Function.thy +++ b/thys/Groebner_Macaulay/Hilbert_Function.thy @@ -1,1388 +1,1391 @@ (* Author: Alexander Maletzky *) section \Direct Decompositions and Hilbert Functions\ theory Hilbert_Function - imports Dube_Prelims Degree_Section "HOL-Library.Permutations" +imports + "HOL-Combinatorics.Permutations" + Dube_Prelims + Degree_Section begin subsection \Direct Decompositions\ text \The main reason for defining \direct_decomp\ in terms of lists rather than sets is that lemma \direct_decomp_direct_decomp\ can be proved easier. At some point one could invest the time to re-define \direct_decomp\ in terms of sets (possibly adding a couple of further assumptions to \direct_decomp_direct_decomp\).\ definition direct_decomp :: "'a set \ 'a::comm_monoid_add set list \ bool" where "direct_decomp A ss \ bij_betw sum_list (listset ss) A" lemma direct_decompI: "inj_on sum_list (listset ss) \ sum_list ` listset ss = A \ direct_decomp A ss" by (simp add: direct_decomp_def bij_betw_def) lemma direct_decompI_alt: "(\qs. qs \ listset ss \ sum_list qs \ A) \ (\a. a \ A \ \!qs\listset ss. a = sum_list qs) \ direct_decomp A ss" by (auto simp: direct_decomp_def intro!: bij_betwI') blast lemma direct_decompD: assumes "direct_decomp A ss" shows "qs \ listset ss \ sum_list qs \ A" and "inj_on sum_list (listset ss)" and "sum_list ` listset ss = A" using assms by (auto simp: direct_decomp_def bij_betw_def) lemma direct_decompE: assumes "direct_decomp A ss" and "a \ A" obtains qs where "qs \ listset ss" and "a = sum_list qs" using assms by (auto simp: direct_decomp_def bij_betw_def) lemma direct_decomp_unique: "direct_decomp A ss \ qs \ listset ss \ qs' \ listset ss \ sum_list qs = sum_list qs' \ qs = qs'" by (auto dest: direct_decompD simp: inj_on_def) lemma direct_decomp_singleton: "direct_decomp A [A]" proof (rule direct_decompI_alt) fix qs assume "qs \ listset [A]" then obtain q where "q \ A" and "qs = [q]" by (rule listset_singletonE) thus "sum_list qs \ A" by simp next fix a assume "a \ A" show "\!qs\listset [A]. a = sum_list qs" proof (intro ex1I conjI allI impI) from \a \ A\ refl show "[a] \ listset [A]" by (rule listset_singletonI) next fix qs assume "qs \ listset [A] \ a = sum_list qs" hence a: "a = sum_list qs" and "qs \ listset [A]" by simp_all from this(2) obtain b where qs: "qs = [b]" by (rule listset_singletonE) with a show "qs = [a]" by simp qed simp_all qed (* TODO: Move. *) lemma mset_bij: assumes "bij_betw f {..i. i < length xs \ xs ! i = ys ! f i" shows "mset xs = mset ys" proof - from assms(1) have 1: "inj_on f {0..i < length xs\ have "\ = map ((!) ys \ f) [0.. f) [0.. = image_mset ((!) ys) (image_mset f (mset_set {0.. = image_mset ((!) ys) (mset_set {0.. = mset (map ((!) ys) [0..f permutes {.. \permute_list f ss2 = ss1\ by (rule mset_eq_permutation) then have f_bij: "bij_betw f {..i. i < length ss2 \ ss1 ! i = ss2 ! f i" by (auto simp add: permutes_imp_bij permute_list_nth) define g where "g = inv_into {.. f ` {.. {.. {.. g ` {.. = {.. listset ss2" then obtain qs1 where qs1_in: "qs1 \ listset ss1" and len_qs1: "length qs1 = length qs2" and *: "\i. i < length qs2 \ qs1 ! i = qs2 ! f i" using f_bij f by (rule listset_permE) blast+ from \qs2 \ listset ss2\ have "length qs2 = length ss2" by (rule listsetD) with f_bij have "bij_betw f {.. \ A" by (rule direct_decompD) finally show "sum_list qs2 \ A" . next fix a assume "a \ A" with assms(1) obtain qs where a: "a = sum_list qs" and qs_in: "qs \ listset ss1" by (rule direct_decompE) from qs_in obtain qs2 where qs2_in: "qs2 \ listset ss2" and len_qs2: "length qs2 = length qs" and 1: "\i. i < length qs \ qs2 ! i = qs ! g i" using g_bij g by (rule listset_permE) blast+ show "\!qs\listset ss2. a = sum_list qs" proof (intro ex1I conjI allI impI) from qs_in have len_qs: "length qs = length ss1" by (rule listsetD) with g_bij have g_bij2: "bij_betw g {.. listset ss2 \ a = sum_list qs'" hence qs'_in: "qs' \ listset ss2" and a': "a = sum_list qs'" by simp_all from this(1) obtain qs1 where qs1_in: "qs1 \ listset ss1" and len_qs1: "length qs1 = length qs'" and 2: "\i. i < length qs' \ qs1 ! i = qs' ! f i" using f_bij f by (rule listset_permE) blast+ from \qs' \ listset ss2\ have "length qs' = length ss2" by (rule listsetD) with f_bij have "bij_betw f {.. {.. g ` {.. = {..i < length qs\ have "qs2 ! i = qs ! g i" by (rule 1) also have "\ = qs1 ! g i" by (simp only: \qs1 = qs\) also from \g i < length qs'\ have "\ = qs' ! f (g i)" by (rule 2) also from \i < length ss1\ have "\ = qs' ! i" by (simp only: f_g) finally show "qs' ! i = qs2 ! i" by (rule sym) qed qed fact qed qed lemma direct_decomp_split_map: "direct_decomp A (map f ss) \ direct_decomp A (map f (filter P ss) @ map f (filter (- P) ss))" proof (rule direct_decomp_perm) show "mset (map f ss) = mset (map f (filter P ss) @ map f (filter (- P) ss))" by simp (metis image_mset_union multiset_partition) qed lemmas direct_decomp_split = direct_decomp_split_map[where f=id, simplified] lemma direct_decomp_direct_decomp: assumes "direct_decomp A (s # ss)" and "direct_decomp s rs" shows "direct_decomp A (ss @ rs)" (is "direct_decomp A ?ss") proof (rule direct_decompI_alt) fix qs assume "qs \ listset ?ss" then obtain qs1 qs2 where qs1: "qs1 \ listset ss" and qs2: "qs2 \ listset rs" and qs: "qs = qs1 @ qs2" by (rule listset_appendE) have "sum_list qs = sum_list ((sum_list qs2) # qs1)" by (simp add: qs add.commute) also from assms(1) have "\ \ A" proof (rule direct_decompD) from assms(2) qs2 have "sum_list qs2 \ s" by (rule direct_decompD) thus "sum_list qs2 # qs1 \ listset (s # ss)" using qs1 refl by (rule listset_ConsI) qed finally show "sum_list qs \ A" . next fix a assume "a \ A" with assms(1) obtain qs1 where qs1_in: "qs1 \ listset (s # ss)" and a: "a = sum_list qs1" by (rule direct_decompE) from qs1_in obtain qs11 qs12 where "qs11 \ s" and qs12_in: "qs12 \ listset ss" and qs1: "qs1 = qs11 # qs12" by (rule listset_ConsE) from assms(2) this(1) obtain qs2 where qs2_in: "qs2 \ listset rs" and qs11: "qs11 = sum_list qs2" by (rule direct_decompE) let ?qs = "qs12 @ qs2" show "\!qs\listset ?ss. a = sum_list qs" proof (intro ex1I conjI allI impI) from qs12_in qs2_in refl show "?qs \ listset ?ss" by (rule listset_appendI) show "a = sum_list ?qs" by (simp add: a qs1 qs11 add.commute) fix qs0 assume "qs0 \ listset ?ss \ a = sum_list qs0" hence qs0_in: "qs0 \ listset ?ss" and a2: "a = sum_list qs0" by simp_all from this(1) obtain qs01 qs02 where qs01_in: "qs01 \ listset ss" and qs02_in: "qs02 \ listset rs" and qs0: "qs0 = qs01 @ qs02" by (rule listset_appendE) note assms(1) moreover from _ qs01_in refl have "(sum_list qs02) # qs01 \ listset (s # ss)" (is "?qs' \ _") proof (rule listset_ConsI) from assms(2) qs02_in show "sum_list qs02 \ s" by (rule direct_decompD) qed moreover note qs1_in moreover from a2 have "sum_list ?qs' = sum_list qs1" by (simp add: qs0 a add.commute) ultimately have "?qs' = qs11 # qs12" unfolding qs1 by (rule direct_decomp_unique) hence "qs11 = sum_list qs02" and 1: "qs01 = qs12" by simp_all from this(1) have "sum_list qs02 = sum_list qs2" by (simp only: qs11) with assms(2) qs02_in qs2_in have "qs02 = qs2" by (rule direct_decomp_unique) thus "qs0 = qs12 @ qs2" by (simp only: 1 qs0) qed qed lemma sum_list_map_times: "sum_list (map ((*) x) xs) = (x::'a::semiring_0) * sum_list xs" by (induct xs) (simp_all add: algebra_simps) lemma direct_decomp_image_times: assumes "direct_decomp (A::'a::semiring_0 set) ss" and "\a b. x * a = x * b \ x \ 0 \ a = b" shows "direct_decomp ((*) x ` A) (map ((`) ((*) x)) ss)" (is "direct_decomp ?A ?ss") proof (rule direct_decompI_alt) fix qs assume "qs \ listset ?ss" then obtain qs0 where qs0_in: "qs0 \ listset ss" and qs: "qs = map ((*) x) qs0" by (rule listset_map_imageE) have "sum_list qs = x * sum_list qs0" by (simp only: qs sum_list_map_times) moreover from assms(1) qs0_in have "sum_list qs0 \ A" by (rule direct_decompD) ultimately show "sum_list qs \ (*) x ` A" by (rule image_eqI) next fix a assume "a \ ?A" then obtain a' where "a' \ A" and a: "a = x * a'" .. from assms(1) this(1) obtain qs' where qs'_in: "qs' \ listset ss" and a': "a' = sum_list qs'" by (rule direct_decompE) define qs where "qs = map ((*) x) qs'" show "\!qs\listset ?ss. a = sum_list qs" proof (intro ex1I conjI allI impI) from qs'_in qs_def show "qs \ listset ?ss" by (rule listset_map_imageI) fix qs0 assume "qs0 \ listset ?ss \ a = sum_list qs0" hence "qs0 \ listset ?ss" and a0: "a = sum_list qs0" by simp_all from this(1) obtain qs1 where qs1_in: "qs1 \ listset ss" and qs0: "qs0 = map ((*) x) qs1" by (rule listset_map_imageE) show "qs0 = qs" proof (cases "x = 0") case True from qs1_in have "length qs1 = length ss" by (rule listsetD) moreover from qs'_in have "length qs' = length ss" by (rule listsetD) ultimately show ?thesis by (simp add: qs_def qs0 list_eq_iff_nth_eq True) next case False have "x * sum_list qs1 = a" by (simp only: a0 qs0 sum_list_map_times) also have "\ = x * sum_list qs'" by (simp only: a' a) finally have "sum_list qs1 = sum_list qs'" using False by (rule assms(2)) with assms(1) qs1_in qs'_in have "qs1 = qs'" by (rule direct_decomp_unique) thus ?thesis by (simp only: qs0 qs_def) qed qed (simp only: a a' qs_def sum_list_map_times) qed lemma direct_decomp_appendD: assumes "direct_decomp A (ss1 @ ss2)" shows "{} \ set ss2 \ direct_decomp (sum_list ` listset ss1) ss1" (is "_ \ ?thesis1") and "{} \ set ss1 \ direct_decomp (sum_list ` listset ss2) ss2" (is "_ \ ?thesis2") and "direct_decomp A [sum_list ` listset ss1, sum_list ` listset ss2]" (is "direct_decomp _ ?ss") proof - have rl: "direct_decomp (sum_list ` listset ts1) ts1" if "direct_decomp A (ts1 @ ts2)" and "{} \ set ts2" for ts1 ts2 proof (intro direct_decompI inj_onI refl) fix qs1 qs2 assume qs1: "qs1 \ listset ts1" and qs2: "qs2 \ listset ts1" assume eq: "sum_list qs1 = sum_list qs2" from that(2) have "listset ts2 \ {}" by (simp add: listset_empty_iff) then obtain qs3 where qs3: "qs3 \ listset ts2" by blast note that(1) moreover from qs1 qs3 refl have "qs1 @ qs3 \ listset (ts1 @ ts2)" by (rule listset_appendI) moreover from qs2 qs3 refl have "qs2 @ qs3 \ listset (ts1 @ ts2)" by (rule listset_appendI) moreover have "sum_list (qs1 @ qs3) = sum_list (qs2 @ qs3)" by (simp add: eq) ultimately have "qs1 @ qs3 = qs2 @ qs3" by (rule direct_decomp_unique) thus "qs1 = qs2" by simp qed { assume "{} \ set ss2" with assms show ?thesis1 by (rule rl) } { from assms have "direct_decomp A (ss2 @ ss1)" by (rule direct_decomp_perm) simp moreover assume "{} \ set ss1" ultimately show ?thesis2 by (rule rl) } show "direct_decomp A ?ss" proof (rule direct_decompI_alt) fix qs assume "qs \ listset ?ss" then obtain q1 q2 where q1: "q1 \ sum_list ` listset ss1" and q2: "q2 \ sum_list ` listset ss2" and qs: "qs = [q1, q2]" by (rule listset_doubletonE) from q1 obtain qs1 where qs1: "qs1 \ listset ss1" and q1: "q1 = sum_list qs1" .. from q2 obtain qs2 where qs2: "qs2 \ listset ss2" and q2: "q2 = sum_list qs2" .. from qs1 qs2 refl have "qs1 @ qs2 \ listset (ss1 @ ss2)" by (rule listset_appendI) with assms have "sum_list (qs1 @ qs2) \ A" by (rule direct_decompD) thus "sum_list qs \ A" by (simp add: qs q1 q2) next fix a assume "a \ A" with assms obtain qs0 where qs0_in: "qs0 \ listset (ss1 @ ss2)" and a: "a = sum_list qs0" by (rule direct_decompE) from this(1) obtain qs1 qs2 where qs1: "qs1 \ listset ss1" and qs2: "qs2 \ listset ss2" and qs0: "qs0 = qs1 @ qs2" by (rule listset_appendE) from qs1 have len_qs1: "length qs1 = length ss1" by (rule listsetD) define qs where "qs = [sum_list qs1, sum_list qs2]" show "\!qs\listset ?ss. a = sum_list qs" proof (intro ex1I conjI) from qs1 have "sum_list qs1 \ sum_list ` listset ss1" by (rule imageI) moreover from qs2 have "sum_list qs2 \ sum_list ` listset ss2" by (rule imageI) ultimately show "qs \ listset ?ss" using qs_def by (rule listset_doubletonI) fix qs' assume "qs' \ listset ?ss \ a = sum_list qs'" hence "qs' \ listset ?ss" and a': "a = sum_list qs'" by simp_all from this(1) obtain q1 q2 where q1: "q1 \ sum_list ` listset ss1" and q2: "q2 \ sum_list ` listset ss2" and qs': "qs' = [q1, q2]" by (rule listset_doubletonE) from q1 obtain qs1' where qs1': "qs1' \ listset ss1" and q1: "q1 = sum_list qs1'" .. from q2 obtain qs2' where qs2': "qs2' \ listset ss2" and q2: "q2 = sum_list qs2'" .. from qs1' have len_qs1': "length qs1' = length ss1" by (rule listsetD) note assms moreover from qs1' qs2' refl have "qs1' @ qs2' \ listset (ss1 @ ss2)" by (rule listset_appendI) moreover note qs0_in moreover have "sum_list (qs1' @ qs2') = sum_list qs0" by (simp add: a' qs' flip: a q1 q2) ultimately have "qs1' @ qs2' = qs0" by (rule direct_decomp_unique) also have "\ = qs1 @ qs2" by fact finally show "qs' = qs" by (simp add: qs_def qs' q1 q2 len_qs1 len_qs1') qed (simp add: qs_def a qs0) qed qed lemma direct_decomp_Cons_zeroI: assumes "direct_decomp A ss" shows "direct_decomp A ({0} # ss)" proof (rule direct_decompI_alt) fix qs assume "qs \ listset ({0} # ss)" then obtain q qs' where "q \ {0}" and "qs' \ listset ss" and "qs = q # qs'" by (rule listset_ConsE) from this(1, 3) have "sum_list qs = sum_list qs'" by simp also from assms \qs' \ listset ss\ have "\ \ A" by (rule direct_decompD) finally show "sum_list qs \ A" . next fix a assume "a \ A" with assms obtain qs' where qs': "qs' \ listset ss" and a: "a = sum_list qs'" by (rule direct_decompE) define qs where "qs = 0 # qs'" show "\!qs. qs \ listset ({0} # ss) \ a = sum_list qs" proof (intro ex1I conjI) from _ qs' qs_def show "qs \ listset ({0} # ss)" by (rule listset_ConsI) simp next fix qs0 assume "qs0 \ listset ({0} # ss) \ a = sum_list qs0" hence "qs0 \ listset ({0} # ss)" and a0: "a = sum_list qs0" by simp_all from this(1) obtain q0 qs0' where "q0 \ {0}" and qs0': "qs0' \ listset ss" and qs0: "qs0 = q0 # qs0'" by (rule listset_ConsE) from this(1, 3) have "sum_list qs0' = sum_list qs'" by (simp add: a0 flip: a) with assms qs0' qs' have "qs0' = qs'" by (rule direct_decomp_unique) with \q0 \ {0}\ show "qs0 = qs" by (simp add: qs_def qs0) qed (simp add: qs_def a) qed lemma direct_decomp_Cons_zeroD: assumes "direct_decomp A ({0} # ss)" shows "direct_decomp A ss" proof - have "direct_decomp {0} []" by (simp add: direct_decomp_def bij_betw_def) with assms have "direct_decomp A (ss @ [])" by (rule direct_decomp_direct_decomp) thus ?thesis by simp qed lemma direct_decomp_Cons_subsetI: assumes "direct_decomp A (s # ss)" and "\s0. s0 \ set ss \ 0 \ s0" shows "s \ A" proof fix x assume "x \ s" moreover from assms(2) have "map (\_. 0) ss \ listset ss" by (induct ss, auto simp del: listset.simps(2) intro: listset_ConsI) ultimately have "x # (map (\_. 0) ss) \ listset (s # ss)" using refl by (rule listset_ConsI) with assms(1) have "sum_list (x # (map (\_. 0) ss)) \ A" by (rule direct_decompD) thus "x \ A" by simp qed lemma direct_decomp_Int_zero: assumes "direct_decomp A ss" and "i < j" and "j < length ss" and "\s. s \ set ss \ 0 \ s" shows "ss ! i \ ss ! j = {0}" proof - from assms(2, 3) have "i < length ss" by (rule less_trans) hence i_in: "ss ! i \ set ss" by simp from assms(3) have j_in: "ss ! j \ set ss" by simp show ?thesis proof show "ss ! i \ ss ! j \ {0}" proof fix x assume "x \ ss ! i \ ss ! j" hence x_i: "x \ ss ! i" and x_j: "x \ ss ! j" by simp_all have 1: "(map (\_. 0) ss)[k := y] \ listset ss" if "k < length ss" and "y \ ss ! k" for k y using assms(4) that proof (induct ss arbitrary: k) case Nil from Nil(2) show ?case by simp next case (Cons s ss) have *: "\s'. s' \ set ss \ 0 \ s'" by (rule Cons.prems) simp show ?case proof (cases k) case k: 0 with Cons.prems(3) have "y \ s" by simp moreover from * have "map (\_. 0) ss \ listset ss" by (induct ss) (auto simp del: listset.simps(2) intro: listset_ConsI) moreover have "(map (\_. 0) (s # ss))[k := y] = y # map (\_. 0) ss" by (simp add: k) ultimately show ?thesis by (rule listset_ConsI) next case k: (Suc k') have "0 \ s" by (rule Cons.prems) simp moreover from * have "(map (\_. 0) ss)[k' := y] \ listset ss" proof (rule Cons.hyps) from Cons.prems(2) show "k' < length ss" by (simp add: k) next from Cons.prems(3) show "y \ ss ! k'" by (simp add: k) qed moreover have "(map (\_. 0) (s # ss))[k := y] = 0 # (map (\_. 0) ss)[k' := y]" by (simp add: k) ultimately show ?thesis by (rule listset_ConsI) qed qed have 2: "sum_list ((map (\_. 0) ss)[k := y]) = y" if "k < length ss" for k and y::'a using that by (induct ss arbitrary: k) (auto simp: add_ac split: nat.split) define qs1 where "qs1 = (map (\_. 0) ss)[i := x]" define qs2 where "qs2 = (map (\_. 0) ss)[j := x]" note assms(1) moreover from \i < length ss\ x_i have "qs1 \ listset ss" unfolding qs1_def by (rule 1) moreover from assms(3) x_j have "qs2 \ listset ss" unfolding qs2_def by (rule 1) thm sum_list_update moreover from \i < length ss\ assms(3) have "sum_list qs1 = sum_list qs2" by (simp add: qs1_def qs2_def 2) ultimately have "qs1 = qs2" by (rule direct_decomp_unique) hence "qs1 ! i = qs2 ! i" by simp with \i < length ss\ assms(2, 3) show "x \ {0}" by (simp add: qs1_def qs2_def) qed next from i_in have "0 \ ss ! i" by (rule assms(4)) moreover from j_in have "0 \ ss ! j" by (rule assms(4)) ultimately show "{0} \ ss ! i \ ss ! j" by simp qed qed corollary direct_decomp_pairwise_zero: assumes "direct_decomp A ss" and "\s. s \ set ss \ 0 \ s" shows "pairwise (\s1 s2. s1 \ s2 = {0}) (set ss)" proof (rule pairwiseI) fix s1 s2 assume "s1 \ set ss" then obtain i where "i < length ss" and s1: "s1 = ss ! i" by (metis in_set_conv_nth) assume "s2 \ set ss" then obtain j where "j < length ss" and s2: "s2 = ss ! j" by (metis in_set_conv_nth) assume "s1 \ s2" hence "i < j \ j < i" by (auto simp: s1 s2) thus "s1 \ s2 = {0}" proof assume "i < j" with assms(1) show ?thesis unfolding s1 s2 using \j < length ss\ assms(2) by (rule direct_decomp_Int_zero) next assume "j < i" with assms(1) have "s2 \ s1 = {0}" unfolding s1 s2 using \i < length ss\ assms(2) by (rule direct_decomp_Int_zero) thus ?thesis by (simp only: Int_commute) qed qed corollary direct_decomp_repeated_eq_zero: assumes "direct_decomp A ss" and "1 < count_list ss X" and "\s. s \ set ss \ 0 \ s" shows "X = {0}" proof - from assms(2) obtain i j where "i < j" and "j < length ss" and 1: "ss ! i = X" and 2: "ss ! j = X" by (rule count_list_gr_1_E) from assms(1) this(1, 2) assms(3) have "ss ! i \ ss ! j = {0}" by (rule direct_decomp_Int_zero) thus ?thesis by (simp add: 1 2) qed corollary direct_decomp_map_Int_zero: assumes "direct_decomp A (map f ss)" and "s1 \ set ss" and "s2 \ set ss" and "s1 \ s2" and "\s. s \ set ss \ 0 \ f s" shows "f s1 \ f s2 = {0}" proof - from assms(2) obtain i where "i < length ss" and s1: "s1 = ss ! i" by (metis in_set_conv_nth) from this(1) have i: "i < length (map f ss)" by simp from assms(3) obtain j where "j < length ss" and s2: "s2 = ss ! j" by (metis in_set_conv_nth) from this(1) have j: "j < length (map f ss)" by simp have *: "0 \ s" if "s \ set (map f ss)" for s proof - from that obtain s' where "s' \ set ss" and s: "s = f s'" unfolding set_map .. from this(1) show "0 \ s" unfolding s by (rule assms(5)) qed show ?thesis proof (rule linorder_cases) assume "i < j" with assms(1) have "(map f ss) ! i \ (map f ss) ! j = {0}" using j * by (rule direct_decomp_Int_zero) with i j show ?thesis by (simp add: s1 s2) next assume "j < i" with assms(1) have "(map f ss) ! j \ (map f ss) ! i = {0}" using i * by (rule direct_decomp_Int_zero) with i j show ?thesis by (simp add: s1 s2 Int_commute) next assume "i = j" with assms(4) show ?thesis by (simp add: s1 s2) qed qed subsection \Direct Decompositions and Vector Spaces\ definition (in vector_space) is_basis :: "'b set \ 'b set \ bool" where "is_basis V B \ (B \ V \ independent B \ V \ span B \ card B = dim V)" definition (in vector_space) some_basis :: "'b set \ 'b set" where "some_basis V = Eps (local.is_basis V)" hide_const (open) real_vector.is_basis real_vector.some_basis context vector_space begin lemma dim_empty [simp]: "dim {} = 0" using dim_span_eq_card_independent independent_empty by fastforce lemma dim_zero [simp]: "dim {0} = 0" using dim_span_eq_card_independent independent_empty by fastforce lemma independent_UnI: assumes "independent A" and "independent B" and "span A \ span B = {0}" shows "independent (A \ B)" proof from span_superset have "A \ B \ span A \ span B" by blast hence "A \ B = {}" unfolding assms(3) using assms(1, 2) dependent_zero by blast assume "dependent (A \ B)" then obtain T u v where "finite T" and "T \ A \ B" and eq: "(\v\T. u v *s v) = 0" and "v \ T" and "u v \ 0" unfolding dependent_explicit by blast define TA where "TA = T \ A" define TB where "TB = T \ B" from \T \ A \ B\ have T: "T = TA \ TB" by (auto simp: TA_def TB_def) from \finite T\ have "finite TA" and "TA \ A" by (simp_all add: TA_def) from \finite T\ have "finite TB" and "TB \ B" by (simp_all add: TB_def) from \A \ B = {}\ \TA \ A\ this(2) have "TA \ TB = {}" by blast have "0 = (\v\TA \ TB. u v *s v)" by (simp only: eq flip: T) also have "\ = (\v\TA. u v *s v) + (\v\TB. u v *s v)" by (rule sum.union_disjoint) fact+ finally have "(\v\TA. u v *s v) = (\v\TB. (- u) v *s v)" (is "?x = ?y") by (simp add: sum_negf eq_neg_iff_add_eq_0) from \finite TB\ \TB \ B\ have "?y \ span B" by (auto simp: span_explicit simp del: uminus_apply) moreover from \finite TA\ \TA \ A\ have "?x \ span A" by (auto simp: span_explicit) ultimately have "?y \ span A \ span B" by (simp add: \?x = ?y\) hence "?x = 0" and "?y = 0" by (simp_all add: \?x = ?y\ assms(3)) from \v \ T\ have "v \ TA \ TB" by (simp only: T) hence "u v = 0" proof assume "v \ TA" with assms(1) \finite TA\ \TA \ A\ \?x = 0\ show "u v = 0" by (rule independentD) next assume "v \ TB" with assms(2) \finite TB\ \TB \ B\ \?y = 0\ have "(- u) v = 0" by (rule independentD) thus "u v = 0" by simp qed with \u v \ 0\ show False .. qed lemma subspace_direct_decomp: assumes "direct_decomp A ss" and "\s. s \ set ss \ subspace s" shows "subspace A" proof (rule subspaceI) let ?qs = "map (\_. 0) ss" from assms(2) have "?qs \ listset ss" by (induct ss) (auto simp del: listset.simps(2) dest: subspace_0 intro: listset_ConsI) with assms(1) have "sum_list ?qs \ A" by (rule direct_decompD) thus "0 \ A" by simp next fix p q assume "p \ A" with assms(1) obtain ps where ps: "ps \ listset ss" and p: "p = sum_list ps" by (rule direct_decompE) assume "q \ A" with assms(1) obtain qs where qs: "qs \ listset ss" and q: "q = sum_list qs" by (rule direct_decompE) from ps qs have l: "length ps = length qs" by (simp only: listsetD) from ps qs have "map2 (+) ps qs \ listset ss" (is "?qs \ _") by (rule listset_closed_map2) (auto dest: assms(2) subspace_add) with assms(1) have "sum_list ?qs \ A" by (rule direct_decompD) thus "p + q \ A" using l by (simp only: p q sum_list_map2_plus) next fix c p assume "p \ A" with assms(1) obtain ps where "ps \ listset ss" and p: "p = sum_list ps" by (rule direct_decompE) from this(1) have "map ((*s) c) ps \ listset ss" (is "?qs \ _") by (rule listset_closed_map) (auto dest: assms(2) subspace_scale) with assms(1) have "sum_list ?qs \ A" by (rule direct_decompD) also have "sum_list ?qs = c *s sum_list ps" by (induct ps) (simp_all add: scale_right_distrib) finally show "c *s p \ A" by (simp only: p) qed lemma is_basis_alt: "subspace V \ is_basis V B \ (independent B \ span B = V)" by (metis (full_types) is_basis_def dim_eq_card span_eq span_eq_iff) lemma is_basis_finite: "is_basis V A \ is_basis V B \ finite A \ finite B" unfolding is_basis_def using independent_span_bound by auto lemma some_basis_is_basis: "is_basis V (some_basis V)" proof - obtain B where "B \ V" and "independent B" and "V \ span B" and "card B = dim V" by (rule basis_exists) hence "is_basis V B" by (simp add: is_basis_def) thus ?thesis unfolding some_basis_def by (rule someI) qed corollary shows some_basis_subset: "some_basis V \ V" and independent_some_basis: "independent (some_basis V)" and span_some_basis_supset: "V \ span (some_basis V)" and card_some_basis: "card (some_basis V) = dim V" using some_basis_is_basis[of V] by (simp_all add: is_basis_def) lemma some_basis_not_zero: "0 \ some_basis V" using independent_some_basis dependent_zero by blast lemma span_some_basis: "subspace V \ span (some_basis V) = V" by (simp add: span_subspace some_basis_subset span_some_basis_supset) lemma direct_decomp_some_basis_pairwise_disjnt: assumes "direct_decomp A ss" and "\s. s \ set ss \ subspace s" shows "pairwise (\s1 s2. disjnt (some_basis s1) (some_basis s2)) (set ss)" proof (rule pairwiseI) fix s1 s2 assume "s1 \ set ss" and "s2 \ set ss" and "s1 \ s2" have "some_basis s1 \ some_basis s2 \ s1 \ s2" using some_basis_subset by blast also from direct_decomp_pairwise_zero have "\ = {0}" proof (rule pairwiseD) fix s assume "s \ set ss" hence "subspace s" by (rule assms(2)) thus "0 \ s" by (rule subspace_0) qed fact+ finally have "some_basis s1 \ some_basis s2 \ {0}" . with some_basis_not_zero show "disjnt (some_basis s1) (some_basis s2)" unfolding disjnt_def by blast qed lemma direct_decomp_span_some_basis: assumes "direct_decomp A ss" and "\s. s \ set ss \ subspace s" shows "span (\(some_basis ` set ss)) = A" proof - from assms(1) have eq0[symmetric]: "sum_list ` listset ss = A" by (rule direct_decompD) show ?thesis unfolding eq0 using assms(2) proof (induct ss) case Nil show ?case by simp next case (Cons s ss) have "subspace s" by (rule Cons.prems) simp hence eq1: "span (some_basis s) = s" by (rule span_some_basis) have "\s'. s' \ set ss \ subspace s'" by (rule Cons.prems) simp hence eq2: "span (\ (some_basis ` set ss)) = sum_list ` listset ss" by (rule Cons.hyps) have "span (\ (some_basis ` set (s # ss))) = {x + y |x y. x \ s \ y \ sum_list ` listset ss}" by (simp add: span_Un eq1 eq2) also have "\ = sum_list ` listset (s # ss)" (is "?A = ?B") proof show "?A \ ?B" proof fix a assume "a \ ?A" then obtain x y where "x \ s" and "y \ sum_list ` listset ss" and a: "a = x + y" by blast from this(2) obtain qs where "qs \ listset ss" and y: "y = sum_list qs" .. from \x \ s\ this(1) refl have "x # qs \ listset (s # ss)" by (rule listset_ConsI) hence "sum_list (x # qs) \ ?B" by (rule imageI) also have "sum_list (x # qs) = a" by (simp add: a y) finally show "a \ ?B" . qed next show "?B \ ?A" proof fix a assume "a \ ?B" then obtain qs' where "qs' \ listset (s # ss)" and a: "a = sum_list qs'" .. from this(1) obtain x qs where "x \ s" and "qs \ listset ss" and qs': "qs' = x # qs" by (rule listset_ConsE) from this(2) have "sum_list qs \ sum_list ` listset ss" by (rule imageI) moreover have "a = x + sum_list qs" by (simp add: a qs') ultimately show "a \ ?A" using \x \ s\ by blast qed qed finally show ?case . qed qed lemma direct_decomp_independent_some_basis: assumes "direct_decomp A ss" and "\s. s \ set ss \ subspace s" shows "independent (\(some_basis ` set ss))" using assms proof (induct ss arbitrary: A) case Nil from independent_empty show ?case by simp next case (Cons s ss) have 1: "\s'. s' \ set ss \ subspace s'" by (rule Cons.prems) simp have "subspace s" by (rule Cons.prems) simp hence "0 \ s" and eq1: "span (some_basis s) = s" by (rule subspace_0, rule span_some_basis) from Cons.prems(1) have *: "direct_decomp A ([s] @ ss)" by simp moreover from \0 \ s\ have "{} \ set [s]" by auto ultimately have 2: "direct_decomp (sum_list ` listset ss) ss" by (rule direct_decomp_appendD) hence eq2: "span (\ (some_basis ` set ss)) = sum_list ` listset ss" using 1 by (rule direct_decomp_span_some_basis) note independent_some_basis[of s] moreover from 2 1 have "independent (\ (some_basis ` set ss))" by (rule Cons.hyps) moreover have "span (some_basis s) \ span (\ (some_basis ` set ss)) = {0}" proof - from * have "direct_decomp A [sum_list ` listset [s], sum_list ` listset ss]" by (rule direct_decomp_appendD) hence "direct_decomp A [s, sum_list ` listset ss]" by (simp add: image_image) moreover have "0 < (1::nat)" by simp moreover have "1 < length [s, sum_list ` listset ss]" by simp ultimately have "[s, sum_list ` listset ss] ! 0 \ [s, sum_list ` listset ss] ! 1 = {0}" by (rule direct_decomp_Int_zero) (auto simp: \0 \ s\ eq2[symmetric] span_zero) thus ?thesis by (simp add: eq1 eq2) qed ultimately have "independent (some_basis s \ (\ (some_basis ` set ss)))" by (rule independent_UnI) thus ?case by simp qed corollary direct_decomp_is_basis: assumes "direct_decomp A ss" and "\s. s \ set ss \ subspace s" shows "is_basis A (\(some_basis ` set ss))" proof - from assms have "subspace A" by (rule subspace_direct_decomp) moreover from assms have "span (\(some_basis ` set ss)) = A" by (rule direct_decomp_span_some_basis) moreover from assms have "independent (\(some_basis ` set ss))" by (rule direct_decomp_independent_some_basis) ultimately show ?thesis by (simp add: is_basis_alt) qed lemma dim_direct_decomp: assumes "direct_decomp A ss" and "finite B" and "A \ span B" and "\s. s \ set ss \ subspace s" shows "dim A = (\s\set ss. dim s)" proof - from assms(1, 4) have "is_basis A (\(some_basis ` set ss))" (is "is_basis A ?B") by (rule direct_decomp_is_basis) hence "dim A = card ?B" and "independent ?B" and "?B \ A" by (simp_all add: is_basis_def) from this(3) assms(3) have "?B \ span B" by (rule subset_trans) with assms(2) \independent ?B\ have "finite ?B" using independent_span_bound by blast note \dim A = card ?B\ also from finite_set have "card ?B = (\s\set ss. card (some_basis s))" proof (intro card_UN_disjoint ballI impI) fix s assume "s \ set ss" with \finite ?B\ show "finite (some_basis s)" by auto next fix s1 s2 have "pairwise (\s t. disjnt (some_basis s) (some_basis t)) (set ss)" using assms(1, 4) by (rule direct_decomp_some_basis_pairwise_disjnt) moreover assume "s1 \ set ss" and "s2 \ set ss" and "s1 \ s2" thm pairwiseD ultimately have "disjnt (some_basis s1) (some_basis s2)" by (rule pairwiseD) thus "some_basis s1 \ some_basis s2 = {}" by (simp only: disjnt_def) qed also from refl card_some_basis have "\ = (\s\set ss. dim s)" by (rule sum.cong) finally show ?thesis . qed end (* vector_space *) subsection \Homogeneous Sets of Polynomials with Fixed Degree\ lemma homogeneous_set_direct_decomp: assumes "direct_decomp A ss" and "\s. s \ set ss \ homogeneous_set s" shows "homogeneous_set A" proof (rule homogeneous_setI) fix a n assume "a \ A" with assms(1) obtain qs where "qs \ listset ss" and a: "a = sum_list qs" by (rule direct_decompE) have "hom_component a n = hom_component (sum_list qs) n" by (simp only: a) also have "\ = sum_list (map (\q. hom_component q n) qs)" by (induct qs) (simp_all add: hom_component_plus) also from assms(1) have "\ \ A" proof (rule direct_decompD) show "map (\q. hom_component q n) qs \ listset ss" proof (rule listset_closed_map) fix s q assume "s \ set ss" hence "homogeneous_set s" by (rule assms(2)) moreover assume "q \ s" ultimately show "hom_component q n \ s" by (rule homogeneous_setD) qed fact qed finally show "hom_component a n \ A" . qed definition hom_deg_set :: "nat \ (('x \\<^sub>0 nat) \\<^sub>0 'a) set \ (('x \\<^sub>0 nat) \\<^sub>0 'a::zero) set" where "hom_deg_set z A = (\a. hom_component a z) ` A" lemma hom_deg_setD: assumes "p \ hom_deg_set z A" shows "homogeneous p" and "p \ 0 \ poly_deg p = z" proof - from assms obtain a where "a \ A" and p: "p = hom_component a z" unfolding hom_deg_set_def .. show *: "homogeneous p" by (simp only: p homogeneous_hom_component) assume "p \ 0" hence "keys p \ {}" by simp then obtain t where "t \ keys p" by blast with * have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg) moreover from \t \ keys p\ have "deg_pm t = z" unfolding p by (rule keys_hom_componentD) ultimately show "poly_deg p = z" by simp qed lemma zero_in_hom_deg_set: assumes "0 \ A" shows "0 \ hom_deg_set z A" proof - have "0 = hom_component 0 z" by simp also from assms have "\ \ hom_deg_set z A" unfolding hom_deg_set_def by (rule imageI) finally show ?thesis . qed lemma hom_deg_set_closed_uminus: assumes "\a. a \ A \ - a \ A" and "p \ hom_deg_set z A" shows "- p \ hom_deg_set z A" proof - from assms(2) obtain a where "a \ A" and p: "p = hom_component a z" unfolding hom_deg_set_def .. from this(1) have "- a \ A" by (rule assms(1)) moreover have "- p = hom_component (- a) z" by (simp add: p) ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI) qed lemma hom_deg_set_closed_plus: assumes "\a1 a2. a1 \ A \ a2 \ A \ a1 + a2 \ A" and "p \ hom_deg_set z A" and "q \ hom_deg_set z A" shows "p + q \ hom_deg_set z A" proof - from assms(2) obtain a1 where "a1 \ A" and p: "p = hom_component a1 z" unfolding hom_deg_set_def .. from assms(3) obtain a2 where "a2 \ A" and q: "q = hom_component a2 z" unfolding hom_deg_set_def .. from \a1 \ A\ this(1) have "a1 + a2 \ A" by (rule assms(1)) moreover have "p + q = hom_component (a1 + a2) z" by (simp only: p q hom_component_plus) ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI) qed lemma hom_deg_set_closed_minus: assumes "\a1 a2. a1 \ A \ a2 \ A \ a1 - a2 \ A" and "p \ hom_deg_set z A" and "q \ hom_deg_set z A" shows "p - q \ hom_deg_set z A" proof - from assms(2) obtain a1 where "a1 \ A" and p: "p = hom_component a1 z" unfolding hom_deg_set_def .. from assms(3) obtain a2 where "a2 \ A" and q: "q = hom_component a2 z" unfolding hom_deg_set_def .. from \a1 \ A\ this(1) have "a1 - a2 \ A" by (rule assms(1)) moreover have "p - q = hom_component (a1 - a2) z" by (simp only: p q hom_component_minus) ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI) qed lemma hom_deg_set_closed_scalar: assumes "\a. a \ A \ c \ a \ A" and "p \ hom_deg_set z A" shows "(c::'a::semiring_0) \ p \ hom_deg_set z A" proof - from assms(2) obtain a where "a \ A" and p: "p = hom_component a z" unfolding hom_deg_set_def .. from this(1) have "c \ a \ A" by (rule assms(1)) moreover have "c \ p = hom_component (c \ a) z" by (simp add: p punit.map_scale_eq_monom_mult hom_component_monom_mult) ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI) qed lemma hom_deg_set_closed_sum: assumes "0 \ A" and "\a1 a2. a1 \ A \ a2 \ A \ a1 + a2 \ A" and "\i. i \ I \ f i \ hom_deg_set z A" shows "sum f I \ hom_deg_set z A" using assms(3) proof (induct I rule: infinite_finite_induct) case (infinite I) with assms(1) show ?case by (simp add: zero_in_hom_deg_set) next case empty with assms(1) show ?case by (simp add: zero_in_hom_deg_set) next case (insert j I) from insert.hyps(1, 2) have "sum f (insert j I) = f j + sum f I" by simp also from assms(2) have "\ \ hom_deg_set z A" proof (intro hom_deg_set_closed_plus insert.hyps) show "f j \ hom_deg_set z A" by (rule insert.prems) simp next fix i assume "i \ I" hence "i \ insert j I" by simp thus "f i \ hom_deg_set z A" by (rule insert.prems) qed finally show ?case . qed lemma hom_deg_set_subset: "homogeneous_set A \ hom_deg_set z A \ A" by (auto dest: homogeneous_setD simp: hom_deg_set_def) lemma Polys_closed_hom_deg_set: assumes "A \ P[X]" shows "hom_deg_set z A \ P[X]" proof fix p assume "p \ hom_deg_set z A" then obtain p' where "p' \ A" and p: "p = hom_component p' z" unfolding hom_deg_set_def .. from this(1) assms have "p' \ P[X]" .. have "keys p \ keys p'" by (simp add: p keys_hom_component) also from \p' \ P[X]\ have "\ \ .[X]" by (rule PolysD) finally show "p \ P[X]" by (rule PolysI) qed lemma hom_deg_set_alt_homogeneous_set: assumes "homogeneous_set A" shows "hom_deg_set z A = {p \ A. homogeneous p \ (p = 0 \ poly_deg p = z)}" (is "?A = ?B") proof show "?A \ ?B" proof fix h assume "h \ ?A" also from assms have "\ \ A" by (rule hom_deg_set_subset) finally show "h \ ?B" using \h \ ?A\ by (auto dest: hom_deg_setD) qed next show "?B \ ?A" proof fix h assume "h \ ?B" hence "h \ A" and "homogeneous h" and "h = 0 \ poly_deg h = z" by simp_all from this(3) show "h \ ?A" proof assume "h = 0" with \h \ A\ have "0 \ A" by simp thus ?thesis unfolding \h = 0\ by (rule zero_in_hom_deg_set) next assume "poly_deg h = z" with \homogeneous h\ have "h = hom_component h z" by (simp add: hom_component_of_homogeneous) with \h \ A\ show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI) qed qed qed lemma hom_deg_set_sum_list_listset: assumes "A = sum_list ` listset ss" shows "hom_deg_set z A = sum_list ` listset (map (hom_deg_set z) ss)" (is "?A = ?B") proof show "?A \ ?B" proof fix h assume "h \ ?A" then obtain a where "a \ A" and h: "h = hom_component a z" unfolding hom_deg_set_def .. from this(1) obtain qs where "qs \ listset ss" and a: "a = sum_list qs" unfolding assms .. have "h = hom_component (sum_list qs) z" by (simp only: a h) also have "\ = sum_list (map (\q. hom_component q z) qs)" by (induct qs) (simp_all add: hom_component_plus) also have "\ \ ?B" proof (rule imageI) show "map (\q. hom_component q z) qs \ listset (map (hom_deg_set z) ss)" unfolding hom_deg_set_def using \qs \ listset ss\ refl by (rule listset_map_imageI) qed finally show "h \ ?B" . qed next show "?B \ ?A" proof fix h assume "h \ ?B" then obtain qs where "qs \ listset (map (hom_deg_set z) ss)" and h: "h = sum_list qs" .. from this(1) obtain qs' where "qs' \ listset ss" and qs: "qs = map (\q. hom_component q z) qs'" unfolding hom_deg_set_def by (rule listset_map_imageE) have "h = sum_list (map (\q. hom_component q z) qs')" by (simp only: h qs) also have "\ = hom_component (sum_list qs') z" by (induct qs') (simp_all add: hom_component_plus) finally have "h = hom_component (sum_list qs') z" . moreover have "sum_list qs' \ A" unfolding assms using \qs' \ listset ss\ by (rule imageI) ultimately show "h \ ?A" unfolding hom_deg_set_def by (rule image_eqI) qed qed lemma direct_decomp_hom_deg_set: assumes "direct_decomp A ss" and "\s. s \ set ss \ homogeneous_set s" shows "direct_decomp (hom_deg_set z A) (map (hom_deg_set z) ss)" proof (rule direct_decompI) from assms(1) have "sum_list ` listset ss = A" by (rule direct_decompD) from this[symmetric] show "sum_list ` listset (map (hom_deg_set z) ss) = hom_deg_set z A" by (simp only: hom_deg_set_sum_list_listset) next from assms(1) have "inj_on sum_list (listset ss)" by (rule direct_decompD) moreover have "listset (map (hom_deg_set z) ss) \ listset ss" proof (rule listset_mono) fix i assume "i < length ss" hence "map (hom_deg_set z) ss ! i = hom_deg_set z (ss ! i)" by simp also from \i < length ss\ have "\ \ ss ! i" by (intro hom_deg_set_subset assms(2) nth_mem) finally show "map (hom_deg_set z) ss ! i \ ss ! i" . qed simp ultimately show "inj_on sum_list (listset (map (hom_deg_set z) ss))" by (rule inj_on_subset) qed subsection \Interpreting Polynomial Rings as Vector Spaces over the Coefficient Field\ text \There is no need to set up any further interpretation, since interpretation \phull\ is exactly what we need.\ lemma subspace_ideal: "phull.subspace (ideal (F::('b::comm_powerprod \\<^sub>0 'a::field) set))" using ideal.span_zero ideal.span_add proof (rule phull.subspaceI) fix c p assume "p \ ideal F" thus "c \ p \ ideal F" unfolding map_scale_eq_times by (rule ideal.span_scale) qed lemma subspace_Polys: "phull.subspace (P[X]::(('x \\<^sub>0 nat) \\<^sub>0 'a::field) set)" using zero_in_Polys Polys_closed_plus Polys_closed_map_scale by (rule phull.subspaceI) lemma subspace_hom_deg_set: assumes "phull.subspace A" shows "phull.subspace (hom_deg_set z A)" (is "phull.subspace ?A") proof (rule phull.subspaceI) from assms have "0 \ A" by (rule phull.subspace_0) thus "0 \ ?A" by (rule zero_in_hom_deg_set) next fix p q assume "p \ ?A" and "q \ ?A" with phull.subspace_add show "p + q \ ?A" by (rule hom_deg_set_closed_plus) (rule assms) next fix c p assume "p \ ?A" with phull.subspace_scale show "c \ p \ ?A" by (rule hom_deg_set_closed_scalar) (rule assms) qed lemma hom_deg_set_Polys_eq_span: "hom_deg_set z P[X] = phull.span (monomial (1::'a::field) ` deg_sect X z)" (is "?A = ?B") proof show "?A \ ?B" proof fix p assume "p \ ?A" also from this have "\ = {p \ P[X]. homogeneous p \ (p = 0 \ poly_deg p = z)}" by (simp only: hom_deg_set_alt_homogeneous_set[OF homogeneous_set_Polys]) finally have "p \ P[X]" and "homogeneous p" and "p \ 0 \ poly_deg p = z" by simp_all thus "p \ ?B" proof (induct p rule: poly_mapping_plus_induct) case 1 from phull.span_zero show ?case . next case (2 p c t) let ?m = "monomial c t" from 2(1) have "t \ keys ?m" by simp hence "t \ keys (?m + p)" using 2(2) by (rule in_keys_plusI1) hence "?m + p \ 0" by auto hence "poly_deg (monomial c t + p) = z" by (rule 2) from 2(4) have "keys (?m + p) \ .[X]" by (rule PolysD) with \t \ keys (?m + p)\ have "t \ .[X]" .. hence "?m \ P[X]" by (rule Polys_closed_monomial) have "t \ deg_sect X z" proof (rule deg_sectI) from 2(5) \t \ keys (?m + p)\ have "deg_pm t = poly_deg (?m + p)" by (rule homogeneousD_poly_deg) also have "\ = z" by fact finally show "deg_pm t = z" . qed fact hence "monomial 1 t \ monomial 1 ` deg_sect X z" by (rule imageI) hence "monomial 1 t \ ?B" by (rule phull.span_base) hence "c \ monomial 1 t \ ?B" by (rule phull.span_scale) hence "?m \ ?B" by simp moreover have "p \ ?B" proof (rule 2) from 2(4) \?m \ P[X]\ have "(?m + p) - ?m \ P[X]" by (rule Polys_closed_minus) thus "p \ P[X]" by simp next have 1: "deg_pm s = z" if "s \ keys p" for s proof - from that 2(2) have "s \ t" by blast hence "s \ keys ?m" by simp with that have "s \ keys (?m + p)" by (rule in_keys_plusI2) with 2(5) have "deg_pm s = poly_deg (?m + p)" by (rule homogeneousD_poly_deg) also have "\ = z" by fact finally show ?thesis . qed show "homogeneous p" by (rule homogeneousI) (simp add: 1) assume "p \ 0" show "poly_deg p = z" proof (rule antisym) show "poly_deg p \ z" by (rule poly_deg_leI) (simp add: 1) next from \p \ 0\ have "keys p \ {}" by simp then obtain s where "s \ keys p" by blast hence "z = deg_pm s" by (simp only: 1) also from \s \ keys p\ have "\ \ poly_deg p" by (rule poly_deg_max_keys) finally show "z \ poly_deg p" . qed qed ultimately show ?case by (rule phull.span_add) qed qed next show "?B \ ?A" proof fix p assume "p \ ?B" then obtain M u where "M \ monomial 1 ` deg_sect X z" and "finite M" and p: "p = (\m\M. u m \ m)" by (auto simp: phull.span_explicit) from this(1) obtain T where "T \ deg_sect X z" and M: "M = monomial 1 ` T" and inj: "inj_on (monomial (1::'a)) T" by (rule subset_imageE_inj) define c where "c = (\t. u (monomial 1 t))" from inj have "p = (\t\T. monomial (c t) t)" by (simp add: p M sum.reindex c_def) also have "\ \ ?A" proof (intro hom_deg_set_closed_sum zero_in_Polys Polys_closed_plus) fix t assume "t \ T" hence "t \ deg_sect X z" using \T \ deg_sect X z\ .. hence "t \ .[X]" and eq: "deg_pm t = z" by (rule deg_sectD)+ from this(1) have "monomial (c t) t \ P[X]" (is "?m \ _") by (rule Polys_closed_monomial) thus "?m \ ?A" by (simp add: hom_deg_set_alt_homogeneous_set[OF homogeneous_set_Polys] poly_deg_monomial monomial_0_iff eq) qed finally show "p \ ?A" . qed qed subsection \(Projective) Hilbert Function\ interpretation phull: vector_space map_scale apply standard subgoal by (fact map_scale_distrib_left) subgoal by (fact map_scale_distrib_right) subgoal by (fact map_scale_assoc) subgoal by (fact map_scale_one_left) done definition Hilbert_fun :: "(('x \\<^sub>0 nat) \\<^sub>0 'a::field) set \ nat \ nat" where "Hilbert_fun A z = phull.dim (hom_deg_set z A)" lemma Hilbert_fun_empty [simp]: "Hilbert_fun {} = 0" by (rule ext) (simp add: Hilbert_fun_def hom_deg_set_def) lemma Hilbert_fun_zero [simp]: "Hilbert_fun {0} = 0" by (rule ext) (simp add: Hilbert_fun_def hom_deg_set_def) lemma Hilbert_fun_direct_decomp: assumes "finite X" and "A \ P[X]" and "direct_decomp (A::(('x::countable \\<^sub>0 nat) \\<^sub>0 'a::field) set) ps" and "\s. s \ set ps \ homogeneous_set s" and "\s. s \ set ps \ phull.subspace s" shows "Hilbert_fun A z = (\p\set ps. Hilbert_fun p z)" proof - from assms(3, 4) have dd: "direct_decomp (hom_deg_set z A) (map (hom_deg_set z) ps)" by (rule direct_decomp_hom_deg_set) have "Hilbert_fun A z = phull.dim (hom_deg_set z A)" by (fact Hilbert_fun_def) also from dd have "\ = sum phull.dim (set (map (hom_deg_set z) ps))" proof (rule phull.dim_direct_decomp) from assms(1) have "finite (deg_sect X z)" by (rule finite_deg_sect) thus "finite (monomial (1::'a) ` deg_sect X z)" by (rule finite_imageI) next from assms(2) have "hom_deg_set z A \ hom_deg_set z P[X]" unfolding hom_deg_set_def by (rule image_mono) thus "hom_deg_set z A \ phull.span (monomial 1 ` deg_sect X z)" by (simp only: hom_deg_set_Polys_eq_span) next fix s assume "s \ set (map (hom_deg_set z) ps)" then obtain s' where "s' \ set ps" and s: "s = hom_deg_set z s'" unfolding set_map .. from this(1) have "phull.subspace s'" by (rule assms(5)) thus "phull.subspace s" unfolding s by (rule subspace_hom_deg_set) qed also have "\ = sum (phull.dim \ hom_deg_set z) (set ps)" unfolding set_map using finite_set proof (rule sum.reindex_nontrivial) fix s1 s2 note dd moreover assume "s1 \ set ps" and "s2 \ set ps" and "s1 \ s2" moreover have "0 \ hom_deg_set z s" if "s \ set ps" for s proof (rule zero_in_hom_deg_set) from that have "phull.subspace s" by (rule assms(5)) thus "0 \ s" by (rule phull.subspace_0) qed ultimately have "hom_deg_set z s1 \ hom_deg_set z s2 = {0}" by (rule direct_decomp_map_Int_zero) moreover assume "hom_deg_set z s1 = hom_deg_set z s2" ultimately show "phull.dim (hom_deg_set z s1) = 0" by simp qed also have "\ = (\p\set ps. Hilbert_fun p z)" by (simp only: o_def Hilbert_fun_def) finally show ?thesis . qed context pm_powerprod begin lemma image_lt_hom_deg_set: assumes "homogeneous_set A" shows "lpp ` (hom_deg_set z A - {0}) = {t \ lpp ` (A - {0}). deg_pm t = z}" (is "?B = ?A") proof (intro set_eqI iffI) fix t assume "t \ ?A" hence "t \ lpp ` (A - {0})" and deg_t[symmetric]: "deg_pm t = z" by simp_all from this(1) obtain p where "p \ A - {0}" and t: "t = lpp p" .. from this(1) have "p \ A" and "p \ 0" by simp_all from this(1) have 1: "hom_component p z \ hom_deg_set z A" (is "?p \ _") unfolding hom_deg_set_def by (rule imageI) from \p \ 0\ have "?p \ 0" and "lpp ?p = t" unfolding t deg_t by (rule hom_component_lpp)+ note this(2)[symmetric] moreover from 1 \?p \ 0\ have "?p \ hom_deg_set z A - {0}" by simp ultimately show "t \ ?B" by (rule image_eqI) next fix t assume "t \ ?B" then obtain p where "p \ hom_deg_set z A - {0}" and t: "t = lpp p" .. from this(1) have "p \ hom_deg_set z A" and "p \ 0" by simp_all with assms have "p \ A" and "homogeneous p" and "poly_deg p = z" by (simp_all add: hom_deg_set_alt_homogeneous_set) from this(1) \p \ 0\ have "p \ A - {0}" by simp hence 1: "t \ lpp ` (A - {0})" using t by (rule rev_image_eqI) from \p \ 0\ have "t \ keys p" unfolding t by (rule punit.lt_in_keys) with \homogeneous p\ have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg) with 1 show "t \ ?A" by (simp add: \poly_deg p = z\) qed lemma Hilbert_fun_alt: assumes "finite X" and "A \ P[X]" and "phull.subspace A" shows "Hilbert_fun A z = card (lpp ` (hom_deg_set z A - {0}))" (is "_ = card ?A") proof - have "?A \ lpp ` (hom_deg_set z A - {0})" by simp then obtain B where sub: "B \ hom_deg_set z A - {0}" and eq1: "?A = lpp ` B" and inj: "inj_on lpp B" by (rule subset_imageE_inj) have "Hilbert_fun A z = phull.dim (hom_deg_set z A)" by (fact Hilbert_fun_def) also have "\ = card B" proof (rule phull.dim_eq_card) show "phull.span B = phull.span (hom_deg_set z A)" proof from sub have "B \ hom_deg_set z A" by blast thus "phull.span B \ phull.span (hom_deg_set z A)" by (rule phull.span_mono) next from assms(3) have "phull.subspace (hom_deg_set z A)" by (rule subspace_hom_deg_set) hence "phull.span (hom_deg_set z A) = hom_deg_set z A" by (simp only: phull.span_eq_iff) also have "\ \ phull.span B" proof (rule ccontr) assume "\ hom_deg_set z A \ phull.span B" then obtain p0 where "p0 \ hom_deg_set z A - phull.span B" (is "_ \ ?B") by blast note assms(1) this moreover have "?B \ P[X]" proof (rule subset_trans) from assms(2) show "hom_deg_set z A \ P[X]" by (rule Polys_closed_hom_deg_set) qed blast ultimately obtain p where "p \ ?B" and p_min: "\q. punit.ord_strict_p q p \ q \ ?B" by (rule punit.ord_p_minimum_dgrad_p_set[OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]) blast from this(1) have "p \ hom_deg_set z A" and "p \ phull.span B" by simp_all from phull.span_zero this(2) have "p \ 0" by blast with \p \ hom_deg_set z A\ have "p \ hom_deg_set z A - {0}" by simp hence "lpp p \ lpp ` (hom_deg_set z A - {0})" by (rule imageI) also have "\ = lpp ` B" by (simp only: eq1) finally obtain b where "b \ B" and eq2: "lpp p = lpp b" .. from this(1) sub have "b \ hom_deg_set z A - {0}" .. hence "b \ hom_deg_set z A" and "b \ 0" by simp_all from this(2) have lcb: "punit.lc b \ 0" by (rule punit.lc_not_0) from \p \ 0\ have lcp: "punit.lc p \ 0" by (rule punit.lc_not_0) from \b \ B\ have "b \ phull.span B" by (rule phull.span_base) hence "(punit.lc p / punit.lc b) \ b \ phull.span B" (is "?b \ _") by (rule phull.span_scale) with \p \ phull.span B\ have "p - ?b \ 0" by auto moreover from lcb lcp \b \ 0\ have "lpp ?b = lpp p" by (simp add: punit.map_scale_eq_monom_mult punit.lt_monom_mult eq2) moreover from lcb have "punit.lc ?b = punit.lc p" by (simp add: punit.map_scale_eq_monom_mult) ultimately have "lpp (p - ?b) \ lpp p" by (rule punit.lt_minus_lessI) hence "punit.ord_strict_p (p - ?b) p" by (rule punit.lt_ord_p) hence "p - ?b \ ?B" by (rule p_min) hence "p - ?b \ hom_deg_set z A \ p - ?b \ phull.span B" by simp thus False proof assume *: "p - ?b \ hom_deg_set z A" from phull.subspace_scale have "?b \ hom_deg_set z A" proof (rule hom_deg_set_closed_scalar) show "phull.subspace A" by fact next show "b \ hom_deg_set z A" by fact qed with phull.subspace_diff \p \ hom_deg_set z A\ have "p - ?b \ hom_deg_set z A" by (rule hom_deg_set_closed_minus) (rule assms(3)) with * show ?thesis .. next assume "p - ?b \ phull.span B" hence "p - ?b + ?b \ phull.span B" using \?b \ phull.span B\ by (rule phull.span_add) hence "p \ phull.span B" by simp with \p \ phull.span B\ show ?thesis .. qed qed finally show "phull.span (hom_deg_set z A) \ phull.span B" . qed next show "phull.independent B" proof assume "phull.dependent B" then obtain B' u b' where "finite B'" and "B' \ B" and "(\b\B'. u b \ b) = 0" and "b' \ B'" and "u b' \ 0" unfolding phull.dependent_explicit by blast define B0 where "B0 = {b \ B'. u b \ 0}" have "B0 \ B'" by (simp add: B0_def) with \finite B'\ have "(\b\B0. u b \ b) = (\b\B'. u b \ b)" by (rule sum.mono_neutral_left) (simp add: B0_def) also have "\ = 0" by fact finally have eq: "(\b\B0. u b \ b) = 0" . define t where "t = ordered_powerprod_lin.Max (lpp ` B0)" from \b' \ B'\ \u b' \ 0\ have "b' \ B0" by (simp add: B0_def) hence "lpp b' \ lpp ` B0" by (rule imageI) hence "lpp ` B0 \ {}" by blast from \B0 \ B'\ \finite B'\ have "finite B0" by (rule finite_subset) hence "finite (lpp ` B0)" by (rule finite_imageI) hence "t \ lpp ` B0" unfolding t_def using \lpp ` B0 \ {}\ by (rule ordered_powerprod_lin.Max_in) then obtain b0 where "b0 \ B0" and t: "t = lpp b0" .. note this(1) moreover from \B0 \ B'\ \B' \ B\ have "B0 \ B" by (rule subset_trans) also have "\ \ hom_deg_set z A - {0}" by fact finally have "b0 \ hom_deg_set z A - {0}" . hence "b0 \ 0" by simp hence "t \ keys b0" unfolding t by (rule punit.lt_in_keys) have "lookup (\b\B0. u b \ b) t = (\b\B0. u b * lookup b t)" by (simp add: lookup_sum) also from \finite B0\ have "\ = (\b\{b0}. u b * lookup b t)" proof (rule sum.mono_neutral_right) from \b0 \ B0\ show "{b0} \ B0" by simp next show "\b\B0 - {b0}. u b * lookup b t = 0" proof fix b assume "b \ B0 - {b0}" hence "b \ B0" and "b \ b0" by simp_all from this(1) have "lpp b \ lpp ` B0" by (rule imageI) with \finite (lpp ` B0)\ have "lpp b \ t" unfolding t_def by (rule ordered_powerprod_lin.Max_ge) have "t \ keys b" proof assume "t \ keys b" hence "t \ lpp b" by (rule punit.lt_max_keys) with \lpp b \ t\ have "lpp b = lpp b0" unfolding t by simp from inj \B0 \ B\ have "inj_on lpp B0" by (rule inj_on_subset) hence "b = b0" using \lpp b = lpp b0\ \b \ B0\ \b0 \ B0\ by (rule inj_onD) with \b \ b0\ show False .. qed thus "u b * lookup b t = 0" by (simp add: in_keys_iff) qed qed also from \t \ keys b0\ \b0 \ B0\ have "\ \ 0" by (simp add: B0_def in_keys_iff) finally show False by (simp add: eq) qed qed also have "\ = card ?A" unfolding eq1 using inj by (rule card_image[symmetric]) finally show ?thesis . qed end (* pm_powerprod *) end (* theory *) diff --git a/thys/Hermite_Lindemann/Misc_HLW.thy b/thys/Hermite_Lindemann/Misc_HLW.thy --- a/thys/Hermite_Lindemann/Misc_HLW.thy +++ b/thys/Hermite_Lindemann/Misc_HLW.thy @@ -1,231 +1,231 @@ (* File: Misc_HLW.thy Author: Manuel Eberl, TU München *) section \Miscellaneous facts\ theory Misc_HLW imports Complex_Main "HOL-Library.Multiset" - "HOL-Library.Permutations" "HOL-Library.FuncSet" "HOL-Library.Groups_Big_Fun" "HOL-Library.Poly_Mapping" "HOL-Library.Landau_Symbols" + "HOL-Combinatorics.Permutations" "HOL-Computational_Algebra.Computational_Algebra" begin lemma set_mset_subset_singletonD: assumes "set_mset A \ {x}" shows "A = replicate_mset (size A) x" using assms by (induction A) auto lemma image_mset_eq_replicate_msetD: assumes "image_mset f A = replicate_mset n y" shows "\x\#A. f x = y" proof - have "f ` set_mset A = set_mset (image_mset f A)" by simp also note assms finally show ?thesis by (auto split: if_splits) qed lemma bij_betw_permutes_compose_left: assumes "\ permutes A" shows "bij_betw (\\. \ \ \) {\. \ permutes A} {\. \ permutes A}" proof (rule bij_betwI) show "(\) \ \ {\. \ permutes A} \ {\. \ permutes A}" by (auto intro: permutes_compose assms) show "(\) (inv_into UNIV \) \ {\. \ permutes A} \ {\. \ permutes A}" by (auto intro: permutes_compose assms permutes_inv) qed (use permutes_inverses[OF assms] in auto) lemma bij_betw_compose_left_perm_Pi: assumes "\ permutes B" shows "bij_betw (\f. (\ \ f)) (A \ B) (A \ B)" proof (rule bij_betwI) have *: "(\f. (\ \ f)) \ (A \ B) \ A \ B" if \: "\ permutes B" for \ by (auto simp: permutes_in_image[OF \]) show "(\f. (\ \ f)) \ (A \ B) \ A \ B" by (rule *) fact show "(\f. (inv_into UNIV \ \ f)) \ (A \ B) \ A \ B" by (intro * permutes_inv) fact qed (auto simp: permutes_inverses[OF assms] fun_eq_iff) lemma bij_betw_compose_left_perm_PiE: assumes "\ permutes B" shows "bij_betw (\f. restrict (\ \ f) A) (A \\<^sub>E B) (A \\<^sub>E B)" proof (rule bij_betwI) have *: "(\f. restrict (\ \ f) A) \ (A \\<^sub>E B) \ A \\<^sub>E B" if \: "\ permutes B" for \ by (auto simp: permutes_in_image[OF \]) show "(\f. restrict (\ \ f) A) \ (A \\<^sub>E B) \ A \\<^sub>E B" by (rule *) fact show "(\f. restrict (inv_into UNIV \ \ f) A) \ (A \\<^sub>E B) \ A \\<^sub>E B" by (intro * permutes_inv) fact qed (auto simp: permutes_inverses[OF assms] fun_eq_iff) lemma bij_betw_image_mset_set: assumes "bij_betw f A B" shows "image_mset f (mset_set A) = mset_set B" using assms by (simp add: bij_betw_def image_mset_mset_set) lemma finite_multisets_of_size: assumes "finite A" shows "finite {X. set_mset X \ A \ size X = n}" proof (rule finite_subset) show "{X. set_mset X \ A \ size X = n} \ mset ` {xs. set xs \ A \ length xs = n}" proof fix X assume X: "X \ {X. set_mset X \ A \ size X = n}" obtain xs where [simp]: "X = mset xs" by (metis ex_mset) thus "X \ mset ` {xs. set xs \ A \ length xs = n}" using X by auto qed next show "finite (mset ` {xs. set xs \ A \ length xs = n})" by (intro finite_imageI finite_lists_length_eq assms) qed lemma sum_mset_image_mset_sum_mset_image_mset: "sum_mset (image_mset g (sum_mset (image_mset f A))) = sum_mset (image_mset (\x. sum_mset (image_mset g (f x))) A)" by (induction A) auto lemma sum_mset_image_mset_singleton: "sum_mset (image_mset (\x. {#f x#}) A) = image_mset f A" by (induction A) auto lemma sum_mset_conv_sum: "sum_mset (image_mset f A) = (\x\set_mset A. of_nat (count A x) * f x)" proof (induction A rule: full_multiset_induct) case (less A) show ?case proof (cases "A = {#}") case False then obtain x where x: "x \# A" by auto define n where "n = count A x" define A' where "A' = filter_mset (\y. y \ x) A" have A_eq: "A = replicate_mset n x + A'" by (intro multiset_eqI) (auto simp: A'_def n_def) have [simp]: "x \# A'" "count A' x = 0" by (auto simp: A'_def) have "n \ 0" using x by (auto simp: n_def) have "sum_mset (image_mset f A) = of_nat n * f x + sum_mset (image_mset f A')" by (simp add: A_eq) also have "A' \# A" unfolding A'_def using x by (simp add: filter_mset_eq_conv subset_mset_def) with less.IH have "sum_mset (image_mset f A') = (\x\set_mset A'. of_nat (count A' x) * f x)" by simp also have "\ = (\x\set_mset A'. of_nat (count A x) * f x)" by (intro sum.cong) (auto simp: A_eq) also have "of_nat n * f x + \ = (\x\insert x (set_mset A'). of_nat (count A x) * f x)" by (subst sum.insert) (auto simp: A_eq) also from \n \ 0\ have "insert x (set_mset A') = set_mset A" by (auto simp: A_eq) finally show ?thesis . qed auto qed lemma sum_mset_conv_Sum_any: "sum_mset (image_mset f A) = Sum_any (\x. of_nat (count A x) * f x)" proof - have "sum_mset (image_mset f A) = (\x\set_mset A. of_nat (count A x) * f x)" by (rule sum_mset_conv_sum) also have "\ = Sum_any (\x. of_nat (count A x) * f x)" proof (rule Sum_any.expand_superset [symmetric]) show "{x. of_nat (count A x) * f x \ 0} \ set_mset A" proof fix x assume "x \ {x. of_nat (count A x) * f x \ 0}" hence "count A x \ 0" by (intro notI) auto thus "x \# A" by auto qed qed auto finally show ?thesis . qed lemma Sum_any_sum_swap: assumes "finite A" "\y. finite {x. f x y \ 0}" shows "Sum_any (\x. sum (f x) A) = (\y\A. Sum_any (\x. f x y))" proof - have "Sum_any (\x. sum (f x) A) = Sum_any (\x. Sum_any (\y. f x y when y \ A))" unfolding when_def by (subst Sum_any.conditionalize) (use assms in simp_all) also have "\ = Sum_any (\y. Sum_any (\x. f x y when y \ A))" by (intro Sum_any.swap[of "(\y\A. {x. f x y \ 0}) \ A"] finite_SigmaI finite_UN_I assms) auto also have "(\y. Sum_any (\x. f x y when y \ A)) = (\y. Sum_any (\x. f x y) when y \ A)" by (auto simp: when_def) also have "Sum_any \ = (\y\A. Sum_any (\x. f x y))" unfolding when_def by (subst Sum_any.conditionalize) (use assms in simp_all) finally show ?thesis . qed lemma (in landau_pair) big_power: assumes "f \ L F g" shows "(\x. f x ^ n) \ L F (\x. g x ^ n)" using big_prod[of "{.._. f" F "\_. g"] assms by simp lemma (in landau_pair) small_power: assumes "f \ l F g" "n > 0" shows "(\x. f x ^ n) \ l F (\x. g x ^ n)" using assms(2,1) by (induction rule: nat_induct_non_zero) (auto intro!: small.mult) lemma pairwise_imp_disjoint_family_on: assumes "pairwise R A" assumes "\m n. m \ A \ n \ A \ R m n \ f m \ f n = {}" shows "disjoint_family_on f A" using assms unfolding disjoint_family_on_def pairwise_def by blast lemma (in comm_monoid_set) If_eq: assumes "y \ A" "finite A" shows "F (\x. g x (if x = y then h1 x else h2 x)) A = f (g y (h1 y)) (F (\x. g x (h2 x)) (A-{y}))" proof - have "F (\x. g x (if x = y then h1 x else h2 x)) A = f (g y (h1 y)) (F (\x. g x (if x = y then h1 x else h2 x)) (A-{y}))" using assms by (subst remove[of _ y]) auto also have "F (\x. g x (if x = y then h1 x else h2 x)) (A-{y}) = F (\x. g x (h2 x)) (A-{y})" by (intro cong) auto finally show ?thesis by simp qed lemma prod_nonzeroI: fixes f :: "'a \ 'b :: {semiring_no_zero_divisors, comm_semiring_1}" assumes "\x. x \ A \ f x \ 0" shows "prod f A \ 0" using assms by (induction rule: infinite_finite_induct) auto lemma frequently_prime_cofinite: "frequently (prime :: nat \ bool) cofinite" unfolding INFM_nat_le by (meson bigger_prime less_imp_le) lemma frequently_eventually_mono: assumes "frequently Q F" "eventually P F" "\x. P x \ Q x \ R x" shows "frequently R F" proof (rule frequently_mp[OF _ assms(1)]) show "eventually (\x. Q x \ R x) F" using assms(2) by eventually_elim (use assms(3) in blast) qed lemma bij_betw_Diff: assumes "bij_betw f A B" "bij_betw f A' B'" "A' \ A" "B' \ B" shows "bij_betw f (A - A') (B - B')" unfolding bij_betw_def proof have "inj_on f A" using assms(1) by (auto simp: bij_betw_def) thus "inj_on f (A - A')" by (rule inj_on_subset) auto have "f ` (A - A') = f ` A - f ` A'" by (intro inj_on_image_set_diff[OF \inj_on f A\]) (use \A' \ A\ in auto) also have "\ = B - B'" using assms(1,2) by (auto simp: bij_betw_def) finally show "f` (A - A') = B - B'" . qed lemma bij_betw_singleton: "bij_betw f {x} {y} \ f x = y" by (auto simp: bij_betw_def) end \ No newline at end of file diff --git a/thys/Hermite_Lindemann/ROOT b/thys/Hermite_Lindemann/ROOT --- a/thys/Hermite_Lindemann/ROOT +++ b/thys/Hermite_Lindemann/ROOT @@ -1,14 +1,15 @@ chapter AFP session Hermite_Lindemann (AFP) = "Pi_Transcendental" + options [timeout = 1200] sessions "HOL-Library" + "HOL-Combinatorics" Algebraic_Numbers Power_Sum_Polynomials theories Hermite_Lindemann document_files "root.tex" "root.bib" diff --git a/thys/Jordan_Normal_Form/Missing_Permutations.thy b/thys/Jordan_Normal_Form/Missing_Permutations.thy --- a/thys/Jordan_Normal_Form/Missing_Permutations.thy +++ b/thys/Jordan_Normal_Form/Missing_Permutations.thy @@ -1,463 +1,463 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Missing Permutations\ text \This theory provides some definitions and lemmas on permutations which we did not find in the Isabelle distribution.\ theory Missing_Permutations imports Missing_Ring - "HOL-Library.Permutations" + "HOL-Combinatorics.Permutations" begin definition signof :: "(nat \ nat) \ 'a :: ring_1" where "signof p = (if sign p = 1 then 1 else - 1)" lemma signof_id[simp]: "signof id = 1" "signof (\ x. x) = 1" unfolding signof_def sign_id id_def[symmetric] by auto lemma signof_inv: "finite S \ p permutes S \ signof (Hilbert_Choice.inv p) = signof p" unfolding signof_def using sign_inverse permutation_permutes by metis lemma signof_pm_one: "signof p \ {1, - 1}" unfolding signof_def by auto lemma signof_compose: assumes "p permutes {0..<(n :: nat)}" and "q permutes {0 ..<(m :: nat)}" shows "signof (p o q) = signof p * signof q" proof - from assms have pp: "permutation p" "permutation q" by (auto simp: permutation_permutes) show "signof (p o q) = signof p * signof q" unfolding signof_def sign_compose[OF pp] by (auto simp: sign_def split: if_splits) qed lemma permutes_funcset: "p permutes A \ (p ` A \ B) = (A \ B)" by (simp add: permutes_image) context comm_monoid begin lemma finprod_permute: assumes p: "p permutes S" and f: "f \ S \ carrier G" shows "finprod G f S = finprod G (f \ p) S" proof - from \p permutes S\ have "inj p" by (rule permutes_inj) then have "inj_on p S" by (auto intro: subset_inj_on) from finprod_reindex[OF _ this, unfolded permutes_image[OF p], OF f] show ?thesis unfolding o_def . qed lemma finprod_singleton_set[simp]: assumes "f a \ carrier G" shows "finprod G f {a} = f a" proof - have "finprod G f {a} = f a \ finprod G f {}" by (rule finprod_insert, insert assms, auto) also have "\ = f a" using assms by auto finally show ?thesis . qed end lemmas (in semiring) finsum_permute = add.finprod_permute lemmas (in semiring) finsum_singleton_set = add.finprod_singleton_set lemma permutes_less[simp]: assumes p: "p permutes {0..<(n :: nat)}" shows "i < n \ p i < n" "i < n \ Hilbert_Choice.inv p i < n" "p (Hilbert_Choice.inv p i) = i" "Hilbert_Choice.inv p (p i) = i" proof - assume i: "i < n" show "p i < n" using permutes_in_image[OF p] i by auto let ?inv = "Hilbert_Choice.inv p" have "\n. ?inv (p n) = n" using permutes_inverses[OF p] by simp thus "?inv i < n" by (metis (no_types) atLeastLessThan_iff f_inv_into_f inv_into_into le0 permutes_image[OF p] i) qed (insert permutes_inverses[OF p], auto) context cring begin lemma finsum_permutations_inverse: assumes f: "f \ {p. p permutes S} \ carrier R" shows "finsum R f {p. p permutes S} = finsum R (\p. f(Hilbert_Choice.inv p)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?inv = "Hilbert_Choice.inv" let ?S = "{p . p permutes S}" have th0: "inj_on ?inv ?S" proof (auto simp add: inj_on_def) fix q r assume q: "q permutes S" and r: "r permutes S" and qr: "?inv q = ?inv r" then have "?inv (?inv q) = ?inv (?inv r)" by simp with permutes_inv_inv[OF q] permutes_inv_inv[OF r] show "q = r" by metis qed have th1: "?inv ` ?S = ?S" using image_inverse_permutations by blast have th2: "?rhs = finsum R (f \ ?inv) ?S" by (simp add: o_def) from finsum_reindex[OF _ th0, of f] show ?thesis unfolding th1 th2 using f . qed lemma finsum_permutations_compose_right: assumes q: "q permutes S" and *: "f \ {p. p permutes S} \ carrier R" shows "finsum R f {p. p permutes S} = finsum R (\p. f(p \ q)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" let ?inv = "Hilbert_Choice.inv" have th0: "?rhs = finsum R (f \ (\p. p \ q)) ?S" by (simp add: o_def) have th1: "inj_on (\p. p \ q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" and r: "r permutes S" and rp: "p \ q = r \ q" then have "p \ (q \ ?inv q) = r \ (q \ ?inv q)" by (simp add: o_assoc) with permutes_surj[OF q, unfolded surj_iff] show "p = r" by simp qed have th3: "(\p. p \ q) ` ?S = ?S" using image_compose_permutations_right[OF q] by auto from finsum_reindex[OF _ th1, of f] show ?thesis unfolding th0 th1 th3 using * . qed end text \The following lemma is slightly generalized from Determinants.thy in HMA.\ lemma finite_bounded_functions: assumes fS: "finite S" shows "finite T \ finite {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = i)}" proof (induct T rule: finite_induct) case empty have th: "{f. \i. f i = i} = {id}" by auto show ?case by (auto simp add: th) next case (insert a T) let ?f = "\(y,g) i. if i = a then y else g i" let ?S = "?f ` (S \ {f. (\i\T. f i \ S) \ (\i. i \ T \ f i = i)})" have "?S = {f. (\i\ insert a T. f i \ S) \ (\i. i \ insert a T \ f i = i)}" apply (auto simp add: image_iff) apply (rule_tac x="x a" in bexI) apply (rule_tac x = "\i. if i = a then i else x i" in exI) apply (insert insert, auto) done with finite_imageI[OF finite_cartesian_product[OF fS insert.hyps(3)], of ?f] show ?case by metis qed lemma finite_bounded_functions': assumes fS: "finite S" shows "finite T \ finite {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = j)}" proof (induct T rule: finite_induct) case empty have th: "{f. \i. f i = j} = {(\ x. j)}" by auto show ?case by (auto simp add: th) next case (insert a T) let ?f = "\(y,g) i. if i = a then y else g i" let ?S = "?f ` (S \ {f. (\i\T. f i \ S) \ (\i. i \ T \ f i = j)})" have "?S = {f. (\i\ insert a T. f i \ S) \ (\i. i \ insert a T \ f i = j)}" apply (auto simp add: image_iff) apply (rule_tac x="x a" in bexI) apply (rule_tac x = "\i. if i = a then j else x i" in exI) apply (insert insert, auto) done with finite_imageI[OF finite_cartesian_product[OF fS insert.hyps(3)], of ?f] show ?case by metis qed context fixes A :: "'a set" and B :: "'b set" and a_to_b :: "'a \ 'b" and b_to_a :: "'b \ 'a" assumes ab: "\ a. a \ A \ a_to_b a \ B" and ba: "\ b. b \ B \ b_to_a b \ A" and ab_ba: "\ a. a \ A \ b_to_a (a_to_b a) = a" and ba_ab: "\ b. b \ B \ a_to_b (b_to_a b) = b" begin qualified lemma permutes_memb: fixes p :: "'b \ 'b" assumes p: "p permutes B" and a: "a \ A" defines "ip \ Hilbert_Choice.inv p" shows "a \ A" "a_to_b a \ B" "ip (a_to_b a) \ B" "p (a_to_b a) \ B" "b_to_a (p (a_to_b a)) \ A" "b_to_a (ip (a_to_b a)) \ A" proof - let ?b = "a_to_b a" from p have ip: "ip permutes B" unfolding ip_def by (rule permutes_inv) note in_ip = permutes_in_image[OF ip] note in_p = permutes_in_image[OF p] show a: "a \ A" by fact show b: "?b \ B" by (rule ab[OF a]) show pb: "p ?b \ B" unfolding in_p by (rule b) show ipb: "ip ?b \ B" unfolding in_ip by (rule b) show "b_to_a (p ?b) \ A" by (rule ba[OF pb]) show "b_to_a (ip ?b) \ A" by (rule ba[OF ipb]) qed lemma permutes_bij_main: "{p . p permutes A} \ (\ p a. if a \ A then b_to_a (p (a_to_b a)) else a) ` {p . p permutes B}" (is "?A \ ?f ` ?B") proof note d = permutes_def let ?g = "\ q b. if b \ B then a_to_b (q (b_to_a b)) else b" let ?inv = "Hilbert_Choice.inv" fix p assume p: "p \ ?f ` ?B" then obtain q where q: "q permutes B" and p: "p = ?f q" by auto let ?iq = "?inv q" from q have iq: "?iq permutes B" by (rule permutes_inv) note in_iq = permutes_in_image[OF iq] note in_q = permutes_in_image[OF q] have qiB: "\ b. b \ B \ q (?iq b) = b" using q by (rule permutes_inverses) have iqB: "\ b. b \ B \ ?iq (q b) = b" using q by (rule permutes_inverses) from q[unfolded d] have q1: "\ b. b \ B \ q b = b" and q2: "\ b. \!b'. q b' = b" by auto note memb = permutes_memb[OF q] show "p \ ?A" unfolding p d proof (rule, intro conjI impI allI, force) fix a show "\!a'. ?f q a' = a" proof (cases "a \ A") case True note a = memb[OF True] let ?a = "b_to_a (?iq (a_to_b a))" show ?thesis proof show "?f q ?a = a" using a by (simp add: ba_ab qiB ab_ba) next fix a' assume id: "?f q a' = a" show "a' = ?a" proof (cases "a' \ A") case False thus ?thesis using id a by auto next case True note a' = memb[OF this] from id True have "b_to_a (q (a_to_b a')) = a" by simp from arg_cong[OF this, of "a_to_b"] a' a have "q (a_to_b a') = a_to_b a" by (simp add: ba_ab) from arg_cong[OF this, of ?iq] have "a_to_b a' = ?iq (a_to_b a)" unfolding iqB[OF a'(2)] . from arg_cong[OF this, of b_to_a] show ?thesis unfolding ab_ba[OF True] . qed qed next case False note a = this show ?thesis proof show "?f q a = a" using a by simp next fix a' assume id: "?f q a' = a" show "a' = a" proof (cases "a' \ A") case False with id show ?thesis by simp next case True note a' = memb[OF True] with id False show ?thesis by auto qed qed qed qed qed end lemma permutes_bij': assumes ab: "\ a. a \ A \ a_to_b a \ B" and ba: "\ b. b \ B \ b_to_a b \ A" and ab_ba: "\ a. a \ A \ b_to_a (a_to_b a) = a" and ba_ab: "\ b. b \ B \ a_to_b (b_to_a b) = b" shows "{p . p permutes A} = (\ p a. if a \ A then b_to_a (p (a_to_b a)) else a) ` {p . p permutes B}" (is "?A = ?f ` ?B") proof - note one_dir = ab ba ab_ba ba_ab note other_dir = ba ab ba_ab ab_ba let ?g = "(\ p b. if b \ B then a_to_b (p (b_to_a b)) else b)" define PA where "PA = ?A" define f where "f = ?f" define g where "g = ?g" { fix p assume "p \ PA" hence p: "p permutes A" unfolding PA_def by simp from p[unfolded permutes_def] have pnA: "\ a. a \ A \ p a = a" by auto have "?f (?g p) = p" proof (rule ext) fix a show "?f (?g p) a = p a" proof (cases "a \ A") case False thus ?thesis by (simp add: pnA) next case True note a = this hence "p a \ A" unfolding permutes_in_image[OF p] . thus ?thesis using a by (simp add: ab_ba ba_ab ab) qed qed hence "f (g p) = p" unfolding f_def g_def . } hence "f ` g ` PA = PA" by force hence id: "?f ` ?g ` ?A = ?A" unfolding PA_def f_def g_def . have "?f ` ?B \ ?A" by (rule permutes_bij_main[OF one_dir]) moreover have "?g ` ?A \ ?B" by (rule permutes_bij_main[OF ba ab ba_ab ab_ba]) hence "?f ` ?g ` ?A \ ?f ` ?B" by auto hence "?A \ ?f ` ?B" unfolding id . ultimately show ?thesis by blast qed lemma inj_on_nat_permutes: assumes i: "inj_on f (S :: nat set)" and fS: "f \ S \ S" and fin: "finite S" and f: "\ i. i \ S \ f i = i" shows "f permutes S" unfolding permutes_def proof (intro conjI allI impI, rule f) fix y from endo_inj_surj[OF fin _ i] fS have fs: "f ` S = S" by auto show "\!x. f x = y" proof (cases "y \ S") case False thus ?thesis by (intro ex1I[of _ y], insert fS f, auto) next case True with fs obtain x where x: "x \ S" and fx: "f x = y" by force show ?thesis proof (rule ex1I, rule fx) fix x' assume fx': "f x' = y" with True f[of x'] have "x' \ S" by metis from inj_onD[OF i fx[folded fx'] x this] show "x' = x" by simp qed qed qed lemma permutes_pair_eq: assumes p: "p permutes S" shows "{ (p s, s) | s. s \ S } = { (s, Hilbert_Choice.inv p s) | s. s \ S }" (is "?L = ?R") proof show "?L \ ?R" proof fix x assume "x \ ?L" then obtain s where x: "x = (p s, s)" and s: "s \ S" by auto note x also have "(p s, s) = (p s, Hilbert_Choice.inv p (p s))" using permutes_inj[OF p] inv_f_f by auto also have "... \ ?R" using s permutes_in_image[OF p] by auto finally show "x \ ?R". qed show "?R \ ?L" proof fix x assume "x \ ?R" then obtain s where x: "x = (s, Hilbert_Choice.inv p s)" (is "_ = (s, ?ips)") and s: "s \ S" by auto note x also have "(s, ?ips) = (p ?ips, ?ips)" using inv_f_f[OF permutes_inj[OF permutes_inv[OF p]]] using inv_inv_eq[OF permutes_bij[OF p]] by auto also have "... \ ?L" using s permutes_in_image[OF permutes_inv[OF p]] by auto finally show "x \ ?L". qed qed lemma inj_on_finite[simp]: assumes inj: "inj_on f A" shows "finite (f ` A) = finite A" proof assume fin: "finite (f ` A)" show "finite A" proof (cases "card (f ` A) = 0") case True thus ?thesis using fin by auto next case False hence "card A > 0" unfolding card_image[OF inj] by auto thus ?thesis using card.infinite by force qed qed auto lemma permutes_prod: assumes p: "p permutes S" shows "(\s\S. f (p s) s) = (\s\S. f s (Hilbert_Choice.inv p s))" (is "?l = ?r") proof - let ?f = "\(x,y). f x y" let ?ps = "\s. (p s, s)" let ?ips = "\s. (s, Hilbert_Choice.inv p s)" have inj1: "inj_on ?ps S" by (rule inj_onI;auto) have inj2: "inj_on ?ips S" by (rule inj_onI;auto) have "?l = prod ?f (?ps ` S)" using prod.reindex[OF inj1, of ?f] by simp also have "?ps ` S = {(p s, s) |s. s \ S}" by auto also have "... = {(s, Hilbert_Choice.inv p s) | s. s \ S}" unfolding permutes_pair_eq[OF p] by simp also have "... = ?ips ` S" by auto also have "prod ?f ... = ?r" using prod.reindex[OF inj2, of ?f] by simp finally show ?thesis. qed lemma permutes_sum: assumes p: "p permutes S" shows "(\s\S. f (p s) s) = (\s\S. f s (Hilbert_Choice.inv p s))" (is "?l = ?r") proof - let ?f = "\(x,y). f x y" let ?ps = "\s. (p s, s)" let ?ips = "\s. (s, Hilbert_Choice.inv p s)" have inj1: "inj_on ?ps S" by (rule inj_onI;auto) have inj2: "inj_on ?ips S" by (rule inj_onI;auto) have "?l = sum ?f (?ps ` S)" using sum.reindex[OF inj1, of ?f] by simp also have "?ps ` S = {(p s, s) |s. s \ S}" by auto also have "... = {(s, Hilbert_Choice.inv p s) | s. s \ S}" unfolding permutes_pair_eq[OF p] by simp also have "... = ?ips ` S" by auto also have "sum ?f ... = ?r" using sum.reindex[OF inj2, of ?f] by simp finally show ?thesis. qed lemma inv_inj_on_permutes: "inj_on Hilbert_Choice.inv { p. p permutes S }" proof (intro inj_onI, unfold mem_Collect_eq) let ?i = "Hilbert_Choice.inv" fix p q assume p: "p permutes S" and q: "q permutes S" and eq: "?i p = ?i q" have "?i (?i p) = ?i (?i q)" using eq by simp thus "p = q" using inv_inv_eq[OF permutes_bij] p q by metis qed lemma permutes_others: assumes p: "p permutes S" and x: "x \ S" shows "p x = x" using p unfolding permutes_def using x by simp end \ No newline at end of file diff --git a/thys/Jordan_Normal_Form/ROOT b/thys/Jordan_Normal_Form/ROOT --- a/thys/Jordan_Normal_Form/ROOT +++ b/thys/Jordan_Normal_Form/ROOT @@ -1,21 +1,22 @@ chapter AFP session "Jordan_Normal_Form" (AFP) = "JNF-AFP-Lib" + options [timeout = 1200] sessions + "HOL-Combinatorics" Polynomial_Factorization theories Missing_Ring Missing_Permutations theories Matrix_Impl Strassen_Algorithm_Code Matrix_Complexity Jordan_Normal_Form_Existence Jordan_Normal_Form_Uniqueness Spectral_Radius theories DL_Rank_Submatrix document_files "root.bib" "root.tex" diff --git a/thys/Linear_Recurrences/Eulerian_Polynomials.thy b/thys/Linear_Recurrences/Eulerian_Polynomials.thy --- a/thys/Linear_Recurrences/Eulerian_Polynomials.thy +++ b/thys/Linear_Recurrences/Eulerian_Polynomials.thy @@ -1,127 +1,127 @@ (* File: Eulerian_Polynomials.thy Author: Manuel Eberl, TU München *) section \Eulerian polynomials\ theory Eulerian_Polynomials imports Complex_Main + "HOL-Combinatorics.Stirling" "HOL-Computational_Algebra.Computational_Algebra" - "HOL-Library.Stirling" begin text \ The Eulerian polynomials are a sequence of polynomials that is related to the closed forms of the power series \[\sum_{n=0}^\infty n^k X^n\] for a fixed $k$. \ primrec eulerian_poly :: "nat \ 'a :: idom poly" where "eulerian_poly 0 = 1" | "eulerian_poly (Suc n) = (let p = eulerian_poly n in [:0,1,-1:] * pderiv p + p * [:1, of_nat n:])" lemmas eulerian_poly_Suc [simp del] = eulerian_poly.simps(2) lemma eulerian_poly: "fps_of_poly (eulerian_poly k :: 'a :: field poly) = Abs_fps (\n. of_nat (n+1) ^ k) * (1 - fps_X) ^ (k + 1)" proof (induction k) case 0 have "Abs_fps (\_. 1 :: 'a) = inverse (1 - fps_X)" by (rule fps_inverse_unique [symmetric]) (simp add: inverse_mult_eq_1 fps_inverse_gp' [symmetric]) thus ?case by (simp add: inverse_mult_eq_1) next case (Suc k) define p :: "'a fps" where "p = fps_of_poly (eulerian_poly k)" define F :: "'a fps" where "F = Abs_fps (\n. of_nat (n+1) ^ k)" have p: "p = F * (1 - fps_X) ^ (k+1)" by (simp add: p_def Suc F_def) have p': "fps_deriv p = fps_deriv F * (1 - fps_X) ^ (k + 1) - F * (1 - fps_X) ^ k * of_nat (k + 1)" by (simp add: p fps_deriv_power algebra_simps fps_const_neg [symmetric] fps_of_nat del: power_Suc of_nat_Suc fps_const_neg) have "fps_of_poly (eulerian_poly (Suc k)) = (fps_X * fps_deriv F + F) * (1 - fps_X) ^ (Suc k + 1)" apply (simp add: Let_def p_def [symmetric] fps_of_poly_simps eulerian_poly_Suc del: power_Suc) apply (simp add: p p' fps_deriv_power fps_const_neg [symmetric] fps_of_nat del: power_Suc of_nat_Suc fps_const_neg) apply (simp add: algebra_simps) done also have "fps_X * fps_deriv F + F = Abs_fps (\n. of_nat (n + 1) ^ Suc k)" unfolding F_def by (intro fps_ext) (auto simp: algebra_simps) finally show ?case . qed lemma eulerian_poly': "Abs_fps (\n. of_nat (n+1) ^ k) = fps_of_poly (eulerian_poly k :: 'a :: field poly) / (1 - fps_X) ^ (k + 1)" by (subst eulerian_poly) simp lemma eulerian_poly'': assumes k: "k > 0" shows "Abs_fps (\n. of_nat n ^ k) = fps_of_poly (pCons 0 (eulerian_poly k :: 'a :: field poly)) / (1 - fps_X) ^ (k + 1)" proof - from assms have "Abs_fps (\n. of_nat n ^ k :: 'a) = fps_X * Abs_fps (\n. of_nat (n + 1) ^ k)" by (intro fps_ext) (auto simp: of_nat_diff) also have "Abs_fps (\n. of_nat (n + 1) ^ k :: 'a) = fps_of_poly (eulerian_poly k) / (1 - fps_X) ^ (k + 1)" by (rule eulerian_poly') also have "fps_X * \ = fps_of_poly (pCons 0 (eulerian_poly k)) / (1 - fps_X) ^ (k + 1)" by (simp add: fps_of_poly_pCons fps_divide_unit) finally show ?thesis . qed definition fps_monom_poly :: "'a :: field \ nat \ 'a poly" where "fps_monom_poly c k = (if k = 0 then 1 else pcompose (pCons 0 (eulerian_poly k)) [:0,c:])" primrec fps_monom_poly_aux :: "'a :: field \ nat \ 'a poly" where "fps_monom_poly_aux c 0 = [:c:]" | "fps_monom_poly_aux c (Suc k) = (let p = fps_monom_poly_aux c k in [:0,1,-c:] * pderiv p + [:1, of_nat k * c:] * p)" lemma fps_monom_poly_aux: "fps_monom_poly_aux c k = smult c (pcompose (eulerian_poly k) [:0,c:])" by (induction k) (simp_all add: eulerian_poly_Suc Let_def pderiv_pcompose pcompose_pCons pcompose_add pcompose_smult pcompose_uminus smult_add_right pderiv_pCons pderiv_smult algebra_simps one_pCons) lemma fps_monom_poly_code [code]: "fps_monom_poly c k = (if k = 0 then 1 else pCons 0 (fps_monom_poly_aux c k))" by (simp add: fps_monom_poly_def fps_monom_poly_aux pcompose_pCons) lemma fps_monom_aux: "Abs_fps (\n. of_nat n ^ k) = fps_of_poly (fps_monom_poly 1 k) / (1 - fps_X) ^ (k+1)" proof (cases "k = 0") assume [simp]: "k = 0" hence "Abs_fps (\n. of_nat n ^ k :: 'a) = Abs_fps (\_. 1)" by simp also have "\ = 1 / (1 - fps_X)" by (subst gp [symmetric]) simp_all finally show ?thesis by (simp add: fps_monom_poly_def) qed (insert eulerian_poly''[of k, where ?'a = 'a], simp add: fps_monom_poly_def) lemma fps_monom: "Abs_fps (\n. of_nat n ^ k * c ^ n) = fps_of_poly (fps_monom_poly c k) / (1 - fps_const c * fps_X) ^ (k+1)" proof - have "Abs_fps (\n. of_nat n ^ k * c ^ n) = fps_compose (Abs_fps (\n. of_nat n ^ k)) (fps_const c * fps_X)" by (subst fps_compose_linear) (simp add: mult_ac) also have "Abs_fps (\n. of_nat n ^ k) = fps_of_poly (fps_monom_poly 1 k) / (1 - fps_X) ^ (k+1)" by (rule fps_monom_aux) also have "fps_compose \ (fps_const c * fps_X) = (fps_of_poly (fps_monom_poly 1 k) oo fps_const c * fps_X) / ((1 - fps_X) ^ (k + 1) oo fps_const c * fps_X)" by (intro fps_compose_divide_distrib) (simp_all add: fps_compose_power [symmetric] fps_compose_sub_distrib del: power_Suc) also have "fps_of_poly (fps_monom_poly 1 k) oo (fps_const c * fps_X) = fps_of_poly (fps_monom_poly c k)" by (simp add: fps_monom_poly_def fps_of_poly_pcompose fps_of_poly_simps fps_of_poly_pCons mult_ac) also have "((1 - fps_X) ^ (k + 1) oo fps_const c * fps_X) = (1 - fps_const c * fps_X) ^ (k + 1)" by (simp add: fps_compose_power [symmetric] fps_compose_sub_distrib del: power_Suc) finally show ?thesis . qed end diff --git a/thys/Linear_Recurrences/Pochhammer_Polynomials.thy b/thys/Linear_Recurrences/Pochhammer_Polynomials.thy --- a/thys/Linear_Recurrences/Pochhammer_Polynomials.thy +++ b/thys/Linear_Recurrences/Pochhammer_Polynomials.thy @@ -1,43 +1,43 @@ (* File: Pochhammer_Polynomials.thy Author: Manuel Eberl, TU München *) section \Falling factorial as a polynomial\ theory Pochhammer_Polynomials imports Complex_Main - "HOL-Library.Stirling" + "HOL-Combinatorics.Stirling" "HOL-Computational_Algebra.Polynomial" begin definition pochhammer_poly :: "nat \ 'a :: {comm_semiring_1} poly" where "pochhammer_poly n = Poly [of_nat (stirling n k). k \ [0..iThe Inversions of a List\ theory List_Inversions - imports Main "HOL-Library.Permutations" +imports + Main + "HOL-Combinatorics.Permutations" begin subsection \Definition of inversions\ context preorder begin text \ We define inversions as pair of indices w.\,r.\,t.\ a preorder. \ inductive_set inversions :: "'a list \ (nat \ nat) set" for xs :: "'a list" where "i < j \ j < length xs \ less (xs ! j) (xs ! i) \ (i, j) \ inversions xs" lemma inversions_subset: "inversions xs \ Sigma {..i. {i<.. j < length xs \ less (xs ! j) (xs ! i)}" by (auto simp: inversions.simps) lemma inversions_code: "inversions xs = Sigma {..i. Set.filter (\j. less (xs ! j) (xs ! i)) {i<.. Suc 0 \ inversions xs = {}" by (auto simp: inversions_altdef) lemma inversions_imp_less: "z \ inversions xs \ fst z < snd z" "z \ inversions xs \ snd z < length xs" by (auto simp: inversions_altdef) lemma inversions_Nil [simp]: "inversions [] = {}" by (auto simp: inversions_altdef) lemma inversions_Cons: "inversions (x # xs) = (\j. (0, j + 1)) ` {j\{.. map_prod Suc Suc ` inversions xs" (is "_ = ?rhs") proof - have "z \ inversions (x # xs) \ z \ ?rhs" for z by (cases z) (auto simp: inversions_altdef map_prod_def nth_Cons split: nat.splits) thus ?thesis by blast qed text \ The following function returns the inversions between two lists, i.\,e.\ all pairs of an element in the first list with an element in the second list such that the former is greater than the latter. \ definition inversions_between :: "'a list \ 'a list \ (nat \ nat) set" where "inversions_between xs ys = {(i, j) \ {..{.. {.. We can now show the following equality for the inversions of the concatenation of two lists: \ proposition inversions_append: fixes xs ys defines "m \ length xs" and "n \ length ys" shows "inversions (xs @ ys) = inversions xs \ map_prod ((+) m) ((+) m) ` inversions ys \ map_prod id ((+) m) ` inversions_between xs ys" (is "_ = ?rhs") proof - note defs = inversions_altdef inversions_between_def m_def n_def map_prod_def have "z \ inversions (xs @ ys) \ z \ ?rhs" for z proof assume "z \ inversions (xs @ ys)" then obtain i j where [simp]: "z = (i, j)" and ij: "i < j" "j < m + n" "less ((xs @ ys) ! j) ((xs @ ys) ! i)" by (cases z) (auto simp: inversions_altdef m_def n_def) from ij consider "j < m" | "i \ m" | "i < m" "j \ m" by linarith thus "z \ ?rhs" proof cases assume "i < m" "j \ m" define j' where "j' = j - m" have [simp]: "j = m + j'" using \j \ m\ by (simp add: j'_def) from ij and \i < m\ show ?thesis by (auto simp: inversions_altdef map_prod_def inversions_between_def nth_append m_def n_def) next assume "i \ m" define i' j' where "i' = i - m" and "j' = j - m" have [simp]: "i = m + i'" "j = m + j'" using \i < j\ and \i \ m\ by (simp_all add: i'_def j'_def) from ij show ?thesis by (auto simp: inversions_altdef map_prod_def nth_append m_def n_def) qed (use ij in \auto simp: nth_append defs\) qed (auto simp: nth_append defs) thus ?thesis by blast qed subsection \Counting inversions\ text \ We now define versions of @{const inversions} and @{const inversions_between} that only return the \<^emph>\number\ of inversions. \ definition inversion_number :: "'a list \ nat" where "inversion_number xs = card (inversions xs)" definition inversion_number_between where "inversion_number_between xs ys = card (inversions_between xs ys)" lemma inversions_between_code: "inversions_between xs ys = Set.filter (\(i,j). less (ys ! j) (xs ! i)) ({..{.. Suc 0 \ inversion_number xs = 0" by (auto simp: inversion_number_def) lemma inversion_number_between_Nil [simp]: "inversion_number_between [] ys = 0" "inversion_number_between xs [] = 0" by (simp_all add: inversion_number_between_def) text \ We again get the following nice equation for the number of inversions of a concatenation: \ proposition inversion_number_append: "inversion_number (xs @ ys) = inversion_number xs + inversion_number ys + inversion_number_between xs ys" proof - define m n where "m = length xs" and "n = length ys" let ?A = "inversions xs" let ?B = "map_prod ((+) m) ((+) m) ` inversions ys" let ?C = "map_prod id ((+) m) ` inversions_between xs ys" have "inversion_number (xs @ ys) = card (?A \ ?B \ ?C)" by (simp add: inversion_number_def inversions_append m_def) also have "\ = card (?A \ ?B) + card ?C" by (intro card_Un_disjoint finite_inversions finite_inversions_between finite_UnI finite_imageI) (auto simp: inversions_altdef inversions_between_def m_def n_def) also have "card (?A \ ?B) = inversion_number xs + card ?B" unfolding inversion_number_def by (intro card_Un_disjoint finite_inversions finite_UnI finite_imageI) (auto simp: inversions_altdef m_def n_def) also have "card ?B = inversion_number ys" unfolding inversion_number_def by (intro card_image) (auto simp: map_prod_def inj_on_def) also have "card ?C = inversion_number_between xs ys" unfolding inversion_number_between_def by (intro card_image inj_onI) (auto simp: map_prod_def) finally show ?thesis . qed subsection \Stability of inversions between lists under permutations\ text \ A crucial fact for counting list inversions with merge sort is that the number of inversions \<^emph>\between\ two lists does not change when the lists are permuted. This is true because the set of inversions commutes with the act of permuting the list: \ lemma inversions_between_permute1: assumes "\ permutes {.. xs) ys = map_prod (inv \) id ` inversions_between xs ys" proof - from assms have [simp]: "\ i < length xs" if "i < length xs" "\ permutes {.. using permutes_in_image[OF that(2)] that by auto have *: "inv \ permutes {.. i" for i]) qed lemma inversions_between_permute2: assumes "\ permutes {.. ys) = map_prod id (inv \) ` inversions_between xs ys" proof - from assms have [simp]: "\ i < length ys" if "i < length ys" "\ permutes {.. using permutes_in_image[OF that(2)] that by auto have *: "inv \ permutes {.. i" for i]) qed proposition inversions_between_permute: assumes "\1 permutes {..2 permutes {..1 xs) (permute_list \2 ys) = map_prod (inv \1) (inv \2) ` inversions_between xs ys" by (simp add: inversions_between_permute1 inversions_between_permute2 assms map_prod_def image_image case_prod_unfold) corollary inversion_number_between_permute: assumes "\1 permutes {..2 permutes {..1 xs) (permute_list \2 ys) = inversion_number_between xs ys" proof - have "inversion_number_between (permute_list \1 xs) (permute_list \2 ys) = card (map_prod (inv \1) (inv \2) ` inversions_between xs ys)" by (simp add: inversion_number_between_def inversions_between_permute assms) also have "\ = inversion_number_between xs ys" unfolding inversion_number_between_def using assms[THEN permutes_inj_on[OF permutes_inv]] by (intro card_image inj_onI) (auto simp: map_prod_def) finally show ?thesis . qed text \ The following form of the above theorem is nicer to apply since it has the form of a congruence rule. \ corollary inversion_number_between_cong_mset: assumes "mset xs = mset xs'" and "mset ys = mset ys'" shows "inversion_number_between xs ys = inversion_number_between xs' ys'" proof - obtain \1 \2 where \12: "\1 permutes {..1 xs'" "\2 permutes {..2 ys'" using assms[THEN mset_eq_permutation] by metis thus ?thesis by (simp add: inversion_number_between_permute) qed subsection \Inversions between sorted lists\ text \ Another fact that is crucial to the efficient computation of the inversion number is this: If we have two sorted lists, we can reduce computing the inversions by inspecting the first elements and deleting one of them. \ lemma inversions_between_Cons_Cons: assumes "sorted_wrt less_eq (x # xs)" and "sorted_wrt less_eq (y # ys)" shows "inversions_between (x # xs) (y # ys) = (if \less y x then map_prod Suc id ` inversions_between xs (y # ys) else {.. {0} \ map_prod id Suc ` inversions_between (x # xs) ys)" using assms unfolding inversions_between_def map_prod_def by (auto, (auto simp: set_conv_nth nth_Cons less_le_not_le image_iff intro: order_trans split: nat.splits)?) (* A bit fragile, but doing this manually is annoying *) text \ This leads to the following analogous equation for counting the inversions between two sorted lists. Note that a single step of this only takes constant time (assuming we pre-computed the lengths of the lists) so that the entire function runs in linear time. \ lemma inversion_number_between_Cons_Cons: assumes "sorted_wrt less_eq (x # xs)" and "sorted_wrt less_eq (y # ys)" shows "inversion_number_between (x # xs) (y # ys) = (if \less y x then inversion_number_between xs (y # ys) else inversion_number_between (x # xs) ys + length (x # xs))" proof (cases "less y x") case False hence "inversion_number_between (x # xs) (y # ys) = card (map_prod Suc id ` inversions_between xs (y # ys))" by (simp add: inversion_number_between_def inversions_between_Cons_Cons[OF assms]) also have "\ = inversion_number_between xs (y # ys)" unfolding inversion_number_between_def by (intro card_image inj_onI) (auto simp: map_prod_def) finally show ?thesis using False by simp next case True hence "inversion_number_between (x # xs) (y # ys) = card ({.. {0} \ map_prod id Suc ` inversions_between (x # xs) ys)" by (simp add: inversion_number_between_def inversions_between_Cons_Cons[OF assms]) also have "\ = length (x # xs) + card (map_prod id Suc ` inversions_between (x # xs) ys)" by (subst card_Un_disjoint) auto also have "card (map_prod id Suc ` inversions_between (x # xs) ys) = inversion_number_between (x # xs) ys" unfolding inversion_number_between_def by (intro card_image inj_onI) (auto simp: map_prod_def) finally show ?thesis using True by simp qed text \ We now define a function to compute the inversion number between two lists that are assumed to be sorted using the equalities we just derived. \ fun inversion_number_between_sorted :: "'a list \ 'a list \ nat" where "inversion_number_between_sorted [] ys = 0" | "inversion_number_between_sorted xs [] = 0" | "inversion_number_between_sorted (x # xs) (y # ys) = (if \less y x then inversion_number_between_sorted xs (y # ys) else inversion_number_between_sorted (x # xs) ys + length (x # xs))" theorem inversion_number_between_sorted_correct: "sorted_wrt less_eq xs \ sorted_wrt less_eq ys \ inversion_number_between_sorted xs ys = inversion_number_between xs ys" by (induction xs ys rule: inversion_number_between_sorted.induct) (simp_all add: inversion_number_between_Cons_Cons) end subsection \Merge sort\ (* TODO: Could be replaced by mergesort from HOL-Library in Isabelle 2019. *) text \ For convenience, we first define a simple merge sort that does not compute the inversions. At this point, we need to start assuming a linear ordering since the merging function does not work otherwise. \ context linorder begin definition split_list where "split_list xs = (let n = length xs div 2 in (take n xs, drop n xs))" fun merge_lists :: "'a list \ 'a list \ 'a list" where "merge_lists [] ys = ys" | "merge_lists xs [] = xs" | "merge_lists (x # xs) (y # ys) = (if less_eq x y then x # merge_lists xs (y # ys) else y # merge_lists (x # xs) ys)" lemma set_merge_lists [simp]: "set (merge_lists xs ys) = set xs \ set ys" by (induction xs ys rule: merge_lists.induct) auto lemma mset_merge_lists [simp]: "mset (merge_lists xs ys) = mset xs + mset ys" by (induction xs ys rule: merge_lists.induct) auto lemma sorted_merge_lists [simp, intro]: "sorted xs \ sorted ys \ sorted (merge_lists xs ys)" by (induction xs ys rule: merge_lists.induct) auto fun merge_sort :: "'a list \ 'a list" where "merge_sort xs = (if length xs \ 1 then xs else merge_lists (merge_sort (take (length xs div 2) xs)) (merge_sort (drop (length xs div 2) xs)))" lemmas [simp del] = merge_sort.simps lemma merge_sort_trivial [simp]: "length xs \ Suc 0 \ merge_sort xs = xs" by (subst merge_sort.simps) auto theorem mset_merge_sort [simp]: "mset (merge_sort xs) = mset xs" by (induction xs rule: merge_sort.induct) (subst merge_sort.simps, auto simp flip: mset_append) corollary set_merge_sort [simp]: "set (merge_sort xs) = set xs" by (rule mset_eq_setD) simp_all theorem sorted_merge_sort [simp, intro]: "sorted (merge_sort xs)" by (induction xs rule: merge_sort.induct) (subst merge_sort.simps, use sorted01 in auto) lemma inversion_number_between_code: "inversion_number_between xs ys = inversion_number_between_sorted (sort xs) (sort ys)" by (subst inversion_number_between_sorted_correct) (simp_all add: sorted_sorted_wrt [symmetric] cong: inversion_number_between_cong_mset) lemmas (in -) [code_unfold] = inversion_number_between_code subsection \Merge sort with inversion counting\ text \ Finally, we can put together all the components and define a variant of merge sort that counts the number of inversions in the original list: \ function sort_and_count_inversions :: "'a list \ 'a list \ nat" where "sort_and_count_inversions xs = (if length xs \ 1 then (xs, 0) else let (xs1, xs2) = split_list xs; (xs1', m) = sort_and_count_inversions xs1; (xs2', n) = sort_and_count_inversions xs2 in (merge_lists xs1' xs2', m + n + inversion_number_between_sorted xs1' xs2'))" by auto termination by (relation "measure length") (auto simp: split_list_def Let_def) lemmas [simp del] = sort_and_count_inversions.simps text \ The projection of this function to the first component is simply the standard merge sort algorithm that we defined and proved correct before. \ theorem fst_sort_and_count_inversions [simp]: "fst (sort_and_count_inversions xs) = merge_sort xs" by (induction xs rule: length_induct) (subst sort_and_count_inversions.simps, subst merge_sort.simps, simp_all add: split_list_def case_prod_unfold Let_def) text \ The projection to the second component is the inversion number. \ theorem snd_sort_and_count_inversions [simp]: "snd (sort_and_count_inversions xs) = inversion_number xs" proof (induction xs rule: length_induct) case (1 xs) show ?case proof (cases "length xs \ 1") case False have "xs = take (length xs div 2) xs @ drop (length xs div 2) xs" by simp also have "inversion_number \ = snd (sort_and_count_inversions xs)" by (subst inversion_number_append, subst sort_and_count_inversions.simps) (use False 1 in \auto simp: Let_def split_list_def case_prod_unfold inversion_number_between_sorted_correct sorted_sorted_wrt [symmetric] cong: inversion_number_between_cong_mset\) finally show ?thesis .. qed (auto simp: sort_and_count_inversions.simps) qed lemmas (in -) [code_unfold] = snd_sort_and_count_inversions [symmetric] end end \ No newline at end of file diff --git a/thys/List_Inversions/ROOT b/thys/List_Inversions/ROOT --- a/thys/List_Inversions/ROOT +++ b/thys/List_Inversions/ROOT @@ -1,11 +1,11 @@ chapter AFP session "List_Inversions" (AFP) = HOL + options [timeout = 300] sessions - "HOL-Library" + "HOL-Combinatorics" theories List_Inversions document_files "root.tex" "root.bib" diff --git a/thys/MuchAdoAboutTwo/MuchAdoAboutTwo.thy b/thys/MuchAdoAboutTwo/MuchAdoAboutTwo.thy --- a/thys/MuchAdoAboutTwo/MuchAdoAboutTwo.thy +++ b/thys/MuchAdoAboutTwo/MuchAdoAboutTwo.thy @@ -1,2059 +1,2059 @@ (* Title: Much Ado about Two Author: Sascha Böhme , 2007 Maintainer: Sascha Böhme *) section \Much Ado about Two\ (*<*) theory MuchAdoAboutTwo -imports "HOL-Library.Permutations" +imports "HOL-Combinatorics.Permutations" begin (*>*) text \ Due to Donald E. Knuth, it is known for some time that certain sorting functions for lists of arbitrary types can be proved correct by only showing that they are correct for boolean lists (\cite{KnuthSortingAndSearching}, see also \cite{LogicalAbstractionsInHaskell}). This reduction idea, i.e. reducing a proof for an arbitrary type to a proof for a fixed type with a fixed number of values, has also instances in other fields. Recently, in \cite{MuchAdoAboutTwo}, a similar result as Knuth's 0-1-principle is explained for the problem of parallel prefix computation \cite{PrefixSumsAndTheirApplications}. That is the task to compute, for given $x_1, \ldots, x_n$ and an associative operation $\oplus$, the values $x_1$, $x_1 \oplus x_2$, $\ldots$, $x_1 \oplus x_2 \oplus \cdots \oplus x_n$. There are several solutions which optimise this computation, and an obvious question is to ask whether these solutions are correct. One way to answer this question is given in \cite{MuchAdoAboutTwo}. There, a ``0-1-2-principle'' is proved which relates an unspecified solution of the parallel prefix computation, expressed as a function \candidate\, with \scanl1\, a functional representation of the parallel prefix computation. The essence proved in the mentioned paper is as follows: If \candidate\ and \scanl1\ behave identical on all lists over a type which has three elements, then \candidate\ is semantically equivalent to \scanl1\, that is, \candidate\ is a correct solution of the parallel prefix computation. Although it seems that nearly nothing is known about the function \candidate\, it turns out that the type of \candidate\ already suffices for the proof of the paper's result. The key is relational parametricity \cite{TypesAbstractionsAndParametricPolymorphism} in the form of a free theorem \cite{TheoremsForFree}. This, some rewriting and a few properties about list-processing functions thrown in allow to proof the ``0-1-2-principle''. The paper first shows some simple properties and derives a specialisation of the free theorem. The proof of the main theorem itself is split up in two parts. The first, and considerably more complicated part relates lists over a type with three values to lists of integer lists. Here, the paper uses several figures to demonstrate and shorten several proofs. The second part then relates lists of integer list with lists over arbitrary types, and consists of applying the free theorem and some rewriting. The combination of these two parts then yields the theorem. Th article at hand formalises the proofs given in \cite{MuchAdoAboutTwo}, which is called here ``the original paper''. Compared to that paper, there are several differences in this article. The major differences are listed below. A more detailed collection follows thereafter. \begin{itemize} \item The original paper requires lists to be non-empty. Eventhough lists in Isabelle may also be empty, we stick to Isabelle's list datatype instead of declaring a new datatype, due to the huge, already existing theory about lists in Isabelle. As a consequence, however, several modifications become necessary. \item The figure-based proofs of the original paper are replaced by formal proofs. This forms a major part of this article (see Section 6). \item Instead of integers, we restrict ourselves to natural numbers. Thus, several conditions can be simplified since every natural number is greater than or equal to \0\. This decision has no further influence on the proofs because they never consider negative integers. \item Mainly due to differences between Haskell and Isabelle, certain notations are different here compared to the original paper. List concatenation is denoted by \@\ instead of $++$, and in writing down intervals, we use \[0.. instead of \[0..k]\. Moreover, we write \f\ instead of $\oplus$ and \g\ instead of $\otimes$. Functions mapping an element of the three-valued type to an arbitrary type are denoted by \h\. \end{itemize} Whenever we use lemmas from already existing Isabelle theories, we qualify them by their theory name. For example, instead of \map_map\, we write \List.map_map\ to point out that this lemma is taken from Isabelle's list theory. The following comparison shows all differences of this article compared to the original paper. The items below follow the structure of the original paper (and also this article's structure). They also highlight the challenges which needed to be solved in formalising the original paper. \begin{itemize} \item Introductions of several list functions (e.g. \length\, \map\, \take\) are dropped. They exist already in Isabelle's list theory and are be considered familiar to the reader. \item The free theorem given in Lemma 1 of the original paper is not sufficient for later proofs, because the assumption is not appropriate in the context of Isabelle's lists, which may also be empty. Thus, here, Lemma 1 is a derived version of the free theorem given as Lemma 1 in the original paper, and some additional proof-work is done. \item Before proceeding in the original paper's way, we state and proof additional lemmas, which are not part of Isabelle's libraries. These lemmas are not specific to this article and may also be used in other theories. \item Laws 1 to 8 and Lemma 2 of the original paper are explicitly proved. Most of the proofs follow directly from existing results of Isabelle's list theory. To proof Law 7, Law 8 and Lemma 2, more work was necessary, especially for Law 8. \item Lemma 3 and its proof are nearly the same here as in the original paper. Only the additional assumptions of Lemma 1, due to Isabelle's list datatype, have to be shown. \item Lemma 4 is split up in several smaller lemmas, and the order of them tries to follow the structure of the original paper's Lemma 4. For every figure of the original paper, there is now one separate proof. These proofs constitute the major difference in the structure of this article compared to the original paper. The proof of Lemma 4 in the original paper concludes by combining the results of the figure-based proofs to a non-trivial permutation property. These three sentences given in the original paper are split up in five separate lemmas and according proofs, and therefore, they as well form a major difference to the original paper. \item Lemma 5 is mostly identical to the version in the original paper. It has one additional assumption required by Lemma 4. Moreover, the proof is slightly more structured, and some steps needed a bit more argumentation than in the original paper. \item In principle, Proposition 1 is identical to the according proposition in the original paper. However, to fulfill the additional requirement of Lemma 5, an additional lemma was proved. This, however, is only necessary, because we use Isabelle's list type which allows lists to be empty. \item Proposition 2 contains one non-trivial step, which is proved as a seperate lemma. Note that this is not due to any decisions of using special datatypes, but inherent in the proof itself. Apart from that, the proof is identical to the original paper's proof of Proposition 2. \item The final theorem is, as in the original paper, just a combination of Proposition 1 and Proposition 2. Only the assumptions are extended due to Isabelle's list datatype. \end{itemize} \ section \Basic definitions\ fun foldl1 :: "('a \ 'a \ 'a) \ 'a list \ 'a" where "foldl1 f (x # xs) = foldl f x xs" fun scanl1 :: "('a \ 'a \ 'a) \ 'a list \ 'a list" where "scanl1 f xs = map (\k. foldl1 f (take k xs)) [1.. The original paper further relies on associative functions. Thus, we define another predicate to be able to express this condition: \ definition "associative f \ (\x y z. f x (f y z) = f (f x y) z)" text \ The following constant symbols represents our unspecified function. We want to show that this function is semantically equivalent to \scanl1\, provided that the first argument is an associative function. \ consts candidate :: "('a \ 'a \ 'a) \ 'a list \ 'a list" text \ With the final theorem, it suffices to show that \candidate\ behaves like \scanl1\ on all lists of the following type, to conclude that \canditate\ is semantically equivalent to \scanl1\. \ datatype three = Zero | One | Two text \ Although most of the functions mentioned in the original paper already exist in Isabelle's list theory, we still need to define two more functions: \ fun wrap :: "'a \ 'a list" where "wrap x = [x]" fun ups :: "nat \ nat list list" where "ups n = map (\k. [0..A Free Theorem\ text \ The key to proof the final theorem is the following free theorem \cite{TypesAbstractionsAndParametricPolymorphism,TheoremsForFree} of \candidate\. Since there is no proof possible without specifying the underlying (functional) language (which would be beyond the scope of this work), this lemma is expected to hold. As a consequence, all following lemmas and also the final theorem only hold under this provision. \ axiomatization where candidate_free_theorem: "\x y. h (f x y) = g (h x) (h y) \ map h (candidate f zs) = candidate g (map h zs)" text \ In what follows in this section, the previous lemma is specialised to a lemma for non-empty lists. More precisely, we want to restrict the above assumption to be applicable for non-empty lists. This is already possible without modifications when having a list datatype which does not allow for empty lists. However, before being able to also use Isabelle's list datatype, further conditions on \f\ and \zs\ are necessary. To prove the derived lemma, we first introduce a datatype for non-empty lists, and we furthermore define conversion functions to map the new datatype on Isabelle lists and back. \ datatype 'a nel = NE_One 'a | NE_Cons 'a "'a nel" fun n2l :: "'a nel \ 'a list" where "n2l (NE_One x) = [x]" | "n2l (NE_Cons x xs) = x # n2l xs" fun l2n :: "'a list \ 'a nel" where "l2n [x] = NE_One x" | "l2n (x # xs) = (case xs of [] \ NE_One x | (_ # _) \ NE_Cons x (l2n xs))" text \ The following results relate Isabelle lists and non-empty lists: \ lemma non_empty_n2l: "n2l xs \ []" by (cases xs, auto) lemma n2l_l2n_id: "x \ [] \ n2l (l2n x) = x" proof (induct x) case Nil thus ?case by simp next case (Cons x xs) thus ?case by (cases xs, auto) qed lemma n2l_l2n_map_id: assumes "\x. x \ set zs \ x \ []" shows "map (n2l \ l2n) zs = zs" using assms proof (induct zs) case Nil thus ?case by simp next case (Cons z zs) hence "\x. x \ set zs \ x \ []" using List.set_subset_Cons by auto with Cons have IH: "map (n2l \ l2n) zs = zs" by blast have "map (n2l \ l2n) (z # zs) = (n2l \ l2n) z # map (n2l \ l2n) zs" by simp also have "\ = z # map (n2l \ l2n) zs" using Cons and n2l_l2n_id by auto also have "\ = z # zs" using IH by simp finally show ?case . qed text \ Based on the previous lemmas, we can state and proof a specialised version of \candidate\'s free theorem, suitable for our setting as explained before. \ lemma Lemma_1: assumes A1: "\(x::'a list) (y::'a list). x \ [] \ y \ [] \ h (f x y) = g (h x) (h y)" and A2: "\x y. x \ [] \ y \ [] \ f x y \ []" and A3: "\x. x \ set zs \ x \ []" shows "map h (candidate f zs) = candidate g (map h zs)" proof - \ \We define two functions, @{text "fn :: 'a nel \ 'a nel \ 'a nel"} and\ \ \@{text "hn :: 'a nel \ b"}, which wrap @{text f} and @{text h} in the\ \ \setting of non-empty lists.\ let ?fn = "\x y. l2n (f (n2l x) (n2l y))" let ?hn = "h \ n2l" \ \Our new functions fulfill the preconditions of @{text candidate}'s\ \ \free theorem:\ have "\(x::'a nel) (y::'a nel). ?hn (?fn x y) = g (?hn x) (?hn y)" proof - fix x y let ?xl = "n2l (x :: 'a nel)" let ?yl = "n2l (y :: 'a nel)" have "?hn (?fn x y) = h (n2l (l2n (f (n2l x) (n2l y))))" by simp also have "\ = h (f ?xl ?yl)" using A2 [where x="?xl" and y="?yl"] and n2l_l2n_id [where x="f (n2l x) (n2l y)"] and non_empty_n2l [where xs=x] and non_empty_n2l [where xs=y] by simp also have "\ = g (h ?xl) (h ?yl)" using A1 and non_empty_n2l and non_empty_n2l by auto also have "\ = g (?hn x) (?hn y)" by simp finally show "?hn (?fn x y) = g (?hn x) (?hn y)" . qed with candidate_free_theorem [where f="?fn" and h="?hn" and g = g] have ne_free_theorem: "map ?hn (candidate ?fn (map l2n zs)) = candidate g (map ?hn (map l2n zs))" by auto \ \We use @{text candidate}'s free theorem again to show the following\ \ \property:\ have n2l_candidate: "\zs. map n2l (candidate ?fn zs) = candidate f (map n2l zs)" proof - fix zs have "\x y. n2l (?fn x y) = f (n2l x) (n2l y)" proof - fix x y show "n2l (?fn x y) = f (n2l x) (n2l y)" using n2l_l2n_id [where x="f (n2l x) (n2l y)"] and A2 [where x="n2l x" and y="n2l y"] and non_empty_n2l [where xs=x] and non_empty_n2l [where xs=y] by simp qed with candidate_free_theorem [where h=n2l and f="?fn" and g=f] show "map n2l (candidate ?fn zs) = candidate f (map n2l zs)" by simp qed \ \Now, with the previous preparations, we conclude the thesis by the\ \ \following rewriting:\ have "map h (candidate f zs) = map h (candidate f (map (n2l \ l2n) zs))" using n2l_l2n_map_id [where zs=zs] and A3 by simp also have "\ = map h (candidate f (map n2l (map l2n zs)))" using List.map_map [where f=n2l and g=l2n and xs=zs] by simp also have "\= map h (map n2l (candidate ?fn (map l2n zs)))" using n2l_candidate by auto also have "\ = map ?hn (candidate ?fn (map l2n zs))" using List.map_map by auto also have "\ = candidate g (map ?hn (map l2n zs))" using ne_free_theorem by simp also have "\ = candidate g (map ((h \ n2l) \ l2n) zs)" using List.map_map [where f="h \ n2l" and g=l2n] by simp also have "\ = candidate g (map (h \ (n2l \ l2n)) zs)" using Fun.o_assoc [symmetric, where f=h and g=n2l and h=l2n] by simp also have "\ = candidate g (map h (map (n2l \ l2n) zs))" using List.map_map [where f=h and g="n2l \ l2n"] by simp also have "\ = candidate g (map h zs)" using n2l_l2n_map_id [where zs=zs] and A3 by auto finally show ?thesis . qed section \Useful lemmas\ text \ In this section, we state and proof several lemmas, which neither occur in the original paper nor in Isabelle's libraries. \ lemma upt_map_Suc: "k > 0 \ [0..x. P [x]" and A3: "\xs ys. \ xs \ [] ; ys \ [] ; P xs ; P ys \ \ P (xs @ ys)" shows "P zs" proof (induct zs) case Nil with A1 show ?case by simp next case (Cons z zs) hence IH: "P zs" by simp show ?case proof (cases zs) case Nil with A2 show ?thesis by simp next case Cons with IH and A2 and A3 [where xs="[z]" and ys=zs] show ?thesis by auto qed qed lemmas divide_and_conquer = divide_and_conquer_induct [case_names Nil One Partition] lemma all_set_inter_empty_distinct: assumes "\xs ys. js = xs @ ys \ set xs \ set ys = {}" shows "distinct js" using assms proof (induct js rule: divide_and_conquer) case Nil thus ?case by simp next case One thus ?case by simp next case (Partition xs ys) hence P: "\as bs. xs @ ys = as @ bs \ set as \ set bs = {}" by simp have "\xs1 xs2. xs = xs1 @ xs2 \ set xs1 \ set xs2 = {}" proof - fix xs1 xs2 assume "xs = xs1 @ xs2" hence "set xs1 \ set (xs2 @ ys) = {}" using P [where as=xs1 and bs="xs2 @ ys"] by simp thus "set xs1 \ set xs2 = {}" using List.set_append and Set.Int_Un_distrib by auto qed with Partition have distinct_xs: "distinct xs" by simp have "\ys1 ys2. ys = ys1 @ ys2 \ set ys1 \ set ys2 = {}" proof - fix ys1 ys2 assume "ys = ys1 @ ys2" hence "set (xs @ ys1) \ set ys2 = {}" using P [where as="xs @ ys1" and bs=ys2] by simp thus "set ys1 \ set ys2 = {}" using List.set_append and Set.Int_Un_distrib by auto qed with Partition have distinct_ys: "distinct ys" by simp from Partition and distinct_xs and distinct_ys show ?case by simp qed lemma partitions_sorted: assumes "\xs ys x y. \ js = xs @ ys ; x \ set xs ; y \ set ys \ \ x \ y" shows "sorted js" using assms proof (induct js rule: divide_and_conquer) case Nil thus ?case by simp next case One thus ?case by simp next case (Partition xs ys) hence P: "\as bs x y. \ xs @ ys = as @ bs ; x \ set as ; y \ set bs\ \ x \ y" by simp have "\xs1 xs2 x y. \ xs = xs1 @ xs2 ; x \ set xs1 ; y \ set xs2 \ \ x \ y" proof - fix xs1 xs2 assume "xs = xs1 @ xs2" hence "\x y. \ x \ set xs1 ; y \ set (xs2 @ ys) \ \ x \ y" using P [where as=xs1 and bs="xs2 @ ys"] by simp thus "\x y. \ x \ set xs1 ; y \ set xs2 \ \ x \ y" using List.set_append by auto qed with Partition have sorted_xs: "sorted xs" by simp have "\ys1 ys2 x y. \ ys = ys1 @ ys2 ; x \ set ys1 ; y \ set ys2 \ \ x \ y" proof - fix ys1 ys2 assume "ys = ys1 @ ys2" hence "\x y. \ x \ set (xs @ ys1) ; y \ set ys2 \ \ x \ y" using P [where as="xs @ ys1" and bs=ys2] by simp thus "\x y. \ x \ set ys1 ; y \ set ys2 \ \ x \ y" using List.set_append by auto qed with Partition have sorted_ys: "sorted ys" by simp have "\x \ set xs. \y \ set ys. x \ y" using P [where as=xs and bs=ys] by simp with sorted_xs and sorted_ys show ?case using List.sorted_append by auto qed section \Preparatory Material\ text \ In the original paper, the following lemmas L1 to L8 are given without a proof, although it is hinted there that most of them follow from parametricity properties \cite{TypesAbstractionsAndParametricPolymorphism,TheoremsForFree}. Alternatively, most of them can be shown by induction over lists. However, since we are using Isabelle's list datatype, we rely on already existing results. \ lemma L1: "map g (map f xs) = map (g \ f) xs" using List.map_map by auto lemma L2: "length (map f xs) = length xs" using List.length_map by simp lemma L3: "take k (map f xs) = map f (take k xs)" using List.take_map by auto lemma L4: "map f \ wrap = wrap \ f" by (simp add: fun_eq_iff) lemma L5: "map f (xs @ ys) = (map f xs) @ (map f ys)" using List.map_append by simp lemma L6: "k < length xs \ (map f xs) ! k = f (xs ! k)" using List.nth_map by simp lemma L7: "\k. k < length xs \ map (nth xs) [0.. 0" by simp hence "map (nth (x # xs)) [0.. = ((x # xs) ! 0) # (map (nth (x # xs) \ Suc) [0.. = x # map (nth xs) [0.. = x # map (nth xs) [0.. = x # take (k' + 1) xs" using Cons and Suc by simp also have "\ = take (k + 1) (x # xs)" using Suc by simp finally show ?thesis . qed qed text \ In Isabelle's list theory, a similar result for \foldl\ already exists. Therefore, it is easy to prove the following lemma for \foldl1\. Note that this lemma does not occur in the original paper. \ lemma foldl1_append: assumes "xs \ []" shows "foldl1 f (xs @ ys) = foldl1 f (foldl1 f xs # ys)" proof - have non_empty_list: "xs \ [] \ \y ys. xs = y # ys" by (cases xs, auto) with assms obtain x xs' where x_xs_def: "xs = x # xs'" by auto have "foldl1 f (xs @ ys) = foldl f x (xs' @ ys)" using x_xs_def by simp also have "\ = foldl f (foldl f x xs') ys" using List.foldl_append by simp also have "\ = foldl f (foldl1 f (x # xs')) ys" by simp also have "\ = foldl1 f (foldl1 f xs # ys)" using x_xs_def by simp finally show ?thesis . qed text \ This is a special induction scheme suitable for proving L8. It is not mentioned in the original paper. \ lemma foldl1_induct': assumes "\f x. P f [x]" and "\f x y. P f [x, y]" and "\f x y z zs. P f (f x y # z # zs) \ P f (x # y # z # zs)" and "\f. P f []" shows "P f xs" proof (induct xs rule: List.length_induct) fix xs assume A: "\ys::'a list. length ys < length (xs::'a list) \ P f ys" thus "P f xs" proof (cases xs) case Nil with assms show ?thesis by simp next case (Cons x1 xs1) hence xs1: "xs = x1 # xs1" by simp thus ?thesis proof (cases xs1) case Nil with assms and xs1 show ?thesis by simp next case (Cons x2 xs2) hence xs2: "xs1 = x2 # xs2" by simp thus ?thesis proof (cases xs2) case Nil with assms and xs1 and xs2 show ?thesis by simp next case (Cons x3 xs3) hence "xs2 = x3 # xs3" by simp with assms and xs1 xs2 and A show ?thesis by simp qed qed qed qed lemmas foldl1_induct = foldl1_induct' [case_names One Two More Nil] lemma L8: assumes "associative f" and "xs \ []" and "ys \ []" shows "foldl1 f (xs @ ys) = f (foldl1 f xs) (foldl1 f ys)" using assms proof (induct f ys rule: foldl1_induct) case (One f y) have "foldl1 f (xs @ [y]) = foldl1 f (foldl1 f xs # [y])" using foldl1_append [where xs=xs] and One by simp also have "\ = f (foldl1 f xs) y" by simp also have "\ = f (foldl1 f xs) (foldl1 f [y])" by simp finally show ?case . next case (Two f x y) have "foldl1 f (xs @ [x, y]) = foldl1 f (foldl1 f xs # [x, y])" using foldl1_append [where xs=xs] and Two by simp also have "\ = foldl1 f (f (foldl1 f xs) x # [y])" by simp also have "\ = f (f (foldl1 f xs) x) y" by simp also have "\ = f (foldl1 f xs) (f x y)" using Two unfolding associative_def by simp also have "\ = f (foldl1 f xs) (foldl1 f [x, y])" by simp finally show ?case . next case (More f x y z zs) hence IH: "foldl1 f (xs @ f x y # z # zs) = f (foldl1 f xs) (foldl1 f (f x y # z # zs))" by simp have "foldl1 f (xs @ x # y # z # zs) = foldl1 f (foldl1 f xs # x # y # z # zs)" using foldl1_append [where xs=xs] and More by simp also have "\ = foldl1 f (f (foldl1 f xs) x # y # z # zs)" by simp also have "\ = foldl1 f (f (f (foldl1 f xs) x) y # z # zs)" by simp also have "\ = foldl1 f (f (foldl1 f xs) (f x y) # z # zs)" using More unfolding associative_def by simp also have "\ = foldl1 f (foldl1 f xs # f x y # z # zs)" by simp also have "\ = foldl1 f (xs @ f x y # z # zs)" using foldl1_append [where xs=xs] and More by simp also have "\ = f (foldl1 f xs) (foldl1 f (x # y # z # zs))" using IH by simp finally show ?case . next case Nil thus ?case by simp qed text \ The next lemma is applied in several following proofs whenever the equivalence of two lists is shown. \ lemma Lemma_2: assumes "length xs = length ys" and "\k. k < length xs \ xs ! k = ys ! k" shows "xs = ys" using assms by (auto simp: List.list_eq_iff_nth_eq) text \ In the original paper, this lemma and its proof appear inside of Lemma 3. However, this property will be useful also in later proofs and is thus separated. \ lemma foldl1_map: assumes "associative f" and "xs \ []" and "ys \ []" shows "foldl1 f (map h (xs @ ys)) = f (foldl1 f (map h xs)) (foldl1 f (map h ys))" proof - have "foldl1 f (map h (xs @ ys)) = foldl1 f (map h xs @ map h ys)" using L5 by simp also have "\ = f (foldl1 f (map h xs)) (foldl1 f (map h ys))" using assms and L8 [where f=f] by auto finally show ?thesis . qed lemma Lemma_3: fixes f :: "'a \ 'a \ 'a" and h :: "nat \ 'a" assumes "associative f" shows "map (foldl1 f \ map h) (candidate (@) (map wrap [0.. \The following three properties @{text P1}, @{text P2} and @{text P3}\ \ \are preconditions of Lemma 1.\ have P1: "\x y. \ x \ [] ; y \ [] \ \ foldl1 f (map h (x @ y)) = f (foldl1 f (map h x)) (foldl1 f (map h y))" using assms and foldl1_map by blast have P2: "\x y. x \ [] \ y \ [] \ x @ y \ []" by auto have P3: "\x. x \ set (map wrap [0.. x \ []" by auto \ \The proof for the thesis is now equal to the one of the original paper:\ from Lemma_1 [where h="foldl1 f \ map h" and zs="map wrap [0.. map h) (candidate (@) (map wrap [0.. map h) (map wrap [0.. = candidate f (map (foldl1 f \ map h \ wrap) [0.. = candidate f (map (foldl1 f \ wrap \ h) [0.. = candidate f (map h [0..Proving Proposition 1\ subsection \Definitions of Lemma 4\ text \ In the same way as in the original paper, the following two functions are defined: \ fun f1 :: "three \ three \ three" where "f1 x Zero = x" | "f1 Zero One = One" | "f1 x y = Two" fun f2 :: "three \ three \ three" where "f2 x Zero = x" | "f2 x One = One" | "f2 x Two = Two" text \ Both functions are associative as is proved by case analysis: \ lemma f1_assoc: "associative f1" unfolding associative_def proof auto fix x y z show "f1 x (f1 y z) = f1 (f1 x y) z" proof (cases z) case Zero thus ?thesis by simp next case One hence z_One: "z = One" by simp thus ?thesis by (cases y, simp_all, cases x, simp_all) next case Two thus ?thesis by simp qed qed lemma f2_assoc: "associative f2" unfolding associative_def proof auto fix x y z show "f2 x (f2 y z) = f2 (f2 x y) z" by (cases z, auto) qed text \ Next, we define two other functions, again according to the original paper. Note that \h1\ has an extra parameter \k\ which is only implicit in the original paper. \ fun h1 :: "nat \ nat \ nat \ three" where "h1 k i j = (if i = j then One else if j \ k then Zero else Two)" fun h2 :: "nat \ nat \ three" where "h2 i j = (if i = j then One else if i + 1 = j then Two else Zero)" subsection \Figures and Proofs\ text \ In the original paper, this lemma is depicted in (and proved by) Figure~2. Therefore, it carries this unusual name here. \ lemma Figure_2: assumes "i \ k" shows "foldl1 f1 (map (h1 k i) [0..j. j < length (map (h1 k i) [0.. (map (h1 k i) [0.. i \ (map (h1 k i) [0.. ?mr ! j = Zero" using List.nth_append [where xs="replicate i Zero"] by simp have R3: "j > i \ ?mr ! j = Zero" using List.nth_append [where xs="replicate i Zero @ [One]"] and j_k by simp show "(map (h1 k i) [0.. j" thus ?thesis proof (cases "i < j") assume "i < j" with M2 and R3 show ?thesis by simp next assume "\(i < j)" with i_ne_j have "i > j" by simp with M2 and R2 show ?thesis by simp qed qed qed from Q1 Q2 and Lemma_2 show ?thesis by blast qed have P2: "\j. j > 0 \ foldl1 f1 (replicate j Zero) = Zero" proof - fix j assume "(j::nat) > 0" thus "foldl1 f1 (replicate j Zero) = Zero" proof (induct j) case 0 thus ?case by simp next case (Suc j) thus ?case by (cases j, auto) qed qed have P3: "\j. foldl1 f1 ([One] @ replicate j Zero) = One" proof - fix j show "foldl1 f1 ([One] @ replicate j Zero) = One" using L8 [where f=f1 and xs="[One]" and ys="replicate (Suc j) Zero"] and f1_assoc and P2 [where j="Suc j"] by simp qed have "foldl1 f1 ?mr = One" proof (cases i) case 0 thus ?thesis using P3 by simp next case (Suc i) hence "foldl1 f1 (replicate (Suc i) Zero @ [One] @ replicate (k - Suc i) Zero) = f1 (foldl1 f1 (replicate (Suc i) Zero)) (foldl1 f1 ([One] @ replicate (k - Suc i) Zero))" using L8 [where xs="replicate (Suc i) Zero" and ys="[One] @ replicate (k - Suc i) Zero"] and f1_assoc by simp also have "\ = One" using P2 [where j="Suc i"] and P3 [where j="k - Suc i"] by simp finally show ?thesis using Suc by simp qed with P1 show ?thesis by simp qed text \ In the original paper, this lemma is depicted in (and proved by) Figure~3. Therefore, it carries this unusual name here. \ lemma Figure_3: assumes "i < k" shows "foldl1 f2 (map (h2 i) [0..j. j < length (map (h2 i) [0.. (map (h2 i) [0.. j > i + 1 \ (map (h2 i) [0.. ?mr ! j = Zero" using List.nth_append [where xs="replicate i Zero"] by simp have R2: "?mr ! i = One" using List.nth_append [where xs="replicate i Zero"] by simp have R3: "?mr ! (i + 1) = Two" using List.nth_append [where xs="replicate i Zero @ [One]"] by simp have R4: "j > i + 1 \ ?mr ! j = Zero" using List.nth_append [where xs="replicate i Zero @ [One,Two]"] and j_k by simp show "(map (h2 i) [0..(j < i)" hence j_ge_i: "j \ i" by simp thus ?thesis proof (cases "j = i") assume "j = i" with M1 and R2 show ?thesis by simp next assume "\(j = i)" with j_ge_i have j_gt_i: "j > i" by simp thus ?thesis proof (cases "j = i + 1") assume "j = i + 1" with M2 and R3 show ?thesis by simp next assume "\(j = i + 1)" with j_gt_i have "j > i + 1" by simp with M3 and R4 show ?thesis by simp qed qed qed qed from Q1 Q2 and Lemma_2 show ?thesis by blast qed have P2: "\j. j > 0 \ foldl1 f2 (replicate j Zero) = Zero" proof - fix j assume j_0: "(j::nat) > 0" show "foldl1 f2 (replicate j Zero) = Zero" using j_0 proof (induct j) case 0 thus ?case by simp next case (Suc j) thus ?case by (cases j, auto) qed qed have P3: "\j. foldl1 f2 ([One, Two] @ replicate j Zero) = Two" proof - fix j show "foldl1 f2 ([One, Two] @ replicate j Zero) = Two" using L8 [where f=f2 and xs="[One,Two]" and ys="replicate (Suc j) Zero"] and f2_assoc and P2 [where j="Suc j"] by simp qed have "foldl1 f2 ?mr = Two" proof (cases i) case 0 thus ?thesis using P3 by simp next case (Suc i) hence "foldl1 f2 (replicate (Suc i) Zero @ [One, Two] @ replicate (k - Suc i - 1) Zero) = f2 (foldl1 f2 (replicate (Suc i) Zero)) (foldl1 f2 ([One, Two] @ replicate (k - Suc i - 1) Zero))" using L8 [where f=f2 and xs="replicate (Suc i) Zero" and ys="[One, Two] @ replicate (k - Suc i - 1) Zero"] and f2_assoc by simp also have "\ = Two" using P2 [where j="Suc i"] and P3 [where j="k - Suc i - 1"] by simp finally show ?thesis using Suc by simp qed with P1 show ?thesis by simp qed text \ Counterparts of the following two lemmas are shown in the proof of Lemma 4 in the original paper. Since here, the proof of Lemma 4 is seperated in several smaller lemmas, also these two properties are given separately. \ lemma L9: assumes "\ (f :: three \ three \ three) h. associative f \ foldl1 f (map h js) = foldl1 f (map h [0.. k" shows "foldl1 f1 (map (h1 k i) js) = One" using assms and f1_assoc and Figure_2 by auto lemma L10: assumes "\ (f :: three \ three \ three) h. associative f \ foldl1 f (map h js) = foldl1 f (map h [0.. In the original paper, this lemma is depicted in (and proved by) Figure~4. Therefore, it carries this unusual name here. This lemma expresses that every \i \ k\ is contained in \js\ at least once. \ lemma Figure_4: assumes "foldl1 f1 (map (h1 k i) js) = One" and "js \ []" shows "i \ set js" proof (rule ccontr) assume i_not_in_js: "i \ set js" have One_not_in_map_js: "One \ set (map (h1 k i) js)" proof assume "One \ set (map (h1 k i) js)" hence "One \ (h1 k i) ` (set js)" by simp then obtain j where j_def: "j \ set js \ One = h1 k i j" using Set.image_iff [where f="h1 k i"] by auto hence "i = j" by (simp split: if_splits) with i_not_in_js and j_def show False by simp qed have f1_One: "\x y. x \ One \ y \ One \ f1 x y \ One" proof - fix x y assume "x \ One \ y \ One" thus "f1 x y \ One" by (cases y, cases x, auto) qed have "\xs. \ xs \ [] ; One \ set xs \ \ foldl1 f1 xs \ One" proof - fix xs assume A: "(xs :: three list) \ []" thus "One \ set xs \ foldl1 f1 xs \ One" proof (induct xs rule: divide_and_conquer) case Nil thus ?case by simp next case (One x) thus "foldl1 f1 [x] \ One" by simp next case (Partition xs ys) hence "One \ set xs \ One \ set ys" by simp with Partition have "foldl1 f1 xs \ One \ foldl1 f1 ys \ One" by simp with f1_One have "f1 (foldl1 f1 xs) (foldl1 f1 ys) \ One" by simp with L8 [symmetric, where f=f1] and f1_assoc and Partition show "foldl1 f1 (xs @ ys) \ One" by auto qed qed with One_not_in_map_js and assms show False by auto qed text \ In the original paper, this lemma is depicted in (and proved by) Figure~5. Therefore, it carries this unusual name here. This lemma expresses that every \i \ k\ is contained in \js\ at most once. \ lemma Figure_5: assumes "foldl1 f1 (map (h1 k i) js) = One" and "js = xs @ ys" shows "\(i \ set xs \ i \ set ys)" proof (rule ccontr) assume "\\(i \ set xs \ i \ set ys)" hence i_xs_ys: "i \ set xs \ i \ set ys" by simp from i_xs_ys have xs_not_empty: "xs \ []" by auto from i_xs_ys have ys_not_empty: "ys \ []" by auto have f1_Zero: "\x y. x \ Zero \ y \ Zero \ f1 x y \ Zero" proof - fix x y show "x \ Zero \ y \ Zero \ f1 x y \ Zero" by (cases y, simp_all, cases x, simp_all) qed have One_foldl1: "\xs. One \ set xs \ foldl1 f1 xs \ Zero" proof - fix xs assume One_xs: "One \ set xs" thus "foldl1 f1 xs \ Zero" proof (induct xs rule: divide_and_conquer) case Nil thus ?case by simp next case One thus ?case by simp next case (Partition xs ys) hence "One \ set xs \ One \ set ys" by simp with Partition have "foldl1 f1 xs \ Zero \ foldl1 f1 ys \ Zero" by auto with f1_Zero have "f1 (foldl1 f1 xs) (foldl1 f1 ys) \ Zero" by simp thus ?case using L8 [symmetric, where f=f1] and f1_assoc and Partition by auto qed qed have f1_Two: "\x y. x \ Zero \ y \ Zero \ f1 x y = Two" proof - fix x y show "x \ Zero \ y \ Zero \ f1 x y = Two" by (cases y, simp_all, cases x, simp_all) qed from i_xs_ys have "One \ set (map (h1 k i) xs) \ One \ set (map (h1 k i) ys)" by simp hence "foldl1 f1 (map (h1 k i) xs) \ Zero \ foldl1 f1 (map (h1 k i) ys) \ Zero" using One_foldl1 by simp hence "f1 (foldl1 f1 (map (h1 k i) xs)) (foldl1 f1 (map (h1 k i) ys)) = Two" using f1_Two by simp hence "foldl1 f1 (map (h1 k i) (xs @ ys)) = Two" using foldl1_map [symmetric, where h="h1 k i"] and f1_assoc and xs_not_empty and ys_not_empty by auto with assms show False by simp qed text \ In the original paper, this lemma is depicted in (and proved by) Figure~6. Therefore, it carries this unusual name here. This lemma expresses that \js\ contains only elements of \[0... \ lemma Figure_6: assumes "\i. i \ k \ foldl1 f1 (map (h1 k i) js) = One" and "i > k" shows "i \ set js" proof assume i_in_js: "i \ set js" have Two_map: "Two \ set (map (h1 k 0) js)" proof - have "Two = h1 k 0 i" using assms by simp with i_in_js show ?thesis using IntI by (auto split: if_splits) qed have f1_Two: "\x y. x = Two \ y = Two \ f1 x y = Two" proof - fix x y show "x = Two \ y = Two \ f1 x y = Two" by (cases y, auto) qed have "\xs. Two \ set xs \ foldl1 f1 xs = Two" proof - fix xs assume Two_xs: "Two \ set xs" thus "foldl1 f1 xs = Two" using Two_xs proof (induct xs rule: divide_and_conquer) case Nil thus ?case by simp next case One thus ?case by simp next case (Partition xs ys) hence "Two \ set xs \ Two \ set ys" by simp hence "foldl1 f1 xs = Two \ foldl1 f1 ys = Two" using Partition by auto with f1_Two have "f1 (foldl1 f1 xs) (foldl1 f1 ys) = Two" by simp thus "foldl1 f1 (xs @ ys) = Two" using L8 [symmetric, where f=f1] and f1_assoc and Partition by auto qed qed with Two_map have "foldl1 f1 (map (h1 k 0) js) = Two" by simp with assms show False by auto qed text \ In the original paper, this lemma is depicted in (and proved by) Figure~7. Therefore, it carries this unusual name here. This lemma expresses that every \i \ k\ in \js\ is eventually followed by \i + 1\. \ lemma Figure_7: assumes "foldl1 f2 (map (h2 i) js) = Two" and "js = xs @ ys" and "xs \ []" and "i = last xs" shows "(i + 1) \ set ys" proof (rule ccontr) assume Suc_i_not_in_ys: "(i + 1) \ set ys" have last_map_One: "last (map (h2 i) xs) = One" proof - have "last (map (h2 i) xs) = (map (h2 i) xs) ! (length (map (h2 i) xs) - 1)" using List.last_conv_nth [where xs="map (h2 i) xs"] and assms by simp also have "\ = (map (h2 i) xs) ! (length xs - 1)" using L2 by simp also have "\ = h2 i (xs ! (length xs - 1))" using L6 and assms by simp also have "\ = h2 i (last xs)" using List.last_conv_nth [symmetric,where xs=xs] and assms by simp also have "\ = One" using assms by simp finally show ?thesis . qed have "\xs. \ xs \ [] ; last xs = One \ \ foldl1 f2 xs = One" proof - fix xs assume last_xs_One: "last xs = One" assume xs_not_empty: "xs \ []" hence xs_partition: "xs = butlast xs @ [last xs]" by simp show "foldl1 f2 xs = One" proof (cases "butlast xs") case Nil with xs_partition and last_xs_One show ?thesis by simp next case Cons hence butlast_not_empty: "butlast xs \ []" by simp have "foldl1 f2 xs = foldl1 f2 (butlast xs @ [last xs])" using xs_partition by simp also have "\ = f2 (foldl1 f2 (butlast xs)) (foldl1 f2 [last xs])" using L8 [where f=f2] and f2_assoc and butlast_not_empty by simp also have "\ = One" using last_xs_One by simp finally show ?thesis . qed qed with last_map_One have foldl1_map_xs: "foldl1 f2 (map (h2 i) xs) = One" using assms by simp have ys_not_empty: "ys \ []" using foldl1_map_xs and assms by auto have Two_map_ys: "Two \ set (map (h2 i) ys)" proof assume "Two \ set (map (h2 i) ys)" hence "Two \ (h2 i) ` (set ys)" by simp then obtain j where j_def: "j \ set ys \ Two = h2 i j" using Set.image_iff [where f="h2 i"] by auto hence "i + 1 = j" by (simp split: if_splits) with Suc_i_not_in_ys and j_def show False by simp qed have f2_One: "\x y. x \ Two \ y \ Two \ f2 x y \ Two" proof - fix x y show "x \ Two \ y \ Two \ f2 x y \ Two" by (cases y, auto) qed have "\xs. \ xs \ [] ; Two \ set xs \ \ foldl1 f2 xs \ Two" proof - fix xs assume xs_not_empty: "(xs :: three list) \ []" thus "Two \ set xs \ foldl1 f2 xs \ Two" proof (induct xs rule: divide_and_conquer) case Nil thus ?case by simp next case One thus ?case by simp next case (Partition xs ys) hence "Two \ set xs \ Two \ set ys" by simp hence "foldl1 f2 xs \ Two \ foldl1 f2 ys \ Two" using Partition by simp hence "f2 (foldl1 f2 xs) (foldl1 f2 ys) \ Two" using f2_One by simp thus "foldl1 f2 (xs @ ys) \ Two" using L8 [symmetric, where f=f2] and f2_assoc and Partition by simp qed qed with Two_map_ys have foldl1_map_ys: "foldl1 f2 (map (h2 i) ys) \ Two" using ys_not_empty by simp from f2_One have "f2 (foldl1 f2 (map (h2 i) xs)) (foldl1 f2 (map (h2 i) ys)) \ Two" using foldl1_map_xs and foldl1_map_ys by simp hence "foldl1 f2 (map (h2 i) (xs @ ys)) \ Two" using foldl1_map [symmetric, where h="h2 i" and f=f2] and f2_assoc and assms and ys_not_empty by simp with assms show False by simp qed subsection \Permutations and Lemma 4\ text \ In the original paper, the argumentation goes as follows: From \Figure_4\ and \Figure_5\ we can show that \js\ contains every \i \ k\ exactly once, and from \Figure_6\ we can furthermore show that \js\ contains no other elements. Thus, \js\ must be a permutation of \[0... Here, however, the argumentation is different, because we want to use already existing results. Therefore, we show first, that the sets of \js\ and \[0.. are equal using the results of \Figure_4\ and \Figure_6\. Second, we show that \js\ is a distinct list, i.e. no element occurs twice in \js\. Since also \[0.. is distinct, the multisets of \js\ and \[0.. are equal, and therefore, both lists are permutations. \ lemma js_is_a_permutation: assumes A1: "\ (f :: three \ three \ three) h. associative f \ foldl1 f (map h js) = foldl1 f (map h [0.. []" shows "mset js = mset [0..i. i \ k \ foldl1 f1 (map (h1 k i) js) = One" by auto from L9' and Figure_4 and A2 have P1: "\i. i \ k \ i \ set js" by auto from L9' and Figure_5 have P2: "\i xs ys. \ i \ k ; js = xs @ ys \ \ \(i \ set xs \ i \ set ys)" by blast from L9' and Figure_6 have P3: "\i. i > k \ i \ set js" by auto have set_eq: "set [0.. set js" by auto next show "set js \ set [0.. set js" hence "\(j \ set js)" by simp with P3 have "\(j > k)" using HOL.contrapos_nn by auto hence "j \ k" by simp thus "j \ set [0..xs ys. js = xs @ ys \ set xs \ set ys = {}" proof - fix xs ys assume js_xs_ys: "js = xs @ ys" with set_eq have i_xs_ys: "\i. i \ set xs \ i \ set ys \ i \ k" by auto have "\i. i \ k \ (i \ set xs) = (i \ set ys)" proof fix i assume "i \ set xs" moreover assume "i \ k" ultimately show "i \ set ys" using js_xs_ys and P2 by simp next fix i assume "i \ set ys" moreover assume "i \ k" ultimately show "i \ set xs" using js_xs_ys and P2 and P1 by auto qed thus "set xs \ set ys = {}" using i_xs_ys by auto qed with all_set_inter_empty_distinct have "distinct js" using A2 by auto with set_eq show "mset js = mset [0.. The result of \Figure_7\ is too specific. Instead of having that every \i\ is eventually followed by \i + 1\, it more useful to know that every \i\ is followed by all \i + r\, where \r > 0\. This result follows easily by induction from \Figure_7\. \ lemma Figure_7_trans: assumes A1: "\i xs ys. \ i < k ; js = xs @ ys ; xs \ [] ; i = last xs \ \ (i + 1) \ set ys" and A2: "(r::nat) > 0" and A3: "i + r \ k" and A4: "js = xs @ ys" and A5: "xs \ []" and A6: "i = last xs" shows "(i + r) \ set ys" using A2 A3 proof (induct r) case 0 thus ?case by simp next case (Suc r) hence IH: "0 < r \ (i + r) \ set ys" by simp from Suc have i_r_k: "i + Suc r \ k" by simp show ?case proof (cases r) case 0 thus ?thesis using A1 and i_r_k and A4 and A5 and A6 by auto next case Suc with IH have "(i + r) \ set ys" by simp then obtain p where p_def: "p < length ys \ ys ! p = i + r" using List.in_set_conv_nth [where x="i + r"] by auto let ?xs = "xs @ take (p + 1) ys" let ?ys = "drop (p + 1) ys" have "i + r < k" using i_r_k by simp moreover have "js = ?xs @ ?ys" using A4 by simp moreover have "?xs \ []" using A5 by simp moreover have "i + r = last ?xs" using p_def and List.take_Suc_conv_app_nth [where i=p and xs=ys] by simp ultimately have "(i + Suc r) \ set ?ys" using A1 [where i="i + r"] by auto thus "(i + Suc r) \ set ys" using List.set_drop_subset [where xs=ys] by auto qed qed text \ Since we want to use Lemma \partitions_sorted\ to show that \js\ is sorted, we need yet another result which can be obtained using the previous lemma and some further argumentation: \ lemma js_partition_order: assumes A1: "mset js = mset [0..i xs ys. \ i < k ; js = xs @ ys ; xs \ [] ; i = last xs \ \ (i + 1) \ set ys" and A3: "js = xs @ ys" and A4: "i \ set xs" and A5: "j \ set ys" shows "i \ j" proof (rule ccontr) from A1 have A1': \set js = {.. by (metis atLeast_upt mset_eq_setD) assume "\(i \ j)" hence i_j: "i > j" by simp from A5 obtain pj where pj_def: "pj < length ys \ ys ! pj = j" using List.in_set_conv_nth [where x=j] by auto let ?xs = "xs @ take (pj + 1) ys" let ?ys = "drop (pj + 1) ys" let ?r = "i - j" from A1 and A3 have "distinct (xs @ ys)" using distinct_upt mset_eq_imp_distinct_iff by blast hence xs_ys_inter_empty: "set xs \ set ys = {}" by simp from A2 and Figure_7_trans have "\i r xs ys. \ r > 0 ; i + r \ k ; js = xs @ ys ; xs \ [] ; i = last xs \ \ (i + r) \ set ys" by blast moreover from i_j have "?r > 0" by simp moreover have "j + ?r \ k" proof - have "i \ set js" using A4 and A3 by simp hence "i \ set [0.. k" by auto thus ?thesis using i_j by simp qed moreover have "js = ?xs @ ?ys" using A3 by simp moreover have "?xs \ []" using A4 by auto moreover have "j = last (?xs)" using pj_def and List.take_Suc_conv_app_nth [where i=pj and xs=ys] by simp ultimately have "(j + ?r) \ set ?ys" by blast hence "i \ set ys" using i_j and List.set_drop_subset [where xs=ys] by auto with A4 and xs_ys_inter_empty show False by auto qed text \ With the help of the previous lemma, we show now that \js\ equals \[0.., if both lists are permutations and every \i\ is eventually followed by \i + 1\ in \js\. \ lemma js_equals_upt_k: assumes A1: "mset js = mset [0..i xs ys. \ i < k ; js = xs @ ys ; xs \ [] ; i = last xs \ \ (i + 1) \ set ys" shows "js = [0..xs ys x y. \ js = xs @ ys ; x \ set xs ; y \ set ys \ \ x \ y" by blast hence "sorted js" using partitions_sorted by blast moreover have "distinct js" using A1 distinct_upt mset_eq_imp_distinct_iff by blast moreover have "sorted [0.. From all the work done before, we conclude now Lemma 4: \ lemma Lemma_4: assumes "\(f :: three \ three \ three) h. associative f \ foldl1 f (map h js) = foldl1 f (map h [0.. []" shows "js = [0..i xs ys. \ i < k ; js = xs @ ys ; xs \ [] ; i = last xs \ \ (i + 1) \ set ys" by blast ultimately show ?thesis using js_equals_upt_k by auto qed subsection \Lemma 5\ text \ This lemma is a lifting of Lemma 4 to the overall computation of \scanl1\. Its proof follows closely the one given in the original paper. \ lemma Lemma_5: assumes "\(f :: three \ three \ three) h. associative f \ map (foldl1 f \ map h) jss = scanl1 f (map h [0..js. js \ set jss \ js \ []" shows "jss = ups n" proof - have P1: "length jss = length (ups n)" proof - obtain f :: "three \ three \ three" where f_assoc: "associative f" using f1_assoc by auto fix h have "length jss = length (map (foldl1 f \ map h) jss)" by (simp add: L2) also have "\ = length (scanl1 f (map h [0.. = length (map (\k. foldl1 f (take (k + 1) (map h [0.. = length [0.. = length [0.. = length [0.. = length (map (\k. [0.. = length (ups n)" by simp finally show ?thesis . qed have P2: "\k. k < length jss \ jss ! k = (ups n) ! k" proof - fix k assume k_length_jss: "k < length jss" hence non_empty_jss_k: "jss ! k \ []" using assms by simp from k_length_jss have k_length_length: "k < length [1..(f :: three \ three \ three) h. associative f \ foldl1 f (map h (jss ! k)) = foldl1 f (map h [0.. three \ three)" have "foldl1 f (map h (jss ! k)) = (map (foldl1 f \ map h) jss) ! k" using L6 and k_length_jss by auto also have "\ = (scanl1 f (map h [0.. = (map (\k. foldl1 f (take k (map h [0.. = (map (\k. foldl1 f (take k (map h [0.. = (\k. foldl1 f (take k (map h [0.. = foldl1 f (take (k + 1) (map h [0.. = foldl1 f (map h (take (k + 1) [0.. = foldl1 f (map h [0..k. [0.. = (\k. [0.. = [0..Proposition 1\ text \ In the original paper, only non-empty lists where considered, whereas here, the used list datatype allows also for empty lists. Therefore, we need to exclude non-empty lists to have a similar setting as in the original paper. In the case of Proposition 1, we need to show that every list contained in the result of \candidate (@) (map wrap [0.. is non-empty. The idea is to interpret empty lists by the value \Zero\ and non-empty lists by the value \One\, and to apply the assumptions. \ lemma non_empty_candidate_results: assumes "\ (f :: three \ three \ three) (xs :: three list). \ associative f ; xs \ [] \ \ candidate f xs = scanl1 f xs" and "js \ set (candidate (@) (map wrap [0.. []" proof - \ \We define a mapping of lists to values of @{text three} as explained\ \ \above, and a function which behaves like @{text @} on values of\ \ \@{text three}.\ let ?h = "\xs. case xs of [] \ Zero | (_#_) \ One" let ?g = "\x y. if (x = One \ y = One) then One else Zero" have g_assoc: "associative ?g" unfolding associative_def by auto \ \Our defined functions fulfill the requirements of the free theorem of\ \ \@{text candidate}, that is:\ have req_free_theorem: "\xs ys. ?h (xs @ ys) = ?g (?h xs) (?h ys)" proof - fix xs ys show "?h (xs @ ys) = ?g (?h xs) (?h ys)" by (cases xs, simp_all, cases ys, simp_all) qed \ \Before applying the assumptions, we show that @{text candidate}'s\ \ \counterpart @{text scanl1}, applied to a non-empty list, returns only\ \ \a repetition of the value @{text One}.\ have set_scanl1_is_One: "set (scanl1 ?g (map ?h (map wrap [0..x. One) [0..x. One) [0..x. One) ([0.. = map (\x. One) [0..x. One) [Suc n]" by simp also have "\ = replicate (Suc n) One @ replicate 1 One" using Suc by simp also have "\ = replicate (Suc n + 1) One" using List.replicate_add [symmetric, where x=One and n="Suc n" and m=1] by simp finally show ?case . qed have foldl1_One: "\xs. foldl1 ?g (One # xs) = One" proof - fix xs show "foldl1 ?g (One # xs) = One" proof (induct xs rule: measure_induct [where f=length]) fix x assume "\y. length y < length (x::three list) \ foldl1 ?g (One # y) = One" thus "foldl1 ?g (One # x) = One" by (cases x, auto) qed qed have "scanl1 ?g (map ?h (map wrap [0.. wrap) [0.. = scanl1 ?g (map (\x. One) [0.. = scanl1 ?g (replicate (n + 1) One)" using const_One by auto also have "\ = map (\k. foldl1 ?g (take k (replicate (n + 1) One))) [1.. = map (\k. foldl1 ?g (take k (replicate (n + 1) One))) (map Suc [0.. = map ((\k. foldl1 ?g (take k (replicate (n + 1) One))) \ Suc) [0.. = map (\k. foldl1 ?g (replicate (min (k + 1) (n + 1)) One)) [0.. = map (\k. foldl1 ?g (One # replicate (min k n) One)) [0.. = map (\k. One) [0.. = replicate (n + 1) One" using const_One by simp finally show ?thesis using List.set_replicate [where n="n + 1"] by simp qed \ \Thus, with the assumptions and the free theorem of candidate, we show\ \ \that results of @{text candidate}, after applying @{text h}, can only\ \ \have the value @{text One}.\ have "scanl1 ?g (map ?h (map wrap [0.. = map ?h (candidate (@) (map wrap [0..x. x \ set (map ?h (candidate (@) (map wrap [0.. x = One" using set_scanl1_is_One by auto \ \Now, it is easy to conclude the thesis.\ from assms have "?h js \ ?h ` set (candidate (@) (map wrap [0.. []" by auto qed text \ Proposition 1 is very similar to the corresponding one shown in the original paper except of a slight modification due to the choice of using Isabelle's list datatype. Strictly speaking, the requirement that \xs\ must be non-empty in the assumptions of Proposition 1 is not necessary, because only non-empty lists are applied in the proof. However, the additional requirement eases the proof obligations of the final theorem, i.e. this additions allows more (or easier) applications of the final theorem. \ lemma Proposition_1: assumes "\ (f :: three \ three \ three) (xs :: three list). \ associative f ; xs \ [] \ \ candidate f xs = scanl1 f xs" shows "candidate (@) (map wrap [0.. \This addition is necessary because we are using Isabelle's list datatype\ \ \which allows for empty lists.\ from assms and non_empty_candidate_results have non_empty_candidate: "\js. js \ set (candidate (@) (map wrap [0.. js \ []" by auto have "\(f:: three \ three \ three) h. associative f \ map (foldl1 f \ map h) (candidate (@) (map wrap [0.. three \ three)" hence "map (foldl1 f \ map h) (candidate (@) (map wrap [0.. = scanl1 f (map h [0.. map h) (candidate (@) (map wrap [0..Proving Proposition 2\ text \ Before proving Proposition 2, a non-trivial step of that proof is shown first. In the original paper, the argumentation simply applies L7 and the definition of \map\ and \[0... However, since, L7 requires that \k\ must be less than \length [0.. and this does not simply follow for the bound occurrence of \k\, a more complicated proof is necessary. Here, it is shown based on Lemma 2. \ lemma Prop_2_step_L7: "map (\k. foldl1 g (map (nth xs) [0..k. foldl1 g (take (k + 1) xs)) [0..k. foldl1 g (map (nth xs) [0..k. foldl1 g (take (k + 1) xs)) [0..k. k < length (map (\k. foldl1 g (map (nth xs) [0.. (map (\k. foldl1 g (map (nth xs) [0..k. foldl1 g (take (k + 1) xs)) [0..k. foldl1 g (map (nth xs) [0..k. foldl1 g (map (nth xs) [0..k. foldl1 g (map (nth xs) [0.. = foldl1 g (map (nth xs) [0.. = foldl1 g (take (k + 1) xs)" using L7 [where k=k and xs=xs] and k_length' by simp also have "\ = (\k. foldl1 g (take (k + 1) xs)) ([0.. = (map (\k. foldl1 g (take (k + 1) xs)) [0..k. foldl1 g (map (nth xs) [0..k. foldl1 g (take (k + 1) xs)) [0.. Compared to the original paper, here, Proposition 2 has the additional assumption that \xs\ is non-empty. The proof, however, is identical to the the one given in the original paper, except for the non-trivial step shown before. \ lemma Proposition_2: assumes A1: "\ n. candidate (@) (map wrap [0.. []" shows "candidate g xs = scanl1 g xs" proof - \ \First, based on Lemma 2, we show that\ \ \@{term "xs = map (nth xs) [0.. \ \by the following facts P1 and P2.\ have P1: "length xs = length (map (nth xs) [0.. = length (map (nth xs) [0..k. k < length xs \ xs ! k = (map (nth xs) [0.. = (map (nth xs) [0.. \Thus, with some rewriting, we show the thesis.\ hence "candidate g xs = candidate g (map (nth xs) [0.. = map (foldl1 g \ map (nth xs)) (candidate (@) (map wrap [0.. = map (foldl1 g \ map (nth xs)) (ups (length xs - 1))" using A1 [where n="length xs - 1"] and A3 by simp also have "\ = map (foldl1 g \ map (nth xs)) (map (\k. [0.. = map (\k. foldl1 g (map (nth xs) [0.. = map (\k. foldl1 g (take (k + 1) xs)) [0.. = map (\k. foldl1 g (take k xs)) (map (\k. k + 1) [0.. = map (\k. foldl1 g (take k xs)) [1.. = scanl1 g xs" by simp finally show ?thesis . qed section \The Final Result\ text \ Finally, we the main result follows directly from Proposition 1 and Proposition 2. \ theorem The_0_1_2_Principle: assumes "\ (f :: three \ three \ three) (xs :: three list). \ associative f ; xs \ [] \ \ candidate f xs = scanl1 f xs" and "associative g" and "ys \ []" shows "candidate g ys = scanl1 g ys" using Proposition_1 Proposition_2 and assms by blast text \ \section*{Acknowledgments} I thank Janis Voigtl\"ander for sharing a draft of his paper before its publication with me. \ (*<*) end (*>*) diff --git a/thys/MuchAdoAboutTwo/ROOT b/thys/MuchAdoAboutTwo/ROOT --- a/thys/MuchAdoAboutTwo/ROOT +++ b/thys/MuchAdoAboutTwo/ROOT @@ -1,9 +1,11 @@ chapter AFP session MuchAdoAboutTwo (AFP) = "HOL-Library" + options [timeout = 600] + sessions + "HOL-Combinatorics" theories MuchAdoAboutTwo document_files "root.bib" "root.tex" diff --git a/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy b/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy --- a/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy +++ b/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy @@ -1,1034 +1,1034 @@ section \Permutations as Products of Disjoint Cycles\ theory Executable_Permutations imports Graph_Theory.Funpow List_Aux - "HOL-Library.Permutations" "HOL-Library.Rewrite" + "HOL-Combinatorics.Permutations" begin subsection \Cyclic Permutations\ definition list_succ :: "'a list \ 'a \ 'a" where "list_succ xs x = (if x \ set xs then xs ! ((index xs x + 1) mod length xs) else x)" text \ We demonstrate the functions on the following simple lemmas @{lemma "list_succ [1 :: int, 2, 3] 1 = 2" by code_simp} @{lemma "list_succ [1 :: int, 2, 3] 2 = 3" by code_simp} @{lemma "list_succ [1 :: int, 2, 3] 3 = 1" by code_simp} \ lemma list_succ_altdef: "list_succ xs x = (let n = index xs x in if n + 1 = length xs then xs ! 0 else if n + 1 < length xs then xs ! (n + 1) else x)" using index_le_size[of xs x] unfolding list_succ_def index_less_size_conv[symmetric] by (auto simp: Let_def) lemma list_succ_Nil: "list_succ [] = id" by (simp add: list_succ_def fun_eq_iff) lemma list_succ_singleton: "list_succ [x] = list_succ []" by (simp add: fun_eq_iff list_succ_def) lemma list_succ_short: assumes "length xs < 2" shows "list_succ xs = id" using assms by (cases xs) (rename_tac [2] y ys, case_tac [2] ys, auto simp: list_succ_Nil list_succ_singleton) lemma list_succ_simps: "index xs x + 1 = length xs \ list_succ xs x = xs ! 0" "index xs x + 1 < length xs \ list_succ xs x = xs ! (index xs x + 1)" "length xs \ index xs x \ list_succ xs x = x" by (auto simp: list_succ_altdef) lemma list_succ_not_in: assumes "x \ set xs" shows "list_succ xs x = x" using assms by (auto simp: list_succ_def) lemma list_succ_list_succ_rev: assumes "distinct xs" shows "list_succ (rev xs) (list_succ xs x) = x" proof - { assume "index xs x + 1 < length xs" moreover then have "length xs - Suc (Suc (length xs - Suc (Suc (index xs x)))) = index xs x" by linarith ultimately have ?thesis using assms by (simp add: list_succ_def index_rev index_nth_id rev_nth) } moreover { assume A: "index xs x + 1 = length xs" moreover from A have "xs \ []" by auto moreover with A have "last xs = xs ! index xs x" by (cases "length xs") (auto simp: last_conv_nth) ultimately have ?thesis using assms by (auto simp add: list_succ_def rev_nth index_rev index_nth_id last_conv_nth) } moreover { assume A: "index xs x \ length xs" then have "x \ set xs" by (metis index_less less_irrefl) then have ?thesis by (auto simp: list_succ_def) } ultimately show ?thesis by (metis discrete le_less not_less) qed lemma inj_list_succ: "distinct xs \ inj (list_succ xs)" by (metis injI list_succ_list_succ_rev) lemma inv_list_succ_eq: "distinct xs \ inv (list_succ xs) = list_succ (rev xs)" by (metis distinct_rev inj_imp_inv_eq inj_list_succ list_succ_list_succ_rev) lemma bij_list_succ: "distinct xs \ bij (list_succ xs)" by (metis bij_def inj_list_succ distinct_rev list_succ_list_succ_rev surj_def) lemma list_succ_permutes: assumes "distinct xs" shows "list_succ xs permutes set xs" using assms by (auto simp: permutes_conv_has_dom bij_list_succ has_dom_def list_succ_def) lemma permutation_list_succ: assumes "distinct xs" shows "permutation (list_succ xs)" using list_succ_permutes[OF assms] by (auto simp: permutation_permutes) lemma list_succ_nth: assumes "distinct xs" "n < length xs" shows "list_succ xs (xs ! n) = xs ! (Suc n mod length xs)" using assms by (auto simp: list_succ_def index_nth_id) lemma list_succ_last[simp]: assumes "distinct xs" "xs \ []" shows "list_succ xs (last xs) = hd xs" using assms by (auto simp: list_succ_def hd_conv_nth) lemma list_succ_rotate1[simp]: assumes "distinct xs" shows "list_succ (rotate1 xs) = list_succ xs" proof (rule ext) fix y show "list_succ (rotate1 xs) y = list_succ xs y" using assms proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases "x = y") case True then have "index (xs @ [y]) y = length xs" using \distinct (x # xs)\ by (simp add: index_append) with True show ?thesis by (cases "xs=[]") (auto simp: list_succ_def nth_append) next case False then show ?thesis apply (cases "index xs y + 1 < length xs") apply (auto simp:list_succ_def index_append nth_append) by (metis Suc_lessI index_less_size_conv mod_self nth_Cons_0 nth_append nth_append_length) qed qed qed lemma list_succ_rotate[simp]: assumes "distinct xs" shows "list_succ (rotate n xs) = list_succ xs" using assms by (induct n) auto lemma list_succ_in_conv: "list_succ xs x \ set xs \ x \ set xs" by (auto simp: list_succ_def not_nil_if_in_set ) lemma list_succ_in_conv1: assumes "A \ set xs = {}" shows "list_succ xs x \ A \ x \ A" by (metis assms disjoint_iff_not_equal list_succ_in_conv list_succ_not_in) lemma list_succ_commute: assumes "set xs \ set ys = {}" shows "list_succ xs (list_succ ys x) = list_succ ys (list_succ xs x)" proof - have "\x. x \ set xs \ list_succ ys x = x" "\x. x \ set ys \ list_succ xs x = x" using assms by (blast intro: list_succ_not_in)+ then show ?thesis by (cases "x \ set xs \ set ys") (auto simp: list_succ_in_conv list_succ_not_in) qed subsection \Arbitrary Permutations\ fun lists_succ :: "'a list list \ 'a \ 'a" where "lists_succ [] x = x" | "lists_succ (xs # xss) x = list_succ xs (lists_succ xss x)" definition distincts :: "'a list list \ bool" where "distincts xss \ distinct xss \ (\xs \ set xss. distinct xs \ xs \ []) \ (\xs \ set xss. \ys \ set xss. xs \ ys \ set xs \ set ys = {})" lemma distincts_distinct: "distincts xss \ distinct xss" by (auto simp: distincts_def) lemma distincts_Nil[simp]: "distincts []" by (simp add: distincts_def) lemma distincts_single: "distincts [xs] \ distinct xs \ xs \ []" by (auto simp add: distincts_def) lemma distincts_Cons: "distincts (xs # xss) \ xs \ [] \ distinct xs \ distincts xss \ (set xs \ (\ys \ set xss. set ys)) = {}" (is "?L \ ?R") proof assume ?L then show ?R by (auto simp: distincts_def) next assume ?R then have "distinct (xs # xss)" apply (auto simp: disjoint_iff_not_equal distincts_distinct) apply (metis length_greater_0_conv nth_mem) done moreover from \?R\ have "\xs \ set (xs # xss). distinct xs \ xs \ []" by (auto simp: distincts_def) moreover from \?R\ have "\xs' \ set (xs # xss). \ys \ set (xs # xss). xs' \ ys \ set xs' \ set ys = {}" by (simp add: distincts_def) blast ultimately show ?L unfolding distincts_def by (intro conjI) qed lemma distincts_Cons': "distincts (xs # xss) \ xs \ [] \ distinct xs \ distincts xss \ (\ys \ set xss. set xs \ set ys = {})" (is "?L \ ?R") unfolding distincts_Cons by blast lemma distincts_rev: "distincts (map rev xss) \ distincts xss" by (simp add: distincts_def distinct_map) lemma length_distincts: assumes "distincts xss" shows "length xss = card (set ` set xss)" using assms proof (induct xss) case Nil then show ?case by simp next case (Cons xs xss) then have "set xs \ set ` set xss" using equals0I[of "set xs"] by (auto simp: distincts_Cons disjoint_iff_not_equal ) with Cons show ?case by (auto simp add: distincts_Cons) qed lemma distincts_remove1: "distincts xss \ distincts (remove1 xs xss)" by (auto simp: distincts_def) lemma distinct_Cons_remove1: "x \ set xs \ distinct (x # remove1 x xs) = distinct xs" by (induct xs) auto lemma set_Cons_remove1: "x \ set xs \ set (x # remove1 x xs) = set xs" by (induct xs) auto lemma distincts_Cons_remove1: "xs \ set xss \ distincts (xs # remove1 xs xss) = distincts xss" by (simp only: distinct_Cons_remove1 set_Cons_remove1 distincts_def) lemma distincts_inj_on_set: assumes "distincts xss" shows "inj_on set (set xss)" by (rule inj_onI) (metis assms distincts_def inf.idem set_empty) lemma distincts_distinct_set: assumes "distincts xss" shows "distinct (map set xss)" using assms by (auto simp: distinct_map distincts_distinct distincts_inj_on_set) lemma distincts_distinct_nth: assumes "distincts xss" "n < length xss" shows "distinct (xss ! n)" using assms by (auto simp: distincts_def) lemma lists_succ_not_in: assumes "x \ (\xs\set xss. set xs)" shows "lists_succ xss x = x" using assms by (induct xss) (auto simp: list_succ_not_in) lemma lists_succ_in_conv: "lists_succ xss x \ (\xs\set xss. set xs) \ x \ (\xs\set xss. set xs)" by (induct xss) (auto simp: list_succ_in_conv lists_succ_not_in list_succ_not_in) lemma lists_succ_in_conv1: assumes "A \ (\xs\set xss. set xs) = {}" shows "lists_succ xss x \ A \ x \ A" by (metis Int_iff assms emptyE lists_succ_in_conv lists_succ_not_in) lemma lists_succ_Cons_pf: "lists_succ (xs # xss) = list_succ xs o lists_succ xss" by auto lemma lists_succ_Nil_pf: "lists_succ [] = id" by (simp add: fun_eq_iff) lemmas lists_succ_simps_pf = lists_succ_Cons_pf lists_succ_Nil_pf lemma lists_succ_permutes: assumes "distincts xss" shows "lists_succ xss permutes (\xs \ set xss. set xs)" using assms proof (induction xss) case Nil then show ?case by auto next case (Cons xs xss) have "list_succ xs permutes (set xs)" using Cons by (intro list_succ_permutes) (simp add: distincts_def in_set_member) moreover have "lists_succ xss permutes (\ys \ set xss. set ys)" using Cons by (auto simp: Cons distincts_def) ultimately show "lists_succ (xs # xss) permutes (\ys \ set (xs # xss). set ys)" using Cons by (auto simp: lists_succ_Cons_pf intro: permutes_compose permutes_subset) qed lemma bij_lists_succ: "distincts xss \ bij (lists_succ xss)" by (induct xss) (auto simp: lists_succ_simps_pf bij_comp bij_list_succ distincts_Cons) lemma lists_succ_snoc: "lists_succ (xss @ [xs]) = lists_succ xss o list_succ xs" by (induct xss) auto lemma inv_lists_succ_eq: assumes "distincts xss" shows "inv (lists_succ xss) = lists_succ (rev (map rev xss))" proof - have *: "\f g. inv (\b. f (g b)) = inv (f o g)" by (simp add: o_def) have **: "lists_succ [] = id" by auto show ?thesis using assms by (induct xss) (auto simp: * ** lists_succ_snoc lists_succ_Cons_pf o_inv_distrib inv_list_succ_eq distincts_Cons bij_list_succ bij_lists_succ) qed lemma lists_succ_remove1: assumes "distincts xss" "xs \ set xss" shows "lists_succ (xs # remove1 xs xss) = lists_succ xss" using assms proof (induct xss) case Nil then show ?case by simp next case (Cons ys xss) show ?case proof cases assume "xs = ys" then show ?case by simp next assume "xs \ ys" with Cons.prems have inter: "set xs \ set ys = {}" and "xs \ set xss" by (auto simp: distincts_Cons) have dists: "distincts (xs # remove1 xs xss)" "distincts (xs # ys # remove1 xs xss)" using \distincts (ys # xss)\ \xs \ set xss\ by (auto simp: distincts_def) have "list_succ xs \ (list_succ ys \ lists_succ (remove1 xs xss)) = list_succ ys \ (list_succ xs \ lists_succ (remove1 xs xss))" using inter unfolding fun_eq_iff comp_def by (subst list_succ_commute) auto also have "\ = list_succ ys o (lists_succ (xs # remove1 xs xss))" using dists by (simp add: lists_succ_Cons_pf distincts_Cons) also have "\ = list_succ ys o lists_succ xss" using \xs \ set xss\ \distincts (ys # xss)\ by (simp add: distincts_Cons Cons.hyps) finally show "lists_succ (xs # remove1 xs (ys # xss)) = lists_succ (ys # xss)" using Cons dists by (auto simp: lists_succ_Cons_pf distincts_Cons) qed qed lemma lists_succ_no_order: assumes "distincts xss" "distincts yss" "set xss = set yss" shows "lists_succ xss = lists_succ yss" using assms proof (induct xss arbitrary: yss) case Nil then show ?case by simp next case (Cons xs xss) have "xs \ set xss" "xs \ set yss" using Cons.prems by (auto dest: distincts_distinct) have "lists_succ xss = lists_succ (remove1 xs yss)" using Cons.prems \xs \ _\ by (intro Cons.hyps) (auto simp add: distincts_Cons distincts_remove1 distincts_distinct) then have "lists_succ (xs # xss) = lists_succ (xs # remove1 xs yss)" using Cons.prems \xs \ _\ by (simp add: lists_succ_Cons_pf distincts_Cons_remove1) then show ?case using Cons.prems \xs \ _\ by (simp add: lists_succ_remove1) qed section \List Orbits\ text \Computes the orbit of @{term x} under @{term f}\ definition orbit_list :: "('a \ 'a) \ 'a \ 'a list" where "orbit_list f x \ iterate 0 (funpow_dist1 f x x) f x" partial_function (tailrec) orbit_list_impl :: "('a \ 'a) \ 'a \ 'a list \ 'a \ 'a list" where "orbit_list_impl f s acc x = (let x' = f x in if x' = s then rev (x # acc) else orbit_list_impl f s (x # acc) x')" context notes [simp] = length_fold_remove1_le begin text \Computes the list of orbits\ fun orbits_list :: "('a \ 'a) \ 'a list \ 'a list list" where "orbits_list f [] = []" | "orbits_list f (x # xs) = orbit_list f x # orbits_list f (fold remove1 (orbit_list f x) xs)" fun orbits_list_impl :: "('a \ 'a) \ 'a list \ 'a list list" where "orbits_list_impl f [] = []" | "orbits_list_impl f (x # xs) = (let fc = orbit_list_impl f x [] x in fc # orbits_list_impl f (fold remove1 fc xs))" declare orbit_list_impl.simps[code] end abbreviation sset :: "'a list list \ 'a set set" where "sset xss \ set ` set xss" lemma iterate_funpow_step: assumes "f x \ y" "y \ orbit f x" shows "iterate 0 (funpow_dist1 f x y) f x = x # iterate 0 (funpow_dist1 f (f x) y) f (f x)" proof - from assms have A: "y \ orbit f (f x)" by (simp add: orbit_step) have "iterate 0 (funpow_dist1 f x y) f x = x # iterate 1 (funpow_dist1 f x y) f x" (is "_ = _ # ?it") unfolding iterate_def by (rewrite in "\ = _" upt_conv_Cons) auto also have "?it = map (\n. (f ^^ n) x) (map Suc [0.. = map (\n. (f ^^ n) (f x)) [0.. = iterate 0 (funpow_dist1 f (f x) y) f (f x)" unfolding iterate_def unfolding iterate_def by (simp add: funpow_dist_step[OF assms(1) A]) finally show ?thesis . qed lemma orbit_list_impl_conv: assumes "y \ orbit f x" shows "orbit_list_impl f y acc x = rev acc @ iterate 0 (funpow_dist1 f x y) f x" using assms proof (induct n\"funpow_dist1 f x y" arbitrary: x acc) case (Suc x) show ?case proof cases assume "f x = y" then show ?thesis by (subst orbit_list_impl.simps) (simp add: Let_def iterate_def funpow_dist_0) next assume not_y :"f x \ y" have y_in_succ: "y \ orbit f (f x)" by (intro orbit_step Suc.prems not_y) have "orbit_list_impl f y acc x = orbit_list_impl f y (x # acc) (f x)" using not_y by (subst orbit_list_impl.simps) simp also have "\ = rev (x # acc) @ iterate 0 (funpow_dist1 f (f x) y) f (f x)" (is "_ = ?rev @ ?it") by (intro Suc funpow_dist_step not_y y_in_succ) also have "\ = rev acc @ iterate 0 (funpow_dist1 f x y) f x" using not_y Suc.prems by (simp add: iterate_funpow_step) finally show ?thesis . qed qed lemma orbit_list_conv_impl: assumes "x \ orbit f x" shows "orbit_list f x = orbit_list_impl f x [] x" unfolding orbit_list_impl_conv[OF assms] orbit_list_def by simp lemma set_orbit_list: assumes "x \ orbit f x" shows "set (orbit_list f x) = orbit f x" by (simp add: orbit_list_def orbit_conv_funpow_dist1[OF assms] set_iterate) lemma set_orbit_list': assumes "permutation f" shows "set (orbit_list f x) = orbit f x" using assms by (simp add: permutation_self_in_orbit set_orbit_list) lemma distinct_orbit_list: assumes "x \ orbit f x" shows "distinct (orbit_list f x)" by (simp del: upt_Suc add: orbit_list_def iterate_def distinct_map inj_on_funpow_dist1[OF assms]) lemma distinct_orbit_list': assumes "permutation f" shows "distinct (orbit_list f x)" using assms by (simp add: permutation_self_in_orbit distinct_orbit_list) lemma orbits_list_conv_impl: assumes "permutation f" shows "orbits_list f xs = orbits_list_impl f xs" proof (induct "length xs" arbitrary: xs rule: less_induct) case less show ?case using assms by (cases xs) (auto simp: assms less less_Suc_eq_le length_fold_remove1_le orbit_list_conv_impl permutation_self_in_orbit Let_def) qed lemma orbit_list_not_nil[simp]: "orbit_list f x \ []" by (simp add: orbit_list_def) lemma sset_orbits_list: assumes "permutation f" shows "sset (orbits_list f xs) = (orbit f) ` set xs" proof (induct "length xs" arbitrary: xs rule: less_induct) case less show ?case proof (cases xs) case Nil then show ?thesis by simp next case (Cons x' xs') let ?xs'' = "fold remove1 (orbit_list f x') xs'" have A: "sset (orbits_list f ?xs'') = orbit f ` set ?xs''" using Cons by (simp add: less_Suc_eq_le length_fold_remove1_le less.hyps) have B: "set (orbit_list f x') = orbit f x'" by (rule set_orbit_list) (simp add: permutation_self_in_orbit assms) have "orbit f ` set (fold remove1 (orbit_list f x') xs') \ orbit f ` set xs'" using set_fold_remove1[of _ xs'] by auto moreover have "orbit f ` set xs' - {orbit f x'} \ (orbit f ` set (fold remove1 (orbit_list f x') xs'))" (is "?L \ ?R") proof fix A assume "A \ ?L" then obtain y where "A = orbit f y" "y \ set xs'" by auto have "A \ orbit f x'" using \A \ ?L\ by auto from \A = _\ \A \ _\ have "y \ orbit f x'" by (meson assms cyclic_on_orbit orbit_cyclic_eq3 permutation_permutes) with \y \ _\ have "y \ set (fold remove1 (orbit_list f x') xs')" by (auto simp: set_fold_remove1' set_orbit_list permutation_self_in_orbit assms) then show "A \ ?R" using \A = _\ by auto qed ultimately show ?thesis by (auto simp: A B Cons) qed qed subsection \Relation to @{term cyclic_on}\ lemma list_succ_orbit_list: assumes "s \ orbit f s" "\x. x \ orbit f s \ f x = x" shows "list_succ (orbit_list f s) = f" proof - have "distinct (orbit_list f s)" "\x. x \ set (orbit_list f s) \ x = f x" using assms by (simp_all add: distinct_orbit_list set_orbit_list) moreover have "\i. i < length (orbit_list f s) \ orbit_list f s ! (Suc i mod length (orbit_list f s)) = f (orbit_list f s ! i)" using funpow_dist1_prop[OF \s \ orbit f s\] by (auto simp: orbit_list_def funpow_mod_eq) ultimately show ?thesis by (auto simp: list_succ_def fun_eq_iff) qed lemma list_succ_funpow_conv: assumes A: "distinct xs" "x \ set xs" shows "(list_succ xs ^^ n) x = xs ! ((index xs x + n) mod length xs)" proof - have "xs \ []" using assms by auto then show ?thesis by (induct n) (auto simp: hd_conv_nth A index_nth_id list_succ_def mod_simps) qed lemma orbit_list_succ: assumes "distinct xs" "x \ set xs" shows "orbit (list_succ xs) x = set xs" proof (intro set_eqI iffI) fix y assume "y \ orbit (list_succ xs) x" then show "y \ set xs" by induct (auto simp: list_succ_in_conv \x \ set xs\) next fix y assume "y \ set xs" moreover { fix i j have "i < length xs \ j < length xs \ \n. xs ! j = xs ! ((i + n) mod length xs)" using assms by (auto simp: exI[where x="j + (length xs - i)"]) } ultimately show "y \ orbit (list_succ xs) x" using assms by (auto simp: orbit_altdef_permutation permutation_list_succ list_succ_funpow_conv index_nth_id in_set_conv_nth) qed lemma cyclic_on_list_succ: assumes "distinct xs" "xs \ []" shows "cyclic_on (list_succ xs) (set xs)" using assms last_in_set by (auto simp: cyclic_on_def orbit_list_succ) lemma obtain_orbit_list_func: assumes "s \ orbit f s" "\x. x \ orbit f s \ f x = x" obtains xs where "f = list_succ xs" "set xs = orbit f s" "distinct xs" "hd xs = s" proof - { from assms have "f = list_succ (orbit_list f s)" by (simp add: list_succ_orbit_list) moreover have "set (orbit_list f s) = orbit f s" "distinct (orbit_list f s)" by (auto simp: set_orbit_list distinct_orbit_list assms) moreover have "hd (orbit_list f s) = s" by (simp add: orbit_list_def iterate_def hd_map del: upt_Suc) ultimately have "\xs. f = list_succ xs \ set xs = orbit f s \ distinct xs \ hd xs = s" by blast } then show ?thesis by (metis that) qed lemma cyclic_on_obtain_list_succ: assumes "cyclic_on f S" "\x. x \ S \ f x = x" obtains xs where "f = list_succ xs" "set xs = S" "distinct xs" proof - from assms obtain s where s: "s \ orbit f s" "\x. x \ orbit f s \ f x = x" "S = orbit f s" by (auto simp: cyclic_on_def) then show ?thesis by (metis that obtain_orbit_list_func) qed lemma cyclic_on_obtain_list_succ': assumes "cyclic_on f S" "f permutes S" obtains xs where "f = list_succ xs" "set xs = S" "distinct xs" using assms unfolding permutes_def by (metis cyclic_on_obtain_list_succ) lemma list_succ_unique: assumes "s \ orbit f s" "\x. x \ orbit f s \ f x = x" shows "\!xs. f = list_succ xs \ distinct xs \ hd xs = s \ set xs = orbit f s" proof - from assms obtain xs where xs: "f = list_succ xs" "distinct xs" "hd xs = s" "set xs = orbit f s" by (rule obtain_orbit_list_func) moreover { fix zs assume A: "f = list_succ zs" "distinct zs" "hd zs = s" "set zs = orbit f s" then have "zs \ []" using \s \ orbit f s\ by auto from \distinct xs\ \distinct zs\ \set xs = orbit f s\ \set zs = orbit f s\ have len: "length xs = length zs" by (metis distinct_card) { fix n assume "n < length xs" then have "zs ! n = xs ! n" proof (induct n) case 0 with A xs \zs \ []\ show ?case by (simp add: hd_conv_nth nth_rotate_conv_nth) next case (Suc n) then have "list_succ zs (zs ! n) = list_succ xs (xs! n)" using \f = list_succ xs\ \f = list_succ zs\ by simp with \Suc n < _\ show ?case by (simp add:list_succ_nth len \distinct xs\ \distinct zs\) qed } then have "zs = xs" by (metis len nth_equalityI) } ultimately show ?thesis by metis qed lemma distincts_orbits_list: assumes "distinct as" "permutation f" shows "distincts (orbits_list f as)" using assms(1) proof (induct "length as" arbitrary: as rule: less_induct) case less show ?case proof (cases as) case Nil then show ?thesis by simp next case (Cons a as') let ?as' = "fold remove1 (orbit_list f a) as'" from Cons less.prems have A: "distincts (orbits_list f (fold remove1 (orbit_list f a) as'))" by (intro less) (auto simp: distinct_fold_remove1 length_fold_remove1_le less_Suc_eq_le) have B: "set (orbit_list f a) \ \(sset (orbits_list f (fold remove1 (orbit_list f a) as'))) = {}" proof - have "orbit f a \ set (fold remove1 (orbit_list f a) as') = {}" using assms less.prems Cons by (simp add: set_fold_remove1_distinct set_orbit_list') then have "orbit f a \ \ (orbit f ` set (fold remove1 (orbit_list f a) as')) = {}" by auto (metis assms(2) cyclic_on_orbit disjoint_iff_not_equal permutation_self_in_orbit[OF assms(2)] orbit_cyclic_eq3 permutation_permutes) then show ?thesis using assms by (auto simp: set_orbit_list' sset_orbits_list disjoint_iff_not_equal) qed show ?thesis using A B assms by (auto simp: distincts_Cons Cons distinct_orbit_list') qed qed lemma cyclic_on_lists_succ': assumes "distincts xss" shows "A \ sset xss \ cyclic_on (lists_succ xss) A" using assms proof (induction xss arbitrary: A) case Nil then show ?case by auto next case (Cons xs xss A) then have inter: "set xs \ (\ys\set xss. set ys) = {}" by (auto simp: distincts_Cons) note pcp[OF _ _ inter] = permutes_comp_preserves_cyclic1 permutes_comp_preserves_cyclic2 from Cons show "cyclic_on (lists_succ (xs # xss)) A" by (cases "A = set xs") (auto intro: pcp simp: cyclic_on_list_succ list_succ_permutes lists_succ_permutes lists_succ_Cons_pf distincts_Cons) qed lemma cyclic_on_lists_succ: assumes "distincts xss" shows "\xs. xs \ set xss \ cyclic_on (lists_succ xss) (set xs)" using assms by (auto intro: cyclic_on_lists_succ') lemma permutes_as_lists_succ: assumes "distincts xss" assumes ls_eq: "\xs. xs \ set xss \ list_succ xs = perm_restrict f (set xs)" assumes "f permutes (\(sset xss))" shows "f = lists_succ xss" using assms proof (induct xss arbitrary: f) case Nil then show ?case by simp next case (Cons xs xss) let ?sets = "\xss. \ys \ set xss. set ys" have xs: "distinct xs" "xs \ []" using Cons by (auto simp: distincts_Cons) have f_xs: "perm_restrict f (set xs) = list_succ xs" using Cons by simp have co_xs: "cyclic_on (perm_restrict f (set xs)) (set xs)" unfolding f_xs using xs by (rule cyclic_on_list_succ) have perm_xs: "perm_restrict f (set xs) permutes set xs" unfolding f_xs using \distinct xs\ by (rule list_succ_permutes) have perm_xss: "perm_restrict f (?sets xss) permutes (?sets xss)" proof - have "perm_restrict f (?sets (xs # xss) - set xs) permutes (?sets (xs # xss) - set xs)" using Cons co_xs by (intro perm_restrict_diff_cyclic) (auto simp: cyclic_on_perm_restrict) also have "?sets (xs # xss) - set xs = ?sets xss" using Cons by (auto simp: distincts_Cons) finally show ?thesis . qed have f_xss: "perm_restrict f (?sets xss) = lists_succ xss" proof - have *: "\xs. xs \ set xss \ ((\x\set xss. set x) \ set xs) = set xs" by blast with perm_xss Cons.prems show ?thesis by (intro Cons.hyps) (auto simp: distincts_Cons perm_restrict_perm_restrict *) qed from Cons.prems show "f = lists_succ (xs # xss)" by (simp add: lists_succ_Cons_pf distincts_Cons f_xss[symmetric] perm_restrict_union perm_xs perm_xss) qed lemma cyclic_on_obtain_lists_succ: assumes permutes: "f permutes S" and S: "S = \(sset css)" and dists: "distincts css" and cyclic: "\cs. cs \ set css \ cyclic_on f (set cs)" obtains xss where "f = lists_succ xss" "distincts xss" "map set xss = map set css" "map hd xss = map hd css" proof - let ?fc = "\cs. perm_restrict f (set cs)" define some_list where "some_list cs = (SOME xs. ?fc cs = list_succ xs \ set xs = set cs \ distinct xs \ hd xs = hd cs)" for cs { fix cs assume "cs \ set css" then have "cyclic_on (?fc cs) (set cs)" "\x. x \ set cs \ ?fc cs x = x" "hd cs \ set cs" using cyclic dists by (auto simp add: cyclic_on_perm_restrict perm_restrict_def distincts_def) then have "hd cs \ orbit (?fc cs) (hd cs)" "\x. x \ orbit (?fc cs) (hd cs) \ ?fc cs x = x" "hd cs \ set cs" "set cs = orbit (?fc cs) (hd cs)" by (auto simp: cyclic_on_alldef) then have "\xs. ?fc cs = list_succ xs \ set xs = set cs \ distinct xs \ hd xs = hd cs" by (metis obtain_orbit_list_func) then have "?fc cs = list_succ (some_list cs) \ set (some_list cs) = set cs \ distinct (some_list cs) \ hd (some_list cs) = hd cs" unfolding some_list_def by (rule someI_ex) then have "?fc cs = list_succ (some_list cs)" "set (some_list cs) = set cs" "distinct (some_list cs)" "hd (some_list cs) = hd cs" by auto } note sl_cs = this have "\cs. cs \ set css \ cs \ []" using dists by (auto simp: distincts_def) then have some_list_ne: "\cs. cs \ set css \ some_list cs \ []" by (metis set_empty sl_cs(2)) have set: "map set (map some_list css) = map set css" "map hd (map some_list css) = map hd css" using sl_cs(2,4) by (auto simp add: map_idI) have distincts: "distincts (map some_list css)" proof - have c_dist: "\xs ys. \xs\set css; ys\set css; xs \ ys\ \ set xs \ set ys = {}" using dists by (auto simp: distincts_def) have "distinct (map some_list css)" proof - have "inj_on some_list (set css)" using sl_cs(2) c_dist by (intro inj_onI) (metis inf.idem set_empty) with \distincts css\ show ?thesis by (auto simp: distincts_distinct distinct_map) qed moreover have "\xs\set (map some_list css). distinct xs \ xs \ []" using sl_cs(3) some_list_ne by auto moreover from c_dist have "(\xs\set (map some_list css). \ys\set (map some_list css). xs \ ys \ set xs \ set ys = {})" using sl_cs(2) by auto ultimately show ?thesis by (simp add: distincts_def) qed have f: "f = lists_succ (map some_list css)" using distincts proof (rule permutes_as_lists_succ) fix xs assume "xs \ set (map some_list css)" then show "list_succ xs = perm_restrict f (set xs)" using sl_cs(1) sl_cs(2) by auto next have "S = (\xs\set (map some_list css). set xs)" using S sl_cs(2) by auto with permutes show "f permutes \(sset (map some_list css))" by simp qed from f distincts set show ?thesis .. qed subsection \Permutations of a List\ lemma length_remove1_less: assumes "x \ set xs" shows "length (remove1 x xs) < length xs" proof - from assms have "0 < length xs" by auto with assms show ?thesis by (auto simp: length_remove1) qed context notes [simp] = length_remove1_less begin fun permutations :: "'a list \ 'a list list" where permutations_Nil: "permutations [] = [[]]" | permutations_Cons: "permutations xs = [y # ys. y <- xs, ys <- permutations (remove1 y xs)]" end declare permutations_Cons[simp del] text \ The function above returns all permutations of a list. The function below computes only those which yield distinct cyclic permutation functions (cf. @{term list_succ}). \ fun cyc_permutations :: "'a list \ 'a list list" where "cyc_permutations [] = [[]]" | "cyc_permutations (x # xs) = map (Cons x) (permutations xs)" lemma nil_in_permutations[simp]: "[] \ set (permutations xs) \ xs = []" by (induct xs) (auto simp: permutations_Cons) lemma permutations_not_nil: assumes "xs \ []" shows "permutations xs = concat (map (\x. map ((#) x) (permutations (remove1 x xs))) xs)" using assms by (cases xs) (auto simp: permutations_Cons) lemma set_permutations_step: assumes "xs \ []" shows "set (permutations xs) = (\x \ set xs. Cons x ` set (permutations (remove1 x xs)))" using assms by (cases xs) (auto simp: permutations_Cons) lemma in_set_permutations: assumes "distinct xs" shows "ys \ set (permutations xs) \ distinct ys \ set xs = set ys" (is "?L xs ys \ ?R xs ys") using assms proof (induct "length xs" arbitrary: xs ys) case 0 then show ?case by auto next case (Suc n) then have "xs \ []" by auto show ?case proof assume "?L xs ys" then obtain y ys' where "ys = y # ys'" "y \ set xs" "ys' \ set (permutations (remove1 (hd ys) xs))" using \xs \ []\ by (auto simp: permutations_not_nil) moreover then have "?R (remove1 y xs) ys'" using Suc.prems Suc.hyps(2) by (intro Suc.hyps(1)[THEN iffD1]) (auto simp: length_remove1) ultimately show "?R xs ys" using Suc by auto next assume "?R xs ys" with \xs \ []\ obtain y ys' where "ys = y # ys'" "y \ set xs" by (cases ys) auto moreover then have "ys' \ set (permutations (remove1 y xs))" using Suc \?R xs ys\ by (intro Suc.hyps(1)[THEN iffD2]) (auto simp: length_remove1) ultimately show "?L xs ys" using \xs \ []\ by (auto simp: permutations_not_nil) qed qed lemma in_set_cyc_permutations: assumes "distinct xs" shows "ys \ set (cyc_permutations xs) \ distinct ys \ set xs = set ys \ hd ys = hd xs" (is "?L xs ys \ ?R xs ys") proof (cases xs) case (Cons x xs) with assms show ?thesis by (cases ys) (auto simp: in_set_permutations intro!: imageI) qed auto lemma in_set_cyc_permutations_obtain: assumes "distinct xs" "distinct ys" "set xs = set ys" obtains n where "rotate n ys \ set (cyc_permutations xs)" proof (cases xs) case Nil with assms have "rotate 0 ys \ set (cyc_permutations xs)" by auto then show ?thesis .. next case (Cons x xs') let ?ys' = "rotate (index ys x) ys" have "ys \ []" "x \ set ys" using Cons assms by auto then have "distinct ?ys' \ set xs = set ?ys' \ hd ?ys' = hd xs" using assms Cons by (auto simp add: hd_rotate_conv_nth) with \distinct xs\ have "?ys' \ set (cyc_permutations xs)" by (rule in_set_cyc_permutations[THEN iffD2]) then show ?thesis .. qed lemma list_succ_set_cyc_permutations: assumes "distinct xs" "xs \ []" shows "list_succ ` set (cyc_permutations xs) = {f. f permutes set xs \ cyclic_on f (set xs)}" (is "?L = ?R") proof (intro set_eqI iffI) fix f assume "f \ ?L" moreover have "\ys. set xs = set ys \ xs \ [] \ ys \ []" by auto ultimately show "f \ ?R" using assms by (auto simp: in_set_cyc_permutations list_succ_permutes cyclic_on_list_succ) next fix f assume "f \ ?R" then obtain ys where ys: "list_succ ys = f" "distinct ys" "set ys = set xs" by (auto elim: cyclic_on_obtain_list_succ') moreover with \distinct xs\ obtain n where "rotate n ys \ set (cyc_permutations xs)" by (auto elim: in_set_cyc_permutations_obtain) then have "list_succ (rotate n ys) \ ?L" by simp ultimately show "f \ ?L" by simp qed subsection \Enumerating Permutations from List Orbits\ definition cyc_permutationss :: "'a list list \ 'a list list list" where "cyc_permutationss = product_lists o map cyc_permutations" lemma cyc_permutationss_Nil[simp]: "cyc_permutationss [] = [[]]" by (auto simp: cyc_permutationss_def) lemma in_set_cyc_permutationss: assumes "distincts xss" shows "yss \ set (cyc_permutationss xss) \ distincts yss \ map set xss = map set yss \ map hd xss = map hd yss" proof - { assume A: "list_all2 (\x ys. x \ set ys) yss (map cyc_permutations xss)" then have "length yss = length xss" by (auto simp: list_all2_lengthD) then have "\(sset xss) = \(sset yss)" "distincts yss" "map set xss = map set yss" "map hd xss = map hd yss" using A assms by (induct yss xss rule: list_induct2) (auto simp: distincts_Cons in_set_cyc_permutations) } note X = this { assume A: "distincts yss" "map set xss = map set yss" "map hd xss = map hd yss" then have "length yss = length xss" by (auto dest: map_eq_imp_length_eq) then have "list_all2 (\x ys. x \ set ys) yss (map cyc_permutations xss)" using A assms by (induct yss xss rule: list_induct2) (auto simp: distincts_Cons in_set_cyc_permutations) } note Y = this show "?thesis" unfolding cyc_permutationss_def by (auto simp: product_lists_set intro: X Y) qed lemma lists_succ_set_cyc_permutationss: assumes "distincts xss" shows "lists_succ ` set (cyc_permutationss xss) = {f. f permutes \(sset xss) \ (\c \ sset xss. cyclic_on f c)}" (is "?L = ?R") using assms proof (intro set_eqI iffI) fix f assume "f \ ?L" then obtain yss where "yss \ set (cyc_permutationss xss)" "f = lists_succ yss" by (rule imageE) moreover from \yss \ _\ assms have "set (map set xss) = set (map set yss)" by (auto simp: in_set_cyc_permutationss) then have "sset xss = sset yss" by simp ultimately show "f \ ?R" using assms by (auto simp: in_set_cyc_permutationss cyclic_on_lists_succ') (metis lists_succ_permutes) next fix f assume "f \ ?R" then have "f permutes \(sset xss)" "\cs. cs \ set xss \ cyclic_on f (set cs)" by auto from this(1) refl assms this(2) obtain yss where "f = lists_succ yss" "distincts yss" "map set yss = map set xss" "map hd yss = map hd xss" by (rule cyclic_on_obtain_lists_succ) with assms show "f \ ?L" by (auto intro!: imageI simp: in_set_cyc_permutationss) qed subsection \Lists of Permutations\ definition permutationss :: "'a list list \ 'a list list list" where "permutationss = product_lists o map permutations" lemma permutationss_Nil[simp]: "permutationss [] = [[]]" by (auto simp: permutationss_def) lemma permutationss_Cons: "permutationss (xs # xss) = concat (map (\ys. map (Cons ys) (permutationss xss)) (permutations xs))" by (auto simp: permutationss_def) lemma in_set_permutationss: assumes "distincts xss" shows "yss \ set (permutationss xss) \ distincts yss \ map set xss = map set yss" proof - { assume A: "list_all2 (\x ys. x \ set ys) yss (map permutations xss)" then have "length yss = length xss" by (auto simp: list_all2_lengthD) then have "\(sset xss) = \(sset yss)" "distincts yss" "map set xss = map set yss" using A assms by (induct yss xss rule: list_induct2) (auto simp: distincts_Cons in_set_permutations) } note X = this { assume A: "distincts yss" "map set xss = map set yss" then have "length yss = length xss" by (auto dest: map_eq_imp_length_eq) then have "list_all2 (\x ys. x \ set ys) yss (map permutations xss)" using A assms by (induct yss xss rule: list_induct2) (auto simp: in_set_permutations distincts_Cons) } note Y = this show "?thesis" unfolding permutationss_def by (auto simp: product_lists_set intro: X Y) qed lemma set_permutationss: assumes "distincts xss" shows "set (permutationss xss) = {yss. distincts yss \ map set xss = map set yss}" using in_set_permutationss[OF assms] by blast lemma permutationss_complete: assumes "distincts xss" "distincts yss" "xss \ []" and "set ` set xss = set ` set yss" shows "set yss \ set ` set (permutationss xss)" proof - have "length xss = length yss" using assms by (simp add: length_distincts) from \sset xss = _\ have "\yss'. set yss' = set yss \ map set yss' = map set xss" using assms(1-2) proof (induct xss arbitrary: yss) case Nil then show ?case by simp next case (Cons xs xss) from \sset (xs # xss) = sset yss\ obtain ys where ys: "ys \ set yss" "set ys = set xs" by auto (metis imageE insertI1) with \distincts yss\ have "set ys \ sset (remove1 ys yss)" by (fastforce simp: distincts_def) moreover from \distincts (xs # xss)\ have "set xs \ sset xss" by (fastforce simp: distincts_def) ultimately have "sset xss = sset (remove1 ys yss)" using \distincts yss\ \sset (xs # xss) = sset yss\ apply (auto simp: distincts_distinct \set ys = set xs\[symmetric]) apply (smt Diff_insert_absorb \ys \ set yss\ image_insert insert_Diff rev_image_eqI) by (metis \ys \ set yss\ image_eqI insert_Diff insert_iff) then obtain yss' where "set yss' = set (remove1 ys yss) \ map set yss' = map set xss" using Cons by atomize_elim (auto simp: distincts_Cons distincts_remove1) then have "set (ys # yss') = set yss \ map set (ys # yss') = map set (xs # xss)" using ys set_remove1_eq \distincts yss\ by (auto simp: distincts_distinct) then show ?case .. qed then obtain yss' where "set yss' = set yss" "map set yss' = map set xss" by blast then have "distincts yss'" using \distincts xss\ \distincts yss\ unfolding distincts_def by simp_all (metis \length xss = length yss\ card_distinct distinct_card length_map) then have "set yss' \ set ` set (permutationss xss)" using \distincts xss\ \map set yss' = _\ by (auto simp: set_permutationss) then show ?thesis using \set yss' = _\ by auto qed lemma permutations_complete: (* could generalize with multi-sets *) assumes "distinct xs" "distinct ys" "set xs = set ys" shows "ys \ set (permutations xs)" using assms proof (induct "length xs" arbitrary: xs ys) case 0 then show ?case by simp next case (Suc n) from Suc.hyps have "xs \ []" by auto then obtain y ys' where [simp]: "ys = y # ys'" "y \ set xs" using Suc.prems by (cases ys) auto have "ys' \ set (permutations (remove1 y xs))" using Suc.prems \Suc n = _\ by (intro Suc.hyps) (simp_all add: length_remove1 ) then show ?case using \xs \ []\ by (auto simp: set_permutations_step) qed end diff --git a/thys/Planarity_Certificates/Planarity/Graph_Genus.thy b/thys/Planarity_Certificates/Planarity/Graph_Genus.thy --- a/thys/Planarity_Certificates/Planarity/Graph_Genus.thy +++ b/thys/Planarity_Certificates/Planarity/Graph_Genus.thy @@ -1,538 +1,538 @@ theory Graph_Genus imports + "HOL-Combinatorics.Permutations" Graph_Theory.Graph_Theory - "HOL-Library.Permutations" begin lemma nat_diff_mod_right: fixes a b c :: nat assumes "b < a" shows "(a - b) mod c = (a - b mod c) mod c" proof - from assms have b_mod: "b mod c \ a" by (metis mod_less_eq_dividend linear not_le order_trans) have "int ((a - b) mod c) = (int a - int b mod int c) mod int c" using assms by (simp add: zmod_int of_nat_diff mod_simps) also have "\ = int ((a - b mod c) mod c)" using assms b_mod by (simp add: zmod_int [symmetric] of_nat_diff [symmetric]) finally show ?thesis by simp qed lemma inj_on_f_imageI: assumes "inj_on f S" "\t. t \ T \ t \ S" shows "inj_on ((`) f) T" using assms by (auto simp: inj_on_image_eq_iff intro: inj_onI) section \Combinatorial Maps\ lemma (in bidirected_digraph) has_dom_arev: "has_dom arev (arcs G)" using arev_dom by (auto simp: has_dom_def) record 'b pre_map = edge_rev :: "'b \ 'b" edge_succ :: "'b \ 'b" definition edge_pred :: "'b pre_map \ 'b \ 'b" where "edge_pred M = inv (edge_succ M)" locale pre_digraph_map = pre_digraph + fixes M :: "'b pre_map" locale digraph_map = fin_digraph G + pre_digraph_map G M + bidirected_digraph G "edge_rev M" for G M + assumes edge_succ_permutes: "edge_succ M permutes arcs G" assumes edge_succ_cyclic: "\v. v \ verts G \ out_arcs G v \ {} \ cyclic_on (edge_succ M) (out_arcs G v)" lemma (in fin_digraph) digraph_mapI: assumes bidi: "\a. a \ arcs G \ edge_rev M a = a" "\a. a \ arcs G \ edge_rev M a \ a" "\a. a \ arcs G \ edge_rev M (edge_rev M a) = a" "\a. a \ arcs G \ tail G (edge_rev M a) = head G a" assumes edge_succ_permutes: "edge_succ M permutes arcs G" assumes edge_succ_cyclic: "\v. v \ verts G \ out_arcs G v \ {} \ cyclic_on (edge_succ M) (out_arcs G v)" shows "digraph_map G M" using assms by unfold_locales auto lemma (in fin_digraph) digraph_mapI_permutes: assumes bidi: "edge_rev M permutes arcs G" "\a. a \ arcs G \ edge_rev M a \ a" "\a. a \ arcs G \ edge_rev M (edge_rev M a) = a" "\a. a \ arcs G \ tail G (edge_rev M a) = head G a" assumes edge_succ_permutes: "edge_succ M permutes arcs G" assumes edge_succ_cyclic: "\v. v \ verts G \ out_arcs G v \ {} \ cyclic_on (edge_succ M) (out_arcs G v)" shows "digraph_map G M" proof - interpret bidirected_digraph G "edge_rev M" using bidi by unfold_locales (auto simp: permutes_def) show ?thesis using edge_succ_permutes edge_succ_cyclic by unfold_locales qed context digraph_map begin lemma digraph_map[intro]: "digraph_map G M" by unfold_locales lemma permutation_edge_succ: "permutation (edge_succ M)" by (metis edge_succ_permutes finite_arcs permutation_permutes) lemma edge_pred_succ[simp]: "edge_pred M (edge_succ M a) = a" by (metis edge_pred_def edge_succ_permutes permutes_inverses(2)) lemma edge_succ_pred[simp]: "edge_succ M (edge_pred M a) = a" by (metis edge_pred_def edge_succ_permutes permutes_inverses(1)) lemma edge_pred_permutes: "edge_pred M permutes arcs G" unfolding edge_pred_def using edge_succ_permutes by (rule permutes_inv) lemma permutation_edge_pred: "permutation (edge_pred M)" by (metis edge_pred_permutes finite_arcs permutation_permutes) lemma edge_succ_eq_iff[simp]: "\x y. edge_succ M x = edge_succ M y \ x = y" by (metis edge_pred_succ) lemma edge_rev_in_arcs[simp]: "edge_rev M a \ arcs G \ a \ arcs G" by (metis arev_arev arev_permutes_arcs permutes_not_in) lemma edge_succ_in_arcs[simp]: "edge_succ M a \ arcs G \ a \ arcs G" by (metis edge_pred_succ edge_succ_permutes permutes_not_in) lemma edge_pred_in_arcs[simp]: "edge_pred M a \ arcs G \ a \ arcs G" by (metis edge_succ_pred edge_pred_permutes permutes_not_in) lemma tail_edge_succ[simp]: "tail G (edge_succ M a) = tail G a" proof cases assume "a \ arcs G" then have "tail G a \ verts G" by auto moreover then have "out_arcs G (tail G a) \ {}" using \a \ arcs G\ by auto ultimately have "cyclic_on (edge_succ M) (out_arcs G (tail G a))" by (rule edge_succ_cyclic) moreover have "a \ out_arcs G (tail G a)" using \a \ arcs G\ by simp ultimately have "edge_succ M a \ out_arcs G (tail G a)" by (rule cyclic_on_inI) then show ?thesis by simp next assume "a \ arcs G" then show ?thesis using edge_succ_permutes by (simp add: permutes_not_in) qed lemma tail_edge_pred[simp]: "tail G (edge_pred M a) = tail G a" by (metis edge_succ_pred tail_edge_succ) lemma bij_edge_succ[intro]: "bij (edge_succ M)" using edge_succ_permutes by (simp add: permutes_conv_has_dom) lemma edge_pred_cyclic: assumes "v \ verts G" "out_arcs G v \ {}" shows "cyclic_on (edge_pred M) (out_arcs G v)" proof - obtain a where orb_a_eq: "orbit (edge_succ M) a = out_arcs G v" using edge_succ_cyclic[OF assms] by (auto simp: cyclic_on_def) have "cyclic_on (edge_pred M) (orbit (edge_pred M) a)" using permutation_edge_pred by (rule cyclic_on_orbit') also have "orbit (edge_pred M) a = orbit (edge_succ M) a" unfolding edge_pred_def using permutation_edge_succ by (rule orbit_inv_eq) finally show "cyclic_on (edge_pred M) (out_arcs G v)" by (simp add: orb_a_eq) qed definition (in pre_digraph_map) face_cycle_succ :: "'b \ 'b" where "face_cycle_succ \ edge_succ M o edge_rev M" definition (in pre_digraph_map) face_cycle_pred :: "'b \ 'b" where "face_cycle_pred \ edge_rev M o edge_pred M" lemma face_cycle_pred_succ[simp]: shows "face_cycle_pred (face_cycle_succ a) = a" unfolding face_cycle_pred_def face_cycle_succ_def by simp lemma face_cycle_succ_pred[simp]: shows "face_cycle_succ (face_cycle_pred a) = a" unfolding face_cycle_pred_def face_cycle_succ_def by simp lemma tail_face_cycle_succ: "a \ arcs G \ tail G (face_cycle_succ a) = head G a" by (auto simp: face_cycle_succ_def) lemma funpow_prop: assumes "\x. P (f x) \ P x" shows "P ((f ^^ n) x) \ P x" using assms by (induct n) (auto simp: ) lemma face_cycle_succ_no_arc[simp]: "a \ arcs G \ face_cycle_succ a = a" by (auto simp: face_cycle_succ_def permutes_not_in[OF arev_permutes_arcs] permutes_not_in[OF edge_succ_permutes]) lemma funpow_face_cycle_succ_no_arc[simp]: assumes "a \ arcs G" shows "(face_cycle_succ ^^ n) a = a" using assms by (induct n) auto lemma funpow_face_cycle_pred_no_arc[simp]: assumes "a \ arcs G" shows "(face_cycle_pred ^^ n) a = a" using assms by (induct n) (auto simp: face_cycle_pred_def permutes_not_in[OF arev_permutes_arcs] permutes_not_in[OF edge_pred_permutes]) lemma face_cycle_succ_closed[simp]: "face_cycle_succ a \ arcs G \ a \ arcs G" by (metis comp_apply edge_rev_in_arcs edge_succ_in_arcs face_cycle_succ_def) lemma face_cycle_pred_closed[simp]: "face_cycle_pred a \ arcs G \ a \ arcs G" by (metis face_cycle_succ_closed face_cycle_succ_pred) lemma face_cycle_succ_permutes: "face_cycle_succ permutes arcs G" unfolding face_cycle_succ_def using arev_permutes_arcs edge_succ_permutes by (rule permutes_compose) lemma permutation_face_cycle_succ: "permutation face_cycle_succ" using face_cycle_succ_permutes finite_arcs by (metis permutation_permutes) lemma bij_face_cycle_succ: "bij face_cycle_succ" using face_cycle_succ_permutes by (simp add: permutes_conv_has_dom) lemma face_cycle_pred_permutes: "face_cycle_pred permutes arcs G" unfolding face_cycle_pred_def using edge_pred_permutes arev_permutes_arcs by (rule permutes_compose) definition (in pre_digraph_map) face_cycle_set :: "'b \ 'b set" where "face_cycle_set a = orbit face_cycle_succ a" definition (in pre_digraph_map) face_cycle_sets :: "'b set set" where "face_cycle_sets = face_cycle_set ` arcs G" lemma face_cycle_set_altdef: "face_cycle_set a = {(face_cycle_succ ^^ n) a | n. True}" unfolding face_cycle_set_def by (intro orbit_altdef_self_in permutation_self_in_orbit permutation_face_cycle_succ) lemma face_cycle_set_self[simp, intro]: "a \ face_cycle_set a" unfolding face_cycle_set_def using permutation_face_cycle_succ by (rule permutation_self_in_orbit) lemma empty_not_in_face_cycle_sets: "{} \ face_cycle_sets" by (auto simp: face_cycle_sets_def) lemma finite_face_cycle_set[simp, intro]: "finite (face_cycle_set a)" using face_cycle_set_self unfolding face_cycle_set_def by (simp add: finite_orbit) lemma finite_face_cycle_sets[simp, intro]: "finite face_cycle_sets" by (auto simp: face_cycle_sets_def) lemma face_cycle_set_induct[case_names base step, induct set: face_cycle_set]: assumes consume: "a \ face_cycle_set x" and ih_base: "P x" and ih_step: "\y. y \ face_cycle_set x \ P y \ P (face_cycle_succ y)" shows "P a" using consume unfolding face_cycle_set_def by induct (auto simp: ih_step face_cycle_set_def[symmetric] ih_base ) lemma face_cycle_succ_cyclic: "cyclic_on face_cycle_succ (face_cycle_set a)" unfolding face_cycle_set_def using permutation_face_cycle_succ by (rule cyclic_on_orbit') lemma face_cycle_eq: assumes "b \ face_cycle_set a" shows "face_cycle_set b = face_cycle_set a" using assms unfolding face_cycle_set_def by (auto intro: orbit_swap orbit_trans permutation_face_cycle_succ permutation_self_in_orbit) lemma face_cycle_succ_in_arcsI: "\a. a \ arcs G \ face_cycle_succ a \ arcs G" by (auto simp: face_cycle_succ_def) lemma face_cycle_succ_inI: "\x y. x \ face_cycle_set y \ face_cycle_succ x \ face_cycle_set y" by (metis face_cycle_succ_cyclic cyclic_on_inI) lemma face_cycle_succ_inD: "\x y. face_cycle_succ x \ face_cycle_set y \ x \ face_cycle_set y" by (metis face_cycle_eq face_cycle_set_self face_cycle_succ_inI) lemma face_cycle_set_parts: "face_cycle_set a = face_cycle_set b \ face_cycle_set a \ face_cycle_set b = {}" by (metis disjoint_iff_not_equal face_cycle_eq) definition fc_equiv :: "'b \ 'b \ bool" where "fc_equiv a b \ a \ face_cycle_set b" lemma reflp_fc_equiv: "reflp fc_equiv" by (rule reflpI) (simp add: fc_equiv_def) lemma symp_fc_equiv: "symp fc_equiv" using face_cycle_set_parts by (intro sympI) (auto simp: fc_equiv_def) lemma transp_fc_equiv: "transp fc_equiv" using face_cycle_set_parts by (intro transpI) (auto simp: fc_equiv_def) lemma "equivp fc_equiv" by (intro equivpI reflp_fc_equiv symp_fc_equiv transp_fc_equiv) lemma in_face_cycle_setD: assumes "y \ face_cycle_set x" "x \ arcs G" shows "y \ arcs G" using assms by (auto simp: face_cycle_set_def dest: permutes_orbit_subset[OF face_cycle_succ_permutes]) lemma in_face_cycle_setsD: assumes "x \ face_cycle_sets" shows "x \ arcs G" using assms by (auto simp: face_cycle_sets_def dest: in_face_cycle_setD) end definition (in pre_digraph) isolated_verts :: "'a set" where "isolated_verts \ {v \ verts G. out_arcs G v = {}}" definition (in pre_digraph_map) euler_char :: int where "euler_char \ int (card (verts G)) - int (card (arcs G) div 2) + int (card face_cycle_sets)" definition (in pre_digraph_map) euler_genus :: int where "euler_genus \ (int (2 * card sccs) - int (card isolated_verts) - euler_char) div 2" definition comb_planar :: "('a,'b) pre_digraph \ bool" where "comb_planar G \ \M. digraph_map G M \ pre_digraph_map.euler_genus G M = 0" text \Number of isolated vertices is a graph invariant\ context fixes G hom assumes hom: "pre_digraph.digraph_isomorphism G hom" begin interpretation wf_digraph G using hom by (auto simp: pre_digraph.digraph_isomorphism_def) lemma isolated_verts_app_iso[simp]: "pre_digraph.isolated_verts (app_iso hom G) = iso_verts hom ` isolated_verts" using hom by (auto simp: pre_digraph.isolated_verts_def iso_verts_tail inj_image_mem_iff out_arcs_app_iso_eq) lemma card_isolated_verts_iso[simp]: "card (iso_verts hom ` pre_digraph.isolated_verts G) = card isolated_verts" apply (rule card_image) using hom apply (rule digraph_isomorphism_inj_on_verts[THEN subset_inj_on]) apply (auto simp: isolated_verts_def) done end context digraph_map begin lemma face_cycle_succ_neq: assumes "a \ arcs G" "tail G a \ head G a" shows "face_cycle_succ a \ a " proof - from assms have "edge_rev M a \ arcs G" by (subst edge_rev_in_arcs) simp then have "cyclic_on (edge_succ M) (out_arcs G (tail G (edge_rev M a)))" by (intro edge_succ_cyclic) (auto dest: tail_in_verts simp: out_arcs_def intro: exI[where x="edge_rev M a"]) then have "edge_succ M (edge_rev M a) \ (out_arcs G (tail G (edge_rev M a)))" by (rule cyclic_on_inI) (auto simp: \edge_rev M a \ _\[simplified]) moreover have "tail G (edge_succ M (edge_rev M a)) = head G a" using assms by auto then have "edge_succ M (edge_rev M a) \ a" using assms by metis ultimately show ?thesis using assms by (auto simp: face_cycle_succ_def) qed end section \Maps and Isomorphism\ definition (in pre_digraph) "wrap_iso_arcs hom f = perm_restrict (iso_arcs hom o f o iso_arcs (inv_iso hom)) (arcs (app_iso hom G))" definition (in pre_digraph_map) map_iso :: "('a,'b,'a2,'b2) digraph_isomorphism \ 'b2 pre_map" where "map_iso f \ \ edge_rev = wrap_iso_arcs f (edge_rev M) , edge_succ = wrap_iso_arcs f (edge_succ M) \" lemma funcsetI_permutes: assumes "f permutes S" shows "f \ S \ S" by (metis assms funcsetI permutes_in_image) context fixes G hom assumes hom: "pre_digraph.digraph_isomorphism G hom" begin interpretation wf_digraph G using hom by (auto simp: pre_digraph.digraph_isomorphism_def) lemma wrap_iso_arcs_iso_arcs[simp]: assumes "x \ arcs G" shows "wrap_iso_arcs hom f (iso_arcs hom x) = iso_arcs hom (f x)" using assms hom by (auto simp: wrap_iso_arcs_def perm_restrict_def) lemma inj_on_wrap_iso_arcs: assumes dom: "\f. f \ F \ has_dom f (arcs G)" assumes funcset: "F \ arcs G \ arcs G" shows "inj_on (wrap_iso_arcs hom) F" proof (rule inj_onI) fix f g assume F: "f \ F" "g \ F" and eq: "wrap_iso_arcs hom f = wrap_iso_arcs hom g" { fix x assume "x \ arcs G" then have "f x = x" "g x = x" using F dom by (auto simp: has_dom_def) then have "f x = g x" by simp } moreover { fix x assume "x \ arcs G" then have "f x \ arcs G" "g x \ arcs G" using F funcset by auto with digraph_isomorphism_inj_on_arcs[OF hom] _ have "iso_arcs hom (f x) = iso_arcs hom (g x) \ f x = g x" by (rule inj_onD) then have "f x = g x" using assms hom \x \ arcs G\ eq by (auto simp: wrap_iso_arcs_def fun_eq_iff perm_restrict_def split: if_splits) } ultimately show "f = g" by auto qed lemma inj_on_wrap_iso_arcs_f: assumes "A \ arcs G" "f \ A \ A" "B = iso_arcs hom ` A" assumes "inj_on f A" shows "inj_on (wrap_iso_arcs hom f) B" proof (rule inj_onI) fix x y assume in_hom_A: "x \ B" "y \ B" and wia_eq: "wrap_iso_arcs hom f x = wrap_iso_arcs hom f y" from in_hom_A \B = _\ obtain x0 where x0: "x = iso_arcs hom x0" "x0 \ A" by auto from in_hom_A \B = _\ obtain y0 where y0: "y = iso_arcs hom y0" "y0 \ A" by auto have arcs_0: "x0 \ arcs G" "y0 \ arcs G" "f x0 \ arcs G" "f y0 \ arcs G" using x0 y0 \A \ _\ \f \ _\ by auto have "(iso_arcs hom o f o iso_arcs (inv_iso hom)) x = (iso_arcs hom o f o iso_arcs (inv_iso hom)) y" using in_hom_A wia_eq assms(1) \B = _\ by (auto simp: wrap_iso_arcs_def perm_restrict_def split: if_splits) then show "x = y" using hom assms digraph_isomorphism_inj_on_arcs[OF hom] x0 y0 arcs_0 \inj_on f A\ \A \ _\ by (auto dest!: inj_onD) qed lemma wrap_iso_arcs_in_funcsetI: assumes "A \ arcs G" "f \ A \ A" shows "wrap_iso_arcs hom f \ iso_arcs hom ` A \ iso_arcs hom ` A" proof fix x assume "x \ iso_arcs hom ` A" then obtain x0 where "x = iso_arcs hom x0" "x0 \ A" by blast then have "f x0 \ A" using \f \ _\ by auto then show "wrap_iso_arcs hom f x \ iso_arcs hom ` A" unfolding \x = _\ using \x0 \ A\ assms hom by (auto simp: wrap_iso_arcs_def perm_restrict_def) qed lemma wrap_iso_arcs_permutes: assumes "A \ arcs G" "f permutes A" shows "wrap_iso_arcs hom f permutes (iso_arcs hom ` A)" proof - { fix x assume A: "x \ iso_arcs hom ` A" have "wrap_iso_arcs hom f x = x" proof cases assume "x \ iso_arcs hom ` arcs G" then have "iso_arcs (inv_iso hom) x \ A" "x \ arcs (app_iso hom G)" using A hom by (metis arcs_app_iso image_eqI pre_digraph.iso_arcs_iso_inv, simp) then have "f (iso_arcs (inv_iso hom) x) = (iso_arcs (inv_iso hom) x)" using \f permutes A\ by (simp add: permutes_not_in) then show ?thesis using hom assms \x \ arcs _\ by (simp add: wrap_iso_arcs_def perm_restrict_def) next assume "x \ iso_arcs hom ` arcs G" then show ?thesis by (simp add: wrap_iso_arcs_def perm_restrict_def) qed } note not_in_id = this have "f \ A \ A" using assms by (intro funcsetI_permutes) have inj_on_wrap: "inj_on (wrap_iso_arcs hom f) (iso_arcs hom ` A)" using assms \f \ A \ A\ by (intro inj_on_wrap_iso_arcs_f) (auto intro: subset_inj_on permutes_inj) have woa_in_fs: "wrap_iso_arcs hom f \ iso_arcs hom ` A \ iso_arcs hom ` A" using assms \f \ A \ A\ by (intro wrap_iso_arcs_in_funcsetI) { fix x y assume "wrap_iso_arcs hom f x = wrap_iso_arcs hom f y" then have "x = y" apply (cases "x \ iso_arcs hom ` A"; cases "y \ iso_arcs hom ` A") using woa_in_fs inj_on_wrap by (auto dest: inj_onD simp: not_in_id) } note uniqueD = this note \f permutes A\ moreover note not_in_id moreover { fix y have "\x. wrap_iso_arcs hom f x = y" proof cases assume "y \ iso_arcs hom ` A" then obtain y0 where "y0 \ A" "iso_arcs hom y0 = y" by blast with \f permutes A\ obtain x0 where "x0 \ A" "f x0 = y0" unfolding permutes_def by metis moreover then have "\x. x \ arcs G \ iso_arcs hom x0 = iso_arcs hom x \ x0 = x" using assms hom by (auto simp: digraph_isomorphism_def dest: inj_onD) ultimately have "wrap_iso_arcs hom f (iso_arcs hom x0) = y" using \_ = y\ assms hom by (auto simp: wrap_iso_arcs_def perm_restrict_def) then show ?thesis .. qed (metis not_in_id) } ultimately show ?thesis unfolding permutes_def by (auto simp: dest: uniqueD) qed end lemma (in digraph_map) digraph_map_isoI: assumes "digraph_isomorphism hom" shows "digraph_map (app_iso hom G) (map_iso hom)" proof - interpret iG: fin_digraph "app_iso hom G" using assms by (rule fin_digraphI_app_iso) show ?thesis proof (rule iG.digraph_mapI_permutes) show "edge_rev (map_iso hom) permutes arcs (app_iso hom G)" using assms unfolding map_iso_def by (simp add: wrap_iso_arcs_permutes arev_permutes_arcs) next show "edge_succ (map_iso hom) permutes arcs (app_iso hom G)" using assms unfolding map_iso_def by (simp add: wrap_iso_arcs_permutes edge_succ_permutes) next fix a assume A: "a \ arcs (app_iso hom G)" show "tail (app_iso hom G) (edge_rev (map_iso hom) a) = head (app_iso hom G) a" using A assms by (cases rule: in_arcs_app_iso_cases) (auto simp: map_iso_def iso_verts_tail iso_verts_head) show "edge_rev (map_iso hom) (edge_rev (map_iso hom) a) = a" using A assms by (cases rule: in_arcs_app_iso_cases) (auto simp: map_iso_def) show "edge_rev (map_iso hom) a \ a" using A assms by (auto simp: map_iso_def arev_neq) next fix v assume "v \ verts (app_iso hom G)" and oa_hom: "out_arcs (app_iso hom G) v \ {}" then obtain v0 where "v0 \ verts G" "v = iso_verts hom v0" by auto moreover then have oa: "out_arcs G v0 \ {}" using assms oa_hom by (auto simp: out_arcs_def iso_verts_tail) ultimately have cyclic_on_v0: "cyclic_on (edge_succ M) (out_arcs G v0)" by (intro edge_succ_cyclic) from oa_hom obtain a where "a \ out_arcs (app_iso hom G) v" by blast then obtain a0 where "a0 \ arcs G" "a = iso_arcs hom a0" by auto then have "a0 \ out_arcs G v0" using \v = _\ \v0 \ _\ \a \ _\ assms by (simp add: iso_verts_tail) show "cyclic_on (edge_succ (map_iso hom)) (out_arcs (app_iso hom G) v)" proof (rule cyclic_on_singleI) show "a \ out_arcs (app_iso hom G) v" by fact next have "out_arcs (app_iso hom G) v = iso_arcs hom ` out_arcs G v0" unfolding \v = _\ by (rule out_arcs_app_iso_eq) fact+ also have "out_arcs G v0 = orbit (edge_succ M) a0" using cyclic_on_v0 \a0 \ out_arcs G v0\ unfolding cyclic_on_alldef by simp also have "iso_arcs hom ` \ = orbit (edge_succ (map_iso hom)) a" proof - have "\x. x \ orbit (edge_succ M) a0 \ x \ arcs G" using \out_arcs G v0 = _\ by auto then show ?thesis using \out_arcs G v0 = _\ unfolding \a = _\ assms using \a0 \ out_arcs G v0\ by (intro orbit_FOO) (insert assms, auto simp: map_iso_def) qed finally show "out_arcs (app_iso hom G) v = orbit (edge_succ (map_iso hom)) a" . qed qed qed end diff --git a/thys/Planarity_Certificates/Planarity/Permutations_2.thy b/thys/Planarity_Certificates/Planarity/Permutations_2.thy --- a/thys/Planarity_Certificates/Planarity/Permutations_2.thy +++ b/thys/Planarity_Certificates/Planarity/Permutations_2.thy @@ -1,153 +1,153 @@ theory Permutations_2 imports - "HOL-Library.Permutations" + "HOL-Combinatorics.Permutations" Executable_Permutations Graph_Theory.Funpow begin section \Modifying Permutations\ abbreviation funswapid :: "'a \ 'a \ 'a \ 'a" (infix "\\<^sub>F" 90) where "x \\<^sub>F y \ Fun.swap x y id" definition perm_swap :: "'a \ 'a \ ('a \ 'a) \ ('a \ 'a)" where "perm_swap x y f \ x \\<^sub>F y o f o x \\<^sub>F y" definition perm_rem :: "'a \ ('a \ 'a) \ ('a \ 'a)" where "perm_rem x f \ if f x \ x then x \\<^sub>F f x o f else f" text \ An example: @{lemma "perm_rem (2 :: nat) (list_succ [1,2,3,4]) x = list_succ [1,3,4] x" by (auto simp: perm_rem_def Fun.swap_def list_succ_def)} \ lemma perm_swap_id[simp]: "perm_swap a b id = id" by (auto simp: perm_swap_def) lemma perm_rem_permutes: assumes "f permutes S \ {x}" shows "perm_rem x f permutes S" using assms by (auto simp: permutes_def perm_rem_def) (metis swap_id_eq)+ lemma perm_rem_same: assumes "bij f" "f y = y" shows "perm_rem x f y = f y" using assms by (auto simp: perm_rem_def swap_id_eq bij_iff) lemma perm_rem_simps: assumes "bij f" shows "x = y \ perm_rem x f y = x" "f y = x \ perm_rem x f y = f x" "y \ x \ f y \ x \ perm_rem x f y = f y" using assms apply (auto simp: perm_rem_def ) by (metis bij_iff id_apply swap_apply(3)) lemma bij_swap_compose: "bij (x \\<^sub>F y o f) \ bij f" by (metis UNIV_I bij_betw_comp_iff2 bij_betw_id bij_swap_iff subsetI) lemma bij_perm_rem[simp]: "bij (perm_rem x f) \ bij f" by (simp add: perm_rem_def bij_swap_compose) lemma perm_rem_conv: "\f x y. bij f \ perm_rem x f y = ( if x = y then x else if f y = x then f (f y) else f y)" by (auto simp: perm_rem_simps) lemma perm_rem_commutes: assumes "bij f" shows "perm_rem a (perm_rem b f) = perm_rem b (perm_rem a f)" proof - have bij_simp: "\x y. f x = f y \ x = y" using assms by (auto simp: bij_iff) show ?thesis using assms by (auto simp: perm_rem_conv bij_simp fun_eq_iff) qed lemma perm_rem_id[simp]: "perm_rem a id = id" by (simp add: perm_rem_def) lemma bij_eq_iff: assumes "bij f" shows "f x = f y \ x = y" using assms by (metis bij_iff) lemma swap_swap_id[simp]: "(x \\<^sub>F y) ((x \\<^sub>F y) z) = z" by (simp add: swap_id_eq) lemma in_funswapid_image_iff: "\a b x S. x \ (a \\<^sub>F b) ` S \ (a \\<^sub>F b) x \ S" by (metis bij_def bij_id bij_swap_iff inj_image_mem_iff swap_swap_id) lemma perm_swap_comp: "perm_swap a b (f \ g) x = perm_swap a b f (perm_swap a b g x)" by (auto simp: perm_swap_def) lemma bij_perm_swap_iff[simp]: "bij (perm_swap a b f) \ bij f" by (auto simp: perm_swap_def bij_swap_compose bij_comp comp_swap) lemma funpow_perm_swap: "perm_swap a b f ^^ n = perm_swap a b (f ^^ n)" by (induct n) (auto simp: perm_swap_def fun_eq_iff) lemma orbit_perm_swap: "orbit (perm_swap a b f) x = (a \\<^sub>F b) ` orbit f ((a \\<^sub>F b) x)" by (auto simp: orbit_altdef funpow_perm_swap) (auto simp: perm_swap_def) lemma has_dom_perm_swap: "has_dom (perm_swap a b f) S = has_dom f ((a \\<^sub>F b) ` S)" by (auto simp: has_dom_def perm_swap_def inj_image_mem_iff) (metis image_iff swap_swap_id) lemma perm_restrict_dom_subset: assumes "has_dom f A" shows "perm_restrict f A = f" proof - from assms have "\x. x \ A \ f x = x" by (auto simp: has_dom_def) then show ?thesis by (auto simp: perm_restrict_def fun_eq_iff) qed lemma has_domD: "has_dom f S \ x \ S \ f x = x" by (auto simp: has_dom_def) lemma has_domI: "(\x. x \ S \ f x = x) \ has_dom f S" by (auto simp: has_dom_def) lemma perm_swap_permutes2: assumes "f permutes ((x \\<^sub>F y) ` S)" shows "perm_swap x y f permutes S" using assms by (auto simp: perm_swap_def permutes_conv_has_dom has_dom_perm_swap[unfolded perm_swap_def]) (metis bij_swap_iff bij_swap_compose_bij comp_id comp_swap) section \Cyclic Permutations\ lemma cyclic_on_perm_swap: assumes "cyclic_on f S" shows "cyclic_on (perm_swap x y f) ((x \\<^sub>F y) ` S)" using assms by (rule cyclic_on_FOO) (auto simp: perm_swap_def swap_swap_id) lemma orbit_perm_rem: assumes "bij f" "x \ y" shows "orbit (perm_rem y f) x = orbit f x - {y}" (is "?L = ?R") proof (intro set_eqI iffI) fix z assume "z \ ?L" then show "z \ ?R" using assms by induct (auto simp: perm_rem_conv bij_iff intro: orbit.intros) next fix z assume A: "z \ ?R" { assume "z \ orbit f x" then have "(z \ y \ z \ ?L) \ (z = y \ f z \ ?L)" proof induct case base with assms show ?case by (auto intro: orbit_eqI(1) simp: perm_rem_conv) next case (step z) then show ?case using assms by (cases "y = z") (auto intro: orbit_eqI simp: perm_rem_conv) qed } with A show "z \ ?L" by auto qed lemma orbit_perm_rem_eq: assumes "bij f" shows "orbit (perm_rem y f) x = (if x = y then {y} else orbit f x - {y})" using assms by (simp add: orbit_eq_singleton_iff orbit_perm_rem perm_rem_simps) lemma cyclic_on_perm_rem: assumes "cyclic_on f S" "bij f" "S \ {x}" shows "cyclic_on (perm_rem x f) (S - {x})" using assms[unfolded cyclic_on_alldef] by (simp add: cyclic_on_def orbit_perm_rem_eq) auto end diff --git a/thys/Planarity_Certificates/ROOT b/thys/Planarity_Certificates/ROOT --- a/thys/Planarity_Certificates/ROOT +++ b/thys/Planarity_Certificates/ROOT @@ -1,22 +1,23 @@ chapter AFP session Planarity_Certificates (AFP) = Simpl + options [timeout = 600] sessions + "HOL-Combinatorics" "HOL-Eisbach" "List-Index" "Transitive-Closure" Case_Labeling Graph_Theory directories "l4v/lib" "l4v/lib/wp" "Planarity" "Verification" theories [document = false] "l4v/lib/OptionMonadWP" theories Planarity_Certificates document_files "root.tex" "root.bib" diff --git a/thys/Polynomial_Factorization/ROOT b/thys/Polynomial_Factorization/ROOT --- a/thys/Polynomial_Factorization/ROOT +++ b/thys/Polynomial_Factorization/ROOT @@ -1,73 +1,72 @@ chapter AFP session "JNF-AFP-Lib" (AFP) in "Lib" = "HOL-Algebra" + description "Theories from HOL-Library and the Archive of Formal Proofs that are used by this entry" options [timeout = 600] sessions Containers "Abstract-Rewriting" Gauss_Jordan Matrix Polynomial_Interpolation Show VectorSpace "HOL-Cardinals" theories Containers.Set_Impl Matrix.Utility Matrix.Ordered_Semiring "Abstract-Rewriting.SN_Order_Carrier" "Abstract-Rewriting.Relative_Rewriting" Show.Show_Instances VectorSpace.VectorSpace Polynomial_Interpolation.Missing_Polynomial Polynomial_Interpolation.Ring_Hom_Poly "HOL-Library.AList" "HOL-Library.Cardinality" "HOL-Library.Char_ord" "HOL-Library.Code_Binary_Nat" "HOL-Library.Code_Target_Numeral" "HOL-Library.DAList" "HOL-Library.DAList_Multiset" "HOL-Library.Infinite_Set" "HOL-Library.Lattice_Syntax" "HOL-Library.Mapping" "HOL-Library.Monad_Syntax" "HOL-Library.More_List" "HOL-Library.Multiset" - "HOL-Library.Permutations" "HOL-Library.IArray" "HOL-Library.Phantom_Type" "HOL-Library.Ramsey" "HOL-Library.RBT_Impl" "HOL-Library.Simps_Case_Conv" "HOL-Library.While_Combinator" "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" "HOL-Computational_Algebra.Fraction_Field" "HOL-Computational_Algebra.Polynomial" "HOL-Computational_Algebra.Primes" "HOL-Cardinals.Order_Union" "HOL-Cardinals.Wellorder_Extension" session Polynomial_Factorization (AFP) = "JNF-AFP-Lib" + description "Algebraic Numbers" options [timeout = 600] sessions Partial_Function_MR Polynomial_Interpolation Show Sqrt_Babylonian theories Missing_Multiset Missing_List Precomputation Order_Polynomial Explicit_Roots Dvd_Int_Poly Rational_Root_Test Kronecker_Factorization Square_Free_Factorization Rational_Factorization document_files "root.bib" "root.tex" diff --git a/thys/Randomised_Social_Choice/Order_Predicates.thy b/thys/Randomised_Social_Choice/Order_Predicates.thy --- a/thys/Randomised_Social_Choice/Order_Predicates.thy +++ b/thys/Randomised_Social_Choice/Order_Predicates.thy @@ -1,1076 +1,1074 @@ (* Title: Order_Predicates.thy Author: Manuel Eberl, TU München Locales for order relations modelled as predicates (as opposed to sets of pairs). *) section \Order Relations as Binary Predicates\ theory Order_Predicates imports Main "HOL-Library.Disjoint_Sets" - "HOL-Library.Permutations" + "HOL-Combinatorics.Permutations" "List-Index.List_Index" begin - - subsection \Basic Operations on Relations\ text \The type of binary relations\ type_synonym 'a relation = "'a \ 'a \ bool" definition map_relation :: "('a \ 'b) \ 'b relation \ 'a relation" where "map_relation f R = (\x y. R (f x) (f y))" definition restrict_relation :: "'a set \ 'a relation \ 'a relation" where "restrict_relation A R = (\x y. x \ A \ y \ A \ R x y)" lemma restrict_relation_restrict_relation [simp]: "restrict_relation A (restrict_relation B R) = restrict_relation (A \ B) R" by (intro ext) (auto simp add: restrict_relation_def) lemma restrict_relation_empty [simp]: "restrict_relation {} R = (\_ _. False)" by (simp add: restrict_relation_def) lemma restrict_relation_UNIV [simp]: "restrict_relation UNIV R = R" by (simp add: restrict_relation_def) subsection \Preorders\ text \Preorders are reflexive and transitive binary relations.\ locale preorder_on = fixes carrier :: "'a set" fixes le :: "'a relation" assumes not_outside: "le x y \ x \ carrier" "le x y \ y \ carrier" assumes refl: "x \ carrier \ le x x" assumes trans: "le x y \ le y z \ le x z" begin lemma carrier_eq: "carrier = {x. le x x}" using not_outside refl by auto lemma preorder_on_map: "preorder_on (f -` carrier) (map_relation f le)" by unfold_locales (auto dest: not_outside simp: map_relation_def refl elim: trans) lemma preorder_on_restrict: "preorder_on (carrier \ A) (restrict_relation A le)" by unfold_locales (auto simp: restrict_relation_def refl intro: trans not_outside) lemma preorder_on_restrict_subset: "A \ carrier \ preorder_on A (restrict_relation A le)" using preorder_on_restrict[of A] by (simp add: Int_absorb1) lemma restrict_relation_carrier [simp]: "restrict_relation carrier le = le" using not_outside by (intro ext) (auto simp add: restrict_relation_def) end subsection \Total preorders\ text \Total preorders are preorders where any two elements are comparable.\ locale total_preorder_on = preorder_on + assumes total: "x \ carrier \ y \ carrier \ le x y \ le y x" begin lemma total': "\le x y \ x \ carrier \ y \ carrier \ le y x" using total[of x y] by blast lemma total_preorder_on_map: "total_preorder_on (f -` carrier) (map_relation f le)" proof - interpret R': preorder_on "f -` carrier" "map_relation f le" using preorder_on_map[of f] . show ?thesis by unfold_locales (simp add: map_relation_def total) qed lemma total_preorder_on_restrict: "total_preorder_on (carrier \ A) (restrict_relation A le)" proof - interpret R': preorder_on "carrier \ A" "restrict_relation A le" by (rule preorder_on_restrict) from total show ?thesis by unfold_locales (auto simp: restrict_relation_def) qed lemma total_preorder_on_restrict_subset: "A \ carrier \ total_preorder_on A (restrict_relation A le)" using total_preorder_on_restrict[of A] by (simp add: Int_absorb1) end text \Some fancy notation for order relations\ abbreviation (input) weakly_preferred :: "'a \ 'a relation \ 'a \ bool" ("_ \[_] _" [51,10,51] 60) where "a \[R] b \ R a b" definition strongly_preferred ("_ \[_] _" [51,10,51] 60) where "a \[R] b \ (a \[R] b) \ \(b \[R] a)" definition indifferent ("_ \[_] _" [51,10,51] 60) where "a \[R] b \ (a \[R] b) \ (b \[R] a)" abbreviation (input) weakly_not_preferred ("_ \[_] _" [51,10,51] 60) where "a \[R] b \ b \[R] a" term "a \[R] b \ b \[R] a" abbreviation (input) strongly_not_preferred ("_ \[_] _" [51,10,51] 60) where "a \[R] b \ b \[R] a" context preorder_on begin lemma strict_trans: "a \[le] b \ b \[le] c \ a \[le] c" unfolding strongly_preferred_def by (blast intro: trans) lemma weak_strict_trans: "a \[le] b \ b \[le] c \ a \[le] c" unfolding strongly_preferred_def by (blast intro: trans) lemma strict_weak_trans: "a \[le] b \ b \[le] c \ a \[le] c" unfolding strongly_preferred_def by (blast intro: trans) end lemma (in total_preorder_on) not_weakly_preferred_iff: "a \ carrier \ b \ carrier \ \a \[le] b \ b \[le] a" using total[of a b] by (auto simp: strongly_preferred_def) lemma (in total_preorder_on) not_strongly_preferred_iff: "a \ carrier \ b \ carrier \ \a \[le] b \ b \[le] a" using total[of a b] by (auto simp: strongly_preferred_def) subsection \Orders\ locale order_on = preorder_on + assumes antisymmetric: "le x y \ le y x \ x = y" locale linorder_on = order_on carrier le + total_preorder_on carrier le for carrier le subsection \Maximal elements\ text \ Maximal elements are elements in a preorder for which there exists no strictly greater element. \ definition Max_wrt_among :: "'a relation \ 'a set \ 'a set" where "Max_wrt_among R A = {x\A. R x x \ (\y\A. R x y \ R y x)}" lemma Max_wrt_among_cong: assumes "restrict_relation A R = restrict_relation A R'" shows "Max_wrt_among R A = Max_wrt_among R' A" proof - from assms have "R x y \ R' x y" if "x \ A" "y \ A" for x y using that by (auto simp: restrict_relation_def fun_eq_iff) thus ?thesis unfolding Max_wrt_among_def by blast qed definition Max_wrt :: "'a relation \ 'a set" where "Max_wrt R = Max_wrt_among R UNIV" lemma Max_wrt_altdef: "Max_wrt R = {x. R x x \ (\y. R x y \ R y x)}" unfolding Max_wrt_def Max_wrt_among_def by simp context preorder_on begin lemma Max_wrt_among_preorder: "Max_wrt_among le A = {x\carrier \ A. \y\carrier \ A. le x y \ le y x}" unfolding Max_wrt_among_def using not_outside refl by blast lemma Max_wrt_preorder: "Max_wrt le = {x\carrier. \y\carrier. le x y \ le y x}" unfolding Max_wrt_altdef using not_outside refl by blast lemma Max_wrt_among_subset: "Max_wrt_among le A \ carrier" "Max_wrt_among le A \ A" unfolding Max_wrt_among_preorder by auto lemma Max_wrt_subset: "Max_wrt le \ carrier" unfolding Max_wrt_preorder by auto lemma Max_wrt_among_nonempty: assumes "B \ carrier \ {}" "finite (B \ carrier)" shows "Max_wrt_among le B \ {}" proof - define A where "A = B \ carrier" have "A \ carrier" by (simp add: A_def) from assms(2,1)[folded A_def] this have "{x\A. (\y\A. le x y \ le y x)} \ {}" proof (induction A rule: finite_ne_induct) case (singleton x) thus ?case by (auto simp: refl) next case (insert x A) then obtain y where y: "y \ A" "\z. z \ A \ le y z \ le z y" by blast thus ?case using insert.prems by (cases "le y x") (blast intro: trans)+ qed thus ?thesis by (simp add: A_def Max_wrt_among_preorder Int_commute) qed lemma Max_wrt_nonempty: "carrier \ {} \ finite carrier \ Max_wrt le \ {}" using Max_wrt_among_nonempty[of UNIV] by (simp add: Max_wrt_def) lemma Max_wrt_among_map_relation_vimage: "f -` Max_wrt_among le A \ Max_wrt_among (map_relation f le) (f -` A)" by (auto simp: Max_wrt_among_def map_relation_def) lemma Max_wrt_map_relation_vimage: "f -` Max_wrt le \ Max_wrt (map_relation f le)" by (auto simp: Max_wrt_altdef map_relation_def) lemma image_subset_vimage_the_inv_into: assumes "inj_on f A" "B \ A" shows "f ` B \ the_inv_into A f -` B" using assms by (auto simp: the_inv_into_f_f) lemma Max_wrt_among_map_relation_bij_subset: assumes "bij (f :: 'a \ 'b)" shows "f ` Max_wrt_among le A \ Max_wrt_among (map_relation (inv f) le) (f ` A)" using assms Max_wrt_among_map_relation_vimage[of "inv f" A] by (simp add: bij_imp_bij_inv inv_inv_eq bij_vimage_eq_inv_image) lemma Max_wrt_among_map_relation_bij: assumes "bij f" shows "f ` Max_wrt_among le A = Max_wrt_among (map_relation (inv f) le) (f ` A)" proof (intro equalityI Max_wrt_among_map_relation_bij_subset assms) interpret R: preorder_on "f ` carrier" "map_relation (inv f) le" using preorder_on_map[of "inv f"] assms by (simp add: bij_imp_bij_inv bij_vimage_eq_inv_image inv_inv_eq) show "Max_wrt_among (map_relation (inv f) le) (f ` A) \ f ` Max_wrt_among le A" unfolding Max_wrt_among_preorder R.Max_wrt_among_preorder using assms bij_is_inj[OF assms] by (auto simp: map_relation_def inv_f_f image_Int [symmetric]) qed lemma Max_wrt_map_relation_bij: "bij f \ f ` Max_wrt le = Max_wrt (map_relation (inv f) le)" proof - assume bij: "bij f" interpret R: preorder_on "f ` carrier" "map_relation (inv f) le" using preorder_on_map[of "inv f"] bij by (simp add: bij_imp_bij_inv bij_vimage_eq_inv_image inv_inv_eq) from bij show ?thesis unfolding R.Max_wrt_preorder Max_wrt_preorder by (auto simp: map_relation_def inv_f_f bij_is_inj) qed lemma Max_wrt_among_mono: "le x y \ x \ Max_wrt_among le A \ y \ A \ y \ Max_wrt_among le A" using not_outside by (auto simp: Max_wrt_among_preorder intro: trans) lemma Max_wrt_mono: "le x y \ x \ Max_wrt le \ y \ Max_wrt le" unfolding Max_wrt_def using Max_wrt_among_mono[of x y UNIV] by blast end context total_preorder_on begin lemma Max_wrt_among_total_preorder: "Max_wrt_among le A = {x\carrier \ A. \y\carrier \ A. le y x}" unfolding Max_wrt_among_preorder using total by blast lemma Max_wrt_total_preorder: "Max_wrt le = {x\carrier. \y\carrier. le y x}" unfolding Max_wrt_preorder using total by blast lemma decompose_Max: assumes A: "A \ carrier" defines "M \ Max_wrt_among le A" shows "restrict_relation A le = (\x y. x \ A \ y \ M \ (y \ M \ restrict_relation (A - M) le x y))" using A by (intro ext) (auto simp: M_def Max_wrt_among_total_preorder restrict_relation_def Int_absorb1 intro: trans) end subsection \Weak rankings\ inductive of_weak_ranking :: "'alt set list \ 'alt relation" where "i \ j \ i < length xs \ j < length xs \ x \ xs ! i \ y \ xs ! j \ x \[of_weak_ranking xs] y" lemma of_weak_ranking_Nil [simp]: "of_weak_ranking [] = (\_ _. False)" by (intro ext) (simp add: of_weak_ranking.simps) lemma of_weak_ranking_Nil' [code]: "of_weak_ranking [] x y = False" by simp lemma of_weak_ranking_Cons [code]: "x \[of_weak_ranking (z#zs)] y \ x \ z \ y \ \(set (z#zs)) \ x \[of_weak_ranking zs] y" (is "?lhs \ ?rhs") proof assume ?lhs then obtain i j where ij: "i < length (z#zs)" "j < length (z#zs)" "i \ j" "x \ (z#zs) ! i" "y \ (z#zs) ! j" by (blast elim: of_weak_ranking.cases) thus ?rhs by (cases i; cases j) (force intro: of_weak_ranking.intros)+ next assume ?rhs thus ?lhs proof (elim disjE conjE) assume "x \ z" "y \ \(set (z # zs))" then obtain j where "j < length (z # zs)" "y \ (z # zs) ! j" by (subst (asm) set_conv_nth) auto with \x \ z\ show "of_weak_ranking (z # zs) y x" by (intro of_weak_ranking.intros[of 0 j]) auto next assume "of_weak_ranking zs y x" then obtain i j where "i < length zs" "j < length zs" "i \ j" "x \ zs ! i" "y \ zs ! j" by (blast elim: of_weak_ranking.cases) thus "of_weak_ranking (z # zs) y x" by (intro of_weak_ranking.intros[of "Suc i" "Suc j"]) auto qed qed lemma of_weak_ranking_indifference: assumes "A \ set xs" "x \ A" "y \ A" shows "x \[of_weak_ranking xs] y" using assms by (induction xs) (auto simp: of_weak_ranking_Cons) lemma of_weak_ranking_map: "map_relation f (of_weak_ranking xs) = of_weak_ranking (map ((-`) f) xs)" by (intro ext, induction xs) (simp_all add: map_relation_def of_weak_ranking_Cons) lemma of_weak_ranking_permute': assumes "f permutes (\(set xs))" shows "map_relation f (of_weak_ranking xs) = of_weak_ranking (map ((`) (inv f)) xs)" proof - have "map_relation f (of_weak_ranking xs) = of_weak_ranking (map ((-`) f) xs)" by (rule of_weak_ranking_map) also from assms have "map ((-`) f) xs = map ((`) (inv f)) xs" by (intro map_cong refl) (simp_all add: bij_vimage_eq_inv_image permutes_bij) finally show ?thesis . qed lemma of_weak_ranking_permute: assumes "f permutes (\(set xs))" shows "of_weak_ranking (map ((`) f) xs) = map_relation (inv f) (of_weak_ranking xs)" using of_weak_ranking_permute'[OF permutes_inv[OF assms]] assms by (simp add: inv_inv_eq permutes_bij) definition is_weak_ranking where "is_weak_ranking xs \ ({} \ set xs) \ (\i j. i < length xs \ j < length xs \ i \ j \ xs ! i \ xs ! j = {})" definition is_finite_weak_ranking where "is_finite_weak_ranking xs \ is_weak_ranking xs \ (\x\set xs. finite x)" definition weak_ranking :: "'alt relation \ 'alt set list" where "weak_ranking R = (SOME xs. is_weak_ranking xs \ R = of_weak_ranking xs)" lemma is_weak_rankingI [intro?]: assumes "{} \ set xs" "\i j. i < length xs \ j < length xs \ i \ j \ xs ! i \ xs ! j = {}" shows "is_weak_ranking xs" using assms by (auto simp add: is_weak_ranking_def) lemma is_weak_ranking_nonempty: "is_weak_ranking xs \ {} \ set xs" by (simp add: is_weak_ranking_def) lemma is_weak_rankingD: assumes "is_weak_ranking xs" "i < length xs" "j < length xs" "i \ j" shows "xs ! i \ xs ! j = {}" using assms by (simp add: is_weak_ranking_def) lemma is_weak_ranking_iff: "is_weak_ranking xs \ distinct xs \ disjoint (set xs) \ {} \ set xs" proof safe assume wf: "is_weak_ranking xs" from wf show "disjoint (set xs)" by (auto simp: disjoint_def is_weak_ranking_def set_conv_nth) show "distinct xs" proof (subst distinct_conv_nth, safe) fix i j assume ij: "i < length xs" "j < length xs" "i \ j" "xs ! i = xs ! j" then have "xs ! i \ xs ! j = {}" by (intro is_weak_rankingD wf) with ij have "xs ! i = {}" by simp with ij have "{} \ set xs" by (auto simp: set_conv_nth) moreover from wf ij have "{} \ set xs" by (intro is_weak_ranking_nonempty wf) ultimately show False by contradiction qed next assume A: "distinct xs" "disjoint (set xs)" "{} \ set xs" thus "is_weak_ranking xs" by (intro is_weak_rankingI) (auto simp: disjoint_def distinct_conv_nth) qed (simp_all add: is_weak_ranking_nonempty) lemma is_weak_ranking_rev [simp]: "is_weak_ranking (rev xs) \ is_weak_ranking xs" by (simp add: is_weak_ranking_iff) lemma is_weak_ranking_map_inj: assumes "is_weak_ranking xs" "inj_on f (\(set xs))" shows "is_weak_ranking (map ((`) f) xs)" using assms by (auto simp: is_weak_ranking_iff distinct_map inj_on_image disjoint_image) lemma of_weak_ranking_rev [simp]: "of_weak_ranking (rev xs) (x::'a) y \ of_weak_ranking xs y x" proof - have "of_weak_ranking (rev xs) y x" if "of_weak_ranking xs x y" for xs and x y :: 'a proof - from that obtain i j where "i < length xs" "j < length xs" "x \ xs ! i" "y \ xs ! j" "i \ j" by (elim of_weak_ranking.cases) simp_all thus ?thesis by (intro of_weak_ranking.intros[of "length xs - i - 1" "length xs - j - 1"] diff_le_mono2) (auto simp: diff_le_mono2 rev_nth) qed from this[of xs y x] this[of "rev xs" x y] show ?thesis by (intro iffI) simp_all qed lemma is_weak_ranking_Nil [simp, code]: "is_weak_ranking []" by (auto simp: is_weak_ranking_def) lemma is_finite_weak_ranking_Nil [simp, code]: "is_finite_weak_ranking []" by (auto simp: is_finite_weak_ranking_def) lemma is_weak_ranking_Cons_empty [simp]: "\is_weak_ranking ({} # xs)" by (simp add: is_weak_ranking_def) lemma is_finite_weak_ranking_Cons_empty [simp]: "\is_finite_weak_ranking ({} # xs)" by (simp add: is_finite_weak_ranking_def) lemma is_weak_ranking_singleton [simp]: "is_weak_ranking [x] \ x \ {}" by (auto simp add: is_weak_ranking_def) lemma is_finite_weak_ranking_singleton [simp]: "is_finite_weak_ranking [x] \ x \ {} \ finite x" by (auto simp add: is_finite_weak_ranking_def) lemma is_weak_ranking_append: "is_weak_ranking (xs @ ys) \ is_weak_ranking xs \ is_weak_ranking ys \ (set xs \ set ys = {} \ \(set xs) \ \(set ys) = {})" by (simp only: is_weak_ranking_iff) (auto dest: disjointD disjoint_unionD1 disjoint_unionD2 intro: disjoint_union) lemma is_weak_ranking_Cons [code]: "is_weak_ranking (x # xs) \ x \ {} \ is_weak_ranking xs \ x \ \(set xs) = {}" using is_weak_ranking_append[of "[x]" xs] by auto lemma is_finite_weak_ranking_Cons [code]: "is_finite_weak_ranking (x # xs) \ x \ {} \ finite x \ is_finite_weak_ranking xs \ x \ \(set xs) = {}" by (auto simp add: is_finite_weak_ranking_def is_weak_ranking_Cons) primrec is_weak_ranking_aux where "is_weak_ranking_aux A [] \ True" | "is_weak_ranking_aux A (x#xs) \ x \ {} \ A \ x = {} \ is_weak_ranking_aux (A \ x) xs" lemma is_weak_ranking_aux: "is_weak_ranking_aux A xs \ A \ \(set xs) = {} \ is_weak_ranking xs" by (induction xs arbitrary: A) (auto simp: is_weak_ranking_Cons) lemma is_weak_ranking_code [code]: "is_weak_ranking xs \ is_weak_ranking_aux {} xs" by (subst is_weak_ranking_aux) auto lemma of_weak_ranking_altdef: assumes "is_weak_ranking xs" "x \ \(set xs)" "y \ \(set xs)" shows "of_weak_ranking xs x y \ find_index ((\) x) xs \ find_index ((\) y) xs" proof - from assms have A: "find_index ((\) x) xs < length xs" "find_index ((\) y) xs < length xs" by (simp_all add: find_index_less_size_conv) from this[THEN nth_find_index] have B: "x \ xs ! find_index ((\) x) xs" "y \ xs ! find_index ((\) y) xs" . show ?thesis proof assume "of_weak_ranking xs x y" then obtain i j where ij: "j \ i" "i < length xs" "j < length xs" "x \ xs ! i" "y \ xs !j" by (cases rule: of_weak_ranking.cases) simp_all with A B have "i = find_index ((\) x) xs" "j = find_index ((\) y) xs" using assms(1) unfolding is_weak_ranking_def by blast+ with ij show "find_index ((\) x) xs \ find_index ((\) y) xs" by simp next assume "find_index ((\) x) xs \ find_index ((\) y) xs" from this A(2,1) B(2,1) show "of_weak_ranking xs x y" by (rule of_weak_ranking.intros) qed qed lemma total_preorder_of_weak_ranking: assumes "\(set xs) = A" assumes "is_weak_ranking xs" shows "total_preorder_on A (of_weak_ranking xs)" proof fix x y assume "x \[of_weak_ranking xs] y" with assms show "x \ A" "y \ A" by (auto elim!: of_weak_ranking.cases) next fix x assume "x \ A" with assms(1) obtain i where "i < length xs" "x \ xs ! i" by (auto simp: set_conv_nth) thus "x \[of_weak_ranking xs] x" by (auto intro: of_weak_ranking.intros) next fix x y assume "x \ A" "y \ A" with assms(1) obtain i j where ij: "i < length xs" "j < length xs" "x \ xs ! i" "y \ xs ! j" by (auto simp: set_conv_nth) consider "i \ j" | "j \ i" by force thus "x \[of_weak_ranking xs] y \ y \[of_weak_ranking xs] x" by cases (insert ij, (blast intro: of_weak_ranking.intros)+) next fix x y z assume A: "x \[of_weak_ranking xs] y" and B: "y \[of_weak_ranking xs] z" from A obtain i j where ij: "i \ j" "i < length xs" "j < length xs" "x \ xs ! i" "y \ xs ! j" by (auto elim!: of_weak_ranking.cases) moreover from B obtain j' k where j'k: "j' \ k" "j' < length xs" "k < length xs" "y \ xs ! j'" "z \ xs ! k" by (auto elim!: of_weak_ranking.cases) moreover from ij j'k is_weak_rankingD[OF assms(2), of j j'] have "j = j'" by blast ultimately show "x \[of_weak_ranking xs] z" by (auto intro: of_weak_ranking.intros[of k i]) qed lemma restrict_relation_of_weak_ranking_Cons: assumes "is_weak_ranking (A # As)" shows "restrict_relation (\(set As)) (of_weak_ranking (A # As)) = of_weak_ranking As" proof - from assms interpret R: total_preorder_on "\(set As)" "of_weak_ranking As" by (intro total_preorder_of_weak_ranking) (simp_all add: is_weak_ranking_Cons) from assms show ?thesis using R.not_outside by (intro ext) (auto simp: restrict_relation_def of_weak_ranking_Cons is_weak_ranking_Cons) qed lemmas of_weak_ranking_wf = total_preorder_of_weak_ranking is_weak_ranking_code insert_commute (* Test *) lemma "total_preorder_on {1,2,3,4::nat} (of_weak_ranking [{1,3},{2},{4}])" by (simp add: of_weak_ranking_wf) context fixes x :: "'alt set" and xs :: "'alt set list" assumes wf: "is_weak_ranking (x#xs)" begin interpretation R: total_preorder_on "\(set (x#xs))" "of_weak_ranking (x#xs)" by (intro total_preorder_of_weak_ranking) (simp_all add: wf) lemma of_weak_ranking_imp_in_set: assumes "of_weak_ranking xs a b" shows "a \ \(set xs)" "b \ \(set xs)" using assms by (fastforce elim!: of_weak_ranking.cases)+ lemma of_weak_ranking_Cons': assumes "a \ \(set (x#xs))" "b \ \(set (x#xs))" shows "of_weak_ranking (x#xs) a b \ b \ x \ (a \ x \ of_weak_ranking xs a b)" proof assume "of_weak_ranking (x # xs) a b" with wf of_weak_ranking_imp_in_set[of a b] show "(b \ x \ a \ x \ of_weak_ranking xs a b)" by (auto simp: is_weak_ranking_Cons of_weak_ranking_Cons) next assume "b \ x \ a \ x \ of_weak_ranking xs a b" with assms show "of_weak_ranking (x#xs) a b" by (fastforce simp: of_weak_ranking_Cons) qed lemma Max_wrt_among_of_weak_ranking_Cons1: assumes "x \ A = {}" shows "Max_wrt_among (of_weak_ranking (x#xs)) A = Max_wrt_among (of_weak_ranking xs) A" proof - from wf interpret R': total_preorder_on "\(set xs)" "of_weak_ranking xs" by (intro total_preorder_of_weak_ranking) (simp_all add: is_weak_ranking_Cons) from assms show ?thesis by (auto simp: R.Max_wrt_among_total_preorder R'.Max_wrt_among_total_preorder of_weak_ranking_Cons) qed lemma Max_wrt_among_of_weak_ranking_Cons2: assumes "x \ A \ {}" shows "Max_wrt_among (of_weak_ranking (x#xs)) A = x \ A" proof - from wf interpret R': total_preorder_on "\(set xs)" "of_weak_ranking xs" by (intro total_preorder_of_weak_ranking) (simp_all add: is_weak_ranking_Cons) from assms obtain a where "a \ x \ A" by blast with wf R'.not_outside(1)[of a] show ?thesis by (auto simp: R.Max_wrt_among_total_preorder is_weak_ranking_Cons R'.Max_wrt_among_total_preorder of_weak_ranking_Cons) qed lemma Max_wrt_among_of_weak_ranking_Cons: "Max_wrt_among (of_weak_ranking (x#xs)) A = (if x \ A = {} then Max_wrt_among (of_weak_ranking xs) A else x \ A)" using Max_wrt_among_of_weak_ranking_Cons1 Max_wrt_among_of_weak_ranking_Cons2 by simp lemma Max_wrt_of_weak_ranking_Cons: "Max_wrt (of_weak_ranking (x#xs)) = x" using wf by (simp add: is_weak_ranking_Cons Max_wrt_def Max_wrt_among_of_weak_ranking_Cons) end lemma Max_wrt_of_weak_ranking: assumes "is_weak_ranking xs" shows "Max_wrt (of_weak_ranking xs) = (if xs = [] then {} else hd xs)" proof (cases xs) case Nil hence "of_weak_ranking xs = (\_ _. False)" by (intro ext) simp_all with Nil show ?thesis by (simp add: Max_wrt_def Max_wrt_among_def) next case (Cons x xs') with assms show ?thesis by (simp add: Max_wrt_of_weak_ranking_Cons) qed locale finite_total_preorder_on = total_preorder_on + assumes finite_carrier [intro]: "finite carrier" begin lemma finite_total_preorder_on_map: assumes "finite (f -` carrier)" shows "finite_total_preorder_on (f -` carrier) (map_relation f le)" proof - interpret R': total_preorder_on "f -` carrier" "map_relation f le" using total_preorder_on_map[of f] . from assms show ?thesis by unfold_locales simp qed function weak_ranking_aux :: "'a set \ 'a set list" where "weak_ranking_aux {} = []" | "A \ {} \ A \ carrier \ weak_ranking_aux A = Max_wrt_among le A # weak_ranking_aux (A - Max_wrt_among le A)" | "\(A \ carrier) \ weak_ranking_aux A = undefined" by blast simp_all termination proof (relation "Wellfounded.measure card") fix A let ?B = "Max_wrt_among le A" assume A: "A \ {}" "A \ carrier" moreover from A(2) have "finite A" by (rule finite_subset) blast moreover from A have "?B \ {}" "?B \ A" by (intro Max_wrt_among_nonempty Max_wrt_among_subset; force)+ ultimately have "card (A - ?B) < card A" by (intro psubset_card_mono) auto thus "(A - ?B, A) \ measure card" by simp qed simp_all lemma weak_ranking_aux_Union: "A \ carrier \ \(set (weak_ranking_aux A)) = A" proof (induction A rule: weak_ranking_aux.induct [case_names empty nonempty]) case (nonempty A) with Max_wrt_among_subset[of A] show ?case by auto qed simp_all lemma weak_ranking_aux_wf: "A \ carrier \ is_weak_ranking (weak_ranking_aux A)" proof (induction A rule: weak_ranking_aux.induct [case_names empty nonempty]) case (nonempty A) have "is_weak_ranking (Max_wrt_among le A # weak_ranking_aux (A - Max_wrt_among le A))" unfolding is_weak_ranking_Cons proof (intro conjI) from nonempty.prems nonempty.hyps show "Max_wrt_among le A \ {}" by (intro Max_wrt_among_nonempty) auto next from nonempty.prems show "is_weak_ranking (weak_ranking_aux (A - Max_wrt_among le A))" by (intro nonempty.IH) blast next from nonempty.prems nonempty.hyps have "Max_wrt_among le A \ {}" by (intro Max_wrt_among_nonempty) auto moreover from nonempty.prems have "\(set (weak_ranking_aux (A - Max_wrt_among le A))) = A - Max_wrt_among le A" by (intro weak_ranking_aux_Union) auto ultimately show "Max_wrt_among le A \ \(set (weak_ranking_aux (A - Max_wrt_among le A))) = {}" by blast+ qed with nonempty.prems nonempty.hyps show ?case by simp qed simp_all lemma of_weak_ranking_weak_ranking_aux': assumes "A \ carrier" "x \ A" "y \ A" shows "of_weak_ranking (weak_ranking_aux A) x y \ restrict_relation A le x y" using assms proof (induction A rule: weak_ranking_aux.induct [case_names empty nonempty]) case (nonempty A) define M where "M = Max_wrt_among le A" from nonempty.prems nonempty.hyps have M: "M \ A" unfolding M_def by (intro Max_wrt_among_subset) from nonempty.prems have in_MD: "le x y" if "x \ A" "y \ M" for x y using that unfolding M_def Max_wrt_among_total_preorder by (auto simp: Int_absorb1) from nonempty.prems have in_MI: "x \ M" if "y \ M" "x \ A" "le y x" for x y using that unfolding M_def Max_wrt_among_total_preorder by (auto simp: Int_absorb1 intro: trans) from nonempty.prems nonempty.hyps have IH: "of_weak_ranking (weak_ranking_aux (A - M)) x y = restrict_relation (A - M) le x y" if "x \ M" "y \ M" using that unfolding M_def by (intro nonempty.IH) auto from nonempty.prems interpret R': total_preorder_on "A - M" "of_weak_ranking (weak_ranking_aux (A - M))" by (intro total_preorder_of_weak_ranking weak_ranking_aux_wf weak_ranking_aux_Union) auto from nonempty.prems nonempty.hyps M weak_ranking_aux_Union[of A] R'.not_outside[of x y] show ?case by (cases "x \ M"; cases "y \ M") (auto simp: restrict_relation_def of_weak_ranking_Cons IH M_def [symmetric] intro: in_MD dest: in_MI) qed simp_all lemma of_weak_ranking_weak_ranking_aux: "of_weak_ranking (weak_ranking_aux carrier) = le" proof (intro ext) fix x y have "is_weak_ranking (weak_ranking_aux carrier)" by (rule weak_ranking_aux_wf) simp then interpret R: total_preorder_on carrier "of_weak_ranking (weak_ranking_aux carrier)" by (intro total_preorder_of_weak_ranking weak_ranking_aux_wf weak_ranking_aux_Union) (simp_all add: weak_ranking_aux_Union) show "of_weak_ranking (weak_ranking_aux carrier) x y = le x y" proof (cases "x \ carrier \ y \ carrier") case True thus ?thesis using of_weak_ranking_weak_ranking_aux'[of carrier x y] by simp next case False with R.not_outside have "of_weak_ranking (weak_ranking_aux carrier) x y = False" by auto also from not_outside False have "\ = le x y" by auto finally show ?thesis . qed qed lemma weak_ranking_aux_unique': assumes "\(set As) \ carrier" "is_weak_ranking As" "of_weak_ranking As = restrict_relation (\(set As)) le" shows "As = weak_ranking_aux (\(set As))" using assms proof (induction As) case (Cons A As) have "restrict_relation (\(set As)) (of_weak_ranking (A # As)) = of_weak_ranking As" by (intro restrict_relation_of_weak_ranking_Cons Cons.prems) also have eq1: "of_weak_ranking (A # As) = restrict_relation (\(set (A # As))) le" by fact finally have eq: "of_weak_ranking As = restrict_relation (\(set As)) le" by (simp add: Int_absorb2) with Cons.prems have eq2: "weak_ranking_aux (\(set As)) = As" by (intro sym [OF Cons.IH]) (auto simp: is_weak_ranking_Cons) from eq1 have "Max_wrt_among le (\(set (A # As))) = Max_wrt_among (of_weak_ranking (A#As)) (\(set (A#As)))" by (intro Max_wrt_among_cong) simp_all also from Cons.prems have "\ = A" by (subst Max_wrt_among_of_weak_ranking_Cons2) (simp_all add: is_weak_ranking_Cons) finally have Max: "Max_wrt_among le (\(set (A # As))) = A" . moreover from Cons.prems have "A \ {}" by (simp add: is_weak_ranking_Cons) ultimately have "weak_ranking_aux (\(set (A # As))) = A # weak_ranking_aux (A \ \(set As) - A)" using Cons.prems by simp also from Cons.prems have "A \ \(set As) - A = \(set As)" by (auto simp: is_weak_ranking_Cons) also from eq2 have "weak_ranking_aux \ = As" . finally show ?case .. qed simp_all lemma weak_ranking_aux_unique: assumes "is_weak_ranking As" "of_weak_ranking As = le" shows "As = weak_ranking_aux carrier" proof - interpret R: total_preorder_on "\(set As)" "of_weak_ranking As" by (intro total_preorder_of_weak_ranking assms) simp_all from assms have "x \ \(set As) \ x \ carrier" for x using R.not_outside not_outside R.refl[of x] refl[of x] by blast hence eq: "\(set As) = carrier" by blast from assms eq have "As = weak_ranking_aux (\(set As))" by (intro weak_ranking_aux_unique') simp_all with eq show ?thesis by simp qed lemma weak_ranking_total_preorder: "is_weak_ranking (weak_ranking le)" "of_weak_ranking (weak_ranking le) = le" proof - from weak_ranking_aux_wf[of carrier] of_weak_ranking_weak_ranking_aux have "\x. is_weak_ranking x \ le = of_weak_ranking x" by auto hence "is_weak_ranking (weak_ranking le) \ le = of_weak_ranking (weak_ranking le)" unfolding weak_ranking_def by (rule someI_ex) thus "is_weak_ranking (weak_ranking le)" "of_weak_ranking (weak_ranking le) = le" by simp_all qed lemma weak_ranking_altdef: "weak_ranking le = weak_ranking_aux carrier" by (intro weak_ranking_aux_unique weak_ranking_total_preorder) lemma weak_ranking_Union: "\(set (weak_ranking le)) = carrier" by (simp add: weak_ranking_altdef weak_ranking_aux_Union) lemma weak_ranking_unique: assumes "is_weak_ranking As" "of_weak_ranking As = le" shows "As = weak_ranking le" using assms unfolding weak_ranking_altdef by (rule weak_ranking_aux_unique) lemma weak_ranking_permute: assumes "f permutes carrier" shows "weak_ranking (map_relation (inv f) le) = map ((`) f) (weak_ranking le)" proof - from assms have "inv f -` carrier = carrier" by (simp add: permutes_vimage permutes_inv) then interpret R: finite_total_preorder_on "inv f -` carrier" "map_relation (inv f) le" by (intro finite_total_preorder_on_map) (simp_all add: finite_carrier) from assms have "is_weak_ranking (map ((`) f) (weak_ranking le))" by (intro is_weak_ranking_map_inj) (simp_all add: weak_ranking_total_preorder permutes_inj_on) with assms show ?thesis by (intro sym[OF R.weak_ranking_unique]) (simp_all add: of_weak_ranking_permute weak_ranking_Union weak_ranking_total_preorder) qed lemma weak_ranking_index_unique: assumes "is_weak_ranking xs" "i < length xs" "j < length xs" "x \ xs ! i" "x \ xs ! j" shows "i = j" using assms unfolding is_weak_ranking_def by auto lemma weak_ranking_index_unique': assumes "is_weak_ranking xs" "i < length xs" "x \ xs ! i" shows "i = find_index ((\) x) xs" using assms find_index_less_size_conv nth_mem by (intro weak_ranking_index_unique[OF assms(1,2) _ assms(3)] nth_find_index[of "(\) x"]) blast+ lemma weak_ranking_eqclass1: assumes "A \ set (weak_ranking le)" "x \ A" "y \ A" shows "le x y" proof - from assms obtain i where "weak_ranking le ! i = A" "i < length (weak_ranking le)" by (auto simp: set_conv_nth) with assms have "of_weak_ranking (weak_ranking le) x y" by (intro of_weak_ranking.intros[of i i]) auto thus ?thesis by (simp add: weak_ranking_total_preorder) qed lemma weak_ranking_eqclass2: assumes A: "A \ set (weak_ranking le)" "x \ A" and le: "le x y" "le y x" shows "y \ A" proof - define xs where "xs = weak_ranking le" have wf: "is_weak_ranking xs" by (simp add: xs_def weak_ranking_total_preorder) let ?le' = "of_weak_ranking xs" from le have le': "?le' x y" "?le' y x" by (simp_all add: weak_ranking_total_preorder xs_def) from le'(1) obtain i j where ij: "j \ i" "i < length xs" "j < length xs" "x \ xs ! i" "y \ xs ! j" by (cases rule: of_weak_ranking.cases) from le'(2) obtain i' j' where i'j': "j' \ i'" "i' < length xs" "j' < length xs" "x \ xs ! j'" "y \ xs ! i'" by (cases rule: of_weak_ranking.cases) from ij i'j' have eq: "i = j'" "j = i'" by (intro weak_ranking_index_unique[OF wf]; simp)+ moreover from A obtain k where k: "k < length xs" "A = xs ! k" by (auto simp: xs_def set_conv_nth) ultimately have "k = i" using ij i'j' A by (intro weak_ranking_index_unique[OF wf, of _ _ x]) auto with ij i'j' k eq show ?thesis by (auto simp: xs_def) qed lemma hd_weak_ranking: assumes "x \ hd (weak_ranking le)" "y \ carrier" shows "le y x" proof - from weak_ranking_Union assms obtain i where "i < length (weak_ranking le)" "y \ weak_ranking le ! i" by (auto simp: set_conv_nth) moreover from assms(2) weak_ranking_Union have "weak_ranking le \ []" by auto ultimately have "of_weak_ranking (weak_ranking le) y x" using assms(1) by (intro of_weak_ranking.intros[of 0 i]) (auto simp: hd_conv_nth) thus ?thesis by (simp add: weak_ranking_total_preorder) qed lemma last_weak_ranking: assumes "x \ last (weak_ranking le)" "y \ carrier" shows "le x y" proof - from weak_ranking_Union assms obtain i where "i < length (weak_ranking le)" "y \ weak_ranking le ! i" by (auto simp: set_conv_nth) moreover from assms(2) weak_ranking_Union have "weak_ranking le \ []" by auto ultimately have "of_weak_ranking (weak_ranking le) x y" using assms(1) by (intro of_weak_ranking.intros[of i "length (weak_ranking le) - 1"]) (auto simp: last_conv_nth) thus ?thesis by (simp add: weak_ranking_total_preorder) qed text \ The index in weak ranking of a given alternative. An element with index 0 is first-ranked; larger indices correspond to less-preferred alternatives. \ definition weak_ranking_index :: "'a \ nat" where "weak_ranking_index x = find_index (\A. x \ A) (weak_ranking le)" lemma nth_weak_ranking_index: assumes "x \ carrier" shows "weak_ranking_index x < length (weak_ranking le)" "x \ weak_ranking le ! weak_ranking_index x" proof - from assms weak_ranking_Union show "weak_ranking_index x < length (weak_ranking le)" unfolding weak_ranking_index_def by (auto simp add: find_index_less_size_conv) thus "x \ weak_ranking le ! weak_ranking_index x" unfolding weak_ranking_index_def by (rule nth_find_index) qed lemma ranking_index_eqI: "i < length (weak_ranking le) \ x \ weak_ranking le ! i \ weak_ranking_index x = i" using weak_ranking_index_unique'[of "weak_ranking le" i x] by (simp add: weak_ranking_index_def weak_ranking_total_preorder) lemma ranking_index_le_iff [simp]: assumes "x \ carrier" "y \ carrier" shows "weak_ranking_index x \ weak_ranking_index y \ le x y" proof - have "le x y \ of_weak_ranking (weak_ranking le) x y" by (simp add: weak_ranking_total_preorder) also have "\ \ weak_ranking_index x \ weak_ranking_index y" proof assume "weak_ranking_index x \ weak_ranking_index y" thus "of_weak_ranking (weak_ranking le) x y" by (rule of_weak_ranking.intros) (simp_all add: nth_weak_ranking_index assms) next assume "of_weak_ranking (weak_ranking le) x y" then obtain i j where "i \ j" "i < length (weak_ranking le)" "j < length (weak_ranking le)" "x \ weak_ranking le ! j" "y \ weak_ranking le ! i" by (elim of_weak_ranking.cases) blast with ranking_index_eqI[of i] ranking_index_eqI[of j] show "weak_ranking_index x \ weak_ranking_index y" by simp qed finally show ?thesis .. qed end lemma weak_ranking_False [simp]: "weak_ranking (\_ _. False) = []" proof - interpret finite_total_preorder_on "{}" "\_ _. False" by unfold_locales simp_all have "[] = weak_ranking (\_ _. False)" by (rule weak_ranking_unique) simp_all thus ?thesis .. qed lemmas of_weak_ranking_weak_ranking = finite_total_preorder_on.weak_ranking_total_preorder(2) lemma finite_total_preorder_on_iff: "finite_total_preorder_on A R \ total_preorder_on A R \ finite A" by (simp add: finite_total_preorder_on_def finite_total_preorder_on_axioms_def) lemma finite_total_preorder_of_weak_ranking: assumes "\(set xs) = A" "is_finite_weak_ranking xs" shows "finite_total_preorder_on A (of_weak_ranking xs)" proof - from assms(2) have "is_weak_ranking xs" by (simp add: is_finite_weak_ranking_def) from assms(1) and this interpret total_preorder_on A "of_weak_ranking xs" by (rule total_preorder_of_weak_ranking) from assms(2) show ?thesis by unfold_locales (simp add: assms(1)[symmetric] is_finite_weak_ranking_def) qed lemma weak_ranking_of_weak_ranking: assumes "is_finite_weak_ranking xs" shows "weak_ranking (of_weak_ranking xs) = xs" proof - from assms interpret finite_total_preorder_on "\(set xs)" "of_weak_ranking xs" by (intro finite_total_preorder_of_weak_ranking) simp_all from assms show ?thesis by (intro sym[OF weak_ranking_unique]) (simp_all add: is_finite_weak_ranking_def) qed lemma weak_ranking_eqD: assumes "finite_total_preorder_on alts R1" assumes "finite_total_preorder_on alts R2" assumes "weak_ranking R1 = weak_ranking R2" shows "R1 = R2" proof - from assms have "of_weak_ranking (weak_ranking R1) = of_weak_ranking (weak_ranking R2)" by simp with assms(1,2) show ?thesis by (simp add: of_weak_ranking_weak_ranking) qed lemma weak_ranking_eq_iff: assumes "finite_total_preorder_on alts R1" assumes "finite_total_preorder_on alts R2" shows "weak_ranking R1 = weak_ranking R2 \ R1 = R2" using assms weak_ranking_eqD by auto definition preferred_alts :: "'alt relation \ 'alt \ 'alt set" where "preferred_alts R x = {y. y \[R] x}" lemma (in preorder_on) preferred_alts_refl [simp]: "x \ carrier \ x \ preferred_alts le x" by (simp add: preferred_alts_def refl) lemma (in preorder_on) preferred_alts_altdef: "preferred_alts le x = {y\carrier. y \[le] x}" by (auto simp: preferred_alts_def intro: not_outside) lemma (in preorder_on) preferred_alts_subset: "preferred_alts le x \ carrier" unfolding preferred_alts_def using not_outside by blast subsection \Rankings\ (* TODO: Extend theory on rankings. Can probably mostly be based on existing theory on weak rankings. *) definition ranking :: "'a relation \ 'a list" where "ranking R = map the_elem (weak_ranking R)" locale finite_linorder_on = linorder_on + assumes finite_carrier [intro]: "finite carrier" begin sublocale finite_total_preorder_on carrier le by unfold_locales (fact finite_carrier) lemma singleton_weak_ranking: assumes "A \ set (weak_ranking le)" shows "is_singleton A" proof (rule is_singletonI') from assms show "A \ {}" using weak_ranking_total_preorder(1) is_weak_ranking_iff by auto next fix x y assume "x \ A" "y \ A" with assms have "x \[of_weak_ranking (weak_ranking le)] y" "y \[of_weak_ranking (weak_ranking le)] x" by (auto intro!: of_weak_ranking_indifference) with weak_ranking_total_preorder(2) show "x = y" by (intro antisymmetric) simp_all qed lemma weak_ranking_ranking: "weak_ranking le = map (\x. {x}) (ranking le)" unfolding ranking_def map_map o_def proof (rule sym, rule map_idI) fix A assume "A \ set (weak_ranking le)" hence "is_singleton A" by (rule singleton_weak_ranking) thus "{the_elem A} = A" by (auto elim: is_singletonE) qed end end diff --git a/thys/Randomised_Social_Choice/ROOT b/thys/Randomised_Social_Choice/ROOT --- a/thys/Randomised_Social_Choice/ROOT +++ b/thys/Randomised_Social_Choice/ROOT @@ -1,14 +1,15 @@ chapter AFP session Randomised_Social_Choice (AFP) = "HOL-Probability" + options [timeout = 300] sessions + "HOL-Combinatorics" "List-Index" directories "Automation" theories Randomised_Social_Choice "Automation/SDS_Automation" document_files "root.tex" "root.bib" diff --git a/thys/Symmetric_Polynomials/ROOT b/thys/Symmetric_Polynomials/ROOT --- a/thys/Symmetric_Polynomials/ROOT +++ b/thys/Symmetric_Polynomials/ROOT @@ -1,14 +1,14 @@ chapter AFP session Symmetric_Polynomials (AFP) = "HOL-Computational_Algebra" + options [timeout = 600] sessions - "HOL-Library" + "HOL-Combinatorics" Polynomials theories Symmetric_Polynomials Symmetric_Polynomials_Code document_files "root.tex" "root.bib" diff --git a/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy b/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy --- a/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy +++ b/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy @@ -1,2890 +1,2890 @@ (* File: Symmetric_Polynomials.thy Author: Manuel Eberl (TU München) The definition of symmetric polynomials and the elementary symmetric polynomials. Proof of the fundamental theorem of symmetric polynomials and its corollaries. *) section \Symmetric Polynomials\ theory Symmetric_Polynomials imports Vieta "Polynomials.More_MPoly_Type" - "HOL-Library.Permutations" + "HOL-Combinatorics.Permutations" begin subsection \Auxiliary facts\ (* TODO: Many of these facts and definitions should be moved elsewhere, especially the ones on polynomials (leading monomial, mapping, insertion etc.) *) text \ An infinite set has infinitely many infinite subsets. \ lemma infinite_infinite_subsets: assumes "infinite A" shows "infinite {X. X \ A \ infinite X}" proof - have "\k. \X. X \ A \ infinite X \ card (A - X) = k" for k :: nat proof fix k :: nat obtain Y where "finite Y" "card Y = k" "Y \ A" using infinite_arbitrarily_large[of A k] assms by auto moreover from this have "A - (A - Y) = Y" by auto ultimately show "\X. X \ A \ infinite X \ card (A - X) = k" using assms by (intro exI[of _ "A - Y"]) auto qed from choice[OF this] obtain f where f: "\k. f k \ A \ infinite (f k) \ card (A - f k) = k" by blast have "k = l" if "f k = f l" for k l proof (rule ccontr) assume "k \ l" hence "card (A - f k) \ card (A - f l)" using f[of k] f[of l] by auto with \f k = f l\ show False by simp qed hence "inj f" by (auto intro: injI) moreover have "range f \ {X. X \ A \ infinite X}" using f by auto ultimately show ?thesis by (subst infinite_iff_countable_subset) auto qed text \ An infinite set contains infinitely many finite subsets of any fixed nonzero cardinality. \ lemma infinite_card_subsets: assumes "infinite A" "k > 0" shows "infinite {X. X \ A \ finite X \ card X = k}" proof - obtain B where B: "B \ A" "finite B" "card B = k - 1" using infinite_arbitrarily_large[OF assms(1), of "k - 1"] by blast define f where "f = (\x. insert x B)" have "f ` (A - B) \ {X. X \ A \ finite X \ card X = k}" using assms B by (auto simp: f_def) moreover have "inj_on f (A - B)" by (auto intro!: inj_onI simp: f_def) hence "infinite (f ` (A - B))" using assms B by (subst finite_image_iff) auto ultimately show ?thesis by (rule infinite_super) qed lemma comp_bij_eq_iff: assumes "bij f" shows "g \ f = h \ f \ g = h" proof assume *: "g \ f = h \ f" show "g = h" proof fix x obtain y where [simp]: "x = f y" using bij_is_surj[OF assms] by auto have "(g \ f) y = (h \ f) y" by (simp only: *) thus "g x = h x" by simp qed qed auto lemma sum_list_replicate [simp]: "sum_list (replicate n x) = of_nat n * (x :: 'a :: semiring_1)" by (induction n) (auto simp: algebra_simps) lemma ex_subset_of_card: assumes "finite A" "card A \ k" shows "\B. B \ A \ card B = k" using assms proof (induction arbitrary: k rule: finite_induct) case empty thus ?case by auto next case (insert x A k) show ?case proof (cases "k = 0") case True thus ?thesis by (intro exI[of _ "{}"]) auto next case False from insert have "\B\A. card B = k - 1" by (intro insert.IH) auto then obtain B where B: "B \ A" "card B = k - 1" by auto with insert have [simp]: "x \ B" by auto have "insert x B \ insert x A" using B insert by auto moreover have "card (insert x B) = k" using insert B finite_subset[of B A] False by (subst card.insert_remove) auto ultimately show ?thesis by blast qed qed lemma length_sorted_list_of_set [simp]: "length (sorted_list_of_set A) = card A" using distinct_card[of "sorted_list_of_set A"] by (cases "finite A") simp_all lemma upt_add_eq_append': "i \ j \ j \ k \ [i..Subrings and ring homomorphisms\ locale ring_closed = fixes A :: "'a :: comm_ring_1 set" assumes zero_closed [simp]: "0 \ A" assumes one_closed [simp]: "1 \ A" assumes add_closed [simp]: "x \ A \ y \ A \ (x + y) \ A" assumes mult_closed [simp]: "x \ A \ y \ A \ (x * y) \ A" assumes uminus_closed [simp]: "x \ A \ -x \ A" begin lemma minus_closed [simp]: "x \ A \ y \ A \ x - y \ A" using add_closed[of x "-y"] uminus_closed[of y] by auto lemma sum_closed [intro]: "(\x. x \ X \ f x \ A) \ sum f X \ A" by (induction X rule: infinite_finite_induct) auto lemma power_closed [intro]: "x \ A \ x ^ n \ A" by (induction n) auto lemma Sum_any_closed [intro]: "(\x. f x \ A) \ Sum_any f \ A" unfolding Sum_any.expand_set by (rule sum_closed) lemma prod_closed [intro]: "(\x. x \ X \ f x \ A) \ prod f X \ A" by (induction X rule: infinite_finite_induct) auto lemma Prod_any_closed [intro]: "(\x. f x \ A) \ Prod_any f \ A" unfolding Prod_any.expand_set by (rule prod_closed) lemma prod_fun_closed [intro]: "(\x. f x \ A) \ (\x. g x \ A) \ prod_fun f g x \ A" by (auto simp: prod_fun_def when_def intro!: Sum_any_closed mult_closed) lemma of_nat_closed [simp, intro]: "of_nat n \ A" by (induction n) auto lemma of_int_closed [simp, intro]: "of_int n \ A" by (induction n) auto end locale ring_homomorphism = fixes f :: "'a :: comm_ring_1 \ 'b :: comm_ring_1" assumes add[simp]: "f (x + y) = f x + f y" assumes uminus[simp]: "f (-x) = -f x" assumes mult[simp]: "f (x * y) = f x * f y" assumes zero[simp]: "f 0 = 0" assumes one [simp]: "f 1 = 1" begin lemma diff [simp]: "f (x - y) = f x - f y" using add[of x "-y"] by (simp del: add) lemma power [simp]: "f (x ^ n) = f x ^ n" by (induction n) auto lemma sum [simp]: "f (sum g A) = (\x\A. f (g x))" by (induction A rule: infinite_finite_induct) auto lemma prod [simp]: "f (prod g A) = (\x\A. f (g x))" by (induction A rule: infinite_finite_induct) auto end lemma ring_homomorphism_id [intro]: "ring_homomorphism id" by standard auto lemma ring_homomorphism_id' [intro]: "ring_homomorphism (\x. x)" by standard auto lemma ring_homomorphism_of_int [intro]: "ring_homomorphism of_int" by standard auto subsection \Various facts about multivariate polynomials\ lemma poly_mapping_nat_ge_0 [simp]: "(m :: nat \\<^sub>0 nat) \ 0" proof (cases "m = 0") case False hence "Poly_Mapping.lookup m \ Poly_Mapping.lookup 0" by transfer auto hence "\k. Poly_Mapping.lookup m k \ 0" by (auto simp: fun_eq_iff) from LeastI_ex[OF this] Least_le[of "\k. Poly_Mapping.lookup m k \ 0"] show ?thesis by (force simp: less_eq_poly_mapping_def less_fun_def) qed auto lemma poly_mapping_nat_le_0 [simp]: "(m :: nat \\<^sub>0 nat) \ 0 \ m = 0" unfolding less_eq_poly_mapping_def poly_mapping_eq_iff less_fun_def by auto lemma of_nat_diff_poly_mapping_nat: assumes "m \ n" shows "of_nat (m - n) = (of_nat m - of_nat n :: 'a :: monoid_add \\<^sub>0 nat)" by (auto intro!: poly_mapping_eqI simp: lookup_of_nat lookup_minus when_def) lemma mpoly_coeff_transfer [transfer_rule]: "rel_fun cr_mpoly (=) poly_mapping.lookup MPoly_Type.coeff" unfolding MPoly_Type.coeff_def by transfer_prover lemma mapping_of_sum: "(\x\A. mapping_of (f x)) = mapping_of (sum f A)" by (induction A rule: infinite_finite_induct) (auto simp: plus_mpoly.rep_eq zero_mpoly.rep_eq) lemma mapping_of_eq_0_iff [simp]: "mapping_of p = 0 \ p = 0" by transfer auto lemma Sum_any_mapping_of: "Sum_any (\x. mapping_of (f x)) = mapping_of (Sum_any f)" by (simp add: Sum_any.expand_set mapping_of_sum) lemma Sum_any_parametric_cr_mpoly [transfer_rule]: "(rel_fun (rel_fun (=) cr_mpoly) cr_mpoly) Sum_any Sum_any" by (auto simp: rel_fun_def cr_mpoly_def Sum_any_mapping_of) lemma lookup_mult_of_nat [simp]: "lookup (of_nat n * m) k = n * lookup m k" proof - have "of_nat n * m = (\i k = (\i = n * lookup m k" by simp finally show ?thesis . qed lemma mpoly_eqI: assumes "\mon. MPoly_Type.coeff p mon = MPoly_Type.coeff q mon" shows "p = q" using assms by (transfer, transfer) (auto simp: fun_eq_iff) lemma coeff_mpoly_times: "MPoly_Type.coeff (p * q) mon = prod_fun (MPoly_Type.coeff p) (MPoly_Type.coeff q) mon" by (transfer', transfer') auto lemma (in ring_closed) coeff_mult_closed [intro]: "(\x. coeff p x \ A) \ (\x. coeff q x \ A) \ coeff (p * q) x \ A" by (auto simp: coeff_mpoly_times prod_fun_closed) lemma coeff_notin_vars: assumes "\(keys m \ vars p)" shows "coeff p m = 0" using assms unfolding vars_def by transfer' (auto simp: in_keys_iff) lemma finite_coeff_support [intro]: "finite {m. coeff p m \ 0}" by transfer simp lemma insertion_altdef: "insertion f p = Sum_any (\m. coeff p m * Prod_any (\i. f i ^ lookup m i))" by (transfer', transfer') (simp add: insertion_fun_def) lemma mpoly_coeff_uminus [simp]: "coeff (-p) m = -coeff p m" by transfer auto lemma Sum_any_uminus: "Sum_any (\x. -f x :: 'a :: ab_group_add) = -Sum_any f" by (simp add: Sum_any.expand_set sum_negf) lemma insertion_uminus [simp]: "insertion f (-p :: 'a :: comm_ring_1 mpoly) = -insertion f p" by (simp add: insertion_altdef Sum_any_uminus) lemma Sum_any_lookup: "finite {x. g x \ 0} \ Sum_any (\x. lookup (g x) y) = lookup (Sum_any g) y" by (auto simp: Sum_any.expand_set lookup_sum intro!: sum.mono_neutral_left) lemma Sum_any_diff: assumes "finite {x. f x \ 0}" assumes "finite {x. g x \ 0}" shows "Sum_any (\x. f x - g x :: 'a :: ab_group_add) = Sum_any f - Sum_any g" proof - have "{x. f x - g x \ 0} \ {x. f x \ 0} \ {x. g x \ 0}" by auto moreover have "finite ({x. f x \ 0} \ {x. g x \ 0})" by (subst finite_Un) (insert assms, auto) ultimately have "finite {x. f x - g x \ 0}" by (rule finite_subset) with assms show ?thesis by (simp add: algebra_simps Sum_any.distrib [symmetric]) qed lemma insertion_diff: "insertion f (p - q :: 'a :: comm_ring_1 mpoly) = insertion f p - insertion f q" proof (transfer, transfer) fix f :: "nat \ 'a" and p q :: "(nat \\<^sub>0 nat) \ 'a" assume fin: "finite {x. p x \ 0}" "finite {x. q x \ 0}" have "insertion_fun f (\x. p x - q x) = (\m. p m * (\v. f v ^ lookup m v) - q m * (\v. f v ^ lookup m v))" by (simp add: insertion_fun_def algebra_simps Sum_any_diff) also have "\ = (\m. p m * (\v. f v ^ lookup m v)) - (\m. q m * (\v. f v ^ lookup m v))" by (subst Sum_any_diff) (auto intro: finite_subset[OF _ fin(1)] finite_subset[OF _ fin(2)]) also have "\ = insertion_fun f p - insertion_fun f q" by (simp add: insertion_fun_def) finally show "insertion_fun f (\x. p x - q x) = \" . qed lemma insertion_power: "insertion f (p ^ n) = insertion f p ^ n" by (induction n) (simp_all add: insertion_mult) lemma insertion_sum: "insertion f (sum g A) = (\x\A. insertion f (g x))" by (induction A rule: infinite_finite_induct) (auto simp: insertion_add) lemma insertion_prod: "insertion f (prod g A) = (\x\A. insertion f (g x))" by (induction A rule: infinite_finite_induct) (auto simp: insertion_mult) lemma coeff_Var: "coeff (Var i) m = (1 when m = Poly_Mapping.single i 1)" by transfer' (auto simp: Var\<^sub>0_def lookup_single when_def) lemma vars_Var: "vars (Var i :: 'a :: {one,zero} mpoly) = (if (0::'a) = 1 then {} else {i})" unfolding vars_def by (auto simp: Var.rep_eq Var\<^sub>0_def) lemma insertion_Var [simp]: "insertion f (Var i) = f i" proof - have "insertion f (Var i) = (\m. (1 when m = Poly_Mapping.single i 1) * (\i. f i ^ lookup m i))" by (simp add: insertion_altdef coeff_Var) also have "\ = (\j. f j ^ lookup (Poly_Mapping.single i 1) j)" by (subst Sum_any.expand_superset[of "{Poly_Mapping.single i 1}"]) (auto simp: when_def) also have "\ = f i" by (subst Prod_any.expand_superset[of "{i}"]) (auto simp: when_def lookup_single) finally show ?thesis . qed lemma insertion_Sum_any: assumes "finite {x. g x \ 0}" shows "insertion f (Sum_any g) = Sum_any (\x. insertion f (g x))" unfolding Sum_any.expand_set insertion_sum by (intro sum.mono_neutral_right) (auto intro!: finite_subset[OF _ assms]) lemma keys_diff_subset: "keys (f - g) \ keys f \ keys g" by transfer auto lemma keys_empty_iff [simp]: "keys p = {} \ p = 0" by transfer auto lemma mpoly_coeff_0 [simp]: "MPoly_Type.coeff 0 m = 0" by transfer auto lemma lookup_1: "lookup 1 m = (if m = 0 then 1 else 0)" by transfer (simp add: when_def) lemma mpoly_coeff_1: "MPoly_Type.coeff 1 m = (if m = 0 then 1 else 0)" by (simp add: MPoly_Type.coeff_def one_mpoly.rep_eq lookup_1) lemma lookup_Const\<^sub>0: "lookup (Const\<^sub>0 c) m = (if m = 0 then c else 0)" unfolding Const\<^sub>0_def by (simp add: lookup_single when_def) lemma mpoly_coeff_Const: "MPoly_Type.coeff (Const c) m = (if m = 0 then c else 0)" by (simp add: MPoly_Type.coeff_def Const.rep_eq lookup_Const\<^sub>0) lemma coeff_smult [simp]: "coeff (smult c p) m = (c :: 'a :: mult_zero) * coeff p m" by transfer (auto simp: map_lookup) lemma in_keys_mapI: "x \ keys m \ f (lookup m x) \ 0 \ x \ keys (Poly_Mapping.map f m)" by transfer auto lemma keys_uminus [simp]: "keys (-m) = keys m" by transfer auto lemma vars_uminus [simp]: "vars (-p) = vars p" unfolding vars_def by transfer' auto lemma vars_smult: "vars (smult c p) \ vars p" unfolding vars_def by (transfer', transfer') auto lemma vars_0 [simp]: "vars 0 = {}" unfolding vars_def by transfer' simp lemma vars_1 [simp]: "vars 1 = {}" unfolding vars_def by transfer' simp lemma vars_sum: "vars (sum f A) \ (\x\A. vars (f x))" using vars_add by (induction A rule: infinite_finite_induct) auto lemma vars_prod: "vars (prod f A) \ (\x\A. vars (f x))" using vars_mult by (induction A rule: infinite_finite_induct) auto lemma vars_Sum_any: "vars (Sum_any h) \ (\i. vars (h i))" unfolding Sum_any.expand_set by (intro order.trans[OF vars_sum]) auto lemma vars_Prod_any: "vars (Prod_any h) \ (\i. vars (h i))" unfolding Prod_any.expand_set by (intro order.trans[OF vars_prod]) auto lemma vars_power: "vars (p ^ n) \ vars p" using vars_mult by (induction n) auto lemma vars_diff: "vars (p1 - p2) \ vars p1 \ vars p2" unfolding vars_def proof transfer' fix p1 p2 :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" show "\ (keys ` keys (p1 - p2)) \ \(keys ` (keys p1)) \ \(keys ` (keys p2))" using keys_diff_subset[of p1 p2] by (auto simp flip: not_in_keys_iff_lookup_eq_zero) qed lemma insertion_smult [simp]: "insertion f (smult c p) = c * insertion f p" unfolding insertion_altdef by (subst Sum_any_right_distrib) (auto intro: finite_subset[OF _ finite_coeff_support[of p]] simp: mult.assoc) lemma coeff_add [simp]: "coeff (p + q) m = coeff p m + coeff q m" by transfer' (simp add: lookup_add) lemma coeff_diff [simp]: "coeff (p - q) m = coeff p m - coeff q m" by transfer' (simp add: lookup_minus) lemma insertion_monom [simp]: "insertion f (monom m c) = c * Prod_any (\x. f x ^ lookup m x)" proof - have "insertion f (monom m c) = (\m'. (c when m = m') * (\v. f v ^ lookup m' v))" by (simp add: insertion_def insertion_aux_def insertion_fun_def lookup_single) also have "\ = c * (\v. f v ^ lookup m v)" by (subst Sum_any.expand_superset[of "{m}"]) (auto simp: when_def) finally show ?thesis . qed lemma insertion_aux_Const\<^sub>0 [simp]: "insertion_aux f (Const\<^sub>0 c) = c" proof - have "insertion_aux f (Const\<^sub>0 c) = (\m. (c when m = 0) * (\v. f v ^ lookup m v))" by (simp add: Const\<^sub>0_def insertion_aux_def insertion_fun_def lookup_single) also have "\ = (\m\{0}. (c when m = 0) * (\v. f v ^ lookup m v))" by (intro Sum_any.expand_superset) (auto simp: when_def) also have "\ = c" by simp finally show ?thesis . qed lemma insertion_Const [simp]: "insertion f (Const c) = c" by (simp add: insertion_def Const.rep_eq) lemma coeffs_0 [simp]: "coeffs 0 = {}" by transfer auto lemma coeffs_1 [simp]: "coeffs 1 = {1}" by transfer auto lemma coeffs_Const: "coeffs (Const c) = (if c = 0 then {} else {c})" unfolding Const_def Const\<^sub>0_def by transfer' auto lemma coeffs_subset: "coeffs (Const c) \ {c}" by (auto simp: coeffs_Const) lemma keys_Const\<^sub>0: "keys (Const\<^sub>0 c) = (if c = 0 then {} else {0})" unfolding Const\<^sub>0_def by transfer' auto lemma vars_Const [simp]: "vars (Const c) = {}" unfolding vars_def by transfer' (auto simp: keys_Const\<^sub>0) lemma prod_fun_compose_bij: assumes "bij f" and f: "\x y. f (x + y) = f x + f y" shows "prod_fun m1 m2 (f x) = prod_fun (m1 \ f) (m2 \ f) x" proof - have [simp]: "f x = f y \ x = y" for x y using \bij f\ by (auto dest!: bij_is_inj inj_onD) have "prod_fun (m1 \ f) (m2 \ f) x = Sum_any ((\l. m1 l * Sum_any ((\q. m2 q when f x = l + q) \ f)) \ f)" by (simp add: prod_fun_def f(1) [symmetric] o_def) also have "\ = Sum_any ((\l. m1 l * Sum_any ((\q. m2 q when f x = l + q))))" by (simp only: Sum_any.reindex_cong[OF assms(1) refl, symmetric]) also have "\ = prod_fun m1 m2 (f x)" by (simp add: prod_fun_def) finally show ?thesis .. qed lemma add_nat_poly_mapping_zero_iff [simp]: "(a + b :: 'a \\<^sub>0 nat) = 0 \ a = 0 \ b = 0" by transfer (auto simp: fun_eq_iff) lemma prod_fun_nat_0: fixes f g :: "('a \\<^sub>0 nat) \ 'b::semiring_0" shows "prod_fun f g 0 = f 0 * g 0" proof - have "prod_fun f g 0 = (\l. f l * (\q. g q when 0 = l + q))" unfolding prod_fun_def .. also have "(\l. \q. g q when 0 = l + q) = (\l. \q\{0}. g q when 0 = l + q)" by (intro ext Sum_any.expand_superset) (auto simp: when_def) also have "(\l. f l * \ l) = (\l\{0}. f l * \ l)" by (intro ext Sum_any.expand_superset) (auto simp: when_def) finally show ?thesis by simp qed lemma mpoly_coeff_times_0: "coeff (p * q) 0 = coeff p 0 * coeff q 0" by (simp add: coeff_mpoly_times prod_fun_nat_0) lemma mpoly_coeff_prod_0: "coeff (\x\A. f x) 0 = (\x\A. coeff (f x) 0)" by (induction A rule: infinite_finite_induct) (auto simp: mpoly_coeff_times_0 mpoly_coeff_1) lemma mpoly_coeff_power_0: "coeff (p ^ n) 0 = coeff p 0 ^ n" by (induction n) (auto simp: mpoly_coeff_times_0 mpoly_coeff_1) lemma prod_fun_max: fixes f g :: "'a::{linorder, ordered_cancel_comm_monoid_add} \ 'b::semiring_0" assumes zero: "\m. m > a \ f m = 0" "\m. m > b \ g m = 0" assumes fin: "finite {m. f m \ 0}" "finite {m. g m \ 0}" shows "prod_fun f g (a + b) = f a * g b" proof - note fin' = finite_subset[OF _ fin(1)] finite_subset[OF _ fin(2)] have "prod_fun f g (a + b) = (\l. f l * (\q. g q when a + b = l + q))" by (simp add: prod_fun_def Sum_any_right_distrib) also have "\ = (\l. \q. f l * g q when a + b = l + q)" by (subst Sum_any_right_distrib) (auto intro!: Sum_any.cong fin'(2) simp: when_def) also { fix l q assume lq: "a + b = l + q" "(a, b) \ (l, q)" and nz: "f l * g q \ 0" from nz and zero have "l \ a" "q \ b" by (auto intro: leI) moreover from this and lq(2) have "l < a \ q < b" by auto ultimately have "l + q < a + b" by (auto intro: add_less_le_mono add_le_less_mono) with lq(1) have False by simp } hence "(\l. \q. f l * g q when a + b = l + q) = (\l. \q. f l * g q when (a, b) = (l, q))" by (intro Sum_any.cong refl) (auto simp: when_def) also have "\ = (\(l,q). f l * g q when (a, b) = (l, q))" by (intro Sum_any.cartesian_product[of "{(a, b)}"]) auto also have "\ = (\(l,q)\{(a,b)}. f l * g q when (a, b) = (l, q))" by (intro Sum_any.expand_superset) auto also have "\ = f a * g b" by simp finally show ?thesis . qed lemma prod_fun_gt_max_eq_zero: fixes f g :: "'a::{linorder, ordered_cancel_comm_monoid_add} \ 'b::semiring_0" assumes "m > a + b" assumes zero: "\m. m > a \ f m = 0" "\m. m > b \ g m = 0" assumes fin: "finite {m. f m \ 0}" "finite {m. g m \ 0}" shows "prod_fun f g m = 0" proof - note fin' = finite_subset[OF _ fin(1)] finite_subset[OF _ fin(2)] have "prod_fun f g m = (\l. f l * (\q. g q when m = l + q))" by (simp add: prod_fun_def Sum_any_right_distrib) also have "\ = (\l. \q. f l * g q when m = l + q)" by (subst Sum_any_right_distrib) (auto intro!: Sum_any.cong fin'(2) simp: when_def) also { fix l q assume lq: "m = l + q" and nz: "f l * g q \ 0" from nz and zero have "l \ a" "q \ b" by (auto intro: leI) hence "l + q \ a + b" by (intro add_mono) also have "\ < m" by fact finally have "l + q < m" . } hence "(\l. \q. f l * g q when m = l + q) = (\l. \q. f l * g q when False)" by (intro Sum_any.cong refl) (auto simp: when_def) also have "\ = 0" by simp finally show ?thesis . qed subsection \Restricting a monomial to a subset of variables\ lift_definition restrictpm :: "'a set \ ('a \\<^sub>0 'b :: zero) \ ('a \\<^sub>0 'b)" is "\A f x. if x \ A then f x else 0" by (erule finite_subset[rotated]) auto lemma lookup_restrictpm: "lookup (restrictpm A m) x = (if x \ A then lookup m x else 0)" by transfer auto lemma lookup_restrictpm_in [simp]: "x \ A \ lookup (restrictpm A m) x = lookup m x" and lookup_restrict_pm_not_in [simp]: "x \ A \ lookup (restrictpm A m) x = 0" by (simp_all add: lookup_restrictpm) lemma keys_restrictpm [simp]: "keys (restrictpm A m) = keys m \ A" by transfer auto lemma restrictpm_add: "restrictpm X (m1 + m2) = restrictpm X m1 + restrictpm X m2" by transfer auto lemma restrictpm_id [simp]: "keys m \ X \ restrictpm X m = m" by transfer (auto simp: fun_eq_iff) lemma restrictpm_orthogonal [simp]: "keys m \ -X \ restrictpm X m = 0" by transfer (auto simp: fun_eq_iff) lemma restrictpm_add_disjoint: "X \ Y = {} \ restrictpm X m + restrictpm Y m = restrictpm (X \ Y) m" by transfer (auto simp: fun_eq_iff) lemma restrictpm_add_complements: "restrictpm X m + restrictpm (-X) m = m" "restrictpm (-X) m + restrictpm X m = m" by (subst restrictpm_add_disjoint; force)+ subsection \Mapping over a polynomial\ lift_definition map_mpoly :: "('a :: zero \ 'b :: zero) \ 'a mpoly \ 'b mpoly" is "\(f :: 'a \ 'b) (p :: (nat \\<^sub>0 nat) \\<^sub>0 'a). Poly_Mapping.map f p" . lift_definition mapm_mpoly :: "((nat \\<^sub>0 nat) \ 'a :: zero \ 'b :: zero) \ 'a mpoly \ 'b mpoly" is "\(f :: (nat \\<^sub>0 nat) \ 'a \ 'b) (p :: (nat \\<^sub>0 nat) \\<^sub>0 'a). Poly_Mapping.mapp f p" . lemma poly_mapping_map_conv_mapp: "Poly_Mapping.map f = Poly_Mapping.mapp (\_. f)" by (auto simp: Poly_Mapping.mapp_def Poly_Mapping.map_def map_fun_def o_def fun_eq_iff when_def in_keys_iff cong: if_cong) lemma map_mpoly_conv_mapm_mpoly: "map_mpoly f = mapm_mpoly (\_. f)" by transfer' (auto simp: poly_mapping_map_conv_mapp) lemma map_mpoly_comp: "f 0 = 0 \ map_mpoly f (map_mpoly g p) = map_mpoly (f \ g) p" by (transfer', transfer') (auto simp: when_def fun_eq_iff) lemma mapp_mapp: "(\x. f x 0 = 0) \ Poly_Mapping.mapp f (Poly_Mapping.mapp g m) = Poly_Mapping.mapp (\x y. f x (g x y)) m" by transfer' (auto simp: fun_eq_iff lookup_mapp in_keys_iff) lemma mapm_mpoly_comp: "(\x. f x 0 = 0) \ mapm_mpoly f (mapm_mpoly g p) = mapm_mpoly (\m c. f m (g m c)) p" by transfer' (simp add: mapp_mapp) lemma coeff_map_mpoly: "coeff (map_mpoly f p) m = (if coeff p m = 0 then 0 else f (coeff p m))" by (transfer, transfer) auto lemma coeff_map_mpoly' [simp]: "f 0 = 0 \ coeff (map_mpoly f p) m = f (coeff p m)" by (subst coeff_map_mpoly) auto lemma coeff_mapm_mpoly: "coeff (mapm_mpoly f p) m = (if coeff p m = 0 then 0 else f m (coeff p m))" by (transfer, transfer') (auto simp: in_keys_iff) lemma coeff_mapm_mpoly' [simp]: "(\m. f m 0 = 0) \ coeff (mapm_mpoly f p) m = f m (coeff p m)" by (subst coeff_mapm_mpoly) auto lemma vars_map_mpoly_subset: "vars (map_mpoly f p) \ vars p" unfolding vars_def by (transfer', transfer') (auto simp: map_mpoly.rep_eq) lemma coeff_sum [simp]: "coeff (sum f A) m = (\x\A. coeff (f x) m)" by (induction A rule: infinite_finite_induct) auto lemma coeff_Sum_any: "finite {x. f x \ 0} \ coeff (Sum_any f) m = Sum_any (\x. coeff (f x) m)" by (auto simp add: Sum_any.expand_set intro!: sum.mono_neutral_right) lemma Sum_any_zeroI: "(\x. f x = 0) \ Sum_any f = 0" by (auto simp: Sum_any.expand_set) lemma insertion_Prod_any: "finite {x. g x \ 1} \ insertion f (Prod_any g) = Prod_any (\x. insertion f (g x))" by (auto simp: Prod_any.expand_set insertion_prod intro!: prod.mono_neutral_right) lemma insertion_insertion: "insertion g (insertion k p) = insertion (\x. insertion g (k x)) (map_mpoly (insertion g) p)" (is "?lhs = ?rhs") proof - have "insertion g (insertion k p) = (\x. insertion g (coeff p x) * insertion g (\i. k i ^ lookup x i))" unfolding insertion_altdef[of k p] by (subst insertion_Sum_any) (auto intro: finite_subset[OF _ finite_coeff_support[of p]] simp: insertion_mult) also have "\ = (\x. insertion g (coeff p x) * (\i. insertion g (k i) ^ lookup x i))" proof (intro Sum_any.cong) fix x show "insertion g (coeff p x) * insertion g (\i. k i ^ lookup x i) = insertion g (coeff p x) * (\i. insertion g (k i) ^ lookup x i)" by (subst insertion_Prod_any) (auto simp: insertion_power intro!: finite_subset[OF _ finite_lookup[of x]] Nat.gr0I) qed also have "\ = insertion (\x. insertion g (k x)) (map_mpoly (insertion g) p)" unfolding insertion_altdef[of _ "map_mpoly f p" for f] by auto finally show ?thesis . qed lemma insertion_substitute_linear: "insertion (\i. c i * f i) p = insertion f (mapm_mpoly (\m d. Prod_any (\i. c i ^ lookup m i) * d) p)" unfolding insertion_altdef proof (intro Sum_any.cong, goal_cases) case (1 m) have "coeff (mapm_mpoly (\m. (*) (\i. c i ^ lookup m i)) p) m * (\i. f i ^ lookup m i) = MPoly_Type.coeff p m * ((\i. c i ^ lookup m i) * (\i. f i ^ lookup m i))" by (simp add: mult_ac) also have "(\i. c i ^ lookup m i) * (\i. f i ^ lookup m i) = (\i. (c i * f i) ^ lookup m i)" by (subst Prod_any.distrib [symmetric]) (auto simp: power_mult_distrib intro!: finite_subset[OF _ finite_lookup[of m]] Nat.gr0I) finally show ?case by simp qed lemma vars_mapm_mpoly_subset: "vars (mapm_mpoly f p) \ vars p" unfolding vars_def using keys_mapp_subset[of f] by (auto simp: mapm_mpoly.rep_eq) lemma map_mpoly_cong: assumes "\m. f (coeff p m) = g (coeff p m)" "p = q" shows "map_mpoly f p = map_mpoly g q" using assms by (intro mpoly_eqI) (auto simp: coeff_map_mpoly) subsection \The leading monomial and leading coefficient\ text \ The leading monomial of a multivariate polynomial is the one with the largest monomial w.\,r.\,t.\ the monomial ordering induced by the standard variable ordering. The leading coefficient is the coefficient of the leading monomial. As a convention, the leading monomial of the zero polynomial is defined to be the same as that of any non-constant zero polynomial, i.\,e.\ the monomial $X_1^0 \ldots X_n^0$. \ lift_definition lead_monom :: "'a :: zero mpoly \ (nat \\<^sub>0 nat)" is "\f :: (nat \\<^sub>0 nat) \\<^sub>0 'a. Max (insert 0 (keys f))" . lemma lead_monom_geI [intro]: assumes "coeff p m \ 0" shows "m \ lead_monom p" using assms by (auto simp: lead_monom_def coeff_def in_keys_iff) lemma coeff_gt_lead_monom_zero [simp]: assumes "m > lead_monom p" shows "coeff p m = 0" using lead_monom_geI[of p m] assms by force lemma lead_monom_nonzero_eq: assumes "p \ 0" shows "lead_monom p = Max (keys (mapping_of p))" using assms by transfer (simp add: max_def) lemma lead_monom_0 [simp]: "lead_monom 0 = 0" by (simp add: lead_monom_def zero_mpoly.rep_eq) lemma lead_monom_1 [simp]: "lead_monom 1 = 0" by (simp add: lead_monom_def one_mpoly.rep_eq) lemma lead_monom_Const [simp]: "lead_monom (Const c) = 0" by (simp add: lead_monom_def Const.rep_eq Const\<^sub>0_def) lemma lead_monom_uminus [simp]: "lead_monom (-p) = lead_monom p" by (simp add: lead_monom_def uminus_mpoly.rep_eq) lemma keys_mult_const [simp]: fixes c :: "'a :: {semiring_0, semiring_no_zero_divisors}" assumes "c \ 0" shows "keys (Poly_Mapping.map ((*) c) p) = keys p" using assms by transfer auto lemma lead_monom_eq_0_iff: "lead_monom p = 0 \ vars p = {}" unfolding vars_def by transfer' (auto simp: Max_eq_iff) lemma lead_monom_monom: "lead_monom (monom m c) = (if c = 0 then 0 else m)" by (auto simp add: lead_monom_def monom.rep_eq Const\<^sub>0_def max_def ) lemma lead_monom_monom' [simp]: "c \ 0 \ lead_monom (monom m c) = m" by (simp add: lead_monom_monom) lemma lead_monom_numeral [simp]: "lead_monom (numeral n) = 0" unfolding monom_numeral[symmetric] by (subst lead_monom_monom) auto lemma lead_monom_add: "lead_monom (p + q) \ max (lead_monom p) (lead_monom q)" proof transfer fix p q :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" show "Max (insert 0 (keys (p + q))) \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof (rule Max.boundedI) fix m assume m: "m \ insert 0 (keys (p + q))" thus "m \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof assume "m \ keys (p + q)" with keys_add[of p q] have "m \ keys p \ m \ keys q" by (auto simp: in_keys_iff plus_poly_mapping.rep_eq) thus ?thesis by (auto simp: le_max_iff_disj) qed auto qed auto qed lemma lead_monom_diff: "lead_monom (p - q) \ max (lead_monom p) (lead_monom q)" proof transfer fix p q :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" show "Max (insert 0 (keys (p - q))) \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof (rule Max.boundedI) fix m assume m: "m \ insert 0 (keys (p - q))" thus "m \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof assume "m \ keys (p - q)" with keys_diff_subset[of p q] have "m \ keys p \ m \ keys q" by auto thus ?thesis by (auto simp: le_max_iff_disj) qed auto qed auto qed definition lead_coeff where "lead_coeff p = coeff p (lead_monom p)" lemma vars_empty_iff: "vars p = {} \ p = Const (lead_coeff p)" proof assume "vars p = {}" hence [simp]: "lead_monom p = 0" by (simp add: lead_monom_eq_0_iff) have [simp]: "mon \ 0 \ (mon > (0 :: nat \\<^sub>0 nat))" for mon by (auto simp: order.strict_iff_order) thus "p = Const (lead_coeff p)" by (intro mpoly_eqI) (auto simp: mpoly_coeff_Const lead_coeff_def) next assume "p = Const (lead_coeff p)" also have "vars \ = {}" by simp finally show "vars p = {}" . qed lemma lead_coeff_0 [simp]: "lead_coeff 0 = 0" by (simp add: lead_coeff_def) lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" by (simp add: lead_coeff_def mpoly_coeff_1) lemma lead_coeff_Const [simp]: "lead_coeff (Const c) = c" by (simp add: lead_coeff_def mpoly_coeff_Const) lemma lead_coeff_monom [simp]: "lead_coeff (monom p c) = c" by (simp add: lead_coeff_def coeff_monom when_def lead_monom_monom) lemma lead_coeff_nonzero [simp]: "p \ 0 \ lead_coeff p \ 0" unfolding lead_coeff_def lead_monom_def by (cases "keys (mapping_of p) = {}") (auto simp: coeff_def max_def) lemma fixes c :: "'a :: semiring_0" assumes "c * lead_coeff p \ 0" shows lead_monom_smult [simp]: "lead_monom (smult c p) = lead_monom p" and lead_coeff_smult [simp]: "lead_coeff (smult c p) = c * lead_coeff p" proof - from assms have *: "keys (mapping_of p) \ {}" by auto from assms have "coeff (MPoly_Type.smult c p) (lead_monom p) \ 0" by (simp add: lead_coeff_def) hence smult_nz: "MPoly_Type.smult c p \ 0" by (auto simp del: coeff_smult) with assms have **: "keys (mapping_of (smult c p)) \ {}" by simp have "Max (keys (mapping_of (smult c p))) = Max (keys (mapping_of p))" proof (safe intro!: antisym Max.coboundedI) have "lookup (mapping_of p) (Max (keys (mapping_of p))) = lead_coeff p" using * by (simp add: lead_coeff_def lead_monom_def max_def coeff_def) with assms show "Max (keys (mapping_of p)) \ keys (mapping_of (smult c p))" using * by (auto simp: smult.rep_eq intro!: in_keys_mapI) from smult_nz have "lead_coeff (smult c p) \ 0" by (intro lead_coeff_nonzero) auto hence "coeff p (Max (keys (mapping_of (smult c p)))) \ 0" using assms * ** by (auto simp: lead_coeff_def lead_monom_def max_def) thus "Max (keys (mapping_of (smult c p))) \ keys (mapping_of p)" by (auto simp: smult.rep_eq coeff_def in_keys_iff) qed auto with * ** show "lead_monom (smult c p) = lead_monom p" by (simp add: lead_monom_def max_def) thus "lead_coeff (smult c p) = c * lead_coeff p" by (simp add: lead_coeff_def) qed lemma lead_coeff_mult_aux: "coeff (p * q) (lead_monom p + lead_monom q) = lead_coeff p * lead_coeff q" proof (cases "p = 0 \ q = 0") case False define a b where "a = lead_monom p" and "b = lead_monom q" have "coeff (p * q) (a + b) = coeff p a * coeff q b" unfolding coeff_mpoly_times by (rule prod_fun_max) (insert False, auto simp: a_def b_def) thus ?thesis by (simp add: a_def b_def lead_coeff_def) qed auto lemma lead_monom_mult_le: "lead_monom (p * q) \ lead_monom p + lead_monom q" proof (cases "p * q = 0") case False show ?thesis proof (intro leI notI) assume "lead_monom p + lead_monom q < lead_monom (p * q)" hence "lead_coeff (p * q) = 0" unfolding lead_coeff_def coeff_mpoly_times by (rule prod_fun_gt_max_eq_zero) auto with False show False by simp qed qed auto lemma lead_monom_mult: assumes "lead_coeff p * lead_coeff q \ 0" shows "lead_monom (p * q) = lead_monom p + lead_monom q" by (intro antisym lead_monom_mult_le lead_monom_geI) (insert assms, auto simp: lead_coeff_mult_aux) lemma lead_coeff_mult: assumes "lead_coeff p * lead_coeff q \ 0" shows "lead_coeff (p * q) = lead_coeff p * lead_coeff q" using assms by (simp add: lead_monom_mult lead_coeff_mult_aux lead_coeff_def) lemma keys_lead_monom_subset: "keys (lead_monom p) \ vars p" proof (cases "p = 0") case False hence "lead_coeff p \ 0" by simp hence "coeff p (lead_monom p) \ 0" unfolding lead_coeff_def . thus ?thesis unfolding vars_def by transfer' (auto simp: max_def in_keys_iff) qed auto lemma assumes "(\i\A. lead_coeff (f i)) \ 0" shows lead_monom_prod: "lead_monom (\i\A. f i) = (\i\A. lead_monom (f i))" (is ?th1) and lead_coeff_prod: "lead_coeff (\i\A. f i) = (\i\A. lead_coeff (f i))" (is ?th2) proof - have "?th1 \ ?th2" using assms proof (induction A rule: infinite_finite_induct) case (insert x A) from insert have nz: "lead_coeff (f x) \ 0" "(\i\A. lead_coeff (f i)) \ 0" by auto note IH = insert.IH[OF this(2)] from insert have nz': "lead_coeff (f x) * lead_coeff (\i\A. f i) \ 0" by (subst IH) auto from insert.prems insert.hyps nz nz' show ?case by (auto simp: lead_monom_mult lead_coeff_mult IH) qed auto thus ?th1 ?th2 by blast+ qed lemma lead_monom_sum_le: "(\x. x \ X \ lead_monom (h x) \ ub) \ lead_monom (sum h X) \ ub" by (induction X rule: infinite_finite_induct) (auto intro!: order.trans[OF lead_monom_add]) text \ The leading monomial of a sum where the leading monomial the summands are distinct is simply the maximum of the leading monomials. \ lemma lead_monom_sum: assumes "inj_on (lead_monom \ h) X" and "finite X" and "X \ {}" and "\x. x \ X \ h x \ 0" defines "m \ Max ((lead_monom \ h) ` X)" shows "lead_monom (\x\X. h x) = m" proof (rule antisym) show "lead_monom (sum h X) \ m" unfolding m_def using assms by (intro lead_monom_sum_le Max_ge finite_imageI) auto next from assms have "m \ (lead_monom \ h) ` X" unfolding m_def by (intro Max_in finite_imageI) auto then obtain x where x: "x \ X" "m = lead_monom (h x)" by auto have "coeff (\x\X. h x) m = (\x\X. coeff (h x) m)" by simp also have "\ = (\x\{x}. coeff (h x) m)" proof (intro sum.mono_neutral_right ballI) fix y assume y: "y \ X - {x}" hence "(lead_monom \ h) y \ m" using assms unfolding m_def by (intro Max_ge finite_imageI) auto moreover have "(lead_monom \ h) y \ (lead_monom \ h) x" using \x \ X\ y inj_onD[OF assms(1), of x y] by auto ultimately have "lead_monom (h y) < m" using x by auto thus "coeff (h y) m = 0" by simp qed (insert x assms, auto) also have "\ = coeff (h x) m" by simp also have "\ = lead_coeff (h x)" using x by (simp add: lead_coeff_def) also have "\ \ 0" using assms x by auto finally show "lead_monom (sum h X) \ m" by (intro lead_monom_geI) qed lemma lead_coeff_eq_0_iff [simp]: "lead_coeff p = 0 \ p = 0" by (cases "p = 0") auto lemma fixes f :: "_ \ 'a :: semidom mpoly" assumes "\i. i \ A \ f i \ 0" shows lead_monom_prod' [simp]: "lead_monom (\i\A. f i) = (\i\A. lead_monom (f i))" (is ?th1) and lead_coeff_prod' [simp]: "lead_coeff (\i\A. f i) = (\i\A. lead_coeff (f i))" (is ?th2) proof - from assms have "(\i\A. lead_coeff (f i)) \ 0" by (cases "finite A") auto thus ?th1 ?th2 by (simp_all add: lead_monom_prod lead_coeff_prod) qed lemma fixes p :: "'a :: comm_semiring_1 mpoly" assumes "lead_coeff p ^ n \ 0" shows lead_monom_power: "lead_monom (p ^ n) = of_nat n * lead_monom p" and lead_coeff_power: "lead_coeff (p ^ n) = lead_coeff p ^ n" using assms lead_monom_prod[of "\_. p" "{.._. p" "{.. 0" shows lead_monom_power' [simp]: "lead_monom (p ^ n) = of_nat n * lead_monom p" and lead_coeff_power' [simp]: "lead_coeff (p ^ n) = lead_coeff p ^ n" using assms lead_monom_prod'[of "{.._. p"] lead_coeff_prod'[of "{.._. p"] by simp_all subsection \Turning a set of variables into a monomial\ text \ Given a finite set $\{X_1,\ldots,X_n\}$ of variables, the following is the monomial $X_1\ldots X_n$: \ lift_definition monom_of_set :: "nat set \ (nat \\<^sub>0 nat)" is "\X x. if finite X \ x \ X then 1 else 0" by auto lemma lookup_monom_of_set: "Poly_Mapping.lookup (monom_of_set X) i = (if finite X \ i \ X then 1 else 0)" by transfer auto lemma lookup_monom_of_set_1 [simp]: "finite X \ i \ X \ Poly_Mapping.lookup (monom_of_set X) i = 1" and lookup_monom_of_set_0 [simp]: "i \ X \ Poly_Mapping.lookup (monom_of_set X) i = 0" by (simp_all add: lookup_monom_of_set) lemma keys_monom_of_set: "keys (monom_of_set X) = (if finite X then X else {})" by transfer auto lemma keys_monom_of_set_finite [simp]: "finite X \ keys (monom_of_set X) = X" by (simp add: keys_monom_of_set) lemma monom_of_set_eq_iff [simp]: "finite X \ finite Y \ monom_of_set X = monom_of_set Y \ X = Y" by transfer (auto simp: fun_eq_iff) lemma monom_of_set_empty [simp]: "monom_of_set {} = 0" by transfer auto lemma monom_of_set_eq_zero_iff [simp]: "monom_of_set X = 0 \ infinite X \ X = {}" by transfer (auto simp: fun_eq_iff) lemma zero_eq_monom_of_set_iff [simp]: "0 = monom_of_set X \ infinite X \ X = {}" by transfer (auto simp: fun_eq_iff) subsection \Permuting the variables of a polynomial\ text \ Next, we define the operation of permuting the variables of a monomial and polynomial. \ lift_definition permutep :: "('a \ 'a) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b :: zero)" is "\f p. if bij f then p \ f else p" proof - fix f :: "'a \ 'a" and g :: "'a \ 'b" assume *: "finite {x. g x \ 0}" show "finite {x. (if bij f then g \ f else g) x \ 0}" proof (cases "bij f") case True with * have "finite (f -` {x. g x \ 0})" by (intro finite_vimageI) (auto dest: bij_is_inj) with True show ?thesis by auto qed (use * in auto) qed lift_definition mpoly_map_vars :: "(nat \ nat) \ 'a :: zero mpoly \ 'a mpoly" is "\f p. permutep (permutep f) p" . lemma keys_permutep: "bij f \ keys (permutep f m) = f -` keys m" by transfer auto lemma permutep_id'' [simp]: "permutep id = id" by transfer' (auto simp: fun_eq_iff) lemma permutep_id''' [simp]: "permutep (\x. x) = id" by transfer' (auto simp: fun_eq_iff) lemma permutep_0 [simp]: "permutep f 0 = 0" by transfer auto lemma permutep_single: "bij f \ permutep f (Poly_Mapping.single a b) = Poly_Mapping.single (inv_into UNIV f a) b" by transfer (auto simp: fun_eq_iff when_def inv_f_f surj_f_inv_f bij_is_inj bij_is_surj) lemma mpoly_map_vars_id [simp]: "mpoly_map_vars id = id" by transfer auto lemma mpoly_map_vars_id' [simp]: "mpoly_map_vars (\x. x) = id" by transfer auto lemma lookup_permutep: "Poly_Mapping.lookup (permutep f m) x = (if bij f then Poly_Mapping.lookup m (f x) else Poly_Mapping.lookup m x)" by transfer auto lemma inj_permutep [intro]: "inj (permutep (f :: 'a \ 'a) :: _ \ 'a \\<^sub>0 'b :: zero)" unfolding inj_def proof (transfer, safe) fix f :: "'a \ 'a" and x y :: "'a \ 'b" assume eq: "(if bij f then x \ f else x) = (if bij f then y \ f else y)" show "x = y" proof (cases "bij f") case True show ?thesis proof fix t :: 'a from \bij f\ obtain s where "t = f s" by (auto dest!: bij_is_surj) with eq and True show "x t = y t" by (auto simp: fun_eq_iff) qed qed (use eq in auto) qed lemma surj_permutep [intro]: "surj (permutep (f :: 'a \ 'a) :: _ \ 'a \\<^sub>0 'b :: zero)" unfolding surj_def proof (transfer, safe) fix f :: "'a \ 'a" and y :: "'a \ 'b" assume fin: "finite {t. y t \ 0}" show "\x\{f. finite {x. f x \ 0}}. y = (if bij f then x \ f else x)" proof (cases "bij f") case True with fin have "finite (the_inv f -` {t. y t \ 0})" by (intro finite_vimageI) (auto simp: bij_is_inj bij_betw_the_inv_into) moreover have "y \ the_inv f \ f = y" using True by (simp add: fun_eq_iff the_inv_f_f bij_is_inj) ultimately show ?thesis by (intro bexI[of _ "y \ the_inv f"]) (auto simp: True) qed (use fin in auto) qed lemma bij_permutep [intro]: "bij (permutep f)" using inj_permutep[of f] surj_permutep[of f] by (simp add: bij_def) lemma mpoly_map_vars_map_mpoly: "mpoly_map_vars f (map_mpoly g p) = map_mpoly g (mpoly_map_vars f p)" by (transfer', transfer') (auto simp: fun_eq_iff) lemma coeff_mpoly_map_vars: fixes f :: "nat \ nat" and p :: "'a :: zero mpoly" assumes "bij f" shows "MPoly_Type.coeff (mpoly_map_vars f p) mon = MPoly_Type.coeff p (permutep f mon)" using assms by transfer' (simp add: lookup_permutep bij_permutep) lemma permutep_monom_of_set: assumes "bij f" shows "permutep f (monom_of_set A) = monom_of_set (f -` A)" using assms by transfer (auto simp: fun_eq_iff bij_is_inj finite_vimage_iff) lemma permutep_comp: "bij f \ bij g \ permutep (f \ g) = permutep g \ permutep f" by transfer' (auto simp: fun_eq_iff bij_comp) lemma permutep_comp': "bij f \ bij g \ permutep (f \ g) mon = permutep g (permutep f mon)" by transfer (auto simp: fun_eq_iff bij_comp) lemma mpoly_map_vars_comp: "bij f \ bij g \ mpoly_map_vars f (mpoly_map_vars g p) = mpoly_map_vars (f \ g) p" by transfer (auto simp: bij_permutep permutep_comp) lemma permutep_id [simp]: "permutep id mon = mon" by transfer auto lemma permutep_id' [simp]: "permutep (\x. x) mon = mon" by transfer auto lemma inv_permutep [simp]: fixes f :: "'a \ 'a" assumes "bij f" shows "inv_into UNIV (permutep f) = permutep (inv_into UNIV f)" proof fix m :: "'a \\<^sub>0 'b" show "inv_into UNIV (permutep f) m = permutep (inv_into UNIV f) m" using permutep_comp'[of "inv_into UNIV f" f m] assms inj_iff[of f] by (intro inv_f_eq) (auto simp: bij_imp_bij_inv bij_is_inj) qed lemma mpoly_map_vars_monom: "bij f \ mpoly_map_vars f (monom m c) = monom (permutep (inv_into UNIV f) m) c" by transfer' (simp add: permutep_single bij_permutep) lemma vars_mpoly_map_vars: fixes f :: "nat \ nat" and p :: "'a :: zero mpoly" assumes "bij f" shows "vars (mpoly_map_vars f p) = f ` vars p" using assms unfolding vars_def proof transfer' fix f :: "nat \ nat" and p :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" assume f: "bij f" have eq: "f (inv_into UNIV f x) = x" for x using f by (subst surj_f_inv_f[of f]) (auto simp: bij_is_surj) show "\ (keys ` keys (permutep (permutep f) p)) = f ` \ (keys ` keys p)" proof safe fix m x assume mx: "m \ keys (permutep (permutep f) p)" "x \ keys m" from mx have "permutep f m \ keys p" by (auto simp: keys_permutep bij_permutep f) with mx have "f (inv_into UNIV f x) \ f ` (\m\keys p. keys m)" by (intro imageI) (auto intro!: bexI[of _ "permutep f m"] simp: keys_permutep f eq) with eq show "x \ f ` (\m\keys p. keys m)" by simp next fix m x assume mx: "m \ keys p" "x \ keys m" from mx have "permutep id m \ keys p" by simp also have "id = inv_into UNIV f \ f" using f by (intro ext) (auto simp: bij_is_inj inv_f_f) also have "permutep \ m = permutep f (permutep (inv_into UNIV f) m)" by (simp add: permutep_comp f bij_imp_bij_inv) finally have **: "permutep f (permutep (inv_into UNIV f) m) \ keys p" . moreover from f mx have "f x \ keys (permutep (inv_into UNIV f) m)" by (auto simp: keys_permutep bij_imp_bij_inv inv_f_f bij_is_inj) ultimately show "f x \ \ (keys ` keys (permutep (permutep f) p))" using f by (auto simp: keys_permutep bij_permutep) qed qed lemma permutep_eq_monom_of_set_iff [simp]: assumes "bij f" shows "permutep f mon = monom_of_set A \ mon = monom_of_set (f ` A)" proof assume eq: "permutep f mon = monom_of_set A" have "permutep (inv_into UNIV f) (permutep f mon) = monom_of_set (inv_into UNIV f -` A)" using assms by (simp add: eq bij_imp_bij_inv assms permutep_monom_of_set) also have "inv_into UNIV f -` A = f ` A" using assms by (force simp: bij_is_surj image_iff inv_f_f bij_is_inj surj_f_inv_f) also have "permutep (inv_into UNIV f) (permutep f mon) = permutep (f \ inv_into UNIV f) mon" using assms by (simp add: permutep_comp bij_imp_bij_inv) also have "f \ inv_into UNIV f = id" by (subst surj_iff [symmetric]) (insert assms, auto simp: bij_is_surj) finally show "mon = monom_of_set (f ` A)" by simp qed (insert assms, auto simp: permutep_monom_of_set inj_vimage_image_eq bij_is_inj) lemma permutep_monom_of_set_permutes [simp]: assumes "\ permutes A" shows "permutep \ (monom_of_set A) = monom_of_set A" using assms by transfer (auto simp: if_splits fun_eq_iff permutes_in_image) lemma mpoly_map_vars_0 [simp]: "mpoly_map_vars f 0 = 0" by (transfer, transfer') (simp add: o_def) lemma permutep_eq_0_iff [simp]: "permutep f m = 0 \ m = 0" proof transfer fix f :: "'a \ 'a" and m :: "'a \ 'b" assume "finite {x. m x \ 0}" show "((if bij f then m \ f else m) = (\k. 0)) = (m = (\k. 0))" proof (cases "bij f") case True hence "(\x. m (f x) = 0) \ (\x. m x = 0)" using bij_iff[of f] by metis with True show ?thesis by (auto simp: fun_eq_iff) qed (auto simp: fun_eq_iff) qed lemma mpoly_map_vars_1 [simp]: "mpoly_map_vars f 1 = 1" by (transfer, transfer') (auto simp: o_def fun_eq_iff when_def) lemma permutep_Const\<^sub>0 [simp]: "(\x. f x = 0 \ x = 0) \ permutep f (Const\<^sub>0 c) = Const\<^sub>0 c" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_add [simp]: "permutep f (m1 + m2) = permutep f m1 + permutep f m2" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_diff [simp]: "permutep f (m1 - m2) = permutep f m1 - permutep f m2" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_uminus [simp]: "permutep f (-m) = -permutep f m" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_mult [simp]: "(\x y. f (x + y) = f x + f y) \ permutep f (m1 * m2) = permutep f m1 * permutep f m2" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff prod_fun_compose_bij) lemma mpoly_map_vars_Const [simp]: "mpoly_map_vars f (Const c) = Const c" by transfer (auto simp: o_def fun_eq_iff when_def) lemma mpoly_map_vars_add [simp]: "mpoly_map_vars f (p + q) = mpoly_map_vars f p + mpoly_map_vars f q" by transfer simp lemma mpoly_map_vars_diff [simp]: "mpoly_map_vars f (p - q) = mpoly_map_vars f p - mpoly_map_vars f q" by transfer simp lemma mpoly_map_vars_uminus [simp]: "mpoly_map_vars f (-p) = -mpoly_map_vars f p" by transfer simp lemma permutep_smult: "permutep (permutep f) (Poly_Mapping.map ((*) c) p) = Poly_Mapping.map ((*) c) (permutep (permutep f) p)" by transfer' (auto split: if_splits simp: fun_eq_iff) lemma mpoly_map_vars_smult [simp]: "mpoly_map_vars f (smult c p) = smult c (mpoly_map_vars f p)" by transfer (simp add: permutep_smult) lemma mpoly_map_vars_mult [simp]: "mpoly_map_vars f (p * q) = mpoly_map_vars f p * mpoly_map_vars f q" by transfer simp lemma mpoly_map_vars_sum [simp]: "mpoly_map_vars f (sum g A) = (\x\A. mpoly_map_vars f (g x))" by (induction A rule: infinite_finite_induct) auto lemma mpoly_map_vars_prod [simp]: "mpoly_map_vars f (prod g A) = (\x\A. mpoly_map_vars f (g x))" by (induction A rule: infinite_finite_induct) auto lemma mpoly_map_vars_eq_0_iff [simp]: "mpoly_map_vars f p = 0 \ p = 0" by transfer auto lemma permutep_eq_iff [simp]: "permutep f p = permutep f q \ p = q" by transfer (auto simp: comp_bij_eq_iff) lemma mpoly_map_vars_Sum_any [simp]: "mpoly_map_vars f (Sum_any g) = Sum_any (\x. mpoly_map_vars f (g x))" by (simp add: Sum_any.expand_set) lemma mpoly_map_vars_power [simp]: "mpoly_map_vars f (p ^ n) = mpoly_map_vars f p ^ n" by (induction n) auto lemma mpoly_map_vars_monom_single [simp]: assumes "bij f" shows "mpoly_map_vars f (monom (Poly_Mapping.single i n) c) = monom (Poly_Mapping.single (f i) n) c" using assms by (simp add: mpoly_map_vars_monom permutep_single bij_imp_bij_inv inv_inv_eq) lemma insertion_mpoly_map_vars: assumes "bij f" shows "insertion g (mpoly_map_vars f p) = insertion (g \ f) p" proof - have "insertion g (mpoly_map_vars f p) = (\m. coeff p (permutep f m) * (\i. g i ^ lookup m i))" using assms by (simp add: insertion_altdef coeff_mpoly_map_vars) also have "\ = Sum_any (\m. coeff p (permutep f m) * Prod_any (\i. g (f i) ^ lookup m (f i)))" by (intro Sum_any.cong arg_cong[where ?f = "\y. x * y" for x] Prod_any.reindex_cong[OF assms]) (auto simp: o_def) also have "\ = Sum_any (\m. coeff p m * (\i. g (f i) ^ lookup m i))" by (intro Sum_any.reindex_cong [OF bij_permutep[of f], symmetric]) (auto simp: o_def lookup_permutep assms) also have "\ = insertion (g \ f) p" by (simp add: insertion_altdef) finally show ?thesis . qed lemma permutep_cong: assumes "f permutes (-keys p)" "g permutes (-keys p)" "p = q" shows "permutep f p = permutep g q" proof (intro poly_mapping_eqI) fix k :: 'a show "lookup (permutep f p) k = lookup (permutep g q) k" proof (cases "k \ keys p") case False with assms have "f k \ keys p" "g k \ keys p" using permutes_in_image[of _ "-keys p" k] by auto thus ?thesis using assms by (auto simp: lookup_permutep permutes_bij in_keys_iff) qed (insert assms, auto simp: lookup_permutep permutes_bij permutes_not_in) qed lemma mpoly_map_vars_cong: assumes "f permutes (-vars p)" "g permutes (-vars q)" "p = q" shows "mpoly_map_vars f p = mpoly_map_vars g (q :: 'a :: zero mpoly)" proof (intro mpoly_eqI) fix mon :: "nat \\<^sub>0 nat" show "coeff (mpoly_map_vars f p) mon = coeff (mpoly_map_vars g q) mon" proof (cases "keys mon \ vars p") case True with assms have "permutep f mon = permutep g mon" by (intro permutep_cong assms(1,2)[THEN permutes_subset]) auto thus ?thesis using assms by (simp add: coeff_mpoly_map_vars permutes_bij) next case False hence "\(keys mon \ f ` vars q)" "\(keys mon \ g ` vars q)" using assms by (auto simp: subset_iff permutes_not_in) thus ?thesis using assms by (subst (1 2) coeff_notin_vars) (auto simp: coeff_notin_vars vars_mpoly_map_vars permutes_bij) qed qed subsection \Symmetric polynomials\ text \ A polynomial is symmetric on a set of variables if it is invariant under any permutation of that set. \ definition symmetric_mpoly :: "nat set \ 'a :: zero mpoly \ bool" where "symmetric_mpoly A p = (\\. \ permutes A \ mpoly_map_vars \ p = p)" lemma symmetric_mpoly_empty [simp, intro]: "symmetric_mpoly {} p" by (simp add: symmetric_mpoly_def) text \ A polynomial is trivially symmetric on any set of variables that do not occur in it. \ lemma symmetric_mpoly_orthogonal: assumes "vars p \ A = {}" shows "symmetric_mpoly A p" unfolding symmetric_mpoly_def proof safe fix \ assume \: "\ permutes A" with assms have "\ x = x" if "x \ vars p" for x using that permutes_not_in[of \ A x] by auto from assms have "mpoly_map_vars \ p = mpoly_map_vars id p" by (intro mpoly_map_vars_cong permutes_subset[OF \] permutes_id) auto also have "\ = p" by simp finally show "mpoly_map_vars \ p = p" . qed lemma symmetric_mpoly_monom [intro]: assumes "keys m \ A = {}" shows "symmetric_mpoly A (monom m c)" using assms vars_monom_subset[of m c] by (intro symmetric_mpoly_orthogonal) auto lemma symmetric_mpoly_subset: assumes "symmetric_mpoly A p" "B \ A" shows "symmetric_mpoly B p" unfolding symmetric_mpoly_def proof safe fix \ assume "\ permutes B" with assms have "\ permutes A" using permutes_subset by blast with assms show "mpoly_map_vars \ p = p" by (auto simp: symmetric_mpoly_def) qed text \ If a polynomial is symmetric over some set of variables, that set must either be a subset of the variables occurring in the polynomial or disjoint from it. \ lemma symmetric_mpoly_imp_orthogonal_or_subset: assumes "symmetric_mpoly A p" shows "vars p \ A = {} \ A \ vars p" proof (rule ccontr) assume "\(vars p \ A = {} \ A \ vars p)" then obtain x y where xy: "x \ vars p \ A" "y \ A - vars p" by auto define \ where "\ = Fun.swap x y id" from xy have \: "\ permutes A" unfolding \_def by (intro permutes_swap_id) auto from xy have "y \ \ ` vars p" by (auto simp: \_def Fun.swap_def) also from \ have "\ ` vars p = vars (mpoly_map_vars \ p)" by (auto simp: vars_mpoly_map_vars permutes_bij) also have "mpoly_map_vars \ p = p" using assms \ by (simp add: symmetric_mpoly_def) finally show False using xy by auto qed text \ Symmetric polynomials are closed under ring operations. \ lemma symmetric_mpoly_add [intro]: "symmetric_mpoly A p \ symmetric_mpoly A q \ symmetric_mpoly A (p + q)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_diff [intro]: "symmetric_mpoly A p \ symmetric_mpoly A q \ symmetric_mpoly A (p - q)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_uminus [intro]: "symmetric_mpoly A p \ symmetric_mpoly A (-p)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_uminus_iff [simp]: "symmetric_mpoly A (-p) \ symmetric_mpoly A p" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_smult [intro]: "symmetric_mpoly A p \ symmetric_mpoly A (smult c p)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_mult [intro]: "symmetric_mpoly A p \ symmetric_mpoly A q \ symmetric_mpoly A (p * q)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_0 [simp, intro]: "symmetric_mpoly A 0" and symmetric_mpoly_1 [simp, intro]: "symmetric_mpoly A 1" and symmetric_mpoly_Const [simp, intro]: "symmetric_mpoly A (Const c)" by (simp_all add: symmetric_mpoly_def) lemma symmetric_mpoly_power [intro]: "symmetric_mpoly A p \ symmetric_mpoly A (p ^ n)" by (induction n) (auto intro!: symmetric_mpoly_mult) lemma symmetric_mpoly_sum [intro]: "(\i. i \ B \ symmetric_mpoly A (f i)) \ symmetric_mpoly A (sum f B)" by (induction B rule: infinite_finite_induct) (auto intro!: symmetric_mpoly_add) lemma symmetric_mpoly_prod [intro]: "(\i. i \ B \ symmetric_mpoly A (f i)) \ symmetric_mpoly A (prod f B)" by (induction B rule: infinite_finite_induct) (auto intro!: symmetric_mpoly_mult) text \ An symmetric sum or product over polynomials yields a symmetric polynomial: \ lemma symmetric_mpoly_symmetric_sum: assumes "g permutes X" assumes "\x \. x \ X \ \ permutes A \ mpoly_map_vars \ (f x) = f (g x)" shows "symmetric_mpoly A (\x\X. f x)" unfolding symmetric_mpoly_def proof safe fix \ assume \: "\ permutes A" have "mpoly_map_vars \ (sum f X) = (\x\X. mpoly_map_vars \ (f x))" by simp also have "\ = (\x\X. f (g x))" by (intro sum.cong assms \ refl) also have "\ = (\x\g`X. f x)" using assms by (subst sum.reindex) (auto simp: permutes_inj_on) also have "g ` X = X" using assms by (simp add: permutes_image) finally show "mpoly_map_vars \ (sum f X) = sum f X" . qed lemma symmetric_mpoly_symmetric_prod: assumes "g permutes X" assumes "\x \. x \ X \ \ permutes A \ mpoly_map_vars \ (f x) = f (g x)" shows "symmetric_mpoly A (\x\X. f x)" unfolding symmetric_mpoly_def proof safe fix \ assume \: "\ permutes A" have "mpoly_map_vars \ (prod f X) = (\x\X. mpoly_map_vars \ (f x))" by simp also have "\ = (\x\X. f (g x))" by (intro prod.cong assms \ refl) also have "\ = (\x\g`X. f x)" using assms by (subst prod.reindex) (auto simp: permutes_inj_on) also have "g ` X = X" using assms by (simp add: permutes_image) finally show "mpoly_map_vars \ (prod f X) = prod f X" . qed text \ If $p$ is a polynomial that is symmetric on some subset of variables $A$, then for the leading monomial of $p$, the exponents of these variables are decreasing w.\,r.\,t.\ the variable ordering. \ theorem lookup_lead_monom_decreasing: assumes "symmetric_mpoly A p" defines "m \ lead_monom p" assumes "i \ A" "j \ A" "i \ j" shows "lookup m i \ lookup m j" proof (cases "p = 0") case [simp]: False show ?thesis proof (intro leI notI) assume less: "lookup m i < lookup m j" define \ where "\ = Fun.swap i j id" from assms have \: "\ permutes A" unfolding \_def by (intro permutes_swap_id) auto have [simp]: "\ \ \ = id" "\ i = j" "\ j = i" "\k. k \ i \ k \ j \ \ k = k" by (auto simp: \_def Fun.swap_def fun_eq_iff) have "0 \ lead_coeff p" by simp also have "lead_coeff p = MPoly_Type.coeff (mpoly_map_vars \ p) (permutep \ m)" using \ by (simp add: lead_coeff_def m_def coeff_mpoly_map_vars permutes_bij permutep_comp' [symmetric]) also have "mpoly_map_vars \ p = p" using \ assms by (simp add: symmetric_mpoly_def) finally have "permutep \ m \ m" by (auto simp: m_def) moreover have "lookup m i < lookup (permutep \ m) i" and "(\k m) k)" using assms \ less by (auto simp: lookup_permutep permutes_bij) hence "m < permutep \ m" by (auto simp: less_poly_mapping_def less_fun_def) ultimately show False by simp qed qed (auto simp: m_def) subsection \The elementary symmetric polynomials\ text \ The $k$-th elementary symmetric polynomial for a finite set of variables $A$, with $k$ ranging between 1 and $|A|$, is the sum of the product of all subsets of $A$ with cardinality $k$: \ lift_definition sym_mpoly_aux :: "nat set \ nat \ (nat \\<^sub>0 nat) \\<^sub>0 'a :: {zero_neq_one}" is "\X k mon. if finite X \ (\Y. Y \ X \ card Y = k \ mon = monom_of_set Y) then 1 else 0" proof - fix k :: nat and X :: "nat set" show "finite {x. (if finite X \ (\Y\X. card Y = k \ x = monom_of_set Y) then 1 else 0) \ (0::'a)}" (is "finite ?A") proof (cases "finite X") case True have "?A \ monom_of_set ` Pow X" by auto moreover from True have "finite (monom_of_set ` Pow X)" by simp ultimately show ?thesis by (rule finite_subset) qed auto qed lemma lookup_sym_mpoly_aux: "Poly_Mapping.lookup (sym_mpoly_aux X k) mon = (if finite X \ (\Y. Y \ X \ card Y = k \ mon = monom_of_set Y) then 1 else 0)" by transfer' simp lemma lookup_sym_mpoly_aux_monom_of_set [simp]: assumes "finite X" "Y \ X" "card Y = k" shows "Poly_Mapping.lookup (sym_mpoly_aux X k) (monom_of_set Y) = 1" using assms by (auto simp: lookup_sym_mpoly_aux) lemma keys_sym_mpoly_aux: "m \ keys (sym_mpoly_aux A k) \ keys m \ A" by transfer' (auto split: if_splits simp: keys_monom_of_set) lift_definition sym_mpoly :: "nat set \ nat \ 'a :: {zero_neq_one} mpoly" is "sym_mpoly_aux" . lemma vars_sym_mpoly_subset: "vars (sym_mpoly A k) \ A" using keys_sym_mpoly_aux by (auto simp: vars_def sym_mpoly.rep_eq) lemma coeff_sym_mpoly: "MPoly_Type.coeff (sym_mpoly X k) mon = (if finite X \ (\Y. Y \ X \ card Y = k \ mon = monom_of_set Y) then 1 else 0)" by transfer' (simp add: lookup_sym_mpoly_aux) lemma sym_mpoly_infinite: "\finite A \ sym_mpoly A k = 0" by (transfer, transfer) auto lemma sym_mpoly_altdef: "sym_mpoly A k = (\X | X \ A \ card X = k. monom (monom_of_set X) 1)" proof (cases "finite A") case False hence *: "infinite {X. X \ A \ infinite X}" by (rule infinite_infinite_subsets) have "infinite {X. X \ A \ card X = 0}" by (rule infinite_super[OF _ *]) auto moreover have **: "infinite {X. X \ A \ finite X \ card X = k}" if "k \ 0" using that infinite_card_subsets[of A k] False by auto have "infinite {X. X \ A \ card X = k}" if "k \ 0" by (rule infinite_super[OF _ **[OF that]]) auto ultimately show ?thesis using False by (cases "k = 0") (simp_all add: sym_mpoly_infinite) next case True show ?thesis proof (intro mpoly_eqI, goal_cases) case (1 m) show ?case proof (cases "\X. X \ A \ card X = k \ m = monom_of_set X") case False thus ?thesis by (auto simp: coeff_sym_mpoly coeff_sum coeff_monom) next case True then obtain X where X: "X \ A" "card X = k" "m = monom_of_set X" by blast have "coeff (\X | X \ A \ card X = k. monom (monom_of_set X) 1) m = (\X\{X}. 1)" unfolding coeff_sum proof (intro sum.mono_neutral_cong_right ballI) fix Y assume Y: "Y \ {X. X \ A \ card X = k} - {X}" hence "X = Y" if "monom_of_set X = monom_of_set Y" using that finite_subset[OF X(1)] finite_subset[of Y A] \finite A\ by auto thus "coeff (monom (monom_of_set Y) 1) m = 0" using X Y by (auto simp: coeff_monom when_def ) qed (insert X \finite A\, auto simp: coeff_monom) thus ?thesis using \finite A\ by (auto simp: coeff_sym_mpoly coeff_sum coeff_monom) qed qed qed lemma coeff_sym_mpoly_monom_of_set [simp]: assumes "finite X" "Y \ X" "card Y = k" shows "MPoly_Type.coeff (sym_mpoly X k) (monom_of_set Y) = 1" using assms by (auto simp: coeff_sym_mpoly) lemma coeff_sym_mpoly_0: "coeff (sym_mpoly X k) 0 = (if finite X \ k = 0 then 1 else 0)" proof - consider "finite X" "k = 0" | "finite X" "k \ 0" | "infinite X" by blast thus ?thesis proof cases assume "finite X" "k = 0" hence "coeff (sym_mpoly X k) (monom_of_set {}) = 1" by (subst coeff_sym_mpoly_monom_of_set) auto thus ?thesis unfolding monom_of_set_empty using \finite X\ \k = 0\ by simp next assume "finite X" "k \ 0" hence "\(\Y. finite Y \ Y \ X \ card Y = k \ monom_of_set Y = 0)" by auto thus ?thesis using \k \ 0\ by (auto simp: coeff_sym_mpoly) next assume "infinite X" thus ?thesis by (simp add: coeff_sym_mpoly) qed qed lemma symmetric_sym_mpoly [intro]: assumes "A \ B" shows "symmetric_mpoly A (sym_mpoly B k :: 'a :: zero_neq_one mpoly)" unfolding symmetric_mpoly_def proof (safe intro!: mpoly_eqI) fix \ and mon :: "nat \\<^sub>0 nat" assume \: "\ permutes A" from \ have \': "\ permutes B" by (rule permutes_subset) fact from \ have "MPoly_Type.coeff (mpoly_map_vars \ (sym_mpoly B k :: 'a mpoly)) mon = MPoly_Type.coeff (sym_mpoly B k :: 'a mpoly) (permutep \ mon)" by (simp add: coeff_mpoly_map_vars permutes_bij) also have "\ = 1 \ MPoly_Type.coeff (sym_mpoly B k :: 'a mpoly) mon = 1" (is "?lhs = 1 \ ?rhs = 1") proof assume "?rhs = 1" then obtain Y where "finite B" and Y: "Y \ B" "card Y = k" "mon = monom_of_set Y" by (auto simp: coeff_sym_mpoly split: if_splits) with \' have "\ -` Y \ B" "card (\ -` Y) = k" "permutep \ mon = monom_of_set (\ -` Y)" by (auto simp: permutes_in_image card_vimage_inj permutep_monom_of_set permutes_bij permutes_inj permutes_surj) thus "?lhs = 1" using \finite B\ by (auto simp: coeff_sym_mpoly) next assume "?lhs = 1" then obtain Y where "finite B" and Y: "Y \ B" "card Y = k" "permutep \ mon = monom_of_set Y" by (auto simp: coeff_sym_mpoly split: if_splits) from Y(1) have "inj_on \ Y" using inj_on_subset[of \ UNIV Y] \' by (auto simp: permutes_inj) with Y \' have "\ ` Y \ B" "card (\ ` Y) = k" "mon = monom_of_set (\ ` Y)" by (auto simp: permutes_in_image card_image permutep_monom_of_set permutes_bij permutes_inj permutes_surj) thus "?rhs = 1" using \finite B\ by (auto simp: coeff_sym_mpoly) qed hence "?lhs = ?rhs" by (auto simp: coeff_sym_mpoly split: if_splits) finally show "MPoly_Type.coeff (mpoly_map_vars \ (sym_mpoly B k :: 'a mpoly)) mon = MPoly_Type.coeff (sym_mpoly B k :: 'a mpoly) mon" . qed lemma insertion_sym_mpoly: assumes "finite X" shows "insertion f (sym_mpoly X k) = (\Y | Y \ X \ card Y = k. prod f Y)" using assms proof (transfer, transfer) fix f :: "nat \ 'a" and k :: nat and X :: "nat set" assume X: "finite X" have "insertion_fun f (\mon. if finite X \ (\Y\X. card Y = k \ mon = monom_of_set Y) then 1 else 0) = (\m. (\v. f v ^ poly_mapping.lookup m v) when (\Y\X. card Y = k \ m = monom_of_set Y))" by (auto simp add: insertion_fun_def X when_def intro!: Sum_any.cong) also have "\ = (\m | \Y\Pow X. card Y = k \ m = monom_of_set Y. (\v. f v ^ poly_mapping.lookup m v) when (\Y\X. card Y = k \ m = monom_of_set Y))" by (rule Sum_any.expand_superset) (use X in auto) also have "\ = (\m | \Y\Pow X. card Y = k \ m = monom_of_set Y. (\v. f v ^ poly_mapping.lookup m v))" by (intro sum.cong) (auto simp: when_def) also have "\ = (\Y | Y \ X \ card Y = k. (\v. f v ^ poly_mapping.lookup (monom_of_set Y) v))" by (rule sum.reindex_bij_witness[of _ monom_of_set keys]) (auto simp: finite_subset[OF _ X]) also have "\ = (\Y | Y \ X \ card Y = k. \v\Y. f v)" proof (intro sum.cong when_cong refl, goal_cases) case (1 Y) hence "finite Y" by (auto dest: finite_subset[OF _ X]) with 1 have "(\v. f v ^ poly_mapping.lookup (monom_of_set Y) v) = (\v::nat. if v \ Y then f v else 1)" by (intro Prod_any.cong) (auto simp: lookup_monom_of_set) also have "\ = (\v\Y. f v)" by (rule Prod_any.conditionalize [symmetric]) fact+ finally show ?case . qed finally show "insertion_fun f (\mon. if finite X \ (\Y\X. card Y = k \ mon = monom_of_set Y) then 1 else 0) = (\Y | Y \ X \ card Y = k. prod f Y)" . qed lemma sym_mpoly_nz [simp]: assumes "finite A" "k \ card A" shows "sym_mpoly A k \ (0 :: 'a :: zero_neq_one mpoly)" proof - from assms obtain B where B: "B \ A" "card B = k" using ex_subset_of_card by blast with assms have "coeff (sym_mpoly A k :: 'a mpoly) (monom_of_set B) = 1" by (intro coeff_sym_mpoly_monom_of_set) thus ?thesis by auto qed lemma coeff_sym_mpoly_0_or_1: "coeff (sym_mpoly A k) m \ {0, 1}" by (transfer, transfer) auto lemma lead_coeff_sym_mpoly [simp]: assumes "finite A" "k \ card A" shows "lead_coeff (sym_mpoly A k) = 1" proof - from assms have "lead_coeff (sym_mpoly A k) \ 0" by simp thus ?thesis using coeff_sym_mpoly_0_or_1[of A k "lead_monom (sym_mpoly A k)"] unfolding lead_coeff_def by blast qed lemma lead_monom_sym_mpoly: assumes "sorted xs" "distinct xs" "k \ length xs" shows "lead_monom (sym_mpoly (set xs) k :: 'a :: zero_neq_one mpoly) = monom_of_set (set (take k xs))" (is "lead_monom ?p = _") proof - let ?m = "lead_monom ?p" have sym: "symmetric_mpoly (set xs) (sym_mpoly (set xs) k)" by (intro symmetric_sym_mpoly) auto from assms have [simp]: "card (set xs) = length xs" by (subst distinct_card) auto from assms have "lead_coeff ?p = 1" by (subst lead_coeff_sym_mpoly) auto then obtain X where X: "X \ set xs" "card X = k" "?m = monom_of_set X" unfolding lead_coeff_def by (subst (asm) coeff_sym_mpoly) (auto split: if_splits) define ys where "ys = map (\x. if x \ X then 1 else 0 :: nat) xs" have [simp]: "length ys = length xs" by (simp add: ys_def) have ys_altdef: "ys = map (lookup ?m) xs" unfolding ys_def using X finite_subset[OF X(1)] by (intro map_cong) (auto simp: lookup_monom_of_set) define i where "i = Min (insert (length xs) {i. i < length xs \ ys ! i = 0})" have "i \ length xs" by (auto simp: i_def) have in_X: "xs ! j \ X" if "j < i" for j using that unfolding i_def by (auto simp: ys_def) have not_in_X: "xs ! j \ X" if "i \ j" "j < length xs" for j proof - have ne: "{i. i < length xs \ ys ! i = 0} \ {}" proof assume [simp]: "{i. i < length xs \ ys ! i = 0} = {}" from that show False by (simp add: i_def) qed hence "Min {i. i < length xs \ ys ! i = 0} \ {i. i < length xs \ ys ! i = 0}" using that by (intro Min_in) auto also have "Min {i. i < length xs \ ys ! i = 0} = i" unfolding i_def using ne by (subst Min_insert) (auto simp: min_def) finally have i: "ys ! i = 0" "i < length xs" by simp_all have "lookup ?m (xs ! j) \ lookup ?m (xs ! i)" using that assms by (intro lookup_lead_monom_decreasing[OF sym]) (auto intro!: sorted_nth_mono simp: set_conv_nth) also have "\ = 0" using i by (simp add: ys_altdef) finally show ?thesis using that X finite_subset[OF X(1)] by (auto simp: lookup_monom_of_set) qed from X have "k = card X" by simp also have "X = (\i. xs ! i) ` {i. i < length xs \ xs ! i \ X}" using X by (auto simp: set_conv_nth) also have "card \ = (\i | i < length xs \ xs ! i \ X. 1)" using assms by (subst card_image) (auto intro!: inj_on_nth) also have "\ = (\i | i < length xs. if xs ! i \ X then 1 else 0)" by (intro sum.mono_neutral_cong_left) auto also have "\ = sum_list ys" by (auto simp: sum_list_sum_nth ys_def intro!: sum.cong) also have "ys = take i ys @ drop i ys" by simp also have "sum_list \ = sum_list (take i ys) + sum_list (drop i ys)" by (subst sum_list_append) auto also have "take i ys = replicate i 1" using \i \ length xs\ in_X by (intro replicate_eqI) (auto simp: ys_def set_conv_nth) also have "sum_list \ = i" by simp also have "drop i ys = replicate (length ys - i) 0" using \i \ length xs\ not_in_X by (intro replicate_eqI) (auto simp: ys_def set_conv_nth) also have "sum_list \ = 0" by simp finally have "i = k" by simp have "X = set (filter (\x. x \ X) xs)" using X by auto also have "xs = take i xs @ drop i xs" by simp also note filter_append also have "filter (\x. x \ X) (take i xs) = take i xs" using in_X by (intro filter_True) (auto simp: set_conv_nth) also have "filter (\x. x \ X) (drop i xs) = []" using not_in_X by (intro filter_False) (auto simp: set_conv_nth) finally have "X = set (take i xs)" by simp with \i = k\ and X show ?thesis by simp qed subsection \Induction on the leading monomial\ text \ We show that the monomial ordering for a fixed set of variables is well-founded, so we can perform induction on the leading monomial of a polynomial. \ definition monom_less_on where "monom_less_on A = {(m1, m2). m1 < m2 \ keys m1 \ A \ keys m2 \ A}" lemma wf_monom_less_on: assumes "finite A" shows "wf (monom_less_on A :: ((nat \\<^sub>0 'b :: {zero, wellorder}) \ _) set)" proof (rule wf_subset) define n where "n = Suc (Max (insert 0 A))" have less_n: "k < n" if "k \ A" for k using that assms by (auto simp: n_def less_Suc_eq_le Max_ge_iff) define f :: "(nat \\<^sub>0 'b) \ 'b list" where "f = (\m. map (lookup m) [0.. inv_image (lexn {(x, y). x < y} n) f" proof safe fix m1 m2 :: "nat \\<^sub>0 'b" assume "(m1, m2) \ monom_less_on A" hence m12: "m1 < m2" "keys m1 \ A" "keys m2 \ A" by (auto simp: monom_less_on_def) then obtain k where k: "lookup m1 k < lookup m2 k" "\i(lookup m1 k = 0 \ lookup m2 k = 0)" proof (intro notI) assume "lookup m1 k = 0 \ lookup m2 k = 0" hence [simp]: "lookup m1 k = 0" "lookup m2 k = 0" by blast+ from k(1) show False by simp qed hence "k \ A" using m12 by (auto simp: in_keys_iff) hence "k < n" by (simp add: less_n) define as where "as = map (lookup m1) [0..k < n\ by (simp flip: upt_conv_Cons upt_add_eq_append') have [simp]: "length as = k" "length bs1 = n - Suc k" "length bs2 = n - Suc k" by (simp_all add: as_def bs1_def bs2_def) have "f m1 = as @ [lookup m1 k] @ bs1" unfolding f_def by (subst decomp) (simp add: as_def bs1_def) moreover have "f m2 = as @ [lookup m2 k] @ bs2" unfolding f_def using k by (subst decomp) (simp add: as_def bs2_def) ultimately show "(m1, m2) \ inv_image (lexn {(x,y). x < y} n) f" using k(1) \k < n\ unfolding lexn_conv by fastforce qed qed lemma lead_monom_induct [consumes 2, case_names less]: fixes p :: "'a :: zero mpoly" assumes fin: "finite A" and vars: "vars p \ A" assumes IH: "\p. vars p \ A \ (\p'. vars p' \ A \ lead_monom p' < lead_monom p \ P p') \ P p" shows "P p" using assms(2) proof (induct m \ "lead_monom p" arbitrary: p rule: wf_induct_rule[OF wf_monom_less_on[OF fin]]) case (1 p) show ?case proof (rule IH) fix p' :: "'a mpoly" assume *: "vars p' \ A" "lead_monom p' < lead_monom p" show "P p'" by (rule 1) (insert * "1.prems" keys_lead_monom_subset, auto simp: monom_less_on_def) qed (insert 1, auto) qed lemma lead_monom_induct' [case_names less]: fixes p :: "'a :: zero mpoly" assumes IH: "\p. (\p'. vars p' \ vars p \ lead_monom p' < lead_monom p \ P p') \ P p" shows "P p" proof - have "finite (vars p)" "vars p \ vars p" by (auto simp: vars_finite) thus ?thesis by (induction rule: lead_monom_induct) (use IH in blast) qed subsection \The fundamental theorem of symmetric polynomials\ lemma lead_coeff_sym_mpoly_powerprod: assumes "finite A" "\x. x \ X \ f x \ {1..card A}" shows "lead_coeff (\x\X. sym_mpoly A (f (x::'a)) ^ g x) = 1" proof - have eq: "lead_coeff (sym_mpoly A (f x) ^ g x :: 'b mpoly) = 1" if "x \ X" for x using that assms by (subst lead_coeff_power) (auto simp: lead_coeff_sym_mpoly assms) hence "(\x\X. lead_coeff (sym_mpoly A (f x) ^ g x :: 'b mpoly)) = (\x\X. 1)" by (intro prod.cong eq refl) also have "\ = 1" by simp finally have eq': "(\x\X. lead_coeff (sym_mpoly A (f x) ^ g x :: 'b mpoly)) = 1" . show ?thesis by (subst lead_coeff_prod) (auto simp: eq eq') qed context fixes A :: "nat set" and xs n f and decr :: "'a :: comm_ring_1 mpoly \ bool" defines "xs \ sorted_list_of_set A" defines "n \ card A" defines "f \ (\i. if i < n then xs ! i else 0)" defines "decr \ (\p. \i\A. \j\A. i \ j \ lookup (lead_monom p) i \ lookup (lead_monom p) j)" begin text \ The computation of the witness for the fundamental theorem works like this: Given some polynomial $p$ (that is assumed to be symmetric in the variables in $A$), we inspect its leading monomial, which is of the form $c X_1^{i_1}\ldots X_n{i_n}$ where the $A = \{X_1,\ldots, X_n\}$, $c$ contains only variables not in $A$, and the sequence $i_j$ is decreasing. The latter holds because $p$ is symmetric. Now, we form the polynomial $q := c e_1^{i_1 - i_2} e_2^{i_2 - i_3} \ldots e_n^{i_n}$, which has the same leading term as $p$. Then $p - q$ has a smaller leading monomial, so by induction, we can assume it to be of the required form and obtain a witness for $p - q$. Now, we only need to add $c Y_1^{i_1 - i_2} \ldots Y_n^{i_n}$ to that witness and we obtain a witness for $p$. \ definition fund_sym_step_coeff :: "'a mpoly \ 'a mpoly" where "fund_sym_step_coeff p = monom (restrictpm (-A) (lead_monom p)) (lead_coeff p)" definition fund_sym_step_monom :: "'a mpoly \ (nat \\<^sub>0 nat)" where "fund_sym_step_monom p = ( let g = (\i. if i < n then lookup (lead_monom p) (f i) else 0) in (\i 'a mpoly" where "fund_sym_step_poly p = ( let g = (\i. if i < n then lookup (lead_monom p) (f i) else 0) in fund_sym_step_coeff p * (\i The following function computes the witness, with the convention that it returns a constant polynomial if the input was not symmetric: \ function (domintros) fund_sym_poly_wit :: "'a :: comm_ring_1 mpoly \ 'a mpoly mpoly" where "fund_sym_poly_wit p = (if \symmetric_mpoly A p \ lead_monom p = 0 \ vars p \ A = {} then Const p else fund_sym_poly_wit (p - fund_sym_step_poly p) + monom (fund_sym_step_monom p) (fund_sym_step_coeff p))" by auto lemma coeff_fund_sym_step_coeff: "coeff (fund_sym_step_coeff p) m \ {lead_coeff p, 0}" by (auto simp: fund_sym_step_coeff_def coeff_monom when_def) lemma vars_fund_sym_step_coeff: "vars (fund_sym_step_coeff p) \ vars p - A" unfolding fund_sym_step_coeff_def using keys_lead_monom_subset[of p] by (intro order.trans[OF vars_monom_subset]) auto lemma keys_fund_sym_step_monom: "keys (fund_sym_step_monom p) \ {1..n}" unfolding fund_sym_step_monom_def Let_def by (intro order.trans[OF keys_sum] UN_least, subst keys_single) auto lemma coeff_fund_sym_step_poly: assumes C: "\m. coeff p m \ C" and "ring_closed C" shows "coeff (fund_sym_step_poly p) m \ C" proof - interpret ring_closed C by fact have *: "\m. coeff (p ^ x) m \ C" if "\m. coeff p m \ C" for p x using that by (induction x) (auto simp: coeff_mpoly_times mpoly_coeff_1 intro!: prod_fun_closed) have **: "\m. coeff (prod f X) m \ C" if "\i m. i \ X \ coeff (f i) m \ C" for X and f :: "nat \ _" using that by (induction X rule: infinite_finite_induct) (auto simp: coeff_mpoly_times mpoly_coeff_1 intro!: prod_fun_closed) show ?thesis using C unfolding fund_sym_step_poly_def Let_def fund_sym_step_coeff_def coeff_mpoly_times by (intro prod_fun_closed) (auto simp: coeff_monom when_def lead_coeff_def coeff_sym_mpoly intro!: * **) qed text \ We now show various relevant properties of the subtracted polynomial: \<^enum> Its leading term is the same as that of the input polynomial. \<^enum> It contains now new variables. \<^enum> It is symmetric in the variables in \A\. \ lemma fund_sym_step_poly: shows "finite A \ p \ 0 \ decr p \ lead_monom (fund_sym_step_poly p) = lead_monom p" and "finite A \ p \ 0 \ decr p \ lead_coeff (fund_sym_step_poly p) = lead_coeff p" and "finite A \ p \ 0 \ decr p \ fund_sym_step_poly p = fund_sym_step_coeff p * (\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" and "vars (fund_sym_step_poly p) \ vars p \ A" and "symmetric_mpoly A (fund_sym_step_poly p)" proof - define g where "g = (\i. if i < n then lookup (lead_monom p) (f i) else 0)" define q where "q = (\i vars p \ A" using keys_lead_monom_subset[of p] vars_monom_subset[of "restrictpm (-A) (lead_monom p)" "lead_coeff p"] unfolding c_def q_def by (intro order.trans[OF vars_mult] order.trans[OF vars_prod] order.trans[OF vars_power] Un_least UN_least order.trans[OF vars_sym_mpoly_subset]) auto thus "vars (fund_sym_step_poly p) \ vars p \ A" by simp have "symmetric_mpoly A (c * q)" unfolding c_def q_def by (intro symmetric_mpoly_mult symmetric_mpoly_monom symmetric_mpoly_prod symmetric_mpoly_power symmetric_sym_mpoly) auto thus "symmetric_mpoly A (fund_sym_step_poly p)" by simp assume finite: "finite A" and [simp]: "p \ 0" and "decr p" have "set xs = A" "distinct xs" and [simp]: "length xs = n" using finite by (auto simp: xs_def n_def) have [simp]: "lead_coeff c = lead_coeff p" "lead_monom c = restrictpm (- A) (lead_monom p)" by (simp_all add: c_def lead_monom_monom) hence f_range [simp]: "f i \ A" if "i < n" for i using that \set xs = A\ by (auto simp: f_def set_conv_nth) have "sorted xs" by (simp add: xs_def) hence f_mono: "f i \ f j" if "i \ j" "j < n" for i j using that by (auto simp: f_def n_def intro: sorted_nth_mono) hence g_mono: "g i \ g j" if "i \ j" for i j unfolding g_def using that using \decr p\ by (auto simp: decr_def) have *: "(\iifinite A\ by (intro prod.cong) (auto simp: n_def lead_coeff_power) hence "lead_coeff q = (\i = (\ifinite A\ by (intro prod.cong) (auto simp: lead_coeff_power n_def) finally have [simp]: "lead_coeff q = 1" by simp have "lead_monom q = (\i = (\ifinite A\ by (intro sum.cong) (auto simp: lead_monom_power n_def) also have "\ = (\iset xs = A\) also from 1 have "\ = monom_of_set (set (take (Suc i) xs))" by (subst lead_monom_sym_mpoly) (auto simp: xs_def n_def) finally show ?case by simp qed finally have lead_monom_q: "lead_monom q = (\i = lead_monom p" (is "?S = _") proof (intro poly_mapping_eqI) fix i :: nat show "lookup (lead_monom c + lead_monom q) i = lookup (lead_monom p) i" proof (cases "i \ A") case False hence "lookup (lead_monom c + lead_monom q) i = lookup (lead_monom p) i + (\jset xs = A\ dest!: in_set_takeD) finally show ?thesis by simp next case True with \set xs = A\ obtain m where m: "i = xs ! m" "m < n" by (auto simp: set_conv_nth) have "lookup (lead_monom c + lead_monom q) i = (\j = (\j | j < n \ i \ set (take (Suc j) xs). g j - g (Suc j))" by (intro sum.mono_neutral_cong_right) auto also have "{j. j < n \ i \ set (take (Suc j) xs)} = {m..distinct xs\ by (force simp: set_conv_nth nth_eq_iff_index_eq) also have "(\j\\. g j - g (Suc j)) = (\j\\. g j) - (\j\\. g (Suc j))" by (subst sum_subtractf_nat) (auto intro!: g_mono) also have "(\j\{m..j\{m<..n}. g j)" by (intro sum.reindex_bij_witness[of _ "\j. j - 1" Suc]) auto also have "\ = (\j\{m<..j\{m.. = (\j\{m..j\\. g j) = lookup (lead_monom p) i" using m by (auto simp: g_def not_less le_Suc_eq f_def) finally show ?thesis . qed qed finally show "lead_monom (fund_sym_step_poly p) = lead_monom p" by simp show "lead_coeff (fund_sym_step_poly p) = lead_coeff p" by (simp add: lead_coeff_mult) have *: "lookup (fund_sym_step_monom p) k = (if k \ {1..n} then g (k - 1) - g k else 0)" for k proof - have "lookup (fund_sym_step_monom p) k = (\x\(if k \ {1..n} then {k - 1} else {}). g (k - 1) - g k)" unfolding fund_sym_step_monom_def lookup_sum Let_def by (intro sum.mono_neutral_cong_right) (auto simp: g_def lookup_single when_def split: if_splits) thus ?thesis by simp qed hence "(\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x :: 'a mpoly) = (\x\{1..n}. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" by (intro Prod_any.expand_superset) auto also have "\ = (\xi. i - 1"]) auto also have "\ = q" unfolding q_def by (intro prod.cong) (auto simp: *) finally show "fund_sym_step_poly p = fund_sym_step_coeff p * (\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" by (simp add: c_def q_def f_def g_def fund_sym_step_monom_def fund_sym_step_coeff_def) qed text \ If the input is well-formed, a single step of the procedure always decreases the leading monomial. \ lemma lead_monom_fund_sym_step_poly_less: assumes "finite A" and "lead_monom p \ 0" and "decr p" shows "lead_monom (p - fund_sym_step_poly p) < lead_monom p" proof (cases "p = fund_sym_step_poly p") case True thus ?thesis using assms by (auto simp: order.strict_iff_order) next case False from assms have [simp]: "p \ 0" by auto let ?q = "fund_sym_step_poly p" and ?m = "lead_monom p" have "coeff (p - ?q) ?m = 0" using fund_sym_step_poly[of p] assms by (simp add: lead_coeff_def) moreover have "lead_coeff (p - ?q) \ 0" using False by auto ultimately have "lead_monom (p - ?q) \ ?m" unfolding lead_coeff_def by auto moreover have "lead_monom (p - ?q) \ ?m" using fund_sym_step_poly[of p] assms by (intro order.trans[OF lead_monom_diff] max.boundedI) auto ultimately show ?thesis by (auto simp: order.strict_iff_order) qed text \ Finally, we prove that the witness is indeed well-defined for all inputs. \ lemma fund_sym_poly_wit_dom_aux: assumes "finite B" "vars p \ B" "A \ B" shows "fund_sym_poly_wit_dom p" using assms(1-3) proof (induction p rule: lead_monom_induct) case (less p) have [simp]: "finite A" by (rule finite_subset[of _ B]) fact+ show ?case proof (cases "lead_monom p = 0 \ \symmetric_mpoly A p") case False hence [simp]: "p \ 0" by auto note decr = lookup_lead_monom_decreasing[of A p] have "vars (p - fund_sym_step_poly p) \ B" using fund_sym_step_poly[of p] decr False less.prems less.hyps \A \ B\ by (intro order.trans[OF vars_diff]) auto hence "fund_sym_poly_wit_dom (p - local.fund_sym_step_poly p)" using False less.prems less.hyps decr by (intro less.IH fund_sym_step_poly symmetric_mpoly_diff lead_monom_fund_sym_step_poly_less) (auto simp: decr_def) thus ?thesis using fund_sym_poly_wit.domintros by blast qed (auto intro: fund_sym_poly_wit.domintros) qed lemma fund_sym_poly_wit_dom [intro]: "fund_sym_poly_wit_dom p" proof - consider "\symmetric_mpoly A p" | "vars p \ A = {}" | "symmetric_mpoly A p" "A \ vars p" using symmetric_mpoly_imp_orthogonal_or_subset[of A p] by blast thus ?thesis proof cases assume "symmetric_mpoly A p" "A \ vars p" thus ?thesis using fund_sym_poly_wit_dom_aux[of "vars p" p] by (auto simp: vars_finite) qed (auto intro: fund_sym_poly_wit.domintros) qed termination fund_sym_poly_wit by (intro allI fund_sym_poly_wit_dom) (*<*) lemmas [simp del] = fund_sym_poly_wit.simps (*>*) text \ Next, we prove that our witness indeed fulfils all the properties stated by the fundamental theorem: \<^enum> If the original polynomial was in $R[X_1,\ldots,X_n,\ldots, X_m]$ where the $X_1$ to $X_n$ are the symmetric variables, then the witness is a polynomial in $R[X_{n+1},\ldots,X_m][Y_1,\ldots,Y_n]$. This means that its coefficients are polynomials in the variables of the original polynomial, minus the symmetric ones, and the (new and independent) variables of the witness polynomial range from $1$ to $n$. \<^enum> Substituting the \i\-th symmetric polynomial $e_i(X_1,\ldots,X_n)$ for the $Y_i$ variable for every \i\ yields the original polynomial. \<^enum> The coefficient ring $R$ need not be the entire type; if the coefficients of the original polynomial are in some subring, then the coefficients of the coefficients of the witness also do. \ lemma fund_sym_poly_wit_coeffs_aux: assumes "finite B" "vars p \ B" "symmetric_mpoly A p" "A \ B" shows "vars (coeff (fund_sym_poly_wit p) m) \ B - A" using assms proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) show ?case proof (cases "lead_monom p = 0 \ vars p \ A = {}") case False have "vars (p - fund_sym_step_poly p) \ B" using "1.prems" fund_sym_step_poly[of p] by (intro order.trans[OF vars_diff]) auto with 1 False have "vars (coeff (fund_sym_poly_wit (p - fund_sym_step_poly p)) m) \ B - A" by (intro 1 symmetric_mpoly_diff fund_sym_step_poly) auto hence "vars (coeff (fund_sym_poly_wit (p - fund_sym_step_poly p) + monom (fund_sym_step_monom p) (fund_sym_step_coeff p)) m) \ B - A" unfolding coeff_add coeff_monom using vars_fund_sym_step_coeff[of p] "1.prems" by (intro order.trans[OF vars_add] Un_least order.trans[OF vars_monom_subset]) (auto simp: when_def) thus ?thesis using "1.prems" False unfolding fund_sym_poly_wit.simps[of p] by simp qed (insert "1.prems", auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const lead_monom_eq_0_iff) qed lemma fund_sym_poly_wit_coeffs: assumes "symmetric_mpoly A p" shows "vars (coeff (fund_sym_poly_wit p) m) \ vars p - A" proof (cases "A \ vars p") case True with fund_sym_poly_wit_coeffs_aux[of "vars p" p m] assms show ?thesis by (auto simp: vars_finite) next case False hence "vars p \ A = {}" using symmetric_mpoly_imp_orthogonal_or_subset[OF assms] by auto thus ?thesis by (auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const) qed lemma fund_sym_poly_wit_vars: "vars (fund_sym_poly_wit p) \ {1..n}" proof (cases "symmetric_mpoly A p \ A \ vars p") case True define B where "B = vars p" have "finite B" "vars p \ B" "symmetric_mpoly A p" "A \ B" using True unfolding B_def by (auto simp: vars_finite) thus ?thesis proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) show ?case proof (cases "lead_monom p = 0 \ vars p \ A = {}") case False have "vars (p - fund_sym_step_poly p) \ B" using "1.prems" fund_sym_step_poly[of p] by (intro order.trans[OF vars_diff]) auto hence "vars (local.fund_sym_poly_wit (p - local.fund_sym_step_poly p)) \ {1..n}" using False "1.prems" by (intro 1 symmetric_mpoly_diff fund_sym_step_poly) (auto simp: lead_monom_eq_0_iff) hence "vars (fund_sym_poly_wit (p - fund_sym_step_poly p) + monom (fund_sym_step_monom p) (local.fund_sym_step_coeff p)) \ {1..n}" by (intro order.trans[OF vars_add] Un_least order.trans[OF vars_monom_subset] keys_fund_sym_step_monom) auto thus ?thesis using "1.prems" False unfolding fund_sym_poly_wit.simps[of p] by simp qed (insert "1.prems", auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const lead_monom_eq_0_iff) qed next case False then consider "\symmetric_mpoly A p" | "symmetric_mpoly A p" "vars p \ A = {}" using symmetric_mpoly_imp_orthogonal_or_subset[of A p] by auto thus ?thesis by cases (auto simp: fund_sym_poly_wit.simps[of p]) qed lemma fund_sym_poly_wit_insertion_aux: assumes "finite B" "vars p \ B" "symmetric_mpoly A p" "A \ B" shows "insertion (sym_mpoly A) (fund_sym_poly_wit p) = p" using assms proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) from "1.prems" have "decr p" using lookup_lead_monom_decreasing[of A p] by (auto simp: decr_def) show ?case proof (cases "lead_monom p = 0 \ vars p \ A = {}") case False have "vars (p - fund_sym_step_poly p) \ B" using "1.prems" fund_sym_step_poly[of p] by (intro order.trans[OF vars_diff]) auto hence "insertion (sym_mpoly A) (fund_sym_poly_wit (p - fund_sym_step_poly p)) = p - fund_sym_step_poly p" using 1 False by (intro 1 symmetric_mpoly_diff fund_sym_step_poly) auto moreover have "fund_sym_step_poly p = fund_sym_step_coeff p * (\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" using "1.prems" finite_subset[of A B] False \decr p\ by (intro fund_sym_step_poly) auto ultimately show ?thesis unfolding fund_sym_poly_wit.simps[of p] by (auto simp: insertion_add) qed (auto simp: fund_sym_poly_wit.simps[of p]) qed lemma fund_sym_poly_wit_insertion: assumes "symmetric_mpoly A p" shows "insertion (sym_mpoly A) (fund_sym_poly_wit p) = p" proof (cases "A \ vars p") case False hence "vars p \ A = {}" using symmetric_mpoly_imp_orthogonal_or_subset[OF assms] by auto thus ?thesis by (auto simp: fund_sym_poly_wit.simps[of p]) next case True with fund_sym_poly_wit_insertion_aux[of "vars p" p] assms show ?thesis by (auto simp: vars_finite) qed lemma fund_sym_poly_wit_coeff: assumes "\m. coeff p m \ C" "ring_closed C" shows "\m m'. coeff (coeff (fund_sym_poly_wit p) m) m' \ C" using assms(1) proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) interpret ring_closed C by fact show ?case proof (cases "\symmetric_mpoly A p \ lead_monom p = 0 \ vars p \ A = {}") case True thus ?thesis using "1.prems" by (auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const) next case False have *: "\m m'. coeff (coeff (fund_sym_poly_wit (p - fund_sym_step_poly p)) m) m' \ C" using False "1.prems" assms coeff_fund_sym_step_poly [of p] by (intro 1) auto show ?thesis proof (intro allI, goal_cases) case (1 m m') thus ?case using * False coeff_fund_sym_step_coeff[of p m'] "1.prems" by (auto simp: fund_sym_poly_wit.simps[of p] coeff_monom lead_coeff_def when_def) qed qed qed subsection \Uniqueness\ text \ Next, we show that the polynomial representation of a symmetric polynomial in terms of the elementary symmetric polynomials not only exists, but is unique. The key property here is that products of powers of elementary symmetric polynomials uniquely determine the exponent vectors, i.\,e.\ if $e_1, \ldots, e_n$ are the elementary symmetric polynomials, $a = (a_1,\ldots, a_n)$ and $b = (b_1,\ldots,b_n)$ are vectors of natural numbers, then: \[e_1^{a_1}\ldots e_n^{a_n} = e_1^{b_1}\ldots e_n^{b_n} \longleftrightarrow a = b\] We show this now. \ lemma lead_monom_sym_mpoly_prod: assumes "finite A" shows "lead_monom (\i = 1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i = 1..n. of_nat (h i) * lead_monom (sym_mpoly A i :: 'a mpoly))" proof - have "(\i=1..n. lead_coeff (sym_mpoly A i ^ h i :: 'a mpoly)) = 1" using assms unfolding n_def by (intro prod.neutral allI) (auto simp: lead_coeff_power) hence "lead_monom (\i=1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i=1..n. lead_monom (sym_mpoly A i ^ h i :: 'a mpoly))" by (subst lead_monom_prod) auto also have "\ = (\i=1..n. of_nat (h i) * lead_monom (sym_mpoly A i :: 'a mpoly))" by (intro sum.cong refl, subst lead_monom_power) (auto simp: lead_coeff_power assms n_def) finally show ?thesis . qed lemma lead_monom_sym_mpoly_prod_notin: assumes "finite A" "k \ A" shows "lookup (lead_monom (\i=1..n. sym_mpoly A i ^ h i :: 'a mpoly)) k = 0" proof - have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) have "lead_monom (\i = 1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i = 1..n. of_nat (h i) * lead_monom (sym_mpoly (set xs) i :: 'a mpoly))" by (subst lead_monom_sym_mpoly_prod) (use xs assms in auto) also have "lookup \ k = 0" unfolding lookup_sum by (intro sum.neutral ballI, subst lead_monom_sym_mpoly) (insert xs assms, auto simp: xs lead_monom_sym_mpoly lookup_monom_of_set set_conv_nth) finally show ?thesis . qed lemma lead_monom_sym_mpoly_prod_in: assumes "finite A" "k < n" shows "lookup (lead_monom (\i=1..n. sym_mpoly A i ^ h i :: 'a mpoly)) (xs ! k) = (\i=k+1..n. h i)" proof - have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) have "lead_monom (\i = 1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i = 1..n. of_nat (h i) * lead_monom (sym_mpoly (set xs) i :: 'a mpoly))" by (subst lead_monom_sym_mpoly_prod) (use xs assms in simp_all) also have "\ = (\i=1..n. of_nat (h i) * monom_of_set (set (take i xs)))" using xs by (intro sum.cong refl, subst lead_monom_sym_mpoly) auto also have "lookup \ (xs ! k) = (\i | i \ {1..n} \ xs ! k \ set (take i xs). h i)" unfolding lookup_sum lookup_monom_of_set by (intro sum.mono_neutral_cong_right) auto also have "{i. i \ {1..n} \ xs ! k \ set (take i xs)} = {k+1..n}" proof (intro equalityI subsetI) fix i assume i: "i \ {k+1..n}" hence "take i xs ! k = xs ! k" "k < n" "k < i" using assms by auto with i show "i \ {i. i \ {1..n} \ xs ! k \ set (take i xs)}" by (force simp: set_conv_nth) qed (insert assms xs, auto simp: set_conv_nth Suc_le_eq nth_eq_iff_index_eq) finally show ?thesis . qed lemma lead_monom_sym_poly_powerprod_inj: assumes "lead_monom (\i. sym_mpoly A i ^ lookup m1 i :: 'a mpoly) = lead_monom (\i. sym_mpoly A i ^ lookup m2 i :: 'a mpoly)" assumes "finite A" "keys m1 \ {1..n}" "keys m2 \ {1..n}" shows "m1 = m2" proof (rule poly_mapping_eqI) fix k :: nat have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) from assms(3,4) have *: "i \ {1..n}" if "lookup m1 i \ 0 \ lookup m2 i \ 0" for i using that by (auto simp: subset_iff in_keys_iff) have **: "(\i. sym_mpoly A i ^ lookup m i :: 'a mpoly) = (\i=1..n. sym_mpoly A i ^ lookup m i :: 'a mpoly)" if "m \ {m1, m2}" for m using that * by (intro Prod_any.expand_superset subsetI * ) (auto intro!: Nat.gr0I) have ***: "lead_monom (\i=1..n. sym_mpoly A i ^ lookup m1 i :: 'a mpoly) = lead_monom (\i=1..n. sym_mpoly A i ^ lookup m2 i :: 'a mpoly)" using assms by (simp add: ** ) have sum_eq: "sum (lookup m1) {Suc k..n} = sum (lookup m2) {Suc k..n}" if "k < n" for k using arg_cong[OF ***, of "\m. lookup m (xs ! k)"] \finite A\ that by (subst (asm) (1 2) lead_monom_sym_mpoly_prod_in) auto show "lookup m1 k = lookup m2 k" proof (cases "k \ {1..n}") case False hence "lookup m1 k = 0" "lookup m2 k = 0" using assms by (auto simp: in_keys_iff) thus ?thesis by simp next case True thus ?thesis proof (induction "n - k" arbitrary: k rule: less_induct) case (less l) have "sum (lookup m1) {Suc (l - 1)..n} = sum (lookup m2) {Suc (l - 1)..n}" using less by (intro sum_eq) auto also have "{Suc (l - 1)..n} = insert l {Suc l..n}" using less by auto also have "sum (lookup m1) \ = lookup m1 l + (\i=Suc l..n. lookup m1 i)" by (subst sum.insert) auto also have "(\i=Suc l..n. lookup m1 i) = (\i=Suc l..n. lookup m2 i)" by (intro sum.cong less) auto also have "sum (lookup m2) (insert l {Suc l..n}) = lookup m2 l + (\i=Suc l..n. lookup m2 i)" by (subst sum.insert) auto finally show "lookup m1 l = lookup m2 l" by simp qed qed qed text \ We now show uniqueness by first showing that the zero polynomial has a unique representation. We fix some polynomial $p$ with $p(e_1,\ldots, e_n) = 0$ and then show, by contradiction, that $p = 0$. We have \[p(e_1,\ldots,e_n) = \sum c_{a_1,\ldots,a_n} e_1^{a_1}\ldots e_n^{a_n}\] and due to the injectivity of products of powers of elementary symmetric polynomials, the leading term of that sum is precisely the leading term of the summand with the biggest leading monomial, since summands cannot cancel each other. However, we also know that $p(e_1,\ldots,e_n) = 0$, so it follows that all summands must have leading term 0, and it is then easy to see that they must all be identically 0. \ lemma sym_mpoly_representation_unique_aux: fixes p :: "'a mpoly mpoly" assumes "finite A" "insertion (sym_mpoly A) p = 0" "\m. vars (coeff p m) \ A = {}" "vars p \ {1..n}" shows "p = 0" proof (rule ccontr) assume p: "p \ 0" have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) define h where "h = (\m. coeff p m * (\i. sym_mpoly A i ^ lookup m i))" define M where "M = {m. coeff p m \ 0}" define maxm where "maxm = Max ((lead_monom \ h) ` M)" have "finite M" by (auto intro!: finite_subset[OF _ finite_coeff_support[of p]] simp: h_def M_def) have keys_subset: "keys m \ {1..n}" if "coeff p m \ 0" for m using that assms coeff_notin_vars[of m p] by blast have lead_coeff: "lead_coeff (h m) = lead_coeff (coeff p m)" (is ?th1) and lead_monom: "lead_monom (h m) = lead_monom (coeff p m) + lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)" (is ?th2) if [simp]: "coeff p m \ 0" for m proof - have "(\i. sym_mpoly A i ^ lookup m i :: 'a mpoly) = (\i | lookup m i \ 0. sym_mpoly A i ^ lookup m i :: 'a mpoly)" by (intro Prod_any.expand_superset) (auto intro!: Nat.gr0I) also have "lead_coeff \ = 1" using assms keys_subset[of m] by (intro lead_coeff_sym_mpoly_powerprod) (auto simp: in_keys_iff subset_iff n_def) finally have eq: "lead_coeff (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly) = 1" . thus ?th1 unfolding h_def using \coeff p m \ 0\ by (subst lead_coeff_mult) auto show ?th2 unfolding h_def by (subst lead_monom_mult) (auto simp: eq) qed have "insertion (sym_mpoly A) p = (\m\M. h m)" unfolding insertion_altdef h_def M_def by (intro Sum_any.expand_superset) auto also have "lead_monom \ = maxm" unfolding maxm_def proof (rule lead_monom_sum) from p obtain m where "coeff p m \ 0" using mpoly_eqI[of p 0] by auto hence "m \ M" using \coeff p m \ 0\ lead_coeff[of m] by (auto simp: M_def) thus "M \ {}" by auto next have restrict_lead_monom: "restrictpm A (lead_monom (h m)) = lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)" if [simp]: "coeff p m \ 0" for m proof - have "restrictpm A (lead_monom (h m)) = restrictpm A (lead_monom (coeff p m)) + restrictpm A (lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly))" by (auto simp: lead_monom restrictpm_add) also have "restrictpm A (lead_monom (coeff p m)) = 0" using assms by (intro restrictpm_orthogonal order.trans[OF keys_lead_monom_subset]) auto also have "restrictpm A (lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)) = lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)" by (intro restrictpm_id order.trans[OF keys_lead_monom_subset] order.trans[OF vars_Prod_any] UN_least order.trans[OF vars_power] vars_sym_mpoly_subset) finally show ?thesis by simp qed show "inj_on (lead_monom \ h) M" proof fix m1 m2 assume m12: "m1 \ M" "m2 \ M" "(lead_monom \ h) m1 = (lead_monom \ h) m2" hence [simp]: "coeff p m1 \ 0" "coeff p m2 \ 0" by (auto simp: M_def h_def) have "restrictpm A (lead_monom (h m1)) = restrictpm A (lead_monom (h m2))" using m12 by simp hence "lead_monom (\i. sym_mpoly A i ^ lookup m1 i :: 'a mpoly) = lead_monom (\i. sym_mpoly A i ^ lookup m2 i :: 'a mpoly)" by (simp add: restrict_lead_monom) thus "m1 = m2" by (rule lead_monom_sym_poly_powerprod_inj) (use \finite A\ keys_subset[of m1] keys_subset[of m2] in auto) qed next fix m assume "m \ M" hence "lead_coeff (h m) = lead_coeff (coeff p m)" by (simp add: lead_coeff M_def) with \m \ M\ show "h m \ 0" by (auto simp: M_def) qed fact+ finally have "maxm = 0" by (simp add: assms) have only_zero: "m = 0" if "m \ M" for m proof - from that have nz [simp]: "coeff p m \ 0" by (auto simp: M_def h_def) from that have "(lead_monom \ h) m \ maxm" using \finite M\ unfolding maxm_def by (intro Max_ge imageI finite_imageI) with \maxm = 0\ have [simp]: "lead_monom (h m) = 0" by simp have lookup_nzD: "k \ {1..n}" if "lookup m k \ 0" for k using keys_subset[of m] that by (auto simp: in_keys_iff subset_iff) have "lead_monom (coeff p m) + 0 \ lead_monom (h m)" unfolding lead_monom[OF nz] by (intro add_left_mono) auto also have "\ = 0" by simp finally have lead_monom_0: "lead_monom (coeff p m) = 0" by simp have "sum (lookup m) {1..n} = 0" proof (rule ccontr) assume "sum (lookup m) {1..n} \ 0" hence "sum (lookup m) {1..n} > 0" by presburger have "0 \ lead_coeff (MPoly_Type.coeff p m)" by auto also have "lead_coeff (MPoly_Type.coeff p m) = lead_coeff (h m)" by (simp add: lead_coeff) also have "lead_coeff (h m) = coeff (h m) 0" by (simp add: lead_coeff_def) also have "\ = coeff (coeff p m) 0 * coeff (\i. sym_mpoly A i ^ lookup m i) 0" by (simp add: h_def mpoly_coeff_times_0) also have "(\i. sym_mpoly A i ^ lookup m i) = (\i=1..n. sym_mpoly A i ^ lookup m i)" by (intro Prod_any.expand_superset subsetI lookup_nzD) (auto intro!: Nat.gr0I) also have "coeff \ 0 = (\i=1..n. 0 ^ lookup m i)" unfolding mpoly_coeff_prod_0 mpoly_coeff_power_0 by (intro prod.cong) (auto simp: coeff_sym_mpoly_0) also have "\ = 0 ^ (\i=1..n. lookup m i)" by (simp add: power_sum) also have "\ = 0" using zero_power[OF \sum (lookup m) {1..n} > 0\] by simp finally show False by auto qed hence "lookup m k = 0" for k using keys_subset[of m] by (cases "k \ {1..n}") (auto simp: in_keys_iff) thus "m = 0" by (intro poly_mapping_eqI) auto qed have "0 = insertion (sym_mpoly A) p" using assms by simp also have "insertion (sym_mpoly A) p = (\m\M. h m)" by fact also have "\ = (\m\{0}. h m)" using only_zero by (intro sum.mono_neutral_left) (auto simp: h_def M_def) also have "\ = coeff p 0" by (simp add: h_def) finally have "0 \ M" by (auto simp: M_def) with only_zero have "M = {}" by auto hence "p = 0" by (intro mpoly_eqI) (auto simp: M_def) with \p \ 0\ show False by contradiction qed text \ The general uniqueness theorem now follows easily. This essentially shows that the substitution $Y_i \mapsto e_i(X_1,\ldots,X_n)$ is an isomorphism between the ring $R[Y_1,\ldots, Y_n]$ and the ring $R[X_1,\ldots,X_n]^{S_n}$ of symmetric polynomials. \ theorem sym_mpoly_representation_unique: fixes p :: "'a mpoly mpoly" assumes "finite A" "insertion (sym_mpoly A) p = insertion (sym_mpoly A) q" "\m. vars (coeff p m) \ A = {}" "\m. vars (coeff q m) \ A = {}" "vars p \ {1..n}" "vars q \ {1..n}" shows "p = q" proof - have "p - q = 0" proof (rule sym_mpoly_representation_unique_aux) fix m show "vars (coeff (p - q) m) \ A = {}" using vars_diff[of "coeff p m" "coeff q m"] assms(3,4)[of m] by auto qed (insert assms vars_diff[of p q], auto simp: insertion_diff) thus ?thesis by simp qed theorem eq_fund_sym_poly_witI: fixes p :: "'a mpoly" and q :: "'a mpoly mpoly" assumes "finite A" "symmetric_mpoly A p" "insertion (sym_mpoly A) q = p" "\m. vars (coeff q m) \ A = {}" "vars q \ {1..n}" shows "q = fund_sym_poly_wit p" using fund_sym_poly_wit_insertion[of p] fund_sym_poly_wit_vars[of p] fund_sym_poly_wit_coeffs[of p] by (intro sym_mpoly_representation_unique) (insert assms, auto simp: fund_sym_poly_wit_insertion) subsection \A recursive characterisation of symmetry\ text \ In a similar spirit to the proof of the fundamental theorem, we obtain a nice recursive and executable characterisation of symmetry. \ (*<*) lemmas [fundef_cong] = disj_cong conj_cong (*>*) function (domintros) check_symmetric_mpoly where "check_symmetric_mpoly p \ (vars p \ A = {} \ A \ vars p \ decr p \ check_symmetric_mpoly (p - fund_sym_step_poly p))" by auto lemma check_symmetric_mpoly_dom_aux: assumes "finite B" "vars p \ B" "A \ B" shows "check_symmetric_mpoly_dom p" using assms(1-3) proof (induction p rule: lead_monom_induct) case (less p) have [simp]: "finite A" by (rule finite_subset[of _ B]) fact+ show ?case proof (cases "lead_monom p = 0 \ \decr p") case False hence [simp]: "p \ 0" by auto have "vars (p - fund_sym_step_poly p) \ B" using fund_sym_step_poly[of p] False less.prems less.hyps \A \ B\ by (intro order.trans[OF vars_diff]) auto hence "check_symmetric_mpoly_dom (p - local.fund_sym_step_poly p)" using False less.prems less.hyps by (intro less.IH fund_sym_step_poly symmetric_mpoly_diff lead_monom_fund_sym_step_poly_less) (auto simp: decr_def) thus ?thesis using check_symmetric_mpoly.domintros by blast qed (auto intro: check_symmetric_mpoly.domintros simp: lead_monom_eq_0_iff) qed lemma check_symmetric_mpoly_dom [intro]: "check_symmetric_mpoly_dom p" proof - show ?thesis proof (cases "A \ vars p") assume "A \ vars p" thus ?thesis using check_symmetric_mpoly_dom_aux[of "vars p" p] by (auto simp: vars_finite) qed (auto intro: check_symmetric_mpoly.domintros) qed termination check_symmetric_mpoly by (intro allI check_symmetric_mpoly_dom) lemmas [simp del] = check_symmetric_mpoly.simps lemma check_symmetric_mpoly_correct: "check_symmetric_mpoly p \ symmetric_mpoly A p" proof (induction p rule: check_symmetric_mpoly.induct) case (1 p) have "symmetric_mpoly A (p - fund_sym_step_poly p) \ symmetric_mpoly A p" (is "?lhs = ?rhs") proof assume ?rhs thus ?lhs by (intro symmetric_mpoly_diff fund_sym_step_poly) next assume ?lhs hence "symmetric_mpoly A (p - fund_sym_step_poly p + fund_sym_step_poly p)" by (intro symmetric_mpoly_add fund_sym_step_poly) thus ?rhs by simp qed moreover have "decr p" if "symmetric_mpoly A p" using lookup_lead_monom_decreasing[of A p] that by (auto simp: decr_def) ultimately show "check_symmetric_mpoly p \ symmetric_mpoly A p" using 1 symmetric_mpoly_imp_orthogonal_or_subset[of A p] by (auto simp: Let_def check_symmetric_mpoly.simps[of p] intro: symmetric_mpoly_orthogonal) qed end subsection \Symmetric functions of roots of a univariate polynomial\ text \ Consider a factored polynomial \[p(X) = c_n X^n + c_{n-1} X^{n-1} + \ldots + c_1X + c_0 = (X - x_1)\ldots(X - x_n)\ .\] where $c_n$ is a unit. Then any symmetric polynomial expression $q(x_1, \ldots, x_n)$ in the roots $x_i$ can be written as a polynomial expression $q'(c_0,\ldots, c_{n-1})$ in the $c_i$. Moreover, if the coefficients of $q$ and the inverse of $c_n$ all lie in some subring, the coefficients of $q'$ do as well. \ context fixes C :: "'b :: comm_ring_1 set" and A :: "nat set" and root :: "nat \ 'a :: comm_ring_1" and l :: "'a \ 'b" and q :: "'b mpoly" and n :: nat defines "n \ card A" assumes C: "ring_closed C" "\m. coeff q m \ C" assumes l: "ring_homomorphism l" assumes finite: "finite A" assumes sym: "symmetric_mpoly A q" and vars: "vars q \ A" begin interpretation ring_closed C by fact interpretation ring_homomorphism l by fact theorem symmetric_poly_of_roots_conv_poly_of_coeffs: assumes c: "cinv * l c = 1" "cinv \ C" assumes "p = Polynomial.smult c (\i\A. [:-root i, 1:])" obtains q' where "vars q' \ {0..m. coeff q' m \ C" and "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" proof - define q' where "q' = fund_sym_poly_wit A q" define q'' where "q'' = mapm_mpoly (\m x. (\i. (cinv * l (- 1) ^ i) ^ lookup m i) * insertion (\_. 0) x) q'" define reindex where "reindex = (\i. if i \ n then n - i else i)" have "bij reindex" by (intro bij_betwI[of reindex _ _ reindex]) (auto simp: reindex_def) have "vars q' \ {1..n}" unfolding q'_def n_def by (intro fund_sym_poly_wit_vars) hence "vars q'' \ {1..n}" unfolding q''_def using vars_mapm_mpoly_subset by auto have "insertion (l \ root) (insertion (sym_mpoly A) q') = insertion (\n. insertion (l \ root) (sym_mpoly A n)) (map_mpoly (insertion (l \ root)) q')" by (rule insertion_insertion) also have "insertion (sym_mpoly A) q' = q" unfolding q'_def by (intro fund_sym_poly_wit_insertion sym) also have "insertion (\i. insertion (l \ root) (sym_mpoly A i)) (map_mpoly (insertion (l \ root)) q') = insertion (\i. cinv * l ((- 1) ^ i) * l (poly.coeff p (n - i))) (map_mpoly (insertion (l \ root)) q')" proof (intro insertion_irrelevant_vars, goal_cases) case (1 i) hence "i \ vars q'" using vars_map_mpoly_subset by auto also have "\ \ {1..n}" unfolding q'_def n_def by (intro fund_sym_poly_wit_vars) finally have i: "i \ {1..n}" . have "insertion (l \ root) (sym_mpoly A i) = l (\Y | Y \ A \ card Y = i. prod root Y)" using \finite A\ by (simp add: insertion_sym_mpoly) also have "\ = cinv * l (c * (\Y | Y \ A \ card Y = i. prod root Y))" unfolding mult mult.assoc[symmetric] \cinv * l c = 1\ by simp also have "c * (\Y | Y \ A \ card Y = i. prod root Y) = ((-1) ^ i * poly.coeff p (n - i))" using coeff_poly_from_roots[of A "n - i" root] i assms finite by (auto simp: n_def minus_one_power_iff) finally show ?case by (simp add: o_def) qed also have "map_mpoly (insertion (l \ root)) q' = map_mpoly (insertion (\_. 0)) q'" using fund_sym_poly_wit_coeffs[OF sym] vars by (intro map_mpoly_cong insertion_irrelevant_vars) (auto simp: q'_def) also have "insertion (\i. cinv * l ((- 1) ^ i) * l (poly.coeff p (n - i))) \ = insertion (\i. l (poly.coeff p (n - i))) q''" unfolding insertion_substitute_linear map_mpoly_conv_mapm_mpoly q''_def by (subst mapm_mpoly_comp) auto also have "\ = insertion (l \ poly.coeff p) (mpoly_map_vars reindex q'')" using \bij reindex\ and \vars q'' \ {1..n}\ by (subst insertion_mpoly_map_vars) (auto simp: o_def reindex_def intro!: insertion_irrelevant_vars) finally have "insertion (l \ root) q = insertion (l \ poly.coeff p) (mpoly_map_vars reindex q'')" . moreover have "coeff (mpoly_map_vars reindex q'') m \ C" for m unfolding q''_def q'_def using \bij reindex\ fund_sym_poly_wit_coeff[of q C A] C \cinv \ C\ by (auto simp: coeff_mpoly_map_vars intro!: mult_closed Prod_any_closed power_closed Sum_any_closed) moreover have "vars (mpoly_map_vars reindex q'') \ {0..bij reindex\ and \vars q'' \ {1..n}\ by (subst vars_mpoly_map_vars) (auto simp: reindex_def subset_iff)+ ultimately show ?thesis using that[of "mpoly_map_vars reindex q''"] by auto qed corollary symmetric_poly_of_roots_conv_poly_of_coeffs_monic: assumes "p = (\i\A. [:-root i, 1:])" obtains q' where "vars q' \ {0..m. coeff q' m \ C" and "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" proof - obtain q' where "vars q' \ {0..m. coeff q' m \ C" and "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" by (rule symmetric_poly_of_roots_conv_poly_of_coeffs[of 1 1 p]) (use assms in auto) thus ?thesis by (intro that[of q']) auto qed text \ As a corollary, we obtain the following: Let $R, S$ be rings with $R\subseteq S$. Consider a polynomial $p\in R[X]$ whose leading coefficient $c$ is a unit in $R$ and that has a full set of roots $x_1,\ldots, x_n \in S$, i.\,e.\ $p(X) = c(X - x_1) \ldots (X - x_n)$. Let $q \in R[X_1,\ldots, X_n]$ be some symmetric polynomial expression in the roots. Then $q(x_1, \ldots, x_n) \in R$. A typical use case is $R = \mathbb{Q}$ and $S = \mathbb{C}$, i.\,e.\ any symmetric polynomial expression with rational coefficients in the roots of a rational polynomial is again rational. Similarly, any symmetric polynomial expression with integer coefficients in the roots of a monic integer polynomial is agan an integer. This is remarkable, since the roots themselves are usually not rational (possibly not even real). This particular fact is a key ingredient used in the standard proof that $\pi$ is transcendental. \ corollary symmetric_poly_of_roots_in_subring: assumes "cinv * l c = 1" "cinv \ C" assumes "p = Polynomial.smult c (\i\A. [:-root i, 1:])" assumes "\i. l (poly.coeff p i) \ C" shows "insertion (\x. l (root x)) q \ C" proof - obtain q' where q': "vars q' \ {0..m. coeff q' m \ C" "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" by (rule symmetric_poly_of_roots_conv_poly_of_coeffs[of cinv c p]) (use assms in simp_all) have "insertion (l \ poly.coeff p) q' \ C" using C assms unfolding insertion_altdef by (intro Sum_any_closed mult_closed q' Prod_any_closed power_closed) auto also have "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" by fact finally show ?thesis by (simp add: o_def) qed corollary symmetric_poly_of_roots_in_subring_monic: assumes "p = (\i\A. [:-root i, 1:])" assumes "\i. l (poly.coeff p i) \ C" shows "insertion (\x. l (root x)) q \ C" proof - interpret ring_closed C by fact interpret ring_homomorphism l by fact show ?thesis by (rule symmetric_poly_of_roots_in_subring[of 1 1 p]) (use assms in auto) qed end end \ No newline at end of file diff --git a/thys/Twelvefold_Way/Preliminaries.thy b/thys/Twelvefold_Way/Preliminaries.thy --- a/thys/Twelvefold_Way/Preliminaries.thy +++ b/thys/Twelvefold_Way/Preliminaries.thy @@ -1,604 +1,604 @@ (* Author: Lukas Bulwahn *) section \Preliminaries\ theory Preliminaries imports Main "HOL-Library.Multiset" "HOL-Library.FuncSet" - "HOL-Library.Permutations" + "HOL-Combinatorics.Permutations" "HOL-ex.Birthday_Paradox" Card_Partitions.Card_Partitions Bell_Numbers_Spivey.Bell_Numbers Card_Multisets.Card_Multisets Card_Number_Partitions.Card_Number_Partitions begin subsection \Additions to Finite Set Theory\ lemma subset_with_given_card_exists: assumes "n \ card A" shows "\B \ A. card B = n" using assms proof (induct n) case 0 then show ?case by auto next case (Suc n) from this obtain B where "B \ A" "card B = n" by auto from this \B \ A\ \card B = n\ have "card B < card A" using Suc.prems by linarith from \Suc n \ card A\ card.infinite have "finite A" by force from this \B \ A\ finite_subset have "finite B" by blast from \card B < card A\ \B \ A\ obtain a where "a \ A" "a \ B" by (metis less_irrefl subsetI subset_antisym) have "insert a B \ A" "card (insert a B) = Suc n" using \finite B\ \a \ A\ \a \ B\ \B \ A\ \card B = n\ by auto then show ?case by blast qed subsection \Additions to Equiv Relation Theory\ lemmas univ_commute' = univ_commute[unfolded Equiv_Relations.proj_def] lemma univ_predicate_impl_forall: assumes "equiv A R" assumes "P respects R" assumes "X \ A // R" assumes "univ P X" shows "\x\X. P x" proof - from assms(1,3) obtain x where "x \ X" by (metis equiv_class_self quotientE) from \x \ X\ assms(1,3) have "X = R `` {x}" by (metis Image_singleton_iff equiv_class_eq quotientE) from assms(1,2,4) this show ?thesis using equiv_class_eq_iff univ_commute' by fastforce qed lemma univ_preserves_predicate: assumes "equiv A r" assumes "P respects r" shows "{x \ A. P x} // r = {X \ A // r. univ P X}" proof show "{x \ A. P x} // r \ {X \ A // r. univ P X}" proof fix X assume "X \ {x \ A. P x} // r" from this obtain x where "x \ {x \ A. P x}" and "X = r `` {x}" using quotientE by blast have "X \ A // r" using \X = r `` {x}\ \x \ {x \ A. P x}\ by (auto intro: quotientI) moreover have "univ P X" using \X = r `` {x}\ \x \ {x \ A. P x}\ assms by (simp add: proj_def[symmetric] univ_commute) ultimately show "X \ {X \ A // r. univ P X}" by auto qed next show "{X \ A // r. univ P X} \ {x \ A. P x} // r" proof fix X assume "X \ {X \ A // r. univ P X}" from this have "X \ A // r" and "univ P X" by auto from \X \ A // r\ obtain x where "x \ A" and "X = r `` {x}" using quotientE by blast have "x \ {x \ A. P x}" using \x \ A\ \X = r `` {x}\ \univ P X\ assms by (simp add: proj_def[symmetric] univ_commute) from this show "X \ {x \ A. P x} // r" using \X = r `` {x}\ by (auto intro: quotientI) qed qed lemma Union_quotient_restricted: assumes "equiv A r" assumes "P respects r" shows "\({x \ A. P x} // r) = {x \ A. P x}" proof show "\({x \ A. P x} // r) \ {x \ A. P x}" proof fix x assume "x \ \({x \ A. P x} // r)" from this obtain X where "x \ X" and "X \ {x \ A. P x} // r" by blast from this obtain x' where "X = r `` {x'}" and "x' \ {x \ A. P x}" using quotientE by blast from this \x \ X\ have "x \ A" using \equiv A r\ by (simp add: equiv_class_eq_iff) moreover from \X = r `` {x'}\ \x \ X\ \x' \ {x \ A. P x}\ have "P x" using \P respects r\ congruentD by fastforce ultimately show "x \ {x \ A. P x}" by auto qed next show "{x \ A. P x} \ \({x \ A. P x} // r)" proof fix x assume "x \ {x \ A. P x}" from this have "x \ r `` {x}" using \equiv A r\ equiv_class_self by fastforce from \x \ {x \ A. P x}\ have "r `` {x} \ {x \ A. P x} // r" by (auto intro: quotientI) from this \x \ r `` {x}\ show "x \ \({x \ A. P x} // r)" by auto qed qed lemma finite_equiv_implies_finite_carrier: assumes "equiv A R" assumes "finite (A // R)" assumes "\X \ A // R. finite X" shows "finite A" proof - from \equiv A R\ have "A = \(A // R)" by (simp add: Union_quotient) from this \finite (A // R)\ \\X \ A // R. finite X\ show "finite A" using finite_Union by fastforce qed lemma finite_quotient_iff: assumes "equiv A R" shows "finite A \ (finite (A // R) \ (\X \ A // R. finite X))" using assms by (meson equiv_type finite_equiv_class finite_equiv_implies_finite_carrier finite_quotient) subsubsection \Counting Sets by Splitting into Equivalence Classes\ lemma card_equiv_class_restricted: assumes "finite {x \ A. P x}" assumes "equiv A R" assumes "P respects R" shows "card {x \ A. P x} = sum card ({x \ A. P x} // R)" proof - have "card {x \ A. P x} = card (\({x \ A. P x} // R))" using \equiv A R\ \P respects R\ by (simp add: Union_quotient_restricted) also have "card (\({x \ A. P x} // R)) = (\C\{x \ A. P x} // R. card C)" proof - from \finite {x \ A. P x}\ have "finite ({x \ A. P x} // R)" using \equiv A R\ by (metis finite_imageI proj_image) moreover from \finite {x \ A. P x}\ have "\C\{x \ A. P x} // R. finite C" using \equiv A R\ \P respects R\ Union_quotient_restricted Union_upper finite_subset by fastforce moreover have "\C1 \ {x \ A. P x} // R. \C2 \ {x \ A. P x} // R. C1 \ C2 \ C1 \ C2 = {}" using \equiv A R\ quotient_disj by (metis (no_types, lifting) mem_Collect_eq quotientE quotientI) ultimately show ?thesis by (subst card_Union_disjoint) (auto simp: pairwise_def disjnt_def) qed finally show ?thesis . qed lemma card_equiv_class_restricted_same_size: assumes "equiv A R" assumes "P respects R" assumes "\F. F \ {x \ A. P x} // R \ card F = k" shows "card {x \ A. P x} = k * card ({x \ A. P x} // R)" proof cases assume "finite {x \ A. P x}" have "card {x \ A. P x} = sum card ({x \ A. P x} // R)" using \finite {x \ A. P x}\ \equiv A R\ \P respects R\ by (simp add: card_equiv_class_restricted) also have "sum card ({x \ A. P x} // R) = k * card ({x \ A. P x} // R)" by (simp add: \\F. F \ {x \ A. P x} // R \ card F = k\) finally show ?thesis . next assume "infinite {x \ A. P x}" from this have "infinite (\({a \ A. P a} // R))" using \equiv A R\ \P respects R\ by (simp add: Union_quotient_restricted) from this have "infinite ({x \ A. P x} // R) \ (\X \ {x \ A. P x} // R. infinite X)" by auto from this show ?thesis proof assume "infinite ({x \ A. P x} // R)" from this \infinite {x \ A. P x}\ show ?thesis by simp next assume "\X \ {x \ A. P x} // R. infinite X" from this \infinite {x \ A. P x}\ show ?thesis using \\F. F \ {x \ A. P x} // R \ card F = k\ card.infinite by auto qed qed lemma card_equiv_class: assumes "finite A" assumes "equiv A R" shows "card A = sum card (A // R)" proof - have "(\x. True) respects R" by (simp add: congruentI) from \finite A\ \equiv A R\ this show ?thesis using card_equiv_class_restricted[where P="\x. True"] by auto qed lemma card_equiv_class_same_size: assumes "equiv A R" assumes "\F. F \ A // R \ card F = k" shows "card A = k * card (A // R)" proof - have "(\x. True) respects R" by (simp add: congruentI) from \equiv A R\ \\F. F \ A // R \ card F = k\ this show ?thesis using card_equiv_class_restricted_same_size[where P="\x. True"] by auto qed subsection \Additions to FuncSet Theory\ lemma finite_same_card_bij_on_ext_funcset: assumes "finite A" "finite B" "card A = card B" shows "\f. f \ A \\<^sub>E B \ bij_betw f A B" proof - from assms obtain f' where f': "bij_betw f' A B" using finite_same_card_bij by auto define f where "\x. f x = (if x \ A then f' x else undefined)" have "f \ A \\<^sub>E B" using f' unfolding f_def by (auto simp add: bij_betwE) moreover have "bij_betw f A B" proof - have "bij_betw f' A B \ bij_betw f A B" unfolding f_def by (auto intro!: bij_betw_cong) from this \bij_betw f' A B\ show ?thesis by auto qed ultimately show ?thesis by auto qed lemma card_extensional_funcset: assumes "finite A" shows "card (A \\<^sub>E B) = card B ^ card A" using assms by (simp add: card_PiE prod_constant) lemma bij_betw_implies_inj_on_and_card_eq: assumes "finite B" assumes "f \ A \\<^sub>E B" shows "bij_betw f A B \ inj_on f A \ card A = card B" proof assume "bij_betw f A B" from this show "inj_on f A \ card A = card B" by (simp add: bij_betw_imp_inj_on bij_betw_same_card) next assume "inj_on f A \ card A = card B" from this have "inj_on f A" and "card A = card B" by auto from \f \ A \\<^sub>E B\ have "f ` A \ B" by auto from \inj_on f A\ have "card (f ` A) = card A" by (simp add: card_image) from \f ` A \ B\ \card A = card B\ this have "f ` A = B" by (simp add: \finite B\ card_subset_eq) from \inj_on f A\ this show "bij_betw f A B" by (rule bij_betw_imageI) qed lemma bij_betw_implies_surj_on_and_card_eq: assumes "finite A" assumes "f \ A \\<^sub>E B" shows "bij_betw f A B \ f ` A = B \ card A = card B" proof assume "bij_betw f A B" show "f ` A = B \ card A = card B" using \bij_betw f A B\ bij_betw_imp_surj_on bij_betw_same_card by blast next assume "f ` A = B \ card A = card B" from this have "f ` A = B" and "card A = card B" by auto from this have "inj_on f A" by (simp add: \finite A\ inj_on_iff_eq_card) from this \f ` A = B\ show "bij_betw f A B" by (rule bij_betw_imageI) qed subsection \Additions to Permutations Theory\ lemma assumes "f \ A \\<^sub>E B" "f ` A = B" assumes "p permutes B" "(\x. f' x = p (f x))" shows "(\b. {x\A. f x = b}) ` B = (\b. {x\A. f' x = b}) ` B" proof show "(\b. {x \ A. f x = b}) ` B \ (\b. {x \ A. f' x = b}) ` B" proof fix X assume "X \ (\b. {x \ A. f x = b}) ` B" from this obtain b where X_eq: "X = {x \ A. f x = b}" and "b \ B" by blast from assms(3, 4) have "\x. f x = b \ f' x = p b" by (metis permutes_def) from \p permutes B\ X_eq this have "X = {x \ A. f' x = p b}" using Collect_cong by auto moreover from \b \ B\ \p permutes B\ have "p b \ B" by (simp add: permutes_in_image) ultimately show "X \ (\b. {x \ A. f' x = b}) ` B" by blast qed next show "(\b. {x \ A. f' x = b}) ` B \ (\b. {x \ A. f x = b}) ` B" proof fix X assume "X \ (\b. {x \ A. f' x = b}) ` B" from this obtain b where X_eq: "X = {x \ A. f' x = b}" and "b \ B" by blast from assms(3, 4) have "\x. f' x = b \ f x = inv p b" by (auto simp add: permutes_inverses(1, 2)) from \p permutes B\ X_eq this have "X = {x \ A. f x = inv p b}" using Collect_cong by auto moreover from \b \ B\ \p permutes B\ have "inv p b \ B" by (simp add: permutes_in_image permutes_inv) ultimately show "X \ (\b. {x \ A. f x = b}) ` B" by blast qed qed subsection \Additions to List Theory\ text \ The theorem @{thm [source] card_lists_length_eq} contains the superfluous assumption @{term "finite A"}. Here, we derive that fact without that unnecessary assumption. \ lemma lists_length_eq_Suc_eq_image_Cons: "{xs. set xs \ A \ length xs = Suc n} = (\(x, xs). x#xs) ` (A \ {xs. set xs \ A \ length xs = n})" (is "?A = ?B") proof show "?A \ ?B" proof fix xs assume "xs \ ?A" from this show "xs \ ?B" by (cases xs) auto qed next show "?B \ ?A" by auto qed lemma lists_length_eq_Suc_eq_empty_iff: "{xs. set xs \ A \ length xs = Suc n} = {} \ A = {}" proof (induct n) case 0 have "{xs. set xs \ A \ length xs = Suc 0} = {x#[] |x. x \ A}" proof show "{[x] |x. x \ A} \ {xs. set xs \ A \ length xs = Suc 0}" by auto next show "{xs. set xs \ A \ length xs = Suc 0} \ {[x] |x. x \ A}" proof fix xs assume "xs \ {xs. set xs \ A \ length xs = Suc 0}" from this have "set xs \ A \ length xs = Suc 0" by simp from this have "\x. xs = [x] \ x \ A" by (metis Suc_length_conv insert_subset length_0_conv list.set(2)) from this show "xs \ {[x] |x. x \ A}" by simp qed qed then show ?case by simp next case (Suc n) from this show ?case by (auto simp only: lists_length_eq_Suc_eq_image_Cons) qed lemma lists_length_eq_eq_empty_iff: "{xs. set xs \ A \ length xs = n} = {} \ (A = {} \ n > 0)" proof (cases n) case 0 then show ?thesis by auto next case (Suc n) then show ?thesis by (auto simp only: lists_length_eq_Suc_eq_empty_iff) qed lemma finite_lists_length_eq_iff: "finite {xs. set xs \ A \ length xs = n} \ (finite A \ n = 0)" proof assume "finite {xs. set xs \ A \ length xs = n}" from this show "finite A \ n = 0" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "inj (\(x, xs). x#xs)" by (auto intro: inj_onI) from this Suc(2) have "finite (A \ {xs. set xs \ A \ length xs = n})" using finite_imageD inj_on_subset subset_UNIV lists_length_eq_Suc_eq_image_Cons[of A n] by fastforce from this have "finite A" by (cases "A = {}") (auto simp only: lists_length_eq_eq_empty_iff dest: finite_cartesian_productD1) from this show ?case by auto qed next assume "finite A \ n = 0" from this show "finite {xs. set xs \ A \ length xs = n}" by (auto intro: finite_lists_length_eq) qed lemma card_lists_length_eq: shows "card {xs. set xs \ B \ length xs = n} = card B ^ n" proof cases assume "finite B" then show ?thesis by (rule card_lists_length_eq) next assume "infinite B" then show ?thesis proof cases assume "n = 0" from this have "{xs. set xs \ B \ length xs = n} = {[]}" by auto from this \n = 0\ show ?thesis by simp next assume "n \ 0" from this \infinite B\ have "infinite {xs. set xs \ B \ length xs = n}" by (simp add: finite_lists_length_eq_iff) from this \infinite B\ show ?thesis by auto qed qed subsection \Additions to Disjoint Set Theory\ lemma bij_betw_congI: assumes "bij_betw f A A'" assumes "\a \ A. f a = g a" shows "bij_betw g A A'" using assms bij_betw_cong by fastforce lemma disjoint_family_onI[intro]: assumes "\m n. m \ S \ n \ S \ m \ n \ A m \ A n = {}" shows "disjoint_family_on A S" using assms unfolding disjoint_family_on_def by simp text \ The following lemma is not needed for this development, but is useful and could be moved to Disjoint Set theory or Equiv Relation theory if translated from set partitions to equivalence relations. \ lemma infinite_partition_on: assumes "infinite A" shows "infinite {P. partition_on A P}" proof - from \infinite A\ obtain x where "x \ A" by (meson finite.intros(1) finite_subset subsetI) from \infinite A\ have "infinite (A - {x})" by (simp add: infinite_remove) define singletons_except_one where "singletons_except_one = (\a'. (\a. if a = a' then {a, x} else {a}) ` (A - {x}))" have "infinite (singletons_except_one ` (A - {x}))" proof - have "inj_on singletons_except_one (A - {x})" unfolding singletons_except_one_def by (rule inj_onI) auto from \infinite (A - {x})\ this show ?thesis using finite_imageD by blast qed moreover have "singletons_except_one ` (A - {x}) \ {P. partition_on A P}" proof fix P assume "P \ singletons_except_one ` (A - {x})" from this obtain a' where "a' \ A - {x}" and P: "P = singletons_except_one a'" by blast have "partition_on A ((\a. if a = a' then {a, x} else {a}) ` (A - {x}))" using \x \ A\ \a' \ A - {x}\ by (auto intro: partition_onI) from this have "partition_on A P" unfolding P singletons_except_one_def . from this show "P \ {P. partition_on A P}" .. qed ultimately show ?thesis by (simp add: infinite_super) qed lemma finitely_many_partition_on_iff: "finite {P. partition_on A P} \ finite A" using finitely_many_partition_on infinite_partition_on by blast subsection \Additions to Multiset Theory\ lemma mset_set_subseteq_mset_set: assumes "finite B" "A \ B" shows "mset_set A \# mset_set B" proof - from \A \ B\ \finite B\ have "finite A" using finite_subset by blast { fix x have "count (mset_set A) x \ count (mset_set B) x" using \finite A\ \finite B\ \A \ B\ by (metis count_mset_set(1, 3) eq_iff subsetCE zero_le_one) } from this show "mset_set A \# mset_set B" using mset_subset_eqI by blast qed lemma mset_set_set_mset: assumes "M \# mset_set A" shows "mset_set (set_mset M) = M" proof - { fix x from \M \# mset_set A\ have "count M x \ count (mset_set A) x" by (simp add: mset_subset_eq_count) from this have "count (mset_set (set_mset M)) x = count M x" by (metis count_eq_zero_iff count_greater_eq_one_iff count_mset_set dual_order.antisym dual_order.trans finite_set_mset) } from this show ?thesis by (simp add: multiset_eq_iff) qed lemma mset_set_set_mset': assumes "\x. count M x \ 1" shows "mset_set (set_mset M) = M" proof - { fix x from assms have "count M x = 0 \ count M x = 1" by (auto elim: le_SucE) from this have "count (mset_set (set_mset M)) x = count M x" by (metis count_eq_zero_iff count_mset_set(1,3) finite_set_mset) } from this show ?thesis by (simp add: multiset_eq_iff) qed lemma card_set_mset: assumes "M \# mset_set A" shows "card (set_mset M) = size M" using assms by (metis mset_set_set_mset size_mset_set) lemma card_set_mset': assumes "\x. count M x \ 1" shows "card (set_mset M) = size M" using assms by (metis mset_set_set_mset' size_mset_set) lemma count_mset_set_leq: assumes "finite A" shows "count (mset_set A) x \ 1" using assms by (metis count_mset_set(1,3) eq_iff zero_le_one) lemma count_mset_set_leq': assumes "finite A" shows "count (mset_set A) x \ Suc 0" using assms count_mset_set_leq by fastforce lemma msubset_mset_set_iff: assumes "finite A" shows "set_mset M \ A \ (\x. count M x \ 1) \ (M \# mset_set A)" proof assume "set_mset M \ A \ (\x. count M x \ 1)" from this assms show "M \# mset_set A" by (metis count_inI count_mset_set(1) le0 mset_subset_eqI subsetCE) next assume "M \# mset_set A" from this assms have "set_mset M \ A" using mset_subset_eqD by fastforce moreover { fix x from \M \# mset_set A\ have "count M x \ count (mset_set A) x" by (simp add: mset_subset_eq_count) from this \finite A\ have "count M x \ 1" by (meson count_mset_set_leq le_trans) } ultimately show "set_mset M \ A \ (\x. count M x \ 1)" by simp qed lemma image_mset_fun_upd: assumes "x \# M" shows "image_mset (f(x := y)) M = image_mset f M" using assms by (induct M) auto subsection \Additions to Number Partitions Theory\ lemma Partition_diag: shows "Partition n n = 1" by (cases n) (auto simp only: Partition_diag Partition.simps(1)) subsection \Cardinality Theorems with Iverson Function\ definition iverson :: "bool \ nat" where "iverson b = (if b then 1 else 0)" lemma card_partition_on_size1_eq_iverson: assumes "finite A" shows "card {P. partition_on A P \ card P \ k \ (\X\P. card X = 1)} = iverson (card A \ k)" proof (cases "card A \ k") case True from this \finite A\ show ?thesis unfolding iverson_def using card_partition_on_size1_eq_1 by fastforce next case False from this \finite A\ show ?thesis unfolding iverson_def using card_partition_on_size1_eq_0 by fastforce qed lemma card_number_partitions_with_only_parts_1: "card {N. (\n. n\# N \ n = 1) \ number_partition n N \ size N \ x} = iverson (n \ x)" proof - show ?thesis proof cases assume "n \ x" from this show ?thesis using card_number_partitions_with_only_parts_1_eq_1 unfolding iverson_def by auto next assume "\ n \ x" from this show ?thesis using card_number_partitions_with_only_parts_1_eq_0 unfolding iverson_def by auto qed qed end diff --git a/thys/Twelvefold_Way/ROOT b/thys/Twelvefold_Way/ROOT --- a/thys/Twelvefold_Way/ROOT +++ b/thys/Twelvefold_Way/ROOT @@ -1,16 +1,17 @@ chapter AFP session Twelvefold_Way (AFP) = "HOL-Library" + options [timeout = 300] sessions "HOL-ex" "HOL-Eisbach" + "HOL-Combinatorics" Bell_Numbers_Spivey Card_Multisets Card_Number_Partitions Card_Partitions theories Twelvefold_Way document_files "root.bib" "root.tex"