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 |
|
|
17 |
subsection {* Negation Normal Form *}
|
|
18 |
|
|
19 |
subsubsection {* de Morgan laws *}
|
|
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 |
|
|
41 |
subsubsection {* Pushing in the existential quantifiers *}
|
|
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 |
|
|
53 |
subsubsection {* Pushing in the universal quantifiers *}
|
|
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 |
|
|
66 |
ML {*
|
26342
|
67 |
val mini_ss = @{simpset} addsimps @{thms mini_simps};
|
|
68 |
val mini_tac = rtac @{thm ccontr} THEN' asm_full_simp_tac mini_ss;
|
19820
|
69 |
*}
|
|
70 |
|
|
71 |
end
|