src/HOL/Hilbert_Choice.thy
changeset 39943 0ef551d47783
parent 39900 549c00e0e89b
child 39950 f3c4849868b8
--- a/src/HOL/Hilbert_Choice.thy	Mon Oct 04 21:50:32 2010 +0200
+++ b/src/HOL/Hilbert_Choice.thy	Mon Oct 04 21:55:54 2010 +0200
@@ -7,8 +7,7 @@
 
 theory Hilbert_Choice
 imports Nat Wellfounded Plain
-uses ("Tools/meson.ML")
-     ("Tools/choice_specification.ML")
+uses ("Tools/choice_specification.ML")
 begin
 
 subsection {* Hilbert's epsilon *}
@@ -81,16 +80,6 @@
 
 subsection{*Axiom of Choice, Proved Using the Description Operator*}
 
-ML {*
-structure Meson_Choices = Named_Thms
-(
-  val name = "meson_choice"
-  val description = "choice axioms for MESON's (and Metis's) skolemizer"
-)
-*}
-
-setup Meson_Choices.setup
-
 lemma choice [meson_choice]: "\<forall>x. \<exists>y. Q x y ==> \<exists>f. \<forall>x. Q x (f x)"
 by (fast elim: someI)
 
@@ -451,128 +440,6 @@
   done
 
 
-subsection {* The Meson proof procedure *}
-
-subsubsection {* Negation Normal Form *}
-
-text {* de Morgan laws *}
-
-lemma meson_not_conjD: "~(P&Q) ==> ~P | ~Q"
-  and meson_not_disjD: "~(P|Q) ==> ~P & ~Q"
-  and meson_not_notD: "~~P ==> P"
-  and meson_not_allD: "!!P. ~(\<forall>x. P(x)) ==> \<exists>x. ~P(x)"
-  and meson_not_exD: "!!P. ~(\<exists>x. P(x)) ==> \<forall>x. ~P(x)"
-  by fast+
-
-text {* Removal of @{text "-->"} and @{text "<->"} (positive and
-negative occurrences) *}
-
-lemma meson_imp_to_disjD: "P-->Q ==> ~P | Q"
-  and meson_not_impD: "~(P-->Q) ==> P & ~Q"
-  and meson_iff_to_disjD: "P=Q ==> (~P | Q) & (~Q | P)"
-  and meson_not_iffD: "~(P=Q) ==> (P | Q) & (~P | ~Q)"
-    -- {* Much more efficient than @{prop "(P & ~Q) | (Q & ~P)"} for computing CNF *}
-  and meson_not_refl_disj_D: "x ~= x | P ==> P"
-  by fast+
-
-
-subsubsection {* Pulling out the existential quantifiers *}
-
-text {* Conjunction *}
-
-lemma meson_conj_exD1: "!!P Q. (\<exists>x. P(x)) & Q ==> \<exists>x. P(x) & Q"
-  and meson_conj_exD2: "!!P Q. P & (\<exists>x. Q(x)) ==> \<exists>x. P & Q(x)"
-  by fast+
-
-
-text {* Disjunction *}
-
-lemma meson_disj_exD: "!!P Q. (\<exists>x. P(x)) | (\<exists>x. Q(x)) ==> \<exists>x. P(x) | Q(x)"
-  -- {* DO NOT USE with forall-Skolemization: makes fewer schematic variables!! *}
-  -- {* With ex-Skolemization, makes fewer Skolem constants *}
-  and meson_disj_exD1: "!!P Q. (\<exists>x. P(x)) | Q ==> \<exists>x. P(x) | Q"
-  and meson_disj_exD2: "!!P Q. P | (\<exists>x. Q(x)) ==> \<exists>x. P | Q(x)"
-  by fast+
-
-
-subsubsection {* Generating clauses for the Meson Proof Procedure *}
-
-text {* Disjunctions *}
-
-lemma meson_disj_assoc: "(P|Q)|R ==> P|(Q|R)"
-  and meson_disj_comm: "P|Q ==> Q|P"
-  and meson_disj_FalseD1: "False|P ==> P"
-  and meson_disj_FalseD2: "P|False ==> P"
-  by fast+
-
-
-subsection{*Lemmas for Meson, the Model Elimination Procedure*}
-
-text{* Generation of contrapositives *}
-
-text{*Inserts negated disjunct after removing the negation; P is a literal.
-  Model elimination requires assuming the negation of every attempted subgoal,
-  hence the negated disjuncts.*}
-lemma make_neg_rule: "~P|Q ==> ((~P==>P) ==> Q)"
-by blast
-
-text{*Version for Plaisted's "Postive refinement" of the Meson procedure*}
-lemma make_refined_neg_rule: "~P|Q ==> (P ==> Q)"
-by blast
-
-text{*@{term P} should be a literal*}
-lemma make_pos_rule: "P|Q ==> ((P==>~P) ==> Q)"
-by blast
-
-text{*Versions of @{text make_neg_rule} and @{text make_pos_rule} that don't
-insert new assumptions, for ordinary resolution.*}
-
-lemmas make_neg_rule' = make_refined_neg_rule
-
-lemma make_pos_rule': "[|P|Q; ~P|] ==> Q"
-by blast
-
-text{* Generation of a goal clause -- put away the final literal *}
-
-lemma make_neg_goal: "~P ==> ((~P==>P) ==> False)"
-by blast
-
-lemma make_pos_goal: "P ==> ((P==>~P) ==> False)"
-by blast
-
-
-subsubsection{* Lemmas for Forward Proof*}
-
-text{*There is a similarity to congruence rules*}
-
-(*NOTE: could handle conjunctions (faster?) by
-    nf(th RS conjunct2) RS (nf(th RS conjunct1) RS conjI) *)
-lemma conj_forward: "[| P'&Q';  P' ==> P;  Q' ==> Q |] ==> P&Q"
-by blast
-
-lemma disj_forward: "[| P'|Q';  P' ==> P;  Q' ==> Q |] ==> P|Q"
-by blast
-
-(*Version of @{text disj_forward} for removal of duplicate literals*)
-lemma disj_forward2:
-    "[| P'|Q';  P' ==> P;  [| Q'; P==>False |] ==> Q |] ==> P|Q"
-apply blast 
-done
-
-lemma all_forward: "[| \<forall>x. P'(x);  !!x. P'(x) ==> P(x) |] ==> \<forall>x. P(x)"
-by blast
-
-lemma ex_forward: "[| \<exists>x. P'(x);  !!x. P'(x) ==> P(x) |] ==> \<exists>x. P(x)"
-by blast
-
-
-subsection {* Meson package *}
-
-use "Tools/meson.ML"
-
-setup Meson.setup
-
-
 subsection {* Specification package -- Hilbertized version *}
 
 lemma exE_some: "[| Ex P ; c == Eps P |] ==> P c"
@@ -580,5 +447,4 @@
 
 use "Tools/choice_specification.ML"
 
-
 end