author | nipkow |
Sun, 05 Jan 2003 21:03:14 +0100 | |
changeset 13771 | 6cd59cc885a1 |
parent 13615 | 449a70d88b38 |
child 13807 | a28a8fbc76d4 |
permissions | -rw-r--r-- |
2469 | 1 |
(* Title: ZF/AC/OrdQuant.thy |
2 |
ID: $Id$ |
|
3 |
Authors: Krzysztof Grabczewski and L C Paulson |
|
4 |
*) |
|
5 |
||
13253 | 6 |
header {*Special quantifiers*} |
7 |
||
12620 | 8 |
theory OrdQuant = Ordinal: |
2469 | 9 |
|
13253 | 10 |
subsection {*Quantifiers and union operator for ordinals*} |
11 |
||
12620 | 12 |
constdefs |
13298 | 13 |
|
2469 | 14 |
(* Ordinal Quantifiers *) |
12620 | 15 |
oall :: "[i, i => o] => o" |
16 |
"oall(A, P) == ALL x. x<A --> P(x)" |
|
13298 | 17 |
|
12620 | 18 |
oex :: "[i, i => o] => o" |
19 |
"oex(A, P) == EX x. x<A & P(x)" |
|
2469 | 20 |
|
21 |
(* Ordinal Union *) |
|
12620 | 22 |
OUnion :: "[i, i => i] => i" |
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
23 |
"OUnion(i,B) == {z: \<Union>x\<in>i. B(x). Ord(i)}" |
13298 | 24 |
|
2469 | 25 |
syntax |
12620 | 26 |
"@oall" :: "[idt, i, o] => o" ("(3ALL _<_./ _)" 10) |
27 |
"@oex" :: "[idt, i, o] => o" ("(3EX _<_./ _)" 10) |
|
28 |
"@OUNION" :: "[idt, i, i] => i" ("(3UN _<_./ _)" 10) |
|
2469 | 29 |
|
30 |
translations |
|
31 |
"ALL x<a. P" == "oall(a, %x. P)" |
|
32 |
"EX x<a. P" == "oex(a, %x. P)" |
|
33 |
"UN x<a. B" == "OUnion(a, %x. B)" |
|
34 |
||
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
6093
diff
changeset
|
35 |
syntax (xsymbols) |
12620 | 36 |
"@oall" :: "[idt, i, o] => o" ("(3\<forall>_<_./ _)" 10) |
37 |
"@oex" :: "[idt, i, o] => o" ("(3\<exists>_<_./ _)" 10) |
|
38 |
"@OUNION" :: "[idt, i, i] => i" ("(3\<Union>_<_./ _)" 10) |
|
39 |
||
40 |
||
13302 | 41 |
subsubsection {*simplification of the new quantifiers*} |
12825 | 42 |
|
43 |
||
13169 | 44 |
(*MOST IMPORTANT that this is added to the simpset BEFORE Ord_atomize |
13298 | 45 |
is proved. Ord_atomize would convert this rule to |
12825 | 46 |
x < 0 ==> P(x) == True, which causes dire effects!*) |
47 |
lemma [simp]: "(ALL x<0. P(x))" |
|
13298 | 48 |
by (simp add: oall_def) |
12825 | 49 |
|
50 |
lemma [simp]: "~(EX x<0. P(x))" |
|
13298 | 51 |
by (simp add: oex_def) |
12825 | 52 |
|
53 |
lemma [simp]: "(ALL x<succ(i). P(x)) <-> (Ord(i) --> P(i) & (ALL x<i. P(x)))" |
|
13298 | 54 |
apply (simp add: oall_def le_iff) |
55 |
apply (blast intro: lt_Ord2) |
|
12825 | 56 |
done |
57 |
||
58 |
lemma [simp]: "(EX x<succ(i). P(x)) <-> (Ord(i) & (P(i) | (EX x<i. P(x))))" |
|
13298 | 59 |
apply (simp add: oex_def le_iff) |
60 |
apply (blast intro: lt_Ord2) |
|
12825 | 61 |
done |
62 |
||
13302 | 63 |
subsubsection {*Union over ordinals*} |
13118 | 64 |
|
12620 | 65 |
lemma Ord_OUN [intro,simp]: |
13162
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
paulson
parents:
13149
diff
changeset
|
66 |
"[| !!x. x<A ==> Ord(B(x)) |] ==> Ord(\<Union>x<A. B(x))" |
13298 | 67 |
by (simp add: OUnion_def ltI Ord_UN) |
12620 | 68 |
|
69 |
lemma OUN_upper_lt: |
|
13162
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
paulson
parents:
13149
diff
changeset
|
70 |
"[| a<A; i < b(a); Ord(\<Union>x<A. b(x)) |] ==> i < (\<Union>x<A. b(x))" |
12620 | 71 |
by (unfold OUnion_def lt_def, blast ) |
72 |
||
73 |
lemma OUN_upper_le: |
|
13162
660a71e712af
New theorems from Constructible, and moving some Isar material from Main
paulson
parents:
13149
diff
changeset
|
74 |
"[| a<A; i\<le>b(a); Ord(\<Union>x<A. b(x)) |] ==> i \<le> (\<Union>x<A. b(x))" |
12820 | 75 |
apply (unfold OUnion_def, auto) |
12620 | 76 |
apply (rule UN_upper_le ) |
13298 | 77 |
apply (auto simp add: lt_def) |
12620 | 78 |
done |
2469 | 79 |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
80 |
lemma Limit_OUN_eq: "Limit(i) ==> (\<Union>x<i. x) = i" |
12620 | 81 |
by (simp add: OUnion_def Limit_Union_eq Limit_is_Ord) |
82 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
83 |
(* No < version; consider (\<Union>i\<in>nat.i)=nat *) |
12620 | 84 |
lemma OUN_least: |
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
85 |
"(!!x. x<A ==> B(x) \<subseteq> C) ==> (\<Union>x<A. B(x)) \<subseteq> C" |
12620 | 86 |
by (simp add: OUnion_def UN_least ltI) |
87 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
88 |
(* No < version; consider (\<Union>i\<in>nat.i)=nat *) |
12620 | 89 |
lemma OUN_least_le: |
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
90 |
"[| Ord(i); !!x. x<A ==> b(x) \<le> i |] ==> (\<Union>x<A. b(x)) \<le> i" |
12620 | 91 |
by (simp add: OUnion_def UN_least_le ltI Ord_0_le) |
92 |
||
93 |
lemma le_implies_OUN_le_OUN: |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
94 |
"[| !!x. x<A ==> c(x) \<le> d(x) |] ==> (\<Union>x<A. c(x)) \<le> (\<Union>x<A. d(x))" |
12620 | 95 |
by (blast intro: OUN_least_le OUN_upper_le le_Ord2 Ord_OUN) |
96 |
||
97 |
lemma OUN_UN_eq: |
|
98 |
"(!!x. x:A ==> Ord(B(x))) |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
99 |
==> (\<Union>z < (\<Union>x\<in>A. B(x)). C(z)) = (\<Union>x\<in>A. \<Union>z < B(x). C(z))" |
13298 | 100 |
by (simp add: OUnion_def) |
12620 | 101 |
|
102 |
lemma OUN_Union_eq: |
|
103 |
"(!!x. x:X ==> Ord(x)) |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
104 |
==> (\<Union>z < Union(X). C(z)) = (\<Union>x\<in>X. \<Union>z < x. C(z))" |
13298 | 105 |
by (simp add: OUnion_def) |
12620 | 106 |
|
12763 | 107 |
(*So that rule_format will get rid of ALL x<A...*) |
108 |
lemma atomize_oall [symmetric, rulify]: |
|
109 |
"(!!x. x<A ==> P(x)) == Trueprop (ALL x<A. P(x))" |
|
110 |
by (simp add: oall_def atomize_all atomize_imp) |
|
111 |
||
13302 | 112 |
subsubsection {*universal quantifier for ordinals*} |
13169 | 113 |
|
114 |
lemma oallI [intro!]: |
|
115 |
"[| !!x. x<A ==> P(x) |] ==> ALL x<A. P(x)" |
|
13298 | 116 |
by (simp add: oall_def) |
13169 | 117 |
|
118 |
lemma ospec: "[| ALL x<A. P(x); x<A |] ==> P(x)" |
|
13298 | 119 |
by (simp add: oall_def) |
13169 | 120 |
|
121 |
lemma oallE: |
|
122 |
"[| ALL x<A. P(x); P(x) ==> Q; ~x<A ==> Q |] ==> Q" |
|
13298 | 123 |
by (simp add: oall_def, blast) |
13169 | 124 |
|
125 |
lemma rev_oallE [elim]: |
|
126 |
"[| ALL x<A. P(x); ~x<A ==> Q; P(x) ==> Q |] ==> Q" |
|
13298 | 127 |
by (simp add: oall_def, blast) |
13169 | 128 |
|
129 |
||
130 |
(*Trival rewrite rule; (ALL x<a.P)<->P holds only if a is not 0!*) |
|
131 |
lemma oall_simp [simp]: "(ALL x<a. True) <-> True" |
|
13170 | 132 |
by blast |
13169 | 133 |
|
134 |
(*Congruence rule for rewriting*) |
|
135 |
lemma oall_cong [cong]: |
|
13298 | 136 |
"[| a=a'; !!x. x<a' ==> P(x) <-> P'(x) |] |
13289
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
paulson
parents:
13253
diff
changeset
|
137 |
==> oall(a, %x. P(x)) <-> oall(a', %x. P'(x))" |
13169 | 138 |
by (simp add: oall_def) |
139 |
||
140 |
||
13302 | 141 |
subsubsection {*existential quantifier for ordinals*} |
13169 | 142 |
|
143 |
lemma oexI [intro]: |
|
144 |
"[| P(x); x<A |] ==> EX x<A. P(x)" |
|
13298 | 145 |
apply (simp add: oex_def, blast) |
13169 | 146 |
done |
147 |
||
148 |
(*Not of the general form for such rules; ~EX has become ALL~ *) |
|
149 |
lemma oexCI: |
|
150 |
"[| ALL x<A. ~P(x) ==> P(a); a<A |] ==> EX x<A. P(x)" |
|
13298 | 151 |
apply (simp add: oex_def, blast) |
13169 | 152 |
done |
153 |
||
154 |
lemma oexE [elim!]: |
|
155 |
"[| EX x<A. P(x); !!x. [| x<A; P(x) |] ==> Q |] ==> Q" |
|
13298 | 156 |
apply (simp add: oex_def, blast) |
13169 | 157 |
done |
158 |
||
159 |
lemma oex_cong [cong]: |
|
13298 | 160 |
"[| a=a'; !!x. x<a' ==> P(x) <-> P'(x) |] |
13289
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
paulson
parents:
13253
diff
changeset
|
161 |
==> oex(a, %x. P(x)) <-> oex(a', %x. P'(x))" |
13169 | 162 |
apply (simp add: oex_def cong add: conj_cong) |
163 |
done |
|
164 |
||
165 |
||
13302 | 166 |
subsubsection {*Rules for Ordinal-Indexed Unions*} |
13169 | 167 |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
168 |
lemma OUN_I [intro]: "[| a<i; b: B(a) |] ==> b: (\<Union>z<i. B(z))" |
13170 | 169 |
by (unfold OUnion_def lt_def, blast) |
13169 | 170 |
|
171 |
lemma OUN_E [elim!]: |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
172 |
"[| b : (\<Union>z<i. B(z)); !!a.[| b: B(a); a<i |] ==> R |] ==> R" |
13170 | 173 |
apply (unfold OUnion_def lt_def, blast) |
13169 | 174 |
done |
175 |
||
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
176 |
lemma OUN_iff: "b : (\<Union>x<i. B(x)) <-> (EX x<i. b : B(x))" |
13170 | 177 |
by (unfold OUnion_def oex_def lt_def, blast) |
13169 | 178 |
|
179 |
lemma OUN_cong [cong]: |
|
13615
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
paulson
parents:
13462
diff
changeset
|
180 |
"[| i=j; !!x. x<j ==> C(x)=D(x) |] ==> (\<Union>x<i. C(x)) = (\<Union>x<j. D(x))" |
13169 | 181 |
by (simp add: OUnion_def lt_def OUN_iff) |
182 |
||
13298 | 183 |
lemma lt_induct: |
13169 | 184 |
"[| i<k; !!x.[| x<k; ALL y<x. P(y) |] ==> P(x) |] ==> P(i)" |
185 |
apply (simp add: lt_def oall_def) |
|
13298 | 186 |
apply (erule conjE) |
187 |
apply (erule Ord_induct, assumption, blast) |
|
13169 | 188 |
done |
189 |
||
13253 | 190 |
|
191 |
subsection {*Quantification over a class*} |
|
192 |
||
193 |
constdefs |
|
194 |
"rall" :: "[i=>o, i=>o] => o" |
|
195 |
"rall(M, P) == ALL x. M(x) --> P(x)" |
|
196 |
||
197 |
"rex" :: "[i=>o, i=>o] => o" |
|
198 |
"rex(M, P) == EX x. M(x) & P(x)" |
|
199 |
||
200 |
syntax |
|
201 |
"@rall" :: "[pttrn, i=>o, o] => o" ("(3ALL _[_]./ _)" 10) |
|
202 |
"@rex" :: "[pttrn, i=>o, o] => o" ("(3EX _[_]./ _)" 10) |
|
203 |
||
204 |
syntax (xsymbols) |
|
205 |
"@rall" :: "[pttrn, i=>o, o] => o" ("(3\<forall>_[_]./ _)" 10) |
|
206 |
"@rex" :: "[pttrn, i=>o, o] => o" ("(3\<exists>_[_]./ _)" 10) |
|
207 |
||
208 |
translations |
|
209 |
"ALL x[M]. P" == "rall(M, %x. P)" |
|
210 |
"EX x[M]. P" == "rex(M, %x. P)" |
|
211 |
||
13298 | 212 |
|
213 |
subsubsection{*Relativized universal quantifier*} |
|
13253 | 214 |
|
215 |
lemma rallI [intro!]: "[| !!x. M(x) ==> P(x) |] ==> ALL x[M]. P(x)" |
|
216 |
by (simp add: rall_def) |
|
217 |
||
218 |
lemma rspec: "[| ALL x[M]. P(x); M(x) |] ==> P(x)" |
|
219 |
by (simp add: rall_def) |
|
220 |
||
221 |
(*Instantiates x first: better for automatic theorem proving?*) |
|
13298 | 222 |
lemma rev_rallE [elim]: |
13253 | 223 |
"[| ALL x[M]. P(x); ~ M(x) ==> Q; P(x) ==> Q |] ==> Q" |
13298 | 224 |
by (simp add: rall_def, blast) |
13253 | 225 |
|
226 |
lemma rallE: "[| ALL x[M]. P(x); P(x) ==> Q; ~ M(x) ==> Q |] ==> Q" |
|
227 |
by blast |
|
228 |
||
229 |
(*Trival rewrite rule; (ALL x[M].P)<->P holds only if A is nonempty!*) |
|
230 |
lemma rall_triv [simp]: "(ALL x[M]. P) <-> ((EX x. M(x)) --> P)" |
|
231 |
by (simp add: rall_def) |
|
232 |
||
233 |
(*Congruence rule for rewriting*) |
|
234 |
lemma rall_cong [cong]: |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13302
diff
changeset
|
235 |
"(!!x. M(x) ==> P(x) <-> P'(x)) ==> (ALL x[M]. P(x)) <-> (ALL x[M]. P'(x))" |
13253 | 236 |
by (simp add: rall_def) |
237 |
||
13298 | 238 |
|
239 |
subsubsection{*Relativized existential quantifier*} |
|
13253 | 240 |
|
241 |
lemma rexI [intro]: "[| P(x); M(x) |] ==> EX x[M]. P(x)" |
|
242 |
by (simp add: rex_def, blast) |
|
243 |
||
244 |
(*The best argument order when there is only one M(x)*) |
|
245 |
lemma rev_rexI: "[| M(x); P(x) |] ==> EX x[M]. P(x)" |
|
246 |
by blast |
|
247 |
||
248 |
(*Not of the general form for such rules; ~EX has become ALL~ *) |
|
249 |
lemma rexCI: "[| ALL x[M]. ~P(x) ==> P(a); M(a) |] ==> EX x[M]. P(x)" |
|
250 |
by blast |
|
251 |
||
252 |
lemma rexE [elim!]: "[| EX x[M]. P(x); !!x. [| M(x); P(x) |] ==> Q |] ==> Q" |
|
253 |
by (simp add: rex_def, blast) |
|
254 |
||
255 |
(*We do not even have (EX x[M]. True) <-> True unless A is nonempty!!*) |
|
256 |
lemma rex_triv [simp]: "(EX x[M]. P) <-> ((EX x. M(x)) & P)" |
|
257 |
by (simp add: rex_def) |
|
258 |
||
259 |
lemma rex_cong [cong]: |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13302
diff
changeset
|
260 |
"(!!x. M(x) ==> P(x) <-> P'(x)) ==> (EX x[M]. P(x)) <-> (EX x[M]. P'(x))" |
13253 | 261 |
by (simp add: rex_def cong: conj_cong) |
262 |
||
13289
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
paulson
parents:
13253
diff
changeset
|
263 |
lemma rall_is_ball [simp]: "(\<forall>x[%z. z\<in>A]. P(x)) <-> (\<forall>x\<in>A. P(x))" |
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
paulson
parents:
13253
diff
changeset
|
264 |
by blast |
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
paulson
parents:
13253
diff
changeset
|
265 |
|
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
paulson
parents:
13253
diff
changeset
|
266 |
lemma rex_is_bex [simp]: "(\<exists>x[%z. z\<in>A]. P(x)) <-> (\<exists>x\<in>A. P(x))" |
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
paulson
parents:
13253
diff
changeset
|
267 |
by blast |
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
paulson
parents:
13253
diff
changeset
|
268 |
|
13253 | 269 |
lemma atomize_rall: "(!!x. M(x) ==> P(x)) == Trueprop (ALL x[M]. P(x))"; |
270 |
by (simp add: rall_def atomize_all atomize_imp) |
|
271 |
||
272 |
declare atomize_rall [symmetric, rulify] |
|
273 |
||
274 |
lemma rall_simps1: |
|
275 |
"(ALL x[M]. P(x) & Q) <-> (ALL x[M]. P(x)) & ((ALL x[M]. False) | Q)" |
|
276 |
"(ALL x[M]. P(x) | Q) <-> ((ALL x[M]. P(x)) | Q)" |
|
277 |
"(ALL x[M]. P(x) --> Q) <-> ((EX x[M]. P(x)) --> Q)" |
|
13298 | 278 |
"(~(ALL x[M]. P(x))) <-> (EX x[M]. ~P(x))" |
13253 | 279 |
by blast+ |
280 |
||
281 |
lemma rall_simps2: |
|
282 |
"(ALL x[M]. P & Q(x)) <-> ((ALL x[M]. False) | P) & (ALL x[M]. Q(x))" |
|
283 |
"(ALL x[M]. P | Q(x)) <-> (P | (ALL x[M]. Q(x)))" |
|
284 |
"(ALL x[M]. P --> Q(x)) <-> (P --> (ALL x[M]. Q(x)))" |
|
285 |
by blast+ |
|
286 |
||
13289
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
paulson
parents:
13253
diff
changeset
|
287 |
lemmas rall_simps [simp] = rall_simps1 rall_simps2 |
13253 | 288 |
|
289 |
lemma rall_conj_distrib: |
|
290 |
"(ALL x[M]. P(x) & Q(x)) <-> ((ALL x[M]. P(x)) & (ALL x[M]. Q(x)))" |
|
291 |
by blast |
|
292 |
||
293 |
lemma rex_simps1: |
|
294 |
"(EX x[M]. P(x) & Q) <-> ((EX x[M]. P(x)) & Q)" |
|
295 |
"(EX x[M]. P(x) | Q) <-> (EX x[M]. P(x)) | ((EX x[M]. True) & Q)" |
|
296 |
"(EX x[M]. P(x) --> Q) <-> ((ALL x[M]. P(x)) --> ((EX x[M]. True) & Q))" |
|
297 |
"(~(EX x[M]. P(x))) <-> (ALL x[M]. ~P(x))" |
|
298 |
by blast+ |
|
299 |
||
300 |
lemma rex_simps2: |
|
301 |
"(EX x[M]. P & Q(x)) <-> (P & (EX x[M]. Q(x)))" |
|
302 |
"(EX x[M]. P | Q(x)) <-> ((EX x[M]. True) & P) | (EX x[M]. Q(x))" |
|
303 |
"(EX x[M]. P --> Q(x)) <-> (((ALL x[M]. False) | P) --> (EX x[M]. Q(x)))" |
|
304 |
by blast+ |
|
305 |
||
13289
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
paulson
parents:
13253
diff
changeset
|
306 |
lemmas rex_simps [simp] = rex_simps1 rex_simps2 |
13253 | 307 |
|
308 |
lemma rex_disj_distrib: |
|
309 |
"(EX x[M]. P(x) | Q(x)) <-> ((EX x[M]. P(x)) | (EX x[M]. Q(x)))" |
|
310 |
by blast |
|
311 |
||
312 |
||
13298 | 313 |
subsubsection{*One-point rule for bounded quantifiers*} |
13253 | 314 |
|
315 |
lemma rex_triv_one_point1 [simp]: "(EX x[M]. x=a) <-> ( M(a))" |
|
316 |
by blast |
|
317 |
||
318 |
lemma rex_triv_one_point2 [simp]: "(EX x[M]. a=x) <-> ( M(a))" |
|
319 |
by blast |
|
320 |
||
321 |
lemma rex_one_point1 [simp]: "(EX x[M]. x=a & P(x)) <-> ( M(a) & P(a))" |
|
322 |
by blast |
|
323 |
||
324 |
lemma rex_one_point2 [simp]: "(EX x[M]. a=x & P(x)) <-> ( M(a) & P(a))" |
|
325 |
by blast |
|
326 |
||
327 |
lemma rall_one_point1 [simp]: "(ALL x[M]. x=a --> P(x)) <-> ( M(a) --> P(a))" |
|
328 |
by blast |
|
329 |
||
330 |
lemma rall_one_point2 [simp]: "(ALL x[M]. a=x --> P(x)) <-> ( M(a) --> P(a))" |
|
331 |
by blast |
|
332 |
||
333 |
||
13298 | 334 |
subsubsection{*Sets as Classes*} |
335 |
||
13302 | 336 |
constdefs setclass :: "[i,i] => o" ("**_" [40] 40) |
13362 | 337 |
"setclass(A) == %x. x : A" |
13298 | 338 |
|
13362 | 339 |
lemma setclass_iff [simp]: "setclass(A,x) <-> x : A" |
340 |
by (simp add: setclass_def) |
|
13298 | 341 |
|
342 |
lemma rall_setclass_is_ball [simp]: "(\<forall>x[**A]. P(x)) <-> (\<forall>x\<in>A. P(x))" |
|
343 |
by auto |
|
344 |
||
345 |
lemma rex_setclass_is_bex [simp]: "(\<exists>x[**A]. P(x)) <-> (\<exists>x\<in>A. P(x))" |
|
346 |
by auto |
|
347 |
||
348 |
||
13169 | 349 |
ML |
350 |
{* |
|
351 |
val oall_def = thm "oall_def" |
|
352 |
val oex_def = thm "oex_def" |
|
353 |
val OUnion_def = thm "OUnion_def" |
|
354 |
||
355 |
val oallI = thm "oallI"; |
|
356 |
val ospec = thm "ospec"; |
|
357 |
val oallE = thm "oallE"; |
|
358 |
val rev_oallE = thm "rev_oallE"; |
|
359 |
val oall_simp = thm "oall_simp"; |
|
360 |
val oall_cong = thm "oall_cong"; |
|
361 |
val oexI = thm "oexI"; |
|
362 |
val oexCI = thm "oexCI"; |
|
363 |
val oexE = thm "oexE"; |
|
364 |
val oex_cong = thm "oex_cong"; |
|
365 |
val OUN_I = thm "OUN_I"; |
|
366 |
val OUN_E = thm "OUN_E"; |
|
367 |
val OUN_iff = thm "OUN_iff"; |
|
368 |
val OUN_cong = thm "OUN_cong"; |
|
369 |
val lt_induct = thm "lt_induct"; |
|
370 |
||
13253 | 371 |
val rall_def = thm "rall_def" |
372 |
val rex_def = thm "rex_def" |
|
373 |
||
374 |
val rallI = thm "rallI"; |
|
375 |
val rspec = thm "rspec"; |
|
376 |
val rallE = thm "rallE"; |
|
377 |
val rev_oallE = thm "rev_oallE"; |
|
378 |
val rall_cong = thm "rall_cong"; |
|
379 |
val rexI = thm "rexI"; |
|
380 |
val rexCI = thm "rexCI"; |
|
381 |
val rexE = thm "rexE"; |
|
382 |
val rex_cong = thm "rex_cong"; |
|
383 |
||
13169 | 384 |
val Ord_atomize = |
13253 | 385 |
atomize ([("OrdQuant.oall", [ospec]),("OrdQuant.rall", [rspec])]@ |
13298 | 386 |
ZF_conn_pairs, |
13253 | 387 |
ZF_mem_pairs); |
13169 | 388 |
simpset_ref() := simpset() setmksimps (map mk_eq o Ord_atomize o gen_all); |
389 |
*} |
|
390 |
||
13462 | 391 |
text {* Setting up the one-point-rule simproc *} |
13253 | 392 |
|
13462 | 393 |
ML_setup {* |
394 |
local |
|
13253 | 395 |
|
13462 | 396 |
val prove_rex_tac = rewtac rex_def THEN Quantifier1.prove_one_point_ex_tac; |
13253 | 397 |
val rearrange_bex = Quantifier1.rearrange_bex prove_rex_tac; |
398 |
||
13462 | 399 |
val prove_rall_tac = rewtac rall_def THEN Quantifier1.prove_one_point_all_tac; |
13253 | 400 |
val rearrange_ball = Quantifier1.rearrange_ball prove_rall_tac; |
401 |
||
402 |
in |
|
403 |
||
13462 | 404 |
val defREX_regroup = Simplifier.simproc (Theory.sign_of (the_context ())) |
405 |
"defined REX" ["EX x[M]. P(x) & Q(x)"] rearrange_bex; |
|
406 |
val defRALL_regroup = Simplifier.simproc (Theory.sign_of (the_context ())) |
|
407 |
"defined RALL" ["ALL x[M]. P(x) --> Q(x)"] rearrange_ball; |
|
13253 | 408 |
|
409 |
end; |
|
13462 | 410 |
|
411 |
Addsimprocs [defRALL_regroup,defREX_regroup]; |
|
13253 | 412 |
*} |
413 |
||
2469 | 414 |
end |