# HG changeset patch # User wenzelm # Date 1027897036 -7200 # Node ID 99e52e78eb655db6a8c056dbaa5f82fc7ca64154 # Parent b429fd98549ccb6fe57f19e631a19ec1959445f9 eliminate open locales and special ML code; diff -r b429fd98549c -r 99e52e78eb65 src/ZF/Constructible/Datatype_absolute.thy --- a/src/ZF/Constructible/Datatype_absolute.thy Sun Jul 28 21:09:37 2002 +0200 +++ b/src/ZF/Constructible/Datatype_absolute.thy Mon Jul 29 00:57:16 2002 +0200 @@ -362,7 +362,7 @@ is_formula :: "[i=>o,i] => o" "is_formula(M,Z) == \p[M]. p \ Z <-> mem_formula(M,p)" -locale (open) M_datatypes = M_wfrank + +locale M_datatypes = M_wfrank + assumes list_replacement1: "M(A) ==> iterates_replacement(M, is_list_functor(M,A), 0)" and list_replacement2: @@ -506,7 +506,7 @@ "is_eclose(M,A,Z) == \u[M]. u \ Z <-> mem_eclose(M,A,u)" -locale (open) M_eclose = M_datatypes + +locale M_eclose = M_datatypes + assumes eclose_replacement1: "M(A) ==> iterates_replacement(M, big_union(M), A)" and eclose_replacement2: diff -r b429fd98549c -r 99e52e78eb65 src/ZF/Constructible/L_axioms.thy --- a/src/ZF/Constructible/L_axioms.thy Sun Jul 28 21:09:37 2002 +0200 +++ b/src/ZF/Constructible/L_axioms.thy Mon Jul 29 00:57:16 2002 +0200 @@ -75,110 +75,65 @@ lemma Lset_cont: "cont_Ord(Lset)" by (simp add: cont_Ord_def Limit_Lset_eq OUnion_def Limit_is_Ord) -lemmas Pair_in_Lset = Formula.Pair_in_LLimit; +lemmas Pair_in_Lset = Formula.Pair_in_LLimit -lemmas L_nat = Ord_in_L [OF Ord_nat]; +lemmas L_nat = Ord_in_L [OF Ord_nat] -ML -{* -val transL = thm "transL"; -val nonempty = thm "nonempty"; -val upair_ax = thm "upair_ax"; -val Union_ax = thm "Union_ax"; -val power_ax = thm "power_ax"; -val replacement = thm "replacement"; -val L_nat = thm "L_nat"; - -fun kill_flex_triv_prems st = Seq.hd ((REPEAT_FIRST assume_tac) st); - -fun triv_axioms_L th = - kill_flex_triv_prems - ([transL, nonempty, upair_ax, Union_ax, power_ax, replacement, L_nat] - MRS (inst "M" "L" th)); +theorem M_triv_axioms_L: "PROP M_triv_axioms(L)" + apply (rule M_triv_axioms.intro) + apply (erule (1) transL) + apply (rule nonempty) + apply (rule upair_ax) + apply (rule Union_ax) + apply (rule power_ax) + apply (rule replacement) + apply (rule L_nat) + done -bind_thm ("rall_abs", triv_axioms_L (thm "M_triv_axioms.rall_abs")); -bind_thm ("rex_abs", triv_axioms_L (thm "M_triv_axioms.rex_abs")); -bind_thm ("ball_iff_equiv", triv_axioms_L (thm "M_triv_axioms.ball_iff_equiv")); -bind_thm ("M_equalityI", triv_axioms_L (thm "M_triv_axioms.M_equalityI")); -bind_thm ("empty_abs", triv_axioms_L (thm "M_triv_axioms.empty_abs")); -bind_thm ("subset_abs", triv_axioms_L (thm "M_triv_axioms.subset_abs")); -bind_thm ("upair_abs", triv_axioms_L (thm "M_triv_axioms.upair_abs")); -bind_thm ("upair_in_M_iff", triv_axioms_L (thm "M_triv_axioms.upair_in_M_iff")); -bind_thm ("singleton_in_M_iff", triv_axioms_L (thm "M_triv_axioms.singleton_in_M_iff")); -bind_thm ("pair_abs", triv_axioms_L (thm "M_triv_axioms.pair_abs")); -bind_thm ("pair_in_M_iff", triv_axioms_L (thm "M_triv_axioms.pair_in_M_iff")); -bind_thm ("pair_components_in_M", triv_axioms_L (thm "M_triv_axioms.pair_components_in_M")); -bind_thm ("cartprod_abs", triv_axioms_L (thm "M_triv_axioms.cartprod_abs")); -bind_thm ("union_abs", triv_axioms_L (thm "M_triv_axioms.union_abs")); -bind_thm ("inter_abs", triv_axioms_L (thm "M_triv_axioms.inter_abs")); -bind_thm ("setdiff_abs", triv_axioms_L (thm "M_triv_axioms.setdiff_abs")); -bind_thm ("Union_abs", triv_axioms_L (thm "M_triv_axioms.Union_abs")); -bind_thm ("Union_closed", triv_axioms_L (thm "M_triv_axioms.Union_closed")); -bind_thm ("Un_closed", triv_axioms_L (thm "M_triv_axioms.Un_closed")); -bind_thm ("cons_closed", triv_axioms_L (thm "M_triv_axioms.cons_closed")); -bind_thm ("successor_abs", triv_axioms_L (thm "M_triv_axioms.successor_abs")); -bind_thm ("succ_in_M_iff", triv_axioms_L (thm "M_triv_axioms.succ_in_M_iff")); -bind_thm ("separation_closed", triv_axioms_L (thm "M_triv_axioms.separation_closed")); -bind_thm ("strong_replacementI", triv_axioms_L (thm "M_triv_axioms.strong_replacementI")); -bind_thm ("strong_replacement_closed", triv_axioms_L (thm "M_triv_axioms.strong_replacement_closed")); -bind_thm ("RepFun_closed", triv_axioms_L (thm "M_triv_axioms.RepFun_closed")); -bind_thm ("lam_closed", triv_axioms_L (thm "M_triv_axioms.lam_closed")); -bind_thm ("image_abs", triv_axioms_L (thm "M_triv_axioms.image_abs")); -bind_thm ("powerset_Pow", triv_axioms_L (thm "M_triv_axioms.powerset_Pow")); -bind_thm ("powerset_imp_subset_Pow", triv_axioms_L (thm "M_triv_axioms.powerset_imp_subset_Pow")); -bind_thm ("nat_into_M", triv_axioms_L (thm "M_triv_axioms.nat_into_M")); -bind_thm ("nat_case_closed", triv_axioms_L (thm "M_triv_axioms.nat_case_closed")); -bind_thm ("Inl_in_M_iff", triv_axioms_L (thm "M_triv_axioms.Inl_in_M_iff")); -bind_thm ("Inr_in_M_iff", triv_axioms_L (thm "M_triv_axioms.Inr_in_M_iff")); -bind_thm ("lt_closed", triv_axioms_L (thm "M_triv_axioms.lt_closed")); -bind_thm ("transitive_set_abs", triv_axioms_L (thm "M_triv_axioms.transitive_set_abs")); -bind_thm ("ordinal_abs", triv_axioms_L (thm "M_triv_axioms.ordinal_abs")); -bind_thm ("limit_ordinal_abs", triv_axioms_L (thm "M_triv_axioms.limit_ordinal_abs")); -bind_thm ("successor_ordinal_abs", triv_axioms_L (thm "M_triv_axioms.successor_ordinal_abs")); -bind_thm ("finite_ordinal_abs", triv_axioms_L (thm "M_triv_axioms.finite_ordinal_abs")); -bind_thm ("omega_abs", triv_axioms_L (thm "M_triv_axioms.omega_abs")); -bind_thm ("number1_abs", triv_axioms_L (thm "M_triv_axioms.number1_abs")); -bind_thm ("number1_abs", triv_axioms_L (thm "M_triv_axioms.number1_abs")); -bind_thm ("number3_abs", triv_axioms_L (thm "M_triv_axioms.number3_abs")); -*} - -declare rall_abs [simp] -declare rex_abs [simp] -declare empty_abs [simp] -declare subset_abs [simp] -declare upair_abs [simp] -declare upair_in_M_iff [iff] -declare singleton_in_M_iff [iff] -declare pair_abs [simp] -declare pair_in_M_iff [iff] -declare cartprod_abs [simp] -declare union_abs [simp] -declare inter_abs [simp] -declare setdiff_abs [simp] -declare Union_abs [simp] -declare Union_closed [intro,simp] -declare Un_closed [intro,simp] -declare cons_closed [intro,simp] -declare successor_abs [simp] -declare succ_in_M_iff [iff] -declare separation_closed [intro,simp] -declare strong_replacementI -declare strong_replacement_closed [intro,simp] -declare RepFun_closed [intro,simp] -declare lam_closed [intro,simp] -declare image_abs [simp] -declare nat_into_M [intro] -declare Inl_in_M_iff [iff] -declare Inr_in_M_iff [iff] -declare transitive_set_abs [simp] -declare ordinal_abs [simp] -declare limit_ordinal_abs [simp] -declare successor_ordinal_abs [simp] -declare finite_ordinal_abs [simp] -declare omega_abs [simp] -declare number1_abs [simp] -declare number1_abs [simp] -declare number3_abs [simp] +lemmas rall_abs [simp] = M_triv_axioms.rall_abs [OF M_triv_axioms_L] + and rex_abs [simp] = M_triv_axioms.rex_abs [OF M_triv_axioms_L] + and ball_iff_equiv = M_triv_axioms.ball_iff_equiv [OF M_triv_axioms_L] + and M_equalityI = M_triv_axioms.M_equalityI [OF M_triv_axioms_L] + and empty_abs [simp] = M_triv_axioms.empty_abs [OF M_triv_axioms_L] + and subset_abs [simp] = M_triv_axioms.subset_abs [OF M_triv_axioms_L] + and upair_abs [simp] = M_triv_axioms.upair_abs [OF M_triv_axioms_L] + and upair_in_M_iff [iff] = M_triv_axioms.upair_in_M_iff [OF M_triv_axioms_L] + and singleton_in_M_iff [iff] = M_triv_axioms.singleton_in_M_iff [OF M_triv_axioms_L] + and pair_abs [simp] = M_triv_axioms.pair_abs [OF M_triv_axioms_L] + and pair_in_M_iff [iff] = M_triv_axioms.pair_in_M_iff [OF M_triv_axioms_L] + and pair_components_in_M = M_triv_axioms.pair_components_in_M [OF M_triv_axioms_L] + and cartprod_abs [simp] = M_triv_axioms.cartprod_abs [OF M_triv_axioms_L] + and union_abs [simp] = M_triv_axioms.union_abs [OF M_triv_axioms_L] + and inter_abs [simp] = M_triv_axioms.inter_abs [OF M_triv_axioms_L] + and setdiff_abs [simp] = M_triv_axioms.setdiff_abs [OF M_triv_axioms_L] + and Union_abs [simp] = M_triv_axioms.Union_abs [OF M_triv_axioms_L] + and Union_closed [intro, simp] = M_triv_axioms.Union_closed [OF M_triv_axioms_L] + and Un_closed [intro, simp] = M_triv_axioms.Un_closed [OF M_triv_axioms_L] + and cons_closed [intro, simp] = M_triv_axioms.cons_closed [OF M_triv_axioms_L] + and successor_abs [simp] = M_triv_axioms.successor_abs [OF M_triv_axioms_L] + and succ_in_M_iff [iff] = M_triv_axioms.succ_in_M_iff [OF M_triv_axioms_L] + and separation_closed [intro, simp] = M_triv_axioms.separation_closed [OF M_triv_axioms_L] + and strong_replacementI = M_triv_axioms.strong_replacementI [OF M_triv_axioms_L] + and strong_replacement_closed [intro, simp] = M_triv_axioms.strong_replacement_closed [OF M_triv_axioms_L] + and RepFun_closed [intro, simp] = M_triv_axioms.RepFun_closed [OF M_triv_axioms_L] + and lam_closed [intro, simp] = M_triv_axioms.lam_closed [OF M_triv_axioms_L] + and image_abs [simp] = M_triv_axioms.image_abs [OF M_triv_axioms_L] + and powerset_Pow = M_triv_axioms.powerset_Pow [OF M_triv_axioms_L] + and powerset_imp_subset_Pow = M_triv_axioms.powerset_imp_subset_Pow [OF M_triv_axioms_L] + and nat_into_M [intro] = M_triv_axioms.nat_into_M [OF M_triv_axioms_L] + and nat_case_closed = M_triv_axioms.nat_case_closed [OF M_triv_axioms_L] + and Inl_in_M_iff [iff] = M_triv_axioms.Inl_in_M_iff [OF M_triv_axioms_L] + and Inr_in_M_iff [iff] = M_triv_axioms.Inr_in_M_iff [OF M_triv_axioms_L] + and lt_closed = M_triv_axioms.lt_closed [OF M_triv_axioms_L] + and transitive_set_abs [simp] = M_triv_axioms.transitive_set_abs [OF M_triv_axioms_L] + and ordinal_abs [simp] = M_triv_axioms.ordinal_abs [OF M_triv_axioms_L] + and limit_ordinal_abs [simp] = M_triv_axioms.limit_ordinal_abs [OF M_triv_axioms_L] + and successor_ordinal_abs [simp] = M_triv_axioms.successor_ordinal_abs [OF M_triv_axioms_L] + and finite_ordinal_abs = M_triv_axioms.finite_ordinal_abs [OF M_triv_axioms_L] + and omega_abs [simp] = M_triv_axioms.omega_abs [OF M_triv_axioms_L] + and number1_abs [simp] = M_triv_axioms.number1_abs [OF M_triv_axioms_L] + and number2_abs [simp] = M_triv_axioms.number2_abs [OF M_triv_axioms_L] + and number3_abs [simp] = M_triv_axioms.number3_abs [OF M_triv_axioms_L] subsection{*Instantiation of the locale @{text reflection}*} @@ -260,8 +215,9 @@ apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def) apply (elim meta_exE) apply (rule meta_exI) -apply (rule reflection.Ex_reflection [OF Lset_mono_le Lset_cont Pair_in_Lset], - assumption+) +apply (rule reflection.Ex_reflection + [OF reflection.intro, OF Lset_mono_le Lset_cont Pair_in_Lset], + assumption+) done theorem All_reflection: @@ -270,7 +226,8 @@ apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def) apply (elim meta_exE) apply (rule meta_exI) -apply (rule reflection.All_reflection [OF Lset_mono_le Lset_cont Pair_in_Lset], +apply (rule reflection.All_reflection + [OF reflection.intro, OF Lset_mono_le Lset_cont Pair_in_Lset], assumption+) done @@ -308,7 +265,7 @@ apply (drule ReflectsD, assumption, blast) done -lemma Collect_mem_eq: "{x\A. x\B} = A \ B"; +lemma Collect_mem_eq: "{x\A. x\B} = A \ B" by blast diff -r b429fd98549c -r 99e52e78eb65 src/ZF/Constructible/Normal.thy --- a/src/ZF/Constructible/Normal.thy Sun Jul 28 21:09:37 2002 +0200 +++ b/src/ZF/Constructible/Normal.thy Mon Jul 29 00:57:16 2002 +0200 @@ -72,7 +72,7 @@ c.u.*} text{*The constructions below come from Kunen, \emph{Set Theory}, page 78.*} -locale (open) cub_family = +locale cub_family = fixes P and A fixes next_greater -- "the next ordinal satisfying class @{term A}" fixes sup_greater -- "sup of those ordinals over all @{term A}" @@ -177,7 +177,7 @@ "(!!a. a\A ==> Closed_Unbounded(P(a))) ==> Closed_Unbounded(\x. \a\A. P(a, x))" apply (case_tac "A=0", simp) -apply (rule cub_family.Closed_Unbounded_INT) +apply (rule cub_family.Closed_Unbounded_INT [OF cub_family.intro]) apply (simp_all add: Closed_Unbounded_def) done diff -r b429fd98549c -r 99e52e78eb65 src/ZF/Constructible/Rec_Separation.thy --- a/src/ZF/Constructible/Rec_Separation.thy Sun Jul 28 21:09:37 2002 +0200 +++ b/src/ZF/Constructible/Rec_Separation.thy Mon Jul 29 00:57:16 2002 +0200 @@ -6,7 +6,7 @@ "M_trancl"}, @{text "M_wfrank"} and @{text "M_datatypes"}*} lemma eq_succ_imp_lt: "[|i = succ(j); Ord(i)|] ==> jnnat[M]. \n[M]. \n'[M]. + \nnat[M]. \n[M]. \n'[M]. omega(M,nnat) & n\nnat & successor(M,n,n') & (\f[M]. typed_function(M,n',A,f) & - (\x[M]. \y[M]. \zero[M]. pair(M,x,y,p) & empty(M,zero) & - fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) & - (\j[M]. j\n --> - (\fj[M]. \sj[M]. \fsj[M]. \ffp[M]. - fun_apply(M,f,j,fj) & successor(M,j,sj) & - fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \ r)))"*) + (\x[M]. \y[M]. \zero[M]. pair(M,x,y,p) & empty(M,zero) & + fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) & + (\j[M]. j\n --> + (\fj[M]. \sj[M]. \fsj[M]. \ffp[M]. + fun_apply(M,f,j,fj) & successor(M,j,sj) & + fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \ r)))"*) constdefs rtran_closure_mem_fm :: "[i,i,i]=>i" - "rtran_closure_mem_fm(A,r,p) == + "rtran_closure_mem_fm(A,r,p) == Exists(Exists(Exists( And(omega_fm(2), And(Member(1,2), And(succ_fm(1,0), Exists(And(typed_function_fm(1, A#+4, 0), - And(Exists(Exists(Exists( - And(pair_fm(2,1,p#+7), - And(empty_fm(0), - And(fun_apply_fm(3,0,2), fun_apply_fm(3,5,1))))))), - Forall(Implies(Member(0,3), - Exists(Exists(Exists(Exists( - And(fun_apply_fm(5,4,3), - And(succ_fm(4,2), - And(fun_apply_fm(5,2,1), - And(pair_fm(3,1,0), Member(0,r#+9))))))))))))))))))))" + And(Exists(Exists(Exists( + And(pair_fm(2,1,p#+7), + And(empty_fm(0), + And(fun_apply_fm(3,0,2), fun_apply_fm(3,5,1))))))), + Forall(Implies(Member(0,3), + Exists(Exists(Exists(Exists( + And(fun_apply_fm(5,4,3), + And(succ_fm(4,2), + And(fun_apply_fm(5,2,1), + And(pair_fm(3,1,0), Member(0,r#+9))))))))))))))))))))" lemma rtran_closure_mem_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> rtran_closure_mem_fm(x,y,z) \ formula" -by (simp add: rtran_closure_mem_fm_def) +by (simp add: rtran_closure_mem_fm_def) lemma arity_rtran_closure_mem_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(rtran_closure_mem_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: rtran_closure_mem_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: rtran_closure_mem_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_rtran_closure_mem_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, rtran_closure_mem_fm(x,y,z), env) <-> + ==> sats(A, rtran_closure_mem_fm(x,y,z), env) <-> rtran_closure_mem(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: rtran_closure_mem_fm_def rtran_closure_mem_def) lemma rtran_closure_mem_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> rtran_closure_mem(**A, x, y, z) <-> sats(A, rtran_closure_mem_fm(i,j,k), env)" by (simp add: sats_rtran_closure_mem_fm) theorem rtran_closure_mem_reflection: - "REFLECTS[\x. rtran_closure_mem(L,f(x),g(x),h(x)), + "REFLECTS[\x. rtran_closure_mem(L,f(x),g(x),h(x)), \i x. rtran_closure_mem(**Lset(i),f(x),g(x),h(x))]" apply (simp only: rtran_closure_mem_def setclass_simps) -apply (intro FOL_reflections function_reflections fun_plus_reflections) +apply (intro FOL_reflections function_reflections fun_plus_reflections) done text{*Separation for @{term "rtrancl(r)"}.*} lemma rtrancl_separation: "[| L(r); L(A) |] ==> separation (L, rtran_closure_mem(L,A,r))" -apply (rule separation_CollectI) -apply (rule_tac A="{r,A,z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{r,A,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF rtran_closure_mem_reflection], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) @@ -89,38 +89,38 @@ subsubsection{*Reflexive/Transitive Closure, Internalized*} -(* "rtran_closure(M,r,s) == +(* "rtran_closure(M,r,s) == \A[M]. is_field(M,r,A) --> - (\p[M]. p \ s <-> rtran_closure_mem(M,A,r,p))" *) + (\p[M]. p \ s <-> rtran_closure_mem(M,A,r,p))" *) constdefs rtran_closure_fm :: "[i,i]=>i" - "rtran_closure_fm(r,s) == + "rtran_closure_fm(r,s) == Forall(Implies(field_fm(succ(r),0), Forall(Iff(Member(0,succ(succ(s))), rtran_closure_mem_fm(1,succ(succ(r)),0)))))" lemma rtran_closure_type [TC]: "[| x \ nat; y \ nat |] ==> rtran_closure_fm(x,y) \ formula" -by (simp add: rtran_closure_fm_def) +by (simp add: rtran_closure_fm_def) lemma arity_rtran_closure_fm [simp]: - "[| x \ nat; y \ nat |] + "[| x \ nat; y \ nat |] ==> arity(rtran_closure_fm(x,y)) = succ(x) \ succ(y)" by (simp add: rtran_closure_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_rtran_closure_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, rtran_closure_fm(x,y), env) <-> + ==> sats(A, rtran_closure_fm(x,y), env) <-> rtran_closure(**A, nth(x,env), nth(y,env))" by (simp add: rtran_closure_fm_def rtran_closure_def) lemma rtran_closure_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; + "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> rtran_closure(**A, x, y) <-> sats(A, rtran_closure_fm(i,j), env)" by simp theorem rtran_closure_reflection: - "REFLECTS[\x. rtran_closure(L,f(x),g(x)), + "REFLECTS[\x. rtran_closure(L,f(x),g(x)), \i x. rtran_closure(**Lset(i),f(x),g(x))]" apply (simp only: rtran_closure_def setclass_simps) apply (intro FOL_reflections function_reflections rtran_closure_mem_reflection) @@ -132,35 +132,35 @@ (* "tran_closure(M,r,t) == \s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" *) constdefs tran_closure_fm :: "[i,i]=>i" - "tran_closure_fm(r,s) == + "tran_closure_fm(r,s) == Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))" lemma tran_closure_type [TC]: "[| x \ nat; y \ nat |] ==> tran_closure_fm(x,y) \ formula" -by (simp add: tran_closure_fm_def) +by (simp add: tran_closure_fm_def) lemma arity_tran_closure_fm [simp]: - "[| x \ nat; y \ nat |] + "[| x \ nat; y \ nat |] ==> arity(tran_closure_fm(x,y)) = succ(x) \ succ(y)" by (simp add: tran_closure_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_tran_closure_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, tran_closure_fm(x,y), env) <-> + ==> sats(A, tran_closure_fm(x,y), env) <-> tran_closure(**A, nth(x,env), nth(y,env))" by (simp add: tran_closure_fm_def tran_closure_def) lemma tran_closure_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; + "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> tran_closure(**A, x, y) <-> sats(A, tran_closure_fm(i,j), env)" by simp theorem tran_closure_reflection: - "REFLECTS[\x. tran_closure(L,f(x),g(x)), + "REFLECTS[\x. tran_closure(L,f(x),g(x)), \i x. tran_closure(**Lset(i),f(x),g(x))]" apply (simp only: tran_closure_def setclass_simps) -apply (intro FOL_reflections function_reflections +apply (intro FOL_reflections function_reflections rtran_closure_reflection composition_reflection) done @@ -168,60 +168,62 @@ subsection{*Separation for the Proof of @{text "wellfounded_on_trancl"}*} lemma wellfounded_trancl_reflects: - "REFLECTS[\x. \w[L]. \wx[L]. \rp[L]. - w \ Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \ rp, - \i x. \w \ Lset(i). \wx \ Lset(i). \rp \ Lset(i). + "REFLECTS[\x. \w[L]. \wx[L]. \rp[L]. + w \ Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \ rp, + \i x. \w \ Lset(i). \wx \ Lset(i). \rp \ Lset(i). w \ Z & pair(**Lset(i),w,x,wx) & tran_closure(**Lset(i),r,rp) & wx \ rp]" -by (intro FOL_reflections function_reflections fun_plus_reflections +by (intro FOL_reflections function_reflections fun_plus_reflections tran_closure_reflection) lemma wellfounded_trancl_separation: - "[| L(r); L(Z) |] ==> - separation (L, \x. - \w[L]. \wx[L]. \rp[L]. - w \ Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \ rp)" -apply (rule separation_CollectI) -apply (rule_tac A="{r,Z,z}" in subset_LsetE, blast ) + "[| L(r); L(Z) |] ==> + separation (L, \x. + \w[L]. \wx[L]. \rp[L]. + w \ Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \ rp)" +apply (rule separation_CollectI) +apply (rule_tac A="{r,Z,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF wellfounded_trancl_reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule bex_iff_sats conj_iff_sats)+ -apply (rule_tac env = "[w,u,r,Z]" in mem_iff_sats) +apply (rule_tac env = "[w,u,r,Z]" in mem_iff_sats) apply (rule sep_rules tran_closure_iff_sats | simp)+ done subsubsection{*Instantiating the locale @{text M_trancl}*} -ML -{* -val rtrancl_separation = thm "rtrancl_separation"; -val wellfounded_trancl_separation = thm "wellfounded_trancl_separation"; + +theorem M_trancl_axioms_L: "M_trancl_axioms(L)" + apply (rule M_trancl_axioms.intro) + apply (assumption | rule + rtrancl_separation wellfounded_trancl_separation)+ + done - -val m_trancl = [rtrancl_separation, wellfounded_trancl_separation]; - -fun trancl_L th = - kill_flex_triv_prems (m_trancl MRS (axioms_L th)); +theorem M_trancl_L: "PROP M_trancl(L)" + apply (rule M_trancl.intro) + apply (rule M_triv_axioms_L) + apply (rule M_axioms_axioms_L) + apply (rule M_trancl_axioms_L) + done -bind_thm ("iterates_abs", trancl_L (thm "M_trancl.iterates_abs")); -bind_thm ("rtran_closure_rtrancl", trancl_L (thm "M_trancl.rtran_closure_rtrancl")); -bind_thm ("rtrancl_closed", trancl_L (thm "M_trancl.rtrancl_closed")); -bind_thm ("rtrancl_abs", trancl_L (thm "M_trancl.rtrancl_abs")); -bind_thm ("trancl_closed", trancl_L (thm "M_trancl.trancl_closed")); -bind_thm ("trancl_abs", trancl_L (thm "M_trancl.trancl_abs")); -bind_thm ("wellfounded_on_trancl", trancl_L (thm "M_trancl.wellfounded_on_trancl")); -bind_thm ("wellfounded_trancl", trancl_L (thm "M_trancl.wellfounded_trancl")); -bind_thm ("wfrec_relativize", trancl_L (thm "M_trancl.wfrec_relativize")); -bind_thm ("trans_wfrec_relativize", trancl_L (thm "M_trancl.trans_wfrec_relativize")); -bind_thm ("trans_wfrec_abs", trancl_L (thm "M_trancl.trans_wfrec_abs")); -bind_thm ("trans_eq_pair_wfrec_iff", trancl_L (thm "M_trancl.trans_eq_pair_wfrec_iff")); -bind_thm ("eq_pair_wfrec_iff", trancl_L (thm "M_trancl.eq_pair_wfrec_iff")); -*} +lemmas iterates_abs = M_trancl.iterates_abs [OF M_trancl_L] + and rtran_closure_rtrancl = M_trancl.rtran_closure_rtrancl [OF M_trancl_L] + and rtrancl_closed = M_trancl.rtrancl_closed [OF M_trancl_L] + and rtrancl_abs = M_trancl.rtrancl_abs [OF M_trancl_L] + and trancl_closed = M_trancl.trancl_closed [OF M_trancl_L] + and trancl_abs = M_trancl.trancl_abs [OF M_trancl_L] + and wellfounded_on_trancl = M_trancl.wellfounded_on_trancl [OF M_trancl_L] + and wellfounded_trancl = M_trancl.wellfounded_trancl [OF M_trancl_L] + and wfrec_relativize = M_trancl.wfrec_relativize [OF M_trancl_L] + and trans_wfrec_relativize = M_trancl.trans_wfrec_relativize [OF M_trancl_L] + and trans_wfrec_abs = M_trancl.trans_wfrec_abs [OF M_trancl_L] + and trans_eq_pair_wfrec_iff = M_trancl.trans_eq_pair_wfrec_iff [OF M_trancl_L] + and eq_pair_wfrec_iff = M_trancl.eq_pair_wfrec_iff [OF M_trancl_L] declare rtrancl_closed [intro,simp] declare rtrancl_abs [simp] @@ -232,17 +234,17 @@ subsection{*Well-Founded Recursion!*} (* M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" - "M_is_recfun(M,MH,r,a,f) == - \z[M]. z \ f <-> + "M_is_recfun(M,MH,r,a,f) == + \z[M]. z \ f <-> 5 4 3 2 1 0 - (\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) & + (\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))" *) constdefs is_recfun_fm :: "[[i,i,i]=>i, i, i, i]=>i" - "is_recfun_fm(p,r,a,f) == + "is_recfun_fm(p,r,a,f) == Forall(Iff(Member(0,succ(f)), Exists(Exists(Exists(Exists(Exists(Exists( And(pair_fm(5,4,6), @@ -254,39 +256,39 @@ lemma is_recfun_type_0: - "[| !!x y z. [| x \ nat; y \ nat; z \ nat |] ==> p(x,y,z) \ formula; - x \ nat; y \ nat; z \ nat |] + "[| !!x y z. [| x \ nat; y \ nat; z \ nat |] ==> p(x,y,z) \ formula; + x \ nat; y \ nat; z \ nat |] ==> is_recfun_fm(p,x,y,z) \ formula" apply (unfold is_recfun_fm_def) (*FIXME: FIND OUT why simp loops!*) apply typecheck -by simp +by simp lemma is_recfun_type [TC]: - "[| p(5,0,4) \ formula; - x \ nat; y \ nat; z \ nat |] + "[| p(5,0,4) \ formula; + x \ nat; y \ nat; z \ nat |] ==> is_recfun_fm(p,x,y,z) \ formula" -by (simp add: is_recfun_fm_def) +by (simp add: is_recfun_fm_def) lemma arity_is_recfun_fm [simp]: - "[| arity(p(5,0,4)) le 8; x \ nat; y \ nat; z \ nat |] + "[| arity(p(5,0,4)) le 8; x \ nat; y \ nat; z \ nat |] ==> arity(is_recfun_fm(p,x,y,z)) = succ(x) \ succ(y) \ succ(z)" -apply (frule lt_nat_in_nat, simp) -apply (simp add: is_recfun_fm_def succ_Un_distrib [symmetric] ) -apply (subst subset_Un_iff2 [of "arity(p(5,0,4))", THEN iffD1]) -apply (rule le_imp_subset) -apply (erule le_trans, simp) -apply (simp add: succ_Un_distrib [symmetric] Un_ac) +apply (frule lt_nat_in_nat, simp) +apply (simp add: is_recfun_fm_def succ_Un_distrib [symmetric] ) +apply (subst subset_Un_iff2 [of "arity(p(5,0,4))", THEN iffD1]) +apply (rule le_imp_subset) +apply (erule le_trans, simp) +apply (simp add: succ_Un_distrib [symmetric] Un_ac) done lemma sats_is_recfun_fm: - assumes MH_iff_sats: - "!!x y z env. - [| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> MH(nth(x,env), nth(y,env), nth(z,env)) <-> sats(A, p(x,y,z), env)" - shows + assumes MH_iff_sats: + "!!x y z env. + [| x \ nat; y \ nat; z \ nat; env \ list(A)|] + ==> MH(nth(x,env), nth(y,env), nth(z,env)) <-> sats(A, p(x,y,z), env)" + shows "[|x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, is_recfun_fm(p,x,y,z), env) <-> + ==> sats(A, is_recfun_fm(p,x,y,z), env) <-> M_is_recfun(**A, MH, nth(x,env), nth(y,env), nth(z,env))" by (simp add: is_recfun_fm_def M_is_recfun_def MH_iff_sats [THEN iff_sym]) @@ -294,20 +296,20 @@ "[| (!!x y z env. [| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> MH(nth(x,env), nth(y,env), nth(z,env)) <-> sats(A, p(x,y,z), env)); - nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> M_is_recfun(**A, MH, x, y, z) <-> sats(A, is_recfun_fm(p,i,j,k), env)" + ==> M_is_recfun(**A, MH, x, y, z) <-> sats(A, is_recfun_fm(p,i,j,k), env)" by (simp add: sats_is_recfun_fm [of A MH]) theorem is_recfun_reflection: assumes MH_reflection: - "!!f g h. REFLECTS[\x. MH(L, f(x), g(x), h(x)), + "!!f g h. REFLECTS[\x. MH(L, f(x), g(x), h(x)), \i x. MH(**Lset(i), f(x), g(x), h(x))]" - shows "REFLECTS[\x. M_is_recfun(L, MH(L), f(x), g(x), h(x)), + shows "REFLECTS[\x. M_is_recfun(L, MH(L), f(x), g(x), h(x)), \i x. M_is_recfun(**Lset(i), MH(**Lset(i)), f(x), g(x), h(x))]" apply (simp (no_asm_use) only: M_is_recfun_def setclass_simps) -apply (intro FOL_reflections function_reflections - restriction_reflection MH_reflection) +apply (intro FOL_reflections function_reflections + restriction_reflection MH_reflection) done text{*Currently, @{text sats}-theorems for higher-order operators don't seem @@ -315,12 +317,12 @@ of the @{text MH}-term.*} theorem is_wfrec_reflection: assumes MH_reflection: - "!!f g h. REFLECTS[\x. MH(L, f(x), g(x), h(x)), + "!!f g h. REFLECTS[\x. MH(L, f(x), g(x), h(x)), \i x. MH(**Lset(i), f(x), g(x), h(x))]" - shows "REFLECTS[\x. is_wfrec(L, MH(L), f(x), g(x), h(x)), + shows "REFLECTS[\x. is_wfrec(L, MH(L), f(x), g(x), h(x)), \i x. is_wfrec(**Lset(i), MH(**Lset(i)), f(x), g(x), h(x))]" apply (simp (no_asm_use) only: is_wfrec_def setclass_simps) -apply (intro FOL_reflections MH_reflection is_recfun_reflection) +apply (intro FOL_reflections MH_reflection is_recfun_reflection) done subsection{*The Locale @{text "M_wfrank"}*} @@ -331,23 +333,23 @@ "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), + ~ (\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) +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 separation_CollectI) -apply (rule_tac A="{r,z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{r,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF wfrank_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2, clarify) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule ball_iff_sats imp_iff_sats)+ apply (rule_tac env="[rplus,u,r]" in tran_closure_iff_sats) apply (rule sep_rules is_recfun_iff_sats | simp)+ @@ -357,14 +359,14 @@ subsubsection{*Replacement for @{term "wfrank"}*} lemma wfrank_replacement_Reflects: - "REFLECTS[\z. \x[L]. x \ A & + "REFLECTS[\z. \x[L]. x \ A & (\rplus[L]. tran_closure(L,r,rplus) --> - (\y[L]. \f[L]. pair(L,x,y,z) & + (\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 & + \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) & + (\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 @@ -373,24 +375,24 @@ lemma wfrank_strong_replacement: "L(r) ==> - strong_replacement(L, \x z. + strong_replacement(L, \x z. \rplus[L]. tran_closure(L,r,rplus) --> - (\y[L]. \f[L]. pair(L,x,y,z) & + (\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 strong_replacementI) apply (rule rallI) -apply (rename_tac B) -apply (rule separation_CollectI) -apply (rule_tac A="{B,r,z}" in subset_LsetE, blast ) +apply (rename_tac B) +apply (rule separation_CollectI) +apply (rule_tac A="{B,r,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF wfrank_replacement_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule bex_iff_sats ball_iff_sats conj_iff_sats)+ -apply (rule_tac env = "[x,u,B,r]" in mem_iff_sats) +apply (rule_tac env = "[x,u,B,r]" in mem_iff_sats) apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+ done @@ -398,36 +400,36 @@ 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]. + "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). + \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), + 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 +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]. + \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 separation_CollectI) -apply (rule_tac A="{r,z}" in subset_LsetE, blast ) + ordinal(L,rangef)))" +apply (rule separation_CollectI) +apply (rule_tac A="{r,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF Ord_wfrank_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2, clarify) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule ball_iff_sats imp_iff_sats)+ apply (rule_tac env="[rplus,u,r]" in tran_closure_iff_sats) apply (rule sep_rules is_recfun_iff_sats | simp)+ @@ -435,40 +437,40 @@ subsubsection{*Instantiating the locale @{text M_wfrank}*} -ML -{* -val wfrank_separation = thm "wfrank_separation"; -val wfrank_strong_replacement = thm "wfrank_strong_replacement"; -val Ord_wfrank_separation = thm "Ord_wfrank_separation"; + +theorem 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 -val m_wfrank = - [wfrank_separation, wfrank_strong_replacement, Ord_wfrank_separation]; - -fun wfrank_L th = - kill_flex_triv_prems (m_wfrank MRS (trancl_L th)); - - +theorem M_wfrank_L: "PROP M_wfrank(L)" + apply (rule M_wfrank.intro) + apply (rule M_triv_axioms_L) + apply (rule M_axioms_axioms_L) + apply (rule M_trancl_axioms_L) + apply (rule M_wfrank_axioms_L) + done -bind_thm ("iterates_closed", wfrank_L (thm "M_wfrank.iterates_closed")); -bind_thm ("exists_wfrank", wfrank_L (thm "M_wfrank.exists_wfrank")); -bind_thm ("M_wellfoundedrank", wfrank_L (thm "M_wfrank.M_wellfoundedrank")); -bind_thm ("Ord_wfrank_range", wfrank_L (thm "M_wfrank.Ord_wfrank_range")); -bind_thm ("Ord_range_wellfoundedrank", wfrank_L (thm "M_wfrank.Ord_range_wellfoundedrank")); -bind_thm ("function_wellfoundedrank", wfrank_L (thm "M_wfrank.function_wellfoundedrank")); -bind_thm ("domain_wellfoundedrank", wfrank_L (thm "M_wfrank.domain_wellfoundedrank")); -bind_thm ("wellfoundedrank_type", wfrank_L (thm "M_wfrank.wellfoundedrank_type")); -bind_thm ("Ord_wellfoundedrank", wfrank_L (thm "M_wfrank.Ord_wellfoundedrank")); -bind_thm ("wellfoundedrank_eq", wfrank_L (thm "M_wfrank.wellfoundedrank_eq")); -bind_thm ("wellfoundedrank_lt", wfrank_L (thm "M_wfrank.wellfoundedrank_lt")); -bind_thm ("wellfounded_imp_subset_rvimage", wfrank_L (thm "M_wfrank.wellfounded_imp_subset_rvimage")); -bind_thm ("wellfounded_imp_wf", wfrank_L (thm "M_wfrank.wellfounded_imp_wf")); -bind_thm ("wellfounded_on_imp_wf_on", wfrank_L (thm "M_wfrank.wellfounded_on_imp_wf_on")); -bind_thm ("wf_abs", wfrank_L (thm "M_wfrank.wf_abs")); -bind_thm ("wf_on_abs", wfrank_L (thm "M_wfrank.wf_on_abs")); -bind_thm ("wfrec_replacement_iff", wfrank_L (thm "M_wfrank.wfrec_replacement_iff")); -bind_thm ("trans_wfrec_closed", wfrank_L (thm "M_wfrank.trans_wfrec_closed")); -bind_thm ("wfrec_closed", wfrank_L (thm "M_wfrank.wfrec_closed")); -*} +lemmas iterates_closed = M_wfrank.iterates_closed [OF M_wfrank_L] + and 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] + and wfrec_replacement_iff = M_wfrank.wfrec_replacement_iff [OF M_wfrank_L] + and trans_wfrec_closed = M_wfrank.trans_wfrec_closed [OF M_wfrank_L] + and wfrec_closed = M_wfrank.wfrec_closed [OF M_wfrank_L] declare iterates_closed [intro,simp] declare Ord_wfrank_range [rule_format] @@ -481,9 +483,9 @@ subsubsection{*Binary Products, Internalized*} constdefs cartprod_fm :: "[i,i,i]=>i" -(* "cartprod(M,A,B,z) == - \u[M]. u \ z <-> (\x[M]. x\A & (\y[M]. y\B & pair(M,x,y,u)))" *) - "cartprod_fm(A,B,z) == +(* "cartprod(M,A,B,z) == + \u[M]. u \ z <-> (\x[M]. x\A & (\y[M]. y\B & pair(M,x,y,u)))" *) + "cartprod_fm(A,B,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(A))), Exists(And(Member(0,succ(succ(succ(B)))), @@ -491,74 +493,74 @@ lemma cartprod_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> cartprod_fm(x,y,z) \ formula" -by (simp add: cartprod_fm_def) +by (simp add: cartprod_fm_def) lemma arity_cartprod_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(cartprod_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: cartprod_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: cartprod_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_cartprod_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, cartprod_fm(x,y,z), env) <-> + ==> sats(A, cartprod_fm(x,y,z), env) <-> cartprod(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: cartprod_fm_def cartprod_def) lemma cartprod_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> cartprod(**A, x, y, z) <-> sats(A, cartprod_fm(i,j,k), env)" by (simp add: sats_cartprod_fm) theorem cartprod_reflection: - "REFLECTS[\x. cartprod(L,f(x),g(x),h(x)), + "REFLECTS[\x. cartprod(L,f(x),g(x),h(x)), \i x. cartprod(**Lset(i),f(x),g(x),h(x))]" apply (simp only: cartprod_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done subsubsection{*Binary Sums, Internalized*} -(* "is_sum(M,A,B,Z) == - \A0[M]. \n1[M]. \s1[M]. \B1[M]. +(* "is_sum(M,A,B,Z) == + \A0[M]. \n1[M]. \s1[M]. \B1[M]. 3 2 1 0 number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) & cartprod(M,s1,B,B1) & union(M,A0,B1,Z)" *) constdefs sum_fm :: "[i,i,i]=>i" - "sum_fm(A,B,Z) == + "sum_fm(A,B,Z) == Exists(Exists(Exists(Exists( - And(number1_fm(2), + And(number1_fm(2), And(cartprod_fm(2,A#+4,3), And(upair_fm(2,2,1), And(cartprod_fm(1,B#+4,0), union_fm(3,0,Z#+4)))))))))" lemma sum_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> sum_fm(x,y,z) \ formula" -by (simp add: sum_fm_def) +by (simp add: sum_fm_def) lemma arity_sum_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(sum_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: sum_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: sum_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_sum_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, sum_fm(x,y,z), env) <-> + ==> sats(A, sum_fm(x,y,z), env) <-> is_sum(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: sum_fm_def is_sum_def) lemma sum_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> is_sum(**A, x, y, z) <-> sats(A, sum_fm(i,j,k), env)" by simp theorem sum_reflection: - "REFLECTS[\x. is_sum(L,f(x),g(x),h(x)), + "REFLECTS[\x. is_sum(L,f(x),g(x),h(x)), \i x. is_sum(**Lset(i),f(x),g(x),h(x))]" apply (simp only: is_sum_def setclass_simps) -apply (intro FOL_reflections function_reflections cartprod_reflection) +apply (intro FOL_reflections function_reflections cartprod_reflection) done @@ -570,11 +572,11 @@ lemma quasinat_type [TC]: "x \ nat ==> quasinat_fm(x) \ formula" -by (simp add: quasinat_fm_def) +by (simp add: quasinat_fm_def) lemma arity_quasinat_fm [simp]: "x \ nat ==> arity(quasinat_fm(x)) = succ(x)" -by (simp add: quasinat_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: quasinat_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_quasinat_fm [simp]: "[| x \ nat; env \ list(A)|] @@ -582,85 +584,85 @@ by (simp add: quasinat_fm_def is_quasinat_def) lemma quasinat_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; + "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> is_quasinat(**A, x) <-> sats(A, quasinat_fm(i), env)" by simp theorem quasinat_reflection: - "REFLECTS[\x. is_quasinat(L,f(x)), + "REFLECTS[\x. is_quasinat(L,f(x)), \i x. is_quasinat(**Lset(i),f(x))]" apply (simp only: is_quasinat_def setclass_simps) -apply (intro FOL_reflections function_reflections) +apply (intro FOL_reflections function_reflections) done subsubsection{*The Operator @{term is_nat_case}*} (* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" - "is_nat_case(M, a, is_b, k, z) == + "is_nat_case(M, a, is_b, k, z) == (empty(M,k) --> z=a) & (\m[M]. successor(M,m,k) --> is_b(m,z)) & (is_quasinat(M,k) | empty(M,z))" *) text{*The formula @{term is_b} has free variables 1 and 0.*} constdefs is_nat_case_fm :: "[i, [i,i]=>i, i, i]=>i" - "is_nat_case_fm(a,is_b,k,z) == + "is_nat_case_fm(a,is_b,k,z) == And(Implies(empty_fm(k), Equal(z,a)), - And(Forall(Implies(succ_fm(0,succ(k)), + And(Forall(Implies(succ_fm(0,succ(k)), Forall(Implies(Equal(0,succ(succ(z))), is_b(1,0))))), Or(quasinat_fm(k), empty_fm(z))))" lemma is_nat_case_type [TC]: - "[| is_b(1,0) \ formula; - x \ nat; y \ nat; z \ nat |] + "[| is_b(1,0) \ formula; + x \ nat; y \ nat; z \ nat |] ==> is_nat_case_fm(x,is_b,y,z) \ formula" -by (simp add: is_nat_case_fm_def) +by (simp add: is_nat_case_fm_def) lemma arity_is_nat_case_fm [simp]: - "[| is_b(1,0) \ formula; x \ nat; y \ nat; z \ nat |] - ==> arity(is_nat_case_fm(x,is_b,y,z)) = - succ(x) \ succ(y) \ succ(z) \ (arity(is_b(1,0)) #- 2)" -apply (subgoal_tac "arity(is_b(1,0)) \ nat") + "[| is_b(1,0) \ formula; x \ nat; y \ nat; z \ nat |] + ==> arity(is_nat_case_fm(x,is_b,y,z)) = + succ(x) \ succ(y) \ succ(z) \ (arity(is_b(1,0)) #- 2)" +apply (subgoal_tac "arity(is_b(1,0)) \ nat") apply typecheck (*FIXME: could nat_diff_split work?*) apply (auto simp add: diff_def raw_diff_succ is_nat_case_fm_def nat_imp_quasinat succ_Un_distrib [symmetric] Un_ac - split: split_nat_case) + split: split_nat_case) done lemma sats_is_nat_case_fm: - assumes is_b_iff_sats: - "!!a b. [| a \ A; b \ A|] + assumes is_b_iff_sats: + "!!a b. [| a \ A; b \ A|] ==> is_b(a,b) <-> sats(A, p(1,0), Cons(b, Cons(a,env)))" - shows + shows "[|x \ nat; y \ nat; z < length(env); env \ list(A)|] - ==> sats(A, is_nat_case_fm(x,p,y,z), env) <-> + ==> sats(A, is_nat_case_fm(x,p,y,z), env) <-> is_nat_case(**A, nth(x,env), is_b, nth(y,env), nth(z,env))" -apply (frule lt_length_in_nat, assumption) +apply (frule lt_length_in_nat, assumption) apply (simp add: is_nat_case_fm_def is_nat_case_def is_b_iff_sats [THEN iff_sym]) done lemma is_nat_case_iff_sats: - "[| (!!a b. [| a \ A; b \ A|] + "[| (!!a b. [| a \ A; b \ A|] ==> is_b(a,b) <-> sats(A, p(1,0), Cons(b, Cons(a,env)))); - nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k < length(env); env \ list(A)|] - ==> is_nat_case(**A, x, is_b, y, z) <-> sats(A, is_nat_case_fm(i,p,j,k), env)" + ==> is_nat_case(**A, x, is_b, y, z) <-> sats(A, is_nat_case_fm(i,p,j,k), env)" by (simp add: sats_is_nat_case_fm [of A is_b]) text{*The second argument of @{term is_b} gives it direct access to @{term x}, - which is essential for handling free variable references. Without this + which is essential for handling free variable references. Without this argument, we cannot prove reflection for @{term iterates_MH}.*} theorem is_nat_case_reflection: assumes is_b_reflection: - "!!h f g. REFLECTS[\x. is_b(L, h(x), f(x), g(x)), + "!!h f g. REFLECTS[\x. is_b(L, h(x), f(x), g(x)), \i x. is_b(**Lset(i), h(x), f(x), g(x))]" - shows "REFLECTS[\x. is_nat_case(L, f(x), is_b(L,x), g(x), h(x)), + shows "REFLECTS[\x. is_nat_case(L, f(x), is_b(L,x), g(x), h(x)), \i x. is_nat_case(**Lset(i), f(x), is_b(**Lset(i), x), g(x), h(x))]" apply (simp (no_asm_use) only: is_nat_case_def setclass_simps) -apply (intro FOL_reflections function_reflections - restriction_reflection is_b_reflection quasinat_reflection) +apply (intro FOL_reflections function_reflections + restriction_reflection is_b_reflection quasinat_reflection) done @@ -672,117 +674,117 @@ is_nat_case(M, v, \m u. \gm[M]. fun_apply(M,g,m,gm) & isF(gm,u), n, z)" *) constdefs iterates_MH_fm :: "[[i,i]=>i, i, i, i, i]=>i" - "iterates_MH_fm(isF,v,n,g,z) == - is_nat_case_fm(v, - \m u. Exists(And(fun_apply_fm(succ(succ(succ(g))),succ(m),0), - Forall(Implies(Equal(0,succ(succ(u))), isF(1,0))))), + "iterates_MH_fm(isF,v,n,g,z) == + is_nat_case_fm(v, + \m u. Exists(And(fun_apply_fm(succ(succ(succ(g))),succ(m),0), + Forall(Implies(Equal(0,succ(succ(u))), isF(1,0))))), n, z)" lemma iterates_MH_type [TC]: - "[| p(1,0) \ formula; - v \ nat; x \ nat; y \ nat; z \ nat |] + "[| p(1,0) \ formula; + v \ nat; x \ nat; y \ nat; z \ nat |] ==> iterates_MH_fm(p,v,x,y,z) \ formula" -by (simp add: iterates_MH_fm_def) +by (simp add: iterates_MH_fm_def) lemma arity_iterates_MH_fm [simp]: - "[| p(1,0) \ formula; - v \ nat; x \ nat; y \ nat; z \ nat |] - ==> arity(iterates_MH_fm(p,v,x,y,z)) = + "[| p(1,0) \ formula; + v \ nat; x \ nat; y \ nat; z \ nat |] + ==> arity(iterates_MH_fm(p,v,x,y,z)) = succ(v) \ succ(x) \ succ(y) \ succ(z) \ (arity(p(1,0)) #- 4)" apply (subgoal_tac "arity(p(1,0)) \ nat") apply typecheck apply (simp add: nat_imp_quasinat iterates_MH_fm_def Un_ac split: split_nat_case, clarify) apply (rename_tac i j) -apply (drule eq_succ_imp_eq_m1, simp) +apply (drule eq_succ_imp_eq_m1, simp) apply (drule eq_succ_imp_eq_m1, simp) apply (simp add: diff_Un_distrib succ_Un_distrib Un_ac diff_diff_left) done lemma sats_iterates_MH_fm: - assumes is_F_iff_sats: - "!!a b c d. [| a \ A; b \ A; c \ A; d \ A|] + assumes is_F_iff_sats: + "!!a b c d. [| a \ A; b \ A; c \ A; d \ A|] ==> is_F(a,b) <-> sats(A, p(1,0), Cons(b, Cons(a, Cons(c, Cons(d,env)))))" - shows + shows "[|v \ nat; x \ nat; y \ nat; z < length(env); env \ list(A)|] - ==> sats(A, iterates_MH_fm(p,v,x,y,z), env) <-> + ==> sats(A, iterates_MH_fm(p,v,x,y,z), env) <-> iterates_MH(**A, is_F, nth(v,env), nth(x,env), nth(y,env), nth(z,env))" -by (simp add: iterates_MH_fm_def iterates_MH_def sats_is_nat_case_fm +by (simp add: iterates_MH_fm_def iterates_MH_def sats_is_nat_case_fm is_F_iff_sats [symmetric]) lemma iterates_MH_iff_sats: - "[| (!!a b c d. [| a \ A; b \ A; c \ A; d \ A|] + "[| (!!a b c d. [| a \ A; b \ A; c \ A; d \ A|] ==> is_F(a,b) <-> sats(A, p(1,0), Cons(b, Cons(a, Cons(c, Cons(d,env)))))); - nth(i',env) = v; nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + nth(i',env) = v; nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i' \ nat; i \ nat; j \ nat; k < length(env); env \ list(A)|] - ==> iterates_MH(**A, is_F, v, x, y, z) <-> + ==> iterates_MH(**A, is_F, v, x, y, z) <-> sats(A, iterates_MH_fm(p,i',i,j,k), env)" -apply (rule iff_sym) -apply (rule iff_trans) -apply (rule sats_iterates_MH_fm [of A is_F], blast, simp_all) +apply (rule iff_sym) +apply (rule iff_trans) +apply (rule sats_iterates_MH_fm [of A is_F], blast, simp_all) done theorem iterates_MH_reflection: assumes p_reflection: - "!!f g h. REFLECTS[\x. p(L, f(x), g(x)), + "!!f g h. REFLECTS[\x. p(L, f(x), g(x)), \i x. p(**Lset(i), f(x), g(x))]" - shows "REFLECTS[\x. iterates_MH(L, p(L), e(x), f(x), g(x), h(x)), + shows "REFLECTS[\x. iterates_MH(L, p(L), e(x), f(x), g(x), h(x)), \i x. iterates_MH(**Lset(i), p(**Lset(i)), e(x), f(x), g(x), h(x))]" apply (simp (no_asm_use) only: iterates_MH_def) txt{*Must be careful: simplifying with @{text setclass_simps} above would change @{text "\gm[**Lset(i)]"} into @{text "\gm \ Lset(i)"}, when it would no longer match rule @{text is_nat_case_reflection}. *} -apply (rule is_nat_case_reflection) +apply (rule is_nat_case_reflection) apply (simp (no_asm_use) only: setclass_simps) apply (intro FOL_reflections function_reflections is_nat_case_reflection - restriction_reflection p_reflection) + restriction_reflection p_reflection) done -subsection{*@{term L} is Closed Under the Operator @{term list}*} +subsection{*@{term L} is Closed Under the Operator @{term list}*} subsubsection{*The List Functor, Internalized*} constdefs list_functor_fm :: "[i,i,i]=>i" -(* "is_list_functor(M,A,X,Z) == - \n1[M]. \AX[M]. +(* "is_list_functor(M,A,X,Z) == + \n1[M]. \AX[M]. number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)" *) - "list_functor_fm(A,X,Z) == + "list_functor_fm(A,X,Z) == Exists(Exists( - And(number1_fm(1), + And(number1_fm(1), And(cartprod_fm(A#+2,X#+2,0), sum_fm(1,0,Z#+2)))))" lemma list_functor_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> list_functor_fm(x,y,z) \ formula" -by (simp add: list_functor_fm_def) +by (simp add: list_functor_fm_def) lemma arity_list_functor_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(list_functor_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: list_functor_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: list_functor_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_list_functor_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, list_functor_fm(x,y,z), env) <-> + ==> sats(A, list_functor_fm(x,y,z), env) <-> is_list_functor(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: list_functor_fm_def is_list_functor_def) lemma list_functor_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> is_list_functor(**A, x, y, z) <-> sats(A, list_functor_fm(i,j,k), env)" by simp theorem list_functor_reflection: - "REFLECTS[\x. is_list_functor(L,f(x),g(x),h(x)), + "REFLECTS[\x. is_list_functor(L,f(x),g(x),h(x)), \i x. is_list_functor(**Lset(i),f(x),g(x),h(x))]" apply (simp only: is_list_functor_def setclass_simps) apply (intro FOL_reflections number1_reflection - cartprod_reflection sum_reflection) + cartprod_reflection sum_reflection) done @@ -793,29 +795,29 @@ [\x. \u[L]. u \ B \ (\y[L]. pair(L,u,y,x) \ is_wfrec(L, iterates_MH(L, is_list_functor(L,A), 0), memsn, u, y)), \i x. \u \ Lset(i). u \ B \ (\y \ Lset(i). pair(**Lset(i), u, y, x) \ - is_wfrec(**Lset(i), - iterates_MH(**Lset(i), + is_wfrec(**Lset(i), + iterates_MH(**Lset(i), is_list_functor(**Lset(i), A), 0), memsn, u, y))]" -by (intro FOL_reflections function_reflections is_wfrec_reflection - iterates_MH_reflection list_functor_reflection) +by (intro FOL_reflections function_reflections is_wfrec_reflection + iterates_MH_reflection list_functor_reflection) -lemma list_replacement1: +lemma list_replacement1: "L(A) ==> iterates_replacement(L, is_list_functor(L,A), 0)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) -apply (rule strong_replacementI) +apply (rule strong_replacementI) apply (rule rallI) -apply (rename_tac B) -apply (rule separation_CollectI) -apply (insert nonempty) -apply (subgoal_tac "L(Memrel(succ(n)))") -apply (rule_tac A="{B,A,n,z,0,Memrel(succ(n))}" in subset_LsetE, blast ) +apply (rename_tac B) +apply (rule separation_CollectI) +apply (insert nonempty) +apply (subgoal_tac "L(Memrel(succ(n)))") +apply (rule_tac A="{B,A,n,z,0,Memrel(succ(n))}" in subset_LsetE, blast ) apply (rule ReflectsE [OF list_replacement1_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2 Memrel_closed) -apply (elim conjE) +apply (elim conjE) apply (rule DPow_LsetI) -apply (rename_tac v) +apply (rename_tac v) apply (rule bex_iff_sats conj_iff_sats)+ apply (rule_tac env = "[u,v,A,n,B,0,Memrel(succ(n))]" in mem_iff_sats) apply (rule sep_rules | simp)+ @@ -832,34 +834,34 @@ is_wfrec (L, iterates_MH (L, is_list_functor(L, A), 0), msn, u, x)), \i x. \u \ Lset(i). u \ B \ u \ nat \ - (\sn \ Lset(i). \msn \ Lset(i). + (\sn \ Lset(i). \msn \ Lset(i). successor(**Lset(i), u, sn) \ membership(**Lset(i), sn, msn) \ - is_wfrec (**Lset(i), + is_wfrec (**Lset(i), iterates_MH (**Lset(i), is_list_functor(**Lset(i), A), 0), msn, u, x))]" -by (intro FOL_reflections function_reflections is_wfrec_reflection - iterates_MH_reflection list_functor_reflection) +by (intro FOL_reflections function_reflections is_wfrec_reflection + iterates_MH_reflection list_functor_reflection) -lemma list_replacement2: - "L(A) ==> strong_replacement(L, - \n y. n\nat & +lemma list_replacement2: + "L(A) ==> strong_replacement(L, + \n y. n\nat & (\sn[L]. \msn[L]. successor(L,n,sn) & membership(L,sn,msn) & - is_wfrec(L, iterates_MH(L,is_list_functor(L,A), 0), + is_wfrec(L, iterates_MH(L,is_list_functor(L,A), 0), msn, n, y)))" -apply (rule strong_replacementI) +apply (rule strong_replacementI) apply (rule rallI) -apply (rename_tac B) -apply (rule separation_CollectI) -apply (insert nonempty) -apply (rule_tac A="{A,B,z,0,nat}" in subset_LsetE) -apply (blast intro: L_nat) +apply (rename_tac B) +apply (rule separation_CollectI) +apply (insert nonempty) +apply (rule_tac A="{A,B,z,0,nat}" in subset_LsetE) +apply (blast intro: L_nat) apply (rule ReflectsE [OF list_replacement2_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac v) +apply (rename_tac v) apply (rule bex_iff_sats conj_iff_sats)+ apply (rule_tac env = "[u,v,A,B,0,nat]" in mem_iff_sats) apply (rule sep_rules | simp)+ @@ -868,21 +870,21 @@ done -subsection{*@{term L} is Closed Under the Operator @{term formula}*} +subsection{*@{term L} is Closed Under the Operator @{term formula}*} subsubsection{*The Formula Functor, Internalized*} constdefs formula_functor_fm :: "[i,i]=>i" -(* "is_formula_functor(M,X,Z) == - \nat'[M]. \natnat[M]. \natnatsum[M]. \XX[M]. \X3[M]. +(* "is_formula_functor(M,X,Z) == + \nat'[M]. \natnat[M]. \natnatsum[M]. \XX[M]. \X3[M]. 4 3 2 1 0 - omega(M,nat') & cartprod(M,nat',nat',natnat) & + omega(M,nat') & cartprod(M,nat',nat',natnat) & is_sum(M,natnat,natnat,natnatsum) & - cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) & - is_sum(M,natnatsum,X3,Z)" *) - "formula_functor_fm(X,Z) == + cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) & + is_sum(M,natnatsum,X3,Z)" *) + "formula_functor_fm(X,Z) == Exists(Exists(Exists(Exists(Exists( - And(omega_fm(4), + And(omega_fm(4), And(cartprod_fm(4,4,3), And(sum_fm(3,3,2), And(cartprod_fm(X#+5,X#+5,1), @@ -890,26 +892,26 @@ lemma formula_functor_type [TC]: "[| x \ nat; y \ nat |] ==> formula_functor_fm(x,y) \ formula" -by (simp add: formula_functor_fm_def) +by (simp add: formula_functor_fm_def) lemma sats_formula_functor_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, formula_functor_fm(x,y), env) <-> + ==> sats(A, formula_functor_fm(x,y), env) <-> is_formula_functor(**A, nth(x,env), nth(y,env))" by (simp add: formula_functor_fm_def is_formula_functor_def) lemma formula_functor_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; + "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> is_formula_functor(**A, x, y) <-> sats(A, formula_functor_fm(i,j), env)" by simp theorem formula_functor_reflection: - "REFLECTS[\x. is_formula_functor(L,f(x),g(x)), + "REFLECTS[\x. is_formula_functor(L,f(x),g(x)), \i x. is_formula_functor(**Lset(i),f(x),g(x))]" apply (simp only: is_formula_functor_def setclass_simps) apply (intro FOL_reflections omega_reflection - cartprod_reflection sum_reflection) + cartprod_reflection sum_reflection) done subsubsection{*Instances of Replacement for Formulas*} @@ -919,28 +921,28 @@ [\x. \u[L]. u \ B \ (\y[L]. pair(L,u,y,x) \ is_wfrec(L, iterates_MH(L, is_formula_functor(L), 0), memsn, u, y)), \i x. \u \ Lset(i). u \ B \ (\y \ Lset(i). pair(**Lset(i), u, y, x) \ - is_wfrec(**Lset(i), - iterates_MH(**Lset(i), + is_wfrec(**Lset(i), + iterates_MH(**Lset(i), is_formula_functor(**Lset(i)), 0), memsn, u, y))]" -by (intro FOL_reflections function_reflections is_wfrec_reflection - iterates_MH_reflection formula_functor_reflection) +by (intro FOL_reflections function_reflections is_wfrec_reflection + iterates_MH_reflection formula_functor_reflection) -lemma formula_replacement1: +lemma formula_replacement1: "iterates_replacement(L, is_formula_functor(L), 0)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) -apply (rule strong_replacementI) +apply (rule strong_replacementI) apply (rule rallI) -apply (rename_tac B) -apply (rule separation_CollectI) -apply (insert nonempty) -apply (subgoal_tac "L(Memrel(succ(n)))") -apply (rule_tac A="{B,n,z,0,Memrel(succ(n))}" in subset_LsetE, blast ) +apply (rename_tac B) +apply (rule separation_CollectI) +apply (insert nonempty) +apply (subgoal_tac "L(Memrel(succ(n)))") +apply (rule_tac A="{B,n,z,0,Memrel(succ(n))}" in subset_LsetE, blast ) apply (rule ReflectsE [OF formula_replacement1_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2 Memrel_closed) apply (rule DPow_LsetI) -apply (rename_tac v) +apply (rename_tac v) apply (rule bex_iff_sats conj_iff_sats)+ apply (rule_tac env = "[u,v,n,B,0,Memrel(succ(n))]" in mem_iff_sats) apply (rule sep_rules | simp)+ @@ -957,34 +959,34 @@ is_wfrec (L, iterates_MH (L, is_formula_functor(L), 0), msn, u, x)), \i x. \u \ Lset(i). u \ B \ u \ nat \ - (\sn \ Lset(i). \msn \ Lset(i). + (\sn \ Lset(i). \msn \ Lset(i). successor(**Lset(i), u, sn) \ membership(**Lset(i), sn, msn) \ - is_wfrec (**Lset(i), + is_wfrec (**Lset(i), iterates_MH (**Lset(i), is_formula_functor(**Lset(i)), 0), msn, u, x))]" -by (intro FOL_reflections function_reflections is_wfrec_reflection - iterates_MH_reflection formula_functor_reflection) +by (intro FOL_reflections function_reflections is_wfrec_reflection + iterates_MH_reflection formula_functor_reflection) -lemma formula_replacement2: - "strong_replacement(L, - \n y. n\nat & +lemma formula_replacement2: + "strong_replacement(L, + \n y. n\nat & (\sn[L]. \msn[L]. successor(L,n,sn) & membership(L,sn,msn) & - is_wfrec(L, iterates_MH(L,is_formula_functor(L), 0), + is_wfrec(L, iterates_MH(L,is_formula_functor(L), 0), msn, n, y)))" -apply (rule strong_replacementI) +apply (rule strong_replacementI) apply (rule rallI) -apply (rename_tac B) -apply (rule separation_CollectI) -apply (insert nonempty) -apply (rule_tac A="{B,z,0,nat}" in subset_LsetE) -apply (blast intro: L_nat) +apply (rename_tac B) +apply (rule separation_CollectI) +apply (insert nonempty) +apply (rule_tac A="{B,z,0,nat}" in subset_LsetE) +apply (blast intro: L_nat) apply (rule ReflectsE [OF formula_replacement2_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac v) +apply (rename_tac v) apply (rule bex_iff_sats conj_iff_sats)+ apply (rule_tac env = "[u,v,B,0,nat]" in mem_iff_sats) apply (rule sep_rules | simp)+ @@ -1006,7 +1008,7 @@ lemma Inl_type [TC]: "[| x \ nat; z \ nat |] ==> Inl_fm(x,z) \ formula" -by (simp add: Inl_fm_def) +by (simp add: Inl_fm_def) lemma sats_Inl_fm [simp]: "[| x \ nat; z \ nat; env \ list(A)|] @@ -1014,16 +1016,16 @@ by (simp add: Inl_fm_def is_Inl_def) lemma Inl_iff_sats: - "[| nth(i,env) = x; nth(k,env) = z; + "[| nth(i,env) = x; nth(k,env) = z; i \ nat; k \ nat; env \ list(A)|] ==> is_Inl(**A, x, z) <-> sats(A, Inl_fm(i,k), env)" by simp theorem Inl_reflection: - "REFLECTS[\x. is_Inl(L,f(x),h(x)), + "REFLECTS[\x. is_Inl(L,f(x),h(x)), \i x. is_Inl(**Lset(i),f(x),h(x))]" apply (simp only: is_Inl_def setclass_simps) -apply (intro FOL_reflections function_reflections) +apply (intro FOL_reflections function_reflections) done @@ -1035,7 +1037,7 @@ lemma Inr_type [TC]: "[| x \ nat; z \ nat |] ==> Inr_fm(x,z) \ formula" -by (simp add: Inr_fm_def) +by (simp add: Inr_fm_def) lemma sats_Inr_fm [simp]: "[| x \ nat; z \ nat; env \ list(A)|] @@ -1043,16 +1045,16 @@ by (simp add: Inr_fm_def is_Inr_def) lemma Inr_iff_sats: - "[| nth(i,env) = x; nth(k,env) = z; + "[| nth(i,env) = x; nth(k,env) = z; i \ nat; k \ nat; env \ list(A)|] ==> is_Inr(**A, x, z) <-> sats(A, Inr_fm(i,k), env)" by simp theorem Inr_reflection: - "REFLECTS[\x. is_Inr(L,f(x),h(x)), + "REFLECTS[\x. is_Inr(L,f(x),h(x)), \i x. is_Inr(**Lset(i),f(x),h(x))]" apply (simp only: is_Inr_def setclass_simps) -apply (intro FOL_reflections function_reflections) +apply (intro FOL_reflections function_reflections) done @@ -1062,9 +1064,9 @@ constdefs Nil_fm :: "i=>i" "Nil_fm(x) == Exists(And(empty_fm(0), Inl_fm(0,succ(x))))" - + lemma Nil_type [TC]: "x \ nat ==> Nil_fm(x) \ formula" -by (simp add: Nil_fm_def) +by (simp add: Nil_fm_def) lemma sats_Nil_fm [simp]: "[| x \ nat; env \ list(A)|] @@ -1077,10 +1079,10 @@ by simp theorem Nil_reflection: - "REFLECTS[\x. is_Nil(L,f(x)), + "REFLECTS[\x. is_Nil(L,f(x)), \i x. is_Nil(**Lset(i),f(x))]" apply (simp only: is_Nil_def setclass_simps) -apply (intro FOL_reflections function_reflections Inl_reflection) +apply (intro FOL_reflections function_reflections Inl_reflection) done @@ -1089,30 +1091,30 @@ (* "is_Cons(M,a,l,Z) == \p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" *) constdefs Cons_fm :: "[i,i,i]=>i" - "Cons_fm(a,l,Z) == + "Cons_fm(a,l,Z) == Exists(And(pair_fm(succ(a),succ(l),0), Inr_fm(0,succ(Z))))" lemma Cons_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> Cons_fm(x,y,z) \ formula" -by (simp add: Cons_fm_def) +by (simp add: Cons_fm_def) lemma sats_Cons_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, Cons_fm(x,y,z), env) <-> + ==> sats(A, Cons_fm(x,y,z), env) <-> is_Cons(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: Cons_fm_def is_Cons_def) lemma Cons_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==>is_Cons(**A, x, y, z) <-> sats(A, Cons_fm(i,j,k), env)" by simp theorem Cons_reflection: - "REFLECTS[\x. is_Cons(L,f(x),g(x),h(x)), + "REFLECTS[\x. is_Cons(L,f(x),g(x),h(x)), \i x. is_Cons(**Lset(i),f(x),g(x),h(x))]" apply (simp only: is_Cons_def setclass_simps) -apply (intro FOL_reflections pair_reflection Inr_reflection) +apply (intro FOL_reflections pair_reflection Inr_reflection) done subsubsection{*The Formula @{term is_quasilist}, Internalized*} @@ -1120,11 +1122,11 @@ (* is_quasilist(M,xs) == is_Nil(M,z) | (\x[M]. \l[M]. is_Cons(M,x,l,z))" *) constdefs quasilist_fm :: "i=>i" - "quasilist_fm(x) == + "quasilist_fm(x) == Or(Nil_fm(x), Exists(Exists(Cons_fm(1,0,succ(succ(x))))))" - + lemma quasilist_type [TC]: "x \ nat ==> quasilist_fm(x) \ formula" -by (simp add: quasilist_fm_def) +by (simp add: quasilist_fm_def) lemma sats_quasilist_fm [simp]: "[| x \ nat; env \ list(A)|] @@ -1137,10 +1139,10 @@ by simp theorem quasilist_reflection: - "REFLECTS[\x. is_quasilist(L,f(x)), + "REFLECTS[\x. is_quasilist(L,f(x)), \i x. is_quasilist(**Lset(i),f(x))]" apply (simp only: is_quasilist_def setclass_simps) -apply (intro FOL_reflections Nil_reflection Cons_reflection) +apply (intro FOL_reflections Nil_reflection Cons_reflection) done @@ -1149,19 +1151,19 @@ subsubsection{*The Formula @{term is_tl}, Internalized*} -(* "is_tl(M,xs,T) == +(* "is_tl(M,xs,T) == (is_Nil(M,xs) --> T=xs) & (\x[M]. \l[M]. ~ is_Cons(M,x,l,xs) | T=l) & (is_quasilist(M,xs) | empty(M,T))" *) constdefs tl_fm :: "[i,i]=>i" - "tl_fm(xs,T) == + "tl_fm(xs,T) == And(Implies(Nil_fm(xs), Equal(T,xs)), And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(T#+2,0)))), Or(quasilist_fm(xs), empty_fm(T))))" lemma tl_type [TC]: "[| x \ nat; y \ nat |] ==> tl_fm(x,y) \ formula" -by (simp add: tl_fm_def) +by (simp add: tl_fm_def) lemma sats_tl_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] @@ -1175,11 +1177,11 @@ by simp theorem tl_reflection: - "REFLECTS[\x. is_tl(L,f(x),g(x)), + "REFLECTS[\x. is_tl(L,f(x),g(x)), \i x. is_tl(**Lset(i),f(x),g(x))]" apply (simp only: is_tl_def setclass_simps) apply (intro FOL_reflections Nil_reflection Cons_reflection - quasilist_reflection empty_reflection) + quasilist_reflection empty_reflection) done @@ -1190,27 +1192,27 @@ [\x. \u[L]. u \ B \ (\y[L]. pair(L,u,y,x) \ is_wfrec(L, iterates_MH(L, is_tl(L), z), memsn, u, y)), \i x. \u \ Lset(i). u \ B \ (\y \ Lset(i). pair(**Lset(i), u, y, x) \ - is_wfrec(**Lset(i), - iterates_MH(**Lset(i), + is_wfrec(**Lset(i), + iterates_MH(**Lset(i), is_tl(**Lset(i)), z), memsn, u, y))]" -by (intro FOL_reflections function_reflections is_wfrec_reflection - iterates_MH_reflection list_functor_reflection tl_reflection) +by (intro FOL_reflections function_reflections is_wfrec_reflection + iterates_MH_reflection list_functor_reflection tl_reflection) -lemma nth_replacement: +lemma nth_replacement: "L(w) ==> iterates_replacement(L, %l t. is_tl(L,l,t), w)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) -apply (rule strong_replacementI) -apply (rule rallI) -apply (rule separation_CollectI) -apply (subgoal_tac "L(Memrel(succ(n)))") -apply (rule_tac A="{A,n,w,z,Memrel(succ(n))}" in subset_LsetE, blast ) +apply (rule strong_replacementI) +apply (rule rallI) +apply (rule separation_CollectI) +apply (subgoal_tac "L(Memrel(succ(n)))") +apply (rule_tac A="{A,n,w,z,Memrel(succ(n))}" in subset_LsetE, blast ) apply (rule ReflectsE [OF nth_replacement_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2 Memrel_closed) -apply (elim conjE) +apply (elim conjE) apply (rule DPow_LsetI) -apply (rename_tac v) +apply (rename_tac v) apply (rule bex_iff_sats conj_iff_sats)+ apply (rule_tac env = "[u,v,A,z,w,Memrel(succ(n))]" in mem_iff_sats) apply (rule sep_rules | simp)+ @@ -1221,27 +1223,29 @@ subsubsection{*Instantiating the locale @{text M_datatypes}*} -ML -{* -val list_replacement1 = thm "list_replacement1"; -val list_replacement2 = thm "list_replacement2"; -val formula_replacement1 = thm "formula_replacement1"; -val formula_replacement2 = thm "formula_replacement2"; -val nth_replacement = thm "nth_replacement"; + +theorem M_datatypes_axioms_L: "M_datatypes_axioms(L)" + apply (rule M_datatypes_axioms.intro) + apply (assumption | rule + list_replacement1 list_replacement2 + formula_replacement1 formula_replacement2 + nth_replacement)+ + done -val m_datatypes = [list_replacement1, list_replacement2, - formula_replacement1, formula_replacement2, - nth_replacement]; - -fun datatypes_L th = - kill_flex_triv_prems (m_datatypes MRS (wfrank_L th)); +theorem M_datatypes_L: "PROP M_datatypes(L)" + apply (rule M_datatypes.intro) + apply (rule M_triv_axioms_L) + apply (rule M_axioms_axioms_L) + apply (rule M_trancl_axioms_L) + apply (rule M_wfrank_axioms_L) + apply (rule M_datatypes_axioms_L) + done -bind_thm ("list_closed", datatypes_L (thm "M_datatypes.list_closed")); -bind_thm ("formula_closed", datatypes_L (thm "M_datatypes.formula_closed")); -bind_thm ("list_abs", datatypes_L (thm "M_datatypes.list_abs")); -bind_thm ("formula_abs", datatypes_L (thm "M_datatypes.formula_abs")); -bind_thm ("nth_abs", datatypes_L (thm "M_datatypes.nth_abs")); -*} +lemmas list_closed = M_datatypes.list_closed [OF M_datatypes_L] + and formula_closed = M_datatypes.formula_closed [OF M_datatypes_L] + and list_abs = M_datatypes.list_abs [OF M_datatypes_L] + and formula_abs = M_datatypes.formula_abs [OF M_datatypes_L] + and nth_abs = M_datatypes.nth_abs [OF M_datatypes_L] declare list_closed [intro,simp] declare formula_closed [intro,simp] @@ -1250,8 +1254,7 @@ declare nth_abs [simp] - -subsection{*@{term L} is Closed Under the Operator @{term eclose}*} +subsection{*@{term L} is Closed Under the Operator @{term eclose}*} subsubsection{*Instances of Replacement for @{term eclose}*} @@ -1260,28 +1263,28 @@ [\x. \u[L]. u \ B \ (\y[L]. pair(L,u,y,x) \ is_wfrec(L, iterates_MH(L, big_union(L), A), memsn, u, y)), \i x. \u \ Lset(i). u \ B \ (\y \ Lset(i). pair(**Lset(i), u, y, x) \ - is_wfrec(**Lset(i), - iterates_MH(**Lset(i), big_union(**Lset(i)), A), + is_wfrec(**Lset(i), + iterates_MH(**Lset(i), big_union(**Lset(i)), A), memsn, u, y))]" -by (intro FOL_reflections function_reflections is_wfrec_reflection - iterates_MH_reflection) +by (intro FOL_reflections function_reflections is_wfrec_reflection + iterates_MH_reflection) -lemma eclose_replacement1: +lemma eclose_replacement1: "L(A) ==> iterates_replacement(L, big_union(L), A)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) -apply (rule strong_replacementI) +apply (rule strong_replacementI) apply (rule rallI) -apply (rename_tac B) -apply (rule separation_CollectI) -apply (subgoal_tac "L(Memrel(succ(n)))") -apply (rule_tac A="{B,A,n,z,Memrel(succ(n))}" in subset_LsetE, blast ) +apply (rename_tac B) +apply (rule separation_CollectI) +apply (subgoal_tac "L(Memrel(succ(n)))") +apply (rule_tac A="{B,A,n,z,Memrel(succ(n))}" in subset_LsetE, blast ) apply (rule ReflectsE [OF eclose_replacement1_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2 Memrel_closed) -apply (elim conjE) +apply (elim conjE) apply (rule DPow_LsetI) -apply (rename_tac v) +apply (rename_tac v) apply (rule bex_iff_sats conj_iff_sats)+ apply (rule_tac env = "[u,v,A,n,B,Memrel(succ(n))]" in mem_iff_sats) apply (rule sep_rules | simp)+ @@ -1298,33 +1301,33 @@ is_wfrec (L, iterates_MH (L, big_union(L), A), msn, u, x)), \i x. \u \ Lset(i). u \ B \ u \ nat \ - (\sn \ Lset(i). \msn \ Lset(i). + (\sn \ Lset(i). \msn \ Lset(i). successor(**Lset(i), u, sn) \ membership(**Lset(i), sn, msn) \ - is_wfrec (**Lset(i), + is_wfrec (**Lset(i), iterates_MH (**Lset(i), big_union(**Lset(i)), A), msn, u, x))]" -by (intro FOL_reflections function_reflections is_wfrec_reflection - iterates_MH_reflection) +by (intro FOL_reflections function_reflections is_wfrec_reflection + iterates_MH_reflection) -lemma eclose_replacement2: - "L(A) ==> strong_replacement(L, - \n y. n\nat & +lemma eclose_replacement2: + "L(A) ==> strong_replacement(L, + \n y. n\nat & (\sn[L]. \msn[L]. successor(L,n,sn) & membership(L,sn,msn) & - is_wfrec(L, iterates_MH(L,big_union(L), A), + is_wfrec(L, iterates_MH(L,big_union(L), A), msn, n, y)))" -apply (rule strong_replacementI) +apply (rule strong_replacementI) apply (rule rallI) -apply (rename_tac B) -apply (rule separation_CollectI) -apply (rule_tac A="{A,B,z,nat}" in subset_LsetE) -apply (blast intro: L_nat) +apply (rename_tac B) +apply (rule separation_CollectI) +apply (rule_tac A="{A,B,z,nat}" in subset_LsetE) +apply (blast intro: L_nat) apply (rule ReflectsE [OF eclose_replacement2_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac v) +apply (rename_tac v) apply (rule bex_iff_sats conj_iff_sats)+ apply (rule_tac env = "[u,v,A,B,nat]" in mem_iff_sats) apply (rule sep_rules | simp)+ @@ -1334,22 +1337,23 @@ subsubsection{*Instantiating the locale @{text M_eclose}*} -ML -{* -val eclose_replacement1 = thm "eclose_replacement1"; -val eclose_replacement2 = thm "eclose_replacement2"; -val m_eclose = [eclose_replacement1, eclose_replacement2]; +theorem M_eclose_axioms_L: "M_eclose_axioms(L)" + apply (rule M_eclose_axioms.intro) + apply (assumption | rule eclose_replacement1 eclose_replacement2)+ + done -fun eclose_L th = - kill_flex_triv_prems (m_eclose MRS (datatypes_L th)); +theorem M_eclose_L: "PROP M_eclose(L)" + apply (rule M_eclose.intro) + apply (rule M_triv_axioms_L) + apply (rule M_axioms_axioms_L) + apply (rule M_trancl_axioms_L) + apply (rule M_wfrank_axioms_L) + apply (rule M_datatypes_axioms_L) + apply (rule M_eclose_axioms_L) + done -bind_thm ("eclose_closed", eclose_L (thm "M_eclose.eclose_closed")); -bind_thm ("eclose_abs", eclose_L (thm "M_eclose.eclose_abs")); -*} - -declare eclose_closed [intro,simp] -declare eclose_abs [intro,simp] - +lemmas eclose_closed [intro, simp] = M_eclose.eclose_closed [OF M_eclose_L] + and eclose_abs [intro, simp] = M_eclose.eclose_abs [OF M_eclose_L] end diff -r b429fd98549c -r 99e52e78eb65 src/ZF/Constructible/Reflection.thy --- a/src/ZF/Constructible/Reflection.thy Sun Jul 28 21:09:37 2002 +0200 +++ b/src/ZF/Constructible/Reflection.thy Mon Jul 29 00:57:16 2002 +0200 @@ -25,7 +25,7 @@ ultimately the @{text ex_reflection} proof is packaged up using the predicate @{text Reflects}. *} -locale (open) reflection = +locale reflection = fixes Mset and M and Reflects assumes Mset_mono_le : "mono_le_subset(Mset)" and Mset_cont : "cont_Ord(Mset)" @@ -124,7 +124,7 @@ text{*Locale for the induction hypothesis*} -locale (open) ex_reflection = reflection + +locale ex_reflection = reflection + fixes P --"the original formula" fixes Q --"the reflected formula" fixes Cl --"the class of reflecting ordinals" @@ -170,16 +170,19 @@ !!a. [| Cl(a); Ord(a) |] ==> \x\Mset(a). P(x) <-> Q(a,x) |] ==> (\z. M(z) \ P()) <-> (\z\Mset(a). Q(a,))" apply (unfold ClEx_def FF_def F0_def M_def) -apply (rule Reflection.ZF_ClEx_iff [of Mset Cl], - simp_all add: Mset_mono_le Mset_cont Pair_in_Mset) +apply (rule ex_reflection.ZF_ClEx_iff + [OF ex_reflection.intro, OF reflection.intro ex_reflection_axioms.intro, + of Mset Cl]) +apply (simp_all add: Mset_mono_le Mset_cont Pair_in_Mset) done lemma (in reflection) Closed_Unbounded_ClEx: "(!!a. [| Cl(a); Ord(a) |] ==> \x\Mset(a). P(x) <-> Q(a,x)) ==> Closed_Unbounded(ClEx(P))" apply (unfold ClEx_def FF_def F0_def M_def) -apply (rule Reflection.ZF_Closed_Unbounded_ClEx, - simp_all add: Mset_mono_le Mset_cont Pair_in_Mset) +apply (rule ex_reflection.ZF_Closed_Unbounded_ClEx + [OF ex_reflection.intro, OF reflection.intro ex_reflection_axioms.intro]) +apply (simp_all add: Mset_mono_le Mset_cont Pair_in_Mset) done subsection{*Packaging the Quantifier Reflection Rules*} diff -r b429fd98549c -r 99e52e78eb65 src/ZF/Constructible/Relative.thy --- a/src/ZF/Constructible/Relative.thy Sun Jul 28 21:09:37 2002 +0200 +++ b/src/ZF/Constructible/Relative.thy Mon Jul 29 00:57:16 2002 +0200 @@ -448,7 +448,7 @@ text{*The class M is assumed to be transitive and to satisfy some relativized ZF axioms*} -locale (open) M_triv_axioms = +locale M_triv_axioms = fixes M assumes transM: "[| y\x; M(x) |] ==> M(y)" and nonempty [simp]: "M(0)" @@ -821,7 +821,7 @@ "M(a) ==> number1(M,a) <-> a = 1" by (simp add: number1_def) -lemma (in M_triv_axioms) number1_abs [simp]: +lemma (in M_triv_axioms) number2_abs [simp]: "M(a) ==> number2(M,a) <-> a = succ(1)" by (simp add: number2_def) @@ -854,7 +854,7 @@ subsection{*Some instances of separation and strong replacement*} -locale (open) M_axioms = M_triv_axioms + +locale M_axioms = M_triv_axioms + assumes Inter_separation: "M(A) ==> separation(M, \x. \y[M]. y\A --> x\y)" and cartprod_separation: diff -r b429fd98549c -r 99e52e78eb65 src/ZF/Constructible/Separation.thy --- a/src/ZF/Constructible/Separation.thy Sun Jul 28 21:09:37 2002 +0200 +++ b/src/ZF/Constructible/Separation.thy Mon Jul 29 00:57:16 2002 +0200 @@ -9,39 +9,39 @@ by simp lemmas nth_rules = nth_0 nth_ConsI nat_0I nat_succI -lemmas sep_rules = nth_0 nth_ConsI FOL_iff_sats function_iff_sats +lemmas sep_rules = nth_0 nth_ConsI FOL_iff_sats function_iff_sats fun_plus_iff_sats lemma Collect_conj_in_DPow: - "[| {x\A. P(x)} \ DPow(A); {x\A. Q(x)} \ DPow(A) |] + "[| {x\A. P(x)} \ DPow(A); {x\A. Q(x)} \ DPow(A) |] ==> {x\A. P(x) & Q(x)} \ DPow(A)" -by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric]) +by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric]) lemma Collect_conj_in_DPow_Lset: "[|z \ Lset(j); {x \ Lset(j). P(x)} \ DPow(Lset(j))|] ==> {x \ Lset(j). x \ z & P(x)} \ DPow(Lset(j))" apply (frule mem_Lset_imp_subset_Lset) -apply (simp add: Collect_conj_in_DPow Collect_mem_eq +apply (simp add: Collect_conj_in_DPow Collect_mem_eq subset_Int_iff2 elem_subset_in_DPow) done lemma separation_CollectI: "(\z. L(z) ==> L({x \ z . P(x)})) ==> separation(L, \x. P(x))" -apply (unfold separation_def, clarify) -apply (rule_tac x="{x\z. P(x)}" in rexI) +apply (unfold separation_def, clarify) +apply (rule_tac x="{x\z. P(x)}" in rexI) apply simp_all done text{*Reduces the original comprehension to the reflected one*} lemma reflection_imp_L_separation: "[| \x\Lset(j). P(x) <-> Q(x); - {x \ Lset(j) . Q(x)} \ DPow(Lset(j)); + {x \ Lset(j) . Q(x)} \ DPow(Lset(j)); Ord(j); z \ Lset(j)|] ==> L({x \ z . P(x)})" apply (rule_tac i = "succ(j)" in L_I) prefer 2 apply simp apply (subgoal_tac "{x \ z. P(x)} = {x \ Lset(j). x \ z & (Q(x))}") prefer 2 - apply (blast dest: mem_Lset_imp_subset_Lset) + apply (blast dest: mem_Lset_imp_subset_Lset) apply (simp add: Lset_succ Collect_conj_in_DPow_Lset) done @@ -49,20 +49,20 @@ subsection{*Separation for Intersection*} lemma Inter_Reflects: - "REFLECTS[\x. \y[L]. y\A --> x \ y, + "REFLECTS[\x. \y[L]. y\A --> x \ y, \i x. \y\Lset(i). y\A --> x \ y]" -by (intro FOL_reflections) +by (intro FOL_reflections) lemma Inter_separation: "L(A) ==> separation(L, \x. \y[L]. y\A --> x\y)" -apply (rule separation_CollectI) -apply (rule_tac A="{A,z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{A,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF Inter_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2, clarify) -apply (rule DPow_LsetI) -apply (rule ball_iff_sats) +apply (rule DPow_LsetI) +apply (rule ball_iff_sats) apply (rule imp_iff_sats) apply (rule_tac [2] i=1 and j=0 and env="[y,x,A]" in mem_iff_sats) apply (rule_tac i=0 and j=2 in mem_iff_sats) @@ -73,22 +73,22 @@ lemma cartprod_Reflects: "REFLECTS[\z. \x[L]. x\A & (\y[L]. y\B & pair(L,x,y,z)), - \i z. \x\Lset(i). x\A & (\y\Lset(i). y\B & + \i z. \x\Lset(i). x\A & (\y\Lset(i). y\B & pair(**Lset(i),x,y,z))]" by (intro FOL_reflections function_reflections) lemma cartprod_separation: - "[| L(A); L(B) |] + "[| L(A); L(B) |] ==> separation(L, \z. \x[L]. x\A & (\y[L]. y\B & pair(L,x,y,z)))" -apply (rule separation_CollectI) -apply (rule_tac A="{A,B,z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{A,B,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF cartprod_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) - apply (simp_all add: lt_Ord2, clarify) + apply (simp_all add: lt_Ord2, clarify) apply (rule DPow_LsetI) -apply (rename_tac u) -apply (rule bex_iff_sats) +apply (rename_tac u) +apply (rule bex_iff_sats) apply (rule conj_iff_sats) apply (rule_tac i=0 and j=2 and env="[x,u,A,B]" in mem_iff_sats, simp_all) apply (rule sep_rules | simp)+ @@ -102,16 +102,16 @@ by (intro FOL_reflections function_reflections) lemma image_separation: - "[| L(A); L(r) |] + "[| L(A); L(r) |] ==> separation(L, \y. \p[L]. p\r & (\x[L]. x\A & pair(L,x,y,p)))" -apply (rule separation_CollectI) -apply (rule_tac A="{A,r,z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{A,r,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF image_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2, clarify) apply (rule DPow_LsetI) -apply (rule bex_iff_sats) +apply (rule bex_iff_sats) apply (rule conj_iff_sats) apply (rule_tac env="[p,y,A,r]" in mem_iff_sats) apply (rule sep_rules | simp)+ @@ -122,22 +122,22 @@ lemma converse_Reflects: "REFLECTS[\z. \p[L]. p\r & (\x[L]. \y[L]. pair(L,x,y,p) & pair(L,y,x,z)), - \i z. \p\Lset(i). p\r & (\x\Lset(i). \y\Lset(i). + \i z. \p\Lset(i). p\r & (\x\Lset(i). \y\Lset(i). pair(**Lset(i),x,y,p) & pair(**Lset(i),y,x,z))]" by (intro FOL_reflections function_reflections) lemma converse_separation: - "L(r) ==> separation(L, + "L(r) ==> separation(L, \z. \p[L]. p\r & (\x[L]. \y[L]. pair(L,x,y,p) & pair(L,y,x,z)))" -apply (rule separation_CollectI) -apply (rule_tac A="{r,z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{r,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF converse_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2, clarify) apply (rule DPow_LsetI) -apply (rename_tac u) -apply (rule bex_iff_sats) +apply (rename_tac u) +apply (rule bex_iff_sats) apply (rule conj_iff_sats) apply (rule_tac i=0 and j=2 and env="[p,u,r]" in mem_iff_sats, simp_all) apply (rule sep_rules | simp)+ @@ -153,15 +153,15 @@ lemma restrict_separation: "L(A) ==> separation(L, \z. \x[L]. x\A & (\y[L]. pair(L,x,y,z)))" -apply (rule separation_CollectI) -apply (rule_tac A="{A,z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{A,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF restrict_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2, clarify) apply (rule DPow_LsetI) -apply (rename_tac u) -apply (rule bex_iff_sats) +apply (rename_tac u) +apply (rule bex_iff_sats) apply (rule conj_iff_sats) apply (rule_tac i=0 and j=2 and env="[x,u,A]" in mem_iff_sats, simp_all) apply (rule sep_rules | simp)+ @@ -171,29 +171,29 @@ subsection{*Separation for Composition*} lemma comp_Reflects: - "REFLECTS[\xz. \x[L]. \y[L]. \z[L]. \xy[L]. \yz[L]. - pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & + "REFLECTS[\xz. \x[L]. \y[L]. \z[L]. \xy[L]. \yz[L]. + pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & xy\s & yz\r, - \i xz. \x\Lset(i). \y\Lset(i). \z\Lset(i). \xy\Lset(i). \yz\Lset(i). - pair(**Lset(i),x,z,xz) & pair(**Lset(i),x,y,xy) & + \i xz. \x\Lset(i). \y\Lset(i). \z\Lset(i). \xy\Lset(i). \yz\Lset(i). + pair(**Lset(i),x,z,xz) & pair(**Lset(i),x,y,xy) & pair(**Lset(i),y,z,yz) & xy\s & yz\r]" by (intro FOL_reflections function_reflections) lemma comp_separation: "[| L(r); L(s) |] - ==> separation(L, \xz. \x[L]. \y[L]. \z[L]. \xy[L]. \yz[L]. - pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & + ==> separation(L, \xz. \x[L]. \y[L]. \z[L]. \xy[L]. \yz[L]. + pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & xy\s & yz\r)" -apply (rule separation_CollectI) -apply (rule_tac A="{r,s,z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{r,s,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF comp_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2, clarify) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule bex_iff_sats)+ -apply (rename_tac x y z) +apply (rename_tac x y z) apply (rule conj_iff_sats) apply (rule_tac env="[z,y,x,u,r,s]" in pair_iff_sats) apply (rule sep_rules | simp)+ @@ -208,17 +208,17 @@ lemma pred_separation: "[| L(r); L(x) |] ==> separation(L, \y. \p[L]. p\r & pair(L,y,x,p))" -apply (rule separation_CollectI) -apply (rule_tac A="{r,x,z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{r,x,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF pred_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2, clarify) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule bex_iff_sats) apply (rule conj_iff_sats) -apply (rule_tac env = "[p,u,r,x]" in mem_iff_sats) +apply (rule_tac env = "[p,u,r,x]" in mem_iff_sats) apply (rule sep_rules | simp)+ done @@ -232,50 +232,50 @@ lemma Memrel_separation: "separation(L, \z. \x[L]. \y[L]. pair(L,x,y,z) & x \ y)" -apply (rule separation_CollectI) -apply (rule_tac A="{z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF Memrel_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule bex_iff_sats conj_iff_sats)+ -apply (rule_tac env = "[y,x,u]" in pair_iff_sats) +apply (rule_tac env = "[y,x,u]" in pair_iff_sats) apply (rule sep_rules | simp)+ done subsection{*Replacement for FunSpace*} - + lemma funspace_succ_Reflects: - "REFLECTS[\z. \p[L]. p\A & (\f[L]. \b[L]. \nb[L]. \cnbf[L]. - pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) & - upair(L,cnbf,cnbf,z)), - \i z. \p \ Lset(i). p\A & (\f \ Lset(i). \b \ Lset(i). - \nb \ Lset(i). \cnbf \ Lset(i). - pair(**Lset(i),f,b,p) & pair(**Lset(i),n,b,nb) & - is_cons(**Lset(i),nb,f,cnbf) & upair(**Lset(i),cnbf,cnbf,z))]" + "REFLECTS[\z. \p[L]. p\A & (\f[L]. \b[L]. \nb[L]. \cnbf[L]. + pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) & + upair(L,cnbf,cnbf,z)), + \i z. \p \ Lset(i). p\A & (\f \ Lset(i). \b \ Lset(i). + \nb \ Lset(i). \cnbf \ Lset(i). + pair(**Lset(i),f,b,p) & pair(**Lset(i),n,b,nb) & + is_cons(**Lset(i),nb,f,cnbf) & upair(**Lset(i),cnbf,cnbf,z))]" by (intro FOL_reflections function_reflections) lemma funspace_succ_replacement: - "L(n) ==> - strong_replacement(L, \p z. \f[L]. \b[L]. \nb[L]. \cnbf[L]. + "L(n) ==> + strong_replacement(L, \p z. \f[L]. \b[L]. \nb[L]. \cnbf[L]. pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) & upair(L,cnbf,cnbf,z))" -apply (rule strong_replacementI) -apply (rule rallI) -apply (rule separation_CollectI) -apply (rule_tac A="{n,A,z}" in subset_LsetE, blast ) +apply (rule strong_replacementI) +apply (rule rallI) +apply (rule separation_CollectI) +apply (rule_tac A="{n,A,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF funspace_succ_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule bex_iff_sats) apply (rule conj_iff_sats) -apply (rule_tac env = "[p,u,n,A]" in mem_iff_sats) +apply (rule_tac env = "[p,u,n,A]" in mem_iff_sats) apply (rule sep_rules | simp)+ done @@ -283,26 +283,26 @@ subsection{*Separation for Order-Isomorphisms*} lemma well_ord_iso_Reflects: - "REFLECTS[\x. x\A --> + "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). + \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 separation_CollectI) -apply (rule_tac A="{A,f,r,z}" in subset_LsetE, blast ) + "[| 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 separation_CollectI) +apply (rule_tac A="{A,f,r,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF well_ord_iso_Reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule imp_iff_sats) -apply (rule_tac env = "[u,A,f,r]" in mem_iff_sats) +apply (rule_tac env = "[u,A,f,r]" in mem_iff_sats) apply (rule sep_rules | simp)+ done @@ -310,31 +310,31 @@ subsection{*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)]" + "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 separation_CollectI) -apply (rule_tac A="{A,r,z}" in subset_LsetE, blast ) + "[| 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 separation_CollectI) +apply (rule_tac A="{A,r,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF obase_reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule bex_iff_sats) apply (rule conj_iff_sats) -apply (rule_tac env = "[x,u,A,r]" in ordinal_iff_sats) +apply (rule_tac env = "[x,u,A,r]" in ordinal_iff_sats) apply (rule sep_rules | simp)+ done @@ -342,33 +342,33 @@ subsection{*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)))]" + "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 separation_CollectI) -apply (rule_tac A="{A,r,z}" in subset_LsetE, blast ) + "[| 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 separation_CollectI) +apply (rule_tac A="{A,r,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF obase_equals_reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule imp_iff_sats ball_iff_sats disj_iff_sats not_iff_sats)+ -apply (rule_tac env = "[u,A,r]" in mem_iff_sats) +apply (rule_tac env = "[u,A,r]" in mem_iff_sats) apply (rule sep_rules | simp)+ done @@ -376,35 +376,35 @@ subsection{*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) & + "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) & + \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) |] + "[| 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) + \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 rallI) -apply (rename_tac B) -apply (rule separation_CollectI) -apply (rule_tac A="{A,B,r,z}" in subset_LsetE, blast ) +apply (rename_tac B) +apply (rule separation_CollectI) +apply (rule_tac A="{A,B,r,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF omap_reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule bex_iff_sats conj_iff_sats)+ -apply (rule_tac env = "[a,u,A,B,r]" in mem_iff_sats) +apply (rule_tac env = "[a,u,A,B,r]" in mem_iff_sats) apply (rule sep_rules | simp)+ done @@ -412,34 +412,34 @@ subsection{*Separation for a Theorem about @{term "obase"}*} lemma is_recfun_reflects: - "REFLECTS[\x. \xa[L]. \xb[L]. - pair(L,x,a,xa) & xa \ r & pair(L,x,b,xb) & xb \ r & - (\fx[L]. \gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & + "REFLECTS[\x. \xa[L]. \xb[L]. + pair(L,x,a,xa) & xa \ r & pair(L,x,b,xb) & xb \ r & + (\fx[L]. \gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & fx \ gx), - \i x. \xa \ Lset(i). \xb \ Lset(i). + \i x. \xa \ Lset(i). \xb \ Lset(i). pair(**Lset(i),x,a,xa) & xa \ r & pair(**Lset(i),x,b,xb) & xb \ r & - (\fx \ Lset(i). \gx \ Lset(i). fun_apply(**Lset(i),f,x,fx) & + (\fx \ Lset(i). \gx \ Lset(i). fun_apply(**Lset(i),f,x,fx) & fun_apply(**Lset(i),g,x,gx) & fx \ gx)]" by (intro FOL_reflections function_reflections fun_plus_reflections) lemma is_recfun_separation: --{*for well-founded recursion*} - "[| L(r); L(f); L(g); L(a); L(b) |] - ==> separation(L, - \x. \xa[L]. \xb[L]. - pair(L,x,a,xa) & xa \ r & pair(L,x,b,xb) & xb \ r & - (\fx[L]. \gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & + "[| L(r); L(f); L(g); L(a); L(b) |] + ==> separation(L, + \x. \xa[L]. \xb[L]. + pair(L,x,a,xa) & xa \ r & pair(L,x,b,xb) & xb \ r & + (\fx[L]. \gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & fx \ gx))" -apply (rule separation_CollectI) -apply (rule_tac A="{r,f,g,a,b,z}" in subset_LsetE, blast ) +apply (rule separation_CollectI) +apply (rule_tac A="{r,f,g,a,b,z}" in subset_LsetE, blast ) apply (rule ReflectsE [OF is_recfun_reflects], assumption) -apply (drule subset_Lset_ltD, assumption) +apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2) apply (rule DPow_LsetI) -apply (rename_tac u) +apply (rename_tac u) apply (rule bex_iff_sats conj_iff_sats)+ -apply (rule_tac env = "[xa,u,r,f,g,a,b]" in pair_iff_sats) +apply (rule_tac env = "[xa,u,r,f,g,a,b]" in pair_iff_sats) apply (rule sep_rules | simp)+ done @@ -448,144 +448,128 @@ text{*Separation (and Strong Replacement) for basic set-theoretic constructions such as intersection, Cartesian Product and image.*} -ML -{* -val Inter_separation = thm "Inter_separation"; -val cartprod_separation = thm "cartprod_separation"; -val image_separation = thm "image_separation"; -val converse_separation = thm "converse_separation"; -val restrict_separation = thm "restrict_separation"; -val comp_separation = thm "comp_separation"; -val pred_separation = thm "pred_separation"; -val Memrel_separation = thm "Memrel_separation"; -val funspace_succ_replacement = thm "funspace_succ_replacement"; -val well_ord_iso_separation = thm "well_ord_iso_separation"; -val obase_separation = thm "obase_separation"; -val obase_equals_separation = thm "obase_equals_separation"; -val omap_replacement = thm "omap_replacement"; -val is_recfun_separation = thm "is_recfun_separation"; - -val m_axioms = - [Inter_separation, cartprod_separation, image_separation, - converse_separation, restrict_separation, comp_separation, - pred_separation, Memrel_separation, funspace_succ_replacement, - well_ord_iso_separation, obase_separation, - obase_equals_separation, omap_replacement, is_recfun_separation] - -fun axioms_L th = - kill_flex_triv_prems (m_axioms MRS (triv_axioms_L th)); +theorem M_axioms_axioms_L: "M_axioms_axioms(L)" + apply (rule M_axioms_axioms.intro) + apply (assumption | rule + Inter_separation cartprod_separation image_separation + converse_separation restrict_separation + comp_separation pred_separation Memrel_separation + funspace_succ_replacement well_ord_iso_separation + obase_separation obase_equals_separation + omap_replacement is_recfun_separation)+ + done + +theorem M_axioms_L: "PROP M_axioms(L)" + apply (rule M_axioms.intro) + apply (rule M_triv_axioms_L) + apply (rule M_axioms_axioms_L) + done -bind_thm ("cartprod_iff", axioms_L (thm "M_axioms.cartprod_iff")); -bind_thm ("cartprod_closed", axioms_L (thm "M_axioms.cartprod_closed")); -bind_thm ("sum_closed", axioms_L (thm "M_axioms.sum_closed")); -bind_thm ("M_converse_iff", axioms_L (thm "M_axioms.M_converse_iff")); -bind_thm ("converse_closed", axioms_L (thm "M_axioms.converse_closed")); -bind_thm ("converse_abs", axioms_L (thm "M_axioms.converse_abs")); -bind_thm ("image_closed", axioms_L (thm "M_axioms.image_closed")); -bind_thm ("vimage_abs", axioms_L (thm "M_axioms.vimage_abs")); -bind_thm ("vimage_closed", axioms_L (thm "M_axioms.vimage_closed")); -bind_thm ("domain_abs", axioms_L (thm "M_axioms.domain_abs")); -bind_thm ("domain_closed", axioms_L (thm "M_axioms.domain_closed")); -bind_thm ("range_abs", axioms_L (thm "M_axioms.range_abs")); -bind_thm ("range_closed", axioms_L (thm "M_axioms.range_closed")); -bind_thm ("field_abs", axioms_L (thm "M_axioms.field_abs")); -bind_thm ("field_closed", axioms_L (thm "M_axioms.field_closed")); -bind_thm ("relation_abs", axioms_L (thm "M_axioms.relation_abs")); -bind_thm ("function_abs", axioms_L (thm "M_axioms.function_abs")); -bind_thm ("apply_closed", axioms_L (thm "M_axioms.apply_closed")); -bind_thm ("apply_abs", axioms_L (thm "M_axioms.apply_abs")); -bind_thm ("typed_function_abs", axioms_L (thm "M_axioms.typed_function_abs")); -bind_thm ("injection_abs", axioms_L (thm "M_axioms.injection_abs")); -bind_thm ("surjection_abs", axioms_L (thm "M_axioms.surjection_abs")); -bind_thm ("bijection_abs", axioms_L (thm "M_axioms.bijection_abs")); -bind_thm ("M_comp_iff", axioms_L (thm "M_axioms.M_comp_iff")); -bind_thm ("comp_closed", axioms_L (thm "M_axioms.comp_closed")); -bind_thm ("composition_abs", axioms_L (thm "M_axioms.composition_abs")); -bind_thm ("restriction_is_function", axioms_L (thm "M_axioms.restriction_is_function")); -bind_thm ("restriction_abs", axioms_L (thm "M_axioms.restriction_abs")); -bind_thm ("M_restrict_iff", axioms_L (thm "M_axioms.M_restrict_iff")); -bind_thm ("restrict_closed", axioms_L (thm "M_axioms.restrict_closed")); -bind_thm ("Inter_abs", axioms_L (thm "M_axioms.Inter_abs")); -bind_thm ("Inter_closed", axioms_L (thm "M_axioms.Inter_closed")); -bind_thm ("Int_closed", axioms_L (thm "M_axioms.Int_closed")); -bind_thm ("finite_fun_closed", axioms_L (thm "M_axioms.finite_fun_closed")); -bind_thm ("is_funspace_abs", axioms_L (thm "M_axioms.is_funspace_abs")); -bind_thm ("succ_fun_eq2", axioms_L (thm "M_axioms.succ_fun_eq2")); -bind_thm ("funspace_succ", axioms_L (thm "M_axioms.funspace_succ")); -bind_thm ("finite_funspace_closed", axioms_L (thm "M_axioms.finite_funspace_closed")); -*} +lemmas cartprod_iff = M_axioms.cartprod_iff [OF M_axioms_L] + and cartprod_closed = M_axioms.cartprod_closed [OF M_axioms_L] + and sum_closed = M_axioms.sum_closed [OF M_axioms_L] + and M_converse_iff = M_axioms.M_converse_iff [OF M_axioms_L] + and converse_closed = M_axioms.converse_closed [OF M_axioms_L] + and converse_abs = M_axioms.converse_abs [OF M_axioms_L] + and image_closed = M_axioms.image_closed [OF M_axioms_L] + and vimage_abs = M_axioms.vimage_abs [OF M_axioms_L] + and vimage_closed = M_axioms.vimage_closed [OF M_axioms_L] + and domain_abs = M_axioms.domain_abs [OF M_axioms_L] + and domain_closed = M_axioms.domain_closed [OF M_axioms_L] + and range_abs = M_axioms.range_abs [OF M_axioms_L] + and range_closed = M_axioms.range_closed [OF M_axioms_L] + and field_abs = M_axioms.field_abs [OF M_axioms_L] + and field_closed = M_axioms.field_closed [OF M_axioms_L] + and relation_abs = M_axioms.relation_abs [OF M_axioms_L] + and function_abs = M_axioms.function_abs [OF M_axioms_L] + and apply_closed = M_axioms.apply_closed [OF M_axioms_L] + and apply_abs = M_axioms.apply_abs [OF M_axioms_L] + and typed_function_abs = M_axioms.typed_function_abs [OF M_axioms_L] + and injection_abs = M_axioms.injection_abs [OF M_axioms_L] + and surjection_abs = M_axioms.surjection_abs [OF M_axioms_L] + and bijection_abs = M_axioms.bijection_abs [OF M_axioms_L] + and M_comp_iff = M_axioms.M_comp_iff [OF M_axioms_L] + and comp_closed = M_axioms.comp_closed [OF M_axioms_L] + and composition_abs = M_axioms.composition_abs [OF M_axioms_L] + and restriction_is_function = M_axioms.restriction_is_function [OF M_axioms_L] + and restriction_abs = M_axioms.restriction_abs [OF M_axioms_L] + and M_restrict_iff = M_axioms.M_restrict_iff [OF M_axioms_L] + and restrict_closed = M_axioms.restrict_closed [OF M_axioms_L] + and Inter_abs = M_axioms.Inter_abs [OF M_axioms_L] + and Inter_closed = M_axioms.Inter_closed [OF M_axioms_L] + and Int_closed = M_axioms.Int_closed [OF M_axioms_L] + and finite_fun_closed = M_axioms.finite_fun_closed [OF M_axioms_L] + and is_funspace_abs = M_axioms.is_funspace_abs [OF M_axioms_L] + and succ_fun_eq2 = M_axioms.succ_fun_eq2 [OF M_axioms_L] + and funspace_succ = M_axioms.funspace_succ [OF M_axioms_L] + and finite_funspace_closed = M_axioms.finite_funspace_closed [OF M_axioms_L] -ML -{* -bind_thm ("is_recfun_equal", axioms_L (thm "M_axioms.is_recfun_equal")); -bind_thm ("is_recfun_cut", axioms_L (thm "M_axioms.is_recfun_cut")); -bind_thm ("is_recfun_functional", axioms_L (thm "M_axioms.is_recfun_functional")); -bind_thm ("is_recfun_relativize", axioms_L (thm "M_axioms.is_recfun_relativize")); -bind_thm ("is_recfun_restrict", axioms_L (thm "M_axioms.is_recfun_restrict")); -bind_thm ("univalent_is_recfun", axioms_L (thm "M_axioms.univalent_is_recfun")); -bind_thm ("exists_is_recfun_indstep", axioms_L (thm "M_axioms.exists_is_recfun_indstep")); -bind_thm ("wellfounded_exists_is_recfun", axioms_L (thm "M_axioms.wellfounded_exists_is_recfun")); -bind_thm ("wf_exists_is_recfun", axioms_L (thm "M_axioms.wf_exists_is_recfun")); -bind_thm ("is_recfun_abs", axioms_L (thm "M_axioms.is_recfun_abs")); -bind_thm ("irreflexive_abs", axioms_L (thm "M_axioms.irreflexive_abs")); -bind_thm ("transitive_rel_abs", axioms_L (thm "M_axioms.transitive_rel_abs")); -bind_thm ("linear_rel_abs", axioms_L (thm "M_axioms.linear_rel_abs")); -bind_thm ("wellordered_is_trans_on", axioms_L (thm "M_axioms.wellordered_is_trans_on")); -bind_thm ("wellordered_is_linear", axioms_L (thm "M_axioms.wellordered_is_linear")); -bind_thm ("wellordered_is_wellfounded_on", axioms_L (thm "M_axioms.wellordered_is_wellfounded_on")); -bind_thm ("wellfounded_imp_wellfounded_on", axioms_L (thm "M_axioms.wellfounded_imp_wellfounded_on")); -bind_thm ("wellfounded_on_subset_A", axioms_L (thm "M_axioms.wellfounded_on_subset_A")); -bind_thm ("wellfounded_on_iff_wellfounded", axioms_L (thm "M_axioms.wellfounded_on_iff_wellfounded")); -bind_thm ("wellfounded_on_imp_wellfounded", axioms_L (thm "M_axioms.wellfounded_on_imp_wellfounded")); -bind_thm ("wellfounded_on_field_imp_wellfounded", axioms_L (thm "M_axioms.wellfounded_on_field_imp_wellfounded")); -bind_thm ("wellfounded_iff_wellfounded_on_field", axioms_L (thm "M_axioms.wellfounded_iff_wellfounded_on_field")); -bind_thm ("wellfounded_induct", axioms_L (thm "M_axioms.wellfounded_induct")); -bind_thm ("wellfounded_on_induct", axioms_L (thm "M_axioms.wellfounded_on_induct")); -bind_thm ("wellfounded_on_induct2", axioms_L (thm "M_axioms.wellfounded_on_induct2")); -bind_thm ("linear_imp_relativized", axioms_L (thm "M_axioms.linear_imp_relativized")); -bind_thm ("trans_on_imp_relativized", axioms_L (thm "M_axioms.trans_on_imp_relativized")); -bind_thm ("wf_on_imp_relativized", axioms_L (thm "M_axioms.wf_on_imp_relativized")); -bind_thm ("wf_imp_relativized", axioms_L (thm "M_axioms.wf_imp_relativized")); -bind_thm ("well_ord_imp_relativized", axioms_L (thm "M_axioms.well_ord_imp_relativized")); -bind_thm ("order_isomorphism_abs", axioms_L (thm "M_axioms.order_isomorphism_abs")); -bind_thm ("pred_set_abs", axioms_L (thm "M_axioms.pred_set_abs")); -*} +lemmas is_recfun_equal = M_axioms.is_recfun_equal [OF M_axioms_L] + and is_recfun_cut = M_axioms.is_recfun_cut [OF M_axioms_L] + and is_recfun_functional = M_axioms.is_recfun_functional [OF M_axioms_L] + and is_recfun_relativize = M_axioms.is_recfun_relativize [OF M_axioms_L] + and is_recfun_restrict = M_axioms.is_recfun_restrict [OF M_axioms_L] + and univalent_is_recfun = M_axioms.univalent_is_recfun [OF M_axioms_L] + and exists_is_recfun_indstep = M_axioms.exists_is_recfun_indstep [OF M_axioms_L] + and wellfounded_exists_is_recfun = M_axioms.wellfounded_exists_is_recfun [OF M_axioms_L] + and wf_exists_is_recfun = M_axioms.wf_exists_is_recfun [OF M_axioms_L] + and is_recfun_abs = M_axioms.is_recfun_abs [OF M_axioms_L] + and irreflexive_abs = M_axioms.irreflexive_abs [OF M_axioms_L] + and transitive_rel_abs = M_axioms.transitive_rel_abs [OF M_axioms_L] + and linear_rel_abs = M_axioms.linear_rel_abs [OF M_axioms_L] + and wellordered_is_trans_on = M_axioms.wellordered_is_trans_on [OF M_axioms_L] + and wellordered_is_linear = M_axioms.wellordered_is_linear [OF M_axioms_L] + and wellordered_is_wellfounded_on = M_axioms.wellordered_is_wellfounded_on [OF M_axioms_L] + and wellfounded_imp_wellfounded_on = M_axioms.wellfounded_imp_wellfounded_on [OF M_axioms_L] + and wellfounded_on_subset_A = M_axioms.wellfounded_on_subset_A [OF M_axioms_L] + and wellfounded_on_iff_wellfounded = M_axioms.wellfounded_on_iff_wellfounded [OF M_axioms_L] + and wellfounded_on_imp_wellfounded = M_axioms.wellfounded_on_imp_wellfounded [OF M_axioms_L] + and wellfounded_on_field_imp_wellfounded = M_axioms.wellfounded_on_field_imp_wellfounded [OF M_axioms_L] + and wellfounded_iff_wellfounded_on_field = M_axioms.wellfounded_iff_wellfounded_on_field [OF M_axioms_L] + and wellfounded_induct = M_axioms.wellfounded_induct [OF M_axioms_L] + and wellfounded_on_induct = M_axioms.wellfounded_on_induct [OF M_axioms_L] + and wellfounded_on_induct2 = M_axioms.wellfounded_on_induct2 [OF M_axioms_L] + and linear_imp_relativized = M_axioms.linear_imp_relativized [OF M_axioms_L] + and trans_on_imp_relativized = M_axioms.trans_on_imp_relativized [OF M_axioms_L] + and wf_on_imp_relativized = M_axioms.wf_on_imp_relativized [OF M_axioms_L] + and wf_imp_relativized = M_axioms.wf_imp_relativized [OF M_axioms_L] + and well_ord_imp_relativized = M_axioms.well_ord_imp_relativized [OF M_axioms_L] + and order_isomorphism_abs = M_axioms.order_isomorphism_abs [OF M_axioms_L] + and pred_set_abs = M_axioms.pred_set_abs [OF M_axioms_L] -ML -{* -bind_thm ("pred_closed", axioms_L (thm "M_axioms.pred_closed")); -bind_thm ("membership_abs", axioms_L (thm "M_axioms.membership_abs")); -bind_thm ("M_Memrel_iff", axioms_L (thm "M_axioms.M_Memrel_iff")); -bind_thm ("Memrel_closed", axioms_L (thm "M_axioms.Memrel_closed")); -bind_thm ("wellordered_iso_predD", axioms_L (thm "M_axioms.wellordered_iso_predD")); -bind_thm ("wellordered_iso_pred_eq", axioms_L (thm "M_axioms.wellordered_iso_pred_eq")); -bind_thm ("wellfounded_on_asym", axioms_L (thm "M_axioms.wellfounded_on_asym")); -bind_thm ("wellordered_asym", axioms_L (thm "M_axioms.wellordered_asym")); -bind_thm ("ord_iso_pred_imp_lt", axioms_L (thm "M_axioms.ord_iso_pred_imp_lt")); -bind_thm ("obase_iff", axioms_L (thm "M_axioms.obase_iff")); -bind_thm ("omap_iff", axioms_L (thm "M_axioms.omap_iff")); -bind_thm ("omap_unique", axioms_L (thm "M_axioms.omap_unique")); -bind_thm ("omap_yields_Ord", axioms_L (thm "M_axioms.omap_yields_Ord")); -bind_thm ("otype_iff", axioms_L (thm "M_axioms.otype_iff")); -bind_thm ("otype_eq_range", axioms_L (thm "M_axioms.otype_eq_range")); -bind_thm ("Ord_otype", axioms_L (thm "M_axioms.Ord_otype")); -bind_thm ("domain_omap", axioms_L (thm "M_axioms.domain_omap")); -bind_thm ("omap_subset", axioms_L (thm "M_axioms.omap_subset")); -bind_thm ("omap_funtype", axioms_L (thm "M_axioms.omap_funtype")); -bind_thm ("wellordered_omap_bij", axioms_L (thm "M_axioms.wellordered_omap_bij")); -bind_thm ("omap_ord_iso", axioms_L (thm "M_axioms.omap_ord_iso")); -bind_thm ("Ord_omap_image_pred", axioms_L (thm "M_axioms.Ord_omap_image_pred")); -bind_thm ("restrict_omap_ord_iso", axioms_L (thm "M_axioms.restrict_omap_ord_iso")); -bind_thm ("obase_equals", axioms_L (thm "M_axioms.obase_equals")); -bind_thm ("omap_ord_iso_otype", axioms_L (thm "M_axioms.omap_ord_iso_otype")); -bind_thm ("obase_exists", axioms_L (thm "M_axioms.obase_exists")); -bind_thm ("omap_exists", axioms_L (thm "M_axioms.omap_exists")); -bind_thm ("otype_exists", axioms_L (thm "M_axioms.otype_exists")); -bind_thm ("omap_ord_iso_otype", axioms_L (thm "M_axioms.omap_ord_iso_otype")); -bind_thm ("ordertype_exists", axioms_L (thm "M_axioms.ordertype_exists")); -bind_thm ("relativized_imp_well_ord", axioms_L (thm "M_axioms.relativized_imp_well_ord")); -bind_thm ("well_ord_abs", axioms_L (thm "M_axioms.well_ord_abs")); -*} +lemmas pred_closed = M_axioms.pred_closed [OF M_axioms_L] + and membership_abs = M_axioms.membership_abs [OF M_axioms_L] + and M_Memrel_iff = M_axioms.M_Memrel_iff [OF M_axioms_L] + and Memrel_closed = M_axioms.Memrel_closed [OF M_axioms_L] + and wellordered_iso_predD = M_axioms.wellordered_iso_predD [OF M_axioms_L] + and wellordered_iso_pred_eq = M_axioms.wellordered_iso_pred_eq [OF M_axioms_L] + and wellfounded_on_asym = M_axioms.wellfounded_on_asym [OF M_axioms_L] + and wellordered_asym = M_axioms.wellordered_asym [OF M_axioms_L] + and ord_iso_pred_imp_lt = M_axioms.ord_iso_pred_imp_lt [OF M_axioms_L] + and obase_iff = M_axioms.obase_iff [OF M_axioms_L] + and omap_iff = M_axioms.omap_iff [OF M_axioms_L] + and omap_unique = M_axioms.omap_unique [OF M_axioms_L] + and omap_yields_Ord = M_axioms.omap_yields_Ord [OF M_axioms_L] + and otype_iff = M_axioms.otype_iff [OF M_axioms_L] + and otype_eq_range = M_axioms.otype_eq_range [OF M_axioms_L] + and Ord_otype = M_axioms.Ord_otype [OF M_axioms_L] + and domain_omap = M_axioms.domain_omap [OF M_axioms_L] + and omap_subset = M_axioms.omap_subset [OF M_axioms_L] + and omap_funtype = M_axioms.omap_funtype [OF M_axioms_L] + and wellordered_omap_bij = M_axioms.wellordered_omap_bij [OF M_axioms_L] + and omap_ord_iso = M_axioms.omap_ord_iso [OF M_axioms_L] + and Ord_omap_image_pred = M_axioms.Ord_omap_image_pred [OF M_axioms_L] + and restrict_omap_ord_iso = M_axioms.restrict_omap_ord_iso [OF M_axioms_L] + and obase_equals = M_axioms.obase_equals [OF M_axioms_L] + and omap_ord_iso_otype = M_axioms.omap_ord_iso_otype [OF M_axioms_L] + and obase_exists = M_axioms.obase_exists [OF M_axioms_L] + and omap_exists = M_axioms.omap_exists [OF M_axioms_L] + and otype_exists = M_axioms.otype_exists [OF M_axioms_L] + and omap_ord_iso_otype' = M_axioms.omap_ord_iso_otype' [OF M_axioms_L] + and ordertype_exists = M_axioms.ordertype_exists [OF M_axioms_L] + and relativized_imp_well_ord = M_axioms.relativized_imp_well_ord [OF M_axioms_L] + and well_ord_abs = M_axioms.well_ord_abs [OF M_axioms_L] + declare cartprod_closed [intro,simp] declare sum_closed [intro,simp] @@ -614,7 +598,6 @@ declare Inter_abs [simp] declare Inter_closed [intro,simp] declare Int_closed [intro,simp] -declare finite_fun_closed [rule_format] declare is_funspace_abs [simp] declare finite_funspace_closed [intro,simp] diff -r b429fd98549c -r 99e52e78eb65 src/ZF/Constructible/WF_absolute.thy --- a/src/ZF/Constructible/WF_absolute.thy Sun Jul 28 21:09:37 2002 +0200 +++ b/src/ZF/Constructible/WF_absolute.thy Mon Jul 29 00:57:16 2002 +0200 @@ -143,7 +143,7 @@ by (simp add: rtran_closure_mem_def Ord_succ_mem_iff nat_0_le [THEN ltD]) -locale (open) M_trancl = M_axioms + +locale M_trancl = M_axioms + assumes rtrancl_separation: "[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))" and wellfounded_trancl_separation: @@ -231,7 +231,7 @@ rank function.*} -locale (open) M_wfrank = M_trancl + +locale M_wfrank = M_trancl + assumes wfrank_separation: "M(r) ==> separation (M, \x. @@ -317,7 +317,7 @@ "[| wellfounded(M,r); a\A; M(r); M(A) |] ==> \f[M]. is_recfun(r^+, a, %x f. range(f), f) --> Ord(range(f))" apply (drule wellfounded_trancl, assumption) -apply (rule wellfounded_induct, assumption+) +apply (rule wellfounded_induct, assumption, erule (1) transM) apply simp apply (blast intro: Ord_wfrank_separation', clarify) txt{*The reasoning in both cases is that we get @{term y} such that @@ -445,7 +445,8 @@ apply (subgoal_tac "a\A & b\A") prefer 2 apply blast apply (simp add: lt_def Ord_wellfoundedrank, clarify) -apply (frule exists_wfrank [of concl: _ b], assumption+, clarify) +apply (frule exists_wfrank [of concl: _ b], erule (1) transM, assumption) +apply clarify apply (rename_tac fb) apply (frule is_recfun_restrict [of concl: "r^+" a]) apply (rule trans_trancl, assumption) diff -r b429fd98549c -r 99e52e78eb65 src/ZF/Constructible/WFrec.thy --- a/src/ZF/Constructible/WFrec.thy Sun Jul 28 21:09:37 2002 +0200 +++ b/src/ZF/Constructible/WFrec.thy Mon Jul 29 00:57:16 2002 +0200 @@ -378,7 +378,7 @@ fun_apply(M,f,j,fj) & fj = k" -locale (open) M_ord_arith = M_axioms + +locale M_ord_arith = M_axioms + assumes oadd_strong_replacement: "[| M(i); M(j) |] ==> strong_replacement(M, diff -r b429fd98549c -r 99e52e78eb65 src/ZF/Constructible/Wellorderings.thy --- a/src/ZF/Constructible/Wellorderings.thy Sun Jul 28 21:09:37 2002 +0200 +++ b/src/ZF/Constructible/Wellorderings.thy Mon Jul 29 00:57:16 2002 +0200 @@ -603,7 +603,7 @@ apply blast+ done -theorem (in M_axioms) omap_ord_iso_otype: +theorem (in M_axioms) omap_ord_iso_otype': "[| wellordered(M,A,r); M(A); M(r) |] ==> \f[M]. (\i[M]. Ord(i) & f \ ord_iso(A, r, i, Memrel(i)))" apply (insert obase_exists [of A r] omap_exists [of A r] otype_exists [of A r], simp, clarify) @@ -619,7 +619,7 @@ ==> \f[M]. (\i[M]. Ord(i) & f \ ord_iso(A, r, i, Memrel(i)))" apply (insert obase_exists [of A r] omap_exists [of A r] otype_exists [of A r], simp, clarify) apply (rename_tac i) -apply (subgoal_tac "Ord(i)", blast intro: omap_ord_iso_otype) +apply (subgoal_tac "Ord(i)", blast intro: omap_ord_iso_otype') apply (rule Ord_otype) apply (force simp add: otype_def range_closed) apply (simp_all add: wellordered_is_trans_on)