blanchet@39941: (* Title: HOL/Meson.thy blanchet@39944: Author: Lawrence C. Paulson, Cambridge University Computer Laboratory blanchet@39944: Author: Tobias Nipkow, TU Muenchen blanchet@39944: Author: Jasmin Blanchette, TU Muenchen blanchet@39941: Copyright 2001 University of Cambridge blanchet@39941: *) blanchet@39941: blanchet@39947: header {* MESON Proof Method *} blanchet@39941: blanchet@39941: theory Meson blanchet@39946: imports Datatype blanchet@39941: uses ("Tools/Meson/meson.ML") blanchet@39941: ("Tools/Meson/meson_clausify.ML") blanchet@39948: ("Tools/Meson/meson_tactic.ML") blanchet@39941: begin blanchet@39941: blanchet@39941: section {* Negation Normal Form *} blanchet@39941: blanchet@39941: text {* de Morgan laws *} blanchet@39941: blanchet@39953: lemma not_conjD: "~(P&Q) ==> ~P | ~Q" blanchet@39953: and not_disjD: "~(P|Q) ==> ~P & ~Q" blanchet@39953: and not_notD: "~~P ==> P" blanchet@39953: and not_allD: "!!P. ~(\x. P(x)) ==> \x. ~P(x)" blanchet@39953: and not_exD: "!!P. ~(\x. P(x)) ==> \x. ~P(x)" blanchet@39941: by fast+ blanchet@39941: blanchet@39941: text {* Removal of @{text "-->"} and @{text "<->"} (positive and blanchet@39941: negative occurrences) *} blanchet@39941: blanchet@39953: lemma imp_to_disjD: "P-->Q ==> ~P | Q" blanchet@39953: and not_impD: "~(P-->Q) ==> P & ~Q" blanchet@39953: and iff_to_disjD: "P=Q ==> (~P | Q) & (~Q | P)" blanchet@39953: and not_iffD: "~(P=Q) ==> (P | Q) & (~P | ~Q)" blanchet@39941: -- {* Much more efficient than @{prop "(P & ~Q) | (Q & ~P)"} for computing CNF *} blanchet@39953: and not_refl_disj_D: "x ~= x | P ==> P" blanchet@39941: by fast+ blanchet@39941: blanchet@39941: blanchet@39941: section {* Pulling out the existential quantifiers *} blanchet@39941: blanchet@39941: text {* Conjunction *} blanchet@39941: blanchet@39953: lemma conj_exD1: "!!P Q. (\x. P(x)) & Q ==> \x. P(x) & Q" blanchet@39953: and conj_exD2: "!!P Q. P & (\x. Q(x)) ==> \x. P & Q(x)" blanchet@39941: by fast+ blanchet@39941: blanchet@39941: blanchet@39941: text {* Disjunction *} blanchet@39941: blanchet@39953: lemma disj_exD: "!!P Q. (\x. P(x)) | (\x. Q(x)) ==> \x. P(x) | Q(x)" blanchet@39941: -- {* DO NOT USE with forall-Skolemization: makes fewer schematic variables!! *} blanchet@39941: -- {* With ex-Skolemization, makes fewer Skolem constants *} blanchet@39953: and disj_exD1: "!!P Q. (\x. P(x)) | Q ==> \x. P(x) | Q" blanchet@39953: and disj_exD2: "!!P Q. P | (\x. Q(x)) ==> \x. P | Q(x)" blanchet@39941: by fast+ blanchet@39941: blanchet@39953: lemma disj_assoc: "(P|Q)|R ==> P|(Q|R)" blanchet@39953: and disj_comm: "P|Q ==> Q|P" blanchet@39953: and disj_FalseD1: "False|P ==> P" blanchet@39953: and disj_FalseD2: "P|False ==> P" blanchet@39941: by fast+ blanchet@39941: blanchet@39941: blanchet@39941: text{* Generation of contrapositives *} blanchet@39941: blanchet@39941: text{*Inserts negated disjunct after removing the negation; P is a literal. blanchet@39941: Model elimination requires assuming the negation of every attempted subgoal, blanchet@39941: hence the negated disjuncts.*} blanchet@39941: lemma make_neg_rule: "~P|Q ==> ((~P==>P) ==> Q)" blanchet@39941: by blast blanchet@39941: blanchet@39941: text{*Version for Plaisted's "Postive refinement" of the Meson procedure*} blanchet@39941: lemma make_refined_neg_rule: "~P|Q ==> (P ==> Q)" blanchet@39941: by blast blanchet@39941: blanchet@39941: text{*@{term P} should be a literal*} blanchet@39941: lemma make_pos_rule: "P|Q ==> ((P==>~P) ==> Q)" blanchet@39941: by blast blanchet@39941: blanchet@39941: text{*Versions of @{text make_neg_rule} and @{text make_pos_rule} that don't blanchet@39941: insert new assumptions, for ordinary resolution.*} blanchet@39941: blanchet@39941: lemmas make_neg_rule' = make_refined_neg_rule blanchet@39941: blanchet@39941: lemma make_pos_rule': "[|P|Q; ~P|] ==> Q" blanchet@39941: by blast blanchet@39941: blanchet@39941: text{* Generation of a goal clause -- put away the final literal *} blanchet@39941: blanchet@39941: lemma make_neg_goal: "~P ==> ((~P==>P) ==> False)" blanchet@39941: by blast blanchet@39941: blanchet@39941: lemma make_pos_goal: "P ==> ((P==>~P) ==> False)" blanchet@39941: by blast blanchet@39941: blanchet@39941: blanchet@39941: section {* Lemmas for Forward Proof *} blanchet@39941: blanchet@39941: text{*There is a similarity to congruence rules*} blanchet@39941: blanchet@39941: (*NOTE: could handle conjunctions (faster?) by blanchet@39941: nf(th RS conjunct2) RS (nf(th RS conjunct1) RS conjI) *) blanchet@39941: lemma conj_forward: "[| P'&Q'; P' ==> P; Q' ==> Q |] ==> P&Q" blanchet@39941: by blast blanchet@39941: blanchet@39941: lemma disj_forward: "[| P'|Q'; P' ==> P; Q' ==> Q |] ==> P|Q" blanchet@39941: by blast blanchet@39941: blanchet@39941: (*Version of @{text disj_forward} for removal of duplicate literals*) blanchet@39941: lemma disj_forward2: blanchet@39941: "[| P'|Q'; P' ==> P; [| Q'; P==>False |] ==> Q |] ==> P|Q" blanchet@39941: apply blast blanchet@39941: done blanchet@39941: blanchet@39941: lemma all_forward: "[| \x. P'(x); !!x. P'(x) ==> P(x) |] ==> \x. P(x)" blanchet@39941: by blast blanchet@39941: blanchet@39941: lemma ex_forward: "[| \x. P'(x); !!x. P'(x) ==> P(x) |] ==> \x. P(x)" blanchet@39941: by blast blanchet@39941: blanchet@39941: blanchet@39941: section {* Clausification helper *} blanchet@39941: blanchet@39941: lemma TruepropI: "P \ Q \ Trueprop P \ Trueprop Q" blanchet@39941: by simp blanchet@39941: blanchet@39941: blanchet@39941: text{* Combinator translation helpers *} blanchet@39941: blanchet@39941: definition COMBI :: "'a \ 'a" where blanchet@39941: [no_atp]: "COMBI P = P" blanchet@39941: blanchet@39941: definition COMBK :: "'a \ 'b \ 'a" where blanchet@39941: [no_atp]: "COMBK P Q = P" blanchet@39941: blanchet@39941: definition COMBB :: "('b => 'c) \ ('a => 'b) \ 'a \ 'c" where [no_atp]: blanchet@39941: "COMBB P Q R = P (Q R)" blanchet@39941: blanchet@39941: definition COMBC :: "('a \ 'b \ 'c) \ 'b \ 'a \ 'c" where blanchet@39941: [no_atp]: "COMBC P Q R = P R Q" blanchet@39941: blanchet@39941: definition COMBS :: "('a \ 'b \ 'c) \ ('a \ 'b) \ 'a \ 'c" where blanchet@39941: [no_atp]: "COMBS P Q R = P R (Q R)" blanchet@39941: blanchet@39941: lemma abs_S [no_atp]: "\x. (f x) (g x) \ COMBS f g" blanchet@39941: apply (rule eq_reflection) blanchet@39941: apply (rule ext) blanchet@39941: apply (simp add: COMBS_def) blanchet@39941: done blanchet@39941: blanchet@39941: lemma abs_I [no_atp]: "\x. x \ COMBI" blanchet@39941: apply (rule eq_reflection) blanchet@39941: apply (rule ext) blanchet@39941: apply (simp add: COMBI_def) blanchet@39941: done blanchet@39941: blanchet@39941: lemma abs_K [no_atp]: "\x. y \ COMBK y" blanchet@39941: apply (rule eq_reflection) blanchet@39941: apply (rule ext) blanchet@39941: apply (simp add: COMBK_def) blanchet@39941: done blanchet@39941: blanchet@39941: lemma abs_B [no_atp]: "\x. a (g x) \ COMBB a g" blanchet@39941: apply (rule eq_reflection) blanchet@39941: apply (rule ext) blanchet@39941: apply (simp add: COMBB_def) blanchet@39941: done blanchet@39941: blanchet@39941: lemma abs_C [no_atp]: "\x. (f x) b \ COMBC f b" blanchet@39941: apply (rule eq_reflection) blanchet@39941: apply (rule ext) blanchet@39941: apply (simp add: COMBC_def) blanchet@39941: done blanchet@39941: blanchet@39941: blanchet@39941: section {* Skolemization helpers *} blanchet@39941: blanchet@39941: definition skolem :: "'a \ 'a" where blanchet@39941: [no_atp]: "skolem = (\x. x)" blanchet@39941: blanchet@39941: lemma skolem_COMBK_iff: "P \ skolem (COMBK P (i\nat))" blanchet@39941: unfolding skolem_def COMBK_def by (rule refl) blanchet@39941: blanchet@39941: lemmas skolem_COMBK_I = iffD1 [OF skolem_COMBK_iff] blanchet@39941: lemmas skolem_COMBK_D = iffD2 [OF skolem_COMBK_iff] blanchet@39941: blanchet@39941: blanchet@39941: section {* Meson package *} blanchet@39941: blanchet@39941: use "Tools/Meson/meson.ML" blanchet@39941: use "Tools/Meson/meson_clausify.ML" blanchet@39948: use "Tools/Meson/meson_tactic.ML" blanchet@39941: blanchet@39941: setup {* blanchet@39950: Meson.setup blanchet@39948: #> Meson_Tactic.setup blanchet@39941: *} blanchet@39941: blanchet@39953: hide_const (open) COMBI COMBK COMBB COMBC COMBS skolem blanchet@39953: hide_fact (open) not_conjD not_disjD not_notD not_allD not_exD imp_to_disjD blanchet@39953: not_impD iff_to_disjD not_iffD not_refl_disj_D conj_exD1 conj_exD2 disj_exD blanchet@39953: disj_exD1 disj_exD2 disj_assoc disj_comm disj_FalseD1 disj_FalseD2 TruepropI blanchet@39953: COMBI_def COMBK_def COMBB_def COMBC_def COMBS_def abs_I abs_K abs_B abs_C blanchet@39953: abs_S skolem_def skolem_COMBK_iff skolem_COMBK_I skolem_COMBK_D blanchet@39953: blanchet@39941: end