author | wenzelm |
Sat, 23 May 2015 17:19:37 +0200 | |
changeset 60299 | 5ae2a2e74c93 |
parent 58963 | 26bf09b95dda |
child 60770 | 240563fbf41d |
permissions | -rw-r--r-- |
26408 | 1 |
(* Title: FOLP/ex/Quantifiers_Int.thy |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
3 |
Copyright 1991 University of Cambridge |
|
4 |
||
5 |
First-Order Logic: quantifier examples (intuitionistic and classical) |
|
35762 | 6 |
Needs declarations of the theory "thy" and the tactic "tac". |
26408 | 7 |
*) |
8 |
||
9 |
theory Quantifiers_Int |
|
10 |
imports IFOLP |
|
11 |
begin |
|
12 |
||
36319 | 13 |
schematic_lemma "?p : (ALL x y. P(x,y)) --> (ALL y x. P(x,y))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
14 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 15 |
|
36319 | 16 |
schematic_lemma "?p : (EX x y. P(x,y)) --> (EX y x. P(x,y))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
17 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 18 |
|
19 |
||
20 |
(*Converse is false*) |
|
36319 | 21 |
schematic_lemma "?p : (ALL x. P(x)) | (ALL x. Q(x)) --> (ALL x. P(x) | Q(x))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
22 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 23 |
|
36319 | 24 |
schematic_lemma "?p : (ALL x. P-->Q(x)) <-> (P--> (ALL x. Q(x)))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
25 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 26 |
|
27 |
||
36319 | 28 |
schematic_lemma "?p : (ALL x. P(x)-->Q) <-> ((EX x. P(x)) --> Q)" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
29 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 30 |
|
31 |
||
32 |
text "Some harder ones" |
|
33 |
||
36319 | 34 |
schematic_lemma "?p : (EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
35 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 36 |
|
37 |
(*Converse is false*) |
|
36319 | 38 |
schematic_lemma "?p : (EX x. P(x)&Q(x)) --> (EX x. P(x)) & (EX x. Q(x))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
39 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 40 |
|
41 |
||
42 |
text "Basic test of quantifier reasoning" |
|
43 |
(*TRUE*) |
|
36319 | 44 |
schematic_lemma "?p : (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
45 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 46 |
|
36319 | 47 |
schematic_lemma "?p : (ALL x. Q(x)) --> (EX x. Q(x))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
48 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 49 |
|
50 |
||
51 |
text "The following should fail, as they are false!" |
|
52 |
||
36319 | 53 |
schematic_lemma "?p : (ALL x. EX y. Q(x,y)) --> (EX y. ALL x. Q(x,y))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
54 |
apply (tactic {* IntPr.fast_tac @{context} 1 *})? |
26408 | 55 |
oops |
56 |
||
36319 | 57 |
schematic_lemma "?p : (EX x. Q(x)) --> (ALL x. Q(x))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
58 |
apply (tactic {* IntPr.fast_tac @{context} 1 *})? |
26408 | 59 |
oops |
60 |
||
36319 | 61 |
schematic_lemma "?p : P(?a) --> (ALL x. P(x))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
62 |
apply (tactic {* IntPr.fast_tac @{context} 1 *})? |
26408 | 63 |
oops |
64 |
||
36319 | 65 |
schematic_lemma "?p : (P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
66 |
apply (tactic {* IntPr.fast_tac @{context} 1 *})? |
26408 | 67 |
oops |
68 |
||
69 |
||
70 |
text "Back to things that are provable..." |
|
71 |
||
36319 | 72 |
schematic_lemma "?p : (ALL x. P(x)-->Q(x)) & (EX x. P(x)) --> (EX x. Q(x))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
73 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 74 |
|
75 |
||
76 |
(*An example of why exI should be delayed as long as possible*) |
|
36319 | 77 |
schematic_lemma "?p : (P --> (EX x. Q(x))) & P --> (EX x. Q(x))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
78 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 79 |
|
36319 | 80 |
schematic_lemma "?p : (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
81 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 82 |
|
36319 | 83 |
schematic_lemma "?p : (ALL x. Q(x)) --> (EX x. Q(x))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
84 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 85 |
|
86 |
||
87 |
text "Some slow ones" |
|
88 |
||
89 |
(*Principia Mathematica *11.53 *) |
|
36319 | 90 |
schematic_lemma "?p : (ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
91 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 92 |
|
93 |
(*Principia Mathematica *11.55 *) |
|
36319 | 94 |
schematic_lemma "?p : (EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
95 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 96 |
|
97 |
(*Principia Mathematica *11.61 *) |
|
36319 | 98 |
schematic_lemma "?p : (EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))" |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
36319
diff
changeset
|
99 |
by (tactic {* IntPr.fast_tac @{context} 1 *}) |
26408 | 100 |
|
101 |
end |