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