| author | blanchet | 
| Fri, 15 Feb 2013 10:13:04 +0100 | |
| changeset 51146 | 754127b3af23 | 
| parent 50227 | 01d545993e8c | 
| child 51376 | 8e38ff09864a | 
| permissions | -rw-r--r-- | 
| 47660 | 1  | 
(* Title: HOL/Quotient_Examples/Lift_FSet.thy  | 
2  | 
Author: Brian Huffman, TU Munich  | 
|
3  | 
*)  | 
|
4  | 
||
5  | 
header {* Lifting and transfer with a finite set type *}
 | 
|
6  | 
||
7  | 
theory Lift_FSet  | 
|
8  | 
imports "~~/src/HOL/Library/Quotient_List"  | 
|
9  | 
begin  | 
|
10  | 
||
11  | 
subsection {* Equivalence relation and quotient type definition *}
 | 
|
12  | 
||
13  | 
definition list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"  | 
|
14  | 
where [simp]: "list_eq xs ys \<longleftrightarrow> set xs = set ys"  | 
|
15  | 
||
16  | 
lemma reflp_list_eq: "reflp list_eq"  | 
|
17  | 
unfolding reflp_def by simp  | 
|
18  | 
||
19  | 
lemma symp_list_eq: "symp list_eq"  | 
|
20  | 
unfolding symp_def by simp  | 
|
21  | 
||
22  | 
lemma transp_list_eq: "transp list_eq"  | 
|
23  | 
unfolding transp_def by simp  | 
|
24  | 
||
25  | 
lemma equivp_list_eq: "equivp list_eq"  | 
|
26  | 
by (intro equivpI reflp_list_eq symp_list_eq transp_list_eq)  | 
|
27  | 
||
28  | 
quotient_type 'a fset = "'a list" / "list_eq"  | 
|
29  | 
by (rule equivp_list_eq)  | 
|
30  | 
||
31  | 
subsection {* Lifted constant definitions *}
 | 
|
32  | 
||
33  | 
lift_definition fnil :: "'a fset" is "[]"  | 
|
34  | 
by simp  | 
|
35  | 
||
36  | 
lift_definition fcons :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is Cons  | 
|
37  | 
by simp  | 
|
38  | 
||
39  | 
lift_definition fappend :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is append  | 
|
40  | 
by simp  | 
|
41  | 
||
42  | 
lift_definition fmap :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" is map
 | 
|
43  | 
by simp  | 
|
44  | 
||
45  | 
lift_definition ffilter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is filter
 | 
|
46  | 
by simp  | 
|
47  | 
||
48  | 
lift_definition fset :: "'a fset \<Rightarrow> 'a set" is set  | 
|
49  | 
by simp  | 
|
50  | 
||
51  | 
text {* Constants with nested types (like concat) yield a more
 | 
|
52  | 
complicated proof obligation. *}  | 
|
53  | 
||
54  | 
lemma list_all2_cr_fset:  | 
|
55  | 
"list_all2 cr_fset xs ys \<longleftrightarrow> map abs_fset xs = ys"  | 
|
56  | 
unfolding cr_fset_def  | 
|
57  | 
apply safe  | 
|
58  | 
apply (erule list_all2_induct, simp, simp)  | 
|
59  | 
apply (simp add: list_all2_map2 List.list_all2_refl)  | 
|
60  | 
done  | 
|
61  | 
||
62  | 
lemma abs_fset_eq_iff: "abs_fset xs = abs_fset ys \<longleftrightarrow> list_eq xs ys"  | 
|
63  | 
using Quotient_rel [OF Quotient_fset] by simp  | 
|
64  | 
||
65  | 
lift_definition fconcat :: "'a fset fset \<Rightarrow> 'a fset" is concat  | 
|
66  | 
proof -  | 
|
67  | 
fix xss yss :: "'a list list"  | 
|
68  | 
assume "(list_all2 cr_fset OO list_eq OO (list_all2 cr_fset)\<inverse>\<inverse>) xss yss"  | 
|
69  | 
then obtain uss vss where  | 
|
70  | 
"list_all2 cr_fset xss uss" and "list_eq uss vss" and  | 
|
71  | 
"list_all2 cr_fset yss vss" by clarsimp  | 
|
72  | 
hence "list_eq (map abs_fset xss) (map abs_fset yss)"  | 
|
73  | 
unfolding list_all2_cr_fset by simp  | 
|
74  | 
thus "list_eq (concat xss) (concat yss)"  | 
|
75  | 
apply (simp add: set_eq_iff image_def)  | 
|
76  | 
apply safe  | 
|
77  | 
apply (rename_tac xs, drule_tac x="abs_fset xs" in spec)  | 
|
78  | 
apply (drule iffD1, fast, clarsimp simp add: abs_fset_eq_iff, fast)  | 
|
79  | 
apply (rename_tac xs, drule_tac x="abs_fset xs" in spec)  | 
|
80  | 
apply (drule iffD2, fast, clarsimp simp add: abs_fset_eq_iff, fast)  | 
|
81  | 
done  | 
|
82  | 
qed  | 
|
83  | 
||
| 
47937
 
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
 
kuncar 
parents: 
47676 
diff
changeset
 | 
84  | 
text {* We can export code: *}
 | 
| 
 
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
 
kuncar 
parents: 
47676 
diff
changeset
 | 
85  | 
|
| 
 
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
 
kuncar 
parents: 
47676 
diff
changeset
 | 
86  | 
export_code fnil fcons fappend fmap ffilter fset in SML  | 
| 
 
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
 
kuncar 
parents: 
47676 
diff
changeset
 | 
87  | 
|
| 47660 | 88  | 
text {* Note that the generated transfer rule contains a composition
 | 
89  | 
of relations. The transfer rule is not yet very useful in this form. *}  | 
|
90  | 
||
91  | 
lemma "(list_all2 cr_fset OO cr_fset ===> cr_fset) concat fconcat"  | 
|
92  | 
by (fact fconcat.transfer)  | 
|
93  | 
||
94  | 
||
95  | 
subsection {* Using transfer with type @{text "fset"} *}
 | 
|
96  | 
||
97  | 
text {* The correspondence relation @{text "cr_fset"} can only relate
 | 
|
98  | 
  @{text "list"} and @{text "fset"} types with the same element type.
 | 
|
99  | 
  To relate nested types like @{text "'a list list"} and
 | 
|
100  | 
  @{text "'a fset fset"}, we define a parameterized version of the
 | 
|
| 50227 | 101  | 
  correspondence relation, @{text "pcr_fset"}. *}
 | 
102  | 
||
103  | 
thm pcr_fset_def  | 
|
| 47660 | 104  | 
|
| 50227 | 105  | 
lemma right_unique_pcr_fset [transfer_rule]:  | 
106  | 
"right_unique A \<Longrightarrow> right_unique (pcr_fset A)"  | 
|
107  | 
unfolding pcr_fset_def  | 
|
| 47660 | 108  | 
by (intro right_unique_OO right_unique_list_all2 fset.right_unique)  | 
109  | 
||
| 50227 | 110  | 
lemma right_total_pcr_fset [transfer_rule]:  | 
111  | 
"right_total A \<Longrightarrow> right_total (pcr_fset A)"  | 
|
112  | 
unfolding pcr_fset_def  | 
|
| 47660 | 113  | 
by (intro right_total_OO right_total_list_all2 fset.right_total)  | 
114  | 
||
| 50227 | 115  | 
lemma bi_total_pcr_fset [transfer_rule]:  | 
116  | 
"bi_total A \<Longrightarrow> bi_total (pcr_fset A)"  | 
|
117  | 
unfolding pcr_fset_def  | 
|
| 47660 | 118  | 
by (intro bi_total_OO bi_total_list_all2 fset.bi_total)  | 
119  | 
||
| 50227 | 120  | 
text {* Transfer rules for @{text "pcr_fset"} can be derived from the
 | 
| 47660 | 121  | 
  existing transfer rules for @{text "cr_fset"} together with the
 | 
122  | 
transfer rules for the polymorphic raw constants. *}  | 
|
123  | 
||
124  | 
text {* Note that the proofs below all have a similar structure and
 | 
|
125  | 
could potentially be automated. *}  | 
|
126  | 
||
127  | 
lemma fnil_transfer [transfer_rule]:  | 
|
| 50227 | 128  | 
"(pcr_fset A) [] fnil"  | 
129  | 
unfolding pcr_fset_def  | 
|
| 47660 | 130  | 
apply (rule relcomppI)  | 
131  | 
apply (rule Nil_transfer)  | 
|
132  | 
apply (rule fnil.transfer)  | 
|
133  | 
done  | 
|
134  | 
||
135  | 
lemma fcons_transfer [transfer_rule]:  | 
|
| 50227 | 136  | 
"(A ===> pcr_fset A ===> pcr_fset A) Cons fcons"  | 
137  | 
unfolding pcr_fset_def  | 
|
| 47660 | 138  | 
apply (intro fun_relI)  | 
139  | 
apply (elim relcomppE)  | 
|
140  | 
apply (rule relcomppI)  | 
|
141  | 
apply (erule (1) Cons_transfer [THEN fun_relD, THEN fun_relD])  | 
|
142  | 
apply (erule fcons.transfer [THEN fun_relD, THEN fun_relD, OF refl])  | 
|
143  | 
done  | 
|
144  | 
||
145  | 
lemma fappend_transfer [transfer_rule]:  | 
|
| 50227 | 146  | 
"(pcr_fset A ===> pcr_fset A ===> pcr_fset A) append fappend"  | 
147  | 
unfolding pcr_fset_def  | 
|
| 47660 | 148  | 
apply (intro fun_relI)  | 
149  | 
apply (elim relcomppE)  | 
|
150  | 
apply (rule relcomppI)  | 
|
151  | 
apply (erule (1) append_transfer [THEN fun_relD, THEN fun_relD])  | 
|
152  | 
apply (erule (1) fappend.transfer [THEN fun_relD, THEN fun_relD])  | 
|
153  | 
done  | 
|
154  | 
||
155  | 
lemma fmap_transfer [transfer_rule]:  | 
|
| 50227 | 156  | 
"((A ===> B) ===> pcr_fset A ===> pcr_fset B) map fmap"  | 
157  | 
unfolding pcr_fset_def  | 
|
| 47660 | 158  | 
apply (intro fun_relI)  | 
159  | 
apply (elim relcomppE)  | 
|
160  | 
apply (rule relcomppI)  | 
|
161  | 
apply (erule (1) map_transfer [THEN fun_relD, THEN fun_relD])  | 
|
| 
47676
 
ec235f564444
adapt to changes in generated transfer rules (cf. 4483c004499a)
 
huffman 
parents: 
47660 
diff
changeset
 | 
162  | 
apply (erule fmap.transfer [THEN fun_relD, THEN fun_relD, unfolded relator_eq, OF refl])  | 
| 47660 | 163  | 
done  | 
164  | 
||
165  | 
lemma ffilter_transfer [transfer_rule]:  | 
|
| 50227 | 166  | 
"((A ===> op =) ===> pcr_fset A ===> pcr_fset A) filter ffilter"  | 
167  | 
unfolding pcr_fset_def  | 
|
| 47660 | 168  | 
apply (intro fun_relI)  | 
169  | 
apply (elim relcomppE)  | 
|
170  | 
apply (rule relcomppI)  | 
|
171  | 
apply (erule (1) filter_transfer [THEN fun_relD, THEN fun_relD])  | 
|
| 
47676
 
ec235f564444
adapt to changes in generated transfer rules (cf. 4483c004499a)
 
huffman 
parents: 
47660 
diff
changeset
 | 
172  | 
apply (erule ffilter.transfer [THEN fun_relD, THEN fun_relD, unfolded relator_eq, OF refl])  | 
| 47660 | 173  | 
done  | 
174  | 
||
175  | 
lemma fset_transfer [transfer_rule]:  | 
|
| 50227 | 176  | 
"(pcr_fset A ===> set_rel A) set fset"  | 
177  | 
unfolding pcr_fset_def  | 
|
| 47660 | 178  | 
apply (intro fun_relI)  | 
179  | 
apply (elim relcomppE)  | 
|
| 
47676
 
ec235f564444
adapt to changes in generated transfer rules (cf. 4483c004499a)
 
huffman 
parents: 
47660 
diff
changeset
 | 
180  | 
apply (drule fset.transfer [THEN fun_relD, unfolded relator_eq])  | 
| 47660 | 181  | 
apply (erule subst)  | 
182  | 
apply (erule set_transfer [THEN fun_relD])  | 
|
183  | 
done  | 
|
184  | 
||
185  | 
lemma fconcat_transfer [transfer_rule]:  | 
|
| 50227 | 186  | 
"(pcr_fset (pcr_fset A) ===> pcr_fset A) concat fconcat"  | 
187  | 
unfolding pcr_fset_def  | 
|
| 47660 | 188  | 
unfolding list_all2_OO  | 
189  | 
apply (intro fun_relI)  | 
|
190  | 
apply (elim relcomppE)  | 
|
191  | 
apply (rule relcomppI)  | 
|
192  | 
apply (erule concat_transfer [THEN fun_relD])  | 
|
193  | 
apply (rule fconcat.transfer [THEN fun_relD])  | 
|
194  | 
apply (erule (1) relcomppI)  | 
|
195  | 
done  | 
|
196  | 
||
197  | 
lemma list_eq_transfer [transfer_rule]:  | 
|
198  | 
assumes [transfer_rule]: "bi_unique A"  | 
|
199  | 
shows "(list_all2 A ===> list_all2 A ===> op =) list_eq list_eq"  | 
|
200  | 
unfolding list_eq_def [abs_def] by transfer_prover  | 
|
201  | 
||
202  | 
lemma fset_eq_transfer [transfer_rule]:  | 
|
203  | 
assumes "bi_unique A"  | 
|
| 50227 | 204  | 
shows "(pcr_fset A ===> pcr_fset A ===> op =) list_eq (op =)"  | 
205  | 
unfolding pcr_fset_def  | 
|
| 47660 | 206  | 
apply (intro fun_relI)  | 
207  | 
apply (elim relcomppE)  | 
|
208  | 
apply (rule trans)  | 
|
209  | 
apply (erule (1) list_eq_transfer [THEN fun_relD, THEN fun_relD, OF assms])  | 
|
210  | 
apply (erule (1) fset.rel_eq_transfer [THEN fun_relD, THEN fun_relD])  | 
|
211  | 
done  | 
|
212  | 
||
213  | 
text {* We don't need the original transfer rules any more: *}
 | 
|
214  | 
||
215  | 
lemmas [transfer_rule del] =  | 
|
216  | 
fset.bi_total fset.right_total fset.right_unique  | 
|
217  | 
fnil.transfer fcons.transfer fappend.transfer fmap.transfer  | 
|
218  | 
ffilter.transfer fset.transfer fconcat.transfer fset.rel_eq_transfer  | 
|
219  | 
||
220  | 
subsection {* Transfer examples *}
 | 
|
221  | 
||
222  | 
text {* The @{text "transfer"} method replaces equality on @{text
 | 
|
223  | 
  "fset"} with the @{text "list_eq"} relation on lists, which is
 | 
|
224  | 
logically equivalent. *}  | 
|
225  | 
||
226  | 
lemma "fmap f (fmap g xs) = fmap (f \<circ> g) xs"  | 
|
227  | 
apply transfer  | 
|
228  | 
apply simp  | 
|
229  | 
done  | 
|
230  | 
||
231  | 
text {* The @{text "transfer'"} variant can replace equality on @{text
 | 
|
232  | 
  "fset"} with equality on @{text "list"}, which is logically stronger
 | 
|
233  | 
but sometimes more convenient. *}  | 
|
234  | 
||
235  | 
lemma "fmap f (fmap g xs) = fmap (f \<circ> g) xs"  | 
|
236  | 
apply transfer'  | 
|
237  | 
apply (rule map_map)  | 
|
238  | 
done  | 
|
239  | 
||
240  | 
lemma "ffilter p (fmap f xs) = fmap f (ffilter (p \<circ> f) xs)"  | 
|
241  | 
apply transfer'  | 
|
242  | 
apply (rule filter_map)  | 
|
243  | 
done  | 
|
244  | 
||
245  | 
lemma "ffilter p (ffilter q xs) = ffilter (\<lambda>x. q x \<and> p x) xs"  | 
|
246  | 
apply transfer'  | 
|
247  | 
apply (rule filter_filter)  | 
|
248  | 
done  | 
|
249  | 
||
250  | 
lemma "fset (fcons x xs) = insert x (fset xs)"  | 
|
251  | 
apply transfer  | 
|
252  | 
apply (rule set.simps)  | 
|
253  | 
done  | 
|
254  | 
||
255  | 
lemma "fset (fappend xs ys) = fset xs \<union> fset ys"  | 
|
256  | 
apply transfer  | 
|
257  | 
apply (rule set_append)  | 
|
258  | 
done  | 
|
259  | 
||
260  | 
lemma "fset (fconcat xss) = (\<Union>xs\<in>fset xss. fset xs)"  | 
|
261  | 
apply transfer  | 
|
262  | 
apply (rule set_concat)  | 
|
263  | 
done  | 
|
264  | 
||
265  | 
lemma "\<forall>x\<in>fset xs. f x = g x \<Longrightarrow> fmap f xs = fmap g xs"  | 
|
266  | 
apply transfer  | 
|
267  | 
apply (simp cong: map_cong del: set_map)  | 
|
268  | 
done  | 
|
269  | 
||
270  | 
lemma "fnil = fconcat xss \<longleftrightarrow> (\<forall>xs\<in>fset xss. xs = fnil)"  | 
|
271  | 
apply transfer  | 
|
272  | 
apply simp  | 
|
273  | 
done  | 
|
274  | 
||
275  | 
lemma "fconcat (fmap (\<lambda>x. fcons x fnil) xs) = xs"  | 
|
276  | 
apply transfer'  | 
|
277  | 
apply simp  | 
|
278  | 
done  | 
|
279  | 
||
280  | 
lemma concat_map_concat: "concat (map concat xsss) = concat (concat xsss)"  | 
|
281  | 
by (induct xsss, simp_all)  | 
|
282  | 
||
283  | 
lemma "fconcat (fmap fconcat xss) = fconcat (fconcat xss)"  | 
|
284  | 
apply transfer'  | 
|
285  | 
apply (rule concat_map_concat)  | 
|
286  | 
done  | 
|
287  | 
||
288  | 
end  |