193 |
193 |
194 instance .. |
194 instance .. |
195 |
195 |
196 end |
196 end |
197 |
197 |
|
198 |
|
199 section \<open>alist is a BNF\<close> |
|
200 |
|
201 lift_definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> ('k, 'a) alist \<Rightarrow> ('k, 'b) alist" |
|
202 is "\<lambda>f xs. List.map (map_prod id f) xs" by simp |
|
203 |
|
204 lift_definition set :: "('k, 'v) alist => 'v set" |
|
205 is "\<lambda>xs. snd ` List.set xs" . |
|
206 |
|
207 lift_definition rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('k, 'a) alist \<Rightarrow> ('k, 'b) alist \<Rightarrow> bool" |
|
208 is "\<lambda>R xs ys. list_all2 (rel_prod op = R) xs ys" . |
|
209 |
|
210 bnf "('k, 'v) alist" |
|
211 map: map |
|
212 sets: set |
|
213 bd: natLeq |
|
214 wits: empty |
|
215 rel: rel |
|
216 proof (unfold OO_Grp_alt) |
|
217 show "map id = id" by (rule ext, transfer) (simp add: prod.map_id0) |
|
218 next |
|
219 fix f g |
|
220 show "map (g \<circ> f) = map g \<circ> map f" |
|
221 by (rule ext, transfer) (simp add: prod.map_comp) |
|
222 next |
|
223 fix x f g |
|
224 assume "(\<And>z. z \<in> set x \<Longrightarrow> f z = g z)" |
|
225 then show "map f x = map g x" by transfer force |
|
226 next |
|
227 fix f |
|
228 show "set \<circ> map f = op ` f \<circ> set" |
|
229 by (rule ext, transfer) (simp add: image_image) |
|
230 next |
|
231 fix x |
|
232 show "ordLeq3 (card_of (set x)) natLeq" |
|
233 by transfer (auto simp: finite_iff_ordLess_natLeq[symmetric] intro: ordLess_imp_ordLeq) |
|
234 next |
|
235 fix R S |
|
236 show "rel R OO rel S \<le> rel (R OO S)" |
|
237 by (rule predicate2I, transfer) |
|
238 (auto simp: list.rel_compp prod.rel_compp[of "op =", unfolded eq_OO]) |
|
239 next |
|
240 fix R |
|
241 show "rel R = (\<lambda>x y. \<exists>z. z \<in> {x. set x \<subseteq> {(x, y). R x y}} \<and> map fst z = x \<and> map snd z = y)" |
|
242 unfolding fun_eq_iff by transfer (fastforce simp: list.in_rel o_def intro: |
|
243 exI[of _ "List.map (\<lambda>p. ((fst p, fst (snd p)), (fst p, snd (snd p)))) z" for z] |
|
244 exI[of _ "List.map (\<lambda>p. (fst (fst p), snd (fst p), snd (snd p))) z" for z]) |
|
245 next |
|
246 fix z assume "z \<in> set empty" |
|
247 then show False by transfer simp |
|
248 qed (simp_all add: natLeq_cinfinite natLeq_card_order) |
|
249 |
198 hide_const valterm_empty valterm_update random_aux_alist |
250 hide_const valterm_empty valterm_update random_aux_alist |
199 |
251 |
200 hide_fact (open) lookup_def empty_def update_def delete_def map_entry_def filter_def map_default_def |
252 hide_fact (open) lookup_def empty_def update_def delete_def map_entry_def filter_def map_default_def |
201 hide_const (open) impl_of lookup empty update delete map_entry filter map_default |
253 hide_const (open) impl_of lookup empty update delete map_entry filter map_default map set rel |
202 |
254 |
203 end |
255 end |