diff -r 56fcd826f90c -r c874ff5658dc src/HOL/Codatatype/BNF_Def.thy --- a/src/HOL/Codatatype/BNF_Def.thy Wed Sep 12 06:27:36 2012 +0200 +++ b/src/HOL/Codatatype/BNF_Def.thy Wed Sep 12 06:27:48 2012 +0200 @@ -14,6 +14,119 @@ "bnf_def" :: thy_goal begin +lemma collect_o: "collect F o g = collect ((\f. f o g) ` F)" +by (rule ext) (auto simp only: o_apply collect_def) + +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 + +definition csquare where +"csquare A f1 f2 p1 p2 \ (\ a \ A. f1 (p1 a) = f2 (p2 a))" + +(* 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 + +lemma Id_alt: "Id = Gr UNIV id" +unfolding Gr_def by auto + +lemma Gr_UNIV_id: "f = id \ (Gr UNIV f)^-1 O Gr UNIV f = Gr UNIV f" +unfolding Gr_def by auto + +lemma Gr_mono: "A \ B \ Gr A f \ Gr B f" +unfolding Gr_def by auto + +lemma wpull_Gr: +"wpull (Gr A f) A (f ` A) f id fst snd" +unfolding wpull_def Gr_def by auto + +definition "pick_middle P Q a c = (SOME b. (a,b) \ P \ (b,c) \ Q)" + +lemma pick_middle: +assumes "(a,c) \ P O Q" +shows "(a, pick_middle P Q a c) \ P \ (pick_middle P Q a c, c) \ Q" +unfolding pick_middle_def apply(rule someI_ex) +using assms unfolding relcomp_def by auto + +definition fstO where "fstO P Q ac = (fst ac, pick_middle P Q (fst ac) (snd ac))" +definition sndO where "sndO P Q ac = (pick_middle 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 pick_middle 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 pick_middle 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 pick_middle by auto + +lemma wppull_fstO_sndO: +shows "wppull (P O Q) P Q snd fst fst snd (fstO P Q) (sndO P Q)" +using pick_middle unfolding wppull_def fstO_def sndO_def relcomp_def by auto + +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 pointfreeE: "f o g = f' o g' \ f (g x) = f' (g' x)" +unfolding o_def fun_eq_iff by simp + ML_file "Tools/bnf_def_tactics.ML" ML_file"Tools/bnf_def.ML"