--- a/src/ZF/Constructible/Rec_Separation.thy Sun Jul 28 21:09:37 2002 +0200
+++ b/src/ZF/Constructible/Rec_Separation.thy Mon Jul 29 00:57:16 2002 +0200
@@ -6,7 +6,7 @@
"M_trancl"}, @{text "M_wfrank"} and @{text "M_datatypes"}*}
lemma eq_succ_imp_lt: "[|i = succ(j); Ord(i)|] ==> j<i"
-by simp
+by simp
subsection{*The Locale @{text "M_trancl"}*}
@@ -15,69 +15,69 @@
text{*First, The Defining Formula*}
(* "rtran_closure_mem(M,A,r,p) ==
- \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M].
+ \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M].
omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
(\<exists>f[M]. typed_function(M,n',A,f) &
- (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
- fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
- (\<forall>j[M]. j\<in>n -->
- (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M].
- fun_apply(M,f,j,fj) & successor(M,j,sj) &
- fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"*)
+ (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
+ fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
+ (\<forall>j[M]. j\<in>n -->
+ (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M].
+ fun_apply(M,f,j,fj) & successor(M,j,sj) &
+ fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"*)
constdefs rtran_closure_mem_fm :: "[i,i,i]=>i"
- "rtran_closure_mem_fm(A,r,p) ==
+ "rtran_closure_mem_fm(A,r,p) ==
Exists(Exists(Exists(
And(omega_fm(2),
And(Member(1,2),
And(succ_fm(1,0),
Exists(And(typed_function_fm(1, A#+4, 0),
- And(Exists(Exists(Exists(
- And(pair_fm(2,1,p#+7),
- And(empty_fm(0),
- And(fun_apply_fm(3,0,2), fun_apply_fm(3,5,1))))))),
- Forall(Implies(Member(0,3),
- Exists(Exists(Exists(Exists(
- And(fun_apply_fm(5,4,3),
- And(succ_fm(4,2),
- And(fun_apply_fm(5,2,1),
- And(pair_fm(3,1,0), Member(0,r#+9))))))))))))))))))))"
+ And(Exists(Exists(Exists(
+ And(pair_fm(2,1,p#+7),
+ And(empty_fm(0),
+ And(fun_apply_fm(3,0,2), fun_apply_fm(3,5,1))))))),
+ Forall(Implies(Member(0,3),
+ Exists(Exists(Exists(Exists(
+ And(fun_apply_fm(5,4,3),
+ And(succ_fm(4,2),
+ And(fun_apply_fm(5,2,1),
+ And(pair_fm(3,1,0), Member(0,r#+9))))))))))))))))))))"
lemma rtran_closure_mem_type [TC]:
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> rtran_closure_mem_fm(x,y,z) \<in> formula"
-by (simp add: rtran_closure_mem_fm_def)
+by (simp add: rtran_closure_mem_fm_def)
lemma arity_rtran_closure_mem_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |]
+ "[| x \<in> nat; y \<in> nat; z \<in> nat |]
==> arity(rtran_closure_mem_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
-by (simp add: rtran_closure_mem_fm_def succ_Un_distrib [symmetric] Un_ac)
+by (simp add: rtran_closure_mem_fm_def succ_Un_distrib [symmetric] Un_ac)
lemma sats_rtran_closure_mem_fm [simp]:
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, rtran_closure_mem_fm(x,y,z), env) <->
+ ==> sats(A, rtran_closure_mem_fm(x,y,z), env) <->
rtran_closure_mem(**A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: rtran_closure_mem_fm_def rtran_closure_mem_def)
lemma rtran_closure_mem_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
==> rtran_closure_mem(**A, x, y, z) <-> sats(A, rtran_closure_mem_fm(i,j,k), env)"
by (simp add: sats_rtran_closure_mem_fm)
theorem rtran_closure_mem_reflection:
- "REFLECTS[\<lambda>x. rtran_closure_mem(L,f(x),g(x),h(x)),
+ "REFLECTS[\<lambda>x. rtran_closure_mem(L,f(x),g(x),h(x)),
\<lambda>i x. rtran_closure_mem(**Lset(i),f(x),g(x),h(x))]"
apply (simp only: rtran_closure_mem_def setclass_simps)
-apply (intro FOL_reflections function_reflections fun_plus_reflections)
+apply (intro FOL_reflections function_reflections fun_plus_reflections)
done
text{*Separation for @{term "rtrancl(r)"}.*}
lemma rtrancl_separation:
"[| L(r); L(A) |] ==> separation (L, rtran_closure_mem(L,A,r))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{r,A,z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{r,A,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF rtran_closure_mem_reflection], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
@@ -89,38 +89,38 @@
subsubsection{*Reflexive/Transitive Closure, Internalized*}
-(* "rtran_closure(M,r,s) ==
+(* "rtran_closure(M,r,s) ==
\<forall>A[M]. is_field(M,r,A) -->
- (\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))" *)
+ (\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))" *)
constdefs rtran_closure_fm :: "[i,i]=>i"
- "rtran_closure_fm(r,s) ==
+ "rtran_closure_fm(r,s) ==
Forall(Implies(field_fm(succ(r),0),
Forall(Iff(Member(0,succ(succ(s))),
rtran_closure_mem_fm(1,succ(succ(r)),0)))))"
lemma rtran_closure_type [TC]:
"[| x \<in> nat; y \<in> nat |] ==> rtran_closure_fm(x,y) \<in> formula"
-by (simp add: rtran_closure_fm_def)
+by (simp add: rtran_closure_fm_def)
lemma arity_rtran_closure_fm [simp]:
- "[| x \<in> nat; y \<in> nat |]
+ "[| x \<in> nat; y \<in> nat |]
==> arity(rtran_closure_fm(x,y)) = succ(x) \<union> succ(y)"
by (simp add: rtran_closure_fm_def succ_Un_distrib [symmetric] Un_ac)
lemma sats_rtran_closure_fm [simp]:
"[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, rtran_closure_fm(x,y), env) <->
+ ==> sats(A, rtran_closure_fm(x,y), env) <->
rtran_closure(**A, nth(x,env), nth(y,env))"
by (simp add: rtran_closure_fm_def rtran_closure_def)
lemma rtran_closure_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
+ "[| nth(i,env) = x; nth(j,env) = y;
i \<in> nat; j \<in> nat; env \<in> list(A)|]
==> rtran_closure(**A, x, y) <-> sats(A, rtran_closure_fm(i,j), env)"
by simp
theorem rtran_closure_reflection:
- "REFLECTS[\<lambda>x. rtran_closure(L,f(x),g(x)),
+ "REFLECTS[\<lambda>x. rtran_closure(L,f(x),g(x)),
\<lambda>i x. rtran_closure(**Lset(i),f(x),g(x))]"
apply (simp only: rtran_closure_def setclass_simps)
apply (intro FOL_reflections function_reflections rtran_closure_mem_reflection)
@@ -132,35 +132,35 @@
(* "tran_closure(M,r,t) ==
\<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" *)
constdefs tran_closure_fm :: "[i,i]=>i"
- "tran_closure_fm(r,s) ==
+ "tran_closure_fm(r,s) ==
Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))"
lemma tran_closure_type [TC]:
"[| x \<in> nat; y \<in> nat |] ==> tran_closure_fm(x,y) \<in> formula"
-by (simp add: tran_closure_fm_def)
+by (simp add: tran_closure_fm_def)
lemma arity_tran_closure_fm [simp]:
- "[| x \<in> nat; y \<in> nat |]
+ "[| x \<in> nat; y \<in> nat |]
==> arity(tran_closure_fm(x,y)) = succ(x) \<union> succ(y)"
by (simp add: tran_closure_fm_def succ_Un_distrib [symmetric] Un_ac)
lemma sats_tran_closure_fm [simp]:
"[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, tran_closure_fm(x,y), env) <->
+ ==> sats(A, tran_closure_fm(x,y), env) <->
tran_closure(**A, nth(x,env), nth(y,env))"
by (simp add: tran_closure_fm_def tran_closure_def)
lemma tran_closure_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
+ "[| nth(i,env) = x; nth(j,env) = y;
i \<in> nat; j \<in> nat; env \<in> list(A)|]
==> tran_closure(**A, x, y) <-> sats(A, tran_closure_fm(i,j), env)"
by simp
theorem tran_closure_reflection:
- "REFLECTS[\<lambda>x. tran_closure(L,f(x),g(x)),
+ "REFLECTS[\<lambda>x. tran_closure(L,f(x),g(x)),
\<lambda>i x. tran_closure(**Lset(i),f(x),g(x))]"
apply (simp only: tran_closure_def setclass_simps)
-apply (intro FOL_reflections function_reflections
+apply (intro FOL_reflections function_reflections
rtran_closure_reflection composition_reflection)
done
@@ -168,60 +168,62 @@
subsection{*Separation for the Proof of @{text "wellfounded_on_trancl"}*}
lemma wellfounded_trancl_reflects:
- "REFLECTS[\<lambda>x. \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].
- w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp,
- \<lambda>i x. \<exists>w \<in> Lset(i). \<exists>wx \<in> Lset(i). \<exists>rp \<in> Lset(i).
+ "REFLECTS[\<lambda>x. \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].
+ w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp,
+ \<lambda>i x. \<exists>w \<in> Lset(i). \<exists>wx \<in> Lset(i). \<exists>rp \<in> Lset(i).
w \<in> Z & pair(**Lset(i),w,x,wx) & tran_closure(**Lset(i),r,rp) &
wx \<in> rp]"
-by (intro FOL_reflections function_reflections fun_plus_reflections
+by (intro FOL_reflections function_reflections fun_plus_reflections
tran_closure_reflection)
lemma wellfounded_trancl_separation:
- "[| L(r); L(Z) |] ==>
- separation (L, \<lambda>x.
- \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].
- w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp)"
-apply (rule separation_CollectI)
-apply (rule_tac A="{r,Z,z}" in subset_LsetE, blast )
+ "[| L(r); L(Z) |] ==>
+ separation (L, \<lambda>x.
+ \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].
+ w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp)"
+apply (rule separation_CollectI)
+apply (rule_tac A="{r,Z,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF wellfounded_trancl_reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule bex_iff_sats conj_iff_sats)+
-apply (rule_tac env = "[w,u,r,Z]" in mem_iff_sats)
+apply (rule_tac env = "[w,u,r,Z]" in mem_iff_sats)
apply (rule sep_rules tran_closure_iff_sats | simp)+
done
subsubsection{*Instantiating the locale @{text M_trancl}*}
-ML
-{*
-val rtrancl_separation = thm "rtrancl_separation";
-val wellfounded_trancl_separation = thm "wellfounded_trancl_separation";
+
+theorem M_trancl_axioms_L: "M_trancl_axioms(L)"
+ apply (rule M_trancl_axioms.intro)
+ apply (assumption | rule
+ rtrancl_separation wellfounded_trancl_separation)+
+ done
-
-val m_trancl = [rtrancl_separation, wellfounded_trancl_separation];
-
-fun trancl_L th =
- kill_flex_triv_prems (m_trancl MRS (axioms_L th));
+theorem M_trancl_L: "PROP M_trancl(L)"
+ apply (rule M_trancl.intro)
+ apply (rule M_triv_axioms_L)
+ apply (rule M_axioms_axioms_L)
+ apply (rule M_trancl_axioms_L)
+ done
-bind_thm ("iterates_abs", trancl_L (thm "M_trancl.iterates_abs"));
-bind_thm ("rtran_closure_rtrancl", trancl_L (thm "M_trancl.rtran_closure_rtrancl"));
-bind_thm ("rtrancl_closed", trancl_L (thm "M_trancl.rtrancl_closed"));
-bind_thm ("rtrancl_abs", trancl_L (thm "M_trancl.rtrancl_abs"));
-bind_thm ("trancl_closed", trancl_L (thm "M_trancl.trancl_closed"));
-bind_thm ("trancl_abs", trancl_L (thm "M_trancl.trancl_abs"));
-bind_thm ("wellfounded_on_trancl", trancl_L (thm "M_trancl.wellfounded_on_trancl"));
-bind_thm ("wellfounded_trancl", trancl_L (thm "M_trancl.wellfounded_trancl"));
-bind_thm ("wfrec_relativize", trancl_L (thm "M_trancl.wfrec_relativize"));
-bind_thm ("trans_wfrec_relativize", trancl_L (thm "M_trancl.trans_wfrec_relativize"));
-bind_thm ("trans_wfrec_abs", trancl_L (thm "M_trancl.trans_wfrec_abs"));
-bind_thm ("trans_eq_pair_wfrec_iff", trancl_L (thm "M_trancl.trans_eq_pair_wfrec_iff"));
-bind_thm ("eq_pair_wfrec_iff", trancl_L (thm "M_trancl.eq_pair_wfrec_iff"));
-*}
+lemmas iterates_abs = M_trancl.iterates_abs [OF M_trancl_L]
+ and rtran_closure_rtrancl = M_trancl.rtran_closure_rtrancl [OF M_trancl_L]
+ and rtrancl_closed = M_trancl.rtrancl_closed [OF M_trancl_L]
+ and rtrancl_abs = M_trancl.rtrancl_abs [OF M_trancl_L]
+ and trancl_closed = M_trancl.trancl_closed [OF M_trancl_L]
+ and trancl_abs = M_trancl.trancl_abs [OF M_trancl_L]
+ and wellfounded_on_trancl = M_trancl.wellfounded_on_trancl [OF M_trancl_L]
+ and wellfounded_trancl = M_trancl.wellfounded_trancl [OF M_trancl_L]
+ and wfrec_relativize = M_trancl.wfrec_relativize [OF M_trancl_L]
+ and trans_wfrec_relativize = M_trancl.trans_wfrec_relativize [OF M_trancl_L]
+ and trans_wfrec_abs = M_trancl.trans_wfrec_abs [OF M_trancl_L]
+ and trans_eq_pair_wfrec_iff = M_trancl.trans_eq_pair_wfrec_iff [OF M_trancl_L]
+ and eq_pair_wfrec_iff = M_trancl.eq_pair_wfrec_iff [OF M_trancl_L]
declare rtrancl_closed [intro,simp]
declare rtrancl_abs [simp]
@@ -232,17 +234,17 @@
subsection{*Well-Founded Recursion!*}
(* M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
- "M_is_recfun(M,MH,r,a,f) ==
- \<forall>z[M]. z \<in> f <->
+ "M_is_recfun(M,MH,r,a,f) ==
+ \<forall>z[M]. z \<in> f <->
5 4 3 2 1 0
- (\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M].
- pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) &
+ (\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M].
+ pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) &
pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) &
xa \<in> r & MH(x, f_r_sx, y))"
*)
constdefs is_recfun_fm :: "[[i,i,i]=>i, i, i, i]=>i"
- "is_recfun_fm(p,r,a,f) ==
+ "is_recfun_fm(p,r,a,f) ==
Forall(Iff(Member(0,succ(f)),
Exists(Exists(Exists(Exists(Exists(Exists(
And(pair_fm(5,4,6),
@@ -254,39 +256,39 @@
lemma is_recfun_type_0:
- "[| !!x y z. [| x \<in> nat; y \<in> nat; z \<in> nat |] ==> p(x,y,z) \<in> formula;
- x \<in> nat; y \<in> nat; z \<in> nat |]
+ "[| !!x y z. [| x \<in> nat; y \<in> nat; z \<in> nat |] ==> p(x,y,z) \<in> formula;
+ x \<in> nat; y \<in> nat; z \<in> nat |]
==> is_recfun_fm(p,x,y,z) \<in> formula"
apply (unfold is_recfun_fm_def)
(*FIXME: FIND OUT why simp loops!*)
apply typecheck
-by simp
+by simp
lemma is_recfun_type [TC]:
- "[| p(5,0,4) \<in> formula;
- x \<in> nat; y \<in> nat; z \<in> nat |]
+ "[| p(5,0,4) \<in> formula;
+ x \<in> nat; y \<in> nat; z \<in> nat |]
==> is_recfun_fm(p,x,y,z) \<in> formula"
-by (simp add: is_recfun_fm_def)
+by (simp add: is_recfun_fm_def)
lemma arity_is_recfun_fm [simp]:
- "[| arity(p(5,0,4)) le 8; x \<in> nat; y \<in> nat; z \<in> nat |]
+ "[| arity(p(5,0,4)) le 8; x \<in> nat; y \<in> nat; z \<in> nat |]
==> arity(is_recfun_fm(p,x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
-apply (frule lt_nat_in_nat, simp)
-apply (simp add: is_recfun_fm_def succ_Un_distrib [symmetric] )
-apply (subst subset_Un_iff2 [of "arity(p(5,0,4))", THEN iffD1])
-apply (rule le_imp_subset)
-apply (erule le_trans, simp)
-apply (simp add: succ_Un_distrib [symmetric] Un_ac)
+apply (frule lt_nat_in_nat, simp)
+apply (simp add: is_recfun_fm_def succ_Un_distrib [symmetric] )
+apply (subst subset_Un_iff2 [of "arity(p(5,0,4))", THEN iffD1])
+apply (rule le_imp_subset)
+apply (erule le_trans, simp)
+apply (simp add: succ_Un_distrib [symmetric] Un_ac)
done
lemma sats_is_recfun_fm:
- assumes MH_iff_sats:
- "!!x y z env.
- [| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> MH(nth(x,env), nth(y,env), nth(z,env)) <-> sats(A, p(x,y,z), env)"
- shows
+ assumes MH_iff_sats:
+ "!!x y z env.
+ [| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
+ ==> MH(nth(x,env), nth(y,env), nth(z,env)) <-> sats(A, p(x,y,z), env)"
+ shows
"[|x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, is_recfun_fm(p,x,y,z), env) <->
+ ==> sats(A, is_recfun_fm(p,x,y,z), env) <->
M_is_recfun(**A, MH, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: is_recfun_fm_def M_is_recfun_def MH_iff_sats [THEN iff_sym])
@@ -294,20 +296,20 @@
"[| (!!x y z env. [| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
==> MH(nth(x,env), nth(y,env), nth(z,env)) <->
sats(A, p(x,y,z), env));
- nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> M_is_recfun(**A, MH, x, y, z) <-> sats(A, is_recfun_fm(p,i,j,k), env)"
+ ==> M_is_recfun(**A, MH, x, y, z) <-> sats(A, is_recfun_fm(p,i,j,k), env)"
by (simp add: sats_is_recfun_fm [of A MH])
theorem is_recfun_reflection:
assumes MH_reflection:
- "!!f g h. REFLECTS[\<lambda>x. MH(L, f(x), g(x), h(x)),
+ "!!f g h. REFLECTS[\<lambda>x. MH(L, f(x), g(x), h(x)),
\<lambda>i x. MH(**Lset(i), f(x), g(x), h(x))]"
- shows "REFLECTS[\<lambda>x. M_is_recfun(L, MH(L), f(x), g(x), h(x)),
+ shows "REFLECTS[\<lambda>x. M_is_recfun(L, MH(L), f(x), g(x), h(x)),
\<lambda>i x. M_is_recfun(**Lset(i), MH(**Lset(i)), f(x), g(x), h(x))]"
apply (simp (no_asm_use) only: M_is_recfun_def setclass_simps)
-apply (intro FOL_reflections function_reflections
- restriction_reflection MH_reflection)
+apply (intro FOL_reflections function_reflections
+ restriction_reflection MH_reflection)
done
text{*Currently, @{text sats}-theorems for higher-order operators don't seem
@@ -315,12 +317,12 @@
of the @{text MH}-term.*}
theorem is_wfrec_reflection:
assumes MH_reflection:
- "!!f g h. REFLECTS[\<lambda>x. MH(L, f(x), g(x), h(x)),
+ "!!f g h. REFLECTS[\<lambda>x. MH(L, f(x), g(x), h(x)),
\<lambda>i x. MH(**Lset(i), f(x), g(x), h(x))]"
- shows "REFLECTS[\<lambda>x. is_wfrec(L, MH(L), f(x), g(x), h(x)),
+ shows "REFLECTS[\<lambda>x. is_wfrec(L, MH(L), f(x), g(x), h(x)),
\<lambda>i x. is_wfrec(**Lset(i), MH(**Lset(i)), f(x), g(x), h(x))]"
apply (simp (no_asm_use) only: is_wfrec_def setclass_simps)
-apply (intro FOL_reflections MH_reflection is_recfun_reflection)
+apply (intro FOL_reflections MH_reflection is_recfun_reflection)
done
subsection{*The Locale @{text "M_wfrank"}*}
@@ -331,23 +333,23 @@
"REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) -->
~ (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)),
\<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(**Lset(i),r,rplus) -->
- ~ (\<exists>f \<in> Lset(i).
- M_is_recfun(**Lset(i), %x f y. is_range(**Lset(i),f,y),
+ ~ (\<exists>f \<in> Lset(i).
+ M_is_recfun(**Lset(i), %x f y. is_range(**Lset(i),f,y),
rplus, x, f))]"
-by (intro FOL_reflections function_reflections is_recfun_reflection tran_closure_reflection)
+by (intro FOL_reflections function_reflections is_recfun_reflection tran_closure_reflection)
lemma wfrank_separation:
"L(r) ==>
separation (L, \<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) -->
~ (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{r,z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{r,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF wfrank_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2, clarify)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule ball_iff_sats imp_iff_sats)+
apply (rule_tac env="[rplus,u,r]" in tran_closure_iff_sats)
apply (rule sep_rules is_recfun_iff_sats | simp)+
@@ -357,14 +359,14 @@
subsubsection{*Replacement for @{term "wfrank"}*}
lemma wfrank_replacement_Reflects:
- "REFLECTS[\<lambda>z. \<exists>x[L]. x \<in> A &
+ "REFLECTS[\<lambda>z. \<exists>x[L]. x \<in> A &
(\<forall>rplus[L]. tran_closure(L,r,rplus) -->
- (\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z) &
+ (\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z) &
M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) &
is_range(L,f,y))),
- \<lambda>i z. \<exists>x \<in> Lset(i). x \<in> A &
+ \<lambda>i z. \<exists>x \<in> Lset(i). x \<in> A &
(\<forall>rplus \<in> Lset(i). tran_closure(**Lset(i),r,rplus) -->
- (\<exists>y \<in> Lset(i). \<exists>f \<in> Lset(i). pair(**Lset(i),x,y,z) &
+ (\<exists>y \<in> Lset(i). \<exists>f \<in> Lset(i). pair(**Lset(i),x,y,z) &
M_is_recfun(**Lset(i), %x f y. is_range(**Lset(i),f,y), rplus, x, f) &
is_range(**Lset(i),f,y)))]"
by (intro FOL_reflections function_reflections fun_plus_reflections
@@ -373,24 +375,24 @@
lemma wfrank_strong_replacement:
"L(r) ==>
- strong_replacement(L, \<lambda>x z.
+ strong_replacement(L, \<lambda>x z.
\<forall>rplus[L]. tran_closure(L,r,rplus) -->
- (\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z) &
+ (\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z) &
M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) &
is_range(L,f,y)))"
-apply (rule strong_replacementI)
+apply (rule strong_replacementI)
apply (rule rallI)
-apply (rename_tac B)
-apply (rule separation_CollectI)
-apply (rule_tac A="{B,r,z}" in subset_LsetE, blast )
+apply (rename_tac B)
+apply (rule separation_CollectI)
+apply (rule_tac A="{B,r,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF wfrank_replacement_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule bex_iff_sats ball_iff_sats conj_iff_sats)+
-apply (rule_tac env = "[x,u,B,r]" in mem_iff_sats)
+apply (rule_tac env = "[x,u,B,r]" in mem_iff_sats)
apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+
done
@@ -398,36 +400,36 @@
subsubsection{*Separation for Proving @{text Ord_wfrank_range}*}
lemma Ord_wfrank_Reflects:
- "REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) -->
- ~ (\<forall>f[L]. \<forall>rangef[L].
+ "REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) -->
+ ~ (\<forall>f[L]. \<forall>rangef[L].
is_range(L,f,rangef) -->
M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) -->
ordinal(L,rangef)),
- \<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(**Lset(i),r,rplus) -->
- ~ (\<forall>f \<in> Lset(i). \<forall>rangef \<in> Lset(i).
+ \<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(**Lset(i),r,rplus) -->
+ ~ (\<forall>f \<in> Lset(i). \<forall>rangef \<in> Lset(i).
is_range(**Lset(i),f,rangef) -->
- M_is_recfun(**Lset(i), \<lambda>x f y. is_range(**Lset(i),f,y),
+ M_is_recfun(**Lset(i), \<lambda>x f y. is_range(**Lset(i),f,y),
rplus, x, f) -->
ordinal(**Lset(i),rangef))]"
-by (intro FOL_reflections function_reflections is_recfun_reflection
+by (intro FOL_reflections function_reflections is_recfun_reflection
tran_closure_reflection ordinal_reflection)
lemma Ord_wfrank_separation:
"L(r) ==>
separation (L, \<lambda>x.
- \<forall>rplus[L]. tran_closure(L,r,rplus) -->
- ~ (\<forall>f[L]. \<forall>rangef[L].
+ \<forall>rplus[L]. tran_closure(L,r,rplus) -->
+ ~ (\<forall>f[L]. \<forall>rangef[L].
is_range(L,f,rangef) -->
M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) -->
- ordinal(L,rangef)))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{r,z}" in subset_LsetE, blast )
+ ordinal(L,rangef)))"
+apply (rule separation_CollectI)
+apply (rule_tac A="{r,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF Ord_wfrank_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2, clarify)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule ball_iff_sats imp_iff_sats)+
apply (rule_tac env="[rplus,u,r]" in tran_closure_iff_sats)
apply (rule sep_rules is_recfun_iff_sats | simp)+
@@ -435,40 +437,40 @@
subsubsection{*Instantiating the locale @{text M_wfrank}*}
-ML
-{*
-val wfrank_separation = thm "wfrank_separation";
-val wfrank_strong_replacement = thm "wfrank_strong_replacement";
-val Ord_wfrank_separation = thm "Ord_wfrank_separation";
+
+theorem M_wfrank_axioms_L: "M_wfrank_axioms(L)"
+ apply (rule M_wfrank_axioms.intro)
+ apply (assumption | rule
+ wfrank_separation wfrank_strong_replacement Ord_wfrank_separation)+
+ done
-val m_wfrank =
- [wfrank_separation, wfrank_strong_replacement, Ord_wfrank_separation];
-
-fun wfrank_L th =
- kill_flex_triv_prems (m_wfrank MRS (trancl_L th));
-
-
+theorem M_wfrank_L: "PROP M_wfrank(L)"
+ apply (rule M_wfrank.intro)
+ apply (rule M_triv_axioms_L)
+ apply (rule M_axioms_axioms_L)
+ apply (rule M_trancl_axioms_L)
+ apply (rule M_wfrank_axioms_L)
+ done
-bind_thm ("iterates_closed", wfrank_L (thm "M_wfrank.iterates_closed"));
-bind_thm ("exists_wfrank", wfrank_L (thm "M_wfrank.exists_wfrank"));
-bind_thm ("M_wellfoundedrank", wfrank_L (thm "M_wfrank.M_wellfoundedrank"));
-bind_thm ("Ord_wfrank_range", wfrank_L (thm "M_wfrank.Ord_wfrank_range"));
-bind_thm ("Ord_range_wellfoundedrank", wfrank_L (thm "M_wfrank.Ord_range_wellfoundedrank"));
-bind_thm ("function_wellfoundedrank", wfrank_L (thm "M_wfrank.function_wellfoundedrank"));
-bind_thm ("domain_wellfoundedrank", wfrank_L (thm "M_wfrank.domain_wellfoundedrank"));
-bind_thm ("wellfoundedrank_type", wfrank_L (thm "M_wfrank.wellfoundedrank_type"));
-bind_thm ("Ord_wellfoundedrank", wfrank_L (thm "M_wfrank.Ord_wellfoundedrank"));
-bind_thm ("wellfoundedrank_eq", wfrank_L (thm "M_wfrank.wellfoundedrank_eq"));
-bind_thm ("wellfoundedrank_lt", wfrank_L (thm "M_wfrank.wellfoundedrank_lt"));
-bind_thm ("wellfounded_imp_subset_rvimage", wfrank_L (thm "M_wfrank.wellfounded_imp_subset_rvimage"));
-bind_thm ("wellfounded_imp_wf", wfrank_L (thm "M_wfrank.wellfounded_imp_wf"));
-bind_thm ("wellfounded_on_imp_wf_on", wfrank_L (thm "M_wfrank.wellfounded_on_imp_wf_on"));
-bind_thm ("wf_abs", wfrank_L (thm "M_wfrank.wf_abs"));
-bind_thm ("wf_on_abs", wfrank_L (thm "M_wfrank.wf_on_abs"));
-bind_thm ("wfrec_replacement_iff", wfrank_L (thm "M_wfrank.wfrec_replacement_iff"));
-bind_thm ("trans_wfrec_closed", wfrank_L (thm "M_wfrank.trans_wfrec_closed"));
-bind_thm ("wfrec_closed", wfrank_L (thm "M_wfrank.wfrec_closed"));
-*}
+lemmas iterates_closed = M_wfrank.iterates_closed [OF M_wfrank_L]
+ and exists_wfrank = M_wfrank.exists_wfrank [OF M_wfrank_L]
+ and M_wellfoundedrank = M_wfrank.M_wellfoundedrank [OF M_wfrank_L]
+ and Ord_wfrank_range = M_wfrank.Ord_wfrank_range [OF M_wfrank_L]
+ and Ord_range_wellfoundedrank = M_wfrank.Ord_range_wellfoundedrank [OF M_wfrank_L]
+ and function_wellfoundedrank = M_wfrank.function_wellfoundedrank [OF M_wfrank_L]
+ and domain_wellfoundedrank = M_wfrank.domain_wellfoundedrank [OF M_wfrank_L]
+ and wellfoundedrank_type = M_wfrank.wellfoundedrank_type [OF M_wfrank_L]
+ and Ord_wellfoundedrank = M_wfrank.Ord_wellfoundedrank [OF M_wfrank_L]
+ and wellfoundedrank_eq = M_wfrank.wellfoundedrank_eq [OF M_wfrank_L]
+ and wellfoundedrank_lt = M_wfrank.wellfoundedrank_lt [OF M_wfrank_L]
+ and wellfounded_imp_subset_rvimage = M_wfrank.wellfounded_imp_subset_rvimage [OF M_wfrank_L]
+ and wellfounded_imp_wf = M_wfrank.wellfounded_imp_wf [OF M_wfrank_L]
+ and wellfounded_on_imp_wf_on = M_wfrank.wellfounded_on_imp_wf_on [OF M_wfrank_L]
+ and wf_abs = M_wfrank.wf_abs [OF M_wfrank_L]
+ and wf_on_abs = M_wfrank.wf_on_abs [OF M_wfrank_L]
+ and wfrec_replacement_iff = M_wfrank.wfrec_replacement_iff [OF M_wfrank_L]
+ and trans_wfrec_closed = M_wfrank.trans_wfrec_closed [OF M_wfrank_L]
+ and wfrec_closed = M_wfrank.wfrec_closed [OF M_wfrank_L]
declare iterates_closed [intro,simp]
declare Ord_wfrank_range [rule_format]
@@ -481,9 +483,9 @@
subsubsection{*Binary Products, Internalized*}
constdefs cartprod_fm :: "[i,i,i]=>i"
-(* "cartprod(M,A,B,z) ==
- \<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))" *)
- "cartprod_fm(A,B,z) ==
+(* "cartprod(M,A,B,z) ==
+ \<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))" *)
+ "cartprod_fm(A,B,z) ==
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(A))),
Exists(And(Member(0,succ(succ(succ(B)))),
@@ -491,74 +493,74 @@
lemma cartprod_type [TC]:
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> cartprod_fm(x,y,z) \<in> formula"
-by (simp add: cartprod_fm_def)
+by (simp add: cartprod_fm_def)
lemma arity_cartprod_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |]
+ "[| x \<in> nat; y \<in> nat; z \<in> nat |]
==> arity(cartprod_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
-by (simp add: cartprod_fm_def succ_Un_distrib [symmetric] Un_ac)
+by (simp add: cartprod_fm_def succ_Un_distrib [symmetric] Un_ac)
lemma sats_cartprod_fm [simp]:
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, cartprod_fm(x,y,z), env) <->
+ ==> sats(A, cartprod_fm(x,y,z), env) <->
cartprod(**A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: cartprod_fm_def cartprod_def)
lemma cartprod_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
==> cartprod(**A, x, y, z) <-> sats(A, cartprod_fm(i,j,k), env)"
by (simp add: sats_cartprod_fm)
theorem cartprod_reflection:
- "REFLECTS[\<lambda>x. cartprod(L,f(x),g(x),h(x)),
+ "REFLECTS[\<lambda>x. cartprod(L,f(x),g(x),h(x)),
\<lambda>i x. cartprod(**Lset(i),f(x),g(x),h(x))]"
apply (simp only: cartprod_def setclass_simps)
-apply (intro FOL_reflections pair_reflection)
+apply (intro FOL_reflections pair_reflection)
done
subsubsection{*Binary Sums, Internalized*}
-(* "is_sum(M,A,B,Z) ==
- \<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M].
+(* "is_sum(M,A,B,Z) ==
+ \<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M].
3 2 1 0
number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) &
cartprod(M,s1,B,B1) & union(M,A0,B1,Z)" *)
constdefs sum_fm :: "[i,i,i]=>i"
- "sum_fm(A,B,Z) ==
+ "sum_fm(A,B,Z) ==
Exists(Exists(Exists(Exists(
- And(number1_fm(2),
+ And(number1_fm(2),
And(cartprod_fm(2,A#+4,3),
And(upair_fm(2,2,1),
And(cartprod_fm(1,B#+4,0), union_fm(3,0,Z#+4)))))))))"
lemma sum_type [TC]:
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> sum_fm(x,y,z) \<in> formula"
-by (simp add: sum_fm_def)
+by (simp add: sum_fm_def)
lemma arity_sum_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |]
+ "[| x \<in> nat; y \<in> nat; z \<in> nat |]
==> arity(sum_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
-by (simp add: sum_fm_def succ_Un_distrib [symmetric] Un_ac)
+by (simp add: sum_fm_def succ_Un_distrib [symmetric] Un_ac)
lemma sats_sum_fm [simp]:
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, sum_fm(x,y,z), env) <->
+ ==> sats(A, sum_fm(x,y,z), env) <->
is_sum(**A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: sum_fm_def is_sum_def)
lemma sum_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
==> is_sum(**A, x, y, z) <-> sats(A, sum_fm(i,j,k), env)"
by simp
theorem sum_reflection:
- "REFLECTS[\<lambda>x. is_sum(L,f(x),g(x),h(x)),
+ "REFLECTS[\<lambda>x. is_sum(L,f(x),g(x),h(x)),
\<lambda>i x. is_sum(**Lset(i),f(x),g(x),h(x))]"
apply (simp only: is_sum_def setclass_simps)
-apply (intro FOL_reflections function_reflections cartprod_reflection)
+apply (intro FOL_reflections function_reflections cartprod_reflection)
done
@@ -570,11 +572,11 @@
lemma quasinat_type [TC]:
"x \<in> nat ==> quasinat_fm(x) \<in> formula"
-by (simp add: quasinat_fm_def)
+by (simp add: quasinat_fm_def)
lemma arity_quasinat_fm [simp]:
"x \<in> nat ==> arity(quasinat_fm(x)) = succ(x)"
-by (simp add: quasinat_fm_def succ_Un_distrib [symmetric] Un_ac)
+by (simp add: quasinat_fm_def succ_Un_distrib [symmetric] Un_ac)
lemma sats_quasinat_fm [simp]:
"[| x \<in> nat; env \<in> list(A)|]
@@ -582,85 +584,85 @@
by (simp add: quasinat_fm_def is_quasinat_def)
lemma quasinat_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
+ "[| nth(i,env) = x; nth(j,env) = y;
i \<in> nat; env \<in> list(A)|]
==> is_quasinat(**A, x) <-> sats(A, quasinat_fm(i), env)"
by simp
theorem quasinat_reflection:
- "REFLECTS[\<lambda>x. is_quasinat(L,f(x)),
+ "REFLECTS[\<lambda>x. is_quasinat(L,f(x)),
\<lambda>i x. is_quasinat(**Lset(i),f(x))]"
apply (simp only: is_quasinat_def setclass_simps)
-apply (intro FOL_reflections function_reflections)
+apply (intro FOL_reflections function_reflections)
done
subsubsection{*The Operator @{term is_nat_case}*}
(* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o"
- "is_nat_case(M, a, is_b, k, z) ==
+ "is_nat_case(M, a, is_b, k, z) ==
(empty(M,k) --> z=a) &
(\<forall>m[M]. successor(M,m,k) --> is_b(m,z)) &
(is_quasinat(M,k) | empty(M,z))" *)
text{*The formula @{term is_b} has free variables 1 and 0.*}
constdefs is_nat_case_fm :: "[i, [i,i]=>i, i, i]=>i"
- "is_nat_case_fm(a,is_b,k,z) ==
+ "is_nat_case_fm(a,is_b,k,z) ==
And(Implies(empty_fm(k), Equal(z,a)),
- And(Forall(Implies(succ_fm(0,succ(k)),
+ And(Forall(Implies(succ_fm(0,succ(k)),
Forall(Implies(Equal(0,succ(succ(z))), is_b(1,0))))),
Or(quasinat_fm(k), empty_fm(z))))"
lemma is_nat_case_type [TC]:
- "[| is_b(1,0) \<in> formula;
- x \<in> nat; y \<in> nat; z \<in> nat |]
+ "[| is_b(1,0) \<in> formula;
+ x \<in> nat; y \<in> nat; z \<in> nat |]
==> is_nat_case_fm(x,is_b,y,z) \<in> formula"
-by (simp add: is_nat_case_fm_def)
+by (simp add: is_nat_case_fm_def)
lemma arity_is_nat_case_fm [simp]:
- "[| is_b(1,0) \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> arity(is_nat_case_fm(x,is_b,y,z)) =
- succ(x) \<union> succ(y) \<union> succ(z) \<union> (arity(is_b(1,0)) #- 2)"
-apply (subgoal_tac "arity(is_b(1,0)) \<in> nat")
+ "[| is_b(1,0) \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat |]
+ ==> arity(is_nat_case_fm(x,is_b,y,z)) =
+ succ(x) \<union> succ(y) \<union> succ(z) \<union> (arity(is_b(1,0)) #- 2)"
+apply (subgoal_tac "arity(is_b(1,0)) \<in> nat")
apply typecheck
(*FIXME: could nat_diff_split work?*)
apply (auto simp add: diff_def raw_diff_succ is_nat_case_fm_def nat_imp_quasinat
succ_Un_distrib [symmetric] Un_ac
- split: split_nat_case)
+ split: split_nat_case)
done
lemma sats_is_nat_case_fm:
- assumes is_b_iff_sats:
- "!!a b. [| a \<in> A; b \<in> A|]
+ assumes is_b_iff_sats:
+ "!!a b. [| a \<in> A; b \<in> A|]
==> is_b(a,b) <-> sats(A, p(1,0), Cons(b, Cons(a,env)))"
- shows
+ shows
"[|x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A)|]
- ==> sats(A, is_nat_case_fm(x,p,y,z), env) <->
+ ==> sats(A, is_nat_case_fm(x,p,y,z), env) <->
is_nat_case(**A, nth(x,env), is_b, nth(y,env), nth(z,env))"
-apply (frule lt_length_in_nat, assumption)
+apply (frule lt_length_in_nat, assumption)
apply (simp add: is_nat_case_fm_def is_nat_case_def is_b_iff_sats [THEN iff_sym])
done
lemma is_nat_case_iff_sats:
- "[| (!!a b. [| a \<in> A; b \<in> A|]
+ "[| (!!a b. [| a \<in> A; b \<in> A|]
==> is_b(a,b) <-> sats(A, p(1,0), Cons(b, Cons(a,env))));
- nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
i \<in> nat; j \<in> nat; k < length(env); env \<in> list(A)|]
- ==> is_nat_case(**A, x, is_b, y, z) <-> sats(A, is_nat_case_fm(i,p,j,k), env)"
+ ==> is_nat_case(**A, x, is_b, y, z) <-> sats(A, is_nat_case_fm(i,p,j,k), env)"
by (simp add: sats_is_nat_case_fm [of A is_b])
text{*The second argument of @{term is_b} gives it direct access to @{term x},
- which is essential for handling free variable references. Without this
+ which is essential for handling free variable references. Without this
argument, we cannot prove reflection for @{term iterates_MH}.*}
theorem is_nat_case_reflection:
assumes is_b_reflection:
- "!!h f g. REFLECTS[\<lambda>x. is_b(L, h(x), f(x), g(x)),
+ "!!h f g. REFLECTS[\<lambda>x. is_b(L, h(x), f(x), g(x)),
\<lambda>i x. is_b(**Lset(i), h(x), f(x), g(x))]"
- shows "REFLECTS[\<lambda>x. is_nat_case(L, f(x), is_b(L,x), g(x), h(x)),
+ shows "REFLECTS[\<lambda>x. is_nat_case(L, f(x), is_b(L,x), g(x), h(x)),
\<lambda>i x. is_nat_case(**Lset(i), f(x), is_b(**Lset(i), x), g(x), h(x))]"
apply (simp (no_asm_use) only: is_nat_case_def setclass_simps)
-apply (intro FOL_reflections function_reflections
- restriction_reflection is_b_reflection quasinat_reflection)
+apply (intro FOL_reflections function_reflections
+ restriction_reflection is_b_reflection quasinat_reflection)
done
@@ -672,117 +674,117 @@
is_nat_case(M, v, \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) & isF(gm,u),
n, z)" *)
constdefs iterates_MH_fm :: "[[i,i]=>i, i, i, i, i]=>i"
- "iterates_MH_fm(isF,v,n,g,z) ==
- is_nat_case_fm(v,
- \<lambda>m u. Exists(And(fun_apply_fm(succ(succ(succ(g))),succ(m),0),
- Forall(Implies(Equal(0,succ(succ(u))), isF(1,0))))),
+ "iterates_MH_fm(isF,v,n,g,z) ==
+ is_nat_case_fm(v,
+ \<lambda>m u. Exists(And(fun_apply_fm(succ(succ(succ(g))),succ(m),0),
+ Forall(Implies(Equal(0,succ(succ(u))), isF(1,0))))),
n, z)"
lemma iterates_MH_type [TC]:
- "[| p(1,0) \<in> formula;
- v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |]
+ "[| p(1,0) \<in> formula;
+ v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |]
==> iterates_MH_fm(p,v,x,y,z) \<in> formula"
-by (simp add: iterates_MH_fm_def)
+by (simp add: iterates_MH_fm_def)
lemma arity_iterates_MH_fm [simp]:
- "[| p(1,0) \<in> formula;
- v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> arity(iterates_MH_fm(p,v,x,y,z)) =
+ "[| p(1,0) \<in> formula;
+ v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |]
+ ==> arity(iterates_MH_fm(p,v,x,y,z)) =
succ(v) \<union> succ(x) \<union> succ(y) \<union> succ(z) \<union> (arity(p(1,0)) #- 4)"
apply (subgoal_tac "arity(p(1,0)) \<in> nat")
apply typecheck
apply (simp add: nat_imp_quasinat iterates_MH_fm_def Un_ac
split: split_nat_case, clarify)
apply (rename_tac i j)
-apply (drule eq_succ_imp_eq_m1, simp)
+apply (drule eq_succ_imp_eq_m1, simp)
apply (drule eq_succ_imp_eq_m1, simp)
apply (simp add: diff_Un_distrib succ_Un_distrib Un_ac diff_diff_left)
done
lemma sats_iterates_MH_fm:
- assumes is_F_iff_sats:
- "!!a b c d. [| a \<in> A; b \<in> A; c \<in> A; d \<in> A|]
+ assumes is_F_iff_sats:
+ "!!a b c d. [| a \<in> A; b \<in> A; c \<in> A; d \<in> A|]
==> is_F(a,b) <->
sats(A, p(1,0), Cons(b, Cons(a, Cons(c, Cons(d,env)))))"
- shows
+ shows
"[|v \<in> nat; x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A)|]
- ==> sats(A, iterates_MH_fm(p,v,x,y,z), env) <->
+ ==> sats(A, iterates_MH_fm(p,v,x,y,z), env) <->
iterates_MH(**A, is_F, nth(v,env), nth(x,env), nth(y,env), nth(z,env))"
-by (simp add: iterates_MH_fm_def iterates_MH_def sats_is_nat_case_fm
+by (simp add: iterates_MH_fm_def iterates_MH_def sats_is_nat_case_fm
is_F_iff_sats [symmetric])
lemma iterates_MH_iff_sats:
- "[| (!!a b c d. [| a \<in> A; b \<in> A; c \<in> A; d \<in> A|]
+ "[| (!!a b c d. [| a \<in> A; b \<in> A; c \<in> A; d \<in> A|]
==> is_F(a,b) <->
sats(A, p(1,0), Cons(b, Cons(a, Cons(c, Cons(d,env))))));
- nth(i',env) = v; nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ nth(i',env) = v; nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
i' \<in> nat; i \<in> nat; j \<in> nat; k < length(env); env \<in> list(A)|]
- ==> iterates_MH(**A, is_F, v, x, y, z) <->
+ ==> iterates_MH(**A, is_F, v, x, y, z) <->
sats(A, iterates_MH_fm(p,i',i,j,k), env)"
-apply (rule iff_sym)
-apply (rule iff_trans)
-apply (rule sats_iterates_MH_fm [of A is_F], blast, simp_all)
+apply (rule iff_sym)
+apply (rule iff_trans)
+apply (rule sats_iterates_MH_fm [of A is_F], blast, simp_all)
done
theorem iterates_MH_reflection:
assumes p_reflection:
- "!!f g h. REFLECTS[\<lambda>x. p(L, f(x), g(x)),
+ "!!f g h. REFLECTS[\<lambda>x. p(L, f(x), g(x)),
\<lambda>i x. p(**Lset(i), f(x), g(x))]"
- shows "REFLECTS[\<lambda>x. iterates_MH(L, p(L), e(x), f(x), g(x), h(x)),
+ shows "REFLECTS[\<lambda>x. iterates_MH(L, p(L), e(x), f(x), g(x), h(x)),
\<lambda>i x. iterates_MH(**Lset(i), p(**Lset(i)), e(x), f(x), g(x), h(x))]"
apply (simp (no_asm_use) only: iterates_MH_def)
txt{*Must be careful: simplifying with @{text setclass_simps} above would
change @{text "\<exists>gm[**Lset(i)]"} into @{text "\<exists>gm \<in> Lset(i)"}, when
it would no longer match rule @{text is_nat_case_reflection}. *}
-apply (rule is_nat_case_reflection)
+apply (rule is_nat_case_reflection)
apply (simp (no_asm_use) only: setclass_simps)
apply (intro FOL_reflections function_reflections is_nat_case_reflection
- restriction_reflection p_reflection)
+ restriction_reflection p_reflection)
done
-subsection{*@{term L} is Closed Under the Operator @{term list}*}
+subsection{*@{term L} is Closed Under the Operator @{term list}*}
subsubsection{*The List Functor, Internalized*}
constdefs list_functor_fm :: "[i,i,i]=>i"
-(* "is_list_functor(M,A,X,Z) ==
- \<exists>n1[M]. \<exists>AX[M].
+(* "is_list_functor(M,A,X,Z) ==
+ \<exists>n1[M]. \<exists>AX[M].
number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)" *)
- "list_functor_fm(A,X,Z) ==
+ "list_functor_fm(A,X,Z) ==
Exists(Exists(
- And(number1_fm(1),
+ And(number1_fm(1),
And(cartprod_fm(A#+2,X#+2,0), sum_fm(1,0,Z#+2)))))"
lemma list_functor_type [TC]:
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> list_functor_fm(x,y,z) \<in> formula"
-by (simp add: list_functor_fm_def)
+by (simp add: list_functor_fm_def)
lemma arity_list_functor_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |]
+ "[| x \<in> nat; y \<in> nat; z \<in> nat |]
==> arity(list_functor_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
-by (simp add: list_functor_fm_def succ_Un_distrib [symmetric] Un_ac)
+by (simp add: list_functor_fm_def succ_Un_distrib [symmetric] Un_ac)
lemma sats_list_functor_fm [simp]:
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, list_functor_fm(x,y,z), env) <->
+ ==> sats(A, list_functor_fm(x,y,z), env) <->
is_list_functor(**A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: list_functor_fm_def is_list_functor_def)
lemma list_functor_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
==> is_list_functor(**A, x, y, z) <-> sats(A, list_functor_fm(i,j,k), env)"
by simp
theorem list_functor_reflection:
- "REFLECTS[\<lambda>x. is_list_functor(L,f(x),g(x),h(x)),
+ "REFLECTS[\<lambda>x. is_list_functor(L,f(x),g(x),h(x)),
\<lambda>i x. is_list_functor(**Lset(i),f(x),g(x),h(x))]"
apply (simp only: is_list_functor_def setclass_simps)
apply (intro FOL_reflections number1_reflection
- cartprod_reflection sum_reflection)
+ cartprod_reflection sum_reflection)
done
@@ -793,29 +795,29 @@
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>
is_wfrec(L, iterates_MH(L, is_list_functor(L,A), 0), memsn, u, y)),
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(**Lset(i), u, y, x) \<and>
- is_wfrec(**Lset(i),
- iterates_MH(**Lset(i),
+ is_wfrec(**Lset(i),
+ iterates_MH(**Lset(i),
is_list_functor(**Lset(i), A), 0), memsn, u, y))]"
-by (intro FOL_reflections function_reflections is_wfrec_reflection
- iterates_MH_reflection list_functor_reflection)
+by (intro FOL_reflections function_reflections is_wfrec_reflection
+ iterates_MH_reflection list_functor_reflection)
-lemma list_replacement1:
+lemma list_replacement1:
"L(A) ==> iterates_replacement(L, is_list_functor(L,A), 0)"
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
-apply (rule strong_replacementI)
+apply (rule strong_replacementI)
apply (rule rallI)
-apply (rename_tac B)
-apply (rule separation_CollectI)
-apply (insert nonempty)
-apply (subgoal_tac "L(Memrel(succ(n)))")
-apply (rule_tac A="{B,A,n,z,0,Memrel(succ(n))}" in subset_LsetE, blast )
+apply (rename_tac B)
+apply (rule separation_CollectI)
+apply (insert nonempty)
+apply (subgoal_tac "L(Memrel(succ(n)))")
+apply (rule_tac A="{B,A,n,z,0,Memrel(succ(n))}" in subset_LsetE, blast )
apply (rule ReflectsE [OF list_replacement1_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2 Memrel_closed)
-apply (elim conjE)
+apply (elim conjE)
apply (rule DPow_LsetI)
-apply (rename_tac v)
+apply (rename_tac v)
apply (rule bex_iff_sats conj_iff_sats)+
apply (rule_tac env = "[u,v,A,n,B,0,Memrel(succ(n))]" in mem_iff_sats)
apply (rule sep_rules | simp)+
@@ -832,34 +834,34 @@
is_wfrec (L, iterates_MH (L, is_list_functor(L, A), 0),
msn, u, x)),
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and>
- (\<exists>sn \<in> Lset(i). \<exists>msn \<in> Lset(i).
+ (\<exists>sn \<in> Lset(i). \<exists>msn \<in> Lset(i).
successor(**Lset(i), u, sn) \<and> membership(**Lset(i), sn, msn) \<and>
- is_wfrec (**Lset(i),
+ is_wfrec (**Lset(i),
iterates_MH (**Lset(i), is_list_functor(**Lset(i), A), 0),
msn, u, x))]"
-by (intro FOL_reflections function_reflections is_wfrec_reflection
- iterates_MH_reflection list_functor_reflection)
+by (intro FOL_reflections function_reflections is_wfrec_reflection
+ iterates_MH_reflection list_functor_reflection)
-lemma list_replacement2:
- "L(A) ==> strong_replacement(L,
- \<lambda>n y. n\<in>nat &
+lemma list_replacement2:
+ "L(A) ==> strong_replacement(L,
+ \<lambda>n y. n\<in>nat &
(\<exists>sn[L]. \<exists>msn[L]. successor(L,n,sn) & membership(L,sn,msn) &
- is_wfrec(L, iterates_MH(L,is_list_functor(L,A), 0),
+ is_wfrec(L, iterates_MH(L,is_list_functor(L,A), 0),
msn, n, y)))"
-apply (rule strong_replacementI)
+apply (rule strong_replacementI)
apply (rule rallI)
-apply (rename_tac B)
-apply (rule separation_CollectI)
-apply (insert nonempty)
-apply (rule_tac A="{A,B,z,0,nat}" in subset_LsetE)
-apply (blast intro: L_nat)
+apply (rename_tac B)
+apply (rule separation_CollectI)
+apply (insert nonempty)
+apply (rule_tac A="{A,B,z,0,nat}" in subset_LsetE)
+apply (blast intro: L_nat)
apply (rule ReflectsE [OF list_replacement2_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac v)
+apply (rename_tac v)
apply (rule bex_iff_sats conj_iff_sats)+
apply (rule_tac env = "[u,v,A,B,0,nat]" in mem_iff_sats)
apply (rule sep_rules | simp)+
@@ -868,21 +870,21 @@
done
-subsection{*@{term L} is Closed Under the Operator @{term formula}*}
+subsection{*@{term L} is Closed Under the Operator @{term formula}*}
subsubsection{*The Formula Functor, Internalized*}
constdefs formula_functor_fm :: "[i,i]=>i"
-(* "is_formula_functor(M,X,Z) ==
- \<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M].
+(* "is_formula_functor(M,X,Z) ==
+ \<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M].
4 3 2 1 0
- omega(M,nat') & cartprod(M,nat',nat',natnat) &
+ omega(M,nat') & cartprod(M,nat',nat',natnat) &
is_sum(M,natnat,natnat,natnatsum) &
- cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) &
- is_sum(M,natnatsum,X3,Z)" *)
- "formula_functor_fm(X,Z) ==
+ cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) &
+ is_sum(M,natnatsum,X3,Z)" *)
+ "formula_functor_fm(X,Z) ==
Exists(Exists(Exists(Exists(Exists(
- And(omega_fm(4),
+ And(omega_fm(4),
And(cartprod_fm(4,4,3),
And(sum_fm(3,3,2),
And(cartprod_fm(X#+5,X#+5,1),
@@ -890,26 +892,26 @@
lemma formula_functor_type [TC]:
"[| x \<in> nat; y \<in> nat |] ==> formula_functor_fm(x,y) \<in> formula"
-by (simp add: formula_functor_fm_def)
+by (simp add: formula_functor_fm_def)
lemma sats_formula_functor_fm [simp]:
"[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, formula_functor_fm(x,y), env) <->
+ ==> sats(A, formula_functor_fm(x,y), env) <->
is_formula_functor(**A, nth(x,env), nth(y,env))"
by (simp add: formula_functor_fm_def is_formula_functor_def)
lemma formula_functor_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
+ "[| nth(i,env) = x; nth(j,env) = y;
i \<in> nat; j \<in> nat; env \<in> list(A)|]
==> is_formula_functor(**A, x, y) <-> sats(A, formula_functor_fm(i,j), env)"
by simp
theorem formula_functor_reflection:
- "REFLECTS[\<lambda>x. is_formula_functor(L,f(x),g(x)),
+ "REFLECTS[\<lambda>x. is_formula_functor(L,f(x),g(x)),
\<lambda>i x. is_formula_functor(**Lset(i),f(x),g(x))]"
apply (simp only: is_formula_functor_def setclass_simps)
apply (intro FOL_reflections omega_reflection
- cartprod_reflection sum_reflection)
+ cartprod_reflection sum_reflection)
done
subsubsection{*Instances of Replacement for Formulas*}
@@ -919,28 +921,28 @@
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>
is_wfrec(L, iterates_MH(L, is_formula_functor(L), 0), memsn, u, y)),
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(**Lset(i), u, y, x) \<and>
- is_wfrec(**Lset(i),
- iterates_MH(**Lset(i),
+ is_wfrec(**Lset(i),
+ iterates_MH(**Lset(i),
is_formula_functor(**Lset(i)), 0), memsn, u, y))]"
-by (intro FOL_reflections function_reflections is_wfrec_reflection
- iterates_MH_reflection formula_functor_reflection)
+by (intro FOL_reflections function_reflections is_wfrec_reflection
+ iterates_MH_reflection formula_functor_reflection)
-lemma formula_replacement1:
+lemma formula_replacement1:
"iterates_replacement(L, is_formula_functor(L), 0)"
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
-apply (rule strong_replacementI)
+apply (rule strong_replacementI)
apply (rule rallI)
-apply (rename_tac B)
-apply (rule separation_CollectI)
-apply (insert nonempty)
-apply (subgoal_tac "L(Memrel(succ(n)))")
-apply (rule_tac A="{B,n,z,0,Memrel(succ(n))}" in subset_LsetE, blast )
+apply (rename_tac B)
+apply (rule separation_CollectI)
+apply (insert nonempty)
+apply (subgoal_tac "L(Memrel(succ(n)))")
+apply (rule_tac A="{B,n,z,0,Memrel(succ(n))}" in subset_LsetE, blast )
apply (rule ReflectsE [OF formula_replacement1_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2 Memrel_closed)
apply (rule DPow_LsetI)
-apply (rename_tac v)
+apply (rename_tac v)
apply (rule bex_iff_sats conj_iff_sats)+
apply (rule_tac env = "[u,v,n,B,0,Memrel(succ(n))]" in mem_iff_sats)
apply (rule sep_rules | simp)+
@@ -957,34 +959,34 @@
is_wfrec (L, iterates_MH (L, is_formula_functor(L), 0),
msn, u, x)),
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and>
- (\<exists>sn \<in> Lset(i). \<exists>msn \<in> Lset(i).
+ (\<exists>sn \<in> Lset(i). \<exists>msn \<in> Lset(i).
successor(**Lset(i), u, sn) \<and> membership(**Lset(i), sn, msn) \<and>
- is_wfrec (**Lset(i),
+ is_wfrec (**Lset(i),
iterates_MH (**Lset(i), is_formula_functor(**Lset(i)), 0),
msn, u, x))]"
-by (intro FOL_reflections function_reflections is_wfrec_reflection
- iterates_MH_reflection formula_functor_reflection)
+by (intro FOL_reflections function_reflections is_wfrec_reflection
+ iterates_MH_reflection formula_functor_reflection)
-lemma formula_replacement2:
- "strong_replacement(L,
- \<lambda>n y. n\<in>nat &
+lemma formula_replacement2:
+ "strong_replacement(L,
+ \<lambda>n y. n\<in>nat &
(\<exists>sn[L]. \<exists>msn[L]. successor(L,n,sn) & membership(L,sn,msn) &
- is_wfrec(L, iterates_MH(L,is_formula_functor(L), 0),
+ is_wfrec(L, iterates_MH(L,is_formula_functor(L), 0),
msn, n, y)))"
-apply (rule strong_replacementI)
+apply (rule strong_replacementI)
apply (rule rallI)
-apply (rename_tac B)
-apply (rule separation_CollectI)
-apply (insert nonempty)
-apply (rule_tac A="{B,z,0,nat}" in subset_LsetE)
-apply (blast intro: L_nat)
+apply (rename_tac B)
+apply (rule separation_CollectI)
+apply (insert nonempty)
+apply (rule_tac A="{B,z,0,nat}" in subset_LsetE)
+apply (blast intro: L_nat)
apply (rule ReflectsE [OF formula_replacement2_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac v)
+apply (rename_tac v)
apply (rule bex_iff_sats conj_iff_sats)+
apply (rule_tac env = "[u,v,B,0,nat]" in mem_iff_sats)
apply (rule sep_rules | simp)+
@@ -1006,7 +1008,7 @@
lemma Inl_type [TC]:
"[| x \<in> nat; z \<in> nat |] ==> Inl_fm(x,z) \<in> formula"
-by (simp add: Inl_fm_def)
+by (simp add: Inl_fm_def)
lemma sats_Inl_fm [simp]:
"[| x \<in> nat; z \<in> nat; env \<in> list(A)|]
@@ -1014,16 +1016,16 @@
by (simp add: Inl_fm_def is_Inl_def)
lemma Inl_iff_sats:
- "[| nth(i,env) = x; nth(k,env) = z;
+ "[| nth(i,env) = x; nth(k,env) = z;
i \<in> nat; k \<in> nat; env \<in> list(A)|]
==> is_Inl(**A, x, z) <-> sats(A, Inl_fm(i,k), env)"
by simp
theorem Inl_reflection:
- "REFLECTS[\<lambda>x. is_Inl(L,f(x),h(x)),
+ "REFLECTS[\<lambda>x. is_Inl(L,f(x),h(x)),
\<lambda>i x. is_Inl(**Lset(i),f(x),h(x))]"
apply (simp only: is_Inl_def setclass_simps)
-apply (intro FOL_reflections function_reflections)
+apply (intro FOL_reflections function_reflections)
done
@@ -1035,7 +1037,7 @@
lemma Inr_type [TC]:
"[| x \<in> nat; z \<in> nat |] ==> Inr_fm(x,z) \<in> formula"
-by (simp add: Inr_fm_def)
+by (simp add: Inr_fm_def)
lemma sats_Inr_fm [simp]:
"[| x \<in> nat; z \<in> nat; env \<in> list(A)|]
@@ -1043,16 +1045,16 @@
by (simp add: Inr_fm_def is_Inr_def)
lemma Inr_iff_sats:
- "[| nth(i,env) = x; nth(k,env) = z;
+ "[| nth(i,env) = x; nth(k,env) = z;
i \<in> nat; k \<in> nat; env \<in> list(A)|]
==> is_Inr(**A, x, z) <-> sats(A, Inr_fm(i,k), env)"
by simp
theorem Inr_reflection:
- "REFLECTS[\<lambda>x. is_Inr(L,f(x),h(x)),
+ "REFLECTS[\<lambda>x. is_Inr(L,f(x),h(x)),
\<lambda>i x. is_Inr(**Lset(i),f(x),h(x))]"
apply (simp only: is_Inr_def setclass_simps)
-apply (intro FOL_reflections function_reflections)
+apply (intro FOL_reflections function_reflections)
done
@@ -1062,9 +1064,9 @@
constdefs Nil_fm :: "i=>i"
"Nil_fm(x) == Exists(And(empty_fm(0), Inl_fm(0,succ(x))))"
-
+
lemma Nil_type [TC]: "x \<in> nat ==> Nil_fm(x) \<in> formula"
-by (simp add: Nil_fm_def)
+by (simp add: Nil_fm_def)
lemma sats_Nil_fm [simp]:
"[| x \<in> nat; env \<in> list(A)|]
@@ -1077,10 +1079,10 @@
by simp
theorem Nil_reflection:
- "REFLECTS[\<lambda>x. is_Nil(L,f(x)),
+ "REFLECTS[\<lambda>x. is_Nil(L,f(x)),
\<lambda>i x. is_Nil(**Lset(i),f(x))]"
apply (simp only: is_Nil_def setclass_simps)
-apply (intro FOL_reflections function_reflections Inl_reflection)
+apply (intro FOL_reflections function_reflections Inl_reflection)
done
@@ -1089,30 +1091,30 @@
(* "is_Cons(M,a,l,Z) == \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" *)
constdefs Cons_fm :: "[i,i,i]=>i"
- "Cons_fm(a,l,Z) ==
+ "Cons_fm(a,l,Z) ==
Exists(And(pair_fm(succ(a),succ(l),0), Inr_fm(0,succ(Z))))"
lemma Cons_type [TC]:
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> Cons_fm(x,y,z) \<in> formula"
-by (simp add: Cons_fm_def)
+by (simp add: Cons_fm_def)
lemma sats_Cons_fm [simp]:
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> 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 \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
==>is_Cons(**A, x, y, z) <-> sats(A, Cons_fm(i,j,k), env)"
by simp
theorem Cons_reflection:
- "REFLECTS[\<lambda>x. is_Cons(L,f(x),g(x),h(x)),
+ "REFLECTS[\<lambda>x. is_Cons(L,f(x),g(x),h(x)),
\<lambda>i x. is_Cons(**Lset(i),f(x),g(x),h(x))]"
apply (simp only: is_Cons_def setclass_simps)
-apply (intro FOL_reflections pair_reflection Inr_reflection)
+apply (intro FOL_reflections pair_reflection Inr_reflection)
done
subsubsection{*The Formula @{term is_quasilist}, Internalized*}
@@ -1120,11 +1122,11 @@
(* is_quasilist(M,xs) == is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))" *)
constdefs quasilist_fm :: "i=>i"
- "quasilist_fm(x) ==
+ "quasilist_fm(x) ==
Or(Nil_fm(x), Exists(Exists(Cons_fm(1,0,succ(succ(x))))))"
-
+
lemma quasilist_type [TC]: "x \<in> nat ==> quasilist_fm(x) \<in> formula"
-by (simp add: quasilist_fm_def)
+by (simp add: quasilist_fm_def)
lemma sats_quasilist_fm [simp]:
"[| x \<in> nat; env \<in> list(A)|]
@@ -1137,10 +1139,10 @@
by simp
theorem quasilist_reflection:
- "REFLECTS[\<lambda>x. is_quasilist(L,f(x)),
+ "REFLECTS[\<lambda>x. is_quasilist(L,f(x)),
\<lambda>i x. is_quasilist(**Lset(i),f(x))]"
apply (simp only: is_quasilist_def setclass_simps)
-apply (intro FOL_reflections Nil_reflection Cons_reflection)
+apply (intro FOL_reflections Nil_reflection Cons_reflection)
done
@@ -1149,19 +1151,19 @@
subsubsection{*The Formula @{term is_tl}, Internalized*}
-(* "is_tl(M,xs,T) ==
+(* "is_tl(M,xs,T) ==
(is_Nil(M,xs) --> T=xs) &
(\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | T=l) &
(is_quasilist(M,xs) | empty(M,T))" *)
constdefs tl_fm :: "[i,i]=>i"
- "tl_fm(xs,T) ==
+ "tl_fm(xs,T) ==
And(Implies(Nil_fm(xs), Equal(T,xs)),
And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(T#+2,0)))),
Or(quasilist_fm(xs), empty_fm(T))))"
lemma tl_type [TC]:
"[| x \<in> nat; y \<in> nat |] ==> tl_fm(x,y) \<in> formula"
-by (simp add: tl_fm_def)
+by (simp add: tl_fm_def)
lemma sats_tl_fm [simp]:
"[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
@@ -1175,11 +1177,11 @@
by simp
theorem tl_reflection:
- "REFLECTS[\<lambda>x. is_tl(L,f(x),g(x)),
+ "REFLECTS[\<lambda>x. is_tl(L,f(x),g(x)),
\<lambda>i x. is_tl(**Lset(i),f(x),g(x))]"
apply (simp only: is_tl_def setclass_simps)
apply (intro FOL_reflections Nil_reflection Cons_reflection
- quasilist_reflection empty_reflection)
+ quasilist_reflection empty_reflection)
done
@@ -1190,27 +1192,27 @@
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>
is_wfrec(L, iterates_MH(L, is_tl(L), z), memsn, u, y)),
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(**Lset(i), u, y, x) \<and>
- is_wfrec(**Lset(i),
- iterates_MH(**Lset(i),
+ is_wfrec(**Lset(i),
+ iterates_MH(**Lset(i),
is_tl(**Lset(i)), z), memsn, u, y))]"
-by (intro FOL_reflections function_reflections is_wfrec_reflection
- iterates_MH_reflection list_functor_reflection tl_reflection)
+by (intro FOL_reflections function_reflections is_wfrec_reflection
+ iterates_MH_reflection list_functor_reflection tl_reflection)
-lemma nth_replacement:
+lemma nth_replacement:
"L(w) ==> iterates_replacement(L, %l t. is_tl(L,l,t), w)"
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
-apply (rule strong_replacementI)
-apply (rule rallI)
-apply (rule separation_CollectI)
-apply (subgoal_tac "L(Memrel(succ(n)))")
-apply (rule_tac A="{A,n,w,z,Memrel(succ(n))}" in subset_LsetE, blast )
+apply (rule strong_replacementI)
+apply (rule rallI)
+apply (rule separation_CollectI)
+apply (subgoal_tac "L(Memrel(succ(n)))")
+apply (rule_tac A="{A,n,w,z,Memrel(succ(n))}" in subset_LsetE, blast )
apply (rule ReflectsE [OF nth_replacement_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2 Memrel_closed)
-apply (elim conjE)
+apply (elim conjE)
apply (rule DPow_LsetI)
-apply (rename_tac v)
+apply (rename_tac v)
apply (rule bex_iff_sats conj_iff_sats)+
apply (rule_tac env = "[u,v,A,z,w,Memrel(succ(n))]" in mem_iff_sats)
apply (rule sep_rules | simp)+
@@ -1221,27 +1223,29 @@
subsubsection{*Instantiating the locale @{text M_datatypes}*}
-ML
-{*
-val list_replacement1 = thm "list_replacement1";
-val list_replacement2 = thm "list_replacement2";
-val formula_replacement1 = thm "formula_replacement1";
-val formula_replacement2 = thm "formula_replacement2";
-val nth_replacement = thm "nth_replacement";
+
+theorem M_datatypes_axioms_L: "M_datatypes_axioms(L)"
+ apply (rule M_datatypes_axioms.intro)
+ apply (assumption | rule
+ list_replacement1 list_replacement2
+ formula_replacement1 formula_replacement2
+ nth_replacement)+
+ done
-val m_datatypes = [list_replacement1, list_replacement2,
- formula_replacement1, formula_replacement2,
- nth_replacement];
-
-fun datatypes_L th =
- kill_flex_triv_prems (m_datatypes MRS (wfrank_L th));
+theorem M_datatypes_L: "PROP M_datatypes(L)"
+ apply (rule M_datatypes.intro)
+ apply (rule M_triv_axioms_L)
+ apply (rule M_axioms_axioms_L)
+ apply (rule M_trancl_axioms_L)
+ apply (rule M_wfrank_axioms_L)
+ apply (rule M_datatypes_axioms_L)
+ done
-bind_thm ("list_closed", datatypes_L (thm "M_datatypes.list_closed"));
-bind_thm ("formula_closed", datatypes_L (thm "M_datatypes.formula_closed"));
-bind_thm ("list_abs", datatypes_L (thm "M_datatypes.list_abs"));
-bind_thm ("formula_abs", datatypes_L (thm "M_datatypes.formula_abs"));
-bind_thm ("nth_abs", datatypes_L (thm "M_datatypes.nth_abs"));
-*}
+lemmas list_closed = M_datatypes.list_closed [OF M_datatypes_L]
+ and formula_closed = M_datatypes.formula_closed [OF M_datatypes_L]
+ and list_abs = M_datatypes.list_abs [OF M_datatypes_L]
+ and formula_abs = M_datatypes.formula_abs [OF M_datatypes_L]
+ and nth_abs = M_datatypes.nth_abs [OF M_datatypes_L]
declare list_closed [intro,simp]
declare formula_closed [intro,simp]
@@ -1250,8 +1254,7 @@
declare nth_abs [simp]
-
-subsection{*@{term L} is Closed Under the Operator @{term eclose}*}
+subsection{*@{term L} is Closed Under the Operator @{term eclose}*}
subsubsection{*Instances of Replacement for @{term eclose}*}
@@ -1260,28 +1263,28 @@
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>
is_wfrec(L, iterates_MH(L, big_union(L), A), memsn, u, y)),
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(**Lset(i), u, y, x) \<and>
- is_wfrec(**Lset(i),
- iterates_MH(**Lset(i), big_union(**Lset(i)), A),
+ is_wfrec(**Lset(i),
+ iterates_MH(**Lset(i), big_union(**Lset(i)), A),
memsn, u, y))]"
-by (intro FOL_reflections function_reflections is_wfrec_reflection
- iterates_MH_reflection)
+by (intro FOL_reflections function_reflections is_wfrec_reflection
+ iterates_MH_reflection)
-lemma eclose_replacement1:
+lemma eclose_replacement1:
"L(A) ==> iterates_replacement(L, big_union(L), A)"
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
-apply (rule strong_replacementI)
+apply (rule strong_replacementI)
apply (rule rallI)
-apply (rename_tac B)
-apply (rule separation_CollectI)
-apply (subgoal_tac "L(Memrel(succ(n)))")
-apply (rule_tac A="{B,A,n,z,Memrel(succ(n))}" in subset_LsetE, blast )
+apply (rename_tac B)
+apply (rule separation_CollectI)
+apply (subgoal_tac "L(Memrel(succ(n)))")
+apply (rule_tac A="{B,A,n,z,Memrel(succ(n))}" in subset_LsetE, blast )
apply (rule ReflectsE [OF eclose_replacement1_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2 Memrel_closed)
-apply (elim conjE)
+apply (elim conjE)
apply (rule DPow_LsetI)
-apply (rename_tac v)
+apply (rename_tac v)
apply (rule bex_iff_sats conj_iff_sats)+
apply (rule_tac env = "[u,v,A,n,B,Memrel(succ(n))]" in mem_iff_sats)
apply (rule sep_rules | simp)+
@@ -1298,33 +1301,33 @@
is_wfrec (L, iterates_MH (L, big_union(L), A),
msn, u, x)),
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and>
- (\<exists>sn \<in> Lset(i). \<exists>msn \<in> Lset(i).
+ (\<exists>sn \<in> Lset(i). \<exists>msn \<in> Lset(i).
successor(**Lset(i), u, sn) \<and> membership(**Lset(i), sn, msn) \<and>
- is_wfrec (**Lset(i),
+ is_wfrec (**Lset(i),
iterates_MH (**Lset(i), big_union(**Lset(i)), A),
msn, u, x))]"
-by (intro FOL_reflections function_reflections is_wfrec_reflection
- iterates_MH_reflection)
+by (intro FOL_reflections function_reflections is_wfrec_reflection
+ iterates_MH_reflection)
-lemma eclose_replacement2:
- "L(A) ==> strong_replacement(L,
- \<lambda>n y. n\<in>nat &
+lemma eclose_replacement2:
+ "L(A) ==> strong_replacement(L,
+ \<lambda>n y. n\<in>nat &
(\<exists>sn[L]. \<exists>msn[L]. successor(L,n,sn) & membership(L,sn,msn) &
- is_wfrec(L, iterates_MH(L,big_union(L), A),
+ is_wfrec(L, iterates_MH(L,big_union(L), A),
msn, n, y)))"
-apply (rule strong_replacementI)
+apply (rule strong_replacementI)
apply (rule rallI)
-apply (rename_tac B)
-apply (rule separation_CollectI)
-apply (rule_tac A="{A,B,z,nat}" in subset_LsetE)
-apply (blast intro: L_nat)
+apply (rename_tac B)
+apply (rule separation_CollectI)
+apply (rule_tac A="{A,B,z,nat}" in subset_LsetE)
+apply (blast intro: L_nat)
apply (rule ReflectsE [OF eclose_replacement2_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac v)
+apply (rename_tac v)
apply (rule bex_iff_sats conj_iff_sats)+
apply (rule_tac env = "[u,v,A,B,nat]" in mem_iff_sats)
apply (rule sep_rules | simp)+
@@ -1334,22 +1337,23 @@
subsubsection{*Instantiating the locale @{text M_eclose}*}
-ML
-{*
-val eclose_replacement1 = thm "eclose_replacement1";
-val eclose_replacement2 = thm "eclose_replacement2";
-val m_eclose = [eclose_replacement1, eclose_replacement2];
+theorem M_eclose_axioms_L: "M_eclose_axioms(L)"
+ apply (rule M_eclose_axioms.intro)
+ apply (assumption | rule eclose_replacement1 eclose_replacement2)+
+ done
-fun eclose_L th =
- kill_flex_triv_prems (m_eclose MRS (datatypes_L th));
+theorem M_eclose_L: "PROP M_eclose(L)"
+ apply (rule M_eclose.intro)
+ apply (rule M_triv_axioms_L)
+ apply (rule M_axioms_axioms_L)
+ apply (rule M_trancl_axioms_L)
+ apply (rule M_wfrank_axioms_L)
+ apply (rule M_datatypes_axioms_L)
+ apply (rule M_eclose_axioms_L)
+ done
-bind_thm ("eclose_closed", eclose_L (thm "M_eclose.eclose_closed"));
-bind_thm ("eclose_abs", eclose_L (thm "M_eclose.eclose_abs"));
-*}
-
-declare eclose_closed [intro,simp]
-declare eclose_abs [intro,simp]
-
+lemmas eclose_closed [intro, simp] = M_eclose.eclose_closed [OF M_eclose_L]
+ and eclose_abs [intro, simp] = M_eclose.eclose_abs [OF M_eclose_L]
end
--- a/src/ZF/Constructible/Separation.thy Sun Jul 28 21:09:37 2002 +0200
+++ b/src/ZF/Constructible/Separation.thy Mon Jul 29 00:57:16 2002 +0200
@@ -9,39 +9,39 @@
by simp
lemmas nth_rules = nth_0 nth_ConsI nat_0I nat_succI
-lemmas sep_rules = nth_0 nth_ConsI FOL_iff_sats function_iff_sats
+lemmas sep_rules = nth_0 nth_ConsI FOL_iff_sats function_iff_sats
fun_plus_iff_sats
lemma Collect_conj_in_DPow:
- "[| {x\<in>A. P(x)} \<in> DPow(A); {x\<in>A. Q(x)} \<in> DPow(A) |]
+ "[| {x\<in>A. P(x)} \<in> DPow(A); {x\<in>A. Q(x)} \<in> DPow(A) |]
==> {x\<in>A. P(x) & Q(x)} \<in> DPow(A)"
-by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric])
+by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric])
lemma Collect_conj_in_DPow_Lset:
"[|z \<in> Lset(j); {x \<in> Lset(j). P(x)} \<in> DPow(Lset(j))|]
==> {x \<in> Lset(j). x \<in> z & P(x)} \<in> DPow(Lset(j))"
apply (frule mem_Lset_imp_subset_Lset)
-apply (simp add: Collect_conj_in_DPow Collect_mem_eq
+apply (simp add: Collect_conj_in_DPow Collect_mem_eq
subset_Int_iff2 elem_subset_in_DPow)
done
lemma separation_CollectI:
"(\<And>z. L(z) ==> L({x \<in> z . P(x)})) ==> separation(L, \<lambda>x. P(x))"
-apply (unfold separation_def, clarify)
-apply (rule_tac x="{x\<in>z. P(x)}" in rexI)
+apply (unfold separation_def, clarify)
+apply (rule_tac x="{x\<in>z. P(x)}" in rexI)
apply simp_all
done
text{*Reduces the original comprehension to the reflected one*}
lemma reflection_imp_L_separation:
"[| \<forall>x\<in>Lset(j). P(x) <-> Q(x);
- {x \<in> Lset(j) . Q(x)} \<in> DPow(Lset(j));
+ {x \<in> Lset(j) . Q(x)} \<in> DPow(Lset(j));
Ord(j); z \<in> Lset(j)|] ==> L({x \<in> z . P(x)})"
apply (rule_tac i = "succ(j)" in L_I)
prefer 2 apply simp
apply (subgoal_tac "{x \<in> z. P(x)} = {x \<in> Lset(j). x \<in> z & (Q(x))}")
prefer 2
- apply (blast dest: mem_Lset_imp_subset_Lset)
+ apply (blast dest: mem_Lset_imp_subset_Lset)
apply (simp add: Lset_succ Collect_conj_in_DPow_Lset)
done
@@ -49,20 +49,20 @@
subsection{*Separation for Intersection*}
lemma Inter_Reflects:
- "REFLECTS[\<lambda>x. \<forall>y[L]. y\<in>A --> x \<in> y,
+ "REFLECTS[\<lambda>x. \<forall>y[L]. y\<in>A --> x \<in> y,
\<lambda>i x. \<forall>y\<in>Lset(i). y\<in>A --> x \<in> y]"
-by (intro FOL_reflections)
+by (intro FOL_reflections)
lemma Inter_separation:
"L(A) ==> separation(L, \<lambda>x. \<forall>y[L]. y\<in>A --> x\<in>y)"
-apply (rule separation_CollectI)
-apply (rule_tac A="{A,z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{A,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF Inter_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2, clarify)
-apply (rule DPow_LsetI)
-apply (rule ball_iff_sats)
+apply (rule DPow_LsetI)
+apply (rule ball_iff_sats)
apply (rule imp_iff_sats)
apply (rule_tac [2] i=1 and j=0 and env="[y,x,A]" in mem_iff_sats)
apply (rule_tac i=0 and j=2 in mem_iff_sats)
@@ -73,22 +73,22 @@
lemma cartprod_Reflects:
"REFLECTS[\<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)),
- \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). y\<in>B &
+ \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). y\<in>B &
pair(**Lset(i),x,y,z))]"
by (intro FOL_reflections function_reflections)
lemma cartprod_separation:
- "[| L(A); L(B) |]
+ "[| L(A); L(B) |]
==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{A,B,z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{A,B,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF cartprod_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
- apply (simp_all add: lt_Ord2, clarify)
+ apply (simp_all add: lt_Ord2, clarify)
apply (rule DPow_LsetI)
-apply (rename_tac u)
-apply (rule bex_iff_sats)
+apply (rename_tac u)
+apply (rule bex_iff_sats)
apply (rule conj_iff_sats)
apply (rule_tac i=0 and j=2 and env="[x,u,A,B]" in mem_iff_sats, simp_all)
apply (rule sep_rules | simp)+
@@ -102,16 +102,16 @@
by (intro FOL_reflections function_reflections)
lemma image_separation:
- "[| L(A); L(r) |]
+ "[| L(A); L(r) |]
==> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{A,r,z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{A,r,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF image_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2, clarify)
apply (rule DPow_LsetI)
-apply (rule bex_iff_sats)
+apply (rule bex_iff_sats)
apply (rule conj_iff_sats)
apply (rule_tac env="[p,y,A,r]" in mem_iff_sats)
apply (rule sep_rules | simp)+
@@ -122,22 +122,22 @@
lemma converse_Reflects:
"REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)),
- \<lambda>i z. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i).
+ \<lambda>i z. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i).
pair(**Lset(i),x,y,p) & pair(**Lset(i),y,x,z))]"
by (intro FOL_reflections function_reflections)
lemma converse_separation:
- "L(r) ==> separation(L,
+ "L(r) ==> separation(L,
\<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{r,z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{r,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF converse_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2, clarify)
apply (rule DPow_LsetI)
-apply (rename_tac u)
-apply (rule bex_iff_sats)
+apply (rename_tac u)
+apply (rule bex_iff_sats)
apply (rule conj_iff_sats)
apply (rule_tac i=0 and j=2 and env="[p,u,r]" in mem_iff_sats, simp_all)
apply (rule sep_rules | simp)+
@@ -153,15 +153,15 @@
lemma restrict_separation:
"L(A) ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{A,z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{A,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF restrict_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2, clarify)
apply (rule DPow_LsetI)
-apply (rename_tac u)
-apply (rule bex_iff_sats)
+apply (rename_tac u)
+apply (rule bex_iff_sats)
apply (rule conj_iff_sats)
apply (rule_tac i=0 and j=2 and env="[x,u,A]" in mem_iff_sats, simp_all)
apply (rule sep_rules | simp)+
@@ -171,29 +171,29 @@
subsection{*Separation for Composition*}
lemma comp_Reflects:
- "REFLECTS[\<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
- pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
+ "REFLECTS[\<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
+ pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
xy\<in>s & yz\<in>r,
- \<lambda>i xz. \<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i). \<exists>z\<in>Lset(i). \<exists>xy\<in>Lset(i). \<exists>yz\<in>Lset(i).
- pair(**Lset(i),x,z,xz) & pair(**Lset(i),x,y,xy) &
+ \<lambda>i xz. \<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i). \<exists>z\<in>Lset(i). \<exists>xy\<in>Lset(i). \<exists>yz\<in>Lset(i).
+ pair(**Lset(i),x,z,xz) & pair(**Lset(i),x,y,xy) &
pair(**Lset(i),y,z,yz) & xy\<in>s & yz\<in>r]"
by (intro FOL_reflections function_reflections)
lemma comp_separation:
"[| L(r); L(s) |]
- ==> separation(L, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
- pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
+ ==> separation(L, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
+ pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
xy\<in>s & yz\<in>r)"
-apply (rule separation_CollectI)
-apply (rule_tac A="{r,s,z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{r,s,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF comp_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2, clarify)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule bex_iff_sats)+
-apply (rename_tac x y z)
+apply (rename_tac x y z)
apply (rule conj_iff_sats)
apply (rule_tac env="[z,y,x,u,r,s]" in pair_iff_sats)
apply (rule sep_rules | simp)+
@@ -208,17 +208,17 @@
lemma pred_separation:
"[| L(r); L(x) |] ==> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & pair(L,y,x,p))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{r,x,z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{r,x,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF pred_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2, clarify)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule bex_iff_sats)
apply (rule conj_iff_sats)
-apply (rule_tac env = "[p,u,r,x]" in mem_iff_sats)
+apply (rule_tac env = "[p,u,r,x]" in mem_iff_sats)
apply (rule sep_rules | simp)+
done
@@ -232,50 +232,50 @@
lemma Memrel_separation:
"separation(L, \<lambda>z. \<exists>x[L]. \<exists>y[L]. pair(L,x,y,z) & x \<in> y)"
-apply (rule separation_CollectI)
-apply (rule_tac A="{z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF Memrel_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule bex_iff_sats conj_iff_sats)+
-apply (rule_tac env = "[y,x,u]" in pair_iff_sats)
+apply (rule_tac env = "[y,x,u]" in pair_iff_sats)
apply (rule sep_rules | simp)+
done
subsection{*Replacement for FunSpace*}
-
+
lemma funspace_succ_Reflects:
- "REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>A & (\<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L].
- pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) &
- upair(L,cnbf,cnbf,z)),
- \<lambda>i z. \<exists>p \<in> Lset(i). p\<in>A & (\<exists>f \<in> Lset(i). \<exists>b \<in> Lset(i).
- \<exists>nb \<in> Lset(i). \<exists>cnbf \<in> Lset(i).
- pair(**Lset(i),f,b,p) & pair(**Lset(i),n,b,nb) &
- is_cons(**Lset(i),nb,f,cnbf) & upair(**Lset(i),cnbf,cnbf,z))]"
+ "REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>A & (\<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L].
+ pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) &
+ upair(L,cnbf,cnbf,z)),
+ \<lambda>i z. \<exists>p \<in> Lset(i). p\<in>A & (\<exists>f \<in> Lset(i). \<exists>b \<in> Lset(i).
+ \<exists>nb \<in> Lset(i). \<exists>cnbf \<in> Lset(i).
+ pair(**Lset(i),f,b,p) & pair(**Lset(i),n,b,nb) &
+ is_cons(**Lset(i),nb,f,cnbf) & upair(**Lset(i),cnbf,cnbf,z))]"
by (intro FOL_reflections function_reflections)
lemma funspace_succ_replacement:
- "L(n) ==>
- strong_replacement(L, \<lambda>p z. \<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L].
+ "L(n) ==>
+ strong_replacement(L, \<lambda>p z. \<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L].
pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) &
upair(L,cnbf,cnbf,z))"
-apply (rule strong_replacementI)
-apply (rule rallI)
-apply (rule separation_CollectI)
-apply (rule_tac A="{n,A,z}" in subset_LsetE, blast )
+apply (rule strong_replacementI)
+apply (rule rallI)
+apply (rule separation_CollectI)
+apply (rule_tac A="{n,A,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF funspace_succ_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule bex_iff_sats)
apply (rule conj_iff_sats)
-apply (rule_tac env = "[p,u,n,A]" in mem_iff_sats)
+apply (rule_tac env = "[p,u,n,A]" in mem_iff_sats)
apply (rule sep_rules | simp)+
done
@@ -283,26 +283,26 @@
subsection{*Separation for Order-Isomorphisms*}
lemma well_ord_iso_Reflects:
- "REFLECTS[\<lambda>x. x\<in>A -->
+ "REFLECTS[\<lambda>x. x\<in>A -->
(\<exists>y[L]. \<exists>p[L]. fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r),
- \<lambda>i x. x\<in>A --> (\<exists>y \<in> Lset(i). \<exists>p \<in> Lset(i).
+ \<lambda>i x. x\<in>A --> (\<exists>y \<in> Lset(i). \<exists>p \<in> Lset(i).
fun_apply(**Lset(i),f,x,y) & pair(**Lset(i),y,x,p) & p \<in> r)]"
by (intro FOL_reflections function_reflections)
lemma well_ord_iso_separation:
- "[| L(A); L(f); L(r) |]
- ==> separation (L, \<lambda>x. x\<in>A --> (\<exists>y[L]. (\<exists>p[L].
- fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r)))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{A,f,r,z}" in subset_LsetE, blast )
+ "[| L(A); L(f); L(r) |]
+ ==> separation (L, \<lambda>x. x\<in>A --> (\<exists>y[L]. (\<exists>p[L].
+ fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r)))"
+apply (rule separation_CollectI)
+apply (rule_tac A="{A,f,r,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF well_ord_iso_Reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule imp_iff_sats)
-apply (rule_tac env = "[u,A,f,r]" in mem_iff_sats)
+apply (rule_tac env = "[u,A,f,r]" in mem_iff_sats)
apply (rule sep_rules | simp)+
done
@@ -310,31 +310,31 @@
subsection{*Separation for @{term "obase"}*}
lemma obase_reflects:
- "REFLECTS[\<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
- ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
- order_isomorphism(L,par,r,x,mx,g),
- \<lambda>i a. \<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i). \<exists>par \<in> Lset(i).
- ordinal(**Lset(i),x) & membership(**Lset(i),x,mx) & pred_set(**Lset(i),A,a,r,par) &
- order_isomorphism(**Lset(i),par,r,x,mx,g)]"
+ "REFLECTS[\<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
+ ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
+ order_isomorphism(L,par,r,x,mx,g),
+ \<lambda>i a. \<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i). \<exists>par \<in> Lset(i).
+ ordinal(**Lset(i),x) & membership(**Lset(i),x,mx) & pred_set(**Lset(i),A,a,r,par) &
+ order_isomorphism(**Lset(i),par,r,x,mx,g)]"
by (intro FOL_reflections function_reflections fun_plus_reflections)
lemma obase_separation:
--{*part of the order type formalization*}
- "[| L(A); L(r) |]
- ==> separation(L, \<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
- ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
- order_isomorphism(L,par,r,x,mx,g))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{A,r,z}" in subset_LsetE, blast )
+ "[| L(A); L(r) |]
+ ==> separation(L, \<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
+ ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
+ order_isomorphism(L,par,r,x,mx,g))"
+apply (rule separation_CollectI)
+apply (rule_tac A="{A,r,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF obase_reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule bex_iff_sats)
apply (rule conj_iff_sats)
-apply (rule_tac env = "[x,u,A,r]" in ordinal_iff_sats)
+apply (rule_tac env = "[x,u,A,r]" in ordinal_iff_sats)
apply (rule sep_rules | simp)+
done
@@ -342,33 +342,33 @@
subsection{*Separation for a Theorem about @{term "obase"}*}
lemma obase_equals_reflects:
- "REFLECTS[\<lambda>x. x\<in>A --> ~(\<exists>y[L]. \<exists>g[L].
- ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
- membership(L,y,my) & pred_set(L,A,x,r,pxr) &
- order_isomorphism(L,pxr,r,y,my,g))),
- \<lambda>i x. x\<in>A --> ~(\<exists>y \<in> Lset(i). \<exists>g \<in> Lset(i).
- ordinal(**Lset(i),y) & (\<exists>my \<in> Lset(i). \<exists>pxr \<in> Lset(i).
- membership(**Lset(i),y,my) & pred_set(**Lset(i),A,x,r,pxr) &
- order_isomorphism(**Lset(i),pxr,r,y,my,g)))]"
+ "REFLECTS[\<lambda>x. x\<in>A --> ~(\<exists>y[L]. \<exists>g[L].
+ ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
+ membership(L,y,my) & pred_set(L,A,x,r,pxr) &
+ order_isomorphism(L,pxr,r,y,my,g))),
+ \<lambda>i x. x\<in>A --> ~(\<exists>y \<in> Lset(i). \<exists>g \<in> Lset(i).
+ ordinal(**Lset(i),y) & (\<exists>my \<in> Lset(i). \<exists>pxr \<in> Lset(i).
+ membership(**Lset(i),y,my) & pred_set(**Lset(i),A,x,r,pxr) &
+ order_isomorphism(**Lset(i),pxr,r,y,my,g)))]"
by (intro FOL_reflections function_reflections fun_plus_reflections)
lemma obase_equals_separation:
- "[| L(A); L(r) |]
- ==> separation (L, \<lambda>x. x\<in>A --> ~(\<exists>y[L]. \<exists>g[L].
- ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
- membership(L,y,my) & pred_set(L,A,x,r,pxr) &
- order_isomorphism(L,pxr,r,y,my,g))))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{A,r,z}" in subset_LsetE, blast )
+ "[| L(A); L(r) |]
+ ==> separation (L, \<lambda>x. x\<in>A --> ~(\<exists>y[L]. \<exists>g[L].
+ ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
+ membership(L,y,my) & pred_set(L,A,x,r,pxr) &
+ order_isomorphism(L,pxr,r,y,my,g))))"
+apply (rule separation_CollectI)
+apply (rule_tac A="{A,r,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF obase_equals_reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule imp_iff_sats ball_iff_sats disj_iff_sats not_iff_sats)+
-apply (rule_tac env = "[u,A,r]" in mem_iff_sats)
+apply (rule_tac env = "[u,A,r]" in mem_iff_sats)
apply (rule sep_rules | simp)+
done
@@ -376,35 +376,35 @@
subsection{*Replacement for @{term "omap"}*}
lemma omap_reflects:
- "REFLECTS[\<lambda>z. \<exists>a[L]. a\<in>B & (\<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
- ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) &
+ "REFLECTS[\<lambda>z. \<exists>a[L]. a\<in>B & (\<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
+ ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) &
pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g)),
- \<lambda>i z. \<exists>a \<in> Lset(i). a\<in>B & (\<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i).
- \<exists>par \<in> Lset(i).
- ordinal(**Lset(i),x) & pair(**Lset(i),a,x,z) &
- membership(**Lset(i),x,mx) & pred_set(**Lset(i),A,a,r,par) &
+ \<lambda>i z. \<exists>a \<in> Lset(i). a\<in>B & (\<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i).
+ \<exists>par \<in> Lset(i).
+ ordinal(**Lset(i),x) & pair(**Lset(i),a,x,z) &
+ membership(**Lset(i),x,mx) & pred_set(**Lset(i),A,a,r,par) &
order_isomorphism(**Lset(i),par,r,x,mx,g))]"
by (intro FOL_reflections function_reflections fun_plus_reflections)
lemma omap_replacement:
- "[| L(A); L(r) |]
+ "[| L(A); L(r) |]
==> strong_replacement(L,
- \<lambda>a z. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
- ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) &
- pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g))"
-apply (rule strong_replacementI)
+ \<lambda>a z. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
+ ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) &
+ pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g))"
+apply (rule strong_replacementI)
apply (rule rallI)
-apply (rename_tac B)
-apply (rule separation_CollectI)
-apply (rule_tac A="{A,B,r,z}" in subset_LsetE, blast )
+apply (rename_tac B)
+apply (rule separation_CollectI)
+apply (rule_tac A="{A,B,r,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF omap_reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule bex_iff_sats conj_iff_sats)+
-apply (rule_tac env = "[a,u,A,B,r]" in mem_iff_sats)
+apply (rule_tac env = "[a,u,A,B,r]" in mem_iff_sats)
apply (rule sep_rules | simp)+
done
@@ -412,34 +412,34 @@
subsection{*Separation for a Theorem about @{term "obase"}*}
lemma is_recfun_reflects:
- "REFLECTS[\<lambda>x. \<exists>xa[L]. \<exists>xb[L].
- pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r &
- (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) &
+ "REFLECTS[\<lambda>x. \<exists>xa[L]. \<exists>xb[L].
+ pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r &
+ (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) &
fx \<noteq> gx),
- \<lambda>i x. \<exists>xa \<in> Lset(i). \<exists>xb \<in> Lset(i).
+ \<lambda>i x. \<exists>xa \<in> Lset(i). \<exists>xb \<in> Lset(i).
pair(**Lset(i),x,a,xa) & xa \<in> r & pair(**Lset(i),x,b,xb) & xb \<in> r &
- (\<exists>fx \<in> Lset(i). \<exists>gx \<in> Lset(i). fun_apply(**Lset(i),f,x,fx) &
+ (\<exists>fx \<in> Lset(i). \<exists>gx \<in> Lset(i). fun_apply(**Lset(i),f,x,fx) &
fun_apply(**Lset(i),g,x,gx) & fx \<noteq> gx)]"
by (intro FOL_reflections function_reflections fun_plus_reflections)
lemma is_recfun_separation:
--{*for well-founded recursion*}
- "[| L(r); L(f); L(g); L(a); L(b) |]
- ==> separation(L,
- \<lambda>x. \<exists>xa[L]. \<exists>xb[L].
- pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r &
- (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) &
+ "[| L(r); L(f); L(g); L(a); L(b) |]
+ ==> separation(L,
+ \<lambda>x. \<exists>xa[L]. \<exists>xb[L].
+ pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r &
+ (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) &
fx \<noteq> gx))"
-apply (rule separation_CollectI)
-apply (rule_tac A="{r,f,g,a,b,z}" in subset_LsetE, blast )
+apply (rule separation_CollectI)
+apply (rule_tac A="{r,f,g,a,b,z}" in subset_LsetE, blast )
apply (rule ReflectsE [OF is_recfun_reflects], assumption)
-apply (drule subset_Lset_ltD, assumption)
+apply (drule subset_Lset_ltD, assumption)
apply (erule reflection_imp_L_separation)
apply (simp_all add: lt_Ord2)
apply (rule DPow_LsetI)
-apply (rename_tac u)
+apply (rename_tac u)
apply (rule bex_iff_sats conj_iff_sats)+
-apply (rule_tac env = "[xa,u,r,f,g,a,b]" in pair_iff_sats)
+apply (rule_tac env = "[xa,u,r,f,g,a,b]" in pair_iff_sats)
apply (rule sep_rules | simp)+
done
@@ -448,144 +448,128 @@
text{*Separation (and Strong Replacement) for basic set-theoretic constructions
such as intersection, Cartesian Product and image.*}
-ML
-{*
-val Inter_separation = thm "Inter_separation";
-val cartprod_separation = thm "cartprod_separation";
-val image_separation = thm "image_separation";
-val converse_separation = thm "converse_separation";
-val restrict_separation = thm "restrict_separation";
-val comp_separation = thm "comp_separation";
-val pred_separation = thm "pred_separation";
-val Memrel_separation = thm "Memrel_separation";
-val funspace_succ_replacement = thm "funspace_succ_replacement";
-val well_ord_iso_separation = thm "well_ord_iso_separation";
-val obase_separation = thm "obase_separation";
-val obase_equals_separation = thm "obase_equals_separation";
-val omap_replacement = thm "omap_replacement";
-val is_recfun_separation = thm "is_recfun_separation";
-
-val m_axioms =
- [Inter_separation, cartprod_separation, image_separation,
- converse_separation, restrict_separation, comp_separation,
- pred_separation, Memrel_separation, funspace_succ_replacement,
- well_ord_iso_separation, obase_separation,
- obase_equals_separation, omap_replacement, is_recfun_separation]
-
-fun axioms_L th =
- kill_flex_triv_prems (m_axioms MRS (triv_axioms_L th));
+theorem M_axioms_axioms_L: "M_axioms_axioms(L)"
+ apply (rule M_axioms_axioms.intro)
+ apply (assumption | rule
+ Inter_separation cartprod_separation image_separation
+ converse_separation restrict_separation
+ comp_separation pred_separation Memrel_separation
+ funspace_succ_replacement well_ord_iso_separation
+ obase_separation obase_equals_separation
+ omap_replacement is_recfun_separation)+
+ done
+
+theorem M_axioms_L: "PROP M_axioms(L)"
+ apply (rule M_axioms.intro)
+ apply (rule M_triv_axioms_L)
+ apply (rule M_axioms_axioms_L)
+ done
-bind_thm ("cartprod_iff", axioms_L (thm "M_axioms.cartprod_iff"));
-bind_thm ("cartprod_closed", axioms_L (thm "M_axioms.cartprod_closed"));
-bind_thm ("sum_closed", axioms_L (thm "M_axioms.sum_closed"));
-bind_thm ("M_converse_iff", axioms_L (thm "M_axioms.M_converse_iff"));
-bind_thm ("converse_closed", axioms_L (thm "M_axioms.converse_closed"));
-bind_thm ("converse_abs", axioms_L (thm "M_axioms.converse_abs"));
-bind_thm ("image_closed", axioms_L (thm "M_axioms.image_closed"));
-bind_thm ("vimage_abs", axioms_L (thm "M_axioms.vimage_abs"));
-bind_thm ("vimage_closed", axioms_L (thm "M_axioms.vimage_closed"));
-bind_thm ("domain_abs", axioms_L (thm "M_axioms.domain_abs"));
-bind_thm ("domain_closed", axioms_L (thm "M_axioms.domain_closed"));
-bind_thm ("range_abs", axioms_L (thm "M_axioms.range_abs"));
-bind_thm ("range_closed", axioms_L (thm "M_axioms.range_closed"));
-bind_thm ("field_abs", axioms_L (thm "M_axioms.field_abs"));
-bind_thm ("field_closed", axioms_L (thm "M_axioms.field_closed"));
-bind_thm ("relation_abs", axioms_L (thm "M_axioms.relation_abs"));
-bind_thm ("function_abs", axioms_L (thm "M_axioms.function_abs"));
-bind_thm ("apply_closed", axioms_L (thm "M_axioms.apply_closed"));
-bind_thm ("apply_abs", axioms_L (thm "M_axioms.apply_abs"));
-bind_thm ("typed_function_abs", axioms_L (thm "M_axioms.typed_function_abs"));
-bind_thm ("injection_abs", axioms_L (thm "M_axioms.injection_abs"));
-bind_thm ("surjection_abs", axioms_L (thm "M_axioms.surjection_abs"));
-bind_thm ("bijection_abs", axioms_L (thm "M_axioms.bijection_abs"));
-bind_thm ("M_comp_iff", axioms_L (thm "M_axioms.M_comp_iff"));
-bind_thm ("comp_closed", axioms_L (thm "M_axioms.comp_closed"));
-bind_thm ("composition_abs", axioms_L (thm "M_axioms.composition_abs"));
-bind_thm ("restriction_is_function", axioms_L (thm "M_axioms.restriction_is_function"));
-bind_thm ("restriction_abs", axioms_L (thm "M_axioms.restriction_abs"));
-bind_thm ("M_restrict_iff", axioms_L (thm "M_axioms.M_restrict_iff"));
-bind_thm ("restrict_closed", axioms_L (thm "M_axioms.restrict_closed"));
-bind_thm ("Inter_abs", axioms_L (thm "M_axioms.Inter_abs"));
-bind_thm ("Inter_closed", axioms_L (thm "M_axioms.Inter_closed"));
-bind_thm ("Int_closed", axioms_L (thm "M_axioms.Int_closed"));
-bind_thm ("finite_fun_closed", axioms_L (thm "M_axioms.finite_fun_closed"));
-bind_thm ("is_funspace_abs", axioms_L (thm "M_axioms.is_funspace_abs"));
-bind_thm ("succ_fun_eq2", axioms_L (thm "M_axioms.succ_fun_eq2"));
-bind_thm ("funspace_succ", axioms_L (thm "M_axioms.funspace_succ"));
-bind_thm ("finite_funspace_closed", axioms_L (thm "M_axioms.finite_funspace_closed"));
-*}
+lemmas cartprod_iff = M_axioms.cartprod_iff [OF M_axioms_L]
+ and cartprod_closed = M_axioms.cartprod_closed [OF M_axioms_L]
+ and sum_closed = M_axioms.sum_closed [OF M_axioms_L]
+ and M_converse_iff = M_axioms.M_converse_iff [OF M_axioms_L]
+ and converse_closed = M_axioms.converse_closed [OF M_axioms_L]
+ and converse_abs = M_axioms.converse_abs [OF M_axioms_L]
+ and image_closed = M_axioms.image_closed [OF M_axioms_L]
+ and vimage_abs = M_axioms.vimage_abs [OF M_axioms_L]
+ and vimage_closed = M_axioms.vimage_closed [OF M_axioms_L]
+ and domain_abs = M_axioms.domain_abs [OF M_axioms_L]
+ and domain_closed = M_axioms.domain_closed [OF M_axioms_L]
+ and range_abs = M_axioms.range_abs [OF M_axioms_L]
+ and range_closed = M_axioms.range_closed [OF M_axioms_L]
+ and field_abs = M_axioms.field_abs [OF M_axioms_L]
+ and field_closed = M_axioms.field_closed [OF M_axioms_L]
+ and relation_abs = M_axioms.relation_abs [OF M_axioms_L]
+ and function_abs = M_axioms.function_abs [OF M_axioms_L]
+ and apply_closed = M_axioms.apply_closed [OF M_axioms_L]
+ and apply_abs = M_axioms.apply_abs [OF M_axioms_L]
+ and typed_function_abs = M_axioms.typed_function_abs [OF M_axioms_L]
+ and injection_abs = M_axioms.injection_abs [OF M_axioms_L]
+ and surjection_abs = M_axioms.surjection_abs [OF M_axioms_L]
+ and bijection_abs = M_axioms.bijection_abs [OF M_axioms_L]
+ and M_comp_iff = M_axioms.M_comp_iff [OF M_axioms_L]
+ and comp_closed = M_axioms.comp_closed [OF M_axioms_L]
+ and composition_abs = M_axioms.composition_abs [OF M_axioms_L]
+ and restriction_is_function = M_axioms.restriction_is_function [OF M_axioms_L]
+ and restriction_abs = M_axioms.restriction_abs [OF M_axioms_L]
+ and M_restrict_iff = M_axioms.M_restrict_iff [OF M_axioms_L]
+ and restrict_closed = M_axioms.restrict_closed [OF M_axioms_L]
+ and Inter_abs = M_axioms.Inter_abs [OF M_axioms_L]
+ and Inter_closed = M_axioms.Inter_closed [OF M_axioms_L]
+ and Int_closed = M_axioms.Int_closed [OF M_axioms_L]
+ and finite_fun_closed = M_axioms.finite_fun_closed [OF M_axioms_L]
+ and is_funspace_abs = M_axioms.is_funspace_abs [OF M_axioms_L]
+ and succ_fun_eq2 = M_axioms.succ_fun_eq2 [OF M_axioms_L]
+ and funspace_succ = M_axioms.funspace_succ [OF M_axioms_L]
+ and finite_funspace_closed = M_axioms.finite_funspace_closed [OF M_axioms_L]
-ML
-{*
-bind_thm ("is_recfun_equal", axioms_L (thm "M_axioms.is_recfun_equal"));
-bind_thm ("is_recfun_cut", axioms_L (thm "M_axioms.is_recfun_cut"));
-bind_thm ("is_recfun_functional", axioms_L (thm "M_axioms.is_recfun_functional"));
-bind_thm ("is_recfun_relativize", axioms_L (thm "M_axioms.is_recfun_relativize"));
-bind_thm ("is_recfun_restrict", axioms_L (thm "M_axioms.is_recfun_restrict"));
-bind_thm ("univalent_is_recfun", axioms_L (thm "M_axioms.univalent_is_recfun"));
-bind_thm ("exists_is_recfun_indstep", axioms_L (thm "M_axioms.exists_is_recfun_indstep"));
-bind_thm ("wellfounded_exists_is_recfun", axioms_L (thm "M_axioms.wellfounded_exists_is_recfun"));
-bind_thm ("wf_exists_is_recfun", axioms_L (thm "M_axioms.wf_exists_is_recfun"));
-bind_thm ("is_recfun_abs", axioms_L (thm "M_axioms.is_recfun_abs"));
-bind_thm ("irreflexive_abs", axioms_L (thm "M_axioms.irreflexive_abs"));
-bind_thm ("transitive_rel_abs", axioms_L (thm "M_axioms.transitive_rel_abs"));
-bind_thm ("linear_rel_abs", axioms_L (thm "M_axioms.linear_rel_abs"));
-bind_thm ("wellordered_is_trans_on", axioms_L (thm "M_axioms.wellordered_is_trans_on"));
-bind_thm ("wellordered_is_linear", axioms_L (thm "M_axioms.wellordered_is_linear"));
-bind_thm ("wellordered_is_wellfounded_on", axioms_L (thm "M_axioms.wellordered_is_wellfounded_on"));
-bind_thm ("wellfounded_imp_wellfounded_on", axioms_L (thm "M_axioms.wellfounded_imp_wellfounded_on"));
-bind_thm ("wellfounded_on_subset_A", axioms_L (thm "M_axioms.wellfounded_on_subset_A"));
-bind_thm ("wellfounded_on_iff_wellfounded", axioms_L (thm "M_axioms.wellfounded_on_iff_wellfounded"));
-bind_thm ("wellfounded_on_imp_wellfounded", axioms_L (thm "M_axioms.wellfounded_on_imp_wellfounded"));
-bind_thm ("wellfounded_on_field_imp_wellfounded", axioms_L (thm "M_axioms.wellfounded_on_field_imp_wellfounded"));
-bind_thm ("wellfounded_iff_wellfounded_on_field", axioms_L (thm "M_axioms.wellfounded_iff_wellfounded_on_field"));
-bind_thm ("wellfounded_induct", axioms_L (thm "M_axioms.wellfounded_induct"));
-bind_thm ("wellfounded_on_induct", axioms_L (thm "M_axioms.wellfounded_on_induct"));
-bind_thm ("wellfounded_on_induct2", axioms_L (thm "M_axioms.wellfounded_on_induct2"));
-bind_thm ("linear_imp_relativized", axioms_L (thm "M_axioms.linear_imp_relativized"));
-bind_thm ("trans_on_imp_relativized", axioms_L (thm "M_axioms.trans_on_imp_relativized"));
-bind_thm ("wf_on_imp_relativized", axioms_L (thm "M_axioms.wf_on_imp_relativized"));
-bind_thm ("wf_imp_relativized", axioms_L (thm "M_axioms.wf_imp_relativized"));
-bind_thm ("well_ord_imp_relativized", axioms_L (thm "M_axioms.well_ord_imp_relativized"));
-bind_thm ("order_isomorphism_abs", axioms_L (thm "M_axioms.order_isomorphism_abs"));
-bind_thm ("pred_set_abs", axioms_L (thm "M_axioms.pred_set_abs"));
-*}
+lemmas is_recfun_equal = M_axioms.is_recfun_equal [OF M_axioms_L]
+ and is_recfun_cut = M_axioms.is_recfun_cut [OF M_axioms_L]
+ and is_recfun_functional = M_axioms.is_recfun_functional [OF M_axioms_L]
+ and is_recfun_relativize = M_axioms.is_recfun_relativize [OF M_axioms_L]
+ and is_recfun_restrict = M_axioms.is_recfun_restrict [OF M_axioms_L]
+ and univalent_is_recfun = M_axioms.univalent_is_recfun [OF M_axioms_L]
+ and exists_is_recfun_indstep = M_axioms.exists_is_recfun_indstep [OF M_axioms_L]
+ and wellfounded_exists_is_recfun = M_axioms.wellfounded_exists_is_recfun [OF M_axioms_L]
+ and wf_exists_is_recfun = M_axioms.wf_exists_is_recfun [OF M_axioms_L]
+ and is_recfun_abs = M_axioms.is_recfun_abs [OF M_axioms_L]
+ and irreflexive_abs = M_axioms.irreflexive_abs [OF M_axioms_L]
+ and transitive_rel_abs = M_axioms.transitive_rel_abs [OF M_axioms_L]
+ and linear_rel_abs = M_axioms.linear_rel_abs [OF M_axioms_L]
+ and wellordered_is_trans_on = M_axioms.wellordered_is_trans_on [OF M_axioms_L]
+ and wellordered_is_linear = M_axioms.wellordered_is_linear [OF M_axioms_L]
+ and wellordered_is_wellfounded_on = M_axioms.wellordered_is_wellfounded_on [OF M_axioms_L]
+ and wellfounded_imp_wellfounded_on = M_axioms.wellfounded_imp_wellfounded_on [OF M_axioms_L]
+ and wellfounded_on_subset_A = M_axioms.wellfounded_on_subset_A [OF M_axioms_L]
+ and wellfounded_on_iff_wellfounded = M_axioms.wellfounded_on_iff_wellfounded [OF M_axioms_L]
+ and wellfounded_on_imp_wellfounded = M_axioms.wellfounded_on_imp_wellfounded [OF M_axioms_L]
+ and wellfounded_on_field_imp_wellfounded = M_axioms.wellfounded_on_field_imp_wellfounded [OF M_axioms_L]
+ and wellfounded_iff_wellfounded_on_field = M_axioms.wellfounded_iff_wellfounded_on_field [OF M_axioms_L]
+ and wellfounded_induct = M_axioms.wellfounded_induct [OF M_axioms_L]
+ and wellfounded_on_induct = M_axioms.wellfounded_on_induct [OF M_axioms_L]
+ and wellfounded_on_induct2 = M_axioms.wellfounded_on_induct2 [OF M_axioms_L]
+ and linear_imp_relativized = M_axioms.linear_imp_relativized [OF M_axioms_L]
+ and trans_on_imp_relativized = M_axioms.trans_on_imp_relativized [OF M_axioms_L]
+ and wf_on_imp_relativized = M_axioms.wf_on_imp_relativized [OF M_axioms_L]
+ and wf_imp_relativized = M_axioms.wf_imp_relativized [OF M_axioms_L]
+ and well_ord_imp_relativized = M_axioms.well_ord_imp_relativized [OF M_axioms_L]
+ and order_isomorphism_abs = M_axioms.order_isomorphism_abs [OF M_axioms_L]
+ and pred_set_abs = M_axioms.pred_set_abs [OF M_axioms_L]
-ML
-{*
-bind_thm ("pred_closed", axioms_L (thm "M_axioms.pred_closed"));
-bind_thm ("membership_abs", axioms_L (thm "M_axioms.membership_abs"));
-bind_thm ("M_Memrel_iff", axioms_L (thm "M_axioms.M_Memrel_iff"));
-bind_thm ("Memrel_closed", axioms_L (thm "M_axioms.Memrel_closed"));
-bind_thm ("wellordered_iso_predD", axioms_L (thm "M_axioms.wellordered_iso_predD"));
-bind_thm ("wellordered_iso_pred_eq", axioms_L (thm "M_axioms.wellordered_iso_pred_eq"));
-bind_thm ("wellfounded_on_asym", axioms_L (thm "M_axioms.wellfounded_on_asym"));
-bind_thm ("wellordered_asym", axioms_L (thm "M_axioms.wellordered_asym"));
-bind_thm ("ord_iso_pred_imp_lt", axioms_L (thm "M_axioms.ord_iso_pred_imp_lt"));
-bind_thm ("obase_iff", axioms_L (thm "M_axioms.obase_iff"));
-bind_thm ("omap_iff", axioms_L (thm "M_axioms.omap_iff"));
-bind_thm ("omap_unique", axioms_L (thm "M_axioms.omap_unique"));
-bind_thm ("omap_yields_Ord", axioms_L (thm "M_axioms.omap_yields_Ord"));
-bind_thm ("otype_iff", axioms_L (thm "M_axioms.otype_iff"));
-bind_thm ("otype_eq_range", axioms_L (thm "M_axioms.otype_eq_range"));
-bind_thm ("Ord_otype", axioms_L (thm "M_axioms.Ord_otype"));
-bind_thm ("domain_omap", axioms_L (thm "M_axioms.domain_omap"));
-bind_thm ("omap_subset", axioms_L (thm "M_axioms.omap_subset"));
-bind_thm ("omap_funtype", axioms_L (thm "M_axioms.omap_funtype"));
-bind_thm ("wellordered_omap_bij", axioms_L (thm "M_axioms.wellordered_omap_bij"));
-bind_thm ("omap_ord_iso", axioms_L (thm "M_axioms.omap_ord_iso"));
-bind_thm ("Ord_omap_image_pred", axioms_L (thm "M_axioms.Ord_omap_image_pred"));
-bind_thm ("restrict_omap_ord_iso", axioms_L (thm "M_axioms.restrict_omap_ord_iso"));
-bind_thm ("obase_equals", axioms_L (thm "M_axioms.obase_equals"));
-bind_thm ("omap_ord_iso_otype", axioms_L (thm "M_axioms.omap_ord_iso_otype"));
-bind_thm ("obase_exists", axioms_L (thm "M_axioms.obase_exists"));
-bind_thm ("omap_exists", axioms_L (thm "M_axioms.omap_exists"));
-bind_thm ("otype_exists", axioms_L (thm "M_axioms.otype_exists"));
-bind_thm ("omap_ord_iso_otype", axioms_L (thm "M_axioms.omap_ord_iso_otype"));
-bind_thm ("ordertype_exists", axioms_L (thm "M_axioms.ordertype_exists"));
-bind_thm ("relativized_imp_well_ord", axioms_L (thm "M_axioms.relativized_imp_well_ord"));
-bind_thm ("well_ord_abs", axioms_L (thm "M_axioms.well_ord_abs"));
-*}
+lemmas pred_closed = M_axioms.pred_closed [OF M_axioms_L]
+ and membership_abs = M_axioms.membership_abs [OF M_axioms_L]
+ and M_Memrel_iff = M_axioms.M_Memrel_iff [OF M_axioms_L]
+ and Memrel_closed = M_axioms.Memrel_closed [OF M_axioms_L]
+ and wellordered_iso_predD = M_axioms.wellordered_iso_predD [OF M_axioms_L]
+ and wellordered_iso_pred_eq = M_axioms.wellordered_iso_pred_eq [OF M_axioms_L]
+ and wellfounded_on_asym = M_axioms.wellfounded_on_asym [OF M_axioms_L]
+ and wellordered_asym = M_axioms.wellordered_asym [OF M_axioms_L]
+ and ord_iso_pred_imp_lt = M_axioms.ord_iso_pred_imp_lt [OF M_axioms_L]
+ and obase_iff = M_axioms.obase_iff [OF M_axioms_L]
+ and omap_iff = M_axioms.omap_iff [OF M_axioms_L]
+ and omap_unique = M_axioms.omap_unique [OF M_axioms_L]
+ and omap_yields_Ord = M_axioms.omap_yields_Ord [OF M_axioms_L]
+ and otype_iff = M_axioms.otype_iff [OF M_axioms_L]
+ and otype_eq_range = M_axioms.otype_eq_range [OF M_axioms_L]
+ and Ord_otype = M_axioms.Ord_otype [OF M_axioms_L]
+ and domain_omap = M_axioms.domain_omap [OF M_axioms_L]
+ and omap_subset = M_axioms.omap_subset [OF M_axioms_L]
+ and omap_funtype = M_axioms.omap_funtype [OF M_axioms_L]
+ and wellordered_omap_bij = M_axioms.wellordered_omap_bij [OF M_axioms_L]
+ and omap_ord_iso = M_axioms.omap_ord_iso [OF M_axioms_L]
+ and Ord_omap_image_pred = M_axioms.Ord_omap_image_pred [OF M_axioms_L]
+ and restrict_omap_ord_iso = M_axioms.restrict_omap_ord_iso [OF M_axioms_L]
+ and obase_equals = M_axioms.obase_equals [OF M_axioms_L]
+ and omap_ord_iso_otype = M_axioms.omap_ord_iso_otype [OF M_axioms_L]
+ and obase_exists = M_axioms.obase_exists [OF M_axioms_L]
+ and omap_exists = M_axioms.omap_exists [OF M_axioms_L]
+ and otype_exists = M_axioms.otype_exists [OF M_axioms_L]
+ and omap_ord_iso_otype' = M_axioms.omap_ord_iso_otype' [OF M_axioms_L]
+ and ordertype_exists = M_axioms.ordertype_exists [OF M_axioms_L]
+ and relativized_imp_well_ord = M_axioms.relativized_imp_well_ord [OF M_axioms_L]
+ and well_ord_abs = M_axioms.well_ord_abs [OF M_axioms_L]
+
declare cartprod_closed [intro,simp]
declare sum_closed [intro,simp]
@@ -614,7 +598,6 @@
declare Inter_abs [simp]
declare Inter_closed [intro,simp]
declare Int_closed [intro,simp]
-declare finite_fun_closed [rule_format]
declare is_funspace_abs [simp]
declare finite_funspace_closed [intro,simp]