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