# HG changeset patch # User paulson # Date 1026204293 -7200 # Node ID 23de7b3af4536039af99266e4a535aceeb075d9d # Parent 3f475e54875c62255e423933b9a97f07c60bbe1e More Separation proofs diff -r 3f475e54875c -r 23de7b3af453 src/ZF/Constructible/Relative.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, \x. \x,a\ \ r & \x,b\ \ r & f`x \ g`x)" + --{*for well-founded recursion*} + "[| M(r); M(f); M(g); M(a); M(b) |] + ==> separation(M, + \x. \xa[M]. \xb[M]. + pair(M,x,a,xa) & xa \ r & pair(M,x,b,xb) & xb \ r & + (\fx[M]. \gx[M]. fun_apply(M,f,x,fx) & fun_apply(M,g,x,gx) & + fx \ gx))" lemma (in M_axioms) cartprod_iff_lemma: "[| M(C); \u[M]. u \ C <-> (\x\A. \y\B. u = {{x}, {x,y}}); diff -r 3f475e54875c -r 23de7b3af453 src/ZF/Constructible/Separation.thy --- 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[\x. x\A --> ~(\y[L]. \g[L]. diff -r 3f475e54875c -r 23de7b3af453 src/ZF/Constructible/WFrec.thy --- 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 \ r -`` {a} \ range(f); g \ r -`` {b} \ range(g); + M(r); M(f); M(g); M(a); M(b) |] + ==> separation(M, \x. \ (\x, a\ \ r \ \x, b\ \ r \ 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)+