author | haftmann |
Fri, 17 Jun 2005 16:12:49 +0200 | |
changeset 16417 | 9bc16273c2d4 |
parent 13721 | 2cf506c09946 |
child 21233 | 5a5c8ea5f66a |
permissions | -rw-r--r-- |
13505 | 1 |
(* Title: ZF/Constructible/WFrec.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
*) |
|
5 |
||
13306 | 6 |
header{*Relativized Well-Founded Recursion*} |
7 |
||
16417 | 8 |
theory WFrec imports Wellorderings begin |
13223 | 9 |
|
10 |
||
13506 | 11 |
subsection{*General Lemmas*} |
12 |
||
13254 | 13 |
(*Many of these might be useful in WF.thy*) |
13223 | 14 |
|
13269 | 15 |
lemma apply_recfun2: |
16 |
"[| is_recfun(r,a,H,f); <x,i>:f |] ==> i = H(x, restrict(f,r-``{x}))" |
|
17 |
apply (frule apply_recfun) |
|
18 |
apply (blast dest: is_recfun_type fun_is_rel) |
|
19 |
apply (simp add: function_apply_equality [OF _ is_recfun_imp_function]) |
|
13223 | 20 |
done |
21 |
||
22 |
text{*Expresses @{text is_recfun} as a recursion equation*} |
|
23 |
lemma is_recfun_iff_equation: |
|
24 |
"is_recfun(r,a,H,f) <-> |
|
25 |
f \<in> r -`` {a} \<rightarrow> range(f) & |
|
26 |
(\<forall>x \<in> r-``{a}. f`x = H(x, restrict(f, r-``{x})))" |
|
27 |
apply (rule iffI) |
|
28 |
apply (simp add: is_recfun_type apply_recfun Ball_def vimage_singleton_iff, |
|
29 |
clarify) |
|
30 |
apply (simp add: is_recfun_def) |
|
31 |
apply (rule fun_extension) |
|
32 |
apply assumption |
|
33 |
apply (fast intro: lam_type, simp) |
|
34 |
done |
|
35 |
||
13245 | 36 |
lemma is_recfun_imp_in_r: "[|is_recfun(r,a,H,f); \<langle>x,i\<rangle> \<in> f|] ==> \<langle>x, a\<rangle> \<in> r" |
13269 | 37 |
by (blast dest: is_recfun_type fun_is_rel) |
13245 | 38 |
|
13254 | 39 |
lemma trans_Int_eq: |
40 |
"[| trans(r); <y,x> \<in> r |] ==> r -`` {x} \<inter> r -`` {y} = r -`` {y}" |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
41 |
by (blast intro: transD) |
13223 | 42 |
|
13254 | 43 |
lemma is_recfun_restrict_idem: |
44 |
"is_recfun(r,a,H,f) ==> restrict(f, r -`` {a}) = f" |
|
45 |
apply (drule is_recfun_type) |
|
46 |
apply (auto simp add: Pi_iff subset_Sigma_imp_relation restrict_idem) |
|
47 |
done |
|
48 |
||
49 |
lemma is_recfun_cong_lemma: |
|
50 |
"[| is_recfun(r,a,H,f); r = r'; a = a'; f = f'; |
|
51 |
!!x g. [| <x,a'> \<in> r'; relation(g); domain(g) <= r' -``{x} |] |
|
52 |
==> H(x,g) = H'(x,g) |] |
|
53 |
==> is_recfun(r',a',H',f')" |
|
54 |
apply (simp add: is_recfun_def) |
|
55 |
apply (erule trans) |
|
56 |
apply (rule lam_cong) |
|
57 |
apply (simp_all add: vimage_singleton_iff Int_lower2) |
|
58 |
done |
|
59 |
||
60 |
text{*For @{text is_recfun} we need only pay attention to functions |
|
61 |
whose domains are initial segments of @{term r}.*} |
|
62 |
lemma is_recfun_cong: |
|
63 |
"[| r = r'; a = a'; f = f'; |
|
64 |
!!x g. [| <x,a'> \<in> r'; relation(g); domain(g) <= r' -``{x} |] |
|
65 |
==> H(x,g) = H'(x,g) |] |
|
66 |
==> is_recfun(r,a,H,f) <-> is_recfun(r',a',H',f')" |
|
67 |
apply (rule iffI) |
|
68 |
txt{*Messy: fast and blast don't work for some reason*} |
|
69 |
apply (erule is_recfun_cong_lemma, auto) |
|
70 |
apply (erule is_recfun_cong_lemma) |
|
71 |
apply (blast intro: sym)+ |
|
72 |
done |
|
13223 | 73 |
|
13506 | 74 |
subsection{*Reworking of the Recursion Theory Within @{term M}*} |
75 |
||
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
76 |
lemma (in M_basic) is_recfun_separation': |
13319 | 77 |
"[| f \<in> r -`` {a} \<rightarrow> range(f); g \<in> r -`` {b} \<rightarrow> range(g); |
78 |
M(r); M(f); M(g); M(a); M(b) |] |
|
79 |
==> separation(M, \<lambda>x. \<not> (\<langle>x, a\<rangle> \<in> r \<longrightarrow> \<langle>x, b\<rangle> \<in> r \<longrightarrow> f ` x = g ` x))" |
|
80 |
apply (insert is_recfun_separation [of r f g a b]) |
|
13352 | 81 |
apply (simp add: vimage_singleton_iff) |
13319 | 82 |
done |
13223 | 83 |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
84 |
text{*Stated using @{term "trans(r)"} rather than |
13223 | 85 |
@{term "transitive_rel(M,A,r)"} because the latter rewrites to |
86 |
the former anyway, by @{text transitive_rel_abs}. |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
87 |
As always, theorems should be expressed in simplified form. |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
88 |
The last three M-premises are redundant because of @{term "M(r)"}, |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
89 |
but without them we'd have to undertake |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
90 |
more work to set up the induction formula.*} |
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
91 |
lemma (in M_basic) is_recfun_equal [rule_format]: |
13223 | 92 |
"[|is_recfun(r,a,H,f); is_recfun(r,b,H,g); |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
93 |
wellfounded(M,r); trans(r); |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
94 |
M(f); M(g); M(r); M(x); M(a); M(b) |] |
13223 | 95 |
==> <x,a> \<in> r --> <x,b> \<in> r --> f`x=g`x" |
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13319
diff
changeset
|
96 |
apply (frule_tac f=f in is_recfun_type) |
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13319
diff
changeset
|
97 |
apply (frule_tac f=g in is_recfun_type) |
13223 | 98 |
apply (simp add: is_recfun_def) |
13254 | 99 |
apply (erule_tac a=x in wellfounded_induct, assumption+) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
100 |
txt{*Separation to justify the induction*} |
13319 | 101 |
apply (blast intro: is_recfun_separation') |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
102 |
txt{*Now the inductive argument itself*} |
13254 | 103 |
apply clarify |
13223 | 104 |
apply (erule ssubst)+ |
105 |
apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def) |
|
106 |
apply (rename_tac x1) |
|
107 |
apply (rule_tac t="%z. H(x1,z)" in subst_context) |
|
13721 | 108 |
apply (subgoal_tac "\<forall>y \<in> r-``{x1}. ALL z. <y,z>\<in>f <-> <y,z>\<in>g") |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
109 |
apply (blast intro: transD) |
13223 | 110 |
apply (simp add: apply_iff) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
111 |
apply (blast intro: transD sym) |
13223 | 112 |
done |
113 |
||
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
114 |
lemma (in M_basic) is_recfun_cut: |
13223 | 115 |
"[|is_recfun(r,a,H,f); is_recfun(r,b,H,g); |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
116 |
wellfounded(M,r); trans(r); |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
117 |
M(f); M(g); M(r); <b,a> \<in> r |] |
13223 | 118 |
==> restrict(f, r-``{b}) = g" |
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13319
diff
changeset
|
119 |
apply (frule_tac f=f in is_recfun_type) |
13223 | 120 |
apply (rule fun_extension) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
121 |
apply (blast intro: transD restrict_type2) |
13223 | 122 |
apply (erule is_recfun_type, simp) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
123 |
apply (blast intro: is_recfun_equal transD dest: transM) |
13223 | 124 |
done |
125 |
||
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
126 |
lemma (in M_basic) is_recfun_functional: |
13223 | 127 |
"[|is_recfun(r,a,H,f); is_recfun(r,a,H,g); |
13268 | 128 |
wellfounded(M,r); trans(r); M(f); M(g); M(r) |] ==> f=g" |
13223 | 129 |
apply (rule fun_extension) |
130 |
apply (erule is_recfun_type)+ |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
131 |
apply (blast intro!: is_recfun_equal dest: transM) |
13254 | 132 |
done |
13223 | 133 |
|
13295 | 134 |
text{*Tells us that @{text is_recfun} can (in principle) be relativized.*} |
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
135 |
lemma (in M_basic) is_recfun_relativize: |
13254 | 136 |
"[| M(r); M(f); \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
137 |
==> is_recfun(r,a,H,f) <-> |
13254 | 138 |
(\<forall>z[M]. z \<in> f <-> |
139 |
(\<exists>x[M]. <x,a> \<in> r & z = <x, H(x, restrict(f, r-``{x}))>))"; |
|
140 |
apply (simp add: is_recfun_def lam_def) |
|
13223 | 141 |
apply (safe intro!: equalityI) |
13254 | 142 |
apply (drule equalityD1 [THEN subsetD], assumption) |
143 |
apply (blast dest: pair_components_in_M) |
|
144 |
apply (blast elim!: equalityE dest: pair_components_in_M) |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13564
diff
changeset
|
145 |
apply (frule transM, assumption) |
13223 | 146 |
apply simp |
147 |
apply blast |
|
13254 | 148 |
apply (subgoal_tac "is_function(M,f)") |
149 |
txt{*We use @{term "is_function"} rather than @{term "function"} because |
|
150 |
the subgoal's easier to prove with relativized quantifiers!*} |
|
151 |
prefer 2 apply (simp add: is_function_def) |
|
13223 | 152 |
apply (frule pair_components_in_M, assumption) |
13254 | 153 |
apply (simp add: is_recfun_imp_function function_restrictI) |
13223 | 154 |
done |
155 |
||
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
156 |
lemma (in M_basic) is_recfun_restrict: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
157 |
"[| wellfounded(M,r); trans(r); is_recfun(r,x,H,f); \<langle>y,x\<rangle> \<in> r; |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
158 |
M(r); M(f); |
13254 | 159 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] |
13223 | 160 |
==> is_recfun(r, y, H, restrict(f, r -`` {y}))" |
161 |
apply (frule pair_components_in_M, assumption, clarify) |
|
13254 | 162 |
apply (simp (no_asm_simp) add: is_recfun_relativize restrict_iff |
163 |
trans_Int_eq) |
|
13223 | 164 |
apply safe |
165 |
apply (simp_all add: vimage_singleton_iff is_recfun_type [THEN apply_iff]) |
|
166 |
apply (frule_tac x=xa in pair_components_in_M, assumption) |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
167 |
apply (frule_tac x=xa in apply_recfun, blast intro: transD) |
13247 | 168 |
apply (simp add: is_recfun_type [THEN apply_iff] |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
169 |
is_recfun_imp_function function_restrictI) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
170 |
apply (blast intro: apply_recfun dest: transD) |
13223 | 171 |
done |
172 |
||
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
173 |
lemma (in M_basic) restrict_Y_lemma: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
174 |
"[| wellfounded(M,r); trans(r); M(r); |
13254 | 175 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)); M(Y); |
13299 | 176 |
\<forall>b[M]. |
13223 | 177 |
b \<in> Y <-> |
13299 | 178 |
(\<exists>x[M]. <x,a1> \<in> r & |
179 |
(\<exists>y[M]. b = \<langle>x,y\<rangle> & (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g)))); |
|
180 |
\<langle>x,a1\<rangle> \<in> r; is_recfun(r,x,H,f); M(f) |] |
|
13223 | 181 |
==> restrict(Y, r -`` {x}) = f" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
182 |
apply (subgoal_tac "\<forall>y \<in> r-``{x}. \<forall>z. <y,z>:Y <-> <y,z>:f") |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
183 |
apply (simp (no_asm_simp) add: restrict_def) |
13254 | 184 |
apply (thin_tac "rall(M,?P)")+ --{*essential for efficiency*} |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
185 |
apply (frule is_recfun_type [THEN fun_is_rel], blast) |
13223 | 186 |
apply (frule pair_components_in_M, assumption, clarify) |
187 |
apply (rule iffI) |
|
13505 | 188 |
apply (frule_tac y="<y,z>" in transM, assumption) |
13223 | 189 |
apply (clarsimp simp add: vimage_singleton_iff is_recfun_type [THEN apply_iff] |
190 |
apply_recfun is_recfun_cut) |
|
191 |
txt{*Opposite inclusion: something in f, show in Y*} |
|
13293 | 192 |
apply (frule_tac y="<y,z>" in transM, assumption) |
193 |
apply (simp add: vimage_singleton_iff) |
|
194 |
apply (rule conjI) |
|
195 |
apply (blast dest: transD) |
|
13268 | 196 |
apply (rule_tac x="restrict(f, r -`` {y})" in rexI) |
197 |
apply (simp_all add: is_recfun_restrict |
|
198 |
apply_recfun is_recfun_type [THEN apply_iff]) |
|
13223 | 199 |
done |
200 |
||
13245 | 201 |
text{*For typical applications of Replacement for recursive definitions*} |
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
202 |
lemma (in M_basic) univalent_is_recfun: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
203 |
"[|wellfounded(M,r); trans(r); M(r)|] |
13268 | 204 |
==> univalent (M, A, \<lambda>x p. |
13293 | 205 |
\<exists>y[M]. p = \<langle>x,y\<rangle> & (\<exists>f[M]. is_recfun(r,x,H,f) & y = H(x,f)))" |
13245 | 206 |
apply (simp add: univalent_def) |
207 |
apply (blast dest: is_recfun_functional) |
|
208 |
done |
|
209 |
||
13299 | 210 |
|
13223 | 211 |
text{*Proof of the inductive step for @{text exists_is_recfun}, since |
212 |
we must prove two versions.*} |
|
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
213 |
lemma (in M_basic) exists_is_recfun_indstep: |
13268 | 214 |
"[|\<forall>y. \<langle>y, a1\<rangle> \<in> r --> (\<exists>f[M]. is_recfun(r, y, H, f)); |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
215 |
wellfounded(M,r); trans(r); M(r); M(a1); |
13268 | 216 |
strong_replacement(M, \<lambda>x z. |
217 |
\<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); |
|
13254 | 218 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] |
13268 | 219 |
==> \<exists>f[M]. is_recfun(r,a1,H,f)" |
13223 | 220 |
apply (drule_tac A="r-``{a1}" in strong_replacementD) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
221 |
apply blast |
13223 | 222 |
txt{*Discharge the "univalent" obligation of Replacement*} |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
223 |
apply (simp add: univalent_is_recfun) |
13223 | 224 |
txt{*Show that the constructed object satisfies @{text is_recfun}*} |
225 |
apply clarify |
|
13268 | 226 |
apply (rule_tac x=Y in rexI) |
13254 | 227 |
txt{*Unfold only the top-level occurrence of @{term is_recfun}*} |
228 |
apply (simp (no_asm_simp) add: is_recfun_relativize [of concl: _ a1]) |
|
13268 | 229 |
txt{*The big iff-formula defining @{term Y} is now redundant*} |
13254 | 230 |
apply safe |
13299 | 231 |
apply (simp add: vimage_singleton_iff restrict_Y_lemma [of r H _ a1]) |
13223 | 232 |
txt{*one more case*} |
13254 | 233 |
apply (simp (no_asm_simp) add: Bex_def vimage_singleton_iff) |
13223 | 234 |
apply (drule_tac x1=x in spec [THEN mp], assumption, clarify) |
13268 | 235 |
apply (rename_tac f) |
236 |
apply (rule_tac x=f in rexI) |
|
13293 | 237 |
apply (simp_all add: restrict_Y_lemma [of r H]) |
13299 | 238 |
txt{*FIXME: should not be needed!*} |
239 |
apply (subst restrict_Y_lemma [of r H]) |
|
240 |
apply (simp add: vimage_singleton_iff)+ |
|
241 |
apply blast+ |
|
13223 | 242 |
done |
243 |
||
244 |
text{*Relativized version, when we have the (currently weaker) premise |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
245 |
@{term "wellfounded(M,r)"}*} |
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
246 |
lemma (in M_basic) wellfounded_exists_is_recfun: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
247 |
"[|wellfounded(M,r); trans(r); |
13268 | 248 |
separation(M, \<lambda>x. ~ (\<exists>f[M]. is_recfun(r, x, H, f))); |
249 |
strong_replacement(M, \<lambda>x z. |
|
250 |
\<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
251 |
M(r); M(a); |
13254 | 252 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] |
13268 | 253 |
==> \<exists>f[M]. is_recfun(r,a,H,f)" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
254 |
apply (rule wellfounded_induct, assumption+, clarify) |
13223 | 255 |
apply (rule exists_is_recfun_indstep, assumption+) |
256 |
done |
|
257 |
||
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
258 |
lemma (in M_basic) wf_exists_is_recfun [rule_format]: |
13268 | 259 |
"[|wf(r); trans(r); M(r); |
260 |
strong_replacement(M, \<lambda>x z. |
|
261 |
\<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); |
|
13254 | 262 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] |
13268 | 263 |
==> M(a) --> (\<exists>f[M]. is_recfun(r,a,H,f))" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
264 |
apply (rule wf_induct, assumption+) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
265 |
apply (frule wf_imp_relativized) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
266 |
apply (intro impI) |
13268 | 267 |
apply (rule exists_is_recfun_indstep) |
268 |
apply (blast dest: transM del: rev_rallE, assumption+) |
|
13223 | 269 |
done |
270 |
||
13506 | 271 |
|
272 |
subsection{*Relativization of the ZF Predicate @{term is_recfun}*} |
|
273 |
||
13223 | 274 |
constdefs |
13353 | 275 |
M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" |
13352 | 276 |
"M_is_recfun(M,MH,r,a,f) == |
13254 | 277 |
\<forall>z[M]. z \<in> f <-> |
278 |
(\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M]. |
|
279 |
pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) & |
|
280 |
pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) & |
|
13348 | 281 |
xa \<in> r & MH(x, f_r_sx, y))" |
13223 | 282 |
|
13353 | 283 |
is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o" |
284 |
"is_wfrec(M,MH,r,a,z) == |
|
285 |
\<exists>f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)" |
|
286 |
||
287 |
wfrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" |
|
288 |
"wfrec_replacement(M,MH,r) == |
|
289 |
strong_replacement(M, |
|
290 |
\<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))" |
|
291 |
||
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
292 |
lemma (in M_basic) is_recfun_abs: |
13350 | 293 |
"[| \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)); M(r); M(a); M(f); |
13634 | 294 |
relation2(M,MH,H) |] |
13352 | 295 |
==> M_is_recfun(M,MH,r,a,f) <-> is_recfun(r,a,H,f)" |
13634 | 296 |
apply (simp add: M_is_recfun_def relation2_def is_recfun_relativize) |
13254 | 297 |
apply (rule rall_cong) |
298 |
apply (blast dest: transM) |
|
13223 | 299 |
done |
300 |
||
301 |
lemma M_is_recfun_cong [cong]: |
|
302 |
"[| r = r'; a = a'; f = f'; |
|
13348 | 303 |
!!x g y. [| M(x); M(g); M(y) |] ==> MH(x,g,y) <-> MH'(x,g,y) |] |
13352 | 304 |
==> M_is_recfun(M,MH,r,a,f) <-> M_is_recfun(M,MH',r',a',f')" |
13223 | 305 |
by (simp add: M_is_recfun_def) |
306 |
||
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
307 |
lemma (in M_basic) is_wfrec_abs: |
13353 | 308 |
"[| \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)); |
13634 | 309 |
relation2(M,MH,H); M(r); M(a); M(z) |] |
13353 | 310 |
==> is_wfrec(M,MH,r,a,z) <-> |
311 |
(\<exists>g[M]. is_recfun(r,a,H,g) & z = H(a,g))" |
|
13634 | 312 |
by (simp add: is_wfrec_def relation2_def is_recfun_abs) |
13223 | 313 |
|
13353 | 314 |
text{*Relating @{term wfrec_replacement} to native constructs*} |
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
315 |
lemma (in M_basic) wfrec_replacement': |
13353 | 316 |
"[|wfrec_replacement(M,MH,r); |
317 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)); |
|
13634 | 318 |
relation2(M,MH,H); M(r)|] |
13353 | 319 |
==> strong_replacement(M, \<lambda>x z. \<exists>y[M]. |
320 |
pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g)))" |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13564
diff
changeset
|
321 |
by (simp add: wfrec_replacement_def is_wfrec_abs) |
13353 | 322 |
|
323 |
lemma wfrec_replacement_cong [cong]: |
|
324 |
"[| !!x y z. [| M(x); M(y); M(z) |] ==> MH(x,y,z) <-> MH'(x,y,z); |
|
325 |
r=r' |] |
|
326 |
==> wfrec_replacement(M, %x y. MH(x,y), r) <-> |
|
327 |
wfrec_replacement(M, %x y. MH'(x,y), r')" |
|
328 |
by (simp add: is_wfrec_def wfrec_replacement_def) |
|
329 |
||
330 |
||
13223 | 331 |
end |
332 |