1(*  Title:      ZF/Constructible/Separation.thy
2    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
3*)
4
5section\<open>Early Instances of Separation and Strong Replacement\<close>
6
7theory Separation imports L_axioms WF_absolute begin
8
9text\<open>This theory proves all instances needed for locale \<open>M_basic\<close>\<close>
10
11text\<open>Helps us solve for de Bruijn indices!\<close>
12lemma nth_ConsI: "[|nth(n,l) = x; n \<in> nat|] ==> nth(succ(n), Cons(a,l)) = x"
13by simp
14
15lemmas nth_rules = nth_0 nth_ConsI nat_0I nat_succI
16lemmas sep_rules = nth_0 nth_ConsI FOL_iff_sats function_iff_sats
17                   fun_plus_iff_sats
18
19lemma Collect_conj_in_DPow:
20     "[| {x\<in>A. P(x)} \<in> DPow(A);  {x\<in>A. Q(x)} \<in> DPow(A) |]
21      ==> {x\<in>A. P(x) & Q(x)} \<in> DPow(A)"
22by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric])
23
24lemma Collect_conj_in_DPow_Lset:
25     "[|z \<in> Lset(j); {x \<in> Lset(j). P(x)} \<in> DPow(Lset(j))|]
26      ==> {x \<in> Lset(j). x \<in> z & P(x)} \<in> DPow(Lset(j))"
27apply (frule mem_Lset_imp_subset_Lset)
28apply (simp add: Collect_conj_in_DPow Collect_mem_eq
29                 subset_Int_iff2 elem_subset_in_DPow)
30done
31
32lemma separation_CollectI:
33     "(\<And>z. L(z) ==> L({x \<in> z . P(x)})) ==> separation(L, \<lambda>x. P(x))"
34apply (unfold separation_def, clarify)
35apply (rule_tac x="{x\<in>z. P(x)}" in rexI)
36apply simp_all
37done
38
39text\<open>Reduces the original comprehension to the reflected one\<close>
40lemma reflection_imp_L_separation:
41      "[| \<forall>x\<in>Lset(j). P(x) \<longleftrightarrow> Q(x);
42          {x \<in> Lset(j) . Q(x)} \<in> DPow(Lset(j));
43          Ord(j);  z \<in> Lset(j)|] ==> L({x \<in> z . P(x)})"
44apply (rule_tac i = "succ(j)" in L_I)
45 prefer 2 apply simp
46apply (subgoal_tac "{x \<in> z. P(x)} = {x \<in> Lset(j). x \<in> z & (Q(x))}")
47 prefer 2
48 apply (blast dest: mem_Lset_imp_subset_Lset)
49apply (simp add: Lset_succ Collect_conj_in_DPow_Lset)
50done
51
52text\<open>Encapsulates the standard proof script for proving instances of 
53      Separation.\<close>
54lemma gen_separation:
55 assumes reflection: "REFLECTS [P,Q]"
56     and Lu:         "L(u)"
57     and collI: "!!j. u \<in> Lset(j)
58                \<Longrightarrow> Collect(Lset(j), Q(j)) \<in> DPow(Lset(j))"
59 shows "separation(L,P)"
60apply (rule separation_CollectI)
61apply (rule_tac A="{u,z}" in subset_LsetE, blast intro: Lu)
62apply (rule ReflectsE [OF reflection], assumption)
63apply (drule subset_Lset_ltD, assumption)
64apply (erule reflection_imp_L_separation)
65  apply (simp_all add: lt_Ord2, clarify)
66apply (rule collI, assumption)
67done
68
69text\<open>As above, but typically \<^term>\<open>u\<close> is a finite enumeration such as
70  \<^term>\<open>{a,b}\<close>; thus the new subgoal gets the assumption
71  \<^term>\<open>{a,b} \<subseteq> Lset(i)\<close>, which is logically equivalent to 
72  \<^term>\<open>a \<in> Lset(i)\<close> and \<^term>\<open>b \<in> Lset(i)\<close>.\<close>
73lemma gen_separation_multi:
74 assumes reflection: "REFLECTS [P,Q]"
75     and Lu:         "L(u)"
76     and collI: "!!j. u \<subseteq> Lset(j)
77                \<Longrightarrow> Collect(Lset(j), Q(j)) \<in> DPow(Lset(j))"
78 shows "separation(L,P)"
79apply (rule gen_separation [OF reflection Lu])
80apply (drule mem_Lset_imp_subset_Lset)
81apply (erule collI) 
82done
83
84
85subsection\<open>Separation for Intersection\<close>
86
87lemma Inter_Reflects:
88     "REFLECTS[\<lambda>x. \<forall>y[L]. y\<in>A \<longrightarrow> x \<in> y,
89               \<lambda>i x. \<forall>y\<in>Lset(i). y\<in>A \<longrightarrow> x \<in> y]"
90by (intro FOL_reflections)
91
92lemma Inter_separation:
93     "L(A) ==> separation(L, \<lambda>x. \<forall>y[L]. y\<in>A \<longrightarrow> x\<in>y)"
94apply (rule gen_separation [OF Inter_Reflects], simp)
95apply (rule DPow_LsetI)
96 txt\<open>I leave this one example of a manual proof.  The tedium of manually
97      instantiating \<^term>\<open>i\<close>, \<^term>\<open>j\<close> and \<^term>\<open>env\<close> is obvious.\<close>
98apply (rule ball_iff_sats)
99apply (rule imp_iff_sats)
100apply (rule_tac [2] i=1 and j=0 and env="[y,x,A]" in mem_iff_sats)
101apply (rule_tac i=0 and j=2 in mem_iff_sats)
102apply (simp_all add: succ_Un_distrib [symmetric])
103done
104
105subsection\<open>Separation for Set Difference\<close>
106
107lemma Diff_Reflects:
108     "REFLECTS[\<lambda>x. x \<notin> B, \<lambda>i x. x \<notin> B]"
109by (intro FOL_reflections)  
110
111lemma Diff_separation:
112     "L(B) ==> separation(L, \<lambda>x. x \<notin> B)"
113apply (rule gen_separation [OF Diff_Reflects], simp)
114apply (rule_tac env="[B]" in DPow_LsetI)
115apply (rule sep_rules | simp)+
116done
117
118subsection\<open>Separation for Cartesian Product\<close>
119
120lemma cartprod_Reflects:
121     "REFLECTS[\<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)),
122                \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). y\<in>B &
123                                   pair(##Lset(i),x,y,z))]"
124by (intro FOL_reflections function_reflections)
125
126lemma cartprod_separation:
127     "[| L(A); L(B) |]
128      ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)))"
129apply (rule gen_separation_multi [OF cartprod_Reflects, of "{A,B}"], auto)
130apply (rule_tac env="[A,B]" in DPow_LsetI)
131apply (rule sep_rules | simp)+
132done
133
134subsection\<open>Separation for Image\<close>
135
136lemma image_Reflects:
137     "REFLECTS[\<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)),
138           \<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))]"
139by (intro FOL_reflections function_reflections)
140
141lemma image_separation:
142     "[| L(A); L(r) |]
143      ==> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)))"
144apply (rule gen_separation_multi [OF image_Reflects, of "{A,r}"], auto)
145apply (rule_tac env="[A,r]" in DPow_LsetI)
146apply (rule sep_rules | simp)+
147done
148
149
150subsection\<open>Separation for Converse\<close>
151
152lemma converse_Reflects:
153  "REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)),
154     \<lambda>i z. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i).
155                     pair(##Lset(i),x,y,p) & pair(##Lset(i),y,x,z))]"
156by (intro FOL_reflections function_reflections)
157
158lemma converse_separation:
159     "L(r) ==> separation(L,
160         \<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)))"
161apply (rule gen_separation [OF converse_Reflects], simp)
162apply (rule_tac env="[r]" in DPow_LsetI)
163apply (rule sep_rules | simp)+
164done
165
166
167subsection\<open>Separation for Restriction\<close>
168
169lemma restrict_Reflects:
170     "REFLECTS[\<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)),
171        \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). pair(##Lset(i),x,y,z))]"
172by (intro FOL_reflections function_reflections)
173
174lemma restrict_separation:
175   "L(A) ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)))"
176apply (rule gen_separation [OF restrict_Reflects], simp)
177apply (rule_tac env="[A]" in DPow_LsetI)
178apply (rule sep_rules | simp)+
179done
180
181
182subsection\<open>Separation for Composition\<close>
183
184lemma comp_Reflects:
185     "REFLECTS[\<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
186                  pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
187                  xy\<in>s & yz\<in>r,
188        \<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).
189                  pair(##Lset(i),x,z,xz) & pair(##Lset(i),x,y,xy) &
190                  pair(##Lset(i),y,z,yz) & xy\<in>s & yz\<in>r]"
191by (intro FOL_reflections function_reflections)
192
193lemma comp_separation:
194     "[| L(r); L(s) |]
195      ==> separation(L, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
196                  pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
197                  xy\<in>s & yz\<in>r)"
198apply (rule gen_separation_multi [OF comp_Reflects, of "{r,s}"], auto)
199txt\<open>Subgoals after applying general ``separation'' rule:
200     @{subgoals[display,indent=0,margin=65]}\<close>
201apply (rule_tac env="[r,s]" in DPow_LsetI)
202txt\<open>Subgoals ready for automatic synthesis of a formula:
203     @{subgoals[display,indent=0,margin=65]}\<close>
204apply (rule sep_rules | simp)+
205done
206
207
208subsection\<open>Separation for Predecessors in an Order\<close>
209
210lemma pred_Reflects:
211     "REFLECTS[\<lambda>y. \<exists>p[L]. p\<in>r & pair(L,y,x,p),
212                    \<lambda>i y. \<exists>p \<in> Lset(i). p\<in>r & pair(##Lset(i),y,x,p)]"
213by (intro FOL_reflections function_reflections)
214
215lemma pred_separation:
216     "[| L(r); L(x) |] ==> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & pair(L,y,x,p))"
217apply (rule gen_separation_multi [OF pred_Reflects, of "{r,x}"], auto)
218apply (rule_tac env="[r,x]" in DPow_LsetI)
219apply (rule sep_rules | simp)+
220done
221
222
223subsection\<open>Separation for the Membership Relation\<close>
224
225lemma Memrel_Reflects:
226     "REFLECTS[\<lambda>z. \<exists>x[L]. \<exists>y[L]. pair(L,x,y,z) & x \<in> y,
227            \<lambda>i z. \<exists>x \<in> Lset(i). \<exists>y \<in> Lset(i). pair(##Lset(i),x,y,z) & x \<in> y]"
228by (intro FOL_reflections function_reflections)
229
230lemma Memrel_separation:
231     "separation(L, \<lambda>z. \<exists>x[L]. \<exists>y[L]. pair(L,x,y,z) & x \<in> y)"
232apply (rule gen_separation [OF Memrel_Reflects nonempty])
233apply (rule_tac env="[]" in DPow_LsetI)
234apply (rule sep_rules | simp)+
235done
236
237
238subsection\<open>Replacement for FunSpace\<close>
239
240lemma funspace_succ_Reflects:
241 "REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>A & (\<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L].
242            pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) &
243            upair(L,cnbf,cnbf,z)),
244        \<lambda>i z. \<exists>p \<in> Lset(i). p\<in>A & (\<exists>f \<in> Lset(i). \<exists>b \<in> Lset(i).
245              \<exists>nb \<in> Lset(i). \<exists>cnbf \<in> Lset(i).
246                pair(##Lset(i),f,b,p) & pair(##Lset(i),n,b,nb) &
247                is_cons(##Lset(i),nb,f,cnbf) & upair(##Lset(i),cnbf,cnbf,z))]"
248by (intro FOL_reflections function_reflections)
249
250lemma funspace_succ_replacement:
251     "L(n) ==>
252      strong_replacement(L, \<lambda>p z. \<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L].
253                pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) &
254                upair(L,cnbf,cnbf,z))"
255apply (rule strong_replacementI)
256apply (rule_tac u="{n,B}" in gen_separation_multi [OF funspace_succ_Reflects], 
257       auto)
258apply (rule_tac env="[n,B]" in DPow_LsetI)
259apply (rule sep_rules | simp)+
260done
261
262
263subsection\<open>Separation for a Theorem about \<^term>\<open>is_recfun\<close>\<close>
264
265lemma is_recfun_reflects:
266  "REFLECTS[\<lambda>x. \<exists>xa[L]. \<exists>xb[L].
267                pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r &
268                (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) &
269                                   fx \<noteq> gx),
270   \<lambda>i x. \<exists>xa \<in> Lset(i). \<exists>xb \<in> Lset(i).
271          pair(##Lset(i),x,a,xa) & xa \<in> r & pair(##Lset(i),x,b,xb) & xb \<in> r &
272                (\<exists>fx \<in> Lset(i). \<exists>gx \<in> Lset(i). fun_apply(##Lset(i),f,x,fx) &
273                  fun_apply(##Lset(i),g,x,gx) & fx \<noteq> gx)]"
274by (intro FOL_reflections function_reflections fun_plus_reflections)
275
276lemma is_recfun_separation:
277     \<comment> \<open>for well-founded recursion\<close>
278     "[| L(r); L(f); L(g); L(a); L(b) |]
279     ==> separation(L,
280            \<lambda>x. \<exists>xa[L]. \<exists>xb[L].
281                pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r &
282                (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) &
283                                   fx \<noteq> gx))"
284apply (rule gen_separation_multi [OF is_recfun_reflects, of "{r,f,g,a,b}"], 
285            auto)
286apply (rule_tac env="[r,f,g,a,b]" in DPow_LsetI)
287apply (rule sep_rules | simp)+
288done
289
290
291subsection\<open>Instantiating the locale \<open>M_basic\<close>\<close>
292text\<open>Separation (and Strong Replacement) for basic set-theoretic constructions
293such as intersection, Cartesian Product and image.\<close>
294
295lemma M_basic_axioms_L: "M_basic_axioms(L)"
296  apply (rule M_basic_axioms.intro)
297       apply (assumption | rule
298         Inter_separation Diff_separation cartprod_separation image_separation
299         converse_separation restrict_separation
300         comp_separation pred_separation Memrel_separation
301         funspace_succ_replacement is_recfun_separation power_ax)+
302  done
303
304theorem M_basic_L: " M_basic(L)"
305by (rule M_basic.intro [OF M_trivial_L M_basic_axioms_L])
306
307interpretation L: M_basic L by (rule M_basic_L)
308
309
310end
311