Constructible: some separation axioms
authorpaulson
Thu, 04 Jul 2002 16:59:54 +0200
changeset 13298 b4f370679c65
parent 13297 e4ae0732e2be
child 13299 3a932abf97e8
Constructible: some separation axioms
src/ZF/Constructible/Formula.thy
src/ZF/Constructible/L_axioms.thy
src/ZF/Constructible/Relative.thy
src/ZF/Constructible/Wellorderings.thy
src/ZF/OrdQuant.thy
--- a/src/ZF/Constructible/Formula.thy	Thu Jul 04 16:48:21 2002 +0200
+++ b/src/ZF/Constructible/Formula.thy	Thu Jul 04 16:59:54 2002 +0200
@@ -95,7 +95,7 @@
 
 declare satisfies.simps [simp del]; 
 
-subsubsection{*Dividing line between primitive and derived connectives*}
+subsection{*Dividing line between primitive and derived connectives*}
 
 lemma sats_Or_iff [simp]:
   "env \<in> list(A) 
@@ -125,6 +125,11 @@
        ==> (x\<in>y) <-> sats(A, Member(i,j), env)" 
 by (simp add: satisfies.simps)
 
+lemma equal_iff_sats:
+      "[| nth(i,env) = x; nth(j,env) = y; env \<in> list(A)|]
+       ==> (x=y) <-> sats(A, Equal(i,j), env)" 
+by (simp add: satisfies.simps)
+
 lemma conj_iff_sats:
       "[| P <-> sats(A,p,env); Q <-> sats(A,q,env); env \<in> list(A)|]
        ==> (P & Q) <-> sats(A, And(p,q), env)"
@@ -161,22 +166,6 @@
 by (simp add: sats_Exists_iff) 
 
 
-
-(*pretty but unnecessary
-constdefs sat     :: "[i,i] => o"
-  "sat(A,p) == satisfies(A,p)`[] = 1"
-
-syntax "_sat"  :: "[i,i] => o"    (infixl "|=" 50)
-translations "A |= p" == "sat(A,p)"
-
-lemma [simp]: "(A |= Neg(p)) <-> ~ (A |= p)"
-by (simp add: sat_def)
-
-lemma [simp]: "(A |= And(p,q)) <-> A|=p & A|=q"
-by (simp add: sat_def)
-*) 
-
-
 constdefs incr_var :: "[i,i]=>i"
     "incr_var(x,lev) == if x<lev then x else succ(x)"
 
@@ -355,7 +344,7 @@
 done
 
 
-(**** TRYING INCR_BV1 AGAIN ****)
+subsection{*Renaming all but the first bound variable*}
 
 constdefs incr_bv1 :: "i => i"
     "incr_bv1(p) == incr_bv(p)`1"
@@ -398,7 +387,7 @@
    ==> arity(incr_bv1^n(p)) =
          (if 1 < arity(p) then n #+ arity(p) else arity(p))"
 apply (induct_tac n) 
-apply (simp_all add: arity_incr_bv1_eq )
+apply (simp_all add: arity_incr_bv1_eq)
 apply (simp add: not_lt_iff_le)
 apply (blast intro: le_trans add_le_self2 arity_type) 
 done
@@ -520,6 +509,76 @@
 oops
 *)
 
+
+subsection{*Internalized formulas for basic concepts*}
+
+subsubsection{*The subset relation*}
+
+lemma lt_length_in_nat:
+   "[|x < length(xs); xs \<in> list(A)|] ==> x \<in> nat"
+apply (frule lt_nat_in_nat, typecheck) 
+done
+
+constdefs subset_fm :: "[i,i]=>i"
+    "subset_fm(x,y) == Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))"
+
+lemma subset_type [TC]: "[| x \<in> nat; y \<in> nat |] ==> subset_fm(x,y) \<in> formula"
+by (simp add: subset_fm_def) 
+
+lemma arity_subset_fm [simp]:
+     "[| x \<in> nat; y \<in> nat |] ==> arity(subset_fm(x,y)) = succ(x) \<union> succ(y)"
+by (simp add: subset_fm_def succ_Un_distrib [symmetric]) 
+
+lemma sats_subset_fm [simp]:
+   "[|x < length(env); y \<in> nat; env \<in> list(A); Transset(A)|]
+    ==> sats(A, subset_fm(x,y), env) <-> nth(x,env) \<subseteq> nth(y,env)"
+apply (frule lt_length_in_nat, assumption)  
+apply (simp add: subset_fm_def Transset_def) 
+apply (blast intro: nth_type) 
+done
+
+subsubsection{*Transitive sets*}
+
+constdefs transset_fm :: "i=>i"
+   "transset_fm(x) == Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))"
+
+lemma transset_type [TC]: "x \<in> nat ==> transset_fm(x) \<in> formula"
+by (simp add: transset_fm_def) 
+
+lemma arity_transset_fm [simp]:
+     "x \<in> nat ==> arity(transset_fm(x)) = succ(x)"
+by (simp add: transset_fm_def succ_Un_distrib [symmetric]) 
+
+lemma sats_transset_fm [simp]:
+   "[|x < length(env); env \<in> list(A); Transset(A)|]
+    ==> sats(A, transset_fm(x), env) <-> Transset(nth(x,env))"
+apply (frule lt_nat_in_nat, erule length_type) 
+apply (simp add: transset_fm_def Transset_def) 
+apply (blast intro: nth_type) 
+done
+
+subsubsection{*Ordinals*}
+
+constdefs ordinal_fm :: "i=>i"
+   "ordinal_fm(x) == 
+      And(transset_fm(x), Forall(Implies(Member(0,succ(x)), transset_fm(0))))"
+
+lemma ordinal_type [TC]: "x \<in> nat ==> ordinal_fm(x) \<in> formula"
+by (simp add: ordinal_fm_def) 
+
+lemma arity_ordinal_fm [simp]:
+     "x \<in> nat ==> arity(ordinal_fm(x)) = succ(x)"
+by (simp add: ordinal_fm_def succ_Un_distrib [symmetric]) 
+
+lemma sats_ordinal_fm [simp]:
+   "[|x < length(env); env \<in> list(A); Transset(A)|]
+    ==> sats(A, ordinal_fm(x), env) <-> Ord(nth(x,env))"
+apply (frule lt_nat_in_nat, erule length_type) 
+apply (simp add: ordinal_fm_def Ord_def Transset_def)
+apply (blast intro: nth_type) 
+done
+
+
 subsection{* Constant Lset: Levels of the Constructible Universe *}
 
 constdefs Lset :: "i=>i"
@@ -661,63 +720,7 @@
 done
 
 
-
-text{*Kunen's VI, 1.9 (b)*}
-
-constdefs subset_fm :: "[i,i]=>i"
-    "subset_fm(x,y) == Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))"
-
-lemma subset_type [TC]: "[| x \<in> nat; y \<in> nat |] ==> subset_fm(x,y) \<in> formula"
-by (simp add: subset_fm_def) 
-
-lemma arity_subset_fm [simp]:
-     "[| x \<in> nat; y \<in> nat |] ==> arity(subset_fm(x,y)) = succ(x) \<union> succ(y)"
-by (simp add: subset_fm_def succ_Un_distrib [symmetric]) 
-
-lemma sats_subset_fm [simp]:
-   "[|x < length(env); y \<in> nat; env \<in> list(A); Transset(A)|]
-    ==> sats(A, subset_fm(x,y), env) <-> nth(x,env) \<subseteq> nth(y,env)"
-apply (frule lt_nat_in_nat, erule length_type) 
-apply (simp add: subset_fm_def Transset_def) 
-apply (blast intro: nth_type ) 
-done
-
-constdefs transset_fm :: "i=>i"
-   "transset_fm(x) == Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))"
-
-lemma transset_type [TC]: "x \<in> nat ==> transset_fm(x) \<in> formula"
-by (simp add: transset_fm_def) 
-
-lemma arity_transset_fm [simp]:
-     "x \<in> nat ==> arity(transset_fm(x)) = succ(x)"
-by (simp add: transset_fm_def succ_Un_distrib [symmetric]) 
-
-lemma sats_transset_fm [simp]:
-   "[|x < length(env); env \<in> list(A); Transset(A)|]
-    ==> sats(A, transset_fm(x), env) <-> Transset(nth(x,env))"
-apply (frule lt_nat_in_nat, erule length_type) 
-apply (simp add: transset_fm_def Transset_def) 
-apply (blast intro: nth_type ) 
-done
-
-constdefs ordinal_fm :: "i=>i"
-   "ordinal_fm(x) == 
-      And(transset_fm(x), Forall(Implies(Member(0,succ(x)), transset_fm(0))))"
-
-lemma ordinal_type [TC]: "x \<in> nat ==> ordinal_fm(x) \<in> formula"
-by (simp add: ordinal_fm_def) 
-
-lemma arity_ordinal_fm [simp]:
-     "x \<in> nat ==> arity(ordinal_fm(x)) = succ(x)"
-by (simp add: ordinal_fm_def succ_Un_distrib [symmetric]) 
-
-lemma sats_ordinal_fm [simp]:
-   "[|x < length(env); env \<in> list(A); Transset(A)|]
-    ==> sats(A, ordinal_fm(x), env) <-> Ord(nth(x,env))"
-apply (frule lt_nat_in_nat, erule length_type) 
-apply (simp add: ordinal_fm_def Ord_def Transset_def)
-apply (blast intro: nth_type ) 
-done
+subsection{*Constructible Ordinals: Kunen's VI, 1.9 (b)*}
 
 text{*The subset consisting of the ordinals is definable.*}
 lemma Ords_in_DPow: "Transset(A) ==> {x \<in> A. Ord(x)} \<in> DPow(A)"
--- a/src/ZF/Constructible/L_axioms.thy	Thu Jul 04 16:48:21 2002 +0200
+++ b/src/ZF/Constructible/L_axioms.thy	Thu Jul 04 16:59:54 2002 +0200
@@ -288,30 +288,344 @@
 by blast
 
 
+subsection{*Internalized formulas for some relativized ones*}
+
+subsubsection{*Unordered pairs*}
+
+constdefs upair_fm :: "[i,i,i]=>i"
+    "upair_fm(x,y,z) == 
+       And(Member(x,z), 
+           And(Member(y,z),
+               Forall(Implies(Member(0,succ(z)), 
+                              Or(Equal(0,succ(x)), Equal(0,succ(y)))))))"
+
+lemma upair_type [TC]:
+     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> upair_fm(x,y,z) \<in> formula"
+by (simp add: upair_fm_def) 
+
+lemma arity_upair_fm [simp]:
+     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
+      ==> arity(upair_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
+by (simp add: upair_fm_def succ_Un_distrib [symmetric] Un_ac) 
+
+lemma sats_upair_fm [simp]:
+   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
+    ==> sats(A, upair_fm(x,y,z), env) <-> 
+            upair(**A, nth(x,env), nth(y,env), nth(z,env))"
+by (simp add: upair_fm_def upair_def)
+
+lemma upair_iff_sats:
+      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
+          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
+       ==> upair(**A, x, y, z) <-> sats(A, upair_fm(i,j,k), env)"
+by (simp add: sats_upair_fm)
+
+text{*Useful? At least it refers to "real" unordered pairs*}
+lemma sats_upair_fm2 [simp]:
+   "[| x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A); Transset(A)|]
+    ==> sats(A, upair_fm(x,y,z), env) <-> 
+        nth(z,env) = {nth(x,env), nth(y,env)}"
+apply (frule lt_length_in_nat, assumption)  
+apply (simp add: upair_fm_def Transset_def, auto) 
+apply (blast intro: nth_type) 
+done
+
+subsubsection{*Ordered pairs*}
+
+constdefs pair_fm :: "[i,i,i]=>i"
+    "pair_fm(x,y,z) == 
+       Exists(And(upair_fm(succ(x),succ(x),0),
+              Exists(And(upair_fm(succ(succ(x)),succ(succ(y)),0),
+                         upair_fm(1,0,succ(succ(z)))))))"
+
+lemma pair_type [TC]:
+     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> pair_fm(x,y,z) \<in> formula"
+by (simp add: pair_fm_def) 
+
+lemma arity_pair_fm [simp]:
+     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
+      ==> arity(pair_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
+by (simp add: pair_fm_def succ_Un_distrib [symmetric] Un_ac) 
+
+lemma sats_pair_fm [simp]:
+   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
+    ==> sats(A, pair_fm(x,y,z), env) <-> 
+        pair(**A, nth(x,env), nth(y,env), nth(z,env))"
+by (simp add: pair_fm_def pair_def)
+
+lemma pair_iff_sats:
+      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
+          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
+       ==> pair(**A, x, y, z) <-> sats(A, pair_fm(i,j,k), env)"
+by (simp add: sats_pair_fm)
+
+
+
+subsection{*Proving instances of Separation using Reflection!*}
+
+text{*Helps us solve for de Bruijn indices!*}
+lemma nth_ConsI: "[|nth(n,l) = x; n \<in> nat|] ==> nth(succ(n), Cons(a,l)) = x"
+by simp
+
+
+lemma Collect_conj_in_DPow:
+     "[| {x\<in>A. P(x)} \<in> DPow(A);  {x\<in>A. Q(x)} \<in> DPow(A) |] 
+      ==> {x\<in>A. P(x) & Q(x)} \<in> DPow(A)"
+by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric]) 
+
+lemma Collect_conj_in_DPow_Lset:
+     "[|z \<in> Lset(j); {x \<in> Lset(j). P(x)} \<in> DPow(Lset(j))|]
+      ==> {x \<in> Lset(j). x \<in> z & P(x)} \<in> DPow(Lset(j))"
+apply (frule mem_Lset_imp_subset_Lset)
+apply (simp add: Collect_conj_in_DPow Collect_mem_eq 
+                 subset_Int_iff2 elem_subset_in_DPow)
+done
+
+lemma separation_CollectI:
+     "(\<And>z. L(z) ==> L({x \<in> z . P(x)})) ==> separation(L, \<lambda>x. P(x))"
+apply (unfold separation_def, clarify) 
+apply (rule_tac x="{x\<in>z. P(x)}" in rexI) 
+apply simp_all
+done
+
+text{*Reduces the original comprehension to the reflected one*}
+lemma reflection_imp_L_separation:
+      "[| \<forall>x\<in>Lset(j). P(x) <-> Q(x);
+          {x \<in> Lset(j) . Q(x)} \<in> DPow(Lset(j)); 
+          Ord(j);  z \<in> Lset(j)|] ==> L({x \<in> z . P(x)})"
+apply (rule_tac i = "succ(j)" in L_I)
+ prefer 2 apply simp
+apply (subgoal_tac "{x \<in> z. P(x)} = {x \<in> Lset(j). x \<in> z & (Q(x))}")
+ prefer 2
+ apply (blast dest: mem_Lset_imp_subset_Lset) 
+apply (simp add: Lset_succ Collect_conj_in_DPow_Lset)
+done
+
+
+subsubsection{*Separation for Intersection*}
+
+lemma Inter_Reflects:
+     "L_Reflects(?Cl, \<lambda>x. \<forall>y[L]. y\<in>A --> x \<in> y, 
+               \<lambda>i x. \<forall>y\<in>Lset(i). y\<in>A --> x \<in> y)"
+by fast
+
+lemma Inter_separation:
+     "L(A) ==> separation(L, \<lambda>x. \<forall>y[L]. y\<in>A --> x\<in>y)"
+apply (rule separation_CollectI) 
+apply (rule_tac A="{A,z}" in subset_LsetE, blast ) 
+apply (rule ReflectsE [OF Inter_Reflects], assumption)
+apply (drule subset_Lset_ltD, assumption) 
+apply (erule reflection_imp_L_separation)
+  apply (simp_all add: lt_Ord2, clarify)
+apply (rule DPowI2) 
+apply (rule ball_iff_sats) 
+apply (rule imp_iff_sats)
+apply (rule_tac [2] i=1 and j=0 and env="[y,x,A]" in mem_iff_sats)
+apply (rule_tac i=0 and j=2 in mem_iff_sats)
+apply (simp_all add: succ_Un_distrib [symmetric])
+done
+
+subsubsection{*Separation for Cartesian Product*}
+
+text{*The @{text simplified} attribute tidies up the reflecting class.*}
+theorem upair_reflection [simplified,intro]:
+     "L_Reflects(?Cl, \<lambda>x. upair(L,f(x),g(x),h(x)), 
+                    \<lambda>i x. upair(**Lset(i),f(x),g(x),h(x)))" 
+by (simp add: upair_def, fast) 
+
+theorem pair_reflection [simplified,intro]:
+     "L_Reflects(?Cl, \<lambda>x. pair(L,f(x),g(x),h(x)), 
+                    \<lambda>i x. pair(**Lset(i),f(x),g(x),h(x)))"
+by (simp only: pair_def rex_setclass_is_bex, fast) 
+
+lemma cartprod_Reflects [simplified]:
+     "L_Reflects(?Cl, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)),
+                \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). y\<in>B & 
+                               pair(**Lset(i),x,y,z)))"
+by fast
+
+lemma cartprod_separation:
+     "[| L(A); L(B) |] 
+      ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)))"
+apply (rule separation_CollectI) 
+apply (rule_tac A="{A,B,z}" in subset_LsetE, blast ) 
+apply (rule ReflectsE [OF cartprod_Reflects], assumption)
+apply (drule subset_Lset_ltD, assumption) 
+apply (erule reflection_imp_L_separation)
+  apply (simp_all add: lt_Ord2, clarify) 
+apply (rule DPowI2)
+apply (rename_tac u)  
+apply (rule bex_iff_sats) 
+apply (rule conj_iff_sats)
+apply (rule_tac i=0 and j=2 and env="[x,u,A,B]" in mem_iff_sats, simp_all)
+apply (rule bex_iff_sats) 
+apply (rule conj_iff_sats)
+apply (rule mem_iff_sats)
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI, simp_all)
+apply (rule_tac i=1 and j=0 and k=2 in pair_iff_sats)
+apply (simp_all add: succ_Un_distrib [symmetric])
+done
+
+subsubsection{*Separation for Image*}
+
+text{*No @{text simplified} here: it simplifies the occurrence of 
+      the predicate @{term pair}!*}
+lemma image_Reflects:
+     "L_Reflects(?Cl, \<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)),
+           \<lambda>i y. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). x\<in>A & pair(**Lset(i),x,y,p)))"
+by fast
+
+
+lemma image_separation:
+     "[| L(A); L(r) |] 
+      ==> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)))"
+apply (rule separation_CollectI) 
+apply (rule_tac A="{A,r,z}" in subset_LsetE, blast ) 
+apply (rule ReflectsE [OF image_Reflects], assumption)
+apply (drule subset_Lset_ltD, assumption) 
+apply (erule reflection_imp_L_separation)
+  apply (simp_all add: lt_Ord2, clarify)
+apply (rule DPowI2)
+apply (rule bex_iff_sats) 
+apply (rule conj_iff_sats)
+apply (rule_tac env="[p,y,A,r]" in mem_iff_sats)
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI, simp_all)
+apply (rule bex_iff_sats) 
+apply (rule conj_iff_sats)
+apply (rule mem_iff_sats)
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI, simp_all)
+apply (rule pair_iff_sats)
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI)
+apply (simp_all add: succ_Un_distrib [symmetric])
+done
+
+
+subsubsection{*Separation for Converse*}
+
+lemma converse_Reflects:
+     "L_Reflects(?Cl, 
+        \<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)),
+     \<lambda>i z. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i). 
+                     pair(**Lset(i),x,y,p) & pair(**Lset(i),y,x,z)))"
+by fast
+
+lemma converse_separation:
+     "L(r) ==> separation(L, 
+         \<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)))"
+apply (rule separation_CollectI) 
+apply (rule_tac A="{r,z}" in subset_LsetE, blast ) 
+apply (rule ReflectsE [OF converse_Reflects], assumption)
+apply (drule subset_Lset_ltD, assumption) 
+apply (erule reflection_imp_L_separation)
+  apply (simp_all add: lt_Ord2, clarify)
+apply (rule DPowI2)
+apply (rename_tac u) 
+apply (rule bex_iff_sats) 
+apply (rule conj_iff_sats)
+apply (rule_tac i=0 and j="2" and env="[p,u,r]" in mem_iff_sats, simp_all)
+apply (rule bex_iff_sats) 
+apply (rule bex_iff_sats) 
+apply (rule conj_iff_sats)
+apply (rule_tac i=1 and j=0 and k=2 in pair_iff_sats, simp_all)
+apply (rule pair_iff_sats)
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI)
+apply (simp_all add: succ_Un_distrib [symmetric])
+done
+
+
+subsubsection{*Separation for Restriction*}
+
+lemma restrict_Reflects:
+     "L_Reflects(?Cl, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)),
+        \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). pair(**Lset(i),x,y,z)))"
+by fast
+
+lemma restrict_separation:
+   "L(A) ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)))"
+apply (rule separation_CollectI) 
+apply (rule_tac A="{A,z}" in subset_LsetE, blast ) 
+apply (rule ReflectsE [OF restrict_Reflects], assumption)
+apply (drule subset_Lset_ltD, assumption) 
+apply (erule reflection_imp_L_separation)
+  apply (simp_all add: lt_Ord2, clarify)
+apply (rule DPowI2)
+apply (rename_tac u) 
+apply (rule bex_iff_sats) 
+apply (rule conj_iff_sats)
+apply (rule_tac i=0 and j="2" and env="[x,u,A]" in mem_iff_sats, simp_all)
+apply (rule bex_iff_sats) 
+apply (rule_tac i=1 and j=0 and k=2 in pair_iff_sats)
+apply (simp_all add: succ_Un_distrib [symmetric])
+done
+
+
+subsubsection{*Separation for Composition*}
+
+lemma comp_Reflects:
+     "L_Reflects(?Cl, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L]. 
+		  pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & 
+                  xy\<in>s & yz\<in>r,
+        \<lambda>i xz. \<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i). \<exists>z\<in>Lset(i). \<exists>xy\<in>Lset(i). \<exists>yz\<in>Lset(i). 
+		  pair(**Lset(i),x,z,xz) & pair(**Lset(i),x,y,xy) & 
+                  pair(**Lset(i),y,z,yz) & xy\<in>s & yz\<in>r)"
+by fast
+
+lemma comp_separation:
+     "[| L(r); L(s) |]
+      ==> separation(L, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L]. 
+		  pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & 
+                  xy\<in>s & yz\<in>r)"
+apply (rule separation_CollectI) 
+apply (rule_tac A="{r,s,z}" in subset_LsetE, blast ) 
+apply (rule ReflectsE [OF comp_Reflects], assumption)
+apply (drule subset_Lset_ltD, assumption) 
+apply (erule reflection_imp_L_separation)
+  apply (simp_all add: lt_Ord2, clarify)
+apply (rule DPowI2)
+apply (rename_tac u) 
+apply (rule bex_iff_sats)+
+apply (rename_tac x y z)  
+apply (rule conj_iff_sats)
+apply (rule_tac env="[z,y,x,u,r,s]" in pair_iff_sats)
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI, simp_all)
+apply (rule bex_iff_sats) 
+apply (rule conj_iff_sats)
+apply (rule pair_iff_sats)
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI, simp_all)
+apply (rule bex_iff_sats) 
+apply (rule conj_iff_sats)
+apply (rule pair_iff_sats)
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI, simp_all) 
+apply (rule conj_iff_sats)
+apply (rule mem_iff_sats) 
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI, simp) 
+apply (rule mem_iff_sats) 
+apply (blast intro: nth_0 nth_ConsI) 
+apply (blast intro: nth_0 nth_ConsI)
+apply (simp_all add: succ_Un_distrib [symmetric])
+done
+
+
+
+
 end
 
 (*
 
-  and cartprod_separation:
-     "[| L(A); L(B) |] 
-      ==> separation(L, \<lambda>z. \<exists>x\<in>A. \<exists>y\<in>B. L(x) & L(y) & pair(L,x,y,z))"
-  and image_separation:
-     "[| L(A); L(r) |] 
-      ==> separation(L, \<lambda>y. \<exists>p\<in>r. L(p) & (\<exists>x\<in>A. L(x) & pair(L,x,y,p)))"
-  and vimage_separation:
-     "[| L(A); L(r) |] 
-      ==> separation(L, \<lambda>x. \<exists>p\<in>r. L(p) & (\<exists>y\<in>A. L(x) & pair(L,x,y,p)))"
-  and converse_separation:
-     "L(r) ==> separation(L, \<lambda>z. \<exists>p\<in>r. L(p) & (\<exists>x y. L(x) & L(y) & 
-				     pair(L,x,y,p) & pair(L,y,x,z)))"
-  and restrict_separation:
-     "L(A) 
-      ==> separation(L, \<lambda>z. \<exists>x\<in>A. L(x) & (\<exists>y. L(y) & pair(L,x,y,z)))"
-  and comp_separation:
-     "[| L(r); L(s) |]
-      ==> separation(L, \<lambda>xz. \<exists>x y z. L(x) & L(y) & L(z) &
-			   (\<exists>xy\<in>s. \<exists>yz\<in>r. L(xy) & L(yz) & 
-		  pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz)))"
   and pred_separation:
      "[| L(r); L(x) |] ==> separation(L, \<lambda>y. \<exists>p\<in>r. L(p) & pair(L,y,x,p))"
   and Memrel_separation:
--- a/src/ZF/Constructible/Relative.thy	Thu Jul 04 16:48:21 2002 +0200
+++ b/src/ZF/Constructible/Relative.thy	Thu Jul 04 16:59:54 2002 +0200
@@ -9,10 +9,10 @@
     "empty(M,z) == \<forall>x[M]. x \<notin> z"
 
   subset :: "[i=>o,i,i] => o"
-    "subset(M,A,B) == \<forall>x\<in>A. M(x) --> x \<in> B"
+    "subset(M,A,B) == \<forall>x[M]. x\<in>A --> x \<in> B"
 
   upair :: "[i=>o,i,i,i] => o"
-    "upair(M,a,b,z) == a \<in> z & b \<in> z & (\<forall>x\<in>z. M(x) --> x = a | x = b)"
+    "upair(M,a,b,z) == a \<in> z & b \<in> z & (\<forall>x[M]. x\<in>z --> x = a | x = b)"
 
   pair :: "[i=>o,i,i,i] => o"
     "pair(M,a,b,z) == \<exists>x[M]. upair(M,a,a,x) & 
@@ -34,35 +34,34 @@
     "setdiff(M,a,b,z) == \<forall>x[M]. x \<in> z <-> x \<in> a & x \<notin> b"
 
   big_union :: "[i=>o,i,i] => o"
-    "big_union(M,A,z) == \<forall>x[M]. x \<in> z <-> (\<exists>y\<in>A. M(y) & x \<in> y)"
+    "big_union(M,A,z) == \<forall>x[M]. x \<in> z <-> (\<exists>y[M]. y\<in>A & x \<in> y)"
 
   big_inter :: "[i=>o,i,i] => o"
     "big_inter(M,A,z) == 
              (A=0 --> z=0) &
-	     (A\<noteq>0 --> (\<forall>x[M]. x \<in> z <-> (\<forall>y\<in>A. M(y) --> x \<in> y)))"
+	     (A\<noteq>0 --> (\<forall>x[M]. x \<in> z <-> (\<forall>y[M]. y\<in>A --> x \<in> y)))"
 
   cartprod :: "[i=>o,i,i,i] => o"
     "cartprod(M,A,B,z) == 
-	\<forall>u[M]. u \<in> z <-> (\<exists>x\<in>A. M(x) & (\<exists>y\<in>B. M(y) & pair(M,x,y,u)))"
+	\<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))"
 
   is_converse :: "[i=>o,i,i] => o"
     "is_converse(M,r,z) == 
 	\<forall>x. M(x) --> 
             (x \<in> z <-> 
-             (\<exists>w\<in>r. M(w) & 
-              (\<exists>u[M]. \<exists>v[M]. pair(M,u,v,w) & pair(M,v,u,x))))"
+             (\<exists>w[M]. w\<in>r & (\<exists>u[M]. \<exists>v[M]. pair(M,u,v,w) & pair(M,v,u,x))))"
 
   pre_image :: "[i=>o,i,i,i] => o"
     "pre_image(M,r,A,z) == 
-	\<forall>x. M(x) --> (x \<in> z <-> (\<exists>w\<in>r. M(w) & (\<exists>y\<in>A. M(y) & pair(M,x,y,w))))"
+	\<forall>x. M(x) --> (x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w))))"
 
   is_domain :: "[i=>o,i,i] => o"
     "is_domain(M,r,z) == 
-	\<forall>x. M(x) --> (x \<in> z <-> (\<exists>w\<in>r. M(w) & (\<exists>y. M(y) & pair(M,x,y,w))))"
+	\<forall>x. M(x) --> (x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w))))"
 
   image :: "[i=>o,i,i,i] => o"
     "image(M,r,A,z) == 
-        \<forall>y. M(y) --> (y \<in> z <-> (\<exists>w\<in>r. M(w) & (\<exists>x\<in>A. M(x) & pair(M,x,y,w))))"
+        \<forall>y. M(y) --> (y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w))))"
 
   is_range :: "[i=>o,i,i] => o"
     --{*the cleaner 
@@ -70,16 +69,16 @@
       unfortunately needs an instance of separation in order to prove 
         @{term "M(converse(r))"}.*}
     "is_range(M,r,z) == 
-	\<forall>y. M(y) --> (y \<in> z <-> (\<exists>w\<in>r. M(w) & (\<exists>x. M(x) & pair(M,x,y,w))))"
+	\<forall>y. M(y) --> (y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w))))"
 
   is_field :: "[i=>o,i,i] => o"
     "is_field(M,r,z) == 
-	\<exists>dr. M(dr) & is_domain(M,r,dr) & 
-            (\<exists>rr. M(rr) & is_range(M,r,rr) & union(M,dr,rr,z))"
+	\<exists>dr[M]. is_domain(M,r,dr) & 
+            (\<exists>rr[M]. is_range(M,r,rr) & union(M,dr,rr,z))"
 
   is_relation :: "[i=>o,i] => o"
     "is_relation(M,r) == 
-        (\<forall>z\<in>r. M(z) --> (\<exists>x y. M(x) & M(y) & pair(M,x,y,z)))"
+        (\<forall>z[M]. z\<in>r --> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))"
 
   is_function :: "[i=>o,i] => o"
     "is_function(M,r) == 
@@ -713,13 +712,13 @@
      "M(A) ==> separation(M, \<lambda>x. \<forall>y[M]. y\<in>A --> x\<in>y)"
   and cartprod_separation:
      "[| M(A); M(B) |] 
-      ==> separation(M, \<lambda>z. \<exists>x\<in>A. \<exists>y\<in>B. M(x) & M(y) & pair(M,x,y,z))"
+      ==> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,z)))"
   and image_separation:
      "[| M(A); M(r) |] 
       ==> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,p)))"
   and converse_separation:
-     "M(r) ==> separation(M, \<lambda>z. \<exists>p\<in>r. 
-                    M(p) & (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,p) & pair(M,y,x,z)))"
+     "M(r) ==> separation(M, 
+         \<lambda>z. \<exists>p[M]. p\<in>r & (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,p) & pair(M,y,x,z)))"
   and restrict_separation:
      "M(A) ==> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A & (\<exists>y[M]. pair(M,x,y,z)))"
   and comp_separation:
@@ -728,9 +727,9 @@
 		  pair(M,x,z,xz) & pair(M,x,y,xy) & pair(M,y,z,yz) & 
                   xy\<in>s & yz\<in>r)"
   and pred_separation:
-     "[| M(r); M(x) |] ==> separation(M, \<lambda>y. \<exists>p\<in>r. M(p) & pair(M,y,x,p))"
+     "[| M(r); M(x) |] ==> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & pair(M,y,x,p))"
   and Memrel_separation:
-     "separation(M, \<lambda>z. \<exists>x y. M(x) & M(y) & pair(M,x,y,z) & x \<in> y)"
+     "separation(M, \<lambda>z. \<exists>x[M]. \<exists>y[M]. pair(M,x,y,z) & x \<in> y)"
   and obase_separation:
      --{*part of the order type formalization*}
      "[| M(A); M(r) |] 
--- a/src/ZF/Constructible/Wellorderings.thy	Thu Jul 04 16:48:21 2002 +0200
+++ b/src/ZF/Constructible/Wellorderings.thy	Thu Jul 04 16:59:54 2002 +0200
@@ -192,7 +192,7 @@
 
 lemma (in M_axioms) M_Memrel_iff:
      "M(A) ==> 
-      Memrel(A) = {z \<in> A*A. \<exists>x. M(x) \<and> (\<exists>y. M(y) \<and> z = \<langle>x,y\<rangle> \<and> x \<in> y)}"
+      Memrel(A) = {z \<in> A*A. \<exists>x[M]. \<exists>y[M]. z = \<langle>x,y\<rangle> & x \<in> y}"
 apply (simp add: Memrel_def) 
 apply (blast dest: transM)
 done 
--- a/src/ZF/OrdQuant.thy	Thu Jul 04 16:48:21 2002 +0200
+++ b/src/ZF/OrdQuant.thy	Thu Jul 04 16:59:54 2002 +0200
@@ -10,18 +10,18 @@
 subsection {*Quantifiers and union operator for ordinals*}
 
 constdefs
-  
+
   (* Ordinal Quantifiers *)
   oall :: "[i, i => o] => o"
     "oall(A, P) == ALL x. x<A --> P(x)"
-  
+
   oex :: "[i, i => o] => o"
     "oex(A, P)  == EX x. x<A & P(x)"
 
   (* Ordinal Union *)
   OUnion :: "[i, i => i] => i"
     "OUnion(i,B) == {z: UN x:i. B(x). Ord(i)}"
-  
+
 syntax
   "@oall"     :: "[idt, i, o] => o"        ("(3ALL _<_./ _)" 10)
   "@oex"      :: "[idt, i, o] => o"        ("(3EX _<_./ _)" 10)
@@ -42,29 +42,29 @@
 
 
 (*MOST IMPORTANT that this is added to the simpset BEFORE Ord_atomize
-  is proved.  Ord_atomize would convert this rule to 
+  is proved.  Ord_atomize would convert this rule to
     x < 0 ==> P(x) == True, which causes dire effects!*)
 lemma [simp]: "(ALL x<0. P(x))"
-by (simp add: oall_def) 
+by (simp add: oall_def)
 
 lemma [simp]: "~(EX x<0. P(x))"
-by (simp add: oex_def) 
+by (simp add: oex_def)
 
 lemma [simp]: "(ALL x<succ(i). P(x)) <-> (Ord(i) --> P(i) & (ALL x<i. P(x)))"
-apply (simp add: oall_def le_iff) 
-apply (blast intro: lt_Ord2) 
+apply (simp add: oall_def le_iff)
+apply (blast intro: lt_Ord2)
 done
 
 lemma [simp]: "(EX x<succ(i). P(x)) <-> (Ord(i) & (P(i) | (EX x<i. P(x))))"
-apply (simp add: oex_def le_iff) 
-apply (blast intro: lt_Ord2) 
+apply (simp add: oex_def le_iff)
+apply (blast intro: lt_Ord2)
 done
 
 (** Union over ordinals **)
 
 lemma Ord_OUN [intro,simp]:
      "[| !!x. x<A ==> Ord(B(x)) |] ==> Ord(\<Union>x<A. B(x))"
-by (simp add: OUnion_def ltI Ord_UN) 
+by (simp add: OUnion_def ltI Ord_UN)
 
 lemma OUN_upper_lt:
      "[| a<A;  i < b(a);  Ord(\<Union>x<A. b(x)) |] ==> i < (\<Union>x<A. b(x))"
@@ -74,7 +74,7 @@
      "[| a<A;  i\<le>b(a);  Ord(\<Union>x<A. b(x)) |] ==> i \<le> (\<Union>x<A. b(x))"
 apply (unfold OUnion_def, auto)
 apply (rule UN_upper_le )
-apply (auto simp add: lt_def) 
+apply (auto simp add: lt_def)
 done
 
 lemma Limit_OUN_eq: "Limit(i) ==> (UN x<i. x) = i"
@@ -97,12 +97,12 @@
 lemma OUN_UN_eq:
      "(!!x. x:A ==> Ord(B(x)))
       ==> (UN z < (UN x:A. B(x)). C(z)) = (UN  x:A. UN z < B(x). C(z))"
-by (simp add: OUnion_def) 
+by (simp add: OUnion_def)
 
 lemma OUN_Union_eq:
      "(!!x. x:X ==> Ord(x))
       ==> (UN z < Union(X). C(z)) = (UN x:X. UN z < x. C(z))"
-by (simp add: OUnion_def) 
+by (simp add: OUnion_def)
 
 (*So that rule_format will get rid of ALL x<A...*)
 lemma atomize_oall [symmetric, rulify]:
@@ -113,20 +113,18 @@
 
 lemma oallI [intro!]:
     "[| !!x. x<A ==> P(x) |] ==> ALL x<A. P(x)"
-by (simp add: oall_def) 
+by (simp add: oall_def)
 
 lemma ospec: "[| ALL x<A. P(x);  x<A |] ==> P(x)"
-by (simp add: oall_def) 
+by (simp add: oall_def)
 
 lemma oallE:
     "[| ALL x<A. P(x);  P(x) ==> Q;  ~x<A ==> Q |] ==> Q"
-apply (simp add: oall_def, blast) 
-done
+by (simp add: oall_def, blast)
 
 lemma rev_oallE [elim]:
     "[| ALL x<A. P(x);  ~x<A ==> Q;  P(x) ==> Q |] ==> Q"
-apply (simp add: oall_def, blast)  
-done
+by (simp add: oall_def, blast)
 
 
 (*Trival rewrite rule;   (ALL x<a.P)<->P holds only if a is not 0!*)
@@ -135,7 +133,7 @@
 
 (*Congruence rule for rewriting*)
 lemma oall_cong [cong]:
-    "[| a=a';  !!x. x<a' ==> P(x) <-> P'(x) |] 
+    "[| a=a';  !!x. x<a' ==> P(x) <-> P'(x) |]
      ==> oall(a, %x. P(x)) <-> oall(a', %x. P'(x))"
 by (simp add: oall_def)
 
@@ -144,22 +142,22 @@
 
 lemma oexI [intro]:
     "[| P(x);  x<A |] ==> EX x<A. P(x)"
-apply (simp add: oex_def, blast) 
+apply (simp add: oex_def, blast)
 done
 
 (*Not of the general form for such rules; ~EX has become ALL~ *)
 lemma oexCI:
    "[| ALL x<A. ~P(x) ==> P(a);  a<A |] ==> EX x<A. P(x)"
-apply (simp add: oex_def, blast) 
+apply (simp add: oex_def, blast)
 done
 
 lemma oexE [elim!]:
     "[| EX x<A. P(x);  !!x. [| x<A; P(x) |] ==> Q |] ==> Q"
-apply (simp add: oex_def, blast) 
+apply (simp add: oex_def, blast)
 done
 
 lemma oex_cong [cong]:
-    "[| a=a';  !!x. x<a' ==> P(x) <-> P'(x) |] 
+    "[| a=a';  !!x. x<a' ==> P(x) <-> P'(x) |]
      ==> oex(a, %x. P(x)) <-> oex(a', %x. P'(x))"
 apply (simp add: oex_def cong add: conj_cong)
 done
@@ -182,11 +180,11 @@
     "[| i=j;  !!x. x<j ==> C(x)=D(x) |] ==> (UN x<i. C(x)) = (UN x<j. D(x))"
 by (simp add: OUnion_def lt_def OUN_iff)
 
-lemma lt_induct: 
+lemma lt_induct:
     "[| i<k;  !!x.[| x<k;  ALL y<x. P(y) |] ==> P(x) |]  ==>  P(i)"
 apply (simp add: lt_def oall_def)
-apply (erule conjE) 
-apply (erule Ord_induct, assumption, blast) 
+apply (erule conjE)
+apply (erule Ord_induct, assumption, blast)
 done
 
 
@@ -211,7 +209,8 @@
   "ALL x[M]. P"  == "rall(M, %x. P)"
   "EX x[M]. P"   == "rex(M, %x. P)"
 
-(*** Relativized universal quantifier ***)
+
+subsubsection{*Relativized universal quantifier*}
 
 lemma rallI [intro!]: "[| !!x. M(x) ==> P(x) |] ==> ALL x[M]. P(x)"
 by (simp add: rall_def)
@@ -220,9 +219,9 @@
 by (simp add: rall_def)
 
 (*Instantiates x first: better for automatic theorem proving?*)
-lemma rev_rallE [elim]: 
+lemma rev_rallE [elim]:
     "[| ALL x[M]. P(x);  ~ M(x) ==> Q;  P(x) ==> Q |] ==> Q"
-by (simp add: rall_def, blast) 
+by (simp add: rall_def, blast)
 
 lemma rallE: "[| ALL x[M]. P(x);  P(x) ==> Q;  ~ M(x) ==> Q |] ==> Q"
 by blast
@@ -233,11 +232,12 @@
 
 (*Congruence rule for rewriting*)
 lemma rall_cong [cong]:
-    "(!!x. M(x) ==> P(x) <-> P'(x)) 
+    "(!!x. M(x) ==> P(x) <-> P'(x))
      ==> rall(M, %x. P(x)) <-> rall(M, %x. P'(x))"
 by (simp add: rall_def)
 
-(*** Relativized existential quantifier ***)
+
+subsubsection{*Relativized existential quantifier*}
 
 lemma rexI [intro]: "[| P(x); M(x) |] ==> EX x[M]. P(x)"
 by (simp add: rex_def, blast)
@@ -258,7 +258,7 @@
 by (simp add: rex_def)
 
 lemma rex_cong [cong]:
-    "(!!x. M(x) ==> P(x) <-> P'(x)) 
+    "(!!x. M(x) ==> P(x) <-> P'(x))
      ==> rex(M, %x. P(x)) <-> rex(M, %x. P'(x))"
 by (simp add: rex_def cong: conj_cong)
 
@@ -277,7 +277,7 @@
      "(ALL x[M]. P(x) & Q)   <-> (ALL x[M]. P(x)) & ((ALL x[M]. False) | Q)"
      "(ALL x[M]. P(x) | Q)   <-> ((ALL x[M]. P(x)) | Q)"
      "(ALL x[M]. P(x) --> Q) <-> ((EX x[M]. P(x)) --> Q)"
-     "(~(ALL x[M]. P(x))) <-> (EX x[M]. ~P(x))" 
+     "(~(ALL x[M]. P(x))) <-> (EX x[M]. ~P(x))"
 by blast+
 
 lemma rall_simps2:
@@ -312,7 +312,7 @@
 by blast
 
 
-(** One-point rule for bounded quantifiers: see HOL/Set.ML **)
+subsubsection{*One-point rule for bounded quantifiers*}
 
 lemma rex_triv_one_point1 [simp]: "(EX x[M]. x=a) <-> ( M(a))"
 by blast
@@ -333,6 +333,20 @@
 by blast
 
 
+subsubsection{*Sets as Classes*}
+
+constdefs setclass :: "[i,i] => o"       ("**_")
+   "setclass(A,x) == x : A"
+
+declare setclass_def [simp]
+
+lemma rall_setclass_is_ball [simp]: "(\<forall>x[**A]. P(x)) <-> (\<forall>x\<in>A. P(x))"
+by auto
+
+lemma rex_setclass_is_bex [simp]: "(\<exists>x[**A]. P(x)) <-> (\<exists>x\<in>A. P(x))"
+by auto
+
+
 ML
 {*
 val oall_def = thm "oall_def"
@@ -370,7 +384,7 @@
 
 val Ord_atomize =
     atomize ([("OrdQuant.oall", [ospec]),("OrdQuant.rall", [rspec])]@
-                 ZF_conn_pairs, 
+                 ZF_conn_pairs,
              ZF_mem_pairs);
 simpset_ref() := simpset() setmksimps (map mk_eq o Ord_atomize o gen_all);
 *}
@@ -391,7 +405,7 @@
 val all_pattern = Thm.read_cterm (Theory.sign_of (the_context ()))
                                  ("ALL x[M]. P(x) --> Q(x)", FOLogic.oT)
 
-val prove_rall_tac = rewtac rall_def THEN 
+val prove_rall_tac = rewtac rall_def THEN
                      Quantifier1.prove_one_point_all_tac;
 
 val rearrange_ball = Quantifier1.rearrange_ball prove_rall_tac;