blanchet@58128: (* Title: HOL/BNF_Fixpoint_Base.thy blanchet@53311: Author: Lorenz Panny, TU Muenchen blanchet@49308: Author: Dmitriy Traytel, TU Muenchen blanchet@49308: Author: Jasmin Blanchette, TU Muenchen blanchet@57698: Author: Martin Desharnais, TU Muenchen blanchet@57698: Copyright 2012, 2013, 2014 blanchet@49308: blanchet@58352: Shared fixpoint operations on bounded natural functors. blanchet@49308: *) blanchet@49308: blanchet@58352: header {* Shared Fixpoint Operations on Bounded Natural Functors *} blanchet@49308: blanchet@58128: theory BNF_Fixpoint_Base blanchet@58128: imports BNF_Composition Basic_BNFs blanchet@49308: begin blanchet@49308: desharna@57525: lemma False_imp_eq_True: "(False \ Q) \ Trueprop True" desharna@57525: by default simp_all desharna@57525: desharna@57525: lemma conj_imp_eq_imp_imp: "(P \ Q \ PROP R) \ (P \ Q \ PROP R)" desharna@57525: by default simp_all desharna@57525: blanchet@49590: lemma mp_conj: "(P \ Q) \ R \ P \ R \ Q" blanchet@58159: by auto blanchet@49590: desharna@57303: lemma predicate2D_conj: "P \ Q \ R \ P x y \ R \ Q x y" desharna@57302: by auto desharna@57302: blanchet@49591: lemma eq_sym_Unity_conv: "(x = (() = ())) = x" blanchet@58159: by blast blanchet@49585: blanchet@55414: lemma case_unit_Unity: "(case u of () \ f) = f" blanchet@58159: by (cases u) (hypsubst, rule unit.case) blanchet@49312: blanchet@55414: lemma case_prod_Pair_iden: "(case p of (x, y) \ (x, y)) = p" blanchet@58159: by simp blanchet@49539: blanchet@49335: lemma unit_all_impI: "(P () \ Q ()) \ \x. P x \ Q x" blanchet@58159: by simp blanchet@49335: blanchet@49683: lemma pointfree_idE: "f \ g = id \ f (g x) = x" blanchet@58159: unfolding comp_def fun_eq_iff by simp blanchet@49312: blanchet@49312: lemma o_bij: blanchet@49683: assumes gf: "g \ f = id" and fg: "f \ g = id" blanchet@49312: shows "bij f" blanchet@49312: unfolding bij_def inj_on_def surj_def proof safe blanchet@49312: fix a1 a2 assume "f a1 = f a2" blanchet@49312: hence "g ( f a1) = g (f a2)" by simp blanchet@49312: thus "a1 = a2" using gf unfolding fun_eq_iff by simp blanchet@49312: next blanchet@49312: fix b blanchet@49312: have "b = f (g b)" blanchet@49312: using fg unfolding fun_eq_iff by simp blanchet@49312: thus "EX a. b = f a" by blast blanchet@49312: qed blanchet@49312: blanchet@58159: lemma ssubst_mem: "\t = s; s \ X\ \ t \ X" blanchet@58159: by simp blanchet@49312: blanchet@55414: lemma case_sum_step: blanchet@58159: "case_sum (case_sum f' g') g (Inl p) = case_sum f' g' p" blanchet@58159: "case_sum f (case_sum f' g') (Inr p) = case_sum f' g' p" blanchet@58159: by auto blanchet@49312: blanchet@49312: lemma obj_one_pointE: "\x. s = x \ P \ P" blanchet@58159: by blast blanchet@49312: traytel@55803: lemma type_copy_obj_one_point_absE: traytel@55811: assumes "type_definition Rep Abs UNIV" "\x. s = Abs x \ P" shows P traytel@55811: using type_definition.Rep_inverse[OF assms(1)] traytel@55811: by (intro mp[OF spec[OF assms(2), of "Rep s"]]) simp traytel@55803: blanchet@49312: lemma obj_sumE_f: traytel@55811: assumes "\x. s = f (Inl x) \ P" "\x. s = f (Inr x) \ P" traytel@55811: shows "\x. s = f x \ P" traytel@55811: proof traytel@55811: fix x from assms show "s = f x \ P" by (cases x) auto traytel@55811: qed blanchet@49312: blanchet@55414: lemma case_sum_if: blanchet@58159: "case_sum f g (if p then Inl x else Inr y) = (if p then f x else g y)" blanchet@58159: by simp blanchet@49312: blanchet@49429: lemma prod_set_simps: blanchet@58159: "fsts (x, y) = {x}" blanchet@58159: "snds (x, y) = {y}" blanchet@58159: unfolding fsts_def snds_def by simp+ blanchet@49429: blanchet@49429: lemma sum_set_simps: blanchet@58159: "setl (Inl x) = {x}" blanchet@58159: "setl (Inr x) = {}" blanchet@58159: "setr (Inl x) = {}" blanchet@58159: "setr (Inr x) = {x}" blanchet@58159: unfolding sum_set_defs by simp+ blanchet@49429: desharna@57301: lemma Inl_Inr_False: "(Inl x = Inr y) = False" desharna@57301: by simp desharna@57301: desharna@57471: lemma Inr_Inl_False: "(Inr x = Inl y) = False" desharna@57471: by simp desharna@57471: traytel@52505: lemma spec2: "\x y. P x y \ P x y" blanchet@58159: by blast traytel@52505: blanchet@56640: lemma rewriteR_comp_comp: "\g \ h = r\ \ f \ g \ h = f \ r" blanchet@55066: unfolding comp_def fun_eq_iff by auto traytel@52913: blanchet@56640: lemma rewriteR_comp_comp2: "\g \ h = r1 \ r2; f \ r1 = l\ \ f \ g \ h = l \ r2" blanchet@55066: unfolding comp_def fun_eq_iff by auto traytel@52913: blanchet@56640: lemma rewriteL_comp_comp: "\f \ g = l\ \ f \ (g \ h) = l \ h" blanchet@55066: unfolding comp_def fun_eq_iff by auto traytel@52913: blanchet@56640: lemma rewriteL_comp_comp2: "\f \ g = l1 \ l2; l2 \ h = r\ \ f \ (g \ h) = l1 \ r" blanchet@55066: unfolding comp_def fun_eq_iff by auto traytel@52913: wenzelm@57641: lemma convol_o: "\f, g\ \ h = \f \ h, g \ h\" traytel@52913: unfolding convol_def by auto traytel@52913: wenzelm@57641: lemma map_prod_o_convol: "map_prod h1 h2 \ \f, g\ = \h1 \ f, h2 \ g\" traytel@52913: unfolding convol_def by auto traytel@52913: wenzelm@57641: lemma map_prod_o_convol_id: "(map_prod f id \ \id, g\) x = \id \ f, g\ x" blanchet@55932: unfolding map_prod_o_convol id_comp comp_id .. traytel@52913: blanchet@56640: lemma o_case_sum: "h \ case_sum f g = case_sum (h \ f) (h \ g)" blanchet@55066: unfolding comp_def by (auto split: sum.splits) traytel@52913: blanchet@56640: lemma case_sum_o_map_sum: "case_sum f g \ map_sum h1 h2 = case_sum (f \ h1) (g \ h2)" blanchet@55066: unfolding comp_def by (auto split: sum.splits) traytel@52913: blanchet@56640: lemma case_sum_o_map_sum_id: "(case_sum id g \ map_sum f id) x = case_sum (f \ id) g x" blanchet@55931: unfolding case_sum_o_map_sum id_comp comp_id .. traytel@52913: blanchet@55945: lemma rel_fun_def_butlast: blanchet@55945: "rel_fun R (rel_fun S T) f g = (\x y. R x y \ (rel_fun S T) (f x) (g y))" blanchet@55945: unfolding rel_fun_def .. traytel@52731: traytel@53105: lemma subst_eq_imp: "(\a b. a = b \ P a b) \ (\a. P a a)" traytel@53105: by auto traytel@53105: traytel@53105: lemma eq_subset: "op = \ (\a b. P a b \ a = b)" traytel@53105: by auto traytel@53105: blanchet@53308: lemma eq_le_Grp_id_iff: "(op = \ Grp (Collect R) id) = (All R)" blanchet@53308: unfolding Grp_def id_apply by blast blanchet@53308: blanchet@53308: lemma Grp_id_mono_subst: "(\x y. Grp P id x y \ Grp Q id (f x) (f y)) \ blanchet@53308: (\x. x \ P \ f x \ Q)" blanchet@53308: unfolding Grp_def by rule auto blanchet@53308: traytel@55803: lemma vimage2p_mono: "vimage2p f g R x y \ R \ S \ vimage2p f g S x y" traytel@55803: unfolding vimage2p_def by blast traytel@55803: traytel@55803: lemma vimage2p_refl: "(\x. R x x) \ vimage2p f f R x x" traytel@55803: unfolding vimage2p_def by auto traytel@55803: traytel@55803: lemma traytel@55803: assumes "type_definition Rep Abs UNIV" blanchet@56640: shows type_copy_Rep_o_Abs: "Rep \ Abs = id" and type_copy_Abs_o_Rep: "Abs \ Rep = id" traytel@55803: unfolding fun_eq_iff comp_apply id_apply traytel@55803: type_definition.Abs_inverse[OF assms UNIV_I] type_definition.Rep_inverse[OF assms] by simp_all traytel@55803: traytel@55803: lemma type_copy_map_comp0_undo: traytel@55803: assumes "type_definition Rep Abs UNIV" traytel@55803: "type_definition Rep' Abs' UNIV" traytel@55803: "type_definition Rep'' Abs'' UNIV" blanchet@56640: shows "Abs' \ M \ Rep'' = (Abs' \ M1 \ Rep) \ (Abs \ M2 \ Rep'') \ M1 \ M2 = M" traytel@55803: by (rule sym) (auto simp: fun_eq_iff type_definition.Abs_inject[OF assms(2) UNIV_I UNIV_I] traytel@55803: type_definition.Abs_inverse[OF assms(1) UNIV_I] traytel@55803: type_definition.Abs_inverse[OF assms(3) UNIV_I] dest: spec[of _ "Abs'' x" for x]) traytel@55803: blanchet@55854: lemma vimage2p_id: "vimage2p id id R = R" blanchet@55854: unfolding vimage2p_def by auto blanchet@55854: blanchet@56640: lemma vimage2p_comp: "vimage2p (f1 \ f2) (g1 \ g2) = vimage2p f2 g2 \ vimage2p f1 g1" traytel@55803: unfolding fun_eq_iff vimage2p_def o_apply by simp traytel@55803: blanchet@56650: lemma fun_cong_unused_0: "f = (\x. g) \ f (\x. 0) = g" blanchet@56650: by (erule arg_cong) blanchet@56650: blanchet@56684: lemma inj_on_convol_ident: "inj_on (\x. (x, f x)) X" blanchet@56650: unfolding inj_on_def by simp blanchet@56650: blanchet@56650: lemma case_prod_app: "case_prod f x y = case_prod (\l r. f l r y) x" blanchet@56650: by (case_tac x) simp blanchet@56650: blanchet@56650: lemma case_sum_map_sum: "case_sum l r (map_sum f g x) = case_sum (l \ f) (r \ g) x" blanchet@56650: by (case_tac x) simp+ blanchet@56650: blanchet@56650: lemma case_prod_map_prod: "case_prod h (map_prod f g x) = case_prod (\l r. h (f l) (g r)) x" blanchet@56650: by (case_tac x) simp+ blanchet@56650: blanchet@56650: lemma prod_inj_map: "inj f \ inj g \ inj (map_prod f g)" blanchet@56650: by (simp add: inj_on_def) blanchet@56650: desharna@57489: lemma eq_ifI: "(P \ t = u1) \ (\ P \ t = u2) \ t = (if P then u1 else u2)" desharna@57489: by simp desharna@57489: blanchet@55062: ML_file "Tools/BNF/bnf_fp_util.ML" blanchet@55062: ML_file "Tools/BNF/bnf_fp_def_sugar_tactics.ML" blanchet@55062: ML_file "Tools/BNF/bnf_fp_def_sugar.ML" blanchet@55062: ML_file "Tools/BNF/bnf_fp_n2m_tactics.ML" blanchet@55062: ML_file "Tools/BNF/bnf_fp_n2m.ML" blanchet@55062: ML_file "Tools/BNF/bnf_fp_n2m_sugar.ML" blanchet@55702: blanchet@49308: end