31 \<close> |
31 \<close> |
32 locale reflection = |
32 locale reflection = |
33 fixes Mset and M and Reflects |
33 fixes Mset and M and Reflects |
34 assumes Mset_mono_le : "mono_le_subset(Mset)" |
34 assumes Mset_mono_le : "mono_le_subset(Mset)" |
35 and Mset_cont : "cont_Ord(Mset)" |
35 and Mset_cont : "cont_Ord(Mset)" |
36 and Pair_in_Mset : "[| x \<in> Mset(a); y \<in> Mset(a); Limit(a) |] |
36 and Pair_in_Mset : "\<lbrakk>x \<in> Mset(a); y \<in> Mset(a); Limit(a)\<rbrakk> |
37 ==> <x,y> \<in> Mset(a)" |
37 \<Longrightarrow> <x,y> \<in> Mset(a)" |
38 defines "M(x) == \<exists>a. Ord(a) & x \<in> Mset(a)" |
38 defines "M(x) \<equiv> \<exists>a. Ord(a) & x \<in> Mset(a)" |
39 and "Reflects(Cl,P,Q) == Closed_Unbounded(Cl) & |
39 and "Reflects(Cl,P,Q) \<equiv> Closed_Unbounded(Cl) & |
40 (\<forall>a. Cl(a) \<longrightarrow> (\<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)))" |
40 (\<forall>a. Cl(a) \<longrightarrow> (\<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)))" |
41 fixes F0 \<comment> \<open>ordinal for a specific value \<^term>\<open>y\<close>\<close> |
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> |
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> |
43 fixes ClEx \<comment> \<open>Reflecting ordinals for the formula \<^term>\<open>\<exists>z. P\<close>\<close> |
44 defines "F0(P,y) == \<mu> b. (\<exists>z. M(z) & P(<y,z>)) \<longrightarrow> |
44 defines "F0(P,y) \<equiv> \<mu> b. (\<exists>z. M(z) & P(<y,z>)) \<longrightarrow> |
45 (\<exists>z\<in>Mset(b). P(<y,z>))" |
45 (\<exists>z\<in>Mset(b). P(<y,z>))" |
46 and "FF(P) == \<lambda>a. \<Union>y\<in>Mset(a). F0(P,y)" |
46 and "FF(P) \<equiv> \<lambda>a. \<Union>y\<in>Mset(a). F0(P,y)" |
47 and "ClEx(P,a) == Limit(a) & normalize(FF(P),a) = a" |
47 and "ClEx(P,a) \<equiv> Limit(a) & normalize(FF(P),a) = a" |
48 |
48 |
49 lemma (in reflection) Mset_mono: "i\<le>j ==> Mset(i) \<subseteq> Mset(j)" |
49 begin |
50 apply (insert Mset_mono_le) |
50 |
51 apply (simp add: mono_le_subset_def leI) |
51 lemma Mset_mono: "i\<le>j \<Longrightarrow> Mset(i) \<subseteq> Mset(j)" |
52 done |
52 using Mset_mono_le by (simp add: mono_le_subset_def leI) |
53 |
53 |
54 text\<open>Awkward: we need a version of \<open>ClEx_def\<close> as an equality |
54 text\<open>Awkward: we need a version of \<open>ClEx_def\<close> as an equality |
55 at the level of classes, which do not really exist\<close> |
55 at the level of classes, which do not really exist\<close> |
56 lemma (in reflection) ClEx_eq: |
56 lemma ClEx_eq: |
57 "ClEx(P) == \<lambda>a. Limit(a) & normalize(FF(P),a) = a" |
57 "ClEx(P) \<equiv> \<lambda>a. Limit(a) & normalize(FF(P),a) = a" |
58 by (simp add: ClEx_def [symmetric]) |
58 by (simp add: ClEx_def [symmetric]) |
59 |
59 |
60 |
60 |
61 subsection\<open>Easy Cases of the Reflection Theorem\<close> |
61 subsection\<open>Easy Cases of the Reflection Theorem\<close> |
62 |
62 |
63 theorem (in reflection) Triv_reflection [intro]: |
63 theorem Triv_reflection [intro]: |
64 "Reflects(Ord, P, \<lambda>a x. P(x))" |
64 "Reflects(Ord, P, \<lambda>a x. P(x))" |
65 by (simp add: Reflects_def) |
65 by (simp add: Reflects_def) |
66 |
66 |
67 theorem (in reflection) Not_reflection [intro]: |
67 theorem Not_reflection [intro]: |
68 "Reflects(Cl,P,Q) ==> Reflects(Cl, \<lambda>x. ~P(x), \<lambda>a x. ~Q(a,x))" |
68 "Reflects(Cl,P,Q) \<Longrightarrow> Reflects(Cl, \<lambda>x. \<not>P(x), \<lambda>a x. \<not>Q(a,x))" |
69 by (simp add: Reflects_def) |
69 by (simp add: Reflects_def) |
70 |
70 |
71 theorem (in reflection) And_reflection [intro]: |
71 theorem And_reflection [intro]: |
72 "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] |
72 "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk> |
73 ==> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) & P'(x), |
73 \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) & P'(x), |
74 \<lambda>a x. Q(a,x) & Q'(a,x))" |
74 \<lambda>a x. Q(a,x) & Q'(a,x))" |
75 by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
75 by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
76 |
76 |
77 theorem (in reflection) Or_reflection [intro]: |
77 theorem Or_reflection [intro]: |
78 "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] |
78 "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk> |
79 ==> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) | P'(x), |
79 \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) | P'(x), |
80 \<lambda>a x. Q(a,x) | Q'(a,x))" |
80 \<lambda>a x. Q(a,x) | Q'(a,x))" |
81 by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
81 by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
82 |
82 |
83 theorem (in reflection) Imp_reflection [intro]: |
83 theorem Imp_reflection [intro]: |
84 "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] |
84 "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk> |
85 ==> Reflects(\<lambda>a. Cl(a) & C'(a), |
85 \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a), |
86 \<lambda>x. P(x) \<longrightarrow> P'(x), |
86 \<lambda>x. P(x) \<longrightarrow> P'(x), |
87 \<lambda>a x. Q(a,x) \<longrightarrow> Q'(a,x))" |
87 \<lambda>a x. Q(a,x) \<longrightarrow> Q'(a,x))" |
88 by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
88 by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
89 |
89 |
90 theorem (in reflection) Iff_reflection [intro]: |
90 theorem Iff_reflection [intro]: |
91 "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] |
91 "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk> |
92 ==> Reflects(\<lambda>a. Cl(a) & C'(a), |
92 \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a), |
93 \<lambda>x. P(x) \<longleftrightarrow> P'(x), |
93 \<lambda>x. P(x) \<longleftrightarrow> P'(x), |
94 \<lambda>a x. Q(a,x) \<longleftrightarrow> Q'(a,x))" |
94 \<lambda>a x. Q(a,x) \<longleftrightarrow> Q'(a,x))" |
95 by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
95 by (simp add: Reflects_def Closed_Unbounded_Int, blast) |
96 |
96 |
97 subsection\<open>Reflection for Existential Quantifiers\<close> |
97 subsection\<open>Reflection for Existential Quantifiers\<close> |
98 |
98 |
99 lemma (in reflection) F0_works: |
99 lemma F0_works: |
100 "[| y\<in>Mset(a); Ord(a); M(z); P(<y,z>) |] ==> \<exists>z\<in>Mset(F0(P,y)). P(<y,z>)" |
100 "\<lbrakk>y\<in>Mset(a); Ord(a); M(z); P(<y,z>)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(F0(P,y)). P(<y,z>)" |
101 apply (unfold F0_def M_def, clarify) |
101 apply (unfold F0_def M_def, clarify) |
102 apply (rule LeastI2) |
102 apply (rule LeastI2) |
103 apply (blast intro: Mset_mono [THEN subsetD]) |
103 apply (blast intro: Mset_mono [THEN subsetD]) |
104 apply (blast intro: lt_Ord2, blast) |
104 apply (blast intro: lt_Ord2, blast) |
105 done |
105 done |
106 |
106 |
107 lemma (in reflection) Ord_F0 [intro,simp]: "Ord(F0(P,y))" |
107 lemma Ord_F0 [intro,simp]: "Ord(F0(P,y))" |
108 by (simp add: F0_def) |
108 by (simp add: F0_def) |
109 |
109 |
110 lemma (in reflection) Ord_FF [intro,simp]: "Ord(FF(P,y))" |
110 lemma Ord_FF [intro,simp]: "Ord(FF(P,y))" |
111 by (simp add: FF_def) |
111 by (simp add: FF_def) |
112 |
112 |
113 lemma (in reflection) cont_Ord_FF: "cont_Ord(FF(P))" |
113 lemma cont_Ord_FF: "cont_Ord(FF(P))" |
114 apply (insert Mset_cont) |
114 apply (insert Mset_cont) |
115 apply (simp add: cont_Ord_def FF_def, blast) |
115 apply (simp add: cont_Ord_def FF_def, blast) |
116 done |
116 done |
117 |
117 |
118 text\<open>Recall that \<^term>\<open>F0\<close> depends upon \<^term>\<open>y\<in>Mset(a)\<close>, |
118 text\<open>Recall that \<^term>\<open>F0\<close> depends upon \<^term>\<open>y\<in>Mset(a)\<close>, |
119 while \<^term>\<open>FF\<close> depends only upon \<^term>\<open>a\<close>.\<close> |
119 while \<^term>\<open>FF\<close> depends only upon \<^term>\<open>a\<close>.\<close> |
120 lemma (in reflection) FF_works: |
120 lemma FF_works: |
121 "[| M(z); y\<in>Mset(a); P(<y,z>); Ord(a) |] ==> \<exists>z\<in>Mset(FF(P,a)). P(<y,z>)" |
121 "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Ord(a)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(FF(P,a)). P(<y,z>)" |
122 apply (simp add: FF_def) |
122 apply (simp add: FF_def) |
123 apply (simp_all add: cont_Ord_Union [of concl: Mset] |
123 apply (simp_all add: cont_Ord_Union [of concl: Mset] |
124 Mset_cont Mset_mono_le not_emptyI) |
124 Mset_cont Mset_mono_le not_emptyI) |
125 apply (blast intro: F0_works) |
125 apply (blast intro: F0_works) |
126 done |
126 done |
127 |
127 |
128 lemma (in reflection) FFN_works: |
128 lemma FFN_works: |
129 "[| M(z); y\<in>Mset(a); P(<y,z>); Ord(a) |] |
129 "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Ord(a)\<rbrakk> |
130 ==> \<exists>z\<in>Mset(normalize(FF(P),a)). P(<y,z>)" |
130 \<Longrightarrow> \<exists>z\<in>Mset(normalize(FF(P),a)). P(<y,z>)" |
131 apply (drule FF_works [of concl: P], assumption+) |
131 apply (drule FF_works [of concl: P], assumption+) |
132 apply (blast intro: cont_Ord_FF le_normalize [THEN Mset_mono, THEN subsetD]) |
132 apply (blast intro: cont_Ord_FF le_normalize [THEN Mset_mono, THEN subsetD]) |
133 done |
133 done |
134 |
134 |
|
135 end |
135 |
136 |
136 text\<open>Locale for the induction hypothesis\<close> |
137 text\<open>Locale for the induction hypothesis\<close> |
137 |
138 |
138 locale ex_reflection = reflection + |
139 locale ex_reflection = reflection + |
139 fixes P \<comment> \<open>the original formula\<close> |
140 fixes P \<comment> \<open>the original formula\<close> |
140 fixes Q \<comment> \<open>the reflected formula\<close> |
141 fixes Q \<comment> \<open>the reflected formula\<close> |
141 fixes Cl \<comment> \<open>the class of reflecting ordinals\<close> |
142 fixes Cl \<comment> \<open>the class of reflecting ordinals\<close> |
142 assumes Cl_reflects: "[| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)" |
143 assumes Cl_reflects: "\<lbrakk>Cl(a); Ord(a)\<rbrakk> \<Longrightarrow> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)" |
143 |
144 |
144 lemma (in ex_reflection) ClEx_downward: |
145 begin |
145 "[| M(z); y\<in>Mset(a); P(<y,z>); Cl(a); ClEx(P,a) |] |
146 |
146 ==> \<exists>z\<in>Mset(a). Q(a,<y,z>)" |
147 lemma ClEx_downward: |
|
148 "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Cl(a); ClEx(P,a)\<rbrakk> |
|
149 \<Longrightarrow> \<exists>z\<in>Mset(a). Q(a,<y,z>)" |
147 apply (simp add: ClEx_def, clarify) |
150 apply (simp add: ClEx_def, clarify) |
148 apply (frule Limit_is_Ord) |
151 apply (frule Limit_is_Ord) |
149 apply (frule FFN_works [of concl: P], assumption+) |
152 apply (frule FFN_works [of concl: P], assumption+) |
150 apply (drule Cl_reflects, assumption+) |
153 apply (drule Cl_reflects, assumption+) |
151 apply (auto simp add: Limit_is_Ord Pair_in_Mset) |
154 apply (auto simp add: Limit_is_Ord Pair_in_Mset) |
152 done |
155 done |
153 |
156 |
154 lemma (in ex_reflection) ClEx_upward: |
157 lemma ClEx_upward: |
155 "[| z\<in>Mset(a); y\<in>Mset(a); Q(a,<y,z>); Cl(a); ClEx(P,a) |] |
158 "\<lbrakk>z\<in>Mset(a); y\<in>Mset(a); Q(a,<y,z>); Cl(a); ClEx(P,a)\<rbrakk> |
156 ==> \<exists>z. M(z) & P(<y,z>)" |
159 \<Longrightarrow> \<exists>z. M(z) & P(<y,z>)" |
157 apply (simp add: ClEx_def M_def) |
160 apply (simp add: ClEx_def M_def) |
158 apply (blast dest: Cl_reflects |
161 apply (blast dest: Cl_reflects |
159 intro: Limit_is_Ord Pair_in_Mset) |
162 intro: Limit_is_Ord Pair_in_Mset) |
160 done |
163 done |
161 |
164 |
162 text\<open>Class \<open>ClEx\<close> indeed consists of reflecting ordinals...\<close> |
165 text\<open>Class \<open>ClEx\<close> indeed consists of reflecting ordinals...\<close> |
163 lemma (in ex_reflection) ZF_ClEx_iff: |
166 lemma ZF_ClEx_iff: |
164 "[| y\<in>Mset(a); Cl(a); ClEx(P,a) |] |
167 "\<lbrakk>y\<in>Mset(a); Cl(a); ClEx(P,a)\<rbrakk> |
165 ==> (\<exists>z. M(z) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))" |
168 \<Longrightarrow> (\<exists>z. M(z) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))" |
166 by (blast intro: dest: ClEx_downward ClEx_upward) |
169 by (blast intro: dest: ClEx_downward ClEx_upward) |
167 |
170 |
168 text\<open>...and it is closed and unbounded\<close> |
171 text\<open>...and it is closed and unbounded\<close> |
169 lemma (in ex_reflection) ZF_Closed_Unbounded_ClEx: |
172 lemma ZF_Closed_Unbounded_ClEx: |
170 "Closed_Unbounded(ClEx(P))" |
173 "Closed_Unbounded(ClEx(P))" |
171 apply (simp add: ClEx_eq) |
174 apply (simp add: ClEx_eq) |
172 apply (fast intro: Closed_Unbounded_Int Normal_imp_fp_Closed_Unbounded |
175 apply (fast intro: Closed_Unbounded_Int Normal_imp_fp_Closed_Unbounded |
173 Closed_Unbounded_Limit Normal_normalize) |
176 Closed_Unbounded_Limit Normal_normalize) |
174 done |
177 done |
175 |
178 |
|
179 end |
|
180 |
176 text\<open>The same two theorems, exported to locale \<open>reflection\<close>.\<close> |
181 text\<open>The same two theorems, exported to locale \<open>reflection\<close>.\<close> |
177 |
182 |
|
183 context reflection |
|
184 begin |
|
185 |
178 text\<open>Class \<open>ClEx\<close> indeed consists of reflecting ordinals...\<close> |
186 text\<open>Class \<open>ClEx\<close> indeed consists of reflecting ordinals...\<close> |
179 lemma (in reflection) ClEx_iff: |
187 lemma ClEx_iff: |
180 "[| y\<in>Mset(a); Cl(a); ClEx(P,a); |
188 "\<lbrakk>y\<in>Mset(a); Cl(a); ClEx(P,a); |
181 !!a. [| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x) |] |
189 \<And>a. \<lbrakk>Cl(a); Ord(a)\<rbrakk> \<Longrightarrow> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)\<rbrakk> |
182 ==> (\<exists>z. M(z) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))" |
190 \<Longrightarrow> (\<exists>z. M(z) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))" |
183 apply (unfold ClEx_def FF_def F0_def M_def) |
191 apply (unfold ClEx_def FF_def F0_def M_def) |
184 apply (rule ex_reflection.ZF_ClEx_iff |
192 apply (rule ex_reflection.ZF_ClEx_iff |
185 [OF ex_reflection.intro, OF reflection.intro ex_reflection_axioms.intro, |
193 [OF ex_reflection.intro, OF reflection.intro ex_reflection_axioms.intro, |
186 of Mset Cl]) |
194 of Mset Cl]) |
187 apply (simp_all add: Mset_mono_le Mset_cont Pair_in_Mset) |
195 apply (simp_all add: Mset_mono_le Mset_cont Pair_in_Mset) |
192 apply (fold ClEx_def FF_def F0_def) |
200 apply (fold ClEx_def FF_def F0_def) |
193 apply (rule ex_reflection.intro, assumption) |
201 apply (rule ex_reflection.intro, assumption) |
194 apply (simp add: ex_reflection_axioms.intro, assumption+) |
202 apply (simp add: ex_reflection_axioms.intro, assumption+) |
195 *) |
203 *) |
196 |
204 |
197 lemma (in reflection) Closed_Unbounded_ClEx: |
205 lemma Closed_Unbounded_ClEx: |
198 "(!!a. [| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)) |
206 "(\<And>a. \<lbrakk>Cl(a); Ord(a)\<rbrakk> \<Longrightarrow> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)) |
199 ==> Closed_Unbounded(ClEx(P))" |
207 \<Longrightarrow> Closed_Unbounded(ClEx(P))" |
200 apply (unfold ClEx_eq FF_def F0_def M_def) |
208 apply (unfold ClEx_eq FF_def F0_def M_def) |
201 apply (rule ex_reflection.ZF_Closed_Unbounded_ClEx [of Mset _ _ Cl]) |
209 apply (rule ex_reflection.ZF_Closed_Unbounded_ClEx [of Mset _ _ Cl]) |
202 apply (rule ex_reflection.intro, rule reflection_axioms) |
210 apply (rule ex_reflection.intro, rule reflection_axioms) |
203 apply (blast intro: ex_reflection_axioms.intro) |
211 apply (blast intro: ex_reflection_axioms.intro) |
204 done |
212 done |
205 |
213 |
206 subsection\<open>Packaging the Quantifier Reflection Rules\<close> |
214 subsection\<open>Packaging the Quantifier Reflection Rules\<close> |
207 |
215 |
208 lemma (in reflection) Ex_reflection_0: |
216 lemma Ex_reflection_0: |
209 "Reflects(Cl,P0,Q0) |
217 "Reflects(Cl,P0,Q0) |
210 ==> Reflects(\<lambda>a. Cl(a) & ClEx(P0,a), |
218 \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(P0,a), |
211 \<lambda>x. \<exists>z. M(z) & P0(<x,z>), |
219 \<lambda>x. \<exists>z. M(z) & P0(<x,z>), |
212 \<lambda>a x. \<exists>z\<in>Mset(a). Q0(a,<x,z>))" |
220 \<lambda>a x. \<exists>z\<in>Mset(a). Q0(a,<x,z>))" |
213 apply (simp add: Reflects_def) |
221 apply (simp add: Reflects_def) |
214 apply (intro conjI Closed_Unbounded_Int) |
222 apply (intro conjI Closed_Unbounded_Int) |
215 apply blast |
223 apply blast |
216 apply (rule Closed_Unbounded_ClEx [of Cl P0 Q0], blast, clarify) |
224 apply (rule Closed_Unbounded_ClEx [of Cl P0 Q0], blast, clarify) |
217 apply (rule_tac Cl=Cl in ClEx_iff, assumption+, blast) |
225 apply (rule_tac Cl=Cl in ClEx_iff, assumption+, blast) |
218 done |
226 done |
219 |
227 |
220 lemma (in reflection) All_reflection_0: |
228 lemma All_reflection_0: |
221 "Reflects(Cl,P0,Q0) |
229 "Reflects(Cl,P0,Q0) |
222 ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x.~P0(x), a), |
230 \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x.\<not>P0(x), a), |
223 \<lambda>x. \<forall>z. M(z) \<longrightarrow> P0(<x,z>), |
231 \<lambda>x. \<forall>z. M(z) \<longrightarrow> P0(<x,z>), |
224 \<lambda>a x. \<forall>z\<in>Mset(a). Q0(a,<x,z>))" |
232 \<lambda>a x. \<forall>z\<in>Mset(a). Q0(a,<x,z>))" |
225 apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not) |
233 apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not) |
226 apply (rule Not_reflection, drule Not_reflection, simp) |
234 apply (rule Not_reflection, drule Not_reflection, simp) |
227 apply (erule Ex_reflection_0) |
235 apply (erule Ex_reflection_0) |
228 done |
236 done |
229 |
237 |
230 theorem (in reflection) Ex_reflection [intro]: |
238 theorem Ex_reflection [intro]: |
231 "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
239 "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
232 ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a), |
240 \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a), |
233 \<lambda>x. \<exists>z. M(z) & P(x,z), |
241 \<lambda>x. \<exists>z. M(z) & P(x,z), |
234 \<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))" |
242 \<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))" |
235 by (rule Ex_reflection_0 [of _ " \<lambda>x. P(fst(x),snd(x))" |
243 by (rule Ex_reflection_0 [of _ " \<lambda>x. P(fst(x),snd(x))" |
236 "\<lambda>a x. Q(a,fst(x),snd(x))", simplified]) |
244 "\<lambda>a x. Q(a,fst(x),snd(x))", simplified]) |
237 |
245 |
238 theorem (in reflection) All_reflection [intro]: |
246 theorem All_reflection [intro]: |
239 "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
247 "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
240 ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. ~P(fst(x),snd(x)), a), |
248 \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. \<not>P(fst(x),snd(x)), a), |
241 \<lambda>x. \<forall>z. M(z) \<longrightarrow> P(x,z), |
249 \<lambda>x. \<forall>z. M(z) \<longrightarrow> P(x,z), |
242 \<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))" |
250 \<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))" |
243 by (rule All_reflection_0 [of _ "\<lambda>x. P(fst(x),snd(x))" |
251 by (rule All_reflection_0 [of _ "\<lambda>x. P(fst(x),snd(x))" |
244 "\<lambda>a x. Q(a,fst(x),snd(x))", simplified]) |
252 "\<lambda>a x. Q(a,fst(x),snd(x))", simplified]) |
245 |
253 |
246 text\<open>And again, this time using class-bounded quantifiers\<close> |
254 text\<open>And again, this time using class-bounded quantifiers\<close> |
247 |
255 |
248 theorem (in reflection) Rex_reflection [intro]: |
256 theorem Rex_reflection [intro]: |
249 "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
257 "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
250 ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a), |
258 \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a), |
251 \<lambda>x. \<exists>z[M]. P(x,z), |
259 \<lambda>x. \<exists>z[M]. P(x,z), |
252 \<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))" |
260 \<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))" |
253 by (unfold rex_def, blast) |
261 by (unfold rex_def, blast) |
254 |
262 |
255 theorem (in reflection) Rall_reflection [intro]: |
263 theorem Rall_reflection [intro]: |
256 "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
264 "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))) |
257 ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. ~P(fst(x),snd(x)), a), |
265 \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. \<not>P(fst(x),snd(x)), a), |
258 \<lambda>x. \<forall>z[M]. P(x,z), |
266 \<lambda>x. \<forall>z[M]. P(x,z), |
259 \<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))" |
267 \<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))" |
260 by (unfold rall_def, blast) |
268 by (unfold rall_def, blast) |
261 |
269 |
262 |
270 |
266 subsection\<open>Simple Examples of Reflection\<close> |
274 subsection\<open>Simple Examples of Reflection\<close> |
267 |
275 |
268 text\<open>Example 1: reflecting a simple formula. The reflecting class is first |
276 text\<open>Example 1: reflecting a simple formula. The reflecting class is first |
269 given as the variable \<open>?Cl\<close> and later retrieved from the final |
277 given as the variable \<open>?Cl\<close> and later retrieved from the final |
270 proof state.\<close> |
278 proof state.\<close> |
271 schematic_goal (in reflection) |
279 schematic_goal |
272 "Reflects(?Cl, |
280 "Reflects(?Cl, |
273 \<lambda>x. \<exists>y. M(y) & x \<in> y, |
281 \<lambda>x. \<exists>y. M(y) & x \<in> y, |
274 \<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)" |
282 \<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)" |
275 by fast |
283 by fast |
276 |
284 |
277 text\<open>Problem here: there needs to be a conjunction (class intersection) |
285 text\<open>Problem here: there needs to be a conjunction (class intersection) |
278 in the class of reflecting ordinals. The \<^term>\<open>Ord(a)\<close> is redundant, |
286 in the class of reflecting ordinals. The \<^term>\<open>Ord(a)\<close> is redundant, |
279 though harmless.\<close> |
287 though harmless.\<close> |
280 lemma (in reflection) |
288 lemma |
281 "Reflects(\<lambda>a. Ord(a) & ClEx(\<lambda>x. fst(x) \<in> snd(x), a), |
289 "Reflects(\<lambda>a. Ord(a) & ClEx(\<lambda>x. fst(x) \<in> snd(x), a), |
282 \<lambda>x. \<exists>y. M(y) & x \<in> y, |
290 \<lambda>x. \<exists>y. M(y) & x \<in> y, |
283 \<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)" |
291 \<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)" |
284 by fast |
292 by fast |
285 |
293 |
286 |
294 |
287 text\<open>Example 2\<close> |
295 text\<open>Example 2\<close> |
288 schematic_goal (in reflection) |
296 schematic_goal |
289 "Reflects(?Cl, |
297 "Reflects(?Cl, |
290 \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y), |
298 \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y), |
291 \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)" |
299 \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)" |
292 by fast |
300 by fast |
293 |
301 |
294 text\<open>Example 2'. We give the reflecting class explicitly.\<close> |
302 text\<open>Example 2'. We give the reflecting class explicitly.\<close> |
295 lemma (in reflection) |
303 lemma |
296 "Reflects |
304 "Reflects |
297 (\<lambda>a. (Ord(a) & |
305 (\<lambda>a. (Ord(a) & |
298 ClEx(\<lambda>x. ~ (snd(x) \<subseteq> fst(fst(x)) \<longrightarrow> snd(x) \<in> snd(fst(x))), a)) & |
306 ClEx(\<lambda>x. \<not> (snd(x) \<subseteq> fst(fst(x)) \<longrightarrow> snd(x) \<in> snd(fst(x))), a)) & |
299 ClEx(\<lambda>x. \<forall>z. M(z) \<longrightarrow> z \<subseteq> fst(x) \<longrightarrow> z \<in> snd(x), a), |
307 ClEx(\<lambda>x. \<forall>z. M(z) \<longrightarrow> z \<subseteq> fst(x) \<longrightarrow> z \<in> snd(x), a), |
300 \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y), |
308 \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y), |
301 \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)" |
309 \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)" |
302 by fast |
310 by fast |
303 |
311 |
304 text\<open>Example 2''. We expand the subset relation.\<close> |
312 text\<open>Example 2''. We expand the subset relation.\<close> |
305 schematic_goal (in reflection) |
313 schematic_goal |
306 "Reflects(?Cl, |
314 "Reflects(?Cl, |
307 \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> (\<forall>w. M(w) \<longrightarrow> w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y), |
315 \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> (\<forall>w. M(w) \<longrightarrow> w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y), |
308 \<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)" |
316 \<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)" |
309 by fast |
317 by fast |
310 |
318 |
311 text\<open>Example 2'''. Single-step version, to reveal the reflecting class.\<close> |
319 text\<open>Example 2'''. Single-step version, to reveal the reflecting class.\<close> |
312 schematic_goal (in reflection) |
320 schematic_goal |
313 "Reflects(?Cl, |
321 "Reflects(?Cl, |
314 \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y), |
322 \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y), |
315 \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)" |
323 \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)" |
316 apply (rule Ex_reflection) |
324 apply (rule Ex_reflection) |
317 txt\<open> |
325 txt\<open> |