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