author | blanchet |
Wed, 06 Nov 2013 22:50:12 +0100 | |
changeset 54284 | 0b53378080d9 |
parent 32960 | 69916a850301 |
child 58871 | c399ae4b836f |
permissions | -rw-r--r-- |
13505 | 1 |
(* Title: ZF/Constructible/MetaExists.thy |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
3 |
*) |
|
4 |
||
13314 | 5 |
header{*The meta-existential quantifier*} |
6 |
||
16417 | 7 |
theory MetaExists imports Main begin |
13314 | 8 |
|
9 |
text{*Allows quantification over any term having sort @{text logic}. Used to |
|
10 |
quantify over classes. Yields a proposition rather than a FOL formula.*} |
|
11 |
||
21233 | 12 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
13 |
ex :: "(('a::{}) => prop) => prop" (binder "?? " 0) where |
13314 | 14 |
"ex(P) == (!!Q. (!!x. PROP P(x) ==> PROP Q) ==> PROP Q)" |
15 |
||
21233 | 16 |
notation (xsymbols) |
17 |
ex (binder "\<Or>" 0) |
|
13314 | 18 |
|
19 |
lemma meta_exI: "PROP P(x) ==> (?? x. PROP P(x))" |
|
13315 | 20 |
proof (unfold ex_def) |
13314 | 21 |
assume P: "PROP P(x)" |
13315 | 22 |
fix Q |
23 |
assume PQ: "\<And>x. PROP P(x) \<Longrightarrow> PROP Q" |
|
24 |
from P show "PROP Q" by (rule PQ) |
|
13314 | 25 |
qed |
26 |
||
27 |
lemma meta_exE: "[| ?? x. PROP P(x); !!x. PROP P(x) ==> PROP R |] ==> PROP R" |
|
13315 | 28 |
proof (unfold ex_def) |
13314 | 29 |
assume QPQ: "\<And>Q. (\<And>x. PROP P(x) \<Longrightarrow> PROP Q) \<Longrightarrow> PROP Q" |
30 |
assume PR: "\<And>x. PROP P(x) \<Longrightarrow> PROP R" |
|
31 |
from PR show "PROP R" by (rule QPQ) |
|
32 |
qed |
|
33 |
||
34 |
end |