diff -r 7392193f9ecf -r c69069242a51 src/HOL/Dense_Linear_Order.thy --- a/src/HOL/Dense_Linear_Order.thy Sun Jul 22 17:53:42 2007 +0200 +++ b/src/HOL/Dense_Linear_Order.thy Sun Jul 22 17:53:45 2007 +0200 @@ -10,12 +10,144 @@ imports Finite_Set uses "Tools/Qelim/qelim.ML" + "Tools/Qelim/langford_data.ML" "Tools/Qelim/ferrante_rackoff_data.ML" + ("Tools/Qelim/langford.ML") ("Tools/Qelim/ferrante_rackoff.ML") begin + +setup Langford_Data.setup setup Ferrante_Rackoff_Data.setup +section {* The classical QE after Langford for dense linear orders *} + +locale dense_linear_order = Linorder + + assumes gt_ex: "\x. \y. x \ y" + and lt_ex: "\x. \y. y \ x" + and dense: "\x y. x \ y \ (\z. x \ z \ z \ y)" +begin + +lemma dlo_qe_bnds: + assumes ne: "L \ {}" and neU: "U \ {}" and fL: "finite L" and fU: "finite U" + shows "(\x. (\y \ L. y \ x) \ (\y \ U. x \ y)) \ (\ l \ L. \u \ U. l \ u)" +proof (simp only: atomize_eq, rule iffI) + assume H: "\x. (\y\L. y \<^loc>< x) \ (\y\U. x \<^loc>< y)" + then obtain x where xL: "\y\L. y \<^loc>< x" and xU: "\y\U. x \<^loc>< y" by blast + {fix l u assume l: "l \ L" and u: "u \ U" + from less_trans[OF xL[rule_format, OF l] xU[rule_format, OF u]] + have "l \ u" .} + thus "\l\L. \u\U. l \<^loc>< u" by blast +next + assume H: "\l\L. \u\U. l \<^loc>< u" + let ?ML = "Max L" + let ?MU = "Min U" + from fL ne have th1: "?ML \ L" and th1': "\l\L. l \ ?ML" by auto + from fU neU have th2: "?MU \ U" and th2': "\u\U. ?MU \ u" by auto + from th1 th2 H have "?ML \ ?MU" by auto + with dense obtain w where th3: "?ML \ w" and th4: "w \ ?MU" by blast + from th3 th1' have "\l \ L. l \ w" + apply auto + apply (erule_tac x="l" in ballE) + by (auto intro: le_less_trans) + + moreover from th4 th2' have "\u \ U. w \ u" + apply auto + apply (erule_tac x="u" in ballE) + by (auto intro: less_le_trans) + ultimately show "\x. (\y\L. y \<^loc>< x) \ (\y\U. x \<^loc>< y)" by auto +qed + +lemma dlo_qe_noub: + assumes ne: "L \ {}" and fL: "finite L" + shows "(\x. (\y \ L. y \ x) \ (\y \ {}. x \ y)) \ True" +proof(simp add: atomize_eq) + from gt_ex[rule_format, of "Max L"] obtain M where M: "Max L \ M" by blast + from ne fL have "\x \ L. x \ Max L" by simp + with M have "\x\L. x \ M" by (auto intro: le_less_trans) + thus "\x. \y\L. y \<^loc>< x" by blast +qed + +lemma dlo_qe_nolb: + assumes ne: "U \ {}" and fU: "finite U" + shows "(\x. (\y \ {}. y \ x) \ (\y \ U. x \ y)) \ True" +proof(simp add: atomize_eq) + from lt_ex[rule_format, of "Min U"] obtain M where M: "M \ Min U" by blast + from ne fU have "\x \ U. Min U \ x" by simp + with M have "\x\U. M \ x" by (auto intro: less_le_trans) + thus "\x. \y\U. x \<^loc>< y" by blast +qed + +lemma gather_simps: + shows + "(\x. (\y \ L. y \ x) \ (\y \ U. x \ y) \ x \ u \ P x) \ (\x. (\y \ L. y \ x) \ (\y \ (insert u U). x \ y) \ P x)" + and "(\x. (\y \ L. y \ x) \ (\y \ U. x \ y) \ l \ x \ P x) \ (\x. (\y \ (insert l L). y \ x) \ (\y \ U. x \ y) \ P x)" + "(\x. (\y \ L. y \ x) \ (\y \ U. x \ y) \ x \ u) \ (\x. (\y \ L. y \ x) \ (\y \ (insert u U). x \ y))" + and "(\x. (\y \ L. y \ x) \ (\y \ U. x \ y) \ l \ x) \ (\x. (\y \ (insert l L). y \ x) \ (\y \ U. x \ y))" by auto + +lemma + gather_start: "(\x. P x) \ (\x. (\y \ {}. y \<^loc>< x) \ (\y\ {}. x \ y) \ P x)" + by simp + +lemma exists_neq: "\(x::'a). x \ t" "\(x::'a). t \ x" + using gt_ex[rule_format, of t] by auto + +lemma less_not_permute: "\ (x \ y \ y \ x)" by (simp add: not_less linear) + +lemmas dlo_simps = order_refl less_irrefl not_less not_le exists_neq + le_less neq_iff linear less_not_permute + +lemma axiom: "dense_linear_order (op \) (op \)" . +lemma atoms: includes meta_term_syntax + shows "TERM (op \ :: 'a \ _)" and "TERM (op \)" and "TERM (op = :: 'a \ _)" . + +declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms] +declare dlo_simps[langfordsimp] +end + (* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *) + +lemma dnf: + "(P & (Q | R)) = ((P&Q) | (P&R))" + "((Q | R) & P) = ((Q&P) | (R&P))" + by blast+ + +lemmas weak_dnf_simps = simp_thms dnf + + +lemma nnf_simps: + "(\(P \ Q)) = (\P \ \Q)" "(\(P \ Q)) = (\P \ \Q)" "(P \ Q) = (\P \ Q)" + "(P = Q) = ((P \ Q) \ (\P \ \ Q))" "(\ \(P)) = P" + by blast+ + +lemma ex_distrib: "(\x. P x \ Q x) \ ((\x. P x) \ (\x. Q x))" by blast + +lemmas dnf_simps = weak_dnf_simps nnf_simps ex_distrib + +use "Tools/Qelim/langford.ML" + +method_setup dlo = {* + Method.ctxt_args (Method.SIMPLE_METHOD' o LangfordQE.dlo_tac) +*} "Langford's algorithm for quantifier elimination in dense linear orders" + +interpretation dlo_ordring_class: dense_linear_order["op \ :: 'a::{ordered_field} \ _" "op <"] +apply unfold_locales +apply auto +apply (rule_tac x = "x + 1" in exI, simp) +apply (rule_tac x = "x - 1" in exI, simp) +apply (rule_tac x = "(x + y) / (1 + 1)" in exI) +apply (rule conjI) +apply (rule less_half_sum, simp) +apply (rule gt_half_sum, simp) +done + +lemma (in dense_linear_order) "\a b. (\x. a \ x \ x\ b) \ (a \ b)" + by dlo + +lemma "\(a::'a::ordered_field) b. (\x. a < x \ x< b) \ (a < b)" + by dlo + +section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields -- see Arith_Tools.thy *} + context Linorder begin @@ -291,10 +423,17 @@ end -locale dense_linear_order = linorder_no_lb + linorder_no_ub + +locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub + fixes between assumes between_less: "\x y. x \ y \ x \ between x y \ between x y \ y" and between_same: "\x. between x x = x" + +interpretation constr_dense_linear_order < dense_linear_order + apply unfold_locales + using gt_ex lt_ex between_less + by (auto, rule_tac x="between x y" in exI, simp) + +context constr_dense_linear_order begin lemma rinf_U: @@ -374,21 +513,21 @@ lemmas npi_thms = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P lemmas lin_dense_thms = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P -lemma ferrack_axiom: "dense_linear_order less_eq less between" by fact +lemma ferrack_axiom: "constr_dense_linear_order less_eq less between" by fact lemma atoms: includes meta_term_syntax shows "TERM (op \ :: 'a \ _)" and "TERM (op \)" and "TERM (op = :: 'a \ _)" . -declare ferrack_axiom [dlo minf: minf_thms pinf: pinf_thms +declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms nmi: nmi_thms npi: npi_thms lindense: lin_dense_thms qe: fr_eq atoms: atoms] declaration {* let +fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}] fun generic_whatis phi = let val [lt, le] = map (Morphism.term phi) (ProofContext.read_term_pats @{typ "dummy"} @{context} ["op \", "op \"]) (* FIXME avoid read? *) - val le = Morphism.term phi @{term "op \"} fun h x t = case term_of t of Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq @@ -406,7 +545,7 @@ else Ferrante_Rackoff_Data.Nox | _ => Ferrante_Rackoff_Data.Nox in h end - val ss = K (HOL_ss addsimps [@{thm "not_less"}, @{thm "not_le"}]) + fun ss phi = HOL_ss addsimps (simps phi) in Ferrante_Rackoff_Data.funs @{thm "ferrack_axiom"} {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss} @@ -415,9 +554,10 @@ end + use "Tools/Qelim/ferrante_rackoff.ML" -method_setup dlo = {* +method_setup ferrack = {* Method.ctxt_args (Method.SIMPLE_METHOD' o FerranteRackoff.dlo_tac) *} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"