--- 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 \<in> nat; y \<in> nat; z \<in> nat |] ==> upair_fm(x,y,z) \<in> formula"
+by (simp add: upair_fm_def)
+
+lemma arity_upair_fm [simp]:
+ "[| x \<in> nat; y \<in> nat; z \<in> nat |]
+ ==> arity(upair_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
+by (simp add: upair_fm_def succ_Un_distrib [symmetric] Un_ac)
+
+lemma sats_upair_fm [simp]:
+ "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> 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 \<in> nat; j \<in> nat; k \<in> nat; env \<in> 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 \<in> nat; y \<in> nat; z < length(env); env \<in> 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 \<in> nat; y \<in> nat; z \<in> nat |] ==> pair_fm(x,y,z) \<in> formula"
+by (simp add: pair_fm_def)
+
+lemma arity_pair_fm [simp]:
+ "[| x \<in> nat; y \<in> nat; z \<in> nat |]
+ ==> arity(pair_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
+by (simp add: pair_fm_def succ_Un_distrib [symmetric] Un_ac)
+
+lemma sats_pair_fm [simp]:
+ "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> 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 \<in> nat; j \<in> nat; k \<in> nat; env \<in> 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 \<in> nat|] ==> nth(succ(n), Cons(a,l)) = x"
+by simp
+
+
+lemma Collect_conj_in_DPow:
+ "[| {x\<in>A. P(x)} \<in> DPow(A); {x\<in>A. Q(x)} \<in> DPow(A) |]
+ ==> {x\<in>A. P(x) & Q(x)} \<in> DPow(A)"
+by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric])
+
+lemma Collect_conj_in_DPow_Lset:
+ "[|z \<in> Lset(j); {x \<in> Lset(j). P(x)} \<in> DPow(Lset(j))|]
+ ==> {x \<in> Lset(j). x \<in> z & P(x)} \<in> 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:
+ "(\<And>z. L(z) ==> L({x \<in> z . P(x)})) ==> separation(L, \<lambda>x. P(x))"
+apply (unfold separation_def, clarify)
+apply (rule_tac x="{x\<in>z. P(x)}" in rexI)
+apply simp_all
+done
+
+text{*Reduces the original comprehension to the reflected one*}
+lemma reflection_imp_L_separation:
+ "[| \<forall>x\<in>Lset(j). P(x) <-> Q(x);
+ {x \<in> Lset(j) . Q(x)} \<in> DPow(Lset(j));
+ Ord(j); z \<in> Lset(j)|] ==> L({x \<in> z . P(x)})"
+apply (rule_tac i = "succ(j)" in L_I)
+ prefer 2 apply simp
+apply (subgoal_tac "{x \<in> z. P(x)} = {x \<in> Lset(j). x \<in> 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, \<lambda>x. \<forall>y[L]. y\<in>A --> x \<in> y,
+ \<lambda>i x. \<forall>y\<in>Lset(i). y\<in>A --> x \<in> y)"
+by fast
+
+lemma Inter_separation:
+ "L(A) ==> separation(L, \<lambda>x. \<forall>y[L]. y\<in>A --> x\<in>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, \<lambda>x. upair(L,f(x),g(x),h(x)),
+ \<lambda>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, \<lambda>x. pair(L,f(x),g(x),h(x)),
+ \<lambda>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, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)),
+ \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). y\<in>B &
+ pair(**Lset(i),x,y,z)))"
+by fast
+
+lemma cartprod_separation:
+ "[| L(A); L(B) |]
+ ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>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, \<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)),
+ \<lambda>i y. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). x\<in>A & pair(**Lset(i),x,y,p)))"
+by fast
+
+
+lemma image_separation:
+ "[| L(A); L(r) |]
+ ==> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>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,
+ \<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)),
+ \<lambda>i z. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i).
+ pair(**Lset(i),x,y,p) & pair(**Lset(i),y,x,z)))"
+by fast
+
+lemma converse_separation:
+ "L(r) ==> separation(L,
+ \<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>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, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)),
+ \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). pair(**Lset(i),x,y,z)))"
+by fast
+
+lemma restrict_separation:
+ "L(A) ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>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, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
+ pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
+ xy\<in>s & yz\<in>r,
+ \<lambda>i xz. \<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i). \<exists>z\<in>Lset(i). \<exists>xy\<in>Lset(i). \<exists>yz\<in>Lset(i).
+ pair(**Lset(i),x,z,xz) & pair(**Lset(i),x,y,xy) &
+ pair(**Lset(i),y,z,yz) & xy\<in>s & yz\<in>r)"
+by fast
+
+lemma comp_separation:
+ "[| L(r); L(s) |]
+ ==> separation(L, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
+ pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
+ xy\<in>s & yz\<in>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, \<lambda>z. \<exists>x\<in>A. \<exists>y\<in>B. L(x) & L(y) & pair(L,x,y,z))"
- and image_separation:
- "[| L(A); L(r) |]
- ==> separation(L, \<lambda>y. \<exists>p\<in>r. L(p) & (\<exists>x\<in>A. L(x) & pair(L,x,y,p)))"
- and vimage_separation:
- "[| L(A); L(r) |]
- ==> separation(L, \<lambda>x. \<exists>p\<in>r. L(p) & (\<exists>y\<in>A. L(x) & pair(L,x,y,p)))"
- and converse_separation:
- "L(r) ==> separation(L, \<lambda>z. \<exists>p\<in>r. L(p) & (\<exists>x y. L(x) & L(y) &
- pair(L,x,y,p) & pair(L,y,x,z)))"
- and restrict_separation:
- "L(A)
- ==> separation(L, \<lambda>z. \<exists>x\<in>A. L(x) & (\<exists>y. L(y) & pair(L,x,y,z)))"
- and comp_separation:
- "[| L(r); L(s) |]
- ==> separation(L, \<lambda>xz. \<exists>x y z. L(x) & L(y) & L(z) &
- (\<exists>xy\<in>s. \<exists>yz\<in>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, \<lambda>y. \<exists>p\<in>r. L(p) & pair(L,y,x,p))"
and Memrel_separation: