Up to index of Isabelle/ZF/Constructible
theory Rank_Separation(* Title: ZF/Constructible/Rank_Separation.thy ID: $Id: Rank_Separation.thy,v 1.4 2005/06/17 14:15:10 haftmann Exp $ Author: Lawrence C Paulson, Cambridge University Computer Laboratory *) header {*Separation for Facts About Order Types, Rank Functions and Well-Founded Relations*} theory Rank_Separation imports Rank Rec_Separation begin text{*This theory proves all instances needed for locales @{text "M_ordertype"} and @{text "M_wfrank"}. But the material is not needed for proving the relative consistency of AC.*} subsection{*The Locale @{text "M_ordertype"}*} subsubsection{*Separation for Order-Isomorphisms*} lemma well_ord_iso_Reflects: "REFLECTS[λx. x∈A --> (∃y[L]. ∃p[L]. fun_apply(L,f,x,y) & pair(L,y,x,p) & p ∈ r), λi x. x∈A --> (∃y ∈ Lset(i). ∃p ∈ Lset(i). fun_apply(##Lset(i),f,x,y) & pair(##Lset(i),y,x,p) & p ∈ r)]" by (intro FOL_reflections function_reflections) lemma well_ord_iso_separation: "[| L(A); L(f); L(r) |] ==> separation (L, λx. x∈A --> (∃y[L]. (∃p[L]. fun_apply(L,f,x,y) & pair(L,y,x,p) & p ∈ r)))" apply (rule gen_separation_multi [OF well_ord_iso_Reflects, of "{A,f,r}"], auto) apply (rule_tac env="[A,f,r]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsubsection{*Separation for @{term "obase"}*} lemma obase_reflects: "REFLECTS[λa. ∃x[L]. ∃g[L]. ∃mx[L]. ∃par[L]. ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g), λi a. ∃x ∈ Lset(i). ∃g ∈ Lset(i). ∃mx ∈ Lset(i). ∃par ∈ Lset(i). ordinal(##Lset(i),x) & membership(##Lset(i),x,mx) & pred_set(##Lset(i),A,a,r,par) & order_isomorphism(##Lset(i),par,r,x,mx,g)]" by (intro FOL_reflections function_reflections fun_plus_reflections) lemma obase_separation: --{*part of the order type formalization*} "[| L(A); L(r) |] ==> separation(L, λa. ∃x[L]. ∃g[L]. ∃mx[L]. ∃par[L]. ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g))" apply (rule gen_separation_multi [OF obase_reflects, of "{A,r}"], auto) apply (rule_tac env="[A,r]" in DPow_LsetI) apply (rule ordinal_iff_sats sep_rules | simp)+ done subsubsection{*Separation for a Theorem about @{term "obase"}*} lemma obase_equals_reflects: "REFLECTS[λx. x∈A --> ~(∃y[L]. ∃g[L]. ordinal(L,y) & (∃my[L]. ∃pxr[L]. membership(L,y,my) & pred_set(L,A,x,r,pxr) & order_isomorphism(L,pxr,r,y,my,g))), λi x. x∈A --> ~(∃y ∈ Lset(i). ∃g ∈ Lset(i). ordinal(##Lset(i),y) & (∃my ∈ Lset(i). ∃pxr ∈ Lset(i). membership(##Lset(i),y,my) & pred_set(##Lset(i),A,x,r,pxr) & order_isomorphism(##Lset(i),pxr,r,y,my,g)))]" by (intro FOL_reflections function_reflections fun_plus_reflections) lemma obase_equals_separation: "[| L(A); L(r) |] ==> separation (L, λx. x∈A --> ~(∃y[L]. ∃g[L]. ordinal(L,y) & (∃my[L]. ∃pxr[L]. membership(L,y,my) & pred_set(L,A,x,r,pxr) & order_isomorphism(L,pxr,r,y,my,g))))" apply (rule gen_separation_multi [OF obase_equals_reflects, of "{A,r}"], auto) apply (rule_tac env="[A,r]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsubsection{*Replacement for @{term "omap"}*} lemma omap_reflects: "REFLECTS[λz. ∃a[L]. a∈B & (∃x[L]. ∃g[L]. ∃mx[L]. ∃par[L]. ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) & pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g)), λi z. ∃a ∈ Lset(i). a∈B & (∃x ∈ Lset(i). ∃g ∈ Lset(i). ∃mx ∈ Lset(i). ∃par ∈ Lset(i). ordinal(##Lset(i),x) & pair(##Lset(i),a,x,z) & membership(##Lset(i),x,mx) & pred_set(##Lset(i),A,a,r,par) & order_isomorphism(##Lset(i),par,r,x,mx,g))]" by (intro FOL_reflections function_reflections fun_plus_reflections) lemma omap_replacement: "[| L(A); L(r) |] ==> strong_replacement(L, λa z. ∃x[L]. ∃g[L]. ∃mx[L]. ∃par[L]. ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) & pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g))" apply (rule strong_replacementI) apply (rule_tac u="{A,r,B}" in gen_separation_multi [OF omap_reflects], auto) apply (rule_tac env="[A,B,r]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Instantiating the locale @{text M_ordertype}*} text{*Separation (and Strong Replacement) for basic set-theoretic constructions such as intersection, Cartesian Product and image.*} lemma M_ordertype_axioms_L: "M_ordertype_axioms(L)" apply (rule M_ordertype_axioms.intro) apply (assumption | rule well_ord_iso_separation obase_separation obase_equals_separation omap_replacement)+ done theorem M_ordertype_L: "PROP M_ordertype(L)" apply (rule M_ordertype.intro) apply (rule M_basic.axioms [OF M_basic_L])+ apply (rule M_ordertype_axioms_L) done subsection{*The Locale @{text "M_wfrank"}*} subsubsection{*Separation for @{term "wfrank"}*} lemma wfrank_Reflects: "REFLECTS[λx. ∀rplus[L]. tran_closure(L,r,rplus) --> ~ (∃f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)), λi x. ∀rplus ∈ Lset(i). tran_closure(##Lset(i),r,rplus) --> ~ (∃f ∈ Lset(i). M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y), rplus, x, f))]" by (intro FOL_reflections function_reflections is_recfun_reflection tran_closure_reflection) lemma wfrank_separation: "L(r) ==> separation (L, λx. ∀rplus[L]. tran_closure(L,r,rplus) --> ~ (∃f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))" apply (rule gen_separation [OF wfrank_Reflects], simp) apply (rule_tac env="[r]" in DPow_LsetI) apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+ done subsubsection{*Replacement for @{term "wfrank"}*} lemma wfrank_replacement_Reflects: "REFLECTS[λz. ∃x[L]. x ∈ A & (∀rplus[L]. tran_closure(L,r,rplus) --> (∃y[L]. ∃f[L]. pair(L,x,y,z) & M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) & is_range(L,f,y))), λi z. ∃x ∈ Lset(i). x ∈ A & (∀rplus ∈ Lset(i). tran_closure(##Lset(i),r,rplus) --> (∃y ∈ Lset(i). ∃f ∈ Lset(i). pair(##Lset(i),x,y,z) & M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y), rplus, x, f) & is_range(##Lset(i),f,y)))]" by (intro FOL_reflections function_reflections fun_plus_reflections is_recfun_reflection tran_closure_reflection) lemma wfrank_strong_replacement: "L(r) ==> strong_replacement(L, λx z. ∀rplus[L]. tran_closure(L,r,rplus) --> (∃y[L]. ∃f[L]. pair(L,x,y,z) & M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) & is_range(L,f,y)))" apply (rule strong_replacementI) apply (rule_tac u="{r,B}" in gen_separation_multi [OF wfrank_replacement_Reflects], auto) apply (rule_tac env="[r,B]" in DPow_LsetI) apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+ done subsubsection{*Separation for Proving @{text Ord_wfrank_range}*} lemma Ord_wfrank_Reflects: "REFLECTS[λx. ∀rplus[L]. tran_closure(L,r,rplus) --> ~ (∀f[L]. ∀rangef[L]. is_range(L,f,rangef) --> M_is_recfun(L, λx f y. is_range(L,f,y), rplus, x, f) --> ordinal(L,rangef)), λi x. ∀rplus ∈ Lset(i). tran_closure(##Lset(i),r,rplus) --> ~ (∀f ∈ Lset(i). ∀rangef ∈ Lset(i). is_range(##Lset(i),f,rangef) --> M_is_recfun(##Lset(i), λx f y. is_range(##Lset(i),f,y), rplus, x, f) --> ordinal(##Lset(i),rangef))]" by (intro FOL_reflections function_reflections is_recfun_reflection tran_closure_reflection ordinal_reflection) lemma Ord_wfrank_separation: "L(r) ==> separation (L, λx. ∀rplus[L]. tran_closure(L,r,rplus) --> ~ (∀f[L]. ∀rangef[L]. is_range(L,f,rangef) --> M_is_recfun(L, λx f y. is_range(L,f,y), rplus, x, f) --> ordinal(L,rangef)))" apply (rule gen_separation [OF Ord_wfrank_Reflects], simp) apply (rule_tac env="[r]" in DPow_LsetI) apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+ done subsubsection{*Instantiating the locale @{text M_wfrank}*} lemma M_wfrank_axioms_L: "M_wfrank_axioms(L)" apply (rule M_wfrank_axioms.intro) apply (assumption | rule wfrank_separation wfrank_strong_replacement Ord_wfrank_separation)+ done theorem M_wfrank_L: "PROP M_wfrank(L)" apply (rule M_wfrank.intro) apply (rule M_trancl.axioms [OF M_trancl_L])+ apply (rule M_wfrank_axioms_L) done lemmas exists_wfrank = M_wfrank.exists_wfrank [OF M_wfrank_L] and M_wellfoundedrank = M_wfrank.M_wellfoundedrank [OF M_wfrank_L] and Ord_wfrank_range = M_wfrank.Ord_wfrank_range [OF M_wfrank_L] and Ord_range_wellfoundedrank = M_wfrank.Ord_range_wellfoundedrank [OF M_wfrank_L] and function_wellfoundedrank = M_wfrank.function_wellfoundedrank [OF M_wfrank_L] and domain_wellfoundedrank = M_wfrank.domain_wellfoundedrank [OF M_wfrank_L] and wellfoundedrank_type = M_wfrank.wellfoundedrank_type [OF M_wfrank_L] and Ord_wellfoundedrank = M_wfrank.Ord_wellfoundedrank [OF M_wfrank_L] and wellfoundedrank_eq = M_wfrank.wellfoundedrank_eq [OF M_wfrank_L] and wellfoundedrank_lt = M_wfrank.wellfoundedrank_lt [OF M_wfrank_L] and wellfounded_imp_subset_rvimage = M_wfrank.wellfounded_imp_subset_rvimage [OF M_wfrank_L] and wellfounded_imp_wf = M_wfrank.wellfounded_imp_wf [OF M_wfrank_L] and wellfounded_on_imp_wf_on = M_wfrank.wellfounded_on_imp_wf_on [OF M_wfrank_L] and wf_abs = M_wfrank.wf_abs [OF M_wfrank_L] and wf_on_abs = M_wfrank.wf_on_abs [OF M_wfrank_L] end
lemma well_ord_iso_Reflects:
REFLECTS [%x. x ∈ A --> (∃y[L]. ∃p[L]. fun_apply(L, f, x, y) ∧ pair(L, y, x, p) ∧ p ∈ r), %i x. x ∈ A --> (∃y∈Lset(i). ∃p∈Lset(i). fun_apply(##Lset(i), f, x, y) ∧ pair(##Lset(i), y, x, p) ∧ p ∈ r)]
lemma well_ord_iso_separation:
[| L(A); L(f); L(r) |] ==> separation (L, %x. x ∈ A --> (∃y[L]. ∃p[L]. fun_apply(L, f, x, y) ∧ pair(L, y, x, p) ∧ p ∈ r))
lemma obase_reflects:
REFLECTS [%a. ∃x[L]. ∃g[L]. ∃mx[L]. ∃par[L]. ordinal(L, x) ∧ membership(L, x, mx) ∧ pred_set(L, A, a, r, par) ∧ order_isomorphism(L, par, r, x, mx, g), %i a. ∃x∈Lset(i). ∃g∈Lset(i). ∃mx∈Lset(i). ∃par∈Lset(i). ordinal(##Lset(i), x) ∧ membership(##Lset(i), x, mx) ∧ pred_set(##Lset(i), A, a, r, par) ∧ order_isomorphism(##Lset(i), par, r, x, mx, g)]
lemma obase_separation:
[| L(A); L(r) |] ==> separation (L, %a. ∃x[L]. ∃g[L]. ∃mx[L]. ∃par[L]. ordinal(L, x) ∧ membership(L, x, mx) ∧ pred_set(L, A, a, r, par) ∧ order_isomorphism(L, par, r, x, mx, g))
lemma obase_equals_reflects:
REFLECTS [%x. x ∈ A --> ¬ (∃y[L]. ∃g[L]. ordinal(L, y) ∧ (∃my[L]. ∃pxr[L]. membership(L, y, my) ∧ pred_set(L, A, x, r, pxr) ∧ order_isomorphism(L, pxr, r, y, my, g))), %i x. x ∈ A --> ¬ (∃y∈Lset(i). ∃g∈Lset(i). ordinal(##Lset(i), y) ∧ (∃my∈Lset(i). ∃pxr∈Lset(i). membership(##Lset(i), y, my) ∧ pred_set(##Lset(i), A, x, r, pxr) ∧ order_isomorphism(##Lset(i), pxr, r, y, my, g)))]
lemma obase_equals_separation:
[| L(A); L(r) |] ==> separation (L, %x. x ∈ A --> ¬ (∃y[L]. ∃g[L]. ordinal(L, y) ∧ (∃my[L]. ∃pxr[L]. membership(L, y, my) ∧ pred_set(L, A, x, r, pxr) ∧ order_isomorphism(L, pxr, r, y, my, g))))
lemma omap_reflects:
REFLECTS [%z. ∃a[L]. a ∈ B ∧ (∃x[L]. ∃g[L]. ∃mx[L]. ∃par[L]. ordinal(L, x) ∧ pair(L, a, x, z) ∧ membership(L, x, mx) ∧ pred_set(L, A, a, r, par) ∧ order_isomorphism(L, par, r, x, mx, g)), %i z. ∃a∈Lset(i). a ∈ B ∧ (∃x∈Lset(i). ∃g∈Lset(i). ∃mx∈Lset(i). ∃par∈Lset(i). ordinal(##Lset(i), x) ∧ pair(##Lset(i), a, x, z) ∧ membership(##Lset(i), x, mx) ∧ pred_set(##Lset(i), A, a, r, par) ∧ order_isomorphism(##Lset(i), par, r, x, mx, g))]
lemma omap_replacement:
[| L(A); L(r) |] ==> strong_replacement (L, %a z. ∃x[L]. ∃g[L]. ∃mx[L]. ∃par[L]. ordinal(L, x) ∧ pair(L, a, x, z) ∧ membership(L, x, mx) ∧ pred_set(L, A, a, r, par) ∧ order_isomorphism(L, par, r, x, mx, g))
lemma M_ordertype_axioms_L:
M_ordertype_axioms(L)
theorem M_ordertype_L:
PROP M_ordertype(L)
lemma wfrank_Reflects:
REFLECTS [%x. ∀rplus[L]. tran_closure(L, r, rplus) --> ¬ (∃f[L]. M_is_recfun(L, %x f y. is_range(L, f, y), rplus, x, f)), %i x. ∀rplus∈Lset(i). tran_closure(##Lset(i), r, rplus) --> ¬ (∃f∈Lset(i). M_is_recfun (##Lset(i), %x f y. is_range(##Lset(i), f, y), rplus, x, f))]
lemma wfrank_separation:
L(r) ==> separation (L, %x. ∀rplus[L]. tran_closure(L, r, rplus) --> ¬ (∃f[L]. M_is_recfun (L, %x f y. is_range(L, f, y), rplus, x, f)))
lemma wfrank_replacement_Reflects:
REFLECTS [%z. ∃x[L]. x ∈ A ∧ (∀rplus[L]. tran_closure(L, r, rplus) --> (∃y[L]. ∃f[L]. pair(L, x, y, z) ∧ M_is_recfun (L, %x f y. is_range(L, f, y), rplus, x, f) ∧ is_range(L, f, y))), %i z. ∃x∈Lset(i). x ∈ A ∧ (∀rplus∈Lset(i). tran_closure(##Lset(i), r, rplus) --> (∃y∈Lset(i). ∃f∈Lset(i). pair(##Lset(i), x, y, z) ∧ M_is_recfun (##Lset(i), %x f y. is_range(##Lset(i), f, y), rplus, x, f) ∧ is_range(##Lset(i), f, y)))]
lemma wfrank_strong_replacement:
L(r) ==> strong_replacement (L, %x z. ∀rplus[L]. tran_closure(L, r, rplus) --> (∃y[L]. ∃f[L]. pair(L, x, y, z) ∧ M_is_recfun (L, %x f y. is_range(L, f, y), rplus, x, f) ∧ is_range(L, f, y)))
lemma Ord_wfrank_Reflects:
REFLECTS [%x. ∀rplus[L]. tran_closure(L, r, rplus) --> ¬ (∀f[L]. ∀rangef[L]. is_range(L, f, rangef) --> M_is_recfun (L, %x f y. is_range(L, f, y), rplus, x, f) --> ordinal(L, rangef)), %i x. ∀rplus∈Lset(i). tran_closure(##Lset(i), r, rplus) --> ¬ (∀f∈Lset(i). ∀rangef∈Lset(i). is_range(##Lset(i), f, rangef) --> M_is_recfun (##Lset(i), %x f y. is_range(##Lset(i), f, y), rplus, x, f) --> ordinal(##Lset(i), rangef))]
lemma Ord_wfrank_separation:
L(r) ==> separation (L, %x. ∀rplus[L]. tran_closure(L, r, rplus) --> ¬ (∀f[L]. ∀rangef[L]. is_range(L, f, rangef) --> M_is_recfun (L, %x f y. is_range(L, f, y), rplus, x, f) --> ordinal(L, rangef)))
lemma M_wfrank_axioms_L:
M_wfrank_axioms(L)
theorem M_wfrank_L:
PROP M_wfrank(L)
lemmas exists_wfrank:
[| wellfounded(L, r); L(a); L(r) |] ==> ∃f[L]. is_recfun(r^+, a, %x f. range(f), f)
and M_wellfoundedrank:
[| wellfounded(L, r); L(r); L(A) |] ==> L(wellfoundedrank(L, r, A))
and Ord_wfrank_range:
[| wellfounded(L, r); a ∈ A; L(r); L(A) |] ==> ∀f[L]. is_recfun(r^+, a, %x f. range(f), f) --> Ord(range(f))
and Ord_range_wellfoundedrank:
[| wellfounded(L, r); r ⊆ A × A; L(r); L(A) |] ==> Ord(range(wellfoundedrank(L, r, A)))
and function_wellfoundedrank:
[| wellfounded(L, r); L(r); L(A) |] ==> function(wellfoundedrank(L, r, A))
and domain_wellfoundedrank:
[| wellfounded(L, r); L(r); L(A) |] ==> domain(wellfoundedrank(L, r, A)) = A
and wellfoundedrank_type:
[| wellfounded(L, r); L(r); L(A) |] ==> wellfoundedrank(L, r, A) ∈ A -> range(wellfoundedrank(L, r, A))
and Ord_wellfoundedrank:
[| wellfounded(L, r); a ∈ A; r ⊆ A × A; L(r); L(A) |] ==> Ord(wellfoundedrank(L, r, A) ` a)
and wellfoundedrank_eq:
[| is_recfun(r^+, a, %x. range, f); wellfounded(L, r); a ∈ A; L(f); L(r); L(A) |] ==> wellfoundedrank(L, r, A) ` a = range(f)
and wellfoundedrank_lt:
[| 〈a, b〉 ∈ r; wellfounded(L, r); r ⊆ A × A; L(r); L(A) |] ==> wellfoundedrank(L, r, A) ` a < wellfoundedrank(L, r, A) ` b
and wellfounded_imp_subset_rvimage:
[| wellfounded(L, r); r ⊆ A × A; L(r); L(A) |] ==> ∃i f. Ord(i) ∧ r ⊆ rvimage(A, f, Memrel(i))
and wellfounded_imp_wf:
[| wellfounded(L, r); relation(r); L(r) |] ==> wf(r)
and wellfounded_on_imp_wf_on:
[| wellfounded_on(L, A, r); relation(r); L(r); L(A) |] ==> wf[A](r)
and wf_abs:
[| relation(r); L(r) |] ==> wellfounded(L, r) <-> wf(r)
and wf_on_abs:
[| relation(r); L(r); L(A) |] ==> wellfounded_on(L, A, r) <-> wf[A](r)
lemmas exists_wfrank:
[| wellfounded(L, r); L(a); L(r) |] ==> ∃f[L]. is_recfun(r^+, a, %x f. range(f), f)
and M_wellfoundedrank:
[| wellfounded(L, r); L(r); L(A) |] ==> L(wellfoundedrank(L, r, A))
and Ord_wfrank_range:
[| wellfounded(L, r); a ∈ A; L(r); L(A) |] ==> ∀f[L]. is_recfun(r^+, a, %x f. range(f), f) --> Ord(range(f))
and Ord_range_wellfoundedrank:
[| wellfounded(L, r); r ⊆ A × A; L(r); L(A) |] ==> Ord(range(wellfoundedrank(L, r, A)))
and function_wellfoundedrank:
[| wellfounded(L, r); L(r); L(A) |] ==> function(wellfoundedrank(L, r, A))
and domain_wellfoundedrank:
[| wellfounded(L, r); L(r); L(A) |] ==> domain(wellfoundedrank(L, r, A)) = A
and wellfoundedrank_type:
[| wellfounded(L, r); L(r); L(A) |] ==> wellfoundedrank(L, r, A) ∈ A -> range(wellfoundedrank(L, r, A))
and Ord_wellfoundedrank:
[| wellfounded(L, r); a ∈ A; r ⊆ A × A; L(r); L(A) |] ==> Ord(wellfoundedrank(L, r, A) ` a)
and wellfoundedrank_eq:
[| is_recfun(r^+, a, %x. range, f); wellfounded(L, r); a ∈ A; L(f); L(r); L(A) |] ==> wellfoundedrank(L, r, A) ` a = range(f)
and wellfoundedrank_lt:
[| 〈a, b〉 ∈ r; wellfounded(L, r); r ⊆ A × A; L(r); L(A) |] ==> wellfoundedrank(L, r, A) ` a < wellfoundedrank(L, r, A) ` b
and wellfounded_imp_subset_rvimage:
[| wellfounded(L, r); r ⊆ A × A; L(r); L(A) |] ==> ∃i f. Ord(i) ∧ r ⊆ rvimage(A, f, Memrel(i))
and wellfounded_imp_wf:
[| wellfounded(L, r); relation(r); L(r) |] ==> wf(r)
and wellfounded_on_imp_wf_on:
[| wellfounded_on(L, A, r); relation(r); L(r); L(A) |] ==> wf[A](r)
and wf_abs:
[| relation(r); L(r) |] ==> wellfounded(L, r) <-> wf(r)
and wf_on_abs:
[| relation(r); L(r); L(A) |] ==> wellfounded_on(L, A, r) <-> wf[A](r)