author | wenzelm |
Thu, 28 Dec 2017 12:13:56 +0100 | |
changeset 67284 | 0094d938c53b |
parent 61489 | b8d375aee0df |
child 68536 | e14848001c4c |
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 |
lemmas ccontr = FalseE [THEN classical] |
|
15 |
||
60770 | 16 |
subsection \<open>Negation Normal Form\<close> |
19820 | 17 |
|
60770 | 18 |
subsubsection \<open>de Morgan laws\<close> |
19820 | 19 |
|
20 |
lemma demorgans: |
|
61489 | 21 |
"\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q" |
22 |
"\<not> (P \<or> Q) \<longleftrightarrow> \<not> P \<and> \<not> Q" |
|
23 |
"\<not> \<not> P \<longleftrightarrow> P" |
|
24 |
"\<And>P. \<not> (\<forall>x. P(x)) \<longleftrightarrow> (\<exists>x. \<not> P(x))" |
|
25 |
"\<And>P. \<not> (\<exists>x. P(x)) \<longleftrightarrow> (\<forall>x. \<not> P(x))" |
|
19820 | 26 |
by blast+ |
27 |
||
28 |
(*** Removal of --> and <-> (positive and negative occurrences) ***) |
|
29 |
(*Last one is important for computing a compact CNF*) |
|
30 |
lemma nnf_simps: |
|
61489 | 31 |
"(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)" |
32 |
"\<not> (P \<longrightarrow> Q) \<longleftrightarrow> (P \<and> \<not> Q)" |
|
33 |
"(P \<longleftrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q) \<and> (\<not> Q \<or> P)" |
|
34 |
"\<not> (P \<longleftrightarrow> Q) \<longleftrightarrow> (P \<or> Q) \<and> (\<not> P \<or> \<not> Q)" |
|
19820 | 35 |
by blast+ |
36 |
||
37 |
||
38 |
(* BEWARE: rewrite rules for <-> can confuse the simplifier!! *) |
|
39 |
||
60770 | 40 |
subsubsection \<open>Pushing in the existential quantifiers\<close> |
19820 | 41 |
|
42 |
lemma ex_simps: |
|
61489 | 43 |
"(\<exists>x. P) \<longleftrightarrow> P" |
44 |
"\<And>P Q. (\<exists>x. P(x) \<and> Q) \<longleftrightarrow> (\<exists>x. P(x)) \<and> Q" |
|
45 |
"\<And>P Q. (\<exists>x. P \<and> Q(x)) \<longleftrightarrow> P \<and> (\<exists>x. Q(x))" |
|
46 |
"\<And>P Q. (\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> (\<exists>x. P(x)) \<or> (\<exists>x. Q(x))" |
|
47 |
"\<And>P Q. (\<exists>x. P(x) \<or> Q) \<longleftrightarrow> (\<exists>x. P(x)) \<or> Q" |
|
48 |
"\<And>P Q. (\<exists>x. P \<or> Q(x)) \<longleftrightarrow> P \<or> (\<exists>x. Q(x))" |
|
19820 | 49 |
by blast+ |
50 |
||
51 |
||
60770 | 52 |
subsubsection \<open>Pushing in the universal quantifiers\<close> |
19820 | 53 |
|
54 |
lemma all_simps: |
|
61489 | 55 |
"(\<forall>x. P) \<longleftrightarrow> P" |
56 |
"\<And>P Q. (\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> (\<forall>x. P(x)) \<and> (\<forall>x. Q(x))" |
|
57 |
"\<And>P Q. (\<forall>x. P(x) \<and> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<and> Q" |
|
58 |
"\<And>P Q. (\<forall>x. P \<and> Q(x)) \<longleftrightarrow> P \<and> (\<forall>x. Q(x))" |
|
59 |
"\<And>P Q. (\<forall>x. P(x) \<or> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<or> Q" |
|
60 |
"\<And>P Q. (\<forall>x. P \<or> Q(x)) \<longleftrightarrow> P \<or> (\<forall>x. Q(x))" |
|
19820 | 61 |
by blast+ |
62 |
||
63 |
lemmas mini_simps = demorgans nnf_simps ex_simps all_simps |
|
64 |
||
60770 | 65 |
ML \<open> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
31974
diff
changeset
|
66 |
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
|
67 |
fun mini_tac ctxt = |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58838
diff
changeset
|
68 |
resolve_tac ctxt @{thms ccontr} THEN' asm_full_simp_tac (put_simpset mini_ss ctxt); |
60770 | 69 |
\<close> |
19820 | 70 |
|
71 |
end |