23914
|
1 |
(* Title: FOL/ex/Quantifiers_Int.thy
|
|
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
3 |
Copyright 1991 University of Cambridge
|
|
4 |
*)
|
|
5 |
|
60770
|
6 |
section \<open>First-Order Logic: quantifier examples (intuitionistic version)\<close>
|
23914
|
7 |
|
|
8 |
theory Quantifiers_Int
|
|
9 |
imports IFOL
|
|
10 |
begin
|
|
11 |
|
61489
|
12 |
lemma "(\<forall>x y. P(x,y)) \<longrightarrow> (\<forall>y x. P(x,y))"
|
51798
|
13 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
14 |
|
61489
|
15 |
lemma "(\<exists>x y. P(x,y)) \<longrightarrow> (\<exists>y x. P(x,y))"
|
51798
|
16 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
17 |
|
|
18 |
|
62020
|
19 |
\<comment> \<open>Converse is false\<close>
|
61489
|
20 |
lemma "(\<forall>x. P(x)) \<or> (\<forall>x. Q(x)) \<longrightarrow> (\<forall>x. P(x) \<or> Q(x))"
|
51798
|
21 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
22 |
|
61489
|
23 |
lemma "(\<forall>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<forall>x. Q(x)))"
|
51798
|
24 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
25 |
|
|
26 |
|
61489
|
27 |
lemma "(\<forall>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<exists>x. P(x)) \<longrightarrow> Q)"
|
51798
|
28 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
29 |
|
|
30 |
|
60770
|
31 |
text \<open>Some harder ones\<close>
|
23914
|
32 |
|
61489
|
33 |
lemma "(\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> (\<exists>x. P(x)) \<or> (\<exists>x. Q(x))"
|
51798
|
34 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
35 |
|
62020
|
36 |
\<comment> \<open>Converse is false\<close>
|
61489
|
37 |
lemma "(\<exists>x. P(x) \<and> Q(x)) \<longrightarrow> (\<exists>x. P(x)) \<and> (\<exists>x. Q(x))"
|
51798
|
38 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
39 |
|
|
40 |
|
60770
|
41 |
text \<open>Basic test of quantifier reasoning\<close>
|
23914
|
42 |
|
62020
|
43 |
\<comment> \<open>TRUE\<close>
|
61489
|
44 |
lemma "(\<exists>y. \<forall>x. Q(x,y)) \<longrightarrow> (\<forall>x. \<exists>y. Q(x,y))"
|
51798
|
45 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
46 |
|
61489
|
47 |
lemma "(\<forall>x. Q(x)) \<longrightarrow> (\<exists>x. Q(x))"
|
51798
|
48 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
49 |
|
|
50 |
|
60770
|
51 |
text \<open>The following should fail, as they are false!\<close>
|
23914
|
52 |
|
61489
|
53 |
lemma "(\<forall>x. \<exists>y. Q(x,y)) \<longrightarrow> (\<exists>y. \<forall>x. Q(x,y))"
|
51798
|
54 |
apply (tactic "IntPr.fast_tac @{context} 1")?
|
23914
|
55 |
oops
|
|
56 |
|
61489
|
57 |
lemma "(\<exists>x. Q(x)) \<longrightarrow> (\<forall>x. Q(x))"
|
51798
|
58 |
apply (tactic "IntPr.fast_tac @{context} 1")?
|
23914
|
59 |
oops
|
|
60 |
|
61489
|
61 |
schematic_goal "P(?a) \<longrightarrow> (\<forall>x. P(x))"
|
51798
|
62 |
apply (tactic "IntPr.fast_tac @{context} 1")?
|
23914
|
63 |
oops
|
|
64 |
|
61489
|
65 |
schematic_goal "(P(?a) \<longrightarrow> (\<forall>x. Q(x))) \<longrightarrow> (\<forall>x. P(x) \<longrightarrow> Q(x))"
|
51798
|
66 |
apply (tactic "IntPr.fast_tac @{context} 1")?
|
23914
|
67 |
oops
|
|
68 |
|
|
69 |
|
60770
|
70 |
text \<open>Back to things that are provable \dots\<close>
|
23914
|
71 |
|
61489
|
72 |
lemma "(\<forall>x. P(x) \<longrightarrow> Q(x)) \<and> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x))"
|
51798
|
73 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
74 |
|
62020
|
75 |
\<comment> \<open>An example of why exI should be delayed as long as possible\<close>
|
61489
|
76 |
lemma "(P \<longrightarrow> (\<exists>x. Q(x))) \<and> P \<longrightarrow> (\<exists>x. Q(x))"
|
51798
|
77 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
78 |
|
61489
|
79 |
schematic_goal "(\<forall>x. P(x) \<longrightarrow> Q(f(x))) \<and> (\<forall>x. Q(x) \<longrightarrow> R(g(x))) \<and> P(d) \<longrightarrow> R(?a)"
|
51798
|
80 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
81 |
|
61489
|
82 |
lemma "(\<forall>x. Q(x)) \<longrightarrow> (\<exists>x. Q(x))"
|
51798
|
83 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
84 |
|
|
85 |
|
60770
|
86 |
text \<open>Some slow ones\<close>
|
23914
|
87 |
|
62020
|
88 |
\<comment> \<open>Principia Mathematica *11.53\<close>
|
61489
|
89 |
lemma "(\<forall>x y. P(x) \<longrightarrow> Q(y)) \<longleftrightarrow> ((\<exists>x. P(x)) \<longrightarrow> (\<forall>y. Q(y)))"
|
51798
|
90 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
91 |
|
|
92 |
(*Principia Mathematica *11.55 *)
|
61489
|
93 |
lemma "(\<exists>x y. P(x) \<and> Q(x,y)) \<longleftrightarrow> (\<exists>x. P(x) \<and> (\<exists>y. Q(x,y)))"
|
51798
|
94 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
95 |
|
|
96 |
(*Principia Mathematica *11.61 *)
|
61489
|
97 |
lemma "(\<exists>y. \<forall>x. P(x) \<longrightarrow> Q(x,y)) \<longrightarrow> (\<forall>x. P(x) \<longrightarrow> (\<exists>y. Q(x,y)))"
|
51798
|
98 |
by (tactic "IntPr.fast_tac @{context} 1")
|
23914
|
99 |
|
|
100 |
end
|