blanchet@58128: (* Title: HOL/BNF_Greatest_Fixpoint.thy blanchet@48975: Author: Dmitriy Traytel, TU Muenchen blanchet@55059: Author: Lorenz Panny, TU Muenchen blanchet@55059: Author: Jasmin Blanchette, TU Muenchen blanchet@57698: Copyright 2012, 2013, 2014 blanchet@48975: blanchet@58352: Greatest fixpoint (codatatype) operation on bounded natural functors. blanchet@48975: *) blanchet@48975: wenzelm@60758: section \Greatest Fixpoint (Codatatype) Operation on Bounded Natural Functors\ blanchet@48975: blanchet@58128: theory BNF_Greatest_Fixpoint blanchet@58128: imports BNF_Fixpoint_Base String blanchet@48975: keywords blanchet@53310: "codatatype" :: thy_decl and panny@53822: "primcorecursive" :: thy_goal and panny@53822: "primcorec" :: thy_decl blanchet@48975: begin blanchet@48975: wenzelm@60758: setup \Sign.const_alias @{binding proj} @{const_name Equiv_Relations.proj}\ blanchet@55024: blanchet@55966: lemma one_pointE: "\\x. s = x \ P\ \ P" blanchet@57896: by simp blanchet@55966: blanchet@55966: lemma obj_sumE: "\\x. s = Inl x \ P; \x. s = Inr x \ P\ \ P" blanchet@57896: by (cases s) auto blanchet@55966: blanchet@54485: lemma not_TrueE: "\ True \ P" blanchet@57896: by (erule notE, rule TrueI) blanchet@54485: blanchet@54485: lemma neq_eq_eq_contradict: "\t \ u; s = t; s = u\ \ P" blanchet@57896: by fast blanchet@54485: blanchet@55414: lemma case_sum_expand_Inr: "f o Inl = g \ f x = case_sum g (f o Inr) x" blanchet@57896: by (auto split: sum.splits) blanchet@49312: blanchet@55414: lemma case_sum_expand_Inr': "f o Inl = g \ h = f o Inr \ case_sum g h = f" blanchet@57896: apply rule blanchet@57896: apply (rule ext, force split: sum.split) blanchet@57896: by (rule ext, metis case_sum_o_inj(2)) traytel@51739: blanchet@49312: lemma converse_Times: "(A \ B) ^-1 = B \ A" blanchet@57896: by fast blanchet@49312: blanchet@49312: lemma equiv_proj: blanchet@57896: assumes e: "equiv A R" and m: "z \ R" blanchet@49312: shows "(proj R o fst) z = (proj R o snd) z" blanchet@49312: proof - blanchet@57896: from m have z: "(fst z, snd z) \ R" by auto traytel@53695: with e have "\x. (fst z, x) \ R \ (snd z, x) \ R" "\x. (snd z, x) \ R \ (fst z, x) \ R" traytel@53695: unfolding equiv_def sym_def trans_def by blast+ traytel@53695: then show ?thesis unfolding proj_def[abs_def] by auto blanchet@49312: qed blanchet@49312: blanchet@49312: (* Operators: *) blanchet@49312: definition image2 where "image2 A f g = {(f a, g a) | a. a \ A}" blanchet@49312: traytel@51447: lemma Id_on_Gr: "Id_on A = Gr A id" blanchet@57896: unfolding Id_on_def Gr_def by auto blanchet@49312: blanchet@49312: lemma image2_eqI: "\b = f x; c = g x; x \ A\ \ (b, c) \ image2 A f g" blanchet@57896: unfolding image2_def by auto blanchet@49312: blanchet@49312: lemma IdD: "(a, b) \ Id \ a = b" blanchet@57896: by auto blanchet@49312: blanchet@49312: lemma image2_Gr: "image2 A f g = (Gr A f)^-1 O (Gr A g)" blanchet@57896: unfolding image2_def Gr_def by auto blanchet@49312: blanchet@49312: lemma GrD1: "(x, fx) \ Gr A f \ x \ A" blanchet@57896: unfolding Gr_def by simp blanchet@49312: blanchet@49312: lemma GrD2: "(x, fx) \ Gr A f \ f x = fx" blanchet@57896: unfolding Gr_def by simp blanchet@49312: wenzelm@61943: lemma Gr_incl: "Gr A f \ A \ B \ f ` A \ B" blanchet@57896: unfolding Gr_def by auto blanchet@49312: blanchet@54485: lemma subset_Collect_iff: "B \ A \ (B \ {x \ A. P x}) = (\x \ B. P x)" blanchet@57896: by blast blanchet@54485: blanchet@54485: lemma subset_CollectI: "B \ A \ (\x. x \ B \ Q x \ P x) \ ({x \ B. Q x} \ {x \ A. P x})" blanchet@57896: by blast blanchet@54485: haftmann@61424: lemma in_rel_Collect_case_prod_eq: "in_rel (Collect (case_prod X)) = X" blanchet@57896: unfolding fun_eq_iff by auto traytel@51893: haftmann@61424: lemma Collect_case_prod_in_rel_leI: "X \ Y \ X \ Collect (case_prod (in_rel Y))" blanchet@57896: by auto traytel@51893: haftmann@61424: lemma Collect_case_prod_in_rel_leE: "X \ Collect (case_prod (in_rel Y)) \ (X \ Y \ R) \ R" blanchet@57896: by force traytel@51893: traytel@51893: lemma conversep_in_rel: "(in_rel R)\\ = in_rel (R\)" blanchet@57896: unfolding fun_eq_iff by auto traytel@51893: traytel@51893: lemma relcompp_in_rel: "in_rel R OO in_rel S = in_rel (R O S)" blanchet@57896: unfolding fun_eq_iff by auto traytel@51893: traytel@51893: lemma in_rel_Gr: "in_rel (Gr A f) = Grp A f" blanchet@57896: unfolding Gr_def Grp_def fun_eq_iff by auto traytel@51893: blanchet@49312: definition relImage where blanchet@57896: "relImage R f \ {(f a1, f a2) | a1 a2. (a1,a2) \ R}" blanchet@49312: blanchet@49312: definition relInvImage where blanchet@57896: "relInvImage A R f \ {(a1, a2) | a1 a2. a1 \ A \ a2 \ A \ (f a1, f a2) \ R}" blanchet@49312: blanchet@49312: lemma relImage_Gr: blanchet@57896: "\R \ A \ A\ \ relImage R f = (Gr A f)^-1 O R O Gr A f" blanchet@57896: unfolding relImage_def Gr_def relcomp_def by auto blanchet@49312: blanchet@49312: lemma relInvImage_Gr: "\R \ B \ B\ \ relInvImage A R f = Gr A f O R O (Gr A f)^-1" blanchet@57896: unfolding Gr_def relcomp_def image_def relInvImage_def by auto blanchet@49312: blanchet@49312: lemma relImage_mono: blanchet@57896: "R1 \ R2 \ relImage R1 f \ relImage R2 f" blanchet@57896: unfolding relImage_def by auto blanchet@49312: blanchet@49312: lemma relInvImage_mono: blanchet@57896: "R1 \ R2 \ relInvImage A R1 f \ relInvImage A R2 f" blanchet@57896: unfolding relInvImage_def by auto blanchet@49312: traytel@51447: lemma relInvImage_Id_on: blanchet@57896: "(\a1 a2. f a1 = f a2 \ a1 = a2) \ relInvImage A (Id_on B) f \ Id" blanchet@57896: unfolding relInvImage_def Id_on_def by auto blanchet@49312: blanchet@49312: lemma relInvImage_UNIV_relImage: blanchet@57896: "R \ relInvImage UNIV (relImage R f) f" blanchet@57896: unfolding relInvImage_def relImage_def by auto blanchet@49312: blanchet@49312: lemma relImage_proj: blanchet@57896: assumes "equiv A R" blanchet@57896: shows "relImage R (proj R) \ Id_on (A//R)" blanchet@57896: unfolding relImage_def Id_on_def blanchet@57896: using proj_iff[OF assms] equiv_class_eq_iff[OF assms] blanchet@57896: by (auto simp: proj_preserves) blanchet@49312: blanchet@49312: lemma relImage_relInvImage: wenzelm@61943: assumes "R \ f ` A \ f ` A" blanchet@57896: shows "relImage (relInvImage A R f) f = R" blanchet@57896: using assms unfolding relImage_def relInvImage_def by fast blanchet@49312: blanchet@49312: lemma subst_Pair: "P x y \ a = (x, y) \ P (fst a) (snd a)" blanchet@57896: by simp blanchet@49312: traytel@55644: lemma fst_diag_id: "(fst \ (%x. (x, x))) z = id z" by simp traytel@55644: lemma snd_diag_id: "(snd \ (%x. (x, x))) z = id z" by simp blanchet@49312: traytel@55644: lemma fst_diag_fst: "fst o ((\x. (x, x)) o fst) = fst" by auto traytel@55644: lemma snd_diag_fst: "snd o ((\x. (x, x)) o fst) = fst" by auto traytel@55644: lemma fst_diag_snd: "fst o ((\x. (x, x)) o snd) = snd" by auto traytel@55644: lemma snd_diag_snd: "snd o ((\x. (x, x)) o snd) = snd" by auto blanchet@49312: blanchet@49312: definition Succ where "Succ Kl kl = {k . kl @ [k] \ Kl}" blanchet@49312: definition Shift where "Shift Kl k = {kl. k # kl \ Kl}" blanchet@49312: definition shift where "shift lab k = (\kl. lab (k # kl))" blanchet@49312: blanchet@49312: lemma empty_Shift: "\[] \ Kl; k \ Succ Kl []\ \ [] \ Shift Kl k" blanchet@57896: unfolding Shift_def Succ_def by simp blanchet@49312: blanchet@49312: lemma SuccD: "k \ Succ Kl kl \ kl @ [k] \ Kl" blanchet@57896: unfolding Succ_def by simp blanchet@49312: blanchet@49312: lemmas SuccE = SuccD[elim_format] blanchet@49312: blanchet@49312: lemma SuccI: "kl @ [k] \ Kl \ k \ Succ Kl kl" blanchet@57896: unfolding Succ_def by simp blanchet@49312: blanchet@49312: lemma ShiftD: "kl \ Shift Kl k \ k # kl \ Kl" blanchet@57896: unfolding Shift_def by simp blanchet@49312: blanchet@49312: lemma Succ_Shift: "Succ (Shift Kl k) kl = Succ Kl (k # kl)" blanchet@57896: unfolding Succ_def Shift_def by auto blanchet@49312: blanchet@49312: lemma length_Cons: "length (x # xs) = Suc (length xs)" blanchet@57896: by simp blanchet@49312: blanchet@49312: lemma length_append_singleton: "length (xs @ [x]) = Suc (length xs)" blanchet@57896: by simp blanchet@49312: blanchet@49312: (*injection into the field of a cardinal*) blanchet@49312: definition "toCard_pred A r f \ inj_on f A \ f ` A \ Field r \ Card_order r" blanchet@49312: definition "toCard A r \ SOME f. toCard_pred A r f" blanchet@49312: blanchet@49312: lemma ex_toCard_pred: blanchet@57896: "\|A| \o r; Card_order r\ \ \ f. toCard_pred A r f" blanchet@57896: unfolding toCard_pred_def blanchet@57896: using card_of_ordLeq[of A "Field r"] blanchet@57896: ordLeq_ordIso_trans[OF _ card_of_unique[of "Field r" r], of "|A|"] blanchet@57896: by blast blanchet@49312: blanchet@49312: lemma toCard_pred_toCard: blanchet@49312: "\|A| \o r; Card_order r\ \ toCard_pred A r (toCard A r)" blanchet@57896: unfolding toCard_def using someI_ex[OF ex_toCard_pred] . blanchet@49312: blanchet@57896: lemma toCard_inj: "\|A| \o r; Card_order r; x \ A; y \ A\ \ toCard A r x = toCard A r y \ x = y" blanchet@57896: using toCard_pred_toCard unfolding inj_on_def toCard_pred_def by blast blanchet@49312: blanchet@49312: definition "fromCard A r k \ SOME b. b \ A \ toCard A r b = k" blanchet@49312: blanchet@49312: lemma fromCard_toCard: blanchet@57896: "\|A| \o r; Card_order r; b \ A\ \ fromCard A r (toCard A r b) = b" blanchet@57896: unfolding fromCard_def by (rule some_equality) (auto simp add: toCard_inj) blanchet@49312: blanchet@49312: lemma Inl_Field_csum: "a \ Field r \ Inl a \ Field (r +c s)" blanchet@57896: unfolding Field_card_of csum_def by auto blanchet@49312: blanchet@49312: lemma Inr_Field_csum: "a \ Field s \ Inr a \ Field (r +c s)" blanchet@57896: unfolding Field_card_of csum_def by auto blanchet@49312: blanchet@55415: lemma rec_nat_0_imp: "f = rec_nat f1 (%n rec. f2 n rec) \ f 0 = f1" blanchet@57896: by auto blanchet@49312: blanchet@55415: lemma rec_nat_Suc_imp: "f = rec_nat f1 (%n rec. f2 n rec) \ f (Suc n) = f2 n (f n)" blanchet@57896: by auto blanchet@49312: blanchet@55413: lemma rec_list_Nil_imp: "f = rec_list f1 (%x xs rec. f2 x xs rec) \ f [] = f1" blanchet@57896: by auto blanchet@49312: blanchet@55413: lemma rec_list_Cons_imp: "f = rec_list f1 (%x xs rec. f2 x xs rec) \ f (x # xs) = f2 x xs (f xs)" blanchet@57896: by auto blanchet@49312: blanchet@49312: lemma not_arg_cong_Inr: "x \ y \ Inr x \ Inr y" blanchet@57896: by simp blanchet@49312: traytel@52731: definition image2p where traytel@52731: "image2p f g R = (\x y. \x' y'. R x' y' \ f x' = x \ g y' = y)" traytel@52731: blanchet@55463: lemma image2pI: "R x y \ image2p f g R (f x) (g y)" traytel@52731: unfolding image2p_def by blast traytel@52731: blanchet@55463: lemma image2pE: "\image2p f g R fx gy; (\x y. fx = f x \ gy = g y \ R x y \ P)\ \ P" traytel@52731: unfolding image2p_def by blast traytel@52731: blanchet@55945: lemma rel_fun_iff_geq_image2p: "rel_fun R S f g = (image2p f g R \ S)" blanchet@55945: unfolding rel_fun_def image2p_def by auto traytel@52731: blanchet@55945: lemma rel_fun_image2p: "rel_fun R (image2p f g R) f g" blanchet@55945: unfolding rel_fun_def image2p_def by auto traytel@52731: blanchet@55022: wenzelm@60758: subsection \Equivalence relations, quotients, and Hilbert's choice\ blanchet@55022: blanchet@55022: lemma equiv_Eps_in: blanchet@55022: "\equiv A r; X \ A//r\ \ Eps (%x. x \ X) \ X" blanchet@57896: apply (rule someI2_ex) blanchet@57896: using in_quotient_imp_non_empty by blast blanchet@55022: blanchet@55022: lemma equiv_Eps_preserves: blanchet@57896: assumes ECH: "equiv A r" and X: "X \ A//r" blanchet@57896: shows "Eps (%x. x \ X) \ A" blanchet@57896: apply (rule in_mono[rule_format]) blanchet@57896: using assms apply (rule in_quotient_imp_subset) blanchet@57896: by (rule equiv_Eps_in) (rule assms)+ blanchet@55022: blanchet@55022: lemma proj_Eps: blanchet@57896: assumes "equiv A r" and "X \ A//r" blanchet@57896: shows "proj r (Eps (%x. x \ X)) = X" blanchet@57896: unfolding proj_def blanchet@57896: proof auto blanchet@55022: fix x assume x: "x \ X" blanchet@55022: thus "(Eps (%x. x \ X), x) \ r" using assms equiv_Eps_in in_quotient_imp_in_rel by fast blanchet@55022: next blanchet@55022: fix x assume "(Eps (%x. x \ X),x) \ r" blanchet@55022: thus "x \ X" using in_quotient_imp_closed[OF assms equiv_Eps_in[OF assms]] by fast blanchet@55022: qed blanchet@55022: blanchet@55022: definition univ where "univ f X == f (Eps (%x. x \ X))" blanchet@55022: blanchet@55022: lemma univ_commute: blanchet@55022: assumes ECH: "equiv A r" and RES: "f respects r" and x: "x \ A" blanchet@55022: shows "(univ f) (proj r x) = f x" blanchet@57896: proof (unfold univ_def) blanchet@55022: have prj: "proj r x \ A//r" using x proj_preserves by fast blanchet@55022: hence "Eps (%y. y \ proj r x) \ A" using ECH equiv_Eps_preserves by fast blanchet@55022: moreover have "proj r (Eps (%y. y \ proj r x)) = proj r x" using ECH prj proj_Eps by fast blanchet@55022: ultimately have "(x, Eps (%y. y \ proj r x)) \ r" using x ECH proj_iff by fast blanchet@55022: thus "f (Eps (%y. y \ proj r x)) = f x" using RES unfolding congruent_def by fastforce blanchet@55022: qed blanchet@55022: blanchet@55022: lemma univ_preserves: blanchet@57991: assumes ECH: "equiv A r" and RES: "f respects r" and PRES: "\x \ A. f x \ B" blanchet@57896: shows "\X \ A//r. univ f X \ B" blanchet@55022: proof blanchet@55022: fix X assume "X \ A//r" blanchet@55022: then obtain x where x: "x \ A" and X: "X = proj r x" using ECH proj_image[of r A] by blast blanchet@57991: hence "univ f X = f x" using ECH RES univ_commute by fastforce blanchet@55022: thus "univ f X \ B" using x PRES by simp blanchet@55022: qed blanchet@55022: blanchet@55062: ML_file "Tools/BNF/bnf_gfp_util.ML" blanchet@55062: ML_file "Tools/BNF/bnf_gfp_tactics.ML" blanchet@55062: ML_file "Tools/BNF/bnf_gfp.ML" blanchet@55538: ML_file "Tools/BNF/bnf_gfp_rec_sugar_tactics.ML" blanchet@55538: ML_file "Tools/BNF/bnf_gfp_rec_sugar.ML" blanchet@49309: blanchet@48975: end