# HG changeset patch # User chaieb # Date 1204119588 -3600 # Node ID 894f3860ebfdc6df206ea10f357df0bbf7bc3e5a # Parent b037fd9016fa57bb9cc5d724cfc41478c1fa2557 Installation of Quantifier elimination for ordered fields moved to Library/Dense_Linear_Order.thy diff -r b037fd9016fa -r 894f3860ebfd src/HOL/Arith_Tools.thy --- a/src/HOL/Arith_Tools.thy Tue Feb 26 20:38:18 2008 +0100 +++ b/src/HOL/Arith_Tools.thy Wed Feb 27 14:39:48 2008 +0100 @@ -7,7 +7,7 @@ header {* Setup of arithmetic tools *} theory Arith_Tools -imports Groebner_Basis Dense_Linear_Order +imports Groebner_Basis uses "~~/src/Provers/Arith/cancel_numeral_factor.ML" "~~/src/Provers/Arith/extract_common_term.ML" @@ -609,332 +609,4 @@ *} declaration{* fieldgb_declaration *} - - -subsection {* Ferrante and Rackoff algorithm over ordered fields *} - -lemma neg_prod_lt:"(c\'a\ordered_field) < 0 \ ((c*x < 0) == (x > 0))" -proof- - assume H: "c < 0" - have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps) - also have "\ = (0 < x)" by simp - finally show "(c*x < 0) == (x > 0)" by simp -qed - -lemma pos_prod_lt:"(c\'a\ordered_field) > 0 \ ((c*x < 0) == (x < 0))" -proof- - assume H: "c > 0" - hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps) - also have "\ = (0 > x)" by simp - finally show "(c*x < 0) == (x < 0)" by simp -qed - -lemma neg_prod_sum_lt: "(c\'a\ordered_field) < 0 \ ((c*x + t< 0) == (x > (- 1/c)*t))" -proof- - assume H: "c < 0" - have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp) - also have "\ = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps) - also have "\ = ((- 1/c)*t < x)" by simp - finally show "(c*x + t < 0) == (x > (- 1/c)*t)" by simp -qed - -lemma pos_prod_sum_lt:"(c\'a\ordered_field) > 0 \ ((c*x + t < 0) == (x < (- 1/c)*t))" -proof- - assume H: "c > 0" - have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp) - also have "\ = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps) - also have "\ = ((- 1/c)*t > x)" by simp - finally show "(c*x + t < 0) == (x < (- 1/c)*t)" by simp -qed - -lemma sum_lt:"((x::'a::pordered_ab_group_add) + t < 0) == (x < - t)" - using less_diff_eq[where a= x and b=t and c=0] by simp - -lemma neg_prod_le:"(c\'a\ordered_field) < 0 \ ((c*x <= 0) == (x >= 0))" -proof- - assume H: "c < 0" - have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps) - also have "\ = (0 <= x)" by simp - finally show "(c*x <= 0) == (x >= 0)" by simp -qed - -lemma pos_prod_le:"(c\'a\ordered_field) > 0 \ ((c*x <= 0) == (x <= 0))" -proof- - assume H: "c > 0" - hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps) - also have "\ = (0 >= x)" by simp - finally show "(c*x <= 0) == (x <= 0)" by simp -qed - -lemma neg_prod_sum_le: "(c\'a\ordered_field) < 0 \ ((c*x + t <= 0) == (x >= (- 1/c)*t))" -proof- - assume H: "c < 0" - have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp) - also have "\ = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps) - also have "\ = ((- 1/c)*t <= x)" by simp - finally show "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp -qed - -lemma pos_prod_sum_le:"(c\'a\ordered_field) > 0 \ ((c*x + t <= 0) == (x <= (- 1/c)*t))" -proof- - assume H: "c > 0" - have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp) - also have "\ = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps) - also have "\ = ((- 1/c)*t >= x)" by simp - finally show "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp -qed - -lemma sum_le:"((x::'a::pordered_ab_group_add) + t <= 0) == (x <= - t)" - using le_diff_eq[where a= x and b=t and c=0] by simp - -lemma nz_prod_eq:"(c\'a\ordered_field) \ 0 \ ((c*x = 0) == (x = 0))" by simp -lemma nz_prod_sum_eq: "(c\'a\ordered_field) \ 0 \ ((c*x + t = 0) == (x = (- 1/c)*t))" -proof- - assume H: "c \ 0" - have "c*x + t = 0 = (c*x = -t)" by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp) - also have "\ = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] ring_simps) - finally show "(c*x + t = 0) == (x = (- 1/c)*t)" by simp -qed -lemma sum_eq:"((x::'a::pordered_ab_group_add) + t = 0) == (x = - t)" - using eq_diff_eq[where a= x and b=t and c=0] by simp - - -interpretation class_ordered_field_dense_linear_order: constr_dense_linear_order - ["op <=" "op <" - "\ x y. 1/2 * ((x::'a::{ordered_field,recpower,number_ring}) + y)"] -proof (unfold_locales, dlo, dlo, auto) - fix x y::'a assume lt: "x < y" - from less_half_sum[OF lt] show "x < (x + y) /2" by simp -next - fix x y::'a assume lt: "x < y" - from gt_half_sum[OF lt] show "(x + y) /2 < y" by simp -qed - -declaration{* -let -fun earlier [] x y = false - | earlier (h::t) x y = - if h aconvc y then false else if h aconvc x then true else earlier t x y; - -fun dest_frac ct = case term_of ct of - Const (@{const_name "HOL.divide"},_) $ a $ b=> - Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b)) - | t => Rat.rat_of_int (snd (HOLogic.dest_number t)) - -fun mk_frac phi cT x = - let val (a, b) = Rat.quotient_of_rat x - in if b = 1 then Numeral.mk_cnumber cT a - else Thm.capply - (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"}) - (Numeral.mk_cnumber cT a)) - (Numeral.mk_cnumber cT b) - end - -fun whatis x ct = case term_of ct of - Const(@{const_name "HOL.plus"}, _)$(Const(@{const_name "HOL.times"},_)$_$y)$_ => - if y aconv term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct]) - else ("Nox",[]) -| Const(@{const_name "HOL.plus"}, _)$y$_ => - if y aconv term_of x then ("x+t",[Thm.dest_arg ct]) - else ("Nox",[]) -| Const(@{const_name "HOL.times"}, _)$_$y => - if y aconv term_of x then ("c*x",[Thm.dest_arg1 ct]) - else ("Nox",[]) -| t => if t aconv term_of x then ("x",[]) else ("Nox",[]); - -fun xnormalize_conv ctxt [] ct = reflexive ct -| xnormalize_conv ctxt (vs as (x::_)) ct = - case term_of ct of - Const(@{const_name HOL.less},_)$_$Const(@{const_name "HOL.zero"},_) => - (case whatis x (Thm.dest_arg1 ct) of - ("c*x+t",[c,t]) => - let - val cr = dest_frac c - val clt = Thm.dest_fun2 ct - val cz = Thm.dest_arg ct - val neg = cr - let - val T = ctyp_of_term x - val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"} - val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv - (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th - in rth end - | ("c*x",[c]) => - let - val cr = dest_frac c - val clt = Thm.dest_fun2 ct - val cz = Thm.dest_arg ct - val neg = cr reflexive ct) - - -| Const(@{const_name HOL.less_eq},_)$_$Const(@{const_name "HOL.zero"},_) => - (case whatis x (Thm.dest_arg1 ct) of - ("c*x+t",[c,t]) => - let - val T = ctyp_of_term x - val cr = dest_frac c - val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"} - val cz = Thm.dest_arg ct - val neg = cr - let - val T = ctyp_of_term x - val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_le"} - val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv - (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th - in rth end - | ("c*x",[c]) => - let - val T = ctyp_of_term x - val cr = dest_frac c - val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"} - val cz = Thm.dest_arg ct - val neg = cr reflexive ct) - -| Const("op =",_)$_$Const(@{const_name "HOL.zero"},_) => - (case whatis x (Thm.dest_arg1 ct) of - ("c*x+t",[c,t]) => - let - val T = ctyp_of_term x - val cr = dest_frac c - val ceq = Thm.dest_fun2 ct - val cz = Thm.dest_arg ct - val cthp = Simplifier.rewrite (local_simpset_of ctxt) - (Thm.capply @{cterm "Trueprop"} - (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz))) - val cth = equal_elim (symmetric cthp) TrueI - val th = implies_elim - (instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) cth - val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv - (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th - in rth end - | ("x+t",[t]) => - let - val T = ctyp_of_term x - val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"} - val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv - (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th - in rth end - | ("c*x",[c]) => - let - val T = ctyp_of_term x - val cr = dest_frac c - val ceq = Thm.dest_fun2 ct - val cz = Thm.dest_arg ct - val cthp = Simplifier.rewrite (local_simpset_of ctxt) - (Thm.capply @{cterm "Trueprop"} - (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz))) - val cth = equal_elim (symmetric cthp) TrueI - val rth = implies_elim - (instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) cth - in rth end - | _ => reflexive ct); - -local - val less_iff_diff_less_0 = mk_meta_eq @{thm "less_iff_diff_less_0"} - val le_iff_diff_le_0 = mk_meta_eq @{thm "le_iff_diff_le_0"} - val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"} -in -fun field_isolate_conv phi ctxt vs ct = case term_of ct of - Const(@{const_name HOL.less},_)$a$b => - let val (ca,cb) = Thm.dest_binop ct - val T = ctyp_of_term ca - val th = instantiate' [SOME T] [SOME ca, SOME cb] less_iff_diff_less_0 - val nth = Conv.fconv_rule - (Conv.arg_conv (Conv.arg1_conv - (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th - val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth)) - in rth end -| Const(@{const_name HOL.less_eq},_)$a$b => - let val (ca,cb) = Thm.dest_binop ct - val T = ctyp_of_term ca - val th = instantiate' [SOME T] [SOME ca, SOME cb] le_iff_diff_le_0 - val nth = Conv.fconv_rule - (Conv.arg_conv (Conv.arg1_conv - (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th - val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth)) - in rth end - -| Const("op =",_)$a$b => - let val (ca,cb) = Thm.dest_binop ct - val T = ctyp_of_term ca - val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0 - val nth = Conv.fconv_rule - (Conv.arg_conv (Conv.arg1_conv - (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th - val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth)) - in rth end -| @{term "Not"} $(Const("op =",_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct -| _ => reflexive ct -end; - -fun classfield_whatis phi = - let - fun h x t = - case term_of t of - Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq - else Ferrante_Rackoff_Data.Nox - | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq - else Ferrante_Rackoff_Data.Nox - | Const(@{const_name HOL.less},_)$y$z => - if term_of x aconv y then Ferrante_Rackoff_Data.Lt - else if term_of x aconv z then Ferrante_Rackoff_Data.Gt - else Ferrante_Rackoff_Data.Nox - | Const (@{const_name HOL.less_eq},_)$y$z => - if term_of x aconv y then Ferrante_Rackoff_Data.Le - else if term_of x aconv z then Ferrante_Rackoff_Data.Ge - else Ferrante_Rackoff_Data.Nox - | _ => Ferrante_Rackoff_Data.Nox - in h end; -fun class_field_ss phi = - HOL_basic_ss addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}]) - addsplits [@{thm "abs_split"},@{thm "split_max"}, @{thm "split_min"}] - -in -Ferrante_Rackoff_Data.funs @{thm "class_ordered_field_dense_linear_order.ferrack_axiom"} - {isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss} end -*} - -end