diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/WFrec.thy --- a/src/ZF/Constructible/WFrec.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/WFrec.thy Tue Sep 27 17:03:23 2022 +0100 @@ -21,7 +21,7 @@ text\Expresses \is_recfun\ as a recursion equation\ lemma is_recfun_iff_equation: "is_recfun(r,a,H,f) \ - f \ r -`` {a} \ range(f) & + f \ r -`` {a} \ range(f) \ (\x \ r-``{a}. f`x = H(x, restrict(f, r-``{x})))" apply (rule iffI) apply (simp add: is_recfun_type apply_recfun Ball_def vimage_singleton_iff, @@ -135,7 +135,7 @@ "\M(r); M(f); \x[M]. \g[M]. function(g) \ M(H(x,g))\ \ is_recfun(r,a,H,f) \ (\z[M]. z \ f \ - (\x[M]. \ r & z = ))" + (\x[M]. \ r \ z = ))" apply (simp add: is_recfun_def lam_def) apply (safe intro!: equalityI) apply (drule equalityD1 [THEN subsetD], assumption) @@ -174,8 +174,8 @@ \x[M]. \g[M]. function(g) \ M(H(x,g)); M(Y); \b[M]. b \ Y \ - (\x[M]. \ r & - (\y[M]. b = \x,y\ & (\g[M]. is_recfun(r,x,H,g) \ y = H(x,g)))); + (\x[M]. \ r \ + (\y[M]. b = \x,y\ \ (\g[M]. is_recfun(r,x,H,g) \ y = H(x,g)))); \x,a1\ \ r; is_recfun(r,x,H,f); M(f)\ \ restrict(Y, r -`` {x}) = f" apply (subgoal_tac "\y \ r-``{x}. \z. :Y \ :f") @@ -201,7 +201,7 @@ lemma (in M_basic) univalent_is_recfun: "\wellfounded(M,r); trans(r); M(r)\ \ univalent (M, A, \x p. - \y[M]. p = \x,y\ & (\f[M]. is_recfun(r,x,H,f) & y = H(x,f)))" + \y[M]. p = \x,y\ \ (\f[M]. is_recfun(r,x,H,f) \ y = H(x,f)))" apply (simp add: univalent_def) apply (blast dest: is_recfun_functional) done @@ -213,7 +213,7 @@ "\\y. \y, a1\ \ r \ (\f[M]. is_recfun(r, y, H, f)); wellfounded(M,r); trans(r); M(r); M(a1); strong_replacement(M, \x z. - \y[M]. \g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); + \y[M]. \g[M]. pair(M,x,y,z) \ is_recfun(r,x,H,g) \ y = H(x,g)); \x[M]. \g[M]. function(g) \ M(H(x,g))\ \ \f[M]. is_recfun(r,a1,H,f)" apply (drule_tac A="r-``{a1}" in strong_replacementD) @@ -246,7 +246,7 @@ "\wellfounded(M,r); trans(r); separation(M, \x. \ (\f[M]. is_recfun(r, x, H, f))); strong_replacement(M, \x z. - \y[M]. \g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); + \y[M]. \g[M]. pair(M,x,y,z) \ is_recfun(r,x,H,g) \ y = H(x,g)); M(r); M(a); \x[M]. \g[M]. function(g) \ M(H(x,g))\ \ \f[M]. is_recfun(r,a,H,f)" @@ -257,7 +257,7 @@ lemma (in M_basic) wf_exists_is_recfun [rule_format]: "\wf(r); trans(r); M(r); strong_replacement(M, \x z. - \y[M]. \g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); + \y[M]. \g[M]. pair(M,x,y,z) \ is_recfun(r,x,H,g) \ y = H(x,g)); \x[M]. \g[M]. function(g) \ M(H(x,g))\ \ M(a) \ (\f[M]. is_recfun(r,a,H,f))" apply (rule wf_induct, assumption+) @@ -275,20 +275,20 @@ "M_is_recfun(M,MH,r,a,f) \ \z[M]. z \ f \ (\x[M]. \y[M]. \xa[M]. \sx[M]. \r_sx[M]. \f_r_sx[M]. - pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) & - pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) & - xa \ r & MH(x, f_r_sx, y))" + pair(M,x,y,z) \ pair(M,x,a,xa) \ upair(M,x,x,sx) \ + pre_image(M,r,sx,r_sx) \ restriction(M,f,r_sx,f_r_sx) \ + xa \ r \ MH(x, f_r_sx, y))" definition is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where "is_wfrec(M,MH,r,a,z) \ - \f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)" + \f[M]. M_is_recfun(M,MH,r,a,f) \ MH(a,f,z)" definition wfrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where "wfrec_replacement(M,MH,r) \ strong_replacement(M, - \x z. \y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))" + \x z. \y[M]. pair(M,x,y,z) \ is_wfrec(M,MH,r,x,y))" lemma (in M_basic) is_recfun_abs: "\\x[M]. \g[M]. function(g) \ M(H(x,g)); M(r); M(a); M(f); @@ -309,7 +309,7 @@ "\\x[M]. \g[M]. function(g) \ M(H(x,g)); relation2(M,MH,H); M(r); M(a); M(z)\ \ is_wfrec(M,MH,r,a,z) \ - (\g[M]. is_recfun(r,a,H,g) & z = H(a,g))" + (\g[M]. is_recfun(r,a,H,g) \ z = H(a,g))" by (simp add: is_wfrec_def relation2_def is_recfun_abs) text\Relating \<^term>\wfrec_replacement\ to native constructs\ @@ -318,7 +318,7 @@ \x[M]. \g[M]. function(g) \ M(H(x,g)); relation2(M,MH,H); M(r)\ \ strong_replacement(M, \x z. \y[M]. - pair(M,x,y,z) & (\g[M]. is_recfun(r,x,H,g) & y = H(x,g)))" + pair(M,x,y,z) \ (\g[M]. is_recfun(r,x,H,g) \ y = H(x,g)))" by (simp add: wfrec_replacement_def is_wfrec_abs) lemma wfrec_replacement_cong [cong]: