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