| author | haftmann | 
| Wed, 17 Feb 2016 21:51:56 +0100 | |
| changeset 62345 | e66d7841d5a2 | 
| parent 62324 | ae44f16dcea5 | 
| child 63040 | eb4ddd18d635 | 
| permissions | -rw-r--r-- | 
| 60247 | 1  | 
(* Title: HOL/Cardinals/Bounded_Set.thy  | 
| 59747 | 2  | 
Author: Dmitriy Traytel, TU Muenchen  | 
3  | 
Copyright 2015  | 
|
4  | 
||
5  | 
Bounded powerset type.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
section \<open>Sets Strictly Bounded by an Infinite Cardinal\<close>  | 
|
9  | 
||
10  | 
theory Bounded_Set  | 
|
11  | 
imports Cardinals  | 
|
12  | 
begin  | 
|
13  | 
||
14  | 
typedef ('a, 'k) bset ("_ set[_]" [22, 21] 21) =
 | 
|
15  | 
  "{A :: 'a set. |A| <o natLeq +c |UNIV :: 'k set|}"
 | 
|
16  | 
morphisms set_bset Abs_bset  | 
|
17  | 
  by (rule exI[of _ "{}"]) (auto simp: card_of_empty4 csum_def)
 | 
|
18  | 
||
19  | 
setup_lifting type_definition_bset  | 
|
20  | 
||
21  | 
lift_definition map_bset ::  | 
|
22  | 
  "('a \<Rightarrow> 'b) \<Rightarrow> 'a set['k] \<Rightarrow> 'b set['k]" is image
 | 
|
23  | 
using card_of_image ordLeq_ordLess_trans by blast  | 
|
24  | 
||
25  | 
lift_definition rel_bset ::  | 
|
26  | 
  "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a set['k] \<Rightarrow> 'b set['k] \<Rightarrow> bool" is rel_set
 | 
|
27  | 
.  | 
|
28  | 
||
29  | 
lift_definition bempty :: "'a set['k]" is "{}"
 | 
|
30  | 
by (auto simp: card_of_empty4 csum_def)  | 
|
31  | 
||
32  | 
lift_definition binsert :: "'a \<Rightarrow> 'a set['k] \<Rightarrow> 'a set['k]" is "insert"  | 
|
33  | 
using infinite_card_of_insert ordIso_ordLess_trans Field_card_of Field_natLeq UNIV_Plus_UNIV  | 
|
34  | 
csum_def finite_Plus_UNIV_iff finite_insert finite_ordLess_infinite2 infinite_UNIV_nat by metis  | 
|
35  | 
||
36  | 
definition bsingleton where  | 
|
37  | 
"bsingleton x = binsert x bempty"  | 
|
38  | 
||
39  | 
lemma set_bset_to_set_bset: "|A| <o natLeq +c |UNIV :: 'k set| \<Longrightarrow>  | 
|
40  | 
set_bset (the_inv set_bset A :: 'a set['k]) = A"  | 
|
41  | 
apply (rule f_the_inv_into_f[unfolded inj_on_def])  | 
|
42  | 
apply (simp add: set_bset_inject range_eqI Abs_bset_inverse[symmetric])  | 
|
43  | 
apply (rule range_eqI Abs_bset_inverse[symmetric] CollectI)+  | 
|
44  | 
.  | 
|
45  | 
||
46  | 
lemma rel_bset_aux_infinite:  | 
|
47  | 
fixes a :: "'a set['k]" and b :: "'b set['k]"  | 
|
48  | 
shows "(\<forall>t \<in> set_bset a. \<exists>u \<in> set_bset b. R t u) \<and> (\<forall>u \<in> set_bset b. \<exists>t \<in> set_bset a. R t u) \<longleftrightarrow>  | 
|
49  | 
   ((BNF_Def.Grp {a. set_bset a \<subseteq> {(a, b). R a b}} (map_bset fst))\<inverse>\<inverse> OO
 | 
|
50  | 
    BNF_Def.Grp {a. set_bset a \<subseteq> {(a, b). R a b}} (map_bset snd)) a b" (is "?L \<longleftrightarrow> ?R")
 | 
|
51  | 
proof  | 
|
52  | 
assume ?L  | 
|
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
60247 
diff
changeset
 | 
53  | 
  def R' \<equiv> "the_inv set_bset (Collect (case_prod R) \<inter> (set_bset a \<times> set_bset b)) :: ('a \<times> 'b) set['k]"
 | 
| 59747 | 54  | 
(is "the_inv set_bset ?L'")  | 
55  | 
have "|?L'| <o natLeq +c |UNIV :: 'k set|"  | 
|
56  | 
unfolding csum_def Field_natLeq  | 
|
57  | 
by (intro ordLeq_ordLess_trans[OF card_of_mono1[OF Int_lower2]]  | 
|
58  | 
card_of_Times_ordLess_infinite)  | 
|
59  | 
(simp, (transfer, simp add: csum_def Field_natLeq)+)  | 
|
60  | 
hence *: "set_bset R' = ?L'" unfolding R'_def by (intro set_bset_to_set_bset)  | 
|
61  | 
show ?R unfolding Grp_def relcompp.simps conversep.simps  | 
|
62  | 
proof (intro CollectI case_prodI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)  | 
|
63  | 
from * show "a = map_bset fst R'" using conjunct1[OF `?L`]  | 
|
64  | 
by (transfer, auto simp add: image_def Int_def split: prod.splits)  | 
|
65  | 
from * show "b = map_bset snd R'" using conjunct2[OF `?L`]  | 
|
66  | 
by (transfer, auto simp add: image_def Int_def split: prod.splits)  | 
|
67  | 
qed (auto simp add: *)  | 
|
68  | 
next  | 
|
69  | 
assume ?R thus ?L unfolding Grp_def relcompp.simps conversep.simps  | 
|
70  | 
by transfer force  | 
|
71  | 
qed  | 
|
72  | 
||
73  | 
bnf "'a set['k]"  | 
|
74  | 
map: map_bset  | 
|
75  | 
sets: set_bset  | 
|
76  | 
bd: "natLeq +c |UNIV :: 'k set|"  | 
|
77  | 
wits: bempty  | 
|
78  | 
rel: rel_bset  | 
|
79  | 
proof -  | 
|
80  | 
show "map_bset id = id" by (rule ext, transfer) simp  | 
|
81  | 
next  | 
|
82  | 
fix f g  | 
|
83  | 
show "map_bset (f o g) = map_bset f o map_bset g" by (rule ext, transfer) auto  | 
|
84  | 
next  | 
|
85  | 
fix X f g  | 
|
86  | 
assume "\<And>z. z \<in> set_bset X \<Longrightarrow> f z = g z"  | 
|
87  | 
then show "map_bset f X = map_bset g X" by transfer force  | 
|
88  | 
next  | 
|
89  | 
fix f  | 
|
90  | 
show "set_bset \<circ> map_bset f = op ` f \<circ> set_bset" by (rule ext, transfer) auto  | 
|
91  | 
next  | 
|
92  | 
fix X :: "'a set['k]"  | 
|
93  | 
show "|set_bset X| \<le>o natLeq +c |UNIV :: 'k set|"  | 
|
94  | 
by transfer (blast dest: ordLess_imp_ordLeq)  | 
|
95  | 
next  | 
|
96  | 
fix R S  | 
|
97  | 
show "rel_bset R OO rel_bset S \<le> rel_bset (R OO S)"  | 
|
98  | 
by (rule predicate2I, transfer) (auto simp: rel_set_OO[symmetric])  | 
|
99  | 
next  | 
|
100  | 
fix R :: "'a \<Rightarrow> 'b \<Rightarrow> bool"  | 
|
| 62324 | 101  | 
  show "rel_bset R = ((\<lambda>x y. \<exists>z. set_bset z \<subseteq> {(x, y). R x y} \<and>
 | 
102  | 
map_bset fst z = x \<and> map_bset snd z = y) :: 'a set['k] \<Rightarrow> 'b set['k] \<Rightarrow> bool)"  | 
|
103  | 
by (simp add: rel_bset_def map_fun_def o_def rel_set_def  | 
|
104  | 
rel_bset_aux_infinite[unfolded OO_Grp_alt])  | 
|
| 59747 | 105  | 
next  | 
106  | 
fix x  | 
|
107  | 
assume "x \<in> set_bset bempty"  | 
|
108  | 
then show False by transfer simp  | 
|
109  | 
qed (simp_all add: card_order_csum natLeq_card_order cinfinite_csum natLeq_cinfinite)  | 
|
110  | 
||
111  | 
||
112  | 
lemma map_bset_bempty[simp]: "map_bset f bempty = bempty"  | 
|
113  | 
by transfer auto  | 
|
114  | 
||
115  | 
lemma map_bset_binsert[simp]: "map_bset f (binsert x X) = binsert (f x) (map_bset f X)"  | 
|
116  | 
by transfer auto  | 
|
117  | 
||
118  | 
lemma map_bset_bsingleton: "map_bset f (bsingleton x) = bsingleton (f x)"  | 
|
119  | 
unfolding bsingleton_def by simp  | 
|
120  | 
||
121  | 
lemma bempty_not_binsert: "bempty \<noteq> binsert x X" "binsert x X \<noteq> bempty"  | 
|
122  | 
by (transfer, auto)+  | 
|
123  | 
||
124  | 
lemma bempty_not_bsingleton[simp]: "bempty \<noteq> bsingleton x" "bsingleton x \<noteq> bempty"  | 
|
125  | 
unfolding bsingleton_def by (simp_all add: bempty_not_binsert)  | 
|
126  | 
||
127  | 
lemma bsingleton_inj[simp]: "bsingleton x = bsingleton y \<longleftrightarrow> x = y"  | 
|
128  | 
unfolding bsingleton_def by transfer auto  | 
|
129  | 
||
130  | 
lemma rel_bsingleton[simp]:  | 
|
131  | 
"rel_bset R (bsingleton x1) (bsingleton x2) = R x1 x2"  | 
|
132  | 
unfolding bsingleton_def  | 
|
133  | 
by transfer (auto simp: rel_set_def)  | 
|
134  | 
||
135  | 
lemma rel_bset_bsingleton[simp]:  | 
|
136  | 
"rel_bset R (bsingleton x1) = (\<lambda>X. X \<noteq> bempty \<and> (\<forall>x2\<in>set_bset X. R x1 x2))"  | 
|
137  | 
"rel_bset R X (bsingleton x2) = (X \<noteq> bempty \<and> (\<forall>x1\<in>set_bset X. R x1 x2))"  | 
|
138  | 
unfolding bsingleton_def fun_eq_iff  | 
|
139  | 
by (transfer, force simp add: rel_set_def)+  | 
|
140  | 
||
141  | 
lemma rel_bset_bempty[simp]:  | 
|
142  | 
"rel_bset R bempty X = (X = bempty)"  | 
|
143  | 
"rel_bset R Y bempty = (Y = bempty)"  | 
|
144  | 
by (transfer, simp add: rel_set_def)+  | 
|
145  | 
||
146  | 
definition bset_of_option where  | 
|
147  | 
"bset_of_option = case_option bempty bsingleton"  | 
|
148  | 
||
149  | 
lift_definition bgraph :: "('a \<Rightarrow> 'b option) \<Rightarrow> ('a \<times> 'b) set['a set]" is
 | 
|
150  | 
  "\<lambda>f. {(a, b). f a = Some b}"
 | 
|
151  | 
proof -  | 
|
152  | 
fix f :: "'a \<Rightarrow> 'b option"  | 
|
153  | 
  have "|{(a, b). f a = Some b}| \<le>o |UNIV :: 'a set|"
 | 
|
154  | 
by (rule surj_imp_ordLeq[of _ "\<lambda>x. (x, the (f x))"]) auto  | 
|
155  | 
also have "|UNIV :: 'a set| <o |UNIV :: 'a set set|"  | 
|
156  | 
by simp  | 
|
157  | 
also have "|UNIV :: 'a set set| \<le>o natLeq +c |UNIV :: 'a set set|"  | 
|
158  | 
by (rule ordLeq_csum2) simp  | 
|
159  | 
  finally show "|{(a, b). f a = Some b}| <o natLeq +c |UNIV :: 'a set set|" .
 | 
|
160  | 
qed  | 
|
161  | 
||
162  | 
lemma rel_bset_False[simp]: "rel_bset (\<lambda>x y. False) x y = (x = bempty \<and> y = bempty)"  | 
|
163  | 
by transfer (auto simp: rel_set_def)  | 
|
164  | 
||
165  | 
lemma rel_bset_of_option[simp]:  | 
|
166  | 
"rel_bset R (bset_of_option x1) (bset_of_option x2) = rel_option R x1 x2"  | 
|
167  | 
unfolding bset_of_option_def bsingleton_def[abs_def]  | 
|
168  | 
by transfer (auto simp: rel_set_def split: option.splits)  | 
|
169  | 
||
170  | 
lemma rel_bgraph[simp]:  | 
|
171  | 
"rel_bset (rel_prod (op =) R) (bgraph f1) (bgraph f2) = rel_fun (op =) (rel_option R) f1 f2"  | 
|
172  | 
apply transfer  | 
|
173  | 
apply (auto simp: rel_fun_def rel_option_iff rel_set_def split: option.splits)  | 
|
174  | 
using option.collapse apply fastforce+  | 
|
175  | 
done  | 
|
176  | 
||
177  | 
lemma set_bset_bsingleton[simp]:  | 
|
178  | 
  "set_bset (bsingleton x) = {x}"
 | 
|
179  | 
unfolding bsingleton_def by transfer auto  | 
|
180  | 
||
181  | 
lemma binsert_absorb[simp]: "binsert a (binsert a x) = binsert a x"  | 
|
182  | 
by transfer simp  | 
|
183  | 
||
184  | 
lemma map_bset_eq_bempty_iff[simp]: "map_bset f X = bempty \<longleftrightarrow> X = bempty"  | 
|
185  | 
by transfer auto  | 
|
186  | 
||
187  | 
lemma map_bset_eq_bsingleton_iff[simp]:  | 
|
188  | 
  "map_bset f X = bsingleton x \<longleftrightarrow> (set_bset X \<noteq> {} \<and> (\<forall>y \<in> set_bset X. f y = x))"
 | 
|
189  | 
unfolding bsingleton_def by transfer auto  | 
|
190  | 
||
191  | 
lift_definition bCollect :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set['a set]" is Collect
 | 
|
192  | 
apply (rule ordLeq_ordLess_trans[OF card_of_mono1[OF subset_UNIV]])  | 
|
193  | 
apply (rule ordLess_ordLeq_trans[OF card_of_set_type])  | 
|
194  | 
apply (rule ordLeq_csum2[OF card_of_Card_order])  | 
|
195  | 
done  | 
|
196  | 
||
197  | 
lift_definition bmember :: "'a \<Rightarrow> 'a set['k] \<Rightarrow> bool" is "op \<in>" .  | 
|
198  | 
||
199  | 
lemma bmember_bCollect[simp]: "bmember a (bCollect P) = P a"  | 
|
200  | 
by transfer simp  | 
|
201  | 
||
202  | 
lemma bset_eq_iff: "A = B \<longleftrightarrow> (\<forall>a. bmember a A = bmember a B)"  | 
|
203  | 
by transfer auto  | 
|
204  | 
||
205  | 
(* FIXME: Lifting does not work with dead variables,  | 
|
206  | 
that is why we are hiding the below setup in a locale*)  | 
|
207  | 
locale bset_lifting  | 
|
208  | 
begin  | 
|
209  | 
||
210  | 
declare bset.rel_eq[relator_eq]  | 
|
211  | 
declare bset.rel_mono[relator_mono]  | 
|
212  | 
declare bset.rel_compp[symmetric, relator_distr]  | 
|
213  | 
declare bset.rel_transfer[transfer_rule]  | 
|
214  | 
||
215  | 
lemma bset_quot_map[quot_map]: "Quotient R Abs Rep T \<Longrightarrow>  | 
|
216  | 
Quotient (rel_bset R) (map_bset Abs) (map_bset Rep) (rel_bset T)"  | 
|
217  | 
unfolding Quotient_alt_def5 bset.rel_Grp[of UNIV, simplified, symmetric]  | 
|
218  | 
bset.rel_conversep[symmetric] bset.rel_compp[symmetric]  | 
|
219  | 
by (auto elim: bset.rel_mono_strong)  | 
|
220  | 
||
221  | 
lemma set_relator_eq_onp [relator_eq_onp]:  | 
|
222  | 
"rel_bset (eq_onp P) = eq_onp (\<lambda>A. Ball (set_bset A) P)"  | 
|
223  | 
unfolding fun_eq_iff eq_onp_def by transfer (auto simp: rel_set_def)  | 
|
224  | 
||
225  | 
end  | 
|
226  | 
||
227  | 
end  |