# HG changeset patch # User blanchet # Date 1347376179 -7200 # Node ID c057e1b39f162cfdecf7c97f48c12661f7d529fe # Parent 3d87f4fd0d504807b6fe6fa2de254a150e22cee6 renamed "BNF_Library" to "BNF_Util" diff -r 3d87f4fd0d50 -r c057e1b39f16 src/HOL/Codatatype/BNF_Def.thy --- a/src/HOL/Codatatype/BNF_Def.thy Tue Sep 11 17:06:27 2012 +0200 +++ b/src/HOL/Codatatype/BNF_Def.thy Tue Sep 11 17:09:39 2012 +0200 @@ -8,7 +8,7 @@ header {* Definition of Bounded Natural Functors *} theory BNF_Def -imports BNF_Library +imports BNF_Util keywords "print_bnfs" :: diag and diff -r 3d87f4fd0d50 -r c057e1b39f16 src/HOL/Codatatype/BNF_Library.thy --- a/src/HOL/Codatatype/BNF_Library.thy Tue Sep 11 17:06:27 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,853 +0,0 @@ -(* Title: HOL/Codatatype/BNF_Library.thy - Author: Dmitriy Traytel, TU Muenchen - Copyright 2012 - -Library for bounded natural functors. -*) - -header {* Library for Bounded Natural Functors *} - -theory BNF_Library -imports - "../Ordinals_and_Cardinals/Cardinal_Arithmetic" - "~~/src/HOL/Library/Prefix_Order" - Equiv_Relations_More -begin - -lemma iffI_np: "\x \ \ y; \ x \ y\ \ \ x \ y" -by (erule iffI) (erule contrapos_pn) - -lemma all_unit_eq: "(\x. PROP P x) \ PROP P ()" by simp - -lemma all_prod_eq: "(\x. PROP P x) \ (\a b. PROP P (a, b))" by clarsimp - -lemma False_imp_eq: "(False \ P) \ Trueprop True" -by presburger - -lemma case_unit: "(case u of () => f) = f" -by (cases u) (hypsubst, rule unit.cases) - -lemma All_point_1: "(\z. z = b \ phi z) \ Trueprop (phi b)" -by presburger - -lemma subset_Collect_iff: "B \ A \ (B \ {x \ A. P x}) = (\x \ B. P x)" -by blast - -lemma subset_CollectI: "B \ A \ (\x. x \ B \ Q x \ P x) \ ({x \ B. Q x} \ {x \ A. P x})" -by blast - -lemma mem_Collect_eq_split: "{(x, y). (x, y) \ X} = X" -by simp - -lemma image_comp: "image (f o g) = image f o image g" -by (rule ext) (auto simp only: o_apply image_def) - -lemma empty_natural: "(\_. {}) o f = image g o (\_. {})" -by (rule ext) simp - -lemma Union_natural: "Union o image (image f) = image f o Union" -by (rule ext) (auto simp only: o_apply) - -lemma in_Union_o_assoc: "x \ (Union o gset o gmap) A \ x \ (Union o (gset o gmap)) A" -by (unfold o_assoc) - -lemma comp_single_set_bd: - assumes fbd_Card_order: "Card_order fbd" and - fset_bd: "\x. |fset x| \o fbd" and - gset_bd: "\x. |gset x| \o gbd" - shows "|\fset ` gset x| \o gbd *c fbd" -apply (subst sym[OF SUP_def]) -apply (rule ordLeq_transitive) -apply (rule card_of_UNION_Sigma) -apply (subst SIGMA_CSUM) -apply (rule ordLeq_transitive) -apply (rule card_of_Csum_Times') -apply (rule fbd_Card_order) -apply (rule ballI) -apply (rule fset_bd) -apply (rule ordLeq_transitive) -apply (rule cprod_mono1) -apply (rule gset_bd) -apply (rule ordIso_imp_ordLeq) -apply (rule ordIso_refl) -apply (rule Card_order_cprod) -done - -lemma Union_image_insert: "\f ` insert a B = f a \ \f ` B" -by simp - -lemma Union_image_empty: "A \ \f ` {} = A" -by simp - -definition collect where - "collect F x = (\f \ F. f x)" - -lemma collect_o: "collect F o g = collect ((\f. f o g) ` F)" -by (rule ext) (auto simp only: o_apply collect_def) - -lemma image_o_collect: "collect ((\f. image g o f) ` F) = image g o collect F" -by (rule ext) (auto simp add: collect_def) - -lemma conj_subset_def: "A \ {x. P x \ Q x} = (A \ {x. P x} \ A \ {x. Q x})" -by blast - -lemma subset_emptyI: "(\x. x \ A \ False) \ A \ {}" -by blast - -lemma rev_bspec: "a \ A \ \z \ A. P z \ P a" -by simp - -lemma Un_cong: "\A = B; C = D\ \ A \ C = B \ D" -by simp - -lemma UN_image_subset: "\f ` g x \ X = (g x \ {x. f x \ X})" -by blast - -lemma image_Collect_subsetI: - "(\x. P x \ f x \ B) \ f ` {x. P x} \ B" -by blast - -lemma comp_set_bd_Union_o_collect: "|\\(\f. f x) ` X| \o hbd \ |(Union \ collect X) x| \o hbd" -by (unfold o_apply collect_def SUP_def) - -lemma sum_case_comp_Inl: -"sum_case f g \ Inl = f" -unfolding comp_def by simp - -lemma sum_case_expand_Inr: "f o Inl = g \ f x = sum_case g (f o Inr) x" -by (auto split: sum.splits) - -lemma converse_mono: -"R1 ^-1 \ R2 ^-1 \ R1 \ R2" -unfolding converse_def by auto - -lemma converse_shift: -"R1 \ R2 ^-1 \ R1 ^-1 \ R2" -unfolding converse_def by auto - -lemma converse_Times: "(A \ B) ^-1 = B \ A" -by auto - -lemma equiv_triv1: -assumes "equiv A R" and "(a, b) \ R" and "(a, c) \ R" -shows "(b, c) \ R" -using assms unfolding equiv_def sym_def trans_def by blast - -lemma equiv_triv2: -assumes "equiv A R" and "(a, b) \ R" and "(b, c) \ R" -shows "(a, c) \ R" -using assms unfolding equiv_def trans_def by blast - -lemma equiv_proj: - assumes e: "equiv A R" and "z \ R" - shows "(proj R o fst) z = (proj R o snd) z" -proof - - from assms(2) have z: "(fst z, snd z) \ R" by auto - have P: "\x. (fst z, x) \ R \ (snd z, x) \ R" by (erule equiv_triv1[OF e z]) - have "\x. (snd z, x) \ R \ (fst z, x) \ R" by (erule equiv_triv2[OF e z]) - with P show ?thesis unfolding proj_def[abs_def] by auto -qed - - -section{* Weak pullbacks: *} - -definition csquare where -"csquare A f1 f2 p1 p2 \ (\ a \ A. f1 (p1 a) = f2 (p2 a))" - -definition wpull where -"wpull A B1 B2 f1 f2 p1 p2 \ - (\ b1 b2. b1 \ B1 \ b2 \ B2 \ f1 b1 = f2 b2 \ (\ a \ A. p1 a = b1 \ p2 a = b2))" - -lemma wpull_cong: -"\A' = A; B1' = B1; B2' = B2; wpull A B1 B2 f1 f2 p1 p2\ \ wpull A' B1' B2' f1 f2 p1 p2" -by simp - -lemma wpull_id: "wpull UNIV B1 B2 id id id id" -unfolding wpull_def by simp - - -(* Weak pseudo-pullbacks *) - -definition wppull where -"wppull A B1 B2 f1 f2 e1 e2 p1 p2 \ - (\ b1 b2. b1 \ B1 \ b2 \ B2 \ f1 b1 = f2 b2 \ - (\ a \ A. e1 (p1 a) = e1 b1 \ e2 (p2 a) = e2 b2))" - - -(* The pullback of sets *) -definition thePull where -"thePull B1 B2 f1 f2 = {(b1,b2). b1 \ B1 \ b2 \ B2 \ f1 b1 = f2 b2}" - -lemma wpull_thePull: -"wpull (thePull B1 B2 f1 f2) B1 B2 f1 f2 fst snd" -unfolding wpull_def thePull_def by auto - -lemma wppull_thePull: -assumes "wppull A B1 B2 f1 f2 e1 e2 p1 p2" -shows -"\ j. \ a' \ thePull B1 B2 f1 f2. - j a' \ A \ - e1 (p1 (j a')) = e1 (fst a') \ e2 (p2 (j a')) = e2 (snd a')" -(is "\ j. \ a' \ ?A'. ?phi a' (j a')") -proof(rule bchoice[of ?A' ?phi], default) - fix a' assume a': "a' \ ?A'" - hence "fst a' \ B1" unfolding thePull_def by auto - moreover - from a' have "snd a' \ B2" unfolding thePull_def by auto - moreover have "f1 (fst a') = f2 (snd a')" - using a' unfolding csquare_def thePull_def by auto - ultimately show "\ ja'. ?phi a' ja'" - using assms unfolding wppull_def by auto -qed - -lemma wpull_wppull: -assumes wp: "wpull A' B1 B2 f1 f2 p1' p2'" and -1: "\ a' \ A'. j a' \ A \ e1 (p1 (j a')) = e1 (p1' a') \ e2 (p2 (j a')) = e2 (p2' a')" -shows "wppull A B1 B2 f1 f2 e1 e2 p1 p2" -unfolding wppull_def proof safe - fix b1 b2 - assume b1: "b1 \ B1" and b2: "b2 \ B2" and f: "f1 b1 = f2 b2" - then obtain a' where a': "a' \ A'" and b1: "b1 = p1' a'" and b2: "b2 = p2' a'" - using wp unfolding wpull_def by blast - show "\a\A. e1 (p1 a) = e1 b1 \ e2 (p2 a) = e2 b2" - apply(rule bexI[of _ "j a'"]) unfolding b1 b2 using a' 1 by auto -qed - -lemma wppull_id: "\wpull UNIV UNIV UNIV f1 f2 p1 p2; e1 = id; e2 = id\ \ - wppull UNIV UNIV UNIV f1 f2 e1 e2 p1 p2" -by (erule wpull_wppull) auto - - -(* Operators: *) -definition diag where "diag A \ {(a,a) | a. a \ A}" -definition "Gr A f = {(a,f a) | a. a \ A}" -definition image2 where "image2 A f g = {(f a, g a) | a. a \ A}" - -lemma diagI: "x \ A \ (x, x) \ diag A" -unfolding diag_def by simp - -lemma diagE: "(a, b) \ diag A \ a = b" -unfolding diag_def by simp - -lemma diagE': "x \ diag A \ fst x = snd x" -unfolding diag_def by auto - -lemma diag_fst: "x \ diag A \ fst x \ A" -unfolding diag_def by auto - -lemma diag_UNIV: "diag UNIV = Id" -unfolding diag_def by auto - -lemma diag_converse: "diag A = (diag A) ^-1" -unfolding diag_def by auto - -lemma diag_Comp: "diag A = diag A O diag A" -unfolding diag_def by auto - -lemma diag_Gr: "diag A = Gr A id" -unfolding diag_def Gr_def by simp - -lemma diag_UNIV_I: "x = y \ (x, y) \ diag UNIV" -unfolding diag_def by auto - -lemma image2_eqI: "\b = f x; c = g x; x \ A\ \ (b, c) \ image2 A f g" -unfolding image2_def by auto - -lemma Id_def': "Id = {(a,b). a = b}" -by auto - -lemma Id_alt: "Id = Gr UNIV id" -unfolding Gr_def by auto - -lemma Id_subset: "Id \ {(a, b). P a b \ a = b}" -by auto - -lemma IdD: "(a, b) \ Id \ a = b" -by auto - -lemma image2_Gr: "image2 A f g = (Gr A f)^-1 O (Gr A g)" -unfolding image2_def Gr_def by auto - -lemma GrI: "\x \ A; f x = fx\ \ (x, fx) \ Gr A f" -unfolding Gr_def by simp - -lemma GrE: "(x, fx) \ Gr A f \ (x \ A \ f x = fx \ P) \ P" -unfolding Gr_def by simp - -lemma GrD1: "(x, fx) \ Gr A f \ x \ A" -unfolding Gr_def by simp - -lemma GrD2: "(x, fx) \ Gr A f \ f x = fx" -unfolding Gr_def by simp - -lemma Gr_UNIV_id: "f = id \ (Gr UNIV f)^-1 O Gr UNIV f = Gr UNIV f" -unfolding Gr_def by auto - -lemma Gr_fst_snd: "(Gr R fst)^-1 O Gr R snd = R" -unfolding Gr_def by auto - -lemma Gr_mono: "A \ B \ Gr A f \ Gr B f" -unfolding Gr_def by auto - -lemma subst_rel_def: "A = B \ (Gr A f)^-1 O Gr A g = (Gr B f)^-1 O Gr B g" -by simp - -lemma abs_pred_def: "\\x y. (x, y) \ rel = pred x y\ \ rel = Collect (split pred)" -by auto - -lemma Collect_split_cong: "Collect (split pred) = Collect (split pred') \ pred = pred'" -by blast - -lemma pred_def_abs: "rel = Collect (split pred) \ pred = (\x y. (x, y) \ rel)" -by auto - -lemma wpull_Gr: -"wpull (Gr A f) A (f ` A) f id fst snd" -unfolding wpull_def Gr_def by auto - -lemma Gr_incl: "Gr A f \ A <*> B \ f ` A \ B" -unfolding Gr_def by auto - -lemma equiv_Image: "equiv A R \ (\a b. (a, b) \ R \ a \ A \ b \ A \ R `` {a} = R `` {b})" -unfolding equiv_def refl_on_def Image_def by (auto intro: transD symD) - -definition relImage where -"relImage R f \ {(f a1, f a2) | a1 a2. (a1,a2) \ R}" - -definition relInvImage where -"relInvImage A R f \ {(a1, a2) | a1 a2. a1 \ A \ a2 \ A \ (f a1, f a2) \ R}" - -lemma relImage_Gr: -"\R \ A \ A\ \ relImage R f = (Gr A f)^-1 O R O Gr A f" -unfolding relImage_def Gr_def relcomp_def by auto - -lemma relInvImage_Gr: "\R \ B \ B\ \ relInvImage A R f = Gr A f O R O (Gr A f)^-1" -unfolding Gr_def relcomp_def image_def relInvImage_def by auto - -lemma relImage_mono: -"R1 \ R2 \ relImage R1 f \ relImage R2 f" -unfolding relImage_def by auto - -lemma relInvImage_mono: -"R1 \ R2 \ relInvImage A R1 f \ relInvImage A R2 f" -unfolding relInvImage_def by auto - -lemma relInvImage_diag: -"(\a1 a2. f a1 = f a2 \ a1 = a2) \ relInvImage A (diag B) f \ Id" -unfolding relInvImage_def diag_def by auto - -lemma relInvImage_UNIV_relImage: -"R \ relInvImage UNIV (relImage R f) f" -unfolding relInvImage_def relImage_def by auto - -lemma relImage_proj: -assumes "equiv A R" -shows "relImage R (proj R) \ diag (A//R)" -unfolding relImage_def diag_def apply safe -using proj_iff[OF assms] -by (metis assms equiv_Image proj_def proj_preserves) - -lemma relImage_relInvImage: -assumes "R \ f ` A <*> f ` A" -shows "relImage (relInvImage A R f) f = R" -using assms unfolding relImage_def relInvImage_def by fastforce - - -(* Relation composition as a weak pseudo-pullback *) - -(* pick middle *) -definition "pickM P Q a c = (SOME b. (a,b) \ P \ (b,c) \ Q)" - -lemma pickM: -assumes "(a,c) \ P O Q" -shows "(a, pickM P Q a c) \ P \ (pickM P Q a c, c) \ Q" -unfolding pickM_def apply(rule someI_ex) -using assms unfolding relcomp_def by auto - -definition fstO where "fstO P Q ac = (fst ac, pickM P Q (fst ac) (snd ac))" -definition sndO where "sndO P Q ac = (pickM P Q (fst ac) (snd ac), snd ac)" - -lemma fstO_in: "ac \ P O Q \ fstO P Q ac \ P" -by (metis assms fstO_def pickM surjective_pairing) - -lemma fst_fstO: "fst bc = (fst \ fstO P Q) bc" -unfolding comp_def fstO_def by simp - -lemma snd_sndO: "snd bc = (snd \ sndO P Q) bc" -unfolding comp_def sndO_def by simp - -lemma sndO_in: "ac \ P O Q \ sndO P Q ac \ Q" -by (metis assms sndO_def pickM surjective_pairing) - -lemma csquare_fstO_sndO: -"csquare (P O Q) snd fst (fstO P Q) (sndO P Q)" -unfolding csquare_def fstO_def sndO_def using pickM by auto - -lemma wppull_fstO_sndO: -shows "wppull (P O Q) P Q snd fst fst snd (fstO P Q) (sndO P Q)" -using pickM unfolding wppull_def fstO_def sndO_def relcomp_def by auto - -lemma subst_Pair: "P x y \ a = (x, y) \ P (fst a) (snd a)" -by simp - -lemma snd_fst_flip: "snd xy = (fst o (%(x, y). (y, x))) xy" -by (simp split: prod.split) - -lemma fst_snd_flip: "fst xy = (snd o (%(x, y). (y, x))) xy" -by (simp split: prod.split) - -lemma flip_rel: "A \ (R ^-1) \ (%(x, y). (y, x)) ` A \ R" -by auto - -lemma fst_diag_id: "(fst \ (%x. (x, x))) z = id z" -by simp - -lemma snd_diag_id: "(snd \ (%x. (x, x))) z = id z" -by simp - -lemma fst_snd: "\snd x = (y, z)\ \ fst (snd x) = y" -by simp - -lemma snd_snd: "\snd x = (y, z)\ \ snd (snd x) = z" -by simp - -lemma fstI: "x = (y, z) \ fst x = y" -by simp - -lemma sndI: "x = (y, z) \ snd x = z" -by simp - -lemma Collect_restrict: "{x. x \ X \ P x} \ X" -by auto - -lemma Collect_restrict': "{(x, y) | x y. phi x y \ P x y} \ {(x, y) | x y. phi x y}" -by auto - -lemma prop_restrict: "\x \ Z; Z \ {x. x \ X \ P x}\ \ P x" -by auto - -lemma underS_I: "\i \ j; (i, j) \ R\ \ i \ rel.underS R j" -unfolding rel.underS_def by simp - -lemma underS_E: "i \ rel.underS R j \ i \ j \ (i, j) \ R" -unfolding rel.underS_def by simp - -lemma underS_Field: "i \ rel.underS R j \ i \ Field R" -unfolding rel.underS_def Field_def by auto - -lemma FieldI2: "(i, j) \ R \ j \ Field R" -unfolding Field_def by auto - - -subsection {* Convolution product *} - -definition convol ("<_ , _>") where -" \ %a. (f a, g a)" - -lemma fst_convol: -"fst o = f" -apply(rule ext) -unfolding convol_def by simp - -lemma snd_convol: -"snd o = g" -apply(rule ext) -unfolding convol_def by simp - -lemma fst_convol': "fst ( x) = f x" -using fst_convol unfolding convol_def by simp - -lemma snd_convol': "snd ( x) = g x" -using snd_convol unfolding convol_def by simp - -lemma image_convolD: "\(a, b) \ ` X\ \ \x. x \ X \ a = f x \ b = g x" -unfolding convol_def by auto - -lemma convol_expand_snd: "fst o f = g \ = f" -unfolding convol_def by auto - -subsection{* Facts about functions *} - -lemma pointfreeE: "f o g = f' o g' \ f (g x) = f' (g' x)" -unfolding o_def fun_eq_iff by simp - -lemma pointfree_idE: "f o g = id \ f (g x) = x" -unfolding o_def fun_eq_iff by simp - -definition inver where - "inver g f A = (ALL a : A. g (f a) = a)" - -lemma bij_betw_iff_ex: - "bij_betw f A B = (EX g. g ` B = A \ inver g f A \ inver f g B)" (is "?L = ?R") -proof (rule iffI) - assume ?L - hence f: "f ` A = B" and inj_f: "inj_on f A" unfolding bij_betw_def by auto - let ?phi = "% b a. a : A \ f a = b" - have "ALL b : B. EX a. ?phi b a" using f by blast - then obtain g where g: "ALL b : B. g b : A \ f (g b) = b" - using bchoice[of B ?phi] by blast - hence gg: "ALL b : f ` A. g b : A \ f (g b) = b" using f by blast - have gf: "inver g f A" unfolding inver_def by (metis gg imageI inj_f the_inv_into_f_f) - moreover have "g ` B \ A \ inver f g B" using g unfolding inver_def by blast - moreover have "A \ g ` B" - proof safe - fix a assume a: "a : A" - hence "f a : B" using f by auto - moreover have "a = g (f a)" using a gf unfolding inver_def by auto - ultimately show "a : g ` B" by blast - qed - ultimately show ?R by blast -next - assume ?R - then obtain g where g: "g ` B = A \ inver g f A \ inver f g B" by blast - show ?L unfolding bij_betw_def - proof safe - show "inj_on f A" unfolding inj_on_def - proof safe - fix a1 a2 assume a: "a1 : A" "a2 : A" and "f a1 = f a2" - hence "g (f a1) = g (f a2)" by simp - thus "a1 = a2" using a g unfolding inver_def by simp - qed - next - fix a assume "a : A" - then obtain b where b: "b : B" and a: "a = g b" using g by blast - hence "b = f (g b)" using g unfolding inver_def by auto - thus "f a : B" unfolding a using b by simp - next - fix b assume "b : B" - hence "g b : A \ b = f (g b)" using g unfolding inver_def by auto - thus "b : f ` A" by auto - qed -qed - -lemma bijI: "\\x y. (f x = f y) = (x = y); \y. \x. y = f x\ \ bij f" -unfolding bij_def inj_on_def by auto blast - -lemma bij_betw_ex_weakE: - "\bij_betw f A B\ \ \g. g ` B \ A \ inver g f A \ inver f g B" -by (auto simp only: bij_betw_iff_ex) - -lemma inver_surj: "\g ` B \ A; f ` A \ B; inver g f A\ \ g ` B = A" -unfolding inver_def by auto (rule rev_image_eqI, auto) - -lemma inver_mono: "\A \ B; inver f g B\ \ inver f g A" -unfolding inver_def by auto - -lemma inver_pointfree: "inver f g A = (\a \ A. (f o g) a = a)" -unfolding inver_def by simp - -lemma bij_betwE: "bij_betw f A B \ \a\A. f a \ B" -unfolding bij_betw_def by auto - -lemma bij_betw_imageE: "bij_betw f A B \ f ` A = B" -unfolding bij_betw_def by auto - -lemma inverE: "\inver f f' A; x \ A\ \ f (f' x) = x" -unfolding inver_def by auto - -lemma bij_betw_inver1: "bij_betw f A B \ inver (inv_into A f) f A" -unfolding bij_betw_def inver_def by auto - -lemma bij_betw_inver2: "bij_betw f A B \ inver f (inv_into A f) B" -unfolding bij_betw_def inver_def by auto - -lemma bij_betwI: "\bij_betw g B A; inver g f A; inver f g B\ \ bij_betw f A B" -by (metis bij_betw_iff_ex bij_betw_imageE) - -lemma bij_betwI': - "\\x y. \x \ X; y \ X\ \ (f x = f y) = (x = y); - \x. x \ X \ f x \ Y; - \y. y \ Y \ \x \ X. y = f x\ \ bij_betw f X Y" -unfolding bij_betw_def inj_on_def by auto (metis rev_image_eqI) - -lemma o_bij: - assumes gf: "g o f = id" and fg: "f o g = id" - shows "bij f" -unfolding bij_def inj_on_def surj_def proof safe - fix a1 a2 assume "f a1 = f a2" - hence "g ( f a1) = g (f a2)" by simp - thus "a1 = a2" using gf unfolding fun_eq_iff by simp -next - fix b - have "b = f (g b)" - using fg unfolding fun_eq_iff by simp - thus "EX a. b = f a" by blast -qed - -lemma surj_fun_eq: - assumes surj_on: "f ` X = UNIV" and eq_on: "\x \ X. (g1 o f) x = (g2 o f) x" - shows "g1 = g2" -proof (rule ext) - fix y - from surj_on obtain x where "x \ X" and "y = f x" by blast - thus "g1 y = g2 y" using eq_on by simp -qed - -lemma Card_order_wo_rel: "Card_order r \ wo_rel r" -unfolding wo_rel_def card_order_on_def by blast - -lemma Cinfinite_limit: "\x \ Field r; Cinfinite r\ \ - \y \ Field r. x \ y \ (x, y) \ r" -unfolding cinfinite_def by (auto simp add: infinite_Card_order_limit) - -lemma Card_order_trans: - "\Card_order r; x \ y; (x, y) \ r; y \ z; (y, z) \ r\ \ x \ z \ (x, z) \ r" -unfolding card_order_on_def well_order_on_def linear_order_on_def - partial_order_on_def preorder_on_def trans_def antisym_def by blast - -lemma Cinfinite_limit2: - assumes x1: "x1 \ Field r" and x2: "x2 \ Field r" and r: "Cinfinite r" - shows "\y \ Field r. (x1 \ y \ (x1, y) \ r) \ (x2 \ y \ (x2, y) \ r)" -proof - - from r have trans: "trans r" and total: "Total r" and antisym: "antisym r" - unfolding card_order_on_def well_order_on_def linear_order_on_def - partial_order_on_def preorder_on_def by auto - obtain y1 where y1: "y1 \ Field r" "x1 \ y1" "(x1, y1) \ r" - using Cinfinite_limit[OF x1 r] by blast - obtain y2 where y2: "y2 \ Field r" "x2 \ y2" "(x2, y2) \ r" - using Cinfinite_limit[OF x2 r] by blast - show ?thesis - proof (cases "y1 = y2") - case True with y1 y2 show ?thesis by blast - next - case False - with y1(1) y2(1) total have "(y1, y2) \ r \ (y2, y1) \ r" - unfolding total_on_def by auto - thus ?thesis - proof - assume *: "(y1, y2) \ r" - with trans y1(3) have "(x1, y2) \ r" unfolding trans_def by blast - with False y1 y2 * antisym show ?thesis by (cases "x1 = y2") (auto simp: antisym_def) - next - assume *: "(y2, y1) \ r" - with trans y2(3) have "(x2, y1) \ r" unfolding trans_def by blast - with False y1 y2 * antisym show ?thesis by (cases "x2 = y1") (auto simp: antisym_def) - qed - qed -qed - -lemma Cinfinite_limit_finite: "\finite X; X \ Field r; Cinfinite r\ - \ \y \ Field r. \x \ X. (x \ y \ (x, y) \ r)" -proof (induct X rule: finite_induct) - case empty thus ?case unfolding cinfinite_def using ex_in_conv[of "Field r"] finite.emptyI by auto -next - case (insert x X) - then obtain y where y: "y \ Field r" "\x \ X. (x \ y \ (x, y) \ r)" by blast - then obtain z where z: "z \ Field r" "x \ z \ (x, z) \ r" "y \ z \ (y, z) \ r" - using Cinfinite_limit2[OF _ y(1) insert(5), of x] insert(4) by blast - show ?case by (metis Card_order_trans insert(5) insertE y(2) z) -qed - -lemma insert_subsetI: "\x \ A; X \ A\ \ insert x X \ A" -by auto - -(*helps resolution*) -lemma well_order_induct_imp: - "wo_rel r \ (\x. \y. y \ x \ (y, x) \ r \ y \ Field r \ P y \ x \ Field r \ P x) \ - x \ Field r \ P x" -by (erule wo_rel.well_order_induct) - -lemma meta_spec2: - assumes "(\x y. PROP P x y)" - shows "PROP P x y" -by (rule `(\x y. PROP P x y)`) - -(*Extended Sublist*) - -definition prefCl where - "prefCl Kl = (\ kl1 kl2. kl1 \ kl2 \ kl2 \ Kl \ kl1 \ Kl)" -definition PrefCl where - "PrefCl A n = (\kl kl'. kl \ A n \ kl' \ kl \ (\m\n. kl' \ A m))" - -lemma prefCl_UN: - "\\n. PrefCl A n\ \ prefCl (\n. A n)" -unfolding prefCl_def PrefCl_def by fastforce - -definition Succ where "Succ Kl kl = {k . kl @ [k] \ Kl}" -definition Shift where "Shift Kl k = {kl. k # kl \ Kl}" -definition shift where "shift lab k = (\kl. lab (k # kl))" - -lemmas sh_def = Shift_def shift_def - -lemma empty_Shift: "\[] \ Kl; k \ Succ Kl []\ \ [] \ Shift Kl k" -unfolding Shift_def Succ_def by simp - -lemma Shift_clists: "Kl \ Field (clists r) \ Shift Kl k \ Field (clists r)" -unfolding Shift_def clists_def Field_card_of by auto - -lemma Shift_prefCl: "prefCl Kl \ prefCl (Shift Kl k)" -unfolding prefCl_def Shift_def -proof safe - fix kl1 kl2 - assume "\kl1 kl2. kl1 \ kl2 \ kl2 \ Kl \ kl1 \ Kl" - "kl1 \ kl2" "k # kl2 \ Kl" - thus "k # kl1 \ Kl" using Cons_prefix_Cons[of k kl1 k kl2] by blast -qed - -lemma not_in_Shift: "kl \ Shift Kl x \ x # kl \ Kl" -unfolding Shift_def by simp - -lemma prefCl_Succ: "\prefCl Kl; k # kl \ Kl\ \ k \ Succ Kl []" -unfolding Succ_def proof - assume "prefCl Kl" "k # kl \ Kl" - moreover have "k # [] \ k # kl" by auto - ultimately have "k # [] \ Kl" unfolding prefCl_def by blast - thus "[] @ [k] \ Kl" by simp -qed - -lemma SuccD: "k \ Succ Kl kl \ kl @ [k] \ Kl" -unfolding Succ_def by simp - -lemmas SuccE = SuccD[elim_format] - -lemma SuccI: "kl @ [k] \ Kl \ k \ Succ Kl kl" -unfolding Succ_def by simp - -lemma ShiftD: "kl \ Shift Kl k \ k # kl \ Kl" -unfolding Shift_def by simp - -lemma Succ_Shift: "Succ (Shift Kl k) kl = Succ Kl (k # kl)" -unfolding Succ_def Shift_def by auto - -lemma ShiftI: "k # kl \ Kl \ kl \ Shift Kl k" -unfolding Shift_def by simp - -lemma Func_cexp: "|Func A B| =o |B| ^c |A|" -unfolding cexp_def Field_card_of by (simp only: card_of_refl) - -lemma clists_bound: "A \ Field (cpow (clists r)) - {{}} \ |A| \o clists r" -unfolding cpow_def clists_def Field_card_of by (auto simp: card_of_mono1) - -lemma cpow_clists_czero: "\A \ Field (cpow (clists r)) - {{}}; |A| =o czero\ \ False" -unfolding cpow_def clists_def -by (auto simp add: card_of_ordIso_czero_iff_empty[symmetric]) - (erule notE, erule ordIso_transitive, rule czero_ordIso) - -lemma incl_UNION_I: -assumes "i \ I" and "A \ F i" -shows "A \ UNION I F" -using assms by auto - -lemma Nil_clists: "{[]} \ Field (clists r)" -unfolding clists_def Field_card_of by auto - -lemma Cons_clists: - "\x \ Field r; xs \ Field (clists r)\ \ x # xs \ Field (clists r)" -unfolding clists_def Field_card_of by auto - -lemma length_Cons: "length (x#xs) = Suc (length xs)" -by simp - -lemma length_append_singleton: "length (xs @ [x]) = Suc (length xs)" -by simp - -(*injection into the field of a cardinal*) -definition "toCard_pred A r f \ inj_on f A \ f ` A \ Field r \ Card_order r" -definition "toCard A r \ SOME f. toCard_pred A r f" - -lemma ex_toCard_pred: -"\|A| \o r; Card_order r\ \ \ f. toCard_pred A r f" -unfolding toCard_pred_def -using card_of_ordLeq[of A "Field r"] - ordLeq_ordIso_trans[OF _ card_of_unique[of "Field r" r], of "|A|"] -by blast - -lemma toCard_pred_toCard: - "\|A| \o r; Card_order r\ \ toCard_pred A r (toCard A r)" -unfolding toCard_def using someI_ex[OF ex_toCard_pred] . - -lemma toCard_inj: "\|A| \o r; Card_order r; x \ A; y \ A\ \ - toCard A r x = toCard A r y \ x = y" -using toCard_pred_toCard unfolding inj_on_def toCard_pred_def by blast - -lemma toCard: "\|A| \o r; Card_order r; b \ A\ \ toCard A r b \ Field r" -using toCard_pred_toCard unfolding toCard_pred_def by blast - -definition "fromCard A r k \ SOME b. b \ A \ toCard A r b = k" - -lemma fromCard_toCard: -"\|A| \o r; Card_order r; b \ A\ \ fromCard A r (toCard A r b) = b" -unfolding fromCard_def by (rule some_equality) (auto simp add: toCard_inj) - -(* pick according to the weak pullback *) -definition pickWP_pred where -"pickWP_pred A p1 p2 b1 b2 a \ - a \ A \ p1 a = b1 \ p2 a = b2" - -definition pickWP where -"pickWP A p1 p2 b1 b2 \ - SOME a. pickWP_pred A p1 p2 b1 b2 a" - -lemma pickWP_pred: -assumes "wpull A B1 B2 f1 f2 p1 p2" and -"b1 \ B1" and "b2 \ B2" and "f1 b1 = f2 b2" -shows "\ a. pickWP_pred A p1 p2 b1 b2 a" -using assms unfolding wpull_def pickWP_pred_def by blast - -lemma pickWP_pred_pickWP: -assumes "wpull A B1 B2 f1 f2 p1 p2" and -"b1 \ B1" and "b2 \ B2" and "f1 b1 = f2 b2" -shows "pickWP_pred A p1 p2 b1 b2 (pickWP A p1 p2 b1 b2)" -unfolding pickWP_def using assms by(rule someI_ex[OF pickWP_pred]) - -lemma pickWP: -assumes "wpull A B1 B2 f1 f2 p1 p2" and -"b1 \ B1" and "b2 \ B2" and "f1 b1 = f2 b2" -shows "pickWP A p1 p2 b1 b2 \ A" - "p1 (pickWP A p1 p2 b1 b2) = b1" - "p2 (pickWP A p1 p2 b1 b2) = b2" -using assms pickWP_pred_pickWP unfolding pickWP_pred_def by fastforce+ - -lemma ssubst_mem: "\t = s; s \ X\ \ t \ X" by simp - -lemma Inl_Field_csum: "a \ Field r \ Inl a \ Field (r +c s)" -unfolding Field_card_of csum_def by auto - -lemma Inr_Field_csum: "a \ Field s \ Inr a \ Field (r +c s)" -unfolding Field_card_of csum_def by auto - -lemma nat_rec_0: "f = nat_rec f1 (%n rec. f2 n rec) \ f 0 = f1" -by auto - -lemma nat_rec_Suc: "f = nat_rec f1 (%n rec. f2 n rec) \ f (Suc n) = f2 n (f n)" -by auto - -lemma list_rec_Nil: "f = list_rec f1 (%x xs rec. f2 x xs rec) \ f [] = f1" -by auto - -lemma list_rec_Cons: "f = list_rec f1 (%x xs rec. f2 x xs rec) \ f (x # xs) = f2 x xs (f xs)" -by auto - -lemma sum_case_cong: "p = q \ sum_case f g p = sum_case f g q" -by simp - -lemma sum_case_step: - "sum_case (sum_case f' g') g (Inl p) = sum_case f' g' p" - "sum_case f (sum_case f' g') (Inr p) = sum_case f' g' p" -by auto - -lemma one_pointE: "\\x. s = x \ P\ \ P" -by simp - -lemma obj_one_pointE: "\x. s = x \ P \ P" -by blast - -lemma obj_sumE_f: -"\\x. s = f (Inl x) \ P; \x. s = f (Inr x) \ P\ \ \x. s = f x \ P" -by (metis sum.exhaust) - -lemma obj_sumE: "\\x. s = Inl x \ P; \x. s = Inr x \ P\ \ P" -by (cases s) auto - -lemma obj_sum_step: - "\\x. s = f (Inr (Inl x)) \ P; \x. s = f (Inr (Inr x)) \ P\ \ \x. s = f (Inr x) \ P" -by (metis obj_sumE) - -lemma sum_case_if: -"sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)" -by simp - -lemma not_arg_cong_Inr: "x \ y \ Inr x \ Inr y" -by simp - -end diff -r 3d87f4fd0d50 -r c057e1b39f16 src/HOL/Codatatype/BNF_Util.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Codatatype/BNF_Util.thy Tue Sep 11 17:09:39 2012 +0200 @@ -0,0 +1,854 @@ +(* Title: HOL/Codatatype/BNF_Util.thy + Author: Dmitriy Traytel, TU Muenchen + Author: Jasmin Blanchette, TU Muenchen + Copyright 2012 + +Library for bounded natural functors. +*) + +header {* Library for Bounded Natural Functors *} + +theory BNF_Util +imports + "../Ordinals_and_Cardinals/Cardinal_Arithmetic" + "~~/src/HOL/Library/Prefix_Order" + Equiv_Relations_More +begin + +lemma iffI_np: "\x \ \ y; \ x \ y\ \ \ x \ y" +by (erule iffI) (erule contrapos_pn) + +lemma all_unit_eq: "(\x. PROP P x) \ PROP P ()" by simp + +lemma all_prod_eq: "(\x. PROP P x) \ (\a b. PROP P (a, b))" by clarsimp + +lemma False_imp_eq: "(False \ P) \ Trueprop True" +by presburger + +lemma case_unit: "(case u of () => f) = f" +by (cases u) (hypsubst, rule unit.cases) + +lemma All_point_1: "(\z. z = b \ phi z) \ Trueprop (phi b)" +by presburger + +lemma subset_Collect_iff: "B \ A \ (B \ {x \ A. P x}) = (\x \ B. P x)" +by blast + +lemma subset_CollectI: "B \ A \ (\x. x \ B \ Q x \ P x) \ ({x \ B. Q x} \ {x \ A. P x})" +by blast + +lemma mem_Collect_eq_split: "{(x, y). (x, y) \ X} = X" +by simp + +lemma image_comp: "image (f o g) = image f o image g" +by (rule ext) (auto simp only: o_apply image_def) + +lemma empty_natural: "(\_. {}) o f = image g o (\_. {})" +by (rule ext) simp + +lemma Union_natural: "Union o image (image f) = image f o Union" +by (rule ext) (auto simp only: o_apply) + +lemma in_Union_o_assoc: "x \ (Union o gset o gmap) A \ x \ (Union o (gset o gmap)) A" +by (unfold o_assoc) + +lemma comp_single_set_bd: + assumes fbd_Card_order: "Card_order fbd" and + fset_bd: "\x. |fset x| \o fbd" and + gset_bd: "\x. |gset x| \o gbd" + shows "|\fset ` gset x| \o gbd *c fbd" +apply (subst sym[OF SUP_def]) +apply (rule ordLeq_transitive) +apply (rule card_of_UNION_Sigma) +apply (subst SIGMA_CSUM) +apply (rule ordLeq_transitive) +apply (rule card_of_Csum_Times') +apply (rule fbd_Card_order) +apply (rule ballI) +apply (rule fset_bd) +apply (rule ordLeq_transitive) +apply (rule cprod_mono1) +apply (rule gset_bd) +apply (rule ordIso_imp_ordLeq) +apply (rule ordIso_refl) +apply (rule Card_order_cprod) +done + +lemma Union_image_insert: "\f ` insert a B = f a \ \f ` B" +by simp + +lemma Union_image_empty: "A \ \f ` {} = A" +by simp + +definition collect where + "collect F x = (\f \ F. f x)" + +lemma collect_o: "collect F o g = collect ((\f. f o g) ` F)" +by (rule ext) (auto simp only: o_apply collect_def) + +lemma image_o_collect: "collect ((\f. image g o f) ` F) = image g o collect F" +by (rule ext) (auto simp add: collect_def) + +lemma conj_subset_def: "A \ {x. P x \ Q x} = (A \ {x. P x} \ A \ {x. Q x})" +by blast + +lemma subset_emptyI: "(\x. x \ A \ False) \ A \ {}" +by blast + +lemma rev_bspec: "a \ A \ \z \ A. P z \ P a" +by simp + +lemma Un_cong: "\A = B; C = D\ \ A \ C = B \ D" +by simp + +lemma UN_image_subset: "\f ` g x \ X = (g x \ {x. f x \ X})" +by blast + +lemma image_Collect_subsetI: + "(\x. P x \ f x \ B) \ f ` {x. P x} \ B" +by blast + +lemma comp_set_bd_Union_o_collect: "|\\(\f. f x) ` X| \o hbd \ |(Union \ collect X) x| \o hbd" +by (unfold o_apply collect_def SUP_def) + +lemma sum_case_comp_Inl: +"sum_case f g \ Inl = f" +unfolding comp_def by simp + +lemma sum_case_expand_Inr: "f o Inl = g \ f x = sum_case g (f o Inr) x" +by (auto split: sum.splits) + +lemma converse_mono: +"R1 ^-1 \ R2 ^-1 \ R1 \ R2" +unfolding converse_def by auto + +lemma converse_shift: +"R1 \ R2 ^-1 \ R1 ^-1 \ R2" +unfolding converse_def by auto + +lemma converse_Times: "(A \ B) ^-1 = B \ A" +by auto + +lemma equiv_triv1: +assumes "equiv A R" and "(a, b) \ R" and "(a, c) \ R" +shows "(b, c) \ R" +using assms unfolding equiv_def sym_def trans_def by blast + +lemma equiv_triv2: +assumes "equiv A R" and "(a, b) \ R" and "(b, c) \ R" +shows "(a, c) \ R" +using assms unfolding equiv_def trans_def by blast + +lemma equiv_proj: + assumes e: "equiv A R" and "z \ R" + shows "(proj R o fst) z = (proj R o snd) z" +proof - + from assms(2) have z: "(fst z, snd z) \ R" by auto + have P: "\x. (fst z, x) \ R \ (snd z, x) \ R" by (erule equiv_triv1[OF e z]) + have "\x. (snd z, x) \ R \ (fst z, x) \ R" by (erule equiv_triv2[OF e z]) + with P show ?thesis unfolding proj_def[abs_def] by auto +qed + + +section{* Weak pullbacks: *} + +definition csquare where +"csquare A f1 f2 p1 p2 \ (\ a \ A. f1 (p1 a) = f2 (p2 a))" + +definition wpull where +"wpull A B1 B2 f1 f2 p1 p2 \ + (\ b1 b2. b1 \ B1 \ b2 \ B2 \ f1 b1 = f2 b2 \ (\ a \ A. p1 a = b1 \ p2 a = b2))" + +lemma wpull_cong: +"\A' = A; B1' = B1; B2' = B2; wpull A B1 B2 f1 f2 p1 p2\ \ wpull A' B1' B2' f1 f2 p1 p2" +by simp + +lemma wpull_id: "wpull UNIV B1 B2 id id id id" +unfolding wpull_def by simp + + +(* Weak pseudo-pullbacks *) + +definition wppull where +"wppull A B1 B2 f1 f2 e1 e2 p1 p2 \ + (\ b1 b2. b1 \ B1 \ b2 \ B2 \ f1 b1 = f2 b2 \ + (\ a \ A. e1 (p1 a) = e1 b1 \ e2 (p2 a) = e2 b2))" + + +(* The pullback of sets *) +definition thePull where +"thePull B1 B2 f1 f2 = {(b1,b2). b1 \ B1 \ b2 \ B2 \ f1 b1 = f2 b2}" + +lemma wpull_thePull: +"wpull (thePull B1 B2 f1 f2) B1 B2 f1 f2 fst snd" +unfolding wpull_def thePull_def by auto + +lemma wppull_thePull: +assumes "wppull A B1 B2 f1 f2 e1 e2 p1 p2" +shows +"\ j. \ a' \ thePull B1 B2 f1 f2. + j a' \ A \ + e1 (p1 (j a')) = e1 (fst a') \ e2 (p2 (j a')) = e2 (snd a')" +(is "\ j. \ a' \ ?A'. ?phi a' (j a')") +proof(rule bchoice[of ?A' ?phi], default) + fix a' assume a': "a' \ ?A'" + hence "fst a' \ B1" unfolding thePull_def by auto + moreover + from a' have "snd a' \ B2" unfolding thePull_def by auto + moreover have "f1 (fst a') = f2 (snd a')" + using a' unfolding csquare_def thePull_def by auto + ultimately show "\ ja'. ?phi a' ja'" + using assms unfolding wppull_def by auto +qed + +lemma wpull_wppull: +assumes wp: "wpull A' B1 B2 f1 f2 p1' p2'" and +1: "\ a' \ A'. j a' \ A \ e1 (p1 (j a')) = e1 (p1' a') \ e2 (p2 (j a')) = e2 (p2' a')" +shows "wppull A B1 B2 f1 f2 e1 e2 p1 p2" +unfolding wppull_def proof safe + fix b1 b2 + assume b1: "b1 \ B1" and b2: "b2 \ B2" and f: "f1 b1 = f2 b2" + then obtain a' where a': "a' \ A'" and b1: "b1 = p1' a'" and b2: "b2 = p2' a'" + using wp unfolding wpull_def by blast + show "\a\A. e1 (p1 a) = e1 b1 \ e2 (p2 a) = e2 b2" + apply(rule bexI[of _ "j a'"]) unfolding b1 b2 using a' 1 by auto +qed + +lemma wppull_id: "\wpull UNIV UNIV UNIV f1 f2 p1 p2; e1 = id; e2 = id\ \ + wppull UNIV UNIV UNIV f1 f2 e1 e2 p1 p2" +by (erule wpull_wppull) auto + + +(* Operators: *) +definition diag where "diag A \ {(a,a) | a. a \ A}" +definition "Gr A f = {(a,f a) | a. a \ A}" +definition image2 where "image2 A f g = {(f a, g a) | a. a \ A}" + +lemma diagI: "x \ A \ (x, x) \ diag A" +unfolding diag_def by simp + +lemma diagE: "(a, b) \ diag A \ a = b" +unfolding diag_def by simp + +lemma diagE': "x \ diag A \ fst x = snd x" +unfolding diag_def by auto + +lemma diag_fst: "x \ diag A \ fst x \ A" +unfolding diag_def by auto + +lemma diag_UNIV: "diag UNIV = Id" +unfolding diag_def by auto + +lemma diag_converse: "diag A = (diag A) ^-1" +unfolding diag_def by auto + +lemma diag_Comp: "diag A = diag A O diag A" +unfolding diag_def by auto + +lemma diag_Gr: "diag A = Gr A id" +unfolding diag_def Gr_def by simp + +lemma diag_UNIV_I: "x = y \ (x, y) \ diag UNIV" +unfolding diag_def by auto + +lemma image2_eqI: "\b = f x; c = g x; x \ A\ \ (b, c) \ image2 A f g" +unfolding image2_def by auto + +lemma Id_def': "Id = {(a,b). a = b}" +by auto + +lemma Id_alt: "Id = Gr UNIV id" +unfolding Gr_def by auto + +lemma Id_subset: "Id \ {(a, b). P a b \ a = b}" +by auto + +lemma IdD: "(a, b) \ Id \ a = b" +by auto + +lemma image2_Gr: "image2 A f g = (Gr A f)^-1 O (Gr A g)" +unfolding image2_def Gr_def by auto + +lemma GrI: "\x \ A; f x = fx\ \ (x, fx) \ Gr A f" +unfolding Gr_def by simp + +lemma GrE: "(x, fx) \ Gr A f \ (x \ A \ f x = fx \ P) \ P" +unfolding Gr_def by simp + +lemma GrD1: "(x, fx) \ Gr A f \ x \ A" +unfolding Gr_def by simp + +lemma GrD2: "(x, fx) \ Gr A f \ f x = fx" +unfolding Gr_def by simp + +lemma Gr_UNIV_id: "f = id \ (Gr UNIV f)^-1 O Gr UNIV f = Gr UNIV f" +unfolding Gr_def by auto + +lemma Gr_fst_snd: "(Gr R fst)^-1 O Gr R snd = R" +unfolding Gr_def by auto + +lemma Gr_mono: "A \ B \ Gr A f \ Gr B f" +unfolding Gr_def by auto + +lemma subst_rel_def: "A = B \ (Gr A f)^-1 O Gr A g = (Gr B f)^-1 O Gr B g" +by simp + +lemma abs_pred_def: "\\x y. (x, y) \ rel = pred x y\ \ rel = Collect (split pred)" +by auto + +lemma Collect_split_cong: "Collect (split pred) = Collect (split pred') \ pred = pred'" +by blast + +lemma pred_def_abs: "rel = Collect (split pred) \ pred = (\x y. (x, y) \ rel)" +by auto + +lemma wpull_Gr: +"wpull (Gr A f) A (f ` A) f id fst snd" +unfolding wpull_def Gr_def by auto + +lemma Gr_incl: "Gr A f \ A <*> B \ f ` A \ B" +unfolding Gr_def by auto + +lemma equiv_Image: "equiv A R \ (\a b. (a, b) \ R \ a \ A \ b \ A \ R `` {a} = R `` {b})" +unfolding equiv_def refl_on_def Image_def by (auto intro: transD symD) + +definition relImage where +"relImage R f \ {(f a1, f a2) | a1 a2. (a1,a2) \ R}" + +definition relInvImage where +"relInvImage A R f \ {(a1, a2) | a1 a2. a1 \ A \ a2 \ A \ (f a1, f a2) \ R}" + +lemma relImage_Gr: +"\R \ A \ A\ \ relImage R f = (Gr A f)^-1 O R O Gr A f" +unfolding relImage_def Gr_def relcomp_def by auto + +lemma relInvImage_Gr: "\R \ B \ B\ \ relInvImage A R f = Gr A f O R O (Gr A f)^-1" +unfolding Gr_def relcomp_def image_def relInvImage_def by auto + +lemma relImage_mono: +"R1 \ R2 \ relImage R1 f \ relImage R2 f" +unfolding relImage_def by auto + +lemma relInvImage_mono: +"R1 \ R2 \ relInvImage A R1 f \ relInvImage A R2 f" +unfolding relInvImage_def by auto + +lemma relInvImage_diag: +"(\a1 a2. f a1 = f a2 \ a1 = a2) \ relInvImage A (diag B) f \ Id" +unfolding relInvImage_def diag_def by auto + +lemma relInvImage_UNIV_relImage: +"R \ relInvImage UNIV (relImage R f) f" +unfolding relInvImage_def relImage_def by auto + +lemma relImage_proj: +assumes "equiv A R" +shows "relImage R (proj R) \ diag (A//R)" +unfolding relImage_def diag_def apply safe +using proj_iff[OF assms] +by (metis assms equiv_Image proj_def proj_preserves) + +lemma relImage_relInvImage: +assumes "R \ f ` A <*> f ` A" +shows "relImage (relInvImage A R f) f = R" +using assms unfolding relImage_def relInvImage_def by fastforce + + +(* Relation composition as a weak pseudo-pullback *) + +(* pick middle *) +definition "pickM P Q a c = (SOME b. (a,b) \ P \ (b,c) \ Q)" + +lemma pickM: +assumes "(a,c) \ P O Q" +shows "(a, pickM P Q a c) \ P \ (pickM P Q a c, c) \ Q" +unfolding pickM_def apply(rule someI_ex) +using assms unfolding relcomp_def by auto + +definition fstO where "fstO P Q ac = (fst ac, pickM P Q (fst ac) (snd ac))" +definition sndO where "sndO P Q ac = (pickM P Q (fst ac) (snd ac), snd ac)" + +lemma fstO_in: "ac \ P O Q \ fstO P Q ac \ P" +by (metis assms fstO_def pickM surjective_pairing) + +lemma fst_fstO: "fst bc = (fst \ fstO P Q) bc" +unfolding comp_def fstO_def by simp + +lemma snd_sndO: "snd bc = (snd \ sndO P Q) bc" +unfolding comp_def sndO_def by simp + +lemma sndO_in: "ac \ P O Q \ sndO P Q ac \ Q" +by (metis assms sndO_def pickM surjective_pairing) + +lemma csquare_fstO_sndO: +"csquare (P O Q) snd fst (fstO P Q) (sndO P Q)" +unfolding csquare_def fstO_def sndO_def using pickM by auto + +lemma wppull_fstO_sndO: +shows "wppull (P O Q) P Q snd fst fst snd (fstO P Q) (sndO P Q)" +using pickM unfolding wppull_def fstO_def sndO_def relcomp_def by auto + +lemma subst_Pair: "P x y \ a = (x, y) \ P (fst a) (snd a)" +by simp + +lemma snd_fst_flip: "snd xy = (fst o (%(x, y). (y, x))) xy" +by (simp split: prod.split) + +lemma fst_snd_flip: "fst xy = (snd o (%(x, y). (y, x))) xy" +by (simp split: prod.split) + +lemma flip_rel: "A \ (R ^-1) \ (%(x, y). (y, x)) ` A \ R" +by auto + +lemma fst_diag_id: "(fst \ (%x. (x, x))) z = id z" +by simp + +lemma snd_diag_id: "(snd \ (%x. (x, x))) z = id z" +by simp + +lemma fst_snd: "\snd x = (y, z)\ \ fst (snd x) = y" +by simp + +lemma snd_snd: "\snd x = (y, z)\ \ snd (snd x) = z" +by simp + +lemma fstI: "x = (y, z) \ fst x = y" +by simp + +lemma sndI: "x = (y, z) \ snd x = z" +by simp + +lemma Collect_restrict: "{x. x \ X \ P x} \ X" +by auto + +lemma Collect_restrict': "{(x, y) | x y. phi x y \ P x y} \ {(x, y) | x y. phi x y}" +by auto + +lemma prop_restrict: "\x \ Z; Z \ {x. x \ X \ P x}\ \ P x" +by auto + +lemma underS_I: "\i \ j; (i, j) \ R\ \ i \ rel.underS R j" +unfolding rel.underS_def by simp + +lemma underS_E: "i \ rel.underS R j \ i \ j \ (i, j) \ R" +unfolding rel.underS_def by simp + +lemma underS_Field: "i \ rel.underS R j \ i \ Field R" +unfolding rel.underS_def Field_def by auto + +lemma FieldI2: "(i, j) \ R \ j \ Field R" +unfolding Field_def by auto + + +subsection {* Convolution product *} + +definition convol ("<_ , _>") where +" \ %a. (f a, g a)" + +lemma fst_convol: +"fst o = f" +apply(rule ext) +unfolding convol_def by simp + +lemma snd_convol: +"snd o = g" +apply(rule ext) +unfolding convol_def by simp + +lemma fst_convol': "fst ( x) = f x" +using fst_convol unfolding convol_def by simp + +lemma snd_convol': "snd ( x) = g x" +using snd_convol unfolding convol_def by simp + +lemma image_convolD: "\(a, b) \ ` X\ \ \x. x \ X \ a = f x \ b = g x" +unfolding convol_def by auto + +lemma convol_expand_snd: "fst o f = g \ = f" +unfolding convol_def by auto + +subsection{* Facts about functions *} + +lemma pointfreeE: "f o g = f' o g' \ f (g x) = f' (g' x)" +unfolding o_def fun_eq_iff by simp + +lemma pointfree_idE: "f o g = id \ f (g x) = x" +unfolding o_def fun_eq_iff by simp + +definition inver where + "inver g f A = (ALL a : A. g (f a) = a)" + +lemma bij_betw_iff_ex: + "bij_betw f A B = (EX g. g ` B = A \ inver g f A \ inver f g B)" (is "?L = ?R") +proof (rule iffI) + assume ?L + hence f: "f ` A = B" and inj_f: "inj_on f A" unfolding bij_betw_def by auto + let ?phi = "% b a. a : A \ f a = b" + have "ALL b : B. EX a. ?phi b a" using f by blast + then obtain g where g: "ALL b : B. g b : A \ f (g b) = b" + using bchoice[of B ?phi] by blast + hence gg: "ALL b : f ` A. g b : A \ f (g b) = b" using f by blast + have gf: "inver g f A" unfolding inver_def by (metis gg imageI inj_f the_inv_into_f_f) + moreover have "g ` B \ A \ inver f g B" using g unfolding inver_def by blast + moreover have "A \ g ` B" + proof safe + fix a assume a: "a : A" + hence "f a : B" using f by auto + moreover have "a = g (f a)" using a gf unfolding inver_def by auto + ultimately show "a : g ` B" by blast + qed + ultimately show ?R by blast +next + assume ?R + then obtain g where g: "g ` B = A \ inver g f A \ inver f g B" by blast + show ?L unfolding bij_betw_def + proof safe + show "inj_on f A" unfolding inj_on_def + proof safe + fix a1 a2 assume a: "a1 : A" "a2 : A" and "f a1 = f a2" + hence "g (f a1) = g (f a2)" by simp + thus "a1 = a2" using a g unfolding inver_def by simp + qed + next + fix a assume "a : A" + then obtain b where b: "b : B" and a: "a = g b" using g by blast + hence "b = f (g b)" using g unfolding inver_def by auto + thus "f a : B" unfolding a using b by simp + next + fix b assume "b : B" + hence "g b : A \ b = f (g b)" using g unfolding inver_def by auto + thus "b : f ` A" by auto + qed +qed + +lemma bijI: "\\x y. (f x = f y) = (x = y); \y. \x. y = f x\ \ bij f" +unfolding bij_def inj_on_def by auto blast + +lemma bij_betw_ex_weakE: + "\bij_betw f A B\ \ \g. g ` B \ A \ inver g f A \ inver f g B" +by (auto simp only: bij_betw_iff_ex) + +lemma inver_surj: "\g ` B \ A; f ` A \ B; inver g f A\ \ g ` B = A" +unfolding inver_def by auto (rule rev_image_eqI, auto) + +lemma inver_mono: "\A \ B; inver f g B\ \ inver f g A" +unfolding inver_def by auto + +lemma inver_pointfree: "inver f g A = (\a \ A. (f o g) a = a)" +unfolding inver_def by simp + +lemma bij_betwE: "bij_betw f A B \ \a\A. f a \ B" +unfolding bij_betw_def by auto + +lemma bij_betw_imageE: "bij_betw f A B \ f ` A = B" +unfolding bij_betw_def by auto + +lemma inverE: "\inver f f' A; x \ A\ \ f (f' x) = x" +unfolding inver_def by auto + +lemma bij_betw_inver1: "bij_betw f A B \ inver (inv_into A f) f A" +unfolding bij_betw_def inver_def by auto + +lemma bij_betw_inver2: "bij_betw f A B \ inver f (inv_into A f) B" +unfolding bij_betw_def inver_def by auto + +lemma bij_betwI: "\bij_betw g B A; inver g f A; inver f g B\ \ bij_betw f A B" +by (metis bij_betw_iff_ex bij_betw_imageE) + +lemma bij_betwI': + "\\x y. \x \ X; y \ X\ \ (f x = f y) = (x = y); + \x. x \ X \ f x \ Y; + \y. y \ Y \ \x \ X. y = f x\ \ bij_betw f X Y" +unfolding bij_betw_def inj_on_def by auto (metis rev_image_eqI) + +lemma o_bij: + assumes gf: "g o f = id" and fg: "f o g = id" + shows "bij f" +unfolding bij_def inj_on_def surj_def proof safe + fix a1 a2 assume "f a1 = f a2" + hence "g ( f a1) = g (f a2)" by simp + thus "a1 = a2" using gf unfolding fun_eq_iff by simp +next + fix b + have "b = f (g b)" + using fg unfolding fun_eq_iff by simp + thus "EX a. b = f a" by blast +qed + +lemma surj_fun_eq: + assumes surj_on: "f ` X = UNIV" and eq_on: "\x \ X. (g1 o f) x = (g2 o f) x" + shows "g1 = g2" +proof (rule ext) + fix y + from surj_on obtain x where "x \ X" and "y = f x" by blast + thus "g1 y = g2 y" using eq_on by simp +qed + +lemma Card_order_wo_rel: "Card_order r \ wo_rel r" +unfolding wo_rel_def card_order_on_def by blast + +lemma Cinfinite_limit: "\x \ Field r; Cinfinite r\ \ + \y \ Field r. x \ y \ (x, y) \ r" +unfolding cinfinite_def by (auto simp add: infinite_Card_order_limit) + +lemma Card_order_trans: + "\Card_order r; x \ y; (x, y) \ r; y \ z; (y, z) \ r\ \ x \ z \ (x, z) \ r" +unfolding card_order_on_def well_order_on_def linear_order_on_def + partial_order_on_def preorder_on_def trans_def antisym_def by blast + +lemma Cinfinite_limit2: + assumes x1: "x1 \ Field r" and x2: "x2 \ Field r" and r: "Cinfinite r" + shows "\y \ Field r. (x1 \ y \ (x1, y) \ r) \ (x2 \ y \ (x2, y) \ r)" +proof - + from r have trans: "trans r" and total: "Total r" and antisym: "antisym r" + unfolding card_order_on_def well_order_on_def linear_order_on_def + partial_order_on_def preorder_on_def by auto + obtain y1 where y1: "y1 \ Field r" "x1 \ y1" "(x1, y1) \ r" + using Cinfinite_limit[OF x1 r] by blast + obtain y2 where y2: "y2 \ Field r" "x2 \ y2" "(x2, y2) \ r" + using Cinfinite_limit[OF x2 r] by blast + show ?thesis + proof (cases "y1 = y2") + case True with y1 y2 show ?thesis by blast + next + case False + with y1(1) y2(1) total have "(y1, y2) \ r \ (y2, y1) \ r" + unfolding total_on_def by auto + thus ?thesis + proof + assume *: "(y1, y2) \ r" + with trans y1(3) have "(x1, y2) \ r" unfolding trans_def by blast + with False y1 y2 * antisym show ?thesis by (cases "x1 = y2") (auto simp: antisym_def) + next + assume *: "(y2, y1) \ r" + with trans y2(3) have "(x2, y1) \ r" unfolding trans_def by blast + with False y1 y2 * antisym show ?thesis by (cases "x2 = y1") (auto simp: antisym_def) + qed + qed +qed + +lemma Cinfinite_limit_finite: "\finite X; X \ Field r; Cinfinite r\ + \ \y \ Field r. \x \ X. (x \ y \ (x, y) \ r)" +proof (induct X rule: finite_induct) + case empty thus ?case unfolding cinfinite_def using ex_in_conv[of "Field r"] finite.emptyI by auto +next + case (insert x X) + then obtain y where y: "y \ Field r" "\x \ X. (x \ y \ (x, y) \ r)" by blast + then obtain z where z: "z \ Field r" "x \ z \ (x, z) \ r" "y \ z \ (y, z) \ r" + using Cinfinite_limit2[OF _ y(1) insert(5), of x] insert(4) by blast + show ?case by (metis Card_order_trans insert(5) insertE y(2) z) +qed + +lemma insert_subsetI: "\x \ A; X \ A\ \ insert x X \ A" +by auto + +(*helps resolution*) +lemma well_order_induct_imp: + "wo_rel r \ (\x. \y. y \ x \ (y, x) \ r \ y \ Field r \ P y \ x \ Field r \ P x) \ + x \ Field r \ P x" +by (erule wo_rel.well_order_induct) + +lemma meta_spec2: + assumes "(\x y. PROP P x y)" + shows "PROP P x y" +by (rule `(\x y. PROP P x y)`) + +(*Extended Sublist*) + +definition prefCl where + "prefCl Kl = (\ kl1 kl2. kl1 \ kl2 \ kl2 \ Kl \ kl1 \ Kl)" +definition PrefCl where + "PrefCl A n = (\kl kl'. kl \ A n \ kl' \ kl \ (\m\n. kl' \ A m))" + +lemma prefCl_UN: + "\\n. PrefCl A n\ \ prefCl (\n. A n)" +unfolding prefCl_def PrefCl_def by fastforce + +definition Succ where "Succ Kl kl = {k . kl @ [k] \ Kl}" +definition Shift where "Shift Kl k = {kl. k # kl \ Kl}" +definition shift where "shift lab k = (\kl. lab (k # kl))" + +lemmas sh_def = Shift_def shift_def + +lemma empty_Shift: "\[] \ Kl; k \ Succ Kl []\ \ [] \ Shift Kl k" +unfolding Shift_def Succ_def by simp + +lemma Shift_clists: "Kl \ Field (clists r) \ Shift Kl k \ Field (clists r)" +unfolding Shift_def clists_def Field_card_of by auto + +lemma Shift_prefCl: "prefCl Kl \ prefCl (Shift Kl k)" +unfolding prefCl_def Shift_def +proof safe + fix kl1 kl2 + assume "\kl1 kl2. kl1 \ kl2 \ kl2 \ Kl \ kl1 \ Kl" + "kl1 \ kl2" "k # kl2 \ Kl" + thus "k # kl1 \ Kl" using Cons_prefix_Cons[of k kl1 k kl2] by blast +qed + +lemma not_in_Shift: "kl \ Shift Kl x \ x # kl \ Kl" +unfolding Shift_def by simp + +lemma prefCl_Succ: "\prefCl Kl; k # kl \ Kl\ \ k \ Succ Kl []" +unfolding Succ_def proof + assume "prefCl Kl" "k # kl \ Kl" + moreover have "k # [] \ k # kl" by auto + ultimately have "k # [] \ Kl" unfolding prefCl_def by blast + thus "[] @ [k] \ Kl" by simp +qed + +lemma SuccD: "k \ Succ Kl kl \ kl @ [k] \ Kl" +unfolding Succ_def by simp + +lemmas SuccE = SuccD[elim_format] + +lemma SuccI: "kl @ [k] \ Kl \ k \ Succ Kl kl" +unfolding Succ_def by simp + +lemma ShiftD: "kl \ Shift Kl k \ k # kl \ Kl" +unfolding Shift_def by simp + +lemma Succ_Shift: "Succ (Shift Kl k) kl = Succ Kl (k # kl)" +unfolding Succ_def Shift_def by auto + +lemma ShiftI: "k # kl \ Kl \ kl \ Shift Kl k" +unfolding Shift_def by simp + +lemma Func_cexp: "|Func A B| =o |B| ^c |A|" +unfolding cexp_def Field_card_of by (simp only: card_of_refl) + +lemma clists_bound: "A \ Field (cpow (clists r)) - {{}} \ |A| \o clists r" +unfolding cpow_def clists_def Field_card_of by (auto simp: card_of_mono1) + +lemma cpow_clists_czero: "\A \ Field (cpow (clists r)) - {{}}; |A| =o czero\ \ False" +unfolding cpow_def clists_def +by (auto simp add: card_of_ordIso_czero_iff_empty[symmetric]) + (erule notE, erule ordIso_transitive, rule czero_ordIso) + +lemma incl_UNION_I: +assumes "i \ I" and "A \ F i" +shows "A \ UNION I F" +using assms by auto + +lemma Nil_clists: "{[]} \ Field (clists r)" +unfolding clists_def Field_card_of by auto + +lemma Cons_clists: + "\x \ Field r; xs \ Field (clists r)\ \ x # xs \ Field (clists r)" +unfolding clists_def Field_card_of by auto + +lemma length_Cons: "length (x#xs) = Suc (length xs)" +by simp + +lemma length_append_singleton: "length (xs @ [x]) = Suc (length xs)" +by simp + +(*injection into the field of a cardinal*) +definition "toCard_pred A r f \ inj_on f A \ f ` A \ Field r \ Card_order r" +definition "toCard A r \ SOME f. toCard_pred A r f" + +lemma ex_toCard_pred: +"\|A| \o r; Card_order r\ \ \ f. toCard_pred A r f" +unfolding toCard_pred_def +using card_of_ordLeq[of A "Field r"] + ordLeq_ordIso_trans[OF _ card_of_unique[of "Field r" r], of "|A|"] +by blast + +lemma toCard_pred_toCard: + "\|A| \o r; Card_order r\ \ toCard_pred A r (toCard A r)" +unfolding toCard_def using someI_ex[OF ex_toCard_pred] . + +lemma toCard_inj: "\|A| \o r; Card_order r; x \ A; y \ A\ \ + toCard A r x = toCard A r y \ x = y" +using toCard_pred_toCard unfolding inj_on_def toCard_pred_def by blast + +lemma toCard: "\|A| \o r; Card_order r; b \ A\ \ toCard A r b \ Field r" +using toCard_pred_toCard unfolding toCard_pred_def by blast + +definition "fromCard A r k \ SOME b. b \ A \ toCard A r b = k" + +lemma fromCard_toCard: +"\|A| \o r; Card_order r; b \ A\ \ fromCard A r (toCard A r b) = b" +unfolding fromCard_def by (rule some_equality) (auto simp add: toCard_inj) + +(* pick according to the weak pullback *) +definition pickWP_pred where +"pickWP_pred A p1 p2 b1 b2 a \ + a \ A \ p1 a = b1 \ p2 a = b2" + +definition pickWP where +"pickWP A p1 p2 b1 b2 \ + SOME a. pickWP_pred A p1 p2 b1 b2 a" + +lemma pickWP_pred: +assumes "wpull A B1 B2 f1 f2 p1 p2" and +"b1 \ B1" and "b2 \ B2" and "f1 b1 = f2 b2" +shows "\ a. pickWP_pred A p1 p2 b1 b2 a" +using assms unfolding wpull_def pickWP_pred_def by blast + +lemma pickWP_pred_pickWP: +assumes "wpull A B1 B2 f1 f2 p1 p2" and +"b1 \ B1" and "b2 \ B2" and "f1 b1 = f2 b2" +shows "pickWP_pred A p1 p2 b1 b2 (pickWP A p1 p2 b1 b2)" +unfolding pickWP_def using assms by(rule someI_ex[OF pickWP_pred]) + +lemma pickWP: +assumes "wpull A B1 B2 f1 f2 p1 p2" and +"b1 \ B1" and "b2 \ B2" and "f1 b1 = f2 b2" +shows "pickWP A p1 p2 b1 b2 \ A" + "p1 (pickWP A p1 p2 b1 b2) = b1" + "p2 (pickWP A p1 p2 b1 b2) = b2" +using assms pickWP_pred_pickWP unfolding pickWP_pred_def by fastforce+ + +lemma ssubst_mem: "\t = s; s \ X\ \ t \ X" by simp + +lemma Inl_Field_csum: "a \ Field r \ Inl a \ Field (r +c s)" +unfolding Field_card_of csum_def by auto + +lemma Inr_Field_csum: "a \ Field s \ Inr a \ Field (r +c s)" +unfolding Field_card_of csum_def by auto + +lemma nat_rec_0: "f = nat_rec f1 (%n rec. f2 n rec) \ f 0 = f1" +by auto + +lemma nat_rec_Suc: "f = nat_rec f1 (%n rec. f2 n rec) \ f (Suc n) = f2 n (f n)" +by auto + +lemma list_rec_Nil: "f = list_rec f1 (%x xs rec. f2 x xs rec) \ f [] = f1" +by auto + +lemma list_rec_Cons: "f = list_rec f1 (%x xs rec. f2 x xs rec) \ f (x # xs) = f2 x xs (f xs)" +by auto + +lemma sum_case_cong: "p = q \ sum_case f g p = sum_case f g q" +by simp + +lemma sum_case_step: + "sum_case (sum_case f' g') g (Inl p) = sum_case f' g' p" + "sum_case f (sum_case f' g') (Inr p) = sum_case f' g' p" +by auto + +lemma one_pointE: "\\x. s = x \ P\ \ P" +by simp + +lemma obj_one_pointE: "\x. s = x \ P \ P" +by blast + +lemma obj_sumE_f: +"\\x. s = f (Inl x) \ P; \x. s = f (Inr x) \ P\ \ \x. s = f x \ P" +by (metis sum.exhaust) + +lemma obj_sumE: "\\x. s = Inl x \ P; \x. s = Inr x \ P\ \ P" +by (cases s) auto + +lemma obj_sum_step: + "\\x. s = f (Inr (Inl x)) \ P; \x. s = f (Inr (Inr x)) \ P\ \ \x. s = f (Inr x) \ P" +by (metis obj_sumE) + +lemma sum_case_if: +"sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)" +by simp + +lemma not_arg_cong_Inr: "x \ y \ Inr x \ Inr y" +by simp + +end diff -r 3d87f4fd0d50 -r c057e1b39f16 src/HOL/Codatatype/Tools/bnf_util.ML --- a/src/HOL/Codatatype/Tools/bnf_util.ML Tue Sep 11 17:06:27 2012 +0200 +++ b/src/HOL/Codatatype/Tools/bnf_util.ML Tue Sep 11 17:09:39 2012 +0200 @@ -2,7 +2,7 @@ Author: Dmitriy Traytel, TU Muenchen Copyright 2012 -General library functions. +Library for bounded natural functors. *) signature BNF_UTIL =