blanchet@55075: (* Title: HOL/Basic_BNFs.thy blanchet@48975: Author: Dmitriy Traytel, TU Muenchen blanchet@48975: Author: Andrei Popescu, TU Muenchen blanchet@48975: Author: Jasmin Blanchette, TU Muenchen blanchet@48975: Copyright 2012 blanchet@48975: blanchet@49309: Registration of basic types as bounded natural functors. blanchet@48975: *) blanchet@48975: blanchet@49309: header {* Registration of Basic Types as Bounded Natural Functors *} blanchet@48975: blanchet@48975: theory Basic_BNFs blanchet@49310: imports BNF_Def blanchet@48975: begin blanchet@48975: blanchet@49451: definition setl :: "'a + 'b \ 'a set" where blanchet@49451: "setl x = (case x of Inl z => {z} | _ => {})" blanchet@48975: blanchet@49451: definition setr :: "'a + 'b \ 'b set" where blanchet@49451: "setr x = (case x of Inr z => {z} | _ => {})" blanchet@48975: blanchet@49451: lemmas sum_set_defs = setl_def[abs_def] setr_def[abs_def] blanchet@48975: blanchet@55083: definition blanchet@55943: rel_sum :: "('a \ 'c \ bool) \ ('b \ 'd \ bool) \ 'a + 'b \ 'c + 'd \ bool" blanchet@55083: where blanchet@55943: "rel_sum R1 R2 x y = blanchet@55083: (case (x, y) of (Inl x, Inl y) \ R1 x y blanchet@55083: | (Inr x, Inr y) \ R2 x y blanchet@55083: | _ \ False)" blanchet@55083: blanchet@55943: lemma rel_sum_simps[simp]: blanchet@55943: "rel_sum R1 R2 (Inl a1) (Inl b1) = R1 a1 b1" blanchet@55943: "rel_sum R1 R2 (Inl a1) (Inr b2) = False" blanchet@55943: "rel_sum R1 R2 (Inr a2) (Inl b1) = False" blanchet@55943: "rel_sum R1 R2 (Inr a2) (Inr b2) = R2 a2 b2" blanchet@55943: unfolding rel_sum_def by simp_all blanchet@55083: traytel@54421: bnf "'a + 'b" blanchet@55931: map: map_sum traytel@54421: sets: setl setr traytel@54421: bd: natLeq traytel@54421: wits: Inl Inr blanchet@55943: rel: rel_sum blanchet@48975: proof - blanchet@55931: show "map_sum id id = id" by (rule map_sum.id) blanchet@48975: next blanchet@54486: fix f1 :: "'o \ 's" and f2 :: "'p \ 't" and g1 :: "'s \ 'q" and g2 :: "'t \ 'r" blanchet@55931: show "map_sum (g1 o f1) (g2 o f2) = map_sum g1 g2 o map_sum f1 f2" blanchet@55931: by (rule map_sum.comp[symmetric]) blanchet@48975: next blanchet@54486: fix x and f1 :: "'o \ 'q" and f2 :: "'p \ 'r" and g1 g2 blanchet@49451: assume a1: "\z. z \ setl x \ f1 z = g1 z" and blanchet@49451: a2: "\z. z \ setr x \ f2 z = g2 z" blanchet@55931: thus "map_sum f1 f2 x = map_sum g1 g2 x" blanchet@48975: proof (cases x) blanchet@49451: case Inl thus ?thesis using a1 by (clarsimp simp: setl_def) blanchet@48975: next blanchet@49451: case Inr thus ?thesis using a2 by (clarsimp simp: setr_def) blanchet@48975: qed blanchet@48975: next blanchet@54486: fix f1 :: "'o \ 'q" and f2 :: "'p \ 'r" blanchet@55931: show "setl o map_sum f1 f2 = image f1 o setl" blanchet@49451: by (rule ext, unfold o_apply) (simp add: setl_def split: sum.split) blanchet@48975: next blanchet@54486: fix f1 :: "'o \ 'q" and f2 :: "'p \ 'r" blanchet@55931: show "setr o map_sum f1 f2 = image f2 o setr" blanchet@49451: by (rule ext, unfold o_apply) (simp add: setr_def split: sum.split) blanchet@48975: next blanchet@48975: show "card_order natLeq" by (rule natLeq_card_order) blanchet@48975: next blanchet@48975: show "cinfinite natLeq" by (rule natLeq_cinfinite) blanchet@48975: next blanchet@54486: fix x :: "'o + 'p" blanchet@49451: show "|setl x| \o natLeq" blanchet@48975: apply (rule ordLess_imp_ordLeq) blanchet@48975: apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) blanchet@49451: by (simp add: setl_def split: sum.split) blanchet@48975: next blanchet@54486: fix x :: "'o + 'p" blanchet@49451: show "|setr x| \o natLeq" blanchet@48975: apply (rule ordLess_imp_ordLeq) blanchet@48975: apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) blanchet@49451: by (simp add: setr_def split: sum.split) blanchet@48975: next traytel@54841: fix R1 R2 S1 S2 blanchet@55943: show "rel_sum R1 R2 OO rel_sum S1 S2 \ rel_sum (R1 OO S1) (R2 OO S2)" blanchet@55943: by (auto simp: rel_sum_def split: sum.splits) blanchet@49453: next blanchet@49453: fix R S blanchet@55943: show "rel_sum R S = blanchet@55931: (Grp {x. setl x \ Collect (split R) \ setr x \ Collect (split S)} (map_sum fst fst))\\ OO blanchet@55931: Grp {x. setl x \ Collect (split R) \ setr x \ Collect (split S)} (map_sum snd snd)" blanchet@55943: unfolding setl_def setr_def rel_sum_def Grp_def relcompp.simps conversep.simps fun_eq_iff blanchet@49453: by (fastforce split: sum.splits) blanchet@48975: qed (auto simp: sum_set_defs) blanchet@48975: blanchet@48975: definition fsts :: "'a \ 'b \ 'a set" where blanchet@48975: "fsts x = {fst x}" blanchet@48975: blanchet@48975: definition snds :: "'a \ 'b \ 'b set" where blanchet@48975: "snds x = {snd x}" blanchet@48975: blanchet@48975: lemmas prod_set_defs = fsts_def[abs_def] snds_def[abs_def] blanchet@48975: blanchet@55083: definition blanchet@55944: rel_prod :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ 'a \ 'c \ 'b \ 'd \ bool" blanchet@55083: where blanchet@55944: "rel_prod R1 R2 = (\(a, b) (c, d). R1 a c \ R2 b d)" blanchet@55083: blanchet@55944: lemma rel_prod_apply [simp]: blanchet@55944: "rel_prod R1 R2 (a, b) (c, d) \ R1 a c \ R2 b d" blanchet@55944: by (simp add: rel_prod_def) blanchet@55083: traytel@54421: bnf "'a \ 'b" blanchet@55932: map: map_prod traytel@54421: sets: fsts snds traytel@54421: bd: natLeq blanchet@55944: rel: rel_prod blanchet@48975: proof (unfold prod_set_defs) blanchet@55932: show "map_prod id id = id" by (rule map_prod.id) blanchet@48975: next blanchet@48975: fix f1 f2 g1 g2 blanchet@55932: show "map_prod (g1 o f1) (g2 o f2) = map_prod g1 g2 o map_prod f1 f2" blanchet@55932: by (rule map_prod.comp[symmetric]) blanchet@48975: next blanchet@48975: fix x f1 f2 g1 g2 blanchet@48975: assume "\z. z \ {fst x} \ f1 z = g1 z" "\z. z \ {snd x} \ f2 z = g2 z" blanchet@55932: thus "map_prod f1 f2 x = map_prod g1 g2 x" by (cases x) simp blanchet@48975: next blanchet@48975: fix f1 f2 blanchet@55932: show "(\x. {fst x}) o map_prod f1 f2 = image f1 o (\x. {fst x})" blanchet@48975: by (rule ext, unfold o_apply) simp blanchet@48975: next blanchet@48975: fix f1 f2 blanchet@55932: show "(\x. {snd x}) o map_prod f1 f2 = image f2 o (\x. {snd x})" blanchet@48975: by (rule ext, unfold o_apply) simp blanchet@48975: next traytel@52635: show "card_order natLeq" by (rule natLeq_card_order) blanchet@48975: next traytel@52635: show "cinfinite natLeq" by (rule natLeq_cinfinite) blanchet@48975: next blanchet@48975: fix x traytel@52635: show "|{fst x}| \o natLeq" traytel@55811: by (rule ordLess_imp_ordLeq) (simp add: finite_iff_ordLess_natLeq[symmetric]) blanchet@48975: next traytel@52635: fix x traytel@52635: show "|{snd x}| \o natLeq" traytel@55811: by (rule ordLess_imp_ordLeq) (simp add: finite_iff_ordLess_natLeq[symmetric]) blanchet@48975: next traytel@54841: fix R1 R2 S1 S2 blanchet@55944: show "rel_prod R1 R2 OO rel_prod S1 S2 \ rel_prod (R1 OO S1) (R2 OO S2)" by auto blanchet@49453: next blanchet@49453: fix R S blanchet@55944: show "rel_prod R S = blanchet@55932: (Grp {x. {fst x} \ Collect (split R) \ {snd x} \ Collect (split S)} (map_prod fst fst))\\ OO blanchet@55932: Grp {x. {fst x} \ Collect (split R) \ {snd x} \ Collect (split S)} (map_prod snd snd)" blanchet@55944: unfolding prod_set_defs rel_prod_def Grp_def relcompp.simps conversep.simps fun_eq_iff blanchet@49453: by auto traytel@54189: qed blanchet@48975: traytel@54421: bnf "'a \ 'b" traytel@54421: map: "op \" traytel@54421: sets: range traytel@54421: bd: "natLeq +c |UNIV :: 'a set|" blanchet@55945: rel: "rel_fun op =" blanchet@48975: proof blanchet@48975: fix f show "id \ f = id f" by simp blanchet@48975: next blanchet@48975: fix f g show "op \ (g \ f) = op \ g \ op \ f" blanchet@48975: unfolding comp_def[abs_def] .. blanchet@48975: next blanchet@48975: fix x f g blanchet@48975: assume "\z. z \ range x \ f z = g z" blanchet@48975: thus "f \ x = g \ x" by auto blanchet@48975: next blanchet@48975: fix f show "range \ op \ f = op ` f \ range" haftmann@56077: by (auto simp add: fun_eq_iff) blanchet@48975: next blanchet@48975: show "card_order (natLeq +c |UNIV| )" (is "_ (_ +c ?U)") blanchet@48975: apply (rule card_order_csum) blanchet@48975: apply (rule natLeq_card_order) blanchet@48975: by (rule card_of_card_order_on) blanchet@48975: (* *) blanchet@48975: show "cinfinite (natLeq +c ?U)" blanchet@48975: apply (rule cinfinite_csum) blanchet@48975: apply (rule disjI1) blanchet@48975: by (rule natLeq_cinfinite) blanchet@48975: next blanchet@48975: fix f :: "'d => 'a" blanchet@48975: have "|range f| \o | (UNIV::'d set) |" (is "_ \o ?U") by (rule card_of_image) blanchet@54486: also have "?U \o natLeq +c ?U" by (rule ordLeq_csum2) (rule card_of_Card_order) blanchet@48975: finally show "|range f| \o natLeq +c ?U" . blanchet@48975: next traytel@54841: fix R S blanchet@55945: show "rel_fun op = R OO rel_fun op = S \ rel_fun op = (R OO S)" by (auto simp: rel_fun_def) blanchet@49453: next blanchet@49463: fix R blanchet@55945: show "rel_fun op = R = traytel@51893: (Grp {x. range x \ Collect (split R)} (op \ fst))\\ OO traytel@51893: Grp {x. range x \ Collect (split R)} (op \ snd)" blanchet@55945: unfolding rel_fun_def Grp_def fun_eq_iff relcompp.simps conversep.simps subset_iff image_iff traytel@55811: comp_apply mem_Collect_eq split_beta bex_UNIV traytel@55811: proof (safe, unfold fun_eq_iff[symmetric]) traytel@55811: fix x xa a b c xb y aa ba traytel@55811: assume *: "x = a" "xa = c" "a = ba" "b = aa" "c = (\x. snd (b x))" "ba = (\x. fst (aa x))" and traytel@55811: **: "\t. (\x. t = aa x) \ R (fst t) (snd t)" traytel@55811: show "R (x y) (xa y)" unfolding * by (rule mp[OF spec[OF **]]) blast traytel@55811: qed force traytel@54189: qed traytel@54191: blanchet@48975: end