| author | blanchet | 
| Thu, 13 Mar 2014 14:48:05 +0100 | |
| changeset 56102 | 439dda276b3f | 
| parent 55945 | e96383acecf9 | 
| child 56237 | 69a9dfe71aed | 
| permissions | -rw-r--r-- | 
| 55059 | 1 | (* Title: HOL/BNF_LFP.thy | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 2 | Author: Dmitriy Traytel, TU Muenchen | 
| 53305 | 3 | Author: Lorenz Panny, TU Muenchen | 
| 4 | Author: Jasmin Blanchette, TU Muenchen | |
| 5 | Copyright 2012, 2013 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 6 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 7 | Least fixed point operation on bounded natural functors. | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 8 | *) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 9 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 10 | header {* Least Fixed Point Operation on Bounded Natural Functors *}
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 11 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 12 | theory BNF_LFP | 
| 53311 | 13 | imports BNF_FP_Base | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 14 | keywords | 
| 53305 | 15 | "datatype_new" :: thy_decl and | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55571diff
changeset | 16 | "datatype_compat" :: thy_decl | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 17 | begin | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 18 | |
| 49312 | 19 | lemma subset_emptyI: "(\<And>x. x \<in> A \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> {}"
 | 
| 20 | by blast | |
| 21 | ||
| 22 | lemma image_Collect_subsetI: | |
| 23 |   "(\<And>x. P x \<Longrightarrow> f x \<in> B) \<Longrightarrow> f ` {x. P x} \<subseteq> B"
 | |
| 24 | by blast | |
| 25 | ||
| 26 | lemma Collect_restrict: "{x. x \<in> X \<and> P x} \<subseteq> X"
 | |
| 27 | by auto | |
| 28 | ||
| 29 | lemma prop_restrict: "\<lbrakk>x \<in> Z; Z \<subseteq> {x. x \<in> X \<and> P x}\<rbrakk> \<Longrightarrow> P x"
 | |
| 30 | by auto | |
| 31 | ||
| 55023 
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
 blanchet parents: 
54841diff
changeset | 32 | lemma underS_I: "\<lbrakk>i \<noteq> j; (i, j) \<in> R\<rbrakk> \<Longrightarrow> i \<in> underS R j" | 
| 
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
 blanchet parents: 
54841diff
changeset | 33 | unfolding underS_def by simp | 
| 49312 | 34 | |
| 55023 
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
 blanchet parents: 
54841diff
changeset | 35 | lemma underS_E: "i \<in> underS R j \<Longrightarrow> i \<noteq> j \<and> (i, j) \<in> R" | 
| 
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
 blanchet parents: 
54841diff
changeset | 36 | unfolding underS_def by simp | 
| 49312 | 37 | |
| 55023 
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
 blanchet parents: 
54841diff
changeset | 38 | lemma underS_Field: "i \<in> underS R j \<Longrightarrow> i \<in> Field R" | 
| 
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
 blanchet parents: 
54841diff
changeset | 39 | unfolding underS_def Field_def by auto | 
| 49312 | 40 | |
| 41 | lemma FieldI2: "(i, j) \<in> R \<Longrightarrow> j \<in> Field R" | |
| 42 | unfolding Field_def by auto | |
| 43 | ||
| 44 | lemma fst_convol': "fst (<f, g> x) = f x" | |
| 45 | using fst_convol unfolding convol_def by simp | |
| 46 | ||
| 47 | lemma snd_convol': "snd (<f, g> x) = g x" | |
| 48 | using snd_convol unfolding convol_def by simp | |
| 49 | ||
| 50 | lemma convol_expand_snd: "fst o f = g \<Longrightarrow> <g, snd o f> = f" | |
| 51 | unfolding convol_def by auto | |
| 52 | ||
| 55811 | 53 | lemma convol_expand_snd': | 
| 54 | assumes "(fst o f = g)" | |
| 55 | shows "h = snd o f \<longleftrightarrow> <g, h> = f" | |
| 56 | proof - | |
| 57 | from assms have *: "<g, snd o f> = f" by (rule convol_expand_snd) | |
| 58 | then have "h = snd o f \<longleftrightarrow> h = snd o <g, snd o f>" by simp | |
| 59 | moreover have "\<dots> \<longleftrightarrow> h = snd o f" by (simp add: snd_convol) | |
| 60 | moreover have "\<dots> \<longleftrightarrow> <g, h> = f" by (subst (2) *[symmetric]) (auto simp: convol_def fun_eq_iff) | |
| 61 | ultimately show ?thesis by simp | |
| 62 | qed | |
| 51739 
3514b90d0a8b
(co)rec is (just as the (un)fold) the unique morphism;
 traytel parents: 
49635diff
changeset | 63 | |
| 49312 | 64 | definition inver where | 
| 65 | "inver g f A = (ALL a : A. g (f a) = a)" | |
| 66 | ||
| 67 | lemma bij_betw_iff_ex: | |
| 68 | "bij_betw f A B = (EX g. g ` B = A \<and> inver g f A \<and> inver f g B)" (is "?L = ?R") | |
| 69 | proof (rule iffI) | |
| 70 | assume ?L | |
| 71 | hence f: "f ` A = B" and inj_f: "inj_on f A" unfolding bij_betw_def by auto | |
| 72 | let ?phi = "% b a. a : A \<and> f a = b" | |
| 73 | have "ALL b : B. EX a. ?phi b a" using f by blast | |
| 74 | then obtain g where g: "ALL b : B. g b : A \<and> f (g b) = b" | |
| 75 | using bchoice[of B ?phi] by blast | |
| 76 | hence gg: "ALL b : f ` A. g b : A \<and> f (g b) = b" using f by blast | |
| 49326 | 77 | have gf: "inver g f A" unfolding inver_def | 
| 55811 | 78 | proof | 
| 79 | fix a assume "a \<in> A" | |
| 80 | then show "g (f a) = a" using the_inv_into_f_f[OF inj_f, of "g (f a)"] | |
| 81 | the_inv_into_f_f[OF inj_f, of a] gg imageI[of a A f] by auto | |
| 82 | qed | |
| 49312 | 83 | moreover have "g ` B \<le> A \<and> inver f g B" using g unfolding inver_def by blast | 
| 84 | moreover have "A \<le> g ` B" | |
| 85 | proof safe | |
| 86 | fix a assume a: "a : A" | |
| 87 | hence "f a : B" using f by auto | |
| 88 | moreover have "a = g (f a)" using a gf unfolding inver_def by auto | |
| 89 | ultimately show "a : g ` B" by blast | |
| 90 | qed | |
| 91 | ultimately show ?R by blast | |
| 92 | next | |
| 93 | assume ?R | |
| 94 | then obtain g where g: "g ` B = A \<and> inver g f A \<and> inver f g B" by blast | |
| 95 | show ?L unfolding bij_betw_def | |
| 96 | proof safe | |
| 97 | show "inj_on f A" unfolding inj_on_def | |
| 98 | proof safe | |
| 99 | fix a1 a2 assume a: "a1 : A" "a2 : A" and "f a1 = f a2" | |
| 100 | hence "g (f a1) = g (f a2)" by simp | |
| 101 | thus "a1 = a2" using a g unfolding inver_def by simp | |
| 102 | qed | |
| 103 | next | |
| 104 | fix a assume "a : A" | |
| 105 | then obtain b where b: "b : B" and a: "a = g b" using g by blast | |
| 106 | hence "b = f (g b)" using g unfolding inver_def by auto | |
| 107 | thus "f a : B" unfolding a using b by simp | |
| 108 | next | |
| 109 | fix b assume "b : B" | |
| 110 | hence "g b : A \<and> b = f (g b)" using g unfolding inver_def by auto | |
| 111 | thus "b : f ` A" by auto | |
| 112 | qed | |
| 113 | qed | |
| 114 | ||
| 115 | lemma bij_betw_ex_weakE: | |
| 116 | "\<lbrakk>bij_betw f A B\<rbrakk> \<Longrightarrow> \<exists>g. g ` B \<subseteq> A \<and> inver g f A \<and> inver f g B" | |
| 117 | by (auto simp only: bij_betw_iff_ex) | |
| 118 | ||
| 119 | lemma inver_surj: "\<lbrakk>g ` B \<subseteq> A; f ` A \<subseteq> B; inver g f A\<rbrakk> \<Longrightarrow> g ` B = A" | |
| 120 | unfolding inver_def by auto (rule rev_image_eqI, auto) | |
| 121 | ||
| 122 | lemma inver_mono: "\<lbrakk>A \<subseteq> B; inver f g B\<rbrakk> \<Longrightarrow> inver f g A" | |
| 123 | unfolding inver_def by auto | |
| 124 | ||
| 125 | lemma inver_pointfree: "inver f g A = (\<forall>a \<in> A. (f o g) a = a)" | |
| 126 | unfolding inver_def by simp | |
| 127 | ||
| 128 | lemma bij_betwE: "bij_betw f A B \<Longrightarrow> \<forall>a\<in>A. f a \<in> B" | |
| 129 | unfolding bij_betw_def by auto | |
| 130 | ||
| 131 | lemma bij_betw_imageE: "bij_betw f A B \<Longrightarrow> f ` A = B" | |
| 132 | unfolding bij_betw_def by auto | |
| 133 | ||
| 134 | lemma inverE: "\<lbrakk>inver f f' A; x \<in> A\<rbrakk> \<Longrightarrow> f (f' x) = x" | |
| 135 | unfolding inver_def by auto | |
| 136 | ||
| 137 | lemma bij_betw_inver1: "bij_betw f A B \<Longrightarrow> inver (inv_into A f) f A" | |
| 138 | unfolding bij_betw_def inver_def by auto | |
| 139 | ||
| 140 | lemma bij_betw_inver2: "bij_betw f A B \<Longrightarrow> inver f (inv_into A f) B" | |
| 141 | unfolding bij_betw_def inver_def by auto | |
| 142 | ||
| 143 | lemma bij_betwI: "\<lbrakk>bij_betw g B A; inver g f A; inver f g B\<rbrakk> \<Longrightarrow> bij_betw f A B" | |
| 49326 | 144 | by (drule bij_betw_imageE, unfold bij_betw_iff_ex) blast | 
| 49312 | 145 | |
| 146 | lemma bij_betwI': | |
| 147 | "\<lbrakk>\<And>x y. \<lbrakk>x \<in> X; y \<in> X\<rbrakk> \<Longrightarrow> (f x = f y) = (x = y); | |
| 148 | \<And>x. x \<in> X \<Longrightarrow> f x \<in> Y; | |
| 149 | \<And>y. y \<in> Y \<Longrightarrow> \<exists>x \<in> X. y = f x\<rbrakk> \<Longrightarrow> bij_betw f X Y" | |
| 53695 | 150 | unfolding bij_betw_def inj_on_def by blast | 
| 49312 | 151 | |
| 152 | lemma surj_fun_eq: | |
| 153 | assumes surj_on: "f ` X = UNIV" and eq_on: "\<forall>x \<in> X. (g1 o f) x = (g2 o f) x" | |
| 154 | shows "g1 = g2" | |
| 155 | proof (rule ext) | |
| 156 | fix y | |
| 157 | from surj_on obtain x where "x \<in> X" and "y = f x" by blast | |
| 158 | thus "g1 y = g2 y" using eq_on by simp | |
| 159 | qed | |
| 160 | ||
| 161 | lemma Card_order_wo_rel: "Card_order r \<Longrightarrow> wo_rel r" | |
| 49514 | 162 | unfolding wo_rel_def card_order_on_def by blast | 
| 49312 | 163 | |
| 164 | lemma Cinfinite_limit: "\<lbrakk>x \<in> Field r; Cinfinite r\<rbrakk> \<Longrightarrow> | |
| 165 | \<exists>y \<in> Field r. x \<noteq> y \<and> (x, y) \<in> r" | |
| 166 | unfolding cinfinite_def by (auto simp add: infinite_Card_order_limit) | |
| 167 | ||
| 168 | lemma Card_order_trans: | |
| 169 | "\<lbrakk>Card_order r; x \<noteq> y; (x, y) \<in> r; y \<noteq> z; (y, z) \<in> r\<rbrakk> \<Longrightarrow> x \<noteq> z \<and> (x, z) \<in> r" | |
| 170 | unfolding card_order_on_def well_order_on_def linear_order_on_def | |
| 171 | partial_order_on_def preorder_on_def trans_def antisym_def by blast | |
| 172 | ||
| 173 | lemma Cinfinite_limit2: | |
| 174 | assumes x1: "x1 \<in> Field r" and x2: "x2 \<in> Field r" and r: "Cinfinite r" | |
| 175 | shows "\<exists>y \<in> Field r. (x1 \<noteq> y \<and> (x1, y) \<in> r) \<and> (x2 \<noteq> y \<and> (x2, y) \<in> r)" | |
| 176 | proof - | |
| 177 | from r have trans: "trans r" and total: "Total r" and antisym: "antisym r" | |
| 178 | unfolding card_order_on_def well_order_on_def linear_order_on_def | |
| 179 | partial_order_on_def preorder_on_def by auto | |
| 180 | obtain y1 where y1: "y1 \<in> Field r" "x1 \<noteq> y1" "(x1, y1) \<in> r" | |
| 181 | using Cinfinite_limit[OF x1 r] by blast | |
| 182 | obtain y2 where y2: "y2 \<in> Field r" "x2 \<noteq> y2" "(x2, y2) \<in> r" | |
| 183 | using Cinfinite_limit[OF x2 r] by blast | |
| 184 | show ?thesis | |
| 185 | proof (cases "y1 = y2") | |
| 186 | case True with y1 y2 show ?thesis by blast | |
| 187 | next | |
| 188 | case False | |
| 189 | with y1(1) y2(1) total have "(y1, y2) \<in> r \<or> (y2, y1) \<in> r" | |
| 190 | unfolding total_on_def by auto | |
| 191 | thus ?thesis | |
| 192 | proof | |
| 193 | assume *: "(y1, y2) \<in> r" | |
| 194 | with trans y1(3) have "(x1, y2) \<in> r" unfolding trans_def by blast | |
| 195 | with False y1 y2 * antisym show ?thesis by (cases "x1 = y2") (auto simp: antisym_def) | |
| 196 | next | |
| 197 | assume *: "(y2, y1) \<in> r" | |
| 198 | with trans y2(3) have "(x2, y1) \<in> r" unfolding trans_def by blast | |
| 199 | with False y1 y2 * antisym show ?thesis by (cases "x2 = y1") (auto simp: antisym_def) | |
| 200 | qed | |
| 201 | qed | |
| 202 | qed | |
| 203 | ||
| 204 | lemma Cinfinite_limit_finite: "\<lbrakk>finite X; X \<subseteq> Field r; Cinfinite r\<rbrakk> | |
| 205 | \<Longrightarrow> \<exists>y \<in> Field r. \<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)" | |
| 206 | proof (induct X rule: finite_induct) | |
| 207 | case empty thus ?case unfolding cinfinite_def using ex_in_conv[of "Field r"] finite.emptyI by auto | |
| 208 | next | |
| 209 | case (insert x X) | |
| 210 | then obtain y where y: "y \<in> Field r" "\<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)" by blast | |
| 211 | then obtain z where z: "z \<in> Field r" "x \<noteq> z \<and> (x, z) \<in> r" "y \<noteq> z \<and> (y, z) \<in> r" | |
| 212 | using Cinfinite_limit2[OF _ y(1) insert(5), of x] insert(4) by blast | |
| 49326 | 213 | show ?case | 
| 214 | apply (intro bexI ballI) | |
| 215 | apply (erule insertE) | |
| 216 | apply hypsubst | |
| 217 | apply (rule z(2)) | |
| 218 | using Card_order_trans[OF insert(5)[THEN conjunct2]] y(2) z(3) | |
| 219 | apply blast | |
| 220 | apply (rule z(1)) | |
| 221 | done | |
| 49312 | 222 | qed | 
| 223 | ||
| 224 | lemma insert_subsetI: "\<lbrakk>x \<in> A; X \<subseteq> A\<rbrakk> \<Longrightarrow> insert x X \<subseteq> A" | |
| 225 | by auto | |
| 226 | ||
| 227 | (*helps resolution*) | |
| 228 | lemma well_order_induct_imp: | |
| 229 | "wo_rel r \<Longrightarrow> (\<And>x. \<forall>y. y \<noteq> x \<and> (y, x) \<in> r \<longrightarrow> y \<in> Field r \<longrightarrow> P y \<Longrightarrow> x \<in> Field r \<longrightarrow> P x) \<Longrightarrow> | |
| 230 | x \<in> Field r \<longrightarrow> P x" | |
| 231 | by (erule wo_rel.well_order_induct) | |
| 232 | ||
| 233 | lemma meta_spec2: | |
| 234 | assumes "(\<And>x y. PROP P x y)" | |
| 235 | shows "PROP P x y" | |
| 55084 | 236 | by (rule assms) | 
| 49312 | 237 | |
| 54841 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 traytel parents: 
54246diff
changeset | 238 | lemma nchotomy_relcomppE: | 
| 55811 | 239 | assumes "\<And>y. \<exists>x. y = f x" "(r OO s) a c" "\<And>b. r a (f b) \<Longrightarrow> s (f b) c \<Longrightarrow> P" | 
| 240 | shows P | |
| 241 | proof (rule relcompp.cases[OF assms(2)], hypsubst) | |
| 242 | fix b assume "r a b" "s b c" | |
| 243 | moreover from assms(1) obtain b' where "b = f b'" by blast | |
| 244 | ultimately show P by (blast intro: assms(3)) | |
| 245 | qed | |
| 54841 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 traytel parents: 
54246diff
changeset | 246 | |
| 55945 | 247 | lemma vimage2p_rel_fun: "rel_fun (vimage2p f g R) R f g" | 
| 248 | unfolding rel_fun_def vimage2p_def by auto | |
| 52731 | 249 | |
| 250 | lemma predicate2D_vimage2p: "\<lbrakk>R \<le> vimage2p f g S; R x y\<rbrakk> \<Longrightarrow> S (f x) (g y)" | |
| 251 | unfolding vimage2p_def by auto | |
| 252 | ||
| 55945 | 253 | lemma id_transfer: "rel_fun A A id id" | 
| 254 | unfolding rel_fun_def by simp | |
| 55084 | 255 | |
| 55770 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 traytel parents: 
55575diff
changeset | 256 | lemma ssubst_Pair_rhs: "\<lbrakk>(r, s) \<in> R; s' = s\<rbrakk> \<Longrightarrow> (r, s') \<in> R" | 
| 55851 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 blanchet parents: 
55811diff
changeset | 257 | by (rule ssubst) | 
| 55770 
f2cf7f92c9ac
intermediate typedef for the type of the bound (local to lfp)
 traytel parents: 
55575diff
changeset | 258 | |
| 55062 | 259 | ML_file "Tools/BNF/bnf_lfp_util.ML" | 
| 260 | ML_file "Tools/BNF/bnf_lfp_tactics.ML" | |
| 261 | ML_file "Tools/BNF/bnf_lfp.ML" | |
| 262 | ML_file "Tools/BNF/bnf_lfp_compat.ML" | |
| 55571 | 263 | ML_file "Tools/BNF/bnf_lfp_rec_sugar_more.ML" | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: 
49308diff
changeset | 264 | |
| 55084 | 265 | hide_fact (open) id_transfer | 
| 266 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 267 | end |