# HG changeset patch # User haftmann # Date 1410027156 -7200 # Node ID 4fd7f47ead6ce1e8b8b9cc310ea696b60481be20 # Parent 1b3fbfb859804c250100b817bf82d1bb35ab04b9 theory about sum and product on function bodies diff -r 1b3fbfb85980 -r 4fd7f47ead6c src/HOL/Library/Groups_Big_Fun.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Groups_Big_Fun.thy Sat Sep 06 20:12:36 2014 +0200 @@ -0,0 +1,340 @@ +(* Author: Florian Haftmann, TU Muenchen *) + +header \Big sum and product over function bodies\ + +theory Groups_Big_Fun +imports + Main + "~~/src/Tools/Permanent_Interpretation" +begin + +subsection \Abstract product\ + +no_notation times (infixl "*" 70) +no_notation Groups.one ("1") + +locale comm_monoid_fun = comm_monoid +begin + +definition G :: "('b \ 'a) \ 'a" +where + expand_set: "G g = comm_monoid_set.F f 1 g {a. g a \ 1}" + +interpretation F!: comm_monoid_set f 1 + .. + +lemma expand_superset: + assumes "finite A" and "{a. g a \ 1} \ A" + shows "G g = F.F g A" + apply (simp add: expand_set) + apply (rule F.same_carrierI [of A]) + apply (simp_all add: assms) + done + +lemma conditionalize: + assumes "finite A" + shows "F.F g A = G (\a. if a \ A then g a else 1)" + using assms + apply (simp add: expand_set) + apply (rule F.same_carrierI [of A]) + apply auto + done + +lemma neutral [simp]: + "G (\a. 1) = 1" + by (simp add: expand_set) + +lemma update [simp]: + assumes "finite {a. g a \ 1}" + assumes "g a = 1" + shows "G (g(a := b)) = b * G g" +proof (cases "b = 1") + case True with `g a = 1` show ?thesis + by (simp add: expand_set) (rule F.cong, auto) +next + case False + moreover have "{a'. a' \ a \ g a' \ 1} = insert a {a. g a \ 1}" + by auto + moreover from `g a = 1` have "a \ {a. g a \ 1}" + by simp + moreover have "F.F (\a'. if a' = a then b else g a') {a. g a \ 1} = F.F g {a. g a \ 1}" + by (rule F.cong) (auto simp add: `g a = 1`) + ultimately show ?thesis using `finite {a. g a \ 1}` by (simp add: expand_set) +qed + +lemma infinite [simp]: + "\ finite {a. g a \ 1} \ G g = 1" + by (simp add: expand_set) + +lemma cong: + assumes "\a. g a = h a" + shows "G g = G h" + using assms by (simp add: expand_set) + +lemma strong_cong [cong]: + assumes "\a. g a = h a" + shows "G (\a. g a) = G (\a. h a)" + using assms by (fact cong) + +lemma not_neutral_obtains_not_neutral: + assumes "G g \ 1" + obtains a where "g a \ 1" + using assms by (auto elim: F.not_neutral_contains_not_neutral simp add: expand_set) + +lemma reindex_cong: + assumes "bij l" + assumes "g \ l = h" + shows "G g = G h" +proof - + from assms have unfold: "h = g \ l" by simp + from `bij l` have "inj l" by (rule bij_is_inj) + then have "inj_on l {a. h a \ 1}" by (rule subset_inj_on) simp + moreover from `bij l` have "{a. g a \ 1} = l ` {a. h a \ 1}" + by (auto simp add: image_Collect unfold elim: bij_pointE) + moreover have "\x. x \ {a. h a \ 1} \ g (l x) = h x" + by (simp add: unfold) + ultimately have "F.F g {a. g a \ 1} = F.F h {a. h a \ 1}" + by (rule F.reindex_cong) + then show ?thesis by (simp add: expand_set) +qed + +lemma distrib: + assumes "finite {a. g a \ 1}" and "finite {a. h a \ 1}" + shows "G (\a. g a * h a) = G g * G h" +proof - + from assms have "finite ({a. g a \ 1} \ {a. h a \ 1})" by simp + moreover have "{a. g a * h a \ 1} \ {a. g a \ 1} \ {a. h a \ 1}" + by auto (drule sym, simp) + ultimately show ?thesis + using assms + by (simp add: expand_superset [of "{a. g a \ 1} \ {a. h a \ 1}"] F.distrib) +qed + +lemma commute: + assumes "finite C" + assumes subset: "{a. \b. g a b \ 1} \ {b. \a. g a b \ 1} \ C" (is "?A \ ?B \ C") + shows "G (\a. G (g a)) = G (\b. G (\a. g a b))" +proof - + from `finite C` subset + have "finite ({a. \b. g a b \ 1} \ {b. \a. g a b \ 1})" + by (rule rev_finite_subset) + then have fins: + "finite {b. \a. g a b \ 1}" "finite {a. \b. g a b \ 1}" + by (auto simp add: finite_cartesian_product_iff) + have subsets: "\a. {b. g a b \ 1} \ {b. \a. g a b \ 1}" + "\b. {a. g a b \ 1} \ {a. \b. g a b \ 1}" + "{a. F.F (g a) {b. \a. g a b \ 1} \ 1} \ {a. \b. g a b \ 1}" + "{a. F.F (\aa. g aa a) {a. \b. g a b \ 1} \ 1} \ {b. \a. g a b \ 1}" + by (auto elim: F.not_neutral_contains_not_neutral) + from F.commute have + "F.F (\a. F.F (g a) {b. \a. g a b \ 1}) {a. \b. g a b \ 1} = + F.F (\b. F.F (\a. g a b) {a. \b. g a b \ 1}) {b. \a. g a b \ 1}" . + with subsets fins have "G (\a. F.F (g a) {b. \a. g a b \ 1}) = + G (\b. F.F (\a. g a b) {a. \b. g a b \ 1})" + by (auto simp add: expand_superset [of "{b. \a. g a b \ 1}"] + expand_superset [of "{a. \b. g a b \ 1}"]) + with subsets fins show ?thesis + by (auto simp add: expand_superset [of "{b. \a. g a b \ 1}"] + expand_superset [of "{a. \b. g a b \ 1}"]) +qed + +lemma cartesian_product: + assumes "finite C" + assumes subset: "{a. \b. g a b \ 1} \ {b. \a. g a b \ 1} \ C" (is "?A \ ?B \ C") + shows "G (\a. G (g a)) = G (\(a, b). g a b)" +proof - + from subset `finite C` have fin_prod: "finite (?A \ ?B)" + by (rule finite_subset) + from fin_prod have "finite ?A" and "finite ?B" + by (auto simp add: finite_cartesian_product_iff) + have *: "G (\a. G (g a)) = + (F.F (\a. F.F (g a) {b. \a. g a b \ 1}) {a. \b. g a b \ 1})" + apply (subst expand_superset [of "?B"]) + apply (rule `finite ?B`) + apply auto + apply (subst expand_superset [of "?A"]) + apply (rule `finite ?A`) + apply auto + apply (erule F.not_neutral_contains_not_neutral) + apply auto + done + have "{p. (case p of (a, b) \ g a b) \ 1} \ ?A \ ?B" + by auto + with subset have **: "{p. (case p of (a, b) \ g a b) \ 1} \ C" + by blast + show ?thesis + apply (simp add: *) + apply (simp add: F.cartesian_product) + apply (subst expand_superset [of C]) + apply (rule `finite C`) + apply (simp_all add: **) + apply (rule F.same_carrierI [of C]) + apply (rule `finite C`) + apply (simp_all add: subset) + apply auto + done +qed + +lemma cartesian_product2: + assumes fin: "finite D" + assumes subset: "{(a, b). \c. g a b c \ 1} \ {c. \a b. g a b c \ 1} \ D" (is "?AB \ ?C \ D") + shows "G (\(a, b). G (g a b)) = G (\(a, b, c). g a b c)" +proof - + have bij: "bij (\(a, b, c). ((a, b), c))" + by (auto intro!: bijI injI simp add: image_def) + have "{p. \c. g (fst p) (snd p) c \ 1} \ {c. \p. g (fst p) (snd p) c \ 1} \ D" + by auto (insert subset, auto) + with fin have "G (\p. G (g (fst p) (snd p))) = G (\(p, c). g (fst p) (snd p) c)" + by (rule cartesian_product) + then have "G (\(a, b). G (g a b)) = G (\((a, b), c). g a b c)" + by (auto simp add: split_def) + also have "G (\((a, b), c). g a b c) = G (\(a, b, c). g a b c)" + using bij by (rule reindex_cong [of "\(a, b, c). ((a, b), c)"]) (simp add: fun_eq_iff) + finally show ?thesis . +qed + +lemma delta [simp]: + "G (\b. if b = a then g b else 1) = g a" +proof - + have "{b. (if b = a then g b else 1) \ 1} \ {a}" by auto + then show ?thesis by (simp add: expand_superset [of "{a}"]) +qed + +lemma delta' [simp]: + "G (\b. if a = b then g b else 1) = g a" +proof - + have "(\b. if a = b then g b else 1) = (\b. if b = a then g b else 1)" + by (simp add: fun_eq_iff) + then have "G (\b. if a = b then g b else 1) = G (\b. if b = a then g b else 1)" + by (simp cong del: strong_cong) + then show ?thesis by simp +qed + +end + +notation times (infixl "*" 70) +notation Groups.one ("1") + + +subsection \Concrete sum\ + +context comm_monoid_add +begin + +definition Sum_any :: "('b \ 'a) \ 'a" +where + "Sum_any = comm_monoid_fun.G plus 0" + +permanent_interpretation Sum_any!: comm_monoid_fun plus 0 +where + "comm_monoid_fun.G plus 0 = Sum_any" and + "comm_monoid_set.F plus 0 = setsum" +proof - + show "comm_monoid_fun plus 0" .. + then interpret Sum_any!: comm_monoid_fun plus 0 . + from Sum_any_def show "comm_monoid_fun.G plus 0 = Sum_any" by rule + from setsum_def show "comm_monoid_set.F plus 0 = setsum" by rule +qed + +end + +syntax + "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" ("(3SUM _. _)" [0, 10] 10) +syntax (xsymbols) + "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" ("(3\_. _)" [0, 10] 10) +syntax (HTML output) + "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" ("(3\_. _)" [0, 10] 10) + +translations + "\a. b" == "CONST Sum_any (\a. b)" + +lemma Sum_any_left_distrib: + fixes r :: "'a :: semiring_0" + assumes "finite {a. g a \ 0}" + shows "Sum_any g * r = (\n. g n * r)" +proof - + note assms + moreover have "{a. g a * r \ 0} \ {a. g a \ 0}" by auto + ultimately show ?thesis + by (simp add: setsum_left_distrib Sum_any.expand_superset [of "{a. g a \ 0}"]) +qed + +lemma Sum_any_right_distrib: + fixes r :: "'a :: semiring_0" + assumes "finite {a. g a \ 0}" + shows "r * Sum_any g = (\n. r * g n)" +proof - + note assms + moreover have "{a. r * g a \ 0} \ {a. g a \ 0}" by auto + ultimately show ?thesis + by (simp add: setsum_right_distrib Sum_any.expand_superset [of "{a. g a \ 0}"]) +qed + +lemma Sum_any_product: + fixes f g :: "'b \ 'a::semiring_0" + assumes "finite {a. f a \ 0}" and "finite {b. g b \ 0}" + shows "Sum_any f * Sum_any g = (\a. \b. f a * g b)" +proof - + have subset_f: "{a. (\b. f a * g b) \ 0} \ {a. f a \ 0}" + by rule (simp, rule, auto) + moreover have subset_g: "\a. {b. f a * g b \ 0} \ {b. g b \ 0}" + by rule (simp, rule, auto) + ultimately show ?thesis using assms + by (auto simp add: Sum_any.expand_set [of f] Sum_any.expand_set [of g] + Sum_any.expand_superset [of "{a. f a \ 0}"] Sum_any.expand_superset [of "{b. g b \ 0}"] + setsum_product) +qed + + +subsection \Concrete product\ + +context comm_monoid_mult +begin + +definition Prod_any :: "('b \ 'a) \ 'a" +where + "Prod_any = comm_monoid_fun.G times 1" + +permanent_interpretation Prod_any!: comm_monoid_fun times 1 +where + "comm_monoid_fun.G times 1 = Prod_any" and + "comm_monoid_set.F times 1 = setprod" +proof - + show "comm_monoid_fun times 1" .. + then interpret Prod_any!: comm_monoid_fun times 1 . + from Prod_any_def show "comm_monoid_fun.G times 1 = Prod_any" by rule + from setprod_def show "comm_monoid_set.F times 1 = setprod" by rule +qed + +end + +syntax + "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3PROD _. _)" [0, 10] 10) +syntax (xsymbols) + "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3\_. _)" [0, 10] 10) +syntax (HTML output) + "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3\_. _)" [0, 10] 10) + +translations + "\a. b" == "CONST Prod_any (\a. b)" + +lemma Prod_any_zero: + fixes f :: "'b \ 'a :: comm_semiring_1" + assumes "finite {a. f a \ 1}" + assumes "f a = 0" + shows "(\a. f a) = 0" +proof - + from `f a = 0` have "f a \ 1" by simp + with `f a = 0` have "\a. f a \ 1 \ f a = 0" by blast + with `finite {a. f a \ 1}` show ?thesis + by (simp add: Prod_any.expand_set setprod_zero) +qed + +lemma Prod_any_not_zero: + fixes f :: "'b \ 'a :: comm_semiring_1" + assumes "finite {a. f a \ 1}" + assumes "(\a. f a) \ 0" + shows "f a \ 0" + using assms Prod_any_zero [of f] by blast + +end diff -r 1b3fbfb85980 -r 4fd7f47ead6c src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Sat Sep 06 20:12:34 2014 +0200 +++ b/src/HOL/Library/Library.thy Sat Sep 06 20:12:36 2014 +0200 @@ -27,6 +27,7 @@ Function_Growth Fundamental_Theorem_Algebra Fun_Lexorder + Groups_Big_Fun Indicator_Function Infinite_Set Inner_Product