Theory Discipline_Base
theory Discipline_Base
imports
"ZF-Constructible.Rank"
ZF_Miscellanea
"Relativization"
begin
declare [[syntax_ambiguity_warning = false]]
subsection‹Discipline of relativization of basic concepts›
definition
is_singleton :: "[i⇒o,i,i] ⇒ o" where
"is_singleton(A,x,z) ≡ ∃c[A]. empty(A,c) ∧ is_cons(A,x,c,z)"
lemma (in M_trivial) singleton_abs[simp] :
"⟦ M(x) ; M(s) ⟧ ⟹ is_singleton(M,x,s) ⟷ s = {x}"
unfolding is_singleton_def using nonempty by simp
synthesize "singleton" from_definition "is_singleton"
lemma (in M_trivial) singleton_closed [simp]:
"M(x) ⟹ M({x})"
by simp
lemma (in M_trivial) Upair_closed[simp]: "M(a) ⟹ M(b) ⟹ M(Upair(a,b))"
using Upair_eq_cons by simp
lemma (in M_trivial) upair_closed[simp] : "M(x) ⟹ M(y) ⟹ M({x,y})"
by simp
text‹The following named theorems gather instances of transitivity
that arise from closure theorems›
named_theorems trans_closed
definition
is_hcomp :: "[i⇒o,i⇒i⇒o,i⇒i⇒o,i,i] ⇒ o" where
"is_hcomp(M,is_f,is_g,a,w) ≡ ∃z[M]. is_g(a,z) ∧ is_f(z,w)"
lemma (in M_trivial) is_hcomp_abs:
assumes
is_f_abs:"⋀a z. M(a) ⟹ M(z) ⟹ is_f(a,z) ⟷ z = f(a)" and
is_g_abs:"⋀a z. M(a) ⟹ M(z) ⟹ is_g(a,z) ⟷ z = g(a)" and
g_closed:"⋀a. M(a) ⟹ M(g(a))"
"M(a)" "M(w)"
shows
"is_hcomp(M,is_f,is_g,a,w) ⟷ w = f(g(a))"
unfolding is_hcomp_def using assms by simp
definition
hcomp_fm :: "[i⇒i⇒i,i⇒i⇒i,i,i] ⇒ i" where
"hcomp_fm(pf,pg,a,w) ≡ Exists(And(pg(succ(a),0),pf(0,succ(w))))"
lemma sats_hcomp_fm:
assumes
f_iff_sats:"⋀a b z. a∈nat ⟹ b∈nat ⟹ z∈M ⟹
is_f(nth(a,Cons(z,env)),nth(b,Cons(z,env))) ⟷ sats(M,pf(a,b),Cons(z,env))"
and
g_iff_sats:"⋀a b z. a∈nat ⟹ b∈nat ⟹ z∈M ⟹
is_g(nth(a,Cons(z,env)),nth(b,Cons(z,env))) ⟷ sats(M,pg(a,b),Cons(z,env))"
and
"a∈nat" "w∈nat" "env∈list(M)"
shows
"sats(M,hcomp_fm(pf,pg,a,w),env) ⟷ is_hcomp(##M,is_f,is_g,nth(a,env),nth(w,env))"
proof -
have "sats(M, pf(0, succ(w)), Cons(x, env)) ⟷ is_f(x,nth(w,env))" if "x∈M" "w∈nat" for x w
using f_iff_sats[of 0 "succ(w)" x] that by simp
moreover
have "sats(M, pg(succ(a), 0), Cons(x, env)) ⟷ is_g(nth(a,env),x)" if "x∈M" "a∈nat" for x a
using g_iff_sats[of "succ(a)" 0 x] that by simp
ultimately
show ?thesis unfolding hcomp_fm_def is_hcomp_def using assms by simp
qed
definition
hcomp_r :: "[i⇒o,[i⇒o,i,i]⇒o,[i⇒o,i,i]⇒o,i,i] ⇒ o" where
"hcomp_r(M,is_f,is_g,a,w) ≡ ∃z[M]. is_g(M,a,z) ∧ is_f(M,z,w)"
definition
is_hcomp2_2 :: "[i⇒o,[i⇒o,i,i,i]⇒o,[i⇒o,i,i,i]⇒o,[i⇒o,i,i,i]⇒o,i,i,i] ⇒ o" where
"is_hcomp2_2(M,is_f,is_g1,is_g2,a,b,w) ≡ ∃g1ab[M]. ∃g2ab[M].
is_g1(M,a,b,g1ab) ∧ is_g2(M,a,b,g2ab) ∧ is_f(M,g1ab,g2ab,w)"
lemma (in M_trivial) hcomp_abs:
assumes
is_f_abs:"⋀a z. M(a) ⟹ M(z) ⟹ is_f(M,a,z) ⟷ z = f(a)" and
is_g_abs:"⋀a z. M(a) ⟹ M(z) ⟹ is_g(M,a,z) ⟷ z = g(a)" and
g_closed:"⋀a. M(a) ⟹ M(g(a))"
"M(a)" "M(w)"
shows
"hcomp_r(M,is_f,is_g,a,w) ⟷ w = f(g(a))"
unfolding hcomp_r_def using assms by simp
lemma hcomp_uniqueness:
assumes
uniq_is_f:
"⋀r d d'. M(r) ⟹ M(d) ⟹ M(d') ⟹ is_f(M, r, d) ⟹ is_f(M, r, d') ⟹
d = d'"
and
uniq_is_g:
"⋀r d d'. M(r) ⟹ M(d) ⟹ M(d') ⟹ is_g(M, r, d) ⟹ is_g(M, r, d') ⟹
d = d'"
and
"M(a)" "M(w)" "M(w')"
"hcomp_r(M,is_f,is_g,a,w)"
"hcomp_r(M,is_f,is_g,a,w')"
shows
"w=w'"
proof -
from assms
obtain z z' where "is_g(M, a, z)" "is_g(M, a, z')"
"is_f(M,z,w)" "is_f(M,z',w')"
"M(z)" "M(z')"
unfolding hcomp_r_def by blast
moreover from this and uniq_is_g and ‹M(a)›
have "z=z'" by blast
moreover note uniq_is_f and ‹M(w)› ‹M(w')›
ultimately
show ?thesis by blast
qed
lemma hcomp_witness:
assumes
wit_is_f: "⋀r. M(r) ⟹ ∃d[M]. is_f(M,r,d)" and
wit_is_g: "⋀r. M(r) ⟹ ∃d[M]. is_g(M,r,d)" and
"M(a)"
shows
"∃w[M]. hcomp_r(M,is_f,is_g,a,w)"
proof -
from ‹M(a)› and wit_is_g
obtain z where "is_g(M,a,z)" "M(z)" by blast
moreover from this and wit_is_f
obtain w where "is_f(M,z,w)" "M(w)" by blast
ultimately
show ?thesis
using assms unfolding hcomp_r_def by auto
qed
lemma (in M_trivial) hcomp2_2_abs:
assumes
is_f_abs:"⋀r1 r2 z. M(r1) ⟹ M(r2) ⟹ M(z) ⟹ is_f(M,r1,r2,z) ⟷ z = f(r1,r2)" and
is_g1_abs:"⋀r1 r2 z. M(r1) ⟹ M(r2) ⟹ M(z) ⟹ is_g1(M,r1,r2,z) ⟷ z = g1(r1,r2)" and
is_g2_abs:"⋀r1 r2 z. M(r1) ⟹ M(r2) ⟹ M(z) ⟹ is_g2(M,r1,r2,z) ⟷ z = g2(r1,r2)" and
types: "M(a)" "M(b)" "M(w)" "M(g1(a,b))" "M(g2(a,b))"
shows
"is_hcomp2_2(M,is_f,is_g1,is_g2,a,b,w) ⟷ w = f(g1(a,b),g2(a,b))"
unfolding is_hcomp2_2_def using assms
by simp
lemma hcomp2_2_uniqueness:
assumes
uniq_is_f:
"⋀r1 r2 d d'. M(r1) ⟹ M(r2) ⟹ M(d) ⟹ M(d') ⟹
is_f(M, r1, r2 , d) ⟹ is_f(M, r1, r2, d') ⟹ d = d'"
and
uniq_is_g1:
"⋀r1 r2 d d'. M(r1) ⟹ M(r2)⟹ M(d) ⟹ M(d') ⟹ is_g1(M, r1,r2, d) ⟹ is_g1(M, r1,r2, d') ⟹
d = d'"
and
uniq_is_g2:
"⋀r1 r2 d d'. M(r1) ⟹ M(r2)⟹ M(d) ⟹ M(d') ⟹ is_g2(M, r1,r2, d) ⟹ is_g2(M, r1,r2, d') ⟹
d = d'"
and
"M(a)" "M(b)" "M(w)" "M(w')"
"is_hcomp2_2(M,is_f,is_g1,is_g2,a,b,w)"
"is_hcomp2_2(M,is_f,is_g1,is_g2,a,b,w')"
shows
"w=w'"
proof -
from assms
obtain z z' y y' where "is_g1(M, a,b, z)" "is_g1(M, a,b, z')"
"is_g2(M, a,b, y)" "is_g2(M, a,b, y')"
"is_f(M,z,y,w)" "is_f(M,z',y',w')"
"M(z)" "M(z')" "M(y)" "M(y')"
unfolding is_hcomp2_2_def by force
moreover from this and uniq_is_g1 uniq_is_g2 and ‹M(a)› ‹M(b)›
have "z=z'" "y=y'" by blast+
moreover note uniq_is_f and ‹M(w)› ‹M(w')›
ultimately
show ?thesis by blast
qed
lemma hcomp2_2_witness:
assumes
wit_is_f: "⋀r1 r2. M(r1) ⟹ M(r2) ⟹ ∃d[M]. is_f(M,r1,r2,d)" and
wit_is_g1: "⋀r1 r2. M(r1) ⟹ M(r2) ⟹ ∃d[M]. is_g1(M,r1,r2,d)" and
wit_is_g2: "⋀r1 r2. M(r1) ⟹ M(r2) ⟹ ∃d[M]. is_g2(M,r1,r2,d)" and
"M(a)" "M(b)"
shows
"∃w[M]. is_hcomp2_2(M,is_f,is_g1,is_g2,a,b,w)"
proof -
from ‹M(a)› ‹M(b)› and wit_is_g1
obtain g1a where "is_g1(M,a,b,g1a)" "M(g1a)" by blast
moreover from ‹M(a)› ‹M(b)› and wit_is_g2
obtain g2a where "is_g2(M,a,b,g2a)" "M(g2a)" by blast
moreover from calculation and wit_is_f
obtain w where "is_f(M,g1a,g2a,w)" "M(w)" by blast
ultimately
show ?thesis
using assms unfolding is_hcomp2_2_def by auto
qed
lemma (in M_trivial) extensionality_trans:
assumes
"M(d) ∧ (∀x[M]. x∈d ⟷ P(x))"
"M(d') ∧ (∀x[M]. x∈d' ⟷ P(x))"
shows
"d=d'"
proof -
from assms
have "∀x. x∈d ⟷ P(x) ∧ M(x)"
using transM[of _ d] by auto
moreover from assms
have "∀x. x∈d' ⟷ P(x) ∧ M(x)"
using transM[of _ d'] by auto
ultimately
show ?thesis by auto
qed
definition
lt_rel :: "[i⇒o,i,i] ⇒ o" where
"lt_rel(M,a,b) ≡ a∈b ∧ ordinal(M,b)"
lemma (in M_trans) lt_abs[absolut]: "M(a) ⟹ M(b) ⟹ lt_rel(M,a,b) ⟷ a<b"
unfolding lt_rel_def lt_def by auto
definition
le_rel :: "[i⇒o,i,i] ⇒ o" where
"le_rel(M,a,b) ≡ ∃sb[M]. successor(M,b,sb) ∧ lt_rel(M,a,sb)"
lemma (in M_trivial) le_abs[absolut]: "M(a) ⟹ M(b) ⟹ le_rel(M,a,b) ⟷ a≤b"
unfolding le_rel_def by (simp add:absolut)
subsection‹Discipline for \<^term>‹Pow››
definition
is_Pow :: "[i⇒o,i,i] ⇒ o" where
"is_Pow(M,A,z) ≡ M(z) ∧ (∀x[M]. x ∈ z ⟷ subset(M,x,A))"
definition
Pow_rel :: "[i⇒o,i] ⇒ i" (‹Pow⇗_⇖'(_')›) where
"Pow_rel(M,r) ≡ THE d. is_Pow(M,r,d)"
abbreviation
Pow_r_set :: "[i,i] ⇒ i" (‹Pow⇗_⇖'(_')›) where
"Pow_r_set(M) ≡ Pow_rel(##M)"
context M_basic
begin
lemma is_Pow_uniqueness:
assumes
"M(r)"
"is_Pow(M,r,d)" "is_Pow(M,r,d')"
shows
"d=d'"
using assms extensionality_trans
unfolding is_Pow_def
by simp
lemma is_Pow_witness: "M(r) ⟹ ∃d[M]. is_Pow(M,r,d)"
using power_ax unfolding power_ax_def powerset_def is_Pow_def
by simp
lemma is_Pow_closed : "⟦ M(r);is_Pow(M,r,d) ⟧ ⟹ M(d)"
unfolding is_Pow_def by simp
lemma Pow_rel_closed[intro,simp]: "M(r) ⟹ M(Pow_rel(M,r))"
unfolding Pow_rel_def
using is_Pow_closed theI[OF ex1I[of "λd. is_Pow(M,r,d)"], OF _ is_Pow_uniqueness[of r]]
is_Pow_witness
by fastforce
lemmas trans_Pow_rel_closed[trans_closed] = transM[OF _ Pow_rel_closed]
text‹The proof of \<^term>‹f_rel_iff› lemma is schematic and it can reused by copy-paste
replacing appropriately.›
lemma Pow_rel_iff:
assumes "M(r)" "M(d)"
shows "is_Pow(M,r,d) ⟷ d = Pow_rel(M,r)"
proof (intro iffI)
assume "d = Pow_rel(M,r)"
with assms
show "is_Pow(M, r, d)"
using is_Pow_uniqueness[of r] is_Pow_witness
theI[OF ex1I[of "λd. is_Pow(M,r,d)"], OF _ is_Pow_uniqueness[of r]]
unfolding Pow_rel_def
by auto
next
assume "is_Pow(M, r, d)"
with assms
show "d = Pow_rel(M,r)"
using is_Pow_uniqueness unfolding Pow_rel_def
by (auto del:the_equality intro:the_equality[symmetric])
qed
text‹The next "def\_" result really corresponds to @{thm Pow_iff}›
lemma def_Pow_rel: "M(A) ⟹ M(r) ⟹ A∈Pow_rel(M,r) ⟷ A ⊆ r"
using Pow_rel_iff[OF _ Pow_rel_closed, of r r]
unfolding is_Pow_def by simp
lemma Pow_rel_char: "M(r) ⟹ Pow_rel(M,r) = {A∈Pow(r). M(A)}"
proof -
assume "M(r)"
moreover from this
have "x ∈ Pow_rel(M,r) ⟹ x⊆r" "M(x) ⟹ x ⊆ r ⟹ x ∈ Pow_rel(M,r)" for x
using def_Pow_rel by (auto intro!:trans_closed)
ultimately
show ?thesis
using trans_closed by blast
qed
lemma mem_Pow_rel_abs: "M(a) ⟹ M(r) ⟹ a ∈ Pow_rel(M,r) ⟷ a ∈ Pow(r)"
using Pow_rel_char by simp
end
subsection‹Discipline for \<^term>‹PiP››
definition
PiP_rel:: "[i⇒o,i,i]⇒o" where
"PiP_rel(M,A,f) ≡ ∃df[M]. is_domain(M,f,df) ∧ subset(M,A,df) ∧
is_function(M,f)"
context M_basic
begin
lemma def_PiP_rel:
assumes
"M(A)" "M(f)"
shows
"PiP_rel(M,A,f) ⟷ A ⊆ domain(f) ∧ function(f)"
using assms unfolding PiP_rel_def by simp
end
definition
Sigfun :: "[i,i⇒i]⇒i" where
"Sigfun(x,B) ≡ ⋃y∈B(x). {⟨x,y⟩}"
lemma Sigma_Sigfun: "Sigma(A,B) = ⋃ {Sigfun(x,B) . x∈A}"
unfolding Sigma_def Sigfun_def ..
definition
is_Sigfun :: "[i⇒o,i,i⇒i,i]⇒o" where
"is_Sigfun(M,x,B,Sd) ≡ M(Sd) ∧ (∃RB[M]. is_Replace(M,B(x),λy z. z={⟨x,y⟩},RB)
∧ big_union(M,RB,Sd))"
context M_trivial
begin
lemma is_Sigfun_abs:
assumes
"strong_replacement(M,λy z. z={⟨x,y⟩})"
"M(x)" "M(B(x))" "M(Sd)"
shows
"is_Sigfun(M,x,B,Sd) ⟷ Sd = Sigfun(x,B)"
proof -
have "⋃{z . y ∈ B(x), z = {⟨x, y⟩}} = (⋃y∈B(x). {⟨x, y⟩})" by auto
then
show ?thesis
using assms transM[OF _ ‹M(B(x))›] Replace_abs
unfolding is_Sigfun_def Sigfun_def by auto
qed
lemma Sigfun_closed:
assumes
"strong_replacement(M, λy z. y ∈ B(x) ∧ z = {⟨x, y⟩})"
"M(x)" "M(B(x))"
shows
"M(Sigfun(x,B))"
using assms transM[OF _ ‹M(B(x))›] RepFun_closed2
unfolding Sigfun_def by simp
lemmas trans_Sigfun_closed[trans_closed] = transM[OF _ Sigfun_closed]
end
definition
is_Sigma :: "[i⇒o,i,i⇒i,i]⇒o" where
"is_Sigma(M,A,B,S) ≡ M(S) ∧ (∃RSf[M].
is_Replace(M,A,λx z. z=Sigfun(x,B),RSf) ∧ big_union(M,RSf,S))"
locale M_Pi = M_basic +
assumes
Pi_separation: "M(A) ⟹ separation(M, PiP_rel(M,A))"
and
Pi_replacement:
"M(x) ⟹ M(y) ⟹
strong_replacement(M, λya z. ya ∈ y ∧ z = {⟨x, ya⟩})"
"M(y) ⟹
strong_replacement(M, λx z. z = (⋃xa∈y. {⟨x, xa⟩}))"
locale M_Pi_assumptions = M_Pi +
fixes A B
assumes
Pi_assumptions:
"M(A)"
"⋀x. x∈A ⟹ M(B(x))"
"∀x∈A. strong_replacement(M, λy z. y ∈ B(x) ∧ z = {⟨x, y⟩})"
"strong_replacement(M,λx z. z=Sigfun(x,B))"
begin
lemma Sigma_abs[simp]:
assumes
"M(S)"
shows
"is_Sigma(M,A,B,S) ⟷ S = Sigma(A,B)"
proof -
have "⋃{z . x ∈ A, z = Sigfun(x, B)} = (⋃x∈A. Sigfun(x, B))"
by auto
with assms
show ?thesis
using Replace_abs[of A _ "λx z. z=Sigfun(x,B)"]
Sigfun_closed Sigma_Sigfun[of A B] transM[of _ A]
Pi_assumptions is_Sigfun_abs
unfolding is_Sigma_def by simp
qed
lemma Sigma_closed[intro,simp]: "M(Sigma(A,B))"
proof -
have "(⋃x∈A. Sigfun(x, B)) = ⋃{z . x ∈ A, z = Sigfun(x, B)}"
by auto
then
show ?thesis
using Sigma_Sigfun[of A B] transM[of _ A]
Sigfun_closed Pi_assumptions
by simp
qed
lemmas trans_Sigma_closed[trans_closed] = transM[OF _ Sigma_closed]
end
subsection‹Discipline for \<^term>‹Pi››
definition
is_Pi :: "[i⇒o,i,i⇒i,i]⇒o" where
"is_Pi(M,A,B,I) ≡ M(I) ∧ (∃S[M]. ∃PS[M]. is_Sigma(M,A,B,S) ∧
is_Pow(M,S,PS) ∧
is_Collect(M,PS,PiP_rel(M,A),I))"
definition
Pi_rel :: "[i⇒o,i,i⇒i] ⇒ i" (‹Pi⇗_⇖'(_,_')›) where
"Pi_rel(M,A,B) ≡ THE d. is_Pi(M,A,B,d)"
abbreviation
Pi_r_set :: "[i,i,i⇒i] ⇒ i" (‹Pi⇗_⇖'(_,_')›) where
"Pi_r_set(M,A,B) ≡ Pi_rel(##M,A,B)"
context M_Pi_assumptions
begin
lemma is_Pi_uniqueness:
assumes
"is_Pi(M,A,B,d)" "is_Pi(M,A,B,d')"
shows
"d=d'"
using assms Pi_assumptions extensionality_trans
Pow_rel_iff
unfolding is_Pi_def by simp
lemma is_Pi_witness: "∃d[M]. is_Pi(M,A,B,d)"
using Pow_rel_iff Pi_separation Pi_assumptions
unfolding is_Pi_def by simp
lemma is_Pi_closed : "is_Pi(M,A,B,d) ⟹ M(d)"
unfolding is_Pi_def by simp
lemma Pi_rel_closed[intro,simp]: "M(Pi_rel(M,A,B))"
proof -
have "is_Pi(M, A, B, THE xa. is_Pi(M, A, B, xa))"
using Pi_assumptions
theI[OF ex1I[of "is_Pi(M,A,B)"], OF _ is_Pi_uniqueness]
is_Pi_witness is_Pi_closed
by auto
then show ?thesis
using is_Pi_closed
unfolding Pi_rel_def
by simp
qed
lemmas trans_Pi_rel_closed[trans_closed] = transM[OF _ Pi_rel_closed]
lemma Pi_rel_iff:
assumes "M(d)"
shows "is_Pi(M,A,B,d) ⟷ d = Pi_rel(M,A,B)"
proof (intro iffI)
assume "d = Pi_rel(M,A,B)"
moreover
note assms
moreover from this
obtain e where "M(e)" "is_Pi(M,A,B,e)"
using is_Pi_witness by blast
ultimately
show "is_Pi(M, A, B, d)"
using is_Pi_uniqueness is_Pi_witness is_Pi_closed
theI[OF ex1I[of "is_Pi(M,A,B)"], OF _ is_Pi_uniqueness, of e]
unfolding Pi_rel_def
by simp
next
assume "is_Pi(M, A, B, d)"
with assms
show "d = Pi_rel(M,A,B)"
using is_Pi_uniqueness is_Pi_closed unfolding Pi_rel_def
by (blast del:the_equality intro:the_equality[symmetric])
qed
lemma def_Pi_rel:
"Pi_rel(M,A,B) = {f∈Pow_rel(M,Sigma(A,B)). A⊆domain(f) ∧ function(f)}"
proof -
have "Pi_rel(M,A, B) ⊆ Pow_rel(M,Sigma(A,B))"
using Pi_assumptions Pi_rel_iff[of "Pi_rel(M,A,B)"] Pow_rel_iff
unfolding is_Pi_def by auto
moreover
have "f ∈ Pi_rel(M,A, B) ⟹ A⊆domain(f) ∧ function(f)" for f
using Pi_assumptions Pi_rel_iff[of "Pi_rel(M,A,B)"]
def_PiP_rel[of A f] trans_closed Pow_rel_iff
unfolding is_Pi_def by simp
moreover
have "f ∈ Pow_rel(M,Sigma(A,B)) ⟹ A⊆domain(f) ∧ function(f) ⟹ f ∈ Pi_rel(M,A, B)" for f
using Pi_rel_iff[of "Pi_rel(M,A,B)"] Pi_assumptions
def_PiP_rel[of A f] trans_closed Pow_rel_iff
unfolding is_Pi_def by simp
ultimately
show ?thesis by force
qed
lemma Pi_rel_char: "Pi_rel(M,A,B) = {f∈Pi(A,B). M(f)}"
using Pi_assumptions def_Pi_rel Pow_rel_char[OF Sigma_closed] unfolding Pi_def
by fastforce
lemma mem_Pi_rel_abs:
assumes "M(f)"
shows "f ∈ Pi_rel(M,A,B) ⟷ f ∈ Pi(A,B)"
using assms Pi_rel_char by simp
end
text‹The next locale (and similar ones below) are used to
show the relationship between versions of simple (i.e.
$\Sigma_1^{\mathit{ZF}}$, $\Pi_1^{\mathit{ZF}}$) concepts in two
different transitive models.›
locale M_N_Pi_assumptions = M:M_Pi_assumptions + N:M_Pi_assumptions N for N +
assumes
M_imp_N:"M(x) ⟹ N(x)"
begin
lemma Pi_rel_transfer: "Pi⇗M⇖(A,B) ⊆ Pi⇗N⇖(A,B)"
using M.Pi_rel_char N.Pi_rel_char M_imp_N by auto
end
locale M_Pi_assumptions_0 = M_Pi_assumptions _ 0
begin
text‹This is used in the proof of \<^term>‹AC_Pi_rel››
lemma Pi_rel_empty1[simp]: "Pi⇗M⇖(0,B) = {0}"
using Pi_assumptions Pow_rel_char
by (unfold def_Pi_rel function_def) (auto)
end
context M_Pi_assumptions
begin
subsection‹Auxiliary ported results on \<^term>‹Pi_rel›, now unused›
lemma Pi_rel_iff':
assumes types:"M(f)"
shows
"f ∈ Pi_rel(M,A,B) ⟷ function(f) ∧ f ⊆ Sigma(A,B) ∧ A ⊆ domain(f)"
using assms Pow_rel_char
by (simp add:def_Pi_rel, blast)
lemma lam_type_M:
assumes "M(A)" "⋀x. x∈A ⟹ M(B(x))"
"⋀x. x ∈ A ⟹ b(x)∈B(x)" "strong_replacement(M,λx y. y=⟨x, b(x)⟩) "
shows "(λx∈A. b(x)) ∈ Pi_rel(M,A,B)"
proof (auto simp add: lam_def def_Pi_rel function_def)
from assms
have "M({⟨x, b(x)⟩ . x ∈ A})"
using Pi_assumptions transM[OF _ ‹M(A)›]
by (rule_tac RepFun_closed, auto intro!:transM[OF _ ‹⋀x. x∈A ⟹ M(B(x))›])
with assms
show "{⟨x, b(x)⟩ . x ∈ A} ∈ Pow⇗M⇖(Sigma(A, B))"
using Pow_rel_char by auto
qed
end
locale M_Pi_assumptions2 = M_Pi_assumptions +
PiC: M_Pi_assumptions _ _ C for C
begin
lemma Pi_rel_type:
assumes "f ∈ Pi⇗M⇖(A,C)" "⋀x. x ∈ A ⟹ f`x ∈ B(x)"
and types: "M(f)"
shows "f ∈ Pi⇗M⇖(A,B)"
using assms Pi_assumptions
by (simp only: Pi_rel_iff' PiC.Pi_rel_iff')
(blast dest: function_apply_equality)
lemma Pi_rel_weaken_type:
assumes "f ∈ Pi⇗M⇖(A,B)" "⋀x. x ∈ A ⟹ B(x) ⊆ C(x)"
and types: "M(f)"
shows "f ∈ Pi⇗M⇖(A,C)"
using assms Pi_assumptions
by (simp only: Pi_rel_iff' PiC.Pi_rel_iff')
(blast intro: Pi_rel_type dest: apply_type)
end
end