# HG changeset patch # User wenzelm # Date 1006796037 -3600 # Node ID b344486c33e2800fe4fd8bd55708d11fd82bb5a0 # Parent 2ce7b42b0a64c19b994ba1247a08750ba020a26e tuned; meson lemmas from Tools/meson.ML; diff -r 2ce7b42b0a64 -r b344486c33e2 src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy Mon Nov 26 18:33:21 2001 +0100 +++ b/src/HOL/Hilbert_Choice.thy Mon Nov 26 18:33:57 2001 +0100 @@ -9,199 +9,247 @@ theory Hilbert_Choice = NatArith files ("Hilbert_Choice_lemmas.ML") ("meson_lemmas.ML") ("Tools/meson.ML"): + +subsection {* Hilbert's epsilon *} + consts Eps :: "('a => bool) => 'a" - syntax (input) - "_Eps" :: "[pttrn, bool] => 'a" ("(3\_./ _)" [0, 10] 10) - + "_Eps" :: "[pttrn, bool] => 'a" ("(3\_./ _)" [0, 10] 10) syntax (HOL) - "_Eps" :: "[pttrn, bool] => 'a" ("(3@ _./ _)" [0, 10] 10) - + "_Eps" :: "[pttrn, bool] => 'a" ("(3@ _./ _)" [0, 10] 10) syntax - "_Eps" :: "[pttrn, bool] => 'a" ("(3SOME _./ _)" [0, 10] 10) - + "_Eps" :: "[pttrn, bool] => 'a" ("(3SOME _./ _)" [0, 10] 10) translations - "SOME x. P" == "Eps (%x. P)" + "SOME x. P" == "Eps (%x. P)" -axioms - someI: "P (x::'a) ==> P (SOME x. P x)" +axioms + someI: "P (x::'a) ==> P (SOME x. P x)" -(*used in TFL*) -lemma tfl_some: "\P x. P x --> P (Eps P)" - by (blast intro: someI) - +constdefs + inv :: "('a => 'b) => ('b => 'a)" + "inv(f :: 'a => 'b) == %y. SOME x. f x = y" -constdefs - inv :: "('a => 'b) => ('b => 'a)" - "inv(f::'a=>'b) == % y. @x. f(x)=y" - - Inv :: "['a set, 'a => 'b] => ('b => 'a)" - "Inv A f == (% x. (@ y. y : A & f y = x))" + Inv :: "'a set => ('a => 'b) => ('b => 'a)" + "Inv A f == %x. SOME y. y : A & f y = x" use "Hilbert_Choice_lemmas.ML" +lemma tfl_some: "\P x. P x --> P (Eps P)" + -- {* dynamically-scoped fact for TFL *} + by (blast intro: someI) -(** Least value operator **) + +subsection {* Least value operator *} constdefs - LeastM :: "['a => 'b::ord, 'a => bool] => 'a" - "LeastM m P == @x. P x & (ALL y. P y --> m x <= m y)" + LeastM :: "['a => 'b::ord, 'a => bool] => 'a" + "LeastM m P == SOME x. P x & (ALL y. P y --> m x <= m y)" syntax - "@LeastM" :: "[pttrn, 'a=>'b::ord, bool] => 'a" ("LEAST _ WRT _. _" [0,4,10]10) - + "_LeastM" :: "[pttrn, 'a => 'b::ord, bool] => 'a" ("LEAST _ WRT _. _" [0, 4, 10] 10) translations - "LEAST x WRT m. P" == "LeastM m (%x. P)" + "LEAST x WRT m. P" == "LeastM m (%x. P)" lemma LeastMI2: - "[| P x; !!y. P y ==> m x <= m y; - !!x. [| P x; \y. P y --> m x \ m y |] ==> Q x |] - ==> Q (LeastM m P)"; -apply (unfold LeastM_def) -apply (rule someI2_ex) -apply blast -apply blast -done + "P x ==> (!!y. P y ==> m x <= m y) + ==> (!!x. P x ==> \y. P y --> m x \ m y ==> Q x) + ==> Q (LeastM m P)" + apply (unfold LeastM_def) + apply (rule someI2_ex) + apply blast + apply blast + done lemma LeastM_equality: - "[| P k; !!x. P x ==> m k <= m x |] ==> m (LEAST x WRT m. P x) = - (m k::'a::order)"; -apply (rule LeastMI2) -apply assumption -apply blast -apply (blast intro!: order_antisym) -done + "P k ==> (!!x. P x ==> m k <= m x) + ==> m (LEAST x WRT m. P x) = (m k::'a::order)" + apply (rule LeastMI2) + apply assumption + apply blast + apply (blast intro!: order_antisym) + done lemma wf_linord_ex_has_least: - "[|wf r; ALL x y. ((x,y):r^+) = ((y,x)~:r^*); P k|] \ -\ ==> EX x. P x & (!y. P y --> (m x,m y):r^*)" -apply (drule wf_trancl [THEN wf_eq_minimal [THEN iffD1]]) -apply (drule_tac x = "m`Collect P" in spec) -apply force -done + "wf r ==> ALL x y. ((x,y):r^+) = ((y,x)~:r^*) ==> P k + ==> EX x. P x & (!y. P y --> (m x,m y):r^*)" + apply (drule wf_trancl [THEN wf_eq_minimal [THEN iffD1]]) + apply (drule_tac x = "m`Collect P" in spec) + apply force + done -(* successor of obsolete nonempty_has_least *) lemma ex_has_least_nat: - "P k ==> EX x. P x & (ALL y. P y --> m x <= (m y::nat))" -apply (simp only: pred_nat_trancl_eq_le [symmetric]) -apply (rule wf_pred_nat [THEN wf_linord_ex_has_least]) -apply (simp (no_asm) add: less_eq not_le_iff_less pred_nat_trancl_eq_le) -apply assumption -done + "P k ==> EX x. P x & (ALL y. P y --> m x <= (m y::nat))" + apply (simp only: pred_nat_trancl_eq_le [symmetric]) + apply (rule wf_pred_nat [THEN wf_linord_ex_has_least]) + apply (simp add: less_eq not_le_iff_less pred_nat_trancl_eq_le) + apply assumption + done -lemma LeastM_nat_lemma: - "P k ==> P (LeastM m P) & (ALL y. P y --> m (LeastM m P) <= (m y::nat))" -apply (unfold LeastM_def) -apply (rule someI_ex) -apply (erule ex_has_least_nat) -done +lemma LeastM_nat_lemma: + "P k ==> P (LeastM m P) & (ALL y. P y --> m (LeastM m P) <= (m y::nat))" + apply (unfold LeastM_def) + apply (rule someI_ex) + apply (erule ex_has_least_nat) + done lemmas LeastM_natI = LeastM_nat_lemma [THEN conjunct1, standard] lemma LeastM_nat_le: "P x ==> m (LeastM m P) <= (m x::nat)" -apply (rule LeastM_nat_lemma [THEN conjunct2, THEN spec, THEN mp]) -apply assumption -apply assumption -done + apply (rule LeastM_nat_lemma [THEN conjunct2, THEN spec, THEN mp]) + apply assumption + apply assumption + done -(** Greatest value operator **) +subsection {* Greatest value operator *} constdefs - GreatestM :: "['a => 'b::ord, 'a => bool] => 'a" - "GreatestM m P == @x. P x & (ALL y. P y --> m y <= m x)" - - Greatest :: "('a::ord => bool) => 'a" (binder "GREATEST " 10) - "Greatest == GreatestM (%x. x)" + GreatestM :: "['a => 'b::ord, 'a => bool] => 'a" + "GreatestM m P == SOME x. P x & (ALL y. P y --> m y <= m x)" + + Greatest :: "('a::ord => bool) => 'a" (binder "GREATEST " 10) + "Greatest == GreatestM (%x. x)" syntax - "@GreatestM" :: "[pttrn, 'a=>'b::ord, bool] => 'a" - ("GREATEST _ WRT _. _" [0,4,10]10) + "_GreatestM" :: "[pttrn, 'a=>'b::ord, bool] => 'a" + ("GREATEST _ WRT _. _" [0, 4, 10] 10) translations - "GREATEST x WRT m. P" == "GreatestM m (%x. P)" + "GREATEST x WRT m. P" == "GreatestM m (%x. P)" lemma GreatestMI2: - "[| P x; - !!y. P y ==> m y <= m x; - !!x. [| P x; \y. P y --> m y \ m x |] ==> Q x |] - ==> Q (GreatestM m P)"; -apply (unfold GreatestM_def) -apply (rule someI2_ex) -apply blast -apply blast -done + "P x ==> (!!y. P y ==> m y <= m x) + ==> (!!x. P x ==> \y. P y --> m y \ m x ==> Q x) + ==> Q (GreatestM m P)" + apply (unfold GreatestM_def) + apply (rule someI2_ex) + apply blast + apply blast + done lemma GreatestM_equality: - "[| P k; !!x. P x ==> m x <= m k |] - ==> m (GREATEST x WRT m. P x) = (m k::'a::order)"; -apply (rule_tac m=m in GreatestMI2) -apply assumption -apply blast -apply (blast intro!: order_antisym) -done + "P k ==> (!!x. P x ==> m x <= m k) + ==> m (GREATEST x WRT m. P x) = (m k::'a::order)" + apply (rule_tac m = m in GreatestMI2) + apply assumption + apply blast + apply (blast intro!: order_antisym) + done lemma Greatest_equality: - "[| P (k::'a::order); !!x. P x ==> x <= k |] ==> (GREATEST x. P x) = k"; -apply (unfold Greatest_def) -apply (erule GreatestM_equality) -apply blast -done + "P (k::'a::order) ==> (!!x. P x ==> x <= k) ==> (GREATEST x. P x) = k" + apply (unfold Greatest_def) + apply (erule GreatestM_equality) + apply blast + done lemma ex_has_greatest_nat_lemma: - "[|P k; ALL x. P x --> (EX y. P y & ~ ((m y::nat) <= m x))|] - ==> EX y. P y & ~ (m y < m k + n)" -apply (induct_tac "n") -apply force -(*ind step*) -apply (force simp add: le_Suc_eq) -done + "P k ==> ALL x. P x --> (EX y. P y & ~ ((m y::nat) <= m x)) + ==> EX y. P y & ~ (m y < m k + n)" + apply (induct_tac n) + apply force + apply (force simp add: le_Suc_eq) + done -lemma ex_has_greatest_nat: "[|P k; ! y. P y --> m y < b|] - ==> ? x. P x & (! y. P y --> (m y::nat) <= m x)" -apply (rule ccontr) -apply (cut_tac P = "P" and n = "b - m k" in ex_has_greatest_nat_lemma) -apply (subgoal_tac [3] "m k <= b") -apply auto -done +lemma ex_has_greatest_nat: + "P k ==> ALL y. P y --> m y < b + ==> EX x. P x & (ALL y. P y --> (m y::nat) <= m x)" + apply (rule ccontr) + apply (cut_tac P = P and n = "b - m k" in ex_has_greatest_nat_lemma) + apply (subgoal_tac [3] "m k <= b") + apply auto + done -lemma GreatestM_nat_lemma: - "[|P k; ! y. P y --> m y < b|] - ==> P (GreatestM m P) & (!y. P y --> (m y::nat) <= m (GreatestM m P))" -apply (unfold GreatestM_def) -apply (rule someI_ex) -apply (erule ex_has_greatest_nat) -apply assumption -done +lemma GreatestM_nat_lemma: + "P k ==> ALL y. P y --> m y < b + ==> P (GreatestM m P) & (ALL y. P y --> (m y::nat) <= m (GreatestM m P))" + apply (unfold GreatestM_def) + apply (rule someI_ex) + apply (erule ex_has_greatest_nat) + apply assumption + done lemmas GreatestM_natI = GreatestM_nat_lemma [THEN conjunct1, standard] -lemma GreatestM_nat_le: "[|P x; ! y. P y --> m y < b|] - ==> (m x::nat) <= m (GreatestM m P)" -apply (blast dest: GreatestM_nat_lemma [THEN conjunct2, THEN spec]) -done +lemma GreatestM_nat_le: + "P x ==> ALL y. P y --> m y < b + ==> (m x::nat) <= m (GreatestM m P)" + apply (blast dest: GreatestM_nat_lemma [THEN conjunct2, THEN spec]) + done + + +text {* \medskip Specialization to @{text GREATEST}. *} + +lemma GreatestI: "P (k::nat) ==> ALL y. P y --> y < b ==> P (GREATEST x. P x)" + apply (unfold Greatest_def) + apply (rule GreatestM_natI) + apply auto + done -(** Specialization to GREATEST **) +lemma Greatest_le: + "P x ==> ALL y. P y --> y < b ==> (x::nat) <= (GREATEST x. P x)" + apply (unfold Greatest_def) + apply (rule GreatestM_nat_le) + apply auto + done + + +subsection {* The Meson proof procedure *} -lemma GreatestI: - "[|P (k::nat); ! y. P y --> y < b|] ==> P (GREATEST x. P x)" +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. ~(ALL x. P(x)) ==> EX x. ~P(x)" + and meson_not_exD: "!!P. ~(EX x. P(x)) ==> ALL x. ~P(x)" + by fast+ -apply (unfold Greatest_def) -apply (rule GreatestM_natI) -apply auto -done +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 *} + by fast+ + + +subsubsection {* Pulling out the existential quantifiers *} + +text {* Conjunction *} + +lemma meson_conj_exD1: "!!P Q. (EX x. P(x)) & Q ==> EX x. P(x) & Q" + and meson_conj_exD2: "!!P Q. P & (EX x. Q(x)) ==> EX x. P & Q(x)" + by fast+ + -lemma Greatest_le: - "[|P x; ! y. P y --> y < b|] ==> (x::nat) <= (GREATEST x. P x)" -apply (unfold Greatest_def) -apply (rule GreatestM_nat_le) -apply auto -done +text {* Disjunction *} + +lemma meson_disj_exD: "!!P Q. (EX x. P(x)) | (EX x. Q(x)) ==> EX 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. (EX x. P(x)) | Q ==> EX x. P(x) | Q" + and meson_disj_exD2: "!!P Q. P | (EX x. Q(x)) ==> EX 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+ use "meson_lemmas.ML" use "Tools/meson.ML"