|
14052
|
1 |
(* Title: ZF/UNITY/MultusetSum.thy
|
|
|
2 |
ID: $Id$
|
|
|
3 |
Author: Sidi O Ehmety
|
|
|
4 |
Copyright: 2002 University of Cambridge
|
|
|
5 |
Setsum for multisets.
|
|
|
6 |
*)
|
|
|
7 |
|
|
|
8 |
Goal "mset_of(normalize(M)) <= mset_of(M)";
|
|
|
9 |
by (auto_tac (claset(), simpset() addsimps [mset_of_def,normalize_def]));
|
|
|
10 |
qed "mset_of_normalize";
|
|
|
11 |
|
|
|
12 |
Goalw [general_setsum_def]
|
|
|
13 |
"general_setsum(0, B, e, f, g) = e";
|
|
|
14 |
by Auto_tac;
|
|
|
15 |
qed "general_setsum_0";
|
|
|
16 |
Addsimps [general_setsum_0];
|
|
|
17 |
|
|
|
18 |
Goalw [general_setsum_def]
|
|
|
19 |
"[| C:Fin(A); a:A; a~:C; e:B; ALL x:A. g(x):B; lcomm(B, B, f) |] ==> \
|
|
|
20 |
\ general_setsum(cons(a, C), B, e, f, g) = \
|
|
|
21 |
\ f(g(a), general_setsum(C, B, e, f,g))";
|
|
|
22 |
by (auto_tac (claset(), simpset() addsimps [Fin_into_Finite RS Finite_cons,
|
|
|
23 |
cons_absorb]));
|
|
|
24 |
by (blast_tac (claset() addDs [Fin_into_Finite]) 2);
|
|
|
25 |
by (resolve_tac [fold_cons] 1);
|
|
|
26 |
by (auto_tac (claset(), simpset() addsimps [lcomm_def]));
|
|
|
27 |
qed "general_setsum_cons";
|
|
|
28 |
Addsimps [general_setsum_cons];
|
|
|
29 |
|
|
|
30 |
(** lcomm **)
|
|
|
31 |
|
|
|
32 |
Goalw [lcomm_def]
|
|
|
33 |
"[| C<=A; lcomm(A, B, f) |] ==> lcomm(C, B,f)";
|
|
|
34 |
by Auto_tac;
|
|
|
35 |
by (Blast_tac 1);
|
|
|
36 |
qed "lcomm_mono";
|
|
|
37 |
|
|
|
38 |
Goalw [lcomm_def]
|
|
|
39 |
"lcomm(Mult(A), Mult(A), op +#)";
|
|
|
40 |
by (auto_tac (claset(), simpset()
|
|
|
41 |
addsimps [Mult_iff_multiset, munion_lcommute]));
|
|
|
42 |
qed "munion_lcomm";
|
|
|
43 |
Addsimps [munion_lcomm];
|
|
|
44 |
|
|
|
45 |
(** msetsum **)
|
|
|
46 |
|
|
|
47 |
Goal
|
|
|
48 |
"[| C:Fin(A); ALL x:A. multiset(g(x))& mset_of(g(x))<=B |] \
|
|
|
49 |
\ ==> multiset(general_setsum(C, B -||> nat - {0}, 0, \\<lambda>u v. u +# v, g))";
|
|
|
50 |
by (etac Fin_induct 1);
|
|
|
51 |
by Auto_tac;
|
|
|
52 |
by (stac general_setsum_cons 1);
|
|
|
53 |
by (auto_tac (claset(), simpset() addsimps [Mult_iff_multiset]));
|
|
|
54 |
qed "multiset_general_setsum";
|
|
|
55 |
|
|
|
56 |
Goalw [msetsum_def] "msetsum(g, 0, B) = 0";
|
|
|
57 |
by Auto_tac;
|
|
|
58 |
qed "msetsum_0";
|
|
|
59 |
Addsimps [msetsum_0];
|
|
|
60 |
|
|
|
61 |
Goalw [msetsum_def]
|
|
|
62 |
"[| C:Fin(A); a~:C; a:A; ALL x:A. multiset(g(x)) & mset_of(g(x))<=B |] \
|
|
|
63 |
\ ==> msetsum(g, cons(a, C), B) = g(a) +# msetsum(g, C, B)";
|
|
|
64 |
by (stac general_setsum_cons 1);
|
|
|
65 |
by (auto_tac (claset(), simpset() addsimps [multiset_general_setsum, Mult_iff_multiset]));
|
|
|
66 |
qed "msetsum_cons";
|
|
|
67 |
Addsimps [msetsum_cons];
|
|
|
68 |
|
|
|
69 |
(* msetsum type *)
|
|
|
70 |
|
|
|
71 |
Goal "multiset(msetsum(g, C, B))";
|
|
|
72 |
by (asm_full_simp_tac (simpset() addsimps [msetsum_def]) 1);
|
|
|
73 |
qed "msetsum_multiset";
|
|
|
74 |
|
|
|
75 |
Goal
|
|
|
76 |
"[| C:Fin(A); ALL x:A. multiset(g(x)) & mset_of(g(x))<=B |] \
|
|
|
77 |
\ ==> mset_of(msetsum(g, C, B))<=B";
|
|
|
78 |
by (etac Fin_induct 1);
|
|
|
79 |
by Auto_tac;
|
|
|
80 |
qed "mset_of_msetsum";
|
|
|
81 |
|
|
|
82 |
|
|
|
83 |
|
|
|
84 |
(*The reversed orientation looks more natural, but LOOPS as a simprule!*)
|
|
|
85 |
Goal
|
|
|
86 |
"[| C:Fin(A); D:Fin(A); ALL x:A. multiset(g(x)) & mset_of(g(x))<=B |] \
|
|
|
87 |
\ ==> msetsum(g, C Un D, B) +# msetsum(g, C Int D, B) \
|
|
|
88 |
\ = msetsum(g, C, B) +# msetsum(g, D, B)";
|
|
|
89 |
by (etac Fin_induct 1);
|
|
|
90 |
by (subgoal_tac "cons(x, y) Un D = cons(x, y Un D)" 2);
|
|
|
91 |
by (auto_tac (claset(), simpset() addsimps [msetsum_multiset]));
|
|
|
92 |
by (subgoal_tac "y Un D:Fin(A) & y Int D : Fin(A)" 1);
|
|
|
93 |
by (Clarify_tac 1);
|
|
|
94 |
by (case_tac "x:D" 1);
|
|
|
95 |
by (subgoal_tac "cons(x, y) Int D = y Int D" 2);
|
|
|
96 |
by (subgoal_tac "cons(x, y) Int D = cons(x, y Int D)" 1);
|
|
|
97 |
by (ALLGOALS(asm_simp_tac (simpset() addsimps [cons_absorb,
|
|
|
98 |
munion_assoc, msetsum_multiset])));
|
|
|
99 |
by (asm_simp_tac (simpset() addsimps [munion_lcommute, msetsum_multiset]) 1);
|
|
|
100 |
by Auto_tac;
|
|
|
101 |
qed "msetsum_Un_Int";
|
|
|
102 |
|
|
|
103 |
|
|
|
104 |
Goal "[| C:Fin(A); D:Fin(A); C Int D = 0; \
|
|
|
105 |
\ ALL x:A. multiset(g(x)) & mset_of(g(x))<=B |] \
|
|
|
106 |
\ ==> msetsum(g, C Un D, B) = msetsum(g, C, B) +# msetsum(g,D, B)";
|
|
|
107 |
by (stac (msetsum_Un_Int RS sym) 1);
|
|
|
108 |
by (auto_tac (claset(), simpset() addsimps [msetsum_multiset]));
|
|
|
109 |
qed "msetsum_Un_disjoint";
|
|
|
110 |
|
|
|
111 |
Goal "I:Fin(A) ==> (ALL i:I. C(i):Fin(B)) --> (UN i:I. C(i)):Fin(B)";
|
|
|
112 |
by (etac Fin_induct 1);
|
|
|
113 |
by Auto_tac;
|
|
|
114 |
qed_spec_mp "UN_Fin_lemma";
|
|
|
115 |
|
|
|
116 |
Goal "!!I. [| I:Fin(K); ALL i:K. C(i):Fin(A) |] ==> \
|
|
|
117 |
\ (ALL x:A. multiset(f(x)) & mset_of(f(x))<=B) --> \
|
|
|
118 |
\ (ALL i:I. ALL j:I. i~=j --> C(i) Int C(j) = 0) --> \
|
|
|
119 |
\ msetsum(f, UN i:I. C(i), B) = msetsum (%i. msetsum(f, C(i),B), I, B)";
|
|
|
120 |
by (etac Fin_induct 1);
|
|
|
121 |
by (ALLGOALS(Clarify_tac));
|
|
|
122 |
by Auto_tac;
|
|
|
123 |
by (subgoal_tac "ALL i:y. x ~= i" 1);
|
|
|
124 |
by (Blast_tac 2);
|
|
|
125 |
by (subgoal_tac "C(x) Int (UN i:y. C(i)) = 0" 1);
|
|
|
126 |
by (Blast_tac 2);
|
|
|
127 |
by (subgoal_tac " (UN i:y. C(i)):Fin(A) & C(x):Fin(A)" 1);
|
|
|
128 |
by (blast_tac (claset() addIs [UN_Fin_lemma] addDs [FinD]) 2);
|
|
|
129 |
by (Clarify_tac 1);
|
|
|
130 |
by (asm_simp_tac (simpset() addsimps [msetsum_Un_disjoint]) 1);
|
|
|
131 |
by (subgoal_tac "ALL x:K. multiset(msetsum(f, C(x), B)) &\
|
|
|
132 |
\ mset_of(msetsum(f, C(x), B)) <= B" 1);
|
|
|
133 |
by (Asm_simp_tac 1);
|
|
|
134 |
by (Clarify_tac 1);
|
|
|
135 |
by (dres_inst_tac [("x", "xa")] bspec 1);
|
|
|
136 |
by (ALLGOALS(asm_simp_tac (simpset() addsimps [msetsum_multiset, mset_of_msetsum])));
|
|
|
137 |
qed_spec_mp "msetsum_UN_disjoint";
|
|
|
138 |
|
|
|
139 |
Goal
|
|
|
140 |
"[| C:Fin(A); \
|
|
|
141 |
\ ALL x:A. multiset(f(x)) & mset_of(f(x))<=B; \
|
|
|
142 |
\ ALL x:A. multiset(g(x)) & mset_of(g(x))<=B |] ==>\
|
|
|
143 |
\ msetsum(%x. f(x) +# g(x), C, B) = msetsum(f, C, B) +# msetsum(g, C, B)";
|
|
|
144 |
by (subgoal_tac "ALL x:A. multiset(f(x) +# g(x)) & mset_of(f(x) +# g(x))<=B" 1);
|
|
|
145 |
by (etac Fin_induct 1);
|
|
|
146 |
by (ALLGOALS(Asm_simp_tac));
|
|
|
147 |
by (resolve_tac [trans] 1);
|
|
|
148 |
by (resolve_tac [msetsum_cons] 1);
|
|
|
149 |
by (assume_tac 1);
|
|
|
150 |
by (auto_tac (claset(), simpset() addsimps [msetsum_multiset, munion_assoc]));
|
|
|
151 |
by (asm_simp_tac (simpset() addsimps [msetsum_multiset, munion_lcommute]) 1);
|
|
|
152 |
qed "msetsum_addf";
|
|
|
153 |
|
|
|
154 |
|
|
|
155 |
val prems = Goal
|
|
|
156 |
"[| C=D; !!x. x:D ==> f(x) = g(x) |] ==> \
|
|
|
157 |
\ msetsum(f, C, B) = msetsum(g, D, B)";
|
|
|
158 |
by (asm_full_simp_tac (simpset() addsimps [msetsum_def, general_setsum_def]@prems addcongs [fold_cong]) 1);
|
|
|
159 |
qed "msetsum_cong";
|
|
|
160 |
|
|
|
161 |
Goal "[| multiset(M); multiset(N) |] ==> M +# N -# N = M";
|
|
|
162 |
by (asm_simp_tac (simpset() addsimps [multiset_equality]) 1);
|
|
|
163 |
qed "multiset_union_diff";
|
|
|
164 |
|
|
|
165 |
|
|
|
166 |
Goal "[| C:Fin(A); D:Fin(A); \
|
|
|
167 |
\ ALL x:A. multiset(f(x)) & mset_of(f(x))<=B |] \
|
|
|
168 |
\ ==> msetsum(f, C Un D, B) = \
|
|
|
169 |
\ msetsum(f, C, B) +# msetsum(f, D, B) -# msetsum(f, C Int D, B)";
|
|
|
170 |
by (subgoal_tac "C Un D:Fin(A) & C Int D:Fin(A)" 1);
|
|
|
171 |
by (Clarify_tac 1);
|
|
|
172 |
by (stac (msetsum_Un_Int RS sym) 1);
|
|
|
173 |
by (ALLGOALS(asm_simp_tac (simpset() addsimps
|
|
|
174 |
[msetsum_multiset,multiset_union_diff])));
|
|
|
175 |
by (blast_tac (claset() addDs [FinD]) 1);
|
|
|
176 |
qed "msetsum_Un";
|
|
|
177 |
|
|
|
178 |
|
|
|
179 |
Goalw [nsetsum_def] "nsetsum(g, 0)=0";
|
|
|
180 |
by Auto_tac;
|
|
|
181 |
qed "nsetsum_0";
|
|
|
182 |
Addsimps [nsetsum_0];
|
|
|
183 |
|
|
|
184 |
Goalw [nsetsum_def, general_setsum_def]
|
|
|
185 |
"[| Finite(C); x~:C |] \
|
|
|
186 |
\ ==> nsetsum(g, cons(x, C))= g(x) #+ nsetsum(g, C)";
|
|
|
187 |
by (auto_tac (claset(), simpset() addsimps [Finite_cons]));
|
|
|
188 |
by (res_inst_tac [("A", "cons(x, C)")] fold_cons 1);
|
|
|
189 |
by (auto_tac (claset() addIs [Finite_cons_lemma], simpset()));
|
|
|
190 |
qed "nsetsum_cons";
|
|
|
191 |
Addsimps [nsetsum_cons];
|
|
|
192 |
|
|
|
193 |
Goal "nsetsum(g, C):nat";
|
|
|
194 |
by (case_tac "Finite(C)" 1);
|
|
|
195 |
by (asm_simp_tac (simpset() addsimps [nsetsum_def, general_setsum_def]) 2);
|
|
|
196 |
by (etac Finite_induct 1);
|
|
|
197 |
by Auto_tac;
|
|
|
198 |
qed "nsetsum_type";
|
|
|
199 |
Addsimps [nsetsum_type];
|
|
|
200 |
AddTCs [nsetsum_type];
|