More Separation proofs
authorpaulson
Tue, 09 Jul 2002 10:44:53 +0200
changeset 13319 23de7b3af453
parent 13318 3f475e54875c
child 13320 2c6ee189ae63
More Separation proofs
src/ZF/Constructible/Relative.thy
src/ZF/Constructible/Separation.thy
src/ZF/Constructible/WFrec.thy
--- a/src/ZF/Constructible/Relative.thy	Tue Jul 09 10:44:41 2002 +0200
+++ b/src/ZF/Constructible/Relative.thy	Tue Jul 09 10:44:53 2002 +0200
@@ -764,9 +764,13 @@
 	     ordinal(M,x) & pair(M,a,x,z) & membership(M,x,mx) & 
 	     pred_set(M,A,a,r,par) & order_isomorphism(M,par,r,x,mx,g))"
   and is_recfun_separation:
-     --{*for well-founded recursion.  NEEDS RELATIVIZATION*}
-     "[| M(A); M(f); M(g); M(a); M(b) |] 
-     ==> separation(M, \<lambda>x. \<langle>x,a\<rangle> \<in> r & \<langle>x,b\<rangle> \<in> r & f`x \<noteq> g`x)"
+     --{*for well-founded recursion*}
+     "[| M(r); M(f); M(g); M(a); M(b) |] 
+     ==> separation(M, 
+            \<lambda>x. \<exists>xa[M]. \<exists>xb[M]. 
+                pair(M,x,a,xa) & xa \<in> r & pair(M,x,b,xb) & xb \<in> r & 
+                (\<exists>fx[M]. \<exists>gx[M]. fun_apply(M,f,x,fx) & fun_apply(M,g,x,gx) & 
+                                   fx \<noteq> gx))"
 
 lemma (in M_axioms) cartprod_iff_lemma:
      "[| M(C);  \<forall>u[M]. u \<in> C <-> (\<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}}); 
--- a/src/ZF/Constructible/Separation.thy	Tue Jul 09 10:44:41 2002 +0200
+++ b/src/ZF/Constructible/Separation.thy	Tue Jul 09 10:44:53 2002 +0200
@@ -349,7 +349,7 @@
 done
 
 
-subsection{*Separation for @{term "well_ord_iso"}*}
+subsection{*Separation for a Theorem about @{term "obase"}*}
 
 lemma obase_equals_reflects:
   "REFLECTS[\<lambda>x. x\<in>A --> ~(\<exists>y[L]. \<exists>g[L]. 
--- a/src/ZF/Constructible/WFrec.thy	Tue Jul 09 10:44:41 2002 +0200
+++ b/src/ZF/Constructible/WFrec.thy	Tue Jul 09 10:44:53 2002 +0200
@@ -64,6 +64,13 @@
 apply (blast intro: sym)+
 done
 
+lemma (in M_axioms) is_recfun_separation':
+    "[| f \<in> r -`` {a} \<rightarrow> range(f); g \<in> r -`` {b} \<rightarrow> range(g);
+        M(r); M(f); M(g); M(a); M(b) |] 
+     ==> separation(M, \<lambda>x. \<not> (\<langle>x, a\<rangle> \<in> r \<longrightarrow> \<langle>x, b\<rangle> \<in> r \<longrightarrow> f ` x = g ` x))"
+apply (insert is_recfun_separation [of r f g a b]) 
+apply (simp add: typed_apply_abs vimage_singleton_iff)
+done
 
 text{*Stated using @{term "trans(r)"} rather than
       @{term "transitive_rel(M,A,r)"} because the latter rewrites to
@@ -82,7 +89,7 @@
 apply (simp add: is_recfun_def)
 apply (erule_tac a=x in wellfounded_induct, assumption+)
 txt{*Separation to justify the induction*}
- apply (force intro: is_recfun_separation)
+ apply (blast intro: is_recfun_separation') 
 txt{*Now the inductive argument itself*}
 apply clarify 
 apply (erule ssubst)+