# HG changeset patch # User paulson # Date 1031669491 -7200 # Node ID 1500a2e48d443d3c27e6e386a0ab01f913df3228 # Parent 7d6c9817c4329e20da104f1fba17839dcddfcbd8 renamed M_triv_axioms to M_trivial and M_axioms to M_basic diff -r 7d6c9817c432 -r 1500a2e48d44 src/ZF/Constructible/Datatype_absolute.thy --- a/src/ZF/Constructible/Datatype_absolute.thy Tue Sep 10 16:47:17 2002 +0200 +++ b/src/ZF/Constructible/Datatype_absolute.thy Tue Sep 10 16:51:31 2002 +0200 @@ -126,13 +126,13 @@ \n[M]. n\nat --> wfrec_replacement(M, iterates_MH(M,isF,v), Memrel(succ(n)))" -lemma (in M_axioms) iterates_MH_abs: +lemma (in M_basic) iterates_MH_abs: "[| relativize1(M,isF,F); M(n); M(g); M(z) |] ==> iterates_MH(M,isF,v,n,g,z) <-> z = nat_case(v, \m. F(g`m), n)" by (simp add: nat_case_abs [of _ "\m. F(g ` m)"] relativize1_def iterates_MH_def) -lemma (in M_axioms) iterates_imp_wfrec_replacement: +lemma (in M_basic) iterates_imp_wfrec_replacement: "[|relativize1(M,isF,F); n \ nat; iterates_replacement(M,isF,v)|] ==> wfrec_replacement(M, \n f z. z = nat_case(v, \m. F(f`m), n), Memrel(succ(n)))" @@ -210,7 +210,7 @@ \n1[M]. \AX[M]. number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)" -lemma (in M_axioms) list_functor_abs [simp]: +lemma (in M_basic) list_functor_abs [simp]: "[| M(A); M(X); M(Z) |] ==> is_list_functor(M,A,X,Z) <-> (Z = {0} + A*X)" by (simp add: is_list_functor_def singleton_0 nat_into_M) @@ -266,7 +266,7 @@ cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) & is_sum(M,natnatsum,X3,Z)" -lemma (in M_axioms) formula_functor_abs [simp]: +lemma (in M_basic) formula_functor_abs [simp]: "[| M(X); M(Z) |] ==> is_formula_functor(M,X,Z) <-> Z = ((nat*nat) + (nat*nat)) + (X*X + X)" @@ -323,7 +323,7 @@ text{*Express @{term list_rec} without using @{term rank} or @{term Vset}, neither of which is absolute.*} -lemma (in M_triv_axioms) list_rec_eq: +lemma (in M_trivial) list_rec_eq: "l \ list(A) ==> list_rec(a,g,l) = transrec (succ(length(l)), @@ -728,7 +728,7 @@ done text{*Proof is trivial since @{term length} returns natural numbers.*} -lemma (in M_triv_axioms) length_closed [intro,simp]: +lemma (in M_trivial) length_closed [intro,simp]: "l \ list(A) ==> M(length(l))" by (simp add: nat_into_M) @@ -744,7 +744,7 @@ apply (simp add: tl'_Cons iterates_commute) done -lemma (in M_axioms) iterates_tl'_closed: +lemma (in M_basic) iterates_tl'_closed: "[|n \ nat; M(x)|] ==> M(tl'^n (x))" apply (induct_tac n, simp) apply (simp add: tl'_Cons tl'_closed) @@ -785,11 +785,11 @@ "is_Member(M,x,y,Z) == \p[M]. \u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)" -lemma (in M_triv_axioms) Member_abs [simp]: +lemma (in M_trivial) Member_abs [simp]: "[|M(x); M(y); M(Z)|] ==> is_Member(M,x,y,Z) <-> (Z = Member(x,y))" by (simp add: is_Member_def Member_def) -lemma (in M_triv_axioms) Member_in_M_iff [iff]: +lemma (in M_trivial) Member_in_M_iff [iff]: "M(Member(x,y)) <-> M(x) & M(y)" by (simp add: Member_def) @@ -799,11 +799,11 @@ "is_Equal(M,x,y,Z) == \p[M]. \u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)" -lemma (in M_triv_axioms) Equal_abs [simp]: +lemma (in M_trivial) Equal_abs [simp]: "[|M(x); M(y); M(Z)|] ==> is_Equal(M,x,y,Z) <-> (Z = Equal(x,y))" by (simp add: is_Equal_def Equal_def) -lemma (in M_triv_axioms) Equal_in_M_iff [iff]: "M(Equal(x,y)) <-> M(x) & M(y)" +lemma (in M_trivial) Equal_in_M_iff [iff]: "M(Equal(x,y)) <-> M(x) & M(y)" by (simp add: Equal_def) constdefs @@ -812,11 +812,11 @@ "is_Nand(M,x,y,Z) == \p[M]. \u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)" -lemma (in M_triv_axioms) Nand_abs [simp]: +lemma (in M_trivial) Nand_abs [simp]: "[|M(x); M(y); M(Z)|] ==> is_Nand(M,x,y,Z) <-> (Z = Nand(x,y))" by (simp add: is_Nand_def Nand_def) -lemma (in M_triv_axioms) Nand_in_M_iff [iff]: "M(Nand(x,y)) <-> M(x) & M(y)" +lemma (in M_trivial) Nand_in_M_iff [iff]: "M(Nand(x,y)) <-> M(x) & M(y)" by (simp add: Nand_def) constdefs @@ -824,11 +824,11 @@ --{* because @{term "Forall(x) \ Inr(Inr(p))"}*} "is_Forall(M,p,Z) == \u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" -lemma (in M_triv_axioms) Forall_abs [simp]: +lemma (in M_trivial) Forall_abs [simp]: "[|M(x); M(Z)|] ==> is_Forall(M,x,Z) <-> (Z = Forall(x))" by (simp add: is_Forall_def Forall_def) -lemma (in M_triv_axioms) Forall_in_M_iff [iff]: "M(Forall(x)) <-> M(x)" +lemma (in M_trivial) Forall_in_M_iff [iff]: "M(Forall(x)) <-> M(x)" by (simp add: Forall_def) @@ -890,7 +890,7 @@ done text{*Proof is trivial since @{term depth} returns natural numbers.*} -lemma (in M_triv_axioms) depth_closed [intro,simp]: +lemma (in M_trivial) depth_closed [intro,simp]: "p \ formula ==> M(depth(p))" by (simp add: nat_into_M) @@ -916,7 +916,7 @@ text{*Unfold @{term formula_rec} to @{term formula_rec_case}. Express @{term formula_rec} without using @{term rank} or @{term Vset}, neither of which is absolute.*} -lemma (in M_triv_axioms) formula_rec_eq: +lemma (in M_trivial) formula_rec_eq: "p \ formula ==> formula_rec(a,b,c,d,p) = transrec (succ(depth(p)), diff -r 7d6c9817c432 -r 1500a2e48d44 src/ZF/Constructible/L_axioms.thy --- a/src/ZF/Constructible/L_axioms.thy Tue Sep 10 16:47:17 2002 +0200 +++ b/src/ZF/Constructible/L_axioms.thy Tue Sep 10 16:51:31 2002 +0200 @@ -8,7 +8,7 @@ theory L_axioms = Formula + Relative + Reflection + MetaExists: -text {* The class L satisfies the premises of locale @{text M_triv_axioms} *} +text {* The class L satisfies the premises of locale @{text M_trivial} *} lemma transL: "[| y\x; L(x) |] ==> L(y)" apply (insert Transset_Lset) @@ -80,7 +80,7 @@ apply (simp_all add: Replace_iff univalent_def, blast) done -subsection{*Instantiating the locale @{text M_triv_axioms}*} +subsection{*Instantiating the locale @{text M_trivial}*} text{*No instances of Separation yet.*} lemma Lset_mono_le: "mono_le_subset(Lset)" @@ -93,8 +93,8 @@ lemmas L_nat = Ord_in_L [OF Ord_nat] -theorem M_triv_axioms_L: "PROP M_triv_axioms(L)" - apply (rule M_triv_axioms.intro) +theorem M_trivial_L: "PROP M_trivial(L)" + apply (rule M_trivial.intro) apply (erule (1) transL) apply (rule nonempty) apply (rule upair_ax) @@ -104,50 +104,50 @@ apply (rule L_nat) done -lemmas rall_abs = M_triv_axioms.rall_abs [OF M_triv_axioms_L] - and rex_abs = 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 = M_triv_axioms.empty_abs [OF M_triv_axioms_L] - and subset_abs = M_triv_axioms.subset_abs [OF M_triv_axioms_L] - and upair_abs = M_triv_axioms.upair_abs [OF M_triv_axioms_L] - and upair_in_M_iff = M_triv_axioms.upair_in_M_iff [OF M_triv_axioms_L] - and singleton_in_M_iff = M_triv_axioms.singleton_in_M_iff [OF M_triv_axioms_L] - and pair_abs = M_triv_axioms.pair_abs [OF M_triv_axioms_L] - and pair_in_M_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 = M_triv_axioms.cartprod_abs [OF M_triv_axioms_L] - and union_abs = M_triv_axioms.union_abs [OF M_triv_axioms_L] - and inter_abs = M_triv_axioms.inter_abs [OF M_triv_axioms_L] - and setdiff_abs = M_triv_axioms.setdiff_abs [OF M_triv_axioms_L] - and Union_abs = M_triv_axioms.Union_abs [OF M_triv_axioms_L] - and Union_closed = M_triv_axioms.Union_closed [OF M_triv_axioms_L] - and Un_closed = M_triv_axioms.Un_closed [OF M_triv_axioms_L] - and cons_closed = M_triv_axioms.cons_closed [OF M_triv_axioms_L] - and successor_abs = M_triv_axioms.successor_abs [OF M_triv_axioms_L] - and succ_in_M_iff = M_triv_axioms.succ_in_M_iff [OF M_triv_axioms_L] - and separation_closed = 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 = M_triv_axioms.strong_replacement_closed [OF M_triv_axioms_L] - and RepFun_closed = M_triv_axioms.RepFun_closed [OF M_triv_axioms_L] - and lam_closed = M_triv_axioms.lam_closed [OF M_triv_axioms_L] - and image_abs = 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 = 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 = M_triv_axioms.Inl_in_M_iff [OF M_triv_axioms_L] - and Inr_in_M_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 = M_triv_axioms.transitive_set_abs [OF M_triv_axioms_L] - and ordinal_abs = M_triv_axioms.ordinal_abs [OF M_triv_axioms_L] - and limit_ordinal_abs = M_triv_axioms.limit_ordinal_abs [OF M_triv_axioms_L] - and successor_ordinal_abs = 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 = M_triv_axioms.omega_abs [OF M_triv_axioms_L] - and number1_abs = M_triv_axioms.number1_abs [OF M_triv_axioms_L] - and number2_abs = M_triv_axioms.number2_abs [OF M_triv_axioms_L] - and number3_abs = M_triv_axioms.number3_abs [OF M_triv_axioms_L] +lemmas rall_abs = M_trivial.rall_abs [OF M_trivial_L] + and rex_abs = M_trivial.rex_abs [OF M_trivial_L] + and ball_iff_equiv = M_trivial.ball_iff_equiv [OF M_trivial_L] + and M_equalityI = M_trivial.M_equalityI [OF M_trivial_L] + and empty_abs = M_trivial.empty_abs [OF M_trivial_L] + and subset_abs = M_trivial.subset_abs [OF M_trivial_L] + and upair_abs = M_trivial.upair_abs [OF M_trivial_L] + and upair_in_M_iff = M_trivial.upair_in_M_iff [OF M_trivial_L] + and singleton_in_M_iff = M_trivial.singleton_in_M_iff [OF M_trivial_L] + and pair_abs = M_trivial.pair_abs [OF M_trivial_L] + and pair_in_M_iff = M_trivial.pair_in_M_iff [OF M_trivial_L] + and pair_components_in_M = M_trivial.pair_components_in_M [OF M_trivial_L] + and cartprod_abs = M_trivial.cartprod_abs [OF M_trivial_L] + and union_abs = M_trivial.union_abs [OF M_trivial_L] + and inter_abs = M_trivial.inter_abs [OF M_trivial_L] + and setdiff_abs = M_trivial.setdiff_abs [OF M_trivial_L] + and Union_abs = M_trivial.Union_abs [OF M_trivial_L] + and Union_closed = M_trivial.Union_closed [OF M_trivial_L] + and Un_closed = M_trivial.Un_closed [OF M_trivial_L] + and cons_closed = M_trivial.cons_closed [OF M_trivial_L] + and successor_abs = M_trivial.successor_abs [OF M_trivial_L] + and succ_in_M_iff = M_trivial.succ_in_M_iff [OF M_trivial_L] + and separation_closed = M_trivial.separation_closed [OF M_trivial_L] + and strong_replacementI = M_trivial.strong_replacementI [OF M_trivial_L] + and strong_replacement_closed = M_trivial.strong_replacement_closed [OF M_trivial_L] + and RepFun_closed = M_trivial.RepFun_closed [OF M_trivial_L] + and lam_closed = M_trivial.lam_closed [OF M_trivial_L] + and image_abs = M_trivial.image_abs [OF M_trivial_L] + and powerset_Pow = M_trivial.powerset_Pow [OF M_trivial_L] + and powerset_imp_subset_Pow = M_trivial.powerset_imp_subset_Pow [OF M_trivial_L] + and nat_into_M = M_trivial.nat_into_M [OF M_trivial_L] + and nat_case_closed = M_trivial.nat_case_closed [OF M_trivial_L] + and Inl_in_M_iff = M_trivial.Inl_in_M_iff [OF M_trivial_L] + and Inr_in_M_iff = M_trivial.Inr_in_M_iff [OF M_trivial_L] + and lt_closed = M_trivial.lt_closed [OF M_trivial_L] + and transitive_set_abs = M_trivial.transitive_set_abs [OF M_trivial_L] + and ordinal_abs = M_trivial.ordinal_abs [OF M_trivial_L] + and limit_ordinal_abs = M_trivial.limit_ordinal_abs [OF M_trivial_L] + and successor_ordinal_abs = M_trivial.successor_ordinal_abs [OF M_trivial_L] + and finite_ordinal_abs = M_trivial.finite_ordinal_abs [OF M_trivial_L] + and omega_abs = M_trivial.omega_abs [OF M_trivial_L] + and number1_abs = M_trivial.number1_abs [OF M_trivial_L] + and number2_abs = M_trivial.number2_abs [OF M_trivial_L] + and number3_abs = M_trivial.number3_abs [OF M_trivial_L] declare rall_abs [simp] declare rex_abs [simp] diff -r 7d6c9817c432 -r 1500a2e48d44 src/ZF/Constructible/Rec_Separation.thy --- a/src/ZF/Constructible/Rec_Separation.thy Tue Sep 10 16:47:17 2002 +0200 +++ b/src/ZF/Constructible/Rec_Separation.thy Tue Sep 10 16:51:31 2002 +0200 @@ -212,7 +212,7 @@ theorem M_trancl_L: "PROP M_trancl(L)" by (rule M_trancl.intro - [OF M_triv_axioms_L M_axioms_axioms_L M_trancl_axioms_L]) + [OF M_trivial_L M_basic_axioms_L M_trancl_axioms_L]) lemmas iterates_abs = M_trancl.iterates_abs [OF M_trancl_L] and rtran_closure_rtrancl = M_trancl.rtran_closure_rtrancl [OF M_trancl_L] diff -r 7d6c9817c432 -r 1500a2e48d44 src/ZF/Constructible/Relative.thy --- a/src/ZF/Constructible/Relative.thy Tue Sep 10 16:47:17 2002 +0200 +++ b/src/ZF/Constructible/Relative.thy Tue Sep 10 16:51:31 2002 +0200 @@ -465,7 +465,7 @@ text{*The class M is assumed to be transitive and to satisfy some relativized ZF axioms*} -locale M_triv_axioms = +locale M_trivial = fixes M assumes transM: "[| y\x; M(x) |] ==> M(y)" and nonempty [simp]: "M(0)" @@ -475,73 +475,73 @@ and replacement: "replacement(M,P)" and M_nat [iff]: "M(nat)" (*i.e. the axiom of infinity*) -lemma (in M_triv_axioms) rall_abs [simp]: +lemma (in M_trivial) rall_abs [simp]: "M(A) ==> (\x[M]. x\A --> P(x)) <-> (\x\A. P(x))" by (blast intro: transM) -lemma (in M_triv_axioms) rex_abs [simp]: +lemma (in M_trivial) rex_abs [simp]: "M(A) ==> (\x[M]. x\A & P(x)) <-> (\x\A. P(x))" by (blast intro: transM) -lemma (in M_triv_axioms) ball_iff_equiv: +lemma (in M_trivial) ball_iff_equiv: "M(A) ==> (\x[M]. (x\A <-> P(x))) <-> (\x\A. P(x)) & (\x. P(x) --> M(x) --> x\A)" by (blast intro: transM) text{*Simplifies proofs of equalities when there's an iff-equality available for rewriting, universally quantified over M. *} -lemma (in M_triv_axioms) M_equalityI: +lemma (in M_trivial) M_equalityI: "[| !!x. M(x) ==> x\A <-> x\B; M(A); M(B) |] ==> A=B" by (blast intro!: equalityI dest: transM) subsubsection{*Trivial Absoluteness Proofs: Empty Set, Pairs, etc.*} -lemma (in M_triv_axioms) empty_abs [simp]: +lemma (in M_trivial) empty_abs [simp]: "M(z) ==> empty(M,z) <-> z=0" apply (simp add: empty_def) apply (blast intro: transM) done -lemma (in M_triv_axioms) subset_abs [simp]: +lemma (in M_trivial) subset_abs [simp]: "M(A) ==> subset(M,A,B) <-> A \ B" apply (simp add: subset_def) apply (blast intro: transM) done -lemma (in M_triv_axioms) upair_abs [simp]: +lemma (in M_trivial) upair_abs [simp]: "M(z) ==> upair(M,a,b,z) <-> z={a,b}" apply (simp add: upair_def) apply (blast intro: transM) done -lemma (in M_triv_axioms) upair_in_M_iff [iff]: +lemma (in M_trivial) upair_in_M_iff [iff]: "M({a,b}) <-> M(a) & M(b)" apply (insert upair_ax, simp add: upair_ax_def) apply (blast intro: transM) done -lemma (in M_triv_axioms) singleton_in_M_iff [iff]: +lemma (in M_trivial) singleton_in_M_iff [iff]: "M({a}) <-> M(a)" by (insert upair_in_M_iff [of a a], simp) -lemma (in M_triv_axioms) pair_abs [simp]: +lemma (in M_trivial) pair_abs [simp]: "M(z) ==> pair(M,a,b,z) <-> z=" apply (simp add: pair_def ZF.Pair_def) apply (blast intro: transM) done -lemma (in M_triv_axioms) pair_in_M_iff [iff]: +lemma (in M_trivial) pair_in_M_iff [iff]: "M() <-> M(a) & M(b)" by (simp add: ZF.Pair_def) -lemma (in M_triv_axioms) pair_components_in_M: +lemma (in M_trivial) pair_components_in_M: "[| \ A; M(A) |] ==> M(x) & M(y)" apply (simp add: Pair_def) apply (blast dest: transM) done -lemma (in M_triv_axioms) cartprod_abs [simp]: +lemma (in M_trivial) cartprod_abs [simp]: "[| M(A); M(B); M(z) |] ==> cartprod(M,A,B,z) <-> z = A*B" apply (simp add: cartprod_def) apply (rule iffI) @@ -551,51 +551,51 @@ subsubsection{*Absoluteness for Unions and Intersections*} -lemma (in M_triv_axioms) union_abs [simp]: +lemma (in M_trivial) union_abs [simp]: "[| M(a); M(b); M(z) |] ==> union(M,a,b,z) <-> z = a Un b" apply (simp add: union_def) apply (blast intro: transM) done -lemma (in M_triv_axioms) inter_abs [simp]: +lemma (in M_trivial) inter_abs [simp]: "[| M(a); M(b); M(z) |] ==> inter(M,a,b,z) <-> z = a Int b" apply (simp add: inter_def) apply (blast intro: transM) done -lemma (in M_triv_axioms) setdiff_abs [simp]: +lemma (in M_trivial) setdiff_abs [simp]: "[| M(a); M(b); M(z) |] ==> setdiff(M,a,b,z) <-> z = a-b" apply (simp add: setdiff_def) apply (blast intro: transM) done -lemma (in M_triv_axioms) Union_abs [simp]: +lemma (in M_trivial) Union_abs [simp]: "[| M(A); M(z) |] ==> big_union(M,A,z) <-> z = Union(A)" apply (simp add: big_union_def) apply (blast intro!: equalityI dest: transM) done -lemma (in M_triv_axioms) Union_closed [intro,simp]: +lemma (in M_trivial) Union_closed [intro,simp]: "M(A) ==> M(Union(A))" by (insert Union_ax, simp add: Union_ax_def) -lemma (in M_triv_axioms) Un_closed [intro,simp]: +lemma (in M_trivial) Un_closed [intro,simp]: "[| M(A); M(B) |] ==> M(A Un B)" by (simp only: Un_eq_Union, blast) -lemma (in M_triv_axioms) cons_closed [intro,simp]: +lemma (in M_trivial) cons_closed [intro,simp]: "[| M(a); M(A) |] ==> M(cons(a,A))" by (subst cons_eq [symmetric], blast) -lemma (in M_triv_axioms) cons_abs [simp]: +lemma (in M_trivial) cons_abs [simp]: "[| M(b); M(z) |] ==> is_cons(M,a,b,z) <-> z = cons(a,b)" by (simp add: is_cons_def, blast intro: transM) -lemma (in M_triv_axioms) successor_abs [simp]: +lemma (in M_trivial) successor_abs [simp]: "[| M(a); M(z) |] ==> successor(M,a,z) <-> z = succ(a)" by (simp add: successor_def, blast) -lemma (in M_triv_axioms) succ_in_M_iff [iff]: +lemma (in M_trivial) succ_in_M_iff [iff]: "M(succ(a)) <-> M(a)" apply (simp add: succ_def) apply (blast intro: transM) @@ -603,7 +603,7 @@ subsubsection{*Absoluteness for Separation and Replacement*} -lemma (in M_triv_axioms) separation_closed [intro,simp]: +lemma (in M_trivial) separation_closed [intro,simp]: "[| separation(M,P); M(A) |] ==> M(Collect(A,P))" apply (insert separation, simp add: separation_def) apply (drule rspec, assumption, clarify) @@ -615,14 +615,14 @@ "separation(M,P) <-> (\z[M]. \y[M]. is_Collect(M,z,P,y))" by (simp add: separation_def is_Collect_def) -lemma (in M_triv_axioms) Collect_abs [simp]: +lemma (in M_trivial) Collect_abs [simp]: "[| M(A); M(z) |] ==> is_Collect(M,A,P,z) <-> z = Collect(A,P)" apply (simp add: is_Collect_def) apply (blast intro!: equalityI dest: transM) done text{*Probably the premise and conclusion are equivalent*} -lemma (in M_triv_axioms) strong_replacementI [rule_format]: +lemma (in M_trivial) strong_replacementI [rule_format]: "[| \A[M]. separation(M, %u. \x[M]. x\A & P(x,u)) |] ==> strong_replacement(M,P)" apply (simp add: strong_replacement_def, clarify) @@ -645,7 +645,7 @@ is_Replace(M, A', %x y. P'(x,y), z')" by (simp add: is_Replace_def) -lemma (in M_triv_axioms) univalent_Replace_iff: +lemma (in M_trivial) univalent_Replace_iff: "[| M(A); univalent(M,A,P); !!x y. [| x\A; P(x,y) |] ==> M(y) |] ==> u \ Replace(A,P) <-> (\x. x\A & P(x,u))" @@ -654,7 +654,7 @@ done (*The last premise expresses that P takes M to M*) -lemma (in M_triv_axioms) strong_replacement_closed [intro,simp]: +lemma (in M_trivial) strong_replacement_closed [intro,simp]: "[| strong_replacement(M,P); M(A); univalent(M,A,P); !!x y. [| x\A; P(x,y) |] ==> M(y) |] ==> M(Replace(A,P))" apply (simp add: strong_replacement_def) @@ -666,7 +666,7 @@ apply (blast dest: transM) done -lemma (in M_triv_axioms) Replace_abs: +lemma (in M_trivial) Replace_abs: "[| M(A); M(z); univalent(M,A,P); strong_replacement(M, P); !!x y. [| x\A; P(x,y) |] ==> M(y) |] ==> is_Replace(M,A,P,z) <-> z = Replace(A,P)" @@ -683,7 +683,7 @@ nonconstructible set. So we cannot assume that M(X) implies M(RepFun(X,f)) even for f : M -> M. *) -lemma (in M_triv_axioms) RepFun_closed: +lemma (in M_trivial) RepFun_closed: "[| strong_replacement(M, \x y. y = f(x)); M(A); \x\A. M(f(x)) |] ==> M(RepFun(A,f))" apply (simp add: RepFun_def) @@ -696,7 +696,7 @@ text{*Better than @{text RepFun_closed} when having the formula @{term "x\A"} makes relativization easier.*} -lemma (in M_triv_axioms) RepFun_closed2: +lemma (in M_trivial) RepFun_closed2: "[| strong_replacement(M, \x y. x\A & y = f(x)); M(A); \x\A. M(f(x)) |] ==> M(RepFun(A, %x. f(x)))" apply (simp add: RepFun_def) @@ -712,20 +712,20 @@ \p[M]. p \ z <-> (\u[M]. \v[M]. u\A & pair(M,u,v,p) & is_b(u,v))" -lemma (in M_triv_axioms) lam_closed: +lemma (in M_trivial) lam_closed: "[| strong_replacement(M, \x y. y = ); M(A); \x\A. M(b(x)) |] ==> M(\x\A. b(x))" by (simp add: lam_def, blast intro: RepFun_closed dest: transM) text{*Better than @{text lam_closed}: has the formula @{term "x\A"}*} -lemma (in M_triv_axioms) lam_closed2: +lemma (in M_trivial) lam_closed2: "[|strong_replacement(M, \x y. x\A & y = \x, b(x)\); M(A); \m[M]. m\A --> M(b(m))|] ==> M(Lambda(A,b))" apply (simp add: lam_def) apply (blast intro: RepFun_closed2 dest: transM) done -lemma (in M_triv_axioms) lambda_abs2 [simp]: +lemma (in M_trivial) lambda_abs2 [simp]: "[| strong_replacement(M, \x y. x\A & y = \x, b(x)\); Relativize1(M,A,is_b,b); M(A); \m[M]. m\A --> M(b(m)); M(z) |] ==> is_lambda(M,A,is_b,z) <-> z = Lambda(A,b)" @@ -744,7 +744,7 @@ is_lambda(M, A', %x y. is_b'(x,y), z')" by (simp add: is_lambda_def) -lemma (in M_triv_axioms) image_abs [simp]: +lemma (in M_trivial) image_abs [simp]: "[| M(r); M(A); M(z) |] ==> image(M,r,A,z) <-> z = r``A" apply (simp add: image_def) apply (rule iffI) @@ -754,13 +754,13 @@ text{*What about @{text Pow_abs}? Powerset is NOT absolute! This result is one direction of absoluteness.*} -lemma (in M_triv_axioms) powerset_Pow: +lemma (in M_trivial) powerset_Pow: "powerset(M, x, Pow(x))" by (simp add: powerset_def) text{*But we can't prove that the powerset in @{text M} includes the real powerset.*} -lemma (in M_triv_axioms) powerset_imp_subset_Pow: +lemma (in M_trivial) powerset_imp_subset_Pow: "[| powerset(M,x,y); M(y) |] ==> y <= Pow(x)" apply (simp add: powerset_def) apply (blast dest: transM) @@ -768,22 +768,22 @@ subsubsection{*Absoluteness for the Natural Numbers*} -lemma (in M_triv_axioms) nat_into_M [intro]: +lemma (in M_trivial) nat_into_M [intro]: "n \ nat ==> M(n)" by (induct n rule: nat_induct, simp_all) -lemma (in M_triv_axioms) nat_case_closed [intro,simp]: +lemma (in M_trivial) nat_case_closed [intro,simp]: "[|M(k); M(a); \m[M]. M(b(m))|] ==> M(nat_case(a,b,k))" apply (case_tac "k=0", simp) apply (case_tac "\m. k = succ(m)", force) apply (simp add: nat_case_def) done -lemma (in M_triv_axioms) quasinat_abs [simp]: +lemma (in M_trivial) quasinat_abs [simp]: "M(z) ==> is_quasinat(M,z) <-> quasinat(z)" by (auto simp add: is_quasinat_def quasinat_def) -lemma (in M_triv_axioms) nat_case_abs [simp]: +lemma (in M_trivial) nat_case_abs [simp]: "[| relativize1(M,is_b,b); M(k); M(z) |] ==> is_nat_case(M,a,is_b,k,z) <-> z = nat_case(a,b,k)" apply (case_tac "quasinat(k)") @@ -808,26 +808,26 @@ subsection{*Absoluteness for Ordinals*} text{*These results constitute Theorem IV 5.1 of Kunen (page 126).*} -lemma (in M_triv_axioms) lt_closed: +lemma (in M_trivial) lt_closed: "[| j M(j)" by (blast dest: ltD intro: transM) -lemma (in M_triv_axioms) transitive_set_abs [simp]: +lemma (in M_trivial) transitive_set_abs [simp]: "M(a) ==> transitive_set(M,a) <-> Transset(a)" by (simp add: transitive_set_def Transset_def) -lemma (in M_triv_axioms) ordinal_abs [simp]: +lemma (in M_trivial) ordinal_abs [simp]: "M(a) ==> ordinal(M,a) <-> Ord(a)" by (simp add: ordinal_def Ord_def) -lemma (in M_triv_axioms) limit_ordinal_abs [simp]: +lemma (in M_trivial) limit_ordinal_abs [simp]: "M(a) ==> limit_ordinal(M,a) <-> Limit(a)" apply (unfold Limit_def limit_ordinal_def) apply (simp add: Ord_0_lt_iff) apply (simp add: lt_def, blast) done -lemma (in M_triv_axioms) successor_ordinal_abs [simp]: +lemma (in M_trivial) successor_ordinal_abs [simp]: "M(a) ==> successor_ordinal(M,a) <-> Ord(a) & (\b[M]. a = succ(b))" apply (simp add: successor_ordinal_def, safe) apply (drule Ord_cases_disj, auto) @@ -840,7 +840,7 @@ lemma naturals_not_limit: "a \ nat ==> ~ Limit(a)" by (induct a rule: nat_induct, auto) -lemma (in M_triv_axioms) finite_ordinal_abs [simp]: +lemma (in M_trivial) finite_ordinal_abs [simp]: "M(a) ==> finite_ordinal(M,a) <-> a \ nat" apply (simp add: finite_ordinal_def) apply (blast intro: finite_Ord_is_nat intro: nat_into_Ord @@ -856,21 +856,21 @@ apply (erule nat_le_Limit) done -lemma (in M_triv_axioms) omega_abs [simp]: +lemma (in M_trivial) omega_abs [simp]: "M(a) ==> omega(M,a) <-> a = nat" apply (simp add: omega_def) apply (blast intro: Limit_non_Limit_implies_nat dest: naturals_not_limit) done -lemma (in M_triv_axioms) number1_abs [simp]: +lemma (in M_trivial) number1_abs [simp]: "M(a) ==> number1(M,a) <-> a = 1" by (simp add: number1_def) -lemma (in M_triv_axioms) number2_abs [simp]: +lemma (in M_trivial) number2_abs [simp]: "M(a) ==> number2(M,a) <-> a = succ(1)" by (simp add: number2_def) -lemma (in M_triv_axioms) number3_abs [simp]: +lemma (in M_trivial) number3_abs [simp]: "M(a) ==> number3(M,a) <-> a = succ(succ(1))" by (simp add: number3_def) @@ -893,13 +893,13 @@ natnumber :: "[i=>o,i,i] => o" "natnumber(M,n,x) == natnumber_aux(M,n)`x = 1" - lemma (in M_triv_axioms) [simp]: + lemma (in M_trivial) [simp]: "natnumber(M,0,x) == x=0" *) subsection{*Some instances of separation and strong replacement*} -locale M_axioms = M_triv_axioms + +locale M_basic = M_trivial + assumes Inter_separation: "M(A) ==> separation(M, \x. \y[M]. y\A --> x\y)" and Diff_separation: @@ -960,7 +960,7 @@ (\fx[M]. \gx[M]. fun_apply(M,f,x,fx) & fun_apply(M,g,x,gx) & fx \ gx))" -lemma (in M_axioms) cartprod_iff_lemma: +lemma (in M_basic) cartprod_iff_lemma: "[| M(C); \u[M]. u \ C <-> (\x\A. \y\B. u = {{x}, {x,y}}); powerset(M, A \ B, p1); powerset(M, p1, p2); M(p2) |] ==> C = {u \ p2 . \x\A. \y\B. u = {{x}, {x,y}}}" @@ -973,7 +973,7 @@ apply (frule transM, assumption, force) done -lemma (in M_axioms) cartprod_iff: +lemma (in M_basic) cartprod_iff: "[| M(A); M(B); M(C) |] ==> cartprod(M,A,B,C) <-> (\p1 p2. M(p1) & M(p2) & powerset(M,A Un B,p1) & powerset(M,p1,p2) & @@ -991,7 +991,7 @@ apply (blast intro: cartprod_iff_lemma) done -lemma (in M_axioms) cartprod_closed_lemma: +lemma (in M_basic) cartprod_closed_lemma: "[| M(A); M(B) |] ==> \C[M]. cartprod(M,A,B,C)" apply (simp del: cartprod_abs add: cartprod_iff) apply (insert power_ax, simp add: power_ax_def) @@ -1008,38 +1008,38 @@ text{*All the lemmas above are necessary because Powerset is not absolute. I should have used Replacement instead!*} -lemma (in M_axioms) cartprod_closed [intro,simp]: +lemma (in M_basic) cartprod_closed [intro,simp]: "[| M(A); M(B) |] ==> M(A*B)" by (frule cartprod_closed_lemma, assumption, force) -lemma (in M_axioms) sum_closed [intro,simp]: +lemma (in M_basic) sum_closed [intro,simp]: "[| M(A); M(B) |] ==> M(A+B)" by (simp add: sum_def) -lemma (in M_axioms) sum_abs [simp]: +lemma (in M_basic) sum_abs [simp]: "[| M(A); M(B); M(Z) |] ==> is_sum(M,A,B,Z) <-> (Z = A+B)" by (simp add: is_sum_def sum_def singleton_0 nat_into_M) -lemma (in M_triv_axioms) Inl_in_M_iff [iff]: +lemma (in M_trivial) Inl_in_M_iff [iff]: "M(Inl(a)) <-> M(a)" by (simp add: Inl_def) -lemma (in M_triv_axioms) Inl_abs [simp]: +lemma (in M_trivial) Inl_abs [simp]: "M(Z) ==> is_Inl(M,a,Z) <-> (Z = Inl(a))" by (simp add: is_Inl_def Inl_def) -lemma (in M_triv_axioms) Inr_in_M_iff [iff]: +lemma (in M_trivial) Inr_in_M_iff [iff]: "M(Inr(a)) <-> M(a)" by (simp add: Inr_def) -lemma (in M_triv_axioms) Inr_abs [simp]: +lemma (in M_trivial) Inr_abs [simp]: "M(Z) ==> is_Inr(M,a,Z) <-> (Z = Inr(a))" by (simp add: is_Inr_def Inr_def) subsubsection {*converse of a relation*} -lemma (in M_axioms) M_converse_iff: +lemma (in M_basic) M_converse_iff: "M(r) ==> converse(r) = {z \ Union(Union(r)) * Union(Union(r)). @@ -1050,13 +1050,13 @@ apply (blast dest: transM) done -lemma (in M_axioms) converse_closed [intro,simp]: +lemma (in M_basic) converse_closed [intro,simp]: "M(r) ==> M(converse(r))" apply (simp add: M_converse_iff) apply (insert converse_separation [of r], simp) done -lemma (in M_axioms) converse_abs [simp]: +lemma (in M_basic) converse_abs [simp]: "[| M(r); M(z) |] ==> is_converse(M,r,z) <-> z = converse(r)" apply (simp add: is_converse_def) apply (rule iffI) @@ -1069,105 +1069,105 @@ subsubsection {*image, preimage, domain, range*} -lemma (in M_axioms) image_closed [intro,simp]: +lemma (in M_basic) image_closed [intro,simp]: "[| M(A); M(r) |] ==> M(r``A)" apply (simp add: image_iff_Collect) apply (insert image_separation [of A r], simp) done -lemma (in M_axioms) vimage_abs [simp]: +lemma (in M_basic) vimage_abs [simp]: "[| M(r); M(A); M(z) |] ==> pre_image(M,r,A,z) <-> z = r-``A" apply (simp add: pre_image_def) apply (rule iffI) apply (blast intro!: equalityI dest: transM, blast) done -lemma (in M_axioms) vimage_closed [intro,simp]: +lemma (in M_basic) vimage_closed [intro,simp]: "[| M(A); M(r) |] ==> M(r-``A)" by (simp add: vimage_def) subsubsection{*Domain, range and field*} -lemma (in M_axioms) domain_abs [simp]: +lemma (in M_basic) domain_abs [simp]: "[| M(r); M(z) |] ==> is_domain(M,r,z) <-> z = domain(r)" apply (simp add: is_domain_def) apply (blast intro!: equalityI dest: transM) done -lemma (in M_axioms) domain_closed [intro,simp]: +lemma (in M_basic) domain_closed [intro,simp]: "M(r) ==> M(domain(r))" apply (simp add: domain_eq_vimage) done -lemma (in M_axioms) range_abs [simp]: +lemma (in M_basic) range_abs [simp]: "[| M(r); M(z) |] ==> is_range(M,r,z) <-> z = range(r)" apply (simp add: is_range_def) apply (blast intro!: equalityI dest: transM) done -lemma (in M_axioms) range_closed [intro,simp]: +lemma (in M_basic) range_closed [intro,simp]: "M(r) ==> M(range(r))" apply (simp add: range_eq_image) done -lemma (in M_axioms) field_abs [simp]: +lemma (in M_basic) field_abs [simp]: "[| M(r); M(z) |] ==> is_field(M,r,z) <-> z = field(r)" by (simp add: domain_closed range_closed is_field_def field_def) -lemma (in M_axioms) field_closed [intro,simp]: +lemma (in M_basic) field_closed [intro,simp]: "M(r) ==> M(field(r))" by (simp add: domain_closed range_closed Un_closed field_def) subsubsection{*Relations, functions and application*} -lemma (in M_axioms) relation_abs [simp]: +lemma (in M_basic) relation_abs [simp]: "M(r) ==> is_relation(M,r) <-> relation(r)" apply (simp add: is_relation_def relation_def) apply (blast dest!: bspec dest: pair_components_in_M)+ done -lemma (in M_axioms) function_abs [simp]: +lemma (in M_basic) function_abs [simp]: "M(r) ==> is_function(M,r) <-> function(r)" apply (simp add: is_function_def function_def, safe) apply (frule transM, assumption) apply (blast dest: pair_components_in_M)+ done -lemma (in M_axioms) apply_closed [intro,simp]: +lemma (in M_basic) apply_closed [intro,simp]: "[|M(f); M(a)|] ==> M(f`a)" by (simp add: apply_def) -lemma (in M_axioms) apply_abs [simp]: +lemma (in M_basic) apply_abs [simp]: "[| M(f); M(x); M(y) |] ==> fun_apply(M,f,x,y) <-> f`x = y" apply (simp add: fun_apply_def apply_def, blast) done -lemma (in M_axioms) typed_function_abs [simp]: +lemma (in M_basic) typed_function_abs [simp]: "[| M(A); M(f) |] ==> typed_function(M,A,B,f) <-> f \ A -> B" apply (auto simp add: typed_function_def relation_def Pi_iff) apply (blast dest: pair_components_in_M)+ done -lemma (in M_axioms) injection_abs [simp]: +lemma (in M_basic) injection_abs [simp]: "[| M(A); M(f) |] ==> injection(M,A,B,f) <-> f \ inj(A,B)" apply (simp add: injection_def apply_iff inj_def apply_closed) apply (blast dest: transM [of _ A]) done -lemma (in M_axioms) surjection_abs [simp]: +lemma (in M_basic) surjection_abs [simp]: "[| M(A); M(B); M(f) |] ==> surjection(M,A,B,f) <-> f \ surj(A,B)" by (simp add: surjection_def surj_def) -lemma (in M_axioms) bijection_abs [simp]: +lemma (in M_basic) bijection_abs [simp]: "[| M(A); M(B); M(f) |] ==> bijection(M,A,B,f) <-> f \ bij(A,B)" by (simp add: bijection_def bij_def) subsubsection{*Composition of relations*} -lemma (in M_axioms) M_comp_iff: +lemma (in M_basic) M_comp_iff: "[| M(r); M(s) |] ==> r O s = {xz \ domain(s) * range(r). @@ -1179,13 +1179,13 @@ apply (blast dest: transM)+ done -lemma (in M_axioms) comp_closed [intro,simp]: +lemma (in M_basic) comp_closed [intro,simp]: "[| M(r); M(s) |] ==> M(r O s)" apply (simp add: M_comp_iff) apply (insert comp_separation [of r s], simp) done -lemma (in M_axioms) composition_abs [simp]: +lemma (in M_basic) composition_abs [simp]: "[| M(r); M(s); M(t) |] ==> composition(M,r,s,t) <-> t = r O s" apply safe @@ -1200,7 +1200,7 @@ done text{*no longer needed*} -lemma (in M_axioms) restriction_is_function: +lemma (in M_basic) restriction_is_function: "[| restriction(M,f,A,z); function(f); M(f); M(A); M(z) |] ==> function(z)" apply (rotate_tac 1) @@ -1208,7 +1208,7 @@ apply (unfold function_def, blast) done -lemma (in M_axioms) restriction_abs [simp]: +lemma (in M_basic) restriction_abs [simp]: "[| M(f); M(A); M(z) |] ==> restriction(M,f,A,z) <-> z = restrict(f,A)" apply (simp add: ball_iff_equiv restriction_def restrict_def) @@ -1216,39 +1216,39 @@ done -lemma (in M_axioms) M_restrict_iff: +lemma (in M_basic) M_restrict_iff: "M(r) ==> restrict(r,A) = {z \ r . \x\A. \y[M]. z = \x, y\}" by (simp add: restrict_def, blast dest: transM) -lemma (in M_axioms) restrict_closed [intro,simp]: +lemma (in M_basic) restrict_closed [intro,simp]: "[| M(A); M(r) |] ==> M(restrict(r,A))" apply (simp add: M_restrict_iff) apply (insert restrict_separation [of A], simp) done -lemma (in M_axioms) Inter_abs [simp]: +lemma (in M_basic) Inter_abs [simp]: "[| M(A); M(z) |] ==> big_inter(M,A,z) <-> z = Inter(A)" apply (simp add: big_inter_def Inter_def) apply (blast intro!: equalityI dest: transM) done -lemma (in M_axioms) Inter_closed [intro,simp]: +lemma (in M_basic) Inter_closed [intro,simp]: "M(A) ==> M(Inter(A))" by (insert Inter_separation, simp add: Inter_def) -lemma (in M_axioms) Int_closed [intro,simp]: +lemma (in M_basic) Int_closed [intro,simp]: "[| M(A); M(B) |] ==> M(A Int B)" apply (subgoal_tac "M({A,B})") apply (frule Inter_closed, force+) done -lemma (in M_axioms) Diff_closed [intro,simp]: +lemma (in M_basic) Diff_closed [intro,simp]: "[|M(A); M(B)|] ==> M(A-B)" by (insert Diff_separation, simp add: Diff_def) subsubsection{*Some Facts About Separation Axioms*} -lemma (in M_axioms) separation_conj: +lemma (in M_basic) separation_conj: "[|separation(M,P); separation(M,Q)|] ==> separation(M, \z. P(z) & Q(z))" by (simp del: separation_closed add: separation_iff Collect_Int_Collect_eq [symmetric]) @@ -1262,24 +1262,24 @@ "A - Collect(A,P) = Collect(A, %x. ~ P(x))" by blast -lemma (in M_triv_axioms) Collect_rall_eq: +lemma (in M_trivial) Collect_rall_eq: "M(Y) ==> Collect(A, %x. \y[M]. y\Y --> P(x,y)) = (if Y=0 then A else (\y \ Y. {x \ A. P(x,y)}))" apply simp apply (blast intro!: equalityI dest: transM) done -lemma (in M_axioms) separation_disj: +lemma (in M_basic) separation_disj: "[|separation(M,P); separation(M,Q)|] ==> separation(M, \z. P(z) | Q(z))" by (simp del: separation_closed add: separation_iff Collect_Un_Collect_eq [symmetric]) -lemma (in M_axioms) separation_neg: +lemma (in M_basic) separation_neg: "separation(M,P) ==> separation(M, \z. ~P(z))" by (simp del: separation_closed add: separation_iff Diff_Collect_eq [symmetric]) -lemma (in M_axioms) separation_imp: +lemma (in M_basic) separation_imp: "[|separation(M,P); separation(M,Q)|] ==> separation(M, \z. P(z) --> Q(z))" by (simp add: separation_neg separation_disj not_disj_iff_imp [symmetric]) @@ -1287,7 +1287,7 @@ text{*This result is a hint of how little can be done without the Reflection Theorem. The quantifier has to be bounded by a set. We also need another instance of Separation!*} -lemma (in M_axioms) separation_rall: +lemma (in M_basic) separation_rall: "[|M(Y); \y[M]. separation(M, \x. P(x,y)); \z[M]. strong_replacement(M, \x y. y = {u \ z . P(u,x)})|] ==> separation(M, \x. \y[M]. y\Y --> P(x,y))" @@ -1300,7 +1300,7 @@ subsubsection{*Functions and function space*} text{*M contains all finite functions*} -lemma (in M_axioms) finite_fun_closed_lemma [rule_format]: +lemma (in M_basic) finite_fun_closed_lemma [rule_format]: "[| n \ nat; M(A) |] ==> \f \ n -> A. M(f)" apply (induct_tac n, simp) apply (rule ballI) @@ -1312,13 +1312,13 @@ apply (blast intro: apply_funtype transM restrict_type2) done -lemma (in M_axioms) finite_fun_closed [rule_format]: +lemma (in M_basic) finite_fun_closed [rule_format]: "[| f \ n -> A; n \ nat; M(A) |] ==> M(f)" by (blast intro: finite_fun_closed_lemma) text{*The assumption @{term "M(A->B)"} is unusual, but essential: in all but trivial cases, A->B cannot be expected to belong to @{term M}.*} -lemma (in M_axioms) is_funspace_abs [simp]: +lemma (in M_basic) is_funspace_abs [simp]: "[|M(A); M(B); M(F); M(A->B)|] ==> is_funspace(M,A,B,F) <-> F = A->B"; apply (simp add: is_funspace_def) apply (rule iffI) @@ -1327,7 +1327,7 @@ apply simp_all done -lemma (in M_axioms) succ_fun_eq2: +lemma (in M_basic) succ_fun_eq2: "[|M(B); M(n->B)|] ==> succ(n) -> B = \{z. p \ (n->B)*B, \f[M]. \b[M]. p = & z = {cons(, f)}}" @@ -1335,7 +1335,7 @@ apply (blast dest: transM) done -lemma (in M_axioms) funspace_succ: +lemma (in M_basic) funspace_succ: "[|M(n); M(B); M(n->B) |] ==> M(succ(n) -> B)" apply (insert funspace_succ_replacement [of n], simp) apply (force simp add: succ_fun_eq2 univalent_def) @@ -1343,7 +1343,7 @@ text{*@{term M} contains all finite function spaces. Needed to prove the absoluteness of transitive closure.*} -lemma (in M_axioms) finite_funspace_closed [intro,simp]: +lemma (in M_basic) finite_funspace_closed [intro,simp]: "[|n\nat; M(B)|] ==> M(n->B)" apply (induct_tac n, simp) apply (simp add: funspace_succ nat_into_M) @@ -1368,37 +1368,37 @@ "is_or(M,a,b,z) == (number1(M,a) & number1(M,z)) | (~number1(M,a) & z=b)" -lemma (in M_triv_axioms) bool_of_o_abs [simp]: +lemma (in M_trivial) bool_of_o_abs [simp]: "M(z) ==> is_bool_of_o(M,P,z) <-> z = bool_of_o(P)" by (simp add: is_bool_of_o_def bool_of_o_def) -lemma (in M_triv_axioms) not_abs [simp]: +lemma (in M_trivial) not_abs [simp]: "[| M(a); M(z)|] ==> is_not(M,a,z) <-> z = not(a)" by (simp add: Bool.not_def cond_def is_not_def) -lemma (in M_triv_axioms) and_abs [simp]: +lemma (in M_trivial) and_abs [simp]: "[| M(a); M(b); M(z)|] ==> is_and(M,a,b,z) <-> z = a and b" by (simp add: Bool.and_def cond_def is_and_def) -lemma (in M_triv_axioms) or_abs [simp]: +lemma (in M_trivial) or_abs [simp]: "[| M(a); M(b); M(z)|] ==> is_or(M,a,b,z) <-> z = a or b" by (simp add: Bool.or_def cond_def is_or_def) -lemma (in M_triv_axioms) bool_of_o_closed [intro,simp]: +lemma (in M_trivial) bool_of_o_closed [intro,simp]: "M(bool_of_o(P))" by (simp add: bool_of_o_def) -lemma (in M_triv_axioms) and_closed [intro,simp]: +lemma (in M_trivial) and_closed [intro,simp]: "[| M(p); M(q) |] ==> M(p and q)" by (simp add: and_def cond_def) -lemma (in M_triv_axioms) or_closed [intro,simp]: +lemma (in M_trivial) or_closed [intro,simp]: "[| M(p); M(q) |] ==> M(p or q)" by (simp add: or_def cond_def) -lemma (in M_triv_axioms) not_closed [intro,simp]: +lemma (in M_trivial) not_closed [intro,simp]: "M(p) ==> M(not(p))" by (simp add: Bool.not_def cond_def) @@ -1416,16 +1416,16 @@ "is_Cons(M,a,l,Z) == \p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" -lemma (in M_triv_axioms) Nil_in_M [intro,simp]: "M(Nil)" +lemma (in M_trivial) Nil_in_M [intro,simp]: "M(Nil)" by (simp add: Nil_def) -lemma (in M_triv_axioms) Nil_abs [simp]: "M(Z) ==> is_Nil(M,Z) <-> (Z = Nil)" +lemma (in M_trivial) Nil_abs [simp]: "M(Z) ==> is_Nil(M,Z) <-> (Z = Nil)" by (simp add: is_Nil_def Nil_def) -lemma (in M_triv_axioms) Cons_in_M_iff [iff]: "M(Cons(a,l)) <-> M(a) & M(l)" +lemma (in M_trivial) Cons_in_M_iff [iff]: "M(Cons(a,l)) <-> M(a) & M(l)" by (simp add: Cons_def) -lemma (in M_triv_axioms) Cons_abs [simp]: +lemma (in M_trivial) Cons_abs [simp]: "[|M(a); M(l); M(Z)|] ==> is_Cons(M,a,l,Z) <-> (Z = Cons(a,l))" by (simp add: is_Cons_def Cons_def) @@ -1499,18 +1499,18 @@ "xs \ list(A) ==>list_case'(a,b,xs) = list_case(a,b,xs)" by (erule list.cases, simp_all) -lemma (in M_axioms) list_case'_closed [intro,simp]: +lemma (in M_basic) list_case'_closed [intro,simp]: "[|M(k); M(a); \x[M]. \y[M]. M(b(x,y))|] ==> M(list_case'(a,b,k))" apply (case_tac "quasilist(k)") apply (simp add: quasilist_def, force) apply (simp add: non_list_case) done -lemma (in M_triv_axioms) quasilist_abs [simp]: +lemma (in M_trivial) quasilist_abs [simp]: "M(z) ==> is_quasilist(M,z) <-> quasilist(z)" by (auto simp add: is_quasilist_def quasilist_def) -lemma (in M_triv_axioms) list_case_abs [simp]: +lemma (in M_trivial) list_case_abs [simp]: "[| relativize2(M,is_b,b); M(k); M(z) |] ==> is_list_case(M,a,is_b,k,z) <-> z = list_case'(a,b,k)" apply (case_tac "quasilist(k)") @@ -1525,14 +1525,14 @@ subsubsection{*The Modified Operators @{term hd'} and @{term tl'}*} -lemma (in M_triv_axioms) is_hd_Nil: "is_hd(M,[],Z) <-> empty(M,Z)" +lemma (in M_trivial) is_hd_Nil: "is_hd(M,[],Z) <-> empty(M,Z)" by (simp add: is_hd_def) -lemma (in M_triv_axioms) is_hd_Cons: +lemma (in M_trivial) is_hd_Cons: "[|M(a); M(l)|] ==> is_hd(M,Cons(a,l),Z) <-> Z = a" by (force simp add: is_hd_def) -lemma (in M_triv_axioms) hd_abs [simp]: +lemma (in M_trivial) hd_abs [simp]: "[|M(x); M(y)|] ==> is_hd(M,x,y) <-> y = hd'(x)" apply (simp add: hd'_def) apply (intro impI conjI) @@ -1541,14 +1541,14 @@ apply (elim disjE exE, auto) done -lemma (in M_triv_axioms) is_tl_Nil: "is_tl(M,[],Z) <-> Z = []" +lemma (in M_trivial) is_tl_Nil: "is_tl(M,[],Z) <-> Z = []" by (simp add: is_tl_def) -lemma (in M_triv_axioms) is_tl_Cons: +lemma (in M_trivial) is_tl_Cons: "[|M(a); M(l)|] ==> is_tl(M,Cons(a,l),Z) <-> Z = l" by (force simp add: is_tl_def) -lemma (in M_triv_axioms) tl_abs [simp]: +lemma (in M_trivial) tl_abs [simp]: "[|M(x); M(y)|] ==> is_tl(M,x,y) <-> y = tl'(x)" apply (simp add: tl'_def) apply (intro impI conjI) @@ -1557,7 +1557,7 @@ apply (elim disjE exE, auto) done -lemma (in M_triv_axioms) relativize1_tl: "relativize1(M, is_tl(M), tl')" +lemma (in M_trivial) relativize1_tl: "relativize1(M, is_tl(M), tl')" by (simp add: relativize1_def) lemma hd'_Nil: "hd'([]) = 0" @@ -1577,7 +1577,7 @@ apply (simp_all add: tl'_Nil) done -lemma (in M_axioms) tl'_closed: "M(x) ==> M(tl'(x))" +lemma (in M_basic) tl'_closed: "M(x) ==> M(tl'(x))" apply (simp add: tl'_def) apply (force simp add: quasilist_def) done diff -r 7d6c9817c432 -r 1500a2e48d44 src/ZF/Constructible/Separation.thy --- a/src/ZF/Constructible/Separation.thy Tue Sep 10 16:47:17 2002 +0200 +++ b/src/ZF/Constructible/Separation.thy Tue Sep 10 16:51:31 2002 +0200 @@ -8,7 +8,7 @@ theory Separation = L_axioms + WF_absolute: -text{*This theory proves all instances needed for locale @{text "M_axioms"}*} +text{*This theory proves all instances needed for locale @{text "M_basic"}*} text{*Helps us solve for de Bruijn indices!*} lemma nth_ConsI: "[|nth(n,l) = x; n \ nat|] ==> nth(succ(n), Cons(a,l)) = x" @@ -470,12 +470,12 @@ done -subsection{*Instantiating the locale @{text M_axioms}*} +subsection{*Instantiating the locale @{text M_basic}*} text{*Separation (and Strong Replacement) for basic set-theoretic constructions such as intersection, Cartesian Product and image.*} -lemma M_axioms_axioms_L: "M_axioms_axioms(L)" - apply (rule M_axioms_axioms.intro) +lemma M_basic_axioms_L: "M_basic_axioms(L)" + apply (rule M_basic_axioms.intro) apply (assumption | rule Inter_separation Diff_separation cartprod_separation image_separation converse_separation restrict_separation @@ -485,114 +485,114 @@ omap_replacement is_recfun_separation)+ done -theorem M_axioms_L: "PROP M_axioms(L)" -by (rule M_axioms.intro [OF M_triv_axioms_L M_axioms_axioms_L]) +theorem M_basic_L: "PROP M_basic(L)" +by (rule M_basic.intro [OF M_trivial_L M_basic_axioms_L]) -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] +lemmas cartprod_iff = M_basic.cartprod_iff [OF M_basic_L] + and cartprod_closed = M_basic.cartprod_closed [OF M_basic_L] + and sum_closed = M_basic.sum_closed [OF M_basic_L] + and M_converse_iff = M_basic.M_converse_iff [OF M_basic_L] + and converse_closed = M_basic.converse_closed [OF M_basic_L] + and converse_abs = M_basic.converse_abs [OF M_basic_L] + and image_closed = M_basic.image_closed [OF M_basic_L] + and vimage_abs = M_basic.vimage_abs [OF M_basic_L] + and vimage_closed = M_basic.vimage_closed [OF M_basic_L] + and domain_abs = M_basic.domain_abs [OF M_basic_L] + and domain_closed = M_basic.domain_closed [OF M_basic_L] + and range_abs = M_basic.range_abs [OF M_basic_L] + and range_closed = M_basic.range_closed [OF M_basic_L] + and field_abs = M_basic.field_abs [OF M_basic_L] + and field_closed = M_basic.field_closed [OF M_basic_L] + and relation_abs = M_basic.relation_abs [OF M_basic_L] + and function_abs = M_basic.function_abs [OF M_basic_L] + and apply_closed = M_basic.apply_closed [OF M_basic_L] + and apply_abs = M_basic.apply_abs [OF M_basic_L] + and typed_function_abs = M_basic.typed_function_abs [OF M_basic_L] + and injection_abs = M_basic.injection_abs [OF M_basic_L] + and surjection_abs = M_basic.surjection_abs [OF M_basic_L] + and bijection_abs = M_basic.bijection_abs [OF M_basic_L] + and M_comp_iff = M_basic.M_comp_iff [OF M_basic_L] + and comp_closed = M_basic.comp_closed [OF M_basic_L] + and composition_abs = M_basic.composition_abs [OF M_basic_L] + and restriction_is_function = M_basic.restriction_is_function [OF M_basic_L] + and restriction_abs = M_basic.restriction_abs [OF M_basic_L] + and M_restrict_iff = M_basic.M_restrict_iff [OF M_basic_L] + and restrict_closed = M_basic.restrict_closed [OF M_basic_L] + and Inter_abs = M_basic.Inter_abs [OF M_basic_L] + and Inter_closed = M_basic.Inter_closed [OF M_basic_L] + and Int_closed = M_basic.Int_closed [OF M_basic_L] + and finite_fun_closed = M_basic.finite_fun_closed [OF M_basic_L] + and is_funspace_abs = M_basic.is_funspace_abs [OF M_basic_L] + and succ_fun_eq2 = M_basic.succ_fun_eq2 [OF M_basic_L] + and funspace_succ = M_basic.funspace_succ [OF M_basic_L] + and finite_funspace_closed = M_basic.finite_funspace_closed [OF M_basic_L] -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] +lemmas is_recfun_equal = M_basic.is_recfun_equal [OF M_basic_L] + and is_recfun_cut = M_basic.is_recfun_cut [OF M_basic_L] + and is_recfun_functional = M_basic.is_recfun_functional [OF M_basic_L] + and is_recfun_relativize = M_basic.is_recfun_relativize [OF M_basic_L] + and is_recfun_restrict = M_basic.is_recfun_restrict [OF M_basic_L] + and univalent_is_recfun = M_basic.univalent_is_recfun [OF M_basic_L] + and exists_is_recfun_indstep = M_basic.exists_is_recfun_indstep [OF M_basic_L] + and wellfounded_exists_is_recfun = M_basic.wellfounded_exists_is_recfun [OF M_basic_L] + and wf_exists_is_recfun = M_basic.wf_exists_is_recfun [OF M_basic_L] + and is_recfun_abs = M_basic.is_recfun_abs [OF M_basic_L] + and irreflexive_abs = M_basic.irreflexive_abs [OF M_basic_L] + and transitive_rel_abs = M_basic.transitive_rel_abs [OF M_basic_L] + and linear_rel_abs = M_basic.linear_rel_abs [OF M_basic_L] + and wellordered_is_trans_on = M_basic.wellordered_is_trans_on [OF M_basic_L] + and wellordered_is_linear = M_basic.wellordered_is_linear [OF M_basic_L] + and wellordered_is_wellfounded_on = M_basic.wellordered_is_wellfounded_on [OF M_basic_L] + and wellfounded_imp_wellfounded_on = M_basic.wellfounded_imp_wellfounded_on [OF M_basic_L] + and wellfounded_on_subset_A = M_basic.wellfounded_on_subset_A [OF M_basic_L] + and wellfounded_on_iff_wellfounded = M_basic.wellfounded_on_iff_wellfounded [OF M_basic_L] + and wellfounded_on_imp_wellfounded = M_basic.wellfounded_on_imp_wellfounded [OF M_basic_L] + and wellfounded_on_field_imp_wellfounded = M_basic.wellfounded_on_field_imp_wellfounded [OF M_basic_L] + and wellfounded_iff_wellfounded_on_field = M_basic.wellfounded_iff_wellfounded_on_field [OF M_basic_L] + and wellfounded_induct = M_basic.wellfounded_induct [OF M_basic_L] + and wellfounded_on_induct = M_basic.wellfounded_on_induct [OF M_basic_L] + and wellfounded_on_induct2 = M_basic.wellfounded_on_induct2 [OF M_basic_L] + and linear_imp_relativized = M_basic.linear_imp_relativized [OF M_basic_L] + and trans_on_imp_relativized = M_basic.trans_on_imp_relativized [OF M_basic_L] + and wf_on_imp_relativized = M_basic.wf_on_imp_relativized [OF M_basic_L] + and wf_imp_relativized = M_basic.wf_imp_relativized [OF M_basic_L] + and well_ord_imp_relativized = M_basic.well_ord_imp_relativized [OF M_basic_L] + and order_isomorphism_abs = M_basic.order_isomorphism_abs [OF M_basic_L] + and pred_set_abs = M_basic.pred_set_abs [OF M_basic_L] -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] +lemmas pred_closed = M_basic.pred_closed [OF M_basic_L] + and membership_abs = M_basic.membership_abs [OF M_basic_L] + and M_Memrel_iff = M_basic.M_Memrel_iff [OF M_basic_L] + and Memrel_closed = M_basic.Memrel_closed [OF M_basic_L] + and wellordered_iso_predD = M_basic.wellordered_iso_predD [OF M_basic_L] + and wellordered_iso_pred_eq = M_basic.wellordered_iso_pred_eq [OF M_basic_L] + and wellfounded_on_asym = M_basic.wellfounded_on_asym [OF M_basic_L] + and wellordered_asym = M_basic.wellordered_asym [OF M_basic_L] + and ord_iso_pred_imp_lt = M_basic.ord_iso_pred_imp_lt [OF M_basic_L] + and obase_iff = M_basic.obase_iff [OF M_basic_L] + and omap_iff = M_basic.omap_iff [OF M_basic_L] + and omap_unique = M_basic.omap_unique [OF M_basic_L] + and omap_yields_Ord = M_basic.omap_yields_Ord [OF M_basic_L] + and otype_iff = M_basic.otype_iff [OF M_basic_L] + and otype_eq_range = M_basic.otype_eq_range [OF M_basic_L] + and Ord_otype = M_basic.Ord_otype [OF M_basic_L] + and domain_omap = M_basic.domain_omap [OF M_basic_L] + and omap_subset = M_basic.omap_subset [OF M_basic_L] + and omap_funtype = M_basic.omap_funtype [OF M_basic_L] + and wellordered_omap_bij = M_basic.wellordered_omap_bij [OF M_basic_L] + and omap_ord_iso = M_basic.omap_ord_iso [OF M_basic_L] + and Ord_omap_image_pred = M_basic.Ord_omap_image_pred [OF M_basic_L] + and restrict_omap_ord_iso = M_basic.restrict_omap_ord_iso [OF M_basic_L] + and obase_equals = M_basic.obase_equals [OF M_basic_L] + and omap_ord_iso_otype = M_basic.omap_ord_iso_otype [OF M_basic_L] + and obase_exists = M_basic.obase_exists [OF M_basic_L] + and omap_exists = M_basic.omap_exists [OF M_basic_L] + and otype_exists = M_basic.otype_exists [OF M_basic_L] + and omap_ord_iso_otype' = M_basic.omap_ord_iso_otype' [OF M_basic_L] + and ordertype_exists = M_basic.ordertype_exists [OF M_basic_L] + and relativized_imp_well_ord = M_basic.relativized_imp_well_ord [OF M_basic_L] + and well_ord_abs = M_basic.well_ord_abs [OF M_basic_L] declare cartprod_closed [intro, simp] declare sum_closed [intro, simp] diff -r 7d6c9817c432 -r 1500a2e48d44 src/ZF/Constructible/WF_absolute.thy --- a/src/ZF/Constructible/WF_absolute.thy Tue Sep 10 16:47:17 2002 +0200 +++ b/src/ZF/Constructible/WF_absolute.thy Tue Sep 10 16:51:31 2002 +0200 @@ -139,7 +139,7 @@ "tran_closure(M,r,t) == \s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" -lemma (in M_axioms) rtran_closure_mem_iff: +lemma (in M_basic) rtran_closure_mem_iff: "[|M(A); M(r); M(p)|] ==> rtran_closure_mem(M,A,r,p) <-> (\n[M]. n\nat & @@ -149,7 +149,7 @@ by (simp add: rtran_closure_mem_def Ord_succ_mem_iff nat_0_le [THEN ltD]) -locale M_trancl = M_axioms + +locale M_trancl = M_basic + assumes rtrancl_separation: "[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))" and wellfounded_trancl_separation: diff -r 7d6c9817c432 -r 1500a2e48d44 src/ZF/Constructible/WFrec.thy --- a/src/ZF/Constructible/WFrec.thy Tue Sep 10 16:47:17 2002 +0200 +++ b/src/ZF/Constructible/WFrec.thy Tue Sep 10 16:51:31 2002 +0200 @@ -74,7 +74,7 @@ subsection{*Reworking of the Recursion Theory Within @{term M}*} -lemma (in M_axioms) is_recfun_separation': +lemma (in M_basic) is_recfun_separation': "[| f \ r -`` {a} \ range(f); g \ r -`` {b} \ range(g); M(r); M(f); M(g); M(a); M(b) |] ==> separation(M, \x. \ (\x, a\ \ r \ \x, b\ \ r \ f ` x = g ` x))" @@ -89,7 +89,7 @@ The last three M-premises are redundant because of @{term "M(r)"}, but without them we'd have to undertake more work to set up the induction formula.*} -lemma (in M_axioms) is_recfun_equal [rule_format]: +lemma (in M_basic) is_recfun_equal [rule_format]: "[|is_recfun(r,a,H,f); is_recfun(r,b,H,g); wellfounded(M,r); trans(r); M(f); M(g); M(r); M(x); M(a); M(b) |] @@ -112,7 +112,7 @@ apply (blast intro: transD sym) done -lemma (in M_axioms) is_recfun_cut: +lemma (in M_basic) is_recfun_cut: "[|is_recfun(r,a,H,f); is_recfun(r,b,H,g); wellfounded(M,r); trans(r); M(f); M(g); M(r); \ r |] @@ -124,7 +124,7 @@ apply (blast intro: is_recfun_equal transD dest: transM) done -lemma (in M_axioms) is_recfun_functional: +lemma (in M_basic) is_recfun_functional: "[|is_recfun(r,a,H,f); is_recfun(r,a,H,g); wellfounded(M,r); trans(r); M(f); M(g); M(r) |] ==> f=g" apply (rule fun_extension) @@ -133,7 +133,7 @@ done text{*Tells us that @{text is_recfun} can (in principle) be relativized.*} -lemma (in M_axioms) is_recfun_relativize: +lemma (in M_basic) is_recfun_relativize: "[| M(r); M(f); \x[M]. \g[M]. function(g) --> M(H(x,g)) |] ==> is_recfun(r,a,H,f) <-> (\z[M]. z \ f <-> @@ -154,7 +154,7 @@ apply (simp add: is_recfun_imp_function function_restrictI) done -lemma (in M_axioms) is_recfun_restrict: +lemma (in M_basic) is_recfun_restrict: "[| wellfounded(M,r); trans(r); is_recfun(r,x,H,f); \y,x\ \ r; M(r); M(f); \x[M]. \g[M]. function(g) --> M(H(x,g)) |] @@ -171,7 +171,7 @@ apply (blast intro: apply_recfun dest: transD) done -lemma (in M_axioms) restrict_Y_lemma: +lemma (in M_basic) restrict_Y_lemma: "[| wellfounded(M,r); trans(r); M(r); \x[M]. \g[M]. function(g) --> M(H(x,g)); M(Y); \b[M]. @@ -201,7 +201,7 @@ done text{*For typical applications of Replacement for recursive definitions*} -lemma (in M_axioms) univalent_is_recfun: +lemma (in M_basic) univalent_is_recfun: "[|wellfounded(M,r); trans(r); M(r)|] ==> univalent (M, A, \x p. \y[M]. p = \x,y\ & (\f[M]. is_recfun(r,x,H,f) & y = H(x,f)))" @@ -212,7 +212,7 @@ text{*Proof of the inductive step for @{text exists_is_recfun}, since we must prove two versions.*} -lemma (in M_axioms) exists_is_recfun_indstep: +lemma (in M_basic) exists_is_recfun_indstep: "[|\y. \y, a1\ \ r --> (\f[M]. is_recfun(r, y, H, f)); wellfounded(M,r); trans(r); M(r); M(a1); strong_replacement(M, \x z. @@ -245,7 +245,7 @@ text{*Relativized version, when we have the (currently weaker) premise @{term "wellfounded(M,r)"}*} -lemma (in M_axioms) wellfounded_exists_is_recfun: +lemma (in M_basic) wellfounded_exists_is_recfun: "[|wellfounded(M,r); trans(r); separation(M, \x. ~ (\f[M]. is_recfun(r, x, H, f))); strong_replacement(M, \x z. @@ -257,7 +257,7 @@ apply (rule exists_is_recfun_indstep, assumption+) done -lemma (in M_axioms) wf_exists_is_recfun [rule_format]: +lemma (in M_basic) wf_exists_is_recfun [rule_format]: "[|wf(r); trans(r); M(r); strong_replacement(M, \x z. \y[M]. \g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); @@ -291,7 +291,7 @@ strong_replacement(M, \x z. \y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))" -lemma (in M_axioms) is_recfun_abs: +lemma (in M_basic) is_recfun_abs: "[| \x[M]. \g[M]. function(g) --> M(H(x,g)); M(r); M(a); M(f); relativize2(M,MH,H) |] ==> M_is_recfun(M,MH,r,a,f) <-> is_recfun(r,a,H,f)" @@ -306,7 +306,7 @@ ==> M_is_recfun(M,MH,r,a,f) <-> M_is_recfun(M,MH',r',a',f')" by (simp add: M_is_recfun_def) -lemma (in M_axioms) is_wfrec_abs: +lemma (in M_basic) is_wfrec_abs: "[| \x[M]. \g[M]. function(g) --> M(H(x,g)); relativize2(M,MH,H); M(r); M(a); M(z) |] ==> is_wfrec(M,MH,r,a,z) <-> @@ -314,7 +314,7 @@ by (simp add: is_wfrec_def relativize2_def is_recfun_abs) text{*Relating @{term wfrec_replacement} to native constructs*} -lemma (in M_axioms) wfrec_replacement': +lemma (in M_basic) wfrec_replacement': "[|wfrec_replacement(M,MH,r); \x[M]. \g[M]. function(g) --> M(H(x,g)); relativize2(M,MH,H); M(r)|] @@ -381,7 +381,7 @@ fun_apply(M,f,j,fj) & fj = k" -locale M_ord_arith = M_axioms + +locale M_ord_arith = M_basic + assumes oadd_strong_replacement: "[| M(i); M(j) |] ==> strong_replacement(M, diff -r 7d6c9817c432 -r 1500a2e48d44 src/ZF/Constructible/Wellorderings.thy --- a/src/ZF/Constructible/Wellorderings.thy Tue Sep 10 16:47:17 2002 +0200 +++ b/src/ZF/Constructible/Wellorderings.thy Tue Sep 10 16:51:31 2002 +0200 @@ -49,63 +49,63 @@ subsubsection {*Trivial absoluteness proofs*} -lemma (in M_axioms) irreflexive_abs [simp]: +lemma (in M_basic) irreflexive_abs [simp]: "M(A) ==> irreflexive(M,A,r) <-> irrefl(A,r)" by (simp add: irreflexive_def irrefl_def) -lemma (in M_axioms) transitive_rel_abs [simp]: +lemma (in M_basic) transitive_rel_abs [simp]: "M(A) ==> transitive_rel(M,A,r) <-> trans[A](r)" by (simp add: transitive_rel_def trans_on_def) -lemma (in M_axioms) linear_rel_abs [simp]: +lemma (in M_basic) linear_rel_abs [simp]: "M(A) ==> linear_rel(M,A,r) <-> linear(A,r)" by (simp add: linear_rel_def linear_def) -lemma (in M_axioms) wellordered_is_trans_on: +lemma (in M_basic) wellordered_is_trans_on: "[| wellordered(M,A,r); M(A) |] ==> trans[A](r)" by (auto simp add: wellordered_def) -lemma (in M_axioms) wellordered_is_linear: +lemma (in M_basic) wellordered_is_linear: "[| wellordered(M,A,r); M(A) |] ==> linear(A,r)" by (auto simp add: wellordered_def) -lemma (in M_axioms) wellordered_is_wellfounded_on: +lemma (in M_basic) wellordered_is_wellfounded_on: "[| wellordered(M,A,r); M(A) |] ==> wellfounded_on(M,A,r)" by (auto simp add: wellordered_def) -lemma (in M_axioms) wellfounded_imp_wellfounded_on: +lemma (in M_basic) wellfounded_imp_wellfounded_on: "[| wellfounded(M,r); M(A) |] ==> wellfounded_on(M,A,r)" by (auto simp add: wellfounded_def wellfounded_on_def) -lemma (in M_axioms) wellfounded_on_subset_A: +lemma (in M_basic) wellfounded_on_subset_A: "[| wellfounded_on(M,A,r); B<=A |] ==> wellfounded_on(M,B,r)" by (simp add: wellfounded_on_def, blast) subsubsection {*Well-founded relations*} -lemma (in M_axioms) wellfounded_on_iff_wellfounded: +lemma (in M_basic) wellfounded_on_iff_wellfounded: "wellfounded_on(M,A,r) <-> wellfounded(M, r \ A*A)" apply (simp add: wellfounded_on_def wellfounded_def, safe) apply blast apply (drule_tac x=x in rspec, assumption, blast) done -lemma (in M_axioms) wellfounded_on_imp_wellfounded: +lemma (in M_basic) wellfounded_on_imp_wellfounded: "[|wellfounded_on(M,A,r); r \ A*A|] ==> wellfounded(M,r)" by (simp add: wellfounded_on_iff_wellfounded subset_Int_iff) -lemma (in M_axioms) wellfounded_on_field_imp_wellfounded: +lemma (in M_basic) wellfounded_on_field_imp_wellfounded: "wellfounded_on(M, field(r), r) ==> wellfounded(M,r)" by (simp add: wellfounded_def wellfounded_on_iff_wellfounded, fast) -lemma (in M_axioms) wellfounded_iff_wellfounded_on_field: +lemma (in M_basic) wellfounded_iff_wellfounded_on_field: "M(r) ==> wellfounded(M,r) <-> wellfounded_on(M, field(r), r)" by (blast intro: wellfounded_imp_wellfounded_on wellfounded_on_field_imp_wellfounded) (*Consider the least z in domain(r) such that P(z) does not hold...*) -lemma (in M_axioms) wellfounded_induct: +lemma (in M_basic) wellfounded_induct: "[| wellfounded(M,r); M(a); M(r); separation(M, \x. ~P(x)); \x. M(x) & (\y. \ r --> P(y)) --> P(x) |] ==> P(a)"; @@ -114,7 +114,7 @@ apply (blast dest: transM)+ done -lemma (in M_axioms) wellfounded_on_induct: +lemma (in M_basic) wellfounded_on_induct: "[| a\A; wellfounded_on(M,A,r); M(A); separation(M, \x. x\A --> ~P(x)); \x\A. M(x) & (\y\A. \ r --> P(y)) --> P(x) |] @@ -126,7 +126,7 @@ text{*The assumption @{term "r \ A*A"} justifies strengthening the induction hypothesis by removing the restriction to @{term A}.*} -lemma (in M_axioms) wellfounded_on_induct2: +lemma (in M_basic) wellfounded_on_induct2: "[| a\A; wellfounded_on(M,A,r); M(A); r \ A*A; separation(M, \x. x\A --> ~P(x)); \x\A. M(x) & (\y. \ r --> P(y)) --> P(x) |] @@ -136,27 +136,27 @@ subsubsection {*Kunen's lemma IV 3.14, page 123*} -lemma (in M_axioms) linear_imp_relativized: +lemma (in M_basic) linear_imp_relativized: "linear(A,r) ==> linear_rel(M,A,r)" by (simp add: linear_def linear_rel_def) -lemma (in M_axioms) trans_on_imp_relativized: +lemma (in M_basic) trans_on_imp_relativized: "trans[A](r) ==> transitive_rel(M,A,r)" by (unfold transitive_rel_def trans_on_def, blast) -lemma (in M_axioms) wf_on_imp_relativized: +lemma (in M_basic) wf_on_imp_relativized: "wf[A](r) ==> wellfounded_on(M,A,r)" apply (simp add: wellfounded_on_def wf_def wf_on_def, clarify) apply (drule_tac x=x in spec, blast) done -lemma (in M_axioms) wf_imp_relativized: +lemma (in M_basic) wf_imp_relativized: "wf(r) ==> wellfounded(M,r)" apply (simp add: wellfounded_def wf_def, clarify) apply (drule_tac x=x in spec, blast) done -lemma (in M_axioms) well_ord_imp_relativized: +lemma (in M_basic) well_ord_imp_relativized: "well_ord(A,r) ==> wellordered(M,A,r)" by (simp add: wellordered_def well_ord_def tot_ord_def part_ord_def linear_imp_relativized trans_on_imp_relativized wf_on_imp_relativized) @@ -164,24 +164,24 @@ subsection{* Relativized versions of order-isomorphisms and order types *} -lemma (in M_axioms) order_isomorphism_abs [simp]: +lemma (in M_basic) order_isomorphism_abs [simp]: "[| M(A); M(B); M(f) |] ==> order_isomorphism(M,A,r,B,s,f) <-> f \ ord_iso(A,r,B,s)" by (simp add: apply_closed order_isomorphism_def ord_iso_def) -lemma (in M_axioms) pred_set_abs [simp]: +lemma (in M_basic) pred_set_abs [simp]: "[| M(r); M(B) |] ==> pred_set(M,A,x,r,B) <-> B = Order.pred(A,x,r)" apply (simp add: pred_set_def Order.pred_def) apply (blast dest: transM) done -lemma (in M_axioms) pred_closed [intro,simp]: +lemma (in M_basic) pred_closed [intro,simp]: "[| M(A); M(r); M(x) |] ==> M(Order.pred(A,x,r))" apply (simp add: Order.pred_def) apply (insert pred_separation [of r x], simp) done -lemma (in M_axioms) membership_abs [simp]: +lemma (in M_basic) membership_abs [simp]: "[| M(r); M(A) |] ==> membership(M,A,r) <-> r = Memrel(A)" apply (simp add: membership_def Memrel_def, safe) apply (rule equalityI) @@ -194,14 +194,14 @@ apply auto done -lemma (in M_axioms) M_Memrel_iff: +lemma (in M_basic) M_Memrel_iff: "M(A) ==> Memrel(A) = {z \ A*A. \x[M]. \y[M]. z = \x,y\ & x \ y}" apply (simp add: Memrel_def) apply (blast dest: transM) done -lemma (in M_axioms) Memrel_closed [intro,simp]: +lemma (in M_basic) Memrel_closed [intro,simp]: "M(A) ==> M(Memrel(A))" apply (simp add: M_Memrel_iff) apply (insert Memrel_separation, simp) @@ -233,7 +233,7 @@ text{*Inductive argument for Kunen's Lemma 6.1, etc. Simple proof from Halmos, page 72*} -lemma (in M_axioms) wellordered_iso_subset_lemma: +lemma (in M_basic) wellordered_iso_subset_lemma: "[| wellordered(M,A,r); f \ ord_iso(A,r, A',r); A'<= A; y \ A; M(A); M(f); M(r) |] ==> ~ \ r" apply (unfold wellordered_def ord_iso_def) @@ -247,7 +247,7 @@ text{*Kunen's Lemma 6.1: there's no order-isomorphism to an initial segment of a well-ordering*} -lemma (in M_axioms) wellordered_iso_predD: +lemma (in M_basic) wellordered_iso_predD: "[| wellordered(M,A,r); f \ ord_iso(A, r, Order.pred(A,x,r), r); M(A); M(f); M(r) |] ==> x \ A" apply (rule notI) @@ -260,7 +260,7 @@ done -lemma (in M_axioms) wellordered_iso_pred_eq_lemma: +lemma (in M_basic) wellordered_iso_pred_eq_lemma: "[| f \ \Order.pred(A,y,r), r\ \ \Order.pred(A,x,r), r\; wellordered(M,A,r); x\A; y\A; M(A); M(f); M(r) |] ==> \ r" apply (frule wellordered_is_trans_on, assumption) @@ -273,7 +273,7 @@ text{*Simple consequence of Lemma 6.1*} -lemma (in M_axioms) wellordered_iso_pred_eq: +lemma (in M_basic) wellordered_iso_pred_eq: "[| wellordered(M,A,r); f \ ord_iso(Order.pred(A,a,r), r, Order.pred(A,c,r), r); M(A); M(f); M(r); a\A; c\A |] ==> a=c" @@ -285,21 +285,21 @@ apply (blast dest: wellordered_iso_pred_eq_lemma)+ done -lemma (in M_axioms) wellfounded_on_asym: +lemma (in M_basic) wellfounded_on_asym: "[| wellfounded_on(M,A,r); \r; a\A; x\A; M(A) |] ==> \r" apply (simp add: wellfounded_on_def) apply (drule_tac x="{x,a}" in rspec) apply (blast dest: transM)+ done -lemma (in M_axioms) wellordered_asym: +lemma (in M_basic) wellordered_asym: "[| wellordered(M,A,r); \r; a\A; x\A; M(A) |] ==> \r" by (simp add: wellordered_def, blast dest: wellfounded_on_asym) text{*Surely a shorter proof using lemmas in @{text Order}? Like @{text well_ord_iso_preserving}?*} -lemma (in M_axioms) ord_iso_pred_imp_lt: +lemma (in M_basic) ord_iso_pred_imp_lt: "[| f \ ord_iso(Order.pred(A,x,r), r, i, Memrel(i)); g \ ord_iso(Order.pred(A,y,r), r, j, Memrel(j)); wellordered(M,A,r); x \ A; y \ A; M(A); M(r); M(f); M(g); M(j); @@ -372,7 +372,7 @@ -lemma (in M_axioms) obase_iff: +lemma (in M_basic) obase_iff: "[| M(A); M(r); M(z) |] ==> obase(M,A,r,z) <-> z = {a\A. \x[M]. \g[M]. Ord(x) & @@ -387,7 +387,7 @@ text{*Can also be proved with the premise @{term "M(z)"} instead of @{term "M(f)"}, but that version is less useful.*} -lemma (in M_axioms) omap_iff: +lemma (in M_basic) omap_iff: "[| omap(M,A,r,f); M(A); M(r); M(f) |] ==> z \ f <-> (\a\A. \x[M]. \g[M]. z = & Ord(x) & @@ -400,18 +400,18 @@ apply (blast dest: transM)+ done -lemma (in M_axioms) omap_unique: +lemma (in M_basic) omap_unique: "[| omap(M,A,r,f); omap(M,A,r,f'); M(A); M(r); M(f); M(f') |] ==> f' = f" apply (rule equality_iffI) apply (simp add: omap_iff) done -lemma (in M_axioms) omap_yields_Ord: +lemma (in M_basic) omap_yields_Ord: "[| omap(M,A,r,f); \a,x\ \ f; M(a); M(x) |] ==> Ord(x)" apply (simp add: omap_def, blast) done -lemma (in M_axioms) otype_iff: +lemma (in M_basic) otype_iff: "[| otype(M,A,r,i); M(A); M(r); M(i) |] ==> x \ i <-> (M(x) & Ord(x) & @@ -423,7 +423,7 @@ apply (simp add: omap_iff, blast) done -lemma (in M_axioms) otype_eq_range: +lemma (in M_basic) otype_eq_range: "[| omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(i) |] ==> i = range(f)" apply (auto simp add: otype_def omap_iff) @@ -431,7 +431,7 @@ done -lemma (in M_axioms) Ord_otype: +lemma (in M_basic) Ord_otype: "[| otype(M,A,r,i); trans[A](r); M(A); M(r); M(i) |] ==> Ord(i)" apply (rotate_tac 1) apply (rule OrdI) @@ -452,7 +452,7 @@ apply (blast intro: restrict_ord_iso ord_iso_sym ltI dest: transM) done -lemma (in M_axioms) domain_omap: +lemma (in M_basic) domain_omap: "[| omap(M,A,r,f); obase(M,A,r,B); M(A); M(r); M(B); M(f) |] ==> domain(f) = B" apply (rotate_tac 2) @@ -461,7 +461,7 @@ apply (simp add: domain_iff omap_iff, blast) done -lemma (in M_axioms) omap_subset: +lemma (in M_basic) omap_subset: "[| omap(M,A,r,f); obase(M,A,r,B); otype(M,A,r,i); M(A); M(r); M(f); M(B); M(i) |] ==> f \ B * i" apply (rotate_tac 3, clarify) @@ -469,7 +469,7 @@ apply (force simp add: otype_iff) done -lemma (in M_axioms) omap_funtype: +lemma (in M_basic) omap_funtype: "[| omap(M,A,r,f); obase(M,A,r,B); otype(M,A,r,i); M(A); M(r); M(f); M(B); M(i) |] ==> f \ B -> i" apply (rotate_tac 3) @@ -478,7 +478,7 @@ done -lemma (in M_axioms) wellordered_omap_bij: +lemma (in M_basic) wellordered_omap_bij: "[| wellordered(M,A,r); omap(M,A,r,f); obase(M,A,r,B); otype(M,A,r,i); M(A); M(r); M(f); M(B); M(i) |] ==> f \ bij(B,i)" apply (insert omap_funtype [of A r f B i]) @@ -492,7 +492,7 @@ text{*This is not the final result: we must show @{term "oB(A,r) = A"}*} -lemma (in M_axioms) omap_ord_iso: +lemma (in M_basic) omap_ord_iso: "[| wellordered(M,A,r); omap(M,A,r,f); obase(M,A,r,B); otype(M,A,r,i); M(A); M(r); M(f); M(B); M(i) |] ==> f \ ord_iso(B,r,i,Memrel(i))" apply (rule ord_isoI) @@ -513,7 +513,7 @@ apply (blast dest: ord_iso_pred_imp_lt ltD elim: mem_asym) done -lemma (in M_axioms) Ord_omap_image_pred: +lemma (in M_basic) Ord_omap_image_pred: "[| wellordered(M,A,r); omap(M,A,r,f); obase(M,A,r,B); otype(M,A,r,i); M(A); M(r); M(f); M(B); M(i); b \ A |] ==> Ord(f `` Order.pred(A,b,r))" apply (frule wellordered_is_trans_on, assumption) @@ -539,7 +539,7 @@ apply (auto simp add: obase_iff) done -lemma (in M_axioms) restrict_omap_ord_iso: +lemma (in M_basic) restrict_omap_ord_iso: "[| wellordered(M,A,r); omap(M,A,r,f); obase(M,A,r,B); otype(M,A,r,i); D \ B; M(A); M(r); M(f); M(B); M(i) |] ==> restrict(f,D) \ (\D,r\ \ \f``D, Memrel(f``D)\)" @@ -550,7 +550,7 @@ apply (drule ord_iso_sym, simp) done -lemma (in M_axioms) obase_equals: +lemma (in M_basic) obase_equals: "[| wellordered(M,A,r); omap(M,A,r,f); obase(M,A,r,B); otype(M,A,r,i); M(A); M(r); M(f); M(B); M(i) |] ==> B = A" apply (rotate_tac 4) @@ -570,21 +570,21 @@ text{*Main result: @{term om} gives the order-isomorphism @{term "\A,r\ \ \i, Memrel(i)\"} *} -theorem (in M_axioms) omap_ord_iso_otype: +theorem (in M_basic) omap_ord_iso_otype: "[| wellordered(M,A,r); omap(M,A,r,f); obase(M,A,r,B); otype(M,A,r,i); M(A); M(r); M(f); M(B); M(i) |] ==> f \ ord_iso(A, r, i, Memrel(i))" apply (frule omap_ord_iso, assumption+) apply (frule obase_equals, assumption+, blast) done -lemma (in M_axioms) obase_exists: +lemma (in M_basic) obase_exists: "[| M(A); M(r) |] ==> \z[M]. obase(M,A,r,z)" apply (simp add: obase_def) apply (insert obase_separation [of A r]) apply (simp add: separation_def) done -lemma (in M_axioms) omap_exists: +lemma (in M_basic) omap_exists: "[| M(A); M(r) |] ==> \z[M]. omap(M,A,r,z)" apply (insert obase_exists [of A r]) apply (simp add: omap_def) @@ -601,7 +601,7 @@ declare rall_simps [simp] rex_simps [simp] -lemma (in M_axioms) otype_exists: +lemma (in M_basic) otype_exists: "[| wellordered(M,A,r); M(A); M(r) |] ==> \i[M]. otype(M,A,r,i)" apply (insert omap_exists [of A r]) apply (simp add: otype_def, safe) @@ -609,7 +609,7 @@ apply blast+ done -theorem (in M_axioms) omap_ord_iso_otype': +theorem (in M_basic) 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) @@ -620,7 +620,7 @@ apply (simp_all add: wellordered_is_trans_on) done -lemma (in M_axioms) ordertype_exists: +lemma (in M_basic) ordertype_exists: "[| 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) @@ -632,7 +632,7 @@ done -lemma (in M_axioms) relativized_imp_well_ord: +lemma (in M_basic) relativized_imp_well_ord: "[| wellordered(M,A,r); M(A); M(r) |] ==> well_ord(A,r)" apply (insert ordertype_exists [of A r], simp) apply (blast intro: well_ord_ord_iso well_ord_Memrel) @@ -641,13 +641,13 @@ subsection {*Kunen's theorem 5.4, poage 127*} text{*(a) The notion of Wellordering is absolute*} -theorem (in M_axioms) well_ord_abs [simp]: +theorem (in M_basic) well_ord_abs [simp]: "[| M(A); M(r) |] ==> wellordered(M,A,r) <-> well_ord(A,r)" by (blast intro: well_ord_imp_relativized relativized_imp_well_ord) text{*(b) Order types are absolute*} -lemma (in M_axioms) +lemma (in M_basic) "[| wellordered(M,A,r); f \ ord_iso(A, r, i, Memrel(i)); M(A); M(r); M(f); M(i); Ord(i) |] ==> i = ordertype(A,r)" by (blast intro: Ord_ordertype relativized_imp_well_ord ordertype_ord_iso