# HG changeset patch # User paulson # Date 1024912643 -7200 # Node ID f96bd927dd37ef5f320872e302214c00d73e395d # Parent 0ffc4403f81178e92394cf54643bd273f5e3b820 towards absoluteness of wf diff -r 0ffc4403f811 -r f96bd927dd37 src/ZF/Constructible/WF_absolute.thy --- a/src/ZF/Constructible/WF_absolute.thy Mon Jun 24 11:56:27 2002 +0200 +++ b/src/ZF/Constructible/WF_absolute.thy Mon Jun 24 11:57:23 2002 +0200 @@ -1,19 +1,12 @@ -theory WF_absolute = WF_extras + WFrec: - +theory WF_absolute = WFrec: subsection{*Transitive closure without fixedpoints*} -(*Ordinal.thy: just after succ_le_iff?*) -lemma Ord_succ_mem_iff: "Ord(j) ==> succ(i) \ succ(j) <-> i\j" -apply (insert succ_le_iff [of i j]) -apply (simp add: lt_def) -done - constdefs rtrancl_alt :: "[i,i]=>i" "rtrancl_alt(A,r) == {p \ A*A. \n\nat. \f \ succ(n) -> A. - \x y. p = & f`0 = x & f`n = y & + (\x y. p = & f`0 = x & f`n = y) & (\i\n. \ r)}" lemma alt_rtrancl_lemma1 [rule_format]: @@ -37,8 +30,7 @@ lemma rtrancl_subset_rtrancl_alt: "r^* <= rtrancl_alt(field(r),r)" apply (simp add: rtrancl_alt_def, clarify) -apply (frule rtrancl_type [THEN subsetD], clarify) -apply simp +apply (frule rtrancl_type [THEN subsetD], clarify, simp) apply (erule rtrancl_induct) txt{*Base case, trivial*} apply (rule_tac x=0 in bexI) @@ -60,151 +52,314 @@ intro: rtrancl_alt_subset_rtrancl rtrancl_subset_rtrancl_alt) +constdefs + + rtran_closure :: "[i=>o,i,i] => o" + "rtran_closure(M,r,s) == + \A. M(A) --> is_field(M,r,A) --> + (\p. M(p) --> + (p \ s <-> + (\n\nat. M(n) & + (\n'. M(n') & successor(M,n,n') & + (\f. M(f) & typed_function(M,n',A,f) & + (\x\A. M(x) & (\y\A. M(y) & pair(M,x,y,p) & + fun_apply(M,f,0,x) & fun_apply(M,f,n,y))) & + (\i\n. M(i) --> + (\i'. M(i') --> successor(M,i,i') --> + (\fi. M(fi) --> fun_apply(M,f,i,fi) --> + (\fi'. M(fi') --> fun_apply(M,f,i',fi') --> + (\q. M(q) --> pair(M,fi,fi',q) --> q \ r))))))))))" + + tran_closure :: "[i=>o,i,i] => o" + "tran_closure(M,r,t) == + \s. M(s) & rtran_closure(M,r,s) & composition(M,r,s,t)" + + +locale M_trancl = M_axioms + +(*THEY NEED RELATIVIZATION*) + assumes rtrancl_separation: + "[| M(r); M(A) |] ==> + separation + (M, \p. \n\nat. \f \ succ(n) -> A. + (\x y. p = & f`0 = x & f`n = y) & + (\i\n. \ r))" + and wellfounded_trancl_separation: + "[| M(r); M(Z) |] ==> separation (M, \x. \w. M(w) & \w,x\ \ r^+ & w \ Z)" + + +lemma (in M_trancl) rtran_closure_rtrancl: + "M(r) ==> rtran_closure(M,r,rtrancl(r))" +apply (simp add: rtran_closure_def rtrancl_alt_eq_rtrancl [symmetric] + rtrancl_alt_def field_closed typed_apply_abs apply_closed + Ord_succ_mem_iff M_nat nat_0_le [THEN ltD], clarify) +apply (rule iffI) + apply clarify + apply simp + apply (rename_tac n f) + apply (rule_tac x=n in bexI) + apply (rule_tac x=f in exI) + apply simp + apply (blast dest: finite_fun_closed dest: transM) + apply assumption +apply clarify +apply (simp add: nat_0_le [THEN ltD] apply_funtype, blast) +done + +lemma (in M_trancl) rtrancl_closed [intro,simp]: + "M(r) ==> M(rtrancl(r))" +apply (insert rtrancl_separation [of r "field(r)"]) +apply (simp add: rtrancl_alt_eq_rtrancl [symmetric] + rtrancl_alt_def field_closed typed_apply_abs apply_closed + Ord_succ_mem_iff M_nat + nat_0_le [THEN ltD] leI [THEN ltD] ltI apply_funtype) +done + +lemma (in M_trancl) rtrancl_abs [simp]: + "[| M(r); M(z) |] ==> rtran_closure(M,r,z) <-> z = rtrancl(r)" +apply (rule iffI) + txt{*Proving the right-to-left implication*} + prefer 2 apply (blast intro: rtran_closure_rtrancl) +apply (rule M_equalityI) +apply (simp add: rtran_closure_def rtrancl_alt_eq_rtrancl [symmetric] + rtrancl_alt_def field_closed typed_apply_abs apply_closed + Ord_succ_mem_iff M_nat + nat_0_le [THEN ltD] leI [THEN ltD] ltI apply_funtype) + prefer 2 apply assumption + prefer 2 apply blast +apply (rule iffI, clarify) +apply (simp add: nat_0_le [THEN ltD] apply_funtype, blast, clarify) +apply simp + apply (rename_tac n f) + apply (rule_tac x=n in bexI) + apply (rule_tac x=f in exI) + apply (blast dest!: finite_fun_closed, assumption) +done + + +lemma (in M_trancl) trancl_closed [intro,simp]: + "M(r) ==> M(trancl(r))" +by (simp add: trancl_def comp_closed rtrancl_closed) + +lemma (in M_trancl) trancl_abs [simp]: + "[| M(r); M(z) |] ==> tran_closure(M,r,z) <-> z = trancl(r)" +by (simp add: tran_closure_def trancl_def) + + +text{*Alternative proof of @{text wf_on_trancl}; inspiration for the + relativized version. Original version is on theory WF.*} +lemma "[| wf[A](r); r-``A <= A |] ==> wf[A](r^+)" +apply (simp add: wf_on_def wf_def) +apply (safe intro!: equalityI) +apply (drule_tac x = "{x\A. \w. \w,x\ \ r^+ & w \ Z}" in spec) +apply (blast elim: tranclE) +done + + +lemma (in M_trancl) wellfounded_on_trancl: + "[| wellfounded_on(M,A,r); r-``A <= A; M(r); M(A) |] + ==> wellfounded_on(M,A,r^+)" +apply (simp add: wellfounded_on_def) +apply (safe intro!: equalityI) +apply (rename_tac Z x) +apply (subgoal_tac "M({x\A. \w. M(w) & \w,x\ \ r^+ & w \ Z})") + prefer 2 + apply (simp add: wellfounded_trancl_separation) +apply (drule_tac x = "{x\A. \w. M(w) & \w,x\ \ r^+ & w \ Z}" in spec) +apply safe +apply (blast dest: transM, simp) +apply (rename_tac y w) +apply (drule_tac x=w in bspec, assumption, clarify) +apply (erule tranclE) + apply (blast dest: transM) (*transM is needed to prove M(xa)*) + apply blast +done + + text{*Relativized to M: Every well-founded relation is a subset of some inverse image of an ordinal. Key step is the construction (in M) of a rank function.*} (*NEEDS RELATIVIZATION*) -locale M_recursion = M_axioms + +locale M_recursion = M_trancl + assumes wfrank_separation': - "[| M(r); M(a); r \ A*A |] ==> + "[| M(r); M(A) |] ==> separation (M, \x. x \ A --> - ~(\f. M(f) \ - is_recfun(r, x, %x f. \y \ r-``{x}. succ(f`y), f)))" + ~(\f. M(f) \ is_recfun(r^+, x, %x f. range(f), f)))" and wfrank_strong_replacement': - "[| M(r); M(a); r \ A*A |] ==> - strong_replacement(M, \x z. \y g. M(y) & M(g) & - pair(M,x,y,z) & - is_recfun(r, x, %x f. \y \ r-``{x}. succ(f`y), f) & - y = (\y \ r-``{x}. succ(g`y)))" + "M(r) ==> + strong_replacement(M, \x z. \y f. M(y) & M(f) & + pair(M,x,y,z) & is_recfun(r^+, x, %x f. range(f), f) & + y = range(f))" + and Ord_wfrank_separation: + "[| M(r); M(A) |] ==> + separation (M, \x. x \ A \ + \ (\f. M(f) \ + is_recfun(r^+, x, \x. range, f) \ Ord(range(f))))" - -constdefs (*????????????????NEEDED?*) - is_wfrank_fun :: "[i=>o,i,i,i] => o" - "is_wfrank_fun(M,r,a,f) == - function(f) & domain(f) = r-``{a} & - (\x. M(x) --> \ r --> f`x = (\y \ r-``{x}. succ(f`y)))" - - - +constdefs + wellfoundedrank :: "[i=>o,i,i] => i" + "wellfoundedrank(M,r,A) == + {p. x\A, \y f. M(y) & M(f) & + p = & is_recfun(r^+, x, %x f. range(f), f) & + y = range(f)}" lemma (in M_recursion) exists_wfrank: "[| wellfounded(M,r); r \ A*A; a\A; M(r); M(A) |] - ==> \f. M(f) & is_recfun(r, a, %x g. (\y \ r-``{x}. succ(g`y)), f)" -apply (rule exists_is_recfun [of A r]) -apply (erule wellfounded_imp_wellfounded_on) -apply assumption; -apply (rule trans_Memrel [THEN trans_imp_trans_on], simp) -apply (rule succI1) -apply (blast intro: wfrank_separation') -apply (blast intro: wfrank_strong_replacement') -apply (simp_all add: Memrel_type Memrel_closed Un_closed image_closed) + ==> \f. M(f) & is_recfun(r^+, a, %x f. range(f), f)" +apply (rule wellfounded_exists_is_recfun [of A]) +apply (blast intro: wellfounded_on_trancl wellfounded_imp_wellfounded_on) +apply (rule trans_trancl [THEN trans_imp_trans_on], assumption+) +apply (simp_all add: trancl_subset_times) done -lemma (in M_recursion) exists_wfrank_fun: - "[| Ord(j); M(i); M(j) |] ==> \f. M(f) & is_wfrank_fun(M,i,succ(j),f)" -apply (rule exists_wfrank [THEN exE]) -apply (erule Ord_succ, assumption, simp) -apply (rename_tac f, clarify) -apply (frule is_recfun_type) -apply (rule_tac x=f in exI) -apply (simp add: fun_is_function domain_of_fun lt_Memrel apply_recfun lt_def - is_wfrank_fun_eq Ord_trans [OF _ succI1]) +lemma (in M_recursion) M_wellfoundedrank: + "[| wellfounded(M,r); r \ A*A; M(r); M(A) |] + ==> M(wellfoundedrank(M,r,A))" +apply (insert wfrank_strong_replacement' [of r]) +apply (simp add: wellfoundedrank_def) +apply (rule strong_replacement_closed) + apply assumption+ + apply (rule univalent_is_recfun) + apply (blast intro: wellfounded_on_trancl wellfounded_imp_wellfounded_on) + apply (rule trans_on_trancl) + apply (simp add: trancl_subset_times) + apply blast+ done -lemma (in M_recursion) is_wfrank_fun_apply: - "[| x < j; M(i); M(j); M(f); is_wfrank_fun(M,r,a,f) |] - ==> f`x = i Un (\k\x. {f ` k})" -apply (simp add: is_wfrank_fun_eq lt_Ord2) -apply (frule lt_closed, simp) -apply (subgoal_tac "x <= domain(f)") - apply (simp add: Ord_trans [OF _ succI1] image_function) - apply (blast intro: elim:); -apply (blast intro: dest!: leI [THEN le_imp_subset] ) -done - -lemma (in M_recursion) is_wfrank_fun_eq_wfrank [rule_format]: - "[| is_wfrank_fun(M,i,J,f); M(i); M(J); M(f); Ord(i); Ord(j) |] - ==> j f`j = i++j" -apply (erule_tac i=j in trans_induct, clarify) -apply (subgoal_tac "\k\x. k fun_apply(M,f,j,k) <-> f`j = k" -by (auto simp add: lt_def is_wfrank_fun_eq subsetD apply_abs) - -lemma (in M_recursion) Ord_wfrank_abs: - "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_wfrank(M,r,a,k) <-> k = i++j" -apply (simp add: is_wfrank_def wfrank_abs_fun_apply_iff is_wfrank_fun_eq_wfrank) -apply (frule exists_wfrank_fun [of j i], blast+) +lemma (in M_recursion) Ord_wfrank_range [rule_format]: + "[| wellfounded(M,r); r \ A*A; a\A; M(r); M(A) |] + ==> \f. M(f) --> is_recfun(r^+, a, %x f. range(f), f) --> Ord(range(f))" +apply (subgoal_tac "wellfounded_on(M, A, r^+)") + prefer 2 + apply (blast intro: wellfounded_on_trancl wellfounded_imp_wellfounded_on) +apply (erule wellfounded_on_induct2, assumption+) +apply (simp add: trancl_subset_times) +apply (blast intro: Ord_wfrank_separation, clarify) +txt{*The reasoning in both cases is that we get @{term y} such that + @{term "\y, x\ \ r^+"}. We find that + @{term "f`y = restrict(f, r^+ -`` {y})"}. *} +apply (rule OrdI [OF _ Ord_is_Transset]) + txt{*An ordinal is a transitive set...*} + apply (simp add: Transset_def) + apply clarify + apply (frule apply_recfun2, assumption) + apply (force simp add: restrict_iff) +txt{*...of ordinals. This second case requires the induction hyp.*} +apply clarify +apply (rename_tac i y) +apply (frule apply_recfun2, assumption) +apply (frule is_recfun_imp_in_r, assumption) +apply (frule is_recfun_restrict) + (*simp_all won't work*) + apply (simp add: trans_on_trancl trancl_subset_times)+ +apply (drule spec [THEN mp], assumption) +apply (subgoal_tac "M(restrict(f, r^+ -`` {y}))") + apply (drule_tac x="restrict(f, r^+ -`` {y})" in spec) + apply (simp add: function_apply_equality [OF _ is_recfun_imp_function]) +apply (blast dest: pair_components_in_M) done -lemma (in M_recursion) wfrank_abs: - "[| M(i); M(j); M(k) |] ==> is_wfrank(M,r,a,k) <-> k = i++j" -apply (case_tac "Ord(i) & Ord(j)") - apply (simp add: Ord_wfrank_abs) -apply (auto simp add: is_wfrank_def wfrank_eq_if_raw_wfrank) +lemma (in M_recursion) Ord_range_wellfoundedrank: + "[| wellfounded(M,r); r \ A*A; M(r); M(A) |] + ==> Ord (range(wellfoundedrank(M,r,A)))" +apply (subgoal_tac "wellfounded_on(M, A, r^+)") + prefer 2 + apply (blast intro: wellfounded_on_trancl wellfounded_imp_wellfounded_on) +apply (frule trancl_subset_times) +apply (simp add: wellfoundedrank_def) +apply (rule OrdI [OF _ Ord_is_Transset]) + prefer 2 + txt{*by our previous result the range consists of ordinals.*} + apply (blast intro: Ord_wfrank_range) +txt{*We still must show that the range is a transitive set.*} +apply (simp add: Transset_def, clarify) +apply simp +apply (rename_tac x i f u) +apply (frule is_recfun_imp_in_r, assumption) +apply (subgoal_tac "M(u) & M(i) & M(x)") + prefer 2 apply (blast dest: transM, clarify) +apply (rule_tac a=u in rangeI) +apply (rule ReplaceI) + apply (rule_tac x=i in exI, simp) + apply (rule_tac x="restrict(f, r^+ -`` {u})" in exI) + apply (blast intro: is_recfun_restrict trans_on_trancl dest: apply_recfun2) + apply blast +txt{*Unicity requirement of Replacement*} +apply clarify +apply (frule apply_recfun2, assumption) +apply (simp add: trans_on_trancl is_recfun_cut)+ done -lemma (in M_recursion) wfrank_closed [intro]: - "[| M(i); M(j) |] ==> M(i++j)" -apply (simp add: wfrank_eq_if_raw_wfrank, clarify) -apply (simp add: raw_wfrank_eq_wfrank) -apply (frule exists_wfrank_fun [of j i], auto) -apply (simp add: apply_closed is_wfrank_fun_eq_wfrank [symmetric]) +lemma (in M_recursion) function_wellfoundedrank: + "[| wellfounded(M,r); r \ A*A; M(r); M(A)|] + ==> function(wellfoundedrank(M,r,A))" +apply (simp add: wellfoundedrank_def function_def, clarify) +txt{*Uniqueness: repeated below!*} +apply (drule is_recfun_functional, assumption) + apply (blast intro: wellfounded_on_trancl wellfounded_imp_wellfounded_on) + apply (simp_all add: trancl_subset_times + trans_trancl [THEN trans_imp_trans_on]) +apply (blast dest: transM) done - - -constdefs - wfrank :: "[i,i]=>i" - "wfrank(r,a) == wfrec(r, a, %x f. \y \ r-``{x}. succ(f`y))" - -constdefs - wftype :: "i=>i" - "wftype(r) == \y \ range(r). succ(wfrank(r,y))" - -lemma (in M_axioms) wfrank: "wellfounded(M,r) ==> wfrank(r,a) = (\y \ r-``{a}. succ(wfrank(r,y)))" -by (subst wfrank_def [THEN def_wfrec], simp_all) - -lemma (in M_axioms) Ord_wfrank: "wellfounded(M,r) ==> Ord(wfrank(r,a))" -apply (rule_tac a="a" in wf_induct, assumption) -apply (subst wfrank, assumption) -apply (rule Ord_succ [THEN Ord_UN], blast) +lemma (in M_recursion) domain_wellfoundedrank: + "[| wellfounded(M,r); r \ A*A; M(r); M(A)|] + ==> domain(wellfoundedrank(M,r,A)) = A" +apply (simp add: wellfoundedrank_def function_def) +apply (rule equalityI, auto) +apply (frule transM, assumption) +apply (frule exists_wfrank, assumption+) +apply clarify +apply (rule domainI) +apply (rule ReplaceI) +apply (rule_tac x="range(f)" in exI) +apply simp +apply (rule_tac x=f in exI, blast) +apply assumption +txt{*Uniqueness (for Replacement): repeated above!*} +apply clarify +apply (drule is_recfun_functional, assumption) + apply (blast intro: wellfounded_on_trancl wellfounded_imp_wellfounded_on) + apply (simp_all add: trancl_subset_times + trans_trancl [THEN trans_imp_trans_on]) done -lemma (in M_axioms) wfrank_lt: "[|wellfounded(M,r); \ r|] ==> wfrank(r,a) < wfrank(r,b)" -apply (rule_tac a1 = "b" in wfrank [THEN ssubst], assumption) -apply (rule UN_I [THEN ltI]) -apply (simp add: Ord_wfrank vimage_iff)+ -done - -lemma (in M_axioms) Ord_wftype: "wellfounded(M,r) ==> Ord(wftype(r))" -by (simp add: wftype_def Ord_wfrank) - -lemma (in M_axioms) wftypeI: "\wellfounded(M,r); x \ field(r)\ \ wfrank(r,x) \ wftype(r)" -apply (simp add: wftype_def) -apply (blast intro: wfrank_lt [THEN ltD]) +lemma (in M_recursion) wellfoundedrank_type: + "[| wellfounded(M,r); r \ A*A; M(r); M(A)|] + ==> wellfoundedrank(M,r,A) \ A -> range(wellfoundedrank(M,r,A))" +apply (frule function_wellfoundedrank, assumption+) +apply (frule function_imp_Pi) + apply (simp add: wellfoundedrank_def relation_def) + apply blast +apply (simp add: domain_wellfoundedrank) done +lemma (in M_recursion) Ord_wellfoundedrank: + "[| wellfounded(M,r); a \ A; r \ A*A; M(r); M(A) |] + ==> Ord(wellfoundedrank(M,r,A) ` a)" +by (blast intro: apply_funtype [OF wellfoundedrank_type] + Ord_in_Ord [OF Ord_range_wellfoundedrank]) -lemma (in M_axioms) wf_imp_subset_rvimage: - "[|wellfounded(M,r); r \ A*A|] ==> \i f. Ord(i) & r <= rvimage(A, f, Memrel(i))" -apply (rule_tac x="wftype(r)" in exI) -apply (rule_tac x="\x\A. wfrank(r,x)" in exI) -apply (simp add: Ord_wftype, clarify) -apply (frule subsetD, assumption, clarify) -apply (simp add: rvimage_iff wfrank_lt [THEN ltD]) -apply (blast intro: wftypeI ) +lemma (in M_recursion) wellfoundedrank_eq: + "[| is_recfun(r^+, a, %x. range, f); + wellfounded(M,r); a \ A; r \ A*A; M(f); M(r); M(A)|] + ==> wellfoundedrank(M,r,A) ` a = range(f)" +apply (rule apply_equality) + prefer 2 apply (blast intro: wellfoundedrank_type ) +apply (simp add: wellfoundedrank_def) +apply (rule ReplaceI) + apply (rule_tac x="range(f)" in exI) + apply blast + apply assumption +txt{*Unicity requirement of Replacement*} +apply clarify +apply (drule is_recfun_functional, assumption) + apply (blast intro: wellfounded_on_trancl wellfounded_imp_wellfounded_on) + apply (simp_all add: trancl_subset_times + trans_trancl [THEN trans_imp_trans_on]) +apply (blast dest: transM) done - - - end