1(*  Title:      ZF/Constructible/WF_absolute.thy
2    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
3*)
4
5section \<open>Absoluteness of Well-Founded Recursion\<close>
6
7theory WF_absolute imports WFrec begin
8
9subsection\<open>Transitive closure without fixedpoints\<close>
10
11definition
12  rtrancl_alt :: "[i,i]=>i" where
13    "rtrancl_alt(A,r) ==
14       {p \<in> A*A. \<exists>n\<in>nat. \<exists>f \<in> succ(n) -> A.
15                 (\<exists>x y. p = <x,y> &  f`0 = x & f`n = y) &
16                       (\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)}"
17
18lemma alt_rtrancl_lemma1 [rule_format]:
19    "n \<in> nat
20     ==> \<forall>f \<in> succ(n) -> field(r).
21         (\<forall>i\<in>n. \<langle>f`i, f ` succ(i)\<rangle> \<in> r) \<longrightarrow> \<langle>f`0, f`n\<rangle> \<in> r^*"
22apply (induct_tac n)
23apply (simp_all add: apply_funtype rtrancl_refl, clarify)
24apply (rename_tac n f)
25apply (rule rtrancl_into_rtrancl)
26 prefer 2 apply assumption
27apply (drule_tac x="restrict(f,succ(n))" in bspec)
28 apply (blast intro: restrict_type2)
29apply (simp add: Ord_succ_mem_iff nat_0_le [THEN ltD] leI [THEN ltD] ltI)
30done
31
32lemma rtrancl_alt_subset_rtrancl: "rtrancl_alt(field(r),r) \<subseteq> r^*"
33apply (simp add: rtrancl_alt_def)
34apply (blast intro: alt_rtrancl_lemma1)
35done
36
37lemma rtrancl_subset_rtrancl_alt: "r^* \<subseteq> rtrancl_alt(field(r),r)"
38apply (simp add: rtrancl_alt_def, clarify)
39apply (frule rtrancl_type [THEN subsetD], clarify, simp)
40apply (erule rtrancl_induct)
41 txt\<open>Base case, trivial\<close>
42 apply (rule_tac x=0 in bexI)
43  apply (rule_tac x="\<lambda>x\<in>1. xa" in bexI)
44   apply simp_all
45txt\<open>Inductive step\<close>
46apply clarify
47apply (rename_tac n f)
48apply (rule_tac x="succ(n)" in bexI)
49 apply (rule_tac x="\<lambda>i\<in>succ(succ(n)). if i=succ(n) then z else f`i" in bexI)
50  apply (simp add: Ord_succ_mem_iff nat_0_le [THEN ltD] leI [THEN ltD] ltI)
51  apply (blast intro: mem_asym)
52 apply typecheck
53 apply auto
54done
55
56lemma rtrancl_alt_eq_rtrancl: "rtrancl_alt(field(r),r) = r^*"
57by (blast del: subsetI
58          intro: rtrancl_alt_subset_rtrancl rtrancl_subset_rtrancl_alt)
59
60
61definition
62  rtran_closure_mem :: "[i=>o,i,i,i] => o" where
63    \<comment> \<open>The property of belonging to \<open>rtran_closure(r)\<close>\<close>
64    "rtran_closure_mem(M,A,r,p) ==
65              \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. 
66               omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
67               (\<exists>f[M]. typed_function(M,n',A,f) &
68                (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
69                  fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
70                  (\<forall>j[M]. j\<in>n \<longrightarrow> 
71                    (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. 
72                      fun_apply(M,f,j,fj) & successor(M,j,sj) &
73                      fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"
74
75definition
76  rtran_closure :: "[i=>o,i,i] => o" where
77    "rtran_closure(M,r,s) == 
78        \<forall>A[M]. is_field(M,r,A) \<longrightarrow>
79         (\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))"
80
81definition
82  tran_closure :: "[i=>o,i,i] => o" where
83    "tran_closure(M,r,t) ==
84         \<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)"
85    
86locale M_trancl = M_basic +
87  assumes rtrancl_separation:
88         "[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))"
89      and wellfounded_trancl_separation:
90         "[| M(r); M(Z) |] ==> 
91          separation (M, \<lambda>x. 
92              \<exists>w[M]. \<exists>wx[M]. \<exists>rp[M]. 
93               w \<in> Z & pair(M,w,x,wx) & tran_closure(M,r,rp) & wx \<in> rp)"
94      and M_nat [iff] : "M(nat)"
95
96lemma (in M_trancl) rtran_closure_mem_iff:
97     "[|M(A); M(r); M(p)|]
98      ==> rtran_closure_mem(M,A,r,p) \<longleftrightarrow>
99          (\<exists>n[M]. n\<in>nat & 
100           (\<exists>f[M]. f \<in> succ(n) -> A &
101            (\<exists>x[M]. \<exists>y[M]. p = <x,y> & f`0 = x & f`n = y) &
102                           (\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)))"
103  apply (simp add: rtran_closure_mem_def Ord_succ_mem_iff nat_0_le [THEN ltD]) 
104done
105
106lemma (in M_trancl) rtran_closure_rtrancl:
107     "M(r) ==> rtran_closure(M,r,rtrancl(r))"
108apply (simp add: rtran_closure_def rtran_closure_mem_iff 
109                 rtrancl_alt_eq_rtrancl [symmetric] rtrancl_alt_def)
110apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) 
111done
112
113lemma (in M_trancl) rtrancl_closed [intro,simp]:
114     "M(r) ==> M(rtrancl(r))"
115apply (insert rtrancl_separation [of r "field(r)"])
116apply (simp add: rtrancl_alt_eq_rtrancl [symmetric]
117                 rtrancl_alt_def rtran_closure_mem_iff)
118done
119
120lemma (in M_trancl) rtrancl_abs [simp]:
121     "[| M(r); M(z) |] ==> rtran_closure(M,r,z) \<longleftrightarrow> z = rtrancl(r)"
122apply (rule iffI)
123 txt\<open>Proving the right-to-left implication\<close>
124 prefer 2 apply (blast intro: rtran_closure_rtrancl)
125apply (rule M_equalityI)
126apply (simp add: rtran_closure_def rtrancl_alt_eq_rtrancl [symmetric]
127                 rtrancl_alt_def rtran_closure_mem_iff)
128apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) 
129done
130
131lemma (in M_trancl) trancl_closed [intro,simp]:
132     "M(r) ==> M(trancl(r))"
133by (simp add: trancl_def)
134
135lemma (in M_trancl) trancl_abs [simp]:
136     "[| M(r); M(z) |] ==> tran_closure(M,r,z) \<longleftrightarrow> z = trancl(r)"
137by (simp add: tran_closure_def trancl_def)
138
139lemma (in M_trancl) wellfounded_trancl_separation':
140     "[| M(r); M(Z) |] ==> separation (M, \<lambda>x. \<exists>w[M]. w \<in> Z & <w,x> \<in> r^+)"
141by (insert wellfounded_trancl_separation [of r Z], simp) 
142
143text\<open>Alternative proof of \<open>wf_on_trancl\<close>; inspiration for the
144      relativized version.  Original version is on theory WF.\<close>
145lemma "[| wf[A](r);  r-``A \<subseteq> A |] ==> wf[A](r^+)"
146apply (simp add: wf_on_def wf_def)
147apply (safe)
148apply (drule_tac x = "{x\<in>A. \<exists>w. \<langle>w,x\<rangle> \<in> r^+ & w \<in> Z}" in spec)
149apply (blast elim: tranclE)
150done
151
152lemma (in M_trancl) wellfounded_on_trancl:
153     "[| wellfounded_on(M,A,r);  r-``A \<subseteq> A; M(r); M(A) |]
154      ==> wellfounded_on(M,A,r^+)"
155apply (simp add: wellfounded_on_def)
156apply (safe intro!: equalityI)
157apply (rename_tac Z x)
158apply (subgoal_tac "M({x\<in>A. \<exists>w[M]. w \<in> Z & \<langle>w,x\<rangle> \<in> r^+})")
159 prefer 2
160 apply (blast intro: wellfounded_trancl_separation') 
161apply (drule_tac x = "{x\<in>A. \<exists>w[M]. w \<in> Z & \<langle>w,x\<rangle> \<in> r^+}" in rspec, safe)
162apply (blast dest: transM, simp)
163apply (rename_tac y w)
164apply (drule_tac x=w in bspec, assumption, clarify)
165apply (erule tranclE)
166  apply (blast dest: transM)   (*transM is needed to prove M(xa)*)
167 apply blast
168done
169
170lemma (in M_trancl) wellfounded_trancl:
171     "[|wellfounded(M,r); M(r)|] ==> wellfounded(M,r^+)"
172apply (simp add: wellfounded_iff_wellfounded_on_field)
173apply (rule wellfounded_on_subset_A, erule wellfounded_on_trancl)
174   apply blast
175  apply (simp_all add: trancl_type [THEN field_rel_subset])
176done
177
178
179text\<open>Absoluteness for wfrec-defined functions.\<close>
180
181(*first use is_recfun, then M_is_recfun*)
182
183lemma (in M_trancl) wfrec_relativize:
184  "[|wf(r); M(a); M(r);  
185     strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
186          pair(M,x,y,z) & 
187          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
188          y = H(x, restrict(g, r -`` {x}))); 
189     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
190   ==> wfrec(r,a,H) = z \<longleftrightarrow> 
191       (\<exists>f[M]. is_recfun(r^+, a, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & 
192            z = H(a,restrict(f,r-``{a})))"
193apply (frule wf_trancl) 
194apply (simp add: wftrec_def wfrec_def, safe)
195 apply (frule wf_exists_is_recfun 
196              [of concl: "r^+" a "\<lambda>x f. H(x, restrict(f, r -`` {x}))"]) 
197      apply (simp_all add: trans_trancl function_restrictI trancl_subset_times)
198 apply (clarify, rule_tac x=x in rexI) 
199 apply (simp_all add: the_recfun_eq trans_trancl trancl_subset_times)
200done
201
202
203text\<open>Assuming \<^term>\<open>r\<close> is transitive simplifies the occurrences of \<open>H\<close>.
204      The premise \<^term>\<open>relation(r)\<close> is necessary 
205      before we can replace \<^term>\<open>r^+\<close> by \<^term>\<open>r\<close>.\<close>
206theorem (in M_trancl) trans_wfrec_relativize:
207  "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);
208     wfrec_replacement(M,MH,r);  relation2(M,MH,H);
209     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
210   ==> wfrec(r,a,H) = z \<longleftrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))" 
211apply (frule wfrec_replacement', assumption+) 
212apply (simp cong: is_recfun_cong
213           add: wfrec_relativize trancl_eq_r
214                is_recfun_restrict_idem domain_restrict_idem)
215done
216
217theorem (in M_trancl) trans_wfrec_abs:
218  "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);  M(z);
219     wfrec_replacement(M,MH,r);  relation2(M,MH,H);
220     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
221   ==> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> z=wfrec(r,a,H)" 
222by (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast) 
223
224
225lemma (in M_trancl) trans_eq_pair_wfrec_iff:
226  "[|wf(r);  trans(r); relation(r); M(r);  M(y); 
227     wfrec_replacement(M,MH,r);  relation2(M,MH,H);
228     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
229   ==> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
230       (\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
231apply safe 
232 apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) 
233txt\<open>converse direction\<close>
234apply (rule sym)
235apply (simp add: trans_wfrec_relativize, blast) 
236done
237
238
239subsection\<open>M is closed under well-founded recursion\<close>
240
241text\<open>Lemma with the awkward premise mentioning \<open>wfrec\<close>.\<close>
242lemma (in M_trancl) wfrec_closed_lemma [rule_format]:
243     "[|wf(r); M(r); 
244        strong_replacement(M, \<lambda>x y. y = \<langle>x, wfrec(r, x, H)\<rangle>);
245        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
246      ==> M(a) \<longrightarrow> M(wfrec(r,a,H))"
247apply (rule_tac a=a in wf_induct, assumption+)
248apply (subst wfrec, assumption, clarify)
249apply (drule_tac x1=x and x="\<lambda>x\<in>r -`` {x}. wfrec(r, x, H)" 
250       in rspec [THEN rspec]) 
251apply (simp_all add: function_lam) 
252apply (blast intro: lam_closed dest: pair_components_in_M) 
253done
254
255text\<open>Eliminates one instance of replacement.\<close>
256lemma (in M_trancl) wfrec_replacement_iff:
257     "strong_replacement(M, \<lambda>x z. 
258          \<exists>y[M]. pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))) \<longleftrightarrow>
259      strong_replacement(M, 
260           \<lambda>x y. \<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
261apply simp 
262apply (rule strong_replacement_cong, blast) 
263done
264
265text\<open>Useful version for transitive relations\<close>
266theorem (in M_trancl) trans_wfrec_closed:
267     "[|wf(r); trans(r); relation(r); M(r); M(a);
268       wfrec_replacement(M,MH,r);  relation2(M,MH,H);
269        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
270      ==> M(wfrec(r,a,H))"
271apply (frule wfrec_replacement', assumption+) 
272apply (frule wfrec_replacement_iff [THEN iffD1]) 
273apply (rule wfrec_closed_lemma, assumption+) 
274apply (simp_all add: wfrec_replacement_iff trans_eq_pair_wfrec_iff) 
275done
276
277subsection\<open>Absoluteness without assuming transitivity\<close>
278lemma (in M_trancl) eq_pair_wfrec_iff:
279  "[|wf(r);  M(r);  M(y); 
280     strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
281          pair(M,x,y,z) & 
282          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
283          y = H(x, restrict(g, r -`` {x}))); 
284     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
285   ==> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
286       (\<exists>f[M]. is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & 
287            y = <x, H(x,restrict(f,r-``{x}))>)"
288apply safe  
289 apply (simp add: wfrec_relativize [THEN iff_sym, of concl: _ x]) 
290txt\<open>converse direction\<close>
291apply (rule sym)
292apply (simp add: wfrec_relativize, blast) 
293done
294
295text\<open>Full version not assuming transitivity, but maybe not very useful.\<close>
296theorem (in M_trancl) wfrec_closed:
297     "[|wf(r); M(r); M(a);
298        wfrec_replacement(M,MH,r^+);  
299        relation2(M,MH, \<lambda>x f. H(x, restrict(f, r -`` {x})));
300        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
301      ==> M(wfrec(r,a,H))"
302apply (frule wfrec_replacement' 
303               [of MH "r^+" "\<lambda>x f. H(x, restrict(f, r -`` {x}))"])
304   prefer 4
305   apply (frule wfrec_replacement_iff [THEN iffD1]) 
306   apply (rule wfrec_closed_lemma, assumption+) 
307     apply (simp_all add: eq_pair_wfrec_iff func.function_restrictI) 
308done
309
310end
311