src/HOL/Library/Groups_Big_Fun.thy
author wenzelm
Mon, 11 Jul 2016 10:43:27 +0200
changeset 63433 aa03b0487bf5
parent 63290 9ac558ab0906
child 63918 6bf55e6e0b75
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
     2
58881
b9556a055632 modernized header;
wenzelm
parents: 58437
diff changeset
     3
section \<open>Big sum and product over function bodies\<close>
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
     4
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
     5
theory Groups_Big_Fun
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
     6
imports
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
     7
  Main
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
     8
begin
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
     9
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    10
subsection \<open>Abstract product\<close>
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    11
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    12
locale comm_monoid_fun = comm_monoid
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    13
begin
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    14
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    15
definition G :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    16
where
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    17
  expand_set: "G g = comm_monoid_set.F f \<^bold>1 g {a. g a \<noteq> \<^bold>1}"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    18
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    19
interpretation F: comm_monoid_set f "\<^bold>1"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    20
  ..
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    21
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    22
lemma expand_superset:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    23
  assumes "finite A" and "{a. g a \<noteq> \<^bold>1} \<subseteq> A"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    24
  shows "G g = F.F g A"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    25
  apply (simp add: expand_set)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    26
  apply (rule F.same_carrierI [of A])
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    27
  apply (simp_all add: assms)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    28
  done
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    29
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    30
lemma conditionalize:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    31
  assumes "finite A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    32
  shows "F.F g A = G (\<lambda>a. if a \<in> A then g a else \<^bold>1)"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    33
  using assms
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    34
  apply (simp add: expand_set)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    35
  apply (rule F.same_carrierI [of A])
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    36
  apply auto
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    37
  done
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    38
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    39
lemma neutral [simp]:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    40
  "G (\<lambda>a. \<^bold>1) = \<^bold>1"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    41
  by (simp add: expand_set)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    42
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    43
lemma update [simp]:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    44
  assumes "finite {a. g a \<noteq> \<^bold>1}"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    45
  assumes "g a = \<^bold>1"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    46
  shows "G (g(a := b)) = b \<^bold>* G g"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    47
proof (cases "b = \<^bold>1")
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    48
  case True with \<open>g a = \<^bold>1\<close> show ?thesis
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    49
    by (simp add: expand_set) (rule F.cong, auto)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    50
next
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    51
  case False
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    52
  moreover have "{a'. a' \<noteq> a \<longrightarrow> g a' \<noteq> \<^bold>1} = insert a {a. g a \<noteq> \<^bold>1}"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    53
    by auto
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    54
  moreover from \<open>g a = \<^bold>1\<close> have "a \<notin> {a. g a \<noteq> \<^bold>1}"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    55
    by simp
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    56
  moreover have "F.F (\<lambda>a'. if a' = a then b else g a') {a. g a \<noteq> \<^bold>1} = F.F g {a. g a \<noteq> \<^bold>1}"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    57
    by (rule F.cong) (auto simp add: \<open>g a = \<^bold>1\<close>)
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    58
  ultimately show ?thesis using \<open>finite {a. g a \<noteq> \<^bold>1}\<close> by (simp add: expand_set)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    59
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    60
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    61
lemma infinite [simp]:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    62
  "\<not> finite {a. g a \<noteq> \<^bold>1} \<Longrightarrow> G g = \<^bold>1"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    63
  by (simp add: expand_set)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    64
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    65
lemma cong:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    66
  assumes "\<And>a. g a = h a"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    67
  shows "G g = G h"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    68
  using assms by (simp add: expand_set)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    69
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    70
lemma strong_cong [cong]:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    71
  assumes "\<And>a. g a = h a"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    72
  shows "G (\<lambda>a. g a) = G (\<lambda>a. h a)"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    73
  using assms by (fact cong)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    74
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    75
lemma not_neutral_obtains_not_neutral:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    76
  assumes "G g \<noteq> \<^bold>1"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    77
  obtains a where "g a \<noteq> \<^bold>1"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    78
  using assms by (auto elim: F.not_neutral_contains_not_neutral simp add: expand_set)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    79
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    80
lemma reindex_cong:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    81
  assumes "bij l"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    82
  assumes "g \<circ> l = h"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    83
  shows "G g = G h"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    84
proof -
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    85
  from assms have unfold: "h = g \<circ> l" by simp
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    86
  from \<open>bij l\<close> have "inj l" by (rule bij_is_inj)
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    87
  then have "inj_on l {a. h a \<noteq> \<^bold>1}" by (rule subset_inj_on) simp
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    88
  moreover from \<open>bij l\<close> have "{a. g a \<noteq> \<^bold>1} = l ` {a. h a \<noteq> \<^bold>1}"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    89
    by (auto simp add: image_Collect unfold elim: bij_pointE)
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    90
  moreover have "\<And>x. x \<in> {a. h a \<noteq> \<^bold>1} \<Longrightarrow> g (l x) = h x"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    91
    by (simp add: unfold)
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    92
  ultimately have "F.F g {a. g a \<noteq> \<^bold>1} = F.F h {a. h a \<noteq> \<^bold>1}"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    93
    by (rule F.reindex_cong)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    94
  then show ?thesis by (simp add: expand_set)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    95
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    96
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
    97
lemma distrib:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    98
  assumes "finite {a. g a \<noteq> \<^bold>1}" and "finite {a. h a \<noteq> \<^bold>1}"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
    99
  shows "G (\<lambda>a. g a \<^bold>* h a) = G g \<^bold>* G h"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   100
proof -
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   101
  from assms have "finite ({a. g a \<noteq> \<^bold>1} \<union> {a. h a \<noteq> \<^bold>1})" by simp
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   102
  moreover have "{a. g a \<^bold>* h a \<noteq> \<^bold>1} \<subseteq> {a. g a \<noteq> \<^bold>1} \<union> {a. h a \<noteq> \<^bold>1}"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   103
    by auto (drule sym, simp)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   104
  ultimately show ?thesis
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   105
    using assms
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   106
    by (simp add: expand_superset [of "{a. g a \<noteq> \<^bold>1} \<union> {a. h a \<noteq> \<^bold>1}"] F.distrib)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   107
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   108
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   109
lemma commute:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   110
  assumes "finite C"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   111
  assumes subset: "{a. \<exists>b. g a b \<noteq> \<^bold>1} \<times> {b. \<exists>a. g a b \<noteq> \<^bold>1} \<subseteq> C" (is "?A \<times> ?B \<subseteq> C")
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   112
  shows "G (\<lambda>a. G (g a)) = G (\<lambda>b. G (\<lambda>a. g a b))"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   113
proof -
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   114
  from \<open>finite C\<close> subset
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   115
    have "finite ({a. \<exists>b. g a b \<noteq> \<^bold>1} \<times> {b. \<exists>a. g a b \<noteq> \<^bold>1})"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   116
    by (rule rev_finite_subset)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   117
  then have fins:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   118
    "finite {b. \<exists>a. g a b \<noteq> \<^bold>1}" "finite {a. \<exists>b. g a b \<noteq> \<^bold>1}"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   119
    by (auto simp add: finite_cartesian_product_iff)
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   120
  have subsets: "\<And>a. {b. g a b \<noteq> \<^bold>1} \<subseteq> {b. \<exists>a. g a b \<noteq> \<^bold>1}"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   121
    "\<And>b. {a. g a b \<noteq> \<^bold>1} \<subseteq> {a. \<exists>b. g a b \<noteq> \<^bold>1}"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   122
    "{a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1} \<noteq> \<^bold>1} \<subseteq> {a. \<exists>b. g a b \<noteq> \<^bold>1}"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   123
    "{a. F.F (\<lambda>aa. g aa a) {a. \<exists>b. g a b \<noteq> \<^bold>1} \<noteq> \<^bold>1} \<subseteq> {b. \<exists>a. g a b \<noteq> \<^bold>1}"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   124
    by (auto elim: F.not_neutral_contains_not_neutral)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   125
  from F.commute have
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   126
    "F.F (\<lambda>a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1}) {a. \<exists>b. g a b \<noteq> \<^bold>1} =
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   127
      F.F (\<lambda>b. F.F (\<lambda>a. g a b) {a. \<exists>b. g a b \<noteq> \<^bold>1}) {b. \<exists>a. g a b \<noteq> \<^bold>1}" .
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   128
  with subsets fins have "G (\<lambda>a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1}) =
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   129
    G (\<lambda>b. F.F (\<lambda>a. g a b) {a. \<exists>b. g a b \<noteq> \<^bold>1})"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   130
    by (auto simp add: expand_superset [of "{b. \<exists>a. g a b \<noteq> \<^bold>1}"]
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   131
      expand_superset [of "{a. \<exists>b. g a b \<noteq> \<^bold>1}"])
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   132
  with subsets fins show ?thesis
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   133
    by (auto simp add: expand_superset [of "{b. \<exists>a. g a b \<noteq> \<^bold>1}"]
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   134
      expand_superset [of "{a. \<exists>b. g a b \<noteq> \<^bold>1}"])
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   135
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   136
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   137
lemma cartesian_product:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   138
  assumes "finite C"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   139
  assumes subset: "{a. \<exists>b. g a b \<noteq> \<^bold>1} \<times> {b. \<exists>a. g a b \<noteq> \<^bold>1} \<subseteq> C" (is "?A \<times> ?B \<subseteq> C")
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   140
  shows "G (\<lambda>a. G (g a)) = G (\<lambda>(a, b). g a b)"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   141
proof -
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   142
  from subset \<open>finite C\<close> have fin_prod: "finite (?A \<times> ?B)"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   143
    by (rule finite_subset)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   144
  from fin_prod have "finite ?A" and "finite ?B"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   145
    by (auto simp add: finite_cartesian_product_iff)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   146
  have *: "G (\<lambda>a. G (g a)) =
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   147
    (F.F (\<lambda>a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1}) {a. \<exists>b. g a b \<noteq> \<^bold>1})"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   148
    apply (subst expand_superset [of "?B"])
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   149
    apply (rule \<open>finite ?B\<close>)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   150
    apply auto
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   151
    apply (subst expand_superset [of "?A"])
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   152
    apply (rule \<open>finite ?A\<close>)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   153
    apply auto
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   154
    apply (erule F.not_neutral_contains_not_neutral)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   155
    apply auto
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   156
    done
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   157
  have "{p. (case p of (a, b) \<Rightarrow> g a b) \<noteq> \<^bold>1} \<subseteq> ?A \<times> ?B"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   158
    by auto
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   159
  with subset have **: "{p. (case p of (a, b) \<Rightarrow> g a b) \<noteq> \<^bold>1} \<subseteq> C"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   160
    by blast
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   161
  show ?thesis
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   162
    apply (simp add: *)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   163
    apply (simp add: F.cartesian_product)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   164
    apply (subst expand_superset [of C])
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   165
    apply (rule \<open>finite C\<close>)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   166
    apply (simp_all add: **)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   167
    apply (rule F.same_carrierI [of C])
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   168
    apply (rule \<open>finite C\<close>)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   169
    apply (simp_all add: subset)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   170
    apply auto
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   171
    done
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   172
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   173
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   174
lemma cartesian_product2:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   175
  assumes fin: "finite D"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   176
  assumes subset: "{(a, b). \<exists>c. g a b c \<noteq> \<^bold>1} \<times> {c. \<exists>a b. g a b c \<noteq> \<^bold>1} \<subseteq> D" (is "?AB \<times> ?C \<subseteq> D")
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   177
  shows "G (\<lambda>(a, b). G (g a b)) = G (\<lambda>(a, b, c). g a b c)"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   178
proof -
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   179
  have bij: "bij (\<lambda>(a, b, c). ((a, b), c))"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   180
    by (auto intro!: bijI injI simp add: image_def)
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   181
  have "{p. \<exists>c. g (fst p) (snd p) c \<noteq> \<^bold>1} \<times> {c. \<exists>p. g (fst p) (snd p) c \<noteq> \<^bold>1} \<subseteq> D"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61378
diff changeset
   182
    by auto (insert subset, blast)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   183
  with fin have "G (\<lambda>p. G (g (fst p) (snd p))) = G (\<lambda>(p, c). g (fst p) (snd p) c)"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   184
    by (rule cartesian_product)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   185
  then have "G (\<lambda>(a, b). G (g a b)) = G (\<lambda>((a, b), c). g a b c)"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   186
    by (auto simp add: split_def)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   187
  also have "G (\<lambda>((a, b), c). g a b c) = G (\<lambda>(a, b, c). g a b c)"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   188
    using bij by (rule reindex_cong [of "\<lambda>(a, b, c). ((a, b), c)"]) (simp add: fun_eq_iff)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   189
  finally show ?thesis .
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   190
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   191
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   192
lemma delta [simp]:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   193
  "G (\<lambda>b. if b = a then g b else \<^bold>1) = g a"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   194
proof -
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   195
  have "{b. (if b = a then g b else \<^bold>1) \<noteq> \<^bold>1} \<subseteq> {a}" by auto
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   196
  then show ?thesis by (simp add: expand_superset [of "{a}"])
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   197
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   198
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   199
lemma delta' [simp]:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   200
  "G (\<lambda>b. if a = b then g b else \<^bold>1) = g a"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   201
proof -
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   202
  have "(\<lambda>b. if a = b then g b else \<^bold>1) = (\<lambda>b. if b = a then g b else \<^bold>1)"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   203
    by (simp add: fun_eq_iff)
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61955
diff changeset
   204
  then have "G (\<lambda>b. if a = b then g b else \<^bold>1) = G (\<lambda>b. if b = a then g b else \<^bold>1)"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   205
    by (simp cong del: strong_cong)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   206
  then show ?thesis by simp
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   207
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   208
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   209
end
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   210
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   211
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   212
subsection \<open>Concrete sum\<close>
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   213
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   214
context comm_monoid_add
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   215
begin
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   216
61776
57bb7da5c867 modernized
haftmann
parents: 61671
diff changeset
   217
sublocale Sum_any: comm_monoid_fun plus 0
63433
wenzelm
parents: 63290
diff changeset
   218
  defines Sum_any = Sum_any.G
wenzelm
parents: 63290
diff changeset
   219
  rewrites "comm_monoid_set.F plus 0 = setsum"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   220
proof -
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   221
  show "comm_monoid_fun plus 0" ..
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61424
diff changeset
   222
  then interpret Sum_any: comm_monoid_fun plus 0 .
61776
57bb7da5c867 modernized
haftmann
parents: 61671
diff changeset
   223
  from setsum_def show "comm_monoid_set.F plus 0 = setsum" by (auto intro: sym)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   224
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   225
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   226
end
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   227
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61776
diff changeset
   228
syntax (ASCII)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   229
  "_Sum_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"    ("(3SUM _. _)" [0, 10] 10)
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61776
diff changeset
   230
syntax
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   231
  "_Sum_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"    ("(3\<Sum>_. _)" [0, 10] 10)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   232
translations
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61776
diff changeset
   233
  "\<Sum>a. b" \<rightleftharpoons> "CONST Sum_any (\<lambda>a. b)"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   234
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   235
lemma Sum_any_left_distrib:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   236
  fixes r :: "'a :: semiring_0"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   237
  assumes "finite {a. g a \<noteq> 0}"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   238
  shows "Sum_any g * r = (\<Sum>n. g n * r)"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   239
proof -
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   240
  note assms
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   241
  moreover have "{a. g a * r \<noteq> 0} \<subseteq> {a. g a \<noteq> 0}" by auto
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   242
  ultimately show ?thesis
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   243
    by (simp add: setsum_left_distrib Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   244
qed  
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   245
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   246
lemma Sum_any_right_distrib:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   247
  fixes r :: "'a :: semiring_0"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   248
  assumes "finite {a. g a \<noteq> 0}"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   249
  shows "r * Sum_any g = (\<Sum>n. r * g n)"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   250
proof -
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   251
  note assms
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   252
  moreover have "{a. r * g a \<noteq> 0} \<subseteq> {a. g a \<noteq> 0}" by auto
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   253
  ultimately show ?thesis
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   254
    by (simp add: setsum_right_distrib Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   255
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   256
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   257
lemma Sum_any_product:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   258
  fixes f g :: "'b \<Rightarrow> 'a::semiring_0"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   259
  assumes "finite {a. f a \<noteq> 0}" and "finite {b. g b \<noteq> 0}"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   260
  shows "Sum_any f * Sum_any g = (\<Sum>a. \<Sum>b. f a * g b)"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   261
proof -
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   262
  have subset_f: "{a. (\<Sum>b. f a * g b) \<noteq> 0} \<subseteq> {a. f a \<noteq> 0}"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   263
    by rule (simp, rule, auto)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   264
  moreover have subset_g: "\<And>a. {b. f a * g b \<noteq> 0} \<subseteq> {b. g b \<noteq> 0}"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   265
    by rule (simp, rule, auto)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   266
  ultimately show ?thesis using assms
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   267
    by (auto simp add: Sum_any.expand_set [of f] Sum_any.expand_set [of g]
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   268
      Sum_any.expand_superset [of "{a. f a \<noteq> 0}"] Sum_any.expand_superset [of "{b. g b \<noteq> 0}"]
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   269
      setsum_product)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   270
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   271
58437
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   272
lemma Sum_any_eq_zero_iff [simp]: 
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   273
  fixes f :: "'a \<Rightarrow> nat"
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   274
  assumes "finite {a. f a \<noteq> 0}"
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   275
  shows "Sum_any f = 0 \<longleftrightarrow> f = (\<lambda>_. 0)"
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   276
  using assms by (simp add: Sum_any.expand_set fun_eq_iff)
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   277
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   278
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   279
subsection \<open>Concrete product\<close>
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   280
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   281
context comm_monoid_mult
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   282
begin
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   283
61776
57bb7da5c867 modernized
haftmann
parents: 61671
diff changeset
   284
sublocale Prod_any: comm_monoid_fun times 1
63433
wenzelm
parents: 63290
diff changeset
   285
  defines Prod_any = Prod_any.G
wenzelm
parents: 63290
diff changeset
   286
  rewrites "comm_monoid_set.F times 1 = setprod"
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   287
proof -
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   288
  show "comm_monoid_fun times 1" ..
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61424
diff changeset
   289
  then interpret Prod_any: comm_monoid_fun times 1 .
61776
57bb7da5c867 modernized
haftmann
parents: 61671
diff changeset
   290
  from setprod_def show "comm_monoid_set.F times 1 = setprod" by (auto intro: sym)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   291
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   292
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   293
end
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   294
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61776
diff changeset
   295
syntax (ASCII)
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61776
diff changeset
   296
  "_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3PROD _. _)" [0, 10] 10)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   297
syntax
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61776
diff changeset
   298
  "_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3\<Prod>_. _)" [0, 10] 10)
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   299
translations
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   300
  "\<Prod>a. b" == "CONST Prod_any (\<lambda>a. b)"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   301
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   302
lemma Prod_any_zero:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   303
  fixes f :: "'b \<Rightarrow> 'a :: comm_semiring_1"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   304
  assumes "finite {a. f a \<noteq> 1}"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   305
  assumes "f a = 0"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   306
  shows "(\<Prod>a. f a) = 0"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   307
proof -
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   308
  from \<open>f a = 0\<close> have "f a \<noteq> 1" by simp
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   309
  with \<open>f a = 0\<close> have "\<exists>a. f a \<noteq> 1 \<and> f a = 0" by blast
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   310
  with \<open>finite {a. f a \<noteq> 1}\<close> show ?thesis
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   311
    by (simp add: Prod_any.expand_set setprod_zero)
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   312
qed
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   313
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   314
lemma Prod_any_not_zero:
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   315
  fixes f :: "'b \<Rightarrow> 'a :: comm_semiring_1"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   316
  assumes "finite {a. f a \<noteq> 1}"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   317
  assumes "(\<Prod>a. f a) \<noteq> 0"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   318
  shows "f a \<noteq> 0"
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   319
  using assms Prod_any_zero [of f] by blast
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   320
58437
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   321
lemma power_Sum_any:
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   322
  assumes "finite {a. f a \<noteq> 0}"
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   323
  shows "c ^ (\<Sum>a. f a) = (\<Prod>a. c ^ f a)"
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   324
proof -
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   325
  have "{a. c ^ f a \<noteq> 1} \<subseteq> {a. f a \<noteq> 0}"
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   326
    by (auto intro: ccontr)
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   327
  with assms show ?thesis
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   328
    by (simp add: Sum_any.expand_set Prod_any.expand_superset power_setsum)
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   329
qed
8d124c73c37a added lemmas
haftmann
parents: 58197
diff changeset
   330
58197
4fd7f47ead6c theory about sum and product on function bodies
haftmann
parents:
diff changeset
   331
end