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