src/ZF/Constructible/Datatype_absolute.thy
changeset 13350 626b79677dfa
parent 13348 374d05460db4
child 13352 3cd767f8d78b
--- a/src/ZF/Constructible/Datatype_absolute.thy	Thu Jul 11 16:57:14 2002 +0200
+++ b/src/ZF/Constructible/Datatype_absolute.thy	Thu Jul 11 17:18:28 2002 +0200
@@ -105,8 +105,7 @@
             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 nat_case_closed)
-
+              wf_Memrel trans_Memrel relation_Memrel)
 
 lemma (in M_wfrank) iterates_closed [intro,simp]:
   "[|n \<in> nat; M(v); \<forall>x[M]. M(F(x));
@@ -121,32 +120,54 @@
               wf_Memrel trans_Memrel relation_Memrel nat_case_closed)
 
 
+constdefs
+  is_list_functor :: "[i=>o,i,i,i] => o"
+    "is_list_functor(M,A,X,Z) == 
+        \<exists>n1[M]. \<exists>AX[M]. 
+         number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)"
+
+  is_list_case :: "[i=>o,i,i,i,i] => o"
+    "is_list_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)
+
+
 locale M_datatypes = M_wfrank +
-(*THEY NEED RELATIVIZATION*)
   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) &
-		     is_recfun (memr, x,
-				\<lambda>n f. nat_case(0, \<lambda>m. {0} + A \<times> f`m, n), g) &
-		     y = nat_case(0, \<lambda>m. {0} + A \<times> g`m, x))"
+       "[|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, memr, x,
+	              \<lambda>n f z. z = nat_case(0, \<lambda>m. {0} + A * f`m, n), g) &
+		 is_nat_case(M, 0, 
+                      \<lambda>m u. is_list_functor(M,A,g`m,u), x, y))"
+(*THEY NEED RELATIVIZATION*)
       and list_replacement2: 
-           "M(A) ==> strong_replacement(M, \<lambda>x y. y = (\<lambda>X. {0} + A \<times> X)^x (0))"
+           "M(A) ==> strong_replacement(M, \<lambda>x y. y = (\<lambda>X. {0} + A * X)^x (0))"
+
 
 
 lemma (in M_datatypes) list_replacement1':
   "[|M(A); n \<in> nat|]
    ==> strong_replacement
-	  (M, \<lambda>x y. \<exists>z[M]. y = \<langle>x,z\<rangle> &
+	  (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 \<times> f`m, n), g) &
- 	       z = nat_case(0, \<lambda>m. {0} + A \<times> g ` m, x)))"
-by (insert list_replacement1, simp add: nat_into_M) 
+		          \<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 apply_abs
+                 is_recfun_abs [of "\<lambda>n f. nat_case(0, \<lambda>m. {0} + A * f`m, n)"])
+done
 
 lemma (in M_datatypes) list_replacement2': 
-  "M(A) ==> strong_replacement(M, \<lambda>x y. y = (\<lambda>X. {0} + A \<times> X)^x (0))"
+  "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)