src/ZF/Constructible/WF_absolute.thy
changeset 76213 e44d86131648
parent 71417 89d05db6dd1f
child 76214 0c18df79b1c8
equal deleted inserted replaced
76212:f2094906e491 76213:e44d86131648
     8 
     8 
     9 subsection\<open>Transitive closure without fixedpoints\<close>
     9 subsection\<open>Transitive closure without fixedpoints\<close>
    10 
    10 
    11 definition
    11 definition
    12   rtrancl_alt :: "[i,i]=>i" where
    12   rtrancl_alt :: "[i,i]=>i" where
    13     "rtrancl_alt(A,r) ==
    13     "rtrancl_alt(A,r) \<equiv>
    14        {p \<in> A*A. \<exists>n\<in>nat. \<exists>f \<in> succ(n) -> A.
    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) &
    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)}"
    16                        (\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)}"
    17 
    17 
    18 lemma alt_rtrancl_lemma1 [rule_format]:
    18 lemma alt_rtrancl_lemma1 [rule_format]:
    19     "n \<in> nat
    19     "n \<in> nat
    20      ==> \<forall>f \<in> succ(n) -> field(r).
    20      \<Longrightarrow> \<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^*"
    21          (\<forall>i\<in>n. \<langle>f`i, f ` succ(i)\<rangle> \<in> r) \<longrightarrow> \<langle>f`0, f`n\<rangle> \<in> r^*"
    22 apply (induct_tac n)
    22 apply (induct_tac n)
    23 apply (simp_all add: apply_funtype rtrancl_refl, clarify)
    23 apply (simp_all add: apply_funtype rtrancl_refl, clarify)
    24 apply (rename_tac n f)
    24 apply (rename_tac n f)
    25 apply (rule rtrancl_into_rtrancl)
    25 apply (rule rtrancl_into_rtrancl)
    59 
    59 
    60 
    60 
    61 definition
    61 definition
    62   rtran_closure_mem :: "[i=>o,i,i,i] => o" where
    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>
    63     \<comment> \<open>The property of belonging to \<open>rtran_closure(r)\<close>\<close>
    64     "rtran_closure_mem(M,A,r,p) ==
    64     "rtran_closure_mem(M,A,r,p) \<equiv>
    65               \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. 
    65               \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. 
    66                omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
    66                omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
    67                (\<exists>f[M]. typed_function(M,n',A,f) &
    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) &
    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)) &
    69                   fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
    72                       fun_apply(M,f,j,fj) & successor(M,j,sj) &
    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)))"
    73                       fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"
    74 
    74 
    75 definition
    75 definition
    76   rtran_closure :: "[i=>o,i,i] => o" where
    76   rtran_closure :: "[i=>o,i,i] => o" where
    77     "rtran_closure(M,r,s) == 
    77     "rtran_closure(M,r,s) \<equiv> 
    78         \<forall>A[M]. is_field(M,r,A) \<longrightarrow>
    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))"
    79          (\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))"
    80 
    80 
    81 definition
    81 definition
    82   tran_closure :: "[i=>o,i,i] => o" where
    82   tran_closure :: "[i=>o,i,i] => o" where
    83     "tran_closure(M,r,t) ==
    83     "tran_closure(M,r,t) \<equiv>
    84          \<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)"
    84          \<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)"
    85     
    85     
    86 locale M_trancl = M_basic +
    86 locale M_trancl = M_basic +
    87   assumes rtrancl_separation:
    87   assumes rtrancl_separation:
    88          "[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))"
    88          "\<lbrakk>M(r); M(A)\<rbrakk> \<Longrightarrow> separation (M, rtran_closure_mem(M,A,r))"
    89       and wellfounded_trancl_separation:
    89       and wellfounded_trancl_separation:
    90          "[| M(r); M(Z) |] ==> 
    90          "\<lbrakk>M(r); M(Z)\<rbrakk> \<Longrightarrow> 
    91           separation (M, \<lambda>x. 
    91           separation (M, \<lambda>x. 
    92               \<exists>w[M]. \<exists>wx[M]. \<exists>rp[M]. 
    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)"
    93                w \<in> Z & pair(M,w,x,wx) & tran_closure(M,r,rp) & wx \<in> rp)"
    94       and M_nat [iff] : "M(nat)"
    94       and M_nat [iff] : "M(nat)"
    95 
    95 
    96 lemma (in M_trancl) rtran_closure_mem_iff:
    96 lemma (in M_trancl) rtran_closure_mem_iff:
    97      "[|M(A); M(r); M(p)|]
    97      "\<lbrakk>M(A); M(r); M(p)\<rbrakk>
    98       ==> rtran_closure_mem(M,A,r,p) \<longleftrightarrow>
    98       \<Longrightarrow> rtran_closure_mem(M,A,r,p) \<longleftrightarrow>
    99           (\<exists>n[M]. n\<in>nat & 
    99           (\<exists>n[M]. n\<in>nat & 
   100            (\<exists>f[M]. f \<in> succ(n) -> A &
   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) &
   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)))"
   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]) 
   103   apply (simp add: rtran_closure_mem_def Ord_succ_mem_iff nat_0_le [THEN ltD]) 
   104 done
   104 done
   105 
   105 
   106 lemma (in M_trancl) rtran_closure_rtrancl:
   106 lemma (in M_trancl) rtran_closure_rtrancl:
   107      "M(r) ==> rtran_closure(M,r,rtrancl(r))"
   107      "M(r) \<Longrightarrow> rtran_closure(M,r,rtrancl(r))"
   108 apply (simp add: rtran_closure_def rtran_closure_mem_iff 
   108 apply (simp add: rtran_closure_def rtran_closure_mem_iff 
   109                  rtrancl_alt_eq_rtrancl [symmetric] rtrancl_alt_def)
   109                  rtrancl_alt_eq_rtrancl [symmetric] rtrancl_alt_def)
   110 apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) 
   110 apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) 
   111 done
   111 done
   112 
   112 
   113 lemma (in M_trancl) rtrancl_closed [intro,simp]:
   113 lemma (in M_trancl) rtrancl_closed [intro,simp]:
   114      "M(r) ==> M(rtrancl(r))"
   114      "M(r) \<Longrightarrow> M(rtrancl(r))"
   115 apply (insert rtrancl_separation [of r "field(r)"])
   115 apply (insert rtrancl_separation [of r "field(r)"])
   116 apply (simp add: rtrancl_alt_eq_rtrancl [symmetric]
   116 apply (simp add: rtrancl_alt_eq_rtrancl [symmetric]
   117                  rtrancl_alt_def rtran_closure_mem_iff)
   117                  rtrancl_alt_def rtran_closure_mem_iff)
   118 done
   118 done
   119 
   119 
   120 lemma (in M_trancl) rtrancl_abs [simp]:
   120 lemma (in M_trancl) rtrancl_abs [simp]:
   121      "[| M(r); M(z) |] ==> rtran_closure(M,r,z) \<longleftrightarrow> z = rtrancl(r)"
   121      "\<lbrakk>M(r); M(z)\<rbrakk> \<Longrightarrow> rtran_closure(M,r,z) \<longleftrightarrow> z = rtrancl(r)"
   122 apply (rule iffI)
   122 apply (rule iffI)
   123  txt\<open>Proving the right-to-left implication\<close>
   123  txt\<open>Proving the right-to-left implication\<close>
   124  prefer 2 apply (blast intro: rtran_closure_rtrancl)
   124  prefer 2 apply (blast intro: rtran_closure_rtrancl)
   125 apply (rule M_equalityI)
   125 apply (rule M_equalityI)
   126 apply (simp add: rtran_closure_def rtrancl_alt_eq_rtrancl [symmetric]
   126 apply (simp add: rtran_closure_def rtrancl_alt_eq_rtrancl [symmetric]
   127                  rtrancl_alt_def rtran_closure_mem_iff)
   127                  rtrancl_alt_def rtran_closure_mem_iff)
   128 apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) 
   128 apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) 
   129 done
   129 done
   130 
   130 
   131 lemma (in M_trancl) trancl_closed [intro,simp]:
   131 lemma (in M_trancl) trancl_closed [intro,simp]:
   132      "M(r) ==> M(trancl(r))"
   132      "M(r) \<Longrightarrow> M(trancl(r))"
   133 by (simp add: trancl_def)
   133 by (simp add: trancl_def)
   134 
   134 
   135 lemma (in M_trancl) trancl_abs [simp]:
   135 lemma (in M_trancl) trancl_abs [simp]:
   136      "[| M(r); M(z) |] ==> tran_closure(M,r,z) \<longleftrightarrow> z = trancl(r)"
   136      "\<lbrakk>M(r); M(z)\<rbrakk> \<Longrightarrow> tran_closure(M,r,z) \<longleftrightarrow> z = trancl(r)"
   137 by (simp add: tran_closure_def trancl_def)
   137 by (simp add: tran_closure_def trancl_def)
   138 
   138 
   139 lemma (in M_trancl) wellfounded_trancl_separation':
   139 lemma (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^+)"
   140      "\<lbrakk>M(r); M(Z)\<rbrakk> \<Longrightarrow> separation (M, \<lambda>x. \<exists>w[M]. w \<in> Z & <w,x> \<in> r^+)"
   141 by (insert wellfounded_trancl_separation [of r Z], simp) 
   141 by (insert wellfounded_trancl_separation [of r Z], simp) 
   142 
   142 
   143 text\<open>Alternative proof of \<open>wf_on_trancl\<close>; inspiration for the
   143 text\<open>Alternative proof of \<open>wf_on_trancl\<close>; inspiration for the
   144       relativized version.  Original version is on theory WF.\<close>
   144       relativized version.  Original version is on theory WF.\<close>
   145 lemma "[| wf[A](r);  r-``A \<subseteq> A |] ==> wf[A](r^+)"
   145 lemma "\<lbrakk>wf[A](r);  r-``A \<subseteq> A\<rbrakk> \<Longrightarrow> wf[A](r^+)"
   146 apply (simp add: wf_on_def wf_def)
   146 apply (simp add: wf_on_def wf_def)
   147 apply (safe)
   147 apply (safe)
   148 apply (drule_tac x = "{x\<in>A. \<exists>w. \<langle>w,x\<rangle> \<in> r^+ & w \<in> Z}" in spec)
   148 apply (drule_tac x = "{x\<in>A. \<exists>w. \<langle>w,x\<rangle> \<in> r^+ & w \<in> Z}" in spec)
   149 apply (blast elim: tranclE)
   149 apply (blast elim: tranclE)
   150 done
   150 done
   151 
   151 
   152 lemma (in M_trancl) wellfounded_on_trancl:
   152 lemma (in M_trancl) wellfounded_on_trancl:
   153      "[| wellfounded_on(M,A,r);  r-``A \<subseteq> A; M(r); M(A) |]
   153      "\<lbrakk>wellfounded_on(M,A,r);  r-``A \<subseteq> A; M(r); M(A)\<rbrakk>
   154       ==> wellfounded_on(M,A,r^+)"
   154       \<Longrightarrow> wellfounded_on(M,A,r^+)"
   155 apply (simp add: wellfounded_on_def)
   155 apply (simp add: wellfounded_on_def)
   156 apply (safe intro!: equalityI)
   156 apply (safe intro!: equalityI)
   157 apply (rename_tac Z x)
   157 apply (rename_tac Z x)
   158 apply (subgoal_tac "M({x\<in>A. \<exists>w[M]. w \<in> Z & \<langle>w,x\<rangle> \<in> r^+})")
   158 apply (subgoal_tac "M({x\<in>A. \<exists>w[M]. w \<in> Z & \<langle>w,x\<rangle> \<in> r^+})")
   159  prefer 2
   159  prefer 2
   166   apply (blast dest: transM)   (*transM is needed to prove M(xa)*)
   166   apply (blast dest: transM)   (*transM is needed to prove M(xa)*)
   167  apply blast
   167  apply blast
   168 done
   168 done
   169 
   169 
   170 lemma (in M_trancl) wellfounded_trancl:
   170 lemma (in M_trancl) wellfounded_trancl:
   171      "[|wellfounded(M,r); M(r)|] ==> wellfounded(M,r^+)"
   171      "\<lbrakk>wellfounded(M,r); M(r)\<rbrakk> \<Longrightarrow> wellfounded(M,r^+)"
   172 apply (simp add: wellfounded_iff_wellfounded_on_field)
   172 apply (simp add: wellfounded_iff_wellfounded_on_field)
   173 apply (rule wellfounded_on_subset_A, erule wellfounded_on_trancl)
   173 apply (rule wellfounded_on_subset_A, erule wellfounded_on_trancl)
   174    apply blast
   174    apply blast
   175   apply (simp_all add: trancl_type [THEN field_rel_subset])
   175   apply (simp_all add: trancl_type [THEN field_rel_subset])
   176 done
   176 done
   179 text\<open>Absoluteness for wfrec-defined functions.\<close>
   179 text\<open>Absoluteness for wfrec-defined functions.\<close>
   180 
   180 
   181 (*first use is_recfun, then M_is_recfun*)
   181 (*first use is_recfun, then M_is_recfun*)
   182 
   182 
   183 lemma (in M_trancl) wfrec_relativize:
   183 lemma (in M_trancl) wfrec_relativize:
   184   "[|wf(r); M(a); M(r);  
   184   "\<lbrakk>wf(r); M(a); M(r);  
   185      strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
   185      strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
   186           pair(M,x,y,z) & 
   186           pair(M,x,y,z) & 
   187           is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
   187           is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
   188           y = H(x, restrict(g, r -`` {x}))); 
   188           y = H(x, restrict(g, r -`` {x}))); 
   189      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
   189      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
   190    ==> wfrec(r,a,H) = z \<longleftrightarrow> 
   190    \<Longrightarrow> wfrec(r,a,H) = z \<longleftrightarrow> 
   191        (\<exists>f[M]. is_recfun(r^+, a, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & 
   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})))"
   192             z = H(a,restrict(f,r-``{a})))"
   193 apply (frule wf_trancl) 
   193 apply (frule wf_trancl) 
   194 apply (simp add: wftrec_def wfrec_def, safe)
   194 apply (simp add: wftrec_def wfrec_def, safe)
   195  apply (frule wf_exists_is_recfun 
   195  apply (frule wf_exists_is_recfun 
   202 
   202 
   203 text\<open>Assuming \<^term>\<open>r\<close> is transitive simplifies the occurrences of \<open>H\<close>.
   203 text\<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 
   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>
   205       before we can replace \<^term>\<open>r^+\<close> by \<^term>\<open>r\<close>.\<close>
   206 theorem (in M_trancl) trans_wfrec_relativize:
   206 theorem (in M_trancl) trans_wfrec_relativize:
   207   "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);
   207   "\<lbrakk>wf(r);  trans(r);  relation(r);  M(r);  M(a);
   208      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
   208      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
   209      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
   209      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
   210    ==> wfrec(r,a,H) = z \<longleftrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))" 
   210    \<Longrightarrow> wfrec(r,a,H) = z \<longleftrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))" 
   211 apply (frule wfrec_replacement', assumption+) 
   211 apply (frule wfrec_replacement', assumption+) 
   212 apply (simp cong: is_recfun_cong
   212 apply (simp cong: is_recfun_cong
   213            add: wfrec_relativize trancl_eq_r
   213            add: wfrec_relativize trancl_eq_r
   214                 is_recfun_restrict_idem domain_restrict_idem)
   214                 is_recfun_restrict_idem domain_restrict_idem)
   215 done
   215 done
   216 
   216 
   217 theorem (in M_trancl) trans_wfrec_abs:
   217 theorem (in M_trancl) trans_wfrec_abs:
   218   "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);  M(z);
   218   "\<lbrakk>wf(r);  trans(r);  relation(r);  M(r);  M(a);  M(z);
   219      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
   219      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
   220      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
   220      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
   221    ==> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> z=wfrec(r,a,H)" 
   221    \<Longrightarrow> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> z=wfrec(r,a,H)" 
   222 by (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast) 
   222 by (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast) 
   223 
   223 
   224 
   224 
   225 lemma (in M_trancl) trans_eq_pair_wfrec_iff:
   225 lemma (in M_trancl) trans_eq_pair_wfrec_iff:
   226   "[|wf(r);  trans(r); relation(r); M(r);  M(y); 
   226   "\<lbrakk>wf(r);  trans(r); relation(r); M(r);  M(y); 
   227      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
   227      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
   228      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
   228      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
   229    ==> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
   229    \<Longrightarrow> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
   230        (\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
   230        (\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
   231 apply safe 
   231 apply safe 
   232  apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) 
   232  apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) 
   233 txt\<open>converse direction\<close>
   233 txt\<open>converse direction\<close>
   234 apply (rule sym)
   234 apply (rule sym)
   238 
   238 
   239 subsection\<open>M is closed under well-founded recursion\<close>
   239 subsection\<open>M is closed under well-founded recursion\<close>
   240 
   240 
   241 text\<open>Lemma with the awkward premise mentioning \<open>wfrec\<close>.\<close>
   241 text\<open>Lemma with the awkward premise mentioning \<open>wfrec\<close>.\<close>
   242 lemma (in M_trancl) wfrec_closed_lemma [rule_format]:
   242 lemma (in M_trancl) wfrec_closed_lemma [rule_format]:
   243      "[|wf(r); M(r); 
   243      "\<lbrakk>wf(r); M(r); 
   244         strong_replacement(M, \<lambda>x y. y = \<langle>x, wfrec(r, x, H)\<rangle>);
   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)) |] 
   245         \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
   246       ==> M(a) \<longrightarrow> M(wfrec(r,a,H))"
   246       \<Longrightarrow> M(a) \<longrightarrow> M(wfrec(r,a,H))"
   247 apply (rule_tac a=a in wf_induct, assumption+)
   247 apply (rule_tac a=a in wf_induct, assumption+)
   248 apply (subst wfrec, assumption, clarify)
   248 apply (subst wfrec, assumption, clarify)
   249 apply (drule_tac x1=x and x="\<lambda>x\<in>r -`` {x}. wfrec(r, x, H)" 
   249 apply (drule_tac x1=x and x="\<lambda>x\<in>r -`` {x}. wfrec(r, x, H)" 
   250        in rspec [THEN rspec]) 
   250        in rspec [THEN rspec]) 
   251 apply (simp_all add: function_lam) 
   251 apply (simp_all add: function_lam) 
   262 apply (rule strong_replacement_cong, blast) 
   262 apply (rule strong_replacement_cong, blast) 
   263 done
   263 done
   264 
   264 
   265 text\<open>Useful version for transitive relations\<close>
   265 text\<open>Useful version for transitive relations\<close>
   266 theorem (in M_trancl) trans_wfrec_closed:
   266 theorem (in M_trancl) trans_wfrec_closed:
   267      "[|wf(r); trans(r); relation(r); M(r); M(a);
   267      "\<lbrakk>wf(r); trans(r); relation(r); M(r); M(a);
   268        wfrec_replacement(M,MH,r);  relation2(M,MH,H);
   268        wfrec_replacement(M,MH,r);  relation2(M,MH,H);
   269         \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
   269         \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
   270       ==> M(wfrec(r,a,H))"
   270       \<Longrightarrow> M(wfrec(r,a,H))"
   271 apply (frule wfrec_replacement', assumption+) 
   271 apply (frule wfrec_replacement', assumption+) 
   272 apply (frule wfrec_replacement_iff [THEN iffD1]) 
   272 apply (frule wfrec_replacement_iff [THEN iffD1]) 
   273 apply (rule wfrec_closed_lemma, assumption+) 
   273 apply (rule wfrec_closed_lemma, assumption+) 
   274 apply (simp_all add: wfrec_replacement_iff trans_eq_pair_wfrec_iff) 
   274 apply (simp_all add: wfrec_replacement_iff trans_eq_pair_wfrec_iff) 
   275 done
   275 done
   276 
   276 
   277 subsection\<open>Absoluteness without assuming transitivity\<close>
   277 subsection\<open>Absoluteness without assuming transitivity\<close>
   278 lemma (in M_trancl) eq_pair_wfrec_iff:
   278 lemma (in M_trancl) eq_pair_wfrec_iff:
   279   "[|wf(r);  M(r);  M(y); 
   279   "\<lbrakk>wf(r);  M(r);  M(y); 
   280      strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
   280      strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
   281           pair(M,x,y,z) & 
   281           pair(M,x,y,z) & 
   282           is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
   282           is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
   283           y = H(x, restrict(g, r -`` {x}))); 
   283           y = H(x, restrict(g, r -`` {x}))); 
   284      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
   284      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
   285    ==> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
   285    \<Longrightarrow> 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) & 
   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}))>)"
   287             y = <x, H(x,restrict(f,r-``{x}))>)"
   288 apply safe  
   288 apply safe  
   289  apply (simp add: wfrec_relativize [THEN iff_sym, of concl: _ x]) 
   289  apply (simp add: wfrec_relativize [THEN iff_sym, of concl: _ x]) 
   290 txt\<open>converse direction\<close>
   290 txt\<open>converse direction\<close>
   292 apply (simp add: wfrec_relativize, blast) 
   292 apply (simp add: wfrec_relativize, blast) 
   293 done
   293 done
   294 
   294 
   295 text\<open>Full version not assuming transitivity, but maybe not very useful.\<close>
   295 text\<open>Full version not assuming transitivity, but maybe not very useful.\<close>
   296 theorem (in M_trancl) wfrec_closed:
   296 theorem (in M_trancl) wfrec_closed:
   297      "[|wf(r); M(r); M(a);
   297      "\<lbrakk>wf(r); M(r); M(a);
   298         wfrec_replacement(M,MH,r^+);  
   298         wfrec_replacement(M,MH,r^+);  
   299         relation2(M,MH, \<lambda>x f. H(x, restrict(f, r -`` {x})));
   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)) |] 
   300         \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
   301       ==> M(wfrec(r,a,H))"
   301       \<Longrightarrow> M(wfrec(r,a,H))"
   302 apply (frule wfrec_replacement' 
   302 apply (frule wfrec_replacement' 
   303                [of MH "r^+" "\<lambda>x f. H(x, restrict(f, r -`` {x}))"])
   303                [of MH "r^+" "\<lambda>x f. H(x, restrict(f, r -`` {x}))"])
   304    prefer 4
   304    prefer 4
   305    apply (frule wfrec_replacement_iff [THEN iffD1]) 
   305    apply (frule wfrec_replacement_iff [THEN iffD1]) 
   306    apply (rule wfrec_closed_lemma, assumption+) 
   306    apply (rule wfrec_closed_lemma, assumption+)