diff -r 95f1e700b712 -r 57bf0cecb366 src/ZF/Constructible/WFrec.thy --- a/src/ZF/Constructible/WFrec.thy Tue Mar 06 16:46:27 2012 +0000 +++ b/src/ZF/Constructible/WFrec.thy Tue Mar 06 17:01:37 2012 +0000 @@ -20,7 +20,7 @@ text{*Expresses @{text is_recfun} as a recursion equation*} lemma is_recfun_iff_equation: - "is_recfun(r,a,H,f) <-> + "is_recfun(r,a,H,f) \ f \ r -`` {a} \ range(f) & (\x \ r-``{a}. f`x = H(x, restrict(f, r-``{x})))" apply (rule iffI) @@ -47,7 +47,7 @@ lemma is_recfun_cong_lemma: "[| is_recfun(r,a,H,f); r = r'; a = a'; f = f'; - !!x g. [| \ r'; relation(g); domain(g) <= r' -``{x} |] + !!x g. [| \ r'; relation(g); domain(g) \ r' -``{x} |] ==> H(x,g) = H'(x,g) |] ==> is_recfun(r',a',H',f')" apply (simp add: is_recfun_def) @@ -60,9 +60,9 @@ whose domains are initial segments of @{term r}.*} lemma is_recfun_cong: "[| r = r'; a = a'; f = f'; - !!x g. [| \ r'; relation(g); domain(g) <= r' -``{x} |] + !!x g. [| \ r'; relation(g); domain(g) \ r' -``{x} |] ==> H(x,g) = H'(x,g) |] - ==> is_recfun(r,a,H,f) <-> is_recfun(r',a',H',f')" + ==> is_recfun(r,a,H,f) \ is_recfun(r',a',H',f')" apply (rule iffI) txt{*Messy: fast and blast don't work for some reason*} apply (erule is_recfun_cong_lemma, auto) @@ -91,7 +91,7 @@ "[|is_recfun(r,a,H,f); is_recfun(r,b,H,g); wellfounded(M,r); trans(r); M(f); M(g); M(r); M(x); M(a); M(b) |] - ==> \ r --> \ r --> f`x=g`x" + ==> \ r \ \ r \ f`x=g`x" apply (frule_tac f=f in is_recfun_type) apply (frule_tac f=g in is_recfun_type) apply (simp add: is_recfun_def) @@ -104,7 +104,7 @@ apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def) apply (rename_tac x1) apply (rule_tac t="%z. H(x1,z)" in subst_context) -apply (subgoal_tac "\y \ r-``{x1}. ALL z. \f <-> \g") +apply (subgoal_tac "\y \ r-``{x1}. \z. \f \ \g") apply (blast intro: transD) apply (simp add: apply_iff) apply (blast intro: transD sym) @@ -132,9 +132,9 @@ text{*Tells us that @{text is_recfun} can (in principle) be relativized.*} lemma (in M_basic) is_recfun_relativize: - "[| M(r); M(f); \x[M]. \g[M]. function(g) --> M(H(x,g)) |] - ==> is_recfun(r,a,H,f) <-> - (\z[M]. z \ f <-> + "[| 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 = ))"; apply (simp add: is_recfun_def lam_def) apply (safe intro!: equalityI) @@ -155,7 +155,7 @@ lemma (in M_basic) is_recfun_restrict: "[| wellfounded(M,r); trans(r); is_recfun(r,x,H,f); \y,x\ \ r; M(r); M(f); - \x[M]. \g[M]. function(g) --> M(H(x,g)) |] + \x[M]. \g[M]. function(g) \ M(H(x,g)) |] ==> is_recfun(r, y, H, restrict(f, r -`` {y}))" apply (frule pair_components_in_M, assumption, clarify) apply (simp (no_asm_simp) add: is_recfun_relativize restrict_iff @@ -171,14 +171,14 @@ lemma (in M_basic) restrict_Y_lemma: "[| wellfounded(M,r); trans(r); M(r); - \x[M]. \g[M]. function(g) --> M(H(x,g)); M(Y); + \x[M]. \g[M]. function(g) \ M(H(x,g)); M(Y); \b[M]. - b \ Y <-> + b \ Y \ (\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") +apply (subgoal_tac "\y \ r-``{x}. \z. :Y \ :f") apply (simp (no_asm_simp) add: restrict_def) apply (thin_tac "rall(M,?P)")+ --{*essential for efficiency*} apply (frule is_recfun_type [THEN fun_is_rel], blast) @@ -210,11 +210,11 @@ text{*Proof of the inductive step for @{text exists_is_recfun}, since we must prove two versions.*} lemma (in M_basic) exists_is_recfun_indstep: - "[|\y. \y, a1\ \ r --> (\f[M]. is_recfun(r, y, H, f)); + "[|\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)); - \x[M]. \g[M]. function(g) --> M(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) apply blast @@ -248,7 +248,7 @@ strong_replacement(M, \x z. \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)) |] + \x[M]. \g[M]. function(g) \ M(H(x,g)) |] ==> \f[M]. is_recfun(r,a,H,f)" apply (rule wellfounded_induct, assumption+, clarify) apply (rule exists_is_recfun_indstep, assumption+) @@ -258,8 +258,8 @@ "[|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)); - \x[M]. \g[M]. function(g) --> M(H(x,g)) |] - ==> M(a) --> (\f[M]. is_recfun(r,a,H,f))" + \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+) apply (frule wf_imp_relativized) apply (intro impI) @@ -273,7 +273,7 @@ definition M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where "M_is_recfun(M,MH,r,a,f) == - \z[M]. z \ 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) & @@ -291,9 +291,9 @@ \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); + "[| \x[M]. \g[M]. function(g) \ M(H(x,g)); M(r); M(a); M(f); relation2(M,MH,H) |] - ==> M_is_recfun(M,MH,r,a,f) <-> is_recfun(r,a,H,f)" + ==> M_is_recfun(M,MH,r,a,f) \ is_recfun(r,a,H,f)" apply (simp add: M_is_recfun_def relation2_def is_recfun_relativize) apply (rule rall_cong) apply (blast dest: transM) @@ -301,30 +301,30 @@ lemma M_is_recfun_cong [cong]: "[| r = r'; a = a'; f = f'; - !!x g y. [| M(x); M(g); M(y) |] ==> MH(x,g,y) <-> MH'(x,g,y) |] - ==> M_is_recfun(M,MH,r,a,f) <-> M_is_recfun(M,MH',r',a',f')" + !!x g y. [| M(x); M(g); M(y) |] ==> MH(x,g,y) \ MH'(x,g,y) |] + ==> 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_basic) is_wfrec_abs: - "[| \x[M]. \g[M]. function(g) --> M(H(x,g)); + "[| \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) <-> + ==> is_wfrec(M,MH,r,a,z) \ (\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*} lemma (in M_basic) wfrec_replacement': "[|wfrec_replacement(M,MH,r); - \x[M]. \g[M]. function(g) --> M(H(x,g)); + \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)))" by (simp add: wfrec_replacement_def is_wfrec_abs) lemma wfrec_replacement_cong [cong]: - "[| !!x y z. [| M(x); M(y); M(z) |] ==> MH(x,y,z) <-> MH'(x,y,z); + "[| !!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) \ wfrec_replacement(M, %x y. MH'(x,y), r')" by (simp add: is_wfrec_def wfrec_replacement_def)