paulson@11451: (* Title: HOL/Hilbert_Choice.thy paulson@11451: ID: $Id$ paulson@11451: Author: Lawrence C Paulson paulson@11451: Copyright 2001 University of Cambridge wenzelm@12023: *) paulson@11451: wenzelm@12023: header {* Hilbert's epsilon-operator and everything to do with the Axiom of Choice *} paulson@11451: paulson@11451: theory Hilbert_Choice = NatArith paulson@11451: files ("Hilbert_Choice_lemmas.ML") ("meson_lemmas.ML") ("Tools/meson.ML"): paulson@11451: wenzelm@12298: wenzelm@12298: subsection {* Hilbert's epsilon *} wenzelm@12298: paulson@11451: consts paulson@11451: Eps :: "('a => bool) => 'a" paulson@11451: paulson@11451: syntax (input) wenzelm@12298: "_Eps" :: "[pttrn, bool] => 'a" ("(3\_./ _)" [0, 10] 10) paulson@11451: syntax (HOL) wenzelm@12298: "_Eps" :: "[pttrn, bool] => 'a" ("(3@ _./ _)" [0, 10] 10) paulson@11451: syntax wenzelm@12298: "_Eps" :: "[pttrn, bool] => 'a" ("(3SOME _./ _)" [0, 10] 10) paulson@11451: translations nipkow@13764: "SOME x. P" == "Eps (%x. P)" nipkow@13763: nipkow@13763: print_translation {* nipkow@13763: (* to avoid eta-contraction of body *) nipkow@13763: [("Eps", fn [Abs abs] => nipkow@13763: let val (x,t) = atomic_abs_tr' abs nipkow@13763: in Syntax.const "_Eps" $ x $ t end)] nipkow@13763: *} paulson@11451: wenzelm@12298: axioms wenzelm@12298: someI: "P (x::'a) ==> P (SOME x. P x)" paulson@11451: paulson@11451: wenzelm@12298: constdefs wenzelm@12298: inv :: "('a => 'b) => ('b => 'a)" wenzelm@12298: "inv(f :: 'a => 'b) == %y. SOME x. f x = y" paulson@11454: wenzelm@12298: Inv :: "'a set => ('a => 'b) => ('b => 'a)" wenzelm@12298: "Inv A f == %x. SOME y. y : A & f y = x" paulson@11451: paulson@11451: paulson@11451: use "Hilbert_Choice_lemmas.ML" wenzelm@12372: declare someI_ex [elim?]; wenzelm@12372: paulson@13585: lemma Inv_mem: "[| f ` A = B; x \ B |] ==> Inv A f x \ A" paulson@13585: apply (unfold Inv_def) paulson@13585: apply (fast intro: someI2) paulson@13585: done paulson@11451: wenzelm@12298: lemma tfl_some: "\P x. P x --> P (Eps P)" wenzelm@12298: -- {* dynamically-scoped fact for TFL *} wenzelm@12298: by (blast intro: someI) paulson@11451: wenzelm@12298: wenzelm@12298: subsection {* Least value operator *} paulson@11451: paulson@11451: constdefs wenzelm@12298: LeastM :: "['a => 'b::ord, 'a => bool] => 'a" wenzelm@12298: "LeastM m P == SOME x. P x & (ALL y. P y --> m x <= m y)" paulson@11451: paulson@11451: syntax wenzelm@12298: "_LeastM" :: "[pttrn, 'a => 'b::ord, bool] => 'a" ("LEAST _ WRT _. _" [0, 4, 10] 10) paulson@11451: translations wenzelm@12298: "LEAST x WRT m. P" == "LeastM m (%x. P)" paulson@11451: paulson@11451: lemma LeastMI2: wenzelm@12298: "P x ==> (!!y. P y ==> m x <= m y) wenzelm@12298: ==> (!!x. P x ==> \y. P y --> m x \ m y ==> Q x) wenzelm@12298: ==> Q (LeastM m P)" wenzelm@12298: apply (unfold LeastM_def) wenzelm@12298: apply (rule someI2_ex) wenzelm@12298: apply blast wenzelm@12298: apply blast wenzelm@12298: done paulson@11451: paulson@11451: lemma LeastM_equality: wenzelm@12298: "P k ==> (!!x. P x ==> m k <= m x) wenzelm@12298: ==> m (LEAST x WRT m. P x) = (m k::'a::order)" wenzelm@12298: apply (rule LeastMI2) wenzelm@12298: apply assumption wenzelm@12298: apply blast wenzelm@12298: apply (blast intro!: order_antisym) wenzelm@12298: done paulson@11451: paulson@11454: lemma wf_linord_ex_has_least: wenzelm@12298: "wf r ==> ALL x y. ((x,y):r^+) = ((y,x)~:r^*) ==> P k wenzelm@12298: ==> EX x. P x & (!y. P y --> (m x,m y):r^*)" wenzelm@12298: apply (drule wf_trancl [THEN wf_eq_minimal [THEN iffD1]]) wenzelm@12298: apply (drule_tac x = "m`Collect P" in spec) wenzelm@12298: apply force wenzelm@12298: done paulson@11454: paulson@11454: lemma ex_has_least_nat: wenzelm@12298: "P k ==> EX x. P x & (ALL y. P y --> m x <= (m y::nat))" wenzelm@12298: apply (simp only: pred_nat_trancl_eq_le [symmetric]) wenzelm@12298: apply (rule wf_pred_nat [THEN wf_linord_ex_has_least]) wenzelm@12298: apply (simp add: less_eq not_le_iff_less pred_nat_trancl_eq_le) wenzelm@12298: apply assumption wenzelm@12298: done paulson@11454: wenzelm@12298: lemma LeastM_nat_lemma: wenzelm@12298: "P k ==> P (LeastM m P) & (ALL y. P y --> m (LeastM m P) <= (m y::nat))" wenzelm@12298: apply (unfold LeastM_def) wenzelm@12298: apply (rule someI_ex) wenzelm@12298: apply (erule ex_has_least_nat) wenzelm@12298: done paulson@11454: paulson@11454: lemmas LeastM_natI = LeastM_nat_lemma [THEN conjunct1, standard] paulson@11454: paulson@11454: lemma LeastM_nat_le: "P x ==> m (LeastM m P) <= (m x::nat)" wenzelm@12298: apply (rule LeastM_nat_lemma [THEN conjunct2, THEN spec, THEN mp]) wenzelm@12298: apply assumption wenzelm@12298: apply assumption wenzelm@12298: done paulson@11454: paulson@11451: wenzelm@12298: subsection {* Greatest value operator *} paulson@11451: paulson@11451: constdefs wenzelm@12298: GreatestM :: "['a => 'b::ord, 'a => bool] => 'a" wenzelm@12298: "GreatestM m P == SOME x. P x & (ALL y. P y --> m y <= m x)" wenzelm@12298: wenzelm@12298: Greatest :: "('a::ord => bool) => 'a" (binder "GREATEST " 10) wenzelm@12298: "Greatest == GreatestM (%x. x)" paulson@11451: paulson@11451: syntax wenzelm@12298: "_GreatestM" :: "[pttrn, 'a=>'b::ord, bool] => 'a" wenzelm@12298: ("GREATEST _ WRT _. _" [0, 4, 10] 10) paulson@11451: paulson@11451: translations wenzelm@12298: "GREATEST x WRT m. P" == "GreatestM m (%x. P)" paulson@11451: paulson@11451: lemma GreatestMI2: wenzelm@12298: "P x ==> (!!y. P y ==> m y <= m x) wenzelm@12298: ==> (!!x. P x ==> \y. P y --> m y \ m x ==> Q x) wenzelm@12298: ==> Q (GreatestM m P)" wenzelm@12298: apply (unfold GreatestM_def) wenzelm@12298: apply (rule someI2_ex) wenzelm@12298: apply blast wenzelm@12298: apply blast wenzelm@12298: done paulson@11451: paulson@11451: lemma GreatestM_equality: wenzelm@12298: "P k ==> (!!x. P x ==> m x <= m k) wenzelm@12298: ==> m (GREATEST x WRT m. P x) = (m k::'a::order)" wenzelm@12298: apply (rule_tac m = m in GreatestMI2) wenzelm@12298: apply assumption wenzelm@12298: apply blast wenzelm@12298: apply (blast intro!: order_antisym) wenzelm@12298: done paulson@11451: paulson@11451: lemma Greatest_equality: wenzelm@12298: "P (k::'a::order) ==> (!!x. P x ==> x <= k) ==> (GREATEST x. P x) = k" wenzelm@12298: apply (unfold Greatest_def) wenzelm@12298: apply (erule GreatestM_equality) wenzelm@12298: apply blast wenzelm@12298: done paulson@11451: paulson@11451: lemma ex_has_greatest_nat_lemma: wenzelm@12298: "P k ==> ALL x. P x --> (EX y. P y & ~ ((m y::nat) <= m x)) wenzelm@12298: ==> EX y. P y & ~ (m y < m k + n)" wenzelm@12298: apply (induct_tac n) wenzelm@12298: apply force wenzelm@12298: apply (force simp add: le_Suc_eq) wenzelm@12298: done paulson@11451: wenzelm@12298: lemma ex_has_greatest_nat: wenzelm@12298: "P k ==> ALL y. P y --> m y < b wenzelm@12298: ==> EX x. P x & (ALL y. P y --> (m y::nat) <= m x)" wenzelm@12298: apply (rule ccontr) wenzelm@12298: apply (cut_tac P = P and n = "b - m k" in ex_has_greatest_nat_lemma) wenzelm@12298: apply (subgoal_tac [3] "m k <= b") wenzelm@12298: apply auto wenzelm@12298: done paulson@11451: wenzelm@12298: lemma GreatestM_nat_lemma: wenzelm@12298: "P k ==> ALL y. P y --> m y < b wenzelm@12298: ==> P (GreatestM m P) & (ALL y. P y --> (m y::nat) <= m (GreatestM m P))" wenzelm@12298: apply (unfold GreatestM_def) wenzelm@12298: apply (rule someI_ex) wenzelm@12298: apply (erule ex_has_greatest_nat) wenzelm@12298: apply assumption wenzelm@12298: done paulson@11451: paulson@11451: lemmas GreatestM_natI = GreatestM_nat_lemma [THEN conjunct1, standard] paulson@11451: wenzelm@12298: lemma GreatestM_nat_le: wenzelm@12298: "P x ==> ALL y. P y --> m y < b wenzelm@12298: ==> (m x::nat) <= m (GreatestM m P)" wenzelm@12298: apply (blast dest: GreatestM_nat_lemma [THEN conjunct2, THEN spec]) wenzelm@12298: done wenzelm@12298: wenzelm@12298: wenzelm@12298: text {* \medskip Specialization to @{text GREATEST}. *} wenzelm@12298: wenzelm@12298: lemma GreatestI: "P (k::nat) ==> ALL y. P y --> y < b ==> P (GREATEST x. P x)" wenzelm@12298: apply (unfold Greatest_def) wenzelm@12298: apply (rule GreatestM_natI) wenzelm@12298: apply auto wenzelm@12298: done paulson@11451: wenzelm@12298: lemma Greatest_le: wenzelm@12298: "P x ==> ALL y. P y --> y < b ==> (x::nat) <= (GREATEST x. P x)" wenzelm@12298: apply (unfold Greatest_def) wenzelm@12298: apply (rule GreatestM_nat_le) wenzelm@12298: apply auto wenzelm@12298: done wenzelm@12298: wenzelm@12298: wenzelm@12298: subsection {* The Meson proof procedure *} paulson@11451: wenzelm@12298: subsubsection {* Negation Normal Form *} wenzelm@12298: wenzelm@12298: text {* de Morgan laws *} wenzelm@12298: wenzelm@12298: lemma meson_not_conjD: "~(P&Q) ==> ~P | ~Q" wenzelm@12298: and meson_not_disjD: "~(P|Q) ==> ~P & ~Q" wenzelm@12298: and meson_not_notD: "~~P ==> P" wenzelm@12298: and meson_not_allD: "!!P. ~(ALL x. P(x)) ==> EX x. ~P(x)" wenzelm@12298: and meson_not_exD: "!!P. ~(EX x. P(x)) ==> ALL x. ~P(x)" wenzelm@12298: by fast+ paulson@11451: wenzelm@12298: text {* Removal of @{text "-->"} and @{text "<->"} (positive and wenzelm@12298: negative occurrences) *} wenzelm@12298: wenzelm@12298: lemma meson_imp_to_disjD: "P-->Q ==> ~P | Q" wenzelm@12298: and meson_not_impD: "~(P-->Q) ==> P & ~Q" wenzelm@12298: and meson_iff_to_disjD: "P=Q ==> (~P | Q) & (~Q | P)" wenzelm@12298: and meson_not_iffD: "~(P=Q) ==> (P | Q) & (~P | ~Q)" wenzelm@12298: -- {* Much more efficient than @{prop "(P & ~Q) | (Q & ~P)"} for computing CNF *} wenzelm@12298: by fast+ wenzelm@12298: wenzelm@12298: wenzelm@12298: subsubsection {* Pulling out the existential quantifiers *} wenzelm@12298: wenzelm@12298: text {* Conjunction *} wenzelm@12298: wenzelm@12298: lemma meson_conj_exD1: "!!P Q. (EX x. P(x)) & Q ==> EX x. P(x) & Q" wenzelm@12298: and meson_conj_exD2: "!!P Q. P & (EX x. Q(x)) ==> EX x. P & Q(x)" wenzelm@12298: by fast+ wenzelm@12298: paulson@11451: wenzelm@12298: text {* Disjunction *} wenzelm@12298: wenzelm@12298: lemma meson_disj_exD: "!!P Q. (EX x. P(x)) | (EX x. Q(x)) ==> EX x. P(x) | Q(x)" wenzelm@12298: -- {* DO NOT USE with forall-Skolemization: makes fewer schematic variables!! *} wenzelm@12298: -- {* With ex-Skolemization, makes fewer Skolem constants *} wenzelm@12298: and meson_disj_exD1: "!!P Q. (EX x. P(x)) | Q ==> EX x. P(x) | Q" wenzelm@12298: and meson_disj_exD2: "!!P Q. P | (EX x. Q(x)) ==> EX x. P | Q(x)" wenzelm@12298: by fast+ wenzelm@12298: paulson@11451: wenzelm@12298: subsubsection {* Generating clauses for the Meson Proof Procedure *} wenzelm@12298: wenzelm@12298: text {* Disjunctions *} wenzelm@12298: wenzelm@12298: lemma meson_disj_assoc: "(P|Q)|R ==> P|(Q|R)" wenzelm@12298: and meson_disj_comm: "P|Q ==> Q|P" wenzelm@12298: and meson_disj_FalseD1: "False|P ==> P" wenzelm@12298: and meson_disj_FalseD2: "P|False ==> P" wenzelm@12298: by fast+ paulson@11451: paulson@11451: use "meson_lemmas.ML" paulson@11451: use "Tools/meson.ML" paulson@11451: setup meson_setup paulson@11451: paulson@11451: end