towards relativization of "iterates" and "wfrec"
authorpaulson
Fri, 12 Jul 2002 16:41:39 +0200
changeset 13353 1800e7134d2e
parent 13352 3cd767f8d78b
child 13354 8c8eafb63746
towards relativization of "iterates" and "wfrec"
src/ZF/Constructible/Datatype_absolute.thy
src/ZF/Constructible/Relative.thy
src/ZF/Constructible/WF_absolute.thy
src/ZF/Constructible/WFrec.thy
--- a/src/ZF/Constructible/Datatype_absolute.thy	Fri Jul 12 11:24:40 2002 +0200
+++ b/src/ZF/Constructible/Datatype_absolute.thy	Fri Jul 12 16:41:39 2002 +0200
@@ -92,33 +92,58 @@
 
 subsection {*Absoluteness for "Iterates"*}
 
-lemma (in M_trancl) iterates_relativize:
-  "[|n \<in> nat; M(v); \<forall>x[M]. M(F(x));
-     strong_replacement(M, 
-       \<lambda>x z. \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) &
-              M_is_recfun(M, \<lambda>n f z. z = nat_case(v, \<lambda>m. F(f`m), n), 
-                          Memrel(succ(n)), x, g) &
-              y = nat_case(v, \<lambda>m. F(g`m), x))|] 
-   ==> iterates(F,n,v) = z <-> 
-       (\<exists>g[M]. is_recfun(Memrel(succ(n)), n, 
-                             \<lambda>n g. nat_case(v, \<lambda>m. F(g`m), n), g) &
-            z = nat_case(v, \<lambda>m. F(g`m), n))"
-by (simp add: iterates_nat_def recursor_def transrec_def 
-              eclose_sing_Ord_eq trans_wfrec_relativize nat_into_M
-              wf_Memrel trans_Memrel relation_Memrel
-              is_recfun_abs [of "\<lambda>n g. nat_case(v, \<lambda>m. F(g`m), n)"])
+constdefs
+
+  iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o"
+   "iterates_MH(M,isF,v,n,g,z) ==
+        is_nat_case(M, v, \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) & isF(gm,u),
+                    n, z)"
+
+  iterates_replacement :: "[i=>o, [i,i]=>o, i] => o"
+   "iterates_replacement(M,isF,v) ==
+      \<forall>n[M]. n\<in>nat -->
+         wfrec_replacement(M, iterates_MH(M,isF,v), Memrel(succ(n)))"
+
+lemma (in M_axioms) iterates_MH_abs:
+  "[| relativize1(M,isF,F); M(n); M(g); M(z) |] 
+   ==> iterates_MH(M,isF,v,n,g,z) <-> z = nat_case(v, \<lambda>m. F(g`m), n)"
+apply (simp add: iterates_MH_def) 
+apply (rule nat_case_abs) 
+apply (simp_all add: relativize1_def) 
+done
+
+
+lemma (in M_axioms) iterates_imp_wfrec_replacement:
+  "[|relativize1(M,isF,F); n \<in> nat; iterates_replacement(M,isF,v)|] 
+   ==> wfrec_replacement(M, \<lambda>n f z. z = nat_case(v, \<lambda>m. F(f`m), n), 
+                       Memrel(succ(n)))" 
+by (simp add: iterates_replacement_def iterates_MH_abs)
+
+theorem (in M_trancl) iterates_abs:
+  "[| iterates_replacement(M,isF,v); relativize1(M,isF,F);
+      n \<in> nat; M(v); M(z); \<forall>x[M]. M(F(x)) |] 
+   ==> is_wfrec(M, iterates_MH(M,isF,v), Memrel(succ(n)), n, z) <->
+       z = iterates(F,n,v)" 
+apply (frule iterates_imp_wfrec_replacement, assumption+)
+apply (simp add: wf_Memrel trans_Memrel relation_Memrel nat_into_M
+                 relativize2_def iterates_MH_abs 
+                 iterates_nat_def recursor_def transrec_def 
+                 eclose_sing_Ord_eq nat_into_M
+         trans_wfrec_abs [of _ _ _ _ "\<lambda>n g. nat_case(v, \<lambda>m. F(g`m), n)"])
+done
+
 
 lemma (in M_wfrank) iterates_closed [intro,simp]:
-  "[|n \<in> nat; M(v); \<forall>x[M]. M(F(x));
-     strong_replacement(M, 
-       \<lambda>x z. \<exists>y[M]. \<exists>g[M]. pair(M, x, y, z) &
-              is_recfun (Memrel(succ(n)), x,
-                         \<lambda>n f. nat_case(v, \<lambda>m. F(f`m), n), g) &
-              y = nat_case(v, \<lambda>m. F(g`m), x))|] 
+  "[| iterates_replacement(M,isF,v); relativize1(M,isF,F);
+      n \<in> nat; M(v); \<forall>x[M]. M(F(x)) |] 
    ==> M(iterates(F,n,v))"
-by (simp add: iterates_nat_def recursor_def transrec_def 
-              eclose_sing_Ord_eq trans_wfrec_closed nat_into_M
-              wf_Memrel trans_Memrel relation_Memrel nat_case_closed)
+apply (frule iterates_imp_wfrec_replacement, assumption+)
+apply (simp add: wf_Memrel trans_Memrel relation_Memrel nat_into_M
+                 relativize2_def iterates_MH_abs 
+                 iterates_nat_def recursor_def transrec_def 
+                 eclose_sing_Ord_eq nat_into_M
+         trans_wfrec_closed [of _ _ _ "\<lambda>n g. nat_case(v, \<lambda>m. F(g`m), n)"])
+done
 
 
 constdefs
@@ -127,60 +152,43 @@
         \<exists>n1[M]. \<exists>AX[M]. 
          number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)"
 
-  list_functor_case :: "[i=>o,i,i,i,i] => o"
-    --{*Abbreviation for the definition of lists below*}
-    "list_functor_case(M,A,g,x,y) == 
-        is_nat_case(M, 0, 
-             \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) & is_list_functor(M,A,gm,u),
-             x, y)"
-
 lemma (in M_axioms) list_functor_abs [simp]: 
      "[| M(A); M(X); M(Z) |] ==> is_list_functor(M,A,X,Z) <-> (Z = {0} + A*X)"
 by (simp add: is_list_functor_def singleton_0 nat_into_M)
 
-lemma (in M_axioms) list_functor_case_abs: 
-     "[| M(A); M(n); M(y); M(g) |] 
-      ==> list_functor_case(M,A,g,n,y) <-> 
-          y = nat_case(0, \<lambda>m. {0} + A * g`m, n)"
-by (simp add: list_functor_case_def nat_into_M)
-
 
 locale M_datatypes = M_wfrank +
-  assumes list_replacement1: 
-       "[|M(A); n \<in> nat|] ==> 
-	strong_replacement(M, 
-	  \<lambda>x z. \<exists>y[M]. \<exists>g[M]. \<exists>sucn[M]. \<exists>memr[M]. 
-		 pair(M,x,y,z) & successor(M,n,sucn) & 
-		 membership(M,sucn,memr) &
-		 M_is_recfun(M, \<lambda>n f z. list_functor_case(M,A,f,n,z), 
-                             memr, x, g) &
-                 list_functor_case(M,A,g,x,y))"
-(*THEY NEED RELATIVIZATION*)
-      and list_replacement2: 
-           "M(A) ==> strong_replacement(M, \<lambda>x y. y = (\<lambda>X. {0} + A * X)^x (0))"
+ assumes list_replacement1: 
+   "M(A) ==> \<exists>zero[M]. empty(M,zero) & 
+		       iterates_replacement(M, is_list_functor(M,A), zero)"
+  and list_replacement2: 
+   "M(A) ==> 
+    \<exists>zero[M]. empty(M,zero) & 
+	      strong_replacement(M, 
+         \<lambda>n y. n\<in>nat & 
+               (\<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) & membership(M,sn,msn) &
+               is_wfrec(M, iterates_MH(M,is_list_functor(M,A), zero), 
+                        msn, n, y)))"
 
-
-
-lemma (in M_datatypes) list_replacement1':
-  "[|M(A); n \<in> nat|]
-   ==> strong_replacement
-	  (M, \<lambda>x z. \<exists>y[M]. z = \<langle>x,y\<rangle> &
-               (\<exists>g[M]. is_recfun (Memrel(succ(n)), x,
-		          \<lambda>n f. nat_case(0, \<lambda>m. {0} + A * f`m, n), g) &
- 	       y = nat_case(0, \<lambda>m. {0} + A * g ` m, x)))"
-apply (insert list_replacement1 [of A n], simp add: nat_into_M)
-apply (simp add: nat_into_M list_functor_case_abs
-                 is_recfun_abs [of "\<lambda>n f. nat_case(0, \<lambda>m. {0} + A * f`m, n)"])
+lemma (in M_datatypes) list_replacement1': 
+   "M(A) ==> iterates_replacement(M, is_list_functor(M,A), 0)"
+apply (insert list_replacement1, simp) 
 done
 
 lemma (in M_datatypes) list_replacement2': 
-  "M(A) ==> strong_replacement(M, \<lambda>x y. y = (\<lambda>X. {0} + A * X)^x (0))"
-by (insert list_replacement2, simp add: nat_into_M) 
-
+  "M(A) ==> strong_replacement(M, \<lambda>n y. n\<in>nat & y = (\<lambda>X. {0} + A * X)^n (0))"
+apply (insert list_replacement2 [of A]) 
+apply (rule strong_replacement_cong [THEN iffD1])  
+apply (rule conj_cong [OF iff_refl iterates_abs [of "is_list_functor(M,A)"]]) 
+apply (simp_all add: list_replacement1' relativize1_def) 
+done
 
 lemma (in M_datatypes) list_closed [intro,simp]:
      "M(A) ==> M(list(A))"
-by (simp add: list_eq_Union list_replacement1' list_replacement2')
+apply (insert list_replacement1)
+by  (simp add: RepFun_closed2 list_eq_Union 
+               list_replacement2' relativize1_def
+               iterates_closed [of "is_list_functor(M,A)"])
 
 
 end
--- a/src/ZF/Constructible/Relative.thy	Fri Jul 12 11:24:40 2002 +0200
+++ b/src/ZF/Constructible/Relative.thy	Fri Jul 12 16:41:39 2002 +0200
@@ -179,6 +179,16 @@
        (\<forall>m[M]. successor(M,m,k) --> is_b(m,z)) &
        (is_quasinat(M,k) | z=0)"
 
+  relativize1 :: "[i=>o, [i,i]=>o, i=>i] => o"
+    "relativize1(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. is_f(x,y) <-> y = f(x)"
+
+  relativize2 :: "[i=>o, [i,i,i]=>o, [i,i]=>i] => o"
+    "relativize2(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. is_f(x,y,z) <-> z = f(x,y)"
+
+  relativize3 :: "[i=>o, [i,i,i,i]=>o, [i,i,i]=>i] => o"
+    "relativize3(M,is_f,f) == 
+       \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M]. is_f(x,y,z,u) <-> u = f(x,y,z)"
+
 
 subsection {*The relativized ZF axioms*}
 constdefs
@@ -584,7 +594,7 @@
   nonconstructible set.  So we cannot assume that M(X) implies M(RepFun(X,f))
   even for f : M -> M.
 *)
-lemma (in M_triv_axioms) RepFun_closed [intro,simp]:
+lemma (in M_triv_axioms) RepFun_closed:
      "[| strong_replacement(M, \<lambda>x y. y = f(x)); M(A); \<forall>x\<in>A. M(f(x)) |]
       ==> M(RepFun(A,f))"
 apply (simp add: RepFun_def) 
@@ -592,10 +602,23 @@
 apply (auto dest: transM  simp add: univalent_def) 
 done
 
+lemma Replace_conj_eq: "{y . x \<in> A, x\<in>A & y=f(x)} = {y . x\<in>A, y=f(x)}"
+by simp
+
+text{*Better than @{text RepFun_closed} when having the formula @{term "x\<in>A"}
+      makes relativization easier.*}
+lemma (in M_triv_axioms) RepFun_closed2:
+     "[| strong_replacement(M, \<lambda>x y. x\<in>A & y = f(x)); M(A); \<forall>x\<in>A. M(f(x)) |]
+      ==> M(RepFun(A, %x. f(x)))"
+apply (simp add: RepFun_def)
+apply (frule strong_replacement_closed, assumption)
+apply (auto dest: transM  simp add: Replace_conj_eq univalent_def) 
+done
+
 lemma (in M_triv_axioms) lam_closed [intro,simp]:
      "[| strong_replacement(M, \<lambda>x y. y = <x,b(x)>); M(A); \<forall>x\<in>A. M(b(x)) |]
       ==> M(\<lambda>x\<in>A. b(x))"
-by (simp add: lam_def, blast dest: transM) 
+by (simp add: lam_def, blast intro: RepFun_closed dest: transM) 
 
 lemma (in M_triv_axioms) image_abs [simp]: 
      "[| M(r); M(A); M(z) |] ==> image(M,r,A,z) <-> z = r``A"
@@ -635,19 +658,18 @@
 by (auto simp add: is_quasinat_def quasinat_def)
 
 lemma (in M_triv_axioms) nat_case_abs [simp]: 
-  assumes b_abs: "!!x y. M(x) --> M(y) --> (is_b(x,y) <-> y = b(x))"
-  shows
-     "[| M(k); M(z) |] ==> is_nat_case(M,a,is_b,k,z) <-> z = nat_case(a,b,k)"
+     "[| relativize1(M,is_b,b); M(k); M(z) |] 
+      ==> is_nat_case(M,a,is_b,k,z) <-> z = nat_case(a,b,k)"
 apply (case_tac "quasinat(k)") 
  prefer 2 
  apply (simp add: is_nat_case_def non_nat_case) 
  apply (force simp add: quasinat_def) 
 apply (simp add: quasinat_def is_nat_case_def)
 apply (elim disjE exE) 
- apply (simp_all add: b_abs) 
+ apply (simp_all add: relativize1_def) 
 done
 
-(*Needed?  surely better to replace by nat_case?*)
+(*Needed?  surely better to replace is_nat_case by nat_case?*)
 lemma (in M_triv_axioms) is_nat_case_cong [cong]:
      "[| a = a'; k = k';  z = z';  M(z');
        !!x y. [| M(x); M(y) |] ==> is_b(x,y) <-> is_b'(x,y) |]
@@ -980,8 +1002,7 @@
 
 lemma (in M_axioms) apply_abs [simp]: 
      "[| M(f); M(x); M(y) |] ==> fun_apply(M,f,x,y) <-> f`x = y"
-apply (simp add: fun_apply_def apply_def)
-apply (blast intro: elim:); 
+apply (simp add: fun_apply_def apply_def, blast) 
 done
 
 lemma (in M_axioms) typed_function_abs [simp]: 
--- a/src/ZF/Constructible/WF_absolute.thy	Fri Jul 12 11:24:40 2002 +0200
+++ b/src/ZF/Constructible/WF_absolute.thy	Fri Jul 12 16:41:39 2002 +0200
@@ -261,7 +261,7 @@
       separation
 	   (M, \<lambda>x. ~ (\<exists>f[M]. is_recfun(r^+, x, %x f. range(f), f)))"
 apply (insert wfrank_separation [of r])
-apply (simp add: is_recfun_abs [of "%x. range"])
+apply (simp add: relativize2_def is_recfun_abs [of "%x. range"])
 done
 
 lemma (in M_wfrank) wfrank_strong_replacement':
@@ -270,7 +270,7 @@
 		  pair(M,x,y,z) & is_recfun(r^+, x, %x f. range(f), f) &
 		  y = range(f))"
 apply (insert wfrank_strong_replacement [of r])
-apply (simp add: is_recfun_abs [of "%x. range"])
+apply (simp add: relativize2_def is_recfun_abs [of "%x. range"])
 done
 
 lemma (in M_wfrank) Ord_wfrank_separation':
@@ -278,7 +278,7 @@
       separation (M, \<lambda>x. 
          ~ (\<forall>f[M]. is_recfun(r^+, x, \<lambda>x. range, f) --> Ord(range(f))))" 
 apply (insert Ord_wfrank_separation [of r])
-apply (simp add: is_recfun_abs [of "%x. range"])
+apply (simp add: relativize2_def is_recfun_abs [of "%x. range"])
 done
 
 text{*This function, defined using replacement, is a rank function for
@@ -524,19 +524,26 @@
       before we can replace @{term "r^+"} by @{term r}. *}
 theorem (in M_trancl) trans_wfrec_relativize:
   "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);
-     strong_replacement(M, \<lambda>x z. \<exists>y[M]. 
-                pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))); 
+     wfrec_replacement(M,MH,r);  relativize2(M,MH,H);
      \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
    ==> wfrec(r,a,H) = z <-> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))" 
-by (simp cong: is_recfun_cong
-         add: wfrec_relativize trancl_eq_r
-               is_recfun_restrict_idem domain_restrict_idem)
+apply (frule wfrec_replacement', assumption+) 
+apply (simp cong: is_recfun_cong
+           add: wfrec_relativize trancl_eq_r
+                is_recfun_restrict_idem domain_restrict_idem)
+done
 
+theorem (in M_trancl) trans_wfrec_abs:
+  "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);  M(z);
+     wfrec_replacement(M,MH,r);  relativize2(M,MH,H);
+     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
+   ==> is_wfrec(M,MH,r,a,z) <-> z=wfrec(r,a,H)" 
+apply (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast) 
+done
 
 lemma (in M_trancl) trans_eq_pair_wfrec_iff:
   "[|wf(r);  trans(r); relation(r); M(r);  M(y); 
-     strong_replacement(M, \<lambda>x z. \<exists>y[M]. 
-                pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))); 
+     wfrec_replacement(M,MH,r);  relativize2(M,MH,H);
      \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
    ==> y = <x, wfrec(r, x, H)> <-> 
        (\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
@@ -566,8 +573,8 @@
 
 text{*Eliminates one instance of replacement.*}
 lemma (in M_wfrank) wfrec_replacement_iff:
-     "strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M]. 
-                pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)) <->
+     "strong_replacement(M, \<lambda>x z. 
+          \<exists>y[M]. pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))) <->
       strong_replacement(M, 
            \<lambda>x y. \<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
 apply simp 
@@ -577,11 +584,10 @@
 text{*Useful version for transitive relations*}
 theorem (in M_wfrank) trans_wfrec_closed:
      "[|wf(r); trans(r); relation(r); M(r); M(a);
-        strong_replacement(M, 
-             \<lambda>x z. \<exists>y[M]. \<exists>g[M].
-                    pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
+       wfrec_replacement(M,MH,r);  relativize2(M,MH,H);
         \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] 
       ==> M(wfrec(r,a,H))"
+apply (frule wfrec_replacement', assumption+) 
 apply (frule wfrec_replacement_iff [THEN iffD1]) 
 apply (rule wfrec_closed_lemma, assumption+) 
 apply (simp_all add: wfrec_replacement_iff trans_eq_pair_wfrec_iff) 
@@ -621,15 +627,16 @@
 text{*Full version not assuming transitivity, but maybe not very useful.*}
 theorem (in M_wfrank) wfrec_closed:
      "[|wf(r); M(r); M(a);
-     strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
-          pair(M,x,y,z) & 
-          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
-          y = H(x, restrict(g, r -`` {x}))); 
+        wfrec_replacement(M,MH,r^+);  
+        relativize2(M,MH, \<lambda>x f. H(x, restrict(f, r -`` {x})));
         \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] 
       ==> M(wfrec(r,a,H))"
-apply (frule wfrec_replacement_iff [THEN iffD1]) 
-apply (rule wfrec_closed_lemma, assumption+) 
-apply (simp_all add: eq_pair_wfrec_iff) 
+apply (frule wfrec_replacement' 
+               [of MH "r^+" "\<lambda>x f. H(x, restrict(f, r -`` {x}))"])
+   prefer 4
+   apply (frule wfrec_replacement_iff [THEN iffD1]) 
+   apply (rule wfrec_closed_lemma, assumption+) 
+     apply (simp_all add: eq_pair_wfrec_iff func.function_restrictI) 
 done
 
 end
--- a/src/ZF/Constructible/WFrec.thy	Fri Jul 12 11:24:40 2002 +0200
+++ b/src/ZF/Constructible/WFrec.thy	Fri Jul 12 16:41:39 2002 +0200
@@ -275,7 +275,7 @@
 done
 
 constdefs
- M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
+  M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
    "M_is_recfun(M,MH,r,a,f) == 
      \<forall>z[M]. z \<in> f <-> 
             (\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M]. 
@@ -283,11 +283,20 @@
                pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) &
                xa \<in> r & MH(x, f_r_sx, y))"
 
+  is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
+   "is_wfrec(M,MH,r,a,z) == 
+      \<exists>f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)"
+
+  wfrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o"
+   "wfrec_replacement(M,MH,r) ==
+        strong_replacement(M, 
+             \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))"
+
 lemma (in M_axioms) is_recfun_abs:
      "[| \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g));  M(r); M(a); M(f); 
-         \<forall>x g y. M(x) --> M(g) --> M(y) --> MH(x,g,y) <-> y = H(x,g) |] 
+         relativize2(M,MH,H) |] 
       ==> M_is_recfun(M,MH,r,a,f) <-> is_recfun(r,a,H,f)"
-apply (simp add: M_is_recfun_def is_recfun_relativize)
+apply (simp add: M_is_recfun_def relativize2_def is_recfun_relativize)
 apply (rule rall_cong)
 apply (blast dest: transM)
 done
@@ -298,7 +307,33 @@
       ==> M_is_recfun(M,MH,r,a,f) <-> M_is_recfun(M,MH',r',a',f')"
 by (simp add: M_is_recfun_def) 
 
+lemma (in M_axioms) is_wfrec_abs:
+     "[| \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)); 
+         relativize2(M,MH,H);  M(r); M(a); M(z) |]
+      ==> is_wfrec(M,MH,r,a,z) <-> 
+          (\<exists>g[M]. is_recfun(r,a,H,g) & z = H(a,g))"
+by (simp add: is_wfrec_def relativize2_def is_recfun_abs)
 
+text{*Relating @{term wfrec_replacement} to native constructs*}
+lemma (in M_axioms) wfrec_replacement':
+  "[|wfrec_replacement(M,MH,r);
+     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)); 
+     relativize2(M,MH,H);  M(r)|] 
+   ==> strong_replacement(M, \<lambda>x z. \<exists>y[M]. 
+                pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g)))"
+apply (rotate_tac 1) 
+apply (simp add: wfrec_replacement_def is_wfrec_abs) 
+done
+
+lemma wfrec_replacement_cong [cong]:
+     "[| !!x y z. [| M(x); M(y); M(z) |] ==> MH(x,y,z) <-> MH'(x,y,z);
+         r=r' |] 
+      ==> wfrec_replacement(M, %x y. MH(x,y), r) <-> 
+          wfrec_replacement(M, %x y. MH'(x,y), r')" 
+by (simp add: is_wfrec_def wfrec_replacement_def) 
+
+
+(*FIXME: update to use new techniques!!*)
 constdefs
  (*This expresses ordinal addition in the language of ZF.  It also 
    provides an abbreviation that can be used in the instance of strong
@@ -367,7 +402,7 @@
 	f \<in> a \<rightarrow> range(f) & (\<forall>x. M(x) --> x < a --> f`x = i Un f``x)"
 apply (frule lt_Ord) 
 apply (simp add: is_oadd_fun_def Memrel_closed Un_closed 
-             is_recfun_abs [of "%x g. i Un g``x"]
+             relativize2_def is_recfun_abs [of "%x g. i Un g``x"]
              image_closed is_recfun_iff_equation  
              Ball_def lt_trans [OF ltI, of _ a] lt_Memrel)
 apply (simp add: lt_def) 
@@ -382,7 +417,7 @@
 		  (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. i Un g``x,g) & 
 		  y = i Un g``x))" 
 apply (insert oadd_strong_replacement [of i j]) 
-apply (simp add: is_oadd_fun_def is_recfun_abs [of "%x g. i Un g``x"])  
+apply (simp add: is_oadd_fun_def relativize2_def is_recfun_abs [of "%x g. i Un g``x"])  
 done