13505
|
1 |
(* Title: ZF/Constructible/MetaExists.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
*)
|
|
5 |
|
13314
|
6 |
header{*The meta-existential quantifier*}
|
|
7 |
|
16417
|
8 |
theory MetaExists imports Main begin
|
13314
|
9 |
|
|
10 |
text{*Allows quantification over any term having sort @{text logic}. Used to
|
|
11 |
quantify over classes. Yields a proposition rather than a FOL formula.*}
|
|
12 |
|
|
13 |
constdefs
|
14854
|
14 |
ex :: "(('a::{}) => prop) => prop" (binder "?? " 0)
|
13314
|
15 |
"ex(P) == (!!Q. (!!x. PROP P(x) ==> PROP Q) ==> PROP Q)"
|
|
16 |
|
|
17 |
syntax (xsymbols)
|
|
18 |
"?? " :: "[idts, o] => o" ("(3\<Or>_./ _)" [0, 0] 0)
|
|
19 |
|
|
20 |
lemma meta_exI: "PROP P(x) ==> (?? x. PROP P(x))"
|
13315
|
21 |
proof (unfold ex_def)
|
13314
|
22 |
assume P: "PROP P(x)"
|
13315
|
23 |
fix Q
|
|
24 |
assume PQ: "\<And>x. PROP P(x) \<Longrightarrow> PROP Q"
|
|
25 |
from P show "PROP Q" by (rule PQ)
|
13314
|
26 |
qed
|
|
27 |
|
|
28 |
lemma meta_exE: "[| ?? x. PROP P(x); !!x. PROP P(x) ==> PROP R |] ==> PROP R"
|
13315
|
29 |
proof (unfold ex_def)
|
13314
|
30 |
assume QPQ: "\<And>Q. (\<And>x. PROP P(x) \<Longrightarrow> PROP Q) \<Longrightarrow> PROP Q"
|
|
31 |
assume PR: "\<And>x. PROP P(x) \<Longrightarrow> PROP R"
|
|
32 |
from PR show "PROP R" by (rule QPQ)
|
|
33 |
qed
|
|
34 |
|
|
35 |
end
|