# HG changeset patch # User wenzelm # Date 1027958873 -7200 # Node ID 2232810416fcbc1cd60aa0b458e4a2f69dbdc646 # Parent 99e52e78eb655db6a8c056dbaa5f82fc7ca64154 tuned; diff -r 99e52e78eb65 -r 2232810416fc src/ZF/Constructible/L_axioms.thy --- a/src/ZF/Constructible/L_axioms.thy Mon Jul 29 00:57:16 2002 +0200 +++ b/src/ZF/Constructible/L_axioms.thy Mon Jul 29 18:07:53 2002 +0200 @@ -1,37 +1,38 @@ -header {*The ZF Axioms (Except Separation) in L*} + +header {* The ZF Axioms (Except Separation) in L *} theory L_axioms = Formula + Relative + Reflection + MetaExists: text {* The class L satisfies the premises of locale @{text M_triv_axioms} *} lemma transL: "[| y\x; L(x) |] ==> L(y)" -apply (insert Transset_Lset) -apply (simp add: Transset_def L_def, blast) +apply (insert Transset_Lset) +apply (simp add: Transset_def L_def, blast) done lemma nonempty: "L(0)" -apply (simp add: L_def) -apply (blast intro: zero_in_Lset) +apply (simp add: L_def) +apply (blast intro: zero_in_Lset) done lemma upair_ax: "upair_ax(L)" apply (simp add: upair_ax_def upair_def, clarify) -apply (rule_tac x="{x,y}" in rexI) -apply (simp_all add: doubleton_in_L) +apply (rule_tac x="{x,y}" in rexI) +apply (simp_all add: doubleton_in_L) done lemma Union_ax: "Union_ax(L)" apply (simp add: Union_ax_def big_union_def, clarify) -apply (rule_tac x="Union(x)" in rexI) -apply (simp_all add: Union_in_L, auto) -apply (blast intro: transL) +apply (rule_tac x="Union(x)" in rexI) +apply (simp_all add: Union_in_L, auto) +apply (blast intro: transL) done lemma power_ax: "power_ax(L)" apply (simp add: power_ax_def powerset_def Relative.subset_def, clarify) -apply (rule_tac x="{y \ Pow(x). L(y)}" in rexI) +apply (rule_tac x="{y \ Pow(x). L(y)}" in rexI) apply (simp_all add: LPow_in_L, auto) -apply (blast intro: transL) +apply (blast intro: transL) done subsubsection{*For L to satisfy Replacement *} @@ -40,40 +41,40 @@ there too!*) lemma LReplace_in_Lset: - "[|X \ Lset(i); univalent(L,X,Q); Ord(i)|] + "[|X \ Lset(i); univalent(L,X,Q); Ord(i)|] ==> \j. Ord(j) & Replace(X, %x y. Q(x,y) & L(y)) \ Lset(j)" -apply (rule_tac x="\y \ Replace(X, %x y. Q(x,y) & L(y)). succ(lrank(y))" +apply (rule_tac x="\y \ Replace(X, %x y. Q(x,y) & L(y)). succ(lrank(y))" in exI) apply simp -apply clarify -apply (rule_tac a=x in UN_I) - apply (simp_all add: Replace_iff univalent_def) -apply (blast dest: transL L_I) +apply clarify +apply (rule_tac a=x in UN_I) + apply (simp_all add: Replace_iff univalent_def) +apply (blast dest: transL L_I) done -lemma LReplace_in_L: - "[|L(X); univalent(L,X,Q)|] +lemma LReplace_in_L: + "[|L(X); univalent(L,X,Q)|] ==> \Y. L(Y) & Replace(X, %x y. Q(x,y) & L(y)) \ Y" -apply (drule L_D, clarify) +apply (drule L_D, clarify) apply (drule LReplace_in_Lset, assumption+) apply (blast intro: L_I Lset_in_Lset_succ) done lemma replacement: "replacement(L,P)" apply (simp add: replacement_def, clarify) -apply (frule LReplace_in_L, assumption+, clarify) -apply (rule_tac x=Y in rexI) -apply (simp_all add: Replace_iff univalent_def, blast) +apply (frule LReplace_in_L, assumption+, clarify) +apply (rule_tac x=Y in rexI) +apply (simp_all add: Replace_iff univalent_def, blast) done subsection{*Instantiating the locale @{text M_triv_axioms}*} text{*No instances of Separation yet.*} lemma Lset_mono_le: "mono_le_subset(Lset)" -by (simp add: mono_le_subset_def le_imp_subset Lset_mono) +by (simp add: mono_le_subset_def le_imp_subset Lset_mono) lemma Lset_cont: "cont_Ord(Lset)" -by (simp add: cont_Ord_def Limit_Lset_eq OUnion_def Limit_is_Ord) +by (simp add: cont_Ord_def Limit_Lset_eq OUnion_def Limit_is_Ord) lemmas Pair_in_Lset = Formula.Pair_in_LLimit @@ -90,50 +91,88 @@ apply (rule L_nat) done -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] +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 [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 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 [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 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 [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 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 [intro] = M_triv_axioms.nat_into_M [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 [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 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 [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 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 [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] + 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] + +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 number2_abs [simp] +declare number3_abs [simp] subsection{*Instantiation of the locale @{text reflection}*} @@ -151,7 +190,7 @@ text{*We must use the meta-existential quantifier; otherwise the reflection - terms become enormous!*} + terms become enormous!*} constdefs L_Reflects :: "[i=>o,[i,i]=>o] => prop" ("(3REFLECTS/ [_,/ _])") "REFLECTS[P,Q] == (??Cl. Closed_Unbounded(Cl) & @@ -160,60 +199,60 @@ theorem Triv_reflection: "REFLECTS[P, \a x. P(x)]" -apply (simp add: L_Reflects_def) -apply (rule meta_exI) -apply (rule Closed_Unbounded_Ord) +apply (simp add: L_Reflects_def) +apply (rule meta_exI) +apply (rule Closed_Unbounded_Ord) done theorem Not_reflection: "REFLECTS[P,Q] ==> REFLECTS[\x. ~P(x), \a x. ~Q(a,x)]" -apply (unfold L_Reflects_def) -apply (erule meta_exE) -apply (rule_tac x=Cl in meta_exI, simp) +apply (unfold L_Reflects_def) +apply (erule meta_exE) +apply (rule_tac x=Cl in meta_exI, simp) done theorem And_reflection: - "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] + "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] ==> REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" -apply (unfold L_Reflects_def) -apply (elim meta_exE) -apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) -apply (simp add: Closed_Unbounded_Int, blast) +apply (unfold L_Reflects_def) +apply (elim meta_exE) +apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) +apply (simp add: Closed_Unbounded_Int, blast) done theorem Or_reflection: - "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] + "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] ==> REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" -apply (unfold L_Reflects_def) -apply (elim meta_exE) -apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) -apply (simp add: Closed_Unbounded_Int, blast) +apply (unfold L_Reflects_def) +apply (elim meta_exE) +apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) +apply (simp add: Closed_Unbounded_Int, blast) done theorem Imp_reflection: - "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] + "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] ==> REFLECTS[\x. P(x) --> P'(x), \a x. Q(a,x) --> Q'(a,x)]" -apply (unfold L_Reflects_def) -apply (elim meta_exE) -apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) -apply (simp add: Closed_Unbounded_Int, blast) +apply (unfold L_Reflects_def) +apply (elim meta_exE) +apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) +apply (simp add: Closed_Unbounded_Int, blast) done theorem Iff_reflection: - "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] + "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] ==> REFLECTS[\x. P(x) <-> P'(x), \a x. Q(a,x) <-> Q'(a,x)]" -apply (unfold L_Reflects_def) -apply (elim meta_exE) -apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) -apply (simp add: Closed_Unbounded_Int, blast) +apply (unfold L_Reflects_def) +apply (elim meta_exE) +apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) +apply (simp add: Closed_Unbounded_Int, blast) done theorem Ex_reflection: "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] ==> REFLECTS[\x. \z. L(z) \ P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" -apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def) -apply (elim meta_exE) +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 reflection.intro, OF Lset_mono_le Lset_cont Pair_in_Lset], @@ -222,9 +261,9 @@ theorem All_reflection: "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] - ==> REFLECTS[\x. \z. L(z) --> P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" -apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def) -apply (elim meta_exE) + ==> REFLECTS[\x. \z. L(z) --> P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" +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 reflection.intro, OF Lset_mono_le Lset_cont Pair_in_Lset], @@ -234,35 +273,35 @@ theorem Rex_reflection: "REFLECTS[ \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] ==> REFLECTS[\x. \z[L]. P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" -apply (unfold rex_def) +apply (unfold rex_def) apply (intro And_reflection Ex_reflection, assumption) done theorem Rall_reflection: "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] - ==> REFLECTS[\x. \z[L]. P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" -apply (unfold rall_def) + ==> REFLECTS[\x. \z[L]. P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" +apply (unfold rall_def) apply (intro Imp_reflection All_reflection, assumption) done -lemmas FOL_reflections = +lemmas FOL_reflections = Triv_reflection Not_reflection And_reflection Or_reflection Imp_reflection Iff_reflection Ex_reflection All_reflection Rex_reflection Rall_reflection lemma ReflectsD: - "[|REFLECTS[P,Q]; Ord(i)|] + "[|REFLECTS[P,Q]; Ord(i)|] ==> \j. ix \ Lset(j). P(x) <-> Q(j,x))" -apply (unfold L_Reflects_def Closed_Unbounded_def) -apply (elim meta_exE, clarify) -apply (blast dest!: UnboundedD) +apply (unfold L_Reflects_def Closed_Unbounded_def) +apply (elim meta_exE, clarify) +apply (blast dest!: UnboundedD) done lemma ReflectsE: "[| REFLECTS[P,Q]; Ord(i); !!j. [|ix \ Lset(j). P(x) <-> Q(j,x)|] ==> R |] ==> R" -apply (drule ReflectsD, assumption, blast) +apply (drule ReflectsD, assumption, blast) done lemma Collect_mem_eq: "{x\A. x\B} = A \ B" @@ -301,11 +340,11 @@ lemma empty_type [TC]: "x \ nat ==> empty_fm(x) \ formula" -by (simp add: empty_fm_def) +by (simp add: empty_fm_def) lemma arity_empty_fm [simp]: "x \ nat ==> arity(empty_fm(x)) = succ(x)" -by (simp add: empty_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: empty_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_empty_fm [simp]: "[| x \ nat; env \ list(A)|] @@ -313,16 +352,16 @@ by (simp add: empty_fm_def empty_def) lemma empty_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; + "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> empty(**A, x) <-> sats(A, empty_fm(i), env)" by simp theorem empty_reflection: - "REFLECTS[\x. empty(L,f(x)), + "REFLECTS[\x. empty(L,f(x)), \i x. empty(**Lset(i),f(x))]" apply (simp only: empty_def setclass_simps) -apply (intro FOL_reflections) +apply (intro FOL_reflections) done text{*Not used. But maybe useful?*} @@ -330,38 +369,38 @@ "[| n \ nat; env \ list(A); Transset(A)|] ==> sats(A, empty_fm(n), env) <-> nth(n,env) = 0" apply (simp add: empty_fm_def empty_def Transset_def, auto) -apply (case_tac "n < length(env)") -apply (frule nth_type, assumption+, blast) -apply (simp_all add: not_lt_iff_le nth_eq_0) +apply (case_tac "n < length(env)") +apply (frule nth_type, assumption+, blast) +apply (simp_all add: not_lt_iff_le nth_eq_0) done subsubsection{*Unordered Pairs, Internalized*} constdefs upair_fm :: "[i,i,i]=>i" - "upair_fm(x,y,z) == - And(Member(x,z), + "upair_fm(x,y,z) == + And(Member(x,z), And(Member(y,z), - Forall(Implies(Member(0,succ(z)), + Forall(Implies(Member(0,succ(z)), Or(Equal(0,succ(x)), Equal(0,succ(y)))))))" lemma upair_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> upair_fm(x,y,z) \ formula" -by (simp add: upair_fm_def) +by (simp add: upair_fm_def) lemma arity_upair_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(upair_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: upair_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: upair_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_upair_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, upair_fm(x,y,z), env) <-> + ==> sats(A, upair_fm(x,y,z), env) <-> upair(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: upair_fm_def upair_def) lemma upair_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)|] ==> upair(**A, x, y, z) <-> sats(A, upair_fm(i,j,k), env)" by (simp add: sats_upair_fm) @@ -369,127 +408,127 @@ text{*Useful? At least it refers to "real" unordered pairs*} lemma sats_upair_fm2 [simp]: "[| x \ nat; y \ nat; z < length(env); env \ list(A); Transset(A)|] - ==> sats(A, upair_fm(x,y,z), env) <-> + ==> sats(A, upair_fm(x,y,z), env) <-> nth(z,env) = {nth(x,env), nth(y,env)}" -apply (frule lt_length_in_nat, assumption) -apply (simp add: upair_fm_def Transset_def, auto) -apply (blast intro: nth_type) +apply (frule lt_length_in_nat, assumption) +apply (simp add: upair_fm_def Transset_def, auto) +apply (blast intro: nth_type) done theorem upair_reflection: - "REFLECTS[\x. upair(L,f(x),g(x),h(x)), - \i x. upair(**Lset(i),f(x),g(x),h(x))]" + "REFLECTS[\x. upair(L,f(x),g(x),h(x)), + \i x. upair(**Lset(i),f(x),g(x),h(x))]" apply (simp add: upair_def) -apply (intro FOL_reflections) +apply (intro FOL_reflections) done subsubsection{*Ordered pairs, Internalized*} constdefs pair_fm :: "[i,i,i]=>i" - "pair_fm(x,y,z) == + "pair_fm(x,y,z) == Exists(And(upair_fm(succ(x),succ(x),0), Exists(And(upair_fm(succ(succ(x)),succ(succ(y)),0), upair_fm(1,0,succ(succ(z)))))))" lemma pair_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> pair_fm(x,y,z) \ formula" -by (simp add: pair_fm_def) +by (simp add: pair_fm_def) lemma arity_pair_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(pair_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: pair_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: pair_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_pair_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, pair_fm(x,y,z), env) <-> + ==> sats(A, pair_fm(x,y,z), env) <-> pair(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: pair_fm_def pair_def) lemma pair_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)|] ==> pair(**A, x, y, z) <-> sats(A, pair_fm(i,j,k), env)" by (simp add: sats_pair_fm) theorem pair_reflection: - "REFLECTS[\x. pair(L,f(x),g(x),h(x)), + "REFLECTS[\x. pair(L,f(x),g(x),h(x)), \i x. pair(**Lset(i),f(x),g(x),h(x))]" apply (simp only: pair_def setclass_simps) -apply (intro FOL_reflections upair_reflection) +apply (intro FOL_reflections upair_reflection) done subsubsection{*Binary Unions, Internalized*} constdefs union_fm :: "[i,i,i]=>i" - "union_fm(x,y,z) == + "union_fm(x,y,z) == Forall(Iff(Member(0,succ(z)), Or(Member(0,succ(x)),Member(0,succ(y)))))" lemma union_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> union_fm(x,y,z) \ formula" -by (simp add: union_fm_def) +by (simp add: union_fm_def) lemma arity_union_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(union_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: union_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: union_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_union_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, union_fm(x,y,z), env) <-> + ==> sats(A, union_fm(x,y,z), env) <-> union(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: union_fm_def union_def) lemma union_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)|] ==> union(**A, x, y, z) <-> sats(A, union_fm(i,j,k), env)" by (simp add: sats_union_fm) theorem union_reflection: - "REFLECTS[\x. union(L,f(x),g(x),h(x)), + "REFLECTS[\x. union(L,f(x),g(x),h(x)), \i x. union(**Lset(i),f(x),g(x),h(x))]" apply (simp only: union_def setclass_simps) -apply (intro FOL_reflections) +apply (intro FOL_reflections) done subsubsection{*Set ``Cons,'' Internalized*} constdefs cons_fm :: "[i,i,i]=>i" - "cons_fm(x,y,z) == + "cons_fm(x,y,z) == Exists(And(upair_fm(succ(x),succ(x),0), union_fm(0,succ(y),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 arity_cons_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(cons_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: cons_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: cons_fm_def succ_Un_distrib [symmetric] Un_ac) 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 upair_reflection union_reflection) +apply (intro FOL_reflections upair_reflection union_reflection) done @@ -500,30 +539,30 @@ lemma succ_type [TC]: "[| x \ nat; y \ nat |] ==> succ_fm(x,y) \ formula" -by (simp add: succ_fm_def) +by (simp add: succ_fm_def) lemma arity_succ_fm [simp]: - "[| x \ nat; y \ nat |] + "[| x \ nat; y \ nat |] ==> arity(succ_fm(x,y)) = succ(x) \ succ(y)" by (simp add: succ_fm_def) lemma sats_succ_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, succ_fm(x,y), env) <-> + ==> sats(A, succ_fm(x,y), env) <-> successor(**A, nth(x,env), nth(y,env))" by (simp add: succ_fm_def successor_def) lemma successor_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)|] ==> successor(**A, x, y) <-> sats(A, succ_fm(i,j), env)" by simp theorem successor_reflection: - "REFLECTS[\x. successor(L,f(x),g(x)), + "REFLECTS[\x. successor(L,f(x),g(x)), \i x. successor(**Lset(i),f(x),g(x))]" apply (simp only: successor_def setclass_simps) -apply (intro cons_reflection) +apply (intro cons_reflection) done @@ -535,11 +574,11 @@ lemma number1_type [TC]: "x \ nat ==> number1_fm(x) \ formula" -by (simp add: number1_fm_def) +by (simp add: number1_fm_def) lemma arity_number1_fm [simp]: "x \ nat ==> arity(number1_fm(x)) = succ(x)" -by (simp add: number1_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: number1_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_number1_fm [simp]: "[| x \ nat; env \ list(A)|] @@ -547,13 +586,13 @@ by (simp add: number1_fm_def number1_def) lemma number1_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; + "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> number1(**A, x) <-> sats(A, number1_fm(i), env)" by simp theorem number1_reflection: - "REFLECTS[\x. number1(L,f(x)), + "REFLECTS[\x. number1(L,f(x)), \i x. number1(**Lset(i),f(x))]" apply (simp only: number1_def setclass_simps) apply (intro FOL_reflections empty_reflection successor_reflection) @@ -564,36 +603,36 @@ (* "big_union(M,A,z) == \x[M]. x \ z <-> (\y[M]. y\A & x \ y)" *) constdefs big_union_fm :: "[i,i]=>i" - "big_union_fm(A,z) == + "big_union_fm(A,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(A))), Member(1,0)))))" lemma big_union_type [TC]: "[| x \ nat; y \ nat |] ==> big_union_fm(x,y) \ formula" -by (simp add: big_union_fm_def) +by (simp add: big_union_fm_def) lemma arity_big_union_fm [simp]: - "[| x \ nat; y \ nat |] + "[| x \ nat; y \ nat |] ==> arity(big_union_fm(x,y)) = succ(x) \ succ(y)" by (simp add: big_union_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_big_union_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, big_union_fm(x,y), env) <-> + ==> sats(A, big_union_fm(x,y), env) <-> big_union(**A, nth(x,env), nth(y,env))" by (simp add: big_union_fm_def big_union_def) lemma big_union_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)|] ==> big_union(**A, x, y) <-> sats(A, big_union_fm(i,j), env)" by simp theorem big_union_reflection: - "REFLECTS[\x. big_union(L,f(x),g(x)), + "REFLECTS[\x. big_union(L,f(x),g(x)), \i x. big_union(**Lset(i),f(x),g(x))]" apply (simp only: big_union_def setclass_simps) -apply (intro FOL_reflections) +apply (intro FOL_reflections) done @@ -604,26 +643,26 @@ lemma sats_subset_fm': "[|x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, subset_fm(x,y), env) <-> subset(**A, nth(x,env), nth(y,env))" -by (simp add: subset_fm_def Relative.subset_def) + ==> sats(A, subset_fm(x,y), env) <-> subset(**A, nth(x,env), nth(y,env))" +by (simp add: subset_fm_def Relative.subset_def) theorem subset_reflection: - "REFLECTS[\x. subset(L,f(x),g(x)), - \i x. subset(**Lset(i),f(x),g(x))]" + "REFLECTS[\x. subset(L,f(x),g(x)), + \i x. subset(**Lset(i),f(x),g(x))]" apply (simp only: Relative.subset_def setclass_simps) -apply (intro FOL_reflections) +apply (intro FOL_reflections) done lemma sats_transset_fm': "[|x \ nat; env \ list(A)|] ==> sats(A, transset_fm(x), env) <-> transitive_set(**A, nth(x,env))" -by (simp add: sats_subset_fm' transset_fm_def transitive_set_def) +by (simp add: sats_subset_fm' transset_fm_def transitive_set_def) theorem transitive_set_reflection: "REFLECTS[\x. transitive_set(L,f(x)), \i x. transitive_set(**Lset(i),f(x))]" apply (simp only: transitive_set_def setclass_simps) -apply (intro FOL_reflections subset_reflection) +apply (intro FOL_reflections subset_reflection) done lemma sats_ordinal_fm': @@ -639,14 +678,14 @@ theorem ordinal_reflection: "REFLECTS[\x. ordinal(L,f(x)), \i x. ordinal(**Lset(i),f(x))]" apply (simp only: ordinal_def setclass_simps) -apply (intro FOL_reflections transitive_set_reflection) +apply (intro FOL_reflections transitive_set_reflection) done subsubsection{*Membership Relation, Internalized*} constdefs Memrel_fm :: "[i,i]=>i" - "Memrel_fm(A,r) == + "Memrel_fm(A,r) == Forall(Iff(Member(0,succ(r)), Exists(And(Member(0,succ(succ(A))), Exists(And(Member(0,succ(succ(succ(A)))), @@ -655,36 +694,36 @@ lemma Memrel_type [TC]: "[| x \ nat; y \ nat |] ==> Memrel_fm(x,y) \ formula" -by (simp add: Memrel_fm_def) +by (simp add: Memrel_fm_def) lemma arity_Memrel_fm [simp]: - "[| x \ nat; y \ nat |] + "[| x \ nat; y \ nat |] ==> arity(Memrel_fm(x,y)) = succ(x) \ succ(y)" -by (simp add: Memrel_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: Memrel_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_Memrel_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, Memrel_fm(x,y), env) <-> + ==> sats(A, Memrel_fm(x,y), env) <-> membership(**A, nth(x,env), nth(y,env))" by (simp add: Memrel_fm_def membership_def) lemma Memrel_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)|] ==> membership(**A, x, y) <-> sats(A, Memrel_fm(i,j), env)" by simp theorem membership_reflection: - "REFLECTS[\x. membership(L,f(x),g(x)), + "REFLECTS[\x. membership(L,f(x),g(x)), \i x. membership(**Lset(i),f(x),g(x))]" apply (simp only: membership_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done subsubsection{*Predecessor Set, Internalized*} constdefs pred_set_fm :: "[i,i,i,i]=>i" - "pred_set_fm(A,x,r,B) == + "pred_set_fm(A,x,r,B) == Forall(Iff(Member(0,succ(B)), Exists(And(Member(0,succ(succ(r))), And(Member(1,succ(succ(A))), @@ -692,148 +731,148 @@ lemma pred_set_type [TC]: - "[| A \ nat; x \ nat; r \ nat; B \ nat |] + "[| A \ nat; x \ nat; r \ nat; B \ nat |] ==> pred_set_fm(A,x,r,B) \ formula" -by (simp add: pred_set_fm_def) +by (simp add: pred_set_fm_def) lemma arity_pred_set_fm [simp]: - "[| A \ nat; x \ nat; r \ nat; B \ nat |] + "[| A \ nat; x \ nat; r \ nat; B \ nat |] ==> arity(pred_set_fm(A,x,r,B)) = succ(A) \ succ(x) \ succ(r) \ succ(B)" -by (simp add: pred_set_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: pred_set_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_pred_set_fm [simp]: "[| U \ nat; x \ nat; r \ nat; B \ nat; env \ list(A)|] - ==> sats(A, pred_set_fm(U,x,r,B), env) <-> + ==> sats(A, pred_set_fm(U,x,r,B), env) <-> pred_set(**A, nth(U,env), nth(x,env), nth(r,env), nth(B,env))" by (simp add: pred_set_fm_def pred_set_def) lemma pred_set_iff_sats: - "[| nth(i,env) = U; nth(j,env) = x; nth(k,env) = r; nth(l,env) = B; + "[| nth(i,env) = U; nth(j,env) = x; nth(k,env) = r; nth(l,env) = B; i \ nat; j \ nat; k \ nat; l \ nat; env \ list(A)|] ==> pred_set(**A,U,x,r,B) <-> sats(A, pred_set_fm(i,j,k,l), env)" by (simp add: sats_pred_set_fm) theorem pred_set_reflection: - "REFLECTS[\x. pred_set(L,f(x),g(x),h(x),b(x)), - \i x. pred_set(**Lset(i),f(x),g(x),h(x),b(x))]" + "REFLECTS[\x. pred_set(L,f(x),g(x),h(x),b(x)), + \i x. pred_set(**Lset(i),f(x),g(x),h(x),b(x))]" apply (simp only: pred_set_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done subsubsection{*Domain of a Relation, Internalized*} -(* "is_domain(M,r,z) == - \x[M]. (x \ z <-> (\w[M]. w\r & (\y[M]. pair(M,x,y,w))))" *) +(* "is_domain(M,r,z) == + \x[M]. (x \ z <-> (\w[M]. w\r & (\y[M]. pair(M,x,y,w))))" *) constdefs domain_fm :: "[i,i]=>i" - "domain_fm(r,z) == + "domain_fm(r,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(pair_fm(2,0,1))))))" lemma domain_type [TC]: "[| x \ nat; y \ nat |] ==> domain_fm(x,y) \ formula" -by (simp add: domain_fm_def) +by (simp add: domain_fm_def) lemma arity_domain_fm [simp]: - "[| x \ nat; y \ nat |] + "[| x \ nat; y \ nat |] ==> arity(domain_fm(x,y)) = succ(x) \ succ(y)" -by (simp add: domain_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: domain_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_domain_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, domain_fm(x,y), env) <-> + ==> sats(A, domain_fm(x,y), env) <-> is_domain(**A, nth(x,env), nth(y,env))" by (simp add: domain_fm_def is_domain_def) lemma domain_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_domain(**A, x, y) <-> sats(A, domain_fm(i,j), env)" by simp theorem domain_reflection: - "REFLECTS[\x. is_domain(L,f(x),g(x)), + "REFLECTS[\x. is_domain(L,f(x),g(x)), \i x. is_domain(**Lset(i),f(x),g(x))]" apply (simp only: is_domain_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done subsubsection{*Range of a Relation, Internalized*} -(* "is_range(M,r,z) == - \y[M]. (y \ z <-> (\w[M]. w\r & (\x[M]. pair(M,x,y,w))))" *) +(* "is_range(M,r,z) == + \y[M]. (y \ z <-> (\w[M]. w\r & (\x[M]. pair(M,x,y,w))))" *) constdefs range_fm :: "[i,i]=>i" - "range_fm(r,z) == + "range_fm(r,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(pair_fm(0,2,1))))))" lemma range_type [TC]: "[| x \ nat; y \ nat |] ==> range_fm(x,y) \ formula" -by (simp add: range_fm_def) +by (simp add: range_fm_def) lemma arity_range_fm [simp]: - "[| x \ nat; y \ nat |] + "[| x \ nat; y \ nat |] ==> arity(range_fm(x,y)) = succ(x) \ succ(y)" -by (simp add: range_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: range_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_range_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, range_fm(x,y), env) <-> + ==> sats(A, range_fm(x,y), env) <-> is_range(**A, nth(x,env), nth(y,env))" by (simp add: range_fm_def is_range_def) lemma range_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_range(**A, x, y) <-> sats(A, range_fm(i,j), env)" by simp theorem range_reflection: - "REFLECTS[\x. is_range(L,f(x),g(x)), + "REFLECTS[\x. is_range(L,f(x),g(x)), \i x. is_range(**Lset(i),f(x),g(x))]" apply (simp only: is_range_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done - + subsubsection{*Field of a Relation, Internalized*} -(* "is_field(M,r,z) == - \dr[M]. is_domain(M,r,dr) & +(* "is_field(M,r,z) == + \dr[M]. is_domain(M,r,dr) & (\rr[M]. is_range(M,r,rr) & union(M,dr,rr,z))" *) constdefs field_fm :: "[i,i]=>i" - "field_fm(r,z) == - Exists(And(domain_fm(succ(r),0), - Exists(And(range_fm(succ(succ(r)),0), + "field_fm(r,z) == + Exists(And(domain_fm(succ(r),0), + Exists(And(range_fm(succ(succ(r)),0), union_fm(1,0,succ(succ(z)))))))" lemma field_type [TC]: "[| x \ nat; y \ nat |] ==> field_fm(x,y) \ formula" -by (simp add: field_fm_def) +by (simp add: field_fm_def) lemma arity_field_fm [simp]: - "[| x \ nat; y \ nat |] + "[| x \ nat; y \ nat |] ==> arity(field_fm(x,y)) = succ(x) \ succ(y)" -by (simp add: field_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: field_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_field_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, field_fm(x,y), env) <-> + ==> sats(A, field_fm(x,y), env) <-> is_field(**A, nth(x,env), nth(y,env))" by (simp add: field_fm_def is_field_def) lemma field_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_field(**A, x, y) <-> sats(A, field_fm(i,j), env)" by simp theorem field_reflection: - "REFLECTS[\x. is_field(L,f(x),g(x)), + "REFLECTS[\x. is_field(L,f(x),g(x)), \i x. is_field(**Lset(i),f(x),g(x))]" apply (simp only: is_field_def setclass_simps) apply (intro FOL_reflections domain_reflection range_reflection @@ -843,140 +882,140 @@ subsubsection{*Image under a Relation, Internalized*} -(* "image(M,r,A,z) == +(* "image(M,r,A,z) == \y[M]. (y \ z <-> (\w[M]. w\r & (\x[M]. x\A & pair(M,x,y,w))))" *) constdefs image_fm :: "[i,i,i]=>i" - "image_fm(r,A,z) == + "image_fm(r,A,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(And(Member(0,succ(succ(succ(A)))), - pair_fm(0,2,1)))))))" + pair_fm(0,2,1)))))))" lemma image_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> image_fm(x,y,z) \ formula" -by (simp add: image_fm_def) +by (simp add: image_fm_def) lemma arity_image_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(image_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: image_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: image_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_image_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, image_fm(x,y,z), env) <-> + ==> sats(A, image_fm(x,y,z), env) <-> image(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: image_fm_def Relative.image_def) lemma image_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)|] ==> image(**A, x, y, z) <-> sats(A, image_fm(i,j,k), env)" by (simp add: sats_image_fm) theorem image_reflection: - "REFLECTS[\x. image(L,f(x),g(x),h(x)), + "REFLECTS[\x. image(L,f(x),g(x),h(x)), \i x. image(**Lset(i),f(x),g(x),h(x))]" apply (simp only: Relative.image_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done subsubsection{*Pre-Image under a Relation, Internalized*} -(* "pre_image(M,r,A,z) == - \x[M]. x \ z <-> (\w[M]. w\r & (\y[M]. y\A & pair(M,x,y,w)))" *) +(* "pre_image(M,r,A,z) == + \x[M]. x \ z <-> (\w[M]. w\r & (\y[M]. y\A & pair(M,x,y,w)))" *) constdefs pre_image_fm :: "[i,i,i]=>i" - "pre_image_fm(r,A,z) == + "pre_image_fm(r,A,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(And(Member(0,succ(succ(succ(A)))), - pair_fm(2,0,1)))))))" + pair_fm(2,0,1)))))))" lemma pre_image_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> pre_image_fm(x,y,z) \ formula" -by (simp add: pre_image_fm_def) +by (simp add: pre_image_fm_def) lemma arity_pre_image_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(pre_image_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: pre_image_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: pre_image_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_pre_image_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, pre_image_fm(x,y,z), env) <-> + ==> sats(A, pre_image_fm(x,y,z), env) <-> pre_image(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: pre_image_fm_def Relative.pre_image_def) lemma pre_image_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)|] ==> pre_image(**A, x, y, z) <-> sats(A, pre_image_fm(i,j,k), env)" by (simp add: sats_pre_image_fm) theorem pre_image_reflection: - "REFLECTS[\x. pre_image(L,f(x),g(x),h(x)), + "REFLECTS[\x. pre_image(L,f(x),g(x),h(x)), \i x. pre_image(**Lset(i),f(x),g(x),h(x))]" apply (simp only: Relative.pre_image_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done subsubsection{*Function Application, Internalized*} -(* "fun_apply(M,f,x,y) == - (\xs[M]. \fxs[M]. +(* "fun_apply(M,f,x,y) == + (\xs[M]. \fxs[M]. upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))" *) constdefs fun_apply_fm :: "[i,i,i]=>i" - "fun_apply_fm(f,x,y) == + "fun_apply_fm(f,x,y) == Exists(Exists(And(upair_fm(succ(succ(x)), succ(succ(x)), 1), - And(image_fm(succ(succ(f)), 1, 0), + And(image_fm(succ(succ(f)), 1, 0), big_union_fm(0,succ(succ(y)))))))" lemma fun_apply_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> fun_apply_fm(x,y,z) \ formula" -by (simp add: fun_apply_fm_def) +by (simp add: fun_apply_fm_def) lemma arity_fun_apply_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(fun_apply_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: fun_apply_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: fun_apply_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_fun_apply_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, fun_apply_fm(x,y,z), env) <-> + ==> sats(A, fun_apply_fm(x,y,z), env) <-> fun_apply(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: fun_apply_fm_def fun_apply_def) lemma fun_apply_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)|] ==> fun_apply(**A, x, y, z) <-> sats(A, fun_apply_fm(i,j,k), env)" by simp theorem fun_apply_reflection: - "REFLECTS[\x. fun_apply(L,f(x),g(x),h(x)), - \i x. fun_apply(**Lset(i),f(x),g(x),h(x))]" + "REFLECTS[\x. fun_apply(L,f(x),g(x),h(x)), + \i x. fun_apply(**Lset(i),f(x),g(x),h(x))]" apply (simp only: fun_apply_def setclass_simps) apply (intro FOL_reflections upair_reflection image_reflection - big_union_reflection) + big_union_reflection) done subsubsection{*The Concept of Relation, Internalized*} -(* "is_relation(M,r) == +(* "is_relation(M,r) == (\z[M]. z\r --> (\x[M]. \y[M]. pair(M,x,y,z)))" *) constdefs relation_fm :: "i=>i" - "relation_fm(r) == + "relation_fm(r) == Forall(Implies(Member(0,succ(r)), Exists(Exists(pair_fm(1,0,2)))))" lemma relation_type [TC]: "[| x \ nat |] ==> relation_fm(x) \ formula" -by (simp add: relation_fm_def) +by (simp add: relation_fm_def) lemma arity_relation_fm [simp]: "x \ nat ==> arity(relation_fm(x)) = succ(x)" -by (simp add: relation_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: relation_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_relation_fm [simp]: "[| x \ nat; env \ list(A)|] @@ -984,26 +1023,26 @@ by (simp add: relation_fm_def is_relation_def) lemma relation_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; + "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> is_relation(**A, x) <-> sats(A, relation_fm(i), env)" by simp theorem is_relation_reflection: - "REFLECTS[\x. is_relation(L,f(x)), + "REFLECTS[\x. is_relation(L,f(x)), \i x. is_relation(**Lset(i),f(x))]" apply (simp only: is_relation_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done subsubsection{*The Concept of Function, Internalized*} -(* "is_function(M,r) == - \x[M]. \y[M]. \y'[M]. \p[M]. \p'[M]. +(* "is_function(M,r) == + \x[M]. \y[M]. \y'[M]. \p[M]. \p'[M]. pair(M,x,y,p) --> pair(M,x,y',p') --> p\r --> p'\r --> y=y'" *) constdefs function_fm :: "i=>i" - "function_fm(r) == + "function_fm(r) == Forall(Forall(Forall(Forall(Forall( Implies(pair_fm(4,3,1), Implies(pair_fm(4,2,0), @@ -1012,11 +1051,11 @@ lemma function_type [TC]: "[| x \ nat |] ==> function_fm(x) \ formula" -by (simp add: function_fm_def) +by (simp add: function_fm_def) lemma arity_function_fm [simp]: "x \ nat ==> arity(function_fm(x)) = succ(x)" -by (simp add: function_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: function_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_function_fm [simp]: "[| x \ nat; env \ list(A)|] @@ -1024,27 +1063,27 @@ by (simp add: function_fm_def is_function_def) lemma function_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; + "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> is_function(**A, x) <-> sats(A, function_fm(i), env)" by simp theorem is_function_reflection: - "REFLECTS[\x. is_function(L,f(x)), + "REFLECTS[\x. is_function(L,f(x)), \i x. is_function(**Lset(i),f(x))]" apply (simp only: is_function_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done subsubsection{*Typed Functions, Internalized*} -(* "typed_function(M,A,B,r) == +(* "typed_function(M,A,B,r) == is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) & (\u[M]. u\r --> (\x[M]. \y[M]. pair(M,x,y,u) --> y\B))" *) constdefs typed_function_fm :: "[i,i,i]=>i" - "typed_function_fm(A,B,r) == + "typed_function_fm(A,B,r) == And(function_fm(r), And(relation_fm(r), And(domain_fm(r,A), @@ -1053,64 +1092,64 @@ lemma typed_function_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> typed_function_fm(x,y,z) \ formula" -by (simp add: typed_function_fm_def) +by (simp add: typed_function_fm_def) lemma arity_typed_function_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(typed_function_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: typed_function_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: typed_function_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_typed_function_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, typed_function_fm(x,y,z), env) <-> + ==> sats(A, typed_function_fm(x,y,z), env) <-> typed_function(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: typed_function_fm_def typed_function_def) lemma typed_function_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)|] ==> typed_function(**A, x, y, z) <-> sats(A, typed_function_fm(i,j,k), env)" by simp -lemmas function_reflections = +lemmas function_reflections = empty_reflection number1_reflection - upair_reflection pair_reflection union_reflection - big_union_reflection cons_reflection successor_reflection + upair_reflection pair_reflection union_reflection + big_union_reflection cons_reflection successor_reflection fun_apply_reflection subset_reflection - transitive_set_reflection membership_reflection - pred_set_reflection domain_reflection range_reflection field_reflection + transitive_set_reflection membership_reflection + pred_set_reflection domain_reflection range_reflection field_reflection image_reflection pre_image_reflection - is_relation_reflection is_function_reflection + is_relation_reflection is_function_reflection -lemmas function_iff_sats = - empty_iff_sats number1_iff_sats - upair_iff_sats pair_iff_sats union_iff_sats - cons_iff_sats successor_iff_sats +lemmas function_iff_sats = + empty_iff_sats number1_iff_sats + upair_iff_sats pair_iff_sats union_iff_sats + cons_iff_sats successor_iff_sats fun_apply_iff_sats Memrel_iff_sats - pred_set_iff_sats domain_iff_sats range_iff_sats field_iff_sats - image_iff_sats pre_image_iff_sats - relation_iff_sats function_iff_sats + pred_set_iff_sats domain_iff_sats range_iff_sats field_iff_sats + image_iff_sats pre_image_iff_sats + relation_iff_sats function_iff_sats theorem typed_function_reflection: - "REFLECTS[\x. typed_function(L,f(x),g(x),h(x)), + "REFLECTS[\x. typed_function(L,f(x),g(x),h(x)), \i x. typed_function(**Lset(i),f(x),g(x),h(x))]" apply (simp only: typed_function_def setclass_simps) -apply (intro FOL_reflections function_reflections) +apply (intro FOL_reflections function_reflections) done subsubsection{*Composition of Relations, Internalized*} -(* "composition(M,r,s,t) == - \p[M]. p \ t <-> - (\x[M]. \y[M]. \z[M]. \xy[M]. \yz[M]. - pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) & +(* "composition(M,r,s,t) == + \p[M]. p \ t <-> + (\x[M]. \y[M]. \z[M]. \xy[M]. \yz[M]. + pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) & xy \ s & yz \ r)" *) constdefs composition_fm :: "[i,i,i]=>i" - "composition_fm(r,s,t) == + "composition_fm(r,s,t) == Forall(Iff(Member(0,succ(t)), - Exists(Exists(Exists(Exists(Exists( + Exists(Exists(Exists(Exists(Exists( And(pair_fm(4,2,5), And(pair_fm(4,3,1), And(pair_fm(3,2,0), @@ -1118,41 +1157,41 @@ lemma composition_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> composition_fm(x,y,z) \ formula" -by (simp add: composition_fm_def) +by (simp add: composition_fm_def) lemma arity_composition_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(composition_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: composition_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: composition_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_composition_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, composition_fm(x,y,z), env) <-> + ==> sats(A, composition_fm(x,y,z), env) <-> composition(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: composition_fm_def composition_def) lemma composition_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)|] ==> composition(**A, x, y, z) <-> sats(A, composition_fm(i,j,k), env)" by simp theorem composition_reflection: - "REFLECTS[\x. composition(L,f(x),g(x),h(x)), + "REFLECTS[\x. composition(L,f(x),g(x),h(x)), \i x. composition(**Lset(i),f(x),g(x),h(x))]" apply (simp only: composition_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done subsubsection{*Injections, Internalized*} -(* "injection(M,A,B,f) == - typed_function(M,A,B,f) & - (\x[M]. \x'[M]. \y[M]. \p[M]. \p'[M]. +(* "injection(M,A,B,f) == + typed_function(M,A,B,f) & + (\x[M]. \x'[M]. \y[M]. \p[M]. \p'[M]. pair(M,x,y,p) --> pair(M,x',y,p') --> p\f --> p'\f --> x=x')" *) constdefs injection_fm :: "[i,i,i]=>i" - "injection_fm(A,B,f) == + "injection_fm(A,B,f) == And(typed_function_fm(A,B,f), Forall(Forall(Forall(Forall(Forall( Implies(pair_fm(4,2,1), @@ -1163,41 +1202,41 @@ lemma injection_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> injection_fm(x,y,z) \ formula" -by (simp add: injection_fm_def) +by (simp add: injection_fm_def) lemma arity_injection_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(injection_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: injection_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: injection_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_injection_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, injection_fm(x,y,z), env) <-> + ==> sats(A, injection_fm(x,y,z), env) <-> injection(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: injection_fm_def injection_def) lemma injection_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)|] ==> injection(**A, x, y, z) <-> sats(A, injection_fm(i,j,k), env)" by simp theorem injection_reflection: - "REFLECTS[\x. injection(L,f(x),g(x),h(x)), + "REFLECTS[\x. injection(L,f(x),g(x),h(x)), \i x. injection(**Lset(i),f(x),g(x),h(x))]" apply (simp only: injection_def setclass_simps) -apply (intro FOL_reflections function_reflections typed_function_reflection) +apply (intro FOL_reflections function_reflections typed_function_reflection) done subsubsection{*Surjections, Internalized*} (* surjection :: "[i=>o,i,i,i] => o" - "surjection(M,A,B,f) == + "surjection(M,A,B,f) == typed_function(M,A,B,f) & (\y[M]. y\B --> (\x[M]. x\A & fun_apply(M,f,x,y)))" *) constdefs surjection_fm :: "[i,i,i]=>i" - "surjection_fm(A,B,f) == + "surjection_fm(A,B,f) == And(typed_function_fm(A,B,f), Forall(Implies(Member(0,succ(B)), Exists(And(Member(0,succ(succ(A))), @@ -1205,30 +1244,30 @@ lemma surjection_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> surjection_fm(x,y,z) \ formula" -by (simp add: surjection_fm_def) +by (simp add: surjection_fm_def) lemma arity_surjection_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(surjection_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: surjection_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: surjection_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_surjection_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, surjection_fm(x,y,z), env) <-> + ==> sats(A, surjection_fm(x,y,z), env) <-> surjection(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: surjection_fm_def surjection_def) lemma surjection_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)|] ==> surjection(**A, x, y, z) <-> sats(A, surjection_fm(i,j,k), env)" by simp theorem surjection_reflection: - "REFLECTS[\x. surjection(L,f(x),g(x),h(x)), + "REFLECTS[\x. surjection(L,f(x),g(x),h(x)), \i x. surjection(**Lset(i),f(x),g(x),h(x))]" apply (simp only: surjection_def setclass_simps) -apply (intro FOL_reflections function_reflections typed_function_reflection) +apply (intro FOL_reflections function_reflections typed_function_reflection) done @@ -1242,40 +1281,40 @@ lemma bijection_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> bijection_fm(x,y,z) \ formula" -by (simp add: bijection_fm_def) +by (simp add: bijection_fm_def) lemma arity_bijection_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(bijection_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: bijection_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: bijection_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_bijection_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, bijection_fm(x,y,z), env) <-> + ==> sats(A, bijection_fm(x,y,z), env) <-> bijection(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: bijection_fm_def bijection_def) lemma bijection_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)|] ==> bijection(**A, x, y, z) <-> sats(A, bijection_fm(i,j,k), env)" by simp theorem bijection_reflection: - "REFLECTS[\x. bijection(L,f(x),g(x),h(x)), + "REFLECTS[\x. bijection(L,f(x),g(x),h(x)), \i x. bijection(**Lset(i),f(x),g(x),h(x))]" apply (simp only: bijection_def setclass_simps) -apply (intro And_reflection injection_reflection surjection_reflection) +apply (intro And_reflection injection_reflection surjection_reflection) done subsubsection{*Restriction of a Relation, Internalized*} -(* "restriction(M,r,A,z) == - \x[M]. x \ z <-> (x \ r & (\u[M]. u\A & (\v[M]. pair(M,u,v,x))))" *) +(* "restriction(M,r,A,z) == + \x[M]. x \ z <-> (x \ r & (\u[M]. u\A & (\v[M]. pair(M,u,v,x))))" *) constdefs restriction_fm :: "[i,i,i]=>i" - "restriction_fm(r,A,z) == + "restriction_fm(r,A,z) == Forall(Iff(Member(0,succ(z)), And(Member(0,succ(r)), Exists(And(Member(0,succ(succ(A))), @@ -1283,111 +1322,111 @@ lemma restriction_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> restriction_fm(x,y,z) \ formula" -by (simp add: restriction_fm_def) +by (simp add: restriction_fm_def) lemma arity_restriction_fm [simp]: - "[| x \ nat; y \ nat; z \ nat |] + "[| x \ nat; y \ nat; z \ nat |] ==> arity(restriction_fm(x,y,z)) = succ(x) \ succ(y) \ succ(z)" -by (simp add: restriction_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: restriction_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_restriction_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, restriction_fm(x,y,z), env) <-> + ==> sats(A, restriction_fm(x,y,z), env) <-> restriction(**A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: restriction_fm_def restriction_def) lemma restriction_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)|] ==> restriction(**A, x, y, z) <-> sats(A, restriction_fm(i,j,k), env)" by simp theorem restriction_reflection: - "REFLECTS[\x. restriction(L,f(x),g(x),h(x)), + "REFLECTS[\x. restriction(L,f(x),g(x),h(x)), \i x. restriction(**Lset(i),f(x),g(x),h(x))]" apply (simp only: restriction_def setclass_simps) -apply (intro FOL_reflections pair_reflection) +apply (intro FOL_reflections pair_reflection) done subsubsection{*Order-Isomorphisms, Internalized*} (* order_isomorphism :: "[i=>o,i,i,i,i,i] => o" - "order_isomorphism(M,A,r,B,s,f) == - bijection(M,A,B,f) & + "order_isomorphism(M,A,r,B,s,f) == + bijection(M,A,B,f) & (\x[M]. x\A --> (\y[M]. y\A --> (\p[M]. \fx[M]. \fy[M]. \q[M]. - pair(M,x,y,p) --> fun_apply(M,f,x,fx) --> fun_apply(M,f,y,fy) --> + pair(M,x,y,p) --> fun_apply(M,f,x,fx) --> fun_apply(M,f,y,fy) --> pair(M,fx,fy,q) --> (p\r <-> q\s))))" *) constdefs order_isomorphism_fm :: "[i,i,i,i,i]=>i" - "order_isomorphism_fm(A,r,B,s,f) == - And(bijection_fm(A,B,f), + "order_isomorphism_fm(A,r,B,s,f) == + And(bijection_fm(A,B,f), Forall(Implies(Member(0,succ(A)), Forall(Implies(Member(0,succ(succ(A))), Forall(Forall(Forall(Forall( Implies(pair_fm(5,4,3), Implies(fun_apply_fm(f#+6,5,2), Implies(fun_apply_fm(f#+6,4,1), - Implies(pair_fm(2,1,0), + Implies(pair_fm(2,1,0), Iff(Member(3,r#+6), Member(0,s#+6)))))))))))))))" lemma order_isomorphism_type [TC]: - "[| A \ nat; r \ nat; B \ nat; s \ nat; f \ nat |] + "[| A \ nat; r \ nat; B \ nat; s \ nat; f \ nat |] ==> order_isomorphism_fm(A,r,B,s,f) \ formula" -by (simp add: order_isomorphism_fm_def) +by (simp add: order_isomorphism_fm_def) lemma arity_order_isomorphism_fm [simp]: - "[| A \ nat; r \ nat; B \ nat; s \ nat; f \ nat |] - ==> arity(order_isomorphism_fm(A,r,B,s,f)) = - succ(A) \ succ(r) \ succ(B) \ succ(s) \ succ(f)" -by (simp add: order_isomorphism_fm_def succ_Un_distrib [symmetric] Un_ac) + "[| A \ nat; r \ nat; B \ nat; s \ nat; f \ nat |] + ==> arity(order_isomorphism_fm(A,r,B,s,f)) = + succ(A) \ succ(r) \ succ(B) \ succ(s) \ succ(f)" +by (simp add: order_isomorphism_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_order_isomorphism_fm [simp]: "[| U \ nat; r \ nat; B \ nat; s \ nat; f \ nat; env \ list(A)|] - ==> sats(A, order_isomorphism_fm(U,r,B,s,f), env) <-> - order_isomorphism(**A, nth(U,env), nth(r,env), nth(B,env), + ==> sats(A, order_isomorphism_fm(U,r,B,s,f), env) <-> + order_isomorphism(**A, nth(U,env), nth(r,env), nth(B,env), nth(s,env), nth(f,env))" by (simp add: order_isomorphism_fm_def order_isomorphism_def) lemma order_isomorphism_iff_sats: - "[| nth(i,env) = U; nth(j,env) = r; nth(k,env) = B; nth(j',env) = s; - nth(k',env) = f; + "[| nth(i,env) = U; nth(j,env) = r; nth(k,env) = B; nth(j',env) = s; + nth(k',env) = f; i \ nat; j \ nat; k \ nat; j' \ nat; k' \ nat; env \ list(A)|] - ==> order_isomorphism(**A,U,r,B,s,f) <-> - sats(A, order_isomorphism_fm(i,j,k,j',k'), env)" + ==> order_isomorphism(**A,U,r,B,s,f) <-> + sats(A, order_isomorphism_fm(i,j,k,j',k'), env)" by simp theorem order_isomorphism_reflection: - "REFLECTS[\x. order_isomorphism(L,f(x),g(x),h(x),g'(x),h'(x)), + "REFLECTS[\x. order_isomorphism(L,f(x),g(x),h(x),g'(x),h'(x)), \i x. order_isomorphism(**Lset(i),f(x),g(x),h(x),g'(x),h'(x))]" apply (simp only: order_isomorphism_def setclass_simps) -apply (intro FOL_reflections function_reflections bijection_reflection) +apply (intro FOL_reflections function_reflections bijection_reflection) done subsubsection{*Limit Ordinals, Internalized*} text{*A limit ordinal is a non-empty, successor-closed ordinal*} -(* "limit_ordinal(M,a) == - ordinal(M,a) & ~ empty(M,a) & +(* "limit_ordinal(M,a) == + ordinal(M,a) & ~ empty(M,a) & (\x[M]. x\a --> (\y[M]. y\a & successor(M,x,y)))" *) constdefs limit_ordinal_fm :: "i=>i" - "limit_ordinal_fm(x) == + "limit_ordinal_fm(x) == And(ordinal_fm(x), And(Neg(empty_fm(x)), - Forall(Implies(Member(0,succ(x)), + Forall(Implies(Member(0,succ(x)), Exists(And(Member(0,succ(succ(x))), succ_fm(1,0)))))))" lemma limit_ordinal_type [TC]: "x \ nat ==> limit_ordinal_fm(x) \ formula" -by (simp add: limit_ordinal_fm_def) +by (simp add: limit_ordinal_fm_def) lemma arity_limit_ordinal_fm [simp]: "x \ nat ==> arity(limit_ordinal_fm(x)) = succ(x)" -by (simp add: limit_ordinal_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: limit_ordinal_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_limit_ordinal_fm [simp]: "[| x \ nat; env \ list(A)|] @@ -1395,35 +1434,35 @@ by (simp add: limit_ordinal_fm_def limit_ordinal_def sats_ordinal_fm') lemma limit_ordinal_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; + "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> limit_ordinal(**A, x) <-> sats(A, limit_ordinal_fm(i), env)" by simp theorem limit_ordinal_reflection: - "REFLECTS[\x. limit_ordinal(L,f(x)), + "REFLECTS[\x. limit_ordinal(L,f(x)), \i x. limit_ordinal(**Lset(i),f(x))]" apply (simp only: limit_ordinal_def setclass_simps) -apply (intro FOL_reflections ordinal_reflection - empty_reflection successor_reflection) +apply (intro FOL_reflections ordinal_reflection + empty_reflection successor_reflection) done subsubsection{*Omega: The Set of Natural Numbers*} (* omega(M,a) == limit_ordinal(M,a) & (\x[M]. x\a --> ~ limit_ordinal(M,x)) *) constdefs omega_fm :: "i=>i" - "omega_fm(x) == + "omega_fm(x) == And(limit_ordinal_fm(x), Forall(Implies(Member(0,succ(x)), Neg(limit_ordinal_fm(0)))))" lemma omega_type [TC]: "x \ nat ==> omega_fm(x) \ formula" -by (simp add: omega_fm_def) +by (simp add: omega_fm_def) lemma arity_omega_fm [simp]: "x \ nat ==> arity(omega_fm(x)) = succ(x)" -by (simp add: omega_fm_def succ_Un_distrib [symmetric] Un_ac) +by (simp add: omega_fm_def succ_Un_distrib [symmetric] Un_ac) lemma sats_omega_fm [simp]: "[| x \ nat; env \ list(A)|] @@ -1431,16 +1470,16 @@ by (simp add: omega_fm_def omega_def) lemma omega_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; + "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> omega(**A, x) <-> sats(A, omega_fm(i), env)" by simp theorem omega_reflection: - "REFLECTS[\x. omega(L,f(x)), + "REFLECTS[\x. omega(L,f(x)), \i x. omega(**Lset(i),f(x))]" apply (simp only: omega_def setclass_simps) -apply (intro FOL_reflections limit_ordinal_reflection) +apply (intro FOL_reflections limit_ordinal_reflection) done @@ -1451,10 +1490,10 @@ order_isomorphism_reflection ordinal_reflection limit_ordinal_reflection omega_reflection -lemmas fun_plus_iff_sats = - typed_function_iff_sats composition_iff_sats - injection_iff_sats surjection_iff_sats - bijection_iff_sats restriction_iff_sats +lemmas fun_plus_iff_sats = + typed_function_iff_sats composition_iff_sats + injection_iff_sats surjection_iff_sats + bijection_iff_sats restriction_iff_sats order_isomorphism_iff_sats ordinal_iff_sats limit_ordinal_iff_sats omega_iff_sats diff -r 99e52e78eb65 -r 2232810416fc src/ZF/Constructible/Rec_Separation.thy --- a/src/ZF/Constructible/Rec_Separation.thy Mon Jul 29 00:57:16 2002 +0200 +++ b/src/ZF/Constructible/Rec_Separation.thy Mon Jul 29 18:07:53 2002 +0200 @@ -1,4 +1,5 @@ -header{*Separation for Facts About Recursion*} + +header {*Separation for Facts About Recursion*} theory Rec_Separation = Separation + Datatype_absolute: @@ -198,19 +199,14 @@ subsubsection{*Instantiating the locale @{text M_trancl}*} -theorem M_trancl_axioms_L: "M_trancl_axioms(L)" +theorem M_trancl_L: "PROP M_trancl(L)" + apply (rule M_trancl.intro) + apply (rule M_axioms.axioms [OF M_axioms_L])+ apply (rule M_trancl_axioms.intro) apply (assumption | rule rtrancl_separation wellfounded_trancl_separation)+ done -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 - 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] @@ -438,18 +434,12 @@ subsubsection{*Instantiating the locale @{text M_wfrank}*} -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 - 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) + apply (rule M_trancl.axioms [OF M_trancl_L])+ + apply (rule M_wfrank_axioms.intro) + apply (assumption | rule + wfrank_separation wfrank_strong_replacement Ord_wfrank_separation)+ done lemmas iterates_closed = M_wfrank.iterates_closed [OF M_wfrank_L] @@ -1224,7 +1214,9 @@ subsubsection{*Instantiating the locale @{text M_datatypes}*} -theorem M_datatypes_axioms_L: "M_datatypes_axioms(L)" +theorem M_datatypes_L: "PROP M_datatypes(L)" + apply (rule M_datatypes.intro) + apply (rule M_wfrank.axioms [OF M_wfrank_L])+ apply (rule M_datatypes_axioms.intro) apply (assumption | rule list_replacement1 list_replacement2 @@ -1232,15 +1224,6 @@ nth_replacement)+ done -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 - 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] @@ -1338,19 +1321,11 @@ subsubsection{*Instantiating the locale @{text M_eclose}*} -theorem M_eclose_axioms_L: "M_eclose_axioms(L)" - apply (rule M_eclose_axioms.intro) - apply (assumption | rule eclose_replacement1 eclose_replacement2)+ - done - 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) + apply (rule M_datatypes.axioms [OF M_datatypes_L])+ + apply (rule M_eclose_axioms.intro) + apply (assumption | rule eclose_replacement1 eclose_replacement2)+ done lemmas eclose_closed [intro, simp] = M_eclose.eclose_closed [OF M_eclose_L] diff -r 99e52e78eb65 -r 2232810416fc src/ZF/Constructible/Separation.thy --- a/src/ZF/Constructible/Separation.thy Mon Jul 29 00:57:16 2002 +0200 +++ b/src/ZF/Constructible/Separation.thy Mon Jul 29 18:07:53 2002 +0200 @@ -448,7 +448,9 @@ text{*Separation (and Strong Replacement) for basic set-theoretic constructions such as intersection, Cartesian Product and image.*} -theorem M_axioms_axioms_L: "M_axioms_axioms(L)" +theorem M_axioms_L: "PROP M_axioms(L)" + apply (rule M_axioms.intro) + apply (rule M_triv_axioms_L) apply (rule M_axioms_axioms.intro) apply (assumption | rule Inter_separation cartprod_separation image_separation @@ -458,12 +460,6 @@ 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 lemmas cartprod_iff = M_axioms.cartprod_iff [OF M_axioms_L] and cartprod_closed = M_axioms.cartprod_closed [OF M_axioms_L] @@ -570,35 +566,34 @@ 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] -declare converse_closed [intro,simp] +declare cartprod_closed [intro, simp] +declare sum_closed [intro, simp] +declare converse_closed [intro, simp] declare converse_abs [simp] -declare image_closed [intro,simp] +declare image_closed [intro, simp] declare vimage_abs [simp] -declare vimage_closed [intro,simp] +declare vimage_closed [intro, simp] declare domain_abs [simp] -declare domain_closed [intro,simp] +declare domain_closed [intro, simp] declare range_abs [simp] -declare range_closed [intro,simp] +declare range_closed [intro, simp] declare field_abs [simp] -declare field_closed [intro,simp] +declare field_closed [intro, simp] declare relation_abs [simp] declare function_abs [simp] -declare apply_closed [intro,simp] +declare apply_closed [intro, simp] declare typed_function_abs [simp] declare injection_abs [simp] declare surjection_abs [simp] declare bijection_abs [simp] -declare comp_closed [intro,simp] +declare comp_closed [intro, simp] declare composition_abs [simp] declare restriction_abs [simp] -declare restrict_closed [intro,simp] +declare restrict_closed [intro, simp] declare Inter_abs [simp] -declare Inter_closed [intro,simp] -declare Int_closed [intro,simp] +declare Inter_closed [intro, simp] +declare Int_closed [intro, simp] declare is_funspace_abs [simp] -declare finite_funspace_closed [intro,simp] +declare finite_funspace_closed [intro, simp] end