194 |
197 |
195 subsection \<open>Quickcheck generators\<close> |
198 subsection \<open>Quickcheck generators\<close> |
196 |
199 |
197 quickcheck_generator dlist predicate: distinct constructors: Dlist.empty, Dlist.insert |
200 quickcheck_generator dlist predicate: distinct constructors: Dlist.empty, Dlist.insert |
198 |
201 |
199 end |
202 subsection \<open>BNF instance\<close> |
|
203 |
|
204 context begin |
|
205 |
|
206 qualified fun wpull :: "('a \<times> 'b) list \<Rightarrow> ('b \<times> 'c) list \<Rightarrow> ('a \<times> 'c) list" |
|
207 where |
|
208 "wpull [] ys = []" |
|
209 | "wpull xs [] = []" |
|
210 | "wpull ((a, b) # xs) ((b', c) # ys) = |
|
211 (if b \<in> snd ` set xs then |
|
212 (a, the (map_of (rev ((b', c) # ys)) b)) # wpull xs ((b', c) # ys) |
|
213 else if b' \<in> fst ` set ys then |
|
214 (the (map_of (map prod.swap (rev ((a, b) # xs))) b'), c) # wpull ((a, b) # xs) ys |
|
215 else (a, c) # wpull xs ys)" |
|
216 |
|
217 qualified lemma wpull_eq_Nil_iff [simp]: "wpull xs ys = [] \<longleftrightarrow> xs = [] \<or> ys = []" |
|
218 by(cases "(xs, ys)" rule: wpull.cases) simp_all |
|
219 |
|
220 qualified lemma wpull_induct |
|
221 [consumes 1, |
|
222 case_names Nil left[xs eq in_set IH] right[xs ys eq in_set IH] step[xs ys eq IH] ]: |
|
223 assumes eq: "remdups (map snd xs) = remdups (map fst ys)" |
|
224 and Nil: "P [] []" |
|
225 and left: "\<And>a b xs b' c ys. |
|
226 \<lbrakk> b \<in> snd ` set xs; remdups (map snd xs) = remdups (map fst ((b', c) # ys)); |
|
227 (b, the (map_of (rev ((b', c) # ys)) b)) \<in> set ((b', c) # ys); P xs ((b', c) # ys) \<rbrakk> |
|
228 \<Longrightarrow> P ((a, b) # xs) ((b', c) # ys)" |
|
229 and right: "\<And>a b xs b' c ys. |
|
230 \<lbrakk> b \<notin> snd ` set xs; b' \<in> fst ` set ys; |
|
231 remdups (map snd ((a, b) # xs)) = remdups (map fst ys); |
|
232 (the (map_of (map prod.swap (rev ((a, b) #xs))) b'), b') \<in> set ((a, b) # xs); |
|
233 P ((a, b) # xs) ys \<rbrakk> |
|
234 \<Longrightarrow> P ((a, b) # xs) ((b', c) # ys)" |
|
235 and step: "\<And>a b xs c ys. |
|
236 \<lbrakk> b \<notin> snd ` set xs; b \<notin> fst ` set ys; remdups (map snd xs) = remdups (map fst ys); |
|
237 P xs ys \<rbrakk> |
|
238 \<Longrightarrow> P ((a, b) # xs) ((b, c) # ys)" |
|
239 shows "P xs ys" |
|
240 using eq |
|
241 proof(induction xs ys rule: wpull.induct) |
|
242 case 1 thus ?case by(simp add: Nil) |
|
243 next |
|
244 case 2 thus ?case by(simp split: split_if_asm) |
|
245 next |
|
246 case Cons: (3 a b xs b' c ys) |
|
247 let ?xs = "(a, b) # xs" and ?ys = "(b', c) # ys" |
|
248 consider (xs) "b \<in> snd ` set xs" | (ys) "b \<notin> snd ` set xs" "b' \<in> fst ` set ys" |
|
249 | (step) "b \<notin> snd ` set xs" "b' \<notin> fst ` set ys" by auto |
|
250 thus ?case |
|
251 proof cases |
|
252 case xs |
|
253 with Cons.prems have eq: "remdups (map snd xs) = remdups (map fst ?ys)" by auto |
|
254 from xs eq have "b \<in> fst ` set ?ys" by (metis list.set_map set_remdups) |
|
255 hence "map_of (rev ?ys) b \<noteq> None" unfolding map_of_eq_None_iff by auto |
|
256 then obtain c' where "map_of (rev ?ys) b = Some c'" by blast |
|
257 then have "(b, the (map_of (rev ?ys) b)) \<in> set ?ys" by(auto dest: map_of_SomeD split: split_if_asm) |
|
258 from xs eq this Cons.IH(1)[OF xs eq] show ?thesis by(rule left) |
|
259 next |
|
260 case ys |
|
261 from ys Cons.prems have eq: "remdups (map snd ?xs) = remdups (map fst ys)" by auto |
|
262 from ys eq have "b' \<in> snd ` set ?xs" by (metis list.set_map set_remdups) |
|
263 hence "map_of (map prod.swap (rev ?xs)) b' \<noteq> None" |
|
264 unfolding map_of_eq_None_iff by(auto simp add: image_image) |
|
265 then obtain a' where "map_of (map prod.swap (rev ?xs)) b' = Some a'" by blast |
|
266 then have "(the (map_of (map prod.swap (rev ?xs)) b'), b') \<in> set ?xs" |
|
267 by(auto dest: map_of_SomeD split: split_if_asm) |
|
268 from ys eq this Cons.IH(2)[OF ys eq] show ?thesis by(rule right) |
|
269 next |
|
270 case *: step |
|
271 hence "remdups (map snd xs) = remdups (map fst ys)" "b = b'" using Cons.prems by auto |
|
272 from * this(1) Cons.IH(3)[OF * this(1)] show ?thesis unfolding \<open>b = b'\<close> by(rule step) |
|
273 qed |
|
274 qed |
|
275 |
|
276 qualified lemma set_wpull_subset: |
|
277 assumes "remdups (map snd xs) = remdups (map fst ys)" |
|
278 shows "set (wpull xs ys) \<subseteq> set xs O set ys" |
|
279 using assms by(induction xs ys rule: wpull_induct) auto |
|
280 |
|
281 qualified lemma set_fst_wpull: |
|
282 assumes "remdups (map snd xs) = remdups (map fst ys)" |
|
283 shows "fst ` set (wpull xs ys) = fst ` set xs" |
|
284 using assms by(induction xs ys rule: wpull_induct)(auto intro: rev_image_eqI) |
|
285 |
|
286 qualified lemma set_snd_wpull: |
|
287 assumes "remdups (map snd xs) = remdups (map fst ys)" |
|
288 shows "snd ` set (wpull xs ys) = snd ` set ys" |
|
289 using assms by(induction xs ys rule: wpull_induct)(auto intro: rev_image_eqI) |
|
290 |
|
291 qualified lemma wpull: |
|
292 assumes "distinct xs" |
|
293 and "distinct ys" |
|
294 and "set xs \<subseteq> {(x, y). R x y}" |
|
295 and "set ys \<subseteq> {(x, y). S x y}" |
|
296 and eq: "remdups (map snd xs) = remdups (map fst ys)" |
|
297 shows "\<exists>zs. distinct zs \<and> set zs \<subseteq> {(x, y). (R OO S) x y} \<and> |
|
298 remdups (map fst zs) = remdups (map fst xs) \<and> remdups (map snd zs) = remdups (map snd ys)" |
|
299 proof(intro exI conjI) |
|
300 let ?zs = "remdups (wpull xs ys)" |
|
301 show "distinct ?zs" by simp |
|
302 show "set ?zs \<subseteq> {(x, y). (R OO S) x y}" using assms(3-4) set_wpull_subset[OF eq] by fastforce |
|
303 show "remdups (map fst ?zs) = remdups (map fst xs)" unfolding remdups_map_remdups using eq |
|
304 by(induction xs ys rule: wpull_induct)(auto simp add: set_fst_wpull intro: rev_image_eqI) |
|
305 show "remdups (map snd ?zs) = remdups (map snd ys)" unfolding remdups_map_remdups using eq |
|
306 by(induction xs ys rule: wpull_induct)(auto simp add: set_snd_wpull intro: rev_image_eqI) |
|
307 qed |
|
308 |
|
309 qualified lift_definition set :: "'a dlist \<Rightarrow> 'a set" is List.set . |
|
310 |
|
311 qualified lemma map_transfer [transfer_rule]: |
|
312 "(rel_fun op = (rel_fun (pcr_dlist op =) (pcr_dlist op =))) (\<lambda>f x. remdups (List.map f x)) Dlist.map" |
|
313 by(simp add: rel_fun_def dlist.pcr_cr_eq cr_dlist_def Dlist.map_def remdups_remdups) |
|
314 |
|
315 bnf "'a dlist" |
|
316 map: Dlist.map |
|
317 sets: set |
|
318 bd: natLeq |
|
319 wits: Dlist.empty |
|
320 unfolding OO_Grp_alt mem_Collect_eq |
|
321 subgoal by(rule ext)(simp add: dlist_eq_iff) |
|
322 subgoal by(rule ext)(simp add: dlist_eq_iff remdups_map_remdups) |
|
323 subgoal by(simp add: dlist_eq_iff set_def cong: list.map_cong) |
|
324 subgoal by(simp add: set_def fun_eq_iff) |
|
325 subgoal by(simp add: natLeq_card_order) |
|
326 subgoal by(simp add: natLeq_cinfinite) |
|
327 subgoal by(rule ordLess_imp_ordLeq)(simp add: finite_iff_ordLess_natLeq[symmetric] set_def) |
|
328 subgoal by(rule predicate2I)(transfer; auto simp add: wpull) |
|
329 subgoal by(rule refl) |
|
330 subgoal by(simp add: set_def) |
|
331 done |
|
332 |
|
333 lifting_update dlist.lifting |
|
334 lifting_forget dlist.lifting |
|
335 |
|
336 end |
|
337 |
|
338 end |