blanchet@50323: (* Title: HOL/Codatatype/BNF_FP.thy blanchet@50323: Author: Dmitriy Traytel, TU Muenchen blanchet@50323: Author: Jasmin Blanchette, TU Muenchen blanchet@50323: Copyright 2012 blanchet@50323: blanchet@50323: Composition of bounded natural functors. blanchet@50323: *) blanchet@50323: blanchet@50323: header {* Composition of Bounded Natural Functors *} blanchet@50323: blanchet@50323: theory BNF_FP blanchet@50323: imports BNF_Comp BNF_Wrap blanchet@50323: keywords blanchet@50323: "defaults" blanchet@50323: begin blanchet@50323: blanchet@50327: lemma case_unit: "(case u of () => f) = f" blanchet@50327: by (cases u) (hypsubst, rule unit.cases) blanchet@50327: blanchet@50350: lemma unit_all_impI: "(P () \ Q ()) \ \x. P x \ Q x" blanchet@50350: by simp blanchet@50350: blanchet@50350: lemma prod_all_impI: "(\x y. P (x, y) \ Q (x, y)) \ \x. P x \ Q x" blanchet@50350: by clarify blanchet@50350: blanchet@50350: lemma prod_all_impI_step: "(\x. \y. P (x, y) \ Q (x, y)) \ \x. P x \ Q x" blanchet@50350: by auto blanchet@50350: blanchet@50383: lemma all_unit_eq: "(\x. PROP P x) \ PROP P ()" blanchet@50383: by simp blanchet@50327: blanchet@50383: lemma all_prod_eq: "(\x. PROP P x) \ (\a b. PROP P (a, b))" blanchet@50383: by clarsimp blanchet@50327: blanchet@50327: lemma rev_bspec: "a \ A \ \z \ A. P z \ P a" blanchet@50327: by simp blanchet@50327: blanchet@50327: lemma Un_cong: "\A = B; C = D\ \ A \ C = B \ D" blanchet@50327: by simp blanchet@50327: blanchet@50327: definition convol ("<_ , _>") where blanchet@50327: " \ %a. (f a, g a)" blanchet@50327: blanchet@50327: lemma fst_convol: blanchet@50327: "fst o = f" blanchet@50327: apply(rule ext) blanchet@50327: unfolding convol_def by simp blanchet@50327: blanchet@50327: lemma snd_convol: blanchet@50327: "snd o = g" blanchet@50327: apply(rule ext) blanchet@50327: unfolding convol_def by simp blanchet@50327: blanchet@50327: lemma pointfree_idE: "f o g = id \ f (g x) = x" blanchet@50327: unfolding o_def fun_eq_iff by simp blanchet@50327: blanchet@50327: lemma o_bij: blanchet@50327: assumes gf: "g o f = id" and fg: "f o g = id" blanchet@50327: shows "bij f" blanchet@50327: unfolding bij_def inj_on_def surj_def proof safe blanchet@50327: fix a1 a2 assume "f a1 = f a2" blanchet@50327: hence "g ( f a1) = g (f a2)" by simp blanchet@50327: thus "a1 = a2" using gf unfolding fun_eq_iff by simp blanchet@50327: next blanchet@50327: fix b blanchet@50327: have "b = f (g b)" blanchet@50327: using fg unfolding fun_eq_iff by simp blanchet@50327: thus "EX a. b = f a" by blast blanchet@50327: qed blanchet@50327: blanchet@50327: lemma ssubst_mem: "\t = s; s \ X\ \ t \ X" by simp blanchet@50327: blanchet@50327: lemma sum_case_step: blanchet@50327: "sum_case (sum_case f' g') g (Inl p) = sum_case f' g' p" blanchet@50327: "sum_case f (sum_case f' g') (Inr p) = sum_case f' g' p" blanchet@50327: by auto blanchet@50327: blanchet@50327: lemma one_pointE: "\\x. s = x \ P\ \ P" blanchet@50327: by simp blanchet@50327: blanchet@50327: lemma obj_one_pointE: "\x. s = x \ P \ P" blanchet@50327: by blast blanchet@50327: blanchet@50340: lemma obj_sumE_f': blanchet@50340: "\\x. s = f (Inl x) \ P; \x. s = f (Inr x) \ P\ \ s = f x \ P" blanchet@50340: by (cases x) blast+ blanchet@50340: blanchet@50327: lemma obj_sumE_f: blanchet@50327: "\\x. s = f (Inl x) \ P; \x. s = f (Inr x) \ P\ \ \x. s = f x \ P" blanchet@50340: by (rule allI) (rule obj_sumE_f') blanchet@50327: blanchet@50327: lemma obj_sumE: "\\x. s = Inl x \ P; \x. s = Inr x \ P\ \ P" blanchet@50327: by (cases s) auto blanchet@50327: blanchet@50340: lemma obj_sum_step': blanchet@50340: "\\x. s = f (Inr (Inl x)) \ P; \x. s = f (Inr (Inr x)) \ P\ \ s = f (Inr x) \ P" blanchet@50340: by (cases x) blast+ blanchet@50340: blanchet@50327: lemma obj_sum_step: blanchet@50340: "\\x. s = f (Inr (Inl x)) \ P; \x. s = f (Inr (Inr x)) \ P\ \ \x. s = f (Inr x) \ P" blanchet@50340: by (rule allI) (rule obj_sum_step') blanchet@50327: blanchet@50327: lemma sum_case_if: blanchet@50327: "sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)" blanchet@50327: by simp blanchet@50327: blanchet@50441: lemma UN_compreh_bex_eq_empty: blanchet@50441: "\{y. \x\A. y = {}} = {}" blanchet@50441: by blast blanchet@50441: blanchet@50441: lemma UN_compreh_bex_eq_singleton: blanchet@50441: "\{y. \x\A. y = {f x}} = {y. \x\A. y = f x}" blanchet@50441: by blast blanchet@50441: blanchet@50441: lemma mem_UN_comprehI: blanchet@50441: "z \ {y. \x\A. y = f x} \ z \ \{y. \x\A. y = {f x}}" blanchet@50441: "z \ {y. \x\A. y = f x} \ B \ z \ \{y. \x\A. y = {f x}} \ B" blanchet@50441: "z \ \{y. \x\A. y = F x} \ \{y. \x\A. y = G x} \ z \ \{y. \x\A. y = F x \ G x}" blanchet@50441: "z \ \{y. \x\A. y = F x} \ (\{y. \x\A. y = G x} \ B) \ z \ \{y. \x\A. y = F x \ G x} \ B" blanchet@50383: by blast+ blanchet@50383: blanchet@50441: lemma mem_UN_comprehI': blanchet@50441: "\y. (\x\A. y = F x) \ z \ y \ z \ \{y. \x\A. y = {y. \y'\F x. y = y'}}" blanchet@50441: by blast blanchet@50441: blanchet@50443: lemma mem_UN_compreh_eq: "(z : \{y. \x\A. y = F x}) = (\x\A. z : F x)" blanchet@50443: by blast blanchet@50443: blanchet@50443: lemma eq_UN_compreh_Un: "(xa = \{y. \x\A. y = c_set1 x \ c_set2 x}) = blanchet@50443: (xa = (\{y. \x\A. y = c_set1 x} \ \{y. \x\A. y = c_set2 x}))" blanchet@50443: by blast blanchet@50443: blanchet@50442: lemma mem_compreh_eq_iff: blanchet@50443: "z \ {y. \x\A. y = f x} = (\ x. x \ A & z \ {f x})" blanchet@50442: by blast blanchet@50442: blanchet@50442: lemma ex_mem_singleton: "(\y. y \ A \ y \ {x}) = (x \ A)" blanchet@50442: by simp blanchet@50442: blanchet@50444: lemma prod_set_simps: blanchet@50444: "fsts (x, y) = {x}" blanchet@50444: "snds (x, y) = {y}" blanchet@50444: unfolding fsts_def snds_def by simp+ blanchet@50444: blanchet@50444: lemma sum_set_simps: blanchet@50444: "sum_setl (Inl x) = {x}" blanchet@50444: "sum_setl (Inr x) = {}" blanchet@50444: "sum_setr (Inl x) = {}" blanchet@50444: "sum_setr (Inr x) = {x}" blanchet@50444: unfolding sum_setl_def sum_setr_def by simp+ blanchet@50444: blanchet@50442: lemma induct_set_step: blanchet@50442: "\b \ A; c \ F b\ \ \x. x \ A \ c \ F x" blanchet@50442: "\B \ A; c \ f B\ \ \C. (\a \ A. C = f a) \c \ C" blanchet@50442: by blast+ blanchet@50383: blanchet@50324: ML_file "Tools/bnf_fp_util.ML" blanchet@50324: ML_file "Tools/bnf_fp_sugar_tactics.ML" blanchet@50324: ML_file "Tools/bnf_fp_sugar.ML" blanchet@50324: blanchet@50323: end