author | wenzelm |
Fri, 11 Jan 2013 22:01:49 +0100 | |
changeset 50836 | c95af99e003b |
parent 46822 | 95f1e700b712 |
child 60770 | 240563fbf41d |
permissions | -rw-r--r-- |
12610 | 1 |
(* Title: ZF/Induct/FoldSet.thy |
12089
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
2 |
Author: Sidi O Ehmety, Cambridge University Computer Laboratory |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
3 |
|
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
4 |
A "fold" functional for finite sets. For n non-negative we have |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
5 |
fold f e {x1,...,xn} = f x1 (... (f xn e)) where f is at |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
6 |
least left-commutative. |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
7 |
*) |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
8 |
|
16417 | 9 |
theory FoldSet imports Main begin |
12089
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
10 |
|
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
11 |
consts fold_set :: "[i, i, [i,i]=>i, i] => i" |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
12 |
|
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
13 |
inductive |
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
14 |
domains "fold_set(A, B, f,e)" \<subseteq> "Fin(A)*B" |
14071 | 15 |
intros |
16 |
emptyI: "e\<in>B ==> <0, e>\<in>fold_set(A, B, f,e)" |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
17 |
consI: "[| x\<in>A; x \<notin>C; <C,y> \<in> fold_set(A, B,f,e); f(x,y):B |] |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
18 |
==> <cons(x,C), f(x,y)>\<in>fold_set(A, B, f, e)" |
14071 | 19 |
type_intros Fin.intros |
12089
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
20 |
|
24893 | 21 |
definition |
22 |
fold :: "[i, [i,i]=>i, i, i] => i" ("fold[_]'(_,_,_')") where |
|
14071 | 23 |
"fold[B](f,e, A) == THE x. <A, x>\<in>fold_set(A, B, f,e)" |
12089
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
24 |
|
24893 | 25 |
definition |
26 |
setsum :: "[i=>i, i] => i" where |
|
14071 | 27 |
"setsum(g, C) == if Finite(C) then |
28 |
fold[int](%x y. g(x) $+ y, #0, C) else #0" |
|
29 |
||
30 |
(** foldSet **) |
|
31 |
||
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
32 |
inductive_cases empty_fold_setE: "<0, x> \<in> fold_set(A, B, f,e)" |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
33 |
inductive_cases cons_fold_setE: "<cons(x,C), y> \<in> fold_set(A, B, f,e)" |
14071 | 34 |
|
35 |
(* add-hoc lemmas *) |
|
36 |
||
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
37 |
lemma cons_lemma1: "[| x\<notin>C; x\<notin>B |] ==> cons(x,B)=cons(x,C) \<longleftrightarrow> B = C" |
14071 | 38 |
by (auto elim: equalityE) |
39 |
||
40 |
lemma cons_lemma2: "[| cons(x, B)=cons(y, C); x\<noteq>y; x\<notin>B; y\<notin>C |] |
|
41 |
==> B - {y} = C-{x} & x\<in>C & y\<in>B" |
|
42 |
apply (auto elim: equalityE) |
|
43 |
done |
|
44 |
||
45 |
(* fold_set monotonicity *) |
|
46 |
lemma fold_set_mono_lemma: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
47 |
"<C, x> \<in> fold_set(A, B, f, e) |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
48 |
==> \<forall>D. A<=D \<longrightarrow> <C, x> \<in> fold_set(D, B, f, e)" |
14071 | 49 |
apply (erule fold_set.induct) |
50 |
apply (auto intro: fold_set.intros) |
|
51 |
done |
|
52 |
||
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
53 |
lemma fold_set_mono: " C<=A ==> fold_set(C, B, f, e) \<subseteq> fold_set(A, B, f, e)" |
14071 | 54 |
apply clarify |
55 |
apply (frule fold_set.dom_subset [THEN subsetD], clarify) |
|
56 |
apply (auto dest: fold_set_mono_lemma) |
|
57 |
done |
|
58 |
||
59 |
lemma fold_set_lemma: |
|
60 |
"<C, x>\<in>fold_set(A, B, f, e) ==> <C, x>\<in>fold_set(C, B, f, e) & C<=A" |
|
61 |
apply (erule fold_set.induct) |
|
62 |
apply (auto intro!: fold_set.intros intro: fold_set_mono [THEN subsetD]) |
|
63 |
done |
|
64 |
||
65 |
(* Proving that fold_set is deterministic *) |
|
66 |
lemma Diff1_fold_set: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
67 |
"[| <C-{x},y> \<in> fold_set(A, B, f,e); x\<in>C; x\<in>A; f(x, y):B |] |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
68 |
==> <C, f(x, y)> \<in> fold_set(A, B, f, e)" |
14071 | 69 |
apply (frule fold_set.dom_subset [THEN subsetD]) |
70 |
apply (erule cons_Diff [THEN subst], rule fold_set.intros, auto) |
|
71 |
done |
|
72 |
||
73 |
||
74 |
locale fold_typing = |
|
75 |
fixes A and B and e and f |
|
76 |
assumes ftype [intro,simp]: "[|x \<in> A; y \<in> B|] ==> f(x,y) \<in> B" |
|
77 |
and etype [intro,simp]: "e \<in> B" |
|
78 |
and fcomm: "[|x \<in> A; y \<in> A; z \<in> B|] ==> f(x, f(y, z))=f(y, f(x, z))" |
|
79 |
||
80 |
||
81 |
lemma (in fold_typing) Fin_imp_fold_set: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
82 |
"C\<in>Fin(A) ==> (\<exists>x. <C, x> \<in> fold_set(A, B, f,e))" |
14071 | 83 |
apply (erule Fin_induct) |
84 |
apply (auto dest: fold_set.dom_subset [THEN subsetD] |
|
85 |
intro: fold_set.intros etype ftype) |
|
86 |
done |
|
87 |
||
88 |
lemma Diff_sing_imp: |
|
89 |
"[|C - {b} = D - {a}; a \<noteq> b; b \<in> C|] ==> C = cons(b,D) - {a}" |
|
90 |
by (blast elim: equalityE) |
|
91 |
||
92 |
lemma (in fold_typing) fold_set_determ_lemma [rule_format]: |
|
93 |
"n\<in>nat |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
94 |
==> \<forall>C. |C|<n \<longrightarrow> |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
95 |
(\<forall>x. <C, x> \<in> fold_set(A, B, f,e)\<longrightarrow> |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
96 |
(\<forall>y. <C, y> \<in> fold_set(A, B, f,e) \<longrightarrow> y=x))" |
14071 | 97 |
apply (erule nat_induct) |
98 |
apply (auto simp add: le_iff) |
|
99 |
apply (erule fold_set.cases) |
|
100 |
apply (force elim!: empty_fold_setE) |
|
101 |
apply (erule fold_set.cases) |
|
102 |
apply (force elim!: empty_fold_setE, clarify) |
|
103 |
(*force simplification of "|C| < |cons(...)|"*) |
|
104 |
apply (frule_tac a = Ca in fold_set.dom_subset [THEN subsetD, THEN SigmaD1]) |
|
105 |
apply (frule_tac a = Cb in fold_set.dom_subset [THEN subsetD, THEN SigmaD1]) |
|
106 |
apply (simp add: Fin_into_Finite [THEN Finite_imp_cardinal_cons]) |
|
107 |
apply (case_tac "x=xb", auto) |
|
108 |
apply (simp add: cons_lemma1, blast) |
|
109 |
txt{*case @{term "x\<noteq>xb"}*} |
|
110 |
apply (drule cons_lemma2, safe) |
|
111 |
apply (frule Diff_sing_imp, assumption+) |
|
112 |
txt{** LEVEL 17*} |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
113 |
apply (subgoal_tac "|Ca| \<le> |Cb|") |
14071 | 114 |
prefer 2 |
115 |
apply (rule succ_le_imp_le) |
|
116 |
apply (simp add: Fin_into_Finite Finite_imp_succ_cardinal_Diff |
|
117 |
Fin_into_Finite [THEN Finite_imp_cardinal_cons]) |
|
118 |
apply (rule_tac C1 = "Ca-{xb}" in Fin_imp_fold_set [THEN exE]) |
|
119 |
apply (blast intro: Diff_subset [THEN Fin_subset]) |
|
120 |
txt{** LEVEL 24 **} |
|
121 |
apply (frule Diff1_fold_set, blast, blast) |
|
122 |
apply (blast dest!: ftype fold_set.dom_subset [THEN subsetD]) |
|
123 |
apply (subgoal_tac "ya = f(xb,xa) ") |
|
124 |
prefer 2 apply (blast del: equalityCE) |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
125 |
apply (subgoal_tac "<Cb-{x}, xa> \<in> fold_set(A,B,f,e)") |
14071 | 126 |
prefer 2 apply simp |
127 |
apply (subgoal_tac "yb = f (x, xa) ") |
|
128 |
apply (drule_tac [2] C = Cb in Diff1_fold_set, simp_all) |
|
129 |
apply (blast intro: fcomm dest!: fold_set.dom_subset [THEN subsetD]) |
|
130 |
apply (blast intro: ftype dest!: fold_set.dom_subset [THEN subsetD], blast) |
|
131 |
done |
|
132 |
||
133 |
lemma (in fold_typing) fold_set_determ: |
|
134 |
"[| <C, x>\<in>fold_set(A, B, f, e); |
|
135 |
<C, y>\<in>fold_set(A, B, f, e)|] ==> y=x" |
|
136 |
apply (frule fold_set.dom_subset [THEN subsetD], clarify) |
|
137 |
apply (drule Fin_into_Finite) |
|
138 |
apply (unfold Finite_def, clarify) |
|
139 |
apply (rule_tac n = "succ (n)" in fold_set_determ_lemma) |
|
140 |
apply (auto intro: eqpoll_imp_lepoll [THEN lepoll_cardinal_le]) |
|
141 |
done |
|
142 |
||
143 |
(** The fold function **) |
|
144 |
||
145 |
lemma (in fold_typing) fold_equality: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
146 |
"<C,y> \<in> fold_set(A,B,f,e) ==> fold[B](f,e,C) = y" |
14071 | 147 |
apply (unfold fold_def) |
148 |
apply (frule fold_set.dom_subset [THEN subsetD], clarify) |
|
149 |
apply (rule the_equality) |
|
150 |
apply (rule_tac [2] A=C in fold_typing.fold_set_determ) |
|
151 |
apply (force dest: fold_set_lemma) |
|
152 |
apply (auto dest: fold_set_lemma) |
|
153 |
apply (simp add: fold_typing_def, auto) |
|
154 |
apply (auto dest: fold_set_lemma intro: ftype etype fcomm) |
|
155 |
done |
|
156 |
||
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
157 |
lemma fold_0 [simp]: "e \<in> B ==> fold[B](f,e,0) = e" |
14071 | 158 |
apply (unfold fold_def) |
159 |
apply (blast elim!: empty_fold_setE intro: fold_set.intros) |
|
160 |
done |
|
161 |
||
162 |
text{*This result is the right-to-left direction of the subsequent result*} |
|
163 |
lemma (in fold_typing) fold_set_imp_cons: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
164 |
"[| <C, y> \<in> fold_set(C, B, f, e); C \<in> Fin(A); c \<in> A; c\<notin>C |] |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
165 |
==> <cons(c, C), f(c,y)> \<in> fold_set(cons(c, C), B, f, e)" |
14071 | 166 |
apply (frule FinD [THEN fold_set_mono, THEN subsetD]) |
167 |
apply assumption |
|
168 |
apply (frule fold_set.dom_subset [of A, THEN subsetD]) |
|
169 |
apply (blast intro!: fold_set.consI intro: fold_set_mono [THEN subsetD]) |
|
170 |
done |
|
171 |
||
172 |
lemma (in fold_typing) fold_cons_lemma [rule_format]: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
173 |
"[| C \<in> Fin(A); c \<in> A; c\<notin>C |] |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
174 |
==> <cons(c, C), v> \<in> fold_set(cons(c, C), B, f, e) \<longleftrightarrow> |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
175 |
(\<exists>y. <C, y> \<in> fold_set(C, B, f, e) & v = f(c, y))" |
14071 | 176 |
apply auto |
177 |
prefer 2 apply (blast intro: fold_set_imp_cons) |
|
178 |
apply (frule_tac Fin.consI [of c, THEN FinD, THEN fold_set_mono, THEN subsetD], assumption+) |
|
179 |
apply (frule_tac fold_set.dom_subset [of A, THEN subsetD]) |
|
180 |
apply (drule FinD) |
|
181 |
apply (rule_tac A1 = "cons(c,C)" and f1=f and B1=B and C1=C and e1=e in fold_typing.Fin_imp_fold_set [THEN exE]) |
|
182 |
apply (blast intro: fold_typing.intro ftype etype fcomm) |
|
183 |
apply (blast intro: Fin_subset [of _ "cons(c,C)"] Finite_into_Fin |
|
184 |
dest: Fin_into_Finite) |
|
185 |
apply (rule_tac x = x in exI) |
|
186 |
apply (auto intro: fold_set.intros) |
|
187 |
apply (drule_tac fold_set_lemma [of C], blast) |
|
188 |
apply (blast intro!: fold_set.consI |
|
189 |
intro: fold_set_determ fold_set_mono [THEN subsetD] |
|
190 |
dest: fold_set.dom_subset [THEN subsetD]) |
|
191 |
done |
|
192 |
||
193 |
lemma (in fold_typing) fold_cons: |
|
194 |
"[| C\<in>Fin(A); c\<in>A; c\<notin>C|] |
|
195 |
==> fold[B](f, e, cons(c, C)) = f(c, fold[B](f, e, C))" |
|
196 |
apply (unfold fold_def) |
|
197 |
apply (simp add: fold_cons_lemma) |
|
198 |
apply (rule the_equality, auto) |
|
199 |
apply (subgoal_tac [2] "\<langle>C, y\<rangle> \<in> fold_set(A, B, f, e)") |
|
200 |
apply (drule Fin_imp_fold_set) |
|
201 |
apply (auto dest: fold_set_lemma simp add: fold_def [symmetric] fold_equality) |
|
202 |
apply (blast intro: fold_set_mono [THEN subsetD] dest!: FinD) |
|
203 |
done |
|
204 |
||
205 |
lemma (in fold_typing) fold_type [simp,TC]: |
|
206 |
"C\<in>Fin(A) ==> fold[B](f,e,C):B" |
|
207 |
apply (erule Fin_induct) |
|
208 |
apply (simp_all add: fold_cons ftype etype) |
|
209 |
done |
|
210 |
||
211 |
lemma (in fold_typing) fold_commute [rule_format]: |
|
212 |
"[| C\<in>Fin(A); c\<in>A |] |
|
213 |
==> (\<forall>y\<in>B. f(c, fold[B](f, y, C)) = fold[B](f, f(c, y), C))" |
|
214 |
apply (erule Fin_induct) |
|
215 |
apply (simp_all add: fold_typing.fold_cons [of A B _ f] |
|
216 |
fold_typing.fold_type [of A B _ f] |
|
217 |
fold_typing_def fcomm) |
|
218 |
done |
|
219 |
||
220 |
lemma (in fold_typing) fold_nest_Un_Int: |
|
221 |
"[| C\<in>Fin(A); D\<in>Fin(A) |] |
|
222 |
==> fold[B](f, fold[B](f, e, D), C) = |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
223 |
fold[B](f, fold[B](f, e, (C \<inter> D)), C \<union> D)" |
14071 | 224 |
apply (erule Fin_induct, auto) |
225 |
apply (simp add: Un_cons Int_cons_left fold_type fold_commute |
|
226 |
fold_typing.fold_cons [of A _ _ f] |
|
227 |
fold_typing_def fcomm cons_absorb) |
|
228 |
done |
|
229 |
||
230 |
lemma (in fold_typing) fold_nest_Un_disjoint: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
231 |
"[| C\<in>Fin(A); D\<in>Fin(A); C \<inter> D = 0 |] |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
232 |
==> fold[B](f,e,C \<union> D) = fold[B](f, fold[B](f,e,D), C)" |
14071 | 233 |
by (simp add: fold_nest_Un_Int) |
234 |
||
235 |
lemma Finite_cons_lemma: "Finite(C) ==> C\<in>Fin(cons(c, C))" |
|
236 |
apply (drule Finite_into_Fin) |
|
237 |
apply (blast intro: Fin_mono [THEN subsetD]) |
|
238 |
done |
|
239 |
||
240 |
subsection{*The Operator @{term setsum}*} |
|
241 |
||
242 |
lemma setsum_0 [simp]: "setsum(g, 0) = #0" |
|
243 |
by (simp add: setsum_def) |
|
244 |
||
245 |
lemma setsum_cons [simp]: |
|
246 |
"Finite(C) ==> |
|
247 |
setsum(g, cons(c,C)) = |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
248 |
(if c \<in> C then setsum(g,C) else g(c) $+ setsum(g,C))" |
14071 | 249 |
apply (auto simp add: setsum_def Finite_cons cons_absorb) |
250 |
apply (rule_tac A = "cons (c, C)" in fold_typing.fold_cons) |
|
251 |
apply (auto intro: fold_typing.intro Finite_cons_lemma) |
|
252 |
done |
|
253 |
||
254 |
lemma setsum_K0: "setsum((%i. #0), C) = #0" |
|
255 |
apply (case_tac "Finite (C) ") |
|
256 |
prefer 2 apply (simp add: setsum_def) |
|
257 |
apply (erule Finite_induct, auto) |
|
258 |
done |
|
259 |
||
260 |
(*The reversed orientation looks more natural, but LOOPS as a simprule!*) |
|
261 |
lemma setsum_Un_Int: |
|
262 |
"[| Finite(C); Finite(D) |] |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
263 |
==> setsum(g, C \<union> D) $+ setsum(g, C \<inter> D) |
14071 | 264 |
= setsum(g, C) $+ setsum(g, D)" |
265 |
apply (erule Finite_induct) |
|
266 |
apply (simp_all add: Int_cons_right cons_absorb Un_cons Int_commute Finite_Un |
|
267 |
Int_lower1 [THEN subset_Finite]) |
|
268 |
done |
|
269 |
||
270 |
lemma setsum_type [simp,TC]: "setsum(g, C):int" |
|
271 |
apply (case_tac "Finite (C) ") |
|
272 |
prefer 2 apply (simp add: setsum_def) |
|
273 |
apply (erule Finite_induct, auto) |
|
274 |
done |
|
275 |
||
276 |
lemma setsum_Un_disjoint: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
277 |
"[| Finite(C); Finite(D); C \<inter> D = 0 |] |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
278 |
==> setsum(g, C \<union> D) = setsum(g, C) $+ setsum(g,D)" |
14071 | 279 |
apply (subst setsum_Un_Int [symmetric]) |
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
280 |
apply (subgoal_tac [3] "Finite (C \<union> D) ") |
14071 | 281 |
apply (auto intro: Finite_Un) |
282 |
done |
|
283 |
||
284 |
lemma Finite_RepFun [rule_format (no_asm)]: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
285 |
"Finite(I) ==> (\<forall>i\<in>I. Finite(C(i))) \<longrightarrow> Finite(RepFun(I, C))" |
14071 | 286 |
apply (erule Finite_induct, auto) |
287 |
done |
|
288 |
||
289 |
lemma setsum_UN_disjoint [rule_format (no_asm)]: |
|
290 |
"Finite(I) |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
291 |
==> (\<forall>i\<in>I. Finite(C(i))) \<longrightarrow> |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
292 |
(\<forall>i\<in>I. \<forall>j\<in>I. i\<noteq>j \<longrightarrow> C(i) \<inter> C(j) = 0) \<longrightarrow> |
14071 | 293 |
setsum(f, \<Union>i\<in>I. C(i)) = setsum (%i. setsum(f, C(i)), I)" |
294 |
apply (erule Finite_induct, auto) |
|
295 |
apply (subgoal_tac "\<forall>i\<in>B. x \<noteq> i") |
|
296 |
prefer 2 apply blast |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
297 |
apply (subgoal_tac "C (x) \<inter> (\<Union>i\<in>B. C (i)) = 0") |
14071 | 298 |
prefer 2 apply blast |
299 |
apply (subgoal_tac "Finite (\<Union>i\<in>B. C (i)) & Finite (C (x)) & Finite (B) ") |
|
300 |
apply (simp (no_asm_simp) add: setsum_Un_disjoint) |
|
301 |
apply (auto intro: Finite_Union Finite_RepFun) |
|
302 |
done |
|
303 |
||
304 |
||
305 |
lemma setsum_addf: "setsum(%x. f(x) $+ g(x),C) = setsum(f, C) $+ setsum(g, C)" |
|
306 |
apply (case_tac "Finite (C) ") |
|
307 |
prefer 2 apply (simp add: setsum_def) |
|
308 |
apply (erule Finite_induct, auto) |
|
309 |
done |
|
310 |
||
311 |
||
312 |
lemma fold_set_cong: |
|
313 |
"[| A=A'; B=B'; e=e'; (\<forall>x\<in>A'. \<forall>y\<in>B'. f(x,y) = f'(x,y)) |] |
|
314 |
==> fold_set(A,B,f,e) = fold_set(A',B',f',e')" |
|
315 |
apply (simp add: fold_set_def) |
|
316 |
apply (intro refl iff_refl lfp_cong Collect_cong disj_cong ex_cong, auto) |
|
317 |
done |
|
318 |
||
319 |
lemma fold_cong: |
|
320 |
"[| B=B'; A=A'; e=e'; |
|
321 |
!!x y. [|x\<in>A'; y\<in>B'|] ==> f(x,y) = f'(x,y) |] ==> |
|
322 |
fold[B](f,e,A) = fold[B'](f', e', A')" |
|
323 |
apply (simp add: fold_def) |
|
324 |
apply (subst fold_set_cong) |
|
325 |
apply (rule_tac [5] refl, simp_all) |
|
326 |
done |
|
327 |
||
328 |
lemma setsum_cong: |
|
329 |
"[| A=B; !!x. x\<in>B ==> f(x) = g(x) |] ==> |
|
330 |
setsum(f, A) = setsum(g, B)" |
|
331 |
by (simp add: setsum_def cong add: fold_cong) |
|
332 |
||
333 |
lemma setsum_Un: |
|
334 |
"[| Finite(A); Finite(B) |] |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
335 |
==> setsum(f, A \<union> B) = |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
336 |
setsum(f, A) $+ setsum(f, B) $- setsum(f, A \<inter> B)" |
14071 | 337 |
apply (subst setsum_Un_Int [symmetric], auto) |
338 |
done |
|
339 |
||
340 |
||
341 |
lemma setsum_zneg_or_0 [rule_format (no_asm)]: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
342 |
"Finite(A) ==> (\<forall>x\<in>A. g(x) $<= #0) \<longrightarrow> setsum(g, A) $<= #0" |
14071 | 343 |
apply (erule Finite_induct) |
344 |
apply (auto intro: zneg_or_0_add_zneg_or_0_imp_zneg_or_0) |
|
345 |
done |
|
346 |
||
347 |
lemma setsum_succD_lemma [rule_format]: |
|
348 |
"Finite(A) |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
349 |
==> \<forall>n\<in>nat. setsum(f,A) = $# succ(n) \<longrightarrow> (\<exists>a\<in>A. #0 $< f(a))" |
14071 | 350 |
apply (erule Finite_induct) |
351 |
apply (auto simp del: int_of_0 int_of_succ simp add: not_zless_iff_zle int_of_0 [symmetric]) |
|
352 |
apply (subgoal_tac "setsum (f, B) $<= #0") |
|
353 |
apply simp_all |
|
354 |
prefer 2 apply (blast intro: setsum_zneg_or_0) |
|
355 |
apply (subgoal_tac "$# 1 $<= f (x) $+ setsum (f, B) ") |
|
356 |
apply (drule zdiff_zle_iff [THEN iffD2]) |
|
357 |
apply (subgoal_tac "$# 1 $<= $# 1 $- setsum (f,B) ") |
|
358 |
apply (drule_tac x = "$# 1" in zle_trans) |
|
359 |
apply (rule_tac [2] j = "#1" in zless_zle_trans, auto) |
|
360 |
done |
|
361 |
||
362 |
lemma setsum_succD: |
|
363 |
"[| setsum(f, A) = $# succ(n); n\<in>nat |]==> \<exists>a\<in>A. #0 $< f(a)" |
|
364 |
apply (case_tac "Finite (A) ") |
|
365 |
apply (blast intro: setsum_succD_lemma) |
|
366 |
apply (unfold setsum_def) |
|
367 |
apply (auto simp del: int_of_0 int_of_succ simp add: int_succ_int_1 [symmetric] int_of_0 [symmetric]) |
|
368 |
done |
|
369 |
||
370 |
lemma g_zpos_imp_setsum_zpos [rule_format]: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
371 |
"Finite(A) ==> (\<forall>x\<in>A. #0 $<= g(x)) \<longrightarrow> #0 $<= setsum(g, A)" |
14071 | 372 |
apply (erule Finite_induct) |
373 |
apply (simp (no_asm)) |
|
374 |
apply (auto intro: zpos_add_zpos_imp_zpos) |
|
375 |
done |
|
376 |
||
377 |
lemma g_zpos_imp_setsum_zpos2 [rule_format]: |
|
378 |
"[| Finite(A); \<forall>x. #0 $<= g(x) |] ==> #0 $<= setsum(g, A)" |
|
379 |
apply (erule Finite_induct) |
|
380 |
apply (auto intro: zpos_add_zpos_imp_zpos) |
|
381 |
done |
|
382 |
||
383 |
lemma g_zspos_imp_setsum_zspos [rule_format]: |
|
384 |
"Finite(A) |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
385 |
==> (\<forall>x\<in>A. #0 $< g(x)) \<longrightarrow> A \<noteq> 0 \<longrightarrow> (#0 $< setsum(g, A))" |
14071 | 386 |
apply (erule Finite_induct) |
387 |
apply (auto intro: zspos_add_zspos_imp_zspos) |
|
388 |
done |
|
389 |
||
390 |
lemma setsum_Diff [rule_format]: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
391 |
"Finite(A) ==> \<forall>a. M(a) = #0 \<longrightarrow> setsum(M, A) = setsum(M, A-{a})" |
14071 | 392 |
apply (erule Finite_induct) |
393 |
apply (simp_all add: Diff_cons_eq Finite_Diff) |
|
394 |
done |
|
395 |
||
12089
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff
changeset
|
396 |
end |