| author | wenzelm | 
| Wed, 01 Aug 2007 16:55:39 +0200 | |
| changeset 24111 | 20e74aa5f56b | 
| parent 23315 | df3a7e9ebadb | 
| child 24124 | 4399175e3014 | 
| permissions | -rw-r--r-- | 
| 19203 | 1 | (* Title: HOL/ZF/HOLZF.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Steven Obua | |
| 4 | ||
| 5 | Axiomatizes the ZFC universe as an HOL type. | |
| 6 | See "Partizan Games in Isabelle/HOLZF", available from http://www4.in.tum.de/~obua/partizan | |
| 7 | *) | |
| 8 | ||
| 9 | theory HOLZF | |
| 10 | imports Helper | |
| 11 | begin | |
| 12 | ||
| 13 | typedecl ZF | |
| 14 | ||
| 22690 
0b08f218f260
replaced axioms/finalconsts by proper axiomatization;
 wenzelm parents: 
20565diff
changeset | 15 | axiomatization | 
| 
0b08f218f260
replaced axioms/finalconsts by proper axiomatization;
 wenzelm parents: 
20565diff
changeset | 16 | Empty :: ZF and | 
| 
0b08f218f260
replaced axioms/finalconsts by proper axiomatization;
 wenzelm parents: 
20565diff
changeset | 17 | Elem :: "ZF \<Rightarrow> ZF \<Rightarrow> bool" and | 
| 
0b08f218f260
replaced axioms/finalconsts by proper axiomatization;
 wenzelm parents: 
20565diff
changeset | 18 | Sum :: "ZF \<Rightarrow> ZF" and | 
| 
0b08f218f260
replaced axioms/finalconsts by proper axiomatization;
 wenzelm parents: 
20565diff
changeset | 19 | Power :: "ZF \<Rightarrow> ZF" and | 
| 
0b08f218f260
replaced axioms/finalconsts by proper axiomatization;
 wenzelm parents: 
20565diff
changeset | 20 | Repl :: "ZF \<Rightarrow> (ZF \<Rightarrow> ZF) \<Rightarrow> ZF" and | 
| 19203 | 21 | Inf :: ZF | 
| 22 | ||
| 23 | constdefs | |
| 24 | Upair:: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" | |
| 25 | "Upair a b == Repl (Power (Power Empty)) (% x. if x = Empty then a else b)" | |
| 26 | Singleton:: "ZF \<Rightarrow> ZF" | |
| 27 | "Singleton x == Upair x x" | |
| 28 | union :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" | |
| 29 | "union A B == Sum (Upair A B)" | |
| 30 | SucNat:: "ZF \<Rightarrow> ZF" | |
| 31 | "SucNat x == union x (Singleton x)" | |
| 32 | subset :: "ZF \<Rightarrow> ZF \<Rightarrow> bool" | |
| 33 | "subset A B == ! x. Elem x A \<longrightarrow> Elem x B" | |
| 34 | ||
| 35 | axioms | |
| 36 | Empty: "Not (Elem x Empty)" | |
| 37 | Ext: "(x = y) = (! z. Elem z x = Elem z y)" | |
| 38 | Sum: "Elem z (Sum x) = (? y. Elem z y & Elem y x)" | |
| 39 | Power: "Elem y (Power x) = (subset y x)" | |
| 40 | Repl: "Elem b (Repl A f) = (? a. Elem a A & b = f a)" | |
| 41 | Regularity: "A \<noteq> Empty \<longrightarrow> (? x. Elem x A & (! y. Elem y x \<longrightarrow> Not (Elem y A)))" | |
| 42 | Infinity: "Elem Empty Inf & (! x. Elem x Inf \<longrightarrow> Elem (SucNat x) Inf)" | |
| 43 | ||
| 44 | constdefs | |
| 45 | Sep:: "ZF \<Rightarrow> (ZF \<Rightarrow> bool) \<Rightarrow> ZF" | |
| 46 | "Sep A p == (if (!x. Elem x A \<longrightarrow> Not (p x)) then Empty else | |
| 47 | (let z = (\<some> x. Elem x A & p x) in | |
| 48 | let f = % x. (if p x then x else z) in Repl A f))" | |
| 49 | ||
| 50 | thm Power[unfolded subset_def] | |
| 51 | ||
| 52 | theorem Sep: "Elem b (Sep A p) = (Elem b A & p b)" | |
| 53 | apply (auto simp add: Sep_def Empty) | |
| 54 | apply (auto simp add: Let_def Repl) | |
| 55 | apply (rule someI2, auto)+ | |
| 56 | done | |
| 57 | ||
| 58 | lemma subset_empty: "subset Empty A" | |
| 59 | by (simp add: subset_def Empty) | |
| 60 | ||
| 61 | theorem Upair: "Elem x (Upair a b) = (x = a | x = b)" | |
| 62 | apply (auto simp add: Upair_def Repl) | |
| 63 | apply (rule exI[where x=Empty]) | |
| 64 | apply (simp add: Power subset_empty) | |
| 65 | apply (rule exI[where x="Power Empty"]) | |
| 66 | apply (auto) | |
| 67 | apply (auto simp add: Ext Power subset_def Empty) | |
| 68 | apply (drule spec[where x=Empty], simp add: Empty)+ | |
| 69 | done | |
| 70 | ||
| 71 | lemma Singleton: "Elem x (Singleton y) = (x = y)" | |
| 72 | by (simp add: Singleton_def Upair) | |
| 73 | ||
| 74 | constdefs | |
| 75 | Opair:: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" | |
| 76 | "Opair a b == Upair (Upair a a) (Upair a b)" | |
| 77 | ||
| 78 | lemma Upair_singleton: "(Upair a a = Upair c d) = (a = c & a = d)" | |
| 79 | by (auto simp add: Ext[where x="Upair a a"] Upair) | |
| 80 | ||
| 81 | lemma Upair_fsteq: "(Upair a b = Upair a c) = ((a = b & a = c) | (b = c))" | |
| 82 | by (auto simp add: Ext[where x="Upair a b"] Upair) | |
| 83 | ||
| 84 | lemma Upair_comm: "Upair a b = Upair b a" | |
| 85 | by (auto simp add: Ext Upair) | |
| 86 | ||
| 87 | theorem Opair: "(Opair a b = Opair c d) = (a = c & b = d)" | |
| 88 | proof - | |
| 89 | have fst: "(Opair a b = Opair c d) \<Longrightarrow> a = c" | |
| 90 | apply (simp add: Opair_def) | |
| 91 | apply (simp add: Ext[where x="Upair (Upair a a) (Upair a b)"]) | |
| 92 | apply (drule spec[where x="Upair a a"]) | |
| 93 | apply (auto simp add: Upair Upair_singleton) | |
| 94 | done | |
| 95 | show ?thesis | |
| 96 | apply (auto) | |
| 97 | apply (erule fst) | |
| 98 | apply (frule fst) | |
| 99 | apply (auto simp add: Opair_def Upair_fsteq) | |
| 100 | done | |
| 101 | qed | |
| 102 | ||
| 103 | constdefs | |
| 104 | Replacement :: "ZF \<Rightarrow> (ZF \<Rightarrow> ZF option) \<Rightarrow> ZF" | |
| 105 | "Replacement A f == Repl (Sep A (% a. f a \<noteq> None)) (the o f)" | |
| 106 | ||
| 107 | theorem Replacement: "Elem y (Replacement A f) = (? x. Elem x A & f x = Some y)" | |
| 108 | by (auto simp add: Replacement_def Repl Sep) | |
| 109 | ||
| 110 | constdefs | |
| 111 | Fst :: "ZF \<Rightarrow> ZF" | |
| 112 | "Fst q == SOME x. ? y. q = Opair x y" | |
| 113 | Snd :: "ZF \<Rightarrow> ZF" | |
| 114 | "Snd q == SOME y. ? x. q = Opair x y" | |
| 115 | ||
| 116 | theorem Fst: "Fst (Opair x y) = x" | |
| 117 | apply (simp add: Fst_def) | |
| 118 | apply (rule someI2) | |
| 119 | apply (simp_all add: Opair) | |
| 120 | done | |
| 121 | ||
| 122 | theorem Snd: "Snd (Opair x y) = y" | |
| 123 | apply (simp add: Snd_def) | |
| 124 | apply (rule someI2) | |
| 125 | apply (simp_all add: Opair) | |
| 126 | done | |
| 127 | ||
| 128 | constdefs | |
| 129 | isOpair :: "ZF \<Rightarrow> bool" | |
| 130 | "isOpair q == ? x y. q = Opair x y" | |
| 131 | ||
| 132 | lemma isOpair: "isOpair (Opair x y) = True" | |
| 133 | by (auto simp add: isOpair_def) | |
| 134 | ||
| 135 | lemma FstSnd: "isOpair x \<Longrightarrow> Opair (Fst x) (Snd x) = x" | |
| 136 | by (auto simp add: isOpair_def Fst Snd) | |
| 137 | ||
| 138 | constdefs | |
| 139 | CartProd :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" | |
| 140 | "CartProd A B == Sum(Repl A (% a. Repl B (% b. Opair a b)))" | |
| 141 | ||
| 142 | lemma CartProd: "Elem x (CartProd A B) = (? a b. Elem a A & Elem b B & x = (Opair a b))" | |
| 143 | apply (auto simp add: CartProd_def Sum Repl) | |
| 144 | apply (rule_tac x="Repl B (Opair a)" in exI) | |
| 145 | apply (auto simp add: Repl) | |
| 146 | done | |
| 147 | ||
| 148 | constdefs | |
| 149 | explode :: "ZF \<Rightarrow> ZF set" | |
| 150 |   "explode z == { x. Elem x z }"
 | |
| 151 | ||
| 152 | lemma explode_Empty: "(explode x = {}) = (x = Empty)"
 | |
| 153 | by (auto simp add: explode_def Ext Empty) | |
| 154 | ||
| 155 | lemma explode_Elem: "(x \<in> explode X) = (Elem x X)" | |
| 156 | by (simp add: explode_def) | |
| 157 | ||
| 158 | lemma Elem_explode_in: "\<lbrakk> Elem a A; explode A \<subseteq> B\<rbrakk> \<Longrightarrow> a \<in> B" | |
| 159 | by (auto simp add: explode_def) | |
| 160 | ||
| 161 | lemma explode_CartProd_eq: "explode (CartProd a b) = (% (x,y). Opair x y) ` ((explode a) \<times> (explode b))" | |
| 162 | by (simp add: explode_def expand_set_eq CartProd image_def) | |
| 163 | ||
| 164 | lemma explode_Repl_eq: "explode (Repl A f) = image f (explode A)" | |
| 165 | by (simp add: explode_def Repl image_def) | |
| 166 | ||
| 167 | constdefs | |
| 168 | Domain :: "ZF \<Rightarrow> ZF" | |
| 169 | "Domain f == Replacement f (% p. if isOpair p then Some (Fst p) else None)" | |
| 170 | Range :: "ZF \<Rightarrow> ZF" | |
| 171 | "Range f == Replacement f (% p. if isOpair p then Some (Snd p) else None)" | |
| 172 | ||
| 173 | theorem Domain: "Elem x (Domain f) = (? y. Elem (Opair x y) f)" | |
| 174 | apply (auto simp add: Domain_def Replacement) | |
| 175 | apply (rule_tac x="Snd x" in exI) | |
| 176 | apply (simp add: FstSnd) | |
| 177 | apply (rule_tac x="Opair x y" in exI) | |
| 178 | apply (simp add: isOpair Fst) | |
| 179 | done | |
| 180 | ||
| 181 | theorem Range: "Elem y (Range f) = (? x. Elem (Opair x y) f)" | |
| 182 | apply (auto simp add: Range_def Replacement) | |
| 183 | apply (rule_tac x="Fst x" in exI) | |
| 184 | apply (simp add: FstSnd) | |
| 185 | apply (rule_tac x="Opair x y" in exI) | |
| 186 | apply (simp add: isOpair Snd) | |
| 187 | done | |
| 188 | ||
| 189 | theorem union: "Elem x (union A B) = (Elem x A | Elem x B)" | |
| 190 | by (auto simp add: union_def Sum Upair) | |
| 191 | ||
| 192 | constdefs | |
| 193 | Field :: "ZF \<Rightarrow> ZF" | |
| 194 | "Field A == union (Domain A) (Range A)" | |
| 195 | ||
| 196 | constdefs | |
| 197 |   "\<acute>"         :: "ZF \<Rightarrow> ZF => ZF"    (infixl 90) --{*function application*} 
 | |
| 198 | app_def: "f \<acute> x == (THE y. Elem (Opair x y) f)" | |
| 199 | ||
| 200 | constdefs | |
| 201 | isFun :: "ZF \<Rightarrow> bool" | |
| 202 | "isFun f == (! x y1 y2. Elem (Opair x y1) f & Elem (Opair x y2) f \<longrightarrow> y1 = y2)" | |
| 203 | ||
| 204 | constdefs | |
| 205 | Lambda :: "ZF \<Rightarrow> (ZF \<Rightarrow> ZF) \<Rightarrow> ZF" | |
| 206 | "Lambda A f == Repl A (% x. Opair x (f x))" | |
| 207 | ||
| 208 | lemma Lambda_app: "Elem x A \<Longrightarrow> (Lambda A f)\<acute>x = f x" | |
| 209 | by (simp add: app_def Lambda_def Repl Opair) | |
| 210 | ||
| 211 | lemma isFun_Lambda: "isFun (Lambda A f)" | |
| 212 | by (auto simp add: isFun_def Lambda_def Repl Opair) | |
| 213 | ||
| 214 | lemma domain_Lambda: "Domain (Lambda A f) = A" | |
| 215 | apply (auto simp add: Domain_def) | |
| 216 | apply (subst Ext) | |
| 217 | apply (auto simp add: Replacement) | |
| 218 | apply (simp add: Lambda_def Repl) | |
| 219 | apply (auto simp add: Fst) | |
| 220 | apply (simp add: Lambda_def Repl) | |
| 221 | apply (rule_tac x="Opair z (f z)" in exI) | |
| 222 | apply (auto simp add: Fst isOpair_def) | |
| 223 | done | |
| 224 | ||
| 225 | lemma Lambda_ext: "(Lambda s f = Lambda t g) = (s = t & (! x. Elem x s \<longrightarrow> f x = g x))" | |
| 226 | proof - | |
| 227 | have "Lambda s f = Lambda t g \<Longrightarrow> s = t" | |
| 228 | apply (subst domain_Lambda[where A = s and f = f, symmetric]) | |
| 229 | apply (subst domain_Lambda[where A = t and f = g, symmetric]) | |
| 230 | apply auto | |
| 231 | done | |
| 232 | then show ?thesis | |
| 233 | apply auto | |
| 234 | apply (subst Lambda_app[where f=f, symmetric], simp) | |
| 235 | apply (subst Lambda_app[where f=g, symmetric], simp) | |
| 236 | apply auto | |
| 237 | apply (auto simp add: Lambda_def Repl Ext) | |
| 238 | apply (auto simp add: Ext[symmetric]) | |
| 239 | done | |
| 240 | qed | |
| 241 | ||
| 242 | constdefs | |
| 243 | PFun :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" | |
| 244 | "PFun A B == Sep (Power (CartProd A B)) isFun" | |
| 245 | Fun :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" | |
| 246 | "Fun A B == Sep (PFun A B) (\<lambda> f. Domain f = A)" | |
| 247 | ||
| 248 | lemma Fun_Range: "Elem f (Fun U V) \<Longrightarrow> subset (Range f) V" | |
| 249 | apply (simp add: Fun_def Sep PFun_def Power subset_def CartProd) | |
| 250 | apply (auto simp add: Domain Range) | |
| 251 | apply (erule_tac x="Opair xa x" in allE) | |
| 252 | apply (auto simp add: Opair) | |
| 253 | done | |
| 254 | ||
| 255 | lemma Elem_Elem_PFun: "Elem F (PFun U V) \<Longrightarrow> Elem p F \<Longrightarrow> isOpair p & Elem (Fst p) U & Elem (Snd p) V" | |
| 256 | apply (simp add: PFun_def Sep Power subset_def, clarify) | |
| 257 | apply (erule_tac x=p in allE) | |
| 258 | apply (auto simp add: CartProd isOpair Fst Snd) | |
| 259 | done | |
| 260 | ||
| 261 | lemma Fun_implies_PFun[simp]: "Elem f (Fun U V) \<Longrightarrow> Elem f (PFun U V)" | |
| 262 | by (simp add: Fun_def Sep) | |
| 263 | ||
| 264 | lemma Elem_Elem_Fun: "Elem F (Fun U V) \<Longrightarrow> Elem p F \<Longrightarrow> isOpair p & Elem (Fst p) U & Elem (Snd p) V" | |
| 265 | by (auto simp add: Elem_Elem_PFun dest: Fun_implies_PFun) | |
| 266 | ||
| 267 | lemma PFun_inj: "Elem F (PFun U V) \<Longrightarrow> Elem x F \<Longrightarrow> Elem y F \<Longrightarrow> Fst x = Fst y \<Longrightarrow> Snd x = Snd y" | |
| 268 | apply (frule Elem_Elem_PFun[where p=x], simp) | |
| 269 | apply (frule Elem_Elem_PFun[where p=y], simp) | |
| 270 | apply (subgoal_tac "isFun F") | |
| 271 | apply (simp add: isFun_def isOpair_def) | |
| 272 | apply (auto simp add: Fst Snd, blast) | |
| 273 | apply (auto simp add: PFun_def Sep) | |
| 274 | done | |
| 275 | ||
| 276 | ML {* simp_depth_limit := 2 *}
 | |
| 277 | lemma Fun_total: "\<lbrakk>Elem F (Fun U V); Elem a U\<rbrakk> \<Longrightarrow> \<exists>x. Elem (Opair a x) F" | |
| 278 | by (auto simp add: Fun_def Sep Domain) | |
| 279 | ML {* simp_depth_limit := 1000 *}
 | |
| 280 | ||
| 281 | ||
| 282 | lemma unique_fun_value: "\<lbrakk>isFun f; Elem x (Domain f)\<rbrakk> \<Longrightarrow> ?! y. Elem (Opair x y) f" | |
| 283 | by (auto simp add: Domain isFun_def) | |
| 284 | ||
| 285 | lemma fun_value_in_range: "\<lbrakk>isFun f; Elem x (Domain f)\<rbrakk> \<Longrightarrow> Elem (f\<acute>x) (Range f)" | |
| 286 | apply (auto simp add: Range) | |
| 287 | apply (drule unique_fun_value) | |
| 288 | apply simp | |
| 289 | apply (simp add: app_def) | |
| 290 | apply (rule exI[where x=x]) | |
| 291 | apply (auto simp add: the_equality) | |
| 292 | done | |
| 293 | ||
| 294 | lemma fun_range_witness: "\<lbrakk>isFun f; Elem y (Range f)\<rbrakk> \<Longrightarrow> ? x. Elem x (Domain f) & f\<acute>x = y" | |
| 295 | apply (auto simp add: Range) | |
| 296 | apply (rule_tac x="x" in exI) | |
| 297 | apply (auto simp add: app_def the_equality isFun_def Domain) | |
| 298 | done | |
| 299 | ||
| 300 | lemma Elem_Fun_Lambda: "Elem F (Fun U V) \<Longrightarrow> ? f. F = Lambda U f" | |
| 301 | apply (rule exI[where x= "% x. (THE y. Elem (Opair x y) F)"]) | |
| 302 | apply (simp add: Ext Lambda_def Repl Domain) | |
| 303 | apply (simp add: Ext[symmetric]) | |
| 304 | apply auto | |
| 305 | apply (frule Elem_Elem_Fun) | |
| 306 | apply auto | |
| 307 | apply (rule_tac x="Fst z" in exI) | |
| 308 | apply (simp add: isOpair_def) | |
| 309 | apply (auto simp add: Fst Snd Opair) | |
| 310 | apply (rule theI2') | |
| 311 | apply auto | |
| 312 | apply (drule Fun_implies_PFun) | |
| 313 | apply (drule_tac x="Opair x ya" and y="Opair x yb" in PFun_inj) | |
| 314 | apply (auto simp add: Fst Snd) | |
| 315 | apply (drule Fun_implies_PFun) | |
| 316 | apply (drule_tac x="Opair x y" and y="Opair x ya" in PFun_inj) | |
| 317 | apply (auto simp add: Fst Snd) | |
| 318 | apply (rule theI2') | |
| 319 | apply (auto simp add: Fun_total) | |
| 320 | apply (drule Fun_implies_PFun) | |
| 321 | apply (drule_tac x="Opair a x" and y="Opair a y" in PFun_inj) | |
| 322 | apply (auto simp add: Fst Snd) | |
| 323 | done | |
| 324 | ||
| 325 | lemma Elem_Lambda_Fun: "Elem (Lambda A f) (Fun U V) = (A = U & (! x. Elem x A \<longrightarrow> Elem (f x) V))" | |
| 326 | proof - | |
| 327 | have "Elem (Lambda A f) (Fun U V) \<Longrightarrow> A = U" | |
| 328 | by (simp add: Fun_def Sep domain_Lambda) | |
| 329 | then show ?thesis | |
| 330 | apply auto | |
| 331 | apply (drule Fun_Range) | |
| 332 | apply (subgoal_tac "f x = ((Lambda U f) \<acute> x)") | |
| 333 | prefer 2 | |
| 334 | apply (simp add: Lambda_app) | |
| 335 | apply simp | |
| 336 | apply (subgoal_tac "Elem (Lambda U f \<acute> x) (Range (Lambda U f))") | |
| 337 | apply (simp add: subset_def) | |
| 338 | apply (rule fun_value_in_range) | |
| 339 | apply (simp_all add: isFun_Lambda domain_Lambda) | |
| 340 | apply (simp add: Fun_def Sep PFun_def Power domain_Lambda isFun_Lambda) | |
| 341 | apply (auto simp add: subset_def CartProd) | |
| 342 | apply (rule_tac x="Fst x" in exI) | |
| 343 | apply (auto simp add: Lambda_def Repl Fst) | |
| 344 | done | |
| 345 | qed | |
| 346 | ||
| 347 | ||
| 348 | constdefs | |
| 349 | is_Elem_of :: "(ZF * ZF) set" | |
| 350 |   "is_Elem_of == { (a,b) | a b. Elem a b }"
 | |
| 351 | ||
| 352 | lemma cond_wf_Elem: | |
| 353 | assumes hyps:"\<forall>x. (\<forall>y. Elem y x \<longrightarrow> Elem y U \<longrightarrow> P y) \<longrightarrow> Elem x U \<longrightarrow> P x" "Elem a U" | |
| 354 | shows "P a" | |
| 355 | proof - | |
| 356 |   {
 | |
| 357 | fix P | |
| 358 | fix U | |
| 359 | fix a | |
| 360 | assume P_induct: "(\<forall>x. (\<forall>y. Elem y x \<longrightarrow> Elem y U \<longrightarrow> P y) \<longrightarrow> (Elem x U \<longrightarrow> P x))" | |
| 361 | assume a_in_U: "Elem a U" | |
| 362 | have "P a" | |
| 363 | proof - | |
| 364 | term "P" | |
| 365 | term Sep | |
| 366 | let ?Z = "Sep U (Not o P)" | |
| 367 | have "?Z = Empty \<longrightarrow> P a" by (simp add: Ext Sep Empty a_in_U) | |
| 368 | moreover have "?Z \<noteq> Empty \<longrightarrow> False" | |
| 369 | proof | |
| 370 | assume not_empty: "?Z \<noteq> Empty" | |
| 371 | note thereis_x = Regularity[where A="?Z", simplified not_empty, simplified] | |
| 372 | then obtain x where x_def: "Elem x ?Z & (! y. Elem y x \<longrightarrow> Not (Elem y ?Z))" .. | |
| 373 | then have x_induct:"! y. Elem y x \<longrightarrow> Elem y U \<longrightarrow> P y" by (simp add: Sep) | |
| 374 | have "Elem x U \<longrightarrow> P x" | |
| 375 | by (rule impE[OF spec[OF P_induct, where x=x], OF x_induct], assumption) | |
| 376 | moreover have "Elem x U & Not(P x)" | |
| 377 | apply (insert x_def) | |
| 378 | apply (simp add: Sep) | |
| 379 | done | |
| 380 | ultimately show "False" by auto | |
| 381 | qed | |
| 382 | ultimately show "P a" by auto | |
| 383 | qed | |
| 384 | } | |
| 385 | with hyps show ?thesis by blast | |
| 386 | qed | |
| 387 | ||
| 388 | lemma cond2_wf_Elem: | |
| 389 | assumes | |
| 390 | special_P: "? U. ! x. Not(Elem x U) \<longrightarrow> (P x)" | |
| 391 | and P_induct: "\<forall>x. (\<forall>y. Elem y x \<longrightarrow> P y) \<longrightarrow> P x" | |
| 392 | shows | |
| 393 | "P a" | |
| 394 | proof - | |
| 395 | have "? U Q. P = (\<lambda> x. (Elem x U \<longrightarrow> Q x))" | |
| 396 | proof - | |
| 397 | from special_P obtain U where U:"! x. Not(Elem x U) \<longrightarrow> (P x)" .. | |
| 398 | show ?thesis | |
| 399 | apply (rule_tac exI[where x=U]) | |
| 400 | apply (rule exI[where x="P"]) | |
| 401 | apply (rule ext) | |
| 402 | apply (auto simp add: U) | |
| 403 | done | |
| 404 | qed | |
| 405 | then obtain U where "? Q. P = (\<lambda> x. (Elem x U \<longrightarrow> Q x))" .. | |
| 406 | then obtain Q where UQ: "P = (\<lambda> x. (Elem x U \<longrightarrow> Q x))" .. | |
| 407 | show ?thesis | |
| 408 | apply (auto simp add: UQ) | |
| 409 | apply (rule cond_wf_Elem) | |
| 410 | apply (rule P_induct[simplified UQ]) | |
| 411 | apply simp | |
| 412 | done | |
| 413 | qed | |
| 414 | ||
| 415 | consts | |
| 416 | nat2Nat :: "nat \<Rightarrow> ZF" | |
| 417 | ||
| 418 | primrec | |
| 419 | nat2Nat_0[intro]: "nat2Nat 0 = Empty" | |
| 420 | nat2Nat_Suc[intro]: "nat2Nat (Suc n) = SucNat (nat2Nat n)" | |
| 421 | ||
| 422 | constdefs | |
| 423 | Nat2nat :: "ZF \<Rightarrow> nat" | |
| 424 | "Nat2nat == inv nat2Nat" | |
| 425 | ||
| 426 | lemma Elem_nat2Nat_inf[intro]: "Elem (nat2Nat n) Inf" | |
| 427 | apply (induct n) | |
| 428 | apply (simp_all add: Infinity) | |
| 429 | done | |
| 430 | ||
| 431 | constdefs | |
| 432 | Nat :: ZF | |
| 433 | "Nat == Sep Inf (\<lambda> N. ? n. nat2Nat n = N)" | |
| 434 | ||
| 435 | lemma Elem_nat2Nat_Nat[intro]: "Elem (nat2Nat n) Nat" | |
| 436 | by (auto simp add: Nat_def Sep) | |
| 437 | ||
| 438 | lemma Elem_Empty_Nat: "Elem Empty Nat" | |
| 439 | by (auto simp add: Nat_def Sep Infinity) | |
| 440 | ||
| 441 | lemma Elem_SucNat_Nat: "Elem N Nat \<Longrightarrow> Elem (SucNat N) Nat" | |
| 442 | by (auto simp add: Nat_def Sep Infinity) | |
| 443 | ||
| 444 | lemma no_infinite_Elem_down_chain: | |
| 445 | "Not (? f. isFun f & Domain f = Nat & (! N. Elem N Nat \<longrightarrow> Elem (f\<acute>(SucNat N)) (f\<acute>N)))" | |
| 446 | proof - | |
| 447 |   {
 | |
| 448 | fix f | |
| 449 | assume f:"isFun f & Domain f = Nat & (! N. Elem N Nat \<longrightarrow> Elem (f\<acute>(SucNat N)) (f\<acute>N))" | |
| 450 | let ?r = "Range f" | |
| 451 | have "?r \<noteq> Empty" | |
| 452 | apply (auto simp add: Ext Empty) | |
| 453 | apply (rule exI[where x="f\<acute>Empty"]) | |
| 454 | apply (rule fun_value_in_range) | |
| 455 | apply (auto simp add: f Elem_Empty_Nat) | |
| 456 | done | |
| 457 | then have "? x. Elem x ?r & (! y. Elem y x \<longrightarrow> Not(Elem y ?r))" | |
| 458 | by (simp add: Regularity) | |
| 459 | then obtain x where x: "Elem x ?r & (! y. Elem y x \<longrightarrow> Not(Elem y ?r))" .. | |
| 460 | then have "? N. Elem N (Domain f) & f\<acute>N = x" | |
| 461 | apply (rule_tac fun_range_witness) | |
| 462 | apply (simp_all add: f) | |
| 463 | done | |
| 464 | then have "? N. Elem N Nat & f\<acute>N = x" | |
| 465 | by (simp add: f) | |
| 466 | then obtain N where N: "Elem N Nat & f\<acute>N = x" .. | |
| 467 | from N have N': "Elem N Nat" by auto | |
| 468 | let ?y = "f\<acute>(SucNat N)" | |
| 469 | have Elem_y_r: "Elem ?y ?r" | |
| 470 | by (simp_all add: f Elem_SucNat_Nat N fun_value_in_range) | |
| 471 | have "Elem ?y (f\<acute>N)" by (auto simp add: f N') | |
| 472 | then have "Elem ?y x" by (simp add: N) | |
| 473 | with x have "Not (Elem ?y ?r)" by auto | |
| 474 | with Elem_y_r have "False" by auto | |
| 475 | } | |
| 476 | then show ?thesis by auto | |
| 477 | qed | |
| 478 | ||
| 479 | lemma Upair_nonEmpty: "Upair a b \<noteq> Empty" | |
| 480 | by (auto simp add: Ext Empty Upair) | |
| 481 | ||
| 482 | lemma Singleton_nonEmpty: "Singleton x \<noteq> Empty" | |
| 483 | by (auto simp add: Singleton_def Upair_nonEmpty) | |
| 484 | ||
| 485 | lemma antisym_Elem: "Not(Elem a b & Elem b a)" | |
| 486 | proof - | |
| 487 |   {
 | |
| 488 | fix a b | |
| 489 | assume ab: "Elem a b" | |
| 490 | assume ba: "Elem b a" | |
| 491 | let ?Z = "Upair a b" | |
| 492 | have "?Z \<noteq> Empty" by (simp add: Upair_nonEmpty) | |
| 493 | then have "? x. Elem x ?Z & (! y. Elem y x \<longrightarrow> Not(Elem y ?Z))" | |
| 494 | by (simp add: Regularity) | |
| 495 | then obtain x where x:"Elem x ?Z & (! y. Elem y x \<longrightarrow> Not(Elem y ?Z))" .. | |
| 496 | then have "x = a \<or> x = b" by (simp add: Upair) | |
| 497 | moreover have "x = a \<longrightarrow> Not (Elem b ?Z)" | |
| 498 | by (auto simp add: x ba) | |
| 499 | moreover have "x = b \<longrightarrow> Not (Elem a ?Z)" | |
| 500 | by (auto simp add: x ab) | |
| 501 | ultimately have "False" | |
| 502 | by (auto simp add: Upair) | |
| 503 | } | |
| 504 | then show ?thesis by auto | |
| 505 | qed | |
| 506 | ||
| 507 | lemma irreflexiv_Elem: "Not(Elem a a)" | |
| 508 | by (simp add: antisym_Elem[of a a, simplified]) | |
| 509 | ||
| 510 | lemma antisym_Elem: "Elem a b \<Longrightarrow> Not (Elem b a)" | |
| 511 | apply (insert antisym_Elem[of a b]) | |
| 512 | apply auto | |
| 513 | done | |
| 514 | ||
| 515 | consts | |
| 516 | NatInterval :: "nat \<Rightarrow> nat \<Rightarrow> ZF" | |
| 517 | ||
| 518 | primrec | |
| 519 | "NatInterval n 0 = Singleton (nat2Nat n)" | |
| 520 | "NatInterval n (Suc m) = union (NatInterval n m) (Singleton (nat2Nat (n+m+1)))" | |
| 521 | ||
| 522 | lemma n_Elem_NatInterval[rule_format]: "! q. q <= m \<longrightarrow> Elem (nat2Nat (n+q)) (NatInterval n m)" | |
| 523 | apply (induct m) | |
| 524 | apply (auto simp add: Singleton union) | |
| 525 | apply (case_tac "q <= m") | |
| 526 | apply auto | |
| 527 | apply (subgoal_tac "q = Suc m") | |
| 528 | apply auto | |
| 529 | done | |
| 530 | ||
| 531 | lemma NatInterval_not_Empty: "NatInterval n m \<noteq> Empty" | |
| 532 | by (auto intro: n_Elem_NatInterval[where q = 0, simplified] simp add: Empty Ext) | |
| 533 | ||
| 534 | lemma increasing_nat2Nat[rule_format]: "0 < n \<longrightarrow> Elem (nat2Nat (n - 1)) (nat2Nat n)" | |
| 535 | apply (case_tac "? m. n = Suc m") | |
| 536 | apply (auto simp add: SucNat_def union Singleton) | |
| 537 | apply (drule spec[where x="n - 1"]) | |
| 538 | apply arith | |
| 539 | done | |
| 540 | ||
| 541 | lemma represent_NatInterval[rule_format]: "Elem x (NatInterval n m) \<longrightarrow> (? u. n \<le> u & u \<le> n+m & nat2Nat u = x)" | |
| 542 | apply (induct m) | |
| 543 | apply (auto simp add: Singleton union) | |
| 544 | apply (rule_tac x="Suc (n+m)" in exI) | |
| 545 | apply auto | |
| 546 | done | |
| 547 | ||
| 548 | lemma inj_nat2Nat: "inj nat2Nat" | |
| 549 | proof - | |
| 550 |   {
 | |
| 551 | fix n m :: nat | |
| 552 | assume nm: "nat2Nat n = nat2Nat (n+m)" | |
| 553 | assume mg0: "0 < m" | |
| 554 | let ?Z = "NatInterval n m" | |
| 555 | have "?Z \<noteq> Empty" by (simp add: NatInterval_not_Empty) | |
| 556 | then have "? x. (Elem x ?Z) & (! y. Elem y x \<longrightarrow> Not (Elem y ?Z))" | |
| 557 | by (auto simp add: Regularity) | |
| 558 | then obtain x where x:"Elem x ?Z & (! y. Elem y x \<longrightarrow> Not (Elem y ?Z))" .. | |
| 559 | then have "? u. n \<le> u & u \<le> n+m & nat2Nat u = x" | |
| 560 | by (simp add: represent_NatInterval) | |
| 561 | then obtain u where u: "n \<le> u & u \<le> n+m & nat2Nat u = x" .. | |
| 562 | have "n < u \<longrightarrow> False" | |
| 563 | proof | |
| 564 | assume n_less_u: "n < u" | |
| 565 | let ?y = "nat2Nat (u - 1)" | |
| 566 | have "Elem ?y (nat2Nat u)" | |
| 567 | apply (rule increasing_nat2Nat) | |
| 568 | apply (insert n_less_u) | |
| 569 | apply arith | |
| 570 | done | |
| 571 | with u have "Elem ?y x" by auto | |
| 572 | with x have "Not (Elem ?y ?Z)" by auto | |
| 573 | moreover have "Elem ?y ?Z" | |
| 574 | apply (insert n_Elem_NatInterval[where q = "u - n - 1" and n=n and m=m]) | |
| 575 | apply (insert n_less_u) | |
| 576 | apply (insert u) | |
| 577 | apply auto | |
| 578 | done | |
| 579 | ultimately show False by auto | |
| 580 | qed | |
| 581 | moreover have "u = n \<longrightarrow> False" | |
| 582 | proof | |
| 583 | assume "u = n" | |
| 584 | with u have "nat2Nat n = x" by auto | |
| 585 | then have nm_eq_x: "nat2Nat (n+m) = x" by (simp add: nm) | |
| 586 | let ?y = "nat2Nat (n+m - 1)" | |
| 587 | have "Elem ?y (nat2Nat (n+m))" | |
| 588 | apply (rule increasing_nat2Nat) | |
| 589 | apply (insert mg0) | |
| 590 | apply arith | |
| 591 | done | |
| 592 | with nm_eq_x have "Elem ?y x" by auto | |
| 593 | with x have "Not (Elem ?y ?Z)" by auto | |
| 594 | moreover have "Elem ?y ?Z" | |
| 595 | apply (insert n_Elem_NatInterval[where q = "m - 1" and n=n and m=m]) | |
| 596 | apply (insert mg0) | |
| 597 | apply auto | |
| 598 | done | |
| 599 | ultimately show False by auto | |
| 600 | qed | |
| 601 | ultimately have "False" using u by arith | |
| 602 | } | |
| 603 | note lemma_nat2Nat = this | |
| 23315 | 604 | have th:"\<And>x y. \<not> (x < y \<and> (\<forall>(m\<Colon>nat). y \<noteq> x + m))" by presburger | 
| 605 | have th': "\<And>x y. \<not> (x \<noteq> y \<and> (\<not> x < y) \<and> (\<forall>(m\<Colon>nat). x \<noteq> y + m))" by presburger | |
| 19203 | 606 | show ?thesis | 
| 607 | apply (auto simp add: inj_on_def) | |
| 608 | apply (case_tac "x = y") | |
| 609 | apply auto | |
| 610 | apply (case_tac "x < y") | |
| 611 | apply (case_tac "? m. y = x + m & 0 < m") | |
| 23315 | 612 | apply (auto intro: lemma_nat2Nat) | 
| 19203 | 613 | apply (case_tac "y < x") | 
| 614 | apply (case_tac "? m. x = y + m & 0 < m") | |
| 23315 | 615 | apply simp | 
| 616 | apply simp | |
| 617 | using th apply blast | |
| 618 | apply (case_tac "? m. x = y + m") | |
| 619 | apply (auto intro: lemma_nat2Nat) | |
| 19203 | 620 | apply (drule sym) | 
| 23315 | 621 | using lemma_nat2Nat apply blast | 
| 622 | using th' apply blast | |
| 19203 | 623 | done | 
| 624 | qed | |
| 625 | ||
| 626 | lemma Nat2nat_nat2Nat[simp]: "Nat2nat (nat2Nat n) = n" | |
| 627 | by (simp add: Nat2nat_def inv_f_f[OF inj_nat2Nat]) | |
| 628 | ||
| 629 | lemma nat2Nat_Nat2nat[simp]: "Elem n Nat \<Longrightarrow> nat2Nat (Nat2nat n) = n" | |
| 630 | apply (simp add: Nat2nat_def) | |
| 631 | apply (rule_tac f_inv_f) | |
| 632 | apply (auto simp add: image_def Nat_def Sep) | |
| 633 | done | |
| 634 | ||
| 635 | lemma Nat2nat_SucNat: "Elem N Nat \<Longrightarrow> Nat2nat (SucNat N) = Suc (Nat2nat N)" | |
| 636 | apply (auto simp add: Nat_def Sep Nat2nat_def) | |
| 637 | apply (auto simp add: inv_f_f[OF inj_nat2Nat]) | |
| 638 | apply (simp only: nat2Nat.simps[symmetric]) | |
| 639 | apply (simp only: inv_f_f[OF inj_nat2Nat]) | |
| 640 | done | |
| 641 | ||
| 642 | ||
| 643 | (*lemma Elem_induct: "(\<And>x. \<forall>y. Elem y x \<longrightarrow> P y \<Longrightarrow> P x) \<Longrightarrow> P a" | |
| 644 | by (erule wf_induct[OF wf_is_Elem_of, simplified is_Elem_of_def, simplified])*) | |
| 645 | ||
| 646 | lemma Elem_Opair_exists: "? z. Elem x z & Elem y z & Elem z (Opair x y)" | |
| 647 | apply (rule exI[where x="Upair x y"]) | |
| 648 | by (simp add: Upair Opair_def) | |
| 649 | ||
| 650 | lemma UNIV_is_not_in_ZF: "UNIV \<noteq> explode R" | |
| 651 | proof | |
| 652 |   let ?Russell = "{ x. Not(Elem x x) }"
 | |
| 653 | have "?Russell = UNIV" by (simp add: irreflexiv_Elem) | |
| 654 | moreover assume "UNIV = explode R" | |
| 655 | ultimately have russell: "?Russell = explode R" by simp | |
| 656 | then show "False" | |
| 657 | proof(cases "Elem R R") | |
| 658 | case True | |
| 659 | then show ?thesis | |
| 660 | by (insert irreflexiv_Elem, auto) | |
| 661 | next | |
| 662 | case False | |
| 663 | then have "R \<in> ?Russell" by auto | |
| 664 | then have "Elem R R" by (simp add: russell explode_def) | |
| 665 | with False show ?thesis by auto | |
| 666 | qed | |
| 667 | qed | |
| 668 | ||
| 669 | constdefs | |
| 670 | SpecialR :: "(ZF * ZF) set" | |
| 671 |   "SpecialR \<equiv> { (x, y) . x \<noteq> Empty \<and> y = Empty}"
 | |
| 672 | ||
| 673 | lemma "wf SpecialR" | |
| 674 | apply (subst wf_def) | |
| 675 | apply (auto simp add: SpecialR_def) | |
| 676 | done | |
| 677 | ||
| 678 | constdefs | |
| 679 |   Ext :: "('a * 'b) set \<Rightarrow> 'b \<Rightarrow> 'a set"
 | |
| 680 |   "Ext R y \<equiv> { x . (x, y) \<in> R }" 
 | |
| 681 | ||
| 682 | lemma Ext_Elem: "Ext is_Elem_of = explode" | |
| 683 | by (auto intro: ext simp add: Ext_def is_Elem_of_def explode_def) | |
| 684 | ||
| 685 | lemma "Ext SpecialR Empty \<noteq> explode z" | |
| 686 | proof | |
| 687 |   have "Ext SpecialR Empty = UNIV - {Empty}"
 | |
| 688 | by (auto simp add: Ext_def SpecialR_def) | |
| 689 | moreover assume "Ext SpecialR Empty = explode z" | |
| 690 | ultimately have "UNIV = explode(union z (Singleton Empty)) " | |
| 691 | by (auto simp add: explode_def union Singleton) | |
| 692 | then show "False" by (simp add: UNIV_is_not_in_ZF) | |
| 693 | qed | |
| 694 | ||
| 695 | constdefs | |
| 696 | implode :: "ZF set \<Rightarrow> ZF" | |
| 697 | "implode == inv explode" | |
| 698 | ||
| 699 | lemma inj_explode: "inj explode" | |
| 700 | by (auto simp add: inj_on_def explode_def Ext) | |
| 701 | ||
| 702 | lemma implode_explode[simp]: "implode (explode x) = x" | |
| 703 | by (simp add: implode_def inj_explode) | |
| 704 | ||
| 705 | constdefs | |
| 706 | regular :: "(ZF * ZF) set \<Rightarrow> bool" | |
| 707 | "regular R == ! A. A \<noteq> Empty \<longrightarrow> (? x. Elem x A & (! y. (y, x) \<in> R \<longrightarrow> Not (Elem y A)))" | |
| 20565 | 708 | set_like :: "(ZF * ZF) set \<Rightarrow> bool" | 
| 709 | "set_like R == ! y. Ext R y \<in> range explode" | |
| 19203 | 710 | wfzf :: "(ZF * ZF) set \<Rightarrow> bool" | 
| 20565 | 711 | "wfzf R == regular R & set_like R" | 
| 19203 | 712 | |
| 713 | lemma regular_Elem: "regular is_Elem_of" | |
| 714 | by (simp add: regular_def is_Elem_of_def Regularity) | |
| 715 | ||
| 20565 | 716 | lemma set_like_Elem: "set_like is_Elem_of" | 
| 717 | by (auto simp add: set_like_def image_def Ext_Elem) | |
| 19203 | 718 | |
| 719 | lemma wfzf_is_Elem_of: "wfzf is_Elem_of" | |
| 20565 | 720 | by (auto simp add: wfzf_def regular_Elem set_like_Elem) | 
| 19203 | 721 | |
| 722 | constdefs | |
| 723 | SeqSum :: "(nat \<Rightarrow> ZF) \<Rightarrow> ZF" | |
| 724 | "SeqSum f == Sum (Repl Nat (f o Nat2nat))" | |
| 725 | ||
| 726 | lemma SeqSum: "Elem x (SeqSum f) = (? n. Elem x (f n))" | |
| 727 | apply (auto simp add: SeqSum_def Sum Repl) | |
| 728 | apply (rule_tac x = "f n" in exI) | |
| 729 | apply auto | |
| 730 | done | |
| 731 | ||
| 732 | constdefs | |
| 733 | Ext_ZF :: "(ZF * ZF) set \<Rightarrow> ZF \<Rightarrow> ZF" | |
| 734 | "Ext_ZF R s == implode (Ext R s)" | |
| 735 | ||
| 736 | lemma Elem_implode: "A \<in> range explode \<Longrightarrow> Elem x (implode A) = (x \<in> A)" | |
| 737 | apply (auto) | |
| 738 | apply (simp_all add: explode_def) | |
| 739 | done | |
| 740 | ||
| 20565 | 741 | lemma Elem_Ext_ZF: "set_like R \<Longrightarrow> Elem x (Ext_ZF R s) = ((x,s) \<in> R)" | 
| 19203 | 742 | apply (simp add: Ext_ZF_def) | 
| 743 | apply (subst Elem_implode) | |
| 20565 | 744 | apply (simp add: set_like_def) | 
| 19203 | 745 | apply (simp add: Ext_def) | 
| 746 | done | |
| 747 | ||
| 748 | consts | |
| 749 | Ext_ZF_n :: "(ZF * ZF) set \<Rightarrow> ZF \<Rightarrow> nat \<Rightarrow> ZF" | |
| 750 | ||
| 751 | primrec | |
| 752 | "Ext_ZF_n R s 0 = Ext_ZF R s" | |
| 753 | "Ext_ZF_n R s (Suc n) = Sum (Repl (Ext_ZF_n R s n) (Ext_ZF R))" | |
| 754 | ||
| 755 | constdefs | |
| 756 | Ext_ZF_hull :: "(ZF * ZF) set \<Rightarrow> ZF \<Rightarrow> ZF" | |
| 757 | "Ext_ZF_hull R s == SeqSum (Ext_ZF_n R s)" | |
| 758 | ||
| 759 | lemma Elem_Ext_ZF_hull: | |
| 20565 | 760 | assumes set_like_R: "set_like R" | 
| 19203 | 761 | shows "Elem x (Ext_ZF_hull R S) = (? n. Elem x (Ext_ZF_n R S n))" | 
| 762 | by (simp add: Ext_ZF_hull_def SeqSum) | |
| 763 | ||
| 764 | lemma Elem_Elem_Ext_ZF_hull: | |
| 20565 | 765 | assumes set_like_R: "set_like R" | 
| 19203 | 766 | and x_hull: "Elem x (Ext_ZF_hull R S)" | 
| 767 | and y_R_x: "(y, x) \<in> R" | |
| 768 | shows "Elem y (Ext_ZF_hull R S)" | |
| 769 | proof - | |
| 20565 | 770 | from Elem_Ext_ZF_hull[OF set_like_R] x_hull | 
| 19203 | 771 | have "? n. Elem x (Ext_ZF_n R S n)" by auto | 
| 772 | then obtain n where n:"Elem x (Ext_ZF_n R S n)" .. | |
| 773 | with y_R_x have "Elem y (Ext_ZF_n R S (Suc n))" | |
| 774 | apply (auto simp add: Repl Sum) | |
| 775 | apply (rule_tac x="Ext_ZF R x" in exI) | |
| 20565 | 776 | apply (auto simp add: Elem_Ext_ZF[OF set_like_R]) | 
| 19203 | 777 | done | 
| 20565 | 778 | with Elem_Ext_ZF_hull[OF set_like_R, where x=y] show ?thesis | 
| 19203 | 779 | by (auto simp del: Ext_ZF_n.simps) | 
| 780 | qed | |
| 781 | ||
| 782 | lemma wfzf_minimal: | |
| 783 |   assumes hyps: "wfzf R" "C \<noteq> {}"
 | |
| 784 | shows "\<exists>x. x \<in> C \<and> (\<forall>y. (y, x) \<in> R \<longrightarrow> y \<notin> C)" | |
| 785 | proof - | |
| 786 | from hyps have "\<exists>S. S \<in> C" by auto | |
| 787 | then obtain S where S:"S \<in> C" by auto | |
| 788 | let ?T = "Sep (Ext_ZF_hull R S) (\<lambda> s. s \<in> C)" | |
| 20565 | 789 | from hyps have set_like_R: "set_like R" by (simp add: wfzf_def) | 
| 19203 | 790 | show ?thesis | 
| 791 | proof (cases "?T = Empty") | |
| 792 | case True | |
| 793 | then have "\<forall> z. \<not> (Elem z (Sep (Ext_ZF R S) (\<lambda> s. s \<in> C)))" | |
| 794 | apply (auto simp add: Ext Empty Sep Ext_ZF_hull_def SeqSum) | |
| 795 | apply (erule_tac x="z" in allE, auto) | |
| 796 | apply (erule_tac x=0 in allE, auto) | |
| 797 | done | |
| 798 | then show ?thesis | |
| 799 | apply (rule_tac exI[where x=S]) | |
| 800 | apply (auto simp add: Sep Empty S) | |
| 801 | apply (erule_tac x=y in allE) | |
| 20565 | 802 | apply (simp add: set_like_R Elem_Ext_ZF) | 
| 19203 | 803 | done | 
| 804 | next | |
| 805 | case False | |
| 806 | from hyps have regular_R: "regular R" by (simp add: wfzf_def) | |
| 807 | from | |
| 808 | regular_R[simplified regular_def, rule_format, OF False, simplified Sep] | |
| 20565 | 809 | Elem_Elem_Ext_ZF_hull[OF set_like_R] | 
| 19203 | 810 | show ?thesis by blast | 
| 811 | qed | |
| 812 | qed | |
| 813 | ||
| 814 | lemma wfzf_implies_wf: "wfzf R \<Longrightarrow> wf R" | |
| 815 | proof (subst wf_def, rule allI) | |
| 816 | assume wfzf: "wfzf R" | |
| 817 | fix P :: "ZF \<Rightarrow> bool" | |
| 818 |   let ?C = "{x. P x}"
 | |
| 819 |   {
 | |
| 820 | assume induct: "(\<forall>x. (\<forall>y. (y, x) \<in> R \<longrightarrow> P y) \<longrightarrow> P x)" | |
| 821 |     let ?C = "{x. \<not> (P x)}"
 | |
| 822 |     have "?C = {}"
 | |
| 823 | proof (rule ccontr) | |
| 824 |       assume C: "?C \<noteq> {}"
 | |
| 825 | from | |
| 826 | wfzf_minimal[OF wfzf C] | |
| 827 | obtain x where x: "x \<in> ?C \<and> (\<forall>y. (y, x) \<in> R \<longrightarrow> y \<notin> ?C)" .. | |
| 828 | then have "P x" | |
| 829 | apply (rule_tac induct[rule_format]) | |
| 830 | apply auto | |
| 831 | done | |
| 832 | with x show "False" by auto | |
| 833 | qed | |
| 834 | then have "! x. P x" by auto | |
| 835 | } | |
| 836 | then show "(\<forall>x. (\<forall>y. (y, x) \<in> R \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> (! x. P x)" by blast | |
| 837 | qed | |
| 838 | ||
| 839 | lemma wf_is_Elem_of: "wf is_Elem_of" | |
| 840 | by (auto simp add: wfzf_is_Elem_of wfzf_implies_wf) | |
| 841 | ||
| 842 | lemma in_Ext_RTrans_implies_Elem_Ext_ZF_hull: | |
| 20565 | 843 | "set_like R \<Longrightarrow> x \<in> (Ext (R^+) s) \<Longrightarrow> Elem x (Ext_ZF_hull R s)" | 
| 19203 | 844 | apply (simp add: Ext_def Elem_Ext_ZF_hull) | 
| 845 | apply (erule converse_trancl_induct[where r="R"]) | |
| 846 | apply (rule exI[where x=0]) | |
| 847 | apply (simp add: Elem_Ext_ZF) | |
| 848 | apply auto | |
| 849 | apply (rule_tac x="Suc n" in exI) | |
| 850 | apply (simp add: Sum Repl) | |
| 851 | apply (rule_tac x="Ext_ZF R z" in exI) | |
| 852 | apply (auto simp add: Elem_Ext_ZF) | |
| 853 | done | |
| 854 | ||
| 20565 | 855 | lemma implodeable_Ext_trancl: "set_like R \<Longrightarrow> set_like (R^+)" | 
| 856 | apply (subst set_like_def) | |
| 19203 | 857 | apply (auto simp add: image_def) | 
| 858 | apply (rule_tac x="Sep (Ext_ZF_hull R y) (\<lambda> z. z \<in> (Ext (R^+) y))" in exI) | |
| 859 | apply (auto simp add: explode_def Sep set_ext | |
| 860 | in_Ext_RTrans_implies_Elem_Ext_ZF_hull) | |
| 861 | done | |
| 862 | ||
| 863 | lemma Elem_Ext_ZF_hull_implies_in_Ext_RTrans[rule_format]: | |
| 20565 | 864 | "set_like R \<Longrightarrow> ! x. Elem x (Ext_ZF_n R s n) \<longrightarrow> x \<in> (Ext (R^+) s)" | 
| 19203 | 865 | apply (induct_tac n) | 
| 866 | apply (auto simp add: Elem_Ext_ZF Ext_def Sum Repl) | |
| 867 | done | |
| 868 | ||
| 20565 | 869 | lemma "set_like R \<Longrightarrow> Ext_ZF (R^+) s = Ext_ZF_hull R s" | 
| 19203 | 870 | apply (frule implodeable_Ext_trancl) | 
| 871 | apply (auto simp add: Ext) | |
| 872 | apply (erule in_Ext_RTrans_implies_Elem_Ext_ZF_hull) | |
| 873 | apply (simp add: Elem_Ext_ZF Ext_def) | |
| 874 | apply (auto simp add: Elem_Ext_ZF Elem_Ext_ZF_hull) | |
| 875 | apply (erule Elem_Ext_ZF_hull_implies_in_Ext_RTrans[simplified Ext_def, simplified], assumption) | |
| 876 | done | |
| 877 | ||
| 878 | lemma wf_implies_regular: "wf R \<Longrightarrow> regular R" | |
| 879 | proof (simp add: regular_def, rule allI) | |
| 880 | assume wf: "wf R" | |
| 881 | fix A | |
| 882 | show "A \<noteq> Empty \<longrightarrow> (\<exists>x. Elem x A \<and> (\<forall>y. (y, x) \<in> R \<longrightarrow> \<not> Elem y A))" | |
| 883 | proof | |
| 884 | assume A: "A \<noteq> Empty" | |
| 885 | then have "? x. x \<in> explode A" | |
| 886 | by (auto simp add: explode_def Ext Empty) | |
| 887 | then obtain x where x:"x \<in> explode A" .. | |
| 888 | from iffD1[OF wf_eq_minimal wf, rule_format, where Q="explode A", OF x] | |
| 889 | obtain z where "z \<in> explode A \<and> (\<forall>y. (y, z) \<in> R \<longrightarrow> y \<notin> explode A)" by auto | |
| 890 | then show "\<exists>x. Elem x A \<and> (\<forall>y. (y, x) \<in> R \<longrightarrow> \<not> Elem y A)" | |
| 891 | apply (rule_tac exI[where x = z]) | |
| 892 | apply (simp add: explode_def) | |
| 893 | done | |
| 894 | qed | |
| 895 | qed | |
| 896 | ||
| 20565 | 897 | lemma wf_eq_wfzf: "(wf R \<and> set_like R) = wfzf R" | 
| 19203 | 898 | apply (auto simp add: wfzf_implies_wf) | 
| 899 | apply (auto simp add: wfzf_def wf_implies_regular) | |
| 900 | done | |
| 901 | ||
| 902 | lemma wfzf_trancl: "wfzf R \<Longrightarrow> wfzf (R^+)" | |
| 903 | by (auto simp add: wf_eq_wfzf[symmetric] implodeable_Ext_trancl wf_trancl) | |
| 904 | ||
| 905 | lemma Ext_subset_mono: "R \<subseteq> S \<Longrightarrow> Ext R y \<subseteq> Ext S y" | |
| 906 | by (auto simp add: Ext_def) | |
| 907 | ||
| 20565 | 908 | lemma set_like_subset: "set_like R \<Longrightarrow> S \<subseteq> R \<Longrightarrow> set_like S" | 
| 909 | apply (auto simp add: set_like_def) | |
| 19203 | 910 | apply (erule_tac x=y in allE) | 
| 911 | apply (drule_tac y=y in Ext_subset_mono) | |
| 912 | apply (auto simp add: image_def) | |
| 913 | apply (rule_tac x="Sep x (% z. z \<in> (Ext S y))" in exI) | |
| 914 | apply (auto simp add: explode_def Sep) | |
| 915 | done | |
| 916 | ||
| 917 | lemma wfzf_subset: "wfzf S \<Longrightarrow> R \<subseteq> S \<Longrightarrow> wfzf R" | |
| 20565 | 918 | by (auto intro: set_like_subset wf_subset simp add: wf_eq_wfzf[symmetric]) | 
| 19203 | 919 | |
| 920 | end |