author | haftmann |
Sun, 27 Jul 2025 17:52:06 +0200 | |
changeset 82909 | e4fae2227594 |
parent 76217 | 8655344f1cf6 |
permissions | -rw-r--r-- |
13505 | 1 |
(* Title: ZF/Constructible/Reflection.thy |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
3 |
*) |
|
4 |
||
60770 | 5 |
section \<open>The Reflection Theorem\<close> |
13223 | 6 |
|
16417 | 7 |
theory Reflection imports Normal begin |
13223 | 8 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
9 |
lemma all_iff_not_ex_not: "(\<forall>x. P(x)) \<longleftrightarrow> (\<not> (\<exists>x. \<not> P(x)))" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
10 |
by blast |
13223 | 11 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
12 |
lemma ball_iff_not_bex_not: "(\<forall>x\<in>A. P(x)) \<longleftrightarrow> (\<not> (\<exists>x\<in>A. \<not> P(x)))" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
13 |
by blast |
13223 | 14 |
|
60770 | 15 |
text\<open>From the notes of A. S. Kechris, page 6, and from |
13223 | 16 |
Andrzej Mostowski, \emph{Constructible Sets with Applications}, |
60770 | 17 |
North-Holland, 1969, page 23.\<close> |
13223 | 18 |
|
19 |
||
60770 | 20 |
subsection\<open>Basic Definitions\<close> |
13223 | 21 |
|
61798 | 22 |
text\<open>First part: the cumulative hierarchy defining the class \<open>M\<close>. |
23 |
To avoid handling multiple arguments, we assume that \<open>Mset(l)\<close> is |
|
24 |
closed under ordered pairing provided \<open>l\<close> is limit. Possibly this |
|
69593 | 25 |
could be avoided: the induction hypothesis \<^term>\<open>Cl_reflects\<close> |
61798 | 26 |
(in locale \<open>ex_reflection\<close>) could be weakened to |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
27 |
\<^term>\<open>\<forall>y\<in>Mset(a). \<forall>z\<in>Mset(a). P(\<langle>y,z\<rangle>) \<longleftrightarrow> Q(a,\<langle>y,z\<rangle>)\<close>, removing most |
69593 | 28 |
uses of \<^term>\<open>Pair_in_Mset\<close>. But there isn't much point in doing so, since |
61798 | 29 |
ultimately the \<open>ex_reflection\<close> proof is packaged up using the |
30 |
predicate \<open>Reflects\<close>. |
|
60770 | 31 |
\<close> |
13428 | 32 |
locale reflection = |
13223 | 33 |
fixes Mset and M and Reflects |
34 |
assumes Mset_mono_le : "mono_le_subset(Mset)" |
|
35 |
and Mset_cont : "cont_Ord(Mset)" |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
36 |
and Pair_in_Mset : "\<lbrakk>x \<in> Mset(a); y \<in> Mset(a); Limit(a)\<rbrakk> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
37 |
\<Longrightarrow> \<langle>x,y\<rangle> \<in> Mset(a)" |
76214 | 38 |
defines "M(x) \<equiv> \<exists>a. Ord(a) \<and> x \<in> Mset(a)" |
39 |
and "Reflects(Cl,P,Q) \<equiv> Closed_Unbounded(Cl) \<and> |
|
46823 | 40 |
(\<forall>a. Cl(a) \<longrightarrow> (\<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)))" |
69593 | 41 |
fixes F0 \<comment> \<open>ordinal for a specific value \<^term>\<open>y\<close>\<close> |
42 |
fixes FF \<comment> \<open>sup over the whole level, \<^term>\<open>y\<in>Mset(a)\<close>\<close> |
|
43 |
fixes ClEx \<comment> \<open>Reflecting ordinals for the formula \<^term>\<open>\<exists>z. P\<close>\<close> |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
44 |
defines "F0(P,y) \<equiv> \<mu> b. (\<exists>z. M(z) \<and> P(\<langle>y,z\<rangle>)) \<longrightarrow> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
45 |
(\<exists>z\<in>Mset(b). P(\<langle>y,z\<rangle>))" |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
46 |
and "FF(P) \<equiv> \<lambda>a. \<Union>y\<in>Mset(a). F0(P,y)" |
76214 | 47 |
and "ClEx(P,a) \<equiv> Limit(a) \<and> normalize(FF(P),a) = a" |
13223 | 48 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
49 |
begin |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
50 |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
51 |
lemma Mset_mono: "i\<le>j \<Longrightarrow> Mset(i) \<subseteq> Mset(j)" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
52 |
using Mset_mono_le by (simp add: mono_le_subset_def leI) |
13223 | 53 |
|
61798 | 54 |
text\<open>Awkward: we need a version of \<open>ClEx_def\<close> as an equality |
60770 | 55 |
at the level of classes, which do not really exist\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
56 |
lemma ClEx_eq: |
76214 | 57 |
"ClEx(P) \<equiv> \<lambda>a. Limit(a) \<and> normalize(FF(P),a) = a" |
46823 | 58 |
by (simp add: ClEx_def [symmetric]) |
13434 | 59 |
|
60 |
||
60770 | 61 |
subsection\<open>Easy Cases of the Reflection Theorem\<close> |
13223 | 62 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
63 |
theorem Triv_reflection [intro]: |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
64 |
"Reflects(Ord, P, \<lambda>a x. P(x))" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
65 |
by (simp add: Reflects_def) |
13223 | 66 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
67 |
theorem Not_reflection [intro]: |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
68 |
"Reflects(Cl,P,Q) \<Longrightarrow> Reflects(Cl, \<lambda>x. \<not>P(x), \<lambda>a x. \<not>Q(a,x))" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
69 |
by (simp add: Reflects_def) |
13223 | 70 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
71 |
theorem And_reflection [intro]: |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
72 |
"\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk> |
76214 | 73 |
\<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> C'(a), \<lambda>x. P(x) \<and> P'(x), |
74 |
\<lambda>a x. Q(a,x) \<and> Q'(a,x))" |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
75 |
by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
13223 | 76 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
77 |
theorem Or_reflection [intro]: |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
78 |
"\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk> |
76214 | 79 |
\<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> C'(a), \<lambda>x. P(x) \<or> P'(x), |
80 |
\<lambda>a x. Q(a,x) \<or> Q'(a,x))" |
|
13223 | 81 |
by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
82 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
83 |
theorem Imp_reflection [intro]: |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
84 |
"\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk> |
76214 | 85 |
\<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> C'(a), |
46823 | 86 |
\<lambda>x. P(x) \<longrightarrow> P'(x), |
87 |
\<lambda>a x. Q(a,x) \<longrightarrow> Q'(a,x))" |
|
13223 | 88 |
by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
89 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
90 |
theorem Iff_reflection [intro]: |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
91 |
"\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk> |
76214 | 92 |
\<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> C'(a), |
46823 | 93 |
\<lambda>x. P(x) \<longleftrightarrow> P'(x), |
94 |
\<lambda>a x. Q(a,x) \<longleftrightarrow> Q'(a,x))" |
|
95 |
by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
|
13223 | 96 |
|
60770 | 97 |
subsection\<open>Reflection for Existential Quantifiers\<close> |
13223 | 98 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
99 |
lemma F0_works: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
100 |
"\<lbrakk>y\<in>Mset(a); Ord(a); M(z); P(\<langle>y,z\<rangle>)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(F0(P,y)). P(\<langle>y,z\<rangle>)" |
76214 | 101 |
unfolding F0_def M_def |
102 |
apply clarify |
|
103 |
apply (rule LeastI2) |
|
104 |
apply (blast intro: Mset_mono [THEN subsetD]) |
|
105 |
apply (blast intro: lt_Ord2, blast) |
|
106 |
done |
|
13223 | 107 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
108 |
lemma Ord_F0 [intro,simp]: "Ord(F0(P,y))" |
76214 | 109 |
by (simp add: F0_def) |
13223 | 110 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
111 |
lemma Ord_FF [intro,simp]: "Ord(FF(P,y))" |
76214 | 112 |
by (simp add: FF_def) |
13223 | 113 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
114 |
lemma cont_Ord_FF: "cont_Ord(FF(P))" |
76214 | 115 |
using Mset_cont by (simp add: cont_Ord_def FF_def, blast) |
13223 | 116 |
|
69593 | 117 |
text\<open>Recall that \<^term>\<open>F0\<close> depends upon \<^term>\<open>y\<in>Mset(a)\<close>, |
118 |
while \<^term>\<open>FF\<close> depends only upon \<^term>\<open>a\<close>.\<close> |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
119 |
lemma FF_works: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
120 |
"\<lbrakk>M(z); y\<in>Mset(a); P(\<langle>y,z\<rangle>); Ord(a)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(FF(P,a)). P(\<langle>y,z\<rangle>)" |
76214 | 121 |
apply (simp add: FF_def) |
122 |
apply (simp_all add: cont_Ord_Union [of concl: Mset] |
|
123 |
Mset_cont Mset_mono_le not_emptyI) |
|
124 |
apply (blast intro: F0_works) |
|
125 |
done |
|
13223 | 126 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
127 |
lemma FFN_works: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
128 |
"\<lbrakk>M(z); y\<in>Mset(a); P(\<langle>y,z\<rangle>); Ord(a)\<rbrakk> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
129 |
\<Longrightarrow> \<exists>z\<in>Mset(normalize(FF(P),a)). P(\<langle>y,z\<rangle>)" |
46823 | 130 |
apply (drule FF_works [of concl: P], assumption+) |
13223 | 131 |
apply (blast intro: cont_Ord_FF le_normalize [THEN Mset_mono, THEN subsetD]) |
132 |
done |
|
133 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
134 |
end |
13223 | 135 |
|
60770 | 136 |
text\<open>Locale for the induction hypothesis\<close> |
13223 | 137 |
|
13428 | 138 |
locale ex_reflection = reflection + |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61980
diff
changeset
|
139 |
fixes P \<comment> \<open>the original formula\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61980
diff
changeset
|
140 |
fixes Q \<comment> \<open>the reflected formula\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61980
diff
changeset
|
141 |
fixes Cl \<comment> \<open>the class of reflecting ordinals\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
142 |
assumes Cl_reflects: "\<lbrakk>Cl(a); Ord(a)\<rbrakk> \<Longrightarrow> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
143 |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
144 |
begin |
13223 | 145 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
146 |
lemma ClEx_downward: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
147 |
"\<lbrakk>M(z); y\<in>Mset(a); P(\<langle>y,z\<rangle>); Cl(a); ClEx(P,a)\<rbrakk> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
148 |
\<Longrightarrow> \<exists>z\<in>Mset(a). Q(a,\<langle>y,z\<rangle>)" |
46823 | 149 |
apply (simp add: ClEx_def, clarify) |
150 |
apply (frule Limit_is_Ord) |
|
151 |
apply (frule FFN_works [of concl: P], assumption+) |
|
152 |
apply (drule Cl_reflects, assumption+) |
|
13223 | 153 |
apply (auto simp add: Limit_is_Ord Pair_in_Mset) |
154 |
done |
|
155 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
156 |
lemma ClEx_upward: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
157 |
"\<lbrakk>z\<in>Mset(a); y\<in>Mset(a); Q(a,\<langle>y,z\<rangle>); Cl(a); ClEx(P,a)\<rbrakk> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
158 |
\<Longrightarrow> \<exists>z. M(z) \<and> P(\<langle>y,z\<rangle>)" |
13223 | 159 |
apply (simp add: ClEx_def M_def) |
160 |
apply (blast dest: Cl_reflects |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
23464
diff
changeset
|
161 |
intro: Limit_is_Ord Pair_in_Mset) |
13223 | 162 |
done |
163 |
||
61798 | 164 |
text\<open>Class \<open>ClEx\<close> indeed consists of reflecting ordinals...\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
165 |
lemma ZF_ClEx_iff: |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
166 |
"\<lbrakk>y\<in>Mset(a); Cl(a); ClEx(P,a)\<rbrakk> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
167 |
\<Longrightarrow> (\<exists>z. M(z) \<and> P(\<langle>y,z\<rangle>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,\<langle>y,z\<rangle>))" |
46823 | 168 |
by (blast intro: dest: ClEx_downward ClEx_upward) |
13223 | 169 |
|
60770 | 170 |
text\<open>...and it is closed and unbounded\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
171 |
lemma ZF_Closed_Unbounded_ClEx: |
13223 | 172 |
"Closed_Unbounded(ClEx(P))" |
13434 | 173 |
apply (simp add: ClEx_eq) |
13223 | 174 |
apply (fast intro: Closed_Unbounded_Int Normal_imp_fp_Closed_Unbounded |
175 |
Closed_Unbounded_Limit Normal_normalize) |
|
176 |
done |
|
177 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
178 |
end |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
179 |
|
61798 | 180 |
text\<open>The same two theorems, exported to locale \<open>reflection\<close>.\<close> |
13223 | 181 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
182 |
context reflection |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
183 |
begin |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
184 |
|
61798 | 185 |
text\<open>Class \<open>ClEx\<close> indeed consists of reflecting ordinals...\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
186 |
lemma ClEx_iff: |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
187 |
"\<lbrakk>y\<in>Mset(a); Cl(a); ClEx(P,a); |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
188 |
\<And>a. \<lbrakk>Cl(a); Ord(a)\<rbrakk> \<Longrightarrow> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)\<rbrakk> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
189 |
\<Longrightarrow> (\<exists>z. M(z) \<and> P(\<langle>y,z\<rangle>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,\<langle>y,z\<rangle>))" |
76217 | 190 |
unfolding ClEx_def FF_def F0_def M_def |
13428 | 191 |
apply (rule ex_reflection.ZF_ClEx_iff |
192 |
[OF ex_reflection.intro, OF reflection.intro ex_reflection_axioms.intro, |
|
193 |
of Mset Cl]) |
|
194 |
apply (simp_all add: Mset_mono_le Mset_cont Pair_in_Mset) |
|
13223 | 195 |
done |
196 |
||
13434 | 197 |
(*Alternative proof, less unfolding: |
198 |
apply (rule Reflection.ZF_ClEx_iff [of Mset _ _ Cl, folded M_def]) |
|
199 |
apply (fold ClEx_def FF_def F0_def) |
|
200 |
apply (rule ex_reflection.intro, assumption) |
|
201 |
apply (simp add: ex_reflection_axioms.intro, assumption+) |
|
202 |
*) |
|
203 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
204 |
lemma Closed_Unbounded_ClEx: |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
205 |
"(\<And>a. \<lbrakk>Cl(a); Ord(a)\<rbrakk> \<Longrightarrow> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)) |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
206 |
\<Longrightarrow> Closed_Unbounded(ClEx(P))" |
76217 | 207 |
unfolding ClEx_eq FF_def F0_def M_def |
21232 | 208 |
apply (rule ex_reflection.ZF_Closed_Unbounded_ClEx [of Mset _ _ Cl]) |
23464 | 209 |
apply (rule ex_reflection.intro, rule reflection_axioms) |
13434 | 210 |
apply (blast intro: ex_reflection_axioms.intro) |
13223 | 211 |
done |
212 |
||
60770 | 213 |
subsection\<open>Packaging the Quantifier Reflection Rules\<close> |
13292 | 214 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
215 |
lemma Ex_reflection_0: |
46823 | 216 |
"Reflects(Cl,P0,Q0) |
76214 | 217 |
\<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(P0,a), |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
218 |
\<lambda>x. \<exists>z. M(z) \<and> P0(\<langle>x,z\<rangle>), |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
219 |
\<lambda>a x. \<exists>z\<in>Mset(a). Q0(a,\<langle>x,z\<rangle>))" |
46823 | 220 |
apply (simp add: Reflects_def) |
13223 | 221 |
apply (intro conjI Closed_Unbounded_Int) |
46823 | 222 |
apply blast |
223 |
apply (rule Closed_Unbounded_ClEx [of Cl P0 Q0], blast, clarify) |
|
224 |
apply (rule_tac Cl=Cl in ClEx_iff, assumption+, blast) |
|
13223 | 225 |
done |
226 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
227 |
lemma All_reflection_0: |
46823 | 228 |
"Reflects(Cl,P0,Q0) |
76214 | 229 |
\<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(\<lambda>x.\<not>P0(x), a), |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
230 |
\<lambda>x. \<forall>z. M(z) \<longrightarrow> P0(\<langle>x,z\<rangle>), |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
231 |
\<lambda>a x. \<forall>z\<in>Mset(a). Q0(a,\<langle>x,z\<rangle>))" |
46823 | 232 |
apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not) |
233 |
apply (rule Not_reflection, drule Not_reflection, simp) |
|
13223 | 234 |
apply (erule Ex_reflection_0) |
235 |
done |
|
236 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
237 |
theorem Ex_reflection [intro]: |
46823 | 238 |
"Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
76214 | 239 |
\<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(\<lambda>x. P(fst(x),snd(x)), a), |
240 |
\<lambda>x. \<exists>z. M(z) \<and> P(x,z), |
|
13223 | 241 |
\<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))" |
46823 | 242 |
by (rule Ex_reflection_0 [of _ " \<lambda>x. P(fst(x),snd(x))" |
13223 | 243 |
"\<lambda>a x. Q(a,fst(x),snd(x))", simplified]) |
244 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
245 |
theorem All_reflection [intro]: |
13223 | 246 |
"Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
76214 | 247 |
\<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(\<lambda>x. \<not>P(fst(x),snd(x)), a), |
46823 | 248 |
\<lambda>x. \<forall>z. M(z) \<longrightarrow> P(x,z), |
249 |
\<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))" |
|
250 |
by (rule All_reflection_0 [of _ "\<lambda>x. P(fst(x),snd(x))" |
|
13223 | 251 |
"\<lambda>a x. Q(a,fst(x),snd(x))", simplified]) |
252 |
||
60770 | 253 |
text\<open>And again, this time using class-bounded quantifiers\<close> |
13292 | 254 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
255 |
theorem Rex_reflection [intro]: |
46823 | 256 |
"Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
76214 | 257 |
\<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(\<lambda>x. P(fst(x),snd(x)), a), |
46823 | 258 |
\<lambda>x. \<exists>z[M]. P(x,z), |
13292 | 259 |
\<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))" |
46823 | 260 |
by (unfold rex_def, blast) |
13292 | 261 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
262 |
theorem Rall_reflection [intro]: |
13292 | 263 |
"Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
76214 | 264 |
\<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(\<lambda>x. \<not>P(fst(x),snd(x)), a), |
46823 | 265 |
\<lambda>x. \<forall>z[M]. P(x,z), |
266 |
\<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))" |
|
267 |
by (unfold rall_def, blast) |
|
13292 | 268 |
|
269 |
||
60770 | 270 |
text\<open>No point considering bounded quantifiers, where reflection is trivial.\<close> |
13223 | 271 |
|
272 |
||
60770 | 273 |
subsection\<open>Simple Examples of Reflection\<close> |
13223 | 274 |
|
60770 | 275 |
text\<open>Example 1: reflecting a simple formula. The reflecting class is first |
61798 | 276 |
given as the variable \<open>?Cl\<close> and later retrieved from the final |
60770 | 277 |
proof state.\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
278 |
schematic_goal |
13223 | 279 |
"Reflects(?Cl, |
76214 | 280 |
\<lambda>x. \<exists>y. M(y) \<and> x \<in> y, |
13223 | 281 |
\<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)" |
282 |
by fast |
|
283 |
||
60770 | 284 |
text\<open>Problem here: there needs to be a conjunction (class intersection) |
69593 | 285 |
in the class of reflecting ordinals. The \<^term>\<open>Ord(a)\<close> is redundant, |
60770 | 286 |
though harmless.\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
287 |
lemma |
76214 | 288 |
"Reflects(\<lambda>a. Ord(a) \<and> ClEx(\<lambda>x. fst(x) \<in> snd(x), a), |
289 |
\<lambda>x. \<exists>y. M(y) \<and> x \<in> y, |
|
46823 | 290 |
\<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)" |
13223 | 291 |
by fast |
292 |
||
293 |
||
60770 | 294 |
text\<open>Example 2\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
295 |
schematic_goal |
13223 | 296 |
"Reflects(?Cl, |
76214 | 297 |
\<lambda>x. \<exists>y. M(y) \<and> (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y), |
46823 | 298 |
\<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)" |
13223 | 299 |
by fast |
300 |
||
60770 | 301 |
text\<open>Example 2'. We give the reflecting class explicitly.\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
302 |
lemma |
13223 | 303 |
"Reflects |
76214 | 304 |
(\<lambda>a. (Ord(a) \<and> |
305 |
ClEx(\<lambda>x. \<not> (snd(x) \<subseteq> fst(fst(x)) \<longrightarrow> snd(x) \<in> snd(fst(x))), a)) \<and> |
|
46823 | 306 |
ClEx(\<lambda>x. \<forall>z. M(z) \<longrightarrow> z \<subseteq> fst(x) \<longrightarrow> z \<in> snd(x), a), |
76214 | 307 |
\<lambda>x. \<exists>y. M(y) \<and> (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y), |
46823 | 308 |
\<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)" |
13223 | 309 |
by fast |
310 |
||
60770 | 311 |
text\<open>Example 2''. We expand the subset relation.\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
312 |
schematic_goal |
13223 | 313 |
"Reflects(?Cl, |
76214 | 314 |
\<lambda>x. \<exists>y. M(y) \<and> (\<forall>z. M(z) \<longrightarrow> (\<forall>w. M(w) \<longrightarrow> w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y), |
46823 | 315 |
\<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). (\<forall>w\<in>Mset(a). w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y)" |
13223 | 316 |
by fast |
317 |
||
60770 | 318 |
text\<open>Example 2'''. Single-step version, to reveal the reflecting class.\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
319 |
schematic_goal |
13223 | 320 |
"Reflects(?Cl, |
76214 | 321 |
\<lambda>x. \<exists>y. M(y) \<and> (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y), |
46823 | 322 |
\<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)" |
323 |
apply (rule Ex_reflection) |
|
60770 | 324 |
txt\<open> |
13223 | 325 |
@{goals[display,indent=0,margin=60]} |
60770 | 326 |
\<close> |
46823 | 327 |
apply (rule All_reflection) |
60770 | 328 |
txt\<open> |
13223 | 329 |
@{goals[display,indent=0,margin=60]} |
60770 | 330 |
\<close> |
46823 | 331 |
apply (rule Triv_reflection) |
60770 | 332 |
txt\<open> |
13223 | 333 |
@{goals[display,indent=0,margin=60]} |
60770 | 334 |
\<close> |
13223 | 335 |
done |
336 |
||
60770 | 337 |
text\<open>Example 3. Warning: the following examples make sense only |
69593 | 338 |
if \<^term>\<open>P\<close> is quantifier-free, since it is not being relativized.\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
339 |
schematic_goal |
13223 | 340 |
"Reflects(?Cl, |
76214 | 341 |
\<lambda>x. \<exists>y. M(y) \<and> (\<forall>z. M(z) \<longrightarrow> z \<in> y \<longleftrightarrow> z \<in> x \<and> P(z)), |
342 |
\<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<in> y \<longleftrightarrow> z \<in> x \<and> P(z))" |
|
13223 | 343 |
by fast |
344 |
||
60770 | 345 |
text\<open>Example 3'\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
346 |
schematic_goal |
13223 | 347 |
"Reflects(?Cl, |
76214 | 348 |
\<lambda>x. \<exists>y. M(y) \<and> y = Collect(x,P), |
58860 | 349 |
\<lambda>a x. \<exists>y\<in>Mset(a). y = Collect(x,P))" |
13223 | 350 |
by fast |
351 |
||
60770 | 352 |
text\<open>Example 3''\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
353 |
schematic_goal |
13223 | 354 |
"Reflects(?Cl, |
76214 | 355 |
\<lambda>x. \<exists>y. M(y) \<and> y = Replace(x,P), |
58860 | 356 |
\<lambda>a x. \<exists>y\<in>Mset(a). y = Replace(x,P))" |
13223 | 357 |
by fast |
358 |
||
61798 | 359 |
text\<open>Example 4: Axiom of Choice. Possibly wrong, since \<open>\<Pi>\<close> needs |
60770 | 360 |
to be relativized.\<close> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
361 |
schematic_goal |
13223 | 362 |
"Reflects(?Cl, |
76214 | 363 |
\<lambda>A. 0\<notin>A \<longrightarrow> (\<exists>f. M(f) \<and> f \<in> (\<Prod>X \<in> A. X)), |
61980 | 364 |
\<lambda>a A. 0\<notin>A \<longrightarrow> (\<exists>f\<in>Mset(a). f \<in> (\<Prod>X \<in> A. X)))" |
13223 | 365 |
by fast |
366 |
||
367 |
end |
|
368 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
369 |
end |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71417
diff
changeset
|
370 |