1 (* Title: HOL/BNF_Composition.thy
2 Author: Dmitriy Traytel, TU Muenchen
3 Author: Jasmin Blanchette, TU Muenchen
4 Copyright 2012, 2013, 2014
6 Composition of bounded natural functors.
9 section \<open>Composition of Bounded Natural Functors\<close>
11 theory BNF_Composition
14 "copy_bnf" :: thy_decl and
15 "lift_bnf" :: thy_goal
18 lemma ssubst_mem: "\<lbrakk>t = s; s \<in> X\<rbrakk> \<Longrightarrow> t \<in> X"
21 lemma empty_natural: "(\<lambda>_. {}) o f = image g o (\<lambda>_. {})"
24 lemma Union_natural: "Union o image (image f) = image f o Union"
25 by (rule ext) (auto simp only: comp_apply)
27 lemma in_Union_o_assoc: "x \<in> (Union o gset o gmap) A \<Longrightarrow> x \<in> (Union o (gset o gmap)) A"
28 by (unfold comp_assoc)
30 lemma comp_single_set_bd:
31 assumes fbd_Card_order: "Card_order fbd" and
32 fset_bd: "\<And>x. |fset x| \<le>o fbd" and
33 gset_bd: "\<And>x. |gset x| \<le>o gbd"
34 shows "|\<Union>(fset ` gset x)| \<le>o gbd *c fbd"
36 apply (rule ordLeq_transitive)
37 apply (rule card_of_UNION_Sigma)
38 apply (subst SIGMA_CSUM)
39 apply (rule ordLeq_transitive)
40 apply (rule card_of_Csum_Times')
41 apply (rule fbd_Card_order)
44 apply (rule ordLeq_transitive)
45 apply (rule cprod_mono1)
47 apply (rule ordIso_imp_ordLeq)
48 apply (rule ordIso_refl)
49 apply (rule Card_order_cprod)
52 lemma csum_dup: "cinfinite r \<Longrightarrow> Card_order r \<Longrightarrow> p +c p' =o r +c r \<Longrightarrow> p +c p' =o r"
53 apply (erule ordIso_transitive)
54 apply (frule csum_absorb2')
55 apply (erule ordLeq_refl)
58 lemma cprod_dup: "cinfinite r \<Longrightarrow> Card_order r \<Longrightarrow> p *c p' =o r *c r \<Longrightarrow> p *c p' =o r"
59 apply (erule ordIso_transitive)
60 apply (rule cprod_infinite)
63 lemma Union_image_insert: "\<Union>(f ` insert a B) = f a \<union> \<Union>(f ` B)"
66 lemma Union_image_empty: "A \<union> \<Union>(f ` {}) = A"
69 lemma image_o_collect: "collect ((\<lambda>f. image g o f) ` F) = image g o collect F"
70 by (rule ext) (auto simp add: collect_def)
72 lemma conj_subset_def: "A \<subseteq> {x. P x \<and> Q x} = (A \<subseteq> {x. P x} \<and> A \<subseteq> {x. Q x})"
75 lemma UN_image_subset: "\<Union>(f ` g x) \<subseteq> X = (g x \<subseteq> {x. f x \<subseteq> X})"
78 lemma comp_set_bd_Union_o_collect: "|\<Union>\<Union>((\<lambda>f. f x) ` X)| \<le>o hbd \<Longrightarrow> |(Union \<circ> collect X) x| \<le>o hbd"
79 by (unfold comp_apply collect_def) simp
82 "\<lbrakk>A' = A; B1' = B1; B2' = B2; wpull A B1 B2 f1 f2 p1 p2\<rbrakk> \<Longrightarrow> wpull A' B1' B2' f1 f2 p1 p2"
85 lemma Grp_fst_snd: "(Grp (Collect (split R)) fst)^--1 OO Grp (Collect (split R)) snd = R"
86 unfolding Grp_def fun_eq_iff relcompp.simps by auto
88 lemma OO_Grp_cong: "A = B \<Longrightarrow> (Grp A f)^--1 OO Grp A g = (Grp B f)^--1 OO Grp B g"
91 lemma vimage2p_relcompp_mono: "R OO S \<le> T \<Longrightarrow>
92 vimage2p f g R OO vimage2p g h S \<le> vimage2p f h T"
93 unfolding vimage2p_def by auto
95 lemma type_copy_map_cong0: "M (g x) = N (h x) \<Longrightarrow> (f o M o g) x = (f o N o h) x"
98 lemma type_copy_set_bd: "(\<And>y. |S y| \<le>o bd) \<Longrightarrow> |(S o Rep) x| \<le>o bd"
101 lemma vimage2p_cong: "R = S \<Longrightarrow> vimage2p f g R = vimage2p f g S"
106 assumes type_copy: "type_definition Rep Abs UNIV"
109 lemma type_copy_map_id0: "M = id \<Longrightarrow> Abs o M o Rep = id"
110 using type_definition.Rep_inverse[OF type_copy] by auto
112 lemma type_copy_map_comp0: "M = M1 o M2 \<Longrightarrow> f o M o g = (f o M1 o Rep) o (Abs o M2 o g)"
113 using type_definition.Abs_inverse[OF type_copy UNIV_I] by auto
115 lemma type_copy_set_map0: "S o M = image f o S' \<Longrightarrow> (S o Rep) o (Abs o M o g) = image f o (S' o g)"
116 using type_definition.Abs_inverse[OF type_copy UNIV_I] by (auto simp: o_def fun_eq_iff)
118 lemma type_copy_wit: "x \<in> (S o Rep) (Abs y) \<Longrightarrow> x \<in> S y"
119 using type_definition.Abs_inverse[OF type_copy UNIV_I] by auto
121 lemma type_copy_vimage2p_Grp_Rep: "vimage2p f Rep (Grp (Collect P) h) =
122 Grp (Collect (\<lambda>x. P (f x))) (Abs o h o f)"
123 unfolding vimage2p_def Grp_def fun_eq_iff
124 by (auto simp: type_definition.Abs_inverse[OF type_copy UNIV_I]
125 type_definition.Rep_inverse[OF type_copy] dest: sym)
127 lemma type_copy_vimage2p_Grp_Abs:
128 "\<And>h. vimage2p g Abs (Grp (Collect P) h) = Grp (Collect (\<lambda>x. P (g x))) (Rep o h o g)"
129 unfolding vimage2p_def Grp_def fun_eq_iff
130 by (auto simp: type_definition.Abs_inverse[OF type_copy UNIV_I]
131 type_definition.Rep_inverse[OF type_copy] dest: sym)
133 lemma type_copy_ex_RepI: "(\<exists>b. F b) = (\<exists>b. F (Rep b))"
136 show "\<exists>b'. F (Rep b')"
138 from \<open>F b\<close> show "F (Rep (Abs b))" using type_definition.Abs_inverse[OF type_copy] by auto
142 lemma vimage2p_relcompp_converse:
143 "vimage2p f g (R^--1 OO S) = (vimage2p Rep f R)^--1 OO vimage2p Rep g S"
144 unfolding vimage2p_def relcompp.simps conversep.simps fun_eq_iff image_def
145 by (auto simp: type_copy_ex_RepI)
150 map: "id :: 'a \<Rightarrow> 'a"
152 rel: "op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool"
153 by (auto simp add: Grp_def natLeq_card_order natLeq_cinfinite)
155 definition id_bnf :: "'a \<Rightarrow> 'a" where
156 "id_bnf \<equiv> (\<lambda>x. x)"
158 lemma id_bnf_apply: "id_bnf x = x"
159 unfolding id_bnf_def by simp
162 map: "id_bnf :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
163 sets: "\<lambda>x. {x}"
165 rel: "id_bnf :: ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
167 apply (auto simp: Grp_def fun_eq_iff relcompp.simps natLeq_card_order natLeq_cinfinite)
168 apply (rule ordLess_imp_ordLeq[OF finite_ordLess_infinite[OF _ natLeq_Well_order]])
169 apply (auto simp add: Field_card_of Field_natLeq card_of_well_order_on)[3]
172 lemma type_definition_id_bnf_UNIV: "type_definition id_bnf id_bnf UNIV"
173 unfolding id_bnf_def by unfold_locales auto
175 ML_file "Tools/BNF/bnf_comp_tactics.ML"
176 ML_file "Tools/BNF/bnf_comp.ML"
177 ML_file "Tools/BNF/bnf_lift.ML"
180 DEADID.inj_map DEADID.inj_map_strong DEADID.map_comp DEADID.map_cong DEADID.map_cong0
181 DEADID.map_cong_simp DEADID.map_id DEADID.map_id0 DEADID.map_ident DEADID.map_transfer
182 DEADID.rel_Grp DEADID.rel_compp DEADID.rel_compp_Grp DEADID.rel_conversep DEADID.rel_eq
183 DEADID.rel_flip DEADID.rel_map DEADID.rel_mono DEADID.rel_transfer
184 ID.inj_map ID.inj_map_strong ID.map_comp ID.map_cong ID.map_cong0 ID.map_cong_simp ID.map_id
185 ID.map_id0 ID.map_ident ID.map_transfer ID.rel_Grp ID.rel_compp ID.rel_compp_Grp ID.rel_conversep
186 ID.rel_eq ID.rel_flip ID.rel_map ID.rel_mono ID.rel_transfer ID.set_map ID.set_transfer