# HG changeset patch # User paulson # Date 1028119187 -7200 # Node ID 8fd1d803a3d3703568830fa5c3a6e460dfb39c47 # Parent 05631e8f0258ebfd5af3e3933e8b8f557cc74dbf tweaks involving Separation diff -r 05631e8f0258 -r 8fd1d803a3d3 src/ZF/Constructible/Relative.thy --- a/src/ZF/Constructible/Relative.thy Wed Jul 31 14:34:08 2002 +0200 +++ b/src/ZF/Constructible/Relative.thy Wed Jul 31 14:39:47 2002 +0200 @@ -2,10 +2,6 @@ theory Relative = Main: -(*????????????????for IFOL.thy????????????????*) -lemma eq_commute: "a=b <-> b=a" -by blast - subsection{* Relativized versions of standard set-theoretic concepts *} constdefs @@ -33,17 +29,20 @@ "successor(M,a,z) == is_cons(M,a,a,z)" number1 :: "[i=>o,i] => o" - "number1(M,a) == (\x[M]. empty(M,x) & successor(M,x,a))" + "number1(M,a) == \x[M]. empty(M,x) & successor(M,x,a)" number2 :: "[i=>o,i] => o" - "number2(M,a) == (\x[M]. number1(M,x) & successor(M,x,a))" + "number2(M,a) == \x[M]. number1(M,x) & successor(M,x,a)" number3 :: "[i=>o,i] => o" - "number3(M,a) == (\x[M]. number2(M,x) & successor(M,x,a))" + "number3(M,a) == \x[M]. number2(M,x) & successor(M,x,a)" powerset :: "[i=>o,i,i] => o" "powerset(M,A,z) == \x[M]. x \ z <-> subset(M,x,A)" + is_Collect :: "[i=>o,i,i=>o,i] => o" + "is_Collect(M,A,P,z) == \x[M]. x \ z <-> x \ A & P(x)" + inter :: "[i=>o,i,i,i] => o" "inter(M,a,b,z) == \x[M]. x \ z <-> x \ a & x \ b" @@ -85,11 +84,11 @@ is_domain :: "[i=>o,i,i] => o" "is_domain(M,r,z) == - \x[M]. (x \ z <-> (\w[M]. w\r & (\y[M]. pair(M,x,y,w))))" + \x[M]. x \ z <-> (\w[M]. w\r & (\y[M]. pair(M,x,y,w)))" image :: "[i=>o,i,i,i] => o" "image(M,r,A,z) == - \y[M]. (y \ z <-> (\w[M]. w\r & (\x[M]. x\A & pair(M,x,y,w))))" + \y[M]. y \ z <-> (\w[M]. w\r & (\x[M]. x\A & pair(M,x,y,w)))" is_range :: "[i=>o,i,i] => o" --{*the cleaner @@ -97,12 +96,12 @@ unfortunately needs an instance of separation in order to prove @{term "M(converse(r))"}.*} "is_range(M,r,z) == - \y[M]. (y \ z <-> (\w[M]. w\r & (\x[M]. pair(M,x,y,w))))" + \y[M]. y \ z <-> (\w[M]. w\r & (\x[M]. pair(M,x,y,w)))" is_field :: "[i=>o,i,i] => o" "is_field(M,r,z) == - \dr[M]. is_domain(M,r,dr) & - (\rr[M]. is_range(M,r,rr) & union(M,dr,rr,z))" + \dr[M]. \rr[M]. is_domain(M,r,dr) & is_range(M,r,rr) & + union(M,dr,rr,z)" is_relation :: "[i=>o,i] => o" "is_relation(M,r) == @@ -594,6 +593,16 @@ apply (blast dest: transM) done +lemma separation_iff: + "separation(M,P) <-> (\z[M]. \y[M]. is_Collect(M,z,P,y))" +by (simp add: separation_def is_Collect_def) + +lemma (in M_triv_axioms) Collect_abs [simp]: + "[| M(A); M(z) |] ==> is_Collect(M,A,P,z) <-> z = Collect(A,P)" +apply (simp add: is_Collect_def) +apply (blast intro!: equalityI dest: transM) +done + text{*Probably the premise and conclusion are equivalent*} lemma (in M_triv_axioms) strong_replacementI [rule_format]: "[| \A[M]. separation(M, %u. \x[M]. x\A & P(x,u)) |] @@ -857,6 +866,8 @@ locale M_axioms = M_triv_axioms + assumes Inter_separation: "M(A) ==> separation(M, \x. \y[M]. y\A --> x\y)" + and Diff_separation: + "M(B) ==> separation(M, \x. x \ B)" and cartprod_separation: "[| M(A); M(B) |] ==> separation(M, \z. \x[M]. x\A & (\y[M]. y\B & pair(M,x,y,z)))" @@ -1195,6 +1206,61 @@ apply (frule Inter_closed, force+) done +lemma (in M_axioms) Diff_closed [intro,simp]: + "[|M(A); M(B)|] ==> M(A-B)" +by (insert Diff_separation, simp add: Diff_def) + +subsubsection{*Some Facts About Separation Axioms*} + +lemma (in M_axioms) separation_conj: + "[|separation(M,P); separation(M,Q)|] ==> separation(M, \z. P(z) & Q(z))" +by (simp del: separation_closed + add: separation_iff Collect_Int_Collect_eq [symmetric]) + +(*???equalities*) +lemma Collect_Un_Collect_eq: + "Collect(A,P) Un Collect(A,Q) = Collect(A, %x. P(x) | Q(x))" +by blast + +lemma Diff_Collect_eq: + "A - Collect(A,P) = Collect(A, %x. ~ P(x))" +by blast + +lemma (in M_triv_axioms) Collect_rall_eq: + "M(Y) ==> Collect(A, %x. \y[M]. y\Y --> P(x,y)) = + (if Y=0 then A else (\y \ Y. {x \ A. P(x,y)}))" +apply simp +apply (blast intro!: equalityI dest: transM) +done + +lemma (in M_axioms) separation_disj: + "[|separation(M,P); separation(M,Q)|] ==> separation(M, \z. P(z) | Q(z))" +by (simp del: separation_closed + add: separation_iff Collect_Un_Collect_eq [symmetric]) + +lemma (in M_axioms) separation_neg: + "separation(M,P) ==> separation(M, \z. ~P(z))" +by (simp del: separation_closed + add: separation_iff Diff_Collect_eq [symmetric]) + +lemma (in M_axioms) separation_imp: + "[|separation(M,P); separation(M,Q)|] + ==> separation(M, \z. P(z) --> Q(z))" +by (simp add: separation_neg separation_disj not_disj_iff_imp [symmetric]) + +text{*This result is a hint of how little can be done without the Reflection + Theorem. The quantifier has to be bounded by a set. We also need another + instance of Separation!*} +lemma (in M_axioms) separation_rall: + "[|M(Y); \y[M]. separation(M, \x. P(x,y)); + \z[M]. strong_replacement(M, \x y. y = {u \ z . P(u,x)})|] + ==> separation(M, \x. \y[M]. y\Y --> P(x,y))" +apply (simp del: separation_closed rall_abs + add: separation_iff Collect_rall_eq) +apply (blast intro!: Inter_closed RepFun_closed dest: transM) +done + + subsubsection{*Functions and function space*} text{*M contains all finite functions*}