author | wenzelm |
Mon, 12 Oct 2015 19:47:23 +0200 | |
changeset 61410 | f569907de061 |
parent 60770 | 240563fbf41d |
child 61489 | b8d375aee0df |
permissions | -rw-r--r-- |
19820 | 1 |
(* Title: FOL/ex/Miniscope.thy |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
3 |
Copyright 1994 University of Cambridge |
|
4 |
||
5 |
Classical First-Order Logic. |
|
6 |
Conversion to nnf/miniscope format: pushing quantifiers in. |
|
7 |
Demonstration of formula rewriting by proof. |
|
8 |
*) |
|
9 |
||
10 |
theory Miniscope |
|
11 |
imports FOL |
|
12 |
begin |
|
13 |
||
14 |
||
15 |
lemmas ccontr = FalseE [THEN classical] |
|
16 |
||
60770 | 17 |
subsection \<open>Negation Normal Form\<close> |
19820 | 18 |
|
60770 | 19 |
subsubsection \<open>de Morgan laws\<close> |
19820 | 20 |
|
21 |
lemma demorgans: |
|
22 |
"~(P&Q) <-> ~P | ~Q" |
|
23 |
"~(P|Q) <-> ~P & ~Q" |
|
24 |
"~~P <-> P" |
|
25 |
"!!P. ~(ALL x. P(x)) <-> (EX x. ~P(x))" |
|
26 |
"!!P. ~(EX x. P(x)) <-> (ALL x. ~P(x))" |
|
27 |
by blast+ |
|
28 |
||
29 |
(*** Removal of --> and <-> (positive and negative occurrences) ***) |
|
30 |
(*Last one is important for computing a compact CNF*) |
|
31 |
lemma nnf_simps: |
|
32 |
"(P-->Q) <-> (~P | Q)" |
|
33 |
"~(P-->Q) <-> (P & ~Q)" |
|
34 |
"(P<->Q) <-> (~P | Q) & (~Q | P)" |
|
35 |
"~(P<->Q) <-> (P | Q) & (~P | ~Q)" |
|
36 |
by blast+ |
|
37 |
||
38 |
||
39 |
(* BEWARE: rewrite rules for <-> can confuse the simplifier!! *) |
|
40 |
||
60770 | 41 |
subsubsection \<open>Pushing in the existential quantifiers\<close> |
19820 | 42 |
|
43 |
lemma ex_simps: |
|
44 |
"(EX x. P) <-> P" |
|
45 |
"!!P Q. (EX x. P(x) & Q) <-> (EX x. P(x)) & Q" |
|
46 |
"!!P Q. (EX x. P & Q(x)) <-> P & (EX x. Q(x))" |
|
47 |
"!!P Q. (EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))" |
|
48 |
"!!P Q. (EX x. P(x) | Q) <-> (EX x. P(x)) | Q" |
|
49 |
"!!P Q. (EX x. P | Q(x)) <-> P | (EX x. Q(x))" |
|
50 |
by blast+ |
|
51 |
||
52 |
||
60770 | 53 |
subsubsection \<open>Pushing in the universal quantifiers\<close> |
19820 | 54 |
|
55 |
lemma all_simps: |
|
56 |
"(ALL x. P) <-> P" |
|
57 |
"!!P Q. (ALL x. P(x) & Q(x)) <-> (ALL x. P(x)) & (ALL x. Q(x))" |
|
58 |
"!!P Q. (ALL x. P(x) & Q) <-> (ALL x. P(x)) & Q" |
|
59 |
"!!P Q. (ALL x. P & Q(x)) <-> P & (ALL x. Q(x))" |
|
60 |
"!!P Q. (ALL x. P(x) | Q) <-> (ALL x. P(x)) | Q" |
|
61 |
"!!P Q. (ALL x. P | Q(x)) <-> P | (ALL x. Q(x))" |
|
62 |
by blast+ |
|
63 |
||
64 |
lemmas mini_simps = demorgans nnf_simps ex_simps all_simps |
|
65 |
||
60770 | 66 |
ML \<open> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
31974
diff
changeset
|
67 |
val mini_ss = simpset_of (@{context} addsimps @{thms mini_simps}); |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58838
diff
changeset
|
68 |
fun mini_tac ctxt = |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58838
diff
changeset
|
69 |
resolve_tac ctxt @{thms ccontr} THEN' asm_full_simp_tac (put_simpset mini_ss ctxt); |
60770 | 70 |
\<close> |
19820 | 71 |
|
72 |
end |