# HG changeset patch # User paulson # Date 1025794794 -7200 # Node ID b4f370679c6516ef57e03fa93e49038003617d66 # Parent e4ae0732e2beb6222de90208fedc7dc271bb2133 Constructible: some separation axioms diff -r e4ae0732e2be -r b4f370679c65 src/ZF/Constructible/Formula.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 \ list(A) @@ -125,6 +125,11 @@ ==> (x\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 \ 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 \ 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 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 \ list(A)|] ==> x \ 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 \ nat; y \ nat |] ==> subset_fm(x,y) \ formula" +by (simp add: subset_fm_def) + +lemma arity_subset_fm [simp]: + "[| x \ nat; y \ nat |] ==> arity(subset_fm(x,y)) = succ(x) \ succ(y)" +by (simp add: subset_fm_def succ_Un_distrib [symmetric]) + +lemma sats_subset_fm [simp]: + "[|x < length(env); y \ nat; env \ list(A); Transset(A)|] + ==> sats(A, subset_fm(x,y), env) <-> nth(x,env) \ 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 \ nat ==> transset_fm(x) \ formula" +by (simp add: transset_fm_def) + +lemma arity_transset_fm [simp]: + "x \ 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 \ 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 \ nat ==> ordinal_fm(x) \ formula" +by (simp add: ordinal_fm_def) + +lemma arity_ordinal_fm [simp]: + "x \ 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 \ 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 \ nat; y \ nat |] ==> subset_fm(x,y) \ formula" -by (simp add: subset_fm_def) - -lemma arity_subset_fm [simp]: - "[| x \ nat; y \ nat |] ==> arity(subset_fm(x,y)) = succ(x) \ succ(y)" -by (simp add: subset_fm_def succ_Un_distrib [symmetric]) - -lemma sats_subset_fm [simp]: - "[|x < length(env); y \ nat; env \ list(A); Transset(A)|] - ==> sats(A, subset_fm(x,y), env) <-> nth(x,env) \ 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 \ nat ==> transset_fm(x) \ formula" -by (simp add: transset_fm_def) - -lemma arity_transset_fm [simp]: - "x \ 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 \ 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 \ nat ==> ordinal_fm(x) \ formula" -by (simp add: ordinal_fm_def) - -lemma arity_ordinal_fm [simp]: - "x \ 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 \ 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 \ A. Ord(x)} \ DPow(A)" diff -r e4ae0732e2be -r b4f370679c65 src/ZF/Constructible/L_axioms.thy --- 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 \ nat; y \ nat; z \ nat |] ==> upair_fm(x,y,z) \ formula" +by (simp add: upair_fm_def) + +lemma arity_upair_fm [simp]: + "[| x \ nat; y \ nat; z \ nat |] + ==> arity(upair_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" +by (simp add: upair_fm_def succ_Un_distrib [symmetric] Un_ac) + +lemma sats_upair_fm [simp]: + "[| x \ nat; y \ nat; z \ nat; env \ 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 \ nat; j \ nat; k \ nat; env \ 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 \ nat; y \ nat; z < length(env); env \ 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 \ nat; y \ nat; z \ nat |] ==> pair_fm(x,y,z) \ formula" +by (simp add: pair_fm_def) + +lemma arity_pair_fm [simp]: + "[| x \ nat; y \ nat; z \ nat |] + ==> arity(pair_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" +by (simp add: pair_fm_def succ_Un_distrib [symmetric] Un_ac) + +lemma sats_pair_fm [simp]: + "[| x \ nat; y \ nat; z \ nat; env \ 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 \ nat; j \ nat; k \ nat; env \ 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 \ nat|] ==> nth(succ(n), Cons(a,l)) = x" +by simp + + +lemma Collect_conj_in_DPow: + "[| {x\A. P(x)} \ DPow(A); {x\A. Q(x)} \ DPow(A) |] + ==> {x\A. P(x) & Q(x)} \ DPow(A)" +by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric]) + +lemma Collect_conj_in_DPow_Lset: + "[|z \ Lset(j); {x \ Lset(j). P(x)} \ DPow(Lset(j))|] + ==> {x \ Lset(j). x \ z & P(x)} \ 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: + "(\z. L(z) ==> L({x \ z . P(x)})) ==> separation(L, \x. P(x))" +apply (unfold separation_def, clarify) +apply (rule_tac x="{x\z. P(x)}" in rexI) +apply simp_all +done + +text{*Reduces the original comprehension to the reflected one*} +lemma reflection_imp_L_separation: + "[| \x\Lset(j). P(x) <-> Q(x); + {x \ Lset(j) . Q(x)} \ DPow(Lset(j)); + Ord(j); z \ Lset(j)|] ==> L({x \ z . P(x)})" +apply (rule_tac i = "succ(j)" in L_I) + prefer 2 apply simp +apply (subgoal_tac "{x \ z. P(x)} = {x \ Lset(j). x \ 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, \x. \y[L]. y\A --> x \ y, + \i x. \y\Lset(i). y\A --> x \ y)" +by fast + +lemma Inter_separation: + "L(A) ==> separation(L, \x. \y[L]. y\A --> x\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, \x. upair(L,f(x),g(x),h(x)), + \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, \x. pair(L,f(x),g(x),h(x)), + \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, \z. \x[L]. x\A & (\y[L]. y\B & pair(L,x,y,z)), + \i z. \x\Lset(i). x\A & (\y\Lset(i). y\B & + pair(**Lset(i),x,y,z)))" +by fast + +lemma cartprod_separation: + "[| L(A); L(B) |] + ==> separation(L, \z. \x[L]. x\A & (\y[L]. y\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, \y. \p[L]. p\r & (\x[L]. x\A & pair(L,x,y,p)), + \i y. \p\Lset(i). p\r & (\x\Lset(i). x\A & pair(**Lset(i),x,y,p)))" +by fast + + +lemma image_separation: + "[| L(A); L(r) |] + ==> separation(L, \y. \p[L]. p\r & (\x[L]. x\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, + \z. \p[L]. p\r & (\x[L]. \y[L]. pair(L,x,y,p) & pair(L,y,x,z)), + \i z. \p\Lset(i). p\r & (\x\Lset(i). \y\Lset(i). + pair(**Lset(i),x,y,p) & pair(**Lset(i),y,x,z)))" +by fast + +lemma converse_separation: + "L(r) ==> separation(L, + \z. \p[L]. p\r & (\x[L]. \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, \z. \x[L]. x\A & (\y[L]. pair(L,x,y,z)), + \i z. \x\Lset(i). x\A & (\y\Lset(i). pair(**Lset(i),x,y,z)))" +by fast + +lemma restrict_separation: + "L(A) ==> separation(L, \z. \x[L]. x\A & (\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, \xz. \x[L]. \y[L]. \z[L]. \xy[L]. \yz[L]. + pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & + xy\s & yz\r, + \i xz. \x\Lset(i). \y\Lset(i). \z\Lset(i). \xy\Lset(i). \yz\Lset(i). + pair(**Lset(i),x,z,xz) & pair(**Lset(i),x,y,xy) & + pair(**Lset(i),y,z,yz) & xy\s & yz\r)" +by fast + +lemma comp_separation: + "[| L(r); L(s) |] + ==> separation(L, \xz. \x[L]. \y[L]. \z[L]. \xy[L]. \yz[L]. + pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & + xy\s & yz\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, \z. \x\A. \y\B. L(x) & L(y) & pair(L,x,y,z))" - and image_separation: - "[| L(A); L(r) |] - ==> separation(L, \y. \p\r. L(p) & (\x\A. L(x) & pair(L,x,y,p)))" - and vimage_separation: - "[| L(A); L(r) |] - ==> separation(L, \x. \p\r. L(p) & (\y\A. L(x) & pair(L,x,y,p)))" - and converse_separation: - "L(r) ==> separation(L, \z. \p\r. L(p) & (\x y. L(x) & L(y) & - pair(L,x,y,p) & pair(L,y,x,z)))" - and restrict_separation: - "L(A) - ==> separation(L, \z. \x\A. L(x) & (\y. L(y) & pair(L,x,y,z)))" - and comp_separation: - "[| L(r); L(s) |] - ==> separation(L, \xz. \x y z. L(x) & L(y) & L(z) & - (\xy\s. \yz\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, \y. \p\r. L(p) & pair(L,y,x,p))" and Memrel_separation: diff -r e4ae0732e2be -r b4f370679c65 src/ZF/Constructible/Relative.thy --- 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) == \x[M]. x \ z" subset :: "[i=>o,i,i] => o" - "subset(M,A,B) == \x\A. M(x) --> x \ B" + "subset(M,A,B) == \x[M]. x\A --> x \ B" upair :: "[i=>o,i,i,i] => o" - "upair(M,a,b,z) == a \ z & b \ z & (\x\z. M(x) --> x = a | x = b)" + "upair(M,a,b,z) == a \ z & b \ z & (\x[M]. x\z --> x = a | x = b)" pair :: "[i=>o,i,i,i] => o" "pair(M,a,b,z) == \x[M]. upair(M,a,a,x) & @@ -34,35 +34,34 @@ "setdiff(M,a,b,z) == \x[M]. x \ z <-> x \ a & x \ b" big_union :: "[i=>o,i,i] => o" - "big_union(M,A,z) == \x[M]. x \ z <-> (\y\A. M(y) & x \ y)" + "big_union(M,A,z) == \x[M]. x \ z <-> (\y[M]. y\A & x \ y)" big_inter :: "[i=>o,i,i] => o" "big_inter(M,A,z) == (A=0 --> z=0) & - (A\0 --> (\x[M]. x \ z <-> (\y\A. M(y) --> x \ y)))" + (A\0 --> (\x[M]. x \ z <-> (\y[M]. y\A --> x \ y)))" cartprod :: "[i=>o,i,i,i] => o" "cartprod(M,A,B,z) == - \u[M]. u \ z <-> (\x\A. M(x) & (\y\B. M(y) & pair(M,x,y,u)))" + \u[M]. u \ z <-> (\x[M]. x\A & (\y[M]. y\B & pair(M,x,y,u)))" is_converse :: "[i=>o,i,i] => o" "is_converse(M,r,z) == \x. M(x) --> (x \ z <-> - (\w\r. M(w) & - (\u[M]. \v[M]. pair(M,u,v,w) & pair(M,v,u,x))))" + (\w[M]. w\r & (\u[M]. \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) == - \x. M(x) --> (x \ z <-> (\w\r. M(w) & (\y\A. M(y) & pair(M,x,y,w))))" + \x. M(x) --> (x \ z <-> (\w[M]. w\r & (\y[M]. y\A & pair(M,x,y,w))))" is_domain :: "[i=>o,i,i] => o" "is_domain(M,r,z) == - \x. M(x) --> (x \ z <-> (\w\r. M(w) & (\y. M(y) & pair(M,x,y,w))))" + \x. M(x) --> (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) --> (y \ z <-> (\w\r. M(w) & (\x\A. M(x) & pair(M,x,y,w))))" + \y. M(y) --> (y \ z <-> (\w[M]. w\r & (\x[M]. x\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) == - \y. M(y) --> (y \ z <-> (\w\r. M(w) & (\x. M(x) & pair(M,x,y,w))))" + \y. M(y) --> (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(dr) & is_domain(M,r,dr) & - (\rr. M(rr) & is_range(M,r,rr) & union(M,dr,rr,z))" + \dr[M]. is_domain(M,r,dr) & + (\rr[M]. is_range(M,r,rr) & union(M,dr,rr,z))" is_relation :: "[i=>o,i] => o" "is_relation(M,r) == - (\z\r. M(z) --> (\x y. M(x) & M(y) & pair(M,x,y,z)))" + (\z[M]. z\r --> (\x[M]. \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, \x. \y[M]. y\A --> x\y)" and cartprod_separation: "[| M(A); M(B) |] - ==> separation(M, \z. \x\A. \y\B. M(x) & M(y) & pair(M,x,y,z))" + ==> separation(M, \z. \x[M]. x\A & (\y[M]. y\B & pair(M,x,y,z)))" and image_separation: "[| M(A); M(r) |] ==> separation(M, \y. \p[M]. p\r & (\x[M]. x\A & pair(M,x,y,p)))" and converse_separation: - "M(r) ==> separation(M, \z. \p\r. - M(p) & (\x[M]. \y[M]. pair(M,x,y,p) & pair(M,y,x,z)))" + "M(r) ==> separation(M, + \z. \p[M]. p\r & (\x[M]. \y[M]. pair(M,x,y,p) & pair(M,y,x,z)))" and restrict_separation: "M(A) ==> separation(M, \z. \x[M]. x\A & (\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\s & yz\r)" and pred_separation: - "[| M(r); M(x) |] ==> separation(M, \y. \p\r. M(p) & pair(M,y,x,p))" + "[| M(r); M(x) |] ==> separation(M, \y. \p[M]. p\r & pair(M,y,x,p))" and Memrel_separation: - "separation(M, \z. \x y. M(x) & M(y) & pair(M,x,y,z) & x \ y)" + "separation(M, \z. \x[M]. \y[M]. pair(M,x,y,z) & x \ y)" and obase_separation: --{*part of the order type formalization*} "[| M(A); M(r) |] diff -r e4ae0732e2be -r b4f370679c65 src/ZF/Constructible/Wellorderings.thy --- 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 \ A*A. \x. M(x) \ (\y. M(y) \ z = \x,y\ \ x \ y)}" + Memrel(A) = {z \ A*A. \x[M]. \y[M]. z = \x,y\ & x \ y}" apply (simp add: Memrel_def) apply (blast dest: transM) done diff -r e4ae0732e2be -r b4f370679c65 src/ZF/OrdQuant.thy --- 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 P(x)" - + oex :: "[i, i => o] => o" "oex(A, P) == EX x. x 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 (Ord(i) --> P(i) & (ALL x (Ord(i) & (P(i) | (EX x Ord(B(x)) |] ==> Ord(\xx i < (\xb(a); Ord(\x i \ (\x (UN x 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 P(x) |] ==> ALL x P(x)" -by (simp add: oall_def) +by (simp add: oall_def) lemma oallE: "[| ALL x Q; ~x Q |] ==> Q" -apply (simp add: oall_def, blast) -done +by (simp add: oall_def, blast) lemma rev_oallE [elim]: "[| ALL x Q; P(x) ==> Q |] ==> Q" -apply (simp add: oall_def, blast) -done +by (simp add: oall_def, blast) (*Trival rewrite rule; (ALL xP holds only if a is not 0!*) @@ -135,7 +133,7 @@ (*Congruence rule for rewriting*) lemma oall_cong [cong]: - "[| a=a'; !!x. x P(x) <-> P'(x) |] + "[| a=a'; !!x. x 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 EX x P(a); a EX x Q |] ==> Q" -apply (simp add: oex_def, blast) +apply (simp add: oex_def, blast) done lemma oex_cong [cong]: - "[| a=a'; !!x. x P(x) <-> P'(x) |] + "[| a=a'; !!x. x 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 C(x)=D(x) |] ==> (UN x 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]: "(\x[**A]. P(x)) <-> (\x\A. P(x))" +by auto + +lemma rex_setclass_is_bex [simp]: "(\x[**A]. P(x)) <-> (\x\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;