author | wenzelm |
Fri, 19 Jun 2020 16:12:32 +0200 | |
changeset 71960 | 6a64205b491a |
parent 69605 | a96320074298 |
child 74610 | 87fc10f5826c |
permissions | -rw-r--r-- |
39941 | 1 |
(* Title: HOL/Meson.thy |
39944 | 2 |
Author: Lawrence C. Paulson, Cambridge University Computer Laboratory |
3 |
Author: Tobias Nipkow, TU Muenchen |
|
4 |
Author: Jasmin Blanchette, TU Muenchen |
|
39941 | 5 |
Copyright 2001 University of Cambridge |
6 |
*) |
|
7 |
||
60758 | 8 |
section \<open>MESON Proof Method\<close> |
39941 | 9 |
|
10 |
theory Meson |
|
54553 | 11 |
imports Nat |
39941 | 12 |
begin |
13 |
||
60758 | 14 |
subsection \<open>Negation Normal Form\<close> |
39941 | 15 |
|
60758 | 16 |
text \<open>de Morgan laws\<close> |
39941 | 17 |
|
67091 | 18 |
lemma not_conjD: "\<not>(P\<and>Q) \<Longrightarrow> \<not>P \<or> \<not>Q" |
19 |
and not_disjD: "\<not>(P\<or>Q) \<Longrightarrow> \<not>P \<and> \<not>Q" |
|
20 |
and not_notD: "\<not>\<not>P \<Longrightarrow> P" |
|
21 |
and not_allD: "\<And>P. \<not>(\<forall>x. P(x)) \<Longrightarrow> \<exists>x. \<not>P(x)" |
|
22 |
and not_exD: "\<And>P. \<not>(\<exists>x. P(x)) \<Longrightarrow> \<forall>x. \<not>P(x)" |
|
39941 | 23 |
by fast+ |
24 |
||
61941 | 25 |
text \<open>Removal of \<open>\<longrightarrow>\<close> and \<open>\<longleftrightarrow>\<close> (positive and negative occurrences)\<close> |
39941 | 26 |
|
67091 | 27 |
lemma imp_to_disjD: "P\<longrightarrow>Q \<Longrightarrow> \<not>P \<or> Q" |
28 |
and not_impD: "\<not>(P\<longrightarrow>Q) \<Longrightarrow> P \<and> \<not>Q" |
|
29 |
and iff_to_disjD: "P=Q \<Longrightarrow> (\<not>P \<or> Q) \<and> (\<not>Q \<or> P)" |
|
30 |
and not_iffD: "\<not>(P=Q) \<Longrightarrow> (P \<or> Q) \<and> (\<not>P \<or> \<not>Q)" |
|
69593 | 31 |
\<comment> \<open>Much more efficient than \<^prop>\<open>(P \<and> \<not>Q) \<or> (Q \<and> \<not>P)\<close> for computing CNF\<close> |
67091 | 32 |
and not_refl_disj_D: "x \<noteq> x \<or> P \<Longrightarrow> P" |
39941 | 33 |
by fast+ |
34 |
||
35 |
||
60758 | 36 |
subsection \<open>Pulling out the existential quantifiers\<close> |
39941 | 37 |
|
60758 | 38 |
text \<open>Conjunction\<close> |
39941 | 39 |
|
67091 | 40 |
lemma conj_exD1: "\<And>P Q. (\<exists>x. P(x)) \<and> Q \<Longrightarrow> \<exists>x. P(x) \<and> Q" |
41 |
and conj_exD2: "\<And>P Q. P \<and> (\<exists>x. Q(x)) \<Longrightarrow> \<exists>x. P \<and> Q(x)" |
|
39941 | 42 |
by fast+ |
43 |
||
44 |
||
60758 | 45 |
text \<open>Disjunction\<close> |
39941 | 46 |
|
67091 | 47 |
lemma disj_exD: "\<And>P Q. (\<exists>x. P(x)) \<or> (\<exists>x. Q(x)) \<Longrightarrow> \<exists>x. P(x) \<or> Q(x)" |
61799 | 48 |
\<comment> \<open>DO NOT USE with forall-Skolemization: makes fewer schematic variables!!\<close> |
49 |
\<comment> \<open>With ex-Skolemization, makes fewer Skolem constants\<close> |
|
67091 | 50 |
and disj_exD1: "\<And>P Q. (\<exists>x. P(x)) \<or> Q \<Longrightarrow> \<exists>x. P(x) \<or> Q" |
51 |
and disj_exD2: "\<And>P Q. P \<or> (\<exists>x. Q(x)) \<Longrightarrow> \<exists>x. P \<or> Q(x)" |
|
39941 | 52 |
by fast+ |
53 |
||
67091 | 54 |
lemma disj_assoc: "(P\<or>Q)\<or>R \<Longrightarrow> P\<or>(Q\<or>R)" |
55 |
and disj_comm: "P\<or>Q \<Longrightarrow> Q\<or>P" |
|
56 |
and disj_FalseD1: "False\<or>P \<Longrightarrow> P" |
|
57 |
and disj_FalseD2: "P\<or>False \<Longrightarrow> P" |
|
39941 | 58 |
by fast+ |
59 |
||
60 |
||
60758 | 61 |
text\<open>Generation of contrapositives\<close> |
39941 | 62 |
|
60758 | 63 |
text\<open>Inserts negated disjunct after removing the negation; P is a literal. |
39941 | 64 |
Model elimination requires assuming the negation of every attempted subgoal, |
60758 | 65 |
hence the negated disjuncts.\<close> |
67091 | 66 |
lemma make_neg_rule: "\<not>P\<or>Q \<Longrightarrow> ((\<not>P\<Longrightarrow>P) \<Longrightarrow> Q)" |
39941 | 67 |
by blast |
68 |
||
60758 | 69 |
text\<open>Version for Plaisted's "Postive refinement" of the Meson procedure\<close> |
67091 | 70 |
lemma make_refined_neg_rule: "\<not>P\<or>Q \<Longrightarrow> (P \<Longrightarrow> Q)" |
39941 | 71 |
by blast |
72 |
||
69593 | 73 |
text\<open>\<^term>\<open>P\<close> should be a literal\<close> |
67091 | 74 |
lemma make_pos_rule: "P\<or>Q \<Longrightarrow> ((P\<Longrightarrow>\<not>P) \<Longrightarrow> Q)" |
39941 | 75 |
by blast |
76 |
||
61799 | 77 |
text\<open>Versions of \<open>make_neg_rule\<close> and \<open>make_pos_rule\<close> that don't |
60758 | 78 |
insert new assumptions, for ordinary resolution.\<close> |
39941 | 79 |
|
80 |
lemmas make_neg_rule' = make_refined_neg_rule |
|
81 |
||
67091 | 82 |
lemma make_pos_rule': "\<lbrakk>P\<or>Q; \<not>P\<rbrakk> \<Longrightarrow> Q" |
39941 | 83 |
by blast |
84 |
||
60758 | 85 |
text\<open>Generation of a goal clause -- put away the final literal\<close> |
39941 | 86 |
|
67091 | 87 |
lemma make_neg_goal: "\<not>P \<Longrightarrow> ((\<not>P\<Longrightarrow>P) \<Longrightarrow> False)" |
39941 | 88 |
by blast |
89 |
||
67091 | 90 |
lemma make_pos_goal: "P \<Longrightarrow> ((P\<Longrightarrow>\<not>P) \<Longrightarrow> False)" |
39941 | 91 |
by blast |
92 |
||
93 |
||
60758 | 94 |
subsection \<open>Lemmas for Forward Proof\<close> |
39941 | 95 |
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
61941
diff
changeset
|
96 |
text\<open>There is a similarity to congruence rules. They are also useful in ordinary proofs.\<close> |
39941 | 97 |
|
98 |
(*NOTE: could handle conjunctions (faster?) by |
|
99 |
nf(th RS conjunct2) RS (nf(th RS conjunct1) RS conjI) *) |
|
67091 | 100 |
lemma conj_forward: "\<lbrakk>P'\<and>Q'; P' \<Longrightarrow> P; Q' \<Longrightarrow> Q \<rbrakk> \<Longrightarrow> P\<and>Q" |
39941 | 101 |
by blast |
102 |
||
67091 | 103 |
lemma disj_forward: "\<lbrakk>P'\<or>Q'; P' \<Longrightarrow> P; Q' \<Longrightarrow> Q \<rbrakk> \<Longrightarrow> P\<or>Q" |
39941 | 104 |
by blast |
105 |
||
67091 | 106 |
lemma imp_forward: "\<lbrakk>P' \<longrightarrow> Q'; P \<Longrightarrow> P'; Q' \<Longrightarrow> Q \<rbrakk> \<Longrightarrow> P \<longrightarrow> Q" |
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
61941
diff
changeset
|
107 |
by blast |
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
61941
diff
changeset
|
108 |
|
69144
f13b82281715
new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
109 |
lemma imp_forward2: "\<lbrakk>P' \<longrightarrow> Q'; P \<Longrightarrow> P'; P' \<Longrightarrow> Q' \<Longrightarrow> Q \<rbrakk> \<Longrightarrow> P \<longrightarrow> Q" |
f13b82281715
new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
110 |
by blast |
f13b82281715
new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
111 |
|
39941 | 112 |
(*Version of @{text disj_forward} for removal of duplicate literals*) |
67091 | 113 |
lemma disj_forward2: "\<lbrakk> P'\<or>Q'; P' \<Longrightarrow> P; \<lbrakk>Q'; P\<Longrightarrow>False\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> P\<or>Q" |
39941 | 114 |
apply blast |
115 |
done |
|
116 |
||
117 |
lemma all_forward: "[| \<forall>x. P'(x); !!x. P'(x) ==> P(x) |] ==> \<forall>x. P(x)" |
|
118 |
by blast |
|
119 |
||
120 |
lemma ex_forward: "[| \<exists>x. P'(x); !!x. P'(x) ==> P(x) |] ==> \<exists>x. P(x)" |
|
121 |
by blast |
|
122 |
||
123 |
||
60758 | 124 |
subsection \<open>Clausification helper\<close> |
39941 | 125 |
|
126 |
lemma TruepropI: "P \<equiv> Q \<Longrightarrow> Trueprop P \<equiv> Trueprop Q" |
|
127 |
by simp |
|
128 |
||
47953
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
blanchet
parents:
42616
diff
changeset
|
129 |
lemma ext_cong_neq: "F g \<noteq> F h \<Longrightarrow> F g \<noteq> F h \<and> (\<exists>x. g x \<noteq> h x)" |
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
blanchet
parents:
42616
diff
changeset
|
130 |
apply (erule contrapos_np) |
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
blanchet
parents:
42616
diff
changeset
|
131 |
apply clarsimp |
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
blanchet
parents:
42616
diff
changeset
|
132 |
apply (rule cong[where f = F]) |
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
blanchet
parents:
42616
diff
changeset
|
133 |
by auto |
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
blanchet
parents:
42616
diff
changeset
|
134 |
|
39941 | 135 |
|
60758 | 136 |
text\<open>Combinator translation helpers\<close> |
39941 | 137 |
|
138 |
definition COMBI :: "'a \<Rightarrow> 'a" where |
|
54148 | 139 |
"COMBI P = P" |
39941 | 140 |
|
141 |
definition COMBK :: "'a \<Rightarrow> 'b \<Rightarrow> 'a" where |
|
54148 | 142 |
"COMBK P Q = P" |
39941 | 143 |
|
54148 | 144 |
definition COMBB :: "('b => 'c) \<Rightarrow> ('a => 'b) \<Rightarrow> 'a \<Rightarrow> 'c" where |
39941 | 145 |
"COMBB P Q R = P (Q R)" |
146 |
||
147 |
definition COMBC :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where |
|
54148 | 148 |
"COMBC P Q R = P R Q" |
39941 | 149 |
|
150 |
definition COMBS :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" where |
|
54148 | 151 |
"COMBS P Q R = P R (Q R)" |
39941 | 152 |
|
54148 | 153 |
lemma abs_S: "\<lambda>x. (f x) (g x) \<equiv> COMBS f g" |
39941 | 154 |
apply (rule eq_reflection) |
155 |
apply (rule ext) |
|
156 |
apply (simp add: COMBS_def) |
|
157 |
done |
|
158 |
||
54148 | 159 |
lemma abs_I: "\<lambda>x. x \<equiv> COMBI" |
39941 | 160 |
apply (rule eq_reflection) |
161 |
apply (rule ext) |
|
162 |
apply (simp add: COMBI_def) |
|
163 |
done |
|
164 |
||
54148 | 165 |
lemma abs_K: "\<lambda>x. y \<equiv> COMBK y" |
39941 | 166 |
apply (rule eq_reflection) |
167 |
apply (rule ext) |
|
168 |
apply (simp add: COMBK_def) |
|
169 |
done |
|
170 |
||
54148 | 171 |
lemma abs_B: "\<lambda>x. a (g x) \<equiv> COMBB a g" |
39941 | 172 |
apply (rule eq_reflection) |
173 |
apply (rule ext) |
|
174 |
apply (simp add: COMBB_def) |
|
175 |
done |
|
176 |
||
54148 | 177 |
lemma abs_C: "\<lambda>x. (f x) b \<equiv> COMBC f b" |
39941 | 178 |
apply (rule eq_reflection) |
179 |
apply (rule ext) |
|
180 |
apply (simp add: COMBC_def) |
|
181 |
done |
|
182 |
||
183 |
||
60758 | 184 |
subsection \<open>Skolemization helpers\<close> |
39941 | 185 |
|
186 |
definition skolem :: "'a \<Rightarrow> 'a" where |
|
54148 | 187 |
"skolem = (\<lambda>x. x)" |
39941 | 188 |
|
61076 | 189 |
lemma skolem_COMBK_iff: "P \<longleftrightarrow> skolem (COMBK P (i::nat))" |
39941 | 190 |
unfolding skolem_def COMBK_def by (rule refl) |
191 |
||
192 |
lemmas skolem_COMBK_I = iffD1 [OF skolem_COMBK_iff] |
|
193 |
lemmas skolem_COMBK_D = iffD2 [OF skolem_COMBK_iff] |
|
194 |
||
195 |
||
60758 | 196 |
subsection \<open>Meson package\<close> |
39941 | 197 |
|
69605 | 198 |
ML_file \<open>Tools/Meson/meson.ML\<close> |
199 |
ML_file \<open>Tools/Meson/meson_clausify.ML\<close> |
|
200 |
ML_file \<open>Tools/Meson/meson_tactic.ML\<close> |
|
39941 | 201 |
|
39953
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents:
39950
diff
changeset
|
202 |
hide_const (open) COMBI COMBK COMBB COMBC COMBS skolem |
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents:
39950
diff
changeset
|
203 |
hide_fact (open) not_conjD not_disjD not_notD not_allD not_exD imp_to_disjD |
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents:
39950
diff
changeset
|
204 |
not_impD iff_to_disjD not_iffD not_refl_disj_D conj_exD1 conj_exD2 disj_exD |
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents:
39950
diff
changeset
|
205 |
disj_exD1 disj_exD2 disj_assoc disj_comm disj_FalseD1 disj_FalseD2 TruepropI |
47953
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
blanchet
parents:
42616
diff
changeset
|
206 |
ext_cong_neq COMBI_def COMBK_def COMBB_def COMBC_def COMBS_def abs_I abs_K |
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
blanchet
parents:
42616
diff
changeset
|
207 |
abs_B abs_C abs_S skolem_def skolem_COMBK_iff skolem_COMBK_I skolem_COMBK_D |
39953
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents:
39950
diff
changeset
|
208 |
|
39941 | 209 |
end |