changeset 2540 | 135ac0fb2686 |
parent 2539 | a8f5611dbd65 |
child 2541 | d7269488c4b5 |
2539:a8f5611dbd65 | 2540:135ac0fb2686 |
---|---|
7 |
7 |
8 theory FSet |
8 theory FSet |
9 imports Quotient_List |
9 imports Quotient_List |
10 begin |
10 begin |
11 |
11 |
12 text {* Definiton of the list equivalence relation *} |
12 text {* Definiton of the equivalence relation over lists *} |
13 |
13 |
14 fun |
14 fun |
15 list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50) |
15 list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50) |
16 where |
16 where |
17 "list_eq xs ys = (set xs = set ys)" |
17 "list_eq xs ys = (set xs = set ys)" |
27 quotient_type |
27 quotient_type |
28 'a fset = "'a list" / "list_eq" |
28 'a fset = "'a list" / "list_eq" |
29 by (rule list_eq_equivp) |
29 by (rule list_eq_equivp) |
30 |
30 |
31 text {* |
31 text {* |
32 Definitions of membership, sublist, cardinality, intersection, |
32 Definitions for membership, sublist, cardinality, |
33 difference and respectful fold over lists |
33 intersection, difference and respectful fold over |
34 lists. |
|
34 *} |
35 *} |
35 |
36 |
36 definition |
37 definition |
37 "memb x xs \<equiv> x \<in> set xs" |
38 "memb x xs \<equiv> x \<in> set xs" |
38 |
39 |
44 |
45 |
45 definition |
46 definition |
46 "inter_list xs ys = [x \<leftarrow> xs. x \<in> set xs \<and> x \<in> set ys]" |
47 "inter_list xs ys = [x \<leftarrow> xs. x \<in> set xs \<and> x \<in> set ys]" |
47 |
48 |
48 definition |
49 definition |
49 "diff_list xs ys \<equiv> [x \<leftarrow> xs. x\<notin>set ys]" |
50 "diff_list xs ys \<equiv> [x \<leftarrow> xs. x \<notin> set ys]" |
50 |
51 |
51 definition |
52 definition |
52 rsp_fold |
53 rsp_fold |
53 where |
54 where |
54 "rsp_fold f \<equiv> \<forall>u v w. (f u (f v w) = f v (f u w))" |
55 "rsp_fold f \<equiv> \<forall>u v w. (f u (f v w) = f v (f u w))" |
55 |
56 |
56 primrec |
57 primrec |
57 ffold_list :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b" |
58 fold_list :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b" |
58 where |
59 where |
59 "ffold_list f z [] = z" |
60 "fold_list f z [] = z" |
60 | "ffold_list f z (a # xs) = |
61 | "fold_list f z (a # xs) = |
61 (if (rsp_fold f) then |
62 (if (rsp_fold f) then |
62 if a \<in> set xs then ffold_list f z xs |
63 if a \<in> set xs then fold_list f z xs |
63 else f a (ffold_list f z xs) |
64 else f a (fold_list f z xs) |
64 else z)" |
65 else z)" |
65 |
66 |
66 |
67 |
68 |
|
67 section {* Quotient composition lemmas *} |
69 section {* Quotient composition lemmas *} |
68 |
70 |
69 lemma list_all2_refl1: |
71 lemma list_all2_refl': |
70 shows "(list_all2 op \<approx>) r r" |
72 shows "(list_all2 op \<approx>) r r" |
71 by (rule list_all2_refl) (metis equivp_def fset_equivp) |
73 by (rule list_all2_refl) (metis equivp_def fset_equivp) |
72 |
74 |
73 lemma compose_list_refl: |
75 lemma compose_list_refl: |
74 shows "(list_all2 op \<approx> OOO op \<approx>) r r" |
76 shows "(list_all2 op \<approx> OOO op \<approx>) r r" |
75 proof |
77 proof |
76 have *: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp]) |
78 have *: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp]) |
77 show "list_all2 op \<approx> r r" by (rule list_all2_refl1) |
79 show "list_all2 op \<approx> r r" by (rule list_all2_refl') |
78 with * show "(op \<approx> OO list_all2 op \<approx>) r r" .. |
80 with * show "(op \<approx> OO list_all2 op \<approx>) r r" .. |
79 qed |
81 qed |
80 |
82 |
81 lemma Quotient_fset_list: |
83 lemma Quotient_fset_list: |
82 shows "Quotient (list_all2 op \<approx>) (map abs_fset) (map rep_fset)" |
84 shows "Quotient (list_all2 op \<approx>) (map abs_fset) (map rep_fset)" |
83 by (fact list_quotient[OF Quotient_fset]) |
85 by (fact list_quotient[OF Quotient_fset]) |
84 |
86 |
85 lemma map_rel_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba" |
87 lemma map_list_eq_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba" |
86 unfolding list_eq.simps |
88 unfolding list_eq.simps |
87 by (simp only: set_map) |
89 by (simp only: set_map) |
88 |
90 |
89 lemma quotient_compose_list[quot_thm]: |
91 lemma quotient_compose_list[quot_thm]: |
90 shows "Quotient ((list_all2 op \<approx>) OOO (op \<approx>)) |
92 shows "Quotient ((list_all2 op \<approx>) OOO (op \<approx>)) |
93 proof (intro conjI allI) |
95 proof (intro conjI allI) |
94 fix a r s |
96 fix a r s |
95 show "abs_fset (map abs_fset (map rep_fset (rep_fset a))) = a" |
97 show "abs_fset (map abs_fset (map rep_fset (rep_fset a))) = a" |
96 by (simp add: abs_o_rep[OF Quotient_fset] Quotient_abs_rep[OF Quotient_fset] map_id) |
98 by (simp add: abs_o_rep[OF Quotient_fset] Quotient_abs_rep[OF Quotient_fset] map_id) |
97 have b: "list_all2 op \<approx> (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))" |
99 have b: "list_all2 op \<approx> (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))" |
98 by (rule list_all2_refl1) |
100 by (rule list_all2_refl') |
99 have c: "(op \<approx> OO list_all2 op \<approx>) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))" |
101 have c: "(op \<approx> OO list_all2 op \<approx>) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))" |
100 by (rule, rule equivp_reflp[OF fset_equivp]) (rule b) |
102 by (rule, rule equivp_reflp[OF fset_equivp]) (rule b) |
101 show "(list_all2 op \<approx> OOO op \<approx>) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))" |
103 show "(list_all2 op \<approx> OOO op \<approx>) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))" |
102 by (rule, rule list_all2_refl1) (rule c) |
104 by (rule, rule list_all2_refl') (rule c) |
103 show "(list_all2 op \<approx> OOO op \<approx>) r s = ((list_all2 op \<approx> OOO op \<approx>) r r \<and> |
105 show "(list_all2 op \<approx> OOO op \<approx>) r s = ((list_all2 op \<approx> OOO op \<approx>) r r \<and> |
104 (list_all2 op \<approx> OOO op \<approx>) s s \<and> abs_fset (map abs_fset r) = abs_fset (map abs_fset s))" |
106 (list_all2 op \<approx> OOO op \<approx>) s s \<and> abs_fset (map abs_fset r) = abs_fset (map abs_fset s))" |
105 proof (intro iffI conjI) |
107 proof (intro iffI conjI) |
106 show "(list_all2 op \<approx> OOO op \<approx>) r r" by (rule compose_list_refl) |
108 show "(list_all2 op \<approx> OOO op \<approx>) r r" by (rule compose_list_refl) |
107 show "(list_all2 op \<approx> OOO op \<approx>) s s" by (rule compose_list_refl) |
109 show "(list_all2 op \<approx> OOO op \<approx>) s s" by (rule compose_list_refl) |
116 have f: "map abs_fset r = map abs_fset b" |
118 have f: "map abs_fset r = map abs_fset b" |
117 using Quotient_rel[OF Quotient_fset_list] c by blast |
119 using Quotient_rel[OF Quotient_fset_list] c by blast |
118 have "map abs_fset ba = map abs_fset s" |
120 have "map abs_fset ba = map abs_fset s" |
119 using Quotient_rel[OF Quotient_fset_list] e by blast |
121 using Quotient_rel[OF Quotient_fset_list] e by blast |
120 then have g: "map abs_fset s = map abs_fset ba" by simp |
122 then have g: "map abs_fset s = map abs_fset ba" by simp |
121 then show "map abs_fset r \<approx> map abs_fset s" using d f map_rel_cong by simp |
123 then show "map abs_fset r \<approx> map abs_fset s" using d f map_list_eq_cong by simp |
122 qed |
124 qed |
123 then show "abs_fset (map abs_fset r) = abs_fset (map abs_fset s)" |
125 then show "abs_fset (map abs_fset r) = abs_fset (map abs_fset s)" |
124 using Quotient_rel[OF Quotient_fset] by blast |
126 using Quotient_rel[OF Quotient_fset] by blast |
125 next |
127 next |
126 assume a: "(list_all2 op \<approx> OOO op \<approx>) r r \<and> (list_all2 op \<approx> OOO op \<approx>) s s |
128 assume a: "(list_all2 op \<approx> OOO op \<approx>) r r \<and> (list_all2 op \<approx> OOO op \<approx>) s s |
127 \<and> abs_fset (map abs_fset r) = abs_fset (map abs_fset s)" |
129 \<and> abs_fset (map abs_fset r) = abs_fset (map abs_fset s)" |
128 then have s: "(list_all2 op \<approx> OOO op \<approx>) s s" by simp |
130 then have s: "(list_all2 op \<approx> OOO op \<approx>) s s" by simp |
129 have d: "map abs_fset r \<approx> map abs_fset s" |
131 have d: "map abs_fset r \<approx> map abs_fset s" |
130 by (subst Quotient_rel[OF Quotient_fset]) (simp add: a) |
132 by (subst Quotient_rel[OF Quotient_fset]) (simp add: a) |
131 have b: "map rep_fset (map abs_fset r) \<approx> map rep_fset (map abs_fset s)" |
133 have b: "map rep_fset (map abs_fset r) \<approx> map rep_fset (map abs_fset s)" |
132 by (rule map_rel_cong[OF d]) |
134 by (rule map_list_eq_cong[OF d]) |
133 have y: "list_all2 op \<approx> (map rep_fset (map abs_fset s)) s" |
135 have y: "list_all2 op \<approx> (map rep_fset (map abs_fset s)) s" |
134 by (fact rep_abs_rsp_left[OF Quotient_fset_list, OF list_all2_refl1[of s]]) |
136 by (fact rep_abs_rsp_left[OF Quotient_fset_list, OF list_all2_refl'[of s]]) |
135 have c: "(op \<approx> OO list_all2 op \<approx>) (map rep_fset (map abs_fset r)) s" |
137 have c: "(op \<approx> OO list_all2 op \<approx>) (map rep_fset (map abs_fset r)) s" |
136 by (rule pred_compI) (rule b, rule y) |
138 by (rule pred_compI) (rule b, rule y) |
137 have z: "list_all2 op \<approx> r (map rep_fset (map abs_fset r))" |
139 have z: "list_all2 op \<approx> r (map rep_fset (map abs_fset r))" |
138 by (fact rep_abs_rsp[OF Quotient_fset_list, OF list_all2_refl1[of r]]) |
140 by (fact rep_abs_rsp[OF Quotient_fset_list, OF list_all2_refl'[of r]]) |
139 then show "(list_all2 op \<approx> OOO op \<approx>) r s" |
141 then show "(list_all2 op \<approx> OOO op \<approx>) r s" |
140 using a c pred_compI by simp |
142 using a c pred_compI by simp |
141 qed |
143 qed |
142 qed |
144 qed |
143 |
145 |
194 |
196 |
195 lemma filter_rsp [quot_respect]: |
197 lemma filter_rsp [quot_respect]: |
196 shows "(op = ===> op \<approx> ===> op \<approx>) filter filter" |
198 shows "(op = ===> op \<approx> ===> op \<approx>) filter filter" |
197 by simp |
199 by simp |
198 |
200 |
199 lemma memb_commute_ffold_list: |
201 lemma memb_commute_fold_list: |
200 "rsp_fold f \<Longrightarrow> h \<in> set b \<Longrightarrow> ffold_list f z b = f h (ffold_list f z (removeAll h b))" |
202 assumes a: "rsp_fold f" |
201 apply (induct b) |
203 and b: "x \<in> set xs" |
202 apply (auto simp add: rsp_fold_def) |
204 shows "fold_list f y xs = f x (fold_list f y (removeAll x xs))" |
203 done |
205 using a b by (induct xs) (auto simp add: rsp_fold_def) |
204 |
206 |
205 lemma ffold_list_rsp_pre: |
207 lemma fold_list_rsp_pre: |
206 "set a = set b \<Longrightarrow> ffold_list f z a = ffold_list f z b" |
208 assumes a: "set xs = set ys" |
207 apply (induct a arbitrary: b) |
209 shows "fold_list f z xs = fold_list f z ys" |
210 using a |
|
211 apply (induct xs arbitrary: ys) |
|
208 apply (simp) |
212 apply (simp) |
209 apply (simp (no_asm_use)) |
213 apply (simp (no_asm_use)) |
210 apply (rule conjI) |
214 apply (rule conjI) |
211 apply (rule_tac [!] impI) |
215 apply (rule_tac [!] impI) |
212 apply (rule_tac [!] conjI) |
216 apply (rule_tac [!] conjI) |
213 apply (rule_tac [!] impI) |
217 apply (rule_tac [!] impI) |
214 apply (metis insert_absorb) |
218 apply (metis insert_absorb) |
215 apply (metis List.insert_def List.set.simps(2) List.set_insert ffold_list.simps(2)) |
219 apply (metis List.insert_def List.set.simps(2) List.set_insert fold_list.simps(2)) |
216 apply (metis Diff_insert_absorb insertI1 memb_commute_ffold_list set_removeAll) |
220 apply (metis Diff_insert_absorb insertI1 memb_commute_fold_list set_removeAll) |
217 apply(drule_tac x="removeAll a1 b" in meta_spec) |
221 apply(drule_tac x="removeAll a ys" in meta_spec) |
218 apply(auto) |
222 apply(auto) |
219 apply(drule meta_mp) |
223 apply(drule meta_mp) |
220 apply(blast) |
224 apply(blast) |
221 by (metis List.set.simps(2) emptyE ffold_list.simps(2) in_listsp_conv_set listsp.simps mem_def) |
225 by (metis List.set.simps(2) emptyE fold_list.simps(2) in_listsp_conv_set listsp.simps mem_def) |
222 |
226 |
223 lemma ffold_list_rsp [quot_respect]: |
227 lemma fold_list_rsp [quot_respect]: |
224 shows "(op = ===> op = ===> op \<approx> ===> op =) ffold_list ffold_list" |
228 shows "(op = ===> op = ===> op \<approx> ===> op =) fold_list fold_list" |
225 unfolding fun_rel_def |
229 unfolding fun_rel_def |
226 by(auto intro: ffold_list_rsp_pre) |
230 by(auto intro: fold_list_rsp_pre) |
227 |
231 |
228 lemma concat_rsp_pre: |
232 lemma concat_rsp_pre: |
229 assumes a: "list_all2 op \<approx> x x'" |
233 assumes a: "list_all2 op \<approx> x x'" |
230 and b: "x' \<approx> y'" |
234 and b: "x' \<approx> y'" |
231 and c: "list_all2 op \<approx> y' y" |
235 and c: "list_all2 op \<approx> y' y" |
278 quotient_definition |
282 quotient_definition |
279 "bot :: 'a fset" |
283 "bot :: 'a fset" |
280 is "Nil :: 'a list" |
284 is "Nil :: 'a list" |
281 |
285 |
282 abbreviation |
286 abbreviation |
283 fempty ("{||}") |
287 empty_fset ("{||}") |
284 where |
288 where |
285 "{||} \<equiv> bot :: 'a fset" |
289 "{||} \<equiv> bot :: 'a fset" |
286 |
290 |
287 quotient_definition |
291 quotient_definition |
288 "less_eq_fset :: ('a fset \<Rightarrow> 'a fset \<Rightarrow> bool)" |
292 "less_eq_fset :: ('a fset \<Rightarrow> 'a fset \<Rightarrow> bool)" |
289 is "sub_list :: ('a list \<Rightarrow> 'a list \<Rightarrow> bool)" |
293 is "sub_list :: ('a list \<Rightarrow> 'a list \<Rightarrow> bool)" |
290 |
294 |
291 abbreviation |
295 abbreviation |
292 f_subset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50) |
296 subset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50) |
293 where |
297 where |
294 "xs |\<subseteq>| ys \<equiv> xs \<le> ys" |
298 "xs |\<subseteq>| ys \<equiv> xs \<le> ys" |
295 |
299 |
296 definition |
300 definition |
297 less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" |
301 less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" |
298 where |
302 where |
299 "xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)" |
303 "xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)" |
300 |
304 |
301 abbreviation |
305 abbreviation |
302 fsubset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50) |
306 psubset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50) |
303 where |
307 where |
304 "xs |\<subset>| ys \<equiv> xs < ys" |
308 "xs |\<subset>| ys \<equiv> xs < ys" |
305 |
309 |
306 quotient_definition |
310 quotient_definition |
307 "sup :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
311 "sup :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
308 is "append :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
312 is "append :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
309 |
313 |
310 abbreviation |
314 abbreviation |
311 funion (infixl "|\<union>|" 65) |
315 union_fset (infixl "|\<union>|" 65) |
312 where |
316 where |
313 "xs |\<union>| ys \<equiv> sup xs (ys::'a fset)" |
317 "xs |\<union>| ys \<equiv> sup xs (ys::'a fset)" |
314 |
318 |
315 quotient_definition |
319 quotient_definition |
316 "inf :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
320 "inf :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
317 is "inter_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
321 is "inter_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
318 |
322 |
319 abbreviation |
323 abbreviation |
320 finter (infixl "|\<inter>|" 65) |
324 inter_fset (infixl "|\<inter>|" 65) |
321 where |
325 where |
322 "xs |\<inter>| ys \<equiv> inf xs (ys::'a fset)" |
326 "xs |\<inter>| ys \<equiv> inf xs (ys::'a fset)" |
323 |
327 |
324 quotient_definition |
328 quotient_definition |
325 "minus :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
329 "minus :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
368 by (descending) (auto simp add: inter_list_def sub_list_def memb_def) |
372 by (descending) (auto simp add: inter_list_def sub_list_def memb_def) |
369 qed |
373 qed |
370 |
374 |
371 end |
375 end |
372 |
376 |
373 quotient_definition |
377 |
374 "finsert :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
378 subsection {* Other constants for fsets *} |
379 |
|
380 quotient_definition |
|
381 "insert_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
|
375 is "Cons" |
382 is "Cons" |
376 |
383 |
377 syntax |
384 syntax |
378 "@Finset" :: "args => 'a fset" ("{|(_)|}") |
385 "@Insert_fset" :: "args => 'a fset" ("{|(_)|}") |
379 |
386 |
380 translations |
387 translations |
381 "{|x, xs|}" == "CONST finsert x {|xs|}" |
388 "{|x, xs|}" == "CONST insert_fset x {|xs|}" |
382 "{|x|}" == "CONST finsert x {||}" |
389 "{|x|}" == "CONST insert_fset x {||}" |
383 |
390 |
384 quotient_definition |
391 quotient_definition |
385 fin (infix "|\<in>|" 50) |
392 in_fset (infix "|\<in>|" 50) |
386 where |
393 where |
387 "fin :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool" is "memb" |
394 "in_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool" is "memb" |
388 |
395 |
389 abbreviation |
396 abbreviation |
390 fnotin :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50) |
397 notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50) |
391 where |
398 where |
392 "x |\<notin>| S \<equiv> \<not> (x |\<in>| S)" |
399 "x |\<notin>| S \<equiv> \<not> (x |\<in>| S)" |
393 |
400 |
394 section {* Other constants on the Quotient Type *} |
401 |
395 |
402 subsection {* Other constants on the Quotient Type *} |
396 quotient_definition |
403 |
397 "fcard :: 'a fset \<Rightarrow> nat" |
404 quotient_definition |
405 "card_fset :: 'a fset \<Rightarrow> nat" |
|
398 is card_list |
406 is card_list |
399 |
407 |
400 quotient_definition |
408 quotient_definition |
401 "fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" |
409 "map_fset :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" |
402 is map |
410 is map |
403 |
411 |
404 quotient_definition |
412 quotient_definition |
405 "fdelete :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
413 "remove_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
406 is removeAll |
414 is removeAll |
407 |
415 |
408 quotient_definition |
416 quotient_definition |
409 "fset :: 'a fset \<Rightarrow> 'a set" |
417 "fset :: 'a fset \<Rightarrow> 'a set" |
410 is "set" |
418 is "set" |
411 |
419 |
412 quotient_definition |
420 quotient_definition |
413 "ffold :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" |
421 "fold_fset :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" |
414 is ffold_list |
422 is fold_list |
415 |
423 |
416 quotient_definition |
424 quotient_definition |
417 "fconcat :: ('a fset) fset \<Rightarrow> 'a fset" |
425 "concat_fset :: ('a fset) fset \<Rightarrow> 'a fset" |
418 is concat |
426 is concat |
419 |
427 |
420 quotient_definition |
428 quotient_definition |
421 "ffilter :: ('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
429 "filter_fset :: ('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
422 is filter |
430 is filter |
423 |
431 |
424 |
432 |
425 section {* Compositional respectfulness and preservation lemmas *} |
433 subsection {* Compositional respectfulness and preservation lemmas *} |
426 |
434 |
427 lemma Nil_rsp2 [quot_respect]: |
435 lemma Nil_rsp2 [quot_respect]: |
428 shows "(list_all2 op \<approx> OOO op \<approx>) Nil Nil" |
436 shows "(list_all2 op \<approx> OOO op \<approx>) Nil Nil" |
429 by (fact compose_list_refl) |
437 by (fact compose_list_refl) |
430 |
438 |
439 |
447 |
440 lemma map_prs [quot_preserve]: |
448 lemma map_prs [quot_preserve]: |
441 shows "(abs_fset \<circ> map f) [] = abs_fset []" |
449 shows "(abs_fset \<circ> map f) [] = abs_fset []" |
442 by simp |
450 by simp |
443 |
451 |
444 lemma finsert_rsp [quot_preserve]: |
452 lemma insert_fset_rsp [quot_preserve]: |
445 "(rep_fset ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) op # = finsert" |
453 "(rep_fset ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) op # = insert_fset" |
446 by (simp add: fun_eq_iff Quotient_abs_rep[OF Quotient_fset] |
454 by (simp add: fun_eq_iff Quotient_abs_rep[OF Quotient_fset] |
447 abs_o_rep[OF Quotient_fset] map_id finsert_def) |
455 abs_o_rep[OF Quotient_fset] map_id insert_fset_def) |
448 |
456 |
449 lemma funion_rsp [quot_preserve]: |
457 lemma union_fset_rsp [quot_preserve]: |
450 "((map rep_fset \<circ> rep_fset) ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) op @ = funion" |
458 "((map rep_fset \<circ> rep_fset) ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) op @ = union_fset" |
451 by (simp add: fun_eq_iff Quotient_abs_rep[OF Quotient_fset] |
459 by (simp add: fun_eq_iff Quotient_abs_rep[OF Quotient_fset] |
452 abs_o_rep[OF Quotient_fset] map_id sup_fset_def) |
460 abs_o_rep[OF Quotient_fset] map_id sup_fset_def) |
453 |
461 |
454 lemma list_all2_app_l: |
462 lemma list_all2_app_l: |
455 assumes a: "reflp R" |
463 assumes a: "reflp R" |
459 |
467 |
460 lemma append_rsp2_pre0: |
468 lemma append_rsp2_pre0: |
461 assumes a:"list_all2 op \<approx> x x'" |
469 assumes a:"list_all2 op \<approx> x x'" |
462 shows "list_all2 op \<approx> (x @ z) (x' @ z)" |
470 shows "list_all2 op \<approx> (x @ z) (x' @ z)" |
463 using a apply (induct x x' rule: list_induct2') |
471 using a apply (induct x x' rule: list_induct2') |
464 by simp_all (rule list_all2_refl1) |
472 by simp_all (rule list_all2_refl') |
465 |
473 |
466 lemma append_rsp2_pre1: |
474 lemma append_rsp2_pre1: |
467 assumes a:"list_all2 op \<approx> x x'" |
475 assumes a:"list_all2 op \<approx> x x'" |
468 shows "list_all2 op \<approx> (z @ x) (z @ x')" |
476 shows "list_all2 op \<approx> (z @ x) (z @ x')" |
469 using a apply (induct x x' arbitrary: z rule: list_induct2') |
477 using a apply (induct x x' arbitrary: z rule: list_induct2') |
470 apply (rule list_all2_refl1) |
478 apply (rule list_all2_refl') |
471 apply (simp_all del: list_eq.simps) |
479 apply (simp_all del: list_eq.simps) |
472 apply (rule list_all2_app_l) |
480 apply (rule list_all2_app_l) |
473 apply (simp_all add: reflp_def) |
481 apply (simp_all add: reflp_def) |
474 done |
482 done |
475 |
483 |
480 apply (rule list_all2_transp[OF fset_equivp]) |
488 apply (rule list_all2_transp[OF fset_equivp]) |
481 apply (rule append_rsp2_pre0) |
489 apply (rule append_rsp2_pre0) |
482 apply (rule a) |
490 apply (rule a) |
483 using b apply (induct z z' rule: list_induct2') |
491 using b apply (induct z z' rule: list_induct2') |
484 apply (simp_all only: append_Nil2) |
492 apply (simp_all only: append_Nil2) |
485 apply (rule list_all2_refl1) |
493 apply (rule list_all2_refl') |
486 apply simp_all |
494 apply simp_all |
487 apply (rule append_rsp2_pre1) |
495 apply (rule append_rsp2_pre1) |
488 apply simp |
496 apply simp |
489 done |
497 done |
490 |
498 |
510 |
518 |
511 |
519 |
512 section {* Lifted theorems *} |
520 section {* Lifted theorems *} |
513 |
521 |
514 |
522 |
515 subsection {* fin *} |
523 subsection {* in_fset *} |
516 |
524 |
517 lemma not_fin_fnil: |
525 lemma notin_empty_fset: |
518 shows "x |\<notin>| {||}" |
526 shows "x |\<notin>| {||}" |
519 by (descending) (simp add: memb_def) |
527 by (descending) (simp add: memb_def) |
520 |
528 |
521 lemma fin_set: |
529 lemma in_fset: |
522 shows "x |\<in>| S \<longleftrightarrow> x \<in> fset S" |
530 shows "x |\<in>| S \<longleftrightarrow> x \<in> fset S" |
523 by (descending) (simp add: memb_def) |
531 by (descending) (simp add: memb_def) |
524 |
532 |
525 lemma fnotin_set: |
533 lemma notin_fset: |
526 shows "x |\<notin>| S \<longleftrightarrow> x \<notin> fset S" |
534 shows "x |\<notin>| S \<longleftrightarrow> x \<notin> fset S" |
527 by (descending) (simp add: memb_def) |
535 by (descending) (simp add: memb_def) |
528 |
536 |
529 lemma fset_eq_iff: |
537 lemma fset_eq_iff: |
530 shows "S = T \<longleftrightarrow> (\<forall>x. (x |\<in>| S) = (x |\<in>| T))" |
538 shows "S = T \<longleftrightarrow> (\<forall>x. (x |\<in>| S) = (x |\<in>| T))" |
531 by (descending) (auto simp add: memb_def) |
539 by (descending) (auto simp add: memb_def) |
532 |
540 |
533 lemma none_fin_fempty: |
541 lemma none_in_empty_fset: |
534 shows "(\<forall>x. x |\<notin>| S) \<longleftrightarrow> S = {||}" |
542 shows "(\<forall>x. x |\<notin>| S) \<longleftrightarrow> S = {||}" |
535 by (descending) (simp add: memb_def) |
543 by (descending) (simp add: memb_def) |
536 |
544 |
537 |
545 |
538 subsection {* finsert *} |
546 subsection {* insert_fset *} |
539 |
547 |
540 lemma fin_finsert_iff[simp]: |
548 lemma in_insert_fset_iff[simp]: |
541 shows "x |\<in>| finsert y S \<longleftrightarrow> x = y \<or> x |\<in>| S" |
549 shows "x |\<in>| insert_fset y S \<longleftrightarrow> x = y \<or> x |\<in>| S" |
542 by (descending) (simp add: memb_def) |
550 by (descending) (simp add: memb_def) |
543 |
551 |
544 lemma |
552 lemma |
545 shows finsertI1: "x |\<in>| finsert x S" |
553 shows insert_fsetI1: "x |\<in>| insert_fset x S" |
546 and finsertI2: "x |\<in>| S \<Longrightarrow> x |\<in>| finsert y S" |
554 and insert_fsetI2: "x |\<in>| S \<Longrightarrow> x |\<in>| insert_fset y S" |
547 by (descending, simp add: memb_def)+ |
555 by (descending, simp add: memb_def)+ |
548 |
556 |
549 lemma finsert_absorb[simp]: |
557 lemma insert_absorb_fset[simp]: |
550 shows "x |\<in>| S \<Longrightarrow> finsert x S = S" |
558 shows "x |\<in>| S \<Longrightarrow> insert_fset x S = S" |
551 by (descending) (auto simp add: memb_def) |
559 by (descending) (auto simp add: memb_def) |
552 |
560 |
553 lemma fempty_not_finsert[simp]: |
561 lemma empty_not_insert_fset[simp]: |
554 shows "{||} \<noteq> finsert x S" |
562 shows "{||} \<noteq> insert_fset x S" |
555 and "finsert x S \<noteq> {||}" |
563 and "insert_fset x S \<noteq> {||}" |
556 by (descending, simp)+ |
564 by (descending, simp)+ |
557 |
565 |
558 lemma finsert_left_comm: |
566 lemma insert_fset_left_comm: |
559 shows "finsert x (finsert y S) = finsert y (finsert x S)" |
567 shows "insert_fset x (insert_fset y S) = insert_fset y (insert_fset x S)" |
560 by (descending) (auto) |
568 by (descending) (auto) |
561 |
569 |
562 lemma finsert_left_idem: |
570 lemma insert_fset_left_idem: |
563 shows "finsert x (finsert x S) = finsert x S" |
571 shows "insert_fset x (insert_fset x S) = insert_fset x S" |
564 by (descending) (auto) |
572 by (descending) (auto) |
565 |
573 |
566 lemma fsingleton_eq[simp]: |
574 lemma singleton_fset_eq[simp]: |
567 shows "{|x|} = {|y|} \<longleftrightarrow> x = y" |
575 shows "{|x|} = {|y|} \<longleftrightarrow> x = y" |
568 by (descending) (auto) |
576 by (descending) (auto) |
569 |
577 |
570 |
578 |
571 (* FIXME: is this in any case a useful lemma *) |
579 (* FIXME: is this in any case a useful lemma *) |
572 lemma fin_mdef: |
580 lemma in_fset_mdef: |
573 shows "x |\<in>| F \<longleftrightarrow> x |\<notin>| (F - {|x|}) \<and> F = finsert x (F - {|x|})" |
581 shows "x |\<in>| F \<longleftrightarrow> x |\<notin>| (F - {|x|}) \<and> F = insert_fset x (F - {|x|})" |
574 by (descending) (auto simp add: memb_def diff_list_def) |
582 by (descending) (auto simp add: memb_def diff_list_def) |
575 |
583 |
576 |
584 |
577 subsection {* fset *} |
585 subsection {* fset *} |
578 |
586 |
579 lemma fset_simps[simp]: |
587 lemma fset_simps [simp]: |
580 "fset {||} = ({} :: 'a set)" |
588 "fset {||} = ({} :: 'a set)" |
581 "fset (finsert (x :: 'a) S) = insert x (fset S)" |
589 "fset (insert_fset (x :: 'a) S) = insert x (fset S)" |
582 by (lifting set.simps) |
590 by (lifting set.simps) |
583 |
591 |
584 lemma finite_fset [simp]: |
592 lemma finite_fset [simp]: |
585 shows "finite (fset S)" |
593 shows "finite (fset S)" |
586 by (descending) (simp) |
594 by (descending) (simp) |
587 |
595 |
588 lemma fset_cong: |
596 lemma fset_cong: |
589 shows "fset S = fset T \<longleftrightarrow> S = T" |
597 shows "fset S = fset T \<longleftrightarrow> S = T" |
590 by (descending) (simp) |
598 by (descending) (simp) |
591 |
599 |
592 lemma ffilter_set [simp]: |
600 lemma filter_fset [simp]: |
593 shows "fset (ffilter P xs) = P \<inter> fset xs" |
601 shows "fset (filter_fset P xs) = P \<inter> fset xs" |
594 by (descending) (auto simp add: mem_def) |
602 by (descending) (auto simp add: mem_def) |
595 |
603 |
596 lemma fdelete_set [simp]: |
604 lemma remove_fset [simp]: |
597 shows "fset (fdelete x xs) = fset xs - {x}" |
605 shows "fset (remove_fset x xs) = fset xs - {x}" |
598 by (descending) (simp) |
606 by (descending) (simp) |
599 |
607 |
600 lemma finter_set [simp]: |
608 lemma inter_fset [simp]: |
601 shows "fset (xs |\<inter>| ys) = fset xs \<inter> fset ys" |
609 shows "fset (xs |\<inter>| ys) = fset xs \<inter> fset ys" |
602 by (descending) (auto simp add: inter_list_def) |
610 by (descending) (auto simp add: inter_list_def) |
603 |
611 |
604 lemma funion_set [simp]: |
612 lemma union_fset [simp]: |
605 shows "fset (xs |\<union>| ys) = fset xs \<union> fset ys" |
613 shows "fset (xs |\<union>| ys) = fset xs \<union> fset ys" |
606 by (lifting set_append) |
614 by (lifting set_append) |
607 |
615 |
608 lemma fminus_set [simp]: |
616 lemma minus_fset [simp]: |
609 shows "fset (xs - ys) = fset xs - fset ys" |
617 shows "fset (xs - ys) = fset xs - fset ys" |
610 by (descending) (auto simp add: diff_list_def) |
618 by (descending) (auto simp add: diff_list_def) |
611 |
619 |
612 |
620 |
613 subsection {* funion *} |
621 subsection {* union_fset *} |
614 |
622 |
615 lemmas [simp] = |
623 lemmas [simp] = |
616 sup_bot_left[where 'a="'a fset", standard] |
624 sup_bot_left[where 'a="'a fset", standard] |
617 sup_bot_right[where 'a="'a fset", standard] |
625 sup_bot_right[where 'a="'a fset", standard] |
618 |
626 |
619 lemma funion_finsert[simp]: |
627 lemma union_insert_fset [simp]: |
620 shows "finsert x S |\<union>| T = finsert x (S |\<union>| T)" |
628 shows "insert_fset x S |\<union>| T = insert_fset x (S |\<union>| T)" |
621 by (lifting append.simps(2)) |
629 by (lifting append.simps(2)) |
622 |
630 |
623 lemma singleton_funion_left: |
631 lemma singleton_union_fset_left: |
624 shows "{|a|} |\<union>| S = finsert a S" |
632 shows "{|a|} |\<union>| S = insert_fset a S" |
625 by simp |
633 by simp |
626 |
634 |
627 lemma singleton_funion_right: |
635 lemma singleton_union_fset_right: |
628 shows "S |\<union>| {|a|} = finsert a S" |
636 shows "S |\<union>| {|a|} = insert_fset a S" |
629 by (subst sup.commute) simp |
637 by (subst sup.commute) simp |
630 |
638 |
631 lemma fin_funion: |
639 lemma in_union_fset: |
632 shows "x |\<in>| S |\<union>| T \<longleftrightarrow> x |\<in>| S \<or> x |\<in>| T" |
640 shows "x |\<in>| S |\<union>| T \<longleftrightarrow> x |\<in>| S \<or> x |\<in>| T" |
633 by (descending) (simp add: memb_def) |
641 by (descending) (simp add: memb_def) |
634 |
642 |
635 |
643 |
636 subsection {* fminus *} |
644 subsection {* minus_fset *} |
637 |
645 |
638 lemma fminus_fin: |
646 lemma minus_in_fset: |
639 shows "x |\<in>| (xs - ys) \<longleftrightarrow> x |\<in>| xs \<and> x |\<notin>| ys" |
647 shows "x |\<in>| (xs - ys) \<longleftrightarrow> x |\<in>| xs \<and> x |\<notin>| ys" |
640 by (descending) (simp add: diff_list_def memb_def) |
648 by (descending) (simp add: diff_list_def memb_def) |
641 |
649 |
642 lemma fminus_red: |
650 lemma minus_insert_fset: |
643 shows "finsert x xs - ys = (if x |\<in>| ys then xs - ys else finsert x (xs - ys))" |
651 shows "insert_fset x xs - ys = (if x |\<in>| ys then xs - ys else insert_fset x (xs - ys))" |
644 by (descending) (auto simp add: diff_list_def memb_def) |
652 by (descending) (auto simp add: diff_list_def memb_def) |
645 |
653 |
646 lemma fminus_red_fin[simp]: |
654 lemma minus_insert_in_fset[simp]: |
647 shows "x |\<in>| ys \<Longrightarrow> finsert x xs - ys = xs - ys" |
655 shows "x |\<in>| ys \<Longrightarrow> insert_fset x xs - ys = xs - ys" |
648 by (simp add: fminus_red) |
656 by (simp add: minus_insert_fset) |
649 |
657 |
650 lemma fminus_red_fnotin[simp]: |
658 lemma minus_insert_notin_fset[simp]: |
651 shows "x |\<notin>| ys \<Longrightarrow> finsert x xs - ys = finsert x (xs - ys)" |
659 shows "x |\<notin>| ys \<Longrightarrow> insert_fset x xs - ys = insert_fset x (xs - ys)" |
652 by (simp add: fminus_red) |
660 by (simp add: minus_insert_fset) |
653 |
661 |
654 lemma fin_fminus_fnotin: |
662 lemma in_fset_minus_fset_notin_fset: |
655 shows "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S" |
663 shows "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S" |
656 unfolding fin_set fminus_set |
664 unfolding in_fset minus_fset |
657 by blast |
665 by blast |
658 |
666 |
659 lemma fin_fnotin_fminus: |
667 lemma in_fset_notin_fset_minus_fset: |
660 shows "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S" |
668 shows "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S" |
661 unfolding fin_set fminus_set |
669 unfolding in_fset minus_fset |
662 by blast |
670 by blast |
663 |
671 |
664 |
672 |
665 section {* fdelete *} |
673 subsection {* remove_fset *} |
666 |
674 |
667 lemma fin_fdelete: |
675 lemma in_remove_fset: |
668 shows "x |\<in>| fdelete y S \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y" |
676 shows "x |\<in>| remove_fset y S \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y" |
669 by (descending) (simp add: memb_def) |
677 by (descending) (simp add: memb_def) |
670 |
678 |
671 lemma fnotin_fdelete: |
679 lemma notin_remove_fset: |
672 shows "x |\<notin>| fdelete x S" |
680 shows "x |\<notin>| remove_fset x S" |
673 by (descending) (simp add: memb_def) |
681 by (descending) (simp add: memb_def) |
674 |
682 |
675 lemma fnotin_fdelete_ident: |
683 lemma notin_remove_ident_fset: |
676 shows "x |\<notin>| S \<Longrightarrow> fdelete x S = S" |
684 shows "x |\<notin>| S \<Longrightarrow> remove_fset x S = S" |
677 by (descending) (simp add: memb_def) |
685 by (descending) (simp add: memb_def) |
678 |
686 |
679 lemma fset_fdelete_cases: |
687 lemma remove_fset_cases: |
680 shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = finsert x (fdelete x S))" |
688 shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = insert_fset x (remove_fset x S))" |
681 by (descending) (auto simp add: memb_def insert_absorb) |
689 by (descending) (auto simp add: memb_def insert_absorb) |
682 |
690 |
683 |
691 |
684 section {* finter *} |
692 subsection {* inter_fset *} |
685 |
693 |
686 lemma finter_empty_l: |
694 lemma inter_empty_fset_l: |
687 shows "{||} |\<inter>| S = {||}" |
695 shows "{||} |\<inter>| S = {||}" |
688 by simp |
696 by simp |
689 |
697 |
690 lemma finter_empty_r: |
698 lemma inter_empty_fset_r: |
691 shows "S |\<inter>| {||} = {||}" |
699 shows "S |\<inter>| {||} = {||}" |
692 by simp |
700 by simp |
693 |
701 |
694 lemma finter_finsert: |
702 lemma inter_insert_fset: |
695 shows "finsert x S |\<inter>| T = (if x |\<in>| T then finsert x (S |\<inter>| T) else S |\<inter>| T)" |
703 shows "insert_fset x S |\<inter>| T = (if x |\<in>| T then insert_fset x (S |\<inter>| T) else S |\<inter>| T)" |
696 by (descending) (auto simp add: inter_list_def memb_def) |
704 by (descending) (auto simp add: inter_list_def memb_def) |
697 |
705 |
698 lemma fin_finter: |
706 lemma in_inter_fset: |
699 shows "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T" |
707 shows "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T" |
700 by (descending) (simp add: inter_list_def memb_def) |
708 by (descending) (simp add: inter_list_def memb_def) |
701 |
709 |
702 |
710 |
703 subsection {* fsubset *} |
711 subsection {* subset_fset and psubset_fset *} |
704 |
712 |
705 lemma fsubseteq_set: |
713 lemma subset_fset: |
706 shows "xs |\<subseteq>| ys \<longleftrightarrow> fset xs \<subseteq> fset ys" |
714 shows "xs |\<subseteq>| ys \<longleftrightarrow> fset xs \<subseteq> fset ys" |
707 by (descending) (simp add: sub_list_def) |
715 by (descending) (simp add: sub_list_def) |
708 |
716 |
709 lemma fsubset_set: |
717 lemma psubset_fset: |
710 shows "xs |\<subset>| ys \<longleftrightarrow> fset xs \<subset> fset ys" |
718 shows "xs |\<subset>| ys \<longleftrightarrow> fset xs \<subset> fset ys" |
711 unfolding less_fset_def |
719 unfolding less_fset_def |
712 by (descending) (auto simp add: sub_list_def) |
720 by (descending) (auto simp add: sub_list_def) |
713 |
721 |
714 lemma fsubseteq_finsert: |
722 lemma subset_insert_fset: |
715 shows "(finsert x xs) |\<subseteq>| ys \<longleftrightarrow> x |\<in>| ys \<and> xs |\<subseteq>| ys" |
723 shows "(insert_fset x xs) |\<subseteq>| ys \<longleftrightarrow> x |\<in>| ys \<and> xs |\<subseteq>| ys" |
716 by (descending) (simp add: sub_list_def memb_def) |
724 by (descending) (simp add: sub_list_def memb_def) |
717 |
725 |
718 lemma fsubset_fin: |
726 lemma subset_in_fset: |
719 shows "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)" |
727 shows "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)" |
720 by (descending) (auto simp add: sub_list_def memb_def) |
728 by (descending) (auto simp add: sub_list_def memb_def) |
721 |
729 |
722 lemma fsubseteq_fempty: |
730 lemma subset_empty_fset: |
723 shows "xs |\<subseteq>| {||} \<longleftrightarrow> xs = {||}" |
731 shows "xs |\<subseteq>| {||} \<longleftrightarrow> xs = {||}" |
724 by (descending) (simp add: sub_list_def) |
732 by (descending) (simp add: sub_list_def) |
725 |
733 |
726 lemma not_fsubset_fnil: |
734 lemma not_psubset_fset_fnil: |
727 shows "\<not> xs |\<subset>| {||}" |
735 shows "\<not> xs |\<subset>| {||}" |
728 by (metis fset_simps(1) fsubset_set not_psubset_empty) |
736 by (metis fset_simps(1) psubset_fset not_psubset_empty) |
729 |
737 |
730 |
738 |
731 section {* fmap *} |
739 subsection {* map_fset *} |
732 |
740 |
733 lemma fmap_simps [simp]: |
741 lemma map_fset_simps [simp]: |
734 shows "fmap f {||} = {||}" |
742 shows "map_fset f {||} = {||}" |
735 and "fmap f (finsert x S) = finsert (f x) (fmap f S)" |
743 and "map_fset f (insert_fset x S) = insert_fset (f x) (map_fset f S)" |
736 by (descending, simp)+ |
744 by (descending, simp)+ |
737 |
745 |
738 lemma fmap_set_image [simp]: |
746 lemma map_fset_image [simp]: |
739 shows "fset (fmap f S) = f ` (fset S)" |
747 shows "fset (map_fset f S) = f ` (fset S)" |
740 by (descending) (simp) |
748 by (descending) (simp) |
741 |
749 |
742 lemma inj_fmap_eq_iff: |
750 lemma inj_map_fset_eq_iff: |
743 shows "inj f \<Longrightarrow> fmap f S = fmap f T \<longleftrightarrow> S = T" |
751 shows "inj f \<Longrightarrow> map_fset f S = map_fset f T \<longleftrightarrow> S = T" |
744 by (descending) (metis inj_vimage_image_eq list_eq.simps set_map) |
752 by (descending) (metis inj_vimage_image_eq list_eq.simps set_map) |
745 |
753 |
746 lemma fmap_funion: |
754 lemma map_union_fset: |
747 shows "fmap f (S |\<union>| T) = fmap f S |\<union>| fmap f T" |
755 shows "map_fset f (S |\<union>| T) = map_fset f S |\<union>| map_fset f T" |
748 by (descending) (simp) |
756 by (descending) (simp) |
749 |
757 |
750 |
758 |
751 subsection {* fcard *} |
759 subsection {* card_fset *} |
752 |
760 |
753 lemma fcard_set: |
761 lemma card_fset: |
754 shows "fcard xs = card (fset xs)" |
762 shows "card_fset xs = card (fset xs)" |
755 by (lifting card_list_def) |
763 by (lifting card_list_def) |
756 |
764 |
757 lemma fcard_finsert_if [simp]: |
765 lemma card_insert_fset_if [simp]: |
758 shows "fcard (finsert x S) = (if x |\<in>| S then fcard S else Suc (fcard S))" |
766 shows "card_fset (insert_fset x S) = (if x |\<in>| S then card_fset S else Suc (card_fset S))" |
759 by (descending) (auto simp add: card_list_def memb_def insert_absorb) |
767 by (descending) (auto simp add: card_list_def memb_def insert_absorb) |
760 |
768 |
761 lemma fcard_0[simp]: |
769 lemma card_fset_0[simp]: |
762 shows "fcard S = 0 \<longleftrightarrow> S = {||}" |
770 shows "card_fset S = 0 \<longleftrightarrow> S = {||}" |
763 by (descending) (simp add: card_list_def) |
771 by (descending) (simp add: card_list_def) |
764 |
772 |
765 lemma fcard_fempty[simp]: |
773 lemma card_empty_fset[simp]: |
766 shows "fcard {||} = 0" |
774 shows "card_fset {||} = 0" |
767 by (simp add: fcard_0) |
775 by (simp add: card_fset_0) |
768 |
776 |
769 lemma fcard_1: |
777 lemma card_fset_1: |
770 shows "fcard S = 1 \<longleftrightarrow> (\<exists>x. S = {|x|})" |
778 shows "card_fset S = 1 \<longleftrightarrow> (\<exists>x. S = {|x|})" |
771 by (descending) (auto simp add: card_list_def card_Suc_eq) |
779 by (descending) (auto simp add: card_list_def card_Suc_eq) |
772 |
780 |
773 lemma fcard_gt_0: |
781 lemma card_fset_gt_0: |
774 shows "x \<in> fset S \<Longrightarrow> 0 < fcard S" |
782 shows "x \<in> fset S \<Longrightarrow> 0 < card_fset S" |
775 by (descending) (auto simp add: card_list_def card_gt_0_iff) |
783 by (descending) (auto simp add: card_list_def card_gt_0_iff) |
776 |
784 |
777 lemma fcard_not_fin: |
785 lemma card_notin_fset: |
778 shows "(x |\<notin>| S) = (fcard (finsert x S) = Suc (fcard S))" |
786 shows "(x |\<notin>| S) = (card_fset (insert_fset x S) = Suc (card_fset S))" |
779 by (descending) (auto simp add: memb_def card_list_def insert_absorb) |
787 by (descending) (auto simp add: memb_def card_list_def insert_absorb) |
780 |
788 |
781 lemma fcard_suc: |
789 lemma card_fset_Suc: |
782 shows "fcard S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = finsert x T \<and> fcard T = n" |
790 shows "card_fset S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = insert_fset x T \<and> card_fset T = n" |
783 apply(descending) |
791 apply(descending) |
784 apply(auto simp add: card_list_def memb_def) |
792 apply(auto simp add: card_list_def memb_def dest!: card_eq_SucD) |
785 apply(drule card_eq_SucD) |
793 by (metis Diff_insert_absorb set_removeAll) |
786 apply(auto) |
794 |
787 apply(rule_tac x="b" in exI) |
795 lemma card_rsemove_fset: |
788 apply(rule_tac x="removeAll b S" in exI) |
796 shows "card_fset (remove_fset y S) = (if y |\<in>| S then card_fset S - 1 else card_fset S)" |
789 apply(auto) |
|
790 done |
|
791 |
|
792 lemma fcard_delete: |
|
793 shows "fcard (fdelete y S) = (if y |\<in>| S then fcard S - 1 else fcard S)" |
|
794 by (descending) (simp add: card_list_def memb_def) |
797 by (descending) (simp add: card_list_def memb_def) |
795 |
798 |
796 lemma fcard_suc_memb: |
799 lemma card_Suc_in_fset: |
797 shows "fcard A = Suc n \<Longrightarrow> \<exists>a. a |\<in>| A" |
800 shows "card_fset A = Suc n \<Longrightarrow> \<exists>a. a |\<in>| A" |
798 apply(descending) |
801 apply(descending) |
799 apply(simp add: card_list_def memb_def) |
802 apply(simp add: card_list_def memb_def) |
800 apply(drule card_eq_SucD) |
803 apply(drule card_eq_SucD) |
801 apply(auto) |
804 apply(auto) |
802 done |
805 done |
803 |
806 |
804 lemma fin_fcard_not_0: |
807 lemma in_fset_card_fset_not_0: |
805 shows "a |\<in>| A \<Longrightarrow> fcard A \<noteq> 0" |
808 shows "a |\<in>| A \<Longrightarrow> card_fset A \<noteq> 0" |
806 by (descending) (auto simp add: card_list_def memb_def) |
809 by (descending) (auto simp add: card_list_def memb_def) |
807 |
810 |
808 lemma fcard_mono: |
811 lemma card_fset_mono: |
809 shows "xs |\<subseteq>| ys \<Longrightarrow> fcard xs \<le> fcard ys" |
812 shows "xs |\<subseteq>| ys \<Longrightarrow> card_fset xs \<le> card_fset ys" |
810 unfolding fcard_set fsubseteq_set |
813 unfolding card_fset psubset_fset |
811 by (simp add: card_mono[OF finite_fset]) |
814 by (simp add: card_mono subset_fset) |
812 |
815 |
813 lemma fcard_fsubset_eq: |
816 lemma card_subset_fset_eq: |
814 shows "xs |\<subseteq>| ys \<Longrightarrow> fcard ys \<le> fcard xs \<Longrightarrow> xs = ys" |
817 shows "xs |\<subseteq>| ys \<Longrightarrow> card_fset ys \<le> card_fset xs \<Longrightarrow> xs = ys" |
815 unfolding fcard_set fsubseteq_set |
818 unfolding card_fset subset_fset |
816 by (auto dest: card_seteq[OF finite_fset] simp add: fset_cong) |
819 by (auto dest: card_seteq[OF finite_fset] simp add: fset_cong) |
817 |
820 |
818 lemma psubset_fcard_mono: |
821 lemma psubset_card_fset_mono: |
819 shows "xs |\<subset>| ys \<Longrightarrow> fcard xs < fcard ys" |
822 shows "xs |\<subset>| ys \<Longrightarrow> card_fset xs < card_fset ys" |
820 unfolding fcard_set fsubset_set |
823 unfolding card_fset subset_fset |
821 by (rule psubset_card_mono[OF finite_fset]) |
824 by (metis finite_fset psubset_fset psubset_card_mono) |
822 |
825 |
823 lemma fcard_funion_finter: |
826 lemma card_union_inter_fset: |
824 shows "fcard xs + fcard ys = fcard (xs |\<union>| ys) + fcard (xs |\<inter>| ys)" |
827 shows "card_fset xs + card_fset ys = card_fset (xs |\<union>| ys) + card_fset (xs |\<inter>| ys)" |
825 unfolding fcard_set funion_set finter_set |
828 unfolding card_fset union_fset inter_fset |
826 by (rule card_Un_Int[OF finite_fset finite_fset]) |
829 by (rule card_Un_Int[OF finite_fset finite_fset]) |
827 |
830 |
828 lemma fcard_funion_disjoint: |
831 lemma card_union_disjoint_fset: |
829 shows "xs |\<inter>| ys = {||} \<Longrightarrow> fcard (xs |\<union>| ys) = fcard xs + fcard ys" |
832 shows "xs |\<inter>| ys = {||} \<Longrightarrow> card_fset (xs |\<union>| ys) = card_fset xs + card_fset ys" |
830 unfolding fcard_set funion_set |
833 unfolding card_fset union_fset |
831 apply (rule card_Un_disjoint[OF finite_fset finite_fset]) |
834 apply (rule card_Un_disjoint[OF finite_fset finite_fset]) |
832 by (metis finter_set fset_simps(1)) |
835 by (metis inter_fset fset_simps(1)) |
833 |
836 |
834 lemma fcard_delete1_less: |
837 lemma card_remove_fset_less1: |
835 shows "x |\<in>| xs \<Longrightarrow> fcard (fdelete x xs) < fcard xs" |
838 shows "x |\<in>| xs \<Longrightarrow> card_fset (remove_fset x xs) < card_fset xs" |
836 unfolding fcard_set fin_set fdelete_set |
839 unfolding card_fset in_fset remove_fset |
837 by (rule card_Diff1_less[OF finite_fset]) |
840 by (rule card_Diff1_less[OF finite_fset]) |
838 |
841 |
839 lemma fcard_delete2_less: |
842 lemma card_rsemove_fset_less2: |
840 shows "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> fcard (fdelete y (fdelete x xs)) < fcard xs" |
843 shows "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> card_fset (remove_fset y (remove_fset x xs)) < card_fset xs" |
841 unfolding fcard_set fdelete_set fin_set |
844 unfolding card_fset remove_fset in_fset |
842 by (rule card_Diff2_less[OF finite_fset]) |
845 by (rule card_Diff2_less[OF finite_fset]) |
843 |
846 |
844 lemma fcard_delete1_le: |
847 lemma card_rsemove_fset_le1: |
845 shows "fcard (fdelete x xs) \<le> fcard xs" |
848 shows "card_fset (remove_fset x xs) \<le> card_fset xs" |
846 unfolding fdelete_set fcard_set |
849 unfolding remove_fset card_fset |
847 by (rule card_Diff1_le[OF finite_fset]) |
850 by (rule card_Diff1_le[OF finite_fset]) |
848 |
851 |
849 lemma fcard_psubset: |
852 lemma card_psubset_fset: |
850 shows "ys |\<subseteq>| xs \<Longrightarrow> fcard ys < fcard xs \<Longrightarrow> ys |\<subset>| xs" |
853 shows "ys |\<subseteq>| xs \<Longrightarrow> card_fset ys < card_fset xs \<Longrightarrow> ys |\<subset>| xs" |
851 unfolding fcard_set fsubseteq_set fsubset_set |
854 unfolding card_fset psubset_fset subset_fset |
852 by (rule card_psubset[OF finite_fset]) |
855 by (rule card_psubset[OF finite_fset]) |
853 |
856 |
854 lemma fcard_fmap_le: |
857 lemma card_map_fset_le: |
855 shows "fcard (fmap f xs) \<le> fcard xs" |
858 shows "card_fset (map_fset f xs) \<le> card_fset xs" |
856 unfolding fcard_set fmap_set_image |
859 unfolding card_fset map_fset_image |
857 by (rule card_image_le[OF finite_fset]) |
860 by (rule card_image_le[OF finite_fset]) |
858 |
861 |
859 lemma fcard_fminus_finsert[simp]: |
862 lemma card_minus_insert_fset[simp]: |
860 assumes "a |\<in>| A" and "a |\<notin>| B" |
863 assumes "a |\<in>| A" and "a |\<notin>| B" |
861 shows "fcard (A - finsert a B) = fcard (A - B) - 1" |
864 shows "card_fset (A - insert_fset a B) = card_fset (A - B) - 1" |
862 using assms |
865 using assms |
863 unfolding fin_set fcard_set fminus_set |
866 unfolding in_fset card_fset minus_fset |
864 by (simp add: card_Diff_insert[OF finite_fset]) |
867 by (simp add: card_Diff_insert[OF finite_fset]) |
865 |
868 |
866 lemma fcard_fminus_fsubset: |
869 lemma card_minus_subset_fset: |
867 assumes "B |\<subseteq>| A" |
870 assumes "B |\<subseteq>| A" |
868 shows "fcard (A - B) = fcard A - fcard B" |
871 shows "card_fset (A - B) = card_fset A - card_fset B" |
869 using assms |
872 using assms |
870 unfolding fsubseteq_set fcard_set fminus_set |
873 unfolding subset_fset card_fset minus_fset |
871 by (rule card_Diff_subset[OF finite_fset]) |
874 by (rule card_Diff_subset[OF finite_fset]) |
872 |
875 |
873 lemma fcard_fminus_subset_finter: |
876 lemma card_minus_subset_inter_fset: |
874 shows "fcard (A - B) = fcard A - fcard (A |\<inter>| B)" |
877 shows "card_fset (A - B) = card_fset A - card_fset (A |\<inter>| B)" |
875 unfolding finter_set fcard_set fminus_set |
878 unfolding inter_fset card_fset minus_fset |
876 by (rule card_Diff_subset_Int) (simp) |
879 by (rule card_Diff_subset_Int) (simp) |
877 |
880 |
878 |
881 |
879 section {* fconcat *} |
882 subsection {* concat_fset *} |
880 |
883 |
881 lemma fconcat_fempty: |
884 lemma concat_empty_fset: |
882 shows "fconcat {||} = {||}" |
885 shows "concat_fset {||} = {||}" |
883 by (lifting concat.simps(1)) |
886 by (lifting concat.simps(1)) |
884 |
887 |
885 lemma fconcat_finsert: |
888 lemma concat_insert_fset: |
886 shows "fconcat (finsert x S) = x |\<union>| fconcat S" |
889 shows "concat_fset (insert_fset x S) = x |\<union>| concat_fset S" |
887 by (lifting concat.simps(2)) |
890 by (lifting concat.simps(2)) |
888 |
891 |
889 lemma fconcat_finter: |
892 lemma concat_inter_fset: |
890 shows "fconcat (xs |\<union>| ys) = fconcat xs |\<union>| fconcat ys" |
893 shows "concat_fset (xs |\<union>| ys) = concat_fset xs |\<union>| concat_fset ys" |
891 by (lifting concat_append) |
894 by (lifting concat_append) |
892 |
895 |
893 |
896 |
894 section {* ffilter *} |
897 subsection {* filter_fset *} |
895 |
898 |
896 lemma subseteq_filter: |
899 lemma subset_filter_fset: |
897 shows "ffilter P xs <= ffilter Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)" |
900 shows "filter_fset P xs |\<subseteq>| filter_fset Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)" |
898 by (descending) (auto simp add: memb_def sub_list_def) |
901 by (descending) (auto simp add: memb_def sub_list_def) |
899 |
902 |
900 lemma eq_ffilter: |
903 lemma eq_filter_fset: |
901 shows "(ffilter P xs = ffilter Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)" |
904 shows "(filter_fset P xs = filter_fset Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)" |
902 by (descending) (auto simp add: memb_def) |
905 by (descending) (auto simp add: memb_def) |
903 |
906 |
904 lemma subset_ffilter: |
907 lemma psubset_filter_fset: |
905 shows "(\<And>x. x |\<in>| xs \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| xs & \<not> P x & Q x) \<Longrightarrow> ffilter P xs < ffilter Q xs" |
908 shows "(\<And>x. x |\<in>| xs \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| xs & \<not> P x & Q x) \<Longrightarrow> |
906 unfolding less_fset_def by (auto simp add: subseteq_filter eq_ffilter) |
909 filter_fset P xs |\<subset>| filter_fset Q xs" |
907 |
910 unfolding less_fset_def by (auto simp add: subset_filter_fset eq_filter_fset) |
908 |
911 |
909 subsection {* ffold *} |
912 |
910 |
913 subsection {* fold_fset *} |
911 lemma ffold_nil: |
914 |
912 shows "ffold f z {||} = z" |
915 lemma fold_empty_fset: |
916 shows "fold_fset f z {||} = z" |
|
913 by (descending) (simp) |
917 by (descending) (simp) |
914 |
918 |
915 lemma ffold_finsert: "ffold f z (finsert a A) = |
919 lemma fold_insert_fset: "fold_fset f z (insert_fset a A) = |
916 (if rsp_fold f then if a |\<in>| A then ffold f z A else f a (ffold f z A) else z)" |
920 (if rsp_fold f then if a |\<in>| A then fold_fset f z A else f a (fold_fset f z A) else z)" |
917 by (descending) (simp add: memb_def) |
921 by (descending) (simp add: memb_def) |
918 |
922 |
919 lemma fin_commute_ffold: |
923 lemma in_commute_fold_fset: |
920 "\<lbrakk>rsp_fold f; h |\<in>| b\<rbrakk> \<Longrightarrow> ffold f z b = f h (ffold f z (fdelete h b))" |
924 "\<lbrakk>rsp_fold f; h |\<in>| b\<rbrakk> \<Longrightarrow> fold_fset f z b = f h (fold_fset f z (remove_fset h b))" |
921 by (descending) (simp add: memb_def memb_commute_ffold_list) |
925 by (descending) (simp add: memb_def memb_commute_fold_list) |
922 |
926 |
923 |
927 |
924 subsection {* Choice in fsets *} |
928 subsection {* Choice in fsets *} |
925 |
929 |
926 lemma fset_choice: |
930 lemma fset_choice: |
932 by (auto simp add: memb_def Ball_def) |
936 by (auto simp add: memb_def Ball_def) |
933 |
937 |
934 |
938 |
935 section {* Induction and Cases rules for fsets *} |
939 section {* Induction and Cases rules for fsets *} |
936 |
940 |
937 lemma fset_exhaust [case_names fempty finsert, cases type: fset]: |
941 lemma fset_exhaust [case_names empty_fset insert_fset, cases type: fset]: |
938 assumes fempty_case: "S = {||} \<Longrightarrow> P" |
942 assumes empty_fset_case: "S = {||} \<Longrightarrow> P" |
939 and finsert_case: "\<And>x S'. S = finsert x S' \<Longrightarrow> P" |
943 and insert_fset_case: "\<And>x S'. S = insert_fset x S' \<Longrightarrow> P" |
940 shows "P" |
944 shows "P" |
941 using assms by (lifting list.exhaust) |
945 using assms by (lifting list.exhaust) |
942 |
946 |
943 lemma fset_induct [case_names fempty finsert]: |
947 lemma fset_induct [case_names empty_fset insert_fset]: |
944 assumes fempty_case: "P {||}" |
948 assumes empty_fset_case: "P {||}" |
945 and finsert_case: "\<And>x S. P S \<Longrightarrow> P (finsert x S)" |
949 and insert_fset_case: "\<And>x S. P S \<Longrightarrow> P (insert_fset x S)" |
946 shows "P S" |
950 shows "P S" |
947 using assms |
951 using assms |
948 by (descending) (blast intro: list.induct) |
952 by (descending) (blast intro: list.induct) |
949 |
953 |
950 lemma fset_induct_stronger [case_names fempty finsert, induct type: fset]: |
954 lemma fset_induct_stronger [case_names empty_fset insert_fset, induct type: fset]: |
951 assumes fempty_case: "P {||}" |
955 assumes empty_fset_case: "P {||}" |
952 and finsert_case: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (finsert x S)" |
956 and insert_fset_case: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (insert_fset x S)" |
953 shows "P S" |
957 shows "P S" |
954 proof(induct S rule: fset_induct) |
958 proof(induct S rule: fset_induct) |
955 case fempty |
959 case empty_fset |
956 show "P {||}" using fempty_case by simp |
960 show "P {||}" using empty_fset_case by simp |
957 next |
961 next |
958 case (finsert x S) |
962 case (insert_fset x S) |
959 have "P S" by fact |
963 have "P S" by fact |
960 then show "P (finsert x S)" using finsert_case |
964 then show "P (insert_fset x S)" using insert_fset_case |
961 by (cases "x |\<in>| S") (simp_all) |
965 by (cases "x |\<in>| S") (simp_all) |
962 qed |
966 qed |
963 |
967 |
964 lemma fcard_induct: |
968 lemma fset_card_induct: |
965 assumes fempty_case: "P {||}" |
969 assumes empty_fset_case: "P {||}" |
966 and fcard_Suc_case: "\<And>S T. Suc (fcard S) = (fcard T) \<Longrightarrow> P S \<Longrightarrow> P T" |
970 and card_fset_Suc_case: "\<And>S T. Suc (card_fset S) = (card_fset T) \<Longrightarrow> P S \<Longrightarrow> P T" |
967 shows "P S" |
971 shows "P S" |
968 proof (induct S) |
972 proof (induct S) |
969 case fempty |
973 case empty_fset |
970 show "P {||}" by (rule fempty_case) |
974 show "P {||}" by (rule empty_fset_case) |
971 next |
975 next |
972 case (finsert x S) |
976 case (insert_fset x S) |
973 have h: "P S" by fact |
977 have h: "P S" by fact |
974 have "x |\<notin>| S" by fact |
978 have "x |\<notin>| S" by fact |
975 then have "Suc (fcard S) = fcard (finsert x S)" |
979 then have "Suc (card_fset S) = card_fset (insert_fset x S)" |
976 using fcard_suc by auto |
980 using card_fset_Suc by auto |
977 then show "P (finsert x S)" |
981 then show "P (insert_fset x S)" |
978 using h fcard_Suc_case by simp |
982 using h card_fset_Suc_case by simp |
979 qed |
983 qed |
980 |
984 |
981 lemma fset_raw_strong_cases: |
985 lemma fset_raw_strong_cases: |
982 obtains "xs = []" |
986 obtains "xs = []" |
983 | x ys where "\<not> memb x ys" and "xs \<approx> x # ys" |
987 | x ys where "\<not> memb x ys" and "xs \<approx> x # ys" |
1013 qed |
1017 qed |
1014 |
1018 |
1015 |
1019 |
1016 lemma fset_strong_cases: |
1020 lemma fset_strong_cases: |
1017 obtains "xs = {||}" |
1021 obtains "xs = {||}" |
1018 | x ys where "x |\<notin>| ys" and "xs = finsert x ys" |
1022 | x ys where "x |\<notin>| ys" and "xs = insert_fset x ys" |
1019 by (lifting fset_raw_strong_cases) |
1023 by (lifting fset_raw_strong_cases) |
1020 |
1024 |
1021 |
1025 |
1022 lemma fset_induct2: |
1026 lemma fset_induct2: |
1023 "P {||} {||} \<Longrightarrow> |
1027 "P {||} {||} \<Longrightarrow> |
1024 (\<And>x xs. x |\<notin>| xs \<Longrightarrow> P (finsert x xs) {||}) \<Longrightarrow> |
1028 (\<And>x xs. x |\<notin>| xs \<Longrightarrow> P (insert_fset x xs) {||}) \<Longrightarrow> |
1025 (\<And>y ys. y |\<notin>| ys \<Longrightarrow> P {||} (finsert y ys)) \<Longrightarrow> |
1029 (\<And>y ys. y |\<notin>| ys \<Longrightarrow> P {||} (insert_fset y ys)) \<Longrightarrow> |
1026 (\<And>x xs y ys. \<lbrakk>P xs ys; x |\<notin>| xs; y |\<notin>| ys\<rbrakk> \<Longrightarrow> P (finsert x xs) (finsert y ys)) \<Longrightarrow> |
1030 (\<And>x xs y ys. \<lbrakk>P xs ys; x |\<notin>| xs; y |\<notin>| ys\<rbrakk> \<Longrightarrow> P (insert_fset x xs) (insert_fset y ys)) \<Longrightarrow> |
1027 P xsa ysa" |
1031 P xsa ysa" |
1028 apply (induct xsa arbitrary: ysa) |
1032 apply (induct xsa arbitrary: ysa) |
1029 apply (induct_tac x rule: fset_induct_stronger) |
1033 apply (induct_tac x rule: fset_induct_stronger) |
1030 apply simp_all |
1034 apply simp_all |
1031 apply (induct_tac xa rule: fset_induct_stronger) |
1035 apply (induct_tac xa rule: fset_induct_stronger) |
1048 | "\<lbrakk>xs1 \<approx>2 xs2; xs2 \<approx>2 xs3\<rbrakk> \<Longrightarrow> xs1 \<approx>2 xs3" |
1052 | "\<lbrakk>xs1 \<approx>2 xs2; xs2 \<approx>2 xs3\<rbrakk> \<Longrightarrow> xs1 \<approx>2 xs3" |
1049 |
1053 |
1050 lemma list_eq2_refl: |
1054 lemma list_eq2_refl: |
1051 shows "xs \<approx>2 xs" |
1055 shows "xs \<approx>2 xs" |
1052 by (induct xs) (auto intro: list_eq2.intros) |
1056 by (induct xs) (auto intro: list_eq2.intros) |
1053 |
|
1054 lemma list_eq2_set: |
|
1055 assumes a: "xs \<approx>2 ys" |
|
1056 shows "set xs = set ys" |
|
1057 using a by (induct) (auto) |
|
1058 |
|
1059 lemma list_eq2_card: |
|
1060 assumes a: "xs \<approx>2 ys" |
|
1061 shows "card_list xs = card_list ys" |
|
1062 using a |
|
1063 apply(induct) |
|
1064 apply(auto simp add: card_list_def) |
|
1065 apply(metis insert_commute) |
|
1066 apply(metis list_eq2_set) |
|
1067 done |
|
1068 |
|
1069 lemma list_eq2_equiv1: |
|
1070 assumes a: "xs \<approx>2 ys" |
|
1071 shows "xs \<approx> ys" |
|
1072 using a by (induct) (auto) |
|
1073 |
|
1074 |
1057 |
1075 lemma cons_delete_list_eq2: |
1058 lemma cons_delete_list_eq2: |
1076 shows "(a # (removeAll a A)) \<approx>2 (if memb a A then A else a # A)" |
1059 shows "(a # (removeAll a A)) \<approx>2 (if memb a A then A else a # A)" |
1077 apply (induct A) |
1060 apply (induct A) |
1078 apply (simp add: memb_def list_eq2_refl) |
1061 apply (simp add: memb_def list_eq2_refl) |
1090 lemma memb_delete_list_eq2: |
1073 lemma memb_delete_list_eq2: |
1091 assumes a: "memb e r" |
1074 assumes a: "memb e r" |
1092 shows "(e # removeAll e r) \<approx>2 r" |
1075 shows "(e # removeAll e r) \<approx>2 r" |
1093 using a cons_delete_list_eq2[of e r] |
1076 using a cons_delete_list_eq2[of e r] |
1094 by simp |
1077 by simp |
1095 |
|
1096 (* |
|
1097 lemma list_eq2_equiv2: |
|
1098 assumes a: "xs \<approx> ys" |
|
1099 shows "xs \<approx>2 ys" |
|
1100 using a |
|
1101 apply(induct xs arbitrary: ys taking: "card o set" rule: measure_induct) |
|
1102 apply(simp) |
|
1103 apply(case_tac x) |
|
1104 apply(simp) |
|
1105 apply(auto intro: list_eq2.intros)[1] |
|
1106 apply(simp) |
|
1107 apply(case_tac "a \<in> set list") |
|
1108 apply(simp add: insert_absorb) |
|
1109 apply(drule_tac x="removeAll a ys" in spec) |
|
1110 apply(drule mp) |
|
1111 apply(simp) |
|
1112 apply(subgoal_tac "0 < card (set ys)") |
|
1113 apply(simp) |
|
1114 apply(metis length_pos_if_in_set length_remdups_card_conv set_remdups) |
|
1115 apply(simp) |
|
1116 apply(clarify) |
|
1117 apply(drule_tac x="removeAll a list" in spec) |
|
1118 apply(drule mp) |
|
1119 apply(simp) |
|
1120 oops |
|
1121 *) |
|
1122 |
1078 |
1123 lemma list_eq2_equiv: |
1079 lemma list_eq2_equiv: |
1124 "(l \<approx> r) \<longleftrightarrow> (list_eq2 l r)" |
1080 "(l \<approx> r) \<longleftrightarrow> (list_eq2 l r)" |
1125 proof |
1081 proof |
1126 show "list_eq2 l r \<Longrightarrow> l \<approx> r" by (induct rule: list_eq2.induct) auto |
1082 show "list_eq2 l r \<Longrightarrow> l \<approx> r" by (induct rule: list_eq2.induct) auto |
1149 then obtain a where e: "memb a l" by auto |
1105 then obtain a where e: "memb a l" by auto |
1150 then have e': "memb a r" using list_eq.simps[simplified memb_def[symmetric], of l r] b |
1106 then have e': "memb a r" using list_eq.simps[simplified memb_def[symmetric], of l r] b |
1151 unfolding memb_def by auto |
1107 unfolding memb_def by auto |
1152 have f: "card_list (removeAll a l) = m" using e d by (simp add: card_list_def memb_def) |
1108 have f: "card_list (removeAll a l) = m" using e d by (simp add: card_list_def memb_def) |
1153 have g: "removeAll a l \<approx> removeAll a r" using removeAll_rsp b by simp |
1109 have g: "removeAll a l \<approx> removeAll a r" using removeAll_rsp b by simp |
1154 have "list_eq2 (removeAll a l) (removeAll a r)" by (rule Suc.hyps[OF f g]) |
1110 have "(removeAll a l) \<approx>2 (removeAll a r)" by (rule Suc.hyps[OF f g]) |
1155 then have h: "list_eq2 (a # removeAll a l) (a # removeAll a r)" by (rule list_eq2.intros(5)) |
1111 then have h: "(a # removeAll a l) \<approx>2 (a # removeAll a r)" by (rule list_eq2.intros(5)) |
1156 have i: "list_eq2 l (a # removeAll a l)" |
1112 have i: "l \<approx>2 (a # removeAll a l)" |
1157 by (rule list_eq2.intros(3)[OF memb_delete_list_eq2[OF e]]) |
1113 by (rule list_eq2.intros(3)[OF memb_delete_list_eq2[OF e]]) |
1158 have "list_eq2 l (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h]) |
1114 have "l \<approx>2 (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h]) |
1159 then show ?case using list_eq2.intros(6)[OF _ memb_delete_list_eq2[OF e']] by simp |
1115 then show ?case using list_eq2.intros(6)[OF _ memb_delete_list_eq2[OF e']] by simp |
1160 qed |
1116 qed |
1161 } |
1117 } |
1162 then show "l \<approx> r \<Longrightarrow> list_eq2 l r" by blast |
1118 then show "l \<approx> r \<Longrightarrow> l \<approx>2 r" by blast |
1163 qed |
1119 qed |
1164 |
1120 |
1165 |
1121 |
1166 (* We cannot write it as "assumes .. shows" since Isabelle changes |
1122 (* We cannot write it as "assumes .. shows" since Isabelle changes |
1167 the quantifiers to schematic variables and reintroduces them in |
1123 the quantifiers to schematic variables and reintroduces them in |
1168 a different order *) |
1124 a different order *) |
1169 lemma fset_eq_cases: |
1125 lemma fset_eq_cases: |
1170 "\<lbrakk>a1 = a2; |
1126 "\<lbrakk>a1 = a2; |
1171 \<And>a b xs. \<lbrakk>a1 = finsert a (finsert b xs); a2 = finsert b (finsert a xs)\<rbrakk> \<Longrightarrow> P; |
1127 \<And>a b xs. \<lbrakk>a1 = insert_fset a (insert_fset b xs); a2 = insert_fset b (insert_fset a xs)\<rbrakk> \<Longrightarrow> P; |
1172 \<lbrakk>a1 = {||}; a2 = {||}\<rbrakk> \<Longrightarrow> P; \<And>xs ys. \<lbrakk>a1 = ys; a2 = xs; xs = ys\<rbrakk> \<Longrightarrow> P; |
1128 \<lbrakk>a1 = {||}; a2 = {||}\<rbrakk> \<Longrightarrow> P; \<And>xs ys. \<lbrakk>a1 = ys; a2 = xs; xs = ys\<rbrakk> \<Longrightarrow> P; |
1173 \<And>a xs. \<lbrakk>a1 = finsert a (finsert a xs); a2 = finsert a xs\<rbrakk> \<Longrightarrow> P; |
1129 \<And>a xs. \<lbrakk>a1 = insert_fset a (insert_fset a xs); a2 = insert_fset a xs\<rbrakk> \<Longrightarrow> P; |
1174 \<And>xs ys a. \<lbrakk>a1 = finsert a xs; a2 = finsert a ys; xs = ys\<rbrakk> \<Longrightarrow> P; |
1130 \<And>xs ys a. \<lbrakk>a1 = insert_fset a xs; a2 = insert_fset a ys; xs = ys\<rbrakk> \<Longrightarrow> P; |
1175 \<And>xs1 xs2 xs3. \<lbrakk>a1 = xs1; a2 = xs3; xs1 = xs2; xs2 = xs3\<rbrakk> \<Longrightarrow> P\<rbrakk> |
1131 \<And>xs1 xs2 xs3. \<lbrakk>a1 = xs1; a2 = xs3; xs1 = xs2; xs2 = xs3\<rbrakk> \<Longrightarrow> P\<rbrakk> |
1176 \<Longrightarrow> P" |
1132 \<Longrightarrow> P" |
1177 by (lifting list_eq2.cases[simplified list_eq2_equiv[symmetric]]) |
1133 by (lifting list_eq2.cases[simplified list_eq2_equiv[symmetric]]) |
1178 |
1134 |
1179 lemma fset_eq_induct: |
1135 lemma fset_eq_induct: |
1180 assumes "x1 = x2" |
1136 assumes "x1 = x2" |
1181 and "\<And>a b xs. P (finsert a (finsert b xs)) (finsert b (finsert a xs))" |
1137 and "\<And>a b xs. P (insert_fset a (insert_fset b xs)) (insert_fset b (insert_fset a xs))" |
1182 and "P {||} {||}" |
1138 and "P {||} {||}" |
1183 and "\<And>xs ys. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P ys xs" |
1139 and "\<And>xs ys. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P ys xs" |
1184 and "\<And>a xs. P (finsert a (finsert a xs)) (finsert a xs)" |
1140 and "\<And>a xs. P (insert_fset a (insert_fset a xs)) (insert_fset a xs)" |
1185 and "\<And>xs ys a. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P (finsert a xs) (finsert a ys)" |
1141 and "\<And>xs ys a. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P (insert_fset a xs) (insert_fset a ys)" |
1186 and "\<And>xs1 xs2 xs3. \<lbrakk>xs1 = xs2; P xs1 xs2; xs2 = xs3; P xs2 xs3\<rbrakk> \<Longrightarrow> P xs1 xs3" |
1142 and "\<And>xs1 xs2 xs3. \<lbrakk>xs1 = xs2; P xs1 xs2; xs2 = xs3; P xs2 xs3\<rbrakk> \<Longrightarrow> P xs1 xs3" |
1187 shows "P x1 x2" |
1143 shows "P x1 x2" |
1188 using assms |
1144 using assms |
1189 by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]]) |
1145 by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]]) |
1190 |
1146 |
1191 |
1147 lemma list_all2_refl'': |
1192 lemma list_all2_refl: |
|
1193 assumes q: "equivp R" |
1148 assumes q: "equivp R" |
1194 shows "(list_all2 R) r r" |
1149 shows "(list_all2 R) r r" |
1195 by (rule list_all2_refl) (metis equivp_def q) |
1150 by (rule list_all2_refl) (metis equivp_def q) |
1196 |
1151 |
1197 lemma compose_list_refl2: |
1152 lemma compose_list_refl2: |
1198 assumes q: "equivp R" |
1153 assumes q: "equivp R" |
1199 shows "(list_all2 R OOO op \<approx>) r r" |
1154 shows "(list_all2 R OOO op \<approx>) r r" |
1200 proof |
1155 proof |
1201 have *: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp]) |
1156 have *: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp]) |
1202 show "list_all2 R r r" by (rule list_all2_refl[OF q]) |
1157 show "list_all2 R r r" by (rule list_all2_refl''[OF q]) |
1203 with * show "(op \<approx> OO list_all2 R) r r" .. |
1158 with * show "(op \<approx> OO list_all2 R) r r" .. |
1204 qed |
1159 qed |
1205 |
1160 |
1206 lemma quotient_compose_list_g: |
1161 lemma quotient_compose_list_g: |
1207 assumes q: "Quotient R Abs Rep" |
1162 assumes q: "Quotient R Abs Rep" |
1212 proof (intro conjI allI) |
1167 proof (intro conjI allI) |
1213 fix a r s |
1168 fix a r s |
1214 show "abs_fset (map Abs (map Rep (rep_fset a))) = a" |
1169 show "abs_fset (map Abs (map Rep (rep_fset a))) = a" |
1215 by (simp add: abs_o_rep[OF q] Quotient_abs_rep[OF Quotient_fset] map_id) |
1170 by (simp add: abs_o_rep[OF q] Quotient_abs_rep[OF Quotient_fset] map_id) |
1216 have b: "list_all2 R (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
1171 have b: "list_all2 R (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
1217 by (rule list_all2_refl[OF e]) |
1172 by (rule list_all2_refl''[OF e]) |
1218 have c: "(op \<approx> OO list_all2 R) (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
1173 have c: "(op \<approx> OO list_all2 R) (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
1219 by (rule, rule equivp_reflp[OF fset_equivp]) (rule b) |
1174 by (rule, rule equivp_reflp[OF fset_equivp]) (rule b) |
1220 show "(list_all2 R OOO op \<approx>) (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
1175 show "(list_all2 R OOO op \<approx>) (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
1221 by (rule, rule list_all2_refl[OF e]) (rule c) |
1176 by (rule, rule list_all2_refl''[OF e]) (rule c) |
1222 show "(list_all2 R OOO op \<approx>) r s = ((list_all2 R OOO op \<approx>) r r \<and> |
1177 show "(list_all2 R OOO op \<approx>) r s = ((list_all2 R OOO op \<approx>) r r \<and> |
1223 (list_all2 R OOO op \<approx>) s s \<and> abs_fset (map Abs r) = abs_fset (map Abs s))" |
1178 (list_all2 R OOO op \<approx>) s s \<and> abs_fset (map Abs r) = abs_fset (map Abs s))" |
1224 proof (intro iffI conjI) |
1179 proof (intro iffI conjI) |
1225 show "(list_all2 R OOO op \<approx>) r r" by (rule compose_list_refl2[OF e]) |
1180 show "(list_all2 R OOO op \<approx>) r r" by (rule compose_list_refl2[OF e]) |
1226 show "(list_all2 R OOO op \<approx>) s s" by (rule compose_list_refl2[OF e]) |
1181 show "(list_all2 R OOO op \<approx>) s s" by (rule compose_list_refl2[OF e]) |
1235 have f: "map Abs r = map Abs b" |
1190 have f: "map Abs r = map Abs b" |
1236 using Quotient_rel[OF list_quotient[OF q]] c by blast |
1191 using Quotient_rel[OF list_quotient[OF q]] c by blast |
1237 have "map Abs ba = map Abs s" |
1192 have "map Abs ba = map Abs s" |
1238 using Quotient_rel[OF list_quotient[OF q]] e by blast |
1193 using Quotient_rel[OF list_quotient[OF q]] e by blast |
1239 then have g: "map Abs s = map Abs ba" by simp |
1194 then have g: "map Abs s = map Abs ba" by simp |
1240 then show "map Abs r \<approx> map Abs s" using d f map_rel_cong by simp |
1195 then show "map Abs r \<approx> map Abs s" using d f map_list_eq_cong by simp |
1241 qed |
1196 qed |
1242 then show "abs_fset (map Abs r) = abs_fset (map Abs s)" |
1197 then show "abs_fset (map Abs r) = abs_fset (map Abs s)" |
1243 using Quotient_rel[OF Quotient_fset] by blast |
1198 using Quotient_rel[OF Quotient_fset] by blast |
1244 next |
1199 next |
1245 assume a: "(list_all2 R OOO op \<approx>) r r \<and> (list_all2 R OOO op \<approx>) s s |
1200 assume a: "(list_all2 R OOO op \<approx>) r r \<and> (list_all2 R OOO op \<approx>) s s |
1246 \<and> abs_fset (map Abs r) = abs_fset (map Abs s)" |
1201 \<and> abs_fset (map Abs r) = abs_fset (map Abs s)" |
1247 then have s: "(list_all2 R OOO op \<approx>) s s" by simp |
1202 then have s: "(list_all2 R OOO op \<approx>) s s" by simp |
1248 have d: "map Abs r \<approx> map Abs s" |
1203 have d: "map Abs r \<approx> map Abs s" |
1249 by (subst Quotient_rel[OF Quotient_fset]) (simp add: a) |
1204 by (subst Quotient_rel[OF Quotient_fset]) (simp add: a) |
1250 have b: "map Rep (map Abs r) \<approx> map Rep (map Abs s)" |
1205 have b: "map Rep (map Abs r) \<approx> map Rep (map Abs s)" |
1251 by (rule map_rel_cong[OF d]) |
1206 by (rule map_list_eq_cong[OF d]) |
1252 have y: "list_all2 R (map Rep (map Abs s)) s" |
1207 have y: "list_all2 R (map Rep (map Abs s)) s" |
1253 by (fact rep_abs_rsp_left[OF list_quotient[OF q], OF list_all2_refl[OF e, of s]]) |
1208 by (fact rep_abs_rsp_left[OF list_quotient[OF q], OF list_all2_refl''[OF e, of s]]) |
1254 have c: "(op \<approx> OO list_all2 R) (map Rep (map Abs r)) s" |
1209 have c: "(op \<approx> OO list_all2 R) (map Rep (map Abs r)) s" |
1255 by (rule pred_compI) (rule b, rule y) |
1210 by (rule pred_compI) (rule b, rule y) |
1256 have z: "list_all2 R r (map Rep (map Abs r))" |
1211 have z: "list_all2 R r (map Rep (map Abs r))" |
1257 by (fact rep_abs_rsp[OF list_quotient[OF q], OF list_all2_refl[OF e, of r]]) |
1212 by (fact rep_abs_rsp[OF list_quotient[OF q], OF list_all2_refl''[OF e, of r]]) |
1258 then show "(list_all2 R OOO op \<approx>) r s" |
1213 then show "(list_all2 R OOO op \<approx>) r s" |
1259 using a c pred_compI by simp |
1214 using a c pred_compI by simp |
1260 qed |
1215 qed |
1261 qed |
1216 qed |
1262 |
1217 |