Up to index of Isabelle/ZF/Constructible
theory Separation(* Title: ZF/Constructible/Separation.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory *) header{*Early Instances of Separation and Strong Replacement*} theory Separation imports L_axioms WF_absolute begin text{*This theory proves all instances needed for locale @{text "M_basic"}*} 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 lemmas nth_rules = nth_0 nth_ConsI nat_0I nat_succI lemmas sep_rules = nth_0 nth_ConsI FOL_iff_sats function_iff_sats fun_plus_iff_sats 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 text{*Encapsulates the standard proof script for proving instances of Separation.*} lemma gen_separation: assumes reflection: "REFLECTS [P,Q]" and Lu: "L(u)" and collI: "!!j. u ∈ Lset(j) ==> Collect(Lset(j), Q(j)) ∈ DPow(Lset(j))" shows "separation(L,P)" apply (rule separation_CollectI) apply (rule_tac A="{u,z}" in subset_LsetE, blast intro: Lu) apply (rule ReflectsE [OF reflection], assumption) apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2, clarify) apply (rule collI, assumption) done text{*As above, but typically @{term u} is a finite enumeration such as @{term "{a,b}"}; thus the new subgoal gets the assumption @{term "{a,b} ⊆ Lset(i)"}, which is logically equivalent to @{term "a ∈ Lset(i)"} and @{term "b ∈ Lset(i)"}.*} lemma gen_separation_multi: assumes reflection: "REFLECTS [P,Q]" and Lu: "L(u)" and collI: "!!j. u ⊆ Lset(j) ==> Collect(Lset(j), Q(j)) ∈ DPow(Lset(j))" shows "separation(L,P)" apply (rule gen_separation [OF reflection Lu]) apply (drule mem_Lset_imp_subset_Lset) apply (erule collI) done subsection{*Separation for Intersection*} lemma Inter_Reflects: "REFLECTS[λx. ∀y[L]. y∈A --> x ∈ y, λi x. ∀y∈Lset(i). y∈A --> x ∈ y]" by (intro FOL_reflections) lemma Inter_separation: "L(A) ==> separation(L, λx. ∀y[L]. y∈A --> x∈y)" apply (rule gen_separation [OF Inter_Reflects], simp) apply (rule DPow_LsetI) txt{*I leave this one example of a manual proof. The tedium of manually instantiating @{term i}, @{term j} and @{term env} is obvious. *} 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 subsection{*Separation for Set Difference*} lemma Diff_Reflects: "REFLECTS[λx. x ∉ B, λi x. x ∉ B]" by (intro FOL_reflections) lemma Diff_separation: "L(B) ==> separation(L, λx. x ∉ B)" apply (rule gen_separation [OF Diff_Reflects], simp) apply (rule_tac env="[B]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for Cartesian Product*} lemma cartprod_Reflects: "REFLECTS[λ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 (intro FOL_reflections function_reflections) 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 gen_separation_multi [OF cartprod_Reflects, of "{A,B}"], auto) apply (rule_tac env="[A,B]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for Image*} lemma image_Reflects: "REFLECTS[λ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 (intro FOL_reflections function_reflections) 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 gen_separation_multi [OF image_Reflects, of "{A,r}"], auto) apply (rule_tac env="[A,r]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for Converse*} lemma converse_Reflects: "REFLECTS[λ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 (intro FOL_reflections function_reflections) 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 gen_separation [OF converse_Reflects], simp) apply (rule_tac env="[r]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for Restriction*} lemma restrict_Reflects: "REFLECTS[λ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 (intro FOL_reflections function_reflections) lemma restrict_separation: "L(A) ==> separation(L, λz. ∃x[L]. x∈A & (∃y[L]. pair(L,x,y,z)))" apply (rule gen_separation [OF restrict_Reflects], simp) apply (rule_tac env="[A]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for Composition*} lemma comp_Reflects: "REFLECTS[λ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 (intro FOL_reflections function_reflections) 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 gen_separation_multi [OF comp_Reflects, of "{r,s}"], auto) txt{*Subgoals after applying general ``separation'' rule: @{subgoals[display,indent=0,margin=65]}*} apply (rule_tac env="[r,s]" in DPow_LsetI) txt{*Subgoals ready for automatic synthesis of a formula: @{subgoals[display,indent=0,margin=65]}*} apply (rule sep_rules | simp)+ done subsection{*Separation for Predecessors in an Order*} lemma pred_Reflects: "REFLECTS[λy. ∃p[L]. p∈r & pair(L,y,x,p), λi y. ∃p ∈ Lset(i). p∈r & pair(##Lset(i),y,x,p)]" by (intro FOL_reflections function_reflections) lemma pred_separation: "[| L(r); L(x) |] ==> separation(L, λy. ∃p[L]. p∈r & pair(L,y,x,p))" apply (rule gen_separation_multi [OF pred_Reflects, of "{r,x}"], auto) apply (rule_tac env="[r,x]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for the Membership Relation*} lemma Memrel_Reflects: "REFLECTS[λz. ∃x[L]. ∃y[L]. pair(L,x,y,z) & x ∈ y, λi z. ∃x ∈ Lset(i). ∃y ∈ Lset(i). pair(##Lset(i),x,y,z) & x ∈ y]" by (intro FOL_reflections function_reflections) lemma Memrel_separation: "separation(L, λz. ∃x[L]. ∃y[L]. pair(L,x,y,z) & x ∈ y)" apply (rule gen_separation [OF Memrel_Reflects nonempty]) apply (rule_tac env="[]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Replacement for FunSpace*} lemma funspace_succ_Reflects: "REFLECTS[λz. ∃p[L]. p∈A & (∃f[L]. ∃b[L]. ∃nb[L]. ∃cnbf[L]. pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) & upair(L,cnbf,cnbf,z)), λi z. ∃p ∈ Lset(i). p∈A & (∃f ∈ Lset(i). ∃b ∈ Lset(i). ∃nb ∈ Lset(i). ∃cnbf ∈ Lset(i). pair(##Lset(i),f,b,p) & pair(##Lset(i),n,b,nb) & is_cons(##Lset(i),nb,f,cnbf) & upair(##Lset(i),cnbf,cnbf,z))]" by (intro FOL_reflections function_reflections) lemma funspace_succ_replacement: "L(n) ==> strong_replacement(L, λp z. ∃f[L]. ∃b[L]. ∃nb[L]. ∃cnbf[L]. pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) & upair(L,cnbf,cnbf,z))" apply (rule strong_replacementI) apply (rule_tac u="{n,B}" in gen_separation_multi [OF funspace_succ_Reflects], auto) apply (rule_tac env="[n,B]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for a Theorem about @{term "is_recfun"}*} lemma is_recfun_reflects: "REFLECTS[λx. ∃xa[L]. ∃xb[L]. pair(L,x,a,xa) & xa ∈ r & pair(L,x,b,xb) & xb ∈ r & (∃fx[L]. ∃gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & fx ≠ gx), λi x. ∃xa ∈ Lset(i). ∃xb ∈ Lset(i). pair(##Lset(i),x,a,xa) & xa ∈ r & pair(##Lset(i),x,b,xb) & xb ∈ r & (∃fx ∈ Lset(i). ∃gx ∈ Lset(i). fun_apply(##Lset(i),f,x,fx) & fun_apply(##Lset(i),g,x,gx) & fx ≠ gx)]" by (intro FOL_reflections function_reflections fun_plus_reflections) lemma is_recfun_separation: --{*for well-founded recursion*} "[| L(r); L(f); L(g); L(a); L(b) |] ==> separation(L, λx. ∃xa[L]. ∃xb[L]. pair(L,x,a,xa) & xa ∈ r & pair(L,x,b,xb) & xb ∈ r & (∃fx[L]. ∃gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & fx ≠ gx))" apply (rule gen_separation_multi [OF is_recfun_reflects, of "{r,f,g,a,b}"], auto) apply (rule_tac env="[r,f,g,a,b]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Instantiating the locale @{text M_basic}*} text{*Separation (and Strong Replacement) for basic set-theoretic constructions such as intersection, Cartesian Product and image.*} lemma M_basic_axioms_L: "M_basic_axioms(L)" apply (rule M_basic_axioms.intro) apply (assumption | rule Inter_separation Diff_separation cartprod_separation image_separation converse_separation restrict_separation comp_separation pred_separation Memrel_separation funspace_succ_replacement is_recfun_separation)+ done theorem M_basic_L: "PROP M_basic(L)" by (rule M_basic.intro [OF M_trivial_L M_basic_axioms_L]) interpretation L?: M_basic L by (rule M_basic_L) end