# HG changeset patch # User paulson # Date 1031669237 -7200 # Node ID 7d6c9817c4329e20da104f1fba17839dcddfcbd8 # Parent 5b71e1408ac451749558b26a886af20b6c840695 tweaks diff -r 5b71e1408ac4 -r 7d6c9817c432 src/ZF/Constructible/L_axioms.thy --- a/src/ZF/Constructible/L_axioms.thy Mon Sep 09 17:28:29 2002 +0200 +++ b/src/ZF/Constructible/L_axioms.thy Tue Sep 10 16:47:17 2002 +0200 @@ -20,26 +20,34 @@ apply (blast intro: zero_in_Lset) done -lemma upair_ax: "upair_ax(L)" +theorem upair_ax: "upair_ax(L)" apply (simp add: upair_ax_def upair_def, clarify) apply (rule_tac x="{x,y}" in rexI) apply (simp_all add: doubleton_in_L) done -lemma Union_ax: "Union_ax(L)" +theorem Union_ax: "Union_ax(L)" apply (simp add: Union_ax_def big_union_def, clarify) apply (rule_tac x="Union(x)" in rexI) apply (simp_all add: Union_in_L, auto) apply (blast intro: transL) done -lemma power_ax: "power_ax(L)" +theorem power_ax: "power_ax(L)" apply (simp add: power_ax_def powerset_def Relative.subset_def, clarify) apply (rule_tac x="{y \ Pow(x). L(y)}" in rexI) apply (simp_all add: LPow_in_L, auto) apply (blast intro: transL) done +text{*We don't actually need @{term L} to satisfy the foundation axiom.*} +theorem foundation_ax: "foundation_ax(L)" +apply (simp add: foundation_ax_def) +apply (rule rallI) +apply (cut_tac A=x in foundation) +apply (blast intro: transL) +done + subsection{*For L to satisfy Replacement *} (*Can't move these to Formula unless the definition of univalent is moved @@ -65,7 +73,7 @@ apply (blast intro: L_I Lset_in_Lset_succ) done -lemma replacement: "replacement(L,P)" +theorem replacement: "replacement(L,P)" apply (simp add: replacement_def, clarify) apply (frule LReplace_in_L, assumption+, clarify) apply (rule_tac x=Y in rexI) diff -r 5b71e1408ac4 -r 7d6c9817c432 src/ZF/Constructible/Reflection.thy --- a/src/ZF/Constructible/Reflection.thy Mon Sep 09 17:28:29 2002 +0200 +++ b/src/ZF/Constructible/Reflection.thy Tue Sep 10 16:47:17 2002 +0200 @@ -37,16 +37,16 @@ and Mset_cont : "cont_Ord(Mset)" and Pair_in_Mset : "[| x \ Mset(a); y \ Mset(a); Limit(a) |] ==> \ Mset(a)" - defines "M(x) == \a. Ord(a) \ x \ Mset(a)" - and "Reflects(Cl,P,Q) == Closed_Unbounded(Cl) \ + defines "M(x) == \a. Ord(a) & x \ Mset(a)" + and "Reflects(Cl,P,Q) == Closed_Unbounded(Cl) & (\a. Cl(a) --> (\x\Mset(a). P(x) <-> Q(a,x)))" fixes F0 --{*ordinal for a specific value @{term y}*} fixes FF --{*sup over the whole level, @{term "y\Mset(a)"}*} fixes ClEx --{*Reflecting ordinals for the formula @{term "\z. P"}*} - defines "F0(P,y) == \b. (\z. M(z) \ P()) --> + defines "F0(P,y) == \b. (\z. M(z) & P()) --> (\z\Mset(b). P())" and "FF(P) == \a. \y\Mset(a). F0(P,y)" - and "ClEx(P,a) == Limit(a) \ normalize(FF(P),a) = a" + and "ClEx(P,a) == Limit(a) & normalize(FF(P),a) = a" lemma (in reflection) Mset_mono: "i\j ==> Mset(i) <= Mset(j)" apply (insert Mset_mono_le) @@ -56,7 +56,7 @@ text{*Awkward: we need a version of @{text ClEx_def} as an equality at the level of classes, which do not really exist*} lemma (in reflection) ClEx_eq: - "ClEx(P) == \a. Limit(a) \ normalize(FF(P),a) = a" + "ClEx(P) == \a. Limit(a) & normalize(FF(P),a) = a" by (simp add: ClEx_def [symmetric]) @@ -72,26 +72,26 @@ theorem (in reflection) And_reflection [intro]: "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] - ==> Reflects(\a. Cl(a) \ C'(a), \x. P(x) \ P'(x), - \a x. Q(a,x) \ Q'(a,x))" + ==> Reflects(\a. Cl(a) & C'(a), \x. P(x) & P'(x), + \a x. Q(a,x) & Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) theorem (in reflection) Or_reflection [intro]: "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] - ==> Reflects(\a. Cl(a) \ C'(a), \x. P(x) \ P'(x), - \a x. Q(a,x) \ Q'(a,x))" + ==> Reflects(\a. Cl(a) & C'(a), \x. P(x) | P'(x), + \a x. Q(a,x) | Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) theorem (in reflection) Imp_reflection [intro]: "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] - ==> Reflects(\a. Cl(a) \ C'(a), + ==> Reflects(\a. Cl(a) & C'(a), \x. P(x) --> P'(x), \a x. Q(a,x) --> Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) theorem (in reflection) Iff_reflection [intro]: "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] - ==> Reflects(\a. Cl(a) \ C'(a), + ==> Reflects(\a. Cl(a) & C'(a), \x. P(x) <-> P'(x), \a x. Q(a,x) <-> Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) @@ -155,7 +155,7 @@ lemma (in ex_reflection) ClEx_upward: "[| z\Mset(a); y\Mset(a); Q(a,); Cl(a); ClEx(P,a) |] - ==> \z. M(z) \ P()" + ==> \z. M(z) & P()" apply (simp add: ClEx_def M_def) apply (blast dest: Cl_reflects intro: Limit_is_Ord Pair_in_Mset) @@ -164,7 +164,7 @@ text{*Class @{text ClEx} indeed consists of reflecting ordinals...*} lemma (in ex_reflection) ZF_ClEx_iff: "[| y\Mset(a); Cl(a); ClEx(P,a) |] - ==> (\z. M(z) \ P()) <-> (\z\Mset(a). Q(a,))" + ==> (\z. M(z) & P()) <-> (\z\Mset(a). Q(a,))" by (blast intro: dest: ClEx_downward ClEx_upward) text{*...and it is closed and unbounded*} @@ -181,7 +181,7 @@ lemma (in reflection) ClEx_iff: "[| y\Mset(a); Cl(a); ClEx(P,a); !!a. [| Cl(a); Ord(a) |] ==> \x\Mset(a). P(x) <-> Q(a,x) |] - ==> (\z. M(z) \ P()) <-> (\z\Mset(a). Q(a,))" + ==> (\z. M(z) & P()) <-> (\z\Mset(a). Q(a,))" apply (unfold ClEx_def FF_def F0_def M_def) apply (rule ex_reflection.ZF_ClEx_iff [OF ex_reflection.intro, OF reflection.intro ex_reflection_axioms.intro, @@ -209,8 +209,8 @@ lemma (in reflection) Ex_reflection_0: "Reflects(Cl,P0,Q0) - ==> Reflects(\a. Cl(a) \ ClEx(P0,a), - \x. \z. M(z) \ P0(), + ==> Reflects(\a. Cl(a) & ClEx(P0,a), + \x. \z. M(z) & P0(), \a x. \z\Mset(a). Q0(a,))" apply (simp add: Reflects_def) apply (intro conjI Closed_Unbounded_Int) @@ -221,7 +221,7 @@ lemma (in reflection) All_reflection_0: "Reflects(Cl,P0,Q0) - ==> Reflects(\a. Cl(a) \ ClEx(\x.~P0(x), a), + ==> Reflects(\a. Cl(a) & ClEx(\x.~P0(x), a), \x. \z. M(z) --> P0(), \a x. \z\Mset(a). Q0(a,))" apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not) @@ -231,15 +231,15 @@ theorem (in reflection) Ex_reflection [intro]: "Reflects(Cl, \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))) - ==> Reflects(\a. Cl(a) \ ClEx(\x. P(fst(x),snd(x)), a), - \x. \z. M(z) \ P(x,z), + ==> Reflects(\a. Cl(a) & ClEx(\x. P(fst(x),snd(x)), a), + \x. \z. M(z) & P(x,z), \a x. \z\Mset(a). Q(a,x,z))" by (rule Ex_reflection_0 [of _ " \x. P(fst(x),snd(x))" "\a x. Q(a,fst(x),snd(x))", simplified]) theorem (in reflection) All_reflection [intro]: "Reflects(Cl, \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))) - ==> Reflects(\a. Cl(a) \ ClEx(\x. ~P(fst(x),snd(x)), a), + ==> Reflects(\a. Cl(a) & ClEx(\x. ~P(fst(x),snd(x)), a), \x. \z. M(z) --> P(x,z), \a x. \z\Mset(a). Q(a,x,z))" by (rule All_reflection_0 [of _ "\x. P(fst(x),snd(x))" @@ -249,14 +249,14 @@ theorem (in reflection) Rex_reflection [intro]: "Reflects(Cl, \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))) - ==> Reflects(\a. Cl(a) \ ClEx(\x. P(fst(x),snd(x)), a), + ==> Reflects(\a. Cl(a) & ClEx(\x. P(fst(x),snd(x)), a), \x. \z[M]. P(x,z), \a x. \z\Mset(a). Q(a,x,z))" by (unfold rex_def, blast) theorem (in reflection) Rall_reflection [intro]: "Reflects(Cl, \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))) - ==> Reflects(\a. Cl(a) \ ClEx(\x. ~P(fst(x),snd(x)), a), + ==> Reflects(\a. Cl(a) & ClEx(\x. ~P(fst(x),snd(x)), a), \x. \z[M]. P(x,z), \a x. \z\Mset(a). Q(a,x,z))" by (unfold rall_def, blast) @@ -272,7 +272,7 @@ proof state.*} lemma (in reflection) "Reflects(?Cl, - \x. \y. M(y) \ x \ y, + \x. \y. M(y) & x \ y, \a x. \y\Mset(a). x \ y)" by fast @@ -280,8 +280,8 @@ in the class of reflecting ordinals. The @{term "Ord(a)"} is redundant, though harmless.*} lemma (in reflection) - "Reflects(\a. Ord(a) \ ClEx(\x. fst(x) \ snd(x), a), - \x. \y. M(y) \ x \ y, + "Reflects(\a. Ord(a) & ClEx(\x. fst(x) \ snd(x), a), + \x. \y. M(y) & x \ y, \a x. \y\Mset(a). x \ y)" by fast @@ -289,31 +289,31 @@ text{*Example 2*} lemma (in reflection) "Reflects(?Cl, - \x. \y. M(y) \ (\z. M(z) --> z \ x --> z \ y), + \x. \y. M(y) & (\z. M(z) --> z \ x --> z \ y), \a x. \y\Mset(a). \z\Mset(a). z \ x --> z \ y)" by fast text{*Example 2'. We give the reflecting class explicitly. *} lemma (in reflection) "Reflects - (\a. (Ord(a) \ - ClEx(\x. ~ (snd(x) \ fst(fst(x)) --> snd(x) \ snd(fst(x))), a)) \ + (\a. (Ord(a) & + ClEx(\x. ~ (snd(x) \ fst(fst(x)) --> snd(x) \ snd(fst(x))), a)) & ClEx(\x. \z. M(z) --> z \ fst(x) --> z \ snd(x), a), - \x. \y. M(y) \ (\z. M(z) --> z \ x --> z \ y), + \x. \y. M(y) & (\z. M(z) --> z \ x --> z \ y), \a x. \y\Mset(a). \z\Mset(a). z \ x --> z \ y)" by fast text{*Example 2''. We expand the subset relation.*} lemma (in reflection) "Reflects(?Cl, - \x. \y. M(y) \ (\z. M(z) --> (\w. M(w) --> w\z --> w\x) --> z\y), + \x. \y. M(y) & (\z. M(z) --> (\w. M(w) --> w\z --> w\x) --> z\y), \a x. \y\Mset(a). \z\Mset(a). (\w\Mset(a). w\z --> w\x) --> z\y)" by fast text{*Example 2'''. Single-step version, to reveal the reflecting class.*} lemma (in reflection) "Reflects(?Cl, - \x. \y. M(y) \ (\z. M(z) --> z \ x --> z \ y), + \x. \y. M(y) & (\z. M(z) --> z \ x --> z \ y), \a x. \y\Mset(a). \z\Mset(a). z \ x --> z \ y)" apply (rule Ex_reflection) txt{* @@ -333,21 +333,21 @@ if @{term P} is quantifier-free, since it is not being relativized.*} lemma (in reflection) "Reflects(?Cl, - \x. \y. M(y) \ (\z. M(z) --> z \ y <-> z \ x \ P(z)), - \a x. \y\Mset(a). \z\Mset(a). z \ y <-> z \ x \ P(z))" + \x. \y. M(y) & (\z. M(z) --> z \ y <-> z \ x & P(z)), + \a x. \y\Mset(a). \z\Mset(a). z \ y <-> z \ x & P(z))" by fast text{*Example 3'*} lemma (in reflection) "Reflects(?Cl, - \x. \y. M(y) \ y = Collect(x,P), + \x. \y. M(y) & y = Collect(x,P), \a x. \y\Mset(a). y = Collect(x,P))"; by fast text{*Example 3''*} lemma (in reflection) "Reflects(?Cl, - \x. \y. M(y) \ y = Replace(x,P), + \x. \y. M(y) & y = Replace(x,P), \a x. \y\Mset(a). y = Replace(x,P))"; by fast @@ -355,7 +355,7 @@ to be relativized.*} lemma (in reflection) "Reflects(?Cl, - \A. 0\A --> (\f. M(f) \ f \ (\X \ A. X)), + \A. 0\A --> (\f. M(f) & f \ (\X \ A. X)), \a A. 0\A --> (\f\Mset(a). f \ (\X \ A. X)))" by fast diff -r 5b71e1408ac4 -r 7d6c9817c432 src/ZF/Constructible/Relative.thy --- a/src/ZF/Constructible/Relative.thy Mon Sep 09 17:28:29 2002 +0200 +++ b/src/ZF/Constructible/Relative.thy Tue Sep 10 16:47:17 2002 +0200 @@ -244,13 +244,15 @@ \x[M]. \y[M]. (\z[M]. z \ x <-> z \ y) --> x=y" separation :: "[i=>o, i=>o] => o" - --{*Big problem: the formula @{text P} should only involve parameters - belonging to @{text M}. Don't see how to enforce that.*} + --{*The formula @{text P} should only involve parameters + belonging to @{text M}. But we can't prove separation as a scheme + anyway. Every instance that we need must individually be assumed + and later proved.*} "separation(M,P) == \z[M]. \y[M]. \x[M]. x \ y <-> x \ z & P(x)" upair_ax :: "(i=>o) => o" - "upair_ax(M) == \x y. M(x) --> M(y) --> (\z[M]. upair(M,x,y,z))" + "upair_ax(M) == \x[M]. \y[M]. \z[M]. upair(M,x,y,z)" Union_ax :: "(i=>o) => o" "Union_ax(M) == \x[M]. \z[M]. big_union(M,x,z)" @@ -260,7 +262,7 @@ univalent :: "[i=>o, i, [i,i]=>o] => o" "univalent(M,A,P) == - (\x[M]. x\A --> (\y z. M(y) --> M(z) --> P(x,y) & P(x,z) --> y=z))" + (\x[M]. x\A --> (\y[M]. \z[M]. P(x,y) & P(x,z) --> y=z))" replacement :: "[i=>o, [i,i]=>o] => o" "replacement(M,P) == @@ -274,7 +276,7 @@ foundation_ax :: "(i=>o) => o" "foundation_ax(M) == - \x[M]. (\y\x. M(y)) --> (\y[M]. y\x & ~(\z[M]. z\x & z \ y))" + \x[M]. (\y[M]. y\x) --> (\y[M]. y\x & ~(\z[M]. z\x & z \ y))" subsection{*A trivial consistency proof for $V_\omega$ *}