| author | wenzelm | 
| Tue, 08 Mar 2016 18:15:16 +0100 | |
| changeset 62559 | 83e815849a91 | 
| parent 62390 | 842917225d56 | 
| child 63040 | eb4ddd18d635 | 
| permissions | -rw-r--r-- | 
| 53953 | 1 | (* Title: HOL/Library/FSet.thy | 
| 2 | Author: Ondrej Kuncar, TU Muenchen | |
| 3 | Author: Cezary Kaliszyk and Christian Urban | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 4 | Author: Andrei Popescu, TU Muenchen | 
| 53953 | 5 | *) | 
| 6 | ||
| 60500 | 7 | section \<open>Type of finite sets defined as a subtype of sets\<close> | 
| 53953 | 8 | |
| 9 | theory FSet | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 10 | imports Conditionally_Complete_Lattices | 
| 53953 | 11 | begin | 
| 12 | ||
| 60500 | 13 | subsection \<open>Definition of the type\<close> | 
| 53953 | 14 | |
| 15 | typedef 'a fset = "{A :: 'a set. finite A}"  morphisms fset Abs_fset
 | |
| 16 | by auto | |
| 17 | ||
| 18 | setup_lifting type_definition_fset | |
| 19 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 20 | |
| 60500 | 21 | subsection \<open>Basic operations and type class instantiations\<close> | 
| 53953 | 22 | |
| 23 | (* FIXME transfer and right_total vs. bi_total *) | |
| 24 | instantiation fset :: (finite) finite | |
| 25 | begin | |
| 60679 | 26 | instance by (standard; transfer; simp) | 
| 53953 | 27 | end | 
| 28 | ||
| 29 | instantiation fset :: (type) "{bounded_lattice_bot, distrib_lattice, minus}"
 | |
| 30 | begin | |
| 31 | ||
| 32 | interpretation lifting_syntax . | |
| 33 | ||
| 34 | lift_definition bot_fset :: "'a fset" is "{}" parametric empty_transfer by simp 
 | |
| 35 | ||
| 36 | lift_definition less_eq_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" is subset_eq parametric subset_transfer | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55414diff
changeset | 37 | . | 
| 53953 | 38 | |
| 39 | definition less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" where "xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)" | |
| 40 | ||
| 41 | lemma less_fset_transfer[transfer_rule]: | |
| 42 | assumes [transfer_rule]: "bi_unique A" | |
| 43 | shows "((pcr_fset A) ===> (pcr_fset A) ===> op =) op \<subset> op <" | |
| 44 | unfolding less_fset_def[abs_def] psubset_eq[abs_def] by transfer_prover | |
| 45 | ||
| 46 | ||
| 47 | lift_definition sup_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is union parametric union_transfer | |
| 48 | by simp | |
| 49 | ||
| 50 | lift_definition inf_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is inter parametric inter_transfer | |
| 51 | by simp | |
| 52 | ||
| 53 | lift_definition minus_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is minus parametric Diff_transfer | |
| 54 | by simp | |
| 55 | ||
| 56 | instance | |
| 60679 | 57 | by (standard; transfer; auto)+ | 
| 53953 | 58 | |
| 59 | end | |
| 60 | ||
| 61 | abbreviation fempty :: "'a fset" ("{||}") where "{||} \<equiv> bot"
 | |
| 62 | abbreviation fsubset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50) where "xs |\<subseteq>| ys \<equiv> xs \<le> ys" | |
| 63 | abbreviation fsubset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50) where "xs |\<subset>| ys \<equiv> xs < ys" | |
| 64 | abbreviation funion :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl "|\<union>|" 65) where "xs |\<union>| ys \<equiv> sup xs ys" | |
| 65 | abbreviation finter :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl "|\<inter>|" 65) where "xs |\<inter>| ys \<equiv> inf xs ys" | |
| 66 | abbreviation fminus :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl "|-|" 65) where "xs |-| ys \<equiv> minus xs ys" | |
| 67 | ||
| 54014 | 68 | instantiation fset :: (equal) equal | 
| 69 | begin | |
| 70 | definition "HOL.equal A B \<longleftrightarrow> A |\<subseteq>| B \<and> B |\<subseteq>| A" | |
| 71 | instance by intro_classes (auto simp add: equal_fset_def) | |
| 72 | end | |
| 73 | ||
| 53953 | 74 | instantiation fset :: (type) conditionally_complete_lattice | 
| 75 | begin | |
| 76 | ||
| 77 | interpretation lifting_syntax . | |
| 78 | ||
| 79 | lemma right_total_Inf_fset_transfer: | |
| 80 | assumes [transfer_rule]: "bi_unique A" and [transfer_rule]: "right_total A" | |
| 55938 | 81 | shows "(rel_set (rel_set A) ===> rel_set A) | 
| 61952 | 82 |     (\<lambda>S. if finite (\<Inter>S \<inter> Collect (Domainp A)) then \<Inter>S \<inter> Collect (Domainp A) else {}) 
 | 
| 53953 | 83 |       (\<lambda>S. if finite (Inf S) then Inf S else {})"
 | 
| 84 | by transfer_prover | |
| 85 | ||
| 86 | lemma Inf_fset_transfer: | |
| 87 | assumes [transfer_rule]: "bi_unique A" and [transfer_rule]: "bi_total A" | |
| 55938 | 88 |   shows "(rel_set (rel_set A) ===> rel_set A) (\<lambda>A. if finite (Inf A) then Inf A else {}) 
 | 
| 53953 | 89 |     (\<lambda>A. if finite (Inf A) then Inf A else {})"
 | 
| 90 | by transfer_prover | |
| 91 | ||
| 92 | lift_definition Inf_fset :: "'a fset set \<Rightarrow> 'a fset" is "\<lambda>A. if finite (Inf A) then Inf A else {}" 
 | |
| 93 | parametric right_total_Inf_fset_transfer Inf_fset_transfer by simp | |
| 94 | ||
| 95 | lemma Sup_fset_transfer: | |
| 96 | assumes [transfer_rule]: "bi_unique A" | |
| 55938 | 97 |   shows "(rel_set (rel_set A) ===> rel_set A) (\<lambda>A. if finite (Sup A) then Sup A else {})
 | 
| 53953 | 98 |   (\<lambda>A. if finite (Sup A) then Sup A else {})" by transfer_prover
 | 
| 99 | ||
| 100 | lift_definition Sup_fset :: "'a fset set \<Rightarrow> 'a fset" is "\<lambda>A. if finite (Sup A) then Sup A else {}"
 | |
| 101 | parametric Sup_fset_transfer by simp | |
| 102 | ||
| 103 | lemma finite_Sup: "\<exists>z. finite z \<and> (\<forall>a. a \<in> X \<longrightarrow> a \<le> z) \<Longrightarrow> finite (Sup X)" | |
| 104 | by (auto intro: finite_subset) | |
| 105 | ||
| 55938 | 106 | lemma transfer_bdd_below[transfer_rule]: "(rel_set (pcr_fset op =) ===> op =) bdd_below bdd_below" | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54014diff
changeset | 107 | by auto | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54014diff
changeset | 108 | |
| 53953 | 109 | instance | 
| 110 | proof | |
| 111 | fix x z :: "'a fset" | |
| 112 | fix X :: "'a fset set" | |
| 113 |   {
 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54014diff
changeset | 114 | assume "x \<in> X" "bdd_below X" | 
| 56646 | 115 | then show "Inf X |\<subseteq>| x" by transfer auto | 
| 53953 | 116 | next | 
| 117 |     assume "X \<noteq> {}" "(\<And>x. x \<in> X \<Longrightarrow> z |\<subseteq>| x)"
 | |
| 118 | then show "z |\<subseteq>| Inf X" by transfer (clarsimp, blast) | |
| 119 | next | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54014diff
changeset | 120 | assume "x \<in> X" "bdd_above X" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54014diff
changeset | 121 | then obtain z where "x \<in> X" "(\<And>x. x \<in> X \<Longrightarrow> x |\<subseteq>| z)" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54014diff
changeset | 122 | by (auto simp: bdd_above_def) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54014diff
changeset | 123 | then show "x |\<subseteq>| Sup X" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54014diff
changeset | 124 | by transfer (auto intro!: finite_Sup) | 
| 53953 | 125 | next | 
| 126 |     assume "X \<noteq> {}" "(\<And>x. x \<in> X \<Longrightarrow> x |\<subseteq>| z)"
 | |
| 127 | then show "Sup X |\<subseteq>| z" by transfer (clarsimp, blast) | |
| 128 | } | |
| 129 | qed | |
| 130 | end | |
| 131 | ||
| 132 | instantiation fset :: (finite) complete_lattice | |
| 133 | begin | |
| 134 | ||
| 60679 | 135 | lift_definition top_fset :: "'a fset" is UNIV parametric right_total_UNIV_transfer UNIV_transfer | 
| 136 | by simp | |
| 53953 | 137 | |
| 60679 | 138 | instance | 
| 139 | by (standard; transfer; auto) | |
| 140 | ||
| 53953 | 141 | end | 
| 142 | ||
| 143 | instantiation fset :: (finite) complete_boolean_algebra | |
| 144 | begin | |
| 145 | ||
| 146 | lift_definition uminus_fset :: "'a fset \<Rightarrow> 'a fset" is uminus | |
| 147 | parametric right_total_Compl_transfer Compl_transfer by simp | |
| 148 | ||
| 60679 | 149 | instance | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 150 | by (standard; transfer) (simp_all add: Diff_eq) | 
| 53953 | 151 | |
| 152 | end | |
| 153 | ||
| 154 | abbreviation fUNIV :: "'a::finite fset" where "fUNIV \<equiv> top" | |
| 155 | abbreviation fuminus :: "'a::finite fset \<Rightarrow> 'a fset" ("|-| _" [81] 80) where "|-| x \<equiv> uminus x"
 | |
| 156 | ||
| 56646 | 157 | declare top_fset.rep_eq[simp] | 
| 158 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 159 | |
| 60500 | 160 | subsection \<open>Other operations\<close> | 
| 53953 | 161 | |
| 162 | lift_definition finsert :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is insert parametric Lifting_Set.insert_transfer | |
| 163 | by simp | |
| 164 | ||
| 165 | syntax | |
| 166 |   "_insert_fset"     :: "args => 'a fset"  ("{|(_)|}")
 | |
| 167 | ||
| 168 | translations | |
| 169 |   "{|x, xs|}" == "CONST finsert x {|xs|}"
 | |
| 170 |   "{|x|}"     == "CONST finsert x {||}"
 | |
| 171 | ||
| 172 | lift_definition fmember :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<in>|" 50) is Set.member | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55414diff
changeset | 173 | parametric member_transfer . | 
| 53953 | 174 | |
| 175 | abbreviation notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50) where "x |\<notin>| S \<equiv> \<not> (x |\<in>| S)" | |
| 176 | ||
| 177 | context | |
| 178 | begin | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 179 | |
| 53953 | 180 | interpretation lifting_syntax . | 
| 181 | ||
| 182 | lift_definition ffilter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is Set.filter 
 | |
| 183 | parametric Lifting_Set.filter_transfer unfolding Set.filter_def by simp | |
| 184 | ||
| 185 | lift_definition fPow :: "'a fset \<Rightarrow> 'a fset fset" is Pow parametric Pow_transfer | |
| 55732 | 186 | by (simp add: finite_subset) | 
| 53953 | 187 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55414diff
changeset | 188 | lift_definition fcard :: "'a fset \<Rightarrow> nat" is card parametric card_transfer . | 
| 53953 | 189 | |
| 190 | lift_definition fimage :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" (infixr "|`|" 90) is image 
 | |
| 191 | parametric image_transfer by simp | |
| 192 | ||
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55414diff
changeset | 193 | lift_definition fthe_elem :: "'a fset \<Rightarrow> 'a" is the_elem . | 
| 53953 | 194 | |
| 55738 | 195 | lift_definition fbind :: "'a fset \<Rightarrow> ('a \<Rightarrow> 'b fset) \<Rightarrow> 'b fset" is Set.bind parametric bind_transfer 
 | 
| 196 | by (simp add: Set.bind_def) | |
| 53953 | 197 | |
| 55732 | 198 | lift_definition ffUnion :: "'a fset fset \<Rightarrow> 'a fset" is Union parametric Union_transfer by simp | 
| 53953 | 199 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55414diff
changeset | 200 | lift_definition fBall :: "'a fset \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" is Ball parametric Ball_transfer .
 | 
| 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55414diff
changeset | 201 | lift_definition fBex :: "'a fset \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" is Bex parametric Bex_transfer .
 | 
| 53953 | 202 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55414diff
changeset | 203 | lift_definition ffold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" is Finite_Set.fold .
 | 
| 53963 | 204 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 205 | |
| 60500 | 206 | subsection \<open>Transferred lemmas from Set.thy\<close> | 
| 53953 | 207 | |
| 208 | lemmas fset_eqI = set_eqI[Transfer.transferred] | |
| 209 | lemmas fset_eq_iff[no_atp] = set_eq_iff[Transfer.transferred] | |
| 210 | lemmas fBallI[intro!] = ballI[Transfer.transferred] | |
| 211 | lemmas fbspec[dest?] = bspec[Transfer.transferred] | |
| 212 | lemmas fBallE[elim] = ballE[Transfer.transferred] | |
| 213 | lemmas fBexI[intro] = bexI[Transfer.transferred] | |
| 214 | lemmas rev_fBexI[intro?] = rev_bexI[Transfer.transferred] | |
| 215 | lemmas fBexCI = bexCI[Transfer.transferred] | |
| 216 | lemmas fBexE[elim!] = bexE[Transfer.transferred] | |
| 217 | lemmas fBall_triv[simp] = ball_triv[Transfer.transferred] | |
| 218 | lemmas fBex_triv[simp] = bex_triv[Transfer.transferred] | |
| 219 | lemmas fBex_triv_one_point1[simp] = bex_triv_one_point1[Transfer.transferred] | |
| 220 | lemmas fBex_triv_one_point2[simp] = bex_triv_one_point2[Transfer.transferred] | |
| 221 | lemmas fBex_one_point1[simp] = bex_one_point1[Transfer.transferred] | |
| 222 | lemmas fBex_one_point2[simp] = bex_one_point2[Transfer.transferred] | |
| 223 | lemmas fBall_one_point1[simp] = ball_one_point1[Transfer.transferred] | |
| 224 | lemmas fBall_one_point2[simp] = ball_one_point2[Transfer.transferred] | |
| 225 | lemmas fBall_conj_distrib = ball_conj_distrib[Transfer.transferred] | |
| 226 | lemmas fBex_disj_distrib = bex_disj_distrib[Transfer.transferred] | |
| 227 | lemmas fBall_cong = ball_cong[Transfer.transferred] | |
| 228 | lemmas fBex_cong = bex_cong[Transfer.transferred] | |
| 53964 | 229 | lemmas fsubsetI[intro!] = subsetI[Transfer.transferred] | 
| 230 | lemmas fsubsetD[elim, intro?] = subsetD[Transfer.transferred] | |
| 231 | lemmas rev_fsubsetD[no_atp,intro?] = rev_subsetD[Transfer.transferred] | |
| 232 | lemmas fsubsetCE[no_atp,elim] = subsetCE[Transfer.transferred] | |
| 233 | lemmas fsubset_eq[no_atp] = subset_eq[Transfer.transferred] | |
| 234 | lemmas contra_fsubsetD[no_atp] = contra_subsetD[Transfer.transferred] | |
| 235 | lemmas fsubset_refl = subset_refl[Transfer.transferred] | |
| 236 | lemmas fsubset_trans = subset_trans[Transfer.transferred] | |
| 53953 | 237 | lemmas fset_rev_mp = set_rev_mp[Transfer.transferred] | 
| 238 | lemmas fset_mp = set_mp[Transfer.transferred] | |
| 53964 | 239 | lemmas fsubset_not_fsubset_eq[code] = subset_not_subset_eq[Transfer.transferred] | 
| 53953 | 240 | lemmas eq_fmem_trans = eq_mem_trans[Transfer.transferred] | 
| 53964 | 241 | lemmas fsubset_antisym[intro!] = subset_antisym[Transfer.transferred] | 
| 53953 | 242 | lemmas fequalityD1 = equalityD1[Transfer.transferred] | 
| 243 | lemmas fequalityD2 = equalityD2[Transfer.transferred] | |
| 244 | lemmas fequalityE = equalityE[Transfer.transferred] | |
| 245 | lemmas fequalityCE[elim] = equalityCE[Transfer.transferred] | |
| 246 | lemmas eqfset_imp_iff = eqset_imp_iff[Transfer.transferred] | |
| 247 | lemmas eqfelem_imp_iff = eqelem_imp_iff[Transfer.transferred] | |
| 248 | lemmas fempty_iff[simp] = empty_iff[Transfer.transferred] | |
| 53964 | 249 | lemmas fempty_fsubsetI[iff] = empty_subsetI[Transfer.transferred] | 
| 53953 | 250 | lemmas equalsffemptyI = equals0I[Transfer.transferred] | 
| 251 | lemmas equalsffemptyD = equals0D[Transfer.transferred] | |
| 252 | lemmas fBall_fempty[simp] = ball_empty[Transfer.transferred] | |
| 253 | lemmas fBex_fempty[simp] = bex_empty[Transfer.transferred] | |
| 254 | lemmas fPow_iff[iff] = Pow_iff[Transfer.transferred] | |
| 255 | lemmas fPowI = PowI[Transfer.transferred] | |
| 256 | lemmas fPowD = PowD[Transfer.transferred] | |
| 257 | lemmas fPow_bottom = Pow_bottom[Transfer.transferred] | |
| 258 | lemmas fPow_top = Pow_top[Transfer.transferred] | |
| 259 | lemmas fPow_not_fempty = Pow_not_empty[Transfer.transferred] | |
| 260 | lemmas finter_iff[simp] = Int_iff[Transfer.transferred] | |
| 261 | lemmas finterI[intro!] = IntI[Transfer.transferred] | |
| 262 | lemmas finterD1 = IntD1[Transfer.transferred] | |
| 263 | lemmas finterD2 = IntD2[Transfer.transferred] | |
| 264 | lemmas finterE[elim!] = IntE[Transfer.transferred] | |
| 265 | lemmas funion_iff[simp] = Un_iff[Transfer.transferred] | |
| 266 | lemmas funionI1[elim?] = UnI1[Transfer.transferred] | |
| 267 | lemmas funionI2[elim?] = UnI2[Transfer.transferred] | |
| 268 | lemmas funionCI[intro!] = UnCI[Transfer.transferred] | |
| 269 | lemmas funionE[elim!] = UnE[Transfer.transferred] | |
| 270 | lemmas fminus_iff[simp] = Diff_iff[Transfer.transferred] | |
| 271 | lemmas fminusI[intro!] = DiffI[Transfer.transferred] | |
| 272 | lemmas fminusD1 = DiffD1[Transfer.transferred] | |
| 273 | lemmas fminusD2 = DiffD2[Transfer.transferred] | |
| 274 | lemmas fminusE[elim!] = DiffE[Transfer.transferred] | |
| 275 | lemmas finsert_iff[simp] = insert_iff[Transfer.transferred] | |
| 276 | lemmas finsertI1 = insertI1[Transfer.transferred] | |
| 277 | lemmas finsertI2 = insertI2[Transfer.transferred] | |
| 278 | lemmas finsertE[elim!] = insertE[Transfer.transferred] | |
| 279 | lemmas finsertCI[intro!] = insertCI[Transfer.transferred] | |
| 53964 | 280 | lemmas fsubset_finsert_iff = subset_insert_iff[Transfer.transferred] | 
| 53953 | 281 | lemmas finsert_ident = insert_ident[Transfer.transferred] | 
| 282 | lemmas fsingletonI[intro!,no_atp] = singletonI[Transfer.transferred] | |
| 283 | lemmas fsingletonD[dest!,no_atp] = singletonD[Transfer.transferred] | |
| 284 | lemmas fsingleton_iff = singleton_iff[Transfer.transferred] | |
| 285 | lemmas fsingleton_inject[dest!] = singleton_inject[Transfer.transferred] | |
| 286 | lemmas fsingleton_finsert_inj_eq[iff,no_atp] = singleton_insert_inj_eq[Transfer.transferred] | |
| 287 | lemmas fsingleton_finsert_inj_eq'[iff,no_atp] = singleton_insert_inj_eq'[Transfer.transferred] | |
| 53964 | 288 | lemmas fsubset_fsingletonD = subset_singletonD[Transfer.transferred] | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
62082diff
changeset | 289 | lemmas fminus_single_finsert = Diff_single_insert[Transfer.transferred] | 
| 53953 | 290 | lemmas fdoubleton_eq_iff = doubleton_eq_iff[Transfer.transferred] | 
| 291 | lemmas funion_fsingleton_iff = Un_singleton_iff[Transfer.transferred] | |
| 292 | lemmas fsingleton_funion_iff = singleton_Un_iff[Transfer.transferred] | |
| 293 | lemmas fimage_eqI[simp, intro] = image_eqI[Transfer.transferred] | |
| 294 | lemmas fimageI = imageI[Transfer.transferred] | |
| 295 | lemmas rev_fimage_eqI = rev_image_eqI[Transfer.transferred] | |
| 296 | lemmas fimageE[elim!] = imageE[Transfer.transferred] | |
| 297 | lemmas Compr_fimage_eq = Compr_image_eq[Transfer.transferred] | |
| 298 | lemmas fimage_funion = image_Un[Transfer.transferred] | |
| 299 | lemmas fimage_iff = image_iff[Transfer.transferred] | |
| 53964 | 300 | lemmas fimage_fsubset_iff[no_atp] = image_subset_iff[Transfer.transferred] | 
| 301 | lemmas fimage_fsubsetI = image_subsetI[Transfer.transferred] | |
| 53953 | 302 | lemmas fimage_ident[simp] = image_ident[Transfer.transferred] | 
| 62390 | 303 | lemmas if_split_fmem1 = if_split_mem1[Transfer.transferred] | 
| 304 | lemmas if_split_fmem2 = if_split_mem2[Transfer.transferred] | |
| 53964 | 305 | lemmas pfsubsetI[intro!,no_atp] = psubsetI[Transfer.transferred] | 
| 306 | lemmas pfsubsetE[elim!,no_atp] = psubsetE[Transfer.transferred] | |
| 307 | lemmas pfsubset_finsert_iff = psubset_insert_iff[Transfer.transferred] | |
| 308 | lemmas pfsubset_eq = psubset_eq[Transfer.transferred] | |
| 309 | lemmas pfsubset_imp_fsubset = psubset_imp_subset[Transfer.transferred] | |
| 310 | lemmas pfsubset_trans = psubset_trans[Transfer.transferred] | |
| 311 | lemmas pfsubsetD = psubsetD[Transfer.transferred] | |
| 312 | lemmas pfsubset_fsubset_trans = psubset_subset_trans[Transfer.transferred] | |
| 313 | lemmas fsubset_pfsubset_trans = subset_psubset_trans[Transfer.transferred] | |
| 314 | lemmas pfsubset_imp_ex_fmem = psubset_imp_ex_mem[Transfer.transferred] | |
| 53953 | 315 | lemmas fimage_fPow_mono = image_Pow_mono[Transfer.transferred] | 
| 316 | lemmas fimage_fPow_surj = image_Pow_surj[Transfer.transferred] | |
| 53964 | 317 | lemmas fsubset_finsertI = subset_insertI[Transfer.transferred] | 
| 318 | lemmas fsubset_finsertI2 = subset_insertI2[Transfer.transferred] | |
| 319 | lemmas fsubset_finsert = subset_insert[Transfer.transferred] | |
| 53953 | 320 | lemmas funion_upper1 = Un_upper1[Transfer.transferred] | 
| 321 | lemmas funion_upper2 = Un_upper2[Transfer.transferred] | |
| 322 | lemmas funion_least = Un_least[Transfer.transferred] | |
| 323 | lemmas finter_lower1 = Int_lower1[Transfer.transferred] | |
| 324 | lemmas finter_lower2 = Int_lower2[Transfer.transferred] | |
| 325 | lemmas finter_greatest = Int_greatest[Transfer.transferred] | |
| 53964 | 326 | lemmas fminus_fsubset = Diff_subset[Transfer.transferred] | 
| 327 | lemmas fminus_fsubset_conv = Diff_subset_conv[Transfer.transferred] | |
| 328 | lemmas fsubset_fempty[simp] = subset_empty[Transfer.transferred] | |
| 329 | lemmas not_pfsubset_fempty[iff] = not_psubset_empty[Transfer.transferred] | |
| 53953 | 330 | lemmas finsert_is_funion = insert_is_Un[Transfer.transferred] | 
| 331 | lemmas finsert_not_fempty[simp] = insert_not_empty[Transfer.transferred] | |
| 332 | lemmas fempty_not_finsert = empty_not_insert[Transfer.transferred] | |
| 333 | lemmas finsert_absorb = insert_absorb[Transfer.transferred] | |
| 334 | lemmas finsert_absorb2[simp] = insert_absorb2[Transfer.transferred] | |
| 335 | lemmas finsert_commute = insert_commute[Transfer.transferred] | |
| 53964 | 336 | lemmas finsert_fsubset[simp] = insert_subset[Transfer.transferred] | 
| 53953 | 337 | lemmas finsert_inter_finsert[simp] = insert_inter_insert[Transfer.transferred] | 
| 338 | lemmas finsert_disjoint[simp,no_atp] = insert_disjoint[Transfer.transferred] | |
| 339 | lemmas disjoint_finsert[simp,no_atp] = disjoint_insert[Transfer.transferred] | |
| 340 | lemmas fimage_fempty[simp] = image_empty[Transfer.transferred] | |
| 341 | lemmas fimage_finsert[simp] = image_insert[Transfer.transferred] | |
| 342 | lemmas fimage_constant = image_constant[Transfer.transferred] | |
| 343 | lemmas fimage_constant_conv = image_constant_conv[Transfer.transferred] | |
| 344 | lemmas fimage_fimage = image_image[Transfer.transferred] | |
| 345 | lemmas finsert_fimage[simp] = insert_image[Transfer.transferred] | |
| 346 | lemmas fimage_is_fempty[iff] = image_is_empty[Transfer.transferred] | |
| 347 | lemmas fempty_is_fimage[iff] = empty_is_image[Transfer.transferred] | |
| 348 | lemmas fimage_cong = image_cong[Transfer.transferred] | |
| 53964 | 349 | lemmas fimage_finter_fsubset = image_Int_subset[Transfer.transferred] | 
| 350 | lemmas fimage_fminus_fsubset = image_diff_subset[Transfer.transferred] | |
| 53953 | 351 | lemmas finter_absorb = Int_absorb[Transfer.transferred] | 
| 352 | lemmas finter_left_absorb = Int_left_absorb[Transfer.transferred] | |
| 353 | lemmas finter_commute = Int_commute[Transfer.transferred] | |
| 354 | lemmas finter_left_commute = Int_left_commute[Transfer.transferred] | |
| 355 | lemmas finter_assoc = Int_assoc[Transfer.transferred] | |
| 356 | lemmas finter_ac = Int_ac[Transfer.transferred] | |
| 357 | lemmas finter_absorb1 = Int_absorb1[Transfer.transferred] | |
| 358 | lemmas finter_absorb2 = Int_absorb2[Transfer.transferred] | |
| 359 | lemmas finter_fempty_left = Int_empty_left[Transfer.transferred] | |
| 360 | lemmas finter_fempty_right = Int_empty_right[Transfer.transferred] | |
| 361 | lemmas disjoint_iff_fnot_equal = disjoint_iff_not_equal[Transfer.transferred] | |
| 362 | lemmas finter_funion_distrib = Int_Un_distrib[Transfer.transferred] | |
| 363 | lemmas finter_funion_distrib2 = Int_Un_distrib2[Transfer.transferred] | |
| 53964 | 364 | lemmas finter_fsubset_iff[no_atp, simp] = Int_subset_iff[Transfer.transferred] | 
| 53953 | 365 | lemmas funion_absorb = Un_absorb[Transfer.transferred] | 
| 366 | lemmas funion_left_absorb = Un_left_absorb[Transfer.transferred] | |
| 367 | lemmas funion_commute = Un_commute[Transfer.transferred] | |
| 368 | lemmas funion_left_commute = Un_left_commute[Transfer.transferred] | |
| 369 | lemmas funion_assoc = Un_assoc[Transfer.transferred] | |
| 370 | lemmas funion_ac = Un_ac[Transfer.transferred] | |
| 371 | lemmas funion_absorb1 = Un_absorb1[Transfer.transferred] | |
| 372 | lemmas funion_absorb2 = Un_absorb2[Transfer.transferred] | |
| 373 | lemmas funion_fempty_left = Un_empty_left[Transfer.transferred] | |
| 374 | lemmas funion_fempty_right = Un_empty_right[Transfer.transferred] | |
| 375 | lemmas funion_finsert_left[simp] = Un_insert_left[Transfer.transferred] | |
| 376 | lemmas funion_finsert_right[simp] = Un_insert_right[Transfer.transferred] | |
| 377 | lemmas finter_finsert_left = Int_insert_left[Transfer.transferred] | |
| 378 | lemmas finter_finsert_left_ifffempty[simp] = Int_insert_left_if0[Transfer.transferred] | |
| 379 | lemmas finter_finsert_left_if1[simp] = Int_insert_left_if1[Transfer.transferred] | |
| 380 | lemmas finter_finsert_right = Int_insert_right[Transfer.transferred] | |
| 381 | lemmas finter_finsert_right_ifffempty[simp] = Int_insert_right_if0[Transfer.transferred] | |
| 382 | lemmas finter_finsert_right_if1[simp] = Int_insert_right_if1[Transfer.transferred] | |
| 383 | lemmas funion_finter_distrib = Un_Int_distrib[Transfer.transferred] | |
| 384 | lemmas funion_finter_distrib2 = Un_Int_distrib2[Transfer.transferred] | |
| 385 | lemmas funion_finter_crazy = Un_Int_crazy[Transfer.transferred] | |
| 53964 | 386 | lemmas fsubset_funion_eq = subset_Un_eq[Transfer.transferred] | 
| 53953 | 387 | lemmas funion_fempty[iff] = Un_empty[Transfer.transferred] | 
| 53964 | 388 | lemmas funion_fsubset_iff[no_atp, simp] = Un_subset_iff[Transfer.transferred] | 
| 53953 | 389 | lemmas funion_fminus_finter = Un_Diff_Int[Transfer.transferred] | 
| 390 | lemmas fminus_finter2 = Diff_Int2[Transfer.transferred] | |
| 391 | lemmas funion_finter_assoc_eq = Un_Int_assoc_eq[Transfer.transferred] | |
| 392 | lemmas fBall_funion = ball_Un[Transfer.transferred] | |
| 393 | lemmas fBex_funion = bex_Un[Transfer.transferred] | |
| 394 | lemmas fminus_eq_fempty_iff[simp,no_atp] = Diff_eq_empty_iff[Transfer.transferred] | |
| 395 | lemmas fminus_cancel[simp] = Diff_cancel[Transfer.transferred] | |
| 396 | lemmas fminus_idemp[simp] = Diff_idemp[Transfer.transferred] | |
| 397 | lemmas fminus_triv = Diff_triv[Transfer.transferred] | |
| 398 | lemmas fempty_fminus[simp] = empty_Diff[Transfer.transferred] | |
| 399 | lemmas fminus_fempty[simp] = Diff_empty[Transfer.transferred] | |
| 400 | lemmas fminus_finsertffempty[simp,no_atp] = Diff_insert0[Transfer.transferred] | |
| 401 | lemmas fminus_finsert = Diff_insert[Transfer.transferred] | |
| 402 | lemmas fminus_finsert2 = Diff_insert2[Transfer.transferred] | |
| 403 | lemmas finsert_fminus_if = insert_Diff_if[Transfer.transferred] | |
| 404 | lemmas finsert_fminus1[simp] = insert_Diff1[Transfer.transferred] | |
| 405 | lemmas finsert_fminus_single[simp] = insert_Diff_single[Transfer.transferred] | |
| 406 | lemmas finsert_fminus = insert_Diff[Transfer.transferred] | |
| 407 | lemmas fminus_finsert_absorb = Diff_insert_absorb[Transfer.transferred] | |
| 408 | lemmas fminus_disjoint[simp] = Diff_disjoint[Transfer.transferred] | |
| 409 | lemmas fminus_partition = Diff_partition[Transfer.transferred] | |
| 410 | lemmas double_fminus = double_diff[Transfer.transferred] | |
| 411 | lemmas funion_fminus_cancel[simp] = Un_Diff_cancel[Transfer.transferred] | |
| 412 | lemmas funion_fminus_cancel2[simp] = Un_Diff_cancel2[Transfer.transferred] | |
| 413 | lemmas fminus_funion = Diff_Un[Transfer.transferred] | |
| 414 | lemmas fminus_finter = Diff_Int[Transfer.transferred] | |
| 415 | lemmas funion_fminus = Un_Diff[Transfer.transferred] | |
| 416 | lemmas finter_fminus = Int_Diff[Transfer.transferred] | |
| 417 | lemmas fminus_finter_distrib = Diff_Int_distrib[Transfer.transferred] | |
| 418 | lemmas fminus_finter_distrib2 = Diff_Int_distrib2[Transfer.transferred] | |
| 419 | lemmas fUNIV_bool[no_atp] = UNIV_bool[Transfer.transferred] | |
| 420 | lemmas fPow_fempty[simp] = Pow_empty[Transfer.transferred] | |
| 421 | lemmas fPow_finsert = Pow_insert[Transfer.transferred] | |
| 53964 | 422 | lemmas funion_fPow_fsubset = Un_Pow_subset[Transfer.transferred] | 
| 53953 | 423 | lemmas fPow_finter_eq[simp] = Pow_Int_eq[Transfer.transferred] | 
| 53964 | 424 | lemmas fset_eq_fsubset = set_eq_subset[Transfer.transferred] | 
| 425 | lemmas fsubset_iff[no_atp] = subset_iff[Transfer.transferred] | |
| 426 | lemmas fsubset_iff_pfsubset_eq = subset_iff_psubset_eq[Transfer.transferred] | |
| 53953 | 427 | lemmas all_not_fin_conv[simp] = all_not_in_conv[Transfer.transferred] | 
| 428 | lemmas ex_fin_conv = ex_in_conv[Transfer.transferred] | |
| 429 | lemmas fimage_mono = image_mono[Transfer.transferred] | |
| 430 | lemmas fPow_mono = Pow_mono[Transfer.transferred] | |
| 431 | lemmas finsert_mono = insert_mono[Transfer.transferred] | |
| 432 | lemmas funion_mono = Un_mono[Transfer.transferred] | |
| 433 | lemmas finter_mono = Int_mono[Transfer.transferred] | |
| 434 | lemmas fminus_mono = Diff_mono[Transfer.transferred] | |
| 435 | lemmas fin_mono = in_mono[Transfer.transferred] | |
| 436 | lemmas fthe_felem_eq[simp] = the_elem_eq[Transfer.transferred] | |
| 437 | lemmas fLeast_mono = Least_mono[Transfer.transferred] | |
| 438 | lemmas fbind_fbind = bind_bind[Transfer.transferred] | |
| 439 | lemmas fempty_fbind[simp] = empty_bind[Transfer.transferred] | |
| 440 | lemmas nonfempty_fbind_const = nonempty_bind_const[Transfer.transferred] | |
| 441 | lemmas fbind_const = bind_const[Transfer.transferred] | |
| 442 | lemmas ffmember_filter[simp] = member_filter[Transfer.transferred] | |
| 443 | lemmas fequalityI = equalityI[Transfer.transferred] | |
| 444 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 445 | |
| 60500 | 446 | subsection \<open>Additional lemmas\<close> | 
| 53953 | 447 | |
| 61585 | 448 | subsubsection \<open>\<open>fsingleton\<close>\<close> | 
| 53953 | 449 | |
| 450 | lemmas fsingletonE = fsingletonD [elim_format] | |
| 451 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 452 | |
| 61585 | 453 | subsubsection \<open>\<open>femepty\<close>\<close> | 
| 53953 | 454 | |
| 455 | lemma fempty_ffilter[simp]: "ffilter (\<lambda>_. False) A = {||}"
 | |
| 456 | by transfer auto | |
| 457 | ||
| 458 | (* FIXME, transferred doesn't work here *) | |
| 459 | lemma femptyE [elim!]: "a |\<in>| {||} \<Longrightarrow> P"
 | |
| 460 | by simp | |
| 461 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 462 | |
| 61585 | 463 | subsubsection \<open>\<open>fset\<close>\<close> | 
| 53953 | 464 | |
| 53963 | 465 | lemmas fset_simps[simp] = bot_fset.rep_eq finsert.rep_eq | 
| 53953 | 466 | |
| 467 | lemma finite_fset [simp]: | |
| 468 | shows "finite (fset S)" | |
| 469 | by transfer simp | |
| 470 | ||
| 53963 | 471 | lemmas fset_cong = fset_inject | 
| 53953 | 472 | |
| 473 | lemma filter_fset [simp]: | |
| 474 | shows "fset (ffilter P xs) = Collect P \<inter> fset xs" | |
| 475 | by transfer auto | |
| 476 | ||
| 53963 | 477 | lemma notin_fset: "x |\<notin>| S \<longleftrightarrow> x \<notin> fset S" by (simp add: fmember.rep_eq) | 
| 478 | ||
| 479 | lemmas inter_fset[simp] = inf_fset.rep_eq | |
| 53953 | 480 | |
| 53963 | 481 | lemmas union_fset[simp] = sup_fset.rep_eq | 
| 53953 | 482 | |
| 53963 | 483 | lemmas minus_fset[simp] = minus_fset.rep_eq | 
| 53953 | 484 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 485 | |
| 61585 | 486 | subsubsection \<open>\<open>filter_fset\<close>\<close> | 
| 53953 | 487 | |
| 488 | lemma subset_ffilter: | |
| 489 | "ffilter P A |\<subseteq>| ffilter Q A = (\<forall> x. x |\<in>| A \<longrightarrow> P x \<longrightarrow> Q x)" | |
| 490 | by transfer auto | |
| 491 | ||
| 492 | lemma eq_ffilter: | |
| 493 | "(ffilter P A = ffilter Q A) = (\<forall>x. x |\<in>| A \<longrightarrow> P x = Q x)" | |
| 494 | by transfer auto | |
| 495 | ||
| 53964 | 496 | lemma pfsubset_ffilter: | 
| 53953 | 497 | "(\<And>x. x |\<in>| A \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| A & \<not> P x & Q x) \<Longrightarrow> | 
| 498 | ffilter P A |\<subset>| ffilter Q A" | |
| 499 | unfolding less_fset_def by (auto simp add: subset_ffilter eq_ffilter) | |
| 500 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 501 | |
| 61585 | 502 | subsubsection \<open>\<open>finsert\<close>\<close> | 
| 53953 | 503 | |
| 504 | (* FIXME, transferred doesn't work here *) | |
| 505 | lemma set_finsert: | |
| 506 | assumes "x |\<in>| A" | |
| 507 | obtains B where "A = finsert x B" and "x |\<notin>| B" | |
| 508 | using assms by transfer (metis Set.set_insert finite_insert) | |
| 509 | ||
| 510 | lemma mk_disjoint_finsert: "a |\<in>| A \<Longrightarrow> \<exists>B. A = finsert a B \<and> a |\<notin>| B" | |
| 511 |   by (rule_tac x = "A |-| {|a|}" in exI, blast)
 | |
| 512 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 513 | |
| 61585 | 514 | subsubsection \<open>\<open>fimage\<close>\<close> | 
| 53953 | 515 | |
| 516 | lemma subset_fimage_iff: "(B |\<subseteq>| f|`|A) = (\<exists> AA. AA |\<subseteq>| A \<and> B = f|`|AA)" | |
| 517 | by transfer (metis mem_Collect_eq rev_finite_subset subset_image_iff) | |
| 518 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 519 | |
| 60500 | 520 | subsubsection \<open>bounded quantification\<close> | 
| 53953 | 521 | |
| 522 | lemma bex_simps [simp, no_atp]: | |
| 523 | "\<And>A P Q. fBex A (\<lambda>x. P x \<and> Q) = (fBex A P \<and> Q)" | |
| 524 | "\<And>A P Q. fBex A (\<lambda>x. P \<and> Q x) = (P \<and> fBex A Q)" | |
| 525 |   "\<And>P. fBex {||} P = False" 
 | |
| 526 | "\<And>a B P. fBex (finsert a B) P = (P a \<or> fBex B P)" | |
| 527 | "\<And>A P f. fBex (f |`| A) P = fBex A (\<lambda>x. P (f x))" | |
| 528 | "\<And>A P. (\<not> fBex A P) = fBall A (\<lambda>x. \<not> P x)" | |
| 529 | by auto | |
| 530 | ||
| 531 | lemma ball_simps [simp, no_atp]: | |
| 532 | "\<And>A P Q. fBall A (\<lambda>x. P x \<or> Q) = (fBall A P \<or> Q)" | |
| 533 | "\<And>A P Q. fBall A (\<lambda>x. P \<or> Q x) = (P \<or> fBall A Q)" | |
| 534 | "\<And>A P Q. fBall A (\<lambda>x. P \<longrightarrow> Q x) = (P \<longrightarrow> fBall A Q)" | |
| 535 | "\<And>A P Q. fBall A (\<lambda>x. P x \<longrightarrow> Q) = (fBex A P \<longrightarrow> Q)" | |
| 536 |   "\<And>P. fBall {||} P = True"
 | |
| 537 | "\<And>a B P. fBall (finsert a B) P = (P a \<and> fBall B P)" | |
| 538 | "\<And>A P f. fBall (f |`| A) P = fBall A (\<lambda>x. P (f x))" | |
| 539 | "\<And>A P. (\<not> fBall A P) = fBex A (\<lambda>x. \<not> P x)" | |
| 540 | by auto | |
| 541 | ||
| 542 | lemma atomize_fBall: | |
| 543 | "(\<And>x. x |\<in>| A ==> P x) == Trueprop (fBall A (\<lambda>x. P x))" | |
| 544 | apply (simp only: atomize_all atomize_imp) | |
| 545 | apply (rule equal_intr_rule) | |
| 546 | by (transfer, simp)+ | |
| 547 | ||
| 53963 | 548 | end | 
| 549 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 550 | |
| 61585 | 551 | subsubsection \<open>\<open>fcard\<close>\<close> | 
| 53963 | 552 | |
| 53964 | 553 | (* FIXME: improve transferred to handle bounded meta quantification *) | 
| 554 | ||
| 53963 | 555 | lemma fcard_fempty: | 
| 556 |   "fcard {||} = 0"
 | |
| 557 | by transfer (rule card_empty) | |
| 558 | ||
| 559 | lemma fcard_finsert_disjoint: | |
| 560 | "x |\<notin>| A \<Longrightarrow> fcard (finsert x A) = Suc (fcard A)" | |
| 561 | by transfer (rule card_insert_disjoint) | |
| 562 | ||
| 563 | lemma fcard_finsert_if: | |
| 564 | "fcard (finsert x A) = (if x |\<in>| A then fcard A else Suc (fcard A))" | |
| 565 | by transfer (rule card_insert_if) | |
| 566 | ||
| 567 | lemma card_0_eq [simp, no_atp]: | |
| 568 |   "fcard A = 0 \<longleftrightarrow> A = {||}"
 | |
| 569 | by transfer (rule card_0_eq) | |
| 570 | ||
| 571 | lemma fcard_Suc_fminus1: | |
| 572 |   "x |\<in>| A \<Longrightarrow> Suc (fcard (A |-| {|x|})) = fcard A"
 | |
| 573 | by transfer (rule card_Suc_Diff1) | |
| 574 | ||
| 575 | lemma fcard_fminus_fsingleton: | |
| 576 |   "x |\<in>| A \<Longrightarrow> fcard (A |-| {|x|}) = fcard A - 1"
 | |
| 577 | by transfer (rule card_Diff_singleton) | |
| 578 | ||
| 579 | lemma fcard_fminus_fsingleton_if: | |
| 580 |   "fcard (A |-| {|x|}) = (if x |\<in>| A then fcard A - 1 else fcard A)"
 | |
| 581 | by transfer (rule card_Diff_singleton_if) | |
| 582 | ||
| 583 | lemma fcard_fminus_finsert[simp]: | |
| 584 | assumes "a |\<in>| A" and "a |\<notin>| B" | |
| 585 | shows "fcard (A |-| finsert a B) = fcard (A |-| B) - 1" | |
| 586 | using assms by transfer (rule card_Diff_insert) | |
| 587 | ||
| 588 | lemma fcard_finsert: "fcard (finsert x A) = Suc (fcard (A |-| {|x|}))"
 | |
| 589 | by transfer (rule card_insert) | |
| 590 | ||
| 591 | lemma fcard_finsert_le: "fcard A \<le> fcard (finsert x A)" | |
| 592 | by transfer (rule card_insert_le) | |
| 593 | ||
| 594 | lemma fcard_mono: | |
| 595 | "A |\<subseteq>| B \<Longrightarrow> fcard A \<le> fcard B" | |
| 596 | by transfer (rule card_mono) | |
| 597 | ||
| 598 | lemma fcard_seteq: "A |\<subseteq>| B \<Longrightarrow> fcard B \<le> fcard A \<Longrightarrow> A = B" | |
| 599 | by transfer (rule card_seteq) | |
| 600 | ||
| 601 | lemma pfsubset_fcard_mono: "A |\<subset>| B \<Longrightarrow> fcard A < fcard B" | |
| 602 | by transfer (rule psubset_card_mono) | |
| 603 | ||
| 604 | lemma fcard_funion_finter: | |
| 605 | "fcard A + fcard B = fcard (A |\<union>| B) + fcard (A |\<inter>| B)" | |
| 606 | by transfer (rule card_Un_Int) | |
| 607 | ||
| 608 | lemma fcard_funion_disjoint: | |
| 609 |   "A |\<inter>| B = {||} \<Longrightarrow> fcard (A |\<union>| B) = fcard A + fcard B"
 | |
| 610 | by transfer (rule card_Un_disjoint) | |
| 611 | ||
| 612 | lemma fcard_funion_fsubset: | |
| 613 | "B |\<subseteq>| A \<Longrightarrow> fcard (A |-| B) = fcard A - fcard B" | |
| 614 | by transfer (rule card_Diff_subset) | |
| 615 | ||
| 616 | lemma diff_fcard_le_fcard_fminus: | |
| 617 | "fcard A - fcard B \<le> fcard(A |-| B)" | |
| 618 | by transfer (rule diff_card_le_card_Diff) | |
| 619 | ||
| 620 | lemma fcard_fminus1_less: "x |\<in>| A \<Longrightarrow> fcard (A |-| {|x|}) < fcard A"
 | |
| 621 | by transfer (rule card_Diff1_less) | |
| 622 | ||
| 623 | lemma fcard_fminus2_less: | |
| 624 |   "x |\<in>| A \<Longrightarrow> y |\<in>| A \<Longrightarrow> fcard (A |-| {|x|} |-| {|y|}) < fcard A"
 | |
| 625 | by transfer (rule card_Diff2_less) | |
| 626 | ||
| 627 | lemma fcard_fminus1_le: "fcard (A |-| {|x|}) \<le> fcard A"
 | |
| 628 | by transfer (rule card_Diff1_le) | |
| 629 | ||
| 630 | lemma fcard_pfsubset: "A |\<subseteq>| B \<Longrightarrow> fcard A < fcard B \<Longrightarrow> A < B" | |
| 631 | by transfer (rule card_psubset) | |
| 632 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 633 | |
| 61585 | 634 | subsubsection \<open>\<open>ffold\<close>\<close> | 
| 53963 | 635 | |
| 636 | (* FIXME: improve transferred to handle bounded meta quantification *) | |
| 637 | ||
| 638 | context comp_fun_commute | |
| 639 | begin | |
| 640 | lemmas ffold_empty[simp] = fold_empty[Transfer.transferred] | |
| 641 | ||
| 642 | lemma ffold_finsert [simp]: | |
| 643 | assumes "x |\<notin>| A" | |
| 644 | shows "ffold f z (finsert x A) = f x (ffold f z A)" | |
| 645 | using assms by (transfer fixing: f) (rule fold_insert) | |
| 646 | ||
| 647 | lemma ffold_fun_left_comm: | |
| 648 | "f x (ffold f z A) = ffold f (f x z) A" | |
| 649 | by (transfer fixing: f) (rule fold_fun_left_comm) | |
| 650 | ||
| 651 | lemma ffold_finsert2: | |
| 56646 | 652 | "x |\<notin>| A \<Longrightarrow> ffold f z (finsert x A) = ffold f (f x z) A" | 
| 53963 | 653 | by (transfer fixing: f) (rule fold_insert2) | 
| 654 | ||
| 655 | lemma ffold_rec: | |
| 656 | assumes "x |\<in>| A" | |
| 657 |     shows "ffold f z A = f x (ffold f z (A |-| {|x|}))"
 | |
| 658 | using assms by (transfer fixing: f) (rule fold_rec) | |
| 659 | ||
| 660 | lemma ffold_finsert_fremove: | |
| 661 |     "ffold f z (finsert x A) = f x (ffold f z (A |-| {|x|}))"
 | |
| 662 | by (transfer fixing: f) (rule fold_insert_remove) | |
| 663 | end | |
| 664 | ||
| 665 | lemma ffold_fimage: | |
| 666 | assumes "inj_on g (fset A)" | |
| 667 | shows "ffold f z (g |`| A) = ffold (f \<circ> g) z A" | |
| 668 | using assms by transfer' (rule fold_image) | |
| 669 | ||
| 670 | lemma ffold_cong: | |
| 671 | assumes "comp_fun_commute f" "comp_fun_commute g" | |
| 672 | "\<And>x. x |\<in>| A \<Longrightarrow> f x = g x" | |
| 673 | and "s = t" and "A = B" | |
| 674 | shows "ffold f s A = ffold g t B" | |
| 675 | using assms by transfer (metis Finite_Set.fold_cong) | |
| 676 | ||
| 677 | context comp_fun_idem | |
| 678 | begin | |
| 679 | ||
| 680 | lemma ffold_finsert_idem: | |
| 56646 | 681 | "ffold f z (finsert x A) = f x (ffold f z A)" | 
| 53963 | 682 | by (transfer fixing: f) (rule fold_insert_idem) | 
| 683 | ||
| 684 | declare ffold_finsert [simp del] ffold_finsert_idem [simp] | |
| 685 | ||
| 686 | lemma ffold_finsert_idem2: | |
| 687 | "ffold f z (finsert x A) = ffold f (f x z) A" | |
| 688 | by (transfer fixing: f) (rule fold_insert_idem2) | |
| 689 | ||
| 690 | end | |
| 691 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 692 | |
| 60500 | 693 | subsection \<open>Choice in fsets\<close> | 
| 53953 | 694 | |
| 695 | lemma fset_choice: | |
| 696 | assumes "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y)" | |
| 697 | shows "\<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)" | |
| 698 | using assms by transfer metis | |
| 699 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 700 | |
| 60500 | 701 | subsection \<open>Induction and Cases rules for fsets\<close> | 
| 53953 | 702 | |
| 703 | lemma fset_exhaust [case_names empty insert, cases type: fset]: | |
| 704 |   assumes fempty_case: "S = {||} \<Longrightarrow> P" 
 | |
| 705 | and finsert_case: "\<And>x S'. S = finsert x S' \<Longrightarrow> P" | |
| 706 | shows "P" | |
| 707 | using assms by transfer blast | |
| 708 | ||
| 709 | lemma fset_induct [case_names empty insert]: | |
| 710 |   assumes fempty_case: "P {||}"
 | |
| 711 | and finsert_case: "\<And>x S. P S \<Longrightarrow> P (finsert x S)" | |
| 712 | shows "P S" | |
| 713 | proof - | |
| 714 | (* FIXME transfer and right_total vs. bi_total *) | |
| 715 | note Domainp_forall_transfer[transfer_rule] | |
| 716 | show ?thesis | |
| 717 | using assms by transfer (auto intro: finite_induct) | |
| 718 | qed | |
| 719 | ||
| 720 | lemma fset_induct_stronger [case_names empty insert, induct type: fset]: | |
| 721 |   assumes empty_fset_case: "P {||}"
 | |
| 722 | and insert_fset_case: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (finsert x S)" | |
| 723 | shows "P S" | |
| 724 | proof - | |
| 725 | (* FIXME transfer and right_total vs. bi_total *) | |
| 726 | note Domainp_forall_transfer[transfer_rule] | |
| 727 | show ?thesis | |
| 728 | using assms by transfer (auto intro: finite_induct) | |
| 729 | qed | |
| 730 | ||
| 731 | lemma fset_card_induct: | |
| 732 |   assumes empty_fset_case: "P {||}"
 | |
| 733 | and card_fset_Suc_case: "\<And>S T. Suc (fcard S) = (fcard T) \<Longrightarrow> P S \<Longrightarrow> P T" | |
| 734 | shows "P S" | |
| 735 | proof (induct S) | |
| 736 | case empty | |
| 737 |   show "P {||}" by (rule empty_fset_case)
 | |
| 738 | next | |
| 739 | case (insert x S) | |
| 740 | have h: "P S" by fact | |
| 741 | have "x |\<notin>| S" by fact | |
| 742 | then have "Suc (fcard S) = fcard (finsert x S)" | |
| 743 | by transfer auto | |
| 744 | then show "P (finsert x S)" | |
| 745 | using h card_fset_Suc_case by simp | |
| 746 | qed | |
| 747 | ||
| 748 | lemma fset_strong_cases: | |
| 749 |   obtains "xs = {||}"
 | |
| 750 | | ys x where "x |\<notin>| ys" and "xs = finsert x ys" | |
| 751 | by transfer blast | |
| 752 | ||
| 753 | lemma fset_induct2: | |
| 754 |   "P {||} {||} \<Longrightarrow>
 | |
| 755 |   (\<And>x xs. x |\<notin>| xs \<Longrightarrow> P (finsert x xs) {||}) \<Longrightarrow>
 | |
| 756 |   (\<And>y ys. y |\<notin>| ys \<Longrightarrow> P {||} (finsert y ys)) \<Longrightarrow>
 | |
| 757 | (\<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> | |
| 758 | P xsa ysa" | |
| 759 | apply (induct xsa arbitrary: ysa) | |
| 760 | apply (induct_tac x rule: fset_induct_stronger) | |
| 761 | apply simp_all | |
| 762 | apply (induct_tac xa rule: fset_induct_stronger) | |
| 763 | apply simp_all | |
| 764 | done | |
| 765 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 766 | |
| 60500 | 767 | subsection \<open>Setup for Lifting/Transfer\<close> | 
| 53953 | 768 | |
| 60500 | 769 | subsubsection \<open>Relator and predicator properties\<close> | 
| 53953 | 770 | |
| 55938 | 771 | lift_definition rel_fset :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'b fset \<Rightarrow> bool" is rel_set
 | 
| 772 | parametric rel_set_transfer . | |
| 53953 | 773 | |
| 55933 | 774 | lemma rel_fset_alt_def: "rel_fset R = (\<lambda>A B. (\<forall>x.\<exists>y. x|\<in>|A \<longrightarrow> y|\<in>|B \<and> R x y) | 
| 53953 | 775 | \<and> (\<forall>y. \<exists>x. y|\<in>|B \<longrightarrow> x|\<in>|A \<and> R x y))" | 
| 776 | apply (rule ext)+ | |
| 777 | apply transfer' | |
| 55938 | 778 | apply (subst rel_set_def[unfolded fun_eq_iff]) | 
| 53953 | 779 | by blast | 
| 780 | ||
| 55938 | 781 | lemma finite_rel_set: | 
| 53953 | 782 | assumes fin: "finite X" "finite Z" | 
| 55938 | 783 | assumes R_S: "rel_set (R OO S) X Z" | 
| 784 | shows "\<exists>Y. finite Y \<and> rel_set R X Y \<and> rel_set S Y Z" | |
| 53953 | 785 | proof - | 
| 786 | obtain f where f: "\<forall>x\<in>X. R x (f x) \<and> (\<exists>z\<in>Z. S (f x) z)" | |
| 787 | apply atomize_elim | |
| 788 | apply (subst bchoice_iff[symmetric]) | |
| 55938 | 789 | using R_S[unfolded rel_set_def OO_def] by blast | 
| 53953 | 790 | |
| 56646 | 791 | obtain g where g: "\<forall>z\<in>Z. S (g z) z \<and> (\<exists>x\<in>X. R x (g z))" | 
| 53953 | 792 | apply atomize_elim | 
| 793 | apply (subst bchoice_iff[symmetric]) | |
| 55938 | 794 | using R_S[unfolded rel_set_def OO_def] by blast | 
| 53953 | 795 | |
| 796 | let ?Y = "f ` X \<union> g ` Z" | |
| 797 | have "finite ?Y" by (simp add: fin) | |
| 55938 | 798 | moreover have "rel_set R X ?Y" | 
| 799 | unfolding rel_set_def | |
| 53953 | 800 | using f g by clarsimp blast | 
| 55938 | 801 | moreover have "rel_set S ?Y Z" | 
| 802 | unfolding rel_set_def | |
| 53953 | 803 | using f g by clarsimp blast | 
| 804 | ultimately show ?thesis by metis | |
| 805 | qed | |
| 806 | ||
| 60500 | 807 | subsubsection \<open>Transfer rules for the Transfer package\<close> | 
| 53953 | 808 | |
| 60500 | 809 | text \<open>Unconditional transfer rules\<close> | 
| 53953 | 810 | |
| 53963 | 811 | context | 
| 812 | begin | |
| 813 | ||
| 814 | interpretation lifting_syntax . | |
| 815 | ||
| 53953 | 816 | lemmas fempty_transfer [transfer_rule] = empty_transfer[Transfer.transferred] | 
| 817 | ||
| 818 | lemma finsert_transfer [transfer_rule]: | |
| 55933 | 819 | "(A ===> rel_fset A ===> rel_fset A) finsert finsert" | 
| 55945 | 820 | unfolding rel_fun_def rel_fset_alt_def by blast | 
| 53953 | 821 | |
| 822 | lemma funion_transfer [transfer_rule]: | |
| 55933 | 823 | "(rel_fset A ===> rel_fset A ===> rel_fset A) funion funion" | 
| 55945 | 824 | unfolding rel_fun_def rel_fset_alt_def by blast | 
| 53953 | 825 | |
| 826 | lemma ffUnion_transfer [transfer_rule]: | |
| 55933 | 827 | "(rel_fset (rel_fset A) ===> rel_fset A) ffUnion ffUnion" | 
| 55945 | 828 | unfolding rel_fun_def rel_fset_alt_def by transfer (simp, fast) | 
| 53953 | 829 | |
| 830 | lemma fimage_transfer [transfer_rule]: | |
| 55933 | 831 | "((A ===> B) ===> rel_fset A ===> rel_fset B) fimage fimage" | 
| 55945 | 832 | unfolding rel_fun_def rel_fset_alt_def by simp blast | 
| 53953 | 833 | |
| 834 | lemma fBall_transfer [transfer_rule]: | |
| 55933 | 835 | "(rel_fset A ===> (A ===> op =) ===> op =) fBall fBall" | 
| 55945 | 836 | unfolding rel_fset_alt_def rel_fun_def by blast | 
| 53953 | 837 | |
| 838 | lemma fBex_transfer [transfer_rule]: | |
| 55933 | 839 | "(rel_fset A ===> (A ===> op =) ===> op =) fBex fBex" | 
| 55945 | 840 | unfolding rel_fset_alt_def rel_fun_def by blast | 
| 53953 | 841 | |
| 842 | (* FIXME transfer doesn't work here *) | |
| 843 | lemma fPow_transfer [transfer_rule]: | |
| 55933 | 844 | "(rel_fset A ===> rel_fset (rel_fset A)) fPow fPow" | 
| 55945 | 845 | unfolding rel_fun_def | 
| 846 | using Pow_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] | |
| 53953 | 847 | by blast | 
| 848 | ||
| 55933 | 849 | lemma rel_fset_transfer [transfer_rule]: | 
| 850 | "((A ===> B ===> op =) ===> rel_fset A ===> rel_fset B ===> op =) | |
| 851 | rel_fset rel_fset" | |
| 55945 | 852 | unfolding rel_fun_def | 
| 853 | using rel_set_transfer[unfolded rel_fun_def,rule_format, Transfer.transferred, where A = A and B = B] | |
| 53953 | 854 | by simp | 
| 855 | ||
| 856 | lemma bind_transfer [transfer_rule]: | |
| 55933 | 857 | "(rel_fset A ===> (A ===> rel_fset B) ===> rel_fset B) fbind fbind" | 
| 55945 | 858 | using assms unfolding rel_fun_def | 
| 859 | using bind_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast | |
| 53953 | 860 | |
| 60500 | 861 | text \<open>Rules requiring bi-unique, bi-total or right-total relations\<close> | 
| 53953 | 862 | |
| 863 | lemma fmember_transfer [transfer_rule]: | |
| 864 | assumes "bi_unique A" | |
| 55933 | 865 | shows "(A ===> rel_fset A ===> op =) (op |\<in>|) (op |\<in>|)" | 
| 55945 | 866 | using assms unfolding rel_fun_def rel_fset_alt_def bi_unique_def by metis | 
| 53953 | 867 | |
| 868 | lemma finter_transfer [transfer_rule]: | |
| 869 | assumes "bi_unique A" | |
| 55933 | 870 | shows "(rel_fset A ===> rel_fset A ===> rel_fset A) finter finter" | 
| 55945 | 871 | using assms unfolding rel_fun_def | 
| 872 | using inter_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast | |
| 53953 | 873 | |
| 53963 | 874 | lemma fminus_transfer [transfer_rule]: | 
| 53953 | 875 | assumes "bi_unique A" | 
| 55933 | 876 | shows "(rel_fset A ===> rel_fset A ===> rel_fset A) (op |-|) (op |-|)" | 
| 55945 | 877 | using assms unfolding rel_fun_def | 
| 878 | using Diff_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast | |
| 53953 | 879 | |
| 880 | lemma fsubset_transfer [transfer_rule]: | |
| 881 | assumes "bi_unique A" | |
| 55933 | 882 | shows "(rel_fset A ===> rel_fset A ===> op =) (op |\<subseteq>|) (op |\<subseteq>|)" | 
| 55945 | 883 | using assms unfolding rel_fun_def | 
| 884 | using subset_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast | |
| 53953 | 885 | |
| 886 | lemma fSup_transfer [transfer_rule]: | |
| 55938 | 887 | "bi_unique A \<Longrightarrow> (rel_set (rel_fset A) ===> rel_fset A) Sup Sup" | 
| 55945 | 888 | using assms unfolding rel_fun_def | 
| 53953 | 889 | apply clarify | 
| 890 | apply transfer' | |
| 55945 | 891 | using Sup_fset_transfer[unfolded rel_fun_def] by blast | 
| 53953 | 892 | |
| 893 | (* FIXME: add right_total_fInf_transfer *) | |
| 894 | ||
| 895 | lemma fInf_transfer [transfer_rule]: | |
| 896 | assumes "bi_unique A" and "bi_total A" | |
| 55938 | 897 | shows "(rel_set (rel_fset A) ===> rel_fset A) Inf Inf" | 
| 55945 | 898 | using assms unfolding rel_fun_def | 
| 53953 | 899 | apply clarify | 
| 900 | apply transfer' | |
| 55945 | 901 | using Inf_fset_transfer[unfolded rel_fun_def] by blast | 
| 53953 | 902 | |
| 903 | lemma ffilter_transfer [transfer_rule]: | |
| 904 | assumes "bi_unique A" | |
| 55933 | 905 | shows "((A ===> op=) ===> rel_fset A ===> rel_fset A) ffilter ffilter" | 
| 55945 | 906 | using assms unfolding rel_fun_def | 
| 907 | using Lifting_Set.filter_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast | |
| 53953 | 908 | |
| 909 | lemma card_transfer [transfer_rule]: | |
| 55933 | 910 | "bi_unique A \<Longrightarrow> (rel_fset A ===> op =) fcard fcard" | 
| 55945 | 911 | using assms unfolding rel_fun_def | 
| 912 | using card_transfer[unfolded rel_fun_def, rule_format, Transfer.transferred] by blast | |
| 53953 | 913 | |
| 914 | end | |
| 915 | ||
| 916 | lifting_update fset.lifting | |
| 917 | lifting_forget fset.lifting | |
| 918 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 919 | |
| 60500 | 920 | subsection \<open>BNF setup\<close> | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 921 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 922 | context | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 923 | includes fset.lifting | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 924 | begin | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 925 | |
| 55933 | 926 | lemma rel_fset_alt: | 
| 927 | "rel_fset R a b \<longleftrightarrow> (\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and> (\<forall>t \<in> fset b. \<exists>u \<in> fset a. R u t)" | |
| 55938 | 928 | by transfer (simp add: rel_set_def) | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 929 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 930 | lemma fset_to_fset: "finite A \<Longrightarrow> fset (the_inv fset A) = A" | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 931 | apply (rule f_the_inv_into_f[unfolded inj_on_def]) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 932 | apply (simp add: fset_inject) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 933 | apply (rule range_eqI Abs_fset_inverse[symmetric] CollectI)+ | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 934 | . | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 935 | |
| 55933 | 936 | lemma rel_fset_aux: | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 937 | "(\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and> (\<forall>u \<in> fset b. \<exists>t \<in> fset a. R t u) \<longleftrightarrow> | 
| 57398 | 938 |  ((BNF_Def.Grp {a. fset a \<subseteq> {(a, b). R a b}} (fimage fst))\<inverse>\<inverse> OO
 | 
| 939 |   BNF_Def.Grp {a. fset a \<subseteq> {(a, b). R a b}} (fimage snd)) a b" (is "?L = ?R")
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 940 | proof | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 941 | assume ?L | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
60712diff
changeset | 942 | def R' \<equiv> "the_inv fset (Collect (case_prod R) \<inter> (fset a \<times> fset b))" (is "the_inv fset ?L'") | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 943 | have "finite ?L'" by (intro finite_Int[OF disjI2] finite_cartesian_product) (transfer, simp)+ | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 944 | hence *: "fset R' = ?L'" unfolding R'_def by (intro fset_to_fset) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 945 | show ?R unfolding Grp_def relcompp.simps conversep.simps | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55129diff
changeset | 946 | proof (intro CollectI case_prodI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl) | 
| 60500 | 947 | from * show "a = fimage fst R'" using conjunct1[OF \<open>?L\<close>] | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 948 | by (transfer, auto simp add: image_def Int_def split: prod.splits) | 
| 60500 | 949 | from * show "b = fimage snd R'" using conjunct2[OF \<open>?L\<close>] | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 950 | by (transfer, auto simp add: image_def Int_def split: prod.splits) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 951 | qed (auto simp add: *) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 952 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 953 | assume ?R thus ?L unfolding Grp_def relcompp.simps conversep.simps | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 954 | apply (simp add: subset_eq Ball_def) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 955 | apply (rule conjI) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 956 | apply (transfer, clarsimp, metis snd_conv) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 957 | by (transfer, clarsimp, metis fst_conv) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 958 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 959 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 960 | bnf "'a fset" | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 961 | map: fimage | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 962 | sets: fset | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 963 | bd: natLeq | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 964 |   wits: "{||}"
 | 
| 55933 | 965 | rel: rel_fset | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 966 | apply - | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 967 | apply transfer' apply simp | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 968 | apply transfer' apply force | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 969 | apply transfer apply force | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 970 | apply transfer' apply force | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 971 | apply (rule natLeq_card_order) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 972 | apply (rule natLeq_cinfinite) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 973 | apply transfer apply (metis ordLess_imp_ordLeq finite_iff_ordLess_natLeq) | 
| 55933 | 974 | apply (fastforce simp: rel_fset_alt) | 
| 62324 | 975 | apply (simp add: Grp_def relcompp.simps conversep.simps fun_eq_iff rel_fset_alt | 
| 976 | rel_fset_aux[unfolded OO_Grp_alt]) | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 977 | apply transfer apply simp | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 978 | done | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 979 | |
| 55938 | 980 | lemma rel_fset_fset: "rel_set \<chi> (fset A1) (fset A2) = rel_fset \<chi> A1 A2" | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 981 | by transfer (rule refl) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 982 | |
| 53953 | 983 | end | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 984 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 985 | lemmas [simp] = fset.map_comp fset.map_id fset.set_map | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 986 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 987 | |
| 60500 | 988 | subsection \<open>Size setup\<close> | 
| 56646 | 989 | |
| 990 | context includes fset.lifting begin | |
| 991 | lift_definition size_fset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a fset \<Rightarrow> nat" is "\<lambda>f. setsum (Suc \<circ> f)" .
 | |
| 992 | end | |
| 993 | ||
| 994 | instantiation fset :: (type) size begin | |
| 995 | definition size_fset where | |
| 996 | size_fset_overloaded_def: "size_fset = FSet.size_fset (\<lambda>_. 0)" | |
| 997 | instance .. | |
| 998 | end | |
| 999 | ||
| 1000 | lemmas size_fset_simps[simp] = | |
| 1001 | size_fset_def[THEN meta_eq_to_obj_eq, THEN fun_cong, THEN fun_cong, | |
| 1002 | unfolded map_fun_def comp_def id_apply] | |
| 1003 | ||
| 1004 | lemmas size_fset_overloaded_simps[simp] = | |
| 1005 | size_fset_simps[of "\<lambda>_. 0", unfolded add_0_left add_0_right, | |
| 1006 | folded size_fset_overloaded_def] | |
| 1007 | ||
| 1008 | lemma fset_size_o_map: "inj f \<Longrightarrow> size_fset g \<circ> fimage f = size_fset (g \<circ> f)" | |
| 60228 
32dd7adba5a4
tuned proof; forget the transfer rule for size_fset
 kuncar parents: 
58881diff
changeset | 1009 | apply (subst fun_eq_iff) | 
| 
32dd7adba5a4
tuned proof; forget the transfer rule for size_fset
 kuncar parents: 
58881diff
changeset | 1010 | including fset.lifting by transfer (auto intro: setsum.reindex_cong subset_inj_on) | 
| 
32dd7adba5a4
tuned proof; forget the transfer rule for size_fset
 kuncar parents: 
58881diff
changeset | 1011 | |
| 60500 | 1012 | setup \<open> | 
| 56651 | 1013 | BNF_LFP_Size.register_size_global @{type_name fset} @{const_name size_fset}
 | 
| 62082 | 1014 |   @{thm size_fset_overloaded_def} @{thms size_fset_simps size_fset_overloaded_simps}
 | 
| 1015 |   @{thms fset_size_o_map}
 | |
| 60500 | 1016 | \<close> | 
| 56646 | 1017 | |
| 60228 
32dd7adba5a4
tuned proof; forget the transfer rule for size_fset
 kuncar parents: 
58881diff
changeset | 1018 | lifting_update fset.lifting | 
| 
32dd7adba5a4
tuned proof; forget the transfer rule for size_fset
 kuncar parents: 
58881diff
changeset | 1019 | lifting_forget fset.lifting | 
| 56646 | 1020 | |
| 60500 | 1021 | subsection \<open>Advanced relator customization\<close> | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1022 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1023 | (* Set vs. sum relators: *) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1024 | |
| 55943 | 1025 | lemma rel_set_rel_sum[simp]: | 
| 1026 | "rel_set (rel_sum \<chi> \<phi>) A1 A2 \<longleftrightarrow> | |
| 55938 | 1027 | rel_set \<chi> (Inl -` A1) (Inl -` A2) \<and> rel_set \<phi> (Inr -` A1) (Inr -` A2)" | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1028 | (is "?L \<longleftrightarrow> ?Rl \<and> ?Rr") | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1029 | proof safe | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1030 | assume L: "?L" | 
| 55938 | 1031 | show ?Rl unfolding rel_set_def Bex_def vimage_eq proof safe | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1032 | fix l1 assume "Inl l1 \<in> A1" | 
| 55943 | 1033 | then obtain a2 where a2: "a2 \<in> A2" and "rel_sum \<chi> \<phi> (Inl l1) a2" | 
| 55938 | 1034 | using L unfolding rel_set_def by auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1035 | then obtain l2 where "a2 = Inl l2 \<and> \<chi> l1 l2" by (cases a2, auto) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1036 | thus "\<exists> l2. Inl l2 \<in> A2 \<and> \<chi> l1 l2" using a2 by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1037 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1038 | fix l2 assume "Inl l2 \<in> A2" | 
| 55943 | 1039 | then obtain a1 where a1: "a1 \<in> A1" and "rel_sum \<chi> \<phi> a1 (Inl l2)" | 
| 55938 | 1040 | using L unfolding rel_set_def by auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1041 | then obtain l1 where "a1 = Inl l1 \<and> \<chi> l1 l2" by (cases a1, auto) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1042 | thus "\<exists> l1. Inl l1 \<in> A1 \<and> \<chi> l1 l2" using a1 by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1043 | qed | 
| 55938 | 1044 | show ?Rr unfolding rel_set_def Bex_def vimage_eq proof safe | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1045 | fix r1 assume "Inr r1 \<in> A1" | 
| 55943 | 1046 | then obtain a2 where a2: "a2 \<in> A2" and "rel_sum \<chi> \<phi> (Inr r1) a2" | 
| 55938 | 1047 | using L unfolding rel_set_def by auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1048 | then obtain r2 where "a2 = Inr r2 \<and> \<phi> r1 r2" by (cases a2, auto) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1049 | thus "\<exists> r2. Inr r2 \<in> A2 \<and> \<phi> r1 r2" using a2 by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1050 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1051 | fix r2 assume "Inr r2 \<in> A2" | 
| 55943 | 1052 | then obtain a1 where a1: "a1 \<in> A1" and "rel_sum \<chi> \<phi> a1 (Inr r2)" | 
| 55938 | 1053 | using L unfolding rel_set_def by auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1054 | then obtain r1 where "a1 = Inr r1 \<and> \<phi> r1 r2" by (cases a1, auto) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1055 | thus "\<exists> r1. Inr r1 \<in> A1 \<and> \<phi> r1 r2" using a1 by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1056 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1057 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1058 | assume Rl: "?Rl" and Rr: "?Rr" | 
| 55938 | 1059 | show ?L unfolding rel_set_def Bex_def vimage_eq proof safe | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1060 | fix a1 assume a1: "a1 \<in> A1" | 
| 55943 | 1061 | show "\<exists> a2. a2 \<in> A2 \<and> rel_sum \<chi> \<phi> a1 a2" | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1062 | proof(cases a1) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1063 | case (Inl l1) then obtain l2 where "Inl l2 \<in> A2 \<and> \<chi> l1 l2" | 
| 55938 | 1064 | using Rl a1 unfolding rel_set_def by blast | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1065 | thus ?thesis unfolding Inl by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1066 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1067 | case (Inr r1) then obtain r2 where "Inr r2 \<in> A2 \<and> \<phi> r1 r2" | 
| 55938 | 1068 | using Rr a1 unfolding rel_set_def by blast | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1069 | thus ?thesis unfolding Inr by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1070 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1071 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1072 | fix a2 assume a2: "a2 \<in> A2" | 
| 55943 | 1073 | show "\<exists> a1. a1 \<in> A1 \<and> rel_sum \<chi> \<phi> a1 a2" | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1074 | proof(cases a2) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1075 | case (Inl l2) then obtain l1 where "Inl l1 \<in> A1 \<and> \<chi> l1 l2" | 
| 55938 | 1076 | using Rl a2 unfolding rel_set_def by blast | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1077 | thus ?thesis unfolding Inl by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1078 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1079 | case (Inr r2) then obtain r1 where "Inr r1 \<in> A1 \<and> \<phi> r1 r2" | 
| 55938 | 1080 | using Rr a2 unfolding rel_set_def by blast | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1081 | thus ?thesis unfolding Inr by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1082 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1083 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1084 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1085 | |
| 60712 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1086 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1087 | subsection \<open>Quickcheck setup\<close> | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1088 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1089 | text \<open>Setup adapted from sets.\<close> | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1090 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1091 | notation Quickcheck_Exhaustive.orelse (infixr "orelse" 55) | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1092 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1093 | definition (in term_syntax) [code_unfold]: | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1094 | "valterm_femptyset = Code_Evaluation.valtermify ({||} :: ('a :: typerep) fset)"
 | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1095 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1096 | definition (in term_syntax) [code_unfold]: | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1097 | "valtermify_finsert x s = Code_Evaluation.valtermify finsert {\<cdot>} (x :: ('a :: typerep * _)) {\<cdot>} s"
 | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1098 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1099 | instantiation fset :: (exhaustive) exhaustive | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1100 | begin | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1101 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1102 | fun exhaustive_fset where | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1103 | "exhaustive_fset f i = (if i = 0 then None else (f {||} orelse exhaustive_fset (\<lambda>A. f A orelse Quickcheck_Exhaustive.exhaustive (\<lambda>x. if x |\<in>| A then None else f (finsert x A)) (i - 1)) (i - 1)))"
 | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1104 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1105 | instance .. | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1106 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54258diff
changeset | 1107 | end | 
| 60712 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1108 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1109 | instantiation fset :: (full_exhaustive) full_exhaustive | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1110 | begin | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1111 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1112 | fun full_exhaustive_fset where | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1113 | "full_exhaustive_fset f i = (if i = 0 then None else (f valterm_femptyset orelse full_exhaustive_fset (\<lambda>A. f A orelse Quickcheck_Exhaustive.full_exhaustive (\<lambda>x. if fst x |\<in>| fst A then None else f (valtermify_finsert x A)) (i - 1)) (i - 1)))" | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1114 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1115 | instance .. | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1116 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1117 | end | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1118 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1119 | no_notation Quickcheck_Exhaustive.orelse (infixr "orelse" 55) | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1120 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1121 | notation scomp (infixl "\<circ>\<rightarrow>" 60) | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1122 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1123 | instantiation fset :: (random) random | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1124 | begin | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1125 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1126 | fun random_aux_fset :: "natural \<Rightarrow> natural \<Rightarrow> natural \<times> natural \<Rightarrow> ('a fset \<times> (unit \<Rightarrow> term)) \<times> natural \<times> natural" where
 | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1127 | "random_aux_fset 0 j = Quickcheck_Random.collapse (Random.select_weight [(1, Pair valterm_femptyset)])" | | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1128 | "random_aux_fset (Code_Numeral.Suc i) j = | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1129 | Quickcheck_Random.collapse (Random.select_weight | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1130 | [(1, Pair valterm_femptyset), | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1131 | (Code_Numeral.Suc i, | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1132 | Quickcheck_Random.random j \<circ>\<rightarrow> (\<lambda>x. random_aux_fset i j \<circ>\<rightarrow> (\<lambda>s. Pair (valtermify_finsert x s))))])" | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1133 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1134 | lemma [code]: | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1135 | "random_aux_fset i j = | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1136 | Quickcheck_Random.collapse (Random.select_weight [(1, Pair valterm_femptyset), | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1137 | (i, Quickcheck_Random.random j \<circ>\<rightarrow> (\<lambda>x. random_aux_fset (i - 1) j \<circ>\<rightarrow> (\<lambda>s. Pair (valtermify_finsert x s))))])" | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1138 | proof (induct i rule: natural.induct) | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1139 | case zero | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1140 | show ?case by (subst select_weight_drop_zero[symmetric]) (simp add: less_natural_def) | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1141 | next | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1142 | case (Suc i) | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1143 | show ?case by (simp only: random_aux_fset.simps Suc_natural_minus_one) | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1144 | qed | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1145 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1146 | definition "random_fset i = random_aux_fset i i" | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1147 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1148 | instance .. | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1149 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1150 | end | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1151 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1152 | no_notation scomp (infixl "\<circ>\<rightarrow>" 60) | 
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1153 | |
| 
3ba16d28449d
Quickcheck setup for finite sets
 Lars Hupel <lars.hupel@mytum.de> parents: 
60679diff
changeset | 1154 | end |