# HG changeset patch # User haftmann # Date 1228316324 -3600 # Node ID 15a4b2cf8c34418bb8dffa7b5e7b48434d3e72e4 # Parent 1860f016886d3fdad5873e2ca985fccbaad50341 made repository layout more coherent with logical distribution structure; stripped some $Id$s diff -r 1860f016886d -r 15a4b2cf8c34 NEWS --- a/NEWS Wed Dec 03 09:53:58 2008 +0100 +++ b/NEWS Wed Dec 03 15:58:44 2008 +0100 @@ -58,6 +58,10 @@ *** Pure *** +* Module moves in repository: + src/Pure/Tools/value.ML ~> src/Tools/ + src/Pure/Tools/quickcheck.ML ~> src/Tools/ + * Slightly more coherent Pure syntax, with updated documentation in isar-ref manual. Removed locales meta_term_syntax and meta_conjunction_syntax: TERM and &&& (formerly &&) are now permanent, @@ -133,6 +137,50 @@ *** HOL *** +* Made repository layout more coherent with logical +distribution structure: + + src/HOL/Library/RType.thy ~> src/HOL/Typerep.thy + src/HOL/Library/Code_Message.thy ~> src/HOL/ + src/HOL/Library/GCD.thy ~> src/HOL/ + src/HOL/Library/Order_Relation.thy ~> src/HOL/ + src/HOL/Library/Parity.thy ~> src/HOL/ + src/HOL/Library/Univ_Poly.thy ~> src/HOL/ + src/HOL/Real/ContNotDenum.thy ~> src/HOL/ + src/HOL/Real/Lubs.thy ~> src/HOL/ + src/HOL/Real/PReal.thy ~> src/HOL/ + src/HOL/Real/Rational.thy ~> src/HOL/ + src/HOL/Real/RComplete.thy ~> src/HOL/ + src/HOL/Real/RealDef.thy ~> src/HOL/ + src/HOL/Real/RealPow.thy ~> src/HOL/ + src/HOL/Real/Real.thy ~> src/HOL/ + src/HOL/Complex/Complex_Main.thy ~> src/HOL/ + src/HOL/Complex/Complex.thy ~> src/HOL/ + src/HOL/Complex/FrechetDeriv.thy ~> src/HOL/ + src/HOL/Hyperreal/Deriv.thy ~> src/HOL/ + src/HOL/Hyperreal/Fact.thy ~> src/HOL/ + src/HOL/Hyperreal/Integration.thy ~> src/HOL/ + src/HOL/Hyperreal/Lim.thy ~> src/HOL/ + src/HOL/Hyperreal/Ln.thy ~> src/HOL/ + src/HOL/Hyperreal/Log.thy ~> src/HOL/ + src/HOL/Hyperreal/MacLaurin.thy ~> src/HOL/ + src/HOL/Hyperreal/NthRoot.thy ~> src/HOL/ + src/HOL/Hyperreal/Series.thy ~> src/HOL/ + src/HOL/Hyperreal/Taylor.thy ~> src/HOL/ + src/HOL/Hyperreal/Transcendental.thy ~> src/HOL/ + src/HOL/Real/Float ~> src/HOL/Library/ + + src/HOL/arith_data.ML ~> src/HOL/Tools + src/HOL/hologic.ML ~> src/HOL/Tools + src/HOL/simpdata.ML ~> src/HOL/Tools + src/HOL/int_arith1.ML ~> src/HOL/Tools/int_arith.ML + src/HOL/int_factor_simprocs.ML ~> src/HOL/Tools + src/HOL/nat_simprocs.ML ~> src/HOL/Tools + src/HOL/Real/float_arith.ML ~> src/HOL/Tools + src/HOL/Real/float_syntax.ML ~> src/HOL/Tools + src/HOL/Real/rat_arith.ML ~> src/HOL/Tools + src/HOL/Real/real_arith.ML ~> src/HOL/Tools + * If methods "eval" and "evaluation" encounter a structured proof state with !!/==>, only the conclusion is evaluated to True (if possible), avoiding strange error messages. diff -r 1860f016886d -r 15a4b2cf8c34 doc-src/TutorialI/Types/Numbers.thy --- a/doc-src/TutorialI/Types/Numbers.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/doc-src/TutorialI/Types/Numbers.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,4 +1,3 @@ -(* ID: $Id$ *) theory Numbers imports Complex_Main begin diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Arith_Tools.thy --- a/src/HOL/Arith_Tools.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Arith_Tools.thy Wed Dec 03 15:58:44 2008 +0100 @@ -11,8 +11,8 @@ uses "~~/src/Provers/Arith/cancel_numeral_factor.ML" "~~/src/Provers/Arith/extract_common_term.ML" - "int_factor_simprocs.ML" - "nat_simprocs.ML" + "Tools/int_factor_simprocs.ML" + "Tools/nat_simprocs.ML" "Tools/Qelim/qelim.ML" begin diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Code_Eval.thy --- a/src/HOL/Code_Eval.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Code_Eval.thy Wed Dec 03 15:58:44 2008 +0100 @@ -6,7 +6,7 @@ header {* Term evaluation using the generic code generator *} theory Code_Eval -imports Plain "~~/src/HOL/Library/RType" +imports Plain Typerep begin subsection {* Term representation *} diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Code_Message.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Code_Message.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,58 @@ +(* ID: $Id$ + Author: Florian Haftmann, TU Muenchen +*) + +header {* Monolithic strings (message strings) for code generation *} + +theory Code_Message +imports Plain "~~/src/HOL/List" +begin + +subsection {* Datatype of messages *} + +datatype message_string = STR string + +lemmas [code del] = message_string.recs message_string.cases + +lemma [code]: "size (s\message_string) = 0" + by (cases s) simp_all + +lemma [code]: "message_string_size (s\message_string) = 0" + by (cases s) simp_all + +subsection {* ML interface *} + +ML {* +structure Message_String = +struct + +fun mk s = @{term STR} $ HOLogic.mk_string s; + +end; +*} + + +subsection {* Code serialization *} + +code_type message_string + (SML "string") + (OCaml "string") + (Haskell "String") + +setup {* + fold (fn target => add_literal_message @{const_name STR} target) + ["SML", "OCaml", "Haskell"] +*} + +code_reserved SML string +code_reserved OCaml string + +code_instance message_string :: eq + (Haskell -) + +code_const "eq_class.eq \ message_string \ message_string \ bool" + (SML "!((_ : string) = _)") + (OCaml "!((_ : string) = _)") + (Haskell infixl 4 "==") + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,718 @@ +(* Title: Complex.thy + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 +*) + +header {* Complex Numbers: Rectangular and Polar Representations *} + +theory Complex +imports Transcendental +begin + +datatype complex = Complex real real + +primrec + Re :: "complex \ real" +where + Re: "Re (Complex x y) = x" + +primrec + Im :: "complex \ real" +where + Im: "Im (Complex x y) = y" + +lemma complex_surj [simp]: "Complex (Re z) (Im z) = z" + by (induct z) simp + +lemma complex_equality [intro?]: "\Re x = Re y; Im x = Im y\ \ x = y" + by (induct x, induct y) simp + +lemma expand_complex_eq: "x = y \ Re x = Re y \ Im x = Im y" + by (induct x, induct y) simp + +lemmas complex_Re_Im_cancel_iff = expand_complex_eq + + +subsection {* Addition and Subtraction *} + +instantiation complex :: ab_group_add +begin + +definition + complex_zero_def: "0 = Complex 0 0" + +definition + complex_add_def: "x + y = Complex (Re x + Re y) (Im x + Im y)" + +definition + complex_minus_def: "- x = Complex (- Re x) (- Im x)" + +definition + complex_diff_def: "x - (y\complex) = x + - y" + +lemma Complex_eq_0 [simp]: "Complex a b = 0 \ a = 0 \ b = 0" + by (simp add: complex_zero_def) + +lemma complex_Re_zero [simp]: "Re 0 = 0" + by (simp add: complex_zero_def) + +lemma complex_Im_zero [simp]: "Im 0 = 0" + by (simp add: complex_zero_def) + +lemma complex_add [simp]: + "Complex a b + Complex c d = Complex (a + c) (b + d)" + by (simp add: complex_add_def) + +lemma complex_Re_add [simp]: "Re (x + y) = Re x + Re y" + by (simp add: complex_add_def) + +lemma complex_Im_add [simp]: "Im (x + y) = Im x + Im y" + by (simp add: complex_add_def) + +lemma complex_minus [simp]: + "- (Complex a b) = Complex (- a) (- b)" + by (simp add: complex_minus_def) + +lemma complex_Re_minus [simp]: "Re (- x) = - Re x" + by (simp add: complex_minus_def) + +lemma complex_Im_minus [simp]: "Im (- x) = - Im x" + by (simp add: complex_minus_def) + +lemma complex_diff [simp]: + "Complex a b - Complex c d = Complex (a - c) (b - d)" + by (simp add: complex_diff_def) + +lemma complex_Re_diff [simp]: "Re (x - y) = Re x - Re y" + by (simp add: complex_diff_def) + +lemma complex_Im_diff [simp]: "Im (x - y) = Im x - Im y" + by (simp add: complex_diff_def) + +instance + by intro_classes (simp_all add: complex_add_def complex_diff_def) + +end + + + +subsection {* Multiplication and Division *} + +instantiation complex :: "{field, division_by_zero}" +begin + +definition + complex_one_def: "1 = Complex 1 0" + +definition + complex_mult_def: "x * y = + Complex (Re x * Re y - Im x * Im y) (Re x * Im y + Im x * Re y)" + +definition + complex_inverse_def: "inverse x = + Complex (Re x / ((Re x)\ + (Im x)\)) (- Im x / ((Re x)\ + (Im x)\))" + +definition + complex_divide_def: "x / (y\complex) = x * inverse y" + +lemma Complex_eq_1 [simp]: "(Complex a b = 1) = (a = 1 \ b = 0)" + by (simp add: complex_one_def) + +lemma complex_Re_one [simp]: "Re 1 = 1" + by (simp add: complex_one_def) + +lemma complex_Im_one [simp]: "Im 1 = 0" + by (simp add: complex_one_def) + +lemma complex_mult [simp]: + "Complex a b * Complex c d = Complex (a * c - b * d) (a * d + b * c)" + by (simp add: complex_mult_def) + +lemma complex_Re_mult [simp]: "Re (x * y) = Re x * Re y - Im x * Im y" + by (simp add: complex_mult_def) + +lemma complex_Im_mult [simp]: "Im (x * y) = Re x * Im y + Im x * Re y" + by (simp add: complex_mult_def) + +lemma complex_inverse [simp]: + "inverse (Complex a b) = Complex (a / (a\ + b\)) (- b / (a\ + b\))" + by (simp add: complex_inverse_def) + +lemma complex_Re_inverse: + "Re (inverse x) = Re x / ((Re x)\ + (Im x)\)" + by (simp add: complex_inverse_def) + +lemma complex_Im_inverse: + "Im (inverse x) = - Im x / ((Re x)\ + (Im x)\)" + by (simp add: complex_inverse_def) + +instance + by intro_classes (simp_all add: complex_mult_def + right_distrib left_distrib right_diff_distrib left_diff_distrib + complex_inverse_def complex_divide_def + power2_eq_square add_divide_distrib [symmetric] + expand_complex_eq) + +end + + +subsection {* Exponentiation *} + +instantiation complex :: recpower +begin + +primrec power_complex where + complexpow_0: "z ^ 0 = (1\complex)" + | complexpow_Suc: "z ^ Suc n = (z\complex) * z ^ n" + +instance by intro_classes simp_all + +end + + +subsection {* Numerals and Arithmetic *} + +instantiation complex :: number_ring +begin + +definition number_of_complex where + complex_number_of_def: "number_of w = (of_int w \ complex)" + +instance + by intro_classes (simp only: complex_number_of_def) + +end + +lemma complex_Re_of_nat [simp]: "Re (of_nat n) = of_nat n" +by (induct n) simp_all + +lemma complex_Im_of_nat [simp]: "Im (of_nat n) = 0" +by (induct n) simp_all + +lemma complex_Re_of_int [simp]: "Re (of_int z) = of_int z" +by (cases z rule: int_diff_cases) simp + +lemma complex_Im_of_int [simp]: "Im (of_int z) = 0" +by (cases z rule: int_diff_cases) simp + +lemma complex_Re_number_of [simp]: "Re (number_of v) = number_of v" +unfolding number_of_eq by (rule complex_Re_of_int) + +lemma complex_Im_number_of [simp]: "Im (number_of v) = 0" +unfolding number_of_eq by (rule complex_Im_of_int) + +lemma Complex_eq_number_of [simp]: + "(Complex a b = number_of w) = (a = number_of w \ b = 0)" +by (simp add: expand_complex_eq) + + +subsection {* Scalar Multiplication *} + +instantiation complex :: real_field +begin + +definition + complex_scaleR_def: "scaleR r x = Complex (r * Re x) (r * Im x)" + +lemma complex_scaleR [simp]: + "scaleR r (Complex a b) = Complex (r * a) (r * b)" + unfolding complex_scaleR_def by simp + +lemma complex_Re_scaleR [simp]: "Re (scaleR r x) = r * Re x" + unfolding complex_scaleR_def by simp + +lemma complex_Im_scaleR [simp]: "Im (scaleR r x) = r * Im x" + unfolding complex_scaleR_def by simp + +instance +proof + fix a b :: real and x y :: complex + show "scaleR a (x + y) = scaleR a x + scaleR a y" + by (simp add: expand_complex_eq right_distrib) + show "scaleR (a + b) x = scaleR a x + scaleR b x" + by (simp add: expand_complex_eq left_distrib) + show "scaleR a (scaleR b x) = scaleR (a * b) x" + by (simp add: expand_complex_eq mult_assoc) + show "scaleR 1 x = x" + by (simp add: expand_complex_eq) + show "scaleR a x * y = scaleR a (x * y)" + by (simp add: expand_complex_eq ring_simps) + show "x * scaleR a y = scaleR a (x * y)" + by (simp add: expand_complex_eq ring_simps) +qed + +end + + +subsection{* Properties of Embedding from Reals *} + +abbreviation + complex_of_real :: "real \ complex" where + "complex_of_real \ of_real" + +lemma complex_of_real_def: "complex_of_real r = Complex r 0" +by (simp add: of_real_def complex_scaleR_def) + +lemma Re_complex_of_real [simp]: "Re (complex_of_real z) = z" +by (simp add: complex_of_real_def) + +lemma Im_complex_of_real [simp]: "Im (complex_of_real z) = 0" +by (simp add: complex_of_real_def) + +lemma Complex_add_complex_of_real [simp]: + "Complex x y + complex_of_real r = Complex (x+r) y" +by (simp add: complex_of_real_def) + +lemma complex_of_real_add_Complex [simp]: + "complex_of_real r + Complex x y = Complex (r+x) y" +by (simp add: complex_of_real_def) + +lemma Complex_mult_complex_of_real: + "Complex x y * complex_of_real r = Complex (x*r) (y*r)" +by (simp add: complex_of_real_def) + +lemma complex_of_real_mult_Complex: + "complex_of_real r * Complex x y = Complex (r*x) (r*y)" +by (simp add: complex_of_real_def) + + +subsection {* Vector Norm *} + +instantiation complex :: real_normed_field +begin + +definition + complex_norm_def: "norm z = sqrt ((Re z)\ + (Im z)\)" + +abbreviation + cmod :: "complex \ real" where + "cmod \ norm" + +definition + complex_sgn_def: "sgn x = x /\<^sub>R cmod x" + +lemmas cmod_def = complex_norm_def + +lemma complex_norm [simp]: "cmod (Complex x y) = sqrt (x\ + y\)" + by (simp add: complex_norm_def) + +instance +proof + fix r :: real and x y :: complex + show "0 \ norm x" + by (induct x) simp + show "(norm x = 0) = (x = 0)" + by (induct x) simp + show "norm (x + y) \ norm x + norm y" + by (induct x, induct y) + (simp add: real_sqrt_sum_squares_triangle_ineq) + show "norm (scaleR r x) = \r\ * norm x" + by (induct x) + (simp add: power_mult_distrib right_distrib [symmetric] real_sqrt_mult) + show "norm (x * y) = norm x * norm y" + by (induct x, induct y) + (simp add: real_sqrt_mult [symmetric] power2_eq_square ring_simps) + show "sgn x = x /\<^sub>R cmod x" by(simp add: complex_sgn_def) +qed + +end + +lemma cmod_unit_one [simp]: "cmod (Complex (cos a) (sin a)) = 1" +by simp + +lemma cmod_complex_polar [simp]: + "cmod (complex_of_real r * Complex (cos a) (sin a)) = abs r" +by (simp add: norm_mult) + +lemma complex_Re_le_cmod: "Re x \ cmod x" +unfolding complex_norm_def +by (rule real_sqrt_sum_squares_ge1) + +lemma complex_mod_minus_le_complex_mod [simp]: "- cmod x \ cmod x" +by (rule order_trans [OF _ norm_ge_zero], simp) + +lemma complex_mod_triangle_ineq2 [simp]: "cmod(b + a) - cmod b \ cmod a" +by (rule ord_le_eq_trans [OF norm_triangle_ineq2], simp) + +lemmas real_sum_squared_expand = power2_sum [where 'a=real] + +lemma abs_Re_le_cmod: "\Re x\ \ cmod x" +by (cases x) simp + +lemma abs_Im_le_cmod: "\Im x\ \ cmod x" +by (cases x) simp + +subsection {* Completeness of the Complexes *} + +interpretation Re: bounded_linear ["Re"] +apply (unfold_locales, simp, simp) +apply (rule_tac x=1 in exI) +apply (simp add: complex_norm_def) +done + +interpretation Im: bounded_linear ["Im"] +apply (unfold_locales, simp, simp) +apply (rule_tac x=1 in exI) +apply (simp add: complex_norm_def) +done + +lemma LIMSEQ_Complex: + "\X ----> a; Y ----> b\ \ (\n. Complex (X n) (Y n)) ----> Complex a b" +apply (rule LIMSEQ_I) +apply (subgoal_tac "0 < r / sqrt 2") +apply (drule_tac r="r / sqrt 2" in LIMSEQ_D, safe) +apply (drule_tac r="r / sqrt 2" in LIMSEQ_D, safe) +apply (rename_tac M N, rule_tac x="max M N" in exI, safe) +apply (simp add: real_sqrt_sum_squares_less) +apply (simp add: divide_pos_pos) +done + +instance complex :: banach +proof + fix X :: "nat \ complex" + assume X: "Cauchy X" + from Re.Cauchy [OF X] have 1: "(\n. Re (X n)) ----> lim (\n. Re (X n))" + by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) + from Im.Cauchy [OF X] have 2: "(\n. Im (X n)) ----> lim (\n. Im (X n))" + by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) + have "X ----> Complex (lim (\n. Re (X n))) (lim (\n. Im (X n)))" + using LIMSEQ_Complex [OF 1 2] by simp + thus "convergent X" + by (rule convergentI) +qed + + +subsection {* The Complex Number @{term "\"} *} + +definition + "ii" :: complex ("\") where + i_def: "ii \ Complex 0 1" + +lemma complex_Re_i [simp]: "Re ii = 0" +by (simp add: i_def) + +lemma complex_Im_i [simp]: "Im ii = 1" +by (simp add: i_def) + +lemma Complex_eq_i [simp]: "(Complex x y = ii) = (x = 0 \ y = 1)" +by (simp add: i_def) + +lemma complex_i_not_zero [simp]: "ii \ 0" +by (simp add: expand_complex_eq) + +lemma complex_i_not_one [simp]: "ii \ 1" +by (simp add: expand_complex_eq) + +lemma complex_i_not_number_of [simp]: "ii \ number_of w" +by (simp add: expand_complex_eq) + +lemma i_mult_Complex [simp]: "ii * Complex a b = Complex (- b) a" +by (simp add: expand_complex_eq) + +lemma Complex_mult_i [simp]: "Complex a b * ii = Complex (- b) a" +by (simp add: expand_complex_eq) + +lemma i_complex_of_real [simp]: "ii * complex_of_real r = Complex 0 r" +by (simp add: i_def complex_of_real_def) + +lemma complex_of_real_i [simp]: "complex_of_real r * ii = Complex 0 r" +by (simp add: i_def complex_of_real_def) + +lemma i_squared [simp]: "ii * ii = -1" +by (simp add: i_def) + +lemma power2_i [simp]: "ii\ = -1" +by (simp add: power2_eq_square) + +lemma inverse_i [simp]: "inverse ii = - ii" +by (rule inverse_unique, simp) + + +subsection {* Complex Conjugation *} + +definition + cnj :: "complex \ complex" where + "cnj z = Complex (Re z) (- Im z)" + +lemma complex_cnj [simp]: "cnj (Complex a b) = Complex a (- b)" +by (simp add: cnj_def) + +lemma complex_Re_cnj [simp]: "Re (cnj x) = Re x" +by (simp add: cnj_def) + +lemma complex_Im_cnj [simp]: "Im (cnj x) = - Im x" +by (simp add: cnj_def) + +lemma complex_cnj_cancel_iff [simp]: "(cnj x = cnj y) = (x = y)" +by (simp add: expand_complex_eq) + +lemma complex_cnj_cnj [simp]: "cnj (cnj z) = z" +by (simp add: cnj_def) + +lemma complex_cnj_zero [simp]: "cnj 0 = 0" +by (simp add: expand_complex_eq) + +lemma complex_cnj_zero_iff [iff]: "(cnj z = 0) = (z = 0)" +by (simp add: expand_complex_eq) + +lemma complex_cnj_add: "cnj (x + y) = cnj x + cnj y" +by (simp add: expand_complex_eq) + +lemma complex_cnj_diff: "cnj (x - y) = cnj x - cnj y" +by (simp add: expand_complex_eq) + +lemma complex_cnj_minus: "cnj (- x) = - cnj x" +by (simp add: expand_complex_eq) + +lemma complex_cnj_one [simp]: "cnj 1 = 1" +by (simp add: expand_complex_eq) + +lemma complex_cnj_mult: "cnj (x * y) = cnj x * cnj y" +by (simp add: expand_complex_eq) + +lemma complex_cnj_inverse: "cnj (inverse x) = inverse (cnj x)" +by (simp add: complex_inverse_def) + +lemma complex_cnj_divide: "cnj (x / y) = cnj x / cnj y" +by (simp add: complex_divide_def complex_cnj_mult complex_cnj_inverse) + +lemma complex_cnj_power: "cnj (x ^ n) = cnj x ^ n" +by (induct n, simp_all add: complex_cnj_mult) + +lemma complex_cnj_of_nat [simp]: "cnj (of_nat n) = of_nat n" +by (simp add: expand_complex_eq) + +lemma complex_cnj_of_int [simp]: "cnj (of_int z) = of_int z" +by (simp add: expand_complex_eq) + +lemma complex_cnj_number_of [simp]: "cnj (number_of w) = number_of w" +by (simp add: expand_complex_eq) + +lemma complex_cnj_scaleR: "cnj (scaleR r x) = scaleR r (cnj x)" +by (simp add: expand_complex_eq) + +lemma complex_mod_cnj [simp]: "cmod (cnj z) = cmod z" +by (simp add: complex_norm_def) + +lemma complex_cnj_complex_of_real [simp]: "cnj (of_real x) = of_real x" +by (simp add: expand_complex_eq) + +lemma complex_cnj_i [simp]: "cnj ii = - ii" +by (simp add: expand_complex_eq) + +lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re z)" +by (simp add: expand_complex_eq) + +lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im z) * ii" +by (simp add: expand_complex_eq) + +lemma complex_mult_cnj: "z * cnj z = complex_of_real ((Re z)\ + (Im z)\)" +by (simp add: expand_complex_eq power2_eq_square) + +lemma complex_mod_mult_cnj: "cmod (z * cnj z) = (cmod z)\" +by (simp add: norm_mult power2_eq_square) + +interpretation cnj: bounded_linear ["cnj"] +apply (unfold_locales) +apply (rule complex_cnj_add) +apply (rule complex_cnj_scaleR) +apply (rule_tac x=1 in exI, simp) +done + + +subsection{*The Functions @{term sgn} and @{term arg}*} + +text {*------------ Argand -------------*} + +definition + arg :: "complex => real" where + "arg z = (SOME a. Re(sgn z) = cos a & Im(sgn z) = sin a & -pi < a & a \ pi)" + +lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)" +by (simp add: complex_sgn_def divide_inverse scaleR_conv_of_real mult_commute) + +lemma i_mult_eq: "ii * ii = complex_of_real (-1)" +by (simp add: i_def complex_of_real_def) + +lemma i_mult_eq2 [simp]: "ii * ii = -(1::complex)" +by (simp add: i_def complex_one_def) + +lemma complex_eq_cancel_iff2 [simp]: + "(Complex x y = complex_of_real xa) = (x = xa & y = 0)" +by (simp add: complex_of_real_def) + +lemma Re_sgn [simp]: "Re(sgn z) = Re(z)/cmod z" +by (simp add: complex_sgn_def divide_inverse) + +lemma Im_sgn [simp]: "Im(sgn z) = Im(z)/cmod z" +by (simp add: complex_sgn_def divide_inverse) + +lemma complex_inverse_complex_split: + "inverse(complex_of_real x + ii * complex_of_real y) = + complex_of_real(x/(x ^ 2 + y ^ 2)) - + ii * complex_of_real(y/(x ^ 2 + y ^ 2))" +by (simp add: complex_of_real_def i_def diff_minus divide_inverse) + +(*----------------------------------------------------------------------------*) +(* Many of the theorems below need to be moved elsewhere e.g. Transc. Also *) +(* many of the theorems are not used - so should they be kept? *) +(*----------------------------------------------------------------------------*) + +lemma cos_arg_i_mult_zero_pos: + "0 < y ==> cos (arg(Complex 0 y)) = 0" +apply (simp add: arg_def abs_if) +apply (rule_tac a = "pi/2" in someI2, auto) +apply (rule order_less_trans [of _ 0], auto) +done + +lemma cos_arg_i_mult_zero_neg: + "y < 0 ==> cos (arg(Complex 0 y)) = 0" +apply (simp add: arg_def abs_if) +apply (rule_tac a = "- pi/2" in someI2, auto) +apply (rule order_trans [of _ 0], auto) +done + +lemma cos_arg_i_mult_zero [simp]: + "y \ 0 ==> cos (arg(Complex 0 y)) = 0" +by (auto simp add: linorder_neq_iff cos_arg_i_mult_zero_pos cos_arg_i_mult_zero_neg) + + +subsection{*Finally! Polar Form for Complex Numbers*} + +definition + + (* abbreviation for (cos a + i sin a) *) + cis :: "real => complex" where + "cis a = Complex (cos a) (sin a)" + +definition + (* abbreviation for r*(cos a + i sin a) *) + rcis :: "[real, real] => complex" where + "rcis r a = complex_of_real r * cis a" + +definition + (* e ^ (x + iy) *) + expi :: "complex => complex" where + "expi z = complex_of_real(exp (Re z)) * cis (Im z)" + +lemma complex_split_polar: + "\r a. z = complex_of_real r * (Complex (cos a) (sin a))" +apply (induct z) +apply (auto simp add: polar_Ex complex_of_real_mult_Complex) +done + +lemma rcis_Ex: "\r a. z = rcis r a" +apply (induct z) +apply (simp add: rcis_def cis_def polar_Ex complex_of_real_mult_Complex) +done + +lemma Re_rcis [simp]: "Re(rcis r a) = r * cos a" +by (simp add: rcis_def cis_def) + +lemma Im_rcis [simp]: "Im(rcis r a) = r * sin a" +by (simp add: rcis_def cis_def) + +lemma sin_cos_squared_add2_mult: "(r * cos a)\ + (r * sin a)\ = r\" +proof - + have "(r * cos a)\ + (r * sin a)\ = r\ * ((cos a)\ + (sin a)\)" + by (simp only: power_mult_distrib right_distrib) + thus ?thesis by simp +qed + +lemma complex_mod_rcis [simp]: "cmod(rcis r a) = abs r" +by (simp add: rcis_def cis_def sin_cos_squared_add2_mult) + +lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))" +by (simp add: cmod_def power2_eq_square) + +lemma complex_In_mult_cnj_zero [simp]: "Im (z * cnj z) = 0" +by simp + + +(*---------------------------------------------------------------------------*) +(* (r1 * cis a) * (r2 * cis b) = r1 * r2 * cis (a + b) *) +(*---------------------------------------------------------------------------*) + +lemma cis_rcis_eq: "cis a = rcis 1 a" +by (simp add: rcis_def) + +lemma rcis_mult: "rcis r1 a * rcis r2 b = rcis (r1*r2) (a + b)" +by (simp add: rcis_def cis_def cos_add sin_add right_distrib right_diff_distrib + complex_of_real_def) + +lemma cis_mult: "cis a * cis b = cis (a + b)" +by (simp add: cis_rcis_eq rcis_mult) + +lemma cis_zero [simp]: "cis 0 = 1" +by (simp add: cis_def complex_one_def) + +lemma rcis_zero_mod [simp]: "rcis 0 a = 0" +by (simp add: rcis_def) + +lemma rcis_zero_arg [simp]: "rcis r 0 = complex_of_real r" +by (simp add: rcis_def) + +lemma complex_of_real_minus_one: + "complex_of_real (-(1::real)) = -(1::complex)" +by (simp add: complex_of_real_def complex_one_def) + +lemma complex_i_mult_minus [simp]: "ii * (ii * x) = - x" +by (simp add: mult_assoc [symmetric]) + + +lemma cis_real_of_nat_Suc_mult: + "cis (real (Suc n) * a) = cis a * cis (real n * a)" +by (simp add: cis_def real_of_nat_Suc left_distrib cos_add sin_add right_distrib) + +lemma DeMoivre: "(cis a) ^ n = cis (real n * a)" +apply (induct_tac "n") +apply (auto simp add: cis_real_of_nat_Suc_mult) +done + +lemma DeMoivre2: "(rcis r a) ^ n = rcis (r ^ n) (real n * a)" +by (simp add: rcis_def power_mult_distrib DeMoivre) + +lemma cis_inverse [simp]: "inverse(cis a) = cis (-a)" +by (simp add: cis_def complex_inverse_complex_split diff_minus) + +lemma rcis_inverse: "inverse(rcis r a) = rcis (1/r) (-a)" +by (simp add: divide_inverse rcis_def) + +lemma cis_divide: "cis a / cis b = cis (a - b)" +by (simp add: complex_divide_def cis_mult real_diff_def) + +lemma rcis_divide: "rcis r1 a / rcis r2 b = rcis (r1/r2) (a - b)" +apply (simp add: complex_divide_def) +apply (case_tac "r2=0", simp) +apply (simp add: rcis_inverse rcis_mult real_diff_def) +done + +lemma Re_cis [simp]: "Re(cis a) = cos a" +by (simp add: cis_def) + +lemma Im_cis [simp]: "Im(cis a) = sin a" +by (simp add: cis_def) + +lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re(cis a ^ n)" +by (auto simp add: DeMoivre) + +lemma sin_n_Im_cis_pow_n: "sin (real n * a) = Im(cis a ^ n)" +by (auto simp add: DeMoivre) + +lemma expi_add: "expi(a + b) = expi(a) * expi(b)" +by (simp add: expi_def exp_add cis_mult [symmetric] mult_ac) + +lemma expi_zero [simp]: "expi (0::complex) = 1" +by (simp add: expi_def) + +lemma complex_expi_Ex: "\a r. z = complex_of_real r * expi a" +apply (insert rcis_Ex [of z]) +apply (auto simp add: expi_def rcis_def mult_assoc [symmetric]) +apply (rule_tac x = "ii * complex_of_real a" in exI, auto) +done + +lemma expi_two_pi_i [simp]: "expi((2::complex) * complex_of_real pi * ii) = 1" +by (simp add: expi_def cis_def) + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/Complex.thy --- a/src/HOL/Complex/Complex.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,719 +0,0 @@ -(* Title: Complex.thy - ID: $Id$ - Author: Jacques D. Fleuriot - Copyright: 2001 University of Edinburgh - Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 -*) - -header {* Complex Numbers: Rectangular and Polar Representations *} - -theory Complex -imports "../Hyperreal/Transcendental" -begin - -datatype complex = Complex real real - -primrec - Re :: "complex \ real" -where - Re: "Re (Complex x y) = x" - -primrec - Im :: "complex \ real" -where - Im: "Im (Complex x y) = y" - -lemma complex_surj [simp]: "Complex (Re z) (Im z) = z" - by (induct z) simp - -lemma complex_equality [intro?]: "\Re x = Re y; Im x = Im y\ \ x = y" - by (induct x, induct y) simp - -lemma expand_complex_eq: "x = y \ Re x = Re y \ Im x = Im y" - by (induct x, induct y) simp - -lemmas complex_Re_Im_cancel_iff = expand_complex_eq - - -subsection {* Addition and Subtraction *} - -instantiation complex :: ab_group_add -begin - -definition - complex_zero_def: "0 = Complex 0 0" - -definition - complex_add_def: "x + y = Complex (Re x + Re y) (Im x + Im y)" - -definition - complex_minus_def: "- x = Complex (- Re x) (- Im x)" - -definition - complex_diff_def: "x - (y\complex) = x + - y" - -lemma Complex_eq_0 [simp]: "Complex a b = 0 \ a = 0 \ b = 0" - by (simp add: complex_zero_def) - -lemma complex_Re_zero [simp]: "Re 0 = 0" - by (simp add: complex_zero_def) - -lemma complex_Im_zero [simp]: "Im 0 = 0" - by (simp add: complex_zero_def) - -lemma complex_add [simp]: - "Complex a b + Complex c d = Complex (a + c) (b + d)" - by (simp add: complex_add_def) - -lemma complex_Re_add [simp]: "Re (x + y) = Re x + Re y" - by (simp add: complex_add_def) - -lemma complex_Im_add [simp]: "Im (x + y) = Im x + Im y" - by (simp add: complex_add_def) - -lemma complex_minus [simp]: - "- (Complex a b) = Complex (- a) (- b)" - by (simp add: complex_minus_def) - -lemma complex_Re_minus [simp]: "Re (- x) = - Re x" - by (simp add: complex_minus_def) - -lemma complex_Im_minus [simp]: "Im (- x) = - Im x" - by (simp add: complex_minus_def) - -lemma complex_diff [simp]: - "Complex a b - Complex c d = Complex (a - c) (b - d)" - by (simp add: complex_diff_def) - -lemma complex_Re_diff [simp]: "Re (x - y) = Re x - Re y" - by (simp add: complex_diff_def) - -lemma complex_Im_diff [simp]: "Im (x - y) = Im x - Im y" - by (simp add: complex_diff_def) - -instance - by intro_classes (simp_all add: complex_add_def complex_diff_def) - -end - - - -subsection {* Multiplication and Division *} - -instantiation complex :: "{field, division_by_zero}" -begin - -definition - complex_one_def: "1 = Complex 1 0" - -definition - complex_mult_def: "x * y = - Complex (Re x * Re y - Im x * Im y) (Re x * Im y + Im x * Re y)" - -definition - complex_inverse_def: "inverse x = - Complex (Re x / ((Re x)\ + (Im x)\)) (- Im x / ((Re x)\ + (Im x)\))" - -definition - complex_divide_def: "x / (y\complex) = x * inverse y" - -lemma Complex_eq_1 [simp]: "(Complex a b = 1) = (a = 1 \ b = 0)" - by (simp add: complex_one_def) - -lemma complex_Re_one [simp]: "Re 1 = 1" - by (simp add: complex_one_def) - -lemma complex_Im_one [simp]: "Im 1 = 0" - by (simp add: complex_one_def) - -lemma complex_mult [simp]: - "Complex a b * Complex c d = Complex (a * c - b * d) (a * d + b * c)" - by (simp add: complex_mult_def) - -lemma complex_Re_mult [simp]: "Re (x * y) = Re x * Re y - Im x * Im y" - by (simp add: complex_mult_def) - -lemma complex_Im_mult [simp]: "Im (x * y) = Re x * Im y + Im x * Re y" - by (simp add: complex_mult_def) - -lemma complex_inverse [simp]: - "inverse (Complex a b) = Complex (a / (a\ + b\)) (- b / (a\ + b\))" - by (simp add: complex_inverse_def) - -lemma complex_Re_inverse: - "Re (inverse x) = Re x / ((Re x)\ + (Im x)\)" - by (simp add: complex_inverse_def) - -lemma complex_Im_inverse: - "Im (inverse x) = - Im x / ((Re x)\ + (Im x)\)" - by (simp add: complex_inverse_def) - -instance - by intro_classes (simp_all add: complex_mult_def - right_distrib left_distrib right_diff_distrib left_diff_distrib - complex_inverse_def complex_divide_def - power2_eq_square add_divide_distrib [symmetric] - expand_complex_eq) - -end - - -subsection {* Exponentiation *} - -instantiation complex :: recpower -begin - -primrec power_complex where - complexpow_0: "z ^ 0 = (1\complex)" - | complexpow_Suc: "z ^ Suc n = (z\complex) * z ^ n" - -instance by intro_classes simp_all - -end - - -subsection {* Numerals and Arithmetic *} - -instantiation complex :: number_ring -begin - -definition number_of_complex where - complex_number_of_def: "number_of w = (of_int w \ complex)" - -instance - by intro_classes (simp only: complex_number_of_def) - -end - -lemma complex_Re_of_nat [simp]: "Re (of_nat n) = of_nat n" -by (induct n) simp_all - -lemma complex_Im_of_nat [simp]: "Im (of_nat n) = 0" -by (induct n) simp_all - -lemma complex_Re_of_int [simp]: "Re (of_int z) = of_int z" -by (cases z rule: int_diff_cases) simp - -lemma complex_Im_of_int [simp]: "Im (of_int z) = 0" -by (cases z rule: int_diff_cases) simp - -lemma complex_Re_number_of [simp]: "Re (number_of v) = number_of v" -unfolding number_of_eq by (rule complex_Re_of_int) - -lemma complex_Im_number_of [simp]: "Im (number_of v) = 0" -unfolding number_of_eq by (rule complex_Im_of_int) - -lemma Complex_eq_number_of [simp]: - "(Complex a b = number_of w) = (a = number_of w \ b = 0)" -by (simp add: expand_complex_eq) - - -subsection {* Scalar Multiplication *} - -instantiation complex :: real_field -begin - -definition - complex_scaleR_def: "scaleR r x = Complex (r * Re x) (r * Im x)" - -lemma complex_scaleR [simp]: - "scaleR r (Complex a b) = Complex (r * a) (r * b)" - unfolding complex_scaleR_def by simp - -lemma complex_Re_scaleR [simp]: "Re (scaleR r x) = r * Re x" - unfolding complex_scaleR_def by simp - -lemma complex_Im_scaleR [simp]: "Im (scaleR r x) = r * Im x" - unfolding complex_scaleR_def by simp - -instance -proof - fix a b :: real and x y :: complex - show "scaleR a (x + y) = scaleR a x + scaleR a y" - by (simp add: expand_complex_eq right_distrib) - show "scaleR (a + b) x = scaleR a x + scaleR b x" - by (simp add: expand_complex_eq left_distrib) - show "scaleR a (scaleR b x) = scaleR (a * b) x" - by (simp add: expand_complex_eq mult_assoc) - show "scaleR 1 x = x" - by (simp add: expand_complex_eq) - show "scaleR a x * y = scaleR a (x * y)" - by (simp add: expand_complex_eq ring_simps) - show "x * scaleR a y = scaleR a (x * y)" - by (simp add: expand_complex_eq ring_simps) -qed - -end - - -subsection{* Properties of Embedding from Reals *} - -abbreviation - complex_of_real :: "real \ complex" where - "complex_of_real \ of_real" - -lemma complex_of_real_def: "complex_of_real r = Complex r 0" -by (simp add: of_real_def complex_scaleR_def) - -lemma Re_complex_of_real [simp]: "Re (complex_of_real z) = z" -by (simp add: complex_of_real_def) - -lemma Im_complex_of_real [simp]: "Im (complex_of_real z) = 0" -by (simp add: complex_of_real_def) - -lemma Complex_add_complex_of_real [simp]: - "Complex x y + complex_of_real r = Complex (x+r) y" -by (simp add: complex_of_real_def) - -lemma complex_of_real_add_Complex [simp]: - "complex_of_real r + Complex x y = Complex (r+x) y" -by (simp add: complex_of_real_def) - -lemma Complex_mult_complex_of_real: - "Complex x y * complex_of_real r = Complex (x*r) (y*r)" -by (simp add: complex_of_real_def) - -lemma complex_of_real_mult_Complex: - "complex_of_real r * Complex x y = Complex (r*x) (r*y)" -by (simp add: complex_of_real_def) - - -subsection {* Vector Norm *} - -instantiation complex :: real_normed_field -begin - -definition - complex_norm_def: "norm z = sqrt ((Re z)\ + (Im z)\)" - -abbreviation - cmod :: "complex \ real" where - "cmod \ norm" - -definition - complex_sgn_def: "sgn x = x /\<^sub>R cmod x" - -lemmas cmod_def = complex_norm_def - -lemma complex_norm [simp]: "cmod (Complex x y) = sqrt (x\ + y\)" - by (simp add: complex_norm_def) - -instance -proof - fix r :: real and x y :: complex - show "0 \ norm x" - by (induct x) simp - show "(norm x = 0) = (x = 0)" - by (induct x) simp - show "norm (x + y) \ norm x + norm y" - by (induct x, induct y) - (simp add: real_sqrt_sum_squares_triangle_ineq) - show "norm (scaleR r x) = \r\ * norm x" - by (induct x) - (simp add: power_mult_distrib right_distrib [symmetric] real_sqrt_mult) - show "norm (x * y) = norm x * norm y" - by (induct x, induct y) - (simp add: real_sqrt_mult [symmetric] power2_eq_square ring_simps) - show "sgn x = x /\<^sub>R cmod x" by(simp add: complex_sgn_def) -qed - -end - -lemma cmod_unit_one [simp]: "cmod (Complex (cos a) (sin a)) = 1" -by simp - -lemma cmod_complex_polar [simp]: - "cmod (complex_of_real r * Complex (cos a) (sin a)) = abs r" -by (simp add: norm_mult) - -lemma complex_Re_le_cmod: "Re x \ cmod x" -unfolding complex_norm_def -by (rule real_sqrt_sum_squares_ge1) - -lemma complex_mod_minus_le_complex_mod [simp]: "- cmod x \ cmod x" -by (rule order_trans [OF _ norm_ge_zero], simp) - -lemma complex_mod_triangle_ineq2 [simp]: "cmod(b + a) - cmod b \ cmod a" -by (rule ord_le_eq_trans [OF norm_triangle_ineq2], simp) - -lemmas real_sum_squared_expand = power2_sum [where 'a=real] - -lemma abs_Re_le_cmod: "\Re x\ \ cmod x" -by (cases x) simp - -lemma abs_Im_le_cmod: "\Im x\ \ cmod x" -by (cases x) simp - -subsection {* Completeness of the Complexes *} - -interpretation Re: bounded_linear ["Re"] -apply (unfold_locales, simp, simp) -apply (rule_tac x=1 in exI) -apply (simp add: complex_norm_def) -done - -interpretation Im: bounded_linear ["Im"] -apply (unfold_locales, simp, simp) -apply (rule_tac x=1 in exI) -apply (simp add: complex_norm_def) -done - -lemma LIMSEQ_Complex: - "\X ----> a; Y ----> b\ \ (\n. Complex (X n) (Y n)) ----> Complex a b" -apply (rule LIMSEQ_I) -apply (subgoal_tac "0 < r / sqrt 2") -apply (drule_tac r="r / sqrt 2" in LIMSEQ_D, safe) -apply (drule_tac r="r / sqrt 2" in LIMSEQ_D, safe) -apply (rename_tac M N, rule_tac x="max M N" in exI, safe) -apply (simp add: real_sqrt_sum_squares_less) -apply (simp add: divide_pos_pos) -done - -instance complex :: banach -proof - fix X :: "nat \ complex" - assume X: "Cauchy X" - from Re.Cauchy [OF X] have 1: "(\n. Re (X n)) ----> lim (\n. Re (X n))" - by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) - from Im.Cauchy [OF X] have 2: "(\n. Im (X n)) ----> lim (\n. Im (X n))" - by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) - have "X ----> Complex (lim (\n. Re (X n))) (lim (\n. Im (X n)))" - using LIMSEQ_Complex [OF 1 2] by simp - thus "convergent X" - by (rule convergentI) -qed - - -subsection {* The Complex Number @{term "\"} *} - -definition - "ii" :: complex ("\") where - i_def: "ii \ Complex 0 1" - -lemma complex_Re_i [simp]: "Re ii = 0" -by (simp add: i_def) - -lemma complex_Im_i [simp]: "Im ii = 1" -by (simp add: i_def) - -lemma Complex_eq_i [simp]: "(Complex x y = ii) = (x = 0 \ y = 1)" -by (simp add: i_def) - -lemma complex_i_not_zero [simp]: "ii \ 0" -by (simp add: expand_complex_eq) - -lemma complex_i_not_one [simp]: "ii \ 1" -by (simp add: expand_complex_eq) - -lemma complex_i_not_number_of [simp]: "ii \ number_of w" -by (simp add: expand_complex_eq) - -lemma i_mult_Complex [simp]: "ii * Complex a b = Complex (- b) a" -by (simp add: expand_complex_eq) - -lemma Complex_mult_i [simp]: "Complex a b * ii = Complex (- b) a" -by (simp add: expand_complex_eq) - -lemma i_complex_of_real [simp]: "ii * complex_of_real r = Complex 0 r" -by (simp add: i_def complex_of_real_def) - -lemma complex_of_real_i [simp]: "complex_of_real r * ii = Complex 0 r" -by (simp add: i_def complex_of_real_def) - -lemma i_squared [simp]: "ii * ii = -1" -by (simp add: i_def) - -lemma power2_i [simp]: "ii\ = -1" -by (simp add: power2_eq_square) - -lemma inverse_i [simp]: "inverse ii = - ii" -by (rule inverse_unique, simp) - - -subsection {* Complex Conjugation *} - -definition - cnj :: "complex \ complex" where - "cnj z = Complex (Re z) (- Im z)" - -lemma complex_cnj [simp]: "cnj (Complex a b) = Complex a (- b)" -by (simp add: cnj_def) - -lemma complex_Re_cnj [simp]: "Re (cnj x) = Re x" -by (simp add: cnj_def) - -lemma complex_Im_cnj [simp]: "Im (cnj x) = - Im x" -by (simp add: cnj_def) - -lemma complex_cnj_cancel_iff [simp]: "(cnj x = cnj y) = (x = y)" -by (simp add: expand_complex_eq) - -lemma complex_cnj_cnj [simp]: "cnj (cnj z) = z" -by (simp add: cnj_def) - -lemma complex_cnj_zero [simp]: "cnj 0 = 0" -by (simp add: expand_complex_eq) - -lemma complex_cnj_zero_iff [iff]: "(cnj z = 0) = (z = 0)" -by (simp add: expand_complex_eq) - -lemma complex_cnj_add: "cnj (x + y) = cnj x + cnj y" -by (simp add: expand_complex_eq) - -lemma complex_cnj_diff: "cnj (x - y) = cnj x - cnj y" -by (simp add: expand_complex_eq) - -lemma complex_cnj_minus: "cnj (- x) = - cnj x" -by (simp add: expand_complex_eq) - -lemma complex_cnj_one [simp]: "cnj 1 = 1" -by (simp add: expand_complex_eq) - -lemma complex_cnj_mult: "cnj (x * y) = cnj x * cnj y" -by (simp add: expand_complex_eq) - -lemma complex_cnj_inverse: "cnj (inverse x) = inverse (cnj x)" -by (simp add: complex_inverse_def) - -lemma complex_cnj_divide: "cnj (x / y) = cnj x / cnj y" -by (simp add: complex_divide_def complex_cnj_mult complex_cnj_inverse) - -lemma complex_cnj_power: "cnj (x ^ n) = cnj x ^ n" -by (induct n, simp_all add: complex_cnj_mult) - -lemma complex_cnj_of_nat [simp]: "cnj (of_nat n) = of_nat n" -by (simp add: expand_complex_eq) - -lemma complex_cnj_of_int [simp]: "cnj (of_int z) = of_int z" -by (simp add: expand_complex_eq) - -lemma complex_cnj_number_of [simp]: "cnj (number_of w) = number_of w" -by (simp add: expand_complex_eq) - -lemma complex_cnj_scaleR: "cnj (scaleR r x) = scaleR r (cnj x)" -by (simp add: expand_complex_eq) - -lemma complex_mod_cnj [simp]: "cmod (cnj z) = cmod z" -by (simp add: complex_norm_def) - -lemma complex_cnj_complex_of_real [simp]: "cnj (of_real x) = of_real x" -by (simp add: expand_complex_eq) - -lemma complex_cnj_i [simp]: "cnj ii = - ii" -by (simp add: expand_complex_eq) - -lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re z)" -by (simp add: expand_complex_eq) - -lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im z) * ii" -by (simp add: expand_complex_eq) - -lemma complex_mult_cnj: "z * cnj z = complex_of_real ((Re z)\ + (Im z)\)" -by (simp add: expand_complex_eq power2_eq_square) - -lemma complex_mod_mult_cnj: "cmod (z * cnj z) = (cmod z)\" -by (simp add: norm_mult power2_eq_square) - -interpretation cnj: bounded_linear ["cnj"] -apply (unfold_locales) -apply (rule complex_cnj_add) -apply (rule complex_cnj_scaleR) -apply (rule_tac x=1 in exI, simp) -done - - -subsection{*The Functions @{term sgn} and @{term arg}*} - -text {*------------ Argand -------------*} - -definition - arg :: "complex => real" where - "arg z = (SOME a. Re(sgn z) = cos a & Im(sgn z) = sin a & -pi < a & a \ pi)" - -lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)" -by (simp add: complex_sgn_def divide_inverse scaleR_conv_of_real mult_commute) - -lemma i_mult_eq: "ii * ii = complex_of_real (-1)" -by (simp add: i_def complex_of_real_def) - -lemma i_mult_eq2 [simp]: "ii * ii = -(1::complex)" -by (simp add: i_def complex_one_def) - -lemma complex_eq_cancel_iff2 [simp]: - "(Complex x y = complex_of_real xa) = (x = xa & y = 0)" -by (simp add: complex_of_real_def) - -lemma Re_sgn [simp]: "Re(sgn z) = Re(z)/cmod z" -by (simp add: complex_sgn_def divide_inverse) - -lemma Im_sgn [simp]: "Im(sgn z) = Im(z)/cmod z" -by (simp add: complex_sgn_def divide_inverse) - -lemma complex_inverse_complex_split: - "inverse(complex_of_real x + ii * complex_of_real y) = - complex_of_real(x/(x ^ 2 + y ^ 2)) - - ii * complex_of_real(y/(x ^ 2 + y ^ 2))" -by (simp add: complex_of_real_def i_def diff_minus divide_inverse) - -(*----------------------------------------------------------------------------*) -(* Many of the theorems below need to be moved elsewhere e.g. Transc. Also *) -(* many of the theorems are not used - so should they be kept? *) -(*----------------------------------------------------------------------------*) - -lemma cos_arg_i_mult_zero_pos: - "0 < y ==> cos (arg(Complex 0 y)) = 0" -apply (simp add: arg_def abs_if) -apply (rule_tac a = "pi/2" in someI2, auto) -apply (rule order_less_trans [of _ 0], auto) -done - -lemma cos_arg_i_mult_zero_neg: - "y < 0 ==> cos (arg(Complex 0 y)) = 0" -apply (simp add: arg_def abs_if) -apply (rule_tac a = "- pi/2" in someI2, auto) -apply (rule order_trans [of _ 0], auto) -done - -lemma cos_arg_i_mult_zero [simp]: - "y \ 0 ==> cos (arg(Complex 0 y)) = 0" -by (auto simp add: linorder_neq_iff cos_arg_i_mult_zero_pos cos_arg_i_mult_zero_neg) - - -subsection{*Finally! Polar Form for Complex Numbers*} - -definition - - (* abbreviation for (cos a + i sin a) *) - cis :: "real => complex" where - "cis a = Complex (cos a) (sin a)" - -definition - (* abbreviation for r*(cos a + i sin a) *) - rcis :: "[real, real] => complex" where - "rcis r a = complex_of_real r * cis a" - -definition - (* e ^ (x + iy) *) - expi :: "complex => complex" where - "expi z = complex_of_real(exp (Re z)) * cis (Im z)" - -lemma complex_split_polar: - "\r a. z = complex_of_real r * (Complex (cos a) (sin a))" -apply (induct z) -apply (auto simp add: polar_Ex complex_of_real_mult_Complex) -done - -lemma rcis_Ex: "\r a. z = rcis r a" -apply (induct z) -apply (simp add: rcis_def cis_def polar_Ex complex_of_real_mult_Complex) -done - -lemma Re_rcis [simp]: "Re(rcis r a) = r * cos a" -by (simp add: rcis_def cis_def) - -lemma Im_rcis [simp]: "Im(rcis r a) = r * sin a" -by (simp add: rcis_def cis_def) - -lemma sin_cos_squared_add2_mult: "(r * cos a)\ + (r * sin a)\ = r\" -proof - - have "(r * cos a)\ + (r * sin a)\ = r\ * ((cos a)\ + (sin a)\)" - by (simp only: power_mult_distrib right_distrib) - thus ?thesis by simp -qed - -lemma complex_mod_rcis [simp]: "cmod(rcis r a) = abs r" -by (simp add: rcis_def cis_def sin_cos_squared_add2_mult) - -lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))" -by (simp add: cmod_def power2_eq_square) - -lemma complex_In_mult_cnj_zero [simp]: "Im (z * cnj z) = 0" -by simp - - -(*---------------------------------------------------------------------------*) -(* (r1 * cis a) * (r2 * cis b) = r1 * r2 * cis (a + b) *) -(*---------------------------------------------------------------------------*) - -lemma cis_rcis_eq: "cis a = rcis 1 a" -by (simp add: rcis_def) - -lemma rcis_mult: "rcis r1 a * rcis r2 b = rcis (r1*r2) (a + b)" -by (simp add: rcis_def cis_def cos_add sin_add right_distrib right_diff_distrib - complex_of_real_def) - -lemma cis_mult: "cis a * cis b = cis (a + b)" -by (simp add: cis_rcis_eq rcis_mult) - -lemma cis_zero [simp]: "cis 0 = 1" -by (simp add: cis_def complex_one_def) - -lemma rcis_zero_mod [simp]: "rcis 0 a = 0" -by (simp add: rcis_def) - -lemma rcis_zero_arg [simp]: "rcis r 0 = complex_of_real r" -by (simp add: rcis_def) - -lemma complex_of_real_minus_one: - "complex_of_real (-(1::real)) = -(1::complex)" -by (simp add: complex_of_real_def complex_one_def) - -lemma complex_i_mult_minus [simp]: "ii * (ii * x) = - x" -by (simp add: mult_assoc [symmetric]) - - -lemma cis_real_of_nat_Suc_mult: - "cis (real (Suc n) * a) = cis a * cis (real n * a)" -by (simp add: cis_def real_of_nat_Suc left_distrib cos_add sin_add right_distrib) - -lemma DeMoivre: "(cis a) ^ n = cis (real n * a)" -apply (induct_tac "n") -apply (auto simp add: cis_real_of_nat_Suc_mult) -done - -lemma DeMoivre2: "(rcis r a) ^ n = rcis (r ^ n) (real n * a)" -by (simp add: rcis_def power_mult_distrib DeMoivre) - -lemma cis_inverse [simp]: "inverse(cis a) = cis (-a)" -by (simp add: cis_def complex_inverse_complex_split diff_minus) - -lemma rcis_inverse: "inverse(rcis r a) = rcis (1/r) (-a)" -by (simp add: divide_inverse rcis_def) - -lemma cis_divide: "cis a / cis b = cis (a - b)" -by (simp add: complex_divide_def cis_mult real_diff_def) - -lemma rcis_divide: "rcis r1 a / rcis r2 b = rcis (r1/r2) (a - b)" -apply (simp add: complex_divide_def) -apply (case_tac "r2=0", simp) -apply (simp add: rcis_inverse rcis_mult real_diff_def) -done - -lemma Re_cis [simp]: "Re(cis a) = cos a" -by (simp add: cis_def) - -lemma Im_cis [simp]: "Im(cis a) = sin a" -by (simp add: cis_def) - -lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re(cis a ^ n)" -by (auto simp add: DeMoivre) - -lemma sin_n_Im_cis_pow_n: "sin (real n * a) = Im(cis a ^ n)" -by (auto simp add: DeMoivre) - -lemma expi_add: "expi(a + b) = expi(a) * expi(b)" -by (simp add: expi_def exp_add cis_mult [symmetric] mult_ac) - -lemma expi_zero [simp]: "expi (0::complex) = 1" -by (simp add: expi_def) - -lemma complex_expi_Ex: "\a r. z = complex_of_real r * expi a" -apply (insert rcis_Ex [of z]) -apply (auto simp add: expi_def rcis_def mult_assoc [symmetric]) -apply (rule_tac x = "ii * complex_of_real a" in exI, auto) -done - -lemma expi_two_pi_i [simp]: "expi((2::complex) * complex_of_real pi * ii) = 1" -by (simp add: expi_def cis_def) - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/Complex_Main.thy --- a/src/HOL/Complex/Complex_Main.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,22 +0,0 @@ -(* Title: HOL/Complex/Complex_Main.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 2003 University of Cambridge -*) - -header{*Comprehensive Complex Theory*} - -theory Complex_Main -imports - "../Main" - "../Real/ContNotDenum" - "../Real/Real" - Fundamental_Theorem_Algebra - "../Hyperreal/Log" - "../Hyperreal/Ln" - "../Hyperreal/Taylor" - "../Hyperreal/Integration" - "../Hyperreal/FrechetDeriv" -begin - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/Fundamental_Theorem_Algebra.thy --- a/src/HOL/Complex/Fundamental_Theorem_Algebra.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Complex/Fundamental_Theorem_Algebra.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,12 +1,11 @@ (* Title: Fundamental_Theorem_Algebra.thy - ID: $Id$ Author: Amine Chaieb *) header{*Fundamental Theorem of Algebra*} theory Fundamental_Theorem_Algebra -imports "~~/src/HOL/Library/Univ_Poly" "~~/src/HOL/Library/Dense_Linear_Order" Complex +imports "~~/src/HOL/Univ_Poly" "~~/src/HOL/Library/Dense_Linear_Order" "~~/src/HOL/Complex" begin subsection {* Square root of complex numbers *} diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/ex/Arithmetic_Series_Complex.thy --- a/src/HOL/Complex/ex/Arithmetic_Series_Complex.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,24 +0,0 @@ -(* Title: HOL/Complex/ex/Arithmetic_Series_Complex - ID: $Id$ - Author: Benjamin Porter, 2006 -*) - - -header {* Arithmetic Series for Reals *} - -theory Arithmetic_Series_Complex -imports Complex_Main -begin - -lemma arith_series_real: - "(2::real) * (\i\{..i\{.. g ----> 0 ==> f ----> (0::real)" - apply (simp add: LIMSEQ_def bigo_alt_def) - apply clarify - apply (drule_tac x = "r / c" in spec) - apply (drule mp) - apply (erule divide_pos_pos) - apply assumption - apply clarify - apply (rule_tac x = no in exI) - apply (rule allI) - apply (drule_tac x = n in spec)+ - apply (rule impI) - apply (drule mp) - apply assumption - apply (rule order_le_less_trans) - apply assumption - apply (rule order_less_le_trans) - apply (subgoal_tac "c * abs(g n) < c * (r / c)") - apply assumption - apply (erule mult_strict_left_mono) - apply assumption - apply simp -done - -lemma bigo_LIMSEQ2: "f =o g +o O(h) ==> h ----> 0 ==> f ----> a - ==> g ----> (a::real)" - apply (drule set_plus_imp_minus) - apply (drule bigo_LIMSEQ1) - apply assumption - apply (simp only: fun_diff_def) - apply (erule LIMSEQ_diff_approach_zero2) - apply assumption -done - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/ex/BinEx.thy --- a/src/HOL/Complex/ex/BinEx.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,399 +0,0 @@ -(* Title: HOL/Complex/ex/BinEx.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1999 University of Cambridge -*) - -header {* Binary arithmetic examples *} - -theory BinEx -imports Complex_Main -begin - -text {* - Examples of performing binary arithmetic by simplification. This time - we use the reals, though the representation is just of integers. -*} - -subsection{*Real Arithmetic*} - -subsubsection {*Addition *} - -lemma "(1359::real) + -2468 = -1109" -by simp - -lemma "(93746::real) + -46375 = 47371" -by simp - - -subsubsection {*Negation *} - -lemma "- (65745::real) = -65745" -by simp - -lemma "- (-54321::real) = 54321" -by simp - - -subsubsection {*Multiplication *} - -lemma "(-84::real) * 51 = -4284" -by simp - -lemma "(255::real) * 255 = 65025" -by simp - -lemma "(1359::real) * -2468 = -3354012" -by simp - - -subsubsection {*Inequalities *} - -lemma "(89::real) * 10 \ 889" -by simp - -lemma "(13::real) < 18 - 4" -by simp - -lemma "(-345::real) < -242 + -100" -by simp - -lemma "(13557456::real) < 18678654" -by simp - -lemma "(999999::real) \ (1000001 + 1) - 2" -by simp - -lemma "(1234567::real) \ 1234567" -by simp - - -subsubsection {*Powers *} - -lemma "2 ^ 15 = (32768::real)" -by simp - -lemma "-3 ^ 7 = (-2187::real)" -by simp - -lemma "13 ^ 7 = (62748517::real)" -by simp - -lemma "3 ^ 15 = (14348907::real)" -by simp - -lemma "-5 ^ 11 = (-48828125::real)" -by simp - - -subsubsection {*Tests *} - -lemma "(x + y = x) = (y = (0::real))" -by arith - -lemma "(x + y = y) = (x = (0::real))" -by arith - -lemma "(x + y = (0::real)) = (x = -y)" -by arith - -lemma "(x + y = (0::real)) = (y = -x)" -by arith - -lemma "((x + y) < (x + z)) = (y < (z::real))" -by arith - -lemma "((x + z) < (y + z)) = (x < (y::real))" -by arith - -lemma "(\ x < y) = (y \ (x::real))" -by arith - -lemma "\ (x < y \ y < (x::real))" -by arith - -lemma "(x::real) < y ==> \ y < x" -by arith - -lemma "((x::real) \ y) = (x < y \ y < x)" -by arith - -lemma "(\ x \ y) = (y < (x::real))" -by arith - -lemma "x \ y \ y \ (x::real)" -by arith - -lemma "x \ y \ y < (x::real)" -by arith - -lemma "x < y \ y \ (x::real)" -by arith - -lemma "x \ (x::real)" -by arith - -lemma "((x::real) \ y) = (x < y \ x = y)" -by arith - -lemma "((x::real) \ y \ y \ x) = (x = y)" -by arith - -lemma "\(x < y \ y \ (x::real))" -by arith - -lemma "\(x \ y \ y < (x::real))" -by arith - -lemma "(-x < (0::real)) = (0 < x)" -by arith - -lemma "((0::real) < -x) = (x < 0)" -by arith - -lemma "(-x \ (0::real)) = (0 \ x)" -by arith - -lemma "((0::real) \ -x) = (x \ 0)" -by arith - -lemma "(x::real) = y \ x < y \ y < x" -by arith - -lemma "(x::real) = 0 \ 0 < x \ 0 < -x" -by arith - -lemma "(0::real) \ x \ 0 \ -x" -by arith - -lemma "((x::real) + y \ x + z) = (y \ z)" -by arith - -lemma "((x::real) + z \ y + z) = (x \ y)" -by arith - -lemma "(w::real) < x \ y < z ==> w + y < x + z" -by arith - -lemma "(w::real) \ x \ y \ z ==> w + y \ x + z" -by arith - -lemma "(0::real) \ x \ 0 \ y ==> 0 \ x + y" -by arith - -lemma "(0::real) < x \ 0 < y ==> 0 < x + y" -by arith - -lemma "(-x < y) = (0 < x + (y::real))" -by arith - -lemma "(x < -y) = (x + y < (0::real))" -by arith - -lemma "(y < x + -z) = (y + z < (x::real))" -by arith - -lemma "(x + -y < z) = (x < z + (y::real))" -by arith - -lemma "x \ y ==> x < y + (1::real)" -by arith - -lemma "(x - y) + y = (x::real)" -by arith - -lemma "y + (x - y) = (x::real)" -by arith - -lemma "x - x = (0::real)" -by arith - -lemma "(x - y = 0) = (x = (y::real))" -by arith - -lemma "((0::real) \ x + x) = (0 \ x)" -by arith - -lemma "(-x \ x) = ((0::real) \ x)" -by arith - -lemma "(x \ -x) = (x \ (0::real))" -by arith - -lemma "(-x = (0::real)) = (x = 0)" -by arith - -lemma "-(x - y) = y - (x::real)" -by arith - -lemma "((0::real) < x - y) = (y < x)" -by arith - -lemma "((0::real) \ x - y) = (y \ x)" -by arith - -lemma "(x + y) - x = (y::real)" -by arith - -lemma "(-x = y) = (x = (-y::real))" -by arith - -lemma "x < (y::real) ==> \(x = y)" -by arith - -lemma "(x \ x + y) = ((0::real) \ y)" -by arith - -lemma "(y \ x + y) = ((0::real) \ x)" -by arith - -lemma "(x < x + y) = ((0::real) < y)" -by arith - -lemma "(y < x + y) = ((0::real) < x)" -by arith - -lemma "(x - y) - x = (-y::real)" -by arith - -lemma "(x + y < z) = (x < z - (y::real))" -by arith - -lemma "(x - y < z) = (x < z + (y::real))" -by arith - -lemma "(x < y - z) = (x + z < (y::real))" -by arith - -lemma "(x \ y - z) = (x + z \ (y::real))" -by arith - -lemma "(x - y \ z) = (x \ z + (y::real))" -by arith - -lemma "(-x < -y) = (y < (x::real))" -by arith - -lemma "(-x \ -y) = (y \ (x::real))" -by arith - -lemma "(a + b) - (c + d) = (a - c) + (b - (d::real))" -by arith - -lemma "(0::real) - x = -x" -by arith - -lemma "x - (0::real) = x" -by arith - -lemma "w \ x \ y < z ==> w + y < x + (z::real)" -by arith - -lemma "w < x \ y \ z ==> w + y < x + (z::real)" -by arith - -lemma "(0::real) \ x \ 0 < y ==> 0 < x + (y::real)" -by arith - -lemma "(0::real) < x \ 0 \ y ==> 0 < x + y" -by arith - -lemma "-x - y = -(x + (y::real))" -by arith - -lemma "x - (-y) = x + (y::real)" -by arith - -lemma "-x - -y = y - (x::real)" -by arith - -lemma "(a - b) + (b - c) = a - (c::real)" -by arith - -lemma "(x = y - z) = (x + z = (y::real))" -by arith - -lemma "(x - y = z) = (x = z + (y::real))" -by arith - -lemma "x - (x - y) = (y::real)" -by arith - -lemma "x - (x + y) = -(y::real)" -by arith - -lemma "x = y ==> x \ (y::real)" -by arith - -lemma "(0::real) < x ==> \(x = 0)" -by arith - -lemma "(x + y) * (x - y) = (x * x) - (y * y)" - oops - -lemma "(-x = -y) = (x = (y::real))" -by arith - -lemma "(-x < -y) = (y < (x::real))" -by arith - -lemma "!!a::real. a \ b ==> c \ d ==> x + y < z ==> a + c \ b + d" -by (tactic "fast_arith_tac @{context} 1") - -lemma "!!a::real. a < b ==> c < d ==> a - d \ b + (-c)" -by (tactic "fast_arith_tac @{context} 1") - -lemma "!!a::real. a \ b ==> b + b \ c ==> a + a \ c" -by (tactic "fast_arith_tac @{context} 1") - -lemma "!!a::real. a + b \ i + j ==> a \ b ==> i \ j ==> a + a \ j + j" -by (tactic "fast_arith_tac @{context} 1") - -lemma "!!a::real. a + b < i + j ==> a < b ==> i < j ==> a + a < j + j" -by (tactic "fast_arith_tac @{context} 1") - -lemma "!!a::real. a + b + c \ i + j + k \ a \ b \ b \ c \ i \ j \ j \ k --> a + a + a \ k + k + k" -by arith - -lemma "!!a::real. a + b + c + d \ i + j + k + l ==> a \ b ==> b \ c - ==> c \ d ==> i \ j ==> j \ k ==> k \ l ==> a \ l" -by (tactic "fast_arith_tac @{context} 1") - -lemma "!!a::real. a + b + c + d \ i + j + k + l ==> a \ b ==> b \ c - ==> c \ d ==> i \ j ==> j \ k ==> k \ l ==> a + a + a + a \ l + l + l + l" -by (tactic "fast_arith_tac @{context} 1") - -lemma "!!a::real. a + b + c + d \ i + j + k + l ==> a \ b ==> b \ c - ==> c \ d ==> i \ j ==> j \ k ==> k \ l ==> a + a + a + a + a \ l + l + l + l + i" -by (tactic "fast_arith_tac @{context} 1") - -lemma "!!a::real. a + b + c + d \ i + j + k + l ==> a \ b ==> b \ c - ==> c \ d ==> i \ j ==> j \ k ==> k \ l ==> a + a + a + a + a + a \ l + l + l + l + i + l" -by (tactic "fast_arith_tac @{context} 1") - - -subsection{*Complex Arithmetic*} - -lemma "(1359 + 93746*ii) - (2468 + 46375*ii) = -1109 + 47371*ii" -by simp - -lemma "- (65745 + -47371*ii) = -65745 + 47371*ii" -by simp - -text{*Multiplication requires distributive laws. Perhaps versions instantiated -to literal constants should be added to the simpset.*} - -lemma "(1 + ii) * (1 - ii) = 2" -by (simp add: ring_distribs) - -lemma "(1 + 2*ii) * (1 + 3*ii) = -5 + 5*ii" -by (simp add: ring_distribs) - -lemma "(-84 + 255*ii) + (51 * 255*ii) = -84 + 13260 * ii" -by (simp add: ring_distribs) - -text{*No inequalities or linear arithmetic: the complex numbers are unordered!*} - -text{*No powers (not supported yet)*} - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/ex/HarmonicSeries.thy --- a/src/HOL/Complex/ex/HarmonicSeries.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,323 +0,0 @@ -(* Title: HOL/Library/HarmonicSeries.thy - ID: $Id$ - Author: Benjamin Porter, 2006 -*) - -header {* Divergence of the Harmonic Series *} - -theory HarmonicSeries -imports Complex_Main -begin - -section {* Abstract *} - -text {* The following document presents a proof of the Divergence of -Harmonic Series theorem formalised in the Isabelle/Isar theorem -proving system. - -{\em Theorem:} The series $\sum_{n=1}^{\infty} \frac{1}{n}$ does not -converge to any number. - -{\em Informal Proof:} - The informal proof is based on the following auxillary lemmas: - \begin{itemize} - \item{aux: $\sum_{n=2^m-1}^{2^m} \frac{1}{n} \geq \frac{1}{2}$} - \item{aux2: $\sum_{n=1}^{2^M} \frac{1}{n} = 1 + \sum_{m=1}^{M} \sum_{n=2^m-1}^{2^m} \frac{1}{n}$} - \end{itemize} - - From {\em aux} and {\em aux2} we can deduce that $\sum_{n=1}^{2^M} - \frac{1}{n} \geq 1 + \frac{M}{2}$ for all $M$. - Now for contradiction, assume that $\sum_{n=1}^{\infty} \frac{1}{n} - = s$ for some $s$. Because $\forall n. \frac{1}{n} > 0$ all the - partial sums in the series must be less than $s$. However with our - deduction above we can choose $N > 2*s - 2$ and thus - $\sum_{n=1}^{2^N} \frac{1}{n} > s$. This leads to a contradiction - and hence $\sum_{n=1}^{\infty} \frac{1}{n}$ is not summable. - QED. -*} - -section {* Formal Proof *} - -lemma two_pow_sub: - "0 < m \ (2::nat)^m - 2^(m - 1) = 2^(m - 1)" - by (induct m) auto - -text {* We first prove the following auxillary lemma. This lemma -simply states that the finite sums: $\frac{1}{2}$, $\frac{1}{3} + -\frac{1}{4}$, $\frac{1}{5} + \frac{1}{6} + \frac{1}{7} + \frac{1}{8}$ -etc. are all greater than or equal to $\frac{1}{2}$. We do this by -observing that each term in the sum is greater than or equal to the -last term, e.g. $\frac{1}{3} > \frac{1}{4}$ and thus $\frac{1}{3} + -\frac{1}{4} > \frac{1}{4} + \frac{1}{4} = \frac{1}{2}$. *} - -lemma harmonic_aux: - "\m>0. (\n\{(2::nat)^(m - 1)+1..2^m}. 1/real n) \ 1/2" - (is "\m>0. (\n\(?S m). 1/real n) \ 1/2") -proof - fix m::nat - obtain tm where tmdef: "tm = (2::nat)^m" by simp - { - assume mgt0: "0 < m" - have "\x. x\(?S m) \ 1/(real x) \ 1/(real tm)" - proof - - fix x::nat - assume xs: "x\(?S m)" - have xgt0: "x>0" - proof - - from xs have - "x \ 2^(m - 1) + 1" by auto - moreover with mgt0 have - "2^(m - 1) + 1 \ (1::nat)" by auto - ultimately have - "x \ 1" by (rule xtrans) - thus ?thesis by simp - qed - moreover from xs have "x \ 2^m" by auto - ultimately have - "inverse (real x) \ inverse (real ((2::nat)^m))" by simp - moreover - from xgt0 have "real x \ 0" by simp - then have - "inverse (real x) = 1 / (real x)" - by (rule nonzero_inverse_eq_divide) - moreover from mgt0 have "real tm \ 0" by (simp add: tmdef) - then have - "inverse (real tm) = 1 / (real tm)" - by (rule nonzero_inverse_eq_divide) - ultimately show - "1/(real x) \ 1/(real tm)" by (auto simp add: tmdef) - qed - then have - "(\n\(?S m). 1 / real n) \ (\n\(?S m). 1/(real tm))" - by (rule setsum_mono) - moreover have - "(\n\(?S m). 1/(real tm)) = 1/2" - proof - - have - "(\n\(?S m). 1/(real tm)) = - (1/(real tm))*(\n\(?S m). 1)" - by simp - also have - "\ = ((1/(real tm)) * real (card (?S m)))" - by (simp add: real_of_card real_of_nat_def) - also have - "\ = ((1/(real tm)) * real (tm - (2^(m - 1))))" - by (simp add: tmdef) - also from mgt0 have - "\ = ((1/(real tm)) * real ((2::nat)^(m - 1)))" - by (auto simp: tmdef dest: two_pow_sub) - also have - "\ = (real (2::nat))^(m - 1) / (real (2::nat))^m" - by (simp add: tmdef realpow_real_of_nat [symmetric]) - also from mgt0 have - "\ = (real (2::nat))^(m - 1) / (real (2::nat))^((m - 1) + 1)" - by auto - also have "\ = 1/2" by simp - finally show ?thesis . - qed - ultimately have - "(\n\(?S m). 1 / real n) \ 1/2" - by - (erule subst) - } - thus "0 < m \ 1 / 2 \ (\n\(?S m). 1 / real n)" by simp -qed - -text {* We then show that the sum of a finite number of terms from the -harmonic series can be regrouped in increasing powers of 2. For -example: $1 + \frac{1}{2} + \frac{1}{3} + \frac{1}{4} + \frac{1}{5} + -\frac{1}{6} + \frac{1}{7} + \frac{1}{8} = 1 + (\frac{1}{2}) + -(\frac{1}{3} + \frac{1}{4}) + (\frac{1}{5} + \frac{1}{6} + \frac{1}{7} -+ \frac{1}{8})$. *} - -lemma harmonic_aux2 [rule_format]: - "0 (\n\{1..(2::nat)^M}. 1/real n) = - (1 + (\m\{1..M}. \n\{(2::nat)^(m - 1)+1..2^m}. 1/real n))" - (is "0 ?LHS M = ?RHS M") -proof (induct M) - case 0 show ?case by simp -next - case (Suc M) - have ant: "0 < Suc M" by fact - { - have suc: "?LHS (Suc M) = ?RHS (Suc M)" - proof cases -- "show that LHS = c and RHS = c, and thus LHS = RHS" - assume mz: "M=0" - { - then have - "?LHS (Suc M) = ?LHS 1" by simp - also have - "\ = (\n\{(1::nat)..2}. 1/real n)" by simp - also have - "\ = ((\n\{Suc 1..2}. 1/real n) + 1/(real (1::nat)))" - by (subst setsum_head) - (auto simp: atLeastSucAtMost_greaterThanAtMost) - also have - "\ = ((\n\{2..2::nat}. 1/real n) + 1/(real (1::nat)))" - by (simp add: nat_number) - also have - "\ = 1/(real (2::nat)) + 1/(real (1::nat))" by simp - finally have - "?LHS (Suc M) = 1/2 + 1" by simp - } - moreover - { - from mz have - "?RHS (Suc M) = ?RHS 1" by simp - also have - "\ = (\n\{((2::nat)^0)+1..2^1}. 1/real n) + 1" - by simp - also have - "\ = (\n\{2::nat..2}. 1/real n) + 1" - proof - - have "(2::nat)^0 = 1" by simp - then have "(2::nat)^0+1 = 2" by simp - moreover have "(2::nat)^1 = 2" by simp - ultimately have "{((2::nat)^0)+1..2^1} = {2::nat..2}" by auto - thus ?thesis by simp - qed - also have - "\ = 1/2 + 1" - by simp - finally have - "?RHS (Suc M) = 1/2 + 1" by simp - } - ultimately show "?LHS (Suc M) = ?RHS (Suc M)" by simp - next - assume mnz: "M\0" - then have mgtz: "M>0" by simp - with Suc have suc: - "(?LHS M) = (?RHS M)" by blast - have - "(?LHS (Suc M)) = - ((?LHS M) + (\n\{(2::nat)^M+1..2^(Suc M)}. 1 / real n))" - proof - - have - "{1..(2::nat)^(Suc M)} = - {1..(2::nat)^M}\{(2::nat)^M+1..(2::nat)^(Suc M)}" - by auto - moreover have - "{1..(2::nat)^M}\{(2::nat)^M+1..(2::nat)^(Suc M)} = {}" - by auto - moreover have - "finite {1..(2::nat)^M}" and "finite {(2::nat)^M+1..(2::nat)^(Suc M)}" - by auto - ultimately show ?thesis - by (auto intro: setsum_Un_disjoint) - qed - moreover - { - have - "(?RHS (Suc M)) = - (1 + (\m\{1..M}. \n\{(2::nat)^(m - 1)+1..2^m}. 1/real n) + - (\n\{(2::nat)^(Suc M - 1)+1..2^(Suc M)}. 1/real n))" by simp - also have - "\ = (?RHS M) + (\n\{(2::nat)^M+1..2^(Suc M)}. 1/real n)" - by simp - also from suc have - "\ = (?LHS M) + (\n\{(2::nat)^M+1..2^(Suc M)}. 1/real n)" - by simp - finally have - "(?RHS (Suc M)) = \" by simp - } - ultimately show "?LHS (Suc M) = ?RHS (Suc M)" by simp - qed - } - thus ?case by simp -qed - -text {* Using @{thm [source] harmonic_aux} and @{thm [source] harmonic_aux2} we now show -that each group sum is greater than or equal to $\frac{1}{2}$ and thus -the finite sum is bounded below by a value proportional to the number -of elements we choose. *} - -lemma harmonic_aux3 [rule_format]: - shows "\(M::nat). (\n\{1..(2::nat)^M}. 1 / real n) \ 1 + (real M)/2" - (is "\M. ?P M \ _") -proof (rule allI, cases) - fix M::nat - assume "M=0" - then show "?P M \ 1 + (real M)/2" by simp -next - fix M::nat - assume "M\0" - then have "M > 0" by simp - then have - "(?P M) = - (1 + (\m\{1..M}. \n\{(2::nat)^(m - 1)+1..2^m}. 1/real n))" - by (rule harmonic_aux2) - also have - "\ \ (1 + (\m\{1..M}. 1/2))" - proof - - let ?f = "(\x. 1/2)" - let ?g = "(\x. (\n\{(2::nat)^(x - 1)+1..2^x}. 1/real n))" - from harmonic_aux have "\x. x\{1..M} \ ?f x \ ?g x" by simp - then have "(\m\{1..M}. ?g m) \ (\m\{1..M}. ?f m)" by (rule setsum_mono) - thus ?thesis by simp - qed - finally have "(?P M) \ (1 + (\m\{1..M}. 1/2))" . - moreover - { - have - "(\m\{1..M}. (1::real)/2) = 1/2 * (\m\{1..M}. 1)" - by auto - also have - "\ = 1/2*(real (card {1..M}))" - by (simp only: real_of_card[symmetric]) - also have - "\ = 1/2*(real M)" by simp - also have - "\ = (real M)/2" by simp - finally have "(\m\{1..M}. (1::real)/2) = (real M)/2" . - } - ultimately show "(?P M) \ (1 + (real M)/2)" by simp -qed - -text {* The final theorem shows that as we take more and more elements -(see @{thm [source] harmonic_aux3}) we get an ever increasing sum. By assuming -the sum converges, the lemma @{thm [source] series_pos_less} ( @{thm -series_pos_less} ) states that each sum is bounded above by the -series' limit. This contradicts our first statement and thus we prove -that the harmonic series is divergent. *} - -theorem DivergenceOfHarmonicSeries: - shows "\summable (\n. 1/real (Suc n))" - (is "\summable ?f") -proof -- "by contradiction" - let ?s = "suminf ?f" -- "let ?s equal the sum of the harmonic series" - assume sf: "summable ?f" - then obtain n::nat where ndef: "n = nat \2 * ?s\" by simp - then have ngt: "1 + real n/2 > ?s" - proof - - have "\n. 0 \ ?f n" by simp - with sf have "?s \ 0" - by - (rule suminf_0_le, simp_all) - then have cgt0: "\2*?s\ \ 0" by simp - - from ndef have "n = nat \(2*?s)\" . - then have "real n = real (nat \2*?s\)" by simp - with cgt0 have "real n = real \2*?s\" - by (auto dest: real_nat_eq_real) - then have "real n \ 2*(?s)" by simp - then have "real n/2 \ (?s)" by simp - then show "1 + real n/2 > (?s)" by simp - qed - - obtain j where jdef: "j = (2::nat)^n" by simp - have "\m\j. 0 < ?f m" by simp - with sf have "(\i\{0..i\{1..i\{1..< Suc ((2::nat)^n)}. 1 / (real i)) < ?s" by simp - then have - "(\i\{1..(2::nat)^n}. 1 / (real i)) < ?s" - by (simp only: atLeastLessThanSuc_atLeastAtMost) - moreover from harmonic_aux3 have - "(\i\{1..(2::nat)^n}. 1 / (real i)) \ 1 + real n/2" by simp - moreover from ngt have "1 + real n/2 > ?s" by simp - ultimately show False by simp -qed - -end \ No newline at end of file diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/ex/MIR.thy --- a/src/HOL/Complex/ex/MIR.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,5933 +0,0 @@ -(* Title: Complex/ex/MIR.thy - Author: Amine Chaieb -*) - -theory MIR -imports List Real Code_Integer Efficient_Nat -uses ("mirtac.ML") -begin - -section {* Quantifier elimination for @{text "\ (0, 1, +, floor, <)"} *} - -declare real_of_int_floor_cancel [simp del] - -primrec alluopairs:: "'a list \ ('a \ 'a) list" where - "alluopairs [] = []" -| "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)" - -lemma alluopairs_set1: "set (alluopairs xs) \ {(x,y). x\ set xs \ y\ set xs}" -by (induct xs, auto) - -lemma alluopairs_set: - "\x\ set xs ; y \ set xs\ \ (x,y) \ set (alluopairs xs) \ (y,x) \ set (alluopairs xs) " -by (induct xs, auto) - -lemma alluopairs_ex: - assumes Pc: "\ x y. P x y = P y x" - shows "(\ x \ set xs. \ y \ set xs. P x y) = (\ (x,y) \ set (alluopairs xs). P x y)" -proof - assume "\x\set xs. \y\set xs. P x y" - then obtain x y where x: "x \ set xs" and y:"y \ set xs" and P: "P x y" by blast - from alluopairs_set[OF x y] P Pc show"\(x, y)\set (alluopairs xs). P x y" - by auto -next - assume "\(x, y)\set (alluopairs xs). P x y" - then obtain "x" and "y" where xy:"(x,y) \ set (alluopairs xs)" and P: "P x y" by blast+ - from xy have "x \ set xs \ y\ set xs" using alluopairs_set1 by blast - with P show "\x\set xs. \y\set xs. P x y" by blast -qed - - (* generate a list from i to j*) -consts iupt :: "int \ int \ int list" -recdef iupt "measure (\ (i,j). nat (j-i +1))" - "iupt (i,j) = (if j (x#xs) ! n = xs ! (n - 1)" -using Nat.gr0_conv_Suc -by clarsimp - - -lemma myl: "\ (a::'a::{pordered_ab_group_add}) (b::'a). (a \ b) = (0 \ b - a)" -proof(clarify) - fix x y ::"'a" - have "(x \ y) = (x - y \ 0)" by (simp only: le_iff_diff_le_0[where a="x" and b="y"]) - also have "\ = (- (y - x) \ 0)" by simp - also have "\ = (0 \ y - x)" by (simp only: neg_le_0_iff_le[where a="y-x"]) - finally show "(x \ y) = (0 \ y - x)" . -qed - -lemma myless: "\ (a::'a::{pordered_ab_group_add}) (b::'a). (a < b) = (0 < b - a)" -proof(clarify) - fix x y ::"'a" - have "(x < y) = (x - y < 0)" by (simp only: less_iff_diff_less_0[where a="x" and b="y"]) - also have "\ = (- (y - x) < 0)" by simp - also have "\ = (0 < y - x)" by (simp only: neg_less_0_iff_less[where a="y-x"]) - finally show "(x < y) = (0 < y - x)" . -qed - -lemma myeq: "\ (a::'a::{pordered_ab_group_add}) (b::'a). (a = b) = (0 = b - a)" - by auto - - (* Maybe should be added to the library \ *) -lemma floor_int_eq: "(real n\ x \ x < real (n+1)) = (floor x = n)" -proof( auto) - assume lb: "real n \ x" - and ub: "x < real n + 1" - have "real (floor x) \ x" by simp - hence "real (floor x) < real (n + 1) " using ub by arith - hence "floor x < n+1" by simp - moreover from lb have "n \ floor x" using floor_mono2[where x="real n" and y="x"] - by simp ultimately show "floor x = n" by simp -qed - -(* Periodicity of dvd *) -lemma dvd_period: - assumes advdd: "(a::int) dvd d" - shows "(a dvd (x + t)) = (a dvd ((x+ c*d) + t))" - using advdd -proof- - {fix x k - from inf_period(3)[OF advdd, rule_format, where x=x and k="-k"] - have " ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp} - hence "\x.\k. ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp - then show ?thesis by simp -qed - - (* The Divisibility relation between reals *) -definition - rdvd:: "real \ real \ bool" (infixl "rdvd" 50) -where - rdvd_def: "x rdvd y \ (\k\int. y = x * real k)" - -lemma int_rdvd_real: - shows "real (i::int) rdvd x = (i dvd (floor x) \ real (floor x) = x)" (is "?l = ?r") -proof - assume "?l" - hence th: "\ k. x=real (i*k)" by (simp add: rdvd_def) - hence th': "real (floor x) = x" by (auto simp del: real_of_int_mult) - with th have "\ k. real (floor x) = real (i*k)" by simp - hence "\ k. floor x = i*k" by (simp only: real_of_int_inject) - thus ?r using th' by (simp add: dvd_def) -next - assume "?r" hence "(i\int) dvd \x\real\" .. - hence "\ k. real (floor x) = real (i*k)" - by (simp only: real_of_int_inject) (simp add: dvd_def) - thus ?l using prems by (simp add: rdvd_def) -qed - -lemma int_rdvd_iff: "(real (i::int) rdvd real t) = (i dvd t)" -by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only :real_of_int_mult[symmetric]) - - -lemma rdvd_abs1: - "(abs (real d) rdvd t) = (real (d ::int) rdvd t)" -proof - assume d: "real d rdvd t" - from d int_rdvd_real have d2: "d dvd (floor t)" and ti: "real (floor t) = t" by auto - - from iffD2[OF zdvd_abs1] d2 have "(abs d) dvd (floor t)" by blast - with ti int_rdvd_real[symmetric] have "real (abs d) rdvd t" by blast - thus "abs (real d) rdvd t" by simp -next - assume "abs (real d) rdvd t" hence "real (abs d) rdvd t" by simp - with int_rdvd_real[where i="abs d" and x="t"] have d2: "abs d dvd floor t" and ti: "real (floor t) =t" by auto - from iffD1[OF zdvd_abs1] d2 have "d dvd floor t" by blast - with ti int_rdvd_real[symmetric] show "real d rdvd t" by blast -qed - -lemma rdvd_minus: "(real (d::int) rdvd t) = (real d rdvd -t)" - apply (auto simp add: rdvd_def) - apply (rule_tac x="-k" in exI, simp) - apply (rule_tac x="-k" in exI, simp) -done - -lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)" -by (auto simp add: rdvd_def) - -lemma rdvd_mult: - assumes knz: "k\0" - shows "(real (n::int) * real (k::int) rdvd x * real k) = (real n rdvd x)" -using knz by (simp add:rdvd_def) - -lemma rdvd_trans: assumes mn:"m rdvd n" and nk:"n rdvd k" - shows "m rdvd k" -proof- - from rdvd_def mn obtain c where nmc:"n = m * real (c::int)" by auto - from rdvd_def nk obtain c' where nkc:"k = n * real (c'::int)" by auto - hence "k = m * real (c * c')" using nmc by simp - thus ?thesis using rdvd_def by blast -qed - - (*********************************************************************************) - (**** SHADOW SYNTAX AND SEMANTICS ****) - (*********************************************************************************) - -datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num - | Mul int num | Floor num| CF int num num - - (* A size for num to make inductive proofs simpler*) -primrec num_size :: "num \ nat" where - "num_size (C c) = 1" -| "num_size (Bound n) = 1" -| "num_size (Neg a) = 1 + num_size a" -| "num_size (Add a b) = 1 + num_size a + num_size b" -| "num_size (Sub a b) = 3 + num_size a + num_size b" -| "num_size (CN n c a) = 4 + num_size a " -| "num_size (CF c a b) = 4 + num_size a + num_size b" -| "num_size (Mul c a) = 1 + num_size a" -| "num_size (Floor a) = 1 + num_size a" - - (* Semantics of numeral terms (num) *) -primrec Inum :: "real list \ num \ real" where - "Inum bs (C c) = (real c)" -| "Inum bs (Bound n) = bs!n" -| "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)" -| "Inum bs (Neg a) = -(Inum bs a)" -| "Inum bs (Add a b) = Inum bs a + Inum bs b" -| "Inum bs (Sub a b) = Inum bs a - Inum bs b" -| "Inum bs (Mul c a) = (real c) * Inum bs a" -| "Inum bs (Floor a) = real (floor (Inum bs a))" -| "Inum bs (CF c a b) = real c * real (floor (Inum bs a)) + Inum bs b" -definition "isint t bs \ real (floor (Inum bs t)) = Inum bs t" - -lemma isint_iff: "isint n bs = (real (floor (Inum bs n)) = Inum bs n)" -by (simp add: isint_def) - -lemma isint_Floor: "isint (Floor n) bs" - by (simp add: isint_iff) - -lemma isint_Mul: "isint e bs \ isint (Mul c e) bs" -proof- - let ?e = "Inum bs e" - let ?fe = "floor ?e" - assume be: "isint e bs" hence efe:"real ?fe = ?e" by (simp add: isint_iff) - have "real ((floor (Inum bs (Mul c e)))) = real (floor (real (c * ?fe)))" using efe by simp - also have "\ = real (c* ?fe)" by (simp only: floor_real_of_int) - also have "\ = real c * ?e" using efe by simp - finally show ?thesis using isint_iff by simp -qed - -lemma isint_neg: "isint e bs \ isint (Neg e) bs" -proof- - let ?I = "\ t. Inum bs t" - assume ie: "isint e bs" - hence th: "real (floor (?I e)) = ?I e" by (simp add: isint_def) - have "real (floor (?I (Neg e))) = real (floor (- (real (floor (?I e)))))" by (simp add: th) - also have "\ = - real (floor (?I e))" by(simp add: floor_minus_real_of_int) - finally show "isint (Neg e) bs" by (simp add: isint_def th) -qed - -lemma isint_sub: - assumes ie: "isint e bs" shows "isint (Sub (C c) e) bs" -proof- - let ?I = "\ t. Inum bs t" - from ie have th: "real (floor (?I e)) = ?I e" by (simp add: isint_def) - have "real (floor (?I (Sub (C c) e))) = real (floor ((real (c -floor (?I e)))))" by (simp add: th) - also have "\ = real (c- floor (?I e))" by(simp add: floor_minus_real_of_int) - finally show "isint (Sub (C c) e) bs" by (simp add: isint_def th) -qed - -lemma isint_add: assumes - ai:"isint a bs" and bi: "isint b bs" shows "isint (Add a b) bs" -proof- - let ?a = "Inum bs a" - let ?b = "Inum bs b" - from ai bi isint_iff have "real (floor (?a + ?b)) = real (floor (real (floor ?a) + real (floor ?b)))" by simp - also have "\ = real (floor ?a) + real (floor ?b)" by simp - also have "\ = ?a + ?b" using ai bi isint_iff by simp - finally show "isint (Add a b) bs" by (simp add: isint_iff) -qed - -lemma isint_c: "isint (C j) bs" - by (simp add: isint_iff) - - - (* FORMULAE *) -datatype fm = - T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num| - NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm - - - (* A size for fm *) -fun fmsize :: "fm \ nat" where - "fmsize (NOT p) = 1 + fmsize p" -| "fmsize (And p q) = 1 + fmsize p + fmsize q" -| "fmsize (Or p q) = 1 + fmsize p + fmsize q" -| "fmsize (Imp p q) = 3 + fmsize p + fmsize q" -| "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)" -| "fmsize (E p) = 1 + fmsize p" -| "fmsize (A p) = 4+ fmsize p" -| "fmsize (Dvd i t) = 2" -| "fmsize (NDvd i t) = 2" -| "fmsize p = 1" - (* several lemmas about fmsize *) -lemma fmsize_pos: "fmsize p > 0" -by (induct p rule: fmsize.induct) simp_all - - (* Semantics of formulae (fm) *) -primrec Ifm ::"real list \ fm \ bool" where - "Ifm bs T = True" -| "Ifm bs F = False" -| "Ifm bs (Lt a) = (Inum bs a < 0)" -| "Ifm bs (Gt a) = (Inum bs a > 0)" -| "Ifm bs (Le a) = (Inum bs a \ 0)" -| "Ifm bs (Ge a) = (Inum bs a \ 0)" -| "Ifm bs (Eq a) = (Inum bs a = 0)" -| "Ifm bs (NEq a) = (Inum bs a \ 0)" -| "Ifm bs (Dvd i b) = (real i rdvd Inum bs b)" -| "Ifm bs (NDvd i b) = (\(real i rdvd Inum bs b))" -| "Ifm bs (NOT p) = (\ (Ifm bs p))" -| "Ifm bs (And p q) = (Ifm bs p \ Ifm bs q)" -| "Ifm bs (Or p q) = (Ifm bs p \ Ifm bs q)" -| "Ifm bs (Imp p q) = ((Ifm bs p) \ (Ifm bs q))" -| "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)" -| "Ifm bs (E p) = (\ x. Ifm (x#bs) p)" -| "Ifm bs (A p) = (\ x. Ifm (x#bs) p)" - -consts prep :: "fm \ fm" -recdef prep "measure fmsize" - "prep (E T) = T" - "prep (E F) = F" - "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))" - "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))" - "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" - "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))" - "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))" - "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))" - "prep (E p) = E (prep p)" - "prep (A (And p q)) = And (prep (A p)) (prep (A q))" - "prep (A p) = prep (NOT (E (NOT p)))" - "prep (NOT (NOT p)) = prep p" - "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))" - "prep (NOT (A p)) = prep (E (NOT p))" - "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))" - "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))" - "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))" - "prep (NOT p) = NOT (prep p)" - "prep (Or p q) = Or (prep p) (prep q)" - "prep (And p q) = And (prep p) (prep q)" - "prep (Imp p q) = prep (Or (NOT p) q)" - "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))" - "prep p = p" -(hints simp add: fmsize_pos) -lemma prep: "\ bs. Ifm bs (prep p) = Ifm bs p" -by (induct p rule: prep.induct, auto) - - - (* Quantifier freeness *) -fun qfree:: "fm \ bool" where - "qfree (E p) = False" - | "qfree (A p) = False" - | "qfree (NOT p) = qfree p" - | "qfree (And p q) = (qfree p \ qfree q)" - | "qfree (Or p q) = (qfree p \ qfree q)" - | "qfree (Imp p q) = (qfree p \ qfree q)" - | "qfree (Iff p q) = (qfree p \ qfree q)" - | "qfree p = True" - - (* Boundedness and substitution *) -primrec numbound0 :: "num \ bool" (* a num is INDEPENDENT of Bound 0 *) where - "numbound0 (C c) = True" - | "numbound0 (Bound n) = (n>0)" - | "numbound0 (CN n i a) = (n > 0 \ numbound0 a)" - | "numbound0 (Neg a) = numbound0 a" - | "numbound0 (Add a b) = (numbound0 a \ numbound0 b)" - | "numbound0 (Sub a b) = (numbound0 a \ numbound0 b)" - | "numbound0 (Mul i a) = numbound0 a" - | "numbound0 (Floor a) = numbound0 a" - | "numbound0 (CF c a b) = (numbound0 a \ numbound0 b)" - -lemma numbound0_I: - assumes nb: "numbound0 a" - shows "Inum (b#bs) a = Inum (b'#bs) a" - using nb by (induct a) (auto simp add: nth_pos2) - -lemma numbound0_gen: - assumes nb: "numbound0 t" and ti: "isint t (x#bs)" - shows "\ y. isint t (y#bs)" -using nb ti -proof(clarify) - fix y - from numbound0_I[OF nb, where bs="bs" and b="y" and b'="x"] ti[simplified isint_def] - show "isint t (y#bs)" - by (simp add: isint_def) -qed - -primrec bound0:: "fm \ bool" (* A Formula is independent of Bound 0 *) where - "bound0 T = True" - | "bound0 F = True" - | "bound0 (Lt a) = numbound0 a" - | "bound0 (Le a) = numbound0 a" - | "bound0 (Gt a) = numbound0 a" - | "bound0 (Ge a) = numbound0 a" - | "bound0 (Eq a) = numbound0 a" - | "bound0 (NEq a) = numbound0 a" - | "bound0 (Dvd i a) = numbound0 a" - | "bound0 (NDvd i a) = numbound0 a" - | "bound0 (NOT p) = bound0 p" - | "bound0 (And p q) = (bound0 p \ bound0 q)" - | "bound0 (Or p q) = (bound0 p \ bound0 q)" - | "bound0 (Imp p q) = ((bound0 p) \ (bound0 q))" - | "bound0 (Iff p q) = (bound0 p \ bound0 q)" - | "bound0 (E p) = False" - | "bound0 (A p) = False" - -lemma bound0_I: - assumes bp: "bound0 p" - shows "Ifm (b#bs) p = Ifm (b'#bs) p" - using bp numbound0_I [where b="b" and bs="bs" and b'="b'"] - by (induct p) (auto simp add: nth_pos2) - -primrec numsubst0:: "num \ num \ num" (* substitute a num into a num for Bound 0 *) where - "numsubst0 t (C c) = (C c)" - | "numsubst0 t (Bound n) = (if n=0 then t else Bound n)" - | "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))" - | "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)" - | "numsubst0 t (Neg a) = Neg (numsubst0 t a)" - | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)" - | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" - | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)" - | "numsubst0 t (Floor a) = Floor (numsubst0 t a)" - -lemma numsubst0_I: - shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t" - by (induct t) (simp_all add: nth_pos2) - -lemma numsubst0_I': - assumes nb: "numbound0 a" - shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t" - by (induct t) (simp_all add: nth_pos2 numbound0_I[OF nb, where b="b" and b'="b'"]) - -primrec subst0:: "num \ fm \ fm" (* substitue a num into a formula for Bound 0 *) where - "subst0 t T = T" - | "subst0 t F = F" - | "subst0 t (Lt a) = Lt (numsubst0 t a)" - | "subst0 t (Le a) = Le (numsubst0 t a)" - | "subst0 t (Gt a) = Gt (numsubst0 t a)" - | "subst0 t (Ge a) = Ge (numsubst0 t a)" - | "subst0 t (Eq a) = Eq (numsubst0 t a)" - | "subst0 t (NEq a) = NEq (numsubst0 t a)" - | "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)" - | "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)" - | "subst0 t (NOT p) = NOT (subst0 t p)" - | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)" - | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)" - | "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)" - | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)" - -lemma subst0_I: assumes qfp: "qfree p" - shows "Ifm (b#bs) (subst0 a p) = Ifm ((Inum (b#bs) a)#bs) p" - using qfp numsubst0_I[where b="b" and bs="bs" and a="a"] - by (induct p) (simp_all add: nth_pos2 ) - -consts - decrnum:: "num \ num" - decr :: "fm \ fm" - -recdef decrnum "measure size" - "decrnum (Bound n) = Bound (n - 1)" - "decrnum (Neg a) = Neg (decrnum a)" - "decrnum (Add a b) = Add (decrnum a) (decrnum b)" - "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)" - "decrnum (Mul c a) = Mul c (decrnum a)" - "decrnum (Floor a) = Floor (decrnum a)" - "decrnum (CN n c a) = CN (n - 1) c (decrnum a)" - "decrnum (CF c a b) = CF c (decrnum a) (decrnum b)" - "decrnum a = a" - -recdef decr "measure size" - "decr (Lt a) = Lt (decrnum a)" - "decr (Le a) = Le (decrnum a)" - "decr (Gt a) = Gt (decrnum a)" - "decr (Ge a) = Ge (decrnum a)" - "decr (Eq a) = Eq (decrnum a)" - "decr (NEq a) = NEq (decrnum a)" - "decr (Dvd i a) = Dvd i (decrnum a)" - "decr (NDvd i a) = NDvd i (decrnum a)" - "decr (NOT p) = NOT (decr p)" - "decr (And p q) = And (decr p) (decr q)" - "decr (Or p q) = Or (decr p) (decr q)" - "decr (Imp p q) = Imp (decr p) (decr q)" - "decr (Iff p q) = Iff (decr p) (decr q)" - "decr p = p" - -lemma decrnum: assumes nb: "numbound0 t" - shows "Inum (x#bs) t = Inum bs (decrnum t)" - using nb by (induct t rule: decrnum.induct, simp_all add: nth_pos2) - -lemma decr: assumes nb: "bound0 p" - shows "Ifm (x#bs) p = Ifm bs (decr p)" - using nb - by (induct p rule: decr.induct, simp_all add: nth_pos2 decrnum) - -lemma decr_qf: "bound0 p \ qfree (decr p)" -by (induct p, simp_all) - -consts - isatom :: "fm \ bool" (* test for atomicity *) -recdef isatom "measure size" - "isatom T = True" - "isatom F = True" - "isatom (Lt a) = True" - "isatom (Le a) = True" - "isatom (Gt a) = True" - "isatom (Ge a) = True" - "isatom (Eq a) = True" - "isatom (NEq a) = True" - "isatom (Dvd i b) = True" - "isatom (NDvd i b) = True" - "isatom p = False" - -lemma numsubst0_numbound0: assumes nb: "numbound0 t" - shows "numbound0 (numsubst0 t a)" -using nb by (induct a, auto) - -lemma subst0_bound0: assumes qf: "qfree p" and nb: "numbound0 t" - shows "bound0 (subst0 t p)" -using qf numsubst0_numbound0[OF nb] by (induct p, auto) - -lemma bound0_qf: "bound0 p \ qfree p" -by (induct p, simp_all) - - -definition djf:: "('a \ fm) \ 'a \ fm \ fm" where - "djf f p q = (if q=T then T else if q=F then f p else - (let fp = f p in case fp of T \ T | F \ q | _ \ Or fp q))" - -definition evaldjf:: "('a \ fm) \ 'a list \ fm" where - "evaldjf f ps = foldr (djf f) ps F" - -lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)" -by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) -(cases "f p", simp_all add: Let_def djf_def) - -lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\ p \ set ps. Ifm bs (f p))" - by(induct ps, simp_all add: evaldjf_def djf_Or) - -lemma evaldjf_bound0: - assumes nb: "\ x\ set xs. bound0 (f x)" - shows "bound0 (evaldjf f xs)" - using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) - -lemma evaldjf_qf: - assumes nb: "\ x\ set xs. qfree (f x)" - shows "qfree (evaldjf f xs)" - using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) - -consts - disjuncts :: "fm \ fm list" - conjuncts :: "fm \ fm list" -recdef disjuncts "measure size" - "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)" - "disjuncts F = []" - "disjuncts p = [p]" - -recdef conjuncts "measure size" - "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)" - "conjuncts T = []" - "conjuncts p = [p]" -lemma disjuncts: "(\ q\ set (disjuncts p). Ifm bs q) = Ifm bs p" -by(induct p rule: disjuncts.induct, auto) -lemma conjuncts: "(\ q\ set (conjuncts p). Ifm bs q) = Ifm bs p" -by(induct p rule: conjuncts.induct, auto) - -lemma disjuncts_nb: "bound0 p \ \ q\ set (disjuncts p). bound0 q" -proof- - assume nb: "bound0 p" - hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto) - thus ?thesis by (simp only: list_all_iff) -qed -lemma conjuncts_nb: "bound0 p \ \ q\ set (conjuncts p). bound0 q" -proof- - assume nb: "bound0 p" - hence "list_all bound0 (conjuncts p)" by (induct p rule:conjuncts.induct,auto) - thus ?thesis by (simp only: list_all_iff) -qed - -lemma disjuncts_qf: "qfree p \ \ q\ set (disjuncts p). qfree q" -proof- - assume qf: "qfree p" - hence "list_all qfree (disjuncts p)" - by (induct p rule: disjuncts.induct, auto) - thus ?thesis by (simp only: list_all_iff) -qed -lemma conjuncts_qf: "qfree p \ \ q\ set (conjuncts p). qfree q" -proof- - assume qf: "qfree p" - hence "list_all qfree (conjuncts p)" - by (induct p rule: conjuncts.induct, auto) - thus ?thesis by (simp only: list_all_iff) -qed - -constdefs DJ :: "(fm \ fm) \ fm \ fm" - "DJ f p \ evaldjf f (disjuncts p)" - -lemma DJ: assumes fdj: "\ p q. f (Or p q) = Or (f p) (f q)" - and fF: "f F = F" - shows "Ifm bs (DJ f p) = Ifm bs (f p)" -proof- - have "Ifm bs (DJ f p) = (\ q \ set (disjuncts p). Ifm bs (f q))" - by (simp add: DJ_def evaldjf_ex) - also have "\ = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto) - finally show ?thesis . -qed - -lemma DJ_qf: assumes - fqf: "\ p. qfree p \ qfree (f p)" - shows "\p. qfree p \ qfree (DJ f p) " -proof(clarify) - fix p assume qf: "qfree p" - have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def) - from disjuncts_qf[OF qf] have "\ q\ set (disjuncts p). qfree q" . - with fqf have th':"\ q\ set (disjuncts p). qfree (f q)" by blast - - from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp -qed - -lemma DJ_qe: assumes qe: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" - shows "\ bs p. qfree p \ qfree (DJ qe p) \ (Ifm bs ((DJ qe p)) = Ifm bs (E p))" -proof(clarify) - fix p::fm and bs - assume qf: "qfree p" - from qe have qth: "\ p. qfree p \ qfree (qe p)" by blast - from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto - have "Ifm bs (DJ qe p) = (\ q\ set (disjuncts p). Ifm bs (qe q))" - by (simp add: DJ_def evaldjf_ex) - also have "\ = (\ q \ set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto - also have "\ = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto) - finally show "qfree (DJ qe p) \ Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast -qed - (* Simplification *) - - (* Algebraic simplifications for nums *) -consts bnds:: "num \ nat list" - lex_ns:: "nat list \ nat list \ bool" -recdef bnds "measure size" - "bnds (Bound n) = [n]" - "bnds (CN n c a) = n#(bnds a)" - "bnds (Neg a) = bnds a" - "bnds (Add a b) = (bnds a)@(bnds b)" - "bnds (Sub a b) = (bnds a)@(bnds b)" - "bnds (Mul i a) = bnds a" - "bnds (Floor a) = bnds a" - "bnds (CF c a b) = (bnds a)@(bnds b)" - "bnds a = []" -recdef lex_ns "measure (\ (xs,ys). length xs + length ys)" - "lex_ns ([], ms) = True" - "lex_ns (ns, []) = False" - "lex_ns (n#ns, m#ms) = (n ((n = m) \ lex_ns (ns,ms))) " -constdefs lex_bnd :: "num \ num \ bool" - "lex_bnd t s \ lex_ns (bnds t, bnds s)" - -consts - numgcdh:: "num \ int \ int" - reducecoeffh:: "num \ int \ num" - dvdnumcoeff:: "num \ int \ bool" -consts maxcoeff:: "num \ int" -recdef maxcoeff "measure size" - "maxcoeff (C i) = abs i" - "maxcoeff (CN n c t) = max (abs c) (maxcoeff t)" - "maxcoeff (CF c t s) = max (abs c) (maxcoeff s)" - "maxcoeff t = 1" - -lemma maxcoeff_pos: "maxcoeff t \ 0" - apply (induct t rule: maxcoeff.induct, auto) - done - -recdef numgcdh "measure size" - "numgcdh (C i) = (\g. zgcd i g)" - "numgcdh (CN n c t) = (\g. zgcd c (numgcdh t g))" - "numgcdh (CF c s t) = (\g. zgcd c (numgcdh t g))" - "numgcdh t = (\g. 1)" - -definition - numgcd :: "num \ int" -where - numgcd_def: "numgcd t = numgcdh t (maxcoeff t)" - -recdef reducecoeffh "measure size" - "reducecoeffh (C i) = (\ g. C (i div g))" - "reducecoeffh (CN n c t) = (\ g. CN n (c div g) (reducecoeffh t g))" - "reducecoeffh (CF c s t) = (\ g. CF (c div g) s (reducecoeffh t g))" - "reducecoeffh t = (\g. t)" - -definition - reducecoeff :: "num \ num" -where - reducecoeff_def: "reducecoeff t = - (let g = numgcd t in - if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)" - -recdef dvdnumcoeff "measure size" - "dvdnumcoeff (C i) = (\ g. g dvd i)" - "dvdnumcoeff (CN n c t) = (\ g. g dvd c \ (dvdnumcoeff t g))" - "dvdnumcoeff (CF c s t) = (\ g. g dvd c \ (dvdnumcoeff t g))" - "dvdnumcoeff t = (\g. False)" - -lemma dvdnumcoeff_trans: - assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'" - shows "dvdnumcoeff t g" - using dgt' gdg - by (induct t rule: dvdnumcoeff.induct, simp_all add: gdg zdvd_trans[OF gdg]) - -declare zdvd_trans [trans add] - -lemma natabs0: "(nat (abs x) = 0) = (x = 0)" -by arith - -lemma numgcd0: - assumes g0: "numgcd t = 0" - shows "Inum bs t = 0" -proof- - have "\x. numgcdh t x= 0 \ Inum bs t = 0" - by (induct t rule: numgcdh.induct, auto simp add: zgcd_def gcd_zero natabs0 max_def maxcoeff_pos) - thus ?thesis using g0[simplified numgcd_def] by blast -qed - -lemma numgcdh_pos: assumes gp: "g \ 0" shows "numgcdh t g \ 0" - using gp - by (induct t rule: numgcdh.induct, auto simp add: zgcd_def) - -lemma numgcd_pos: "numgcd t \0" - by (simp add: numgcd_def numgcdh_pos maxcoeff_pos) - -lemma reducecoeffh: - assumes gt: "dvdnumcoeff t g" and gp: "g > 0" - shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t" - using gt -proof(induct t rule: reducecoeffh.induct) - case (1 i) hence gd: "g dvd i" by simp - from gp have gnz: "g \ 0" by simp - from prems show ?case by (simp add: real_of_int_div[OF gnz gd]) -next - case (2 n c t) hence gd: "g dvd c" by simp - from gp have gnz: "g \ 0" by simp - from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps) -next - case (3 c s t) hence gd: "g dvd c" by simp - from gp have gnz: "g \ 0" by simp - from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps) -qed (auto simp add: numgcd_def gp) -consts ismaxcoeff:: "num \ int \ bool" -recdef ismaxcoeff "measure size" - "ismaxcoeff (C i) = (\ x. abs i \ x)" - "ismaxcoeff (CN n c t) = (\x. abs c \ x \ (ismaxcoeff t x))" - "ismaxcoeff (CF c s t) = (\x. abs c \ x \ (ismaxcoeff t x))" - "ismaxcoeff t = (\x. True)" - -lemma ismaxcoeff_mono: "ismaxcoeff t c \ c \ c' \ ismaxcoeff t c'" -by (induct t rule: ismaxcoeff.induct, auto) - -lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)" -proof (induct t rule: maxcoeff.induct) - case (2 n c t) - hence H:"ismaxcoeff t (maxcoeff t)" . - have thh: "maxcoeff t \ max (abs c) (maxcoeff t)" by (simp add: le_maxI2) - from ismaxcoeff_mono[OF H thh] show ?case by (simp add: le_maxI1) -next - case (3 c t s) - hence H1:"ismaxcoeff s (maxcoeff s)" by auto - have thh1: "maxcoeff s \ max \c\ (maxcoeff s)" by (simp add: max_def) - from ismaxcoeff_mono[OF H1 thh1] show ?case by (simp add: le_maxI1) -qed simp_all - -lemma zgcd_gt1: "zgcd i j > 1 \ ((abs i > 1 \ abs j > 1) \ (abs i = 0 \ abs j > 1) \ (abs i > 1 \ abs j = 0))" - apply (unfold zgcd_def) - apply (cases "i = 0", simp_all) - apply (cases "j = 0", simp_all) - apply (cases "abs i = 1", simp_all) - apply (cases "abs j = 1", simp_all) - apply auto - done -lemma numgcdh0:"numgcdh t m = 0 \ m =0" - by (induct t rule: numgcdh.induct, auto simp add:zgcd0) - -lemma dvdnumcoeff_aux: - assumes "ismaxcoeff t m" and mp:"m \ 0" and "numgcdh t m > 1" - shows "dvdnumcoeff t (numgcdh t m)" -using prems -proof(induct t rule: numgcdh.induct) - case (2 n c t) - let ?g = "numgcdh t m" - from prems have th:"zgcd c ?g > 1" by simp - from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] - have "(abs c > 1 \ ?g > 1) \ (abs c = 0 \ ?g > 1) \ (abs c > 1 \ ?g = 0)" by simp - moreover {assume "abs c > 1" and gp: "?g > 1" with prems - have th: "dvdnumcoeff t ?g" by simp - have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) - from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)} - moreover {assume "abs c = 0 \ ?g > 1" - with prems have th: "dvdnumcoeff t ?g" by simp - have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) - from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1) - hence ?case by simp } - moreover {assume "abs c > 1" and g0:"?g = 0" - from numgcdh0[OF g0] have "m=0". with prems have ?case by simp } - ultimately show ?case by blast -next - case (3 c s t) - let ?g = "numgcdh t m" - from prems have th:"zgcd c ?g > 1" by simp - from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] - have "(abs c > 1 \ ?g > 1) \ (abs c = 0 \ ?g > 1) \ (abs c > 1 \ ?g = 0)" by simp - moreover {assume "abs c > 1" and gp: "?g > 1" with prems - have th: "dvdnumcoeff t ?g" by simp - have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) - from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)} - moreover {assume "abs c = 0 \ ?g > 1" - with prems have th: "dvdnumcoeff t ?g" by simp - have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) - from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1) - hence ?case by simp } - moreover {assume "abs c > 1" and g0:"?g = 0" - from numgcdh0[OF g0] have "m=0". with prems have ?case by simp } - ultimately show ?case by blast -qed(auto simp add: zgcd_zdvd1) - -lemma dvdnumcoeff_aux2: - assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \ numgcd t > 0" - using prems -proof (simp add: numgcd_def) - let ?mc = "maxcoeff t" - let ?g = "numgcdh t ?mc" - have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff) - have th2: "?mc \ 0" by (rule maxcoeff_pos) - assume H: "numgcdh t ?mc > 1" - from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" . -qed - -lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" -proof- - let ?g = "numgcd t" - have "?g \ 0" by (simp add: numgcd_pos) - hence "?g = 0 \ ?g = 1 \ ?g > 1" by auto - moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} - moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} - moreover { assume g1:"?g > 1" - from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+ - from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis - by (simp add: reducecoeff_def Let_def)} - ultimately show ?thesis by blast -qed - -lemma reducecoeffh_numbound0: "numbound0 t \ numbound0 (reducecoeffh t g)" -by (induct t rule: reducecoeffh.induct, auto) - -lemma reducecoeff_numbound0: "numbound0 t \ numbound0 (reducecoeff t)" -using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def) - -consts - simpnum:: "num \ num" - numadd:: "num \ num \ num" - nummul:: "num \ int \ num" - -recdef numadd "measure (\ (t,s). size t + size s)" - "numadd (CN n1 c1 r1,CN n2 c2 r2) = - (if n1=n2 then - (let c = c1 + c2 - in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2)))) - else if n1 \ n2 then CN n1 c1 (numadd (r1,CN n2 c2 r2)) - else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))" - "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))" - "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" - "numadd (CF c1 t1 r1,CF c2 t2 r2) = - (if t1 = t2 then - (let c=c1+c2; s= numadd(r1,r2) in (if c=0 then s else CF c t1 s)) - else if lex_bnd t1 t2 then CF c1 t1 (numadd(r1,CF c2 t2 r2)) - else CF c2 t2 (numadd(CF c1 t1 r1,r2)))" - "numadd (CF c1 t1 r1,C c) = CF c1 t1 (numadd (r1, C c))" - "numadd (C c,CF c1 t1 r1) = CF c1 t1 (numadd (r1, C c))" - "numadd (C b1, C b2) = C (b1+b2)" - "numadd (a,b) = Add a b" - -lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)" -apply (induct t s rule: numadd.induct, simp_all add: Let_def) - apply (case_tac "c1+c2 = 0",case_tac "n1 \ n2", simp_all) - apply (case_tac "n1 = n2", simp_all add: ring_simps) - apply (simp only: left_distrib[symmetric]) - apply simp -apply (case_tac "lex_bnd t1 t2", simp_all) - apply (case_tac "c1+c2 = 0") - by (case_tac "t1 = t2", simp_all add: ring_simps left_distrib[symmetric] real_of_int_mult[symmetric] real_of_int_add[symmetric]del: real_of_int_mult real_of_int_add left_distrib) - -lemma numadd_nb[simp]: "\ numbound0 t ; numbound0 s\ \ numbound0 (numadd (t,s))" -by (induct t s rule: numadd.induct, auto simp add: Let_def) - -recdef nummul "measure size" - "nummul (C j) = (\ i. C (i*j))" - "nummul (CN n c t) = (\ i. CN n (c*i) (nummul t i))" - "nummul (CF c t s) = (\ i. CF (c*i) t (nummul s i))" - "nummul (Mul c t) = (\ i. nummul t (i*c))" - "nummul t = (\ i. Mul i t)" - -lemma nummul[simp]: "\ i. Inum bs (nummul t i) = Inum bs (Mul i t)" -by (induct t rule: nummul.induct, auto simp add: ring_simps) - -lemma nummul_nb[simp]: "\ i. numbound0 t \ numbound0 (nummul t i)" -by (induct t rule: nummul.induct, auto) - -constdefs numneg :: "num \ num" - "numneg t \ nummul t (- 1)" - -constdefs numsub :: "num \ num \ num" - "numsub s t \ (if s = t then C 0 else numadd (s,numneg t))" - -lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)" -using numneg_def nummul by simp - -lemma numneg_nb[simp]: "numbound0 t \ numbound0 (numneg t)" -using numneg_def by simp - -lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)" -using numsub_def by simp - -lemma numsub_nb[simp]: "\ numbound0 t ; numbound0 s\ \ numbound0 (numsub t s)" -using numsub_def by simp - -lemma isint_CF: assumes si: "isint s bs" shows "isint (CF c t s) bs" -proof- - have cti: "isint (Mul c (Floor t)) bs" by (simp add: isint_Mul isint_Floor) - - have "?thesis = isint (Add (Mul c (Floor t)) s) bs" by (simp add: isint_def) - also have "\" by (simp add: isint_add cti si) - finally show ?thesis . -qed - -consts split_int:: "num \ num\num" -recdef split_int "measure num_size" - "split_int (C c) = (C 0, C c)" - "split_int (CN n c b) = - (let (bv,bi) = split_int b - in (CN n c bv, bi))" - "split_int (CF c a b) = - (let (bv,bi) = split_int b - in (bv, CF c a bi))" - "split_int a = (a,C 0)" - -lemma split_int:"\ tv ti. split_int t = (tv,ti) \ (Inum bs (Add tv ti) = Inum bs t) \ isint ti bs" -proof (induct t rule: split_int.induct) - case (2 c n b tv ti) - let ?bv = "fst (split_int b)" - let ?bi = "snd (split_int b)" - have "split_int b = (?bv,?bi)" by simp - with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+ - from prems(2) have tibi: "ti = ?bi" by (simp add: Let_def split_def) - from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def) -next - case (3 c a b tv ti) - let ?bv = "fst (split_int b)" - let ?bi = "snd (split_int b)" - have "split_int b = (?bv,?bi)" by simp - with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+ - from prems(2) have tibi: "ti = CF c a ?bi" by (simp add: Let_def split_def) - from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def isint_Floor isint_add isint_Mul isint_CF) -qed (auto simp add: Let_def isint_iff isint_Floor isint_add isint_Mul split_def ring_simps) - -lemma split_int_nb: "numbound0 t \ numbound0 (fst (split_int t)) \ numbound0 (snd (split_int t)) " -by (induct t rule: split_int.induct, auto simp add: Let_def split_def) - -definition - numfloor:: "num \ num" -where - numfloor_def: "numfloor t = (let (tv,ti) = split_int t in - (case tv of C i \ numadd (tv,ti) - | _ \ numadd(CF 1 tv (C 0),ti)))" - -lemma numfloor[simp]: "Inum bs (numfloor t) = Inum bs (Floor t)" (is "?n t = ?N (Floor t)") -proof- - let ?tv = "fst (split_int t)" - let ?ti = "snd (split_int t)" - have tvti:"split_int t = (?tv,?ti)" by simp - {assume H: "\ v. ?tv \ C v" - hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)" - by (cases ?tv, auto simp add: numfloor_def Let_def split_def numadd) - from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ - hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp - also have "\ = real (floor (?N ?tv) + (floor (?N ?ti)))" - by (simp,subst tii[simplified isint_iff, symmetric]) simp - also have "\ = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) - finally have ?thesis using th1 by simp} - moreover {fix v assume H:"?tv = C v" - from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ - hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp - also have "\ = real (floor (?N ?tv) + (floor (?N ?ti)))" - by (simp,subst tii[simplified isint_iff, symmetric]) simp - also have "\ = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) - finally have ?thesis by (simp add: H numfloor_def Let_def split_def numadd) } - ultimately show ?thesis by auto -qed - -lemma numfloor_nb[simp]: "numbound0 t \ numbound0 (numfloor t)" - using split_int_nb[where t="t"] - by (cases "fst(split_int t)" , auto simp add: numfloor_def Let_def split_def numadd_nb) - -recdef simpnum "measure num_size" - "simpnum (C j) = C j" - "simpnum (Bound n) = CN n 1 (C 0)" - "simpnum (Neg t) = numneg (simpnum t)" - "simpnum (Add t s) = numadd (simpnum t,simpnum s)" - "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)" - "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)" - "simpnum (Floor t) = numfloor (simpnum t)" - "simpnum (CN n c t) = (if c=0 then simpnum t else CN n c (simpnum t))" - "simpnum (CF c t s) = simpnum(Add (Mul c (Floor t)) s)" - -lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t" -by (induct t rule: simpnum.induct, auto) - -lemma simpnum_numbound0[simp]: - "numbound0 t \ numbound0 (simpnum t)" -by (induct t rule: simpnum.induct, auto) - -consts nozerocoeff:: "num \ bool" -recdef nozerocoeff "measure size" - "nozerocoeff (C c) = True" - "nozerocoeff (CN n c t) = (c\0 \ nozerocoeff t)" - "nozerocoeff (CF c s t) = (c \ 0 \ nozerocoeff t)" - "nozerocoeff (Mul c t) = (c\0 \ nozerocoeff t)" - "nozerocoeff t = True" - -lemma numadd_nz : "nozerocoeff a \ nozerocoeff b \ nozerocoeff (numadd (a,b))" -by (induct a b rule: numadd.induct,auto simp add: Let_def) - -lemma nummul_nz : "\ i. i\0 \ nozerocoeff a \ nozerocoeff (nummul a i)" - by (induct a rule: nummul.induct,auto simp add: Let_def numadd_nz) - -lemma numneg_nz : "nozerocoeff a \ nozerocoeff (numneg a)" -by (simp add: numneg_def nummul_nz) - -lemma numsub_nz: "nozerocoeff a \ nozerocoeff b \ nozerocoeff (numsub a b)" -by (simp add: numsub_def numneg_nz numadd_nz) - -lemma split_int_nz: "nozerocoeff t \ nozerocoeff (fst (split_int t)) \ nozerocoeff (snd (split_int t))" -by (induct t rule: split_int.induct,auto simp add: Let_def split_def) - -lemma numfloor_nz: "nozerocoeff t \ nozerocoeff (numfloor t)" -by (simp add: numfloor_def Let_def split_def) -(cases "fst (split_int t)", simp_all add: split_int_nz numadd_nz) - -lemma simpnum_nz: "nozerocoeff (simpnum t)" -by(induct t rule: simpnum.induct, auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz numfloor_nz) - -lemma maxcoeff_nz: "nozerocoeff t \ maxcoeff t = 0 \ t = C 0" -proof (induct t rule: maxcoeff.induct) - case (2 n c t) - hence cnz: "c \0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+ - have "max (abs c) (maxcoeff t) \ abs c" by (simp add: le_maxI1) - with cnz have "max (abs c) (maxcoeff t) > 0" by arith - with prems show ?case by simp -next - case (3 c s t) - hence cnz: "c \0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+ - have "max (abs c) (maxcoeff t) \ abs c" by (simp add: le_maxI1) - with cnz have "max (abs c) (maxcoeff t) > 0" by arith - with prems show ?case by simp -qed auto - -lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0" -proof- - from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def) - from numgcdh0[OF th] have th:"maxcoeff t = 0" . - from maxcoeff_nz[OF nz th] show ?thesis . -qed - -constdefs simp_num_pair:: "(num \ int) \ num \ int" - "simp_num_pair \ (\ (t,n). (if n = 0 then (C 0, 0) else - (let t' = simpnum t ; g = numgcd t' in - if g > 1 then (let g' = zgcd n g in - if g' = 1 then (t',n) - else (reducecoeffh t' g', n div g')) - else (t',n))))" - -lemma simp_num_pair_ci: - shows "((\ (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((\ (t,n). Inum bs t / real n) (t,n))" - (is "?lhs = ?rhs") -proof- - let ?t' = "simpnum t" - let ?g = "numgcd ?t'" - let ?g' = "zgcd n ?g" - {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)} - moreover - { assume nnz: "n \ 0" - {assume "\ ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def)} - moreover - {assume g1:"?g>1" hence g0: "?g > 0" by simp - from zgcd0 g1 nnz have gp0: "?g' \ 0" by simp - hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith - hence "?g'= 1 \ ?g' > 1" by arith - moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)} - moreover {assume g'1:"?g'>1" - from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" .. - let ?tt = "reducecoeffh ?t' ?g'" - let ?t = "Inum bs ?tt" - have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) - have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1) - have gpdgp: "?g' dvd ?g'" by simp - from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] - have th2:"real ?g' * ?t = Inum bs ?t'" by simp - from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def) - also have "\ = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp - also have "\ = (Inum bs ?t' / real n)" - using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp - finally have "?lhs = Inum bs t / real n" by simp - then have ?thesis using prems by (simp add: simp_num_pair_def)} - ultimately have ?thesis by blast} - ultimately have ?thesis by blast} - ultimately show ?thesis by blast -qed - -lemma simp_num_pair_l: assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')" - shows "numbound0 t' \ n' >0" -proof- - let ?t' = "simpnum t" - let ?g = "numgcd ?t'" - let ?g' = "zgcd n ?g" - {assume nz: "n = 0" hence ?thesis using prems by (simp add: Let_def simp_num_pair_def)} - moreover - { assume nnz: "n \ 0" - {assume "\ ?g > 1" hence ?thesis using prems by (auto simp add: Let_def simp_num_pair_def)} - moreover - {assume g1:"?g>1" hence g0: "?g > 0" by simp - from zgcd0 g1 nnz have gp0: "?g' \ 0" by simp - hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith - hence "?g'= 1 \ ?g' > 1" by arith - moreover {assume "?g'=1" hence ?thesis using prems - by (auto simp add: Let_def simp_num_pair_def)} - moreover {assume g'1:"?g'>1" - have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) - have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1) - have gpdgp: "?g' dvd ?g'" by simp - from zdvd_imp_le[OF gpdd np] have g'n: "?g' \ n" . - from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]] - have "n div ?g' >0" by simp - hence ?thesis using prems - by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)} - ultimately have ?thesis by blast} - ultimately have ?thesis by blast} - ultimately show ?thesis by blast -qed - -consts not:: "fm \ fm" -recdef not "measure size" - "not (NOT p) = p" - "not T = F" - "not F = T" - "not (Lt t) = Ge t" - "not (Le t) = Gt t" - "not (Gt t) = Le t" - "not (Ge t) = Lt t" - "not (Eq t) = NEq t" - "not (NEq t) = Eq t" - "not (Dvd i t) = NDvd i t" - "not (NDvd i t) = Dvd i t" - "not (And p q) = Or (not p) (not q)" - "not (Or p q) = And (not p) (not q)" - "not p = NOT p" -lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)" -by (induct p) auto -lemma not_qf[simp]: "qfree p \ qfree (not p)" -by (induct p, auto) -lemma not_nb[simp]: "bound0 p \ bound0 (not p)" -by (induct p, auto) - -constdefs conj :: "fm \ fm \ fm" - "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else - if p = q then p else And p q)" -lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)" -by (cases "p=F \ q=F",simp_all add: conj_def) (cases p,simp_all) - -lemma conj_qf[simp]: "\qfree p ; qfree q\ \ qfree (conj p q)" -using conj_def by auto -lemma conj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (conj p q)" -using conj_def by auto - -constdefs disj :: "fm \ fm \ fm" - "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p - else if p=q then p else Or p q)" - -lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)" -by (cases "p=T \ q=T",simp_all add: disj_def) (cases p,simp_all) -lemma disj_qf[simp]: "\qfree p ; qfree q\ \ qfree (disj p q)" -using disj_def by auto -lemma disj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (disj p q)" -using disj_def by auto - -constdefs imp :: "fm \ fm \ fm" - "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p - else Imp p q)" -lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)" -by (cases "p=F \ q=T",simp_all add: imp_def) -lemma imp_qf[simp]: "\qfree p ; qfree q\ \ qfree (imp p q)" -using imp_def by (cases "p=F \ q=T",simp_all add: imp_def) -lemma imp_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (imp p q)" -using imp_def by (cases "p=F \ q=T \ p=q",simp_all add: imp_def) - -constdefs iff :: "fm \ fm \ fm" - "iff p q \ (if (p = q) then T else if (p = not q \ not p = q) then F else - if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else - Iff p q)" -lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)" - by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp add:not) -(cases "not p= q", auto simp add:not) -lemma iff_qf[simp]: "\qfree p ; qfree q\ \ qfree (iff p q)" - by (unfold iff_def,cases "p=q", auto) -lemma iff_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (iff p q)" -using iff_def by (unfold iff_def,cases "p=q", auto) - -consts check_int:: "num \ bool" -recdef check_int "measure size" - "check_int (C i) = True" - "check_int (Floor t) = True" - "check_int (Mul i t) = check_int t" - "check_int (Add t s) = (check_int t \ check_int s)" - "check_int (Neg t) = check_int t" - "check_int (CF c t s) = check_int s" - "check_int t = False" -lemma check_int: "check_int t \ isint t bs" -by (induct t, auto simp add: isint_add isint_Floor isint_Mul isint_neg isint_c isint_CF) - -lemma rdvd_left1_int: "real \t\ = t \ 1 rdvd t" - by (simp add: rdvd_def,rule_tac x="\t\" in exI) simp - -lemma rdvd_reduce: - assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0" - shows "real (d::int) rdvd real (c::int)*t = (real (d div g) rdvd real (c div g)*t)" -proof - assume d: "real d rdvd real c * t" - from d rdvd_def obtain k where k_def: "real c * t = real d* real (k::int)" by auto - from gd dvd_def obtain kd where kd_def: "d = g * kd" by auto - from gc dvd_def obtain kc where kc_def: "c = g * kc" by auto - from k_def kd_def kc_def have "real g * real kc * t = real g * real kd * real k" by simp - hence "real kc * t = real kd * real k" using gp by simp - hence th:"real kd rdvd real kc * t" using rdvd_def by blast - from kd_def gp have th':"kd = d div g" by simp - from kc_def gp have "kc = c div g" by simp - with th th' show "real (d div g) rdvd real (c div g) * t" by simp -next - assume d: "real (d div g) rdvd real (c div g) * t" - from gp have gnz: "g \ 0" by simp - thus "real d rdvd real c * t" using d rdvd_mult[OF gnz, where n="d div g" and x="real (c div g) * t"] real_of_int_div[OF gnz gd] real_of_int_div[OF gnz gc] by simp -qed - -constdefs simpdvd:: "int \ num \ (int \ num)" - "simpdvd d t \ - (let g = numgcd t in - if g > 1 then (let g' = zgcd d g in - if g' = 1 then (d, t) - else (d div g',reducecoeffh t g')) - else (d, t))" -lemma simpdvd: - assumes tnz: "nozerocoeff t" and dnz: "d \ 0" - shows "Ifm bs (Dvd (fst (simpdvd d t)) (snd (simpdvd d t))) = Ifm bs (Dvd d t)" -proof- - let ?g = "numgcd t" - let ?g' = "zgcd d ?g" - {assume "\ ?g > 1" hence ?thesis by (simp add: Let_def simpdvd_def)} - moreover - {assume g1:"?g>1" hence g0: "?g > 0" by simp - from zgcd0 g1 dnz have gp0: "?g' \ 0" by simp - hence g'p: "?g' > 0" using zgcd_pos[where i="d" and j="numgcd t"] by arith - hence "?g'= 1 \ ?g' > 1" by arith - moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simpdvd_def)} - moreover {assume g'1:"?g'>1" - from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" .. - let ?tt = "reducecoeffh t ?g'" - let ?t = "Inum bs ?tt" - have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) - have gpdd: "?g' dvd d" by (simp add: zgcd_zdvd1) - have gpdgp: "?g' dvd ?g'" by simp - from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] - have th2:"real ?g' * ?t = Inum bs t" by simp - from prems have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)" - by (simp add: simpdvd_def Let_def) - also have "\ = (real d rdvd (Inum bs t))" - using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified zdiv_self[OF gp0]] - th2[symmetric] by simp - finally have ?thesis by simp } - ultimately have ?thesis by blast - } - ultimately show ?thesis by blast -qed - -consts simpfm :: "fm \ fm" -recdef simpfm "measure fmsize" - "simpfm (And p q) = conj (simpfm p) (simpfm q)" - "simpfm (Or p q) = disj (simpfm p) (simpfm q)" - "simpfm (Imp p q) = imp (simpfm p) (simpfm q)" - "simpfm (Iff p q) = iff (simpfm p) (simpfm q)" - "simpfm (NOT p) = not (simpfm p)" - "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \ if (v < 0) then T else F - | _ \ Lt (reducecoeff a'))" - "simpfm (Le a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Le (reducecoeff a'))" - "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \ if (v > 0) then T else F | _ \ Gt (reducecoeff a'))" - "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Ge (reducecoeff a'))" - "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \ if (v = 0) then T else F | _ \ Eq (reducecoeff a'))" - "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ NEq (reducecoeff a'))" - "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a) - else if (abs i = 1) \ check_int a then T - else let a' = simpnum a in case a' of C v \ if (i dvd v) then T else F | _ \ (let (d,t) = simpdvd i a' in Dvd d t))" - "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) - else if (abs i = 1) \ check_int a then F - else let a' = simpnum a in case a' of C v \ if (\(i dvd v)) then T else F | _ \ (let (d,t) = simpdvd i a' in NDvd d t))" - "simpfm p = p" - -lemma simpfm[simp]: "Ifm bs (simpfm p) = Ifm bs p" -proof(induct p rule: simpfm.induct) - case (6 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume H:"\ (\ v. ?sa = C v)" - let ?g = "numgcd ?sa" - let ?rsa = "reducecoeff ?sa" - let ?r = "Inum bs ?rsa" - have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) - {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} - with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a < 0 = (real ?g * ?r < real ?g * 0)" by simp - also have "\ = (?r < 0)" using gp - by (simp only: mult_less_cancel_left) simp - finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} - ultimately show ?case by blast -next - case (7 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume H:"\ (\ v. ?sa = C v)" - let ?g = "numgcd ?sa" - let ?rsa = "reducecoeff ?sa" - let ?r = "Inum bs ?rsa" - have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) - {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} - with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a \ 0 = (real ?g * ?r \ real ?g * 0)" by simp - also have "\ = (?r \ 0)" using gp - by (simp only: mult_le_cancel_left) simp - finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} - ultimately show ?case by blast -next - case (8 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume H:"\ (\ v. ?sa = C v)" - let ?g = "numgcd ?sa" - let ?rsa = "reducecoeff ?sa" - let ?r = "Inum bs ?rsa" - have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) - {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} - with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a > 0 = (real ?g * ?r > real ?g * 0)" by simp - also have "\ = (?r > 0)" using gp - by (simp only: mult_less_cancel_left) simp - finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} - ultimately show ?case by blast -next - case (9 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume H:"\ (\ v. ?sa = C v)" - let ?g = "numgcd ?sa" - let ?rsa = "reducecoeff ?sa" - let ?r = "Inum bs ?rsa" - have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) - {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} - with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a \ 0 = (real ?g * ?r \ real ?g * 0)" by simp - also have "\ = (?r \ 0)" using gp - by (simp only: mult_le_cancel_left) simp - finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} - ultimately show ?case by blast -next - case (10 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume H:"\ (\ v. ?sa = C v)" - let ?g = "numgcd ?sa" - let ?rsa = "reducecoeff ?sa" - let ?r = "Inum bs ?rsa" - have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) - {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} - with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a = 0 = (real ?g * ?r = 0)" by simp - also have "\ = (?r = 0)" using gp - by (simp add: mult_eq_0_iff) - finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} - ultimately show ?case by blast -next - case (11 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume H:"\ (\ v. ?sa = C v)" - let ?g = "numgcd ?sa" - let ?rsa = "reducecoeff ?sa" - let ?r = "Inum bs ?rsa" - have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) - {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} - with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a \ 0 = (real ?g * ?r \ 0)" by simp - also have "\ = (?r \ 0)" using gp - by (simp add: mult_eq_0_iff) - finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} - ultimately show ?case by blast -next - case (12 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp - have "i=0 \ (abs i = 1 \ check_int a) \ (i\0 \ ((abs i \ 1) \ (\ check_int a)))" by auto - {assume "i=0" hence ?case using "12.hyps" by (simp add: rdvd_left_0_eq Let_def)} - moreover - {assume ai1: "abs i = 1" and ai: "check_int a" - hence "i=1 \ i= - 1" by arith - moreover {assume i1: "i = 1" - from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] - have ?case using i1 ai by simp } - moreover {assume i1: "i = - 1" - from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] - rdvd_abs1[where d="- 1" and t="Inum bs a"] - have ?case using i1 ai by simp } - ultimately have ?case by blast} - moreover - {assume inz: "i\0" and cond: "(abs i \ 1) \ (\ check_int a)" - {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond - by (cases "abs i = 1", auto simp add: int_rdvd_iff) } - moreover {assume H:"\ (\ v. ?sa = C v)" - hence th: "simpfm (Dvd i a) = Dvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond by (cases ?sa, auto simp add: Let_def split_def) - from simpnum_nz have nz:"nozerocoeff ?sa" by simp - from simpdvd [OF nz inz] th have ?case using sa by simp} - ultimately have ?case by blast} - ultimately show ?case by blast -next - case (13 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp - have "i=0 \ (abs i = 1 \ check_int a) \ (i\0 \ ((abs i \ 1) \ (\ check_int a)))" by auto - {assume "i=0" hence ?case using "13.hyps" by (simp add: rdvd_left_0_eq Let_def)} - moreover - {assume ai1: "abs i = 1" and ai: "check_int a" - hence "i=1 \ i= - 1" by arith - moreover {assume i1: "i = 1" - from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] - have ?case using i1 ai by simp } - moreover {assume i1: "i = - 1" - from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] - rdvd_abs1[where d="- 1" and t="Inum bs a"] - have ?case using i1 ai by simp } - ultimately have ?case by blast} - moreover - {assume inz: "i\0" and cond: "(abs i \ 1) \ (\ check_int a)" - {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond - by (cases "abs i = 1", auto simp add: int_rdvd_iff) } - moreover {assume H:"\ (\ v. ?sa = C v)" - hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond - by (cases ?sa, auto simp add: Let_def split_def) - from simpnum_nz have nz:"nozerocoeff ?sa" by simp - from simpdvd [OF nz inz] th have ?case using sa by simp} - ultimately have ?case by blast} - ultimately show ?case by blast -qed (induct p rule: simpfm.induct, simp_all) - -lemma simpdvd_numbound0: "numbound0 t \ numbound0 (snd (simpdvd d t))" - by (simp add: simpdvd_def Let_def split_def reducecoeffh_numbound0) - -lemma simpfm_bound0[simp]: "bound0 p \ bound0 (simpfm p)" -proof(induct p rule: simpfm.induct) - case (6 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) -next - case (7 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) -next - case (8 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) -next - case (9 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) -next - case (10 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) -next - case (11 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) -next - case (12 i a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def) -next - case (13 i a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def) -qed(auto simp add: disj_def imp_def iff_def conj_def) - -lemma simpfm_qf[simp]: "qfree p \ qfree (simpfm p)" -by (induct p rule: simpfm.induct, auto simp add: Let_def) -(case_tac "simpnum a",auto simp add: split_def Let_def)+ - - - (* Generic quantifier elimination *) - -constdefs list_conj :: "fm list \ fm" - "list_conj ps \ foldr conj ps T" -lemma list_conj: "Ifm bs (list_conj ps) = (\p\ set ps. Ifm bs p)" - by (induct ps, auto simp add: list_conj_def) -lemma list_conj_qf: " \p\ set ps. qfree p \ qfree (list_conj ps)" - by (induct ps, auto simp add: list_conj_def) -lemma list_conj_nb: " \p\ set ps. bound0 p \ bound0 (list_conj ps)" - by (induct ps, auto simp add: list_conj_def) -constdefs CJNB:: "(fm \ fm) \ fm \ fm" - "CJNB f p \ (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs - in conj (decr (list_conj yes)) (f (list_conj no)))" - -lemma CJNB_qe: - assumes qe: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" - shows "\ bs p. qfree p \ qfree (CJNB qe p) \ (Ifm bs ((CJNB qe p)) = Ifm bs (E p))" -proof(clarify) - fix bs p - assume qfp: "qfree p" - let ?cjs = "conjuncts p" - let ?yes = "fst (partition bound0 ?cjs)" - let ?no = "snd (partition bound0 ?cjs)" - let ?cno = "list_conj ?no" - let ?cyes = "list_conj ?yes" - have part: "partition bound0 ?cjs = (?yes,?no)" by simp - from partition_P[OF part] have "\ q\ set ?yes. bound0 q" by blast - hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb) - hence yes_qf: "qfree (decr ?cyes )" by (simp add: decr_qf) - from conjuncts_qf[OF qfp] partition_set[OF part] - have " \q\ set ?no. qfree q" by auto - hence no_qf: "qfree ?cno"by (simp add: list_conj_qf) - with qe have cno_qf:"qfree (qe ?cno )" - and noE: "Ifm bs (qe ?cno) = Ifm bs (E ?cno)" by blast+ - from cno_qf yes_qf have qf: "qfree (CJNB qe p)" - by (simp add: CJNB_def Let_def conj_qf split_def) - {fix bs - from conjuncts have "Ifm bs p = (\q\ set ?cjs. Ifm bs q)" by blast - also have "\ = ((\q\ set ?yes. Ifm bs q) \ (\q\ set ?no. Ifm bs q))" - using partition_set[OF part] by auto - finally have "Ifm bs p = ((Ifm bs ?cyes) \ (Ifm bs ?cno))" using list_conj by simp} - hence "Ifm bs (E p) = (\x. (Ifm (x#bs) ?cyes) \ (Ifm (x#bs) ?cno))" by simp - also fix y have "\ = (\x. (Ifm (y#bs) ?cyes) \ (Ifm (x#bs) ?cno))" - using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast - also have "\ = (Ifm bs (decr ?cyes) \ Ifm bs (E ?cno))" - by (auto simp add: decr[OF yes_nb]) - also have "\ = (Ifm bs (conj (decr ?cyes) (qe ?cno)))" - using qe[rule_format, OF no_qf] by auto - finally have "Ifm bs (E p) = Ifm bs (CJNB qe p)" - by (simp add: Let_def CJNB_def split_def) - with qf show "qfree (CJNB qe p) \ Ifm bs (CJNB qe p) = Ifm bs (E p)" by blast -qed - -consts qelim :: "fm \ (fm \ fm) \ fm" -recdef qelim "measure fmsize" - "qelim (E p) = (\ qe. DJ (CJNB qe) (qelim p qe))" - "qelim (A p) = (\ qe. not (qe ((qelim (NOT p) qe))))" - "qelim (NOT p) = (\ qe. not (qelim p qe))" - "qelim (And p q) = (\ qe. conj (qelim p qe) (qelim q qe))" - "qelim (Or p q) = (\ qe. disj (qelim p qe) (qelim q qe))" - "qelim (Imp p q) = (\ qe. disj (qelim (NOT p) qe) (qelim q qe))" - "qelim (Iff p q) = (\ qe. iff (qelim p qe) (qelim q qe))" - "qelim p = (\ y. simpfm p)" - -lemma qelim_ci: - assumes qe_inv: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" - shows "\ bs. qfree (qelim p qe) \ (Ifm bs (qelim p qe) = Ifm bs p)" -using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]] -by(induct p rule: qelim.induct) -(auto simp del: simpfm.simps) - - -text {* The @{text "\"} Part *} -text{* Linearity for fm where Bound 0 ranges over @{text "\"} *} -consts - zsplit0 :: "num \ int \ num" (* splits the bounded from the unbounded part*) -recdef zsplit0 "measure num_size" - "zsplit0 (C c) = (0,C c)" - "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))" - "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)" - "zsplit0 (CF c a b) = zsplit0 (Add (Mul c (Floor a)) b)" - "zsplit0 (Neg a) = (let (i',a') = zsplit0 a in (-i', Neg a'))" - "zsplit0 (Add a b) = (let (ia,a') = zsplit0 a ; - (ib,b') = zsplit0 b - in (ia+ib, Add a' b'))" - "zsplit0 (Sub a b) = (let (ia,a') = zsplit0 a ; - (ib,b') = zsplit0 b - in (ia-ib, Sub a' b'))" - "zsplit0 (Mul i a) = (let (i',a') = zsplit0 a in (i*i', Mul i a'))" - "zsplit0 (Floor a) = (let (i',a') = zsplit0 a in (i',Floor a'))" -(hints simp add: Let_def) - -lemma zsplit0_I: - shows "\ n a. zsplit0 t = (n,a) \ (Inum ((real (x::int)) #bs) (CN 0 n a) = Inum (real x #bs) t) \ numbound0 a" - (is "\ n a. ?S t = (n,a) \ (?I x (CN 0 n a) = ?I x t) \ ?N a") -proof(induct t rule: zsplit0.induct) - case (1 c n a) thus ?case by auto -next - case (2 m n a) thus ?case by (cases "m=0") auto -next - case (3 n i a n a') thus ?case by auto -next - case (4 c a b n a') thus ?case by auto -next - case (5 t n a) - let ?nt = "fst (zsplit0 t)" - let ?at = "snd (zsplit0 t)" - have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \ n=-?nt" using prems - by (simp add: Let_def split_def) - from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast - from th2[simplified] th[simplified] show ?case by simp -next - case (6 s t n a) - let ?ns = "fst (zsplit0 s)" - let ?as = "snd (zsplit0 s)" - let ?nt = "fst (zsplit0 t)" - let ?at = "snd (zsplit0 t)" - have abjs: "zsplit0 s = (?ns,?as)" by simp - moreover have abjt: "zsplit0 t = (?nt,?at)" by simp - ultimately have th: "a=Add ?as ?at \ n=?ns + ?nt" using prems - by (simp add: Let_def split_def) - from abjs[symmetric] have bluddy: "\ x y. (x,y) = zsplit0 s" by blast - from prems have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real x # bs) (CN 0 xa xb) = Inum (real x # bs) t \ numbound0 xb)" by simp - with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast - from abjs prems have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \ ?N ?as" by blast - from th3[simplified] th2[simplified] th[simplified] show ?case - by (simp add: left_distrib) -next - case (7 s t n a) - let ?ns = "fst (zsplit0 s)" - let ?as = "snd (zsplit0 s)" - let ?nt = "fst (zsplit0 t)" - let ?at = "snd (zsplit0 t)" - have abjs: "zsplit0 s = (?ns,?as)" by simp - moreover have abjt: "zsplit0 t = (?nt,?at)" by simp - ultimately have th: "a=Sub ?as ?at \ n=?ns - ?nt" using prems - by (simp add: Let_def split_def) - from abjs[symmetric] have bluddy: "\ x y. (x,y) = zsplit0 s" by blast - from prems have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real x # bs) (CN 0 xa xb) = Inum (real x # bs) t \ numbound0 xb)" by simp - with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast - from abjs prems have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \ ?N ?as" by blast - from th3[simplified] th2[simplified] th[simplified] show ?case - by (simp add: left_diff_distrib) -next - case (8 i t n a) - let ?nt = "fst (zsplit0 t)" - let ?at = "snd (zsplit0 t)" - have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at \ n=i*?nt" using prems - by (simp add: Let_def split_def) - from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast - hence " ?I x (Mul i t) = (real i) * ?I x (CN 0 ?nt ?at)" by simp - also have "\ = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: right_distrib) - finally show ?case using th th2 by simp -next - case (9 t n a) - let ?nt = "fst (zsplit0 t)" - let ?at = "snd (zsplit0 t)" - have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a= Floor ?at \ n=?nt" using prems - by (simp add: Let_def split_def) - from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast - hence na: "?N a" using th by simp - have th': "(real ?nt)*(real x) = real (?nt * x)" by simp - have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp - also have "\ = real (floor ((real ?nt)* real(x) + ?I x ?at))" by simp - also have "\ = real (floor (?I x ?at + real (?nt* x)))" by (simp add: add_ac) - also have "\ = real (floor (?I x ?at) + (?nt* x))" - using floor_add[where x="?I x ?at" and a="?nt* x"] by simp - also have "\ = real (?nt)*(real x) + real (floor (?I x ?at))" by (simp add: add_ac) - finally have "?I x (Floor t) = ?I x (CN 0 n a)" using th by simp - with na show ?case by simp -qed - -consts - iszlfm :: "fm \ real list \ bool" (* Linearity test for fm *) - zlfm :: "fm \ fm" (* Linearity transformation for fm *) -recdef iszlfm "measure size" - "iszlfm (And p q) = (\ bs. iszlfm p bs \ iszlfm q bs)" - "iszlfm (Or p q) = (\ bs. iszlfm p bs \ iszlfm q bs)" - "iszlfm (Eq (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" - "iszlfm (NEq (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" - "iszlfm (Lt (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" - "iszlfm (Le (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" - "iszlfm (Gt (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" - "iszlfm (Ge (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" - "iszlfm (Dvd i (CN 0 c e)) = - (\ bs. c>0 \ i>0 \ numbound0 e \ isint e bs)" - "iszlfm (NDvd i (CN 0 c e))= - (\ bs. c>0 \ i>0 \ numbound0 e \ isint e bs)" - "iszlfm p = (\ bs. isatom p \ (bound0 p))" - -lemma zlin_qfree: "iszlfm p bs \ qfree p" - by (induct p rule: iszlfm.induct) auto - -lemma iszlfm_gen: - assumes lp: "iszlfm p (x#bs)" - shows "\ y. iszlfm p (y#bs)" -proof - fix y - show "iszlfm p (y#bs)" - using lp - by(induct p rule: iszlfm.induct, simp_all add: numbound0_gen[rule_format, where x="x" and y="y"]) -qed - -lemma conj_zl[simp]: "iszlfm p bs \ iszlfm q bs \ iszlfm (conj p q) bs" - using conj_def by (cases p,auto) -lemma disj_zl[simp]: "iszlfm p bs \ iszlfm q bs \ iszlfm (disj p q) bs" - using disj_def by (cases p,auto) -lemma not_zl[simp]: "iszlfm p bs \ iszlfm (not p) bs" - by (induct p rule:iszlfm.induct ,auto) - -recdef zlfm "measure fmsize" - "zlfm (And p q) = conj (zlfm p) (zlfm q)" - "zlfm (Or p q) = disj (zlfm p) (zlfm q)" - "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)" - "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))" - "zlfm (Lt a) = (let (c,r) = zsplit0 a in - if c=0 then Lt r else - if c>0 then Or (Lt (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r))) - else Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))" - "zlfm (Le a) = (let (c,r) = zsplit0 a in - if c=0 then Le r else - if c>0 then Or (Le (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r))) - else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))" - "zlfm (Gt a) = (let (c,r) = zsplit0 a in - if c=0 then Gt r else - if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) - else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))" - "zlfm (Ge a) = (let (c,r) = zsplit0 a in - if c=0 then Ge r else - if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) - else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))" - "zlfm (Eq a) = (let (c,r) = zsplit0 a in - if c=0 then Eq r else - if c>0 then (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Eq (Add (Floor (Neg r)) r))) - else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))" - "zlfm (NEq a) = (let (c,r) = zsplit0 a in - if c=0 then NEq r else - if c>0 then (Or (NEq (CN 0 c (Neg (Floor (Neg r))))) (NEq (Add (Floor (Neg r)) r))) - else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))" - "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a) - else (let (c,r) = zsplit0 a in - if c=0 then Dvd (abs i) r else - if c>0 then And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 c (Floor r))) - else And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 (-c) (Neg (Floor r))))))" - "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a) - else (let (c,r) = zsplit0 a in - if c=0 then NDvd (abs i) r else - if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 c (Floor r))) - else Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 (-c) (Neg (Floor r))))))" - "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))" - "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))" - "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))" - "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))" - "zlfm (NOT (NOT p)) = zlfm p" - "zlfm (NOT T) = F" - "zlfm (NOT F) = T" - "zlfm (NOT (Lt a)) = zlfm (Ge a)" - "zlfm (NOT (Le a)) = zlfm (Gt a)" - "zlfm (NOT (Gt a)) = zlfm (Le a)" - "zlfm (NOT (Ge a)) = zlfm (Lt a)" - "zlfm (NOT (Eq a)) = zlfm (NEq a)" - "zlfm (NOT (NEq a)) = zlfm (Eq a)" - "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)" - "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)" - "zlfm p = p" (hints simp add: fmsize_pos) - -lemma split_int_less_real: - "(real (a::int) < b) = (a < floor b \ (a = floor b \ real (floor b) < b))" -proof( auto) - assume alb: "real a < b" and agb: "\ a < floor b" - from agb have "floor b \ a" by simp hence th: "b < real a + 1" by (simp only: floor_le_eq) - from floor_eq[OF alb th] show "a= floor b" by simp -next - assume alb: "a < floor b" - hence "real a < real (floor b)" by simp - moreover have "real (floor b) \ b" by simp ultimately show "real a < b" by arith -qed - -lemma split_int_less_real': - "(real (a::int) + b < 0) = (real a - real (floor(-b)) < 0 \ (real a - real (floor (-b)) = 0 \ real (floor (-b)) + b < 0))" -proof- - have "(real a + b <0) = (real a < -b)" by arith - with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith -qed - -lemma split_int_gt_real': - "(real (a::int) + b > 0) = (real a + real (floor b) > 0 \ (real a + real (floor b) = 0 \ real (floor b) - b < 0))" -proof- - have th: "(real a + b >0) = (real (-a) + (-b)< 0)" by arith - show ?thesis using myless[rule_format, where b="real (floor b)"] - by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) - (simp add: ring_simps diff_def[symmetric],arith) -qed - -lemma split_int_le_real: - "(real (a::int) \ b) = (a \ floor b \ (a = floor b \ real (floor b) < b))" -proof( auto) - assume alb: "real a \ b" and agb: "\ a \ floor b" - from alb have "floor (real a) \ floor b " by (simp only: floor_mono2) - hence "a \ floor b" by simp with agb show "False" by simp -next - assume alb: "a \ floor b" - hence "real a \ real (floor b)" by (simp only: floor_mono2) - also have "\\ b" by simp finally show "real a \ b" . -qed - -lemma split_int_le_real': - "(real (a::int) + b \ 0) = (real a - real (floor(-b)) \ 0 \ (real a - real (floor (-b)) = 0 \ real (floor (-b)) + b < 0))" -proof- - have "(real a + b \0) = (real a \ -b)" by arith - with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith -qed - -lemma split_int_ge_real': - "(real (a::int) + b \ 0) = (real a + real (floor b) \ 0 \ (real a + real (floor b) = 0 \ real (floor b) - b < 0))" -proof- - have th: "(real a + b \0) = (real (-a) + (-b) \ 0)" by arith - show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"]) - (simp add: ring_simps diff_def[symmetric],arith) -qed - -lemma split_int_eq_real: "(real (a::int) = b) = ( a = floor b \ b = real (floor b))" (is "?l = ?r") -by auto - -lemma split_int_eq_real': "(real (a::int) + b = 0) = ( a - floor (-b) = 0 \ real (floor (-b)) + b = 0)" (is "?l = ?r") -proof- - have "?l = (real a = -b)" by arith - with split_int_eq_real[where a="a" and b="-b"] show ?thesis by simp arith -qed - -lemma zlfm_I: - assumes qfp: "qfree p" - shows "(Ifm (real i #bs) (zlfm p) = Ifm (real i# bs) p) \ iszlfm (zlfm p) (real (i::int) #bs)" - (is "(?I (?l p) = ?I p) \ ?L (?l p)") -using qfp -proof(induct p rule: zlfm.induct) - case (5 a) - let ?c = "fst (zsplit0 a)" - let ?r = "snd (zsplit0 a)" - have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" - have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith - moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] - by (cases "?r", simp_all add: Let_def split_def,case_tac "nat", simp_all)} - moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) - also have "\ = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def diff_def) - finally have ?case using l by simp} - moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) - also from cn cnz have "\ = (?I (?l (Lt a)))" by (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac, arith) - finally have ?case using l by simp} - ultimately show ?case by blast -next - case (6 a) - let ?c = "fst (zsplit0 a)" - let ?r = "snd (zsplit0 a)" - have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" - have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith - moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] - by (cases "?r", simp_all add: Let_def split_def, case_tac "nat",simp_all)} - moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Le a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) - also have "\ = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_def) - finally have ?case using l by simp} - moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Le a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) - also from cn cnz have "\ = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac ,arith) - finally have ?case using l by simp} - ultimately show ?case by blast -next - case (7 a) - let ?c = "fst (zsplit0 a)" - let ?r = "snd (zsplit0 a)" - have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" - have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith - moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] - by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} - moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) - also have "\ = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_def) - finally have ?case using l by simp} - moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) - also from cn cnz have "\ = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac, arith) - finally have ?case using l by simp} - ultimately show ?case by blast -next - case (8 a) - let ?c = "fst (zsplit0 a)" - let ?r = "snd (zsplit0 a)" - have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" - have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith - moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] - by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} - moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Ge a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) - also have "\ = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_def) - finally have ?case using l by simp} - moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Ge a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) - also from cn cnz have "\ = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac, arith) - finally have ?case using l by simp} - ultimately show ?case by blast -next - case (9 a) - let ?c = "fst (zsplit0 a)" - let ?r = "snd (zsplit0 a)" - have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" - have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith - moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] - by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} - moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) - also have "\ = (?I (?l (Eq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult) - finally have ?case using l by simp} - moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) - also from cn cnz have "\ = (?I (?l (Eq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult,arith) - finally have ?case using l by simp} - ultimately show ?case by blast -next - case (10 a) - let ?c = "fst (zsplit0 a)" - let ?r = "snd (zsplit0 a)" - have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" - have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith - moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] - by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} - moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (NEq a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) - also have "\ = (?I (?l (NEq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult) - finally have ?case using l by simp} - moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (NEq a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) - also from cn cnz have "\ = (?I (?l (NEq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult,arith) - finally have ?case using l by simp} - ultimately show ?case by blast -next - case (11 j a) - let ?c = "fst (zsplit0 a)" - let ?r = "snd (zsplit0 a)" - have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" - have "j=0 \ (j\0 \ ?c = 0) \ (j\0 \ ?c >0 \ ?c\0) \ (j\ 0 \ ?c<0 \ ?c\0)" by arith - moreover - {assume "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) - hence ?case using prems by (simp del: zlfm.simps add: rdvd_left_0_eq)} - moreover - {assume "?c=0" and "j\0" hence ?case - using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"] - by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} - moreover - {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" - using Ia by (simp add: Let_def split_def) - also have "\ = (real (abs j) rdvd real (?c*i) + (?N ?r))" - by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp - also have "\ = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ - (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" - by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) - also have "\ = (?I (?l (Dvd j a)))" using cp cnz jnz - by (simp add: Let_def split_def int_rdvd_iff[symmetric] - del: real_of_int_mult) (auto simp add: add_ac) - finally have ?case using l jnz by simp } - moreover - {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" - using Ia by (simp add: Let_def split_def) - also have "\ = (real (abs j) rdvd real (?c*i) + (?N ?r))" - by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp - also have "\ = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ - (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" - by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) - also have "\ = (?I (?l (Dvd j a)))" using cn cnz jnz - using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric] - by (simp add: Let_def split_def int_rdvd_iff[symmetric] - del: real_of_int_mult) (auto simp add: add_ac) - finally have ?case using l jnz by blast } - ultimately show ?case by blast -next - case (12 j a) - let ?c = "fst (zsplit0 a)" - let ?r = "snd (zsplit0 a)" - have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" - have "j=0 \ (j\0 \ ?c = 0) \ (j\0 \ ?c >0 \ ?c\0) \ (j\ 0 \ ?c<0 \ ?c\0)" by arith - moreover - {assume "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) - hence ?case using prems by (simp del: zlfm.simps add: rdvd_left_0_eq)} - moreover - {assume "?c=0" and "j\0" hence ?case - using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"] - by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} - moreover - {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (NDvd j a) = (\ (real j rdvd (real (?c * i) + (?N ?r))))" - using Ia by (simp add: Let_def split_def) - also have "\ = (\ (real (abs j) rdvd real (?c*i) + (?N ?r)))" - by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp - also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ - (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" - by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) - also have "\ = (?I (?l (NDvd j a)))" using cp cnz jnz - by (simp add: Let_def split_def int_rdvd_iff[symmetric] - del: real_of_int_mult) (auto simp add: add_ac) - finally have ?case using l jnz by simp } - moreover - {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" - by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (NDvd j a) = (\ (real j rdvd (real (?c * i) + (?N ?r))))" - using Ia by (simp add: Let_def split_def) - also have "\ = (\ (real (abs j) rdvd real (?c*i) + (?N ?r)))" - by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp - also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ - (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" - by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) - also have "\ = (?I (?l (NDvd j a)))" using cn cnz jnz - using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric] - by (simp add: Let_def split_def int_rdvd_iff[symmetric] - del: real_of_int_mult) (auto simp add: add_ac) - finally have ?case using l jnz by blast } - ultimately show ?case by blast -qed auto - -text{* plusinf : Virtual substitution of @{text "+\"} - minusinf: Virtual substitution of @{text "-\"} - @{text "\"} Compute lcm @{text "d| Dvd d c*x+t \ p"} - @{text "d\"} checks if a given l divides all the ds above*} - -consts - plusinf:: "fm \ fm" - minusinf:: "fm \ fm" - \ :: "fm \ int" - d\ :: "fm \ int \ bool" - -recdef minusinf "measure size" - "minusinf (And p q) = conj (minusinf p) (minusinf q)" - "minusinf (Or p q) = disj (minusinf p) (minusinf q)" - "minusinf (Eq (CN 0 c e)) = F" - "minusinf (NEq (CN 0 c e)) = T" - "minusinf (Lt (CN 0 c e)) = T" - "minusinf (Le (CN 0 c e)) = T" - "minusinf (Gt (CN 0 c e)) = F" - "minusinf (Ge (CN 0 c e)) = F" - "minusinf p = p" - -lemma minusinf_qfree: "qfree p \ qfree (minusinf p)" - by (induct p rule: minusinf.induct, auto) - -recdef plusinf "measure size" - "plusinf (And p q) = conj (plusinf p) (plusinf q)" - "plusinf (Or p q) = disj (plusinf p) (plusinf q)" - "plusinf (Eq (CN 0 c e)) = F" - "plusinf (NEq (CN 0 c e)) = T" - "plusinf (Lt (CN 0 c e)) = F" - "plusinf (Le (CN 0 c e)) = F" - "plusinf (Gt (CN 0 c e)) = T" - "plusinf (Ge (CN 0 c e)) = T" - "plusinf p = p" - -recdef \ "measure size" - "\ (And p q) = zlcm (\ p) (\ q)" - "\ (Or p q) = zlcm (\ p) (\ q)" - "\ (Dvd i (CN 0 c e)) = i" - "\ (NDvd i (CN 0 c e)) = i" - "\ p = 1" - -recdef d\ "measure size" - "d\ (And p q) = (\ d. d\ p d \ d\ q d)" - "d\ (Or p q) = (\ d. d\ p d \ d\ q d)" - "d\ (Dvd i (CN 0 c e)) = (\ d. i dvd d)" - "d\ (NDvd i (CN 0 c e)) = (\ d. i dvd d)" - "d\ p = (\ d. True)" - -lemma delta_mono: - assumes lin: "iszlfm p bs" - and d: "d dvd d'" - and ad: "d\ p d" - shows "d\ p d'" - using lin ad d -proof(induct p rule: iszlfm.induct) - case (9 i c e) thus ?case using d - by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"]) -next - case (10 i c e) thus ?case using d - by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"]) -qed simp_all - -lemma \ : assumes lin:"iszlfm p bs" - shows "d\ p (\ p) \ \ p >0" -using lin -proof (induct p rule: iszlfm.induct) - case (1 p q) - let ?d = "\ (And p q)" - from prems zlcm_pos have dp: "?d >0" by simp - have d1: "\ p dvd \ (And p q)" using prems by simp - hence th: "d\ p ?d" - using delta_mono prems by (auto simp del: dvd_zlcm_self1) - have "\ q dvd \ (And p q)" using prems by simp - hence th': "d\ q ?d" using delta_mono prems by (auto simp del: dvd_zlcm_self2) - from th th' dp show ?case by simp -next - case (2 p q) - let ?d = "\ (And p q)" - from prems zlcm_pos have dp: "?d >0" by simp - have "\ p dvd \ (And p q)" using prems by simp hence th: "d\ p ?d" using delta_mono prems - by (auto simp del: dvd_zlcm_self1) - have "\ q dvd \ (And p q)" using prems by simp hence th': "d\ q ?d" using delta_mono prems by (auto simp del: dvd_zlcm_self2) - from th th' dp show ?case by simp -qed simp_all - - -lemma minusinf_inf: - assumes linp: "iszlfm p (a # bs)" - shows "\ (z::int). \ x < z. Ifm ((real x)#bs) (minusinf p) = Ifm ((real x)#bs) p" - (is "?P p" is "\ (z::int). \ x < z. ?I x (?M p) = ?I x p") -using linp -proof (induct p rule: minusinf.induct) - case (1 f g) - from prems have "?P f" by simp - then obtain z1 where z1_def: "\ x < z1. ?I x (?M f) = ?I x f" by blast - from prems have "?P g" by simp - then obtain z2 where z2_def: "\ x < z2. ?I x (?M g) = ?I x g" by blast - let ?z = "min z1 z2" - from z1_def z2_def have "\ x < ?z. ?I x (?M (And f g)) = ?I x (And f g)" by simp - thus ?case by blast -next - case (2 f g) from prems have "?P f" by simp - then obtain z1 where z1_def: "\ x < z1. ?I x (?M f) = ?I x f" by blast - from prems have "?P g" by simp - then obtain z2 where z2_def: "\ x < z2. ?I x (?M g) = ?I x g" by blast - let ?z = "min z1 z2" - from z1_def z2_def have "\ x < ?z. ?I x (?M (Or f g)) = ?I x (Or f g)" by simp - thus ?case by blast -next - case (3 c e) - from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp - from prems have nbe: "numbound0 e" by simp - fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Eq (CN 0 c e))) = ?I x (Eq (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) - fix x - assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" - by (simp only: real_mult_less_mono2[OF rcpos th1]) - hence "real c * real x + Inum (y # bs) e \ 0"using rcpos by simp - thus "real c * real x + Inum (real x # bs) e \ 0" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] by simp - qed - thus ?case by blast -next - case (4 c e) - from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp - from prems have nbe: "numbound0 e" by simp - fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (NEq (CN 0 c e))) = ?I x (NEq (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) - fix x - assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" - by (simp only: real_mult_less_mono2[OF rcpos th1]) - hence "real c * real x + Inum (y # bs) e \ 0"using rcpos by simp - thus "real c * real x + Inum (real x # bs) e \ 0" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] by simp - qed - thus ?case by blast -next - case (5 c e) - from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp - from prems have nbe: "numbound0 e" by simp - fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Lt (CN 0 c e))) = ?I x (Lt (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) - fix x - assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" - by (simp only: real_mult_less_mono2[OF rcpos th1]) - thus "real c * real x + Inum (real x # bs) e < 0" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp - qed - thus ?case by blast -next - case (6 c e) - from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp - from prems have nbe: "numbound0 e" by simp - fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Le (CN 0 c e))) = ?I x (Le (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) - fix x - assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" - by (simp only: real_mult_less_mono2[OF rcpos th1]) - thus "real c * real x + Inum (real x # bs) e \ 0" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp - qed - thus ?case by blast -next - case (7 c e) - from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp - from prems have nbe: "numbound0 e" by simp - fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Gt (CN 0 c e))) = ?I x (Gt (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) - fix x - assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" - by (simp only: real_mult_less_mono2[OF rcpos th1]) - thus "\ (real c * real x + Inum (real x # bs) e>0)" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp - qed - thus ?case by blast -next - case (8 c e) - from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp - from prems have nbe: "numbound0 e" by simp - fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Ge (CN 0 c e))) = ?I x (Ge (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) - fix x - assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" - by (simp only: real_mult_less_mono2[OF rcpos th1]) - thus "\ real c * real x + Inum (real x # bs) e \ 0" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp - qed - thus ?case by blast -qed simp_all - -lemma minusinf_repeats: - assumes d: "d\ p d" and linp: "iszlfm p (a # bs)" - shows "Ifm ((real(x - k*d))#bs) (minusinf p) = Ifm (real x #bs) (minusinf p)" -using linp d -proof(induct p rule: iszlfm.induct) - case (9 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ - hence "\ k. d=i*k" by (simp add: dvd_def) - then obtain "di" where di_def: "d=i*di" by blast - show ?case - proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI) - assume - "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" - (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") - hence "\ (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" - by (simp add: ring_simps di_def) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))" - by (simp add: ring_simps) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast - thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp - next - assume - "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e") - hence "\ (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: ring_simps) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)" - by blast - thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp - qed -next - case (10 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ - hence "\ k. d=i*k" by (simp add: dvd_def) - then obtain "di" where di_def: "d=i*di" by blast - show ?case - proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI) - assume - "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" - (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") - hence "\ (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" - by (simp add: ring_simps di_def) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))" - by (simp add: ring_simps) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast - thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp - next - assume - "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e") - hence "\ (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: ring_simps) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)" - by blast - thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp - qed -qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="real(x - k*d)" and b'="real x"] simp del: real_of_int_mult real_of_int_diff) - -lemma minusinf_ex: - assumes lin: "iszlfm p (real (a::int) #bs)" - and exmi: "\ (x::int). Ifm (real x#bs) (minusinf p)" (is "\ x. ?P1 x") - shows "\ (x::int). Ifm (real x#bs) p" (is "\ x. ?P x") -proof- - let ?d = "\ p" - from \ [OF lin] have dpos: "?d >0" by simp - from \ [OF lin] have alld: "d\ p ?d" by simp - from minusinf_repeats[OF alld lin] have th1:"\ x k. ?P1 x = ?P1 (x - (k * ?d))" by simp - from minusinf_inf[OF lin] have th2:"\ z. \ x. x (?P x = ?P1 x)" by blast - from minusinfinity [OF dpos th1 th2] exmi show ?thesis by blast -qed - -lemma minusinf_bex: - assumes lin: "iszlfm p (real (a::int) #bs)" - shows "(\ (x::int). Ifm (real x#bs) (minusinf p)) = - (\ (x::int)\ {1..\ p}. Ifm (real x#bs) (minusinf p))" - (is "(\ x. ?P x) = _") -proof- - let ?d = "\ p" - from \ [OF lin] have dpos: "?d >0" by simp - from \ [OF lin] have alld: "d\ p ?d" by simp - from minusinf_repeats[OF alld lin] have th1:"\ x k. ?P x = ?P (x - (k * ?d))" by simp - from periodic_finite_ex[OF dpos th1] show ?thesis by blast -qed - -lemma dvd1_eq1: "x >0 \ (x::int) dvd 1 = (x = 1)" by auto - -consts - a\ :: "fm \ int \ fm" (* adjusts the coeffitients of a formula *) - d\ :: "fm \ int \ bool" (* tests if all coeffs c of c divide a given l*) - \ :: "fm \ int" (* computes the lcm of all coefficients of x*) - \ :: "fm \ num list" - \ :: "fm \ num list" - -recdef a\ "measure size" - "a\ (And p q) = (\ k. And (a\ p k) (a\ q k))" - "a\ (Or p q) = (\ k. Or (a\ p k) (a\ q k))" - "a\ (Eq (CN 0 c e)) = (\ k. Eq (CN 0 1 (Mul (k div c) e)))" - "a\ (NEq (CN 0 c e)) = (\ k. NEq (CN 0 1 (Mul (k div c) e)))" - "a\ (Lt (CN 0 c e)) = (\ k. Lt (CN 0 1 (Mul (k div c) e)))" - "a\ (Le (CN 0 c e)) = (\ k. Le (CN 0 1 (Mul (k div c) e)))" - "a\ (Gt (CN 0 c e)) = (\ k. Gt (CN 0 1 (Mul (k div c) e)))" - "a\ (Ge (CN 0 c e)) = (\ k. Ge (CN 0 1 (Mul (k div c) e)))" - "a\ (Dvd i (CN 0 c e)) =(\ k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" - "a\ (NDvd i (CN 0 c e))=(\ k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" - "a\ p = (\ k. p)" - -recdef d\ "measure size" - "d\ (And p q) = (\ k. (d\ p k) \ (d\ q k))" - "d\ (Or p q) = (\ k. (d\ p k) \ (d\ q k))" - "d\ (Eq (CN 0 c e)) = (\ k. c dvd k)" - "d\ (NEq (CN 0 c e)) = (\ k. c dvd k)" - "d\ (Lt (CN 0 c e)) = (\ k. c dvd k)" - "d\ (Le (CN 0 c e)) = (\ k. c dvd k)" - "d\ (Gt (CN 0 c e)) = (\ k. c dvd k)" - "d\ (Ge (CN 0 c e)) = (\ k. c dvd k)" - "d\ (Dvd i (CN 0 c e)) =(\ k. c dvd k)" - "d\ (NDvd i (CN 0 c e))=(\ k. c dvd k)" - "d\ p = (\ k. True)" - -recdef \ "measure size" - "\ (And p q) = zlcm (\ p) (\ q)" - "\ (Or p q) = zlcm (\ p) (\ q)" - "\ (Eq (CN 0 c e)) = c" - "\ (NEq (CN 0 c e)) = c" - "\ (Lt (CN 0 c e)) = c" - "\ (Le (CN 0 c e)) = c" - "\ (Gt (CN 0 c e)) = c" - "\ (Ge (CN 0 c e)) = c" - "\ (Dvd i (CN 0 c e)) = c" - "\ (NDvd i (CN 0 c e))= c" - "\ p = 1" - -recdef \ "measure size" - "\ (And p q) = (\ p @ \ q)" - "\ (Or p q) = (\ p @ \ q)" - "\ (Eq (CN 0 c e)) = [Sub (C -1) e]" - "\ (NEq (CN 0 c e)) = [Neg e]" - "\ (Lt (CN 0 c e)) = []" - "\ (Le (CN 0 c e)) = []" - "\ (Gt (CN 0 c e)) = [Neg e]" - "\ (Ge (CN 0 c e)) = [Sub (C -1) e]" - "\ p = []" - -recdef \ "measure size" - "\ (And p q) = (\ p @ \ q)" - "\ (Or p q) = (\ p @ \ q)" - "\ (Eq (CN 0 c e)) = [Add (C -1) e]" - "\ (NEq (CN 0 c e)) = [e]" - "\ (Lt (CN 0 c e)) = [e]" - "\ (Le (CN 0 c e)) = [Add (C -1) e]" - "\ (Gt (CN 0 c e)) = []" - "\ (Ge (CN 0 c e)) = []" - "\ p = []" -consts mirror :: "fm \ fm" -recdef mirror "measure size" - "mirror (And p q) = And (mirror p) (mirror q)" - "mirror (Or p q) = Or (mirror p) (mirror q)" - "mirror (Eq (CN 0 c e)) = Eq (CN 0 c (Neg e))" - "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))" - "mirror (Lt (CN 0 c e)) = Gt (CN 0 c (Neg e))" - "mirror (Le (CN 0 c e)) = Ge (CN 0 c (Neg e))" - "mirror (Gt (CN 0 c e)) = Lt (CN 0 c (Neg e))" - "mirror (Ge (CN 0 c e)) = Le (CN 0 c (Neg e))" - "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))" - "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))" - "mirror p = p" - -lemma mirror\\: - assumes lp: "iszlfm p (a#bs)" - shows "(Inum (real (i::int)#bs)) ` set (\ p) = (Inum (real i#bs)) ` set (\ (mirror p))" -using lp -by (induct p rule: mirror.induct, auto) - -lemma mirror: - assumes lp: "iszlfm p (a#bs)" - shows "Ifm (real (x::int)#bs) (mirror p) = Ifm (real (- x)#bs) p" -using lp -proof(induct p rule: iszlfm.induct) - case (9 j c e) - have th: "(real j rdvd real c * real x - Inum (real x # bs) e) = - (real j rdvd - (real c * real x - Inum (real x # bs) e))" - by (simp only: rdvd_minus[symmetric]) - from prems show ?case - by (simp add: ring_simps th[simplified ring_simps] - numbound0_I[where bs="bs" and b'="real x" and b="- real x"]) -next - case (10 j c e) - have th: "(real j rdvd real c * real x - Inum (real x # bs) e) = - (real j rdvd - (real c * real x - Inum (real x # bs) e))" - by (simp only: rdvd_minus[symmetric]) - from prems show ?case - by (simp add: ring_simps th[simplified ring_simps] - numbound0_I[where bs="bs" and b'="real x" and b="- real x"]) -qed (auto simp add: numbound0_I[where bs="bs" and b="real x" and b'="- real x"] nth_pos2) - -lemma mirror_l: "iszlfm p (a#bs) \ iszlfm (mirror p) (a#bs)" -by (induct p rule: mirror.induct, auto simp add: isint_neg) - -lemma mirror_d\: "iszlfm p (a#bs) \ d\ p 1 - \ iszlfm (mirror p) (a#bs) \ d\ (mirror p) 1" -by (induct p rule: mirror.induct, auto simp add: isint_neg) - -lemma mirror_\: "iszlfm p (a#bs) \ \ (mirror p) = \ p" -by (induct p rule: mirror.induct,auto) - - -lemma mirror_ex: - assumes lp: "iszlfm p (real (i::int)#bs)" - shows "(\ (x::int). Ifm (real x#bs) (mirror p)) = (\ (x::int). Ifm (real x#bs) p)" - (is "(\ x. ?I x ?mp) = (\ x. ?I x p)") -proof(auto) - fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast - thus "\ x. ?I x p" by blast -next - fix x assume "?I x p" hence "?I (- x) ?mp" - using mirror[OF lp, where x="- x", symmetric] by auto - thus "\ x. ?I x ?mp" by blast -qed - -lemma \_numbound0: assumes lp: "iszlfm p bs" - shows "\ b\ set (\ p). numbound0 b" - using lp by (induct p rule: \.induct,auto) - -lemma d\_mono: - assumes linp: "iszlfm p (a #bs)" - and dr: "d\ p l" - and d: "l dvd l'" - shows "d\ p l'" -using dr linp zdvd_trans[where n="l" and k="l'", simplified d] -by (induct p rule: iszlfm.induct) simp_all - -lemma \_l: assumes lp: "iszlfm p (a#bs)" - shows "\ b\ set (\ p). numbound0 b \ isint b (a#bs)" -using lp -by(induct p rule: \.induct, auto simp add: isint_add isint_c) - -lemma \: - assumes linp: "iszlfm p (a #bs)" - shows "\ p > 0 \ d\ p (\ p)" -using linp -proof(induct p rule: iszlfm.induct) - case (1 p q) - from prems have dl1: "\ p dvd zlcm (\ p) (\ q)" by simp - from prems have dl2: "\ q dvd zlcm (\ p) (\ q)" by simp - from prems d\_mono[where p = "p" and l="\ p" and l'="zlcm (\ p) (\ q)"] - d\_mono[where p = "q" and l="\ q" and l'="zlcm (\ p) (\ q)"] - dl1 dl2 show ?case by (auto simp add: zlcm_pos) -next - case (2 p q) - from prems have dl1: "\ p dvd zlcm (\ p) (\ q)" by simp - from prems have dl2: "\ q dvd zlcm (\ p) (\ q)" by simp - from prems d\_mono[where p = "p" and l="\ p" and l'="zlcm (\ p) (\ q)"] - d\_mono[where p = "q" and l="\ q" and l'="zlcm (\ p) (\ q)"] - dl1 dl2 show ?case by (auto simp add: zlcm_pos) -qed (auto simp add: zlcm_pos) - -lemma a\: assumes linp: "iszlfm p (a #bs)" and d: "d\ p l" and lp: "l > 0" - shows "iszlfm (a\ p l) (a #bs) \ d\ (a\ p l) 1 \ (Ifm (real (l * x) #bs) (a\ p l) = Ifm ((real x)#bs) p)" -using linp d -proof (induct p rule: iszlfm.induct) - case (5 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ - from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) - from cp have cnz: "c \ 0" by simp - have "c div c\ l div c" - by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" - by (simp add: zdiv_self[OF cnz]) - have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] - by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0\real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e < 0)" - by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) < (real (l div c)) * 0)" by (simp add: ring_simps) - also have "\ = (real c * real x + Inum (real x # bs) e < 0)" - using mult_less_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp -next - case (6 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ - from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) - from cp have cnz: "c \ 0" by simp - have "c div c\ l div c" - by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" - by (simp add: zdiv_self[OF cnz]) - have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] - by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e \ (0\real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \ 0)" - by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (real (l div c)) * 0)" by (simp add: ring_simps) - also have "\ = (real c * real x + Inum (real x # bs) e \ 0)" - using mult_le_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp -next - case (7 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ - from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) - from cp have cnz: "c \ 0" by simp - have "c div c\ l div c" - by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" - by (simp add: zdiv_self[OF cnz]) - have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] - by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0\real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e > 0)" - by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) > (real (l div c)) * 0)" by (simp add: ring_simps) - also have "\ = (real c * real x + Inum (real x # bs) e > 0)" - using zero_less_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp -next - case (8 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ - from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) - from cp have cnz: "c \ 0" by simp - have "c div c\ l div c" - by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" - by (simp add: zdiv_self[OF cnz]) - have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] - by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e \ (0\real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \ 0)" - by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (real (l div c)) * 0)" by (simp add: ring_simps) - also have "\ = (real c * real x + Inum (real x # bs) e \ 0)" - using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp -next - case (3 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ - from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) - from cp have cnz: "c \ 0" by simp - have "c div c\ l div c" - by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" - by (simp add: zdiv_self[OF cnz]) - have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] - by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0\real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = 0)" - by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) = (real (l div c)) * 0)" by (simp add: ring_simps) - also have "\ = (real c * real x + Inum (real x # bs) e = 0)" - using mult_eq_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp -next - case (4 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ - from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) - from cp have cnz: "c \ 0" by simp - have "c div c\ l div c" - by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" - by (simp add: zdiv_self[OF cnz]) - have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] - by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e \ (0\real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \ 0)" - by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (real (l div c)) * 0)" by (simp add: ring_simps) - also have "\ = (real c * real x + Inum (real x # bs) e \ 0)" - using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp -next - case (9 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+ - from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) - from cp have cnz: "c \ 0" by simp - have "c div c\ l div c" - by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" - by (simp add: zdiv_self[OF cnz]) - have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] - by simp - hence "(\ (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\ (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)" by simp - also have "\ = (\ (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: ring_simps) - also fix k have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)" - using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp - also have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp -next - case (10 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+ - from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) - from cp have cnz: "c \ 0" by simp - have "c div c\ l div c" - by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" - by (simp add: zdiv_self[OF cnz]) - have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] - by simp - hence "(\ (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\ (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)" by simp - also have "\ = (\ (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: ring_simps) - also fix k have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)" - using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp - also have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp -qed (simp_all add: nth_pos2 numbound0_I[where bs="bs" and b="real (l * x)" and b'="real x"] isint_Mul del: real_of_int_mult) - -lemma a\_ex: assumes linp: "iszlfm p (a#bs)" and d: "d\ p l" and lp: "l>0" - shows "(\ x. l dvd x \ Ifm (real x #bs) (a\ p l)) = (\ (x::int). Ifm (real x#bs) p)" - (is "(\ x. l dvd x \ ?P x) = (\ x. ?P' x)") -proof- - have "(\ x. l dvd x \ ?P x) = (\ (x::int). ?P (l*x))" - using unity_coeff_ex[where l="l" and P="?P", simplified] by simp - also have "\ = (\ (x::int). ?P' x)" using a\[OF linp d lp] by simp - finally show ?thesis . -qed - -lemma \: - assumes lp: "iszlfm p (a#bs)" - and u: "d\ p 1" - and d: "d\ p d" - and dp: "d > 0" - and nob: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real x = b + real j)" - and p: "Ifm (real x#bs) p" (is "?P x") - shows "?P (x - d)" -using lp u d dp nob p -proof(induct p rule: iszlfm.induct) - case (5 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ - with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] prems - show ?case by (simp del: real_of_int_minus) -next - case (6 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ - with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] prems - show ?case by (simp del: real_of_int_minus) -next - case (7 c e) hence p: "Ifm (real x #bs) (Gt (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" and ie1:"isint e (a#bs)" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" - from ie1 have ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="a#bs"] - numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"] - by (simp add: isint_iff) - {assume "real (x-d) +?e > 0" hence ?case using c1 - numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] - by (simp del: real_of_int_minus)} - moreover - {assume H: "\ real (x-d) + ?e > 0" - let ?v="Neg e" - have vb: "?v \ set (\ (Gt (CN 0 c e)))" by simp - from prems(11)[simplified simp_thms Inum.simps \.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] - have nob: "\ (\ j\ {1 ..d}. real x = - ?e + real j)" by auto - from H p have "real x + ?e > 0 \ real x + ?e \ real d" by (simp add: c1) - hence "real (x + floor ?e) > real (0::int) \ real (x + floor ?e) \ real d" - using ie by simp - hence "x + floor ?e \ 1 \ x + floor ?e \ d" by simp - hence "\ (j::int) \ {1 .. d}. j = x + floor ?e" by simp - hence "\ (j::int) \ {1 .. d}. real x = real (- floor ?e + j)" - by (simp only: real_of_int_inject) (simp add: ring_simps) - hence "\ (j::int) \ {1 .. d}. real x = - ?e + real j" - by (simp add: ie[simplified isint_iff]) - with nob have ?case by auto} - ultimately show ?case by blast -next - case (8 c e) hence p: "Ifm (real x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" - and ie1:"isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" - from ie1 have ie: "real (floor ?e) = ?e" using numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real x)#bs"] - by (simp add: isint_iff) - {assume "real (x-d) +?e \ 0" hence ?case using c1 - numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] - by (simp del: real_of_int_minus)} - moreover - {assume H: "\ real (x-d) + ?e \ 0" - let ?v="Sub (C -1) e" - have vb: "?v \ set (\ (Ge (CN 0 c e)))" by simp - from prems(11)[simplified simp_thms Inum.simps \.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] - have nob: "\ (\ j\ {1 ..d}. real x = - ?e - 1 + real j)" by auto - from H p have "real x + ?e \ 0 \ real x + ?e < real d" by (simp add: c1) - hence "real (x + floor ?e) \ real (0::int) \ real (x + floor ?e) < real d" - using ie by simp - hence "x + floor ?e +1 \ 1 \ x + floor ?e + 1 \ d" by simp - hence "\ (j::int) \ {1 .. d}. j = x + floor ?e + 1" by simp - hence "\ (j::int) \ {1 .. d}. x= - floor ?e - 1 + j" by (simp add: ring_simps) - hence "\ (j::int) \ {1 .. d}. real x= real (- floor ?e - 1 + j)" - by (simp only: real_of_int_inject) - hence "\ (j::int) \ {1 .. d}. real x= - ?e - 1 + real j" - by (simp add: ie[simplified isint_iff]) - with nob have ?case by simp } - ultimately show ?case by blast -next - case (3 c e) hence p: "Ifm (real x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" - and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" - let ?v="(Sub (C -1) e)" - have vb: "?v \ set (\ (Eq (CN 0 c e)))" by simp - from p have "real x= - ?e" by (simp add: c1) with prems(11) show ?case using dp - by simp (erule ballE[where x="1"], - simp_all add:ring_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"]) -next - case (4 c e)hence p: "Ifm (real x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" - and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" - let ?v="Neg e" - have vb: "?v \ set (\ (NEq (CN 0 c e)))" by simp - {assume "real x - real d + Inum ((real (x -d)) # bs) e \ 0" - hence ?case by (simp add: c1)} - moreover - {assume H: "real x - real d + Inum ((real (x -d)) # bs) e = 0" - hence "real x = - Inum ((real (x -d)) # bs) e + real d" by simp - hence "real x = - Inum (a # bs) e + real d" - by (simp add: numbound0_I[OF bn,where b="real x - real d"and b'="a"and bs="bs"]) - with prems(11) have ?case using dp by simp} - ultimately show ?case by blast -next - case (9 j c e) hence p: "Ifm (real x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" - and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" - from prems have "isint e (a #bs)" by simp - hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real x)#bs"] numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] - by (simp add: isint_iff) - from prems have id: "j dvd d" by simp - from c1 ie[symmetric] have "?p x = (real j rdvd real (x+ floor ?e))" by simp - also have "\ = (j dvd x + floor ?e)" - using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp - also have "\ = (j dvd x - d + floor ?e)" - using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp - also have "\ = (real j rdvd real (x - d + floor ?e))" - using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified] - ie by simp - also have "\ = (real j rdvd real x - real d + ?e)" - using ie by simp - finally show ?case - using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp -next - case (10 j c e) hence p: "Ifm (real x #bs) (NDvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" - from prems have "isint e (a#bs)" by simp - hence ie: "real (floor ?e) = ?e" using numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real x)#bs"] - by (simp add: isint_iff) - from prems have id: "j dvd d" by simp - from c1 ie[symmetric] have "?p x = (\ real j rdvd real (x+ floor ?e))" by simp - also have "\ = (\ j dvd x + floor ?e)" - using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp - also have "\ = (\ j dvd x - d + floor ?e)" - using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp - also have "\ = (\ real j rdvd real (x - d + floor ?e))" - using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified] - ie by simp - also have "\ = (\ real j rdvd real x - real d + ?e)" - using ie by simp - finally show ?case using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp -qed (auto simp add: numbound0_I[where bs="bs" and b="real (x - d)" and b'="real x"] nth_pos2 simp del: real_of_int_diff) - -lemma \': - assumes lp: "iszlfm p (a #bs)" - and u: "d\ p 1" - and d: "d\ p d" - and dp: "d > 0" - shows "\ x. \(\(j::int) \ {1 .. d}. \ b\ set(\ p). Ifm ((Inum (a#bs) b + real j) #bs) p) \ Ifm (real x#bs) p \ Ifm (real (x - d)#bs) p" (is "\ x. ?b \ ?P x \ ?P (x - d)") -proof(clarify) - fix x - assume nb:"?b" and px: "?P x" - hence nb2: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real x = b + real j)" - by auto - from \[OF lp u d dp nb2 px] show "?P (x -d )" . -qed - -lemma \_int: assumes lp: "iszlfm p bs" - shows "\ b\ set (\ p). isint b bs" -using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub) - -lemma cpmi_eq: "0 < D \ (EX z::int. ALL x. x < z --> (P x = P1 x)) -==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) -==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D)))) -==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))" -apply(rule iffI) -prefer 2 -apply(drule minusinfinity) -apply assumption+ -apply(fastsimp) -apply clarsimp -apply(subgoal_tac "!!k. 0<=k \ !x. P x \ P (x - k*D)") -apply(frule_tac x = x and z=z in decr_lemma) -apply(subgoal_tac "P1(x - (\x - z\ + 1) * D)") -prefer 2 -apply(subgoal_tac "0 <= (\x - z\ + 1)") -prefer 2 apply arith - apply fastsimp -apply(drule (1) periodic_finite_ex) -apply blast -apply(blast dest:decr_mult_lemma) -done - - -theorem cp_thm: - assumes lp: "iszlfm p (a #bs)" - and u: "d\ p 1" - and d: "d\ p d" - and dp: "d > 0" - shows "(\ (x::int). Ifm (real x #bs) p) = (\ j\ {1.. d}. Ifm (real j #bs) (minusinf p) \ (\ b \ set (\ p). Ifm ((Inum (a#bs) b + real j) #bs) p))" - (is "(\ (x::int). ?P (real x)) = (\ j\ ?D. ?M j \ (\ b\ ?B. ?P (?I b + real j)))") -proof- - from minusinf_inf[OF lp] - have th: "\(z::int). \x_int[OF lp] isint_iff[where bs="a # bs"] have B: "\ b\ ?B. real (floor (?I b)) = ?I b" by simp - from B[rule_format] - have "(\j\?D. \b\ ?B. ?P (?I b + real j)) = (\j\?D. \b\ ?B. ?P (real (floor (?I b)) + real j))" - by simp - also have "\ = (\j\?D. \b\ ?B. ?P (real (floor (?I b) + j)))" by simp - also have"\ = (\ j \ ?D. \ b \ ?B'. ?P (real (b + j)))" by blast - finally have BB': - "(\j\?D. \b\ ?B. ?P (?I b + real j)) = (\ j \ ?D. \ b \ ?B'. ?P (real (b + j)))" - by blast - hence th2: "\ x. \ (\ j \ ?D. \ b \ ?B'. ?P (real (b + j))) \ ?P (real x) \ ?P (real (x - d))" using \'[OF lp u d dp] by blast - from minusinf_repeats[OF d lp] - have th3: "\ x k. ?M x = ?M (x-k*d)" by simp - from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast -qed - - (* Reddy and Loveland *) - - -consts - \ :: "fm \ (num \ int) list" (* Compute the Reddy and Loveland Bset*) - \\:: "fm \ num \ int \ fm" (* Performs the modified substitution of Reddy and Loveland*) - \\ :: "fm \ (num\int) list" - a\ :: "fm \ int \ fm" -recdef \ "measure size" - "\ (And p q) = (\ p @ \ q)" - "\ (Or p q) = (\ p @ \ q)" - "\ (Eq (CN 0 c e)) = [(Sub (C -1) e,c)]" - "\ (NEq (CN 0 c e)) = [(Neg e,c)]" - "\ (Lt (CN 0 c e)) = []" - "\ (Le (CN 0 c e)) = []" - "\ (Gt (CN 0 c e)) = [(Neg e, c)]" - "\ (Ge (CN 0 c e)) = [(Sub (C (-1)) e, c)]" - "\ p = []" - -recdef \\ "measure size" - "\\ (And p q) = (\ (t,k). And (\\ p (t,k)) (\\ q (t,k)))" - "\\ (Or p q) = (\ (t,k). Or (\\ p (t,k)) (\\ q (t,k)))" - "\\ (Eq (CN 0 c e)) = (\ (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e)) - else (Eq (Add (Mul c t) (Mul k e))))" - "\\ (NEq (CN 0 c e)) = (\ (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e)) - else (NEq (Add (Mul c t) (Mul k e))))" - "\\ (Lt (CN 0 c e)) = (\ (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e)) - else (Lt (Add (Mul c t) (Mul k e))))" - "\\ (Le (CN 0 c e)) = (\ (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e)) - else (Le (Add (Mul c t) (Mul k e))))" - "\\ (Gt (CN 0 c e)) = (\ (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e)) - else (Gt (Add (Mul c t) (Mul k e))))" - "\\ (Ge (CN 0 c e)) = (\ (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e)) - else (Ge (Add (Mul c t) (Mul k e))))" - "\\ (Dvd i (CN 0 c e)) =(\ (t,k). if k dvd c then (Dvd i (Add (Mul (c div k) t) e)) - else (Dvd (i*k) (Add (Mul c t) (Mul k e))))" - "\\ (NDvd i (CN 0 c e))=(\ (t,k). if k dvd c then (NDvd i (Add (Mul (c div k) t) e)) - else (NDvd (i*k) (Add (Mul c t) (Mul k e))))" - "\\ p = (\ (t,k). p)" - -recdef \\ "measure size" - "\\ (And p q) = (\\ p @ \\ q)" - "\\ (Or p q) = (\\ p @ \\ q)" - "\\ (Eq (CN 0 c e)) = [(Add (C -1) e,c)]" - "\\ (NEq (CN 0 c e)) = [(e,c)]" - "\\ (Lt (CN 0 c e)) = [(e,c)]" - "\\ (Le (CN 0 c e)) = [(Add (C -1) e,c)]" - "\\ p = []" - - (* Simulates normal substituion by modifying the formula see correctness theorem *) - -recdef a\ "measure size" - "a\ (And p q) = (\ k. And (a\ p k) (a\ q k))" - "a\ (Or p q) = (\ k. Or (a\ p k) (a\ q k))" - "a\ (Eq (CN 0 c e)) = (\ k. if k dvd c then (Eq (CN 0 (c div k) e)) - else (Eq (CN 0 c (Mul k e))))" - "a\ (NEq (CN 0 c e)) = (\ k. if k dvd c then (NEq (CN 0 (c div k) e)) - else (NEq (CN 0 c (Mul k e))))" - "a\ (Lt (CN 0 c e)) = (\ k. if k dvd c then (Lt (CN 0 (c div k) e)) - else (Lt (CN 0 c (Mul k e))))" - "a\ (Le (CN 0 c e)) = (\ k. if k dvd c then (Le (CN 0 (c div k) e)) - else (Le (CN 0 c (Mul k e))))" - "a\ (Gt (CN 0 c e)) = (\ k. if k dvd c then (Gt (CN 0 (c div k) e)) - else (Gt (CN 0 c (Mul k e))))" - "a\ (Ge (CN 0 c e)) = (\ k. if k dvd c then (Ge (CN 0 (c div k) e)) - else (Ge (CN 0 c (Mul k e))))" - "a\ (Dvd i (CN 0 c e)) = (\ k. if k dvd c then (Dvd i (CN 0 (c div k) e)) - else (Dvd (i*k) (CN 0 c (Mul k e))))" - "a\ (NDvd i (CN 0 c e)) = (\ k. if k dvd c then (NDvd i (CN 0 (c div k) e)) - else (NDvd (i*k) (CN 0 c (Mul k e))))" - "a\ p = (\ k. p)" - -constdefs \ :: "fm \ int \ num \ fm" - "\ p k t \ And (Dvd k t) (\\ p (t,k))" - -lemma \\: - assumes linp: "iszlfm p (real (x::int)#bs)" - and kpos: "real k > 0" - and tnb: "numbound0 t" - and tint: "isint t (real x#bs)" - and kdt: "k dvd floor (Inum (b'#bs) t)" - shows "Ifm (real x#bs) (\\ p (t,k)) = - (Ifm ((real ((floor (Inum (b'#bs) t)) div k))#bs) p)" - (is "?I (real x) (?s p) = (?I (real ((floor (?N b' t)) div k)) p)" is "_ = (?I ?tk p)") -using linp kpos tnb -proof(induct p rule: \\.induct) - case (3 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - {assume kdc: "k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } - moreover - {assume "\ k dvd c" - from kpos have knz: "k\0" by simp hence knz': "real k \ 0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have "?I (real x) (?s (Eq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k = 0)" - using real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) - also have "\ = (?I ?tk (Eq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] - by (simp add: ti) - finally have ?case . } - ultimately show ?case by blast -next - case (4 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - {assume kdc: "k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } - moreover - {assume "\ k dvd c" - from kpos have knz: "k\0" by simp hence knz': "real k \ 0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have "?I (real x) (?s (NEq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \ 0)" - using real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) - also have "\ = (?I ?tk (NEq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] - by (simp add: ti) - finally have ?case . } - ultimately show ?case by blast -next - case (5 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - {assume kdc: "k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } - moreover - {assume "\ k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have "?I (real x) (?s (Lt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k < 0)" - using real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) - also have "\ = (?I ?tk (Lt (CN 0 c e)))" using pos_less_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] - by (simp add: ti) - finally have ?case . } - ultimately show ?case by blast -next - case (6 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - {assume kdc: "k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } - moreover - {assume "\ k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have "?I (real x) (?s (Le (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \ 0)" - using real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) - also have "\ = (?I ?tk (Le (CN 0 c e)))" using pos_le_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] - by (simp add: ti) - finally have ?case . } - ultimately show ?case by blast -next - case (7 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - {assume kdc: "k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } - moreover - {assume "\ k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have "?I (real x) (?s (Gt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k > 0)" - using real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) - also have "\ = (?I ?tk (Gt (CN 0 c e)))" using pos_divide_less_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] - by (simp add: ti) - finally have ?case . } - ultimately show ?case by blast -next - case (8 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - {assume kdc: "k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } - moreover - {assume "\ k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have "?I (real x) (?s (Ge (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \ 0)" - using real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) - also have "\ = (?I ?tk (Ge (CN 0 c e)))" using pos_divide_le_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] - by (simp add: ti) - finally have ?case . } - ultimately show ?case by blast -next - case (9 i c e) from prems have cp: "c > 0" and nb: "numbound0 e" by auto - {assume kdc: "k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } - moreover - {assume "\ k dvd c" - from kpos have knz: "k\0" by simp hence knz': "real k \ 0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have "?I (real x) (?s (Dvd i (CN 0 c e))) = (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k)" - using real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) - also have "\ = (?I ?tk (Dvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] - by (simp add: ti) - finally have ?case . } - ultimately show ?case by blast -next - case (10 i c e) from prems have cp: "c > 0" and nb: "numbound0 e" by auto - {assume kdc: "k dvd c" - from kpos have knz: "k\0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } - moreover - {assume "\ k dvd c" - from kpos have knz: "k\0" by simp hence knz': "real k \ 0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from prems have "?I (real x) (?s (NDvd i (CN 0 c e))) = (\ (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k))" - using real_of_int_div[OF knz kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) - also have "\ = (?I ?tk (NDvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] - by (simp add: ti) - finally have ?case . } - ultimately show ?case by blast -qed (simp_all add: nth_pos2 bound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"] numbound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"]) - - -lemma a\: - assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "real k > 0" - shows "Ifm (real (x*k)#bs) (a\ p k) = Ifm (real x#bs) p" (is "?I (x*k) (?f p k) = ?I x p") -using lp bound0_I[where bs="bs" and b="real (x*k)" and b'="real x"] numbound0_I[where bs="bs" and b="real (x*k)" and b'="real x"] -proof(induct p rule: a\.induct) - case (3 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp - {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } - moreover - {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} - ultimately show ?case by blast -next - case (4 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp - {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } - moreover - {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} - ultimately show ?case by blast -next - case (5 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp - {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } - moreover - {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_less_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} - ultimately show ?case by blast -next - case (6 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp - {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } - moreover - {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_le_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} - ultimately show ?case by blast -next - case (7 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp - {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } - moreover - {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_less_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} - ultimately show ?case by blast -next - case (8 c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp - {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } - moreover - {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_le_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} - ultimately show ?case by blast -next - case (9 i c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp - {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } - moreover - {assume "\ k dvd c" - hence "Ifm (real (x*k)#bs) (a\ (Dvd i (CN 0 c e)) k) = - (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k)" - using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] - by (simp add: ring_simps) - also have "\ = (Ifm (real x#bs) (Dvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"]) - finally have ?case . } - ultimately show ?case by blast -next - case (10 i c e) - from prems have cp: "c > 0" and nb: "numbound0 e" by auto - from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp - {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } - moreover - {assume "\ k dvd c" - hence "Ifm (real (x*k)#bs) (a\ (NDvd i (CN 0 c e)) k) = - (\ (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k))" - using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] - by (simp add: ring_simps) - also have "\ = (Ifm (real x#bs) (NDvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"]) - finally have ?case . } - ultimately show ?case by blast -qed (simp_all add: nth_pos2) - -lemma a\_ex: - assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "k > 0" - shows "(\ (x::int). real k rdvd real x \ Ifm (real x#bs) (a\ p k)) = - (\ (x::int). Ifm (real x#bs) p)" (is "(\ x. ?D x \ ?P' x) = (\ x. ?P x)") -proof- - have "(\ x. ?D x \ ?P' x) = (\ x. k dvd x \ ?P' x)" using int_rdvd_iff by simp - also have "\ = (\x. ?P' (x*k))" using unity_coeff_ex[where P="?P'" and l="k", simplified] - by (simp add: ring_simps) - also have "\ = (\ x. ?P x)" using a\ iszlfm_gen[OF lp] kp by auto - finally show ?thesis . -qed - -lemma \\': assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "k > 0" and nb: "numbound0 t" - shows "Ifm (real x#bs) (\\ p (t,k)) = Ifm ((Inum (real x#bs) t)#bs) (a\ p k)" -using lp -by(induct p rule: \\.induct, simp_all add: - numbound0_I[OF nb, where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] - numbound0_I[where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] - bound0_I[where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] nth_pos2 cong: imp_cong) - -lemma \\_nb: assumes lp:"iszlfm p (a#bs)" and nb: "numbound0 t" - shows "bound0 (\\ p (t,k))" - using lp - by (induct p rule: iszlfm.induct, auto simp add: nb) - -lemma \_l: - assumes lp: "iszlfm p (real (i::int)#bs)" - shows "\ (b,k) \ set (\ p). k >0 \ numbound0 b \ isint b (real i#bs)" -using lp by (induct p rule: \.induct, auto simp add: isint_sub isint_neg) - -lemma \\_l: - assumes lp: "iszlfm p (real (i::int)#bs)" - shows "\ (b,k) \ set (\\ p). k >0 \ numbound0 b \ isint b (real i#bs)" -using lp isint_add [OF isint_c[where j="- 1"],where bs="real i#bs"] - by (induct p rule: \\.induct, auto) - -lemma zminusinf_\: - assumes lp: "iszlfm p (real (i::int)#bs)" - and nmi: "\ (Ifm (real i#bs) (minusinf p))" (is "\ (Ifm (real i#bs) (?M p))") - and ex: "Ifm (real i#bs) p" (is "?I i p") - shows "\ (e,c) \ set (\ p). real (c*i) > Inum (real i#bs) e" (is "\ (e,c) \ ?R p. real (c*i) > ?N i e") - using lp nmi ex -by (induct p rule: minusinf.induct, auto) - - -lemma \_And: "Ifm bs (\ (And p q) k t) = Ifm bs (And (\ p k t) (\ q k t))" -using \_def by auto -lemma \_Or: "Ifm bs (\ (Or p q) k t) = Ifm bs (Or (\ p k t) (\ q k t))" -using \_def by auto - -lemma \: assumes lp: "iszlfm p (real (i::int) #bs)" - and pi: "Ifm (real i#bs) p" - and d: "d\ p d" - and dp: "d > 0" - and nob: "\(e,c) \ set (\ p). \ j\ {1 .. c*d}. real (c*i) \ Inum (real i#bs) e + real j" - (is "\(e,c) \ set (\ p). \ j\ {1 .. c*d}. _ \ ?N i e + _") - shows "Ifm (real(i - d)#bs) p" - using lp pi d nob -proof(induct p rule: iszlfm.induct) - case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" - and pi: "real (c*i) = - 1 - ?N i e + real (1::int)" and nob: "\ j\ {1 .. c*d}. real (c*i) \ -1 - ?N i e + real j" - by simp+ - from mult_strict_left_mono[OF dp cp] have one:"1 \ {1 .. c*d}" by auto - from nob[rule_format, where j="1", OF one] pi show ?case by simp -next - case (4 c e) - hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" - and nob: "\ j\ {1 .. c*d}. real (c*i) \ - ?N i e + real j" - by simp+ - {assume "real (c*i) \ - ?N i e + real (c*d)" - with numbound0_I[OF nb, where bs="bs" and b="real i - real d" and b'="real i"] - have ?case by (simp add: ring_simps)} - moreover - {assume pi: "real (c*i) = - ?N i e + real (c*d)" - from mult_strict_left_mono[OF dp cp] have d: "(c*d) \ {1 .. c*d}" by simp - from nob[rule_format, where j="c*d", OF d] pi have ?case by simp } - ultimately show ?case by blast -next - case (5 c e) hence cp: "c > 0" by simp - from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] - real_of_int_mult] - show ?case using prems dp - by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] - ring_simps) -next - case (6 c e) hence cp: "c > 0" by simp - from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] - real_of_int_mult] - show ?case using prems dp - by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] - ring_simps) -next - case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" - and nob: "\ j\ {1 .. c*d}. real (c*i) \ - ?N i e + real j" - and pi: "real (c*i) + ?N i e > 0" and cp': "real c >0" - by simp+ - let ?fe = "floor (?N i e)" - from pi cp have th:"(real i +?N i e / real c)*real c > 0" by (simp add: ring_simps) - from pi ei[simplified isint_iff] have "real (c*i + ?fe) > real (0::int)" by simp - hence pi': "c*i + ?fe > 0" by (simp only: real_of_int_less_iff[symmetric]) - have "real (c*i) + ?N i e > real (c*d) \ real (c*i) + ?N i e \ real (c*d)" by auto - moreover - {assume "real (c*i) + ?N i e > real (c*d)" hence ?case - by (simp add: ring_simps - numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} - moreover - {assume H:"real (c*i) + ?N i e \ real (c*d)" - with ei[simplified isint_iff] have "real (c*i + ?fe) \ real (c*d)" by simp - hence pid: "c*i + ?fe \ c*d" by (simp only: real_of_int_le_iff) - with pi' have "\ j1\ {1 .. c*d}. c*i + ?fe = j1" by auto - hence "\ j1\ {1 .. c*d}. real (c*i) = - ?N i e + real j1" - by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] ring_simps) - with nob have ?case by blast } - ultimately show ?case by blast -next - case (8 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" - and nob: "\ j\ {1 .. c*d}. real (c*i) \ - 1 - ?N i e + real j" - and pi: "real (c*i) + ?N i e \ 0" and cp': "real c >0" - by simp+ - let ?fe = "floor (?N i e)" - from pi cp have th:"(real i +?N i e / real c)*real c \ 0" by (simp add: ring_simps) - from pi ei[simplified isint_iff] have "real (c*i + ?fe) \ real (0::int)" by simp - hence pi': "c*i + 1 + ?fe \ 1" by (simp only: real_of_int_le_iff[symmetric]) - have "real (c*i) + ?N i e \ real (c*d) \ real (c*i) + ?N i e < real (c*d)" by auto - moreover - {assume "real (c*i) + ?N i e \ real (c*d)" hence ?case - by (simp add: ring_simps - numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} - moreover - {assume H:"real (c*i) + ?N i e < real (c*d)" - with ei[simplified isint_iff] have "real (c*i + ?fe) < real (c*d)" by simp - hence pid: "c*i + 1 + ?fe \ c*d" by (simp only: real_of_int_le_iff) - with pi' have "\ j1\ {1 .. c*d}. c*i + 1+ ?fe = j1" by auto - hence "\ j1\ {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1" - by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] ring_simps real_of_one) - hence "\ j1\ {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1" - by (simp only: ring_simps diff_def[symmetric]) - hence "\ j1\ {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1" - by (simp only: add_ac diff_def) - with nob have ?case by blast } - ultimately show ?case by blast -next - case (9 j c e) hence p: "real j rdvd real (c*i) + ?N i e" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ - let ?e = "Inum (real i # bs) e" - from prems have "isint e (real i #bs)" by simp - hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real i)#bs"] numbound0_I[OF bn,where b="real i" and b'="real i" and bs="bs"] - by (simp add: isint_iff) - from prems have id: "j dvd d" by simp - from ie[symmetric] have "?p i = (real j rdvd real (c*i+ floor ?e))" by simp - also have "\ = (j dvd c*i + floor ?e)" - using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp - also have "\ = (j dvd c*i - c*d + floor ?e)" - using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp - also have "\ = (real j rdvd real (c*i - c*d + floor ?e))" - using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified] - ie by simp - also have "\ = (real j rdvd real (c*(i - d)) + ?e)" - using ie by (simp add:ring_simps) - finally show ?case - using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p - by (simp add: ring_simps) -next - case (10 j c e) hence p: "\ (real j rdvd real (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ - let ?e = "Inum (real i # bs) e" - from prems have "isint e (real i #bs)" by simp - hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real i)#bs"] numbound0_I[OF bn,where b="real i" and b'="real i" and bs="bs"] - by (simp add: isint_iff) - from prems have id: "j dvd d" by simp - from ie[symmetric] have "?p i = (\ (real j rdvd real (c*i+ floor ?e)))" by simp - also have "\ = Not (j dvd c*i + floor ?e)" - using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp - also have "\ = Not (j dvd c*i - c*d + floor ?e)" - using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp - also have "\ = Not (real j rdvd real (c*i - c*d + floor ?e))" - using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified] - ie by simp - also have "\ = Not (real j rdvd real (c*(i - d)) + ?e)" - using ie by (simp add:ring_simps) - finally show ?case - using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p - by (simp add: ring_simps) -qed(auto simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] nth_pos2) - -lemma \_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t" - shows "bound0 (\ p k t)" - using \\_nb[OF lp nb] nb by (simp add: \_def) - -lemma \': assumes lp: "iszlfm p (a #bs)" - and d: "d\ p d" - and dp: "d > 0" - shows "\ x. \(\ (e,c) \ set(\ p). \(j::int) \ {1 .. c*d}. Ifm (a #bs) (\ p c (Add e (C j)))) \ Ifm (real x#bs) p \ Ifm (real (x - d)#bs) p" (is "\ x. ?b x \ ?P x \ ?P (x - d)") -proof(clarify) - fix x - assume nob1:"?b x" and px: "?P x" - from iszlfm_gen[OF lp, rule_format, where y="real x"] have lp': "iszlfm p (real x#bs)". - have nob: "\(e, c)\set (\ p). \j\{1..c * d}. real (c * x) \ Inum (real x # bs) e + real j" - proof(clarify) - fix e c j assume ecR: "(e,c) \ set (\ p)" and jD: "j\ {1 .. c*d}" - and cx: "real (c*x) = Inum (real x#bs) e + real j" - let ?e = "Inum (real x#bs) e" - let ?fe = "floor ?e" - from \_l[OF lp'] ecR have ei:"isint e (real x#bs)" and cp:"c>0" and nb:"numbound0 e" - by auto - from numbound0_gen [OF nb ei, rule_format,where y="a"] have "isint e (a#bs)" . - from cx ei[simplified isint_iff] have "real (c*x) = real (?fe + j)" by simp - hence cx: "c*x = ?fe + j" by (simp only: real_of_int_inject) - hence cdej:"c dvd ?fe + j" by (simp add: dvd_def) (rule_tac x="x" in exI, simp) - hence "real c rdvd real (?fe + j)" by (simp only: int_rdvd_iff) - hence rcdej: "real c rdvd ?e + real j" by (simp add: ei[simplified isint_iff]) - from cx have "(c*x) div c = (?fe + j) div c" by simp - with cp have "x = (?fe + j) div c" by simp - with px have th: "?P ((?fe + j) div c)" by auto - from cp have cp': "real c > 0" by simp - from cdej have cdej': "c dvd floor (Inum (real x#bs) (Add e (C j)))" by simp - from nb have nb': "numbound0 (Add e (C j))" by simp - have ji: "isint (C j) (real x#bs)" by (simp add: isint_def) - from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real x#bs)" . - from th \\[where b'="real x", OF lp' cp' nb' ei' cdej',symmetric] - have "Ifm (real x#bs) (\\ p (Add e (C j), c))" by simp - with rcdej have th: "Ifm (real x#bs) (\ p c (Add e (C j)))" by (simp add: \_def) - from th bound0_I[OF \_nb[OF lp nb', where k="c"],where bs="bs" and b="real x" and b'="a"] - have "Ifm (a#bs) (\ p c (Add e (C j)))" by blast - with ecR jD nob1 show "False" by blast - qed - from \[OF lp' px d dp nob] show "?P (x -d )" . -qed - - -lemma rl_thm: - assumes lp: "iszlfm p (real (i::int)#bs)" - shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real j#bs) (minusinf p)) \ (\ (e,c) \ set (\ p). \ j\ {1 .. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" - (is "(\(x::int). ?P x) = ((\ j\ {1.. \ p}. ?MP j)\(\ (e,c) \ ?R. \ j\ _. ?SP c e j))" - is "?lhs = (?MD \ ?RD)" is "?lhs = ?rhs") -proof- - let ?d= "\ p" - from \[OF lp] have d:"d\ p ?d" and dp: "?d > 0" by auto - { assume H:"?MD" hence th:"\ (x::int). ?MP x" by blast - from H minusinf_ex[OF lp th] have ?thesis by blast} - moreover - { fix e c j assume exR:"(e,c) \ ?R" and jD:"j\ {1 .. c*?d}" and spx:"?SP c e j" - from exR \_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real i#bs)" and cp: "c > 0" - by auto - have "isint (C j) (real i#bs)" by (simp add: isint_iff) - with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real i"]] - have eji:"isint (Add e (C j)) (real i#bs)" by simp - from nb have nb': "numbound0 (Add e (C j))" by simp - from spx bound0_I[OF \_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real i"] - have spx': "Ifm (real i # bs) (\ p c (Add e (C j)))" by blast - from spx' have rcdej:"real c rdvd (Inum (real i#bs) (Add e (C j)))" - and sr:"Ifm (real i#bs) (\\ p (Add e (C j),c))" by (simp add: \_def)+ - from rcdej eji[simplified isint_iff] - have "real c rdvd real (floor (Inum (real i#bs) (Add e (C j))))" by simp - hence cdej:"c dvd floor (Inum (real i#bs) (Add e (C j)))" by (simp only: int_rdvd_iff) - from cp have cp': "real c > 0" by simp - from \\[OF lp cp' nb' eji cdej] spx' have "?P (\Inum (real i # bs) (Add e (C j))\ div c)" - by (simp add: \_def) - hence ?lhs by blast - with exR jD spx have ?thesis by blast} - moreover - { fix x assume px: "?P x" and nob: "\ ?RD" - from iszlfm_gen [OF lp,rule_format, where y="a"] have lp':"iszlfm p (a#bs)" . - from \'[OF lp' d dp, rule_format, OF nob] have th:"\ x. ?P x \ ?P (x - ?d)" by blast - from minusinf_inf[OF lp] obtain z where z:"\ x 0" by arith - from decr_lemma[OF dp,where x="x" and z="z"] - decr_mult_lemma[OF dp th zp, rule_format, OF px] z have th:"\ x. ?MP x" by auto - with minusinf_bex[OF lp] px nob have ?thesis by blast} - ultimately show ?thesis by blast -qed - -lemma mirror_\\: assumes lp: "iszlfm p (a#bs)" - shows "(\ (t,k). (Inum (a#bs) t, k)) ` set (\\ p) = (\ (t,k). (Inum (a#bs) t,k)) ` set (\ (mirror p))" -using lp -by (induct p rule: mirror.induct, simp_all add: split_def image_Un ) - -text {* The @{text "\"} part*} - -text{* Linearity for fm where Bound 0 ranges over @{text "\"}*} -consts - isrlfm :: "fm \ bool" (* Linearity test for fm *) -recdef isrlfm "measure size" - "isrlfm (And p q) = (isrlfm p \ isrlfm q)" - "isrlfm (Or p q) = (isrlfm p \ isrlfm q)" - "isrlfm (Eq (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm (NEq (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm (Lt (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm (Le (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm (Gt (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm (Ge (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm p = (isatom p \ (bound0 p))" - -constdefs fp :: "fm \ int \ num \ int \ fm" - "fp p n s j \ (if n > 0 then - (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j))))) - (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1)))))))) - else - (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C j))))) - (Gt (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C (j + 1)))))))))" - - (* splits the bounded from the unbounded part*) -consts rsplit0 :: "num \ (fm \ int \ num) list" -recdef rsplit0 "measure num_size" - "rsplit0 (Bound 0) = [(T,1,C 0)]" - "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b - in map (\ ((p,n,t),(q,m,s)). (And p q, n+m, Add t s)) [(a,b). a\acs,b\bcs])" - "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))" - "rsplit0 (Neg a) = map (\ (p,n,s). (p,-n,Neg s)) (rsplit0 a)" - "rsplit0 (Floor a) = foldl (op @) [] (map - (\ (p,n,s). if n=0 then [(p,0,Floor s)] - else (map (\ j. (fp p n s j, 0, Add (Floor s) (C j))) (if n > 0 then iupt (0,n) else iupt(n,0)))) - (rsplit0 a))" - "rsplit0 (CN 0 c a) = map (\ (p,n,s). (p,n+c,s)) (rsplit0 a)" - "rsplit0 (CN m c a) = map (\ (p,n,s). (p,n,CN m c s)) (rsplit0 a)" - "rsplit0 (CF c t s) = rsplit0 (Add (Mul c (Floor t)) s)" - "rsplit0 (Mul c a) = map (\ (p,n,s). (p,c*n,Mul c s)) (rsplit0 a)" - "rsplit0 t = [(T,0,t)]" - -lemma not_rl[simp]: "isrlfm p \ isrlfm (not p)" - by (induct p rule: isrlfm.induct, auto) -lemma conj_rl[simp]: "isrlfm p \ isrlfm q \ isrlfm (conj p q)" - using conj_def by (cases p, auto) -lemma disj_rl[simp]: "isrlfm p \ isrlfm q \ isrlfm (disj p q)" - using disj_def by (cases p, auto) - - -lemma rsplit0_cs: - shows "\ (p,n,s) \ set (rsplit0 t). - (Ifm (x#bs) p \ (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \ numbound0 s \ isrlfm p" - (is "\ (p,n,s) \ ?SS t. (?I p \ ?N t = ?N (CN 0 n s)) \ _ \ _ ") -proof(induct t rule: rsplit0.induct) - case (5 a) - let ?p = "\ (p,n,s) j. fp p n s j" - let ?f = "(\ (p,n,s) j. (?p (p,n,s) j, (0::int),Add (Floor s) (C j)))" - let ?J = "\ n. if n>0 then iupt (0,n) else iupt (n,0)" - let ?ff=" (\ (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))" - have int_cases: "\ (i::int). i= 0 \ i < 0 \ i > 0" by arith - have U1: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) = - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)]))" by auto - have U2': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n>0}. - ?ff (p,n,s) = map (?f(p,n,s)) (iupt(0,n))" by auto - hence U2: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) = - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). - set (map (?f(p,n,s)) (iupt(0,n)))))" - proof- - fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g - assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" - thus "(UNION M (\ (a,b,c). set (f (a,b,c)))) = (UNION M (\ (a,b,c). set (g a b c)))" - by (auto simp add: split_def) - qed - have U3': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n<0}. ?ff (p,n,s) = map (?f(p,n,s)) (iupt(n,0))" - by auto - hence U3: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))) = - (UNION {(p,n,s). (p,n,s)\ ?SS a\n<0} (\(p,n,s). set (map (?f(p,n,s)) (iupt(n,0)))))" - proof- - fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g - assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" - thus "(UNION M (\ (a,b,c). set (f (a,b,c)))) = (UNION M (\ (a,b,c). set (g a b c)))" - by (auto simp add: split_def) - qed - have "?SS (Floor a) = UNION (?SS a) (\x. set (?ff x))" - by (auto simp add: foldl_conv_concat) - also have "\ = UNION (?SS a) (\ (p,n,s). set (?ff (p,n,s)))" by auto - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))))" - using int_cases[rule_format] by blast - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)])) Un - (UNION {(p,n,s). (p,n,s)\ ?SS a\n>0} (\(p,n,s). set(map(?f(p,n,s)) (iupt(0,n))))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). - set (map (?f(p,n,s)) (iupt(n,0))))))" by (simp only: U1 U2 U3) - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). (?f(p,n,s)) ` {n .. 0})))" - by (simp only: set_map iupt_set set.simps) - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast - finally - have FS: "?SS (Floor a) = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast - show ?case - proof(simp only: FS, clarsimp simp del: Ifm.simps Inum.simps, -) - fix p n s - let ?ths = "(?I p \ (?N (Floor a) = ?N (CN 0 n s))) \ numbound0 s \ isrlfm p" - assume "(\ba. (p, 0, ba) \ set (rsplit0 a) \ n = 0 \ s = Floor ba) \ - (\ab ac ba. - (ab, ac, ba) \ set (rsplit0 a) \ - 0 < ac \ - (\j. p = fp ab ac ba j \ - n = 0 \ s = Add (Floor ba) (C j) \ 0 \ j \ j \ ac)) \ - (\ab ac ba. - (ab, ac, ba) \ set (rsplit0 a) \ - ac < 0 \ - (\j. p = fp ab ac ba j \ - n = 0 \ s = Add (Floor ba) (C j) \ ac \ j \ j \ 0))" - moreover - {fix s' - assume "(p, 0, s') \ ?SS a" and "n = 0" and "s = Floor s'" - hence ?ths using prems by auto} - moreover - { fix p' n' s' j - assume pns: "(p', n', s') \ ?SS a" - and np: "0 < n'" - and p_def: "p = ?p (p',n',s') j" - and n0: "n = 0" - and s_def: "s = (Add (Floor s') (C j))" - and jp: "0 \ j" and jn: "j \ n'" - from prems pns have H:"(Ifm ((x\real) # (bs\real list)) p' \ - Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \ - numbound0 s' \ isrlfm p'" by blast - hence nb: "numbound0 s'" by simp - from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numsub_nb) - let ?nxs = "CN 0 n' s'" - let ?l = "floor (?N s') + j" - from H - have "?I (?p (p',n',s') j) \ - (((?N ?nxs \ real ?l) \ (?N ?nxs < real (?l + 1))) \ (?N a = ?N ?nxs ))" - by (simp add: fp_def np ring_simps numsub numadd numfloor) - also have "\ \ ((floor (?N ?nxs) = ?l) \ (?N a = ?N ?nxs ))" - using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp - moreover - have "\ \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp - ultimately have "?I (?p (p',n',s') j) \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" - by blast - with s_def n0 p_def nb nf have ?ths by auto} - moreover - {fix p' n' s' j - assume pns: "(p', n', s') \ ?SS a" - and np: "n' < 0" - and p_def: "p = ?p (p',n',s') j" - and n0: "n = 0" - and s_def: "s = (Add (Floor s') (C j))" - and jp: "n' \ j" and jn: "j \ 0" - from prems pns have H:"(Ifm ((x\real) # (bs\real list)) p' \ - Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \ - numbound0 s' \ isrlfm p'" by blast - hence nb: "numbound0 s'" by simp - from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numneg_nb) - let ?nxs = "CN 0 n' s'" - let ?l = "floor (?N s') + j" - from H - have "?I (?p (p',n',s') j) \ - (((?N ?nxs \ real ?l) \ (?N ?nxs < real (?l + 1))) \ (?N a = ?N ?nxs ))" - by (simp add: np fp_def ring_simps numneg numfloor numadd numsub) - also have "\ \ ((floor (?N ?nxs) = ?l) \ (?N a = ?N ?nxs ))" - using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp - moreover - have "\ \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp - ultimately have "?I (?p (p',n',s') j) \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" - by blast - with s_def n0 p_def nb nf have ?ths by auto} - ultimately show ?ths by auto - qed -next - case (3 a b) then show ?case - apply auto - apply (erule_tac x = "(aa, aaa, ba)" in ballE) apply simp_all - apply (erule_tac x = "(ab, ac, baa)" in ballE) apply simp_all - done -qed (auto simp add: Let_def split_def ring_simps conj_rl) - -lemma real_in_int_intervals: - assumes xb: "real m \ x \ x < real ((n::int) + 1)" - shows "\ j\ {m.. n}. real j \ x \ x < real (j+1)" (is "\ j\ ?N. ?P j") -by (rule bexI[where P="?P" and x="floor x" and A="?N"]) -(auto simp add: floor_less_eq[where x="x" and a="n+1", simplified] xb[simplified] floor_mono2[where x="real m" and y="x", OF conjunct1[OF xb], simplified floor_real_of_int[where n="m"]]) - -lemma rsplit0_complete: - assumes xp:"0 \ x" and x1:"x < 1" - shows "\ (p,n,s) \ set (rsplit0 t). Ifm (x#bs) p" (is "\ (p,n,s) \ ?SS t. ?I p") -proof(induct t rule: rsplit0.induct) - case (2 a b) - from prems have "\ (pa,na,sa) \ ?SS a. ?I pa" by auto - then obtain "pa" "na" "sa" where pa: "(pa,na,sa)\ ?SS a \ ?I pa" by blast - from prems have "\ (pb,nb,sb) \ ?SS b. ?I pb" by auto - then obtain "pb" "nb" "sb" where pb: "(pb,nb,sb)\ ?SS b \ ?I pb" by blast - from pa pb have th: "((pa,na,sa),(pb,nb,sb)) \ set[(x,y). x\rsplit0 a, y\rsplit0 b]" - by (auto) - let ?f="(\ ((p,n,t),(q,m,s)). (And p q, n+m, Add t s))" - from imageI[OF th, where f="?f"] have "?f ((pa,na,sa),(pb,nb,sb)) \ ?SS (Add a b)" - by (simp add: Let_def) - hence "(And pa pb, na +nb, Add sa sb) \ ?SS (Add a b)" by simp - moreover from pa pb have "?I (And pa pb)" by simp - ultimately show ?case by blast -next - case (5 a) - let ?p = "\ (p,n,s) j. fp p n s j" - let ?f = "(\ (p,n,s) j. (?p (p,n,s) j, (0::int),(Add (Floor s) (C j))))" - let ?J = "\ n. if n>0 then iupt (0,n) else iupt (n,0)" - let ?ff=" (\ (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))" - have int_cases: "\ (i::int). i= 0 \ i < 0 \ i > 0" by arith - have U1: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)]))" by auto - have U2': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n>0}. ?ff (p,n,s) = map (?f(p,n,s)) (iupt(0,n))" - by auto - hence U2: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (map (?f(p,n,s)) (iupt(0,n)))))" - proof- - fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g - assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" - thus "(UNION M (\ (a,b,c). set (f (a,b,c)))) = (UNION M (\ (a,b,c). set (g a b c)))" - by (auto simp add: split_def) - qed - have U3': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n<0}. ?ff (p,n,s) = map (?f(p,n,s)) (iupt(n,0))" - by auto - hence U3: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (map (?f(p,n,s)) (iupt(n,0)))))" - proof- - fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g - assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" - thus "(UNION M (\ (a,b,c). set (f (a,b,c)))) = (UNION M (\ (a,b,c). set (g a b c)))" - by (auto simp add: split_def) - qed - - have "?SS (Floor a) = UNION (?SS a) (\x. set (?ff x))" by (auto simp add: foldl_conv_concat) - also have "\ = UNION (?SS a) (\ (p,n,s). set (?ff (p,n,s)))" by auto - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))))" - using int_cases[rule_format] by blast - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)])) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (map (?f(p,n,s)) (iupt(0,n))))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (map (?f(p,n,s)) (iupt(n,0))))))" by (simp only: U1 U2 U3) - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). (?f(p,n,s)) ` {n .. 0})))" - by (simp only: set_map iupt_set set.simps) - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast - finally - have FS: "?SS (Floor a) = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast - from prems have "\ (p,n,s) \ ?SS a. ?I p" by auto - then obtain "p" "n" "s" where pns: "(p,n,s) \ ?SS a \ ?I p" by blast - let ?N = "\ t. Inum (x#bs) t" - from rsplit0_cs[rule_format] pns have ans:"(?N a = ?N (CN 0 n s)) \ numbound0 s \ isrlfm p" - by auto - - have "n=0 \ n >0 \ n <0" by arith - moreover {assume "n=0" hence ?case using pns by (simp only: FS) auto } - moreover - { - assume np: "n > 0" - from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) \ ?N s" by simp - also from mult_left_mono[OF xp] np have "?N s \ real n * x + ?N s" by simp - finally have "?N (Floor s) \ real n * x + ?N s" . - moreover - {from mult_strict_left_mono[OF x1] np - have "real n *x + ?N s < real n + ?N s" by simp - also from real_of_int_floor_add_one_gt[where r="?N s"] - have "\ < real n + ?N (Floor s) + 1" by simp - finally have "real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp} - ultimately have "?N (Floor s) \ real n *x + ?N s\ real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp - hence th: "0 \ real n *x + ?N s - ?N (Floor s) \ real n *x + ?N s - ?N (Floor s) < real (n+1)" by simp - from real_in_int_intervals th have "\ j\ {0 .. n}. real j \ real n *x + ?N s - ?N (Floor s)\ real n *x + ?N s - ?N (Floor s) < real (j+1)" by simp - - hence "\ j\ {0 .. n}. 0 \ real n *x + ?N s - ?N (Floor s) - real j \ real n *x + ?N s - ?N (Floor s) - real (j+1) < 0" - by(simp only: myl[rule_format, where b="real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) - hence "\ j\ {0.. n}. ?I (?p (p,n,s) j)" - using pns by (simp add: fp_def np ring_simps numsub numadd) - then obtain "j" where j_def: "j\ {0 .. n} \ ?I (?p (p,n,s) j)" by blast - hence "\x \ {?p (p,n,s) j |j. 0\ j \ j \ n }. ?I x" by auto - hence ?case using pns - by (simp only: FS,simp add: bex_Un) - (rule disjI2, rule disjI1,rule exI [where x="p"], - rule exI [where x="n"],rule exI [where x="s"],simp_all add: np) - } - moreover - { assume nn: "n < 0" hence np: "-n >0" by simp - from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) + 1 > ?N s" by simp - moreover from mult_left_mono_neg[OF xp] nn have "?N s \ real n * x + ?N s" by simp - ultimately have "?N (Floor s) + 1 > real n * x + ?N s" by arith - moreover - {from mult_strict_left_mono_neg[OF x1, where c="real n"] nn - have "real n *x + ?N s \ real n + ?N s" by simp - moreover from real_of_int_floor_le[where r="?N s"] have "real n + ?N s \ real n + ?N (Floor s)" by simp - ultimately have "real n *x + ?N s \ ?N (Floor s) + real n" - by (simp only: ring_simps)} - ultimately have "?N (Floor s) + real n \ real n *x + ?N s\ real n *x + ?N s < ?N (Floor s) + real (1::int)" by simp - hence th: "real n \ real n *x + ?N s - ?N (Floor s) \ real n *x + ?N s - ?N (Floor s) < real (1::int)" by simp - have th1: "\ (a::real). (- a > 0) = (a < 0)" by auto - have th2: "\ (a::real). (0 \ - a) = (a \ 0)" by auto - from real_in_int_intervals th have "\ j\ {n .. 0}. real j \ real n *x + ?N s - ?N (Floor s)\ real n *x + ?N s - ?N (Floor s) < real (j+1)" by simp - - hence "\ j\ {n .. 0}. 0 \ real n *x + ?N s - ?N (Floor s) - real j \ real n *x + ?N s - ?N (Floor s) - real (j+1) < 0" - by(simp only: myl[rule_format, where b="real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) - hence "\ j\ {n .. 0}. 0 \ - (real n *x + ?N s - ?N (Floor s) - real j) \ - (real n *x + ?N s - ?N (Floor s) - real (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format]) - hence "\ j\ {n.. 0}. ?I (?p (p,n,s) j)" - using pns by (simp add: fp_def nn diff_def add_ac mult_ac numfloor numadd numneg - del: diff_less_0_iff_less diff_le_0_iff_le) - then obtain "j" where j_def: "j\ {n .. 0} \ ?I (?p (p,n,s) j)" by blast - hence "\x \ {?p (p,n,s) j |j. n\ j \ j \ 0 }. ?I x" by auto - hence ?case using pns - by (simp only: FS,simp add: bex_Un) - (rule disjI2, rule disjI2,rule exI [where x="p"], - rule exI [where x="n"],rule exI [where x="s"],simp_all add: nn) - } - ultimately show ?case by blast -qed (auto simp add: Let_def split_def) - - (* Linearize a formula where Bound 0 ranges over [0,1) *) - -constdefs rsplit :: "(int \ num \ fm) \ num \ fm" - "rsplit f a \ foldr disj (map (\ (\, n, s). conj \ (f n s)) (rsplit0 a)) F" - -lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\ x \ set xs. Ifm bs (f x))" -by(induct xs, simp_all) - -lemma foldr_conj_map: "Ifm bs (foldr conj (map f xs) T) = (\ x \ set xs. Ifm bs (f x))" -by(induct xs, simp_all) - -lemma foldr_disj_map_rlfm: - assumes lf: "\ n s. numbound0 s \ isrlfm (f n s)" - and \: "\ (\,n,s) \ set xs. numbound0 s \ isrlfm \" - shows "isrlfm (foldr disj (map (\ (\, n, s). conj \ (f n s)) xs) F)" -using lf \ by (induct xs, auto) - -lemma rsplit_ex: "Ifm bs (rsplit f a) = (\ (\,n,s) \ set (rsplit0 a). Ifm bs (conj \ (f n s)))" -using foldr_disj_map[where xs="rsplit0 a"] rsplit_def by (simp add: split_def) - -lemma rsplit_l: assumes lf: "\ n s. numbound0 s \ isrlfm (f n s)" - shows "isrlfm (rsplit f a)" -proof- - from rsplit0_cs[where t="a"] have th: "\ (\,n,s) \ set (rsplit0 a). numbound0 s \ isrlfm \" by blast - from foldr_disj_map_rlfm[OF lf th] rsplit_def show ?thesis by simp -qed - -lemma rsplit: - assumes xp: "x \ 0" and x1: "x < 1" - and f: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ (Ifm (x#bs) (f n s) = Ifm (x#bs) (g a))" - shows "Ifm (x#bs) (rsplit f a) = Ifm (x#bs) (g a)" -proof(auto) - let ?I = "\x p. Ifm (x#bs) p" - let ?N = "\ x t. Inum (x#bs) t" - assume "?I x (rsplit f a)" - hence "\ (\,n,s) \ set (rsplit0 a). ?I x (And \ (f n s))" using rsplit_ex by simp - then obtain "\" "n" "s" where fnsS:"(\,n,s) \ set (rsplit0 a)" and "?I x (And \ (f n s))" by blast - hence \: "?I x \" and fns: "?I x (f n s)" by auto - from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \ - have th: "(?N x a = ?N x (CN 0 n s)) \ numbound0 s" by auto - from f[rule_format, OF th] fns show "?I x (g a)" by simp -next - let ?I = "\x p. Ifm (x#bs) p" - let ?N = "\ x t. Inum (x#bs) t" - assume ga: "?I x (g a)" - from rsplit0_complete[OF xp x1, where bs="bs" and t="a"] - obtain "\" "n" "s" where fnsS:"(\,n,s) \ set (rsplit0 a)" and fx: "?I x \" by blast - from rsplit0_cs[where t="a" and x="x" and bs="bs"] fnsS fx - have ans: "?N x a = ?N x (CN 0 n s)" and nb: "numbound0 s" by auto - with ga f have "?I x (f n s)" by auto - with rsplit_ex fnsS fx show "?I x (rsplit f a)" by auto -qed - -definition lt :: "int \ num \ fm" where - lt_def: "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) - else (Gt (CN 0 (-c) (Neg t))))" - -definition le :: "int \ num \ fm" where - le_def: "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) - else (Ge (CN 0 (-c) (Neg t))))" - -definition gt :: "int \ num \ fm" where - gt_def: "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) - else (Lt (CN 0 (-c) (Neg t))))" - -definition ge :: "int \ num \ fm" where - ge_def: "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) - else (Le (CN 0 (-c) (Neg t))))" - -definition eq :: "int \ num \ fm" where - eq_def: "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) - else (Eq (CN 0 (-c) (Neg t))))" - -definition neq :: "int \ num \ fm" where - neq_def: "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) - else (NEq (CN 0 (-c) (Neg t))))" - -lemma lt_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (lt n s) = Ifm (x#bs) (Lt a)" - (is "\ a n s . ?N a = ?N (CN 0 n s) \ _\ ?I (lt n s) = ?I (Lt a)") -proof(clarify) - fix a n s - assume H: "?N a = ?N (CN 0 n s)" - show "?I (lt n s) = ?I (Lt a)" using H by (cases "n=0", (simp add: lt_def)) - (cases "n > 0", simp_all add: lt_def ring_simps myless[rule_format, where b="0"]) -qed - -lemma lt_l: "isrlfm (rsplit lt a)" - by (rule rsplit_l[where f="lt" and a="a"], auto simp add: lt_def, - case_tac s, simp_all, case_tac "nat", simp_all) - -lemma le_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (le n s) = Ifm (x#bs) (Le a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (le n s) = ?I (Le a)") -proof(clarify) - fix a n s - assume H: "?N a = ?N (CN 0 n s)" - show "?I (le n s) = ?I (Le a)" using H by (cases "n=0", (simp add: le_def)) - (cases "n > 0", simp_all add: le_def ring_simps myl[rule_format, where b="0"]) -qed - -lemma le_l: "isrlfm (rsplit le a)" - by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def) -(case_tac s, simp_all, case_tac "nat",simp_all) - -lemma gt_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (gt n s) = Ifm (x#bs) (Gt a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (gt n s) = ?I (Gt a)") -proof(clarify) - fix a n s - assume H: "?N a = ?N (CN 0 n s)" - show "?I (gt n s) = ?I (Gt a)" using H by (cases "n=0", (simp add: gt_def)) - (cases "n > 0", simp_all add: gt_def ring_simps myless[rule_format, where b="0"]) -qed -lemma gt_l: "isrlfm (rsplit gt a)" - by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def) -(case_tac s, simp_all, case_tac "nat", simp_all) - -lemma ge_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (ge n s) = Ifm (x#bs) (Ge a)" (is "\ a n s . ?N a = ?N (CN 0 n s) \ _ \ ?I (ge n s) = ?I (Ge a)") -proof(clarify) - fix a n s - assume H: "?N a = ?N (CN 0 n s)" - show "?I (ge n s) = ?I (Ge a)" using H by (cases "n=0", (simp add: ge_def)) - (cases "n > 0", simp_all add: ge_def ring_simps myl[rule_format, where b="0"]) -qed -lemma ge_l: "isrlfm (rsplit ge a)" - by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def) -(case_tac s, simp_all, case_tac "nat", simp_all) - -lemma eq_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (eq n s) = Ifm (x#bs) (Eq a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (eq n s) = ?I (Eq a)") -proof(clarify) - fix a n s - assume H: "?N a = ?N (CN 0 n s)" - show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def ring_simps) -qed -lemma eq_l: "isrlfm (rsplit eq a)" - by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def) -(case_tac s, simp_all, case_tac"nat", simp_all) - -lemma neq_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (neq n s) = Ifm (x#bs) (NEq a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (neq n s) = ?I (NEq a)") -proof(clarify) - fix a n s bs - assume H: "?N a = ?N (CN 0 n s)" - show "?I (neq n s) = ?I (NEq a)" using H by (auto simp add: neq_def ring_simps) -qed - -lemma neq_l: "isrlfm (rsplit neq a)" - by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def) -(case_tac s, simp_all, case_tac"nat", simp_all) - -lemma small_le: - assumes u0:"0 \ u" and u1: "u < 1" - shows "(-u \ real (n::int)) = (0 \ n)" -using u0 u1 by auto - -lemma small_lt: - assumes u0:"0 \ u" and u1: "u < 1" - shows "(real (n::int) < real (m::int) - u) = (n < m)" -using u0 u1 by auto - -lemma rdvd01_cs: - assumes up: "u \ 0" and u1: "u<1" and np: "real n > 0" - shows "(real (i::int) rdvd real (n::int) * u - s) = (\ j\ {0 .. n - 1}. real n * u = s - real (floor s) + real j \ real i rdvd real (j - floor s))" (is "?lhs = ?rhs") -proof- - let ?ss = "s - real (floor s)" - from real_of_int_floor_add_one_gt[where r="s", simplified myless[rule_format,where a="s"]] - real_of_int_floor_le[where r="s"] have ss0:"?ss \ 0" and ss1:"?ss < 1" - by (auto simp add: myl[rule_format, where b="s", symmetric] myless[rule_format, where a="?ss"]) - from np have n0: "real n \ 0" by simp - from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] - have nu0:"real n * u - s \ -s" and nun:"real n * u -s < real n - s" by auto - from int_rdvd_real[where i="i" and x="real (n::int) * u - s"] - have "real i rdvd real n * u - s = - (i dvd floor (real n * u -s) \ (real (floor (real n * u - s)) = real n * u - s ))" - (is "_ = (?DE)" is "_ = (?D \ ?E)") by simp - also have "\ = (?DE \ real(floor (real n * u - s) + floor s)\ -?ss - \ real(floor (real n * u - s) + floor s)< real n - ?ss)" (is "_=(?DE \real ?a \ _ \ real ?a < _)") - using nu0 nun by auto - also have "\ = (?DE \ ?a \ 0 \ ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1]) - also have "\ = (?DE \ (\ j\ {0 .. (n - 1)}. ?a = j))" by simp - also have "\ = (?DE \ (\ j\ {0 .. (n - 1)}. real (\real n * u - s\) = real j - real \s\ ))" - by (simp only: ring_simps real_of_int_diff[symmetric] real_of_int_inject del: real_of_int_diff) - also have "\ = ((\ j\ {0 .. (n - 1)}. real n * u - s = real j - real \s\ \ real i rdvd real n * u - s))" using int_rdvd_iff[where i="i" and t="\real n * u - s\"] - by (auto cong: conj_cong) - also have "\ = ?rhs" by(simp cong: conj_cong) (simp add: ring_simps ) - finally show ?thesis . -qed - -definition - DVDJ:: "int \ int \ num \ fm" -where - DVDJ_def: "DVDJ i n s = (foldr disj (map (\ j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))) (iupt(0,n - 1))) F)" - -definition - NDVDJ:: "int \ int \ num \ fm" -where - NDVDJ_def: "NDVDJ i n s = (foldr conj (map (\ j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor (Neg s))))) (iupt(0,n - 1))) T)" - -lemma DVDJ_DVD: - assumes xp:"x\ 0" and x1: "x < 1" and np:"real n > 0" - shows "Ifm (x#bs) (DVDJ i n s) = Ifm (x#bs) (Dvd i (CN 0 n s))" -proof- - let ?f = "\ j. conj (Eq(CN 0 n (Add s (Sub(Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))" - let ?s= "Inum (x#bs) s" - from foldr_disj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"] - have "Ifm (x#bs) (DVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" - by (simp add: iupt_set np DVDJ_def del: iupt.simps) - also have "\ = (\ j\ {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \ real i rdvd real (j - floor (- ?s)))" by (simp add: ring_simps diff_def[symmetric]) - also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] - have "\ = (real i rdvd real n * x - (-?s))" by simp - finally show ?thesis by simp -qed - -lemma NDVDJ_NDVD: - assumes xp:"x\ 0" and x1: "x < 1" and np:"real n > 0" - shows "Ifm (x#bs) (NDVDJ i n s) = Ifm (x#bs) (NDvd i (CN 0 n s))" -proof- - let ?f = "\ j. disj(NEq(CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor(Neg s))))" - let ?s= "Inum (x#bs) s" - from foldr_conj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"] - have "Ifm (x#bs) (NDVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" - by (simp add: iupt_set np NDVDJ_def del: iupt.simps) - also have "\ = (\ (\ j\ {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \ real i rdvd real (j - floor (- ?s))))" by (simp add: ring_simps diff_def[symmetric]) - also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] - have "\ = (\ (real i rdvd real n * x - (-?s)))" by simp - finally show ?thesis by simp -qed - -lemma foldr_disj_map_rlfm2: - assumes lf: "\ n . isrlfm (f n)" - shows "isrlfm (foldr disj (map f xs) F)" -using lf by (induct xs, auto) -lemma foldr_And_map_rlfm2: - assumes lf: "\ n . isrlfm (f n)" - shows "isrlfm (foldr conj (map f xs) T)" -using lf by (induct xs, auto) - -lemma DVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s" - shows "isrlfm (DVDJ i n s)" -proof- - let ?f="\j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) - (Dvd i (Sub (C j) (Floor (Neg s))))" - have th: "\ j. isrlfm (?f j)" using nb np by auto - from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis by simp -qed - -lemma NDVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s" - shows "isrlfm (NDVDJ i n s)" -proof- - let ?f="\j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) - (NDvd i (Sub (C j) (Floor (Neg s))))" - have th: "\ j. isrlfm (?f j)" using nb np by auto - from NDVDJ_def foldr_And_map_rlfm2[OF th] show ?thesis by auto -qed - -definition DVD :: "int \ int \ num \ fm" where - DVD_def: "DVD i c t = - (if i=0 then eq c t else - if c = 0 then (Dvd i t) else if c >0 then DVDJ (abs i) c t else DVDJ (abs i) (-c) (Neg t))" - -definition NDVD :: "int \ int \ num \ fm" where - "NDVD i c t = - (if i=0 then neq c t else - if c = 0 then (NDvd i t) else if c >0 then NDVDJ (abs i) c t else NDVDJ (abs i) (-c) (Neg t))" - -lemma DVD_mono: - assumes xp: "0\ x" and x1: "x < 1" - shows "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (DVD i n s) = Ifm (x#bs) (Dvd i a)" - (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (DVD i n s) = ?I (Dvd i a)") -proof(clarify) - fix a n s - assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s" - let ?th = "?I (DVD i n s) = ?I (Dvd i a)" - have "i=0 \ (i\0 \ n=0) \ (i\0 \ n < 0) \ (i\0 \ n > 0)" by arith - moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]] - by (simp add: DVD_def rdvd_left_0_eq)} - moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H DVD_def) } - moreover {assume inz: "i\0" and "n<0" hence ?th - by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 - rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } - moreover {assume inz: "i\0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)} - ultimately show ?th by blast -qed - -lemma NDVD_mono: assumes xp: "0\ x" and x1: "x < 1" - shows "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (NDVD i n s) = Ifm (x#bs) (NDvd i a)" - (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (NDVD i n s) = ?I (NDvd i a)") -proof(clarify) - fix a n s - assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s" - let ?th = "?I (NDVD i n s) = ?I (NDvd i a)" - have "i=0 \ (i\0 \ n=0) \ (i\0 \ n < 0) \ (i\0 \ n > 0)" by arith - moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]] - by (simp add: NDVD_def rdvd_left_0_eq)} - moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H NDVD_def) } - moreover {assume inz: "i\0" and "n<0" hence ?th - by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 - rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } - moreover {assume inz: "i\0" and "n>0" hence ?th - by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)} - ultimately show ?th by blast -qed - -lemma DVD_l: "isrlfm (rsplit (DVD i) a)" - by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l) -(case_tac s, simp_all, case_tac "nat", simp_all) - -lemma NDVD_l: "isrlfm (rsplit (NDVD i) a)" - by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l) -(case_tac s, simp_all, case_tac "nat", simp_all) - -consts rlfm :: "fm \ fm" -recdef rlfm "measure fmsize" - "rlfm (And p q) = conj (rlfm p) (rlfm q)" - "rlfm (Or p q) = disj (rlfm p) (rlfm q)" - "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)" - "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))" - "rlfm (Lt a) = rsplit lt a" - "rlfm (Le a) = rsplit le a" - "rlfm (Gt a) = rsplit gt a" - "rlfm (Ge a) = rsplit ge a" - "rlfm (Eq a) = rsplit eq a" - "rlfm (NEq a) = rsplit neq a" - "rlfm (Dvd i a) = rsplit (\ t. DVD i t) a" - "rlfm (NDvd i a) = rsplit (\ t. NDVD i t) a" - "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))" - "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))" - "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))" - "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))" - "rlfm (NOT (NOT p)) = rlfm p" - "rlfm (NOT T) = F" - "rlfm (NOT F) = T" - "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))" - "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))" - "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))" - "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))" - "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))" - "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))" - "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))" - "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))" - "rlfm p = p" (hints simp add: fmsize_pos) - -lemma bound0at_l : "\isatom p ; bound0 p\ \ isrlfm p" - by (induct p rule: isrlfm.induct, auto) -lemma zgcd_le1: assumes ip: "0 < i" shows "zgcd i j \ i" -proof- - from zgcd_zdvd1 have th: "zgcd i j dvd i" by blast - from zdvd_imp_le[OF th ip] show ?thesis . -qed - - -lemma simpfm_rl: "isrlfm p \ isrlfm (simpfm p)" -proof (induct p) - case (Lt a) - hence "bound0 (Lt a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" - by (cases a,simp_all, case_tac "nat", simp_all) - moreover - {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt a))" - using simpfm_bound0 by blast - have "isatom (simpfm (Lt a))" by (cases "simpnum a", auto simp add: Let_def) - with bn bound0at_l have ?case by blast} - moreover - {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" - { - assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" - with numgcd_pos[where t="CN 0 c (simpnum e)"] - have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" - by (simp add: numgcd_def zgcd_le1) - from prems have th': "c\0" by auto - from prems have cp: "c \ 0" by simp - from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] - have "0 < c div numgcd (CN 0 c (simpnum e))" by simp - } - with prems have ?case - by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} - ultimately show ?case by blast -next - case (Le a) - hence "bound0 (Le a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" - by (cases a,simp_all, case_tac "nat", simp_all) - moreover - {assume "bound0 (Le a)" hence bn:"bound0 (simpfm (Le a))" - using simpfm_bound0 by blast - have "isatom (simpfm (Le a))" by (cases "simpnum a", auto simp add: Let_def) - with bn bound0at_l have ?case by blast} - moreover - {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" - { - assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" - with numgcd_pos[where t="CN 0 c (simpnum e)"] - have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" - by (simp add: numgcd_def zgcd_le1) - from prems have th': "c\0" by auto - from prems have cp: "c \ 0" by simp - from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] - have "0 < c div numgcd (CN 0 c (simpnum e))" by simp - } - with prems have ?case - by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} - ultimately show ?case by blast -next - case (Gt a) - hence "bound0 (Gt a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" - by (cases a,simp_all, case_tac "nat", simp_all) - moreover - {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt a))" - using simpfm_bound0 by blast - have "isatom (simpfm (Gt a))" by (cases "simpnum a", auto simp add: Let_def) - with bn bound0at_l have ?case by blast} - moreover - {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" - { - assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" - with numgcd_pos[where t="CN 0 c (simpnum e)"] - have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" - by (simp add: numgcd_def zgcd_le1) - from prems have th': "c\0" by auto - from prems have cp: "c \ 0" by simp - from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] - have "0 < c div numgcd (CN 0 c (simpnum e))" by simp - } - with prems have ?case - by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} - ultimately show ?case by blast -next - case (Ge a) - hence "bound0 (Ge a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" - by (cases a,simp_all, case_tac "nat", simp_all) - moreover - {assume "bound0 (Ge a)" hence bn:"bound0 (simpfm (Ge a))" - using simpfm_bound0 by blast - have "isatom (simpfm (Ge a))" by (cases "simpnum a", auto simp add: Let_def) - with bn bound0at_l have ?case by blast} - moreover - {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" - { - assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" - with numgcd_pos[where t="CN 0 c (simpnum e)"] - have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" - by (simp add: numgcd_def zgcd_le1) - from prems have th': "c\0" by auto - from prems have cp: "c \ 0" by simp - from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] - have "0 < c div numgcd (CN 0 c (simpnum e))" by simp - } - with prems have ?case - by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} - ultimately show ?case by blast -next - case (Eq a) - hence "bound0 (Eq a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" - by (cases a,simp_all, case_tac "nat", simp_all) - moreover - {assume "bound0 (Eq a)" hence bn:"bound0 (simpfm (Eq a))" - using simpfm_bound0 by blast - have "isatom (simpfm (Eq a))" by (cases "simpnum a", auto simp add: Let_def) - with bn bound0at_l have ?case by blast} - moreover - {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" - { - assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" - with numgcd_pos[where t="CN 0 c (simpnum e)"] - have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" - by (simp add: numgcd_def zgcd_le1) - from prems have th': "c\0" by auto - from prems have cp: "c \ 0" by simp - from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] - have "0 < c div numgcd (CN 0 c (simpnum e))" by simp - } - with prems have ?case - by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} - ultimately show ?case by blast -next - case (NEq a) - hence "bound0 (NEq a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" - by (cases a,simp_all, case_tac "nat", simp_all) - moreover - {assume "bound0 (NEq a)" hence bn:"bound0 (simpfm (NEq a))" - using simpfm_bound0 by blast - have "isatom (simpfm (NEq a))" by (cases "simpnum a", auto simp add: Let_def) - with bn bound0at_l have ?case by blast} - moreover - {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" - { - assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" - with numgcd_pos[where t="CN 0 c (simpnum e)"] - have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" - by (simp add: numgcd_def zgcd_le1) - from prems have th': "c\0" by auto - from prems have cp: "c \ 0" by simp - from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] - have "0 < c div numgcd (CN 0 c (simpnum e))" by simp - } - with prems have ?case - by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} - ultimately show ?case by blast -next - case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i a))" - using simpfm_bound0 by blast - have "isatom (simpfm (Dvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def) - with bn bound0at_l show ?case by blast -next - case (NDvd i a) hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i a))" - using simpfm_bound0 by blast - have "isatom (simpfm (NDvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def) - with bn bound0at_l show ?case by blast -qed(auto simp add: conj_def imp_def disj_def iff_def Let_def simpfm_bound0 numadd_nb numneg_nb) - -lemma rlfm_I: - assumes qfp: "qfree p" - and xp: "0 \ x" and x1: "x < 1" - shows "(Ifm (x#bs) (rlfm p) = Ifm (x# bs) p) \ isrlfm (rlfm p)" - using qfp -by (induct p rule: rlfm.induct) -(auto simp add: rsplit[OF xp x1 lt_mono] lt_l rsplit[OF xp x1 le_mono] le_l rsplit[OF xp x1 gt_mono] gt_l - rsplit[OF xp x1 ge_mono] ge_l rsplit[OF xp x1 eq_mono] eq_l rsplit[OF xp x1 neq_mono] neq_l - rsplit[OF xp x1 DVD_mono[OF xp x1]] DVD_l rsplit[OF xp x1 NDVD_mono[OF xp x1]] NDVD_l simpfm_rl) -lemma rlfm_l: - assumes qfp: "qfree p" - shows "isrlfm (rlfm p)" - using qfp lt_l gt_l ge_l le_l eq_l neq_l DVD_l NDVD_l -by (induct p rule: rlfm.induct,auto simp add: simpfm_rl) - - (* Operations needed for Ferrante and Rackoff *) -lemma rminusinf_inf: - assumes lp: "isrlfm p" - shows "\ z. \ x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\ z. \ x. ?P z x p") -using lp -proof (induct p rule: minusinf.induct) - case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto -next - case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto -next - case (3 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - hence "real c * x + ?e \ 0" by simp - with xz have "?P ?z x (Eq (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp - thus ?case by blast -next - case (4 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - hence "real c * x + ?e \ 0" by simp - with xz have "?P ?z x (NEq (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp - thus ?case by blast -next - case (5 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - with xz have "?P ?z x (Lt (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp - thus ?case by blast -next - case (6 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - with xz have "?P ?z x (Le (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (Le (CN 0 c e))" by simp - thus ?case by blast -next - case (7 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - with xz have "?P ?z x (Gt (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp - thus ?case by blast -next - case (8 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - with xz have "?P ?z x (Ge (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp - thus ?case by blast -qed simp_all - -lemma rplusinf_inf: - assumes lp: "isrlfm p" - shows "\ z. \ x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\ z. \ x. ?P z x p") -using lp -proof (induct p rule: isrlfm.induct) - case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto -next - case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto -next - case (3 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - hence "real c * x + ?e \ 0" by simp - with xz have "?P ?z x (Eq (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp - thus ?case by blast -next - case (4 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - hence "real c * x + ?e \ 0" by simp - with xz have "?P ?z x (NEq (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp - thus ?case by blast -next - case (5 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - with xz have "?P ?z x (Lt (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp - thus ?case by blast -next - case (6 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - with xz have "?P ?z x (Le (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (Le (CN 0 c e))" by simp - thus ?case by blast -next - case (7 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - with xz have "?P ?z x (Gt (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp - thus ?case by blast -next - case (8 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - with xz have "?P ?z x (Ge (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp - thus ?case by blast -qed simp_all - -lemma rminusinf_bound0: - assumes lp: "isrlfm p" - shows "bound0 (minusinf p)" - using lp - by (induct p rule: minusinf.induct) simp_all - -lemma rplusinf_bound0: - assumes lp: "isrlfm p" - shows "bound0 (plusinf p)" - using lp - by (induct p rule: plusinf.induct) simp_all - -lemma rminusinf_ex: - assumes lp: "isrlfm p" - and ex: "Ifm (a#bs) (minusinf p)" - shows "\ x. Ifm (x#bs) p" -proof- - from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex - have th: "\ x. Ifm (x#bs) (minusinf p)" by auto - from rminusinf_inf[OF lp, where bs="bs"] - obtain z where z_def: "\x x. Ifm (x#bs) p" -proof- - from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex - have th: "\ x. Ifm (x#bs) (plusinf p)" by auto - from rplusinf_inf[OF lp, where bs="bs"] - obtain z where z_def: "\x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast - from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp - moreover have "z + 1 > z" by simp - ultimately show ?thesis using z_def by auto -qed - -consts - \:: "fm \ (num \ int) list" - \ :: "fm \ (num \ int) \ fm " -recdef \ "measure size" - "\ (And p q) = (\ p @ \ q)" - "\ (Or p q) = (\ p @ \ q)" - "\ (Eq (CN 0 c e)) = [(Neg e,c)]" - "\ (NEq (CN 0 c e)) = [(Neg e,c)]" - "\ (Lt (CN 0 c e)) = [(Neg e,c)]" - "\ (Le (CN 0 c e)) = [(Neg e,c)]" - "\ (Gt (CN 0 c e)) = [(Neg e,c)]" - "\ (Ge (CN 0 c e)) = [(Neg e,c)]" - "\ p = []" - -recdef \ "measure size" - "\ (And p q) = (\ (t,n). And (\ p (t,n)) (\ q (t,n)))" - "\ (Or p q) = (\ (t,n). Or (\ p (t,n)) (\ q (t,n)))" - "\ (Eq (CN 0 c e)) = (\ (t,n). Eq (Add (Mul c t) (Mul n e)))" - "\ (NEq (CN 0 c e)) = (\ (t,n). NEq (Add (Mul c t) (Mul n e)))" - "\ (Lt (CN 0 c e)) = (\ (t,n). Lt (Add (Mul c t) (Mul n e)))" - "\ (Le (CN 0 c e)) = (\ (t,n). Le (Add (Mul c t) (Mul n e)))" - "\ (Gt (CN 0 c e)) = (\ (t,n). Gt (Add (Mul c t) (Mul n e)))" - "\ (Ge (CN 0 c e)) = (\ (t,n). Ge (Add (Mul c t) (Mul n e)))" - "\ p = (\ (t,n). p)" - -lemma \_I: assumes lp: "isrlfm p" - and np: "real n > 0" and nbt: "numbound0 t" - shows "(Ifm (x#bs) (\ p (t,n)) = Ifm (((Inum (x#bs) t)/(real n))#bs) p) \ bound0 (\ p (t,n))" (is "(?I x (\ p (t,n)) = ?I ?u p) \ ?B p" is "(_ = ?I (?t/?n) p) \ _" is "(_ = ?I (?N x t /_) p) \ _") - using lp -proof(induct p rule: \.induct) - case (5 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - have "?I ?u (Lt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) < 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)" - by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) < 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -next - case (6 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) \ 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -next - case (7 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)" - by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) > 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -next - case (8 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) \ 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -next - case (3 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - from np have np: "real n \ 0" by simp - have "?I ?u (Eq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) = 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)" - by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) = 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -next - case (4 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - from np have np: "real n \ 0" by simp - have "?I ?u (NEq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) \ 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"] nth_pos2) - -lemma \_l: - assumes lp: "isrlfm p" - shows "\ (t,k) \ set (\ p). numbound0 t \ k >0" -using lp -by(induct p rule: \.induct) auto - -lemma rminusinf_\: - assumes lp: "isrlfm p" - and nmi: "\ (Ifm (a#bs) (minusinf p))" (is "\ (Ifm (a#bs) (?M p))") - and ex: "Ifm (x#bs) p" (is "?I x p") - shows "\ (s,m) \ set (\ p). x \ Inum (a#bs) s / real m" (is "\ (s,m) \ ?U p. x \ ?N a s / real m") -proof- - have "\ (s,m) \ set (\ p). real m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real m *x \ ?N a s") - using lp nmi ex - by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) - then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real m * x \ ?N a s" by blast - from \_l[OF lp] smU have mp: "real m > 0" by auto - from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" - by (auto simp add: mult_commute) - thus ?thesis using smU by auto -qed - -lemma rplusinf_\: - assumes lp: "isrlfm p" - and nmi: "\ (Ifm (a#bs) (plusinf p))" (is "\ (Ifm (a#bs) (?M p))") - and ex: "Ifm (x#bs) p" (is "?I x p") - shows "\ (s,m) \ set (\ p). x \ Inum (a#bs) s / real m" (is "\ (s,m) \ ?U p. x \ ?N a s / real m") -proof- - have "\ (s,m) \ set (\ p). real m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real m *x \ ?N a s") - using lp nmi ex - by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) - then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real m * x \ ?N a s" by blast - from \_l[OF lp] smU have mp: "real m > 0" by auto - from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" - by (auto simp add: mult_commute) - thus ?thesis using smU by auto -qed - -lemma lin_dense: - assumes lp: "isrlfm p" - and noS: "\ t. l < t \ t< u \ t \ (\ (t,n). Inum (x#bs) t / real n) ` set (\ p)" - (is "\ t. _ \ _ \ t \ (\ (t,n). ?N x t / real n ) ` (?U p)") - and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p" - and ly: "l < y" and yu: "y < u" - shows "Ifm (y#bs) p" -using lp px noS -proof (induct p rule: isrlfm.induct) - case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ - from prems have "x * real c + ?N x e < 0" by (simp add: ring_simps) - hence pxc: "x < (- ?N x e) / real c" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"]) - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y < (-?N x e)/ real c" - hence "y * real c < - ?N x e" - by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e < 0" by (simp add: ring_simps) - hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y > (- ?N x e) / real c" - with yu have eu: "u > (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ l" by (cases "(- ?N x e) / real c > l", auto) - with lx pxc have "False" by auto - hence ?case by simp } - ultimately show ?case by blast -next - case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp + - from prems have "x * real c + ?N x e \ 0" by (simp add: ring_simps) - hence pxc: "x \ (- ?N x e) / real c" - by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"]) - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y < (-?N x e)/ real c" - hence "y * real c < - ?N x e" - by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e < 0" by (simp add: ring_simps) - hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y > (- ?N x e) / real c" - with yu have eu: "u > (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ l" by (cases "(- ?N x e) / real c > l", auto) - with lx pxc have "False" by auto - hence ?case by simp } - ultimately show ?case by blast -next - case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ - from prems have "x * real c + ?N x e > 0" by (simp add: ring_simps) - hence pxc: "x > (- ?N x e) / real c" - by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"]) - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y > (-?N x e)/ real c" - hence "y * real c > - ?N x e" - by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e > 0" by (simp add: ring_simps) - hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y < (- ?N x e) / real c" - with ly have eu: "l < (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ u" by (cases "(- ?N x e) / real c > l", auto) - with xu pxc have "False" by auto - hence ?case by simp } - ultimately show ?case by blast -next - case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ - from prems have "x * real c + ?N x e \ 0" by (simp add: ring_simps) - hence pxc: "x \ (- ?N x e) / real c" - by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"]) - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y > (-?N x e)/ real c" - hence "y * real c > - ?N x e" - by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e > 0" by (simp add: ring_simps) - hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y < (- ?N x e) / real c" - with ly have eu: "l < (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ u" by (cases "(- ?N x e) / real c > l", auto) - with xu pxc have "False" by auto - hence ?case by simp } - ultimately show ?case by blast -next - case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ - from cp have cnz: "real c \ 0" by simp - from prems have "x * real c + ?N x e = 0" by (simp add: ring_simps) - hence pxc: "x = (- ?N x e) / real c" - by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"]) - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with lx xu have yne: "x \ - ?N x e / real c" by auto - with pxc show ?case by simp -next - case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ - from cp have cnz: "real c \ 0" by simp - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y* real c \ -?N x e" - by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp - hence "y* real c + ?N x e \ 0" by (simp add: ring_simps) - thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] - by (simp add: ring_simps) -qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="y" and b'="x"]) - -lemma finite_set_intervals: - assumes px: "P (x::real)" - and lx: "l \ x" and xu: "x \ u" - and linS: "l\ S" and uinS: "u \ S" - and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" - shows "\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a \ x \ x \ b \ P x" -proof- - let ?Mx = "{y. y\ S \ y \ x}" - let ?xM = "{y. y\ S \ x \ y}" - let ?a = "Max ?Mx" - let ?b = "Min ?xM" - have MxS: "?Mx \ S" by blast - hence fMx: "finite ?Mx" using fS finite_subset by auto - from lx linS have linMx: "l \ ?Mx" by blast - hence Mxne: "?Mx \ {}" by blast - have xMS: "?xM \ S" by blast - hence fxM: "finite ?xM" using fS finite_subset by auto - from xu uinS have linxM: "u \ ?xM" by blast - hence xMne: "?xM \ {}" by blast - have ax:"?a \ x" using Mxne fMx by auto - have xb:"x \ ?b" using xMne fxM by auto - have "?a \ ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \ S" using MxS by blast - have "?b \ ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \ S" using xMS by blast - have noy:"\ y. ?a < y \ y < ?b \ y \ S" - proof(clarsimp) - fix y - assume ay: "?a < y" and yb: "y < ?b" and yS: "y \ S" - from yS have "y\ ?Mx \ y\ ?xM" by auto - moreover {assume "y \ ?Mx" hence "y \ ?a" using Mxne fMx by auto with ay have "False" by simp} - moreover {assume "y \ ?xM" hence "y \ ?b" using xMne fxM by auto with yb have "False" by simp} - ultimately show "False" by blast - qed - from ainS binS noy ax xb px show ?thesis by blast -qed - -lemma finite_set_intervals2: - assumes px: "P (x::real)" - and lx: "l \ x" and xu: "x \ u" - and linS: "l\ S" and uinS: "u \ S" - and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" - shows "(\ s\ S. P s) \ (\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a < x \ x < b \ P x)" -proof- - from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su] - obtain a and b where - as: "a\ S" and bs: "b\ S" and noS:"\y. a < y \ y < b \ y \ S" and axb: "a \ x \ x \ b \ P x" by auto - from axb have "x= a \ x= b \ (a < x \ x < b)" by auto - thus ?thesis using px as bs noS by blast -qed - -lemma rinf_\: - assumes lp: "isrlfm p" - and nmi: "\ (Ifm (x#bs) (minusinf p))" (is "\ (Ifm (x#bs) (?M p))") - and npi: "\ (Ifm (x#bs) (plusinf p))" (is "\ (Ifm (x#bs) (?P p))") - and ex: "\ x. Ifm (x#bs) p" (is "\ x. ?I x p") - shows "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p" -proof- - let ?N = "\ x t. Inum (x#bs) t" - let ?U = "set (\ p)" - from ex obtain a where pa: "?I a p" by blast - from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi - have nmi': "\ (?I a (?M p))" by simp - from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi - have npi': "\ (?I a (?P p))" by simp - have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). ?I ((?N a l/real n + ?N a s /real m) / 2) p" - proof- - let ?M = "(\ (t,c). ?N a t / real c) ` ?U" - have fM: "finite ?M" by auto - from rminusinf_\[OF lp nmi pa] rplusinf_\[OF lp npi pa] - have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). a \ ?N x l / real n \ a \ ?N x s / real m" by blast - then obtain "t" "n" "s" "m" where - tnU: "(t,n) \ ?U" and smU: "(s,m) \ ?U" - and xs1: "a \ ?N x s / real m" and tx1: "a \ ?N x t / real n" by blast - from \_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 have xs: "a \ ?N a s / real m" and tx: "a \ ?N a t / real n" by auto - from tnU have Mne: "?M \ {}" by auto - hence Une: "?U \ {}" by simp - let ?l = "Min ?M" - let ?u = "Max ?M" - have linM: "?l \ ?M" using fM Mne by simp - have uinM: "?u \ ?M" using fM Mne by simp - have tnM: "?N a t / real n \ ?M" using tnU by auto - have smM: "?N a s / real m \ ?M" using smU by auto - have lM: "\ t\ ?M. ?l \ t" using Mne fM by auto - have Mu: "\ t\ ?M. t \ ?u" using Mne fM by auto - have "?l \ ?N a t / real n" using tnM Mne by simp hence lx: "?l \ a" using tx by simp - have "?N a s / real m \ ?u" using smM Mne by simp hence xu: "a \ ?u" using xs by simp - from finite_set_intervals2[where P="\ x. ?I x p",OF pa lx xu linM uinM fM lM Mu] - have "(\ s\ ?M. ?I s p) \ - (\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p)" . - moreover { fix u assume um: "u\ ?M" and pu: "?I u p" - hence "\ (tu,nu) \ ?U. u = ?N a tu / real nu" by auto - then obtain "tu" "nu" where tuU: "(tu,nu) \ ?U" and tuu:"u= ?N a tu / real nu" by blast - have "(u + u) / 2 = u" by auto with pu tuu - have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" by simp - with tuU have ?thesis by blast} - moreover{ - assume "\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p" - then obtain t1 and t2 where t1M: "t1 \ ?M" and t2M: "t2\ ?M" - and noM: "\ y. t1 < y \ y < t2 \ y \ ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p" - by blast - from t1M have "\ (t1u,t1n) \ ?U. t1 = ?N a t1u / real t1n" by auto - then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \ ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast - from t2M have "\ (t2u,t2n) \ ?U. t2 = ?N a t2u / real t2n" by auto - then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \ ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast - from t1x xt2 have t1t2: "t1 < t2" by simp - let ?u = "(t1 + t2) / 2" - from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto - from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" . - with t1uU t2uU t1u t2u have ?thesis by blast} - ultimately show ?thesis by blast - qed - then obtain "l" "n" "s" "m" where lnU: "(l,n) \ ?U" and smU:"(s,m) \ ?U" - and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast - from lnU smU \_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto - from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] - numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu - have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp - with lnU smU - show ?thesis by auto -qed - (* The Ferrante - Rackoff Theorem *) - -theorem fr_eq: - assumes lp: "isrlfm p" - shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,n) \ set (\ p). \ (s,m) \ set (\ p). Ifm ((((Inum (x#bs) t)/ real n + (Inum (x#bs) s) / real m) /2)#bs) p))" - (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") -proof - assume px: "\ x. ?I x p" - have "?M \ ?P \ (\ ?M \ \ ?P)" by blast - moreover {assume "?M \ ?P" hence "?D" by blast} - moreover {assume nmi: "\ ?M" and npi: "\ ?P" - from rinf_\[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast} - ultimately show "?D" by blast -next - assume "?D" - moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} - moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } - moreover {assume f:"?F" hence "?E" by blast} - ultimately show "?E" by blast -qed - - -lemma fr_eq\: - assumes lp: "isrlfm p" - shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,k) \ set (\ p). \ (s,l) \ set (\ p). Ifm (x#bs) (\ p (Add(Mul l t) (Mul k s) , 2*k*l))))" - (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") -proof - assume px: "\ x. ?I x p" - have "?M \ ?P \ (\ ?M \ \ ?P)" by blast - moreover {assume "?M \ ?P" hence "?D" by blast} - moreover {assume nmi: "\ ?M" and npi: "\ ?P" - let ?f ="\ (t,n). Inum (x#bs) t / real n" - let ?N = "\ t. Inum (x#bs) t" - {fix t n s m assume "(t,n)\ set (\ p)" and "(s,m) \ set (\ p)" - with \_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" - by auto - let ?st = "Add (Mul m t) (Mul n s)" - from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute) - from tnb snb have st_nb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" - using mnp mp np by (simp add: ring_simps add_divide_distrib) - from \_I[OF lp mnp st_nb, where x="x" and bs="bs"] - have "?I x (\ p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])} - with rinf_\[OF lp nmi npi px] have "?F" by blast hence "?D" by blast} - ultimately show "?D" by blast -next - assume "?D" - moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} - moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } - moreover {fix t k s l assume "(t,k) \ set (\ p)" and "(s,l) \ set (\ p)" - and px:"?I x (\ p (Add (Mul l t) (Mul k s), 2*k*l))" - with \_l[OF lp] have tnb: "numbound0 t" and np:"real k > 0" and snb: "numbound0 s" and mp:"real l > 0" by auto - let ?st = "Add (Mul l t) (Mul k s)" - from mult_pos_pos[OF np mp] have mnp: "real (2*k*l) > 0" - by (simp add: mult_commute) - from tnb snb have st_nb: "numbound0 ?st" by simp - from \_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto} - ultimately show "?E" by blast -qed - -text{* The overall Part *} - -lemma real_ex_int_real01: - shows "(\ (x::real). P x) = (\ (i::int) (u::real). 0\ u \ u< 1 \ P (real i + u))" -proof(auto) - fix x - assume Px: "P x" - let ?i = "floor x" - let ?u = "x - real ?i" - have "x = real ?i + ?u" by simp - hence "P (real ?i + ?u)" using Px by simp - moreover have "real ?i \ x" using real_of_int_floor_le by simp hence "0 \ ?u" by arith - moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith - ultimately show "(\ (i::int) (u::real). 0\ u \ u< 1 \ P (real i + u))" by blast -qed - -consts exsplitnum :: "num \ num" - exsplit :: "fm \ fm" -recdef exsplitnum "measure size" - "exsplitnum (C c) = (C c)" - "exsplitnum (Bound 0) = Add (Bound 0) (Bound 1)" - "exsplitnum (Bound n) = Bound (n+1)" - "exsplitnum (Neg a) = Neg (exsplitnum a)" - "exsplitnum (Add a b) = Add (exsplitnum a) (exsplitnum b) " - "exsplitnum (Sub a b) = Sub (exsplitnum a) (exsplitnum b) " - "exsplitnum (Mul c a) = Mul c (exsplitnum a)" - "exsplitnum (Floor a) = Floor (exsplitnum a)" - "exsplitnum (CN 0 c a) = CN 0 c (Add (Mul c (Bound 1)) (exsplitnum a))" - "exsplitnum (CN n c a) = CN (n+1) c (exsplitnum a)" - "exsplitnum (CF c s t) = CF c (exsplitnum s) (exsplitnum t)" - -recdef exsplit "measure size" - "exsplit (Lt a) = Lt (exsplitnum a)" - "exsplit (Le a) = Le (exsplitnum a)" - "exsplit (Gt a) = Gt (exsplitnum a)" - "exsplit (Ge a) = Ge (exsplitnum a)" - "exsplit (Eq a) = Eq (exsplitnum a)" - "exsplit (NEq a) = NEq (exsplitnum a)" - "exsplit (Dvd i a) = Dvd i (exsplitnum a)" - "exsplit (NDvd i a) = NDvd i (exsplitnum a)" - "exsplit (And p q) = And (exsplit p) (exsplit q)" - "exsplit (Or p q) = Or (exsplit p) (exsplit q)" - "exsplit (Imp p q) = Imp (exsplit p) (exsplit q)" - "exsplit (Iff p q) = Iff (exsplit p) (exsplit q)" - "exsplit (NOT p) = NOT (exsplit p)" - "exsplit p = p" - -lemma exsplitnum: - "Inum (x#y#bs) (exsplitnum t) = Inum ((x+y) #bs) t" - by(induct t rule: exsplitnum.induct) (simp_all add: ring_simps) - -lemma exsplit: - assumes qfp: "qfree p" - shows "Ifm (x#y#bs) (exsplit p) = Ifm ((x+y)#bs) p" -using qfp exsplitnum[where x="x" and y="y" and bs="bs"] -by(induct p rule: exsplit.induct) simp_all - -lemma splitex: - assumes qf: "qfree p" - shows "(Ifm bs (E p)) = (\ (i::int). Ifm (real i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs") -proof- - have "?rhs = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm (x#(real i)#bs) (exsplit p))" - by (simp add: myless[rule_format, where b="1"] myless[rule_format, where b="0"] add_ac diff_def) - also have "\ = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm ((real i + x) #bs) p)" - by (simp only: exsplit[OF qf] add_ac) - also have "\ = (\ x. Ifm (x#bs) p)" - by (simp only: real_ex_int_real01[where P="\ x. Ifm (x#bs) p"]) - finally show ?thesis by simp -qed - - (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *) - -constdefs ferrack01:: "fm \ fm" - "ferrack01 p \ (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p); - U = remdups(map simp_num_pair - (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) - (alluopairs (\ p')))) - in decr (evaldjf (\ p') U ))" - -lemma fr_eq_01: - assumes qf: "qfree p" - shows "(\ x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (\ (t,n) \ set (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). \ (s,m) \ set (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). Ifm (x#bs) (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) (Add (Mul m t) (Mul n s), 2*n*m)))" - (is "(\ x. ?I x ?q) = ?F") -proof- - let ?rq = "rlfm ?q" - let ?M = "?I x (minusinf ?rq)" - let ?P = "?I x (plusinf ?rq)" - have MF: "?M = False" - apply (simp add: Let_def reducecoeff_def numgcd_def zgcd_def rsplit_def ge_def lt_def conj_def disj_def) - by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all) - have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def zgcd_def rsplit_def ge_def lt_def conj_def disj_def) - by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all) - have "(\ x. ?I x ?q ) = - ((?I x (minusinf ?rq)) \ (?I x (plusinf ?rq )) \ (\ (t,n) \ set (\ ?rq). \ (s,m) \ set (\ ?rq ). ?I x (\ ?rq (Add (Mul m t) (Mul n s), 2*n*m))))" - (is "(\ x. ?I x ?q) = (?M \ ?P \ ?F)" is "?E = ?D") - proof - assume "\ x. ?I x ?q" - then obtain x where qx: "?I x ?q" by blast - hence xp: "0\ x" and x1: "x< 1" and px: "?I x p" - by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf]) - from qx have "?I x ?rq " - by (simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1]) - hence lqx: "?I x ?rq " using simpfm[where p="?rq" and bs="x#bs"] by auto - from qf have qfq:"isrlfm ?rq" - by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1]) - with lqx fr_eq\[OF qfq] show "?M \ ?P \ ?F" by blast - next - assume D: "?D" - let ?U = "set (\ ?rq )" - from MF PF D have "?F" by auto - then obtain t n s m where aU:"(t,n) \ ?U" and bU:"(s,m)\ ?U" and rqx: "?I x (\ ?rq (Add (Mul m t) (Mul n s), 2*n*m))" by blast - from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf] - by (auto simp add: rsplit_def lt_def ge_def) - from aU bU \_l[OF lrq] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" by (auto simp add: split_def) - let ?st = "Add (Mul m t) (Mul n s)" - from tnb snb have stnb: "numbound0 ?st" by simp - from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute) - from conjunct1[OF \_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx - have "\ x. ?I x ?rq" by auto - thus "?E" - using rlfm_I[OF qf] by (auto simp add: rsplit_def lt_def ge_def) - qed - with MF PF show ?thesis by blast -qed - -lemma \_cong_aux: - assumes Ul: "\ (t,n) \ set U. numbound0 t \ n >0" - shows "((\ (t,n). Inum (x#bs) t /real n) ` (set (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs U)))) = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (set U \ set U))" - (is "?lhs = ?rhs") -proof(auto) - fix t n s m - assume "((t,n),(s,m)) \ set (alluopairs U)" - hence th: "((t,n),(s,m)) \ (set U \ set U)" - using alluopairs_set1[where xs="U"] by blast - let ?N = "\ t. Inum (x#bs) t" - let ?st= "Add (Mul m t) (Mul n s)" - from Ul th have mnz: "m \ 0" by auto - from Ul th have nnz: "n \ 0" by auto - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" - using mnz nnz by (simp add: ring_simps add_divide_distrib) - - thus "(real m * Inum (x # bs) t + real n * Inum (x # bs) s) / - (2 * real n * real m) - \ (\((t, n), s, m). - (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) ` - (set U \ set U)"using mnz nnz th - apply (auto simp add: th add_divide_distrib ring_simps split_def image_def) - by (rule_tac x="(s,m)" in bexI,simp_all) - (rule_tac x="(t,n)" in bexI,simp_all) -next - fix t n s m - assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" - let ?N = "\ t. Inum (x#bs) t" - let ?st= "Add (Mul m t) (Mul n s)" - from Ul smU have mnz: "m \ 0" by auto - from Ul tnU have nnz: "n \ 0" by auto - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" - using mnz nnz by (simp add: ring_simps add_divide_distrib) - let ?P = "\ (t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" - have Pc:"\ a b. ?P a b = ?P b a" - by auto - from Ul alluopairs_set1 have Up:"\ ((t,n),(s,m)) \ set (alluopairs U). n \ 0 \ m \ 0" by blast - from alluopairs_ex[OF Pc, where xs="U"] tnU smU - have th':"\ ((t',n'),(s',m')) \ set (alluopairs U). ?P (t',n') (s',m')" - by blast - then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \ set (alluopairs U)" - and Pts': "?P (t',n') (s',m')" by blast - from ts'_U Up have mnz': "m' \ 0" and nnz': "n'\ 0" by auto - let ?st' = "Add (Mul m' t') (Mul n' s')" - have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')" - using mnz' nnz' by (simp add: ring_simps add_divide_distrib) - from Pts' have - "(Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" by simp - also have "\ = ((\(t, n). Inum (x # bs) t / real n) ((\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st') - finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2 - \ (\(t, n). Inum (x # bs) t / real n) ` - (\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ` - set (alluopairs U)" - using ts'_U by blast -qed - -lemma \_cong: - assumes lp: "isrlfm p" - and UU': "((\ (t,n). Inum (x#bs) t /real n) ` U') = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (U \ U))" (is "?f ` U' = ?g ` (U\U)") - and U: "\ (t,n) \ U. numbound0 t \ n > 0" - and U': "\ (t,n) \ U'. numbound0 t \ n > 0" - shows "(\ (t,n) \ U. \ (s,m) \ U. Ifm (x#bs) (\ p (Add (Mul m t) (Mul n s),2*n*m))) = (\ (t,n) \ U'. Ifm (x#bs) (\ p (t,n)))" - (is "?lhs = ?rhs") -proof - assume ?lhs - then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and - Pst: "Ifm (x#bs) (\ p (Add (Mul m t) (Mul n s),2*n*m))" by blast - let ?N = "\ t. Inum (x#bs) t" - from tnU smU U have tnb: "numbound0 t" and np: "n > 0" - and snb: "numbound0 s" and mp:"m > 0" by auto - let ?st= "Add (Mul m t) (Mul n s)" - from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) - from tnb snb have stnb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" - using mp np by (simp add: ring_simps add_divide_distrib) - from tnU smU UU' have "?g ((t,n),(s,m)) \ ?f ` U'" by blast - hence "\ (t',n') \ U'. ?g ((t,n),(s,m)) = ?f (t',n')" - by auto (rule_tac x="(a,b)" in bexI, auto) - then obtain t' n' where tnU': "(t',n') \ U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast - from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto - from \_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst - have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp - from conjunct1[OF \_I[OF lp np' tnb', where bs="bs" and x="x"], symmetric] th[simplified split_def fst_conv snd_conv,symmetric] Pst2[simplified st[symmetric]] - have "Ifm (x # bs) (\ p (t', n')) " by (simp only: st) - then show ?rhs using tnU' by auto -next - assume ?rhs - then obtain t' n' where tnU': "(t',n') \ U'" and Pt': "Ifm (x # bs) (\ p (t', n'))" - by blast - from tnU' UU' have "?f (t',n') \ ?g ` (U\U)" by blast - hence "\ ((t,n),(s,m)) \ (U\U). ?f (t',n') = ?g ((t,n),(s,m))" - by auto (rule_tac x="(a,b)" in bexI, auto) - then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and - th: "?f (t',n') = ?g((t,n),(s,m)) "by blast - let ?N = "\ t. Inum (x#bs) t" - from tnU smU U have tnb: "numbound0 t" and np: "n > 0" - and snb: "numbound0 s" and mp:"m > 0" by auto - let ?st= "Add (Mul m t) (Mul n s)" - from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) - from tnb snb have stnb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" - using mp np by (simp add: ring_simps add_divide_distrib) - from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto - from \_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt' - have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp - with \_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast -qed - -lemma ferrack01: - assumes qf: "qfree p" - shows "((\ x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (Ifm bs (ferrack01 p))) \ qfree (ferrack01 p)" (is "(?lhs = ?rhs) \ _") -proof- - let ?I = "\ x p. Ifm (x#bs) p" - fix x - let ?N = "\ t. Inum (x#bs) t" - let ?q = "rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)" - let ?U = "\ ?q" - let ?Up = "alluopairs ?U" - let ?g = "\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)" - let ?S = "map ?g ?Up" - let ?SS = "map simp_num_pair ?S" - let ?Y = "remdups ?SS" - let ?f= "(\ (t,n). ?N t / real n)" - let ?h = "\ ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2" - let ?F = "\ p. \ a \ set (\ p). \ b \ set (\ p). ?I x (\ p (?g(a,b)))" - let ?ep = "evaldjf (\ ?q) ?Y" - from rlfm_l[OF qf] have lq: "isrlfm ?q" - by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def zgcd_def) - from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \ (set ?U \ set ?U)" by simp - from \_l[OF lq] have U_l: "\ (t,n) \ set ?U. numbound0 t \ n > 0" . - from U_l UpU - have Up_: "\ ((t,n),(s,m)) \ set ?Up. numbound0 t \ n> 0 \ numbound0 s \ m > 0" by auto - hence Snb: "\ (t,n) \ set ?S. numbound0 t \ n > 0 " - by (auto simp add: mult_pos_pos) - have Y_l: "\ (t,n) \ set ?Y. numbound0 t \ n > 0" - proof- - { fix t n assume tnY: "(t,n) \ set ?Y" - hence "(t,n) \ set ?SS" by simp - hence "\ (t',n') \ set ?S. simp_num_pair (t',n') = (t,n)" - by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all) - then obtain t' n' where tn'S: "(t',n') \ set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast - from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto - from simp_num_pair_l[OF tnb np tns] - have "numbound0 t \ n > 0" . } - thus ?thesis by blast - qed - - have YU: "(?f ` set ?Y) = (?h ` (set ?U \ set ?U))" - proof- - from simp_num_pair_ci[where bs="x#bs"] have - "\x. (?f o simp_num_pair) x = ?f x" by auto - hence th: "?f o simp_num_pair = ?f" using ext by blast - have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_compose) - also have "\ = (?f ` set ?S)" by (simp add: th) - also have "\ = ((?f o ?g) ` set ?Up)" - by (simp only: set_map o_def image_compose[symmetric]) - also have "\ = (?h ` (set ?U \ set ?U))" - using \_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_compose[symmetric]] by blast - finally show ?thesis . - qed - have "\ (t,n) \ set ?Y. bound0 (\ ?q (t,n))" - proof- - { fix t n assume tnY: "(t,n) \ set ?Y" - with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto - from \_I[OF lq np tnb] - have "bound0 (\ ?q (t,n))" by simp} - thus ?thesis by blast - qed - hence ep_nb: "bound0 ?ep" using evaldjf_bound0[where xs="?Y" and f="\ ?q"] - by auto - - from fr_eq_01[OF qf, where bs="bs" and x="x"] have "?lhs = ?F ?q" - by (simp only: split_def fst_conv snd_conv) - also have "\ = (\ (t,n) \ set ?Y. ?I x (\ ?q (t,n)))" using \_cong[OF lq YU U_l Y_l] - by (simp only: split_def fst_conv snd_conv) - also have "\ = (Ifm (x#bs) ?ep)" - using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="\ ?q",symmetric] - by (simp only: split_def pair_collapse) - also have "\ = (Ifm bs (decr ?ep))" using decr[OF ep_nb] by blast - finally have lr: "?lhs = ?rhs" by (simp only: ferrack01_def Let_def) - from decr_qf[OF ep_nb] have "qfree (ferrack01 p)" by (simp only: Let_def ferrack01_def) - with lr show ?thesis by blast -qed - -lemma cp_thm': - assumes lp: "iszlfm p (real (i::int)#bs)" - and up: "d\ p 1" and dd: "d\ p d" and dp: "d > 0" - shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. d}. Ifm (real j#bs) (minusinf p)) \ (\ j\ {1.. d}. \ b\ (Inum (real i#bs)) ` set (\ p). Ifm ((b+real j)#bs) p))" - using cp_thm[OF lp up dd dp] by auto - -constdefs unit:: "fm \ fm \ num list \ int" - "unit p \ (let p' = zlfm p ; l = \ p' ; q = And (Dvd l (CN 0 1 (C 0))) (a\ p' l); d = \ q; - B = remdups (map simpnum (\ q)) ; a = remdups (map simpnum (\ q)) - in if length B \ length a then (q,B,d) else (mirror q, a,d))" - -lemma unit: assumes qf: "qfree p" - shows "\ q B d. unit p = (q,B,d) \ ((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ (Inum (real i#bs)) ` set B = (Inum (real i#bs)) ` set (\ q) \ d\ q 1 \ d\ q d \ d >0 \ iszlfm q (real (i::int)#bs) \ (\ b\ set B. numbound0 b)" -proof- - fix q B d - assume qBd: "unit p = (q,B,d)" - let ?thes = "((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ - Inum (real i#bs) ` set B = Inum (real i#bs) ` set (\ q) \ - d\ q 1 \ d\ q d \ 0 < d \ iszlfm q (real i # bs) \ (\ b\ set B. numbound0 b)" - let ?I = "\ (x::int) p. Ifm (real x#bs) p" - let ?p' = "zlfm p" - let ?l = "\ ?p'" - let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a\ ?p' ?l)" - let ?d = "\ ?q" - let ?B = "set (\ ?q)" - let ?B'= "remdups (map simpnum (\ ?q))" - let ?A = "set (\ ?q)" - let ?A'= "remdups (map simpnum (\ ?q))" - from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] - have pp': "\ i. ?I i ?p' = ?I i p" by auto - from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]] - have lp': "\ (i::int). iszlfm ?p' (real i#bs)" by simp - hence lp'': "iszlfm ?p' (real (i::int)#bs)" by simp - from lp' \[where p="?p'" and bs="bs"] have lp: "?l >0" and dl: "d\ ?p' ?l" by auto - from a\_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp' - have pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by (simp add: int_rdvd_iff) - from lp'' lp a\[OF lp'' dl lp] have lq:"iszlfm ?q (real i#bs)" and uq: "d\ ?q 1" - by (auto simp add: isint_def) - from \[OF lq] have dp:"?d >0" and dd: "d\ ?q ?d" by blast+ - let ?N = "\ t. Inum (real (i::int)#bs) t" - have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_compose) - also have "\ = ?N ` ?B" using simpnum_ci[where bs="real i #bs"] by auto - finally have BB': "?N ` set ?B' = ?N ` ?B" . - have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_compose) - also have "\ = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] by auto - finally have AA': "?N ` set ?A' = ?N ` ?A" . - from \_numbound0[OF lq] have B_nb:"\ b\ set ?B'. numbound0 b" - by (simp add: simpnum_numbound0) - from \_l[OF lq] have A_nb: "\ b\ set ?A'. numbound0 b" - by (simp add: simpnum_numbound0) - {assume "length ?B' \ length ?A'" - hence q:"q=?q" and "B = ?B'" and d:"d = ?d" - using qBd by (auto simp add: Let_def unit_def) - with BB' B_nb have b: "?N ` (set B) = ?N ` set (\ q)" - and bn: "\b\ set B. numbound0 b" by simp+ - with pq_ex dp uq dd lq q d have ?thes by simp} - moreover - {assume "\ (length ?B' \ length ?A')" - hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" - using qBd by (auto simp add: Let_def unit_def) - with AA' mirror\\[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\ q)" - and bn: "\b\ set B. numbound0 b" by simp+ - from mirror_ex[OF lq] pq_ex q - have pqm_eq:"(\ (x::int). ?I x p) = (\ (x::int). ?I x q)" by simp - from lq uq q mirror_d\ [where p="?q" and bs="bs" and a="real i"] - have lq': "iszlfm q (real i#bs)" and uq: "d\ q 1" by auto - from \[OF lq'] mirror_\[OF lq] q d have dq:"d\ q d " by auto - from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp - } - ultimately show ?thes by blast -qed - (* Cooper's Algorithm *) - -constdefs cooper :: "fm \ fm" - "cooper p \ - (let (q,B,d) = unit p; js = iupt (1,d); - mq = simpfm (minusinf q); - md = evaldjf (\ j. simpfm (subst0 (C j) mq)) js - in if md = T then T else - (let qd = evaldjf (\ t. simpfm (subst0 t q)) - (remdups (map (\ (b,j). simpnum (Add b (C j))) - [(b,j). b\B,j\js])) - in decr (disj md qd)))" -lemma cooper: assumes qf: "qfree p" - shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (cooper p))) \ qfree (cooper p)" - (is "(?lhs = ?rhs) \ _") -proof- - - let ?I = "\ (x::int) p. Ifm (real x#bs) p" - let ?q = "fst (unit p)" - let ?B = "fst (snd(unit p))" - let ?d = "snd (snd (unit p))" - let ?js = "iupt (1,?d)" - let ?mq = "minusinf ?q" - let ?smq = "simpfm ?mq" - let ?md = "evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js" - fix i - let ?N = "\ t. Inum (real (i::int)#bs) t" - let ?bjs = "[(b,j). b\?B,j\?js]" - let ?sbjs = "map (\ (b,j). simpnum (Add b (C j))) ?bjs" - let ?qd = "evaldjf (\ t. simpfm (subst0 t ?q)) (remdups ?sbjs)" - have qbf:"unit p = (?q,?B,?d)" by simp - from unit[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and - B:"?N ` set ?B = ?N ` set (\ ?q)" and - uq:"d\ ?q 1" and dd: "d\ ?q ?d" and dp: "?d > 0" and - lq: "iszlfm ?q (real i#bs)" and - Bn: "\ b\ set ?B. numbound0 b" by auto - from zlin_qfree[OF lq] have qfq: "qfree ?q" . - from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". - have jsnb: "\ j \ set ?js. numbound0 (C j)" by simp - hence "\ j\ set ?js. bound0 (subst0 (C j) ?smq)" - by (auto simp only: subst0_bound0[OF qfmq]) - hence th: "\ j\ set ?js. bound0 (simpfm (subst0 (C j) ?smq))" - by (auto simp add: simpfm_bound0) - from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp - from Bn jsnb have "\ (b,j) \ set ?bjs. numbound0 (Add b (C j))" - by simp - hence "\ (b,j) \ set ?bjs. numbound0 (simpnum (Add b (C j)))" - using simpnum_numbound0 by blast - hence "\ t \ set ?sbjs. numbound0 t" by simp - hence "\ t \ set (remdups ?sbjs). bound0 (subst0 t ?q)" - using subst0_bound0[OF qfq] by auto - hence th': "\ t \ set (remdups ?sbjs). bound0 (simpfm (subst0 t ?q))" - using simpfm_bound0 by blast - from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp - from mdb qdb - have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \ ?qd=T", simp_all) - from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B - have "?lhs = (\ j\ {1.. ?d}. ?I j ?mq \ (\ b\ ?N ` set ?B. Ifm ((b+ real j)#bs) ?q))" by auto - also have "\ = ((\ j\ set ?js. ?I j ?smq) \ (\ (b,j) \ (?N ` set ?B \ set ?js). Ifm ((b+ real j)#bs) ?q))" apply (simp only: iupt_set simpfm) by auto - also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ (\ (b,j). ?N (Add b (C j))) ` set ?bjs. Ifm (t #bs) ?q))" by simp - also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ (\ (b,j). ?N (simpnum (Add b (C j)))) ` set ?bjs. Ifm (t #bs) ?q))" by (simp only: simpnum_ci) - also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ set ?sbjs. Ifm (?N t #bs) ?q))" - by (auto simp add: split_def) - also have "\ = ((\ j\ set ?js. (\ j. ?I i (simpfm (subst0 (C j) ?smq))) j) \ (\ t \ set (remdups ?sbjs). (\ t. ?I i (simpfm (subst0 t ?q))) t))" by (simp only: simpfm subst0_I[OF qfq] simpfm Inum.simps subst0_I[OF qfmq] set_remdups) - also have "\ = ((?I i (evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js)) \ (?I i (evaldjf (\ t. simpfm (subst0 t ?q)) (remdups ?sbjs))))" by (simp only: evaldjf_ex) - finally have mdqd: "?lhs = (?I i (disj ?md ?qd))" by (simp add: disj) - hence mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" using decr [OF mdqdb] by simp - {assume mdT: "?md = T" - hence cT:"cooper p = T" - by (simp only: cooper_def unit_def split_def Let_def if_True) simp - from mdT mdqd have lhs:"?lhs" by (auto simp add: disj) - from mdT have "?rhs" by (simp add: cooper_def unit_def split_def) - with lhs cT have ?thesis by simp } - moreover - {assume mdT: "?md \ T" hence "cooper p = decr (disj ?md ?qd)" - by (simp only: cooper_def unit_def split_def Let_def if_False) - with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } - ultimately show ?thesis by blast -qed - -lemma DJcooper: - assumes qf: "qfree p" - shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ cooper p))) \ qfree (DJ cooper p)" -proof- - from cooper have cqf: "\ p. qfree p \ qfree (cooper p)" by blast - from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast - have "Ifm bs (DJ cooper p) = (\ q\ set (disjuncts p). Ifm bs (cooper q))" - by (simp add: DJ_def evaldjf_ex) - also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real x#bs) q)" - using cooper disjuncts_qf[OF qf] by blast - also have "\ = (\ (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto) - finally show ?thesis using thqf by blast -qed - - (* Redy and Loveland *) - -lemma \\_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'" - shows "Ifm (a#bs) (\\ p (t,c)) = Ifm (a#bs) (\\ p (t',c))" - using lp - by (induct p rule: iszlfm.induct, auto simp add: tt') - -lemma \_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'" - shows "Ifm (a#bs) (\ p c t) = Ifm (a#bs) (\ p c t')" - by (simp add: \_def tt' \\_cong[OF lp tt']) - -lemma \_cong: assumes lp: "iszlfm p (a#bs)" - and RR: "(\(b,k). (Inum (a#bs) b,k)) ` R = (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" - shows "(\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))) = (\ (e,c) \ set (\ p). \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j))))" - (is "?lhs = ?rhs") -proof - let ?d = "\ p" - assume ?lhs then obtain e c j where ecR: "(e,c) \ R" and jD:"j \ {1 .. c*?d}" - and px: "Ifm (a#bs) (\ p c (Add e (C j)))" (is "?sp c e j") by blast - from ecR have "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` R" by auto - hence "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" using RR by simp - hence "\ (e',c') \ set (\ p). Inum (a#bs) e = Inum (a#bs) e' \ c = c'" by auto - then obtain e' c' where ecRo:"(e',c') \ set (\ p)" and ee':"Inum (a#bs) e = Inum (a#bs) e'" - and cc':"c = c'" by blast - from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp - - from \_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp - from ecRo jD px' cc' show ?rhs apply auto - by (rule_tac x="(e', c')" in bexI,simp_all) - (rule_tac x="j" in bexI, simp_all add: cc'[symmetric]) -next - let ?d = "\ p" - assume ?rhs then obtain e c j where ecR: "(e,c) \ set (\ p)" and jD:"j \ {1 .. c*?d}" - and px: "Ifm (a#bs) (\ p c (Add e (C j)))" (is "?sp c e j") by blast - from ecR have "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" by auto - hence "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` R" using RR by simp - hence "\ (e',c') \ R. Inum (a#bs) e = Inum (a#bs) e' \ c = c'" by auto - then obtain e' c' where ecRo:"(e',c') \ R" and ee':"Inum (a#bs) e = Inum (a#bs) e'" - and cc':"c = c'" by blast - from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp - from \_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp - from ecRo jD px' cc' show ?lhs apply auto - by (rule_tac x="(e', c')" in bexI,simp_all) - (rule_tac x="j" in bexI, simp_all add: cc'[symmetric]) -qed - -lemma rl_thm': - assumes lp: "iszlfm p (real (i::int)#bs)" - and R: "(\(b,k). (Inum (a#bs) b,k)) ` R = (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" - shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real j#bs) (minusinf p)) \ (\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" - using rl_thm[OF lp] \_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp - -constdefs chooset:: "fm \ fm \ ((num\int) list) \ int" - "chooset p \ (let q = zlfm p ; d = \ q; - B = remdups (map (\ (t,k). (simpnum t,k)) (\ q)) ; - a = remdups (map (\ (t,k). (simpnum t,k)) (\\ q)) - in if length B \ length a then (q,B,d) else (mirror q, a,d))" - -lemma chooset: assumes qf: "qfree p" - shows "\ q B d. chooset p = (q,B,d) \ ((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ ((\(t,k). (Inum (real i#bs) t,k)) ` set B = (\(t,k). (Inum (real i#bs) t,k)) ` set (\ q)) \ (\ q = d) \ d >0 \ iszlfm q (real (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" -proof- - fix q B d - assume qBd: "chooset p = (q,B,d)" - let ?thes = "((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ ((\(t,k). (Inum (real i#bs) t,k)) ` set B = (\(t,k). (Inum (real i#bs) t,k)) ` set (\ q)) \ (\ q = d) \ d >0 \ iszlfm q (real (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" - let ?I = "\ (x::int) p. Ifm (real x#bs) p" - let ?q = "zlfm p" - let ?d = "\ ?q" - let ?B = "set (\ ?q)" - let ?f = "\ (t,k). (simpnum t,k)" - let ?B'= "remdups (map ?f (\ ?q))" - let ?A = "set (\\ ?q)" - let ?A'= "remdups (map ?f (\\ ?q))" - from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] - have pp': "\ i. ?I i ?q = ?I i p" by auto - hence pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by simp - from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real i"] - have lq: "iszlfm ?q (real (i::int)#bs)" . - from \[OF lq] have dp:"?d >0" by blast - let ?N = "\ (t,c). (Inum (real (i::int)#bs) t,c)" - have "?N ` set ?B' = ((?N o ?f) ` ?B)" by (simp add: split_def image_compose) - also have "\ = ?N ` ?B" - by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def) - finally have BB': "?N ` set ?B' = ?N ` ?B" . - have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_compose) - also have "\ = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] - by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def) - finally have AA': "?N ` set ?A' = ?N ` ?A" . - from \_l[OF lq] have B_nb:"\ (e,c)\ set ?B'. numbound0 e \ c > 0" - by (simp add: simpnum_numbound0 split_def) - from \\_l[OF lq] have A_nb: "\ (e,c)\ set ?A'. numbound0 e \ c > 0" - by (simp add: simpnum_numbound0 split_def) - {assume "length ?B' \ length ?A'" - hence q:"q=?q" and "B = ?B'" and d:"d = ?d" - using qBd by (auto simp add: Let_def chooset_def) - with BB' B_nb have b: "?N ` (set B) = ?N ` set (\ q)" - and bn: "\(e,c)\ set B. numbound0 e \ c > 0" by auto - with pq_ex dp lq q d have ?thes by simp} - moreover - {assume "\ (length ?B' \ length ?A')" - hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" - using qBd by (auto simp add: Let_def chooset_def) - with AA' mirror_\\[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\ q)" - and bn: "\(e,c)\ set B. numbound0 e \ c > 0" by auto - from mirror_ex[OF lq] pq_ex q - have pqm_eq:"(\ (x::int). ?I x p) = (\ (x::int). ?I x q)" by simp - from lq q mirror_l [where p="?q" and bs="bs" and a="real i"] - have lq': "iszlfm q (real i#bs)" by auto - from mirror_\[OF lq] pqm_eq b bn lq' dp q dp d have ?thes by simp - } - ultimately show ?thes by blast -qed - -constdefs stage:: "fm \ int \ (num \ int) \ fm" - "stage p d \ (\ (e,c). evaldjf (\ j. simpfm (\ p c (Add e (C j)))) (iupt (1,c*d)))" -lemma stage: - shows "Ifm bs (stage p d (e,c)) = (\ j\{1 .. c*d}. Ifm bs (\ p c (Add e (C j))))" - by (unfold stage_def split_def ,simp only: evaldjf_ex iupt_set simpfm) simp - -lemma stage_nb: assumes lp: "iszlfm p (a#bs)" and cp: "c >0" and nb:"numbound0 e" - shows "bound0 (stage p d (e,c))" -proof- - let ?f = "\ j. simpfm (\ p c (Add e (C j)))" - have th: "\ j\ set (iupt(1,c*d)). bound0 (?f j)" - proof - fix j - from nb have nb':"numbound0 (Add e (C j))" by simp - from simpfm_bound0[OF \_nb[OF lp nb', where k="c"]] - show "bound0 (simpfm (\ p c (Add e (C j))))" . - qed - from evaldjf_bound0[OF th] show ?thesis by (unfold stage_def split_def) simp -qed - -constdefs redlove:: "fm \ fm" - "redlove p \ - (let (q,B,d) = chooset p; - mq = simpfm (minusinf q); - md = evaldjf (\ j. simpfm (subst0 (C j) mq)) (iupt (1,d)) - in if md = T then T else - (let qd = evaldjf (stage q d) B - in decr (disj md qd)))" - -lemma redlove: assumes qf: "qfree p" - shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (redlove p))) \ qfree (redlove p)" - (is "(?lhs = ?rhs) \ _") -proof- - - let ?I = "\ (x::int) p. Ifm (real x#bs) p" - let ?q = "fst (chooset p)" - let ?B = "fst (snd(chooset p))" - let ?d = "snd (snd (chooset p))" - let ?js = "iupt (1,?d)" - let ?mq = "minusinf ?q" - let ?smq = "simpfm ?mq" - let ?md = "evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js" - fix i - let ?N = "\ (t,k). (Inum (real (i::int)#bs) t,k)" - let ?qd = "evaldjf (stage ?q ?d) ?B" - have qbf:"chooset p = (?q,?B,?d)" by simp - from chooset[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and - B:"?N ` set ?B = ?N ` set (\ ?q)" and dd: "\ ?q = ?d" and dp: "?d > 0" and - lq: "iszlfm ?q (real i#bs)" and - Bn: "\ (e,c)\ set ?B. numbound0 e \ c > 0" by auto - from zlin_qfree[OF lq] have qfq: "qfree ?q" . - from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". - have jsnb: "\ j \ set ?js. numbound0 (C j)" by simp - hence "\ j\ set ?js. bound0 (subst0 (C j) ?smq)" - by (auto simp only: subst0_bound0[OF qfmq]) - hence th: "\ j\ set ?js. bound0 (simpfm (subst0 (C j) ?smq))" - by (auto simp add: simpfm_bound0) - from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp - from Bn stage_nb[OF lq] have th:"\ x \ set ?B. bound0 (stage ?q ?d x)" by auto - from evaldjf_bound0[OF th] have qdb: "bound0 ?qd" . - from mdb qdb - have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \ ?qd=T", simp_all) - from trans [OF pq_ex rl_thm'[OF lq B]] dd - have "?lhs = ((\ j\ {1.. ?d}. ?I j ?mq) \ (\ (e,c)\ set ?B. \ j\ {1 .. c*?d}. Ifm (real i#bs) (\ ?q c (Add e (C j)))))" by auto - also have "\ = ((\ j\ {1.. ?d}. ?I j ?smq) \ (\ (e,c)\ set ?B. ?I i (stage ?q ?d (e,c) )))" - by (simp add: simpfm stage split_def) - also have "\ = ((\ j\ {1 .. ?d}. ?I i (subst0 (C j) ?smq)) \ ?I i ?qd)" - by (simp add: evaldjf_ex subst0_I[OF qfmq]) - finally have mdqd:"?lhs = (?I i ?md \ ?I i ?qd)" by (simp only: evaldjf_ex iupt_set simpfm) - also have "\ = (?I i (disj ?md ?qd))" by (simp add: disj) - also have "\ = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) - finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" . - {assume mdT: "?md = T" - hence cT:"redlove p = T" by (simp add: redlove_def Let_def chooset_def split_def) - from mdT have lhs:"?lhs" using mdqd by simp - from mdT have "?rhs" by (simp add: redlove_def chooset_def split_def) - with lhs cT have ?thesis by simp } - moreover - {assume mdT: "?md \ T" hence "redlove p = decr (disj ?md ?qd)" - by (simp add: redlove_def chooset_def split_def Let_def) - with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } - ultimately show ?thesis by blast -qed - -lemma DJredlove: - assumes qf: "qfree p" - shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ redlove p))) \ qfree (DJ redlove p)" -proof- - from redlove have cqf: "\ p. qfree p \ qfree (redlove p)" by blast - from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast - have "Ifm bs (DJ redlove p) = (\ q\ set (disjuncts p). Ifm bs (redlove q))" - by (simp add: DJ_def evaldjf_ex) - also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real x#bs) q)" - using redlove disjuncts_qf[OF qf] by blast - also have "\ = (\ (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto) - finally show ?thesis using thqf by blast -qed - - -lemma exsplit_qf: assumes qf: "qfree p" - shows "qfree (exsplit p)" -using qf by (induct p rule: exsplit.induct, auto) - -definition mircfr :: "fm \ fm" where - "mircfr = DJ cooper o ferrack01 o simpfm o exsplit" - -definition mirlfr :: "fm \ fm" where - "mirlfr = DJ redlove o ferrack01 o simpfm o exsplit" - -lemma mircfr: "\ bs p. qfree p \ qfree (mircfr p) \ Ifm bs (mircfr p) = Ifm bs (E p)" -proof(clarsimp simp del: Ifm.simps) - fix bs p - assume qf: "qfree p" - show "qfree (mircfr p)\(Ifm bs (mircfr p) = Ifm bs (E p))" (is "_ \ (?lhs = ?rhs)") - proof- - let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" - have "?rhs = (\ (i::int). \ x. Ifm (x#real i#bs) ?es)" - using splitex[OF qf] by simp - with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ - with DJcooper[OF qf'] show ?thesis by (simp add: mircfr_def) - qed -qed - -lemma mirlfr: "\ bs p. qfree p \ qfree(mirlfr p) \ Ifm bs (mirlfr p) = Ifm bs (E p)" -proof(clarsimp simp del: Ifm.simps) - fix bs p - assume qf: "qfree p" - show "qfree (mirlfr p)\(Ifm bs (mirlfr p) = Ifm bs (E p))" (is "_ \ (?lhs = ?rhs)") - proof- - let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" - have "?rhs = (\ (i::int). \ x. Ifm (x#real i#bs) ?es)" - using splitex[OF qf] by simp - with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ - with DJredlove[OF qf'] show ?thesis by (simp add: mirlfr_def) - qed -qed - -definition mircfrqe:: "fm \ fm" where - "mircfrqe p = qelim (prep p) mircfr" - -definition mirlfrqe:: "fm \ fm" where - "mirlfrqe p = qelim (prep p) mirlfr" - -theorem mircfrqe: "(Ifm bs (mircfrqe p) = Ifm bs p) \ qfree (mircfrqe p)" - using qelim_ci[OF mircfr] prep by (auto simp add: mircfrqe_def) - -theorem mirlfrqe: "(Ifm bs (mirlfrqe p) = Ifm bs p) \ qfree (mirlfrqe p)" - using qelim_ci[OF mirlfr] prep by (auto simp add: mirlfrqe_def) - -definition - "test1 (u\unit) = mircfrqe (A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0)))))))" - -definition - "test2 (u\unit) = mircfrqe (A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0)))))" - -definition - "test3 (u\unit) = mirlfrqe (A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0)))))))" - -definition - "test4 (u\unit) = mirlfrqe (A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0)))))" - -definition - "test5 (u\unit) = mircfrqe (A(E(And (Ge(Sub (Bound 1) (Bound 0))) (Eq (Add (Floor (Bound 1)) (Floor (Neg(Bound 0))))))))" - -ML {* @{code test1} () *} -ML {* @{code test2} () *} -ML {* @{code test3} () *} -ML {* @{code test4} () *} -ML {* @{code test5} () *} - -(*export_code mircfrqe mirlfrqe - in SML module_name Mir file "raw_mir.ML"*) - -oracle mirfr_oracle = {* fn (proofs, ct) => -let - -fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t - of NONE => error "Variable not found in the list!" - | SOME n => @{code Bound} n) - | num_of_term vs @{term "real (0::int)"} = @{code C} 0 - | num_of_term vs @{term "real (1::int)"} = @{code C} 1 - | num_of_term vs @{term "0::real"} = @{code C} 0 - | num_of_term vs @{term "1::real"} = @{code C} 1 - | num_of_term vs (Bound i) = @{code Bound} i - | num_of_term vs (@{term "uminus :: real \ real"} $ t') = @{code Neg} (num_of_term vs t') - | num_of_term vs (@{term "op + :: real \ real \ real"} $ t1 $ t2) = - @{code Add} (num_of_term vs t1, num_of_term vs t2) - | num_of_term vs (@{term "op - :: real \ real \ real"} $ t1 $ t2) = - @{code Sub} (num_of_term vs t1, num_of_term vs t2) - | num_of_term vs (@{term "op * :: real \ real \ real"} $ t1 $ t2) = - (case (num_of_term vs t1) - of @{code C} i => @{code Mul} (i, num_of_term vs t2) - | _ => error "num_of_term: unsupported Multiplication") - | num_of_term vs (@{term "real :: int \ real"} $ (@{term "number_of :: int \ int"} $ t')) = - @{code C} (HOLogic.dest_numeral t') - | num_of_term vs (@{term "real :: int \ real"} $ (@{term "floor :: real \ int"} $ t')) = - @{code Floor} (num_of_term vs t') - | num_of_term vs (@{term "real :: int \ real"} $ (@{term "ceiling :: real \ int"} $ t')) = - @{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t'))) - | num_of_term vs (@{term "number_of :: int \ real"} $ t') = - @{code C} (HOLogic.dest_numeral t') - | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t); - -fun fm_of_term vs @{term True} = @{code T} - | fm_of_term vs @{term False} = @{code F} - | fm_of_term vs (@{term "op < :: real \ real \ bool"} $ t1 $ t2) = - @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) - | fm_of_term vs (@{term "op \ :: real \ real \ bool"} $ t1 $ t2) = - @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) - | fm_of_term vs (@{term "op = :: real \ real \ bool"} $ t1 $ t2) = - @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) - | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \ real"} $ (@{term "number_of :: int \ int"} $ t1)) $ t2) = - @{code Dvd} (HOLogic.dest_numeral t1, num_of_term vs t2) - | fm_of_term vs (@{term "op = :: bool \ bool \ bool"} $ t1 $ t2) = - @{code Iff} (fm_of_term vs t1, fm_of_term vs t2) - | fm_of_term vs (@{term "op &"} $ t1 $ t2) = - @{code And} (fm_of_term vs t1, fm_of_term vs t2) - | fm_of_term vs (@{term "op |"} $ t1 $ t2) = - @{code Or} (fm_of_term vs t1, fm_of_term vs t2) - | fm_of_term vs (@{term "op -->"} $ t1 $ t2) = - @{code Imp} (fm_of_term vs t1, fm_of_term vs t2) - | fm_of_term vs (@{term "Not"} $ t') = - @{code NOT} (fm_of_term vs t') - | fm_of_term vs (Const ("Ex", _) $ Abs (xn, xT, p)) = - @{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) - | fm_of_term vs (Const ("All", _) $ Abs (xn, xT, p)) = - @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) - | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t); - -fun term_of_num vs (@{code C} i) = @{term "real :: int \ real"} $ HOLogic.mk_number HOLogic.intT i - | term_of_num vs (@{code Bound} n) = fst (the (find_first (fn (_, m) => n = m) vs)) - | term_of_num vs (@{code Neg} (@{code Floor} (@{code Neg} t'))) = - @{term "real :: int \ real"} $ (@{term "ceiling :: real \ int"} $ term_of_num vs t') - | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \ real"} $ term_of_num vs t' - | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: real \ real \ real"} $ - term_of_num vs t1 $ term_of_num vs t2 - | term_of_num vs (@{code Sub} (t1, t2)) = @{term "op - :: real \ real \ real"} $ - term_of_num vs t1 $ term_of_num vs t2 - | term_of_num vs (@{code Mul} (i, t2)) = @{term "op * :: real \ real \ real"} $ - term_of_num vs (@{code C} i) $ term_of_num vs t2 - | term_of_num vs (@{code Floor} t) = @{term "real :: int \ real"} $ (@{term "floor :: real \ int"} $ term_of_num vs t) - | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t)) - | term_of_num vs (@{code CF} (c, t, s)) = term_of_num vs (@{code Add} (@{code Mul} (c, @{code Floor} t), s)); - -fun term_of_fm vs @{code T} = HOLogic.true_const - | term_of_fm vs @{code F} = HOLogic.false_const - | term_of_fm vs (@{code Lt} t) = - @{term "op < :: real \ real \ bool"} $ term_of_num vs t $ @{term "0::real"} - | term_of_fm vs (@{code Le} t) = - @{term "op \ :: real \ real \ bool"} $ term_of_num vs t $ @{term "0::real"} - | term_of_fm vs (@{code Gt} t) = - @{term "op < :: real \ real \ bool"} $ @{term "0::real"} $ term_of_num vs t - | term_of_fm vs (@{code Ge} t) = - @{term "op \ :: real \ real \ bool"} $ @{term "0::real"} $ term_of_num vs t - | term_of_fm vs (@{code Eq} t) = - @{term "op = :: real \ real \ bool"} $ term_of_num vs t $ @{term "0::real"} - | term_of_fm vs (@{code NEq} t) = - term_of_fm vs (@{code NOT} (@{code Eq} t)) - | term_of_fm vs (@{code Dvd} (i, t)) = - @{term "op rdvd"} $ term_of_num vs (@{code C} i) $ term_of_num vs t - | term_of_fm vs (@{code NDvd} (i, t)) = - term_of_fm vs (@{code NOT} (@{code Dvd} (i, t))) - | term_of_fm vs (@{code NOT} t') = - HOLogic.Not $ term_of_fm vs t' - | term_of_fm vs (@{code And} (t1, t2)) = - HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2 - | term_of_fm vs (@{code Or} (t1, t2)) = - HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2 - | term_of_fm vs (@{code Imp} (t1, t2)) = - HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2 - | term_of_fm vs (@{code Iff} (t1, t2)) = - @{term "op = :: bool \ bool \ bool"} $ term_of_fm vs t1 $ term_of_fm vs t2; - -in - let - val thy = Thm.theory_of_cterm ct; - val t = Thm.term_of ct; - val fs = term_frees t; - val vs = fs ~~ (0 upto (length fs - 1)); - val qe = if proofs then @{code mirlfrqe} else @{code mircfrqe}; - val t' = (term_of_fm vs o qe o fm_of_term vs) t; - in (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end -end; -*} - -use "mirtac.ML" -setup "MirTac.setup" - -lemma "ALL (x::real). (\x\ = \x\ = (x = real \x\))" -apply mir -done - -lemma "ALL (x::real). real (2::int)*x - (real (1::int)) < real \x\ + real \x\ \ real \x\ + real \x\ \ real (2::int)*x + (real (1::int))" -apply mir -done - -lemma "ALL (x::real). 2*\x\ \ \2*x\ \ \2*x\ \ 2*\x+1\" -apply mir -done - -lemma "ALL (x::real). \y \ x. (\x\ = \y\)" -apply mir -done - -lemma "ALL x y. \x\ = \y\ \ 0 \ abs (y - x) \ abs (y - x) \ 1" -apply mir -done - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/ex/ReflectedFerrack.thy --- a/src/HOL/Complex/ex/ReflectedFerrack.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2101 +0,0 @@ -(* Title: Complex/ex/ReflectedFerrack.thy - Author: Amine Chaieb -*) - -theory ReflectedFerrack -imports Main GCD Real Efficient_Nat -uses ("linrtac.ML") -begin - -section {* Quantifier elimination for @{text "\ (0, 1, +, <)"} *} - - (*********************************************************************************) - (* SOME GENERAL STUFF< HAS TO BE MOVED IN SOME LIB *) - (*********************************************************************************) - -consts alluopairs:: "'a list \ ('a \ 'a) list" -primrec - "alluopairs [] = []" - "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)" - -lemma alluopairs_set1: "set (alluopairs xs) \ {(x,y). x\ set xs \ y\ set xs}" -by (induct xs, auto) - -lemma alluopairs_set: - "\x\ set xs ; y \ set xs\ \ (x,y) \ set (alluopairs xs) \ (y,x) \ set (alluopairs xs) " -by (induct xs, auto) - -lemma alluopairs_ex: - assumes Pc: "\ x y. P x y = P y x" - shows "(\ x \ set xs. \ y \ set xs. P x y) = (\ (x,y) \ set (alluopairs xs). P x y)" -proof - assume "\x\set xs. \y\set xs. P x y" - then obtain x y where x: "x \ set xs" and y:"y \ set xs" and P: "P x y" by blast - from alluopairs_set[OF x y] P Pc show"\(x, y)\set (alluopairs xs). P x y" - by auto -next - assume "\(x, y)\set (alluopairs xs). P x y" - then obtain "x" and "y" where xy:"(x,y) \ set (alluopairs xs)" and P: "P x y" by blast+ - from xy have "x \ set xs \ y\ set xs" using alluopairs_set1 by blast - with P show "\x\set xs. \y\set xs. P x y" by blast -qed - -lemma nth_pos2: "0 < n \ (x#xs) ! n = xs ! (n - 1)" -using Nat.gr0_conv_Suc -by clarsimp - -lemma filter_length: "length (List.filter P xs) < Suc (length xs)" - apply (induct xs, auto) done - -consts remdps:: "'a list \ 'a list" - -recdef remdps "measure size" - "remdps [] = []" - "remdps (x#xs) = (x#(remdps (List.filter (\ y. y \ x) xs)))" -(hints simp add: filter_length[rule_format]) - -lemma remdps_set[simp]: "set (remdps xs) = set xs" - by (induct xs rule: remdps.induct, auto) - - - - (*********************************************************************************) - (**** SHADOW SYNTAX AND SEMANTICS ****) - (*********************************************************************************) - -datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num - | Mul int num - - (* A size for num to make inductive proofs simpler*) -consts num_size :: "num \ nat" -primrec - "num_size (C c) = 1" - "num_size (Bound n) = 1" - "num_size (Neg a) = 1 + num_size a" - "num_size (Add a b) = 1 + num_size a + num_size b" - "num_size (Sub a b) = 3 + num_size a + num_size b" - "num_size (Mul c a) = 1 + num_size a" - "num_size (CN n c a) = 3 + num_size a " - - (* Semantics of numeral terms (num) *) -consts Inum :: "real list \ num \ real" -primrec - "Inum bs (C c) = (real c)" - "Inum bs (Bound n) = bs!n" - "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)" - "Inum bs (Neg a) = -(Inum bs a)" - "Inum bs (Add a b) = Inum bs a + Inum bs b" - "Inum bs (Sub a b) = Inum bs a - Inum bs b" - "Inum bs (Mul c a) = (real c) * Inum bs a" - (* FORMULAE *) -datatype fm = - T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| - NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm - - - (* A size for fm *) -consts fmsize :: "fm \ nat" -recdef fmsize "measure size" - "fmsize (NOT p) = 1 + fmsize p" - "fmsize (And p q) = 1 + fmsize p + fmsize q" - "fmsize (Or p q) = 1 + fmsize p + fmsize q" - "fmsize (Imp p q) = 3 + fmsize p + fmsize q" - "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)" - "fmsize (E p) = 1 + fmsize p" - "fmsize (A p) = 4+ fmsize p" - "fmsize p = 1" - (* several lemmas about fmsize *) -lemma fmsize_pos: "fmsize p > 0" -by (induct p rule: fmsize.induct) simp_all - - (* Semantics of formulae (fm) *) -consts Ifm ::"real list \ fm \ bool" -primrec - "Ifm bs T = True" - "Ifm bs F = False" - "Ifm bs (Lt a) = (Inum bs a < 0)" - "Ifm bs (Gt a) = (Inum bs a > 0)" - "Ifm bs (Le a) = (Inum bs a \ 0)" - "Ifm bs (Ge a) = (Inum bs a \ 0)" - "Ifm bs (Eq a) = (Inum bs a = 0)" - "Ifm bs (NEq a) = (Inum bs a \ 0)" - "Ifm bs (NOT p) = (\ (Ifm bs p))" - "Ifm bs (And p q) = (Ifm bs p \ Ifm bs q)" - "Ifm bs (Or p q) = (Ifm bs p \ Ifm bs q)" - "Ifm bs (Imp p q) = ((Ifm bs p) \ (Ifm bs q))" - "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)" - "Ifm bs (E p) = (\ x. Ifm (x#bs) p)" - "Ifm bs (A p) = (\ x. Ifm (x#bs) p)" - -lemma IfmLeSub: "\ Inum bs s = s' ; Inum bs t = t' \ \ Ifm bs (Le (Sub s t)) = (s' \ t')" -apply simp -done - -lemma IfmLtSub: "\ Inum bs s = s' ; Inum bs t = t' \ \ Ifm bs (Lt (Sub s t)) = (s' < t')" -apply simp -done -lemma IfmEqSub: "\ Inum bs s = s' ; Inum bs t = t' \ \ Ifm bs (Eq (Sub s t)) = (s' = t')" -apply simp -done -lemma IfmNOT: " (Ifm bs p = P) \ (Ifm bs (NOT p) = (\P))" -apply simp -done -lemma IfmAnd: " \ Ifm bs p = P ; Ifm bs q = Q\ \ (Ifm bs (And p q) = (P \ Q))" -apply simp -done -lemma IfmOr: " \ Ifm bs p = P ; Ifm bs q = Q\ \ (Ifm bs (Or p q) = (P \ Q))" -apply simp -done -lemma IfmImp: " \ Ifm bs p = P ; Ifm bs q = Q\ \ (Ifm bs (Imp p q) = (P \ Q))" -apply simp -done -lemma IfmIff: " \ Ifm bs p = P ; Ifm bs q = Q\ \ (Ifm bs (Iff p q) = (P = Q))" -apply simp -done - -lemma IfmE: " (!! x. Ifm (x#bs) p = P x) \ (Ifm bs (E p) = (\x. P x))" -apply simp -done -lemma IfmA: " (!! x. Ifm (x#bs) p = P x) \ (Ifm bs (A p) = (\x. P x))" -apply simp -done - -consts not:: "fm \ fm" -recdef not "measure size" - "not (NOT p) = p" - "not T = F" - "not F = T" - "not p = NOT p" -lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)" -by (cases p) auto - -constdefs conj :: "fm \ fm \ fm" - "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else - if p = q then p else And p q)" -lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)" -by (cases "p=F \ q=F",simp_all add: conj_def) (cases p,simp_all) - -constdefs disj :: "fm \ fm \ fm" - "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p - else if p=q then p else Or p q)" - -lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)" -by (cases "p=T \ q=T",simp_all add: disj_def) (cases p,simp_all) - -constdefs imp :: "fm \ fm \ fm" - "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p - else Imp p q)" -lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)" -by (cases "p=F \ q=T",simp_all add: imp_def) - -constdefs iff :: "fm \ fm \ fm" - "iff p q \ (if (p = q) then T else if (p = NOT q \ NOT p = q) then F else - if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else - Iff p q)" -lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)" - by (unfold iff_def,cases "p=q", simp,cases "p=NOT q", simp) (cases "NOT p= q", auto) - -lemma conj_simps: - "conj F Q = F" - "conj P F = F" - "conj T Q = Q" - "conj P T = P" - "conj P P = P" - "P \ T \ P \ F \ Q \ T \ Q \ F \ P \ Q \ conj P Q = And P Q" - by (simp_all add: conj_def) - -lemma disj_simps: - "disj T Q = T" - "disj P T = T" - "disj F Q = Q" - "disj P F = P" - "disj P P = P" - "P \ T \ P \ F \ Q \ T \ Q \ F \ P \ Q \ disj P Q = Or P Q" - by (simp_all add: disj_def) -lemma imp_simps: - "imp F Q = T" - "imp P T = T" - "imp T Q = Q" - "imp P F = not P" - "imp P P = T" - "P \ T \ P \ F \ P \ Q \ Q \ T \ Q \ F \ imp P Q = Imp P Q" - by (simp_all add: imp_def) -lemma trivNOT: "p \ NOT p" "NOT p \ p" -apply (induct p, auto) -done - -lemma iff_simps: - "iff p p = T" - "iff p (NOT p) = F" - "iff (NOT p) p = F" - "iff p F = not p" - "iff F p = not p" - "p \ NOT T \ iff T p = p" - "p\ NOT T \ iff p T = p" - "p\q \ p\ NOT q \ q\ NOT p \ p\ F \ q\ F \ p \ T \ q \ T \ iff p q = Iff p q" - using trivNOT - by (simp_all add: iff_def, cases p, auto) - (* Quantifier freeness *) -consts qfree:: "fm \ bool" -recdef qfree "measure size" - "qfree (E p) = False" - "qfree (A p) = False" - "qfree (NOT p) = qfree p" - "qfree (And p q) = (qfree p \ qfree q)" - "qfree (Or p q) = (qfree p \ qfree q)" - "qfree (Imp p q) = (qfree p \ qfree q)" - "qfree (Iff p q) = (qfree p \ qfree q)" - "qfree p = True" - - (* Boundedness and substitution *) -consts - numbound0:: "num \ bool" (* a num is INDEPENDENT of Bound 0 *) - bound0:: "fm \ bool" (* A Formula is independent of Bound 0 *) -primrec - "numbound0 (C c) = True" - "numbound0 (Bound n) = (n>0)" - "numbound0 (CN n c a) = (n\0 \ numbound0 a)" - "numbound0 (Neg a) = numbound0 a" - "numbound0 (Add a b) = (numbound0 a \ numbound0 b)" - "numbound0 (Sub a b) = (numbound0 a \ numbound0 b)" - "numbound0 (Mul i a) = numbound0 a" -lemma numbound0_I: - assumes nb: "numbound0 a" - shows "Inum (b#bs) a = Inum (b'#bs) a" -using nb -by (induct a rule: numbound0.induct,auto simp add: nth_pos2) - -primrec - "bound0 T = True" - "bound0 F = True" - "bound0 (Lt a) = numbound0 a" - "bound0 (Le a) = numbound0 a" - "bound0 (Gt a) = numbound0 a" - "bound0 (Ge a) = numbound0 a" - "bound0 (Eq a) = numbound0 a" - "bound0 (NEq a) = numbound0 a" - "bound0 (NOT p) = bound0 p" - "bound0 (And p q) = (bound0 p \ bound0 q)" - "bound0 (Or p q) = (bound0 p \ bound0 q)" - "bound0 (Imp p q) = ((bound0 p) \ (bound0 q))" - "bound0 (Iff p q) = (bound0 p \ bound0 q)" - "bound0 (E p) = False" - "bound0 (A p) = False" - -lemma bound0_I: - assumes bp: "bound0 p" - shows "Ifm (b#bs) p = Ifm (b'#bs) p" -using bp numbound0_I[where b="b" and bs="bs" and b'="b'"] -by (induct p rule: bound0.induct) (auto simp add: nth_pos2) - -lemma not_qf[simp]: "qfree p \ qfree (not p)" -by (cases p, auto) -lemma not_bn[simp]: "bound0 p \ bound0 (not p)" -by (cases p, auto) - - -lemma conj_qf[simp]: "\qfree p ; qfree q\ \ qfree (conj p q)" -using conj_def by auto -lemma conj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (conj p q)" -using conj_def by auto - -lemma disj_qf[simp]: "\qfree p ; qfree q\ \ qfree (disj p q)" -using disj_def by auto -lemma disj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (disj p q)" -using disj_def by auto - -lemma imp_qf[simp]: "\qfree p ; qfree q\ \ qfree (imp p q)" -using imp_def by (cases "p=F \ q=T",simp_all add: imp_def) -lemma imp_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (imp p q)" -using imp_def by (cases "p=F \ q=T \ p=q",simp_all add: imp_def) - -lemma iff_qf[simp]: "\qfree p ; qfree q\ \ qfree (iff p q)" - by (unfold iff_def,cases "p=q", auto) -lemma iff_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (iff p q)" -using iff_def by (unfold iff_def,cases "p=q", auto) - -consts - decrnum:: "num \ num" - decr :: "fm \ fm" - -recdef decrnum "measure size" - "decrnum (Bound n) = Bound (n - 1)" - "decrnum (Neg a) = Neg (decrnum a)" - "decrnum (Add a b) = Add (decrnum a) (decrnum b)" - "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)" - "decrnum (Mul c a) = Mul c (decrnum a)" - "decrnum (CN n c a) = CN (n - 1) c (decrnum a)" - "decrnum a = a" - -recdef decr "measure size" - "decr (Lt a) = Lt (decrnum a)" - "decr (Le a) = Le (decrnum a)" - "decr (Gt a) = Gt (decrnum a)" - "decr (Ge a) = Ge (decrnum a)" - "decr (Eq a) = Eq (decrnum a)" - "decr (NEq a) = NEq (decrnum a)" - "decr (NOT p) = NOT (decr p)" - "decr (And p q) = conj (decr p) (decr q)" - "decr (Or p q) = disj (decr p) (decr q)" - "decr (Imp p q) = imp (decr p) (decr q)" - "decr (Iff p q) = iff (decr p) (decr q)" - "decr p = p" - -lemma decrnum: assumes nb: "numbound0 t" - shows "Inum (x#bs) t = Inum bs (decrnum t)" - using nb by (induct t rule: decrnum.induct, simp_all add: nth_pos2) - -lemma decr: assumes nb: "bound0 p" - shows "Ifm (x#bs) p = Ifm bs (decr p)" - using nb - by (induct p rule: decr.induct, simp_all add: nth_pos2 decrnum) - -lemma decr_qf: "bound0 p \ qfree (decr p)" -by (induct p, simp_all) - -consts - isatom :: "fm \ bool" (* test for atomicity *) -recdef isatom "measure size" - "isatom T = True" - "isatom F = True" - "isatom (Lt a) = True" - "isatom (Le a) = True" - "isatom (Gt a) = True" - "isatom (Ge a) = True" - "isatom (Eq a) = True" - "isatom (NEq a) = True" - "isatom p = False" - -lemma bound0_qf: "bound0 p \ qfree p" -by (induct p, simp_all) - -constdefs djf:: "('a \ fm) \ 'a \ fm \ fm" - "djf f p q \ (if q=T then T else if q=F then f p else - (let fp = f p in case fp of T \ T | F \ q | _ \ Or (f p) q))" -constdefs evaldjf:: "('a \ fm) \ 'a list \ fm" - "evaldjf f ps \ foldr (djf f) ps F" - -lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)" -by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) -(cases "f p", simp_all add: Let_def djf_def) - - -lemma djf_simps: - "djf f p T = T" - "djf f p F = f p" - "q\T \ q\F \ djf f p q = (let fp = f p in case fp of T \ T | F \ q | _ \ Or (f p) q)" - by (simp_all add: djf_def) - -lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\ p \ set ps. Ifm bs (f p))" - by(induct ps, simp_all add: evaldjf_def djf_Or) - -lemma evaldjf_bound0: - assumes nb: "\ x\ set xs. bound0 (f x)" - shows "bound0 (evaldjf f xs)" - using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) - -lemma evaldjf_qf: - assumes nb: "\ x\ set xs. qfree (f x)" - shows "qfree (evaldjf f xs)" - using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) - -consts disjuncts :: "fm \ fm list" -recdef disjuncts "measure size" - "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)" - "disjuncts F = []" - "disjuncts p = [p]" - -lemma disjuncts: "(\ q\ set (disjuncts p). Ifm bs q) = Ifm bs p" -by(induct p rule: disjuncts.induct, auto) - -lemma disjuncts_nb: "bound0 p \ \ q\ set (disjuncts p). bound0 q" -proof- - assume nb: "bound0 p" - hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto) - thus ?thesis by (simp only: list_all_iff) -qed - -lemma disjuncts_qf: "qfree p \ \ q\ set (disjuncts p). qfree q" -proof- - assume qf: "qfree p" - hence "list_all qfree (disjuncts p)" - by (induct p rule: disjuncts.induct, auto) - thus ?thesis by (simp only: list_all_iff) -qed - -constdefs DJ :: "(fm \ fm) \ fm \ fm" - "DJ f p \ evaldjf f (disjuncts p)" - -lemma DJ: assumes fdj: "\ p q. Ifm bs (f (Or p q)) = Ifm bs (Or (f p) (f q))" - and fF: "f F = F" - shows "Ifm bs (DJ f p) = Ifm bs (f p)" -proof- - have "Ifm bs (DJ f p) = (\ q \ set (disjuncts p). Ifm bs (f q))" - by (simp add: DJ_def evaldjf_ex) - also have "\ = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto) - finally show ?thesis . -qed - -lemma DJ_qf: assumes - fqf: "\ p. qfree p \ qfree (f p)" - shows "\p. qfree p \ qfree (DJ f p) " -proof(clarify) - fix p assume qf: "qfree p" - have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def) - from disjuncts_qf[OF qf] have "\ q\ set (disjuncts p). qfree q" . - with fqf have th':"\ q\ set (disjuncts p). qfree (f q)" by blast - - from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp -qed - -lemma DJ_qe: assumes qe: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" - shows "\ bs p. qfree p \ qfree (DJ qe p) \ (Ifm bs ((DJ qe p)) = Ifm bs (E p))" -proof(clarify) - fix p::fm and bs - assume qf: "qfree p" - from qe have qth: "\ p. qfree p \ qfree (qe p)" by blast - from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto - have "Ifm bs (DJ qe p) = (\ q\ set (disjuncts p). Ifm bs (qe q))" - by (simp add: DJ_def evaldjf_ex) - also have "\ = (\ q \ set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto - also have "\ = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto) - finally show "qfree (DJ qe p) \ Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast -qed - (* Simplification *) -consts - numgcd :: "num \ int" - numgcdh:: "num \ int \ int" - reducecoeffh:: "num \ int \ num" - reducecoeff :: "num \ num" - dvdnumcoeff:: "num \ int \ bool" -consts maxcoeff:: "num \ int" -recdef maxcoeff "measure size" - "maxcoeff (C i) = abs i" - "maxcoeff (CN n c t) = max (abs c) (maxcoeff t)" - "maxcoeff t = 1" - -lemma maxcoeff_pos: "maxcoeff t \ 0" - by (induct t rule: maxcoeff.induct, auto) - -recdef numgcdh "measure size" - "numgcdh (C i) = (\g. zgcd i g)" - "numgcdh (CN n c t) = (\g. zgcd c (numgcdh t g))" - "numgcdh t = (\g. 1)" -defs numgcd_def [code]: "numgcd t \ numgcdh t (maxcoeff t)" - -recdef reducecoeffh "measure size" - "reducecoeffh (C i) = (\ g. C (i div g))" - "reducecoeffh (CN n c t) = (\ g. CN n (c div g) (reducecoeffh t g))" - "reducecoeffh t = (\g. t)" - -defs reducecoeff_def: "reducecoeff t \ - (let g = numgcd t in - if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)" - -recdef dvdnumcoeff "measure size" - "dvdnumcoeff (C i) = (\ g. g dvd i)" - "dvdnumcoeff (CN n c t) = (\ g. g dvd c \ (dvdnumcoeff t g))" - "dvdnumcoeff t = (\g. False)" - -lemma dvdnumcoeff_trans: - assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'" - shows "dvdnumcoeff t g" - using dgt' gdg - by (induct t rule: dvdnumcoeff.induct, simp_all add: gdg zdvd_trans[OF gdg]) - -declare zdvd_trans [trans add] - -lemma natabs0: "(nat (abs x) = 0) = (x = 0)" -by arith - -lemma numgcd0: - assumes g0: "numgcd t = 0" - shows "Inum bs t = 0" - using g0[simplified numgcd_def] - by (induct t rule: numgcdh.induct, auto simp add: zgcd_def gcd_zero natabs0 max_def maxcoeff_pos) - -lemma numgcdh_pos: assumes gp: "g \ 0" shows "numgcdh t g \ 0" - using gp - by (induct t rule: numgcdh.induct, auto simp add: zgcd_def) - -lemma numgcd_pos: "numgcd t \0" - by (simp add: numgcd_def numgcdh_pos maxcoeff_pos) - -lemma reducecoeffh: - assumes gt: "dvdnumcoeff t g" and gp: "g > 0" - shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t" - using gt -proof(induct t rule: reducecoeffh.induct) - case (1 i) hence gd: "g dvd i" by simp - from gp have gnz: "g \ 0" by simp - from prems show ?case by (simp add: real_of_int_div[OF gnz gd]) -next - case (2 n c t) hence gd: "g dvd c" by simp - from gp have gnz: "g \ 0" by simp - from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps) -qed (auto simp add: numgcd_def gp) -consts ismaxcoeff:: "num \ int \ bool" -recdef ismaxcoeff "measure size" - "ismaxcoeff (C i) = (\ x. abs i \ x)" - "ismaxcoeff (CN n c t) = (\x. abs c \ x \ (ismaxcoeff t x))" - "ismaxcoeff t = (\x. True)" - -lemma ismaxcoeff_mono: "ismaxcoeff t c \ c \ c' \ ismaxcoeff t c'" -by (induct t rule: ismaxcoeff.induct, auto) - -lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)" -proof (induct t rule: maxcoeff.induct) - case (2 n c t) - hence H:"ismaxcoeff t (maxcoeff t)" . - have thh: "maxcoeff t \ max (abs c) (maxcoeff t)" by (simp add: le_maxI2) - from ismaxcoeff_mono[OF H thh] show ?case by (simp add: le_maxI1) -qed simp_all - -lemma zgcd_gt1: "zgcd i j > 1 \ ((abs i > 1 \ abs j > 1) \ (abs i = 0 \ abs j > 1) \ (abs i > 1 \ abs j = 0))" - apply (cases "abs i = 0", simp_all add: zgcd_def) - apply (cases "abs j = 0", simp_all) - apply (cases "abs i = 1", simp_all) - apply (cases "abs j = 1", simp_all) - apply auto - done -lemma numgcdh0:"numgcdh t m = 0 \ m =0" - by (induct t rule: numgcdh.induct, auto simp add:zgcd0) - -lemma dvdnumcoeff_aux: - assumes "ismaxcoeff t m" and mp:"m \ 0" and "numgcdh t m > 1" - shows "dvdnumcoeff t (numgcdh t m)" -using prems -proof(induct t rule: numgcdh.induct) - case (2 n c t) - let ?g = "numgcdh t m" - from prems have th:"zgcd c ?g > 1" by simp - from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] - have "(abs c > 1 \ ?g > 1) \ (abs c = 0 \ ?g > 1) \ (abs c > 1 \ ?g = 0)" by simp - moreover {assume "abs c > 1" and gp: "?g > 1" with prems - have th: "dvdnumcoeff t ?g" by simp - have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) - from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)} - moreover {assume "abs c = 0 \ ?g > 1" - with prems have th: "dvdnumcoeff t ?g" by simp - have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) - from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1) - hence ?case by simp } - moreover {assume "abs c > 1" and g0:"?g = 0" - from numgcdh0[OF g0] have "m=0". with prems have ?case by simp } - ultimately show ?case by blast -qed(auto simp add: zgcd_zdvd1) - -lemma dvdnumcoeff_aux2: - assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \ numgcd t > 0" - using prems -proof (simp add: numgcd_def) - let ?mc = "maxcoeff t" - let ?g = "numgcdh t ?mc" - have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff) - have th2: "?mc \ 0" by (rule maxcoeff_pos) - assume H: "numgcdh t ?mc > 1" - from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" . -qed - -lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" -proof- - let ?g = "numgcd t" - have "?g \ 0" by (simp add: numgcd_pos) - hence "?g = 0 \ ?g = 1 \ ?g > 1" by auto - moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} - moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} - moreover { assume g1:"?g > 1" - from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+ - from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis - by (simp add: reducecoeff_def Let_def)} - ultimately show ?thesis by blast -qed - -lemma reducecoeffh_numbound0: "numbound0 t \ numbound0 (reducecoeffh t g)" -by (induct t rule: reducecoeffh.induct, auto) - -lemma reducecoeff_numbound0: "numbound0 t \ numbound0 (reducecoeff t)" -using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def) - -consts - simpnum:: "num \ num" - numadd:: "num \ num \ num" - nummul:: "num \ int \ num" -recdef numadd "measure (\ (t,s). size t + size s)" - "numadd (CN n1 c1 r1,CN n2 c2 r2) = - (if n1=n2 then - (let c = c1 + c2 - in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2)))) - else if n1 \ n2 then (CN n1 c1 (numadd (r1,CN n2 c2 r2))) - else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))" - "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))" - "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" - "numadd (C b1, C b2) = C (b1+b2)" - "numadd (a,b) = Add a b" - -lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)" -apply (induct t s rule: numadd.induct, simp_all add: Let_def) -apply (case_tac "c1+c2 = 0",case_tac "n1 \ n2", simp_all) -apply (case_tac "n1 = n2", simp_all add: ring_simps) -by (simp only: left_distrib[symmetric],simp) - -lemma numadd_nb[simp]: "\ numbound0 t ; numbound0 s\ \ numbound0 (numadd (t,s))" -by (induct t s rule: numadd.induct, auto simp add: Let_def) - -recdef nummul "measure size" - "nummul (C j) = (\ i. C (i*j))" - "nummul (CN n c a) = (\ i. CN n (i*c) (nummul a i))" - "nummul t = (\ i. Mul i t)" - -lemma nummul[simp]: "\ i. Inum bs (nummul t i) = Inum bs (Mul i t)" -by (induct t rule: nummul.induct, auto simp add: ring_simps) - -lemma nummul_nb[simp]: "\ i. numbound0 t \ numbound0 (nummul t i)" -by (induct t rule: nummul.induct, auto ) - -constdefs numneg :: "num \ num" - "numneg t \ nummul t (- 1)" - -constdefs numsub :: "num \ num \ num" - "numsub s t \ (if s = t then C 0 else numadd (s,numneg t))" - -lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)" -using numneg_def by simp - -lemma numneg_nb[simp]: "numbound0 t \ numbound0 (numneg t)" -using numneg_def by simp - -lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)" -using numsub_def by simp - -lemma numsub_nb[simp]: "\ numbound0 t ; numbound0 s\ \ numbound0 (numsub t s)" -using numsub_def by simp - -recdef simpnum "measure size" - "simpnum (C j) = C j" - "simpnum (Bound n) = CN n 1 (C 0)" - "simpnum (Neg t) = numneg (simpnum t)" - "simpnum (Add t s) = numadd (simpnum t,simpnum s)" - "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)" - "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)" - "simpnum (CN n c t) = (if c = 0 then simpnum t else numadd (CN n c (C 0),simpnum t))" - -lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t" -by (induct t rule: simpnum.induct, auto simp add: numneg numadd numsub nummul) - -lemma simpnum_numbound0[simp]: - "numbound0 t \ numbound0 (simpnum t)" -by (induct t rule: simpnum.induct, auto) - -consts nozerocoeff:: "num \ bool" -recdef nozerocoeff "measure size" - "nozerocoeff (C c) = True" - "nozerocoeff (CN n c t) = (c\0 \ nozerocoeff t)" - "nozerocoeff t = True" - -lemma numadd_nz : "nozerocoeff a \ nozerocoeff b \ nozerocoeff (numadd (a,b))" -by (induct a b rule: numadd.induct,auto simp add: Let_def) - -lemma nummul_nz : "\ i. i\0 \ nozerocoeff a \ nozerocoeff (nummul a i)" -by (induct a rule: nummul.induct,auto simp add: Let_def numadd_nz) - -lemma numneg_nz : "nozerocoeff a \ nozerocoeff (numneg a)" -by (simp add: numneg_def nummul_nz) - -lemma numsub_nz: "nozerocoeff a \ nozerocoeff b \ nozerocoeff (numsub a b)" -by (simp add: numsub_def numneg_nz numadd_nz) - -lemma simpnum_nz: "nozerocoeff (simpnum t)" -by(induct t rule: simpnum.induct, auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz) - -lemma maxcoeff_nz: "nozerocoeff t \ maxcoeff t = 0 \ t = C 0" -proof (induct t rule: maxcoeff.induct) - case (2 n c t) - hence cnz: "c \0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+ - have "max (abs c) (maxcoeff t) \ abs c" by (simp add: le_maxI1) - with cnz have "max (abs c) (maxcoeff t) > 0" by arith - with prems show ?case by simp -qed auto - -lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0" -proof- - from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def) - from numgcdh0[OF th] have th:"maxcoeff t = 0" . - from maxcoeff_nz[OF nz th] show ?thesis . -qed - -constdefs simp_num_pair:: "(num \ int) \ num \ int" - "simp_num_pair \ (\ (t,n). (if n = 0 then (C 0, 0) else - (let t' = simpnum t ; g = numgcd t' in - if g > 1 then (let g' = zgcd n g in - if g' = 1 then (t',n) - else (reducecoeffh t' g', n div g')) - else (t',n))))" - -lemma simp_num_pair_ci: - shows "((\ (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((\ (t,n). Inum bs t / real n) (t,n))" - (is "?lhs = ?rhs") -proof- - let ?t' = "simpnum t" - let ?g = "numgcd ?t'" - let ?g' = "zgcd n ?g" - {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)} - moreover - { assume nnz: "n \ 0" - {assume "\ ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def simpnum_ci)} - moreover - {assume g1:"?g>1" hence g0: "?g > 0" by simp - from zgcd0 g1 nnz have gp0: "?g' \ 0" by simp - hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith - hence "?g'= 1 \ ?g' > 1" by arith - moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def simpnum_ci)} - moreover {assume g'1:"?g'>1" - from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" .. - let ?tt = "reducecoeffh ?t' ?g'" - let ?t = "Inum bs ?tt" - have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) - have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1) - have gpdgp: "?g' dvd ?g'" by simp - from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] - have th2:"real ?g' * ?t = Inum bs ?t'" by simp - from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def) - also have "\ = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp - also have "\ = (Inum bs ?t' / real n)" - using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp - finally have "?lhs = Inum bs t / real n" by (simp add: simpnum_ci) - then have ?thesis using prems by (simp add: simp_num_pair_def)} - ultimately have ?thesis by blast} - ultimately have ?thesis by blast} - ultimately show ?thesis by blast -qed - -lemma simp_num_pair_l: assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')" - shows "numbound0 t' \ n' >0" -proof- - let ?t' = "simpnum t" - let ?g = "numgcd ?t'" - let ?g' = "zgcd n ?g" - {assume nz: "n = 0" hence ?thesis using prems by (simp add: Let_def simp_num_pair_def)} - moreover - { assume nnz: "n \ 0" - {assume "\ ?g > 1" hence ?thesis using prems by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)} - moreover - {assume g1:"?g>1" hence g0: "?g > 0" by simp - from zgcd0 g1 nnz have gp0: "?g' \ 0" by simp - hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith - hence "?g'= 1 \ ?g' > 1" by arith - moreover {assume "?g'=1" hence ?thesis using prems - by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)} - moreover {assume g'1:"?g'>1" - have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) - have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1) - have gpdgp: "?g' dvd ?g'" by simp - from zdvd_imp_le[OF gpdd np] have g'n: "?g' \ n" . - from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]] - have "n div ?g' >0" by simp - hence ?thesis using prems - by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0 simpnum_numbound0)} - ultimately have ?thesis by blast} - ultimately have ?thesis by blast} - ultimately show ?thesis by blast -qed - -consts simpfm :: "fm \ fm" -recdef simpfm "measure fmsize" - "simpfm (And p q) = conj (simpfm p) (simpfm q)" - "simpfm (Or p q) = disj (simpfm p) (simpfm q)" - "simpfm (Imp p q) = imp (simpfm p) (simpfm q)" - "simpfm (Iff p q) = iff (simpfm p) (simpfm q)" - "simpfm (NOT p) = not (simpfm p)" - "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \ if (v < 0) then T else F - | _ \ Lt a')" - "simpfm (Le a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Le a')" - "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \ if (v > 0) then T else F | _ \ Gt a')" - "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Ge a')" - "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \ if (v = 0) then T else F | _ \ Eq a')" - "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ NEq a')" - "simpfm p = p" -lemma simpfm: "Ifm bs (simpfm p) = Ifm bs p" -proof(induct p rule: simpfm.induct) - case (6 a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa - by (cases ?sa, simp_all add: Let_def)} - ultimately show ?case by blast -next - case (7 a) let ?sa = "simpnum a" - from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa - by (cases ?sa, simp_all add: Let_def)} - ultimately show ?case by blast -next - case (8 a) let ?sa = "simpnum a" - from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa - by (cases ?sa, simp_all add: Let_def)} - ultimately show ?case by blast -next - case (9 a) let ?sa = "simpnum a" - from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa - by (cases ?sa, simp_all add: Let_def)} - ultimately show ?case by blast -next - case (10 a) let ?sa = "simpnum a" - from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa - by (cases ?sa, simp_all add: Let_def)} - ultimately show ?case by blast -next - case (11 a) let ?sa = "simpnum a" - from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp - {fix v assume "?sa = C v" hence ?case using sa by simp } - moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa - by (cases ?sa, simp_all add: Let_def)} - ultimately show ?case by blast -qed (induct p rule: simpfm.induct, simp_all add: conj disj imp iff not) - - -lemma simpfm_bound0: "bound0 p \ bound0 (simpfm p)" -proof(induct p rule: simpfm.induct) - case (6 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def) -next - case (7 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def) -next - case (8 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def) -next - case (9 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def) -next - case (10 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def) -next - case (11 a) hence nb: "numbound0 a" by simp - hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) - thus ?case by (cases "simpnum a", auto simp add: Let_def) -qed(auto simp add: disj_def imp_def iff_def conj_def not_bn) - -lemma simpfm_qf: "qfree p \ qfree (simpfm p)" -by (induct p rule: simpfm.induct, auto simp add: disj_qf imp_qf iff_qf conj_qf not_qf Let_def) - (case_tac "simpnum a",auto)+ - -consts prep :: "fm \ fm" -recdef prep "measure fmsize" - "prep (E T) = T" - "prep (E F) = F" - "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))" - "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))" - "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" - "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))" - "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))" - "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))" - "prep (E p) = E (prep p)" - "prep (A (And p q)) = conj (prep (A p)) (prep (A q))" - "prep (A p) = prep (NOT (E (NOT p)))" - "prep (NOT (NOT p)) = prep p" - "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))" - "prep (NOT (A p)) = prep (E (NOT p))" - "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))" - "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))" - "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))" - "prep (NOT p) = not (prep p)" - "prep (Or p q) = disj (prep p) (prep q)" - "prep (And p q) = conj (prep p) (prep q)" - "prep (Imp p q) = prep (Or (NOT p) q)" - "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))" - "prep p = p" -(hints simp add: fmsize_pos) -lemma prep: "\ bs. Ifm bs (prep p) = Ifm bs p" -by (induct p rule: prep.induct, auto) - - (* Generic quantifier elimination *) -consts qelim :: "fm \ (fm \ fm) \ fm" -recdef qelim "measure fmsize" - "qelim (E p) = (\ qe. DJ qe (qelim p qe))" - "qelim (A p) = (\ qe. not (qe ((qelim (NOT p) qe))))" - "qelim (NOT p) = (\ qe. not (qelim p qe))" - "qelim (And p q) = (\ qe. conj (qelim p qe) (qelim q qe))" - "qelim (Or p q) = (\ qe. disj (qelim p qe) (qelim q qe))" - "qelim (Imp p q) = (\ qe. imp (qelim p qe) (qelim q qe))" - "qelim (Iff p q) = (\ qe. iff (qelim p qe) (qelim q qe))" - "qelim p = (\ y. simpfm p)" - -lemma qelim_ci: - assumes qe_inv: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" - shows "\ bs. qfree (qelim p qe) \ (Ifm bs (qelim p qe) = Ifm bs p)" -using qe_inv DJ_qe[OF qe_inv] -by(induct p rule: qelim.induct) -(auto simp add: not disj conj iff imp not_qf disj_qf conj_qf imp_qf iff_qf - simpfm simpfm_qf simp del: simpfm.simps) - -consts - plusinf:: "fm \ fm" (* Virtual substitution of +\*) - minusinf:: "fm \ fm" (* Virtual substitution of -\*) -recdef minusinf "measure size" - "minusinf (And p q) = conj (minusinf p) (minusinf q)" - "minusinf (Or p q) = disj (minusinf p) (minusinf q)" - "minusinf (Eq (CN 0 c e)) = F" - "minusinf (NEq (CN 0 c e)) = T" - "minusinf (Lt (CN 0 c e)) = T" - "minusinf (Le (CN 0 c e)) = T" - "minusinf (Gt (CN 0 c e)) = F" - "minusinf (Ge (CN 0 c e)) = F" - "minusinf p = p" - -recdef plusinf "measure size" - "plusinf (And p q) = conj (plusinf p) (plusinf q)" - "plusinf (Or p q) = disj (plusinf p) (plusinf q)" - "plusinf (Eq (CN 0 c e)) = F" - "plusinf (NEq (CN 0 c e)) = T" - "plusinf (Lt (CN 0 c e)) = F" - "plusinf (Le (CN 0 c e)) = F" - "plusinf (Gt (CN 0 c e)) = T" - "plusinf (Ge (CN 0 c e)) = T" - "plusinf p = p" - -consts - isrlfm :: "fm \ bool" (* Linearity test for fm *) -recdef isrlfm "measure size" - "isrlfm (And p q) = (isrlfm p \ isrlfm q)" - "isrlfm (Or p q) = (isrlfm p \ isrlfm q)" - "isrlfm (Eq (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm (NEq (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm (Lt (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm (Le (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm (Gt (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm (Ge (CN 0 c e)) = (c>0 \ numbound0 e)" - "isrlfm p = (isatom p \ (bound0 p))" - - (* splits the bounded from the unbounded part*) -consts rsplit0 :: "num \ int \ num" -recdef rsplit0 "measure num_size" - "rsplit0 (Bound 0) = (1,C 0)" - "rsplit0 (Add a b) = (let (ca,ta) = rsplit0 a ; (cb,tb) = rsplit0 b - in (ca+cb, Add ta tb))" - "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))" - "rsplit0 (Neg a) = (let (c,t) = rsplit0 a in (-c,Neg t))" - "rsplit0 (Mul c a) = (let (ca,ta) = rsplit0 a in (c*ca,Mul c ta))" - "rsplit0 (CN 0 c a) = (let (ca,ta) = rsplit0 a in (c+ca,ta))" - "rsplit0 (CN n c a) = (let (ca,ta) = rsplit0 a in (ca,CN n c ta))" - "rsplit0 t = (0,t)" -lemma rsplit0: - shows "Inum bs ((split (CN 0)) (rsplit0 t)) = Inum bs t \ numbound0 (snd (rsplit0 t))" -proof (induct t rule: rsplit0.induct) - case (2 a b) - let ?sa = "rsplit0 a" let ?sb = "rsplit0 b" - let ?ca = "fst ?sa" let ?cb = "fst ?sb" - let ?ta = "snd ?sa" let ?tb = "snd ?sb" - from prems have nb: "numbound0 (snd(rsplit0 (Add a b)))" - by(cases "rsplit0 a",auto simp add: Let_def split_def) - have "Inum bs ((split (CN 0)) (rsplit0 (Add a b))) = - Inum bs ((split (CN 0)) ?sa)+Inum bs ((split (CN 0)) ?sb)" - by (simp add: Let_def split_def ring_simps) - also have "\ = Inum bs a + Inum bs b" using prems by (cases "rsplit0 a", simp_all) - finally show ?case using nb by simp -qed(auto simp add: Let_def split_def ring_simps , simp add: right_distrib[symmetric]) - - (* Linearize a formula*) -definition - lt :: "int \ num \ fm" -where - "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) - else (Gt (CN 0 (-c) (Neg t))))" - -definition - le :: "int \ num \ fm" -where - "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) - else (Ge (CN 0 (-c) (Neg t))))" - -definition - gt :: "int \ num \ fm" -where - "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) - else (Lt (CN 0 (-c) (Neg t))))" - -definition - ge :: "int \ num \ fm" -where - "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) - else (Le (CN 0 (-c) (Neg t))))" - -definition - eq :: "int \ num \ fm" -where - "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) - else (Eq (CN 0 (-c) (Neg t))))" - -definition - neq :: "int \ num \ fm" -where - "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) - else (NEq (CN 0 (-c) (Neg t))))" - -lemma lt: "numnoabs t \ Ifm bs (split lt (rsplit0 t)) = Ifm bs (Lt t) \ isrlfm (split lt (rsplit0 t))" -using rsplit0[where bs = "bs" and t="t"] -by (auto simp add: lt_def split_def,cases "snd(rsplit0 t)",auto,case_tac "nat",auto) - -lemma le: "numnoabs t \ Ifm bs (split le (rsplit0 t)) = Ifm bs (Le t) \ isrlfm (split le (rsplit0 t))" -using rsplit0[where bs = "bs" and t="t"] -by (auto simp add: le_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) - -lemma gt: "numnoabs t \ Ifm bs (split gt (rsplit0 t)) = Ifm bs (Gt t) \ isrlfm (split gt (rsplit0 t))" -using rsplit0[where bs = "bs" and t="t"] -by (auto simp add: gt_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) - -lemma ge: "numnoabs t \ Ifm bs (split ge (rsplit0 t)) = Ifm bs (Ge t) \ isrlfm (split ge (rsplit0 t))" -using rsplit0[where bs = "bs" and t="t"] -by (auto simp add: ge_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) - -lemma eq: "numnoabs t \ Ifm bs (split eq (rsplit0 t)) = Ifm bs (Eq t) \ isrlfm (split eq (rsplit0 t))" -using rsplit0[where bs = "bs" and t="t"] -by (auto simp add: eq_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) - -lemma neq: "numnoabs t \ Ifm bs (split neq (rsplit0 t)) = Ifm bs (NEq t) \ isrlfm (split neq (rsplit0 t))" -using rsplit0[where bs = "bs" and t="t"] -by (auto simp add: neq_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) - -lemma conj_lin: "isrlfm p \ isrlfm q \ isrlfm (conj p q)" -by (auto simp add: conj_def) -lemma disj_lin: "isrlfm p \ isrlfm q \ isrlfm (disj p q)" -by (auto simp add: disj_def) - -consts rlfm :: "fm \ fm" -recdef rlfm "measure fmsize" - "rlfm (And p q) = conj (rlfm p) (rlfm q)" - "rlfm (Or p q) = disj (rlfm p) (rlfm q)" - "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)" - "rlfm (Iff p q) = disj (conj (rlfm p) (rlfm q)) (conj (rlfm (NOT p)) (rlfm (NOT q)))" - "rlfm (Lt a) = split lt (rsplit0 a)" - "rlfm (Le a) = split le (rsplit0 a)" - "rlfm (Gt a) = split gt (rsplit0 a)" - "rlfm (Ge a) = split ge (rsplit0 a)" - "rlfm (Eq a) = split eq (rsplit0 a)" - "rlfm (NEq a) = split neq (rsplit0 a)" - "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))" - "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))" - "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))" - "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))" - "rlfm (NOT (NOT p)) = rlfm p" - "rlfm (NOT T) = F" - "rlfm (NOT F) = T" - "rlfm (NOT (Lt a)) = rlfm (Ge a)" - "rlfm (NOT (Le a)) = rlfm (Gt a)" - "rlfm (NOT (Gt a)) = rlfm (Le a)" - "rlfm (NOT (Ge a)) = rlfm (Lt a)" - "rlfm (NOT (Eq a)) = rlfm (NEq a)" - "rlfm (NOT (NEq a)) = rlfm (Eq a)" - "rlfm p = p" (hints simp add: fmsize_pos) - -lemma rlfm_I: - assumes qfp: "qfree p" - shows "(Ifm bs (rlfm p) = Ifm bs p) \ isrlfm (rlfm p)" - using qfp -by (induct p rule: rlfm.induct, auto simp add: lt le gt ge eq neq conj disj conj_lin disj_lin) - - (* Operations needed for Ferrante and Rackoff *) -lemma rminusinf_inf: - assumes lp: "isrlfm p" - shows "\ z. \ x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\ z. \ x. ?P z x p") -using lp -proof (induct p rule: minusinf.induct) - case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto -next - case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto -next - case (3 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - hence "real c * x + ?e \ 0" by simp - with xz have "?P ?z x (Eq (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp - thus ?case by blast -next - case (4 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - hence "real c * x + ?e \ 0" by simp - with xz have "?P ?z x (NEq (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp - thus ?case by blast -next - case (5 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - with xz have "?P ?z x (Lt (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp - thus ?case by blast -next - case (6 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - with xz have "?P ?z x (Le (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (Le (CN 0 c e))" by simp - thus ?case by blast -next - case (7 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - with xz have "?P ?z x (Gt (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp - thus ?case by blast -next - case (8 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x < ?z" - hence "(real c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) - hence "real c * x + ?e < 0" by arith - with xz have "?P ?z x (Ge (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp - thus ?case by blast -qed simp_all - -lemma rplusinf_inf: - assumes lp: "isrlfm p" - shows "\ z. \ x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\ z. \ x. ?P z x p") -using lp -proof (induct p rule: isrlfm.induct) - case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto -next - case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto -next - case (3 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - hence "real c * x + ?e \ 0" by simp - with xz have "?P ?z x (Eq (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp - thus ?case by blast -next - case (4 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - hence "real c * x + ?e \ 0" by simp - with xz have "?P ?z x (NEq (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp - thus ?case by blast -next - case (5 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - with xz have "?P ?z x (Lt (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp - thus ?case by blast -next - case (6 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - with xz have "?P ?z x (Le (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (Le (CN 0 c e))" by simp - thus ?case by blast -next - case (7 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - with xz have "?P ?z x (Gt (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp - thus ?case by blast -next - case (8 c e) - from prems have nb: "numbound0 e" by simp - from prems have cp: "real c > 0" by simp - fix a - let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" - {fix x - assume xz: "x > ?z" - with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: mult_ac) - hence "real c * x + ?e > 0" by arith - with xz have "?P ?z x (Ge (CN 0 c e))" - using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } - hence "\ x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp - thus ?case by blast -qed simp_all - -lemma rminusinf_bound0: - assumes lp: "isrlfm p" - shows "bound0 (minusinf p)" - using lp - by (induct p rule: minusinf.induct) simp_all - -lemma rplusinf_bound0: - assumes lp: "isrlfm p" - shows "bound0 (plusinf p)" - using lp - by (induct p rule: plusinf.induct) simp_all - -lemma rminusinf_ex: - assumes lp: "isrlfm p" - and ex: "Ifm (a#bs) (minusinf p)" - shows "\ x. Ifm (x#bs) p" -proof- - from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex - have th: "\ x. Ifm (x#bs) (minusinf p)" by auto - from rminusinf_inf[OF lp, where bs="bs"] - obtain z where z_def: "\x x. Ifm (x#bs) p" -proof- - from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex - have th: "\ x. Ifm (x#bs) (plusinf p)" by auto - from rplusinf_inf[OF lp, where bs="bs"] - obtain z where z_def: "\x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast - from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp - moreover have "z + 1 > z" by simp - ultimately show ?thesis using z_def by auto -qed - -consts - uset:: "fm \ (num \ int) list" - usubst :: "fm \ (num \ int) \ fm " -recdef uset "measure size" - "uset (And p q) = (uset p @ uset q)" - "uset (Or p q) = (uset p @ uset q)" - "uset (Eq (CN 0 c e)) = [(Neg e,c)]" - "uset (NEq (CN 0 c e)) = [(Neg e,c)]" - "uset (Lt (CN 0 c e)) = [(Neg e,c)]" - "uset (Le (CN 0 c e)) = [(Neg e,c)]" - "uset (Gt (CN 0 c e)) = [(Neg e,c)]" - "uset (Ge (CN 0 c e)) = [(Neg e,c)]" - "uset p = []" -recdef usubst "measure size" - "usubst (And p q) = (\ (t,n). And (usubst p (t,n)) (usubst q (t,n)))" - "usubst (Or p q) = (\ (t,n). Or (usubst p (t,n)) (usubst q (t,n)))" - "usubst (Eq (CN 0 c e)) = (\ (t,n). Eq (Add (Mul c t) (Mul n e)))" - "usubst (NEq (CN 0 c e)) = (\ (t,n). NEq (Add (Mul c t) (Mul n e)))" - "usubst (Lt (CN 0 c e)) = (\ (t,n). Lt (Add (Mul c t) (Mul n e)))" - "usubst (Le (CN 0 c e)) = (\ (t,n). Le (Add (Mul c t) (Mul n e)))" - "usubst (Gt (CN 0 c e)) = (\ (t,n). Gt (Add (Mul c t) (Mul n e)))" - "usubst (Ge (CN 0 c e)) = (\ (t,n). Ge (Add (Mul c t) (Mul n e)))" - "usubst p = (\ (t,n). p)" - -lemma usubst_I: assumes lp: "isrlfm p" - and np: "real n > 0" and nbt: "numbound0 t" - shows "(Ifm (x#bs) (usubst p (t,n)) = Ifm (((Inum (x#bs) t)/(real n))#bs) p) \ bound0 (usubst p (t,n))" (is "(?I x (usubst p (t,n)) = ?I ?u p) \ ?B p" is "(_ = ?I (?t/?n) p) \ _" is "(_ = ?I (?N x t /_) p) \ _") - using lp -proof(induct p rule: usubst.induct) - case (5 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - have "?I ?u (Lt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) < 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)" - by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) < 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -next - case (6 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) \ 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -next - case (7 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)" - by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) > 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -next - case (8 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) \ 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -next - case (3 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - from np have np: "real n \ 0" by simp - have "?I ?u (Eq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) = 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)" - by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) = 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -next - case (4 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ - from np have np: "real n \ 0" by simp - have "?I ?u (NEq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" - using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" - and b="0", simplified divide_zero_left]) (simp only: ring_simps) - also have "\ = (real c *?t + ?n* (?N x e) \ 0)" - using np by simp - finally show ?case using nbt nb by (simp add: ring_simps) -qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"] nth_pos2) - -lemma uset_l: - assumes lp: "isrlfm p" - shows "\ (t,k) \ set (uset p). numbound0 t \ k >0" -using lp -by(induct p rule: uset.induct,auto) - -lemma rminusinf_uset: - assumes lp: "isrlfm p" - and nmi: "\ (Ifm (a#bs) (minusinf p))" (is "\ (Ifm (a#bs) (?M p))") - and ex: "Ifm (x#bs) p" (is "?I x p") - shows "\ (s,m) \ set (uset p). x \ Inum (a#bs) s / real m" (is "\ (s,m) \ ?U p. x \ ?N a s / real m") -proof- - have "\ (s,m) \ set (uset p). real m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real m *x \ ?N a s") - using lp nmi ex - by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) - then obtain s m where smU: "(s,m) \ set (uset p)" and mx: "real m * x \ ?N a s" by blast - from uset_l[OF lp] smU have mp: "real m > 0" by auto - from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" - by (auto simp add: mult_commute) - thus ?thesis using smU by auto -qed - -lemma rplusinf_uset: - assumes lp: "isrlfm p" - and nmi: "\ (Ifm (a#bs) (plusinf p))" (is "\ (Ifm (a#bs) (?M p))") - and ex: "Ifm (x#bs) p" (is "?I x p") - shows "\ (s,m) \ set (uset p). x \ Inum (a#bs) s / real m" (is "\ (s,m) \ ?U p. x \ ?N a s / real m") -proof- - have "\ (s,m) \ set (uset p). real m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real m *x \ ?N a s") - using lp nmi ex - by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) - then obtain s m where smU: "(s,m) \ set (uset p)" and mx: "real m * x \ ?N a s" by blast - from uset_l[OF lp] smU have mp: "real m > 0" by auto - from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" - by (auto simp add: mult_commute) - thus ?thesis using smU by auto -qed - -lemma lin_dense: - assumes lp: "isrlfm p" - and noS: "\ t. l < t \ t< u \ t \ (\ (t,n). Inum (x#bs) t / real n) ` set (uset p)" - (is "\ t. _ \ _ \ t \ (\ (t,n). ?N x t / real n ) ` (?U p)") - and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p" - and ly: "l < y" and yu: "y < u" - shows "Ifm (y#bs) p" -using lp px noS -proof (induct p rule: isrlfm.induct) - case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ - from prems have "x * real c + ?N x e < 0" by (simp add: ring_simps) - hence pxc: "x < (- ?N x e) / real c" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"]) - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y < (-?N x e)/ real c" - hence "y * real c < - ?N x e" - by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e < 0" by (simp add: ring_simps) - hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y > (- ?N x e) / real c" - with yu have eu: "u > (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ l" by (cases "(- ?N x e) / real c > l", auto) - with lx pxc have "False" by auto - hence ?case by simp } - ultimately show ?case by blast -next - case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp + - from prems have "x * real c + ?N x e \ 0" by (simp add: ring_simps) - hence pxc: "x \ (- ?N x e) / real c" - by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"]) - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y < (-?N x e)/ real c" - hence "y * real c < - ?N x e" - by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e < 0" by (simp add: ring_simps) - hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y > (- ?N x e) / real c" - with yu have eu: "u > (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ l" by (cases "(- ?N x e) / real c > l", auto) - with lx pxc have "False" by auto - hence ?case by simp } - ultimately show ?case by blast -next - case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ - from prems have "x * real c + ?N x e > 0" by (simp add: ring_simps) - hence pxc: "x > (- ?N x e) / real c" - by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"]) - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y > (-?N x e)/ real c" - hence "y * real c > - ?N x e" - by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e > 0" by (simp add: ring_simps) - hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y < (- ?N x e) / real c" - with ly have eu: "l < (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ u" by (cases "(- ?N x e) / real c > l", auto) - with xu pxc have "False" by auto - hence ?case by simp } - ultimately show ?case by blast -next - case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ - from prems have "x * real c + ?N x e \ 0" by (simp add: ring_simps) - hence pxc: "x \ (- ?N x e) / real c" - by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"]) - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y > (-?N x e)/ real c" - hence "y * real c > - ?N x e" - by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e > 0" by (simp add: ring_simps) - hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y < (- ?N x e) / real c" - with ly have eu: "l < (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ u" by (cases "(- ?N x e) / real c > l", auto) - with xu pxc have "False" by auto - hence ?case by simp } - ultimately show ?case by blast -next - case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ - from cp have cnz: "real c \ 0" by simp - from prems have "x * real c + ?N x e = 0" by (simp add: ring_simps) - hence pxc: "x = (- ?N x e) / real c" - by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"]) - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with lx xu have yne: "x \ - ?N x e / real c" by auto - with pxc show ?case by simp -next - case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ - from cp have cnz: "real c \ 0" by simp - from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y* real c \ -?N x e" - by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp - hence "y* real c + ?N x e \ 0" by (simp add: ring_simps) - thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] - by (simp add: ring_simps) -qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="y" and b'="x"]) - -lemma finite_set_intervals: - assumes px: "P (x::real)" - and lx: "l \ x" and xu: "x \ u" - and linS: "l\ S" and uinS: "u \ S" - and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" - shows "\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a \ x \ x \ b \ P x" -proof- - let ?Mx = "{y. y\ S \ y \ x}" - let ?xM = "{y. y\ S \ x \ y}" - let ?a = "Max ?Mx" - let ?b = "Min ?xM" - have MxS: "?Mx \ S" by blast - hence fMx: "finite ?Mx" using fS finite_subset by auto - from lx linS have linMx: "l \ ?Mx" by blast - hence Mxne: "?Mx \ {}" by blast - have xMS: "?xM \ S" by blast - hence fxM: "finite ?xM" using fS finite_subset by auto - from xu uinS have linxM: "u \ ?xM" by blast - hence xMne: "?xM \ {}" by blast - have ax:"?a \ x" using Mxne fMx by auto - have xb:"x \ ?b" using xMne fxM by auto - have "?a \ ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \ S" using MxS by blast - have "?b \ ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \ S" using xMS by blast - have noy:"\ y. ?a < y \ y < ?b \ y \ S" - proof(clarsimp) - fix y - assume ay: "?a < y" and yb: "y < ?b" and yS: "y \ S" - from yS have "y\ ?Mx \ y\ ?xM" by auto - moreover {assume "y \ ?Mx" hence "y \ ?a" using Mxne fMx by auto with ay have "False" by simp} - moreover {assume "y \ ?xM" hence "y \ ?b" using xMne fxM by auto with yb have "False" by simp} - ultimately show "False" by blast - qed - from ainS binS noy ax xb px show ?thesis by blast -qed - -lemma finite_set_intervals2: - assumes px: "P (x::real)" - and lx: "l \ x" and xu: "x \ u" - and linS: "l\ S" and uinS: "u \ S" - and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" - shows "(\ s\ S. P s) \ (\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a < x \ x < b \ P x)" -proof- - from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su] - obtain a and b where - as: "a\ S" and bs: "b\ S" and noS:"\y. a < y \ y < b \ y \ S" and axb: "a \ x \ x \ b \ P x" by auto - from axb have "x= a \ x= b \ (a < x \ x < b)" by auto - thus ?thesis using px as bs noS by blast -qed - -lemma rinf_uset: - assumes lp: "isrlfm p" - and nmi: "\ (Ifm (x#bs) (minusinf p))" (is "\ (Ifm (x#bs) (?M p))") - and npi: "\ (Ifm (x#bs) (plusinf p))" (is "\ (Ifm (x#bs) (?P p))") - and ex: "\ x. Ifm (x#bs) p" (is "\ x. ?I x p") - shows "\ (l,n) \ set (uset p). \ (s,m) \ set (uset p). ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p" -proof- - let ?N = "\ x t. Inum (x#bs) t" - let ?U = "set (uset p)" - from ex obtain a where pa: "?I a p" by blast - from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi - have nmi': "\ (?I a (?M p))" by simp - from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi - have npi': "\ (?I a (?P p))" by simp - have "\ (l,n) \ set (uset p). \ (s,m) \ set (uset p). ?I ((?N a l/real n + ?N a s /real m) / 2) p" - proof- - let ?M = "(\ (t,c). ?N a t / real c) ` ?U" - have fM: "finite ?M" by auto - from rminusinf_uset[OF lp nmi pa] rplusinf_uset[OF lp npi pa] - have "\ (l,n) \ set (uset p). \ (s,m) \ set (uset p). a \ ?N x l / real n \ a \ ?N x s / real m" by blast - then obtain "t" "n" "s" "m" where - tnU: "(t,n) \ ?U" and smU: "(s,m) \ ?U" - and xs1: "a \ ?N x s / real m" and tx1: "a \ ?N x t / real n" by blast - from uset_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 have xs: "a \ ?N a s / real m" and tx: "a \ ?N a t / real n" by auto - from tnU have Mne: "?M \ {}" by auto - hence Une: "?U \ {}" by simp - let ?l = "Min ?M" - let ?u = "Max ?M" - have linM: "?l \ ?M" using fM Mne by simp - have uinM: "?u \ ?M" using fM Mne by simp - have tnM: "?N a t / real n \ ?M" using tnU by auto - have smM: "?N a s / real m \ ?M" using smU by auto - have lM: "\ t\ ?M. ?l \ t" using Mne fM by auto - have Mu: "\ t\ ?M. t \ ?u" using Mne fM by auto - have "?l \ ?N a t / real n" using tnM Mne by simp hence lx: "?l \ a" using tx by simp - have "?N a s / real m \ ?u" using smM Mne by simp hence xu: "a \ ?u" using xs by simp - from finite_set_intervals2[where P="\ x. ?I x p",OF pa lx xu linM uinM fM lM Mu] - have "(\ s\ ?M. ?I s p) \ - (\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p)" . - moreover { fix u assume um: "u\ ?M" and pu: "?I u p" - hence "\ (tu,nu) \ ?U. u = ?N a tu / real nu" by auto - then obtain "tu" "nu" where tuU: "(tu,nu) \ ?U" and tuu:"u= ?N a tu / real nu" by blast - have "(u + u) / 2 = u" by auto with pu tuu - have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" by simp - with tuU have ?thesis by blast} - moreover{ - assume "\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p" - then obtain t1 and t2 where t1M: "t1 \ ?M" and t2M: "t2\ ?M" - and noM: "\ y. t1 < y \ y < t2 \ y \ ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p" - by blast - from t1M have "\ (t1u,t1n) \ ?U. t1 = ?N a t1u / real t1n" by auto - then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \ ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast - from t2M have "\ (t2u,t2n) \ ?U. t2 = ?N a t2u / real t2n" by auto - then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \ ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast - from t1x xt2 have t1t2: "t1 < t2" by simp - let ?u = "(t1 + t2) / 2" - from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto - from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" . - with t1uU t2uU t1u t2u have ?thesis by blast} - ultimately show ?thesis by blast - qed - then obtain "l" "n" "s" "m" where lnU: "(l,n) \ ?U" and smU:"(s,m) \ ?U" - and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast - from lnU smU uset_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto - from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] - numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu - have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp - with lnU smU - show ?thesis by auto -qed - (* The Ferrante - Rackoff Theorem *) - -theorem fr_eq: - assumes lp: "isrlfm p" - shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,n) \ set (uset p). \ (s,m) \ set (uset p). Ifm ((((Inum (x#bs) t)/ real n + (Inum (x#bs) s) / real m) /2)#bs) p))" - (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") -proof - assume px: "\ x. ?I x p" - have "?M \ ?P \ (\ ?M \ \ ?P)" by blast - moreover {assume "?M \ ?P" hence "?D" by blast} - moreover {assume nmi: "\ ?M" and npi: "\ ?P" - from rinf_uset[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast} - ultimately show "?D" by blast -next - assume "?D" - moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} - moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } - moreover {assume f:"?F" hence "?E" by blast} - ultimately show "?E" by blast -qed - - -lemma fr_equsubst: - assumes lp: "isrlfm p" - shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,k) \ set (uset p). \ (s,l) \ set (uset p). Ifm (x#bs) (usubst p (Add(Mul l t) (Mul k s) , 2*k*l))))" - (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") -proof - assume px: "\ x. ?I x p" - have "?M \ ?P \ (\ ?M \ \ ?P)" by blast - moreover {assume "?M \ ?P" hence "?D" by blast} - moreover {assume nmi: "\ ?M" and npi: "\ ?P" - let ?f ="\ (t,n). Inum (x#bs) t / real n" - let ?N = "\ t. Inum (x#bs) t" - {fix t n s m assume "(t,n)\ set (uset p)" and "(s,m) \ set (uset p)" - with uset_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" - by auto - let ?st = "Add (Mul m t) (Mul n s)" - from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute) - from tnb snb have st_nb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" - using mnp mp np by (simp add: ring_simps add_divide_distrib) - from usubst_I[OF lp mnp st_nb, where x="x" and bs="bs"] - have "?I x (usubst p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])} - with rinf_uset[OF lp nmi npi px] have "?F" by blast hence "?D" by blast} - ultimately show "?D" by blast -next - assume "?D" - moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} - moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } - moreover {fix t k s l assume "(t,k) \ set (uset p)" and "(s,l) \ set (uset p)" - and px:"?I x (usubst p (Add (Mul l t) (Mul k s), 2*k*l))" - with uset_l[OF lp] have tnb: "numbound0 t" and np:"real k > 0" and snb: "numbound0 s" and mp:"real l > 0" by auto - let ?st = "Add (Mul l t) (Mul k s)" - from mult_pos_pos[OF np mp] have mnp: "real (2*k*l) > 0" - by (simp add: mult_commute) - from tnb snb have st_nb: "numbound0 ?st" by simp - from usubst_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto} - ultimately show "?E" by blast -qed - - - (* Implement the right hand side of Ferrante and Rackoff's Theorem. *) -constdefs ferrack:: "fm \ fm" - "ferrack p \ (let p' = rlfm (simpfm p); mp = minusinf p'; pp = plusinf p' - in if (mp = T \ pp = T) then T else - (let U = remdps(map simp_num_pair - (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) - (alluopairs (uset p')))) - in decr (disj mp (disj pp (evaldjf (simpfm o (usubst p')) U)))))" - -lemma uset_cong_aux: - assumes Ul: "\ (t,n) \ set U. numbound0 t \ n >0" - shows "((\ (t,n). Inum (x#bs) t /real n) ` (set (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs U)))) = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (set U \ set U))" - (is "?lhs = ?rhs") -proof(auto) - fix t n s m - assume "((t,n),(s,m)) \ set (alluopairs U)" - hence th: "((t,n),(s,m)) \ (set U \ set U)" - using alluopairs_set1[where xs="U"] by blast - let ?N = "\ t. Inum (x#bs) t" - let ?st= "Add (Mul m t) (Mul n s)" - from Ul th have mnz: "m \ 0" by auto - from Ul th have nnz: "n \ 0" by auto - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" - using mnz nnz by (simp add: ring_simps add_divide_distrib) - - thus "(real m * Inum (x # bs) t + real n * Inum (x # bs) s) / - (2 * real n * real m) - \ (\((t, n), s, m). - (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) ` - (set U \ set U)"using mnz nnz th - apply (auto simp add: th add_divide_distrib ring_simps split_def image_def) - by (rule_tac x="(s,m)" in bexI,simp_all) - (rule_tac x="(t,n)" in bexI,simp_all) -next - fix t n s m - assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" - let ?N = "\ t. Inum (x#bs) t" - let ?st= "Add (Mul m t) (Mul n s)" - from Ul smU have mnz: "m \ 0" by auto - from Ul tnU have nnz: "n \ 0" by auto - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" - using mnz nnz by (simp add: ring_simps add_divide_distrib) - let ?P = "\ (t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" - have Pc:"\ a b. ?P a b = ?P b a" - by auto - from Ul alluopairs_set1 have Up:"\ ((t,n),(s,m)) \ set (alluopairs U). n \ 0 \ m \ 0" by blast - from alluopairs_ex[OF Pc, where xs="U"] tnU smU - have th':"\ ((t',n'),(s',m')) \ set (alluopairs U). ?P (t',n') (s',m')" - by blast - then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \ set (alluopairs U)" - and Pts': "?P (t',n') (s',m')" by blast - from ts'_U Up have mnz': "m' \ 0" and nnz': "n'\ 0" by auto - let ?st' = "Add (Mul m' t') (Mul n' s')" - have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')" - using mnz' nnz' by (simp add: ring_simps add_divide_distrib) - from Pts' have - "(Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" by simp - also have "\ = ((\(t, n). Inum (x # bs) t / real n) ((\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st') - finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2 - \ (\(t, n). Inum (x # bs) t / real n) ` - (\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ` - set (alluopairs U)" - using ts'_U by blast -qed - -lemma uset_cong: - assumes lp: "isrlfm p" - and UU': "((\ (t,n). Inum (x#bs) t /real n) ` U') = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (U \ U))" (is "?f ` U' = ?g ` (U\U)") - and U: "\ (t,n) \ U. numbound0 t \ n > 0" - and U': "\ (t,n) \ U'. numbound0 t \ n > 0" - shows "(\ (t,n) \ U. \ (s,m) \ U. Ifm (x#bs) (usubst p (Add (Mul m t) (Mul n s),2*n*m))) = (\ (t,n) \ U'. Ifm (x#bs) (usubst p (t,n)))" - (is "?lhs = ?rhs") -proof - assume ?lhs - then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and - Pst: "Ifm (x#bs) (usubst p (Add (Mul m t) (Mul n s),2*n*m))" by blast - let ?N = "\ t. Inum (x#bs) t" - from tnU smU U have tnb: "numbound0 t" and np: "n > 0" - and snb: "numbound0 s" and mp:"m > 0" by auto - let ?st= "Add (Mul m t) (Mul n s)" - from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) - from tnb snb have stnb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" - using mp np by (simp add: ring_simps add_divide_distrib) - from tnU smU UU' have "?g ((t,n),(s,m)) \ ?f ` U'" by blast - hence "\ (t',n') \ U'. ?g ((t,n),(s,m)) = ?f (t',n')" - by auto (rule_tac x="(a,b)" in bexI, auto) - then obtain t' n' where tnU': "(t',n') \ U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast - from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto - from usubst_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst - have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp - from conjunct1[OF usubst_I[OF lp np' tnb', where bs="bs" and x="x"], symmetric] th[simplified split_def fst_conv snd_conv,symmetric] Pst2[simplified st[symmetric]] - have "Ifm (x # bs) (usubst p (t', n')) " by (simp only: st) - then show ?rhs using tnU' by auto -next - assume ?rhs - then obtain t' n' where tnU': "(t',n') \ U'" and Pt': "Ifm (x # bs) (usubst p (t', n'))" - by blast - from tnU' UU' have "?f (t',n') \ ?g ` (U\U)" by blast - hence "\ ((t,n),(s,m)) \ (U\U). ?f (t',n') = ?g ((t,n),(s,m))" - by auto (rule_tac x="(a,b)" in bexI, auto) - then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and - th: "?f (t',n') = ?g((t,n),(s,m)) "by blast - let ?N = "\ t. Inum (x#bs) t" - from tnU smU U have tnb: "numbound0 t" and np: "n > 0" - and snb: "numbound0 s" and mp:"m > 0" by auto - let ?st= "Add (Mul m t) (Mul n s)" - from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) - from tnb snb have stnb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" - using mp np by (simp add: ring_simps add_divide_distrib) - from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto - from usubst_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt' - have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp - with usubst_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast -qed - -lemma ferrack: - assumes qf: "qfree p" - shows "qfree (ferrack p) \ ((Ifm bs (ferrack p)) = (\ x. Ifm (x#bs) p))" - (is "_ \ (?rhs = ?lhs)") -proof- - let ?I = "\ x p. Ifm (x#bs) p" - fix x - let ?N = "\ t. Inum (x#bs) t" - let ?q = "rlfm (simpfm p)" - let ?U = "uset ?q" - let ?Up = "alluopairs ?U" - let ?g = "\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)" - let ?S = "map ?g ?Up" - let ?SS = "map simp_num_pair ?S" - let ?Y = "remdps ?SS" - let ?f= "(\ (t,n). ?N t / real n)" - let ?h = "\ ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2" - let ?F = "\ p. \ a \ set (uset p). \ b \ set (uset p). ?I x (usubst p (?g(a,b)))" - let ?ep = "evaldjf (simpfm o (usubst ?q)) ?Y" - from rlfm_I[OF simpfm_qf[OF qf]] have lq: "isrlfm ?q" by blast - from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \ (set ?U \ set ?U)" by simp - from uset_l[OF lq] have U_l: "\ (t,n) \ set ?U. numbound0 t \ n > 0" . - from U_l UpU - have "\ ((t,n),(s,m)) \ set ?Up. numbound0 t \ n> 0 \ numbound0 s \ m > 0" by auto - hence Snb: "\ (t,n) \ set ?S. numbound0 t \ n > 0 " - by (auto simp add: mult_pos_pos) - have Y_l: "\ (t,n) \ set ?Y. numbound0 t \ n > 0" - proof- - { fix t n assume tnY: "(t,n) \ set ?Y" - hence "(t,n) \ set ?SS" by simp - hence "\ (t',n') \ set ?S. simp_num_pair (t',n') = (t,n)" - by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all) - then obtain t' n' where tn'S: "(t',n') \ set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast - from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto - from simp_num_pair_l[OF tnb np tns] - have "numbound0 t \ n > 0" . } - thus ?thesis by blast - qed - - have YU: "(?f ` set ?Y) = (?h ` (set ?U \ set ?U))" - proof- - from simp_num_pair_ci[where bs="x#bs"] have - "\x. (?f o simp_num_pair) x = ?f x" by auto - hence th: "?f o simp_num_pair = ?f" using ext by blast - have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_compose) - also have "\ = (?f ` set ?S)" by (simp add: th) - also have "\ = ((?f o ?g) ` set ?Up)" - by (simp only: set_map o_def image_compose[symmetric]) - also have "\ = (?h ` (set ?U \ set ?U))" - using uset_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_compose[symmetric]] by blast - finally show ?thesis . - qed - have "\ (t,n) \ set ?Y. bound0 (simpfm (usubst ?q (t,n)))" - proof- - { fix t n assume tnY: "(t,n) \ set ?Y" - with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto - from usubst_I[OF lq np tnb] - have "bound0 (usubst ?q (t,n))" by simp hence "bound0 (simpfm (usubst ?q (t,n)))" - using simpfm_bound0 by simp} - thus ?thesis by blast - qed - hence ep_nb: "bound0 ?ep" using evaldjf_bound0[where xs="?Y" and f="simpfm o (usubst ?q)"] by auto - let ?mp = "minusinf ?q" - let ?pp = "plusinf ?q" - let ?M = "?I x ?mp" - let ?P = "?I x ?pp" - let ?res = "disj ?mp (disj ?pp ?ep)" - from rminusinf_bound0[OF lq] rplusinf_bound0[OF lq] ep_nb - have nbth: "bound0 ?res" by auto - - from conjunct1[OF rlfm_I[OF simpfm_qf[OF qf]]] simpfm - - have th: "?lhs = (\ x. ?I x ?q)" by auto - from th fr_equsubst[OF lq, where bs="bs" and x="x"] have lhfr: "?lhs = (?M \ ?P \ ?F ?q)" - by (simp only: split_def fst_conv snd_conv) - also have "\ = (?M \ ?P \ (\ (t,n) \ set ?Y. ?I x (simpfm (usubst ?q (t,n)))))" - using uset_cong[OF lq YU U_l Y_l] by (simp only: split_def fst_conv snd_conv simpfm) - also have "\ = (Ifm (x#bs) ?res)" - using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="simpfm o (usubst ?q)",symmetric] - by (simp add: split_def pair_collapse) - finally have lheq: "?lhs = (Ifm bs (decr ?res))" using decr[OF nbth] by blast - hence lr: "?lhs = ?rhs" apply (unfold ferrack_def Let_def) - by (cases "?mp = T \ ?pp = T", auto) (simp add: disj_def)+ - from decr_qf[OF nbth] have "qfree (ferrack p)" by (auto simp add: Let_def ferrack_def) - with lr show ?thesis by blast -qed - -definition linrqe:: "fm \ fm" where - "linrqe p = qelim (prep p) ferrack" - -theorem linrqe: "Ifm bs (linrqe p) = Ifm bs p \ qfree (linrqe p)" -using ferrack qelim_ci prep -unfolding linrqe_def by auto - -definition ferrack_test :: "unit \ fm" where - "ferrack_test u = linrqe (A (A (Imp (Lt (Sub (Bound 1) (Bound 0))) - (E (Eq (Sub (Add (Bound 0) (Bound 2)) (Bound 1)))))))" - -ML {* @{code ferrack_test} () *} - -oracle linr_oracle = {* -let - -fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t - of NONE => error "Variable not found in the list!" - | SOME n => @{code Bound} n) - | num_of_term vs @{term "real (0::int)"} = @{code C} 0 - | num_of_term vs @{term "real (1::int)"} = @{code C} 1 - | num_of_term vs @{term "0::real"} = @{code C} 0 - | num_of_term vs @{term "1::real"} = @{code C} 1 - | num_of_term vs (Bound i) = @{code Bound} i - | num_of_term vs (@{term "uminus :: real \ real"} $ t') = @{code Neg} (num_of_term vs t') - | num_of_term vs (@{term "op + :: real \ real \ real"} $ t1 $ t2) = @{code Add} (num_of_term vs t1, num_of_term vs t2) - | num_of_term vs (@{term "op - :: real \ real \ real"} $ t1 $ t2) = @{code Sub} (num_of_term vs t1, num_of_term vs t2) - | num_of_term vs (@{term "op * :: real \ real \ real"} $ t1 $ t2) = (case (num_of_term vs t1) - of @{code C} i => @{code Mul} (i, num_of_term vs t2) - | _ => error "num_of_term: unsupported Multiplication") - | num_of_term vs (@{term "real :: int \ real"} $ (@{term "number_of :: int \ int"} $ t')) = @{code C} (HOLogic.dest_numeral t') - | num_of_term vs (@{term "number_of :: int \ real"} $ t') = @{code C} (HOLogic.dest_numeral t') - | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t); - -fun fm_of_term vs @{term True} = @{code T} - | fm_of_term vs @{term False} = @{code F} - | fm_of_term vs (@{term "op < :: real \ real \ bool"} $ t1 $ t2) = @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) - | fm_of_term vs (@{term "op \ :: real \ real \ bool"} $ t1 $ t2) = @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) - | fm_of_term vs (@{term "op = :: real \ real \ bool"} $ t1 $ t2) = @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) - | fm_of_term vs (@{term "op \ :: bool \ bool \ bool"} $ t1 $ t2) = @{code Iff} (fm_of_term vs t1, fm_of_term vs t2) - | fm_of_term vs (@{term "op &"} $ t1 $ t2) = @{code And} (fm_of_term vs t1, fm_of_term vs t2) - | fm_of_term vs (@{term "op |"} $ t1 $ t2) = @{code Or} (fm_of_term vs t1, fm_of_term vs t2) - | fm_of_term vs (@{term "op -->"} $ t1 $ t2) = @{code Imp} (fm_of_term vs t1, fm_of_term vs t2) - | fm_of_term vs (@{term "Not"} $ t') = @{code NOT} (fm_of_term vs t') - | fm_of_term vs (Const ("Ex", _) $ Abs (xn, xT, p)) = - @{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) - | fm_of_term vs (Const ("All", _) $ Abs (xn, xT, p)) = - @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) - | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t); - -fun term_of_num vs (@{code C} i) = @{term "real :: int \ real"} $ HOLogic.mk_number HOLogic.intT i - | term_of_num vs (@{code Bound} n) = fst (the (find_first (fn (_, m) => n = m) vs)) - | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \ real"} $ term_of_num vs t' - | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: real \ real \ real"} $ - term_of_num vs t1 $ term_of_num vs t2 - | term_of_num vs (@{code Sub} (t1, t2)) = @{term "op - :: real \ real \ real"} $ - term_of_num vs t1 $ term_of_num vs t2 - | term_of_num vs (@{code Mul} (i, t2)) = @{term "op * :: real \ real \ real"} $ - term_of_num vs (@{code C} i) $ term_of_num vs t2 - | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t)); - -fun term_of_fm vs @{code T} = HOLogic.true_const - | term_of_fm vs @{code F} = HOLogic.false_const - | term_of_fm vs (@{code Lt} t) = @{term "op < :: real \ real \ bool"} $ - term_of_num vs t $ @{term "0::real"} - | term_of_fm vs (@{code Le} t) = @{term "op \ :: real \ real \ bool"} $ - term_of_num vs t $ @{term "0::real"} - | term_of_fm vs (@{code Gt} t) = @{term "op < :: real \ real \ bool"} $ - @{term "0::real"} $ term_of_num vs t - | term_of_fm vs (@{code Ge} t) = @{term "op \ :: real \ real \ bool"} $ - @{term "0::real"} $ term_of_num vs t - | term_of_fm vs (@{code Eq} t) = @{term "op = :: real \ real \ bool"} $ - term_of_num vs t $ @{term "0::real"} - | term_of_fm vs (@{code NEq} t) = term_of_fm vs (@{code NOT} (@{code Eq} t)) - | term_of_fm vs (@{code NOT} t') = HOLogic.Not $ term_of_fm vs t' - | term_of_fm vs (@{code And} (t1, t2)) = HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2 - | term_of_fm vs (@{code Or} (t1, t2)) = HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2 - | term_of_fm vs (@{code Imp} (t1, t2)) = HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2 - | term_of_fm vs (@{code Iff} (t1, t2)) = @{term "op \ :: bool \ bool \ bool"} $ - term_of_fm vs t1 $ term_of_fm vs t2 - | term_of_fm vs _ = error "If this is raised, Isabelle/HOL or generate_code is inconsistent."; - -in fn ct => - let - val thy = Thm.theory_of_cterm ct; - val t = Thm.term_of ct; - val fs = term_frees t; - val vs = fs ~~ (0 upto (length fs - 1)); - val res = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, term_of_fm vs (@{code linrqe} (fm_of_term vs t)))); - in Thm.cterm_of thy res end -end; -*} - -use "linrtac.ML" -setup LinrTac.setup - -lemma - fixes x :: real - shows "2 * x \ 2 * x \ 2 * x \ 2 * x + 1" -apply rferrack -done - -lemma - fixes x :: real - shows "\y \ x. x = y + 1" -apply rferrack -done - -lemma - fixes x :: real - shows "\ (\z. x + z = x + z + 1)" -apply rferrack -done - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/ex/Sqrt.thy --- a/src/HOL/Complex/ex/Sqrt.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,99 +0,0 @@ -(* Title: HOL/Hyperreal/ex/Sqrt.thy - ID: $Id$ - Author: Markus Wenzel, TU Muenchen - -*) - -header {* Square roots of primes are irrational *} - -theory Sqrt -imports Primes Complex_Main -begin - -text {* The definition and the key representation theorem for the set of -rational numbers has been moved to a core theory. *} - -declare Rats_abs_nat_div_natE[elim?] - -subsection {* Main theorem *} - -text {* - The square root of any prime number (including @{text 2}) is - irrational. -*} - -theorem sqrt_prime_irrational: - assumes "prime p" - shows "sqrt (real p) \ \" -proof - from `prime p` have p: "1 < p" by (simp add: prime_def) - assume "sqrt (real p) \ \" - then obtain m n where - n: "n \ 0" and sqrt_rat: "\sqrt (real p)\ = real m / real n" - and gcd: "gcd m n = 1" .. - have eq: "m\ = p * n\" - proof - - from n and sqrt_rat have "real m = \sqrt (real p)\ * real n" by simp - then have "real (m\) = (sqrt (real p))\ * real (n\)" - by (auto simp add: power2_eq_square) - also have "(sqrt (real p))\ = real p" by simp - also have "\ * real (n\) = real (p * n\)" by simp - finally show ?thesis .. - qed - have "p dvd m \ p dvd n" - proof - from eq have "p dvd m\" .. - with `prime p` show "p dvd m" by (rule prime_dvd_power_two) - then obtain k where "m = p * k" .. - with eq have "p * n\ = p\ * k\" by (auto simp add: power2_eq_square mult_ac) - with p have "n\ = p * k\" by (simp add: power2_eq_square) - then have "p dvd n\" .. - with `prime p` show "p dvd n" by (rule prime_dvd_power_two) - qed - then have "p dvd gcd m n" .. - with gcd have "p dvd 1" by simp - then have "p \ 1" by (simp add: dvd_imp_le) - with p show False by simp -qed - -corollary "sqrt (real (2::nat)) \ \" - by (rule sqrt_prime_irrational) (rule two_is_prime) - - -subsection {* Variations *} - -text {* - Here is an alternative version of the main proof, using mostly - linear forward-reasoning. While this results in less top-down - structure, it is probably closer to proofs seen in mathematics. -*} - -theorem - assumes "prime p" - shows "sqrt (real p) \ \" -proof - from `prime p` have p: "1 < p" by (simp add: prime_def) - assume "sqrt (real p) \ \" - then obtain m n where - n: "n \ 0" and sqrt_rat: "\sqrt (real p)\ = real m / real n" - and gcd: "gcd m n = 1" .. - from n and sqrt_rat have "real m = \sqrt (real p)\ * real n" by simp - then have "real (m\) = (sqrt (real p))\ * real (n\)" - by (auto simp add: power2_eq_square) - also have "(sqrt (real p))\ = real p" by simp - also have "\ * real (n\) = real (p * n\)" by simp - finally have eq: "m\ = p * n\" .. - then have "p dvd m\" .. - with `prime p` have dvd_m: "p dvd m" by (rule prime_dvd_power_two) - then obtain k where "m = p * k" .. - with eq have "p * n\ = p\ * k\" by (auto simp add: power2_eq_square mult_ac) - with p have "n\ = p * k\" by (simp add: power2_eq_square) - then have "p dvd n\" .. - with `prime p` have "p dvd n" by (rule prime_dvd_power_two) - with dvd_m have "p dvd gcd m n" by (rule gcd_greatest) - with gcd have "p dvd 1" by simp - then have "p \ 1" by (simp add: dvd_imp_le) - with p show False by simp -qed - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/ex/Sqrt_Script.thy --- a/src/HOL/Complex/ex/Sqrt_Script.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,71 +0,0 @@ -(* Title: HOL/Hyperreal/ex/Sqrt_Script.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 2001 University of Cambridge -*) - -header {* Square roots of primes are irrational (script version) *} - -theory Sqrt_Script -imports Primes Complex_Main -begin - -text {* - \medskip Contrast this linear Isabelle/Isar script with Markus - Wenzel's more mathematical version. -*} - -subsection {* Preliminaries *} - -lemma prime_nonzero: "prime p \ p \ 0" - by (force simp add: prime_def) - -lemma prime_dvd_other_side: - "n * n = p * (k * k) \ prime p \ p dvd n" - apply (subgoal_tac "p dvd n * n", blast dest: prime_dvd_mult) - apply auto - done - -lemma reduction: "prime p \ - 0 < k \ k * k = p * (j * j) \ k < p * j \ 0 < j" - apply (rule ccontr) - apply (simp add: linorder_not_less) - apply (erule disjE) - apply (frule mult_le_mono, assumption) - apply auto - apply (force simp add: prime_def) - done - -lemma rearrange: "(j::nat) * (p * j) = k * k \ k * k = p * (j * j)" - by (simp add: mult_ac) - -lemma prime_not_square: - "prime p \ (\k. 0 < k \ m * m \ p * (k * k))" - apply (induct m rule: nat_less_induct) - apply clarify - apply (frule prime_dvd_other_side, assumption) - apply (erule dvdE) - apply (simp add: nat_mult_eq_cancel_disj prime_nonzero) - apply (blast dest: rearrange reduction) - done - - -subsection {* Main theorem *} - -text {* - The square root of any prime number (including @{text 2}) is - irrational. -*} - -theorem prime_sqrt_irrational: - "prime p \ x * x = real p \ 0 \ x \ x \ \" - apply (rule notI) - apply (erule Rats_abs_nat_div_natE) - apply (simp del: real_of_nat_mult - add: real_abs_def divide_eq_eq prime_not_square real_of_nat_mult [symmetric]) - done - -lemmas two_sqrt_irrational = - prime_sqrt_irrational [OF two_is_prime] - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/ex/document/root.tex diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/ex/linrtac.ML --- a/src/HOL/Complex/ex/linrtac.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,109 +0,0 @@ -structure LinrTac = -struct - -val trace = ref false; -fun trace_msg s = if !trace then tracing s else (); - -val ferrack_ss = let val ths = [@{thm real_of_int_inject}, @{thm real_of_int_less_iff}, - @{thm real_of_int_le_iff}] - in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths) - end; - -val binarith = - @{thms normalize_bin_simps} @ @{thms pred_bin_simps} @ @{thms succ_bin_simps} @ - @{thms add_bin_simps} @ @{thms minus_bin_simps} @ @{thms mult_bin_simps}; -val comp_arith = binarith @ simp_thms - -val zdvd_int = @{thm zdvd_int}; -val zdiff_int_split = @{thm zdiff_int_split}; -val all_nat = @{thm all_nat}; -val ex_nat = @{thm ex_nat}; -val number_of1 = @{thm number_of1}; -val number_of2 = @{thm number_of2}; -val split_zdiv = @{thm split_zdiv}; -val split_zmod = @{thm split_zmod}; -val mod_div_equality' = @{thm mod_div_equality'}; -val split_div' = @{thm split_div'}; -val Suc_plus1 = @{thm Suc_plus1}; -val imp_le_cong = @{thm imp_le_cong}; -val conj_le_cong = @{thm conj_le_cong}; -val nat_mod_add_eq = @{thm mod_add1_eq} RS sym; -val nat_mod_add_left_eq = @{thm mod_add_left_eq} RS sym; -val nat_mod_add_right_eq = @{thm mod_add_right_eq} RS sym; -val int_mod_add_eq = @{thm zmod_zadd1_eq} RS sym; -val int_mod_add_left_eq = @{thm zmod_zadd_left_eq} RS sym; -val int_mod_add_right_eq = @{thm zmod_zadd_right_eq} RS sym; -val nat_div_add_eq = @{thm div_add1_eq} RS sym; -val int_div_add_eq = @{thm zdiv_zadd1_eq} RS sym; -val ZDIVISION_BY_ZERO_MOD = @{thm DIVISION_BY_ZERO} RS conjunct2; -val ZDIVISION_BY_ZERO_DIV = @{thm DIVISION_BY_ZERO} RS conjunct1; - -fun prepare_for_linr sg q fm = - let - val ps = Logic.strip_params fm - val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm) - val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm) - fun mk_all ((s, T), (P,n)) = - if 0 mem loose_bnos P then - (HOLogic.all_const T $ Abs (s, T, P), n) - else (incr_boundvars ~1 P, n-1) - fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t; - val rhs = hs -(* val (rhs,irhs) = List.partition (relevant (rev ps)) hs *) - val np = length ps - val (fm',np) = foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n))) - (foldr HOLogic.mk_imp c rhs, np) ps - val (vs, _) = List.partition (fn t => q orelse (type_of t) = HOLogic.natT) - (term_frees fm' @ term_vars fm'); - val fm2 = foldr mk_all2 fm' vs - in (fm2, np + length vs, length rhs) end; - -(*Object quantifier to meta --*) -fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ; - -(* object implication to meta---*) -fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp; - - -fun linr_tac ctxt q i = - (ObjectLogic.atomize_prems_tac i) - THEN (REPEAT_DETERM (split_tac [@{thm split_min}, @{thm split_max}, @{thm abs_split}] i)) - THEN (fn st => - let - val g = List.nth (prems_of st, i - 1) - val thy = ProofContext.theory_of ctxt - (* Transform the term*) - val (t,np,nh) = prepare_for_linr thy q g - (* Some simpsets for dealing with mod div abs and nat*) - val simpset0 = Simplifier.theory_context thy HOL_basic_ss addsimps comp_arith - val ct = cterm_of thy (HOLogic.mk_Trueprop t) - (* Theorem for the nat --> int transformation *) - val pre_thm = Seq.hd (EVERY - [simp_tac simpset0 1, - TRY (simp_tac (Simplifier.theory_context thy ferrack_ss) 1)] - (trivial ct)) - fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) - (* The result of the quantifier elimination *) - val (th, tac) = case (prop_of pre_thm) of - Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ => - let val pth = linr_oracle (cterm_of thy (Pattern.eta_long [] t1)) - in - (trace_msg ("calling procedure with term:\n" ^ - Syntax.string_of_term ctxt t1); - ((pth RS iffD2) RS pre_thm, - assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))) - end - | _ => (pre_thm, assm_tac i) - in (rtac (((mp_step nh) o (spec_step np)) th) i - THEN tac) st - end handle Subscript => no_tac st); - -fun linr_meth src = - Method.syntax (Args.mode "no_quantify") src - #> (fn (q, ctxt) => Method.SIMPLE_METHOD' (linr_tac ctxt (not q))); - -val setup = - Method.add_method ("rferrack", linr_meth, - "decision procedure for linear real arithmetic"); - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex/ex/mirtac.ML --- a/src/HOL/Complex/ex/mirtac.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,169 +0,0 @@ -(* Title: HOL/Complex/ex/mirtac.ML - ID: $Id$ - Author: Amine Chaieb, TU Muenchen -*) - -structure MirTac = -struct - -val trace = ref false; -fun trace_msg s = if !trace then tracing s else (); - -val mir_ss = -let val ths = map thm ["real_of_int_inject", "real_of_int_less_iff", "real_of_int_le_iff"] -in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths) -end; - -val nT = HOLogic.natT; - val nat_arith = map thm ["add_nat_number_of", "diff_nat_number_of", - "mult_nat_number_of", "eq_nat_number_of", "less_nat_number_of"]; - - val comp_arith = (map thm ["Let_def", "if_False", "if_True", "add_0", - "add_Suc", "add_number_of_left", "mult_number_of_left", - "Suc_eq_add_numeral_1"])@ - (map (fn s => thm s RS sym) ["numeral_1_eq_1", "numeral_0_eq_0"]) - @ @{thms arith_simps} @ nat_arith @ @{thms rel_simps} - val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, - @{thm "real_of_nat_number_of"}, - @{thm "real_of_nat_Suc"}, @{thm "real_of_nat_one"}, @{thm "real_of_one"}, - @{thm "real_of_int_zero"}, @{thm "real_of_nat_zero"}, - @{thm "Ring_and_Field.divide_zero"}, - @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"}, - @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"}, - @{thm "diff_def"}, @{thm "minus_divide_left"}] -val comp_ths = ths @ comp_arith @ simp_thms - - -val zdvd_int = @{thm "zdvd_int"}; -val zdiff_int_split = @{thm "zdiff_int_split"}; -val all_nat = @{thm "all_nat"}; -val ex_nat = @{thm "ex_nat"}; -val number_of1 = @{thm "number_of1"}; -val number_of2 = @{thm "number_of2"}; -val split_zdiv = @{thm "split_zdiv"}; -val split_zmod = @{thm "split_zmod"}; -val mod_div_equality' = @{thm "mod_div_equality'"}; -val split_div' = @{thm "split_div'"}; -val Suc_plus1 = @{thm "Suc_plus1"}; -val imp_le_cong = @{thm "imp_le_cong"}; -val conj_le_cong = @{thm "conj_le_cong"}; -val nat_mod_add_eq = @{thm "mod_add1_eq"} RS sym; -val nat_mod_add_left_eq = @{thm "mod_add_left_eq"} RS sym; -val nat_mod_add_right_eq = @{thm "mod_add_right_eq"} RS sym; -val int_mod_add_eq = @{thm "zmod_zadd1_eq"} RS sym; -val int_mod_add_left_eq = @{thm "zmod_zadd_left_eq"} RS sym; -val int_mod_add_right_eq = @{thm "zmod_zadd_right_eq"} RS sym; -val nat_div_add_eq = @{thm "div_add1_eq"} RS sym; -val int_div_add_eq = @{thm "zdiv_zadd1_eq"} RS sym; -val ZDIVISION_BY_ZERO_MOD = @{thm "DIVISION_BY_ZERO"} RS conjunct2; -val ZDIVISION_BY_ZERO_DIV = @{thm "DIVISION_BY_ZERO"} RS conjunct1; - -fun prepare_for_mir thy q fm = - let - val ps = Logic.strip_params fm - val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm) - val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm) - fun mk_all ((s, T), (P,n)) = - if 0 mem loose_bnos P then - (HOLogic.all_const T $ Abs (s, T, P), n) - else (incr_boundvars ~1 P, n-1) - fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t; - val rhs = hs -(* val (rhs,irhs) = List.partition (relevant (rev ps)) hs *) - val np = length ps - val (fm',np) = foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n))) - (foldr HOLogic.mk_imp c rhs, np) ps - val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT) - (term_frees fm' @ term_vars fm'); - val fm2 = foldr mk_all2 fm' vs - in (fm2, np + length vs, length rhs) end; - -(*Object quantifier to meta --*) -fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ; - -(* object implication to meta---*) -fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp; - - -fun mir_tac ctxt q i = - (ObjectLogic.atomize_prems_tac i) - THEN (simp_tac (HOL_basic_ss addsimps [@{thm "abs_ge_zero"}] addsimps simp_thms) i) - THEN (REPEAT_DETERM (split_tac [@{thm "split_min"}, @{thm "split_max"},@{thm "abs_split"}] i)) - THEN (fn st => - let - val g = List.nth (prems_of st, i - 1) - val thy = ProofContext.theory_of ctxt - (* Transform the term*) - val (t,np,nh) = prepare_for_mir thy q g - (* Some simpsets for dealing with mod div abs and nat*) - val mod_div_simpset = HOL_basic_ss - addsimps [refl,nat_mod_add_eq, - @{thm "mod_self"}, @{thm "zmod_self"}, - @{thm "zdiv_zero"},@{thm "zmod_zero"},@{thm "div_0"}, @{thm "mod_0"}, - @{thm "zdiv_1"}, @{thm "zmod_1"}, @{thm "div_1"}, @{thm "mod_1"}, - @{thm "Suc_plus1"}] - addsimps @{thms add_ac} - addsimprocs [cancel_div_mod_proc] - val simpset0 = HOL_basic_ss - addsimps [mod_div_equality', Suc_plus1] - addsimps comp_ths - addsplits [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, @{thm "split_min"}, @{thm "split_max"}] - (* Simp rules for changing (n::int) to int n *) - val simpset1 = HOL_basic_ss - addsimps [@{thm "nat_number_of_def"}, @{thm "zdvd_int"}] @ map (fn r => r RS sym) - [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, - @{thm "zmult_int"}] - addsplits [@{thm "zdiff_int_split"}] - (*simp rules for elimination of int n*) - - val simpset2 = HOL_basic_ss - addsimps [@{thm "nat_0_le"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "number_of1"}, - @{thm "number_of2"}, @{thm "int_0"}, @{thm "int_1"}] - addcongs [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}] - (* simp rules for elimination of abs *) - val ct = cterm_of thy (HOLogic.mk_Trueprop t) - (* Theorem for the nat --> int transformation *) - val pre_thm = Seq.hd (EVERY - [simp_tac mod_div_simpset 1, simp_tac simpset0 1, - TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), TRY (simp_tac mir_ss 1)] - (trivial ct)) - fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) - (* The result of the quantifier elimination *) - val (th, tac) = case (prop_of pre_thm) of - Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ => - let val pth = - (* If quick_and_dirty then run without proof generation as oracle*) - if !quick_and_dirty - then mirfr_oracle (false, cterm_of thy (Pattern.eta_long [] t1)) - else mirfr_oracle (true, cterm_of thy (Pattern.eta_long [] t1)) - in - (trace_msg ("calling procedure with term:\n" ^ - Syntax.string_of_term ctxt t1); - ((pth RS iffD2) RS pre_thm, - assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))) - end - | _ => (pre_thm, assm_tac i) - in (rtac (((mp_step nh) o (spec_step np)) th) i - THEN tac) st - end handle Subscript => no_tac st); - -fun mir_args meth = - let val parse_flag = - Args.$$$ "no_quantify" >> (K (K false)); - in - Method.simple_args - (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> - curry (Library.foldl op |>) true) - (fn q => fn ctxt => meth ctxt q 1) - end; - -fun mir_method ctxt q i = Method.METHOD (fn facts => - Method.insert_tac facts 1 THEN mir_tac ctxt q i); - -val setup = - Method.add_method ("mir", - mir_args mir_method, - "decision procedure for MIR arithmetic"); - - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Complex_Main.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex_Main.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,21 @@ +(* Title: HOL/Complex_Main.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 2003 University of Cambridge +*) + +header{*Comprehensive Complex Theory*} + +theory Complex_Main +imports + Main + ContNotDenum + Real + "~~/src/HOL/Complex/Fundamental_Theorem_Algebra" + Log + Ln + Taylor + Integration + FrechetDeriv +begin + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ContNotDenum.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ContNotDenum.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,579 @@ +(* Title : HOL/ContNonDenum + Author : Benjamin Porter, Monash University, NICTA, 2005 +*) + +header {* Non-denumerability of the Continuum. *} + +theory ContNotDenum +imports RComplete Hilbert_Choice +begin + +subsection {* Abstract *} + +text {* The following document presents a proof that the Continuum is +uncountable. It is formalised in the Isabelle/Isar theorem proving +system. + +{\em Theorem:} The Continuum @{text "\"} is not denumerable. In other +words, there does not exist a function f:@{text "\\\"} such that f is +surjective. + +{\em Outline:} An elegant informal proof of this result uses Cantor's +Diagonalisation argument. The proof presented here is not this +one. First we formalise some properties of closed intervals, then we +prove the Nested Interval Property. This property relies on the +completeness of the Real numbers and is the foundation for our +argument. Informally it states that an intersection of countable +closed intervals (where each successive interval is a subset of the +last) is non-empty. We then assume a surjective function f:@{text +"\\\"} exists and find a real x such that x is not in the range of f +by generating a sequence of closed intervals then using the NIP. *} + +subsection {* Closed Intervals *} + +text {* This section formalises some properties of closed intervals. *} + +subsubsection {* Definition *} + +definition + closed_int :: "real \ real \ real set" where + "closed_int x y = {z. x \ z \ z \ y}" + +subsubsection {* Properties *} + +lemma closed_int_subset: + assumes xy: "x1 \ x0" "y1 \ y0" + shows "closed_int x1 y1 \ closed_int x0 y0" +proof - + { + fix x::real + assume "x \ closed_int x1 y1" + hence "x \ x1 \ x \ y1" by (simp add: closed_int_def) + with xy have "x \ x0 \ x \ y0" by auto + hence "x \ closed_int x0 y0" by (simp add: closed_int_def) + } + thus ?thesis by auto +qed + +lemma closed_int_least: + assumes a: "a \ b" + shows "a \ closed_int a b \ (\x \ closed_int a b. a \ x)" +proof + from a have "a\{x. a\x \ x\b}" by simp + thus "a \ closed_int a b" by (unfold closed_int_def) +next + have "\x\{x. a\x \ x\b}. a\x" by simp + thus "\x \ closed_int a b. a \ x" by (unfold closed_int_def) +qed + +lemma closed_int_most: + assumes a: "a \ b" + shows "b \ closed_int a b \ (\x \ closed_int a b. x \ b)" +proof + from a have "b\{x. a\x \ x\b}" by simp + thus "b \ closed_int a b" by (unfold closed_int_def) +next + have "\x\{x. a\x \ x\b}. x\b" by simp + thus "\x \ closed_int a b. x\b" by (unfold closed_int_def) +qed + +lemma closed_not_empty: + shows "a \ b \ \x. x \ closed_int a b" + by (auto dest: closed_int_least) + +lemma closed_mem: + assumes "a \ c" and "c \ b" + shows "c \ closed_int a b" + using assms unfolding closed_int_def by auto + +lemma closed_subset: + assumes ac: "a \ b" "c \ d" + assumes closed: "closed_int a b \ closed_int c d" + shows "b \ c" +proof - + from closed have "\x\closed_int a b. x\closed_int c d" by auto + hence "\x. a\x \ x\b \ c\x \ x\d" by (unfold closed_int_def, auto) + with ac have "c\b \ b\d" by simp + thus ?thesis by auto +qed + + +subsection {* Nested Interval Property *} + +theorem NIP: + fixes f::"nat \ real set" + assumes subset: "\n. f (Suc n) \ f n" + and closed: "\n. \a b. f n = closed_int a b \ a \ b" + shows "(\n. f n) \ {}" +proof - + let ?g = "\n. (SOME c. c\(f n) \ (\x\(f n). c \ x))" + have ne: "\n. \x. x\(f n)" + proof + fix n + from closed have "\a b. f n = closed_int a b \ a \ b" by simp + then obtain a and b where fn: "f n = closed_int a b \ a \ b" by auto + hence "a \ b" .. + with closed_not_empty have "\x. x\closed_int a b" by simp + with fn show "\x. x\(f n)" by simp + qed + + have gdef: "\n. (?g n)\(f n) \ (\x\(f n). (?g n)\x)" + proof + fix n + from closed have "\a b. f n = closed_int a b \ a \ b" .. + then obtain a and b where ff: "f n = closed_int a b" and "a \ b" by auto + hence "a \ b" by simp + hence "a\closed_int a b \ (\x\closed_int a b. a \ x)" by (rule closed_int_least) + with ff have "a\(f n) \ (\x\(f n). a \ x)" by simp + hence "\c. c\(f n) \ (\x\(f n). c \ x)" .. + thus "(?g n)\(f n) \ (\x\(f n). (?g n)\x)" by (rule someI_ex) + qed + + -- "A denotes the set of all left-most points of all the intervals ..." + moreover obtain A where Adef: "A = ?g ` \" by simp + ultimately have "\x. x\A" + proof - + have "(0::nat) \ \" by simp + moreover have "?g 0 = ?g 0" by simp + ultimately have "?g 0 \ ?g ` \" by (rule rev_image_eqI) + with Adef have "?g 0 \ A" by simp + thus ?thesis .. + qed + + -- "Now show that A is bounded above ..." + moreover have "\y. isUb (UNIV::real set) A y" + proof - + { + fix n + from ne have ex: "\x. x\(f n)" .. + from gdef have "(?g n)\(f n) \ (\x\(f n). (?g n)\x)" by simp + moreover + from closed have "\a b. f n = closed_int a b \ a \ b" .. + then obtain a and b where "f n = closed_int a b \ a \ b" by auto + hence "b\(f n) \ (\x\(f n). x \ b)" using closed_int_most by blast + ultimately have "\x\(f n). (?g n) \ b" by simp + with ex have "(?g n) \ b" by auto + hence "\b. (?g n) \ b" by auto + } + hence aux: "\n. \b. (?g n) \ b" .. + + have fs: "\n::nat. f n \ f 0" + proof (rule allI, induct_tac n) + show "f 0 \ f 0" by simp + next + fix n + assume "f n \ f 0" + moreover from subset have "f (Suc n) \ f n" .. + ultimately show "f (Suc n) \ f 0" by simp + qed + have "\n. (?g n)\(f 0)" + proof + fix n + from gdef have "(?g n)\(f n) \ (\x\(f n). (?g n)\x)" by simp + hence "?g n \ f n" .. + with fs show "?g n \ f 0" by auto + qed + moreover from closed + obtain a and b where "f 0 = closed_int a b" and alb: "a \ b" by blast + ultimately have "\n. ?g n \ closed_int a b" by auto + with alb have "\n. ?g n \ b" using closed_int_most by blast + with Adef have "\y\A. y\b" by auto + hence "A *<= b" by (unfold setle_def) + moreover have "b \ (UNIV::real set)" by simp + ultimately have "A *<= b \ b \ (UNIV::real set)" by simp + hence "isUb (UNIV::real set) A b" by (unfold isUb_def) + thus ?thesis by auto + qed + -- "by the Axiom Of Completeness, A has a least upper bound ..." + ultimately have "\t. isLub UNIV A t" by (rule reals_complete) + + -- "denote this least upper bound as t ..." + then obtain t where tdef: "isLub UNIV A t" .. + + -- "and finally show that this least upper bound is in all the intervals..." + have "\n. t \ f n" + proof + fix n::nat + from closed obtain a and b where + int: "f n = closed_int a b" and alb: "a \ b" by blast + + have "t \ a" + proof - + have "a \ A" + proof - + (* by construction *) + from alb int have ain: "a\f n \ (\x\f n. a \ x)" + using closed_int_least by blast + moreover have "\e. e\f n \ (\x\f n. e \ x) \ e = a" + proof clarsimp + fix e + assume ein: "e \ f n" and lt: "\x\f n. e \ x" + from lt ain have aux: "\x\f n. a \ x \ e \ x" by auto + + from ein aux have "a \ e \ e \ e" by auto + moreover from ain aux have "a \ a \ e \ a" by auto + ultimately show "e = a" by simp + qed + hence "\e. e\f n \ (\x\f n. e \ x) \ e = a" by simp + ultimately have "(?g n) = a" by (rule some_equality) + moreover + { + have "n = of_nat n" by simp + moreover have "of_nat n \ \" by simp + ultimately have "n \ \" + apply - + apply (subst(asm) eq_sym_conv) + apply (erule subst) + . + } + with Adef have "(?g n) \ A" by auto + ultimately show ?thesis by simp + qed + with tdef show "a \ t" by (rule isLubD2) + qed + moreover have "t \ b" + proof - + have "isUb UNIV A b" + proof - + { + from alb int have + ain: "b\f n \ (\x\f n. x \ b)" using closed_int_most by blast + + have subsetd: "\m. \n. f (n + m) \ f n" + proof (rule allI, induct_tac m) + show "\n. f (n + 0) \ f n" by simp + next + fix m n + assume pp: "\p. f (p + n) \ f p" + { + fix p + from pp have "f (p + n) \ f p" by simp + moreover from subset have "f (Suc (p + n)) \ f (p + n)" by auto + hence "f (p + (Suc n)) \ f (p + n)" by simp + ultimately have "f (p + (Suc n)) \ f p" by simp + } + thus "\p. f (p + Suc n) \ f p" .. + qed + have subsetm: "\\ \. \ \ \ \ (f \) \ (f \)" + proof ((rule allI)+, rule impI) + fix \::nat and \::nat + assume "\ \ \" + hence "\k. \ = \ + k" by (simp only: le_iff_add) + then obtain k where "\ = \ + k" .. + moreover + from subsetd have "f (\ + k) \ f \" by simp + ultimately show "f \ \ f \" by auto + qed + + fix m + { + assume "m \ n" + with subsetm have "f m \ f n" by simp + with ain have "\x\f m. x \ b" by auto + moreover + from gdef have "?g m \ f m \ (\x\f m. ?g m \ x)" by simp + ultimately have "?g m \ b" by auto + } + moreover + { + assume "\(m \ n)" + hence "m < n" by simp + with subsetm have sub: "(f n) \ (f m)" by simp + from closed obtain ma and mb where + "f m = closed_int ma mb \ ma \ mb" by blast + hence one: "ma \ mb" and fm: "f m = closed_int ma mb" by auto + from one alb sub fm int have "ma \ b" using closed_subset by blast + moreover have "(?g m) = ma" + proof - + from gdef have "?g m \ f m \ (\x\f m. ?g m \ x)" .. + moreover from one have + "ma \ closed_int ma mb \ (\x\closed_int ma mb. ma \ x)" + by (rule closed_int_least) + with fm have "ma\f m \ (\x\f m. ma \ x)" by simp + ultimately have "ma \ ?g m \ ?g m \ ma" by auto + thus "?g m = ma" by auto + qed + ultimately have "?g m \ b" by simp + } + ultimately have "?g m \ b" by (rule case_split) + } + with Adef have "\y\A. y\b" by auto + hence "A *<= b" by (unfold setle_def) + moreover have "b \ (UNIV::real set)" by simp + ultimately have "A *<= b \ b \ (UNIV::real set)" by simp + thus "isUb (UNIV::real set) A b" by (unfold isUb_def) + qed + with tdef show "t \ b" by (rule isLub_le_isUb) + qed + ultimately have "t \ closed_int a b" by (rule closed_mem) + with int show "t \ f n" by simp + qed + hence "t \ (\n. f n)" by auto + thus ?thesis by auto +qed + +subsection {* Generating the intervals *} + +subsubsection {* Existence of non-singleton closed intervals *} + +text {* This lemma asserts that given any non-singleton closed +interval (a,b) and any element c, there exists a closed interval that +is a subset of (a,b) and that does not contain c and is a +non-singleton itself. *} + +lemma closed_subset_ex: + fixes c::real + assumes alb: "a < b" + shows + "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ c \ (closed_int ka kb)" +proof - + { + assume clb: "c < b" + { + assume cla: "c < a" + from alb cla clb have "c \ closed_int a b" by (unfold closed_int_def, auto) + with alb have + "a < b \ closed_int a b \ closed_int a b \ c \ closed_int a b" + by auto + hence + "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ c \ (closed_int ka kb)" + by auto + } + moreover + { + assume ncla: "\(c < a)" + with clb have cdef: "a \ c \ c < b" by simp + obtain ka where kadef: "ka = (c + b)/2" by blast + + from kadef clb have kalb: "ka < b" by auto + moreover from kadef cdef have kagc: "ka > c" by simp + ultimately have "c\(closed_int ka b)" by (unfold closed_int_def, auto) + moreover from cdef kagc have "ka \ a" by simp + hence "closed_int ka b \ closed_int a b" by (unfold closed_int_def, auto) + ultimately have + "ka < b \ closed_int ka b \ closed_int a b \ c \ closed_int ka b" + using kalb by auto + hence + "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ c \ (closed_int ka kb)" + by auto + + } + ultimately have + "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ c \ (closed_int ka kb)" + by (rule case_split) + } + moreover + { + assume "\ (c < b)" + hence cgeb: "c \ b" by simp + + obtain kb where kbdef: "kb = (a + b)/2" by blast + with alb have kblb: "kb < b" by auto + with kbdef cgeb have "a < kb \ kb < c" by auto + moreover hence "c \ (closed_int a kb)" by (unfold closed_int_def, auto) + moreover from kblb have + "closed_int a kb \ closed_int a b" by (unfold closed_int_def, auto) + ultimately have + "a < kb \ closed_int a kb \ closed_int a b \ c\closed_int a kb" + by simp + hence + "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ c \ (closed_int ka kb)" + by auto + } + ultimately show ?thesis by (rule case_split) +qed + +subsection {* newInt: Interval generation *} + +text {* Given a function f:@{text "\\\"}, newInt (Suc n) f returns a +closed interval such that @{text "newInt (Suc n) f \ newInt n f"} and +does not contain @{text "f (Suc n)"}. With the base case defined such +that @{text "(f 0)\newInt 0 f"}. *} + +subsubsection {* Definition *} + +primrec newInt :: "nat \ (nat \ real) \ (real set)" where + "newInt 0 f = closed_int (f 0 + 1) (f 0 + 2)" + | "newInt (Suc n) f = + (SOME e. (\e1 e2. + e1 < e2 \ + e = closed_int e1 e2 \ + e \ (newInt n f) \ + (f (Suc n)) \ e) + )" + +declare newInt.simps [code del] + +subsubsection {* Properties *} + +text {* We now show that every application of newInt returns an +appropriate interval. *} + +lemma newInt_ex: + "\a b. a < b \ + newInt (Suc n) f = closed_int a b \ + newInt (Suc n) f \ newInt n f \ + f (Suc n) \ newInt (Suc n) f" +proof (induct n) + case 0 + + let ?e = "SOME e. \e1 e2. + e1 < e2 \ + e = closed_int e1 e2 \ + e \ closed_int (f 0 + 1) (f 0 + 2) \ + f (Suc 0) \ e" + + have "newInt (Suc 0) f = ?e" by auto + moreover + have "f 0 + 1 < f 0 + 2" by simp + with closed_subset_ex have + "\ka kb. ka < kb \ closed_int ka kb \ closed_int (f 0 + 1) (f 0 + 2) \ + f (Suc 0) \ (closed_int ka kb)" . + hence + "\e. \ka kb. ka < kb \ e = closed_int ka kb \ + e \ closed_int (f 0 + 1) (f 0 + 2) \ f (Suc 0) \ e" by simp + hence + "\ka kb. ka < kb \ ?e = closed_int ka kb \ + ?e \ closed_int (f 0 + 1) (f 0 + 2) \ f (Suc 0) \ ?e" + by (rule someI_ex) + ultimately have "\e1 e2. e1 < e2 \ + newInt (Suc 0) f = closed_int e1 e2 \ + newInt (Suc 0) f \ closed_int (f 0 + 1) (f 0 + 2) \ + f (Suc 0) \ newInt (Suc 0) f" by simp + thus + "\a b. a < b \ newInt (Suc 0) f = closed_int a b \ + newInt (Suc 0) f \ newInt 0 f \ f (Suc 0) \ newInt (Suc 0) f" + by simp +next + case (Suc n) + hence "\a b. + a < b \ + newInt (Suc n) f = closed_int a b \ + newInt (Suc n) f \ newInt n f \ + f (Suc n) \ newInt (Suc n) f" by simp + then obtain a and b where ab: "a < b \ + newInt (Suc n) f = closed_int a b \ + newInt (Suc n) f \ newInt n f \ + f (Suc n) \ newInt (Suc n) f" by auto + hence cab: "closed_int a b = newInt (Suc n) f" by simp + + let ?e = "SOME e. \e1 e2. + e1 < e2 \ + e = closed_int e1 e2 \ + e \ closed_int a b \ + f (Suc (Suc n)) \ e" + from cab have ni: "newInt (Suc (Suc n)) f = ?e" by auto + + from ab have "a < b" by simp + with closed_subset_ex have + "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ + f (Suc (Suc n)) \ closed_int ka kb" . + hence + "\e. \ka kb. ka < kb \ e = closed_int ka kb \ + closed_int ka kb \ closed_int a b \ f (Suc (Suc n)) \ closed_int ka kb" + by simp + hence + "\e. \ka kb. ka < kb \ e = closed_int ka kb \ + e \ closed_int a b \ f (Suc (Suc n)) \ e" by simp + hence + "\ka kb. ka < kb \ ?e = closed_int ka kb \ + ?e \ closed_int a b \ f (Suc (Suc n)) \ ?e" by (rule someI_ex) + with ab ni show + "\ka kb. ka < kb \ + newInt (Suc (Suc n)) f = closed_int ka kb \ + newInt (Suc (Suc n)) f \ newInt (Suc n) f \ + f (Suc (Suc n)) \ newInt (Suc (Suc n)) f" by auto +qed + +lemma newInt_subset: + "newInt (Suc n) f \ newInt n f" + using newInt_ex by auto + + +text {* Another fundamental property is that no element in the range +of f is in the intersection of all closed intervals generated by +newInt. *} + +lemma newInt_inter: + "\n. f n \ (\n. newInt n f)" +proof + fix n::nat + { + assume n0: "n = 0" + moreover have "newInt 0 f = closed_int (f 0 + 1) (f 0 + 2)" by simp + ultimately have "f n \ newInt n f" by (unfold closed_int_def, simp) + } + moreover + { + assume "\ n = 0" + hence "n > 0" by simp + then obtain m where ndef: "n = Suc m" by (auto simp add: gr0_conv_Suc) + + from newInt_ex have + "\a b. a < b \ (newInt (Suc m) f) = closed_int a b \ + newInt (Suc m) f \ newInt m f \ f (Suc m) \ newInt (Suc m) f" . + then have "f (Suc m) \ newInt (Suc m) f" by auto + with ndef have "f n \ newInt n f" by simp + } + ultimately have "f n \ newInt n f" by (rule case_split) + thus "f n \ (\n. newInt n f)" by auto +qed + + +lemma newInt_notempty: + "(\n. newInt n f) \ {}" +proof - + let ?g = "\n. newInt n f" + have "\n. ?g (Suc n) \ ?g n" + proof + fix n + show "?g (Suc n) \ ?g n" by (rule newInt_subset) + qed + moreover have "\n. \a b. ?g n = closed_int a b \ a \ b" + proof + fix n::nat + { + assume "n = 0" + then have + "?g n = closed_int (f 0 + 1) (f 0 + 2) \ (f 0 + 1 \ f 0 + 2)" + by simp + hence "\a b. ?g n = closed_int a b \ a \ b" by blast + } + moreover + { + assume "\ n = 0" + then have "n > 0" by simp + then obtain m where nd: "n = Suc m" by (auto simp add: gr0_conv_Suc) + + have + "\a b. a < b \ (newInt (Suc m) f) = closed_int a b \ + (newInt (Suc m) f) \ (newInt m f) \ (f (Suc m)) \ (newInt (Suc m) f)" + by (rule newInt_ex) + then obtain a and b where + "a < b \ (newInt (Suc m) f) = closed_int a b" by auto + with nd have "?g n = closed_int a b \ a \ b" by auto + hence "\a b. ?g n = closed_int a b \ a \ b" by blast + } + ultimately show "\a b. ?g n = closed_int a b \ a \ b" by (rule case_split) + qed + ultimately show ?thesis by (rule NIP) +qed + + +subsection {* Final Theorem *} + +theorem real_non_denum: + shows "\ (\f::nat\real. surj f)" +proof -- "by contradiction" + assume "\f::nat\real. surj f" + then obtain f::"nat\real" where "surj f" by auto + hence rangeF: "range f = UNIV" by (rule surj_range) + -- "We now produce a real number x that is not in the range of f, using the properties of newInt. " + have "\x. x \ (\n. newInt n f)" using newInt_notempty by blast + moreover have "\n. f n \ (\n. newInt n f)" by (rule newInt_inter) + ultimately obtain x where "x \ (\n. newInt n f)" and "\n. f n \ x" by blast + moreover from rangeF have "x \ range f" by simp + ultimately show False by blast +qed + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Deriv.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Deriv.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,1725 @@ +(* Title : Deriv.thy + ID : $Id$ + Author : Jacques D. Fleuriot + Copyright : 1998 University of Cambridge + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 + GMVT by Benjamin Porter, 2005 +*) + +header{* Differentiation *} + +theory Deriv +imports Lim Univ_Poly +begin + +text{*Standard Definitions*} + +definition + deriv :: "['a::real_normed_field \ 'a, 'a, 'a] \ bool" + --{*Differentiation: D is derivative of function f at x*} + ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where + "DERIV f x :> D = ((%h. (f(x + h) - f x) / h) -- 0 --> D)" + +definition + differentiable :: "['a::real_normed_field \ 'a, 'a] \ bool" + (infixl "differentiable" 60) where + "f differentiable x = (\D. DERIV f x :> D)" + + +consts + Bolzano_bisect :: "[real*real=>bool, real, real, nat] => (real*real)" +primrec + "Bolzano_bisect P a b 0 = (a,b)" + "Bolzano_bisect P a b (Suc n) = + (let (x,y) = Bolzano_bisect P a b n + in if P(x, (x+y)/2) then ((x+y)/2, y) + else (x, (x+y)/2))" + + +subsection {* Derivatives *} + +lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --> D)" +by (simp add: deriv_def) + +lemma DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --> D" +by (simp add: deriv_def) + +lemma DERIV_const [simp]: "DERIV (\x. k) x :> 0" +by (simp add: deriv_def) + +lemma DERIV_ident [simp]: "DERIV (\x. x) x :> 1" +by (simp add: deriv_def cong: LIM_cong) + +lemma add_diff_add: + fixes a b c d :: "'a::ab_group_add" + shows "(a + c) - (b + d) = (a - b) + (c - d)" +by simp + +lemma DERIV_add: + "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x + g x) x :> D + E" +by (simp only: deriv_def add_diff_add add_divide_distrib LIM_add) + +lemma DERIV_minus: + "DERIV f x :> D \ DERIV (\x. - f x) x :> - D" +by (simp only: deriv_def minus_diff_minus divide_minus_left LIM_minus) + +lemma DERIV_diff: + "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x - g x) x :> D - E" +by (simp only: diff_def DERIV_add DERIV_minus) + +lemma DERIV_add_minus: + "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x + - g x) x :> D + - E" +by (simp only: DERIV_add DERIV_minus) + +lemma DERIV_isCont: "DERIV f x :> D \ isCont f x" +proof (unfold isCont_iff) + assume "DERIV f x :> D" + hence "(\h. (f(x+h) - f(x)) / h) -- 0 --> D" + by (rule DERIV_D) + hence "(\h. (f(x+h) - f(x)) / h * h) -- 0 --> D * 0" + by (intro LIM_mult LIM_ident) + hence "(\h. (f(x+h) - f(x)) * (h / h)) -- 0 --> 0" + by simp + hence "(\h. f(x+h) - f(x)) -- 0 --> 0" + by (simp cong: LIM_cong) + thus "(\h. f(x+h)) -- 0 --> f(x)" + by (simp add: LIM_def) +qed + +lemma DERIV_mult_lemma: + fixes a b c d :: "'a::real_field" + shows "(a * b - c * d) / h = a * ((b - d) / h) + ((a - c) / h) * d" +by (simp add: diff_minus add_divide_distrib [symmetric] ring_distribs) + +lemma DERIV_mult': + assumes f: "DERIV f x :> D" + assumes g: "DERIV g x :> E" + shows "DERIV (\x. f x * g x) x :> f x * E + D * g x" +proof (unfold deriv_def) + from f have "isCont f x" + by (rule DERIV_isCont) + hence "(\h. f(x+h)) -- 0 --> f x" + by (simp only: isCont_iff) + hence "(\h. f(x+h) * ((g(x+h) - g x) / h) + + ((f(x+h) - f x) / h) * g x) + -- 0 --> f x * E + D * g x" + by (intro LIM_add LIM_mult LIM_const DERIV_D f g) + thus "(\h. (f(x+h) * g(x+h) - f x * g x) / h) + -- 0 --> f x * E + D * g x" + by (simp only: DERIV_mult_lemma) +qed + +lemma DERIV_mult: + "[| DERIV f x :> Da; DERIV g x :> Db |] + ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" +by (drule (1) DERIV_mult', simp only: mult_commute add_commute) + +lemma DERIV_unique: + "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E" +apply (simp add: deriv_def) +apply (blast intro: LIM_unique) +done + +text{*Differentiation of finite sum*} + +lemma DERIV_sumr [rule_format (no_asm)]: + "(\r. m \ r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x)) + --> DERIV (%x. \n=m.. (\r=m.. D) = + ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)" +apply (rule iffI) +apply (drule_tac k="- a" in LIM_offset) +apply (simp add: diff_minus) +apply (drule_tac k="a" in LIM_offset) +apply (simp add: add_commute) +done + +lemma DERIV_iff2: "(DERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --> D)" +by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff) + +lemma inverse_diff_inverse: + "\(a::'a::division_ring) \ 0; b \ 0\ + \ inverse a - inverse b = - (inverse a * (a - b) * inverse b)" +by (simp add: ring_simps) + +lemma DERIV_inverse_lemma: + "\a \ 0; b \ (0::'a::real_normed_field)\ + \ (inverse a - inverse b) / h + = - (inverse a * ((a - b) / h) * inverse b)" +by (simp add: inverse_diff_inverse) + +lemma DERIV_inverse': + assumes der: "DERIV f x :> D" + assumes neq: "f x \ 0" + shows "DERIV (\x. inverse (f x)) x :> - (inverse (f x) * D * inverse (f x))" + (is "DERIV _ _ :> ?E") +proof (unfold DERIV_iff2) + from der have lim_f: "f -- x --> f x" + by (rule DERIV_isCont [unfolded isCont_def]) + + from neq have "0 < norm (f x)" by simp + with LIM_D [OF lim_f] obtain s + where s: "0 < s" + and less_fx: "\z. \z \ x; norm (z - x) < s\ + \ norm (f z - f x) < norm (f x)" + by fast + + show "(\z. (inverse (f z) - inverse (f x)) / (z - x)) -- x --> ?E" + proof (rule LIM_equal2 [OF s]) + fix z + assume "z \ x" "norm (z - x) < s" + hence "norm (f z - f x) < norm (f x)" by (rule less_fx) + hence "f z \ 0" by auto + thus "(inverse (f z) - inverse (f x)) / (z - x) = + - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x))" + using neq by (rule DERIV_inverse_lemma) + next + from der have "(\z. (f z - f x) / (z - x)) -- x --> D" + by (unfold DERIV_iff2) + thus "(\z. - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x))) + -- x --> ?E" + by (intro LIM_mult LIM_inverse LIM_minus LIM_const lim_f neq) + qed +qed + +lemma DERIV_divide: + "\DERIV f x :> D; DERIV g x :> E; g x \ 0\ + \ DERIV (\x. f x / g x) x :> (D * g x - f x * E) / (g x * g x)" +apply (subgoal_tac "f x * - (inverse (g x) * E * inverse (g x)) + + D * inverse (g x) = (D * g x - f x * E) / (g x * g x)") +apply (erule subst) +apply (unfold divide_inverse) +apply (erule DERIV_mult') +apply (erule (1) DERIV_inverse') +apply (simp add: ring_distribs nonzero_inverse_mult_distrib) +apply (simp add: mult_ac) +done + +lemma DERIV_power_Suc: + fixes f :: "'a \ 'a::{real_normed_field,recpower}" + assumes f: "DERIV f x :> D" + shows "DERIV (\x. f x ^ Suc n) x :> (1 + of_nat n) * (D * f x ^ n)" +proof (induct n) +case 0 + show ?case by (simp add: power_Suc f) +case (Suc k) + from DERIV_mult' [OF f Suc] show ?case + apply (simp only: of_nat_Suc ring_distribs mult_1_left) + apply (simp only: power_Suc right_distrib mult_ac add_ac) + done +qed + +lemma DERIV_power: + fixes f :: "'a \ 'a::{real_normed_field,recpower}" + assumes f: "DERIV f x :> D" + shows "DERIV (\x. f x ^ n) x :> of_nat n * (D * f x ^ (n - Suc 0))" +by (cases "n", simp, simp add: DERIV_power_Suc f) + + +(* ------------------------------------------------------------------------ *) +(* Caratheodory formulation of derivative at a point: standard proof *) +(* ------------------------------------------------------------------------ *) + +lemma CARAT_DERIV: + "(DERIV f x :> l) = + (\g. (\z. f z - f x = g z * (z-x)) & isCont g x & g x = l)" + (is "?lhs = ?rhs") +proof + assume der: "DERIV f x :> l" + show "\g. (\z. f z - f x = g z * (z-x)) \ isCont g x \ g x = l" + proof (intro exI conjI) + let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" + show "\z. f z - f x = ?g z * (z-x)" by simp + show "isCont ?g x" using der + by (simp add: isCont_iff DERIV_iff diff_minus + cong: LIM_equal [rule_format]) + show "?g x = l" by simp + qed +next + assume "?rhs" + then obtain g where + "(\z. f z - f x = g z * (z-x))" and "isCont g x" and "g x = l" by blast + thus "(DERIV f x :> l)" + by (auto simp add: isCont_iff DERIV_iff cong: LIM_cong) +qed + +lemma DERIV_chain': + assumes f: "DERIV f x :> D" + assumes g: "DERIV g (f x) :> E" + shows "DERIV (\x. g (f x)) x :> E * D" +proof (unfold DERIV_iff2) + obtain d where d: "\y. g y - g (f x) = d y * (y - f x)" + and cont_d: "isCont d (f x)" and dfx: "d (f x) = E" + using CARAT_DERIV [THEN iffD1, OF g] by fast + from f have "f -- x --> f x" + by (rule DERIV_isCont [unfolded isCont_def]) + with cont_d have "(\z. d (f z)) -- x --> d (f x)" + by (rule isCont_LIM_compose) + hence "(\z. d (f z) * ((f z - f x) / (z - x))) + -- x --> d (f x) * D" + by (rule LIM_mult [OF _ f [unfolded DERIV_iff2]]) + thus "(\z. (g (f z) - g (f x)) / (z - x)) -- x --> E * D" + by (simp add: d dfx real_scaleR_def) +qed + +(* let's do the standard proof though theorem *) +(* LIM_mult2 follows from a NS proof *) + +lemma DERIV_cmult: + "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D" +by (drule DERIV_mult' [OF DERIV_const], simp) + +(* standard version *) +lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db" +by (drule (1) DERIV_chain', simp add: o_def real_scaleR_def mult_commute) + +lemma DERIV_chain2: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db" +by (auto dest: DERIV_chain simp add: o_def) + +(*derivative of linear multiplication*) +lemma DERIV_cmult_Id [simp]: "DERIV (op * c) x :> c" +by (cut_tac c = c and x = x in DERIV_ident [THEN DERIV_cmult], simp) + +lemma DERIV_pow: "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" +apply (cut_tac DERIV_power [OF DERIV_ident]) +apply (simp add: real_scaleR_def real_of_nat_def) +done + +text{*Power of -1*} + +lemma DERIV_inverse: + fixes x :: "'a::{real_normed_field,recpower}" + shows "x \ 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))" +by (drule DERIV_inverse' [OF DERIV_ident]) (simp add: power_Suc) + +text{*Derivative of inverse*} +lemma DERIV_inverse_fun: + fixes x :: "'a::{real_normed_field,recpower}" + shows "[| DERIV f x :> d; f(x) \ 0 |] + ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" +by (drule (1) DERIV_inverse') (simp add: mult_ac power_Suc nonzero_inverse_mult_distrib) + +text{*Derivative of quotient*} +lemma DERIV_quotient: + fixes x :: "'a::{real_normed_field,recpower}" + shows "[| DERIV f x :> d; DERIV g x :> e; g(x) \ 0 |] + ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ Suc (Suc 0))" +by (drule (2) DERIV_divide) (simp add: mult_commute power_Suc) + + +subsection {* Differentiability predicate *} + +lemma differentiableD: "f differentiable x ==> \D. DERIV f x :> D" +by (simp add: differentiable_def) + +lemma differentiableI: "DERIV f x :> D ==> f differentiable x" +by (force simp add: differentiable_def) + +lemma differentiable_const: "(\z. a) differentiable x" + apply (unfold differentiable_def) + apply (rule_tac x=0 in exI) + apply simp + done + +lemma differentiable_sum: + assumes "f differentiable x" + and "g differentiable x" + shows "(\x. f x + g x) differentiable x" +proof - + from prems have "\D. DERIV f x :> D" by (unfold differentiable_def) + then obtain df where "DERIV f x :> df" .. + moreover from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) + then obtain dg where "DERIV g x :> dg" .. + ultimately have "DERIV (\x. f x + g x) x :> df + dg" by (rule DERIV_add) + hence "\D. DERIV (\x. f x + g x) x :> D" by auto + thus ?thesis by (fold differentiable_def) +qed + +lemma differentiable_diff: + assumes "f differentiable x" + and "g differentiable x" + shows "(\x. f x - g x) differentiable x" +proof - + from prems have "f differentiable x" by simp + moreover + from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) + then obtain dg where "DERIV g x :> dg" .. + then have "DERIV (\x. - g x) x :> -dg" by (rule DERIV_minus) + hence "\D. DERIV (\x. - g x) x :> D" by auto + hence "(\x. - g x) differentiable x" by (fold differentiable_def) + ultimately + show ?thesis + by (auto simp: diff_def dest: differentiable_sum) +qed + +lemma differentiable_mult: + assumes "f differentiable x" + and "g differentiable x" + shows "(\x. f x * g x) differentiable x" +proof - + from prems have "\D. DERIV f x :> D" by (unfold differentiable_def) + then obtain df where "DERIV f x :> df" .. + moreover from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) + then obtain dg where "DERIV g x :> dg" .. + ultimately have "DERIV (\x. f x * g x) x :> df * g x + dg * f x" by (simp add: DERIV_mult) + hence "\D. DERIV (\x. f x * g x) x :> D" by auto + thus ?thesis by (fold differentiable_def) +qed + + +subsection {* Nested Intervals and Bisection *} + +text{*Lemmas about nested intervals and proof by bisection (cf.Harrison). + All considerably tidied by lcp.*} + +lemma lemma_f_mono_add [rule_format (no_asm)]: "(\n. (f::nat=>real) n \ f (Suc n)) --> f m \ f(m + no)" +apply (induct "no") +apply (auto intro: order_trans) +done + +lemma f_inc_g_dec_Beq_f: "[| \n. f(n) \ f(Suc n); + \n. g(Suc n) \ g(n); + \n. f(n) \ g(n) |] + ==> Bseq (f :: nat \ real)" +apply (rule_tac k = "f 0" and K = "g 0" in BseqI2, rule allI) +apply (induct_tac "n") +apply (auto intro: order_trans) +apply (rule_tac y = "g (Suc na)" in order_trans) +apply (induct_tac [2] "na") +apply (auto intro: order_trans) +done + +lemma f_inc_g_dec_Beq_g: "[| \n. f(n) \ f(Suc n); + \n. g(Suc n) \ g(n); + \n. f(n) \ g(n) |] + ==> Bseq (g :: nat \ real)" +apply (subst Bseq_minus_iff [symmetric]) +apply (rule_tac g = "%x. - (f x)" in f_inc_g_dec_Beq_f) +apply auto +done + +lemma f_inc_imp_le_lim: + fixes f :: "nat \ real" + shows "\\n. f n \ f (Suc n); convergent f\ \ f n \ lim f" +apply (rule linorder_not_less [THEN iffD1]) +apply (auto simp add: convergent_LIMSEQ_iff LIMSEQ_iff monoseq_Suc) +apply (drule real_less_sum_gt_zero) +apply (drule_tac x = "f n + - lim f" in spec, safe) +apply (drule_tac P = "%na. no\na --> ?Q na" and x = "no + n" in spec, auto) +apply (subgoal_tac "lim f \ f (no + n) ") +apply (drule_tac no=no and m=n in lemma_f_mono_add) +apply (auto simp add: add_commute) +apply (induct_tac "no") +apply simp +apply (auto intro: order_trans simp add: diff_minus abs_if) +done + +lemma lim_uminus: "convergent g ==> lim (%x. - g x) = - (lim g)" +apply (rule LIMSEQ_minus [THEN limI]) +apply (simp add: convergent_LIMSEQ_iff) +done + +lemma g_dec_imp_lim_le: + fixes g :: "nat \ real" + shows "\\n. g (Suc n) \ g(n); convergent g\ \ lim g \ g n" +apply (subgoal_tac "- (g n) \ - (lim g) ") +apply (cut_tac [2] f = "%x. - (g x)" in f_inc_imp_le_lim) +apply (auto simp add: lim_uminus convergent_minus_iff [symmetric]) +done + +lemma lemma_nest: "[| \n. f(n) \ f(Suc n); + \n. g(Suc n) \ g(n); + \n. f(n) \ g(n) |] + ==> \l m :: real. l \ m & ((\n. f(n) \ l) & f ----> l) & + ((\n. m \ g(n)) & g ----> m)" +apply (subgoal_tac "monoseq f & monoseq g") +prefer 2 apply (force simp add: LIMSEQ_iff monoseq_Suc) +apply (subgoal_tac "Bseq f & Bseq g") +prefer 2 apply (blast intro: f_inc_g_dec_Beq_f f_inc_g_dec_Beq_g) +apply (auto dest!: Bseq_monoseq_convergent simp add: convergent_LIMSEQ_iff) +apply (rule_tac x = "lim f" in exI) +apply (rule_tac x = "lim g" in exI) +apply (auto intro: LIMSEQ_le) +apply (auto simp add: f_inc_imp_le_lim g_dec_imp_lim_le convergent_LIMSEQ_iff) +done + +lemma lemma_nest_unique: "[| \n. f(n) \ f(Suc n); + \n. g(Suc n) \ g(n); + \n. f(n) \ g(n); + (%n. f(n) - g(n)) ----> 0 |] + ==> \l::real. ((\n. f(n) \ l) & f ----> l) & + ((\n. l \ g(n)) & g ----> l)" +apply (drule lemma_nest, auto) +apply (subgoal_tac "l = m") +apply (drule_tac [2] X = f in LIMSEQ_diff) +apply (auto intro: LIMSEQ_unique) +done + +text{*The universal quantifiers below are required for the declaration + of @{text Bolzano_nest_unique} below.*} + +lemma Bolzano_bisect_le: + "a \ b ==> \n. fst (Bolzano_bisect P a b n) \ snd (Bolzano_bisect P a b n)" +apply (rule allI) +apply (induct_tac "n") +apply (auto simp add: Let_def split_def) +done + +lemma Bolzano_bisect_fst_le_Suc: "a \ b ==> + \n. fst(Bolzano_bisect P a b n) \ fst (Bolzano_bisect P a b (Suc n))" +apply (rule allI) +apply (induct_tac "n") +apply (auto simp add: Bolzano_bisect_le Let_def split_def) +done + +lemma Bolzano_bisect_Suc_le_snd: "a \ b ==> + \n. snd(Bolzano_bisect P a b (Suc n)) \ snd (Bolzano_bisect P a b n)" +apply (rule allI) +apply (induct_tac "n") +apply (auto simp add: Bolzano_bisect_le Let_def split_def) +done + +lemma eq_divide_2_times_iff: "((x::real) = y / (2 * z)) = (2 * x = y/z)" +apply (auto) +apply (drule_tac f = "%u. (1/2) *u" in arg_cong) +apply (simp) +done + +lemma Bolzano_bisect_diff: + "a \ b ==> + snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = + (b-a) / (2 ^ n)" +apply (induct "n") +apply (auto simp add: eq_divide_2_times_iff add_divide_distrib Let_def split_def) +done + +lemmas Bolzano_nest_unique = + lemma_nest_unique + [OF Bolzano_bisect_fst_le_Suc Bolzano_bisect_Suc_le_snd Bolzano_bisect_le] + + +lemma not_P_Bolzano_bisect: + assumes P: "!!a b c. [| P(a,b); P(b,c); a \ b; b \ c|] ==> P(a,c)" + and notP: "~ P(a,b)" + and le: "a \ b" + shows "~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" +proof (induct n) + case 0 show ?case using notP by simp + next + case (Suc n) + thus ?case + by (auto simp del: surjective_pairing [symmetric] + simp add: Let_def split_def Bolzano_bisect_le [OF le] + P [of "fst (Bolzano_bisect P a b n)" _ "snd (Bolzano_bisect P a b n)"]) +qed + +(*Now we re-package P_prem as a formula*) +lemma not_P_Bolzano_bisect': + "[| \a b c. P(a,b) & P(b,c) & a \ b & b \ c --> P(a,c); + ~ P(a,b); a \ b |] ==> + \n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" +by (blast elim!: not_P_Bolzano_bisect [THEN [2] rev_notE]) + + + +lemma lemma_BOLZANO: + "[| \a b c. P(a,b) & P(b,c) & a \ b & b \ c --> P(a,c); + \x. \d::real. 0 < d & + (\a b. a \ x & x \ b & (b-a) < d --> P(a,b)); + a \ b |] + ==> P(a,b)" +apply (rule Bolzano_nest_unique [where P1=P, THEN exE], assumption+) +apply (rule LIMSEQ_minus_cancel) +apply (simp (no_asm_simp) add: Bolzano_bisect_diff LIMSEQ_divide_realpow_zero) +apply (rule ccontr) +apply (drule not_P_Bolzano_bisect', assumption+) +apply (rename_tac "l") +apply (drule_tac x = l in spec, clarify) +apply (simp add: LIMSEQ_def) +apply (drule_tac P = "%r. 0 ?Q r" and x = "d/2" in spec) +apply (drule_tac P = "%r. 0 ?Q r" and x = "d/2" in spec) +apply (drule real_less_half_sum, auto) +apply (drule_tac x = "fst (Bolzano_bisect P a b (no + noa))" in spec) +apply (drule_tac x = "snd (Bolzano_bisect P a b (no + noa))" in spec) +apply safe +apply (simp_all (no_asm_simp)) +apply (rule_tac y = "abs (fst (Bolzano_bisect P a b (no + noa)) - l) + abs (snd (Bolzano_bisect P a b (no + noa)) - l)" in order_le_less_trans) +apply (simp (no_asm_simp) add: abs_if) +apply (rule real_sum_of_halves [THEN subst]) +apply (rule add_strict_mono) +apply (simp_all add: diff_minus [symmetric]) +done + + +lemma lemma_BOLZANO2: "((\a b c. (a \ b & b \ c & P(a,b) & P(b,c)) --> P(a,c)) & + (\x. \d::real. 0 < d & + (\a b. a \ x & x \ b & (b-a) < d --> P(a,b)))) + --> (\a b. a \ b --> P(a,b))" +apply clarify +apply (blast intro: lemma_BOLZANO) +done + + +subsection {* Intermediate Value Theorem *} + +text {*Prove Contrapositive by Bisection*} + +lemma IVT: "[| f(a::real) \ (y::real); y \ f(b); + a \ b; + (\x. a \ x & x \ b --> isCont f x) |] + ==> \x. a \ x & x \ b & f(x) = y" +apply (rule contrapos_pp, assumption) +apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> ~ (f (u) \ y & y \ f (v))" in lemma_BOLZANO2) +apply safe +apply simp_all +apply (simp add: isCont_iff LIM_def) +apply (rule ccontr) +apply (subgoal_tac "a \ x & x \ b") + prefer 2 + apply simp + apply (drule_tac P = "%d. 0 ?P d" and x = 1 in spec, arith) +apply (drule_tac x = x in spec)+ +apply simp +apply (drule_tac P = "%r. ?P r --> (\s>0. ?Q r s) " and x = "\y - f x\" in spec) +apply safe +apply simp +apply (drule_tac x = s in spec, clarify) +apply (cut_tac x = "f x" and y = y in linorder_less_linear, safe) +apply (drule_tac x = "ba-x" in spec) +apply (simp_all add: abs_if) +apply (drule_tac x = "aa-x" in spec) +apply (case_tac "x \ aa", simp_all) +done + +lemma IVT2: "[| f(b::real) \ (y::real); y \ f(a); + a \ b; + (\x. a \ x & x \ b --> isCont f x) + |] ==> \x. a \ x & x \ b & f(x) = y" +apply (subgoal_tac "- f a \ -y & -y \ - f b", clarify) +apply (drule IVT [where f = "%x. - f x"], assumption) +apply (auto intro: isCont_minus) +done + +(*HOL style here: object-level formulations*) +lemma IVT_objl: "(f(a::real) \ (y::real) & y \ f(b) & a \ b & + (\x. a \ x & x \ b --> isCont f x)) + --> (\x. a \ x & x \ b & f(x) = y)" +apply (blast intro: IVT) +done + +lemma IVT2_objl: "(f(b::real) \ (y::real) & y \ f(a) & a \ b & + (\x. a \ x & x \ b --> isCont f x)) + --> (\x. a \ x & x \ b & f(x) = y)" +apply (blast intro: IVT2) +done + +text{*By bisection, function continuous on closed interval is bounded above*} + +lemma isCont_bounded: + "[| a \ b; \x. a \ x & x \ b --> isCont f x |] + ==> \M::real. \x::real. a \ x & x \ b --> f(x) \ M" +apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> (\M. \x. u \ x & x \ v --> f x \ M)" in lemma_BOLZANO2) +apply safe +apply simp_all +apply (rename_tac x xa ya M Ma) +apply (cut_tac x = M and y = Ma in linorder_linear, safe) +apply (rule_tac x = Ma in exI, clarify) +apply (cut_tac x = xb and y = xa in linorder_linear, force) +apply (rule_tac x = M in exI, clarify) +apply (cut_tac x = xb and y = xa in linorder_linear, force) +apply (case_tac "a \ x & x \ b") +apply (rule_tac [2] x = 1 in exI) +prefer 2 apply force +apply (simp add: LIM_def isCont_iff) +apply (drule_tac x = x in spec, auto) +apply (erule_tac V = "\M. \x. a \ x & x \ b & ~ f x \ M" in thin_rl) +apply (drule_tac x = 1 in spec, auto) +apply (rule_tac x = s in exI, clarify) +apply (rule_tac x = "\f x\ + 1" in exI, clarify) +apply (drule_tac x = "xa-x" in spec) +apply (auto simp add: abs_ge_self) +done + +text{*Refine the above to existence of least upper bound*} + +lemma lemma_reals_complete: "((\x. x \ S) & (\y. isUb UNIV S (y::real))) --> + (\t. isLub UNIV S t)" +by (blast intro: reals_complete) + +lemma isCont_has_Ub: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] + ==> \M::real. (\x::real. a \ x & x \ b --> f(x) \ M) & + (\N. N < M --> (\x. a \ x & x \ b & N < f(x)))" +apply (cut_tac S = "Collect (%y. \x. a \ x & x \ b & y = f x)" + in lemma_reals_complete) +apply auto +apply (drule isCont_bounded, assumption) +apply (auto simp add: isUb_def leastP_def isLub_def setge_def setle_def) +apply (rule exI, auto) +apply (auto dest!: spec simp add: linorder_not_less) +done + +text{*Now show that it attains its upper bound*} + +lemma isCont_eq_Ub: + assumes le: "a \ b" + and con: "\x::real. a \ x & x \ b --> isCont f x" + shows "\M::real. (\x. a \ x & x \ b --> f(x) \ M) & + (\x. a \ x & x \ b & f(x) = M)" +proof - + from isCont_has_Ub [OF le con] + obtain M where M1: "\x. a \ x \ x \ b \ f x \ M" + and M2: "!!N. N \x. a \ x \ x \ b \ N < f x" by blast + show ?thesis + proof (intro exI, intro conjI) + show " \x. a \ x \ x \ b \ f x \ M" by (rule M1) + show "\x. a \ x \ x \ b \ f x = M" + proof (rule ccontr) + assume "\ (\x. a \ x \ x \ b \ f x = M)" + with M1 have M3: "\x. a \ x & x \ b --> f x < M" + by (fastsimp simp add: linorder_not_le [symmetric]) + hence "\x. a \ x & x \ b --> isCont (%x. inverse (M - f x)) x" + by (auto simp add: isCont_inverse isCont_diff con) + from isCont_bounded [OF le this] + obtain k where k: "!!x. a \ x & x \ b --> inverse (M - f x) \ k" by auto + have Minv: "!!x. a \ x & x \ b --> 0 < inverse (M - f (x))" + by (simp add: M3 compare_rls) + have "!!x. a \ x & x \ b --> inverse (M - f x) < k+1" using k + by (auto intro: order_le_less_trans [of _ k]) + with Minv + have "!!x. a \ x & x \ b --> inverse(k+1) < inverse(inverse(M - f x))" + by (intro strip less_imp_inverse_less, simp_all) + hence invlt: "!!x. a \ x & x \ b --> inverse(k+1) < M - f x" + by simp + have "M - inverse (k+1) < M" using k [of a] Minv [of a] le + by (simp, arith) + from M2 [OF this] + obtain x where ax: "a \ x & x \ b & M - inverse(k+1) < f x" .. + thus False using invlt [of x] by force + qed + qed +qed + + +text{*Same theorem for lower bound*} + +lemma isCont_eq_Lb: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] + ==> \M::real. (\x::real. a \ x & x \ b --> M \ f(x)) & + (\x. a \ x & x \ b & f(x) = M)" +apply (subgoal_tac "\x. a \ x & x \ b --> isCont (%x. - (f x)) x") +prefer 2 apply (blast intro: isCont_minus) +apply (drule_tac f = "(%x. - (f x))" in isCont_eq_Ub) +apply safe +apply auto +done + + +text{*Another version.*} + +lemma isCont_Lb_Ub: "[|a \ b; \x. a \ x & x \ b --> isCont f x |] + ==> \L M::real. (\x::real. a \ x & x \ b --> L \ f(x) & f(x) \ M) & + (\y. L \ y & y \ M --> (\x. a \ x & x \ b & (f(x) = y)))" +apply (frule isCont_eq_Lb) +apply (frule_tac [2] isCont_eq_Ub) +apply (assumption+, safe) +apply (rule_tac x = "f x" in exI) +apply (rule_tac x = "f xa" in exI, simp, safe) +apply (cut_tac x = x and y = xa in linorder_linear, safe) +apply (cut_tac f = f and a = x and b = xa and y = y in IVT_objl) +apply (cut_tac [2] f = f and a = xa and b = x and y = y in IVT2_objl, safe) +apply (rule_tac [2] x = xb in exI) +apply (rule_tac [4] x = xb in exI, simp_all) +done + + +text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*} + +lemma DERIV_left_inc: + fixes f :: "real => real" + assumes der: "DERIV f x :> l" + and l: "0 < l" + shows "\d > 0. \h > 0. h < d --> f(x) < f(x + h)" +proof - + from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]] + have "\s > 0. (\z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < l)" + by (simp add: diff_minus) + then obtain s + where s: "0 < s" + and all: "!!z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < l" + by auto + thus ?thesis + proof (intro exI conjI strip) + show "0 real" + assumes der: "DERIV f x :> l" + and l: "l < 0" + shows "\d > 0. \h > 0. h < d --> f(x) < f(x-h)" +proof - + from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]] + have "\s > 0. (\z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < -l)" + by (simp add: diff_minus) + then obtain s + where s: "0 < s" + and all: "!!z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < -l" + by auto + thus ?thesis + proof (intro exI conjI strip) + show "0 real" + assumes der: "DERIV f x :> l" + and d: "0 < d" + and le: "\y. \x-y\ < d --> f(y) \ f(x)" + shows "l = 0" +proof (cases rule: linorder_cases [of l 0]) + case equal thus ?thesis . +next + case less + from DERIV_left_dec [OF der less] + obtain d' where d': "0 < d'" + and lt: "\h > 0. h < d' \ f x < f (x-h)" by blast + from real_lbound_gt_zero [OF d d'] + obtain e where "0 < e \ e < d \ e < d'" .. + with lt le [THEN spec [where x="x-e"]] + show ?thesis by (auto simp add: abs_if) +next + case greater + from DERIV_left_inc [OF der greater] + obtain d' where d': "0 < d'" + and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast + from real_lbound_gt_zero [OF d d'] + obtain e where "0 < e \ e < d \ e < d'" .. + with lt le [THEN spec [where x="x+e"]] + show ?thesis by (auto simp add: abs_if) +qed + + +text{*Similar theorem for a local minimum*} +lemma DERIV_local_min: + fixes f :: "real => real" + shows "[| DERIV f x :> l; 0 < d; \y. \x-y\ < d --> f(x) \ f(y) |] ==> l = 0" +by (drule DERIV_minus [THEN DERIV_local_max], auto) + + +text{*In particular, if a function is locally flat*} +lemma DERIV_local_const: + fixes f :: "real => real" + shows "[| DERIV f x :> l; 0 < d; \y. \x-y\ < d --> f(x) = f(y) |] ==> l = 0" +by (auto dest!: DERIV_local_max) + +text{*Lemma about introducing open ball in open interval*} +lemma lemma_interval_lt: + "[| a < x; x < b |] + ==> \d::real. 0 < d & (\y. \x-y\ < d --> a < y & y < b)" + +apply (simp add: abs_less_iff) +apply (insert linorder_linear [of "x-a" "b-x"], safe) +apply (rule_tac x = "x-a" in exI) +apply (rule_tac [2] x = "b-x" in exI, auto) +done + +lemma lemma_interval: "[| a < x; x < b |] ==> + \d::real. 0 < d & (\y. \x-y\ < d --> a \ y & y \ b)" +apply (drule lemma_interval_lt, auto) +apply (auto intro!: exI) +done + +text{*Rolle's Theorem. + If @{term f} is defined and continuous on the closed interval + @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"}, + and @{term "f(a) = f(b)"}, + then there exists @{text "x0 \ (a,b)"} such that @{term "f'(x0) = 0"}*} +theorem Rolle: + assumes lt: "a < b" + and eq: "f(a) = f(b)" + and con: "\x. a \ x & x \ b --> isCont f x" + and dif [rule_format]: "\x. a < x & x < b --> f differentiable x" + shows "\z::real. a < z & z < b & DERIV f z :> 0" +proof - + have le: "a \ b" using lt by simp + from isCont_eq_Ub [OF le con] + obtain x where x_max: "\z. a \ z \ z \ b \ f z \ f x" + and alex: "a \ x" and xleb: "x \ b" + by blast + from isCont_eq_Lb [OF le con] + obtain x' where x'_min: "\z. a \ z \ z \ b \ f x' \ f z" + and alex': "a \ x'" and x'leb: "x' \ b" + by blast + show ?thesis + proof cases + assume axb: "a < x & x < b" + --{*@{term f} attains its maximum within the interval*} + hence ax: "ay. \x-y\ < d \ a \ y \ y \ b" + by blast + hence bound': "\y. \x-y\ < d \ f y \ f x" using x_max + by blast + from differentiableD [OF dif [OF axb]] + obtain l where der: "DERIV f x :> l" .. + have "l=0" by (rule DERIV_local_max [OF der d bound']) + --{*the derivative at a local maximum is zero*} + thus ?thesis using ax xb der by auto + next + assume notaxb: "~ (a < x & x < b)" + hence xeqab: "x=a | x=b" using alex xleb by arith + hence fb_eq_fx: "f b = f x" by (auto simp add: eq) + show ?thesis + proof cases + assume ax'b: "a < x' & x' < b" + --{*@{term f} attains its minimum within the interval*} + hence ax': "ay. \x'-y\ < d \ a \ y \ y \ b" + by blast + hence bound': "\y. \x'-y\ < d \ f x' \ f y" using x'_min + by blast + from differentiableD [OF dif [OF ax'b]] + obtain l where der: "DERIV f x' :> l" .. + have "l=0" by (rule DERIV_local_min [OF der d bound']) + --{*the derivative at a local minimum is zero*} + thus ?thesis using ax' x'b der by auto + next + assume notax'b: "~ (a < x' & x' < b)" + --{*@{term f} is constant througout the interval*} + hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith + hence fb_eq_fx': "f b = f x'" by (auto simp add: eq) + from dense [OF lt] + obtain r where ar: "a < r" and rb: "r < b" by blast + from lemma_interval [OF ar rb] + obtain d where d: "0y. \r-y\ < d \ a \ y \ y \ b" + by blast + have eq_fb: "\z. a \ z --> z \ b --> f z = f b" + proof (clarify) + fix z::real + assume az: "a \ z" and zb: "z \ b" + show "f z = f b" + proof (rule order_antisym) + show "f z \ f b" by (simp add: fb_eq_fx x_max az zb) + show "f b \ f z" by (simp add: fb_eq_fx' x'_min az zb) + qed + qed + have bound': "\y. \r-y\ < d \ f r = f y" + proof (intro strip) + fix y::real + assume lt: "\r-y\ < d" + hence "f y = f b" by (simp add: eq_fb bound) + thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le) + qed + from differentiableD [OF dif [OF conjI [OF ar rb]]] + obtain l where der: "DERIV f r :> l" .. + have "l=0" by (rule DERIV_local_const [OF der d bound']) + --{*the derivative of a constant function is zero*} + thus ?thesis using ar rb der by auto + qed + qed +qed + + +subsection{*Mean Value Theorem*} + +lemma lemma_MVT: + "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)" +proof cases + assume "a=b" thus ?thesis by simp +next + assume "a\b" + hence ba: "b-a \ 0" by arith + show ?thesis + by (rule real_mult_left_cancel [OF ba, THEN iffD1], + simp add: right_diff_distrib, + simp add: left_diff_distrib) +qed + +theorem MVT: + assumes lt: "a < b" + and con: "\x. a \ x & x \ b --> isCont f x" + and dif [rule_format]: "\x. a < x & x < b --> f differentiable x" + shows "\l z::real. a < z & z < b & DERIV f z :> l & + (f(b) - f(a) = (b-a) * l)" +proof - + let ?F = "%x. f x - ((f b - f a) / (b-a)) * x" + have contF: "\x. a \ x \ x \ b \ isCont ?F x" using con + by (fast intro: isCont_diff isCont_const isCont_mult isCont_ident) + have difF: "\x. a < x \ x < b \ ?F differentiable x" + proof (clarify) + fix x::real + assume ax: "a < x" and xb: "x < b" + from differentiableD [OF dif [OF conjI [OF ax xb]]] + obtain l where der: "DERIV f x :> l" .. + show "?F differentiable x" + by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"], + blast intro: DERIV_diff DERIV_cmult_Id der) + qed + from Rolle [where f = ?F, OF lt lemma_MVT contF difF] + obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0" + by blast + have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)" + by (rule DERIV_cmult_Id) + hence derF: "DERIV (\x. ?F x + (f b - f a) / (b - a) * x) z + :> 0 + (f b - f a) / (b - a)" + by (rule DERIV_add [OF der]) + show ?thesis + proof (intro exI conjI) + show "a < z" using az . + show "z < b" using zb . + show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by (simp) + show "DERIV f z :> ((f b - f a)/(b-a))" using derF by simp + qed +qed + + +text{*A function is constant if its derivative is 0 over an interval.*} + +lemma DERIV_isconst_end: + fixes f :: "real => real" + shows "[| a < b; + \x. a \ x & x \ b --> isCont f x; + \x. a < x & x < b --> DERIV f x :> 0 |] + ==> f b = f a" +apply (drule MVT, assumption) +apply (blast intro: differentiableI) +apply (auto dest!: DERIV_unique simp add: diff_eq_eq) +done + +lemma DERIV_isconst1: + fixes f :: "real => real" + shows "[| a < b; + \x. a \ x & x \ b --> isCont f x; + \x. a < x & x < b --> DERIV f x :> 0 |] + ==> \x. a \ x & x \ b --> f x = f a" +apply safe +apply (drule_tac x = a in order_le_imp_less_or_eq, safe) +apply (drule_tac b = x in DERIV_isconst_end, auto) +done + +lemma DERIV_isconst2: + fixes f :: "real => real" + shows "[| a < b; + \x. a \ x & x \ b --> isCont f x; + \x. a < x & x < b --> DERIV f x :> 0; + a \ x; x \ b |] + ==> f x = f a" +apply (blast dest: DERIV_isconst1) +done + +lemma DERIV_isconst_all: + fixes f :: "real => real" + shows "\x. DERIV f x :> 0 ==> f(x) = f(y)" +apply (rule linorder_cases [of x y]) +apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+ +done + +lemma DERIV_const_ratio_const: + fixes f :: "real => real" + shows "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k" +apply (rule linorder_cases [of a b], auto) +apply (drule_tac [!] f = f in MVT) +apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def) +apply (auto dest: DERIV_unique simp add: ring_distribs diff_minus) +done + +lemma DERIV_const_ratio_const2: + fixes f :: "real => real" + shows "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k" +apply (rule_tac c1 = "b-a" in real_mult_right_cancel [THEN iffD1]) +apply (auto dest!: DERIV_const_ratio_const simp add: mult_assoc) +done + +lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)" +by (simp) + +lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)" +by (simp) + +text{*Gallileo's "trick": average velocity = av. of end velocities*} + +lemma DERIV_const_average: + fixes v :: "real => real" + assumes neq: "a \ (b::real)" + and der: "\x. DERIV v x :> k" + shows "v ((a + b)/2) = (v a + v b)/2" +proof (cases rule: linorder_cases [of a b]) + case equal with neq show ?thesis by simp +next + case less + have "(v b - v a) / (b - a) = k" + by (rule DERIV_const_ratio_const2 [OF neq der]) + hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp + moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" + by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) + ultimately show ?thesis using neq by force +next + case greater + have "(v b - v a) / (b - a) = k" + by (rule DERIV_const_ratio_const2 [OF neq der]) + hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp + moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" + by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) + ultimately show ?thesis using neq by (force simp add: add_commute) +qed + + +text{*Dull lemma: an continuous injection on an interval must have a +strict maximum at an end point, not in the middle.*} + +lemma lemma_isCont_inj: + fixes f :: "real \ real" + assumes d: "0 < d" + and inj [rule_format]: "\z. \z-x\ \ d --> g(f z) = z" + and cont: "\z. \z-x\ \ d --> isCont f z" + shows "\z. \z-x\ \ d & f x < f z" +proof (rule ccontr) + assume "~ (\z. \z-x\ \ d & f x < f z)" + hence all [rule_format]: "\z. \z - x\ \ d --> f z \ f x" by auto + show False + proof (cases rule: linorder_le_cases [of "f(x-d)" "f(x+d)"]) + case le + from d cont all [of "x+d"] + have flef: "f(x+d) \ f x" + and xlex: "x - d \ x" + and cont': "\z. x - d \ z \ z \ x \ isCont f z" + by (auto simp add: abs_if) + from IVT [OF le flef xlex cont'] + obtain x' where "x-d \ x'" "x' \ x" "f x' = f(x+d)" by blast + moreover + hence "g(f x') = g (f(x+d))" by simp + ultimately show False using d inj [of x'] inj [of "x+d"] + by (simp add: abs_le_iff) + next + case ge + from d cont all [of "x-d"] + have flef: "f(x-d) \ f x" + and xlex: "x \ x+d" + and cont': "\z. x \ z \ z \ x+d \ isCont f z" + by (auto simp add: abs_if) + from IVT2 [OF ge flef xlex cont'] + obtain x' where "x \ x'" "x' \ x+d" "f x' = f(x-d)" by blast + moreover + hence "g(f x') = g (f(x-d))" by simp + ultimately show False using d inj [of x'] inj [of "x-d"] + by (simp add: abs_le_iff) + qed +qed + + +text{*Similar version for lower bound.*} + +lemma lemma_isCont_inj2: + fixes f g :: "real \ real" + shows "[|0 < d; \z. \z-x\ \ d --> g(f z) = z; + \z. \z-x\ \ d --> isCont f z |] + ==> \z. \z-x\ \ d & f z < f x" +apply (insert lemma_isCont_inj + [where f = "%x. - f x" and g = "%y. g(-y)" and x = x and d = d]) +apply (simp add: isCont_minus linorder_not_le) +done + +text{*Show there's an interval surrounding @{term "f(x)"} in +@{text "f[[x - d, x + d]]"} .*} + +lemma isCont_inj_range: + fixes f :: "real \ real" + assumes d: "0 < d" + and inj: "\z. \z-x\ \ d --> g(f z) = z" + and cont: "\z. \z-x\ \ d --> isCont f z" + shows "\e>0. \y. \y - f x\ \ e --> (\z. \z-x\ \ d & f z = y)" +proof - + have "x-d \ x+d" "\z. x-d \ z \ z \ x+d \ isCont f z" using cont d + by (auto simp add: abs_le_iff) + from isCont_Lb_Ub [OF this] + obtain L M + where all1 [rule_format]: "\z. x-d \ z \ z \ x+d \ L \ f z \ f z \ M" + and all2 [rule_format]: + "\y. L \ y \ y \ M \ (\z. x-d \ z \ z \ x+d \ f z = y)" + by auto + with d have "L \ f x & f x \ M" by simp + moreover have "L \ f x" + proof - + from lemma_isCont_inj2 [OF d inj cont] + obtain u where "\u - x\ \ d" "f u < f x" by auto + thus ?thesis using all1 [of u] by arith + qed + moreover have "f x \ M" + proof - + from lemma_isCont_inj [OF d inj cont] + obtain u where "\u - x\ \ d" "f x < f u" by auto + thus ?thesis using all1 [of u] by arith + qed + ultimately have "L < f x & f x < M" by arith + hence "0 < f x - L" "0 < M - f x" by arith+ + from real_lbound_gt_zero [OF this] + obtain e where e: "0 < e" "e < f x - L" "e < M - f x" by auto + thus ?thesis + proof (intro exI conjI) + show "0y. \y - f x\ \ e \ (\z. \z - x\ \ d \ f z = y)" + proof (intro strip) + fix y::real + assume "\y - f x\ \ e" + with e have "L \ y \ y \ M" by arith + from all2 [OF this] + obtain z where "x - d \ z" "z \ x + d" "f z = y" by blast + thus "\z. \z - x\ \ d \ f z = y" + by (force simp add: abs_le_iff) + qed + qed +qed + + +text{*Continuity of inverse function*} + +lemma isCont_inverse_function: + fixes f g :: "real \ real" + assumes d: "0 < d" + and inj: "\z. \z-x\ \ d --> g(f z) = z" + and cont: "\z. \z-x\ \ d --> isCont f z" + shows "isCont g (f x)" +proof (simp add: isCont_iff LIM_eq) + show "\r. 0 < r \ + (\s>0. \z. z\0 \ \z\ < s \ \g(f x + z) - g(f x)\ < r)" + proof (intro strip) + fix r::real + assume r: "0 e < d" by blast + with inj cont + have e_simps: "\z. \z-x\ \ e --> g (f z) = z" + "\z. \z-x\ \ e --> isCont f z" by auto + from isCont_inj_range [OF e this] + obtain e' where e': "0 < e'" + and all: "\y. \y - f x\ \ e' \ (\z. \z - x\ \ e \ f z = y)" + by blast + show "\s>0. \z. z\0 \ \z\ < s \ \g(f x + z) - g(f x)\ < r" + proof (intro exI conjI) + show "0z. z \ 0 \ \z\ < e' \ \g (f x + z) - g (f x)\ < r" + proof (intro strip) + fix z::real + assume z: "z \ 0 \ \z\ < e'" + with e e_lt e_simps all [rule_format, of "f x + z"] + show "\g (f x + z) - g (f x)\ < r" by force + qed + qed + qed +qed + +text {* Derivative of inverse function *} + +lemma DERIV_inverse_function: + fixes f g :: "real \ real" + assumes der: "DERIV f (g x) :> D" + assumes neq: "D \ 0" + assumes a: "a < x" and b: "x < b" + assumes inj: "\y. a < y \ y < b \ f (g y) = y" + assumes cont: "isCont g x" + shows "DERIV g x :> inverse D" +unfolding DERIV_iff2 +proof (rule LIM_equal2) + show "0 < min (x - a) (b - x)" + using a b by arith +next + fix y + assume "norm (y - x) < min (x - a) (b - x)" + hence "a < y" and "y < b" + by (simp_all add: abs_less_iff) + thus "(g y - g x) / (y - x) = + inverse ((f (g y) - x) / (g y - g x))" + by (simp add: inj) +next + have "(\z. (f z - f (g x)) / (z - g x)) -- g x --> D" + by (rule der [unfolded DERIV_iff2]) + hence 1: "(\z. (f z - x) / (z - g x)) -- g x --> D" + using inj a b by simp + have 2: "\d>0. \y. y \ x \ norm (y - x) < d \ g y \ g x" + proof (safe intro!: exI) + show "0 < min (x - a) (b - x)" + using a b by simp + next + fix y + assume "norm (y - x) < min (x - a) (b - x)" + hence y: "a < y" "y < b" + by (simp_all add: abs_less_iff) + assume "g y = g x" + hence "f (g y) = f (g x)" by simp + hence "y = x" using inj y a b by simp + also assume "y \ x" + finally show False by simp + qed + have "(\y. (f (g y) - x) / (g y - g x)) -- x --> D" + using cont 1 2 by (rule isCont_LIM_compose2) + thus "(\y. inverse ((f (g y) - x) / (g y - g x))) + -- x --> inverse D" + using neq by (rule LIM_inverse) +qed + +theorem GMVT: + fixes a b :: real + assumes alb: "a < b" + and fc: "\x. a \ x \ x \ b \ isCont f x" + and fd: "\x. a < x \ x < b \ f differentiable x" + and gc: "\x. a \ x \ x \ b \ isCont g x" + and gd: "\x. a < x \ x < b \ g differentiable x" + shows "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ ((f b - f a) * g'c) = ((g b - g a) * f'c)" +proof - + let ?h = "\x. (f b - f a)*(g x) - (g b - g a)*(f x)" + from prems have "a < b" by simp + moreover have "\x. a \ x \ x \ b \ isCont ?h x" + proof - + have "\x. a <= x \ x <= b \ isCont (\x. f b - f a) x" by simp + with gc have "\x. a <= x \ x <= b \ isCont (\x. (f b - f a) * g x) x" + by (auto intro: isCont_mult) + moreover + have "\x. a <= x \ x <= b \ isCont (\x. g b - g a) x" by simp + with fc have "\x. a <= x \ x <= b \ isCont (\x. (g b - g a) * f x) x" + by (auto intro: isCont_mult) + ultimately show ?thesis + by (fastsimp intro: isCont_diff) + qed + moreover + have "\x. a < x \ x < b \ ?h differentiable x" + proof - + have "\x. a < x \ x < b \ (\x. f b - f a) differentiable x" by (simp add: differentiable_const) + with gd have "\x. a < x \ x < b \ (\x. (f b - f a) * g x) differentiable x" by (simp add: differentiable_mult) + moreover + have "\x. a < x \ x < b \ (\x. g b - g a) differentiable x" by (simp add: differentiable_const) + with fd have "\x. a < x \ x < b \ (\x. (g b - g a) * f x) differentiable x" by (simp add: differentiable_mult) + ultimately show ?thesis by (simp add: differentiable_diff) + qed + ultimately have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" by (rule MVT) + then obtain l where ldef: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. + then obtain c where cdef: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" .. + + from cdef have cint: "a < c \ c < b" by auto + with gd have "g differentiable c" by simp + hence "\D. DERIV g c :> D" by (rule differentiableD) + then obtain g'c where g'cdef: "DERIV g c :> g'c" .. + + from cdef have "a < c \ c < b" by auto + with fd have "f differentiable c" by simp + hence "\D. DERIV f c :> D" by (rule differentiableD) + then obtain f'c where f'cdef: "DERIV f c :> f'c" .. + + from cdef have "DERIV ?h c :> l" by auto + moreover + { + have "DERIV (\x. (f b - f a) * g x) c :> g'c * (f b - f a)" + apply (insert DERIV_const [where k="f b - f a"]) + apply (drule meta_spec [of _ c]) + apply (drule DERIV_mult [OF _ g'cdef]) + by simp + moreover have "DERIV (\x. (g b - g a) * f x) c :> f'c * (g b - g a)" + apply (insert DERIV_const [where k="g b - g a"]) + apply (drule meta_spec [of _ c]) + apply (drule DERIV_mult [OF _ f'cdef]) + by simp + ultimately have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" + by (simp add: DERIV_diff) + } + ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) + + { + from cdef have "?h b - ?h a = (b - a) * l" by auto + also with leq have "\ = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp + finally have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp + } + moreover + { + have "?h b - ?h a = + ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - + ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" + by (simp add: mult_ac add_ac right_diff_distrib) + hence "?h b - ?h a = 0" by auto + } + ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto + with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp + hence "g'c * (f b - f a) = f'c * (g b - g a)" by simp + hence "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: mult_ac) + + with g'cdef f'cdef cint show ?thesis by auto +qed + +lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" +by auto + +subsection {* Derivatives of univariate polynomials *} + + + +primrec pderiv_aux :: "nat => real list => real list" where + pderiv_aux_Nil: "pderiv_aux n [] = []" +| pderiv_aux_Cons: "pderiv_aux n (h#t) = + (real n * h)#(pderiv_aux (Suc n) t)" + +definition + pderiv :: "real list => real list" where + "pderiv p = (if p = [] then [] else pderiv_aux 1 (tl p))" + + +text{*The derivative*} + +lemma pderiv_Nil: "pderiv [] = []" + +apply (simp add: pderiv_def) +done +declare pderiv_Nil [simp] + +lemma pderiv_singleton: "pderiv [c] = []" +by (simp add: pderiv_def) +declare pderiv_singleton [simp] + +lemma pderiv_Cons: "pderiv (h#t) = pderiv_aux 1 t" +by (simp add: pderiv_def) + +lemma DERIV_cmult2: "DERIV f x :> D ==> DERIV (%x. (f x) * c :: real) x :> D * c" +by (simp add: DERIV_cmult mult_commute [of _ c]) + +lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" +by (rule lemma_DERIV_subst, rule DERIV_pow, simp) +declare DERIV_pow2 [simp] DERIV_pow [simp] + +lemma lemma_DERIV_poly1: "\n. DERIV (%x. (x ^ (Suc n) * poly p x)) x :> + x ^ n * poly (pderiv_aux (Suc n) p) x " +apply (induct "p") +apply (auto intro!: DERIV_add DERIV_cmult2 + simp add: pderiv_def right_distrib real_mult_assoc [symmetric] + simp del: realpow_Suc) +apply (subst mult_commute) +apply (simp del: realpow_Suc) +apply (simp add: mult_commute realpow_Suc [symmetric] del: realpow_Suc) +done + +lemma lemma_DERIV_poly: "DERIV (%x. (x ^ (Suc n) * poly p x)) x :> + x ^ n * poly (pderiv_aux (Suc n) p) x " +by (simp add: lemma_DERIV_poly1 del: realpow_Suc) + +lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: real) x :> D" +by (rule lemma_DERIV_subst, rule DERIV_add, auto) + +lemma poly_DERIV[simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x" +apply (induct "p") +apply (auto simp add: pderiv_Cons) +apply (rule DERIV_add_const) +apply (rule lemma_DERIV_subst) +apply (rule lemma_DERIV_poly [where n=0, simplified], simp) +done + + +text{* Consequences of the derivative theorem above*} + +lemma poly_differentiable[simp]: "(%x. poly p x) differentiable (x::real)" +apply (simp add: differentiable_def) +apply (blast intro: poly_DERIV) +done + +lemma poly_isCont[simp]: "isCont (%x. poly p x) (x::real)" +by (rule poly_DERIV [THEN DERIV_isCont]) + +lemma poly_IVT_pos: "[| a < b; poly p (a::real) < 0; 0 < poly p b |] + ==> \x. a < x & x < b & (poly p x = 0)" +apply (cut_tac f = "%x. poly p x" and a = a and b = b and y = 0 in IVT_objl) +apply (auto simp add: order_le_less) +done + +lemma poly_IVT_neg: "[| (a::real) < b; 0 < poly p a; poly p b < 0 |] + ==> \x. a < x & x < b & (poly p x = 0)" +apply (insert poly_IVT_pos [where p = "-- p" ]) +apply (simp add: poly_minus neg_less_0_iff_less) +done + +lemma poly_MVT: "a < b ==> + \x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)" +apply (drule_tac f = "poly p" in MVT, auto) +apply (rule_tac x = z in exI) +apply (auto simp add: real_mult_left_cancel poly_DERIV [THEN DERIV_unique]) +done + +text{*Lemmas for Derivatives*} + +lemma lemma_poly_pderiv_aux_add: "\p2 n. poly (pderiv_aux n (p1 +++ p2)) x = + poly (pderiv_aux n p1 +++ pderiv_aux n p2) x" +apply (induct "p1", simp, clarify) +apply (case_tac "p2") +apply (auto simp add: right_distrib) +done + +lemma poly_pderiv_aux_add: "poly (pderiv_aux n (p1 +++ p2)) x = + poly (pderiv_aux n p1 +++ pderiv_aux n p2) x" +apply (simp add: lemma_poly_pderiv_aux_add) +done + +lemma lemma_poly_pderiv_aux_cmult: "\n. poly (pderiv_aux n (c %* p)) x = poly (c %* pderiv_aux n p) x" +apply (induct "p") +apply (auto simp add: poly_cmult mult_ac) +done + +lemma poly_pderiv_aux_cmult: "poly (pderiv_aux n (c %* p)) x = poly (c %* pderiv_aux n p) x" +by (simp add: lemma_poly_pderiv_aux_cmult) + +lemma poly_pderiv_aux_minus: + "poly (pderiv_aux n (-- p)) x = poly (-- pderiv_aux n p) x" +apply (simp add: poly_minus_def poly_pderiv_aux_cmult) +done + +lemma lemma_poly_pderiv_aux_mult1: "\n. poly (pderiv_aux (Suc n) p) x = poly ((pderiv_aux n p) +++ p) x" +apply (induct "p") +apply (auto simp add: real_of_nat_Suc left_distrib) +done + +lemma lemma_poly_pderiv_aux_mult: "poly (pderiv_aux (Suc n) p) x = poly ((pderiv_aux n p) +++ p) x" +by (simp add: lemma_poly_pderiv_aux_mult1) + +lemma lemma_poly_pderiv_add: "\q. poly (pderiv (p +++ q)) x = poly (pderiv p +++ pderiv q) x" +apply (induct "p", simp, clarify) +apply (case_tac "q") +apply (auto simp add: poly_pderiv_aux_add poly_add pderiv_def) +done + +lemma poly_pderiv_add: "poly (pderiv (p +++ q)) x = poly (pderiv p +++ pderiv q) x" +by (simp add: lemma_poly_pderiv_add) + +lemma poly_pderiv_cmult: "poly (pderiv (c %* p)) x = poly (c %* (pderiv p)) x" +apply (induct "p") +apply (auto simp add: poly_pderiv_aux_cmult poly_cmult pderiv_def) +done + +lemma poly_pderiv_minus: "poly (pderiv (--p)) x = poly (--(pderiv p)) x" +by (simp add: poly_minus_def poly_pderiv_cmult) + +lemma lemma_poly_mult_pderiv: + "poly (pderiv (h#t)) x = poly ((0 # (pderiv t)) +++ t) x" +apply (simp add: pderiv_def) +apply (induct "t") +apply (auto simp add: poly_add lemma_poly_pderiv_aux_mult) +done + +lemma poly_pderiv_mult: "\q. poly (pderiv (p *** q)) x = + poly (p *** (pderiv q) +++ q *** (pderiv p)) x" +apply (induct "p") +apply (auto simp add: poly_add poly_cmult poly_pderiv_cmult poly_pderiv_add poly_mult) +apply (rule lemma_poly_mult_pderiv [THEN ssubst]) +apply (rule lemma_poly_mult_pderiv [THEN ssubst]) +apply (rule poly_add [THEN ssubst]) +apply (rule poly_add [THEN ssubst]) +apply (simp (no_asm_simp) add: poly_mult right_distrib add_ac mult_ac) +done + +lemma poly_pderiv_exp: "poly (pderiv (p %^ (Suc n))) x = + poly ((real (Suc n)) %* (p %^ n) *** pderiv p) x" +apply (induct "n") +apply (auto simp add: poly_add poly_pderiv_cmult poly_cmult poly_pderiv_mult + real_of_nat_zero poly_mult real_of_nat_Suc + right_distrib left_distrib mult_ac) +done + +lemma poly_pderiv_exp_prime: "poly (pderiv ([-a, 1] %^ (Suc n))) x = + poly (real (Suc n) %* ([-a, 1] %^ n)) x" +apply (simp add: poly_pderiv_exp poly_mult del: pexp_Suc) +apply (simp add: poly_cmult pderiv_def) +done + + +lemma real_mult_zero_disj_iff[simp]: "(x * y = 0) = (x = (0::real) | y = 0)" +by simp + +lemma pderiv_aux_iszero [rule_format, simp]: + "\n. list_all (%c. c = 0) (pderiv_aux (Suc n) p) = list_all (%c. c = 0) p" +by (induct "p", auto) + +lemma pderiv_aux_iszero_num: "(number_of n :: nat) \ 0 + ==> (list_all (%c. c = 0) (pderiv_aux (number_of n) p) = + list_all (%c. c = 0) p)" +unfolding neq0_conv +apply (rule_tac n1 = "number_of n" and m1 = 0 in less_imp_Suc_add [THEN exE], force) +apply (rule_tac n1 = "0 + x" in pderiv_aux_iszero [THEN subst]) +apply (simp (no_asm_simp) del: pderiv_aux_iszero) +done + +instance real:: idom_char_0 +apply (intro_classes) +done + +instance real:: recpower_idom_char_0 +apply (intro_classes) +done + +lemma pderiv_iszero [rule_format]: + "poly (pderiv p) = poly [] --> (\h. poly p = poly [h])" +apply (simp add: poly_zero) +apply (induct "p", force) +apply (simp add: pderiv_Cons pderiv_aux_iszero_num del: poly_Cons) +apply (auto simp add: poly_zero [symmetric]) +done + +lemma pderiv_zero_obj: "poly p = poly [] --> (poly (pderiv p) = poly [])" +apply (simp add: poly_zero) +apply (induct "p", force) +apply (simp add: pderiv_Cons pderiv_aux_iszero_num del: poly_Cons) +done + +lemma pderiv_zero: "poly p = poly [] ==> (poly (pderiv p) = poly [])" +by (blast elim: pderiv_zero_obj [THEN impE]) +declare pderiv_zero [simp] + +lemma poly_pderiv_welldef: "poly p = poly q ==> (poly (pderiv p) = poly (pderiv q))" +apply (cut_tac p = "p +++ --q" in pderiv_zero_obj) +apply (simp add: fun_eq poly_add poly_minus poly_pderiv_add poly_pderiv_minus del: pderiv_zero) +done + +lemma lemma_order_pderiv [rule_format]: + "\p q a. 0 < n & + poly (pderiv p) \ poly [] & + poly p = poly ([- a, 1] %^ n *** q) & ~ [- a, 1] divides q + --> n = Suc (order a (pderiv p))" +apply (induct "n", safe) +apply (rule order_unique_lemma, rule conjI, assumption) +apply (subgoal_tac "\r. r divides (pderiv p) = r divides (pderiv ([-a, 1] %^ Suc n *** q))") +apply (drule_tac [2] poly_pderiv_welldef) + prefer 2 apply (simp add: divides_def del: pmult_Cons pexp_Suc) +apply (simp del: pmult_Cons pexp_Suc) +apply (rule conjI) +apply (simp add: divides_def fun_eq del: pmult_Cons pexp_Suc) +apply (rule_tac x = "[-a, 1] *** (pderiv q) +++ real (Suc n) %* q" in exI) +apply (simp add: poly_pderiv_mult poly_pderiv_exp_prime poly_add poly_mult poly_cmult right_distrib mult_ac del: pmult_Cons pexp_Suc) +apply (simp add: poly_mult right_distrib left_distrib mult_ac del: pmult_Cons) +apply (erule_tac V = "\r. r divides pderiv p = r divides pderiv ([- a, 1] %^ Suc n *** q)" in thin_rl) +apply (unfold divides_def) +apply (simp (no_asm) add: poly_pderiv_mult poly_pderiv_exp_prime fun_eq poly_add poly_mult del: pmult_Cons pexp_Suc) +apply (rule contrapos_np, assumption) +apply (rotate_tac 3, erule contrapos_np) +apply (simp del: pmult_Cons pexp_Suc, safe) +apply (rule_tac x = "inverse (real (Suc n)) %* (qa +++ -- (pderiv q))" in exI) +apply (subgoal_tac "poly ([-a, 1] %^ n *** q) = poly ([-a, 1] %^ n *** ([-a, 1] *** (inverse (real (Suc n)) %* (qa +++ -- (pderiv q))))) ") +apply (drule poly_mult_left_cancel [THEN iffD1], simp) +apply (simp add: fun_eq poly_mult poly_add poly_cmult poly_minus del: pmult_Cons mult_cancel_left, safe) +apply (rule_tac c1 = "real (Suc n)" in real_mult_left_cancel [THEN iffD1]) +apply (simp (no_asm)) +apply (subgoal_tac "real (Suc n) * (poly ([- a, 1] %^ n) xa * poly q xa) = + (poly qa xa + - poly (pderiv q) xa) * + (poly ([- a, 1] %^ n) xa * + ((- a + xa) * (inverse (real (Suc n)) * real (Suc n))))") +apply (simp only: mult_ac) +apply (rotate_tac 2) +apply (drule_tac x = xa in spec) +apply (simp add: left_distrib mult_ac del: pmult_Cons) +done + +lemma order_pderiv: "[| poly (pderiv p) \ poly []; order a p \ 0 |] + ==> (order a p = Suc (order a (pderiv p)))" +apply (case_tac "poly p = poly []") +apply (auto dest: pderiv_zero) +apply (drule_tac a = a and p = p in order_decomp) +using neq0_conv +apply (blast intro: lemma_order_pderiv) +done + +text{*Now justify the standard squarefree decomposition, i.e. f / gcd(f,f'). *} + +lemma poly_squarefree_decomp_order: "[| poly (pderiv p) \ poly []; + poly p = poly (q *** d); + poly (pderiv p) = poly (e *** d); + poly d = poly (r *** p +++ s *** pderiv p) + |] ==> order a q = (if order a p = 0 then 0 else 1)" +apply (subgoal_tac "order a p = order a q + order a d") +apply (rule_tac [2] s = "order a (q *** d)" in trans) +prefer 2 apply (blast intro: order_poly) +apply (rule_tac [2] order_mult) + prefer 2 apply force +apply (case_tac "order a p = 0", simp) +apply (subgoal_tac "order a (pderiv p) = order a e + order a d") +apply (rule_tac [2] s = "order a (e *** d)" in trans) +prefer 2 apply (blast intro: order_poly) +apply (rule_tac [2] order_mult) + prefer 2 apply force +apply (case_tac "poly p = poly []") +apply (drule_tac p = p in pderiv_zero, simp) +apply (drule order_pderiv, assumption) +apply (subgoal_tac "order a (pderiv p) \ order a d") +apply (subgoal_tac [2] " ([-a, 1] %^ (order a (pderiv p))) divides d") + prefer 2 apply (simp add: poly_entire order_divides) +apply (subgoal_tac [2] " ([-a, 1] %^ (order a (pderiv p))) divides p & ([-a, 1] %^ (order a (pderiv p))) divides (pderiv p) ") + prefer 3 apply (simp add: order_divides) + prefer 2 apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe) +apply (rule_tac x = "r *** qa +++ s *** qaa" in exI) +apply (simp add: fun_eq poly_add poly_mult left_distrib right_distrib mult_ac del: pexp_Suc pmult_Cons, auto) +done + + +lemma poly_squarefree_decomp_order2: "[| poly (pderiv p) \ poly []; + poly p = poly (q *** d); + poly (pderiv p) = poly (e *** d); + poly d = poly (r *** p +++ s *** pderiv p) + |] ==> \a. order a q = (if order a p = 0 then 0 else 1)" +apply (blast intro: poly_squarefree_decomp_order) +done + +lemma order_pderiv2: "[| poly (pderiv p) \ poly []; order a p \ 0 |] + ==> (order a (pderiv p) = n) = (order a p = Suc n)" +apply (auto dest: order_pderiv) +done + +lemma rsquarefree_roots: + "rsquarefree p = (\a. ~(poly p a = 0 & poly (pderiv p) a = 0))" +apply (simp add: rsquarefree_def) +apply (case_tac "poly p = poly []", simp, simp) +apply (case_tac "poly (pderiv p) = poly []") +apply simp +apply (drule pderiv_iszero, clarify) +apply (subgoal_tac "\a. order a p = order a [h]") +apply (simp add: fun_eq) +apply (rule allI) +apply (cut_tac p = "[h]" and a = a in order_root) +apply (simp add: fun_eq) +apply (blast intro: order_poly) +apply (auto simp add: order_root order_pderiv2) +apply (erule_tac x="a" in allE, simp) +done + +lemma poly_squarefree_decomp: "[| poly (pderiv p) \ poly []; + poly p = poly (q *** d); + poly (pderiv p) = poly (e *** d); + poly d = poly (r *** p +++ s *** pderiv p) + |] ==> rsquarefree q & (\a. (poly q a = 0) = (poly p a = 0))" +apply (frule poly_squarefree_decomp_order2, assumption+) +apply (case_tac "poly p = poly []") +apply (blast dest: pderiv_zero) +apply (simp (no_asm) add: rsquarefree_def order_root del: pmult_Cons) +apply (simp add: poly_entire del: pmult_Cons) +done + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Fact.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Fact.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,76 @@ +(* Title : Fact.thy + Author : Jacques D. Fleuriot + Copyright : 1998 University of Cambridge + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 +*) + +header{*Factorial Function*} + +theory Fact +imports RealDef +begin + +consts fact :: "nat => nat" +primrec + fact_0: "fact 0 = 1" + fact_Suc: "fact (Suc n) = (Suc n) * fact n" + + +lemma fact_gt_zero [simp]: "0 < fact n" +by (induct n) auto + +lemma fact_not_eq_zero [simp]: "fact n \ 0" +by simp + +lemma real_of_nat_fact_not_zero [simp]: "real (fact n) \ 0" +by auto + +lemma real_of_nat_fact_gt_zero [simp]: "0 < real(fact n)" +by auto + +lemma real_of_nat_fact_ge_zero [simp]: "0 \ real(fact n)" +by simp + +lemma fact_ge_one [simp]: "1 \ fact n" +by (induct n) auto + +lemma fact_mono: "m \ n ==> fact m \ fact n" +apply (drule le_imp_less_or_eq) +apply (auto dest!: less_imp_Suc_add) +apply (induct_tac k, auto) +done + +text{*Note that @{term "fact 0 = fact 1"}*} +lemma fact_less_mono: "[| 0 < m; m < n |] ==> fact m < fact n" +apply (drule_tac m = m in less_imp_Suc_add, auto) +apply (induct_tac k, auto) +done + +lemma inv_real_of_nat_fact_gt_zero [simp]: "0 < inverse (real (fact n))" +by (auto simp add: positive_imp_inverse_positive) + +lemma inv_real_of_nat_fact_ge_zero [simp]: "0 \ inverse (real (fact n))" +by (auto intro: order_less_imp_le) + +lemma fact_diff_Suc [rule_format]: + "n < Suc m ==> fact (Suc m - n) = (Suc m - n) * fact (m - n)" +apply (induct n arbitrary: m) +apply auto +apply (drule_tac x = "m - 1" in meta_spec, auto) +done + +lemma fact_num0 [simp]: "fact 0 = 1" +by auto + +lemma fact_num_eq_if: "fact m = (if m=0 then 1 else m * fact (m - 1))" +by (cases m) auto + +lemma fact_add_num_eq_if: + "fact (m + n) = (if m + n = 0 then 1 else (m + n) * fact (m + n - 1))" +by (cases "m + n") auto + +lemma fact_add_num_eq_if2: + "fact (m + n) = (if m = 0 then fact n else (m + n) * fact ((m - 1) + n))" +by (cases m) auto + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/FrechetDeriv.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/FrechetDeriv.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,503 @@ +(* Title : FrechetDeriv.thy + ID : $Id$ + Author : Brian Huffman +*) + +header {* Frechet Derivative *} + +theory FrechetDeriv +imports Lim +begin + +definition + fderiv :: + "['a::real_normed_vector \ 'b::real_normed_vector, 'a, 'a \ 'b] \ bool" + -- {* Frechet derivative: D is derivative of function f at x *} + ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where + "FDERIV f x :> D = (bounded_linear D \ + (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0)" + +lemma FDERIV_I: + "\bounded_linear D; (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0\ + \ FDERIV f x :> D" +by (simp add: fderiv_def) + +lemma FDERIV_D: + "FDERIV f x :> D \ (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0" +by (simp add: fderiv_def) + +lemma FDERIV_bounded_linear: "FDERIV f x :> D \ bounded_linear D" +by (simp add: fderiv_def) + +lemma bounded_linear_zero: + "bounded_linear (\x::'a::real_normed_vector. 0::'b::real_normed_vector)" +proof + show "(0::'b) = 0 + 0" by simp + fix r show "(0::'b) = scaleR r 0" by simp + have "\x::'a. norm (0::'b) \ norm x * 0" by simp + thus "\K. \x::'a. norm (0::'b) \ norm x * K" .. +qed + +lemma FDERIV_const: "FDERIV (\x. k) x :> (\h. 0)" +by (simp add: fderiv_def bounded_linear_zero) + +lemma bounded_linear_ident: + "bounded_linear (\x::'a::real_normed_vector. x)" +proof + fix x y :: 'a show "x + y = x + y" by simp + fix r and x :: 'a show "scaleR r x = scaleR r x" by simp + have "\x::'a. norm x \ norm x * 1" by simp + thus "\K. \x::'a. norm x \ norm x * K" .. +qed + +lemma FDERIV_ident: "FDERIV (\x. x) x :> (\h. h)" +by (simp add: fderiv_def bounded_linear_ident) + +subsection {* Addition *} + +lemma add_diff_add: + fixes a b c d :: "'a::ab_group_add" + shows "(a + c) - (b + d) = (a - b) + (c - d)" +by simp + +lemma bounded_linear_add: + assumes "bounded_linear f" + assumes "bounded_linear g" + shows "bounded_linear (\x. f x + g x)" +proof - + interpret f: bounded_linear [f] by fact + interpret g: bounded_linear [g] by fact + show ?thesis apply (unfold_locales) + apply (simp only: f.add g.add add_ac) + apply (simp only: f.scaleR g.scaleR scaleR_right_distrib) + apply (rule f.pos_bounded [THEN exE], rename_tac Kf) + apply (rule g.pos_bounded [THEN exE], rename_tac Kg) + apply (rule_tac x="Kf + Kg" in exI, safe) + apply (subst right_distrib) + apply (rule order_trans [OF norm_triangle_ineq]) + apply (rule add_mono, erule spec, erule spec) + done +qed + +lemma norm_ratio_ineq: + fixes x y :: "'a::real_normed_vector" + fixes h :: "'b::real_normed_vector" + shows "norm (x + y) / norm h \ norm x / norm h + norm y / norm h" +apply (rule ord_le_eq_trans) +apply (rule divide_right_mono) +apply (rule norm_triangle_ineq) +apply (rule norm_ge_zero) +apply (rule add_divide_distrib) +done + +lemma FDERIV_add: + assumes f: "FDERIV f x :> F" + assumes g: "FDERIV g x :> G" + shows "FDERIV (\x. f x + g x) x :> (\h. F h + G h)" +proof (rule FDERIV_I) + show "bounded_linear (\h. F h + G h)" + apply (rule bounded_linear_add) + apply (rule FDERIV_bounded_linear [OF f]) + apply (rule FDERIV_bounded_linear [OF g]) + done +next + have f': "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" + using f by (rule FDERIV_D) + have g': "(\h. norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" + using g by (rule FDERIV_D) + from f' g' + have "(\h. norm (f (x + h) - f x - F h) / norm h + + norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" + by (rule LIM_add_zero) + thus "(\h. norm (f (x + h) + g (x + h) - (f x + g x) - (F h + G h)) + / norm h) -- 0 --> 0" + apply (rule real_LIM_sandwich_zero) + apply (simp add: divide_nonneg_pos) + apply (simp only: add_diff_add) + apply (rule norm_ratio_ineq) + done +qed + +subsection {* Subtraction *} + +lemma bounded_linear_minus: + assumes "bounded_linear f" + shows "bounded_linear (\x. - f x)" +proof - + interpret f: bounded_linear [f] by fact + show ?thesis apply (unfold_locales) + apply (simp add: f.add) + apply (simp add: f.scaleR) + apply (simp add: f.bounded) + done +qed + +lemma FDERIV_minus: + "FDERIV f x :> F \ FDERIV (\x. - f x) x :> (\h. - F h)" +apply (rule FDERIV_I) +apply (rule bounded_linear_minus) +apply (erule FDERIV_bounded_linear) +apply (simp only: fderiv_def minus_diff_minus norm_minus_cancel) +done + +lemma FDERIV_diff: + "\FDERIV f x :> F; FDERIV g x :> G\ + \ FDERIV (\x. f x - g x) x :> (\h. F h - G h)" +by (simp only: diff_minus FDERIV_add FDERIV_minus) + +subsection {* Continuity *} + +lemma FDERIV_isCont: + assumes f: "FDERIV f x :> F" + shows "isCont f x" +proof - + from f interpret F: bounded_linear ["F"] by (rule FDERIV_bounded_linear) + have "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" + by (rule FDERIV_D [OF f]) + hence "(\h. norm (f (x + h) - f x - F h) / norm h * norm h) -- 0 --> 0" + by (intro LIM_mult_zero LIM_norm_zero LIM_ident) + hence "(\h. norm (f (x + h) - f x - F h)) -- 0 --> 0" + by (simp cong: LIM_cong) + hence "(\h. f (x + h) - f x - F h) -- 0 --> 0" + by (rule LIM_norm_zero_cancel) + hence "(\h. f (x + h) - f x - F h + F h) -- 0 --> 0" + by (intro LIM_add_zero F.LIM_zero LIM_ident) + hence "(\h. f (x + h) - f x) -- 0 --> 0" + by simp + thus "isCont f x" + unfolding isCont_iff by (rule LIM_zero_cancel) +qed + +subsection {* Composition *} + +lemma real_divide_cancel_lemma: + fixes a b c :: real + shows "(b = 0 \ a = 0) \ (a / b) * (b / c) = a / c" +by simp + +lemma bounded_linear_compose: + assumes "bounded_linear f" + assumes "bounded_linear g" + shows "bounded_linear (\x. f (g x))" +proof - + interpret f: bounded_linear [f] by fact + interpret g: bounded_linear [g] by fact + show ?thesis proof (unfold_locales) + fix x y show "f (g (x + y)) = f (g x) + f (g y)" + by (simp only: f.add g.add) + next + fix r x show "f (g (scaleR r x)) = scaleR r (f (g x))" + by (simp only: f.scaleR g.scaleR) + next + from f.pos_bounded + obtain Kf where f: "\x. norm (f x) \ norm x * Kf" and Kf: "0 < Kf" by fast + from g.pos_bounded + obtain Kg where g: "\x. norm (g x) \ norm x * Kg" by fast + show "\K. \x. norm (f (g x)) \ norm x * K" + proof (intro exI allI) + fix x + have "norm (f (g x)) \ norm (g x) * Kf" + using f . + also have "\ \ (norm x * Kg) * Kf" + using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) + also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" + by (rule mult_assoc) + finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . + qed + qed +qed + +lemma FDERIV_compose: + fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" + fixes g :: "'b::real_normed_vector \ 'c::real_normed_vector" + assumes f: "FDERIV f x :> F" + assumes g: "FDERIV g (f x) :> G" + shows "FDERIV (\x. g (f x)) x :> (\h. G (F h))" +proof (rule FDERIV_I) + from FDERIV_bounded_linear [OF g] FDERIV_bounded_linear [OF f] + show "bounded_linear (\h. G (F h))" + by (rule bounded_linear_compose) +next + let ?Rf = "\h. f (x + h) - f x - F h" + let ?Rg = "\k. g (f x + k) - g (f x) - G k" + let ?k = "\h. f (x + h) - f x" + let ?Nf = "\h. norm (?Rf h) / norm h" + let ?Ng = "\h. norm (?Rg (?k h)) / norm (?k h)" + from f interpret F: bounded_linear ["F"] by (rule FDERIV_bounded_linear) + from g interpret G: bounded_linear ["G"] by (rule FDERIV_bounded_linear) + from F.bounded obtain kF where kF: "\x. norm (F x) \ norm x * kF" by fast + from G.bounded obtain kG where kG: "\x. norm (G x) \ norm x * kG" by fast + + let ?fun2 = "\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)" + + show "(\h. norm (g (f (x + h)) - g (f x) - G (F h)) / norm h) -- 0 --> 0" + proof (rule real_LIM_sandwich_zero) + have Nf: "?Nf -- 0 --> 0" + using FDERIV_D [OF f] . + + have Ng1: "isCont (\k. norm (?Rg k) / norm k) 0" + by (simp add: isCont_def FDERIV_D [OF g]) + have Ng2: "?k -- 0 --> 0" + apply (rule LIM_zero) + apply (fold isCont_iff) + apply (rule FDERIV_isCont [OF f]) + done + have Ng: "?Ng -- 0 --> 0" + using isCont_LIM_compose [OF Ng1 Ng2] by simp + + have "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) + -- 0 --> 0 * kG + 0 * (0 + kF)" + by (intro LIM_add LIM_mult LIM_const Nf Ng) + thus "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) -- 0 --> 0" + by simp + next + fix h::'a assume h: "h \ 0" + thus "0 \ norm (g (f (x + h)) - g (f x) - G (F h)) / norm h" + by (simp add: divide_nonneg_pos) + next + fix h::'a assume h: "h \ 0" + have "g (f (x + h)) - g (f x) - G (F h) = G (?Rf h) + ?Rg (?k h)" + by (simp add: G.diff) + hence "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h + = norm (G (?Rf h) + ?Rg (?k h)) / norm h" + by (rule arg_cong) + also have "\ \ norm (G (?Rf h)) / norm h + norm (?Rg (?k h)) / norm h" + by (rule norm_ratio_ineq) + also have "\ \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" + proof (rule add_mono) + show "norm (G (?Rf h)) / norm h \ ?Nf h * kG" + apply (rule ord_le_eq_trans) + apply (rule divide_right_mono [OF kG norm_ge_zero]) + apply simp + done + next + have "norm (?Rg (?k h)) / norm h = ?Ng h * (norm (?k h) / norm h)" + apply (rule real_divide_cancel_lemma [symmetric]) + apply (simp add: G.zero) + done + also have "\ \ ?Ng h * (?Nf h + kF)" + proof (rule mult_left_mono) + have "norm (?k h) / norm h = norm (?Rf h + F h) / norm h" + by simp + also have "\ \ ?Nf h + norm (F h) / norm h" + by (rule norm_ratio_ineq) + also have "\ \ ?Nf h + kF" + apply (rule add_left_mono) + apply (subst pos_divide_le_eq, simp add: h) + apply (subst mult_commute) + apply (rule kF) + done + finally show "norm (?k h) / norm h \ ?Nf h + kF" . + next + show "0 \ ?Ng h" + apply (case_tac "f (x + h) - f x = 0", simp) + apply (rule divide_nonneg_pos [OF norm_ge_zero]) + apply simp + done + qed + finally show "norm (?Rg (?k h)) / norm h \ ?Ng h * (?Nf h + kF)" . + qed + finally show "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h + \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" . + qed +qed + +subsection {* Product Rule *} + +lemma (in bounded_bilinear) FDERIV_lemma: + "a' ** b' - a ** b - (a ** B + A ** b) + = a ** (b' - b - B) + (a' - a - A) ** b' + A ** (b' - b)" +by (simp add: diff_left diff_right) + +lemma (in bounded_bilinear) FDERIV: + fixes x :: "'d::real_normed_vector" + assumes f: "FDERIV f x :> F" + assumes g: "FDERIV g x :> G" + shows "FDERIV (\x. f x ** g x) x :> (\h. f x ** G h + F h ** g x)" +proof (rule FDERIV_I) + show "bounded_linear (\h. f x ** G h + F h ** g x)" + apply (rule bounded_linear_add) + apply (rule bounded_linear_compose [OF bounded_linear_right]) + apply (rule FDERIV_bounded_linear [OF g]) + apply (rule bounded_linear_compose [OF bounded_linear_left]) + apply (rule FDERIV_bounded_linear [OF f]) + done +next + from bounded_linear.bounded [OF FDERIV_bounded_linear [OF f]] + obtain KF where norm_F: "\x. norm (F x) \ norm x * KF" by fast + + from pos_bounded obtain K where K: "0 < K" and norm_prod: + "\a b. norm (a ** b) \ norm a * norm b * K" by fast + + let ?Rf = "\h. f (x + h) - f x - F h" + let ?Rg = "\h. g (x + h) - g x - G h" + + let ?fun1 = "\h. + norm (f x ** ?Rg h + ?Rf h ** g (x + h) + F h ** (g (x + h) - g x)) / + norm h" + + let ?fun2 = "\h. + norm (f x) * (norm (?Rg h) / norm h) * K + + norm (?Rf h) / norm h * norm (g (x + h)) * K + + KF * norm (g (x + h) - g x) * K" + + have "?fun1 -- 0 --> 0" + proof (rule real_LIM_sandwich_zero) + from f g isCont_iff [THEN iffD1, OF FDERIV_isCont [OF g]] + have "?fun2 -- 0 --> + norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K" + by (intro LIM_add LIM_mult LIM_const LIM_norm LIM_zero FDERIV_D) + thus "?fun2 -- 0 --> 0" + by simp + next + fix h::'d assume "h \ 0" + thus "0 \ ?fun1 h" + by (simp add: divide_nonneg_pos) + next + fix h::'d assume "h \ 0" + have "?fun1 h \ (norm (f x) * norm (?Rg h) * K + + norm (?Rf h) * norm (g (x + h)) * K + + norm h * KF * norm (g (x + h) - g x) * K) / norm h" + by (intro + divide_right_mono mult_mono' + order_trans [OF norm_triangle_ineq add_mono] + order_trans [OF norm_prod mult_right_mono] + mult_nonneg_nonneg order_refl norm_ge_zero norm_F + K [THEN order_less_imp_le] + ) + also have "\ = ?fun2 h" + by (simp add: add_divide_distrib) + finally show "?fun1 h \ ?fun2 h" . + qed + thus "(\h. + norm (f (x + h) ** g (x + h) - f x ** g x - (f x ** G h + F h ** g x)) + / norm h) -- 0 --> 0" + by (simp only: FDERIV_lemma) +qed + +lemmas FDERIV_mult = bounded_bilinear_locale.mult.prod.FDERIV + +lemmas FDERIV_scaleR = bounded_bilinear_locale.scaleR.prod.FDERIV + + +subsection {* Powers *} + +lemma FDERIV_power_Suc: + fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}" + shows "FDERIV (\x. x ^ Suc n) x :> (\h. (1 + of_nat n) * x ^ n * h)" + apply (induct n) + apply (simp add: power_Suc FDERIV_ident) + apply (drule FDERIV_mult [OF FDERIV_ident]) + apply (simp only: of_nat_Suc left_distrib mult_1_left) + apply (simp only: power_Suc right_distrib add_ac mult_ac) +done + +lemma FDERIV_power: + fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}" + shows "FDERIV (\x. x ^ n) x :> (\h. of_nat n * x ^ (n - 1) * h)" + apply (cases n) + apply (simp add: FDERIV_const) + apply (simp add: FDERIV_power_Suc) + done + + +subsection {* Inverse *} + +lemma inverse_diff_inverse: + "\(a::'a::division_ring) \ 0; b \ 0\ + \ inverse a - inverse b = - (inverse a * (a - b) * inverse b)" +by (simp add: right_diff_distrib left_diff_distrib mult_assoc) + +lemmas bounded_linear_mult_const = + bounded_bilinear_locale.mult.prod.bounded_linear_left [THEN bounded_linear_compose] + +lemmas bounded_linear_const_mult = + bounded_bilinear_locale.mult.prod.bounded_linear_right [THEN bounded_linear_compose] + +lemma FDERIV_inverse: + fixes x :: "'a::real_normed_div_algebra" + assumes x: "x \ 0" + shows "FDERIV inverse x :> (\h. - (inverse x * h * inverse x))" + (is "FDERIV ?inv _ :> _") +proof (rule FDERIV_I) + show "bounded_linear (\h. - (?inv x * h * ?inv x))" + apply (rule bounded_linear_minus) + apply (rule bounded_linear_mult_const) + apply (rule bounded_linear_const_mult) + apply (rule bounded_linear_ident) + done +next + show "(\h. norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h) + -- 0 --> 0" + proof (rule LIM_equal2) + show "0 < norm x" using x by simp + next + fix h::'a + assume 1: "h \ 0" + assume "norm (h - 0) < norm x" + hence "h \ -x" by clarsimp + hence 2: "x + h \ 0" + apply (rule contrapos_nn) + apply (rule sym) + apply (erule equals_zero_I) + done + show "norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h + = norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" + apply (subst inverse_diff_inverse [OF 2 x]) + apply (subst minus_diff_minus) + apply (subst norm_minus_cancel) + apply (simp add: left_diff_distrib) + done + next + show "(\h. norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h) + -- 0 --> 0" + proof (rule real_LIM_sandwich_zero) + show "(\h. norm (?inv (x + h) - ?inv x) * norm (?inv x)) + -- 0 --> 0" + apply (rule LIM_mult_left_zero) + apply (rule LIM_norm_zero) + apply (rule LIM_zero) + apply (rule LIM_offset_zero) + apply (rule LIM_inverse) + apply (rule LIM_ident) + apply (rule x) + done + next + fix h::'a assume h: "h \ 0" + show "0 \ norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" + apply (rule divide_nonneg_pos) + apply (rule norm_ge_zero) + apply (simp add: h) + done + next + fix h::'a assume h: "h \ 0" + have "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h + \ norm (?inv (x + h) - ?inv x) * norm h * norm (?inv x) / norm h" + apply (rule divide_right_mono [OF _ norm_ge_zero]) + apply (rule order_trans [OF norm_mult_ineq]) + apply (rule mult_right_mono [OF _ norm_ge_zero]) + apply (rule norm_mult_ineq) + done + also have "\ = norm (?inv (x + h) - ?inv x) * norm (?inv x)" + by simp + finally show "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h + \ norm (?inv (x + h) - ?inv x) * norm (?inv x)" . + qed + qed +qed + +subsection {* Alternate definition *} + +lemma field_fderiv_def: + fixes x :: "'a::real_normed_field" shows + "FDERIV f x :> (\h. h * D) = (\h. (f (x + h) - f x) / h) -- 0 --> D" + apply (unfold fderiv_def) + apply (simp add: bounded_bilinear_locale.mult.prod.bounded_linear_left) + apply (simp cong: LIM_cong add: nonzero_norm_divide [symmetric]) + apply (subst diff_divide_distrib) + apply (subst times_divide_eq_left [symmetric]) + apply (simp cong: LIM_cong) + apply (simp add: LIM_norm_zero_iff LIM_zero_iff) +done + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/GCD.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/GCD.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,783 @@ +(* Title: HOL/GCD.thy + ID: $Id$ + Author: Christophe Tabacznyj and Lawrence C Paulson + Copyright 1996 University of Cambridge +*) + +header {* The Greatest Common Divisor *} + +theory GCD +imports Plain Presburger +begin + +text {* + See \cite{davenport92}. \bigskip +*} + +subsection {* Specification of GCD on nats *} + +definition + is_gcd :: "nat \ nat \ nat \ bool" where -- {* @{term gcd} as a relation *} + [code del]: "is_gcd m n p \ p dvd m \ p dvd n \ + (\d. d dvd m \ d dvd n \ d dvd p)" + +text {* Uniqueness *} + +lemma is_gcd_unique: "is_gcd a b m \ is_gcd a b n \ m = n" + by (simp add: is_gcd_def) (blast intro: dvd_anti_sym) + +text {* Connection to divides relation *} + +lemma is_gcd_dvd: "is_gcd a b m \ k dvd a \ k dvd b \ k dvd m" + by (auto simp add: is_gcd_def) + +text {* Commutativity *} + +lemma is_gcd_commute: "is_gcd m n k = is_gcd n m k" + by (auto simp add: is_gcd_def) + + +subsection {* GCD on nat by Euclid's algorithm *} + +fun + gcd :: "nat => nat => nat" +where + "gcd m n = (if n = 0 then m else gcd n (m mod n))" +lemma gcd_induct [case_names "0" rec]: + fixes m n :: nat + assumes "\m. P m 0" + and "\m n. 0 < n \ P n (m mod n) \ P m n" + shows "P m n" +proof (induct m n rule: gcd.induct) + case (1 m n) with assms show ?case by (cases "n = 0") simp_all +qed + +lemma gcd_0 [simp, algebra]: "gcd m 0 = m" + by simp + +lemma gcd_0_left [simp,algebra]: "gcd 0 m = m" + by simp + +lemma gcd_non_0: "n > 0 \ gcd m n = gcd n (m mod n)" + by simp + +lemma gcd_1 [simp, algebra]: "gcd m (Suc 0) = 1" + by simp + +declare gcd.simps [simp del] + +text {* + \medskip @{term "gcd m n"} divides @{text m} and @{text n}. The + conjunctions don't seem provable separately. +*} + +lemma gcd_dvd1 [iff, algebra]: "gcd m n dvd m" + and gcd_dvd2 [iff, algebra]: "gcd m n dvd n" + apply (induct m n rule: gcd_induct) + apply (simp_all add: gcd_non_0) + apply (blast dest: dvd_mod_imp_dvd) + done + +text {* + \medskip Maximality: for all @{term m}, @{term n}, @{term k} + naturals, if @{term k} divides @{term m} and @{term k} divides + @{term n} then @{term k} divides @{term "gcd m n"}. +*} + +lemma gcd_greatest: "k dvd m \ k dvd n \ k dvd gcd m n" + by (induct m n rule: gcd_induct) (simp_all add: gcd_non_0 dvd_mod) + +text {* + \medskip Function gcd yields the Greatest Common Divisor. +*} + +lemma is_gcd: "is_gcd m n (gcd m n) " + by (simp add: is_gcd_def gcd_greatest) + + +subsection {* Derived laws for GCD *} + +lemma gcd_greatest_iff [iff, algebra]: "k dvd gcd m n \ k dvd m \ k dvd n" + by (blast intro!: gcd_greatest intro: dvd_trans) + +lemma gcd_zero[algebra]: "gcd m n = 0 \ m = 0 \ n = 0" + by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff) + +lemma gcd_commute: "gcd m n = gcd n m" + apply (rule is_gcd_unique) + apply (rule is_gcd) + apply (subst is_gcd_commute) + apply (simp add: is_gcd) + done + +lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)" + apply (rule is_gcd_unique) + apply (rule is_gcd) + apply (simp add: is_gcd_def) + apply (blast intro: dvd_trans) + done + +lemma gcd_1_left [simp, algebra]: "gcd (Suc 0) m = 1" + by (simp add: gcd_commute) + +text {* + \medskip Multiplication laws +*} + +lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)" + -- {* \cite[page 27]{davenport92} *} + apply (induct m n rule: gcd_induct) + apply simp + apply (case_tac "k = 0") + apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2) + done + +lemma gcd_mult [simp, algebra]: "gcd k (k * n) = k" + apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric]) + done + +lemma gcd_self [simp, algebra]: "gcd k k = k" + apply (rule gcd_mult [of k 1, simplified]) + done + +lemma relprime_dvd_mult: "gcd k n = 1 ==> k dvd m * n ==> k dvd m" + apply (insert gcd_mult_distrib2 [of m k n]) + apply simp + apply (erule_tac t = m in ssubst) + apply simp + done + +lemma relprime_dvd_mult_iff: "gcd k n = 1 ==> (k dvd m * n) = (k dvd m)" + by (auto intro: relprime_dvd_mult dvd_mult2) + +lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n" + apply (rule dvd_anti_sym) + apply (rule gcd_greatest) + apply (rule_tac n = k in relprime_dvd_mult) + apply (simp add: gcd_assoc) + apply (simp add: gcd_commute) + apply (simp_all add: mult_commute) + apply (blast intro: dvd_mult) + done + + +text {* \medskip Addition laws *} + +lemma gcd_add1 [simp, algebra]: "gcd (m + n) n = gcd m n" + by (cases "n = 0") (auto simp add: gcd_non_0) + +lemma gcd_add2 [simp, algebra]: "gcd m (m + n) = gcd m n" +proof - + have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute) + also have "... = gcd (n + m) m" by (simp add: add_commute) + also have "... = gcd n m" by simp + also have "... = gcd m n" by (rule gcd_commute) + finally show ?thesis . +qed + +lemma gcd_add2' [simp, algebra]: "gcd m (n + m) = gcd m n" + apply (subst add_commute) + apply (rule gcd_add2) + done + +lemma gcd_add_mult[algebra]: "gcd m (k * m + n) = gcd m n" + by (induct k) (simp_all add: add_assoc) + +lemma gcd_dvd_prod: "gcd m n dvd m * n" + using mult_dvd_mono [of 1] by auto + +text {* + \medskip Division by gcd yields rrelatively primes. +*} + +lemma div_gcd_relprime: + assumes nz: "a \ 0 \ b \ 0" + shows "gcd (a div gcd a b) (b div gcd a b) = 1" +proof - + let ?g = "gcd a b" + let ?a' = "a div ?g" + let ?b' = "b div ?g" + let ?g' = "gcd ?a' ?b'" + have dvdg: "?g dvd a" "?g dvd b" by simp_all + have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all + from dvdg dvdg' obtain ka kb ka' kb' where + kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'" + unfolding dvd_def by blast + then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all + then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b" + by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)] + dvd_mult_div_cancel [OF dvdg(2)] dvd_def) + have "?g \ 0" using nz by (simp add: gcd_zero) + then have gp: "?g > 0" by simp + from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" . + with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp +qed + + +lemma gcd_unique: "d dvd a\d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" +proof(auto) + assume H: "d dvd a" "d dvd b" "\e. e dvd a \ e dvd b \ e dvd d" + from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b] + have th: "gcd a b dvd d" by blast + from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]] show "d = gcd a b" by blast +qed + +lemma gcd_eq: assumes H: "\d. d dvd x \ d dvd y \ d dvd u \ d dvd v" + shows "gcd x y = gcd u v" +proof- + from H have "\d. d dvd x \ d dvd y \ d dvd gcd u v" by simp + with gcd_unique[of "gcd u v" x y] show ?thesis by auto +qed + +lemma ind_euclid: + assumes c: " \a b. P (a::nat) b \ P b a" and z: "\a. P a 0" + and add: "\a b. P a b \ P a (a + b)" + shows "P a b" +proof(induct n\"a+b" arbitrary: a b rule: nat_less_induct) + fix n a b + assume H: "\m < n. \a b. m = a + b \ P a b" "n = a + b" + have "a = b \ a < b \ b < a" by arith + moreover {assume eq: "a= b" + from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq by simp} + moreover + {assume lt: "a < b" + hence "a + b - a < n \ a = 0" using H(2) by arith + moreover + {assume "a =0" with z c have "P a b" by blast } + moreover + {assume ab: "a + b - a < n" + have th0: "a + b - a = a + (b - a)" using lt by arith + from add[rule_format, OF H(1)[rule_format, OF ab th0]] + have "P a b" by (simp add: th0[symmetric])} + ultimately have "P a b" by blast} + moreover + {assume lt: "a > b" + hence "b + a - b < n \ b = 0" using H(2) by arith + moreover + {assume "b =0" with z c have "P a b" by blast } + moreover + {assume ab: "b + a - b < n" + have th0: "b + a - b = b + (a - b)" using lt by arith + from add[rule_format, OF H(1)[rule_format, OF ab th0]] + have "P b a" by (simp add: th0[symmetric]) + hence "P a b" using c by blast } + ultimately have "P a b" by blast} +ultimately show "P a b" by blast +qed + +lemma bezout_lemma: + assumes ex: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x = b * y + d \ b * x = a * y + d)" + shows "\d x y. d dvd a \ d dvd a + b \ (a * x = (a + b) * y + d \ (a + b) * x = a * y + d)" +using ex +apply clarsimp +apply (rule_tac x="d" in exI, simp add: dvd_add) +apply (case_tac "a * x = b * y + d" , simp_all) +apply (rule_tac x="x + y" in exI) +apply (rule_tac x="y" in exI) +apply algebra +apply (rule_tac x="x" in exI) +apply (rule_tac x="x + y" in exI) +apply algebra +done + +lemma bezout_add: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x = b * y + d \ b * x = a * y + d)" +apply(induct a b rule: ind_euclid) +apply blast +apply clarify +apply (rule_tac x="a" in exI, simp add: dvd_add) +apply clarsimp +apply (rule_tac x="d" in exI) +apply (case_tac "a * x = b * y + d", simp_all add: dvd_add) +apply (rule_tac x="x+y" in exI) +apply (rule_tac x="y" in exI) +apply algebra +apply (rule_tac x="x" in exI) +apply (rule_tac x="x+y" in exI) +apply algebra +done + +lemma bezout: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x - b * y = d \ b * x - a * y = d)" +using bezout_add[of a b] +apply clarsimp +apply (rule_tac x="d" in exI, simp) +apply (rule_tac x="x" in exI) +apply (rule_tac x="y" in exI) +apply auto +done + + +text {* We can get a stronger version with a nonzeroness assumption. *} +lemma divides_le: "m dvd n ==> m <= n \ n = (0::nat)" by (auto simp add: dvd_def) + +lemma bezout_add_strong: assumes nz: "a \ (0::nat)" + shows "\d x y. d dvd a \ d dvd b \ a * x = b * y + d" +proof- + from nz have ap: "a > 0" by simp + from bezout_add[of a b] + have "(\d x y. d dvd a \ d dvd b \ a * x = b * y + d) \ (\d x y. d dvd a \ d dvd b \ b * x = a * y + d)" by blast + moreover + {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d" + from H have ?thesis by blast } + moreover + {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d" + {assume b0: "b = 0" with H have ?thesis by simp} + moreover + {assume b: "b \ 0" hence bp: "b > 0" by simp + from divides_le[OF H(2)] b have "d < b \ d = b" using le_less by blast + moreover + {assume db: "d=b" + from prems have ?thesis apply simp + apply (rule exI[where x = b], simp) + apply (rule exI[where x = b]) + by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)} + moreover + {assume db: "d < b" + {assume "x=0" hence ?thesis using prems by simp } + moreover + {assume x0: "x \ 0" hence xp: "x > 0" by simp + + from db have "d \ b - 1" by simp + hence "d*b \ b*(b - 1)" by simp + with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"] + have dble: "d*b \ x*b*(b - 1)" using bp by simp + from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra + hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp + hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)" + by (simp only: diff_add_assoc[OF dble, of d, symmetric]) + hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d" + by (simp only: diff_mult_distrib2 add_commute mult_ac) + hence ?thesis using H(1,2) + apply - + apply (rule exI[where x=d], simp) + apply (rule exI[where x="(b - 1) * y"]) + by (rule exI[where x="x*(b - 1) - d"], simp)} + ultimately have ?thesis by blast} + ultimately have ?thesis by blast} + ultimately have ?thesis by blast} + ultimately show ?thesis by blast +qed + + +lemma bezout_gcd: "\x y. a * x - b * y = gcd a b \ b * x - a * y = gcd a b" +proof- + let ?g = "gcd a b" + from bezout[of a b] obtain d x y where d: "d dvd a" "d dvd b" "a * x - b * y = d \ b * x - a * y = d" by blast + from d(1,2) have "d dvd ?g" by simp + then obtain k where k: "?g = d*k" unfolding dvd_def by blast + from d(3) have "(a * x - b * y)*k = d*k \ (b * x - a * y)*k = d*k" by blast + hence "a * x * k - b * y*k = d*k \ b * x * k - a * y*k = d*k" + by (algebra add: diff_mult_distrib) + hence "a * (x * k) - b * (y*k) = ?g \ b * (x * k) - a * (y*k) = ?g" + by (simp add: k mult_assoc) + thus ?thesis by blast +qed + +lemma bezout_gcd_strong: assumes a: "a \ 0" + shows "\x y. a * x = b * y + gcd a b" +proof- + let ?g = "gcd a b" + from bezout_add_strong[OF a, of b] + obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast + from d(1,2) have "d dvd ?g" by simp + then obtain k where k: "?g = d*k" unfolding dvd_def by blast + from d(3) have "a * x * k = (b * y + d) *k " by algebra + hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k) + thus ?thesis by blast +qed + +lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b" +by(simp add: gcd_mult_distrib2 mult_commute) + +lemma gcd_bezout: "(\x y. a * x - b * y = d \ b * x - a * y = d) \ gcd a b dvd d" + (is "?lhs \ ?rhs") +proof- + let ?g = "gcd a b" + {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast + from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \ b * x - a * y = ?g" + by blast + hence "(a * x - b * y)*k = ?g*k \ (b * x - a * y)*k = ?g*k" by auto + hence "a * x*k - b * y*k = ?g*k \ b * x * k - a * y*k = ?g*k" + by (simp only: diff_mult_distrib) + hence "a * (x*k) - b * (y*k) = d \ b * (x * k) - a * (y*k) = d" + by (simp add: k[symmetric] mult_assoc) + hence ?lhs by blast} + moreover + {fix x y assume H: "a * x - b * y = d \ b * x - a * y = d" + have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y" + using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all + from dvd_diff[OF dv(1,2)] dvd_diff[OF dv(3,4)] H + have ?rhs by auto} + ultimately show ?thesis by blast +qed + +lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd a b dvd d" +proof- + let ?g = "gcd a b" + have dv: "?g dvd a*x" "?g dvd b * y" + using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all + from dvd_add[OF dv] H + show ?thesis by auto +qed + +lemma gcd_mult': "gcd b (a * b) = b" +by (simp add: gcd_mult mult_commute[of a b]) + +lemma gcd_add: "gcd(a + b) b = gcd a b" + "gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b" +apply (simp_all add: gcd_add1) +by (simp add: gcd_commute gcd_add1) + +lemma gcd_sub: "b <= a ==> gcd(a - b) b = gcd a b" "a <= b ==> gcd a (b - a) = gcd a b" +proof- + {fix a b assume H: "b \ (a::nat)" + hence th: "a - b + b = a" by arith + from gcd_add(1)[of "a - b" b] th have "gcd(a - b) b = gcd a b" by simp} + note th = this +{ + assume ab: "b \ a" + from th[OF ab] show "gcd (a - b) b = gcd a b" by blast +next + assume ab: "a \ b" + from th[OF ab] show "gcd a (b - a) = gcd a b" + by (simp add: gcd_commute)} +qed + + +subsection {* LCM defined by GCD *} + + +definition + lcm :: "nat \ nat \ nat" +where + lcm_def: "lcm m n = m * n div gcd m n" + +lemma prod_gcd_lcm: + "m * n = gcd m n * lcm m n" + unfolding lcm_def by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod]) + +lemma lcm_0 [simp]: "lcm m 0 = 0" + unfolding lcm_def by simp + +lemma lcm_1 [simp]: "lcm m 1 = m" + unfolding lcm_def by simp + +lemma lcm_0_left [simp]: "lcm 0 n = 0" + unfolding lcm_def by simp + +lemma lcm_1_left [simp]: "lcm 1 m = m" + unfolding lcm_def by simp + +lemma dvd_pos: + fixes n m :: nat + assumes "n > 0" and "m dvd n" + shows "m > 0" +using assms by (cases m) auto + +lemma lcm_least: + assumes "m dvd k" and "n dvd k" + shows "lcm m n dvd k" +proof (cases k) + case 0 then show ?thesis by auto +next + case (Suc _) then have pos_k: "k > 0" by auto + from assms dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto + with gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp + from assms obtain p where k_m: "k = m * p" using dvd_def by blast + from assms obtain q where k_n: "k = n * q" using dvd_def by blast + from pos_k k_m have pos_p: "p > 0" by auto + from pos_k k_n have pos_q: "q > 0" by auto + have "k * k * gcd q p = k * gcd (k * q) (k * p)" + by (simp add: mult_ac gcd_mult_distrib2) + also have "\ = k * gcd (m * p * q) (n * q * p)" + by (simp add: k_m [symmetric] k_n [symmetric]) + also have "\ = k * p * q * gcd m n" + by (simp add: mult_ac gcd_mult_distrib2) + finally have "(m * p) * (n * q) * gcd q p = k * p * q * gcd m n" + by (simp only: k_m [symmetric] k_n [symmetric]) + then have "p * q * m * n * gcd q p = p * q * k * gcd m n" + by (simp add: mult_ac) + with pos_p pos_q have "m * n * gcd q p = k * gcd m n" + by simp + with prod_gcd_lcm [of m n] + have "lcm m n * gcd q p * gcd m n = k * gcd m n" + by (simp add: mult_ac) + with pos_gcd have "lcm m n * gcd q p = k" by simp + then show ?thesis using dvd_def by auto +qed + +lemma lcm_dvd1 [iff]: + "m dvd lcm m n" +proof (cases m) + case 0 then show ?thesis by simp +next + case (Suc _) + then have mpos: "m > 0" by simp + show ?thesis + proof (cases n) + case 0 then show ?thesis by simp + next + case (Suc _) + then have npos: "n > 0" by simp + have "gcd m n dvd n" by simp + then obtain k where "n = gcd m n * k" using dvd_def by auto + then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n" by (simp add: mult_ac) + also have "\ = m * k" using mpos npos gcd_zero by simp + finally show ?thesis by (simp add: lcm_def) + qed +qed + +lemma lcm_dvd2 [iff]: + "n dvd lcm m n" +proof (cases n) + case 0 then show ?thesis by simp +next + case (Suc _) + then have npos: "n > 0" by simp + show ?thesis + proof (cases m) + case 0 then show ?thesis by simp + next + case (Suc _) + then have mpos: "m > 0" by simp + have "gcd m n dvd m" by simp + then obtain k where "m = gcd m n * k" using dvd_def by auto + then have "m * n div gcd m n = (gcd m n * k) * n div gcd m n" by (simp add: mult_ac) + also have "\ = n * k" using mpos npos gcd_zero by simp + finally show ?thesis by (simp add: lcm_def) + qed +qed + +lemma gcd_add1_eq: "gcd (m + k) k = gcd (m + k) m" + by (simp add: gcd_commute) + +lemma gcd_diff2: "m \ n ==> gcd n (n - m) = gcd n m" + apply (subgoal_tac "n = m + (n - m)") + apply (erule ssubst, rule gcd_add1_eq, simp) + done + + +subsection {* GCD and LCM on integers *} + +definition + zgcd :: "int \ int \ int" where + "zgcd i j = int (gcd (nat (abs i)) (nat (abs j)))" + +lemma zgcd_zdvd1 [iff,simp, algebra]: "zgcd i j dvd i" + by (simp add: zgcd_def int_dvd_iff) + +lemma zgcd_zdvd2 [iff,simp, algebra]: "zgcd i j dvd j" + by (simp add: zgcd_def int_dvd_iff) + +lemma zgcd_pos: "zgcd i j \ 0" + by (simp add: zgcd_def) + +lemma zgcd0 [simp,algebra]: "(zgcd i j = 0) = (i = 0 \ j = 0)" + by (simp add: zgcd_def gcd_zero) arith + +lemma zgcd_commute: "zgcd i j = zgcd j i" + unfolding zgcd_def by (simp add: gcd_commute) + +lemma zgcd_zminus [simp, algebra]: "zgcd (- i) j = zgcd i j" + unfolding zgcd_def by simp + +lemma zgcd_zminus2 [simp, algebra]: "zgcd i (- j) = zgcd i j" + unfolding zgcd_def by simp + + (* should be solved by algebra*) +lemma zrelprime_dvd_mult: "zgcd i j = 1 \ i dvd k * j \ i dvd k" + unfolding zgcd_def +proof - + assume "int (gcd (nat \i\) (nat \j\)) = 1" "i dvd k * j" + then have g: "gcd (nat \i\) (nat \j\) = 1" by simp + from `i dvd k * j` obtain h where h: "k*j = i*h" unfolding dvd_def by blast + have th: "nat \i\ dvd nat \k\ * nat \j\" + unfolding dvd_def + by (rule_tac x= "nat \h\" in exI, simp add: h nat_abs_mult_distrib [symmetric]) + from relprime_dvd_mult [OF g th] obtain h' where h': "nat \k\ = nat \i\ * h'" + unfolding dvd_def by blast + from h' have "int (nat \k\) = int (nat \i\ * h')" by simp + then have "\k\ = \i\ * int h'" by (simp add: int_mult) + then show ?thesis + apply (subst zdvd_abs1 [symmetric]) + apply (subst zdvd_abs2 [symmetric]) + apply (unfold dvd_def) + apply (rule_tac x = "int h'" in exI, simp) + done +qed + +lemma int_nat_abs: "int (nat (abs x)) = abs x" by arith + +lemma zgcd_greatest: + assumes "k dvd m" and "k dvd n" + shows "k dvd zgcd m n" +proof - + let ?k' = "nat \k\" + let ?m' = "nat \m\" + let ?n' = "nat \n\" + from `k dvd m` and `k dvd n` have dvd': "?k' dvd ?m'" "?k' dvd ?n'" + unfolding zdvd_int by (simp_all only: int_nat_abs zdvd_abs1 zdvd_abs2) + from gcd_greatest [OF dvd'] have "int (nat \k\) dvd zgcd m n" + unfolding zgcd_def by (simp only: zdvd_int) + then have "\k\ dvd zgcd m n" by (simp only: int_nat_abs) + then show "k dvd zgcd m n" by (simp add: zdvd_abs1) +qed + +lemma div_zgcd_relprime: + assumes nz: "a \ 0 \ b \ 0" + shows "zgcd (a div (zgcd a b)) (b div (zgcd a b)) = 1" +proof - + from nz have nz': "nat \a\ \ 0 \ nat \b\ \ 0" by arith + let ?g = "zgcd a b" + let ?a' = "a div ?g" + let ?b' = "b div ?g" + let ?g' = "zgcd ?a' ?b'" + have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2) + have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2) + from dvdg dvdg' obtain ka kb ka' kb' where + kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'" + unfolding dvd_def by blast + then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all + then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b" + by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)] + zdvd_mult_div_cancel [OF dvdg(2)] dvd_def) + have "?g \ 0" using nz by simp + then have gp: "?g \ 0" using zgcd_pos[where i="a" and j="b"] by arith + from zgcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" . + with zdvd_mult_cancel1 [OF gp] have "\?g'\ = 1" by simp + with zgcd_pos show "?g' = 1" by simp +qed + +lemma zgcd_0 [simp, algebra]: "zgcd m 0 = abs m" + by (simp add: zgcd_def abs_if) + +lemma zgcd_0_left [simp, algebra]: "zgcd 0 m = abs m" + by (simp add: zgcd_def abs_if) + +lemma zgcd_non_0: "0 < n ==> zgcd m n = zgcd n (m mod n)" + apply (frule_tac b = n and a = m in pos_mod_sign) + apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib) + apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if) + apply (frule_tac a = m in pos_mod_bound) + apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle) + done + +lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)" + apply (case_tac "n = 0", simp add: DIVISION_BY_ZERO) + apply (auto simp add: linorder_neq_iff zgcd_non_0) + apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0, auto) + done + +lemma zgcd_1 [simp, algebra]: "zgcd m 1 = 1" + by (simp add: zgcd_def abs_if) + +lemma zgcd_0_1_iff [simp, algebra]: "zgcd 0 m = 1 \ \m\ = 1" + by (simp add: zgcd_def abs_if) + +lemma zgcd_greatest_iff[algebra]: "k dvd zgcd m n = (k dvd m \ k dvd n)" + by (simp add: zgcd_def abs_if int_dvd_iff dvd_int_iff nat_dvd_iff) + +lemma zgcd_1_left [simp, algebra]: "zgcd 1 m = 1" + by (simp add: zgcd_def gcd_1_left) + +lemma zgcd_assoc: "zgcd (zgcd k m) n = zgcd k (zgcd m n)" + by (simp add: zgcd_def gcd_assoc) + +lemma zgcd_left_commute: "zgcd k (zgcd m n) = zgcd m (zgcd k n)" + apply (rule zgcd_commute [THEN trans]) + apply (rule zgcd_assoc [THEN trans]) + apply (rule zgcd_commute [THEN arg_cong]) + done + +lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute + -- {* addition is an AC-operator *} + +lemma zgcd_zmult_distrib2: "0 \ k ==> k * zgcd m n = zgcd (k * m) (k * n)" + by (simp del: minus_mult_right [symmetric] + add: minus_mult_right nat_mult_distrib zgcd_def abs_if + mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric]) + +lemma zgcd_zmult_distrib2_abs: "zgcd (k * m) (k * n) = abs k * zgcd m n" + by (simp add: abs_if zgcd_zmult_distrib2) + +lemma zgcd_self [simp]: "0 \ m ==> zgcd m m = m" + by (cut_tac k = m and m = 1 and n = 1 in zgcd_zmult_distrib2, simp_all) + +lemma zgcd_zmult_eq_self [simp]: "0 \ k ==> zgcd k (k * n) = k" + by (cut_tac k = k and m = 1 and n = n in zgcd_zmult_distrib2, simp_all) + +lemma zgcd_zmult_eq_self2 [simp]: "0 \ k ==> zgcd (k * n) k = k" + by (cut_tac k = k and m = n and n = 1 in zgcd_zmult_distrib2, simp_all) + + +definition "zlcm i j = int (lcm(nat(abs i)) (nat(abs j)))" + +lemma dvd_zlcm_self1[simp, algebra]: "i dvd zlcm i j" +by(simp add:zlcm_def dvd_int_iff) + +lemma dvd_zlcm_self2[simp, algebra]: "j dvd zlcm i j" +by(simp add:zlcm_def dvd_int_iff) + + +lemma dvd_imp_dvd_zlcm1: + assumes "k dvd i" shows "k dvd (zlcm i j)" +proof - + have "nat(abs k) dvd nat(abs i)" using `k dvd i` + by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric] zdvd_abs1) + thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans) +qed + +lemma dvd_imp_dvd_zlcm2: + assumes "k dvd j" shows "k dvd (zlcm i j)" +proof - + have "nat(abs k) dvd nat(abs j)" using `k dvd j` + by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric] zdvd_abs1) + thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans) +qed + + +lemma zdvd_self_abs1: "(d::int) dvd (abs d)" +by (case_tac "d <0", simp_all) + +lemma zdvd_self_abs2: "(abs (d::int)) dvd d" +by (case_tac "d<0", simp_all) + +(* lcm a b is positive for positive a and b *) + +lemma lcm_pos: + assumes mpos: "m > 0" + and npos: "n>0" + shows "lcm m n > 0" +proof(rule ccontr, simp add: lcm_def gcd_zero) +assume h:"m*n div gcd m n = 0" +from mpos npos have "gcd m n \ 0" using gcd_zero by simp +hence gcdp: "gcd m n > 0" by simp +with h +have "m*n < gcd m n" + by (cases "m * n < gcd m n") (auto simp add: div_if[OF gcdp, where m="m*n"]) +moreover +have "gcd m n dvd m" by simp + with mpos dvd_imp_le have t1:"gcd m n \ m" by simp + with npos have t1:"gcd m n *n \ m*n" by simp + have "gcd m n \ gcd m n*n" using npos by simp + with t1 have "gcd m n \ m*n" by arith +ultimately show "False" by simp +qed + +lemma zlcm_pos: + assumes anz: "a \ 0" + and bnz: "b \ 0" + shows "0 < zlcm a b" +proof- + let ?na = "nat (abs a)" + let ?nb = "nat (abs b)" + have nap: "?na >0" using anz by simp + have nbp: "?nb >0" using bnz by simp + have "0 < lcm ?na ?nb" by (rule lcm_pos[OF nap nbp]) + thus ?thesis by (simp add: zlcm_def) +qed + +lemma zgcd_code [code]: + "zgcd k l = \if l = 0 then k else zgcd l (\k\ mod \l\)\" + by (simp add: zgcd_def gcd.simps [of "nat \k\"] nat_mod_distrib) + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/HOL.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/HOL.thy - ID: $Id$ Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson *) @@ -8,7 +7,7 @@ theory HOL imports Pure uses - ("hologic.ML") + ("Tools/hologic.ML") "~~/src/Tools/IsaPlanner/zipper.ML" "~~/src/Tools/IsaPlanner/isand.ML" "~~/src/Tools/IsaPlanner/rw_tools.ML" @@ -22,7 +21,7 @@ "~~/src/Provers/coherent.ML" "~~/src/Provers/eqsubst.ML" "~~/src/Provers/quantifier1.ML" - ("simpdata.ML") + ("Tools/simpdata.ML") "~~/src/Tools/random_word.ML" "~~/src/Tools/atomize_elim.ML" "~~/src/Tools/induct.ML" @@ -801,7 +800,7 @@ and [sym] = sym not_sym and [Pure.elim?] = iffD1 iffD2 impE -use "hologic.ML" +use "Tools/hologic.ML" subsubsection {* Atomizing meta-level connectives *} @@ -1285,7 +1284,7 @@ "(\x y. P x y) = (\y x. P x y)" by blast -use "simpdata.ML" +use "Tools/simpdata.ML" ML {* open Simpdata *} setup {* diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/Deriv.thy --- a/src/HOL/Hyperreal/Deriv.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1725 +0,0 @@ -(* Title : Deriv.thy - ID : $Id$ - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Conversion to Isar and new proofs by Lawrence C Paulson, 2004 - GMVT by Benjamin Porter, 2005 -*) - -header{* Differentiation *} - -theory Deriv -imports Lim Univ_Poly -begin - -text{*Standard Definitions*} - -definition - deriv :: "['a::real_normed_field \ 'a, 'a, 'a] \ bool" - --{*Differentiation: D is derivative of function f at x*} - ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where - "DERIV f x :> D = ((%h. (f(x + h) - f x) / h) -- 0 --> D)" - -definition - differentiable :: "['a::real_normed_field \ 'a, 'a] \ bool" - (infixl "differentiable" 60) where - "f differentiable x = (\D. DERIV f x :> D)" - - -consts - Bolzano_bisect :: "[real*real=>bool, real, real, nat] => (real*real)" -primrec - "Bolzano_bisect P a b 0 = (a,b)" - "Bolzano_bisect P a b (Suc n) = - (let (x,y) = Bolzano_bisect P a b n - in if P(x, (x+y)/2) then ((x+y)/2, y) - else (x, (x+y)/2))" - - -subsection {* Derivatives *} - -lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --> D)" -by (simp add: deriv_def) - -lemma DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --> D" -by (simp add: deriv_def) - -lemma DERIV_const [simp]: "DERIV (\x. k) x :> 0" -by (simp add: deriv_def) - -lemma DERIV_ident [simp]: "DERIV (\x. x) x :> 1" -by (simp add: deriv_def cong: LIM_cong) - -lemma add_diff_add: - fixes a b c d :: "'a::ab_group_add" - shows "(a + c) - (b + d) = (a - b) + (c - d)" -by simp - -lemma DERIV_add: - "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x + g x) x :> D + E" -by (simp only: deriv_def add_diff_add add_divide_distrib LIM_add) - -lemma DERIV_minus: - "DERIV f x :> D \ DERIV (\x. - f x) x :> - D" -by (simp only: deriv_def minus_diff_minus divide_minus_left LIM_minus) - -lemma DERIV_diff: - "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x - g x) x :> D - E" -by (simp only: diff_def DERIV_add DERIV_minus) - -lemma DERIV_add_minus: - "\DERIV f x :> D; DERIV g x :> E\ \ DERIV (\x. f x + - g x) x :> D + - E" -by (simp only: DERIV_add DERIV_minus) - -lemma DERIV_isCont: "DERIV f x :> D \ isCont f x" -proof (unfold isCont_iff) - assume "DERIV f x :> D" - hence "(\h. (f(x+h) - f(x)) / h) -- 0 --> D" - by (rule DERIV_D) - hence "(\h. (f(x+h) - f(x)) / h * h) -- 0 --> D * 0" - by (intro LIM_mult LIM_ident) - hence "(\h. (f(x+h) - f(x)) * (h / h)) -- 0 --> 0" - by simp - hence "(\h. f(x+h) - f(x)) -- 0 --> 0" - by (simp cong: LIM_cong) - thus "(\h. f(x+h)) -- 0 --> f(x)" - by (simp add: LIM_def) -qed - -lemma DERIV_mult_lemma: - fixes a b c d :: "'a::real_field" - shows "(a * b - c * d) / h = a * ((b - d) / h) + ((a - c) / h) * d" -by (simp add: diff_minus add_divide_distrib [symmetric] ring_distribs) - -lemma DERIV_mult': - assumes f: "DERIV f x :> D" - assumes g: "DERIV g x :> E" - shows "DERIV (\x. f x * g x) x :> f x * E + D * g x" -proof (unfold deriv_def) - from f have "isCont f x" - by (rule DERIV_isCont) - hence "(\h. f(x+h)) -- 0 --> f x" - by (simp only: isCont_iff) - hence "(\h. f(x+h) * ((g(x+h) - g x) / h) + - ((f(x+h) - f x) / h) * g x) - -- 0 --> f x * E + D * g x" - by (intro LIM_add LIM_mult LIM_const DERIV_D f g) - thus "(\h. (f(x+h) * g(x+h) - f x * g x) / h) - -- 0 --> f x * E + D * g x" - by (simp only: DERIV_mult_lemma) -qed - -lemma DERIV_mult: - "[| DERIV f x :> Da; DERIV g x :> Db |] - ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" -by (drule (1) DERIV_mult', simp only: mult_commute add_commute) - -lemma DERIV_unique: - "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E" -apply (simp add: deriv_def) -apply (blast intro: LIM_unique) -done - -text{*Differentiation of finite sum*} - -lemma DERIV_sumr [rule_format (no_asm)]: - "(\r. m \ r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x)) - --> DERIV (%x. \n=m.. (\r=m.. D) = - ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)" -apply (rule iffI) -apply (drule_tac k="- a" in LIM_offset) -apply (simp add: diff_minus) -apply (drule_tac k="a" in LIM_offset) -apply (simp add: add_commute) -done - -lemma DERIV_iff2: "(DERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --> D)" -by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff) - -lemma inverse_diff_inverse: - "\(a::'a::division_ring) \ 0; b \ 0\ - \ inverse a - inverse b = - (inverse a * (a - b) * inverse b)" -by (simp add: ring_simps) - -lemma DERIV_inverse_lemma: - "\a \ 0; b \ (0::'a::real_normed_field)\ - \ (inverse a - inverse b) / h - = - (inverse a * ((a - b) / h) * inverse b)" -by (simp add: inverse_diff_inverse) - -lemma DERIV_inverse': - assumes der: "DERIV f x :> D" - assumes neq: "f x \ 0" - shows "DERIV (\x. inverse (f x)) x :> - (inverse (f x) * D * inverse (f x))" - (is "DERIV _ _ :> ?E") -proof (unfold DERIV_iff2) - from der have lim_f: "f -- x --> f x" - by (rule DERIV_isCont [unfolded isCont_def]) - - from neq have "0 < norm (f x)" by simp - with LIM_D [OF lim_f] obtain s - where s: "0 < s" - and less_fx: "\z. \z \ x; norm (z - x) < s\ - \ norm (f z - f x) < norm (f x)" - by fast - - show "(\z. (inverse (f z) - inverse (f x)) / (z - x)) -- x --> ?E" - proof (rule LIM_equal2 [OF s]) - fix z - assume "z \ x" "norm (z - x) < s" - hence "norm (f z - f x) < norm (f x)" by (rule less_fx) - hence "f z \ 0" by auto - thus "(inverse (f z) - inverse (f x)) / (z - x) = - - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x))" - using neq by (rule DERIV_inverse_lemma) - next - from der have "(\z. (f z - f x) / (z - x)) -- x --> D" - by (unfold DERIV_iff2) - thus "(\z. - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x))) - -- x --> ?E" - by (intro LIM_mult LIM_inverse LIM_minus LIM_const lim_f neq) - qed -qed - -lemma DERIV_divide: - "\DERIV f x :> D; DERIV g x :> E; g x \ 0\ - \ DERIV (\x. f x / g x) x :> (D * g x - f x * E) / (g x * g x)" -apply (subgoal_tac "f x * - (inverse (g x) * E * inverse (g x)) + - D * inverse (g x) = (D * g x - f x * E) / (g x * g x)") -apply (erule subst) -apply (unfold divide_inverse) -apply (erule DERIV_mult') -apply (erule (1) DERIV_inverse') -apply (simp add: ring_distribs nonzero_inverse_mult_distrib) -apply (simp add: mult_ac) -done - -lemma DERIV_power_Suc: - fixes f :: "'a \ 'a::{real_normed_field,recpower}" - assumes f: "DERIV f x :> D" - shows "DERIV (\x. f x ^ Suc n) x :> (1 + of_nat n) * (D * f x ^ n)" -proof (induct n) -case 0 - show ?case by (simp add: power_Suc f) -case (Suc k) - from DERIV_mult' [OF f Suc] show ?case - apply (simp only: of_nat_Suc ring_distribs mult_1_left) - apply (simp only: power_Suc right_distrib mult_ac add_ac) - done -qed - -lemma DERIV_power: - fixes f :: "'a \ 'a::{real_normed_field,recpower}" - assumes f: "DERIV f x :> D" - shows "DERIV (\x. f x ^ n) x :> of_nat n * (D * f x ^ (n - Suc 0))" -by (cases "n", simp, simp add: DERIV_power_Suc f) - - -(* ------------------------------------------------------------------------ *) -(* Caratheodory formulation of derivative at a point: standard proof *) -(* ------------------------------------------------------------------------ *) - -lemma CARAT_DERIV: - "(DERIV f x :> l) = - (\g. (\z. f z - f x = g z * (z-x)) & isCont g x & g x = l)" - (is "?lhs = ?rhs") -proof - assume der: "DERIV f x :> l" - show "\g. (\z. f z - f x = g z * (z-x)) \ isCont g x \ g x = l" - proof (intro exI conjI) - let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" - show "\z. f z - f x = ?g z * (z-x)" by simp - show "isCont ?g x" using der - by (simp add: isCont_iff DERIV_iff diff_minus - cong: LIM_equal [rule_format]) - show "?g x = l" by simp - qed -next - assume "?rhs" - then obtain g where - "(\z. f z - f x = g z * (z-x))" and "isCont g x" and "g x = l" by blast - thus "(DERIV f x :> l)" - by (auto simp add: isCont_iff DERIV_iff cong: LIM_cong) -qed - -lemma DERIV_chain': - assumes f: "DERIV f x :> D" - assumes g: "DERIV g (f x) :> E" - shows "DERIV (\x. g (f x)) x :> E * D" -proof (unfold DERIV_iff2) - obtain d where d: "\y. g y - g (f x) = d y * (y - f x)" - and cont_d: "isCont d (f x)" and dfx: "d (f x) = E" - using CARAT_DERIV [THEN iffD1, OF g] by fast - from f have "f -- x --> f x" - by (rule DERIV_isCont [unfolded isCont_def]) - with cont_d have "(\z. d (f z)) -- x --> d (f x)" - by (rule isCont_LIM_compose) - hence "(\z. d (f z) * ((f z - f x) / (z - x))) - -- x --> d (f x) * D" - by (rule LIM_mult [OF _ f [unfolded DERIV_iff2]]) - thus "(\z. (g (f z) - g (f x)) / (z - x)) -- x --> E * D" - by (simp add: d dfx real_scaleR_def) -qed - -(* let's do the standard proof though theorem *) -(* LIM_mult2 follows from a NS proof *) - -lemma DERIV_cmult: - "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D" -by (drule DERIV_mult' [OF DERIV_const], simp) - -(* standard version *) -lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db" -by (drule (1) DERIV_chain', simp add: o_def real_scaleR_def mult_commute) - -lemma DERIV_chain2: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db" -by (auto dest: DERIV_chain simp add: o_def) - -(*derivative of linear multiplication*) -lemma DERIV_cmult_Id [simp]: "DERIV (op * c) x :> c" -by (cut_tac c = c and x = x in DERIV_ident [THEN DERIV_cmult], simp) - -lemma DERIV_pow: "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" -apply (cut_tac DERIV_power [OF DERIV_ident]) -apply (simp add: real_scaleR_def real_of_nat_def) -done - -text{*Power of -1*} - -lemma DERIV_inverse: - fixes x :: "'a::{real_normed_field,recpower}" - shows "x \ 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))" -by (drule DERIV_inverse' [OF DERIV_ident]) (simp add: power_Suc) - -text{*Derivative of inverse*} -lemma DERIV_inverse_fun: - fixes x :: "'a::{real_normed_field,recpower}" - shows "[| DERIV f x :> d; f(x) \ 0 |] - ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" -by (drule (1) DERIV_inverse') (simp add: mult_ac power_Suc nonzero_inverse_mult_distrib) - -text{*Derivative of quotient*} -lemma DERIV_quotient: - fixes x :: "'a::{real_normed_field,recpower}" - shows "[| DERIV f x :> d; DERIV g x :> e; g(x) \ 0 |] - ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ Suc (Suc 0))" -by (drule (2) DERIV_divide) (simp add: mult_commute power_Suc) - - -subsection {* Differentiability predicate *} - -lemma differentiableD: "f differentiable x ==> \D. DERIV f x :> D" -by (simp add: differentiable_def) - -lemma differentiableI: "DERIV f x :> D ==> f differentiable x" -by (force simp add: differentiable_def) - -lemma differentiable_const: "(\z. a) differentiable x" - apply (unfold differentiable_def) - apply (rule_tac x=0 in exI) - apply simp - done - -lemma differentiable_sum: - assumes "f differentiable x" - and "g differentiable x" - shows "(\x. f x + g x) differentiable x" -proof - - from prems have "\D. DERIV f x :> D" by (unfold differentiable_def) - then obtain df where "DERIV f x :> df" .. - moreover from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) - then obtain dg where "DERIV g x :> dg" .. - ultimately have "DERIV (\x. f x + g x) x :> df + dg" by (rule DERIV_add) - hence "\D. DERIV (\x. f x + g x) x :> D" by auto - thus ?thesis by (fold differentiable_def) -qed - -lemma differentiable_diff: - assumes "f differentiable x" - and "g differentiable x" - shows "(\x. f x - g x) differentiable x" -proof - - from prems have "f differentiable x" by simp - moreover - from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) - then obtain dg where "DERIV g x :> dg" .. - then have "DERIV (\x. - g x) x :> -dg" by (rule DERIV_minus) - hence "\D. DERIV (\x. - g x) x :> D" by auto - hence "(\x. - g x) differentiable x" by (fold differentiable_def) - ultimately - show ?thesis - by (auto simp: diff_def dest: differentiable_sum) -qed - -lemma differentiable_mult: - assumes "f differentiable x" - and "g differentiable x" - shows "(\x. f x * g x) differentiable x" -proof - - from prems have "\D. DERIV f x :> D" by (unfold differentiable_def) - then obtain df where "DERIV f x :> df" .. - moreover from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) - then obtain dg where "DERIV g x :> dg" .. - ultimately have "DERIV (\x. f x * g x) x :> df * g x + dg * f x" by (simp add: DERIV_mult) - hence "\D. DERIV (\x. f x * g x) x :> D" by auto - thus ?thesis by (fold differentiable_def) -qed - - -subsection {* Nested Intervals and Bisection *} - -text{*Lemmas about nested intervals and proof by bisection (cf.Harrison). - All considerably tidied by lcp.*} - -lemma lemma_f_mono_add [rule_format (no_asm)]: "(\n. (f::nat=>real) n \ f (Suc n)) --> f m \ f(m + no)" -apply (induct "no") -apply (auto intro: order_trans) -done - -lemma f_inc_g_dec_Beq_f: "[| \n. f(n) \ f(Suc n); - \n. g(Suc n) \ g(n); - \n. f(n) \ g(n) |] - ==> Bseq (f :: nat \ real)" -apply (rule_tac k = "f 0" and K = "g 0" in BseqI2, rule allI) -apply (induct_tac "n") -apply (auto intro: order_trans) -apply (rule_tac y = "g (Suc na)" in order_trans) -apply (induct_tac [2] "na") -apply (auto intro: order_trans) -done - -lemma f_inc_g_dec_Beq_g: "[| \n. f(n) \ f(Suc n); - \n. g(Suc n) \ g(n); - \n. f(n) \ g(n) |] - ==> Bseq (g :: nat \ real)" -apply (subst Bseq_minus_iff [symmetric]) -apply (rule_tac g = "%x. - (f x)" in f_inc_g_dec_Beq_f) -apply auto -done - -lemma f_inc_imp_le_lim: - fixes f :: "nat \ real" - shows "\\n. f n \ f (Suc n); convergent f\ \ f n \ lim f" -apply (rule linorder_not_less [THEN iffD1]) -apply (auto simp add: convergent_LIMSEQ_iff LIMSEQ_iff monoseq_Suc) -apply (drule real_less_sum_gt_zero) -apply (drule_tac x = "f n + - lim f" in spec, safe) -apply (drule_tac P = "%na. no\na --> ?Q na" and x = "no + n" in spec, auto) -apply (subgoal_tac "lim f \ f (no + n) ") -apply (drule_tac no=no and m=n in lemma_f_mono_add) -apply (auto simp add: add_commute) -apply (induct_tac "no") -apply simp -apply (auto intro: order_trans simp add: diff_minus abs_if) -done - -lemma lim_uminus: "convergent g ==> lim (%x. - g x) = - (lim g)" -apply (rule LIMSEQ_minus [THEN limI]) -apply (simp add: convergent_LIMSEQ_iff) -done - -lemma g_dec_imp_lim_le: - fixes g :: "nat \ real" - shows "\\n. g (Suc n) \ g(n); convergent g\ \ lim g \ g n" -apply (subgoal_tac "- (g n) \ - (lim g) ") -apply (cut_tac [2] f = "%x. - (g x)" in f_inc_imp_le_lim) -apply (auto simp add: lim_uminus convergent_minus_iff [symmetric]) -done - -lemma lemma_nest: "[| \n. f(n) \ f(Suc n); - \n. g(Suc n) \ g(n); - \n. f(n) \ g(n) |] - ==> \l m :: real. l \ m & ((\n. f(n) \ l) & f ----> l) & - ((\n. m \ g(n)) & g ----> m)" -apply (subgoal_tac "monoseq f & monoseq g") -prefer 2 apply (force simp add: LIMSEQ_iff monoseq_Suc) -apply (subgoal_tac "Bseq f & Bseq g") -prefer 2 apply (blast intro: f_inc_g_dec_Beq_f f_inc_g_dec_Beq_g) -apply (auto dest!: Bseq_monoseq_convergent simp add: convergent_LIMSEQ_iff) -apply (rule_tac x = "lim f" in exI) -apply (rule_tac x = "lim g" in exI) -apply (auto intro: LIMSEQ_le) -apply (auto simp add: f_inc_imp_le_lim g_dec_imp_lim_le convergent_LIMSEQ_iff) -done - -lemma lemma_nest_unique: "[| \n. f(n) \ f(Suc n); - \n. g(Suc n) \ g(n); - \n. f(n) \ g(n); - (%n. f(n) - g(n)) ----> 0 |] - ==> \l::real. ((\n. f(n) \ l) & f ----> l) & - ((\n. l \ g(n)) & g ----> l)" -apply (drule lemma_nest, auto) -apply (subgoal_tac "l = m") -apply (drule_tac [2] X = f in LIMSEQ_diff) -apply (auto intro: LIMSEQ_unique) -done - -text{*The universal quantifiers below are required for the declaration - of @{text Bolzano_nest_unique} below.*} - -lemma Bolzano_bisect_le: - "a \ b ==> \n. fst (Bolzano_bisect P a b n) \ snd (Bolzano_bisect P a b n)" -apply (rule allI) -apply (induct_tac "n") -apply (auto simp add: Let_def split_def) -done - -lemma Bolzano_bisect_fst_le_Suc: "a \ b ==> - \n. fst(Bolzano_bisect P a b n) \ fst (Bolzano_bisect P a b (Suc n))" -apply (rule allI) -apply (induct_tac "n") -apply (auto simp add: Bolzano_bisect_le Let_def split_def) -done - -lemma Bolzano_bisect_Suc_le_snd: "a \ b ==> - \n. snd(Bolzano_bisect P a b (Suc n)) \ snd (Bolzano_bisect P a b n)" -apply (rule allI) -apply (induct_tac "n") -apply (auto simp add: Bolzano_bisect_le Let_def split_def) -done - -lemma eq_divide_2_times_iff: "((x::real) = y / (2 * z)) = (2 * x = y/z)" -apply (auto) -apply (drule_tac f = "%u. (1/2) *u" in arg_cong) -apply (simp) -done - -lemma Bolzano_bisect_diff: - "a \ b ==> - snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = - (b-a) / (2 ^ n)" -apply (induct "n") -apply (auto simp add: eq_divide_2_times_iff add_divide_distrib Let_def split_def) -done - -lemmas Bolzano_nest_unique = - lemma_nest_unique - [OF Bolzano_bisect_fst_le_Suc Bolzano_bisect_Suc_le_snd Bolzano_bisect_le] - - -lemma not_P_Bolzano_bisect: - assumes P: "!!a b c. [| P(a,b); P(b,c); a \ b; b \ c|] ==> P(a,c)" - and notP: "~ P(a,b)" - and le: "a \ b" - shows "~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" -proof (induct n) - case 0 show ?case using notP by simp - next - case (Suc n) - thus ?case - by (auto simp del: surjective_pairing [symmetric] - simp add: Let_def split_def Bolzano_bisect_le [OF le] - P [of "fst (Bolzano_bisect P a b n)" _ "snd (Bolzano_bisect P a b n)"]) -qed - -(*Now we re-package P_prem as a formula*) -lemma not_P_Bolzano_bisect': - "[| \a b c. P(a,b) & P(b,c) & a \ b & b \ c --> P(a,c); - ~ P(a,b); a \ b |] ==> - \n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" -by (blast elim!: not_P_Bolzano_bisect [THEN [2] rev_notE]) - - - -lemma lemma_BOLZANO: - "[| \a b c. P(a,b) & P(b,c) & a \ b & b \ c --> P(a,c); - \x. \d::real. 0 < d & - (\a b. a \ x & x \ b & (b-a) < d --> P(a,b)); - a \ b |] - ==> P(a,b)" -apply (rule Bolzano_nest_unique [where P1=P, THEN exE], assumption+) -apply (rule LIMSEQ_minus_cancel) -apply (simp (no_asm_simp) add: Bolzano_bisect_diff LIMSEQ_divide_realpow_zero) -apply (rule ccontr) -apply (drule not_P_Bolzano_bisect', assumption+) -apply (rename_tac "l") -apply (drule_tac x = l in spec, clarify) -apply (simp add: LIMSEQ_def) -apply (drule_tac P = "%r. 0 ?Q r" and x = "d/2" in spec) -apply (drule_tac P = "%r. 0 ?Q r" and x = "d/2" in spec) -apply (drule real_less_half_sum, auto) -apply (drule_tac x = "fst (Bolzano_bisect P a b (no + noa))" in spec) -apply (drule_tac x = "snd (Bolzano_bisect P a b (no + noa))" in spec) -apply safe -apply (simp_all (no_asm_simp)) -apply (rule_tac y = "abs (fst (Bolzano_bisect P a b (no + noa)) - l) + abs (snd (Bolzano_bisect P a b (no + noa)) - l)" in order_le_less_trans) -apply (simp (no_asm_simp) add: abs_if) -apply (rule real_sum_of_halves [THEN subst]) -apply (rule add_strict_mono) -apply (simp_all add: diff_minus [symmetric]) -done - - -lemma lemma_BOLZANO2: "((\a b c. (a \ b & b \ c & P(a,b) & P(b,c)) --> P(a,c)) & - (\x. \d::real. 0 < d & - (\a b. a \ x & x \ b & (b-a) < d --> P(a,b)))) - --> (\a b. a \ b --> P(a,b))" -apply clarify -apply (blast intro: lemma_BOLZANO) -done - - -subsection {* Intermediate Value Theorem *} - -text {*Prove Contrapositive by Bisection*} - -lemma IVT: "[| f(a::real) \ (y::real); y \ f(b); - a \ b; - (\x. a \ x & x \ b --> isCont f x) |] - ==> \x. a \ x & x \ b & f(x) = y" -apply (rule contrapos_pp, assumption) -apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> ~ (f (u) \ y & y \ f (v))" in lemma_BOLZANO2) -apply safe -apply simp_all -apply (simp add: isCont_iff LIM_def) -apply (rule ccontr) -apply (subgoal_tac "a \ x & x \ b") - prefer 2 - apply simp - apply (drule_tac P = "%d. 0 ?P d" and x = 1 in spec, arith) -apply (drule_tac x = x in spec)+ -apply simp -apply (drule_tac P = "%r. ?P r --> (\s>0. ?Q r s) " and x = "\y - f x\" in spec) -apply safe -apply simp -apply (drule_tac x = s in spec, clarify) -apply (cut_tac x = "f x" and y = y in linorder_less_linear, safe) -apply (drule_tac x = "ba-x" in spec) -apply (simp_all add: abs_if) -apply (drule_tac x = "aa-x" in spec) -apply (case_tac "x \ aa", simp_all) -done - -lemma IVT2: "[| f(b::real) \ (y::real); y \ f(a); - a \ b; - (\x. a \ x & x \ b --> isCont f x) - |] ==> \x. a \ x & x \ b & f(x) = y" -apply (subgoal_tac "- f a \ -y & -y \ - f b", clarify) -apply (drule IVT [where f = "%x. - f x"], assumption) -apply (auto intro: isCont_minus) -done - -(*HOL style here: object-level formulations*) -lemma IVT_objl: "(f(a::real) \ (y::real) & y \ f(b) & a \ b & - (\x. a \ x & x \ b --> isCont f x)) - --> (\x. a \ x & x \ b & f(x) = y)" -apply (blast intro: IVT) -done - -lemma IVT2_objl: "(f(b::real) \ (y::real) & y \ f(a) & a \ b & - (\x. a \ x & x \ b --> isCont f x)) - --> (\x. a \ x & x \ b & f(x) = y)" -apply (blast intro: IVT2) -done - -text{*By bisection, function continuous on closed interval is bounded above*} - -lemma isCont_bounded: - "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. \x::real. a \ x & x \ b --> f(x) \ M" -apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> (\M. \x. u \ x & x \ v --> f x \ M)" in lemma_BOLZANO2) -apply safe -apply simp_all -apply (rename_tac x xa ya M Ma) -apply (cut_tac x = M and y = Ma in linorder_linear, safe) -apply (rule_tac x = Ma in exI, clarify) -apply (cut_tac x = xb and y = xa in linorder_linear, force) -apply (rule_tac x = M in exI, clarify) -apply (cut_tac x = xb and y = xa in linorder_linear, force) -apply (case_tac "a \ x & x \ b") -apply (rule_tac [2] x = 1 in exI) -prefer 2 apply force -apply (simp add: LIM_def isCont_iff) -apply (drule_tac x = x in spec, auto) -apply (erule_tac V = "\M. \x. a \ x & x \ b & ~ f x \ M" in thin_rl) -apply (drule_tac x = 1 in spec, auto) -apply (rule_tac x = s in exI, clarify) -apply (rule_tac x = "\f x\ + 1" in exI, clarify) -apply (drule_tac x = "xa-x" in spec) -apply (auto simp add: abs_ge_self) -done - -text{*Refine the above to existence of least upper bound*} - -lemma lemma_reals_complete: "((\x. x \ S) & (\y. isUb UNIV S (y::real))) --> - (\t. isLub UNIV S t)" -by (blast intro: reals_complete) - -lemma isCont_has_Ub: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. (\x::real. a \ x & x \ b --> f(x) \ M) & - (\N. N < M --> (\x. a \ x & x \ b & N < f(x)))" -apply (cut_tac S = "Collect (%y. \x. a \ x & x \ b & y = f x)" - in lemma_reals_complete) -apply auto -apply (drule isCont_bounded, assumption) -apply (auto simp add: isUb_def leastP_def isLub_def setge_def setle_def) -apply (rule exI, auto) -apply (auto dest!: spec simp add: linorder_not_less) -done - -text{*Now show that it attains its upper bound*} - -lemma isCont_eq_Ub: - assumes le: "a \ b" - and con: "\x::real. a \ x & x \ b --> isCont f x" - shows "\M::real. (\x. a \ x & x \ b --> f(x) \ M) & - (\x. a \ x & x \ b & f(x) = M)" -proof - - from isCont_has_Ub [OF le con] - obtain M where M1: "\x. a \ x \ x \ b \ f x \ M" - and M2: "!!N. N \x. a \ x \ x \ b \ N < f x" by blast - show ?thesis - proof (intro exI, intro conjI) - show " \x. a \ x \ x \ b \ f x \ M" by (rule M1) - show "\x. a \ x \ x \ b \ f x = M" - proof (rule ccontr) - assume "\ (\x. a \ x \ x \ b \ f x = M)" - with M1 have M3: "\x. a \ x & x \ b --> f x < M" - by (fastsimp simp add: linorder_not_le [symmetric]) - hence "\x. a \ x & x \ b --> isCont (%x. inverse (M - f x)) x" - by (auto simp add: isCont_inverse isCont_diff con) - from isCont_bounded [OF le this] - obtain k where k: "!!x. a \ x & x \ b --> inverse (M - f x) \ k" by auto - have Minv: "!!x. a \ x & x \ b --> 0 < inverse (M - f (x))" - by (simp add: M3 compare_rls) - have "!!x. a \ x & x \ b --> inverse (M - f x) < k+1" using k - by (auto intro: order_le_less_trans [of _ k]) - with Minv - have "!!x. a \ x & x \ b --> inverse(k+1) < inverse(inverse(M - f x))" - by (intro strip less_imp_inverse_less, simp_all) - hence invlt: "!!x. a \ x & x \ b --> inverse(k+1) < M - f x" - by simp - have "M - inverse (k+1) < M" using k [of a] Minv [of a] le - by (simp, arith) - from M2 [OF this] - obtain x where ax: "a \ x & x \ b & M - inverse(k+1) < f x" .. - thus False using invlt [of x] by force - qed - qed -qed - - -text{*Same theorem for lower bound*} - -lemma isCont_eq_Lb: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \M::real. (\x::real. a \ x & x \ b --> M \ f(x)) & - (\x. a \ x & x \ b & f(x) = M)" -apply (subgoal_tac "\x. a \ x & x \ b --> isCont (%x. - (f x)) x") -prefer 2 apply (blast intro: isCont_minus) -apply (drule_tac f = "(%x. - (f x))" in isCont_eq_Ub) -apply safe -apply auto -done - - -text{*Another version.*} - -lemma isCont_Lb_Ub: "[|a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \L M::real. (\x::real. a \ x & x \ b --> L \ f(x) & f(x) \ M) & - (\y. L \ y & y \ M --> (\x. a \ x & x \ b & (f(x) = y)))" -apply (frule isCont_eq_Lb) -apply (frule_tac [2] isCont_eq_Ub) -apply (assumption+, safe) -apply (rule_tac x = "f x" in exI) -apply (rule_tac x = "f xa" in exI, simp, safe) -apply (cut_tac x = x and y = xa in linorder_linear, safe) -apply (cut_tac f = f and a = x and b = xa and y = y in IVT_objl) -apply (cut_tac [2] f = f and a = xa and b = x and y = y in IVT2_objl, safe) -apply (rule_tac [2] x = xb in exI) -apply (rule_tac [4] x = xb in exI, simp_all) -done - - -text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*} - -lemma DERIV_left_inc: - fixes f :: "real => real" - assumes der: "DERIV f x :> l" - and l: "0 < l" - shows "\d > 0. \h > 0. h < d --> f(x) < f(x + h)" -proof - - from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]] - have "\s > 0. (\z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < l)" - by (simp add: diff_minus) - then obtain s - where s: "0 < s" - and all: "!!z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < l" - by auto - thus ?thesis - proof (intro exI conjI strip) - show "0 real" - assumes der: "DERIV f x :> l" - and l: "l < 0" - shows "\d > 0. \h > 0. h < d --> f(x) < f(x-h)" -proof - - from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]] - have "\s > 0. (\z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < -l)" - by (simp add: diff_minus) - then obtain s - where s: "0 < s" - and all: "!!z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < -l" - by auto - thus ?thesis - proof (intro exI conjI strip) - show "0 real" - assumes der: "DERIV f x :> l" - and d: "0 < d" - and le: "\y. \x-y\ < d --> f(y) \ f(x)" - shows "l = 0" -proof (cases rule: linorder_cases [of l 0]) - case equal thus ?thesis . -next - case less - from DERIV_left_dec [OF der less] - obtain d' where d': "0 < d'" - and lt: "\h > 0. h < d' \ f x < f (x-h)" by blast - from real_lbound_gt_zero [OF d d'] - obtain e where "0 < e \ e < d \ e < d'" .. - with lt le [THEN spec [where x="x-e"]] - show ?thesis by (auto simp add: abs_if) -next - case greater - from DERIV_left_inc [OF der greater] - obtain d' where d': "0 < d'" - and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast - from real_lbound_gt_zero [OF d d'] - obtain e where "0 < e \ e < d \ e < d'" .. - with lt le [THEN spec [where x="x+e"]] - show ?thesis by (auto simp add: abs_if) -qed - - -text{*Similar theorem for a local minimum*} -lemma DERIV_local_min: - fixes f :: "real => real" - shows "[| DERIV f x :> l; 0 < d; \y. \x-y\ < d --> f(x) \ f(y) |] ==> l = 0" -by (drule DERIV_minus [THEN DERIV_local_max], auto) - - -text{*In particular, if a function is locally flat*} -lemma DERIV_local_const: - fixes f :: "real => real" - shows "[| DERIV f x :> l; 0 < d; \y. \x-y\ < d --> f(x) = f(y) |] ==> l = 0" -by (auto dest!: DERIV_local_max) - -text{*Lemma about introducing open ball in open interval*} -lemma lemma_interval_lt: - "[| a < x; x < b |] - ==> \d::real. 0 < d & (\y. \x-y\ < d --> a < y & y < b)" - -apply (simp add: abs_less_iff) -apply (insert linorder_linear [of "x-a" "b-x"], safe) -apply (rule_tac x = "x-a" in exI) -apply (rule_tac [2] x = "b-x" in exI, auto) -done - -lemma lemma_interval: "[| a < x; x < b |] ==> - \d::real. 0 < d & (\y. \x-y\ < d --> a \ y & y \ b)" -apply (drule lemma_interval_lt, auto) -apply (auto intro!: exI) -done - -text{*Rolle's Theorem. - If @{term f} is defined and continuous on the closed interval - @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"}, - and @{term "f(a) = f(b)"}, - then there exists @{text "x0 \ (a,b)"} such that @{term "f'(x0) = 0"}*} -theorem Rolle: - assumes lt: "a < b" - and eq: "f(a) = f(b)" - and con: "\x. a \ x & x \ b --> isCont f x" - and dif [rule_format]: "\x. a < x & x < b --> f differentiable x" - shows "\z::real. a < z & z < b & DERIV f z :> 0" -proof - - have le: "a \ b" using lt by simp - from isCont_eq_Ub [OF le con] - obtain x where x_max: "\z. a \ z \ z \ b \ f z \ f x" - and alex: "a \ x" and xleb: "x \ b" - by blast - from isCont_eq_Lb [OF le con] - obtain x' where x'_min: "\z. a \ z \ z \ b \ f x' \ f z" - and alex': "a \ x'" and x'leb: "x' \ b" - by blast - show ?thesis - proof cases - assume axb: "a < x & x < b" - --{*@{term f} attains its maximum within the interval*} - hence ax: "ay. \x-y\ < d \ a \ y \ y \ b" - by blast - hence bound': "\y. \x-y\ < d \ f y \ f x" using x_max - by blast - from differentiableD [OF dif [OF axb]] - obtain l where der: "DERIV f x :> l" .. - have "l=0" by (rule DERIV_local_max [OF der d bound']) - --{*the derivative at a local maximum is zero*} - thus ?thesis using ax xb der by auto - next - assume notaxb: "~ (a < x & x < b)" - hence xeqab: "x=a | x=b" using alex xleb by arith - hence fb_eq_fx: "f b = f x" by (auto simp add: eq) - show ?thesis - proof cases - assume ax'b: "a < x' & x' < b" - --{*@{term f} attains its minimum within the interval*} - hence ax': "ay. \x'-y\ < d \ a \ y \ y \ b" - by blast - hence bound': "\y. \x'-y\ < d \ f x' \ f y" using x'_min - by blast - from differentiableD [OF dif [OF ax'b]] - obtain l where der: "DERIV f x' :> l" .. - have "l=0" by (rule DERIV_local_min [OF der d bound']) - --{*the derivative at a local minimum is zero*} - thus ?thesis using ax' x'b der by auto - next - assume notax'b: "~ (a < x' & x' < b)" - --{*@{term f} is constant througout the interval*} - hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith - hence fb_eq_fx': "f b = f x'" by (auto simp add: eq) - from dense [OF lt] - obtain r where ar: "a < r" and rb: "r < b" by blast - from lemma_interval [OF ar rb] - obtain d where d: "0y. \r-y\ < d \ a \ y \ y \ b" - by blast - have eq_fb: "\z. a \ z --> z \ b --> f z = f b" - proof (clarify) - fix z::real - assume az: "a \ z" and zb: "z \ b" - show "f z = f b" - proof (rule order_antisym) - show "f z \ f b" by (simp add: fb_eq_fx x_max az zb) - show "f b \ f z" by (simp add: fb_eq_fx' x'_min az zb) - qed - qed - have bound': "\y. \r-y\ < d \ f r = f y" - proof (intro strip) - fix y::real - assume lt: "\r-y\ < d" - hence "f y = f b" by (simp add: eq_fb bound) - thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le) - qed - from differentiableD [OF dif [OF conjI [OF ar rb]]] - obtain l where der: "DERIV f r :> l" .. - have "l=0" by (rule DERIV_local_const [OF der d bound']) - --{*the derivative of a constant function is zero*} - thus ?thesis using ar rb der by auto - qed - qed -qed - - -subsection{*Mean Value Theorem*} - -lemma lemma_MVT: - "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)" -proof cases - assume "a=b" thus ?thesis by simp -next - assume "a\b" - hence ba: "b-a \ 0" by arith - show ?thesis - by (rule real_mult_left_cancel [OF ba, THEN iffD1], - simp add: right_diff_distrib, - simp add: left_diff_distrib) -qed - -theorem MVT: - assumes lt: "a < b" - and con: "\x. a \ x & x \ b --> isCont f x" - and dif [rule_format]: "\x. a < x & x < b --> f differentiable x" - shows "\l z::real. a < z & z < b & DERIV f z :> l & - (f(b) - f(a) = (b-a) * l)" -proof - - let ?F = "%x. f x - ((f b - f a) / (b-a)) * x" - have contF: "\x. a \ x \ x \ b \ isCont ?F x" using con - by (fast intro: isCont_diff isCont_const isCont_mult isCont_ident) - have difF: "\x. a < x \ x < b \ ?F differentiable x" - proof (clarify) - fix x::real - assume ax: "a < x" and xb: "x < b" - from differentiableD [OF dif [OF conjI [OF ax xb]]] - obtain l where der: "DERIV f x :> l" .. - show "?F differentiable x" - by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"], - blast intro: DERIV_diff DERIV_cmult_Id der) - qed - from Rolle [where f = ?F, OF lt lemma_MVT contF difF] - obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0" - by blast - have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)" - by (rule DERIV_cmult_Id) - hence derF: "DERIV (\x. ?F x + (f b - f a) / (b - a) * x) z - :> 0 + (f b - f a) / (b - a)" - by (rule DERIV_add [OF der]) - show ?thesis - proof (intro exI conjI) - show "a < z" using az . - show "z < b" using zb . - show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by (simp) - show "DERIV f z :> ((f b - f a)/(b-a))" using derF by simp - qed -qed - - -text{*A function is constant if its derivative is 0 over an interval.*} - -lemma DERIV_isconst_end: - fixes f :: "real => real" - shows "[| a < b; - \x. a \ x & x \ b --> isCont f x; - \x. a < x & x < b --> DERIV f x :> 0 |] - ==> f b = f a" -apply (drule MVT, assumption) -apply (blast intro: differentiableI) -apply (auto dest!: DERIV_unique simp add: diff_eq_eq) -done - -lemma DERIV_isconst1: - fixes f :: "real => real" - shows "[| a < b; - \x. a \ x & x \ b --> isCont f x; - \x. a < x & x < b --> DERIV f x :> 0 |] - ==> \x. a \ x & x \ b --> f x = f a" -apply safe -apply (drule_tac x = a in order_le_imp_less_or_eq, safe) -apply (drule_tac b = x in DERIV_isconst_end, auto) -done - -lemma DERIV_isconst2: - fixes f :: "real => real" - shows "[| a < b; - \x. a \ x & x \ b --> isCont f x; - \x. a < x & x < b --> DERIV f x :> 0; - a \ x; x \ b |] - ==> f x = f a" -apply (blast dest: DERIV_isconst1) -done - -lemma DERIV_isconst_all: - fixes f :: "real => real" - shows "\x. DERIV f x :> 0 ==> f(x) = f(y)" -apply (rule linorder_cases [of x y]) -apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+ -done - -lemma DERIV_const_ratio_const: - fixes f :: "real => real" - shows "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k" -apply (rule linorder_cases [of a b], auto) -apply (drule_tac [!] f = f in MVT) -apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def) -apply (auto dest: DERIV_unique simp add: ring_distribs diff_minus) -done - -lemma DERIV_const_ratio_const2: - fixes f :: "real => real" - shows "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k" -apply (rule_tac c1 = "b-a" in real_mult_right_cancel [THEN iffD1]) -apply (auto dest!: DERIV_const_ratio_const simp add: mult_assoc) -done - -lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)" -by (simp) - -lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)" -by (simp) - -text{*Gallileo's "trick": average velocity = av. of end velocities*} - -lemma DERIV_const_average: - fixes v :: "real => real" - assumes neq: "a \ (b::real)" - and der: "\x. DERIV v x :> k" - shows "v ((a + b)/2) = (v a + v b)/2" -proof (cases rule: linorder_cases [of a b]) - case equal with neq show ?thesis by simp -next - case less - have "(v b - v a) / (b - a) = k" - by (rule DERIV_const_ratio_const2 [OF neq der]) - hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp - moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" - by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) - ultimately show ?thesis using neq by force -next - case greater - have "(v b - v a) / (b - a) = k" - by (rule DERIV_const_ratio_const2 [OF neq der]) - hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp - moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" - by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) - ultimately show ?thesis using neq by (force simp add: add_commute) -qed - - -text{*Dull lemma: an continuous injection on an interval must have a -strict maximum at an end point, not in the middle.*} - -lemma lemma_isCont_inj: - fixes f :: "real \ real" - assumes d: "0 < d" - and inj [rule_format]: "\z. \z-x\ \ d --> g(f z) = z" - and cont: "\z. \z-x\ \ d --> isCont f z" - shows "\z. \z-x\ \ d & f x < f z" -proof (rule ccontr) - assume "~ (\z. \z-x\ \ d & f x < f z)" - hence all [rule_format]: "\z. \z - x\ \ d --> f z \ f x" by auto - show False - proof (cases rule: linorder_le_cases [of "f(x-d)" "f(x+d)"]) - case le - from d cont all [of "x+d"] - have flef: "f(x+d) \ f x" - and xlex: "x - d \ x" - and cont': "\z. x - d \ z \ z \ x \ isCont f z" - by (auto simp add: abs_if) - from IVT [OF le flef xlex cont'] - obtain x' where "x-d \ x'" "x' \ x" "f x' = f(x+d)" by blast - moreover - hence "g(f x') = g (f(x+d))" by simp - ultimately show False using d inj [of x'] inj [of "x+d"] - by (simp add: abs_le_iff) - next - case ge - from d cont all [of "x-d"] - have flef: "f(x-d) \ f x" - and xlex: "x \ x+d" - and cont': "\z. x \ z \ z \ x+d \ isCont f z" - by (auto simp add: abs_if) - from IVT2 [OF ge flef xlex cont'] - obtain x' where "x \ x'" "x' \ x+d" "f x' = f(x-d)" by blast - moreover - hence "g(f x') = g (f(x-d))" by simp - ultimately show False using d inj [of x'] inj [of "x-d"] - by (simp add: abs_le_iff) - qed -qed - - -text{*Similar version for lower bound.*} - -lemma lemma_isCont_inj2: - fixes f g :: "real \ real" - shows "[|0 < d; \z. \z-x\ \ d --> g(f z) = z; - \z. \z-x\ \ d --> isCont f z |] - ==> \z. \z-x\ \ d & f z < f x" -apply (insert lemma_isCont_inj - [where f = "%x. - f x" and g = "%y. g(-y)" and x = x and d = d]) -apply (simp add: isCont_minus linorder_not_le) -done - -text{*Show there's an interval surrounding @{term "f(x)"} in -@{text "f[[x - d, x + d]]"} .*} - -lemma isCont_inj_range: - fixes f :: "real \ real" - assumes d: "0 < d" - and inj: "\z. \z-x\ \ d --> g(f z) = z" - and cont: "\z. \z-x\ \ d --> isCont f z" - shows "\e>0. \y. \y - f x\ \ e --> (\z. \z-x\ \ d & f z = y)" -proof - - have "x-d \ x+d" "\z. x-d \ z \ z \ x+d \ isCont f z" using cont d - by (auto simp add: abs_le_iff) - from isCont_Lb_Ub [OF this] - obtain L M - where all1 [rule_format]: "\z. x-d \ z \ z \ x+d \ L \ f z \ f z \ M" - and all2 [rule_format]: - "\y. L \ y \ y \ M \ (\z. x-d \ z \ z \ x+d \ f z = y)" - by auto - with d have "L \ f x & f x \ M" by simp - moreover have "L \ f x" - proof - - from lemma_isCont_inj2 [OF d inj cont] - obtain u where "\u - x\ \ d" "f u < f x" by auto - thus ?thesis using all1 [of u] by arith - qed - moreover have "f x \ M" - proof - - from lemma_isCont_inj [OF d inj cont] - obtain u where "\u - x\ \ d" "f x < f u" by auto - thus ?thesis using all1 [of u] by arith - qed - ultimately have "L < f x & f x < M" by arith - hence "0 < f x - L" "0 < M - f x" by arith+ - from real_lbound_gt_zero [OF this] - obtain e where e: "0 < e" "e < f x - L" "e < M - f x" by auto - thus ?thesis - proof (intro exI conjI) - show "0y. \y - f x\ \ e \ (\z. \z - x\ \ d \ f z = y)" - proof (intro strip) - fix y::real - assume "\y - f x\ \ e" - with e have "L \ y \ y \ M" by arith - from all2 [OF this] - obtain z where "x - d \ z" "z \ x + d" "f z = y" by blast - thus "\z. \z - x\ \ d \ f z = y" - by (force simp add: abs_le_iff) - qed - qed -qed - - -text{*Continuity of inverse function*} - -lemma isCont_inverse_function: - fixes f g :: "real \ real" - assumes d: "0 < d" - and inj: "\z. \z-x\ \ d --> g(f z) = z" - and cont: "\z. \z-x\ \ d --> isCont f z" - shows "isCont g (f x)" -proof (simp add: isCont_iff LIM_eq) - show "\r. 0 < r \ - (\s>0. \z. z\0 \ \z\ < s \ \g(f x + z) - g(f x)\ < r)" - proof (intro strip) - fix r::real - assume r: "0 e < d" by blast - with inj cont - have e_simps: "\z. \z-x\ \ e --> g (f z) = z" - "\z. \z-x\ \ e --> isCont f z" by auto - from isCont_inj_range [OF e this] - obtain e' where e': "0 < e'" - and all: "\y. \y - f x\ \ e' \ (\z. \z - x\ \ e \ f z = y)" - by blast - show "\s>0. \z. z\0 \ \z\ < s \ \g(f x + z) - g(f x)\ < r" - proof (intro exI conjI) - show "0z. z \ 0 \ \z\ < e' \ \g (f x + z) - g (f x)\ < r" - proof (intro strip) - fix z::real - assume z: "z \ 0 \ \z\ < e'" - with e e_lt e_simps all [rule_format, of "f x + z"] - show "\g (f x + z) - g (f x)\ < r" by force - qed - qed - qed -qed - -text {* Derivative of inverse function *} - -lemma DERIV_inverse_function: - fixes f g :: "real \ real" - assumes der: "DERIV f (g x) :> D" - assumes neq: "D \ 0" - assumes a: "a < x" and b: "x < b" - assumes inj: "\y. a < y \ y < b \ f (g y) = y" - assumes cont: "isCont g x" - shows "DERIV g x :> inverse D" -unfolding DERIV_iff2 -proof (rule LIM_equal2) - show "0 < min (x - a) (b - x)" - using a b by arith -next - fix y - assume "norm (y - x) < min (x - a) (b - x)" - hence "a < y" and "y < b" - by (simp_all add: abs_less_iff) - thus "(g y - g x) / (y - x) = - inverse ((f (g y) - x) / (g y - g x))" - by (simp add: inj) -next - have "(\z. (f z - f (g x)) / (z - g x)) -- g x --> D" - by (rule der [unfolded DERIV_iff2]) - hence 1: "(\z. (f z - x) / (z - g x)) -- g x --> D" - using inj a b by simp - have 2: "\d>0. \y. y \ x \ norm (y - x) < d \ g y \ g x" - proof (safe intro!: exI) - show "0 < min (x - a) (b - x)" - using a b by simp - next - fix y - assume "norm (y - x) < min (x - a) (b - x)" - hence y: "a < y" "y < b" - by (simp_all add: abs_less_iff) - assume "g y = g x" - hence "f (g y) = f (g x)" by simp - hence "y = x" using inj y a b by simp - also assume "y \ x" - finally show False by simp - qed - have "(\y. (f (g y) - x) / (g y - g x)) -- x --> D" - using cont 1 2 by (rule isCont_LIM_compose2) - thus "(\y. inverse ((f (g y) - x) / (g y - g x))) - -- x --> inverse D" - using neq by (rule LIM_inverse) -qed - -theorem GMVT: - fixes a b :: real - assumes alb: "a < b" - and fc: "\x. a \ x \ x \ b \ isCont f x" - and fd: "\x. a < x \ x < b \ f differentiable x" - and gc: "\x. a \ x \ x \ b \ isCont g x" - and gd: "\x. a < x \ x < b \ g differentiable x" - shows "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ ((f b - f a) * g'c) = ((g b - g a) * f'c)" -proof - - let ?h = "\x. (f b - f a)*(g x) - (g b - g a)*(f x)" - from prems have "a < b" by simp - moreover have "\x. a \ x \ x \ b \ isCont ?h x" - proof - - have "\x. a <= x \ x <= b \ isCont (\x. f b - f a) x" by simp - with gc have "\x. a <= x \ x <= b \ isCont (\x. (f b - f a) * g x) x" - by (auto intro: isCont_mult) - moreover - have "\x. a <= x \ x <= b \ isCont (\x. g b - g a) x" by simp - with fc have "\x. a <= x \ x <= b \ isCont (\x. (g b - g a) * f x) x" - by (auto intro: isCont_mult) - ultimately show ?thesis - by (fastsimp intro: isCont_diff) - qed - moreover - have "\x. a < x \ x < b \ ?h differentiable x" - proof - - have "\x. a < x \ x < b \ (\x. f b - f a) differentiable x" by (simp add: differentiable_const) - with gd have "\x. a < x \ x < b \ (\x. (f b - f a) * g x) differentiable x" by (simp add: differentiable_mult) - moreover - have "\x. a < x \ x < b \ (\x. g b - g a) differentiable x" by (simp add: differentiable_const) - with fd have "\x. a < x \ x < b \ (\x. (g b - g a) * f x) differentiable x" by (simp add: differentiable_mult) - ultimately show ?thesis by (simp add: differentiable_diff) - qed - ultimately have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" by (rule MVT) - then obtain l where ldef: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. - then obtain c where cdef: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" .. - - from cdef have cint: "a < c \ c < b" by auto - with gd have "g differentiable c" by simp - hence "\D. DERIV g c :> D" by (rule differentiableD) - then obtain g'c where g'cdef: "DERIV g c :> g'c" .. - - from cdef have "a < c \ c < b" by auto - with fd have "f differentiable c" by simp - hence "\D. DERIV f c :> D" by (rule differentiableD) - then obtain f'c where f'cdef: "DERIV f c :> f'c" .. - - from cdef have "DERIV ?h c :> l" by auto - moreover - { - have "DERIV (\x. (f b - f a) * g x) c :> g'c * (f b - f a)" - apply (insert DERIV_const [where k="f b - f a"]) - apply (drule meta_spec [of _ c]) - apply (drule DERIV_mult [OF _ g'cdef]) - by simp - moreover have "DERIV (\x. (g b - g a) * f x) c :> f'c * (g b - g a)" - apply (insert DERIV_const [where k="g b - g a"]) - apply (drule meta_spec [of _ c]) - apply (drule DERIV_mult [OF _ f'cdef]) - by simp - ultimately have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" - by (simp add: DERIV_diff) - } - ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) - - { - from cdef have "?h b - ?h a = (b - a) * l" by auto - also with leq have "\ = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp - finally have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp - } - moreover - { - have "?h b - ?h a = - ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - - ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" - by (simp add: mult_ac add_ac right_diff_distrib) - hence "?h b - ?h a = 0" by auto - } - ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto - with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp - hence "g'c * (f b - f a) = f'c * (g b - g a)" by simp - hence "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: mult_ac) - - with g'cdef f'cdef cint show ?thesis by auto -qed - -lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" -by auto - -subsection {* Derivatives of univariate polynomials *} - - - -primrec pderiv_aux :: "nat => real list => real list" where - pderiv_aux_Nil: "pderiv_aux n [] = []" -| pderiv_aux_Cons: "pderiv_aux n (h#t) = - (real n * h)#(pderiv_aux (Suc n) t)" - -definition - pderiv :: "real list => real list" where - "pderiv p = (if p = [] then [] else pderiv_aux 1 (tl p))" - - -text{*The derivative*} - -lemma pderiv_Nil: "pderiv [] = []" - -apply (simp add: pderiv_def) -done -declare pderiv_Nil [simp] - -lemma pderiv_singleton: "pderiv [c] = []" -by (simp add: pderiv_def) -declare pderiv_singleton [simp] - -lemma pderiv_Cons: "pderiv (h#t) = pderiv_aux 1 t" -by (simp add: pderiv_def) - -lemma DERIV_cmult2: "DERIV f x :> D ==> DERIV (%x. (f x) * c :: real) x :> D * c" -by (simp add: DERIV_cmult mult_commute [of _ c]) - -lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" -by (rule lemma_DERIV_subst, rule DERIV_pow, simp) -declare DERIV_pow2 [simp] DERIV_pow [simp] - -lemma lemma_DERIV_poly1: "\n. DERIV (%x. (x ^ (Suc n) * poly p x)) x :> - x ^ n * poly (pderiv_aux (Suc n) p) x " -apply (induct "p") -apply (auto intro!: DERIV_add DERIV_cmult2 - simp add: pderiv_def right_distrib real_mult_assoc [symmetric] - simp del: realpow_Suc) -apply (subst mult_commute) -apply (simp del: realpow_Suc) -apply (simp add: mult_commute realpow_Suc [symmetric] del: realpow_Suc) -done - -lemma lemma_DERIV_poly: "DERIV (%x. (x ^ (Suc n) * poly p x)) x :> - x ^ n * poly (pderiv_aux (Suc n) p) x " -by (simp add: lemma_DERIV_poly1 del: realpow_Suc) - -lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: real) x :> D" -by (rule lemma_DERIV_subst, rule DERIV_add, auto) - -lemma poly_DERIV[simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x" -apply (induct "p") -apply (auto simp add: pderiv_Cons) -apply (rule DERIV_add_const) -apply (rule lemma_DERIV_subst) -apply (rule lemma_DERIV_poly [where n=0, simplified], simp) -done - - -text{* Consequences of the derivative theorem above*} - -lemma poly_differentiable[simp]: "(%x. poly p x) differentiable (x::real)" -apply (simp add: differentiable_def) -apply (blast intro: poly_DERIV) -done - -lemma poly_isCont[simp]: "isCont (%x. poly p x) (x::real)" -by (rule poly_DERIV [THEN DERIV_isCont]) - -lemma poly_IVT_pos: "[| a < b; poly p (a::real) < 0; 0 < poly p b |] - ==> \x. a < x & x < b & (poly p x = 0)" -apply (cut_tac f = "%x. poly p x" and a = a and b = b and y = 0 in IVT_objl) -apply (auto simp add: order_le_less) -done - -lemma poly_IVT_neg: "[| (a::real) < b; 0 < poly p a; poly p b < 0 |] - ==> \x. a < x & x < b & (poly p x = 0)" -apply (insert poly_IVT_pos [where p = "-- p" ]) -apply (simp add: poly_minus neg_less_0_iff_less) -done - -lemma poly_MVT: "a < b ==> - \x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)" -apply (drule_tac f = "poly p" in MVT, auto) -apply (rule_tac x = z in exI) -apply (auto simp add: real_mult_left_cancel poly_DERIV [THEN DERIV_unique]) -done - -text{*Lemmas for Derivatives*} - -lemma lemma_poly_pderiv_aux_add: "\p2 n. poly (pderiv_aux n (p1 +++ p2)) x = - poly (pderiv_aux n p1 +++ pderiv_aux n p2) x" -apply (induct "p1", simp, clarify) -apply (case_tac "p2") -apply (auto simp add: right_distrib) -done - -lemma poly_pderiv_aux_add: "poly (pderiv_aux n (p1 +++ p2)) x = - poly (pderiv_aux n p1 +++ pderiv_aux n p2) x" -apply (simp add: lemma_poly_pderiv_aux_add) -done - -lemma lemma_poly_pderiv_aux_cmult: "\n. poly (pderiv_aux n (c %* p)) x = poly (c %* pderiv_aux n p) x" -apply (induct "p") -apply (auto simp add: poly_cmult mult_ac) -done - -lemma poly_pderiv_aux_cmult: "poly (pderiv_aux n (c %* p)) x = poly (c %* pderiv_aux n p) x" -by (simp add: lemma_poly_pderiv_aux_cmult) - -lemma poly_pderiv_aux_minus: - "poly (pderiv_aux n (-- p)) x = poly (-- pderiv_aux n p) x" -apply (simp add: poly_minus_def poly_pderiv_aux_cmult) -done - -lemma lemma_poly_pderiv_aux_mult1: "\n. poly (pderiv_aux (Suc n) p) x = poly ((pderiv_aux n p) +++ p) x" -apply (induct "p") -apply (auto simp add: real_of_nat_Suc left_distrib) -done - -lemma lemma_poly_pderiv_aux_mult: "poly (pderiv_aux (Suc n) p) x = poly ((pderiv_aux n p) +++ p) x" -by (simp add: lemma_poly_pderiv_aux_mult1) - -lemma lemma_poly_pderiv_add: "\q. poly (pderiv (p +++ q)) x = poly (pderiv p +++ pderiv q) x" -apply (induct "p", simp, clarify) -apply (case_tac "q") -apply (auto simp add: poly_pderiv_aux_add poly_add pderiv_def) -done - -lemma poly_pderiv_add: "poly (pderiv (p +++ q)) x = poly (pderiv p +++ pderiv q) x" -by (simp add: lemma_poly_pderiv_add) - -lemma poly_pderiv_cmult: "poly (pderiv (c %* p)) x = poly (c %* (pderiv p)) x" -apply (induct "p") -apply (auto simp add: poly_pderiv_aux_cmult poly_cmult pderiv_def) -done - -lemma poly_pderiv_minus: "poly (pderiv (--p)) x = poly (--(pderiv p)) x" -by (simp add: poly_minus_def poly_pderiv_cmult) - -lemma lemma_poly_mult_pderiv: - "poly (pderiv (h#t)) x = poly ((0 # (pderiv t)) +++ t) x" -apply (simp add: pderiv_def) -apply (induct "t") -apply (auto simp add: poly_add lemma_poly_pderiv_aux_mult) -done - -lemma poly_pderiv_mult: "\q. poly (pderiv (p *** q)) x = - poly (p *** (pderiv q) +++ q *** (pderiv p)) x" -apply (induct "p") -apply (auto simp add: poly_add poly_cmult poly_pderiv_cmult poly_pderiv_add poly_mult) -apply (rule lemma_poly_mult_pderiv [THEN ssubst]) -apply (rule lemma_poly_mult_pderiv [THEN ssubst]) -apply (rule poly_add [THEN ssubst]) -apply (rule poly_add [THEN ssubst]) -apply (simp (no_asm_simp) add: poly_mult right_distrib add_ac mult_ac) -done - -lemma poly_pderiv_exp: "poly (pderiv (p %^ (Suc n))) x = - poly ((real (Suc n)) %* (p %^ n) *** pderiv p) x" -apply (induct "n") -apply (auto simp add: poly_add poly_pderiv_cmult poly_cmult poly_pderiv_mult - real_of_nat_zero poly_mult real_of_nat_Suc - right_distrib left_distrib mult_ac) -done - -lemma poly_pderiv_exp_prime: "poly (pderiv ([-a, 1] %^ (Suc n))) x = - poly (real (Suc n) %* ([-a, 1] %^ n)) x" -apply (simp add: poly_pderiv_exp poly_mult del: pexp_Suc) -apply (simp add: poly_cmult pderiv_def) -done - - -lemma real_mult_zero_disj_iff[simp]: "(x * y = 0) = (x = (0::real) | y = 0)" -by simp - -lemma pderiv_aux_iszero [rule_format, simp]: - "\n. list_all (%c. c = 0) (pderiv_aux (Suc n) p) = list_all (%c. c = 0) p" -by (induct "p", auto) - -lemma pderiv_aux_iszero_num: "(number_of n :: nat) \ 0 - ==> (list_all (%c. c = 0) (pderiv_aux (number_of n) p) = - list_all (%c. c = 0) p)" -unfolding neq0_conv -apply (rule_tac n1 = "number_of n" and m1 = 0 in less_imp_Suc_add [THEN exE], force) -apply (rule_tac n1 = "0 + x" in pderiv_aux_iszero [THEN subst]) -apply (simp (no_asm_simp) del: pderiv_aux_iszero) -done - -instance real:: idom_char_0 -apply (intro_classes) -done - -instance real:: recpower_idom_char_0 -apply (intro_classes) -done - -lemma pderiv_iszero [rule_format]: - "poly (pderiv p) = poly [] --> (\h. poly p = poly [h])" -apply (simp add: poly_zero) -apply (induct "p", force) -apply (simp add: pderiv_Cons pderiv_aux_iszero_num del: poly_Cons) -apply (auto simp add: poly_zero [symmetric]) -done - -lemma pderiv_zero_obj: "poly p = poly [] --> (poly (pderiv p) = poly [])" -apply (simp add: poly_zero) -apply (induct "p", force) -apply (simp add: pderiv_Cons pderiv_aux_iszero_num del: poly_Cons) -done - -lemma pderiv_zero: "poly p = poly [] ==> (poly (pderiv p) = poly [])" -by (blast elim: pderiv_zero_obj [THEN impE]) -declare pderiv_zero [simp] - -lemma poly_pderiv_welldef: "poly p = poly q ==> (poly (pderiv p) = poly (pderiv q))" -apply (cut_tac p = "p +++ --q" in pderiv_zero_obj) -apply (simp add: fun_eq poly_add poly_minus poly_pderiv_add poly_pderiv_minus del: pderiv_zero) -done - -lemma lemma_order_pderiv [rule_format]: - "\p q a. 0 < n & - poly (pderiv p) \ poly [] & - poly p = poly ([- a, 1] %^ n *** q) & ~ [- a, 1] divides q - --> n = Suc (order a (pderiv p))" -apply (induct "n", safe) -apply (rule order_unique_lemma, rule conjI, assumption) -apply (subgoal_tac "\r. r divides (pderiv p) = r divides (pderiv ([-a, 1] %^ Suc n *** q))") -apply (drule_tac [2] poly_pderiv_welldef) - prefer 2 apply (simp add: divides_def del: pmult_Cons pexp_Suc) -apply (simp del: pmult_Cons pexp_Suc) -apply (rule conjI) -apply (simp add: divides_def fun_eq del: pmult_Cons pexp_Suc) -apply (rule_tac x = "[-a, 1] *** (pderiv q) +++ real (Suc n) %* q" in exI) -apply (simp add: poly_pderiv_mult poly_pderiv_exp_prime poly_add poly_mult poly_cmult right_distrib mult_ac del: pmult_Cons pexp_Suc) -apply (simp add: poly_mult right_distrib left_distrib mult_ac del: pmult_Cons) -apply (erule_tac V = "\r. r divides pderiv p = r divides pderiv ([- a, 1] %^ Suc n *** q)" in thin_rl) -apply (unfold divides_def) -apply (simp (no_asm) add: poly_pderiv_mult poly_pderiv_exp_prime fun_eq poly_add poly_mult del: pmult_Cons pexp_Suc) -apply (rule contrapos_np, assumption) -apply (rotate_tac 3, erule contrapos_np) -apply (simp del: pmult_Cons pexp_Suc, safe) -apply (rule_tac x = "inverse (real (Suc n)) %* (qa +++ -- (pderiv q))" in exI) -apply (subgoal_tac "poly ([-a, 1] %^ n *** q) = poly ([-a, 1] %^ n *** ([-a, 1] *** (inverse (real (Suc n)) %* (qa +++ -- (pderiv q))))) ") -apply (drule poly_mult_left_cancel [THEN iffD1], simp) -apply (simp add: fun_eq poly_mult poly_add poly_cmult poly_minus del: pmult_Cons mult_cancel_left, safe) -apply (rule_tac c1 = "real (Suc n)" in real_mult_left_cancel [THEN iffD1]) -apply (simp (no_asm)) -apply (subgoal_tac "real (Suc n) * (poly ([- a, 1] %^ n) xa * poly q xa) = - (poly qa xa + - poly (pderiv q) xa) * - (poly ([- a, 1] %^ n) xa * - ((- a + xa) * (inverse (real (Suc n)) * real (Suc n))))") -apply (simp only: mult_ac) -apply (rotate_tac 2) -apply (drule_tac x = xa in spec) -apply (simp add: left_distrib mult_ac del: pmult_Cons) -done - -lemma order_pderiv: "[| poly (pderiv p) \ poly []; order a p \ 0 |] - ==> (order a p = Suc (order a (pderiv p)))" -apply (case_tac "poly p = poly []") -apply (auto dest: pderiv_zero) -apply (drule_tac a = a and p = p in order_decomp) -using neq0_conv -apply (blast intro: lemma_order_pderiv) -done - -text{*Now justify the standard squarefree decomposition, i.e. f / gcd(f,f'). *} - -lemma poly_squarefree_decomp_order: "[| poly (pderiv p) \ poly []; - poly p = poly (q *** d); - poly (pderiv p) = poly (e *** d); - poly d = poly (r *** p +++ s *** pderiv p) - |] ==> order a q = (if order a p = 0 then 0 else 1)" -apply (subgoal_tac "order a p = order a q + order a d") -apply (rule_tac [2] s = "order a (q *** d)" in trans) -prefer 2 apply (blast intro: order_poly) -apply (rule_tac [2] order_mult) - prefer 2 apply force -apply (case_tac "order a p = 0", simp) -apply (subgoal_tac "order a (pderiv p) = order a e + order a d") -apply (rule_tac [2] s = "order a (e *** d)" in trans) -prefer 2 apply (blast intro: order_poly) -apply (rule_tac [2] order_mult) - prefer 2 apply force -apply (case_tac "poly p = poly []") -apply (drule_tac p = p in pderiv_zero, simp) -apply (drule order_pderiv, assumption) -apply (subgoal_tac "order a (pderiv p) \ order a d") -apply (subgoal_tac [2] " ([-a, 1] %^ (order a (pderiv p))) divides d") - prefer 2 apply (simp add: poly_entire order_divides) -apply (subgoal_tac [2] " ([-a, 1] %^ (order a (pderiv p))) divides p & ([-a, 1] %^ (order a (pderiv p))) divides (pderiv p) ") - prefer 3 apply (simp add: order_divides) - prefer 2 apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe) -apply (rule_tac x = "r *** qa +++ s *** qaa" in exI) -apply (simp add: fun_eq poly_add poly_mult left_distrib right_distrib mult_ac del: pexp_Suc pmult_Cons, auto) -done - - -lemma poly_squarefree_decomp_order2: "[| poly (pderiv p) \ poly []; - poly p = poly (q *** d); - poly (pderiv p) = poly (e *** d); - poly d = poly (r *** p +++ s *** pderiv p) - |] ==> \a. order a q = (if order a p = 0 then 0 else 1)" -apply (blast intro: poly_squarefree_decomp_order) -done - -lemma order_pderiv2: "[| poly (pderiv p) \ poly []; order a p \ 0 |] - ==> (order a (pderiv p) = n) = (order a p = Suc n)" -apply (auto dest: order_pderiv) -done - -lemma rsquarefree_roots: - "rsquarefree p = (\a. ~(poly p a = 0 & poly (pderiv p) a = 0))" -apply (simp add: rsquarefree_def) -apply (case_tac "poly p = poly []", simp, simp) -apply (case_tac "poly (pderiv p) = poly []") -apply simp -apply (drule pderiv_iszero, clarify) -apply (subgoal_tac "\a. order a p = order a [h]") -apply (simp add: fun_eq) -apply (rule allI) -apply (cut_tac p = "[h]" and a = a in order_root) -apply (simp add: fun_eq) -apply (blast intro: order_poly) -apply (auto simp add: order_root order_pderiv2) -apply (erule_tac x="a" in allE, simp) -done - -lemma poly_squarefree_decomp: "[| poly (pderiv p) \ poly []; - poly p = poly (q *** d); - poly (pderiv p) = poly (e *** d); - poly d = poly (r *** p +++ s *** pderiv p) - |] ==> rsquarefree q & (\a. (poly q a = 0) = (poly p a = 0))" -apply (frule poly_squarefree_decomp_order2, assumption+) -apply (case_tac "poly p = poly []") -apply (blast dest: pderiv_zero) -apply (simp (no_asm) add: rsquarefree_def order_root del: pmult_Cons) -apply (simp add: poly_entire del: pmult_Cons) -done - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/Fact.thy --- a/src/HOL/Hyperreal/Fact.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,76 +0,0 @@ -(* Title : Fact.thy - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Conversion to Isar and new proofs by Lawrence C Paulson, 2004 -*) - -header{*Factorial Function*} - -theory Fact -imports "../Real/RealDef" -begin - -consts fact :: "nat => nat" -primrec - fact_0: "fact 0 = 1" - fact_Suc: "fact (Suc n) = (Suc n) * fact n" - - -lemma fact_gt_zero [simp]: "0 < fact n" -by (induct n) auto - -lemma fact_not_eq_zero [simp]: "fact n \ 0" -by simp - -lemma real_of_nat_fact_not_zero [simp]: "real (fact n) \ 0" -by auto - -lemma real_of_nat_fact_gt_zero [simp]: "0 < real(fact n)" -by auto - -lemma real_of_nat_fact_ge_zero [simp]: "0 \ real(fact n)" -by simp - -lemma fact_ge_one [simp]: "1 \ fact n" -by (induct n) auto - -lemma fact_mono: "m \ n ==> fact m \ fact n" -apply (drule le_imp_less_or_eq) -apply (auto dest!: less_imp_Suc_add) -apply (induct_tac k, auto) -done - -text{*Note that @{term "fact 0 = fact 1"}*} -lemma fact_less_mono: "[| 0 < m; m < n |] ==> fact m < fact n" -apply (drule_tac m = m in less_imp_Suc_add, auto) -apply (induct_tac k, auto) -done - -lemma inv_real_of_nat_fact_gt_zero [simp]: "0 < inverse (real (fact n))" -by (auto simp add: positive_imp_inverse_positive) - -lemma inv_real_of_nat_fact_ge_zero [simp]: "0 \ inverse (real (fact n))" -by (auto intro: order_less_imp_le) - -lemma fact_diff_Suc [rule_format]: - "n < Suc m ==> fact (Suc m - n) = (Suc m - n) * fact (m - n)" -apply (induct n arbitrary: m) -apply auto -apply (drule_tac x = "m - 1" in meta_spec, auto) -done - -lemma fact_num0 [simp]: "fact 0 = 1" -by auto - -lemma fact_num_eq_if: "fact m = (if m=0 then 1 else m * fact (m - 1))" -by (cases m) auto - -lemma fact_add_num_eq_if: - "fact (m + n) = (if m + n = 0 then 1 else (m + n) * fact (m + n - 1))" -by (cases "m + n") auto - -lemma fact_add_num_eq_if2: - "fact (m + n) = (if m = 0 then fact n else (m + n) * fact ((m - 1) + n))" -by (cases m) auto - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/FrechetDeriv.thy --- a/src/HOL/Hyperreal/FrechetDeriv.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,503 +0,0 @@ -(* Title : FrechetDeriv.thy - ID : $Id$ - Author : Brian Huffman -*) - -header {* Frechet Derivative *} - -theory FrechetDeriv -imports Lim -begin - -definition - fderiv :: - "['a::real_normed_vector \ 'b::real_normed_vector, 'a, 'a \ 'b] \ bool" - -- {* Frechet derivative: D is derivative of function f at x *} - ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where - "FDERIV f x :> D = (bounded_linear D \ - (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0)" - -lemma FDERIV_I: - "\bounded_linear D; (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0\ - \ FDERIV f x :> D" -by (simp add: fderiv_def) - -lemma FDERIV_D: - "FDERIV f x :> D \ (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0" -by (simp add: fderiv_def) - -lemma FDERIV_bounded_linear: "FDERIV f x :> D \ bounded_linear D" -by (simp add: fderiv_def) - -lemma bounded_linear_zero: - "bounded_linear (\x::'a::real_normed_vector. 0::'b::real_normed_vector)" -proof - show "(0::'b) = 0 + 0" by simp - fix r show "(0::'b) = scaleR r 0" by simp - have "\x::'a. norm (0::'b) \ norm x * 0" by simp - thus "\K. \x::'a. norm (0::'b) \ norm x * K" .. -qed - -lemma FDERIV_const: "FDERIV (\x. k) x :> (\h. 0)" -by (simp add: fderiv_def bounded_linear_zero) - -lemma bounded_linear_ident: - "bounded_linear (\x::'a::real_normed_vector. x)" -proof - fix x y :: 'a show "x + y = x + y" by simp - fix r and x :: 'a show "scaleR r x = scaleR r x" by simp - have "\x::'a. norm x \ norm x * 1" by simp - thus "\K. \x::'a. norm x \ norm x * K" .. -qed - -lemma FDERIV_ident: "FDERIV (\x. x) x :> (\h. h)" -by (simp add: fderiv_def bounded_linear_ident) - -subsection {* Addition *} - -lemma add_diff_add: - fixes a b c d :: "'a::ab_group_add" - shows "(a + c) - (b + d) = (a - b) + (c - d)" -by simp - -lemma bounded_linear_add: - assumes "bounded_linear f" - assumes "bounded_linear g" - shows "bounded_linear (\x. f x + g x)" -proof - - interpret f: bounded_linear [f] by fact - interpret g: bounded_linear [g] by fact - show ?thesis apply (unfold_locales) - apply (simp only: f.add g.add add_ac) - apply (simp only: f.scaleR g.scaleR scaleR_right_distrib) - apply (rule f.pos_bounded [THEN exE], rename_tac Kf) - apply (rule g.pos_bounded [THEN exE], rename_tac Kg) - apply (rule_tac x="Kf + Kg" in exI, safe) - apply (subst right_distrib) - apply (rule order_trans [OF norm_triangle_ineq]) - apply (rule add_mono, erule spec, erule spec) - done -qed - -lemma norm_ratio_ineq: - fixes x y :: "'a::real_normed_vector" - fixes h :: "'b::real_normed_vector" - shows "norm (x + y) / norm h \ norm x / norm h + norm y / norm h" -apply (rule ord_le_eq_trans) -apply (rule divide_right_mono) -apply (rule norm_triangle_ineq) -apply (rule norm_ge_zero) -apply (rule add_divide_distrib) -done - -lemma FDERIV_add: - assumes f: "FDERIV f x :> F" - assumes g: "FDERIV g x :> G" - shows "FDERIV (\x. f x + g x) x :> (\h. F h + G h)" -proof (rule FDERIV_I) - show "bounded_linear (\h. F h + G h)" - apply (rule bounded_linear_add) - apply (rule FDERIV_bounded_linear [OF f]) - apply (rule FDERIV_bounded_linear [OF g]) - done -next - have f': "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" - using f by (rule FDERIV_D) - have g': "(\h. norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" - using g by (rule FDERIV_D) - from f' g' - have "(\h. norm (f (x + h) - f x - F h) / norm h - + norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" - by (rule LIM_add_zero) - thus "(\h. norm (f (x + h) + g (x + h) - (f x + g x) - (F h + G h)) - / norm h) -- 0 --> 0" - apply (rule real_LIM_sandwich_zero) - apply (simp add: divide_nonneg_pos) - apply (simp only: add_diff_add) - apply (rule norm_ratio_ineq) - done -qed - -subsection {* Subtraction *} - -lemma bounded_linear_minus: - assumes "bounded_linear f" - shows "bounded_linear (\x. - f x)" -proof - - interpret f: bounded_linear [f] by fact - show ?thesis apply (unfold_locales) - apply (simp add: f.add) - apply (simp add: f.scaleR) - apply (simp add: f.bounded) - done -qed - -lemma FDERIV_minus: - "FDERIV f x :> F \ FDERIV (\x. - f x) x :> (\h. - F h)" -apply (rule FDERIV_I) -apply (rule bounded_linear_minus) -apply (erule FDERIV_bounded_linear) -apply (simp only: fderiv_def minus_diff_minus norm_minus_cancel) -done - -lemma FDERIV_diff: - "\FDERIV f x :> F; FDERIV g x :> G\ - \ FDERIV (\x. f x - g x) x :> (\h. F h - G h)" -by (simp only: diff_minus FDERIV_add FDERIV_minus) - -subsection {* Continuity *} - -lemma FDERIV_isCont: - assumes f: "FDERIV f x :> F" - shows "isCont f x" -proof - - from f interpret F: bounded_linear ["F"] by (rule FDERIV_bounded_linear) - have "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" - by (rule FDERIV_D [OF f]) - hence "(\h. norm (f (x + h) - f x - F h) / norm h * norm h) -- 0 --> 0" - by (intro LIM_mult_zero LIM_norm_zero LIM_ident) - hence "(\h. norm (f (x + h) - f x - F h)) -- 0 --> 0" - by (simp cong: LIM_cong) - hence "(\h. f (x + h) - f x - F h) -- 0 --> 0" - by (rule LIM_norm_zero_cancel) - hence "(\h. f (x + h) - f x - F h + F h) -- 0 --> 0" - by (intro LIM_add_zero F.LIM_zero LIM_ident) - hence "(\h. f (x + h) - f x) -- 0 --> 0" - by simp - thus "isCont f x" - unfolding isCont_iff by (rule LIM_zero_cancel) -qed - -subsection {* Composition *} - -lemma real_divide_cancel_lemma: - fixes a b c :: real - shows "(b = 0 \ a = 0) \ (a / b) * (b / c) = a / c" -by simp - -lemma bounded_linear_compose: - assumes "bounded_linear f" - assumes "bounded_linear g" - shows "bounded_linear (\x. f (g x))" -proof - - interpret f: bounded_linear [f] by fact - interpret g: bounded_linear [g] by fact - show ?thesis proof (unfold_locales) - fix x y show "f (g (x + y)) = f (g x) + f (g y)" - by (simp only: f.add g.add) - next - fix r x show "f (g (scaleR r x)) = scaleR r (f (g x))" - by (simp only: f.scaleR g.scaleR) - next - from f.pos_bounded - obtain Kf where f: "\x. norm (f x) \ norm x * Kf" and Kf: "0 < Kf" by fast - from g.pos_bounded - obtain Kg where g: "\x. norm (g x) \ norm x * Kg" by fast - show "\K. \x. norm (f (g x)) \ norm x * K" - proof (intro exI allI) - fix x - have "norm (f (g x)) \ norm (g x) * Kf" - using f . - also have "\ \ (norm x * Kg) * Kf" - using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) - also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" - by (rule mult_assoc) - finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . - qed - qed -qed - -lemma FDERIV_compose: - fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" - fixes g :: "'b::real_normed_vector \ 'c::real_normed_vector" - assumes f: "FDERIV f x :> F" - assumes g: "FDERIV g (f x) :> G" - shows "FDERIV (\x. g (f x)) x :> (\h. G (F h))" -proof (rule FDERIV_I) - from FDERIV_bounded_linear [OF g] FDERIV_bounded_linear [OF f] - show "bounded_linear (\h. G (F h))" - by (rule bounded_linear_compose) -next - let ?Rf = "\h. f (x + h) - f x - F h" - let ?Rg = "\k. g (f x + k) - g (f x) - G k" - let ?k = "\h. f (x + h) - f x" - let ?Nf = "\h. norm (?Rf h) / norm h" - let ?Ng = "\h. norm (?Rg (?k h)) / norm (?k h)" - from f interpret F: bounded_linear ["F"] by (rule FDERIV_bounded_linear) - from g interpret G: bounded_linear ["G"] by (rule FDERIV_bounded_linear) - from F.bounded obtain kF where kF: "\x. norm (F x) \ norm x * kF" by fast - from G.bounded obtain kG where kG: "\x. norm (G x) \ norm x * kG" by fast - - let ?fun2 = "\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)" - - show "(\h. norm (g (f (x + h)) - g (f x) - G (F h)) / norm h) -- 0 --> 0" - proof (rule real_LIM_sandwich_zero) - have Nf: "?Nf -- 0 --> 0" - using FDERIV_D [OF f] . - - have Ng1: "isCont (\k. norm (?Rg k) / norm k) 0" - by (simp add: isCont_def FDERIV_D [OF g]) - have Ng2: "?k -- 0 --> 0" - apply (rule LIM_zero) - apply (fold isCont_iff) - apply (rule FDERIV_isCont [OF f]) - done - have Ng: "?Ng -- 0 --> 0" - using isCont_LIM_compose [OF Ng1 Ng2] by simp - - have "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) - -- 0 --> 0 * kG + 0 * (0 + kF)" - by (intro LIM_add LIM_mult LIM_const Nf Ng) - thus "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) -- 0 --> 0" - by simp - next - fix h::'a assume h: "h \ 0" - thus "0 \ norm (g (f (x + h)) - g (f x) - G (F h)) / norm h" - by (simp add: divide_nonneg_pos) - next - fix h::'a assume h: "h \ 0" - have "g (f (x + h)) - g (f x) - G (F h) = G (?Rf h) + ?Rg (?k h)" - by (simp add: G.diff) - hence "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h - = norm (G (?Rf h) + ?Rg (?k h)) / norm h" - by (rule arg_cong) - also have "\ \ norm (G (?Rf h)) / norm h + norm (?Rg (?k h)) / norm h" - by (rule norm_ratio_ineq) - also have "\ \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" - proof (rule add_mono) - show "norm (G (?Rf h)) / norm h \ ?Nf h * kG" - apply (rule ord_le_eq_trans) - apply (rule divide_right_mono [OF kG norm_ge_zero]) - apply simp - done - next - have "norm (?Rg (?k h)) / norm h = ?Ng h * (norm (?k h) / norm h)" - apply (rule real_divide_cancel_lemma [symmetric]) - apply (simp add: G.zero) - done - also have "\ \ ?Ng h * (?Nf h + kF)" - proof (rule mult_left_mono) - have "norm (?k h) / norm h = norm (?Rf h + F h) / norm h" - by simp - also have "\ \ ?Nf h + norm (F h) / norm h" - by (rule norm_ratio_ineq) - also have "\ \ ?Nf h + kF" - apply (rule add_left_mono) - apply (subst pos_divide_le_eq, simp add: h) - apply (subst mult_commute) - apply (rule kF) - done - finally show "norm (?k h) / norm h \ ?Nf h + kF" . - next - show "0 \ ?Ng h" - apply (case_tac "f (x + h) - f x = 0", simp) - apply (rule divide_nonneg_pos [OF norm_ge_zero]) - apply simp - done - qed - finally show "norm (?Rg (?k h)) / norm h \ ?Ng h * (?Nf h + kF)" . - qed - finally show "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h - \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" . - qed -qed - -subsection {* Product Rule *} - -lemma (in bounded_bilinear) FDERIV_lemma: - "a' ** b' - a ** b - (a ** B + A ** b) - = a ** (b' - b - B) + (a' - a - A) ** b' + A ** (b' - b)" -by (simp add: diff_left diff_right) - -lemma (in bounded_bilinear) FDERIV: - fixes x :: "'d::real_normed_vector" - assumes f: "FDERIV f x :> F" - assumes g: "FDERIV g x :> G" - shows "FDERIV (\x. f x ** g x) x :> (\h. f x ** G h + F h ** g x)" -proof (rule FDERIV_I) - show "bounded_linear (\h. f x ** G h + F h ** g x)" - apply (rule bounded_linear_add) - apply (rule bounded_linear_compose [OF bounded_linear_right]) - apply (rule FDERIV_bounded_linear [OF g]) - apply (rule bounded_linear_compose [OF bounded_linear_left]) - apply (rule FDERIV_bounded_linear [OF f]) - done -next - from bounded_linear.bounded [OF FDERIV_bounded_linear [OF f]] - obtain KF where norm_F: "\x. norm (F x) \ norm x * KF" by fast - - from pos_bounded obtain K where K: "0 < K" and norm_prod: - "\a b. norm (a ** b) \ norm a * norm b * K" by fast - - let ?Rf = "\h. f (x + h) - f x - F h" - let ?Rg = "\h. g (x + h) - g x - G h" - - let ?fun1 = "\h. - norm (f x ** ?Rg h + ?Rf h ** g (x + h) + F h ** (g (x + h) - g x)) / - norm h" - - let ?fun2 = "\h. - norm (f x) * (norm (?Rg h) / norm h) * K + - norm (?Rf h) / norm h * norm (g (x + h)) * K + - KF * norm (g (x + h) - g x) * K" - - have "?fun1 -- 0 --> 0" - proof (rule real_LIM_sandwich_zero) - from f g isCont_iff [THEN iffD1, OF FDERIV_isCont [OF g]] - have "?fun2 -- 0 --> - norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K" - by (intro LIM_add LIM_mult LIM_const LIM_norm LIM_zero FDERIV_D) - thus "?fun2 -- 0 --> 0" - by simp - next - fix h::'d assume "h \ 0" - thus "0 \ ?fun1 h" - by (simp add: divide_nonneg_pos) - next - fix h::'d assume "h \ 0" - have "?fun1 h \ (norm (f x) * norm (?Rg h) * K + - norm (?Rf h) * norm (g (x + h)) * K + - norm h * KF * norm (g (x + h) - g x) * K) / norm h" - by (intro - divide_right_mono mult_mono' - order_trans [OF norm_triangle_ineq add_mono] - order_trans [OF norm_prod mult_right_mono] - mult_nonneg_nonneg order_refl norm_ge_zero norm_F - K [THEN order_less_imp_le] - ) - also have "\ = ?fun2 h" - by (simp add: add_divide_distrib) - finally show "?fun1 h \ ?fun2 h" . - qed - thus "(\h. - norm (f (x + h) ** g (x + h) - f x ** g x - (f x ** G h + F h ** g x)) - / norm h) -- 0 --> 0" - by (simp only: FDERIV_lemma) -qed - -lemmas FDERIV_mult = bounded_bilinear_locale.mult.prod.FDERIV - -lemmas FDERIV_scaleR = bounded_bilinear_locale.scaleR.prod.FDERIV - - -subsection {* Powers *} - -lemma FDERIV_power_Suc: - fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}" - shows "FDERIV (\x. x ^ Suc n) x :> (\h. (1 + of_nat n) * x ^ n * h)" - apply (induct n) - apply (simp add: power_Suc FDERIV_ident) - apply (drule FDERIV_mult [OF FDERIV_ident]) - apply (simp only: of_nat_Suc left_distrib mult_1_left) - apply (simp only: power_Suc right_distrib add_ac mult_ac) -done - -lemma FDERIV_power: - fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}" - shows "FDERIV (\x. x ^ n) x :> (\h. of_nat n * x ^ (n - 1) * h)" - apply (cases n) - apply (simp add: FDERIV_const) - apply (simp add: FDERIV_power_Suc) - done - - -subsection {* Inverse *} - -lemma inverse_diff_inverse: - "\(a::'a::division_ring) \ 0; b \ 0\ - \ inverse a - inverse b = - (inverse a * (a - b) * inverse b)" -by (simp add: right_diff_distrib left_diff_distrib mult_assoc) - -lemmas bounded_linear_mult_const = - bounded_bilinear_locale.mult.prod.bounded_linear_left [THEN bounded_linear_compose] - -lemmas bounded_linear_const_mult = - bounded_bilinear_locale.mult.prod.bounded_linear_right [THEN bounded_linear_compose] - -lemma FDERIV_inverse: - fixes x :: "'a::real_normed_div_algebra" - assumes x: "x \ 0" - shows "FDERIV inverse x :> (\h. - (inverse x * h * inverse x))" - (is "FDERIV ?inv _ :> _") -proof (rule FDERIV_I) - show "bounded_linear (\h. - (?inv x * h * ?inv x))" - apply (rule bounded_linear_minus) - apply (rule bounded_linear_mult_const) - apply (rule bounded_linear_const_mult) - apply (rule bounded_linear_ident) - done -next - show "(\h. norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h) - -- 0 --> 0" - proof (rule LIM_equal2) - show "0 < norm x" using x by simp - next - fix h::'a - assume 1: "h \ 0" - assume "norm (h - 0) < norm x" - hence "h \ -x" by clarsimp - hence 2: "x + h \ 0" - apply (rule contrapos_nn) - apply (rule sym) - apply (erule equals_zero_I) - done - show "norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h - = norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" - apply (subst inverse_diff_inverse [OF 2 x]) - apply (subst minus_diff_minus) - apply (subst norm_minus_cancel) - apply (simp add: left_diff_distrib) - done - next - show "(\h. norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h) - -- 0 --> 0" - proof (rule real_LIM_sandwich_zero) - show "(\h. norm (?inv (x + h) - ?inv x) * norm (?inv x)) - -- 0 --> 0" - apply (rule LIM_mult_left_zero) - apply (rule LIM_norm_zero) - apply (rule LIM_zero) - apply (rule LIM_offset_zero) - apply (rule LIM_inverse) - apply (rule LIM_ident) - apply (rule x) - done - next - fix h::'a assume h: "h \ 0" - show "0 \ norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" - apply (rule divide_nonneg_pos) - apply (rule norm_ge_zero) - apply (simp add: h) - done - next - fix h::'a assume h: "h \ 0" - have "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h - \ norm (?inv (x + h) - ?inv x) * norm h * norm (?inv x) / norm h" - apply (rule divide_right_mono [OF _ norm_ge_zero]) - apply (rule order_trans [OF norm_mult_ineq]) - apply (rule mult_right_mono [OF _ norm_ge_zero]) - apply (rule norm_mult_ineq) - done - also have "\ = norm (?inv (x + h) - ?inv x) * norm (?inv x)" - by simp - finally show "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h - \ norm (?inv (x + h) - ?inv x) * norm (?inv x)" . - qed - qed -qed - -subsection {* Alternate definition *} - -lemma field_fderiv_def: - fixes x :: "'a::real_normed_field" shows - "FDERIV f x :> (\h. h * D) = (\h. (f (x + h) - f x) / h) -- 0 --> D" - apply (unfold fderiv_def) - apply (simp add: bounded_bilinear_locale.mult.prod.bounded_linear_left) - apply (simp cong: LIM_cong add: nonzero_norm_divide [symmetric]) - apply (subst diff_divide_distrib) - apply (subst times_divide_eq_left [symmetric]) - apply (simp cong: LIM_cong) - apply (simp add: LIM_norm_zero_iff LIM_zero_iff) -done - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/Integration.thy --- a/src/HOL/Hyperreal/Integration.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,882 +0,0 @@ -(* ID : $Id$ - Author : Jacques D. Fleuriot - Copyright : 2000 University of Edinburgh - Conversion to Isar and new proofs by Lawrence C Paulson, 2004 -*) - -header{*Theory of Integration*} - -theory Integration -imports MacLaurin -begin - -text{*We follow John Harrison in formalizing the Gauge integral.*} - -definition - --{*Partitions and tagged partitions etc.*} - - partition :: "[(real*real),nat => real] => bool" where - [code del]: "partition = (%(a,b) D. D 0 = a & - (\N. (\n < N. D(n) < D(Suc n)) & - (\n \ N. D(n) = b)))" - -definition - psize :: "(nat => real) => nat" where - [code del]:"psize D = (SOME N. (\n < N. D(n) < D(Suc n)) & - (\n \ N. D(n) = D(N)))" - -definition - tpart :: "[(real*real),((nat => real)*(nat =>real))] => bool" where - [code del]:"tpart = (%(a,b) (D,p). partition(a,b) D & - (\n. D(n) \ p(n) & p(n) \ D(Suc n)))" - - --{*Gauges and gauge-fine divisions*} - -definition - gauge :: "[real => bool, real => real] => bool" where - [code del]:"gauge E g = (\x. E x --> 0 < g(x))" - -definition - fine :: "[real => real, ((nat => real)*(nat => real))] => bool" where - [code del]:"fine = (%g (D,p). \n. n < (psize D) --> D(Suc n) - D(n) < g(p n))" - - --{*Riemann sum*} - -definition - rsum :: "[((nat=>real)*(nat=>real)),real=>real] => real" where - "rsum = (%(D,p) f. \n=0..real,real] => bool" where - [code del]: "Integral = (%(a,b) f k. \e > 0. - (\g. gauge(%x. a \ x & x \ b) g & - (\D p. tpart(a,b) (D,p) & fine(g)(D,p) --> - \rsum(D,p) f - k\ < e)))" - - -lemma partition_zero [simp]: "a = b ==> psize (%n. if n = 0 then a else b) = 0" -by (auto simp add: psize_def) - -lemma partition_one [simp]: "a < b ==> psize (%n. if n = 0 then a else b) = 1" -apply (simp add: psize_def) -apply (rule some_equality, auto) -apply (drule_tac x = 1 in spec, auto) -done - -lemma partition_single [simp]: - "a \ b ==> partition(a,b)(%n. if n = 0 then a else b)" -by (auto simp add: partition_def order_le_less) - -lemma partition_lhs: "partition(a,b) D ==> (D(0) = a)" -by (simp add: partition_def) - -lemma partition: - "(partition(a,b) D) = - ((D 0 = a) & - (\n < psize D. D n < D(Suc n)) & - (\n \ psize D. D n = b))" -apply (simp add: partition_def, auto) -apply (subgoal_tac [!] "psize D = N", auto) -apply (simp_all (no_asm) add: psize_def) -apply (rule_tac [!] some_equality, blast) - prefer 2 apply blast -apply (rule_tac [!] ccontr) -apply (simp_all add: linorder_neq_iff, safe) -apply (drule_tac x = Na in spec) -apply (rotate_tac 3) -apply (drule_tac x = "Suc Na" in spec, simp) -apply (rotate_tac 2) -apply (drule_tac x = N in spec, simp) -apply (drule_tac x = Na in spec) -apply (drule_tac x = "Suc Na" and P = "%n. Na\n \ D n = D Na" in spec, auto) -done - -lemma partition_rhs: "partition(a,b) D ==> (D(psize D) = b)" -by (simp add: partition) - -lemma partition_rhs2: "[|partition(a,b) D; psize D \ n |] ==> (D n = b)" -by (simp add: partition) - -lemma lemma_partition_lt_gen [rule_format]: - "partition(a,b) D & m + Suc d \ n & n \ (psize D) --> D(m) < D(m + Suc d)" -apply (induct "d", auto simp add: partition) -apply (blast dest: Suc_le_lessD intro: less_le_trans order_less_trans) -done - -lemma less_eq_add_Suc: "m < n ==> \d. n = m + Suc d" -by (auto simp add: less_iff_Suc_add) - -lemma partition_lt_gen: - "[|partition(a,b) D; m < n; n \ (psize D)|] ==> D(m) < D(n)" -by (auto dest: less_eq_add_Suc intro: lemma_partition_lt_gen) - -lemma partition_lt: "partition(a,b) D ==> n < (psize D) ==> D(0) < D(Suc n)" -apply (induct "n") -apply (auto simp add: partition) -done - -lemma partition_le: "partition(a,b) D ==> a \ b" -apply (frule partition [THEN iffD1], safe) -apply (drule_tac x = "psize D" and P="%n. psize D \ n --> ?P n" in spec, safe) -apply (case_tac "psize D = 0") -apply (drule_tac [2] n = "psize D - 1" in partition_lt, auto) -done - -lemma partition_gt: "[|partition(a,b) D; n < (psize D)|] ==> D(n) < D(psize D)" -by (auto intro: partition_lt_gen) - -lemma partition_eq: "partition(a,b) D ==> ((a = b) = (psize D = 0))" -apply (frule partition [THEN iffD1], safe) -apply (rotate_tac 2) -apply (drule_tac x = "psize D" in spec) -apply (rule ccontr) -apply (drule_tac n = "psize D - 1" in partition_lt) -apply auto -done - -lemma partition_lb: "partition(a,b) D ==> a \ D(r)" -apply (frule partition [THEN iffD1], safe) -apply (induct "r") -apply (cut_tac [2] y = "Suc r" and x = "psize D" in linorder_le_less_linear) -apply (auto intro: partition_le) -apply (drule_tac x = r in spec) -apply arith; -done - -lemma partition_lb_lt: "[| partition(a,b) D; psize D ~= 0 |] ==> a < D(Suc n)" -apply (rule_tac t = a in partition_lhs [THEN subst], assumption) -apply (cut_tac x = "Suc n" and y = "psize D" in linorder_le_less_linear) -apply (frule partition [THEN iffD1], safe) - apply (blast intro: partition_lt less_le_trans) -apply (rotate_tac 3) -apply (drule_tac x = "Suc n" in spec) -apply (erule impE) -apply (erule less_imp_le) -apply (frule partition_rhs) -apply (drule partition_gt[of _ _ _ 0], arith) -apply (simp (no_asm_simp)) -done - -lemma partition_ub: "partition(a,b) D ==> D(r) \ b" -apply (frule partition [THEN iffD1]) -apply (cut_tac x = "psize D" and y = r in linorder_le_less_linear, safe, blast) -apply (subgoal_tac "\x. D ((psize D) - x) \ b") -apply (rotate_tac 4) -apply (drule_tac x = "psize D - r" in spec) -apply (subgoal_tac "psize D - (psize D - r) = r") -apply simp -apply arith -apply safe -apply (induct_tac "x") -apply (simp (no_asm), blast) -apply (case_tac "psize D - Suc n = 0") -apply (erule_tac V = "\n. psize D \ n --> D n = b" in thin_rl) -apply (simp (no_asm_simp) add: partition_le) -apply (rule order_trans) - prefer 2 apply assumption -apply (subgoal_tac "psize D - n = Suc (psize D - Suc n)") - prefer 2 apply arith -apply (drule_tac x = "psize D - Suc n" in spec, simp) -done - -lemma partition_ub_lt: "[| partition(a,b) D; n < psize D |] ==> D(n) < b" -by (blast intro: partition_rhs [THEN subst] partition_gt) - -lemma lemma_partition_append1: - "[| partition (a, b) D1; partition (b, c) D2 |] - ==> (\n < psize D1 + psize D2. - (if n < psize D1 then D1 n else D2 (n - psize D1)) - < (if Suc n < psize D1 then D1 (Suc n) - else D2 (Suc n - psize D1))) & - (\n \ psize D1 + psize D2. - (if n < psize D1 then D1 n else D2 (n - psize D1)) = - (if psize D1 + psize D2 < psize D1 then D1 (psize D1 + psize D2) - else D2 (psize D1 + psize D2 - psize D1)))" -apply (auto intro: partition_lt_gen) -apply (subgoal_tac "psize D1 = Suc n") -apply (auto intro!: partition_lt_gen simp add: partition_lhs partition_ub_lt) -apply (auto intro!: partition_rhs2 simp add: partition_rhs - split: nat_diff_split) -done - -lemma lemma_psize1: - "[| partition (a, b) D1; partition (b, c) D2; N < psize D1 |] - ==> D1(N) < D2 (psize D2)" -apply (rule_tac y = "D1 (psize D1)" in order_less_le_trans) -apply (erule partition_gt) -apply (auto simp add: partition_rhs partition_le) -done - -lemma lemma_partition_append2: - "[| partition (a, b) D1; partition (b, c) D2 |] - ==> psize (%n. if n < psize D1 then D1 n else D2 (n - psize D1)) = - psize D1 + psize D2" -apply (unfold psize_def - [of "%n. if n < psize D1 then D1 n else D2 (n - psize D1)"]) -apply (rule some1_equality) - prefer 2 apply (blast intro!: lemma_partition_append1) -apply (rule ex1I, rule lemma_partition_append1) -apply (simp_all split: split_if_asm) - txt{*The case @{term "N < psize D1"}*} - apply (drule_tac x = "psize D1 + psize D2" and P="%n. ?P n & ?Q n" in spec) - apply (force dest: lemma_psize1) -apply (rule order_antisym); - txt{*The case @{term "psize D1 \ N"}: - proving @{term "N \ psize D1 + psize D2"}*} - apply (drule_tac x = "psize D1 + psize D2" in spec) - apply (simp add: partition_rhs2) -apply (case_tac "N - psize D1 < psize D2") - prefer 2 apply arith - txt{*Proving @{term "psize D1 + psize D2 \ N"}*} -apply (drule_tac x = "psize D1 + psize D2" and P="%n. N\n --> ?P n" in spec, simp) -apply (drule_tac a = b and b = c in partition_gt, assumption, simp) -done - -lemma tpart_eq_lhs_rhs: "[|psize D = 0; tpart(a,b) (D,p)|] ==> a = b" -by (auto simp add: tpart_def partition_eq) - -lemma tpart_partition: "tpart(a,b) (D,p) ==> partition(a,b) D" -by (simp add: tpart_def) - -lemma partition_append: - "[| tpart(a,b) (D1,p1); fine(g) (D1,p1); - tpart(b,c) (D2,p2); fine(g) (D2,p2) |] - ==> \D p. tpart(a,c) (D,p) & fine(g) (D,p)" -apply (rule_tac x = "%n. if n < psize D1 then D1 n else D2 (n - psize D1)" - in exI) -apply (rule_tac x = "%n. if n < psize D1 then p1 n else p2 (n - psize D1)" - in exI) -apply (case_tac "psize D1 = 0") -apply (auto dest: tpart_eq_lhs_rhs) - prefer 2 -apply (simp add: fine_def - lemma_partition_append2 [OF tpart_partition tpart_partition]) - --{*But must not expand @{term fine} in other subgoals*} -apply auto -apply (subgoal_tac "psize D1 = Suc n") - prefer 2 apply arith -apply (drule tpart_partition [THEN partition_rhs]) -apply (drule tpart_partition [THEN partition_lhs]) -apply (auto split: nat_diff_split) -apply (auto simp add: tpart_def) -defer 1 - apply (subgoal_tac "psize D1 = Suc n") - prefer 2 apply arith - apply (drule partition_rhs) - apply (drule partition_lhs, auto) -apply (simp split: nat_diff_split) -apply (subst partition) -apply (subst (1 2) lemma_partition_append2, assumption+) -apply (rule conjI) -apply (simp add: partition_lhs) -apply (drule lemma_partition_append1) -apply assumption; -apply (simp add: partition_rhs) -done - - -text{*We can always find a division that is fine wrt any gauge*} - -lemma partition_exists: - "[| a \ b; gauge(%x. a \ x & x \ b) g |] - ==> \D p. tpart(a,b) (D,p) & fine g (D,p)" -apply (cut_tac P = "%(u,v). a \ u & v \ b --> - (\D p. tpart (u,v) (D,p) & fine (g) (D,p))" - in lemma_BOLZANO2) -apply safe -apply (blast intro: order_trans)+ -apply (auto intro: partition_append) -apply (case_tac "a \ x & x \ b") -apply (rule_tac [2] x = 1 in exI, auto) -apply (rule_tac x = "g x" in exI) -apply (auto simp add: gauge_def) -apply (rule_tac x = "%n. if n = 0 then aa else ba" in exI) -apply (rule_tac x = "%n. if n = 0 then x else ba" in exI) -apply (auto simp add: tpart_def fine_def) -done - -text{*Lemmas about combining gauges*} - -lemma gauge_min: - "[| gauge(E) g1; gauge(E) g2 |] - ==> gauge(E) (%x. if g1(x) < g2(x) then g1(x) else g2(x))" -by (simp add: gauge_def) - -lemma fine_min: - "fine (%x. if g1(x) < g2(x) then g1(x) else g2(x)) (D,p) - ==> fine(g1) (D,p) & fine(g2) (D,p)" -by (auto simp add: fine_def split: split_if_asm) - - -text{*The integral is unique if it exists*} - -lemma Integral_unique: - "[| a \ b; Integral(a,b) f k1; Integral(a,b) f k2 |] ==> k1 = k2" -apply (simp add: Integral_def) -apply (drule_tac x = "\k1 - k2\ /2" in spec)+ -apply auto -apply (drule gauge_min, assumption) -apply (drule_tac g = "%x. if g x < ga x then g x else ga x" - in partition_exists, assumption, auto) -apply (drule fine_min) -apply (drule spec)+ -apply auto -apply (subgoal_tac "\(rsum (D,p) f - k2) - (rsum (D,p) f - k1)\ < \k1 - k2\") -apply arith -apply (drule add_strict_mono, assumption) -apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] - mult_less_cancel_right) -done - -lemma Integral_zero [simp]: "Integral(a,a) f 0" -apply (auto simp add: Integral_def) -apply (rule_tac x = "%x. 1" in exI) -apply (auto dest: partition_eq simp add: gauge_def tpart_def rsum_def) -done - -lemma sumr_partition_eq_diff_bounds [simp]: - "(\n=0.. b ==> Integral(a,b) (%x. 1) (b - a)" -apply (auto simp add: order_le_less rsum_def Integral_def) -apply (rule_tac x = "%x. b - a" in exI) -apply (auto simp add: gauge_def abs_less_iff tpart_def partition) -done - -lemma Integral_mult_const: "a \ b ==> Integral(a,b) (%x. c) (c*(b - a))" -apply (auto simp add: order_le_less rsum_def Integral_def) -apply (rule_tac x = "%x. b - a" in exI) -apply (auto simp add: setsum_right_distrib [symmetric] gauge_def abs_less_iff - right_diff_distrib [symmetric] partition tpart_def) -done - -lemma Integral_mult: - "[| a \ b; Integral(a,b) f k |] ==> Integral(a,b) (%x. c * f x) (c * k)" -apply (auto simp add: order_le_less - dest: Integral_unique [OF order_refl Integral_zero]) -apply (auto simp add: rsum_def Integral_def setsum_right_distrib[symmetric] mult_assoc) -apply (rule_tac a2 = c in abs_ge_zero [THEN order_le_imp_less_or_eq, THEN disjE]) - prefer 2 apply force -apply (drule_tac x = "e/abs c" in spec, auto) -apply (simp add: zero_less_mult_iff divide_inverse) -apply (rule exI, auto) -apply (drule spec)+ -apply auto -apply (rule_tac z1 = "inverse (abs c)" in real_mult_less_iff1 [THEN iffD1]) -apply (auto simp add: abs_mult divide_inverse [symmetric] right_diff_distrib [symmetric]) -done - -text{*Fundamental theorem of calculus (Part I)*} - -text{*"Straddle Lemma" : Swartz and Thompson: AMM 95(7) 1988 *} - -lemma choiceP: "\x. P(x) --> (\y. Q x y) ==> \f. (\x. P(x) --> Q x (f x))" -by (insert bchoice [of "Collect P" Q], simp) - -(*UNUSED -lemma choice2: "\x. (\y. R(y) & (\z. Q x y z)) ==> - \f fa. (\x. R(f x) & Q x (f x) (fa x))" -*) - - -(* new simplifications e.g. (y < x/n) = (y * n < x) are a real nuisance - they break the original proofs and make new proofs longer!*) -lemma strad1: - "\\xa::real. xa \ x \ \xa - x\ < s \ - \(f xa - f x) / (xa - x) - f' x\ * 2 < e; - 0 < e; a \ x; x \ b; 0 < s\ - \ \z. \z - x\ < s -->\f z - f x - f' x * (z - x)\ * 2 \ e * \z - x\" -apply auto -apply (case_tac "0 < \z - x\") - prefer 2 apply (simp add: zero_less_abs_iff) -apply (drule_tac x = z in spec) -apply (rule_tac z1 = "\inverse (z - x)\" - in real_mult_le_cancel_iff2 [THEN iffD1]) - apply simp -apply (simp del: abs_inverse abs_mult add: abs_mult [symmetric] - mult_assoc [symmetric]) -apply (subgoal_tac "inverse (z - x) * (f z - f x - f' x * (z - x)) - = (f z - f x) / (z - x) - f' x") - apply (simp add: abs_mult [symmetric] mult_ac diff_minus) -apply (subst mult_commute) -apply (simp add: left_distrib diff_minus) -apply (simp add: mult_assoc divide_inverse) -apply (simp add: left_distrib) -done - -lemma lemma_straddle: - "[| \x. a \ x & x \ b --> DERIV f x :> f'(x); 0 < e |] - ==> \g. gauge(%x. a \ x & x \ b) g & - (\x u v. a \ u & u \ x & x \ v & v \ b & (v - u) < g(x) - --> \(f(v) - f(u)) - (f'(x) * (v - u))\ \ e * (v - u))" -apply (simp add: gauge_def) -apply (subgoal_tac "\x. a \ x & x \ b --> - (\d > 0. \u v. u \ x & x \ v & (v - u) < d --> - \(f (v) - f (u)) - (f' (x) * (v - u))\ \ e * (v - u))") -apply (drule choiceP, auto) -apply (drule spec, auto) -apply (auto simp add: DERIV_iff2 LIM_def) -apply (drule_tac x = "e/2" in spec, auto) -apply (frule strad1, assumption+) -apply (rule_tac x = s in exI, auto) -apply (rule_tac x = u and y = v in linorder_cases, auto) -apply (rule_tac y = "\(f (v) - f (x)) - (f' (x) * (v - x))\ + - \(f (x) - f (u)) - (f' (x) * (x - u))\" - in order_trans) -apply (rule abs_triangle_ineq [THEN [2] order_trans]) -apply (simp add: right_diff_distrib) -apply (rule_tac t = "e* (v - u)" in real_sum_of_halves [THEN subst]) -apply (rule add_mono) -apply (rule_tac y = "(e/2) * \v - x\" in order_trans) - prefer 2 apply simp -apply (erule_tac [!] V= "\x'. x' ~= x & \x' - x\ < s --> ?P x'" in thin_rl) -apply (drule_tac x = v in spec, simp add: times_divide_eq) -apply (drule_tac x = u in spec, auto) -apply (subgoal_tac "\f u - f x - f' x * (u - x)\ = \f x - f u - f' x * (x - u)\") -apply (rule order_trans) -apply (auto simp add: abs_le_iff) -apply (simp add: right_diff_distrib) -done - -lemma FTC1: "[|a \ b; \x. a \ x & x \ b --> DERIV f x :> f'(x) |] - ==> Integral(a,b) f' (f(b) - f(a))" -apply (drule order_le_imp_less_or_eq, auto) -apply (auto simp add: Integral_def) -apply (rule ccontr) -apply (subgoal_tac "\e > 0. \g. gauge (%x. a \ x & x \ b) g & (\D p. tpart (a, b) (D, p) & fine g (D, p) --> \rsum (D, p) f' - (f b - f a)\ \ e)") -apply (rotate_tac 3) -apply (drule_tac x = "e/2" in spec, auto) -apply (drule spec, auto) -apply ((drule spec)+, auto) -apply (drule_tac e = "ea/ (b - a)" in lemma_straddle) -apply (auto simp add: zero_less_divide_iff) -apply (rule exI) -apply (auto simp add: tpart_def rsum_def) -apply (subgoal_tac "(\n=0..n=0.. Integral(a,b) f k2" -by simp - -lemma Integral_add: - "[| a \ b; b \ c; Integral(a,b) f' k1; Integral(b,c) f' k2; - \x. a \ x & x \ c --> DERIV f x :> f' x |] - ==> Integral(a,c) f' (k1 + k2)" -apply (rule FTC1 [THEN Integral_subst], auto) -apply (frule FTC1, auto) -apply (frule_tac a = b in FTC1, auto) -apply (drule_tac x = x in spec, auto) -apply (drule_tac ?k2.0 = "f b - f a" in Integral_unique) -apply (drule_tac [3] ?k2.0 = "f c - f b" in Integral_unique, auto) -done - -lemma partition_psize_Least: - "partition(a,b) D ==> psize D = (LEAST n. D(n) = b)" -apply (auto intro!: Least_equality [symmetric] partition_rhs) -apply (auto dest: partition_ub_lt simp add: linorder_not_less [symmetric]) -done - -lemma lemma_partition_bounded: "partition (a, c) D ==> ~ (\n. c < D(n))" -apply safe -apply (drule_tac r = n in partition_ub, auto) -done - -lemma lemma_partition_eq: - "partition (a, c) D ==> D = (%n. if D n < c then D n else c)" -apply (rule ext, auto) -apply (auto dest!: lemma_partition_bounded) -apply (drule_tac x = n in spec, auto) -done - -lemma lemma_partition_eq2: - "partition (a, c) D ==> D = (%n. if D n \ c then D n else c)" -apply (rule ext, auto) -apply (auto dest!: lemma_partition_bounded) -apply (drule_tac x = n in spec, auto) -done - -lemma partition_lt_Suc: - "[| partition(a,b) D; n < psize D |] ==> D n < D (Suc n)" -by (auto simp add: partition) - -lemma tpart_tag_eq: "tpart(a,c) (D,p) ==> p = (%n. if D n < c then p n else c)" -apply (rule ext) -apply (auto simp add: tpart_def) -apply (drule linorder_not_less [THEN iffD1]) -apply (drule_tac r = "Suc n" in partition_ub) -apply (drule_tac x = n in spec, auto) -done - -subsection{*Lemmas for Additivity Theorem of Gauge Integral*} - -lemma lemma_additivity1: - "[| a \ D n; D n < b; partition(a,b) D |] ==> n < psize D" -by (auto simp add: partition linorder_not_less [symmetric]) - -lemma lemma_additivity2: "[| a \ D n; partition(a,D n) D |] ==> psize D \ n" -apply (rule ccontr, drule not_leE) -apply (frule partition [THEN iffD1], safe) -apply (frule_tac r = "Suc n" in partition_ub) -apply (auto dest!: spec) -done - -lemma partition_eq_bound: - "[| partition(a,b) D; psize D < m |] ==> D(m) = D(psize D)" -by (auto simp add: partition) - -lemma partition_ub2: "[| partition(a,b) D; psize D < m |] ==> D(r) \ D(m)" -by (simp add: partition partition_ub) - -lemma tag_point_eq_partition_point: - "[| tpart(a,b) (D,p); psize D \ m |] ==> p(m) = D(m)" -apply (simp add: tpart_def, auto) -apply (drule_tac x = m in spec) -apply (auto simp add: partition_rhs2) -done - -lemma partition_lt_cancel: "[| partition(a,b) D; D m < D n |] ==> m < n" -apply (cut_tac less_linear [of n "psize D"], auto) -apply (cut_tac less_linear [of m n]) -apply (cut_tac less_linear [of m "psize D"]) -apply (auto dest: partition_gt) -apply (drule_tac n = m in partition_lt_gen, auto) -apply (frule partition_eq_bound) -apply (drule_tac [2] partition_gt, auto) -apply (metis dense_linear_order_class.dlo_simps(8) not_less partition_rhs partition_rhs2) -apply (metis le_less_trans dense_linear_order_class.dlo_simps(8) nat_le_linear partition_eq_bound partition_rhs2) -done - -lemma lemma_additivity4_psize_eq: - "[| a \ D n; D n < b; partition (a, b) D |] - ==> psize (%x. if D x < D n then D(x) else D n) = n" -apply (unfold psize_def) -apply (frule lemma_additivity1) -apply (assumption, assumption) -apply (rule some_equality) -apply (auto intro: partition_lt_Suc) -apply (drule_tac n = n in partition_lt_gen, assumption) -apply (arith, arith) -apply (cut_tac x = na and y = "psize D" in less_linear) -apply (auto dest: partition_lt_cancel) -apply (rule_tac x=N and y=n in linorder_cases) -apply (drule_tac x = n and P="%m. N \ m --> ?f m = ?g m" in spec, simp) -apply (drule_tac n = n in partition_lt_gen, auto) -apply (drule_tac x = n in spec) -apply (simp split: split_if_asm) -done - -lemma lemma_psize_left_less_psize: - "partition (a, b) D - ==> psize (%x. if D x < D n then D(x) else D n) \ psize D" -apply (frule_tac r = n in partition_ub) -apply (drule_tac x = "D n" in order_le_imp_less_or_eq) -apply (auto simp add: lemma_partition_eq [symmetric]) -apply (frule_tac r = n in partition_lb) -apply (drule (2) lemma_additivity4_psize_eq) -apply (rule ccontr, auto) -apply (frule_tac not_leE [THEN [2] partition_eq_bound]) -apply (auto simp add: partition_rhs) -done - -lemma lemma_psize_left_less_psize2: - "[| partition(a,b) D; na < psize (%x. if D x < D n then D(x) else D n) |] - ==> na < psize D" -by (erule lemma_psize_left_less_psize [THEN [2] less_le_trans]) - - -lemma lemma_additivity3: - "[| partition(a,b) D; D na < D n; D n < D (Suc na); - n < psize D |] - ==> False" -by (metis not_less_eq partition_lt_cancel real_of_nat_less_iff) - - -lemma psize_const [simp]: "psize (%x. k) = 0" -by (auto simp add: psize_def) - -lemma lemma_additivity3a: - "[| partition(a,b) D; D na < D n; D n < D (Suc na); - na < psize D |] - ==> False" -apply (frule_tac m = n in partition_lt_cancel) -apply (auto intro: lemma_additivity3) -done - -lemma better_lemma_psize_right_eq1: - "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) \ psize D - n" -apply (simp add: psize_def [of "(%x. D (x + n))"]); -apply (rule_tac a = "psize D - n" in someI2, auto) - apply (simp add: partition less_diff_conv) - apply (simp add: le_diff_conv partition_rhs2 split: nat_diff_split) -apply (drule_tac x = "psize D - n" in spec, auto) -apply (frule partition_rhs, safe) -apply (frule partition_lt_cancel, assumption) -apply (drule partition [THEN iffD1], safe) -apply (subgoal_tac "~ D (psize D - n + n) < D (Suc (psize D - n + n))") - apply blast -apply (drule_tac x = "Suc (psize D)" and P="%n. ?P n \ D n = D (psize D)" - in spec) -apply simp -done - -lemma psize_le_n: "partition (a, D n) D ==> psize D \ n" -apply (rule ccontr, drule not_leE) -apply (frule partition_lt_Suc, assumption) -apply (frule_tac r = "Suc n" in partition_ub, auto) -done - -lemma better_lemma_psize_right_eq1a: - "partition(a,D n) D ==> psize (%x. D (x + n)) \ psize D - n" -apply (simp add: psize_def [of "(%x. D (x + n))"]); -apply (rule_tac a = "psize D - n" in someI2, auto) - apply (simp add: partition less_diff_conv) - apply (simp add: le_diff_conv) -apply (case_tac "psize D \ n") - apply (force intro: partition_rhs2) - apply (simp add: partition linorder_not_le) -apply (rule ccontr, drule not_leE) -apply (frule psize_le_n) -apply (drule_tac x = "psize D - n" in spec, simp) -apply (drule partition [THEN iffD1], safe) -apply (drule_tac x = "Suc n" and P="%na. ?s \ na \ D na = D n" in spec, auto) -done - -lemma better_lemma_psize_right_eq: - "partition(a,b) D ==> psize (%x. D (x + n)) \ psize D - n" -apply (frule_tac r1 = n in partition_ub [THEN order_le_imp_less_or_eq]) -apply (blast intro: better_lemma_psize_right_eq1a better_lemma_psize_right_eq1) -done - -lemma lemma_psize_right_eq1: - "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) \ psize D" -apply (simp add: psize_def [of "(%x. D (x + n))"]) -apply (rule_tac a = "psize D - n" in someI2, auto) - apply (simp add: partition less_diff_conv) - apply (subgoal_tac "n \ psize D") - apply (simp add: partition le_diff_conv) - apply (rule ccontr, drule not_leE) - apply (drule_tac less_imp_le [THEN [2] partition_rhs2], assumption, simp) -apply (drule_tac x = "psize D" in spec) -apply (simp add: partition) -done - -(* should be combined with previous theorem; also proof has redundancy *) -lemma lemma_psize_right_eq1a: - "partition(a,D n) D ==> psize (%x. D (x + n)) \ psize D" -apply (simp add: psize_def [of "(%x. D (x + n))"]); -apply (rule_tac a = "psize D - n" in someI2, auto) - apply (simp add: partition less_diff_conv) - apply (case_tac "psize D \ n") - apply (force intro: partition_rhs2 simp add: le_diff_conv) - apply (simp add: partition le_diff_conv) -apply (rule ccontr, drule not_leE) -apply (drule_tac x = "psize D" in spec) -apply (simp add: partition) -done - -lemma lemma_psize_right_eq: - "[| partition(a,b) D |] ==> psize (%x. D (x + n)) \ psize D" -apply (frule_tac r1 = n in partition_ub [THEN order_le_imp_less_or_eq]) -apply (blast intro: lemma_psize_right_eq1a lemma_psize_right_eq1) -done - -lemma tpart_left1: - "[| a \ D n; tpart (a, b) (D, p) |] - ==> tpart(a, D n) (%x. if D x < D n then D(x) else D n, - %x. if D x < D n then p(x) else D n)" -apply (frule_tac r = n in tpart_partition [THEN partition_ub]) -apply (drule_tac x = "D n" in order_le_imp_less_or_eq) -apply (auto simp add: tpart_partition [THEN lemma_partition_eq, symmetric] tpart_tag_eq [symmetric]) -apply (frule_tac tpart_partition [THEN [3] lemma_additivity1]) -apply (auto simp add: tpart_def) -apply (drule_tac [2] linorder_not_less [THEN iffD1, THEN order_le_imp_less_or_eq], auto) - prefer 3 apply (drule_tac x=na in spec, arith) - prefer 2 apply (blast dest: lemma_additivity3) -apply (frule (2) lemma_additivity4_psize_eq) -apply (rule partition [THEN iffD2]) -apply (frule partition [THEN iffD1]) -apply safe -apply (auto simp add: partition_lt_gen) -apply (drule (1) partition_lt_cancel, arith) -done - -lemma fine_left1: - "[| a \ D n; tpart (a, b) (D, p); gauge (%x. a \ x & x \ D n) g; - fine (%x. if x < D n then min (g x) ((D n - x)/ 2) - else if x = D n then min (g (D n)) (ga (D n)) - else min (ga x) ((x - D n)/ 2)) (D, p) |] - ==> fine g - (%x. if D x < D n then D(x) else D n, - %x. if D x < D n then p(x) else D n)" -apply (auto simp add: fine_def tpart_def gauge_def) -apply (frule_tac [!] na=na in lemma_psize_left_less_psize2) -apply (drule_tac [!] x = na in spec, auto) -apply (drule_tac [!] x = na in spec, auto) -apply (auto dest: lemma_additivity3a simp add: split_if_asm) -done - -lemma tpart_right1: - "[| a \ D n; tpart (a, b) (D, p) |] - ==> tpart(D n, b) (%x. D(x + n),%x. p(x + n))" -apply (simp add: tpart_def partition_def, safe) -apply (rule_tac x = "N - n" in exI, auto) -done - -lemma fine_right1: - "[| a \ D n; tpart (a, b) (D, p); gauge (%x. D n \ x & x \ b) ga; - fine (%x. if x < D n then min (g x) ((D n - x)/ 2) - else if x = D n then min (g (D n)) (ga (D n)) - else min (ga x) ((x - D n)/ 2)) (D, p) |] - ==> fine ga (%x. D(x + n),%x. p(x + n))" -apply (auto simp add: fine_def gauge_def) -apply (drule_tac x = "na + n" in spec) -apply (frule_tac n = n in tpart_partition [THEN better_lemma_psize_right_eq], auto) -apply (simp add: tpart_def, safe) -apply (subgoal_tac "D n \ p (na + n)") -apply (drule_tac y = "p (na + n)" in order_le_imp_less_or_eq) -apply safe -apply (simp split: split_if_asm, simp) -apply (drule less_le_trans, assumption) -apply (rotate_tac 5) -apply (drule_tac x = "na + n" in spec, safe) -apply (rule_tac y="D (na + n)" in order_trans) -apply (case_tac "na = 0", auto) -apply (erule partition_lt_gen [THEN order_less_imp_le]) -apply arith -apply arith -done - -lemma rsum_add: "rsum (D, p) (%x. f x + g x) = rsum (D, p) f + rsum(D, p) g" -by (simp add: rsum_def setsum_addf left_distrib) - -text{* Bartle/Sherbert: Theorem 10.1.5 p. 278 *} -lemma Integral_add_fun: - "[| a \ b; Integral(a,b) f k1; Integral(a,b) g k2 |] - ==> Integral(a,b) (%x. f x + g x) (k1 + k2)" -apply (simp add: Integral_def, auto) -apply ((drule_tac x = "e/2" in spec)+) -apply auto -apply (drule gauge_min, assumption) -apply (rule_tac x = " (%x. if ga x < gaa x then ga x else gaa x)" in exI) -apply auto -apply (drule fine_min) -apply ((drule spec)+, auto) -apply (drule_tac a = "\rsum (D, p) f - k1\ * 2" and c = "\rsum (D, p) g - k2\ * 2" in add_strict_mono, assumption) -apply (auto simp only: rsum_add left_distrib [symmetric] - mult_2_right [symmetric] real_mult_less_iff1) -done - -lemma partition_lt_gen2: - "[| partition(a,b) D; r < psize D |] ==> 0 < D (Suc r) - D r" -by (auto simp add: partition) - -lemma lemma_Integral_le: - "[| \x. a \ x & x \ b --> f x \ g x; - tpart(a,b) (D,p) - |] ==> \n \ psize D. f (p n) \ g (p n)" -apply (simp add: tpart_def) -apply (auto, frule partition [THEN iffD1], auto) -apply (drule_tac x = "p n" in spec, auto) -apply (case_tac "n = 0", simp) -apply (rule partition_lt_gen [THEN order_less_le_trans, THEN order_less_imp_le], auto) -apply (drule le_imp_less_or_eq, auto) -apply (drule_tac [2] x = "psize D" in spec, auto) -apply (drule_tac r = "Suc n" in partition_ub) -apply (drule_tac x = n in spec, auto) -done - -lemma lemma_Integral_rsum_le: - "[| \x. a \ x & x \ b --> f x \ g x; - tpart(a,b) (D,p) - |] ==> rsum(D,p) f \ rsum(D,p) g" -apply (simp add: rsum_def) -apply (auto intro!: setsum_mono dest: tpart_partition [THEN partition_lt_gen2] - dest!: lemma_Integral_le) -done - -lemma Integral_le: - "[| a \ b; - \x. a \ x & x \ b --> f(x) \ g(x); - Integral(a,b) f k1; Integral(a,b) g k2 - |] ==> k1 \ k2" -apply (simp add: Integral_def) -apply (rotate_tac 2) -apply (drule_tac x = "\k1 - k2\ /2" in spec) -apply (drule_tac x = "\k1 - k2\ /2" in spec, auto) -apply (drule gauge_min, assumption) -apply (drule_tac g = "%x. if ga x < gaa x then ga x else gaa x" - in partition_exists, assumption, auto) -apply (drule fine_min) -apply (drule_tac x = D in spec, drule_tac x = D in spec) -apply (drule_tac x = p in spec, drule_tac x = p in spec, auto) -apply (frule lemma_Integral_rsum_le, assumption) -apply (subgoal_tac "\(rsum (D,p) f - k1) - (rsum (D,p) g - k2)\ < \k1 - k2\") -apply arith -apply (drule add_strict_mono, assumption) -apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] - real_mult_less_iff1) -done - -lemma Integral_imp_Cauchy: - "(\k. Integral(a,b) f k) ==> - (\e > 0. \g. gauge (%x. a \ x & x \ b) g & - (\D1 D2 p1 p2. - tpart(a,b) (D1, p1) & fine g (D1,p1) & - tpart(a,b) (D2, p2) & fine g (D2,p2) --> - \rsum(D1,p1) f - rsum(D2,p2) f\ < e))" -apply (simp add: Integral_def, auto) -apply (drule_tac x = "e/2" in spec, auto) -apply (rule exI, auto) -apply (frule_tac x = D1 in spec) -apply (frule_tac x = D2 in spec) -apply ((drule spec)+, auto) -apply (erule_tac V = "0 < e" in thin_rl) -apply (drule add_strict_mono, assumption) -apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] - real_mult_less_iff1) -done - -lemma Cauchy_iff2: - "Cauchy X = - (\j. (\M. \m \ M. \n \ M. \X m - X n\ < inverse(real (Suc j))))" -apply (simp add: Cauchy_def, auto) -apply (drule reals_Archimedean, safe) -apply (drule_tac x = n in spec, auto) -apply (rule_tac x = M in exI, auto) -apply (drule_tac x = m in spec, simp) -apply (drule_tac x = na in spec, auto) -done - -lemma partition_exists2: - "[| a \ b; \n. gauge (%x. a \ x & x \ b) (fa n) |] - ==> \n. \D p. tpart (a, b) (D, p) & fine (fa n) (D, p)" -by (blast dest: partition_exists) - -lemma monotonic_anti_derivative: - fixes f g :: "real => real" shows - "[| a \ b; \c. a \ c & c \ b --> f' c \ g' c; - \x. DERIV f x :> f' x; \x. DERIV g x :> g' x |] - ==> f b - f a \ g b - g a" -apply (rule Integral_le, assumption) -apply (auto intro: FTC1) -done - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/Lim.thy --- a/src/HOL/Hyperreal/Lim.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,673 +0,0 @@ -(* Title : Lim.thy - ID : $Id$ - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Conversion to Isar and new proofs by Lawrence C Paulson, 2004 -*) - -header{* Limits and Continuity *} - -theory Lim -imports SEQ -begin - -text{*Standard Definitions*} - -definition - LIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool" - ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where - [code del]: "f -- a --> L = - (\r > 0. \s > 0. \x. x \ a & norm (x - a) < s - --> norm (f x - L) < r)" - -definition - isCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where - "isCont f a = (f -- a --> (f a))" - -definition - isUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where - [code del]: "isUCont f = (\r>0. \s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r)" - - -subsection {* Limits of Functions *} - -subsubsection {* Purely standard proofs *} - -lemma LIM_eq: - "f -- a --> L = - (\r>0.\s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r)" -by (simp add: LIM_def diff_minus) - -lemma LIM_I: - "(!!r. 0 \s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r) - ==> f -- a --> L" -by (simp add: LIM_eq) - -lemma LIM_D: - "[| f -- a --> L; 0 \s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r" -by (simp add: LIM_eq) - -lemma LIM_offset: "f -- a --> L \ (\x. f (x + k)) -- a - k --> L" -apply (rule LIM_I) -apply (drule_tac r="r" in LIM_D, safe) -apply (rule_tac x="s" in exI, safe) -apply (drule_tac x="x + k" in spec) -apply (simp add: compare_rls) -done - -lemma LIM_offset_zero: "f -- a --> L \ (\h. f (a + h)) -- 0 --> L" -by (drule_tac k="a" in LIM_offset, simp add: add_commute) - -lemma LIM_offset_zero_cancel: "(\h. f (a + h)) -- 0 --> L \ f -- a --> L" -by (drule_tac k="- a" in LIM_offset, simp) - -lemma LIM_const [simp]: "(%x. k) -- x --> k" -by (simp add: LIM_def) - -lemma LIM_add: - fixes f g :: "'a::real_normed_vector \ 'b::real_normed_vector" - assumes f: "f -- a --> L" and g: "g -- a --> M" - shows "(%x. f x + g(x)) -- a --> (L + M)" -proof (rule LIM_I) - fix r :: real - assume r: "0 < r" - from LIM_D [OF f half_gt_zero [OF r]] - obtain fs - where fs: "0 < fs" - and fs_lt: "\x. x \ a & norm (x-a) < fs --> norm (f x - L) < r/2" - by blast - from LIM_D [OF g half_gt_zero [OF r]] - obtain gs - where gs: "0 < gs" - and gs_lt: "\x. x \ a & norm (x-a) < gs --> norm (g x - M) < r/2" - by blast - show "\s>0.\x. x \ a \ norm (x-a) < s \ norm (f x + g x - (L + M)) < r" - proof (intro exI conjI strip) - show "0 < min fs gs" by (simp add: fs gs) - fix x :: 'a - assume "x \ a \ norm (x-a) < min fs gs" - hence "x \ a \ norm (x-a) < fs \ norm (x-a) < gs" by simp - with fs_lt gs_lt - have "norm (f x - L) < r/2" and "norm (g x - M) < r/2" by blast+ - hence "norm (f x - L) + norm (g x - M) < r" by arith - thus "norm (f x + g x - (L + M)) < r" - by (blast intro: norm_diff_triangle_ineq order_le_less_trans) - qed -qed - -lemma LIM_add_zero: - "\f -- a --> 0; g -- a --> 0\ \ (\x. f x + g x) -- a --> 0" -by (drule (1) LIM_add, simp) - -lemma minus_diff_minus: - fixes a b :: "'a::ab_group_add" - shows "(- a) - (- b) = - (a - b)" -by simp - -lemma LIM_minus: "f -- a --> L ==> (%x. -f(x)) -- a --> -L" -by (simp only: LIM_eq minus_diff_minus norm_minus_cancel) - -lemma LIM_add_minus: - "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + -g(x)) -- x --> (l + -m)" -by (intro LIM_add LIM_minus) - -lemma LIM_diff: - "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) - g(x)) -- x --> l-m" -by (simp only: diff_minus LIM_add LIM_minus) - -lemma LIM_zero: "f -- a --> l \ (\x. f x - l) -- a --> 0" -by (simp add: LIM_def) - -lemma LIM_zero_cancel: "(\x. f x - l) -- a --> 0 \ f -- a --> l" -by (simp add: LIM_def) - -lemma LIM_zero_iff: "(\x. f x - l) -- a --> 0 = f -- a --> l" -by (simp add: LIM_def) - -lemma LIM_imp_LIM: - assumes f: "f -- a --> l" - assumes le: "\x. x \ a \ norm (g x - m) \ norm (f x - l)" - shows "g -- a --> m" -apply (rule LIM_I, drule LIM_D [OF f], safe) -apply (rule_tac x="s" in exI, safe) -apply (drule_tac x="x" in spec, safe) -apply (erule (1) order_le_less_trans [OF le]) -done - -lemma LIM_norm: "f -- a --> l \ (\x. norm (f x)) -- a --> norm l" -by (erule LIM_imp_LIM, simp add: norm_triangle_ineq3) - -lemma LIM_norm_zero: "f -- a --> 0 \ (\x. norm (f x)) -- a --> 0" -by (drule LIM_norm, simp) - -lemma LIM_norm_zero_cancel: "(\x. norm (f x)) -- a --> 0 \ f -- a --> 0" -by (erule LIM_imp_LIM, simp) - -lemma LIM_norm_zero_iff: "(\x. norm (f x)) -- a --> 0 = f -- a --> 0" -by (rule iffI [OF LIM_norm_zero_cancel LIM_norm_zero]) - -lemma LIM_rabs: "f -- a --> (l::real) \ (\x. \f x\) -- a --> \l\" -by (fold real_norm_def, rule LIM_norm) - -lemma LIM_rabs_zero: "f -- a --> (0::real) \ (\x. \f x\) -- a --> 0" -by (fold real_norm_def, rule LIM_norm_zero) - -lemma LIM_rabs_zero_cancel: "(\x. \f x\) -- a --> (0::real) \ f -- a --> 0" -by (fold real_norm_def, rule LIM_norm_zero_cancel) - -lemma LIM_rabs_zero_iff: "(\x. \f x\) -- a --> (0::real) = f -- a --> 0" -by (fold real_norm_def, rule LIM_norm_zero_iff) - -lemma LIM_const_not_eq: - fixes a :: "'a::real_normed_algebra_1" - shows "k \ L \ \ (\x. k) -- a --> L" -apply (simp add: LIM_eq) -apply (rule_tac x="norm (k - L)" in exI, simp, safe) -apply (rule_tac x="a + of_real (s/2)" in exI, simp add: norm_of_real) -done - -lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] - -lemma LIM_const_eq: - fixes a :: "'a::real_normed_algebra_1" - shows "(\x. k) -- a --> L \ k = L" -apply (rule ccontr) -apply (blast dest: LIM_const_not_eq) -done - -lemma LIM_unique: - fixes a :: "'a::real_normed_algebra_1" - shows "\f -- a --> L; f -- a --> M\ \ L = M" -apply (drule (1) LIM_diff) -apply (auto dest!: LIM_const_eq) -done - -lemma LIM_ident [simp]: "(\x. x) -- a --> a" -by (auto simp add: LIM_def) - -text{*Limits are equal for functions equal except at limit point*} -lemma LIM_equal: - "[| \x. x \ a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)" -by (simp add: LIM_def) - -lemma LIM_cong: - "\a = b; \x. x \ b \ f x = g x; l = m\ - \ ((\x. f x) -- a --> l) = ((\x. g x) -- b --> m)" -by (simp add: LIM_def) - -lemma LIM_equal2: - assumes 1: "0 < R" - assumes 2: "\x. \x \ a; norm (x - a) < R\ \ f x = g x" - shows "g -- a --> l \ f -- a --> l" -apply (unfold LIM_def, safe) -apply (drule_tac x="r" in spec, safe) -apply (rule_tac x="min s R" in exI, safe) -apply (simp add: 1) -apply (simp add: 2) -done - -text{*Two uses in Hyperreal/Transcendental.ML*} -lemma LIM_trans: - "[| (%x. f(x) + -g(x)) -- a --> 0; g -- a --> l |] ==> f -- a --> l" -apply (drule LIM_add, assumption) -apply (auto simp add: add_assoc) -done - -lemma LIM_compose: - assumes g: "g -- l --> g l" - assumes f: "f -- a --> l" - shows "(\x. g (f x)) -- a --> g l" -proof (rule LIM_I) - fix r::real assume r: "0 < r" - obtain s where s: "0 < s" - and less_r: "\y. \y \ l; norm (y - l) < s\ \ norm (g y - g l) < r" - using LIM_D [OF g r] by fast - obtain t where t: "0 < t" - and less_s: "\x. \x \ a; norm (x - a) < t\ \ norm (f x - l) < s" - using LIM_D [OF f s] by fast - - show "\t>0. \x. x \ a \ norm (x - a) < t \ norm (g (f x) - g l) < r" - proof (rule exI, safe) - show "0 < t" using t . - next - fix x assume "x \ a" and "norm (x - a) < t" - hence "norm (f x - l) < s" by (rule less_s) - thus "norm (g (f x) - g l) < r" - using r less_r by (case_tac "f x = l", simp_all) - qed -qed - -lemma LIM_compose2: - assumes f: "f -- a --> b" - assumes g: "g -- b --> c" - assumes inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ b" - shows "(\x. g (f x)) -- a --> c" -proof (rule LIM_I) - fix r :: real - assume r: "0 < r" - obtain s where s: "0 < s" - and less_r: "\y. \y \ b; norm (y - b) < s\ \ norm (g y - c) < r" - using LIM_D [OF g r] by fast - obtain t where t: "0 < t" - and less_s: "\x. \x \ a; norm (x - a) < t\ \ norm (f x - b) < s" - using LIM_D [OF f s] by fast - obtain d where d: "0 < d" - and neq_b: "\x. \x \ a; norm (x - a) < d\ \ f x \ b" - using inj by fast - - show "\t>0. \x. x \ a \ norm (x - a) < t \ norm (g (f x) - c) < r" - proof (safe intro!: exI) - show "0 < min d t" using d t by simp - next - fix x - assume "x \ a" and "norm (x - a) < min d t" - hence "f x \ b" and "norm (f x - b) < s" - using neq_b less_s by simp_all - thus "norm (g (f x) - c) < r" - by (rule less_r) - qed -qed - -lemma LIM_o: "\g -- l --> g l; f -- a --> l\ \ (g \ f) -- a --> g l" -unfolding o_def by (rule LIM_compose) - -lemma real_LIM_sandwich_zero: - fixes f g :: "'a::real_normed_vector \ real" - assumes f: "f -- a --> 0" - assumes 1: "\x. x \ a \ 0 \ g x" - assumes 2: "\x. x \ a \ g x \ f x" - shows "g -- a --> 0" -proof (rule LIM_imp_LIM [OF f]) - fix x assume x: "x \ a" - have "norm (g x - 0) = g x" by (simp add: 1 x) - also have "g x \ f x" by (rule 2 [OF x]) - also have "f x \ \f x\" by (rule abs_ge_self) - also have "\f x\ = norm (f x - 0)" by simp - finally show "norm (g x - 0) \ norm (f x - 0)" . -qed - -text {* Bounded Linear Operators *} - -lemma (in bounded_linear) cont: "f -- a --> f a" -proof (rule LIM_I) - fix r::real assume r: "0 < r" - obtain K where K: "0 < K" and norm_le: "\x. norm (f x) \ norm x * K" - using pos_bounded by fast - show "\s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - f a) < r" - proof (rule exI, safe) - from r K show "0 < r / K" by (rule divide_pos_pos) - next - fix x assume x: "norm (x - a) < r / K" - have "norm (f x - f a) = norm (f (x - a))" by (simp only: diff) - also have "\ \ norm (x - a) * K" by (rule norm_le) - also from K x have "\ < r" by (simp only: pos_less_divide_eq) - finally show "norm (f x - f a) < r" . - qed -qed - -lemma (in bounded_linear) LIM: - "g -- a --> l \ (\x. f (g x)) -- a --> f l" -by (rule LIM_compose [OF cont]) - -lemma (in bounded_linear) LIM_zero: - "g -- a --> 0 \ (\x. f (g x)) -- a --> 0" -by (drule LIM, simp only: zero) - -text {* Bounded Bilinear Operators *} - -lemma (in bounded_bilinear) LIM_prod_zero: - assumes f: "f -- a --> 0" - assumes g: "g -- a --> 0" - shows "(\x. f x ** g x) -- a --> 0" -proof (rule LIM_I) - fix r::real assume r: "0 < r" - obtain K where K: "0 < K" - and norm_le: "\x y. norm (x ** y) \ norm x * norm y * K" - using pos_bounded by fast - from K have K': "0 < inverse K" - by (rule positive_imp_inverse_positive) - obtain s where s: "0 < s" - and norm_f: "\x. \x \ a; norm (x - a) < s\ \ norm (f x) < r" - using LIM_D [OF f r] by auto - obtain t where t: "0 < t" - and norm_g: "\x. \x \ a; norm (x - a) < t\ \ norm (g x) < inverse K" - using LIM_D [OF g K'] by auto - show "\s>0. \x. x \ a \ norm (x - a) < s \ norm (f x ** g x - 0) < r" - proof (rule exI, safe) - from s t show "0 < min s t" by simp - next - fix x assume x: "x \ a" - assume "norm (x - a) < min s t" - hence xs: "norm (x - a) < s" and xt: "norm (x - a) < t" by simp_all - from x xs have 1: "norm (f x) < r" by (rule norm_f) - from x xt have 2: "norm (g x) < inverse K" by (rule norm_g) - have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) - also from 1 2 K have "\ < r * inverse K * K" - by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero) - also from K have "r * inverse K * K = r" by simp - finally show "norm (f x ** g x - 0) < r" by simp - qed -qed - -lemma (in bounded_bilinear) LIM_left_zero: - "f -- a --> 0 \ (\x. f x ** c) -- a --> 0" -by (rule bounded_linear.LIM_zero [OF bounded_linear_left]) - -lemma (in bounded_bilinear) LIM_right_zero: - "f -- a --> 0 \ (\x. c ** f x) -- a --> 0" -by (rule bounded_linear.LIM_zero [OF bounded_linear_right]) - -lemma (in bounded_bilinear) LIM: - "\f -- a --> L; g -- a --> M\ \ (\x. f x ** g x) -- a --> L ** M" -apply (drule LIM_zero) -apply (drule LIM_zero) -apply (rule LIM_zero_cancel) -apply (subst prod_diff_prod) -apply (rule LIM_add_zero) -apply (rule LIM_add_zero) -apply (erule (1) LIM_prod_zero) -apply (erule LIM_left_zero) -apply (erule LIM_right_zero) -done - -lemmas LIM_mult = mult.LIM - -lemmas LIM_mult_zero = mult.LIM_prod_zero - -lemmas LIM_mult_left_zero = mult.LIM_left_zero - -lemmas LIM_mult_right_zero = mult.LIM_right_zero - -lemmas LIM_scaleR = scaleR.LIM - -lemmas LIM_of_real = of_real.LIM - -lemma LIM_power: - fixes f :: "'a::real_normed_vector \ 'b::{recpower,real_normed_algebra}" - assumes f: "f -- a --> l" - shows "(\x. f x ^ n) -- a --> l ^ n" -by (induct n, simp, simp add: power_Suc LIM_mult f) - -subsubsection {* Derived theorems about @{term LIM} *} - -lemma LIM_inverse_lemma: - fixes x :: "'a::real_normed_div_algebra" - assumes r: "0 < r" - assumes x: "norm (x - 1) < min (1/2) (r/2)" - shows "norm (inverse x - 1) < r" -proof - - from r have r2: "0 < r/2" by simp - from x have 0: "x \ 0" by clarsimp - from x have x': "norm (1 - x) < min (1/2) (r/2)" - by (simp only: norm_minus_commute) - hence less1: "norm (1 - x) < r/2" by simp - have "norm (1::'a) - norm x \ norm (1 - x)" by (rule norm_triangle_ineq2) - also from x' have "norm (1 - x) < 1/2" by simp - finally have "1/2 < norm x" by simp - hence "inverse (norm x) < inverse (1/2)" - by (rule less_imp_inverse_less, simp) - hence less2: "norm (inverse x) < 2" - by (simp add: nonzero_norm_inverse 0) - from less1 less2 r2 norm_ge_zero - have "norm (1 - x) * norm (inverse x) < (r/2) * 2" - by (rule mult_strict_mono) - thus "norm (inverse x - 1) < r" - by (simp only: norm_mult [symmetric] left_diff_distrib, simp add: 0) -qed - -lemma LIM_inverse_fun: - assumes a: "a \ (0::'a::real_normed_div_algebra)" - shows "inverse -- a --> inverse a" -proof (rule LIM_equal2) - from a show "0 < norm a" by simp -next - fix x assume "norm (x - a) < norm a" - hence "x \ 0" by auto - with a show "inverse x = inverse (inverse a * x) * inverse a" - by (simp add: nonzero_inverse_mult_distrib - nonzero_imp_inverse_nonzero - nonzero_inverse_inverse_eq mult_assoc) -next - have 1: "inverse -- 1 --> inverse (1::'a)" - apply (rule LIM_I) - apply (rule_tac x="min (1/2) (r/2)" in exI) - apply (simp add: LIM_inverse_lemma) - done - have "(\x. inverse a * x) -- a --> inverse a * a" - by (intro LIM_mult LIM_ident LIM_const) - hence "(\x. inverse a * x) -- a --> 1" - by (simp add: a) - with 1 have "(\x. inverse (inverse a * x)) -- a --> inverse 1" - by (rule LIM_compose) - hence "(\x. inverse (inverse a * x)) -- a --> 1" - by simp - hence "(\x. inverse (inverse a * x) * inverse a) -- a --> 1 * inverse a" - by (intro LIM_mult LIM_const) - thus "(\x. inverse (inverse a * x) * inverse a) -- a --> inverse a" - by simp -qed - -lemma LIM_inverse: - fixes L :: "'a::real_normed_div_algebra" - shows "\f -- a --> L; L \ 0\ \ (\x. inverse (f x)) -- a --> inverse L" -by (rule LIM_inverse_fun [THEN LIM_compose]) - - -subsection {* Continuity *} - -subsubsection {* Purely standard proofs *} - -lemma LIM_isCont_iff: "(f -- a --> f a) = ((\h. f (a + h)) -- 0 --> f a)" -by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) - -lemma isCont_iff: "isCont f x = (\h. f (x + h)) -- 0 --> f x" -by (simp add: isCont_def LIM_isCont_iff) - -lemma isCont_ident [simp]: "isCont (\x. x) a" - unfolding isCont_def by (rule LIM_ident) - -lemma isCont_const [simp]: "isCont (\x. k) a" - unfolding isCont_def by (rule LIM_const) - -lemma isCont_norm: "isCont f a \ isCont (\x. norm (f x)) a" - unfolding isCont_def by (rule LIM_norm) - -lemma isCont_rabs: "isCont f a \ isCont (\x. \f x :: real\) a" - unfolding isCont_def by (rule LIM_rabs) - -lemma isCont_add: "\isCont f a; isCont g a\ \ isCont (\x. f x + g x) a" - unfolding isCont_def by (rule LIM_add) - -lemma isCont_minus: "isCont f a \ isCont (\x. - f x) a" - unfolding isCont_def by (rule LIM_minus) - -lemma isCont_diff: "\isCont f a; isCont g a\ \ isCont (\x. f x - g x) a" - unfolding isCont_def by (rule LIM_diff) - -lemma isCont_mult: - fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" - shows "\isCont f a; isCont g a\ \ isCont (\x. f x * g x) a" - unfolding isCont_def by (rule LIM_mult) - -lemma isCont_inverse: - fixes f :: "'a::real_normed_vector \ 'b::real_normed_div_algebra" - shows "\isCont f a; f a \ 0\ \ isCont (\x. inverse (f x)) a" - unfolding isCont_def by (rule LIM_inverse) - -lemma isCont_LIM_compose: - "\isCont g l; f -- a --> l\ \ (\x. g (f x)) -- a --> g l" - unfolding isCont_def by (rule LIM_compose) - -lemma isCont_LIM_compose2: - assumes f [unfolded isCont_def]: "isCont f a" - assumes g: "g -- f a --> l" - assumes inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ f a" - shows "(\x. g (f x)) -- a --> l" -by (rule LIM_compose2 [OF f g inj]) - -lemma isCont_o2: "\isCont f a; isCont g (f a)\ \ isCont (\x. g (f x)) a" - unfolding isCont_def by (rule LIM_compose) - -lemma isCont_o: "\isCont f a; isCont g (f a)\ \ isCont (g o f) a" - unfolding o_def by (rule isCont_o2) - -lemma (in bounded_linear) isCont: "isCont f a" - unfolding isCont_def by (rule cont) - -lemma (in bounded_bilinear) isCont: - "\isCont f a; isCont g a\ \ isCont (\x. f x ** g x) a" - unfolding isCont_def by (rule LIM) - -lemmas isCont_scaleR = scaleR.isCont - -lemma isCont_of_real: - "isCont f a \ isCont (\x. of_real (f x)) a" - unfolding isCont_def by (rule LIM_of_real) - -lemma isCont_power: - fixes f :: "'a::real_normed_vector \ 'b::{recpower,real_normed_algebra}" - shows "isCont f a \ isCont (\x. f x ^ n) a" - unfolding isCont_def by (rule LIM_power) - -lemma isCont_abs [simp]: "isCont abs (a::real)" -by (rule isCont_rabs [OF isCont_ident]) - - -subsection {* Uniform Continuity *} - -lemma isUCont_isCont: "isUCont f ==> isCont f x" -by (simp add: isUCont_def isCont_def LIM_def, force) - -lemma isUCont_Cauchy: - "\isUCont f; Cauchy X\ \ Cauchy (\n. f (X n))" -unfolding isUCont_def -apply (rule CauchyI) -apply (drule_tac x=e in spec, safe) -apply (drule_tac e=s in CauchyD, safe) -apply (rule_tac x=M in exI, simp) -done - -lemma (in bounded_linear) isUCont: "isUCont f" -unfolding isUCont_def -proof (intro allI impI) - fix r::real assume r: "0 < r" - obtain K where K: "0 < K" and norm_le: "\x. norm (f x) \ norm x * K" - using pos_bounded by fast - show "\s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r" - proof (rule exI, safe) - from r K show "0 < r / K" by (rule divide_pos_pos) - next - fix x y :: 'a - assume xy: "norm (x - y) < r / K" - have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff) - also have "\ \ norm (x - y) * K" by (rule norm_le) - also from K xy have "\ < r" by (simp only: pos_less_divide_eq) - finally show "norm (f x - f y) < r" . - qed -qed - -lemma (in bounded_linear) Cauchy: "Cauchy X \ Cauchy (\n. f (X n))" -by (rule isUCont [THEN isUCont_Cauchy]) - - -subsection {* Relation of LIM and LIMSEQ *} - -lemma LIMSEQ_SEQ_conv1: - fixes a :: "'a::real_normed_vector" - assumes X: "X -- a --> L" - shows "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" -proof (safe intro!: LIMSEQ_I) - fix S :: "nat \ 'a" - fix r :: real - assume rgz: "0 < r" - assume as: "\n. S n \ a" - assume S: "S ----> a" - from LIM_D [OF X rgz] obtain s - where sgz: "0 < s" - and aux: "\x. \x \ a; norm (x - a) < s\ \ norm (X x - L) < r" - by fast - from LIMSEQ_D [OF S sgz] - obtain no where "\n\no. norm (S n - a) < s" by blast - hence "\n\no. norm (X (S n) - L) < r" by (simp add: aux as) - thus "\no. \n\no. norm (X (S n) - L) < r" .. -qed - -lemma LIMSEQ_SEQ_conv2: - fixes a :: real - assumes "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" - shows "X -- a --> L" -proof (rule ccontr) - assume "\ (X -- a --> L)" - hence "\ (\r > 0. \s > 0. \x. x \ a & norm (x - a) < s --> norm (X x - L) < r)" by (unfold LIM_def) - hence "\r > 0. \s > 0. \x. \(x \ a \ \x - a\ < s --> norm (X x - L) < r)" by simp - hence "\r > 0. \s > 0. \x. (x \ a \ \x - a\ < s \ norm (X x - L) \ r)" by (simp add: linorder_not_less) - then obtain r where rdef: "r > 0 \ (\s > 0. \x. (x \ a \ \x - a\ < s \ norm (X x - L) \ r))" by auto - - let ?F = "\n::nat. SOME x. x\a \ \x - a\ < inverse (real (Suc n)) \ norm (X x - L) \ r" - have "\n. \x. x\a \ \x - a\ < inverse (real (Suc n)) \ norm (X x - L) \ r" - using rdef by simp - hence F: "\n. ?F n \ a \ \?F n - a\ < inverse (real (Suc n)) \ norm (X (?F n) - L) \ r" - by (rule someI_ex) - hence F1: "\n. ?F n \ a" - and F2: "\n. \?F n - a\ < inverse (real (Suc n))" - and F3: "\n. norm (X (?F n) - L) \ r" - by fast+ - - have "?F ----> a" - proof (rule LIMSEQ_I, unfold real_norm_def) - fix e::real - assume "0 < e" - (* choose no such that inverse (real (Suc n)) < e *) - then have "\no. inverse (real (Suc no)) < e" by (rule reals_Archimedean) - then obtain m where nodef: "inverse (real (Suc m)) < e" by auto - show "\no. \n. no \ n --> \?F n - a\ < e" - proof (intro exI allI impI) - fix n - assume mlen: "m \ n" - have "\?F n - a\ < inverse (real (Suc n))" - by (rule F2) - also have "inverse (real (Suc n)) \ inverse (real (Suc m))" - using mlen by auto - also from nodef have - "inverse (real (Suc m)) < e" . - finally show "\?F n - a\ < e" . - qed - qed - - moreover have "\n. ?F n \ a" - by (rule allI) (rule F1) - - moreover from prems have "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" by simp - ultimately have "(\n. X (?F n)) ----> L" by simp - - moreover have "\ ((\n. X (?F n)) ----> L)" - proof - - { - fix no::nat - obtain n where "n = no + 1" by simp - then have nolen: "no \ n" by simp - (* We prove this by showing that for any m there is an n\m such that |X (?F n) - L| \ r *) - have "norm (X (?F n) - L) \ r" - by (rule F3) - with nolen have "\n. no \ n \ norm (X (?F n) - L) \ r" by fast - } - then have "(\no. \n. no \ n \ norm (X (?F n) - L) \ r)" by simp - with rdef have "\e>0. (\no. \n. no \ n \ norm (X (?F n) - L) \ e)" by auto - thus ?thesis by (unfold LIMSEQ_def, auto simp add: linorder_not_less) - qed - ultimately show False by simp -qed - -lemma LIMSEQ_SEQ_conv: - "(\S. (\n. S n \ a) \ S ----> (a::real) \ (\n. X (S n)) ----> L) = - (X -- a --> L)" -proof - assume "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" - thus "X -- a --> L" by (rule LIMSEQ_SEQ_conv2) -next - assume "(X -- a --> L)" - thus "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" by (rule LIMSEQ_SEQ_conv1) -qed - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/Ln.thy --- a/src/HOL/Hyperreal/Ln.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,431 +0,0 @@ -(* Title: Ln.thy - Author: Jeremy Avigad - ID: $Id$ -*) - -header {* Properties of ln *} - -theory Ln -imports Transcendental -begin - -lemma exp_first_two_terms: "exp x = 1 + x + suminf (%n. - inverse(real (fact (n+2))) * (x ^ (n+2)))" -proof - - have "exp x = suminf (%n. inverse(real (fact n)) * (x ^ n))" - by (simp add: exp_def) - also from summable_exp have "... = (SUM n : {0..<2}. - inverse(real (fact n)) * (x ^ n)) + suminf (%n. - inverse(real (fact (n+2))) * (x ^ (n+2)))" (is "_ = ?a + _") - by (rule suminf_split_initial_segment) - also have "?a = 1 + x" - by (simp add: numerals) - finally show ?thesis . -qed - -lemma exp_tail_after_first_two_terms_summable: - "summable (%n. inverse(real (fact (n+2))) * (x ^ (n+2)))" -proof - - note summable_exp - thus ?thesis - by (frule summable_ignore_initial_segment) -qed - -lemma aux1: assumes a: "0 <= x" and b: "x <= 1" - shows "inverse (real (fact (n + 2))) * x ^ (n + 2) <= (x^2/2) * ((1/2)^n)" -proof (induct n) - show "inverse (real (fact (0 + 2))) * x ^ (0 + 2) <= - x ^ 2 / 2 * (1 / 2) ^ 0" - by (simp add: real_of_nat_Suc power2_eq_square) -next - fix n - assume c: "inverse (real (fact (n + 2))) * x ^ (n + 2) - <= x ^ 2 / 2 * (1 / 2) ^ n" - show "inverse (real (fact (Suc n + 2))) * x ^ (Suc n + 2) - <= x ^ 2 / 2 * (1 / 2) ^ Suc n" - proof - - have "inverse(real (fact (Suc n + 2))) <= - (1 / 2) *inverse (real (fact (n+2)))" - proof - - have "Suc n + 2 = Suc (n + 2)" by simp - then have "fact (Suc n + 2) = Suc (n + 2) * fact (n + 2)" - by simp - then have "real(fact (Suc n + 2)) = real(Suc (n + 2) * fact (n + 2))" - apply (rule subst) - apply (rule refl) - done - also have "... = real(Suc (n + 2)) * real(fact (n + 2))" - by (rule real_of_nat_mult) - finally have "real (fact (Suc n + 2)) = - real (Suc (n + 2)) * real (fact (n + 2))" . - then have "inverse(real (fact (Suc n + 2))) = - inverse(real (Suc (n + 2))) * inverse(real (fact (n + 2)))" - apply (rule ssubst) - apply (rule inverse_mult_distrib) - done - also have "... <= (1/2) * inverse(real (fact (n + 2)))" - apply (rule mult_right_mono) - apply (subst inverse_eq_divide) - apply simp - apply (rule inv_real_of_nat_fact_ge_zero) - done - finally show ?thesis . - qed - moreover have "x ^ (Suc n + 2) <= x ^ (n + 2)" - apply (simp add: mult_compare_simps) - apply (simp add: prems) - apply (subgoal_tac "0 <= x * (x * x^n)") - apply force - apply (rule mult_nonneg_nonneg, rule a)+ - apply (rule zero_le_power, rule a) - done - ultimately have "inverse (real (fact (Suc n + 2))) * x ^ (Suc n + 2) <= - (1 / 2 * inverse (real (fact (n + 2)))) * x ^ (n + 2)" - apply (rule mult_mono) - apply (rule mult_nonneg_nonneg) - apply simp - apply (subst inverse_nonnegative_iff_nonnegative) - apply (rule real_of_nat_ge_zero) - apply (rule zero_le_power) - apply (rule a) - done - also have "... = 1 / 2 * (inverse (real (fact (n + 2))) * x ^ (n + 2))" - by simp - also have "... <= 1 / 2 * (x ^ 2 / 2 * (1 / 2) ^ n)" - apply (rule mult_left_mono) - apply (rule prems) - apply simp - done - also have "... = x ^ 2 / 2 * (1 / 2 * (1 / 2) ^ n)" - by auto - also have "(1::real) / 2 * (1 / 2) ^ n = (1 / 2) ^ (Suc n)" - by (rule realpow_Suc [THEN sym]) - finally show ?thesis . - qed -qed - -lemma aux2: "(%n. (x::real) ^ 2 / 2 * (1 / 2) ^ n) sums x^2" -proof - - have "(%n. (1 / 2::real)^n) sums (1 / (1 - (1/2)))" - apply (rule geometric_sums) - by (simp add: abs_less_iff) - also have "(1::real) / (1 - 1/2) = 2" - by simp - finally have "(%n. (1 / 2::real)^n) sums 2" . - then have "(%n. x ^ 2 / 2 * (1 / 2) ^ n) sums (x^2 / 2 * 2)" - by (rule sums_mult) - also have "x^2 / 2 * 2 = x^2" - by simp - finally show ?thesis . -qed - -lemma exp_bound: "0 <= (x::real) ==> x <= 1 ==> exp x <= 1 + x + x^2" -proof - - assume a: "0 <= x" - assume b: "x <= 1" - have c: "exp x = 1 + x + suminf (%n. inverse(real (fact (n+2))) * - (x ^ (n+2)))" - by (rule exp_first_two_terms) - moreover have "suminf (%n. inverse(real (fact (n+2))) * (x ^ (n+2))) <= x^2" - proof - - have "suminf (%n. inverse(real (fact (n+2))) * (x ^ (n+2))) <= - suminf (%n. (x^2/2) * ((1/2)^n))" - apply (rule summable_le) - apply (auto simp only: aux1 prems) - apply (rule exp_tail_after_first_two_terms_summable) - by (rule sums_summable, rule aux2) - also have "... = x^2" - by (rule sums_unique [THEN sym], rule aux2) - finally show ?thesis . - qed - ultimately show ?thesis - by auto -qed - -lemma aux4: "0 <= (x::real) ==> x <= 1 ==> exp (x - x^2) <= 1 + x" -proof - - assume a: "0 <= x" and b: "x <= 1" - have "exp (x - x^2) = exp x / exp (x^2)" - by (rule exp_diff) - also have "... <= (1 + x + x^2) / exp (x ^2)" - apply (rule divide_right_mono) - apply (rule exp_bound) - apply (rule a, rule b) - apply simp - done - also have "... <= (1 + x + x^2) / (1 + x^2)" - apply (rule divide_left_mono) - apply (auto simp add: exp_ge_add_one_self_aux) - apply (rule add_nonneg_nonneg) - apply (insert prems, auto) - apply (rule mult_pos_pos) - apply auto - apply (rule add_pos_nonneg) - apply auto - done - also from a have "... <= 1 + x" - by(simp add:field_simps zero_compare_simps) - finally show ?thesis . -qed - -lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==> - x - x^2 <= ln (1 + x)" -proof - - assume a: "0 <= x" and b: "x <= 1" - then have "exp (x - x^2) <= 1 + x" - by (rule aux4) - also have "... = exp (ln (1 + x))" - proof - - from a have "0 < 1 + x" by auto - thus ?thesis - by (auto simp only: exp_ln_iff [THEN sym]) - qed - finally have "exp (x - x ^ 2) <= exp (ln (1 + x))" . - thus ?thesis by (auto simp only: exp_le_cancel_iff) -qed - -lemma ln_one_minus_pos_upper_bound: "0 <= x ==> x < 1 ==> ln (1 - x) <= - x" -proof - - assume a: "0 <= (x::real)" and b: "x < 1" - have "(1 - x) * (1 + x + x^2) = (1 - x^3)" - by (simp add: ring_simps power2_eq_square power3_eq_cube) - also have "... <= 1" - by (auto simp add: a) - finally have "(1 - x) * (1 + x + x ^ 2) <= 1" . - moreover have "0 < 1 + x + x^2" - apply (rule add_pos_nonneg) - apply (insert a, auto) - done - ultimately have "1 - x <= 1 / (1 + x + x^2)" - by (elim mult_imp_le_div_pos) - also have "... <= 1 / exp x" - apply (rule divide_left_mono) - apply (rule exp_bound, rule a) - apply (insert prems, auto) - apply (rule mult_pos_pos) - apply (rule add_pos_nonneg) - apply auto - done - also have "... = exp (-x)" - by (auto simp add: exp_minus real_divide_def) - finally have "1 - x <= exp (- x)" . - also have "1 - x = exp (ln (1 - x))" - proof - - have "0 < 1 - x" - by (insert b, auto) - thus ?thesis - by (auto simp only: exp_ln_iff [THEN sym]) - qed - finally have "exp (ln (1 - x)) <= exp (- x)" . - thus ?thesis by (auto simp only: exp_le_cancel_iff) -qed - -lemma aux5: "x < 1 ==> ln(1 - x) = - ln(1 + x / (1 - x))" -proof - - assume a: "x < 1" - have "ln(1 - x) = - ln(1 / (1 - x))" - proof - - have "ln(1 - x) = - (- ln (1 - x))" - by auto - also have "- ln(1 - x) = ln 1 - ln(1 - x)" - by simp - also have "... = ln(1 / (1 - x))" - apply (rule ln_div [THEN sym]) - by (insert a, auto) - finally show ?thesis . - qed - also have " 1 / (1 - x) = 1 + x / (1 - x)" using a by(simp add:field_simps) - finally show ?thesis . -qed - -lemma ln_one_minus_pos_lower_bound: "0 <= x ==> x <= (1 / 2) ==> - - x - 2 * x^2 <= ln (1 - x)" -proof - - assume a: "0 <= x" and b: "x <= (1 / 2)" - from b have c: "x < 1" - by auto - then have "ln (1 - x) = - ln (1 + x / (1 - x))" - by (rule aux5) - also have "- (x / (1 - x)) <= ..." - proof - - have "ln (1 + x / (1 - x)) <= x / (1 - x)" - apply (rule ln_add_one_self_le_self) - apply (rule divide_nonneg_pos) - by (insert a c, auto) - thus ?thesis - by auto - qed - also have "- (x / (1 - x)) = -x / (1 - x)" - by auto - finally have d: "- x / (1 - x) <= ln (1 - x)" . - have "0 < 1 - x" using prems by simp - hence e: "-x - 2 * x^2 <= - x / (1 - x)" - using mult_right_le_one_le[of "x*x" "2*x"] prems - by(simp add:field_simps power2_eq_square) - from e d show "- x - 2 * x^2 <= ln (1 - x)" - by (rule order_trans) -qed - -lemma exp_ge_add_one_self [simp]: "1 + (x::real) <= exp x" - apply (case_tac "0 <= x") - apply (erule exp_ge_add_one_self_aux) - apply (case_tac "x <= -1") - apply (subgoal_tac "1 + x <= 0") - apply (erule order_trans) - apply simp - apply simp - apply (subgoal_tac "1 + x = exp(ln (1 + x))") - apply (erule ssubst) - apply (subst exp_le_cancel_iff) - apply (subgoal_tac "ln (1 - (- x)) <= - (- x)") - apply simp - apply (rule ln_one_minus_pos_upper_bound) - apply auto -done - -lemma ln_add_one_self_le_self2: "-1 < x ==> ln(1 + x) <= x" - apply (subgoal_tac "x = ln (exp x)") - apply (erule ssubst)back - apply (subst ln_le_cancel_iff) - apply auto -done - -lemma abs_ln_one_plus_x_minus_x_bound_nonneg: - "0 <= x ==> x <= 1 ==> abs(ln (1 + x) - x) <= x^2" -proof - - assume x: "0 <= x" - assume "x <= 1" - from x have "ln (1 + x) <= x" - by (rule ln_add_one_self_le_self) - then have "ln (1 + x) - x <= 0" - by simp - then have "abs(ln(1 + x) - x) = - (ln(1 + x) - x)" - by (rule abs_of_nonpos) - also have "... = x - ln (1 + x)" - by simp - also have "... <= x^2" - proof - - from prems have "x - x^2 <= ln (1 + x)" - by (intro ln_one_plus_pos_lower_bound) - thus ?thesis - by simp - qed - finally show ?thesis . -qed - -lemma abs_ln_one_plus_x_minus_x_bound_nonpos: - "-(1 / 2) <= x ==> x <= 0 ==> abs(ln (1 + x) - x) <= 2 * x^2" -proof - - assume "-(1 / 2) <= x" - assume "x <= 0" - have "abs(ln (1 + x) - x) = x - ln(1 - (-x))" - apply (subst abs_of_nonpos) - apply simp - apply (rule ln_add_one_self_le_self2) - apply (insert prems, auto) - done - also have "... <= 2 * x^2" - apply (subgoal_tac "- (-x) - 2 * (-x)^2 <= ln (1 - (-x))") - apply (simp add: compare_rls) - apply (rule ln_one_minus_pos_lower_bound) - apply (insert prems, auto) - done - finally show ?thesis . -qed - -lemma abs_ln_one_plus_x_minus_x_bound: - "abs x <= 1 / 2 ==> abs(ln (1 + x) - x) <= 2 * x^2" - apply (case_tac "0 <= x") - apply (rule order_trans) - apply (rule abs_ln_one_plus_x_minus_x_bound_nonneg) - apply auto - apply (rule abs_ln_one_plus_x_minus_x_bound_nonpos) - apply auto -done - -lemma DERIV_ln: "0 < x ==> DERIV ln x :> 1 / x" - apply (unfold deriv_def, unfold LIM_def, clarsimp) - apply (rule exI) - apply (rule conjI) - prefer 2 - apply clarsimp - apply (subgoal_tac "(ln (x + xa) - ln x) / xa - (1 / x) = - (ln (1 + xa / x) - xa / x) / xa") - apply (erule ssubst) - apply (subst abs_divide) - apply (rule mult_imp_div_pos_less) - apply force - apply (rule order_le_less_trans) - apply (rule abs_ln_one_plus_x_minus_x_bound) - apply (subst abs_divide) - apply (subst abs_of_pos, assumption) - apply (erule mult_imp_div_pos_le) - apply (subgoal_tac "abs xa < min (x / 2) (r * x^2 / 2)") - apply force - apply assumption - apply (simp add: power2_eq_square mult_compare_simps) - apply (rule mult_imp_div_pos_less) - apply (rule mult_pos_pos, assumption, assumption) - apply (subgoal_tac "xa * xa = abs xa * abs xa") - apply (erule ssubst) - apply (subgoal_tac "abs xa * (abs xa * 2) < abs xa * (r * (x * x))") - apply (simp only: mult_ac) - apply (rule mult_strict_left_mono) - apply (erule conjE, assumption) - apply force - apply simp - apply (subst ln_div [THEN sym]) - apply arith - apply (auto simp add: ring_simps add_frac_eq frac_eq_eq - add_divide_distrib power2_eq_square) - apply (rule mult_pos_pos, assumption)+ - apply assumption -done - -lemma ln_x_over_x_mono: "exp 1 <= x ==> x <= y ==> (ln y / y) <= (ln x / x)" -proof - - assume "exp 1 <= x" and "x <= y" - have a: "0 < x" and b: "0 < y" - apply (insert prems) - apply (subgoal_tac "0 < exp (1::real)") - apply arith - apply auto - apply (subgoal_tac "0 < exp (1::real)") - apply arith - apply auto - done - have "x * ln y - x * ln x = x * (ln y - ln x)" - by (simp add: ring_simps) - also have "... = x * ln(y / x)" - apply (subst ln_div) - apply (rule b, rule a, rule refl) - done - also have "y / x = (x + (y - x)) / x" - by simp - also have "... = 1 + (y - x) / x" using a prems by(simp add:field_simps) - also have "x * ln(1 + (y - x) / x) <= x * ((y - x) / x)" - apply (rule mult_left_mono) - apply (rule ln_add_one_self_le_self) - apply (rule divide_nonneg_pos) - apply (insert prems a, simp_all) - done - also have "... = y - x" using a by simp - also have "... = (y - x) * ln (exp 1)" by simp - also have "... <= (y - x) * ln x" - apply (rule mult_left_mono) - apply (subst ln_le_cancel_iff) - apply force - apply (rule a) - apply (rule prems) - apply (insert prems, simp) - done - also have "... = y * ln x - x * ln x" - by (rule left_diff_distrib) - finally have "x * ln y <= y * ln x" - by arith - then have "ln y <= (y * ln x) / x" using a by(simp add:field_simps) - also have "... = y * (ln x / x)" by simp - finally show ?thesis using b by(simp add:field_simps) -qed - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/Log.thy --- a/src/HOL/Hyperreal/Log.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,276 +0,0 @@ -(* Title : Log.thy - Author : Jacques D. Fleuriot - Additional contributions by Jeremy Avigad - Copyright : 2000,2001 University of Edinburgh -*) - -header{*Logarithms: Standard Version*} - -theory Log -imports Transcendental -begin - -definition - powr :: "[real,real] => real" (infixr "powr" 80) where - --{*exponentation with real exponent*} - "x powr a = exp(a * ln x)" - -definition - log :: "[real,real] => real" where - --{*logarithm of @{term x} to base @{term a}*} - "log a x = ln x / ln a" - - - -lemma powr_one_eq_one [simp]: "1 powr a = 1" -by (simp add: powr_def) - -lemma powr_zero_eq_one [simp]: "x powr 0 = 1" -by (simp add: powr_def) - -lemma powr_one_gt_zero_iff [simp]: "(x powr 1 = x) = (0 < x)" -by (simp add: powr_def) -declare powr_one_gt_zero_iff [THEN iffD2, simp] - -lemma powr_mult: - "[| 0 < x; 0 < y |] ==> (x * y) powr a = (x powr a) * (y powr a)" -by (simp add: powr_def exp_add [symmetric] ln_mult right_distrib) - -lemma powr_gt_zero [simp]: "0 < x powr a" -by (simp add: powr_def) - -lemma powr_ge_pzero [simp]: "0 <= x powr y" -by (rule order_less_imp_le, rule powr_gt_zero) - -lemma powr_not_zero [simp]: "x powr a \ 0" -by (simp add: powr_def) - -lemma powr_divide: - "[| 0 < x; 0 < y |] ==> (x / y) powr a = (x powr a)/(y powr a)" -apply (simp add: divide_inverse positive_imp_inverse_positive powr_mult) -apply (simp add: powr_def exp_minus [symmetric] exp_add [symmetric] ln_inverse) -done - -lemma powr_divide2: "x powr a / x powr b = x powr (a - b)" - apply (simp add: powr_def) - apply (subst exp_diff [THEN sym]) - apply (simp add: left_diff_distrib) -done - -lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)" -by (simp add: powr_def exp_add [symmetric] left_distrib) - -lemma powr_powr: "(x powr a) powr b = x powr (a * b)" -by (simp add: powr_def) - -lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a" -by (simp add: powr_powr real_mult_commute) - -lemma powr_minus: "x powr (-a) = inverse (x powr a)" -by (simp add: powr_def exp_minus [symmetric]) - -lemma powr_minus_divide: "x powr (-a) = 1/(x powr a)" -by (simp add: divide_inverse powr_minus) - -lemma powr_less_mono: "[| a < b; 1 < x |] ==> x powr a < x powr b" -by (simp add: powr_def) - -lemma powr_less_cancel: "[| x powr a < x powr b; 1 < x |] ==> a < b" -by (simp add: powr_def) - -lemma powr_less_cancel_iff [simp]: "1 < x ==> (x powr a < x powr b) = (a < b)" -by (blast intro: powr_less_cancel powr_less_mono) - -lemma powr_le_cancel_iff [simp]: "1 < x ==> (x powr a \ x powr b) = (a \ b)" -by (simp add: linorder_not_less [symmetric]) - -lemma log_ln: "ln x = log (exp(1)) x" -by (simp add: log_def) - -lemma powr_log_cancel [simp]: - "[| 0 < a; a \ 1; 0 < x |] ==> a powr (log a x) = x" -by (simp add: powr_def log_def) - -lemma log_powr_cancel [simp]: "[| 0 < a; a \ 1 |] ==> log a (a powr y) = y" -by (simp add: log_def powr_def) - -lemma log_mult: - "[| 0 < a; a \ 1; 0 < x; 0 < y |] - ==> log a (x * y) = log a x + log a y" -by (simp add: log_def ln_mult divide_inverse left_distrib) - -lemma log_eq_div_ln_mult_log: - "[| 0 < a; a \ 1; 0 < b; b \ 1; 0 < x |] - ==> log a x = (ln b/ln a) * log b x" -by (simp add: log_def divide_inverse) - -text{*Base 10 logarithms*} -lemma log_base_10_eq1: "0 < x ==> log 10 x = (ln (exp 1) / ln 10) * ln x" -by (simp add: log_def) - -lemma log_base_10_eq2: "0 < x ==> log 10 x = (log 10 (exp 1)) * ln x" -by (simp add: log_def) - -lemma log_one [simp]: "log a 1 = 0" -by (simp add: log_def) - -lemma log_eq_one [simp]: "[| 0 < a; a \ 1 |] ==> log a a = 1" -by (simp add: log_def) - -lemma log_inverse: - "[| 0 < a; a \ 1; 0 < x |] ==> log a (inverse x) = - log a x" -apply (rule_tac a1 = "log a x" in add_left_cancel [THEN iffD1]) -apply (simp add: log_mult [symmetric]) -done - -lemma log_divide: - "[|0 < a; a \ 1; 0 < x; 0 < y|] ==> log a (x/y) = log a x - log a y" -by (simp add: log_mult divide_inverse log_inverse) - -lemma log_less_cancel_iff [simp]: - "[| 1 < a; 0 < x; 0 < y |] ==> (log a x < log a y) = (x < y)" -apply safe -apply (rule_tac [2] powr_less_cancel) -apply (drule_tac a = "log a x" in powr_less_mono, auto) -done - -lemma log_le_cancel_iff [simp]: - "[| 1 < a; 0 < x; 0 < y |] ==> (log a x \ log a y) = (x \ y)" -by (simp add: linorder_not_less [symmetric]) - - -lemma powr_realpow: "0 < x ==> x powr (real n) = x^n" - apply (induct n, simp) - apply (subgoal_tac "real(Suc n) = real n + 1") - apply (erule ssubst) - apply (subst powr_add, simp, simp) -done - -lemma powr_realpow2: "0 <= x ==> 0 < n ==> x^n = (if (x = 0) then 0 - else x powr (real n))" - apply (case_tac "x = 0", simp, simp) - apply (rule powr_realpow [THEN sym], simp) -done - -lemma ln_pwr: "0 < x ==> 0 < y ==> ln(x powr y) = y * ln x" -by (unfold powr_def, simp) - -lemma ln_bound: "1 <= x ==> ln x <= x" - apply (subgoal_tac "ln(1 + (x - 1)) <= x - 1") - apply simp - apply (rule ln_add_one_self_le_self, simp) -done - -lemma powr_mono: "a <= b ==> 1 <= x ==> x powr a <= x powr b" - apply (case_tac "x = 1", simp) - apply (case_tac "a = b", simp) - apply (rule order_less_imp_le) - apply (rule powr_less_mono, auto) -done - -lemma ge_one_powr_ge_zero: "1 <= x ==> 0 <= a ==> 1 <= x powr a" - apply (subst powr_zero_eq_one [THEN sym]) - apply (rule powr_mono, assumption+) -done - -lemma powr_less_mono2: "0 < a ==> 0 < x ==> x < y ==> x powr a < - y powr a" - apply (unfold powr_def) - apply (rule exp_less_mono) - apply (rule mult_strict_left_mono) - apply (subst ln_less_cancel_iff, assumption) - apply (rule order_less_trans) - prefer 2 - apply assumption+ -done - -lemma powr_less_mono2_neg: "a < 0 ==> 0 < x ==> x < y ==> y powr a < - x powr a" - apply (unfold powr_def) - apply (rule exp_less_mono) - apply (rule mult_strict_left_mono_neg) - apply (subst ln_less_cancel_iff) - apply assumption - apply (rule order_less_trans) - prefer 2 - apply assumption+ -done - -lemma powr_mono2: "0 <= a ==> 0 < x ==> x <= y ==> x powr a <= y powr a" - apply (case_tac "a = 0", simp) - apply (case_tac "x = y", simp) - apply (rule order_less_imp_le) - apply (rule powr_less_mono2, auto) -done - -lemma ln_powr_bound: "1 <= x ==> 0 < a ==> ln x <= (x powr a) / a" - apply (rule mult_imp_le_div_pos) - apply (assumption) - apply (subst mult_commute) - apply (subst ln_pwr [THEN sym]) - apply auto - apply (rule ln_bound) - apply (erule ge_one_powr_ge_zero) - apply (erule order_less_imp_le) -done - -lemma ln_powr_bound2: "1 < x ==> 0 < a ==> (ln x) powr a <= (a powr a) * x" -proof - - assume "1 < x" and "0 < a" - then have "ln x <= (x powr (1 / a)) / (1 / a)" - apply (intro ln_powr_bound) - apply (erule order_less_imp_le) - apply (rule divide_pos_pos) - apply simp_all - done - also have "... = a * (x powr (1 / a))" - by simp - finally have "(ln x) powr a <= (a * (x powr (1 / a))) powr a" - apply (intro powr_mono2) - apply (rule order_less_imp_le, rule prems) - apply (rule ln_gt_zero) - apply (rule prems) - apply assumption - done - also have "... = (a powr a) * ((x powr (1 / a)) powr a)" - apply (rule powr_mult) - apply (rule prems) - apply (rule powr_gt_zero) - done - also have "(x powr (1 / a)) powr a = x powr ((1 / a) * a)" - by (rule powr_powr) - also have "... = x" - apply simp - apply (subgoal_tac "a ~= 0") - apply (insert prems, auto) - done - finally show ?thesis . -qed - -lemma LIMSEQ_neg_powr: "0 < s ==> (%x. (real x) powr - s) ----> 0" - apply (unfold LIMSEQ_def) - apply clarsimp - apply (rule_tac x = "natfloor(r powr (1 / - s)) + 1" in exI) - apply clarify - proof - - fix r fix n - assume "0 < s" and "0 < r" and "natfloor (r powr (1 / - s)) + 1 <= n" - have "r powr (1 / - s) < real(natfloor(r powr (1 / - s))) + 1" - by (rule real_natfloor_add_one_gt) - also have "... = real(natfloor(r powr (1 / -s)) + 1)" - by simp - also have "... <= real n" - apply (subst real_of_nat_le_iff) - apply (rule prems) - done - finally have "r powr (1 / - s) < real n". - then have "real n powr (- s) < (r powr (1 / - s)) powr - s" - apply (intro powr_less_mono2_neg) - apply (auto simp add: prems) - done - also have "... = r" - by (simp add: powr_powr prems less_imp_neq [THEN not_sym]) - finally show "real n powr - s < r" . - qed - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/MacLaurin.thy --- a/src/HOL/Hyperreal/MacLaurin.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,577 +0,0 @@ -(* ID : $Id$ - Author : Jacques D. Fleuriot - Copyright : 2001 University of Edinburgh - Conversion to Isar and new proofs by Lawrence C Paulson, 2004 -*) - -header{*MacLaurin Series*} - -theory MacLaurin -imports Dense_Linear_Order Transcendental -begin - -subsection{*Maclaurin's Theorem with Lagrange Form of Remainder*} - -text{*This is a very long, messy proof even now that it's been broken down -into lemmas.*} - -lemma Maclaurin_lemma: - "0 < h ==> - \B. f h = (\m=0.. - resolve_tac @{thms deriv_rulesI} i ORELSE - ((rtac (read_instantiate ctxt [(("f", 0), get_fun_name prem)] - @{thm DERIV_chain2}) i) handle DERIV_name => no_tac)); - -fun DERIV_tac ctxt = ALLGOALS (fn i => REPEAT (deriv_tac ctxt i)); - -end -*} - -lemma Maclaurin_lemma2: - "[| \m t. m < n \ 0\t \ t\h \ DERIV (diff m) t :> diff (Suc m) t; - n = Suc k; - difg = - (\m t. diff m t - - ((\p = 0.. - \m t. m < n & 0 \ t & t \ h --> - DERIV (difg m) t :> difg (Suc m) t" -apply clarify -apply (rule DERIV_diff) -apply (simp (no_asm_simp)) -apply (tactic {* DERIV_tac @{context} *}) -apply (tactic {* DERIV_tac @{context} *}) -apply (rule_tac [2] lemma_DERIV_subst) -apply (rule_tac [2] DERIV_quotient) -apply (rule_tac [3] DERIV_const) -apply (rule_tac [2] DERIV_pow) - prefer 3 apply (simp add: fact_diff_Suc) - prefer 2 apply simp -apply (frule_tac m = m in less_add_one, clarify) -apply (simp del: setsum_op_ivl_Suc) -apply (insert sumr_offset4 [of 1]) -apply (simp del: setsum_op_ivl_Suc fact_Suc realpow_Suc) -apply (rule lemma_DERIV_subst) -apply (rule DERIV_add) -apply (rule_tac [2] DERIV_const) -apply (rule DERIV_sumr, clarify) - prefer 2 apply simp -apply (simp (no_asm) add: divide_inverse mult_assoc del: fact_Suc realpow_Suc) -apply (rule DERIV_cmult) -apply (rule lemma_DERIV_subst) -apply (best intro: DERIV_chain2 intro!: DERIV_intros) -apply (subst fact_Suc) -apply (subst real_of_nat_mult) -apply (simp add: mult_ac) -done - - -lemma Maclaurin_lemma3: - fixes difg :: "nat => real => real" shows - "[|\k t. k < Suc m \ 0\t & t\h \ DERIV (difg k) t :> difg (Suc k) t; - \k 0; n < m; 0 < t; - t < h|] - ==> \ta. 0 < ta & ta < t & DERIV (difg (Suc n)) ta :> 0" -apply (rule Rolle, assumption, simp) -apply (drule_tac x = n and P="%k. k difg k 0 = 0" in spec) -apply (rule DERIV_unique) -prefer 2 apply assumption -apply force -apply (metis DERIV_isCont dlo_simps(4) dlo_simps(9) less_trans_Suc nat_less_le not_less_eq real_le_trans) -apply (metis Suc_less_eq differentiableI dlo_simps(7) dlo_simps(8) dlo_simps(9) real_le_trans xt1(8)) -done - -lemma Maclaurin: - "[| 0 < h; n > 0; diff 0 = f; - \m t. m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t |] - ==> \t. 0 < t & - t < h & - f h = - setsum (%m. (diff m 0 / real (fact m)) * h ^ m) {0..g. - g = (%t. f t - (setsum (%m. (diff m 0 / real(fact m)) * t^m) {0..difg. difg = (%m t. diff m t - (setsum (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) {0..ma. ma < n --> (\t. 0 < t & t < h & difg (Suc ma) t = 0) ") - apply (drule_tac x = m and P="%m. m (\t. ?QQ m t)" in spec) - apply (erule impE) - apply (simp (no_asm_simp)) - apply (erule exE) - apply (rule_tac x = t in exI) - apply (simp del: realpow_Suc fact_Suc) -apply (subgoal_tac "\m. m < n --> difg m 0 = 0") - prefer 2 - apply clarify - apply simp - apply (frule_tac m = ma in less_add_one, clarify) - apply (simp del: setsum_op_ivl_Suc) -apply (insert sumr_offset4 [of 1]) -apply (simp del: setsum_op_ivl_Suc fact_Suc realpow_Suc) -apply (subgoal_tac "\m. m < n --> (\t. 0 < t & t < h & DERIV (difg m) t :> 0) ") -apply (rule allI, rule impI) -apply (drule_tac x = ma and P="%m. m (\t. ?QQ m t)" in spec) -apply (erule impE, assumption) -apply (erule exE) -apply (rule_tac x = t in exI) -(* do some tidying up *) -apply (erule_tac [!] V= "difg = (%m t. diff m t - (setsum (%p. diff (m + p) 0 / real (fact p) * t ^ p) {0..0 & diff 0 = f & - (\m t. m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t) - --> (\t. 0 < t & t < h & - f h = (\m=0..m t. - m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t |] - ==> \t. 0 < t & - t \ h & - f h = - (\m=0..m t. - m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t) - --> (\t. 0 < t & - t \ h & - f h = - (\m=0.. 0; diff 0 = f; - \m t. m < n & h \ t & t \ 0 --> DERIV (diff m) t :> diff (Suc m) t |] - ==> \t. h < t & - t < 0 & - f h = - (\m=0..m = 0..m = 0.. 0 & diff 0 = f & - (\m t. - m < n & h \ t & t \ 0 --> DERIV (diff m) t :> diff (Suc m) t)) - --> (\t. h < t & - t < 0 & - f h = - (\m=0..0 \ - diff 0 0 = - (\m = 0..m t. m < n & abs t \ abs x --> DERIV (diff m) t :> diff (Suc m) t |] - ==> \t. abs t \ abs x & - f x = - (\m=0..m x. DERIV (diff m) x :> diff(Suc m) x; - x ~= 0; n > 0 - |] ==> \t. 0 < abs t & abs t < abs x & - f x = (\m=0..m x. DERIV (diff m) x :> diff(Suc m) x) & - x ~= 0 & n > 0 - --> (\t. 0 < abs t & abs t < abs x & - f x = (\m=0.. n \ 0 --> - (\m=0..m x. DERIV (diff m) x :> diff (Suc m) x - |] ==> \t. abs t \ abs x & - f x = (\m=0..m x. DERIV (diff m) x :> diff (Suc m) x) - --> (\t. abs t \ abs x & - f x = (\m=0.. 0 |] - ==> (\t. 0 < abs t & - abs t < abs x & - exp x = (\m=0..t. abs t \ abs x & - exp x = (\m=0..x. a \ x & x \ b --> DERIV f x :> f'(x) |] - ==> \z::real. a < z & z < b & (f b - f a = (b - a) * f'(z))" -apply (drule MVT) -apply (blast intro: DERIV_isCont) -apply (force dest: order_less_imp_le simp add: differentiable_def) -apply (blast dest: DERIV_unique order_less_imp_le) -done - -lemma mod_exhaust_less_4: - "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = (3::nat)" -by auto - -lemma Suc_Suc_mult_two_diff_two [rule_format, simp]: - "n\0 --> Suc (Suc (2 * n - 2)) = 2*n" -by (induct "n", auto) - -lemma lemma_Suc_Suc_4n_diff_2 [rule_format, simp]: - "n\0 --> Suc (Suc (4*n - 2)) = 4*n" -by (induct "n", auto) - -lemma Suc_mult_two_diff_one [rule_format, simp]: - "n\0 --> Suc (2 * n - 1) = 2*n" -by (induct "n", auto) - - -text{*It is unclear why so many variant results are needed.*} - -lemma Maclaurin_sin_expansion2: - "\t. abs t \ abs x & - sin x = - (\m=0..t. sin x = - (\m=0.. 0; 0 < x |] ==> - \t. 0 < t & t < x & - sin x = - (\m=0.. - \t. 0 < t & t \ x & - sin x = - (\m=0..m=0..<(Suc n). - (if even m then -1 ^ (m div 2)/(real (fact m)) else 0) * 0 ^ m) = 1" -by (induct "n", auto) - -lemma Maclaurin_cos_expansion: - "\t. abs t \ abs x & - cos x = - (\m=0.. 0 |] ==> - \t. 0 < t & t < x & - cos x = - (\m=0.. 0 |] ==> - \t. x < t & t < 0 & - cos x = - (\m=0.. (v::real) |] ==> \(x + u) - y\ \ v" -by auto - -lemma Maclaurin_sin_bound: - "abs(sin x - (\m=0.. inverse(real (fact n)) * \x\ ^ n" -proof - - have "!! x (y::real). x \ 1 \ 0 \ y \ x * y \ 1 * y" - by (rule_tac mult_right_mono,simp_all) - note est = this[simplified] - let ?diff = "\(n::nat) x. if n mod 4 = 0 then sin(x) else if n mod 4 = 1 then cos(x) else if n mod 4 = 2 then -sin(x) else -cos(x)" - have diff_0: "?diff 0 = sin" by simp - have DERIV_diff: "\m x. DERIV (?diff m) x :> ?diff (Suc m) x" - apply (clarify) - apply (subst (1 2 3) mod_Suc_eq_Suc_mod) - apply (cut_tac m=m in mod_exhaust_less_4) - apply (safe, simp_all) - apply (rule DERIV_minus, simp) - apply (rule lemma_DERIV_subst, rule DERIV_minus, rule DERIV_cos, simp) - done - from Maclaurin_all_le [OF diff_0 DERIV_diff] - obtain t where t1: "\t\ \ \x\" and - t2: "sin x = (\m = 0..m. ?diff m 0 = (if even m then 0 - else -1 ^ ((m - Suc 0) div 2))" - apply (subst even_even_mod_4_iff) - apply (cut_tac m=m in mod_exhaust_less_4) - apply (elim disjE, simp_all) - apply (safe dest!: mod_eqD, simp_all) - done - show ?thesis - apply (subst t2) - apply (rule sin_bound_lemma) - apply (rule setsum_cong[OF refl]) - apply (subst diff_m_0, simp) - apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono - simp add: est mult_nonneg_nonneg mult_ac divide_inverse - power_abs [symmetric] abs_mult) - done -qed - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/NthRoot.thy --- a/src/HOL/Hyperreal/NthRoot.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,631 +0,0 @@ -(* Title : NthRoot.thy - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Conversion to Isar and new proofs by Lawrence C Paulson, 2004 -*) - -header {* Nth Roots of Real Numbers *} - -theory NthRoot -imports "~~/src/HOL/Library/Parity" "../Hyperreal/Deriv" -begin - -subsection {* Existence of Nth Root *} - -text {* Existence follows from the Intermediate Value Theorem *} - -lemma realpow_pos_nth: - assumes n: "0 < n" - assumes a: "0 < a" - shows "\r>0. r ^ n = (a::real)" -proof - - have "\r\0. r \ (max 1 a) \ r ^ n = a" - proof (rule IVT) - show "0 ^ n \ a" using n a by (simp add: power_0_left) - show "0 \ max 1 a" by simp - from n have n1: "1 \ n" by simp - have "a \ max 1 a ^ 1" by simp - also have "max 1 a ^ 1 \ max 1 a ^ n" - using n1 by (rule power_increasing, simp) - finally show "a \ max 1 a ^ n" . - show "\r. 0 \ r \ r \ max 1 a \ isCont (\x. x ^ n) r" - by (simp add: isCont_power) - qed - then obtain r where r: "0 \ r \ r ^ n = a" by fast - with n a have "r \ 0" by (auto simp add: power_0_left) - with r have "0 < r \ r ^ n = a" by simp - thus ?thesis .. -qed - -(* Used by Integration/RealRandVar.thy in AFP *) -lemma realpow_pos_nth2: "(0::real) < a \ \r>0. r ^ Suc n = a" -by (blast intro: realpow_pos_nth) - -text {* Uniqueness of nth positive root *} - -lemma realpow_pos_nth_unique: - "\0 < n; 0 < a\ \ \!r. 0 < r \ r ^ n = (a::real)" -apply (auto intro!: realpow_pos_nth) -apply (rule_tac n=n in power_eq_imp_eq_base, simp_all) -done - -subsection {* Nth Root *} - -text {* We define roots of negative reals such that - @{term "root n (- x) = - root n x"}. This allows - us to omit side conditions from many theorems. *} - -definition - root :: "[nat, real] \ real" where - "root n x = (if 0 < x then (THE u. 0 < u \ u ^ n = x) else - if x < 0 then - (THE u. 0 < u \ u ^ n = - x) else 0)" - -lemma real_root_zero [simp]: "root n 0 = 0" -unfolding root_def by simp - -lemma real_root_minus: "0 < n \ root n (- x) = - root n x" -unfolding root_def by simp - -lemma real_root_gt_zero: "\0 < n; 0 < x\ \ 0 < root n x" -apply (simp add: root_def) -apply (drule (1) realpow_pos_nth_unique) -apply (erule theI' [THEN conjunct1]) -done - -lemma real_root_pow_pos: (* TODO: rename *) - "\0 < n; 0 < x\ \ root n x ^ n = x" -apply (simp add: root_def) -apply (drule (1) realpow_pos_nth_unique) -apply (erule theI' [THEN conjunct2]) -done - -lemma real_root_pow_pos2 [simp]: (* TODO: rename *) - "\0 < n; 0 \ x\ \ root n x ^ n = x" -by (auto simp add: order_le_less real_root_pow_pos) - -lemma odd_real_root_pow: "odd n \ root n x ^ n = x" -apply (rule_tac x=0 and y=x in linorder_le_cases) -apply (erule (1) real_root_pow_pos2 [OF odd_pos]) -apply (subgoal_tac "root n (- x) ^ n = - x") -apply (simp add: real_root_minus odd_pos) -apply (simp add: odd_pos) -done - -lemma real_root_ge_zero: "\0 < n; 0 \ x\ \ 0 \ root n x" -by (auto simp add: order_le_less real_root_gt_zero) - -lemma real_root_power_cancel: "\0 < n; 0 \ x\ \ root n (x ^ n) = x" -apply (subgoal_tac "0 \ x ^ n") -apply (subgoal_tac "0 \ root n (x ^ n)") -apply (subgoal_tac "root n (x ^ n) ^ n = x ^ n") -apply (erule (3) power_eq_imp_eq_base) -apply (erule (1) real_root_pow_pos2) -apply (erule (1) real_root_ge_zero) -apply (erule zero_le_power) -done - -lemma odd_real_root_power_cancel: "odd n \ root n (x ^ n) = x" -apply (rule_tac x=0 and y=x in linorder_le_cases) -apply (erule (1) real_root_power_cancel [OF odd_pos]) -apply (subgoal_tac "root n ((- x) ^ n) = - x") -apply (simp add: real_root_minus odd_pos) -apply (erule real_root_power_cancel [OF odd_pos], simp) -done - -lemma real_root_pos_unique: - "\0 < n; 0 \ y; y ^ n = x\ \ root n x = y" -by (erule subst, rule real_root_power_cancel) - -lemma odd_real_root_unique: - "\odd n; y ^ n = x\ \ root n x = y" -by (erule subst, rule odd_real_root_power_cancel) - -lemma real_root_one [simp]: "0 < n \ root n 1 = 1" -by (simp add: real_root_pos_unique) - -text {* Root function is strictly monotonic, hence injective *} - -lemma real_root_less_mono_lemma: - "\0 < n; 0 \ x; x < y\ \ root n x < root n y" -apply (subgoal_tac "0 \ y") -apply (subgoal_tac "root n x ^ n < root n y ^ n") -apply (erule power_less_imp_less_base) -apply (erule (1) real_root_ge_zero) -apply simp -apply simp -done - -lemma real_root_less_mono: "\0 < n; x < y\ \ root n x < root n y" -apply (cases "0 \ x") -apply (erule (2) real_root_less_mono_lemma) -apply (cases "0 \ y") -apply (rule_tac y=0 in order_less_le_trans) -apply (subgoal_tac "0 < root n (- x)") -apply (simp add: real_root_minus) -apply (simp add: real_root_gt_zero) -apply (simp add: real_root_ge_zero) -apply (subgoal_tac "root n (- y) < root n (- x)") -apply (simp add: real_root_minus) -apply (simp add: real_root_less_mono_lemma) -done - -lemma real_root_le_mono: "\0 < n; x \ y\ \ root n x \ root n y" -by (auto simp add: order_le_less real_root_less_mono) - -lemma real_root_less_iff [simp]: - "0 < n \ (root n x < root n y) = (x < y)" -apply (cases "x < y") -apply (simp add: real_root_less_mono) -apply (simp add: linorder_not_less real_root_le_mono) -done - -lemma real_root_le_iff [simp]: - "0 < n \ (root n x \ root n y) = (x \ y)" -apply (cases "x \ y") -apply (simp add: real_root_le_mono) -apply (simp add: linorder_not_le real_root_less_mono) -done - -lemma real_root_eq_iff [simp]: - "0 < n \ (root n x = root n y) = (x = y)" -by (simp add: order_eq_iff) - -lemmas real_root_gt_0_iff [simp] = real_root_less_iff [where x=0, simplified] -lemmas real_root_lt_0_iff [simp] = real_root_less_iff [where y=0, simplified] -lemmas real_root_ge_0_iff [simp] = real_root_le_iff [where x=0, simplified] -lemmas real_root_le_0_iff [simp] = real_root_le_iff [where y=0, simplified] -lemmas real_root_eq_0_iff [simp] = real_root_eq_iff [where y=0, simplified] - -lemma real_root_gt_1_iff [simp]: "0 < n \ (1 < root n y) = (1 < y)" -by (insert real_root_less_iff [where x=1], simp) - -lemma real_root_lt_1_iff [simp]: "0 < n \ (root n x < 1) = (x < 1)" -by (insert real_root_less_iff [where y=1], simp) - -lemma real_root_ge_1_iff [simp]: "0 < n \ (1 \ root n y) = (1 \ y)" -by (insert real_root_le_iff [where x=1], simp) - -lemma real_root_le_1_iff [simp]: "0 < n \ (root n x \ 1) = (x \ 1)" -by (insert real_root_le_iff [where y=1], simp) - -lemma real_root_eq_1_iff [simp]: "0 < n \ (root n x = 1) = (x = 1)" -by (insert real_root_eq_iff [where y=1], simp) - -text {* Roots of roots *} - -lemma real_root_Suc_0 [simp]: "root (Suc 0) x = x" -by (simp add: odd_real_root_unique) - -lemma real_root_pos_mult_exp: - "\0 < m; 0 < n; 0 < x\ \ root (m * n) x = root m (root n x)" -by (rule real_root_pos_unique, simp_all add: power_mult) - -lemma real_root_mult_exp: - "\0 < m; 0 < n\ \ root (m * n) x = root m (root n x)" -apply (rule linorder_cases [where x=x and y=0]) -apply (subgoal_tac "root (m * n) (- x) = root m (root n (- x))") -apply (simp add: real_root_minus) -apply (simp_all add: real_root_pos_mult_exp) -done - -lemma real_root_commute: - "\0 < m; 0 < n\ \ root m (root n x) = root n (root m x)" -by (simp add: real_root_mult_exp [symmetric] mult_commute) - -text {* Monotonicity in first argument *} - -lemma real_root_strict_decreasing: - "\0 < n; n < N; 1 < x\ \ root N x < root n x" -apply (subgoal_tac "root n (root N x) ^ n < root N (root n x) ^ N", simp) -apply (simp add: real_root_commute power_strict_increasing - del: real_root_pow_pos2) -done - -lemma real_root_strict_increasing: - "\0 < n; n < N; 0 < x; x < 1\ \ root n x < root N x" -apply (subgoal_tac "root N (root n x) ^ N < root n (root N x) ^ n", simp) -apply (simp add: real_root_commute power_strict_decreasing - del: real_root_pow_pos2) -done - -lemma real_root_decreasing: - "\0 < n; n < N; 1 \ x\ \ root N x \ root n x" -by (auto simp add: order_le_less real_root_strict_decreasing) - -lemma real_root_increasing: - "\0 < n; n < N; 0 \ x; x \ 1\ \ root n x \ root N x" -by (auto simp add: order_le_less real_root_strict_increasing) - -text {* Roots of multiplication and division *} - -lemma real_root_mult_lemma: - "\0 < n; 0 \ x; 0 \ y\ \ root n (x * y) = root n x * root n y" -by (simp add: real_root_pos_unique mult_nonneg_nonneg power_mult_distrib) - -lemma real_root_inverse_lemma: - "\0 < n; 0 \ x\ \ root n (inverse x) = inverse (root n x)" -by (simp add: real_root_pos_unique power_inverse [symmetric]) - -lemma real_root_mult: - assumes n: "0 < n" - shows "root n (x * y) = root n x * root n y" -proof (rule linorder_le_cases, rule_tac [!] linorder_le_cases) - assume "0 \ x" and "0 \ y" - thus ?thesis by (rule real_root_mult_lemma [OF n]) -next - assume "0 \ x" and "y \ 0" - hence "0 \ x" and "0 \ - y" by simp_all - hence "root n (x * - y) = root n x * root n (- y)" - by (rule real_root_mult_lemma [OF n]) - thus ?thesis by (simp add: real_root_minus [OF n]) -next - assume "x \ 0" and "0 \ y" - hence "0 \ - x" and "0 \ y" by simp_all - hence "root n (- x * y) = root n (- x) * root n y" - by (rule real_root_mult_lemma [OF n]) - thus ?thesis by (simp add: real_root_minus [OF n]) -next - assume "x \ 0" and "y \ 0" - hence "0 \ - x" and "0 \ - y" by simp_all - hence "root n (- x * - y) = root n (- x) * root n (- y)" - by (rule real_root_mult_lemma [OF n]) - thus ?thesis by (simp add: real_root_minus [OF n]) -qed - -lemma real_root_inverse: - assumes n: "0 < n" - shows "root n (inverse x) = inverse (root n x)" -proof (rule linorder_le_cases) - assume "0 \ x" - thus ?thesis by (rule real_root_inverse_lemma [OF n]) -next - assume "x \ 0" - hence "0 \ - x" by simp - hence "root n (inverse (- x)) = inverse (root n (- x))" - by (rule real_root_inverse_lemma [OF n]) - thus ?thesis by (simp add: real_root_minus [OF n]) -qed - -lemma real_root_divide: - "0 < n \ root n (x / y) = root n x / root n y" -by (simp add: divide_inverse real_root_mult real_root_inverse) - -lemma real_root_power: - "0 < n \ root n (x ^ k) = root n x ^ k" -by (induct k, simp_all add: real_root_mult) - -lemma real_root_abs: "0 < n \ root n \x\ = \root n x\" -by (simp add: abs_if real_root_minus) - -text {* Continuity and derivatives *} - -lemma isCont_root_pos: - assumes n: "0 < n" - assumes x: "0 < x" - shows "isCont (root n) x" -proof - - have "isCont (root n) (root n x ^ n)" - proof (rule isCont_inverse_function [where f="\a. a ^ n"]) - show "0 < root n x" using n x by simp - show "\z. \z - root n x\ \ root n x \ root n (z ^ n) = z" - by (simp add: abs_le_iff real_root_power_cancel n) - show "\z. \z - root n x\ \ root n x \ isCont (\a. a ^ n) z" - by (simp add: isCont_power) - qed - thus ?thesis using n x by simp -qed - -lemma isCont_root_neg: - "\0 < n; x < 0\ \ isCont (root n) x" -apply (subgoal_tac "isCont (\x. - root n (- x)) x") -apply (simp add: real_root_minus) -apply (rule isCont_o2 [OF isCont_minus [OF isCont_ident]]) -apply (simp add: isCont_minus isCont_root_pos) -done - -lemma isCont_root_zero: - "0 < n \ isCont (root n) 0" -unfolding isCont_def -apply (rule LIM_I) -apply (rule_tac x="r ^ n" in exI, safe) -apply (simp) -apply (simp add: real_root_abs [symmetric]) -apply (rule_tac n="n" in power_less_imp_less_base, simp_all) -done - -lemma isCont_real_root: "0 < n \ isCont (root n) x" -apply (rule_tac x=x and y=0 in linorder_cases) -apply (simp_all add: isCont_root_pos isCont_root_neg isCont_root_zero) -done - -lemma DERIV_real_root: - assumes n: "0 < n" - assumes x: "0 < x" - shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))" -proof (rule DERIV_inverse_function) - show "0 < x" using x . - show "x < x + 1" by simp - show "\y. 0 < y \ y < x + 1 \ root n y ^ n = y" - using n by simp - show "DERIV (\x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)" - by (rule DERIV_pow) - show "real n * root n x ^ (n - Suc 0) \ 0" - using n x by simp - show "isCont (root n) x" - using n by (rule isCont_real_root) -qed - -lemma DERIV_odd_real_root: - assumes n: "odd n" - assumes x: "x \ 0" - shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))" -proof (rule DERIV_inverse_function) - show "x - 1 < x" by simp - show "x < x + 1" by simp - show "\y. x - 1 < y \ y < x + 1 \ root n y ^ n = y" - using n by (simp add: odd_real_root_pow) - show "DERIV (\x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)" - by (rule DERIV_pow) - show "real n * root n x ^ (n - Suc 0) \ 0" - using odd_pos [OF n] x by simp - show "isCont (root n) x" - using odd_pos [OF n] by (rule isCont_real_root) -qed - -subsection {* Square Root *} - -definition - sqrt :: "real \ real" where - "sqrt = root 2" - -lemma pos2: "0 < (2::nat)" by simp - -lemma real_sqrt_unique: "\y\ = x; 0 \ y\ \ sqrt x = y" -unfolding sqrt_def by (rule real_root_pos_unique [OF pos2]) - -lemma real_sqrt_abs [simp]: "sqrt (x\) = \x\" -apply (rule real_sqrt_unique) -apply (rule power2_abs) -apply (rule abs_ge_zero) -done - -lemma real_sqrt_pow2 [simp]: "0 \ x \ (sqrt x)\ = x" -unfolding sqrt_def by (rule real_root_pow_pos2 [OF pos2]) - -lemma real_sqrt_pow2_iff [simp]: "((sqrt x)\ = x) = (0 \ x)" -apply (rule iffI) -apply (erule subst) -apply (rule zero_le_power2) -apply (erule real_sqrt_pow2) -done - -lemma real_sqrt_zero [simp]: "sqrt 0 = 0" -unfolding sqrt_def by (rule real_root_zero) - -lemma real_sqrt_one [simp]: "sqrt 1 = 1" -unfolding sqrt_def by (rule real_root_one [OF pos2]) - -lemma real_sqrt_minus: "sqrt (- x) = - sqrt x" -unfolding sqrt_def by (rule real_root_minus [OF pos2]) - -lemma real_sqrt_mult: "sqrt (x * y) = sqrt x * sqrt y" -unfolding sqrt_def by (rule real_root_mult [OF pos2]) - -lemma real_sqrt_inverse: "sqrt (inverse x) = inverse (sqrt x)" -unfolding sqrt_def by (rule real_root_inverse [OF pos2]) - -lemma real_sqrt_divide: "sqrt (x / y) = sqrt x / sqrt y" -unfolding sqrt_def by (rule real_root_divide [OF pos2]) - -lemma real_sqrt_power: "sqrt (x ^ k) = sqrt x ^ k" -unfolding sqrt_def by (rule real_root_power [OF pos2]) - -lemma real_sqrt_gt_zero: "0 < x \ 0 < sqrt x" -unfolding sqrt_def by (rule real_root_gt_zero [OF pos2]) - -lemma real_sqrt_ge_zero: "0 \ x \ 0 \ sqrt x" -unfolding sqrt_def by (rule real_root_ge_zero [OF pos2]) - -lemma real_sqrt_less_mono: "x < y \ sqrt x < sqrt y" -unfolding sqrt_def by (rule real_root_less_mono [OF pos2]) - -lemma real_sqrt_le_mono: "x \ y \ sqrt x \ sqrt y" -unfolding sqrt_def by (rule real_root_le_mono [OF pos2]) - -lemma real_sqrt_less_iff [simp]: "(sqrt x < sqrt y) = (x < y)" -unfolding sqrt_def by (rule real_root_less_iff [OF pos2]) - -lemma real_sqrt_le_iff [simp]: "(sqrt x \ sqrt y) = (x \ y)" -unfolding sqrt_def by (rule real_root_le_iff [OF pos2]) - -lemma real_sqrt_eq_iff [simp]: "(sqrt x = sqrt y) = (x = y)" -unfolding sqrt_def by (rule real_root_eq_iff [OF pos2]) - -lemmas real_sqrt_gt_0_iff [simp] = real_sqrt_less_iff [where x=0, simplified] -lemmas real_sqrt_lt_0_iff [simp] = real_sqrt_less_iff [where y=0, simplified] -lemmas real_sqrt_ge_0_iff [simp] = real_sqrt_le_iff [where x=0, simplified] -lemmas real_sqrt_le_0_iff [simp] = real_sqrt_le_iff [where y=0, simplified] -lemmas real_sqrt_eq_0_iff [simp] = real_sqrt_eq_iff [where y=0, simplified] - -lemmas real_sqrt_gt_1_iff [simp] = real_sqrt_less_iff [where x=1, simplified] -lemmas real_sqrt_lt_1_iff [simp] = real_sqrt_less_iff [where y=1, simplified] -lemmas real_sqrt_ge_1_iff [simp] = real_sqrt_le_iff [where x=1, simplified] -lemmas real_sqrt_le_1_iff [simp] = real_sqrt_le_iff [where y=1, simplified] -lemmas real_sqrt_eq_1_iff [simp] = real_sqrt_eq_iff [where y=1, simplified] - -lemma isCont_real_sqrt: "isCont sqrt x" -unfolding sqrt_def by (rule isCont_real_root [OF pos2]) - -lemma DERIV_real_sqrt: - "0 < x \ DERIV sqrt x :> inverse (sqrt x) / 2" -unfolding sqrt_def by (rule DERIV_real_root [OF pos2, simplified]) - -lemma not_real_square_gt_zero [simp]: "(~ (0::real) < x*x) = (x = 0)" -apply auto -apply (cut_tac x = x and y = 0 in linorder_less_linear) -apply (simp add: zero_less_mult_iff) -done - -lemma real_sqrt_abs2 [simp]: "sqrt(x*x) = \x\" -apply (subst power2_eq_square [symmetric]) -apply (rule real_sqrt_abs) -done - -lemma real_sqrt_pow2_gt_zero: "0 < x ==> 0 < (sqrt x)\" -by simp (* TODO: delete *) - -lemma real_sqrt_not_eq_zero: "0 < x ==> sqrt x \ 0" -by simp (* TODO: delete *) - -lemma real_inv_sqrt_pow2: "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x" -by (simp add: power_inverse [symmetric]) - -lemma real_sqrt_eq_zero_cancel: "[| 0 \ x; sqrt(x) = 0|] ==> x = 0" -by simp - -lemma real_sqrt_ge_one: "1 \ x ==> 1 \ sqrt x" -by simp - -lemma real_sqrt_two_gt_zero [simp]: "0 < sqrt 2" -by simp - -lemma real_sqrt_two_ge_zero [simp]: "0 \ sqrt 2" -by simp - -lemma real_sqrt_two_gt_one [simp]: "1 < sqrt 2" -by simp - -lemma sqrt_divide_self_eq: - assumes nneg: "0 \ x" - shows "sqrt x / x = inverse (sqrt x)" -proof cases - assume "x=0" thus ?thesis by simp -next - assume nz: "x\0" - hence pos: "0 0" by (simp add: divide_inverse nneg nz) - show "inverse (sqrt x) / (sqrt x / x) = 1" - by (simp add: divide_inverse mult_assoc [symmetric] - power2_eq_square [symmetric] real_inv_sqrt_pow2 pos nz) - qed -qed - -lemma real_divide_square_eq [simp]: "(((r::real) * a) / (r * r)) = a / r" -apply (simp add: divide_inverse) -apply (case_tac "r=0") -apply (auto simp add: mult_ac) -done - -lemma lemma_real_divide_sqrt_less: "0 < u ==> u / sqrt 2 < u" -by (simp add: divide_less_eq mult_compare_simps) - -lemma four_x_squared: - fixes x::real - shows "4 * x\ = (2 * x)\" -by (simp add: power2_eq_square) - -subsection {* Square Root of Sum of Squares *} - -lemma real_sqrt_mult_self_sum_ge_zero [simp]: "0 \ sqrt(x*x + y*y)" -by (rule real_sqrt_ge_zero [OF sum_squares_ge_zero]) - -lemma real_sqrt_sum_squares_ge_zero [simp]: "0 \ sqrt (x\ + y\)" -by simp - -declare real_sqrt_sum_squares_ge_zero [THEN abs_of_nonneg, simp] - -lemma real_sqrt_sum_squares_mult_ge_zero [simp]: - "0 \ sqrt ((x\ + y\)*(xa\ + ya\))" -by (auto intro!: real_sqrt_ge_zero simp add: zero_le_mult_iff) - -lemma real_sqrt_sum_squares_mult_squared_eq [simp]: - "sqrt ((x\ + y\) * (xa\ + ya\)) ^ 2 = (x\ + y\) * (xa\ + ya\)" -by (auto simp add: zero_le_mult_iff) - -lemma real_sqrt_sum_squares_eq_cancel: "sqrt (x\ + y\) = x \ y = 0" -by (drule_tac f = "%x. x\" in arg_cong, simp) - -lemma real_sqrt_sum_squares_eq_cancel2: "sqrt (x\ + y\) = y \ x = 0" -by (drule_tac f = "%x. x\" in arg_cong, simp) - -lemma real_sqrt_sum_squares_ge1 [simp]: "x \ sqrt (x\ + y\)" -by (rule power2_le_imp_le, simp_all) - -lemma real_sqrt_sum_squares_ge2 [simp]: "y \ sqrt (x\ + y\)" -by (rule power2_le_imp_le, simp_all) - -lemma real_sqrt_ge_abs1 [simp]: "\x\ \ sqrt (x\ + y\)" -by (rule power2_le_imp_le, simp_all) - -lemma real_sqrt_ge_abs2 [simp]: "\y\ \ sqrt (x\ + y\)" -by (rule power2_le_imp_le, simp_all) - -lemma le_real_sqrt_sumsq [simp]: "x \ sqrt (x * x + y * y)" -by (simp add: power2_eq_square [symmetric]) - -lemma power2_sum: - fixes x y :: "'a::{number_ring,recpower}" - shows "(x + y)\ = x\ + y\ + 2 * x * y" -by (simp add: ring_distribs power2_eq_square) - -lemma power2_diff: - fixes x y :: "'a::{number_ring,recpower}" - shows "(x - y)\ = x\ + y\ - 2 * x * y" -by (simp add: ring_distribs power2_eq_square) - -lemma real_sqrt_sum_squares_triangle_ineq: - "sqrt ((a + c)\ + (b + d)\) \ sqrt (a\ + b\) + sqrt (c\ + d\)" -apply (rule power2_le_imp_le, simp) -apply (simp add: power2_sum) -apply (simp only: mult_assoc right_distrib [symmetric]) -apply (rule mult_left_mono) -apply (rule power2_le_imp_le) -apply (simp add: power2_sum power_mult_distrib) -apply (simp add: ring_distribs) -apply (subgoal_tac "0 \ b\ * c\ + a\ * d\ - 2 * (a * c) * (b * d)", simp) -apply (rule_tac b="(a * d - b * c)\" in ord_le_eq_trans) -apply (rule zero_le_power2) -apply (simp add: power2_diff power_mult_distrib) -apply (simp add: mult_nonneg_nonneg) -apply simp -apply (simp add: add_increasing) -done - -lemma real_sqrt_sum_squares_less: - "\\x\ < u / sqrt 2; \y\ < u / sqrt 2\ \ sqrt (x\ + y\) < u" -apply (rule power2_less_imp_less, simp) -apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) -apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) -apply (simp add: power_divide) -apply (drule order_le_less_trans [OF abs_ge_zero]) -apply (simp add: zero_less_divide_iff) -done - -text{*Needed for the infinitely close relation over the nonstandard - complex numbers*} -lemma lemma_sqrt_hcomplex_capprox: - "[| 0 < u; x < u/2; y < u/2; 0 \ x; 0 \ y |] ==> sqrt (x\ + y\) < u" -apply (rule_tac y = "u/sqrt 2" in order_le_less_trans) -apply (erule_tac [2] lemma_real_divide_sqrt_less) -apply (rule power2_le_imp_le) -apply (auto simp add: real_0_le_divide_iff power_divide) -apply (rule_tac t = "u\" in real_sum_of_halves [THEN subst]) -apply (rule add_mono) -apply (auto simp add: four_x_squared simp del: realpow_Suc intro: power_mono) -done - -text "Legacy theorem names:" -lemmas real_root_pos2 = real_root_power_cancel -lemmas real_root_pos_pos = real_root_gt_zero [THEN order_less_imp_le] -lemmas real_root_pos_pos_le = real_root_ge_zero -lemmas real_sqrt_mult_distrib = real_sqrt_mult -lemmas real_sqrt_mult_distrib2 = real_sqrt_mult -lemmas real_sqrt_eq_zero_cancel_iff = real_sqrt_eq_0_iff - -(* needed for CauchysMeanTheorem.het_base from AFP *) -lemma real_root_pos: "0 < x \ root (Suc n) (x ^ (Suc n)) = x" -by (rule real_root_power_cancel [OF zero_less_Suc order_less_imp_le]) - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/SEQ.thy --- a/src/HOL/Hyperreal/SEQ.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Hyperreal/SEQ.thy Wed Dec 03 15:58:44 2008 +0100 @@ -9,7 +9,7 @@ header {* Sequences and Convergence *} theory SEQ -imports "../Real/RealVector" "../Real/RComplete" +imports "../Real/RealVector" "../RComplete" begin definition diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/Series.thy --- a/src/HOL/Hyperreal/Series.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,654 +0,0 @@ -(* Title : Series.thy - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - -Converted to Isar and polished by lcp -Converted to setsum and polished yet more by TNN -Additional contributions by Jeremy Avigad -*) - -header{*Finite Summation and Infinite Series*} - -theory Series -imports SEQ -begin - -definition - sums :: "(nat \ 'a::real_normed_vector) \ 'a \ bool" - (infixr "sums" 80) where - "f sums s = (%n. setsum f {0.. s" - -definition - summable :: "(nat \ 'a::real_normed_vector) \ bool" where - "summable f = (\s. f sums s)" - -definition - suminf :: "(nat \ 'a::real_normed_vector) \ 'a" where - "suminf f = (THE s. f sums s)" - -syntax - "_suminf" :: "idt \ 'a \ 'a" ("\_. _" [0, 10] 10) -translations - "\i. b" == "CONST suminf (%i. b)" - - -lemma sumr_diff_mult_const: - "setsum f {0.. f(p) \ K) - \ setsum f {0.. real n * K" -using setsum_bounded[where A = "{0..i=0..<2*n. (-1) ^ Suc i) = (0::real)" -by (induct "n", auto) - -(* FIXME this is an awful lemma! *) -lemma sumr_one_lb_realpow_zero [simp]: - "(\n=Suc 0..m=0..m=0.. 'a::ab_group_add" - shows "(\m=0..f. (\m=0..n f. setsum f {0::nat..m=0.. - (\n=Suc 0 ..< Suc n. if even(n) then 0 else - ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n = - (\n=0.. summable f" -by (simp add: sums_def summable_def, blast) - -lemma summable_sums: "summable f ==> f sums (suminf f)" -apply (simp add: summable_def suminf_def sums_def) -apply (blast intro: theI LIMSEQ_unique) -done - -lemma summable_sumr_LIMSEQ_suminf: - "summable f ==> (%n. setsum f {0.. (suminf f)" -by (rule summable_sums [unfolded sums_def]) - -(*------------------- - sum is unique - ------------------*) -lemma sums_unique: "f sums s ==> (s = suminf f)" -apply (frule sums_summable [THEN summable_sums]) -apply (auto intro!: LIMSEQ_unique simp add: sums_def) -done - -lemma sums_split_initial_segment: "f sums s ==> - (%n. f(n + k)) sums (s - (SUM i = 0..< k. f i))" - apply (unfold sums_def); - apply (simp add: sumr_offset); - apply (rule LIMSEQ_diff_const) - apply (rule LIMSEQ_ignore_initial_segment) - apply assumption -done - -lemma summable_ignore_initial_segment: "summable f ==> - summable (%n. f(n + k))" - apply (unfold summable_def) - apply (auto intro: sums_split_initial_segment) -done - -lemma suminf_minus_initial_segment: "summable f ==> - suminf f = s ==> suminf (%n. f(n + k)) = s - (SUM i = 0..< k. f i)" - apply (frule summable_ignore_initial_segment) - apply (rule sums_unique [THEN sym]) - apply (frule summable_sums) - apply (rule sums_split_initial_segment) - apply auto -done - -lemma suminf_split_initial_segment: "summable f ==> - suminf f = (SUM i = 0..< k. f i) + suminf (%n. f(n + k))" -by (auto simp add: suminf_minus_initial_segment) - -lemma series_zero: - "(\m. n \ m --> f(m) = 0) ==> f sums (setsum f {0..n. 0) sums 0" -unfolding sums_def by (simp add: LIMSEQ_const) - -lemma summable_zero: "summable (\n. 0)" -by (rule sums_zero [THEN sums_summable]) - -lemma suminf_zero: "suminf (\n. 0) = 0" -by (rule sums_zero [THEN sums_unique, symmetric]) - -lemma (in bounded_linear) sums: - "(\n. X n) sums a \ (\n. f (X n)) sums (f a)" -unfolding sums_def by (drule LIMSEQ, simp only: setsum) - -lemma (in bounded_linear) summable: - "summable (\n. X n) \ summable (\n. f (X n))" -unfolding summable_def by (auto intro: sums) - -lemma (in bounded_linear) suminf: - "summable (\n. X n) \ f (\n. X n) = (\n. f (X n))" -by (intro sums_unique sums summable_sums) - -lemma sums_mult: - fixes c :: "'a::real_normed_algebra" - shows "f sums a \ (\n. c * f n) sums (c * a)" -by (rule mult_right.sums) - -lemma summable_mult: - fixes c :: "'a::real_normed_algebra" - shows "summable f \ summable (%n. c * f n)" -by (rule mult_right.summable) - -lemma suminf_mult: - fixes c :: "'a::real_normed_algebra" - shows "summable f \ suminf (\n. c * f n) = c * suminf f"; -by (rule mult_right.suminf [symmetric]) - -lemma sums_mult2: - fixes c :: "'a::real_normed_algebra" - shows "f sums a \ (\n. f n * c) sums (a * c)" -by (rule mult_left.sums) - -lemma summable_mult2: - fixes c :: "'a::real_normed_algebra" - shows "summable f \ summable (\n. f n * c)" -by (rule mult_left.summable) - -lemma suminf_mult2: - fixes c :: "'a::real_normed_algebra" - shows "summable f \ suminf f * c = (\n. f n * c)" -by (rule mult_left.suminf) - -lemma sums_divide: - fixes c :: "'a::real_normed_field" - shows "f sums a \ (\n. f n / c) sums (a / c)" -by (rule divide.sums) - -lemma summable_divide: - fixes c :: "'a::real_normed_field" - shows "summable f \ summable (\n. f n / c)" -by (rule divide.summable) - -lemma suminf_divide: - fixes c :: "'a::real_normed_field" - shows "summable f \ suminf (\n. f n / c) = suminf f / c" -by (rule divide.suminf [symmetric]) - -lemma sums_add: "\X sums a; Y sums b\ \ (\n. X n + Y n) sums (a + b)" -unfolding sums_def by (simp add: setsum_addf LIMSEQ_add) - -lemma summable_add: "\summable X; summable Y\ \ summable (\n. X n + Y n)" -unfolding summable_def by (auto intro: sums_add) - -lemma suminf_add: - "\summable X; summable Y\ \ suminf X + suminf Y = (\n. X n + Y n)" -by (intro sums_unique sums_add summable_sums) - -lemma sums_diff: "\X sums a; Y sums b\ \ (\n. X n - Y n) sums (a - b)" -unfolding sums_def by (simp add: setsum_subtractf LIMSEQ_diff) - -lemma summable_diff: "\summable X; summable Y\ \ summable (\n. X n - Y n)" -unfolding summable_def by (auto intro: sums_diff) - -lemma suminf_diff: - "\summable X; summable Y\ \ suminf X - suminf Y = (\n. X n - Y n)" -by (intro sums_unique sums_diff summable_sums) - -lemma sums_minus: "X sums a ==> (\n. - X n) sums (- a)" -unfolding sums_def by (simp add: setsum_negf LIMSEQ_minus) - -lemma summable_minus: "summable X \ summable (\n. - X n)" -unfolding summable_def by (auto intro: sums_minus) - -lemma suminf_minus: "summable X \ (\n. - X n) = - (\n. X n)" -by (intro sums_unique [symmetric] sums_minus summable_sums) - -lemma sums_group: - "[|summable f; 0 < k |] ==> (%n. setsum f {n*k.. real" - shows "\summable f; \m\n. 0 \ f m\ \ setsum f {0.. suminf f" -apply (drule summable_sums) -apply (simp add: sums_def) -apply (cut_tac k = "setsum f {0.. real" - shows "\summable f; \m\n. 0 < f m\ \ setsum f {0.. real" - shows "\summable f; \n. 0 < f n\ \ 0 < suminf f" -by (drule_tac n="0" in series_pos_less, simp_all) - -lemma suminf_ge_zero: - fixes f :: "nat \ real" - shows "\summable f; \n. 0 \ f n\ \ 0 \ suminf f" -by (drule_tac n="0" in series_pos_le, simp_all) - -lemma sumr_pos_lt_pair: - fixes f :: "nat \ real" - shows "\summable f; - \d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc(Suc 0) * d) + 1))\ - \ setsum f {0.. (\n. x ^ n) sums (1 / (1 - x))" -proof - - assume less_1: "norm x < 1" - hence neq_1: "x \ 1" by auto - hence neq_0: "x - 1 \ 0" by simp - from less_1 have lim_0: "(\n. x ^ n) ----> 0" - by (rule LIMSEQ_power_zero) - hence "(\n. x ^ n / (x - 1) - 1 / (x - 1)) ----> 0 / (x - 1) - 1 / (x - 1)" - using neq_0 by (intro LIMSEQ_divide LIMSEQ_diff LIMSEQ_const) - hence "(\n. (x ^ n - 1) / (x - 1)) ----> 1 / (1 - x)" - by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib) - thus "(\n. x ^ n) sums (1 / (1 - x))" - by (simp add: sums_def geometric_sum neq_1) -qed - -lemma summable_geometric: - fixes x :: "'a::{real_normed_field,recpower}" - shows "norm x < 1 \ summable (\n. x ^ n)" -by (rule geometric_sums [THEN sums_summable]) - -text{*Cauchy-type criterion for convergence of series (c.f. Harrison)*} - -lemma summable_convergent_sumr_iff: - "summable f = convergent (%n. setsum f {0.. f ----> 0" -apply (drule summable_convergent_sumr_iff [THEN iffD1]) -apply (drule convergent_Cauchy) -apply (simp only: Cauchy_def LIMSEQ_def, safe) -apply (drule_tac x="r" in spec, safe) -apply (rule_tac x="M" in exI, safe) -apply (drule_tac x="Suc n" in spec, simp) -apply (drule_tac x="n" in spec, simp) -done - -lemma summable_Cauchy: - "summable (f::nat \ 'a::banach) = - (\e > 0. \N. \m \ N. \n. norm (setsum f {m.. 'b::real_normed_vector" - shows "norm (setsum f A) \ (\i\A. norm (f i))" -apply (case_tac "finite A") -apply (erule finite_induct) -apply simp -apply simp -apply (erule order_trans [OF norm_triangle_ineq add_left_mono]) -apply simp -done - -lemma summable_comparison_test: - fixes f :: "nat \ 'a::banach" - shows "\\N. \n\N. norm (f n) \ g n; summable g\ \ summable f" -apply (simp add: summable_Cauchy, safe) -apply (drule_tac x="e" in spec, safe) -apply (rule_tac x = "N + Na" in exI, safe) -apply (rotate_tac 2) -apply (drule_tac x = m in spec) -apply (auto, rotate_tac 2, drule_tac x = n in spec) -apply (rule_tac y = "\k=m.. 'a::banach" - shows "\\N. \n\N. norm (f n) \ g n; summable g\ - \ summable (\n. norm (f n))" -apply (rule summable_comparison_test) -apply (auto) -done - -lemma summable_rabs_comparison_test: - fixes f :: "nat \ real" - shows "\\N. \n\N. \f n\ \ g n; summable g\ \ summable (\n. \f n\)" -apply (rule summable_comparison_test) -apply (auto) -done - -text{*Summability of geometric series for real algebras*} - -lemma complete_algebra_summable_geometric: - fixes x :: "'a::{real_normed_algebra_1,banach,recpower}" - shows "norm x < 1 \ summable (\n. x ^ n)" -proof (rule summable_comparison_test) - show "\N. \n\N. norm (x ^ n) \ norm x ^ n" - by (simp add: norm_power_ineq) - show "norm x < 1 \ summable (\n. norm x ^ n)" - by (simp add: summable_geometric) -qed - -text{*Limit comparison property for series (c.f. jrh)*} - -lemma summable_le: - fixes f g :: "nat \ real" - shows "\\n. f n \ g n; summable f; summable g\ \ suminf f \ suminf g" -apply (drule summable_sums)+ -apply (simp only: sums_def, erule (1) LIMSEQ_le) -apply (rule exI) -apply (auto intro!: setsum_mono) -done - -lemma summable_le2: - fixes f g :: "nat \ real" - shows "\\n. \f n\ \ g n; summable g\ \ summable f \ suminf f \ suminf g" -apply (subgoal_tac "summable f") -apply (auto intro!: summable_le) -apply (simp add: abs_le_iff) -apply (rule_tac g="g" in summable_comparison_test, simp_all) -done - -(* specialisation for the common 0 case *) -lemma suminf_0_le: - fixes f::"nat\real" - assumes gt0: "\n. 0 \ f n" and sm: "summable f" - shows "0 \ suminf f" -proof - - let ?g = "(\n. (0::real))" - from gt0 have "\n. ?g n \ f n" by simp - moreover have "summable ?g" by (rule summable_zero) - moreover from sm have "summable f" . - ultimately have "suminf ?g \ suminf f" by (rule summable_le) - then show "0 \ suminf f" by (simp add: suminf_zero) -qed - - -text{*Absolute convergence imples normal convergence*} -lemma summable_norm_cancel: - fixes f :: "nat \ 'a::banach" - shows "summable (\n. norm (f n)) \ summable f" -apply (simp only: summable_Cauchy, safe) -apply (drule_tac x="e" in spec, safe) -apply (rule_tac x="N" in exI, safe) -apply (drule_tac x="m" in spec, safe) -apply (rule order_le_less_trans [OF norm_setsum]) -apply (rule order_le_less_trans [OF abs_ge_self]) -apply simp -done - -lemma summable_rabs_cancel: - fixes f :: "nat \ real" - shows "summable (\n. \f n\) \ summable f" -by (rule summable_norm_cancel, simp) - -text{*Absolute convergence of series*} -lemma summable_norm: - fixes f :: "nat \ 'a::banach" - shows "summable (\n. norm (f n)) \ norm (suminf f) \ (\n. norm (f n))" -by (auto intro: LIMSEQ_le LIMSEQ_norm summable_norm_cancel - summable_sumr_LIMSEQ_suminf norm_setsum) - -lemma summable_rabs: - fixes f :: "nat \ real" - shows "summable (\n. \f n\) \ \suminf f\ \ (\n. \f n\)" -by (fold real_norm_def, rule summable_norm) - -subsection{* The Ratio Test*} - -lemma norm_ratiotest_lemma: - fixes x y :: "'a::real_normed_vector" - shows "\c \ 0; norm x \ c * norm y\ \ x = 0" -apply (subgoal_tac "norm x \ 0", simp) -apply (erule order_trans) -apply (simp add: mult_le_0_iff) -done - -lemma rabs_ratiotest_lemma: "[| c \ 0; abs x \ c * abs y |] ==> x = (0::real)" -by (erule norm_ratiotest_lemma, simp) - -lemma le_Suc_ex: "(k::nat) \ l ==> (\n. l = k + n)" -apply (drule le_imp_less_or_eq) -apply (auto dest: less_imp_Suc_add) -done - -lemma le_Suc_ex_iff: "((k::nat) \ l) = (\n. l = k + n)" -by (auto simp add: le_Suc_ex) - -(*All this trouble just to get 0 'a::banach" - shows "\\n\N. norm (f (Suc n)) \ c * norm (f n)\ \ 0 < c \ summable f" -apply (simp (no_asm) add: linorder_not_le [symmetric]) -apply (simp add: summable_Cauchy) -apply (safe, subgoal_tac "\n. N < n --> f (n) = 0") - prefer 2 - apply clarify - apply(erule_tac x = "n - 1" in allE) - apply (simp add:diff_Suc split:nat.splits) - apply (blast intro: norm_ratiotest_lemma) -apply (rule_tac x = "Suc N" in exI, clarify) -apply(simp cong:setsum_ivl_cong) -done - -lemma ratio_test: - fixes f :: "nat \ 'a::banach" - shows "\c < 1; \n\N. norm (f (Suc n)) \ c * norm (f n)\ \ summable f" -apply (frule ratio_test_lemma2, auto) -apply (rule_tac g = "%n. (norm (f N) / (c ^ N))*c ^ n" - in summable_comparison_test) -apply (rule_tac x = N in exI, safe) -apply (drule le_Suc_ex_iff [THEN iffD1]) -apply (auto simp add: power_add field_power_not_zero) -apply (induct_tac "na", auto) -apply (rule_tac y = "c * norm (f (N + n))" in order_trans) -apply (auto intro: mult_right_mono simp add: summable_def) -apply (simp add: mult_ac) -apply (rule_tac x = "norm (f N) * (1/ (1 - c)) / (c ^ N)" in exI) -apply (rule sums_divide) -apply (rule sums_mult) -apply (auto intro!: geometric_sums) -done - -subsection {* Cauchy Product Formula *} - -(* Proof based on Analysis WebNotes: Chapter 07, Class 41 -http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm *) - -lemma setsum_triangle_reindex: - fixes n :: nat - shows "(\(i,j)\{(i,j). i+j < n}. f i j) = (\k=0..i=0..k. f i (k - i))" -proof - - have "(\(i, j)\{(i, j). i + j < n}. f i j) = - (\(k, i)\(SIGMA k:{0..(k,i). (i, k - i)) (SIGMA k:{0..(k,i). (i, k - i)) ` (SIGMA k:{0..a. (\(k, i). f i (k - i)) a = split f ((\(k, i). (i, k - i)) a)" - by clarify - qed - thus ?thesis by (simp add: setsum_Sigma) -qed - -lemma Cauchy_product_sums: - fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" - assumes a: "summable (\k. norm (a k))" - assumes b: "summable (\k. norm (b k))" - shows "(\k. \i=0..k. a i * b (k - i)) sums ((\k. a k) * (\k. b k))" -proof - - let ?S1 = "\n::nat. {0.. {0..m n. m \ n \ ?S1 m \ ?S1 n" by auto - have S2_le_S1: "\n. ?S2 n \ ?S1 n" by auto - have S1_le_S2: "\n. ?S1 (n div 2) \ ?S2 n" by auto - have finite_S1: "\n. finite (?S1 n)" by simp - with S2_le_S1 have finite_S2: "\n. finite (?S2 n)" by (rule finite_subset) - - let ?g = "\(i,j). a i * b j" - let ?f = "\(i,j). norm (a i) * norm (b j)" - have f_nonneg: "\x. 0 \ ?f x" - by (auto simp add: mult_nonneg_nonneg) - hence norm_setsum_f: "\A. norm (setsum ?f A) = setsum ?f A" - unfolding real_norm_def - by (simp only: abs_of_nonneg setsum_nonneg [rule_format]) - - have "(\n. (\k=0..k=0.. (\k. a k) * (\k. b k)" - by (intro LIMSEQ_mult summable_sumr_LIMSEQ_suminf - summable_norm_cancel [OF a] summable_norm_cancel [OF b]) - hence 1: "(\n. setsum ?g (?S1 n)) ----> (\k. a k) * (\k. b k)" - by (simp only: setsum_product setsum_Sigma [rule_format] - finite_atLeastLessThan) - - have "(\n. (\k=0..k=0.. (\k. norm (a k)) * (\k. norm (b k))" - using a b by (intro LIMSEQ_mult summable_sumr_LIMSEQ_suminf) - hence "(\n. setsum ?f (?S1 n)) ----> (\k. norm (a k)) * (\k. norm (b k))" - by (simp only: setsum_product setsum_Sigma [rule_format] - finite_atLeastLessThan) - hence "convergent (\n. setsum ?f (?S1 n))" - by (rule convergentI) - hence Cauchy: "Cauchy (\n. setsum ?f (?S1 n))" - by (rule convergent_Cauchy) - have "Zseq (\n. setsum ?f (?S1 n - ?S2 n))" - proof (rule ZseqI, simp only: norm_setsum_f) - fix r :: real - assume r: "0 < r" - from CauchyD [OF Cauchy r] obtain N - where "\m\N. \n\N. norm (setsum ?f (?S1 m) - setsum ?f (?S1 n)) < r" .. - hence "\m n. \N \ n; n \ m\ \ norm (setsum ?f (?S1 m - ?S1 n)) < r" - by (simp only: setsum_diff finite_S1 S1_mono) - hence N: "\m n. \N \ n; n \ m\ \ setsum ?f (?S1 m - ?S1 n) < r" - by (simp only: norm_setsum_f) - show "\N. \n\N. setsum ?f (?S1 n - ?S2 n) < r" - proof (intro exI allI impI) - fix n assume "2 * N \ n" - hence n: "N \ n div 2" by simp - have "setsum ?f (?S1 n - ?S2 n) \ setsum ?f (?S1 n - ?S1 (n div 2))" - by (intro setsum_mono2 finite_Diff finite_S1 f_nonneg - Diff_mono subset_refl S1_le_S2) - also have "\ < r" - using n div_le_dividend by (rule N) - finally show "setsum ?f (?S1 n - ?S2 n) < r" . - qed - qed - hence "Zseq (\n. setsum ?g (?S1 n - ?S2 n))" - apply (rule Zseq_le [rule_format]) - apply (simp only: norm_setsum_f) - apply (rule order_trans [OF norm_setsum setsum_mono]) - apply (auto simp add: norm_mult_ineq) - done - hence 2: "(\n. setsum ?g (?S1 n) - setsum ?g (?S2 n)) ----> 0" - by (simp only: LIMSEQ_Zseq_iff setsum_diff finite_S1 S2_le_S1 diff_0_right) - - with 1 have "(\n. setsum ?g (?S2 n)) ----> (\k. a k) * (\k. b k)" - by (rule LIMSEQ_diff_approach_zero2) - thus ?thesis by (simp only: sums_def setsum_triangle_reindex) -qed - -lemma Cauchy_product: - fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" - assumes a: "summable (\k. norm (a k))" - assumes b: "summable (\k. norm (b k))" - shows "(\k. a k) * (\k. b k) = (\k. \i=0..k. a i * b (k - i))" -using a b -by (rule Cauchy_product_sums [THEN sums_unique]) - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Hyperreal/Taylor.thy --- a/src/HOL/Hyperreal/Taylor.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,134 +0,0 @@ -(* Title: HOL/Hyperreal/Taylor.thy - ID: $Id$ - Author: Lukas Bulwahn, Bernhard Haeupler, Technische Universitaet Muenchen -*) - -header {* Taylor series *} - -theory Taylor -imports MacLaurin -begin - -text {* -We use MacLaurin and the translation of the expansion point @{text c} to @{text 0} -to prove Taylor's theorem. -*} - -lemma taylor_up: - assumes INIT: "n>0" "diff 0 = f" - and DERIV: "(\ m t. m < n & a \ t & t \ b \ DERIV (diff m) t :> (diff (Suc m) t))" - and INTERV: "a \ c" "c < b" - shows "\ t. c < t & t < b & - f b = setsum (%m. (diff m c / real (fact m)) * (b - c)^m) {0..0" "((\m x. diff m (x + c)) 0) = (\x. f (x + c))" by auto - moreover - have "ALL m t. m < n & 0 <= t & t <= b - c --> DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" - proof (intro strip) - fix m t - assume "m < n & 0 <= t & t <= b - c" - with DERIV and INTERV have "DERIV (diff m) (t + c) :> diff (Suc m) (t + c)" by auto - moreover - from DERIV_ident and DERIV_const have "DERIV (%x. x + c) t :> 1+0" by (rule DERIV_add) - ultimately have "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1+0)" - by (rule DERIV_chain2) - thus "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" by simp - qed - ultimately - have EX:"EX t>0. t < b - c & - f (b - c + c) = (SUM m = 0..m = 0.. f b = (\m = 0..0" "diff 0 = f" - and DERIV: "(\ m t. m < n & a \ t & t \ b \ DERIV (diff m) t :> (diff (Suc m) t))" - and INTERV: "a < c" "c \ b" - shows "\ t. a < t & t < c & - f a = setsum (% m. (diff m c / real (fact m)) * (a - c)^m) {0..0" "((\m x. diff m (x + c)) 0) = (\x. f (x + c))" by auto - moreover - have "ALL m t. m < n & a-c <= t & t <= 0 --> DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" - proof (rule allI impI)+ - fix m t - assume "m < n & a-c <= t & t <= 0" - with DERIV and INTERV have "DERIV (diff m) (t + c) :> diff (Suc m) (t + c)" by auto - moreover - from DERIV_ident and DERIV_const have "DERIV (%x. x + c) t :> 1+0" by (rule DERIV_add) - ultimately have "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1+0)" by (rule DERIV_chain2) - thus "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" by simp - qed - ultimately - have EX: "EX t>a - c. t < 0 & - f (a - c + c) = (SUM m = 0.. f a = (\m = 0..0" "diff 0 = f" - and DERIV: "(\ m t. m < n & a \ t & t \ b \ DERIV (diff m) t :> (diff (Suc m) t))" - and INTERV: "a \ c " "c \ b" "a \ x" "x \ b" "x \ c" - shows "\ t. (if xm t. m < n \ x \ t \ t \ b \ DERIV (diff m) t :> diff (Suc m) t" - by fastsimp - moreover note True - moreover from INTERV have "c \ b" by simp - ultimately have EX: "\t>x. t < c \ f x = - (\m = 0..m t. m < n \ a \ t \ t \ x \ DERIV (diff m) t :> diff (Suc m) t" - by fastsimp - moreover from INTERV have "a \ c" by arith - moreover from False and INTERV have "c < x" by arith - ultimately have EX: "\t>c. t < x \ f x = - (\m = 0.. n \ y ^ (Suc n - p) = (y ^ (n - p)) * y" -proof - - assume "p \ n" - hence "Suc n - p = Suc (n - p)" by (rule Suc_diff_le) - thus ?thesis by (simp add: power_Suc power_commutes) -qed - -lemma lemma_realpow_diff_sumr: - fixes y :: "'a::{recpower,comm_semiring_0}" shows - "(\p=0..p=0..p=0..p=0..p=0..z\ < \x\"}.*} - -lemma powser_insidea: - fixes x z :: "'a::{real_normed_field,banach,recpower}" - assumes 1: "summable (\n. f n * x ^ n)" - assumes 2: "norm z < norm x" - shows "summable (\n. norm (f n * z ^ n))" -proof - - from 2 have x_neq_0: "x \ 0" by clarsimp - from 1 have "(\n. f n * x ^ n) ----> 0" - by (rule summable_LIMSEQ_zero) - hence "convergent (\n. f n * x ^ n)" - by (rule convergentI) - hence "Cauchy (\n. f n * x ^ n)" - by (simp add: Cauchy_convergent_iff) - hence "Bseq (\n. f n * x ^ n)" - by (rule Cauchy_Bseq) - then obtain K where 3: "0 < K" and 4: "\n. norm (f n * x ^ n) \ K" - by (simp add: Bseq_def, safe) - have "\N. \n\N. norm (norm (f n * z ^ n)) \ - K * norm (z ^ n) * inverse (norm (x ^ n))" - proof (intro exI allI impI) - fix n::nat assume "0 \ n" - have "norm (norm (f n * z ^ n)) * norm (x ^ n) = - norm (f n * x ^ n) * norm (z ^ n)" - by (simp add: norm_mult abs_mult) - also have "\ \ K * norm (z ^ n)" - by (simp only: mult_right_mono 4 norm_ge_zero) - also have "\ = K * norm (z ^ n) * (inverse (norm (x ^ n)) * norm (x ^ n))" - by (simp add: x_neq_0) - also have "\ = K * norm (z ^ n) * inverse (norm (x ^ n)) * norm (x ^ n)" - by (simp only: mult_assoc) - finally show "norm (norm (f n * z ^ n)) \ - K * norm (z ^ n) * inverse (norm (x ^ n))" - by (simp add: mult_le_cancel_right x_neq_0) - qed - moreover have "summable (\n. K * norm (z ^ n) * inverse (norm (x ^ n)))" - proof - - from 2 have "norm (norm (z * inverse x)) < 1" - using x_neq_0 - by (simp add: nonzero_norm_divide divide_inverse [symmetric]) - hence "summable (\n. norm (z * inverse x) ^ n)" - by (rule summable_geometric) - hence "summable (\n. K * norm (z * inverse x) ^ n)" - by (rule summable_mult) - thus "summable (\n. K * norm (z ^ n) * inverse (norm (x ^ n)))" - using x_neq_0 - by (simp add: norm_mult nonzero_norm_inverse power_mult_distrib - power_inverse norm_power mult_assoc) - qed - ultimately show "summable (\n. norm (f n * z ^ n))" - by (rule summable_comparison_test) -qed - -lemma powser_inside: - fixes f :: "nat \ 'a::{real_normed_field,banach,recpower}" shows - "[| summable (%n. f(n) * (x ^ n)); norm z < norm x |] - ==> summable (%n. f(n) * (z ^ n))" -by (rule powser_insidea [THEN summable_norm_cancel]) - - -subsection{*Term-by-Term Differentiability of Power Series*} - -definition - diffs :: "(nat => 'a::ring_1) => nat => 'a" where - "diffs c = (%n. of_nat (Suc n) * c(Suc n))" - -text{*Lemma about distributing negation over it*} -lemma diffs_minus: "diffs (%n. - c n) = (%n. - diffs c n)" -by (simp add: diffs_def) - -text{*Show that we can shift the terms down one*} -lemma lemma_diffs: - "(\n=0..n=0..n=0..n=0.. - (%n. of_nat n * c(n) * (x ^ (n - Suc 0))) sums - (\n. (diffs c)(n) * (x ^ n))" -apply (subgoal_tac " (%n. of_nat n * c (n) * (x ^ (n - Suc 0))) ----> 0") -apply (rule_tac [2] LIMSEQ_imp_Suc) -apply (drule summable_sums) -apply (auto simp add: sums_def) -apply (drule_tac X="(\n. \n = 0..p=0..p=0.. (\d. n = m + d + Suc 0)" -by (simp add: less_iff_Suc_add) - -lemma sumdiff: "a + b - (c + d) = a - c + b - (d::real)" -by arith - -lemma sumr_diff_mult_const2: - "setsum f {0..i = 0.. 0" shows - "((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0) = - h * (\p=0..< n - Suc 0. \q=0..< n - Suc 0 - p. - (z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs") -apply (subgoal_tac "h * ?lhs = h * ?rhs", simp add: h) -apply (simp add: right_diff_distrib diff_divide_distrib h) -apply (simp add: mult_assoc [symmetric]) -apply (cases "n", simp) -apply (simp add: lemma_realpow_diff_sumr2 h - right_diff_distrib [symmetric] mult_assoc - del: realpow_Suc setsum_op_ivl_Suc of_nat_Suc) -apply (subst lemma_realpow_rev_sumr) -apply (subst sumr_diff_mult_const2) -apply simp -apply (simp only: lemma_termdiff1 setsum_right_distrib) -apply (rule setsum_cong [OF refl]) -apply (simp add: diff_minus [symmetric] less_iff_Suc_add) -apply (clarify) -apply (simp add: setsum_right_distrib lemma_realpow_diff_sumr2 mult_ac - del: setsum_op_ivl_Suc realpow_Suc) -apply (subst mult_assoc [symmetric], subst power_add [symmetric]) -apply (simp add: mult_ac) -done - -lemma real_setsum_nat_ivl_bounded2: - fixes K :: "'a::ordered_semidom" - assumes f: "\p::nat. p < n \ f p \ K" - assumes K: "0 \ K" - shows "setsum f {0.. of_nat n * K" -apply (rule order_trans [OF setsum_mono]) -apply (rule f, simp) -apply (simp add: mult_right_mono K) -done - -lemma lemma_termdiff3: - fixes h z :: "'a::{real_normed_field,recpower}" - assumes 1: "h \ 0" - assumes 2: "norm z \ K" - assumes 3: "norm (z + h) \ K" - shows "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) - \ of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" -proof - - have "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) = - norm (\p = 0..q = 0.. \ of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2)) * norm h" - proof (rule mult_right_mono [OF _ norm_ge_zero]) - from norm_ge_zero 2 have K: "0 \ K" by (rule order_trans) - have le_Kn: "\i j n. i + j = n \ norm ((z + h) ^ i * z ^ j) \ K ^ n" - apply (erule subst) - apply (simp only: norm_mult norm_power power_add) - apply (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K) - done - show "norm (\p = 0..q = 0.. of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2))" - apply (intro - order_trans [OF norm_setsum] - real_setsum_nat_ivl_bounded2 - mult_nonneg_nonneg - zero_le_imp_of_nat - zero_le_power K) - apply (rule le_Kn, simp) - done - qed - also have "\ = of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" - by (simp only: mult_assoc) - finally show ?thesis . -qed - -lemma lemma_termdiff4: - fixes f :: "'a::{real_normed_field,recpower} \ - 'b::real_normed_vector" - assumes k: "0 < (k::real)" - assumes le: "\h. \h \ 0; norm h < k\ \ norm (f h) \ K * norm h" - shows "f -- 0 --> 0" -proof (simp add: LIM_def, safe) - fix r::real assume r: "0 < r" - have zero_le_K: "0 \ K" - apply (cut_tac k) - apply (cut_tac h="of_real (k/2)" in le, simp) - apply (simp del: of_real_divide) - apply (drule order_trans [OF norm_ge_zero]) - apply (simp add: zero_le_mult_iff) - done - show "\s. 0 < s \ (\x. x \ 0 \ norm x < s \ norm (f x) < r)" - proof (cases) - assume "K = 0" - with k r le have "0 < k \ (\x. x \ 0 \ norm x < k \ norm (f x) < r)" - by simp - thus "\s. 0 < s \ (\x. x \ 0 \ norm x < s \ norm (f x) < r)" .. - next - assume K_neq_zero: "K \ 0" - with zero_le_K have K: "0 < K" by simp - show "\s. 0 < s \ (\x. x \ 0 \ norm x < s \ norm (f x) < r)" - proof (rule exI, safe) - from k r K show "0 < min k (r * inverse K / 2)" - by (simp add: mult_pos_pos positive_imp_inverse_positive) - next - fix x::'a - assume x1: "x \ 0" and x2: "norm x < min k (r * inverse K / 2)" - from x2 have x3: "norm x < k" and x4: "norm x < r * inverse K / 2" - by simp_all - from x1 x3 le have "norm (f x) \ K * norm x" by simp - also from x4 K have "K * norm x < K * (r * inverse K / 2)" - by (rule mult_strict_left_mono) - also have "\ = r / 2" - using K_neq_zero by simp - also have "r / 2 < r" - using r by simp - finally show "norm (f x) < r" . - qed - qed -qed - -lemma lemma_termdiff5: - fixes g :: "'a::{recpower,real_normed_field} \ - nat \ 'b::banach" - assumes k: "0 < (k::real)" - assumes f: "summable f" - assumes le: "\h n. \h \ 0; norm h < k\ \ norm (g h n) \ f n * norm h" - shows "(\h. suminf (g h)) -- 0 --> 0" -proof (rule lemma_termdiff4 [OF k]) - fix h::'a assume "h \ 0" and "norm h < k" - hence A: "\n. norm (g h n) \ f n * norm h" - by (simp add: le) - hence "\N. \n\N. norm (norm (g h n)) \ f n * norm h" - by simp - moreover from f have B: "summable (\n. f n * norm h)" - by (rule summable_mult2) - ultimately have C: "summable (\n. norm (g h n))" - by (rule summable_comparison_test) - hence "norm (suminf (g h)) \ (\n. norm (g h n))" - by (rule summable_norm) - also from A C B have "(\n. norm (g h n)) \ (\n. f n * norm h)" - by (rule summable_le) - also from f have "(\n. f n * norm h) = suminf f * norm h" - by (rule suminf_mult2 [symmetric]) - finally show "norm (suminf (g h)) \ suminf f * norm h" . -qed - - -text{* FIXME: Long proofs*} - -lemma termdiffs_aux: - fixes x :: "'a::{recpower,real_normed_field,banach}" - assumes 1: "summable (\n. diffs (diffs c) n * K ^ n)" - assumes 2: "norm x < norm K" - shows "(\h. \n. c n * (((x + h) ^ n - x ^ n) / h - - of_nat n * x ^ (n - Suc 0))) -- 0 --> 0" -proof - - from dense [OF 2] - obtain r where r1: "norm x < r" and r2: "r < norm K" by fast - from norm_ge_zero r1 have r: "0 < r" - by (rule order_le_less_trans) - hence r_neq_0: "r \ 0" by simp - show ?thesis - proof (rule lemma_termdiff5) - show "0 < r - norm x" using r1 by simp - next - from r r2 have "norm (of_real r::'a) < norm K" - by simp - with 1 have "summable (\n. norm (diffs (diffs c) n * (of_real r ^ n)))" - by (rule powser_insidea) - hence "summable (\n. diffs (diffs (\n. norm (c n))) n * r ^ n)" - using r - by (simp add: diffs_def norm_mult norm_power del: of_nat_Suc) - hence "summable (\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0))" - by (rule diffs_equiv [THEN sums_summable]) - also have "(\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0)) - = (\n. diffs (%m. of_nat (m - Suc 0) * norm (c m) * inverse r) n * (r ^ n))" - apply (rule ext) - apply (simp add: diffs_def) - apply (case_tac n, simp_all add: r_neq_0) - done - finally have "summable - (\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0))" - by (rule diffs_equiv [THEN sums_summable]) - also have - "(\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * - r ^ (n - Suc 0)) = - (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" - apply (rule ext) - apply (case_tac "n", simp) - apply (case_tac "nat", simp) - apply (simp add: r_neq_0) - done - finally show - "summable (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" . - next - fix h::'a and n::nat - assume h: "h \ 0" - assume "norm h < r - norm x" - hence "norm x + norm h < r" by simp - with norm_triangle_ineq have xh: "norm (x + h) < r" - by (rule order_le_less_trans) - show "norm (c n * (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0))) - \ norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2) * norm h" - apply (simp only: norm_mult mult_assoc) - apply (rule mult_left_mono [OF _ norm_ge_zero]) - apply (simp (no_asm) add: mult_assoc [symmetric]) - apply (rule lemma_termdiff3) - apply (rule h) - apply (rule r1 [THEN order_less_imp_le]) - apply (rule xh [THEN order_less_imp_le]) - done - qed -qed - -lemma termdiffs: - fixes K x :: "'a::{recpower,real_normed_field,banach}" - assumes 1: "summable (\n. c n * K ^ n)" - assumes 2: "summable (\n. (diffs c) n * K ^ n)" - assumes 3: "summable (\n. (diffs (diffs c)) n * K ^ n)" - assumes 4: "norm x < norm K" - shows "DERIV (\x. \n. c n * x ^ n) x :> (\n. (diffs c) n * x ^ n)" -proof (simp add: deriv_def, rule LIM_zero_cancel) - show "(\h. (suminf (\n. c n * (x + h) ^ n) - suminf (\n. c n * x ^ n)) / h - - suminf (\n. diffs c n * x ^ n)) -- 0 --> 0" - proof (rule LIM_equal2) - show "0 < norm K - norm x" by (simp add: less_diff_eq 4) - next - fix h :: 'a - assume "h \ 0" - assume "norm (h - 0) < norm K - norm x" - hence "norm x + norm h < norm K" by simp - hence 5: "norm (x + h) < norm K" - by (rule norm_triangle_ineq [THEN order_le_less_trans]) - have A: "summable (\n. c n * x ^ n)" - by (rule powser_inside [OF 1 4]) - have B: "summable (\n. c n * (x + h) ^ n)" - by (rule powser_inside [OF 1 5]) - have C: "summable (\n. diffs c n * x ^ n)" - by (rule powser_inside [OF 2 4]) - show "((\n. c n * (x + h) ^ n) - (\n. c n * x ^ n)) / h - - (\n. diffs c n * x ^ n) = - (\n. c n * (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0)))" - apply (subst sums_unique [OF diffs_equiv [OF C]]) - apply (subst suminf_diff [OF B A]) - apply (subst suminf_divide [symmetric]) - apply (rule summable_diff [OF B A]) - apply (subst suminf_diff) - apply (rule summable_divide) - apply (rule summable_diff [OF B A]) - apply (rule sums_summable [OF diffs_equiv [OF C]]) - apply (rule_tac f="suminf" in arg_cong) - apply (rule ext) - apply (simp add: ring_simps) - done - next - show "(\h. \n. c n * (((x + h) ^ n - x ^ n) / h - - of_nat n * x ^ (n - Suc 0))) -- 0 --> 0" - by (rule termdiffs_aux [OF 3 4]) - qed -qed - - -subsection{*Exponential Function*} - -definition - exp :: "'a \ 'a::{recpower,real_normed_field,banach}" where - "exp x = (\n. x ^ n /\<^sub>R real (fact n))" - -definition - sin :: "real => real" where - "sin x = (\n. (if even(n) then 0 else - (-1 ^ ((n - Suc 0) div 2))/(real (fact n))) * x ^ n)" - -definition - cos :: "real => real" where - "cos x = (\n. (if even(n) then (-1 ^ (n div 2))/(real (fact n)) - else 0) * x ^ n)" - -lemma summable_exp_generic: - fixes x :: "'a::{real_normed_algebra_1,recpower,banach}" - defines S_def: "S \ \n. x ^ n /\<^sub>R real (fact n)" - shows "summable S" -proof - - have S_Suc: "\n. S (Suc n) = (x * S n) /\<^sub>R real (Suc n)" - unfolding S_def by (simp add: power_Suc del: mult_Suc) - obtain r :: real where r0: "0 < r" and r1: "r < 1" - using dense [OF zero_less_one] by fast - obtain N :: nat where N: "norm x < real N * r" - using reals_Archimedean3 [OF r0] by fast - from r1 show ?thesis - proof (rule ratio_test [rule_format]) - fix n :: nat - assume n: "N \ n" - have "norm x \ real N * r" - using N by (rule order_less_imp_le) - also have "real N * r \ real (Suc n) * r" - using r0 n by (simp add: mult_right_mono) - finally have "norm x * norm (S n) \ real (Suc n) * r * norm (S n)" - using norm_ge_zero by (rule mult_right_mono) - hence "norm (x * S n) \ real (Suc n) * r * norm (S n)" - by (rule order_trans [OF norm_mult_ineq]) - hence "norm (x * S n) / real (Suc n) \ r * norm (S n)" - by (simp add: pos_divide_le_eq mult_ac) - thus "norm (S (Suc n)) \ r * norm (S n)" - by (simp add: S_Suc norm_scaleR inverse_eq_divide) - qed -qed - -lemma summable_norm_exp: - fixes x :: "'a::{real_normed_algebra_1,recpower,banach}" - shows "summable (\n. norm (x ^ n /\<^sub>R real (fact n)))" -proof (rule summable_norm_comparison_test [OF exI, rule_format]) - show "summable (\n. norm x ^ n /\<^sub>R real (fact n))" - by (rule summable_exp_generic) -next - fix n show "norm (x ^ n /\<^sub>R real (fact n)) \ norm x ^ n /\<^sub>R real (fact n)" - by (simp add: norm_scaleR norm_power_ineq) -qed - -lemma summable_exp: "summable (%n. inverse (real (fact n)) * x ^ n)" -by (insert summable_exp_generic [where x=x], simp) - -lemma summable_sin: - "summable (%n. - (if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * - x ^ n)" -apply (rule_tac g = "(%n. inverse (real (fact n)) * \x\ ^ n)" in summable_comparison_test) -apply (rule_tac [2] summable_exp) -apply (rule_tac x = 0 in exI) -apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) -done - -lemma summable_cos: - "summable (%n. - (if even n then - -1 ^ (n div 2)/(real (fact n)) else 0) * x ^ n)" -apply (rule_tac g = "(%n. inverse (real (fact n)) * \x\ ^ n)" in summable_comparison_test) -apply (rule_tac [2] summable_exp) -apply (rule_tac x = 0 in exI) -apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) -done - -lemma lemma_STAR_sin: - "(if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0" -by (induct "n", auto) - -lemma lemma_STAR_cos: - "0 < n --> - -1 ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" -by (induct "n", auto) - -lemma lemma_STAR_cos1: - "0 < n --> - (-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" -by (induct "n", auto) - -lemma lemma_STAR_cos2: - "(\n=1..n. x ^ n /\<^sub>R real (fact n)) sums exp x" -unfolding exp_def by (rule summable_exp_generic [THEN summable_sums]) - -lemma sin_converges: - "(%n. (if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * - x ^ n) sums sin(x)" -unfolding sin_def by (rule summable_sin [THEN summable_sums]) - -lemma cos_converges: - "(%n. (if even n then - -1 ^ (n div 2)/(real (fact n)) - else 0) * x ^ n) sums cos(x)" -unfolding cos_def by (rule summable_cos [THEN summable_sums]) - - -subsection{*Formal Derivatives of Exp, Sin, and Cos Series*} - -lemma exp_fdiffs: - "diffs (%n. inverse(real (fact n))) = (%n. inverse(real (fact n)))" -by (simp add: diffs_def mult_assoc [symmetric] real_of_nat_def of_nat_mult - del: mult_Suc of_nat_Suc) - -lemma diffs_of_real: "diffs (\n. of_real (f n)) = (\n. of_real (diffs f n))" -by (simp add: diffs_def) - -lemma sin_fdiffs: - "diffs(%n. if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) - = (%n. if even n then - -1 ^ (n div 2)/(real (fact n)) - else 0)" -by (auto intro!: ext - simp add: diffs_def divide_inverse real_of_nat_def of_nat_mult - simp del: mult_Suc of_nat_Suc) - -lemma sin_fdiffs2: - "diffs(%n. if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) n - = (if even n then - -1 ^ (n div 2)/(real (fact n)) - else 0)" -by (simp only: sin_fdiffs) - -lemma cos_fdiffs: - "diffs(%n. if even n then - -1 ^ (n div 2)/(real (fact n)) else 0) - = (%n. - (if even n then 0 - else -1 ^ ((n - Suc 0)div 2)/(real (fact n))))" -by (auto intro!: ext - simp add: diffs_def divide_inverse odd_Suc_mult_two_ex real_of_nat_def of_nat_mult - simp del: mult_Suc of_nat_Suc) - - -lemma cos_fdiffs2: - "diffs(%n. if even n then - -1 ^ (n div 2)/(real (fact n)) else 0) n - = - (if even n then 0 - else -1 ^ ((n - Suc 0)div 2)/(real (fact n)))" -by (simp only: cos_fdiffs) - -text{*Now at last we can get the derivatives of exp, sin and cos*} - -lemma lemma_sin_minus: - "- sin x = (\n. - ((if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * x ^ n))" -by (auto intro!: sums_unique sums_minus sin_converges) - -lemma lemma_exp_ext: "exp = (\x. \n. x ^ n /\<^sub>R real (fact n))" -by (auto intro!: ext simp add: exp_def) - -lemma DERIV_exp [simp]: "DERIV exp x :> exp(x)" -apply (simp add: exp_def) -apply (subst lemma_exp_ext) -apply (subgoal_tac "DERIV (\u. \n. of_real (inverse (real (fact n))) * u ^ n) x :> (\n. diffs (\n. of_real (inverse (real (fact n)))) n * x ^ n)") -apply (rule_tac [2] K = "of_real (1 + norm x)" in termdiffs) -apply (simp_all only: diffs_of_real scaleR_conv_of_real exp_fdiffs) -apply (rule exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real])+ -apply (simp del: of_real_add) -done - -lemma lemma_sin_ext: - "sin = (%x. \n. - (if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * - x ^ n)" -by (auto intro!: ext simp add: sin_def) - -lemma lemma_cos_ext: - "cos = (%x. \n. - (if even n then -1 ^ (n div 2)/(real (fact n)) else 0) * - x ^ n)" -by (auto intro!: ext simp add: cos_def) - -lemma DERIV_sin [simp]: "DERIV sin x :> cos(x)" -apply (simp add: cos_def) -apply (subst lemma_sin_ext) -apply (auto simp add: sin_fdiffs2 [symmetric]) -apply (rule_tac K = "1 + \x\" in termdiffs) -apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs) -done - -lemma DERIV_cos [simp]: "DERIV cos x :> -sin(x)" -apply (subst lemma_cos_ext) -apply (auto simp add: lemma_sin_minus cos_fdiffs2 [symmetric] minus_mult_left) -apply (rule_tac K = "1 + \x\" in termdiffs) -apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs diffs_minus) -done - -lemma isCont_exp [simp]: "isCont exp x" -by (rule DERIV_exp [THEN DERIV_isCont]) - -lemma isCont_sin [simp]: "isCont sin x" -by (rule DERIV_sin [THEN DERIV_isCont]) - -lemma isCont_cos [simp]: "isCont cos x" -by (rule DERIV_cos [THEN DERIV_isCont]) - - -subsection{*Properties of the Exponential Function*} - -lemma powser_zero: - fixes f :: "nat \ 'a::{real_normed_algebra_1,recpower}" - shows "(\n. f n * 0 ^ n) = f 0" -proof - - have "(\n = 0..<1. f n * 0 ^ n) = (\n. f n * 0 ^ n)" - by (rule sums_unique [OF series_zero], simp add: power_0_left) - thus ?thesis by simp -qed - -lemma exp_zero [simp]: "exp 0 = 1" -unfolding exp_def by (simp add: scaleR_conv_of_real powser_zero) - -lemma setsum_cl_ivl_Suc2: - "(\i=m..Suc n. f i) = (if Suc n < m then 0 else f m + (\i=m..n. f (Suc i)))" -by (simp add: setsum_head_Suc setsum_shift_bounds_cl_Suc_ivl - del: setsum_cl_ivl_Suc) - -lemma exp_series_add: - fixes x y :: "'a::{real_field,recpower}" - defines S_def: "S \ \x n. x ^ n /\<^sub>R real (fact n)" - shows "S (x + y) n = (\i=0..n. S x i * S y (n - i))" -proof (induct n) - case 0 - show ?case - unfolding S_def by simp -next - case (Suc n) - have S_Suc: "\x n. S x (Suc n) = (x * S x n) /\<^sub>R real (Suc n)" - unfolding S_def by (simp add: power_Suc del: mult_Suc) - hence times_S: "\x n. x * S x n = real (Suc n) *\<^sub>R S x (Suc n)" - by simp - - have "real (Suc n) *\<^sub>R S (x + y) (Suc n) = (x + y) * S (x + y) n" - by (simp only: times_S) - also have "\ = (x + y) * (\i=0..n. S x i * S y (n-i))" - by (simp only: Suc) - also have "\ = x * (\i=0..n. S x i * S y (n-i)) - + y * (\i=0..n. S x i * S y (n-i))" - by (rule left_distrib) - also have "\ = (\i=0..n. (x * S x i) * S y (n-i)) - + (\i=0..n. S x i * (y * S y (n-i)))" - by (simp only: setsum_right_distrib mult_ac) - also have "\ = (\i=0..n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n-i))) - + (\i=0..n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i)))" - by (simp add: times_S Suc_diff_le) - also have "(\i=0..n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n-i))) = - (\i=0..Suc n. real i *\<^sub>R (S x i * S y (Suc n-i)))" - by (subst setsum_cl_ivl_Suc2, simp) - also have "(\i=0..n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i))) = - (\i=0..Suc n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i)))" - by (subst setsum_cl_ivl_Suc, simp) - also have "(\i=0..Suc n. real i *\<^sub>R (S x i * S y (Suc n-i))) + - (\i=0..Suc n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i))) = - (\i=0..Suc n. real (Suc n) *\<^sub>R (S x i * S y (Suc n-i)))" - by (simp only: setsum_addf [symmetric] scaleR_left_distrib [symmetric] - real_of_nat_add [symmetric], simp) - also have "\ = real (Suc n) *\<^sub>R (\i=0..Suc n. S x i * S y (Suc n-i))" - by (simp only: scaleR_right.setsum) - finally show - "S (x + y) (Suc n) = (\i=0..Suc n. S x i * S y (Suc n - i))" - by (simp add: scaleR_cancel_left del: setsum_cl_ivl_Suc) -qed - -lemma exp_add: "exp (x + y) = exp x * exp y" -unfolding exp_def -by (simp only: Cauchy_product summable_norm_exp exp_series_add) - -lemma exp_of_real: "exp (of_real x) = of_real (exp x)" -unfolding exp_def -apply (subst of_real.suminf) -apply (rule summable_exp_generic) -apply (simp add: scaleR_conv_of_real) -done - -lemma exp_ge_add_one_self_aux: "0 \ (x::real) ==> (1 + x) \ exp(x)" -apply (drule order_le_imp_less_or_eq, auto) -apply (simp add: exp_def) -apply (rule real_le_trans) -apply (rule_tac [2] n = 2 and f = "(%n. inverse (real (fact n)) * x ^ n)" in series_pos_le) -apply (auto intro: summable_exp simp add: numeral_2_eq_2 zero_le_mult_iff) -done - -lemma exp_gt_one [simp]: "0 < (x::real) ==> 1 < exp x" -apply (rule order_less_le_trans) -apply (rule_tac [2] exp_ge_add_one_self_aux, auto) -done - -lemma DERIV_exp_add_const: "DERIV (%x. exp (x + y)) x :> exp(x + y)" -proof - - have "DERIV (exp \ (\x. x + y)) x :> exp (x + y) * (1+0)" - by (fast intro: DERIV_chain DERIV_add DERIV_exp DERIV_ident DERIV_const) - thus ?thesis by (simp add: o_def) -qed - -lemma DERIV_exp_minus [simp]: "DERIV (%x. exp (-x)) x :> - exp(-x)" -proof - - have "DERIV (exp \ uminus) x :> exp (- x) * - 1" - by (fast intro: DERIV_chain DERIV_minus DERIV_exp DERIV_ident) - thus ?thesis by (simp add: o_def) -qed - -lemma DERIV_exp_exp_zero [simp]: "DERIV (%x. exp (x + y) * exp (- x)) x :> 0" -proof - - have "DERIV (\x. exp (x + y) * exp (- x)) x - :> exp (x + y) * exp (- x) + - exp (- x) * exp (x + y)" - by (fast intro: DERIV_exp_add_const DERIV_exp_minus DERIV_mult) - thus ?thesis by (simp add: mult_commute) -qed - -lemma exp_add_mult_minus [simp]: "exp(x + y)*exp(-x) = exp(y::real)" -proof - - have "\x. DERIV (%x. exp (x + y) * exp (- x)) x :> 0" by simp - hence "exp (x + y) * exp (- x) = exp (0 + y) * exp (- 0)" - by (rule DERIV_isconst_all) - thus ?thesis by simp -qed - -lemma exp_mult_minus [simp]: "exp x * exp(-x) = 1" -by (simp add: exp_add [symmetric]) - -lemma exp_mult_minus2 [simp]: "exp(-x)*exp(x) = 1" -by (simp add: mult_commute) - - -lemma exp_minus: "exp(-x) = inverse(exp(x))" -by (auto intro: inverse_unique [symmetric]) - -text{*Proof: because every exponential can be seen as a square.*} -lemma exp_ge_zero [simp]: "0 \ exp (x::real)" -apply (rule_tac t = x in real_sum_of_halves [THEN subst]) -apply (subst exp_add, auto) -done - -lemma exp_not_eq_zero [simp]: "exp x \ 0" -apply (cut_tac x = x in exp_mult_minus2) -apply (auto simp del: exp_mult_minus2) -done - -lemma exp_gt_zero [simp]: "0 < exp (x::real)" -by (simp add: order_less_le) - -lemma inv_exp_gt_zero [simp]: "0 < inverse(exp x::real)" -by (auto intro: positive_imp_inverse_positive) - -lemma abs_exp_cancel [simp]: "\exp x::real\ = exp x" -by auto - -lemma exp_real_of_nat_mult: "exp(real n * x) = exp(x) ^ n" -apply (induct "n") -apply (auto simp add: real_of_nat_Suc right_distrib exp_add mult_commute) -done - -lemma exp_diff: "exp(x - y) = exp(x)/(exp y)" -apply (simp add: diff_minus divide_inverse) -apply (simp (no_asm) add: exp_add exp_minus) -done - - -lemma exp_less_mono: - fixes x y :: real - assumes xy: "x < y" shows "exp x < exp y" -proof - - from xy have "1 < exp (y + - x)" - by (rule real_less_sum_gt_zero [THEN exp_gt_one]) - hence "exp x * inverse (exp x) < exp y * inverse (exp x)" - by (auto simp add: exp_add exp_minus) - thus ?thesis - by (simp add: divide_inverse [symmetric] pos_less_divide_eq - del: divide_self_if) -qed - -lemma exp_less_cancel: "exp (x::real) < exp y ==> x < y" -apply (simp add: linorder_not_le [symmetric]) -apply (auto simp add: order_le_less exp_less_mono) -done - -lemma exp_less_cancel_iff [iff]: "(exp(x::real) < exp(y)) = (x < y)" -by (auto intro: exp_less_mono exp_less_cancel) - -lemma exp_le_cancel_iff [iff]: "(exp(x::real) \ exp(y)) = (x \ y)" -by (auto simp add: linorder_not_less [symmetric]) - -lemma exp_inj_iff [iff]: "(exp (x::real) = exp y) = (x = y)" -by (simp add: order_eq_iff) - -lemma lemma_exp_total: "1 \ y ==> \x. 0 \ x & x \ y - 1 & exp(x::real) = y" -apply (rule IVT) -apply (auto intro: isCont_exp simp add: le_diff_eq) -apply (subgoal_tac "1 + (y - 1) \ exp (y - 1)") -apply simp -apply (rule exp_ge_add_one_self_aux, simp) -done - -lemma exp_total: "0 < (y::real) ==> \x. exp x = y" -apply (rule_tac x = 1 and y = y in linorder_cases) -apply (drule order_less_imp_le [THEN lemma_exp_total]) -apply (rule_tac [2] x = 0 in exI) -apply (frule_tac [3] real_inverse_gt_one) -apply (drule_tac [4] order_less_imp_le [THEN lemma_exp_total], auto) -apply (rule_tac x = "-x" in exI) -apply (simp add: exp_minus) -done - - -subsection{*Properties of the Logarithmic Function*} - -definition - ln :: "real => real" where - "ln x = (THE u. exp u = x)" - -lemma ln_exp [simp]: "ln (exp x) = x" -by (simp add: ln_def) - -lemma exp_ln [simp]: "0 < x \ exp (ln x) = x" -by (auto dest: exp_total) - -lemma exp_ln_iff [simp]: "(exp (ln x) = x) = (0 < x)" -apply (auto dest: exp_total) -apply (erule subst, simp) -done - -lemma ln_mult: "[| 0 < x; 0 < y |] ==> ln(x * y) = ln(x) + ln(y)" -apply (rule exp_inj_iff [THEN iffD1]) -apply (simp add: exp_add exp_ln mult_pos_pos) -done - -lemma ln_inj_iff[simp]: "[| 0 < x; 0 < y |] ==> (ln x = ln y) = (x = y)" -apply (simp only: exp_ln_iff [symmetric]) -apply (erule subst)+ -apply simp -done - -lemma ln_one[simp]: "ln 1 = 0" -by (rule exp_inj_iff [THEN iffD1], auto) - -lemma ln_inverse: "0 < x ==> ln(inverse x) = - ln x" -apply (rule_tac a1 = "ln x" in add_left_cancel [THEN iffD1]) -apply (auto simp add: positive_imp_inverse_positive ln_mult [symmetric]) -done - -lemma ln_div: - "[|0 < x; 0 < y|] ==> ln(x/y) = ln x - ln y" -apply (simp add: divide_inverse) -apply (auto simp add: positive_imp_inverse_positive ln_mult ln_inverse) -done - -lemma ln_less_cancel_iff[simp]: "[| 0 < x; 0 < y|] ==> (ln x < ln y) = (x < y)" -apply (simp only: exp_ln_iff [symmetric]) -apply (erule subst)+ -apply simp -done - -lemma ln_le_cancel_iff[simp]: "[| 0 < x; 0 < y|] ==> (ln x \ ln y) = (x \ y)" -by (auto simp add: linorder_not_less [symmetric]) - -lemma ln_realpow: "0 < x ==> ln(x ^ n) = real n * ln(x)" -by (auto dest!: exp_total simp add: exp_real_of_nat_mult [symmetric]) - -lemma ln_add_one_self_le_self [simp]: "0 \ x ==> ln(1 + x) \ x" -apply (rule ln_exp [THEN subst]) -apply (rule ln_le_cancel_iff [THEN iffD2]) -apply (auto simp add: exp_ge_add_one_self_aux) -done - -lemma ln_less_self [simp]: "0 < x ==> ln x < x" -apply (rule order_less_le_trans) -apply (rule_tac [2] ln_add_one_self_le_self) -apply (rule ln_less_cancel_iff [THEN iffD2], auto) -done - -lemma ln_ge_zero [simp]: - assumes x: "1 \ x" shows "0 \ ln x" -proof - - have "0 < x" using x by arith - hence "exp 0 \ exp (ln x)" - by (simp add: x) - thus ?thesis by (simp only: exp_le_cancel_iff) -qed - -lemma ln_ge_zero_imp_ge_one: - assumes ln: "0 \ ln x" - and x: "0 < x" - shows "1 \ x" -proof - - from ln have "ln 1 \ ln x" by simp - thus ?thesis by (simp add: x del: ln_one) -qed - -lemma ln_ge_zero_iff [simp]: "0 < x ==> (0 \ ln x) = (1 \ x)" -by (blast intro: ln_ge_zero ln_ge_zero_imp_ge_one) - -lemma ln_less_zero_iff [simp]: "0 < x ==> (ln x < 0) = (x < 1)" -by (insert ln_ge_zero_iff [of x], arith) - -lemma ln_gt_zero: - assumes x: "1 < x" shows "0 < ln x" -proof - - have "0 < x" using x by arith - hence "exp 0 < exp (ln x)" by (simp add: x) - thus ?thesis by (simp only: exp_less_cancel_iff) -qed - -lemma ln_gt_zero_imp_gt_one: - assumes ln: "0 < ln x" - and x: "0 < x" - shows "1 < x" -proof - - from ln have "ln 1 < ln x" by simp - thus ?thesis by (simp add: x del: ln_one) -qed - -lemma ln_gt_zero_iff [simp]: "0 < x ==> (0 < ln x) = (1 < x)" -by (blast intro: ln_gt_zero ln_gt_zero_imp_gt_one) - -lemma ln_eq_zero_iff [simp]: "0 < x ==> (ln x = 0) = (x = 1)" -by (insert ln_less_zero_iff [of x] ln_gt_zero_iff [of x], arith) - -lemma ln_less_zero: "[| 0 < x; x < 1 |] ==> ln x < 0" -by simp - -lemma exp_ln_eq: "exp u = x ==> ln x = u" -by auto - -lemma isCont_ln: "0 < x \ isCont ln x" -apply (subgoal_tac "isCont ln (exp (ln x))", simp) -apply (rule isCont_inverse_function [where f=exp], simp_all) -done - -lemma DERIV_ln: "0 < x \ DERIV ln x :> inverse x" -apply (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"]) -apply (erule lemma_DERIV_subst [OF DERIV_exp exp_ln]) -apply (simp_all add: abs_if isCont_ln) -done - - -subsection{*Basic Properties of the Trigonometric Functions*} - -lemma sin_zero [simp]: "sin 0 = 0" -unfolding sin_def by (simp add: powser_zero) - -lemma cos_zero [simp]: "cos 0 = 1" -unfolding cos_def by (simp add: powser_zero) - -lemma DERIV_sin_sin_mult [simp]: - "DERIV (%x. sin(x)*sin(x)) x :> cos(x) * sin(x) + cos(x) * sin(x)" -by (rule DERIV_mult, auto) - -lemma DERIV_sin_sin_mult2 [simp]: - "DERIV (%x. sin(x)*sin(x)) x :> 2 * cos(x) * sin(x)" -apply (cut_tac x = x in DERIV_sin_sin_mult) -apply (auto simp add: mult_assoc) -done - -lemma DERIV_sin_realpow2 [simp]: - "DERIV (%x. (sin x)\) x :> cos(x) * sin(x) + cos(x) * sin(x)" -by (auto simp add: numeral_2_eq_2 real_mult_assoc [symmetric]) - -lemma DERIV_sin_realpow2a [simp]: - "DERIV (%x. (sin x)\) x :> 2 * cos(x) * sin(x)" -by (auto simp add: numeral_2_eq_2) - -lemma DERIV_cos_cos_mult [simp]: - "DERIV (%x. cos(x)*cos(x)) x :> -sin(x) * cos(x) + -sin(x) * cos(x)" -by (rule DERIV_mult, auto) - -lemma DERIV_cos_cos_mult2 [simp]: - "DERIV (%x. cos(x)*cos(x)) x :> -2 * cos(x) * sin(x)" -apply (cut_tac x = x in DERIV_cos_cos_mult) -apply (auto simp add: mult_ac) -done - -lemma DERIV_cos_realpow2 [simp]: - "DERIV (%x. (cos x)\) x :> -sin(x) * cos(x) + -sin(x) * cos(x)" -by (auto simp add: numeral_2_eq_2 real_mult_assoc [symmetric]) - -lemma DERIV_cos_realpow2a [simp]: - "DERIV (%x. (cos x)\) x :> -2 * cos(x) * sin(x)" -by (auto simp add: numeral_2_eq_2) - -lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" -by auto - -lemma DERIV_cos_realpow2b: "DERIV (%x. (cos x)\) x :> -(2 * cos(x) * sin(x))" -apply (rule lemma_DERIV_subst) -apply (rule DERIV_cos_realpow2a, auto) -done - -(* most useful *) -lemma DERIV_cos_cos_mult3 [simp]: - "DERIV (%x. cos(x)*cos(x)) x :> -(2 * cos(x) * sin(x))" -apply (rule lemma_DERIV_subst) -apply (rule DERIV_cos_cos_mult2, auto) -done - -lemma DERIV_sin_circle_all: - "\x. DERIV (%x. (sin x)\ + (cos x)\) x :> - (2*cos(x)*sin(x) - 2*cos(x)*sin(x))" -apply (simp only: diff_minus, safe) -apply (rule DERIV_add) -apply (auto simp add: numeral_2_eq_2) -done - -lemma DERIV_sin_circle_all_zero [simp]: - "\x. DERIV (%x. (sin x)\ + (cos x)\) x :> 0" -by (cut_tac DERIV_sin_circle_all, auto) - -lemma sin_cos_squared_add [simp]: "((sin x)\) + ((cos x)\) = 1" -apply (cut_tac x = x and y = 0 in DERIV_sin_circle_all_zero [THEN DERIV_isconst_all]) -apply (auto simp add: numeral_2_eq_2) -done - -lemma sin_cos_squared_add2 [simp]: "((cos x)\) + ((sin x)\) = 1" -apply (subst add_commute) -apply (simp (no_asm) del: realpow_Suc) -done - -lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1" -apply (cut_tac x = x in sin_cos_squared_add2) -apply (auto simp add: numeral_2_eq_2) -done - -lemma sin_squared_eq: "(sin x)\ = 1 - (cos x)\" -apply (rule_tac a1 = "(cos x)\" in add_right_cancel [THEN iffD1]) -apply (simp del: realpow_Suc) -done - -lemma cos_squared_eq: "(cos x)\ = 1 - (sin x)\" -apply (rule_tac a1 = "(sin x)\" in add_right_cancel [THEN iffD1]) -apply (simp del: realpow_Suc) -done - -lemma real_gt_one_ge_zero_add_less: "[| 1 < x; 0 \ y |] ==> 1 < x + (y::real)" -by arith - -lemma abs_sin_le_one [simp]: "\sin x\ \ 1" -by (rule power2_le_imp_le, simp_all add: sin_squared_eq) - -lemma sin_ge_minus_one [simp]: "-1 \ sin x" -apply (insert abs_sin_le_one [of x]) -apply (simp add: abs_le_iff del: abs_sin_le_one) -done - -lemma sin_le_one [simp]: "sin x \ 1" -apply (insert abs_sin_le_one [of x]) -apply (simp add: abs_le_iff del: abs_sin_le_one) -done - -lemma abs_cos_le_one [simp]: "\cos x\ \ 1" -by (rule power2_le_imp_le, simp_all add: cos_squared_eq) - -lemma cos_ge_minus_one [simp]: "-1 \ cos x" -apply (insert abs_cos_le_one [of x]) -apply (simp add: abs_le_iff del: abs_cos_le_one) -done - -lemma cos_le_one [simp]: "cos x \ 1" -apply (insert abs_cos_le_one [of x]) -apply (simp add: abs_le_iff del: abs_cos_le_one) -done - -lemma DERIV_fun_pow: "DERIV g x :> m ==> - DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" -apply (rule lemma_DERIV_subst) -apply (rule_tac f = "(%x. x ^ n)" in DERIV_chain2) -apply (rule DERIV_pow, auto) -done - -lemma DERIV_fun_exp: - "DERIV g x :> m ==> DERIV (%x. exp(g x)) x :> exp(g x) * m" -apply (rule lemma_DERIV_subst) -apply (rule_tac f = exp in DERIV_chain2) -apply (rule DERIV_exp, auto) -done - -lemma DERIV_fun_sin: - "DERIV g x :> m ==> DERIV (%x. sin(g x)) x :> cos(g x) * m" -apply (rule lemma_DERIV_subst) -apply (rule_tac f = sin in DERIV_chain2) -apply (rule DERIV_sin, auto) -done - -lemma DERIV_fun_cos: - "DERIV g x :> m ==> DERIV (%x. cos(g x)) x :> -sin(g x) * m" -apply (rule lemma_DERIV_subst) -apply (rule_tac f = cos in DERIV_chain2) -apply (rule DERIV_cos, auto) -done - -lemmas DERIV_intros = DERIV_ident DERIV_const DERIV_cos DERIV_cmult - DERIV_sin DERIV_exp DERIV_inverse DERIV_pow - DERIV_add DERIV_diff DERIV_mult DERIV_minus - DERIV_inverse_fun DERIV_quotient DERIV_fun_pow - DERIV_fun_exp DERIV_fun_sin DERIV_fun_cos - -(* lemma *) -lemma lemma_DERIV_sin_cos_add: - "\x. - DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + - (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2) x :> 0" -apply (safe, rule lemma_DERIV_subst) -apply (best intro!: DERIV_intros intro: DERIV_chain2) - --{*replaces the old @{text DERIV_tac}*} -apply (auto simp add: diff_minus left_distrib right_distrib mult_ac add_ac) -done - -lemma sin_cos_add [simp]: - "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + - (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = 0" -apply (cut_tac y = 0 and x = x and y7 = y - in lemma_DERIV_sin_cos_add [THEN DERIV_isconst_all]) -apply (auto simp add: numeral_2_eq_2) -done - -lemma sin_add: "sin (x + y) = sin x * cos y + cos x * sin y" -apply (cut_tac x = x and y = y in sin_cos_add) -apply (simp del: sin_cos_add) -done - -lemma cos_add: "cos (x + y) = cos x * cos y - sin x * sin y" -apply (cut_tac x = x and y = y in sin_cos_add) -apply (simp del: sin_cos_add) -done - -lemma lemma_DERIV_sin_cos_minus: - "\x. DERIV (%x. (sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2) x :> 0" -apply (safe, rule lemma_DERIV_subst) -apply (best intro!: DERIV_intros intro: DERIV_chain2) -apply (auto simp add: diff_minus left_distrib right_distrib mult_ac add_ac) -done - -lemma sin_cos_minus [simp]: - "(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0" -apply (cut_tac y = 0 and x = x - in lemma_DERIV_sin_cos_minus [THEN DERIV_isconst_all]) -apply simp -done - -lemma sin_minus [simp]: "sin (-x) = -sin(x)" -apply (cut_tac x = x in sin_cos_minus) -apply (simp del: sin_cos_minus) -done - -lemma cos_minus [simp]: "cos (-x) = cos(x)" -apply (cut_tac x = x in sin_cos_minus) -apply (simp del: sin_cos_minus) -done - -lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y" -by (simp add: diff_minus sin_add) - -lemma sin_diff2: "sin (x - y) = cos y * sin x - sin y * cos x" -by (simp add: sin_diff mult_commute) - -lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y" -by (simp add: diff_minus cos_add) - -lemma cos_diff2: "cos (x - y) = cos y * cos x + sin y * sin x" -by (simp add: cos_diff mult_commute) - -lemma sin_double [simp]: "sin(2 * x) = 2* sin x * cos x" -by (cut_tac x = x and y = x in sin_add, auto) - - -lemma cos_double: "cos(2* x) = ((cos x)\) - ((sin x)\)" -apply (cut_tac x = x and y = x in cos_add) -apply (simp add: power2_eq_square) -done - - -subsection{*The Constant Pi*} - -definition - pi :: "real" where - "pi = 2 * (THE x. 0 \ (x::real) & x \ 2 & cos x = 0)" - -text{*Show that there's a least positive @{term x} with @{term "cos(x) = 0"}; - hence define pi.*} - -lemma sin_paired: - "(%n. -1 ^ n /(real (fact (2 * n + 1))) * x ^ (2 * n + 1)) - sums sin x" -proof - - have "(\n. \k = n * 2.. 0 < sin x" -apply (subgoal_tac - "(\n. \k = n * 2..n. -1 ^ n / real (fact (2 * n + 1)) * x ^ (2 * n + 1))") - prefer 2 - apply (rule sin_paired [THEN sums_summable, THEN sums_group], simp) -apply (rotate_tac 2) -apply (drule sin_paired [THEN sums_unique, THEN ssubst]) -apply (auto simp del: fact_Suc realpow_Suc) -apply (frule sums_unique) -apply (auto simp del: fact_Suc realpow_Suc) -apply (rule_tac n1 = 0 in series_pos_less [THEN [2] order_le_less_trans]) -apply (auto simp del: fact_Suc realpow_Suc) -apply (erule sums_summable) -apply (case_tac "m=0") -apply (simp (no_asm_simp)) -apply (subgoal_tac "6 * (x * (x * x) / real (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) < 6 * x") -apply (simp only: mult_less_cancel_left, simp) -apply (simp (no_asm_simp) add: numeral_2_eq_2 [symmetric] mult_assoc [symmetric]) -apply (subgoal_tac "x*x < 2*3", simp) -apply (rule mult_strict_mono) -apply (auto simp add: real_0_less_add_iff real_of_nat_Suc simp del: fact_Suc) -apply (subst fact_Suc) -apply (subst fact_Suc) -apply (subst fact_Suc) -apply (subst fact_Suc) -apply (subst real_of_nat_mult) -apply (subst real_of_nat_mult) -apply (subst real_of_nat_mult) -apply (subst real_of_nat_mult) -apply (simp (no_asm) add: divide_inverse del: fact_Suc) -apply (auto simp add: mult_assoc [symmetric] simp del: fact_Suc) -apply (rule_tac c="real (Suc (Suc (4*m)))" in mult_less_imp_less_right) -apply (auto simp add: mult_assoc simp del: fact_Suc) -apply (rule_tac c="real (Suc (Suc (Suc (4*m))))" in mult_less_imp_less_right) -apply (auto simp add: mult_assoc mult_less_cancel_left simp del: fact_Suc) -apply (subgoal_tac "x * (x * x ^ (4*m)) = (x ^ (4*m)) * (x * x)") -apply (erule ssubst)+ -apply (auto simp del: fact_Suc) -apply (subgoal_tac "0 < x ^ (4 * m) ") - prefer 2 apply (simp only: zero_less_power) -apply (simp (no_asm_simp) add: mult_less_cancel_left) -apply (rule mult_strict_mono) -apply (simp_all (no_asm_simp)) -done - -lemma sin_gt_zero1: "[|0 < x; x < 2 |] ==> 0 < sin x" -by (auto intro: sin_gt_zero) - -lemma cos_double_less_one: "[| 0 < x; x < 2 |] ==> cos (2 * x) < 1" -apply (cut_tac x = x in sin_gt_zero1) -apply (auto simp add: cos_squared_eq cos_double) -done - -lemma cos_paired: - "(%n. -1 ^ n /(real (fact (2 * n))) * x ^ (2 * n)) sums cos x" -proof - - have "(\n. \k = n * 2..n=0..< Suc(Suc(Suc 0)). - (-1 ^ n / (real(fact (2*n))) * 2 ^ (2*n))" - in order_less_trans) -apply (simp (no_asm) add: fact_num_eq_if realpow_num_eq_if del: fact_Suc realpow_Suc) -apply (simp (no_asm) add: mult_assoc del: setsum_op_ivl_Suc) -apply (rule sumr_pos_lt_pair) -apply (erule sums_summable, safe) -apply (simp (no_asm) add: divide_inverse real_0_less_add_iff mult_assoc [symmetric] - del: fact_Suc) -apply (rule real_mult_inverse_cancel2) -apply (rule real_of_nat_fact_gt_zero)+ -apply (simp (no_asm) add: mult_assoc [symmetric] del: fact_Suc) -apply (subst fact_lemma) -apply (subst fact_Suc [of "Suc (Suc (Suc (Suc (Suc (Suc (Suc (4 * d)))))))"]) -apply (simp only: real_of_nat_mult) -apply (rule mult_strict_mono, force) - apply (rule_tac [3] real_of_nat_ge_zero) - prefer 2 apply force -apply (rule real_of_nat_less_iff [THEN iffD2]) -apply (rule fact_less_mono, auto) -done - -lemmas cos_two_neq_zero [simp] = cos_two_less_zero [THEN less_imp_neq] -lemmas cos_two_le_zero [simp] = cos_two_less_zero [THEN order_less_imp_le] - -lemma cos_is_zero: "EX! x. 0 \ x & x \ 2 & cos x = 0" -apply (subgoal_tac "\x. 0 \ x & x \ 2 & cos x = 0") -apply (rule_tac [2] IVT2) -apply (auto intro: DERIV_isCont DERIV_cos) -apply (cut_tac x = xa and y = y in linorder_less_linear) -apply (rule ccontr) -apply (subgoal_tac " (\x. cos differentiable x) & (\x. isCont cos x) ") -apply (auto intro: DERIV_cos DERIV_isCont simp add: differentiable_def) -apply (drule_tac f = cos in Rolle) -apply (drule_tac [5] f = cos in Rolle) -apply (auto dest!: DERIV_cos [THEN DERIV_unique] simp add: differentiable_def) -apply (drule_tac y1 = xa in order_le_less_trans [THEN sin_gt_zero]) -apply (assumption, rule_tac y=y in order_less_le_trans, simp_all) -apply (drule_tac y1 = y in order_le_less_trans [THEN sin_gt_zero], assumption, simp_all) -done - -lemma pi_half: "pi/2 = (THE x. 0 \ x & x \ 2 & cos x = 0)" -by (simp add: pi_def) - -lemma cos_pi_half [simp]: "cos (pi / 2) = 0" -by (simp add: pi_half cos_is_zero [THEN theI']) - -lemma pi_half_gt_zero [simp]: "0 < pi / 2" -apply (rule order_le_neq_trans) -apply (simp add: pi_half cos_is_zero [THEN theI']) -apply (rule notI, drule arg_cong [where f=cos], simp) -done - -lemmas pi_half_neq_zero [simp] = pi_half_gt_zero [THEN less_imp_neq, symmetric] -lemmas pi_half_ge_zero [simp] = pi_half_gt_zero [THEN order_less_imp_le] - -lemma pi_half_less_two [simp]: "pi / 2 < 2" -apply (rule order_le_neq_trans) -apply (simp add: pi_half cos_is_zero [THEN theI']) -apply (rule notI, drule arg_cong [where f=cos], simp) -done - -lemmas pi_half_neq_two [simp] = pi_half_less_two [THEN less_imp_neq] -lemmas pi_half_le_two [simp] = pi_half_less_two [THEN order_less_imp_le] - -lemma pi_gt_zero [simp]: "0 < pi" -by (insert pi_half_gt_zero, simp) - -lemma pi_ge_zero [simp]: "0 \ pi" -by (rule pi_gt_zero [THEN order_less_imp_le]) - -lemma pi_neq_zero [simp]: "pi \ 0" -by (rule pi_gt_zero [THEN less_imp_neq, THEN not_sym]) - -lemma pi_not_less_zero [simp]: "\ pi < 0" -by (simp add: linorder_not_less) - -lemma minus_pi_half_less_zero [simp]: "-(pi/2) < 0" -by auto - -lemma sin_pi_half [simp]: "sin(pi/2) = 1" -apply (cut_tac x = "pi/2" in sin_cos_squared_add2) -apply (cut_tac sin_gt_zero [OF pi_half_gt_zero pi_half_less_two]) -apply (simp add: power2_eq_square) -done - -lemma cos_pi [simp]: "cos pi = -1" -by (cut_tac x = "pi/2" and y = "pi/2" in cos_add, simp) - -lemma sin_pi [simp]: "sin pi = 0" -by (cut_tac x = "pi/2" and y = "pi/2" in sin_add, simp) - -lemma sin_cos_eq: "sin x = cos (pi/2 - x)" -by (simp add: diff_minus cos_add) -declare sin_cos_eq [symmetric, simp] - -lemma minus_sin_cos_eq: "-sin x = cos (x + pi/2)" -by (simp add: cos_add) -declare minus_sin_cos_eq [symmetric, simp] - -lemma cos_sin_eq: "cos x = sin (pi/2 - x)" -by (simp add: diff_minus sin_add) -declare cos_sin_eq [symmetric, simp] - -lemma sin_periodic_pi [simp]: "sin (x + pi) = - sin x" -by (simp add: sin_add) - -lemma sin_periodic_pi2 [simp]: "sin (pi + x) = - sin x" -by (simp add: sin_add) - -lemma cos_periodic_pi [simp]: "cos (x + pi) = - cos x" -by (simp add: cos_add) - -lemma sin_periodic [simp]: "sin (x + 2*pi) = sin x" -by (simp add: sin_add cos_double) - -lemma cos_periodic [simp]: "cos (x + 2*pi) = cos x" -by (simp add: cos_add cos_double) - -lemma cos_npi [simp]: "cos (real n * pi) = -1 ^ n" -apply (induct "n") -apply (auto simp add: real_of_nat_Suc left_distrib) -done - -lemma cos_npi2 [simp]: "cos (pi * real n) = -1 ^ n" -proof - - have "cos (pi * real n) = cos (real n * pi)" by (simp only: mult_commute) - also have "... = -1 ^ n" by (rule cos_npi) - finally show ?thesis . -qed - -lemma sin_npi [simp]: "sin (real (n::nat) * pi) = 0" -apply (induct "n") -apply (auto simp add: real_of_nat_Suc left_distrib) -done - -lemma sin_npi2 [simp]: "sin (pi * real (n::nat)) = 0" -by (simp add: mult_commute [of pi]) - -lemma cos_two_pi [simp]: "cos (2 * pi) = 1" -by (simp add: cos_double) - -lemma sin_two_pi [simp]: "sin (2 * pi) = 0" -by simp - -lemma sin_gt_zero2: "[| 0 < x; x < pi/2 |] ==> 0 < sin x" -apply (rule sin_gt_zero, assumption) -apply (rule order_less_trans, assumption) -apply (rule pi_half_less_two) -done - -lemma sin_less_zero: - assumes lb: "- pi/2 < x" and "x < 0" shows "sin x < 0" -proof - - have "0 < sin (- x)" using prems by (simp only: sin_gt_zero2) - thus ?thesis by simp -qed - -lemma pi_less_4: "pi < 4" -by (cut_tac pi_half_less_two, auto) - -lemma cos_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < cos x" -apply (cut_tac pi_less_4) -apply (cut_tac f = cos and a = 0 and b = x and y = 0 in IVT2_objl, safe, simp_all) -apply (cut_tac cos_is_zero, safe) -apply (rename_tac y z) -apply (drule_tac x = y in spec) -apply (drule_tac x = "pi/2" in spec, simp) -done - -lemma cos_gt_zero_pi: "[| -(pi/2) < x; x < pi/2 |] ==> 0 < cos x" -apply (rule_tac x = x and y = 0 in linorder_cases) -apply (rule cos_minus [THEN subst]) -apply (rule cos_gt_zero) -apply (auto intro: cos_gt_zero) -done - -lemma cos_ge_zero: "[| -(pi/2) \ x; x \ pi/2 |] ==> 0 \ cos x" -apply (auto simp add: order_le_less cos_gt_zero_pi) -apply (subgoal_tac "x = pi/2", auto) -done - -lemma sin_gt_zero_pi: "[| 0 < x; x < pi |] ==> 0 < sin x" -apply (subst sin_cos_eq) -apply (rotate_tac 1) -apply (drule real_sum_of_halves [THEN ssubst]) -apply (auto intro!: cos_gt_zero_pi simp del: sin_cos_eq [symmetric]) -done - -lemma sin_ge_zero: "[| 0 \ x; x \ pi |] ==> 0 \ sin x" -by (auto simp add: order_le_less sin_gt_zero_pi) - -lemma cos_total: "[| -1 \ y; y \ 1 |] ==> EX! x. 0 \ x & x \ pi & (cos x = y)" -apply (subgoal_tac "\x. 0 \ x & x \ pi & cos x = y") -apply (rule_tac [2] IVT2) -apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos) -apply (cut_tac x = xa and y = y in linorder_less_linear) -apply (rule ccontr, auto) -apply (drule_tac f = cos in Rolle) -apply (drule_tac [5] f = cos in Rolle) -apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos - dest!: DERIV_cos [THEN DERIV_unique] - simp add: differentiable_def) -apply (auto dest: sin_gt_zero_pi [OF order_le_less_trans order_less_le_trans]) -done - -lemma sin_total: - "[| -1 \ y; y \ 1 |] ==> EX! x. -(pi/2) \ x & x \ pi/2 & (sin x = y)" -apply (rule ccontr) -apply (subgoal_tac "\x. (- (pi/2) \ x & x \ pi/2 & (sin x = y)) = (0 \ (x + pi/2) & (x + pi/2) \ pi & (cos (x + pi/2) = -y))") -apply (erule contrapos_np) -apply (simp del: minus_sin_cos_eq [symmetric]) -apply (cut_tac y="-y" in cos_total, simp) apply simp -apply (erule ex1E) -apply (rule_tac a = "x - (pi/2)" in ex1I) -apply (simp (no_asm) add: add_assoc) -apply (rotate_tac 3) -apply (drule_tac x = "xa + pi/2" in spec, safe, simp_all) -done - -lemma reals_Archimedean4: - "[| 0 < y; 0 \ x |] ==> \n. real n * y \ x & x < real (Suc n) * y" -apply (auto dest!: reals_Archimedean3) -apply (drule_tac x = x in spec, clarify) -apply (subgoal_tac "x < real(LEAST m::nat. x < real m * y) * y") - prefer 2 apply (erule LeastI) -apply (case_tac "LEAST m::nat. x < real m * y", simp) -apply (subgoal_tac "~ x < real nat * y") - prefer 2 apply (rule not_less_Least, simp, force) -done - -(* Pre Isabelle99-2 proof was simpler- numerals arithmetic - now causes some unwanted re-arrangements of literals! *) -lemma cos_zero_lemma: - "[| 0 \ x; cos x = 0 |] ==> - \n::nat. ~even n & x = real n * (pi/2)" -apply (drule pi_gt_zero [THEN reals_Archimedean4], safe) -apply (subgoal_tac "0 \ x - real n * pi & - (x - real n * pi) \ pi & (cos (x - real n * pi) = 0) ") -apply (auto simp add: compare_rls) - prefer 3 apply (simp add: cos_diff) - prefer 2 apply (simp add: real_of_nat_Suc left_distrib) -apply (simp add: cos_diff) -apply (subgoal_tac "EX! x. 0 \ x & x \ pi & cos x = 0") -apply (rule_tac [2] cos_total, safe) -apply (drule_tac x = "x - real n * pi" in spec) -apply (drule_tac x = "pi/2" in spec) -apply (simp add: cos_diff) -apply (rule_tac x = "Suc (2 * n)" in exI) -apply (simp add: real_of_nat_Suc left_distrib, auto) -done - -lemma sin_zero_lemma: - "[| 0 \ x; sin x = 0 |] ==> - \n::nat. even n & x = real n * (pi/2)" -apply (subgoal_tac "\n::nat. ~ even n & x + pi/2 = real n * (pi/2) ") - apply (clarify, rule_tac x = "n - 1" in exI) - apply (force simp add: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib) -apply (rule cos_zero_lemma) -apply (simp_all add: add_increasing) -done - - -lemma cos_zero_iff: - "(cos x = 0) = - ((\n::nat. ~even n & (x = real n * (pi/2))) | - (\n::nat. ~even n & (x = -(real n * (pi/2)))))" -apply (rule iffI) -apply (cut_tac linorder_linear [of 0 x], safe) -apply (drule cos_zero_lemma, assumption+) -apply (cut_tac x="-x" in cos_zero_lemma, simp, simp) -apply (force simp add: minus_equation_iff [of x]) -apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib) -apply (auto simp add: cos_add) -done - -(* ditto: but to a lesser extent *) -lemma sin_zero_iff: - "(sin x = 0) = - ((\n::nat. even n & (x = real n * (pi/2))) | - (\n::nat. even n & (x = -(real n * (pi/2)))))" -apply (rule iffI) -apply (cut_tac linorder_linear [of 0 x], safe) -apply (drule sin_zero_lemma, assumption+) -apply (cut_tac x="-x" in sin_zero_lemma, simp, simp, safe) -apply (force simp add: minus_equation_iff [of x]) -apply (auto simp add: even_mult_two_ex) -done - - -subsection{*Tangent*} - -definition - tan :: "real => real" where - "tan x = (sin x)/(cos x)" - -lemma tan_zero [simp]: "tan 0 = 0" -by (simp add: tan_def) - -lemma tan_pi [simp]: "tan pi = 0" -by (simp add: tan_def) - -lemma tan_npi [simp]: "tan (real (n::nat) * pi) = 0" -by (simp add: tan_def) - -lemma tan_minus [simp]: "tan (-x) = - tan x" -by (simp add: tan_def minus_mult_left) - -lemma tan_periodic [simp]: "tan (x + 2*pi) = tan x" -by (simp add: tan_def) - -lemma lemma_tan_add1: - "[| cos x \ 0; cos y \ 0 |] - ==> 1 - tan(x)*tan(y) = cos (x + y)/(cos x * cos y)" -apply (simp add: tan_def divide_inverse) -apply (auto simp del: inverse_mult_distrib - simp add: inverse_mult_distrib [symmetric] mult_ac) -apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst]) -apply (auto simp del: inverse_mult_distrib - simp add: mult_assoc left_diff_distrib cos_add) -done - -lemma add_tan_eq: - "[| cos x \ 0; cos y \ 0 |] - ==> tan x + tan y = sin(x + y)/(cos x * cos y)" -apply (simp add: tan_def) -apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst]) -apply (auto simp add: mult_assoc left_distrib) -apply (simp add: sin_add) -done - -lemma tan_add: - "[| cos x \ 0; cos y \ 0; cos (x + y) \ 0 |] - ==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))" -apply (simp (no_asm_simp) add: add_tan_eq lemma_tan_add1) -apply (simp add: tan_def) -done - -lemma tan_double: - "[| cos x \ 0; cos (2 * x) \ 0 |] - ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))" -apply (insert tan_add [of x x]) -apply (simp add: mult_2 [symmetric]) -apply (auto simp add: numeral_2_eq_2) -done - -lemma tan_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < tan x" -by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) - -lemma tan_less_zero: - assumes lb: "- pi/2 < x" and "x < 0" shows "tan x < 0" -proof - - have "0 < tan (- x)" using prems by (simp only: tan_gt_zero) - thus ?thesis by simp -qed - -lemma lemma_DERIV_tan: - "cos x \ 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse((cos x)\)" -apply (rule lemma_DERIV_subst) -apply (best intro!: DERIV_intros intro: DERIV_chain2) -apply (auto simp add: divide_inverse numeral_2_eq_2) -done - -lemma DERIV_tan [simp]: "cos x \ 0 ==> DERIV tan x :> inverse((cos x)\)" -by (auto dest: lemma_DERIV_tan simp add: tan_def [symmetric]) - -lemma isCont_tan [simp]: "cos x \ 0 ==> isCont tan x" -by (rule DERIV_tan [THEN DERIV_isCont]) - -lemma LIM_cos_div_sin [simp]: "(%x. cos(x)/sin(x)) -- pi/2 --> 0" -apply (subgoal_tac "(\x. cos x * inverse (sin x)) -- pi * inverse 2 --> 0*1") -apply (simp add: divide_inverse [symmetric]) -apply (rule LIM_mult) -apply (rule_tac [2] inverse_1 [THEN subst]) -apply (rule_tac [2] LIM_inverse) -apply (simp_all add: divide_inverse [symmetric]) -apply (simp_all only: isCont_def [symmetric] cos_pi_half [symmetric] sin_pi_half [symmetric]) -apply (blast intro!: DERIV_isCont DERIV_sin DERIV_cos)+ -done - -lemma lemma_tan_total: "0 < y ==> \x. 0 < x & x < pi/2 & y < tan x" -apply (cut_tac LIM_cos_div_sin) -apply (simp only: LIM_def) -apply (drule_tac x = "inverse y" in spec, safe, force) -apply (drule_tac ?d1.0 = s in pi_half_gt_zero [THEN [2] real_lbound_gt_zero], safe) -apply (rule_tac x = "(pi/2) - e" in exI) -apply (simp (no_asm_simp)) -apply (drule_tac x = "(pi/2) - e" in spec) -apply (auto simp add: tan_def) -apply (rule inverse_less_iff_less [THEN iffD1]) -apply (auto simp add: divide_inverse) -apply (rule real_mult_order) -apply (subgoal_tac [3] "0 < sin e & 0 < cos e") -apply (auto intro: cos_gt_zero sin_gt_zero2 simp add: mult_commute) -done - -lemma tan_total_pos: "0 \ y ==> \x. 0 \ x & x < pi/2 & tan x = y" -apply (frule order_le_imp_less_or_eq, safe) - prefer 2 apply force -apply (drule lemma_tan_total, safe) -apply (cut_tac f = tan and a = 0 and b = x and y = y in IVT_objl) -apply (auto intro!: DERIV_tan [THEN DERIV_isCont]) -apply (drule_tac y = xa in order_le_imp_less_or_eq) -apply (auto dest: cos_gt_zero) -done - -lemma lemma_tan_total1: "\x. -(pi/2) < x & x < (pi/2) & tan x = y" -apply (cut_tac linorder_linear [of 0 y], safe) -apply (drule tan_total_pos) -apply (cut_tac [2] y="-y" in tan_total_pos, safe) -apply (rule_tac [3] x = "-x" in exI) -apply (auto intro!: exI) -done - -lemma tan_total: "EX! x. -(pi/2) < x & x < (pi/2) & tan x = y" -apply (cut_tac y = y in lemma_tan_total1, auto) -apply (cut_tac x = xa and y = y in linorder_less_linear, auto) -apply (subgoal_tac [2] "\z. y < z & z < xa & DERIV tan z :> 0") -apply (subgoal_tac "\z. xa < z & z < y & DERIV tan z :> 0") -apply (rule_tac [4] Rolle) -apply (rule_tac [2] Rolle) -apply (auto intro!: DERIV_tan DERIV_isCont exI - simp add: differentiable_def) -txt{*Now, simulate TRYALL*} -apply (rule_tac [!] DERIV_tan asm_rl) -apply (auto dest!: DERIV_unique [OF _ DERIV_tan] - simp add: cos_gt_zero_pi [THEN less_imp_neq, THEN not_sym]) -done - - -subsection {* Inverse Trigonometric Functions *} - -definition - arcsin :: "real => real" where - "arcsin y = (THE x. -(pi/2) \ x & x \ pi/2 & sin x = y)" - -definition - arccos :: "real => real" where - "arccos y = (THE x. 0 \ x & x \ pi & cos x = y)" - -definition - arctan :: "real => real" where - "arctan y = (THE x. -(pi/2) < x & x < pi/2 & tan x = y)" - -lemma arcsin: - "[| -1 \ y; y \ 1 |] - ==> -(pi/2) \ arcsin y & - arcsin y \ pi/2 & sin(arcsin y) = y" -unfolding arcsin_def by (rule theI' [OF sin_total]) - -lemma arcsin_pi: - "[| -1 \ y; y \ 1 |] - ==> -(pi/2) \ arcsin y & arcsin y \ pi & sin(arcsin y) = y" -apply (drule (1) arcsin) -apply (force intro: order_trans) -done - -lemma sin_arcsin [simp]: "[| -1 \ y; y \ 1 |] ==> sin(arcsin y) = y" -by (blast dest: arcsin) - -lemma arcsin_bounded: - "[| -1 \ y; y \ 1 |] ==> -(pi/2) \ arcsin y & arcsin y \ pi/2" -by (blast dest: arcsin) - -lemma arcsin_lbound: "[| -1 \ y; y \ 1 |] ==> -(pi/2) \ arcsin y" -by (blast dest: arcsin) - -lemma arcsin_ubound: "[| -1 \ y; y \ 1 |] ==> arcsin y \ pi/2" -by (blast dest: arcsin) - -lemma arcsin_lt_bounded: - "[| -1 < y; y < 1 |] ==> -(pi/2) < arcsin y & arcsin y < pi/2" -apply (frule order_less_imp_le) -apply (frule_tac y = y in order_less_imp_le) -apply (frule arcsin_bounded) -apply (safe, simp) -apply (drule_tac y = "arcsin y" in order_le_imp_less_or_eq) -apply (drule_tac [2] y = "pi/2" in order_le_imp_less_or_eq, safe) -apply (drule_tac [!] f = sin in arg_cong, auto) -done - -lemma arcsin_sin: "[|-(pi/2) \ x; x \ pi/2 |] ==> arcsin(sin x) = x" -apply (unfold arcsin_def) -apply (rule the1_equality) -apply (rule sin_total, auto) -done - -lemma arccos: - "[| -1 \ y; y \ 1 |] - ==> 0 \ arccos y & arccos y \ pi & cos(arccos y) = y" -unfolding arccos_def by (rule theI' [OF cos_total]) - -lemma cos_arccos [simp]: "[| -1 \ y; y \ 1 |] ==> cos(arccos y) = y" -by (blast dest: arccos) - -lemma arccos_bounded: "[| -1 \ y; y \ 1 |] ==> 0 \ arccos y & arccos y \ pi" -by (blast dest: arccos) - -lemma arccos_lbound: "[| -1 \ y; y \ 1 |] ==> 0 \ arccos y" -by (blast dest: arccos) - -lemma arccos_ubound: "[| -1 \ y; y \ 1 |] ==> arccos y \ pi" -by (blast dest: arccos) - -lemma arccos_lt_bounded: - "[| -1 < y; y < 1 |] - ==> 0 < arccos y & arccos y < pi" -apply (frule order_less_imp_le) -apply (frule_tac y = y in order_less_imp_le) -apply (frule arccos_bounded, auto) -apply (drule_tac y = "arccos y" in order_le_imp_less_or_eq) -apply (drule_tac [2] y = pi in order_le_imp_less_or_eq, auto) -apply (drule_tac [!] f = cos in arg_cong, auto) -done - -lemma arccos_cos: "[|0 \ x; x \ pi |] ==> arccos(cos x) = x" -apply (simp add: arccos_def) -apply (auto intro!: the1_equality cos_total) -done - -lemma arccos_cos2: "[|x \ 0; -pi \ x |] ==> arccos(cos x) = -x" -apply (simp add: arccos_def) -apply (auto intro!: the1_equality cos_total) -done - -lemma cos_arcsin: "\-1 \ x; x \ 1\ \ cos (arcsin x) = sqrt (1 - x\)" -apply (subgoal_tac "x\ \ 1") -apply (rule power2_eq_imp_eq) -apply (simp add: cos_squared_eq) -apply (rule cos_ge_zero) -apply (erule (1) arcsin_lbound) -apply (erule (1) arcsin_ubound) -apply simp -apply (subgoal_tac "\x\\ \ 1\", simp) -apply (rule power_mono, simp, simp) -done - -lemma sin_arccos: "\-1 \ x; x \ 1\ \ sin (arccos x) = sqrt (1 - x\)" -apply (subgoal_tac "x\ \ 1") -apply (rule power2_eq_imp_eq) -apply (simp add: sin_squared_eq) -apply (rule sin_ge_zero) -apply (erule (1) arccos_lbound) -apply (erule (1) arccos_ubound) -apply simp -apply (subgoal_tac "\x\\ \ 1\", simp) -apply (rule power_mono, simp, simp) -done - -lemma arctan [simp]: - "- (pi/2) < arctan y & arctan y < pi/2 & tan (arctan y) = y" -unfolding arctan_def by (rule theI' [OF tan_total]) - -lemma tan_arctan: "tan(arctan y) = y" -by auto - -lemma arctan_bounded: "- (pi/2) < arctan y & arctan y < pi/2" -by (auto simp only: arctan) - -lemma arctan_lbound: "- (pi/2) < arctan y" -by auto - -lemma arctan_ubound: "arctan y < pi/2" -by (auto simp only: arctan) - -lemma arctan_tan: - "[|-(pi/2) < x; x < pi/2 |] ==> arctan(tan x) = x" -apply (unfold arctan_def) -apply (rule the1_equality) -apply (rule tan_total, auto) -done - -lemma arctan_zero_zero [simp]: "arctan 0 = 0" -by (insert arctan_tan [of 0], simp) - -lemma cos_arctan_not_zero [simp]: "cos(arctan x) \ 0" -apply (auto simp add: cos_zero_iff) -apply (case_tac "n") -apply (case_tac [3] "n") -apply (cut_tac [2] y = x in arctan_ubound) -apply (cut_tac [4] y = x in arctan_lbound) -apply (auto simp add: real_of_nat_Suc left_distrib mult_less_0_iff) -done - -lemma tan_sec: "cos x \ 0 ==> 1 + tan(x) ^ 2 = inverse(cos x) ^ 2" -apply (rule power_inverse [THEN subst]) -apply (rule_tac c1 = "(cos x)\" in real_mult_right_cancel [THEN iffD1]) -apply (auto dest: field_power_not_zero - simp add: power_mult_distrib left_distrib power_divide tan_def - mult_assoc power_inverse [symmetric] - simp del: realpow_Suc) -done - -lemma isCont_inverse_function2: - fixes f g :: "real \ real" shows - "\a < x; x < b; - \z. a \ z \ z \ b \ g (f z) = z; - \z. a \ z \ z \ b \ isCont f z\ - \ isCont g (f x)" -apply (rule isCont_inverse_function - [where f=f and d="min (x - a) (b - x)"]) -apply (simp_all add: abs_le_iff) -done - -lemma isCont_arcsin: "\-1 < x; x < 1\ \ isCont arcsin x" -apply (subgoal_tac "isCont arcsin (sin (arcsin x))", simp) -apply (rule isCont_inverse_function2 [where f=sin]) -apply (erule (1) arcsin_lt_bounded [THEN conjunct1]) -apply (erule (1) arcsin_lt_bounded [THEN conjunct2]) -apply (fast intro: arcsin_sin, simp) -done - -lemma isCont_arccos: "\-1 < x; x < 1\ \ isCont arccos x" -apply (subgoal_tac "isCont arccos (cos (arccos x))", simp) -apply (rule isCont_inverse_function2 [where f=cos]) -apply (erule (1) arccos_lt_bounded [THEN conjunct1]) -apply (erule (1) arccos_lt_bounded [THEN conjunct2]) -apply (fast intro: arccos_cos, simp) -done - -lemma isCont_arctan: "isCont arctan x" -apply (rule arctan_lbound [of x, THEN dense, THEN exE], clarify) -apply (rule arctan_ubound [of x, THEN dense, THEN exE], clarify) -apply (subgoal_tac "isCont arctan (tan (arctan x))", simp) -apply (erule (1) isCont_inverse_function2 [where f=tan]) -apply (clarify, rule arctan_tan) -apply (erule (1) order_less_le_trans) -apply (erule (1) order_le_less_trans) -apply (clarify, rule isCont_tan) -apply (rule less_imp_neq [symmetric]) -apply (rule cos_gt_zero_pi) -apply (erule (1) order_less_le_trans) -apply (erule (1) order_le_less_trans) -done - -lemma DERIV_arcsin: - "\-1 < x; x < 1\ \ DERIV arcsin x :> inverse (sqrt (1 - x\))" -apply (rule DERIV_inverse_function [where f=sin and a="-1" and b="1"]) -apply (rule lemma_DERIV_subst [OF DERIV_sin]) -apply (simp add: cos_arcsin) -apply (subgoal_tac "\x\\ < 1\", simp) -apply (rule power_strict_mono, simp, simp, simp) -apply assumption -apply assumption -apply simp -apply (erule (1) isCont_arcsin) -done - -lemma DERIV_arccos: - "\-1 < x; x < 1\ \ DERIV arccos x :> inverse (- sqrt (1 - x\))" -apply (rule DERIV_inverse_function [where f=cos and a="-1" and b="1"]) -apply (rule lemma_DERIV_subst [OF DERIV_cos]) -apply (simp add: sin_arccos) -apply (subgoal_tac "\x\\ < 1\", simp) -apply (rule power_strict_mono, simp, simp, simp) -apply assumption -apply assumption -apply simp -apply (erule (1) isCont_arccos) -done - -lemma DERIV_arctan: "DERIV arctan x :> inverse (1 + x\)" -apply (rule DERIV_inverse_function [where f=tan and a="x - 1" and b="x + 1"]) -apply (rule lemma_DERIV_subst [OF DERIV_tan]) -apply (rule cos_arctan_not_zero) -apply (simp add: power_inverse tan_sec [symmetric]) -apply (subgoal_tac "0 < 1 + x\", simp) -apply (simp add: add_pos_nonneg) -apply (simp, simp, simp, rule isCont_arctan) -done - - -subsection {* More Theorems about Sin and Cos *} - -lemma cos_45: "cos (pi / 4) = sqrt 2 / 2" -proof - - let ?c = "cos (pi / 4)" and ?s = "sin (pi / 4)" - have nonneg: "0 \ ?c" - by (rule cos_ge_zero, rule order_trans [where y=0], simp_all) - have "0 = cos (pi / 4 + pi / 4)" - by simp - also have "cos (pi / 4 + pi / 4) = ?c\ - ?s\" - by (simp only: cos_add power2_eq_square) - also have "\ = 2 * ?c\ - 1" - by (simp add: sin_squared_eq) - finally have "?c\ = (sqrt 2 / 2)\" - by (simp add: power_divide) - thus ?thesis - using nonneg by (rule power2_eq_imp_eq) simp -qed - -lemma cos_30: "cos (pi / 6) = sqrt 3 / 2" -proof - - let ?c = "cos (pi / 6)" and ?s = "sin (pi / 6)" - have pos_c: "0 < ?c" - by (rule cos_gt_zero, simp, simp) - have "0 = cos (pi / 6 + pi / 6 + pi / 6)" - by simp - also have "\ = (?c * ?c - ?s * ?s) * ?c - (?s * ?c + ?c * ?s) * ?s" - by (simp only: cos_add sin_add) - also have "\ = ?c * (?c\ - 3 * ?s\)" - by (simp add: ring_simps power2_eq_square) - finally have "?c\ = (sqrt 3 / 2)\" - using pos_c by (simp add: sin_squared_eq power_divide) - thus ?thesis - using pos_c [THEN order_less_imp_le] - by (rule power2_eq_imp_eq) simp -qed - -lemma sin_45: "sin (pi / 4) = sqrt 2 / 2" -proof - - have "sin (pi / 4) = cos (pi / 2 - pi / 4)" by (rule sin_cos_eq) - also have "pi / 2 - pi / 4 = pi / 4" by simp - also have "cos (pi / 4) = sqrt 2 / 2" by (rule cos_45) - finally show ?thesis . -qed - -lemma sin_60: "sin (pi / 3) = sqrt 3 / 2" -proof - - have "sin (pi / 3) = cos (pi / 2 - pi / 3)" by (rule sin_cos_eq) - also have "pi / 2 - pi / 3 = pi / 6" by simp - also have "cos (pi / 6) = sqrt 3 / 2" by (rule cos_30) - finally show ?thesis . -qed - -lemma cos_60: "cos (pi / 3) = 1 / 2" -apply (rule power2_eq_imp_eq) -apply (simp add: cos_squared_eq sin_60 power_divide) -apply (rule cos_ge_zero, rule order_trans [where y=0], simp_all) -done - -lemma sin_30: "sin (pi / 6) = 1 / 2" -proof - - have "sin (pi / 6) = cos (pi / 2 - pi / 6)" by (rule sin_cos_eq) - also have "pi / 2 - pi / 6 = pi / 3" by simp - also have "cos (pi / 3) = 1 / 2" by (rule cos_60) - finally show ?thesis . -qed - -lemma tan_30: "tan (pi / 6) = 1 / sqrt 3" -unfolding tan_def by (simp add: sin_30 cos_30) - -lemma tan_45: "tan (pi / 4) = 1" -unfolding tan_def by (simp add: sin_45 cos_45) - -lemma tan_60: "tan (pi / 3) = sqrt 3" -unfolding tan_def by (simp add: sin_60 cos_60) - -text{*NEEDED??*} -lemma [simp]: - "sin (x + 1 / 2 * real (Suc m) * pi) = - cos (x + 1 / 2 * real (m) * pi)" -by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib, auto) - -text{*NEEDED??*} -lemma [simp]: - "sin (x + real (Suc m) * pi / 2) = - cos (x + real (m) * pi / 2)" -by (simp only: cos_add sin_add real_of_nat_Suc add_divide_distrib left_distrib, auto) - -lemma DERIV_sin_add [simp]: "DERIV (%x. sin (x + k)) xa :> cos (xa + k)" -apply (rule lemma_DERIV_subst) -apply (rule_tac f = sin and g = "%x. x + k" in DERIV_chain2) -apply (best intro!: DERIV_intros intro: DERIV_chain2)+ -apply (simp (no_asm)) -done - -lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi / 2) = (-1) ^ n" -proof - - have "sin ((real n + 1/2) * pi) = cos (real n * pi)" - by (auto simp add: right_distrib sin_add left_distrib mult_ac) - thus ?thesis - by (simp add: real_of_nat_Suc left_distrib add_divide_distrib - mult_commute [of pi]) -qed - -lemma cos_2npi [simp]: "cos (2 * real (n::nat) * pi) = 1" -by (simp add: cos_double mult_assoc power_add [symmetric] numeral_2_eq_2) - -lemma cos_3over2_pi [simp]: "cos (3 / 2 * pi) = 0" -apply (subgoal_tac "cos (pi + pi/2) = 0", simp) -apply (subst cos_add, simp) -done - -lemma sin_2npi [simp]: "sin (2 * real (n::nat) * pi) = 0" -by (auto simp add: mult_assoc) - -lemma sin_3over2_pi [simp]: "sin (3 / 2 * pi) = - 1" -apply (subgoal_tac "sin (pi + pi/2) = - 1", simp) -apply (subst sin_add, simp) -done - -(*NEEDED??*) -lemma [simp]: - "cos(x + 1 / 2 * real(Suc m) * pi) = -sin (x + 1 / 2 * real m * pi)" -apply (simp only: cos_add sin_add real_of_nat_Suc right_distrib left_distrib minus_mult_right, auto) -done - -(*NEEDED??*) -lemma [simp]: "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)" -by (simp only: cos_add sin_add real_of_nat_Suc left_distrib add_divide_distrib, auto) - -lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0" -by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib add_divide_distrib, auto) - -lemma DERIV_cos_add [simp]: "DERIV (%x. cos (x + k)) xa :> - sin (xa + k)" -apply (rule lemma_DERIV_subst) -apply (rule_tac f = cos and g = "%x. x + k" in DERIV_chain2) -apply (best intro!: DERIV_intros intro: DERIV_chain2)+ -apply (simp (no_asm)) -done - -lemma sin_zero_abs_cos_one: "sin x = 0 ==> \cos x\ = 1" -by (auto simp add: sin_zero_iff even_mult_two_ex) - -lemma exp_eq_one_iff [simp]: "(exp (x::real) = 1) = (x = 0)" -apply auto -apply (drule_tac f = ln in arg_cong, auto) -done - -lemma cos_one_sin_zero: "cos x = 1 ==> sin x = 0" -by (cut_tac x = x in sin_cos_squared_add3, auto) - - -subsection {* Existence of Polar Coordinates *} - -lemma cos_x_y_le_one: "\x / sqrt (x\ + y\)\ \ 1" -apply (rule power2_le_imp_le [OF _ zero_le_one]) -apply (simp add: abs_divide power_divide divide_le_eq not_sum_power2_lt_zero) -done - -lemma cos_arccos_abs: "\y\ \ 1 \ cos (arccos y) = y" -by (simp add: abs_le_iff) - -lemma sin_arccos_abs: "\y\ \ 1 \ sin (arccos y) = sqrt (1 - y\)" -by (simp add: sin_arccos abs_le_iff) - -lemmas cos_arccos_lemma1 = cos_arccos_abs [OF cos_x_y_le_one] - -lemmas sin_arccos_lemma1 = sin_arccos_abs [OF cos_x_y_le_one] - -lemma polar_ex1: - "0 < y ==> \r a. x = r * cos a & y = r * sin a" -apply (rule_tac x = "sqrt (x\ + y\)" in exI) -apply (rule_tac x = "arccos (x / sqrt (x\ + y\))" in exI) -apply (simp add: cos_arccos_lemma1) -apply (simp add: sin_arccos_lemma1) -apply (simp add: power_divide) -apply (simp add: real_sqrt_mult [symmetric]) -apply (simp add: right_diff_distrib) -done - -lemma polar_ex2: - "y < 0 ==> \r a. x = r * cos a & y = r * sin a" -apply (insert polar_ex1 [where x=x and y="-y"], simp, clarify) -apply (rule_tac x = r in exI) -apply (rule_tac x = "-a" in exI, simp) -done - -lemma polar_Ex: "\r a. x = r * cos a & y = r * sin a" -apply (rule_tac x=0 and y=y in linorder_cases) -apply (erule polar_ex1) -apply (rule_tac x=x in exI, rule_tac x=0 in exI, simp) -apply (erule polar_ex2) -done - - -subsection {* Theorems about Limits *} - -(* need to rename second isCont_inverse *) - -lemma isCont_inv_fun: - fixes f g :: "real \ real" - shows "[| 0 < d; \z. \z - x\ \ d --> g(f(z)) = z; - \z. \z - x\ \ d --> isCont f z |] - ==> isCont g (f x)" -by (rule isCont_inverse_function) - -lemma isCont_inv_fun_inv: - fixes f g :: "real \ real" - shows "[| 0 < d; - \z. \z - x\ \ d --> g(f(z)) = z; - \z. \z - x\ \ d --> isCont f z |] - ==> \e. 0 < e & - (\y. 0 < \y - f(x)\ & \y - f(x)\ < e --> f(g(y)) = y)" -apply (drule isCont_inj_range) -prefer 2 apply (assumption, assumption, auto) -apply (rule_tac x = e in exI, auto) -apply (rotate_tac 2) -apply (drule_tac x = y in spec, auto) -done - - -text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*} -lemma LIM_fun_gt_zero: - "[| f -- c --> (l::real); 0 < l |] - ==> \r. 0 < r & (\x::real. x \ c & \c - x\ < r --> 0 < f x)" -apply (auto simp add: LIM_def) -apply (drule_tac x = "l/2" in spec, safe, force) -apply (rule_tac x = s in exI) -apply (auto simp only: abs_less_iff) -done - -lemma LIM_fun_less_zero: - "[| f -- c --> (l::real); l < 0 |] - ==> \r. 0 < r & (\x::real. x \ c & \c - x\ < r --> f x < 0)" -apply (auto simp add: LIM_def) -apply (drule_tac x = "-l/2" in spec, safe, force) -apply (rule_tac x = s in exI) -apply (auto simp only: abs_less_iff) -done - - -lemma LIM_fun_not_zero: - "[| f -- c --> (l::real); l \ 0 |] - ==> \r. 0 < r & (\x::real. x \ c & \c - x\ < r --> f x \ 0)" -apply (cut_tac x = l and y = 0 in linorder_less_linear, auto) -apply (drule LIM_fun_less_zero) -apply (drule_tac [3] LIM_fun_gt_zero) -apply force+ -done - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Int.thy --- a/src/HOL/Int.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Int.thy Wed Dec 03 15:58:44 2008 +0100 @@ -16,7 +16,7 @@ ("~~/src/Provers/Arith/assoc_fold.ML") "~~/src/Provers/Arith/cancel_numerals.ML" "~~/src/Provers/Arith/combine_numerals.ML" - ("int_arith1.ML") + ("Tools/int_arith.ML") begin subsection {* The equivalence relation underlying the integers *} @@ -1427,7 +1427,7 @@ lemmas int_int_eq = of_nat_eq_iff [where 'a=int] use "~~/src/Provers/Arith/assoc_fold.ML" -use "int_arith1.ML" +use "Tools/int_arith.ML" declaration {* K int_arith_setup *} diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Integration.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Integration.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,881 @@ +(* Author : Jacques D. Fleuriot + Copyright : 2000 University of Edinburgh + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 +*) + +header{*Theory of Integration*} + +theory Integration +imports MacLaurin +begin + +text{*We follow John Harrison in formalizing the Gauge integral.*} + +definition + --{*Partitions and tagged partitions etc.*} + + partition :: "[(real*real),nat => real] => bool" where + [code del]: "partition = (%(a,b) D. D 0 = a & + (\N. (\n < N. D(n) < D(Suc n)) & + (\n \ N. D(n) = b)))" + +definition + psize :: "(nat => real) => nat" where + [code del]:"psize D = (SOME N. (\n < N. D(n) < D(Suc n)) & + (\n \ N. D(n) = D(N)))" + +definition + tpart :: "[(real*real),((nat => real)*(nat =>real))] => bool" where + [code del]:"tpart = (%(a,b) (D,p). partition(a,b) D & + (\n. D(n) \ p(n) & p(n) \ D(Suc n)))" + + --{*Gauges and gauge-fine divisions*} + +definition + gauge :: "[real => bool, real => real] => bool" where + [code del]:"gauge E g = (\x. E x --> 0 < g(x))" + +definition + fine :: "[real => real, ((nat => real)*(nat => real))] => bool" where + [code del]:"fine = (%g (D,p). \n. n < (psize D) --> D(Suc n) - D(n) < g(p n))" + + --{*Riemann sum*} + +definition + rsum :: "[((nat=>real)*(nat=>real)),real=>real] => real" where + "rsum = (%(D,p) f. \n=0..real,real] => bool" where + [code del]: "Integral = (%(a,b) f k. \e > 0. + (\g. gauge(%x. a \ x & x \ b) g & + (\D p. tpart(a,b) (D,p) & fine(g)(D,p) --> + \rsum(D,p) f - k\ < e)))" + + +lemma partition_zero [simp]: "a = b ==> psize (%n. if n = 0 then a else b) = 0" +by (auto simp add: psize_def) + +lemma partition_one [simp]: "a < b ==> psize (%n. if n = 0 then a else b) = 1" +apply (simp add: psize_def) +apply (rule some_equality, auto) +apply (drule_tac x = 1 in spec, auto) +done + +lemma partition_single [simp]: + "a \ b ==> partition(a,b)(%n. if n = 0 then a else b)" +by (auto simp add: partition_def order_le_less) + +lemma partition_lhs: "partition(a,b) D ==> (D(0) = a)" +by (simp add: partition_def) + +lemma partition: + "(partition(a,b) D) = + ((D 0 = a) & + (\n < psize D. D n < D(Suc n)) & + (\n \ psize D. D n = b))" +apply (simp add: partition_def, auto) +apply (subgoal_tac [!] "psize D = N", auto) +apply (simp_all (no_asm) add: psize_def) +apply (rule_tac [!] some_equality, blast) + prefer 2 apply blast +apply (rule_tac [!] ccontr) +apply (simp_all add: linorder_neq_iff, safe) +apply (drule_tac x = Na in spec) +apply (rotate_tac 3) +apply (drule_tac x = "Suc Na" in spec, simp) +apply (rotate_tac 2) +apply (drule_tac x = N in spec, simp) +apply (drule_tac x = Na in spec) +apply (drule_tac x = "Suc Na" and P = "%n. Na\n \ D n = D Na" in spec, auto) +done + +lemma partition_rhs: "partition(a,b) D ==> (D(psize D) = b)" +by (simp add: partition) + +lemma partition_rhs2: "[|partition(a,b) D; psize D \ n |] ==> (D n = b)" +by (simp add: partition) + +lemma lemma_partition_lt_gen [rule_format]: + "partition(a,b) D & m + Suc d \ n & n \ (psize D) --> D(m) < D(m + Suc d)" +apply (induct "d", auto simp add: partition) +apply (blast dest: Suc_le_lessD intro: less_le_trans order_less_trans) +done + +lemma less_eq_add_Suc: "m < n ==> \d. n = m + Suc d" +by (auto simp add: less_iff_Suc_add) + +lemma partition_lt_gen: + "[|partition(a,b) D; m < n; n \ (psize D)|] ==> D(m) < D(n)" +by (auto dest: less_eq_add_Suc intro: lemma_partition_lt_gen) + +lemma partition_lt: "partition(a,b) D ==> n < (psize D) ==> D(0) < D(Suc n)" +apply (induct "n") +apply (auto simp add: partition) +done + +lemma partition_le: "partition(a,b) D ==> a \ b" +apply (frule partition [THEN iffD1], safe) +apply (drule_tac x = "psize D" and P="%n. psize D \ n --> ?P n" in spec, safe) +apply (case_tac "psize D = 0") +apply (drule_tac [2] n = "psize D - 1" in partition_lt, auto) +done + +lemma partition_gt: "[|partition(a,b) D; n < (psize D)|] ==> D(n) < D(psize D)" +by (auto intro: partition_lt_gen) + +lemma partition_eq: "partition(a,b) D ==> ((a = b) = (psize D = 0))" +apply (frule partition [THEN iffD1], safe) +apply (rotate_tac 2) +apply (drule_tac x = "psize D" in spec) +apply (rule ccontr) +apply (drule_tac n = "psize D - 1" in partition_lt) +apply auto +done + +lemma partition_lb: "partition(a,b) D ==> a \ D(r)" +apply (frule partition [THEN iffD1], safe) +apply (induct "r") +apply (cut_tac [2] y = "Suc r" and x = "psize D" in linorder_le_less_linear) +apply (auto intro: partition_le) +apply (drule_tac x = r in spec) +apply arith; +done + +lemma partition_lb_lt: "[| partition(a,b) D; psize D ~= 0 |] ==> a < D(Suc n)" +apply (rule_tac t = a in partition_lhs [THEN subst], assumption) +apply (cut_tac x = "Suc n" and y = "psize D" in linorder_le_less_linear) +apply (frule partition [THEN iffD1], safe) + apply (blast intro: partition_lt less_le_trans) +apply (rotate_tac 3) +apply (drule_tac x = "Suc n" in spec) +apply (erule impE) +apply (erule less_imp_le) +apply (frule partition_rhs) +apply (drule partition_gt[of _ _ _ 0], arith) +apply (simp (no_asm_simp)) +done + +lemma partition_ub: "partition(a,b) D ==> D(r) \ b" +apply (frule partition [THEN iffD1]) +apply (cut_tac x = "psize D" and y = r in linorder_le_less_linear, safe, blast) +apply (subgoal_tac "\x. D ((psize D) - x) \ b") +apply (rotate_tac 4) +apply (drule_tac x = "psize D - r" in spec) +apply (subgoal_tac "psize D - (psize D - r) = r") +apply simp +apply arith +apply safe +apply (induct_tac "x") +apply (simp (no_asm), blast) +apply (case_tac "psize D - Suc n = 0") +apply (erule_tac V = "\n. psize D \ n --> D n = b" in thin_rl) +apply (simp (no_asm_simp) add: partition_le) +apply (rule order_trans) + prefer 2 apply assumption +apply (subgoal_tac "psize D - n = Suc (psize D - Suc n)") + prefer 2 apply arith +apply (drule_tac x = "psize D - Suc n" in spec, simp) +done + +lemma partition_ub_lt: "[| partition(a,b) D; n < psize D |] ==> D(n) < b" +by (blast intro: partition_rhs [THEN subst] partition_gt) + +lemma lemma_partition_append1: + "[| partition (a, b) D1; partition (b, c) D2 |] + ==> (\n < psize D1 + psize D2. + (if n < psize D1 then D1 n else D2 (n - psize D1)) + < (if Suc n < psize D1 then D1 (Suc n) + else D2 (Suc n - psize D1))) & + (\n \ psize D1 + psize D2. + (if n < psize D1 then D1 n else D2 (n - psize D1)) = + (if psize D1 + psize D2 < psize D1 then D1 (psize D1 + psize D2) + else D2 (psize D1 + psize D2 - psize D1)))" +apply (auto intro: partition_lt_gen) +apply (subgoal_tac "psize D1 = Suc n") +apply (auto intro!: partition_lt_gen simp add: partition_lhs partition_ub_lt) +apply (auto intro!: partition_rhs2 simp add: partition_rhs + split: nat_diff_split) +done + +lemma lemma_psize1: + "[| partition (a, b) D1; partition (b, c) D2; N < psize D1 |] + ==> D1(N) < D2 (psize D2)" +apply (rule_tac y = "D1 (psize D1)" in order_less_le_trans) +apply (erule partition_gt) +apply (auto simp add: partition_rhs partition_le) +done + +lemma lemma_partition_append2: + "[| partition (a, b) D1; partition (b, c) D2 |] + ==> psize (%n. if n < psize D1 then D1 n else D2 (n - psize D1)) = + psize D1 + psize D2" +apply (unfold psize_def + [of "%n. if n < psize D1 then D1 n else D2 (n - psize D1)"]) +apply (rule some1_equality) + prefer 2 apply (blast intro!: lemma_partition_append1) +apply (rule ex1I, rule lemma_partition_append1) +apply (simp_all split: split_if_asm) + txt{*The case @{term "N < psize D1"}*} + apply (drule_tac x = "psize D1 + psize D2" and P="%n. ?P n & ?Q n" in spec) + apply (force dest: lemma_psize1) +apply (rule order_antisym); + txt{*The case @{term "psize D1 \ N"}: + proving @{term "N \ psize D1 + psize D2"}*} + apply (drule_tac x = "psize D1 + psize D2" in spec) + apply (simp add: partition_rhs2) +apply (case_tac "N - psize D1 < psize D2") + prefer 2 apply arith + txt{*Proving @{term "psize D1 + psize D2 \ N"}*} +apply (drule_tac x = "psize D1 + psize D2" and P="%n. N\n --> ?P n" in spec, simp) +apply (drule_tac a = b and b = c in partition_gt, assumption, simp) +done + +lemma tpart_eq_lhs_rhs: "[|psize D = 0; tpart(a,b) (D,p)|] ==> a = b" +by (auto simp add: tpart_def partition_eq) + +lemma tpart_partition: "tpart(a,b) (D,p) ==> partition(a,b) D" +by (simp add: tpart_def) + +lemma partition_append: + "[| tpart(a,b) (D1,p1); fine(g) (D1,p1); + tpart(b,c) (D2,p2); fine(g) (D2,p2) |] + ==> \D p. tpart(a,c) (D,p) & fine(g) (D,p)" +apply (rule_tac x = "%n. if n < psize D1 then D1 n else D2 (n - psize D1)" + in exI) +apply (rule_tac x = "%n. if n < psize D1 then p1 n else p2 (n - psize D1)" + in exI) +apply (case_tac "psize D1 = 0") +apply (auto dest: tpart_eq_lhs_rhs) + prefer 2 +apply (simp add: fine_def + lemma_partition_append2 [OF tpart_partition tpart_partition]) + --{*But must not expand @{term fine} in other subgoals*} +apply auto +apply (subgoal_tac "psize D1 = Suc n") + prefer 2 apply arith +apply (drule tpart_partition [THEN partition_rhs]) +apply (drule tpart_partition [THEN partition_lhs]) +apply (auto split: nat_diff_split) +apply (auto simp add: tpart_def) +defer 1 + apply (subgoal_tac "psize D1 = Suc n") + prefer 2 apply arith + apply (drule partition_rhs) + apply (drule partition_lhs, auto) +apply (simp split: nat_diff_split) +apply (subst partition) +apply (subst (1 2) lemma_partition_append2, assumption+) +apply (rule conjI) +apply (simp add: partition_lhs) +apply (drule lemma_partition_append1) +apply assumption; +apply (simp add: partition_rhs) +done + + +text{*We can always find a division that is fine wrt any gauge*} + +lemma partition_exists: + "[| a \ b; gauge(%x. a \ x & x \ b) g |] + ==> \D p. tpart(a,b) (D,p) & fine g (D,p)" +apply (cut_tac P = "%(u,v). a \ u & v \ b --> + (\D p. tpart (u,v) (D,p) & fine (g) (D,p))" + in lemma_BOLZANO2) +apply safe +apply (blast intro: order_trans)+ +apply (auto intro: partition_append) +apply (case_tac "a \ x & x \ b") +apply (rule_tac [2] x = 1 in exI, auto) +apply (rule_tac x = "g x" in exI) +apply (auto simp add: gauge_def) +apply (rule_tac x = "%n. if n = 0 then aa else ba" in exI) +apply (rule_tac x = "%n. if n = 0 then x else ba" in exI) +apply (auto simp add: tpart_def fine_def) +done + +text{*Lemmas about combining gauges*} + +lemma gauge_min: + "[| gauge(E) g1; gauge(E) g2 |] + ==> gauge(E) (%x. if g1(x) < g2(x) then g1(x) else g2(x))" +by (simp add: gauge_def) + +lemma fine_min: + "fine (%x. if g1(x) < g2(x) then g1(x) else g2(x)) (D,p) + ==> fine(g1) (D,p) & fine(g2) (D,p)" +by (auto simp add: fine_def split: split_if_asm) + + +text{*The integral is unique if it exists*} + +lemma Integral_unique: + "[| a \ b; Integral(a,b) f k1; Integral(a,b) f k2 |] ==> k1 = k2" +apply (simp add: Integral_def) +apply (drule_tac x = "\k1 - k2\ /2" in spec)+ +apply auto +apply (drule gauge_min, assumption) +apply (drule_tac g = "%x. if g x < ga x then g x else ga x" + in partition_exists, assumption, auto) +apply (drule fine_min) +apply (drule spec)+ +apply auto +apply (subgoal_tac "\(rsum (D,p) f - k2) - (rsum (D,p) f - k1)\ < \k1 - k2\") +apply arith +apply (drule add_strict_mono, assumption) +apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] + mult_less_cancel_right) +done + +lemma Integral_zero [simp]: "Integral(a,a) f 0" +apply (auto simp add: Integral_def) +apply (rule_tac x = "%x. 1" in exI) +apply (auto dest: partition_eq simp add: gauge_def tpart_def rsum_def) +done + +lemma sumr_partition_eq_diff_bounds [simp]: + "(\n=0.. b ==> Integral(a,b) (%x. 1) (b - a)" +apply (auto simp add: order_le_less rsum_def Integral_def) +apply (rule_tac x = "%x. b - a" in exI) +apply (auto simp add: gauge_def abs_less_iff tpart_def partition) +done + +lemma Integral_mult_const: "a \ b ==> Integral(a,b) (%x. c) (c*(b - a))" +apply (auto simp add: order_le_less rsum_def Integral_def) +apply (rule_tac x = "%x. b - a" in exI) +apply (auto simp add: setsum_right_distrib [symmetric] gauge_def abs_less_iff + right_diff_distrib [symmetric] partition tpart_def) +done + +lemma Integral_mult: + "[| a \ b; Integral(a,b) f k |] ==> Integral(a,b) (%x. c * f x) (c * k)" +apply (auto simp add: order_le_less + dest: Integral_unique [OF order_refl Integral_zero]) +apply (auto simp add: rsum_def Integral_def setsum_right_distrib[symmetric] mult_assoc) +apply (rule_tac a2 = c in abs_ge_zero [THEN order_le_imp_less_or_eq, THEN disjE]) + prefer 2 apply force +apply (drule_tac x = "e/abs c" in spec, auto) +apply (simp add: zero_less_mult_iff divide_inverse) +apply (rule exI, auto) +apply (drule spec)+ +apply auto +apply (rule_tac z1 = "inverse (abs c)" in real_mult_less_iff1 [THEN iffD1]) +apply (auto simp add: abs_mult divide_inverse [symmetric] right_diff_distrib [symmetric]) +done + +text{*Fundamental theorem of calculus (Part I)*} + +text{*"Straddle Lemma" : Swartz and Thompson: AMM 95(7) 1988 *} + +lemma choiceP: "\x. P(x) --> (\y. Q x y) ==> \f. (\x. P(x) --> Q x (f x))" +by (insert bchoice [of "Collect P" Q], simp) + +(*UNUSED +lemma choice2: "\x. (\y. R(y) & (\z. Q x y z)) ==> + \f fa. (\x. R(f x) & Q x (f x) (fa x))" +*) + + +(* new simplifications e.g. (y < x/n) = (y * n < x) are a real nuisance + they break the original proofs and make new proofs longer!*) +lemma strad1: + "\\xa::real. xa \ x \ \xa - x\ < s \ + \(f xa - f x) / (xa - x) - f' x\ * 2 < e; + 0 < e; a \ x; x \ b; 0 < s\ + \ \z. \z - x\ < s -->\f z - f x - f' x * (z - x)\ * 2 \ e * \z - x\" +apply auto +apply (case_tac "0 < \z - x\") + prefer 2 apply (simp add: zero_less_abs_iff) +apply (drule_tac x = z in spec) +apply (rule_tac z1 = "\inverse (z - x)\" + in real_mult_le_cancel_iff2 [THEN iffD1]) + apply simp +apply (simp del: abs_inverse abs_mult add: abs_mult [symmetric] + mult_assoc [symmetric]) +apply (subgoal_tac "inverse (z - x) * (f z - f x - f' x * (z - x)) + = (f z - f x) / (z - x) - f' x") + apply (simp add: abs_mult [symmetric] mult_ac diff_minus) +apply (subst mult_commute) +apply (simp add: left_distrib diff_minus) +apply (simp add: mult_assoc divide_inverse) +apply (simp add: left_distrib) +done + +lemma lemma_straddle: + "[| \x. a \ x & x \ b --> DERIV f x :> f'(x); 0 < e |] + ==> \g. gauge(%x. a \ x & x \ b) g & + (\x u v. a \ u & u \ x & x \ v & v \ b & (v - u) < g(x) + --> \(f(v) - f(u)) - (f'(x) * (v - u))\ \ e * (v - u))" +apply (simp add: gauge_def) +apply (subgoal_tac "\x. a \ x & x \ b --> + (\d > 0. \u v. u \ x & x \ v & (v - u) < d --> + \(f (v) - f (u)) - (f' (x) * (v - u))\ \ e * (v - u))") +apply (drule choiceP, auto) +apply (drule spec, auto) +apply (auto simp add: DERIV_iff2 LIM_def) +apply (drule_tac x = "e/2" in spec, auto) +apply (frule strad1, assumption+) +apply (rule_tac x = s in exI, auto) +apply (rule_tac x = u and y = v in linorder_cases, auto) +apply (rule_tac y = "\(f (v) - f (x)) - (f' (x) * (v - x))\ + + \(f (x) - f (u)) - (f' (x) * (x - u))\" + in order_trans) +apply (rule abs_triangle_ineq [THEN [2] order_trans]) +apply (simp add: right_diff_distrib) +apply (rule_tac t = "e* (v - u)" in real_sum_of_halves [THEN subst]) +apply (rule add_mono) +apply (rule_tac y = "(e/2) * \v - x\" in order_trans) + prefer 2 apply simp +apply (erule_tac [!] V= "\x'. x' ~= x & \x' - x\ < s --> ?P x'" in thin_rl) +apply (drule_tac x = v in spec, simp add: times_divide_eq) +apply (drule_tac x = u in spec, auto) +apply (subgoal_tac "\f u - f x - f' x * (u - x)\ = \f x - f u - f' x * (x - u)\") +apply (rule order_trans) +apply (auto simp add: abs_le_iff) +apply (simp add: right_diff_distrib) +done + +lemma FTC1: "[|a \ b; \x. a \ x & x \ b --> DERIV f x :> f'(x) |] + ==> Integral(a,b) f' (f(b) - f(a))" +apply (drule order_le_imp_less_or_eq, auto) +apply (auto simp add: Integral_def) +apply (rule ccontr) +apply (subgoal_tac "\e > 0. \g. gauge (%x. a \ x & x \ b) g & (\D p. tpart (a, b) (D, p) & fine g (D, p) --> \rsum (D, p) f' - (f b - f a)\ \ e)") +apply (rotate_tac 3) +apply (drule_tac x = "e/2" in spec, auto) +apply (drule spec, auto) +apply ((drule spec)+, auto) +apply (drule_tac e = "ea/ (b - a)" in lemma_straddle) +apply (auto simp add: zero_less_divide_iff) +apply (rule exI) +apply (auto simp add: tpart_def rsum_def) +apply (subgoal_tac "(\n=0..n=0.. Integral(a,b) f k2" +by simp + +lemma Integral_add: + "[| a \ b; b \ c; Integral(a,b) f' k1; Integral(b,c) f' k2; + \x. a \ x & x \ c --> DERIV f x :> f' x |] + ==> Integral(a,c) f' (k1 + k2)" +apply (rule FTC1 [THEN Integral_subst], auto) +apply (frule FTC1, auto) +apply (frule_tac a = b in FTC1, auto) +apply (drule_tac x = x in spec, auto) +apply (drule_tac ?k2.0 = "f b - f a" in Integral_unique) +apply (drule_tac [3] ?k2.0 = "f c - f b" in Integral_unique, auto) +done + +lemma partition_psize_Least: + "partition(a,b) D ==> psize D = (LEAST n. D(n) = b)" +apply (auto intro!: Least_equality [symmetric] partition_rhs) +apply (auto dest: partition_ub_lt simp add: linorder_not_less [symmetric]) +done + +lemma lemma_partition_bounded: "partition (a, c) D ==> ~ (\n. c < D(n))" +apply safe +apply (drule_tac r = n in partition_ub, auto) +done + +lemma lemma_partition_eq: + "partition (a, c) D ==> D = (%n. if D n < c then D n else c)" +apply (rule ext, auto) +apply (auto dest!: lemma_partition_bounded) +apply (drule_tac x = n in spec, auto) +done + +lemma lemma_partition_eq2: + "partition (a, c) D ==> D = (%n. if D n \ c then D n else c)" +apply (rule ext, auto) +apply (auto dest!: lemma_partition_bounded) +apply (drule_tac x = n in spec, auto) +done + +lemma partition_lt_Suc: + "[| partition(a,b) D; n < psize D |] ==> D n < D (Suc n)" +by (auto simp add: partition) + +lemma tpart_tag_eq: "tpart(a,c) (D,p) ==> p = (%n. if D n < c then p n else c)" +apply (rule ext) +apply (auto simp add: tpart_def) +apply (drule linorder_not_less [THEN iffD1]) +apply (drule_tac r = "Suc n" in partition_ub) +apply (drule_tac x = n in spec, auto) +done + +subsection{*Lemmas for Additivity Theorem of Gauge Integral*} + +lemma lemma_additivity1: + "[| a \ D n; D n < b; partition(a,b) D |] ==> n < psize D" +by (auto simp add: partition linorder_not_less [symmetric]) + +lemma lemma_additivity2: "[| a \ D n; partition(a,D n) D |] ==> psize D \ n" +apply (rule ccontr, drule not_leE) +apply (frule partition [THEN iffD1], safe) +apply (frule_tac r = "Suc n" in partition_ub) +apply (auto dest!: spec) +done + +lemma partition_eq_bound: + "[| partition(a,b) D; psize D < m |] ==> D(m) = D(psize D)" +by (auto simp add: partition) + +lemma partition_ub2: "[| partition(a,b) D; psize D < m |] ==> D(r) \ D(m)" +by (simp add: partition partition_ub) + +lemma tag_point_eq_partition_point: + "[| tpart(a,b) (D,p); psize D \ m |] ==> p(m) = D(m)" +apply (simp add: tpart_def, auto) +apply (drule_tac x = m in spec) +apply (auto simp add: partition_rhs2) +done + +lemma partition_lt_cancel: "[| partition(a,b) D; D m < D n |] ==> m < n" +apply (cut_tac less_linear [of n "psize D"], auto) +apply (cut_tac less_linear [of m n]) +apply (cut_tac less_linear [of m "psize D"]) +apply (auto dest: partition_gt) +apply (drule_tac n = m in partition_lt_gen, auto) +apply (frule partition_eq_bound) +apply (drule_tac [2] partition_gt, auto) +apply (metis dense_linear_order_class.dlo_simps(8) not_less partition_rhs partition_rhs2) +apply (metis le_less_trans dense_linear_order_class.dlo_simps(8) nat_le_linear partition_eq_bound partition_rhs2) +done + +lemma lemma_additivity4_psize_eq: + "[| a \ D n; D n < b; partition (a, b) D |] + ==> psize (%x. if D x < D n then D(x) else D n) = n" +apply (unfold psize_def) +apply (frule lemma_additivity1) +apply (assumption, assumption) +apply (rule some_equality) +apply (auto intro: partition_lt_Suc) +apply (drule_tac n = n in partition_lt_gen, assumption) +apply (arith, arith) +apply (cut_tac x = na and y = "psize D" in less_linear) +apply (auto dest: partition_lt_cancel) +apply (rule_tac x=N and y=n in linorder_cases) +apply (drule_tac x = n and P="%m. N \ m --> ?f m = ?g m" in spec, simp) +apply (drule_tac n = n in partition_lt_gen, auto) +apply (drule_tac x = n in spec) +apply (simp split: split_if_asm) +done + +lemma lemma_psize_left_less_psize: + "partition (a, b) D + ==> psize (%x. if D x < D n then D(x) else D n) \ psize D" +apply (frule_tac r = n in partition_ub) +apply (drule_tac x = "D n" in order_le_imp_less_or_eq) +apply (auto simp add: lemma_partition_eq [symmetric]) +apply (frule_tac r = n in partition_lb) +apply (drule (2) lemma_additivity4_psize_eq) +apply (rule ccontr, auto) +apply (frule_tac not_leE [THEN [2] partition_eq_bound]) +apply (auto simp add: partition_rhs) +done + +lemma lemma_psize_left_less_psize2: + "[| partition(a,b) D; na < psize (%x. if D x < D n then D(x) else D n) |] + ==> na < psize D" +by (erule lemma_psize_left_less_psize [THEN [2] less_le_trans]) + + +lemma lemma_additivity3: + "[| partition(a,b) D; D na < D n; D n < D (Suc na); + n < psize D |] + ==> False" +by (metis not_less_eq partition_lt_cancel real_of_nat_less_iff) + + +lemma psize_const [simp]: "psize (%x. k) = 0" +by (auto simp add: psize_def) + +lemma lemma_additivity3a: + "[| partition(a,b) D; D na < D n; D n < D (Suc na); + na < psize D |] + ==> False" +apply (frule_tac m = n in partition_lt_cancel) +apply (auto intro: lemma_additivity3) +done + +lemma better_lemma_psize_right_eq1: + "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) \ psize D - n" +apply (simp add: psize_def [of "(%x. D (x + n))"]); +apply (rule_tac a = "psize D - n" in someI2, auto) + apply (simp add: partition less_diff_conv) + apply (simp add: le_diff_conv partition_rhs2 split: nat_diff_split) +apply (drule_tac x = "psize D - n" in spec, auto) +apply (frule partition_rhs, safe) +apply (frule partition_lt_cancel, assumption) +apply (drule partition [THEN iffD1], safe) +apply (subgoal_tac "~ D (psize D - n + n) < D (Suc (psize D - n + n))") + apply blast +apply (drule_tac x = "Suc (psize D)" and P="%n. ?P n \ D n = D (psize D)" + in spec) +apply simp +done + +lemma psize_le_n: "partition (a, D n) D ==> psize D \ n" +apply (rule ccontr, drule not_leE) +apply (frule partition_lt_Suc, assumption) +apply (frule_tac r = "Suc n" in partition_ub, auto) +done + +lemma better_lemma_psize_right_eq1a: + "partition(a,D n) D ==> psize (%x. D (x + n)) \ psize D - n" +apply (simp add: psize_def [of "(%x. D (x + n))"]); +apply (rule_tac a = "psize D - n" in someI2, auto) + apply (simp add: partition less_diff_conv) + apply (simp add: le_diff_conv) +apply (case_tac "psize D \ n") + apply (force intro: partition_rhs2) + apply (simp add: partition linorder_not_le) +apply (rule ccontr, drule not_leE) +apply (frule psize_le_n) +apply (drule_tac x = "psize D - n" in spec, simp) +apply (drule partition [THEN iffD1], safe) +apply (drule_tac x = "Suc n" and P="%na. ?s \ na \ D na = D n" in spec, auto) +done + +lemma better_lemma_psize_right_eq: + "partition(a,b) D ==> psize (%x. D (x + n)) \ psize D - n" +apply (frule_tac r1 = n in partition_ub [THEN order_le_imp_less_or_eq]) +apply (blast intro: better_lemma_psize_right_eq1a better_lemma_psize_right_eq1) +done + +lemma lemma_psize_right_eq1: + "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) \ psize D" +apply (simp add: psize_def [of "(%x. D (x + n))"]) +apply (rule_tac a = "psize D - n" in someI2, auto) + apply (simp add: partition less_diff_conv) + apply (subgoal_tac "n \ psize D") + apply (simp add: partition le_diff_conv) + apply (rule ccontr, drule not_leE) + apply (drule_tac less_imp_le [THEN [2] partition_rhs2], assumption, simp) +apply (drule_tac x = "psize D" in spec) +apply (simp add: partition) +done + +(* should be combined with previous theorem; also proof has redundancy *) +lemma lemma_psize_right_eq1a: + "partition(a,D n) D ==> psize (%x. D (x + n)) \ psize D" +apply (simp add: psize_def [of "(%x. D (x + n))"]); +apply (rule_tac a = "psize D - n" in someI2, auto) + apply (simp add: partition less_diff_conv) + apply (case_tac "psize D \ n") + apply (force intro: partition_rhs2 simp add: le_diff_conv) + apply (simp add: partition le_diff_conv) +apply (rule ccontr, drule not_leE) +apply (drule_tac x = "psize D" in spec) +apply (simp add: partition) +done + +lemma lemma_psize_right_eq: + "[| partition(a,b) D |] ==> psize (%x. D (x + n)) \ psize D" +apply (frule_tac r1 = n in partition_ub [THEN order_le_imp_less_or_eq]) +apply (blast intro: lemma_psize_right_eq1a lemma_psize_right_eq1) +done + +lemma tpart_left1: + "[| a \ D n; tpart (a, b) (D, p) |] + ==> tpart(a, D n) (%x. if D x < D n then D(x) else D n, + %x. if D x < D n then p(x) else D n)" +apply (frule_tac r = n in tpart_partition [THEN partition_ub]) +apply (drule_tac x = "D n" in order_le_imp_less_or_eq) +apply (auto simp add: tpart_partition [THEN lemma_partition_eq, symmetric] tpart_tag_eq [symmetric]) +apply (frule_tac tpart_partition [THEN [3] lemma_additivity1]) +apply (auto simp add: tpart_def) +apply (drule_tac [2] linorder_not_less [THEN iffD1, THEN order_le_imp_less_or_eq], auto) + prefer 3 apply (drule_tac x=na in spec, arith) + prefer 2 apply (blast dest: lemma_additivity3) +apply (frule (2) lemma_additivity4_psize_eq) +apply (rule partition [THEN iffD2]) +apply (frule partition [THEN iffD1]) +apply safe +apply (auto simp add: partition_lt_gen) +apply (drule (1) partition_lt_cancel, arith) +done + +lemma fine_left1: + "[| a \ D n; tpart (a, b) (D, p); gauge (%x. a \ x & x \ D n) g; + fine (%x. if x < D n then min (g x) ((D n - x)/ 2) + else if x = D n then min (g (D n)) (ga (D n)) + else min (ga x) ((x - D n)/ 2)) (D, p) |] + ==> fine g + (%x. if D x < D n then D(x) else D n, + %x. if D x < D n then p(x) else D n)" +apply (auto simp add: fine_def tpart_def gauge_def) +apply (frule_tac [!] na=na in lemma_psize_left_less_psize2) +apply (drule_tac [!] x = na in spec, auto) +apply (drule_tac [!] x = na in spec, auto) +apply (auto dest: lemma_additivity3a simp add: split_if_asm) +done + +lemma tpart_right1: + "[| a \ D n; tpart (a, b) (D, p) |] + ==> tpart(D n, b) (%x. D(x + n),%x. p(x + n))" +apply (simp add: tpart_def partition_def, safe) +apply (rule_tac x = "N - n" in exI, auto) +done + +lemma fine_right1: + "[| a \ D n; tpart (a, b) (D, p); gauge (%x. D n \ x & x \ b) ga; + fine (%x. if x < D n then min (g x) ((D n - x)/ 2) + else if x = D n then min (g (D n)) (ga (D n)) + else min (ga x) ((x - D n)/ 2)) (D, p) |] + ==> fine ga (%x. D(x + n),%x. p(x + n))" +apply (auto simp add: fine_def gauge_def) +apply (drule_tac x = "na + n" in spec) +apply (frule_tac n = n in tpart_partition [THEN better_lemma_psize_right_eq], auto) +apply (simp add: tpart_def, safe) +apply (subgoal_tac "D n \ p (na + n)") +apply (drule_tac y = "p (na + n)" in order_le_imp_less_or_eq) +apply safe +apply (simp split: split_if_asm, simp) +apply (drule less_le_trans, assumption) +apply (rotate_tac 5) +apply (drule_tac x = "na + n" in spec, safe) +apply (rule_tac y="D (na + n)" in order_trans) +apply (case_tac "na = 0", auto) +apply (erule partition_lt_gen [THEN order_less_imp_le]) +apply arith +apply arith +done + +lemma rsum_add: "rsum (D, p) (%x. f x + g x) = rsum (D, p) f + rsum(D, p) g" +by (simp add: rsum_def setsum_addf left_distrib) + +text{* Bartle/Sherbert: Theorem 10.1.5 p. 278 *} +lemma Integral_add_fun: + "[| a \ b; Integral(a,b) f k1; Integral(a,b) g k2 |] + ==> Integral(a,b) (%x. f x + g x) (k1 + k2)" +apply (simp add: Integral_def, auto) +apply ((drule_tac x = "e/2" in spec)+) +apply auto +apply (drule gauge_min, assumption) +apply (rule_tac x = " (%x. if ga x < gaa x then ga x else gaa x)" in exI) +apply auto +apply (drule fine_min) +apply ((drule spec)+, auto) +apply (drule_tac a = "\rsum (D, p) f - k1\ * 2" and c = "\rsum (D, p) g - k2\ * 2" in add_strict_mono, assumption) +apply (auto simp only: rsum_add left_distrib [symmetric] + mult_2_right [symmetric] real_mult_less_iff1) +done + +lemma partition_lt_gen2: + "[| partition(a,b) D; r < psize D |] ==> 0 < D (Suc r) - D r" +by (auto simp add: partition) + +lemma lemma_Integral_le: + "[| \x. a \ x & x \ b --> f x \ g x; + tpart(a,b) (D,p) + |] ==> \n \ psize D. f (p n) \ g (p n)" +apply (simp add: tpart_def) +apply (auto, frule partition [THEN iffD1], auto) +apply (drule_tac x = "p n" in spec, auto) +apply (case_tac "n = 0", simp) +apply (rule partition_lt_gen [THEN order_less_le_trans, THEN order_less_imp_le], auto) +apply (drule le_imp_less_or_eq, auto) +apply (drule_tac [2] x = "psize D" in spec, auto) +apply (drule_tac r = "Suc n" in partition_ub) +apply (drule_tac x = n in spec, auto) +done + +lemma lemma_Integral_rsum_le: + "[| \x. a \ x & x \ b --> f x \ g x; + tpart(a,b) (D,p) + |] ==> rsum(D,p) f \ rsum(D,p) g" +apply (simp add: rsum_def) +apply (auto intro!: setsum_mono dest: tpart_partition [THEN partition_lt_gen2] + dest!: lemma_Integral_le) +done + +lemma Integral_le: + "[| a \ b; + \x. a \ x & x \ b --> f(x) \ g(x); + Integral(a,b) f k1; Integral(a,b) g k2 + |] ==> k1 \ k2" +apply (simp add: Integral_def) +apply (rotate_tac 2) +apply (drule_tac x = "\k1 - k2\ /2" in spec) +apply (drule_tac x = "\k1 - k2\ /2" in spec, auto) +apply (drule gauge_min, assumption) +apply (drule_tac g = "%x. if ga x < gaa x then ga x else gaa x" + in partition_exists, assumption, auto) +apply (drule fine_min) +apply (drule_tac x = D in spec, drule_tac x = D in spec) +apply (drule_tac x = p in spec, drule_tac x = p in spec, auto) +apply (frule lemma_Integral_rsum_le, assumption) +apply (subgoal_tac "\(rsum (D,p) f - k1) - (rsum (D,p) g - k2)\ < \k1 - k2\") +apply arith +apply (drule add_strict_mono, assumption) +apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] + real_mult_less_iff1) +done + +lemma Integral_imp_Cauchy: + "(\k. Integral(a,b) f k) ==> + (\e > 0. \g. gauge (%x. a \ x & x \ b) g & + (\D1 D2 p1 p2. + tpart(a,b) (D1, p1) & fine g (D1,p1) & + tpart(a,b) (D2, p2) & fine g (D2,p2) --> + \rsum(D1,p1) f - rsum(D2,p2) f\ < e))" +apply (simp add: Integral_def, auto) +apply (drule_tac x = "e/2" in spec, auto) +apply (rule exI, auto) +apply (frule_tac x = D1 in spec) +apply (frule_tac x = D2 in spec) +apply ((drule spec)+, auto) +apply (erule_tac V = "0 < e" in thin_rl) +apply (drule add_strict_mono, assumption) +apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] + real_mult_less_iff1) +done + +lemma Cauchy_iff2: + "Cauchy X = + (\j. (\M. \m \ M. \n \ M. \X m - X n\ < inverse(real (Suc j))))" +apply (simp add: Cauchy_def, auto) +apply (drule reals_Archimedean, safe) +apply (drule_tac x = n in spec, auto) +apply (rule_tac x = M in exI, auto) +apply (drule_tac x = m in spec, simp) +apply (drule_tac x = na in spec, auto) +done + +lemma partition_exists2: + "[| a \ b; \n. gauge (%x. a \ x & x \ b) (fa n) |] + ==> \n. \D p. tpart (a, b) (D, p) & fine (fa n) (D, p)" +by (blast dest: partition_exists) + +lemma monotonic_anti_derivative: + fixes f g :: "real => real" shows + "[| a \ b; \c. a \ c & c \ b --> f' c \ g' c; + \x. DERIV f x :> f' x; \x. DERIV g x :> g' x |] + ==> f b - f a \ g b - g a" +apply (rule Integral_le, assumption) +apply (auto intro: FTC1) +done + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/IsaMakefile Wed Dec 03 15:58:44 2008 +0100 @@ -99,6 +99,7 @@ SAT.thy \ Set.thy \ Sum_Type.thy \ + Tools/arith_data.ML \ Tools/cnf_funcs.ML \ Tools/datatype_abs_proofs.ML \ Tools/datatype_aux.ML \ @@ -124,6 +125,7 @@ Tools/function_package/pattern_split.ML \ Tools/function_package/size.ML \ Tools/function_package/sum_tree.ML \ + Tools/hologic.ML \ Tools/inductive_codegen.ML \ Tools/inductive_package.ML \ Tools/inductive_realizer.ML \ @@ -139,6 +141,7 @@ Tools/rewrite_hol_proof.ML \ Tools/sat_funcs.ML \ Tools/sat_solver.ML \ + Tools/simpdata.ML \ Tools/split_rule.ML \ Tools/typecopy_package.ML \ Tools/typedef_codegen.ML \ @@ -146,9 +149,6 @@ Transitive_Closure.thy \ Typedef.thy \ Wellfounded.thy \ - arith_data.ML \ - hologic.ML \ - simpdata.ML \ $(SRC)/Provers/Arith/abel_cancel.ML \ $(SRC)/Provers/Arith/cancel_div_mod.ML \ $(SRC)/Provers/Arith/cancel_sums.ML \ @@ -170,6 +170,7 @@ $(SRC)/Tools/IsaPlanner/zipper.ML \ $(SRC)/Tools/atomize_elim.ML \ $(SRC)/Tools/code/code_funcgr.ML \ + $(SRC)/Tools/code/code_funcgr.ML \ $(SRC)/Tools/code/code_name.ML \ $(SRC)/Tools/code/code_printer.ML \ $(SRC)/Tools/code/code_target.ML \ @@ -189,20 +190,18 @@ Arith_Tools.thy \ ATP_Linkup.thy \ Code_Eval.thy \ + Code_Message.thy \ Equiv_Relations.thy \ Groebner_Basis.thy \ Hilbert_Choice.thy \ - int_arith1.ML \ IntDiv.thy \ - int_factor_simprocs.ML \ Int.thy \ - Library/RType.thy \ + Typerep.thy \ List.thy \ Main.thy \ Map.thy \ NatBin.thy \ Nat_Int_Bij.thy \ - nat_simprocs.ML \ Presburger.thy \ Recdef.thy \ Relation_Power.thy \ @@ -213,6 +212,9 @@ $(SRC)/Provers/Arith/combine_numerals.ML \ $(SRC)/Provers/Arith/extract_common_term.ML \ $(SRC)/Tools/Metis/metis.ML \ + Tools/int_arith.ML \ + Tools/int_factor_simprocs.ML \ + Tools/nat_simprocs.ML \ Tools/Groebner_Basis/groebner.ML \ Tools/Groebner_Basis/misc.ML \ Tools/Groebner_Basis/normalizer_data.ML \ @@ -251,38 +253,39 @@ @$(ISABELLE_TOOL) usedir -b -f main.ML -g true $(OUT)/Pure HOL-Main $(OUT)/HOL: ROOT.ML $(MAIN_DEPENDENCIES) \ - Complex/Complex_Main.thy \ - Complex/Complex.thy \ + Complex_Main.thy \ + Complex.thy \ Complex/Fundamental_Theorem_Algebra.thy \ - Hyperreal/Deriv.thy \ - Hyperreal/Fact.thy \ - Hyperreal/Integration.thy \ - Hyperreal/Lim.thy \ - Hyperreal/Ln.thy \ - Hyperreal/Log.thy \ - Hyperreal/MacLaurin.thy \ - Hyperreal/NthRoot.thy \ + Deriv.thy \ + Fact.thy \ + FrechetDeriv.thy \ + Integration.thy \ + Lim.thy \ + Ln.thy \ + Log.thy \ + MacLaurin.thy \ + NthRoot.thy \ Hyperreal/SEQ.thy \ - Hyperreal/Series.thy \ - Hyperreal/Taylor.thy \ - Hyperreal/Transcendental.thy \ + Series.thy \ + Taylor.thy \ + Transcendental.thy \ Library/Dense_Linear_Order.thy \ - Library/GCD.thy \ - Library/Order_Relation.thy \ - Library/Parity.thy \ - Library/Univ_Poly.thy \ - Real/ContNotDenum.thy \ - Real/float_syntax.ML \ - Real/Lubs.thy \ - Real/PReal.thy \ - Real/rat_arith.ML \ - Real/Rational.thy \ - Real/RComplete.thy \ - Real/real_arith.ML \ - Real/RealDef.thy \ - Real/RealPow.thy \ - Real/Real.thy \ + GCD.thy \ + Order_Relation.thy \ + Parity.thy \ + Univ_Poly.thy \ + ContNotDenum.thy \ + Lubs.thy \ + PReal.thy \ + Rational.thy \ + RComplete.thy \ + RealDef.thy \ + RealPow.thy \ + Real.thy \ Real/RealVector.thy \ + Tools/float_syntax.ML \ + Tools/rat_arith.ML \ + Tools/real_arith.ML \ Tools/Qelim/ferrante_rackoff_data.ML \ Tools/Qelim/ferrante_rackoff.ML \ Tools/Qelim/langford_data.ML \ @@ -315,13 +318,12 @@ Library/Binomial.thy Library/Eval_Witness.thy \ Library/Code_Index.thy Library/Code_Char.thy \ Library/Code_Char_chr.thy Library/Code_Integer.thy \ - Library/Code_Message.thy \ Library/Numeral_Type.thy \ - Library/Boolean_Algebra.thy Library/Countable.thy Library/RType.thy \ + Library/Boolean_Algebra.thy Library/Countable.thy \ Library/Heap.thy Library/Heap_Monad.thy Library/Array.thy \ Library/Relational.thy Library/Sublist.thy Library/Subarray.thy \ Library/Ref.thy Library/Imperative_HOL.thy Library/RBT.thy \ - Library/Enum.thy Real/Float.thy $(SRC)/Tools/float.ML + Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML $(SRC)/HOL/Tools/float_arith.ML @cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library @@ -793,11 +795,12 @@ ex/document/root.tex ex/Meson_Test.thy ex/reflection.ML ex/set.thy \ ex/svc_funcs.ML ex/svc_test.thy \ ex/ImperativeQuicksort.thy \ - Complex/ex/BigO_Complex.thy Complex/ex/BinEx.thy \ - Complex/ex/Sqrt.thy \ - Complex/ex/Sqrt_Script.thy Complex/ex/MIR.thy Complex/ex/mirtac.ML \ - Complex/ex/ReflectedFerrack.thy \ - Complex/ex/linrtac.ML + ex/BigO_Complex.thy \ + ex/Arithmetic_Series_Complex.thy ex/HarmonicSeries.thy \ + ex/Sqrt.thy \ + ex/Sqrt_Script.thy ex/MIR.thy ex/mirtac.ML \ + ex/ReflectedFerrack.thy \ + ex/linrtac.ML @$(ISABELLE_TOOL) usedir $(OUT)/HOL ex @@ -942,7 +945,7 @@ HOL-Word: HOL $(OUT)/HOL-Word $(OUT)/HOL-Word: $(OUT)/HOL Word/ROOT.ML Library/Infinite_Set.thy \ - Library/Parity.thy Library/Boolean_Algebra.thy \ + Library/Boolean_Algebra.thy \ Library/Numeral_Type.thy Word/Num_Lemmas.thy Word/TdThs.thy \ Word/Size.thy Word/BinGeneral.thy Word/BinOperations.thy \ Word/BinBoolList.thy Word/BitSyntax.thy Word/WordDefinition.thy \ diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/Code_Message.thy --- a/src/HOL/Library/Code_Message.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,58 +0,0 @@ -(* ID: $Id$ - Author: Florian Haftmann, TU Muenchen -*) - -header {* Monolithic strings (message strings) for code generation *} - -theory Code_Message -imports Plain "~~/src/HOL/List" -begin - -subsection {* Datatype of messages *} - -datatype message_string = STR string - -lemmas [code del] = message_string.recs message_string.cases - -lemma [code]: "size (s\message_string) = 0" - by (cases s) simp_all - -lemma [code]: "message_string_size (s\message_string) = 0" - by (cases s) simp_all - -subsection {* ML interface *} - -ML {* -structure Message_String = -struct - -fun mk s = @{term STR} $ HOLogic.mk_string s; - -end; -*} - - -subsection {* Code serialization *} - -code_type message_string - (SML "string") - (OCaml "string") - (Haskell "String") - -setup {* - fold (fn target => add_literal_message @{const_name STR} target) - ["SML", "OCaml", "Haskell"] -*} - -code_reserved SML string -code_reserved OCaml string - -code_instance message_string :: eq - (Haskell -) - -code_const "eq_class.eq \ message_string \ message_string \ bool" - (SML "!((_ : string) = _)") - (OCaml "!((_ : string) = _)") - (Haskell infixl 4 "==") - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/Commutative_Ring.thy --- a/src/HOL/Library/Commutative_Ring.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Library/Commutative_Ring.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ -(* ID: $Id$ - Author: Bernhard Haeupler +(* Author: Bernhard Haeupler Proving equalities in commutative rings done "right" in Isabelle/HOL. *) @@ -7,7 +6,7 @@ header {* Proving equalities in commutative rings *} theory Commutative_Ring -imports Plain "~~/src/HOL/List" Parity +imports Plain "~~/src/HOL/List" "~~/src/HOL/Parity" uses ("comm_ring.ML") begin diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/Float.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Float.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,568 @@ +(* Title: HOL/Real/Float.thy + Author: Steven Obua +*) + +header {* Floating Point Representation of the Reals *} + +theory Float +imports Complex_Main +uses "~~/src/Tools/float.ML" ("~~/src/HOL/Tools/float_arith.ML") +begin + +definition + pow2 :: "int \ real" where + "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))" + +definition + float :: "int * int \ real" where + "float x = real (fst x) * pow2 (snd x)" + +lemma pow2_0[simp]: "pow2 0 = 1" +by (simp add: pow2_def) + +lemma pow2_1[simp]: "pow2 1 = 2" +by (simp add: pow2_def) + +lemma pow2_neg: "pow2 x = inverse (pow2 (-x))" +by (simp add: pow2_def) + +lemma pow2_add1: "pow2 (1 + a) = 2 * (pow2 a)" +proof - + have h: "! n. nat (2 + int n) - Suc 0 = nat (1 + int n)" by arith + have g: "! a b. a - -1 = a + (1::int)" by arith + have pos: "! n. pow2 (int n + 1) = 2 * pow2 (int n)" + apply (auto, induct_tac n) + apply (simp_all add: pow2_def) + apply (rule_tac m1="2" and n1="nat (2 + int na)" in ssubst[OF realpow_num_eq_if]) + by (auto simp add: h) + show ?thesis + proof (induct a) + case (1 n) + from pos show ?case by (simp add: ring_simps) + next + case (2 n) + show ?case + apply (auto) + apply (subst pow2_neg[of "- int n"]) + apply (subst pow2_neg[of "-1 - int n"]) + apply (auto simp add: g pos) + done + qed +qed + +lemma pow2_add: "pow2 (a+b) = (pow2 a) * (pow2 b)" +proof (induct b) + case (1 n) + show ?case + proof (induct n) + case 0 + show ?case by simp + next + case (Suc m) + show ?case by (auto simp add: ring_simps pow2_add1 prems) + qed +next + case (2 n) + show ?case + proof (induct n) + case 0 + show ?case + apply (auto) + apply (subst pow2_neg[of "a + -1"]) + apply (subst pow2_neg[of "-1"]) + apply (simp) + apply (insert pow2_add1[of "-a"]) + apply (simp add: ring_simps) + apply (subst pow2_neg[of "-a"]) + apply (simp) + done + case (Suc m) + have a: "int m - (a + -2) = 1 + (int m - a + 1)" by arith + have b: "int m - -2 = 1 + (int m + 1)" by arith + show ?case + apply (auto) + apply (subst pow2_neg[of "a + (-2 - int m)"]) + apply (subst pow2_neg[of "-2 - int m"]) + apply (auto simp add: ring_simps) + apply (subst a) + apply (subst b) + apply (simp only: pow2_add1) + apply (subst pow2_neg[of "int m - a + 1"]) + apply (subst pow2_neg[of "int m + 1"]) + apply auto + apply (insert prems) + apply (auto simp add: ring_simps) + done + qed +qed + +lemma "float (a, e) + float (b, e) = float (a + b, e)" +by (simp add: float_def ring_simps) + +definition + int_of_real :: "real \ int" where + "int_of_real x = (SOME y. real y = x)" + +definition + real_is_int :: "real \ bool" where + "real_is_int x = (EX (u::int). x = real u)" + +lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))" +by (auto simp add: real_is_int_def int_of_real_def) + +lemma float_transfer: "real_is_int ((real a)*(pow2 c)) \ float (a, b) = float (int_of_real ((real a)*(pow2 c)), b - c)" +by (simp add: float_def real_is_int_def2 pow2_add[symmetric]) + +lemma pow2_int: "pow2 (int c) = 2^c" +by (simp add: pow2_def) + +lemma float_transfer_nat: "float (a, b) = float (a * 2^c, b - int c)" +by (simp add: float_def pow2_int[symmetric] pow2_add[symmetric]) + +lemma real_is_int_real[simp]: "real_is_int (real (x::int))" +by (auto simp add: real_is_int_def int_of_real_def) + +lemma int_of_real_real[simp]: "int_of_real (real x) = x" +by (simp add: int_of_real_def) + +lemma real_int_of_real[simp]: "real_is_int x \ real (int_of_real x) = x" +by (auto simp add: int_of_real_def real_is_int_def) + +lemma real_is_int_add_int_of_real: "real_is_int a \ real_is_int b \ (int_of_real (a+b)) = (int_of_real a) + (int_of_real b)" +by (auto simp add: int_of_real_def real_is_int_def) + +lemma real_is_int_add[simp]: "real_is_int a \ real_is_int b \ real_is_int (a+b)" +apply (subst real_is_int_def2) +apply (simp add: real_is_int_add_int_of_real real_int_of_real) +done + +lemma int_of_real_sub: "real_is_int a \ real_is_int b \ (int_of_real (a-b)) = (int_of_real a) - (int_of_real b)" +by (auto simp add: int_of_real_def real_is_int_def) + +lemma real_is_int_sub[simp]: "real_is_int a \ real_is_int b \ real_is_int (a-b)" +apply (subst real_is_int_def2) +apply (simp add: int_of_real_sub real_int_of_real) +done + +lemma real_is_int_rep: "real_is_int x \ ?! (a::int). real a = x" +by (auto simp add: real_is_int_def) + +lemma int_of_real_mult: + assumes "real_is_int a" "real_is_int b" + shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)" +proof - + from prems have a: "?! (a'::int). real a' = a" by (rule_tac real_is_int_rep, auto) + from prems have b: "?! (b'::int). real b' = b" by (rule_tac real_is_int_rep, auto) + from a obtain a'::int where a':"a = real a'" by auto + from b obtain b'::int where b':"b = real b'" by auto + have r: "real a' * real b' = real (a' * b')" by auto + show ?thesis + apply (simp add: a' b') + apply (subst r) + apply (simp only: int_of_real_real) + done +qed + +lemma real_is_int_mult[simp]: "real_is_int a \ real_is_int b \ real_is_int (a*b)" +apply (subst real_is_int_def2) +apply (simp add: int_of_real_mult) +done + +lemma real_is_int_0[simp]: "real_is_int (0::real)" +by (simp add: real_is_int_def int_of_real_def) + +lemma real_is_int_1[simp]: "real_is_int (1::real)" +proof - + have "real_is_int (1::real) = real_is_int(real (1::int))" by auto + also have "\ = True" by (simp only: real_is_int_real) + ultimately show ?thesis by auto +qed + +lemma real_is_int_n1: "real_is_int (-1::real)" +proof - + have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto + also have "\ = True" by (simp only: real_is_int_real) + ultimately show ?thesis by auto +qed + +lemma real_is_int_number_of[simp]: "real_is_int ((number_of \ int \ real) x)" +proof - + have neg1: "real_is_int (-1::real)" + proof - + have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto + also have "\ = True" by (simp only: real_is_int_real) + ultimately show ?thesis by auto + qed + + { + fix x :: int + have "real_is_int ((number_of \ int \ real) x)" + unfolding number_of_eq + apply (induct x) + apply (induct_tac n) + apply (simp) + apply (simp) + apply (induct_tac n) + apply (simp add: neg1) + proof - + fix n :: nat + assume rn: "(real_is_int (of_int (- (int (Suc n)))))" + have s: "-(int (Suc (Suc n))) = -1 + - (int (Suc n))" by simp + show "real_is_int (of_int (- (int (Suc (Suc n)))))" + apply (simp only: s of_int_add) + apply (rule real_is_int_add) + apply (simp add: neg1) + apply (simp only: rn) + done + qed + } + note Abs_Bin = this + { + fix x :: int + have "? u. x = u" + apply (rule exI[where x = "x"]) + apply (simp) + done + } + then obtain u::int where "x = u" by auto + with Abs_Bin show ?thesis by auto +qed + +lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)" +by (simp add: int_of_real_def) + +lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)" +proof - + have 1: "(1::real) = real (1::int)" by auto + show ?thesis by (simp only: 1 int_of_real_real) +qed + +lemma int_of_real_number_of[simp]: "int_of_real (number_of b) = number_of b" +proof - + have "real_is_int (number_of b)" by simp + then have uu: "?! u::int. number_of b = real u" by (auto simp add: real_is_int_rep) + then obtain u::int where u:"number_of b = real u" by auto + have "number_of b = real ((number_of b)::int)" + by (simp add: number_of_eq real_of_int_def) + have ub: "number_of b = real ((number_of b)::int)" + by (simp add: number_of_eq real_of_int_def) + from uu u ub have unb: "u = number_of b" + by blast + have "int_of_real (number_of b) = u" by (simp add: u) + with unb show ?thesis by simp +qed + +lemma float_transfer_even: "even a \ float (a, b) = float (a div 2, b+1)" + apply (subst float_transfer[where a="a" and b="b" and c="-1", simplified]) + apply (simp_all add: pow2_def even_def real_is_int_def ring_simps) + apply (auto) +proof - + fix q::int + have a:"b - (-1\int) = (1\int) + b" by arith + show "(float (q, (b - (-1\int)))) = (float (q, ((1\int) + b)))" + by (simp add: a) +qed + +lemma int_div_zdiv: "int (a div b) = (int a) div (int b)" +by (rule zdiv_int) + +lemma int_mod_zmod: "int (a mod b) = (int a) mod (int b)" +by (rule zmod_int) + +lemma abs_div_2_less: "a \ 0 \ a \ -1 \ abs((a::int) div 2) < abs a" +by arith + +function norm_float :: "int \ int \ int \ int" where + "norm_float a b = (if a \ 0 \ even a then norm_float (a div 2) (b + 1) + else if a = 0 then (0, 0) else (a, b))" +by auto + +termination by (relation "measure (nat o abs o fst)") + (auto intro: abs_div_2_less) + +lemma norm_float: "float x = float (split norm_float x)" +proof - + { + fix a b :: int + have norm_float_pair: "float (a, b) = float (norm_float a b)" + proof (induct a b rule: norm_float.induct) + case (1 u v) + show ?case + proof cases + assume u: "u \ 0 \ even u" + with prems have ind: "float (u div 2, v + 1) = float (norm_float (u div 2) (v + 1))" by auto + with u have "float (u,v) = float (u div 2, v+1)" by (simp add: float_transfer_even) + then show ?thesis + apply (subst norm_float.simps) + apply (simp add: ind) + done + next + assume "~(u \ 0 \ even u)" + then show ?thesis + by (simp add: prems float_def) + qed + qed + } + note helper = this + have "? a b. x = (a,b)" by auto + then obtain a b where "x = (a, b)" by blast + then show ?thesis by (simp add: helper) +qed + +lemma float_add_l0: "float (0, e) + x = x" + by (simp add: float_def) + +lemma float_add_r0: "x + float (0, e) = x" + by (simp add: float_def) + +lemma float_add: + "float (a1, e1) + float (a2, e2) = + (if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1) + else float (a1*2^(nat (e1-e2))+a2, e2))" + apply (simp add: float_def ring_simps) + apply (auto simp add: pow2_int[symmetric] pow2_add[symmetric]) + done + +lemma float_add_assoc1: + "(x + float (y1, e1)) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x" + by simp + +lemma float_add_assoc2: + "(float (y1, e1) + x) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x" + by simp + +lemma float_add_assoc3: + "float (y1, e1) + (x + float (y2, e2)) = (float (y1, e1) + float (y2, e2)) + x" + by simp + +lemma float_add_assoc4: + "float (y1, e1) + (float (y2, e2) + x) = (float (y1, e1) + float (y2, e2)) + x" + by simp + +lemma float_mult_l0: "float (0, e) * x = float (0, 0)" + by (simp add: float_def) + +lemma float_mult_r0: "x * float (0, e) = float (0, 0)" + by (simp add: float_def) + +definition + lbound :: "real \ real" +where + "lbound x = min 0 x" + +definition + ubound :: "real \ real" +where + "ubound x = max 0 x" + +lemma lbound: "lbound x \ x" + by (simp add: lbound_def) + +lemma ubound: "x \ ubound x" + by (simp add: ubound_def) + +lemma float_mult: + "float (a1, e1) * float (a2, e2) = + (float (a1 * a2, e1 + e2))" + by (simp add: float_def pow2_add) + +lemma float_minus: + "- (float (a,b)) = float (-a, b)" + by (simp add: float_def) + +lemma zero_less_pow2: + "0 < pow2 x" +proof - + { + fix y + have "0 <= y \ 0 < pow2 y" + by (induct y, induct_tac n, simp_all add: pow2_add) + } + note helper=this + show ?thesis + apply (case_tac "0 <= x") + apply (simp add: helper) + apply (subst pow2_neg) + apply (simp add: helper) + done +qed + +lemma zero_le_float: + "(0 <= float (a,b)) = (0 <= a)" + apply (auto simp add: float_def) + apply (auto simp add: zero_le_mult_iff zero_less_pow2) + apply (insert zero_less_pow2[of b]) + apply (simp_all) + done + +lemma float_le_zero: + "(float (a,b) <= 0) = (a <= 0)" + apply (auto simp add: float_def) + apply (auto simp add: mult_le_0_iff) + apply (insert zero_less_pow2[of b]) + apply auto + done + +lemma float_abs: + "abs (float (a,b)) = (if 0 <= a then (float (a,b)) else (float (-a,b)))" + apply (auto simp add: abs_if) + apply (simp_all add: zero_le_float[symmetric, of a b] float_minus) + done + +lemma float_zero: + "float (0, b) = 0" + by (simp add: float_def) + +lemma float_pprt: + "pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0, b)))" + by (auto simp add: zero_le_float float_le_zero float_zero) + +lemma pprt_lbound: "pprt (lbound x) = float (0, 0)" + apply (simp add: float_def) + apply (rule pprt_eq_0) + apply (simp add: lbound_def) + done + +lemma nprt_ubound: "nprt (ubound x) = float (0, 0)" + apply (simp add: float_def) + apply (rule nprt_eq_0) + apply (simp add: ubound_def) + done + +lemma float_nprt: + "nprt (float (a, b)) = (if 0 <= a then (float (0,b)) else (float (a, b)))" + by (auto simp add: zero_le_float float_le_zero float_zero) + +lemma norm_0_1: "(0::_::number_ring) = Numeral0 & (1::_::number_ring) = Numeral1" + by auto + +lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)" + by simp + +lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)" + by simp + +lemma mult_left_one: "1 * a = (a::'a::semiring_1)" + by simp + +lemma mult_right_one: "a * 1 = (a::'a::semiring_1)" + by simp + +lemma int_pow_0: "(a::int)^(Numeral0) = 1" + by simp + +lemma int_pow_1: "(a::int)^(Numeral1) = a" + by simp + +lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0" + by simp + +lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1" + by simp + +lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0" + by simp + +lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1" + by simp + +lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1" + by simp + +lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1" +proof - + have 1:"((-1)::nat) = 0" + by simp + show ?thesis by (simp add: 1) +qed + +lemma fst_cong: "a=a' \ fst (a,b) = fst (a',b)" + by simp + +lemma snd_cong: "b=b' \ snd (a,b) = snd (a,b')" + by simp + +lemma lift_bool: "x \ x=True" + by simp + +lemma nlift_bool: "~x \ x=False" + by simp + +lemma not_false_eq_true: "(~ False) = True" by simp + +lemma not_true_eq_false: "(~ True) = False" by simp + +lemmas binarith = + normalize_bin_simps + pred_bin_simps succ_bin_simps + add_bin_simps minus_bin_simps mult_bin_simps + +lemma int_eq_number_of_eq: + "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)" + by simp + +lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)" + by (simp only: iszero_number_of_Pls) + +lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))" + by simp + +lemma int_iszero_number_of_Bit0: "iszero ((number_of (Int.Bit0 w))::int) = iszero ((number_of w)::int)" + by simp + +lemma int_iszero_number_of_Bit1: "\ iszero ((number_of (Int.Bit1 w))::int)" + by simp + +lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)" + by simp + +lemma int_not_neg_number_of_Pls: "\ (neg (Numeral0::int))" + by simp + +lemma int_neg_number_of_Min: "neg (-1::int)" + by simp + +lemma int_neg_number_of_Bit0: "neg ((number_of (Int.Bit0 w))::int) = neg ((number_of w)::int)" + by simp + +lemma int_neg_number_of_Bit1: "neg ((number_of (Int.Bit1 w))::int) = neg ((number_of w)::int)" + by simp + +lemma int_le_number_of_eq: "(((number_of x)::int) \ number_of y) = (\ neg ((number_of (y + (uminus x)))::int))" + by simp + +lemmas intarithrel = + int_eq_number_of_eq + lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_Bit0 + lift_bool[OF int_iszero_number_of_Bit1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min] + int_neg_number_of_Bit0 int_neg_number_of_Bit1 int_le_number_of_eq + +lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)" + by simp + +lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (v + (uminus w))" + by simp + +lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (v * w)" + by simp + +lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)" + by simp + +lemmas intarith = int_number_of_add_sym int_number_of_minus_sym int_number_of_diff_sym int_number_of_mult_sym + +lemmas natarith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of + +lemmas powerarith = nat_number_of zpower_number_of_even + zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring] + zpower_Pls zpower_Min + +lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0 + float_minus float_abs zero_le_float float_pprt float_nprt pprt_lbound nprt_ubound + +(* for use with the compute oracle *) +lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false + +use "~~/src/HOL/Tools/float_arith.ML" + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/GCD.thy --- a/src/HOL/Library/GCD.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,783 +0,0 @@ -(* Title: HOL/GCD.thy - ID: $Id$ - Author: Christophe Tabacznyj and Lawrence C Paulson - Copyright 1996 University of Cambridge -*) - -header {* The Greatest Common Divisor *} - -theory GCD -imports Plain "~~/src/HOL/Presburger" -begin - -text {* - See \cite{davenport92}. \bigskip -*} - -subsection {* Specification of GCD on nats *} - -definition - is_gcd :: "nat \ nat \ nat \ bool" where -- {* @{term gcd} as a relation *} - [code del]: "is_gcd m n p \ p dvd m \ p dvd n \ - (\d. d dvd m \ d dvd n \ d dvd p)" - -text {* Uniqueness *} - -lemma is_gcd_unique: "is_gcd a b m \ is_gcd a b n \ m = n" - by (simp add: is_gcd_def) (blast intro: dvd_anti_sym) - -text {* Connection to divides relation *} - -lemma is_gcd_dvd: "is_gcd a b m \ k dvd a \ k dvd b \ k dvd m" - by (auto simp add: is_gcd_def) - -text {* Commutativity *} - -lemma is_gcd_commute: "is_gcd m n k = is_gcd n m k" - by (auto simp add: is_gcd_def) - - -subsection {* GCD on nat by Euclid's algorithm *} - -fun - gcd :: "nat => nat => nat" -where - "gcd m n = (if n = 0 then m else gcd n (m mod n))" -lemma gcd_induct [case_names "0" rec]: - fixes m n :: nat - assumes "\m. P m 0" - and "\m n. 0 < n \ P n (m mod n) \ P m n" - shows "P m n" -proof (induct m n rule: gcd.induct) - case (1 m n) with assms show ?case by (cases "n = 0") simp_all -qed - -lemma gcd_0 [simp, algebra]: "gcd m 0 = m" - by simp - -lemma gcd_0_left [simp,algebra]: "gcd 0 m = m" - by simp - -lemma gcd_non_0: "n > 0 \ gcd m n = gcd n (m mod n)" - by simp - -lemma gcd_1 [simp, algebra]: "gcd m (Suc 0) = 1" - by simp - -declare gcd.simps [simp del] - -text {* - \medskip @{term "gcd m n"} divides @{text m} and @{text n}. The - conjunctions don't seem provable separately. -*} - -lemma gcd_dvd1 [iff, algebra]: "gcd m n dvd m" - and gcd_dvd2 [iff, algebra]: "gcd m n dvd n" - apply (induct m n rule: gcd_induct) - apply (simp_all add: gcd_non_0) - apply (blast dest: dvd_mod_imp_dvd) - done - -text {* - \medskip Maximality: for all @{term m}, @{term n}, @{term k} - naturals, if @{term k} divides @{term m} and @{term k} divides - @{term n} then @{term k} divides @{term "gcd m n"}. -*} - -lemma gcd_greatest: "k dvd m \ k dvd n \ k dvd gcd m n" - by (induct m n rule: gcd_induct) (simp_all add: gcd_non_0 dvd_mod) - -text {* - \medskip Function gcd yields the Greatest Common Divisor. -*} - -lemma is_gcd: "is_gcd m n (gcd m n) " - by (simp add: is_gcd_def gcd_greatest) - - -subsection {* Derived laws for GCD *} - -lemma gcd_greatest_iff [iff, algebra]: "k dvd gcd m n \ k dvd m \ k dvd n" - by (blast intro!: gcd_greatest intro: dvd_trans) - -lemma gcd_zero[algebra]: "gcd m n = 0 \ m = 0 \ n = 0" - by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff) - -lemma gcd_commute: "gcd m n = gcd n m" - apply (rule is_gcd_unique) - apply (rule is_gcd) - apply (subst is_gcd_commute) - apply (simp add: is_gcd) - done - -lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)" - apply (rule is_gcd_unique) - apply (rule is_gcd) - apply (simp add: is_gcd_def) - apply (blast intro: dvd_trans) - done - -lemma gcd_1_left [simp, algebra]: "gcd (Suc 0) m = 1" - by (simp add: gcd_commute) - -text {* - \medskip Multiplication laws -*} - -lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)" - -- {* \cite[page 27]{davenport92} *} - apply (induct m n rule: gcd_induct) - apply simp - apply (case_tac "k = 0") - apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2) - done - -lemma gcd_mult [simp, algebra]: "gcd k (k * n) = k" - apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric]) - done - -lemma gcd_self [simp, algebra]: "gcd k k = k" - apply (rule gcd_mult [of k 1, simplified]) - done - -lemma relprime_dvd_mult: "gcd k n = 1 ==> k dvd m * n ==> k dvd m" - apply (insert gcd_mult_distrib2 [of m k n]) - apply simp - apply (erule_tac t = m in ssubst) - apply simp - done - -lemma relprime_dvd_mult_iff: "gcd k n = 1 ==> (k dvd m * n) = (k dvd m)" - by (auto intro: relprime_dvd_mult dvd_mult2) - -lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n" - apply (rule dvd_anti_sym) - apply (rule gcd_greatest) - apply (rule_tac n = k in relprime_dvd_mult) - apply (simp add: gcd_assoc) - apply (simp add: gcd_commute) - apply (simp_all add: mult_commute) - apply (blast intro: dvd_mult) - done - - -text {* \medskip Addition laws *} - -lemma gcd_add1 [simp, algebra]: "gcd (m + n) n = gcd m n" - by (cases "n = 0") (auto simp add: gcd_non_0) - -lemma gcd_add2 [simp, algebra]: "gcd m (m + n) = gcd m n" -proof - - have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute) - also have "... = gcd (n + m) m" by (simp add: add_commute) - also have "... = gcd n m" by simp - also have "... = gcd m n" by (rule gcd_commute) - finally show ?thesis . -qed - -lemma gcd_add2' [simp, algebra]: "gcd m (n + m) = gcd m n" - apply (subst add_commute) - apply (rule gcd_add2) - done - -lemma gcd_add_mult[algebra]: "gcd m (k * m + n) = gcd m n" - by (induct k) (simp_all add: add_assoc) - -lemma gcd_dvd_prod: "gcd m n dvd m * n" - using mult_dvd_mono [of 1] by auto - -text {* - \medskip Division by gcd yields rrelatively primes. -*} - -lemma div_gcd_relprime: - assumes nz: "a \ 0 \ b \ 0" - shows "gcd (a div gcd a b) (b div gcd a b) = 1" -proof - - let ?g = "gcd a b" - let ?a' = "a div ?g" - let ?b' = "b div ?g" - let ?g' = "gcd ?a' ?b'" - have dvdg: "?g dvd a" "?g dvd b" by simp_all - have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all - from dvdg dvdg' obtain ka kb ka' kb' where - kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'" - unfolding dvd_def by blast - then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all - then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b" - by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)] - dvd_mult_div_cancel [OF dvdg(2)] dvd_def) - have "?g \ 0" using nz by (simp add: gcd_zero) - then have gp: "?g > 0" by simp - from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" . - with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp -qed - - -lemma gcd_unique: "d dvd a\d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" -proof(auto) - assume H: "d dvd a" "d dvd b" "\e. e dvd a \ e dvd b \ e dvd d" - from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b] - have th: "gcd a b dvd d" by blast - from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]] show "d = gcd a b" by blast -qed - -lemma gcd_eq: assumes H: "\d. d dvd x \ d dvd y \ d dvd u \ d dvd v" - shows "gcd x y = gcd u v" -proof- - from H have "\d. d dvd x \ d dvd y \ d dvd gcd u v" by simp - with gcd_unique[of "gcd u v" x y] show ?thesis by auto -qed - -lemma ind_euclid: - assumes c: " \a b. P (a::nat) b \ P b a" and z: "\a. P a 0" - and add: "\a b. P a b \ P a (a + b)" - shows "P a b" -proof(induct n\"a+b" arbitrary: a b rule: nat_less_induct) - fix n a b - assume H: "\m < n. \a b. m = a + b \ P a b" "n = a + b" - have "a = b \ a < b \ b < a" by arith - moreover {assume eq: "a= b" - from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq by simp} - moreover - {assume lt: "a < b" - hence "a + b - a < n \ a = 0" using H(2) by arith - moreover - {assume "a =0" with z c have "P a b" by blast } - moreover - {assume ab: "a + b - a < n" - have th0: "a + b - a = a + (b - a)" using lt by arith - from add[rule_format, OF H(1)[rule_format, OF ab th0]] - have "P a b" by (simp add: th0[symmetric])} - ultimately have "P a b" by blast} - moreover - {assume lt: "a > b" - hence "b + a - b < n \ b = 0" using H(2) by arith - moreover - {assume "b =0" with z c have "P a b" by blast } - moreover - {assume ab: "b + a - b < n" - have th0: "b + a - b = b + (a - b)" using lt by arith - from add[rule_format, OF H(1)[rule_format, OF ab th0]] - have "P b a" by (simp add: th0[symmetric]) - hence "P a b" using c by blast } - ultimately have "P a b" by blast} -ultimately show "P a b" by blast -qed - -lemma bezout_lemma: - assumes ex: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x = b * y + d \ b * x = a * y + d)" - shows "\d x y. d dvd a \ d dvd a + b \ (a * x = (a + b) * y + d \ (a + b) * x = a * y + d)" -using ex -apply clarsimp -apply (rule_tac x="d" in exI, simp add: dvd_add) -apply (case_tac "a * x = b * y + d" , simp_all) -apply (rule_tac x="x + y" in exI) -apply (rule_tac x="y" in exI) -apply algebra -apply (rule_tac x="x" in exI) -apply (rule_tac x="x + y" in exI) -apply algebra -done - -lemma bezout_add: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x = b * y + d \ b * x = a * y + d)" -apply(induct a b rule: ind_euclid) -apply blast -apply clarify -apply (rule_tac x="a" in exI, simp add: dvd_add) -apply clarsimp -apply (rule_tac x="d" in exI) -apply (case_tac "a * x = b * y + d", simp_all add: dvd_add) -apply (rule_tac x="x+y" in exI) -apply (rule_tac x="y" in exI) -apply algebra -apply (rule_tac x="x" in exI) -apply (rule_tac x="x+y" in exI) -apply algebra -done - -lemma bezout: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x - b * y = d \ b * x - a * y = d)" -using bezout_add[of a b] -apply clarsimp -apply (rule_tac x="d" in exI, simp) -apply (rule_tac x="x" in exI) -apply (rule_tac x="y" in exI) -apply auto -done - - -text {* We can get a stronger version with a nonzeroness assumption. *} -lemma divides_le: "m dvd n ==> m <= n \ n = (0::nat)" by (auto simp add: dvd_def) - -lemma bezout_add_strong: assumes nz: "a \ (0::nat)" - shows "\d x y. d dvd a \ d dvd b \ a * x = b * y + d" -proof- - from nz have ap: "a > 0" by simp - from bezout_add[of a b] - have "(\d x y. d dvd a \ d dvd b \ a * x = b * y + d) \ (\d x y. d dvd a \ d dvd b \ b * x = a * y + d)" by blast - moreover - {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d" - from H have ?thesis by blast } - moreover - {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d" - {assume b0: "b = 0" with H have ?thesis by simp} - moreover - {assume b: "b \ 0" hence bp: "b > 0" by simp - from divides_le[OF H(2)] b have "d < b \ d = b" using le_less by blast - moreover - {assume db: "d=b" - from prems have ?thesis apply simp - apply (rule exI[where x = b], simp) - apply (rule exI[where x = b]) - by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)} - moreover - {assume db: "d < b" - {assume "x=0" hence ?thesis using prems by simp } - moreover - {assume x0: "x \ 0" hence xp: "x > 0" by simp - - from db have "d \ b - 1" by simp - hence "d*b \ b*(b - 1)" by simp - with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"] - have dble: "d*b \ x*b*(b - 1)" using bp by simp - from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra - hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp - hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)" - by (simp only: diff_add_assoc[OF dble, of d, symmetric]) - hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d" - by (simp only: diff_mult_distrib2 add_commute mult_ac) - hence ?thesis using H(1,2) - apply - - apply (rule exI[where x=d], simp) - apply (rule exI[where x="(b - 1) * y"]) - by (rule exI[where x="x*(b - 1) - d"], simp)} - ultimately have ?thesis by blast} - ultimately have ?thesis by blast} - ultimately have ?thesis by blast} - ultimately show ?thesis by blast -qed - - -lemma bezout_gcd: "\x y. a * x - b * y = gcd a b \ b * x - a * y = gcd a b" -proof- - let ?g = "gcd a b" - from bezout[of a b] obtain d x y where d: "d dvd a" "d dvd b" "a * x - b * y = d \ b * x - a * y = d" by blast - from d(1,2) have "d dvd ?g" by simp - then obtain k where k: "?g = d*k" unfolding dvd_def by blast - from d(3) have "(a * x - b * y)*k = d*k \ (b * x - a * y)*k = d*k" by blast - hence "a * x * k - b * y*k = d*k \ b * x * k - a * y*k = d*k" - by (algebra add: diff_mult_distrib) - hence "a * (x * k) - b * (y*k) = ?g \ b * (x * k) - a * (y*k) = ?g" - by (simp add: k mult_assoc) - thus ?thesis by blast -qed - -lemma bezout_gcd_strong: assumes a: "a \ 0" - shows "\x y. a * x = b * y + gcd a b" -proof- - let ?g = "gcd a b" - from bezout_add_strong[OF a, of b] - obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast - from d(1,2) have "d dvd ?g" by simp - then obtain k where k: "?g = d*k" unfolding dvd_def by blast - from d(3) have "a * x * k = (b * y + d) *k " by algebra - hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k) - thus ?thesis by blast -qed - -lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b" -by(simp add: gcd_mult_distrib2 mult_commute) - -lemma gcd_bezout: "(\x y. a * x - b * y = d \ b * x - a * y = d) \ gcd a b dvd d" - (is "?lhs \ ?rhs") -proof- - let ?g = "gcd a b" - {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast - from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \ b * x - a * y = ?g" - by blast - hence "(a * x - b * y)*k = ?g*k \ (b * x - a * y)*k = ?g*k" by auto - hence "a * x*k - b * y*k = ?g*k \ b * x * k - a * y*k = ?g*k" - by (simp only: diff_mult_distrib) - hence "a * (x*k) - b * (y*k) = d \ b * (x * k) - a * (y*k) = d" - by (simp add: k[symmetric] mult_assoc) - hence ?lhs by blast} - moreover - {fix x y assume H: "a * x - b * y = d \ b * x - a * y = d" - have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y" - using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all - from dvd_diff[OF dv(1,2)] dvd_diff[OF dv(3,4)] H - have ?rhs by auto} - ultimately show ?thesis by blast -qed - -lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd a b dvd d" -proof- - let ?g = "gcd a b" - have dv: "?g dvd a*x" "?g dvd b * y" - using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all - from dvd_add[OF dv] H - show ?thesis by auto -qed - -lemma gcd_mult': "gcd b (a * b) = b" -by (simp add: gcd_mult mult_commute[of a b]) - -lemma gcd_add: "gcd(a + b) b = gcd a b" - "gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b" -apply (simp_all add: gcd_add1) -by (simp add: gcd_commute gcd_add1) - -lemma gcd_sub: "b <= a ==> gcd(a - b) b = gcd a b" "a <= b ==> gcd a (b - a) = gcd a b" -proof- - {fix a b assume H: "b \ (a::nat)" - hence th: "a - b + b = a" by arith - from gcd_add(1)[of "a - b" b] th have "gcd(a - b) b = gcd a b" by simp} - note th = this -{ - assume ab: "b \ a" - from th[OF ab] show "gcd (a - b) b = gcd a b" by blast -next - assume ab: "a \ b" - from th[OF ab] show "gcd a (b - a) = gcd a b" - by (simp add: gcd_commute)} -qed - - -subsection {* LCM defined by GCD *} - - -definition - lcm :: "nat \ nat \ nat" -where - lcm_def: "lcm m n = m * n div gcd m n" - -lemma prod_gcd_lcm: - "m * n = gcd m n * lcm m n" - unfolding lcm_def by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod]) - -lemma lcm_0 [simp]: "lcm m 0 = 0" - unfolding lcm_def by simp - -lemma lcm_1 [simp]: "lcm m 1 = m" - unfolding lcm_def by simp - -lemma lcm_0_left [simp]: "lcm 0 n = 0" - unfolding lcm_def by simp - -lemma lcm_1_left [simp]: "lcm 1 m = m" - unfolding lcm_def by simp - -lemma dvd_pos: - fixes n m :: nat - assumes "n > 0" and "m dvd n" - shows "m > 0" -using assms by (cases m) auto - -lemma lcm_least: - assumes "m dvd k" and "n dvd k" - shows "lcm m n dvd k" -proof (cases k) - case 0 then show ?thesis by auto -next - case (Suc _) then have pos_k: "k > 0" by auto - from assms dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto - with gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp - from assms obtain p where k_m: "k = m * p" using dvd_def by blast - from assms obtain q where k_n: "k = n * q" using dvd_def by blast - from pos_k k_m have pos_p: "p > 0" by auto - from pos_k k_n have pos_q: "q > 0" by auto - have "k * k * gcd q p = k * gcd (k * q) (k * p)" - by (simp add: mult_ac gcd_mult_distrib2) - also have "\ = k * gcd (m * p * q) (n * q * p)" - by (simp add: k_m [symmetric] k_n [symmetric]) - also have "\ = k * p * q * gcd m n" - by (simp add: mult_ac gcd_mult_distrib2) - finally have "(m * p) * (n * q) * gcd q p = k * p * q * gcd m n" - by (simp only: k_m [symmetric] k_n [symmetric]) - then have "p * q * m * n * gcd q p = p * q * k * gcd m n" - by (simp add: mult_ac) - with pos_p pos_q have "m * n * gcd q p = k * gcd m n" - by simp - with prod_gcd_lcm [of m n] - have "lcm m n * gcd q p * gcd m n = k * gcd m n" - by (simp add: mult_ac) - with pos_gcd have "lcm m n * gcd q p = k" by simp - then show ?thesis using dvd_def by auto -qed - -lemma lcm_dvd1 [iff]: - "m dvd lcm m n" -proof (cases m) - case 0 then show ?thesis by simp -next - case (Suc _) - then have mpos: "m > 0" by simp - show ?thesis - proof (cases n) - case 0 then show ?thesis by simp - next - case (Suc _) - then have npos: "n > 0" by simp - have "gcd m n dvd n" by simp - then obtain k where "n = gcd m n * k" using dvd_def by auto - then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n" by (simp add: mult_ac) - also have "\ = m * k" using mpos npos gcd_zero by simp - finally show ?thesis by (simp add: lcm_def) - qed -qed - -lemma lcm_dvd2 [iff]: - "n dvd lcm m n" -proof (cases n) - case 0 then show ?thesis by simp -next - case (Suc _) - then have npos: "n > 0" by simp - show ?thesis - proof (cases m) - case 0 then show ?thesis by simp - next - case (Suc _) - then have mpos: "m > 0" by simp - have "gcd m n dvd m" by simp - then obtain k where "m = gcd m n * k" using dvd_def by auto - then have "m * n div gcd m n = (gcd m n * k) * n div gcd m n" by (simp add: mult_ac) - also have "\ = n * k" using mpos npos gcd_zero by simp - finally show ?thesis by (simp add: lcm_def) - qed -qed - -lemma gcd_add1_eq: "gcd (m + k) k = gcd (m + k) m" - by (simp add: gcd_commute) - -lemma gcd_diff2: "m \ n ==> gcd n (n - m) = gcd n m" - apply (subgoal_tac "n = m + (n - m)") - apply (erule ssubst, rule gcd_add1_eq, simp) - done - - -subsection {* GCD and LCM on integers *} - -definition - zgcd :: "int \ int \ int" where - "zgcd i j = int (gcd (nat (abs i)) (nat (abs j)))" - -lemma zgcd_zdvd1 [iff,simp, algebra]: "zgcd i j dvd i" - by (simp add: zgcd_def int_dvd_iff) - -lemma zgcd_zdvd2 [iff,simp, algebra]: "zgcd i j dvd j" - by (simp add: zgcd_def int_dvd_iff) - -lemma zgcd_pos: "zgcd i j \ 0" - by (simp add: zgcd_def) - -lemma zgcd0 [simp,algebra]: "(zgcd i j = 0) = (i = 0 \ j = 0)" - by (simp add: zgcd_def gcd_zero) arith - -lemma zgcd_commute: "zgcd i j = zgcd j i" - unfolding zgcd_def by (simp add: gcd_commute) - -lemma zgcd_zminus [simp, algebra]: "zgcd (- i) j = zgcd i j" - unfolding zgcd_def by simp - -lemma zgcd_zminus2 [simp, algebra]: "zgcd i (- j) = zgcd i j" - unfolding zgcd_def by simp - - (* should be solved by algebra*) -lemma zrelprime_dvd_mult: "zgcd i j = 1 \ i dvd k * j \ i dvd k" - unfolding zgcd_def -proof - - assume "int (gcd (nat \i\) (nat \j\)) = 1" "i dvd k * j" - then have g: "gcd (nat \i\) (nat \j\) = 1" by simp - from `i dvd k * j` obtain h where h: "k*j = i*h" unfolding dvd_def by blast - have th: "nat \i\ dvd nat \k\ * nat \j\" - unfolding dvd_def - by (rule_tac x= "nat \h\" in exI, simp add: h nat_abs_mult_distrib [symmetric]) - from relprime_dvd_mult [OF g th] obtain h' where h': "nat \k\ = nat \i\ * h'" - unfolding dvd_def by blast - from h' have "int (nat \k\) = int (nat \i\ * h')" by simp - then have "\k\ = \i\ * int h'" by (simp add: int_mult) - then show ?thesis - apply (subst zdvd_abs1 [symmetric]) - apply (subst zdvd_abs2 [symmetric]) - apply (unfold dvd_def) - apply (rule_tac x = "int h'" in exI, simp) - done -qed - -lemma int_nat_abs: "int (nat (abs x)) = abs x" by arith - -lemma zgcd_greatest: - assumes "k dvd m" and "k dvd n" - shows "k dvd zgcd m n" -proof - - let ?k' = "nat \k\" - let ?m' = "nat \m\" - let ?n' = "nat \n\" - from `k dvd m` and `k dvd n` have dvd': "?k' dvd ?m'" "?k' dvd ?n'" - unfolding zdvd_int by (simp_all only: int_nat_abs zdvd_abs1 zdvd_abs2) - from gcd_greatest [OF dvd'] have "int (nat \k\) dvd zgcd m n" - unfolding zgcd_def by (simp only: zdvd_int) - then have "\k\ dvd zgcd m n" by (simp only: int_nat_abs) - then show "k dvd zgcd m n" by (simp add: zdvd_abs1) -qed - -lemma div_zgcd_relprime: - assumes nz: "a \ 0 \ b \ 0" - shows "zgcd (a div (zgcd a b)) (b div (zgcd a b)) = 1" -proof - - from nz have nz': "nat \a\ \ 0 \ nat \b\ \ 0" by arith - let ?g = "zgcd a b" - let ?a' = "a div ?g" - let ?b' = "b div ?g" - let ?g' = "zgcd ?a' ?b'" - have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2) - have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2) - from dvdg dvdg' obtain ka kb ka' kb' where - kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'" - unfolding dvd_def by blast - then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all - then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b" - by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)] - zdvd_mult_div_cancel [OF dvdg(2)] dvd_def) - have "?g \ 0" using nz by simp - then have gp: "?g \ 0" using zgcd_pos[where i="a" and j="b"] by arith - from zgcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" . - with zdvd_mult_cancel1 [OF gp] have "\?g'\ = 1" by simp - with zgcd_pos show "?g' = 1" by simp -qed - -lemma zgcd_0 [simp, algebra]: "zgcd m 0 = abs m" - by (simp add: zgcd_def abs_if) - -lemma zgcd_0_left [simp, algebra]: "zgcd 0 m = abs m" - by (simp add: zgcd_def abs_if) - -lemma zgcd_non_0: "0 < n ==> zgcd m n = zgcd n (m mod n)" - apply (frule_tac b = n and a = m in pos_mod_sign) - apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib) - apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if) - apply (frule_tac a = m in pos_mod_bound) - apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle) - done - -lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)" - apply (case_tac "n = 0", simp add: DIVISION_BY_ZERO) - apply (auto simp add: linorder_neq_iff zgcd_non_0) - apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0, auto) - done - -lemma zgcd_1 [simp, algebra]: "zgcd m 1 = 1" - by (simp add: zgcd_def abs_if) - -lemma zgcd_0_1_iff [simp, algebra]: "zgcd 0 m = 1 \ \m\ = 1" - by (simp add: zgcd_def abs_if) - -lemma zgcd_greatest_iff[algebra]: "k dvd zgcd m n = (k dvd m \ k dvd n)" - by (simp add: zgcd_def abs_if int_dvd_iff dvd_int_iff nat_dvd_iff) - -lemma zgcd_1_left [simp, algebra]: "zgcd 1 m = 1" - by (simp add: zgcd_def gcd_1_left) - -lemma zgcd_assoc: "zgcd (zgcd k m) n = zgcd k (zgcd m n)" - by (simp add: zgcd_def gcd_assoc) - -lemma zgcd_left_commute: "zgcd k (zgcd m n) = zgcd m (zgcd k n)" - apply (rule zgcd_commute [THEN trans]) - apply (rule zgcd_assoc [THEN trans]) - apply (rule zgcd_commute [THEN arg_cong]) - done - -lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute - -- {* addition is an AC-operator *} - -lemma zgcd_zmult_distrib2: "0 \ k ==> k * zgcd m n = zgcd (k * m) (k * n)" - by (simp del: minus_mult_right [symmetric] - add: minus_mult_right nat_mult_distrib zgcd_def abs_if - mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric]) - -lemma zgcd_zmult_distrib2_abs: "zgcd (k * m) (k * n) = abs k * zgcd m n" - by (simp add: abs_if zgcd_zmult_distrib2) - -lemma zgcd_self [simp]: "0 \ m ==> zgcd m m = m" - by (cut_tac k = m and m = 1 and n = 1 in zgcd_zmult_distrib2, simp_all) - -lemma zgcd_zmult_eq_self [simp]: "0 \ k ==> zgcd k (k * n) = k" - by (cut_tac k = k and m = 1 and n = n in zgcd_zmult_distrib2, simp_all) - -lemma zgcd_zmult_eq_self2 [simp]: "0 \ k ==> zgcd (k * n) k = k" - by (cut_tac k = k and m = n and n = 1 in zgcd_zmult_distrib2, simp_all) - - -definition "zlcm i j = int (lcm(nat(abs i)) (nat(abs j)))" - -lemma dvd_zlcm_self1[simp, algebra]: "i dvd zlcm i j" -by(simp add:zlcm_def dvd_int_iff) - -lemma dvd_zlcm_self2[simp, algebra]: "j dvd zlcm i j" -by(simp add:zlcm_def dvd_int_iff) - - -lemma dvd_imp_dvd_zlcm1: - assumes "k dvd i" shows "k dvd (zlcm i j)" -proof - - have "nat(abs k) dvd nat(abs i)" using `k dvd i` - by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric] zdvd_abs1) - thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans) -qed - -lemma dvd_imp_dvd_zlcm2: - assumes "k dvd j" shows "k dvd (zlcm i j)" -proof - - have "nat(abs k) dvd nat(abs j)" using `k dvd j` - by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric] zdvd_abs1) - thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans) -qed - - -lemma zdvd_self_abs1: "(d::int) dvd (abs d)" -by (case_tac "d <0", simp_all) - -lemma zdvd_self_abs2: "(abs (d::int)) dvd d" -by (case_tac "d<0", simp_all) - -(* lcm a b is positive for positive a and b *) - -lemma lcm_pos: - assumes mpos: "m > 0" - and npos: "n>0" - shows "lcm m n > 0" -proof(rule ccontr, simp add: lcm_def gcd_zero) -assume h:"m*n div gcd m n = 0" -from mpos npos have "gcd m n \ 0" using gcd_zero by simp -hence gcdp: "gcd m n > 0" by simp -with h -have "m*n < gcd m n" - by (cases "m * n < gcd m n") (auto simp add: div_if[OF gcdp, where m="m*n"]) -moreover -have "gcd m n dvd m" by simp - with mpos dvd_imp_le have t1:"gcd m n \ m" by simp - with npos have t1:"gcd m n *n \ m*n" by simp - have "gcd m n \ gcd m n*n" using npos by simp - with t1 have "gcd m n \ m*n" by arith -ultimately show "False" by simp -qed - -lemma zlcm_pos: - assumes anz: "a \ 0" - and bnz: "b \ 0" - shows "0 < zlcm a b" -proof- - let ?na = "nat (abs a)" - let ?nb = "nat (abs b)" - have nap: "?na >0" using anz by simp - have nbp: "?nb >0" using bnz by simp - have "0 < lcm ?na ?nb" by (rule lcm_pos[OF nap nbp]) - thus ?thesis by (simp add: zlcm_def) -qed - -lemma zgcd_code [code]: - "zgcd k l = \if l = 0 then k else zgcd l (\k\ mod \l\)\" - by (simp add: zgcd_def gcd.simps [of "nat \k\"] nat_mod_distrib) - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/Heap.thy --- a/src/HOL/Library/Heap.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Library/Heap.thy Wed Dec 03 15:58:44 2008 +0100 @@ -6,7 +6,7 @@ header {* A polymorphic heap based on cantor encodings *} theory Heap -imports Plain "~~/src/HOL/List" Countable RType +imports Plain "~~/src/HOL/List" Countable Typerep begin subsection {* Representable types *} @@ -41,7 +41,7 @@ begin fun to_nat_typerep :: "typerep \ nat" where - "to_nat_typerep (RType.Typerep c ts) = to_nat (to_nat c, to_nat (map to_nat_typerep ts))" + "to_nat_typerep (Typerep.Typerep c ts) = to_nat (to_nat c, to_nat (map to_nat_typerep ts))" instance proof (rule countable_classI) @@ -52,11 +52,11 @@ case (Typerep c ts) show ?case proof (rule allI, rule impI) fix t' - assume hyp: "to_nat_typerep (RType.Typerep c ts) = to_nat_typerep t'" - then obtain c' ts' where t': "t' = (RType.Typerep c' ts')" + assume hyp: "to_nat_typerep (Typerep.Typerep c ts) = to_nat_typerep t'" + then obtain c' ts' where t': "t' = (Typerep.Typerep c' ts')" by (cases t') auto with Typerep hyp have "c = c'" and "ts = ts'" by simp_all - with t' show "RType.Typerep c ts = t'" by simp + with t' show "Typerep.Typerep c ts = t'" by simp qed next case Nil_typerep then show ?case by simp diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Library/Library.thy Wed Dec 03 15:58:44 2008 +0100 @@ -11,7 +11,6 @@ Code_Char_chr Code_Index Code_Integer - Code_Message Coinductive_List Commutative_Ring Continuity @@ -21,7 +20,7 @@ Enum Eval_Witness Executable_Set - "../Real/Float" + Float FuncSet Imperative_HOL Infinite_Set diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/Order_Relation.thy --- a/src/HOL/Library/Order_Relation.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,131 +0,0 @@ -(* ID : $Id$ - Author : Tobias Nipkow -*) - -header {* Orders as Relations *} - -theory Order_Relation -imports Plain "~~/src/HOL/Hilbert_Choice" "~~/src/HOL/ATP_Linkup" -begin - -text{* This prelude could be moved to theory Relation: *} - -definition "irrefl r \ \x. (x,x) \ r" - -definition "total_on A r \ \x\A.\y\A. x\y \ (x,y)\r \ (y,x)\r" - -abbreviation "total \ total_on UNIV" - - -lemma total_on_empty[simp]: "total_on {} r" -by(simp add:total_on_def) - -lemma refl_on_converse[simp]: "refl A (r^-1) = refl A r" -by(auto simp add:refl_def) - -lemma total_on_converse[simp]: "total_on A (r^-1) = total_on A r" -by (auto simp: total_on_def) - -lemma irrefl_diff_Id[simp]: "irrefl(r-Id)" -by(simp add:irrefl_def) - -declare [[simp_depth_limit = 2]] -lemma trans_diff_Id: " trans r \ antisym r \ trans (r-Id)" -by(simp add: antisym_def trans_def) blast -declare [[simp_depth_limit = 50]] - -lemma total_on_diff_Id[simp]: "total_on A (r-Id) = total_on A r" -by(simp add: total_on_def) - - -subsection{* Orders on a set *} - -definition "preorder_on A r \ refl A r \ trans r" - -definition "partial_order_on A r \ preorder_on A r \ antisym r" - -definition "linear_order_on A r \ partial_order_on A r \ total_on A r" - -definition "strict_linear_order_on A r \ trans r \ irrefl r \ total_on A r" - -definition "well_order_on A r \ linear_order_on A r \ wf(r - Id)" - -lemmas order_on_defs = - preorder_on_def partial_order_on_def linear_order_on_def - strict_linear_order_on_def well_order_on_def - - -lemma preorder_on_empty[simp]: "preorder_on {} {}" -by(simp add:preorder_on_def trans_def) - -lemma partial_order_on_empty[simp]: "partial_order_on {} {}" -by(simp add:partial_order_on_def) - -lemma lnear_order_on_empty[simp]: "linear_order_on {} {}" -by(simp add:linear_order_on_def) - -lemma well_order_on_empty[simp]: "well_order_on {} {}" -by(simp add:well_order_on_def) - - -lemma preorder_on_converse[simp]: "preorder_on A (r^-1) = preorder_on A r" -by (simp add:preorder_on_def) - -lemma partial_order_on_converse[simp]: - "partial_order_on A (r^-1) = partial_order_on A r" -by (simp add: partial_order_on_def) - -lemma linear_order_on_converse[simp]: - "linear_order_on A (r^-1) = linear_order_on A r" -by (simp add: linear_order_on_def) - - -lemma strict_linear_order_on_diff_Id: - "linear_order_on A r \ strict_linear_order_on A (r-Id)" -by(simp add: order_on_defs trans_diff_Id) - - -subsection{* Orders on the field *} - -abbreviation "Refl r \ refl (Field r) r" - -abbreviation "Preorder r \ preorder_on (Field r) r" - -abbreviation "Partial_order r \ partial_order_on (Field r) r" - -abbreviation "Total r \ total_on (Field r) r" - -abbreviation "Linear_order r \ linear_order_on (Field r) r" - -abbreviation "Well_order r \ well_order_on (Field r) r" - - -lemma subset_Image_Image_iff: - "\ Preorder r; A \ Field r; B \ Field r\ \ - r `` A \ r `` B \ (\a\A.\b\B. (b,a):r)" -apply(auto simp add: subset_eq preorder_on_def refl_def Image_def) -apply metis -by(metis trans_def) - -lemma subset_Image1_Image1_iff: - "\ Preorder r; a : Field r; b : Field r\ \ r `` {a} \ r `` {b} \ (b,a):r" -by(simp add:subset_Image_Image_iff) - -lemma Refl_antisym_eq_Image1_Image1_iff: - "\Refl r; antisym r; a:Field r; b:Field r\ \ r `` {a} = r `` {b} \ a=b" -by(simp add: expand_set_eq antisym_def refl_def) metis - -lemma Partial_order_eq_Image1_Image1_iff: - "\Partial_order r; a:Field r; b:Field r\ \ r `` {a} = r `` {b} \ a=b" -by(auto simp:order_on_defs Refl_antisym_eq_Image1_Image1_iff) - - -subsection{* Orders on a type *} - -abbreviation "strict_linear_order \ strict_linear_order_on UNIV" - -abbreviation "linear_order \ linear_order_on UNIV" - -abbreviation "well_order r \ well_order_on UNIV" - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/Parity.thy --- a/src/HOL/Library/Parity.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,424 +0,0 @@ -(* Title: HOL/Library/Parity.thy - ID: $Id$ - Author: Jeremy Avigad, Jacques D. Fleuriot -*) - -header {* Even and Odd for int and nat *} - -theory Parity -imports Plain "~~/src/HOL/Presburger" -begin - -class even_odd = type + - fixes even :: "'a \ bool" - -abbreviation - odd :: "'a\even_odd \ bool" where - "odd x \ \ even x" - -instantiation nat and int :: even_odd -begin - -definition - even_def [presburger]: "even x \ (x\int) mod 2 = 0" - -definition - even_nat_def [presburger]: "even x \ even (int x)" - -instance .. - -end - - -subsection {* Even and odd are mutually exclusive *} - -lemma int_pos_lt_two_imp_zero_or_one: - "0 <= x ==> (x::int) < 2 ==> x = 0 | x = 1" - by presburger - -lemma neq_one_mod_two [simp, presburger]: - "((x::int) mod 2 ~= 0) = (x mod 2 = 1)" by presburger - - -subsection {* Behavior under integer arithmetic operations *} -declare dvd_def[algebra] -lemma nat_even_iff_2_dvd[algebra]: "even (x::nat) \ 2 dvd x" - by (presburger add: even_nat_def even_def) -lemma int_even_iff_2_dvd[algebra]: "even (x::int) \ 2 dvd x" - by presburger - -lemma even_times_anything: "even (x::int) ==> even (x * y)" - by algebra - -lemma anything_times_even: "even (y::int) ==> even (x * y)" by algebra - -lemma odd_times_odd: "odd (x::int) ==> odd y ==> odd (x * y)" - by (simp add: even_def zmod_zmult1_eq) - -lemma even_product[presburger]: "even((x::int) * y) = (even x | even y)" - apply (auto simp add: even_times_anything anything_times_even) - apply (rule ccontr) - apply (auto simp add: odd_times_odd) - done - -lemma even_plus_even: "even (x::int) ==> even y ==> even (x + y)" - by presburger - -lemma even_plus_odd: "even (x::int) ==> odd y ==> odd (x + y)" - by presburger - -lemma odd_plus_even: "odd (x::int) ==> even y ==> odd (x + y)" - by presburger - -lemma odd_plus_odd: "odd (x::int) ==> odd y ==> even (x + y)" by presburger - -lemma even_sum[presburger]: "even ((x::int) + y) = ((even x & even y) | (odd x & odd y))" - by presburger - -lemma even_neg[presburger, algebra]: "even (-(x::int)) = even x" by presburger - -lemma even_difference: - "even ((x::int) - y) = ((even x & even y) | (odd x & odd y))" by presburger - -lemma even_pow_gt_zero: - "even (x::int) ==> 0 < n ==> even (x^n)" - by (induct n) (auto simp add: even_product) - -lemma odd_pow_iff[presburger, algebra]: - "odd ((x::int) ^ n) \ (n = 0 \ odd x)" - apply (induct n, simp_all) - apply presburger - apply (case_tac n, auto) - apply (simp_all add: even_product) - done - -lemma odd_pow: "odd x ==> odd((x::int)^n)" by (simp add: odd_pow_iff) - -lemma even_power[presburger]: "even ((x::int)^n) = (even x & 0 < n)" - apply (auto simp add: even_pow_gt_zero) - apply (erule contrapos_pp, erule odd_pow) - apply (erule contrapos_pp, simp add: even_def) - done - -lemma even_zero[presburger]: "even (0::int)" by presburger - -lemma odd_one[presburger]: "odd (1::int)" by presburger - -lemmas even_odd_simps [simp] = even_def[of "number_of v",standard] even_zero - odd_one even_product even_sum even_neg even_difference even_power - - -subsection {* Equivalent definitions *} - -lemma two_times_even_div_two: "even (x::int) ==> 2 * (x div 2) = x" - by presburger - -lemma two_times_odd_div_two_plus_one: "odd (x::int) ==> - 2 * (x div 2) + 1 = x" by presburger - -lemma even_equiv_def: "even (x::int) = (EX y. x = 2 * y)" by presburger - -lemma odd_equiv_def: "odd (x::int) = (EX y. x = 2 * y + 1)" by presburger - -subsection {* even and odd for nats *} - -lemma pos_int_even_equiv_nat_even: "0 \ x ==> even x = even (nat x)" - by (simp add: even_nat_def) - -lemma even_nat_product[presburger, algebra]: "even((x::nat) * y) = (even x | even y)" - by (simp add: even_nat_def int_mult) - -lemma even_nat_sum[presburger, algebra]: "even ((x::nat) + y) = - ((even x & even y) | (odd x & odd y))" by presburger - -lemma even_nat_difference[presburger, algebra]: - "even ((x::nat) - y) = (x < y | (even x & even y) | (odd x & odd y))" -by presburger - -lemma even_nat_Suc[presburger, algebra]: "even (Suc x) = odd x" by presburger - -lemma even_nat_power[presburger, algebra]: "even ((x::nat)^y) = (even x & 0 < y)" - by (simp add: even_nat_def int_power) - -lemma even_nat_zero[presburger]: "even (0::nat)" by presburger - -lemmas even_odd_nat_simps [simp] = even_nat_def[of "number_of v",standard] - even_nat_zero even_nat_Suc even_nat_product even_nat_sum even_nat_power - - -subsection {* Equivalent definitions *} - -lemma nat_lt_two_imp_zero_or_one: "(x::nat) < Suc (Suc 0) ==> - x = 0 | x = Suc 0" by presburger - -lemma even_nat_mod_two_eq_zero: "even (x::nat) ==> x mod (Suc (Suc 0)) = 0" - by presburger - -lemma odd_nat_mod_two_eq_one: "odd (x::nat) ==> x mod (Suc (Suc 0)) = Suc 0" -by presburger - -lemma even_nat_equiv_def: "even (x::nat) = (x mod Suc (Suc 0) = 0)" - by presburger - -lemma odd_nat_equiv_def: "odd (x::nat) = (x mod Suc (Suc 0) = Suc 0)" - by presburger - -lemma even_nat_div_two_times_two: "even (x::nat) ==> - Suc (Suc 0) * (x div Suc (Suc 0)) = x" by presburger - -lemma odd_nat_div_two_times_two_plus_one: "odd (x::nat) ==> - Suc( Suc (Suc 0) * (x div Suc (Suc 0))) = x" by presburger - -lemma even_nat_equiv_def2: "even (x::nat) = (EX y. x = Suc (Suc 0) * y)" - by presburger - -lemma odd_nat_equiv_def2: "odd (x::nat) = (EX y. x = Suc(Suc (Suc 0) * y))" - by presburger - - -subsection {* Parity and powers *} - -lemma minus_one_even_odd_power: - "(even x --> (- 1::'a::{comm_ring_1,recpower})^x = 1) & - (odd x --> (- 1::'a)^x = - 1)" - apply (induct x) - apply (rule conjI) - apply simp - apply (insert even_nat_zero, blast) - apply (simp add: power_Suc) - done - -lemma minus_one_even_power [simp]: - "even x ==> (- 1::'a::{comm_ring_1,recpower})^x = 1" - using minus_one_even_odd_power by blast - -lemma minus_one_odd_power [simp]: - "odd x ==> (- 1::'a::{comm_ring_1,recpower})^x = - 1" - using minus_one_even_odd_power by blast - -lemma neg_one_even_odd_power: - "(even x --> (-1::'a::{number_ring,recpower})^x = 1) & - (odd x --> (-1::'a)^x = -1)" - apply (induct x) - apply (simp, simp add: power_Suc) - done - -lemma neg_one_even_power [simp]: - "even x ==> (-1::'a::{number_ring,recpower})^x = 1" - using neg_one_even_odd_power by blast - -lemma neg_one_odd_power [simp]: - "odd x ==> (-1::'a::{number_ring,recpower})^x = -1" - using neg_one_even_odd_power by blast - -lemma neg_power_if: - "(-x::'a::{comm_ring_1,recpower}) ^ n = - (if even n then (x ^ n) else -(x ^ n))" - apply (induct n) - apply (simp_all split: split_if_asm add: power_Suc) - done - -lemma zero_le_even_power: "even n ==> - 0 <= (x::'a::{recpower,ordered_ring_strict}) ^ n" - apply (simp add: even_nat_equiv_def2) - apply (erule exE) - apply (erule ssubst) - apply (subst power_add) - apply (rule zero_le_square) - done - -lemma zero_le_odd_power: "odd n ==> - (0 <= (x::'a::{recpower,ordered_idom}) ^ n) = (0 <= x)" - apply (simp add: odd_nat_equiv_def2) - apply (erule exE) - apply (erule ssubst) - apply (subst power_Suc) - apply (subst power_add) - apply (subst zero_le_mult_iff) - apply auto - apply (subgoal_tac "x = 0 & y > 0") - apply (erule conjE, assumption) - apply (subst power_eq_0_iff [symmetric]) - apply (subgoal_tac "0 <= x^y * x^y") - apply simp - apply (rule zero_le_square)+ - done - -lemma zero_le_power_eq[presburger]: "(0 <= (x::'a::{recpower,ordered_idom}) ^ n) = - (even n | (odd n & 0 <= x))" - apply auto - apply (subst zero_le_odd_power [symmetric]) - apply assumption+ - apply (erule zero_le_even_power) - done - -lemma zero_less_power_eq[presburger]: "(0 < (x::'a::{recpower,ordered_idom}) ^ n) = - (n = 0 | (even n & x ~= 0) | (odd n & 0 < x))" - - unfolding order_less_le zero_le_power_eq by auto - -lemma power_less_zero_eq[presburger]: "((x::'a::{recpower,ordered_idom}) ^ n < 0) = - (odd n & x < 0)" - apply (subst linorder_not_le [symmetric])+ - apply (subst zero_le_power_eq) - apply auto - done - -lemma power_le_zero_eq[presburger]: "((x::'a::{recpower,ordered_idom}) ^ n <= 0) = - (n ~= 0 & ((odd n & x <= 0) | (even n & x = 0)))" - apply (subst linorder_not_less [symmetric])+ - apply (subst zero_less_power_eq) - apply auto - done - -lemma power_even_abs: "even n ==> - (abs (x::'a::{recpower,ordered_idom}))^n = x^n" - apply (subst power_abs [symmetric]) - apply (simp add: zero_le_even_power) - done - -lemma zero_less_power_nat_eq[presburger]: "(0 < (x::nat) ^ n) = (n = 0 | 0 < x)" - by (induct n) auto - -lemma power_minus_even [simp]: "even n ==> - (- x)^n = (x^n::'a::{recpower,comm_ring_1})" - apply (subst power_minus) - apply simp - done - -lemma power_minus_odd [simp]: "odd n ==> - (- x)^n = - (x^n::'a::{recpower,comm_ring_1})" - apply (subst power_minus) - apply simp - done - - -subsection {* General Lemmas About Division *} - -lemma Suc_times_mod_eq: "1 Suc (k * m) mod k = 1" -apply (induct "m") -apply (simp_all add: mod_Suc) -done - -declare Suc_times_mod_eq [of "number_of w", standard, simp] - -lemma [simp]: "n div k \ (Suc n) div k" -by (simp add: div_le_mono) - -lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2" -by arith - -lemma div_2_gt_zero [simp]: "(1::nat) < n ==> 0 < n div 2" -by arith - - (* Potential use of algebra : Equality modulo n*) -lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)" -by (simp add: mult_ac add_ac) - -lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n" -proof - - have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp - also have "... = Suc m mod n" by (rule mod_mult_self3) - finally show ?thesis . -qed - -lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n" -apply (subst mod_Suc [of m]) -apply (subst mod_Suc [of "m mod n"], simp) -done - - -subsection {* More Even/Odd Results *} - -lemma even_mult_two_ex: "even(n) = (\m::nat. n = 2*m)" by presburger -lemma odd_Suc_mult_two_ex: "odd(n) = (\m. n = Suc (2*m))" by presburger -lemma even_add [simp]: "even(m + n::nat) = (even m = even n)" by presburger - -lemma odd_add [simp]: "odd(m + n::nat) = (odd m \ odd n)" by presburger - -lemma div_Suc: "Suc a div c = a div c + Suc 0 div c + - (a mod c + Suc 0 mod c) div c" - apply (subgoal_tac "Suc a = a + Suc 0") - apply (erule ssubst) - apply (rule div_add1_eq, simp) - done - -lemma lemma_even_div2 [simp]: "even (n::nat) ==> (n + 1) div 2 = n div 2" by presburger - -lemma lemma_not_even_div2 [simp]: "~even n ==> (n + 1) div 2 = Suc (n div 2)" -by presburger - -lemma even_num_iff: "0 < n ==> even n = (~ even(n - 1 :: nat))" by presburger -lemma even_even_mod_4_iff: "even (n::nat) = even (n mod 4)" by presburger - -lemma lemma_odd_mod_4_div_2: "n mod 4 = (3::nat) ==> odd((n - 1) div 2)" by presburger - -lemma lemma_even_mod_4_div_2: "n mod 4 = (1::nat) ==> even ((n - 1) div 2)" - by presburger - -text {* Simplify, when the exponent is a numeral *} - -lemmas power_0_left_number_of = power_0_left [of "number_of w", standard] -declare power_0_left_number_of [simp] - -lemmas zero_le_power_eq_number_of [simp] = - zero_le_power_eq [of _ "number_of w", standard] - -lemmas zero_less_power_eq_number_of [simp] = - zero_less_power_eq [of _ "number_of w", standard] - -lemmas power_le_zero_eq_number_of [simp] = - power_le_zero_eq [of _ "number_of w", standard] - -lemmas power_less_zero_eq_number_of [simp] = - power_less_zero_eq [of _ "number_of w", standard] - -lemmas zero_less_power_nat_eq_number_of [simp] = - zero_less_power_nat_eq [of _ "number_of w", standard] - -lemmas power_eq_0_iff_number_of [simp] = power_eq_0_iff [of _ "number_of w", standard] - -lemmas power_even_abs_number_of [simp] = power_even_abs [of "number_of w" _, standard] - - -subsection {* An Equivalence for @{term [source] "0 \ a^n"} *} - -lemma even_power_le_0_imp_0: - "a ^ (2*k) \ (0::'a::{ordered_idom,recpower}) ==> a=0" - by (induct k) (auto simp add: zero_le_mult_iff mult_le_0_iff power_Suc) - -lemma zero_le_power_iff[presburger]: - "(0 \ a^n) = (0 \ (a::'a::{ordered_idom,recpower}) | even n)" -proof cases - assume even: "even n" - then obtain k where "n = 2*k" - by (auto simp add: even_nat_equiv_def2 numeral_2_eq_2) - thus ?thesis by (simp add: zero_le_even_power even) -next - assume odd: "odd n" - then obtain k where "n = Suc(2*k)" - by (auto simp add: odd_nat_equiv_def2 numeral_2_eq_2) - thus ?thesis - by (auto simp add: power_Suc zero_le_mult_iff zero_le_even_power - dest!: even_power_le_0_imp_0) -qed - - -subsection {* Miscellaneous *} - -lemma odd_pos: "odd (n::nat) \ 0 < n" by presburger - -lemma [presburger]:"(x + 1) div 2 = x div 2 \ even (x::int)" by presburger -lemma [presburger]: "(x + 1) div 2 = x div 2 + 1 \ odd (x::int)" by presburger -lemma even_plus_one_div_two: "even (x::int) ==> (x + 1) div 2 = x div 2" by presburger -lemma odd_plus_one_div_two: "odd (x::int) ==> (x + 1) div 2 = x div 2 + 1" by presburger - -lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \ even x" by presburger -lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \ even x" by presburger -lemma even_nat_plus_one_div_two: "even (x::nat) ==> - (Suc x) div Suc (Suc 0) = x div Suc (Suc 0)" by presburger - -lemma odd_nat_plus_one_div_two: "odd (x::nat) ==> - (Suc x) div Suc (Suc 0) = Suc (x div Suc (Suc 0))" by presburger - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/Primes.thy --- a/src/HOL/Library/Primes.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Library/Primes.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Library/Primes.thy - ID: $Id$ Author: Amine Chaieb, Christophe Tabacznyj and Lawrence C Paulson Copyright 1996 University of Cambridge *) @@ -7,7 +6,7 @@ header {* Primality on nat *} theory Primes -imports Plain "~~/src/HOL/ATP_Linkup" GCD Parity +imports Plain "~~/src/HOL/ATP_Linkup" "~~/src/HOL/GCD" "~~/src/HOL/Parity" begin definition diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/RType.thy --- a/src/HOL/Library/RType.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,107 +0,0 @@ -(* Title: HOL/Library/RType.thy - ID: $Id$ - Author: Florian Haftmann, TU Muenchen -*) - -header {* Reflecting Pure types into HOL *} - -theory RType -imports Plain "~~/src/HOL/List" "~~/src/HOL/Library/Code_Message" -begin - -datatype typerep = Typerep message_string "typerep list" - -class typerep = - fixes typerep :: "'a\{} itself \ typerep" -begin - -definition - typerep_of :: "'a \ typerep" -where - [simp]: "typerep_of x = typerep TYPE('a)" - -end - -setup {* -let - fun typerep_tr (*"_TYPEREP"*) [ty] = - Lexicon.const @{const_syntax typerep} $ (Lexicon.const "_constrain" $ Lexicon.const "TYPE" $ - (Lexicon.const "itself" $ ty)) - | typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts); - fun typerep_tr' show_sorts (*"typerep"*) - (Type ("fun", [Type ("itself", [T]), _])) (Const (@{const_syntax TYPE}, _) :: ts) = - Term.list_comb (Lexicon.const "_TYPEREP" $ Syntax.term_of_typ show_sorts T, ts) - | typerep_tr' _ T ts = raise Match; -in - Sign.add_syntax_i - [("_TYPEREP", SimpleSyntax.read_typ "type => logic", Delimfix "(1TYPEREP/(1'(_')))")] - #> Sign.add_trfuns ([], [("_TYPEREP", typerep_tr)], [], []) - #> Sign.add_trfunsT [(@{const_syntax typerep}, typerep_tr')] -end -*} - -ML {* -structure Typerep = -struct - -fun mk f (Type (tyco, tys)) = - @{term Typerep} $ Message_String.mk tyco - $ HOLogic.mk_list @{typ typerep} (map (mk f) tys) - | mk f (TFree v) = - f v; - -fun typerep ty = - Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep}) - $ Logic.mk_type ty; - -fun add_def tyco thy = - let - val sorts = replicate (Sign.arity_number thy tyco) @{sort typerep}; - val vs = Name.names Name.context "'a" sorts; - val ty = Type (tyco, map TFree vs); - val lhs = Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep}) - $ Free ("T", Term.itselfT ty); - val rhs = mk (typerep o TFree) ty; - val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); - in - thy - |> TheoryTarget.instantiation ([tyco], vs, @{sort typerep}) - |> `(fn lthy => Syntax.check_term lthy eq) - |-> (fn eq => Specification.definition (NONE, (Attrib.no_binding, eq))) - |> snd - |> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) - |> LocalTheory.exit_global - end; - -fun perhaps_add_def tyco thy = - let - val inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep} - in if inst then thy else add_def tyco thy end; - -end; -*} - -setup {* - Typerep.add_def @{type_name prop} - #> Typerep.add_def @{type_name fun} - #> Typerep.add_def @{type_name itself} - #> Typerep.add_def @{type_name bool} - #> TypedefPackage.interpretation Typerep.perhaps_add_def -*} - -lemma [code]: - "eq_class.eq (Typerep tyco1 tys1) (Typerep tyco2 tys2) \ eq_class.eq tyco1 tyco2 - \ list_all2 eq_class.eq tys1 tys2" - by (auto simp add: equals_eq [symmetric] list_all2_eq [symmetric]) - -code_type typerep - (SML "Term.typ") - -code_const Typerep - (SML "Term.Type/ (_, _)") - -code_reserved SML Term - -hide (open) const typerep Typerep - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/Univ_Poly.thy --- a/src/HOL/Library/Univ_Poly.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1050 +0,0 @@ -(* Title: Univ_Poly.thy - ID: $Id$ - Author: Amine Chaieb -*) - -header{*Univariate Polynomials*} - -theory Univ_Poly -imports Plain "~~/src/HOL/List" -begin - -text{* Application of polynomial as a function. *} - -primrec (in semiring_0) poly :: "'a list => 'a => 'a" where - poly_Nil: "poly [] x = 0" -| poly_Cons: "poly (h#t) x = h + x * poly t x" - - -subsection{*Arithmetic Operations on Polynomials*} - -text{*addition*} - -primrec (in semiring_0) padd :: "'a list \ 'a list \ 'a list" (infixl "+++" 65) -where - padd_Nil: "[] +++ l2 = l2" -| padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t - else (h + hd l2)#(t +++ tl l2))" - -text{*Multiplication by a constant*} -primrec (in semiring_0) cmult :: "'a \ 'a list \ 'a list" (infixl "%*" 70) where - cmult_Nil: "c %* [] = []" -| cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)" - -text{*Multiplication by a polynomial*} -primrec (in semiring_0) pmult :: "'a list \ 'a list \ 'a list" (infixl "***" 70) -where - pmult_Nil: "[] *** l2 = []" -| pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2 - else (h %* l2) +++ ((0) # (t *** l2)))" - -text{*Repeated multiplication by a polynomial*} -primrec (in semiring_0) mulexp :: "nat \ 'a list \ 'a list \ 'a list" where - mulexp_zero: "mulexp 0 p q = q" -| mulexp_Suc: "mulexp (Suc n) p q = p *** mulexp n p q" - -text{*Exponential*} -primrec (in semiring_1) pexp :: "'a list \ nat \ 'a list" (infixl "%^" 80) where - pexp_0: "p %^ 0 = [1]" -| pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)" - -text{*Quotient related value of dividing a polynomial by x + a*} -(* Useful for divisor properties in inductive proofs *) -primrec (in field) "pquot" :: "'a list \ 'a \ 'a list" where - pquot_Nil: "pquot [] a= []" -| pquot_Cons: "pquot (h#t) a = (if t = [] then [h] - else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))" - -text{*normalization of polynomials (remove extra 0 coeff)*} -primrec (in semiring_0) pnormalize :: "'a list \ 'a list" where - pnormalize_Nil: "pnormalize [] = []" -| pnormalize_Cons: "pnormalize (h#p) = (if ( (pnormalize p) = []) - then (if (h = 0) then [] else [h]) - else (h#(pnormalize p)))" - -definition (in semiring_0) "pnormal p = ((pnormalize p = p) \ p \ [])" -definition (in semiring_0) "nonconstant p = (pnormal p \ (\x. p \ [x]))" -text{*Other definitions*} - -definition (in ring_1) - poly_minus :: "'a list => 'a list" ("-- _" [80] 80) where - "-- p = (- 1) %* p" - -definition (in semiring_0) - divides :: "'a list \ 'a list \ bool" (infixl "divides" 70) where - [code del]: "p1 divides p2 = (\q. poly p2 = poly(p1 *** q))" - - --{*order of a polynomial*} -definition (in ring_1) order :: "'a => 'a list => nat" where - "order a p = (SOME n. ([-a, 1] %^ n) divides p & - ~ (([-a, 1] %^ (Suc n)) divides p))" - - --{*degree of a polynomial*} -definition (in semiring_0) degree :: "'a list => nat" where - "degree p = length (pnormalize p) - 1" - - --{*squarefree polynomials --- NB with respect to real roots only.*} -definition (in ring_1) - rsquarefree :: "'a list => bool" where - "rsquarefree p = (poly p \ poly [] & - (\a. (order a p = 0) | (order a p = 1)))" - -context semiring_0 -begin - -lemma padd_Nil2[simp]: "p +++ [] = p" -by (induct p) auto - -lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)" -by auto - -lemma pminus_Nil[simp]: "-- [] = []" -by (simp add: poly_minus_def) - -lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" by simp -end - -lemma (in semiring_1) poly_ident_mult[simp]: "1 %* t = t" by (induct "t", auto) - -lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ ((0)#t) = (a#t)" -by simp - -text{*Handy general properties*} - -lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b" -proof(induct b arbitrary: a) - case Nil thus ?case by auto -next - case (Cons b bs a) thus ?case by (cases a, simp_all add: add_commute) -qed - -lemma (in comm_semiring_0) padd_assoc: "\b c. (a +++ b) +++ c = a +++ (b +++ c)" -apply (induct a arbitrary: b c) -apply (simp, clarify) -apply (case_tac b, simp_all add: add_ac) -done - -lemma (in semiring_0) poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)" -apply (induct p arbitrary: q,simp) -apply (case_tac q, simp_all add: right_distrib) -done - -lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = ((0)#t)" -apply (induct "t", simp) -apply (auto simp add: mult_zero_left poly_ident_mult padd_commut) -apply (case_tac t, auto) -done - -text{*properties of evaluation of polynomials.*} - -lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x" -proof(induct p1 arbitrary: p2) - case Nil thus ?case by simp -next - case (Cons a as p2) thus ?case - by (cases p2, simp_all add: add_ac right_distrib) -qed - -lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x" -apply (induct "p") -apply (case_tac [2] "x=zero") -apply (auto simp add: right_distrib mult_ac) -done - -lemma (in comm_semiring_0) poly_cmult_map: "poly (map (op * c) p) x = c*poly p x" - by (induct p, auto simp add: right_distrib mult_ac) - -lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)" -apply (simp add: poly_minus_def) -apply (auto simp add: poly_cmult minus_mult_left[symmetric]) -done - -lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x" -proof(induct p1 arbitrary: p2) - case Nil thus ?case by simp -next - case (Cons a as p2) - thus ?case by (cases as, - simp_all add: poly_cmult poly_add left_distrib right_distrib mult_ac) -qed - -class recpower_semiring = semiring + recpower -class recpower_semiring_1 = semiring_1 + recpower -class recpower_semiring_0 = semiring_0 + recpower -class recpower_ring = ring + recpower -class recpower_ring_1 = ring_1 + recpower -subclass (in recpower_ring_1) recpower_ring .. -class recpower_comm_semiring_1 = recpower + comm_semiring_1 -class recpower_comm_ring_1 = recpower + comm_ring_1 -subclass (in recpower_comm_ring_1) recpower_comm_semiring_1 .. -class recpower_idom = recpower + idom -subclass (in recpower_idom) recpower_comm_ring_1 .. -class idom_char_0 = idom + ring_char_0 -class recpower_idom_char_0 = recpower + idom_char_0 -subclass (in recpower_idom_char_0) recpower_idom .. - -lemma (in recpower_comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n" -apply (induct "n") -apply (auto simp add: poly_cmult poly_mult power_Suc) -done - -text{*More Polynomial Evaluation Lemmas*} - -lemma (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x" -by simp - -lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x" - by (simp add: poly_mult mult_assoc) - -lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0" -by (induct "p", auto) - -lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x" -apply (induct "n") -apply (auto simp add: poly_mult mult_assoc) -done - -subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides - @{term "p(x)"} *} - -lemma (in comm_ring_1) lemma_poly_linear_rem: "\h. \q r. h#t = [r] +++ [-a, 1] *** q" -proof(induct t) - case Nil - {fix h have "[h] = [h] +++ [- a, 1] *** []" by simp} - thus ?case by blast -next - case (Cons x xs) - {fix h - from Cons.hyps[rule_format, of x] - obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast - have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)" - using qr by(cases q, simp_all add: ring_simps diff_def[symmetric] - minus_mult_left[symmetric] right_minus) - hence "\q r. h#x#xs = [r] +++ [-a, 1] *** q" by blast} - thus ?case by blast -qed - -lemma (in comm_ring_1) poly_linear_rem: "\q r. h#t = [r] +++ [-a, 1] *** q" -by (cut_tac t = t and a = a in lemma_poly_linear_rem, auto) - - -lemma (in comm_ring_1) poly_linear_divides: "(poly p a = 0) = ((p = []) | (\q. p = [-a, 1] *** q))" -proof- - {assume p: "p = []" hence ?thesis by simp} - moreover - {fix x xs assume p: "p = x#xs" - {fix q assume "p = [-a, 1] *** q" hence "poly p a = 0" - by (simp add: poly_add poly_cmult minus_mult_left[symmetric])} - moreover - {assume p0: "poly p a = 0" - from poly_linear_rem[of x xs a] obtain q r - where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast - have "r = 0" using p0 by (simp only: p qr poly_mult poly_add) simp - hence "\q. p = [- a, 1] *** q" using p qr apply - apply (rule exI[where x=q])apply auto apply (cases q) apply auto done} - ultimately have ?thesis using p by blast} - ultimately show ?thesis by (cases p, auto) -qed - -lemma (in semiring_0) lemma_poly_length_mult[simp]: "\h k a. length (k %* p +++ (h # (a %* p))) = Suc (length p)" -by (induct "p", auto) - -lemma (in semiring_0) lemma_poly_length_mult2[simp]: "\h k. length (k %* p +++ (h # p)) = Suc (length p)" -by (induct "p", auto) - -lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)" -by auto - -subsection{*Polynomial length*} - -lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p" -by (induct "p", auto) - -lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)" -apply (induct p1 arbitrary: p2, simp_all) -apply arith -done - -lemma (in semiring_0) poly_root_mult_length[simp]: "length([a,b] *** p) = Suc (length p)" -by (simp add: poly_add_length) - -lemma (in idom) poly_mult_not_eq_poly_Nil[simp]: - "poly (p *** q) x \ poly [] x \ poly p x \ poly [] x \ poly q x \ poly [] x" -by (auto simp add: poly_mult) - -lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \ poly p x = 0 \ poly q x = 0" -by (auto simp add: poly_mult) - -text{*Normalisation Properties*} - -lemma (in semiring_0) poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)" -by (induct "p", auto) - -text{*A nontrivial polynomial of degree n has no more than n roots*} -lemma (in idom) poly_roots_index_lemma: - assumes p: "poly p x \ poly [] x" and n: "length p = n" - shows "\i. \x. poly p x = 0 \ (\m\n. x = i m)" - using p n -proof(induct n arbitrary: p x) - case 0 thus ?case by simp -next - case (Suc n p x) - {assume C: "\i. \x. poly p x = 0 \ (\m\Suc n. x \ i m)" - from Suc.prems have p0: "poly p x \ 0" "p\ []" by auto - from p0(1)[unfolded poly_linear_divides[of p x]] - have "\q. p \ [- x, 1] *** q" by blast - from C obtain a where a: "poly p a = 0" by blast - from a[unfolded poly_linear_divides[of p a]] p0(2) - obtain q where q: "p = [-a, 1] *** q" by blast - have lg: "length q = n" using q Suc.prems(2) by simp - from q p0 have qx: "poly q x \ poly [] x" - by (auto simp add: poly_mult poly_add poly_cmult) - from Suc.hyps[OF qx lg] obtain i where - i: "\x. poly q x = 0 \ (\m\n. x = i m)" by blast - let ?i = "\m. if m = Suc n then a else i m" - from C[of ?i] obtain y where y: "poly p y = 0" "\m\ Suc n. y \ ?i m" - by blast - from y have "y = a \ poly q y = 0" - by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: ring_simps) - with i[rule_format, of y] y(1) y(2) have False apply auto - apply (erule_tac x="m" in allE) - apply auto - done} - thus ?case by blast -qed - - -lemma (in idom) poly_roots_index_length: "poly p x \ poly [] x ==> - \i. \x. (poly p x = 0) --> (\n. n \ length p & x = i n)" -by (blast intro: poly_roots_index_lemma) - -lemma (in idom) poly_roots_finite_lemma1: "poly p x \ poly [] x ==> - \N i. \x. (poly p x = 0) --> (\n. (n::nat) < N & x = i n)" -apply (drule poly_roots_index_length, safe) -apply (rule_tac x = "Suc (length p)" in exI) -apply (rule_tac x = i in exI) -apply (simp add: less_Suc_eq_le) -done - - -lemma (in idom) idom_finite_lemma: - assumes P: "\x. P x --> (\n. n < length j & x = j!n)" - shows "finite {x. P x}" -proof- - let ?M = "{x. P x}" - let ?N = "set j" - have "?M \ ?N" using P by auto - thus ?thesis using finite_subset by auto -qed - - -lemma (in idom) poly_roots_finite_lemma2: "poly p x \ poly [] x ==> - \i. \x. (poly p x = 0) --> x \ set i" -apply (drule poly_roots_index_length, safe) -apply (rule_tac x="map (\n. i n) [0 ..< Suc (length p)]" in exI) -apply (auto simp add: image_iff) -apply (erule_tac x="x" in allE, clarsimp) -by (case_tac "n=length p", auto simp add: order_le_less) - -lemma UNIV_nat_infinite: "\ finite (UNIV :: nat set)" - unfolding finite_conv_nat_seg_image -proof(auto simp add: expand_set_eq image_iff) - fix n::nat and f:: "nat \ nat" - let ?N = "{i. i < n}" - let ?fN = "f ` ?N" - let ?y = "Max ?fN + 1" - from nat_seg_image_imp_finite[of "?fN" "f" n] - have thfN: "finite ?fN" by simp - {assume "n =0" hence "\x. \xa f xa" by auto} - moreover - {assume nz: "n \ 0" - hence thne: "?fN \ {}" by (auto simp add: neq0_conv) - have "\x\ ?fN. Max ?fN \ x" using nz Max_ge_iff[OF thfN thne] by auto - hence "\x\ ?fN. ?y > x" by auto - hence "?y \ ?fN" by auto - hence "\x. \xa f xa" by auto } - ultimately show "\x. \xa f xa" by blast -qed - -lemma (in ring_char_0) UNIV_ring_char_0_infinte: - "\ (finite (UNIV:: 'a set))" -proof - assume F: "finite (UNIV :: 'a set)" - have th0: "of_nat ` UNIV \ UNIV" by simp - from finite_subset[OF th0] have th: "finite (of_nat ` UNIV :: 'a set)" . - have th': "inj_on (of_nat::nat \ 'a) (UNIV)" - unfolding inj_on_def by auto - from finite_imageD[OF th th'] UNIV_nat_infinite - show False by blast -qed - -lemma (in idom_char_0) poly_roots_finite: "(poly p \ poly []) = - finite {x. poly p x = 0}" -proof - assume H: "poly p \ poly []" - show "finite {x. poly p x = (0::'a)}" - using H - apply - - apply (erule contrapos_np, rule ext) - apply (rule ccontr) - apply (clarify dest!: poly_roots_finite_lemma2) - using finite_subset - proof- - fix x i - assume F: "\ finite {x. poly p x = (0\'a)}" - and P: "\x. poly p x = (0\'a) \ x \ set i" - let ?M= "{x. poly p x = (0\'a)}" - from P have "?M \ set i" by auto - with finite_subset F show False by auto - qed -next - assume F: "finite {x. poly p x = (0\'a)}" - show "poly p \ poly []" using F UNIV_ring_char_0_infinte by auto -qed - -text{*Entirety and Cancellation for polynomials*} - -lemma (in idom_char_0) poly_entire_lemma2: - assumes p0: "poly p \ poly []" and q0: "poly q \ poly []" - shows "poly (p***q) \ poly []" -proof- - let ?S = "\p. {x. poly p x = 0}" - have "?S (p *** q) = ?S p \ ?S q" by (auto simp add: poly_mult) - with p0 q0 show ?thesis unfolding poly_roots_finite by auto -qed - -lemma (in idom_char_0) poly_entire: - "poly (p *** q) = poly [] \ poly p = poly [] \ poly q = poly []" -using poly_entire_lemma2[of p q] -by auto (rule ext, simp add: poly_mult)+ - -lemma (in idom_char_0) poly_entire_neg: "(poly (p *** q) \ poly []) = ((poly p \ poly []) & (poly q \ poly []))" -by (simp add: poly_entire) - -lemma fun_eq: " (f = g) = (\x. f x = g x)" -by (auto intro!: ext) - -lemma (in comm_ring_1) poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)" -by (auto simp add: ring_simps poly_add poly_minus_def fun_eq poly_cmult minus_mult_left[symmetric]) - -lemma (in comm_ring_1) poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))" -by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult right_distrib minus_mult_left[symmetric] minus_mult_right[symmetric]) - -subclass (in idom_char_0) comm_ring_1 .. -lemma (in idom_char_0) poly_mult_left_cancel: "(poly (p *** q) = poly (p *** r)) = (poly p = poly [] | poly q = poly r)" -proof- - have "poly (p *** q) = poly (p *** r) \ poly (p *** q +++ -- (p *** r)) = poly []" by (simp only: poly_add_minus_zero_iff) - also have "\ \ poly p = poly [] | poly q = poly r" - by (auto intro: ext simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff) - finally show ?thesis . -qed - -lemma (in recpower_idom) poly_exp_eq_zero[simp]: - "(poly (p %^ n) = poly []) = (poly p = poly [] & n \ 0)" -apply (simp only: fun_eq add: all_simps [symmetric]) -apply (rule arg_cong [where f = All]) -apply (rule ext) -apply (induct n) -apply (auto simp add: poly_exp poly_mult) -done - -lemma (in semiring_1) one_neq_zero[simp]: "1 \ 0" using zero_neq_one by blast -lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a,1] \ poly []" -apply (simp add: fun_eq) -apply (rule_tac x = "minus one a" in exI) -apply (unfold diff_minus) -apply (subst add_commute) -apply (subst add_assoc) -apply simp -done - -lemma (in recpower_idom) poly_exp_prime_eq_zero: "(poly ([a, 1] %^ n) \ poly [])" -by auto - -text{*A more constructive notion of polynomials being trivial*} - -lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] ==> h = 0 & poly t = poly []" -apply(simp add: fun_eq) -apply (case_tac "h = zero") -apply (drule_tac [2] x = zero in spec, auto) -apply (cases "poly t = poly []", simp) -proof- - fix x - assume H: "\x. x = (0\'a) \ poly t x = (0\'a)" and pnz: "poly t \ poly []" - let ?S = "{x. poly t x = 0}" - from H have "\x. x \0 \ poly t x = 0" by blast - hence th: "?S \ UNIV - {0}" by auto - from poly_roots_finite pnz have th': "finite ?S" by blast - from finite_subset[OF th th'] UNIV_ring_char_0_infinte - show "poly t x = (0\'a)" by simp - qed - -lemma (in idom_char_0) poly_zero: "(poly p = poly []) = list_all (%c. c = 0) p" -apply (induct "p", simp) -apply (rule iffI) -apply (drule poly_zero_lemma', auto) -done - -lemma (in idom_char_0) poly_0: "list_all (\c. c = 0) p \ poly p x = 0" - unfolding poly_zero[symmetric] by simp - - - -text{*Basics of divisibility.*} - -lemma (in idom) poly_primes: "([a, 1] divides (p *** q)) = ([a, 1] divides p | [a, 1] divides q)" -apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult left_distrib [symmetric]) -apply (drule_tac x = "uminus a" in spec) -apply (simp add: poly_linear_divides poly_add poly_cmult left_distrib [symmetric]) -apply (cases "p = []") -apply (rule exI[where x="[]"]) -apply simp -apply (cases "q = []") -apply (erule allE[where x="[]"], simp) - -apply clarsimp -apply (cases "\q\'a list. p = a %* q +++ ((0\'a) # q)") -apply (clarsimp simp add: poly_add poly_cmult) -apply (rule_tac x="qa" in exI) -apply (simp add: left_distrib [symmetric]) -apply clarsimp - -apply (auto simp add: right_minus poly_linear_divides poly_add poly_cmult left_distrib [symmetric]) -apply (rule_tac x = "pmult qa q" in exI) -apply (rule_tac [2] x = "pmult p qa" in exI) -apply (auto simp add: poly_add poly_mult poly_cmult mult_ac) -done - -lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p" -apply (simp add: divides_def) -apply (rule_tac x = "[one]" in exI) -apply (auto simp add: poly_mult fun_eq) -done - -lemma (in comm_semiring_1) poly_divides_trans: "[| p divides q; q divides r |] ==> p divides r" -apply (simp add: divides_def, safe) -apply (rule_tac x = "pmult qa qaa" in exI) -apply (auto simp add: poly_mult fun_eq mult_assoc) -done - - -lemma (in recpower_comm_semiring_1) poly_divides_exp: "m \ n ==> (p %^ m) divides (p %^ n)" -apply (auto simp add: le_iff_add) -apply (induct_tac k) -apply (rule_tac [2] poly_divides_trans) -apply (auto simp add: divides_def) -apply (rule_tac x = p in exI) -apply (auto simp add: poly_mult fun_eq mult_ac) -done - -lemma (in recpower_comm_semiring_1) poly_exp_divides: "[| (p %^ n) divides q; m\n |] ==> (p %^ m) divides q" -by (blast intro: poly_divides_exp poly_divides_trans) - -lemma (in comm_semiring_0) poly_divides_add: - "[| p divides q; p divides r |] ==> p divides (q +++ r)" -apply (simp add: divides_def, auto) -apply (rule_tac x = "padd qa qaa" in exI) -apply (auto simp add: poly_add fun_eq poly_mult right_distrib) -done - -lemma (in comm_ring_1) poly_divides_diff: - "[| p divides q; p divides (q +++ r) |] ==> p divides r" -apply (simp add: divides_def, auto) -apply (rule_tac x = "padd qaa (poly_minus qa)" in exI) -apply (auto simp add: poly_add fun_eq poly_mult poly_minus right_diff_distrib compare_rls add_ac) -done - -lemma (in comm_ring_1) poly_divides_diff2: "[| p divides r; p divides (q +++ r) |] ==> p divides q" -apply (erule poly_divides_diff) -apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac) -done - -lemma (in semiring_0) poly_divides_zero: "poly p = poly [] ==> q divides p" -apply (simp add: divides_def) -apply (rule exI[where x="[]"]) -apply (auto simp add: fun_eq poly_mult) -done - -lemma (in semiring_0) poly_divides_zero2[simp]: "q divides []" -apply (simp add: divides_def) -apply (rule_tac x = "[]" in exI) -apply (auto simp add: fun_eq) -done - -text{*At last, we can consider the order of a root.*} - -lemma (in idom_char_0) poly_order_exists_lemma: - assumes lp: "length p = d" and p: "poly p \ poly []" - shows "\n q. p = mulexp n [-a, 1] q \ poly q a \ 0" -using lp p -proof(induct d arbitrary: p) - case 0 thus ?case by simp -next - case (Suc n p) - {assume p0: "poly p a = 0" - from Suc.prems have h: "length p = Suc n" "poly p \ poly []" by blast - hence pN: "p \ []" by - (rule ccontr, simp) - from p0[unfolded poly_linear_divides] pN obtain q where - q: "p = [-a, 1] *** q" by blast - from q h p0 have qh: "length q = n" "poly q \ poly []" - apply - - apply simp - apply (simp only: fun_eq) - apply (rule ccontr) - apply (simp add: fun_eq poly_add poly_cmult minus_mult_left[symmetric]) - done - from Suc.hyps[OF qh] obtain m r where - mr: "q = mulexp m [-a,1] r" "poly r a \ 0" by blast - from mr q have "p = mulexp (Suc m) [-a,1] r \ poly r a \ 0" by simp - hence ?case by blast} - moreover - {assume p0: "poly p a \ 0" - hence ?case using Suc.prems apply simp by (rule exI[where x="0::nat"], simp)} - ultimately show ?case by blast -qed - - -lemma (in recpower_comm_semiring_1) poly_mulexp: "poly (mulexp n p q) x = (poly p x) ^ n * poly q x" -by(induct n, auto simp add: poly_mult power_Suc mult_ac) - -lemma (in comm_semiring_1) divides_left_mult: - assumes d:"(p***q) divides r" shows "p divides r \ q divides r" -proof- - from d obtain t where r:"poly r = poly (p***q *** t)" - unfolding divides_def by blast - hence "poly r = poly (p *** (q *** t))" - "poly r = poly (q *** (p***t))" by(auto simp add: fun_eq poly_mult mult_ac) - thus ?thesis unfolding divides_def by blast -qed - - - -(* FIXME: Tidy up *) - -lemma (in recpower_semiring_1) - zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)" - by (induct n, simp_all add: power_Suc) - -lemma (in recpower_idom_char_0) poly_order_exists: - assumes lp: "length p = d" and p0: "poly p \ poly []" - shows "\n. ([-a, 1] %^ n) divides p & ~(([-a, 1] %^ (Suc n)) divides p)" -proof- -let ?poly = poly -let ?mulexp = mulexp -let ?pexp = pexp -from lp p0 -show ?thesis -apply - -apply (drule poly_order_exists_lemma [where a=a], assumption, clarify) -apply (rule_tac x = n in exI, safe) -apply (unfold divides_def) -apply (rule_tac x = q in exI) -apply (induct_tac "n", simp) -apply (simp (no_asm_simp) add: poly_add poly_cmult poly_mult right_distrib mult_ac) -apply safe -apply (subgoal_tac "?poly (?mulexp n [uminus a, one] q) \ ?poly (pmult (?pexp [uminus a, one] (Suc n)) qa)") -apply simp -apply (induct_tac "n") -apply (simp del: pmult_Cons pexp_Suc) -apply (erule_tac Q = "?poly q a = zero" in contrapos_np) -apply (simp add: poly_add poly_cmult minus_mult_left[symmetric]) -apply (rule pexp_Suc [THEN ssubst]) -apply (rule ccontr) -apply (simp add: poly_mult_left_cancel poly_mult_assoc del: pmult_Cons pexp_Suc) -done -qed - - -lemma (in semiring_1) poly_one_divides[simp]: "[1] divides p" -by (simp add: divides_def, auto) - -lemma (in recpower_idom_char_0) poly_order: "poly p \ poly [] - ==> EX! n. ([-a, 1] %^ n) divides p & - ~(([-a, 1] %^ (Suc n)) divides p)" -apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc) -apply (cut_tac x = y and y = n in less_linear) -apply (drule_tac m = n in poly_exp_divides) -apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides] - simp del: pmult_Cons pexp_Suc) -done - -text{*Order*} - -lemma some1_equalityD: "[| n = (@n. P n); EX! n. P n |] ==> P n" -by (blast intro: someI2) - -lemma (in recpower_idom_char_0) order: - "(([-a, 1] %^ n) divides p & - ~(([-a, 1] %^ (Suc n)) divides p)) = - ((n = order a p) & ~(poly p = poly []))" -apply (unfold order_def) -apply (rule iffI) -apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order) -apply (blast intro!: poly_order [THEN [2] some1_equalityD]) -done - -lemma (in recpower_idom_char_0) order2: "[| poly p \ poly [] |] - ==> ([-a, 1] %^ (order a p)) divides p & - ~(([-a, 1] %^ (Suc(order a p))) divides p)" -by (simp add: order del: pexp_Suc) - -lemma (in recpower_idom_char_0) order_unique: "[| poly p \ poly []; ([-a, 1] %^ n) divides p; - ~(([-a, 1] %^ (Suc n)) divides p) - |] ==> (n = order a p)" -by (insert order [of a n p], auto) - -lemma (in recpower_idom_char_0) order_unique_lemma: "(poly p \ poly [] & ([-a, 1] %^ n) divides p & - ~(([-a, 1] %^ (Suc n)) divides p)) - ==> (n = order a p)" -by (blast intro: order_unique) - -lemma (in ring_1) order_poly: "poly p = poly q ==> order a p = order a q" -by (auto simp add: fun_eq divides_def poly_mult order_def) - -lemma (in semiring_1) pexp_one[simp]: "p %^ (Suc 0) = p" -apply (induct "p") -apply (auto simp add: numeral_1_eq_1) -done - -lemma (in comm_ring_1) lemma_order_root: - " 0 < n & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p - \ poly p a = 0" -apply (induct n arbitrary: a p, blast) -apply (auto simp add: divides_def poly_mult simp del: pmult_Cons) -done - -lemma (in recpower_idom_char_0) order_root: "(poly p a = 0) = ((poly p = poly []) | order a p \ 0)" -proof- - let ?poly = poly - show ?thesis -apply (case_tac "?poly p = ?poly []", auto) -apply (simp add: poly_linear_divides del: pmult_Cons, safe) -apply (drule_tac [!] a = a in order2) -apply (rule ccontr) -apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast) -using neq0_conv -apply (blast intro: lemma_order_root) -done -qed - -lemma (in recpower_idom_char_0) order_divides: "(([-a, 1] %^ n) divides p) = ((poly p = poly []) | n \ order a p)" -proof- - let ?poly = poly - show ?thesis -apply (case_tac "?poly p = ?poly []", auto) -apply (simp add: divides_def fun_eq poly_mult) -apply (rule_tac x = "[]" in exI) -apply (auto dest!: order2 [where a=a] - intro: poly_exp_divides simp del: pexp_Suc) -done -qed - -lemma (in recpower_idom_char_0) order_decomp: - "poly p \ poly [] - ==> \q. (poly p = poly (([-a, 1] %^ (order a p)) *** q)) & - ~([-a, 1] divides q)" -apply (unfold divides_def) -apply (drule order2 [where a = a]) -apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe) -apply (rule_tac x = q in exI, safe) -apply (drule_tac x = qa in spec) -apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons) -done - -text{*Important composition properties of orders.*} -lemma order_mult: "poly (p *** q) \ poly [] - ==> order a (p *** q) = order a p + order (a::'a::{recpower_idom_char_0}) q" -apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order) -apply (auto simp add: poly_entire simp del: pmult_Cons) -apply (drule_tac a = a in order2)+ -apply safe -apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe) -apply (rule_tac x = "qa *** qaa" in exI) -apply (simp add: poly_mult mult_ac del: pmult_Cons) -apply (drule_tac a = a in order_decomp)+ -apply safe -apply (subgoal_tac "[-a,1] divides (qa *** qaa) ") -apply (simp add: poly_primes del: pmult_Cons) -apply (auto simp add: divides_def simp del: pmult_Cons) -apply (rule_tac x = qb in exI) -apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))") -apply (drule poly_mult_left_cancel [THEN iffD1], force) -apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ") -apply (drule poly_mult_left_cancel [THEN iffD1], force) -apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons) -done - -lemma (in recpower_idom_char_0) order_mult: - assumes pq0: "poly (p *** q) \ poly []" - shows "order a (p *** q) = order a p + order a q" -proof- - let ?order = order - let ?divides = "op divides" - let ?poly = poly -from pq0 -show ?thesis -apply (cut_tac a = a and p = "pmult p q" and n = "?order a p + ?order a q" in order) -apply (auto simp add: poly_entire simp del: pmult_Cons) -apply (drule_tac a = a in order2)+ -apply safe -apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe) -apply (rule_tac x = "pmult qa qaa" in exI) -apply (simp add: poly_mult mult_ac del: pmult_Cons) -apply (drule_tac a = a in order_decomp)+ -apply safe -apply (subgoal_tac "?divides [uminus a,one ] (pmult qa qaa) ") -apply (simp add: poly_primes del: pmult_Cons) -apply (auto simp add: divides_def simp del: pmult_Cons) -apply (rule_tac x = qb in exI) -apply (subgoal_tac "?poly (pmult (pexp [uminus a, one] (?order a p)) (pmult qa qaa)) = ?poly (pmult (pexp [uminus a, one] (?order a p)) (pmult [uminus a, one] qb))") -apply (drule poly_mult_left_cancel [THEN iffD1], force) -apply (subgoal_tac "?poly (pmult (pexp [uminus a, one ] (order a q)) (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) = ?poly (pmult (pexp [uminus a, one] (order a q)) (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb))) ") -apply (drule poly_mult_left_cancel [THEN iffD1], force) -apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons) -done -qed - -lemma (in recpower_idom_char_0) order_root2: "poly p \ poly [] ==> (poly p a = 0) = (order a p \ 0)" -by (rule order_root [THEN ssubst], auto) - -lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p" by auto - -lemma (in semiring_0) poly_Nil_zero: "poly [] = poly [0]" -by (simp add: fun_eq) - -lemma (in recpower_idom_char_0) rsquarefree_decomp: - "[| rsquarefree p; poly p a = 0 |] - ==> \q. (poly p = poly ([-a, 1] *** q)) & poly q a \ 0" -apply (simp add: rsquarefree_def, safe) -apply (frule_tac a = a in order_decomp) -apply (drule_tac x = a in spec) -apply (drule_tac a = a in order_root2 [symmetric]) -apply (auto simp del: pmult_Cons) -apply (rule_tac x = q in exI, safe) -apply (simp add: poly_mult fun_eq) -apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1]) -apply (simp add: divides_def del: pmult_Cons, safe) -apply (drule_tac x = "[]" in spec) -apply (auto simp add: fun_eq) -done - - -text{*Normalization of a polynomial.*} - -lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p" -apply (induct "p") -apply (auto simp add: fun_eq) -done - -text{*The degree of a polynomial.*} - -lemma (in semiring_0) lemma_degree_zero: - "list_all (%c. c = 0) p \ pnormalize p = []" -by (induct "p", auto) - -lemma (in idom_char_0) degree_zero: - assumes pN: "poly p = poly []" shows"degree p = 0" -proof- - let ?pn = pnormalize - from pN - show ?thesis - apply (simp add: degree_def) - apply (case_tac "?pn p = []") - apply (auto simp add: poly_zero lemma_degree_zero ) - done -qed - -lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \ x \ 0" by simp -lemma (in semiring_0) pnormalize_pair: "y \ 0 \ (pnormalize [x, y] = [x, y])" by simp -lemma (in semiring_0) pnormal_cons: "pnormal p \ pnormal (c#p)" - unfolding pnormal_def by simp -lemma (in semiring_0) pnormal_tail: "p\[] \ pnormal (c#p) \ pnormal p" - unfolding pnormal_def - apply (cases "pnormalize p = []", auto) - by (cases "c = 0", auto) - - -lemma (in semiring_0) pnormal_last_nonzero: "pnormal p ==> last p \ 0" -proof(induct p) - case Nil thus ?case by (simp add: pnormal_def) -next - case (Cons a as) thus ?case - apply (simp add: pnormal_def) - apply (cases "pnormalize as = []", simp_all) - apply (cases "as = []", simp_all) - apply (cases "a=0", simp_all) - apply (cases "a=0", simp_all) - done -qed - -lemma (in semiring_0) pnormal_length: "pnormal p \ 0 < length p" - unfolding pnormal_def length_greater_0_conv by blast - -lemma (in semiring_0) pnormal_last_length: "\0 < length p ; last p \ 0\ \ pnormal p" - apply (induct p, auto) - apply (case_tac "p = []", auto) - apply (simp add: pnormal_def) - by (rule pnormal_cons, auto) - -lemma (in semiring_0) pnormal_id: "pnormal p \ (0 < length p \ last p \ 0)" - using pnormal_last_length pnormal_length pnormal_last_nonzero by blast - -lemma (in idom_char_0) poly_Cons_eq: "poly (c#cs) = poly (d#ds) \ c=d \ poly cs = poly ds" (is "?lhs \ ?rhs") -proof - assume eq: ?lhs - hence "\x. poly ((c#cs) +++ -- (d#ds)) x = 0" - by (simp only: poly_minus poly_add ring_simps) simp - hence "poly ((c#cs) +++ -- (d#ds)) = poly []" by - (rule ext, simp) - hence "c = d \ list_all (\x. x=0) ((cs +++ -- ds))" - unfolding poly_zero by (simp add: poly_minus_def ring_simps minus_mult_left[symmetric]) - hence "c = d \ (\x. poly (cs +++ -- ds) x = 0)" - unfolding poly_zero[symmetric] by simp - thus ?rhs apply (simp add: poly_minus poly_add ring_simps) apply (rule ext, simp) done -next - assume ?rhs then show ?lhs by - (rule ext,simp) -qed - -lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \ pnormalize p = pnormalize q" -proof(induct q arbitrary: p) - case Nil thus ?case by (simp only: poly_zero lemma_degree_zero) simp -next - case (Cons c cs p) - thus ?case - proof(induct p) - case Nil - hence "poly [] = poly (c#cs)" by blast - then have "poly (c#cs) = poly [] " by simp - thus ?case by (simp only: poly_zero lemma_degree_zero) simp - next - case (Cons d ds) - hence eq: "poly (d # ds) = poly (c # cs)" by blast - hence eq': "\x. poly (d # ds) x = poly (c # cs) x" by simp - hence "poly (d # ds) 0 = poly (c # cs) 0" by blast - hence dc: "d = c" by auto - with eq have "poly ds = poly cs" - unfolding poly_Cons_eq by simp - with Cons.prems have "pnormalize ds = pnormalize cs" by blast - with dc show ?case by simp - qed -qed - -lemma (in idom_char_0) degree_unique: assumes pq: "poly p = poly q" - shows "degree p = degree q" -using pnormalize_unique[OF pq] unfolding degree_def by simp - -lemma (in semiring_0) pnormalize_length: "length (pnormalize p) \ length p" by (induct p, auto) - -lemma (in semiring_0) last_linear_mul_lemma: - "last ((a %* p) +++ (x#(b %* p))) = (if p=[] then x else b*last p)" - -apply (induct p arbitrary: a x b, auto) -apply (subgoal_tac "padd (cmult aa p) (times b a # cmult b p) \ []", simp) -apply (induct_tac p, auto) -done - -lemma (in semiring_1) last_linear_mul: assumes p:"p\[]" shows "last ([a,1] *** p) = last p" -proof- - from p obtain c cs where cs: "p = c#cs" by (cases p, auto) - from cs have eq:"[a,1] *** p = (a %* (c#cs)) +++ (0#(1 %* (c#cs)))" - by (simp add: poly_cmult_distr) - show ?thesis using cs - unfolding eq last_linear_mul_lemma by simp -qed - -lemma (in semiring_0) pnormalize_eq: "last p \ 0 \ pnormalize p = p" - apply (induct p, auto) - apply (case_tac p, auto)+ - done - -lemma (in semiring_0) last_pnormalize: "pnormalize p \ [] \ last (pnormalize p) \ 0" - by (induct p, auto) - -lemma (in semiring_0) pnormal_degree: "last p \ 0 \ degree p = length p - 1" - using pnormalize_eq[of p] unfolding degree_def by simp - -lemma (in semiring_0) poly_Nil_ext: "poly [] = (\x. 0)" by (rule ext) simp - -lemma (in idom_char_0) linear_mul_degree: assumes p: "poly p \ poly []" - shows "degree ([a,1] *** p) = degree p + 1" -proof- - from p have pnz: "pnormalize p \ []" - unfolding poly_zero lemma_degree_zero . - - from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz] - have l0: "last ([a, 1] *** pnormalize p) \ 0" by simp - from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a] - pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz - - - have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1" - by (auto simp add: poly_length_mult) - - have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)" - by (rule ext) (simp add: poly_mult poly_add poly_cmult) - from degree_unique[OF eqs] th - show ?thesis by (simp add: degree_unique[OF poly_normalize]) -qed - -lemma (in idom_char_0) linear_pow_mul_degree: - "degree([a,1] %^n *** p) = (if poly p = poly [] then 0 else degree p + n)" -proof(induct n arbitrary: a p) - case (0 a p) - {assume p: "poly p = poly []" - hence ?case using degree_unique[OF p] by (simp add: degree_def)} - moreover - {assume p: "poly p \ poly []" hence ?case by (auto simp add: poly_Nil_ext) } - ultimately show ?case by blast -next - case (Suc n a p) - have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1] %^ n *** ([a,1] *** p))" - apply (rule ext, simp add: poly_mult poly_add poly_cmult) - by (simp add: mult_ac add_ac right_distrib) - note deq = degree_unique[OF eq] - {assume p: "poly p = poly []" - with eq have eq': "poly ([a,1] %^(Suc n) *** p) = poly []" - by - (rule ext,simp add: poly_mult poly_cmult poly_add) - from degree_unique[OF eq'] p have ?case by (simp add: degree_def)} - moreover - {assume p: "poly p \ poly []" - from p have ap: "poly ([a,1] *** p) \ poly []" - using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto - have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1]%^n *** ([a,1] *** p))" - by (rule ext, simp add: poly_mult poly_add poly_exp poly_cmult mult_ac add_ac right_distrib) - from ap have ap': "(poly ([a,1] *** p) = poly []) = False" by blast - have th0: "degree ([a,1]%^n *** ([a,1] *** p)) = degree ([a,1] *** p) + n" - apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap') - by simp - - from degree_unique[OF eq] ap p th0 linear_mul_degree[OF p, of a] - have ?case by (auto simp del: poly.simps)} - ultimately show ?case by blast -qed - -lemma (in recpower_idom_char_0) order_degree: - assumes p0: "poly p \ poly []" - shows "order a p \ degree p" -proof- - from order2[OF p0, unfolded divides_def] - obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" by blast - {assume "poly q = poly []" - with q p0 have False by (simp add: poly_mult poly_entire)} - with degree_unique[OF q, unfolded linear_pow_mul_degree] - show ?thesis by auto -qed - -text{*Tidier versions of finiteness of roots.*} - -lemma (in idom_char_0) poly_roots_finite_set: "poly p \ poly [] ==> finite {x. poly p x = 0}" -unfolding poly_roots_finite . - -text{*bound for polynomial.*} - -lemma poly_mono: "abs(x) \ k ==> abs(poly p (x::'a::{ordered_idom})) \ poly (map abs p) k" -apply (induct "p", auto) -apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans) -apply (rule abs_triangle_ineq) -apply (auto intro!: mult_mono simp add: abs_mult) -done - -lemma (in semiring_0) poly_Sing: "poly [c] x = c" by simp - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Library/Zorn.thy --- a/src/HOL/Library/Zorn.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Library/Zorn.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title : HOL/Library/Zorn.thy - ID : $Id$ Author : Jacques D. Fleuriot, Tobias Nipkow Description : Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF) The well-ordering theorem @@ -8,7 +7,7 @@ header {* Zorn's Lemma *} theory Zorn -imports Order_Relation +imports "~~/src/HOL/Order_Relation" begin (* Define globally? In Set.thy? *) diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Lim.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Lim.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,672 @@ +(* Title : Lim.thy + Author : Jacques D. Fleuriot + Copyright : 1998 University of Cambridge + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 +*) + +header{* Limits and Continuity *} + +theory Lim +imports "~~/src/HOL/Hyperreal/SEQ" +begin + +text{*Standard Definitions*} + +definition + LIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool" + ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where + [code del]: "f -- a --> L = + (\r > 0. \s > 0. \x. x \ a & norm (x - a) < s + --> norm (f x - L) < r)" + +definition + isCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where + "isCont f a = (f -- a --> (f a))" + +definition + isUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where + [code del]: "isUCont f = (\r>0. \s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r)" + + +subsection {* Limits of Functions *} + +subsubsection {* Purely standard proofs *} + +lemma LIM_eq: + "f -- a --> L = + (\r>0.\s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r)" +by (simp add: LIM_def diff_minus) + +lemma LIM_I: + "(!!r. 0 \s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r) + ==> f -- a --> L" +by (simp add: LIM_eq) + +lemma LIM_D: + "[| f -- a --> L; 0 \s>0.\x. x \ a & norm (x-a) < s --> norm (f x - L) < r" +by (simp add: LIM_eq) + +lemma LIM_offset: "f -- a --> L \ (\x. f (x + k)) -- a - k --> L" +apply (rule LIM_I) +apply (drule_tac r="r" in LIM_D, safe) +apply (rule_tac x="s" in exI, safe) +apply (drule_tac x="x + k" in spec) +apply (simp add: compare_rls) +done + +lemma LIM_offset_zero: "f -- a --> L \ (\h. f (a + h)) -- 0 --> L" +by (drule_tac k="a" in LIM_offset, simp add: add_commute) + +lemma LIM_offset_zero_cancel: "(\h. f (a + h)) -- 0 --> L \ f -- a --> L" +by (drule_tac k="- a" in LIM_offset, simp) + +lemma LIM_const [simp]: "(%x. k) -- x --> k" +by (simp add: LIM_def) + +lemma LIM_add: + fixes f g :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes f: "f -- a --> L" and g: "g -- a --> M" + shows "(%x. f x + g(x)) -- a --> (L + M)" +proof (rule LIM_I) + fix r :: real + assume r: "0 < r" + from LIM_D [OF f half_gt_zero [OF r]] + obtain fs + where fs: "0 < fs" + and fs_lt: "\x. x \ a & norm (x-a) < fs --> norm (f x - L) < r/2" + by blast + from LIM_D [OF g half_gt_zero [OF r]] + obtain gs + where gs: "0 < gs" + and gs_lt: "\x. x \ a & norm (x-a) < gs --> norm (g x - M) < r/2" + by blast + show "\s>0.\x. x \ a \ norm (x-a) < s \ norm (f x + g x - (L + M)) < r" + proof (intro exI conjI strip) + show "0 < min fs gs" by (simp add: fs gs) + fix x :: 'a + assume "x \ a \ norm (x-a) < min fs gs" + hence "x \ a \ norm (x-a) < fs \ norm (x-a) < gs" by simp + with fs_lt gs_lt + have "norm (f x - L) < r/2" and "norm (g x - M) < r/2" by blast+ + hence "norm (f x - L) + norm (g x - M) < r" by arith + thus "norm (f x + g x - (L + M)) < r" + by (blast intro: norm_diff_triangle_ineq order_le_less_trans) + qed +qed + +lemma LIM_add_zero: + "\f -- a --> 0; g -- a --> 0\ \ (\x. f x + g x) -- a --> 0" +by (drule (1) LIM_add, simp) + +lemma minus_diff_minus: + fixes a b :: "'a::ab_group_add" + shows "(- a) - (- b) = - (a - b)" +by simp + +lemma LIM_minus: "f -- a --> L ==> (%x. -f(x)) -- a --> -L" +by (simp only: LIM_eq minus_diff_minus norm_minus_cancel) + +lemma LIM_add_minus: + "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + -g(x)) -- x --> (l + -m)" +by (intro LIM_add LIM_minus) + +lemma LIM_diff: + "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) - g(x)) -- x --> l-m" +by (simp only: diff_minus LIM_add LIM_minus) + +lemma LIM_zero: "f -- a --> l \ (\x. f x - l) -- a --> 0" +by (simp add: LIM_def) + +lemma LIM_zero_cancel: "(\x. f x - l) -- a --> 0 \ f -- a --> l" +by (simp add: LIM_def) + +lemma LIM_zero_iff: "(\x. f x - l) -- a --> 0 = f -- a --> l" +by (simp add: LIM_def) + +lemma LIM_imp_LIM: + assumes f: "f -- a --> l" + assumes le: "\x. x \ a \ norm (g x - m) \ norm (f x - l)" + shows "g -- a --> m" +apply (rule LIM_I, drule LIM_D [OF f], safe) +apply (rule_tac x="s" in exI, safe) +apply (drule_tac x="x" in spec, safe) +apply (erule (1) order_le_less_trans [OF le]) +done + +lemma LIM_norm: "f -- a --> l \ (\x. norm (f x)) -- a --> norm l" +by (erule LIM_imp_LIM, simp add: norm_triangle_ineq3) + +lemma LIM_norm_zero: "f -- a --> 0 \ (\x. norm (f x)) -- a --> 0" +by (drule LIM_norm, simp) + +lemma LIM_norm_zero_cancel: "(\x. norm (f x)) -- a --> 0 \ f -- a --> 0" +by (erule LIM_imp_LIM, simp) + +lemma LIM_norm_zero_iff: "(\x. norm (f x)) -- a --> 0 = f -- a --> 0" +by (rule iffI [OF LIM_norm_zero_cancel LIM_norm_zero]) + +lemma LIM_rabs: "f -- a --> (l::real) \ (\x. \f x\) -- a --> \l\" +by (fold real_norm_def, rule LIM_norm) + +lemma LIM_rabs_zero: "f -- a --> (0::real) \ (\x. \f x\) -- a --> 0" +by (fold real_norm_def, rule LIM_norm_zero) + +lemma LIM_rabs_zero_cancel: "(\x. \f x\) -- a --> (0::real) \ f -- a --> 0" +by (fold real_norm_def, rule LIM_norm_zero_cancel) + +lemma LIM_rabs_zero_iff: "(\x. \f x\) -- a --> (0::real) = f -- a --> 0" +by (fold real_norm_def, rule LIM_norm_zero_iff) + +lemma LIM_const_not_eq: + fixes a :: "'a::real_normed_algebra_1" + shows "k \ L \ \ (\x. k) -- a --> L" +apply (simp add: LIM_eq) +apply (rule_tac x="norm (k - L)" in exI, simp, safe) +apply (rule_tac x="a + of_real (s/2)" in exI, simp add: norm_of_real) +done + +lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] + +lemma LIM_const_eq: + fixes a :: "'a::real_normed_algebra_1" + shows "(\x. k) -- a --> L \ k = L" +apply (rule ccontr) +apply (blast dest: LIM_const_not_eq) +done + +lemma LIM_unique: + fixes a :: "'a::real_normed_algebra_1" + shows "\f -- a --> L; f -- a --> M\ \ L = M" +apply (drule (1) LIM_diff) +apply (auto dest!: LIM_const_eq) +done + +lemma LIM_ident [simp]: "(\x. x) -- a --> a" +by (auto simp add: LIM_def) + +text{*Limits are equal for functions equal except at limit point*} +lemma LIM_equal: + "[| \x. x \ a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)" +by (simp add: LIM_def) + +lemma LIM_cong: + "\a = b; \x. x \ b \ f x = g x; l = m\ + \ ((\x. f x) -- a --> l) = ((\x. g x) -- b --> m)" +by (simp add: LIM_def) + +lemma LIM_equal2: + assumes 1: "0 < R" + assumes 2: "\x. \x \ a; norm (x - a) < R\ \ f x = g x" + shows "g -- a --> l \ f -- a --> l" +apply (unfold LIM_def, safe) +apply (drule_tac x="r" in spec, safe) +apply (rule_tac x="min s R" in exI, safe) +apply (simp add: 1) +apply (simp add: 2) +done + +text{*Two uses in Hyperreal/Transcendental.ML*} +lemma LIM_trans: + "[| (%x. f(x) + -g(x)) -- a --> 0; g -- a --> l |] ==> f -- a --> l" +apply (drule LIM_add, assumption) +apply (auto simp add: add_assoc) +done + +lemma LIM_compose: + assumes g: "g -- l --> g l" + assumes f: "f -- a --> l" + shows "(\x. g (f x)) -- a --> g l" +proof (rule LIM_I) + fix r::real assume r: "0 < r" + obtain s where s: "0 < s" + and less_r: "\y. \y \ l; norm (y - l) < s\ \ norm (g y - g l) < r" + using LIM_D [OF g r] by fast + obtain t where t: "0 < t" + and less_s: "\x. \x \ a; norm (x - a) < t\ \ norm (f x - l) < s" + using LIM_D [OF f s] by fast + + show "\t>0. \x. x \ a \ norm (x - a) < t \ norm (g (f x) - g l) < r" + proof (rule exI, safe) + show "0 < t" using t . + next + fix x assume "x \ a" and "norm (x - a) < t" + hence "norm (f x - l) < s" by (rule less_s) + thus "norm (g (f x) - g l) < r" + using r less_r by (case_tac "f x = l", simp_all) + qed +qed + +lemma LIM_compose2: + assumes f: "f -- a --> b" + assumes g: "g -- b --> c" + assumes inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ b" + shows "(\x. g (f x)) -- a --> c" +proof (rule LIM_I) + fix r :: real + assume r: "0 < r" + obtain s where s: "0 < s" + and less_r: "\y. \y \ b; norm (y - b) < s\ \ norm (g y - c) < r" + using LIM_D [OF g r] by fast + obtain t where t: "0 < t" + and less_s: "\x. \x \ a; norm (x - a) < t\ \ norm (f x - b) < s" + using LIM_D [OF f s] by fast + obtain d where d: "0 < d" + and neq_b: "\x. \x \ a; norm (x - a) < d\ \ f x \ b" + using inj by fast + + show "\t>0. \x. x \ a \ norm (x - a) < t \ norm (g (f x) - c) < r" + proof (safe intro!: exI) + show "0 < min d t" using d t by simp + next + fix x + assume "x \ a" and "norm (x - a) < min d t" + hence "f x \ b" and "norm (f x - b) < s" + using neq_b less_s by simp_all + thus "norm (g (f x) - c) < r" + by (rule less_r) + qed +qed + +lemma LIM_o: "\g -- l --> g l; f -- a --> l\ \ (g \ f) -- a --> g l" +unfolding o_def by (rule LIM_compose) + +lemma real_LIM_sandwich_zero: + fixes f g :: "'a::real_normed_vector \ real" + assumes f: "f -- a --> 0" + assumes 1: "\x. x \ a \ 0 \ g x" + assumes 2: "\x. x \ a \ g x \ f x" + shows "g -- a --> 0" +proof (rule LIM_imp_LIM [OF f]) + fix x assume x: "x \ a" + have "norm (g x - 0) = g x" by (simp add: 1 x) + also have "g x \ f x" by (rule 2 [OF x]) + also have "f x \ \f x\" by (rule abs_ge_self) + also have "\f x\ = norm (f x - 0)" by simp + finally show "norm (g x - 0) \ norm (f x - 0)" . +qed + +text {* Bounded Linear Operators *} + +lemma (in bounded_linear) cont: "f -- a --> f a" +proof (rule LIM_I) + fix r::real assume r: "0 < r" + obtain K where K: "0 < K" and norm_le: "\x. norm (f x) \ norm x * K" + using pos_bounded by fast + show "\s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - f a) < r" + proof (rule exI, safe) + from r K show "0 < r / K" by (rule divide_pos_pos) + next + fix x assume x: "norm (x - a) < r / K" + have "norm (f x - f a) = norm (f (x - a))" by (simp only: diff) + also have "\ \ norm (x - a) * K" by (rule norm_le) + also from K x have "\ < r" by (simp only: pos_less_divide_eq) + finally show "norm (f x - f a) < r" . + qed +qed + +lemma (in bounded_linear) LIM: + "g -- a --> l \ (\x. f (g x)) -- a --> f l" +by (rule LIM_compose [OF cont]) + +lemma (in bounded_linear) LIM_zero: + "g -- a --> 0 \ (\x. f (g x)) -- a --> 0" +by (drule LIM, simp only: zero) + +text {* Bounded Bilinear Operators *} + +lemma (in bounded_bilinear) LIM_prod_zero: + assumes f: "f -- a --> 0" + assumes g: "g -- a --> 0" + shows "(\x. f x ** g x) -- a --> 0" +proof (rule LIM_I) + fix r::real assume r: "0 < r" + obtain K where K: "0 < K" + and norm_le: "\x y. norm (x ** y) \ norm x * norm y * K" + using pos_bounded by fast + from K have K': "0 < inverse K" + by (rule positive_imp_inverse_positive) + obtain s where s: "0 < s" + and norm_f: "\x. \x \ a; norm (x - a) < s\ \ norm (f x) < r" + using LIM_D [OF f r] by auto + obtain t where t: "0 < t" + and norm_g: "\x. \x \ a; norm (x - a) < t\ \ norm (g x) < inverse K" + using LIM_D [OF g K'] by auto + show "\s>0. \x. x \ a \ norm (x - a) < s \ norm (f x ** g x - 0) < r" + proof (rule exI, safe) + from s t show "0 < min s t" by simp + next + fix x assume x: "x \ a" + assume "norm (x - a) < min s t" + hence xs: "norm (x - a) < s" and xt: "norm (x - a) < t" by simp_all + from x xs have 1: "norm (f x) < r" by (rule norm_f) + from x xt have 2: "norm (g x) < inverse K" by (rule norm_g) + have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) + also from 1 2 K have "\ < r * inverse K * K" + by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero) + also from K have "r * inverse K * K = r" by simp + finally show "norm (f x ** g x - 0) < r" by simp + qed +qed + +lemma (in bounded_bilinear) LIM_left_zero: + "f -- a --> 0 \ (\x. f x ** c) -- a --> 0" +by (rule bounded_linear.LIM_zero [OF bounded_linear_left]) + +lemma (in bounded_bilinear) LIM_right_zero: + "f -- a --> 0 \ (\x. c ** f x) -- a --> 0" +by (rule bounded_linear.LIM_zero [OF bounded_linear_right]) + +lemma (in bounded_bilinear) LIM: + "\f -- a --> L; g -- a --> M\ \ (\x. f x ** g x) -- a --> L ** M" +apply (drule LIM_zero) +apply (drule LIM_zero) +apply (rule LIM_zero_cancel) +apply (subst prod_diff_prod) +apply (rule LIM_add_zero) +apply (rule LIM_add_zero) +apply (erule (1) LIM_prod_zero) +apply (erule LIM_left_zero) +apply (erule LIM_right_zero) +done + +lemmas LIM_mult = mult.LIM + +lemmas LIM_mult_zero = mult.LIM_prod_zero + +lemmas LIM_mult_left_zero = mult.LIM_left_zero + +lemmas LIM_mult_right_zero = mult.LIM_right_zero + +lemmas LIM_scaleR = scaleR.LIM + +lemmas LIM_of_real = of_real.LIM + +lemma LIM_power: + fixes f :: "'a::real_normed_vector \ 'b::{recpower,real_normed_algebra}" + assumes f: "f -- a --> l" + shows "(\x. f x ^ n) -- a --> l ^ n" +by (induct n, simp, simp add: power_Suc LIM_mult f) + +subsubsection {* Derived theorems about @{term LIM} *} + +lemma LIM_inverse_lemma: + fixes x :: "'a::real_normed_div_algebra" + assumes r: "0 < r" + assumes x: "norm (x - 1) < min (1/2) (r/2)" + shows "norm (inverse x - 1) < r" +proof - + from r have r2: "0 < r/2" by simp + from x have 0: "x \ 0" by clarsimp + from x have x': "norm (1 - x) < min (1/2) (r/2)" + by (simp only: norm_minus_commute) + hence less1: "norm (1 - x) < r/2" by simp + have "norm (1::'a) - norm x \ norm (1 - x)" by (rule norm_triangle_ineq2) + also from x' have "norm (1 - x) < 1/2" by simp + finally have "1/2 < norm x" by simp + hence "inverse (norm x) < inverse (1/2)" + by (rule less_imp_inverse_less, simp) + hence less2: "norm (inverse x) < 2" + by (simp add: nonzero_norm_inverse 0) + from less1 less2 r2 norm_ge_zero + have "norm (1 - x) * norm (inverse x) < (r/2) * 2" + by (rule mult_strict_mono) + thus "norm (inverse x - 1) < r" + by (simp only: norm_mult [symmetric] left_diff_distrib, simp add: 0) +qed + +lemma LIM_inverse_fun: + assumes a: "a \ (0::'a::real_normed_div_algebra)" + shows "inverse -- a --> inverse a" +proof (rule LIM_equal2) + from a show "0 < norm a" by simp +next + fix x assume "norm (x - a) < norm a" + hence "x \ 0" by auto + with a show "inverse x = inverse (inverse a * x) * inverse a" + by (simp add: nonzero_inverse_mult_distrib + nonzero_imp_inverse_nonzero + nonzero_inverse_inverse_eq mult_assoc) +next + have 1: "inverse -- 1 --> inverse (1::'a)" + apply (rule LIM_I) + apply (rule_tac x="min (1/2) (r/2)" in exI) + apply (simp add: LIM_inverse_lemma) + done + have "(\x. inverse a * x) -- a --> inverse a * a" + by (intro LIM_mult LIM_ident LIM_const) + hence "(\x. inverse a * x) -- a --> 1" + by (simp add: a) + with 1 have "(\x. inverse (inverse a * x)) -- a --> inverse 1" + by (rule LIM_compose) + hence "(\x. inverse (inverse a * x)) -- a --> 1" + by simp + hence "(\x. inverse (inverse a * x) * inverse a) -- a --> 1 * inverse a" + by (intro LIM_mult LIM_const) + thus "(\x. inverse (inverse a * x) * inverse a) -- a --> inverse a" + by simp +qed + +lemma LIM_inverse: + fixes L :: "'a::real_normed_div_algebra" + shows "\f -- a --> L; L \ 0\ \ (\x. inverse (f x)) -- a --> inverse L" +by (rule LIM_inverse_fun [THEN LIM_compose]) + + +subsection {* Continuity *} + +subsubsection {* Purely standard proofs *} + +lemma LIM_isCont_iff: "(f -- a --> f a) = ((\h. f (a + h)) -- 0 --> f a)" +by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) + +lemma isCont_iff: "isCont f x = (\h. f (x + h)) -- 0 --> f x" +by (simp add: isCont_def LIM_isCont_iff) + +lemma isCont_ident [simp]: "isCont (\x. x) a" + unfolding isCont_def by (rule LIM_ident) + +lemma isCont_const [simp]: "isCont (\x. k) a" + unfolding isCont_def by (rule LIM_const) + +lemma isCont_norm: "isCont f a \ isCont (\x. norm (f x)) a" + unfolding isCont_def by (rule LIM_norm) + +lemma isCont_rabs: "isCont f a \ isCont (\x. \f x :: real\) a" + unfolding isCont_def by (rule LIM_rabs) + +lemma isCont_add: "\isCont f a; isCont g a\ \ isCont (\x. f x + g x) a" + unfolding isCont_def by (rule LIM_add) + +lemma isCont_minus: "isCont f a \ isCont (\x. - f x) a" + unfolding isCont_def by (rule LIM_minus) + +lemma isCont_diff: "\isCont f a; isCont g a\ \ isCont (\x. f x - g x) a" + unfolding isCont_def by (rule LIM_diff) + +lemma isCont_mult: + fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" + shows "\isCont f a; isCont g a\ \ isCont (\x. f x * g x) a" + unfolding isCont_def by (rule LIM_mult) + +lemma isCont_inverse: + fixes f :: "'a::real_normed_vector \ 'b::real_normed_div_algebra" + shows "\isCont f a; f a \ 0\ \ isCont (\x. inverse (f x)) a" + unfolding isCont_def by (rule LIM_inverse) + +lemma isCont_LIM_compose: + "\isCont g l; f -- a --> l\ \ (\x. g (f x)) -- a --> g l" + unfolding isCont_def by (rule LIM_compose) + +lemma isCont_LIM_compose2: + assumes f [unfolded isCont_def]: "isCont f a" + assumes g: "g -- f a --> l" + assumes inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ f a" + shows "(\x. g (f x)) -- a --> l" +by (rule LIM_compose2 [OF f g inj]) + +lemma isCont_o2: "\isCont f a; isCont g (f a)\ \ isCont (\x. g (f x)) a" + unfolding isCont_def by (rule LIM_compose) + +lemma isCont_o: "\isCont f a; isCont g (f a)\ \ isCont (g o f) a" + unfolding o_def by (rule isCont_o2) + +lemma (in bounded_linear) isCont: "isCont f a" + unfolding isCont_def by (rule cont) + +lemma (in bounded_bilinear) isCont: + "\isCont f a; isCont g a\ \ isCont (\x. f x ** g x) a" + unfolding isCont_def by (rule LIM) + +lemmas isCont_scaleR = scaleR.isCont + +lemma isCont_of_real: + "isCont f a \ isCont (\x. of_real (f x)) a" + unfolding isCont_def by (rule LIM_of_real) + +lemma isCont_power: + fixes f :: "'a::real_normed_vector \ 'b::{recpower,real_normed_algebra}" + shows "isCont f a \ isCont (\x. f x ^ n) a" + unfolding isCont_def by (rule LIM_power) + +lemma isCont_abs [simp]: "isCont abs (a::real)" +by (rule isCont_rabs [OF isCont_ident]) + + +subsection {* Uniform Continuity *} + +lemma isUCont_isCont: "isUCont f ==> isCont f x" +by (simp add: isUCont_def isCont_def LIM_def, force) + +lemma isUCont_Cauchy: + "\isUCont f; Cauchy X\ \ Cauchy (\n. f (X n))" +unfolding isUCont_def +apply (rule CauchyI) +apply (drule_tac x=e in spec, safe) +apply (drule_tac e=s in CauchyD, safe) +apply (rule_tac x=M in exI, simp) +done + +lemma (in bounded_linear) isUCont: "isUCont f" +unfolding isUCont_def +proof (intro allI impI) + fix r::real assume r: "0 < r" + obtain K where K: "0 < K" and norm_le: "\x. norm (f x) \ norm x * K" + using pos_bounded by fast + show "\s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r" + proof (rule exI, safe) + from r K show "0 < r / K" by (rule divide_pos_pos) + next + fix x y :: 'a + assume xy: "norm (x - y) < r / K" + have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff) + also have "\ \ norm (x - y) * K" by (rule norm_le) + also from K xy have "\ < r" by (simp only: pos_less_divide_eq) + finally show "norm (f x - f y) < r" . + qed +qed + +lemma (in bounded_linear) Cauchy: "Cauchy X \ Cauchy (\n. f (X n))" +by (rule isUCont [THEN isUCont_Cauchy]) + + +subsection {* Relation of LIM and LIMSEQ *} + +lemma LIMSEQ_SEQ_conv1: + fixes a :: "'a::real_normed_vector" + assumes X: "X -- a --> L" + shows "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" +proof (safe intro!: LIMSEQ_I) + fix S :: "nat \ 'a" + fix r :: real + assume rgz: "0 < r" + assume as: "\n. S n \ a" + assume S: "S ----> a" + from LIM_D [OF X rgz] obtain s + where sgz: "0 < s" + and aux: "\x. \x \ a; norm (x - a) < s\ \ norm (X x - L) < r" + by fast + from LIMSEQ_D [OF S sgz] + obtain no where "\n\no. norm (S n - a) < s" by blast + hence "\n\no. norm (X (S n) - L) < r" by (simp add: aux as) + thus "\no. \n\no. norm (X (S n) - L) < r" .. +qed + +lemma LIMSEQ_SEQ_conv2: + fixes a :: real + assumes "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" + shows "X -- a --> L" +proof (rule ccontr) + assume "\ (X -- a --> L)" + hence "\ (\r > 0. \s > 0. \x. x \ a & norm (x - a) < s --> norm (X x - L) < r)" by (unfold LIM_def) + hence "\r > 0. \s > 0. \x. \(x \ a \ \x - a\ < s --> norm (X x - L) < r)" by simp + hence "\r > 0. \s > 0. \x. (x \ a \ \x - a\ < s \ norm (X x - L) \ r)" by (simp add: linorder_not_less) + then obtain r where rdef: "r > 0 \ (\s > 0. \x. (x \ a \ \x - a\ < s \ norm (X x - L) \ r))" by auto + + let ?F = "\n::nat. SOME x. x\a \ \x - a\ < inverse (real (Suc n)) \ norm (X x - L) \ r" + have "\n. \x. x\a \ \x - a\ < inverse (real (Suc n)) \ norm (X x - L) \ r" + using rdef by simp + hence F: "\n. ?F n \ a \ \?F n - a\ < inverse (real (Suc n)) \ norm (X (?F n) - L) \ r" + by (rule someI_ex) + hence F1: "\n. ?F n \ a" + and F2: "\n. \?F n - a\ < inverse (real (Suc n))" + and F3: "\n. norm (X (?F n) - L) \ r" + by fast+ + + have "?F ----> a" + proof (rule LIMSEQ_I, unfold real_norm_def) + fix e::real + assume "0 < e" + (* choose no such that inverse (real (Suc n)) < e *) + then have "\no. inverse (real (Suc no)) < e" by (rule reals_Archimedean) + then obtain m where nodef: "inverse (real (Suc m)) < e" by auto + show "\no. \n. no \ n --> \?F n - a\ < e" + proof (intro exI allI impI) + fix n + assume mlen: "m \ n" + have "\?F n - a\ < inverse (real (Suc n))" + by (rule F2) + also have "inverse (real (Suc n)) \ inverse (real (Suc m))" + using mlen by auto + also from nodef have + "inverse (real (Suc m)) < e" . + finally show "\?F n - a\ < e" . + qed + qed + + moreover have "\n. ?F n \ a" + by (rule allI) (rule F1) + + moreover from prems have "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" by simp + ultimately have "(\n. X (?F n)) ----> L" by simp + + moreover have "\ ((\n. X (?F n)) ----> L)" + proof - + { + fix no::nat + obtain n where "n = no + 1" by simp + then have nolen: "no \ n" by simp + (* We prove this by showing that for any m there is an n\m such that |X (?F n) - L| \ r *) + have "norm (X (?F n) - L) \ r" + by (rule F3) + with nolen have "\n. no \ n \ norm (X (?F n) - L) \ r" by fast + } + then have "(\no. \n. no \ n \ norm (X (?F n) - L) \ r)" by simp + with rdef have "\e>0. (\no. \n. no \ n \ norm (X (?F n) - L) \ e)" by auto + thus ?thesis by (unfold LIMSEQ_def, auto simp add: linorder_not_less) + qed + ultimately show False by simp +qed + +lemma LIMSEQ_SEQ_conv: + "(\S. (\n. S n \ a) \ S ----> (a::real) \ (\n. X (S n)) ----> L) = + (X -- a --> L)" +proof + assume "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" + thus "X -- a --> L" by (rule LIMSEQ_SEQ_conv2) +next + assume "(X -- a --> L)" + thus "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" by (rule LIMSEQ_SEQ_conv1) +qed + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Ln.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Ln.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,430 @@ +(* Title: Ln.thy + Author: Jeremy Avigad +*) + +header {* Properties of ln *} + +theory Ln +imports Transcendental +begin + +lemma exp_first_two_terms: "exp x = 1 + x + suminf (%n. + inverse(real (fact (n+2))) * (x ^ (n+2)))" +proof - + have "exp x = suminf (%n. inverse(real (fact n)) * (x ^ n))" + by (simp add: exp_def) + also from summable_exp have "... = (SUM n : {0..<2}. + inverse(real (fact n)) * (x ^ n)) + suminf (%n. + inverse(real (fact (n+2))) * (x ^ (n+2)))" (is "_ = ?a + _") + by (rule suminf_split_initial_segment) + also have "?a = 1 + x" + by (simp add: numerals) + finally show ?thesis . +qed + +lemma exp_tail_after_first_two_terms_summable: + "summable (%n. inverse(real (fact (n+2))) * (x ^ (n+2)))" +proof - + note summable_exp + thus ?thesis + by (frule summable_ignore_initial_segment) +qed + +lemma aux1: assumes a: "0 <= x" and b: "x <= 1" + shows "inverse (real (fact (n + 2))) * x ^ (n + 2) <= (x^2/2) * ((1/2)^n)" +proof (induct n) + show "inverse (real (fact (0 + 2))) * x ^ (0 + 2) <= + x ^ 2 / 2 * (1 / 2) ^ 0" + by (simp add: real_of_nat_Suc power2_eq_square) +next + fix n + assume c: "inverse (real (fact (n + 2))) * x ^ (n + 2) + <= x ^ 2 / 2 * (1 / 2) ^ n" + show "inverse (real (fact (Suc n + 2))) * x ^ (Suc n + 2) + <= x ^ 2 / 2 * (1 / 2) ^ Suc n" + proof - + have "inverse(real (fact (Suc n + 2))) <= + (1 / 2) *inverse (real (fact (n+2)))" + proof - + have "Suc n + 2 = Suc (n + 2)" by simp + then have "fact (Suc n + 2) = Suc (n + 2) * fact (n + 2)" + by simp + then have "real(fact (Suc n + 2)) = real(Suc (n + 2) * fact (n + 2))" + apply (rule subst) + apply (rule refl) + done + also have "... = real(Suc (n + 2)) * real(fact (n + 2))" + by (rule real_of_nat_mult) + finally have "real (fact (Suc n + 2)) = + real (Suc (n + 2)) * real (fact (n + 2))" . + then have "inverse(real (fact (Suc n + 2))) = + inverse(real (Suc (n + 2))) * inverse(real (fact (n + 2)))" + apply (rule ssubst) + apply (rule inverse_mult_distrib) + done + also have "... <= (1/2) * inverse(real (fact (n + 2)))" + apply (rule mult_right_mono) + apply (subst inverse_eq_divide) + apply simp + apply (rule inv_real_of_nat_fact_ge_zero) + done + finally show ?thesis . + qed + moreover have "x ^ (Suc n + 2) <= x ^ (n + 2)" + apply (simp add: mult_compare_simps) + apply (simp add: prems) + apply (subgoal_tac "0 <= x * (x * x^n)") + apply force + apply (rule mult_nonneg_nonneg, rule a)+ + apply (rule zero_le_power, rule a) + done + ultimately have "inverse (real (fact (Suc n + 2))) * x ^ (Suc n + 2) <= + (1 / 2 * inverse (real (fact (n + 2)))) * x ^ (n + 2)" + apply (rule mult_mono) + apply (rule mult_nonneg_nonneg) + apply simp + apply (subst inverse_nonnegative_iff_nonnegative) + apply (rule real_of_nat_ge_zero) + apply (rule zero_le_power) + apply (rule a) + done + also have "... = 1 / 2 * (inverse (real (fact (n + 2))) * x ^ (n + 2))" + by simp + also have "... <= 1 / 2 * (x ^ 2 / 2 * (1 / 2) ^ n)" + apply (rule mult_left_mono) + apply (rule prems) + apply simp + done + also have "... = x ^ 2 / 2 * (1 / 2 * (1 / 2) ^ n)" + by auto + also have "(1::real) / 2 * (1 / 2) ^ n = (1 / 2) ^ (Suc n)" + by (rule realpow_Suc [THEN sym]) + finally show ?thesis . + qed +qed + +lemma aux2: "(%n. (x::real) ^ 2 / 2 * (1 / 2) ^ n) sums x^2" +proof - + have "(%n. (1 / 2::real)^n) sums (1 / (1 - (1/2)))" + apply (rule geometric_sums) + by (simp add: abs_less_iff) + also have "(1::real) / (1 - 1/2) = 2" + by simp + finally have "(%n. (1 / 2::real)^n) sums 2" . + then have "(%n. x ^ 2 / 2 * (1 / 2) ^ n) sums (x^2 / 2 * 2)" + by (rule sums_mult) + also have "x^2 / 2 * 2 = x^2" + by simp + finally show ?thesis . +qed + +lemma exp_bound: "0 <= (x::real) ==> x <= 1 ==> exp x <= 1 + x + x^2" +proof - + assume a: "0 <= x" + assume b: "x <= 1" + have c: "exp x = 1 + x + suminf (%n. inverse(real (fact (n+2))) * + (x ^ (n+2)))" + by (rule exp_first_two_terms) + moreover have "suminf (%n. inverse(real (fact (n+2))) * (x ^ (n+2))) <= x^2" + proof - + have "suminf (%n. inverse(real (fact (n+2))) * (x ^ (n+2))) <= + suminf (%n. (x^2/2) * ((1/2)^n))" + apply (rule summable_le) + apply (auto simp only: aux1 prems) + apply (rule exp_tail_after_first_two_terms_summable) + by (rule sums_summable, rule aux2) + also have "... = x^2" + by (rule sums_unique [THEN sym], rule aux2) + finally show ?thesis . + qed + ultimately show ?thesis + by auto +qed + +lemma aux4: "0 <= (x::real) ==> x <= 1 ==> exp (x - x^2) <= 1 + x" +proof - + assume a: "0 <= x" and b: "x <= 1" + have "exp (x - x^2) = exp x / exp (x^2)" + by (rule exp_diff) + also have "... <= (1 + x + x^2) / exp (x ^2)" + apply (rule divide_right_mono) + apply (rule exp_bound) + apply (rule a, rule b) + apply simp + done + also have "... <= (1 + x + x^2) / (1 + x^2)" + apply (rule divide_left_mono) + apply (auto simp add: exp_ge_add_one_self_aux) + apply (rule add_nonneg_nonneg) + apply (insert prems, auto) + apply (rule mult_pos_pos) + apply auto + apply (rule add_pos_nonneg) + apply auto + done + also from a have "... <= 1 + x" + by(simp add:field_simps zero_compare_simps) + finally show ?thesis . +qed + +lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==> + x - x^2 <= ln (1 + x)" +proof - + assume a: "0 <= x" and b: "x <= 1" + then have "exp (x - x^2) <= 1 + x" + by (rule aux4) + also have "... = exp (ln (1 + x))" + proof - + from a have "0 < 1 + x" by auto + thus ?thesis + by (auto simp only: exp_ln_iff [THEN sym]) + qed + finally have "exp (x - x ^ 2) <= exp (ln (1 + x))" . + thus ?thesis by (auto simp only: exp_le_cancel_iff) +qed + +lemma ln_one_minus_pos_upper_bound: "0 <= x ==> x < 1 ==> ln (1 - x) <= - x" +proof - + assume a: "0 <= (x::real)" and b: "x < 1" + have "(1 - x) * (1 + x + x^2) = (1 - x^3)" + by (simp add: ring_simps power2_eq_square power3_eq_cube) + also have "... <= 1" + by (auto simp add: a) + finally have "(1 - x) * (1 + x + x ^ 2) <= 1" . + moreover have "0 < 1 + x + x^2" + apply (rule add_pos_nonneg) + apply (insert a, auto) + done + ultimately have "1 - x <= 1 / (1 + x + x^2)" + by (elim mult_imp_le_div_pos) + also have "... <= 1 / exp x" + apply (rule divide_left_mono) + apply (rule exp_bound, rule a) + apply (insert prems, auto) + apply (rule mult_pos_pos) + apply (rule add_pos_nonneg) + apply auto + done + also have "... = exp (-x)" + by (auto simp add: exp_minus real_divide_def) + finally have "1 - x <= exp (- x)" . + also have "1 - x = exp (ln (1 - x))" + proof - + have "0 < 1 - x" + by (insert b, auto) + thus ?thesis + by (auto simp only: exp_ln_iff [THEN sym]) + qed + finally have "exp (ln (1 - x)) <= exp (- x)" . + thus ?thesis by (auto simp only: exp_le_cancel_iff) +qed + +lemma aux5: "x < 1 ==> ln(1 - x) = - ln(1 + x / (1 - x))" +proof - + assume a: "x < 1" + have "ln(1 - x) = - ln(1 / (1 - x))" + proof - + have "ln(1 - x) = - (- ln (1 - x))" + by auto + also have "- ln(1 - x) = ln 1 - ln(1 - x)" + by simp + also have "... = ln(1 / (1 - x))" + apply (rule ln_div [THEN sym]) + by (insert a, auto) + finally show ?thesis . + qed + also have " 1 / (1 - x) = 1 + x / (1 - x)" using a by(simp add:field_simps) + finally show ?thesis . +qed + +lemma ln_one_minus_pos_lower_bound: "0 <= x ==> x <= (1 / 2) ==> + - x - 2 * x^2 <= ln (1 - x)" +proof - + assume a: "0 <= x" and b: "x <= (1 / 2)" + from b have c: "x < 1" + by auto + then have "ln (1 - x) = - ln (1 + x / (1 - x))" + by (rule aux5) + also have "- (x / (1 - x)) <= ..." + proof - + have "ln (1 + x / (1 - x)) <= x / (1 - x)" + apply (rule ln_add_one_self_le_self) + apply (rule divide_nonneg_pos) + by (insert a c, auto) + thus ?thesis + by auto + qed + also have "- (x / (1 - x)) = -x / (1 - x)" + by auto + finally have d: "- x / (1 - x) <= ln (1 - x)" . + have "0 < 1 - x" using prems by simp + hence e: "-x - 2 * x^2 <= - x / (1 - x)" + using mult_right_le_one_le[of "x*x" "2*x"] prems + by(simp add:field_simps power2_eq_square) + from e d show "- x - 2 * x^2 <= ln (1 - x)" + by (rule order_trans) +qed + +lemma exp_ge_add_one_self [simp]: "1 + (x::real) <= exp x" + apply (case_tac "0 <= x") + apply (erule exp_ge_add_one_self_aux) + apply (case_tac "x <= -1") + apply (subgoal_tac "1 + x <= 0") + apply (erule order_trans) + apply simp + apply simp + apply (subgoal_tac "1 + x = exp(ln (1 + x))") + apply (erule ssubst) + apply (subst exp_le_cancel_iff) + apply (subgoal_tac "ln (1 - (- x)) <= - (- x)") + apply simp + apply (rule ln_one_minus_pos_upper_bound) + apply auto +done + +lemma ln_add_one_self_le_self2: "-1 < x ==> ln(1 + x) <= x" + apply (subgoal_tac "x = ln (exp x)") + apply (erule ssubst)back + apply (subst ln_le_cancel_iff) + apply auto +done + +lemma abs_ln_one_plus_x_minus_x_bound_nonneg: + "0 <= x ==> x <= 1 ==> abs(ln (1 + x) - x) <= x^2" +proof - + assume x: "0 <= x" + assume "x <= 1" + from x have "ln (1 + x) <= x" + by (rule ln_add_one_self_le_self) + then have "ln (1 + x) - x <= 0" + by simp + then have "abs(ln(1 + x) - x) = - (ln(1 + x) - x)" + by (rule abs_of_nonpos) + also have "... = x - ln (1 + x)" + by simp + also have "... <= x^2" + proof - + from prems have "x - x^2 <= ln (1 + x)" + by (intro ln_one_plus_pos_lower_bound) + thus ?thesis + by simp + qed + finally show ?thesis . +qed + +lemma abs_ln_one_plus_x_minus_x_bound_nonpos: + "-(1 / 2) <= x ==> x <= 0 ==> abs(ln (1 + x) - x) <= 2 * x^2" +proof - + assume "-(1 / 2) <= x" + assume "x <= 0" + have "abs(ln (1 + x) - x) = x - ln(1 - (-x))" + apply (subst abs_of_nonpos) + apply simp + apply (rule ln_add_one_self_le_self2) + apply (insert prems, auto) + done + also have "... <= 2 * x^2" + apply (subgoal_tac "- (-x) - 2 * (-x)^2 <= ln (1 - (-x))") + apply (simp add: compare_rls) + apply (rule ln_one_minus_pos_lower_bound) + apply (insert prems, auto) + done + finally show ?thesis . +qed + +lemma abs_ln_one_plus_x_minus_x_bound: + "abs x <= 1 / 2 ==> abs(ln (1 + x) - x) <= 2 * x^2" + apply (case_tac "0 <= x") + apply (rule order_trans) + apply (rule abs_ln_one_plus_x_minus_x_bound_nonneg) + apply auto + apply (rule abs_ln_one_plus_x_minus_x_bound_nonpos) + apply auto +done + +lemma DERIV_ln: "0 < x ==> DERIV ln x :> 1 / x" + apply (unfold deriv_def, unfold LIM_def, clarsimp) + apply (rule exI) + apply (rule conjI) + prefer 2 + apply clarsimp + apply (subgoal_tac "(ln (x + xa) - ln x) / xa - (1 / x) = + (ln (1 + xa / x) - xa / x) / xa") + apply (erule ssubst) + apply (subst abs_divide) + apply (rule mult_imp_div_pos_less) + apply force + apply (rule order_le_less_trans) + apply (rule abs_ln_one_plus_x_minus_x_bound) + apply (subst abs_divide) + apply (subst abs_of_pos, assumption) + apply (erule mult_imp_div_pos_le) + apply (subgoal_tac "abs xa < min (x / 2) (r * x^2 / 2)") + apply force + apply assumption + apply (simp add: power2_eq_square mult_compare_simps) + apply (rule mult_imp_div_pos_less) + apply (rule mult_pos_pos, assumption, assumption) + apply (subgoal_tac "xa * xa = abs xa * abs xa") + apply (erule ssubst) + apply (subgoal_tac "abs xa * (abs xa * 2) < abs xa * (r * (x * x))") + apply (simp only: mult_ac) + apply (rule mult_strict_left_mono) + apply (erule conjE, assumption) + apply force + apply simp + apply (subst ln_div [THEN sym]) + apply arith + apply (auto simp add: ring_simps add_frac_eq frac_eq_eq + add_divide_distrib power2_eq_square) + apply (rule mult_pos_pos, assumption)+ + apply assumption +done + +lemma ln_x_over_x_mono: "exp 1 <= x ==> x <= y ==> (ln y / y) <= (ln x / x)" +proof - + assume "exp 1 <= x" and "x <= y" + have a: "0 < x" and b: "0 < y" + apply (insert prems) + apply (subgoal_tac "0 < exp (1::real)") + apply arith + apply auto + apply (subgoal_tac "0 < exp (1::real)") + apply arith + apply auto + done + have "x * ln y - x * ln x = x * (ln y - ln x)" + by (simp add: ring_simps) + also have "... = x * ln(y / x)" + apply (subst ln_div) + apply (rule b, rule a, rule refl) + done + also have "y / x = (x + (y - x)) / x" + by simp + also have "... = 1 + (y - x) / x" using a prems by(simp add:field_simps) + also have "x * ln(1 + (y - x) / x) <= x * ((y - x) / x)" + apply (rule mult_left_mono) + apply (rule ln_add_one_self_le_self) + apply (rule divide_nonneg_pos) + apply (insert prems a, simp_all) + done + also have "... = y - x" using a by simp + also have "... = (y - x) * ln (exp 1)" by simp + also have "... <= (y - x) * ln x" + apply (rule mult_left_mono) + apply (subst ln_le_cancel_iff) + apply force + apply (rule a) + apply (rule prems) + apply (insert prems, simp) + done + also have "... = y * ln x - x * ln x" + by (rule left_diff_distrib) + finally have "x * ln y <= y * ln x" + by arith + then have "ln y <= (y * ln x) / x" using a by(simp add:field_simps) + also have "... = y * (ln x / x)" by simp + finally show ?thesis using b by(simp add:field_simps) +qed + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Log.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Log.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,276 @@ +(* Title : Log.thy + Author : Jacques D. Fleuriot + Additional contributions by Jeremy Avigad + Copyright : 2000,2001 University of Edinburgh +*) + +header{*Logarithms: Standard Version*} + +theory Log +imports Transcendental +begin + +definition + powr :: "[real,real] => real" (infixr "powr" 80) where + --{*exponentation with real exponent*} + "x powr a = exp(a * ln x)" + +definition + log :: "[real,real] => real" where + --{*logarithm of @{term x} to base @{term a}*} + "log a x = ln x / ln a" + + + +lemma powr_one_eq_one [simp]: "1 powr a = 1" +by (simp add: powr_def) + +lemma powr_zero_eq_one [simp]: "x powr 0 = 1" +by (simp add: powr_def) + +lemma powr_one_gt_zero_iff [simp]: "(x powr 1 = x) = (0 < x)" +by (simp add: powr_def) +declare powr_one_gt_zero_iff [THEN iffD2, simp] + +lemma powr_mult: + "[| 0 < x; 0 < y |] ==> (x * y) powr a = (x powr a) * (y powr a)" +by (simp add: powr_def exp_add [symmetric] ln_mult right_distrib) + +lemma powr_gt_zero [simp]: "0 < x powr a" +by (simp add: powr_def) + +lemma powr_ge_pzero [simp]: "0 <= x powr y" +by (rule order_less_imp_le, rule powr_gt_zero) + +lemma powr_not_zero [simp]: "x powr a \ 0" +by (simp add: powr_def) + +lemma powr_divide: + "[| 0 < x; 0 < y |] ==> (x / y) powr a = (x powr a)/(y powr a)" +apply (simp add: divide_inverse positive_imp_inverse_positive powr_mult) +apply (simp add: powr_def exp_minus [symmetric] exp_add [symmetric] ln_inverse) +done + +lemma powr_divide2: "x powr a / x powr b = x powr (a - b)" + apply (simp add: powr_def) + apply (subst exp_diff [THEN sym]) + apply (simp add: left_diff_distrib) +done + +lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)" +by (simp add: powr_def exp_add [symmetric] left_distrib) + +lemma powr_powr: "(x powr a) powr b = x powr (a * b)" +by (simp add: powr_def) + +lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a" +by (simp add: powr_powr real_mult_commute) + +lemma powr_minus: "x powr (-a) = inverse (x powr a)" +by (simp add: powr_def exp_minus [symmetric]) + +lemma powr_minus_divide: "x powr (-a) = 1/(x powr a)" +by (simp add: divide_inverse powr_minus) + +lemma powr_less_mono: "[| a < b; 1 < x |] ==> x powr a < x powr b" +by (simp add: powr_def) + +lemma powr_less_cancel: "[| x powr a < x powr b; 1 < x |] ==> a < b" +by (simp add: powr_def) + +lemma powr_less_cancel_iff [simp]: "1 < x ==> (x powr a < x powr b) = (a < b)" +by (blast intro: powr_less_cancel powr_less_mono) + +lemma powr_le_cancel_iff [simp]: "1 < x ==> (x powr a \ x powr b) = (a \ b)" +by (simp add: linorder_not_less [symmetric]) + +lemma log_ln: "ln x = log (exp(1)) x" +by (simp add: log_def) + +lemma powr_log_cancel [simp]: + "[| 0 < a; a \ 1; 0 < x |] ==> a powr (log a x) = x" +by (simp add: powr_def log_def) + +lemma log_powr_cancel [simp]: "[| 0 < a; a \ 1 |] ==> log a (a powr y) = y" +by (simp add: log_def powr_def) + +lemma log_mult: + "[| 0 < a; a \ 1; 0 < x; 0 < y |] + ==> log a (x * y) = log a x + log a y" +by (simp add: log_def ln_mult divide_inverse left_distrib) + +lemma log_eq_div_ln_mult_log: + "[| 0 < a; a \ 1; 0 < b; b \ 1; 0 < x |] + ==> log a x = (ln b/ln a) * log b x" +by (simp add: log_def divide_inverse) + +text{*Base 10 logarithms*} +lemma log_base_10_eq1: "0 < x ==> log 10 x = (ln (exp 1) / ln 10) * ln x" +by (simp add: log_def) + +lemma log_base_10_eq2: "0 < x ==> log 10 x = (log 10 (exp 1)) * ln x" +by (simp add: log_def) + +lemma log_one [simp]: "log a 1 = 0" +by (simp add: log_def) + +lemma log_eq_one [simp]: "[| 0 < a; a \ 1 |] ==> log a a = 1" +by (simp add: log_def) + +lemma log_inverse: + "[| 0 < a; a \ 1; 0 < x |] ==> log a (inverse x) = - log a x" +apply (rule_tac a1 = "log a x" in add_left_cancel [THEN iffD1]) +apply (simp add: log_mult [symmetric]) +done + +lemma log_divide: + "[|0 < a; a \ 1; 0 < x; 0 < y|] ==> log a (x/y) = log a x - log a y" +by (simp add: log_mult divide_inverse log_inverse) + +lemma log_less_cancel_iff [simp]: + "[| 1 < a; 0 < x; 0 < y |] ==> (log a x < log a y) = (x < y)" +apply safe +apply (rule_tac [2] powr_less_cancel) +apply (drule_tac a = "log a x" in powr_less_mono, auto) +done + +lemma log_le_cancel_iff [simp]: + "[| 1 < a; 0 < x; 0 < y |] ==> (log a x \ log a y) = (x \ y)" +by (simp add: linorder_not_less [symmetric]) + + +lemma powr_realpow: "0 < x ==> x powr (real n) = x^n" + apply (induct n, simp) + apply (subgoal_tac "real(Suc n) = real n + 1") + apply (erule ssubst) + apply (subst powr_add, simp, simp) +done + +lemma powr_realpow2: "0 <= x ==> 0 < n ==> x^n = (if (x = 0) then 0 + else x powr (real n))" + apply (case_tac "x = 0", simp, simp) + apply (rule powr_realpow [THEN sym], simp) +done + +lemma ln_pwr: "0 < x ==> 0 < y ==> ln(x powr y) = y * ln x" +by (unfold powr_def, simp) + +lemma ln_bound: "1 <= x ==> ln x <= x" + apply (subgoal_tac "ln(1 + (x - 1)) <= x - 1") + apply simp + apply (rule ln_add_one_self_le_self, simp) +done + +lemma powr_mono: "a <= b ==> 1 <= x ==> x powr a <= x powr b" + apply (case_tac "x = 1", simp) + apply (case_tac "a = b", simp) + apply (rule order_less_imp_le) + apply (rule powr_less_mono, auto) +done + +lemma ge_one_powr_ge_zero: "1 <= x ==> 0 <= a ==> 1 <= x powr a" + apply (subst powr_zero_eq_one [THEN sym]) + apply (rule powr_mono, assumption+) +done + +lemma powr_less_mono2: "0 < a ==> 0 < x ==> x < y ==> x powr a < + y powr a" + apply (unfold powr_def) + apply (rule exp_less_mono) + apply (rule mult_strict_left_mono) + apply (subst ln_less_cancel_iff, assumption) + apply (rule order_less_trans) + prefer 2 + apply assumption+ +done + +lemma powr_less_mono2_neg: "a < 0 ==> 0 < x ==> x < y ==> y powr a < + x powr a" + apply (unfold powr_def) + apply (rule exp_less_mono) + apply (rule mult_strict_left_mono_neg) + apply (subst ln_less_cancel_iff) + apply assumption + apply (rule order_less_trans) + prefer 2 + apply assumption+ +done + +lemma powr_mono2: "0 <= a ==> 0 < x ==> x <= y ==> x powr a <= y powr a" + apply (case_tac "a = 0", simp) + apply (case_tac "x = y", simp) + apply (rule order_less_imp_le) + apply (rule powr_less_mono2, auto) +done + +lemma ln_powr_bound: "1 <= x ==> 0 < a ==> ln x <= (x powr a) / a" + apply (rule mult_imp_le_div_pos) + apply (assumption) + apply (subst mult_commute) + apply (subst ln_pwr [THEN sym]) + apply auto + apply (rule ln_bound) + apply (erule ge_one_powr_ge_zero) + apply (erule order_less_imp_le) +done + +lemma ln_powr_bound2: "1 < x ==> 0 < a ==> (ln x) powr a <= (a powr a) * x" +proof - + assume "1 < x" and "0 < a" + then have "ln x <= (x powr (1 / a)) / (1 / a)" + apply (intro ln_powr_bound) + apply (erule order_less_imp_le) + apply (rule divide_pos_pos) + apply simp_all + done + also have "... = a * (x powr (1 / a))" + by simp + finally have "(ln x) powr a <= (a * (x powr (1 / a))) powr a" + apply (intro powr_mono2) + apply (rule order_less_imp_le, rule prems) + apply (rule ln_gt_zero) + apply (rule prems) + apply assumption + done + also have "... = (a powr a) * ((x powr (1 / a)) powr a)" + apply (rule powr_mult) + apply (rule prems) + apply (rule powr_gt_zero) + done + also have "(x powr (1 / a)) powr a = x powr ((1 / a) * a)" + by (rule powr_powr) + also have "... = x" + apply simp + apply (subgoal_tac "a ~= 0") + apply (insert prems, auto) + done + finally show ?thesis . +qed + +lemma LIMSEQ_neg_powr: "0 < s ==> (%x. (real x) powr - s) ----> 0" + apply (unfold LIMSEQ_def) + apply clarsimp + apply (rule_tac x = "natfloor(r powr (1 / - s)) + 1" in exI) + apply clarify + proof - + fix r fix n + assume "0 < s" and "0 < r" and "natfloor (r powr (1 / - s)) + 1 <= n" + have "r powr (1 / - s) < real(natfloor(r powr (1 / - s))) + 1" + by (rule real_natfloor_add_one_gt) + also have "... = real(natfloor(r powr (1 / -s)) + 1)" + by simp + also have "... <= real n" + apply (subst real_of_nat_le_iff) + apply (rule prems) + done + finally have "r powr (1 / - s) < real n". + then have "real n powr (- s) < (r powr (1 / - s)) powr - s" + apply (intro powr_less_mono2_neg) + apply (auto simp add: prems) + done + also have "... = r" + by (simp add: powr_powr prems less_imp_neq [THEN not_sym]) + finally show "real n powr - s < r" . + qed + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Lubs.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Lubs.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,113 @@ +(* Title : Lubs.thy + ID : $Id$ + Author : Jacques D. Fleuriot + Copyright : 1998 University of Cambridge +*) + +header{*Definitions of Upper Bounds and Least Upper Bounds*} + +theory Lubs +imports Plain +begin + +text{*Thanks to suggestions by James Margetson*} + +definition + setle :: "['a set, 'a::ord] => bool" (infixl "*<=" 70) where + "S *<= x = (ALL y: S. y <= x)" + +definition + setge :: "['a::ord, 'a set] => bool" (infixl "<=*" 70) where + "x <=* S = (ALL y: S. x <= y)" + +definition + leastP :: "['a =>bool,'a::ord] => bool" where + "leastP P x = (P x & x <=* Collect P)" + +definition + isUb :: "['a set, 'a set, 'a::ord] => bool" where + "isUb R S x = (S *<= x & x: R)" + +definition + isLub :: "['a set, 'a set, 'a::ord] => bool" where + "isLub R S x = leastP (isUb R S) x" + +definition + ubs :: "['a set, 'a::ord set] => 'a set" where + "ubs R S = Collect (isUb R S)" + + + +subsection{*Rules for the Relations @{text "*<="} and @{text "<=*"}*} + +lemma setleI: "ALL y: S. y <= x ==> S *<= x" +by (simp add: setle_def) + +lemma setleD: "[| S *<= x; y: S |] ==> y <= x" +by (simp add: setle_def) + +lemma setgeI: "ALL y: S. x<= y ==> x <=* S" +by (simp add: setge_def) + +lemma setgeD: "[| x <=* S; y: S |] ==> x <= y" +by (simp add: setge_def) + + +subsection{*Rules about the Operators @{term leastP}, @{term ub} + and @{term lub}*} + +lemma leastPD1: "leastP P x ==> P x" +by (simp add: leastP_def) + +lemma leastPD2: "leastP P x ==> x <=* Collect P" +by (simp add: leastP_def) + +lemma leastPD3: "[| leastP P x; y: Collect P |] ==> x <= y" +by (blast dest!: leastPD2 setgeD) + +lemma isLubD1: "isLub R S x ==> S *<= x" +by (simp add: isLub_def isUb_def leastP_def) + +lemma isLubD1a: "isLub R S x ==> x: R" +by (simp add: isLub_def isUb_def leastP_def) + +lemma isLub_isUb: "isLub R S x ==> isUb R S x" +apply (simp add: isUb_def) +apply (blast dest: isLubD1 isLubD1a) +done + +lemma isLubD2: "[| isLub R S x; y : S |] ==> y <= x" +by (blast dest!: isLubD1 setleD) + +lemma isLubD3: "isLub R S x ==> leastP(isUb R S) x" +by (simp add: isLub_def) + +lemma isLubI1: "leastP(isUb R S) x ==> isLub R S x" +by (simp add: isLub_def) + +lemma isLubI2: "[| isUb R S x; x <=* Collect (isUb R S) |] ==> isLub R S x" +by (simp add: isLub_def leastP_def) + +lemma isUbD: "[| isUb R S x; y : S |] ==> y <= x" +by (simp add: isUb_def setle_def) + +lemma isUbD2: "isUb R S x ==> S *<= x" +by (simp add: isUb_def) + +lemma isUbD2a: "isUb R S x ==> x: R" +by (simp add: isUb_def) + +lemma isUbI: "[| S *<= x; x: R |] ==> isUb R S x" +by (simp add: isUb_def) + +lemma isLub_le_isUb: "[| isLub R S x; isUb R S y |] ==> x <= y" +apply (simp add: isLub_def) +apply (blast intro!: leastPD3) +done + +lemma isLub_ubs: "isLub R S x ==> x <=* ubs R S" +apply (simp add: ubs_def isLub_def) +apply (erule leastPD2) +done + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/MacLaurin.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MacLaurin.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,576 @@ +(* Author : Jacques D. Fleuriot + Copyright : 2001 University of Edinburgh + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 +*) + +header{*MacLaurin Series*} + +theory MacLaurin +imports Dense_Linear_Order Transcendental +begin + +subsection{*Maclaurin's Theorem with Lagrange Form of Remainder*} + +text{*This is a very long, messy proof even now that it's been broken down +into lemmas.*} + +lemma Maclaurin_lemma: + "0 < h ==> + \B. f h = (\m=0.. + resolve_tac @{thms deriv_rulesI} i ORELSE + ((rtac (read_instantiate ctxt [(("f", 0), get_fun_name prem)] + @{thm DERIV_chain2}) i) handle DERIV_name => no_tac)); + +fun DERIV_tac ctxt = ALLGOALS (fn i => REPEAT (deriv_tac ctxt i)); + +end +*} + +lemma Maclaurin_lemma2: + "[| \m t. m < n \ 0\t \ t\h \ DERIV (diff m) t :> diff (Suc m) t; + n = Suc k; + difg = + (\m t. diff m t - + ((\p = 0.. + \m t. m < n & 0 \ t & t \ h --> + DERIV (difg m) t :> difg (Suc m) t" +apply clarify +apply (rule DERIV_diff) +apply (simp (no_asm_simp)) +apply (tactic {* DERIV_tac @{context} *}) +apply (tactic {* DERIV_tac @{context} *}) +apply (rule_tac [2] lemma_DERIV_subst) +apply (rule_tac [2] DERIV_quotient) +apply (rule_tac [3] DERIV_const) +apply (rule_tac [2] DERIV_pow) + prefer 3 apply (simp add: fact_diff_Suc) + prefer 2 apply simp +apply (frule_tac m = m in less_add_one, clarify) +apply (simp del: setsum_op_ivl_Suc) +apply (insert sumr_offset4 [of 1]) +apply (simp del: setsum_op_ivl_Suc fact_Suc realpow_Suc) +apply (rule lemma_DERIV_subst) +apply (rule DERIV_add) +apply (rule_tac [2] DERIV_const) +apply (rule DERIV_sumr, clarify) + prefer 2 apply simp +apply (simp (no_asm) add: divide_inverse mult_assoc del: fact_Suc realpow_Suc) +apply (rule DERIV_cmult) +apply (rule lemma_DERIV_subst) +apply (best intro: DERIV_chain2 intro!: DERIV_intros) +apply (subst fact_Suc) +apply (subst real_of_nat_mult) +apply (simp add: mult_ac) +done + + +lemma Maclaurin_lemma3: + fixes difg :: "nat => real => real" shows + "[|\k t. k < Suc m \ 0\t & t\h \ DERIV (difg k) t :> difg (Suc k) t; + \k 0; n < m; 0 < t; + t < h|] + ==> \ta. 0 < ta & ta < t & DERIV (difg (Suc n)) ta :> 0" +apply (rule Rolle, assumption, simp) +apply (drule_tac x = n and P="%k. k difg k 0 = 0" in spec) +apply (rule DERIV_unique) +prefer 2 apply assumption +apply force +apply (metis DERIV_isCont dlo_simps(4) dlo_simps(9) less_trans_Suc nat_less_le not_less_eq real_le_trans) +apply (metis Suc_less_eq differentiableI dlo_simps(7) dlo_simps(8) dlo_simps(9) real_le_trans xt1(8)) +done + +lemma Maclaurin: + "[| 0 < h; n > 0; diff 0 = f; + \m t. m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t |] + ==> \t. 0 < t & + t < h & + f h = + setsum (%m. (diff m 0 / real (fact m)) * h ^ m) {0..g. + g = (%t. f t - (setsum (%m. (diff m 0 / real(fact m)) * t^m) {0..difg. difg = (%m t. diff m t - (setsum (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) {0..ma. ma < n --> (\t. 0 < t & t < h & difg (Suc ma) t = 0) ") + apply (drule_tac x = m and P="%m. m (\t. ?QQ m t)" in spec) + apply (erule impE) + apply (simp (no_asm_simp)) + apply (erule exE) + apply (rule_tac x = t in exI) + apply (simp del: realpow_Suc fact_Suc) +apply (subgoal_tac "\m. m < n --> difg m 0 = 0") + prefer 2 + apply clarify + apply simp + apply (frule_tac m = ma in less_add_one, clarify) + apply (simp del: setsum_op_ivl_Suc) +apply (insert sumr_offset4 [of 1]) +apply (simp del: setsum_op_ivl_Suc fact_Suc realpow_Suc) +apply (subgoal_tac "\m. m < n --> (\t. 0 < t & t < h & DERIV (difg m) t :> 0) ") +apply (rule allI, rule impI) +apply (drule_tac x = ma and P="%m. m (\t. ?QQ m t)" in spec) +apply (erule impE, assumption) +apply (erule exE) +apply (rule_tac x = t in exI) +(* do some tidying up *) +apply (erule_tac [!] V= "difg = (%m t. diff m t - (setsum (%p. diff (m + p) 0 / real (fact p) * t ^ p) {0..0 & diff 0 = f & + (\m t. m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t) + --> (\t. 0 < t & t < h & + f h = (\m=0..m t. + m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t |] + ==> \t. 0 < t & + t \ h & + f h = + (\m=0..m t. + m < n & 0 \ t & t \ h --> DERIV (diff m) t :> diff (Suc m) t) + --> (\t. 0 < t & + t \ h & + f h = + (\m=0.. 0; diff 0 = f; + \m t. m < n & h \ t & t \ 0 --> DERIV (diff m) t :> diff (Suc m) t |] + ==> \t. h < t & + t < 0 & + f h = + (\m=0..m = 0..m = 0.. 0 & diff 0 = f & + (\m t. + m < n & h \ t & t \ 0 --> DERIV (diff m) t :> diff (Suc m) t)) + --> (\t. h < t & + t < 0 & + f h = + (\m=0..0 \ + diff 0 0 = + (\m = 0..m t. m < n & abs t \ abs x --> DERIV (diff m) t :> diff (Suc m) t |] + ==> \t. abs t \ abs x & + f x = + (\m=0..m x. DERIV (diff m) x :> diff(Suc m) x; + x ~= 0; n > 0 + |] ==> \t. 0 < abs t & abs t < abs x & + f x = (\m=0..m x. DERIV (diff m) x :> diff(Suc m) x) & + x ~= 0 & n > 0 + --> (\t. 0 < abs t & abs t < abs x & + f x = (\m=0.. n \ 0 --> + (\m=0..m x. DERIV (diff m) x :> diff (Suc m) x + |] ==> \t. abs t \ abs x & + f x = (\m=0..m x. DERIV (diff m) x :> diff (Suc m) x) + --> (\t. abs t \ abs x & + f x = (\m=0.. 0 |] + ==> (\t. 0 < abs t & + abs t < abs x & + exp x = (\m=0..t. abs t \ abs x & + exp x = (\m=0..x. a \ x & x \ b --> DERIV f x :> f'(x) |] + ==> \z::real. a < z & z < b & (f b - f a = (b - a) * f'(z))" +apply (drule MVT) +apply (blast intro: DERIV_isCont) +apply (force dest: order_less_imp_le simp add: differentiable_def) +apply (blast dest: DERIV_unique order_less_imp_le) +done + +lemma mod_exhaust_less_4: + "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = (3::nat)" +by auto + +lemma Suc_Suc_mult_two_diff_two [rule_format, simp]: + "n\0 --> Suc (Suc (2 * n - 2)) = 2*n" +by (induct "n", auto) + +lemma lemma_Suc_Suc_4n_diff_2 [rule_format, simp]: + "n\0 --> Suc (Suc (4*n - 2)) = 4*n" +by (induct "n", auto) + +lemma Suc_mult_two_diff_one [rule_format, simp]: + "n\0 --> Suc (2 * n - 1) = 2*n" +by (induct "n", auto) + + +text{*It is unclear why so many variant results are needed.*} + +lemma Maclaurin_sin_expansion2: + "\t. abs t \ abs x & + sin x = + (\m=0..t. sin x = + (\m=0.. 0; 0 < x |] ==> + \t. 0 < t & t < x & + sin x = + (\m=0.. + \t. 0 < t & t \ x & + sin x = + (\m=0..m=0..<(Suc n). + (if even m then -1 ^ (m div 2)/(real (fact m)) else 0) * 0 ^ m) = 1" +by (induct "n", auto) + +lemma Maclaurin_cos_expansion: + "\t. abs t \ abs x & + cos x = + (\m=0.. 0 |] ==> + \t. 0 < t & t < x & + cos x = + (\m=0.. 0 |] ==> + \t. x < t & t < 0 & + cos x = + (\m=0.. (v::real) |] ==> \(x + u) - y\ \ v" +by auto + +lemma Maclaurin_sin_bound: + "abs(sin x - (\m=0.. inverse(real (fact n)) * \x\ ^ n" +proof - + have "!! x (y::real). x \ 1 \ 0 \ y \ x * y \ 1 * y" + by (rule_tac mult_right_mono,simp_all) + note est = this[simplified] + let ?diff = "\(n::nat) x. if n mod 4 = 0 then sin(x) else if n mod 4 = 1 then cos(x) else if n mod 4 = 2 then -sin(x) else -cos(x)" + have diff_0: "?diff 0 = sin" by simp + have DERIV_diff: "\m x. DERIV (?diff m) x :> ?diff (Suc m) x" + apply (clarify) + apply (subst (1 2 3) mod_Suc_eq_Suc_mod) + apply (cut_tac m=m in mod_exhaust_less_4) + apply (safe, simp_all) + apply (rule DERIV_minus, simp) + apply (rule lemma_DERIV_subst, rule DERIV_minus, rule DERIV_cos, simp) + done + from Maclaurin_all_le [OF diff_0 DERIV_diff] + obtain t where t1: "\t\ \ \x\" and + t2: "sin x = (\m = 0..m. ?diff m 0 = (if even m then 0 + else -1 ^ ((m - Suc 0) div 2))" + apply (subst even_even_mod_4_iff) + apply (cut_tac m=m in mod_exhaust_less_4) + apply (elim disjE, simp_all) + apply (safe dest!: mod_eqD, simp_all) + done + show ?thesis + apply (subst t2) + apply (rule sin_bound_lemma) + apply (rule setsum_cong[OF refl]) + apply (subst diff_m_0, simp) + apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono + simp add: est mult_nonneg_nonneg mult_ac divide_inverse + power_abs [symmetric] abs_mult) + done +qed + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/NSA/HyperDef.thy --- a/src/HOL/NSA/HyperDef.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/NSA/HyperDef.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ -(* Title : HOL/Hyperreal/HyperDef.thy - ID : $Id$ +(* Title : HOL/NSA/HyperDef.thy Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge Conversion to Isar and new proofs by Lawrence C Paulson, 2004 @@ -8,7 +7,7 @@ header{*Construction of Hyperreals Using Ultrafilters*} theory HyperDef -imports HyperNat "../Real/Real" +imports HyperNat Real uses ("hypreal_arith.ML") begin diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/NSA/Hyperreal.thy --- a/src/HOL/NSA/Hyperreal.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/NSA/Hyperreal.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ -(* Title: HOL/Hyperreal/Hyperreal.thy - ID: $Id$ +(* Title: HOL/NSA/Hyperreal.thy Author: Jacques Fleuriot, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/NSA/NSA.thy --- a/src/HOL/NSA/NSA.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/NSA/NSA.thy Wed Dec 03 15:58:44 2008 +0100 @@ -8,7 +8,7 @@ header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*} theory NSA -imports HyperDef "../Real/RComplete" +imports HyperDef RComplete begin definition diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/NSA/NSCA.thy --- a/src/HOL/NSA/NSCA.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/NSA/NSCA.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,4 +1,4 @@ -(* Title : NSCA.thy +(* Title : NSA/NSCA.thy Author : Jacques D. Fleuriot Copyright : 2001,2002 University of Edinburgh *) @@ -6,7 +6,7 @@ header{*Non-Standard Complex Analysis*} theory NSCA -imports NSComplex "../Hyperreal/HTranscendental" +imports NSComplex HTranscendental begin abbreviation diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Nat.thy --- a/src/HOL/Nat.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Nat.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Nat.thy - ID: $Id$ Author: Tobias Nipkow and Lawrence C Paulson and Markus Wenzel Type "nat" is a linear order, and a datatype; arithmetic operators + - @@ -13,7 +12,7 @@ uses "~~/src/Tools/rat.ML" "~~/src/Provers/Arith/cancel_sums.ML" - ("arith_data.ML") + ("Tools/arith_data.ML") "~~/src/Provers/Arith/fast_lin_arith.ML" ("Tools/lin_arith.ML") begin @@ -1323,7 +1322,7 @@ shows "u = s" using 2 1 by (rule trans) -use "arith_data.ML" +use "Tools/arith_data.ML" declaration {* K ArithData.setup *} use "Tools/lin_arith.ML" diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/NthRoot.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/NthRoot.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,631 @@ +(* Title : NthRoot.thy + Author : Jacques D. Fleuriot + Copyright : 1998 University of Cambridge + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 +*) + +header {* Nth Roots of Real Numbers *} + +theory NthRoot +imports Parity Deriv +begin + +subsection {* Existence of Nth Root *} + +text {* Existence follows from the Intermediate Value Theorem *} + +lemma realpow_pos_nth: + assumes n: "0 < n" + assumes a: "0 < a" + shows "\r>0. r ^ n = (a::real)" +proof - + have "\r\0. r \ (max 1 a) \ r ^ n = a" + proof (rule IVT) + show "0 ^ n \ a" using n a by (simp add: power_0_left) + show "0 \ max 1 a" by simp + from n have n1: "1 \ n" by simp + have "a \ max 1 a ^ 1" by simp + also have "max 1 a ^ 1 \ max 1 a ^ n" + using n1 by (rule power_increasing, simp) + finally show "a \ max 1 a ^ n" . + show "\r. 0 \ r \ r \ max 1 a \ isCont (\x. x ^ n) r" + by (simp add: isCont_power) + qed + then obtain r where r: "0 \ r \ r ^ n = a" by fast + with n a have "r \ 0" by (auto simp add: power_0_left) + with r have "0 < r \ r ^ n = a" by simp + thus ?thesis .. +qed + +(* Used by Integration/RealRandVar.thy in AFP *) +lemma realpow_pos_nth2: "(0::real) < a \ \r>0. r ^ Suc n = a" +by (blast intro: realpow_pos_nth) + +text {* Uniqueness of nth positive root *} + +lemma realpow_pos_nth_unique: + "\0 < n; 0 < a\ \ \!r. 0 < r \ r ^ n = (a::real)" +apply (auto intro!: realpow_pos_nth) +apply (rule_tac n=n in power_eq_imp_eq_base, simp_all) +done + +subsection {* Nth Root *} + +text {* We define roots of negative reals such that + @{term "root n (- x) = - root n x"}. This allows + us to omit side conditions from many theorems. *} + +definition + root :: "[nat, real] \ real" where + "root n x = (if 0 < x then (THE u. 0 < u \ u ^ n = x) else + if x < 0 then - (THE u. 0 < u \ u ^ n = - x) else 0)" + +lemma real_root_zero [simp]: "root n 0 = 0" +unfolding root_def by simp + +lemma real_root_minus: "0 < n \ root n (- x) = - root n x" +unfolding root_def by simp + +lemma real_root_gt_zero: "\0 < n; 0 < x\ \ 0 < root n x" +apply (simp add: root_def) +apply (drule (1) realpow_pos_nth_unique) +apply (erule theI' [THEN conjunct1]) +done + +lemma real_root_pow_pos: (* TODO: rename *) + "\0 < n; 0 < x\ \ root n x ^ n = x" +apply (simp add: root_def) +apply (drule (1) realpow_pos_nth_unique) +apply (erule theI' [THEN conjunct2]) +done + +lemma real_root_pow_pos2 [simp]: (* TODO: rename *) + "\0 < n; 0 \ x\ \ root n x ^ n = x" +by (auto simp add: order_le_less real_root_pow_pos) + +lemma odd_real_root_pow: "odd n \ root n x ^ n = x" +apply (rule_tac x=0 and y=x in linorder_le_cases) +apply (erule (1) real_root_pow_pos2 [OF odd_pos]) +apply (subgoal_tac "root n (- x) ^ n = - x") +apply (simp add: real_root_minus odd_pos) +apply (simp add: odd_pos) +done + +lemma real_root_ge_zero: "\0 < n; 0 \ x\ \ 0 \ root n x" +by (auto simp add: order_le_less real_root_gt_zero) + +lemma real_root_power_cancel: "\0 < n; 0 \ x\ \ root n (x ^ n) = x" +apply (subgoal_tac "0 \ x ^ n") +apply (subgoal_tac "0 \ root n (x ^ n)") +apply (subgoal_tac "root n (x ^ n) ^ n = x ^ n") +apply (erule (3) power_eq_imp_eq_base) +apply (erule (1) real_root_pow_pos2) +apply (erule (1) real_root_ge_zero) +apply (erule zero_le_power) +done + +lemma odd_real_root_power_cancel: "odd n \ root n (x ^ n) = x" +apply (rule_tac x=0 and y=x in linorder_le_cases) +apply (erule (1) real_root_power_cancel [OF odd_pos]) +apply (subgoal_tac "root n ((- x) ^ n) = - x") +apply (simp add: real_root_minus odd_pos) +apply (erule real_root_power_cancel [OF odd_pos], simp) +done + +lemma real_root_pos_unique: + "\0 < n; 0 \ y; y ^ n = x\ \ root n x = y" +by (erule subst, rule real_root_power_cancel) + +lemma odd_real_root_unique: + "\odd n; y ^ n = x\ \ root n x = y" +by (erule subst, rule odd_real_root_power_cancel) + +lemma real_root_one [simp]: "0 < n \ root n 1 = 1" +by (simp add: real_root_pos_unique) + +text {* Root function is strictly monotonic, hence injective *} + +lemma real_root_less_mono_lemma: + "\0 < n; 0 \ x; x < y\ \ root n x < root n y" +apply (subgoal_tac "0 \ y") +apply (subgoal_tac "root n x ^ n < root n y ^ n") +apply (erule power_less_imp_less_base) +apply (erule (1) real_root_ge_zero) +apply simp +apply simp +done + +lemma real_root_less_mono: "\0 < n; x < y\ \ root n x < root n y" +apply (cases "0 \ x") +apply (erule (2) real_root_less_mono_lemma) +apply (cases "0 \ y") +apply (rule_tac y=0 in order_less_le_trans) +apply (subgoal_tac "0 < root n (- x)") +apply (simp add: real_root_minus) +apply (simp add: real_root_gt_zero) +apply (simp add: real_root_ge_zero) +apply (subgoal_tac "root n (- y) < root n (- x)") +apply (simp add: real_root_minus) +apply (simp add: real_root_less_mono_lemma) +done + +lemma real_root_le_mono: "\0 < n; x \ y\ \ root n x \ root n y" +by (auto simp add: order_le_less real_root_less_mono) + +lemma real_root_less_iff [simp]: + "0 < n \ (root n x < root n y) = (x < y)" +apply (cases "x < y") +apply (simp add: real_root_less_mono) +apply (simp add: linorder_not_less real_root_le_mono) +done + +lemma real_root_le_iff [simp]: + "0 < n \ (root n x \ root n y) = (x \ y)" +apply (cases "x \ y") +apply (simp add: real_root_le_mono) +apply (simp add: linorder_not_le real_root_less_mono) +done + +lemma real_root_eq_iff [simp]: + "0 < n \ (root n x = root n y) = (x = y)" +by (simp add: order_eq_iff) + +lemmas real_root_gt_0_iff [simp] = real_root_less_iff [where x=0, simplified] +lemmas real_root_lt_0_iff [simp] = real_root_less_iff [where y=0, simplified] +lemmas real_root_ge_0_iff [simp] = real_root_le_iff [where x=0, simplified] +lemmas real_root_le_0_iff [simp] = real_root_le_iff [where y=0, simplified] +lemmas real_root_eq_0_iff [simp] = real_root_eq_iff [where y=0, simplified] + +lemma real_root_gt_1_iff [simp]: "0 < n \ (1 < root n y) = (1 < y)" +by (insert real_root_less_iff [where x=1], simp) + +lemma real_root_lt_1_iff [simp]: "0 < n \ (root n x < 1) = (x < 1)" +by (insert real_root_less_iff [where y=1], simp) + +lemma real_root_ge_1_iff [simp]: "0 < n \ (1 \ root n y) = (1 \ y)" +by (insert real_root_le_iff [where x=1], simp) + +lemma real_root_le_1_iff [simp]: "0 < n \ (root n x \ 1) = (x \ 1)" +by (insert real_root_le_iff [where y=1], simp) + +lemma real_root_eq_1_iff [simp]: "0 < n \ (root n x = 1) = (x = 1)" +by (insert real_root_eq_iff [where y=1], simp) + +text {* Roots of roots *} + +lemma real_root_Suc_0 [simp]: "root (Suc 0) x = x" +by (simp add: odd_real_root_unique) + +lemma real_root_pos_mult_exp: + "\0 < m; 0 < n; 0 < x\ \ root (m * n) x = root m (root n x)" +by (rule real_root_pos_unique, simp_all add: power_mult) + +lemma real_root_mult_exp: + "\0 < m; 0 < n\ \ root (m * n) x = root m (root n x)" +apply (rule linorder_cases [where x=x and y=0]) +apply (subgoal_tac "root (m * n) (- x) = root m (root n (- x))") +apply (simp add: real_root_minus) +apply (simp_all add: real_root_pos_mult_exp) +done + +lemma real_root_commute: + "\0 < m; 0 < n\ \ root m (root n x) = root n (root m x)" +by (simp add: real_root_mult_exp [symmetric] mult_commute) + +text {* Monotonicity in first argument *} + +lemma real_root_strict_decreasing: + "\0 < n; n < N; 1 < x\ \ root N x < root n x" +apply (subgoal_tac "root n (root N x) ^ n < root N (root n x) ^ N", simp) +apply (simp add: real_root_commute power_strict_increasing + del: real_root_pow_pos2) +done + +lemma real_root_strict_increasing: + "\0 < n; n < N; 0 < x; x < 1\ \ root n x < root N x" +apply (subgoal_tac "root N (root n x) ^ N < root n (root N x) ^ n", simp) +apply (simp add: real_root_commute power_strict_decreasing + del: real_root_pow_pos2) +done + +lemma real_root_decreasing: + "\0 < n; n < N; 1 \ x\ \ root N x \ root n x" +by (auto simp add: order_le_less real_root_strict_decreasing) + +lemma real_root_increasing: + "\0 < n; n < N; 0 \ x; x \ 1\ \ root n x \ root N x" +by (auto simp add: order_le_less real_root_strict_increasing) + +text {* Roots of multiplication and division *} + +lemma real_root_mult_lemma: + "\0 < n; 0 \ x; 0 \ y\ \ root n (x * y) = root n x * root n y" +by (simp add: real_root_pos_unique mult_nonneg_nonneg power_mult_distrib) + +lemma real_root_inverse_lemma: + "\0 < n; 0 \ x\ \ root n (inverse x) = inverse (root n x)" +by (simp add: real_root_pos_unique power_inverse [symmetric]) + +lemma real_root_mult: + assumes n: "0 < n" + shows "root n (x * y) = root n x * root n y" +proof (rule linorder_le_cases, rule_tac [!] linorder_le_cases) + assume "0 \ x" and "0 \ y" + thus ?thesis by (rule real_root_mult_lemma [OF n]) +next + assume "0 \ x" and "y \ 0" + hence "0 \ x" and "0 \ - y" by simp_all + hence "root n (x * - y) = root n x * root n (- y)" + by (rule real_root_mult_lemma [OF n]) + thus ?thesis by (simp add: real_root_minus [OF n]) +next + assume "x \ 0" and "0 \ y" + hence "0 \ - x" and "0 \ y" by simp_all + hence "root n (- x * y) = root n (- x) * root n y" + by (rule real_root_mult_lemma [OF n]) + thus ?thesis by (simp add: real_root_minus [OF n]) +next + assume "x \ 0" and "y \ 0" + hence "0 \ - x" and "0 \ - y" by simp_all + hence "root n (- x * - y) = root n (- x) * root n (- y)" + by (rule real_root_mult_lemma [OF n]) + thus ?thesis by (simp add: real_root_minus [OF n]) +qed + +lemma real_root_inverse: + assumes n: "0 < n" + shows "root n (inverse x) = inverse (root n x)" +proof (rule linorder_le_cases) + assume "0 \ x" + thus ?thesis by (rule real_root_inverse_lemma [OF n]) +next + assume "x \ 0" + hence "0 \ - x" by simp + hence "root n (inverse (- x)) = inverse (root n (- x))" + by (rule real_root_inverse_lemma [OF n]) + thus ?thesis by (simp add: real_root_minus [OF n]) +qed + +lemma real_root_divide: + "0 < n \ root n (x / y) = root n x / root n y" +by (simp add: divide_inverse real_root_mult real_root_inverse) + +lemma real_root_power: + "0 < n \ root n (x ^ k) = root n x ^ k" +by (induct k, simp_all add: real_root_mult) + +lemma real_root_abs: "0 < n \ root n \x\ = \root n x\" +by (simp add: abs_if real_root_minus) + +text {* Continuity and derivatives *} + +lemma isCont_root_pos: + assumes n: "0 < n" + assumes x: "0 < x" + shows "isCont (root n) x" +proof - + have "isCont (root n) (root n x ^ n)" + proof (rule isCont_inverse_function [where f="\a. a ^ n"]) + show "0 < root n x" using n x by simp + show "\z. \z - root n x\ \ root n x \ root n (z ^ n) = z" + by (simp add: abs_le_iff real_root_power_cancel n) + show "\z. \z - root n x\ \ root n x \ isCont (\a. a ^ n) z" + by (simp add: isCont_power) + qed + thus ?thesis using n x by simp +qed + +lemma isCont_root_neg: + "\0 < n; x < 0\ \ isCont (root n) x" +apply (subgoal_tac "isCont (\x. - root n (- x)) x") +apply (simp add: real_root_minus) +apply (rule isCont_o2 [OF isCont_minus [OF isCont_ident]]) +apply (simp add: isCont_minus isCont_root_pos) +done + +lemma isCont_root_zero: + "0 < n \ isCont (root n) 0" +unfolding isCont_def +apply (rule LIM_I) +apply (rule_tac x="r ^ n" in exI, safe) +apply (simp) +apply (simp add: real_root_abs [symmetric]) +apply (rule_tac n="n" in power_less_imp_less_base, simp_all) +done + +lemma isCont_real_root: "0 < n \ isCont (root n) x" +apply (rule_tac x=x and y=0 in linorder_cases) +apply (simp_all add: isCont_root_pos isCont_root_neg isCont_root_zero) +done + +lemma DERIV_real_root: + assumes n: "0 < n" + assumes x: "0 < x" + shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))" +proof (rule DERIV_inverse_function) + show "0 < x" using x . + show "x < x + 1" by simp + show "\y. 0 < y \ y < x + 1 \ root n y ^ n = y" + using n by simp + show "DERIV (\x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)" + by (rule DERIV_pow) + show "real n * root n x ^ (n - Suc 0) \ 0" + using n x by simp + show "isCont (root n) x" + using n by (rule isCont_real_root) +qed + +lemma DERIV_odd_real_root: + assumes n: "odd n" + assumes x: "x \ 0" + shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))" +proof (rule DERIV_inverse_function) + show "x - 1 < x" by simp + show "x < x + 1" by simp + show "\y. x - 1 < y \ y < x + 1 \ root n y ^ n = y" + using n by (simp add: odd_real_root_pow) + show "DERIV (\x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)" + by (rule DERIV_pow) + show "real n * root n x ^ (n - Suc 0) \ 0" + using odd_pos [OF n] x by simp + show "isCont (root n) x" + using odd_pos [OF n] by (rule isCont_real_root) +qed + +subsection {* Square Root *} + +definition + sqrt :: "real \ real" where + "sqrt = root 2" + +lemma pos2: "0 < (2::nat)" by simp + +lemma real_sqrt_unique: "\y\ = x; 0 \ y\ \ sqrt x = y" +unfolding sqrt_def by (rule real_root_pos_unique [OF pos2]) + +lemma real_sqrt_abs [simp]: "sqrt (x\) = \x\" +apply (rule real_sqrt_unique) +apply (rule power2_abs) +apply (rule abs_ge_zero) +done + +lemma real_sqrt_pow2 [simp]: "0 \ x \ (sqrt x)\ = x" +unfolding sqrt_def by (rule real_root_pow_pos2 [OF pos2]) + +lemma real_sqrt_pow2_iff [simp]: "((sqrt x)\ = x) = (0 \ x)" +apply (rule iffI) +apply (erule subst) +apply (rule zero_le_power2) +apply (erule real_sqrt_pow2) +done + +lemma real_sqrt_zero [simp]: "sqrt 0 = 0" +unfolding sqrt_def by (rule real_root_zero) + +lemma real_sqrt_one [simp]: "sqrt 1 = 1" +unfolding sqrt_def by (rule real_root_one [OF pos2]) + +lemma real_sqrt_minus: "sqrt (- x) = - sqrt x" +unfolding sqrt_def by (rule real_root_minus [OF pos2]) + +lemma real_sqrt_mult: "sqrt (x * y) = sqrt x * sqrt y" +unfolding sqrt_def by (rule real_root_mult [OF pos2]) + +lemma real_sqrt_inverse: "sqrt (inverse x) = inverse (sqrt x)" +unfolding sqrt_def by (rule real_root_inverse [OF pos2]) + +lemma real_sqrt_divide: "sqrt (x / y) = sqrt x / sqrt y" +unfolding sqrt_def by (rule real_root_divide [OF pos2]) + +lemma real_sqrt_power: "sqrt (x ^ k) = sqrt x ^ k" +unfolding sqrt_def by (rule real_root_power [OF pos2]) + +lemma real_sqrt_gt_zero: "0 < x \ 0 < sqrt x" +unfolding sqrt_def by (rule real_root_gt_zero [OF pos2]) + +lemma real_sqrt_ge_zero: "0 \ x \ 0 \ sqrt x" +unfolding sqrt_def by (rule real_root_ge_zero [OF pos2]) + +lemma real_sqrt_less_mono: "x < y \ sqrt x < sqrt y" +unfolding sqrt_def by (rule real_root_less_mono [OF pos2]) + +lemma real_sqrt_le_mono: "x \ y \ sqrt x \ sqrt y" +unfolding sqrt_def by (rule real_root_le_mono [OF pos2]) + +lemma real_sqrt_less_iff [simp]: "(sqrt x < sqrt y) = (x < y)" +unfolding sqrt_def by (rule real_root_less_iff [OF pos2]) + +lemma real_sqrt_le_iff [simp]: "(sqrt x \ sqrt y) = (x \ y)" +unfolding sqrt_def by (rule real_root_le_iff [OF pos2]) + +lemma real_sqrt_eq_iff [simp]: "(sqrt x = sqrt y) = (x = y)" +unfolding sqrt_def by (rule real_root_eq_iff [OF pos2]) + +lemmas real_sqrt_gt_0_iff [simp] = real_sqrt_less_iff [where x=0, simplified] +lemmas real_sqrt_lt_0_iff [simp] = real_sqrt_less_iff [where y=0, simplified] +lemmas real_sqrt_ge_0_iff [simp] = real_sqrt_le_iff [where x=0, simplified] +lemmas real_sqrt_le_0_iff [simp] = real_sqrt_le_iff [where y=0, simplified] +lemmas real_sqrt_eq_0_iff [simp] = real_sqrt_eq_iff [where y=0, simplified] + +lemmas real_sqrt_gt_1_iff [simp] = real_sqrt_less_iff [where x=1, simplified] +lemmas real_sqrt_lt_1_iff [simp] = real_sqrt_less_iff [where y=1, simplified] +lemmas real_sqrt_ge_1_iff [simp] = real_sqrt_le_iff [where x=1, simplified] +lemmas real_sqrt_le_1_iff [simp] = real_sqrt_le_iff [where y=1, simplified] +lemmas real_sqrt_eq_1_iff [simp] = real_sqrt_eq_iff [where y=1, simplified] + +lemma isCont_real_sqrt: "isCont sqrt x" +unfolding sqrt_def by (rule isCont_real_root [OF pos2]) + +lemma DERIV_real_sqrt: + "0 < x \ DERIV sqrt x :> inverse (sqrt x) / 2" +unfolding sqrt_def by (rule DERIV_real_root [OF pos2, simplified]) + +lemma not_real_square_gt_zero [simp]: "(~ (0::real) < x*x) = (x = 0)" +apply auto +apply (cut_tac x = x and y = 0 in linorder_less_linear) +apply (simp add: zero_less_mult_iff) +done + +lemma real_sqrt_abs2 [simp]: "sqrt(x*x) = \x\" +apply (subst power2_eq_square [symmetric]) +apply (rule real_sqrt_abs) +done + +lemma real_sqrt_pow2_gt_zero: "0 < x ==> 0 < (sqrt x)\" +by simp (* TODO: delete *) + +lemma real_sqrt_not_eq_zero: "0 < x ==> sqrt x \ 0" +by simp (* TODO: delete *) + +lemma real_inv_sqrt_pow2: "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x" +by (simp add: power_inverse [symmetric]) + +lemma real_sqrt_eq_zero_cancel: "[| 0 \ x; sqrt(x) = 0|] ==> x = 0" +by simp + +lemma real_sqrt_ge_one: "1 \ x ==> 1 \ sqrt x" +by simp + +lemma real_sqrt_two_gt_zero [simp]: "0 < sqrt 2" +by simp + +lemma real_sqrt_two_ge_zero [simp]: "0 \ sqrt 2" +by simp + +lemma real_sqrt_two_gt_one [simp]: "1 < sqrt 2" +by simp + +lemma sqrt_divide_self_eq: + assumes nneg: "0 \ x" + shows "sqrt x / x = inverse (sqrt x)" +proof cases + assume "x=0" thus ?thesis by simp +next + assume nz: "x\0" + hence pos: "0 0" by (simp add: divide_inverse nneg nz) + show "inverse (sqrt x) / (sqrt x / x) = 1" + by (simp add: divide_inverse mult_assoc [symmetric] + power2_eq_square [symmetric] real_inv_sqrt_pow2 pos nz) + qed +qed + +lemma real_divide_square_eq [simp]: "(((r::real) * a) / (r * r)) = a / r" +apply (simp add: divide_inverse) +apply (case_tac "r=0") +apply (auto simp add: mult_ac) +done + +lemma lemma_real_divide_sqrt_less: "0 < u ==> u / sqrt 2 < u" +by (simp add: divide_less_eq mult_compare_simps) + +lemma four_x_squared: + fixes x::real + shows "4 * x\ = (2 * x)\" +by (simp add: power2_eq_square) + +subsection {* Square Root of Sum of Squares *} + +lemma real_sqrt_mult_self_sum_ge_zero [simp]: "0 \ sqrt(x*x + y*y)" +by (rule real_sqrt_ge_zero [OF sum_squares_ge_zero]) + +lemma real_sqrt_sum_squares_ge_zero [simp]: "0 \ sqrt (x\ + y\)" +by simp + +declare real_sqrt_sum_squares_ge_zero [THEN abs_of_nonneg, simp] + +lemma real_sqrt_sum_squares_mult_ge_zero [simp]: + "0 \ sqrt ((x\ + y\)*(xa\ + ya\))" +by (auto intro!: real_sqrt_ge_zero simp add: zero_le_mult_iff) + +lemma real_sqrt_sum_squares_mult_squared_eq [simp]: + "sqrt ((x\ + y\) * (xa\ + ya\)) ^ 2 = (x\ + y\) * (xa\ + ya\)" +by (auto simp add: zero_le_mult_iff) + +lemma real_sqrt_sum_squares_eq_cancel: "sqrt (x\ + y\) = x \ y = 0" +by (drule_tac f = "%x. x\" in arg_cong, simp) + +lemma real_sqrt_sum_squares_eq_cancel2: "sqrt (x\ + y\) = y \ x = 0" +by (drule_tac f = "%x. x\" in arg_cong, simp) + +lemma real_sqrt_sum_squares_ge1 [simp]: "x \ sqrt (x\ + y\)" +by (rule power2_le_imp_le, simp_all) + +lemma real_sqrt_sum_squares_ge2 [simp]: "y \ sqrt (x\ + y\)" +by (rule power2_le_imp_le, simp_all) + +lemma real_sqrt_ge_abs1 [simp]: "\x\ \ sqrt (x\ + y\)" +by (rule power2_le_imp_le, simp_all) + +lemma real_sqrt_ge_abs2 [simp]: "\y\ \ sqrt (x\ + y\)" +by (rule power2_le_imp_le, simp_all) + +lemma le_real_sqrt_sumsq [simp]: "x \ sqrt (x * x + y * y)" +by (simp add: power2_eq_square [symmetric]) + +lemma power2_sum: + fixes x y :: "'a::{number_ring,recpower}" + shows "(x + y)\ = x\ + y\ + 2 * x * y" +by (simp add: ring_distribs power2_eq_square) + +lemma power2_diff: + fixes x y :: "'a::{number_ring,recpower}" + shows "(x - y)\ = x\ + y\ - 2 * x * y" +by (simp add: ring_distribs power2_eq_square) + +lemma real_sqrt_sum_squares_triangle_ineq: + "sqrt ((a + c)\ + (b + d)\) \ sqrt (a\ + b\) + sqrt (c\ + d\)" +apply (rule power2_le_imp_le, simp) +apply (simp add: power2_sum) +apply (simp only: mult_assoc right_distrib [symmetric]) +apply (rule mult_left_mono) +apply (rule power2_le_imp_le) +apply (simp add: power2_sum power_mult_distrib) +apply (simp add: ring_distribs) +apply (subgoal_tac "0 \ b\ * c\ + a\ * d\ - 2 * (a * c) * (b * d)", simp) +apply (rule_tac b="(a * d - b * c)\" in ord_le_eq_trans) +apply (rule zero_le_power2) +apply (simp add: power2_diff power_mult_distrib) +apply (simp add: mult_nonneg_nonneg) +apply simp +apply (simp add: add_increasing) +done + +lemma real_sqrt_sum_squares_less: + "\\x\ < u / sqrt 2; \y\ < u / sqrt 2\ \ sqrt (x\ + y\) < u" +apply (rule power2_less_imp_less, simp) +apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) +apply (drule power_strict_mono [OF _ abs_ge_zero pos2]) +apply (simp add: power_divide) +apply (drule order_le_less_trans [OF abs_ge_zero]) +apply (simp add: zero_less_divide_iff) +done + +text{*Needed for the infinitely close relation over the nonstandard + complex numbers*} +lemma lemma_sqrt_hcomplex_capprox: + "[| 0 < u; x < u/2; y < u/2; 0 \ x; 0 \ y |] ==> sqrt (x\ + y\) < u" +apply (rule_tac y = "u/sqrt 2" in order_le_less_trans) +apply (erule_tac [2] lemma_real_divide_sqrt_less) +apply (rule power2_le_imp_le) +apply (auto simp add: real_0_le_divide_iff power_divide) +apply (rule_tac t = "u\" in real_sum_of_halves [THEN subst]) +apply (rule add_mono) +apply (auto simp add: four_x_squared simp del: realpow_Suc intro: power_mono) +done + +text "Legacy theorem names:" +lemmas real_root_pos2 = real_root_power_cancel +lemmas real_root_pos_pos = real_root_gt_zero [THEN order_less_imp_le] +lemmas real_root_pos_pos_le = real_root_ge_zero +lemmas real_sqrt_mult_distrib = real_sqrt_mult +lemmas real_sqrt_mult_distrib2 = real_sqrt_mult +lemmas real_sqrt_eq_zero_cancel_iff = real_sqrt_eq_0_iff + +(* needed for CauchysMeanTheorem.het_base from AFP *) +lemma real_root_pos: "0 < x \ root (Suc n) (x ^ (Suc n)) = x" +by (rule real_root_power_cancel [OF zero_less_Suc order_less_imp_le]) + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/NumberTheory/EvenOdd.thy --- a/src/HOL/NumberTheory/EvenOdd.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/NumberTheory/EvenOdd.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Quadratic_Reciprocity/EvenOdd.thy - ID: $Id$ Authors: Jeremy Avigad, David Gray, and Adam Kramer *) diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Order_Relation.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Order_Relation.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,131 @@ +(* ID : $Id$ + Author : Tobias Nipkow +*) + +header {* Orders as Relations *} + +theory Order_Relation +imports Plain "~~/src/HOL/Hilbert_Choice" "~~/src/HOL/ATP_Linkup" +begin + +text{* This prelude could be moved to theory Relation: *} + +definition "irrefl r \ \x. (x,x) \ r" + +definition "total_on A r \ \x\A.\y\A. x\y \ (x,y)\r \ (y,x)\r" + +abbreviation "total \ total_on UNIV" + + +lemma total_on_empty[simp]: "total_on {} r" +by(simp add:total_on_def) + +lemma refl_on_converse[simp]: "refl A (r^-1) = refl A r" +by(auto simp add:refl_def) + +lemma total_on_converse[simp]: "total_on A (r^-1) = total_on A r" +by (auto simp: total_on_def) + +lemma irrefl_diff_Id[simp]: "irrefl(r-Id)" +by(simp add:irrefl_def) + +declare [[simp_depth_limit = 2]] +lemma trans_diff_Id: " trans r \ antisym r \ trans (r-Id)" +by(simp add: antisym_def trans_def) blast +declare [[simp_depth_limit = 50]] + +lemma total_on_diff_Id[simp]: "total_on A (r-Id) = total_on A r" +by(simp add: total_on_def) + + +subsection{* Orders on a set *} + +definition "preorder_on A r \ refl A r \ trans r" + +definition "partial_order_on A r \ preorder_on A r \ antisym r" + +definition "linear_order_on A r \ partial_order_on A r \ total_on A r" + +definition "strict_linear_order_on A r \ trans r \ irrefl r \ total_on A r" + +definition "well_order_on A r \ linear_order_on A r \ wf(r - Id)" + +lemmas order_on_defs = + preorder_on_def partial_order_on_def linear_order_on_def + strict_linear_order_on_def well_order_on_def + + +lemma preorder_on_empty[simp]: "preorder_on {} {}" +by(simp add:preorder_on_def trans_def) + +lemma partial_order_on_empty[simp]: "partial_order_on {} {}" +by(simp add:partial_order_on_def) + +lemma lnear_order_on_empty[simp]: "linear_order_on {} {}" +by(simp add:linear_order_on_def) + +lemma well_order_on_empty[simp]: "well_order_on {} {}" +by(simp add:well_order_on_def) + + +lemma preorder_on_converse[simp]: "preorder_on A (r^-1) = preorder_on A r" +by (simp add:preorder_on_def) + +lemma partial_order_on_converse[simp]: + "partial_order_on A (r^-1) = partial_order_on A r" +by (simp add: partial_order_on_def) + +lemma linear_order_on_converse[simp]: + "linear_order_on A (r^-1) = linear_order_on A r" +by (simp add: linear_order_on_def) + + +lemma strict_linear_order_on_diff_Id: + "linear_order_on A r \ strict_linear_order_on A (r-Id)" +by(simp add: order_on_defs trans_diff_Id) + + +subsection{* Orders on the field *} + +abbreviation "Refl r \ refl (Field r) r" + +abbreviation "Preorder r \ preorder_on (Field r) r" + +abbreviation "Partial_order r \ partial_order_on (Field r) r" + +abbreviation "Total r \ total_on (Field r) r" + +abbreviation "Linear_order r \ linear_order_on (Field r) r" + +abbreviation "Well_order r \ well_order_on (Field r) r" + + +lemma subset_Image_Image_iff: + "\ Preorder r; A \ Field r; B \ Field r\ \ + r `` A \ r `` B \ (\a\A.\b\B. (b,a):r)" +apply(auto simp add: subset_eq preorder_on_def refl_def Image_def) +apply metis +by(metis trans_def) + +lemma subset_Image1_Image1_iff: + "\ Preorder r; a : Field r; b : Field r\ \ r `` {a} \ r `` {b} \ (b,a):r" +by(simp add:subset_Image_Image_iff) + +lemma Refl_antisym_eq_Image1_Image1_iff: + "\Refl r; antisym r; a:Field r; b:Field r\ \ r `` {a} = r `` {b} \ a=b" +by(simp add: expand_set_eq antisym_def refl_def) metis + +lemma Partial_order_eq_Image1_Image1_iff: + "\Partial_order r; a:Field r; b:Field r\ \ r `` {a} = r `` {b} \ a=b" +by(auto simp:order_on_defs Refl_antisym_eq_Image1_Image1_iff) + + +subsection{* Orders on a type *} + +abbreviation "strict_linear_order \ strict_linear_order_on UNIV" + +abbreviation "linear_order \ linear_order_on UNIV" + +abbreviation "well_order r \ well_order_on UNIV" + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/PReal.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/PReal.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,1343 @@ +(* Title : PReal.thy + Author : Jacques D. Fleuriot + Copyright : 1998 University of Cambridge + Description : The positive reals as Dedekind sections of positive + rationals. Fundamentals of Abstract Analysis [Gleason- p. 121] + provides some of the definitions. +*) + +header {* Positive real numbers *} + +theory PReal +imports Rational "~~/src/HOL/Library/Dense_Linear_Order" +begin + +text{*Could be generalized and moved to @{text Ring_and_Field}*} +lemma add_eq_exists: "\x. a+x = (b::rat)" +by (rule_tac x="b-a" in exI, simp) + +definition + cut :: "rat set => bool" where + [code del]: "cut A = ({} \ A & + A < {r. 0 < r} & + (\y \ A. ((\z. 0 z \ A) & (\u \ A. y < u))))" + +lemma cut_of_rat: + assumes q: "0 < q" shows "cut {r::rat. 0 < r & r < q}" (is "cut ?A") +proof - + from q have pos: "?A < {r. 0 < r}" by force + have nonempty: "{} \ ?A" + proof + show "{} \ ?A" by simp + show "{} \ ?A" + by (force simp only: q eq_commute [of "{}"] interval_empty_iff) + qed + show ?thesis + by (simp add: cut_def pos nonempty, + blast dest: dense intro: order_less_trans) +qed + + +typedef preal = "{A. cut A}" + by (blast intro: cut_of_rat [OF zero_less_one]) + +definition + preal_of_rat :: "rat => preal" where + "preal_of_rat q = Abs_preal {x::rat. 0 < x & x < q}" + +definition + psup :: "preal set => preal" where + "psup P = Abs_preal (\X \ P. Rep_preal X)" + +definition + add_set :: "[rat set,rat set] => rat set" where + "add_set A B = {w. \x \ A. \y \ B. w = x + y}" + +definition + diff_set :: "[rat set,rat set] => rat set" where + [code del]: "diff_set A B = {w. \x. 0 < w & 0 < x & x \ B & x + w \ A}" + +definition + mult_set :: "[rat set,rat set] => rat set" where + "mult_set A B = {w. \x \ A. \y \ B. w = x * y}" + +definition + inverse_set :: "rat set => rat set" where + [code del]: "inverse_set A = {x. \y. 0 < x & x < y & inverse y \ A}" + +instantiation preal :: "{ord, plus, minus, times, inverse, one}" +begin + +definition + preal_less_def [code del]: + "R < S == Rep_preal R < Rep_preal S" + +definition + preal_le_def [code del]: + "R \ S == Rep_preal R \ Rep_preal S" + +definition + preal_add_def: + "R + S == Abs_preal (add_set (Rep_preal R) (Rep_preal S))" + +definition + preal_diff_def: + "R - S == Abs_preal (diff_set (Rep_preal R) (Rep_preal S))" + +definition + preal_mult_def: + "R * S == Abs_preal (mult_set (Rep_preal R) (Rep_preal S))" + +definition + preal_inverse_def: + "inverse R == Abs_preal (inverse_set (Rep_preal R))" + +definition "R / S = R * inverse (S\preal)" + +definition + preal_one_def: + "1 == preal_of_rat 1" + +instance .. + +end + + +text{*Reduces equality on abstractions to equality on representatives*} +declare Abs_preal_inject [simp] +declare Abs_preal_inverse [simp] + +lemma rat_mem_preal: "0 < q ==> {r::rat. 0 < r & r < q} \ preal" +by (simp add: preal_def cut_of_rat) + +lemma preal_nonempty: "A \ preal ==> \x\A. 0 < x" +by (unfold preal_def cut_def, blast) + +lemma preal_Ex_mem: "A \ preal \ \x. x \ A" +by (drule preal_nonempty, fast) + +lemma preal_imp_psubset_positives: "A \ preal ==> A < {r. 0 < r}" +by (force simp add: preal_def cut_def) + +lemma preal_exists_bound: "A \ preal ==> \x. 0 < x & x \ A" +by (drule preal_imp_psubset_positives, auto) + +lemma preal_exists_greater: "[| A \ preal; y \ A |] ==> \u \ A. y < u" +by (unfold preal_def cut_def, blast) + +lemma preal_downwards_closed: "[| A \ preal; y \ A; 0 < z; z < y |] ==> z \ A" +by (unfold preal_def cut_def, blast) + +text{*Relaxing the final premise*} +lemma preal_downwards_closed': + "[| A \ preal; y \ A; 0 < z; z \ y |] ==> z \ A" +apply (simp add: order_le_less) +apply (blast intro: preal_downwards_closed) +done + +text{*A positive fraction not in a positive real is an upper bound. + Gleason p. 122 - Remark (1)*} + +lemma not_in_preal_ub: + assumes A: "A \ preal" + and notx: "x \ A" + and y: "y \ A" + and pos: "0 < x" + shows "y < x" +proof (cases rule: linorder_cases) + assume "xx. x \ Rep_preal X" +by (rule preal_Ex_mem [OF Rep_preal]) + +lemma Rep_preal_exists_bound: "\x>0. x \ Rep_preal X" +by (rule preal_exists_bound [OF Rep_preal]) + +lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF Rep_preal] + + + +subsection{*@{term preal_of_prat}: the Injection from prat to preal*} + +lemma rat_less_set_mem_preal: "0 < y ==> {u::rat. 0 < u & u < y} \ preal" +by (simp add: preal_def cut_of_rat) + +lemma rat_subset_imp_le: + "[|{u::rat. 0 < u & u < x} \ {u. 0 < u & u < y}; 0 x \ y" +apply (simp add: linorder_not_less [symmetric]) +apply (blast dest: dense intro: order_less_trans) +done + +lemma rat_set_eq_imp_eq: + "[|{u::rat. 0 < u & u < x} = {u. 0 < u & u < y}; + 0 < x; 0 < y|] ==> x = y" +by (blast intro: rat_subset_imp_le order_antisym) + + + +subsection{*Properties of Ordering*} + +instance preal :: order +proof + fix w :: preal + show "w \ w" by (simp add: preal_le_def) +next + fix i j k :: preal + assume "i \ j" and "j \ k" + then show "i \ k" by (simp add: preal_le_def) +next + fix z w :: preal + assume "z \ w" and "w \ z" + then show "z = w" by (simp add: preal_le_def Rep_preal_inject) +next + fix z w :: preal + show "z < w \ z \ w \ \ w \ z" + by (auto simp add: preal_le_def preal_less_def Rep_preal_inject) +qed + +lemma preal_imp_pos: "[|A \ preal; r \ A|] ==> 0 < r" +by (insert preal_imp_psubset_positives, blast) + +instance preal :: linorder +proof + fix x y :: preal + show "x <= y | y <= x" + apply (auto simp add: preal_le_def) + apply (rule ccontr) + apply (blast dest: not_in_Rep_preal_ub intro: preal_imp_pos [OF Rep_preal] + elim: order_less_asym) + done +qed + +instantiation preal :: distrib_lattice +begin + +definition + "(inf \ preal \ preal \ preal) = min" + +definition + "(sup \ preal \ preal \ preal) = max" + +instance + by intro_classes + (auto simp add: inf_preal_def sup_preal_def min_max.sup_inf_distrib1) + +end + +subsection{*Properties of Addition*} + +lemma preal_add_commute: "(x::preal) + y = y + x" +apply (unfold preal_add_def add_set_def) +apply (rule_tac f = Abs_preal in arg_cong) +apply (force simp add: add_commute) +done + +text{*Lemmas for proving that addition of two positive reals gives + a positive real*} + +lemma empty_psubset_nonempty: "a \ A ==> {} \ A" +by blast + +text{*Part 1 of Dedekind sections definition*} +lemma add_set_not_empty: + "[|A \ preal; B \ preal|] ==> {} \ add_set A B" +apply (drule preal_nonempty)+ +apply (auto simp add: add_set_def) +done + +text{*Part 2 of Dedekind sections definition. A structured version of +this proof is @{text preal_not_mem_mult_set_Ex} below.*} +lemma preal_not_mem_add_set_Ex: + "[|A \ preal; B \ preal|] ==> \q>0. q \ add_set A B" +apply (insert preal_exists_bound [of A] preal_exists_bound [of B], auto) +apply (rule_tac x = "x+xa" in exI) +apply (simp add: add_set_def, clarify) +apply (drule (3) not_in_preal_ub)+ +apply (force dest: add_strict_mono) +done + +lemma add_set_not_rat_set: + assumes A: "A \ preal" + and B: "B \ preal" + shows "add_set A B < {r. 0 < r}" +proof + from preal_imp_pos [OF A] preal_imp_pos [OF B] + show "add_set A B \ {r. 0 < r}" by (force simp add: add_set_def) +next + show "add_set A B \ {r. 0 < r}" + by (insert preal_not_mem_add_set_Ex [OF A B], blast) +qed + +text{*Part 3 of Dedekind sections definition*} +lemma add_set_lemma3: + "[|A \ preal; B \ preal; u \ add_set A B; 0 < z; z < u|] + ==> z \ add_set A B" +proof (unfold add_set_def, clarify) + fix x::rat and y::rat + assume A: "A \ preal" + and B: "B \ preal" + and [simp]: "0 < z" + and zless: "z < x + y" + and x: "x \ A" + and y: "y \ B" + have xpos [simp]: "0x' \ A. \y'\B. z = x' + y'" + proof (intro bexI) + show "z = x*?f + y*?f" + by (simp add: left_distrib [symmetric] divide_inverse mult_ac + order_less_imp_not_eq2) + next + show "y * ?f \ B" + proof (rule preal_downwards_closed [OF B y]) + show "0 < y * ?f" + by (simp add: divide_inverse zero_less_mult_iff) + next + show "y * ?f < y" + by (insert mult_strict_left_mono [OF fless ypos], simp) + qed + next + show "x * ?f \ A" + proof (rule preal_downwards_closed [OF A x]) + show "0 < x * ?f" + by (simp add: divide_inverse zero_less_mult_iff) + next + show "x * ?f < x" + by (insert mult_strict_left_mono [OF fless xpos], simp) + qed + qed +qed + +text{*Part 4 of Dedekind sections definition*} +lemma add_set_lemma4: + "[|A \ preal; B \ preal; y \ add_set A B|] ==> \u \ add_set A B. y < u" +apply (auto simp add: add_set_def) +apply (frule preal_exists_greater [of A], auto) +apply (rule_tac x="u + y" in exI) +apply (auto intro: add_strict_left_mono) +done + +lemma mem_add_set: + "[|A \ preal; B \ preal|] ==> add_set A B \ preal" +apply (simp (no_asm_simp) add: preal_def cut_def) +apply (blast intro!: add_set_not_empty add_set_not_rat_set + add_set_lemma3 add_set_lemma4) +done + +lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)" +apply (simp add: preal_add_def mem_add_set Rep_preal) +apply (force simp add: add_set_def add_ac) +done + +instance preal :: ab_semigroup_add +proof + fix a b c :: preal + show "(a + b) + c = a + (b + c)" by (rule preal_add_assoc) + show "a + b = b + a" by (rule preal_add_commute) +qed + +lemma preal_add_left_commute: "x + (y + z) = y + ((x + z)::preal)" +by (rule add_left_commute) + +text{* Positive Real addition is an AC operator *} +lemmas preal_add_ac = preal_add_assoc preal_add_commute preal_add_left_commute + + +subsection{*Properties of Multiplication*} + +text{*Proofs essentially same as for addition*} + +lemma preal_mult_commute: "(x::preal) * y = y * x" +apply (unfold preal_mult_def mult_set_def) +apply (rule_tac f = Abs_preal in arg_cong) +apply (force simp add: mult_commute) +done + +text{*Multiplication of two positive reals gives a positive real.*} + +text{*Lemmas for proving positive reals multiplication set in @{typ preal}*} + +text{*Part 1 of Dedekind sections definition*} +lemma mult_set_not_empty: + "[|A \ preal; B \ preal|] ==> {} \ mult_set A B" +apply (insert preal_nonempty [of A] preal_nonempty [of B]) +apply (auto simp add: mult_set_def) +done + +text{*Part 2 of Dedekind sections definition*} +lemma preal_not_mem_mult_set_Ex: + assumes A: "A \ preal" + and B: "B \ preal" + shows "\q. 0 < q & q \ mult_set A B" +proof - + from preal_exists_bound [OF A] + obtain x where [simp]: "0 < x" "x \ A" by blast + from preal_exists_bound [OF B] + obtain y where [simp]: "0 < y" "y \ B" by blast + show ?thesis + proof (intro exI conjI) + show "0 < x*y" by (simp add: mult_pos_pos) + show "x * y \ mult_set A B" + proof - + { fix u::rat and v::rat + assume "u \ A" and "v \ B" and "x*y = u*v" + moreover + with prems have "uv" + by (blast intro: preal_imp_pos [OF B] order_less_imp_le prems) + moreover + from calculation + have "u*v < x*y" by (blast intro: mult_strict_mono prems) + ultimately have False by force } + thus ?thesis by (auto simp add: mult_set_def) + qed + qed +qed + +lemma mult_set_not_rat_set: + assumes A: "A \ preal" + and B: "B \ preal" + shows "mult_set A B < {r. 0 < r}" +proof + show "mult_set A B \ {r. 0 < r}" + by (force simp add: mult_set_def + intro: preal_imp_pos [OF A] preal_imp_pos [OF B] mult_pos_pos) + show "mult_set A B \ {r. 0 < r}" + using preal_not_mem_mult_set_Ex [OF A B] by blast +qed + + + +text{*Part 3 of Dedekind sections definition*} +lemma mult_set_lemma3: + "[|A \ preal; B \ preal; u \ mult_set A B; 0 < z; z < u|] + ==> z \ mult_set A B" +proof (unfold mult_set_def, clarify) + fix x::rat and y::rat + assume A: "A \ preal" + and B: "B \ preal" + and [simp]: "0 < z" + and zless: "z < x * y" + and x: "x \ A" + and y: "y \ B" + have [simp]: "0x' \ A. \y' \ B. z = x' * y'" + proof + show "\y'\B. z = (z/y) * y'" + proof + show "z = (z/y)*y" + by (simp add: divide_inverse mult_commute [of y] mult_assoc + order_less_imp_not_eq2) + show "y \ B" by fact + qed + next + show "z/y \ A" + proof (rule preal_downwards_closed [OF A x]) + show "0 < z/y" + by (simp add: zero_less_divide_iff) + show "z/y < x" by (simp add: pos_divide_less_eq zless) + qed + qed +qed + +text{*Part 4 of Dedekind sections definition*} +lemma mult_set_lemma4: + "[|A \ preal; B \ preal; y \ mult_set A B|] ==> \u \ mult_set A B. y < u" +apply (auto simp add: mult_set_def) +apply (frule preal_exists_greater [of A], auto) +apply (rule_tac x="u * y" in exI) +apply (auto intro: preal_imp_pos [of A] preal_imp_pos [of B] + mult_strict_right_mono) +done + + +lemma mem_mult_set: + "[|A \ preal; B \ preal|] ==> mult_set A B \ preal" +apply (simp (no_asm_simp) add: preal_def cut_def) +apply (blast intro!: mult_set_not_empty mult_set_not_rat_set + mult_set_lemma3 mult_set_lemma4) +done + +lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)" +apply (simp add: preal_mult_def mem_mult_set Rep_preal) +apply (force simp add: mult_set_def mult_ac) +done + +instance preal :: ab_semigroup_mult +proof + fix a b c :: preal + show "(a * b) * c = a * (b * c)" by (rule preal_mult_assoc) + show "a * b = b * a" by (rule preal_mult_commute) +qed + +lemma preal_mult_left_commute: "x * (y * z) = y * ((x * z)::preal)" +by (rule mult_left_commute) + + +text{* Positive Real multiplication is an AC operator *} +lemmas preal_mult_ac = + preal_mult_assoc preal_mult_commute preal_mult_left_commute + + +text{* Positive real 1 is the multiplicative identity element *} + +lemma preal_mult_1: "(1::preal) * z = z" +unfolding preal_one_def +proof (induct z) + fix A :: "rat set" + assume A: "A \ preal" + have "{w. \u. 0 < u \ u < 1 & (\v \ A. w = u * v)} = A" (is "?lhs = A") + proof + show "?lhs \ A" + proof clarify + fix x::rat and u::rat and v::rat + assume upos: "0 A" + have vpos: "0 A" + by (force intro: preal_downwards_closed [OF A v] mult_pos_pos + upos vpos) + qed + next + show "A \ ?lhs" + proof clarify + fix x::rat + assume x: "x \ A" + have xpos: "0 A" and xlessv: "x < v" .. + have vpos: "0u. 0 < u \ u < 1 \ (\v\A. x = u * v)" + proof (intro exI conjI) + show "0 < x/v" + by (simp add: zero_less_divide_iff xpos vpos) + show "x / v < 1" + by (simp add: pos_divide_less_eq vpos xlessv) + show "\v'\A. x = (x / v) * v'" + proof + show "x = (x/v)*v" + by (simp add: divide_inverse mult_assoc vpos + order_less_imp_not_eq2) + show "v \ A" by fact + qed + qed + qed + qed + thus "preal_of_rat 1 * Abs_preal A = Abs_preal A" + by (simp add: preal_of_rat_def preal_mult_def mult_set_def + rat_mem_preal A) +qed + +instance preal :: comm_monoid_mult +by intro_classes (rule preal_mult_1) + +lemma preal_mult_1_right: "z * (1::preal) = z" +by (rule mult_1_right) + + +subsection{*Distribution of Multiplication across Addition*} + +lemma mem_Rep_preal_add_iff: + "(z \ Rep_preal(R+S)) = (\x \ Rep_preal R. \y \ Rep_preal S. z = x + y)" +apply (simp add: preal_add_def mem_add_set Rep_preal) +apply (simp add: add_set_def) +done + +lemma mem_Rep_preal_mult_iff: + "(z \ Rep_preal(R*S)) = (\x \ Rep_preal R. \y \ Rep_preal S. z = x * y)" +apply (simp add: preal_mult_def mem_mult_set Rep_preal) +apply (simp add: mult_set_def) +done + +lemma distrib_subset1: + "Rep_preal (w * (x + y)) \ Rep_preal (w * x + w * y)" +apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff) +apply (force simp add: right_distrib) +done + +lemma preal_add_mult_distrib_mean: + assumes a: "a \ Rep_preal w" + and b: "b \ Rep_preal w" + and d: "d \ Rep_preal x" + and e: "e \ Rep_preal y" + shows "\c \ Rep_preal w. a * d + b * e = c * (d + e)" +proof + let ?c = "(a*d + b*e)/(d+e)" + have [simp]: "0 Rep_preal w" + proof (cases rule: linorder_le_cases) + assume "a \ b" + hence "?c \ b" + by (simp add: pos_divide_le_eq right_distrib mult_right_mono + order_less_imp_le) + thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal b cpos]) + next + assume "b \ a" + hence "?c \ a" + by (simp add: pos_divide_le_eq right_distrib mult_right_mono + order_less_imp_le) + thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal a cpos]) + qed +qed + +lemma distrib_subset2: + "Rep_preal (w * x + w * y) \ Rep_preal (w * (x + y))" +apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff) +apply (drule_tac w=w and x=x and y=y in preal_add_mult_distrib_mean, auto) +done + +lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)" +apply (rule Rep_preal_inject [THEN iffD1]) +apply (rule equalityI [OF distrib_subset1 distrib_subset2]) +done + +lemma preal_add_mult_distrib: "(((x::preal) + y) * w) = (x * w) + (y * w)" +by (simp add: preal_mult_commute preal_add_mult_distrib2) + +instance preal :: comm_semiring +by intro_classes (rule preal_add_mult_distrib) + + +subsection{*Existence of Inverse, a Positive Real*} + +lemma mem_inv_set_ex: + assumes A: "A \ preal" shows "\x y. 0 < x & x < y & inverse y \ A" +proof - + from preal_exists_bound [OF A] + obtain x where [simp]: "0 A" by blast + show ?thesis + proof (intro exI conjI) + show "0 < inverse (x+1)" + by (simp add: order_less_trans [OF _ less_add_one]) + show "inverse(x+1) < inverse x" + by (simp add: less_imp_inverse_less less_add_one) + show "inverse (inverse x) \ A" + by (simp add: order_less_imp_not_eq2) + qed +qed + +text{*Part 1 of Dedekind sections definition*} +lemma inverse_set_not_empty: + "A \ preal ==> {} \ inverse_set A" +apply (insert mem_inv_set_ex [of A]) +apply (auto simp add: inverse_set_def) +done + +text{*Part 2 of Dedekind sections definition*} + +lemma preal_not_mem_inverse_set_Ex: + assumes A: "A \ preal" shows "\q. 0 < q & q \ inverse_set A" +proof - + from preal_nonempty [OF A] + obtain x where x: "x \ A" and xpos [simp]: "0 inverse_set A" + proof - + { fix y::rat + assume ygt: "inverse x < y" + have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt]) + have iyless: "inverse y < x" + by (simp add: inverse_less_imp_less [of x] ygt) + have "inverse y \ A" + by (simp add: preal_downwards_closed [OF A x] iyless)} + thus ?thesis by (auto simp add: inverse_set_def) + qed + qed +qed + +lemma inverse_set_not_rat_set: + assumes A: "A \ preal" shows "inverse_set A < {r. 0 < r}" +proof + show "inverse_set A \ {r. 0 < r}" by (force simp add: inverse_set_def) +next + show "inverse_set A \ {r. 0 < r}" + by (insert preal_not_mem_inverse_set_Ex [OF A], blast) +qed + +text{*Part 3 of Dedekind sections definition*} +lemma inverse_set_lemma3: + "[|A \ preal; u \ inverse_set A; 0 < z; z < u|] + ==> z \ inverse_set A" +apply (auto simp add: inverse_set_def) +apply (auto intro: order_less_trans) +done + +text{*Part 4 of Dedekind sections definition*} +lemma inverse_set_lemma4: + "[|A \ preal; y \ inverse_set A|] ==> \u \ inverse_set A. y < u" +apply (auto simp add: inverse_set_def) +apply (drule dense [of y]) +apply (blast intro: order_less_trans) +done + + +lemma mem_inverse_set: + "A \ preal ==> inverse_set A \ preal" +apply (simp (no_asm_simp) add: preal_def cut_def) +apply (blast intro!: inverse_set_not_empty inverse_set_not_rat_set + inverse_set_lemma3 inverse_set_lemma4) +done + + +subsection{*Gleason's Lemma 9-3.4, page 122*} + +lemma Gleason9_34_exists: + assumes A: "A \ preal" + and "\x\A. x + u \ A" + and "0 \ z" + shows "\b\A. b + (of_int z) * u \ A" +proof (cases z rule: int_cases) + case (nonneg n) + show ?thesis + proof (simp add: prems, induct n) + case 0 + from preal_nonempty [OF A] + show ?case by force + case (Suc k) + from this obtain b where "b \ A" "b + of_nat k * u \ A" .. + hence "b + of_int (int k)*u + u \ A" by (simp add: prems) + thus ?case by (force simp add: left_distrib add_ac prems) + qed +next + case (neg n) + with prems show ?thesis by simp +qed + +lemma Gleason9_34_contra: + assumes A: "A \ preal" + shows "[|\x\A. x + u \ A; 0 < u; 0 < y; y \ A|] ==> False" +proof (induct u, induct y) + fix a::int and b::int + fix c::int and d::int + assume bpos [simp]: "0 < b" + and dpos [simp]: "0 < d" + and closed: "\x\A. x + (Fract c d) \ A" + and upos: "0 < Fract c d" + and ypos: "0 < Fract a b" + and notin: "Fract a b \ A" + have cpos [simp]: "0 < c" + by (simp add: zero_less_Fract_iff [OF dpos, symmetric] upos) + have apos [simp]: "0 < a" + by (simp add: zero_less_Fract_iff [OF bpos, symmetric] ypos) + let ?k = "a*d" + have frle: "Fract a b \ Fract ?k 1 * (Fract c d)" + proof - + have "?thesis = ((a * d * b * d) \ c * b * (a * d * b * d))" + by (simp add: mult_rat le_rat order_less_imp_not_eq2 mult_ac) + moreover + have "(1 * (a * d * b * d)) \ c * b * (a * d * b * d)" + by (rule mult_mono, + simp_all add: int_one_le_iff_zero_less zero_less_mult_iff + order_less_imp_le) + ultimately + show ?thesis by simp + qed + have k: "0 \ ?k" by (simp add: order_less_imp_le zero_less_mult_iff) + from Gleason9_34_exists [OF A closed k] + obtain z where z: "z \ A" + and mem: "z + of_int ?k * Fract c d \ A" .. + have less: "z + of_int ?k * Fract c d < Fract a b" + by (rule not_in_preal_ub [OF A notin mem ypos]) + have "0 preal" + and upos: "0 < u" + shows "\r \ A. r + u \ A" +proof (rule ccontr, simp) + assume closed: "\r\A. r + u \ A" + from preal_exists_bound [OF A] + obtain y where y: "y \ A" and ypos: "0 < y" by blast + show False + by (rule Gleason9_34_contra [OF A closed upos ypos y]) +qed + + + +subsection{*Gleason's Lemma 9-3.6*} + +lemma lemma_gleason9_36: + assumes A: "A \ preal" + and x: "1 < x" + shows "\r \ A. r*x \ A" +proof - + from preal_nonempty [OF A] + obtain y where y: "y \ A" and ypos: "0r\A. r * x \ A)" + with y have ymem: "y * x \ A" by blast + from ypos mult_strict_left_mono [OF x] + have yless: "y < y*x" by simp + let ?d = "y*x - y" + from yless have dpos: "0 < ?d" and eq: "y + ?d = y*x" by auto + from Gleason9_34 [OF A dpos] + obtain r where r: "r\A" and notin: "r + ?d \ A" .. + have rpos: "0 y + ?d)" + proof + assume le: "r + ?d \ y + ?d" + from ymem have yd: "y + ?d \ A" by (simp add: eq) + have "r + ?d \ A" by (rule preal_downwards_closed' [OF A yd rdpos le]) + with notin show False by simp + qed + hence "y < r" by simp + with ypos have dless: "?d < (r * ?d)/y" + by (simp add: pos_less_divide_eq mult_commute [of ?d] + mult_strict_right_mono dpos) + have "r + ?d < r*x" + proof - + have "r + ?d < r + (r * ?d)/y" by (simp add: dless) + also with ypos have "... = (r/y) * (y + ?d)" + by (simp only: right_distrib divide_inverse mult_ac, simp) + also have "... = r*x" using ypos + by (simp add: times_divide_eq_left) + finally show "r + ?d < r*x" . + qed + with r notin rdpos + show "\r\A. r * x \ A" by (blast dest: preal_downwards_closed [OF A]) + qed +qed + +subsection{*Existence of Inverse: Part 2*} + +lemma mem_Rep_preal_inverse_iff: + "(z \ Rep_preal(inverse R)) = + (0 < z \ (\y. z < y \ inverse y \ Rep_preal R))" +apply (simp add: preal_inverse_def mem_inverse_set Rep_preal) +apply (simp add: inverse_set_def) +done + +lemma Rep_preal_of_rat: + "0 < q ==> Rep_preal (preal_of_rat q) = {x. 0 < x \ x < q}" +by (simp add: preal_of_rat_def rat_mem_preal) + +lemma subset_inverse_mult_lemma: + assumes xpos: "0 < x" and xless: "x < 1" + shows "\r u y. 0 < r & r < y & inverse y \ Rep_preal R & + u \ Rep_preal R & x = r * u" +proof - + from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff) + from lemma_gleason9_36 [OF Rep_preal this] + obtain r where r: "r \ Rep_preal R" + and notin: "r * (inverse x) \ Rep_preal R" .. + have rpos: "0 Rep_preal R" and rless: "r < u" .. + have upos: "0 Rep_preal R" using notin + by (simp add: divide_inverse mult_commute) + show "u \ Rep_preal R" by (rule u) + show "x = x / u * u" using upos + by (simp add: divide_inverse mult_commute) + qed +qed + +lemma subset_inverse_mult: + "Rep_preal(preal_of_rat 1) \ Rep_preal(inverse R * R)" +apply (auto simp add: Bex_def Rep_preal_of_rat mem_Rep_preal_inverse_iff + mem_Rep_preal_mult_iff) +apply (blast dest: subset_inverse_mult_lemma) +done + +lemma inverse_mult_subset_lemma: + assumes rpos: "0 < r" + and rless: "r < y" + and notin: "inverse y \ Rep_preal R" + and q: "q \ Rep_preal R" + shows "r*q < 1" +proof - + have "q < inverse y" using rpos rless + by (simp add: not_in_preal_ub [OF Rep_preal notin] q) + hence "r * q < r/y" using rpos + by (simp add: divide_inverse mult_less_cancel_left) + also have "... \ 1" using rpos rless + by (simp add: pos_divide_le_eq) + finally show ?thesis . +qed + +lemma inverse_mult_subset: + "Rep_preal(inverse R * R) \ Rep_preal(preal_of_rat 1)" +apply (auto simp add: Bex_def Rep_preal_of_rat mem_Rep_preal_inverse_iff + mem_Rep_preal_mult_iff) +apply (simp add: zero_less_mult_iff preal_imp_pos [OF Rep_preal]) +apply (blast intro: inverse_mult_subset_lemma) +done + +lemma preal_mult_inverse: "inverse R * R = (1::preal)" +unfolding preal_one_def +apply (rule Rep_preal_inject [THEN iffD1]) +apply (rule equalityI [OF inverse_mult_subset subset_inverse_mult]) +done + +lemma preal_mult_inverse_right: "R * inverse R = (1::preal)" +apply (rule preal_mult_commute [THEN subst]) +apply (rule preal_mult_inverse) +done + + +text{*Theorems needing @{text Gleason9_34}*} + +lemma Rep_preal_self_subset: "Rep_preal (R) \ Rep_preal(R + S)" +proof + fix r + assume r: "r \ Rep_preal R" + have rpos: "0 Rep_preal S" .. + have ypos: "0 Rep_preal(R + S)" using r y + by (auto simp add: mem_Rep_preal_add_iff) + show "r \ Rep_preal(R + S)" using r ypos rpos + by (simp add: preal_downwards_closed [OF Rep_preal ry]) +qed + +lemma Rep_preal_sum_not_subset: "~ Rep_preal (R + S) \ Rep_preal(R)" +proof - + from mem_Rep_preal_Ex + obtain y where y: "y \ Rep_preal S" .. + have ypos: "0 Rep_preal R" and notin: "r + y \ Rep_preal R" .. + have "r + y \ Rep_preal (R + S)" using r y + by (auto simp add: mem_Rep_preal_add_iff) + thus ?thesis using notin by blast +qed + +lemma Rep_preal_sum_not_eq: "Rep_preal (R + S) \ Rep_preal(R)" +by (insert Rep_preal_sum_not_subset, blast) + +text{*at last, Gleason prop. 9-3.5(iii) page 123*} +lemma preal_self_less_add_left: "(R::preal) < R + S" +apply (unfold preal_less_def less_le) +apply (simp add: Rep_preal_self_subset Rep_preal_sum_not_eq [THEN not_sym]) +done + +lemma preal_self_less_add_right: "(R::preal) < S + R" +by (simp add: preal_add_commute preal_self_less_add_left) + +lemma preal_not_eq_self: "x \ x + (y::preal)" +by (insert preal_self_less_add_left [of x y], auto) + + +subsection{*Subtraction for Positive Reals*} + +text{*Gleason prop. 9-3.5(iv), page 123: proving @{prop "A < B ==> \D. A + D = +B"}. We define the claimed @{term D} and show that it is a positive real*} + +text{*Part 1 of Dedekind sections definition*} +lemma diff_set_not_empty: + "R < S ==> {} \ diff_set (Rep_preal S) (Rep_preal R)" +apply (auto simp add: preal_less_def diff_set_def elim!: equalityE) +apply (frule_tac x1 = S in Rep_preal [THEN preal_exists_greater]) +apply (drule preal_imp_pos [OF Rep_preal], clarify) +apply (cut_tac a=x and b=u in add_eq_exists, force) +done + +text{*Part 2 of Dedekind sections definition*} +lemma diff_set_nonempty: + "\q. 0 < q & q \ diff_set (Rep_preal S) (Rep_preal R)" +apply (cut_tac X = S in Rep_preal_exists_bound) +apply (erule exE) +apply (rule_tac x = x in exI, auto) +apply (simp add: diff_set_def) +apply (auto dest: Rep_preal [THEN preal_downwards_closed]) +done + +lemma diff_set_not_rat_set: + "diff_set (Rep_preal S) (Rep_preal R) < {r. 0 < r}" (is "?lhs < ?rhs") +proof + show "?lhs \ ?rhs" by (auto simp add: diff_set_def) + show "?lhs \ ?rhs" using diff_set_nonempty by blast +qed + +text{*Part 3 of Dedekind sections definition*} +lemma diff_set_lemma3: + "[|R < S; u \ diff_set (Rep_preal S) (Rep_preal R); 0 < z; z < u|] + ==> z \ diff_set (Rep_preal S) (Rep_preal R)" +apply (auto simp add: diff_set_def) +apply (rule_tac x=x in exI) +apply (drule Rep_preal [THEN preal_downwards_closed], auto) +done + +text{*Part 4 of Dedekind sections definition*} +lemma diff_set_lemma4: + "[|R < S; y \ diff_set (Rep_preal S) (Rep_preal R)|] + ==> \u \ diff_set (Rep_preal S) (Rep_preal R). y < u" +apply (auto simp add: diff_set_def) +apply (drule Rep_preal [THEN preal_exists_greater], clarify) +apply (cut_tac a="x+y" and b=u in add_eq_exists, clarify) +apply (rule_tac x="y+xa" in exI) +apply (auto simp add: add_ac) +done + +lemma mem_diff_set: + "R < S ==> diff_set (Rep_preal S) (Rep_preal R) \ preal" +apply (unfold preal_def cut_def) +apply (blast intro!: diff_set_not_empty diff_set_not_rat_set + diff_set_lemma3 diff_set_lemma4) +done + +lemma mem_Rep_preal_diff_iff: + "R < S ==> + (z \ Rep_preal(S-R)) = + (\x. 0 < x & 0 < z & x \ Rep_preal R & x + z \ Rep_preal S)" +apply (simp add: preal_diff_def mem_diff_set Rep_preal) +apply (force simp add: diff_set_def) +done + + +text{*proving that @{term "R + D \ S"}*} + +lemma less_add_left_lemma: + assumes Rless: "R < S" + and a: "a \ Rep_preal R" + and cb: "c + b \ Rep_preal S" + and "c \ Rep_preal R" + and "0 < b" + and "0 < c" + shows "a + b \ Rep_preal S" +proof - + have "0 R + (S-R) \ S" +apply (auto simp add: Bex_def preal_le_def mem_Rep_preal_add_iff + mem_Rep_preal_diff_iff) +apply (blast intro: less_add_left_lemma) +done + +subsection{*proving that @{term "S \ R + D"} --- trickier*} + +lemma lemma_sum_mem_Rep_preal_ex: + "x \ Rep_preal S ==> \e. 0 < e & x + e \ Rep_preal S" +apply (drule Rep_preal [THEN preal_exists_greater], clarify) +apply (cut_tac a=x and b=u in add_eq_exists, auto) +done + +lemma less_add_left_lemma2: + assumes Rless: "R < S" + and x: "x \ Rep_preal S" + and xnot: "x \ Rep_preal R" + shows "\u v z. 0 < v & 0 < z & u \ Rep_preal R & z \ Rep_preal R & + z + v \ Rep_preal S & x = u + v" +proof - + have xpos: "0 Rep_preal S" by blast + from Gleason9_34 [OF Rep_preal epos] + obtain r where r: "r \ Rep_preal R" and notin: "r + e \ Rep_preal R" .. + with x xnot xpos have rless: "r < x" by (blast intro: not_in_Rep_preal_ub) + from add_eq_exists [of r x] + obtain y where eq: "x = r+y" by auto + show ?thesis + proof (intro exI conjI) + show "r \ Rep_preal R" by (rule r) + show "r + e \ Rep_preal R" by (rule notin) + show "r + e + y \ Rep_preal S" using xe eq by (simp add: add_ac) + show "x = r + y" by (simp add: eq) + show "0 < r + e" using epos preal_imp_pos [OF Rep_preal r] + by simp + show "0 < y" using rless eq by arith + qed +qed + +lemma less_add_left_le2: "R < (S::preal) ==> S \ R + (S-R)" +apply (auto simp add: preal_le_def) +apply (case_tac "x \ Rep_preal R") +apply (cut_tac Rep_preal_self_subset [of R], force) +apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_diff_iff) +apply (blast dest: less_add_left_lemma2) +done + +lemma less_add_left: "R < (S::preal) ==> R + (S-R) = S" +by (blast intro: antisym [OF less_add_left_le1 less_add_left_le2]) + +lemma less_add_left_Ex: "R < (S::preal) ==> \D. R + D = S" +by (fast dest: less_add_left) + +lemma preal_add_less2_mono1: "R < (S::preal) ==> R + T < S + T" +apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc) +apply (rule_tac y1 = D in preal_add_commute [THEN subst]) +apply (auto intro: preal_self_less_add_left simp add: preal_add_assoc [symmetric]) +done + +lemma preal_add_less2_mono2: "R < (S::preal) ==> T + R < T + S" +by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of T]) + +lemma preal_add_right_less_cancel: "R + T < S + T ==> R < (S::preal)" +apply (insert linorder_less_linear [of R S], auto) +apply (drule_tac R = S and T = T in preal_add_less2_mono1) +apply (blast dest: order_less_trans) +done + +lemma preal_add_left_less_cancel: "T + R < T + S ==> R < (S::preal)" +by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of T]) + +lemma preal_add_less_cancel_right: "((R::preal) + T < S + T) = (R < S)" +by (blast intro: preal_add_less2_mono1 preal_add_right_less_cancel) + +lemma preal_add_less_cancel_left: "(T + (R::preal) < T + S) = (R < S)" +by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel) + +lemma preal_add_le_cancel_right: "((R::preal) + T \ S + T) = (R \ S)" +by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_right) + +lemma preal_add_le_cancel_left: "(T + (R::preal) \ T + S) = (R \ S)" +by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_left) + +lemma preal_add_less_mono: + "[| x1 < y1; x2 < y2 |] ==> x1 + x2 < y1 + (y2::preal)" +apply (auto dest!: less_add_left_Ex simp add: preal_add_ac) +apply (rule preal_add_assoc [THEN subst]) +apply (rule preal_self_less_add_right) +done + +lemma preal_add_right_cancel: "(R::preal) + T = S + T ==> R = S" +apply (insert linorder_less_linear [of R S], safe) +apply (drule_tac [!] T = T in preal_add_less2_mono1, auto) +done + +lemma preal_add_left_cancel: "C + A = C + B ==> A = (B::preal)" +by (auto intro: preal_add_right_cancel simp add: preal_add_commute) + +lemma preal_add_left_cancel_iff: "(C + A = C + B) = ((A::preal) = B)" +by (fast intro: preal_add_left_cancel) + +lemma preal_add_right_cancel_iff: "(A + C = B + C) = ((A::preal) = B)" +by (fast intro: preal_add_right_cancel) + +lemmas preal_cancels = + preal_add_less_cancel_right preal_add_less_cancel_left + preal_add_le_cancel_right preal_add_le_cancel_left + preal_add_left_cancel_iff preal_add_right_cancel_iff + +instance preal :: ordered_cancel_ab_semigroup_add +proof + fix a b c :: preal + show "a + b = a + c \ b = c" by (rule preal_add_left_cancel) + show "a \ b \ c + a \ c + b" by (simp only: preal_add_le_cancel_left) +qed + + +subsection{*Completeness of type @{typ preal}*} + +text{*Prove that supremum is a cut*} + +text{*Part 1 of Dedekind sections definition*} + +lemma preal_sup_set_not_empty: + "P \ {} ==> {} \ (\X \ P. Rep_preal(X))" +apply auto +apply (cut_tac X = x in mem_Rep_preal_Ex, auto) +done + + +text{*Part 2 of Dedekind sections definition*} + +lemma preal_sup_not_exists: + "\X \ P. X \ Y ==> \q. 0 < q & q \ (\X \ P. Rep_preal(X))" +apply (cut_tac X = Y in Rep_preal_exists_bound) +apply (auto simp add: preal_le_def) +done + +lemma preal_sup_set_not_rat_set: + "\X \ P. X \ Y ==> (\X \ P. Rep_preal(X)) < {r. 0 < r}" +apply (drule preal_sup_not_exists) +apply (blast intro: preal_imp_pos [OF Rep_preal]) +done + +text{*Part 3 of Dedekind sections definition*} +lemma preal_sup_set_lemma3: + "[|P \ {}; \X \ P. X \ Y; u \ (\X \ P. Rep_preal(X)); 0 < z; z < u|] + ==> z \ (\X \ P. Rep_preal(X))" +by (auto elim: Rep_preal [THEN preal_downwards_closed]) + +text{*Part 4 of Dedekind sections definition*} +lemma preal_sup_set_lemma4: + "[|P \ {}; \X \ P. X \ Y; y \ (\X \ P. Rep_preal(X)) |] + ==> \u \ (\X \ P. Rep_preal(X)). y < u" +by (blast dest: Rep_preal [THEN preal_exists_greater]) + +lemma preal_sup: + "[|P \ {}; \X \ P. X \ Y|] ==> (\X \ P. Rep_preal(X)) \ preal" +apply (unfold preal_def cut_def) +apply (blast intro!: preal_sup_set_not_empty preal_sup_set_not_rat_set + preal_sup_set_lemma3 preal_sup_set_lemma4) +done + +lemma preal_psup_le: + "[| \X \ P. X \ Y; x \ P |] ==> x \ psup P" +apply (simp (no_asm_simp) add: preal_le_def) +apply (subgoal_tac "P \ {}") +apply (auto simp add: psup_def preal_sup) +done + +lemma psup_le_ub: "[| P \ {}; \X \ P. X \ Y |] ==> psup P \ Y" +apply (simp (no_asm_simp) add: preal_le_def) +apply (simp add: psup_def preal_sup) +apply (auto simp add: preal_le_def) +done + +text{*Supremum property*} +lemma preal_complete: + "[| P \ {}; \X \ P. X \ Y |] ==> (\X \ P. Z < X) = (Z < psup P)" +apply (simp add: preal_less_def psup_def preal_sup) +apply (auto simp add: preal_le_def) +apply (rename_tac U) +apply (cut_tac x = U and y = Z in linorder_less_linear) +apply (auto simp add: preal_less_def) +done + + +subsection{*The Embedding from @{typ rat} into @{typ preal}*} + +lemma preal_of_rat_add_lemma1: + "[|x < y + z; 0 < x; 0 < y|] ==> x * y * inverse (y + z) < (y::rat)" +apply (frule_tac c = "y * inverse (y + z) " in mult_strict_right_mono) +apply (simp add: zero_less_mult_iff) +apply (simp add: mult_ac) +done + +lemma preal_of_rat_add_lemma2: + assumes "u < x + y" + and "0 < x" + and "0 < y" + and "0 < u" + shows "\v w::rat. w < y & 0 < v & v < x & 0 < w & u = v + w" +proof (intro exI conjI) + show "u * x * inverse(x+y) < x" using prems + by (simp add: preal_of_rat_add_lemma1) + show "u * y * inverse(x+y) < y" using prems + by (simp add: preal_of_rat_add_lemma1 add_commute [of x]) + show "0 < u * x * inverse (x + y)" using prems + by (simp add: zero_less_mult_iff) + show "0 < u * y * inverse (x + y)" using prems + by (simp add: zero_less_mult_iff) + show "u = u * x * inverse (x + y) + u * y * inverse (x + y)" using prems + by (simp add: left_distrib [symmetric] right_distrib [symmetric] mult_ac) +qed + +lemma preal_of_rat_add: + "[| 0 < x; 0 < y|] + ==> preal_of_rat ((x::rat) + y) = preal_of_rat x + preal_of_rat y" +apply (unfold preal_of_rat_def preal_add_def) +apply (simp add: rat_mem_preal) +apply (rule_tac f = Abs_preal in arg_cong) +apply (auto simp add: add_set_def) +apply (blast dest: preal_of_rat_add_lemma2) +done + +lemma preal_of_rat_mult_lemma1: + "[|x < y; 0 < x; 0 < z|] ==> x * z * inverse y < (z::rat)" +apply (frule_tac c = "z * inverse y" in mult_strict_right_mono) +apply (simp add: zero_less_mult_iff) +apply (subgoal_tac "y * (z * inverse y) = z * (y * inverse y)") +apply (simp_all add: mult_ac) +done + +lemma preal_of_rat_mult_lemma2: + assumes xless: "x < y * z" + and xpos: "0 < x" + and ypos: "0 < y" + shows "x * z * inverse y * inverse z < (z::rat)" +proof - + have "0 < y * z" using prems by simp + hence zpos: "0 < z" using prems by (simp add: zero_less_mult_iff) + have "x * z * inverse y * inverse z = x * inverse y * (z * inverse z)" + by (simp add: mult_ac) + also have "... = x/y" using zpos + by (simp add: divide_inverse) + also from xless have "... < z" + by (simp add: pos_divide_less_eq [OF ypos] mult_commute) + finally show ?thesis . +qed + +lemma preal_of_rat_mult_lemma3: + assumes uless: "u < x * y" + and "0 < x" + and "0 < y" + and "0 < u" + shows "\v w::rat. v < x & w < y & 0 < v & 0 < w & u = v * w" +proof - + from dense [OF uless] + obtain r where "u < r" "r < x * y" by blast + thus ?thesis + proof (intro exI conjI) + show "u * x * inverse r < x" using prems + by (simp add: preal_of_rat_mult_lemma1) + show "r * y * inverse x * inverse y < y" using prems + by (simp add: preal_of_rat_mult_lemma2) + show "0 < u * x * inverse r" using prems + by (simp add: zero_less_mult_iff) + show "0 < r * y * inverse x * inverse y" using prems + by (simp add: zero_less_mult_iff) + have "u * x * inverse r * (r * y * inverse x * inverse y) = + u * (r * inverse r) * (x * inverse x) * (y * inverse y)" + by (simp only: mult_ac) + thus "u = u * x * inverse r * (r * y * inverse x * inverse y)" using prems + by simp + qed +qed + +lemma preal_of_rat_mult: + "[| 0 < x; 0 < y|] + ==> preal_of_rat ((x::rat) * y) = preal_of_rat x * preal_of_rat y" +apply (unfold preal_of_rat_def preal_mult_def) +apply (simp add: rat_mem_preal) +apply (rule_tac f = Abs_preal in arg_cong) +apply (auto simp add: zero_less_mult_iff mult_strict_mono mult_set_def) +apply (blast dest: preal_of_rat_mult_lemma3) +done + +lemma preal_of_rat_less_iff: + "[| 0 < x; 0 < y|] ==> (preal_of_rat x < preal_of_rat y) = (x < y)" +by (force simp add: preal_of_rat_def preal_less_def rat_mem_preal) + +lemma preal_of_rat_le_iff: + "[| 0 < x; 0 < y|] ==> (preal_of_rat x \ preal_of_rat y) = (x \ y)" +by (simp add: preal_of_rat_less_iff linorder_not_less [symmetric]) + +lemma preal_of_rat_eq_iff: + "[| 0 < x; 0 < y|] ==> (preal_of_rat x = preal_of_rat y) = (x = y)" +by (simp add: preal_of_rat_le_iff order_eq_iff) + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Parity.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Parity.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,423 @@ +(* Title: HOL/Library/Parity.thy + Author: Jeremy Avigad, Jacques D. Fleuriot +*) + +header {* Even and Odd for int and nat *} + +theory Parity +imports Plain Presburger +begin + +class even_odd = type + + fixes even :: "'a \ bool" + +abbreviation + odd :: "'a\even_odd \ bool" where + "odd x \ \ even x" + +instantiation nat and int :: even_odd +begin + +definition + even_def [presburger]: "even x \ (x\int) mod 2 = 0" + +definition + even_nat_def [presburger]: "even x \ even (int x)" + +instance .. + +end + + +subsection {* Even and odd are mutually exclusive *} + +lemma int_pos_lt_two_imp_zero_or_one: + "0 <= x ==> (x::int) < 2 ==> x = 0 | x = 1" + by presburger + +lemma neq_one_mod_two [simp, presburger]: + "((x::int) mod 2 ~= 0) = (x mod 2 = 1)" by presburger + + +subsection {* Behavior under integer arithmetic operations *} +declare dvd_def[algebra] +lemma nat_even_iff_2_dvd[algebra]: "even (x::nat) \ 2 dvd x" + by (presburger add: even_nat_def even_def) +lemma int_even_iff_2_dvd[algebra]: "even (x::int) \ 2 dvd x" + by presburger + +lemma even_times_anything: "even (x::int) ==> even (x * y)" + by algebra + +lemma anything_times_even: "even (y::int) ==> even (x * y)" by algebra + +lemma odd_times_odd: "odd (x::int) ==> odd y ==> odd (x * y)" + by (simp add: even_def zmod_zmult1_eq) + +lemma even_product[presburger]: "even((x::int) * y) = (even x | even y)" + apply (auto simp add: even_times_anything anything_times_even) + apply (rule ccontr) + apply (auto simp add: odd_times_odd) + done + +lemma even_plus_even: "even (x::int) ==> even y ==> even (x + y)" + by presburger + +lemma even_plus_odd: "even (x::int) ==> odd y ==> odd (x + y)" + by presburger + +lemma odd_plus_even: "odd (x::int) ==> even y ==> odd (x + y)" + by presburger + +lemma odd_plus_odd: "odd (x::int) ==> odd y ==> even (x + y)" by presburger + +lemma even_sum[presburger]: "even ((x::int) + y) = ((even x & even y) | (odd x & odd y))" + by presburger + +lemma even_neg[presburger, algebra]: "even (-(x::int)) = even x" by presburger + +lemma even_difference: + "even ((x::int) - y) = ((even x & even y) | (odd x & odd y))" by presburger + +lemma even_pow_gt_zero: + "even (x::int) ==> 0 < n ==> even (x^n)" + by (induct n) (auto simp add: even_product) + +lemma odd_pow_iff[presburger, algebra]: + "odd ((x::int) ^ n) \ (n = 0 \ odd x)" + apply (induct n, simp_all) + apply presburger + apply (case_tac n, auto) + apply (simp_all add: even_product) + done + +lemma odd_pow: "odd x ==> odd((x::int)^n)" by (simp add: odd_pow_iff) + +lemma even_power[presburger]: "even ((x::int)^n) = (even x & 0 < n)" + apply (auto simp add: even_pow_gt_zero) + apply (erule contrapos_pp, erule odd_pow) + apply (erule contrapos_pp, simp add: even_def) + done + +lemma even_zero[presburger]: "even (0::int)" by presburger + +lemma odd_one[presburger]: "odd (1::int)" by presburger + +lemmas even_odd_simps [simp] = even_def[of "number_of v",standard] even_zero + odd_one even_product even_sum even_neg even_difference even_power + + +subsection {* Equivalent definitions *} + +lemma two_times_even_div_two: "even (x::int) ==> 2 * (x div 2) = x" + by presburger + +lemma two_times_odd_div_two_plus_one: "odd (x::int) ==> + 2 * (x div 2) + 1 = x" by presburger + +lemma even_equiv_def: "even (x::int) = (EX y. x = 2 * y)" by presburger + +lemma odd_equiv_def: "odd (x::int) = (EX y. x = 2 * y + 1)" by presburger + +subsection {* even and odd for nats *} + +lemma pos_int_even_equiv_nat_even: "0 \ x ==> even x = even (nat x)" + by (simp add: even_nat_def) + +lemma even_nat_product[presburger, algebra]: "even((x::nat) * y) = (even x | even y)" + by (simp add: even_nat_def int_mult) + +lemma even_nat_sum[presburger, algebra]: "even ((x::nat) + y) = + ((even x & even y) | (odd x & odd y))" by presburger + +lemma even_nat_difference[presburger, algebra]: + "even ((x::nat) - y) = (x < y | (even x & even y) | (odd x & odd y))" +by presburger + +lemma even_nat_Suc[presburger, algebra]: "even (Suc x) = odd x" by presburger + +lemma even_nat_power[presburger, algebra]: "even ((x::nat)^y) = (even x & 0 < y)" + by (simp add: even_nat_def int_power) + +lemma even_nat_zero[presburger]: "even (0::nat)" by presburger + +lemmas even_odd_nat_simps [simp] = even_nat_def[of "number_of v",standard] + even_nat_zero even_nat_Suc even_nat_product even_nat_sum even_nat_power + + +subsection {* Equivalent definitions *} + +lemma nat_lt_two_imp_zero_or_one: "(x::nat) < Suc (Suc 0) ==> + x = 0 | x = Suc 0" by presburger + +lemma even_nat_mod_two_eq_zero: "even (x::nat) ==> x mod (Suc (Suc 0)) = 0" + by presburger + +lemma odd_nat_mod_two_eq_one: "odd (x::nat) ==> x mod (Suc (Suc 0)) = Suc 0" +by presburger + +lemma even_nat_equiv_def: "even (x::nat) = (x mod Suc (Suc 0) = 0)" + by presburger + +lemma odd_nat_equiv_def: "odd (x::nat) = (x mod Suc (Suc 0) = Suc 0)" + by presburger + +lemma even_nat_div_two_times_two: "even (x::nat) ==> + Suc (Suc 0) * (x div Suc (Suc 0)) = x" by presburger + +lemma odd_nat_div_two_times_two_plus_one: "odd (x::nat) ==> + Suc( Suc (Suc 0) * (x div Suc (Suc 0))) = x" by presburger + +lemma even_nat_equiv_def2: "even (x::nat) = (EX y. x = Suc (Suc 0) * y)" + by presburger + +lemma odd_nat_equiv_def2: "odd (x::nat) = (EX y. x = Suc(Suc (Suc 0) * y))" + by presburger + + +subsection {* Parity and powers *} + +lemma minus_one_even_odd_power: + "(even x --> (- 1::'a::{comm_ring_1,recpower})^x = 1) & + (odd x --> (- 1::'a)^x = - 1)" + apply (induct x) + apply (rule conjI) + apply simp + apply (insert even_nat_zero, blast) + apply (simp add: power_Suc) + done + +lemma minus_one_even_power [simp]: + "even x ==> (- 1::'a::{comm_ring_1,recpower})^x = 1" + using minus_one_even_odd_power by blast + +lemma minus_one_odd_power [simp]: + "odd x ==> (- 1::'a::{comm_ring_1,recpower})^x = - 1" + using minus_one_even_odd_power by blast + +lemma neg_one_even_odd_power: + "(even x --> (-1::'a::{number_ring,recpower})^x = 1) & + (odd x --> (-1::'a)^x = -1)" + apply (induct x) + apply (simp, simp add: power_Suc) + done + +lemma neg_one_even_power [simp]: + "even x ==> (-1::'a::{number_ring,recpower})^x = 1" + using neg_one_even_odd_power by blast + +lemma neg_one_odd_power [simp]: + "odd x ==> (-1::'a::{number_ring,recpower})^x = -1" + using neg_one_even_odd_power by blast + +lemma neg_power_if: + "(-x::'a::{comm_ring_1,recpower}) ^ n = + (if even n then (x ^ n) else -(x ^ n))" + apply (induct n) + apply (simp_all split: split_if_asm add: power_Suc) + done + +lemma zero_le_even_power: "even n ==> + 0 <= (x::'a::{recpower,ordered_ring_strict}) ^ n" + apply (simp add: even_nat_equiv_def2) + apply (erule exE) + apply (erule ssubst) + apply (subst power_add) + apply (rule zero_le_square) + done + +lemma zero_le_odd_power: "odd n ==> + (0 <= (x::'a::{recpower,ordered_idom}) ^ n) = (0 <= x)" + apply (simp add: odd_nat_equiv_def2) + apply (erule exE) + apply (erule ssubst) + apply (subst power_Suc) + apply (subst power_add) + apply (subst zero_le_mult_iff) + apply auto + apply (subgoal_tac "x = 0 & y > 0") + apply (erule conjE, assumption) + apply (subst power_eq_0_iff [symmetric]) + apply (subgoal_tac "0 <= x^y * x^y") + apply simp + apply (rule zero_le_square)+ + done + +lemma zero_le_power_eq[presburger]: "(0 <= (x::'a::{recpower,ordered_idom}) ^ n) = + (even n | (odd n & 0 <= x))" + apply auto + apply (subst zero_le_odd_power [symmetric]) + apply assumption+ + apply (erule zero_le_even_power) + done + +lemma zero_less_power_eq[presburger]: "(0 < (x::'a::{recpower,ordered_idom}) ^ n) = + (n = 0 | (even n & x ~= 0) | (odd n & 0 < x))" + + unfolding order_less_le zero_le_power_eq by auto + +lemma power_less_zero_eq[presburger]: "((x::'a::{recpower,ordered_idom}) ^ n < 0) = + (odd n & x < 0)" + apply (subst linorder_not_le [symmetric])+ + apply (subst zero_le_power_eq) + apply auto + done + +lemma power_le_zero_eq[presburger]: "((x::'a::{recpower,ordered_idom}) ^ n <= 0) = + (n ~= 0 & ((odd n & x <= 0) | (even n & x = 0)))" + apply (subst linorder_not_less [symmetric])+ + apply (subst zero_less_power_eq) + apply auto + done + +lemma power_even_abs: "even n ==> + (abs (x::'a::{recpower,ordered_idom}))^n = x^n" + apply (subst power_abs [symmetric]) + apply (simp add: zero_le_even_power) + done + +lemma zero_less_power_nat_eq[presburger]: "(0 < (x::nat) ^ n) = (n = 0 | 0 < x)" + by (induct n) auto + +lemma power_minus_even [simp]: "even n ==> + (- x)^n = (x^n::'a::{recpower,comm_ring_1})" + apply (subst power_minus) + apply simp + done + +lemma power_minus_odd [simp]: "odd n ==> + (- x)^n = - (x^n::'a::{recpower,comm_ring_1})" + apply (subst power_minus) + apply simp + done + + +subsection {* General Lemmas About Division *} + +lemma Suc_times_mod_eq: "1 Suc (k * m) mod k = 1" +apply (induct "m") +apply (simp_all add: mod_Suc) +done + +declare Suc_times_mod_eq [of "number_of w", standard, simp] + +lemma [simp]: "n div k \ (Suc n) div k" +by (simp add: div_le_mono) + +lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2" +by arith + +lemma div_2_gt_zero [simp]: "(1::nat) < n ==> 0 < n div 2" +by arith + + (* Potential use of algebra : Equality modulo n*) +lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)" +by (simp add: mult_ac add_ac) + +lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n" +proof - + have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp + also have "... = Suc m mod n" by (rule mod_mult_self3) + finally show ?thesis . +qed + +lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n" +apply (subst mod_Suc [of m]) +apply (subst mod_Suc [of "m mod n"], simp) +done + + +subsection {* More Even/Odd Results *} + +lemma even_mult_two_ex: "even(n) = (\m::nat. n = 2*m)" by presburger +lemma odd_Suc_mult_two_ex: "odd(n) = (\m. n = Suc (2*m))" by presburger +lemma even_add [simp]: "even(m + n::nat) = (even m = even n)" by presburger + +lemma odd_add [simp]: "odd(m + n::nat) = (odd m \ odd n)" by presburger + +lemma div_Suc: "Suc a div c = a div c + Suc 0 div c + + (a mod c + Suc 0 mod c) div c" + apply (subgoal_tac "Suc a = a + Suc 0") + apply (erule ssubst) + apply (rule div_add1_eq, simp) + done + +lemma lemma_even_div2 [simp]: "even (n::nat) ==> (n + 1) div 2 = n div 2" by presburger + +lemma lemma_not_even_div2 [simp]: "~even n ==> (n + 1) div 2 = Suc (n div 2)" +by presburger + +lemma even_num_iff: "0 < n ==> even n = (~ even(n - 1 :: nat))" by presburger +lemma even_even_mod_4_iff: "even (n::nat) = even (n mod 4)" by presburger + +lemma lemma_odd_mod_4_div_2: "n mod 4 = (3::nat) ==> odd((n - 1) div 2)" by presburger + +lemma lemma_even_mod_4_div_2: "n mod 4 = (1::nat) ==> even ((n - 1) div 2)" + by presburger + +text {* Simplify, when the exponent is a numeral *} + +lemmas power_0_left_number_of = power_0_left [of "number_of w", standard] +declare power_0_left_number_of [simp] + +lemmas zero_le_power_eq_number_of [simp] = + zero_le_power_eq [of _ "number_of w", standard] + +lemmas zero_less_power_eq_number_of [simp] = + zero_less_power_eq [of _ "number_of w", standard] + +lemmas power_le_zero_eq_number_of [simp] = + power_le_zero_eq [of _ "number_of w", standard] + +lemmas power_less_zero_eq_number_of [simp] = + power_less_zero_eq [of _ "number_of w", standard] + +lemmas zero_less_power_nat_eq_number_of [simp] = + zero_less_power_nat_eq [of _ "number_of w", standard] + +lemmas power_eq_0_iff_number_of [simp] = power_eq_0_iff [of _ "number_of w", standard] + +lemmas power_even_abs_number_of [simp] = power_even_abs [of "number_of w" _, standard] + + +subsection {* An Equivalence for @{term [source] "0 \ a^n"} *} + +lemma even_power_le_0_imp_0: + "a ^ (2*k) \ (0::'a::{ordered_idom,recpower}) ==> a=0" + by (induct k) (auto simp add: zero_le_mult_iff mult_le_0_iff power_Suc) + +lemma zero_le_power_iff[presburger]: + "(0 \ a^n) = (0 \ (a::'a::{ordered_idom,recpower}) | even n)" +proof cases + assume even: "even n" + then obtain k where "n = 2*k" + by (auto simp add: even_nat_equiv_def2 numeral_2_eq_2) + thus ?thesis by (simp add: zero_le_even_power even) +next + assume odd: "odd n" + then obtain k where "n = Suc(2*k)" + by (auto simp add: odd_nat_equiv_def2 numeral_2_eq_2) + thus ?thesis + by (auto simp add: power_Suc zero_le_mult_iff zero_le_even_power + dest!: even_power_le_0_imp_0) +qed + + +subsection {* Miscellaneous *} + +lemma odd_pos: "odd (n::nat) \ 0 < n" by presburger + +lemma [presburger]:"(x + 1) div 2 = x div 2 \ even (x::int)" by presburger +lemma [presburger]: "(x + 1) div 2 = x div 2 + 1 \ odd (x::int)" by presburger +lemma even_plus_one_div_two: "even (x::int) ==> (x + 1) div 2 = x div 2" by presburger +lemma odd_plus_one_div_two: "odd (x::int) ==> (x + 1) div 2 = x div 2 + 1" by presburger + +lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \ even x" by presburger +lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \ even x" by presburger +lemma even_nat_plus_one_div_two: "even (x::nat) ==> + (Suc x) div Suc (Suc 0) = x div Suc (Suc 0)" by presburger + +lemma odd_nat_plus_one_div_two: "odd (x::nat) ==> + (Suc x) div Suc (Suc 0) = Suc (x div Suc (Suc 0))" by presburger + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/RComplete.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/RComplete.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,1333 @@ +(* Title : HOL/RComplete.thy + Author : Jacques D. Fleuriot, University of Edinburgh + Author : Larry Paulson, University of Cambridge + Author : Jeremy Avigad, Carnegie Mellon University + Author : Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen +*) + +header {* Completeness of the Reals; Floor and Ceiling Functions *} + +theory RComplete +imports Lubs RealDef +begin + +lemma real_sum_of_halves: "x/2 + x/2 = (x::real)" + by simp + + +subsection {* Completeness of Positive Reals *} + +text {* + Supremum property for the set of positive reals + + Let @{text "P"} be a non-empty set of positive reals, with an upper + bound @{text "y"}. Then @{text "P"} has a least upper bound + (written @{text "S"}). + + FIXME: Can the premise be weakened to @{text "\x \ P. x\ y"}? +*} + +lemma posreal_complete: + assumes positive_P: "\x \ P. (0::real) < x" + and not_empty_P: "\x. x \ P" + and upper_bound_Ex: "\y. \x \ P. xS. \y. (\x \ P. y < x) = (y < S)" +proof (rule exI, rule allI) + fix y + let ?pP = "{w. real_of_preal w \ P}" + + show "(\x\P. y < x) = (y < real_of_preal (psup ?pP))" + proof (cases "0 < y") + assume neg_y: "\ 0 < y" + show ?thesis + proof + assume "\x\P. y < x" + have "\x. y < real_of_preal x" + using neg_y by (rule real_less_all_real2) + thus "y < real_of_preal (psup ?pP)" .. + next + assume "y < real_of_preal (psup ?pP)" + obtain "x" where x_in_P: "x \ P" using not_empty_P .. + hence "0 < x" using positive_P by simp + hence "y < x" using neg_y by simp + thus "\x \ P. y < x" using x_in_P .. + qed + next + assume pos_y: "0 < y" + + then obtain py where y_is_py: "y = real_of_preal py" + by (auto simp add: real_gt_zero_preal_Ex) + + obtain a where "a \ P" using not_empty_P .. + with positive_P have a_pos: "0 < a" .. + then obtain pa where "a = real_of_preal pa" + by (auto simp add: real_gt_zero_preal_Ex) + hence "pa \ ?pP" using `a \ P` by auto + hence pP_not_empty: "?pP \ {}" by auto + + obtain sup where sup: "\x \ P. x < sup" + using upper_bound_Ex .. + from this and `a \ P` have "a < sup" .. + hence "0 < sup" using a_pos by arith + then obtain possup where "sup = real_of_preal possup" + by (auto simp add: real_gt_zero_preal_Ex) + hence "\X \ ?pP. X \ possup" + using sup by (auto simp add: real_of_preal_lessI) + with pP_not_empty have psup: "\Z. (\X \ ?pP. Z < X) = (Z < psup ?pP)" + by (rule preal_complete) + + show ?thesis + proof + assume "\x \ P. y < x" + then obtain x where x_in_P: "x \ P" and y_less_x: "y < x" .. + hence "0 < x" using pos_y by arith + then obtain px where x_is_px: "x = real_of_preal px" + by (auto simp add: real_gt_zero_preal_Ex) + + have py_less_X: "\X \ ?pP. py < X" + proof + show "py < px" using y_is_py and x_is_px and y_less_x + by (simp add: real_of_preal_lessI) + show "px \ ?pP" using x_in_P and x_is_px by simp + qed + + have "(\X \ ?pP. py < X) ==> (py < psup ?pP)" + using psup by simp + hence "py < psup ?pP" using py_less_X by simp + thus "y < real_of_preal (psup {w. real_of_preal w \ P})" + using y_is_py and pos_y by (simp add: real_of_preal_lessI) + next + assume y_less_psup: "y < real_of_preal (psup ?pP)" + + hence "py < psup ?pP" using y_is_py + by (simp add: real_of_preal_lessI) + then obtain "X" where py_less_X: "py < X" and X_in_pP: "X \ ?pP" + using psup by auto + then obtain x where x_is_X: "x = real_of_preal X" + by (simp add: real_gt_zero_preal_Ex) + hence "y < x" using py_less_X and y_is_py + by (simp add: real_of_preal_lessI) + + moreover have "x \ P" using x_is_X and X_in_pP by simp + + ultimately show "\ x \ P. y < x" .. + qed + qed +qed + +text {* + \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc. +*} + +lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)" + apply (frule isLub_isUb) + apply (frule_tac x = y in isLub_isUb) + apply (blast intro!: order_antisym dest!: isLub_le_isUb) + done + + +text {* + \medskip Completeness theorem for the positive reals (again). +*} + +lemma posreals_complete: + assumes positive_S: "\x \ S. 0 < x" + and not_empty_S: "\x. x \ S" + and upper_bound_Ex: "\u. isUb (UNIV::real set) S u" + shows "\t. isLub (UNIV::real set) S t" +proof + let ?pS = "{w. real_of_preal w \ S}" + + obtain u where "isUb UNIV S u" using upper_bound_Ex .. + hence sup: "\x \ S. x \ u" by (simp add: isUb_def setle_def) + + obtain x where x_in_S: "x \ S" using not_empty_S .. + hence x_gt_zero: "0 < x" using positive_S by simp + have "x \ u" using sup and x_in_S .. + hence "0 < u" using x_gt_zero by arith + + then obtain pu where u_is_pu: "u = real_of_preal pu" + by (auto simp add: real_gt_zero_preal_Ex) + + have pS_less_pu: "\pa \ ?pS. pa \ pu" + proof + fix pa + assume "pa \ ?pS" + then obtain a where "a \ S" and "a = real_of_preal pa" + by simp + moreover hence "a \ u" using sup by simp + ultimately show "pa \ pu" + using sup and u_is_pu by (simp add: real_of_preal_le_iff) + qed + + have "\y \ S. y \ real_of_preal (psup ?pS)" + proof + fix y + assume y_in_S: "y \ S" + hence "0 < y" using positive_S by simp + then obtain py where y_is_py: "y = real_of_preal py" + by (auto simp add: real_gt_zero_preal_Ex) + hence py_in_pS: "py \ ?pS" using y_in_S by simp + with pS_less_pu have "py \ psup ?pS" + by (rule preal_psup_le) + thus "y \ real_of_preal (psup ?pS)" + using y_is_py by (simp add: real_of_preal_le_iff) + qed + + moreover { + fix x + assume x_ub_S: "\y\S. y \ x" + have "real_of_preal (psup ?pS) \ x" + proof - + obtain "s" where s_in_S: "s \ S" using not_empty_S .. + hence s_pos: "0 < s" using positive_S by simp + + hence "\ ps. s = real_of_preal ps" by (simp add: real_gt_zero_preal_Ex) + then obtain "ps" where s_is_ps: "s = real_of_preal ps" .. + hence ps_in_pS: "ps \ {w. real_of_preal w \ S}" using s_in_S by simp + + from x_ub_S have "s \ x" using s_in_S .. + hence "0 < x" using s_pos by simp + hence "\ px. x = real_of_preal px" by (simp add: real_gt_zero_preal_Ex) + then obtain "px" where x_is_px: "x = real_of_preal px" .. + + have "\pe \ ?pS. pe \ px" + proof + fix pe + assume "pe \ ?pS" + hence "real_of_preal pe \ S" by simp + hence "real_of_preal pe \ x" using x_ub_S by simp + thus "pe \ px" using x_is_px by (simp add: real_of_preal_le_iff) + qed + + moreover have "?pS \ {}" using ps_in_pS by auto + ultimately have "(psup ?pS) \ px" by (simp add: psup_le_ub) + thus "real_of_preal (psup ?pS) \ x" using x_is_px by (simp add: real_of_preal_le_iff) + qed + } + ultimately show "isLub UNIV S (real_of_preal (psup ?pS))" + by (simp add: isLub_def leastP_def isUb_def setle_def setge_def) +qed + +text {* + \medskip reals Completeness (again!) +*} + +lemma reals_complete: + assumes notempty_S: "\X. X \ S" + and exists_Ub: "\Y. isUb (UNIV::real set) S Y" + shows "\t. isLub (UNIV :: real set) S t" +proof - + obtain X where X_in_S: "X \ S" using notempty_S .. + obtain Y where Y_isUb: "isUb (UNIV::real set) S Y" + using exists_Ub .. + let ?SHIFT = "{z. \x \S. z = x + (-X) + 1} \ {x. 0 < x}" + + { + fix x + assume "isUb (UNIV::real set) S x" + hence S_le_x: "\ y \ S. y <= x" + by (simp add: isUb_def setle_def) + { + fix s + assume "s \ {z. \x\S. z = x + - X + 1}" + hence "\ x \ S. s = x + -X + 1" .. + then obtain x1 where "x1 \ S" and "s = x1 + (-X) + 1" .. + moreover hence "x1 \ x" using S_le_x by simp + ultimately have "s \ x + - X + 1" by arith + } + then have "isUb (UNIV::real set) ?SHIFT (x + (-X) + 1)" + by (auto simp add: isUb_def setle_def) + } note S_Ub_is_SHIFT_Ub = this + + hence "isUb UNIV ?SHIFT (Y + (-X) + 1)" using Y_isUb by simp + hence "\Z. isUb UNIV ?SHIFT Z" .. + moreover have "\y \ ?SHIFT. 0 < y" by auto + moreover have shifted_not_empty: "\u. u \ ?SHIFT" + using X_in_S and Y_isUb by auto + ultimately obtain t where t_is_Lub: "isLub UNIV ?SHIFT t" + using posreals_complete [of ?SHIFT] by blast + + show ?thesis + proof + show "isLub UNIV S (t + X + (-1))" + proof (rule isLubI2) + { + fix x + assume "isUb (UNIV::real set) S x" + hence "isUb (UNIV::real set) (?SHIFT) (x + (-X) + 1)" + using S_Ub_is_SHIFT_Ub by simp + hence "t \ (x + (-X) + 1)" + using t_is_Lub by (simp add: isLub_le_isUb) + hence "t + X + -1 \ x" by arith + } + then show "(t + X + -1) <=* Collect (isUb UNIV S)" + by (simp add: setgeI) + next + show "isUb UNIV S (t + X + -1)" + proof - + { + fix y + assume y_in_S: "y \ S" + have "y \ t + X + -1" + proof - + obtain "u" where u_in_shift: "u \ ?SHIFT" using shifted_not_empty .. + hence "\ x \ S. u = x + - X + 1" by simp + then obtain "x" where x_and_u: "u = x + - X + 1" .. + have u_le_t: "u \ t" using u_in_shift and t_is_Lub by (simp add: isLubD2) + + show ?thesis + proof cases + assume "y \ x" + moreover have "x = u + X + - 1" using x_and_u by arith + moreover have "u + X + - 1 \ t + X + -1" using u_le_t by arith + ultimately show "y \ t + X + -1" by arith + next + assume "~(y \ x)" + hence x_less_y: "x < y" by arith + + have "x + (-X) + 1 \ ?SHIFT" using x_and_u and u_in_shift by simp + hence "0 < x + (-X) + 1" by simp + hence "0 < y + (-X) + 1" using x_less_y by arith + hence "y + (-X) + 1 \ ?SHIFT" using y_in_S by simp + hence "y + (-X) + 1 \ t" using t_is_Lub by (simp add: isLubD2) + thus ?thesis by simp + qed + qed + } + then show ?thesis by (simp add: isUb_def setle_def) + qed + qed + qed +qed + + +subsection {* The Archimedean Property of the Reals *} + +theorem reals_Archimedean: + assumes x_pos: "0 < x" + shows "\n. inverse (real (Suc n)) < x" +proof (rule ccontr) + assume contr: "\ ?thesis" + have "\n. x * real (Suc n) <= 1" + proof + fix n + from contr have "x \ inverse (real (Suc n))" + by (simp add: linorder_not_less) + hence "x \ (1 / (real (Suc n)))" + by (simp add: inverse_eq_divide) + moreover have "0 \ real (Suc n)" + by (rule real_of_nat_ge_zero) + ultimately have "x * real (Suc n) \ (1 / real (Suc n)) * real (Suc n)" + by (rule mult_right_mono) + thus "x * real (Suc n) \ 1" by simp + qed + hence "{z. \n. z = x * (real (Suc n))} *<= 1" + by (simp add: setle_def, safe, rule spec) + hence "isUb (UNIV::real set) {z. \n. z = x * (real (Suc n))} 1" + by (simp add: isUbI) + hence "\Y. isUb (UNIV::real set) {z. \n. z = x* (real (Suc n))} Y" .. + moreover have "\X. X \ {z. \n. z = x* (real (Suc n))}" by auto + ultimately have "\t. isLub UNIV {z. \n. z = x * real (Suc n)} t" + by (simp add: reals_complete) + then obtain "t" where + t_is_Lub: "isLub UNIV {z. \n. z = x * real (Suc n)} t" .. + + have "\n::nat. x * real n \ t + - x" + proof + fix n + from t_is_Lub have "x * real (Suc n) \ t" + by (simp add: isLubD2) + hence "x * (real n) + x \ t" + by (simp add: right_distrib real_of_nat_Suc) + thus "x * (real n) \ t + - x" by arith + qed + + hence "\m. x * real (Suc m) \ t + - x" by simp + hence "{z. \n. z = x * (real (Suc n))} *<= (t + - x)" + by (auto simp add: setle_def) + hence "isUb (UNIV::real set) {z. \n. z = x * (real (Suc n))} (t + (-x))" + by (simp add: isUbI) + hence "t \ t + - x" + using t_is_Lub by (simp add: isLub_le_isUb) + thus False using x_pos by arith +qed + +text {* + There must be other proofs, e.g. @{text "Suc"} of the largest + integer in the cut representing @{text "x"}. +*} + +lemma reals_Archimedean2: "\n. (x::real) < real (n::nat)" +proof cases + assume "x \ 0" + hence "x < real (1::nat)" by simp + thus ?thesis .. +next + assume "\ x \ 0" + hence x_greater_zero: "0 < x" by simp + hence "0 < inverse x" by simp + then obtain n where "inverse (real (Suc n)) < inverse x" + using reals_Archimedean by blast + hence "inverse (real (Suc n)) * x < inverse x * x" + using x_greater_zero by (rule mult_strict_right_mono) + hence "inverse (real (Suc n)) * x < 1" + using x_greater_zero by simp + hence "real (Suc n) * (inverse (real (Suc n)) * x) < real (Suc n) * 1" + by (rule mult_strict_left_mono) simp + hence "x < real (Suc n)" + by (simp add: ring_simps) + thus "\(n::nat). x < real n" .. +qed + +lemma reals_Archimedean3: + assumes x_greater_zero: "0 < x" + shows "\(y::real). \(n::nat). y < real n * x" +proof + fix y + have x_not_zero: "x \ 0" using x_greater_zero by simp + obtain n where "y * inverse x < real (n::nat)" + using reals_Archimedean2 .. + hence "y * inverse x * x < real n * x" + using x_greater_zero by (simp add: mult_strict_right_mono) + hence "x * inverse x * y < x * real n" + by (simp add: ring_simps) + hence "y < real (n::nat) * x" + using x_not_zero by (simp add: ring_simps) + thus "\(n::nat). y < real n * x" .. +qed + +lemma reals_Archimedean6: + "0 \ r ==> \(n::nat). real (n - 1) \ r & r < real (n)" +apply (insert reals_Archimedean2 [of r], safe) +apply (subgoal_tac "\x::nat. r < real x \ (\y. r < real y \ x \ y)", auto) +apply (rule_tac x = x in exI) +apply (case_tac x, simp) +apply (rename_tac x') +apply (drule_tac x = x' in spec, simp) +apply (rule_tac x="LEAST n. r < real n" in exI, safe) +apply (erule LeastI, erule Least_le) +done + +lemma reals_Archimedean6a: "0 \ r ==> \n. real (n) \ r & r < real (Suc n)" + by (drule reals_Archimedean6) auto + +lemma reals_Archimedean_6b_int: + "0 \ r ==> \n::int. real n \ r & r < real (n+1)" +apply (drule reals_Archimedean6a, auto) +apply (rule_tac x = "int n" in exI) +apply (simp add: real_of_int_real_of_nat real_of_nat_Suc) +done + +lemma reals_Archimedean_6c_int: + "r < 0 ==> \n::int. real n \ r & r < real (n+1)" +apply (rule reals_Archimedean_6b_int [of "-r", THEN exE], simp, auto) +apply (rename_tac n) +apply (drule order_le_imp_less_or_eq, auto) +apply (rule_tac x = "- n - 1" in exI) +apply (rule_tac [2] x = "- n" in exI, auto) +done + + +subsection{*Density of the Rational Reals in the Reals*} + +text{* This density proof is due to Stefan Richter and was ported by TN. The +original source is \emph{Real Analysis} by H.L. Royden. +It employs the Archimedean property of the reals. *} + +lemma Rats_dense_in_nn_real: fixes x::real +assumes "0\x" and "xr \ \. x r "LEAST n. y \ real (Suc n)/real q" + from reals_Archimedean2 obtain n::nat where "y * real q < real n" by auto + with `0 < real q` have ex: "y \ real n/real q" (is "?P n") + by (simp add: pos_less_divide_eq[THEN sym]) + also from assms have "\ y \ real (0::nat) / real q" by simp + ultimately have main: "(LEAST n. y \ real n/real q) = Suc p" + by (unfold p_def) (rule Least_Suc) + also from ex have "?P (LEAST x. ?P x)" by (rule LeastI) + ultimately have suc: "y \ real (Suc p) / real q" by simp + def r \ "real p/real q" + have "x = y-(y-x)" by simp + also from suc q have "\ < real (Suc p)/real q - inverse (real q)" by arith + also have "\ = real p / real q" + by (simp only: inverse_eq_divide real_diff_def real_of_nat_Suc + minus_divide_left add_divide_distrib[THEN sym]) simp + finally have "x ?P p" by (rule not_less_Least) + hence "r \" by simp + with `xr \ \. x r x + real n" by arith + also from `xr \ \. x + real n < r \ r < y + real n" + by(rule Rats_dense_in_nn_real) + then obtain r where "r \ \" and r2: "x + real n < r" + and r3: "r < y + real n" + by blast + have "r - real n = r + real (int n)/real (-1::int)" by simp + also from `r\\` have "r + real (int n)/real (-1::int) \ \" by simp + also from r2 have "x < r - real n" by arith + moreover from r3 have "r - real n < y" by arith + ultimately show ?thesis by fast +qed + + +subsection{*Floor and Ceiling Functions from the Reals to the Integers*} + +definition + floor :: "real => int" where + [code del]: "floor r = (LEAST n::int. r < real (n+1))" + +definition + ceiling :: "real => int" where + "ceiling r = - floor (- r)" + +notation (xsymbols) + floor ("\_\") and + ceiling ("\_\") + +notation (HTML output) + floor ("\_\") and + ceiling ("\_\") + + +lemma number_of_less_real_of_int_iff [simp]: + "((number_of n) < real (m::int)) = (number_of n < m)" +apply auto +apply (rule real_of_int_less_iff [THEN iffD1]) +apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto) +done + +lemma number_of_less_real_of_int_iff2 [simp]: + "(real (m::int) < (number_of n)) = (m < number_of n)" +apply auto +apply (rule real_of_int_less_iff [THEN iffD1]) +apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto) +done + +lemma number_of_le_real_of_int_iff [simp]: + "((number_of n) \ real (m::int)) = (number_of n \ m)" +by (simp add: linorder_not_less [symmetric]) + +lemma number_of_le_real_of_int_iff2 [simp]: + "(real (m::int) \ (number_of n)) = (m \ number_of n)" +by (simp add: linorder_not_less [symmetric]) + +lemma floor_zero [simp]: "floor 0 = 0" +apply (simp add: floor_def del: real_of_int_add) +apply (rule Least_equality) +apply simp_all +done + +lemma floor_real_of_nat_zero [simp]: "floor (real (0::nat)) = 0" +by auto + +lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n" +apply (simp only: floor_def) +apply (rule Least_equality) +apply (drule_tac [2] real_of_int_of_nat_eq [THEN ssubst]) +apply (drule_tac [2] real_of_int_less_iff [THEN iffD1]) +apply simp_all +done + +lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n" +apply (simp only: floor_def) +apply (rule Least_equality) +apply (drule_tac [2] real_of_int_of_nat_eq [THEN ssubst]) +apply (drule_tac [2] real_of_int_minus [THEN sym, THEN subst]) +apply (drule_tac [2] real_of_int_less_iff [THEN iffD1]) +apply simp_all +done + +lemma floor_real_of_int [simp]: "floor (real (n::int)) = n" +apply (simp only: floor_def) +apply (rule Least_equality) +apply auto +done + +lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n" +apply (simp only: floor_def) +apply (rule Least_equality) +apply (drule_tac [2] real_of_int_minus [THEN sym, THEN subst]) +apply auto +done + +lemma real_lb_ub_int: " \n::int. real n \ r & r < real (n+1)" +apply (case_tac "r < 0") +apply (blast intro: reals_Archimedean_6c_int) +apply (simp only: linorder_not_less) +apply (blast intro: reals_Archimedean_6b_int reals_Archimedean_6c_int) +done + +lemma lemma_floor: + assumes a1: "real m \ r" and a2: "r < real n + 1" + shows "m \ (n::int)" +proof - + have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans) + also have "... = real (n + 1)" by simp + finally have "m < n + 1" by (simp only: real_of_int_less_iff) + thus ?thesis by arith +qed + +lemma real_of_int_floor_le [simp]: "real (floor r) \ r" +apply (simp add: floor_def Least_def) +apply (insert real_lb_ub_int [of r], safe) +apply (rule theI2) +apply auto +done + +lemma floor_mono: "x < y ==> floor x \ floor y" +apply (simp add: floor_def Least_def) +apply (insert real_lb_ub_int [of x]) +apply (insert real_lb_ub_int [of y], safe) +apply (rule theI2) +apply (rule_tac [3] theI2) +apply simp +apply (erule conjI) +apply (auto simp add: order_eq_iff int_le_real_less) +done + +lemma floor_mono2: "x \ y ==> floor x \ floor y" +by (auto dest: order_le_imp_less_or_eq simp add: floor_mono) + +lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \ x" +by (auto intro: lemma_floor) + +lemma real_of_int_floor_cancel [simp]: + "(real (floor x) = x) = (\n::int. x = real n)" +apply (simp add: floor_def Least_def) +apply (insert real_lb_ub_int [of x], erule exE) +apply (rule theI2) +apply (auto intro: lemma_floor) +done + +lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n" +apply (simp add: floor_def) +apply (rule Least_equality) +apply (auto intro: lemma_floor) +done + +lemma floor_eq2: "[| real n \ x; x < real n + 1 |] ==> floor x = n" +apply (simp add: floor_def) +apply (rule Least_equality) +apply (auto intro: lemma_floor) +done + +lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n" +apply (rule inj_int [THEN injD]) +apply (simp add: real_of_nat_Suc) +apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"]) +done + +lemma floor_eq4: "[| real n \ x; x < real (Suc n) |] ==> nat(floor x) = n" +apply (drule order_le_imp_less_or_eq) +apply (auto intro: floor_eq3) +done + +lemma floor_number_of_eq [simp]: + "floor(number_of n :: real) = (number_of n :: int)" +apply (subst real_number_of [symmetric]) +apply (rule floor_real_of_int) +done + +lemma floor_one [simp]: "floor 1 = 1" + apply (rule trans) + prefer 2 + apply (rule floor_real_of_int) + apply simp +done + +lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \ real(floor r)" +apply (simp add: floor_def Least_def) +apply (insert real_lb_ub_int [of r], safe) +apply (rule theI2) +apply (auto intro: lemma_floor) +done + +lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)" +apply (simp add: floor_def Least_def) +apply (insert real_lb_ub_int [of r], safe) +apply (rule theI2) +apply (auto intro: lemma_floor) +done + +lemma real_of_int_floor_add_one_ge [simp]: "r \ real(floor r) + 1" +apply (insert real_of_int_floor_ge_diff_one [of r]) +apply (auto simp del: real_of_int_floor_ge_diff_one) +done + +lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1" +apply (insert real_of_int_floor_gt_diff_one [of r]) +apply (auto simp del: real_of_int_floor_gt_diff_one) +done + +lemma le_floor: "real a <= x ==> a <= floor x" + apply (subgoal_tac "a < floor x + 1") + apply arith + apply (subst real_of_int_less_iff [THEN sym]) + apply simp + apply (insert real_of_int_floor_add_one_gt [of x]) + apply arith +done + +lemma real_le_floor: "a <= floor x ==> real a <= x" + apply (rule order_trans) + prefer 2 + apply (rule real_of_int_floor_le) + apply (subst real_of_int_le_iff) + apply assumption +done + +lemma le_floor_eq: "(a <= floor x) = (real a <= x)" + apply (rule iffI) + apply (erule real_le_floor) + apply (erule le_floor) +done + +lemma le_floor_eq_number_of [simp]: + "(number_of n <= floor x) = (number_of n <= x)" +by (simp add: le_floor_eq) + +lemma le_floor_eq_zero [simp]: "(0 <= floor x) = (0 <= x)" +by (simp add: le_floor_eq) + +lemma le_floor_eq_one [simp]: "(1 <= floor x) = (1 <= x)" +by (simp add: le_floor_eq) + +lemma floor_less_eq: "(floor x < a) = (x < real a)" + apply (subst linorder_not_le [THEN sym])+ + apply simp + apply (rule le_floor_eq) +done + +lemma floor_less_eq_number_of [simp]: + "(floor x < number_of n) = (x < number_of n)" +by (simp add: floor_less_eq) + +lemma floor_less_eq_zero [simp]: "(floor x < 0) = (x < 0)" +by (simp add: floor_less_eq) + +lemma floor_less_eq_one [simp]: "(floor x < 1) = (x < 1)" +by (simp add: floor_less_eq) + +lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)" + apply (insert le_floor_eq [of "a + 1" x]) + apply auto +done + +lemma less_floor_eq_number_of [simp]: + "(number_of n < floor x) = (number_of n + 1 <= x)" +by (simp add: less_floor_eq) + +lemma less_floor_eq_zero [simp]: "(0 < floor x) = (1 <= x)" +by (simp add: less_floor_eq) + +lemma less_floor_eq_one [simp]: "(1 < floor x) = (2 <= x)" +by (simp add: less_floor_eq) + +lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)" + apply (insert floor_less_eq [of x "a + 1"]) + apply auto +done + +lemma floor_le_eq_number_of [simp]: + "(floor x <= number_of n) = (x < number_of n + 1)" +by (simp add: floor_le_eq) + +lemma floor_le_eq_zero [simp]: "(floor x <= 0) = (x < 1)" +by (simp add: floor_le_eq) + +lemma floor_le_eq_one [simp]: "(floor x <= 1) = (x < 2)" +by (simp add: floor_le_eq) + +lemma floor_add [simp]: "floor (x + real a) = floor x + a" + apply (subst order_eq_iff) + apply (rule conjI) + prefer 2 + apply (subgoal_tac "floor x + a < floor (x + real a) + 1") + apply arith + apply (subst real_of_int_less_iff [THEN sym]) + apply simp + apply (subgoal_tac "x + real a < real(floor(x + real a)) + 1") + apply (subgoal_tac "real (floor x) <= x") + apply arith + apply (rule real_of_int_floor_le) + apply (rule real_of_int_floor_add_one_gt) + apply (subgoal_tac "floor (x + real a) < floor x + a + 1") + apply arith + apply (subst real_of_int_less_iff [THEN sym]) + apply simp + apply (subgoal_tac "real(floor(x + real a)) <= x + real a") + apply (subgoal_tac "x < real(floor x) + 1") + apply arith + apply (rule real_of_int_floor_add_one_gt) + apply (rule real_of_int_floor_le) +done + +lemma floor_add_number_of [simp]: + "floor (x + number_of n) = floor x + number_of n" + apply (subst floor_add [THEN sym]) + apply simp +done + +lemma floor_add_one [simp]: "floor (x + 1) = floor x + 1" + apply (subst floor_add [THEN sym]) + apply simp +done + +lemma floor_subtract [simp]: "floor (x - real a) = floor x - a" + apply (subst diff_minus)+ + apply (subst real_of_int_minus [THEN sym]) + apply (rule floor_add) +done + +lemma floor_subtract_number_of [simp]: "floor (x - number_of n) = + floor x - number_of n" + apply (subst floor_subtract [THEN sym]) + apply simp +done + +lemma floor_subtract_one [simp]: "floor (x - 1) = floor x - 1" + apply (subst floor_subtract [THEN sym]) + apply simp +done + +lemma ceiling_zero [simp]: "ceiling 0 = 0" +by (simp add: ceiling_def) + +lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n" +by (simp add: ceiling_def) + +lemma ceiling_real_of_nat_zero [simp]: "ceiling (real (0::nat)) = 0" +by auto + +lemma ceiling_floor [simp]: "ceiling (real (floor r)) = floor r" +by (simp add: ceiling_def) + +lemma floor_ceiling [simp]: "floor (real (ceiling r)) = ceiling r" +by (simp add: ceiling_def) + +lemma real_of_int_ceiling_ge [simp]: "r \ real (ceiling r)" +apply (simp add: ceiling_def) +apply (subst le_minus_iff, simp) +done + +lemma ceiling_mono: "x < y ==> ceiling x \ ceiling y" +by (simp add: floor_mono ceiling_def) + +lemma ceiling_mono2: "x \ y ==> ceiling x \ ceiling y" +by (simp add: floor_mono2 ceiling_def) + +lemma real_of_int_ceiling_cancel [simp]: + "(real (ceiling x) = x) = (\n::int. x = real n)" +apply (auto simp add: ceiling_def) +apply (drule arg_cong [where f = uminus], auto) +apply (rule_tac x = "-n" in exI, auto) +done + +lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1" +apply (simp add: ceiling_def) +apply (rule minus_equation_iff [THEN iffD1]) +apply (simp add: floor_eq [where n = "-(n+1)"]) +done + +lemma ceiling_eq2: "[| real n < x; x \ real n + 1 |] ==> ceiling x = n + 1" +by (simp add: ceiling_def floor_eq2 [where n = "-(n+1)"]) + +lemma ceiling_eq3: "[| real n - 1 < x; x \ real n |] ==> ceiling x = n" +by (simp add: ceiling_def floor_eq2 [where n = "-n"]) + +lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n" +by (simp add: ceiling_def) + +lemma ceiling_number_of_eq [simp]: + "ceiling (number_of n :: real) = (number_of n)" +apply (subst real_number_of [symmetric]) +apply (rule ceiling_real_of_int) +done + +lemma ceiling_one [simp]: "ceiling 1 = 1" + by (unfold ceiling_def, simp) + +lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \ r" +apply (rule neg_le_iff_le [THEN iffD1]) +apply (simp add: ceiling_def diff_minus) +done + +lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \ r + 1" +apply (insert real_of_int_ceiling_diff_one_le [of r]) +apply (simp del: real_of_int_ceiling_diff_one_le) +done + +lemma ceiling_le: "x <= real a ==> ceiling x <= a" + apply (unfold ceiling_def) + apply (subgoal_tac "-a <= floor(- x)") + apply simp + apply (rule le_floor) + apply simp +done + +lemma ceiling_le_real: "ceiling x <= a ==> x <= real a" + apply (unfold ceiling_def) + apply (subgoal_tac "real(- a) <= - x") + apply simp + apply (rule real_le_floor) + apply simp +done + +lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)" + apply (rule iffI) + apply (erule ceiling_le_real) + apply (erule ceiling_le) +done + +lemma ceiling_le_eq_number_of [simp]: + "(ceiling x <= number_of n) = (x <= number_of n)" +by (simp add: ceiling_le_eq) + +lemma ceiling_le_zero_eq [simp]: "(ceiling x <= 0) = (x <= 0)" +by (simp add: ceiling_le_eq) + +lemma ceiling_le_eq_one [simp]: "(ceiling x <= 1) = (x <= 1)" +by (simp add: ceiling_le_eq) + +lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)" + apply (subst linorder_not_le [THEN sym])+ + apply simp + apply (rule ceiling_le_eq) +done + +lemma less_ceiling_eq_number_of [simp]: + "(number_of n < ceiling x) = (number_of n < x)" +by (simp add: less_ceiling_eq) + +lemma less_ceiling_eq_zero [simp]: "(0 < ceiling x) = (0 < x)" +by (simp add: less_ceiling_eq) + +lemma less_ceiling_eq_one [simp]: "(1 < ceiling x) = (1 < x)" +by (simp add: less_ceiling_eq) + +lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)" + apply (insert ceiling_le_eq [of x "a - 1"]) + apply auto +done + +lemma ceiling_less_eq_number_of [simp]: + "(ceiling x < number_of n) = (x <= number_of n - 1)" +by (simp add: ceiling_less_eq) + +lemma ceiling_less_eq_zero [simp]: "(ceiling x < 0) = (x <= -1)" +by (simp add: ceiling_less_eq) + +lemma ceiling_less_eq_one [simp]: "(ceiling x < 1) = (x <= 0)" +by (simp add: ceiling_less_eq) + +lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)" + apply (insert less_ceiling_eq [of "a - 1" x]) + apply auto +done + +lemma le_ceiling_eq_number_of [simp]: + "(number_of n <= ceiling x) = (number_of n - 1 < x)" +by (simp add: le_ceiling_eq) + +lemma le_ceiling_eq_zero [simp]: "(0 <= ceiling x) = (-1 < x)" +by (simp add: le_ceiling_eq) + +lemma le_ceiling_eq_one [simp]: "(1 <= ceiling x) = (0 < x)" +by (simp add: le_ceiling_eq) + +lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a" + apply (unfold ceiling_def, simp) + apply (subst real_of_int_minus [THEN sym]) + apply (subst floor_add) + apply simp +done + +lemma ceiling_add_number_of [simp]: "ceiling (x + number_of n) = + ceiling x + number_of n" + apply (subst ceiling_add [THEN sym]) + apply simp +done + +lemma ceiling_add_one [simp]: "ceiling (x + 1) = ceiling x + 1" + apply (subst ceiling_add [THEN sym]) + apply simp +done + +lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a" + apply (subst diff_minus)+ + apply (subst real_of_int_minus [THEN sym]) + apply (rule ceiling_add) +done + +lemma ceiling_subtract_number_of [simp]: "ceiling (x - number_of n) = + ceiling x - number_of n" + apply (subst ceiling_subtract [THEN sym]) + apply simp +done + +lemma ceiling_subtract_one [simp]: "ceiling (x - 1) = ceiling x - 1" + apply (subst ceiling_subtract [THEN sym]) + apply simp +done + +subsection {* Versions for the natural numbers *} + +definition + natfloor :: "real => nat" where + "natfloor x = nat(floor x)" + +definition + natceiling :: "real => nat" where + "natceiling x = nat(ceiling x)" + +lemma natfloor_zero [simp]: "natfloor 0 = 0" + by (unfold natfloor_def, simp) + +lemma natfloor_one [simp]: "natfloor 1 = 1" + by (unfold natfloor_def, simp) + +lemma zero_le_natfloor [simp]: "0 <= natfloor x" + by (unfold natfloor_def, simp) + +lemma natfloor_number_of_eq [simp]: "natfloor (number_of n) = number_of n" + by (unfold natfloor_def, simp) + +lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n" + by (unfold natfloor_def, simp) + +lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x" + by (unfold natfloor_def, simp) + +lemma natfloor_neg: "x <= 0 ==> natfloor x = 0" + apply (unfold natfloor_def) + apply (subgoal_tac "floor x <= floor 0") + apply simp + apply (erule floor_mono2) +done + +lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y" + apply (case_tac "0 <= x") + apply (subst natfloor_def)+ + apply (subst nat_le_eq_zle) + apply force + apply (erule floor_mono2) + apply (subst natfloor_neg) + apply simp + apply simp +done + +lemma le_natfloor: "real x <= a ==> x <= natfloor a" + apply (unfold natfloor_def) + apply (subst nat_int [THEN sym]) + apply (subst nat_le_eq_zle) + apply simp + apply (rule le_floor) + apply simp +done + +lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)" + apply (rule iffI) + apply (rule order_trans) + prefer 2 + apply (erule real_natfloor_le) + apply (subst real_of_nat_le_iff) + apply assumption + apply (erule le_natfloor) +done + +lemma le_natfloor_eq_number_of [simp]: + "~ neg((number_of n)::int) ==> 0 <= x ==> + (number_of n <= natfloor x) = (number_of n <= x)" + apply (subst le_natfloor_eq, assumption) + apply simp +done + +lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)" + apply (case_tac "0 <= x") + apply (subst le_natfloor_eq, assumption, simp) + apply (rule iffI) + apply (subgoal_tac "natfloor x <= natfloor 0") + apply simp + apply (rule natfloor_mono) + apply simp + apply simp +done + +lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n" + apply (unfold natfloor_def) + apply (subst nat_int [THEN sym]);back; + apply (subst eq_nat_nat_iff) + apply simp + apply simp + apply (rule floor_eq2) + apply auto +done + +lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1" + apply (case_tac "0 <= x") + apply (unfold natfloor_def) + apply simp + apply simp_all +done + +lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)" + apply (simp add: compare_rls) + apply (rule real_natfloor_add_one_gt) +done + +lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n" + apply (subgoal_tac "z < real(natfloor z) + 1") + apply arith + apply (rule real_natfloor_add_one_gt) +done + +lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a" + apply (unfold natfloor_def) + apply (subgoal_tac "real a = real (int a)") + apply (erule ssubst) + apply (simp add: nat_add_distrib del: real_of_int_of_nat_eq) + apply simp +done + +lemma natfloor_add_number_of [simp]: + "~neg ((number_of n)::int) ==> 0 <= x ==> + natfloor (x + number_of n) = natfloor x + number_of n" + apply (subst natfloor_add [THEN sym]) + apply simp_all +done + +lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1" + apply (subst natfloor_add [THEN sym]) + apply assumption + apply simp +done + +lemma natfloor_subtract [simp]: "real a <= x ==> + natfloor(x - real a) = natfloor x - a" + apply (unfold natfloor_def) + apply (subgoal_tac "real a = real (int a)") + apply (erule ssubst) + apply (simp del: real_of_int_of_nat_eq) + apply simp +done + +lemma natceiling_zero [simp]: "natceiling 0 = 0" + by (unfold natceiling_def, simp) + +lemma natceiling_one [simp]: "natceiling 1 = 1" + by (unfold natceiling_def, simp) + +lemma zero_le_natceiling [simp]: "0 <= natceiling x" + by (unfold natceiling_def, simp) + +lemma natceiling_number_of_eq [simp]: "natceiling (number_of n) = number_of n" + by (unfold natceiling_def, simp) + +lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n" + by (unfold natceiling_def, simp) + +lemma real_natceiling_ge: "x <= real(natceiling x)" + apply (unfold natceiling_def) + apply (case_tac "x < 0") + apply simp + apply (subst real_nat_eq_real) + apply (subgoal_tac "ceiling 0 <= ceiling x") + apply simp + apply (rule ceiling_mono2) + apply simp + apply simp +done + +lemma natceiling_neg: "x <= 0 ==> natceiling x = 0" + apply (unfold natceiling_def) + apply simp +done + +lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y" + apply (case_tac "0 <= x") + apply (subst natceiling_def)+ + apply (subst nat_le_eq_zle) + apply (rule disjI2) + apply (subgoal_tac "real (0::int) <= real(ceiling y)") + apply simp + apply (rule order_trans) + apply simp + apply (erule order_trans) + apply simp + apply (erule ceiling_mono2) + apply (subst natceiling_neg) + apply simp_all +done + +lemma natceiling_le: "x <= real a ==> natceiling x <= a" + apply (unfold natceiling_def) + apply (case_tac "x < 0") + apply simp + apply (subst nat_int [THEN sym]);back; + apply (subst nat_le_eq_zle) + apply simp + apply (rule ceiling_le) + apply simp +done + +lemma natceiling_le_eq: "0 <= x ==> (natceiling x <= a) = (x <= real a)" + apply (rule iffI) + apply (rule order_trans) + apply (rule real_natceiling_ge) + apply (subst real_of_nat_le_iff) + apply assumption + apply (erule natceiling_le) +done + +lemma natceiling_le_eq_number_of [simp]: + "~ neg((number_of n)::int) ==> 0 <= x ==> + (natceiling x <= number_of n) = (x <= number_of n)" + apply (subst natceiling_le_eq, assumption) + apply simp +done + +lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)" + apply (case_tac "0 <= x") + apply (subst natceiling_le_eq) + apply assumption + apply simp + apply (subst natceiling_neg) + apply simp + apply simp +done + +lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1" + apply (unfold natceiling_def) + apply (simplesubst nat_int [THEN sym]) back back + apply (subgoal_tac "nat(int n) + 1 = nat(int n + 1)") + apply (erule ssubst) + apply (subst eq_nat_nat_iff) + apply (subgoal_tac "ceiling 0 <= ceiling x") + apply simp + apply (rule ceiling_mono2) + apply force + apply force + apply (rule ceiling_eq2) + apply (simp, simp) + apply (subst nat_add_distrib) + apply auto +done + +lemma natceiling_add [simp]: "0 <= x ==> + natceiling (x + real a) = natceiling x + a" + apply (unfold natceiling_def) + apply (subgoal_tac "real a = real (int a)") + apply (erule ssubst) + apply (simp del: real_of_int_of_nat_eq) + apply (subst nat_add_distrib) + apply (subgoal_tac "0 = ceiling 0") + apply (erule ssubst) + apply (erule ceiling_mono2) + apply simp_all +done + +lemma natceiling_add_number_of [simp]: + "~ neg ((number_of n)::int) ==> 0 <= x ==> + natceiling (x + number_of n) = natceiling x + number_of n" + apply (subst natceiling_add [THEN sym]) + apply simp_all +done + +lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1" + apply (subst natceiling_add [THEN sym]) + apply assumption + apply simp +done + +lemma natceiling_subtract [simp]: "real a <= x ==> + natceiling(x - real a) = natceiling x - a" + apply (unfold natceiling_def) + apply (subgoal_tac "real a = real (int a)") + apply (erule ssubst) + apply (simp del: real_of_int_of_nat_eq) + apply simp +done + +lemma natfloor_div_nat: "1 <= x ==> y > 0 ==> + natfloor (x / real y) = natfloor x div y" +proof - + assume "1 <= (x::real)" and "(y::nat) > 0" + have "natfloor x = (natfloor x) div y * y + (natfloor x) mod y" + by simp + then have a: "real(natfloor x) = real ((natfloor x) div y) * real y + + real((natfloor x) mod y)" + by (simp only: real_of_nat_add [THEN sym] real_of_nat_mult [THEN sym]) + have "x = real(natfloor x) + (x - real(natfloor x))" + by simp + then have "x = real ((natfloor x) div y) * real y + + real((natfloor x) mod y) + (x - real(natfloor x))" + by (simp add: a) + then have "x / real y = ... / real y" + by simp + also have "... = real((natfloor x) div y) + real((natfloor x) mod y) / + real y + (x - real(natfloor x)) / real y" + by (auto simp add: ring_simps add_divide_distrib + diff_divide_distrib prems) + finally have "natfloor (x / real y) = natfloor(...)" by simp + also have "... = natfloor(real((natfloor x) mod y) / + real y + (x - real(natfloor x)) / real y + real((natfloor x) div y))" + by (simp add: add_ac) + also have "... = natfloor(real((natfloor x) mod y) / + real y + (x - real(natfloor x)) / real y) + (natfloor x) div y" + apply (rule natfloor_add) + apply (rule add_nonneg_nonneg) + apply (rule divide_nonneg_pos) + apply simp + apply (simp add: prems) + apply (rule divide_nonneg_pos) + apply (simp add: compare_rls) + apply (rule real_natfloor_le) + apply (insert prems, auto) + done + also have "natfloor(real((natfloor x) mod y) / + real y + (x - real(natfloor x)) / real y) = 0" + apply (rule natfloor_eq) + apply simp + apply (rule add_nonneg_nonneg) + apply (rule divide_nonneg_pos) + apply force + apply (force simp add: prems) + apply (rule divide_nonneg_pos) + apply (simp add: compare_rls) + apply (rule real_natfloor_le) + apply (auto simp add: prems) + apply (insert prems, arith) + apply (simp add: add_divide_distrib [THEN sym]) + apply (subgoal_tac "real y = real y - 1 + 1") + apply (erule ssubst) + apply (rule add_le_less_mono) + apply (simp add: compare_rls) + apply (subgoal_tac "real(natfloor x mod y) + 1 = + real(natfloor x mod y + 1)") + apply (erule ssubst) + apply (subst real_of_nat_le_iff) + apply (subgoal_tac "natfloor x mod y < y") + apply arith + apply (rule mod_less_divisor) + apply auto + apply (simp add: compare_rls) + apply (subst add_commute) + apply (rule real_natfloor_add_one_gt) + done + finally show ?thesis by simp +qed + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ROOT.ML --- a/src/HOL/ROOT.ML Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/ROOT.ML Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/ROOT.ML - ID: $Id$ Classical Higher-order Logic -- batteries included. *) diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Rational.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Rational.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,1019 @@ +(* Title: HOL/Rational.thy + Author: Markus Wenzel, TU Muenchen +*) + +header {* Rational numbers *} + +theory Rational +imports Nat_Int_Bij GCD +uses ("Tools/rat_arith.ML") +begin + +subsection {* Rational numbers as quotient *} + +subsubsection {* Construction of the type of rational numbers *} + +definition + ratrel :: "((int \ int) \ (int \ int)) set" where + "ratrel = {(x, y). snd x \ 0 \ snd y \ 0 \ fst x * snd y = fst y * snd x}" + +lemma ratrel_iff [simp]: + "(x, y) \ ratrel \ snd x \ 0 \ snd y \ 0 \ fst x * snd y = fst y * snd x" + by (simp add: ratrel_def) + +lemma refl_ratrel: "refl {x. snd x \ 0} ratrel" + by (auto simp add: refl_def ratrel_def) + +lemma sym_ratrel: "sym ratrel" + by (simp add: ratrel_def sym_def) + +lemma trans_ratrel: "trans ratrel" +proof (rule transI, unfold split_paired_all) + fix a b a' b' a'' b'' :: int + assume A: "((a, b), (a', b')) \ ratrel" + assume B: "((a', b'), (a'', b'')) \ ratrel" + have "b' * (a * b'') = b'' * (a * b')" by simp + also from A have "a * b' = a' * b" by auto + also have "b'' * (a' * b) = b * (a' * b'')" by simp + also from B have "a' * b'' = a'' * b'" by auto + also have "b * (a'' * b') = b' * (a'' * b)" by simp + finally have "b' * (a * b'') = b' * (a'' * b)" . + moreover from B have "b' \ 0" by auto + ultimately have "a * b'' = a'' * b" by simp + with A B show "((a, b), (a'', b'')) \ ratrel" by auto +qed + +lemma equiv_ratrel: "equiv {x. snd x \ 0} ratrel" + by (rule equiv.intro [OF refl_ratrel sym_ratrel trans_ratrel]) + +lemmas UN_ratrel = UN_equiv_class [OF equiv_ratrel] +lemmas UN_ratrel2 = UN_equiv_class2 [OF equiv_ratrel equiv_ratrel] + +lemma equiv_ratrel_iff [iff]: + assumes "snd x \ 0" and "snd y \ 0" + shows "ratrel `` {x} = ratrel `` {y} \ (x, y) \ ratrel" + by (rule eq_equiv_class_iff, rule equiv_ratrel) (auto simp add: assms) + +typedef (Rat) rat = "{x. snd x \ 0} // ratrel" +proof + have "(0::int, 1::int) \ {x. snd x \ 0}" by simp + then show "ratrel `` {(0, 1)} \ {x. snd x \ 0} // ratrel" by (rule quotientI) +qed + +lemma ratrel_in_Rat [simp]: "snd x \ 0 \ ratrel `` {x} \ Rat" + by (simp add: Rat_def quotientI) + +declare Abs_Rat_inject [simp] Abs_Rat_inverse [simp] + + +subsubsection {* Representation and basic operations *} + +definition + Fract :: "int \ int \ rat" where + [code del]: "Fract a b = Abs_Rat (ratrel `` {if b = 0 then (0, 1) else (a, b)})" + +code_datatype Fract + +lemma Rat_cases [case_names Fract, cases type: rat]: + assumes "\a b. q = Fract a b \ b \ 0 \ C" + shows C + using assms by (cases q) (clarsimp simp add: Fract_def Rat_def quotient_def) + +lemma Rat_induct [case_names Fract, induct type: rat]: + assumes "\a b. b \ 0 \ P (Fract a b)" + shows "P q" + using assms by (cases q) simp + +lemma eq_rat: + shows "\a b c d. b \ 0 \ d \ 0 \ Fract a b = Fract c d \ a * d = c * b" + and "\a. Fract a 0 = Fract 0 1" + and "\a c. Fract 0 a = Fract 0 c" + by (simp_all add: Fract_def) + +instantiation rat :: "{comm_ring_1, recpower}" +begin + +definition + Zero_rat_def [code, code unfold]: "0 = Fract 0 1" + +definition + One_rat_def [code, code unfold]: "1 = Fract 1 1" + +definition + add_rat_def [code del]: + "q + r = Abs_Rat (\x \ Rep_Rat q. \y \ Rep_Rat r. + ratrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)})" + +lemma add_rat [simp]: + assumes "b \ 0" and "d \ 0" + shows "Fract a b + Fract c d = Fract (a * d + c * b) (b * d)" +proof - + have "(\x y. ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)}) + respects2 ratrel" + by (rule equiv_ratrel [THEN congruent2_commuteI]) (simp_all add: left_distrib) + with assms show ?thesis by (simp add: Fract_def add_rat_def UN_ratrel2) +qed + +definition + minus_rat_def [code del]: + "- q = Abs_Rat (\x \ Rep_Rat q. ratrel `` {(- fst x, snd x)})" + +lemma minus_rat [simp, code]: "- Fract a b = Fract (- a) b" +proof - + have "(\x. ratrel `` {(- fst x, snd x)}) respects ratrel" + by (simp add: congruent_def) + then show ?thesis by (simp add: Fract_def minus_rat_def UN_ratrel) +qed + +lemma minus_rat_cancel [simp]: "Fract (- a) (- b) = Fract a b" + by (cases "b = 0") (simp_all add: eq_rat) + +definition + diff_rat_def [code del]: "q - r = q + - (r::rat)" + +lemma diff_rat [simp]: + assumes "b \ 0" and "d \ 0" + shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)" + using assms by (simp add: diff_rat_def) + +definition + mult_rat_def [code del]: + "q * r = Abs_Rat (\x \ Rep_Rat q. \y \ Rep_Rat r. + ratrel``{(fst x * fst y, snd x * snd y)})" + +lemma mult_rat [simp]: "Fract a b * Fract c d = Fract (a * c) (b * d)" +proof - + have "(\x y. ratrel `` {(fst x * fst y, snd x * snd y)}) respects2 ratrel" + by (rule equiv_ratrel [THEN congruent2_commuteI]) simp_all + then show ?thesis by (simp add: Fract_def mult_rat_def UN_ratrel2) +qed + +lemma mult_rat_cancel: + assumes "c \ 0" + shows "Fract (c * a) (c * b) = Fract a b" +proof - + from assms have "Fract c c = Fract 1 1" by (simp add: Fract_def) + then show ?thesis by (simp add: mult_rat [symmetric]) +qed + +primrec power_rat +where + rat_power_0: "q ^ 0 = (1\rat)" + | rat_power_Suc: "q ^ Suc n = (q\rat) * (q ^ n)" + +instance proof + fix q r s :: rat show "(q * r) * s = q * (r * s)" + by (cases q, cases r, cases s) (simp add: eq_rat) +next + fix q r :: rat show "q * r = r * q" + by (cases q, cases r) (simp add: eq_rat) +next + fix q :: rat show "1 * q = q" + by (cases q) (simp add: One_rat_def eq_rat) +next + fix q r s :: rat show "(q + r) + s = q + (r + s)" + by (cases q, cases r, cases s) (simp add: eq_rat ring_simps) +next + fix q r :: rat show "q + r = r + q" + by (cases q, cases r) (simp add: eq_rat) +next + fix q :: rat show "0 + q = q" + by (cases q) (simp add: Zero_rat_def eq_rat) +next + fix q :: rat show "- q + q = 0" + by (cases q) (simp add: Zero_rat_def eq_rat) +next + fix q r :: rat show "q - r = q + - r" + by (cases q, cases r) (simp add: eq_rat) +next + fix q r s :: rat show "(q + r) * s = q * s + r * s" + by (cases q, cases r, cases s) (simp add: eq_rat ring_simps) +next + show "(0::rat) \ 1" by (simp add: Zero_rat_def One_rat_def eq_rat) +next + fix q :: rat show "q * 1 = q" + by (cases q) (simp add: One_rat_def eq_rat) +next + fix q :: rat + fix n :: nat + show "q ^ 0 = 1" by simp + show "q ^ (Suc n) = q * (q ^ n)" by simp +qed + +end + +lemma of_nat_rat: "of_nat k = Fract (of_nat k) 1" + by (induct k) (simp_all add: Zero_rat_def One_rat_def) + +lemma of_int_rat: "of_int k = Fract k 1" + by (cases k rule: int_diff_cases) (simp add: of_nat_rat) + +lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k" + by (rule of_nat_rat [symmetric]) + +lemma Fract_of_int_eq: "Fract k 1 = of_int k" + by (rule of_int_rat [symmetric]) + +instantiation rat :: number_ring +begin + +definition + rat_number_of_def [code del]: "number_of w = Fract w 1" + +instance by intro_classes (simp add: rat_number_of_def of_int_rat) + +end + +lemma rat_number_collapse [code post]: + "Fract 0 k = 0" + "Fract 1 1 = 1" + "Fract (number_of k) 1 = number_of k" + "Fract k 0 = 0" + by (cases "k = 0") + (simp_all add: Zero_rat_def One_rat_def number_of_is_id number_of_eq of_int_rat eq_rat Fract_def) + +lemma rat_number_expand [code unfold]: + "0 = Fract 0 1" + "1 = Fract 1 1" + "number_of k = Fract (number_of k) 1" + by (simp_all add: rat_number_collapse) + +lemma iszero_rat [simp]: + "iszero (number_of k :: rat) \ iszero (number_of k :: int)" + by (simp add: iszero_def rat_number_expand number_of_is_id eq_rat) + +lemma Rat_cases_nonzero [case_names Fract 0]: + assumes Fract: "\a b. q = Fract a b \ b \ 0 \ a \ 0 \ C" + assumes 0: "q = 0 \ C" + shows C +proof (cases "q = 0") + case True then show C using 0 by auto +next + case False + then obtain a b where "q = Fract a b" and "b \ 0" by (cases q) auto + moreover with False have "0 \ Fract a b" by simp + with `b \ 0` have "a \ 0" by (simp add: Zero_rat_def eq_rat) + with Fract `q = Fract a b` `b \ 0` show C by auto +qed + + + +subsubsection {* The field of rational numbers *} + +instantiation rat :: "{field, division_by_zero}" +begin + +definition + inverse_rat_def [code del]: + "inverse q = Abs_Rat (\x \ Rep_Rat q. + ratrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)})" + +lemma inverse_rat [simp]: "inverse (Fract a b) = Fract b a" +proof - + have "(\x. ratrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)}) respects ratrel" + by (auto simp add: congruent_def mult_commute) + then show ?thesis by (simp add: Fract_def inverse_rat_def UN_ratrel) +qed + +definition + divide_rat_def [code del]: "q / r = q * inverse (r::rat)" + +lemma divide_rat [simp]: "Fract a b / Fract c d = Fract (a * d) (b * c)" + by (simp add: divide_rat_def) + +instance proof + show "inverse 0 = (0::rat)" by (simp add: rat_number_expand) + (simp add: rat_number_collapse) +next + fix q :: rat + assume "q \ 0" + then show "inverse q * q = 1" by (cases q rule: Rat_cases_nonzero) + (simp_all add: mult_rat inverse_rat rat_number_expand eq_rat) +next + fix q r :: rat + show "q / r = q * inverse r" by (simp add: divide_rat_def) +qed + +end + + +subsubsection {* Various *} + +lemma Fract_add_one: "n \ 0 ==> Fract (m + n) n = Fract m n + 1" + by (simp add: rat_number_expand) + +lemma Fract_of_int_quotient: "Fract k l = of_int k / of_int l" + by (simp add: Fract_of_int_eq [symmetric]) + +lemma Fract_number_of_quotient [code post]: + "Fract (number_of k) (number_of l) = number_of k / number_of l" + unfolding Fract_of_int_quotient number_of_is_id number_of_eq .. + +lemma Fract_1_number_of [code post]: + "Fract 1 (number_of k) = 1 / number_of k" + unfolding Fract_of_int_quotient number_of_eq by simp + +subsubsection {* The ordered field of rational numbers *} + +instantiation rat :: linorder +begin + +definition + le_rat_def [code del]: + "q \ r \ contents (\x \ Rep_Rat q. \y \ Rep_Rat r. + {(fst x * snd y) * (snd x * snd y) \ (fst y * snd x) * (snd x * snd y)})" + +lemma le_rat [simp]: + assumes "b \ 0" and "d \ 0" + shows "Fract a b \ Fract c d \ (a * d) * (b * d) \ (c * b) * (b * d)" +proof - + have "(\x y. {(fst x * snd y) * (snd x * snd y) \ (fst y * snd x) * (snd x * snd y)}) + respects2 ratrel" + proof (clarsimp simp add: congruent2_def) + fix a b a' b' c d c' d'::int + assume neq: "b \ 0" "b' \ 0" "d \ 0" "d' \ 0" + assume eq1: "a * b' = a' * b" + assume eq2: "c * d' = c' * d" + + let ?le = "\a b c d. ((a * d) * (b * d) \ (c * b) * (b * d))" + { + fix a b c d x :: int assume x: "x \ 0" + have "?le a b c d = ?le (a * x) (b * x) c d" + proof - + from x have "0 < x * x" by (auto simp add: zero_less_mult_iff) + hence "?le a b c d = + ((a * d) * (b * d) * (x * x) \ (c * b) * (b * d) * (x * x))" + by (simp add: mult_le_cancel_right) + also have "... = ?le (a * x) (b * x) c d" + by (simp add: mult_ac) + finally show ?thesis . + qed + } note le_factor = this + + let ?D = "b * d" and ?D' = "b' * d'" + from neq have D: "?D \ 0" by simp + from neq have "?D' \ 0" by simp + hence "?le a b c d = ?le (a * ?D') (b * ?D') c d" + by (rule le_factor) + also have "... = ((a * b') * ?D * ?D' * d * d' \ (c * d') * ?D * ?D' * b * b')" + by (simp add: mult_ac) + also have "... = ((a' * b) * ?D * ?D' * d * d' \ (c' * d) * ?D * ?D' * b * b')" + by (simp only: eq1 eq2) + also have "... = ?le (a' * ?D) (b' * ?D) c' d'" + by (simp add: mult_ac) + also from D have "... = ?le a' b' c' d'" + by (rule le_factor [symmetric]) + finally show "?le a b c d = ?le a' b' c' d'" . + qed + with assms show ?thesis by (simp add: Fract_def le_rat_def UN_ratrel2) +qed + +definition + less_rat_def [code del]: "z < (w::rat) \ z \ w \ z \ w" + +lemma less_rat [simp]: + assumes "b \ 0" and "d \ 0" + shows "Fract a b < Fract c d \ (a * d) * (b * d) < (c * b) * (b * d)" + using assms by (simp add: less_rat_def eq_rat order_less_le) + +instance proof + fix q r s :: rat + { + assume "q \ r" and "r \ s" + show "q \ s" + proof (insert prems, induct q, induct r, induct s) + fix a b c d e f :: int + assume neq: "b \ 0" "d \ 0" "f \ 0" + assume 1: "Fract a b \ Fract c d" and 2: "Fract c d \ Fract e f" + show "Fract a b \ Fract e f" + proof - + from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f" + by (auto simp add: zero_less_mult_iff linorder_neq_iff) + have "(a * d) * (b * d) * (f * f) \ (c * b) * (b * d) * (f * f)" + proof - + from neq 1 have "(a * d) * (b * d) \ (c * b) * (b * d)" + by simp + with ff show ?thesis by (simp add: mult_le_cancel_right) + qed + also have "... = (c * f) * (d * f) * (b * b)" by algebra + also have "... \ (e * d) * (d * f) * (b * b)" + proof - + from neq 2 have "(c * f) * (d * f) \ (e * d) * (d * f)" + by simp + with bb show ?thesis by (simp add: mult_le_cancel_right) + qed + finally have "(a * f) * (b * f) * (d * d) \ e * b * (b * f) * (d * d)" + by (simp only: mult_ac) + with dd have "(a * f) * (b * f) \ (e * b) * (b * f)" + by (simp add: mult_le_cancel_right) + with neq show ?thesis by simp + qed + qed + next + assume "q \ r" and "r \ q" + show "q = r" + proof (insert prems, induct q, induct r) + fix a b c d :: int + assume neq: "b \ 0" "d \ 0" + assume 1: "Fract a b \ Fract c d" and 2: "Fract c d \ Fract a b" + show "Fract a b = Fract c d" + proof - + from neq 1 have "(a * d) * (b * d) \ (c * b) * (b * d)" + by simp + also have "... \ (a * d) * (b * d)" + proof - + from neq 2 have "(c * b) * (d * b) \ (a * d) * (d * b)" + by simp + thus ?thesis by (simp only: mult_ac) + qed + finally have "(a * d) * (b * d) = (c * b) * (b * d)" . + moreover from neq have "b * d \ 0" by simp + ultimately have "a * d = c * b" by simp + with neq show ?thesis by (simp add: eq_rat) + qed + qed + next + show "q \ q" + by (induct q) simp + show "(q < r) = (q \ r \ \ r \ q)" + by (induct q, induct r) (auto simp add: le_less mult_commute) + show "q \ r \ r \ q" + by (induct q, induct r) + (simp add: mult_commute, rule linorder_linear) + } +qed + +end + +instantiation rat :: "{distrib_lattice, abs_if, sgn_if}" +begin + +definition + abs_rat_def [code del]: "\q\ = (if q < 0 then -q else (q::rat))" + +lemma abs_rat [simp, code]: "\Fract a b\ = Fract \a\ \b\" + by (auto simp add: abs_rat_def zabs_def Zero_rat_def less_rat not_less le_less minus_rat eq_rat zero_compare_simps) + +definition + sgn_rat_def [code del]: "sgn (q::rat) = (if q = 0 then 0 else if 0 < q then 1 else - 1)" + +lemma sgn_rat [simp, code]: "sgn (Fract a b) = of_int (sgn a * sgn b)" + unfolding Fract_of_int_eq + by (auto simp: zsgn_def sgn_rat_def Zero_rat_def eq_rat) + (auto simp: rat_number_collapse not_less le_less zero_less_mult_iff) + +definition + "(inf \ rat \ rat \ rat) = min" + +definition + "(sup \ rat \ rat \ rat) = max" + +instance by intro_classes + (auto simp add: abs_rat_def sgn_rat_def min_max.sup_inf_distrib1 inf_rat_def sup_rat_def) + +end + +instance rat :: ordered_field +proof + fix q r s :: rat + show "q \ r ==> s + q \ s + r" + proof (induct q, induct r, induct s) + fix a b c d e f :: int + assume neq: "b \ 0" "d \ 0" "f \ 0" + assume le: "Fract a b \ Fract c d" + show "Fract e f + Fract a b \ Fract e f + Fract c d" + proof - + let ?F = "f * f" from neq have F: "0 < ?F" + by (auto simp add: zero_less_mult_iff) + from neq le have "(a * d) * (b * d) \ (c * b) * (b * d)" + by simp + with F have "(a * d) * (b * d) * ?F * ?F \ (c * b) * (b * d) * ?F * ?F" + by (simp add: mult_le_cancel_right) + with neq show ?thesis by (simp add: mult_ac int_distrib) + qed + qed + show "q < r ==> 0 < s ==> s * q < s * r" + proof (induct q, induct r, induct s) + fix a b c d e f :: int + assume neq: "b \ 0" "d \ 0" "f \ 0" + assume le: "Fract a b < Fract c d" + assume gt: "0 < Fract e f" + show "Fract e f * Fract a b < Fract e f * Fract c d" + proof - + let ?E = "e * f" and ?F = "f * f" + from neq gt have "0 < ?E" + by (auto simp add: Zero_rat_def order_less_le eq_rat) + moreover from neq have "0 < ?F" + by (auto simp add: zero_less_mult_iff) + moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)" + by simp + ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F" + by (simp add: mult_less_cancel_right) + with neq show ?thesis + by (simp add: mult_ac) + qed + qed +qed auto + +lemma Rat_induct_pos [case_names Fract, induct type: rat]: + assumes step: "\a b. 0 < b \ P (Fract a b)" + shows "P q" +proof (cases q) + have step': "\a b. b < 0 \ P (Fract a b)" + proof - + fix a::int and b::int + assume b: "b < 0" + hence "0 < -b" by simp + hence "P (Fract (-a) (-b))" by (rule step) + thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b]) + qed + case (Fract a b) + thus "P q" by (force simp add: linorder_neq_iff step step') +qed + +lemma zero_less_Fract_iff: + "0 < b ==> (0 < Fract a b) = (0 < a)" +by (simp add: Zero_rat_def order_less_imp_not_eq2 zero_less_mult_iff) + + +subsection {* Arithmetic setup *} + +use "Tools/rat_arith.ML" +declaration {* K rat_arith_setup *} + + +subsection {* Embedding from Rationals to other Fields *} + +class field_char_0 = field + ring_char_0 + +subclass (in ordered_field) field_char_0 .. + +context field_char_0 +begin + +definition of_rat :: "rat \ 'a" where + [code del]: "of_rat q = contents (\(a,b) \ Rep_Rat q. {of_int a / of_int b})" + +end + +lemma of_rat_congruent: + "(\(a, b). {of_int a / of_int b :: 'a::field_char_0}) respects ratrel" +apply (rule congruent.intro) +apply (clarsimp simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq) +apply (simp only: of_int_mult [symmetric]) +done + +lemma of_rat_rat: "b \ 0 \ of_rat (Fract a b) = of_int a / of_int b" + unfolding Fract_def of_rat_def by (simp add: UN_ratrel of_rat_congruent) + +lemma of_rat_0 [simp]: "of_rat 0 = 0" +by (simp add: Zero_rat_def of_rat_rat) + +lemma of_rat_1 [simp]: "of_rat 1 = 1" +by (simp add: One_rat_def of_rat_rat) + +lemma of_rat_add: "of_rat (a + b) = of_rat a + of_rat b" +by (induct a, induct b, simp add: of_rat_rat add_frac_eq) + +lemma of_rat_minus: "of_rat (- a) = - of_rat a" +by (induct a, simp add: of_rat_rat) + +lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b" +by (simp only: diff_minus of_rat_add of_rat_minus) + +lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b" +apply (induct a, induct b, simp add: of_rat_rat) +apply (simp add: divide_inverse nonzero_inverse_mult_distrib mult_ac) +done + +lemma nonzero_of_rat_inverse: + "a \ 0 \ of_rat (inverse a) = inverse (of_rat a)" +apply (rule inverse_unique [symmetric]) +apply (simp add: of_rat_mult [symmetric]) +done + +lemma of_rat_inverse: + "(of_rat (inverse a)::'a::{field_char_0,division_by_zero}) = + inverse (of_rat a)" +by (cases "a = 0", simp_all add: nonzero_of_rat_inverse) + +lemma nonzero_of_rat_divide: + "b \ 0 \ of_rat (a / b) = of_rat a / of_rat b" +by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse) + +lemma of_rat_divide: + "(of_rat (a / b)::'a::{field_char_0,division_by_zero}) + = of_rat a / of_rat b" +by (cases "b = 0") (simp_all add: nonzero_of_rat_divide) + +lemma of_rat_power: + "(of_rat (a ^ n)::'a::{field_char_0,recpower}) = of_rat a ^ n" +by (induct n) (simp_all add: of_rat_mult power_Suc) + +lemma of_rat_eq_iff [simp]: "(of_rat a = of_rat b) = (a = b)" +apply (induct a, induct b) +apply (simp add: of_rat_rat eq_rat) +apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq) +apply (simp only: of_int_mult [symmetric] of_int_eq_iff) +done + +lemma of_rat_less: + "(of_rat r :: 'a::ordered_field) < of_rat s \ r < s" +proof (induct r, induct s) + fix a b c d :: int + assume not_zero: "b > 0" "d > 0" + then have "b * d > 0" by (rule mult_pos_pos) + have of_int_divide_less_eq: + "(of_int a :: 'a) / of_int b < of_int c / of_int d + \ (of_int a :: 'a) * of_int d < of_int c * of_int b" + using not_zero by (simp add: pos_less_divide_eq pos_divide_less_eq) + show "(of_rat (Fract a b) :: 'a::ordered_field) < of_rat (Fract c d) + \ Fract a b < Fract c d" + using not_zero `b * d > 0` + by (simp add: of_rat_rat of_int_divide_less_eq of_int_mult [symmetric] del: of_int_mult) + (auto intro: mult_strict_right_mono mult_right_less_imp_less) +qed + +lemma of_rat_less_eq: + "(of_rat r :: 'a::ordered_field) \ of_rat s \ r \ s" + unfolding le_less by (auto simp add: of_rat_less) + +lemmas of_rat_eq_0_iff [simp] = of_rat_eq_iff [of _ 0, simplified] + +lemma of_rat_eq_id [simp]: "of_rat = id" +proof + fix a + show "of_rat a = id a" + by (induct a) + (simp add: of_rat_rat Fract_of_int_eq [symmetric]) +qed + +text{*Collapse nested embeddings*} +lemma of_rat_of_nat_eq [simp]: "of_rat (of_nat n) = of_nat n" +by (induct n) (simp_all add: of_rat_add) + +lemma of_rat_of_int_eq [simp]: "of_rat (of_int z) = of_int z" +by (cases z rule: int_diff_cases) (simp add: of_rat_diff) + +lemma of_rat_number_of_eq [simp]: + "of_rat (number_of w) = (number_of w :: 'a::{number_ring,field_char_0})" +by (simp add: number_of_eq) + +lemmas zero_rat = Zero_rat_def +lemmas one_rat = One_rat_def + +abbreviation + rat_of_nat :: "nat \ rat" +where + "rat_of_nat \ of_nat" + +abbreviation + rat_of_int :: "int \ rat" +where + "rat_of_int \ of_int" + +subsection {* The Set of Rational Numbers *} + +context field_char_0 +begin + +definition + Rats :: "'a set" where + [code del]: "Rats = range of_rat" + +notation (xsymbols) + Rats ("\") + +end + +lemma Rats_of_rat [simp]: "of_rat r \ Rats" +by (simp add: Rats_def) + +lemma Rats_of_int [simp]: "of_int z \ Rats" +by (subst of_rat_of_int_eq [symmetric], rule Rats_of_rat) + +lemma Rats_of_nat [simp]: "of_nat n \ Rats" +by (subst of_rat_of_nat_eq [symmetric], rule Rats_of_rat) + +lemma Rats_number_of [simp]: + "(number_of w::'a::{number_ring,field_char_0}) \ Rats" +by (subst of_rat_number_of_eq [symmetric], rule Rats_of_rat) + +lemma Rats_0 [simp]: "0 \ Rats" +apply (unfold Rats_def) +apply (rule range_eqI) +apply (rule of_rat_0 [symmetric]) +done + +lemma Rats_1 [simp]: "1 \ Rats" +apply (unfold Rats_def) +apply (rule range_eqI) +apply (rule of_rat_1 [symmetric]) +done + +lemma Rats_add [simp]: "\a \ Rats; b \ Rats\ \ a + b \ Rats" +apply (auto simp add: Rats_def) +apply (rule range_eqI) +apply (rule of_rat_add [symmetric]) +done + +lemma Rats_minus [simp]: "a \ Rats \ - a \ Rats" +apply (auto simp add: Rats_def) +apply (rule range_eqI) +apply (rule of_rat_minus [symmetric]) +done + +lemma Rats_diff [simp]: "\a \ Rats; b \ Rats\ \ a - b \ Rats" +apply (auto simp add: Rats_def) +apply (rule range_eqI) +apply (rule of_rat_diff [symmetric]) +done + +lemma Rats_mult [simp]: "\a \ Rats; b \ Rats\ \ a * b \ Rats" +apply (auto simp add: Rats_def) +apply (rule range_eqI) +apply (rule of_rat_mult [symmetric]) +done + +lemma nonzero_Rats_inverse: + fixes a :: "'a::field_char_0" + shows "\a \ Rats; a \ 0\ \ inverse a \ Rats" +apply (auto simp add: Rats_def) +apply (rule range_eqI) +apply (erule nonzero_of_rat_inverse [symmetric]) +done + +lemma Rats_inverse [simp]: + fixes a :: "'a::{field_char_0,division_by_zero}" + shows "a \ Rats \ inverse a \ Rats" +apply (auto simp add: Rats_def) +apply (rule range_eqI) +apply (rule of_rat_inverse [symmetric]) +done + +lemma nonzero_Rats_divide: + fixes a b :: "'a::field_char_0" + shows "\a \ Rats; b \ Rats; b \ 0\ \ a / b \ Rats" +apply (auto simp add: Rats_def) +apply (rule range_eqI) +apply (erule nonzero_of_rat_divide [symmetric]) +done + +lemma Rats_divide [simp]: + fixes a b :: "'a::{field_char_0,division_by_zero}" + shows "\a \ Rats; b \ Rats\ \ a / b \ Rats" +apply (auto simp add: Rats_def) +apply (rule range_eqI) +apply (rule of_rat_divide [symmetric]) +done + +lemma Rats_power [simp]: + fixes a :: "'a::{field_char_0,recpower}" + shows "a \ Rats \ a ^ n \ Rats" +apply (auto simp add: Rats_def) +apply (rule range_eqI) +apply (rule of_rat_power [symmetric]) +done + +lemma Rats_cases [cases set: Rats]: + assumes "q \ \" + obtains (of_rat) r where "q = of_rat r" + unfolding Rats_def +proof - + from `q \ \` have "q \ range of_rat" unfolding Rats_def . + then obtain r where "q = of_rat r" .. + then show thesis .. +qed + +lemma Rats_induct [case_names of_rat, induct set: Rats]: + "q \ \ \ (\r. P (of_rat r)) \ P q" + by (rule Rats_cases) auto + + +subsection {* The Rationals are Countably Infinite *} + +definition nat_to_rat_surj :: "nat \ rat" where +"nat_to_rat_surj n = (let (a,b) = nat_to_nat2 n + in Fract (nat_to_int_bij a) (nat_to_int_bij b))" + +lemma surj_nat_to_rat_surj: "surj nat_to_rat_surj" +unfolding surj_def +proof + fix r::rat + show "\n. r = nat_to_rat_surj n" + proof(cases r) + fix i j assume [simp]: "r = Fract i j" and "j \ 0" + have "r = (let m = inv nat_to_int_bij i; n = inv nat_to_int_bij j + in nat_to_rat_surj(nat2_to_nat (m,n)))" + using nat2_to_nat_inj surj_f_inv_f[OF surj_nat_to_int_bij] + by(simp add:Let_def nat_to_rat_surj_def nat_to_nat2_def) + thus "\n. r = nat_to_rat_surj n" by(auto simp:Let_def) + qed +qed + +lemma Rats_eq_range_nat_to_rat_surj: "\ = range nat_to_rat_surj" +by (simp add: Rats_def surj_nat_to_rat_surj surj_range) + +context field_char_0 +begin + +lemma Rats_eq_range_of_rat_o_nat_to_rat_surj: + "\ = range (of_rat o nat_to_rat_surj)" +using surj_nat_to_rat_surj +by (auto simp: Rats_def image_def surj_def) + (blast intro: arg_cong[where f = of_rat]) + +lemma surj_of_rat_nat_to_rat_surj: + "r\\ \ \n. r = of_rat(nat_to_rat_surj n)" +by(simp add: Rats_eq_range_of_rat_o_nat_to_rat_surj image_def) + +end + + +subsection {* Implementation of rational numbers as pairs of integers *} + +lemma Fract_norm: "Fract (a div zgcd a b) (b div zgcd a b) = Fract a b" +proof (cases "a = 0 \ b = 0") + case True then show ?thesis by (auto simp add: eq_rat) +next + let ?c = "zgcd a b" + case False then have "a \ 0" and "b \ 0" by auto + then have "?c \ 0" by simp + then have "Fract ?c ?c = Fract 1 1" by (simp add: eq_rat) + moreover have "Fract (a div ?c * ?c + a mod ?c) (b div ?c * ?c + b mod ?c) = Fract a b" + by (simp add: semiring_div_class.mod_div_equality) + moreover have "a mod ?c = 0" by (simp add: dvd_eq_mod_eq_0 [symmetric]) + moreover have "b mod ?c = 0" by (simp add: dvd_eq_mod_eq_0 [symmetric]) + ultimately show ?thesis + by (simp add: mult_rat [symmetric]) +qed + +definition Fract_norm :: "int \ int \ rat" where + [simp, code del]: "Fract_norm a b = Fract a b" + +lemma [code]: "Fract_norm a b = (if a = 0 \ b = 0 then 0 else let c = zgcd a b in + if b > 0 then Fract (a div c) (b div c) else Fract (- (a div c)) (- (b div c)))" + by (simp add: eq_rat Zero_rat_def Let_def Fract_norm) + +lemma [code]: + "of_rat (Fract a b) = (if b \ 0 then of_int a / of_int b else 0)" + by (cases "b = 0") (simp_all add: rat_number_collapse of_rat_rat) + +instantiation rat :: eq +begin + +definition [code del]: "eq_class.eq (a\rat) b \ a - b = 0" + +instance by default (simp add: eq_rat_def) + +lemma rat_eq_code [code]: + "eq_class.eq (Fract a b) (Fract c d) \ (if b = 0 + then c = 0 \ d = 0 + else if d = 0 + then a = 0 \ b = 0 + else a * d = b * c)" + by (auto simp add: eq eq_rat) + +lemma rat_eq_refl [code nbe]: + "eq_class.eq (r::rat) r \ True" + by (rule HOL.eq_refl) + +end + +lemma le_rat': + assumes "b \ 0" + and "d \ 0" + shows "Fract a b \ Fract c d \ a * \d\ * sgn b \ c * \b\ * sgn d" +proof - + have abs_sgn: "\k::int. \k\ = k * sgn k" unfolding abs_if sgn_if by simp + have "a * d * (b * d) \ c * b * (b * d) \ a * d * (sgn b * sgn d) \ c * b * (sgn b * sgn d)" + proof (cases "b * d > 0") + case True + moreover from True have "sgn b * sgn d = 1" + by (simp add: sgn_times [symmetric] sgn_1_pos) + ultimately show ?thesis by (simp add: mult_le_cancel_right) + next + case False with assms have "b * d < 0" by (simp add: less_le) + moreover from this have "sgn b * sgn d = - 1" + by (simp only: sgn_times [symmetric] sgn_1_neg) + ultimately show ?thesis by (simp add: mult_le_cancel_right) + qed + also have "\ \ a * \d\ * sgn b \ c * \b\ * sgn d" + by (simp add: abs_sgn mult_ac) + finally show ?thesis using assms by simp +qed + +lemma less_rat': + assumes "b \ 0" + and "d \ 0" + shows "Fract a b < Fract c d \ a * \d\ * sgn b < c * \b\ * sgn d" +proof - + have abs_sgn: "\k::int. \k\ = k * sgn k" unfolding abs_if sgn_if by simp + have "a * d * (b * d) < c * b * (b * d) \ a * d * (sgn b * sgn d) < c * b * (sgn b * sgn d)" + proof (cases "b * d > 0") + case True + moreover from True have "sgn b * sgn d = 1" + by (simp add: sgn_times [symmetric] sgn_1_pos) + ultimately show ?thesis by (simp add: mult_less_cancel_right) + next + case False with assms have "b * d < 0" by (simp add: less_le) + moreover from this have "sgn b * sgn d = - 1" + by (simp only: sgn_times [symmetric] sgn_1_neg) + ultimately show ?thesis by (simp add: mult_less_cancel_right) + qed + also have "\ \ a * \d\ * sgn b < c * \b\ * sgn d" + by (simp add: abs_sgn mult_ac) + finally show ?thesis using assms by simp +qed + +lemma rat_less_eq_code [code]: + "Fract a b \ Fract c d \ (if b = 0 + then sgn c * sgn d \ 0 + else if d = 0 + then sgn a * sgn b \ 0 + else a * \d\ * sgn b \ c * \b\ * sgn d)" +by (auto simp add: sgn_times mult_le_0_iff zero_le_mult_iff le_rat' eq_rat simp del: le_rat) + (auto simp add: sgn_times sgn_0_0 le_less sgn_1_pos [symmetric] sgn_1_neg [symmetric]) + +lemma rat_le_eq_code [code]: + "Fract a b < Fract c d \ (if b = 0 + then sgn c * sgn d > 0 + else if d = 0 + then sgn a * sgn b < 0 + else a * \d\ * sgn b < c * \b\ * sgn d)" +by (auto simp add: sgn_times mult_less_0_iff zero_less_mult_iff less_rat' eq_rat simp del: less_rat) + (auto simp add: sgn_times sgn_0_0 sgn_1_pos [symmetric] sgn_1_neg [symmetric], + auto simp add: sgn_1_pos) + +lemma rat_plus_code [code]: + "Fract a b + Fract c d = (if b = 0 + then Fract c d + else if d = 0 + then Fract a b + else Fract_norm (a * d + c * b) (b * d))" + by (simp add: eq_rat, simp add: Zero_rat_def) + +lemma rat_times_code [code]: + "Fract a b * Fract c d = Fract_norm (a * c) (b * d)" + by simp + +lemma rat_minus_code [code]: + "Fract a b - Fract c d = (if b = 0 + then Fract (- c) d + else if d = 0 + then Fract a b + else Fract_norm (a * d - c * b) (b * d))" + by (simp add: eq_rat, simp add: Zero_rat_def) + +lemma rat_inverse_code [code]: + "inverse (Fract a b) = (if b = 0 then Fract 1 0 + else if a < 0 then Fract (- b) (- a) + else Fract b a)" + by (simp add: eq_rat) + +lemma rat_divide_code [code]: + "Fract a b / Fract c d = Fract_norm (a * d) (b * c)" + by simp + +hide (open) const Fract_norm + +text {* Setup for SML code generator *} + +types_code + rat ("(int */ int)") +attach (term_of) {* +fun term_of_rat (p, q) = + let + val rT = Type ("Rational.rat", []) + in + if q = 1 orelse p = 0 then HOLogic.mk_number rT p + else @{term "op / \ rat \ rat \ rat"} $ + HOLogic.mk_number rT p $ HOLogic.mk_number rT q + end; +*} +attach (test) {* +fun gen_rat i = + let + val p = random_range 0 i; + val q = random_range 1 (i + 1); + val g = Integer.gcd p q; + val p' = p div g; + val q' = q div g; + val r = (if one_of [true, false] then p' else ~ p', + if p' = 0 then 0 else q') + in + (r, fn () => term_of_rat r) + end; +*} + +consts_code + Fract ("(_,/ _)") + +consts_code + "of_int :: int \ rat" ("\rat'_of'_int") +attach {* +fun rat_of_int 0 = (0, 0) + | rat_of_int i = (i, 1); +*} + +end \ No newline at end of file diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Real.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,5 @@ +theory Real +imports ContNotDenum "~~/src/HOL/Real/RealVector" +begin + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/ContNotDenum.thy --- a/src/HOL/Real/ContNotDenum.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,580 +0,0 @@ -(* Title : HOL/Real/ContNonDenum - ID : $Id$ - Author : Benjamin Porter, Monash University, NICTA, 2005 -*) - -header {* Non-denumerability of the Continuum. *} - -theory ContNotDenum -imports RComplete "../Hilbert_Choice" -begin - -subsection {* Abstract *} - -text {* The following document presents a proof that the Continuum is -uncountable. It is formalised in the Isabelle/Isar theorem proving -system. - -{\em Theorem:} The Continuum @{text "\"} is not denumerable. In other -words, there does not exist a function f:@{text "\\\"} such that f is -surjective. - -{\em Outline:} An elegant informal proof of this result uses Cantor's -Diagonalisation argument. The proof presented here is not this -one. First we formalise some properties of closed intervals, then we -prove the Nested Interval Property. This property relies on the -completeness of the Real numbers and is the foundation for our -argument. Informally it states that an intersection of countable -closed intervals (where each successive interval is a subset of the -last) is non-empty. We then assume a surjective function f:@{text -"\\\"} exists and find a real x such that x is not in the range of f -by generating a sequence of closed intervals then using the NIP. *} - -subsection {* Closed Intervals *} - -text {* This section formalises some properties of closed intervals. *} - -subsubsection {* Definition *} - -definition - closed_int :: "real \ real \ real set" where - "closed_int x y = {z. x \ z \ z \ y}" - -subsubsection {* Properties *} - -lemma closed_int_subset: - assumes xy: "x1 \ x0" "y1 \ y0" - shows "closed_int x1 y1 \ closed_int x0 y0" -proof - - { - fix x::real - assume "x \ closed_int x1 y1" - hence "x \ x1 \ x \ y1" by (simp add: closed_int_def) - with xy have "x \ x0 \ x \ y0" by auto - hence "x \ closed_int x0 y0" by (simp add: closed_int_def) - } - thus ?thesis by auto -qed - -lemma closed_int_least: - assumes a: "a \ b" - shows "a \ closed_int a b \ (\x \ closed_int a b. a \ x)" -proof - from a have "a\{x. a\x \ x\b}" by simp - thus "a \ closed_int a b" by (unfold closed_int_def) -next - have "\x\{x. a\x \ x\b}. a\x" by simp - thus "\x \ closed_int a b. a \ x" by (unfold closed_int_def) -qed - -lemma closed_int_most: - assumes a: "a \ b" - shows "b \ closed_int a b \ (\x \ closed_int a b. x \ b)" -proof - from a have "b\{x. a\x \ x\b}" by simp - thus "b \ closed_int a b" by (unfold closed_int_def) -next - have "\x\{x. a\x \ x\b}. x\b" by simp - thus "\x \ closed_int a b. x\b" by (unfold closed_int_def) -qed - -lemma closed_not_empty: - shows "a \ b \ \x. x \ closed_int a b" - by (auto dest: closed_int_least) - -lemma closed_mem: - assumes "a \ c" and "c \ b" - shows "c \ closed_int a b" - using assms unfolding closed_int_def by auto - -lemma closed_subset: - assumes ac: "a \ b" "c \ d" - assumes closed: "closed_int a b \ closed_int c d" - shows "b \ c" -proof - - from closed have "\x\closed_int a b. x\closed_int c d" by auto - hence "\x. a\x \ x\b \ c\x \ x\d" by (unfold closed_int_def, auto) - with ac have "c\b \ b\d" by simp - thus ?thesis by auto -qed - - -subsection {* Nested Interval Property *} - -theorem NIP: - fixes f::"nat \ real set" - assumes subset: "\n. f (Suc n) \ f n" - and closed: "\n. \a b. f n = closed_int a b \ a \ b" - shows "(\n. f n) \ {}" -proof - - let ?g = "\n. (SOME c. c\(f n) \ (\x\(f n). c \ x))" - have ne: "\n. \x. x\(f n)" - proof - fix n - from closed have "\a b. f n = closed_int a b \ a \ b" by simp - then obtain a and b where fn: "f n = closed_int a b \ a \ b" by auto - hence "a \ b" .. - with closed_not_empty have "\x. x\closed_int a b" by simp - with fn show "\x. x\(f n)" by simp - qed - - have gdef: "\n. (?g n)\(f n) \ (\x\(f n). (?g n)\x)" - proof - fix n - from closed have "\a b. f n = closed_int a b \ a \ b" .. - then obtain a and b where ff: "f n = closed_int a b" and "a \ b" by auto - hence "a \ b" by simp - hence "a\closed_int a b \ (\x\closed_int a b. a \ x)" by (rule closed_int_least) - with ff have "a\(f n) \ (\x\(f n). a \ x)" by simp - hence "\c. c\(f n) \ (\x\(f n). c \ x)" .. - thus "(?g n)\(f n) \ (\x\(f n). (?g n)\x)" by (rule someI_ex) - qed - - -- "A denotes the set of all left-most points of all the intervals ..." - moreover obtain A where Adef: "A = ?g ` \" by simp - ultimately have "\x. x\A" - proof - - have "(0::nat) \ \" by simp - moreover have "?g 0 = ?g 0" by simp - ultimately have "?g 0 \ ?g ` \" by (rule rev_image_eqI) - with Adef have "?g 0 \ A" by simp - thus ?thesis .. - qed - - -- "Now show that A is bounded above ..." - moreover have "\y. isUb (UNIV::real set) A y" - proof - - { - fix n - from ne have ex: "\x. x\(f n)" .. - from gdef have "(?g n)\(f n) \ (\x\(f n). (?g n)\x)" by simp - moreover - from closed have "\a b. f n = closed_int a b \ a \ b" .. - then obtain a and b where "f n = closed_int a b \ a \ b" by auto - hence "b\(f n) \ (\x\(f n). x \ b)" using closed_int_most by blast - ultimately have "\x\(f n). (?g n) \ b" by simp - with ex have "(?g n) \ b" by auto - hence "\b. (?g n) \ b" by auto - } - hence aux: "\n. \b. (?g n) \ b" .. - - have fs: "\n::nat. f n \ f 0" - proof (rule allI, induct_tac n) - show "f 0 \ f 0" by simp - next - fix n - assume "f n \ f 0" - moreover from subset have "f (Suc n) \ f n" .. - ultimately show "f (Suc n) \ f 0" by simp - qed - have "\n. (?g n)\(f 0)" - proof - fix n - from gdef have "(?g n)\(f n) \ (\x\(f n). (?g n)\x)" by simp - hence "?g n \ f n" .. - with fs show "?g n \ f 0" by auto - qed - moreover from closed - obtain a and b where "f 0 = closed_int a b" and alb: "a \ b" by blast - ultimately have "\n. ?g n \ closed_int a b" by auto - with alb have "\n. ?g n \ b" using closed_int_most by blast - with Adef have "\y\A. y\b" by auto - hence "A *<= b" by (unfold setle_def) - moreover have "b \ (UNIV::real set)" by simp - ultimately have "A *<= b \ b \ (UNIV::real set)" by simp - hence "isUb (UNIV::real set) A b" by (unfold isUb_def) - thus ?thesis by auto - qed - -- "by the Axiom Of Completeness, A has a least upper bound ..." - ultimately have "\t. isLub UNIV A t" by (rule reals_complete) - - -- "denote this least upper bound as t ..." - then obtain t where tdef: "isLub UNIV A t" .. - - -- "and finally show that this least upper bound is in all the intervals..." - have "\n. t \ f n" - proof - fix n::nat - from closed obtain a and b where - int: "f n = closed_int a b" and alb: "a \ b" by blast - - have "t \ a" - proof - - have "a \ A" - proof - - (* by construction *) - from alb int have ain: "a\f n \ (\x\f n. a \ x)" - using closed_int_least by blast - moreover have "\e. e\f n \ (\x\f n. e \ x) \ e = a" - proof clarsimp - fix e - assume ein: "e \ f n" and lt: "\x\f n. e \ x" - from lt ain have aux: "\x\f n. a \ x \ e \ x" by auto - - from ein aux have "a \ e \ e \ e" by auto - moreover from ain aux have "a \ a \ e \ a" by auto - ultimately show "e = a" by simp - qed - hence "\e. e\f n \ (\x\f n. e \ x) \ e = a" by simp - ultimately have "(?g n) = a" by (rule some_equality) - moreover - { - have "n = of_nat n" by simp - moreover have "of_nat n \ \" by simp - ultimately have "n \ \" - apply - - apply (subst(asm) eq_sym_conv) - apply (erule subst) - . - } - with Adef have "(?g n) \ A" by auto - ultimately show ?thesis by simp - qed - with tdef show "a \ t" by (rule isLubD2) - qed - moreover have "t \ b" - proof - - have "isUb UNIV A b" - proof - - { - from alb int have - ain: "b\f n \ (\x\f n. x \ b)" using closed_int_most by blast - - have subsetd: "\m. \n. f (n + m) \ f n" - proof (rule allI, induct_tac m) - show "\n. f (n + 0) \ f n" by simp - next - fix m n - assume pp: "\p. f (p + n) \ f p" - { - fix p - from pp have "f (p + n) \ f p" by simp - moreover from subset have "f (Suc (p + n)) \ f (p + n)" by auto - hence "f (p + (Suc n)) \ f (p + n)" by simp - ultimately have "f (p + (Suc n)) \ f p" by simp - } - thus "\p. f (p + Suc n) \ f p" .. - qed - have subsetm: "\\ \. \ \ \ \ (f \) \ (f \)" - proof ((rule allI)+, rule impI) - fix \::nat and \::nat - assume "\ \ \" - hence "\k. \ = \ + k" by (simp only: le_iff_add) - then obtain k where "\ = \ + k" .. - moreover - from subsetd have "f (\ + k) \ f \" by simp - ultimately show "f \ \ f \" by auto - qed - - fix m - { - assume "m \ n" - with subsetm have "f m \ f n" by simp - with ain have "\x\f m. x \ b" by auto - moreover - from gdef have "?g m \ f m \ (\x\f m. ?g m \ x)" by simp - ultimately have "?g m \ b" by auto - } - moreover - { - assume "\(m \ n)" - hence "m < n" by simp - with subsetm have sub: "(f n) \ (f m)" by simp - from closed obtain ma and mb where - "f m = closed_int ma mb \ ma \ mb" by blast - hence one: "ma \ mb" and fm: "f m = closed_int ma mb" by auto - from one alb sub fm int have "ma \ b" using closed_subset by blast - moreover have "(?g m) = ma" - proof - - from gdef have "?g m \ f m \ (\x\f m. ?g m \ x)" .. - moreover from one have - "ma \ closed_int ma mb \ (\x\closed_int ma mb. ma \ x)" - by (rule closed_int_least) - with fm have "ma\f m \ (\x\f m. ma \ x)" by simp - ultimately have "ma \ ?g m \ ?g m \ ma" by auto - thus "?g m = ma" by auto - qed - ultimately have "?g m \ b" by simp - } - ultimately have "?g m \ b" by (rule case_split) - } - with Adef have "\y\A. y\b" by auto - hence "A *<= b" by (unfold setle_def) - moreover have "b \ (UNIV::real set)" by simp - ultimately have "A *<= b \ b \ (UNIV::real set)" by simp - thus "isUb (UNIV::real set) A b" by (unfold isUb_def) - qed - with tdef show "t \ b" by (rule isLub_le_isUb) - qed - ultimately have "t \ closed_int a b" by (rule closed_mem) - with int show "t \ f n" by simp - qed - hence "t \ (\n. f n)" by auto - thus ?thesis by auto -qed - -subsection {* Generating the intervals *} - -subsubsection {* Existence of non-singleton closed intervals *} - -text {* This lemma asserts that given any non-singleton closed -interval (a,b) and any element c, there exists a closed interval that -is a subset of (a,b) and that does not contain c and is a -non-singleton itself. *} - -lemma closed_subset_ex: - fixes c::real - assumes alb: "a < b" - shows - "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ c \ (closed_int ka kb)" -proof - - { - assume clb: "c < b" - { - assume cla: "c < a" - from alb cla clb have "c \ closed_int a b" by (unfold closed_int_def, auto) - with alb have - "a < b \ closed_int a b \ closed_int a b \ c \ closed_int a b" - by auto - hence - "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ c \ (closed_int ka kb)" - by auto - } - moreover - { - assume ncla: "\(c < a)" - with clb have cdef: "a \ c \ c < b" by simp - obtain ka where kadef: "ka = (c + b)/2" by blast - - from kadef clb have kalb: "ka < b" by auto - moreover from kadef cdef have kagc: "ka > c" by simp - ultimately have "c\(closed_int ka b)" by (unfold closed_int_def, auto) - moreover from cdef kagc have "ka \ a" by simp - hence "closed_int ka b \ closed_int a b" by (unfold closed_int_def, auto) - ultimately have - "ka < b \ closed_int ka b \ closed_int a b \ c \ closed_int ka b" - using kalb by auto - hence - "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ c \ (closed_int ka kb)" - by auto - - } - ultimately have - "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ c \ (closed_int ka kb)" - by (rule case_split) - } - moreover - { - assume "\ (c < b)" - hence cgeb: "c \ b" by simp - - obtain kb where kbdef: "kb = (a + b)/2" by blast - with alb have kblb: "kb < b" by auto - with kbdef cgeb have "a < kb \ kb < c" by auto - moreover hence "c \ (closed_int a kb)" by (unfold closed_int_def, auto) - moreover from kblb have - "closed_int a kb \ closed_int a b" by (unfold closed_int_def, auto) - ultimately have - "a < kb \ closed_int a kb \ closed_int a b \ c\closed_int a kb" - by simp - hence - "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ c \ (closed_int ka kb)" - by auto - } - ultimately show ?thesis by (rule case_split) -qed - -subsection {* newInt: Interval generation *} - -text {* Given a function f:@{text "\\\"}, newInt (Suc n) f returns a -closed interval such that @{text "newInt (Suc n) f \ newInt n f"} and -does not contain @{text "f (Suc n)"}. With the base case defined such -that @{text "(f 0)\newInt 0 f"}. *} - -subsubsection {* Definition *} - -primrec newInt :: "nat \ (nat \ real) \ (real set)" where - "newInt 0 f = closed_int (f 0 + 1) (f 0 + 2)" - | "newInt (Suc n) f = - (SOME e. (\e1 e2. - e1 < e2 \ - e = closed_int e1 e2 \ - e \ (newInt n f) \ - (f (Suc n)) \ e) - )" - -declare newInt.simps [code del] - -subsubsection {* Properties *} - -text {* We now show that every application of newInt returns an -appropriate interval. *} - -lemma newInt_ex: - "\a b. a < b \ - newInt (Suc n) f = closed_int a b \ - newInt (Suc n) f \ newInt n f \ - f (Suc n) \ newInt (Suc n) f" -proof (induct n) - case 0 - - let ?e = "SOME e. \e1 e2. - e1 < e2 \ - e = closed_int e1 e2 \ - e \ closed_int (f 0 + 1) (f 0 + 2) \ - f (Suc 0) \ e" - - have "newInt (Suc 0) f = ?e" by auto - moreover - have "f 0 + 1 < f 0 + 2" by simp - with closed_subset_ex have - "\ka kb. ka < kb \ closed_int ka kb \ closed_int (f 0 + 1) (f 0 + 2) \ - f (Suc 0) \ (closed_int ka kb)" . - hence - "\e. \ka kb. ka < kb \ e = closed_int ka kb \ - e \ closed_int (f 0 + 1) (f 0 + 2) \ f (Suc 0) \ e" by simp - hence - "\ka kb. ka < kb \ ?e = closed_int ka kb \ - ?e \ closed_int (f 0 + 1) (f 0 + 2) \ f (Suc 0) \ ?e" - by (rule someI_ex) - ultimately have "\e1 e2. e1 < e2 \ - newInt (Suc 0) f = closed_int e1 e2 \ - newInt (Suc 0) f \ closed_int (f 0 + 1) (f 0 + 2) \ - f (Suc 0) \ newInt (Suc 0) f" by simp - thus - "\a b. a < b \ newInt (Suc 0) f = closed_int a b \ - newInt (Suc 0) f \ newInt 0 f \ f (Suc 0) \ newInt (Suc 0) f" - by simp -next - case (Suc n) - hence "\a b. - a < b \ - newInt (Suc n) f = closed_int a b \ - newInt (Suc n) f \ newInt n f \ - f (Suc n) \ newInt (Suc n) f" by simp - then obtain a and b where ab: "a < b \ - newInt (Suc n) f = closed_int a b \ - newInt (Suc n) f \ newInt n f \ - f (Suc n) \ newInt (Suc n) f" by auto - hence cab: "closed_int a b = newInt (Suc n) f" by simp - - let ?e = "SOME e. \e1 e2. - e1 < e2 \ - e = closed_int e1 e2 \ - e \ closed_int a b \ - f (Suc (Suc n)) \ e" - from cab have ni: "newInt (Suc (Suc n)) f = ?e" by auto - - from ab have "a < b" by simp - with closed_subset_ex have - "\ka kb. ka < kb \ closed_int ka kb \ closed_int a b \ - f (Suc (Suc n)) \ closed_int ka kb" . - hence - "\e. \ka kb. ka < kb \ e = closed_int ka kb \ - closed_int ka kb \ closed_int a b \ f (Suc (Suc n)) \ closed_int ka kb" - by simp - hence - "\e. \ka kb. ka < kb \ e = closed_int ka kb \ - e \ closed_int a b \ f (Suc (Suc n)) \ e" by simp - hence - "\ka kb. ka < kb \ ?e = closed_int ka kb \ - ?e \ closed_int a b \ f (Suc (Suc n)) \ ?e" by (rule someI_ex) - with ab ni show - "\ka kb. ka < kb \ - newInt (Suc (Suc n)) f = closed_int ka kb \ - newInt (Suc (Suc n)) f \ newInt (Suc n) f \ - f (Suc (Suc n)) \ newInt (Suc (Suc n)) f" by auto -qed - -lemma newInt_subset: - "newInt (Suc n) f \ newInt n f" - using newInt_ex by auto - - -text {* Another fundamental property is that no element in the range -of f is in the intersection of all closed intervals generated by -newInt. *} - -lemma newInt_inter: - "\n. f n \ (\n. newInt n f)" -proof - fix n::nat - { - assume n0: "n = 0" - moreover have "newInt 0 f = closed_int (f 0 + 1) (f 0 + 2)" by simp - ultimately have "f n \ newInt n f" by (unfold closed_int_def, simp) - } - moreover - { - assume "\ n = 0" - hence "n > 0" by simp - then obtain m where ndef: "n = Suc m" by (auto simp add: gr0_conv_Suc) - - from newInt_ex have - "\a b. a < b \ (newInt (Suc m) f) = closed_int a b \ - newInt (Suc m) f \ newInt m f \ f (Suc m) \ newInt (Suc m) f" . - then have "f (Suc m) \ newInt (Suc m) f" by auto - with ndef have "f n \ newInt n f" by simp - } - ultimately have "f n \ newInt n f" by (rule case_split) - thus "f n \ (\n. newInt n f)" by auto -qed - - -lemma newInt_notempty: - "(\n. newInt n f) \ {}" -proof - - let ?g = "\n. newInt n f" - have "\n. ?g (Suc n) \ ?g n" - proof - fix n - show "?g (Suc n) \ ?g n" by (rule newInt_subset) - qed - moreover have "\n. \a b. ?g n = closed_int a b \ a \ b" - proof - fix n::nat - { - assume "n = 0" - then have - "?g n = closed_int (f 0 + 1) (f 0 + 2) \ (f 0 + 1 \ f 0 + 2)" - by simp - hence "\a b. ?g n = closed_int a b \ a \ b" by blast - } - moreover - { - assume "\ n = 0" - then have "n > 0" by simp - then obtain m where nd: "n = Suc m" by (auto simp add: gr0_conv_Suc) - - have - "\a b. a < b \ (newInt (Suc m) f) = closed_int a b \ - (newInt (Suc m) f) \ (newInt m f) \ (f (Suc m)) \ (newInt (Suc m) f)" - by (rule newInt_ex) - then obtain a and b where - "a < b \ (newInt (Suc m) f) = closed_int a b" by auto - with nd have "?g n = closed_int a b \ a \ b" by auto - hence "\a b. ?g n = closed_int a b \ a \ b" by blast - } - ultimately show "\a b. ?g n = closed_int a b \ a \ b" by (rule case_split) - qed - ultimately show ?thesis by (rule NIP) -qed - - -subsection {* Final Theorem *} - -theorem real_non_denum: - shows "\ (\f::nat\real. surj f)" -proof -- "by contradiction" - assume "\f::nat\real. surj f" - then obtain f::"nat\real" where "surj f" by auto - hence rangeF: "range f = UNIV" by (rule surj_range) - -- "We now produce a real number x that is not in the range of f, using the properties of newInt. " - have "\x. x \ (\n. newInt n f)" using newInt_notempty by blast - moreover have "\n. f n \ (\n. newInt n f)" by (rule newInt_inter) - ultimately obtain x where "x \ (\n. newInt n f)" and "\n. f n \ x" by blast - moreover from rangeF have "x \ range f" by simp - ultimately show False by blast -qed - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/Float.thy --- a/src/HOL/Real/Float.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,569 +0,0 @@ -(* Title: HOL/Real/Float.thy - ID: $Id$ - Author: Steven Obua -*) - -header {* Floating Point Representation of the Reals *} - -theory Float -imports Real Parity -uses "~~/src/Tools/float.ML" ("float_arith.ML") -begin - -definition - pow2 :: "int \ real" where - "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))" - -definition - float :: "int * int \ real" where - "float x = real (fst x) * pow2 (snd x)" - -lemma pow2_0[simp]: "pow2 0 = 1" -by (simp add: pow2_def) - -lemma pow2_1[simp]: "pow2 1 = 2" -by (simp add: pow2_def) - -lemma pow2_neg: "pow2 x = inverse (pow2 (-x))" -by (simp add: pow2_def) - -lemma pow2_add1: "pow2 (1 + a) = 2 * (pow2 a)" -proof - - have h: "! n. nat (2 + int n) - Suc 0 = nat (1 + int n)" by arith - have g: "! a b. a - -1 = a + (1::int)" by arith - have pos: "! n. pow2 (int n + 1) = 2 * pow2 (int n)" - apply (auto, induct_tac n) - apply (simp_all add: pow2_def) - apply (rule_tac m1="2" and n1="nat (2 + int na)" in ssubst[OF realpow_num_eq_if]) - by (auto simp add: h) - show ?thesis - proof (induct a) - case (1 n) - from pos show ?case by (simp add: ring_simps) - next - case (2 n) - show ?case - apply (auto) - apply (subst pow2_neg[of "- int n"]) - apply (subst pow2_neg[of "-1 - int n"]) - apply (auto simp add: g pos) - done - qed -qed - -lemma pow2_add: "pow2 (a+b) = (pow2 a) * (pow2 b)" -proof (induct b) - case (1 n) - show ?case - proof (induct n) - case 0 - show ?case by simp - next - case (Suc m) - show ?case by (auto simp add: ring_simps pow2_add1 prems) - qed -next - case (2 n) - show ?case - proof (induct n) - case 0 - show ?case - apply (auto) - apply (subst pow2_neg[of "a + -1"]) - apply (subst pow2_neg[of "-1"]) - apply (simp) - apply (insert pow2_add1[of "-a"]) - apply (simp add: ring_simps) - apply (subst pow2_neg[of "-a"]) - apply (simp) - done - case (Suc m) - have a: "int m - (a + -2) = 1 + (int m - a + 1)" by arith - have b: "int m - -2 = 1 + (int m + 1)" by arith - show ?case - apply (auto) - apply (subst pow2_neg[of "a + (-2 - int m)"]) - apply (subst pow2_neg[of "-2 - int m"]) - apply (auto simp add: ring_simps) - apply (subst a) - apply (subst b) - apply (simp only: pow2_add1) - apply (subst pow2_neg[of "int m - a + 1"]) - apply (subst pow2_neg[of "int m + 1"]) - apply auto - apply (insert prems) - apply (auto simp add: ring_simps) - done - qed -qed - -lemma "float (a, e) + float (b, e) = float (a + b, e)" -by (simp add: float_def ring_simps) - -definition - int_of_real :: "real \ int" where - "int_of_real x = (SOME y. real y = x)" - -definition - real_is_int :: "real \ bool" where - "real_is_int x = (EX (u::int). x = real u)" - -lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))" -by (auto simp add: real_is_int_def int_of_real_def) - -lemma float_transfer: "real_is_int ((real a)*(pow2 c)) \ float (a, b) = float (int_of_real ((real a)*(pow2 c)), b - c)" -by (simp add: float_def real_is_int_def2 pow2_add[symmetric]) - -lemma pow2_int: "pow2 (int c) = 2^c" -by (simp add: pow2_def) - -lemma float_transfer_nat: "float (a, b) = float (a * 2^c, b - int c)" -by (simp add: float_def pow2_int[symmetric] pow2_add[symmetric]) - -lemma real_is_int_real[simp]: "real_is_int (real (x::int))" -by (auto simp add: real_is_int_def int_of_real_def) - -lemma int_of_real_real[simp]: "int_of_real (real x) = x" -by (simp add: int_of_real_def) - -lemma real_int_of_real[simp]: "real_is_int x \ real (int_of_real x) = x" -by (auto simp add: int_of_real_def real_is_int_def) - -lemma real_is_int_add_int_of_real: "real_is_int a \ real_is_int b \ (int_of_real (a+b)) = (int_of_real a) + (int_of_real b)" -by (auto simp add: int_of_real_def real_is_int_def) - -lemma real_is_int_add[simp]: "real_is_int a \ real_is_int b \ real_is_int (a+b)" -apply (subst real_is_int_def2) -apply (simp add: real_is_int_add_int_of_real real_int_of_real) -done - -lemma int_of_real_sub: "real_is_int a \ real_is_int b \ (int_of_real (a-b)) = (int_of_real a) - (int_of_real b)" -by (auto simp add: int_of_real_def real_is_int_def) - -lemma real_is_int_sub[simp]: "real_is_int a \ real_is_int b \ real_is_int (a-b)" -apply (subst real_is_int_def2) -apply (simp add: int_of_real_sub real_int_of_real) -done - -lemma real_is_int_rep: "real_is_int x \ ?! (a::int). real a = x" -by (auto simp add: real_is_int_def) - -lemma int_of_real_mult: - assumes "real_is_int a" "real_is_int b" - shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)" -proof - - from prems have a: "?! (a'::int). real a' = a" by (rule_tac real_is_int_rep, auto) - from prems have b: "?! (b'::int). real b' = b" by (rule_tac real_is_int_rep, auto) - from a obtain a'::int where a':"a = real a'" by auto - from b obtain b'::int where b':"b = real b'" by auto - have r: "real a' * real b' = real (a' * b')" by auto - show ?thesis - apply (simp add: a' b') - apply (subst r) - apply (simp only: int_of_real_real) - done -qed - -lemma real_is_int_mult[simp]: "real_is_int a \ real_is_int b \ real_is_int (a*b)" -apply (subst real_is_int_def2) -apply (simp add: int_of_real_mult) -done - -lemma real_is_int_0[simp]: "real_is_int (0::real)" -by (simp add: real_is_int_def int_of_real_def) - -lemma real_is_int_1[simp]: "real_is_int (1::real)" -proof - - have "real_is_int (1::real) = real_is_int(real (1::int))" by auto - also have "\ = True" by (simp only: real_is_int_real) - ultimately show ?thesis by auto -qed - -lemma real_is_int_n1: "real_is_int (-1::real)" -proof - - have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto - also have "\ = True" by (simp only: real_is_int_real) - ultimately show ?thesis by auto -qed - -lemma real_is_int_number_of[simp]: "real_is_int ((number_of \ int \ real) x)" -proof - - have neg1: "real_is_int (-1::real)" - proof - - have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto - also have "\ = True" by (simp only: real_is_int_real) - ultimately show ?thesis by auto - qed - - { - fix x :: int - have "real_is_int ((number_of \ int \ real) x)" - unfolding number_of_eq - apply (induct x) - apply (induct_tac n) - apply (simp) - apply (simp) - apply (induct_tac n) - apply (simp add: neg1) - proof - - fix n :: nat - assume rn: "(real_is_int (of_int (- (int (Suc n)))))" - have s: "-(int (Suc (Suc n))) = -1 + - (int (Suc n))" by simp - show "real_is_int (of_int (- (int (Suc (Suc n)))))" - apply (simp only: s of_int_add) - apply (rule real_is_int_add) - apply (simp add: neg1) - apply (simp only: rn) - done - qed - } - note Abs_Bin = this - { - fix x :: int - have "? u. x = u" - apply (rule exI[where x = "x"]) - apply (simp) - done - } - then obtain u::int where "x = u" by auto - with Abs_Bin show ?thesis by auto -qed - -lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)" -by (simp add: int_of_real_def) - -lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)" -proof - - have 1: "(1::real) = real (1::int)" by auto - show ?thesis by (simp only: 1 int_of_real_real) -qed - -lemma int_of_real_number_of[simp]: "int_of_real (number_of b) = number_of b" -proof - - have "real_is_int (number_of b)" by simp - then have uu: "?! u::int. number_of b = real u" by (auto simp add: real_is_int_rep) - then obtain u::int where u:"number_of b = real u" by auto - have "number_of b = real ((number_of b)::int)" - by (simp add: number_of_eq real_of_int_def) - have ub: "number_of b = real ((number_of b)::int)" - by (simp add: number_of_eq real_of_int_def) - from uu u ub have unb: "u = number_of b" - by blast - have "int_of_real (number_of b) = u" by (simp add: u) - with unb show ?thesis by simp -qed - -lemma float_transfer_even: "even a \ float (a, b) = float (a div 2, b+1)" - apply (subst float_transfer[where a="a" and b="b" and c="-1", simplified]) - apply (simp_all add: pow2_def even_def real_is_int_def ring_simps) - apply (auto) -proof - - fix q::int - have a:"b - (-1\int) = (1\int) + b" by arith - show "(float (q, (b - (-1\int)))) = (float (q, ((1\int) + b)))" - by (simp add: a) -qed - -lemma int_div_zdiv: "int (a div b) = (int a) div (int b)" -by (rule zdiv_int) - -lemma int_mod_zmod: "int (a mod b) = (int a) mod (int b)" -by (rule zmod_int) - -lemma abs_div_2_less: "a \ 0 \ a \ -1 \ abs((a::int) div 2) < abs a" -by arith - -function norm_float :: "int \ int \ int \ int" where - "norm_float a b = (if a \ 0 \ even a then norm_float (a div 2) (b + 1) - else if a = 0 then (0, 0) else (a, b))" -by auto - -termination by (relation "measure (nat o abs o fst)") - (auto intro: abs_div_2_less) - -lemma norm_float: "float x = float (split norm_float x)" -proof - - { - fix a b :: int - have norm_float_pair: "float (a, b) = float (norm_float a b)" - proof (induct a b rule: norm_float.induct) - case (1 u v) - show ?case - proof cases - assume u: "u \ 0 \ even u" - with prems have ind: "float (u div 2, v + 1) = float (norm_float (u div 2) (v + 1))" by auto - with u have "float (u,v) = float (u div 2, v+1)" by (simp add: float_transfer_even) - then show ?thesis - apply (subst norm_float.simps) - apply (simp add: ind) - done - next - assume "~(u \ 0 \ even u)" - then show ?thesis - by (simp add: prems float_def) - qed - qed - } - note helper = this - have "? a b. x = (a,b)" by auto - then obtain a b where "x = (a, b)" by blast - then show ?thesis by (simp add: helper) -qed - -lemma float_add_l0: "float (0, e) + x = x" - by (simp add: float_def) - -lemma float_add_r0: "x + float (0, e) = x" - by (simp add: float_def) - -lemma float_add: - "float (a1, e1) + float (a2, e2) = - (if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1) - else float (a1*2^(nat (e1-e2))+a2, e2))" - apply (simp add: float_def ring_simps) - apply (auto simp add: pow2_int[symmetric] pow2_add[symmetric]) - done - -lemma float_add_assoc1: - "(x + float (y1, e1)) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x" - by simp - -lemma float_add_assoc2: - "(float (y1, e1) + x) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x" - by simp - -lemma float_add_assoc3: - "float (y1, e1) + (x + float (y2, e2)) = (float (y1, e1) + float (y2, e2)) + x" - by simp - -lemma float_add_assoc4: - "float (y1, e1) + (float (y2, e2) + x) = (float (y1, e1) + float (y2, e2)) + x" - by simp - -lemma float_mult_l0: "float (0, e) * x = float (0, 0)" - by (simp add: float_def) - -lemma float_mult_r0: "x * float (0, e) = float (0, 0)" - by (simp add: float_def) - -definition - lbound :: "real \ real" -where - "lbound x = min 0 x" - -definition - ubound :: "real \ real" -where - "ubound x = max 0 x" - -lemma lbound: "lbound x \ x" - by (simp add: lbound_def) - -lemma ubound: "x \ ubound x" - by (simp add: ubound_def) - -lemma float_mult: - "float (a1, e1) * float (a2, e2) = - (float (a1 * a2, e1 + e2))" - by (simp add: float_def pow2_add) - -lemma float_minus: - "- (float (a,b)) = float (-a, b)" - by (simp add: float_def) - -lemma zero_less_pow2: - "0 < pow2 x" -proof - - { - fix y - have "0 <= y \ 0 < pow2 y" - by (induct y, induct_tac n, simp_all add: pow2_add) - } - note helper=this - show ?thesis - apply (case_tac "0 <= x") - apply (simp add: helper) - apply (subst pow2_neg) - apply (simp add: helper) - done -qed - -lemma zero_le_float: - "(0 <= float (a,b)) = (0 <= a)" - apply (auto simp add: float_def) - apply (auto simp add: zero_le_mult_iff zero_less_pow2) - apply (insert zero_less_pow2[of b]) - apply (simp_all) - done - -lemma float_le_zero: - "(float (a,b) <= 0) = (a <= 0)" - apply (auto simp add: float_def) - apply (auto simp add: mult_le_0_iff) - apply (insert zero_less_pow2[of b]) - apply auto - done - -lemma float_abs: - "abs (float (a,b)) = (if 0 <= a then (float (a,b)) else (float (-a,b)))" - apply (auto simp add: abs_if) - apply (simp_all add: zero_le_float[symmetric, of a b] float_minus) - done - -lemma float_zero: - "float (0, b) = 0" - by (simp add: float_def) - -lemma float_pprt: - "pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0, b)))" - by (auto simp add: zero_le_float float_le_zero float_zero) - -lemma pprt_lbound: "pprt (lbound x) = float (0, 0)" - apply (simp add: float_def) - apply (rule pprt_eq_0) - apply (simp add: lbound_def) - done - -lemma nprt_ubound: "nprt (ubound x) = float (0, 0)" - apply (simp add: float_def) - apply (rule nprt_eq_0) - apply (simp add: ubound_def) - done - -lemma float_nprt: - "nprt (float (a, b)) = (if 0 <= a then (float (0,b)) else (float (a, b)))" - by (auto simp add: zero_le_float float_le_zero float_zero) - -lemma norm_0_1: "(0::_::number_ring) = Numeral0 & (1::_::number_ring) = Numeral1" - by auto - -lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)" - by simp - -lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)" - by simp - -lemma mult_left_one: "1 * a = (a::'a::semiring_1)" - by simp - -lemma mult_right_one: "a * 1 = (a::'a::semiring_1)" - by simp - -lemma int_pow_0: "(a::int)^(Numeral0) = 1" - by simp - -lemma int_pow_1: "(a::int)^(Numeral1) = a" - by simp - -lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0" - by simp - -lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1" - by simp - -lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0" - by simp - -lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1" - by simp - -lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1" - by simp - -lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1" -proof - - have 1:"((-1)::nat) = 0" - by simp - show ?thesis by (simp add: 1) -qed - -lemma fst_cong: "a=a' \ fst (a,b) = fst (a',b)" - by simp - -lemma snd_cong: "b=b' \ snd (a,b) = snd (a,b')" - by simp - -lemma lift_bool: "x \ x=True" - by simp - -lemma nlift_bool: "~x \ x=False" - by simp - -lemma not_false_eq_true: "(~ False) = True" by simp - -lemma not_true_eq_false: "(~ True) = False" by simp - -lemmas binarith = - normalize_bin_simps - pred_bin_simps succ_bin_simps - add_bin_simps minus_bin_simps mult_bin_simps - -lemma int_eq_number_of_eq: - "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)" - by simp - -lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)" - by (simp only: iszero_number_of_Pls) - -lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))" - by simp - -lemma int_iszero_number_of_Bit0: "iszero ((number_of (Int.Bit0 w))::int) = iszero ((number_of w)::int)" - by simp - -lemma int_iszero_number_of_Bit1: "\ iszero ((number_of (Int.Bit1 w))::int)" - by simp - -lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)" - by simp - -lemma int_not_neg_number_of_Pls: "\ (neg (Numeral0::int))" - by simp - -lemma int_neg_number_of_Min: "neg (-1::int)" - by simp - -lemma int_neg_number_of_Bit0: "neg ((number_of (Int.Bit0 w))::int) = neg ((number_of w)::int)" - by simp - -lemma int_neg_number_of_Bit1: "neg ((number_of (Int.Bit1 w))::int) = neg ((number_of w)::int)" - by simp - -lemma int_le_number_of_eq: "(((number_of x)::int) \ number_of y) = (\ neg ((number_of (y + (uminus x)))::int))" - by simp - -lemmas intarithrel = - int_eq_number_of_eq - lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_Bit0 - lift_bool[OF int_iszero_number_of_Bit1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min] - int_neg_number_of_Bit0 int_neg_number_of_Bit1 int_le_number_of_eq - -lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)" - by simp - -lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (v + (uminus w))" - by simp - -lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (v * w)" - by simp - -lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)" - by simp - -lemmas intarith = int_number_of_add_sym int_number_of_minus_sym int_number_of_diff_sym int_number_of_mult_sym - -lemmas natarith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of - -lemmas powerarith = nat_number_of zpower_number_of_even - zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring] - zpower_Pls zpower_Min - -lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0 - float_minus float_abs zero_le_float float_pprt float_nprt pprt_lbound nprt_ubound - -(* for use with the compute oracle *) -lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false - -use "float_arith.ML"; - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/Lubs.thy --- a/src/HOL/Real/Lubs.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,113 +0,0 @@ -(* Title : Lubs.thy - ID : $Id$ - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge -*) - -header{*Definitions of Upper Bounds and Least Upper Bounds*} - -theory Lubs -imports Plain -begin - -text{*Thanks to suggestions by James Margetson*} - -definition - setle :: "['a set, 'a::ord] => bool" (infixl "*<=" 70) where - "S *<= x = (ALL y: S. y <= x)" - -definition - setge :: "['a::ord, 'a set] => bool" (infixl "<=*" 70) where - "x <=* S = (ALL y: S. x <= y)" - -definition - leastP :: "['a =>bool,'a::ord] => bool" where - "leastP P x = (P x & x <=* Collect P)" - -definition - isUb :: "['a set, 'a set, 'a::ord] => bool" where - "isUb R S x = (S *<= x & x: R)" - -definition - isLub :: "['a set, 'a set, 'a::ord] => bool" where - "isLub R S x = leastP (isUb R S) x" - -definition - ubs :: "['a set, 'a::ord set] => 'a set" where - "ubs R S = Collect (isUb R S)" - - - -subsection{*Rules for the Relations @{text "*<="} and @{text "<=*"}*} - -lemma setleI: "ALL y: S. y <= x ==> S *<= x" -by (simp add: setle_def) - -lemma setleD: "[| S *<= x; y: S |] ==> y <= x" -by (simp add: setle_def) - -lemma setgeI: "ALL y: S. x<= y ==> x <=* S" -by (simp add: setge_def) - -lemma setgeD: "[| x <=* S; y: S |] ==> x <= y" -by (simp add: setge_def) - - -subsection{*Rules about the Operators @{term leastP}, @{term ub} - and @{term lub}*} - -lemma leastPD1: "leastP P x ==> P x" -by (simp add: leastP_def) - -lemma leastPD2: "leastP P x ==> x <=* Collect P" -by (simp add: leastP_def) - -lemma leastPD3: "[| leastP P x; y: Collect P |] ==> x <= y" -by (blast dest!: leastPD2 setgeD) - -lemma isLubD1: "isLub R S x ==> S *<= x" -by (simp add: isLub_def isUb_def leastP_def) - -lemma isLubD1a: "isLub R S x ==> x: R" -by (simp add: isLub_def isUb_def leastP_def) - -lemma isLub_isUb: "isLub R S x ==> isUb R S x" -apply (simp add: isUb_def) -apply (blast dest: isLubD1 isLubD1a) -done - -lemma isLubD2: "[| isLub R S x; y : S |] ==> y <= x" -by (blast dest!: isLubD1 setleD) - -lemma isLubD3: "isLub R S x ==> leastP(isUb R S) x" -by (simp add: isLub_def) - -lemma isLubI1: "leastP(isUb R S) x ==> isLub R S x" -by (simp add: isLub_def) - -lemma isLubI2: "[| isUb R S x; x <=* Collect (isUb R S) |] ==> isLub R S x" -by (simp add: isLub_def leastP_def) - -lemma isUbD: "[| isUb R S x; y : S |] ==> y <= x" -by (simp add: isUb_def setle_def) - -lemma isUbD2: "isUb R S x ==> S *<= x" -by (simp add: isUb_def) - -lemma isUbD2a: "isUb R S x ==> x: R" -by (simp add: isUb_def) - -lemma isUbI: "[| S *<= x; x: R |] ==> isUb R S x" -by (simp add: isUb_def) - -lemma isLub_le_isUb: "[| isLub R S x; isUb R S y |] ==> x <= y" -apply (simp add: isLub_def) -apply (blast intro!: leastPD3) -done - -lemma isLub_ubs: "isLub R S x ==> x <=* ubs R S" -apply (simp add: ubs_def isLub_def) -apply (erule leastPD2) -done - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/PReal.thy --- a/src/HOL/Real/PReal.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1343 +0,0 @@ -(* Title : PReal.thy - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Description : The positive reals as Dedekind sections of positive - rationals. Fundamentals of Abstract Analysis [Gleason- p. 121] - provides some of the definitions. -*) - -header {* Positive real numbers *} - -theory PReal -imports Rational "~~/src/HOL/Library/Dense_Linear_Order" -begin - -text{*Could be generalized and moved to @{text Ring_and_Field}*} -lemma add_eq_exists: "\x. a+x = (b::rat)" -by (rule_tac x="b-a" in exI, simp) - -definition - cut :: "rat set => bool" where - [code del]: "cut A = ({} \ A & - A < {r. 0 < r} & - (\y \ A. ((\z. 0 z \ A) & (\u \ A. y < u))))" - -lemma cut_of_rat: - assumes q: "0 < q" shows "cut {r::rat. 0 < r & r < q}" (is "cut ?A") -proof - - from q have pos: "?A < {r. 0 < r}" by force - have nonempty: "{} \ ?A" - proof - show "{} \ ?A" by simp - show "{} \ ?A" - by (force simp only: q eq_commute [of "{}"] interval_empty_iff) - qed - show ?thesis - by (simp add: cut_def pos nonempty, - blast dest: dense intro: order_less_trans) -qed - - -typedef preal = "{A. cut A}" - by (blast intro: cut_of_rat [OF zero_less_one]) - -definition - preal_of_rat :: "rat => preal" where - "preal_of_rat q = Abs_preal {x::rat. 0 < x & x < q}" - -definition - psup :: "preal set => preal" where - "psup P = Abs_preal (\X \ P. Rep_preal X)" - -definition - add_set :: "[rat set,rat set] => rat set" where - "add_set A B = {w. \x \ A. \y \ B. w = x + y}" - -definition - diff_set :: "[rat set,rat set] => rat set" where - [code del]: "diff_set A B = {w. \x. 0 < w & 0 < x & x \ B & x + w \ A}" - -definition - mult_set :: "[rat set,rat set] => rat set" where - "mult_set A B = {w. \x \ A. \y \ B. w = x * y}" - -definition - inverse_set :: "rat set => rat set" where - [code del]: "inverse_set A = {x. \y. 0 < x & x < y & inverse y \ A}" - -instantiation preal :: "{ord, plus, minus, times, inverse, one}" -begin - -definition - preal_less_def [code del]: - "R < S == Rep_preal R < Rep_preal S" - -definition - preal_le_def [code del]: - "R \ S == Rep_preal R \ Rep_preal S" - -definition - preal_add_def: - "R + S == Abs_preal (add_set (Rep_preal R) (Rep_preal S))" - -definition - preal_diff_def: - "R - S == Abs_preal (diff_set (Rep_preal R) (Rep_preal S))" - -definition - preal_mult_def: - "R * S == Abs_preal (mult_set (Rep_preal R) (Rep_preal S))" - -definition - preal_inverse_def: - "inverse R == Abs_preal (inverse_set (Rep_preal R))" - -definition "R / S = R * inverse (S\preal)" - -definition - preal_one_def: - "1 == preal_of_rat 1" - -instance .. - -end - - -text{*Reduces equality on abstractions to equality on representatives*} -declare Abs_preal_inject [simp] -declare Abs_preal_inverse [simp] - -lemma rat_mem_preal: "0 < q ==> {r::rat. 0 < r & r < q} \ preal" -by (simp add: preal_def cut_of_rat) - -lemma preal_nonempty: "A \ preal ==> \x\A. 0 < x" -by (unfold preal_def cut_def, blast) - -lemma preal_Ex_mem: "A \ preal \ \x. x \ A" -by (drule preal_nonempty, fast) - -lemma preal_imp_psubset_positives: "A \ preal ==> A < {r. 0 < r}" -by (force simp add: preal_def cut_def) - -lemma preal_exists_bound: "A \ preal ==> \x. 0 < x & x \ A" -by (drule preal_imp_psubset_positives, auto) - -lemma preal_exists_greater: "[| A \ preal; y \ A |] ==> \u \ A. y < u" -by (unfold preal_def cut_def, blast) - -lemma preal_downwards_closed: "[| A \ preal; y \ A; 0 < z; z < y |] ==> z \ A" -by (unfold preal_def cut_def, blast) - -text{*Relaxing the final premise*} -lemma preal_downwards_closed': - "[| A \ preal; y \ A; 0 < z; z \ y |] ==> z \ A" -apply (simp add: order_le_less) -apply (blast intro: preal_downwards_closed) -done - -text{*A positive fraction not in a positive real is an upper bound. - Gleason p. 122 - Remark (1)*} - -lemma not_in_preal_ub: - assumes A: "A \ preal" - and notx: "x \ A" - and y: "y \ A" - and pos: "0 < x" - shows "y < x" -proof (cases rule: linorder_cases) - assume "xx. x \ Rep_preal X" -by (rule preal_Ex_mem [OF Rep_preal]) - -lemma Rep_preal_exists_bound: "\x>0. x \ Rep_preal X" -by (rule preal_exists_bound [OF Rep_preal]) - -lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF Rep_preal] - - - -subsection{*@{term preal_of_prat}: the Injection from prat to preal*} - -lemma rat_less_set_mem_preal: "0 < y ==> {u::rat. 0 < u & u < y} \ preal" -by (simp add: preal_def cut_of_rat) - -lemma rat_subset_imp_le: - "[|{u::rat. 0 < u & u < x} \ {u. 0 < u & u < y}; 0 x \ y" -apply (simp add: linorder_not_less [symmetric]) -apply (blast dest: dense intro: order_less_trans) -done - -lemma rat_set_eq_imp_eq: - "[|{u::rat. 0 < u & u < x} = {u. 0 < u & u < y}; - 0 < x; 0 < y|] ==> x = y" -by (blast intro: rat_subset_imp_le order_antisym) - - - -subsection{*Properties of Ordering*} - -instance preal :: order -proof - fix w :: preal - show "w \ w" by (simp add: preal_le_def) -next - fix i j k :: preal - assume "i \ j" and "j \ k" - then show "i \ k" by (simp add: preal_le_def) -next - fix z w :: preal - assume "z \ w" and "w \ z" - then show "z = w" by (simp add: preal_le_def Rep_preal_inject) -next - fix z w :: preal - show "z < w \ z \ w \ \ w \ z" - by (auto simp add: preal_le_def preal_less_def Rep_preal_inject) -qed - -lemma preal_imp_pos: "[|A \ preal; r \ A|] ==> 0 < r" -by (insert preal_imp_psubset_positives, blast) - -instance preal :: linorder -proof - fix x y :: preal - show "x <= y | y <= x" - apply (auto simp add: preal_le_def) - apply (rule ccontr) - apply (blast dest: not_in_Rep_preal_ub intro: preal_imp_pos [OF Rep_preal] - elim: order_less_asym) - done -qed - -instantiation preal :: distrib_lattice -begin - -definition - "(inf \ preal \ preal \ preal) = min" - -definition - "(sup \ preal \ preal \ preal) = max" - -instance - by intro_classes - (auto simp add: inf_preal_def sup_preal_def min_max.sup_inf_distrib1) - -end - -subsection{*Properties of Addition*} - -lemma preal_add_commute: "(x::preal) + y = y + x" -apply (unfold preal_add_def add_set_def) -apply (rule_tac f = Abs_preal in arg_cong) -apply (force simp add: add_commute) -done - -text{*Lemmas for proving that addition of two positive reals gives - a positive real*} - -lemma empty_psubset_nonempty: "a \ A ==> {} \ A" -by blast - -text{*Part 1 of Dedekind sections definition*} -lemma add_set_not_empty: - "[|A \ preal; B \ preal|] ==> {} \ add_set A B" -apply (drule preal_nonempty)+ -apply (auto simp add: add_set_def) -done - -text{*Part 2 of Dedekind sections definition. A structured version of -this proof is @{text preal_not_mem_mult_set_Ex} below.*} -lemma preal_not_mem_add_set_Ex: - "[|A \ preal; B \ preal|] ==> \q>0. q \ add_set A B" -apply (insert preal_exists_bound [of A] preal_exists_bound [of B], auto) -apply (rule_tac x = "x+xa" in exI) -apply (simp add: add_set_def, clarify) -apply (drule (3) not_in_preal_ub)+ -apply (force dest: add_strict_mono) -done - -lemma add_set_not_rat_set: - assumes A: "A \ preal" - and B: "B \ preal" - shows "add_set A B < {r. 0 < r}" -proof - from preal_imp_pos [OF A] preal_imp_pos [OF B] - show "add_set A B \ {r. 0 < r}" by (force simp add: add_set_def) -next - show "add_set A B \ {r. 0 < r}" - by (insert preal_not_mem_add_set_Ex [OF A B], blast) -qed - -text{*Part 3 of Dedekind sections definition*} -lemma add_set_lemma3: - "[|A \ preal; B \ preal; u \ add_set A B; 0 < z; z < u|] - ==> z \ add_set A B" -proof (unfold add_set_def, clarify) - fix x::rat and y::rat - assume A: "A \ preal" - and B: "B \ preal" - and [simp]: "0 < z" - and zless: "z < x + y" - and x: "x \ A" - and y: "y \ B" - have xpos [simp]: "0x' \ A. \y'\B. z = x' + y'" - proof (intro bexI) - show "z = x*?f + y*?f" - by (simp add: left_distrib [symmetric] divide_inverse mult_ac - order_less_imp_not_eq2) - next - show "y * ?f \ B" - proof (rule preal_downwards_closed [OF B y]) - show "0 < y * ?f" - by (simp add: divide_inverse zero_less_mult_iff) - next - show "y * ?f < y" - by (insert mult_strict_left_mono [OF fless ypos], simp) - qed - next - show "x * ?f \ A" - proof (rule preal_downwards_closed [OF A x]) - show "0 < x * ?f" - by (simp add: divide_inverse zero_less_mult_iff) - next - show "x * ?f < x" - by (insert mult_strict_left_mono [OF fless xpos], simp) - qed - qed -qed - -text{*Part 4 of Dedekind sections definition*} -lemma add_set_lemma4: - "[|A \ preal; B \ preal; y \ add_set A B|] ==> \u \ add_set A B. y < u" -apply (auto simp add: add_set_def) -apply (frule preal_exists_greater [of A], auto) -apply (rule_tac x="u + y" in exI) -apply (auto intro: add_strict_left_mono) -done - -lemma mem_add_set: - "[|A \ preal; B \ preal|] ==> add_set A B \ preal" -apply (simp (no_asm_simp) add: preal_def cut_def) -apply (blast intro!: add_set_not_empty add_set_not_rat_set - add_set_lemma3 add_set_lemma4) -done - -lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)" -apply (simp add: preal_add_def mem_add_set Rep_preal) -apply (force simp add: add_set_def add_ac) -done - -instance preal :: ab_semigroup_add -proof - fix a b c :: preal - show "(a + b) + c = a + (b + c)" by (rule preal_add_assoc) - show "a + b = b + a" by (rule preal_add_commute) -qed - -lemma preal_add_left_commute: "x + (y + z) = y + ((x + z)::preal)" -by (rule add_left_commute) - -text{* Positive Real addition is an AC operator *} -lemmas preal_add_ac = preal_add_assoc preal_add_commute preal_add_left_commute - - -subsection{*Properties of Multiplication*} - -text{*Proofs essentially same as for addition*} - -lemma preal_mult_commute: "(x::preal) * y = y * x" -apply (unfold preal_mult_def mult_set_def) -apply (rule_tac f = Abs_preal in arg_cong) -apply (force simp add: mult_commute) -done - -text{*Multiplication of two positive reals gives a positive real.*} - -text{*Lemmas for proving positive reals multiplication set in @{typ preal}*} - -text{*Part 1 of Dedekind sections definition*} -lemma mult_set_not_empty: - "[|A \ preal; B \ preal|] ==> {} \ mult_set A B" -apply (insert preal_nonempty [of A] preal_nonempty [of B]) -apply (auto simp add: mult_set_def) -done - -text{*Part 2 of Dedekind sections definition*} -lemma preal_not_mem_mult_set_Ex: - assumes A: "A \ preal" - and B: "B \ preal" - shows "\q. 0 < q & q \ mult_set A B" -proof - - from preal_exists_bound [OF A] - obtain x where [simp]: "0 < x" "x \ A" by blast - from preal_exists_bound [OF B] - obtain y where [simp]: "0 < y" "y \ B" by blast - show ?thesis - proof (intro exI conjI) - show "0 < x*y" by (simp add: mult_pos_pos) - show "x * y \ mult_set A B" - proof - - { fix u::rat and v::rat - assume "u \ A" and "v \ B" and "x*y = u*v" - moreover - with prems have "uv" - by (blast intro: preal_imp_pos [OF B] order_less_imp_le prems) - moreover - from calculation - have "u*v < x*y" by (blast intro: mult_strict_mono prems) - ultimately have False by force } - thus ?thesis by (auto simp add: mult_set_def) - qed - qed -qed - -lemma mult_set_not_rat_set: - assumes A: "A \ preal" - and B: "B \ preal" - shows "mult_set A B < {r. 0 < r}" -proof - show "mult_set A B \ {r. 0 < r}" - by (force simp add: mult_set_def - intro: preal_imp_pos [OF A] preal_imp_pos [OF B] mult_pos_pos) - show "mult_set A B \ {r. 0 < r}" - using preal_not_mem_mult_set_Ex [OF A B] by blast -qed - - - -text{*Part 3 of Dedekind sections definition*} -lemma mult_set_lemma3: - "[|A \ preal; B \ preal; u \ mult_set A B; 0 < z; z < u|] - ==> z \ mult_set A B" -proof (unfold mult_set_def, clarify) - fix x::rat and y::rat - assume A: "A \ preal" - and B: "B \ preal" - and [simp]: "0 < z" - and zless: "z < x * y" - and x: "x \ A" - and y: "y \ B" - have [simp]: "0x' \ A. \y' \ B. z = x' * y'" - proof - show "\y'\B. z = (z/y) * y'" - proof - show "z = (z/y)*y" - by (simp add: divide_inverse mult_commute [of y] mult_assoc - order_less_imp_not_eq2) - show "y \ B" by fact - qed - next - show "z/y \ A" - proof (rule preal_downwards_closed [OF A x]) - show "0 < z/y" - by (simp add: zero_less_divide_iff) - show "z/y < x" by (simp add: pos_divide_less_eq zless) - qed - qed -qed - -text{*Part 4 of Dedekind sections definition*} -lemma mult_set_lemma4: - "[|A \ preal; B \ preal; y \ mult_set A B|] ==> \u \ mult_set A B. y < u" -apply (auto simp add: mult_set_def) -apply (frule preal_exists_greater [of A], auto) -apply (rule_tac x="u * y" in exI) -apply (auto intro: preal_imp_pos [of A] preal_imp_pos [of B] - mult_strict_right_mono) -done - - -lemma mem_mult_set: - "[|A \ preal; B \ preal|] ==> mult_set A B \ preal" -apply (simp (no_asm_simp) add: preal_def cut_def) -apply (blast intro!: mult_set_not_empty mult_set_not_rat_set - mult_set_lemma3 mult_set_lemma4) -done - -lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)" -apply (simp add: preal_mult_def mem_mult_set Rep_preal) -apply (force simp add: mult_set_def mult_ac) -done - -instance preal :: ab_semigroup_mult -proof - fix a b c :: preal - show "(a * b) * c = a * (b * c)" by (rule preal_mult_assoc) - show "a * b = b * a" by (rule preal_mult_commute) -qed - -lemma preal_mult_left_commute: "x * (y * z) = y * ((x * z)::preal)" -by (rule mult_left_commute) - - -text{* Positive Real multiplication is an AC operator *} -lemmas preal_mult_ac = - preal_mult_assoc preal_mult_commute preal_mult_left_commute - - -text{* Positive real 1 is the multiplicative identity element *} - -lemma preal_mult_1: "(1::preal) * z = z" -unfolding preal_one_def -proof (induct z) - fix A :: "rat set" - assume A: "A \ preal" - have "{w. \u. 0 < u \ u < 1 & (\v \ A. w = u * v)} = A" (is "?lhs = A") - proof - show "?lhs \ A" - proof clarify - fix x::rat and u::rat and v::rat - assume upos: "0 A" - have vpos: "0 A" - by (force intro: preal_downwards_closed [OF A v] mult_pos_pos - upos vpos) - qed - next - show "A \ ?lhs" - proof clarify - fix x::rat - assume x: "x \ A" - have xpos: "0 A" and xlessv: "x < v" .. - have vpos: "0u. 0 < u \ u < 1 \ (\v\A. x = u * v)" - proof (intro exI conjI) - show "0 < x/v" - by (simp add: zero_less_divide_iff xpos vpos) - show "x / v < 1" - by (simp add: pos_divide_less_eq vpos xlessv) - show "\v'\A. x = (x / v) * v'" - proof - show "x = (x/v)*v" - by (simp add: divide_inverse mult_assoc vpos - order_less_imp_not_eq2) - show "v \ A" by fact - qed - qed - qed - qed - thus "preal_of_rat 1 * Abs_preal A = Abs_preal A" - by (simp add: preal_of_rat_def preal_mult_def mult_set_def - rat_mem_preal A) -qed - -instance preal :: comm_monoid_mult -by intro_classes (rule preal_mult_1) - -lemma preal_mult_1_right: "z * (1::preal) = z" -by (rule mult_1_right) - - -subsection{*Distribution of Multiplication across Addition*} - -lemma mem_Rep_preal_add_iff: - "(z \ Rep_preal(R+S)) = (\x \ Rep_preal R. \y \ Rep_preal S. z = x + y)" -apply (simp add: preal_add_def mem_add_set Rep_preal) -apply (simp add: add_set_def) -done - -lemma mem_Rep_preal_mult_iff: - "(z \ Rep_preal(R*S)) = (\x \ Rep_preal R. \y \ Rep_preal S. z = x * y)" -apply (simp add: preal_mult_def mem_mult_set Rep_preal) -apply (simp add: mult_set_def) -done - -lemma distrib_subset1: - "Rep_preal (w * (x + y)) \ Rep_preal (w * x + w * y)" -apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff) -apply (force simp add: right_distrib) -done - -lemma preal_add_mult_distrib_mean: - assumes a: "a \ Rep_preal w" - and b: "b \ Rep_preal w" - and d: "d \ Rep_preal x" - and e: "e \ Rep_preal y" - shows "\c \ Rep_preal w. a * d + b * e = c * (d + e)" -proof - let ?c = "(a*d + b*e)/(d+e)" - have [simp]: "0 Rep_preal w" - proof (cases rule: linorder_le_cases) - assume "a \ b" - hence "?c \ b" - by (simp add: pos_divide_le_eq right_distrib mult_right_mono - order_less_imp_le) - thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal b cpos]) - next - assume "b \ a" - hence "?c \ a" - by (simp add: pos_divide_le_eq right_distrib mult_right_mono - order_less_imp_le) - thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal a cpos]) - qed -qed - -lemma distrib_subset2: - "Rep_preal (w * x + w * y) \ Rep_preal (w * (x + y))" -apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff) -apply (drule_tac w=w and x=x and y=y in preal_add_mult_distrib_mean, auto) -done - -lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)" -apply (rule Rep_preal_inject [THEN iffD1]) -apply (rule equalityI [OF distrib_subset1 distrib_subset2]) -done - -lemma preal_add_mult_distrib: "(((x::preal) + y) * w) = (x * w) + (y * w)" -by (simp add: preal_mult_commute preal_add_mult_distrib2) - -instance preal :: comm_semiring -by intro_classes (rule preal_add_mult_distrib) - - -subsection{*Existence of Inverse, a Positive Real*} - -lemma mem_inv_set_ex: - assumes A: "A \ preal" shows "\x y. 0 < x & x < y & inverse y \ A" -proof - - from preal_exists_bound [OF A] - obtain x where [simp]: "0 A" by blast - show ?thesis - proof (intro exI conjI) - show "0 < inverse (x+1)" - by (simp add: order_less_trans [OF _ less_add_one]) - show "inverse(x+1) < inverse x" - by (simp add: less_imp_inverse_less less_add_one) - show "inverse (inverse x) \ A" - by (simp add: order_less_imp_not_eq2) - qed -qed - -text{*Part 1 of Dedekind sections definition*} -lemma inverse_set_not_empty: - "A \ preal ==> {} \ inverse_set A" -apply (insert mem_inv_set_ex [of A]) -apply (auto simp add: inverse_set_def) -done - -text{*Part 2 of Dedekind sections definition*} - -lemma preal_not_mem_inverse_set_Ex: - assumes A: "A \ preal" shows "\q. 0 < q & q \ inverse_set A" -proof - - from preal_nonempty [OF A] - obtain x where x: "x \ A" and xpos [simp]: "0 inverse_set A" - proof - - { fix y::rat - assume ygt: "inverse x < y" - have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt]) - have iyless: "inverse y < x" - by (simp add: inverse_less_imp_less [of x] ygt) - have "inverse y \ A" - by (simp add: preal_downwards_closed [OF A x] iyless)} - thus ?thesis by (auto simp add: inverse_set_def) - qed - qed -qed - -lemma inverse_set_not_rat_set: - assumes A: "A \ preal" shows "inverse_set A < {r. 0 < r}" -proof - show "inverse_set A \ {r. 0 < r}" by (force simp add: inverse_set_def) -next - show "inverse_set A \ {r. 0 < r}" - by (insert preal_not_mem_inverse_set_Ex [OF A], blast) -qed - -text{*Part 3 of Dedekind sections definition*} -lemma inverse_set_lemma3: - "[|A \ preal; u \ inverse_set A; 0 < z; z < u|] - ==> z \ inverse_set A" -apply (auto simp add: inverse_set_def) -apply (auto intro: order_less_trans) -done - -text{*Part 4 of Dedekind sections definition*} -lemma inverse_set_lemma4: - "[|A \ preal; y \ inverse_set A|] ==> \u \ inverse_set A. y < u" -apply (auto simp add: inverse_set_def) -apply (drule dense [of y]) -apply (blast intro: order_less_trans) -done - - -lemma mem_inverse_set: - "A \ preal ==> inverse_set A \ preal" -apply (simp (no_asm_simp) add: preal_def cut_def) -apply (blast intro!: inverse_set_not_empty inverse_set_not_rat_set - inverse_set_lemma3 inverse_set_lemma4) -done - - -subsection{*Gleason's Lemma 9-3.4, page 122*} - -lemma Gleason9_34_exists: - assumes A: "A \ preal" - and "\x\A. x + u \ A" - and "0 \ z" - shows "\b\A. b + (of_int z) * u \ A" -proof (cases z rule: int_cases) - case (nonneg n) - show ?thesis - proof (simp add: prems, induct n) - case 0 - from preal_nonempty [OF A] - show ?case by force - case (Suc k) - from this obtain b where "b \ A" "b + of_nat k * u \ A" .. - hence "b + of_int (int k)*u + u \ A" by (simp add: prems) - thus ?case by (force simp add: left_distrib add_ac prems) - qed -next - case (neg n) - with prems show ?thesis by simp -qed - -lemma Gleason9_34_contra: - assumes A: "A \ preal" - shows "[|\x\A. x + u \ A; 0 < u; 0 < y; y \ A|] ==> False" -proof (induct u, induct y) - fix a::int and b::int - fix c::int and d::int - assume bpos [simp]: "0 < b" - and dpos [simp]: "0 < d" - and closed: "\x\A. x + (Fract c d) \ A" - and upos: "0 < Fract c d" - and ypos: "0 < Fract a b" - and notin: "Fract a b \ A" - have cpos [simp]: "0 < c" - by (simp add: zero_less_Fract_iff [OF dpos, symmetric] upos) - have apos [simp]: "0 < a" - by (simp add: zero_less_Fract_iff [OF bpos, symmetric] ypos) - let ?k = "a*d" - have frle: "Fract a b \ Fract ?k 1 * (Fract c d)" - proof - - have "?thesis = ((a * d * b * d) \ c * b * (a * d * b * d))" - by (simp add: mult_rat le_rat order_less_imp_not_eq2 mult_ac) - moreover - have "(1 * (a * d * b * d)) \ c * b * (a * d * b * d)" - by (rule mult_mono, - simp_all add: int_one_le_iff_zero_less zero_less_mult_iff - order_less_imp_le) - ultimately - show ?thesis by simp - qed - have k: "0 \ ?k" by (simp add: order_less_imp_le zero_less_mult_iff) - from Gleason9_34_exists [OF A closed k] - obtain z where z: "z \ A" - and mem: "z + of_int ?k * Fract c d \ A" .. - have less: "z + of_int ?k * Fract c d < Fract a b" - by (rule not_in_preal_ub [OF A notin mem ypos]) - have "0 preal" - and upos: "0 < u" - shows "\r \ A. r + u \ A" -proof (rule ccontr, simp) - assume closed: "\r\A. r + u \ A" - from preal_exists_bound [OF A] - obtain y where y: "y \ A" and ypos: "0 < y" by blast - show False - by (rule Gleason9_34_contra [OF A closed upos ypos y]) -qed - - - -subsection{*Gleason's Lemma 9-3.6*} - -lemma lemma_gleason9_36: - assumes A: "A \ preal" - and x: "1 < x" - shows "\r \ A. r*x \ A" -proof - - from preal_nonempty [OF A] - obtain y where y: "y \ A" and ypos: "0r\A. r * x \ A)" - with y have ymem: "y * x \ A" by blast - from ypos mult_strict_left_mono [OF x] - have yless: "y < y*x" by simp - let ?d = "y*x - y" - from yless have dpos: "0 < ?d" and eq: "y + ?d = y*x" by auto - from Gleason9_34 [OF A dpos] - obtain r where r: "r\A" and notin: "r + ?d \ A" .. - have rpos: "0 y + ?d)" - proof - assume le: "r + ?d \ y + ?d" - from ymem have yd: "y + ?d \ A" by (simp add: eq) - have "r + ?d \ A" by (rule preal_downwards_closed' [OF A yd rdpos le]) - with notin show False by simp - qed - hence "y < r" by simp - with ypos have dless: "?d < (r * ?d)/y" - by (simp add: pos_less_divide_eq mult_commute [of ?d] - mult_strict_right_mono dpos) - have "r + ?d < r*x" - proof - - have "r + ?d < r + (r * ?d)/y" by (simp add: dless) - also with ypos have "... = (r/y) * (y + ?d)" - by (simp only: right_distrib divide_inverse mult_ac, simp) - also have "... = r*x" using ypos - by (simp add: times_divide_eq_left) - finally show "r + ?d < r*x" . - qed - with r notin rdpos - show "\r\A. r * x \ A" by (blast dest: preal_downwards_closed [OF A]) - qed -qed - -subsection{*Existence of Inverse: Part 2*} - -lemma mem_Rep_preal_inverse_iff: - "(z \ Rep_preal(inverse R)) = - (0 < z \ (\y. z < y \ inverse y \ Rep_preal R))" -apply (simp add: preal_inverse_def mem_inverse_set Rep_preal) -apply (simp add: inverse_set_def) -done - -lemma Rep_preal_of_rat: - "0 < q ==> Rep_preal (preal_of_rat q) = {x. 0 < x \ x < q}" -by (simp add: preal_of_rat_def rat_mem_preal) - -lemma subset_inverse_mult_lemma: - assumes xpos: "0 < x" and xless: "x < 1" - shows "\r u y. 0 < r & r < y & inverse y \ Rep_preal R & - u \ Rep_preal R & x = r * u" -proof - - from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff) - from lemma_gleason9_36 [OF Rep_preal this] - obtain r where r: "r \ Rep_preal R" - and notin: "r * (inverse x) \ Rep_preal R" .. - have rpos: "0 Rep_preal R" and rless: "r < u" .. - have upos: "0 Rep_preal R" using notin - by (simp add: divide_inverse mult_commute) - show "u \ Rep_preal R" by (rule u) - show "x = x / u * u" using upos - by (simp add: divide_inverse mult_commute) - qed -qed - -lemma subset_inverse_mult: - "Rep_preal(preal_of_rat 1) \ Rep_preal(inverse R * R)" -apply (auto simp add: Bex_def Rep_preal_of_rat mem_Rep_preal_inverse_iff - mem_Rep_preal_mult_iff) -apply (blast dest: subset_inverse_mult_lemma) -done - -lemma inverse_mult_subset_lemma: - assumes rpos: "0 < r" - and rless: "r < y" - and notin: "inverse y \ Rep_preal R" - and q: "q \ Rep_preal R" - shows "r*q < 1" -proof - - have "q < inverse y" using rpos rless - by (simp add: not_in_preal_ub [OF Rep_preal notin] q) - hence "r * q < r/y" using rpos - by (simp add: divide_inverse mult_less_cancel_left) - also have "... \ 1" using rpos rless - by (simp add: pos_divide_le_eq) - finally show ?thesis . -qed - -lemma inverse_mult_subset: - "Rep_preal(inverse R * R) \ Rep_preal(preal_of_rat 1)" -apply (auto simp add: Bex_def Rep_preal_of_rat mem_Rep_preal_inverse_iff - mem_Rep_preal_mult_iff) -apply (simp add: zero_less_mult_iff preal_imp_pos [OF Rep_preal]) -apply (blast intro: inverse_mult_subset_lemma) -done - -lemma preal_mult_inverse: "inverse R * R = (1::preal)" -unfolding preal_one_def -apply (rule Rep_preal_inject [THEN iffD1]) -apply (rule equalityI [OF inverse_mult_subset subset_inverse_mult]) -done - -lemma preal_mult_inverse_right: "R * inverse R = (1::preal)" -apply (rule preal_mult_commute [THEN subst]) -apply (rule preal_mult_inverse) -done - - -text{*Theorems needing @{text Gleason9_34}*} - -lemma Rep_preal_self_subset: "Rep_preal (R) \ Rep_preal(R + S)" -proof - fix r - assume r: "r \ Rep_preal R" - have rpos: "0 Rep_preal S" .. - have ypos: "0 Rep_preal(R + S)" using r y - by (auto simp add: mem_Rep_preal_add_iff) - show "r \ Rep_preal(R + S)" using r ypos rpos - by (simp add: preal_downwards_closed [OF Rep_preal ry]) -qed - -lemma Rep_preal_sum_not_subset: "~ Rep_preal (R + S) \ Rep_preal(R)" -proof - - from mem_Rep_preal_Ex - obtain y where y: "y \ Rep_preal S" .. - have ypos: "0 Rep_preal R" and notin: "r + y \ Rep_preal R" .. - have "r + y \ Rep_preal (R + S)" using r y - by (auto simp add: mem_Rep_preal_add_iff) - thus ?thesis using notin by blast -qed - -lemma Rep_preal_sum_not_eq: "Rep_preal (R + S) \ Rep_preal(R)" -by (insert Rep_preal_sum_not_subset, blast) - -text{*at last, Gleason prop. 9-3.5(iii) page 123*} -lemma preal_self_less_add_left: "(R::preal) < R + S" -apply (unfold preal_less_def less_le) -apply (simp add: Rep_preal_self_subset Rep_preal_sum_not_eq [THEN not_sym]) -done - -lemma preal_self_less_add_right: "(R::preal) < S + R" -by (simp add: preal_add_commute preal_self_less_add_left) - -lemma preal_not_eq_self: "x \ x + (y::preal)" -by (insert preal_self_less_add_left [of x y], auto) - - -subsection{*Subtraction for Positive Reals*} - -text{*Gleason prop. 9-3.5(iv), page 123: proving @{prop "A < B ==> \D. A + D = -B"}. We define the claimed @{term D} and show that it is a positive real*} - -text{*Part 1 of Dedekind sections definition*} -lemma diff_set_not_empty: - "R < S ==> {} \ diff_set (Rep_preal S) (Rep_preal R)" -apply (auto simp add: preal_less_def diff_set_def elim!: equalityE) -apply (frule_tac x1 = S in Rep_preal [THEN preal_exists_greater]) -apply (drule preal_imp_pos [OF Rep_preal], clarify) -apply (cut_tac a=x and b=u in add_eq_exists, force) -done - -text{*Part 2 of Dedekind sections definition*} -lemma diff_set_nonempty: - "\q. 0 < q & q \ diff_set (Rep_preal S) (Rep_preal R)" -apply (cut_tac X = S in Rep_preal_exists_bound) -apply (erule exE) -apply (rule_tac x = x in exI, auto) -apply (simp add: diff_set_def) -apply (auto dest: Rep_preal [THEN preal_downwards_closed]) -done - -lemma diff_set_not_rat_set: - "diff_set (Rep_preal S) (Rep_preal R) < {r. 0 < r}" (is "?lhs < ?rhs") -proof - show "?lhs \ ?rhs" by (auto simp add: diff_set_def) - show "?lhs \ ?rhs" using diff_set_nonempty by blast -qed - -text{*Part 3 of Dedekind sections definition*} -lemma diff_set_lemma3: - "[|R < S; u \ diff_set (Rep_preal S) (Rep_preal R); 0 < z; z < u|] - ==> z \ diff_set (Rep_preal S) (Rep_preal R)" -apply (auto simp add: diff_set_def) -apply (rule_tac x=x in exI) -apply (drule Rep_preal [THEN preal_downwards_closed], auto) -done - -text{*Part 4 of Dedekind sections definition*} -lemma diff_set_lemma4: - "[|R < S; y \ diff_set (Rep_preal S) (Rep_preal R)|] - ==> \u \ diff_set (Rep_preal S) (Rep_preal R). y < u" -apply (auto simp add: diff_set_def) -apply (drule Rep_preal [THEN preal_exists_greater], clarify) -apply (cut_tac a="x+y" and b=u in add_eq_exists, clarify) -apply (rule_tac x="y+xa" in exI) -apply (auto simp add: add_ac) -done - -lemma mem_diff_set: - "R < S ==> diff_set (Rep_preal S) (Rep_preal R) \ preal" -apply (unfold preal_def cut_def) -apply (blast intro!: diff_set_not_empty diff_set_not_rat_set - diff_set_lemma3 diff_set_lemma4) -done - -lemma mem_Rep_preal_diff_iff: - "R < S ==> - (z \ Rep_preal(S-R)) = - (\x. 0 < x & 0 < z & x \ Rep_preal R & x + z \ Rep_preal S)" -apply (simp add: preal_diff_def mem_diff_set Rep_preal) -apply (force simp add: diff_set_def) -done - - -text{*proving that @{term "R + D \ S"}*} - -lemma less_add_left_lemma: - assumes Rless: "R < S" - and a: "a \ Rep_preal R" - and cb: "c + b \ Rep_preal S" - and "c \ Rep_preal R" - and "0 < b" - and "0 < c" - shows "a + b \ Rep_preal S" -proof - - have "0 R + (S-R) \ S" -apply (auto simp add: Bex_def preal_le_def mem_Rep_preal_add_iff - mem_Rep_preal_diff_iff) -apply (blast intro: less_add_left_lemma) -done - -subsection{*proving that @{term "S \ R + D"} --- trickier*} - -lemma lemma_sum_mem_Rep_preal_ex: - "x \ Rep_preal S ==> \e. 0 < e & x + e \ Rep_preal S" -apply (drule Rep_preal [THEN preal_exists_greater], clarify) -apply (cut_tac a=x and b=u in add_eq_exists, auto) -done - -lemma less_add_left_lemma2: - assumes Rless: "R < S" - and x: "x \ Rep_preal S" - and xnot: "x \ Rep_preal R" - shows "\u v z. 0 < v & 0 < z & u \ Rep_preal R & z \ Rep_preal R & - z + v \ Rep_preal S & x = u + v" -proof - - have xpos: "0 Rep_preal S" by blast - from Gleason9_34 [OF Rep_preal epos] - obtain r where r: "r \ Rep_preal R" and notin: "r + e \ Rep_preal R" .. - with x xnot xpos have rless: "r < x" by (blast intro: not_in_Rep_preal_ub) - from add_eq_exists [of r x] - obtain y where eq: "x = r+y" by auto - show ?thesis - proof (intro exI conjI) - show "r \ Rep_preal R" by (rule r) - show "r + e \ Rep_preal R" by (rule notin) - show "r + e + y \ Rep_preal S" using xe eq by (simp add: add_ac) - show "x = r + y" by (simp add: eq) - show "0 < r + e" using epos preal_imp_pos [OF Rep_preal r] - by simp - show "0 < y" using rless eq by arith - qed -qed - -lemma less_add_left_le2: "R < (S::preal) ==> S \ R + (S-R)" -apply (auto simp add: preal_le_def) -apply (case_tac "x \ Rep_preal R") -apply (cut_tac Rep_preal_self_subset [of R], force) -apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_diff_iff) -apply (blast dest: less_add_left_lemma2) -done - -lemma less_add_left: "R < (S::preal) ==> R + (S-R) = S" -by (blast intro: antisym [OF less_add_left_le1 less_add_left_le2]) - -lemma less_add_left_Ex: "R < (S::preal) ==> \D. R + D = S" -by (fast dest: less_add_left) - -lemma preal_add_less2_mono1: "R < (S::preal) ==> R + T < S + T" -apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc) -apply (rule_tac y1 = D in preal_add_commute [THEN subst]) -apply (auto intro: preal_self_less_add_left simp add: preal_add_assoc [symmetric]) -done - -lemma preal_add_less2_mono2: "R < (S::preal) ==> T + R < T + S" -by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of T]) - -lemma preal_add_right_less_cancel: "R + T < S + T ==> R < (S::preal)" -apply (insert linorder_less_linear [of R S], auto) -apply (drule_tac R = S and T = T in preal_add_less2_mono1) -apply (blast dest: order_less_trans) -done - -lemma preal_add_left_less_cancel: "T + R < T + S ==> R < (S::preal)" -by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of T]) - -lemma preal_add_less_cancel_right: "((R::preal) + T < S + T) = (R < S)" -by (blast intro: preal_add_less2_mono1 preal_add_right_less_cancel) - -lemma preal_add_less_cancel_left: "(T + (R::preal) < T + S) = (R < S)" -by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel) - -lemma preal_add_le_cancel_right: "((R::preal) + T \ S + T) = (R \ S)" -by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_right) - -lemma preal_add_le_cancel_left: "(T + (R::preal) \ T + S) = (R \ S)" -by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_left) - -lemma preal_add_less_mono: - "[| x1 < y1; x2 < y2 |] ==> x1 + x2 < y1 + (y2::preal)" -apply (auto dest!: less_add_left_Ex simp add: preal_add_ac) -apply (rule preal_add_assoc [THEN subst]) -apply (rule preal_self_less_add_right) -done - -lemma preal_add_right_cancel: "(R::preal) + T = S + T ==> R = S" -apply (insert linorder_less_linear [of R S], safe) -apply (drule_tac [!] T = T in preal_add_less2_mono1, auto) -done - -lemma preal_add_left_cancel: "C + A = C + B ==> A = (B::preal)" -by (auto intro: preal_add_right_cancel simp add: preal_add_commute) - -lemma preal_add_left_cancel_iff: "(C + A = C + B) = ((A::preal) = B)" -by (fast intro: preal_add_left_cancel) - -lemma preal_add_right_cancel_iff: "(A + C = B + C) = ((A::preal) = B)" -by (fast intro: preal_add_right_cancel) - -lemmas preal_cancels = - preal_add_less_cancel_right preal_add_less_cancel_left - preal_add_le_cancel_right preal_add_le_cancel_left - preal_add_left_cancel_iff preal_add_right_cancel_iff - -instance preal :: ordered_cancel_ab_semigroup_add -proof - fix a b c :: preal - show "a + b = a + c \ b = c" by (rule preal_add_left_cancel) - show "a \ b \ c + a \ c + b" by (simp only: preal_add_le_cancel_left) -qed - - -subsection{*Completeness of type @{typ preal}*} - -text{*Prove that supremum is a cut*} - -text{*Part 1 of Dedekind sections definition*} - -lemma preal_sup_set_not_empty: - "P \ {} ==> {} \ (\X \ P. Rep_preal(X))" -apply auto -apply (cut_tac X = x in mem_Rep_preal_Ex, auto) -done - - -text{*Part 2 of Dedekind sections definition*} - -lemma preal_sup_not_exists: - "\X \ P. X \ Y ==> \q. 0 < q & q \ (\X \ P. Rep_preal(X))" -apply (cut_tac X = Y in Rep_preal_exists_bound) -apply (auto simp add: preal_le_def) -done - -lemma preal_sup_set_not_rat_set: - "\X \ P. X \ Y ==> (\X \ P. Rep_preal(X)) < {r. 0 < r}" -apply (drule preal_sup_not_exists) -apply (blast intro: preal_imp_pos [OF Rep_preal]) -done - -text{*Part 3 of Dedekind sections definition*} -lemma preal_sup_set_lemma3: - "[|P \ {}; \X \ P. X \ Y; u \ (\X \ P. Rep_preal(X)); 0 < z; z < u|] - ==> z \ (\X \ P. Rep_preal(X))" -by (auto elim: Rep_preal [THEN preal_downwards_closed]) - -text{*Part 4 of Dedekind sections definition*} -lemma preal_sup_set_lemma4: - "[|P \ {}; \X \ P. X \ Y; y \ (\X \ P. Rep_preal(X)) |] - ==> \u \ (\X \ P. Rep_preal(X)). y < u" -by (blast dest: Rep_preal [THEN preal_exists_greater]) - -lemma preal_sup: - "[|P \ {}; \X \ P. X \ Y|] ==> (\X \ P. Rep_preal(X)) \ preal" -apply (unfold preal_def cut_def) -apply (blast intro!: preal_sup_set_not_empty preal_sup_set_not_rat_set - preal_sup_set_lemma3 preal_sup_set_lemma4) -done - -lemma preal_psup_le: - "[| \X \ P. X \ Y; x \ P |] ==> x \ psup P" -apply (simp (no_asm_simp) add: preal_le_def) -apply (subgoal_tac "P \ {}") -apply (auto simp add: psup_def preal_sup) -done - -lemma psup_le_ub: "[| P \ {}; \X \ P. X \ Y |] ==> psup P \ Y" -apply (simp (no_asm_simp) add: preal_le_def) -apply (simp add: psup_def preal_sup) -apply (auto simp add: preal_le_def) -done - -text{*Supremum property*} -lemma preal_complete: - "[| P \ {}; \X \ P. X \ Y |] ==> (\X \ P. Z < X) = (Z < psup P)" -apply (simp add: preal_less_def psup_def preal_sup) -apply (auto simp add: preal_le_def) -apply (rename_tac U) -apply (cut_tac x = U and y = Z in linorder_less_linear) -apply (auto simp add: preal_less_def) -done - - -subsection{*The Embedding from @{typ rat} into @{typ preal}*} - -lemma preal_of_rat_add_lemma1: - "[|x < y + z; 0 < x; 0 < y|] ==> x * y * inverse (y + z) < (y::rat)" -apply (frule_tac c = "y * inverse (y + z) " in mult_strict_right_mono) -apply (simp add: zero_less_mult_iff) -apply (simp add: mult_ac) -done - -lemma preal_of_rat_add_lemma2: - assumes "u < x + y" - and "0 < x" - and "0 < y" - and "0 < u" - shows "\v w::rat. w < y & 0 < v & v < x & 0 < w & u = v + w" -proof (intro exI conjI) - show "u * x * inverse(x+y) < x" using prems - by (simp add: preal_of_rat_add_lemma1) - show "u * y * inverse(x+y) < y" using prems - by (simp add: preal_of_rat_add_lemma1 add_commute [of x]) - show "0 < u * x * inverse (x + y)" using prems - by (simp add: zero_less_mult_iff) - show "0 < u * y * inverse (x + y)" using prems - by (simp add: zero_less_mult_iff) - show "u = u * x * inverse (x + y) + u * y * inverse (x + y)" using prems - by (simp add: left_distrib [symmetric] right_distrib [symmetric] mult_ac) -qed - -lemma preal_of_rat_add: - "[| 0 < x; 0 < y|] - ==> preal_of_rat ((x::rat) + y) = preal_of_rat x + preal_of_rat y" -apply (unfold preal_of_rat_def preal_add_def) -apply (simp add: rat_mem_preal) -apply (rule_tac f = Abs_preal in arg_cong) -apply (auto simp add: add_set_def) -apply (blast dest: preal_of_rat_add_lemma2) -done - -lemma preal_of_rat_mult_lemma1: - "[|x < y; 0 < x; 0 < z|] ==> x * z * inverse y < (z::rat)" -apply (frule_tac c = "z * inverse y" in mult_strict_right_mono) -apply (simp add: zero_less_mult_iff) -apply (subgoal_tac "y * (z * inverse y) = z * (y * inverse y)") -apply (simp_all add: mult_ac) -done - -lemma preal_of_rat_mult_lemma2: - assumes xless: "x < y * z" - and xpos: "0 < x" - and ypos: "0 < y" - shows "x * z * inverse y * inverse z < (z::rat)" -proof - - have "0 < y * z" using prems by simp - hence zpos: "0 < z" using prems by (simp add: zero_less_mult_iff) - have "x * z * inverse y * inverse z = x * inverse y * (z * inverse z)" - by (simp add: mult_ac) - also have "... = x/y" using zpos - by (simp add: divide_inverse) - also from xless have "... < z" - by (simp add: pos_divide_less_eq [OF ypos] mult_commute) - finally show ?thesis . -qed - -lemma preal_of_rat_mult_lemma3: - assumes uless: "u < x * y" - and "0 < x" - and "0 < y" - and "0 < u" - shows "\v w::rat. v < x & w < y & 0 < v & 0 < w & u = v * w" -proof - - from dense [OF uless] - obtain r where "u < r" "r < x * y" by blast - thus ?thesis - proof (intro exI conjI) - show "u * x * inverse r < x" using prems - by (simp add: preal_of_rat_mult_lemma1) - show "r * y * inverse x * inverse y < y" using prems - by (simp add: preal_of_rat_mult_lemma2) - show "0 < u * x * inverse r" using prems - by (simp add: zero_less_mult_iff) - show "0 < r * y * inverse x * inverse y" using prems - by (simp add: zero_less_mult_iff) - have "u * x * inverse r * (r * y * inverse x * inverse y) = - u * (r * inverse r) * (x * inverse x) * (y * inverse y)" - by (simp only: mult_ac) - thus "u = u * x * inverse r * (r * y * inverse x * inverse y)" using prems - by simp - qed -qed - -lemma preal_of_rat_mult: - "[| 0 < x; 0 < y|] - ==> preal_of_rat ((x::rat) * y) = preal_of_rat x * preal_of_rat y" -apply (unfold preal_of_rat_def preal_mult_def) -apply (simp add: rat_mem_preal) -apply (rule_tac f = Abs_preal in arg_cong) -apply (auto simp add: zero_less_mult_iff mult_strict_mono mult_set_def) -apply (blast dest: preal_of_rat_mult_lemma3) -done - -lemma preal_of_rat_less_iff: - "[| 0 < x; 0 < y|] ==> (preal_of_rat x < preal_of_rat y) = (x < y)" -by (force simp add: preal_of_rat_def preal_less_def rat_mem_preal) - -lemma preal_of_rat_le_iff: - "[| 0 < x; 0 < y|] ==> (preal_of_rat x \ preal_of_rat y) = (x \ y)" -by (simp add: preal_of_rat_less_iff linorder_not_less [symmetric]) - -lemma preal_of_rat_eq_iff: - "[| 0 < x; 0 < y|] ==> (preal_of_rat x = preal_of_rat y) = (x = y)" -by (simp add: preal_of_rat_le_iff order_eq_iff) - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/RComplete.thy --- a/src/HOL/Real/RComplete.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1334 +0,0 @@ -(* Title : HOL/Real/RComplete.thy - ID : $Id$ - Author : Jacques D. Fleuriot, University of Edinburgh - Author : Larry Paulson, University of Cambridge - Author : Jeremy Avigad, Carnegie Mellon University - Author : Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen -*) - -header {* Completeness of the Reals; Floor and Ceiling Functions *} - -theory RComplete -imports Lubs RealDef -begin - -lemma real_sum_of_halves: "x/2 + x/2 = (x::real)" - by simp - - -subsection {* Completeness of Positive Reals *} - -text {* - Supremum property for the set of positive reals - - Let @{text "P"} be a non-empty set of positive reals, with an upper - bound @{text "y"}. Then @{text "P"} has a least upper bound - (written @{text "S"}). - - FIXME: Can the premise be weakened to @{text "\x \ P. x\ y"}? -*} - -lemma posreal_complete: - assumes positive_P: "\x \ P. (0::real) < x" - and not_empty_P: "\x. x \ P" - and upper_bound_Ex: "\y. \x \ P. xS. \y. (\x \ P. y < x) = (y < S)" -proof (rule exI, rule allI) - fix y - let ?pP = "{w. real_of_preal w \ P}" - - show "(\x\P. y < x) = (y < real_of_preal (psup ?pP))" - proof (cases "0 < y") - assume neg_y: "\ 0 < y" - show ?thesis - proof - assume "\x\P. y < x" - have "\x. y < real_of_preal x" - using neg_y by (rule real_less_all_real2) - thus "y < real_of_preal (psup ?pP)" .. - next - assume "y < real_of_preal (psup ?pP)" - obtain "x" where x_in_P: "x \ P" using not_empty_P .. - hence "0 < x" using positive_P by simp - hence "y < x" using neg_y by simp - thus "\x \ P. y < x" using x_in_P .. - qed - next - assume pos_y: "0 < y" - - then obtain py where y_is_py: "y = real_of_preal py" - by (auto simp add: real_gt_zero_preal_Ex) - - obtain a where "a \ P" using not_empty_P .. - with positive_P have a_pos: "0 < a" .. - then obtain pa where "a = real_of_preal pa" - by (auto simp add: real_gt_zero_preal_Ex) - hence "pa \ ?pP" using `a \ P` by auto - hence pP_not_empty: "?pP \ {}" by auto - - obtain sup where sup: "\x \ P. x < sup" - using upper_bound_Ex .. - from this and `a \ P` have "a < sup" .. - hence "0 < sup" using a_pos by arith - then obtain possup where "sup = real_of_preal possup" - by (auto simp add: real_gt_zero_preal_Ex) - hence "\X \ ?pP. X \ possup" - using sup by (auto simp add: real_of_preal_lessI) - with pP_not_empty have psup: "\Z. (\X \ ?pP. Z < X) = (Z < psup ?pP)" - by (rule preal_complete) - - show ?thesis - proof - assume "\x \ P. y < x" - then obtain x where x_in_P: "x \ P" and y_less_x: "y < x" .. - hence "0 < x" using pos_y by arith - then obtain px where x_is_px: "x = real_of_preal px" - by (auto simp add: real_gt_zero_preal_Ex) - - have py_less_X: "\X \ ?pP. py < X" - proof - show "py < px" using y_is_py and x_is_px and y_less_x - by (simp add: real_of_preal_lessI) - show "px \ ?pP" using x_in_P and x_is_px by simp - qed - - have "(\X \ ?pP. py < X) ==> (py < psup ?pP)" - using psup by simp - hence "py < psup ?pP" using py_less_X by simp - thus "y < real_of_preal (psup {w. real_of_preal w \ P})" - using y_is_py and pos_y by (simp add: real_of_preal_lessI) - next - assume y_less_psup: "y < real_of_preal (psup ?pP)" - - hence "py < psup ?pP" using y_is_py - by (simp add: real_of_preal_lessI) - then obtain "X" where py_less_X: "py < X" and X_in_pP: "X \ ?pP" - using psup by auto - then obtain x where x_is_X: "x = real_of_preal X" - by (simp add: real_gt_zero_preal_Ex) - hence "y < x" using py_less_X and y_is_py - by (simp add: real_of_preal_lessI) - - moreover have "x \ P" using x_is_X and X_in_pP by simp - - ultimately show "\ x \ P. y < x" .. - qed - qed -qed - -text {* - \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc. -*} - -lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)" - apply (frule isLub_isUb) - apply (frule_tac x = y in isLub_isUb) - apply (blast intro!: order_antisym dest!: isLub_le_isUb) - done - - -text {* - \medskip Completeness theorem for the positive reals (again). -*} - -lemma posreals_complete: - assumes positive_S: "\x \ S. 0 < x" - and not_empty_S: "\x. x \ S" - and upper_bound_Ex: "\u. isUb (UNIV::real set) S u" - shows "\t. isLub (UNIV::real set) S t" -proof - let ?pS = "{w. real_of_preal w \ S}" - - obtain u where "isUb UNIV S u" using upper_bound_Ex .. - hence sup: "\x \ S. x \ u" by (simp add: isUb_def setle_def) - - obtain x where x_in_S: "x \ S" using not_empty_S .. - hence x_gt_zero: "0 < x" using positive_S by simp - have "x \ u" using sup and x_in_S .. - hence "0 < u" using x_gt_zero by arith - - then obtain pu where u_is_pu: "u = real_of_preal pu" - by (auto simp add: real_gt_zero_preal_Ex) - - have pS_less_pu: "\pa \ ?pS. pa \ pu" - proof - fix pa - assume "pa \ ?pS" - then obtain a where "a \ S" and "a = real_of_preal pa" - by simp - moreover hence "a \ u" using sup by simp - ultimately show "pa \ pu" - using sup and u_is_pu by (simp add: real_of_preal_le_iff) - qed - - have "\y \ S. y \ real_of_preal (psup ?pS)" - proof - fix y - assume y_in_S: "y \ S" - hence "0 < y" using positive_S by simp - then obtain py where y_is_py: "y = real_of_preal py" - by (auto simp add: real_gt_zero_preal_Ex) - hence py_in_pS: "py \ ?pS" using y_in_S by simp - with pS_less_pu have "py \ psup ?pS" - by (rule preal_psup_le) - thus "y \ real_of_preal (psup ?pS)" - using y_is_py by (simp add: real_of_preal_le_iff) - qed - - moreover { - fix x - assume x_ub_S: "\y\S. y \ x" - have "real_of_preal (psup ?pS) \ x" - proof - - obtain "s" where s_in_S: "s \ S" using not_empty_S .. - hence s_pos: "0 < s" using positive_S by simp - - hence "\ ps. s = real_of_preal ps" by (simp add: real_gt_zero_preal_Ex) - then obtain "ps" where s_is_ps: "s = real_of_preal ps" .. - hence ps_in_pS: "ps \ {w. real_of_preal w \ S}" using s_in_S by simp - - from x_ub_S have "s \ x" using s_in_S .. - hence "0 < x" using s_pos by simp - hence "\ px. x = real_of_preal px" by (simp add: real_gt_zero_preal_Ex) - then obtain "px" where x_is_px: "x = real_of_preal px" .. - - have "\pe \ ?pS. pe \ px" - proof - fix pe - assume "pe \ ?pS" - hence "real_of_preal pe \ S" by simp - hence "real_of_preal pe \ x" using x_ub_S by simp - thus "pe \ px" using x_is_px by (simp add: real_of_preal_le_iff) - qed - - moreover have "?pS \ {}" using ps_in_pS by auto - ultimately have "(psup ?pS) \ px" by (simp add: psup_le_ub) - thus "real_of_preal (psup ?pS) \ x" using x_is_px by (simp add: real_of_preal_le_iff) - qed - } - ultimately show "isLub UNIV S (real_of_preal (psup ?pS))" - by (simp add: isLub_def leastP_def isUb_def setle_def setge_def) -qed - -text {* - \medskip reals Completeness (again!) -*} - -lemma reals_complete: - assumes notempty_S: "\X. X \ S" - and exists_Ub: "\Y. isUb (UNIV::real set) S Y" - shows "\t. isLub (UNIV :: real set) S t" -proof - - obtain X where X_in_S: "X \ S" using notempty_S .. - obtain Y where Y_isUb: "isUb (UNIV::real set) S Y" - using exists_Ub .. - let ?SHIFT = "{z. \x \S. z = x + (-X) + 1} \ {x. 0 < x}" - - { - fix x - assume "isUb (UNIV::real set) S x" - hence S_le_x: "\ y \ S. y <= x" - by (simp add: isUb_def setle_def) - { - fix s - assume "s \ {z. \x\S. z = x + - X + 1}" - hence "\ x \ S. s = x + -X + 1" .. - then obtain x1 where "x1 \ S" and "s = x1 + (-X) + 1" .. - moreover hence "x1 \ x" using S_le_x by simp - ultimately have "s \ x + - X + 1" by arith - } - then have "isUb (UNIV::real set) ?SHIFT (x + (-X) + 1)" - by (auto simp add: isUb_def setle_def) - } note S_Ub_is_SHIFT_Ub = this - - hence "isUb UNIV ?SHIFT (Y + (-X) + 1)" using Y_isUb by simp - hence "\Z. isUb UNIV ?SHIFT Z" .. - moreover have "\y \ ?SHIFT. 0 < y" by auto - moreover have shifted_not_empty: "\u. u \ ?SHIFT" - using X_in_S and Y_isUb by auto - ultimately obtain t where t_is_Lub: "isLub UNIV ?SHIFT t" - using posreals_complete [of ?SHIFT] by blast - - show ?thesis - proof - show "isLub UNIV S (t + X + (-1))" - proof (rule isLubI2) - { - fix x - assume "isUb (UNIV::real set) S x" - hence "isUb (UNIV::real set) (?SHIFT) (x + (-X) + 1)" - using S_Ub_is_SHIFT_Ub by simp - hence "t \ (x + (-X) + 1)" - using t_is_Lub by (simp add: isLub_le_isUb) - hence "t + X + -1 \ x" by arith - } - then show "(t + X + -1) <=* Collect (isUb UNIV S)" - by (simp add: setgeI) - next - show "isUb UNIV S (t + X + -1)" - proof - - { - fix y - assume y_in_S: "y \ S" - have "y \ t + X + -1" - proof - - obtain "u" where u_in_shift: "u \ ?SHIFT" using shifted_not_empty .. - hence "\ x \ S. u = x + - X + 1" by simp - then obtain "x" where x_and_u: "u = x + - X + 1" .. - have u_le_t: "u \ t" using u_in_shift and t_is_Lub by (simp add: isLubD2) - - show ?thesis - proof cases - assume "y \ x" - moreover have "x = u + X + - 1" using x_and_u by arith - moreover have "u + X + - 1 \ t + X + -1" using u_le_t by arith - ultimately show "y \ t + X + -1" by arith - next - assume "~(y \ x)" - hence x_less_y: "x < y" by arith - - have "x + (-X) + 1 \ ?SHIFT" using x_and_u and u_in_shift by simp - hence "0 < x + (-X) + 1" by simp - hence "0 < y + (-X) + 1" using x_less_y by arith - hence "y + (-X) + 1 \ ?SHIFT" using y_in_S by simp - hence "y + (-X) + 1 \ t" using t_is_Lub by (simp add: isLubD2) - thus ?thesis by simp - qed - qed - } - then show ?thesis by (simp add: isUb_def setle_def) - qed - qed - qed -qed - - -subsection {* The Archimedean Property of the Reals *} - -theorem reals_Archimedean: - assumes x_pos: "0 < x" - shows "\n. inverse (real (Suc n)) < x" -proof (rule ccontr) - assume contr: "\ ?thesis" - have "\n. x * real (Suc n) <= 1" - proof - fix n - from contr have "x \ inverse (real (Suc n))" - by (simp add: linorder_not_less) - hence "x \ (1 / (real (Suc n)))" - by (simp add: inverse_eq_divide) - moreover have "0 \ real (Suc n)" - by (rule real_of_nat_ge_zero) - ultimately have "x * real (Suc n) \ (1 / real (Suc n)) * real (Suc n)" - by (rule mult_right_mono) - thus "x * real (Suc n) \ 1" by simp - qed - hence "{z. \n. z = x * (real (Suc n))} *<= 1" - by (simp add: setle_def, safe, rule spec) - hence "isUb (UNIV::real set) {z. \n. z = x * (real (Suc n))} 1" - by (simp add: isUbI) - hence "\Y. isUb (UNIV::real set) {z. \n. z = x* (real (Suc n))} Y" .. - moreover have "\X. X \ {z. \n. z = x* (real (Suc n))}" by auto - ultimately have "\t. isLub UNIV {z. \n. z = x * real (Suc n)} t" - by (simp add: reals_complete) - then obtain "t" where - t_is_Lub: "isLub UNIV {z. \n. z = x * real (Suc n)} t" .. - - have "\n::nat. x * real n \ t + - x" - proof - fix n - from t_is_Lub have "x * real (Suc n) \ t" - by (simp add: isLubD2) - hence "x * (real n) + x \ t" - by (simp add: right_distrib real_of_nat_Suc) - thus "x * (real n) \ t + - x" by arith - qed - - hence "\m. x * real (Suc m) \ t + - x" by simp - hence "{z. \n. z = x * (real (Suc n))} *<= (t + - x)" - by (auto simp add: setle_def) - hence "isUb (UNIV::real set) {z. \n. z = x * (real (Suc n))} (t + (-x))" - by (simp add: isUbI) - hence "t \ t + - x" - using t_is_Lub by (simp add: isLub_le_isUb) - thus False using x_pos by arith -qed - -text {* - There must be other proofs, e.g. @{text "Suc"} of the largest - integer in the cut representing @{text "x"}. -*} - -lemma reals_Archimedean2: "\n. (x::real) < real (n::nat)" -proof cases - assume "x \ 0" - hence "x < real (1::nat)" by simp - thus ?thesis .. -next - assume "\ x \ 0" - hence x_greater_zero: "0 < x" by simp - hence "0 < inverse x" by simp - then obtain n where "inverse (real (Suc n)) < inverse x" - using reals_Archimedean by blast - hence "inverse (real (Suc n)) * x < inverse x * x" - using x_greater_zero by (rule mult_strict_right_mono) - hence "inverse (real (Suc n)) * x < 1" - using x_greater_zero by simp - hence "real (Suc n) * (inverse (real (Suc n)) * x) < real (Suc n) * 1" - by (rule mult_strict_left_mono) simp - hence "x < real (Suc n)" - by (simp add: ring_simps) - thus "\(n::nat). x < real n" .. -qed - -lemma reals_Archimedean3: - assumes x_greater_zero: "0 < x" - shows "\(y::real). \(n::nat). y < real n * x" -proof - fix y - have x_not_zero: "x \ 0" using x_greater_zero by simp - obtain n where "y * inverse x < real (n::nat)" - using reals_Archimedean2 .. - hence "y * inverse x * x < real n * x" - using x_greater_zero by (simp add: mult_strict_right_mono) - hence "x * inverse x * y < x * real n" - by (simp add: ring_simps) - hence "y < real (n::nat) * x" - using x_not_zero by (simp add: ring_simps) - thus "\(n::nat). y < real n * x" .. -qed - -lemma reals_Archimedean6: - "0 \ r ==> \(n::nat). real (n - 1) \ r & r < real (n)" -apply (insert reals_Archimedean2 [of r], safe) -apply (subgoal_tac "\x::nat. r < real x \ (\y. r < real y \ x \ y)", auto) -apply (rule_tac x = x in exI) -apply (case_tac x, simp) -apply (rename_tac x') -apply (drule_tac x = x' in spec, simp) -apply (rule_tac x="LEAST n. r < real n" in exI, safe) -apply (erule LeastI, erule Least_le) -done - -lemma reals_Archimedean6a: "0 \ r ==> \n. real (n) \ r & r < real (Suc n)" - by (drule reals_Archimedean6) auto - -lemma reals_Archimedean_6b_int: - "0 \ r ==> \n::int. real n \ r & r < real (n+1)" -apply (drule reals_Archimedean6a, auto) -apply (rule_tac x = "int n" in exI) -apply (simp add: real_of_int_real_of_nat real_of_nat_Suc) -done - -lemma reals_Archimedean_6c_int: - "r < 0 ==> \n::int. real n \ r & r < real (n+1)" -apply (rule reals_Archimedean_6b_int [of "-r", THEN exE], simp, auto) -apply (rename_tac n) -apply (drule order_le_imp_less_or_eq, auto) -apply (rule_tac x = "- n - 1" in exI) -apply (rule_tac [2] x = "- n" in exI, auto) -done - - -subsection{*Density of the Rational Reals in the Reals*} - -text{* This density proof is due to Stefan Richter and was ported by TN. The -original source is \emph{Real Analysis} by H.L. Royden. -It employs the Archimedean property of the reals. *} - -lemma Rats_dense_in_nn_real: fixes x::real -assumes "0\x" and "xr \ \. x r "LEAST n. y \ real (Suc n)/real q" - from reals_Archimedean2 obtain n::nat where "y * real q < real n" by auto - with `0 < real q` have ex: "y \ real n/real q" (is "?P n") - by (simp add: pos_less_divide_eq[THEN sym]) - also from assms have "\ y \ real (0::nat) / real q" by simp - ultimately have main: "(LEAST n. y \ real n/real q) = Suc p" - by (unfold p_def) (rule Least_Suc) - also from ex have "?P (LEAST x. ?P x)" by (rule LeastI) - ultimately have suc: "y \ real (Suc p) / real q" by simp - def r \ "real p/real q" - have "x = y-(y-x)" by simp - also from suc q have "\ < real (Suc p)/real q - inverse (real q)" by arith - also have "\ = real p / real q" - by (simp only: inverse_eq_divide real_diff_def real_of_nat_Suc - minus_divide_left add_divide_distrib[THEN sym]) simp - finally have "x ?P p" by (rule not_less_Least) - hence "r \" by simp - with `xr \ \. x r x + real n" by arith - also from `xr \ \. x + real n < r \ r < y + real n" - by(rule Rats_dense_in_nn_real) - then obtain r where "r \ \" and r2: "x + real n < r" - and r3: "r < y + real n" - by blast - have "r - real n = r + real (int n)/real (-1::int)" by simp - also from `r\\` have "r + real (int n)/real (-1::int) \ \" by simp - also from r2 have "x < r - real n" by arith - moreover from r3 have "r - real n < y" by arith - ultimately show ?thesis by fast -qed - - -subsection{*Floor and Ceiling Functions from the Reals to the Integers*} - -definition - floor :: "real => int" where - [code del]: "floor r = (LEAST n::int. r < real (n+1))" - -definition - ceiling :: "real => int" where - "ceiling r = - floor (- r)" - -notation (xsymbols) - floor ("\_\") and - ceiling ("\_\") - -notation (HTML output) - floor ("\_\") and - ceiling ("\_\") - - -lemma number_of_less_real_of_int_iff [simp]: - "((number_of n) < real (m::int)) = (number_of n < m)" -apply auto -apply (rule real_of_int_less_iff [THEN iffD1]) -apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto) -done - -lemma number_of_less_real_of_int_iff2 [simp]: - "(real (m::int) < (number_of n)) = (m < number_of n)" -apply auto -apply (rule real_of_int_less_iff [THEN iffD1]) -apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto) -done - -lemma number_of_le_real_of_int_iff [simp]: - "((number_of n) \ real (m::int)) = (number_of n \ m)" -by (simp add: linorder_not_less [symmetric]) - -lemma number_of_le_real_of_int_iff2 [simp]: - "(real (m::int) \ (number_of n)) = (m \ number_of n)" -by (simp add: linorder_not_less [symmetric]) - -lemma floor_zero [simp]: "floor 0 = 0" -apply (simp add: floor_def del: real_of_int_add) -apply (rule Least_equality) -apply simp_all -done - -lemma floor_real_of_nat_zero [simp]: "floor (real (0::nat)) = 0" -by auto - -lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n" -apply (simp only: floor_def) -apply (rule Least_equality) -apply (drule_tac [2] real_of_int_of_nat_eq [THEN ssubst]) -apply (drule_tac [2] real_of_int_less_iff [THEN iffD1]) -apply simp_all -done - -lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n" -apply (simp only: floor_def) -apply (rule Least_equality) -apply (drule_tac [2] real_of_int_of_nat_eq [THEN ssubst]) -apply (drule_tac [2] real_of_int_minus [THEN sym, THEN subst]) -apply (drule_tac [2] real_of_int_less_iff [THEN iffD1]) -apply simp_all -done - -lemma floor_real_of_int [simp]: "floor (real (n::int)) = n" -apply (simp only: floor_def) -apply (rule Least_equality) -apply auto -done - -lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n" -apply (simp only: floor_def) -apply (rule Least_equality) -apply (drule_tac [2] real_of_int_minus [THEN sym, THEN subst]) -apply auto -done - -lemma real_lb_ub_int: " \n::int. real n \ r & r < real (n+1)" -apply (case_tac "r < 0") -apply (blast intro: reals_Archimedean_6c_int) -apply (simp only: linorder_not_less) -apply (blast intro: reals_Archimedean_6b_int reals_Archimedean_6c_int) -done - -lemma lemma_floor: - assumes a1: "real m \ r" and a2: "r < real n + 1" - shows "m \ (n::int)" -proof - - have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans) - also have "... = real (n + 1)" by simp - finally have "m < n + 1" by (simp only: real_of_int_less_iff) - thus ?thesis by arith -qed - -lemma real_of_int_floor_le [simp]: "real (floor r) \ r" -apply (simp add: floor_def Least_def) -apply (insert real_lb_ub_int [of r], safe) -apply (rule theI2) -apply auto -done - -lemma floor_mono: "x < y ==> floor x \ floor y" -apply (simp add: floor_def Least_def) -apply (insert real_lb_ub_int [of x]) -apply (insert real_lb_ub_int [of y], safe) -apply (rule theI2) -apply (rule_tac [3] theI2) -apply simp -apply (erule conjI) -apply (auto simp add: order_eq_iff int_le_real_less) -done - -lemma floor_mono2: "x \ y ==> floor x \ floor y" -by (auto dest: order_le_imp_less_or_eq simp add: floor_mono) - -lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \ x" -by (auto intro: lemma_floor) - -lemma real_of_int_floor_cancel [simp]: - "(real (floor x) = x) = (\n::int. x = real n)" -apply (simp add: floor_def Least_def) -apply (insert real_lb_ub_int [of x], erule exE) -apply (rule theI2) -apply (auto intro: lemma_floor) -done - -lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n" -apply (simp add: floor_def) -apply (rule Least_equality) -apply (auto intro: lemma_floor) -done - -lemma floor_eq2: "[| real n \ x; x < real n + 1 |] ==> floor x = n" -apply (simp add: floor_def) -apply (rule Least_equality) -apply (auto intro: lemma_floor) -done - -lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n" -apply (rule inj_int [THEN injD]) -apply (simp add: real_of_nat_Suc) -apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"]) -done - -lemma floor_eq4: "[| real n \ x; x < real (Suc n) |] ==> nat(floor x) = n" -apply (drule order_le_imp_less_or_eq) -apply (auto intro: floor_eq3) -done - -lemma floor_number_of_eq [simp]: - "floor(number_of n :: real) = (number_of n :: int)" -apply (subst real_number_of [symmetric]) -apply (rule floor_real_of_int) -done - -lemma floor_one [simp]: "floor 1 = 1" - apply (rule trans) - prefer 2 - apply (rule floor_real_of_int) - apply simp -done - -lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \ real(floor r)" -apply (simp add: floor_def Least_def) -apply (insert real_lb_ub_int [of r], safe) -apply (rule theI2) -apply (auto intro: lemma_floor) -done - -lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)" -apply (simp add: floor_def Least_def) -apply (insert real_lb_ub_int [of r], safe) -apply (rule theI2) -apply (auto intro: lemma_floor) -done - -lemma real_of_int_floor_add_one_ge [simp]: "r \ real(floor r) + 1" -apply (insert real_of_int_floor_ge_diff_one [of r]) -apply (auto simp del: real_of_int_floor_ge_diff_one) -done - -lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1" -apply (insert real_of_int_floor_gt_diff_one [of r]) -apply (auto simp del: real_of_int_floor_gt_diff_one) -done - -lemma le_floor: "real a <= x ==> a <= floor x" - apply (subgoal_tac "a < floor x + 1") - apply arith - apply (subst real_of_int_less_iff [THEN sym]) - apply simp - apply (insert real_of_int_floor_add_one_gt [of x]) - apply arith -done - -lemma real_le_floor: "a <= floor x ==> real a <= x" - apply (rule order_trans) - prefer 2 - apply (rule real_of_int_floor_le) - apply (subst real_of_int_le_iff) - apply assumption -done - -lemma le_floor_eq: "(a <= floor x) = (real a <= x)" - apply (rule iffI) - apply (erule real_le_floor) - apply (erule le_floor) -done - -lemma le_floor_eq_number_of [simp]: - "(number_of n <= floor x) = (number_of n <= x)" -by (simp add: le_floor_eq) - -lemma le_floor_eq_zero [simp]: "(0 <= floor x) = (0 <= x)" -by (simp add: le_floor_eq) - -lemma le_floor_eq_one [simp]: "(1 <= floor x) = (1 <= x)" -by (simp add: le_floor_eq) - -lemma floor_less_eq: "(floor x < a) = (x < real a)" - apply (subst linorder_not_le [THEN sym])+ - apply simp - apply (rule le_floor_eq) -done - -lemma floor_less_eq_number_of [simp]: - "(floor x < number_of n) = (x < number_of n)" -by (simp add: floor_less_eq) - -lemma floor_less_eq_zero [simp]: "(floor x < 0) = (x < 0)" -by (simp add: floor_less_eq) - -lemma floor_less_eq_one [simp]: "(floor x < 1) = (x < 1)" -by (simp add: floor_less_eq) - -lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)" - apply (insert le_floor_eq [of "a + 1" x]) - apply auto -done - -lemma less_floor_eq_number_of [simp]: - "(number_of n < floor x) = (number_of n + 1 <= x)" -by (simp add: less_floor_eq) - -lemma less_floor_eq_zero [simp]: "(0 < floor x) = (1 <= x)" -by (simp add: less_floor_eq) - -lemma less_floor_eq_one [simp]: "(1 < floor x) = (2 <= x)" -by (simp add: less_floor_eq) - -lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)" - apply (insert floor_less_eq [of x "a + 1"]) - apply auto -done - -lemma floor_le_eq_number_of [simp]: - "(floor x <= number_of n) = (x < number_of n + 1)" -by (simp add: floor_le_eq) - -lemma floor_le_eq_zero [simp]: "(floor x <= 0) = (x < 1)" -by (simp add: floor_le_eq) - -lemma floor_le_eq_one [simp]: "(floor x <= 1) = (x < 2)" -by (simp add: floor_le_eq) - -lemma floor_add [simp]: "floor (x + real a) = floor x + a" - apply (subst order_eq_iff) - apply (rule conjI) - prefer 2 - apply (subgoal_tac "floor x + a < floor (x + real a) + 1") - apply arith - apply (subst real_of_int_less_iff [THEN sym]) - apply simp - apply (subgoal_tac "x + real a < real(floor(x + real a)) + 1") - apply (subgoal_tac "real (floor x) <= x") - apply arith - apply (rule real_of_int_floor_le) - apply (rule real_of_int_floor_add_one_gt) - apply (subgoal_tac "floor (x + real a) < floor x + a + 1") - apply arith - apply (subst real_of_int_less_iff [THEN sym]) - apply simp - apply (subgoal_tac "real(floor(x + real a)) <= x + real a") - apply (subgoal_tac "x < real(floor x) + 1") - apply arith - apply (rule real_of_int_floor_add_one_gt) - apply (rule real_of_int_floor_le) -done - -lemma floor_add_number_of [simp]: - "floor (x + number_of n) = floor x + number_of n" - apply (subst floor_add [THEN sym]) - apply simp -done - -lemma floor_add_one [simp]: "floor (x + 1) = floor x + 1" - apply (subst floor_add [THEN sym]) - apply simp -done - -lemma floor_subtract [simp]: "floor (x - real a) = floor x - a" - apply (subst diff_minus)+ - apply (subst real_of_int_minus [THEN sym]) - apply (rule floor_add) -done - -lemma floor_subtract_number_of [simp]: "floor (x - number_of n) = - floor x - number_of n" - apply (subst floor_subtract [THEN sym]) - apply simp -done - -lemma floor_subtract_one [simp]: "floor (x - 1) = floor x - 1" - apply (subst floor_subtract [THEN sym]) - apply simp -done - -lemma ceiling_zero [simp]: "ceiling 0 = 0" -by (simp add: ceiling_def) - -lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n" -by (simp add: ceiling_def) - -lemma ceiling_real_of_nat_zero [simp]: "ceiling (real (0::nat)) = 0" -by auto - -lemma ceiling_floor [simp]: "ceiling (real (floor r)) = floor r" -by (simp add: ceiling_def) - -lemma floor_ceiling [simp]: "floor (real (ceiling r)) = ceiling r" -by (simp add: ceiling_def) - -lemma real_of_int_ceiling_ge [simp]: "r \ real (ceiling r)" -apply (simp add: ceiling_def) -apply (subst le_minus_iff, simp) -done - -lemma ceiling_mono: "x < y ==> ceiling x \ ceiling y" -by (simp add: floor_mono ceiling_def) - -lemma ceiling_mono2: "x \ y ==> ceiling x \ ceiling y" -by (simp add: floor_mono2 ceiling_def) - -lemma real_of_int_ceiling_cancel [simp]: - "(real (ceiling x) = x) = (\n::int. x = real n)" -apply (auto simp add: ceiling_def) -apply (drule arg_cong [where f = uminus], auto) -apply (rule_tac x = "-n" in exI, auto) -done - -lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1" -apply (simp add: ceiling_def) -apply (rule minus_equation_iff [THEN iffD1]) -apply (simp add: floor_eq [where n = "-(n+1)"]) -done - -lemma ceiling_eq2: "[| real n < x; x \ real n + 1 |] ==> ceiling x = n + 1" -by (simp add: ceiling_def floor_eq2 [where n = "-(n+1)"]) - -lemma ceiling_eq3: "[| real n - 1 < x; x \ real n |] ==> ceiling x = n" -by (simp add: ceiling_def floor_eq2 [where n = "-n"]) - -lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n" -by (simp add: ceiling_def) - -lemma ceiling_number_of_eq [simp]: - "ceiling (number_of n :: real) = (number_of n)" -apply (subst real_number_of [symmetric]) -apply (rule ceiling_real_of_int) -done - -lemma ceiling_one [simp]: "ceiling 1 = 1" - by (unfold ceiling_def, simp) - -lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \ r" -apply (rule neg_le_iff_le [THEN iffD1]) -apply (simp add: ceiling_def diff_minus) -done - -lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \ r + 1" -apply (insert real_of_int_ceiling_diff_one_le [of r]) -apply (simp del: real_of_int_ceiling_diff_one_le) -done - -lemma ceiling_le: "x <= real a ==> ceiling x <= a" - apply (unfold ceiling_def) - apply (subgoal_tac "-a <= floor(- x)") - apply simp - apply (rule le_floor) - apply simp -done - -lemma ceiling_le_real: "ceiling x <= a ==> x <= real a" - apply (unfold ceiling_def) - apply (subgoal_tac "real(- a) <= - x") - apply simp - apply (rule real_le_floor) - apply simp -done - -lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)" - apply (rule iffI) - apply (erule ceiling_le_real) - apply (erule ceiling_le) -done - -lemma ceiling_le_eq_number_of [simp]: - "(ceiling x <= number_of n) = (x <= number_of n)" -by (simp add: ceiling_le_eq) - -lemma ceiling_le_zero_eq [simp]: "(ceiling x <= 0) = (x <= 0)" -by (simp add: ceiling_le_eq) - -lemma ceiling_le_eq_one [simp]: "(ceiling x <= 1) = (x <= 1)" -by (simp add: ceiling_le_eq) - -lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)" - apply (subst linorder_not_le [THEN sym])+ - apply simp - apply (rule ceiling_le_eq) -done - -lemma less_ceiling_eq_number_of [simp]: - "(number_of n < ceiling x) = (number_of n < x)" -by (simp add: less_ceiling_eq) - -lemma less_ceiling_eq_zero [simp]: "(0 < ceiling x) = (0 < x)" -by (simp add: less_ceiling_eq) - -lemma less_ceiling_eq_one [simp]: "(1 < ceiling x) = (1 < x)" -by (simp add: less_ceiling_eq) - -lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)" - apply (insert ceiling_le_eq [of x "a - 1"]) - apply auto -done - -lemma ceiling_less_eq_number_of [simp]: - "(ceiling x < number_of n) = (x <= number_of n - 1)" -by (simp add: ceiling_less_eq) - -lemma ceiling_less_eq_zero [simp]: "(ceiling x < 0) = (x <= -1)" -by (simp add: ceiling_less_eq) - -lemma ceiling_less_eq_one [simp]: "(ceiling x < 1) = (x <= 0)" -by (simp add: ceiling_less_eq) - -lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)" - apply (insert less_ceiling_eq [of "a - 1" x]) - apply auto -done - -lemma le_ceiling_eq_number_of [simp]: - "(number_of n <= ceiling x) = (number_of n - 1 < x)" -by (simp add: le_ceiling_eq) - -lemma le_ceiling_eq_zero [simp]: "(0 <= ceiling x) = (-1 < x)" -by (simp add: le_ceiling_eq) - -lemma le_ceiling_eq_one [simp]: "(1 <= ceiling x) = (0 < x)" -by (simp add: le_ceiling_eq) - -lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a" - apply (unfold ceiling_def, simp) - apply (subst real_of_int_minus [THEN sym]) - apply (subst floor_add) - apply simp -done - -lemma ceiling_add_number_of [simp]: "ceiling (x + number_of n) = - ceiling x + number_of n" - apply (subst ceiling_add [THEN sym]) - apply simp -done - -lemma ceiling_add_one [simp]: "ceiling (x + 1) = ceiling x + 1" - apply (subst ceiling_add [THEN sym]) - apply simp -done - -lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a" - apply (subst diff_minus)+ - apply (subst real_of_int_minus [THEN sym]) - apply (rule ceiling_add) -done - -lemma ceiling_subtract_number_of [simp]: "ceiling (x - number_of n) = - ceiling x - number_of n" - apply (subst ceiling_subtract [THEN sym]) - apply simp -done - -lemma ceiling_subtract_one [simp]: "ceiling (x - 1) = ceiling x - 1" - apply (subst ceiling_subtract [THEN sym]) - apply simp -done - -subsection {* Versions for the natural numbers *} - -definition - natfloor :: "real => nat" where - "natfloor x = nat(floor x)" - -definition - natceiling :: "real => nat" where - "natceiling x = nat(ceiling x)" - -lemma natfloor_zero [simp]: "natfloor 0 = 0" - by (unfold natfloor_def, simp) - -lemma natfloor_one [simp]: "natfloor 1 = 1" - by (unfold natfloor_def, simp) - -lemma zero_le_natfloor [simp]: "0 <= natfloor x" - by (unfold natfloor_def, simp) - -lemma natfloor_number_of_eq [simp]: "natfloor (number_of n) = number_of n" - by (unfold natfloor_def, simp) - -lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n" - by (unfold natfloor_def, simp) - -lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x" - by (unfold natfloor_def, simp) - -lemma natfloor_neg: "x <= 0 ==> natfloor x = 0" - apply (unfold natfloor_def) - apply (subgoal_tac "floor x <= floor 0") - apply simp - apply (erule floor_mono2) -done - -lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y" - apply (case_tac "0 <= x") - apply (subst natfloor_def)+ - apply (subst nat_le_eq_zle) - apply force - apply (erule floor_mono2) - apply (subst natfloor_neg) - apply simp - apply simp -done - -lemma le_natfloor: "real x <= a ==> x <= natfloor a" - apply (unfold natfloor_def) - apply (subst nat_int [THEN sym]) - apply (subst nat_le_eq_zle) - apply simp - apply (rule le_floor) - apply simp -done - -lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)" - apply (rule iffI) - apply (rule order_trans) - prefer 2 - apply (erule real_natfloor_le) - apply (subst real_of_nat_le_iff) - apply assumption - apply (erule le_natfloor) -done - -lemma le_natfloor_eq_number_of [simp]: - "~ neg((number_of n)::int) ==> 0 <= x ==> - (number_of n <= natfloor x) = (number_of n <= x)" - apply (subst le_natfloor_eq, assumption) - apply simp -done - -lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)" - apply (case_tac "0 <= x") - apply (subst le_natfloor_eq, assumption, simp) - apply (rule iffI) - apply (subgoal_tac "natfloor x <= natfloor 0") - apply simp - apply (rule natfloor_mono) - apply simp - apply simp -done - -lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n" - apply (unfold natfloor_def) - apply (subst nat_int [THEN sym]);back; - apply (subst eq_nat_nat_iff) - apply simp - apply simp - apply (rule floor_eq2) - apply auto -done - -lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1" - apply (case_tac "0 <= x") - apply (unfold natfloor_def) - apply simp - apply simp_all -done - -lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)" - apply (simp add: compare_rls) - apply (rule real_natfloor_add_one_gt) -done - -lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n" - apply (subgoal_tac "z < real(natfloor z) + 1") - apply arith - apply (rule real_natfloor_add_one_gt) -done - -lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a" - apply (unfold natfloor_def) - apply (subgoal_tac "real a = real (int a)") - apply (erule ssubst) - apply (simp add: nat_add_distrib del: real_of_int_of_nat_eq) - apply simp -done - -lemma natfloor_add_number_of [simp]: - "~neg ((number_of n)::int) ==> 0 <= x ==> - natfloor (x + number_of n) = natfloor x + number_of n" - apply (subst natfloor_add [THEN sym]) - apply simp_all -done - -lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1" - apply (subst natfloor_add [THEN sym]) - apply assumption - apply simp -done - -lemma natfloor_subtract [simp]: "real a <= x ==> - natfloor(x - real a) = natfloor x - a" - apply (unfold natfloor_def) - apply (subgoal_tac "real a = real (int a)") - apply (erule ssubst) - apply (simp del: real_of_int_of_nat_eq) - apply simp -done - -lemma natceiling_zero [simp]: "natceiling 0 = 0" - by (unfold natceiling_def, simp) - -lemma natceiling_one [simp]: "natceiling 1 = 1" - by (unfold natceiling_def, simp) - -lemma zero_le_natceiling [simp]: "0 <= natceiling x" - by (unfold natceiling_def, simp) - -lemma natceiling_number_of_eq [simp]: "natceiling (number_of n) = number_of n" - by (unfold natceiling_def, simp) - -lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n" - by (unfold natceiling_def, simp) - -lemma real_natceiling_ge: "x <= real(natceiling x)" - apply (unfold natceiling_def) - apply (case_tac "x < 0") - apply simp - apply (subst real_nat_eq_real) - apply (subgoal_tac "ceiling 0 <= ceiling x") - apply simp - apply (rule ceiling_mono2) - apply simp - apply simp -done - -lemma natceiling_neg: "x <= 0 ==> natceiling x = 0" - apply (unfold natceiling_def) - apply simp -done - -lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y" - apply (case_tac "0 <= x") - apply (subst natceiling_def)+ - apply (subst nat_le_eq_zle) - apply (rule disjI2) - apply (subgoal_tac "real (0::int) <= real(ceiling y)") - apply simp - apply (rule order_trans) - apply simp - apply (erule order_trans) - apply simp - apply (erule ceiling_mono2) - apply (subst natceiling_neg) - apply simp_all -done - -lemma natceiling_le: "x <= real a ==> natceiling x <= a" - apply (unfold natceiling_def) - apply (case_tac "x < 0") - apply simp - apply (subst nat_int [THEN sym]);back; - apply (subst nat_le_eq_zle) - apply simp - apply (rule ceiling_le) - apply simp -done - -lemma natceiling_le_eq: "0 <= x ==> (natceiling x <= a) = (x <= real a)" - apply (rule iffI) - apply (rule order_trans) - apply (rule real_natceiling_ge) - apply (subst real_of_nat_le_iff) - apply assumption - apply (erule natceiling_le) -done - -lemma natceiling_le_eq_number_of [simp]: - "~ neg((number_of n)::int) ==> 0 <= x ==> - (natceiling x <= number_of n) = (x <= number_of n)" - apply (subst natceiling_le_eq, assumption) - apply simp -done - -lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)" - apply (case_tac "0 <= x") - apply (subst natceiling_le_eq) - apply assumption - apply simp - apply (subst natceiling_neg) - apply simp - apply simp -done - -lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1" - apply (unfold natceiling_def) - apply (simplesubst nat_int [THEN sym]) back back - apply (subgoal_tac "nat(int n) + 1 = nat(int n + 1)") - apply (erule ssubst) - apply (subst eq_nat_nat_iff) - apply (subgoal_tac "ceiling 0 <= ceiling x") - apply simp - apply (rule ceiling_mono2) - apply force - apply force - apply (rule ceiling_eq2) - apply (simp, simp) - apply (subst nat_add_distrib) - apply auto -done - -lemma natceiling_add [simp]: "0 <= x ==> - natceiling (x + real a) = natceiling x + a" - apply (unfold natceiling_def) - apply (subgoal_tac "real a = real (int a)") - apply (erule ssubst) - apply (simp del: real_of_int_of_nat_eq) - apply (subst nat_add_distrib) - apply (subgoal_tac "0 = ceiling 0") - apply (erule ssubst) - apply (erule ceiling_mono2) - apply simp_all -done - -lemma natceiling_add_number_of [simp]: - "~ neg ((number_of n)::int) ==> 0 <= x ==> - natceiling (x + number_of n) = natceiling x + number_of n" - apply (subst natceiling_add [THEN sym]) - apply simp_all -done - -lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1" - apply (subst natceiling_add [THEN sym]) - apply assumption - apply simp -done - -lemma natceiling_subtract [simp]: "real a <= x ==> - natceiling(x - real a) = natceiling x - a" - apply (unfold natceiling_def) - apply (subgoal_tac "real a = real (int a)") - apply (erule ssubst) - apply (simp del: real_of_int_of_nat_eq) - apply simp -done - -lemma natfloor_div_nat: "1 <= x ==> y > 0 ==> - natfloor (x / real y) = natfloor x div y" -proof - - assume "1 <= (x::real)" and "(y::nat) > 0" - have "natfloor x = (natfloor x) div y * y + (natfloor x) mod y" - by simp - then have a: "real(natfloor x) = real ((natfloor x) div y) * real y + - real((natfloor x) mod y)" - by (simp only: real_of_nat_add [THEN sym] real_of_nat_mult [THEN sym]) - have "x = real(natfloor x) + (x - real(natfloor x))" - by simp - then have "x = real ((natfloor x) div y) * real y + - real((natfloor x) mod y) + (x - real(natfloor x))" - by (simp add: a) - then have "x / real y = ... / real y" - by simp - also have "... = real((natfloor x) div y) + real((natfloor x) mod y) / - real y + (x - real(natfloor x)) / real y" - by (auto simp add: ring_simps add_divide_distrib - diff_divide_distrib prems) - finally have "natfloor (x / real y) = natfloor(...)" by simp - also have "... = natfloor(real((natfloor x) mod y) / - real y + (x - real(natfloor x)) / real y + real((natfloor x) div y))" - by (simp add: add_ac) - also have "... = natfloor(real((natfloor x) mod y) / - real y + (x - real(natfloor x)) / real y) + (natfloor x) div y" - apply (rule natfloor_add) - apply (rule add_nonneg_nonneg) - apply (rule divide_nonneg_pos) - apply simp - apply (simp add: prems) - apply (rule divide_nonneg_pos) - apply (simp add: compare_rls) - apply (rule real_natfloor_le) - apply (insert prems, auto) - done - also have "natfloor(real((natfloor x) mod y) / - real y + (x - real(natfloor x)) / real y) = 0" - apply (rule natfloor_eq) - apply simp - apply (rule add_nonneg_nonneg) - apply (rule divide_nonneg_pos) - apply force - apply (force simp add: prems) - apply (rule divide_nonneg_pos) - apply (simp add: compare_rls) - apply (rule real_natfloor_le) - apply (auto simp add: prems) - apply (insert prems, arith) - apply (simp add: add_divide_distrib [THEN sym]) - apply (subgoal_tac "real y = real y - 1 + 1") - apply (erule ssubst) - apply (rule add_le_less_mono) - apply (simp add: compare_rls) - apply (subgoal_tac "real(natfloor x mod y) + 1 = - real(natfloor x mod y + 1)") - apply (erule ssubst) - apply (subst real_of_nat_le_iff) - apply (subgoal_tac "natfloor x mod y < y") - apply arith - apply (rule mod_less_divisor) - apply auto - apply (simp add: compare_rls) - apply (subst add_commute) - apply (rule real_natfloor_add_one_gt) - done - finally show ?thesis by simp -qed - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/Rational.thy --- a/src/HOL/Real/Rational.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1020 +0,0 @@ -(* Title: HOL/Library/Rational.thy - ID: $Id$ - Author: Markus Wenzel, TU Muenchen -*) - -header {* Rational numbers *} - -theory Rational -imports "../Nat_Int_Bij" "~~/src/HOL/Library/GCD" -uses ("rat_arith.ML") -begin - -subsection {* Rational numbers as quotient *} - -subsubsection {* Construction of the type of rational numbers *} - -definition - ratrel :: "((int \ int) \ (int \ int)) set" where - "ratrel = {(x, y). snd x \ 0 \ snd y \ 0 \ fst x * snd y = fst y * snd x}" - -lemma ratrel_iff [simp]: - "(x, y) \ ratrel \ snd x \ 0 \ snd y \ 0 \ fst x * snd y = fst y * snd x" - by (simp add: ratrel_def) - -lemma refl_ratrel: "refl {x. snd x \ 0} ratrel" - by (auto simp add: refl_def ratrel_def) - -lemma sym_ratrel: "sym ratrel" - by (simp add: ratrel_def sym_def) - -lemma trans_ratrel: "trans ratrel" -proof (rule transI, unfold split_paired_all) - fix a b a' b' a'' b'' :: int - assume A: "((a, b), (a', b')) \ ratrel" - assume B: "((a', b'), (a'', b'')) \ ratrel" - have "b' * (a * b'') = b'' * (a * b')" by simp - also from A have "a * b' = a' * b" by auto - also have "b'' * (a' * b) = b * (a' * b'')" by simp - also from B have "a' * b'' = a'' * b'" by auto - also have "b * (a'' * b') = b' * (a'' * b)" by simp - finally have "b' * (a * b'') = b' * (a'' * b)" . - moreover from B have "b' \ 0" by auto - ultimately have "a * b'' = a'' * b" by simp - with A B show "((a, b), (a'', b'')) \ ratrel" by auto -qed - -lemma equiv_ratrel: "equiv {x. snd x \ 0} ratrel" - by (rule equiv.intro [OF refl_ratrel sym_ratrel trans_ratrel]) - -lemmas UN_ratrel = UN_equiv_class [OF equiv_ratrel] -lemmas UN_ratrel2 = UN_equiv_class2 [OF equiv_ratrel equiv_ratrel] - -lemma equiv_ratrel_iff [iff]: - assumes "snd x \ 0" and "snd y \ 0" - shows "ratrel `` {x} = ratrel `` {y} \ (x, y) \ ratrel" - by (rule eq_equiv_class_iff, rule equiv_ratrel) (auto simp add: assms) - -typedef (Rat) rat = "{x. snd x \ 0} // ratrel" -proof - have "(0::int, 1::int) \ {x. snd x \ 0}" by simp - then show "ratrel `` {(0, 1)} \ {x. snd x \ 0} // ratrel" by (rule quotientI) -qed - -lemma ratrel_in_Rat [simp]: "snd x \ 0 \ ratrel `` {x} \ Rat" - by (simp add: Rat_def quotientI) - -declare Abs_Rat_inject [simp] Abs_Rat_inverse [simp] - - -subsubsection {* Representation and basic operations *} - -definition - Fract :: "int \ int \ rat" where - [code del]: "Fract a b = Abs_Rat (ratrel `` {if b = 0 then (0, 1) else (a, b)})" - -code_datatype Fract - -lemma Rat_cases [case_names Fract, cases type: rat]: - assumes "\a b. q = Fract a b \ b \ 0 \ C" - shows C - using assms by (cases q) (clarsimp simp add: Fract_def Rat_def quotient_def) - -lemma Rat_induct [case_names Fract, induct type: rat]: - assumes "\a b. b \ 0 \ P (Fract a b)" - shows "P q" - using assms by (cases q) simp - -lemma eq_rat: - shows "\a b c d. b \ 0 \ d \ 0 \ Fract a b = Fract c d \ a * d = c * b" - and "\a. Fract a 0 = Fract 0 1" - and "\a c. Fract 0 a = Fract 0 c" - by (simp_all add: Fract_def) - -instantiation rat :: "{comm_ring_1, recpower}" -begin - -definition - Zero_rat_def [code, code unfold]: "0 = Fract 0 1" - -definition - One_rat_def [code, code unfold]: "1 = Fract 1 1" - -definition - add_rat_def [code del]: - "q + r = Abs_Rat (\x \ Rep_Rat q. \y \ Rep_Rat r. - ratrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)})" - -lemma add_rat [simp]: - assumes "b \ 0" and "d \ 0" - shows "Fract a b + Fract c d = Fract (a * d + c * b) (b * d)" -proof - - have "(\x y. ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)}) - respects2 ratrel" - by (rule equiv_ratrel [THEN congruent2_commuteI]) (simp_all add: left_distrib) - with assms show ?thesis by (simp add: Fract_def add_rat_def UN_ratrel2) -qed - -definition - minus_rat_def [code del]: - "- q = Abs_Rat (\x \ Rep_Rat q. ratrel `` {(- fst x, snd x)})" - -lemma minus_rat [simp, code]: "- Fract a b = Fract (- a) b" -proof - - have "(\x. ratrel `` {(- fst x, snd x)}) respects ratrel" - by (simp add: congruent_def) - then show ?thesis by (simp add: Fract_def minus_rat_def UN_ratrel) -qed - -lemma minus_rat_cancel [simp]: "Fract (- a) (- b) = Fract a b" - by (cases "b = 0") (simp_all add: eq_rat) - -definition - diff_rat_def [code del]: "q - r = q + - (r::rat)" - -lemma diff_rat [simp]: - assumes "b \ 0" and "d \ 0" - shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)" - using assms by (simp add: diff_rat_def) - -definition - mult_rat_def [code del]: - "q * r = Abs_Rat (\x \ Rep_Rat q. \y \ Rep_Rat r. - ratrel``{(fst x * fst y, snd x * snd y)})" - -lemma mult_rat [simp]: "Fract a b * Fract c d = Fract (a * c) (b * d)" -proof - - have "(\x y. ratrel `` {(fst x * fst y, snd x * snd y)}) respects2 ratrel" - by (rule equiv_ratrel [THEN congruent2_commuteI]) simp_all - then show ?thesis by (simp add: Fract_def mult_rat_def UN_ratrel2) -qed - -lemma mult_rat_cancel: - assumes "c \ 0" - shows "Fract (c * a) (c * b) = Fract a b" -proof - - from assms have "Fract c c = Fract 1 1" by (simp add: Fract_def) - then show ?thesis by (simp add: mult_rat [symmetric]) -qed - -primrec power_rat -where - rat_power_0: "q ^ 0 = (1\rat)" - | rat_power_Suc: "q ^ Suc n = (q\rat) * (q ^ n)" - -instance proof - fix q r s :: rat show "(q * r) * s = q * (r * s)" - by (cases q, cases r, cases s) (simp add: eq_rat) -next - fix q r :: rat show "q * r = r * q" - by (cases q, cases r) (simp add: eq_rat) -next - fix q :: rat show "1 * q = q" - by (cases q) (simp add: One_rat_def eq_rat) -next - fix q r s :: rat show "(q + r) + s = q + (r + s)" - by (cases q, cases r, cases s) (simp add: eq_rat ring_simps) -next - fix q r :: rat show "q + r = r + q" - by (cases q, cases r) (simp add: eq_rat) -next - fix q :: rat show "0 + q = q" - by (cases q) (simp add: Zero_rat_def eq_rat) -next - fix q :: rat show "- q + q = 0" - by (cases q) (simp add: Zero_rat_def eq_rat) -next - fix q r :: rat show "q - r = q + - r" - by (cases q, cases r) (simp add: eq_rat) -next - fix q r s :: rat show "(q + r) * s = q * s + r * s" - by (cases q, cases r, cases s) (simp add: eq_rat ring_simps) -next - show "(0::rat) \ 1" by (simp add: Zero_rat_def One_rat_def eq_rat) -next - fix q :: rat show "q * 1 = q" - by (cases q) (simp add: One_rat_def eq_rat) -next - fix q :: rat - fix n :: nat - show "q ^ 0 = 1" by simp - show "q ^ (Suc n) = q * (q ^ n)" by simp -qed - -end - -lemma of_nat_rat: "of_nat k = Fract (of_nat k) 1" - by (induct k) (simp_all add: Zero_rat_def One_rat_def) - -lemma of_int_rat: "of_int k = Fract k 1" - by (cases k rule: int_diff_cases) (simp add: of_nat_rat) - -lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k" - by (rule of_nat_rat [symmetric]) - -lemma Fract_of_int_eq: "Fract k 1 = of_int k" - by (rule of_int_rat [symmetric]) - -instantiation rat :: number_ring -begin - -definition - rat_number_of_def [code del]: "number_of w = Fract w 1" - -instance by intro_classes (simp add: rat_number_of_def of_int_rat) - -end - -lemma rat_number_collapse [code post]: - "Fract 0 k = 0" - "Fract 1 1 = 1" - "Fract (number_of k) 1 = number_of k" - "Fract k 0 = 0" - by (cases "k = 0") - (simp_all add: Zero_rat_def One_rat_def number_of_is_id number_of_eq of_int_rat eq_rat Fract_def) - -lemma rat_number_expand [code unfold]: - "0 = Fract 0 1" - "1 = Fract 1 1" - "number_of k = Fract (number_of k) 1" - by (simp_all add: rat_number_collapse) - -lemma iszero_rat [simp]: - "iszero (number_of k :: rat) \ iszero (number_of k :: int)" - by (simp add: iszero_def rat_number_expand number_of_is_id eq_rat) - -lemma Rat_cases_nonzero [case_names Fract 0]: - assumes Fract: "\a b. q = Fract a b \ b \ 0 \ a \ 0 \ C" - assumes 0: "q = 0 \ C" - shows C -proof (cases "q = 0") - case True then show C using 0 by auto -next - case False - then obtain a b where "q = Fract a b" and "b \ 0" by (cases q) auto - moreover with False have "0 \ Fract a b" by simp - with `b \ 0` have "a \ 0" by (simp add: Zero_rat_def eq_rat) - with Fract `q = Fract a b` `b \ 0` show C by auto -qed - - - -subsubsection {* The field of rational numbers *} - -instantiation rat :: "{field, division_by_zero}" -begin - -definition - inverse_rat_def [code del]: - "inverse q = Abs_Rat (\x \ Rep_Rat q. - ratrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)})" - -lemma inverse_rat [simp]: "inverse (Fract a b) = Fract b a" -proof - - have "(\x. ratrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)}) respects ratrel" - by (auto simp add: congruent_def mult_commute) - then show ?thesis by (simp add: Fract_def inverse_rat_def UN_ratrel) -qed - -definition - divide_rat_def [code del]: "q / r = q * inverse (r::rat)" - -lemma divide_rat [simp]: "Fract a b / Fract c d = Fract (a * d) (b * c)" - by (simp add: divide_rat_def) - -instance proof - show "inverse 0 = (0::rat)" by (simp add: rat_number_expand) - (simp add: rat_number_collapse) -next - fix q :: rat - assume "q \ 0" - then show "inverse q * q = 1" by (cases q rule: Rat_cases_nonzero) - (simp_all add: mult_rat inverse_rat rat_number_expand eq_rat) -next - fix q r :: rat - show "q / r = q * inverse r" by (simp add: divide_rat_def) -qed - -end - - -subsubsection {* Various *} - -lemma Fract_add_one: "n \ 0 ==> Fract (m + n) n = Fract m n + 1" - by (simp add: rat_number_expand) - -lemma Fract_of_int_quotient: "Fract k l = of_int k / of_int l" - by (simp add: Fract_of_int_eq [symmetric]) - -lemma Fract_number_of_quotient [code post]: - "Fract (number_of k) (number_of l) = number_of k / number_of l" - unfolding Fract_of_int_quotient number_of_is_id number_of_eq .. - -lemma Fract_1_number_of [code post]: - "Fract 1 (number_of k) = 1 / number_of k" - unfolding Fract_of_int_quotient number_of_eq by simp - -subsubsection {* The ordered field of rational numbers *} - -instantiation rat :: linorder -begin - -definition - le_rat_def [code del]: - "q \ r \ contents (\x \ Rep_Rat q. \y \ Rep_Rat r. - {(fst x * snd y) * (snd x * snd y) \ (fst y * snd x) * (snd x * snd y)})" - -lemma le_rat [simp]: - assumes "b \ 0" and "d \ 0" - shows "Fract a b \ Fract c d \ (a * d) * (b * d) \ (c * b) * (b * d)" -proof - - have "(\x y. {(fst x * snd y) * (snd x * snd y) \ (fst y * snd x) * (snd x * snd y)}) - respects2 ratrel" - proof (clarsimp simp add: congruent2_def) - fix a b a' b' c d c' d'::int - assume neq: "b \ 0" "b' \ 0" "d \ 0" "d' \ 0" - assume eq1: "a * b' = a' * b" - assume eq2: "c * d' = c' * d" - - let ?le = "\a b c d. ((a * d) * (b * d) \ (c * b) * (b * d))" - { - fix a b c d x :: int assume x: "x \ 0" - have "?le a b c d = ?le (a * x) (b * x) c d" - proof - - from x have "0 < x * x" by (auto simp add: zero_less_mult_iff) - hence "?le a b c d = - ((a * d) * (b * d) * (x * x) \ (c * b) * (b * d) * (x * x))" - by (simp add: mult_le_cancel_right) - also have "... = ?le (a * x) (b * x) c d" - by (simp add: mult_ac) - finally show ?thesis . - qed - } note le_factor = this - - let ?D = "b * d" and ?D' = "b' * d'" - from neq have D: "?D \ 0" by simp - from neq have "?D' \ 0" by simp - hence "?le a b c d = ?le (a * ?D') (b * ?D') c d" - by (rule le_factor) - also have "... = ((a * b') * ?D * ?D' * d * d' \ (c * d') * ?D * ?D' * b * b')" - by (simp add: mult_ac) - also have "... = ((a' * b) * ?D * ?D' * d * d' \ (c' * d) * ?D * ?D' * b * b')" - by (simp only: eq1 eq2) - also have "... = ?le (a' * ?D) (b' * ?D) c' d'" - by (simp add: mult_ac) - also from D have "... = ?le a' b' c' d'" - by (rule le_factor [symmetric]) - finally show "?le a b c d = ?le a' b' c' d'" . - qed - with assms show ?thesis by (simp add: Fract_def le_rat_def UN_ratrel2) -qed - -definition - less_rat_def [code del]: "z < (w::rat) \ z \ w \ z \ w" - -lemma less_rat [simp]: - assumes "b \ 0" and "d \ 0" - shows "Fract a b < Fract c d \ (a * d) * (b * d) < (c * b) * (b * d)" - using assms by (simp add: less_rat_def eq_rat order_less_le) - -instance proof - fix q r s :: rat - { - assume "q \ r" and "r \ s" - show "q \ s" - proof (insert prems, induct q, induct r, induct s) - fix a b c d e f :: int - assume neq: "b \ 0" "d \ 0" "f \ 0" - assume 1: "Fract a b \ Fract c d" and 2: "Fract c d \ Fract e f" - show "Fract a b \ Fract e f" - proof - - from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f" - by (auto simp add: zero_less_mult_iff linorder_neq_iff) - have "(a * d) * (b * d) * (f * f) \ (c * b) * (b * d) * (f * f)" - proof - - from neq 1 have "(a * d) * (b * d) \ (c * b) * (b * d)" - by simp - with ff show ?thesis by (simp add: mult_le_cancel_right) - qed - also have "... = (c * f) * (d * f) * (b * b)" by algebra - also have "... \ (e * d) * (d * f) * (b * b)" - proof - - from neq 2 have "(c * f) * (d * f) \ (e * d) * (d * f)" - by simp - with bb show ?thesis by (simp add: mult_le_cancel_right) - qed - finally have "(a * f) * (b * f) * (d * d) \ e * b * (b * f) * (d * d)" - by (simp only: mult_ac) - with dd have "(a * f) * (b * f) \ (e * b) * (b * f)" - by (simp add: mult_le_cancel_right) - with neq show ?thesis by simp - qed - qed - next - assume "q \ r" and "r \ q" - show "q = r" - proof (insert prems, induct q, induct r) - fix a b c d :: int - assume neq: "b \ 0" "d \ 0" - assume 1: "Fract a b \ Fract c d" and 2: "Fract c d \ Fract a b" - show "Fract a b = Fract c d" - proof - - from neq 1 have "(a * d) * (b * d) \ (c * b) * (b * d)" - by simp - also have "... \ (a * d) * (b * d)" - proof - - from neq 2 have "(c * b) * (d * b) \ (a * d) * (d * b)" - by simp - thus ?thesis by (simp only: mult_ac) - qed - finally have "(a * d) * (b * d) = (c * b) * (b * d)" . - moreover from neq have "b * d \ 0" by simp - ultimately have "a * d = c * b" by simp - with neq show ?thesis by (simp add: eq_rat) - qed - qed - next - show "q \ q" - by (induct q) simp - show "(q < r) = (q \ r \ \ r \ q)" - by (induct q, induct r) (auto simp add: le_less mult_commute) - show "q \ r \ r \ q" - by (induct q, induct r) - (simp add: mult_commute, rule linorder_linear) - } -qed - -end - -instantiation rat :: "{distrib_lattice, abs_if, sgn_if}" -begin - -definition - abs_rat_def [code del]: "\q\ = (if q < 0 then -q else (q::rat))" - -lemma abs_rat [simp, code]: "\Fract a b\ = Fract \a\ \b\" - by (auto simp add: abs_rat_def zabs_def Zero_rat_def less_rat not_less le_less minus_rat eq_rat zero_compare_simps) - -definition - sgn_rat_def [code del]: "sgn (q::rat) = (if q = 0 then 0 else if 0 < q then 1 else - 1)" - -lemma sgn_rat [simp, code]: "sgn (Fract a b) = of_int (sgn a * sgn b)" - unfolding Fract_of_int_eq - by (auto simp: zsgn_def sgn_rat_def Zero_rat_def eq_rat) - (auto simp: rat_number_collapse not_less le_less zero_less_mult_iff) - -definition - "(inf \ rat \ rat \ rat) = min" - -definition - "(sup \ rat \ rat \ rat) = max" - -instance by intro_classes - (auto simp add: abs_rat_def sgn_rat_def min_max.sup_inf_distrib1 inf_rat_def sup_rat_def) - -end - -instance rat :: ordered_field -proof - fix q r s :: rat - show "q \ r ==> s + q \ s + r" - proof (induct q, induct r, induct s) - fix a b c d e f :: int - assume neq: "b \ 0" "d \ 0" "f \ 0" - assume le: "Fract a b \ Fract c d" - show "Fract e f + Fract a b \ Fract e f + Fract c d" - proof - - let ?F = "f * f" from neq have F: "0 < ?F" - by (auto simp add: zero_less_mult_iff) - from neq le have "(a * d) * (b * d) \ (c * b) * (b * d)" - by simp - with F have "(a * d) * (b * d) * ?F * ?F \ (c * b) * (b * d) * ?F * ?F" - by (simp add: mult_le_cancel_right) - with neq show ?thesis by (simp add: mult_ac int_distrib) - qed - qed - show "q < r ==> 0 < s ==> s * q < s * r" - proof (induct q, induct r, induct s) - fix a b c d e f :: int - assume neq: "b \ 0" "d \ 0" "f \ 0" - assume le: "Fract a b < Fract c d" - assume gt: "0 < Fract e f" - show "Fract e f * Fract a b < Fract e f * Fract c d" - proof - - let ?E = "e * f" and ?F = "f * f" - from neq gt have "0 < ?E" - by (auto simp add: Zero_rat_def order_less_le eq_rat) - moreover from neq have "0 < ?F" - by (auto simp add: zero_less_mult_iff) - moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)" - by simp - ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F" - by (simp add: mult_less_cancel_right) - with neq show ?thesis - by (simp add: mult_ac) - qed - qed -qed auto - -lemma Rat_induct_pos [case_names Fract, induct type: rat]: - assumes step: "\a b. 0 < b \ P (Fract a b)" - shows "P q" -proof (cases q) - have step': "\a b. b < 0 \ P (Fract a b)" - proof - - fix a::int and b::int - assume b: "b < 0" - hence "0 < -b" by simp - hence "P (Fract (-a) (-b))" by (rule step) - thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b]) - qed - case (Fract a b) - thus "P q" by (force simp add: linorder_neq_iff step step') -qed - -lemma zero_less_Fract_iff: - "0 < b ==> (0 < Fract a b) = (0 < a)" -by (simp add: Zero_rat_def order_less_imp_not_eq2 zero_less_mult_iff) - - -subsection {* Arithmetic setup *} - -use "rat_arith.ML" -declaration {* K rat_arith_setup *} - - -subsection {* Embedding from Rationals to other Fields *} - -class field_char_0 = field + ring_char_0 - -subclass (in ordered_field) field_char_0 .. - -context field_char_0 -begin - -definition of_rat :: "rat \ 'a" where - [code del]: "of_rat q = contents (\(a,b) \ Rep_Rat q. {of_int a / of_int b})" - -end - -lemma of_rat_congruent: - "(\(a, b). {of_int a / of_int b :: 'a::field_char_0}) respects ratrel" -apply (rule congruent.intro) -apply (clarsimp simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq) -apply (simp only: of_int_mult [symmetric]) -done - -lemma of_rat_rat: "b \ 0 \ of_rat (Fract a b) = of_int a / of_int b" - unfolding Fract_def of_rat_def by (simp add: UN_ratrel of_rat_congruent) - -lemma of_rat_0 [simp]: "of_rat 0 = 0" -by (simp add: Zero_rat_def of_rat_rat) - -lemma of_rat_1 [simp]: "of_rat 1 = 1" -by (simp add: One_rat_def of_rat_rat) - -lemma of_rat_add: "of_rat (a + b) = of_rat a + of_rat b" -by (induct a, induct b, simp add: of_rat_rat add_frac_eq) - -lemma of_rat_minus: "of_rat (- a) = - of_rat a" -by (induct a, simp add: of_rat_rat) - -lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b" -by (simp only: diff_minus of_rat_add of_rat_minus) - -lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b" -apply (induct a, induct b, simp add: of_rat_rat) -apply (simp add: divide_inverse nonzero_inverse_mult_distrib mult_ac) -done - -lemma nonzero_of_rat_inverse: - "a \ 0 \ of_rat (inverse a) = inverse (of_rat a)" -apply (rule inverse_unique [symmetric]) -apply (simp add: of_rat_mult [symmetric]) -done - -lemma of_rat_inverse: - "(of_rat (inverse a)::'a::{field_char_0,division_by_zero}) = - inverse (of_rat a)" -by (cases "a = 0", simp_all add: nonzero_of_rat_inverse) - -lemma nonzero_of_rat_divide: - "b \ 0 \ of_rat (a / b) = of_rat a / of_rat b" -by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse) - -lemma of_rat_divide: - "(of_rat (a / b)::'a::{field_char_0,division_by_zero}) - = of_rat a / of_rat b" -by (cases "b = 0") (simp_all add: nonzero_of_rat_divide) - -lemma of_rat_power: - "(of_rat (a ^ n)::'a::{field_char_0,recpower}) = of_rat a ^ n" -by (induct n) (simp_all add: of_rat_mult power_Suc) - -lemma of_rat_eq_iff [simp]: "(of_rat a = of_rat b) = (a = b)" -apply (induct a, induct b) -apply (simp add: of_rat_rat eq_rat) -apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq) -apply (simp only: of_int_mult [symmetric] of_int_eq_iff) -done - -lemma of_rat_less: - "(of_rat r :: 'a::ordered_field) < of_rat s \ r < s" -proof (induct r, induct s) - fix a b c d :: int - assume not_zero: "b > 0" "d > 0" - then have "b * d > 0" by (rule mult_pos_pos) - have of_int_divide_less_eq: - "(of_int a :: 'a) / of_int b < of_int c / of_int d - \ (of_int a :: 'a) * of_int d < of_int c * of_int b" - using not_zero by (simp add: pos_less_divide_eq pos_divide_less_eq) - show "(of_rat (Fract a b) :: 'a::ordered_field) < of_rat (Fract c d) - \ Fract a b < Fract c d" - using not_zero `b * d > 0` - by (simp add: of_rat_rat of_int_divide_less_eq of_int_mult [symmetric] del: of_int_mult) - (auto intro: mult_strict_right_mono mult_right_less_imp_less) -qed - -lemma of_rat_less_eq: - "(of_rat r :: 'a::ordered_field) \ of_rat s \ r \ s" - unfolding le_less by (auto simp add: of_rat_less) - -lemmas of_rat_eq_0_iff [simp] = of_rat_eq_iff [of _ 0, simplified] - -lemma of_rat_eq_id [simp]: "of_rat = id" -proof - fix a - show "of_rat a = id a" - by (induct a) - (simp add: of_rat_rat Fract_of_int_eq [symmetric]) -qed - -text{*Collapse nested embeddings*} -lemma of_rat_of_nat_eq [simp]: "of_rat (of_nat n) = of_nat n" -by (induct n) (simp_all add: of_rat_add) - -lemma of_rat_of_int_eq [simp]: "of_rat (of_int z) = of_int z" -by (cases z rule: int_diff_cases) (simp add: of_rat_diff) - -lemma of_rat_number_of_eq [simp]: - "of_rat (number_of w) = (number_of w :: 'a::{number_ring,field_char_0})" -by (simp add: number_of_eq) - -lemmas zero_rat = Zero_rat_def -lemmas one_rat = One_rat_def - -abbreviation - rat_of_nat :: "nat \ rat" -where - "rat_of_nat \ of_nat" - -abbreviation - rat_of_int :: "int \ rat" -where - "rat_of_int \ of_int" - -subsection {* The Set of Rational Numbers *} - -context field_char_0 -begin - -definition - Rats :: "'a set" where - [code del]: "Rats = range of_rat" - -notation (xsymbols) - Rats ("\") - -end - -lemma Rats_of_rat [simp]: "of_rat r \ Rats" -by (simp add: Rats_def) - -lemma Rats_of_int [simp]: "of_int z \ Rats" -by (subst of_rat_of_int_eq [symmetric], rule Rats_of_rat) - -lemma Rats_of_nat [simp]: "of_nat n \ Rats" -by (subst of_rat_of_nat_eq [symmetric], rule Rats_of_rat) - -lemma Rats_number_of [simp]: - "(number_of w::'a::{number_ring,field_char_0}) \ Rats" -by (subst of_rat_number_of_eq [symmetric], rule Rats_of_rat) - -lemma Rats_0 [simp]: "0 \ Rats" -apply (unfold Rats_def) -apply (rule range_eqI) -apply (rule of_rat_0 [symmetric]) -done - -lemma Rats_1 [simp]: "1 \ Rats" -apply (unfold Rats_def) -apply (rule range_eqI) -apply (rule of_rat_1 [symmetric]) -done - -lemma Rats_add [simp]: "\a \ Rats; b \ Rats\ \ a + b \ Rats" -apply (auto simp add: Rats_def) -apply (rule range_eqI) -apply (rule of_rat_add [symmetric]) -done - -lemma Rats_minus [simp]: "a \ Rats \ - a \ Rats" -apply (auto simp add: Rats_def) -apply (rule range_eqI) -apply (rule of_rat_minus [symmetric]) -done - -lemma Rats_diff [simp]: "\a \ Rats; b \ Rats\ \ a - b \ Rats" -apply (auto simp add: Rats_def) -apply (rule range_eqI) -apply (rule of_rat_diff [symmetric]) -done - -lemma Rats_mult [simp]: "\a \ Rats; b \ Rats\ \ a * b \ Rats" -apply (auto simp add: Rats_def) -apply (rule range_eqI) -apply (rule of_rat_mult [symmetric]) -done - -lemma nonzero_Rats_inverse: - fixes a :: "'a::field_char_0" - shows "\a \ Rats; a \ 0\ \ inverse a \ Rats" -apply (auto simp add: Rats_def) -apply (rule range_eqI) -apply (erule nonzero_of_rat_inverse [symmetric]) -done - -lemma Rats_inverse [simp]: - fixes a :: "'a::{field_char_0,division_by_zero}" - shows "a \ Rats \ inverse a \ Rats" -apply (auto simp add: Rats_def) -apply (rule range_eqI) -apply (rule of_rat_inverse [symmetric]) -done - -lemma nonzero_Rats_divide: - fixes a b :: "'a::field_char_0" - shows "\a \ Rats; b \ Rats; b \ 0\ \ a / b \ Rats" -apply (auto simp add: Rats_def) -apply (rule range_eqI) -apply (erule nonzero_of_rat_divide [symmetric]) -done - -lemma Rats_divide [simp]: - fixes a b :: "'a::{field_char_0,division_by_zero}" - shows "\a \ Rats; b \ Rats\ \ a / b \ Rats" -apply (auto simp add: Rats_def) -apply (rule range_eqI) -apply (rule of_rat_divide [symmetric]) -done - -lemma Rats_power [simp]: - fixes a :: "'a::{field_char_0,recpower}" - shows "a \ Rats \ a ^ n \ Rats" -apply (auto simp add: Rats_def) -apply (rule range_eqI) -apply (rule of_rat_power [symmetric]) -done - -lemma Rats_cases [cases set: Rats]: - assumes "q \ \" - obtains (of_rat) r where "q = of_rat r" - unfolding Rats_def -proof - - from `q \ \` have "q \ range of_rat" unfolding Rats_def . - then obtain r where "q = of_rat r" .. - then show thesis .. -qed - -lemma Rats_induct [case_names of_rat, induct set: Rats]: - "q \ \ \ (\r. P (of_rat r)) \ P q" - by (rule Rats_cases) auto - - -subsection {* The Rationals are Countably Infinite *} - -definition nat_to_rat_surj :: "nat \ rat" where -"nat_to_rat_surj n = (let (a,b) = nat_to_nat2 n - in Fract (nat_to_int_bij a) (nat_to_int_bij b))" - -lemma surj_nat_to_rat_surj: "surj nat_to_rat_surj" -unfolding surj_def -proof - fix r::rat - show "\n. r = nat_to_rat_surj n" - proof(cases r) - fix i j assume [simp]: "r = Fract i j" and "j \ 0" - have "r = (let m = inv nat_to_int_bij i; n = inv nat_to_int_bij j - in nat_to_rat_surj(nat2_to_nat (m,n)))" - using nat2_to_nat_inj surj_f_inv_f[OF surj_nat_to_int_bij] - by(simp add:Let_def nat_to_rat_surj_def nat_to_nat2_def) - thus "\n. r = nat_to_rat_surj n" by(auto simp:Let_def) - qed -qed - -lemma Rats_eq_range_nat_to_rat_surj: "\ = range nat_to_rat_surj" -by (simp add: Rats_def surj_nat_to_rat_surj surj_range) - -context field_char_0 -begin - -lemma Rats_eq_range_of_rat_o_nat_to_rat_surj: - "\ = range (of_rat o nat_to_rat_surj)" -using surj_nat_to_rat_surj -by (auto simp: Rats_def image_def surj_def) - (blast intro: arg_cong[where f = of_rat]) - -lemma surj_of_rat_nat_to_rat_surj: - "r\\ \ \n. r = of_rat(nat_to_rat_surj n)" -by(simp add: Rats_eq_range_of_rat_o_nat_to_rat_surj image_def) - -end - - -subsection {* Implementation of rational numbers as pairs of integers *} - -lemma Fract_norm: "Fract (a div zgcd a b) (b div zgcd a b) = Fract a b" -proof (cases "a = 0 \ b = 0") - case True then show ?thesis by (auto simp add: eq_rat) -next - let ?c = "zgcd a b" - case False then have "a \ 0" and "b \ 0" by auto - then have "?c \ 0" by simp - then have "Fract ?c ?c = Fract 1 1" by (simp add: eq_rat) - moreover have "Fract (a div ?c * ?c + a mod ?c) (b div ?c * ?c + b mod ?c) = Fract a b" - by (simp add: semiring_div_class.mod_div_equality) - moreover have "a mod ?c = 0" by (simp add: dvd_eq_mod_eq_0 [symmetric]) - moreover have "b mod ?c = 0" by (simp add: dvd_eq_mod_eq_0 [symmetric]) - ultimately show ?thesis - by (simp add: mult_rat [symmetric]) -qed - -definition Fract_norm :: "int \ int \ rat" where - [simp, code del]: "Fract_norm a b = Fract a b" - -lemma [code]: "Fract_norm a b = (if a = 0 \ b = 0 then 0 else let c = zgcd a b in - if b > 0 then Fract (a div c) (b div c) else Fract (- (a div c)) (- (b div c)))" - by (simp add: eq_rat Zero_rat_def Let_def Fract_norm) - -lemma [code]: - "of_rat (Fract a b) = (if b \ 0 then of_int a / of_int b else 0)" - by (cases "b = 0") (simp_all add: rat_number_collapse of_rat_rat) - -instantiation rat :: eq -begin - -definition [code del]: "eq_class.eq (a\rat) b \ a - b = 0" - -instance by default (simp add: eq_rat_def) - -lemma rat_eq_code [code]: - "eq_class.eq (Fract a b) (Fract c d) \ (if b = 0 - then c = 0 \ d = 0 - else if d = 0 - then a = 0 \ b = 0 - else a * d = b * c)" - by (auto simp add: eq eq_rat) - -lemma rat_eq_refl [code nbe]: - "eq_class.eq (r::rat) r \ True" - by (rule HOL.eq_refl) - -end - -lemma le_rat': - assumes "b \ 0" - and "d \ 0" - shows "Fract a b \ Fract c d \ a * \d\ * sgn b \ c * \b\ * sgn d" -proof - - have abs_sgn: "\k::int. \k\ = k * sgn k" unfolding abs_if sgn_if by simp - have "a * d * (b * d) \ c * b * (b * d) \ a * d * (sgn b * sgn d) \ c * b * (sgn b * sgn d)" - proof (cases "b * d > 0") - case True - moreover from True have "sgn b * sgn d = 1" - by (simp add: sgn_times [symmetric] sgn_1_pos) - ultimately show ?thesis by (simp add: mult_le_cancel_right) - next - case False with assms have "b * d < 0" by (simp add: less_le) - moreover from this have "sgn b * sgn d = - 1" - by (simp only: sgn_times [symmetric] sgn_1_neg) - ultimately show ?thesis by (simp add: mult_le_cancel_right) - qed - also have "\ \ a * \d\ * sgn b \ c * \b\ * sgn d" - by (simp add: abs_sgn mult_ac) - finally show ?thesis using assms by simp -qed - -lemma less_rat': - assumes "b \ 0" - and "d \ 0" - shows "Fract a b < Fract c d \ a * \d\ * sgn b < c * \b\ * sgn d" -proof - - have abs_sgn: "\k::int. \k\ = k * sgn k" unfolding abs_if sgn_if by simp - have "a * d * (b * d) < c * b * (b * d) \ a * d * (sgn b * sgn d) < c * b * (sgn b * sgn d)" - proof (cases "b * d > 0") - case True - moreover from True have "sgn b * sgn d = 1" - by (simp add: sgn_times [symmetric] sgn_1_pos) - ultimately show ?thesis by (simp add: mult_less_cancel_right) - next - case False with assms have "b * d < 0" by (simp add: less_le) - moreover from this have "sgn b * sgn d = - 1" - by (simp only: sgn_times [symmetric] sgn_1_neg) - ultimately show ?thesis by (simp add: mult_less_cancel_right) - qed - also have "\ \ a * \d\ * sgn b < c * \b\ * sgn d" - by (simp add: abs_sgn mult_ac) - finally show ?thesis using assms by simp -qed - -lemma rat_less_eq_code [code]: - "Fract a b \ Fract c d \ (if b = 0 - then sgn c * sgn d \ 0 - else if d = 0 - then sgn a * sgn b \ 0 - else a * \d\ * sgn b \ c * \b\ * sgn d)" -by (auto simp add: sgn_times mult_le_0_iff zero_le_mult_iff le_rat' eq_rat simp del: le_rat) - (auto simp add: sgn_times sgn_0_0 le_less sgn_1_pos [symmetric] sgn_1_neg [symmetric]) - -lemma rat_le_eq_code [code]: - "Fract a b < Fract c d \ (if b = 0 - then sgn c * sgn d > 0 - else if d = 0 - then sgn a * sgn b < 0 - else a * \d\ * sgn b < c * \b\ * sgn d)" -by (auto simp add: sgn_times mult_less_0_iff zero_less_mult_iff less_rat' eq_rat simp del: less_rat) - (auto simp add: sgn_times sgn_0_0 sgn_1_pos [symmetric] sgn_1_neg [symmetric], - auto simp add: sgn_1_pos) - -lemma rat_plus_code [code]: - "Fract a b + Fract c d = (if b = 0 - then Fract c d - else if d = 0 - then Fract a b - else Fract_norm (a * d + c * b) (b * d))" - by (simp add: eq_rat, simp add: Zero_rat_def) - -lemma rat_times_code [code]: - "Fract a b * Fract c d = Fract_norm (a * c) (b * d)" - by simp - -lemma rat_minus_code [code]: - "Fract a b - Fract c d = (if b = 0 - then Fract (- c) d - else if d = 0 - then Fract a b - else Fract_norm (a * d - c * b) (b * d))" - by (simp add: eq_rat, simp add: Zero_rat_def) - -lemma rat_inverse_code [code]: - "inverse (Fract a b) = (if b = 0 then Fract 1 0 - else if a < 0 then Fract (- b) (- a) - else Fract b a)" - by (simp add: eq_rat) - -lemma rat_divide_code [code]: - "Fract a b / Fract c d = Fract_norm (a * d) (b * c)" - by simp - -hide (open) const Fract_norm - -text {* Setup for SML code generator *} - -types_code - rat ("(int */ int)") -attach (term_of) {* -fun term_of_rat (p, q) = - let - val rT = Type ("Rational.rat", []) - in - if q = 1 orelse p = 0 then HOLogic.mk_number rT p - else @{term "op / \ rat \ rat \ rat"} $ - HOLogic.mk_number rT p $ HOLogic.mk_number rT q - end; -*} -attach (test) {* -fun gen_rat i = - let - val p = random_range 0 i; - val q = random_range 1 (i + 1); - val g = Integer.gcd p q; - val p' = p div g; - val q' = q div g; - val r = (if one_of [true, false] then p' else ~ p', - if p' = 0 then 0 else q') - in - (r, fn () => term_of_rat r) - end; -*} - -consts_code - Fract ("(_,/ _)") - -consts_code - "of_int :: int \ rat" ("\rat'_of'_int") -attach {* -fun rat_of_int 0 = (0, 0) - | rat_of_int i = (i, 1); -*} - -end \ No newline at end of file diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/Real.thy --- a/src/HOL/Real/Real.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6 +0,0 @@ -(* $Id$ *) - -theory Real -imports ContNotDenum RealVector -begin -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/RealDef.thy --- a/src/HOL/Real/RealDef.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1159 +0,0 @@ -(* Title : Real/RealDef.thy - ID : $Id$ - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 - Additional contributions by Jeremy Avigad -*) - -header{*Defining the Reals from the Positive Reals*} - -theory RealDef -imports PReal -uses ("real_arith.ML") -begin - -definition - realrel :: "((preal * preal) * (preal * preal)) set" where - [code del]: "realrel = {p. \x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}" - -typedef (Real) real = "UNIV//realrel" - by (auto simp add: quotient_def) - -definition - (** these don't use the overloaded "real" function: users don't see them **) - real_of_preal :: "preal => real" where - [code del]: "real_of_preal m = Abs_Real (realrel `` {(m + 1, 1)})" - -instantiation real :: "{zero, one, plus, minus, uminus, times, inverse, ord, abs, sgn}" -begin - -definition - real_zero_def [code del]: "0 = Abs_Real(realrel``{(1, 1)})" - -definition - real_one_def [code del]: "1 = Abs_Real(realrel``{(1 + 1, 1)})" - -definition - real_add_def [code del]: "z + w = - contents (\(x,y) \ Rep_Real(z). \(u,v) \ Rep_Real(w). - { Abs_Real(realrel``{(x+u, y+v)}) })" - -definition - real_minus_def [code del]: "- r = contents (\(x,y) \ Rep_Real(r). { Abs_Real(realrel``{(y,x)}) })" - -definition - real_diff_def [code del]: "r - (s::real) = r + - s" - -definition - real_mult_def [code del]: - "z * w = - contents (\(x,y) \ Rep_Real(z). \(u,v) \ Rep_Real(w). - { Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })" - -definition - real_inverse_def [code del]: "inverse (R::real) = (THE S. (R = 0 & S = 0) | S * R = 1)" - -definition - real_divide_def [code del]: "R / (S::real) = R * inverse S" - -definition - real_le_def [code del]: "z \ (w::real) \ - (\x y u v. x+v \ u+y & (x,y) \ Rep_Real z & (u,v) \ Rep_Real w)" - -definition - real_less_def [code del]: "x < (y\real) \ x \ y \ x \ y" - -definition - real_abs_def: "abs (r::real) = (if r < 0 then - r else r)" - -definition - real_sgn_def: "sgn (x::real) = (if x=0 then 0 else if 0 realrel) = (x1 + y2 = x2 + y1)" -by (simp add: realrel_def) - -lemma equiv_realrel: "equiv UNIV realrel" -apply (auto simp add: equiv_def refl_def sym_def trans_def realrel_def) -apply (blast dest: preal_trans_lemma) -done - -text{*Reduces equality of equivalence classes to the @{term realrel} relation: - @{term "(realrel `` {x} = realrel `` {y}) = ((x,y) \ realrel)"} *} -lemmas equiv_realrel_iff = - eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I] - -declare equiv_realrel_iff [simp] - - -lemma realrel_in_real [simp]: "realrel``{(x,y)}: Real" -by (simp add: Real_def realrel_def quotient_def, blast) - -declare Abs_Real_inject [simp] -declare Abs_Real_inverse [simp] - - -text{*Case analysis on the representation of a real number as an equivalence - class of pairs of positive reals.*} -lemma eq_Abs_Real [case_names Abs_Real, cases type: real]: - "(!!x y. z = Abs_Real(realrel``{(x,y)}) ==> P) ==> P" -apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE]) -apply (drule arg_cong [where f=Abs_Real]) -apply (auto simp add: Rep_Real_inverse) -done - - -subsection {* Addition and Subtraction *} - -lemma real_add_congruent2_lemma: - "[|a + ba = aa + b; ab + bc = ac + bb|] - ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))" -apply (simp add: add_assoc) -apply (rule add_left_commute [of ab, THEN ssubst]) -apply (simp add: add_assoc [symmetric]) -apply (simp add: add_ac) -done - -lemma real_add: - "Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) = - Abs_Real (realrel``{(x+u, y+v)})" -proof - - have "(\z w. (\(x,y). (\(u,v). {Abs_Real (realrel `` {(x+u, y+v)})}) w) z) - respects2 realrel" - by (simp add: congruent2_def, blast intro: real_add_congruent2_lemma) - thus ?thesis - by (simp add: real_add_def UN_UN_split_split_eq - UN_equiv_class2 [OF equiv_realrel equiv_realrel]) -qed - -lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})" -proof - - have "(\(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel" - by (simp add: congruent_def add_commute) - thus ?thesis - by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel]) -qed - -instance real :: ab_group_add -proof - fix x y z :: real - show "(x + y) + z = x + (y + z)" - by (cases x, cases y, cases z, simp add: real_add add_assoc) - show "x + y = y + x" - by (cases x, cases y, simp add: real_add add_commute) - show "0 + x = x" - by (cases x, simp add: real_add real_zero_def add_ac) - show "- x + x = 0" - by (cases x, simp add: real_minus real_add real_zero_def add_commute) - show "x - y = x + - y" - by (simp add: real_diff_def) -qed - - -subsection {* Multiplication *} - -lemma real_mult_congruent2_lemma: - "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==> - x * x1 + y * y1 + (x * y2 + y * x2) = - x * x2 + y * y2 + (x * y1 + y * x1)" -apply (simp add: add_left_commute add_assoc [symmetric]) -apply (simp add: add_assoc right_distrib [symmetric]) -apply (simp add: add_commute) -done - -lemma real_mult_congruent2: - "(%p1 p2. - (%(x1,y1). (%(x2,y2). - { Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1) - respects2 realrel" -apply (rule congruent2_commuteI [OF equiv_realrel], clarify) -apply (simp add: mult_commute add_commute) -apply (auto simp add: real_mult_congruent2_lemma) -done - -lemma real_mult: - "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) = - Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})" -by (simp add: real_mult_def UN_UN_split_split_eq - UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2]) - -lemma real_mult_commute: "(z::real) * w = w * z" -by (cases z, cases w, simp add: real_mult add_ac mult_ac) - -lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)" -apply (cases z1, cases z2, cases z3) -apply (simp add: real_mult right_distrib add_ac mult_ac) -done - -lemma real_mult_1: "(1::real) * z = z" -apply (cases z) -apply (simp add: real_mult real_one_def right_distrib - mult_1_right mult_ac add_ac) -done - -lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)" -apply (cases z1, cases z2, cases w) -apply (simp add: real_add real_mult right_distrib add_ac mult_ac) -done - -text{*one and zero are distinct*} -lemma real_zero_not_eq_one: "0 \ (1::real)" -proof - - have "(1::preal) < 1 + 1" - by (simp add: preal_self_less_add_left) - thus ?thesis - by (simp add: real_zero_def real_one_def) -qed - -instance real :: comm_ring_1 -proof - fix x y z :: real - show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc) - show "x * y = y * x" by (rule real_mult_commute) - show "1 * x = x" by (rule real_mult_1) - show "(x + y) * z = x * z + y * z" by (rule real_add_mult_distrib) - show "0 \ (1::real)" by (rule real_zero_not_eq_one) -qed - -subsection {* Inverse and Division *} - -lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0" -by (simp add: real_zero_def add_commute) - -text{*Instead of using an existential quantifier and constructing the inverse -within the proof, we could define the inverse explicitly.*} - -lemma real_mult_inverse_left_ex: "x \ 0 ==> \y. y*x = (1::real)" -apply (simp add: real_zero_def real_one_def, cases x) -apply (cut_tac x = xa and y = y in linorder_less_linear) -apply (auto dest!: less_add_left_Ex simp add: real_zero_iff) -apply (rule_tac - x = "Abs_Real (realrel``{(1, inverse (D) + 1)})" - in exI) -apply (rule_tac [2] - x = "Abs_Real (realrel``{(inverse (D) + 1, 1)})" - in exI) -apply (auto simp add: real_mult preal_mult_inverse_right ring_simps) -done - -lemma real_mult_inverse_left: "x \ 0 ==> inverse(x)*x = (1::real)" -apply (simp add: real_inverse_def) -apply (drule real_mult_inverse_left_ex, safe) -apply (rule theI, assumption, rename_tac z) -apply (subgoal_tac "(z * x) * y = z * (x * y)") -apply (simp add: mult_commute) -apply (rule mult_assoc) -done - - -subsection{*The Real Numbers form a Field*} - -instance real :: field -proof - fix x y z :: real - show "x \ 0 ==> inverse x * x = 1" by (rule real_mult_inverse_left) - show "x / y = x * inverse y" by (simp add: real_divide_def) -qed - - -text{*Inverse of zero! Useful to simplify certain equations*} - -lemma INVERSE_ZERO: "inverse 0 = (0::real)" -by (simp add: real_inverse_def) - -instance real :: division_by_zero -proof - show "inverse 0 = (0::real)" by (rule INVERSE_ZERO) -qed - - -subsection{*The @{text "\"} Ordering*} - -lemma real_le_refl: "w \ (w::real)" -by (cases w, force simp add: real_le_def) - -text{*The arithmetic decision procedure is not set up for type preal. - This lemma is currently unused, but it could simplify the proofs of the - following two lemmas.*} -lemma preal_eq_le_imp_le: - assumes eq: "a+b = c+d" and le: "c \ a" - shows "b \ (d::preal)" -proof - - have "c+d \ a+d" by (simp add: prems) - hence "a+b \ a+d" by (simp add: prems) - thus "b \ d" by simp -qed - -lemma real_le_lemma: - assumes l: "u1 + v2 \ u2 + v1" - and "x1 + v1 = u1 + y1" - and "x2 + v2 = u2 + y2" - shows "x1 + y2 \ x2 + (y1::preal)" -proof - - have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: prems) - hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: add_ac) - also have "... \ (x2+y1) + (u2+v1)" by (simp add: prems) - finally show ?thesis by simp -qed - -lemma real_le: - "(Abs_Real(realrel``{(x1,y1)}) \ Abs_Real(realrel``{(x2,y2)})) = - (x1 + y2 \ x2 + y1)" -apply (simp add: real_le_def) -apply (auto intro: real_le_lemma) -done - -lemma real_le_anti_sym: "[| z \ w; w \ z |] ==> z = (w::real)" -by (cases z, cases w, simp add: real_le) - -lemma real_trans_lemma: - assumes "x + v \ u + y" - and "u + v' \ u' + v" - and "x2 + v2 = u2 + y2" - shows "x + v' \ u' + (y::preal)" -proof - - have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: add_ac) - also have "... \ (u+y) + (u+v')" by (simp add: prems) - also have "... \ (u+y) + (u'+v)" by (simp add: prems) - also have "... = (u'+y) + (u+v)" by (simp add: add_ac) - finally show ?thesis by simp -qed - -lemma real_le_trans: "[| i \ j; j \ k |] ==> i \ (k::real)" -apply (cases i, cases j, cases k) -apply (simp add: real_le) -apply (blast intro: real_trans_lemma) -done - -instance real :: order -proof - fix u v :: real - show "u < v \ u \ v \ \ v \ u" - by (auto simp add: real_less_def intro: real_le_anti_sym) -qed (assumption | rule real_le_refl real_le_trans real_le_anti_sym)+ - -(* Axiom 'linorder_linear' of class 'linorder': *) -lemma real_le_linear: "(z::real) \ w | w \ z" -apply (cases z, cases w) -apply (auto simp add: real_le real_zero_def add_ac) -done - -instance real :: linorder - by (intro_classes, rule real_le_linear) - - -lemma real_le_eq_diff: "(x \ y) = (x-y \ (0::real))" -apply (cases x, cases y) -apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus - add_ac) -apply (simp_all add: add_assoc [symmetric]) -done - -lemma real_add_left_mono: - assumes le: "x \ y" shows "z + x \ z + (y::real)" -proof - - have "z + x - (z + y) = (z + -z) + (x - y)" - by (simp add: diff_minus add_ac) - with le show ?thesis - by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus) -qed - -lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)" -by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus) - -lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))" -by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus) - -lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y" -apply (cases x, cases y) -apply (simp add: linorder_not_le [where 'a = real, symmetric] - linorder_not_le [where 'a = preal] - real_zero_def real_le real_mult) - --{*Reduce to the (simpler) @{text "\"} relation *} -apply (auto dest!: less_add_left_Ex - simp add: add_ac mult_ac - right_distrib preal_self_less_add_left) -done - -lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y" -apply (rule real_sum_gt_zero_less) -apply (drule real_less_sum_gt_zero [of x y]) -apply (drule real_mult_order, assumption) -apply (simp add: right_distrib) -done - -instantiation real :: distrib_lattice -begin - -definition - "(inf \ real \ real \ real) = min" - -definition - "(sup \ real \ real \ real) = max" - -instance - by default (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1) - -end - - -subsection{*The Reals Form an Ordered Field*} - -instance real :: ordered_field -proof - fix x y z :: real - show "x \ y ==> z + x \ z + y" by (rule real_add_left_mono) - show "x < y ==> 0 < z ==> z * x < z * y" by (rule real_mult_less_mono2) - show "\x\ = (if x < 0 then -x else x)" by (simp only: real_abs_def) - show "sgn x = (if x=0 then 0 else if 0m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)" -apply (simp add: real_of_preal_def real_zero_def, cases x) -apply (auto simp add: real_minus add_ac) -apply (cut_tac x = x and y = y in linorder_less_linear) -apply (auto dest!: less_add_left_Ex simp add: add_assoc [symmetric]) -done - -lemma real_of_preal_leD: - "real_of_preal m1 \ real_of_preal m2 ==> m1 \ m2" -by (simp add: real_of_preal_def real_le) - -lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2" -by (auto simp add: real_of_preal_leD linorder_not_le [symmetric]) - -lemma real_of_preal_lessD: - "real_of_preal m1 < real_of_preal m2 ==> m1 < m2" -by (simp add: real_of_preal_def real_le linorder_not_le [symmetric]) - -lemma real_of_preal_less_iff [simp]: - "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)" -by (blast intro: real_of_preal_lessI real_of_preal_lessD) - -lemma real_of_preal_le_iff: - "(real_of_preal m1 \ real_of_preal m2) = (m1 \ m2)" -by (simp add: linorder_not_less [symmetric]) - -lemma real_of_preal_zero_less: "0 < real_of_preal m" -apply (insert preal_self_less_add_left [of 1 m]) -apply (auto simp add: real_zero_def real_of_preal_def - real_less_def real_le_def add_ac) -apply (rule_tac x="m + 1" in exI, rule_tac x="1" in exI) -apply (simp add: add_ac) -done - -lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0" -by (simp add: real_of_preal_zero_less) - -lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m" -proof - - from real_of_preal_minus_less_zero - show ?thesis by (blast dest: order_less_trans) -qed - - -subsection{*Theorems About the Ordering*} - -lemma real_gt_zero_preal_Ex: "(0 < x) = (\y. x = real_of_preal y)" -apply (auto simp add: real_of_preal_zero_less) -apply (cut_tac x = x in real_of_preal_trichotomy) -apply (blast elim!: real_of_preal_not_minus_gt_zero [THEN notE]) -done - -lemma real_gt_preal_preal_Ex: - "real_of_preal z < x ==> \y. x = real_of_preal y" -by (blast dest!: real_of_preal_zero_less [THEN order_less_trans] - intro: real_gt_zero_preal_Ex [THEN iffD1]) - -lemma real_ge_preal_preal_Ex: - "real_of_preal z \ x ==> \y. x = real_of_preal y" -by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex) - -lemma real_less_all_preal: "y \ 0 ==> \x. y < real_of_preal x" -by (auto elim: order_le_imp_less_or_eq [THEN disjE] - intro: real_of_preal_zero_less [THEN [2] order_less_trans] - simp add: real_of_preal_zero_less) - -lemma real_less_all_real2: "~ 0 < y ==> \x. y < real_of_preal x" -by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1]) - - -subsection{*More Lemmas*} - -lemma real_mult_left_cancel: "(c::real) \ 0 ==> (c*a=c*b) = (a=b)" -by auto - -lemma real_mult_right_cancel: "(c::real) \ 0 ==> (a*c=b*c) = (a=b)" -by auto - -lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)" - by (force elim: order_less_asym - simp add: Ring_and_Field.mult_less_cancel_right) - -lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \ y*z) = (x\y)" -apply (simp add: mult_le_cancel_right) -apply (blast intro: elim: order_less_asym) -done - -lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \ z*y) = (x\y)" -by(simp add:mult_commute) - -lemma real_inverse_gt_one: "[| (0::real) < x; x < 1 |] ==> 1 < inverse x" -by (simp add: one_less_inverse_iff) (* TODO: generalize/move *) - - -subsection {* Embedding numbers into the Reals *} - -abbreviation - real_of_nat :: "nat \ real" -where - "real_of_nat \ of_nat" - -abbreviation - real_of_int :: "int \ real" -where - "real_of_int \ of_int" - -abbreviation - real_of_rat :: "rat \ real" -where - "real_of_rat \ of_rat" - -consts - (*overloaded constant for injecting other types into "real"*) - real :: "'a => real" - -defs (overloaded) - real_of_nat_def [code unfold]: "real == real_of_nat" - real_of_int_def [code unfold]: "real == real_of_int" - -lemma real_eq_of_nat: "real = of_nat" - unfolding real_of_nat_def .. - -lemma real_eq_of_int: "real = of_int" - unfolding real_of_int_def .. - -lemma real_of_int_zero [simp]: "real (0::int) = 0" -by (simp add: real_of_int_def) - -lemma real_of_one [simp]: "real (1::int) = (1::real)" -by (simp add: real_of_int_def) - -lemma real_of_int_add [simp]: "real(x + y) = real (x::int) + real y" -by (simp add: real_of_int_def) - -lemma real_of_int_minus [simp]: "real(-x) = -real (x::int)" -by (simp add: real_of_int_def) - -lemma real_of_int_diff [simp]: "real(x - y) = real (x::int) - real y" -by (simp add: real_of_int_def) - -lemma real_of_int_mult [simp]: "real(x * y) = real (x::int) * real y" -by (simp add: real_of_int_def) - -lemma real_of_int_setsum [simp]: "real ((SUM x:A. f x)::int) = (SUM x:A. real(f x))" - apply (subst real_eq_of_int)+ - apply (rule of_int_setsum) -done - -lemma real_of_int_setprod [simp]: "real ((PROD x:A. f x)::int) = - (PROD x:A. real(f x))" - apply (subst real_eq_of_int)+ - apply (rule of_int_setprod) -done - -lemma real_of_int_zero_cancel [simp, algebra, presburger]: "(real x = 0) = (x = (0::int))" -by (simp add: real_of_int_def) - -lemma real_of_int_inject [iff, algebra, presburger]: "(real (x::int) = real y) = (x = y)" -by (simp add: real_of_int_def) - -lemma real_of_int_less_iff [iff, presburger]: "(real (x::int) < real y) = (x < y)" -by (simp add: real_of_int_def) - -lemma real_of_int_le_iff [simp, presburger]: "(real (x::int) \ real y) = (x \ y)" -by (simp add: real_of_int_def) - -lemma real_of_int_gt_zero_cancel_iff [simp, presburger]: "(0 < real (n::int)) = (0 < n)" -by (simp add: real_of_int_def) - -lemma real_of_int_ge_zero_cancel_iff [simp, presburger]: "(0 <= real (n::int)) = (0 <= n)" -by (simp add: real_of_int_def) - -lemma real_of_int_lt_zero_cancel_iff [simp, presburger]: "(real (n::int) < 0) = (n < 0)" -by (simp add: real_of_int_def) - -lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)" -by (simp add: real_of_int_def) - -lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))" -by (auto simp add: abs_if) - -lemma int_less_real_le: "((n::int) < m) = (real n + 1 <= real m)" - apply (subgoal_tac "real n + 1 = real (n + 1)") - apply (simp del: real_of_int_add) - apply auto -done - -lemma int_le_real_less: "((n::int) <= m) = (real n < real m + 1)" - apply (subgoal_tac "real m + 1 = real (m + 1)") - apply (simp del: real_of_int_add) - apply simp -done - -lemma real_of_int_div_aux: "d ~= 0 ==> (real (x::int)) / (real d) = - real (x div d) + (real (x mod d)) / (real d)" -proof - - assume "d ~= 0" - have "x = (x div d) * d + x mod d" - by auto - then have "real x = real (x div d) * real d + real(x mod d)" - by (simp only: real_of_int_mult [THEN sym] real_of_int_add [THEN sym]) - then have "real x / real d = ... / real d" - by simp - then show ?thesis - by (auto simp add: add_divide_distrib ring_simps prems) -qed - -lemma real_of_int_div: "(d::int) ~= 0 ==> d dvd n ==> - real(n div d) = real n / real d" - apply (frule real_of_int_div_aux [of d n]) - apply simp - apply (simp add: zdvd_iff_zmod_eq_0) -done - -lemma real_of_int_div2: - "0 <= real (n::int) / real (x) - real (n div x)" - apply (case_tac "x = 0") - apply simp - apply (case_tac "0 < x") - apply (simp add: compare_rls) - apply (subst real_of_int_div_aux) - apply simp - apply simp - apply (subst zero_le_divide_iff) - apply auto - apply (simp add: compare_rls) - apply (subst real_of_int_div_aux) - apply simp - apply simp - apply (subst zero_le_divide_iff) - apply auto -done - -lemma real_of_int_div3: - "real (n::int) / real (x) - real (n div x) <= 1" - apply(case_tac "x = 0") - apply simp - apply (simp add: compare_rls) - apply (subst real_of_int_div_aux) - apply assumption - apply simp - apply (subst divide_le_eq) - apply clarsimp - apply (rule conjI) - apply (rule impI) - apply (rule order_less_imp_le) - apply simp - apply (rule impI) - apply (rule order_less_imp_le) - apply simp -done - -lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x" -by (insert real_of_int_div2 [of n x], simp) - - -subsection{*Embedding the Naturals into the Reals*} - -lemma real_of_nat_zero [simp]: "real (0::nat) = 0" -by (simp add: real_of_nat_def) - -lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n" -by (simp add: real_of_nat_def) - -(*Not for addsimps: often the LHS is used to represent a positive natural*) -lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_less_iff [iff]: - "(real (n::nat) < real m) = (n < m)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_le_iff [iff]: "(real (n::nat) \ real m) = (n \ m)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_ge_zero [iff]: "0 \ real (n::nat)" -by (simp add: real_of_nat_def zero_le_imp_of_nat) - -lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)" -by (simp add: real_of_nat_def del: of_nat_Suc) - -lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n" -by (simp add: real_of_nat_def of_nat_mult) - -lemma real_of_nat_setsum [simp]: "real ((SUM x:A. f x)::nat) = - (SUM x:A. real(f x))" - apply (subst real_eq_of_nat)+ - apply (rule of_nat_setsum) -done - -lemma real_of_nat_setprod [simp]: "real ((PROD x:A. f x)::nat) = - (PROD x:A. real(f x))" - apply (subst real_eq_of_nat)+ - apply (rule of_nat_setprod) -done - -lemma real_of_card: "real (card A) = setsum (%x.1) A" - apply (subst card_eq_setsum) - apply (subst real_of_nat_setsum) - apply simp -done - -lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_diff: "n \ m ==> real (m - n) = real (m::nat) - real n" -by (simp add: add: real_of_nat_def of_nat_diff) - -lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)" -by (auto simp: real_of_nat_def) - -lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \ 0) = (n = 0)" -by (simp add: add: real_of_nat_def) - -lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0" -by (simp add: add: real_of_nat_def) - -lemma real_of_nat_ge_zero_cancel_iff [simp]: "(0 \ real (n::nat))" -by (simp add: add: real_of_nat_def) - -lemma nat_less_real_le: "((n::nat) < m) = (real n + 1 <= real m)" - apply (subgoal_tac "real n + 1 = real (Suc n)") - apply simp - apply (auto simp add: real_of_nat_Suc) -done - -lemma nat_le_real_less: "((n::nat) <= m) = (real n < real m + 1)" - apply (subgoal_tac "real m + 1 = real (Suc m)") - apply (simp add: less_Suc_eq_le) - apply (simp add: real_of_nat_Suc) -done - -lemma real_of_nat_div_aux: "0 < d ==> (real (x::nat)) / (real d) = - real (x div d) + (real (x mod d)) / (real d)" -proof - - assume "0 < d" - have "x = (x div d) * d + x mod d" - by auto - then have "real x = real (x div d) * real d + real(x mod d)" - by (simp only: real_of_nat_mult [THEN sym] real_of_nat_add [THEN sym]) - then have "real x / real d = \ / real d" - by simp - then show ?thesis - by (auto simp add: add_divide_distrib ring_simps prems) -qed - -lemma real_of_nat_div: "0 < (d::nat) ==> d dvd n ==> - real(n div d) = real n / real d" - apply (frule real_of_nat_div_aux [of d n]) - apply simp - apply (subst dvd_eq_mod_eq_0 [THEN sym]) - apply assumption -done - -lemma real_of_nat_div2: - "0 <= real (n::nat) / real (x) - real (n div x)" -apply(case_tac "x = 0") - apply (simp) -apply (simp add: compare_rls) -apply (subst real_of_nat_div_aux) - apply simp -apply simp -apply (subst zero_le_divide_iff) -apply simp -done - -lemma real_of_nat_div3: - "real (n::nat) / real (x) - real (n div x) <= 1" -apply(case_tac "x = 0") -apply (simp) -apply (simp add: compare_rls) -apply (subst real_of_nat_div_aux) - apply simp -apply simp -done - -lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x" - by (insert real_of_nat_div2 [of n x], simp) - -lemma real_of_int_real_of_nat: "real (int n) = real n" -by (simp add: real_of_nat_def real_of_int_def int_eq_of_nat) - -lemma real_of_int_of_nat_eq [simp]: "real (of_nat n :: int) = real n" -by (simp add: real_of_int_def real_of_nat_def) - -lemma real_nat_eq_real [simp]: "0 <= x ==> real(nat x) = real x" - apply (subgoal_tac "real(int(nat x)) = real(nat x)") - apply force - apply (simp only: real_of_int_real_of_nat) -done - - -subsection{* Rationals *} - -lemma Rats_real_nat[simp]: "real(n::nat) \ \" -by (simp add: real_eq_of_nat) - - -lemma Rats_eq_int_div_int: - "\ = { real(i::int)/real(j::int) |i j. j \ 0}" (is "_ = ?S") -proof - show "\ \ ?S" - proof - fix x::real assume "x : \" - then obtain r where "x = of_rat r" unfolding Rats_def .. - have "of_rat r : ?S" - by (cases r)(auto simp add:of_rat_rat real_eq_of_int) - thus "x : ?S" using `x = of_rat r` by simp - qed -next - show "?S \ \" - proof(auto simp:Rats_def) - fix i j :: int assume "j \ 0" - hence "real i / real j = of_rat(Fract i j)" - by (simp add:of_rat_rat real_eq_of_int) - thus "real i / real j \ range of_rat" by blast - qed -qed - -lemma Rats_eq_int_div_nat: - "\ = { real(i::int)/real(n::nat) |i n. n \ 0}" -proof(auto simp:Rats_eq_int_div_int) - fix i j::int assume "j \ 0" - show "EX (i'::int) (n::nat). real i/real j = real i'/real n \ 00" - hence "real i/real j = real i/real(nat j) \ 00" - hence "real i/real j = real(-i)/real(nat(-j)) \ 00` - by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat) - thus ?thesis by blast - qed -next - fix i::int and n::nat assume "0 < n" - hence "real i/real n = real i/real(int n) \ int n \ 0" by simp - thus "\(i'::int) j::int. real i/real n = real i'/real j \ j \ 0" by blast -qed - -lemma Rats_abs_nat_div_natE: - assumes "x \ \" - obtains m n where "n \ 0" and "\x\ = real m / real n" and "gcd m n = 1" -proof - - from `x \ \` obtain i::int and n::nat where "n \ 0" and "x = real i / real n" - by(auto simp add: Rats_eq_int_div_nat) - hence "\x\ = real(nat(abs i)) / real n" by simp - then obtain m :: nat where x_rat: "\x\ = real m / real n" by blast - let ?gcd = "gcd m n" - from `n\0` have gcd: "?gcd \ 0" by (simp add: gcd_zero) - let ?k = "m div ?gcd" - let ?l = "n div ?gcd" - let ?gcd' = "gcd ?k ?l" - have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m" - by (rule dvd_mult_div_cancel) - have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n" - by (rule dvd_mult_div_cancel) - from `n\0` and gcd_l have "?l \ 0" by (auto iff del: neq0_conv) - moreover - have "\x\ = real ?k / real ?l" - proof - - from gcd have "real ?k / real ?l = - real (?gcd * ?k) / real (?gcd * ?l)" by simp - also from gcd_k and gcd_l have "\ = real m / real n" by simp - also from x_rat have "\ = \x\" .. - finally show ?thesis .. - qed - moreover - have "?gcd' = 1" - proof - - have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)" - by (rule gcd_mult_distrib2) - with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp - with gcd show ?thesis by simp - qed - ultimately show ?thesis .. -qed - - -subsection{*Numerals and Arithmetic*} - -instantiation real :: number_ring -begin - -definition - real_number_of_def [code del]: "number_of w = real_of_int w" - -instance - by intro_classes (simp add: real_number_of_def) - -end - -lemma [code unfold, symmetric, code post]: - "number_of k = real_of_int (number_of k)" - unfolding number_of_is_id real_number_of_def .. - - -text{*Collapse applications of @{term real} to @{term number_of}*} -lemma real_number_of [simp]: "real (number_of v :: int) = number_of v" -by (simp add: real_of_int_def of_int_number_of_eq) - -lemma real_of_nat_number_of [simp]: - "real (number_of v :: nat) = - (if neg (number_of v :: int) then 0 - else (number_of v :: real))" -by (simp add: real_of_int_real_of_nat [symmetric] int_nat_number_of) - - -use "real_arith.ML" -declaration {* K real_arith_setup *} - - -subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*} - -text{*Needed in this non-standard form by Hyperreal/Transcendental*} -lemma real_0_le_divide_iff: - "((0::real) \ x/y) = ((x \ 0 | 0 \ y) & (0 \ x | y \ 0))" -by (simp add: real_divide_def zero_le_mult_iff, auto) - -lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" -by arith - -lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)" -by auto - -lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)" -by auto - -lemma real_0_less_add_iff: "((0::real) < x+y) = (-x < y)" -by auto - -lemma real_add_le_0_iff: "(x+y \ (0::real)) = (y \ -x)" -by auto - -lemma real_0_le_add_iff: "((0::real) \ x+y) = (-x \ y)" -by auto - - -(* -FIXME: we should have this, as for type int, but many proofs would break. -It replaces x+-y by x-y. -declare real_diff_def [symmetric, simp] -*) - -subsubsection{*Density of the Reals*} - -lemma real_lbound_gt_zero: - "[| (0::real) < d1; 0 < d2 |] ==> \e. 0 < e & e < d1 & e < d2" -apply (rule_tac x = " (min d1 d2) /2" in exI) -apply (simp add: min_def) -done - - -text{*Similar results are proved in @{text Ring_and_Field}*} -lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)" - by auto - -lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y" - by auto - - -subsection{*Absolute Value Function for the Reals*} - -lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))" -by (simp add: abs_if) - -(* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *) -lemma abs_le_interval_iff: "(abs x \ r) = (-r\x & x\(r::real))" -by (force simp add: OrderedGroup.abs_le_iff) - -lemma abs_add_one_gt_zero [simp]: "(0::real) < 1 + abs(x)" -by (simp add: abs_if) - -lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)" -by (rule abs_of_nonneg [OF real_of_nat_ge_zero]) - -lemma abs_add_one_not_less_self [simp]: "~ abs(x) + (1::real) < x" -by simp - -lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \ abs(x + -l) + abs(y + -m)" -by simp - -instance real :: lordered_ring -proof - fix a::real - show "abs a = sup a (-a)" - by (auto simp add: real_abs_def sup_real_def) -qed - - -subsection {* Implementation of rational real numbers *} - -definition Ratreal :: "rat \ real" where - [simp]: "Ratreal = of_rat" - -code_datatype Ratreal - -lemma Ratreal_number_collapse [code post]: - "Ratreal 0 = 0" - "Ratreal 1 = 1" - "Ratreal (number_of k) = number_of k" -by simp_all - -lemma zero_real_code [code, code unfold]: - "0 = Ratreal 0" -by simp - -lemma one_real_code [code, code unfold]: - "1 = Ratreal 1" -by simp - -lemma number_of_real_code [code unfold]: - "number_of k = Ratreal (number_of k)" -by simp - -lemma Ratreal_number_of_quotient [code post]: - "Ratreal (number_of r) / Ratreal (number_of s) = number_of r / number_of s" -by simp - -lemma Ratreal_number_of_quotient2 [code post]: - "Ratreal (number_of r / number_of s) = number_of r / number_of s" -unfolding Ratreal_number_of_quotient [symmetric] Ratreal_def of_rat_divide .. - -instantiation real :: eq -begin - -definition "eq_class.eq (x\real) y \ x - y = 0" - -instance by default (simp add: eq_real_def) - -lemma real_eq_code [code]: "eq_class.eq (Ratreal x) (Ratreal y) \ eq_class.eq x y" - by (simp add: eq_real_def eq) - -lemma real_eq_refl [code nbe]: - "eq_class.eq (x::real) x \ True" - by (rule HOL.eq_refl) - -end - -lemma real_less_eq_code [code]: "Ratreal x \ Ratreal y \ x \ y" - by (simp add: of_rat_less_eq) - -lemma real_less_code [code]: "Ratreal x < Ratreal y \ x < y" - by (simp add: of_rat_less) - -lemma real_plus_code [code]: "Ratreal x + Ratreal y = Ratreal (x + y)" - by (simp add: of_rat_add) - -lemma real_times_code [code]: "Ratreal x * Ratreal y = Ratreal (x * y)" - by (simp add: of_rat_mult) - -lemma real_uminus_code [code]: "- Ratreal x = Ratreal (- x)" - by (simp add: of_rat_minus) - -lemma real_minus_code [code]: "Ratreal x - Ratreal y = Ratreal (x - y)" - by (simp add: of_rat_diff) - -lemma real_inverse_code [code]: "inverse (Ratreal x) = Ratreal (inverse x)" - by (simp add: of_rat_inverse) - -lemma real_divide_code [code]: "Ratreal x / Ratreal y = Ratreal (x / y)" - by (simp add: of_rat_divide) - -text {* Setup for SML code generator *} - -types_code - real ("(int */ int)") -attach (term_of) {* -fun term_of_real (p, q) = - let - val rT = HOLogic.realT - in - if q = 1 orelse p = 0 then HOLogic.mk_number rT p - else @{term "op / \ real \ real \ real"} $ - HOLogic.mk_number rT p $ HOLogic.mk_number rT q - end; -*} -attach (test) {* -fun gen_real i = - let - val p = random_range 0 i; - val q = random_range 1 (i + 1); - val g = Integer.gcd p q; - val p' = p div g; - val q' = q div g; - val r = (if one_of [true, false] then p' else ~ p', - if p' = 0 then 0 else q') - in - (r, fn () => term_of_real r) - end; -*} - -consts_code - Ratreal ("(_)") - -consts_code - "of_int :: int \ real" ("\real'_of'_int") -attach {* -fun real_of_int 0 = (0, 0) - | real_of_int i = (i, 1); -*} - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/RealPow.thy --- a/src/HOL/Real/RealPow.thy Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,282 +0,0 @@ -(* Title : HOL/Real/RealPow.thy - ID : $Id$ - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge -*) - -header {* Natural powers theory *} - -theory RealPow -imports RealDef -uses ("float_syntax.ML") -begin - -declare abs_mult_self [simp] - -instantiation real :: recpower -begin - -primrec power_real where - realpow_0: "r ^ 0 = (1\real)" - | realpow_Suc: "r ^ Suc n = (r\real) * r ^ n" - -instance proof - fix z :: real - fix n :: nat - show "z^0 = 1" by simp - show "z^(Suc n) = z * (z^n)" by simp -qed - -end - - -lemma two_realpow_ge_one [simp]: "(1::real) \ 2 ^ n" -by simp - -lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n" -apply (induct "n") -apply (auto simp add: real_of_nat_Suc) -apply (subst mult_2) -apply (rule add_less_le_mono) -apply (auto simp add: two_realpow_ge_one) -done - -lemma realpow_Suc_le_self: "[| 0 \ r; r \ (1::real) |] ==> r ^ Suc n \ r" -by (insert power_decreasing [of 1 "Suc n" r], simp) - -lemma realpow_minus_mult [rule_format]: - "0 < n --> (x::real) ^ (n - 1) * x = x ^ n" -apply (simp split add: nat_diff_split) -done - -lemma realpow_two_mult_inverse [simp]: - "r \ 0 ==> r * inverse r ^Suc (Suc 0) = inverse (r::real)" -by (simp add: real_mult_assoc [symmetric]) - -lemma realpow_two_minus [simp]: "(-x)^Suc (Suc 0) = (x::real)^Suc (Suc 0)" -by simp - -lemma realpow_two_diff: - "(x::real)^Suc (Suc 0) - y^Suc (Suc 0) = (x - y) * (x + y)" -apply (unfold real_diff_def) -apply (simp add: ring_simps) -done - -lemma realpow_two_disj: - "((x::real)^Suc (Suc 0) = y^Suc (Suc 0)) = (x = y | x = -y)" -apply (cut_tac x = x and y = y in realpow_two_diff) -apply (auto simp del: realpow_Suc) -done - -lemma realpow_real_of_nat: "real (m::nat) ^ n = real (m ^ n)" -apply (induct "n") -apply (auto simp add: real_of_nat_one real_of_nat_mult) -done - -lemma realpow_real_of_nat_two_pos [simp] : "0 < real (Suc (Suc 0) ^ n)" -apply (induct "n") -apply (auto simp add: real_of_nat_mult zero_less_mult_iff) -done - -(* used by AFP Integration theory *) -lemma realpow_increasing: - "[|(0::real) \ x; 0 \ y; x ^ Suc n \ y ^ Suc n|] ==> x \ y" - by (rule power_le_imp_le_base) - - -subsection{*Literal Arithmetic Involving Powers, Type @{typ real}*} - -lemma real_of_int_power: "real (x::int) ^ n = real (x ^ n)" -apply (induct "n") -apply (simp_all add: nat_mult_distrib) -done -declare real_of_int_power [symmetric, simp] - -lemma power_real_number_of: - "(number_of v :: real) ^ n = real ((number_of v :: int) ^ n)" -by (simp only: real_number_of [symmetric] real_of_int_power) - -declare power_real_number_of [of _ "number_of w", standard, simp] - - -subsection {* Properties of Squares *} - -lemma sum_squares_ge_zero: - fixes x y :: "'a::ordered_ring_strict" - shows "0 \ x * x + y * y" -by (intro add_nonneg_nonneg zero_le_square) - -lemma not_sum_squares_lt_zero: - fixes x y :: "'a::ordered_ring_strict" - shows "\ x * x + y * y < 0" -by (simp add: linorder_not_less sum_squares_ge_zero) - -lemma sum_nonneg_eq_zero_iff: - fixes x y :: "'a::pordered_ab_group_add" - assumes x: "0 \ x" and y: "0 \ y" - shows "(x + y = 0) = (x = 0 \ y = 0)" -proof (auto) - from y have "x + 0 \ x + y" by (rule add_left_mono) - also assume "x + y = 0" - finally have "x \ 0" by simp - thus "x = 0" using x by (rule order_antisym) -next - from x have "0 + y \ x + y" by (rule add_right_mono) - also assume "x + y = 0" - finally have "y \ 0" by simp - thus "y = 0" using y by (rule order_antisym) -qed - -lemma sum_squares_eq_zero_iff: - fixes x y :: "'a::ordered_ring_strict" - shows "(x * x + y * y = 0) = (x = 0 \ y = 0)" -by (simp add: sum_nonneg_eq_zero_iff) - -lemma sum_squares_le_zero_iff: - fixes x y :: "'a::ordered_ring_strict" - shows "(x * x + y * y \ 0) = (x = 0 \ y = 0)" -by (simp add: order_le_less not_sum_squares_lt_zero sum_squares_eq_zero_iff) - -lemma sum_squares_gt_zero_iff: - fixes x y :: "'a::ordered_ring_strict" - shows "(0 < x * x + y * y) = (x \ 0 \ y \ 0)" -by (simp add: order_less_le sum_squares_ge_zero sum_squares_eq_zero_iff) - -lemma sum_power2_ge_zero: - fixes x y :: "'a::{ordered_idom,recpower}" - shows "0 \ x\ + y\" -unfolding power2_eq_square by (rule sum_squares_ge_zero) - -lemma not_sum_power2_lt_zero: - fixes x y :: "'a::{ordered_idom,recpower}" - shows "\ x\ + y\ < 0" -unfolding power2_eq_square by (rule not_sum_squares_lt_zero) - -lemma sum_power2_eq_zero_iff: - fixes x y :: "'a::{ordered_idom,recpower}" - shows "(x\ + y\ = 0) = (x = 0 \ y = 0)" -unfolding power2_eq_square by (rule sum_squares_eq_zero_iff) - -lemma sum_power2_le_zero_iff: - fixes x y :: "'a::{ordered_idom,recpower}" - shows "(x\ + y\ \ 0) = (x = 0 \ y = 0)" -unfolding power2_eq_square by (rule sum_squares_le_zero_iff) - -lemma sum_power2_gt_zero_iff: - fixes x y :: "'a::{ordered_idom,recpower}" - shows "(0 < x\ + y\) = (x \ 0 \ y \ 0)" -unfolding power2_eq_square by (rule sum_squares_gt_zero_iff) - - -subsection{* Squares of Reals *} - -lemma real_two_squares_add_zero_iff [simp]: - "(x * x + y * y = 0) = ((x::real) = 0 \ y = 0)" -by (rule sum_squares_eq_zero_iff) - -lemma real_sum_squares_cancel: "x * x + y * y = 0 ==> x = (0::real)" -by simp - -lemma real_sum_squares_cancel2: "x * x + y * y = 0 ==> y = (0::real)" -by simp - -lemma real_mult_self_sum_ge_zero: "(0::real) \ x*x + y*y" -by (rule sum_squares_ge_zero) - -lemma real_sum_squares_cancel_a: "x * x = -(y * y) ==> x = (0::real) & y=0" -by (simp add: real_add_eq_0_iff [symmetric]) - -lemma real_squared_diff_one_factored: "x*x - (1::real) = (x + 1)*(x - 1)" -by (simp add: left_distrib right_diff_distrib) - -lemma real_mult_is_one [simp]: "(x*x = (1::real)) = (x = 1 | x = - 1)" -apply auto -apply (drule right_minus_eq [THEN iffD2]) -apply (auto simp add: real_squared_diff_one_factored) -done - -lemma real_sum_squares_not_zero: "x ~= 0 ==> x * x + y * y ~= (0::real)" -by simp - -lemma real_sum_squares_not_zero2: "y ~= 0 ==> x * x + y * y ~= (0::real)" -by simp - -lemma realpow_two_sum_zero_iff [simp]: - "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)" -by (rule sum_power2_eq_zero_iff) - -lemma realpow_two_le_add_order [simp]: "(0::real) \ u ^ 2 + v ^ 2" -by (rule sum_power2_ge_zero) - -lemma realpow_two_le_add_order2 [simp]: "(0::real) \ u ^ 2 + v ^ 2 + w ^ 2" -by (intro add_nonneg_nonneg zero_le_power2) - -lemma real_sum_square_gt_zero: "x ~= 0 ==> (0::real) < x * x + y * y" -by (simp add: sum_squares_gt_zero_iff) - -lemma real_sum_square_gt_zero2: "y ~= 0 ==> (0::real) < x * x + y * y" -by (simp add: sum_squares_gt_zero_iff) - -lemma real_minus_mult_self_le [simp]: "-(u * u) \ (x * (x::real))" -by (rule_tac j = 0 in real_le_trans, auto) - -lemma realpow_square_minus_le [simp]: "-(u ^ 2) \ (x::real) ^ 2" -by (auto simp add: power2_eq_square) - -(* The following theorem is by Benjamin Porter *) -lemma real_sq_order: - fixes x::real - assumes xgt0: "0 \ x" and ygt0: "0 \ y" and sq: "x^2 \ y^2" - shows "x \ y" -proof - - from sq have "x ^ Suc (Suc 0) \ y ^ Suc (Suc 0)" - by (simp only: numeral_2_eq_2) - thus "x \ y" using ygt0 - by (rule power_le_imp_le_base) -qed - - -subsection {*Various Other Theorems*} - -lemma real_le_add_half_cancel: "(x + y/2 \ (y::real)) = (x \ y /2)" -by auto - -lemma real_minus_half_eq [simp]: "(x::real) - x/2 = x/2" -by auto - -lemma real_mult_inverse_cancel: - "[|(0::real) < x; 0 < x1; x1 * y < x * u |] - ==> inverse x * y < inverse x1 * u" -apply (rule_tac c=x in mult_less_imp_less_left) -apply (auto simp add: real_mult_assoc [symmetric]) -apply (simp (no_asm) add: mult_ac) -apply (rule_tac c=x1 in mult_less_imp_less_right) -apply (auto simp add: mult_ac) -done - -lemma real_mult_inverse_cancel2: - "[|(0::real) < x;0 < x1; x1 * y < x * u |] ==> y * inverse x < u * inverse x1" -apply (auto dest: real_mult_inverse_cancel simp add: mult_ac) -done - -lemma inverse_real_of_nat_gt_zero [simp]: "0 < inverse (real (Suc n))" -by simp - -lemma inverse_real_of_nat_ge_zero [simp]: "0 \ inverse (real (Suc n))" -by simp - -lemma realpow_num_eq_if: "(m::real) ^ n = (if n=0 then 1 else m * m ^ (n - 1))" -by (case_tac "n", auto) - -subsection{* Float syntax *} - -syntax "_Float" :: "float_const \ 'a" ("_") - -use "float_syntax.ML" -setup FloatSyntax.setup - -text{* Test: *} -lemma "123.456 = -111.111 + 200 + 30 + 4 + 5/10 + 6/100 + (7/1000::real)" -by simp - -end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/RealVector.thy --- a/src/HOL/Real/RealVector.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Real/RealVector.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,12 +1,11 @@ -(* Title: RealVector.thy - ID: $Id$ +(* Title: HOL/Real/RealVector.thy Author: Brian Huffman *) header {* Vector Spaces and Algebras over the Reals *} theory RealVector -imports RealPow +imports "~~/src/HOL/RealPow" begin subsection {* Locale for additive functions *} diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/float_arith.ML --- a/src/HOL/Real/float_arith.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,222 +0,0 @@ -(* Title: HOL/Real/float_arith.ML - ID: $Id$ - Author: Steven Obua -*) - -signature FLOAT_ARITH = -sig - exception Destruct_floatstr of string - val destruct_floatstr: (char -> bool) -> (char -> bool) -> string -> bool * string * string * bool * string - - exception Floating_point of string - val approx_dec_by_bin: int -> Float.float -> Float.float * Float.float - val approx_decstr_by_bin: int -> string -> Float.float * Float.float - - val mk_float: Float.float -> term - val dest_float: term -> Float.float - - val approx_float: int -> (Float.float * Float.float -> Float.float * Float.float) - -> string -> term * term -end; - -structure FloatArith : FLOAT_ARITH = -struct - -exception Destruct_floatstr of string; - -fun destruct_floatstr isDigit isExp number = - let - val numlist = filter (not o Char.isSpace) (String.explode number) - - fun countsigns ((#"+")::cs) = countsigns cs - | countsigns ((#"-")::cs) = - let - val (positive, rest) = countsigns cs - in - (not positive, rest) - end - | countsigns cs = (true, cs) - - fun readdigits [] = ([], []) - | readdigits (q as c::cs) = - if (isDigit c) then - let - val (digits, rest) = readdigits cs - in - (c::digits, rest) - end - else - ([], q) - - fun readfromexp_helper cs = - let - val (positive, rest) = countsigns cs - val (digits, rest') = readdigits rest - in - case rest' of - [] => (positive, digits) - | _ => raise (Destruct_floatstr number) - end - - fun readfromexp [] = (true, []) - | readfromexp (c::cs) = - if isExp c then - readfromexp_helper cs - else - raise (Destruct_floatstr number) - - fun readfromdot [] = ([], readfromexp []) - | readfromdot ((#".")::cs) = - let - val (digits, rest) = readdigits cs - val exp = readfromexp rest - in - (digits, exp) - end - | readfromdot cs = readfromdot ((#".")::cs) - - val (positive, numlist) = countsigns numlist - val (digits1, numlist) = readdigits numlist - val (digits2, exp) = readfromdot numlist - in - (positive, String.implode digits1, String.implode digits2, fst exp, String.implode (snd exp)) - end - -exception Floating_point of string; - -val ln2_10 = Math.ln 10.0 / Math.ln 2.0; -fun exp5 x = Integer.pow x 5; -fun exp10 x = Integer.pow x 10; -fun exp2 x = Integer.pow x 2; - -fun find_most_significant q r = - let - fun int2real i = - case (Real.fromString o string_of_int) i of - SOME r => r - | NONE => raise (Floating_point "int2real") - fun subtract (q, r) (q', r') = - if r <= r' then - (q - q' * exp10 (r' - r), r) - else - (q * exp10 (r - r') - q', r') - fun bin2dec d = - if 0 <= d then - (exp2 d, 0) - else - (exp5 (~ d), d) - - val L = Real.floor (int2real (IntInf.log2 q) + int2real r * ln2_10) - val L1 = L + 1 - - val (q1, r1) = subtract (q, r) (bin2dec L1) - in - if 0 <= q1 then - let - val (q2, r2) = subtract (q, r) (bin2dec (L1 + 1)) - in - if 0 <= q2 then - raise (Floating_point "find_most_significant") - else - (L1, (q1, r1)) - end - else - let - val (q0, r0) = subtract (q, r) (bin2dec L) - in - if 0 <= q0 then - (L, (q0, r0)) - else - raise (Floating_point "find_most_significant") - end - end - -fun approx_dec_by_bin n (q,r) = - let - fun addseq acc d' [] = acc - | addseq acc d' (d::ds) = addseq (acc + exp2 (d - d')) d' ds - - fun seq2bin [] = (0, 0) - | seq2bin (d::ds) = (addseq 0 d ds + 1, d) - - fun approx d_seq d0 precision (q,r) = - if q = 0 then - let val x = seq2bin d_seq in - (x, x) - end - else - let - val (d, (q', r')) = find_most_significant q r - in - if precision < d0 - d then - let - val d' = d0 - precision - val x1 = seq2bin (d_seq) - val x2 = (fst x1 * exp2 (snd x1 - d') + 1, d') (* = seq2bin (d'::d_seq) *) - in - (x1, x2) - end - else - approx (d::d_seq) d0 precision (q', r') - end - - fun approx_start precision (q, r) = - if q = 0 then - ((0, 0), (0, 0)) - else - let - val (d, (q', r')) = find_most_significant q r - in - if precision <= 0 then - let - val x1 = seq2bin [d] - in - if q' = 0 then - (x1, x1) - else - (x1, seq2bin [d + 1]) - end - else - approx [d] d precision (q', r') - end - in - if 0 <= q then - approx_start n (q,r) - else - let - val ((a1,b1), (a2, b2)) = approx_start n (~ q, r) - in - ((~ a2, b2), (~ a1, b1)) - end - end - -fun approx_decstr_by_bin n decstr = - let - fun str2int s = the_default 0 (Int.fromString s) - fun signint p x = if p then x else ~ x - - val (p, d1, d2, ep, e) = destruct_floatstr Char.isDigit (fn e => e = #"e" orelse e = #"E") decstr - val s = size d2 - - val q = signint p (str2int d1 * exp10 s + str2int d2) - val r = signint ep (str2int e) - s - in - approx_dec_by_bin n (q,r) - end - -fun mk_float (a, b) = @{term "float"} $ - HOLogic.mk_prod (pairself (HOLogic.mk_number HOLogic.intT) (a, b)); - -fun dest_float (Const ("Float.float", _) $ (Const ("Pair", _) $ a $ b)) = - pairself (snd o HOLogic.dest_number) (a, b) - | dest_float t = ((snd o HOLogic.dest_number) t, 0); - -fun approx_float prec f value = - let - val interval = approx_decstr_by_bin prec value - val (flower, fupper) = f interval - in - (mk_float flower, mk_float fupper) - end; - -end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/float_syntax.ML --- a/src/HOL/Real/float_syntax.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,47 +0,0 @@ -(* ID: $Id$ - Authors: Tobias Nipkow, TU Muenchen - -Concrete syntax for floats -*) - -signature FLOAT_SYNTAX = -sig - val setup: theory -> theory -end; - -structure FloatSyntax: FLOAT_SYNTAX = -struct - -(* parse translation *) - -local - -fun mk_number i = - let - fun mk 0 = Syntax.const @{const_name Int.Pls} - | mk ~1 = Syntax.const @{const_name Int.Min} - | mk i = let val (q, r) = Integer.div_mod i 2 - in HOLogic.mk_bit r $ (mk q) end; - in Syntax.const @{const_name Int.number_of} $ mk i end; - -fun mk_frac str = - let - val {mant=i, exp = n} = Syntax.read_float str; - val exp = Syntax.const @{const_name "Power.power"}; - val ten = mk_number 10; - val exp10 = if n=1 then ten else exp $ ten $ (mk_number n); - in (Syntax.const @{const_name "divide"}) $ (mk_number i) $ exp10 end -in - -fun float_tr (*"_Float"*) [t as Const (str, _)] = mk_frac str - | float_tr (*"_Float"*) ts = raise TERM ("float_tr", ts); - -end; - - -(* theory setup *) - -val setup = - Sign.add_trfuns ([], [("_Float", float_tr)], [], []); - -end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/rat_arith.ML --- a/src/HOL/Real/rat_arith.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,51 +0,0 @@ -(* Title: HOL/Real/rat_arith.ML - ID: $Id$ - Author: Lawrence C Paulson - Copyright 2004 University of Cambridge - -Simprocs for common factor cancellation & Rational coefficient handling - -Instantiation of the generic linear arithmetic package for type rat. -*) - -local - -val simprocs = field_cancel_numeral_factors - -val simps = - [@{thm order_less_irrefl}, @{thm neg_less_iff_less}, @{thm True_implies_equals}, - read_instantiate @{context} [(("a", 0), "(number_of ?v)")] @{thm right_distrib}, - @{thm divide_1}, @{thm divide_zero_left}, - @{thm times_divide_eq_right}, @{thm times_divide_eq_left}, - @{thm minus_divide_left} RS sym, @{thm minus_divide_right} RS sym, - @{thm of_int_0}, @{thm of_int_1}, @{thm of_int_add}, - @{thm of_int_minus}, @{thm of_int_diff}, - @{thm of_int_mult}, @{thm of_int_of_nat_eq}] - -val nat_inj_thms = [@{thm of_nat_le_iff} RS iffD2, - @{thm of_nat_eq_iff} RS iffD2] -(* not needed because x < (y::nat) can be rewritten as Suc x <= y: - of_nat_less_iff RS iffD2 *) - -val int_inj_thms = [@{thm of_int_le_iff} RS iffD2, - @{thm of_int_eq_iff} RS iffD2] -(* not needed because x < (y::int) can be rewritten as x + 1 <= y: - of_int_less_iff RS iffD2 *) - -in - -val ratT = Type ("Rational.rat", []) - -val rat_arith_setup = - LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} => - {add_mono_thms = add_mono_thms, - mult_mono_thms = mult_mono_thms, - inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms, - lessD = lessD, (*Can't change lessD: the rats are dense!*) - neqE = neqE, - simpset = simpset addsimps simps - addsimprocs simprocs}) #> - arith_inj_const (@{const_name of_nat}, HOLogic.natT --> ratT) #> - arith_inj_const (@{const_name of_int}, HOLogic.intT --> ratT) - -end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Real/real_arith.ML --- a/src/HOL/Real/real_arith.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,42 +0,0 @@ -(* Title: HOL/Real/real_arith.ML - ID: $Id$ - Author: Tobias Nipkow, TU Muenchen - Copyright 1999 TU Muenchen - -Simprocs for common factor cancellation & Rational coefficient handling - -Instantiation of the generic linear arithmetic package for type real. -*) - -local - -val simps = [@{thm real_of_nat_zero}, @{thm real_of_nat_Suc}, @{thm real_of_nat_add}, - @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one}, - @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff}, - @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq}, - @{thm real_of_nat_number_of}, @{thm real_number_of}] - -val nat_inj_thms = [@{thm real_of_nat_le_iff} RS iffD2, - @{thm real_of_nat_inject} RS iffD2] -(* not needed because x < (y::nat) can be rewritten as Suc x <= y: - real_of_nat_less_iff RS iffD2 *) - -val int_inj_thms = [@{thm real_of_int_le_iff} RS iffD2, - @{thm real_of_int_inject} RS iffD2] -(* not needed because x < (y::int) can be rewritten as x + 1 <= y: - real_of_int_less_iff RS iffD2 *) - -in - -val real_arith_setup = - LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} => - {add_mono_thms = add_mono_thms, - mult_mono_thms = mult_mono_thms, - inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms, - lessD = lessD, (*Can't change lessD: the reals are dense!*) - neqE = neqE, - simpset = simpset addsimps simps}) #> - arith_inj_const (@{const_name real}, HOLogic.natT --> HOLogic.realT) #> - arith_inj_const (@{const_name real}, HOLogic.intT --> HOLogic.realT) - -end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/RealDef.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/RealDef.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,1158 @@ +(* Title : HOL/RealDef.thy + Author : Jacques D. Fleuriot + Copyright : 1998 University of Cambridge + Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 + Additional contributions by Jeremy Avigad +*) + +header{*Defining the Reals from the Positive Reals*} + +theory RealDef +imports PReal +uses ("Tools/real_arith.ML") +begin + +definition + realrel :: "((preal * preal) * (preal * preal)) set" where + [code del]: "realrel = {p. \x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}" + +typedef (Real) real = "UNIV//realrel" + by (auto simp add: quotient_def) + +definition + (** these don't use the overloaded "real" function: users don't see them **) + real_of_preal :: "preal => real" where + [code del]: "real_of_preal m = Abs_Real (realrel `` {(m + 1, 1)})" + +instantiation real :: "{zero, one, plus, minus, uminus, times, inverse, ord, abs, sgn}" +begin + +definition + real_zero_def [code del]: "0 = Abs_Real(realrel``{(1, 1)})" + +definition + real_one_def [code del]: "1 = Abs_Real(realrel``{(1 + 1, 1)})" + +definition + real_add_def [code del]: "z + w = + contents (\(x,y) \ Rep_Real(z). \(u,v) \ Rep_Real(w). + { Abs_Real(realrel``{(x+u, y+v)}) })" + +definition + real_minus_def [code del]: "- r = contents (\(x,y) \ Rep_Real(r). { Abs_Real(realrel``{(y,x)}) })" + +definition + real_diff_def [code del]: "r - (s::real) = r + - s" + +definition + real_mult_def [code del]: + "z * w = + contents (\(x,y) \ Rep_Real(z). \(u,v) \ Rep_Real(w). + { Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })" + +definition + real_inverse_def [code del]: "inverse (R::real) = (THE S. (R = 0 & S = 0) | S * R = 1)" + +definition + real_divide_def [code del]: "R / (S::real) = R * inverse S" + +definition + real_le_def [code del]: "z \ (w::real) \ + (\x y u v. x+v \ u+y & (x,y) \ Rep_Real z & (u,v) \ Rep_Real w)" + +definition + real_less_def [code del]: "x < (y\real) \ x \ y \ x \ y" + +definition + real_abs_def: "abs (r::real) = (if r < 0 then - r else r)" + +definition + real_sgn_def: "sgn (x::real) = (if x=0 then 0 else if 0 realrel) = (x1 + y2 = x2 + y1)" +by (simp add: realrel_def) + +lemma equiv_realrel: "equiv UNIV realrel" +apply (auto simp add: equiv_def refl_def sym_def trans_def realrel_def) +apply (blast dest: preal_trans_lemma) +done + +text{*Reduces equality of equivalence classes to the @{term realrel} relation: + @{term "(realrel `` {x} = realrel `` {y}) = ((x,y) \ realrel)"} *} +lemmas equiv_realrel_iff = + eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I] + +declare equiv_realrel_iff [simp] + + +lemma realrel_in_real [simp]: "realrel``{(x,y)}: Real" +by (simp add: Real_def realrel_def quotient_def, blast) + +declare Abs_Real_inject [simp] +declare Abs_Real_inverse [simp] + + +text{*Case analysis on the representation of a real number as an equivalence + class of pairs of positive reals.*} +lemma eq_Abs_Real [case_names Abs_Real, cases type: real]: + "(!!x y. z = Abs_Real(realrel``{(x,y)}) ==> P) ==> P" +apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE]) +apply (drule arg_cong [where f=Abs_Real]) +apply (auto simp add: Rep_Real_inverse) +done + + +subsection {* Addition and Subtraction *} + +lemma real_add_congruent2_lemma: + "[|a + ba = aa + b; ab + bc = ac + bb|] + ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))" +apply (simp add: add_assoc) +apply (rule add_left_commute [of ab, THEN ssubst]) +apply (simp add: add_assoc [symmetric]) +apply (simp add: add_ac) +done + +lemma real_add: + "Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) = + Abs_Real (realrel``{(x+u, y+v)})" +proof - + have "(\z w. (\(x,y). (\(u,v). {Abs_Real (realrel `` {(x+u, y+v)})}) w) z) + respects2 realrel" + by (simp add: congruent2_def, blast intro: real_add_congruent2_lemma) + thus ?thesis + by (simp add: real_add_def UN_UN_split_split_eq + UN_equiv_class2 [OF equiv_realrel equiv_realrel]) +qed + +lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})" +proof - + have "(\(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel" + by (simp add: congruent_def add_commute) + thus ?thesis + by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel]) +qed + +instance real :: ab_group_add +proof + fix x y z :: real + show "(x + y) + z = x + (y + z)" + by (cases x, cases y, cases z, simp add: real_add add_assoc) + show "x + y = y + x" + by (cases x, cases y, simp add: real_add add_commute) + show "0 + x = x" + by (cases x, simp add: real_add real_zero_def add_ac) + show "- x + x = 0" + by (cases x, simp add: real_minus real_add real_zero_def add_commute) + show "x - y = x + - y" + by (simp add: real_diff_def) +qed + + +subsection {* Multiplication *} + +lemma real_mult_congruent2_lemma: + "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==> + x * x1 + y * y1 + (x * y2 + y * x2) = + x * x2 + y * y2 + (x * y1 + y * x1)" +apply (simp add: add_left_commute add_assoc [symmetric]) +apply (simp add: add_assoc right_distrib [symmetric]) +apply (simp add: add_commute) +done + +lemma real_mult_congruent2: + "(%p1 p2. + (%(x1,y1). (%(x2,y2). + { Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1) + respects2 realrel" +apply (rule congruent2_commuteI [OF equiv_realrel], clarify) +apply (simp add: mult_commute add_commute) +apply (auto simp add: real_mult_congruent2_lemma) +done + +lemma real_mult: + "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) = + Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})" +by (simp add: real_mult_def UN_UN_split_split_eq + UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2]) + +lemma real_mult_commute: "(z::real) * w = w * z" +by (cases z, cases w, simp add: real_mult add_ac mult_ac) + +lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)" +apply (cases z1, cases z2, cases z3) +apply (simp add: real_mult right_distrib add_ac mult_ac) +done + +lemma real_mult_1: "(1::real) * z = z" +apply (cases z) +apply (simp add: real_mult real_one_def right_distrib + mult_1_right mult_ac add_ac) +done + +lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)" +apply (cases z1, cases z2, cases w) +apply (simp add: real_add real_mult right_distrib add_ac mult_ac) +done + +text{*one and zero are distinct*} +lemma real_zero_not_eq_one: "0 \ (1::real)" +proof - + have "(1::preal) < 1 + 1" + by (simp add: preal_self_less_add_left) + thus ?thesis + by (simp add: real_zero_def real_one_def) +qed + +instance real :: comm_ring_1 +proof + fix x y z :: real + show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc) + show "x * y = y * x" by (rule real_mult_commute) + show "1 * x = x" by (rule real_mult_1) + show "(x + y) * z = x * z + y * z" by (rule real_add_mult_distrib) + show "0 \ (1::real)" by (rule real_zero_not_eq_one) +qed + +subsection {* Inverse and Division *} + +lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0" +by (simp add: real_zero_def add_commute) + +text{*Instead of using an existential quantifier and constructing the inverse +within the proof, we could define the inverse explicitly.*} + +lemma real_mult_inverse_left_ex: "x \ 0 ==> \y. y*x = (1::real)" +apply (simp add: real_zero_def real_one_def, cases x) +apply (cut_tac x = xa and y = y in linorder_less_linear) +apply (auto dest!: less_add_left_Ex simp add: real_zero_iff) +apply (rule_tac + x = "Abs_Real (realrel``{(1, inverse (D) + 1)})" + in exI) +apply (rule_tac [2] + x = "Abs_Real (realrel``{(inverse (D) + 1, 1)})" + in exI) +apply (auto simp add: real_mult preal_mult_inverse_right ring_simps) +done + +lemma real_mult_inverse_left: "x \ 0 ==> inverse(x)*x = (1::real)" +apply (simp add: real_inverse_def) +apply (drule real_mult_inverse_left_ex, safe) +apply (rule theI, assumption, rename_tac z) +apply (subgoal_tac "(z * x) * y = z * (x * y)") +apply (simp add: mult_commute) +apply (rule mult_assoc) +done + + +subsection{*The Real Numbers form a Field*} + +instance real :: field +proof + fix x y z :: real + show "x \ 0 ==> inverse x * x = 1" by (rule real_mult_inverse_left) + show "x / y = x * inverse y" by (simp add: real_divide_def) +qed + + +text{*Inverse of zero! Useful to simplify certain equations*} + +lemma INVERSE_ZERO: "inverse 0 = (0::real)" +by (simp add: real_inverse_def) + +instance real :: division_by_zero +proof + show "inverse 0 = (0::real)" by (rule INVERSE_ZERO) +qed + + +subsection{*The @{text "\"} Ordering*} + +lemma real_le_refl: "w \ (w::real)" +by (cases w, force simp add: real_le_def) + +text{*The arithmetic decision procedure is not set up for type preal. + This lemma is currently unused, but it could simplify the proofs of the + following two lemmas.*} +lemma preal_eq_le_imp_le: + assumes eq: "a+b = c+d" and le: "c \ a" + shows "b \ (d::preal)" +proof - + have "c+d \ a+d" by (simp add: prems) + hence "a+b \ a+d" by (simp add: prems) + thus "b \ d" by simp +qed + +lemma real_le_lemma: + assumes l: "u1 + v2 \ u2 + v1" + and "x1 + v1 = u1 + y1" + and "x2 + v2 = u2 + y2" + shows "x1 + y2 \ x2 + (y1::preal)" +proof - + have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: prems) + hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: add_ac) + also have "... \ (x2+y1) + (u2+v1)" by (simp add: prems) + finally show ?thesis by simp +qed + +lemma real_le: + "(Abs_Real(realrel``{(x1,y1)}) \ Abs_Real(realrel``{(x2,y2)})) = + (x1 + y2 \ x2 + y1)" +apply (simp add: real_le_def) +apply (auto intro: real_le_lemma) +done + +lemma real_le_anti_sym: "[| z \ w; w \ z |] ==> z = (w::real)" +by (cases z, cases w, simp add: real_le) + +lemma real_trans_lemma: + assumes "x + v \ u + y" + and "u + v' \ u' + v" + and "x2 + v2 = u2 + y2" + shows "x + v' \ u' + (y::preal)" +proof - + have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: add_ac) + also have "... \ (u+y) + (u+v')" by (simp add: prems) + also have "... \ (u+y) + (u'+v)" by (simp add: prems) + also have "... = (u'+y) + (u+v)" by (simp add: add_ac) + finally show ?thesis by simp +qed + +lemma real_le_trans: "[| i \ j; j \ k |] ==> i \ (k::real)" +apply (cases i, cases j, cases k) +apply (simp add: real_le) +apply (blast intro: real_trans_lemma) +done + +instance real :: order +proof + fix u v :: real + show "u < v \ u \ v \ \ v \ u" + by (auto simp add: real_less_def intro: real_le_anti_sym) +qed (assumption | rule real_le_refl real_le_trans real_le_anti_sym)+ + +(* Axiom 'linorder_linear' of class 'linorder': *) +lemma real_le_linear: "(z::real) \ w | w \ z" +apply (cases z, cases w) +apply (auto simp add: real_le real_zero_def add_ac) +done + +instance real :: linorder + by (intro_classes, rule real_le_linear) + + +lemma real_le_eq_diff: "(x \ y) = (x-y \ (0::real))" +apply (cases x, cases y) +apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus + add_ac) +apply (simp_all add: add_assoc [symmetric]) +done + +lemma real_add_left_mono: + assumes le: "x \ y" shows "z + x \ z + (y::real)" +proof - + have "z + x - (z + y) = (z + -z) + (x - y)" + by (simp add: diff_minus add_ac) + with le show ?thesis + by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus) +qed + +lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)" +by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus) + +lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))" +by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus) + +lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y" +apply (cases x, cases y) +apply (simp add: linorder_not_le [where 'a = real, symmetric] + linorder_not_le [where 'a = preal] + real_zero_def real_le real_mult) + --{*Reduce to the (simpler) @{text "\"} relation *} +apply (auto dest!: less_add_left_Ex + simp add: add_ac mult_ac + right_distrib preal_self_less_add_left) +done + +lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y" +apply (rule real_sum_gt_zero_less) +apply (drule real_less_sum_gt_zero [of x y]) +apply (drule real_mult_order, assumption) +apply (simp add: right_distrib) +done + +instantiation real :: distrib_lattice +begin + +definition + "(inf \ real \ real \ real) = min" + +definition + "(sup \ real \ real \ real) = max" + +instance + by default (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1) + +end + + +subsection{*The Reals Form an Ordered Field*} + +instance real :: ordered_field +proof + fix x y z :: real + show "x \ y ==> z + x \ z + y" by (rule real_add_left_mono) + show "x < y ==> 0 < z ==> z * x < z * y" by (rule real_mult_less_mono2) + show "\x\ = (if x < 0 then -x else x)" by (simp only: real_abs_def) + show "sgn x = (if x=0 then 0 else if 0m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)" +apply (simp add: real_of_preal_def real_zero_def, cases x) +apply (auto simp add: real_minus add_ac) +apply (cut_tac x = x and y = y in linorder_less_linear) +apply (auto dest!: less_add_left_Ex simp add: add_assoc [symmetric]) +done + +lemma real_of_preal_leD: + "real_of_preal m1 \ real_of_preal m2 ==> m1 \ m2" +by (simp add: real_of_preal_def real_le) + +lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2" +by (auto simp add: real_of_preal_leD linorder_not_le [symmetric]) + +lemma real_of_preal_lessD: + "real_of_preal m1 < real_of_preal m2 ==> m1 < m2" +by (simp add: real_of_preal_def real_le linorder_not_le [symmetric]) + +lemma real_of_preal_less_iff [simp]: + "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)" +by (blast intro: real_of_preal_lessI real_of_preal_lessD) + +lemma real_of_preal_le_iff: + "(real_of_preal m1 \ real_of_preal m2) = (m1 \ m2)" +by (simp add: linorder_not_less [symmetric]) + +lemma real_of_preal_zero_less: "0 < real_of_preal m" +apply (insert preal_self_less_add_left [of 1 m]) +apply (auto simp add: real_zero_def real_of_preal_def + real_less_def real_le_def add_ac) +apply (rule_tac x="m + 1" in exI, rule_tac x="1" in exI) +apply (simp add: add_ac) +done + +lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0" +by (simp add: real_of_preal_zero_less) + +lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m" +proof - + from real_of_preal_minus_less_zero + show ?thesis by (blast dest: order_less_trans) +qed + + +subsection{*Theorems About the Ordering*} + +lemma real_gt_zero_preal_Ex: "(0 < x) = (\y. x = real_of_preal y)" +apply (auto simp add: real_of_preal_zero_less) +apply (cut_tac x = x in real_of_preal_trichotomy) +apply (blast elim!: real_of_preal_not_minus_gt_zero [THEN notE]) +done + +lemma real_gt_preal_preal_Ex: + "real_of_preal z < x ==> \y. x = real_of_preal y" +by (blast dest!: real_of_preal_zero_less [THEN order_less_trans] + intro: real_gt_zero_preal_Ex [THEN iffD1]) + +lemma real_ge_preal_preal_Ex: + "real_of_preal z \ x ==> \y. x = real_of_preal y" +by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex) + +lemma real_less_all_preal: "y \ 0 ==> \x. y < real_of_preal x" +by (auto elim: order_le_imp_less_or_eq [THEN disjE] + intro: real_of_preal_zero_less [THEN [2] order_less_trans] + simp add: real_of_preal_zero_less) + +lemma real_less_all_real2: "~ 0 < y ==> \x. y < real_of_preal x" +by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1]) + + +subsection{*More Lemmas*} + +lemma real_mult_left_cancel: "(c::real) \ 0 ==> (c*a=c*b) = (a=b)" +by auto + +lemma real_mult_right_cancel: "(c::real) \ 0 ==> (a*c=b*c) = (a=b)" +by auto + +lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)" + by (force elim: order_less_asym + simp add: Ring_and_Field.mult_less_cancel_right) + +lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \ y*z) = (x\y)" +apply (simp add: mult_le_cancel_right) +apply (blast intro: elim: order_less_asym) +done + +lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \ z*y) = (x\y)" +by(simp add:mult_commute) + +lemma real_inverse_gt_one: "[| (0::real) < x; x < 1 |] ==> 1 < inverse x" +by (simp add: one_less_inverse_iff) (* TODO: generalize/move *) + + +subsection {* Embedding numbers into the Reals *} + +abbreviation + real_of_nat :: "nat \ real" +where + "real_of_nat \ of_nat" + +abbreviation + real_of_int :: "int \ real" +where + "real_of_int \ of_int" + +abbreviation + real_of_rat :: "rat \ real" +where + "real_of_rat \ of_rat" + +consts + (*overloaded constant for injecting other types into "real"*) + real :: "'a => real" + +defs (overloaded) + real_of_nat_def [code unfold]: "real == real_of_nat" + real_of_int_def [code unfold]: "real == real_of_int" + +lemma real_eq_of_nat: "real = of_nat" + unfolding real_of_nat_def .. + +lemma real_eq_of_int: "real = of_int" + unfolding real_of_int_def .. + +lemma real_of_int_zero [simp]: "real (0::int) = 0" +by (simp add: real_of_int_def) + +lemma real_of_one [simp]: "real (1::int) = (1::real)" +by (simp add: real_of_int_def) + +lemma real_of_int_add [simp]: "real(x + y) = real (x::int) + real y" +by (simp add: real_of_int_def) + +lemma real_of_int_minus [simp]: "real(-x) = -real (x::int)" +by (simp add: real_of_int_def) + +lemma real_of_int_diff [simp]: "real(x - y) = real (x::int) - real y" +by (simp add: real_of_int_def) + +lemma real_of_int_mult [simp]: "real(x * y) = real (x::int) * real y" +by (simp add: real_of_int_def) + +lemma real_of_int_setsum [simp]: "real ((SUM x:A. f x)::int) = (SUM x:A. real(f x))" + apply (subst real_eq_of_int)+ + apply (rule of_int_setsum) +done + +lemma real_of_int_setprod [simp]: "real ((PROD x:A. f x)::int) = + (PROD x:A. real(f x))" + apply (subst real_eq_of_int)+ + apply (rule of_int_setprod) +done + +lemma real_of_int_zero_cancel [simp, algebra, presburger]: "(real x = 0) = (x = (0::int))" +by (simp add: real_of_int_def) + +lemma real_of_int_inject [iff, algebra, presburger]: "(real (x::int) = real y) = (x = y)" +by (simp add: real_of_int_def) + +lemma real_of_int_less_iff [iff, presburger]: "(real (x::int) < real y) = (x < y)" +by (simp add: real_of_int_def) + +lemma real_of_int_le_iff [simp, presburger]: "(real (x::int) \ real y) = (x \ y)" +by (simp add: real_of_int_def) + +lemma real_of_int_gt_zero_cancel_iff [simp, presburger]: "(0 < real (n::int)) = (0 < n)" +by (simp add: real_of_int_def) + +lemma real_of_int_ge_zero_cancel_iff [simp, presburger]: "(0 <= real (n::int)) = (0 <= n)" +by (simp add: real_of_int_def) + +lemma real_of_int_lt_zero_cancel_iff [simp, presburger]: "(real (n::int) < 0) = (n < 0)" +by (simp add: real_of_int_def) + +lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)" +by (simp add: real_of_int_def) + +lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))" +by (auto simp add: abs_if) + +lemma int_less_real_le: "((n::int) < m) = (real n + 1 <= real m)" + apply (subgoal_tac "real n + 1 = real (n + 1)") + apply (simp del: real_of_int_add) + apply auto +done + +lemma int_le_real_less: "((n::int) <= m) = (real n < real m + 1)" + apply (subgoal_tac "real m + 1 = real (m + 1)") + apply (simp del: real_of_int_add) + apply simp +done + +lemma real_of_int_div_aux: "d ~= 0 ==> (real (x::int)) / (real d) = + real (x div d) + (real (x mod d)) / (real d)" +proof - + assume "d ~= 0" + have "x = (x div d) * d + x mod d" + by auto + then have "real x = real (x div d) * real d + real(x mod d)" + by (simp only: real_of_int_mult [THEN sym] real_of_int_add [THEN sym]) + then have "real x / real d = ... / real d" + by simp + then show ?thesis + by (auto simp add: add_divide_distrib ring_simps prems) +qed + +lemma real_of_int_div: "(d::int) ~= 0 ==> d dvd n ==> + real(n div d) = real n / real d" + apply (frule real_of_int_div_aux [of d n]) + apply simp + apply (simp add: zdvd_iff_zmod_eq_0) +done + +lemma real_of_int_div2: + "0 <= real (n::int) / real (x) - real (n div x)" + apply (case_tac "x = 0") + apply simp + apply (case_tac "0 < x") + apply (simp add: compare_rls) + apply (subst real_of_int_div_aux) + apply simp + apply simp + apply (subst zero_le_divide_iff) + apply auto + apply (simp add: compare_rls) + apply (subst real_of_int_div_aux) + apply simp + apply simp + apply (subst zero_le_divide_iff) + apply auto +done + +lemma real_of_int_div3: + "real (n::int) / real (x) - real (n div x) <= 1" + apply(case_tac "x = 0") + apply simp + apply (simp add: compare_rls) + apply (subst real_of_int_div_aux) + apply assumption + apply simp + apply (subst divide_le_eq) + apply clarsimp + apply (rule conjI) + apply (rule impI) + apply (rule order_less_imp_le) + apply simp + apply (rule impI) + apply (rule order_less_imp_le) + apply simp +done + +lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x" +by (insert real_of_int_div2 [of n x], simp) + + +subsection{*Embedding the Naturals into the Reals*} + +lemma real_of_nat_zero [simp]: "real (0::nat) = 0" +by (simp add: real_of_nat_def) + +lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)" +by (simp add: real_of_nat_def) + +lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n" +by (simp add: real_of_nat_def) + +(*Not for addsimps: often the LHS is used to represent a positive natural*) +lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)" +by (simp add: real_of_nat_def) + +lemma real_of_nat_less_iff [iff]: + "(real (n::nat) < real m) = (n < m)" +by (simp add: real_of_nat_def) + +lemma real_of_nat_le_iff [iff]: "(real (n::nat) \ real m) = (n \ m)" +by (simp add: real_of_nat_def) + +lemma real_of_nat_ge_zero [iff]: "0 \ real (n::nat)" +by (simp add: real_of_nat_def zero_le_imp_of_nat) + +lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)" +by (simp add: real_of_nat_def del: of_nat_Suc) + +lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n" +by (simp add: real_of_nat_def of_nat_mult) + +lemma real_of_nat_setsum [simp]: "real ((SUM x:A. f x)::nat) = + (SUM x:A. real(f x))" + apply (subst real_eq_of_nat)+ + apply (rule of_nat_setsum) +done + +lemma real_of_nat_setprod [simp]: "real ((PROD x:A. f x)::nat) = + (PROD x:A. real(f x))" + apply (subst real_eq_of_nat)+ + apply (rule of_nat_setprod) +done + +lemma real_of_card: "real (card A) = setsum (%x.1) A" + apply (subst card_eq_setsum) + apply (subst real_of_nat_setsum) + apply simp +done + +lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)" +by (simp add: real_of_nat_def) + +lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)" +by (simp add: real_of_nat_def) + +lemma real_of_nat_diff: "n \ m ==> real (m - n) = real (m::nat) - real n" +by (simp add: add: real_of_nat_def of_nat_diff) + +lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)" +by (auto simp: real_of_nat_def) + +lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \ 0) = (n = 0)" +by (simp add: add: real_of_nat_def) + +lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0" +by (simp add: add: real_of_nat_def) + +lemma real_of_nat_ge_zero_cancel_iff [simp]: "(0 \ real (n::nat))" +by (simp add: add: real_of_nat_def) + +lemma nat_less_real_le: "((n::nat) < m) = (real n + 1 <= real m)" + apply (subgoal_tac "real n + 1 = real (Suc n)") + apply simp + apply (auto simp add: real_of_nat_Suc) +done + +lemma nat_le_real_less: "((n::nat) <= m) = (real n < real m + 1)" + apply (subgoal_tac "real m + 1 = real (Suc m)") + apply (simp add: less_Suc_eq_le) + apply (simp add: real_of_nat_Suc) +done + +lemma real_of_nat_div_aux: "0 < d ==> (real (x::nat)) / (real d) = + real (x div d) + (real (x mod d)) / (real d)" +proof - + assume "0 < d" + have "x = (x div d) * d + x mod d" + by auto + then have "real x = real (x div d) * real d + real(x mod d)" + by (simp only: real_of_nat_mult [THEN sym] real_of_nat_add [THEN sym]) + then have "real x / real d = \ / real d" + by simp + then show ?thesis + by (auto simp add: add_divide_distrib ring_simps prems) +qed + +lemma real_of_nat_div: "0 < (d::nat) ==> d dvd n ==> + real(n div d) = real n / real d" + apply (frule real_of_nat_div_aux [of d n]) + apply simp + apply (subst dvd_eq_mod_eq_0 [THEN sym]) + apply assumption +done + +lemma real_of_nat_div2: + "0 <= real (n::nat) / real (x) - real (n div x)" +apply(case_tac "x = 0") + apply (simp) +apply (simp add: compare_rls) +apply (subst real_of_nat_div_aux) + apply simp +apply simp +apply (subst zero_le_divide_iff) +apply simp +done + +lemma real_of_nat_div3: + "real (n::nat) / real (x) - real (n div x) <= 1" +apply(case_tac "x = 0") +apply (simp) +apply (simp add: compare_rls) +apply (subst real_of_nat_div_aux) + apply simp +apply simp +done + +lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x" + by (insert real_of_nat_div2 [of n x], simp) + +lemma real_of_int_real_of_nat: "real (int n) = real n" +by (simp add: real_of_nat_def real_of_int_def int_eq_of_nat) + +lemma real_of_int_of_nat_eq [simp]: "real (of_nat n :: int) = real n" +by (simp add: real_of_int_def real_of_nat_def) + +lemma real_nat_eq_real [simp]: "0 <= x ==> real(nat x) = real x" + apply (subgoal_tac "real(int(nat x)) = real(nat x)") + apply force + apply (simp only: real_of_int_real_of_nat) +done + + +subsection{* Rationals *} + +lemma Rats_real_nat[simp]: "real(n::nat) \ \" +by (simp add: real_eq_of_nat) + + +lemma Rats_eq_int_div_int: + "\ = { real(i::int)/real(j::int) |i j. j \ 0}" (is "_ = ?S") +proof + show "\ \ ?S" + proof + fix x::real assume "x : \" + then obtain r where "x = of_rat r" unfolding Rats_def .. + have "of_rat r : ?S" + by (cases r)(auto simp add:of_rat_rat real_eq_of_int) + thus "x : ?S" using `x = of_rat r` by simp + qed +next + show "?S \ \" + proof(auto simp:Rats_def) + fix i j :: int assume "j \ 0" + hence "real i / real j = of_rat(Fract i j)" + by (simp add:of_rat_rat real_eq_of_int) + thus "real i / real j \ range of_rat" by blast + qed +qed + +lemma Rats_eq_int_div_nat: + "\ = { real(i::int)/real(n::nat) |i n. n \ 0}" +proof(auto simp:Rats_eq_int_div_int) + fix i j::int assume "j \ 0" + show "EX (i'::int) (n::nat). real i/real j = real i'/real n \ 00" + hence "real i/real j = real i/real(nat j) \ 00" + hence "real i/real j = real(-i)/real(nat(-j)) \ 00` + by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat) + thus ?thesis by blast + qed +next + fix i::int and n::nat assume "0 < n" + hence "real i/real n = real i/real(int n) \ int n \ 0" by simp + thus "\(i'::int) j::int. real i/real n = real i'/real j \ j \ 0" by blast +qed + +lemma Rats_abs_nat_div_natE: + assumes "x \ \" + obtains m n where "n \ 0" and "\x\ = real m / real n" and "gcd m n = 1" +proof - + from `x \ \` obtain i::int and n::nat where "n \ 0" and "x = real i / real n" + by(auto simp add: Rats_eq_int_div_nat) + hence "\x\ = real(nat(abs i)) / real n" by simp + then obtain m :: nat where x_rat: "\x\ = real m / real n" by blast + let ?gcd = "gcd m n" + from `n\0` have gcd: "?gcd \ 0" by (simp add: gcd_zero) + let ?k = "m div ?gcd" + let ?l = "n div ?gcd" + let ?gcd' = "gcd ?k ?l" + have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m" + by (rule dvd_mult_div_cancel) + have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n" + by (rule dvd_mult_div_cancel) + from `n\0` and gcd_l have "?l \ 0" by (auto iff del: neq0_conv) + moreover + have "\x\ = real ?k / real ?l" + proof - + from gcd have "real ?k / real ?l = + real (?gcd * ?k) / real (?gcd * ?l)" by simp + also from gcd_k and gcd_l have "\ = real m / real n" by simp + also from x_rat have "\ = \x\" .. + finally show ?thesis .. + qed + moreover + have "?gcd' = 1" + proof - + have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)" + by (rule gcd_mult_distrib2) + with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp + with gcd show ?thesis by simp + qed + ultimately show ?thesis .. +qed + + +subsection{*Numerals and Arithmetic*} + +instantiation real :: number_ring +begin + +definition + real_number_of_def [code del]: "number_of w = real_of_int w" + +instance + by intro_classes (simp add: real_number_of_def) + +end + +lemma [code unfold, symmetric, code post]: + "number_of k = real_of_int (number_of k)" + unfolding number_of_is_id real_number_of_def .. + + +text{*Collapse applications of @{term real} to @{term number_of}*} +lemma real_number_of [simp]: "real (number_of v :: int) = number_of v" +by (simp add: real_of_int_def of_int_number_of_eq) + +lemma real_of_nat_number_of [simp]: + "real (number_of v :: nat) = + (if neg (number_of v :: int) then 0 + else (number_of v :: real))" +by (simp add: real_of_int_real_of_nat [symmetric] int_nat_number_of) + + +use "Tools/real_arith.ML" +declaration {* K real_arith_setup *} + + +subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*} + +text{*Needed in this non-standard form by Hyperreal/Transcendental*} +lemma real_0_le_divide_iff: + "((0::real) \ x/y) = ((x \ 0 | 0 \ y) & (0 \ x | y \ 0))" +by (simp add: real_divide_def zero_le_mult_iff, auto) + +lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" +by arith + +lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)" +by auto + +lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)" +by auto + +lemma real_0_less_add_iff: "((0::real) < x+y) = (-x < y)" +by auto + +lemma real_add_le_0_iff: "(x+y \ (0::real)) = (y \ -x)" +by auto + +lemma real_0_le_add_iff: "((0::real) \ x+y) = (-x \ y)" +by auto + + +(* +FIXME: we should have this, as for type int, but many proofs would break. +It replaces x+-y by x-y. +declare real_diff_def [symmetric, simp] +*) + +subsubsection{*Density of the Reals*} + +lemma real_lbound_gt_zero: + "[| (0::real) < d1; 0 < d2 |] ==> \e. 0 < e & e < d1 & e < d2" +apply (rule_tac x = " (min d1 d2) /2" in exI) +apply (simp add: min_def) +done + + +text{*Similar results are proved in @{text Ring_and_Field}*} +lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)" + by auto + +lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y" + by auto + + +subsection{*Absolute Value Function for the Reals*} + +lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))" +by (simp add: abs_if) + +(* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *) +lemma abs_le_interval_iff: "(abs x \ r) = (-r\x & x\(r::real))" +by (force simp add: OrderedGroup.abs_le_iff) + +lemma abs_add_one_gt_zero [simp]: "(0::real) < 1 + abs(x)" +by (simp add: abs_if) + +lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)" +by (rule abs_of_nonneg [OF real_of_nat_ge_zero]) + +lemma abs_add_one_not_less_self [simp]: "~ abs(x) + (1::real) < x" +by simp + +lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \ abs(x + -l) + abs(y + -m)" +by simp + +instance real :: lordered_ring +proof + fix a::real + show "abs a = sup a (-a)" + by (auto simp add: real_abs_def sup_real_def) +qed + + +subsection {* Implementation of rational real numbers *} + +definition Ratreal :: "rat \ real" where + [simp]: "Ratreal = of_rat" + +code_datatype Ratreal + +lemma Ratreal_number_collapse [code post]: + "Ratreal 0 = 0" + "Ratreal 1 = 1" + "Ratreal (number_of k) = number_of k" +by simp_all + +lemma zero_real_code [code, code unfold]: + "0 = Ratreal 0" +by simp + +lemma one_real_code [code, code unfold]: + "1 = Ratreal 1" +by simp + +lemma number_of_real_code [code unfold]: + "number_of k = Ratreal (number_of k)" +by simp + +lemma Ratreal_number_of_quotient [code post]: + "Ratreal (number_of r) / Ratreal (number_of s) = number_of r / number_of s" +by simp + +lemma Ratreal_number_of_quotient2 [code post]: + "Ratreal (number_of r / number_of s) = number_of r / number_of s" +unfolding Ratreal_number_of_quotient [symmetric] Ratreal_def of_rat_divide .. + +instantiation real :: eq +begin + +definition "eq_class.eq (x\real) y \ x - y = 0" + +instance by default (simp add: eq_real_def) + +lemma real_eq_code [code]: "eq_class.eq (Ratreal x) (Ratreal y) \ eq_class.eq x y" + by (simp add: eq_real_def eq) + +lemma real_eq_refl [code nbe]: + "eq_class.eq (x::real) x \ True" + by (rule HOL.eq_refl) + +end + +lemma real_less_eq_code [code]: "Ratreal x \ Ratreal y \ x \ y" + by (simp add: of_rat_less_eq) + +lemma real_less_code [code]: "Ratreal x < Ratreal y \ x < y" + by (simp add: of_rat_less) + +lemma real_plus_code [code]: "Ratreal x + Ratreal y = Ratreal (x + y)" + by (simp add: of_rat_add) + +lemma real_times_code [code]: "Ratreal x * Ratreal y = Ratreal (x * y)" + by (simp add: of_rat_mult) + +lemma real_uminus_code [code]: "- Ratreal x = Ratreal (- x)" + by (simp add: of_rat_minus) + +lemma real_minus_code [code]: "Ratreal x - Ratreal y = Ratreal (x - y)" + by (simp add: of_rat_diff) + +lemma real_inverse_code [code]: "inverse (Ratreal x) = Ratreal (inverse x)" + by (simp add: of_rat_inverse) + +lemma real_divide_code [code]: "Ratreal x / Ratreal y = Ratreal (x / y)" + by (simp add: of_rat_divide) + +text {* Setup for SML code generator *} + +types_code + real ("(int */ int)") +attach (term_of) {* +fun term_of_real (p, q) = + let + val rT = HOLogic.realT + in + if q = 1 orelse p = 0 then HOLogic.mk_number rT p + else @{term "op / \ real \ real \ real"} $ + HOLogic.mk_number rT p $ HOLogic.mk_number rT q + end; +*} +attach (test) {* +fun gen_real i = + let + val p = random_range 0 i; + val q = random_range 1 (i + 1); + val g = Integer.gcd p q; + val p' = p div g; + val q' = q div g; + val r = (if one_of [true, false] then p' else ~ p', + if p' = 0 then 0 else q') + in + (r, fn () => term_of_real r) + end; +*} + +consts_code + Ratreal ("(_)") + +consts_code + "of_int :: int \ real" ("\real'_of'_int") +attach {* +fun real_of_int 0 = (0, 0) + | real_of_int i = (i, 1); +*} + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/RealPow.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/RealPow.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,281 @@ +(* Title : HOL/RealPow.thy + Author : Jacques D. Fleuriot + Copyright : 1998 University of Cambridge +*) + +header {* Natural powers theory *} + +theory RealPow +imports RealDef +uses ("Tools/float_syntax.ML") +begin + +declare abs_mult_self [simp] + +instantiation real :: recpower +begin + +primrec power_real where + realpow_0: "r ^ 0 = (1\real)" + | realpow_Suc: "r ^ Suc n = (r\real) * r ^ n" + +instance proof + fix z :: real + fix n :: nat + show "z^0 = 1" by simp + show "z^(Suc n) = z * (z^n)" by simp +qed + +end + + +lemma two_realpow_ge_one [simp]: "(1::real) \ 2 ^ n" +by simp + +lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n" +apply (induct "n") +apply (auto simp add: real_of_nat_Suc) +apply (subst mult_2) +apply (rule add_less_le_mono) +apply (auto simp add: two_realpow_ge_one) +done + +lemma realpow_Suc_le_self: "[| 0 \ r; r \ (1::real) |] ==> r ^ Suc n \ r" +by (insert power_decreasing [of 1 "Suc n" r], simp) + +lemma realpow_minus_mult [rule_format]: + "0 < n --> (x::real) ^ (n - 1) * x = x ^ n" +apply (simp split add: nat_diff_split) +done + +lemma realpow_two_mult_inverse [simp]: + "r \ 0 ==> r * inverse r ^Suc (Suc 0) = inverse (r::real)" +by (simp add: real_mult_assoc [symmetric]) + +lemma realpow_two_minus [simp]: "(-x)^Suc (Suc 0) = (x::real)^Suc (Suc 0)" +by simp + +lemma realpow_two_diff: + "(x::real)^Suc (Suc 0) - y^Suc (Suc 0) = (x - y) * (x + y)" +apply (unfold real_diff_def) +apply (simp add: ring_simps) +done + +lemma realpow_two_disj: + "((x::real)^Suc (Suc 0) = y^Suc (Suc 0)) = (x = y | x = -y)" +apply (cut_tac x = x and y = y in realpow_two_diff) +apply (auto simp del: realpow_Suc) +done + +lemma realpow_real_of_nat: "real (m::nat) ^ n = real (m ^ n)" +apply (induct "n") +apply (auto simp add: real_of_nat_one real_of_nat_mult) +done + +lemma realpow_real_of_nat_two_pos [simp] : "0 < real (Suc (Suc 0) ^ n)" +apply (induct "n") +apply (auto simp add: real_of_nat_mult zero_less_mult_iff) +done + +(* used by AFP Integration theory *) +lemma realpow_increasing: + "[|(0::real) \ x; 0 \ y; x ^ Suc n \ y ^ Suc n|] ==> x \ y" + by (rule power_le_imp_le_base) + + +subsection{*Literal Arithmetic Involving Powers, Type @{typ real}*} + +lemma real_of_int_power: "real (x::int) ^ n = real (x ^ n)" +apply (induct "n") +apply (simp_all add: nat_mult_distrib) +done +declare real_of_int_power [symmetric, simp] + +lemma power_real_number_of: + "(number_of v :: real) ^ n = real ((number_of v :: int) ^ n)" +by (simp only: real_number_of [symmetric] real_of_int_power) + +declare power_real_number_of [of _ "number_of w", standard, simp] + + +subsection {* Properties of Squares *} + +lemma sum_squares_ge_zero: + fixes x y :: "'a::ordered_ring_strict" + shows "0 \ x * x + y * y" +by (intro add_nonneg_nonneg zero_le_square) + +lemma not_sum_squares_lt_zero: + fixes x y :: "'a::ordered_ring_strict" + shows "\ x * x + y * y < 0" +by (simp add: linorder_not_less sum_squares_ge_zero) + +lemma sum_nonneg_eq_zero_iff: + fixes x y :: "'a::pordered_ab_group_add" + assumes x: "0 \ x" and y: "0 \ y" + shows "(x + y = 0) = (x = 0 \ y = 0)" +proof (auto) + from y have "x + 0 \ x + y" by (rule add_left_mono) + also assume "x + y = 0" + finally have "x \ 0" by simp + thus "x = 0" using x by (rule order_antisym) +next + from x have "0 + y \ x + y" by (rule add_right_mono) + also assume "x + y = 0" + finally have "y \ 0" by simp + thus "y = 0" using y by (rule order_antisym) +qed + +lemma sum_squares_eq_zero_iff: + fixes x y :: "'a::ordered_ring_strict" + shows "(x * x + y * y = 0) = (x = 0 \ y = 0)" +by (simp add: sum_nonneg_eq_zero_iff) + +lemma sum_squares_le_zero_iff: + fixes x y :: "'a::ordered_ring_strict" + shows "(x * x + y * y \ 0) = (x = 0 \ y = 0)" +by (simp add: order_le_less not_sum_squares_lt_zero sum_squares_eq_zero_iff) + +lemma sum_squares_gt_zero_iff: + fixes x y :: "'a::ordered_ring_strict" + shows "(0 < x * x + y * y) = (x \ 0 \ y \ 0)" +by (simp add: order_less_le sum_squares_ge_zero sum_squares_eq_zero_iff) + +lemma sum_power2_ge_zero: + fixes x y :: "'a::{ordered_idom,recpower}" + shows "0 \ x\ + y\" +unfolding power2_eq_square by (rule sum_squares_ge_zero) + +lemma not_sum_power2_lt_zero: + fixes x y :: "'a::{ordered_idom,recpower}" + shows "\ x\ + y\ < 0" +unfolding power2_eq_square by (rule not_sum_squares_lt_zero) + +lemma sum_power2_eq_zero_iff: + fixes x y :: "'a::{ordered_idom,recpower}" + shows "(x\ + y\ = 0) = (x = 0 \ y = 0)" +unfolding power2_eq_square by (rule sum_squares_eq_zero_iff) + +lemma sum_power2_le_zero_iff: + fixes x y :: "'a::{ordered_idom,recpower}" + shows "(x\ + y\ \ 0) = (x = 0 \ y = 0)" +unfolding power2_eq_square by (rule sum_squares_le_zero_iff) + +lemma sum_power2_gt_zero_iff: + fixes x y :: "'a::{ordered_idom,recpower}" + shows "(0 < x\ + y\) = (x \ 0 \ y \ 0)" +unfolding power2_eq_square by (rule sum_squares_gt_zero_iff) + + +subsection{* Squares of Reals *} + +lemma real_two_squares_add_zero_iff [simp]: + "(x * x + y * y = 0) = ((x::real) = 0 \ y = 0)" +by (rule sum_squares_eq_zero_iff) + +lemma real_sum_squares_cancel: "x * x + y * y = 0 ==> x = (0::real)" +by simp + +lemma real_sum_squares_cancel2: "x * x + y * y = 0 ==> y = (0::real)" +by simp + +lemma real_mult_self_sum_ge_zero: "(0::real) \ x*x + y*y" +by (rule sum_squares_ge_zero) + +lemma real_sum_squares_cancel_a: "x * x = -(y * y) ==> x = (0::real) & y=0" +by (simp add: real_add_eq_0_iff [symmetric]) + +lemma real_squared_diff_one_factored: "x*x - (1::real) = (x + 1)*(x - 1)" +by (simp add: left_distrib right_diff_distrib) + +lemma real_mult_is_one [simp]: "(x*x = (1::real)) = (x = 1 | x = - 1)" +apply auto +apply (drule right_minus_eq [THEN iffD2]) +apply (auto simp add: real_squared_diff_one_factored) +done + +lemma real_sum_squares_not_zero: "x ~= 0 ==> x * x + y * y ~= (0::real)" +by simp + +lemma real_sum_squares_not_zero2: "y ~= 0 ==> x * x + y * y ~= (0::real)" +by simp + +lemma realpow_two_sum_zero_iff [simp]: + "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)" +by (rule sum_power2_eq_zero_iff) + +lemma realpow_two_le_add_order [simp]: "(0::real) \ u ^ 2 + v ^ 2" +by (rule sum_power2_ge_zero) + +lemma realpow_two_le_add_order2 [simp]: "(0::real) \ u ^ 2 + v ^ 2 + w ^ 2" +by (intro add_nonneg_nonneg zero_le_power2) + +lemma real_sum_square_gt_zero: "x ~= 0 ==> (0::real) < x * x + y * y" +by (simp add: sum_squares_gt_zero_iff) + +lemma real_sum_square_gt_zero2: "y ~= 0 ==> (0::real) < x * x + y * y" +by (simp add: sum_squares_gt_zero_iff) + +lemma real_minus_mult_self_le [simp]: "-(u * u) \ (x * (x::real))" +by (rule_tac j = 0 in real_le_trans, auto) + +lemma realpow_square_minus_le [simp]: "-(u ^ 2) \ (x::real) ^ 2" +by (auto simp add: power2_eq_square) + +(* The following theorem is by Benjamin Porter *) +lemma real_sq_order: + fixes x::real + assumes xgt0: "0 \ x" and ygt0: "0 \ y" and sq: "x^2 \ y^2" + shows "x \ y" +proof - + from sq have "x ^ Suc (Suc 0) \ y ^ Suc (Suc 0)" + by (simp only: numeral_2_eq_2) + thus "x \ y" using ygt0 + by (rule power_le_imp_le_base) +qed + + +subsection {*Various Other Theorems*} + +lemma real_le_add_half_cancel: "(x + y/2 \ (y::real)) = (x \ y /2)" +by auto + +lemma real_minus_half_eq [simp]: "(x::real) - x/2 = x/2" +by auto + +lemma real_mult_inverse_cancel: + "[|(0::real) < x; 0 < x1; x1 * y < x * u |] + ==> inverse x * y < inverse x1 * u" +apply (rule_tac c=x in mult_less_imp_less_left) +apply (auto simp add: real_mult_assoc [symmetric]) +apply (simp (no_asm) add: mult_ac) +apply (rule_tac c=x1 in mult_less_imp_less_right) +apply (auto simp add: mult_ac) +done + +lemma real_mult_inverse_cancel2: + "[|(0::real) < x;0 < x1; x1 * y < x * u |] ==> y * inverse x < u * inverse x1" +apply (auto dest: real_mult_inverse_cancel simp add: mult_ac) +done + +lemma inverse_real_of_nat_gt_zero [simp]: "0 < inverse (real (Suc n))" +by simp + +lemma inverse_real_of_nat_ge_zero [simp]: "0 \ inverse (real (Suc n))" +by simp + +lemma realpow_num_eq_if: "(m::real) ^ n = (if n=0 then 1 else m * m ^ (n - 1))" +by (case_tac "n", auto) + +subsection{* Float syntax *} + +syntax "_Float" :: "float_const \ 'a" ("_") + +use "Tools/float_syntax.ML" +setup FloatSyntax.setup + +text{* Test: *} +lemma "123.456 = -111.111 + 200 + 30 + 4 + 5/10 + 6/100 + (7/1000::real)" +by simp + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Series.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Series.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,654 @@ +(* Title : Series.thy + Author : Jacques D. Fleuriot + Copyright : 1998 University of Cambridge + +Converted to Isar and polished by lcp +Converted to setsum and polished yet more by TNN +Additional contributions by Jeremy Avigad +*) + +header{*Finite Summation and Infinite Series*} + +theory Series +imports "~~/src/HOL/Hyperreal/SEQ" +begin + +definition + sums :: "(nat \ 'a::real_normed_vector) \ 'a \ bool" + (infixr "sums" 80) where + "f sums s = (%n. setsum f {0.. s" + +definition + summable :: "(nat \ 'a::real_normed_vector) \ bool" where + "summable f = (\s. f sums s)" + +definition + suminf :: "(nat \ 'a::real_normed_vector) \ 'a" where + "suminf f = (THE s. f sums s)" + +syntax + "_suminf" :: "idt \ 'a \ 'a" ("\_. _" [0, 10] 10) +translations + "\i. b" == "CONST suminf (%i. b)" + + +lemma sumr_diff_mult_const: + "setsum f {0.. f(p) \ K) + \ setsum f {0.. real n * K" +using setsum_bounded[where A = "{0..i=0..<2*n. (-1) ^ Suc i) = (0::real)" +by (induct "n", auto) + +(* FIXME this is an awful lemma! *) +lemma sumr_one_lb_realpow_zero [simp]: + "(\n=Suc 0..m=0..m=0.. 'a::ab_group_add" + shows "(\m=0..f. (\m=0..n f. setsum f {0::nat..m=0.. + (\n=Suc 0 ..< Suc n. if even(n) then 0 else + ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n = + (\n=0.. summable f" +by (simp add: sums_def summable_def, blast) + +lemma summable_sums: "summable f ==> f sums (suminf f)" +apply (simp add: summable_def suminf_def sums_def) +apply (blast intro: theI LIMSEQ_unique) +done + +lemma summable_sumr_LIMSEQ_suminf: + "summable f ==> (%n. setsum f {0.. (suminf f)" +by (rule summable_sums [unfolded sums_def]) + +(*------------------- + sum is unique + ------------------*) +lemma sums_unique: "f sums s ==> (s = suminf f)" +apply (frule sums_summable [THEN summable_sums]) +apply (auto intro!: LIMSEQ_unique simp add: sums_def) +done + +lemma sums_split_initial_segment: "f sums s ==> + (%n. f(n + k)) sums (s - (SUM i = 0..< k. f i))" + apply (unfold sums_def); + apply (simp add: sumr_offset); + apply (rule LIMSEQ_diff_const) + apply (rule LIMSEQ_ignore_initial_segment) + apply assumption +done + +lemma summable_ignore_initial_segment: "summable f ==> + summable (%n. f(n + k))" + apply (unfold summable_def) + apply (auto intro: sums_split_initial_segment) +done + +lemma suminf_minus_initial_segment: "summable f ==> + suminf f = s ==> suminf (%n. f(n + k)) = s - (SUM i = 0..< k. f i)" + apply (frule summable_ignore_initial_segment) + apply (rule sums_unique [THEN sym]) + apply (frule summable_sums) + apply (rule sums_split_initial_segment) + apply auto +done + +lemma suminf_split_initial_segment: "summable f ==> + suminf f = (SUM i = 0..< k. f i) + suminf (%n. f(n + k))" +by (auto simp add: suminf_minus_initial_segment) + +lemma series_zero: + "(\m. n \ m --> f(m) = 0) ==> f sums (setsum f {0..n. 0) sums 0" +unfolding sums_def by (simp add: LIMSEQ_const) + +lemma summable_zero: "summable (\n. 0)" +by (rule sums_zero [THEN sums_summable]) + +lemma suminf_zero: "suminf (\n. 0) = 0" +by (rule sums_zero [THEN sums_unique, symmetric]) + +lemma (in bounded_linear) sums: + "(\n. X n) sums a \ (\n. f (X n)) sums (f a)" +unfolding sums_def by (drule LIMSEQ, simp only: setsum) + +lemma (in bounded_linear) summable: + "summable (\n. X n) \ summable (\n. f (X n))" +unfolding summable_def by (auto intro: sums) + +lemma (in bounded_linear) suminf: + "summable (\n. X n) \ f (\n. X n) = (\n. f (X n))" +by (intro sums_unique sums summable_sums) + +lemma sums_mult: + fixes c :: "'a::real_normed_algebra" + shows "f sums a \ (\n. c * f n) sums (c * a)" +by (rule mult_right.sums) + +lemma summable_mult: + fixes c :: "'a::real_normed_algebra" + shows "summable f \ summable (%n. c * f n)" +by (rule mult_right.summable) + +lemma suminf_mult: + fixes c :: "'a::real_normed_algebra" + shows "summable f \ suminf (\n. c * f n) = c * suminf f"; +by (rule mult_right.suminf [symmetric]) + +lemma sums_mult2: + fixes c :: "'a::real_normed_algebra" + shows "f sums a \ (\n. f n * c) sums (a * c)" +by (rule mult_left.sums) + +lemma summable_mult2: + fixes c :: "'a::real_normed_algebra" + shows "summable f \ summable (\n. f n * c)" +by (rule mult_left.summable) + +lemma suminf_mult2: + fixes c :: "'a::real_normed_algebra" + shows "summable f \ suminf f * c = (\n. f n * c)" +by (rule mult_left.suminf) + +lemma sums_divide: + fixes c :: "'a::real_normed_field" + shows "f sums a \ (\n. f n / c) sums (a / c)" +by (rule divide.sums) + +lemma summable_divide: + fixes c :: "'a::real_normed_field" + shows "summable f \ summable (\n. f n / c)" +by (rule divide.summable) + +lemma suminf_divide: + fixes c :: "'a::real_normed_field" + shows "summable f \ suminf (\n. f n / c) = suminf f / c" +by (rule divide.suminf [symmetric]) + +lemma sums_add: "\X sums a; Y sums b\ \ (\n. X n + Y n) sums (a + b)" +unfolding sums_def by (simp add: setsum_addf LIMSEQ_add) + +lemma summable_add: "\summable X; summable Y\ \ summable (\n. X n + Y n)" +unfolding summable_def by (auto intro: sums_add) + +lemma suminf_add: + "\summable X; summable Y\ \ suminf X + suminf Y = (\n. X n + Y n)" +by (intro sums_unique sums_add summable_sums) + +lemma sums_diff: "\X sums a; Y sums b\ \ (\n. X n - Y n) sums (a - b)" +unfolding sums_def by (simp add: setsum_subtractf LIMSEQ_diff) + +lemma summable_diff: "\summable X; summable Y\ \ summable (\n. X n - Y n)" +unfolding summable_def by (auto intro: sums_diff) + +lemma suminf_diff: + "\summable X; summable Y\ \ suminf X - suminf Y = (\n. X n - Y n)" +by (intro sums_unique sums_diff summable_sums) + +lemma sums_minus: "X sums a ==> (\n. - X n) sums (- a)" +unfolding sums_def by (simp add: setsum_negf LIMSEQ_minus) + +lemma summable_minus: "summable X \ summable (\n. - X n)" +unfolding summable_def by (auto intro: sums_minus) + +lemma suminf_minus: "summable X \ (\n. - X n) = - (\n. X n)" +by (intro sums_unique [symmetric] sums_minus summable_sums) + +lemma sums_group: + "[|summable f; 0 < k |] ==> (%n. setsum f {n*k.. real" + shows "\summable f; \m\n. 0 \ f m\ \ setsum f {0.. suminf f" +apply (drule summable_sums) +apply (simp add: sums_def) +apply (cut_tac k = "setsum f {0.. real" + shows "\summable f; \m\n. 0 < f m\ \ setsum f {0.. real" + shows "\summable f; \n. 0 < f n\ \ 0 < suminf f" +by (drule_tac n="0" in series_pos_less, simp_all) + +lemma suminf_ge_zero: + fixes f :: "nat \ real" + shows "\summable f; \n. 0 \ f n\ \ 0 \ suminf f" +by (drule_tac n="0" in series_pos_le, simp_all) + +lemma sumr_pos_lt_pair: + fixes f :: "nat \ real" + shows "\summable f; + \d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc(Suc 0) * d) + 1))\ + \ setsum f {0.. (\n. x ^ n) sums (1 / (1 - x))" +proof - + assume less_1: "norm x < 1" + hence neq_1: "x \ 1" by auto + hence neq_0: "x - 1 \ 0" by simp + from less_1 have lim_0: "(\n. x ^ n) ----> 0" + by (rule LIMSEQ_power_zero) + hence "(\n. x ^ n / (x - 1) - 1 / (x - 1)) ----> 0 / (x - 1) - 1 / (x - 1)" + using neq_0 by (intro LIMSEQ_divide LIMSEQ_diff LIMSEQ_const) + hence "(\n. (x ^ n - 1) / (x - 1)) ----> 1 / (1 - x)" + by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib) + thus "(\n. x ^ n) sums (1 / (1 - x))" + by (simp add: sums_def geometric_sum neq_1) +qed + +lemma summable_geometric: + fixes x :: "'a::{real_normed_field,recpower}" + shows "norm x < 1 \ summable (\n. x ^ n)" +by (rule geometric_sums [THEN sums_summable]) + +text{*Cauchy-type criterion for convergence of series (c.f. Harrison)*} + +lemma summable_convergent_sumr_iff: + "summable f = convergent (%n. setsum f {0.. f ----> 0" +apply (drule summable_convergent_sumr_iff [THEN iffD1]) +apply (drule convergent_Cauchy) +apply (simp only: Cauchy_def LIMSEQ_def, safe) +apply (drule_tac x="r" in spec, safe) +apply (rule_tac x="M" in exI, safe) +apply (drule_tac x="Suc n" in spec, simp) +apply (drule_tac x="n" in spec, simp) +done + +lemma summable_Cauchy: + "summable (f::nat \ 'a::banach) = + (\e > 0. \N. \m \ N. \n. norm (setsum f {m.. 'b::real_normed_vector" + shows "norm (setsum f A) \ (\i\A. norm (f i))" +apply (case_tac "finite A") +apply (erule finite_induct) +apply simp +apply simp +apply (erule order_trans [OF norm_triangle_ineq add_left_mono]) +apply simp +done + +lemma summable_comparison_test: + fixes f :: "nat \ 'a::banach" + shows "\\N. \n\N. norm (f n) \ g n; summable g\ \ summable f" +apply (simp add: summable_Cauchy, safe) +apply (drule_tac x="e" in spec, safe) +apply (rule_tac x = "N + Na" in exI, safe) +apply (rotate_tac 2) +apply (drule_tac x = m in spec) +apply (auto, rotate_tac 2, drule_tac x = n in spec) +apply (rule_tac y = "\k=m.. 'a::banach" + shows "\\N. \n\N. norm (f n) \ g n; summable g\ + \ summable (\n. norm (f n))" +apply (rule summable_comparison_test) +apply (auto) +done + +lemma summable_rabs_comparison_test: + fixes f :: "nat \ real" + shows "\\N. \n\N. \f n\ \ g n; summable g\ \ summable (\n. \f n\)" +apply (rule summable_comparison_test) +apply (auto) +done + +text{*Summability of geometric series for real algebras*} + +lemma complete_algebra_summable_geometric: + fixes x :: "'a::{real_normed_algebra_1,banach,recpower}" + shows "norm x < 1 \ summable (\n. x ^ n)" +proof (rule summable_comparison_test) + show "\N. \n\N. norm (x ^ n) \ norm x ^ n" + by (simp add: norm_power_ineq) + show "norm x < 1 \ summable (\n. norm x ^ n)" + by (simp add: summable_geometric) +qed + +text{*Limit comparison property for series (c.f. jrh)*} + +lemma summable_le: + fixes f g :: "nat \ real" + shows "\\n. f n \ g n; summable f; summable g\ \ suminf f \ suminf g" +apply (drule summable_sums)+ +apply (simp only: sums_def, erule (1) LIMSEQ_le) +apply (rule exI) +apply (auto intro!: setsum_mono) +done + +lemma summable_le2: + fixes f g :: "nat \ real" + shows "\\n. \f n\ \ g n; summable g\ \ summable f \ suminf f \ suminf g" +apply (subgoal_tac "summable f") +apply (auto intro!: summable_le) +apply (simp add: abs_le_iff) +apply (rule_tac g="g" in summable_comparison_test, simp_all) +done + +(* specialisation for the common 0 case *) +lemma suminf_0_le: + fixes f::"nat\real" + assumes gt0: "\n. 0 \ f n" and sm: "summable f" + shows "0 \ suminf f" +proof - + let ?g = "(\n. (0::real))" + from gt0 have "\n. ?g n \ f n" by simp + moreover have "summable ?g" by (rule summable_zero) + moreover from sm have "summable f" . + ultimately have "suminf ?g \ suminf f" by (rule summable_le) + then show "0 \ suminf f" by (simp add: suminf_zero) +qed + + +text{*Absolute convergence imples normal convergence*} +lemma summable_norm_cancel: + fixes f :: "nat \ 'a::banach" + shows "summable (\n. norm (f n)) \ summable f" +apply (simp only: summable_Cauchy, safe) +apply (drule_tac x="e" in spec, safe) +apply (rule_tac x="N" in exI, safe) +apply (drule_tac x="m" in spec, safe) +apply (rule order_le_less_trans [OF norm_setsum]) +apply (rule order_le_less_trans [OF abs_ge_self]) +apply simp +done + +lemma summable_rabs_cancel: + fixes f :: "nat \ real" + shows "summable (\n. \f n\) \ summable f" +by (rule summable_norm_cancel, simp) + +text{*Absolute convergence of series*} +lemma summable_norm: + fixes f :: "nat \ 'a::banach" + shows "summable (\n. norm (f n)) \ norm (suminf f) \ (\n. norm (f n))" +by (auto intro: LIMSEQ_le LIMSEQ_norm summable_norm_cancel + summable_sumr_LIMSEQ_suminf norm_setsum) + +lemma summable_rabs: + fixes f :: "nat \ real" + shows "summable (\n. \f n\) \ \suminf f\ \ (\n. \f n\)" +by (fold real_norm_def, rule summable_norm) + +subsection{* The Ratio Test*} + +lemma norm_ratiotest_lemma: + fixes x y :: "'a::real_normed_vector" + shows "\c \ 0; norm x \ c * norm y\ \ x = 0" +apply (subgoal_tac "norm x \ 0", simp) +apply (erule order_trans) +apply (simp add: mult_le_0_iff) +done + +lemma rabs_ratiotest_lemma: "[| c \ 0; abs x \ c * abs y |] ==> x = (0::real)" +by (erule norm_ratiotest_lemma, simp) + +lemma le_Suc_ex: "(k::nat) \ l ==> (\n. l = k + n)" +apply (drule le_imp_less_or_eq) +apply (auto dest: less_imp_Suc_add) +done + +lemma le_Suc_ex_iff: "((k::nat) \ l) = (\n. l = k + n)" +by (auto simp add: le_Suc_ex) + +(*All this trouble just to get 0 'a::banach" + shows "\\n\N. norm (f (Suc n)) \ c * norm (f n)\ \ 0 < c \ summable f" +apply (simp (no_asm) add: linorder_not_le [symmetric]) +apply (simp add: summable_Cauchy) +apply (safe, subgoal_tac "\n. N < n --> f (n) = 0") + prefer 2 + apply clarify + apply(erule_tac x = "n - 1" in allE) + apply (simp add:diff_Suc split:nat.splits) + apply (blast intro: norm_ratiotest_lemma) +apply (rule_tac x = "Suc N" in exI, clarify) +apply(simp cong:setsum_ivl_cong) +done + +lemma ratio_test: + fixes f :: "nat \ 'a::banach" + shows "\c < 1; \n\N. norm (f (Suc n)) \ c * norm (f n)\ \ summable f" +apply (frule ratio_test_lemma2, auto) +apply (rule_tac g = "%n. (norm (f N) / (c ^ N))*c ^ n" + in summable_comparison_test) +apply (rule_tac x = N in exI, safe) +apply (drule le_Suc_ex_iff [THEN iffD1]) +apply (auto simp add: power_add field_power_not_zero) +apply (induct_tac "na", auto) +apply (rule_tac y = "c * norm (f (N + n))" in order_trans) +apply (auto intro: mult_right_mono simp add: summable_def) +apply (simp add: mult_ac) +apply (rule_tac x = "norm (f N) * (1/ (1 - c)) / (c ^ N)" in exI) +apply (rule sums_divide) +apply (rule sums_mult) +apply (auto intro!: geometric_sums) +done + +subsection {* Cauchy Product Formula *} + +(* Proof based on Analysis WebNotes: Chapter 07, Class 41 +http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm *) + +lemma setsum_triangle_reindex: + fixes n :: nat + shows "(\(i,j)\{(i,j). i+j < n}. f i j) = (\k=0..i=0..k. f i (k - i))" +proof - + have "(\(i, j)\{(i, j). i + j < n}. f i j) = + (\(k, i)\(SIGMA k:{0..(k,i). (i, k - i)) (SIGMA k:{0..(k,i). (i, k - i)) ` (SIGMA k:{0..a. (\(k, i). f i (k - i)) a = split f ((\(k, i). (i, k - i)) a)" + by clarify + qed + thus ?thesis by (simp add: setsum_Sigma) +qed + +lemma Cauchy_product_sums: + fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" + assumes a: "summable (\k. norm (a k))" + assumes b: "summable (\k. norm (b k))" + shows "(\k. \i=0..k. a i * b (k - i)) sums ((\k. a k) * (\k. b k))" +proof - + let ?S1 = "\n::nat. {0.. {0..m n. m \ n \ ?S1 m \ ?S1 n" by auto + have S2_le_S1: "\n. ?S2 n \ ?S1 n" by auto + have S1_le_S2: "\n. ?S1 (n div 2) \ ?S2 n" by auto + have finite_S1: "\n. finite (?S1 n)" by simp + with S2_le_S1 have finite_S2: "\n. finite (?S2 n)" by (rule finite_subset) + + let ?g = "\(i,j). a i * b j" + let ?f = "\(i,j). norm (a i) * norm (b j)" + have f_nonneg: "\x. 0 \ ?f x" + by (auto simp add: mult_nonneg_nonneg) + hence norm_setsum_f: "\A. norm (setsum ?f A) = setsum ?f A" + unfolding real_norm_def + by (simp only: abs_of_nonneg setsum_nonneg [rule_format]) + + have "(\n. (\k=0..k=0.. (\k. a k) * (\k. b k)" + by (intro LIMSEQ_mult summable_sumr_LIMSEQ_suminf + summable_norm_cancel [OF a] summable_norm_cancel [OF b]) + hence 1: "(\n. setsum ?g (?S1 n)) ----> (\k. a k) * (\k. b k)" + by (simp only: setsum_product setsum_Sigma [rule_format] + finite_atLeastLessThan) + + have "(\n. (\k=0..k=0.. (\k. norm (a k)) * (\k. norm (b k))" + using a b by (intro LIMSEQ_mult summable_sumr_LIMSEQ_suminf) + hence "(\n. setsum ?f (?S1 n)) ----> (\k. norm (a k)) * (\k. norm (b k))" + by (simp only: setsum_product setsum_Sigma [rule_format] + finite_atLeastLessThan) + hence "convergent (\n. setsum ?f (?S1 n))" + by (rule convergentI) + hence Cauchy: "Cauchy (\n. setsum ?f (?S1 n))" + by (rule convergent_Cauchy) + have "Zseq (\n. setsum ?f (?S1 n - ?S2 n))" + proof (rule ZseqI, simp only: norm_setsum_f) + fix r :: real + assume r: "0 < r" + from CauchyD [OF Cauchy r] obtain N + where "\m\N. \n\N. norm (setsum ?f (?S1 m) - setsum ?f (?S1 n)) < r" .. + hence "\m n. \N \ n; n \ m\ \ norm (setsum ?f (?S1 m - ?S1 n)) < r" + by (simp only: setsum_diff finite_S1 S1_mono) + hence N: "\m n. \N \ n; n \ m\ \ setsum ?f (?S1 m - ?S1 n) < r" + by (simp only: norm_setsum_f) + show "\N. \n\N. setsum ?f (?S1 n - ?S2 n) < r" + proof (intro exI allI impI) + fix n assume "2 * N \ n" + hence n: "N \ n div 2" by simp + have "setsum ?f (?S1 n - ?S2 n) \ setsum ?f (?S1 n - ?S1 (n div 2))" + by (intro setsum_mono2 finite_Diff finite_S1 f_nonneg + Diff_mono subset_refl S1_le_S2) + also have "\ < r" + using n div_le_dividend by (rule N) + finally show "setsum ?f (?S1 n - ?S2 n) < r" . + qed + qed + hence "Zseq (\n. setsum ?g (?S1 n - ?S2 n))" + apply (rule Zseq_le [rule_format]) + apply (simp only: norm_setsum_f) + apply (rule order_trans [OF norm_setsum setsum_mono]) + apply (auto simp add: norm_mult_ineq) + done + hence 2: "(\n. setsum ?g (?S1 n) - setsum ?g (?S2 n)) ----> 0" + by (simp only: LIMSEQ_Zseq_iff setsum_diff finite_S1 S2_le_S1 diff_0_right) + + with 1 have "(\n. setsum ?g (?S2 n)) ----> (\k. a k) * (\k. b k)" + by (rule LIMSEQ_diff_approach_zero2) + thus ?thesis by (simp only: sums_def setsum_triangle_reindex) +qed + +lemma Cauchy_product: + fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" + assumes a: "summable (\k. norm (a k))" + assumes b: "summable (\k. norm (b k))" + shows "(\k. a k) * (\k. b k) = (\k. \i=0..k. a i * b (k - i))" +using a b +by (rule Cauchy_product_sums [THEN sums_unique]) + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Taylor.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Taylor.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,133 @@ +(* Title: HOL/Taylor.thy + Author: Lukas Bulwahn, Bernhard Haeupler, Technische Universitaet Muenchen +*) + +header {* Taylor series *} + +theory Taylor +imports MacLaurin +begin + +text {* +We use MacLaurin and the translation of the expansion point @{text c} to @{text 0} +to prove Taylor's theorem. +*} + +lemma taylor_up: + assumes INIT: "n>0" "diff 0 = f" + and DERIV: "(\ m t. m < n & a \ t & t \ b \ DERIV (diff m) t :> (diff (Suc m) t))" + and INTERV: "a \ c" "c < b" + shows "\ t. c < t & t < b & + f b = setsum (%m. (diff m c / real (fact m)) * (b - c)^m) {0..0" "((\m x. diff m (x + c)) 0) = (\x. f (x + c))" by auto + moreover + have "ALL m t. m < n & 0 <= t & t <= b - c --> DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" + proof (intro strip) + fix m t + assume "m < n & 0 <= t & t <= b - c" + with DERIV and INTERV have "DERIV (diff m) (t + c) :> diff (Suc m) (t + c)" by auto + moreover + from DERIV_ident and DERIV_const have "DERIV (%x. x + c) t :> 1+0" by (rule DERIV_add) + ultimately have "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1+0)" + by (rule DERIV_chain2) + thus "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" by simp + qed + ultimately + have EX:"EX t>0. t < b - c & + f (b - c + c) = (SUM m = 0..m = 0.. f b = (\m = 0..0" "diff 0 = f" + and DERIV: "(\ m t. m < n & a \ t & t \ b \ DERIV (diff m) t :> (diff (Suc m) t))" + and INTERV: "a < c" "c \ b" + shows "\ t. a < t & t < c & + f a = setsum (% m. (diff m c / real (fact m)) * (a - c)^m) {0..0" "((\m x. diff m (x + c)) 0) = (\x. f (x + c))" by auto + moreover + have "ALL m t. m < n & a-c <= t & t <= 0 --> DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" + proof (rule allI impI)+ + fix m t + assume "m < n & a-c <= t & t <= 0" + with DERIV and INTERV have "DERIV (diff m) (t + c) :> diff (Suc m) (t + c)" by auto + moreover + from DERIV_ident and DERIV_const have "DERIV (%x. x + c) t :> 1+0" by (rule DERIV_add) + ultimately have "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1+0)" by (rule DERIV_chain2) + thus "DERIV (%x. diff m (x + c)) t :> diff (Suc m) (t + c)" by simp + qed + ultimately + have EX: "EX t>a - c. t < 0 & + f (a - c + c) = (SUM m = 0.. f a = (\m = 0..0" "diff 0 = f" + and DERIV: "(\ m t. m < n & a \ t & t \ b \ DERIV (diff m) t :> (diff (Suc m) t))" + and INTERV: "a \ c " "c \ b" "a \ x" "x \ b" "x \ c" + shows "\ t. (if xm t. m < n \ x \ t \ t \ b \ DERIV (diff m) t :> diff (Suc m) t" + by fastsimp + moreover note True + moreover from INTERV have "c \ b" by simp + ultimately have EX: "\t>x. t < c \ f x = + (\m = 0..m t. m < n \ a \ t \ t \ x \ DERIV (diff m) t :> diff (Suc m) t" + by fastsimp + moreover from INTERV have "a \ c" by arith + moreover from False and INTERV have "c < x" by arith + ultimately have EX: "\t>c. t < x \ f x = + (\m = 0.. (MetaSimplifier.simpset -> tactic) + -> MetaSimplifier.simpset -> term * term -> thm + val simp_all_tac: thm list -> MetaSimplifier.simpset -> tactic + + val mk_sum: term list -> term + val mk_norm_sum: term list -> term + val dest_sum: term -> term list + + val nat_cancel_sums_add: simproc list + val nat_cancel_sums: simproc list + val setup: Context.generic -> Context.generic +end; + +structure ArithData: ARITH_DATA = +struct + +(** generic proof tools **) + +(* prove conversions *) + +fun prove_conv expand_tac norm_tac ss tu = (* FIXME avoid standard *) + mk_meta_eq (standard (Goal.prove (Simplifier.the_context ss) [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq tu)) + (K (EVERY [expand_tac, norm_tac ss])))); + +(* rewriting *) + +fun simp_all_tac rules = + let val ss0 = HOL_ss addsimps rules + in fn ss => ALLGOALS (simp_tac (Simplifier.inherit_context ss ss0)) end; + + +(** abstract syntax of structure nat: 0, Suc, + **) + +local + +val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; +val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} HOLogic.natT; + +in + +fun mk_sum [] = HOLogic.zero + | mk_sum [t] = t + | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); + +(*normal form of sums: Suc (... (Suc (a + (b + ...))))*) +fun mk_norm_sum ts = + let val (ones, sums) = List.partition (equal HOLogic.Suc_zero) ts in + funpow (length ones) HOLogic.mk_Suc (mk_sum sums) + end; + + +fun dest_sum tm = + if HOLogic.is_zero tm then [] + else + (case try HOLogic.dest_Suc tm of + SOME t => HOLogic.Suc_zero :: dest_sum t + | NONE => + (case try dest_plus tm of + SOME (t, u) => dest_sum t @ dest_sum u + | NONE => [tm])); + +end; + + +(** cancel common summands **) + +structure Sum = +struct + val mk_sum = mk_norm_sum; + val dest_sum = dest_sum; + val prove_conv = prove_conv; + val norm_tac1 = simp_all_tac [@{thm "add_Suc"}, @{thm "add_Suc_right"}, + @{thm "add_0"}, @{thm "add_0_right"}]; + val norm_tac2 = simp_all_tac @{thms add_ac}; + fun norm_tac ss = norm_tac1 ss THEN norm_tac2 ss; +end; + +fun gen_uncancel_tac rule ct = + rtac (instantiate' [] [NONE, SOME ct] (rule RS @{thm subst_equals})) 1; + + +(* nat eq *) + +structure EqCancelSums = CancelSumsFun +(struct + open Sum; + val mk_bal = HOLogic.mk_eq; + val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT; + val uncancel_tac = gen_uncancel_tac @{thm "nat_add_left_cancel"}; +end); + + +(* nat less *) + +structure LessCancelSums = CancelSumsFun +(struct + open Sum; + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less}; + val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT; + val uncancel_tac = gen_uncancel_tac @{thm "nat_add_left_cancel_less"}; +end); + + +(* nat le *) + +structure LeCancelSums = CancelSumsFun +(struct + open Sum; + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq}; + val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT; + val uncancel_tac = gen_uncancel_tac @{thm "nat_add_left_cancel_le"}; +end); + + +(* nat diff *) + +structure DiffCancelSums = CancelSumsFun +(struct + open Sum; + val mk_bal = HOLogic.mk_binop @{const_name HOL.minus}; + val dest_bal = HOLogic.dest_bin @{const_name HOL.minus} HOLogic.natT; + val uncancel_tac = gen_uncancel_tac @{thm "diff_cancel"}; +end); + + +(* prepare nat_cancel simprocs *) + +val nat_cancel_sums_add = + [Simplifier.simproc (the_context ()) "nateq_cancel_sums" + ["(l::nat) + m = n", "(l::nat) = m + n", "Suc m = n", "m = Suc n"] + (K EqCancelSums.proc), + Simplifier.simproc (the_context ()) "natless_cancel_sums" + ["(l::nat) + m < n", "(l::nat) < m + n", "Suc m < n", "m < Suc n"] + (K LessCancelSums.proc), + Simplifier.simproc (the_context ()) "natle_cancel_sums" + ["(l::nat) + m <= n", "(l::nat) <= m + n", "Suc m <= n", "m <= Suc n"] + (K LeCancelSums.proc)]; + +val nat_cancel_sums = nat_cancel_sums_add @ + [Simplifier.simproc (the_context ()) "natdiff_cancel_sums" + ["((l::nat) + m) - n", "(l::nat) - (m + n)", "Suc m - n", "m - Suc n"] + (K DiffCancelSums.proc)]; + +val setup = + Simplifier.map_ss (fn ss => ss addsimprocs nat_cancel_sums); + +end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Tools/float_arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/float_arith.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,222 @@ +(* Title: HOL/Real/float_arith.ML + ID: $Id$ + Author: Steven Obua +*) + +signature FLOAT_ARITH = +sig + exception Destruct_floatstr of string + val destruct_floatstr: (char -> bool) -> (char -> bool) -> string -> bool * string * string * bool * string + + exception Floating_point of string + val approx_dec_by_bin: int -> Float.float -> Float.float * Float.float + val approx_decstr_by_bin: int -> string -> Float.float * Float.float + + val mk_float: Float.float -> term + val dest_float: term -> Float.float + + val approx_float: int -> (Float.float * Float.float -> Float.float * Float.float) + -> string -> term * term +end; + +structure FloatArith : FLOAT_ARITH = +struct + +exception Destruct_floatstr of string; + +fun destruct_floatstr isDigit isExp number = + let + val numlist = filter (not o Char.isSpace) (String.explode number) + + fun countsigns ((#"+")::cs) = countsigns cs + | countsigns ((#"-")::cs) = + let + val (positive, rest) = countsigns cs + in + (not positive, rest) + end + | countsigns cs = (true, cs) + + fun readdigits [] = ([], []) + | readdigits (q as c::cs) = + if (isDigit c) then + let + val (digits, rest) = readdigits cs + in + (c::digits, rest) + end + else + ([], q) + + fun readfromexp_helper cs = + let + val (positive, rest) = countsigns cs + val (digits, rest') = readdigits rest + in + case rest' of + [] => (positive, digits) + | _ => raise (Destruct_floatstr number) + end + + fun readfromexp [] = (true, []) + | readfromexp (c::cs) = + if isExp c then + readfromexp_helper cs + else + raise (Destruct_floatstr number) + + fun readfromdot [] = ([], readfromexp []) + | readfromdot ((#".")::cs) = + let + val (digits, rest) = readdigits cs + val exp = readfromexp rest + in + (digits, exp) + end + | readfromdot cs = readfromdot ((#".")::cs) + + val (positive, numlist) = countsigns numlist + val (digits1, numlist) = readdigits numlist + val (digits2, exp) = readfromdot numlist + in + (positive, String.implode digits1, String.implode digits2, fst exp, String.implode (snd exp)) + end + +exception Floating_point of string; + +val ln2_10 = Math.ln 10.0 / Math.ln 2.0; +fun exp5 x = Integer.pow x 5; +fun exp10 x = Integer.pow x 10; +fun exp2 x = Integer.pow x 2; + +fun find_most_significant q r = + let + fun int2real i = + case (Real.fromString o string_of_int) i of + SOME r => r + | NONE => raise (Floating_point "int2real") + fun subtract (q, r) (q', r') = + if r <= r' then + (q - q' * exp10 (r' - r), r) + else + (q * exp10 (r - r') - q', r') + fun bin2dec d = + if 0 <= d then + (exp2 d, 0) + else + (exp5 (~ d), d) + + val L = Real.floor (int2real (IntInf.log2 q) + int2real r * ln2_10) + val L1 = L + 1 + + val (q1, r1) = subtract (q, r) (bin2dec L1) + in + if 0 <= q1 then + let + val (q2, r2) = subtract (q, r) (bin2dec (L1 + 1)) + in + if 0 <= q2 then + raise (Floating_point "find_most_significant") + else + (L1, (q1, r1)) + end + else + let + val (q0, r0) = subtract (q, r) (bin2dec L) + in + if 0 <= q0 then + (L, (q0, r0)) + else + raise (Floating_point "find_most_significant") + end + end + +fun approx_dec_by_bin n (q,r) = + let + fun addseq acc d' [] = acc + | addseq acc d' (d::ds) = addseq (acc + exp2 (d - d')) d' ds + + fun seq2bin [] = (0, 0) + | seq2bin (d::ds) = (addseq 0 d ds + 1, d) + + fun approx d_seq d0 precision (q,r) = + if q = 0 then + let val x = seq2bin d_seq in + (x, x) + end + else + let + val (d, (q', r')) = find_most_significant q r + in + if precision < d0 - d then + let + val d' = d0 - precision + val x1 = seq2bin (d_seq) + val x2 = (fst x1 * exp2 (snd x1 - d') + 1, d') (* = seq2bin (d'::d_seq) *) + in + (x1, x2) + end + else + approx (d::d_seq) d0 precision (q', r') + end + + fun approx_start precision (q, r) = + if q = 0 then + ((0, 0), (0, 0)) + else + let + val (d, (q', r')) = find_most_significant q r + in + if precision <= 0 then + let + val x1 = seq2bin [d] + in + if q' = 0 then + (x1, x1) + else + (x1, seq2bin [d + 1]) + end + else + approx [d] d precision (q', r') + end + in + if 0 <= q then + approx_start n (q,r) + else + let + val ((a1,b1), (a2, b2)) = approx_start n (~ q, r) + in + ((~ a2, b2), (~ a1, b1)) + end + end + +fun approx_decstr_by_bin n decstr = + let + fun str2int s = the_default 0 (Int.fromString s) + fun signint p x = if p then x else ~ x + + val (p, d1, d2, ep, e) = destruct_floatstr Char.isDigit (fn e => e = #"e" orelse e = #"E") decstr + val s = size d2 + + val q = signint p (str2int d1 * exp10 s + str2int d2) + val r = signint ep (str2int e) - s + in + approx_dec_by_bin n (q,r) + end + +fun mk_float (a, b) = @{term "float"} $ + HOLogic.mk_prod (pairself (HOLogic.mk_number HOLogic.intT) (a, b)); + +fun dest_float (Const ("Float.float", _) $ (Const ("Pair", _) $ a $ b)) = + pairself (snd o HOLogic.dest_number) (a, b) + | dest_float t = ((snd o HOLogic.dest_number) t, 0); + +fun approx_float prec f value = + let + val interval = approx_decstr_by_bin prec value + val (flower, fupper) = f interval + in + (mk_float flower, mk_float fupper) + end; + +end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Tools/float_syntax.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/float_syntax.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,47 @@ +(* ID: $Id$ + Authors: Tobias Nipkow, TU Muenchen + +Concrete syntax for floats +*) + +signature FLOAT_SYNTAX = +sig + val setup: theory -> theory +end; + +structure FloatSyntax: FLOAT_SYNTAX = +struct + +(* parse translation *) + +local + +fun mk_number i = + let + fun mk 0 = Syntax.const @{const_name Int.Pls} + | mk ~1 = Syntax.const @{const_name Int.Min} + | mk i = let val (q, r) = Integer.div_mod i 2 + in HOLogic.mk_bit r $ (mk q) end; + in Syntax.const @{const_name Int.number_of} $ mk i end; + +fun mk_frac str = + let + val {mant=i, exp = n} = Syntax.read_float str; + val exp = Syntax.const @{const_name "Power.power"}; + val ten = mk_number 10; + val exp10 = if n=1 then ten else exp $ ten $ (mk_number n); + in (Syntax.const @{const_name "divide"}) $ (mk_number i) $ exp10 end +in + +fun float_tr (*"_Float"*) [t as Const (str, _)] = mk_frac str + | float_tr (*"_Float"*) ts = raise TERM ("float_tr", ts); + +end; + + +(* theory setup *) + +val setup = + Sign.add_trfuns ([], [("_Float", float_tr)], [], []); + +end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Tools/hologic.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/hologic.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,568 @@ +(* Title: HOL/hologic.ML + ID: $Id$ + Author: Lawrence C Paulson and Markus Wenzel + +Abstract syntax operations for HOL. +*) + +signature HOLOGIC = +sig + val typeS: sort + val typeT: typ + val boolN: string + val boolT: typ + val true_const: term + val false_const: term + val mk_setT: typ -> typ + val dest_setT: typ -> typ + val Trueprop: term + val mk_Trueprop: term -> term + val dest_Trueprop: term -> term + val conj_intr: thm -> thm -> thm + val conj_elim: thm -> thm * thm + val conj_elims: thm -> thm list + val conj: term + val disj: term + val imp: term + val Not: term + val mk_conj: term * term -> term + val mk_disj: term * term -> term + val mk_imp: term * term -> term + val mk_not: term -> term + val dest_conj: term -> term list + val dest_disj: term -> term list + val disjuncts: term -> term list + val dest_imp: term -> term * term + val dest_not: term -> term + val eq_const: typ -> term + val mk_eq: term * term -> term + val dest_eq: term -> term * term + val all_const: typ -> term + val mk_all: string * typ * term -> term + val list_all: (string * typ) list * term -> term + val exists_const: typ -> term + val mk_exists: string * typ * term -> term + val choice_const: typ -> term + val Collect_const: typ -> term + val mk_Collect: string * typ * term -> term + val class_eq: string + val mk_mem: term * term -> term + val dest_mem: term -> term * term + val mk_UNIV: typ -> term + val mk_binop: string -> term * term -> term + val mk_binrel: string -> term * term -> term + val dest_bin: string -> typ -> term -> term * term + val unitT: typ + val is_unitT: typ -> bool + val unit: term + val is_unit: term -> bool + val mk_prodT: typ * typ -> typ + val dest_prodT: typ -> typ * typ + val pair_const: typ -> typ -> term + val mk_prod: term * term -> term + val dest_prod: term -> term * term + val mk_fst: term -> term + val mk_snd: term -> term + val split_const: typ * typ * typ -> term + val mk_split: term -> term + val prodT_factors: typ -> typ list + val mk_tuple: typ -> term list -> term + val dest_tuple: term -> term list + val ap_split: typ -> typ -> term -> term + val prod_factors: term -> int list list + val dest_tuple': int list list -> term -> term list + val prodT_factors': int list list -> typ -> typ list + val ap_split': int list list -> typ -> typ -> term -> term + val mk_tuple': int list list -> typ -> term list -> term + val mk_tupleT: int list list -> typ list -> typ + val strip_split: term -> term * typ list * int list list + val natT: typ + val zero: term + val is_zero: term -> bool + val mk_Suc: term -> term + val dest_Suc: term -> term + val Suc_zero: term + val mk_nat: int -> term + val dest_nat: term -> int + val class_size: string + val size_const: typ -> term + val indexT: typ + val intT: typ + val pls_const: term + val min_const: term + val bit0_const: term + val bit1_const: term + val mk_bit: int -> term + val dest_bit: term -> int + val mk_numeral: int -> term + val dest_numeral: term -> int + val number_of_const: typ -> term + val add_numerals: term -> (term * typ) list -> (term * typ) list + val mk_number: typ -> int -> term + val dest_number: term -> typ * int + val realT: typ + val nibbleT: typ + val mk_nibble: int -> term + val dest_nibble: term -> int + val charT: typ + val mk_char: int -> term + val dest_char: term -> int + val listT: typ -> typ + val nil_const: typ -> term + val cons_const: typ -> term + val mk_list: typ -> term list -> term + val dest_list: term -> term list + val stringT: typ + val mk_string: string -> term + val dest_string: term -> string +end; + +structure HOLogic: HOLOGIC = +struct + +(* HOL syntax *) + +val typeS: sort = ["HOL.type"]; +val typeT = TypeInfer.anyT typeS; + + +(* bool and set *) + +val boolN = "bool"; +val boolT = Type (boolN, []); + +val true_const = Const ("True", boolT); +val false_const = Const ("False", boolT); + +fun mk_setT T = T --> boolT; + +fun dest_setT (Type ("fun", [T, Type ("bool", [])])) = T + | dest_setT T = raise TYPE ("dest_setT: set type expected", [T], []); + + +(* logic *) + +val Trueprop = Const ("Trueprop", boolT --> propT); + +fun mk_Trueprop P = Trueprop $ P; + +fun dest_Trueprop (Const ("Trueprop", _) $ P) = P + | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]); + +fun conj_intr thP thQ = + let + val (P, Q) = pairself (ObjectLogic.dest_judgment o Thm.cprop_of) (thP, thQ) + handle CTERM (msg, _) => raise THM (msg, 0, [thP, thQ]); + val inst = Thm.instantiate ([], [(@{cpat "?P::bool"}, P), (@{cpat "?Q::bool"}, Q)]); + in Drule.implies_elim_list (inst @{thm conjI}) [thP, thQ] end; + +fun conj_elim thPQ = + let + val (P, Q) = Thm.dest_binop (ObjectLogic.dest_judgment (Thm.cprop_of thPQ)) + handle CTERM (msg, _) => raise THM (msg, 0, [thPQ]); + val inst = Thm.instantiate ([], [(@{cpat "?P::bool"}, P), (@{cpat "?Q::bool"}, Q)]); + val thP = Thm.implies_elim (inst @{thm conjunct1}) thPQ; + val thQ = Thm.implies_elim (inst @{thm conjunct2}) thPQ; + in (thP, thQ) end; + +fun conj_elims th = + let val (th1, th2) = conj_elim th + in conj_elims th1 @ conj_elims th2 end handle THM _ => [th]; + +val conj = @{term "op &"} +and disj = @{term "op |"} +and imp = @{term "op -->"} +and Not = @{term "Not"}; + +fun mk_conj (t1, t2) = conj $ t1 $ t2 +and mk_disj (t1, t2) = disj $ t1 $ t2 +and mk_imp (t1, t2) = imp $ t1 $ t2 +and mk_not t = Not $ t; + +fun dest_conj (Const ("op &", _) $ t $ t') = t :: dest_conj t' + | dest_conj t = [t]; + +fun dest_disj (Const ("op |", _) $ t $ t') = t :: dest_disj t' + | dest_disj t = [t]; + +(*Like dest_disj, but flattens disjunctions however nested*) +fun disjuncts_aux (Const ("op |", _) $ t $ t') disjs = disjuncts_aux t (disjuncts_aux t' disjs) + | disjuncts_aux t disjs = t::disjs; + +fun disjuncts t = disjuncts_aux t []; + +fun dest_imp (Const("op -->",_) $ A $ B) = (A, B) + | dest_imp t = raise TERM ("dest_imp", [t]); + +fun dest_not (Const ("Not", _) $ t) = t + | dest_not t = raise TERM ("dest_not", [t]); + +fun eq_const T = Const ("op =", [T, T] ---> boolT); +fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u; + +fun dest_eq (Const ("op =", _) $ lhs $ rhs) = (lhs, rhs) + | dest_eq t = raise TERM ("dest_eq", [t]) + +fun all_const T = Const ("All", [T --> boolT] ---> boolT); +fun mk_all (x, T, P) = all_const T $ absfree (x, T, P); +fun list_all (xs, t) = fold_rev (fn (x, T) => fn P => all_const T $ Abs (x, T, P)) xs t; + +fun exists_const T = Const ("Ex", [T --> boolT] ---> boolT); +fun mk_exists (x, T, P) = exists_const T $ absfree (x, T, P); + +fun choice_const T = Const("Hilbert_Choice.Eps", (T --> boolT) --> T); + +fun Collect_const T = Const ("Collect", [T --> boolT] ---> mk_setT T); +fun mk_Collect (a, T, t) = Collect_const T $ absfree (a, T, t); + +val class_eq = "HOL.eq"; + +fun mk_mem (x, A) = + let val setT = fastype_of A in + Const ("op :", [dest_setT setT, setT] ---> boolT) $ x $ A + end; + +fun dest_mem (Const ("op :", _) $ x $ A) = (x, A) + | dest_mem t = raise TERM ("dest_mem", [t]); + +fun mk_UNIV T = Const ("UNIV", mk_setT T); + + +(* binary operations and relations *) + +fun mk_binop c (t, u) = + let val T = fastype_of t in + Const (c, [T, T] ---> T) $ t $ u + end; + +fun mk_binrel c (t, u) = + let val T = fastype_of t in + Const (c, [T, T] ---> boolT) $ t $ u + end; + +(*destruct the application of a binary operator. The dummyT case is a crude + way of handling polymorphic operators.*) +fun dest_bin c T (tm as Const (c', Type ("fun", [T', _])) $ t $ u) = + if c = c' andalso (T=T' orelse T=dummyT) then (t, u) + else raise TERM ("dest_bin " ^ c, [tm]) + | dest_bin c _ tm = raise TERM ("dest_bin " ^ c, [tm]); + + +(* unit *) + +val unitT = Type ("Product_Type.unit", []); + +fun is_unitT (Type ("Product_Type.unit", [])) = true + | is_unitT _ = false; + +val unit = Const ("Product_Type.Unity", unitT); + +fun is_unit (Const ("Product_Type.Unity", _)) = true + | is_unit _ = false; + + +(* prod *) + +fun mk_prodT (T1, T2) = Type ("*", [T1, T2]); + +fun dest_prodT (Type ("*", [T1, T2])) = (T1, T2) + | dest_prodT T = raise TYPE ("dest_prodT", [T], []); + +fun pair_const T1 T2 = Const ("Pair", [T1, T2] ---> mk_prodT (T1, T2)); + +fun mk_prod (t1, t2) = + let val T1 = fastype_of t1 and T2 = fastype_of t2 in + pair_const T1 T2 $ t1 $ t2 + end; + +fun dest_prod (Const ("Pair", _) $ t1 $ t2) = (t1, t2) + | dest_prod t = raise TERM ("dest_prod", [t]); + +fun mk_fst p = + let val pT = fastype_of p in + Const ("fst", pT --> fst (dest_prodT pT)) $ p + end; + +fun mk_snd p = + let val pT = fastype_of p in + Const ("snd", pT --> snd (dest_prodT pT)) $ p + end; + +fun split_const (A, B, C) = + Const ("split", (A --> B --> C) --> mk_prodT (A, B) --> C); + +fun mk_split t = + (case Term.fastype_of t of + T as (Type ("fun", [A, Type ("fun", [B, C])])) => + Const ("split", T --> mk_prodT (A, B) --> C) $ t + | _ => raise TERM ("mk_split: bad body type", [t])); + +(*Maps the type T1 * ... * Tn to [T1, ..., Tn], however nested*) +fun prodT_factors (Type ("*", [T1, T2])) = prodT_factors T1 @ prodT_factors T2 + | prodT_factors T = [T]; + +(*Makes a nested tuple from a list, following the product type structure*) +fun mk_tuple (Type ("*", [T1, T2])) tms = + mk_prod (mk_tuple T1 tms, + mk_tuple T2 (Library.drop (length (prodT_factors T1), tms))) + | mk_tuple T (t::_) = t; + +fun dest_tuple (Const ("Pair", _) $ t $ u) = dest_tuple t @ dest_tuple u + | dest_tuple t = [t]; + +(*In ap_split S T u, term u expects separate arguments for the factors of S, + with result type T. The call creates a new term expecting one argument + of type S.*) +fun ap_split T T3 u = + let + fun ap (T :: Ts) = + (case T of + Type ("*", [T1, T2]) => + split_const (T1, T2, Ts ---> T3) $ ap (T1 :: T2 :: Ts) + | _ => Abs ("x", T, ap Ts)) + | ap [] = + let val k = length (prodT_factors T) + in list_comb (incr_boundvars k u, map Bound (k - 1 downto 0)) end + in ap [T] end; + + +(* operations on tuples with specific arities *) +(* + an "arity" of a tuple is a list of lists of integers + ("factors"), denoting paths to subterms that are pairs +*) + +fun prod_err s = raise TERM (s ^ ": inconsistent use of products", []); + +fun prod_factors t = + let + fun factors p (Const ("Pair", _) $ t $ u) = + p :: factors (1::p) t @ factors (2::p) u + | factors p _ = [] + in factors [] t end; + +fun dest_tuple' ps = + let + fun dest p t = if p mem ps then (case t of + Const ("Pair", _) $ t $ u => + dest (1::p) t @ dest (2::p) u + | _ => prod_err "dest_tuple'") else [t] + in dest [] end; + +fun prodT_factors' ps = + let + fun factors p T = if p mem ps then (case T of + Type ("*", [T1, T2]) => + factors (1::p) T1 @ factors (2::p) T2 + | _ => prod_err "prodT_factors'") else [T] + in factors [] end; + +(*In ap_split' ps S T u, term u expects separate arguments for the factors of S, + with result type T. The call creates a new term expecting one argument + of type S.*) +fun ap_split' ps T T3 u = + let + fun ap ((p, T) :: pTs) = + if p mem ps then (case T of + Type ("*", [T1, T2]) => + split_const (T1, T2, map snd pTs ---> T3) $ + ap ((1::p, T1) :: (2::p, T2) :: pTs) + | _ => prod_err "ap_split'") + else Abs ("x", T, ap pTs) + | ap [] = + let val k = length ps + in list_comb (incr_boundvars (k + 1) u, map Bound (k downto 0)) end + in ap [([], T)] end; + +fun mk_tuple' ps = + let + fun mk p T ts = + if p mem ps then (case T of + Type ("*", [T1, T2]) => + let + val (t, ts') = mk (1::p) T1 ts; + val (u, ts'') = mk (2::p) T2 ts' + in (pair_const T1 T2 $ t $ u, ts'') end + | _ => prod_err "mk_tuple'") + else (hd ts, tl ts) + in fst oo mk [] end; + +fun mk_tupleT ps = + let + fun mk p Ts = + if p mem ps then + let + val (T, Ts') = mk (1::p) Ts; + val (U, Ts'') = mk (2::p) Ts' + in (mk_prodT (T, U), Ts'') end + else (hd Ts, tl Ts) + in fst o mk [] end; + +fun strip_split t = + let + fun strip [] qs Ts t = (t, Ts, qs) + | strip (p :: ps) qs Ts (Const ("split", _) $ t) = + strip ((1 :: p) :: (2 :: p) :: ps) (p :: qs) Ts t + | strip (p :: ps) qs Ts (Abs (s, T, t)) = strip ps qs (T :: Ts) t + | strip (p :: ps) qs Ts t = strip ps qs + (hd (binder_types (fastype_of1 (Ts, t))) :: Ts) + (incr_boundvars 1 t $ Bound 0) + in strip [[]] [] [] t end; + + +(* nat *) + +val natT = Type ("nat", []); + +val zero = Const ("HOL.zero_class.zero", natT); + +fun is_zero (Const ("HOL.zero_class.zero", _)) = true + | is_zero _ = false; + +fun mk_Suc t = Const ("Suc", natT --> natT) $ t; + +fun dest_Suc (Const ("Suc", _) $ t) = t + | dest_Suc t = raise TERM ("dest_Suc", [t]); + +val Suc_zero = mk_Suc zero; + +fun mk_nat n = + let + fun mk 0 = zero + | mk n = mk_Suc (mk (n - 1)); + in if n < 0 then raise TERM ("mk_nat: negative number", []) else mk n end; + +fun dest_nat (Const ("HOL.zero_class.zero", _)) = 0 + | dest_nat (Const ("Suc", _) $ t) = dest_nat t + 1 + | dest_nat t = raise TERM ("dest_nat", [t]); + +val class_size = "Nat.size"; + +fun size_const T = Const ("Nat.size_class.size", T --> natT); + + +(* index *) + +val indexT = Type ("Code_Index.index", []); + + +(* binary numerals and int -- non-unique representation due to leading zeros/ones! *) + +val intT = Type ("Int.int", []); + +val pls_const = Const ("Int.Pls", intT) +and min_const = Const ("Int.Min", intT) +and bit0_const = Const ("Int.Bit0", intT --> intT) +and bit1_const = Const ("Int.Bit1", intT --> intT); + +fun mk_bit 0 = bit0_const + | mk_bit 1 = bit1_const + | mk_bit _ = raise TERM ("mk_bit", []); + +fun dest_bit (Const ("Int.Bit0", _)) = 0 + | dest_bit (Const ("Int.Bit1", _)) = 1 + | dest_bit t = raise TERM ("dest_bit", [t]); + +fun mk_numeral 0 = pls_const + | mk_numeral ~1 = min_const + | mk_numeral i = + let val (q, r) = Integer.div_mod i 2; + in mk_bit r $ mk_numeral q end; + +fun dest_numeral (Const ("Int.Pls", _)) = 0 + | dest_numeral (Const ("Int.Min", _)) = ~1 + | dest_numeral (Const ("Int.Bit0", _) $ bs) = 2 * dest_numeral bs + | dest_numeral (Const ("Int.Bit1", _) $ bs) = 2 * dest_numeral bs + 1 + | dest_numeral t = raise TERM ("dest_numeral", [t]); + +fun number_of_const T = Const ("Int.number_class.number_of", intT --> T); + +fun add_numerals (Const ("Int.number_class.number_of", Type (_, [_, T])) $ t) = cons (t, T) + | add_numerals (t $ u) = add_numerals t #> add_numerals u + | add_numerals (Abs (_, _, t)) = add_numerals t + | add_numerals _ = I; + +fun mk_number T 0 = Const ("HOL.zero_class.zero", T) + | mk_number T 1 = Const ("HOL.one_class.one", T) + | mk_number T i = number_of_const T $ mk_numeral i; + +fun dest_number (Const ("HOL.zero_class.zero", T)) = (T, 0) + | dest_number (Const ("HOL.one_class.one", T)) = (T, 1) + | dest_number (Const ("Int.number_class.number_of", Type ("fun", [_, T])) $ t) = + (T, dest_numeral t) + | dest_number t = raise TERM ("dest_number", [t]); + + +(* real *) + +val realT = Type ("RealDef.real", []); + + +(* nibble *) + +val nibbleT = Type ("List.nibble", []); + +fun mk_nibble n = + let val s = + if 0 <= n andalso n <= 9 then chr (n + ord "0") + else if 10 <= n andalso n <= 15 then chr (n + ord "A" - 10) + else raise TERM ("mk_nibble", []) + in Const ("List.nibble.Nibble" ^ s, nibbleT) end; + +fun dest_nibble t = + let fun err () = raise TERM ("dest_nibble", [t]) in + (case try (unprefix "List.nibble.Nibble" o fst o Term.dest_Const) t of + NONE => err () + | SOME c => + if size c <> 1 then err () + else if "0" <= c andalso c <= "9" then ord c - ord "0" + else if "A" <= c andalso c <= "F" then ord c - ord "A" + 10 + else err ()) + end; + + +(* char *) + +val charT = Type ("List.char", []); + +fun mk_char n = + if 0 <= n andalso n <= 255 then + Const ("List.char.Char", nibbleT --> nibbleT --> charT) $ + mk_nibble (n div 16) $ mk_nibble (n mod 16) + else raise TERM ("mk_char", []); + +fun dest_char (Const ("List.char.Char", _) $ t $ u) = + dest_nibble t * 16 + dest_nibble u + | dest_char t = raise TERM ("dest_char", [t]); + + +(* list *) + +fun listT T = Type ("List.list", [T]); + +fun nil_const T = Const ("List.list.Nil", listT T); + +fun cons_const T = + let val lT = listT T + in Const ("List.list.Cons", T --> lT --> lT) end; + +fun mk_list T ts = + let + val lT = listT T; + val Nil = Const ("List.list.Nil", lT); + fun Cons t u = Const ("List.list.Cons", T --> lT --> lT) $ t $ u; + in fold_rev Cons ts Nil end; + +fun dest_list (Const ("List.list.Nil", _)) = [] + | dest_list (Const ("List.list.Cons", _) $ t $ u) = t :: dest_list u + | dest_list t = raise TERM ("dest_list", [t]); + + +(* string *) + +val stringT = Type ("List.string", []); + +val mk_string = mk_list charT o map (mk_char o ord) o explode; +val dest_string = implode o map (chr o dest_char) o dest_list; + +end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Tools/int_arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/int_arith.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,588 @@ +(* Title: HOL/int_arith1.ML + ID: $Id$ + Authors: Larry Paulson and Tobias Nipkow + +Simprocs and decision procedure for linear arithmetic. +*) + +structure Int_Numeral_Base_Simprocs = + struct + fun prove_conv tacs ctxt (_: thm list) (t, u) = + if t aconv u then NONE + else + let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u)) + in SOME (Goal.prove ctxt [] [] eq (K (EVERY tacs))) end + + fun prove_conv_nohyps tacs sg = prove_conv tacs sg []; + + fun prep_simproc (name, pats, proc) = + Simplifier.simproc (the_context()) name pats proc; + + fun is_numeral (Const(@{const_name Int.number_of}, _) $ w) = true + | is_numeral _ = false + + fun simplify_meta_eq f_number_of_eq f_eq = + mk_meta_eq ([f_eq, f_number_of_eq] MRS trans) + + (*reorientation simprules using ==, for the following simproc*) + val meta_zero_reorient = @{thm zero_reorient} RS eq_reflection + val meta_one_reorient = @{thm one_reorient} RS eq_reflection + val meta_number_of_reorient = @{thm number_of_reorient} RS eq_reflection + + (*reorientation simplification procedure: reorients (polymorphic) + 0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a Int.*) + fun reorient_proc sg _ (_ $ t $ u) = + case u of + Const(@{const_name HOL.zero}, _) => NONE + | Const(@{const_name HOL.one}, _) => NONE + | Const(@{const_name Int.number_of}, _) $ _ => NONE + | _ => SOME (case t of + Const(@{const_name HOL.zero}, _) => meta_zero_reorient + | Const(@{const_name HOL.one}, _) => meta_one_reorient + | Const(@{const_name Int.number_of}, _) $ _ => meta_number_of_reorient) + + val reorient_simproc = + prep_simproc ("reorient_simproc", ["0=x", "1=x", "number_of w = x"], reorient_proc) + + end; + + +Addsimprocs [Int_Numeral_Base_Simprocs.reorient_simproc]; + + +structure Int_Numeral_Simprocs = +struct + +(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in Int_Numeral_Base_Simprocs + isn't complicated by the abstract 0 and 1.*) +val numeral_syms = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym]; + +(** New term ordering so that AC-rewriting brings numerals to the front **) + +(*Order integers by absolute value and then by sign. The standard integer + ordering is not well-founded.*) +fun num_ord (i,j) = + (case int_ord (abs i, abs j) of + EQUAL => int_ord (Int.sign i, Int.sign j) + | ord => ord); + +(*This resembles Term.term_ord, but it puts binary numerals before other + non-atomic terms.*) +local open Term +in +fun numterm_ord (Abs (_, T, t), Abs(_, U, u)) = + (case numterm_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord) + | numterm_ord + (Const(@{const_name Int.number_of}, _) $ v, Const(@{const_name Int.number_of}, _) $ w) = + num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w) + | numterm_ord (Const(@{const_name Int.number_of}, _) $ _, _) = LESS + | numterm_ord (_, Const(@{const_name Int.number_of}, _) $ _) = GREATER + | numterm_ord (t, u) = + (case int_ord (size_of_term t, size_of_term u) of + EQUAL => + let val (f, ts) = strip_comb t and (g, us) = strip_comb u in + (case hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord) + end + | ord => ord) +and numterms_ord (ts, us) = list_ord numterm_ord (ts, us) +end; + +fun numtermless tu = (numterm_ord tu = LESS); + +(*Defined in this file, but perhaps needed only for Int_Numeral_Base_Simprocs of type nat.*) +val num_ss = HOL_ss settermless numtermless; + + +(** Utilities **) + +fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n; + +fun find_first_numeral past (t::terms) = + ((snd (HOLogic.dest_number t), rev past @ terms) + handle TERM _ => find_first_numeral (t::past) terms) + | find_first_numeral past [] = raise TERM("find_first_numeral", []); + +val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; + +fun mk_minus t = + let val T = Term.fastype_of t + in Const (@{const_name HOL.uminus}, T --> T) $ t end; + +(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) +fun mk_sum T [] = mk_number T 0 + | mk_sum T [t,u] = mk_plus (t, u) + | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); + +(*this version ALWAYS includes a trailing zero*) +fun long_mk_sum T [] = mk_number T 0 + | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); + +val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} Term.dummyT; + +(*decompose additions AND subtractions as a sum*) +fun dest_summing (pos, Const (@{const_name HOL.plus}, _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (pos, u, ts)) + | dest_summing (pos, Const (@{const_name HOL.minus}, _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (not pos, u, ts)) + | dest_summing (pos, t, ts) = + if pos then t::ts else mk_minus t :: ts; + +fun dest_sum t = dest_summing (true, t, []); + +val mk_diff = HOLogic.mk_binop @{const_name HOL.minus}; +val dest_diff = HOLogic.dest_bin @{const_name HOL.minus} Term.dummyT; + +val mk_times = HOLogic.mk_binop @{const_name HOL.times}; + +fun one_of T = Const(@{const_name HOL.one},T); + +(* build product with trailing 1 rather than Numeral 1 in order to avoid the + unnecessary restriction to type class number_ring + which is not required for cancellation of common factors in divisions. +*) +fun mk_prod T = + let val one = one_of T + fun mk [] = one + | mk [t] = t + | mk (t :: ts) = if t = one then mk ts else mk_times (t, mk ts) + in mk end; + +(*This version ALWAYS includes a trailing one*) +fun long_mk_prod T [] = one_of T + | long_mk_prod T (t :: ts) = mk_times (t, mk_prod T ts); + +val dest_times = HOLogic.dest_bin @{const_name HOL.times} Term.dummyT; + +fun dest_prod t = + let val (t,u) = dest_times t + in dest_prod t @ dest_prod u end + handle TERM _ => [t]; + +(*DON'T do the obvious simplifications; that would create special cases*) +fun mk_coeff (k, t) = mk_times (mk_number (Term.fastype_of t) k, t); + +(*Express t as a product of (possibly) a numeral with other sorted terms*) +fun dest_coeff sign (Const (@{const_name HOL.uminus}, _) $ t) = dest_coeff (~sign) t + | dest_coeff sign t = + let val ts = sort Term.term_ord (dest_prod t) + val (n, ts') = find_first_numeral [] ts + handle TERM _ => (1, ts) + in (sign*n, mk_prod (Term.fastype_of t) ts') end; + +(*Find first coefficient-term THAT MATCHES u*) +fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) + | find_first_coeff past u (t::terms) = + let val (n,u') = dest_coeff 1 t + in if u aconv u' then (n, rev past @ terms) + else find_first_coeff (t::past) u terms + end + handle TERM _ => find_first_coeff (t::past) u terms; + +(*Fractions as pairs of ints. Can't use Rat.rat because the representation + needs to preserve negative values in the denominator.*) +fun mk_frac (p, q) = if q = 0 then raise Div else (p, q); + +(*Don't reduce fractions; sums must be proved by rule add_frac_eq. + Fractions are reduced later by the cancel_numeral_factor simproc.*) +fun add_frac ((p1, q1), (p2, q2)) = (p1 * q2 + p2 * q1, q1 * q2); + +val mk_divide = HOLogic.mk_binop @{const_name HOL.divide}; + +(*Build term (p / q) * t*) +fun mk_fcoeff ((p, q), t) = + let val T = Term.fastype_of t + in mk_times (mk_divide (mk_number T p, mk_number T q), t) end; + +(*Express t as a product of a fraction with other sorted terms*) +fun dest_fcoeff sign (Const (@{const_name HOL.uminus}, _) $ t) = dest_fcoeff (~sign) t + | dest_fcoeff sign (Const (@{const_name HOL.divide}, _) $ t $ u) = + let val (p, t') = dest_coeff sign t + val (q, u') = dest_coeff 1 u + in (mk_frac (p, q), mk_divide (t', u')) end + | dest_fcoeff sign t = + let val (p, t') = dest_coeff sign t + val T = Term.fastype_of t + in (mk_frac (p, 1), mk_divide (t', one_of T)) end; + + +(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1, 1*x, x*1, x/1 *) +val add_0s = thms "add_0s"; +val mult_1s = thms "mult_1s" @ [thm"mult_1_left", thm"mult_1_right", thm"divide_1"]; + +(*Simplify inverse Numeral1, a/Numeral1*) +val inverse_1s = [@{thm inverse_numeral_1}]; +val divide_1s = [@{thm divide_numeral_1}]; + +(*To perform binary arithmetic. The "left" rewriting handles patterns + created by the Int_Numeral_Base_Simprocs, such as 3 * (5 * x). *) +val simps = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym, + @{thm add_number_of_left}, @{thm mult_number_of_left}] @ + @{thms arith_simps} @ @{thms rel_simps}; + +(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms + during re-arrangement*) +val non_add_simps = + subtract Thm.eq_thm [@{thm add_number_of_left}, @{thm number_of_add} RS sym] simps; + +(*To evaluate binary negations of coefficients*) +val minus_simps = [@{thm numeral_m1_eq_minus_1} RS sym, @{thm number_of_minus} RS sym] @ + @{thms minus_bin_simps} @ @{thms pred_bin_simps}; + +(*To let us treat subtraction as addition*) +val diff_simps = [@{thm diff_minus}, @{thm minus_add_distrib}, @{thm minus_minus}]; + +(*To let us treat division as multiplication*) +val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}]; + +(*push the unary minus down: - x * y = x * - y *) +val minus_mult_eq_1_to_2 = + [@{thm minus_mult_left} RS sym, @{thm minus_mult_right}] MRS trans |> standard; + +(*to extract again any uncancelled minuses*) +val minus_from_mult_simps = + [@{thm minus_minus}, @{thm minus_mult_left} RS sym, @{thm minus_mult_right} RS sym]; + +(*combine unary minus with numeric literals, however nested within a product*) +val mult_minus_simps = + [@{thm mult_assoc}, @{thm minus_mult_left}, minus_mult_eq_1_to_2]; + +(*Apply the given rewrite (if present) just once*) +fun trans_tac NONE = all_tac + | trans_tac (SOME th) = ALLGOALS (rtac (th RS trans)); + +fun simplify_meta_eq rules = + let val ss0 = HOL_basic_ss addeqcongs [eq_cong2] addsimps rules + in fn ss => simplify (Simplifier.inherit_context ss ss0) o mk_meta_eq end + +structure CancelNumeralsCommon = + struct + val mk_sum = mk_sum + val dest_sum = dest_sum + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val find_first_coeff = find_first_coeff [] + val trans_tac = fn _ => trans_tac + + val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ + diff_simps @ minus_simps @ @{thms add_ac} + val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps + val norm_ss3 = num_ss addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac} + fun norm_tac ss = + ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) + + val numeral_simp_ss = HOL_ss addsimps add_0s @ simps + fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s) + end; + + +structure EqCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" Term.dummyT + val bal_add1 = @{thm eq_add_iff1} RS trans + val bal_add2 = @{thm eq_add_iff2} RS trans +); + +structure LessCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less} + val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT + val bal_add1 = @{thm less_add_iff1} RS trans + val bal_add2 = @{thm less_add_iff2} RS trans +); + +structure LeCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq} + val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT + val bal_add1 = @{thm le_add_iff1} RS trans + val bal_add2 = @{thm le_add_iff2} RS trans +); + +val cancel_numerals = + map Int_Numeral_Base_Simprocs.prep_simproc + [("inteq_cancel_numerals", + ["(l::'a::number_ring) + m = n", + "(l::'a::number_ring) = m + n", + "(l::'a::number_ring) - m = n", + "(l::'a::number_ring) = m - n", + "(l::'a::number_ring) * m = n", + "(l::'a::number_ring) = m * n"], + K EqCancelNumerals.proc), + ("intless_cancel_numerals", + ["(l::'a::{ordered_idom,number_ring}) + m < n", + "(l::'a::{ordered_idom,number_ring}) < m + n", + "(l::'a::{ordered_idom,number_ring}) - m < n", + "(l::'a::{ordered_idom,number_ring}) < m - n", + "(l::'a::{ordered_idom,number_ring}) * m < n", + "(l::'a::{ordered_idom,number_ring}) < m * n"], + K LessCancelNumerals.proc), + ("intle_cancel_numerals", + ["(l::'a::{ordered_idom,number_ring}) + m <= n", + "(l::'a::{ordered_idom,number_ring}) <= m + n", + "(l::'a::{ordered_idom,number_ring}) - m <= n", + "(l::'a::{ordered_idom,number_ring}) <= m - n", + "(l::'a::{ordered_idom,number_ring}) * m <= n", + "(l::'a::{ordered_idom,number_ring}) <= m * n"], + K LeCancelNumerals.proc)]; + + +structure CombineNumeralsData = + struct + type coeff = int + val iszero = (fn x => x = 0) + val add = op + + val mk_sum = long_mk_sum (*to work for e.g. 2*x + 3*x *) + val dest_sum = dest_sum + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val left_distrib = @{thm combine_common_factor} RS trans + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv_nohyps + val trans_tac = fn _ => trans_tac + + val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ + diff_simps @ minus_simps @ @{thms add_ac} + val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps + val norm_ss3 = num_ss addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac} + fun norm_tac ss = + ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) + + val numeral_simp_ss = HOL_ss addsimps add_0s @ simps + fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s) + end; + +structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); + +(*Version for fields, where coefficients can be fractions*) +structure FieldCombineNumeralsData = + struct + type coeff = int * int + val iszero = (fn (p, q) => p = 0) + val add = add_frac + val mk_sum = long_mk_sum + val dest_sum = dest_sum + val mk_coeff = mk_fcoeff + val dest_coeff = dest_fcoeff 1 + val left_distrib = @{thm combine_common_factor} RS trans + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv_nohyps + val trans_tac = fn _ => trans_tac + + val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ + inverse_1s @ divide_simps @ diff_simps @ minus_simps @ @{thms add_ac} + val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps + val norm_ss3 = num_ss addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac} + fun norm_tac ss = + ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) + + val numeral_simp_ss = HOL_ss addsimps add_0s @ simps @ [@{thm add_frac_eq}] + fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s @ divide_1s) + end; + +structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData); + +val combine_numerals = + Int_Numeral_Base_Simprocs.prep_simproc + ("int_combine_numerals", + ["(i::'a::number_ring) + j", "(i::'a::number_ring) - j"], + K CombineNumerals.proc); + +val field_combine_numerals = + Int_Numeral_Base_Simprocs.prep_simproc + ("field_combine_numerals", + ["(i::'a::{number_ring,field,division_by_zero}) + j", + "(i::'a::{number_ring,field,division_by_zero}) - j"], + K FieldCombineNumerals.proc); + +end; + +Addsimprocs Int_Numeral_Simprocs.cancel_numerals; +Addsimprocs [Int_Numeral_Simprocs.combine_numerals]; +Addsimprocs [Int_Numeral_Simprocs.field_combine_numerals]; + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s, by (Simp_tac 1)); + +test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::int)"; + +test "2*u = (u::int)"; +test "(i + j + 12 + (k::int)) - 15 = y"; +test "(i + j + 12 + (k::int)) - 5 = y"; + +test "y - b < (b::int)"; +test "y - (3*b + c) < (b::int) - 2*c"; + +test "(2*x - (u*v) + y) - v*3*u = (w::int)"; +test "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::int)"; +test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::int)"; +test "u*v - (x*u*v + (u*v)*4 + y) = (w::int)"; + +test "(i + j + 12 + (k::int)) = u + 15 + y"; +test "(i + j*2 + 12 + (k::int)) = j + 5 + y"; + +test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::int)"; + +test "a + -(b+c) + b = (d::int)"; +test "a + -(b+c) - b = (d::int)"; + +(*negative numerals*) +test "(i + j + -2 + (k::int)) - (u + 5 + y) = zz"; +test "(i + j + -3 + (k::int)) < u + 5 + y"; +test "(i + j + 3 + (k::int)) < u + -6 + y"; +test "(i + j + -12 + (k::int)) - 15 = y"; +test "(i + j + 12 + (k::int)) - -15 = y"; +test "(i + j + -12 + (k::int)) - -15 = y"; +*) + + +(** Constant folding for multiplication in semirings **) + +(*We do not need folding for addition: combine_numerals does the same thing*) + +structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA = +struct + val assoc_ss = HOL_ss addsimps @{thms mult_ac} + val eq_reflection = eq_reflection +end; + +structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data); + +val assoc_fold_simproc = + Int_Numeral_Base_Simprocs.prep_simproc + ("semiring_assoc_fold", ["(a::'a::comm_semiring_1_cancel) * b"], + K Semiring_Times_Assoc.proc); + +Addsimprocs [assoc_fold_simproc]; + + + + +(*** decision procedure for linear arithmetic ***) + +(*---------------------------------------------------------------------------*) +(* Linear arithmetic *) +(*---------------------------------------------------------------------------*) + +(* +Instantiation of the generic linear arithmetic package for int. +*) + +(* Update parameters of arithmetic prover *) +local + +(* reduce contradictory =/ if s = @{const_name "HOL.one"} then not (t = @{typ int}) + else s mem_string allowed_consts + | a$b => check a andalso check b + | _ => false; + +val conv = + Simplifier.rewrite + (HOL_basic_ss addsimps + ((map (fn th => th RS sym) [@{thm of_int_add}, @{thm of_int_mult}, + @{thm of_int_diff}, @{thm of_int_minus}])@ + [@{thm of_int_less_iff}, @{thm of_int_le_iff}, @{thm of_int_eq_iff}]) + addsimprocs [zero_to_of_int_zero_simproc,one_to_of_int_one_simproc]); + +fun sproc phi ss ct = if check (term_of ct) then SOME (conv ct) else NONE +val lhss' = + [@{cpat "(?x::?'a::ring_char_0) = (?y::?'a)"}, + @{cpat "(?x::?'a::ordered_idom) < (?y::?'a)"}, + @{cpat "(?x::?'a::ordered_idom) <= (?y::?'a)"}] +in +val zero_one_idom_simproc = + make_simproc {lhss = lhss' , name = "zero_one_idom_simproc", + proc = sproc, identifier = []} +end; + +val add_rules = + simp_thms @ @{thms arith_simps} @ @{thms rel_simps} @ @{thms arith_special} @ + [@{thm neg_le_iff_le}, @{thm numeral_0_eq_0}, @{thm numeral_1_eq_1}, + @{thm minus_zero}, @{thm diff_minus}, @{thm left_minus}, @{thm right_minus}, + @{thm mult_zero_left}, @{thm mult_zero_right}, @{thm mult_Bit1}, @{thm mult_1_right}, + @{thm minus_mult_left} RS sym, @{thm minus_mult_right} RS sym, + @{thm minus_add_distrib}, @{thm minus_minus}, @{thm mult_assoc}, + @{thm of_nat_0}, @{thm of_nat_1}, @{thm of_nat_Suc}, @{thm of_nat_add}, + @{thm of_nat_mult}, @{thm of_int_0}, @{thm of_int_1}, @{thm of_int_add}, + @{thm of_int_mult}] + +val nat_inj_thms = [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2] + +val Int_Numeral_Base_Simprocs = assoc_fold_simproc :: zero_one_idom_simproc + :: Int_Numeral_Simprocs.combine_numerals + :: Int_Numeral_Simprocs.cancel_numerals; + +in + +val int_arith_setup = + LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} => + {add_mono_thms = add_mono_thms, + mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono} :: mult_mono_thms, + inj_thms = nat_inj_thms @ inj_thms, + lessD = lessD @ [@{thm zless_imp_add1_zle}], + neqE = neqE, + simpset = simpset addsimps add_rules + addsimprocs Int_Numeral_Base_Simprocs + addcongs [if_weak_cong]}) #> + arith_inj_const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT) #> + arith_discrete @{type_name Int.int} + +end; + +val fast_int_arith_simproc = + Simplifier.simproc (the_context ()) + "fast_int_arith" + ["(m::'a::{ordered_idom,number_ring}) < n", + "(m::'a::{ordered_idom,number_ring}) <= n", + "(m::'a::{ordered_idom,number_ring}) = n"] (K LinArith.lin_arith_simproc); + +Addsimprocs [fast_int_arith_simproc]; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Tools/int_factor_simprocs.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/int_factor_simprocs.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,340 @@ +(* Title: HOL/int_factor_simprocs.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 2000 University of Cambridge + +Factor cancellation simprocs for the integers (and for fields). + +This file can't be combined with int_arith1 because it requires IntDiv.thy. +*) + + +(*To quote from Provers/Arith/cancel_numeral_factor.ML: + +Cancels common coefficients in balanced expressions: + + u*#m ~~ u'*#m' == #n*u ~~ #n'*u' + +where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /) +and d = gcd(m,m') and n=m/d and n'=m'/d. +*) + +val rel_number_of = [@{thm eq_number_of_eq}, @{thm less_number_of_eq_neg}, @{thm le_number_of_eq}]; + +local + open Int_Numeral_Simprocs +in + +structure CancelNumeralFactorCommon = + struct + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val trans_tac = fn _ => trans_tac + + val norm_ss1 = HOL_ss addsimps minus_from_mult_simps @ mult_1s + val norm_ss2 = HOL_ss addsimps simps @ mult_minus_simps + val norm_ss3 = HOL_ss addsimps @{thms mult_ac} + fun norm_tac ss = + ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) + + val numeral_simp_ss = HOL_ss addsimps rel_number_of @ simps + fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val simplify_meta_eq = Int_Numeral_Simprocs.simplify_meta_eq + [@{thm add_0}, @{thm add_0_right}, @{thm mult_zero_left}, + @{thm mult_zero_right}, @{thm mult_Bit1}, @{thm mult_1_right}]; + end + +(*Version for integer division*) +structure IntDivCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop @{const_name Divides.div} + val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT + val cancel = @{thm zdiv_zmult_zmult1} RS trans + val neg_exchanges = false +) + +(*Version for fields*) +structure DivideCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop @{const_name HOL.divide} + val dest_bal = HOLogic.dest_bin @{const_name HOL.divide} Term.dummyT + val cancel = @{thm mult_divide_mult_cancel_left} RS trans + val neg_exchanges = false +) + +structure EqCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" Term.dummyT + val cancel = @{thm mult_cancel_left} RS trans + val neg_exchanges = false +) + +structure LessCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less} + val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT + val cancel = @{thm mult_less_cancel_left} RS trans + val neg_exchanges = true +) + +structure LeCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq} + val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT + val cancel = @{thm mult_le_cancel_left} RS trans + val neg_exchanges = true +) + +val cancel_numeral_factors = + map Int_Numeral_Base_Simprocs.prep_simproc + [("ring_eq_cancel_numeral_factor", + ["(l::'a::{idom,number_ring}) * m = n", + "(l::'a::{idom,number_ring}) = m * n"], + K EqCancelNumeralFactor.proc), + ("ring_less_cancel_numeral_factor", + ["(l::'a::{ordered_idom,number_ring}) * m < n", + "(l::'a::{ordered_idom,number_ring}) < m * n"], + K LessCancelNumeralFactor.proc), + ("ring_le_cancel_numeral_factor", + ["(l::'a::{ordered_idom,number_ring}) * m <= n", + "(l::'a::{ordered_idom,number_ring}) <= m * n"], + K LeCancelNumeralFactor.proc), + ("int_div_cancel_numeral_factors", + ["((l::int) * m) div n", "(l::int) div (m * n)"], + K IntDivCancelNumeralFactor.proc), + ("divide_cancel_numeral_factor", + ["((l::'a::{division_by_zero,field,number_ring}) * m) / n", + "(l::'a::{division_by_zero,field,number_ring}) / (m * n)", + "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"], + K DivideCancelNumeralFactor.proc)]; + +(* referenced by rat_arith.ML *) +val field_cancel_numeral_factors = + map Int_Numeral_Base_Simprocs.prep_simproc + [("field_eq_cancel_numeral_factor", + ["(l::'a::{field,number_ring}) * m = n", + "(l::'a::{field,number_ring}) = m * n"], + K EqCancelNumeralFactor.proc), + ("field_cancel_numeral_factor", + ["((l::'a::{division_by_zero,field,number_ring}) * m) / n", + "(l::'a::{division_by_zero,field,number_ring}) / (m * n)", + "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"], + K DivideCancelNumeralFactor.proc)] + +end; + +Addsimprocs cancel_numeral_factors; + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s; by (Simp_tac 1)); + +test "9*x = 12 * (y::int)"; +test "(9*x) div (12 * (y::int)) = z"; +test "9*x < 12 * (y::int)"; +test "9*x <= 12 * (y::int)"; + +test "-99*x = 132 * (y::int)"; +test "(-99*x) div (132 * (y::int)) = z"; +test "-99*x < 132 * (y::int)"; +test "-99*x <= 132 * (y::int)"; + +test "999*x = -396 * (y::int)"; +test "(999*x) div (-396 * (y::int)) = z"; +test "999*x < -396 * (y::int)"; +test "999*x <= -396 * (y::int)"; + +test "-99*x = -81 * (y::int)"; +test "(-99*x) div (-81 * (y::int)) = z"; +test "-99*x <= -81 * (y::int)"; +test "-99*x < -81 * (y::int)"; + +test "-2 * x = -1 * (y::int)"; +test "-2 * x = -(y::int)"; +test "(-2 * x) div (-1 * (y::int)) = z"; +test "-2 * x < -(y::int)"; +test "-2 * x <= -1 * (y::int)"; +test "-x < -23 * (y::int)"; +test "-x <= -23 * (y::int)"; +*) + +(*And the same examples for fields such as rat or real: +test "0 <= (y::rat) * -2"; +test "9*x = 12 * (y::rat)"; +test "(9*x) / (12 * (y::rat)) = z"; +test "9*x < 12 * (y::rat)"; +test "9*x <= 12 * (y::rat)"; + +test "-99*x = 132 * (y::rat)"; +test "(-99*x) / (132 * (y::rat)) = z"; +test "-99*x < 132 * (y::rat)"; +test "-99*x <= 132 * (y::rat)"; + +test "999*x = -396 * (y::rat)"; +test "(999*x) / (-396 * (y::rat)) = z"; +test "999*x < -396 * (y::rat)"; +test "999*x <= -396 * (y::rat)"; + +test "(- ((2::rat) * x) <= 2 * y)"; +test "-99*x = -81 * (y::rat)"; +test "(-99*x) / (-81 * (y::rat)) = z"; +test "-99*x <= -81 * (y::rat)"; +test "-99*x < -81 * (y::rat)"; + +test "-2 * x = -1 * (y::rat)"; +test "-2 * x = -(y::rat)"; +test "(-2 * x) / (-1 * (y::rat)) = z"; +test "-2 * x < -(y::rat)"; +test "-2 * x <= -1 * (y::rat)"; +test "-x < -23 * (y::rat)"; +test "-x <= -23 * (y::rat)"; +*) + + +(** Declarations for ExtractCommonTerm **) + +local + open Int_Numeral_Simprocs +in + +(*Find first term that matches u*) +fun find_first_t past u [] = raise TERM ("find_first_t", []) + | find_first_t past u (t::terms) = + if u aconv t then (rev past @ terms) + else find_first_t (t::past) u terms + handle TERM _ => find_first_t (t::past) u terms; + +(** Final simplification for the CancelFactor simprocs **) +val simplify_one = Int_Numeral_Simprocs.simplify_meta_eq + [@{thm mult_1_left}, @{thm mult_1_right}, @{thm zdiv_1}, @{thm numeral_1_eq_1}]; + +fun cancel_simplify_meta_eq cancel_th ss th = + simplify_one ss (([th, cancel_th]) MRS trans); + +structure CancelFactorCommon = + struct + val mk_sum = long_mk_prod + val dest_sum = dest_prod + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff + val find_first = find_first_t [] + val trans_tac = fn _ => trans_tac + val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac} + fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) + end; + +(*mult_cancel_left requires a ring with no zero divisors.*) +structure EqCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" Term.dummyT + val simplify_meta_eq = cancel_simplify_meta_eq @{thm mult_cancel_left} +); + +(*zdiv_zmult_zmult1_if is for integer division (div).*) +structure IntDivCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop @{const_name Divides.div} + val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT + val simplify_meta_eq = cancel_simplify_meta_eq @{thm zdiv_zmult_zmult1_if} +); + +structure IntModCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop @{const_name Divides.mod} + val dest_bal = HOLogic.dest_bin @{const_name Divides.mod} HOLogic.intT + val simplify_meta_eq = cancel_simplify_meta_eq @{thm zmod_zmult_zmult1} +); + +structure IntDvdCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} + val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.intT + val simplify_meta_eq = cancel_simplify_meta_eq @{thm zdvd_zmult_cancel_disj} +); + +(*Version for all fields, including unordered ones (type complex).*) +structure DivideCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop @{const_name HOL.divide} + val dest_bal = HOLogic.dest_bin @{const_name HOL.divide} Term.dummyT + val simplify_meta_eq = cancel_simplify_meta_eq @{thm mult_divide_mult_cancel_left_if} +); + +val cancel_factors = + map Int_Numeral_Base_Simprocs.prep_simproc + [("ring_eq_cancel_factor", + ["(l::'a::{idom}) * m = n", + "(l::'a::{idom}) = m * n"], + K EqCancelFactor.proc), + ("int_div_cancel_factor", + ["((l::int) * m) div n", "(l::int) div (m * n)"], + K IntDivCancelFactor.proc), + ("int_mod_cancel_factor", + ["((l::int) * m) mod n", "(l::int) mod (m * n)"], + K IntModCancelFactor.proc), + ("int_dvd_cancel_factor", + ["((l::int) * m) dvd n", "(l::int) dvd (m * n)"], + K IntDvdCancelFactor.proc), + ("divide_cancel_factor", + ["((l::'a::{division_by_zero,field}) * m) / n", + "(l::'a::{division_by_zero,field}) / (m * n)"], + K DivideCancelFactor.proc)]; + +end; + +Addsimprocs cancel_factors; + + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s; by (Asm_simp_tac 1)); + +test "x*k = k*(y::int)"; +test "k = k*(y::int)"; +test "a*(b*c) = (b::int)"; +test "a*(b*c) = d*(b::int)*(x*a)"; + +test "(x*k) div (k*(y::int)) = (uu::int)"; +test "(k) div (k*(y::int)) = (uu::int)"; +test "(a*(b*c)) div ((b::int)) = (uu::int)"; +test "(a*(b*c)) div (d*(b::int)*(x*a)) = (uu::int)"; +*) + +(*And the same examples for fields such as rat or real: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s; by (Asm_simp_tac 1)); + +test "x*k = k*(y::rat)"; +test "k = k*(y::rat)"; +test "a*(b*c) = (b::rat)"; +test "a*(b*c) = d*(b::rat)*(x*a)"; + + +test "(x*k) / (k*(y::rat)) = (uu::rat)"; +test "(k) / (k*(y::rat)) = (uu::rat)"; +test "(a*(b*c)) / ((b::rat)) = (uu::rat)"; +test "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)"; + +(*FIXME: what do we do about this?*) +test "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z"; +*) diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Tools/nat_simprocs.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/nat_simprocs.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,585 @@ +(* Title: HOL/nat_simprocs.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 2000 University of Cambridge + +Simprocs for nat numerals. +*) + +structure Nat_Numeral_Simprocs = +struct + +(*Maps n to #n for n = 0, 1, 2*) +val numeral_syms = + [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym, @{thm numeral_2_eq_2} RS sym]; +val numeral_sym_ss = HOL_ss addsimps numeral_syms; + +fun rename_numerals th = + simplify numeral_sym_ss (Thm.transfer (the_context ()) th); + +(*Utilities*) + +fun mk_number n = HOLogic.number_of_const HOLogic.natT $ HOLogic.mk_numeral n; +fun dest_number t = Int.max (0, snd (HOLogic.dest_number t)); + +fun find_first_numeral past (t::terms) = + ((dest_number t, t, rev past @ terms) + handle TERM _ => find_first_numeral (t::past) terms) + | find_first_numeral past [] = raise TERM("find_first_numeral", []); + +val zero = mk_number 0; +val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; + +(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) +fun mk_sum [] = zero + | mk_sum [t,u] = mk_plus (t, u) + | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); + +(*this version ALWAYS includes a trailing zero*) +fun long_mk_sum [] = HOLogic.zero + | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts); + +val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} HOLogic.natT; + + +(** Other simproc items **) + +val trans_tac = Int_Numeral_Simprocs.trans_tac; + +val bin_simps = + [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym, + @{thm add_nat_number_of}, @{thm nat_number_of_add_left}, + @{thm diff_nat_number_of}, @{thm le_number_of_eq_not_less}, + @{thm mult_nat_number_of}, @{thm nat_number_of_mult_left}, + @{thm less_nat_number_of}, + @{thm Let_number_of}, @{thm nat_number_of}] @ + @{thms arith_simps} @ @{thms rel_simps}; + +fun prep_simproc (name, pats, proc) = + Simplifier.simproc (the_context ()) name pats proc; + + +(*** CancelNumerals simprocs ***) + +val one = mk_number 1; +val mk_times = HOLogic.mk_binop @{const_name HOL.times}; + +fun mk_prod [] = one + | mk_prod [t] = t + | mk_prod (t :: ts) = if t = one then mk_prod ts + else mk_times (t, mk_prod ts); + +val dest_times = HOLogic.dest_bin @{const_name HOL.times} HOLogic.natT; + +fun dest_prod t = + let val (t,u) = dest_times t + in dest_prod t @ dest_prod u end + handle TERM _ => [t]; + +(*DON'T do the obvious simplifications; that would create special cases*) +fun mk_coeff (k,t) = mk_times (mk_number k, t); + +(*Express t as a product of (possibly) a numeral with other factors, sorted*) +fun dest_coeff t = + let val ts = sort Term.term_ord (dest_prod t) + val (n, _, ts') = find_first_numeral [] ts + handle TERM _ => (1, one, ts) + in (n, mk_prod ts') end; + +(*Find first coefficient-term THAT MATCHES u*) +fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) + | find_first_coeff past u (t::terms) = + let val (n,u') = dest_coeff t + in if u aconv u' then (n, rev past @ terms) + else find_first_coeff (t::past) u terms + end + handle TERM _ => find_first_coeff (t::past) u terms; + + +(*Split up a sum into the list of its constituent terms, on the way removing any + Sucs and counting them.*) +fun dest_Suc_sum (Const ("Suc", _) $ t, (k,ts)) = dest_Suc_sum (t, (k+1,ts)) + | dest_Suc_sum (t, (k,ts)) = + let val (t1,t2) = dest_plus t + in dest_Suc_sum (t1, dest_Suc_sum (t2, (k,ts))) end + handle TERM _ => (k, t::ts); + +(*Code for testing whether numerals are already used in the goal*) +fun is_numeral (Const(@{const_name Int.number_of}, _) $ w) = true + | is_numeral _ = false; + +fun prod_has_numeral t = exists is_numeral (dest_prod t); + +(*The Sucs found in the term are converted to a binary numeral. If relaxed is false, + an exception is raised unless the original expression contains at least one + numeral in a coefficient position. This prevents nat_combine_numerals from + introducing numerals to goals.*) +fun dest_Sucs_sum relaxed t = + let val (k,ts) = dest_Suc_sum (t,(0,[])) + in + if relaxed orelse exists prod_has_numeral ts then + if k=0 then ts + else mk_number k :: ts + else raise TERM("Nat_Numeral_Simprocs.dest_Sucs_sum", [t]) + end; + + +(*Simplify 1*n and n*1 to n*) +val add_0s = map rename_numerals [@{thm add_0}, @{thm add_0_right}]; +val mult_1s = map rename_numerals [@{thm nat_mult_1}, @{thm nat_mult_1_right}]; + +(*Final simplification: cancel + and *; replace Numeral0 by 0 and Numeral1 by 1*) + +(*And these help the simproc return False when appropriate, which helps + the arith prover.*) +val contra_rules = [@{thm add_Suc}, @{thm add_Suc_right}, @{thm Zero_not_Suc}, + @{thm Suc_not_Zero}, @{thm le_0_eq}]; + +val simplify_meta_eq = + Int_Numeral_Simprocs.simplify_meta_eq + ([@{thm nat_numeral_0_eq_0}, @{thm numeral_1_eq_Suc_0}, @{thm add_0}, @{thm add_0_right}, + @{thm mult_0}, @{thm mult_0_right}, @{thm mult_1}, @{thm mult_1_right}] @ contra_rules); + + +(*Like HOL_ss but with an ordering that brings numerals to the front + under AC-rewriting.*) +val num_ss = Int_Numeral_Simprocs.num_ss; + +(*** Applying CancelNumeralsFun ***) + +structure CancelNumeralsCommon = + struct + val mk_sum = (fn T:typ => mk_sum) + val dest_sum = dest_Sucs_sum true + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff + val find_first_coeff = find_first_coeff [] + val trans_tac = fn _ => trans_tac + + val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ + [@{thm Suc_eq_add_numeral_1_left}] @ @{thms add_ac} + val norm_ss2 = num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac} + fun norm_tac ss = + ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) + + val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps; + fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)); + val simplify_meta_eq = simplify_meta_eq + end; + + +structure EqCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT + val bal_add1 = @{thm nat_eq_add_iff1} RS trans + val bal_add2 = @{thm nat_eq_add_iff2} RS trans +); + +structure LessCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less} + val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT + val bal_add1 = @{thm nat_less_add_iff1} RS trans + val bal_add2 = @{thm nat_less_add_iff2} RS trans +); + +structure LeCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq} + val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT + val bal_add1 = @{thm nat_le_add_iff1} RS trans + val bal_add2 = @{thm nat_le_add_iff2} RS trans +); + +structure DiffCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop @{const_name HOL.minus} + val dest_bal = HOLogic.dest_bin @{const_name HOL.minus} HOLogic.natT + val bal_add1 = @{thm nat_diff_add_eq1} RS trans + val bal_add2 = @{thm nat_diff_add_eq2} RS trans +); + + +val cancel_numerals = + map prep_simproc + [("nateq_cancel_numerals", + ["(l::nat) + m = n", "(l::nat) = m + n", + "(l::nat) * m = n", "(l::nat) = m * n", + "Suc m = n", "m = Suc n"], + K EqCancelNumerals.proc), + ("natless_cancel_numerals", + ["(l::nat) + m < n", "(l::nat) < m + n", + "(l::nat) * m < n", "(l::nat) < m * n", + "Suc m < n", "m < Suc n"], + K LessCancelNumerals.proc), + ("natle_cancel_numerals", + ["(l::nat) + m <= n", "(l::nat) <= m + n", + "(l::nat) * m <= n", "(l::nat) <= m * n", + "Suc m <= n", "m <= Suc n"], + K LeCancelNumerals.proc), + ("natdiff_cancel_numerals", + ["((l::nat) + m) - n", "(l::nat) - (m + n)", + "(l::nat) * m - n", "(l::nat) - m * n", + "Suc m - n", "m - Suc n"], + K DiffCancelNumerals.proc)]; + + +(*** Applying CombineNumeralsFun ***) + +structure CombineNumeralsData = + struct + type coeff = int + val iszero = (fn x => x = 0) + val add = op + + val mk_sum = (fn T:typ => long_mk_sum) (*to work for 2*x + 3*x *) + val dest_sum = dest_Sucs_sum false + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff + val left_distrib = @{thm left_add_mult_distrib} RS trans + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv_nohyps + val trans_tac = fn _ => trans_tac + + val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_add_numeral_1}] @ @{thms add_ac} + val norm_ss2 = num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac} + fun norm_tac ss = + ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) + + val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps; + fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val simplify_meta_eq = simplify_meta_eq + end; + +structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); + +val combine_numerals = + prep_simproc ("nat_combine_numerals", ["(i::nat) + j", "Suc (i + j)"], K CombineNumerals.proc); + + +(*** Applying CancelNumeralFactorFun ***) + +structure CancelNumeralFactorCommon = + struct + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff + val trans_tac = fn _ => trans_tac + + val norm_ss1 = num_ss addsimps + numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_add_numeral_1_left}] @ @{thms add_ac} + val norm_ss2 = num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac} + fun norm_tac ss = + ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) + THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) + + val numeral_simp_ss = HOL_ss addsimps bin_simps + fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val simplify_meta_eq = simplify_meta_eq + end + +structure DivCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop @{const_name Divides.div} + val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT + val cancel = @{thm nat_mult_div_cancel1} RS trans + val neg_exchanges = false +) + +structure DvdCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} + val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT + val cancel = @{thm nat_mult_dvd_cancel1} RS trans + val neg_exchanges = false +) + +structure EqCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT + val cancel = @{thm nat_mult_eq_cancel1} RS trans + val neg_exchanges = false +) + +structure LessCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less} + val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT + val cancel = @{thm nat_mult_less_cancel1} RS trans + val neg_exchanges = true +) + +structure LeCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq} + val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT + val cancel = @{thm nat_mult_le_cancel1} RS trans + val neg_exchanges = true +) + +val cancel_numeral_factors = + map prep_simproc + [("nateq_cancel_numeral_factors", + ["(l::nat) * m = n", "(l::nat) = m * n"], + K EqCancelNumeralFactor.proc), + ("natless_cancel_numeral_factors", + ["(l::nat) * m < n", "(l::nat) < m * n"], + K LessCancelNumeralFactor.proc), + ("natle_cancel_numeral_factors", + ["(l::nat) * m <= n", "(l::nat) <= m * n"], + K LeCancelNumeralFactor.proc), + ("natdiv_cancel_numeral_factors", + ["((l::nat) * m) div n", "(l::nat) div (m * n)"], + K DivCancelNumeralFactor.proc), + ("natdvd_cancel_numeral_factors", + ["((l::nat) * m) dvd n", "(l::nat) dvd (m * n)"], + K DvdCancelNumeralFactor.proc)]; + + + +(*** Applying ExtractCommonTermFun ***) + +(*this version ALWAYS includes a trailing one*) +fun long_mk_prod [] = one + | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts); + +(*Find first term that matches u*) +fun find_first_t past u [] = raise TERM("find_first_t", []) + | find_first_t past u (t::terms) = + if u aconv t then (rev past @ terms) + else find_first_t (t::past) u terms + handle TERM _ => find_first_t (t::past) u terms; + +(** Final simplification for the CancelFactor simprocs **) +val simplify_one = Int_Numeral_Simprocs.simplify_meta_eq + [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_1}, @{thm numeral_1_eq_Suc_0}]; + +fun cancel_simplify_meta_eq cancel_th ss th = + simplify_one ss (([th, cancel_th]) MRS trans); + +structure CancelFactorCommon = + struct + val mk_sum = (fn T:typ => long_mk_prod) + val dest_sum = dest_prod + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff + val find_first = find_first_t [] + val trans_tac = fn _ => trans_tac + val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac} + fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) + end; + +structure EqCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT + val simplify_meta_eq = cancel_simplify_meta_eq @{thm nat_mult_eq_cancel_disj} +); + +structure LessCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less} + val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT + val simplify_meta_eq = cancel_simplify_meta_eq @{thm nat_mult_less_cancel_disj} +); + +structure LeCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq} + val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT + val simplify_meta_eq = cancel_simplify_meta_eq @{thm nat_mult_le_cancel_disj} +); + +structure DivideCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop @{const_name Divides.div} + val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT + val simplify_meta_eq = cancel_simplify_meta_eq @{thm nat_mult_div_cancel_disj} +); + +structure DvdCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Int_Numeral_Base_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} + val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT + val simplify_meta_eq = cancel_simplify_meta_eq @{thm nat_mult_dvd_cancel_disj} +); + +val cancel_factor = + map prep_simproc + [("nat_eq_cancel_factor", + ["(l::nat) * m = n", "(l::nat) = m * n"], + K EqCancelFactor.proc), + ("nat_less_cancel_factor", + ["(l::nat) * m < n", "(l::nat) < m * n"], + K LessCancelFactor.proc), + ("nat_le_cancel_factor", + ["(l::nat) * m <= n", "(l::nat) <= m * n"], + K LeCancelFactor.proc), + ("nat_divide_cancel_factor", + ["((l::nat) * m) div n", "(l::nat) div (m * n)"], + K DivideCancelFactor.proc), + ("nat_dvd_cancel_factor", + ["((l::nat) * m) dvd n", "(l::nat) dvd (m * n)"], + K DvdCancelFactor.proc)]; + +end; + + +Addsimprocs Nat_Numeral_Simprocs.cancel_numerals; +Addsimprocs [Nat_Numeral_Simprocs.combine_numerals]; +Addsimprocs Nat_Numeral_Simprocs.cancel_numeral_factors; +Addsimprocs Nat_Numeral_Simprocs.cancel_factor; + + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s; by (Simp_tac 1)); + +(*cancel_numerals*) +test "l +( 2) + (2) + 2 + (l + 2) + (oo + 2) = (uu::nat)"; +test "(2*length xs < 2*length xs + j)"; +test "(2*length xs < length xs * 2 + j)"; +test "2*u = (u::nat)"; +test "2*u = Suc (u)"; +test "(i + j + 12 + (k::nat)) - 15 = y"; +test "(i + j + 12 + (k::nat)) - 5 = y"; +test "Suc u - 2 = y"; +test "Suc (Suc (Suc u)) - 2 = y"; +test "(i + j + 2 + (k::nat)) - 1 = y"; +test "(i + j + 1 + (k::nat)) - 2 = y"; + +test "(2*x + (u*v) + y) - v*3*u = (w::nat)"; +test "(2*x*u*v + 5 + (u*v)*4 + y) - v*u*4 = (w::nat)"; +test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::nat)"; +test "Suc (Suc (2*x*u*v + u*4 + y)) - u = w"; +test "Suc ((u*v)*4) - v*3*u = w"; +test "Suc (Suc ((u*v)*3)) - v*3*u = w"; + +test "(i + j + 12 + (k::nat)) = u + 15 + y"; +test "(i + j + 32 + (k::nat)) - (u + 15 + y) = zz"; +test "(i + j + 12 + (k::nat)) = u + 5 + y"; +(*Suc*) +test "(i + j + 12 + k) = Suc (u + y)"; +test "Suc (Suc (Suc (Suc (Suc (u + y))))) <= ((i + j) + 41 + k)"; +test "(i + j + 5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))"; +test "Suc (Suc (Suc (Suc (Suc (u + y))))) - 5 = v"; +test "(i + j + 5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))"; +test "2*y + 3*z + 2*u = Suc (u)"; +test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = Suc (u)"; +test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::nat)"; +test "6 + 2*y + 3*z + 4*u = Suc (vv + 2*u + z)"; +test "(2*n*m) < (3*(m*n)) + (u::nat)"; + +test "(Suc (Suc (Suc (Suc (Suc (Suc (case length (f c) of 0 => 0 | Suc k => k)))))) <= Suc 0)"; + +test "Suc (Suc (Suc (Suc (Suc (Suc (length l1 + length l2)))))) <= length l1"; + +test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length l3)))))) <= length (compT P E A ST mxr e))"; + +test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length (compT P E (A Un \ e) ST mxr c))))))) <= length (compT P E A ST mxr e))"; + + +(*negative numerals: FAIL*) +test "(i + j + -23 + (k::nat)) < u + 15 + y"; +test "(i + j + 3 + (k::nat)) < u + -15 + y"; +test "(i + j + -12 + (k::nat)) - 15 = y"; +test "(i + j + 12 + (k::nat)) - -15 = y"; +test "(i + j + -12 + (k::nat)) - -15 = y"; + +(*combine_numerals*) +test "k + 3*k = (u::nat)"; +test "Suc (i + 3) = u"; +test "Suc (i + j + 3 + k) = u"; +test "k + j + 3*k + j = (u::nat)"; +test "Suc (j*i + i + k + 5 + 3*k + i*j*4) = (u::nat)"; +test "(2*n*m) + (3*(m*n)) = (u::nat)"; +(*negative numerals: FAIL*) +test "Suc (i + j + -3 + k) = u"; + +(*cancel_numeral_factors*) +test "9*x = 12 * (y::nat)"; +test "(9*x) div (12 * (y::nat)) = z"; +test "9*x < 12 * (y::nat)"; +test "9*x <= 12 * (y::nat)"; + +(*cancel_factor*) +test "x*k = k*(y::nat)"; +test "k = k*(y::nat)"; +test "a*(b*c) = (b::nat)"; +test "a*(b*c) = d*(b::nat)*(x*a)"; + +test "x*k < k*(y::nat)"; +test "k < k*(y::nat)"; +test "a*(b*c) < (b::nat)"; +test "a*(b*c) < d*(b::nat)*(x*a)"; + +test "x*k <= k*(y::nat)"; +test "k <= k*(y::nat)"; +test "a*(b*c) <= (b::nat)"; +test "a*(b*c) <= d*(b::nat)*(x*a)"; + +test "(x*k) div (k*(y::nat)) = (uu::nat)"; +test "(k) div (k*(y::nat)) = (uu::nat)"; +test "(a*(b*c)) div ((b::nat)) = (uu::nat)"; +test "(a*(b*c)) div (d*(b::nat)*(x*a)) = (uu::nat)"; +*) + + +(*** Prepare linear arithmetic for nat numerals ***) + +local + +(* reduce contradictory <= to False *) +val add_rules = @{thms ring_distribs} @ + [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, @{thm nat_0}, @{thm nat_1}, + @{thm add_nat_number_of}, @{thm diff_nat_number_of}, @{thm mult_nat_number_of}, + @{thm eq_nat_number_of}, @{thm less_nat_number_of}, @{thm le_number_of_eq_not_less}, + @{thm le_Suc_number_of}, @{thm le_number_of_Suc}, + @{thm less_Suc_number_of}, @{thm less_number_of_Suc}, + @{thm Suc_eq_number_of}, @{thm eq_number_of_Suc}, + @{thm mult_Suc}, @{thm mult_Suc_right}, + @{thm add_Suc}, @{thm add_Suc_right}, + @{thm eq_number_of_0}, @{thm eq_0_number_of}, @{thm less_0_number_of}, + @{thm of_int_number_of_eq}, @{thm of_nat_number_of_eq}, @{thm nat_number_of}, @{thm if_True}, @{thm if_False}]; + +(* Products are multiplied out during proof (re)construction via +ring_distribs. Ideally they should remain atomic. But that is +currently not possible because 1 is replaced by Suc 0, and then some +simprocs start to mess around with products like (n+1)*m. The rule +1 == Suc 0 is necessary for early parts of HOL where numerals and +simprocs are not yet available. But then it is difficult to remove +that rule later on, because it may find its way back in when theories +(and thus lin-arith simpsets) are merged. Otherwise one could turn the +rule around (Suc n = n+1) and see if that helps products being left +alone. *) + +val simprocs = Nat_Numeral_Simprocs.combine_numerals + :: Nat_Numeral_Simprocs.cancel_numerals; + +in + +val nat_simprocs_setup = + LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} => + {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, + inj_thms = inj_thms, lessD = lessD, neqE = neqE, + simpset = simpset addsimps add_rules + addsimprocs simprocs}); + +end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Tools/rat_arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/rat_arith.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,51 @@ +(* Title: HOL/Real/rat_arith.ML + ID: $Id$ + Author: Lawrence C Paulson + Copyright 2004 University of Cambridge + +Simprocs for common factor cancellation & Rational coefficient handling + +Instantiation of the generic linear arithmetic package for type rat. +*) + +local + +val simprocs = field_cancel_numeral_factors + +val simps = + [@{thm order_less_irrefl}, @{thm neg_less_iff_less}, @{thm True_implies_equals}, + read_instantiate @{context} [(("a", 0), "(number_of ?v)")] @{thm right_distrib}, + @{thm divide_1}, @{thm divide_zero_left}, + @{thm times_divide_eq_right}, @{thm times_divide_eq_left}, + @{thm minus_divide_left} RS sym, @{thm minus_divide_right} RS sym, + @{thm of_int_0}, @{thm of_int_1}, @{thm of_int_add}, + @{thm of_int_minus}, @{thm of_int_diff}, + @{thm of_int_mult}, @{thm of_int_of_nat_eq}] + +val nat_inj_thms = [@{thm of_nat_le_iff} RS iffD2, + @{thm of_nat_eq_iff} RS iffD2] +(* not needed because x < (y::nat) can be rewritten as Suc x <= y: + of_nat_less_iff RS iffD2 *) + +val int_inj_thms = [@{thm of_int_le_iff} RS iffD2, + @{thm of_int_eq_iff} RS iffD2] +(* not needed because x < (y::int) can be rewritten as x + 1 <= y: + of_int_less_iff RS iffD2 *) + +in + +val ratT = Type ("Rational.rat", []) + +val rat_arith_setup = + LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} => + {add_mono_thms = add_mono_thms, + mult_mono_thms = mult_mono_thms, + inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms, + lessD = lessD, (*Can't change lessD: the rats are dense!*) + neqE = neqE, + simpset = simpset addsimps simps + addsimprocs simprocs}) #> + arith_inj_const (@{const_name of_nat}, HOLogic.natT --> ratT) #> + arith_inj_const (@{const_name of_int}, HOLogic.intT --> ratT) + +end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Tools/real_arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/real_arith.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,42 @@ +(* Title: HOL/Real/real_arith.ML + ID: $Id$ + Author: Tobias Nipkow, TU Muenchen + Copyright 1999 TU Muenchen + +Simprocs for common factor cancellation & Rational coefficient handling + +Instantiation of the generic linear arithmetic package for type real. +*) + +local + +val simps = [@{thm real_of_nat_zero}, @{thm real_of_nat_Suc}, @{thm real_of_nat_add}, + @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one}, + @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff}, + @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq}, + @{thm real_of_nat_number_of}, @{thm real_number_of}] + +val nat_inj_thms = [@{thm real_of_nat_le_iff} RS iffD2, + @{thm real_of_nat_inject} RS iffD2] +(* not needed because x < (y::nat) can be rewritten as Suc x <= y: + real_of_nat_less_iff RS iffD2 *) + +val int_inj_thms = [@{thm real_of_int_le_iff} RS iffD2, + @{thm real_of_int_inject} RS iffD2] +(* not needed because x < (y::int) can be rewritten as x + 1 <= y: + real_of_int_less_iff RS iffD2 *) + +in + +val real_arith_setup = + LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} => + {add_mono_thms = add_mono_thms, + mult_mono_thms = mult_mono_thms, + inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms, + lessD = lessD, (*Can't change lessD: the reals are dense!*) + neqE = neqE, + simpset = simpset addsimps simps}) #> + arith_inj_const (@{const_name real}, HOLogic.natT --> HOLogic.realT) #> + arith_inj_const (@{const_name real}, HOLogic.intT --> HOLogic.realT) + +end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Tools/simpdata.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/simpdata.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,200 @@ +(* Title: HOL/simpdata.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +Instantiation of the generic simplifier for HOL. +*) + +(** tools setup **) + +structure Quantifier1 = Quantifier1Fun +(struct + (*abstract syntax*) + fun dest_eq ((c as Const("op =",_)) $ s $ t) = SOME (c, s, t) + | dest_eq _ = NONE; + fun dest_conj ((c as Const("op &",_)) $ s $ t) = SOME (c, s, t) + | dest_conj _ = NONE; + fun dest_imp ((c as Const("op -->",_)) $ s $ t) = SOME (c, s, t) + | dest_imp _ = NONE; + val conj = HOLogic.conj + val imp = HOLogic.imp + (*rules*) + val iff_reflection = @{thm eq_reflection} + val iffI = @{thm iffI} + val iff_trans = @{thm trans} + val conjI= @{thm conjI} + val conjE= @{thm conjE} + val impI = @{thm impI} + val mp = @{thm mp} + val uncurry = @{thm uncurry} + val exI = @{thm exI} + val exE = @{thm exE} + val iff_allI = @{thm iff_allI} + val iff_exI = @{thm iff_exI} + val all_comm = @{thm all_comm} + val ex_comm = @{thm ex_comm} +end); + +structure Simpdata = +struct + +fun mk_meta_eq r = r RS @{thm eq_reflection}; +fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r; + +fun mk_eq th = case concl_of th + (*expects Trueprop if not == *) + of Const ("==",_) $ _ $ _ => th + | _ $ (Const ("op =", _) $ _ $ _) => mk_meta_eq th + | _ $ (Const ("Not", _) $ _) => th RS @{thm Eq_FalseI} + | _ => th RS @{thm Eq_TrueI} + +fun mk_eq_True r = + SOME (r RS @{thm meta_eq_to_obj_eq} RS @{thm Eq_TrueI}) handle Thm.THM _ => NONE; + +(* Produce theorems of the form + (P1 =simp=> ... =simp=> Pn => x == y) ==> (P1 =simp=> ... =simp=> Pn => x = y) +*) + +fun lift_meta_eq_to_obj_eq i st = + let + fun count_imp (Const ("HOL.simp_implies", _) $ P $ Q) = 1 + count_imp Q + | count_imp _ = 0; + val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1))) + in if j = 0 then @{thm meta_eq_to_obj_eq} + else + let + val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j); + fun mk_simp_implies Q = foldr (fn (R, S) => + Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Q Ps + val aT = TFree ("'a", HOLogic.typeS); + val x = Free ("x", aT); + val y = Free ("y", aT) + in Goal.prove_global (Thm.theory_of_thm st) [] + [mk_simp_implies (Logic.mk_equals (x, y))] + (mk_simp_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, y)))) + (fn {prems, ...} => EVERY + [rewrite_goals_tac @{thms simp_implies_def}, + REPEAT (ares_tac (@{thm meta_eq_to_obj_eq} :: + map (rewrite_rule @{thms simp_implies_def}) prems) 1)]) + end + end; + +(*Congruence rules for = (instead of ==)*) +fun mk_meta_cong rl = zero_var_indexes + (let val rl' = Seq.hd (TRYALL (fn i => fn st => + rtac (lift_meta_eq_to_obj_eq i st) i st) rl) + in mk_meta_eq rl' handle THM _ => + if can Logic.dest_equals (concl_of rl') then rl' + else error "Conclusion of congruence rules must be =-equality" + end); + +fun mk_atomize pairs = + let + fun atoms thm = + let + fun res th = map (fn rl => th RS rl); (*exception THM*) + fun res_fixed rls = + if Thm.maxidx_of (Thm.adjust_maxidx_thm ~1 thm) = ~1 then res thm rls + else Variable.trade (K (fn [thm'] => res thm' rls)) (Variable.thm_context thm) [thm]; + in + case concl_of thm + of Const ("Trueprop", _) $ p => (case head_of p + of Const (a, _) => (case AList.lookup (op =) pairs a + of SOME rls => (maps atoms (res_fixed rls) handle THM _ => [thm]) + | NONE => [thm]) + | _ => [thm]) + | _ => [thm] + end; + in atoms end; + +fun mksimps pairs = + map_filter (try mk_eq) o mk_atomize pairs o gen_all; + +fun unsafe_solver_tac prems = + (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' + FIRST' [resolve_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: prems), atac, + etac @{thm FalseE}]; + +val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac; + + +(*No premature instantiation of variables during simplification*) +fun safe_solver_tac prems = + (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' + FIRST' [match_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: prems), + eq_assume_tac, ematch_tac @{thms FalseE}]; + +val safe_solver = mk_solver "HOL safe" safe_solver_tac; + +structure SplitterData = +struct + structure Simplifier = Simplifier + val mk_eq = mk_eq + val meta_eq_to_iff = @{thm meta_eq_to_obj_eq} + val iffD = @{thm iffD2} + val disjE = @{thm disjE} + val conjE = @{thm conjE} + val exE = @{thm exE} + val contrapos = @{thm contrapos_nn} + val contrapos2 = @{thm contrapos_pp} + val notnotD = @{thm notnotD} +end; + +structure Splitter = SplitterFun(SplitterData); + +val split_tac = Splitter.split_tac; +val split_inside_tac = Splitter.split_inside_tac; + +val op addsplits = Splitter.addsplits; +val op delsplits = Splitter.delsplits; +val Addsplits = Splitter.Addsplits; +val Delsplits = Splitter.Delsplits; + + +(* integration of simplifier with classical reasoner *) + +structure Clasimp = ClasimpFun + (structure Simplifier = Simplifier and Splitter = Splitter + and Classical = Classical and Blast = Blast + val iffD1 = @{thm iffD1} val iffD2 = @{thm iffD2} val notE = @{thm notE}); +open Clasimp; + +val _ = ML_Antiquote.value "clasimpset" + (Scan.succeed "Clasimp.local_clasimpset_of (ML_Context.the_local_context ())"); + +val mksimps_pairs = + [("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]), + ("All", [@{thm spec}]), ("True", []), ("False", []), + ("HOL.If", [@{thm if_bool_eq_conj} RS @{thm iffD1}])]; + +val HOL_basic_ss = + Simplifier.theory_context (the_context ()) empty_ss + setsubgoaler asm_simp_tac + setSSolver safe_solver + setSolver unsafe_solver + setmksimps (mksimps mksimps_pairs) + setmkeqTrue mk_eq_True + setmkcong mk_meta_cong; + +fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews); + +fun unfold_tac ths = + let val ss0 = Simplifier.clear_ss HOL_basic_ss addsimps ths + in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; + +val defALL_regroup = + Simplifier.simproc (the_context ()) + "defined ALL" ["ALL x. P x"] Quantifier1.rearrange_all; + +val defEX_regroup = + Simplifier.simproc (the_context ()) + "defined EX" ["EX x. P x"] Quantifier1.rearrange_ex; + + +val simpset_simprocs = HOL_basic_ss addsimprocs [defALL_regroup, defEX_regroup] + +end; + +structure Splitter = Simpdata.Splitter; +structure Clasimp = Simpdata.Clasimp; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Transcendental.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Transcendental.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,2249 @@ +(* Title : Transcendental.thy + Author : Jacques D. Fleuriot + Copyright : 1998,1999 University of Cambridge + 1999,2001 University of Edinburgh + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 +*) + +header{*Power Series, Transcendental Functions etc.*} + +theory Transcendental +imports Fact Series Deriv NthRoot +begin + +subsection{*Properties of Power Series*} + +lemma lemma_realpow_diff: + fixes y :: "'a::recpower" + shows "p \ n \ y ^ (Suc n - p) = (y ^ (n - p)) * y" +proof - + assume "p \ n" + hence "Suc n - p = Suc (n - p)" by (rule Suc_diff_le) + thus ?thesis by (simp add: power_Suc power_commutes) +qed + +lemma lemma_realpow_diff_sumr: + fixes y :: "'a::{recpower,comm_semiring_0}" shows + "(\p=0..p=0..p=0..p=0..p=0..z\ < \x\"}.*} + +lemma powser_insidea: + fixes x z :: "'a::{real_normed_field,banach,recpower}" + assumes 1: "summable (\n. f n * x ^ n)" + assumes 2: "norm z < norm x" + shows "summable (\n. norm (f n * z ^ n))" +proof - + from 2 have x_neq_0: "x \ 0" by clarsimp + from 1 have "(\n. f n * x ^ n) ----> 0" + by (rule summable_LIMSEQ_zero) + hence "convergent (\n. f n * x ^ n)" + by (rule convergentI) + hence "Cauchy (\n. f n * x ^ n)" + by (simp add: Cauchy_convergent_iff) + hence "Bseq (\n. f n * x ^ n)" + by (rule Cauchy_Bseq) + then obtain K where 3: "0 < K" and 4: "\n. norm (f n * x ^ n) \ K" + by (simp add: Bseq_def, safe) + have "\N. \n\N. norm (norm (f n * z ^ n)) \ + K * norm (z ^ n) * inverse (norm (x ^ n))" + proof (intro exI allI impI) + fix n::nat assume "0 \ n" + have "norm (norm (f n * z ^ n)) * norm (x ^ n) = + norm (f n * x ^ n) * norm (z ^ n)" + by (simp add: norm_mult abs_mult) + also have "\ \ K * norm (z ^ n)" + by (simp only: mult_right_mono 4 norm_ge_zero) + also have "\ = K * norm (z ^ n) * (inverse (norm (x ^ n)) * norm (x ^ n))" + by (simp add: x_neq_0) + also have "\ = K * norm (z ^ n) * inverse (norm (x ^ n)) * norm (x ^ n)" + by (simp only: mult_assoc) + finally show "norm (norm (f n * z ^ n)) \ + K * norm (z ^ n) * inverse (norm (x ^ n))" + by (simp add: mult_le_cancel_right x_neq_0) + qed + moreover have "summable (\n. K * norm (z ^ n) * inverse (norm (x ^ n)))" + proof - + from 2 have "norm (norm (z * inverse x)) < 1" + using x_neq_0 + by (simp add: nonzero_norm_divide divide_inverse [symmetric]) + hence "summable (\n. norm (z * inverse x) ^ n)" + by (rule summable_geometric) + hence "summable (\n. K * norm (z * inverse x) ^ n)" + by (rule summable_mult) + thus "summable (\n. K * norm (z ^ n) * inverse (norm (x ^ n)))" + using x_neq_0 + by (simp add: norm_mult nonzero_norm_inverse power_mult_distrib + power_inverse norm_power mult_assoc) + qed + ultimately show "summable (\n. norm (f n * z ^ n))" + by (rule summable_comparison_test) +qed + +lemma powser_inside: + fixes f :: "nat \ 'a::{real_normed_field,banach,recpower}" shows + "[| summable (%n. f(n) * (x ^ n)); norm z < norm x |] + ==> summable (%n. f(n) * (z ^ n))" +by (rule powser_insidea [THEN summable_norm_cancel]) + + +subsection{*Term-by-Term Differentiability of Power Series*} + +definition + diffs :: "(nat => 'a::ring_1) => nat => 'a" where + "diffs c = (%n. of_nat (Suc n) * c(Suc n))" + +text{*Lemma about distributing negation over it*} +lemma diffs_minus: "diffs (%n. - c n) = (%n. - diffs c n)" +by (simp add: diffs_def) + +text{*Show that we can shift the terms down one*} +lemma lemma_diffs: + "(\n=0..n=0..n=0..n=0.. + (%n. of_nat n * c(n) * (x ^ (n - Suc 0))) sums + (\n. (diffs c)(n) * (x ^ n))" +apply (subgoal_tac " (%n. of_nat n * c (n) * (x ^ (n - Suc 0))) ----> 0") +apply (rule_tac [2] LIMSEQ_imp_Suc) +apply (drule summable_sums) +apply (auto simp add: sums_def) +apply (drule_tac X="(\n. \n = 0..p=0..p=0.. (\d. n = m + d + Suc 0)" +by (simp add: less_iff_Suc_add) + +lemma sumdiff: "a + b - (c + d) = a - c + b - (d::real)" +by arith + +lemma sumr_diff_mult_const2: + "setsum f {0..i = 0.. 0" shows + "((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0) = + h * (\p=0..< n - Suc 0. \q=0..< n - Suc 0 - p. + (z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs") +apply (subgoal_tac "h * ?lhs = h * ?rhs", simp add: h) +apply (simp add: right_diff_distrib diff_divide_distrib h) +apply (simp add: mult_assoc [symmetric]) +apply (cases "n", simp) +apply (simp add: lemma_realpow_diff_sumr2 h + right_diff_distrib [symmetric] mult_assoc + del: realpow_Suc setsum_op_ivl_Suc of_nat_Suc) +apply (subst lemma_realpow_rev_sumr) +apply (subst sumr_diff_mult_const2) +apply simp +apply (simp only: lemma_termdiff1 setsum_right_distrib) +apply (rule setsum_cong [OF refl]) +apply (simp add: diff_minus [symmetric] less_iff_Suc_add) +apply (clarify) +apply (simp add: setsum_right_distrib lemma_realpow_diff_sumr2 mult_ac + del: setsum_op_ivl_Suc realpow_Suc) +apply (subst mult_assoc [symmetric], subst power_add [symmetric]) +apply (simp add: mult_ac) +done + +lemma real_setsum_nat_ivl_bounded2: + fixes K :: "'a::ordered_semidom" + assumes f: "\p::nat. p < n \ f p \ K" + assumes K: "0 \ K" + shows "setsum f {0.. of_nat n * K" +apply (rule order_trans [OF setsum_mono]) +apply (rule f, simp) +apply (simp add: mult_right_mono K) +done + +lemma lemma_termdiff3: + fixes h z :: "'a::{real_normed_field,recpower}" + assumes 1: "h \ 0" + assumes 2: "norm z \ K" + assumes 3: "norm (z + h) \ K" + shows "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) + \ of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" +proof - + have "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) = + norm (\p = 0..q = 0.. \ of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2)) * norm h" + proof (rule mult_right_mono [OF _ norm_ge_zero]) + from norm_ge_zero 2 have K: "0 \ K" by (rule order_trans) + have le_Kn: "\i j n. i + j = n \ norm ((z + h) ^ i * z ^ j) \ K ^ n" + apply (erule subst) + apply (simp only: norm_mult norm_power power_add) + apply (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K) + done + show "norm (\p = 0..q = 0.. of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2))" + apply (intro + order_trans [OF norm_setsum] + real_setsum_nat_ivl_bounded2 + mult_nonneg_nonneg + zero_le_imp_of_nat + zero_le_power K) + apply (rule le_Kn, simp) + done + qed + also have "\ = of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" + by (simp only: mult_assoc) + finally show ?thesis . +qed + +lemma lemma_termdiff4: + fixes f :: "'a::{real_normed_field,recpower} \ + 'b::real_normed_vector" + assumes k: "0 < (k::real)" + assumes le: "\h. \h \ 0; norm h < k\ \ norm (f h) \ K * norm h" + shows "f -- 0 --> 0" +proof (simp add: LIM_def, safe) + fix r::real assume r: "0 < r" + have zero_le_K: "0 \ K" + apply (cut_tac k) + apply (cut_tac h="of_real (k/2)" in le, simp) + apply (simp del: of_real_divide) + apply (drule order_trans [OF norm_ge_zero]) + apply (simp add: zero_le_mult_iff) + done + show "\s. 0 < s \ (\x. x \ 0 \ norm x < s \ norm (f x) < r)" + proof (cases) + assume "K = 0" + with k r le have "0 < k \ (\x. x \ 0 \ norm x < k \ norm (f x) < r)" + by simp + thus "\s. 0 < s \ (\x. x \ 0 \ norm x < s \ norm (f x) < r)" .. + next + assume K_neq_zero: "K \ 0" + with zero_le_K have K: "0 < K" by simp + show "\s. 0 < s \ (\x. x \ 0 \ norm x < s \ norm (f x) < r)" + proof (rule exI, safe) + from k r K show "0 < min k (r * inverse K / 2)" + by (simp add: mult_pos_pos positive_imp_inverse_positive) + next + fix x::'a + assume x1: "x \ 0" and x2: "norm x < min k (r * inverse K / 2)" + from x2 have x3: "norm x < k" and x4: "norm x < r * inverse K / 2" + by simp_all + from x1 x3 le have "norm (f x) \ K * norm x" by simp + also from x4 K have "K * norm x < K * (r * inverse K / 2)" + by (rule mult_strict_left_mono) + also have "\ = r / 2" + using K_neq_zero by simp + also have "r / 2 < r" + using r by simp + finally show "norm (f x) < r" . + qed + qed +qed + +lemma lemma_termdiff5: + fixes g :: "'a::{recpower,real_normed_field} \ + nat \ 'b::banach" + assumes k: "0 < (k::real)" + assumes f: "summable f" + assumes le: "\h n. \h \ 0; norm h < k\ \ norm (g h n) \ f n * norm h" + shows "(\h. suminf (g h)) -- 0 --> 0" +proof (rule lemma_termdiff4 [OF k]) + fix h::'a assume "h \ 0" and "norm h < k" + hence A: "\n. norm (g h n) \ f n * norm h" + by (simp add: le) + hence "\N. \n\N. norm (norm (g h n)) \ f n * norm h" + by simp + moreover from f have B: "summable (\n. f n * norm h)" + by (rule summable_mult2) + ultimately have C: "summable (\n. norm (g h n))" + by (rule summable_comparison_test) + hence "norm (suminf (g h)) \ (\n. norm (g h n))" + by (rule summable_norm) + also from A C B have "(\n. norm (g h n)) \ (\n. f n * norm h)" + by (rule summable_le) + also from f have "(\n. f n * norm h) = suminf f * norm h" + by (rule suminf_mult2 [symmetric]) + finally show "norm (suminf (g h)) \ suminf f * norm h" . +qed + + +text{* FIXME: Long proofs*} + +lemma termdiffs_aux: + fixes x :: "'a::{recpower,real_normed_field,banach}" + assumes 1: "summable (\n. diffs (diffs c) n * K ^ n)" + assumes 2: "norm x < norm K" + shows "(\h. \n. c n * (((x + h) ^ n - x ^ n) / h + - of_nat n * x ^ (n - Suc 0))) -- 0 --> 0" +proof - + from dense [OF 2] + obtain r where r1: "norm x < r" and r2: "r < norm K" by fast + from norm_ge_zero r1 have r: "0 < r" + by (rule order_le_less_trans) + hence r_neq_0: "r \ 0" by simp + show ?thesis + proof (rule lemma_termdiff5) + show "0 < r - norm x" using r1 by simp + next + from r r2 have "norm (of_real r::'a) < norm K" + by simp + with 1 have "summable (\n. norm (diffs (diffs c) n * (of_real r ^ n)))" + by (rule powser_insidea) + hence "summable (\n. diffs (diffs (\n. norm (c n))) n * r ^ n)" + using r + by (simp add: diffs_def norm_mult norm_power del: of_nat_Suc) + hence "summable (\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0))" + by (rule diffs_equiv [THEN sums_summable]) + also have "(\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0)) + = (\n. diffs (%m. of_nat (m - Suc 0) * norm (c m) * inverse r) n * (r ^ n))" + apply (rule ext) + apply (simp add: diffs_def) + apply (case_tac n, simp_all add: r_neq_0) + done + finally have "summable + (\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0))" + by (rule diffs_equiv [THEN sums_summable]) + also have + "(\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * + r ^ (n - Suc 0)) = + (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" + apply (rule ext) + apply (case_tac "n", simp) + apply (case_tac "nat", simp) + apply (simp add: r_neq_0) + done + finally show + "summable (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" . + next + fix h::'a and n::nat + assume h: "h \ 0" + assume "norm h < r - norm x" + hence "norm x + norm h < r" by simp + with norm_triangle_ineq have xh: "norm (x + h) < r" + by (rule order_le_less_trans) + show "norm (c n * (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0))) + \ norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2) * norm h" + apply (simp only: norm_mult mult_assoc) + apply (rule mult_left_mono [OF _ norm_ge_zero]) + apply (simp (no_asm) add: mult_assoc [symmetric]) + apply (rule lemma_termdiff3) + apply (rule h) + apply (rule r1 [THEN order_less_imp_le]) + apply (rule xh [THEN order_less_imp_le]) + done + qed +qed + +lemma termdiffs: + fixes K x :: "'a::{recpower,real_normed_field,banach}" + assumes 1: "summable (\n. c n * K ^ n)" + assumes 2: "summable (\n. (diffs c) n * K ^ n)" + assumes 3: "summable (\n. (diffs (diffs c)) n * K ^ n)" + assumes 4: "norm x < norm K" + shows "DERIV (\x. \n. c n * x ^ n) x :> (\n. (diffs c) n * x ^ n)" +proof (simp add: deriv_def, rule LIM_zero_cancel) + show "(\h. (suminf (\n. c n * (x + h) ^ n) - suminf (\n. c n * x ^ n)) / h + - suminf (\n. diffs c n * x ^ n)) -- 0 --> 0" + proof (rule LIM_equal2) + show "0 < norm K - norm x" by (simp add: less_diff_eq 4) + next + fix h :: 'a + assume "h \ 0" + assume "norm (h - 0) < norm K - norm x" + hence "norm x + norm h < norm K" by simp + hence 5: "norm (x + h) < norm K" + by (rule norm_triangle_ineq [THEN order_le_less_trans]) + have A: "summable (\n. c n * x ^ n)" + by (rule powser_inside [OF 1 4]) + have B: "summable (\n. c n * (x + h) ^ n)" + by (rule powser_inside [OF 1 5]) + have C: "summable (\n. diffs c n * x ^ n)" + by (rule powser_inside [OF 2 4]) + show "((\n. c n * (x + h) ^ n) - (\n. c n * x ^ n)) / h + - (\n. diffs c n * x ^ n) = + (\n. c n * (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0)))" + apply (subst sums_unique [OF diffs_equiv [OF C]]) + apply (subst suminf_diff [OF B A]) + apply (subst suminf_divide [symmetric]) + apply (rule summable_diff [OF B A]) + apply (subst suminf_diff) + apply (rule summable_divide) + apply (rule summable_diff [OF B A]) + apply (rule sums_summable [OF diffs_equiv [OF C]]) + apply (rule_tac f="suminf" in arg_cong) + apply (rule ext) + apply (simp add: ring_simps) + done + next + show "(\h. \n. c n * (((x + h) ^ n - x ^ n) / h - + of_nat n * x ^ (n - Suc 0))) -- 0 --> 0" + by (rule termdiffs_aux [OF 3 4]) + qed +qed + + +subsection{*Exponential Function*} + +definition + exp :: "'a \ 'a::{recpower,real_normed_field,banach}" where + "exp x = (\n. x ^ n /\<^sub>R real (fact n))" + +definition + sin :: "real => real" where + "sin x = (\n. (if even(n) then 0 else + (-1 ^ ((n - Suc 0) div 2))/(real (fact n))) * x ^ n)" + +definition + cos :: "real => real" where + "cos x = (\n. (if even(n) then (-1 ^ (n div 2))/(real (fact n)) + else 0) * x ^ n)" + +lemma summable_exp_generic: + fixes x :: "'a::{real_normed_algebra_1,recpower,banach}" + defines S_def: "S \ \n. x ^ n /\<^sub>R real (fact n)" + shows "summable S" +proof - + have S_Suc: "\n. S (Suc n) = (x * S n) /\<^sub>R real (Suc n)" + unfolding S_def by (simp add: power_Suc del: mult_Suc) + obtain r :: real where r0: "0 < r" and r1: "r < 1" + using dense [OF zero_less_one] by fast + obtain N :: nat where N: "norm x < real N * r" + using reals_Archimedean3 [OF r0] by fast + from r1 show ?thesis + proof (rule ratio_test [rule_format]) + fix n :: nat + assume n: "N \ n" + have "norm x \ real N * r" + using N by (rule order_less_imp_le) + also have "real N * r \ real (Suc n) * r" + using r0 n by (simp add: mult_right_mono) + finally have "norm x * norm (S n) \ real (Suc n) * r * norm (S n)" + using norm_ge_zero by (rule mult_right_mono) + hence "norm (x * S n) \ real (Suc n) * r * norm (S n)" + by (rule order_trans [OF norm_mult_ineq]) + hence "norm (x * S n) / real (Suc n) \ r * norm (S n)" + by (simp add: pos_divide_le_eq mult_ac) + thus "norm (S (Suc n)) \ r * norm (S n)" + by (simp add: S_Suc norm_scaleR inverse_eq_divide) + qed +qed + +lemma summable_norm_exp: + fixes x :: "'a::{real_normed_algebra_1,recpower,banach}" + shows "summable (\n. norm (x ^ n /\<^sub>R real (fact n)))" +proof (rule summable_norm_comparison_test [OF exI, rule_format]) + show "summable (\n. norm x ^ n /\<^sub>R real (fact n))" + by (rule summable_exp_generic) +next + fix n show "norm (x ^ n /\<^sub>R real (fact n)) \ norm x ^ n /\<^sub>R real (fact n)" + by (simp add: norm_scaleR norm_power_ineq) +qed + +lemma summable_exp: "summable (%n. inverse (real (fact n)) * x ^ n)" +by (insert summable_exp_generic [where x=x], simp) + +lemma summable_sin: + "summable (%n. + (if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * + x ^ n)" +apply (rule_tac g = "(%n. inverse (real (fact n)) * \x\ ^ n)" in summable_comparison_test) +apply (rule_tac [2] summable_exp) +apply (rule_tac x = 0 in exI) +apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) +done + +lemma summable_cos: + "summable (%n. + (if even n then + -1 ^ (n div 2)/(real (fact n)) else 0) * x ^ n)" +apply (rule_tac g = "(%n. inverse (real (fact n)) * \x\ ^ n)" in summable_comparison_test) +apply (rule_tac [2] summable_exp) +apply (rule_tac x = 0 in exI) +apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) +done + +lemma lemma_STAR_sin: + "(if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0" +by (induct "n", auto) + +lemma lemma_STAR_cos: + "0 < n --> + -1 ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" +by (induct "n", auto) + +lemma lemma_STAR_cos1: + "0 < n --> + (-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" +by (induct "n", auto) + +lemma lemma_STAR_cos2: + "(\n=1..n. x ^ n /\<^sub>R real (fact n)) sums exp x" +unfolding exp_def by (rule summable_exp_generic [THEN summable_sums]) + +lemma sin_converges: + "(%n. (if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * + x ^ n) sums sin(x)" +unfolding sin_def by (rule summable_sin [THEN summable_sums]) + +lemma cos_converges: + "(%n. (if even n then + -1 ^ (n div 2)/(real (fact n)) + else 0) * x ^ n) sums cos(x)" +unfolding cos_def by (rule summable_cos [THEN summable_sums]) + + +subsection{*Formal Derivatives of Exp, Sin, and Cos Series*} + +lemma exp_fdiffs: + "diffs (%n. inverse(real (fact n))) = (%n. inverse(real (fact n)))" +by (simp add: diffs_def mult_assoc [symmetric] real_of_nat_def of_nat_mult + del: mult_Suc of_nat_Suc) + +lemma diffs_of_real: "diffs (\n. of_real (f n)) = (\n. of_real (diffs f n))" +by (simp add: diffs_def) + +lemma sin_fdiffs: + "diffs(%n. if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) + = (%n. if even n then + -1 ^ (n div 2)/(real (fact n)) + else 0)" +by (auto intro!: ext + simp add: diffs_def divide_inverse real_of_nat_def of_nat_mult + simp del: mult_Suc of_nat_Suc) + +lemma sin_fdiffs2: + "diffs(%n. if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) n + = (if even n then + -1 ^ (n div 2)/(real (fact n)) + else 0)" +by (simp only: sin_fdiffs) + +lemma cos_fdiffs: + "diffs(%n. if even n then + -1 ^ (n div 2)/(real (fact n)) else 0) + = (%n. - (if even n then 0 + else -1 ^ ((n - Suc 0)div 2)/(real (fact n))))" +by (auto intro!: ext + simp add: diffs_def divide_inverse odd_Suc_mult_two_ex real_of_nat_def of_nat_mult + simp del: mult_Suc of_nat_Suc) + + +lemma cos_fdiffs2: + "diffs(%n. if even n then + -1 ^ (n div 2)/(real (fact n)) else 0) n + = - (if even n then 0 + else -1 ^ ((n - Suc 0)div 2)/(real (fact n)))" +by (simp only: cos_fdiffs) + +text{*Now at last we can get the derivatives of exp, sin and cos*} + +lemma lemma_sin_minus: + "- sin x = (\n. - ((if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * x ^ n))" +by (auto intro!: sums_unique sums_minus sin_converges) + +lemma lemma_exp_ext: "exp = (\x. \n. x ^ n /\<^sub>R real (fact n))" +by (auto intro!: ext simp add: exp_def) + +lemma DERIV_exp [simp]: "DERIV exp x :> exp(x)" +apply (simp add: exp_def) +apply (subst lemma_exp_ext) +apply (subgoal_tac "DERIV (\u. \n. of_real (inverse (real (fact n))) * u ^ n) x :> (\n. diffs (\n. of_real (inverse (real (fact n)))) n * x ^ n)") +apply (rule_tac [2] K = "of_real (1 + norm x)" in termdiffs) +apply (simp_all only: diffs_of_real scaleR_conv_of_real exp_fdiffs) +apply (rule exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real])+ +apply (simp del: of_real_add) +done + +lemma lemma_sin_ext: + "sin = (%x. \n. + (if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * + x ^ n)" +by (auto intro!: ext simp add: sin_def) + +lemma lemma_cos_ext: + "cos = (%x. \n. + (if even n then -1 ^ (n div 2)/(real (fact n)) else 0) * + x ^ n)" +by (auto intro!: ext simp add: cos_def) + +lemma DERIV_sin [simp]: "DERIV sin x :> cos(x)" +apply (simp add: cos_def) +apply (subst lemma_sin_ext) +apply (auto simp add: sin_fdiffs2 [symmetric]) +apply (rule_tac K = "1 + \x\" in termdiffs) +apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs) +done + +lemma DERIV_cos [simp]: "DERIV cos x :> -sin(x)" +apply (subst lemma_cos_ext) +apply (auto simp add: lemma_sin_minus cos_fdiffs2 [symmetric] minus_mult_left) +apply (rule_tac K = "1 + \x\" in termdiffs) +apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs diffs_minus) +done + +lemma isCont_exp [simp]: "isCont exp x" +by (rule DERIV_exp [THEN DERIV_isCont]) + +lemma isCont_sin [simp]: "isCont sin x" +by (rule DERIV_sin [THEN DERIV_isCont]) + +lemma isCont_cos [simp]: "isCont cos x" +by (rule DERIV_cos [THEN DERIV_isCont]) + + +subsection{*Properties of the Exponential Function*} + +lemma powser_zero: + fixes f :: "nat \ 'a::{real_normed_algebra_1,recpower}" + shows "(\n. f n * 0 ^ n) = f 0" +proof - + have "(\n = 0..<1. f n * 0 ^ n) = (\n. f n * 0 ^ n)" + by (rule sums_unique [OF series_zero], simp add: power_0_left) + thus ?thesis by simp +qed + +lemma exp_zero [simp]: "exp 0 = 1" +unfolding exp_def by (simp add: scaleR_conv_of_real powser_zero) + +lemma setsum_cl_ivl_Suc2: + "(\i=m..Suc n. f i) = (if Suc n < m then 0 else f m + (\i=m..n. f (Suc i)))" +by (simp add: setsum_head_Suc setsum_shift_bounds_cl_Suc_ivl + del: setsum_cl_ivl_Suc) + +lemma exp_series_add: + fixes x y :: "'a::{real_field,recpower}" + defines S_def: "S \ \x n. x ^ n /\<^sub>R real (fact n)" + shows "S (x + y) n = (\i=0..n. S x i * S y (n - i))" +proof (induct n) + case 0 + show ?case + unfolding S_def by simp +next + case (Suc n) + have S_Suc: "\x n. S x (Suc n) = (x * S x n) /\<^sub>R real (Suc n)" + unfolding S_def by (simp add: power_Suc del: mult_Suc) + hence times_S: "\x n. x * S x n = real (Suc n) *\<^sub>R S x (Suc n)" + by simp + + have "real (Suc n) *\<^sub>R S (x + y) (Suc n) = (x + y) * S (x + y) n" + by (simp only: times_S) + also have "\ = (x + y) * (\i=0..n. S x i * S y (n-i))" + by (simp only: Suc) + also have "\ = x * (\i=0..n. S x i * S y (n-i)) + + y * (\i=0..n. S x i * S y (n-i))" + by (rule left_distrib) + also have "\ = (\i=0..n. (x * S x i) * S y (n-i)) + + (\i=0..n. S x i * (y * S y (n-i)))" + by (simp only: setsum_right_distrib mult_ac) + also have "\ = (\i=0..n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n-i))) + + (\i=0..n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i)))" + by (simp add: times_S Suc_diff_le) + also have "(\i=0..n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n-i))) = + (\i=0..Suc n. real i *\<^sub>R (S x i * S y (Suc n-i)))" + by (subst setsum_cl_ivl_Suc2, simp) + also have "(\i=0..n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i))) = + (\i=0..Suc n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i)))" + by (subst setsum_cl_ivl_Suc, simp) + also have "(\i=0..Suc n. real i *\<^sub>R (S x i * S y (Suc n-i))) + + (\i=0..Suc n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i))) = + (\i=0..Suc n. real (Suc n) *\<^sub>R (S x i * S y (Suc n-i)))" + by (simp only: setsum_addf [symmetric] scaleR_left_distrib [symmetric] + real_of_nat_add [symmetric], simp) + also have "\ = real (Suc n) *\<^sub>R (\i=0..Suc n. S x i * S y (Suc n-i))" + by (simp only: scaleR_right.setsum) + finally show + "S (x + y) (Suc n) = (\i=0..Suc n. S x i * S y (Suc n - i))" + by (simp add: scaleR_cancel_left del: setsum_cl_ivl_Suc) +qed + +lemma exp_add: "exp (x + y) = exp x * exp y" +unfolding exp_def +by (simp only: Cauchy_product summable_norm_exp exp_series_add) + +lemma exp_of_real: "exp (of_real x) = of_real (exp x)" +unfolding exp_def +apply (subst of_real.suminf) +apply (rule summable_exp_generic) +apply (simp add: scaleR_conv_of_real) +done + +lemma exp_ge_add_one_self_aux: "0 \ (x::real) ==> (1 + x) \ exp(x)" +apply (drule order_le_imp_less_or_eq, auto) +apply (simp add: exp_def) +apply (rule real_le_trans) +apply (rule_tac [2] n = 2 and f = "(%n. inverse (real (fact n)) * x ^ n)" in series_pos_le) +apply (auto intro: summable_exp simp add: numeral_2_eq_2 zero_le_mult_iff) +done + +lemma exp_gt_one [simp]: "0 < (x::real) ==> 1 < exp x" +apply (rule order_less_le_trans) +apply (rule_tac [2] exp_ge_add_one_self_aux, auto) +done + +lemma DERIV_exp_add_const: "DERIV (%x. exp (x + y)) x :> exp(x + y)" +proof - + have "DERIV (exp \ (\x. x + y)) x :> exp (x + y) * (1+0)" + by (fast intro: DERIV_chain DERIV_add DERIV_exp DERIV_ident DERIV_const) + thus ?thesis by (simp add: o_def) +qed + +lemma DERIV_exp_minus [simp]: "DERIV (%x. exp (-x)) x :> - exp(-x)" +proof - + have "DERIV (exp \ uminus) x :> exp (- x) * - 1" + by (fast intro: DERIV_chain DERIV_minus DERIV_exp DERIV_ident) + thus ?thesis by (simp add: o_def) +qed + +lemma DERIV_exp_exp_zero [simp]: "DERIV (%x. exp (x + y) * exp (- x)) x :> 0" +proof - + have "DERIV (\x. exp (x + y) * exp (- x)) x + :> exp (x + y) * exp (- x) + - exp (- x) * exp (x + y)" + by (fast intro: DERIV_exp_add_const DERIV_exp_minus DERIV_mult) + thus ?thesis by (simp add: mult_commute) +qed + +lemma exp_add_mult_minus [simp]: "exp(x + y)*exp(-x) = exp(y::real)" +proof - + have "\x. DERIV (%x. exp (x + y) * exp (- x)) x :> 0" by simp + hence "exp (x + y) * exp (- x) = exp (0 + y) * exp (- 0)" + by (rule DERIV_isconst_all) + thus ?thesis by simp +qed + +lemma exp_mult_minus [simp]: "exp x * exp(-x) = 1" +by (simp add: exp_add [symmetric]) + +lemma exp_mult_minus2 [simp]: "exp(-x)*exp(x) = 1" +by (simp add: mult_commute) + + +lemma exp_minus: "exp(-x) = inverse(exp(x))" +by (auto intro: inverse_unique [symmetric]) + +text{*Proof: because every exponential can be seen as a square.*} +lemma exp_ge_zero [simp]: "0 \ exp (x::real)" +apply (rule_tac t = x in real_sum_of_halves [THEN subst]) +apply (subst exp_add, auto) +done + +lemma exp_not_eq_zero [simp]: "exp x \ 0" +apply (cut_tac x = x in exp_mult_minus2) +apply (auto simp del: exp_mult_minus2) +done + +lemma exp_gt_zero [simp]: "0 < exp (x::real)" +by (simp add: order_less_le) + +lemma inv_exp_gt_zero [simp]: "0 < inverse(exp x::real)" +by (auto intro: positive_imp_inverse_positive) + +lemma abs_exp_cancel [simp]: "\exp x::real\ = exp x" +by auto + +lemma exp_real_of_nat_mult: "exp(real n * x) = exp(x) ^ n" +apply (induct "n") +apply (auto simp add: real_of_nat_Suc right_distrib exp_add mult_commute) +done + +lemma exp_diff: "exp(x - y) = exp(x)/(exp y)" +apply (simp add: diff_minus divide_inverse) +apply (simp (no_asm) add: exp_add exp_minus) +done + + +lemma exp_less_mono: + fixes x y :: real + assumes xy: "x < y" shows "exp x < exp y" +proof - + from xy have "1 < exp (y + - x)" + by (rule real_less_sum_gt_zero [THEN exp_gt_one]) + hence "exp x * inverse (exp x) < exp y * inverse (exp x)" + by (auto simp add: exp_add exp_minus) + thus ?thesis + by (simp add: divide_inverse [symmetric] pos_less_divide_eq + del: divide_self_if) +qed + +lemma exp_less_cancel: "exp (x::real) < exp y ==> x < y" +apply (simp add: linorder_not_le [symmetric]) +apply (auto simp add: order_le_less exp_less_mono) +done + +lemma exp_less_cancel_iff [iff]: "(exp(x::real) < exp(y)) = (x < y)" +by (auto intro: exp_less_mono exp_less_cancel) + +lemma exp_le_cancel_iff [iff]: "(exp(x::real) \ exp(y)) = (x \ y)" +by (auto simp add: linorder_not_less [symmetric]) + +lemma exp_inj_iff [iff]: "(exp (x::real) = exp y) = (x = y)" +by (simp add: order_eq_iff) + +lemma lemma_exp_total: "1 \ y ==> \x. 0 \ x & x \ y - 1 & exp(x::real) = y" +apply (rule IVT) +apply (auto intro: isCont_exp simp add: le_diff_eq) +apply (subgoal_tac "1 + (y - 1) \ exp (y - 1)") +apply simp +apply (rule exp_ge_add_one_self_aux, simp) +done + +lemma exp_total: "0 < (y::real) ==> \x. exp x = y" +apply (rule_tac x = 1 and y = y in linorder_cases) +apply (drule order_less_imp_le [THEN lemma_exp_total]) +apply (rule_tac [2] x = 0 in exI) +apply (frule_tac [3] real_inverse_gt_one) +apply (drule_tac [4] order_less_imp_le [THEN lemma_exp_total], auto) +apply (rule_tac x = "-x" in exI) +apply (simp add: exp_minus) +done + + +subsection{*Properties of the Logarithmic Function*} + +definition + ln :: "real => real" where + "ln x = (THE u. exp u = x)" + +lemma ln_exp [simp]: "ln (exp x) = x" +by (simp add: ln_def) + +lemma exp_ln [simp]: "0 < x \ exp (ln x) = x" +by (auto dest: exp_total) + +lemma exp_ln_iff [simp]: "(exp (ln x) = x) = (0 < x)" +apply (auto dest: exp_total) +apply (erule subst, simp) +done + +lemma ln_mult: "[| 0 < x; 0 < y |] ==> ln(x * y) = ln(x) + ln(y)" +apply (rule exp_inj_iff [THEN iffD1]) +apply (simp add: exp_add exp_ln mult_pos_pos) +done + +lemma ln_inj_iff[simp]: "[| 0 < x; 0 < y |] ==> (ln x = ln y) = (x = y)" +apply (simp only: exp_ln_iff [symmetric]) +apply (erule subst)+ +apply simp +done + +lemma ln_one[simp]: "ln 1 = 0" +by (rule exp_inj_iff [THEN iffD1], auto) + +lemma ln_inverse: "0 < x ==> ln(inverse x) = - ln x" +apply (rule_tac a1 = "ln x" in add_left_cancel [THEN iffD1]) +apply (auto simp add: positive_imp_inverse_positive ln_mult [symmetric]) +done + +lemma ln_div: + "[|0 < x; 0 < y|] ==> ln(x/y) = ln x - ln y" +apply (simp add: divide_inverse) +apply (auto simp add: positive_imp_inverse_positive ln_mult ln_inverse) +done + +lemma ln_less_cancel_iff[simp]: "[| 0 < x; 0 < y|] ==> (ln x < ln y) = (x < y)" +apply (simp only: exp_ln_iff [symmetric]) +apply (erule subst)+ +apply simp +done + +lemma ln_le_cancel_iff[simp]: "[| 0 < x; 0 < y|] ==> (ln x \ ln y) = (x \ y)" +by (auto simp add: linorder_not_less [symmetric]) + +lemma ln_realpow: "0 < x ==> ln(x ^ n) = real n * ln(x)" +by (auto dest!: exp_total simp add: exp_real_of_nat_mult [symmetric]) + +lemma ln_add_one_self_le_self [simp]: "0 \ x ==> ln(1 + x) \ x" +apply (rule ln_exp [THEN subst]) +apply (rule ln_le_cancel_iff [THEN iffD2]) +apply (auto simp add: exp_ge_add_one_self_aux) +done + +lemma ln_less_self [simp]: "0 < x ==> ln x < x" +apply (rule order_less_le_trans) +apply (rule_tac [2] ln_add_one_self_le_self) +apply (rule ln_less_cancel_iff [THEN iffD2], auto) +done + +lemma ln_ge_zero [simp]: + assumes x: "1 \ x" shows "0 \ ln x" +proof - + have "0 < x" using x by arith + hence "exp 0 \ exp (ln x)" + by (simp add: x) + thus ?thesis by (simp only: exp_le_cancel_iff) +qed + +lemma ln_ge_zero_imp_ge_one: + assumes ln: "0 \ ln x" + and x: "0 < x" + shows "1 \ x" +proof - + from ln have "ln 1 \ ln x" by simp + thus ?thesis by (simp add: x del: ln_one) +qed + +lemma ln_ge_zero_iff [simp]: "0 < x ==> (0 \ ln x) = (1 \ x)" +by (blast intro: ln_ge_zero ln_ge_zero_imp_ge_one) + +lemma ln_less_zero_iff [simp]: "0 < x ==> (ln x < 0) = (x < 1)" +by (insert ln_ge_zero_iff [of x], arith) + +lemma ln_gt_zero: + assumes x: "1 < x" shows "0 < ln x" +proof - + have "0 < x" using x by arith + hence "exp 0 < exp (ln x)" by (simp add: x) + thus ?thesis by (simp only: exp_less_cancel_iff) +qed + +lemma ln_gt_zero_imp_gt_one: + assumes ln: "0 < ln x" + and x: "0 < x" + shows "1 < x" +proof - + from ln have "ln 1 < ln x" by simp + thus ?thesis by (simp add: x del: ln_one) +qed + +lemma ln_gt_zero_iff [simp]: "0 < x ==> (0 < ln x) = (1 < x)" +by (blast intro: ln_gt_zero ln_gt_zero_imp_gt_one) + +lemma ln_eq_zero_iff [simp]: "0 < x ==> (ln x = 0) = (x = 1)" +by (insert ln_less_zero_iff [of x] ln_gt_zero_iff [of x], arith) + +lemma ln_less_zero: "[| 0 < x; x < 1 |] ==> ln x < 0" +by simp + +lemma exp_ln_eq: "exp u = x ==> ln x = u" +by auto + +lemma isCont_ln: "0 < x \ isCont ln x" +apply (subgoal_tac "isCont ln (exp (ln x))", simp) +apply (rule isCont_inverse_function [where f=exp], simp_all) +done + +lemma DERIV_ln: "0 < x \ DERIV ln x :> inverse x" +apply (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"]) +apply (erule lemma_DERIV_subst [OF DERIV_exp exp_ln]) +apply (simp_all add: abs_if isCont_ln) +done + + +subsection{*Basic Properties of the Trigonometric Functions*} + +lemma sin_zero [simp]: "sin 0 = 0" +unfolding sin_def by (simp add: powser_zero) + +lemma cos_zero [simp]: "cos 0 = 1" +unfolding cos_def by (simp add: powser_zero) + +lemma DERIV_sin_sin_mult [simp]: + "DERIV (%x. sin(x)*sin(x)) x :> cos(x) * sin(x) + cos(x) * sin(x)" +by (rule DERIV_mult, auto) + +lemma DERIV_sin_sin_mult2 [simp]: + "DERIV (%x. sin(x)*sin(x)) x :> 2 * cos(x) * sin(x)" +apply (cut_tac x = x in DERIV_sin_sin_mult) +apply (auto simp add: mult_assoc) +done + +lemma DERIV_sin_realpow2 [simp]: + "DERIV (%x. (sin x)\) x :> cos(x) * sin(x) + cos(x) * sin(x)" +by (auto simp add: numeral_2_eq_2 real_mult_assoc [symmetric]) + +lemma DERIV_sin_realpow2a [simp]: + "DERIV (%x. (sin x)\) x :> 2 * cos(x) * sin(x)" +by (auto simp add: numeral_2_eq_2) + +lemma DERIV_cos_cos_mult [simp]: + "DERIV (%x. cos(x)*cos(x)) x :> -sin(x) * cos(x) + -sin(x) * cos(x)" +by (rule DERIV_mult, auto) + +lemma DERIV_cos_cos_mult2 [simp]: + "DERIV (%x. cos(x)*cos(x)) x :> -2 * cos(x) * sin(x)" +apply (cut_tac x = x in DERIV_cos_cos_mult) +apply (auto simp add: mult_ac) +done + +lemma DERIV_cos_realpow2 [simp]: + "DERIV (%x. (cos x)\) x :> -sin(x) * cos(x) + -sin(x) * cos(x)" +by (auto simp add: numeral_2_eq_2 real_mult_assoc [symmetric]) + +lemma DERIV_cos_realpow2a [simp]: + "DERIV (%x. (cos x)\) x :> -2 * cos(x) * sin(x)" +by (auto simp add: numeral_2_eq_2) + +lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" +by auto + +lemma DERIV_cos_realpow2b: "DERIV (%x. (cos x)\) x :> -(2 * cos(x) * sin(x))" +apply (rule lemma_DERIV_subst) +apply (rule DERIV_cos_realpow2a, auto) +done + +(* most useful *) +lemma DERIV_cos_cos_mult3 [simp]: + "DERIV (%x. cos(x)*cos(x)) x :> -(2 * cos(x) * sin(x))" +apply (rule lemma_DERIV_subst) +apply (rule DERIV_cos_cos_mult2, auto) +done + +lemma DERIV_sin_circle_all: + "\x. DERIV (%x. (sin x)\ + (cos x)\) x :> + (2*cos(x)*sin(x) - 2*cos(x)*sin(x))" +apply (simp only: diff_minus, safe) +apply (rule DERIV_add) +apply (auto simp add: numeral_2_eq_2) +done + +lemma DERIV_sin_circle_all_zero [simp]: + "\x. DERIV (%x. (sin x)\ + (cos x)\) x :> 0" +by (cut_tac DERIV_sin_circle_all, auto) + +lemma sin_cos_squared_add [simp]: "((sin x)\) + ((cos x)\) = 1" +apply (cut_tac x = x and y = 0 in DERIV_sin_circle_all_zero [THEN DERIV_isconst_all]) +apply (auto simp add: numeral_2_eq_2) +done + +lemma sin_cos_squared_add2 [simp]: "((cos x)\) + ((sin x)\) = 1" +apply (subst add_commute) +apply (simp (no_asm) del: realpow_Suc) +done + +lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1" +apply (cut_tac x = x in sin_cos_squared_add2) +apply (auto simp add: numeral_2_eq_2) +done + +lemma sin_squared_eq: "(sin x)\ = 1 - (cos x)\" +apply (rule_tac a1 = "(cos x)\" in add_right_cancel [THEN iffD1]) +apply (simp del: realpow_Suc) +done + +lemma cos_squared_eq: "(cos x)\ = 1 - (sin x)\" +apply (rule_tac a1 = "(sin x)\" in add_right_cancel [THEN iffD1]) +apply (simp del: realpow_Suc) +done + +lemma real_gt_one_ge_zero_add_less: "[| 1 < x; 0 \ y |] ==> 1 < x + (y::real)" +by arith + +lemma abs_sin_le_one [simp]: "\sin x\ \ 1" +by (rule power2_le_imp_le, simp_all add: sin_squared_eq) + +lemma sin_ge_minus_one [simp]: "-1 \ sin x" +apply (insert abs_sin_le_one [of x]) +apply (simp add: abs_le_iff del: abs_sin_le_one) +done + +lemma sin_le_one [simp]: "sin x \ 1" +apply (insert abs_sin_le_one [of x]) +apply (simp add: abs_le_iff del: abs_sin_le_one) +done + +lemma abs_cos_le_one [simp]: "\cos x\ \ 1" +by (rule power2_le_imp_le, simp_all add: cos_squared_eq) + +lemma cos_ge_minus_one [simp]: "-1 \ cos x" +apply (insert abs_cos_le_one [of x]) +apply (simp add: abs_le_iff del: abs_cos_le_one) +done + +lemma cos_le_one [simp]: "cos x \ 1" +apply (insert abs_cos_le_one [of x]) +apply (simp add: abs_le_iff del: abs_cos_le_one) +done + +lemma DERIV_fun_pow: "DERIV g x :> m ==> + DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = "(%x. x ^ n)" in DERIV_chain2) +apply (rule DERIV_pow, auto) +done + +lemma DERIV_fun_exp: + "DERIV g x :> m ==> DERIV (%x. exp(g x)) x :> exp(g x) * m" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = exp in DERIV_chain2) +apply (rule DERIV_exp, auto) +done + +lemma DERIV_fun_sin: + "DERIV g x :> m ==> DERIV (%x. sin(g x)) x :> cos(g x) * m" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = sin in DERIV_chain2) +apply (rule DERIV_sin, auto) +done + +lemma DERIV_fun_cos: + "DERIV g x :> m ==> DERIV (%x. cos(g x)) x :> -sin(g x) * m" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = cos in DERIV_chain2) +apply (rule DERIV_cos, auto) +done + +lemmas DERIV_intros = DERIV_ident DERIV_const DERIV_cos DERIV_cmult + DERIV_sin DERIV_exp DERIV_inverse DERIV_pow + DERIV_add DERIV_diff DERIV_mult DERIV_minus + DERIV_inverse_fun DERIV_quotient DERIV_fun_pow + DERIV_fun_exp DERIV_fun_sin DERIV_fun_cos + +(* lemma *) +lemma lemma_DERIV_sin_cos_add: + "\x. + DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + + (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2) x :> 0" +apply (safe, rule lemma_DERIV_subst) +apply (best intro!: DERIV_intros intro: DERIV_chain2) + --{*replaces the old @{text DERIV_tac}*} +apply (auto simp add: diff_minus left_distrib right_distrib mult_ac add_ac) +done + +lemma sin_cos_add [simp]: + "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + + (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = 0" +apply (cut_tac y = 0 and x = x and y7 = y + in lemma_DERIV_sin_cos_add [THEN DERIV_isconst_all]) +apply (auto simp add: numeral_2_eq_2) +done + +lemma sin_add: "sin (x + y) = sin x * cos y + cos x * sin y" +apply (cut_tac x = x and y = y in sin_cos_add) +apply (simp del: sin_cos_add) +done + +lemma cos_add: "cos (x + y) = cos x * cos y - sin x * sin y" +apply (cut_tac x = x and y = y in sin_cos_add) +apply (simp del: sin_cos_add) +done + +lemma lemma_DERIV_sin_cos_minus: + "\x. DERIV (%x. (sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2) x :> 0" +apply (safe, rule lemma_DERIV_subst) +apply (best intro!: DERIV_intros intro: DERIV_chain2) +apply (auto simp add: diff_minus left_distrib right_distrib mult_ac add_ac) +done + +lemma sin_cos_minus [simp]: + "(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0" +apply (cut_tac y = 0 and x = x + in lemma_DERIV_sin_cos_minus [THEN DERIV_isconst_all]) +apply simp +done + +lemma sin_minus [simp]: "sin (-x) = -sin(x)" +apply (cut_tac x = x in sin_cos_minus) +apply (simp del: sin_cos_minus) +done + +lemma cos_minus [simp]: "cos (-x) = cos(x)" +apply (cut_tac x = x in sin_cos_minus) +apply (simp del: sin_cos_minus) +done + +lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y" +by (simp add: diff_minus sin_add) + +lemma sin_diff2: "sin (x - y) = cos y * sin x - sin y * cos x" +by (simp add: sin_diff mult_commute) + +lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y" +by (simp add: diff_minus cos_add) + +lemma cos_diff2: "cos (x - y) = cos y * cos x + sin y * sin x" +by (simp add: cos_diff mult_commute) + +lemma sin_double [simp]: "sin(2 * x) = 2* sin x * cos x" +by (cut_tac x = x and y = x in sin_add, auto) + + +lemma cos_double: "cos(2* x) = ((cos x)\) - ((sin x)\)" +apply (cut_tac x = x and y = x in cos_add) +apply (simp add: power2_eq_square) +done + + +subsection{*The Constant Pi*} + +definition + pi :: "real" where + "pi = 2 * (THE x. 0 \ (x::real) & x \ 2 & cos x = 0)" + +text{*Show that there's a least positive @{term x} with @{term "cos(x) = 0"}; + hence define pi.*} + +lemma sin_paired: + "(%n. -1 ^ n /(real (fact (2 * n + 1))) * x ^ (2 * n + 1)) + sums sin x" +proof - + have "(\n. \k = n * 2.. 0 < sin x" +apply (subgoal_tac + "(\n. \k = n * 2..n. -1 ^ n / real (fact (2 * n + 1)) * x ^ (2 * n + 1))") + prefer 2 + apply (rule sin_paired [THEN sums_summable, THEN sums_group], simp) +apply (rotate_tac 2) +apply (drule sin_paired [THEN sums_unique, THEN ssubst]) +apply (auto simp del: fact_Suc realpow_Suc) +apply (frule sums_unique) +apply (auto simp del: fact_Suc realpow_Suc) +apply (rule_tac n1 = 0 in series_pos_less [THEN [2] order_le_less_trans]) +apply (auto simp del: fact_Suc realpow_Suc) +apply (erule sums_summable) +apply (case_tac "m=0") +apply (simp (no_asm_simp)) +apply (subgoal_tac "6 * (x * (x * x) / real (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) < 6 * x") +apply (simp only: mult_less_cancel_left, simp) +apply (simp (no_asm_simp) add: numeral_2_eq_2 [symmetric] mult_assoc [symmetric]) +apply (subgoal_tac "x*x < 2*3", simp) +apply (rule mult_strict_mono) +apply (auto simp add: real_0_less_add_iff real_of_nat_Suc simp del: fact_Suc) +apply (subst fact_Suc) +apply (subst fact_Suc) +apply (subst fact_Suc) +apply (subst fact_Suc) +apply (subst real_of_nat_mult) +apply (subst real_of_nat_mult) +apply (subst real_of_nat_mult) +apply (subst real_of_nat_mult) +apply (simp (no_asm) add: divide_inverse del: fact_Suc) +apply (auto simp add: mult_assoc [symmetric] simp del: fact_Suc) +apply (rule_tac c="real (Suc (Suc (4*m)))" in mult_less_imp_less_right) +apply (auto simp add: mult_assoc simp del: fact_Suc) +apply (rule_tac c="real (Suc (Suc (Suc (4*m))))" in mult_less_imp_less_right) +apply (auto simp add: mult_assoc mult_less_cancel_left simp del: fact_Suc) +apply (subgoal_tac "x * (x * x ^ (4*m)) = (x ^ (4*m)) * (x * x)") +apply (erule ssubst)+ +apply (auto simp del: fact_Suc) +apply (subgoal_tac "0 < x ^ (4 * m) ") + prefer 2 apply (simp only: zero_less_power) +apply (simp (no_asm_simp) add: mult_less_cancel_left) +apply (rule mult_strict_mono) +apply (simp_all (no_asm_simp)) +done + +lemma sin_gt_zero1: "[|0 < x; x < 2 |] ==> 0 < sin x" +by (auto intro: sin_gt_zero) + +lemma cos_double_less_one: "[| 0 < x; x < 2 |] ==> cos (2 * x) < 1" +apply (cut_tac x = x in sin_gt_zero1) +apply (auto simp add: cos_squared_eq cos_double) +done + +lemma cos_paired: + "(%n. -1 ^ n /(real (fact (2 * n))) * x ^ (2 * n)) sums cos x" +proof - + have "(\n. \k = n * 2..n=0..< Suc(Suc(Suc 0)). - (-1 ^ n / (real(fact (2*n))) * 2 ^ (2*n))" + in order_less_trans) +apply (simp (no_asm) add: fact_num_eq_if realpow_num_eq_if del: fact_Suc realpow_Suc) +apply (simp (no_asm) add: mult_assoc del: setsum_op_ivl_Suc) +apply (rule sumr_pos_lt_pair) +apply (erule sums_summable, safe) +apply (simp (no_asm) add: divide_inverse real_0_less_add_iff mult_assoc [symmetric] + del: fact_Suc) +apply (rule real_mult_inverse_cancel2) +apply (rule real_of_nat_fact_gt_zero)+ +apply (simp (no_asm) add: mult_assoc [symmetric] del: fact_Suc) +apply (subst fact_lemma) +apply (subst fact_Suc [of "Suc (Suc (Suc (Suc (Suc (Suc (Suc (4 * d)))))))"]) +apply (simp only: real_of_nat_mult) +apply (rule mult_strict_mono, force) + apply (rule_tac [3] real_of_nat_ge_zero) + prefer 2 apply force +apply (rule real_of_nat_less_iff [THEN iffD2]) +apply (rule fact_less_mono, auto) +done + +lemmas cos_two_neq_zero [simp] = cos_two_less_zero [THEN less_imp_neq] +lemmas cos_two_le_zero [simp] = cos_two_less_zero [THEN order_less_imp_le] + +lemma cos_is_zero: "EX! x. 0 \ x & x \ 2 & cos x = 0" +apply (subgoal_tac "\x. 0 \ x & x \ 2 & cos x = 0") +apply (rule_tac [2] IVT2) +apply (auto intro: DERIV_isCont DERIV_cos) +apply (cut_tac x = xa and y = y in linorder_less_linear) +apply (rule ccontr) +apply (subgoal_tac " (\x. cos differentiable x) & (\x. isCont cos x) ") +apply (auto intro: DERIV_cos DERIV_isCont simp add: differentiable_def) +apply (drule_tac f = cos in Rolle) +apply (drule_tac [5] f = cos in Rolle) +apply (auto dest!: DERIV_cos [THEN DERIV_unique] simp add: differentiable_def) +apply (drule_tac y1 = xa in order_le_less_trans [THEN sin_gt_zero]) +apply (assumption, rule_tac y=y in order_less_le_trans, simp_all) +apply (drule_tac y1 = y in order_le_less_trans [THEN sin_gt_zero], assumption, simp_all) +done + +lemma pi_half: "pi/2 = (THE x. 0 \ x & x \ 2 & cos x = 0)" +by (simp add: pi_def) + +lemma cos_pi_half [simp]: "cos (pi / 2) = 0" +by (simp add: pi_half cos_is_zero [THEN theI']) + +lemma pi_half_gt_zero [simp]: "0 < pi / 2" +apply (rule order_le_neq_trans) +apply (simp add: pi_half cos_is_zero [THEN theI']) +apply (rule notI, drule arg_cong [where f=cos], simp) +done + +lemmas pi_half_neq_zero [simp] = pi_half_gt_zero [THEN less_imp_neq, symmetric] +lemmas pi_half_ge_zero [simp] = pi_half_gt_zero [THEN order_less_imp_le] + +lemma pi_half_less_two [simp]: "pi / 2 < 2" +apply (rule order_le_neq_trans) +apply (simp add: pi_half cos_is_zero [THEN theI']) +apply (rule notI, drule arg_cong [where f=cos], simp) +done + +lemmas pi_half_neq_two [simp] = pi_half_less_two [THEN less_imp_neq] +lemmas pi_half_le_two [simp] = pi_half_less_two [THEN order_less_imp_le] + +lemma pi_gt_zero [simp]: "0 < pi" +by (insert pi_half_gt_zero, simp) + +lemma pi_ge_zero [simp]: "0 \ pi" +by (rule pi_gt_zero [THEN order_less_imp_le]) + +lemma pi_neq_zero [simp]: "pi \ 0" +by (rule pi_gt_zero [THEN less_imp_neq, THEN not_sym]) + +lemma pi_not_less_zero [simp]: "\ pi < 0" +by (simp add: linorder_not_less) + +lemma minus_pi_half_less_zero [simp]: "-(pi/2) < 0" +by auto + +lemma sin_pi_half [simp]: "sin(pi/2) = 1" +apply (cut_tac x = "pi/2" in sin_cos_squared_add2) +apply (cut_tac sin_gt_zero [OF pi_half_gt_zero pi_half_less_two]) +apply (simp add: power2_eq_square) +done + +lemma cos_pi [simp]: "cos pi = -1" +by (cut_tac x = "pi/2" and y = "pi/2" in cos_add, simp) + +lemma sin_pi [simp]: "sin pi = 0" +by (cut_tac x = "pi/2" and y = "pi/2" in sin_add, simp) + +lemma sin_cos_eq: "sin x = cos (pi/2 - x)" +by (simp add: diff_minus cos_add) +declare sin_cos_eq [symmetric, simp] + +lemma minus_sin_cos_eq: "-sin x = cos (x + pi/2)" +by (simp add: cos_add) +declare minus_sin_cos_eq [symmetric, simp] + +lemma cos_sin_eq: "cos x = sin (pi/2 - x)" +by (simp add: diff_minus sin_add) +declare cos_sin_eq [symmetric, simp] + +lemma sin_periodic_pi [simp]: "sin (x + pi) = - sin x" +by (simp add: sin_add) + +lemma sin_periodic_pi2 [simp]: "sin (pi + x) = - sin x" +by (simp add: sin_add) + +lemma cos_periodic_pi [simp]: "cos (x + pi) = - cos x" +by (simp add: cos_add) + +lemma sin_periodic [simp]: "sin (x + 2*pi) = sin x" +by (simp add: sin_add cos_double) + +lemma cos_periodic [simp]: "cos (x + 2*pi) = cos x" +by (simp add: cos_add cos_double) + +lemma cos_npi [simp]: "cos (real n * pi) = -1 ^ n" +apply (induct "n") +apply (auto simp add: real_of_nat_Suc left_distrib) +done + +lemma cos_npi2 [simp]: "cos (pi * real n) = -1 ^ n" +proof - + have "cos (pi * real n) = cos (real n * pi)" by (simp only: mult_commute) + also have "... = -1 ^ n" by (rule cos_npi) + finally show ?thesis . +qed + +lemma sin_npi [simp]: "sin (real (n::nat) * pi) = 0" +apply (induct "n") +apply (auto simp add: real_of_nat_Suc left_distrib) +done + +lemma sin_npi2 [simp]: "sin (pi * real (n::nat)) = 0" +by (simp add: mult_commute [of pi]) + +lemma cos_two_pi [simp]: "cos (2 * pi) = 1" +by (simp add: cos_double) + +lemma sin_two_pi [simp]: "sin (2 * pi) = 0" +by simp + +lemma sin_gt_zero2: "[| 0 < x; x < pi/2 |] ==> 0 < sin x" +apply (rule sin_gt_zero, assumption) +apply (rule order_less_trans, assumption) +apply (rule pi_half_less_two) +done + +lemma sin_less_zero: + assumes lb: "- pi/2 < x" and "x < 0" shows "sin x < 0" +proof - + have "0 < sin (- x)" using prems by (simp only: sin_gt_zero2) + thus ?thesis by simp +qed + +lemma pi_less_4: "pi < 4" +by (cut_tac pi_half_less_two, auto) + +lemma cos_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < cos x" +apply (cut_tac pi_less_4) +apply (cut_tac f = cos and a = 0 and b = x and y = 0 in IVT2_objl, safe, simp_all) +apply (cut_tac cos_is_zero, safe) +apply (rename_tac y z) +apply (drule_tac x = y in spec) +apply (drule_tac x = "pi/2" in spec, simp) +done + +lemma cos_gt_zero_pi: "[| -(pi/2) < x; x < pi/2 |] ==> 0 < cos x" +apply (rule_tac x = x and y = 0 in linorder_cases) +apply (rule cos_minus [THEN subst]) +apply (rule cos_gt_zero) +apply (auto intro: cos_gt_zero) +done + +lemma cos_ge_zero: "[| -(pi/2) \ x; x \ pi/2 |] ==> 0 \ cos x" +apply (auto simp add: order_le_less cos_gt_zero_pi) +apply (subgoal_tac "x = pi/2", auto) +done + +lemma sin_gt_zero_pi: "[| 0 < x; x < pi |] ==> 0 < sin x" +apply (subst sin_cos_eq) +apply (rotate_tac 1) +apply (drule real_sum_of_halves [THEN ssubst]) +apply (auto intro!: cos_gt_zero_pi simp del: sin_cos_eq [symmetric]) +done + +lemma sin_ge_zero: "[| 0 \ x; x \ pi |] ==> 0 \ sin x" +by (auto simp add: order_le_less sin_gt_zero_pi) + +lemma cos_total: "[| -1 \ y; y \ 1 |] ==> EX! x. 0 \ x & x \ pi & (cos x = y)" +apply (subgoal_tac "\x. 0 \ x & x \ pi & cos x = y") +apply (rule_tac [2] IVT2) +apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos) +apply (cut_tac x = xa and y = y in linorder_less_linear) +apply (rule ccontr, auto) +apply (drule_tac f = cos in Rolle) +apply (drule_tac [5] f = cos in Rolle) +apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos + dest!: DERIV_cos [THEN DERIV_unique] + simp add: differentiable_def) +apply (auto dest: sin_gt_zero_pi [OF order_le_less_trans order_less_le_trans]) +done + +lemma sin_total: + "[| -1 \ y; y \ 1 |] ==> EX! x. -(pi/2) \ x & x \ pi/2 & (sin x = y)" +apply (rule ccontr) +apply (subgoal_tac "\x. (- (pi/2) \ x & x \ pi/2 & (sin x = y)) = (0 \ (x + pi/2) & (x + pi/2) \ pi & (cos (x + pi/2) = -y))") +apply (erule contrapos_np) +apply (simp del: minus_sin_cos_eq [symmetric]) +apply (cut_tac y="-y" in cos_total, simp) apply simp +apply (erule ex1E) +apply (rule_tac a = "x - (pi/2)" in ex1I) +apply (simp (no_asm) add: add_assoc) +apply (rotate_tac 3) +apply (drule_tac x = "xa + pi/2" in spec, safe, simp_all) +done + +lemma reals_Archimedean4: + "[| 0 < y; 0 \ x |] ==> \n. real n * y \ x & x < real (Suc n) * y" +apply (auto dest!: reals_Archimedean3) +apply (drule_tac x = x in spec, clarify) +apply (subgoal_tac "x < real(LEAST m::nat. x < real m * y) * y") + prefer 2 apply (erule LeastI) +apply (case_tac "LEAST m::nat. x < real m * y", simp) +apply (subgoal_tac "~ x < real nat * y") + prefer 2 apply (rule not_less_Least, simp, force) +done + +(* Pre Isabelle99-2 proof was simpler- numerals arithmetic + now causes some unwanted re-arrangements of literals! *) +lemma cos_zero_lemma: + "[| 0 \ x; cos x = 0 |] ==> + \n::nat. ~even n & x = real n * (pi/2)" +apply (drule pi_gt_zero [THEN reals_Archimedean4], safe) +apply (subgoal_tac "0 \ x - real n * pi & + (x - real n * pi) \ pi & (cos (x - real n * pi) = 0) ") +apply (auto simp add: compare_rls) + prefer 3 apply (simp add: cos_diff) + prefer 2 apply (simp add: real_of_nat_Suc left_distrib) +apply (simp add: cos_diff) +apply (subgoal_tac "EX! x. 0 \ x & x \ pi & cos x = 0") +apply (rule_tac [2] cos_total, safe) +apply (drule_tac x = "x - real n * pi" in spec) +apply (drule_tac x = "pi/2" in spec) +apply (simp add: cos_diff) +apply (rule_tac x = "Suc (2 * n)" in exI) +apply (simp add: real_of_nat_Suc left_distrib, auto) +done + +lemma sin_zero_lemma: + "[| 0 \ x; sin x = 0 |] ==> + \n::nat. even n & x = real n * (pi/2)" +apply (subgoal_tac "\n::nat. ~ even n & x + pi/2 = real n * (pi/2) ") + apply (clarify, rule_tac x = "n - 1" in exI) + apply (force simp add: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib) +apply (rule cos_zero_lemma) +apply (simp_all add: add_increasing) +done + + +lemma cos_zero_iff: + "(cos x = 0) = + ((\n::nat. ~even n & (x = real n * (pi/2))) | + (\n::nat. ~even n & (x = -(real n * (pi/2)))))" +apply (rule iffI) +apply (cut_tac linorder_linear [of 0 x], safe) +apply (drule cos_zero_lemma, assumption+) +apply (cut_tac x="-x" in cos_zero_lemma, simp, simp) +apply (force simp add: minus_equation_iff [of x]) +apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib) +apply (auto simp add: cos_add) +done + +(* ditto: but to a lesser extent *) +lemma sin_zero_iff: + "(sin x = 0) = + ((\n::nat. even n & (x = real n * (pi/2))) | + (\n::nat. even n & (x = -(real n * (pi/2)))))" +apply (rule iffI) +apply (cut_tac linorder_linear [of 0 x], safe) +apply (drule sin_zero_lemma, assumption+) +apply (cut_tac x="-x" in sin_zero_lemma, simp, simp, safe) +apply (force simp add: minus_equation_iff [of x]) +apply (auto simp add: even_mult_two_ex) +done + + +subsection{*Tangent*} + +definition + tan :: "real => real" where + "tan x = (sin x)/(cos x)" + +lemma tan_zero [simp]: "tan 0 = 0" +by (simp add: tan_def) + +lemma tan_pi [simp]: "tan pi = 0" +by (simp add: tan_def) + +lemma tan_npi [simp]: "tan (real (n::nat) * pi) = 0" +by (simp add: tan_def) + +lemma tan_minus [simp]: "tan (-x) = - tan x" +by (simp add: tan_def minus_mult_left) + +lemma tan_periodic [simp]: "tan (x + 2*pi) = tan x" +by (simp add: tan_def) + +lemma lemma_tan_add1: + "[| cos x \ 0; cos y \ 0 |] + ==> 1 - tan(x)*tan(y) = cos (x + y)/(cos x * cos y)" +apply (simp add: tan_def divide_inverse) +apply (auto simp del: inverse_mult_distrib + simp add: inverse_mult_distrib [symmetric] mult_ac) +apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst]) +apply (auto simp del: inverse_mult_distrib + simp add: mult_assoc left_diff_distrib cos_add) +done + +lemma add_tan_eq: + "[| cos x \ 0; cos y \ 0 |] + ==> tan x + tan y = sin(x + y)/(cos x * cos y)" +apply (simp add: tan_def) +apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst]) +apply (auto simp add: mult_assoc left_distrib) +apply (simp add: sin_add) +done + +lemma tan_add: + "[| cos x \ 0; cos y \ 0; cos (x + y) \ 0 |] + ==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))" +apply (simp (no_asm_simp) add: add_tan_eq lemma_tan_add1) +apply (simp add: tan_def) +done + +lemma tan_double: + "[| cos x \ 0; cos (2 * x) \ 0 |] + ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))" +apply (insert tan_add [of x x]) +apply (simp add: mult_2 [symmetric]) +apply (auto simp add: numeral_2_eq_2) +done + +lemma tan_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < tan x" +by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) + +lemma tan_less_zero: + assumes lb: "- pi/2 < x" and "x < 0" shows "tan x < 0" +proof - + have "0 < tan (- x)" using prems by (simp only: tan_gt_zero) + thus ?thesis by simp +qed + +lemma lemma_DERIV_tan: + "cos x \ 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse((cos x)\)" +apply (rule lemma_DERIV_subst) +apply (best intro!: DERIV_intros intro: DERIV_chain2) +apply (auto simp add: divide_inverse numeral_2_eq_2) +done + +lemma DERIV_tan [simp]: "cos x \ 0 ==> DERIV tan x :> inverse((cos x)\)" +by (auto dest: lemma_DERIV_tan simp add: tan_def [symmetric]) + +lemma isCont_tan [simp]: "cos x \ 0 ==> isCont tan x" +by (rule DERIV_tan [THEN DERIV_isCont]) + +lemma LIM_cos_div_sin [simp]: "(%x. cos(x)/sin(x)) -- pi/2 --> 0" +apply (subgoal_tac "(\x. cos x * inverse (sin x)) -- pi * inverse 2 --> 0*1") +apply (simp add: divide_inverse [symmetric]) +apply (rule LIM_mult) +apply (rule_tac [2] inverse_1 [THEN subst]) +apply (rule_tac [2] LIM_inverse) +apply (simp_all add: divide_inverse [symmetric]) +apply (simp_all only: isCont_def [symmetric] cos_pi_half [symmetric] sin_pi_half [symmetric]) +apply (blast intro!: DERIV_isCont DERIV_sin DERIV_cos)+ +done + +lemma lemma_tan_total: "0 < y ==> \x. 0 < x & x < pi/2 & y < tan x" +apply (cut_tac LIM_cos_div_sin) +apply (simp only: LIM_def) +apply (drule_tac x = "inverse y" in spec, safe, force) +apply (drule_tac ?d1.0 = s in pi_half_gt_zero [THEN [2] real_lbound_gt_zero], safe) +apply (rule_tac x = "(pi/2) - e" in exI) +apply (simp (no_asm_simp)) +apply (drule_tac x = "(pi/2) - e" in spec) +apply (auto simp add: tan_def) +apply (rule inverse_less_iff_less [THEN iffD1]) +apply (auto simp add: divide_inverse) +apply (rule real_mult_order) +apply (subgoal_tac [3] "0 < sin e & 0 < cos e") +apply (auto intro: cos_gt_zero sin_gt_zero2 simp add: mult_commute) +done + +lemma tan_total_pos: "0 \ y ==> \x. 0 \ x & x < pi/2 & tan x = y" +apply (frule order_le_imp_less_or_eq, safe) + prefer 2 apply force +apply (drule lemma_tan_total, safe) +apply (cut_tac f = tan and a = 0 and b = x and y = y in IVT_objl) +apply (auto intro!: DERIV_tan [THEN DERIV_isCont]) +apply (drule_tac y = xa in order_le_imp_less_or_eq) +apply (auto dest: cos_gt_zero) +done + +lemma lemma_tan_total1: "\x. -(pi/2) < x & x < (pi/2) & tan x = y" +apply (cut_tac linorder_linear [of 0 y], safe) +apply (drule tan_total_pos) +apply (cut_tac [2] y="-y" in tan_total_pos, safe) +apply (rule_tac [3] x = "-x" in exI) +apply (auto intro!: exI) +done + +lemma tan_total: "EX! x. -(pi/2) < x & x < (pi/2) & tan x = y" +apply (cut_tac y = y in lemma_tan_total1, auto) +apply (cut_tac x = xa and y = y in linorder_less_linear, auto) +apply (subgoal_tac [2] "\z. y < z & z < xa & DERIV tan z :> 0") +apply (subgoal_tac "\z. xa < z & z < y & DERIV tan z :> 0") +apply (rule_tac [4] Rolle) +apply (rule_tac [2] Rolle) +apply (auto intro!: DERIV_tan DERIV_isCont exI + simp add: differentiable_def) +txt{*Now, simulate TRYALL*} +apply (rule_tac [!] DERIV_tan asm_rl) +apply (auto dest!: DERIV_unique [OF _ DERIV_tan] + simp add: cos_gt_zero_pi [THEN less_imp_neq, THEN not_sym]) +done + + +subsection {* Inverse Trigonometric Functions *} + +definition + arcsin :: "real => real" where + "arcsin y = (THE x. -(pi/2) \ x & x \ pi/2 & sin x = y)" + +definition + arccos :: "real => real" where + "arccos y = (THE x. 0 \ x & x \ pi & cos x = y)" + +definition + arctan :: "real => real" where + "arctan y = (THE x. -(pi/2) < x & x < pi/2 & tan x = y)" + +lemma arcsin: + "[| -1 \ y; y \ 1 |] + ==> -(pi/2) \ arcsin y & + arcsin y \ pi/2 & sin(arcsin y) = y" +unfolding arcsin_def by (rule theI' [OF sin_total]) + +lemma arcsin_pi: + "[| -1 \ y; y \ 1 |] + ==> -(pi/2) \ arcsin y & arcsin y \ pi & sin(arcsin y) = y" +apply (drule (1) arcsin) +apply (force intro: order_trans) +done + +lemma sin_arcsin [simp]: "[| -1 \ y; y \ 1 |] ==> sin(arcsin y) = y" +by (blast dest: arcsin) + +lemma arcsin_bounded: + "[| -1 \ y; y \ 1 |] ==> -(pi/2) \ arcsin y & arcsin y \ pi/2" +by (blast dest: arcsin) + +lemma arcsin_lbound: "[| -1 \ y; y \ 1 |] ==> -(pi/2) \ arcsin y" +by (blast dest: arcsin) + +lemma arcsin_ubound: "[| -1 \ y; y \ 1 |] ==> arcsin y \ pi/2" +by (blast dest: arcsin) + +lemma arcsin_lt_bounded: + "[| -1 < y; y < 1 |] ==> -(pi/2) < arcsin y & arcsin y < pi/2" +apply (frule order_less_imp_le) +apply (frule_tac y = y in order_less_imp_le) +apply (frule arcsin_bounded) +apply (safe, simp) +apply (drule_tac y = "arcsin y" in order_le_imp_less_or_eq) +apply (drule_tac [2] y = "pi/2" in order_le_imp_less_or_eq, safe) +apply (drule_tac [!] f = sin in arg_cong, auto) +done + +lemma arcsin_sin: "[|-(pi/2) \ x; x \ pi/2 |] ==> arcsin(sin x) = x" +apply (unfold arcsin_def) +apply (rule the1_equality) +apply (rule sin_total, auto) +done + +lemma arccos: + "[| -1 \ y; y \ 1 |] + ==> 0 \ arccos y & arccos y \ pi & cos(arccos y) = y" +unfolding arccos_def by (rule theI' [OF cos_total]) + +lemma cos_arccos [simp]: "[| -1 \ y; y \ 1 |] ==> cos(arccos y) = y" +by (blast dest: arccos) + +lemma arccos_bounded: "[| -1 \ y; y \ 1 |] ==> 0 \ arccos y & arccos y \ pi" +by (blast dest: arccos) + +lemma arccos_lbound: "[| -1 \ y; y \ 1 |] ==> 0 \ arccos y" +by (blast dest: arccos) + +lemma arccos_ubound: "[| -1 \ y; y \ 1 |] ==> arccos y \ pi" +by (blast dest: arccos) + +lemma arccos_lt_bounded: + "[| -1 < y; y < 1 |] + ==> 0 < arccos y & arccos y < pi" +apply (frule order_less_imp_le) +apply (frule_tac y = y in order_less_imp_le) +apply (frule arccos_bounded, auto) +apply (drule_tac y = "arccos y" in order_le_imp_less_or_eq) +apply (drule_tac [2] y = pi in order_le_imp_less_or_eq, auto) +apply (drule_tac [!] f = cos in arg_cong, auto) +done + +lemma arccos_cos: "[|0 \ x; x \ pi |] ==> arccos(cos x) = x" +apply (simp add: arccos_def) +apply (auto intro!: the1_equality cos_total) +done + +lemma arccos_cos2: "[|x \ 0; -pi \ x |] ==> arccos(cos x) = -x" +apply (simp add: arccos_def) +apply (auto intro!: the1_equality cos_total) +done + +lemma cos_arcsin: "\-1 \ x; x \ 1\ \ cos (arcsin x) = sqrt (1 - x\)" +apply (subgoal_tac "x\ \ 1") +apply (rule power2_eq_imp_eq) +apply (simp add: cos_squared_eq) +apply (rule cos_ge_zero) +apply (erule (1) arcsin_lbound) +apply (erule (1) arcsin_ubound) +apply simp +apply (subgoal_tac "\x\\ \ 1\", simp) +apply (rule power_mono, simp, simp) +done + +lemma sin_arccos: "\-1 \ x; x \ 1\ \ sin (arccos x) = sqrt (1 - x\)" +apply (subgoal_tac "x\ \ 1") +apply (rule power2_eq_imp_eq) +apply (simp add: sin_squared_eq) +apply (rule sin_ge_zero) +apply (erule (1) arccos_lbound) +apply (erule (1) arccos_ubound) +apply simp +apply (subgoal_tac "\x\\ \ 1\", simp) +apply (rule power_mono, simp, simp) +done + +lemma arctan [simp]: + "- (pi/2) < arctan y & arctan y < pi/2 & tan (arctan y) = y" +unfolding arctan_def by (rule theI' [OF tan_total]) + +lemma tan_arctan: "tan(arctan y) = y" +by auto + +lemma arctan_bounded: "- (pi/2) < arctan y & arctan y < pi/2" +by (auto simp only: arctan) + +lemma arctan_lbound: "- (pi/2) < arctan y" +by auto + +lemma arctan_ubound: "arctan y < pi/2" +by (auto simp only: arctan) + +lemma arctan_tan: + "[|-(pi/2) < x; x < pi/2 |] ==> arctan(tan x) = x" +apply (unfold arctan_def) +apply (rule the1_equality) +apply (rule tan_total, auto) +done + +lemma arctan_zero_zero [simp]: "arctan 0 = 0" +by (insert arctan_tan [of 0], simp) + +lemma cos_arctan_not_zero [simp]: "cos(arctan x) \ 0" +apply (auto simp add: cos_zero_iff) +apply (case_tac "n") +apply (case_tac [3] "n") +apply (cut_tac [2] y = x in arctan_ubound) +apply (cut_tac [4] y = x in arctan_lbound) +apply (auto simp add: real_of_nat_Suc left_distrib mult_less_0_iff) +done + +lemma tan_sec: "cos x \ 0 ==> 1 + tan(x) ^ 2 = inverse(cos x) ^ 2" +apply (rule power_inverse [THEN subst]) +apply (rule_tac c1 = "(cos x)\" in real_mult_right_cancel [THEN iffD1]) +apply (auto dest: field_power_not_zero + simp add: power_mult_distrib left_distrib power_divide tan_def + mult_assoc power_inverse [symmetric] + simp del: realpow_Suc) +done + +lemma isCont_inverse_function2: + fixes f g :: "real \ real" shows + "\a < x; x < b; + \z. a \ z \ z \ b \ g (f z) = z; + \z. a \ z \ z \ b \ isCont f z\ + \ isCont g (f x)" +apply (rule isCont_inverse_function + [where f=f and d="min (x - a) (b - x)"]) +apply (simp_all add: abs_le_iff) +done + +lemma isCont_arcsin: "\-1 < x; x < 1\ \ isCont arcsin x" +apply (subgoal_tac "isCont arcsin (sin (arcsin x))", simp) +apply (rule isCont_inverse_function2 [where f=sin]) +apply (erule (1) arcsin_lt_bounded [THEN conjunct1]) +apply (erule (1) arcsin_lt_bounded [THEN conjunct2]) +apply (fast intro: arcsin_sin, simp) +done + +lemma isCont_arccos: "\-1 < x; x < 1\ \ isCont arccos x" +apply (subgoal_tac "isCont arccos (cos (arccos x))", simp) +apply (rule isCont_inverse_function2 [where f=cos]) +apply (erule (1) arccos_lt_bounded [THEN conjunct1]) +apply (erule (1) arccos_lt_bounded [THEN conjunct2]) +apply (fast intro: arccos_cos, simp) +done + +lemma isCont_arctan: "isCont arctan x" +apply (rule arctan_lbound [of x, THEN dense, THEN exE], clarify) +apply (rule arctan_ubound [of x, THEN dense, THEN exE], clarify) +apply (subgoal_tac "isCont arctan (tan (arctan x))", simp) +apply (erule (1) isCont_inverse_function2 [where f=tan]) +apply (clarify, rule arctan_tan) +apply (erule (1) order_less_le_trans) +apply (erule (1) order_le_less_trans) +apply (clarify, rule isCont_tan) +apply (rule less_imp_neq [symmetric]) +apply (rule cos_gt_zero_pi) +apply (erule (1) order_less_le_trans) +apply (erule (1) order_le_less_trans) +done + +lemma DERIV_arcsin: + "\-1 < x; x < 1\ \ DERIV arcsin x :> inverse (sqrt (1 - x\))" +apply (rule DERIV_inverse_function [where f=sin and a="-1" and b="1"]) +apply (rule lemma_DERIV_subst [OF DERIV_sin]) +apply (simp add: cos_arcsin) +apply (subgoal_tac "\x\\ < 1\", simp) +apply (rule power_strict_mono, simp, simp, simp) +apply assumption +apply assumption +apply simp +apply (erule (1) isCont_arcsin) +done + +lemma DERIV_arccos: + "\-1 < x; x < 1\ \ DERIV arccos x :> inverse (- sqrt (1 - x\))" +apply (rule DERIV_inverse_function [where f=cos and a="-1" and b="1"]) +apply (rule lemma_DERIV_subst [OF DERIV_cos]) +apply (simp add: sin_arccos) +apply (subgoal_tac "\x\\ < 1\", simp) +apply (rule power_strict_mono, simp, simp, simp) +apply assumption +apply assumption +apply simp +apply (erule (1) isCont_arccos) +done + +lemma DERIV_arctan: "DERIV arctan x :> inverse (1 + x\)" +apply (rule DERIV_inverse_function [where f=tan and a="x - 1" and b="x + 1"]) +apply (rule lemma_DERIV_subst [OF DERIV_tan]) +apply (rule cos_arctan_not_zero) +apply (simp add: power_inverse tan_sec [symmetric]) +apply (subgoal_tac "0 < 1 + x\", simp) +apply (simp add: add_pos_nonneg) +apply (simp, simp, simp, rule isCont_arctan) +done + + +subsection {* More Theorems about Sin and Cos *} + +lemma cos_45: "cos (pi / 4) = sqrt 2 / 2" +proof - + let ?c = "cos (pi / 4)" and ?s = "sin (pi / 4)" + have nonneg: "0 \ ?c" + by (rule cos_ge_zero, rule order_trans [where y=0], simp_all) + have "0 = cos (pi / 4 + pi / 4)" + by simp + also have "cos (pi / 4 + pi / 4) = ?c\ - ?s\" + by (simp only: cos_add power2_eq_square) + also have "\ = 2 * ?c\ - 1" + by (simp add: sin_squared_eq) + finally have "?c\ = (sqrt 2 / 2)\" + by (simp add: power_divide) + thus ?thesis + using nonneg by (rule power2_eq_imp_eq) simp +qed + +lemma cos_30: "cos (pi / 6) = sqrt 3 / 2" +proof - + let ?c = "cos (pi / 6)" and ?s = "sin (pi / 6)" + have pos_c: "0 < ?c" + by (rule cos_gt_zero, simp, simp) + have "0 = cos (pi / 6 + pi / 6 + pi / 6)" + by simp + also have "\ = (?c * ?c - ?s * ?s) * ?c - (?s * ?c + ?c * ?s) * ?s" + by (simp only: cos_add sin_add) + also have "\ = ?c * (?c\ - 3 * ?s\)" + by (simp add: ring_simps power2_eq_square) + finally have "?c\ = (sqrt 3 / 2)\" + using pos_c by (simp add: sin_squared_eq power_divide) + thus ?thesis + using pos_c [THEN order_less_imp_le] + by (rule power2_eq_imp_eq) simp +qed + +lemma sin_45: "sin (pi / 4) = sqrt 2 / 2" +proof - + have "sin (pi / 4) = cos (pi / 2 - pi / 4)" by (rule sin_cos_eq) + also have "pi / 2 - pi / 4 = pi / 4" by simp + also have "cos (pi / 4) = sqrt 2 / 2" by (rule cos_45) + finally show ?thesis . +qed + +lemma sin_60: "sin (pi / 3) = sqrt 3 / 2" +proof - + have "sin (pi / 3) = cos (pi / 2 - pi / 3)" by (rule sin_cos_eq) + also have "pi / 2 - pi / 3 = pi / 6" by simp + also have "cos (pi / 6) = sqrt 3 / 2" by (rule cos_30) + finally show ?thesis . +qed + +lemma cos_60: "cos (pi / 3) = 1 / 2" +apply (rule power2_eq_imp_eq) +apply (simp add: cos_squared_eq sin_60 power_divide) +apply (rule cos_ge_zero, rule order_trans [where y=0], simp_all) +done + +lemma sin_30: "sin (pi / 6) = 1 / 2" +proof - + have "sin (pi / 6) = cos (pi / 2 - pi / 6)" by (rule sin_cos_eq) + also have "pi / 2 - pi / 6 = pi / 3" by simp + also have "cos (pi / 3) = 1 / 2" by (rule cos_60) + finally show ?thesis . +qed + +lemma tan_30: "tan (pi / 6) = 1 / sqrt 3" +unfolding tan_def by (simp add: sin_30 cos_30) + +lemma tan_45: "tan (pi / 4) = 1" +unfolding tan_def by (simp add: sin_45 cos_45) + +lemma tan_60: "tan (pi / 3) = sqrt 3" +unfolding tan_def by (simp add: sin_60 cos_60) + +text{*NEEDED??*} +lemma [simp]: + "sin (x + 1 / 2 * real (Suc m) * pi) = + cos (x + 1 / 2 * real (m) * pi)" +by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib, auto) + +text{*NEEDED??*} +lemma [simp]: + "sin (x + real (Suc m) * pi / 2) = + cos (x + real (m) * pi / 2)" +by (simp only: cos_add sin_add real_of_nat_Suc add_divide_distrib left_distrib, auto) + +lemma DERIV_sin_add [simp]: "DERIV (%x. sin (x + k)) xa :> cos (xa + k)" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = sin and g = "%x. x + k" in DERIV_chain2) +apply (best intro!: DERIV_intros intro: DERIV_chain2)+ +apply (simp (no_asm)) +done + +lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi / 2) = (-1) ^ n" +proof - + have "sin ((real n + 1/2) * pi) = cos (real n * pi)" + by (auto simp add: right_distrib sin_add left_distrib mult_ac) + thus ?thesis + by (simp add: real_of_nat_Suc left_distrib add_divide_distrib + mult_commute [of pi]) +qed + +lemma cos_2npi [simp]: "cos (2 * real (n::nat) * pi) = 1" +by (simp add: cos_double mult_assoc power_add [symmetric] numeral_2_eq_2) + +lemma cos_3over2_pi [simp]: "cos (3 / 2 * pi) = 0" +apply (subgoal_tac "cos (pi + pi/2) = 0", simp) +apply (subst cos_add, simp) +done + +lemma sin_2npi [simp]: "sin (2 * real (n::nat) * pi) = 0" +by (auto simp add: mult_assoc) + +lemma sin_3over2_pi [simp]: "sin (3 / 2 * pi) = - 1" +apply (subgoal_tac "sin (pi + pi/2) = - 1", simp) +apply (subst sin_add, simp) +done + +(*NEEDED??*) +lemma [simp]: + "cos(x + 1 / 2 * real(Suc m) * pi) = -sin (x + 1 / 2 * real m * pi)" +apply (simp only: cos_add sin_add real_of_nat_Suc right_distrib left_distrib minus_mult_right, auto) +done + +(*NEEDED??*) +lemma [simp]: "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)" +by (simp only: cos_add sin_add real_of_nat_Suc left_distrib add_divide_distrib, auto) + +lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0" +by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib add_divide_distrib, auto) + +lemma DERIV_cos_add [simp]: "DERIV (%x. cos (x + k)) xa :> - sin (xa + k)" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = cos and g = "%x. x + k" in DERIV_chain2) +apply (best intro!: DERIV_intros intro: DERIV_chain2)+ +apply (simp (no_asm)) +done + +lemma sin_zero_abs_cos_one: "sin x = 0 ==> \cos x\ = 1" +by (auto simp add: sin_zero_iff even_mult_two_ex) + +lemma exp_eq_one_iff [simp]: "(exp (x::real) = 1) = (x = 0)" +apply auto +apply (drule_tac f = ln in arg_cong, auto) +done + +lemma cos_one_sin_zero: "cos x = 1 ==> sin x = 0" +by (cut_tac x = x in sin_cos_squared_add3, auto) + + +subsection {* Existence of Polar Coordinates *} + +lemma cos_x_y_le_one: "\x / sqrt (x\ + y\)\ \ 1" +apply (rule power2_le_imp_le [OF _ zero_le_one]) +apply (simp add: abs_divide power_divide divide_le_eq not_sum_power2_lt_zero) +done + +lemma cos_arccos_abs: "\y\ \ 1 \ cos (arccos y) = y" +by (simp add: abs_le_iff) + +lemma sin_arccos_abs: "\y\ \ 1 \ sin (arccos y) = sqrt (1 - y\)" +by (simp add: sin_arccos abs_le_iff) + +lemmas cos_arccos_lemma1 = cos_arccos_abs [OF cos_x_y_le_one] + +lemmas sin_arccos_lemma1 = sin_arccos_abs [OF cos_x_y_le_one] + +lemma polar_ex1: + "0 < y ==> \r a. x = r * cos a & y = r * sin a" +apply (rule_tac x = "sqrt (x\ + y\)" in exI) +apply (rule_tac x = "arccos (x / sqrt (x\ + y\))" in exI) +apply (simp add: cos_arccos_lemma1) +apply (simp add: sin_arccos_lemma1) +apply (simp add: power_divide) +apply (simp add: real_sqrt_mult [symmetric]) +apply (simp add: right_diff_distrib) +done + +lemma polar_ex2: + "y < 0 ==> \r a. x = r * cos a & y = r * sin a" +apply (insert polar_ex1 [where x=x and y="-y"], simp, clarify) +apply (rule_tac x = r in exI) +apply (rule_tac x = "-a" in exI, simp) +done + +lemma polar_Ex: "\r a. x = r * cos a & y = r * sin a" +apply (rule_tac x=0 and y=y in linorder_cases) +apply (erule polar_ex1) +apply (rule_tac x=x in exI, rule_tac x=0 in exI, simp) +apply (erule polar_ex2) +done + + +subsection {* Theorems about Limits *} + +(* need to rename second isCont_inverse *) + +lemma isCont_inv_fun: + fixes f g :: "real \ real" + shows "[| 0 < d; \z. \z - x\ \ d --> g(f(z)) = z; + \z. \z - x\ \ d --> isCont f z |] + ==> isCont g (f x)" +by (rule isCont_inverse_function) + +lemma isCont_inv_fun_inv: + fixes f g :: "real \ real" + shows "[| 0 < d; + \z. \z - x\ \ d --> g(f(z)) = z; + \z. \z - x\ \ d --> isCont f z |] + ==> \e. 0 < e & + (\y. 0 < \y - f(x)\ & \y - f(x)\ < e --> f(g(y)) = y)" +apply (drule isCont_inj_range) +prefer 2 apply (assumption, assumption, auto) +apply (rule_tac x = e in exI, auto) +apply (rotate_tac 2) +apply (drule_tac x = y in spec, auto) +done + + +text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*} +lemma LIM_fun_gt_zero: + "[| f -- c --> (l::real); 0 < l |] + ==> \r. 0 < r & (\x::real. x \ c & \c - x\ < r --> 0 < f x)" +apply (auto simp add: LIM_def) +apply (drule_tac x = "l/2" in spec, safe, force) +apply (rule_tac x = s in exI) +apply (auto simp only: abs_less_iff) +done + +lemma LIM_fun_less_zero: + "[| f -- c --> (l::real); l < 0 |] + ==> \r. 0 < r & (\x::real. x \ c & \c - x\ < r --> f x < 0)" +apply (auto simp add: LIM_def) +apply (drule_tac x = "-l/2" in spec, safe, force) +apply (rule_tac x = s in exI) +apply (auto simp only: abs_less_iff) +done + + +lemma LIM_fun_not_zero: + "[| f -- c --> (l::real); l \ 0 |] + ==> \r. 0 < r & (\x::real. x \ c & \c - x\ < r --> f x \ 0)" +apply (cut_tac x = l and y = 0 in linorder_less_linear, auto) +apply (drule LIM_fun_less_zero) +apply (drule_tac [3] LIM_fun_gt_zero) +apply force+ +done + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Typerep.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Typerep.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,107 @@ +(* Title: HOL/Library/RType.thy + ID: $Id$ + Author: Florian Haftmann, TU Muenchen +*) + +header {* Reflecting Pure types into HOL *} + +theory Typerep +imports Plain List Code_Message +begin + +datatype typerep = Typerep message_string "typerep list" + +class typerep = + fixes typerep :: "'a\{} itself \ typerep" +begin + +definition + typerep_of :: "'a \ typerep" +where + [simp]: "typerep_of x = typerep TYPE('a)" + +end + +setup {* +let + fun typerep_tr (*"_TYPEREP"*) [ty] = + Lexicon.const @{const_syntax typerep} $ (Lexicon.const "_constrain" $ Lexicon.const "TYPE" $ + (Lexicon.const "itself" $ ty)) + | typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts); + fun typerep_tr' show_sorts (*"typerep"*) + (Type ("fun", [Type ("itself", [T]), _])) (Const (@{const_syntax TYPE}, _) :: ts) = + Term.list_comb (Lexicon.const "_TYPEREP" $ Syntax.term_of_typ show_sorts T, ts) + | typerep_tr' _ T ts = raise Match; +in + Sign.add_syntax_i + [("_TYPEREP", SimpleSyntax.read_typ "type => logic", Delimfix "(1TYPEREP/(1'(_')))")] + #> Sign.add_trfuns ([], [("_TYPEREP", typerep_tr)], [], []) + #> Sign.add_trfunsT [(@{const_syntax typerep}, typerep_tr')] +end +*} + +ML {* +structure Typerep = +struct + +fun mk f (Type (tyco, tys)) = + @{term Typerep} $ Message_String.mk tyco + $ HOLogic.mk_list @{typ typerep} (map (mk f) tys) + | mk f (TFree v) = + f v; + +fun typerep ty = + Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep}) + $ Logic.mk_type ty; + +fun add_def tyco thy = + let + val sorts = replicate (Sign.arity_number thy tyco) @{sort typerep}; + val vs = Name.names Name.context "'a" sorts; + val ty = Type (tyco, map TFree vs); + val lhs = Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep}) + $ Free ("T", Term.itselfT ty); + val rhs = mk (typerep o TFree) ty; + val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); + in + thy + |> TheoryTarget.instantiation ([tyco], vs, @{sort typerep}) + |> `(fn lthy => Syntax.check_term lthy eq) + |-> (fn eq => Specification.definition (NONE, (Attrib.no_binding, eq))) + |> snd + |> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) + |> LocalTheory.exit_global + end; + +fun perhaps_add_def tyco thy = + let + val inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep} + in if inst then thy else add_def tyco thy end; + +end; +*} + +setup {* + Typerep.add_def @{type_name prop} + #> Typerep.add_def @{type_name fun} + #> Typerep.add_def @{type_name itself} + #> Typerep.add_def @{type_name bool} + #> TypedefPackage.interpretation Typerep.perhaps_add_def +*} + +lemma [code]: + "eq_class.eq (Typerep tyco1 tys1) (Typerep tyco2 tys2) \ eq_class.eq tyco1 tyco2 + \ list_all2 eq_class.eq tys1 tys2" + by (auto simp add: equals_eq [symmetric] list_all2_eq [symmetric]) + +code_type typerep + (SML "Term.typ") + +code_const Typerep + (SML "Term.Type/ (_, _)") + +code_reserved SML Term + +hide (open) const typerep Typerep + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Univ_Poly.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Univ_Poly.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,1049 @@ +(* Title: Univ_Poly.thy + Author: Amine Chaieb +*) + +header{*Univariate Polynomials*} + +theory Univ_Poly +imports Plain List +begin + +text{* Application of polynomial as a function. *} + +primrec (in semiring_0) poly :: "'a list => 'a => 'a" where + poly_Nil: "poly [] x = 0" +| poly_Cons: "poly (h#t) x = h + x * poly t x" + + +subsection{*Arithmetic Operations on Polynomials*} + +text{*addition*} + +primrec (in semiring_0) padd :: "'a list \ 'a list \ 'a list" (infixl "+++" 65) +where + padd_Nil: "[] +++ l2 = l2" +| padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t + else (h + hd l2)#(t +++ tl l2))" + +text{*Multiplication by a constant*} +primrec (in semiring_0) cmult :: "'a \ 'a list \ 'a list" (infixl "%*" 70) where + cmult_Nil: "c %* [] = []" +| cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)" + +text{*Multiplication by a polynomial*} +primrec (in semiring_0) pmult :: "'a list \ 'a list \ 'a list" (infixl "***" 70) +where + pmult_Nil: "[] *** l2 = []" +| pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2 + else (h %* l2) +++ ((0) # (t *** l2)))" + +text{*Repeated multiplication by a polynomial*} +primrec (in semiring_0) mulexp :: "nat \ 'a list \ 'a list \ 'a list" where + mulexp_zero: "mulexp 0 p q = q" +| mulexp_Suc: "mulexp (Suc n) p q = p *** mulexp n p q" + +text{*Exponential*} +primrec (in semiring_1) pexp :: "'a list \ nat \ 'a list" (infixl "%^" 80) where + pexp_0: "p %^ 0 = [1]" +| pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)" + +text{*Quotient related value of dividing a polynomial by x + a*} +(* Useful for divisor properties in inductive proofs *) +primrec (in field) "pquot" :: "'a list \ 'a \ 'a list" where + pquot_Nil: "pquot [] a= []" +| pquot_Cons: "pquot (h#t) a = (if t = [] then [h] + else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))" + +text{*normalization of polynomials (remove extra 0 coeff)*} +primrec (in semiring_0) pnormalize :: "'a list \ 'a list" where + pnormalize_Nil: "pnormalize [] = []" +| pnormalize_Cons: "pnormalize (h#p) = (if ( (pnormalize p) = []) + then (if (h = 0) then [] else [h]) + else (h#(pnormalize p)))" + +definition (in semiring_0) "pnormal p = ((pnormalize p = p) \ p \ [])" +definition (in semiring_0) "nonconstant p = (pnormal p \ (\x. p \ [x]))" +text{*Other definitions*} + +definition (in ring_1) + poly_minus :: "'a list => 'a list" ("-- _" [80] 80) where + "-- p = (- 1) %* p" + +definition (in semiring_0) + divides :: "'a list \ 'a list \ bool" (infixl "divides" 70) where + [code del]: "p1 divides p2 = (\q. poly p2 = poly(p1 *** q))" + + --{*order of a polynomial*} +definition (in ring_1) order :: "'a => 'a list => nat" where + "order a p = (SOME n. ([-a, 1] %^ n) divides p & + ~ (([-a, 1] %^ (Suc n)) divides p))" + + --{*degree of a polynomial*} +definition (in semiring_0) degree :: "'a list => nat" where + "degree p = length (pnormalize p) - 1" + + --{*squarefree polynomials --- NB with respect to real roots only.*} +definition (in ring_1) + rsquarefree :: "'a list => bool" where + "rsquarefree p = (poly p \ poly [] & + (\a. (order a p = 0) | (order a p = 1)))" + +context semiring_0 +begin + +lemma padd_Nil2[simp]: "p +++ [] = p" +by (induct p) auto + +lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)" +by auto + +lemma pminus_Nil[simp]: "-- [] = []" +by (simp add: poly_minus_def) + +lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" by simp +end + +lemma (in semiring_1) poly_ident_mult[simp]: "1 %* t = t" by (induct "t", auto) + +lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ ((0)#t) = (a#t)" +by simp + +text{*Handy general properties*} + +lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b" +proof(induct b arbitrary: a) + case Nil thus ?case by auto +next + case (Cons b bs a) thus ?case by (cases a, simp_all add: add_commute) +qed + +lemma (in comm_semiring_0) padd_assoc: "\b c. (a +++ b) +++ c = a +++ (b +++ c)" +apply (induct a arbitrary: b c) +apply (simp, clarify) +apply (case_tac b, simp_all add: add_ac) +done + +lemma (in semiring_0) poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)" +apply (induct p arbitrary: q,simp) +apply (case_tac q, simp_all add: right_distrib) +done + +lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = ((0)#t)" +apply (induct "t", simp) +apply (auto simp add: mult_zero_left poly_ident_mult padd_commut) +apply (case_tac t, auto) +done + +text{*properties of evaluation of polynomials.*} + +lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x" +proof(induct p1 arbitrary: p2) + case Nil thus ?case by simp +next + case (Cons a as p2) thus ?case + by (cases p2, simp_all add: add_ac right_distrib) +qed + +lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x" +apply (induct "p") +apply (case_tac [2] "x=zero") +apply (auto simp add: right_distrib mult_ac) +done + +lemma (in comm_semiring_0) poly_cmult_map: "poly (map (op * c) p) x = c*poly p x" + by (induct p, auto simp add: right_distrib mult_ac) + +lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)" +apply (simp add: poly_minus_def) +apply (auto simp add: poly_cmult minus_mult_left[symmetric]) +done + +lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x" +proof(induct p1 arbitrary: p2) + case Nil thus ?case by simp +next + case (Cons a as p2) + thus ?case by (cases as, + simp_all add: poly_cmult poly_add left_distrib right_distrib mult_ac) +qed + +class recpower_semiring = semiring + recpower +class recpower_semiring_1 = semiring_1 + recpower +class recpower_semiring_0 = semiring_0 + recpower +class recpower_ring = ring + recpower +class recpower_ring_1 = ring_1 + recpower +subclass (in recpower_ring_1) recpower_ring .. +class recpower_comm_semiring_1 = recpower + comm_semiring_1 +class recpower_comm_ring_1 = recpower + comm_ring_1 +subclass (in recpower_comm_ring_1) recpower_comm_semiring_1 .. +class recpower_idom = recpower + idom +subclass (in recpower_idom) recpower_comm_ring_1 .. +class idom_char_0 = idom + ring_char_0 +class recpower_idom_char_0 = recpower + idom_char_0 +subclass (in recpower_idom_char_0) recpower_idom .. + +lemma (in recpower_comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n" +apply (induct "n") +apply (auto simp add: poly_cmult poly_mult power_Suc) +done + +text{*More Polynomial Evaluation Lemmas*} + +lemma (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x" +by simp + +lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x" + by (simp add: poly_mult mult_assoc) + +lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0" +by (induct "p", auto) + +lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x" +apply (induct "n") +apply (auto simp add: poly_mult mult_assoc) +done + +subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides + @{term "p(x)"} *} + +lemma (in comm_ring_1) lemma_poly_linear_rem: "\h. \q r. h#t = [r] +++ [-a, 1] *** q" +proof(induct t) + case Nil + {fix h have "[h] = [h] +++ [- a, 1] *** []" by simp} + thus ?case by blast +next + case (Cons x xs) + {fix h + from Cons.hyps[rule_format, of x] + obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast + have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)" + using qr by(cases q, simp_all add: ring_simps diff_def[symmetric] + minus_mult_left[symmetric] right_minus) + hence "\q r. h#x#xs = [r] +++ [-a, 1] *** q" by blast} + thus ?case by blast +qed + +lemma (in comm_ring_1) poly_linear_rem: "\q r. h#t = [r] +++ [-a, 1] *** q" +by (cut_tac t = t and a = a in lemma_poly_linear_rem, auto) + + +lemma (in comm_ring_1) poly_linear_divides: "(poly p a = 0) = ((p = []) | (\q. p = [-a, 1] *** q))" +proof- + {assume p: "p = []" hence ?thesis by simp} + moreover + {fix x xs assume p: "p = x#xs" + {fix q assume "p = [-a, 1] *** q" hence "poly p a = 0" + by (simp add: poly_add poly_cmult minus_mult_left[symmetric])} + moreover + {assume p0: "poly p a = 0" + from poly_linear_rem[of x xs a] obtain q r + where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast + have "r = 0" using p0 by (simp only: p qr poly_mult poly_add) simp + hence "\q. p = [- a, 1] *** q" using p qr apply - apply (rule exI[where x=q])apply auto apply (cases q) apply auto done} + ultimately have ?thesis using p by blast} + ultimately show ?thesis by (cases p, auto) +qed + +lemma (in semiring_0) lemma_poly_length_mult[simp]: "\h k a. length (k %* p +++ (h # (a %* p))) = Suc (length p)" +by (induct "p", auto) + +lemma (in semiring_0) lemma_poly_length_mult2[simp]: "\h k. length (k %* p +++ (h # p)) = Suc (length p)" +by (induct "p", auto) + +lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)" +by auto + +subsection{*Polynomial length*} + +lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p" +by (induct "p", auto) + +lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)" +apply (induct p1 arbitrary: p2, simp_all) +apply arith +done + +lemma (in semiring_0) poly_root_mult_length[simp]: "length([a,b] *** p) = Suc (length p)" +by (simp add: poly_add_length) + +lemma (in idom) poly_mult_not_eq_poly_Nil[simp]: + "poly (p *** q) x \ poly [] x \ poly p x \ poly [] x \ poly q x \ poly [] x" +by (auto simp add: poly_mult) + +lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \ poly p x = 0 \ poly q x = 0" +by (auto simp add: poly_mult) + +text{*Normalisation Properties*} + +lemma (in semiring_0) poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)" +by (induct "p", auto) + +text{*A nontrivial polynomial of degree n has no more than n roots*} +lemma (in idom) poly_roots_index_lemma: + assumes p: "poly p x \ poly [] x" and n: "length p = n" + shows "\i. \x. poly p x = 0 \ (\m\n. x = i m)" + using p n +proof(induct n arbitrary: p x) + case 0 thus ?case by simp +next + case (Suc n p x) + {assume C: "\i. \x. poly p x = 0 \ (\m\Suc n. x \ i m)" + from Suc.prems have p0: "poly p x \ 0" "p\ []" by auto + from p0(1)[unfolded poly_linear_divides[of p x]] + have "\q. p \ [- x, 1] *** q" by blast + from C obtain a where a: "poly p a = 0" by blast + from a[unfolded poly_linear_divides[of p a]] p0(2) + obtain q where q: "p = [-a, 1] *** q" by blast + have lg: "length q = n" using q Suc.prems(2) by simp + from q p0 have qx: "poly q x \ poly [] x" + by (auto simp add: poly_mult poly_add poly_cmult) + from Suc.hyps[OF qx lg] obtain i where + i: "\x. poly q x = 0 \ (\m\n. x = i m)" by blast + let ?i = "\m. if m = Suc n then a else i m" + from C[of ?i] obtain y where y: "poly p y = 0" "\m\ Suc n. y \ ?i m" + by blast + from y have "y = a \ poly q y = 0" + by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: ring_simps) + with i[rule_format, of y] y(1) y(2) have False apply auto + apply (erule_tac x="m" in allE) + apply auto + done} + thus ?case by blast +qed + + +lemma (in idom) poly_roots_index_length: "poly p x \ poly [] x ==> + \i. \x. (poly p x = 0) --> (\n. n \ length p & x = i n)" +by (blast intro: poly_roots_index_lemma) + +lemma (in idom) poly_roots_finite_lemma1: "poly p x \ poly [] x ==> + \N i. \x. (poly p x = 0) --> (\n. (n::nat) < N & x = i n)" +apply (drule poly_roots_index_length, safe) +apply (rule_tac x = "Suc (length p)" in exI) +apply (rule_tac x = i in exI) +apply (simp add: less_Suc_eq_le) +done + + +lemma (in idom) idom_finite_lemma: + assumes P: "\x. P x --> (\n. n < length j & x = j!n)" + shows "finite {x. P x}" +proof- + let ?M = "{x. P x}" + let ?N = "set j" + have "?M \ ?N" using P by auto + thus ?thesis using finite_subset by auto +qed + + +lemma (in idom) poly_roots_finite_lemma2: "poly p x \ poly [] x ==> + \i. \x. (poly p x = 0) --> x \ set i" +apply (drule poly_roots_index_length, safe) +apply (rule_tac x="map (\n. i n) [0 ..< Suc (length p)]" in exI) +apply (auto simp add: image_iff) +apply (erule_tac x="x" in allE, clarsimp) +by (case_tac "n=length p", auto simp add: order_le_less) + +lemma UNIV_nat_infinite: "\ finite (UNIV :: nat set)" + unfolding finite_conv_nat_seg_image +proof(auto simp add: expand_set_eq image_iff) + fix n::nat and f:: "nat \ nat" + let ?N = "{i. i < n}" + let ?fN = "f ` ?N" + let ?y = "Max ?fN + 1" + from nat_seg_image_imp_finite[of "?fN" "f" n] + have thfN: "finite ?fN" by simp + {assume "n =0" hence "\x. \xa f xa" by auto} + moreover + {assume nz: "n \ 0" + hence thne: "?fN \ {}" by (auto simp add: neq0_conv) + have "\x\ ?fN. Max ?fN \ x" using nz Max_ge_iff[OF thfN thne] by auto + hence "\x\ ?fN. ?y > x" by auto + hence "?y \ ?fN" by auto + hence "\x. \xa f xa" by auto } + ultimately show "\x. \xa f xa" by blast +qed + +lemma (in ring_char_0) UNIV_ring_char_0_infinte: + "\ (finite (UNIV:: 'a set))" +proof + assume F: "finite (UNIV :: 'a set)" + have th0: "of_nat ` UNIV \ UNIV" by simp + from finite_subset[OF th0] have th: "finite (of_nat ` UNIV :: 'a set)" . + have th': "inj_on (of_nat::nat \ 'a) (UNIV)" + unfolding inj_on_def by auto + from finite_imageD[OF th th'] UNIV_nat_infinite + show False by blast +qed + +lemma (in idom_char_0) poly_roots_finite: "(poly p \ poly []) = + finite {x. poly p x = 0}" +proof + assume H: "poly p \ poly []" + show "finite {x. poly p x = (0::'a)}" + using H + apply - + apply (erule contrapos_np, rule ext) + apply (rule ccontr) + apply (clarify dest!: poly_roots_finite_lemma2) + using finite_subset + proof- + fix x i + assume F: "\ finite {x. poly p x = (0\'a)}" + and P: "\x. poly p x = (0\'a) \ x \ set i" + let ?M= "{x. poly p x = (0\'a)}" + from P have "?M \ set i" by auto + with finite_subset F show False by auto + qed +next + assume F: "finite {x. poly p x = (0\'a)}" + show "poly p \ poly []" using F UNIV_ring_char_0_infinte by auto +qed + +text{*Entirety and Cancellation for polynomials*} + +lemma (in idom_char_0) poly_entire_lemma2: + assumes p0: "poly p \ poly []" and q0: "poly q \ poly []" + shows "poly (p***q) \ poly []" +proof- + let ?S = "\p. {x. poly p x = 0}" + have "?S (p *** q) = ?S p \ ?S q" by (auto simp add: poly_mult) + with p0 q0 show ?thesis unfolding poly_roots_finite by auto +qed + +lemma (in idom_char_0) poly_entire: + "poly (p *** q) = poly [] \ poly p = poly [] \ poly q = poly []" +using poly_entire_lemma2[of p q] +by auto (rule ext, simp add: poly_mult)+ + +lemma (in idom_char_0) poly_entire_neg: "(poly (p *** q) \ poly []) = ((poly p \ poly []) & (poly q \ poly []))" +by (simp add: poly_entire) + +lemma fun_eq: " (f = g) = (\x. f x = g x)" +by (auto intro!: ext) + +lemma (in comm_ring_1) poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)" +by (auto simp add: ring_simps poly_add poly_minus_def fun_eq poly_cmult minus_mult_left[symmetric]) + +lemma (in comm_ring_1) poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))" +by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult right_distrib minus_mult_left[symmetric] minus_mult_right[symmetric]) + +subclass (in idom_char_0) comm_ring_1 .. +lemma (in idom_char_0) poly_mult_left_cancel: "(poly (p *** q) = poly (p *** r)) = (poly p = poly [] | poly q = poly r)" +proof- + have "poly (p *** q) = poly (p *** r) \ poly (p *** q +++ -- (p *** r)) = poly []" by (simp only: poly_add_minus_zero_iff) + also have "\ \ poly p = poly [] | poly q = poly r" + by (auto intro: ext simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff) + finally show ?thesis . +qed + +lemma (in recpower_idom) poly_exp_eq_zero[simp]: + "(poly (p %^ n) = poly []) = (poly p = poly [] & n \ 0)" +apply (simp only: fun_eq add: all_simps [symmetric]) +apply (rule arg_cong [where f = All]) +apply (rule ext) +apply (induct n) +apply (auto simp add: poly_exp poly_mult) +done + +lemma (in semiring_1) one_neq_zero[simp]: "1 \ 0" using zero_neq_one by blast +lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a,1] \ poly []" +apply (simp add: fun_eq) +apply (rule_tac x = "minus one a" in exI) +apply (unfold diff_minus) +apply (subst add_commute) +apply (subst add_assoc) +apply simp +done + +lemma (in recpower_idom) poly_exp_prime_eq_zero: "(poly ([a, 1] %^ n) \ poly [])" +by auto + +text{*A more constructive notion of polynomials being trivial*} + +lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] ==> h = 0 & poly t = poly []" +apply(simp add: fun_eq) +apply (case_tac "h = zero") +apply (drule_tac [2] x = zero in spec, auto) +apply (cases "poly t = poly []", simp) +proof- + fix x + assume H: "\x. x = (0\'a) \ poly t x = (0\'a)" and pnz: "poly t \ poly []" + let ?S = "{x. poly t x = 0}" + from H have "\x. x \0 \ poly t x = 0" by blast + hence th: "?S \ UNIV - {0}" by auto + from poly_roots_finite pnz have th': "finite ?S" by blast + from finite_subset[OF th th'] UNIV_ring_char_0_infinte + show "poly t x = (0\'a)" by simp + qed + +lemma (in idom_char_0) poly_zero: "(poly p = poly []) = list_all (%c. c = 0) p" +apply (induct "p", simp) +apply (rule iffI) +apply (drule poly_zero_lemma', auto) +done + +lemma (in idom_char_0) poly_0: "list_all (\c. c = 0) p \ poly p x = 0" + unfolding poly_zero[symmetric] by simp + + + +text{*Basics of divisibility.*} + +lemma (in idom) poly_primes: "([a, 1] divides (p *** q)) = ([a, 1] divides p | [a, 1] divides q)" +apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult left_distrib [symmetric]) +apply (drule_tac x = "uminus a" in spec) +apply (simp add: poly_linear_divides poly_add poly_cmult left_distrib [symmetric]) +apply (cases "p = []") +apply (rule exI[where x="[]"]) +apply simp +apply (cases "q = []") +apply (erule allE[where x="[]"], simp) + +apply clarsimp +apply (cases "\q\'a list. p = a %* q +++ ((0\'a) # q)") +apply (clarsimp simp add: poly_add poly_cmult) +apply (rule_tac x="qa" in exI) +apply (simp add: left_distrib [symmetric]) +apply clarsimp + +apply (auto simp add: right_minus poly_linear_divides poly_add poly_cmult left_distrib [symmetric]) +apply (rule_tac x = "pmult qa q" in exI) +apply (rule_tac [2] x = "pmult p qa" in exI) +apply (auto simp add: poly_add poly_mult poly_cmult mult_ac) +done + +lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p" +apply (simp add: divides_def) +apply (rule_tac x = "[one]" in exI) +apply (auto simp add: poly_mult fun_eq) +done + +lemma (in comm_semiring_1) poly_divides_trans: "[| p divides q; q divides r |] ==> p divides r" +apply (simp add: divides_def, safe) +apply (rule_tac x = "pmult qa qaa" in exI) +apply (auto simp add: poly_mult fun_eq mult_assoc) +done + + +lemma (in recpower_comm_semiring_1) poly_divides_exp: "m \ n ==> (p %^ m) divides (p %^ n)" +apply (auto simp add: le_iff_add) +apply (induct_tac k) +apply (rule_tac [2] poly_divides_trans) +apply (auto simp add: divides_def) +apply (rule_tac x = p in exI) +apply (auto simp add: poly_mult fun_eq mult_ac) +done + +lemma (in recpower_comm_semiring_1) poly_exp_divides: "[| (p %^ n) divides q; m\n |] ==> (p %^ m) divides q" +by (blast intro: poly_divides_exp poly_divides_trans) + +lemma (in comm_semiring_0) poly_divides_add: + "[| p divides q; p divides r |] ==> p divides (q +++ r)" +apply (simp add: divides_def, auto) +apply (rule_tac x = "padd qa qaa" in exI) +apply (auto simp add: poly_add fun_eq poly_mult right_distrib) +done + +lemma (in comm_ring_1) poly_divides_diff: + "[| p divides q; p divides (q +++ r) |] ==> p divides r" +apply (simp add: divides_def, auto) +apply (rule_tac x = "padd qaa (poly_minus qa)" in exI) +apply (auto simp add: poly_add fun_eq poly_mult poly_minus right_diff_distrib compare_rls add_ac) +done + +lemma (in comm_ring_1) poly_divides_diff2: "[| p divides r; p divides (q +++ r) |] ==> p divides q" +apply (erule poly_divides_diff) +apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac) +done + +lemma (in semiring_0) poly_divides_zero: "poly p = poly [] ==> q divides p" +apply (simp add: divides_def) +apply (rule exI[where x="[]"]) +apply (auto simp add: fun_eq poly_mult) +done + +lemma (in semiring_0) poly_divides_zero2[simp]: "q divides []" +apply (simp add: divides_def) +apply (rule_tac x = "[]" in exI) +apply (auto simp add: fun_eq) +done + +text{*At last, we can consider the order of a root.*} + +lemma (in idom_char_0) poly_order_exists_lemma: + assumes lp: "length p = d" and p: "poly p \ poly []" + shows "\n q. p = mulexp n [-a, 1] q \ poly q a \ 0" +using lp p +proof(induct d arbitrary: p) + case 0 thus ?case by simp +next + case (Suc n p) + {assume p0: "poly p a = 0" + from Suc.prems have h: "length p = Suc n" "poly p \ poly []" by blast + hence pN: "p \ []" by - (rule ccontr, simp) + from p0[unfolded poly_linear_divides] pN obtain q where + q: "p = [-a, 1] *** q" by blast + from q h p0 have qh: "length q = n" "poly q \ poly []" + apply - + apply simp + apply (simp only: fun_eq) + apply (rule ccontr) + apply (simp add: fun_eq poly_add poly_cmult minus_mult_left[symmetric]) + done + from Suc.hyps[OF qh] obtain m r where + mr: "q = mulexp m [-a,1] r" "poly r a \ 0" by blast + from mr q have "p = mulexp (Suc m) [-a,1] r \ poly r a \ 0" by simp + hence ?case by blast} + moreover + {assume p0: "poly p a \ 0" + hence ?case using Suc.prems apply simp by (rule exI[where x="0::nat"], simp)} + ultimately show ?case by blast +qed + + +lemma (in recpower_comm_semiring_1) poly_mulexp: "poly (mulexp n p q) x = (poly p x) ^ n * poly q x" +by(induct n, auto simp add: poly_mult power_Suc mult_ac) + +lemma (in comm_semiring_1) divides_left_mult: + assumes d:"(p***q) divides r" shows "p divides r \ q divides r" +proof- + from d obtain t where r:"poly r = poly (p***q *** t)" + unfolding divides_def by blast + hence "poly r = poly (p *** (q *** t))" + "poly r = poly (q *** (p***t))" by(auto simp add: fun_eq poly_mult mult_ac) + thus ?thesis unfolding divides_def by blast +qed + + + +(* FIXME: Tidy up *) + +lemma (in recpower_semiring_1) + zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)" + by (induct n, simp_all add: power_Suc) + +lemma (in recpower_idom_char_0) poly_order_exists: + assumes lp: "length p = d" and p0: "poly p \ poly []" + shows "\n. ([-a, 1] %^ n) divides p & ~(([-a, 1] %^ (Suc n)) divides p)" +proof- +let ?poly = poly +let ?mulexp = mulexp +let ?pexp = pexp +from lp p0 +show ?thesis +apply - +apply (drule poly_order_exists_lemma [where a=a], assumption, clarify) +apply (rule_tac x = n in exI, safe) +apply (unfold divides_def) +apply (rule_tac x = q in exI) +apply (induct_tac "n", simp) +apply (simp (no_asm_simp) add: poly_add poly_cmult poly_mult right_distrib mult_ac) +apply safe +apply (subgoal_tac "?poly (?mulexp n [uminus a, one] q) \ ?poly (pmult (?pexp [uminus a, one] (Suc n)) qa)") +apply simp +apply (induct_tac "n") +apply (simp del: pmult_Cons pexp_Suc) +apply (erule_tac Q = "?poly q a = zero" in contrapos_np) +apply (simp add: poly_add poly_cmult minus_mult_left[symmetric]) +apply (rule pexp_Suc [THEN ssubst]) +apply (rule ccontr) +apply (simp add: poly_mult_left_cancel poly_mult_assoc del: pmult_Cons pexp_Suc) +done +qed + + +lemma (in semiring_1) poly_one_divides[simp]: "[1] divides p" +by (simp add: divides_def, auto) + +lemma (in recpower_idom_char_0) poly_order: "poly p \ poly [] + ==> EX! n. ([-a, 1] %^ n) divides p & + ~(([-a, 1] %^ (Suc n)) divides p)" +apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc) +apply (cut_tac x = y and y = n in less_linear) +apply (drule_tac m = n in poly_exp_divides) +apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides] + simp del: pmult_Cons pexp_Suc) +done + +text{*Order*} + +lemma some1_equalityD: "[| n = (@n. P n); EX! n. P n |] ==> P n" +by (blast intro: someI2) + +lemma (in recpower_idom_char_0) order: + "(([-a, 1] %^ n) divides p & + ~(([-a, 1] %^ (Suc n)) divides p)) = + ((n = order a p) & ~(poly p = poly []))" +apply (unfold order_def) +apply (rule iffI) +apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order) +apply (blast intro!: poly_order [THEN [2] some1_equalityD]) +done + +lemma (in recpower_idom_char_0) order2: "[| poly p \ poly [] |] + ==> ([-a, 1] %^ (order a p)) divides p & + ~(([-a, 1] %^ (Suc(order a p))) divides p)" +by (simp add: order del: pexp_Suc) + +lemma (in recpower_idom_char_0) order_unique: "[| poly p \ poly []; ([-a, 1] %^ n) divides p; + ~(([-a, 1] %^ (Suc n)) divides p) + |] ==> (n = order a p)" +by (insert order [of a n p], auto) + +lemma (in recpower_idom_char_0) order_unique_lemma: "(poly p \ poly [] & ([-a, 1] %^ n) divides p & + ~(([-a, 1] %^ (Suc n)) divides p)) + ==> (n = order a p)" +by (blast intro: order_unique) + +lemma (in ring_1) order_poly: "poly p = poly q ==> order a p = order a q" +by (auto simp add: fun_eq divides_def poly_mult order_def) + +lemma (in semiring_1) pexp_one[simp]: "p %^ (Suc 0) = p" +apply (induct "p") +apply (auto simp add: numeral_1_eq_1) +done + +lemma (in comm_ring_1) lemma_order_root: + " 0 < n & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p + \ poly p a = 0" +apply (induct n arbitrary: a p, blast) +apply (auto simp add: divides_def poly_mult simp del: pmult_Cons) +done + +lemma (in recpower_idom_char_0) order_root: "(poly p a = 0) = ((poly p = poly []) | order a p \ 0)" +proof- + let ?poly = poly + show ?thesis +apply (case_tac "?poly p = ?poly []", auto) +apply (simp add: poly_linear_divides del: pmult_Cons, safe) +apply (drule_tac [!] a = a in order2) +apply (rule ccontr) +apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast) +using neq0_conv +apply (blast intro: lemma_order_root) +done +qed + +lemma (in recpower_idom_char_0) order_divides: "(([-a, 1] %^ n) divides p) = ((poly p = poly []) | n \ order a p)" +proof- + let ?poly = poly + show ?thesis +apply (case_tac "?poly p = ?poly []", auto) +apply (simp add: divides_def fun_eq poly_mult) +apply (rule_tac x = "[]" in exI) +apply (auto dest!: order2 [where a=a] + intro: poly_exp_divides simp del: pexp_Suc) +done +qed + +lemma (in recpower_idom_char_0) order_decomp: + "poly p \ poly [] + ==> \q. (poly p = poly (([-a, 1] %^ (order a p)) *** q)) & + ~([-a, 1] divides q)" +apply (unfold divides_def) +apply (drule order2 [where a = a]) +apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe) +apply (rule_tac x = q in exI, safe) +apply (drule_tac x = qa in spec) +apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons) +done + +text{*Important composition properties of orders.*} +lemma order_mult: "poly (p *** q) \ poly [] + ==> order a (p *** q) = order a p + order (a::'a::{recpower_idom_char_0}) q" +apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order) +apply (auto simp add: poly_entire simp del: pmult_Cons) +apply (drule_tac a = a in order2)+ +apply safe +apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe) +apply (rule_tac x = "qa *** qaa" in exI) +apply (simp add: poly_mult mult_ac del: pmult_Cons) +apply (drule_tac a = a in order_decomp)+ +apply safe +apply (subgoal_tac "[-a,1] divides (qa *** qaa) ") +apply (simp add: poly_primes del: pmult_Cons) +apply (auto simp add: divides_def simp del: pmult_Cons) +apply (rule_tac x = qb in exI) +apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))") +apply (drule poly_mult_left_cancel [THEN iffD1], force) +apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ") +apply (drule poly_mult_left_cancel [THEN iffD1], force) +apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons) +done + +lemma (in recpower_idom_char_0) order_mult: + assumes pq0: "poly (p *** q) \ poly []" + shows "order a (p *** q) = order a p + order a q" +proof- + let ?order = order + let ?divides = "op divides" + let ?poly = poly +from pq0 +show ?thesis +apply (cut_tac a = a and p = "pmult p q" and n = "?order a p + ?order a q" in order) +apply (auto simp add: poly_entire simp del: pmult_Cons) +apply (drule_tac a = a in order2)+ +apply safe +apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe) +apply (rule_tac x = "pmult qa qaa" in exI) +apply (simp add: poly_mult mult_ac del: pmult_Cons) +apply (drule_tac a = a in order_decomp)+ +apply safe +apply (subgoal_tac "?divides [uminus a,one ] (pmult qa qaa) ") +apply (simp add: poly_primes del: pmult_Cons) +apply (auto simp add: divides_def simp del: pmult_Cons) +apply (rule_tac x = qb in exI) +apply (subgoal_tac "?poly (pmult (pexp [uminus a, one] (?order a p)) (pmult qa qaa)) = ?poly (pmult (pexp [uminus a, one] (?order a p)) (pmult [uminus a, one] qb))") +apply (drule poly_mult_left_cancel [THEN iffD1], force) +apply (subgoal_tac "?poly (pmult (pexp [uminus a, one ] (order a q)) (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) = ?poly (pmult (pexp [uminus a, one] (order a q)) (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb))) ") +apply (drule poly_mult_left_cancel [THEN iffD1], force) +apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons) +done +qed + +lemma (in recpower_idom_char_0) order_root2: "poly p \ poly [] ==> (poly p a = 0) = (order a p \ 0)" +by (rule order_root [THEN ssubst], auto) + +lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p" by auto + +lemma (in semiring_0) poly_Nil_zero: "poly [] = poly [0]" +by (simp add: fun_eq) + +lemma (in recpower_idom_char_0) rsquarefree_decomp: + "[| rsquarefree p; poly p a = 0 |] + ==> \q. (poly p = poly ([-a, 1] *** q)) & poly q a \ 0" +apply (simp add: rsquarefree_def, safe) +apply (frule_tac a = a in order_decomp) +apply (drule_tac x = a in spec) +apply (drule_tac a = a in order_root2 [symmetric]) +apply (auto simp del: pmult_Cons) +apply (rule_tac x = q in exI, safe) +apply (simp add: poly_mult fun_eq) +apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1]) +apply (simp add: divides_def del: pmult_Cons, safe) +apply (drule_tac x = "[]" in spec) +apply (auto simp add: fun_eq) +done + + +text{*Normalization of a polynomial.*} + +lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p" +apply (induct "p") +apply (auto simp add: fun_eq) +done + +text{*The degree of a polynomial.*} + +lemma (in semiring_0) lemma_degree_zero: + "list_all (%c. c = 0) p \ pnormalize p = []" +by (induct "p", auto) + +lemma (in idom_char_0) degree_zero: + assumes pN: "poly p = poly []" shows"degree p = 0" +proof- + let ?pn = pnormalize + from pN + show ?thesis + apply (simp add: degree_def) + apply (case_tac "?pn p = []") + apply (auto simp add: poly_zero lemma_degree_zero ) + done +qed + +lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \ x \ 0" by simp +lemma (in semiring_0) pnormalize_pair: "y \ 0 \ (pnormalize [x, y] = [x, y])" by simp +lemma (in semiring_0) pnormal_cons: "pnormal p \ pnormal (c#p)" + unfolding pnormal_def by simp +lemma (in semiring_0) pnormal_tail: "p\[] \ pnormal (c#p) \ pnormal p" + unfolding pnormal_def + apply (cases "pnormalize p = []", auto) + by (cases "c = 0", auto) + + +lemma (in semiring_0) pnormal_last_nonzero: "pnormal p ==> last p \ 0" +proof(induct p) + case Nil thus ?case by (simp add: pnormal_def) +next + case (Cons a as) thus ?case + apply (simp add: pnormal_def) + apply (cases "pnormalize as = []", simp_all) + apply (cases "as = []", simp_all) + apply (cases "a=0", simp_all) + apply (cases "a=0", simp_all) + done +qed + +lemma (in semiring_0) pnormal_length: "pnormal p \ 0 < length p" + unfolding pnormal_def length_greater_0_conv by blast + +lemma (in semiring_0) pnormal_last_length: "\0 < length p ; last p \ 0\ \ pnormal p" + apply (induct p, auto) + apply (case_tac "p = []", auto) + apply (simp add: pnormal_def) + by (rule pnormal_cons, auto) + +lemma (in semiring_0) pnormal_id: "pnormal p \ (0 < length p \ last p \ 0)" + using pnormal_last_length pnormal_length pnormal_last_nonzero by blast + +lemma (in idom_char_0) poly_Cons_eq: "poly (c#cs) = poly (d#ds) \ c=d \ poly cs = poly ds" (is "?lhs \ ?rhs") +proof + assume eq: ?lhs + hence "\x. poly ((c#cs) +++ -- (d#ds)) x = 0" + by (simp only: poly_minus poly_add ring_simps) simp + hence "poly ((c#cs) +++ -- (d#ds)) = poly []" by - (rule ext, simp) + hence "c = d \ list_all (\x. x=0) ((cs +++ -- ds))" + unfolding poly_zero by (simp add: poly_minus_def ring_simps minus_mult_left[symmetric]) + hence "c = d \ (\x. poly (cs +++ -- ds) x = 0)" + unfolding poly_zero[symmetric] by simp + thus ?rhs apply (simp add: poly_minus poly_add ring_simps) apply (rule ext, simp) done +next + assume ?rhs then show ?lhs by - (rule ext,simp) +qed + +lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \ pnormalize p = pnormalize q" +proof(induct q arbitrary: p) + case Nil thus ?case by (simp only: poly_zero lemma_degree_zero) simp +next + case (Cons c cs p) + thus ?case + proof(induct p) + case Nil + hence "poly [] = poly (c#cs)" by blast + then have "poly (c#cs) = poly [] " by simp + thus ?case by (simp only: poly_zero lemma_degree_zero) simp + next + case (Cons d ds) + hence eq: "poly (d # ds) = poly (c # cs)" by blast + hence eq': "\x. poly (d # ds) x = poly (c # cs) x" by simp + hence "poly (d # ds) 0 = poly (c # cs) 0" by blast + hence dc: "d = c" by auto + with eq have "poly ds = poly cs" + unfolding poly_Cons_eq by simp + with Cons.prems have "pnormalize ds = pnormalize cs" by blast + with dc show ?case by simp + qed +qed + +lemma (in idom_char_0) degree_unique: assumes pq: "poly p = poly q" + shows "degree p = degree q" +using pnormalize_unique[OF pq] unfolding degree_def by simp + +lemma (in semiring_0) pnormalize_length: "length (pnormalize p) \ length p" by (induct p, auto) + +lemma (in semiring_0) last_linear_mul_lemma: + "last ((a %* p) +++ (x#(b %* p))) = (if p=[] then x else b*last p)" + +apply (induct p arbitrary: a x b, auto) +apply (subgoal_tac "padd (cmult aa p) (times b a # cmult b p) \ []", simp) +apply (induct_tac p, auto) +done + +lemma (in semiring_1) last_linear_mul: assumes p:"p\[]" shows "last ([a,1] *** p) = last p" +proof- + from p obtain c cs where cs: "p = c#cs" by (cases p, auto) + from cs have eq:"[a,1] *** p = (a %* (c#cs)) +++ (0#(1 %* (c#cs)))" + by (simp add: poly_cmult_distr) + show ?thesis using cs + unfolding eq last_linear_mul_lemma by simp +qed + +lemma (in semiring_0) pnormalize_eq: "last p \ 0 \ pnormalize p = p" + apply (induct p, auto) + apply (case_tac p, auto)+ + done + +lemma (in semiring_0) last_pnormalize: "pnormalize p \ [] \ last (pnormalize p) \ 0" + by (induct p, auto) + +lemma (in semiring_0) pnormal_degree: "last p \ 0 \ degree p = length p - 1" + using pnormalize_eq[of p] unfolding degree_def by simp + +lemma (in semiring_0) poly_Nil_ext: "poly [] = (\x. 0)" by (rule ext) simp + +lemma (in idom_char_0) linear_mul_degree: assumes p: "poly p \ poly []" + shows "degree ([a,1] *** p) = degree p + 1" +proof- + from p have pnz: "pnormalize p \ []" + unfolding poly_zero lemma_degree_zero . + + from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz] + have l0: "last ([a, 1] *** pnormalize p) \ 0" by simp + from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a] + pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz + + + have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1" + by (auto simp add: poly_length_mult) + + have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)" + by (rule ext) (simp add: poly_mult poly_add poly_cmult) + from degree_unique[OF eqs] th + show ?thesis by (simp add: degree_unique[OF poly_normalize]) +qed + +lemma (in idom_char_0) linear_pow_mul_degree: + "degree([a,1] %^n *** p) = (if poly p = poly [] then 0 else degree p + n)" +proof(induct n arbitrary: a p) + case (0 a p) + {assume p: "poly p = poly []" + hence ?case using degree_unique[OF p] by (simp add: degree_def)} + moreover + {assume p: "poly p \ poly []" hence ?case by (auto simp add: poly_Nil_ext) } + ultimately show ?case by blast +next + case (Suc n a p) + have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1] %^ n *** ([a,1] *** p))" + apply (rule ext, simp add: poly_mult poly_add poly_cmult) + by (simp add: mult_ac add_ac right_distrib) + note deq = degree_unique[OF eq] + {assume p: "poly p = poly []" + with eq have eq': "poly ([a,1] %^(Suc n) *** p) = poly []" + by - (rule ext,simp add: poly_mult poly_cmult poly_add) + from degree_unique[OF eq'] p have ?case by (simp add: degree_def)} + moreover + {assume p: "poly p \ poly []" + from p have ap: "poly ([a,1] *** p) \ poly []" + using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto + have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1]%^n *** ([a,1] *** p))" + by (rule ext, simp add: poly_mult poly_add poly_exp poly_cmult mult_ac add_ac right_distrib) + from ap have ap': "(poly ([a,1] *** p) = poly []) = False" by blast + have th0: "degree ([a,1]%^n *** ([a,1] *** p)) = degree ([a,1] *** p) + n" + apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap') + by simp + + from degree_unique[OF eq] ap p th0 linear_mul_degree[OF p, of a] + have ?case by (auto simp del: poly.simps)} + ultimately show ?case by blast +qed + +lemma (in recpower_idom_char_0) order_degree: + assumes p0: "poly p \ poly []" + shows "order a p \ degree p" +proof- + from order2[OF p0, unfolded divides_def] + obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" by blast + {assume "poly q = poly []" + with q p0 have False by (simp add: poly_mult poly_entire)} + with degree_unique[OF q, unfolded linear_pow_mul_degree] + show ?thesis by auto +qed + +text{*Tidier versions of finiteness of roots.*} + +lemma (in idom_char_0) poly_roots_finite_set: "poly p \ poly [] ==> finite {x. poly p x = 0}" +unfolding poly_roots_finite . + +text{*bound for polynomial.*} + +lemma poly_mono: "abs(x) \ k ==> abs(poly p (x::'a::{ordered_idom})) \ poly (map abs p) k" +apply (induct "p", auto) +apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans) +apply (rule abs_triangle_ineq) +apply (auto intro!: mult_mono simp add: abs_mult) +done + +lemma (in semiring_0) poly_Sing: "poly [c] x = c" by simp + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Word/Num_Lemmas.thy --- a/src/HOL/Word/Num_Lemmas.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Word/Num_Lemmas.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* - ID: $Id$ Author: Jeremy Dawson, NICTA *) diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/Word/ROOT.ML --- a/src/HOL/Word/ROOT.ML Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/Word/ROOT.ML Wed Dec 03 15:58:44 2008 +0100 @@ -1,2 +1,2 @@ -no_document use_thys ["Infinite_Set", "Parity"]; +no_document use_thys ["Infinite_Set"]; use_thy "WordMain"; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/arith_data.ML --- a/src/HOL/arith_data.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,157 +0,0 @@ -(* Title: HOL/arith_data.ML - ID: $Id$ - Author: Markus Wenzel, Stefan Berghofer, and Tobias Nipkow - -Basic arithmetic proof tools. -*) - -signature ARITH_DATA = -sig - val prove_conv: tactic -> (MetaSimplifier.simpset -> tactic) - -> MetaSimplifier.simpset -> term * term -> thm - val simp_all_tac: thm list -> MetaSimplifier.simpset -> tactic - - val mk_sum: term list -> term - val mk_norm_sum: term list -> term - val dest_sum: term -> term list - - val nat_cancel_sums_add: simproc list - val nat_cancel_sums: simproc list - val setup: Context.generic -> Context.generic -end; - -structure ArithData: ARITH_DATA = -struct - -(** generic proof tools **) - -(* prove conversions *) - -fun prove_conv expand_tac norm_tac ss tu = (* FIXME avoid standard *) - mk_meta_eq (standard (Goal.prove (Simplifier.the_context ss) [] [] - (HOLogic.mk_Trueprop (HOLogic.mk_eq tu)) - (K (EVERY [expand_tac, norm_tac ss])))); - -(* rewriting *) - -fun simp_all_tac rules = - let val ss0 = HOL_ss addsimps rules - in fn ss => ALLGOALS (simp_tac (Simplifier.inherit_context ss ss0)) end; - - -(** abstract syntax of structure nat: 0, Suc, + **) - -local - -val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; -val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} HOLogic.natT; - -in - -fun mk_sum [] = HOLogic.zero - | mk_sum [t] = t - | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); - -(*normal form of sums: Suc (... (Suc (a + (b + ...))))*) -fun mk_norm_sum ts = - let val (ones, sums) = List.partition (equal HOLogic.Suc_zero) ts in - funpow (length ones) HOLogic.mk_Suc (mk_sum sums) - end; - - -fun dest_sum tm = - if HOLogic.is_zero tm then [] - else - (case try HOLogic.dest_Suc tm of - SOME t => HOLogic.Suc_zero :: dest_sum t - | NONE => - (case try dest_plus tm of - SOME (t, u) => dest_sum t @ dest_sum u - | NONE => [tm])); - -end; - - -(** cancel common summands **) - -structure Sum = -struct - val mk_sum = mk_norm_sum; - val dest_sum = dest_sum; - val prove_conv = prove_conv; - val norm_tac1 = simp_all_tac [@{thm "add_Suc"}, @{thm "add_Suc_right"}, - @{thm "add_0"}, @{thm "add_0_right"}]; - val norm_tac2 = simp_all_tac @{thms add_ac}; - fun norm_tac ss = norm_tac1 ss THEN norm_tac2 ss; -end; - -fun gen_uncancel_tac rule ct = - rtac (instantiate' [] [NONE, SOME ct] (rule RS @{thm subst_equals})) 1; - - -(* nat eq *) - -structure EqCancelSums = CancelSumsFun -(struct - open Sum; - val mk_bal = HOLogic.mk_eq; - val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT; - val uncancel_tac = gen_uncancel_tac @{thm "nat_add_left_cancel"}; -end); - - -(* nat less *) - -structure LessCancelSums = CancelSumsFun -(struct - open Sum; - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less}; - val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT; - val uncancel_tac = gen_uncancel_tac @{thm "nat_add_left_cancel_less"}; -end); - - -(* nat le *) - -structure LeCancelSums = CancelSumsFun -(struct - open Sum; - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq}; - val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT; - val uncancel_tac = gen_uncancel_tac @{thm "nat_add_left_cancel_le"}; -end); - - -(* nat diff *) - -structure DiffCancelSums = CancelSumsFun -(struct - open Sum; - val mk_bal = HOLogic.mk_binop @{const_name HOL.minus}; - val dest_bal = HOLogic.dest_bin @{const_name HOL.minus} HOLogic.natT; - val uncancel_tac = gen_uncancel_tac @{thm "diff_cancel"}; -end); - - -(* prepare nat_cancel simprocs *) - -val nat_cancel_sums_add = - [Simplifier.simproc (the_context ()) "nateq_cancel_sums" - ["(l::nat) + m = n", "(l::nat) = m + n", "Suc m = n", "m = Suc n"] - (K EqCancelSums.proc), - Simplifier.simproc (the_context ()) "natless_cancel_sums" - ["(l::nat) + m < n", "(l::nat) < m + n", "Suc m < n", "m < Suc n"] - (K LessCancelSums.proc), - Simplifier.simproc (the_context ()) "natle_cancel_sums" - ["(l::nat) + m <= n", "(l::nat) <= m + n", "Suc m <= n", "m <= Suc n"] - (K LeCancelSums.proc)]; - -val nat_cancel_sums = nat_cancel_sums_add @ - [Simplifier.simproc (the_context ()) "natdiff_cancel_sums" - ["((l::nat) + m) - n", "(l::nat) - (m + n)", "Suc m - n", "m - Suc n"] - (K DiffCancelSums.proc)]; - -val setup = - Simplifier.map_ss (fn ss => ss addsimprocs nat_cancel_sums); - -end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/Arithmetic_Series_Complex.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Arithmetic_Series_Complex.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,23 @@ +(* Title: HOL/ex/Arithmetic_Series_Complex + Author: Benjamin Porter, 2006 +*) + + +header {* Arithmetic Series for Reals *} + +theory Arithmetic_Series_Complex +imports Complex_Main +begin + +lemma arith_series_real: + "(2::real) * (\i\{..i\{.. g ----> 0 ==> f ----> (0::real)" + apply (simp add: LIMSEQ_def bigo_alt_def) + apply clarify + apply (drule_tac x = "r / c" in spec) + apply (drule mp) + apply (erule divide_pos_pos) + apply assumption + apply clarify + apply (rule_tac x = no in exI) + apply (rule allI) + apply (drule_tac x = n in spec)+ + apply (rule impI) + apply (drule mp) + apply assumption + apply (rule order_le_less_trans) + apply assumption + apply (rule order_less_le_trans) + apply (subgoal_tac "c * abs(g n) < c * (r / c)") + apply assumption + apply (erule mult_strict_left_mono) + apply assumption + apply simp +done + +lemma bigo_LIMSEQ2: "f =o g +o O(h) ==> h ----> 0 ==> f ----> a + ==> g ----> (a::real)" + apply (drule set_plus_imp_minus) + apply (drule bigo_LIMSEQ1) + apply assumption + apply (simp only: fun_diff_def) + apply (erule LIMSEQ_diff_approach_zero2) + apply assumption +done + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/BinEx.thy --- a/src/HOL/ex/BinEx.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/ex/BinEx.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,12 +1,13 @@ (* Title: HOL/ex/BinEx.thy - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge *) header {* Binary arithmetic examples *} -theory BinEx imports Main begin +theory BinEx +imports Complex_Main +begin subsection {* Regression Testing for Cancellation Simprocs *} @@ -387,4 +388,386 @@ lemma "x + y - x + z - x - y - z + x < (1::int)" by simp + +subsection{*Real Arithmetic*} + +subsubsection {*Addition *} + +lemma "(1359::real) + -2468 = -1109" +by simp + +lemma "(93746::real) + -46375 = 47371" +by simp + + +subsubsection {*Negation *} + +lemma "- (65745::real) = -65745" +by simp + +lemma "- (-54321::real) = 54321" +by simp + + +subsubsection {*Multiplication *} + +lemma "(-84::real) * 51 = -4284" +by simp + +lemma "(255::real) * 255 = 65025" +by simp + +lemma "(1359::real) * -2468 = -3354012" +by simp + + +subsubsection {*Inequalities *} + +lemma "(89::real) * 10 \ 889" +by simp + +lemma "(13::real) < 18 - 4" +by simp + +lemma "(-345::real) < -242 + -100" +by simp + +lemma "(13557456::real) < 18678654" +by simp + +lemma "(999999::real) \ (1000001 + 1) - 2" +by simp + +lemma "(1234567::real) \ 1234567" +by simp + + +subsubsection {*Powers *} + +lemma "2 ^ 15 = (32768::real)" +by simp + +lemma "-3 ^ 7 = (-2187::real)" +by simp + +lemma "13 ^ 7 = (62748517::real)" +by simp + +lemma "3 ^ 15 = (14348907::real)" +by simp + +lemma "-5 ^ 11 = (-48828125::real)" +by simp + + +subsubsection {*Tests *} + +lemma "(x + y = x) = (y = (0::real))" +by arith + +lemma "(x + y = y) = (x = (0::real))" +by arith + +lemma "(x + y = (0::real)) = (x = -y)" +by arith + +lemma "(x + y = (0::real)) = (y = -x)" +by arith + +lemma "((x + y) < (x + z)) = (y < (z::real))" +by arith + +lemma "((x + z) < (y + z)) = (x < (y::real))" +by arith + +lemma "(\ x < y) = (y \ (x::real))" +by arith + +lemma "\ (x < y \ y < (x::real))" +by arith + +lemma "(x::real) < y ==> \ y < x" +by arith + +lemma "((x::real) \ y) = (x < y \ y < x)" +by arith + +lemma "(\ x \ y) = (y < (x::real))" +by arith + +lemma "x \ y \ y \ (x::real)" +by arith + +lemma "x \ y \ y < (x::real)" +by arith + +lemma "x < y \ y \ (x::real)" +by arith + +lemma "x \ (x::real)" +by arith + +lemma "((x::real) \ y) = (x < y \ x = y)" +by arith + +lemma "((x::real) \ y \ y \ x) = (x = y)" +by arith + +lemma "\(x < y \ y \ (x::real))" +by arith + +lemma "\(x \ y \ y < (x::real))" +by arith + +lemma "(-x < (0::real)) = (0 < x)" +by arith + +lemma "((0::real) < -x) = (x < 0)" +by arith + +lemma "(-x \ (0::real)) = (0 \ x)" +by arith + +lemma "((0::real) \ -x) = (x \ 0)" +by arith + +lemma "(x::real) = y \ x < y \ y < x" +by arith + +lemma "(x::real) = 0 \ 0 < x \ 0 < -x" +by arith + +lemma "(0::real) \ x \ 0 \ -x" +by arith + +lemma "((x::real) + y \ x + z) = (y \ z)" +by arith + +lemma "((x::real) + z \ y + z) = (x \ y)" +by arith + +lemma "(w::real) < x \ y < z ==> w + y < x + z" +by arith + +lemma "(w::real) \ x \ y \ z ==> w + y \ x + z" +by arith + +lemma "(0::real) \ x \ 0 \ y ==> 0 \ x + y" +by arith + +lemma "(0::real) < x \ 0 < y ==> 0 < x + y" +by arith + +lemma "(-x < y) = (0 < x + (y::real))" +by arith + +lemma "(x < -y) = (x + y < (0::real))" +by arith + +lemma "(y < x + -z) = (y + z < (x::real))" +by arith + +lemma "(x + -y < z) = (x < z + (y::real))" +by arith + +lemma "x \ y ==> x < y + (1::real)" +by arith + +lemma "(x - y) + y = (x::real)" +by arith + +lemma "y + (x - y) = (x::real)" +by arith + +lemma "x - x = (0::real)" +by arith + +lemma "(x - y = 0) = (x = (y::real))" +by arith + +lemma "((0::real) \ x + x) = (0 \ x)" +by arith + +lemma "(-x \ x) = ((0::real) \ x)" +by arith + +lemma "(x \ -x) = (x \ (0::real))" +by arith + +lemma "(-x = (0::real)) = (x = 0)" +by arith + +lemma "-(x - y) = y - (x::real)" +by arith + +lemma "((0::real) < x - y) = (y < x)" +by arith + +lemma "((0::real) \ x - y) = (y \ x)" +by arith + +lemma "(x + y) - x = (y::real)" +by arith + +lemma "(-x = y) = (x = (-y::real))" +by arith + +lemma "x < (y::real) ==> \(x = y)" +by arith + +lemma "(x \ x + y) = ((0::real) \ y)" +by arith + +lemma "(y \ x + y) = ((0::real) \ x)" +by arith + +lemma "(x < x + y) = ((0::real) < y)" +by arith + +lemma "(y < x + y) = ((0::real) < x)" +by arith + +lemma "(x - y) - x = (-y::real)" +by arith + +lemma "(x + y < z) = (x < z - (y::real))" +by arith + +lemma "(x - y < z) = (x < z + (y::real))" +by arith + +lemma "(x < y - z) = (x + z < (y::real))" +by arith + +lemma "(x \ y - z) = (x + z \ (y::real))" +by arith + +lemma "(x - y \ z) = (x \ z + (y::real))" +by arith + +lemma "(-x < -y) = (y < (x::real))" +by arith + +lemma "(-x \ -y) = (y \ (x::real))" +by arith + +lemma "(a + b) - (c + d) = (a - c) + (b - (d::real))" +by arith + +lemma "(0::real) - x = -x" +by arith + +lemma "x - (0::real) = x" +by arith + +lemma "w \ x \ y < z ==> w + y < x + (z::real)" +by arith + +lemma "w < x \ y \ z ==> w + y < x + (z::real)" +by arith + +lemma "(0::real) \ x \ 0 < y ==> 0 < x + (y::real)" +by arith + +lemma "(0::real) < x \ 0 \ y ==> 0 < x + y" +by arith + +lemma "-x - y = -(x + (y::real))" +by arith + +lemma "x - (-y) = x + (y::real)" +by arith + +lemma "-x - -y = y - (x::real)" +by arith + +lemma "(a - b) + (b - c) = a - (c::real)" +by arith + +lemma "(x = y - z) = (x + z = (y::real))" +by arith + +lemma "(x - y = z) = (x = z + (y::real))" +by arith + +lemma "x - (x - y) = (y::real)" +by arith + +lemma "x - (x + y) = -(y::real)" +by arith + +lemma "x = y ==> x \ (y::real)" +by arith + +lemma "(0::real) < x ==> \(x = 0)" +by arith + +lemma "(x + y) * (x - y) = (x * x) - (y * y)" + oops + +lemma "(-x = -y) = (x = (y::real))" +by arith + +lemma "(-x < -y) = (y < (x::real))" +by arith + +lemma "!!a::real. a \ b ==> c \ d ==> x + y < z ==> a + c \ b + d" +by (tactic "fast_arith_tac @{context} 1") + +lemma "!!a::real. a < b ==> c < d ==> a - d \ b + (-c)" +by (tactic "fast_arith_tac @{context} 1") + +lemma "!!a::real. a \ b ==> b + b \ c ==> a + a \ c" +by (tactic "fast_arith_tac @{context} 1") + +lemma "!!a::real. a + b \ i + j ==> a \ b ==> i \ j ==> a + a \ j + j" +by (tactic "fast_arith_tac @{context} 1") + +lemma "!!a::real. a + b < i + j ==> a < b ==> i < j ==> a + a < j + j" +by (tactic "fast_arith_tac @{context} 1") + +lemma "!!a::real. a + b + c \ i + j + k \ a \ b \ b \ c \ i \ j \ j \ k --> a + a + a \ k + k + k" +by arith + +lemma "!!a::real. a + b + c + d \ i + j + k + l ==> a \ b ==> b \ c + ==> c \ d ==> i \ j ==> j \ k ==> k \ l ==> a \ l" +by (tactic "fast_arith_tac @{context} 1") + +lemma "!!a::real. a + b + c + d \ i + j + k + l ==> a \ b ==> b \ c + ==> c \ d ==> i \ j ==> j \ k ==> k \ l ==> a + a + a + a \ l + l + l + l" +by (tactic "fast_arith_tac @{context} 1") + +lemma "!!a::real. a + b + c + d \ i + j + k + l ==> a \ b ==> b \ c + ==> c \ d ==> i \ j ==> j \ k ==> k \ l ==> a + a + a + a + a \ l + l + l + l + i" +by (tactic "fast_arith_tac @{context} 1") + +lemma "!!a::real. a + b + c + d \ i + j + k + l ==> a \ b ==> b \ c + ==> c \ d ==> i \ j ==> j \ k ==> k \ l ==> a + a + a + a + a + a \ l + l + l + l + i + l" +by (tactic "fast_arith_tac @{context} 1") + + +subsection{*Complex Arithmetic*} + +lemma "(1359 + 93746*ii) - (2468 + 46375*ii) = -1109 + 47371*ii" +by simp + +lemma "- (65745 + -47371*ii) = -65745 + 47371*ii" +by simp + +text{*Multiplication requires distributive laws. Perhaps versions instantiated +to literal constants should be added to the simpset.*} + +lemma "(1 + ii) * (1 - ii) = 2" +by (simp add: ring_distribs) + +lemma "(1 + 2*ii) * (1 + 3*ii) = -5 + 5*ii" +by (simp add: ring_distribs) + +lemma "(-84 + 255*ii) + (51 * 255*ii) = -84 + 13260 * ii" +by (simp add: ring_distribs) + +text{*No inequalities or linear arithmetic: the complex numbers are unordered!*} + +text{*No powers (not supported yet)*} + end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/Eval_Examples.thy --- a/src/HOL/ex/Eval_Examples.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/ex/Eval_Examples.thy Wed Dec 03 15:58:44 2008 +0100 @@ -5,7 +5,7 @@ header {* Small examples for evaluation mechanisms *} theory Eval_Examples -imports Code_Eval "~~/src/HOL/Real/Rational" +imports Code_Eval Rational begin text {* evaluation oracle *} diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/ExecutableContent.thy --- a/src/HOL/ex/ExecutableContent.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/ex/ExecutableContent.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ -(* ID: $Id$ - Author: Florian Haftmann, TU Muenchen +(* Author: Florian Haftmann, TU Muenchen *) header {* A huge set of executable constants *} diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/HarmonicSeries.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/HarmonicSeries.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,322 @@ +(* Title: HOL/ex/HarmonicSeries.thy + Author: Benjamin Porter, 2006 +*) + +header {* Divergence of the Harmonic Series *} + +theory HarmonicSeries +imports Complex_Main +begin + +section {* Abstract *} + +text {* The following document presents a proof of the Divergence of +Harmonic Series theorem formalised in the Isabelle/Isar theorem +proving system. + +{\em Theorem:} The series $\sum_{n=1}^{\infty} \frac{1}{n}$ does not +converge to any number. + +{\em Informal Proof:} + The informal proof is based on the following auxillary lemmas: + \begin{itemize} + \item{aux: $\sum_{n=2^m-1}^{2^m} \frac{1}{n} \geq \frac{1}{2}$} + \item{aux2: $\sum_{n=1}^{2^M} \frac{1}{n} = 1 + \sum_{m=1}^{M} \sum_{n=2^m-1}^{2^m} \frac{1}{n}$} + \end{itemize} + + From {\em aux} and {\em aux2} we can deduce that $\sum_{n=1}^{2^M} + \frac{1}{n} \geq 1 + \frac{M}{2}$ for all $M$. + Now for contradiction, assume that $\sum_{n=1}^{\infty} \frac{1}{n} + = s$ for some $s$. Because $\forall n. \frac{1}{n} > 0$ all the + partial sums in the series must be less than $s$. However with our + deduction above we can choose $N > 2*s - 2$ and thus + $\sum_{n=1}^{2^N} \frac{1}{n} > s$. This leads to a contradiction + and hence $\sum_{n=1}^{\infty} \frac{1}{n}$ is not summable. + QED. +*} + +section {* Formal Proof *} + +lemma two_pow_sub: + "0 < m \ (2::nat)^m - 2^(m - 1) = 2^(m - 1)" + by (induct m) auto + +text {* We first prove the following auxillary lemma. This lemma +simply states that the finite sums: $\frac{1}{2}$, $\frac{1}{3} + +\frac{1}{4}$, $\frac{1}{5} + \frac{1}{6} + \frac{1}{7} + \frac{1}{8}$ +etc. are all greater than or equal to $\frac{1}{2}$. We do this by +observing that each term in the sum is greater than or equal to the +last term, e.g. $\frac{1}{3} > \frac{1}{4}$ and thus $\frac{1}{3} + +\frac{1}{4} > \frac{1}{4} + \frac{1}{4} = \frac{1}{2}$. *} + +lemma harmonic_aux: + "\m>0. (\n\{(2::nat)^(m - 1)+1..2^m}. 1/real n) \ 1/2" + (is "\m>0. (\n\(?S m). 1/real n) \ 1/2") +proof + fix m::nat + obtain tm where tmdef: "tm = (2::nat)^m" by simp + { + assume mgt0: "0 < m" + have "\x. x\(?S m) \ 1/(real x) \ 1/(real tm)" + proof - + fix x::nat + assume xs: "x\(?S m)" + have xgt0: "x>0" + proof - + from xs have + "x \ 2^(m - 1) + 1" by auto + moreover with mgt0 have + "2^(m - 1) + 1 \ (1::nat)" by auto + ultimately have + "x \ 1" by (rule xtrans) + thus ?thesis by simp + qed + moreover from xs have "x \ 2^m" by auto + ultimately have + "inverse (real x) \ inverse (real ((2::nat)^m))" by simp + moreover + from xgt0 have "real x \ 0" by simp + then have + "inverse (real x) = 1 / (real x)" + by (rule nonzero_inverse_eq_divide) + moreover from mgt0 have "real tm \ 0" by (simp add: tmdef) + then have + "inverse (real tm) = 1 / (real tm)" + by (rule nonzero_inverse_eq_divide) + ultimately show + "1/(real x) \ 1/(real tm)" by (auto simp add: tmdef) + qed + then have + "(\n\(?S m). 1 / real n) \ (\n\(?S m). 1/(real tm))" + by (rule setsum_mono) + moreover have + "(\n\(?S m). 1/(real tm)) = 1/2" + proof - + have + "(\n\(?S m). 1/(real tm)) = + (1/(real tm))*(\n\(?S m). 1)" + by simp + also have + "\ = ((1/(real tm)) * real (card (?S m)))" + by (simp add: real_of_card real_of_nat_def) + also have + "\ = ((1/(real tm)) * real (tm - (2^(m - 1))))" + by (simp add: tmdef) + also from mgt0 have + "\ = ((1/(real tm)) * real ((2::nat)^(m - 1)))" + by (auto simp: tmdef dest: two_pow_sub) + also have + "\ = (real (2::nat))^(m - 1) / (real (2::nat))^m" + by (simp add: tmdef realpow_real_of_nat [symmetric]) + also from mgt0 have + "\ = (real (2::nat))^(m - 1) / (real (2::nat))^((m - 1) + 1)" + by auto + also have "\ = 1/2" by simp + finally show ?thesis . + qed + ultimately have + "(\n\(?S m). 1 / real n) \ 1/2" + by - (erule subst) + } + thus "0 < m \ 1 / 2 \ (\n\(?S m). 1 / real n)" by simp +qed + +text {* We then show that the sum of a finite number of terms from the +harmonic series can be regrouped in increasing powers of 2. For +example: $1 + \frac{1}{2} + \frac{1}{3} + \frac{1}{4} + \frac{1}{5} + +\frac{1}{6} + \frac{1}{7} + \frac{1}{8} = 1 + (\frac{1}{2}) + +(\frac{1}{3} + \frac{1}{4}) + (\frac{1}{5} + \frac{1}{6} + \frac{1}{7} ++ \frac{1}{8})$. *} + +lemma harmonic_aux2 [rule_format]: + "0 (\n\{1..(2::nat)^M}. 1/real n) = + (1 + (\m\{1..M}. \n\{(2::nat)^(m - 1)+1..2^m}. 1/real n))" + (is "0 ?LHS M = ?RHS M") +proof (induct M) + case 0 show ?case by simp +next + case (Suc M) + have ant: "0 < Suc M" by fact + { + have suc: "?LHS (Suc M) = ?RHS (Suc M)" + proof cases -- "show that LHS = c and RHS = c, and thus LHS = RHS" + assume mz: "M=0" + { + then have + "?LHS (Suc M) = ?LHS 1" by simp + also have + "\ = (\n\{(1::nat)..2}. 1/real n)" by simp + also have + "\ = ((\n\{Suc 1..2}. 1/real n) + 1/(real (1::nat)))" + by (subst setsum_head) + (auto simp: atLeastSucAtMost_greaterThanAtMost) + also have + "\ = ((\n\{2..2::nat}. 1/real n) + 1/(real (1::nat)))" + by (simp add: nat_number) + also have + "\ = 1/(real (2::nat)) + 1/(real (1::nat))" by simp + finally have + "?LHS (Suc M) = 1/2 + 1" by simp + } + moreover + { + from mz have + "?RHS (Suc M) = ?RHS 1" by simp + also have + "\ = (\n\{((2::nat)^0)+1..2^1}. 1/real n) + 1" + by simp + also have + "\ = (\n\{2::nat..2}. 1/real n) + 1" + proof - + have "(2::nat)^0 = 1" by simp + then have "(2::nat)^0+1 = 2" by simp + moreover have "(2::nat)^1 = 2" by simp + ultimately have "{((2::nat)^0)+1..2^1} = {2::nat..2}" by auto + thus ?thesis by simp + qed + also have + "\ = 1/2 + 1" + by simp + finally have + "?RHS (Suc M) = 1/2 + 1" by simp + } + ultimately show "?LHS (Suc M) = ?RHS (Suc M)" by simp + next + assume mnz: "M\0" + then have mgtz: "M>0" by simp + with Suc have suc: + "(?LHS M) = (?RHS M)" by blast + have + "(?LHS (Suc M)) = + ((?LHS M) + (\n\{(2::nat)^M+1..2^(Suc M)}. 1 / real n))" + proof - + have + "{1..(2::nat)^(Suc M)} = + {1..(2::nat)^M}\{(2::nat)^M+1..(2::nat)^(Suc M)}" + by auto + moreover have + "{1..(2::nat)^M}\{(2::nat)^M+1..(2::nat)^(Suc M)} = {}" + by auto + moreover have + "finite {1..(2::nat)^M}" and "finite {(2::nat)^M+1..(2::nat)^(Suc M)}" + by auto + ultimately show ?thesis + by (auto intro: setsum_Un_disjoint) + qed + moreover + { + have + "(?RHS (Suc M)) = + (1 + (\m\{1..M}. \n\{(2::nat)^(m - 1)+1..2^m}. 1/real n) + + (\n\{(2::nat)^(Suc M - 1)+1..2^(Suc M)}. 1/real n))" by simp + also have + "\ = (?RHS M) + (\n\{(2::nat)^M+1..2^(Suc M)}. 1/real n)" + by simp + also from suc have + "\ = (?LHS M) + (\n\{(2::nat)^M+1..2^(Suc M)}. 1/real n)" + by simp + finally have + "(?RHS (Suc M)) = \" by simp + } + ultimately show "?LHS (Suc M) = ?RHS (Suc M)" by simp + qed + } + thus ?case by simp +qed + +text {* Using @{thm [source] harmonic_aux} and @{thm [source] harmonic_aux2} we now show +that each group sum is greater than or equal to $\frac{1}{2}$ and thus +the finite sum is bounded below by a value proportional to the number +of elements we choose. *} + +lemma harmonic_aux3 [rule_format]: + shows "\(M::nat). (\n\{1..(2::nat)^M}. 1 / real n) \ 1 + (real M)/2" + (is "\M. ?P M \ _") +proof (rule allI, cases) + fix M::nat + assume "M=0" + then show "?P M \ 1 + (real M)/2" by simp +next + fix M::nat + assume "M\0" + then have "M > 0" by simp + then have + "(?P M) = + (1 + (\m\{1..M}. \n\{(2::nat)^(m - 1)+1..2^m}. 1/real n))" + by (rule harmonic_aux2) + also have + "\ \ (1 + (\m\{1..M}. 1/2))" + proof - + let ?f = "(\x. 1/2)" + let ?g = "(\x. (\n\{(2::nat)^(x - 1)+1..2^x}. 1/real n))" + from harmonic_aux have "\x. x\{1..M} \ ?f x \ ?g x" by simp + then have "(\m\{1..M}. ?g m) \ (\m\{1..M}. ?f m)" by (rule setsum_mono) + thus ?thesis by simp + qed + finally have "(?P M) \ (1 + (\m\{1..M}. 1/2))" . + moreover + { + have + "(\m\{1..M}. (1::real)/2) = 1/2 * (\m\{1..M}. 1)" + by auto + also have + "\ = 1/2*(real (card {1..M}))" + by (simp only: real_of_card[symmetric]) + also have + "\ = 1/2*(real M)" by simp + also have + "\ = (real M)/2" by simp + finally have "(\m\{1..M}. (1::real)/2) = (real M)/2" . + } + ultimately show "(?P M) \ (1 + (real M)/2)" by simp +qed + +text {* The final theorem shows that as we take more and more elements +(see @{thm [source] harmonic_aux3}) we get an ever increasing sum. By assuming +the sum converges, the lemma @{thm [source] series_pos_less} ( @{thm +series_pos_less} ) states that each sum is bounded above by the +series' limit. This contradicts our first statement and thus we prove +that the harmonic series is divergent. *} + +theorem DivergenceOfHarmonicSeries: + shows "\summable (\n. 1/real (Suc n))" + (is "\summable ?f") +proof -- "by contradiction" + let ?s = "suminf ?f" -- "let ?s equal the sum of the harmonic series" + assume sf: "summable ?f" + then obtain n::nat where ndef: "n = nat \2 * ?s\" by simp + then have ngt: "1 + real n/2 > ?s" + proof - + have "\n. 0 \ ?f n" by simp + with sf have "?s \ 0" + by - (rule suminf_0_le, simp_all) + then have cgt0: "\2*?s\ \ 0" by simp + + from ndef have "n = nat \(2*?s)\" . + then have "real n = real (nat \2*?s\)" by simp + with cgt0 have "real n = real \2*?s\" + by (auto dest: real_nat_eq_real) + then have "real n \ 2*(?s)" by simp + then have "real n/2 \ (?s)" by simp + then show "1 + real n/2 > (?s)" by simp + qed + + obtain j where jdef: "j = (2::nat)^n" by simp + have "\m\j. 0 < ?f m" by simp + with sf have "(\i\{0..i\{1..i\{1..< Suc ((2::nat)^n)}. 1 / (real i)) < ?s" by simp + then have + "(\i\{1..(2::nat)^n}. 1 / (real i)) < ?s" + by (simp only: atLeastLessThanSuc_atLeastAtMost) + moreover from harmonic_aux3 have + "(\i\{1..(2::nat)^n}. 1 / (real i)) \ 1 + real n/2" by simp + moreover from ngt have "1 + real n/2 > ?s" by simp + ultimately show False by simp +qed + +end \ No newline at end of file diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/MIR.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/MIR.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,5933 @@ +(* Title: Complex/ex/MIR.thy + Author: Amine Chaieb +*) + +theory MIR +imports List Real Code_Integer Efficient_Nat +uses ("mirtac.ML") +begin + +section {* Quantifier elimination for @{text "\ (0, 1, +, floor, <)"} *} + +declare real_of_int_floor_cancel [simp del] + +primrec alluopairs:: "'a list \ ('a \ 'a) list" where + "alluopairs [] = []" +| "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)" + +lemma alluopairs_set1: "set (alluopairs xs) \ {(x,y). x\ set xs \ y\ set xs}" +by (induct xs, auto) + +lemma alluopairs_set: + "\x\ set xs ; y \ set xs\ \ (x,y) \ set (alluopairs xs) \ (y,x) \ set (alluopairs xs) " +by (induct xs, auto) + +lemma alluopairs_ex: + assumes Pc: "\ x y. P x y = P y x" + shows "(\ x \ set xs. \ y \ set xs. P x y) = (\ (x,y) \ set (alluopairs xs). P x y)" +proof + assume "\x\set xs. \y\set xs. P x y" + then obtain x y where x: "x \ set xs" and y:"y \ set xs" and P: "P x y" by blast + from alluopairs_set[OF x y] P Pc show"\(x, y)\set (alluopairs xs). P x y" + by auto +next + assume "\(x, y)\set (alluopairs xs). P x y" + then obtain "x" and "y" where xy:"(x,y) \ set (alluopairs xs)" and P: "P x y" by blast+ + from xy have "x \ set xs \ y\ set xs" using alluopairs_set1 by blast + with P show "\x\set xs. \y\set xs. P x y" by blast +qed + + (* generate a list from i to j*) +consts iupt :: "int \ int \ int list" +recdef iupt "measure (\ (i,j). nat (j-i +1))" + "iupt (i,j) = (if j (x#xs) ! n = xs ! (n - 1)" +using Nat.gr0_conv_Suc +by clarsimp + + +lemma myl: "\ (a::'a::{pordered_ab_group_add}) (b::'a). (a \ b) = (0 \ b - a)" +proof(clarify) + fix x y ::"'a" + have "(x \ y) = (x - y \ 0)" by (simp only: le_iff_diff_le_0[where a="x" and b="y"]) + also have "\ = (- (y - x) \ 0)" by simp + also have "\ = (0 \ y - x)" by (simp only: neg_le_0_iff_le[where a="y-x"]) + finally show "(x \ y) = (0 \ y - x)" . +qed + +lemma myless: "\ (a::'a::{pordered_ab_group_add}) (b::'a). (a < b) = (0 < b - a)" +proof(clarify) + fix x y ::"'a" + have "(x < y) = (x - y < 0)" by (simp only: less_iff_diff_less_0[where a="x" and b="y"]) + also have "\ = (- (y - x) < 0)" by simp + also have "\ = (0 < y - x)" by (simp only: neg_less_0_iff_less[where a="y-x"]) + finally show "(x < y) = (0 < y - x)" . +qed + +lemma myeq: "\ (a::'a::{pordered_ab_group_add}) (b::'a). (a = b) = (0 = b - a)" + by auto + + (* Maybe should be added to the library \ *) +lemma floor_int_eq: "(real n\ x \ x < real (n+1)) = (floor x = n)" +proof( auto) + assume lb: "real n \ x" + and ub: "x < real n + 1" + have "real (floor x) \ x" by simp + hence "real (floor x) < real (n + 1) " using ub by arith + hence "floor x < n+1" by simp + moreover from lb have "n \ floor x" using floor_mono2[where x="real n" and y="x"] + by simp ultimately show "floor x = n" by simp +qed + +(* Periodicity of dvd *) +lemma dvd_period: + assumes advdd: "(a::int) dvd d" + shows "(a dvd (x + t)) = (a dvd ((x+ c*d) + t))" + using advdd +proof- + {fix x k + from inf_period(3)[OF advdd, rule_format, where x=x and k="-k"] + have " ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp} + hence "\x.\k. ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp + then show ?thesis by simp +qed + + (* The Divisibility relation between reals *) +definition + rdvd:: "real \ real \ bool" (infixl "rdvd" 50) +where + rdvd_def: "x rdvd y \ (\k\int. y = x * real k)" + +lemma int_rdvd_real: + shows "real (i::int) rdvd x = (i dvd (floor x) \ real (floor x) = x)" (is "?l = ?r") +proof + assume "?l" + hence th: "\ k. x=real (i*k)" by (simp add: rdvd_def) + hence th': "real (floor x) = x" by (auto simp del: real_of_int_mult) + with th have "\ k. real (floor x) = real (i*k)" by simp + hence "\ k. floor x = i*k" by (simp only: real_of_int_inject) + thus ?r using th' by (simp add: dvd_def) +next + assume "?r" hence "(i\int) dvd \x\real\" .. + hence "\ k. real (floor x) = real (i*k)" + by (simp only: real_of_int_inject) (simp add: dvd_def) + thus ?l using prems by (simp add: rdvd_def) +qed + +lemma int_rdvd_iff: "(real (i::int) rdvd real t) = (i dvd t)" +by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only :real_of_int_mult[symmetric]) + + +lemma rdvd_abs1: + "(abs (real d) rdvd t) = (real (d ::int) rdvd t)" +proof + assume d: "real d rdvd t" + from d int_rdvd_real have d2: "d dvd (floor t)" and ti: "real (floor t) = t" by auto + + from iffD2[OF zdvd_abs1] d2 have "(abs d) dvd (floor t)" by blast + with ti int_rdvd_real[symmetric] have "real (abs d) rdvd t" by blast + thus "abs (real d) rdvd t" by simp +next + assume "abs (real d) rdvd t" hence "real (abs d) rdvd t" by simp + with int_rdvd_real[where i="abs d" and x="t"] have d2: "abs d dvd floor t" and ti: "real (floor t) =t" by auto + from iffD1[OF zdvd_abs1] d2 have "d dvd floor t" by blast + with ti int_rdvd_real[symmetric] show "real d rdvd t" by blast +qed + +lemma rdvd_minus: "(real (d::int) rdvd t) = (real d rdvd -t)" + apply (auto simp add: rdvd_def) + apply (rule_tac x="-k" in exI, simp) + apply (rule_tac x="-k" in exI, simp) +done + +lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)" +by (auto simp add: rdvd_def) + +lemma rdvd_mult: + assumes knz: "k\0" + shows "(real (n::int) * real (k::int) rdvd x * real k) = (real n rdvd x)" +using knz by (simp add:rdvd_def) + +lemma rdvd_trans: assumes mn:"m rdvd n" and nk:"n rdvd k" + shows "m rdvd k" +proof- + from rdvd_def mn obtain c where nmc:"n = m * real (c::int)" by auto + from rdvd_def nk obtain c' where nkc:"k = n * real (c'::int)" by auto + hence "k = m * real (c * c')" using nmc by simp + thus ?thesis using rdvd_def by blast +qed + + (*********************************************************************************) + (**** SHADOW SYNTAX AND SEMANTICS ****) + (*********************************************************************************) + +datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num + | Mul int num | Floor num| CF int num num + + (* A size for num to make inductive proofs simpler*) +primrec num_size :: "num \ nat" where + "num_size (C c) = 1" +| "num_size (Bound n) = 1" +| "num_size (Neg a) = 1 + num_size a" +| "num_size (Add a b) = 1 + num_size a + num_size b" +| "num_size (Sub a b) = 3 + num_size a + num_size b" +| "num_size (CN n c a) = 4 + num_size a " +| "num_size (CF c a b) = 4 + num_size a + num_size b" +| "num_size (Mul c a) = 1 + num_size a" +| "num_size (Floor a) = 1 + num_size a" + + (* Semantics of numeral terms (num) *) +primrec Inum :: "real list \ num \ real" where + "Inum bs (C c) = (real c)" +| "Inum bs (Bound n) = bs!n" +| "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)" +| "Inum bs (Neg a) = -(Inum bs a)" +| "Inum bs (Add a b) = Inum bs a + Inum bs b" +| "Inum bs (Sub a b) = Inum bs a - Inum bs b" +| "Inum bs (Mul c a) = (real c) * Inum bs a" +| "Inum bs (Floor a) = real (floor (Inum bs a))" +| "Inum bs (CF c a b) = real c * real (floor (Inum bs a)) + Inum bs b" +definition "isint t bs \ real (floor (Inum bs t)) = Inum bs t" + +lemma isint_iff: "isint n bs = (real (floor (Inum bs n)) = Inum bs n)" +by (simp add: isint_def) + +lemma isint_Floor: "isint (Floor n) bs" + by (simp add: isint_iff) + +lemma isint_Mul: "isint e bs \ isint (Mul c e) bs" +proof- + let ?e = "Inum bs e" + let ?fe = "floor ?e" + assume be: "isint e bs" hence efe:"real ?fe = ?e" by (simp add: isint_iff) + have "real ((floor (Inum bs (Mul c e)))) = real (floor (real (c * ?fe)))" using efe by simp + also have "\ = real (c* ?fe)" by (simp only: floor_real_of_int) + also have "\ = real c * ?e" using efe by simp + finally show ?thesis using isint_iff by simp +qed + +lemma isint_neg: "isint e bs \ isint (Neg e) bs" +proof- + let ?I = "\ t. Inum bs t" + assume ie: "isint e bs" + hence th: "real (floor (?I e)) = ?I e" by (simp add: isint_def) + have "real (floor (?I (Neg e))) = real (floor (- (real (floor (?I e)))))" by (simp add: th) + also have "\ = - real (floor (?I e))" by(simp add: floor_minus_real_of_int) + finally show "isint (Neg e) bs" by (simp add: isint_def th) +qed + +lemma isint_sub: + assumes ie: "isint e bs" shows "isint (Sub (C c) e) bs" +proof- + let ?I = "\ t. Inum bs t" + from ie have th: "real (floor (?I e)) = ?I e" by (simp add: isint_def) + have "real (floor (?I (Sub (C c) e))) = real (floor ((real (c -floor (?I e)))))" by (simp add: th) + also have "\ = real (c- floor (?I e))" by(simp add: floor_minus_real_of_int) + finally show "isint (Sub (C c) e) bs" by (simp add: isint_def th) +qed + +lemma isint_add: assumes + ai:"isint a bs" and bi: "isint b bs" shows "isint (Add a b) bs" +proof- + let ?a = "Inum bs a" + let ?b = "Inum bs b" + from ai bi isint_iff have "real (floor (?a + ?b)) = real (floor (real (floor ?a) + real (floor ?b)))" by simp + also have "\ = real (floor ?a) + real (floor ?b)" by simp + also have "\ = ?a + ?b" using ai bi isint_iff by simp + finally show "isint (Add a b) bs" by (simp add: isint_iff) +qed + +lemma isint_c: "isint (C j) bs" + by (simp add: isint_iff) + + + (* FORMULAE *) +datatype fm = + T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num| + NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm + + + (* A size for fm *) +fun fmsize :: "fm \ nat" where + "fmsize (NOT p) = 1 + fmsize p" +| "fmsize (And p q) = 1 + fmsize p + fmsize q" +| "fmsize (Or p q) = 1 + fmsize p + fmsize q" +| "fmsize (Imp p q) = 3 + fmsize p + fmsize q" +| "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)" +| "fmsize (E p) = 1 + fmsize p" +| "fmsize (A p) = 4+ fmsize p" +| "fmsize (Dvd i t) = 2" +| "fmsize (NDvd i t) = 2" +| "fmsize p = 1" + (* several lemmas about fmsize *) +lemma fmsize_pos: "fmsize p > 0" +by (induct p rule: fmsize.induct) simp_all + + (* Semantics of formulae (fm) *) +primrec Ifm ::"real list \ fm \ bool" where + "Ifm bs T = True" +| "Ifm bs F = False" +| "Ifm bs (Lt a) = (Inum bs a < 0)" +| "Ifm bs (Gt a) = (Inum bs a > 0)" +| "Ifm bs (Le a) = (Inum bs a \ 0)" +| "Ifm bs (Ge a) = (Inum bs a \ 0)" +| "Ifm bs (Eq a) = (Inum bs a = 0)" +| "Ifm bs (NEq a) = (Inum bs a \ 0)" +| "Ifm bs (Dvd i b) = (real i rdvd Inum bs b)" +| "Ifm bs (NDvd i b) = (\(real i rdvd Inum bs b))" +| "Ifm bs (NOT p) = (\ (Ifm bs p))" +| "Ifm bs (And p q) = (Ifm bs p \ Ifm bs q)" +| "Ifm bs (Or p q) = (Ifm bs p \ Ifm bs q)" +| "Ifm bs (Imp p q) = ((Ifm bs p) \ (Ifm bs q))" +| "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)" +| "Ifm bs (E p) = (\ x. Ifm (x#bs) p)" +| "Ifm bs (A p) = (\ x. Ifm (x#bs) p)" + +consts prep :: "fm \ fm" +recdef prep "measure fmsize" + "prep (E T) = T" + "prep (E F) = F" + "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))" + "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))" + "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" + "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))" + "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))" + "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))" + "prep (E p) = E (prep p)" + "prep (A (And p q)) = And (prep (A p)) (prep (A q))" + "prep (A p) = prep (NOT (E (NOT p)))" + "prep (NOT (NOT p)) = prep p" + "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))" + "prep (NOT (A p)) = prep (E (NOT p))" + "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))" + "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))" + "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))" + "prep (NOT p) = NOT (prep p)" + "prep (Or p q) = Or (prep p) (prep q)" + "prep (And p q) = And (prep p) (prep q)" + "prep (Imp p q) = prep (Or (NOT p) q)" + "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))" + "prep p = p" +(hints simp add: fmsize_pos) +lemma prep: "\ bs. Ifm bs (prep p) = Ifm bs p" +by (induct p rule: prep.induct, auto) + + + (* Quantifier freeness *) +fun qfree:: "fm \ bool" where + "qfree (E p) = False" + | "qfree (A p) = False" + | "qfree (NOT p) = qfree p" + | "qfree (And p q) = (qfree p \ qfree q)" + | "qfree (Or p q) = (qfree p \ qfree q)" + | "qfree (Imp p q) = (qfree p \ qfree q)" + | "qfree (Iff p q) = (qfree p \ qfree q)" + | "qfree p = True" + + (* Boundedness and substitution *) +primrec numbound0 :: "num \ bool" (* a num is INDEPENDENT of Bound 0 *) where + "numbound0 (C c) = True" + | "numbound0 (Bound n) = (n>0)" + | "numbound0 (CN n i a) = (n > 0 \ numbound0 a)" + | "numbound0 (Neg a) = numbound0 a" + | "numbound0 (Add a b) = (numbound0 a \ numbound0 b)" + | "numbound0 (Sub a b) = (numbound0 a \ numbound0 b)" + | "numbound0 (Mul i a) = numbound0 a" + | "numbound0 (Floor a) = numbound0 a" + | "numbound0 (CF c a b) = (numbound0 a \ numbound0 b)" + +lemma numbound0_I: + assumes nb: "numbound0 a" + shows "Inum (b#bs) a = Inum (b'#bs) a" + using nb by (induct a) (auto simp add: nth_pos2) + +lemma numbound0_gen: + assumes nb: "numbound0 t" and ti: "isint t (x#bs)" + shows "\ y. isint t (y#bs)" +using nb ti +proof(clarify) + fix y + from numbound0_I[OF nb, where bs="bs" and b="y" and b'="x"] ti[simplified isint_def] + show "isint t (y#bs)" + by (simp add: isint_def) +qed + +primrec bound0:: "fm \ bool" (* A Formula is independent of Bound 0 *) where + "bound0 T = True" + | "bound0 F = True" + | "bound0 (Lt a) = numbound0 a" + | "bound0 (Le a) = numbound0 a" + | "bound0 (Gt a) = numbound0 a" + | "bound0 (Ge a) = numbound0 a" + | "bound0 (Eq a) = numbound0 a" + | "bound0 (NEq a) = numbound0 a" + | "bound0 (Dvd i a) = numbound0 a" + | "bound0 (NDvd i a) = numbound0 a" + | "bound0 (NOT p) = bound0 p" + | "bound0 (And p q) = (bound0 p \ bound0 q)" + | "bound0 (Or p q) = (bound0 p \ bound0 q)" + | "bound0 (Imp p q) = ((bound0 p) \ (bound0 q))" + | "bound0 (Iff p q) = (bound0 p \ bound0 q)" + | "bound0 (E p) = False" + | "bound0 (A p) = False" + +lemma bound0_I: + assumes bp: "bound0 p" + shows "Ifm (b#bs) p = Ifm (b'#bs) p" + using bp numbound0_I [where b="b" and bs="bs" and b'="b'"] + by (induct p) (auto simp add: nth_pos2) + +primrec numsubst0:: "num \ num \ num" (* substitute a num into a num for Bound 0 *) where + "numsubst0 t (C c) = (C c)" + | "numsubst0 t (Bound n) = (if n=0 then t else Bound n)" + | "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))" + | "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)" + | "numsubst0 t (Neg a) = Neg (numsubst0 t a)" + | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)" + | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" + | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)" + | "numsubst0 t (Floor a) = Floor (numsubst0 t a)" + +lemma numsubst0_I: + shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t" + by (induct t) (simp_all add: nth_pos2) + +lemma numsubst0_I': + assumes nb: "numbound0 a" + shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t" + by (induct t) (simp_all add: nth_pos2 numbound0_I[OF nb, where b="b" and b'="b'"]) + +primrec subst0:: "num \ fm \ fm" (* substitue a num into a formula for Bound 0 *) where + "subst0 t T = T" + | "subst0 t F = F" + | "subst0 t (Lt a) = Lt (numsubst0 t a)" + | "subst0 t (Le a) = Le (numsubst0 t a)" + | "subst0 t (Gt a) = Gt (numsubst0 t a)" + | "subst0 t (Ge a) = Ge (numsubst0 t a)" + | "subst0 t (Eq a) = Eq (numsubst0 t a)" + | "subst0 t (NEq a) = NEq (numsubst0 t a)" + | "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)" + | "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)" + | "subst0 t (NOT p) = NOT (subst0 t p)" + | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)" + | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)" + | "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)" + | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)" + +lemma subst0_I: assumes qfp: "qfree p" + shows "Ifm (b#bs) (subst0 a p) = Ifm ((Inum (b#bs) a)#bs) p" + using qfp numsubst0_I[where b="b" and bs="bs" and a="a"] + by (induct p) (simp_all add: nth_pos2 ) + +consts + decrnum:: "num \ num" + decr :: "fm \ fm" + +recdef decrnum "measure size" + "decrnum (Bound n) = Bound (n - 1)" + "decrnum (Neg a) = Neg (decrnum a)" + "decrnum (Add a b) = Add (decrnum a) (decrnum b)" + "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)" + "decrnum (Mul c a) = Mul c (decrnum a)" + "decrnum (Floor a) = Floor (decrnum a)" + "decrnum (CN n c a) = CN (n - 1) c (decrnum a)" + "decrnum (CF c a b) = CF c (decrnum a) (decrnum b)" + "decrnum a = a" + +recdef decr "measure size" + "decr (Lt a) = Lt (decrnum a)" + "decr (Le a) = Le (decrnum a)" + "decr (Gt a) = Gt (decrnum a)" + "decr (Ge a) = Ge (decrnum a)" + "decr (Eq a) = Eq (decrnum a)" + "decr (NEq a) = NEq (decrnum a)" + "decr (Dvd i a) = Dvd i (decrnum a)" + "decr (NDvd i a) = NDvd i (decrnum a)" + "decr (NOT p) = NOT (decr p)" + "decr (And p q) = And (decr p) (decr q)" + "decr (Or p q) = Or (decr p) (decr q)" + "decr (Imp p q) = Imp (decr p) (decr q)" + "decr (Iff p q) = Iff (decr p) (decr q)" + "decr p = p" + +lemma decrnum: assumes nb: "numbound0 t" + shows "Inum (x#bs) t = Inum bs (decrnum t)" + using nb by (induct t rule: decrnum.induct, simp_all add: nth_pos2) + +lemma decr: assumes nb: "bound0 p" + shows "Ifm (x#bs) p = Ifm bs (decr p)" + using nb + by (induct p rule: decr.induct, simp_all add: nth_pos2 decrnum) + +lemma decr_qf: "bound0 p \ qfree (decr p)" +by (induct p, simp_all) + +consts + isatom :: "fm \ bool" (* test for atomicity *) +recdef isatom "measure size" + "isatom T = True" + "isatom F = True" + "isatom (Lt a) = True" + "isatom (Le a) = True" + "isatom (Gt a) = True" + "isatom (Ge a) = True" + "isatom (Eq a) = True" + "isatom (NEq a) = True" + "isatom (Dvd i b) = True" + "isatom (NDvd i b) = True" + "isatom p = False" + +lemma numsubst0_numbound0: assumes nb: "numbound0 t" + shows "numbound0 (numsubst0 t a)" +using nb by (induct a, auto) + +lemma subst0_bound0: assumes qf: "qfree p" and nb: "numbound0 t" + shows "bound0 (subst0 t p)" +using qf numsubst0_numbound0[OF nb] by (induct p, auto) + +lemma bound0_qf: "bound0 p \ qfree p" +by (induct p, simp_all) + + +definition djf:: "('a \ fm) \ 'a \ fm \ fm" where + "djf f p q = (if q=T then T else if q=F then f p else + (let fp = f p in case fp of T \ T | F \ q | _ \ Or fp q))" + +definition evaldjf:: "('a \ fm) \ 'a list \ fm" where + "evaldjf f ps = foldr (djf f) ps F" + +lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)" +by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) +(cases "f p", simp_all add: Let_def djf_def) + +lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\ p \ set ps. Ifm bs (f p))" + by(induct ps, simp_all add: evaldjf_def djf_Or) + +lemma evaldjf_bound0: + assumes nb: "\ x\ set xs. bound0 (f x)" + shows "bound0 (evaldjf f xs)" + using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) + +lemma evaldjf_qf: + assumes nb: "\ x\ set xs. qfree (f x)" + shows "qfree (evaldjf f xs)" + using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) + +consts + disjuncts :: "fm \ fm list" + conjuncts :: "fm \ fm list" +recdef disjuncts "measure size" + "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)" + "disjuncts F = []" + "disjuncts p = [p]" + +recdef conjuncts "measure size" + "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)" + "conjuncts T = []" + "conjuncts p = [p]" +lemma disjuncts: "(\ q\ set (disjuncts p). Ifm bs q) = Ifm bs p" +by(induct p rule: disjuncts.induct, auto) +lemma conjuncts: "(\ q\ set (conjuncts p). Ifm bs q) = Ifm bs p" +by(induct p rule: conjuncts.induct, auto) + +lemma disjuncts_nb: "bound0 p \ \ q\ set (disjuncts p). bound0 q" +proof- + assume nb: "bound0 p" + hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto) + thus ?thesis by (simp only: list_all_iff) +qed +lemma conjuncts_nb: "bound0 p \ \ q\ set (conjuncts p). bound0 q" +proof- + assume nb: "bound0 p" + hence "list_all bound0 (conjuncts p)" by (induct p rule:conjuncts.induct,auto) + thus ?thesis by (simp only: list_all_iff) +qed + +lemma disjuncts_qf: "qfree p \ \ q\ set (disjuncts p). qfree q" +proof- + assume qf: "qfree p" + hence "list_all qfree (disjuncts p)" + by (induct p rule: disjuncts.induct, auto) + thus ?thesis by (simp only: list_all_iff) +qed +lemma conjuncts_qf: "qfree p \ \ q\ set (conjuncts p). qfree q" +proof- + assume qf: "qfree p" + hence "list_all qfree (conjuncts p)" + by (induct p rule: conjuncts.induct, auto) + thus ?thesis by (simp only: list_all_iff) +qed + +constdefs DJ :: "(fm \ fm) \ fm \ fm" + "DJ f p \ evaldjf f (disjuncts p)" + +lemma DJ: assumes fdj: "\ p q. f (Or p q) = Or (f p) (f q)" + and fF: "f F = F" + shows "Ifm bs (DJ f p) = Ifm bs (f p)" +proof- + have "Ifm bs (DJ f p) = (\ q \ set (disjuncts p). Ifm bs (f q))" + by (simp add: DJ_def evaldjf_ex) + also have "\ = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto) + finally show ?thesis . +qed + +lemma DJ_qf: assumes + fqf: "\ p. qfree p \ qfree (f p)" + shows "\p. qfree p \ qfree (DJ f p) " +proof(clarify) + fix p assume qf: "qfree p" + have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def) + from disjuncts_qf[OF qf] have "\ q\ set (disjuncts p). qfree q" . + with fqf have th':"\ q\ set (disjuncts p). qfree (f q)" by blast + + from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp +qed + +lemma DJ_qe: assumes qe: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" + shows "\ bs p. qfree p \ qfree (DJ qe p) \ (Ifm bs ((DJ qe p)) = Ifm bs (E p))" +proof(clarify) + fix p::fm and bs + assume qf: "qfree p" + from qe have qth: "\ p. qfree p \ qfree (qe p)" by blast + from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto + have "Ifm bs (DJ qe p) = (\ q\ set (disjuncts p). Ifm bs (qe q))" + by (simp add: DJ_def evaldjf_ex) + also have "\ = (\ q \ set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto + also have "\ = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto) + finally show "qfree (DJ qe p) \ Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast +qed + (* Simplification *) + + (* Algebraic simplifications for nums *) +consts bnds:: "num \ nat list" + lex_ns:: "nat list \ nat list \ bool" +recdef bnds "measure size" + "bnds (Bound n) = [n]" + "bnds (CN n c a) = n#(bnds a)" + "bnds (Neg a) = bnds a" + "bnds (Add a b) = (bnds a)@(bnds b)" + "bnds (Sub a b) = (bnds a)@(bnds b)" + "bnds (Mul i a) = bnds a" + "bnds (Floor a) = bnds a" + "bnds (CF c a b) = (bnds a)@(bnds b)" + "bnds a = []" +recdef lex_ns "measure (\ (xs,ys). length xs + length ys)" + "lex_ns ([], ms) = True" + "lex_ns (ns, []) = False" + "lex_ns (n#ns, m#ms) = (n ((n = m) \ lex_ns (ns,ms))) " +constdefs lex_bnd :: "num \ num \ bool" + "lex_bnd t s \ lex_ns (bnds t, bnds s)" + +consts + numgcdh:: "num \ int \ int" + reducecoeffh:: "num \ int \ num" + dvdnumcoeff:: "num \ int \ bool" +consts maxcoeff:: "num \ int" +recdef maxcoeff "measure size" + "maxcoeff (C i) = abs i" + "maxcoeff (CN n c t) = max (abs c) (maxcoeff t)" + "maxcoeff (CF c t s) = max (abs c) (maxcoeff s)" + "maxcoeff t = 1" + +lemma maxcoeff_pos: "maxcoeff t \ 0" + apply (induct t rule: maxcoeff.induct, auto) + done + +recdef numgcdh "measure size" + "numgcdh (C i) = (\g. zgcd i g)" + "numgcdh (CN n c t) = (\g. zgcd c (numgcdh t g))" + "numgcdh (CF c s t) = (\g. zgcd c (numgcdh t g))" + "numgcdh t = (\g. 1)" + +definition + numgcd :: "num \ int" +where + numgcd_def: "numgcd t = numgcdh t (maxcoeff t)" + +recdef reducecoeffh "measure size" + "reducecoeffh (C i) = (\ g. C (i div g))" + "reducecoeffh (CN n c t) = (\ g. CN n (c div g) (reducecoeffh t g))" + "reducecoeffh (CF c s t) = (\ g. CF (c div g) s (reducecoeffh t g))" + "reducecoeffh t = (\g. t)" + +definition + reducecoeff :: "num \ num" +where + reducecoeff_def: "reducecoeff t = + (let g = numgcd t in + if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)" + +recdef dvdnumcoeff "measure size" + "dvdnumcoeff (C i) = (\ g. g dvd i)" + "dvdnumcoeff (CN n c t) = (\ g. g dvd c \ (dvdnumcoeff t g))" + "dvdnumcoeff (CF c s t) = (\ g. g dvd c \ (dvdnumcoeff t g))" + "dvdnumcoeff t = (\g. False)" + +lemma dvdnumcoeff_trans: + assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'" + shows "dvdnumcoeff t g" + using dgt' gdg + by (induct t rule: dvdnumcoeff.induct, simp_all add: gdg zdvd_trans[OF gdg]) + +declare zdvd_trans [trans add] + +lemma natabs0: "(nat (abs x) = 0) = (x = 0)" +by arith + +lemma numgcd0: + assumes g0: "numgcd t = 0" + shows "Inum bs t = 0" +proof- + have "\x. numgcdh t x= 0 \ Inum bs t = 0" + by (induct t rule: numgcdh.induct, auto simp add: zgcd_def gcd_zero natabs0 max_def maxcoeff_pos) + thus ?thesis using g0[simplified numgcd_def] by blast +qed + +lemma numgcdh_pos: assumes gp: "g \ 0" shows "numgcdh t g \ 0" + using gp + by (induct t rule: numgcdh.induct, auto simp add: zgcd_def) + +lemma numgcd_pos: "numgcd t \0" + by (simp add: numgcd_def numgcdh_pos maxcoeff_pos) + +lemma reducecoeffh: + assumes gt: "dvdnumcoeff t g" and gp: "g > 0" + shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t" + using gt +proof(induct t rule: reducecoeffh.induct) + case (1 i) hence gd: "g dvd i" by simp + from gp have gnz: "g \ 0" by simp + from prems show ?case by (simp add: real_of_int_div[OF gnz gd]) +next + case (2 n c t) hence gd: "g dvd c" by simp + from gp have gnz: "g \ 0" by simp + from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps) +next + case (3 c s t) hence gd: "g dvd c" by simp + from gp have gnz: "g \ 0" by simp + from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps) +qed (auto simp add: numgcd_def gp) +consts ismaxcoeff:: "num \ int \ bool" +recdef ismaxcoeff "measure size" + "ismaxcoeff (C i) = (\ x. abs i \ x)" + "ismaxcoeff (CN n c t) = (\x. abs c \ x \ (ismaxcoeff t x))" + "ismaxcoeff (CF c s t) = (\x. abs c \ x \ (ismaxcoeff t x))" + "ismaxcoeff t = (\x. True)" + +lemma ismaxcoeff_mono: "ismaxcoeff t c \ c \ c' \ ismaxcoeff t c'" +by (induct t rule: ismaxcoeff.induct, auto) + +lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)" +proof (induct t rule: maxcoeff.induct) + case (2 n c t) + hence H:"ismaxcoeff t (maxcoeff t)" . + have thh: "maxcoeff t \ max (abs c) (maxcoeff t)" by (simp add: le_maxI2) + from ismaxcoeff_mono[OF H thh] show ?case by (simp add: le_maxI1) +next + case (3 c t s) + hence H1:"ismaxcoeff s (maxcoeff s)" by auto + have thh1: "maxcoeff s \ max \c\ (maxcoeff s)" by (simp add: max_def) + from ismaxcoeff_mono[OF H1 thh1] show ?case by (simp add: le_maxI1) +qed simp_all + +lemma zgcd_gt1: "zgcd i j > 1 \ ((abs i > 1 \ abs j > 1) \ (abs i = 0 \ abs j > 1) \ (abs i > 1 \ abs j = 0))" + apply (unfold zgcd_def) + apply (cases "i = 0", simp_all) + apply (cases "j = 0", simp_all) + apply (cases "abs i = 1", simp_all) + apply (cases "abs j = 1", simp_all) + apply auto + done +lemma numgcdh0:"numgcdh t m = 0 \ m =0" + by (induct t rule: numgcdh.induct, auto simp add:zgcd0) + +lemma dvdnumcoeff_aux: + assumes "ismaxcoeff t m" and mp:"m \ 0" and "numgcdh t m > 1" + shows "dvdnumcoeff t (numgcdh t m)" +using prems +proof(induct t rule: numgcdh.induct) + case (2 n c t) + let ?g = "numgcdh t m" + from prems have th:"zgcd c ?g > 1" by simp + from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] + have "(abs c > 1 \ ?g > 1) \ (abs c = 0 \ ?g > 1) \ (abs c > 1 \ ?g = 0)" by simp + moreover {assume "abs c > 1" and gp: "?g > 1" with prems + have th: "dvdnumcoeff t ?g" by simp + have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) + from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)} + moreover {assume "abs c = 0 \ ?g > 1" + with prems have th: "dvdnumcoeff t ?g" by simp + have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) + from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1) + hence ?case by simp } + moreover {assume "abs c > 1" and g0:"?g = 0" + from numgcdh0[OF g0] have "m=0". with prems have ?case by simp } + ultimately show ?case by blast +next + case (3 c s t) + let ?g = "numgcdh t m" + from prems have th:"zgcd c ?g > 1" by simp + from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] + have "(abs c > 1 \ ?g > 1) \ (abs c = 0 \ ?g > 1) \ (abs c > 1 \ ?g = 0)" by simp + moreover {assume "abs c > 1" and gp: "?g > 1" with prems + have th: "dvdnumcoeff t ?g" by simp + have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) + from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)} + moreover {assume "abs c = 0 \ ?g > 1" + with prems have th: "dvdnumcoeff t ?g" by simp + have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) + from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1) + hence ?case by simp } + moreover {assume "abs c > 1" and g0:"?g = 0" + from numgcdh0[OF g0] have "m=0". with prems have ?case by simp } + ultimately show ?case by blast +qed(auto simp add: zgcd_zdvd1) + +lemma dvdnumcoeff_aux2: + assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \ numgcd t > 0" + using prems +proof (simp add: numgcd_def) + let ?mc = "maxcoeff t" + let ?g = "numgcdh t ?mc" + have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff) + have th2: "?mc \ 0" by (rule maxcoeff_pos) + assume H: "numgcdh t ?mc > 1" + from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" . +qed + +lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" +proof- + let ?g = "numgcd t" + have "?g \ 0" by (simp add: numgcd_pos) + hence "?g = 0 \ ?g = 1 \ ?g > 1" by auto + moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} + moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} + moreover { assume g1:"?g > 1" + from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+ + from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis + by (simp add: reducecoeff_def Let_def)} + ultimately show ?thesis by blast +qed + +lemma reducecoeffh_numbound0: "numbound0 t \ numbound0 (reducecoeffh t g)" +by (induct t rule: reducecoeffh.induct, auto) + +lemma reducecoeff_numbound0: "numbound0 t \ numbound0 (reducecoeff t)" +using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def) + +consts + simpnum:: "num \ num" + numadd:: "num \ num \ num" + nummul:: "num \ int \ num" + +recdef numadd "measure (\ (t,s). size t + size s)" + "numadd (CN n1 c1 r1,CN n2 c2 r2) = + (if n1=n2 then + (let c = c1 + c2 + in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2)))) + else if n1 \ n2 then CN n1 c1 (numadd (r1,CN n2 c2 r2)) + else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))" + "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))" + "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" + "numadd (CF c1 t1 r1,CF c2 t2 r2) = + (if t1 = t2 then + (let c=c1+c2; s= numadd(r1,r2) in (if c=0 then s else CF c t1 s)) + else if lex_bnd t1 t2 then CF c1 t1 (numadd(r1,CF c2 t2 r2)) + else CF c2 t2 (numadd(CF c1 t1 r1,r2)))" + "numadd (CF c1 t1 r1,C c) = CF c1 t1 (numadd (r1, C c))" + "numadd (C c,CF c1 t1 r1) = CF c1 t1 (numadd (r1, C c))" + "numadd (C b1, C b2) = C (b1+b2)" + "numadd (a,b) = Add a b" + +lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)" +apply (induct t s rule: numadd.induct, simp_all add: Let_def) + apply (case_tac "c1+c2 = 0",case_tac "n1 \ n2", simp_all) + apply (case_tac "n1 = n2", simp_all add: ring_simps) + apply (simp only: left_distrib[symmetric]) + apply simp +apply (case_tac "lex_bnd t1 t2", simp_all) + apply (case_tac "c1+c2 = 0") + by (case_tac "t1 = t2", simp_all add: ring_simps left_distrib[symmetric] real_of_int_mult[symmetric] real_of_int_add[symmetric]del: real_of_int_mult real_of_int_add left_distrib) + +lemma numadd_nb[simp]: "\ numbound0 t ; numbound0 s\ \ numbound0 (numadd (t,s))" +by (induct t s rule: numadd.induct, auto simp add: Let_def) + +recdef nummul "measure size" + "nummul (C j) = (\ i. C (i*j))" + "nummul (CN n c t) = (\ i. CN n (c*i) (nummul t i))" + "nummul (CF c t s) = (\ i. CF (c*i) t (nummul s i))" + "nummul (Mul c t) = (\ i. nummul t (i*c))" + "nummul t = (\ i. Mul i t)" + +lemma nummul[simp]: "\ i. Inum bs (nummul t i) = Inum bs (Mul i t)" +by (induct t rule: nummul.induct, auto simp add: ring_simps) + +lemma nummul_nb[simp]: "\ i. numbound0 t \ numbound0 (nummul t i)" +by (induct t rule: nummul.induct, auto) + +constdefs numneg :: "num \ num" + "numneg t \ nummul t (- 1)" + +constdefs numsub :: "num \ num \ num" + "numsub s t \ (if s = t then C 0 else numadd (s,numneg t))" + +lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)" +using numneg_def nummul by simp + +lemma numneg_nb[simp]: "numbound0 t \ numbound0 (numneg t)" +using numneg_def by simp + +lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)" +using numsub_def by simp + +lemma numsub_nb[simp]: "\ numbound0 t ; numbound0 s\ \ numbound0 (numsub t s)" +using numsub_def by simp + +lemma isint_CF: assumes si: "isint s bs" shows "isint (CF c t s) bs" +proof- + have cti: "isint (Mul c (Floor t)) bs" by (simp add: isint_Mul isint_Floor) + + have "?thesis = isint (Add (Mul c (Floor t)) s) bs" by (simp add: isint_def) + also have "\" by (simp add: isint_add cti si) + finally show ?thesis . +qed + +consts split_int:: "num \ num\num" +recdef split_int "measure num_size" + "split_int (C c) = (C 0, C c)" + "split_int (CN n c b) = + (let (bv,bi) = split_int b + in (CN n c bv, bi))" + "split_int (CF c a b) = + (let (bv,bi) = split_int b + in (bv, CF c a bi))" + "split_int a = (a,C 0)" + +lemma split_int:"\ tv ti. split_int t = (tv,ti) \ (Inum bs (Add tv ti) = Inum bs t) \ isint ti bs" +proof (induct t rule: split_int.induct) + case (2 c n b tv ti) + let ?bv = "fst (split_int b)" + let ?bi = "snd (split_int b)" + have "split_int b = (?bv,?bi)" by simp + with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+ + from prems(2) have tibi: "ti = ?bi" by (simp add: Let_def split_def) + from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def) +next + case (3 c a b tv ti) + let ?bv = "fst (split_int b)" + let ?bi = "snd (split_int b)" + have "split_int b = (?bv,?bi)" by simp + with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+ + from prems(2) have tibi: "ti = CF c a ?bi" by (simp add: Let_def split_def) + from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def isint_Floor isint_add isint_Mul isint_CF) +qed (auto simp add: Let_def isint_iff isint_Floor isint_add isint_Mul split_def ring_simps) + +lemma split_int_nb: "numbound0 t \ numbound0 (fst (split_int t)) \ numbound0 (snd (split_int t)) " +by (induct t rule: split_int.induct, auto simp add: Let_def split_def) + +definition + numfloor:: "num \ num" +where + numfloor_def: "numfloor t = (let (tv,ti) = split_int t in + (case tv of C i \ numadd (tv,ti) + | _ \ numadd(CF 1 tv (C 0),ti)))" + +lemma numfloor[simp]: "Inum bs (numfloor t) = Inum bs (Floor t)" (is "?n t = ?N (Floor t)") +proof- + let ?tv = "fst (split_int t)" + let ?ti = "snd (split_int t)" + have tvti:"split_int t = (?tv,?ti)" by simp + {assume H: "\ v. ?tv \ C v" + hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)" + by (cases ?tv, auto simp add: numfloor_def Let_def split_def numadd) + from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ + hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp + also have "\ = real (floor (?N ?tv) + (floor (?N ?ti)))" + by (simp,subst tii[simplified isint_iff, symmetric]) simp + also have "\ = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) + finally have ?thesis using th1 by simp} + moreover {fix v assume H:"?tv = C v" + from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ + hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp + also have "\ = real (floor (?N ?tv) + (floor (?N ?ti)))" + by (simp,subst tii[simplified isint_iff, symmetric]) simp + also have "\ = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) + finally have ?thesis by (simp add: H numfloor_def Let_def split_def numadd) } + ultimately show ?thesis by auto +qed + +lemma numfloor_nb[simp]: "numbound0 t \ numbound0 (numfloor t)" + using split_int_nb[where t="t"] + by (cases "fst(split_int t)" , auto simp add: numfloor_def Let_def split_def numadd_nb) + +recdef simpnum "measure num_size" + "simpnum (C j) = C j" + "simpnum (Bound n) = CN n 1 (C 0)" + "simpnum (Neg t) = numneg (simpnum t)" + "simpnum (Add t s) = numadd (simpnum t,simpnum s)" + "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)" + "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)" + "simpnum (Floor t) = numfloor (simpnum t)" + "simpnum (CN n c t) = (if c=0 then simpnum t else CN n c (simpnum t))" + "simpnum (CF c t s) = simpnum(Add (Mul c (Floor t)) s)" + +lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t" +by (induct t rule: simpnum.induct, auto) + +lemma simpnum_numbound0[simp]: + "numbound0 t \ numbound0 (simpnum t)" +by (induct t rule: simpnum.induct, auto) + +consts nozerocoeff:: "num \ bool" +recdef nozerocoeff "measure size" + "nozerocoeff (C c) = True" + "nozerocoeff (CN n c t) = (c\0 \ nozerocoeff t)" + "nozerocoeff (CF c s t) = (c \ 0 \ nozerocoeff t)" + "nozerocoeff (Mul c t) = (c\0 \ nozerocoeff t)" + "nozerocoeff t = True" + +lemma numadd_nz : "nozerocoeff a \ nozerocoeff b \ nozerocoeff (numadd (a,b))" +by (induct a b rule: numadd.induct,auto simp add: Let_def) + +lemma nummul_nz : "\ i. i\0 \ nozerocoeff a \ nozerocoeff (nummul a i)" + by (induct a rule: nummul.induct,auto simp add: Let_def numadd_nz) + +lemma numneg_nz : "nozerocoeff a \ nozerocoeff (numneg a)" +by (simp add: numneg_def nummul_nz) + +lemma numsub_nz: "nozerocoeff a \ nozerocoeff b \ nozerocoeff (numsub a b)" +by (simp add: numsub_def numneg_nz numadd_nz) + +lemma split_int_nz: "nozerocoeff t \ nozerocoeff (fst (split_int t)) \ nozerocoeff (snd (split_int t))" +by (induct t rule: split_int.induct,auto simp add: Let_def split_def) + +lemma numfloor_nz: "nozerocoeff t \ nozerocoeff (numfloor t)" +by (simp add: numfloor_def Let_def split_def) +(cases "fst (split_int t)", simp_all add: split_int_nz numadd_nz) + +lemma simpnum_nz: "nozerocoeff (simpnum t)" +by(induct t rule: simpnum.induct, auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz numfloor_nz) + +lemma maxcoeff_nz: "nozerocoeff t \ maxcoeff t = 0 \ t = C 0" +proof (induct t rule: maxcoeff.induct) + case (2 n c t) + hence cnz: "c \0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+ + have "max (abs c) (maxcoeff t) \ abs c" by (simp add: le_maxI1) + with cnz have "max (abs c) (maxcoeff t) > 0" by arith + with prems show ?case by simp +next + case (3 c s t) + hence cnz: "c \0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+ + have "max (abs c) (maxcoeff t) \ abs c" by (simp add: le_maxI1) + with cnz have "max (abs c) (maxcoeff t) > 0" by arith + with prems show ?case by simp +qed auto + +lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0" +proof- + from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def) + from numgcdh0[OF th] have th:"maxcoeff t = 0" . + from maxcoeff_nz[OF nz th] show ?thesis . +qed + +constdefs simp_num_pair:: "(num \ int) \ num \ int" + "simp_num_pair \ (\ (t,n). (if n = 0 then (C 0, 0) else + (let t' = simpnum t ; g = numgcd t' in + if g > 1 then (let g' = zgcd n g in + if g' = 1 then (t',n) + else (reducecoeffh t' g', n div g')) + else (t',n))))" + +lemma simp_num_pair_ci: + shows "((\ (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((\ (t,n). Inum bs t / real n) (t,n))" + (is "?lhs = ?rhs") +proof- + let ?t' = "simpnum t" + let ?g = "numgcd ?t'" + let ?g' = "zgcd n ?g" + {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)} + moreover + { assume nnz: "n \ 0" + {assume "\ ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def)} + moreover + {assume g1:"?g>1" hence g0: "?g > 0" by simp + from zgcd0 g1 nnz have gp0: "?g' \ 0" by simp + hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith + hence "?g'= 1 \ ?g' > 1" by arith + moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)} + moreover {assume g'1:"?g'>1" + from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" .. + let ?tt = "reducecoeffh ?t' ?g'" + let ?t = "Inum bs ?tt" + have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) + have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1) + have gpdgp: "?g' dvd ?g'" by simp + from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] + have th2:"real ?g' * ?t = Inum bs ?t'" by simp + from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def) + also have "\ = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp + also have "\ = (Inum bs ?t' / real n)" + using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp + finally have "?lhs = Inum bs t / real n" by simp + then have ?thesis using prems by (simp add: simp_num_pair_def)} + ultimately have ?thesis by blast} + ultimately have ?thesis by blast} + ultimately show ?thesis by blast +qed + +lemma simp_num_pair_l: assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')" + shows "numbound0 t' \ n' >0" +proof- + let ?t' = "simpnum t" + let ?g = "numgcd ?t'" + let ?g' = "zgcd n ?g" + {assume nz: "n = 0" hence ?thesis using prems by (simp add: Let_def simp_num_pair_def)} + moreover + { assume nnz: "n \ 0" + {assume "\ ?g > 1" hence ?thesis using prems by (auto simp add: Let_def simp_num_pair_def)} + moreover + {assume g1:"?g>1" hence g0: "?g > 0" by simp + from zgcd0 g1 nnz have gp0: "?g' \ 0" by simp + hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith + hence "?g'= 1 \ ?g' > 1" by arith + moreover {assume "?g'=1" hence ?thesis using prems + by (auto simp add: Let_def simp_num_pair_def)} + moreover {assume g'1:"?g'>1" + have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) + have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1) + have gpdgp: "?g' dvd ?g'" by simp + from zdvd_imp_le[OF gpdd np] have g'n: "?g' \ n" . + from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]] + have "n div ?g' >0" by simp + hence ?thesis using prems + by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)} + ultimately have ?thesis by blast} + ultimately have ?thesis by blast} + ultimately show ?thesis by blast +qed + +consts not:: "fm \ fm" +recdef not "measure size" + "not (NOT p) = p" + "not T = F" + "not F = T" + "not (Lt t) = Ge t" + "not (Le t) = Gt t" + "not (Gt t) = Le t" + "not (Ge t) = Lt t" + "not (Eq t) = NEq t" + "not (NEq t) = Eq t" + "not (Dvd i t) = NDvd i t" + "not (NDvd i t) = Dvd i t" + "not (And p q) = Or (not p) (not q)" + "not (Or p q) = And (not p) (not q)" + "not p = NOT p" +lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)" +by (induct p) auto +lemma not_qf[simp]: "qfree p \ qfree (not p)" +by (induct p, auto) +lemma not_nb[simp]: "bound0 p \ bound0 (not p)" +by (induct p, auto) + +constdefs conj :: "fm \ fm \ fm" + "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else + if p = q then p else And p q)" +lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)" +by (cases "p=F \ q=F",simp_all add: conj_def) (cases p,simp_all) + +lemma conj_qf[simp]: "\qfree p ; qfree q\ \ qfree (conj p q)" +using conj_def by auto +lemma conj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (conj p q)" +using conj_def by auto + +constdefs disj :: "fm \ fm \ fm" + "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p + else if p=q then p else Or p q)" + +lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)" +by (cases "p=T \ q=T",simp_all add: disj_def) (cases p,simp_all) +lemma disj_qf[simp]: "\qfree p ; qfree q\ \ qfree (disj p q)" +using disj_def by auto +lemma disj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (disj p q)" +using disj_def by auto + +constdefs imp :: "fm \ fm \ fm" + "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p + else Imp p q)" +lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)" +by (cases "p=F \ q=T",simp_all add: imp_def) +lemma imp_qf[simp]: "\qfree p ; qfree q\ \ qfree (imp p q)" +using imp_def by (cases "p=F \ q=T",simp_all add: imp_def) +lemma imp_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (imp p q)" +using imp_def by (cases "p=F \ q=T \ p=q",simp_all add: imp_def) + +constdefs iff :: "fm \ fm \ fm" + "iff p q \ (if (p = q) then T else if (p = not q \ not p = q) then F else + if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else + Iff p q)" +lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)" + by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp add:not) +(cases "not p= q", auto simp add:not) +lemma iff_qf[simp]: "\qfree p ; qfree q\ \ qfree (iff p q)" + by (unfold iff_def,cases "p=q", auto) +lemma iff_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (iff p q)" +using iff_def by (unfold iff_def,cases "p=q", auto) + +consts check_int:: "num \ bool" +recdef check_int "measure size" + "check_int (C i) = True" + "check_int (Floor t) = True" + "check_int (Mul i t) = check_int t" + "check_int (Add t s) = (check_int t \ check_int s)" + "check_int (Neg t) = check_int t" + "check_int (CF c t s) = check_int s" + "check_int t = False" +lemma check_int: "check_int t \ isint t bs" +by (induct t, auto simp add: isint_add isint_Floor isint_Mul isint_neg isint_c isint_CF) + +lemma rdvd_left1_int: "real \t\ = t \ 1 rdvd t" + by (simp add: rdvd_def,rule_tac x="\t\" in exI) simp + +lemma rdvd_reduce: + assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0" + shows "real (d::int) rdvd real (c::int)*t = (real (d div g) rdvd real (c div g)*t)" +proof + assume d: "real d rdvd real c * t" + from d rdvd_def obtain k where k_def: "real c * t = real d* real (k::int)" by auto + from gd dvd_def obtain kd where kd_def: "d = g * kd" by auto + from gc dvd_def obtain kc where kc_def: "c = g * kc" by auto + from k_def kd_def kc_def have "real g * real kc * t = real g * real kd * real k" by simp + hence "real kc * t = real kd * real k" using gp by simp + hence th:"real kd rdvd real kc * t" using rdvd_def by blast + from kd_def gp have th':"kd = d div g" by simp + from kc_def gp have "kc = c div g" by simp + with th th' show "real (d div g) rdvd real (c div g) * t" by simp +next + assume d: "real (d div g) rdvd real (c div g) * t" + from gp have gnz: "g \ 0" by simp + thus "real d rdvd real c * t" using d rdvd_mult[OF gnz, where n="d div g" and x="real (c div g) * t"] real_of_int_div[OF gnz gd] real_of_int_div[OF gnz gc] by simp +qed + +constdefs simpdvd:: "int \ num \ (int \ num)" + "simpdvd d t \ + (let g = numgcd t in + if g > 1 then (let g' = zgcd d g in + if g' = 1 then (d, t) + else (d div g',reducecoeffh t g')) + else (d, t))" +lemma simpdvd: + assumes tnz: "nozerocoeff t" and dnz: "d \ 0" + shows "Ifm bs (Dvd (fst (simpdvd d t)) (snd (simpdvd d t))) = Ifm bs (Dvd d t)" +proof- + let ?g = "numgcd t" + let ?g' = "zgcd d ?g" + {assume "\ ?g > 1" hence ?thesis by (simp add: Let_def simpdvd_def)} + moreover + {assume g1:"?g>1" hence g0: "?g > 0" by simp + from zgcd0 g1 dnz have gp0: "?g' \ 0" by simp + hence g'p: "?g' > 0" using zgcd_pos[where i="d" and j="numgcd t"] by arith + hence "?g'= 1 \ ?g' > 1" by arith + moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simpdvd_def)} + moreover {assume g'1:"?g'>1" + from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" .. + let ?tt = "reducecoeffh t ?g'" + let ?t = "Inum bs ?tt" + have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) + have gpdd: "?g' dvd d" by (simp add: zgcd_zdvd1) + have gpdgp: "?g' dvd ?g'" by simp + from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] + have th2:"real ?g' * ?t = Inum bs t" by simp + from prems have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)" + by (simp add: simpdvd_def Let_def) + also have "\ = (real d rdvd (Inum bs t))" + using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified zdiv_self[OF gp0]] + th2[symmetric] by simp + finally have ?thesis by simp } + ultimately have ?thesis by blast + } + ultimately show ?thesis by blast +qed + +consts simpfm :: "fm \ fm" +recdef simpfm "measure fmsize" + "simpfm (And p q) = conj (simpfm p) (simpfm q)" + "simpfm (Or p q) = disj (simpfm p) (simpfm q)" + "simpfm (Imp p q) = imp (simpfm p) (simpfm q)" + "simpfm (Iff p q) = iff (simpfm p) (simpfm q)" + "simpfm (NOT p) = not (simpfm p)" + "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \ if (v < 0) then T else F + | _ \ Lt (reducecoeff a'))" + "simpfm (Le a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Le (reducecoeff a'))" + "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \ if (v > 0) then T else F | _ \ Gt (reducecoeff a'))" + "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Ge (reducecoeff a'))" + "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \ if (v = 0) then T else F | _ \ Eq (reducecoeff a'))" + "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ NEq (reducecoeff a'))" + "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a) + else if (abs i = 1) \ check_int a then T + else let a' = simpnum a in case a' of C v \ if (i dvd v) then T else F | _ \ (let (d,t) = simpdvd i a' in Dvd d t))" + "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) + else if (abs i = 1) \ check_int a then F + else let a' = simpnum a in case a' of C v \ if (\(i dvd v)) then T else F | _ \ (let (d,t) = simpdvd i a' in NDvd d t))" + "simpfm p = p" + +lemma simpfm[simp]: "Ifm bs (simpfm p) = Ifm bs p" +proof(induct p rule: simpfm.induct) + case (6 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume H:"\ (\ v. ?sa = C v)" + let ?g = "numgcd ?sa" + let ?rsa = "reducecoeff ?sa" + let ?r = "Inum bs ?rsa" + have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) + {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} + with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) + hence gp: "real ?g > 0" by simp + have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a < 0 = (real ?g * ?r < real ?g * 0)" by simp + also have "\ = (?r < 0)" using gp + by (simp only: mult_less_cancel_left) simp + finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} + ultimately show ?case by blast +next + case (7 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume H:"\ (\ v. ?sa = C v)" + let ?g = "numgcd ?sa" + let ?rsa = "reducecoeff ?sa" + let ?r = "Inum bs ?rsa" + have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) + {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} + with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) + hence gp: "real ?g > 0" by simp + have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a \ 0 = (real ?g * ?r \ real ?g * 0)" by simp + also have "\ = (?r \ 0)" using gp + by (simp only: mult_le_cancel_left) simp + finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} + ultimately show ?case by blast +next + case (8 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume H:"\ (\ v. ?sa = C v)" + let ?g = "numgcd ?sa" + let ?rsa = "reducecoeff ?sa" + let ?r = "Inum bs ?rsa" + have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) + {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} + with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) + hence gp: "real ?g > 0" by simp + have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a > 0 = (real ?g * ?r > real ?g * 0)" by simp + also have "\ = (?r > 0)" using gp + by (simp only: mult_less_cancel_left) simp + finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} + ultimately show ?case by blast +next + case (9 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume H:"\ (\ v. ?sa = C v)" + let ?g = "numgcd ?sa" + let ?rsa = "reducecoeff ?sa" + let ?r = "Inum bs ?rsa" + have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) + {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} + with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) + hence gp: "real ?g > 0" by simp + have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a \ 0 = (real ?g * ?r \ real ?g * 0)" by simp + also have "\ = (?r \ 0)" using gp + by (simp only: mult_le_cancel_left) simp + finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} + ultimately show ?case by blast +next + case (10 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume H:"\ (\ v. ?sa = C v)" + let ?g = "numgcd ?sa" + let ?rsa = "reducecoeff ?sa" + let ?r = "Inum bs ?rsa" + have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) + {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} + with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) + hence gp: "real ?g > 0" by simp + have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a = 0 = (real ?g * ?r = 0)" by simp + also have "\ = (?r = 0)" using gp + by (simp add: mult_eq_0_iff) + finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} + ultimately show ?case by blast +next + case (11 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume H:"\ (\ v. ?sa = C v)" + let ?g = "numgcd ?sa" + let ?rsa = "reducecoeff ?sa" + let ?r = "Inum bs ?rsa" + have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) + {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} + with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) + hence gp: "real ?g > 0" by simp + have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a \ 0 = (real ?g * ?r \ 0)" by simp + also have "\ = (?r \ 0)" using gp + by (simp add: mult_eq_0_iff) + finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} + ultimately show ?case by blast +next + case (12 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp + have "i=0 \ (abs i = 1 \ check_int a) \ (i\0 \ ((abs i \ 1) \ (\ check_int a)))" by auto + {assume "i=0" hence ?case using "12.hyps" by (simp add: rdvd_left_0_eq Let_def)} + moreover + {assume ai1: "abs i = 1" and ai: "check_int a" + hence "i=1 \ i= - 1" by arith + moreover {assume i1: "i = 1" + from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] + have ?case using i1 ai by simp } + moreover {assume i1: "i = - 1" + from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] + rdvd_abs1[where d="- 1" and t="Inum bs a"] + have ?case using i1 ai by simp } + ultimately have ?case by blast} + moreover + {assume inz: "i\0" and cond: "(abs i \ 1) \ (\ check_int a)" + {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond + by (cases "abs i = 1", auto simp add: int_rdvd_iff) } + moreover {assume H:"\ (\ v. ?sa = C v)" + hence th: "simpfm (Dvd i a) = Dvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond by (cases ?sa, auto simp add: Let_def split_def) + from simpnum_nz have nz:"nozerocoeff ?sa" by simp + from simpdvd [OF nz inz] th have ?case using sa by simp} + ultimately have ?case by blast} + ultimately show ?case by blast +next + case (13 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp + have "i=0 \ (abs i = 1 \ check_int a) \ (i\0 \ ((abs i \ 1) \ (\ check_int a)))" by auto + {assume "i=0" hence ?case using "13.hyps" by (simp add: rdvd_left_0_eq Let_def)} + moreover + {assume ai1: "abs i = 1" and ai: "check_int a" + hence "i=1 \ i= - 1" by arith + moreover {assume i1: "i = 1" + from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] + have ?case using i1 ai by simp } + moreover {assume i1: "i = - 1" + from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] + rdvd_abs1[where d="- 1" and t="Inum bs a"] + have ?case using i1 ai by simp } + ultimately have ?case by blast} + moreover + {assume inz: "i\0" and cond: "(abs i \ 1) \ (\ check_int a)" + {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond + by (cases "abs i = 1", auto simp add: int_rdvd_iff) } + moreover {assume H:"\ (\ v. ?sa = C v)" + hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond + by (cases ?sa, auto simp add: Let_def split_def) + from simpnum_nz have nz:"nozerocoeff ?sa" by simp + from simpdvd [OF nz inz] th have ?case using sa by simp} + ultimately have ?case by blast} + ultimately show ?case by blast +qed (induct p rule: simpfm.induct, simp_all) + +lemma simpdvd_numbound0: "numbound0 t \ numbound0 (snd (simpdvd d t))" + by (simp add: simpdvd_def Let_def split_def reducecoeffh_numbound0) + +lemma simpfm_bound0[simp]: "bound0 p \ bound0 (simpfm p)" +proof(induct p rule: simpfm.induct) + case (6 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) +next + case (7 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) +next + case (8 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) +next + case (9 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) +next + case (10 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) +next + case (11 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) +next + case (12 i a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def) +next + case (13 i a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def) +qed(auto simp add: disj_def imp_def iff_def conj_def) + +lemma simpfm_qf[simp]: "qfree p \ qfree (simpfm p)" +by (induct p rule: simpfm.induct, auto simp add: Let_def) +(case_tac "simpnum a",auto simp add: split_def Let_def)+ + + + (* Generic quantifier elimination *) + +constdefs list_conj :: "fm list \ fm" + "list_conj ps \ foldr conj ps T" +lemma list_conj: "Ifm bs (list_conj ps) = (\p\ set ps. Ifm bs p)" + by (induct ps, auto simp add: list_conj_def) +lemma list_conj_qf: " \p\ set ps. qfree p \ qfree (list_conj ps)" + by (induct ps, auto simp add: list_conj_def) +lemma list_conj_nb: " \p\ set ps. bound0 p \ bound0 (list_conj ps)" + by (induct ps, auto simp add: list_conj_def) +constdefs CJNB:: "(fm \ fm) \ fm \ fm" + "CJNB f p \ (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs + in conj (decr (list_conj yes)) (f (list_conj no)))" + +lemma CJNB_qe: + assumes qe: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" + shows "\ bs p. qfree p \ qfree (CJNB qe p) \ (Ifm bs ((CJNB qe p)) = Ifm bs (E p))" +proof(clarify) + fix bs p + assume qfp: "qfree p" + let ?cjs = "conjuncts p" + let ?yes = "fst (partition bound0 ?cjs)" + let ?no = "snd (partition bound0 ?cjs)" + let ?cno = "list_conj ?no" + let ?cyes = "list_conj ?yes" + have part: "partition bound0 ?cjs = (?yes,?no)" by simp + from partition_P[OF part] have "\ q\ set ?yes. bound0 q" by blast + hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb) + hence yes_qf: "qfree (decr ?cyes )" by (simp add: decr_qf) + from conjuncts_qf[OF qfp] partition_set[OF part] + have " \q\ set ?no. qfree q" by auto + hence no_qf: "qfree ?cno"by (simp add: list_conj_qf) + with qe have cno_qf:"qfree (qe ?cno )" + and noE: "Ifm bs (qe ?cno) = Ifm bs (E ?cno)" by blast+ + from cno_qf yes_qf have qf: "qfree (CJNB qe p)" + by (simp add: CJNB_def Let_def conj_qf split_def) + {fix bs + from conjuncts have "Ifm bs p = (\q\ set ?cjs. Ifm bs q)" by blast + also have "\ = ((\q\ set ?yes. Ifm bs q) \ (\q\ set ?no. Ifm bs q))" + using partition_set[OF part] by auto + finally have "Ifm bs p = ((Ifm bs ?cyes) \ (Ifm bs ?cno))" using list_conj by simp} + hence "Ifm bs (E p) = (\x. (Ifm (x#bs) ?cyes) \ (Ifm (x#bs) ?cno))" by simp + also fix y have "\ = (\x. (Ifm (y#bs) ?cyes) \ (Ifm (x#bs) ?cno))" + using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast + also have "\ = (Ifm bs (decr ?cyes) \ Ifm bs (E ?cno))" + by (auto simp add: decr[OF yes_nb]) + also have "\ = (Ifm bs (conj (decr ?cyes) (qe ?cno)))" + using qe[rule_format, OF no_qf] by auto + finally have "Ifm bs (E p) = Ifm bs (CJNB qe p)" + by (simp add: Let_def CJNB_def split_def) + with qf show "qfree (CJNB qe p) \ Ifm bs (CJNB qe p) = Ifm bs (E p)" by blast +qed + +consts qelim :: "fm \ (fm \ fm) \ fm" +recdef qelim "measure fmsize" + "qelim (E p) = (\ qe. DJ (CJNB qe) (qelim p qe))" + "qelim (A p) = (\ qe. not (qe ((qelim (NOT p) qe))))" + "qelim (NOT p) = (\ qe. not (qelim p qe))" + "qelim (And p q) = (\ qe. conj (qelim p qe) (qelim q qe))" + "qelim (Or p q) = (\ qe. disj (qelim p qe) (qelim q qe))" + "qelim (Imp p q) = (\ qe. disj (qelim (NOT p) qe) (qelim q qe))" + "qelim (Iff p q) = (\ qe. iff (qelim p qe) (qelim q qe))" + "qelim p = (\ y. simpfm p)" + +lemma qelim_ci: + assumes qe_inv: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" + shows "\ bs. qfree (qelim p qe) \ (Ifm bs (qelim p qe) = Ifm bs p)" +using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]] +by(induct p rule: qelim.induct) +(auto simp del: simpfm.simps) + + +text {* The @{text "\"} Part *} +text{* Linearity for fm where Bound 0 ranges over @{text "\"} *} +consts + zsplit0 :: "num \ int \ num" (* splits the bounded from the unbounded part*) +recdef zsplit0 "measure num_size" + "zsplit0 (C c) = (0,C c)" + "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))" + "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)" + "zsplit0 (CF c a b) = zsplit0 (Add (Mul c (Floor a)) b)" + "zsplit0 (Neg a) = (let (i',a') = zsplit0 a in (-i', Neg a'))" + "zsplit0 (Add a b) = (let (ia,a') = zsplit0 a ; + (ib,b') = zsplit0 b + in (ia+ib, Add a' b'))" + "zsplit0 (Sub a b) = (let (ia,a') = zsplit0 a ; + (ib,b') = zsplit0 b + in (ia-ib, Sub a' b'))" + "zsplit0 (Mul i a) = (let (i',a') = zsplit0 a in (i*i', Mul i a'))" + "zsplit0 (Floor a) = (let (i',a') = zsplit0 a in (i',Floor a'))" +(hints simp add: Let_def) + +lemma zsplit0_I: + shows "\ n a. zsplit0 t = (n,a) \ (Inum ((real (x::int)) #bs) (CN 0 n a) = Inum (real x #bs) t) \ numbound0 a" + (is "\ n a. ?S t = (n,a) \ (?I x (CN 0 n a) = ?I x t) \ ?N a") +proof(induct t rule: zsplit0.induct) + case (1 c n a) thus ?case by auto +next + case (2 m n a) thus ?case by (cases "m=0") auto +next + case (3 n i a n a') thus ?case by auto +next + case (4 c a b n a') thus ?case by auto +next + case (5 t n a) + let ?nt = "fst (zsplit0 t)" + let ?at = "snd (zsplit0 t)" + have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \ n=-?nt" using prems + by (simp add: Let_def split_def) + from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast + from th2[simplified] th[simplified] show ?case by simp +next + case (6 s t n a) + let ?ns = "fst (zsplit0 s)" + let ?as = "snd (zsplit0 s)" + let ?nt = "fst (zsplit0 t)" + let ?at = "snd (zsplit0 t)" + have abjs: "zsplit0 s = (?ns,?as)" by simp + moreover have abjt: "zsplit0 t = (?nt,?at)" by simp + ultimately have th: "a=Add ?as ?at \ n=?ns + ?nt" using prems + by (simp add: Let_def split_def) + from abjs[symmetric] have bluddy: "\ x y. (x,y) = zsplit0 s" by blast + from prems have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real x # bs) (CN 0 xa xb) = Inum (real x # bs) t \ numbound0 xb)" by simp + with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast + from abjs prems have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \ ?N ?as" by blast + from th3[simplified] th2[simplified] th[simplified] show ?case + by (simp add: left_distrib) +next + case (7 s t n a) + let ?ns = "fst (zsplit0 s)" + let ?as = "snd (zsplit0 s)" + let ?nt = "fst (zsplit0 t)" + let ?at = "snd (zsplit0 t)" + have abjs: "zsplit0 s = (?ns,?as)" by simp + moreover have abjt: "zsplit0 t = (?nt,?at)" by simp + ultimately have th: "a=Sub ?as ?at \ n=?ns - ?nt" using prems + by (simp add: Let_def split_def) + from abjs[symmetric] have bluddy: "\ x y. (x,y) = zsplit0 s" by blast + from prems have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real x # bs) (CN 0 xa xb) = Inum (real x # bs) t \ numbound0 xb)" by simp + with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast + from abjs prems have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \ ?N ?as" by blast + from th3[simplified] th2[simplified] th[simplified] show ?case + by (simp add: left_diff_distrib) +next + case (8 i t n a) + let ?nt = "fst (zsplit0 t)" + let ?at = "snd (zsplit0 t)" + have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at \ n=i*?nt" using prems + by (simp add: Let_def split_def) + from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast + hence " ?I x (Mul i t) = (real i) * ?I x (CN 0 ?nt ?at)" by simp + also have "\ = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: right_distrib) + finally show ?case using th th2 by simp +next + case (9 t n a) + let ?nt = "fst (zsplit0 t)" + let ?at = "snd (zsplit0 t)" + have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a= Floor ?at \ n=?nt" using prems + by (simp add: Let_def split_def) + from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast + hence na: "?N a" using th by simp + have th': "(real ?nt)*(real x) = real (?nt * x)" by simp + have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp + also have "\ = real (floor ((real ?nt)* real(x) + ?I x ?at))" by simp + also have "\ = real (floor (?I x ?at + real (?nt* x)))" by (simp add: add_ac) + also have "\ = real (floor (?I x ?at) + (?nt* x))" + using floor_add[where x="?I x ?at" and a="?nt* x"] by simp + also have "\ = real (?nt)*(real x) + real (floor (?I x ?at))" by (simp add: add_ac) + finally have "?I x (Floor t) = ?I x (CN 0 n a)" using th by simp + with na show ?case by simp +qed + +consts + iszlfm :: "fm \ real list \ bool" (* Linearity test for fm *) + zlfm :: "fm \ fm" (* Linearity transformation for fm *) +recdef iszlfm "measure size" + "iszlfm (And p q) = (\ bs. iszlfm p bs \ iszlfm q bs)" + "iszlfm (Or p q) = (\ bs. iszlfm p bs \ iszlfm q bs)" + "iszlfm (Eq (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" + "iszlfm (NEq (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" + "iszlfm (Lt (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" + "iszlfm (Le (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" + "iszlfm (Gt (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" + "iszlfm (Ge (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" + "iszlfm (Dvd i (CN 0 c e)) = + (\ bs. c>0 \ i>0 \ numbound0 e \ isint e bs)" + "iszlfm (NDvd i (CN 0 c e))= + (\ bs. c>0 \ i>0 \ numbound0 e \ isint e bs)" + "iszlfm p = (\ bs. isatom p \ (bound0 p))" + +lemma zlin_qfree: "iszlfm p bs \ qfree p" + by (induct p rule: iszlfm.induct) auto + +lemma iszlfm_gen: + assumes lp: "iszlfm p (x#bs)" + shows "\ y. iszlfm p (y#bs)" +proof + fix y + show "iszlfm p (y#bs)" + using lp + by(induct p rule: iszlfm.induct, simp_all add: numbound0_gen[rule_format, where x="x" and y="y"]) +qed + +lemma conj_zl[simp]: "iszlfm p bs \ iszlfm q bs \ iszlfm (conj p q) bs" + using conj_def by (cases p,auto) +lemma disj_zl[simp]: "iszlfm p bs \ iszlfm q bs \ iszlfm (disj p q) bs" + using disj_def by (cases p,auto) +lemma not_zl[simp]: "iszlfm p bs \ iszlfm (not p) bs" + by (induct p rule:iszlfm.induct ,auto) + +recdef zlfm "measure fmsize" + "zlfm (And p q) = conj (zlfm p) (zlfm q)" + "zlfm (Or p q) = disj (zlfm p) (zlfm q)" + "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)" + "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))" + "zlfm (Lt a) = (let (c,r) = zsplit0 a in + if c=0 then Lt r else + if c>0 then Or (Lt (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r))) + else Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))" + "zlfm (Le a) = (let (c,r) = zsplit0 a in + if c=0 then Le r else + if c>0 then Or (Le (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r))) + else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))" + "zlfm (Gt a) = (let (c,r) = zsplit0 a in + if c=0 then Gt r else + if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) + else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))" + "zlfm (Ge a) = (let (c,r) = zsplit0 a in + if c=0 then Ge r else + if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) + else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))" + "zlfm (Eq a) = (let (c,r) = zsplit0 a in + if c=0 then Eq r else + if c>0 then (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Eq (Add (Floor (Neg r)) r))) + else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))" + "zlfm (NEq a) = (let (c,r) = zsplit0 a in + if c=0 then NEq r else + if c>0 then (Or (NEq (CN 0 c (Neg (Floor (Neg r))))) (NEq (Add (Floor (Neg r)) r))) + else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))" + "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a) + else (let (c,r) = zsplit0 a in + if c=0 then Dvd (abs i) r else + if c>0 then And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 c (Floor r))) + else And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 (-c) (Neg (Floor r))))))" + "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a) + else (let (c,r) = zsplit0 a in + if c=0 then NDvd (abs i) r else + if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 c (Floor r))) + else Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 (-c) (Neg (Floor r))))))" + "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))" + "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))" + "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))" + "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))" + "zlfm (NOT (NOT p)) = zlfm p" + "zlfm (NOT T) = F" + "zlfm (NOT F) = T" + "zlfm (NOT (Lt a)) = zlfm (Ge a)" + "zlfm (NOT (Le a)) = zlfm (Gt a)" + "zlfm (NOT (Gt a)) = zlfm (Le a)" + "zlfm (NOT (Ge a)) = zlfm (Lt a)" + "zlfm (NOT (Eq a)) = zlfm (NEq a)" + "zlfm (NOT (NEq a)) = zlfm (Eq a)" + "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)" + "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)" + "zlfm p = p" (hints simp add: fmsize_pos) + +lemma split_int_less_real: + "(real (a::int) < b) = (a < floor b \ (a = floor b \ real (floor b) < b))" +proof( auto) + assume alb: "real a < b" and agb: "\ a < floor b" + from agb have "floor b \ a" by simp hence th: "b < real a + 1" by (simp only: floor_le_eq) + from floor_eq[OF alb th] show "a= floor b" by simp +next + assume alb: "a < floor b" + hence "real a < real (floor b)" by simp + moreover have "real (floor b) \ b" by simp ultimately show "real a < b" by arith +qed + +lemma split_int_less_real': + "(real (a::int) + b < 0) = (real a - real (floor(-b)) < 0 \ (real a - real (floor (-b)) = 0 \ real (floor (-b)) + b < 0))" +proof- + have "(real a + b <0) = (real a < -b)" by arith + with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith +qed + +lemma split_int_gt_real': + "(real (a::int) + b > 0) = (real a + real (floor b) > 0 \ (real a + real (floor b) = 0 \ real (floor b) - b < 0))" +proof- + have th: "(real a + b >0) = (real (-a) + (-b)< 0)" by arith + show ?thesis using myless[rule_format, where b="real (floor b)"] + by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) + (simp add: ring_simps diff_def[symmetric],arith) +qed + +lemma split_int_le_real: + "(real (a::int) \ b) = (a \ floor b \ (a = floor b \ real (floor b) < b))" +proof( auto) + assume alb: "real a \ b" and agb: "\ a \ floor b" + from alb have "floor (real a) \ floor b " by (simp only: floor_mono2) + hence "a \ floor b" by simp with agb show "False" by simp +next + assume alb: "a \ floor b" + hence "real a \ real (floor b)" by (simp only: floor_mono2) + also have "\\ b" by simp finally show "real a \ b" . +qed + +lemma split_int_le_real': + "(real (a::int) + b \ 0) = (real a - real (floor(-b)) \ 0 \ (real a - real (floor (-b)) = 0 \ real (floor (-b)) + b < 0))" +proof- + have "(real a + b \0) = (real a \ -b)" by arith + with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith +qed + +lemma split_int_ge_real': + "(real (a::int) + b \ 0) = (real a + real (floor b) \ 0 \ (real a + real (floor b) = 0 \ real (floor b) - b < 0))" +proof- + have th: "(real a + b \0) = (real (-a) + (-b) \ 0)" by arith + show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"]) + (simp add: ring_simps diff_def[symmetric],arith) +qed + +lemma split_int_eq_real: "(real (a::int) = b) = ( a = floor b \ b = real (floor b))" (is "?l = ?r") +by auto + +lemma split_int_eq_real': "(real (a::int) + b = 0) = ( a - floor (-b) = 0 \ real (floor (-b)) + b = 0)" (is "?l = ?r") +proof- + have "?l = (real a = -b)" by arith + with split_int_eq_real[where a="a" and b="-b"] show ?thesis by simp arith +qed + +lemma zlfm_I: + assumes qfp: "qfree p" + shows "(Ifm (real i #bs) (zlfm p) = Ifm (real i# bs) p) \ iszlfm (zlfm p) (real (i::int) #bs)" + (is "(?I (?l p) = ?I p) \ ?L (?l p)") +using qfp +proof(induct p rule: zlfm.induct) + case (5 a) + let ?c = "fst (zsplit0 a)" + let ?r = "snd (zsplit0 a)" + have spl: "zsplit0 a = (?c,?r)" by simp + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real i#bs) t" + have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith + moreover + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + by (cases "?r", simp_all add: Let_def split_def,case_tac "nat", simp_all)} + moreover + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) + also have "\ = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def diff_def) + finally have ?case using l by simp} + moreover + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) + also from cn cnz have "\ = (?I (?l (Lt a)))" by (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac, arith) + finally have ?case using l by simp} + ultimately show ?case by blast +next + case (6 a) + let ?c = "fst (zsplit0 a)" + let ?r = "snd (zsplit0 a)" + have spl: "zsplit0 a = (?c,?r)" by simp + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real i#bs) t" + have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith + moreover + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + by (cases "?r", simp_all add: Let_def split_def, case_tac "nat",simp_all)} + moreover + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Le a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + also have "\ = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_def) + finally have ?case using l by simp} + moreover + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Le a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + also from cn cnz have "\ = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac ,arith) + finally have ?case using l by simp} + ultimately show ?case by blast +next + case (7 a) + let ?c = "fst (zsplit0 a)" + let ?r = "snd (zsplit0 a)" + have spl: "zsplit0 a = (?c,?r)" by simp + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real i#bs) t" + have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith + moreover + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} + moreover + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) + also have "\ = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_def) + finally have ?case using l by simp} + moreover + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) + also from cn cnz have "\ = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac, arith) + finally have ?case using l by simp} + ultimately show ?case by blast +next + case (8 a) + let ?c = "fst (zsplit0 a)" + let ?r = "snd (zsplit0 a)" + have spl: "zsplit0 a = (?c,?r)" by simp + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real i#bs) t" + have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith + moreover + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} + moreover + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Ge a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + also have "\ = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_def) + finally have ?case using l by simp} + moreover + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Ge a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + also from cn cnz have "\ = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac, arith) + finally have ?case using l by simp} + ultimately show ?case by blast +next + case (9 a) + let ?c = "fst (zsplit0 a)" + let ?r = "snd (zsplit0 a)" + have spl: "zsplit0 a = (?c,?r)" by simp + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real i#bs) t" + have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith + moreover + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} + moreover + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) + also have "\ = (?I (?l (Eq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult) + finally have ?case using l by simp} + moreover + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) + also from cn cnz have "\ = (?I (?l (Eq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult,arith) + finally have ?case using l by simp} + ultimately show ?case by blast +next + case (10 a) + let ?c = "fst (zsplit0 a)" + let ?r = "snd (zsplit0 a)" + have spl: "zsplit0 a = (?c,?r)" by simp + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real i#bs) t" + have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith + moreover + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} + moreover + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (NEq a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + also have "\ = (?I (?l (NEq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult) + finally have ?case using l by simp} + moreover + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (NEq a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + also from cn cnz have "\ = (?I (?l (NEq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult,arith) + finally have ?case using l by simp} + ultimately show ?case by blast +next + case (11 j a) + let ?c = "fst (zsplit0 a)" + let ?r = "snd (zsplit0 a)" + have spl: "zsplit0 a = (?c,?r)" by simp + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real i#bs) t" + have "j=0 \ (j\0 \ ?c = 0) \ (j\0 \ ?c >0 \ ?c\0) \ (j\ 0 \ ?c<0 \ ?c\0)" by arith + moreover + {assume "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) + hence ?case using prems by (simp del: zlfm.simps add: rdvd_left_0_eq)} + moreover + {assume "?c=0" and "j\0" hence ?case + using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"] + by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} + moreover + {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" + using Ia by (simp add: Let_def split_def) + also have "\ = (real (abs j) rdvd real (?c*i) + (?N ?r))" + by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp + also have "\ = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ + (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" + by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) + also have "\ = (?I (?l (Dvd j a)))" using cp cnz jnz + by (simp add: Let_def split_def int_rdvd_iff[symmetric] + del: real_of_int_mult) (auto simp add: add_ac) + finally have ?case using l jnz by simp } + moreover + {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" + using Ia by (simp add: Let_def split_def) + also have "\ = (real (abs j) rdvd real (?c*i) + (?N ?r))" + by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp + also have "\ = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ + (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" + by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) + also have "\ = (?I (?l (Dvd j a)))" using cn cnz jnz + using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric] + by (simp add: Let_def split_def int_rdvd_iff[symmetric] + del: real_of_int_mult) (auto simp add: add_ac) + finally have ?case using l jnz by blast } + ultimately show ?case by blast +next + case (12 j a) + let ?c = "fst (zsplit0 a)" + let ?r = "snd (zsplit0 a)" + have spl: "zsplit0 a = (?c,?r)" by simp + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real i#bs) t" + have "j=0 \ (j\0 \ ?c = 0) \ (j\0 \ ?c >0 \ ?c\0) \ (j\ 0 \ ?c<0 \ ?c\0)" by arith + moreover + {assume "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) + hence ?case using prems by (simp del: zlfm.simps add: rdvd_left_0_eq)} + moreover + {assume "?c=0" and "j\0" hence ?case + using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"] + by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} + moreover + {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (NDvd j a) = (\ (real j rdvd (real (?c * i) + (?N ?r))))" + using Ia by (simp add: Let_def split_def) + also have "\ = (\ (real (abs j) rdvd real (?c*i) + (?N ?r)))" + by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp + also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ + (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" + by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) + also have "\ = (?I (?l (NDvd j a)))" using cp cnz jnz + by (simp add: Let_def split_def int_rdvd_iff[symmetric] + del: real_of_int_mult) (auto simp add: add_ac) + finally have ?case using l jnz by simp } + moreover + {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" + by (simp add: nb Let_def split_def isint_Floor isint_neg) + have "?I (NDvd j a) = (\ (real j rdvd (real (?c * i) + (?N ?r))))" + using Ia by (simp add: Let_def split_def) + also have "\ = (\ (real (abs j) rdvd real (?c*i) + (?N ?r)))" + by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp + also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ + (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" + by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) + also have "\ = (?I (?l (NDvd j a)))" using cn cnz jnz + using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric] + by (simp add: Let_def split_def int_rdvd_iff[symmetric] + del: real_of_int_mult) (auto simp add: add_ac) + finally have ?case using l jnz by blast } + ultimately show ?case by blast +qed auto + +text{* plusinf : Virtual substitution of @{text "+\"} + minusinf: Virtual substitution of @{text "-\"} + @{text "\"} Compute lcm @{text "d| Dvd d c*x+t \ p"} + @{text "d\"} checks if a given l divides all the ds above*} + +consts + plusinf:: "fm \ fm" + minusinf:: "fm \ fm" + \ :: "fm \ int" + d\ :: "fm \ int \ bool" + +recdef minusinf "measure size" + "minusinf (And p q) = conj (minusinf p) (minusinf q)" + "minusinf (Or p q) = disj (minusinf p) (minusinf q)" + "minusinf (Eq (CN 0 c e)) = F" + "minusinf (NEq (CN 0 c e)) = T" + "minusinf (Lt (CN 0 c e)) = T" + "minusinf (Le (CN 0 c e)) = T" + "minusinf (Gt (CN 0 c e)) = F" + "minusinf (Ge (CN 0 c e)) = F" + "minusinf p = p" + +lemma minusinf_qfree: "qfree p \ qfree (minusinf p)" + by (induct p rule: minusinf.induct, auto) + +recdef plusinf "measure size" + "plusinf (And p q) = conj (plusinf p) (plusinf q)" + "plusinf (Or p q) = disj (plusinf p) (plusinf q)" + "plusinf (Eq (CN 0 c e)) = F" + "plusinf (NEq (CN 0 c e)) = T" + "plusinf (Lt (CN 0 c e)) = F" + "plusinf (Le (CN 0 c e)) = F" + "plusinf (Gt (CN 0 c e)) = T" + "plusinf (Ge (CN 0 c e)) = T" + "plusinf p = p" + +recdef \ "measure size" + "\ (And p q) = zlcm (\ p) (\ q)" + "\ (Or p q) = zlcm (\ p) (\ q)" + "\ (Dvd i (CN 0 c e)) = i" + "\ (NDvd i (CN 0 c e)) = i" + "\ p = 1" + +recdef d\ "measure size" + "d\ (And p q) = (\ d. d\ p d \ d\ q d)" + "d\ (Or p q) = (\ d. d\ p d \ d\ q d)" + "d\ (Dvd i (CN 0 c e)) = (\ d. i dvd d)" + "d\ (NDvd i (CN 0 c e)) = (\ d. i dvd d)" + "d\ p = (\ d. True)" + +lemma delta_mono: + assumes lin: "iszlfm p bs" + and d: "d dvd d'" + and ad: "d\ p d" + shows "d\ p d'" + using lin ad d +proof(induct p rule: iszlfm.induct) + case (9 i c e) thus ?case using d + by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"]) +next + case (10 i c e) thus ?case using d + by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"]) +qed simp_all + +lemma \ : assumes lin:"iszlfm p bs" + shows "d\ p (\ p) \ \ p >0" +using lin +proof (induct p rule: iszlfm.induct) + case (1 p q) + let ?d = "\ (And p q)" + from prems zlcm_pos have dp: "?d >0" by simp + have d1: "\ p dvd \ (And p q)" using prems by simp + hence th: "d\ p ?d" + using delta_mono prems by (auto simp del: dvd_zlcm_self1) + have "\ q dvd \ (And p q)" using prems by simp + hence th': "d\ q ?d" using delta_mono prems by (auto simp del: dvd_zlcm_self2) + from th th' dp show ?case by simp +next + case (2 p q) + let ?d = "\ (And p q)" + from prems zlcm_pos have dp: "?d >0" by simp + have "\ p dvd \ (And p q)" using prems by simp hence th: "d\ p ?d" using delta_mono prems + by (auto simp del: dvd_zlcm_self1) + have "\ q dvd \ (And p q)" using prems by simp hence th': "d\ q ?d" using delta_mono prems by (auto simp del: dvd_zlcm_self2) + from th th' dp show ?case by simp +qed simp_all + + +lemma minusinf_inf: + assumes linp: "iszlfm p (a # bs)" + shows "\ (z::int). \ x < z. Ifm ((real x)#bs) (minusinf p) = Ifm ((real x)#bs) p" + (is "?P p" is "\ (z::int). \ x < z. ?I x (?M p) = ?I x p") +using linp +proof (induct p rule: minusinf.induct) + case (1 f g) + from prems have "?P f" by simp + then obtain z1 where z1_def: "\ x < z1. ?I x (?M f) = ?I x f" by blast + from prems have "?P g" by simp + then obtain z2 where z2_def: "\ x < z2. ?I x (?M g) = ?I x g" by blast + let ?z = "min z1 z2" + from z1_def z2_def have "\ x < ?z. ?I x (?M (And f g)) = ?I x (And f g)" by simp + thus ?case by blast +next + case (2 f g) from prems have "?P f" by simp + then obtain z1 where z1_def: "\ x < z1. ?I x (?M f) = ?I x f" by blast + from prems have "?P g" by simp + then obtain z2 where z2_def: "\ x < z2. ?I x (?M g) = ?I x g" by blast + let ?z = "min z1 z2" + from z1_def z2_def have "\ x < ?z. ?I x (?M (Or f g)) = ?I x (Or f g)" by simp + thus ?case by blast +next + case (3 c e) + from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp + from prems have nbe: "numbound0 e" by simp + fix y + have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Eq (CN 0 c e))) = ?I x (Eq (CN 0 c e))" + proof (simp add: less_floor_eq , rule allI, rule impI) + fix x + assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" + hence th1:"real x < - (Inum (y # bs) e / real c)" by simp + with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + by (simp only: real_mult_less_mono2[OF rcpos th1]) + hence "real c * real x + Inum (y # bs) e \ 0"using rcpos by simp + thus "real c * real x + Inum (real x # bs) e \ 0" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] by simp + qed + thus ?case by blast +next + case (4 c e) + from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp + from prems have nbe: "numbound0 e" by simp + fix y + have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (NEq (CN 0 c e))) = ?I x (NEq (CN 0 c e))" + proof (simp add: less_floor_eq , rule allI, rule impI) + fix x + assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" + hence th1:"real x < - (Inum (y # bs) e / real c)" by simp + with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + by (simp only: real_mult_less_mono2[OF rcpos th1]) + hence "real c * real x + Inum (y # bs) e \ 0"using rcpos by simp + thus "real c * real x + Inum (real x # bs) e \ 0" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] by simp + qed + thus ?case by blast +next + case (5 c e) + from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp + from prems have nbe: "numbound0 e" by simp + fix y + have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Lt (CN 0 c e))) = ?I x (Lt (CN 0 c e))" + proof (simp add: less_floor_eq , rule allI, rule impI) + fix x + assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" + hence th1:"real x < - (Inum (y # bs) e / real c)" by simp + with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + by (simp only: real_mult_less_mono2[OF rcpos th1]) + thus "real c * real x + Inum (real x # bs) e < 0" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp + qed + thus ?case by blast +next + case (6 c e) + from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp + from prems have nbe: "numbound0 e" by simp + fix y + have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Le (CN 0 c e))) = ?I x (Le (CN 0 c e))" + proof (simp add: less_floor_eq , rule allI, rule impI) + fix x + assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" + hence th1:"real x < - (Inum (y # bs) e / real c)" by simp + with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + by (simp only: real_mult_less_mono2[OF rcpos th1]) + thus "real c * real x + Inum (real x # bs) e \ 0" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp + qed + thus ?case by blast +next + case (7 c e) + from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp + from prems have nbe: "numbound0 e" by simp + fix y + have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Gt (CN 0 c e))) = ?I x (Gt (CN 0 c e))" + proof (simp add: less_floor_eq , rule allI, rule impI) + fix x + assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" + hence th1:"real x < - (Inum (y # bs) e / real c)" by simp + with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + by (simp only: real_mult_less_mono2[OF rcpos th1]) + thus "\ (real c * real x + Inum (real x # bs) e>0)" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp + qed + thus ?case by blast +next + case (8 c e) + from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp + from prems have nbe: "numbound0 e" by simp + fix y + have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Ge (CN 0 c e))) = ?I x (Ge (CN 0 c e))" + proof (simp add: less_floor_eq , rule allI, rule impI) + fix x + assume A: "real x + (1\real) \ - (Inum (y # bs) e / real c)" + hence th1:"real x < - (Inum (y # bs) e / real c)" by simp + with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + by (simp only: real_mult_less_mono2[OF rcpos th1]) + thus "\ real c * real x + Inum (real x # bs) e \ 0" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp + qed + thus ?case by blast +qed simp_all + +lemma minusinf_repeats: + assumes d: "d\ p d" and linp: "iszlfm p (a # bs)" + shows "Ifm ((real(x - k*d))#bs) (minusinf p) = Ifm (real x #bs) (minusinf p)" +using linp d +proof(induct p rule: iszlfm.induct) + case (9 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ + hence "\ k. d=i*k" by (simp add: dvd_def) + then obtain "di" where di_def: "d=i*di" by blast + show ?case + proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI) + assume + "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" + (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") + hence "\ (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def) + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" + by (simp add: ring_simps di_def) + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))" + by (simp add: ring_simps) + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast + thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp + next + assume + "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e") + hence "\ (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def) + hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp + hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def) + hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: ring_simps) + hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)" + by blast + thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp + qed +next + case (10 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ + hence "\ k. d=i*k" by (simp add: dvd_def) + then obtain "di" where di_def: "d=i*di" by blast + show ?case + proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI) + assume + "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" + (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") + hence "\ (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def) + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" + by (simp add: ring_simps di_def) + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))" + by (simp add: ring_simps) + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast + thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp + next + assume + "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e") + hence "\ (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def) + hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp + hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def) + hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: ring_simps) + hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)" + by blast + thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp + qed +qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="real(x - k*d)" and b'="real x"] simp del: real_of_int_mult real_of_int_diff) + +lemma minusinf_ex: + assumes lin: "iszlfm p (real (a::int) #bs)" + and exmi: "\ (x::int). Ifm (real x#bs) (minusinf p)" (is "\ x. ?P1 x") + shows "\ (x::int). Ifm (real x#bs) p" (is "\ x. ?P x") +proof- + let ?d = "\ p" + from \ [OF lin] have dpos: "?d >0" by simp + from \ [OF lin] have alld: "d\ p ?d" by simp + from minusinf_repeats[OF alld lin] have th1:"\ x k. ?P1 x = ?P1 (x - (k * ?d))" by simp + from minusinf_inf[OF lin] have th2:"\ z. \ x. x (?P x = ?P1 x)" by blast + from minusinfinity [OF dpos th1 th2] exmi show ?thesis by blast +qed + +lemma minusinf_bex: + assumes lin: "iszlfm p (real (a::int) #bs)" + shows "(\ (x::int). Ifm (real x#bs) (minusinf p)) = + (\ (x::int)\ {1..\ p}. Ifm (real x#bs) (minusinf p))" + (is "(\ x. ?P x) = _") +proof- + let ?d = "\ p" + from \ [OF lin] have dpos: "?d >0" by simp + from \ [OF lin] have alld: "d\ p ?d" by simp + from minusinf_repeats[OF alld lin] have th1:"\ x k. ?P x = ?P (x - (k * ?d))" by simp + from periodic_finite_ex[OF dpos th1] show ?thesis by blast +qed + +lemma dvd1_eq1: "x >0 \ (x::int) dvd 1 = (x = 1)" by auto + +consts + a\ :: "fm \ int \ fm" (* adjusts the coeffitients of a formula *) + d\ :: "fm \ int \ bool" (* tests if all coeffs c of c divide a given l*) + \ :: "fm \ int" (* computes the lcm of all coefficients of x*) + \ :: "fm \ num list" + \ :: "fm \ num list" + +recdef a\ "measure size" + "a\ (And p q) = (\ k. And (a\ p k) (a\ q k))" + "a\ (Or p q) = (\ k. Or (a\ p k) (a\ q k))" + "a\ (Eq (CN 0 c e)) = (\ k. Eq (CN 0 1 (Mul (k div c) e)))" + "a\ (NEq (CN 0 c e)) = (\ k. NEq (CN 0 1 (Mul (k div c) e)))" + "a\ (Lt (CN 0 c e)) = (\ k. Lt (CN 0 1 (Mul (k div c) e)))" + "a\ (Le (CN 0 c e)) = (\ k. Le (CN 0 1 (Mul (k div c) e)))" + "a\ (Gt (CN 0 c e)) = (\ k. Gt (CN 0 1 (Mul (k div c) e)))" + "a\ (Ge (CN 0 c e)) = (\ k. Ge (CN 0 1 (Mul (k div c) e)))" + "a\ (Dvd i (CN 0 c e)) =(\ k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" + "a\ (NDvd i (CN 0 c e))=(\ k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" + "a\ p = (\ k. p)" + +recdef d\ "measure size" + "d\ (And p q) = (\ k. (d\ p k) \ (d\ q k))" + "d\ (Or p q) = (\ k. (d\ p k) \ (d\ q k))" + "d\ (Eq (CN 0 c e)) = (\ k. c dvd k)" + "d\ (NEq (CN 0 c e)) = (\ k. c dvd k)" + "d\ (Lt (CN 0 c e)) = (\ k. c dvd k)" + "d\ (Le (CN 0 c e)) = (\ k. c dvd k)" + "d\ (Gt (CN 0 c e)) = (\ k. c dvd k)" + "d\ (Ge (CN 0 c e)) = (\ k. c dvd k)" + "d\ (Dvd i (CN 0 c e)) =(\ k. c dvd k)" + "d\ (NDvd i (CN 0 c e))=(\ k. c dvd k)" + "d\ p = (\ k. True)" + +recdef \ "measure size" + "\ (And p q) = zlcm (\ p) (\ q)" + "\ (Or p q) = zlcm (\ p) (\ q)" + "\ (Eq (CN 0 c e)) = c" + "\ (NEq (CN 0 c e)) = c" + "\ (Lt (CN 0 c e)) = c" + "\ (Le (CN 0 c e)) = c" + "\ (Gt (CN 0 c e)) = c" + "\ (Ge (CN 0 c e)) = c" + "\ (Dvd i (CN 0 c e)) = c" + "\ (NDvd i (CN 0 c e))= c" + "\ p = 1" + +recdef \ "measure size" + "\ (And p q) = (\ p @ \ q)" + "\ (Or p q) = (\ p @ \ q)" + "\ (Eq (CN 0 c e)) = [Sub (C -1) e]" + "\ (NEq (CN 0 c e)) = [Neg e]" + "\ (Lt (CN 0 c e)) = []" + "\ (Le (CN 0 c e)) = []" + "\ (Gt (CN 0 c e)) = [Neg e]" + "\ (Ge (CN 0 c e)) = [Sub (C -1) e]" + "\ p = []" + +recdef \ "measure size" + "\ (And p q) = (\ p @ \ q)" + "\ (Or p q) = (\ p @ \ q)" + "\ (Eq (CN 0 c e)) = [Add (C -1) e]" + "\ (NEq (CN 0 c e)) = [e]" + "\ (Lt (CN 0 c e)) = [e]" + "\ (Le (CN 0 c e)) = [Add (C -1) e]" + "\ (Gt (CN 0 c e)) = []" + "\ (Ge (CN 0 c e)) = []" + "\ p = []" +consts mirror :: "fm \ fm" +recdef mirror "measure size" + "mirror (And p q) = And (mirror p) (mirror q)" + "mirror (Or p q) = Or (mirror p) (mirror q)" + "mirror (Eq (CN 0 c e)) = Eq (CN 0 c (Neg e))" + "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))" + "mirror (Lt (CN 0 c e)) = Gt (CN 0 c (Neg e))" + "mirror (Le (CN 0 c e)) = Ge (CN 0 c (Neg e))" + "mirror (Gt (CN 0 c e)) = Lt (CN 0 c (Neg e))" + "mirror (Ge (CN 0 c e)) = Le (CN 0 c (Neg e))" + "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))" + "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))" + "mirror p = p" + +lemma mirror\\: + assumes lp: "iszlfm p (a#bs)" + shows "(Inum (real (i::int)#bs)) ` set (\ p) = (Inum (real i#bs)) ` set (\ (mirror p))" +using lp +by (induct p rule: mirror.induct, auto) + +lemma mirror: + assumes lp: "iszlfm p (a#bs)" + shows "Ifm (real (x::int)#bs) (mirror p) = Ifm (real (- x)#bs) p" +using lp +proof(induct p rule: iszlfm.induct) + case (9 j c e) + have th: "(real j rdvd real c * real x - Inum (real x # bs) e) = + (real j rdvd - (real c * real x - Inum (real x # bs) e))" + by (simp only: rdvd_minus[symmetric]) + from prems show ?case + by (simp add: ring_simps th[simplified ring_simps] + numbound0_I[where bs="bs" and b'="real x" and b="- real x"]) +next + case (10 j c e) + have th: "(real j rdvd real c * real x - Inum (real x # bs) e) = + (real j rdvd - (real c * real x - Inum (real x # bs) e))" + by (simp only: rdvd_minus[symmetric]) + from prems show ?case + by (simp add: ring_simps th[simplified ring_simps] + numbound0_I[where bs="bs" and b'="real x" and b="- real x"]) +qed (auto simp add: numbound0_I[where bs="bs" and b="real x" and b'="- real x"] nth_pos2) + +lemma mirror_l: "iszlfm p (a#bs) \ iszlfm (mirror p) (a#bs)" +by (induct p rule: mirror.induct, auto simp add: isint_neg) + +lemma mirror_d\: "iszlfm p (a#bs) \ d\ p 1 + \ iszlfm (mirror p) (a#bs) \ d\ (mirror p) 1" +by (induct p rule: mirror.induct, auto simp add: isint_neg) + +lemma mirror_\: "iszlfm p (a#bs) \ \ (mirror p) = \ p" +by (induct p rule: mirror.induct,auto) + + +lemma mirror_ex: + assumes lp: "iszlfm p (real (i::int)#bs)" + shows "(\ (x::int). Ifm (real x#bs) (mirror p)) = (\ (x::int). Ifm (real x#bs) p)" + (is "(\ x. ?I x ?mp) = (\ x. ?I x p)") +proof(auto) + fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast + thus "\ x. ?I x p" by blast +next + fix x assume "?I x p" hence "?I (- x) ?mp" + using mirror[OF lp, where x="- x", symmetric] by auto + thus "\ x. ?I x ?mp" by blast +qed + +lemma \_numbound0: assumes lp: "iszlfm p bs" + shows "\ b\ set (\ p). numbound0 b" + using lp by (induct p rule: \.induct,auto) + +lemma d\_mono: + assumes linp: "iszlfm p (a #bs)" + and dr: "d\ p l" + and d: "l dvd l'" + shows "d\ p l'" +using dr linp zdvd_trans[where n="l" and k="l'", simplified d] +by (induct p rule: iszlfm.induct) simp_all + +lemma \_l: assumes lp: "iszlfm p (a#bs)" + shows "\ b\ set (\ p). numbound0 b \ isint b (a#bs)" +using lp +by(induct p rule: \.induct, auto simp add: isint_add isint_c) + +lemma \: + assumes linp: "iszlfm p (a #bs)" + shows "\ p > 0 \ d\ p (\ p)" +using linp +proof(induct p rule: iszlfm.induct) + case (1 p q) + from prems have dl1: "\ p dvd zlcm (\ p) (\ q)" by simp + from prems have dl2: "\ q dvd zlcm (\ p) (\ q)" by simp + from prems d\_mono[where p = "p" and l="\ p" and l'="zlcm (\ p) (\ q)"] + d\_mono[where p = "q" and l="\ q" and l'="zlcm (\ p) (\ q)"] + dl1 dl2 show ?case by (auto simp add: zlcm_pos) +next + case (2 p q) + from prems have dl1: "\ p dvd zlcm (\ p) (\ q)" by simp + from prems have dl2: "\ q dvd zlcm (\ p) (\ q)" by simp + from prems d\_mono[where p = "p" and l="\ p" and l'="zlcm (\ p) (\ q)"] + d\_mono[where p = "q" and l="\ q" and l'="zlcm (\ p) (\ q)"] + dl1 dl2 show ?case by (auto simp add: zlcm_pos) +qed (auto simp add: zlcm_pos) + +lemma a\: assumes linp: "iszlfm p (a #bs)" and d: "d\ p l" and lp: "l > 0" + shows "iszlfm (a\ p l) (a #bs) \ d\ (a\ p l) 1 \ (Ifm (real (l * x) #bs) (a\ p l) = Ifm ((real x)#bs) p)" +using linp d +proof (induct p rule: iszlfm.induct) + case (5 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ + from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) + from cp have cnz: "c \ 0" by simp + have "c div c\ l div c" + by (simp add: zdiv_mono1[OF clel cp]) + then have ldcp:"0 < l div c" + by (simp add: zdiv_self[OF cnz]) + have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + by simp + hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0\real)) = + (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e < 0)" + by simp + also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) < (real (l div c)) * 0)" by (simp add: ring_simps) + also have "\ = (real c * real x + Inum (real x # bs) e < 0)" + using mult_less_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp +next + case (6 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ + from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) + from cp have cnz: "c \ 0" by simp + have "c div c\ l div c" + by (simp add: zdiv_mono1[OF clel cp]) + then have ldcp:"0 < l div c" + by (simp add: zdiv_self[OF cnz]) + have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + by simp + hence "(real l * real x + real (l div c) * Inum (real x # bs) e \ (0\real)) = + (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \ 0)" + by simp + also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (real (l div c)) * 0)" by (simp add: ring_simps) + also have "\ = (real c * real x + Inum (real x # bs) e \ 0)" + using mult_le_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp +next + case (7 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ + from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) + from cp have cnz: "c \ 0" by simp + have "c div c\ l div c" + by (simp add: zdiv_mono1[OF clel cp]) + then have ldcp:"0 < l div c" + by (simp add: zdiv_self[OF cnz]) + have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + by simp + hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0\real)) = + (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e > 0)" + by simp + also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) > (real (l div c)) * 0)" by (simp add: ring_simps) + also have "\ = (real c * real x + Inum (real x # bs) e > 0)" + using zero_less_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp +next + case (8 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ + from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) + from cp have cnz: "c \ 0" by simp + have "c div c\ l div c" + by (simp add: zdiv_mono1[OF clel cp]) + then have ldcp:"0 < l div c" + by (simp add: zdiv_self[OF cnz]) + have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + by simp + hence "(real l * real x + real (l div c) * Inum (real x # bs) e \ (0\real)) = + (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \ 0)" + by simp + also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (real (l div c)) * 0)" by (simp add: ring_simps) + also have "\ = (real c * real x + Inum (real x # bs) e \ 0)" + using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp +next + case (3 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ + from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) + from cp have cnz: "c \ 0" by simp + have "c div c\ l div c" + by (simp add: zdiv_mono1[OF clel cp]) + then have ldcp:"0 < l div c" + by (simp add: zdiv_self[OF cnz]) + have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + by simp + hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0\real)) = + (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = 0)" + by simp + also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) = (real (l div c)) * 0)" by (simp add: ring_simps) + also have "\ = (real c * real x + Inum (real x # bs) e = 0)" + using mult_eq_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp +next + case (4 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ + from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) + from cp have cnz: "c \ 0" by simp + have "c div c\ l div c" + by (simp add: zdiv_mono1[OF clel cp]) + then have ldcp:"0 < l div c" + by (simp add: zdiv_self[OF cnz]) + have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + by simp + hence "(real l * real x + real (l div c) * Inum (real x # bs) e \ (0\real)) = + (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \ 0)" + by simp + also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (real (l div c)) * 0)" by (simp add: ring_simps) + also have "\ = (real c * real x + Inum (real x # bs) e \ 0)" + using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp +next + case (9 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+ + from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) + from cp have cnz: "c \ 0" by simp + have "c div c\ l div c" + by (simp add: zdiv_mono1[OF clel cp]) + then have ldcp:"0 < l div c" + by (simp add: zdiv_self[OF cnz]) + have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + by simp + hence "(\ (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\ (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)" by simp + also have "\ = (\ (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: ring_simps) + also fix k have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)" + using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp + also have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp + finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp +next + case (10 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+ + from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) + from cp have cnz: "c \ 0" by simp + have "c div c\ l div c" + by (simp add: zdiv_mono1[OF clel cp]) + then have ldcp:"0 < l div c" + by (simp add: zdiv_self[OF cnz]) + have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + by simp + hence "(\ (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\ (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)" by simp + also have "\ = (\ (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: ring_simps) + also fix k have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)" + using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp + also have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp + finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp +qed (simp_all add: nth_pos2 numbound0_I[where bs="bs" and b="real (l * x)" and b'="real x"] isint_Mul del: real_of_int_mult) + +lemma a\_ex: assumes linp: "iszlfm p (a#bs)" and d: "d\ p l" and lp: "l>0" + shows "(\ x. l dvd x \ Ifm (real x #bs) (a\ p l)) = (\ (x::int). Ifm (real x#bs) p)" + (is "(\ x. l dvd x \ ?P x) = (\ x. ?P' x)") +proof- + have "(\ x. l dvd x \ ?P x) = (\ (x::int). ?P (l*x))" + using unity_coeff_ex[where l="l" and P="?P", simplified] by simp + also have "\ = (\ (x::int). ?P' x)" using a\[OF linp d lp] by simp + finally show ?thesis . +qed + +lemma \: + assumes lp: "iszlfm p (a#bs)" + and u: "d\ p 1" + and d: "d\ p d" + and dp: "d > 0" + and nob: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real x = b + real j)" + and p: "Ifm (real x#bs) p" (is "?P x") + shows "?P (x - d)" +using lp u d dp nob p +proof(induct p rule: iszlfm.induct) + case (5 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ + with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] prems + show ?case by (simp del: real_of_int_minus) +next + case (6 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ + with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] prems + show ?case by (simp del: real_of_int_minus) +next + case (7 c e) hence p: "Ifm (real x #bs) (Gt (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" and ie1:"isint e (a#bs)" using dvd1_eq1[where x="c"] by simp+ + let ?e = "Inum (real x # bs) e" + from ie1 have ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="a#bs"] + numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"] + by (simp add: isint_iff) + {assume "real (x-d) +?e > 0" hence ?case using c1 + numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] + by (simp del: real_of_int_minus)} + moreover + {assume H: "\ real (x-d) + ?e > 0" + let ?v="Neg e" + have vb: "?v \ set (\ (Gt (CN 0 c e)))" by simp + from prems(11)[simplified simp_thms Inum.simps \.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] + have nob: "\ (\ j\ {1 ..d}. real x = - ?e + real j)" by auto + from H p have "real x + ?e > 0 \ real x + ?e \ real d" by (simp add: c1) + hence "real (x + floor ?e) > real (0::int) \ real (x + floor ?e) \ real d" + using ie by simp + hence "x + floor ?e \ 1 \ x + floor ?e \ d" by simp + hence "\ (j::int) \ {1 .. d}. j = x + floor ?e" by simp + hence "\ (j::int) \ {1 .. d}. real x = real (- floor ?e + j)" + by (simp only: real_of_int_inject) (simp add: ring_simps) + hence "\ (j::int) \ {1 .. d}. real x = - ?e + real j" + by (simp add: ie[simplified isint_iff]) + with nob have ?case by auto} + ultimately show ?case by blast +next + case (8 c e) hence p: "Ifm (real x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" + and ie1:"isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ + let ?e = "Inum (real x # bs) e" + from ie1 have ie: "real (floor ?e) = ?e" using numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real x)#bs"] + by (simp add: isint_iff) + {assume "real (x-d) +?e \ 0" hence ?case using c1 + numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] + by (simp del: real_of_int_minus)} + moreover + {assume H: "\ real (x-d) + ?e \ 0" + let ?v="Sub (C -1) e" + have vb: "?v \ set (\ (Ge (CN 0 c e)))" by simp + from prems(11)[simplified simp_thms Inum.simps \.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] + have nob: "\ (\ j\ {1 ..d}. real x = - ?e - 1 + real j)" by auto + from H p have "real x + ?e \ 0 \ real x + ?e < real d" by (simp add: c1) + hence "real (x + floor ?e) \ real (0::int) \ real (x + floor ?e) < real d" + using ie by simp + hence "x + floor ?e +1 \ 1 \ x + floor ?e + 1 \ d" by simp + hence "\ (j::int) \ {1 .. d}. j = x + floor ?e + 1" by simp + hence "\ (j::int) \ {1 .. d}. x= - floor ?e - 1 + j" by (simp add: ring_simps) + hence "\ (j::int) \ {1 .. d}. real x= real (- floor ?e - 1 + j)" + by (simp only: real_of_int_inject) + hence "\ (j::int) \ {1 .. d}. real x= - ?e - 1 + real j" + by (simp add: ie[simplified isint_iff]) + with nob have ?case by simp } + ultimately show ?case by blast +next + case (3 c e) hence p: "Ifm (real x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" + and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ + let ?e = "Inum (real x # bs) e" + let ?v="(Sub (C -1) e)" + have vb: "?v \ set (\ (Eq (CN 0 c e)))" by simp + from p have "real x= - ?e" by (simp add: c1) with prems(11) show ?case using dp + by simp (erule ballE[where x="1"], + simp_all add:ring_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"]) +next + case (4 c e)hence p: "Ifm (real x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" + and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ + let ?e = "Inum (real x # bs) e" + let ?v="Neg e" + have vb: "?v \ set (\ (NEq (CN 0 c e)))" by simp + {assume "real x - real d + Inum ((real (x -d)) # bs) e \ 0" + hence ?case by (simp add: c1)} + moreover + {assume H: "real x - real d + Inum ((real (x -d)) # bs) e = 0" + hence "real x = - Inum ((real (x -d)) # bs) e + real d" by simp + hence "real x = - Inum (a # bs) e + real d" + by (simp add: numbound0_I[OF bn,where b="real x - real d"and b'="a"and bs="bs"]) + with prems(11) have ?case using dp by simp} + ultimately show ?case by blast +next + case (9 j c e) hence p: "Ifm (real x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" + and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ + let ?e = "Inum (real x # bs) e" + from prems have "isint e (a #bs)" by simp + hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real x)#bs"] numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] + by (simp add: isint_iff) + from prems have id: "j dvd d" by simp + from c1 ie[symmetric] have "?p x = (real j rdvd real (x+ floor ?e))" by simp + also have "\ = (j dvd x + floor ?e)" + using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp + also have "\ = (j dvd x - d + floor ?e)" + using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp + also have "\ = (real j rdvd real (x - d + floor ?e))" + using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified] + ie by simp + also have "\ = (real j rdvd real x - real d + ?e)" + using ie by simp + finally show ?case + using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp +next + case (10 j c e) hence p: "Ifm (real x #bs) (NDvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ + let ?e = "Inum (real x # bs) e" + from prems have "isint e (a#bs)" by simp + hence ie: "real (floor ?e) = ?e" using numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real x)#bs"] + by (simp add: isint_iff) + from prems have id: "j dvd d" by simp + from c1 ie[symmetric] have "?p x = (\ real j rdvd real (x+ floor ?e))" by simp + also have "\ = (\ j dvd x + floor ?e)" + using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp + also have "\ = (\ j dvd x - d + floor ?e)" + using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp + also have "\ = (\ real j rdvd real (x - d + floor ?e))" + using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified] + ie by simp + also have "\ = (\ real j rdvd real x - real d + ?e)" + using ie by simp + finally show ?case using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp +qed (auto simp add: numbound0_I[where bs="bs" and b="real (x - d)" and b'="real x"] nth_pos2 simp del: real_of_int_diff) + +lemma \': + assumes lp: "iszlfm p (a #bs)" + and u: "d\ p 1" + and d: "d\ p d" + and dp: "d > 0" + shows "\ x. \(\(j::int) \ {1 .. d}. \ b\ set(\ p). Ifm ((Inum (a#bs) b + real j) #bs) p) \ Ifm (real x#bs) p \ Ifm (real (x - d)#bs) p" (is "\ x. ?b \ ?P x \ ?P (x - d)") +proof(clarify) + fix x + assume nb:"?b" and px: "?P x" + hence nb2: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real x = b + real j)" + by auto + from \[OF lp u d dp nb2 px] show "?P (x -d )" . +qed + +lemma \_int: assumes lp: "iszlfm p bs" + shows "\ b\ set (\ p). isint b bs" +using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub) + +lemma cpmi_eq: "0 < D \ (EX z::int. ALL x. x < z --> (P x = P1 x)) +==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) +==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D)))) +==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))" +apply(rule iffI) +prefer 2 +apply(drule minusinfinity) +apply assumption+ +apply(fastsimp) +apply clarsimp +apply(subgoal_tac "!!k. 0<=k \ !x. P x \ P (x - k*D)") +apply(frule_tac x = x and z=z in decr_lemma) +apply(subgoal_tac "P1(x - (\x - z\ + 1) * D)") +prefer 2 +apply(subgoal_tac "0 <= (\x - z\ + 1)") +prefer 2 apply arith + apply fastsimp +apply(drule (1) periodic_finite_ex) +apply blast +apply(blast dest:decr_mult_lemma) +done + + +theorem cp_thm: + assumes lp: "iszlfm p (a #bs)" + and u: "d\ p 1" + and d: "d\ p d" + and dp: "d > 0" + shows "(\ (x::int). Ifm (real x #bs) p) = (\ j\ {1.. d}. Ifm (real j #bs) (minusinf p) \ (\ b \ set (\ p). Ifm ((Inum (a#bs) b + real j) #bs) p))" + (is "(\ (x::int). ?P (real x)) = (\ j\ ?D. ?M j \ (\ b\ ?B. ?P (?I b + real j)))") +proof- + from minusinf_inf[OF lp] + have th: "\(z::int). \x_int[OF lp] isint_iff[where bs="a # bs"] have B: "\ b\ ?B. real (floor (?I b)) = ?I b" by simp + from B[rule_format] + have "(\j\?D. \b\ ?B. ?P (?I b + real j)) = (\j\?D. \b\ ?B. ?P (real (floor (?I b)) + real j))" + by simp + also have "\ = (\j\?D. \b\ ?B. ?P (real (floor (?I b) + j)))" by simp + also have"\ = (\ j \ ?D. \ b \ ?B'. ?P (real (b + j)))" by blast + finally have BB': + "(\j\?D. \b\ ?B. ?P (?I b + real j)) = (\ j \ ?D. \ b \ ?B'. ?P (real (b + j)))" + by blast + hence th2: "\ x. \ (\ j \ ?D. \ b \ ?B'. ?P (real (b + j))) \ ?P (real x) \ ?P (real (x - d))" using \'[OF lp u d dp] by blast + from minusinf_repeats[OF d lp] + have th3: "\ x k. ?M x = ?M (x-k*d)" by simp + from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast +qed + + (* Reddy and Loveland *) + + +consts + \ :: "fm \ (num \ int) list" (* Compute the Reddy and Loveland Bset*) + \\:: "fm \ num \ int \ fm" (* Performs the modified substitution of Reddy and Loveland*) + \\ :: "fm \ (num\int) list" + a\ :: "fm \ int \ fm" +recdef \ "measure size" + "\ (And p q) = (\ p @ \ q)" + "\ (Or p q) = (\ p @ \ q)" + "\ (Eq (CN 0 c e)) = [(Sub (C -1) e,c)]" + "\ (NEq (CN 0 c e)) = [(Neg e,c)]" + "\ (Lt (CN 0 c e)) = []" + "\ (Le (CN 0 c e)) = []" + "\ (Gt (CN 0 c e)) = [(Neg e, c)]" + "\ (Ge (CN 0 c e)) = [(Sub (C (-1)) e, c)]" + "\ p = []" + +recdef \\ "measure size" + "\\ (And p q) = (\ (t,k). And (\\ p (t,k)) (\\ q (t,k)))" + "\\ (Or p q) = (\ (t,k). Or (\\ p (t,k)) (\\ q (t,k)))" + "\\ (Eq (CN 0 c e)) = (\ (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e)) + else (Eq (Add (Mul c t) (Mul k e))))" + "\\ (NEq (CN 0 c e)) = (\ (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e)) + else (NEq (Add (Mul c t) (Mul k e))))" + "\\ (Lt (CN 0 c e)) = (\ (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e)) + else (Lt (Add (Mul c t) (Mul k e))))" + "\\ (Le (CN 0 c e)) = (\ (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e)) + else (Le (Add (Mul c t) (Mul k e))))" + "\\ (Gt (CN 0 c e)) = (\ (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e)) + else (Gt (Add (Mul c t) (Mul k e))))" + "\\ (Ge (CN 0 c e)) = (\ (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e)) + else (Ge (Add (Mul c t) (Mul k e))))" + "\\ (Dvd i (CN 0 c e)) =(\ (t,k). if k dvd c then (Dvd i (Add (Mul (c div k) t) e)) + else (Dvd (i*k) (Add (Mul c t) (Mul k e))))" + "\\ (NDvd i (CN 0 c e))=(\ (t,k). if k dvd c then (NDvd i (Add (Mul (c div k) t) e)) + else (NDvd (i*k) (Add (Mul c t) (Mul k e))))" + "\\ p = (\ (t,k). p)" + +recdef \\ "measure size" + "\\ (And p q) = (\\ p @ \\ q)" + "\\ (Or p q) = (\\ p @ \\ q)" + "\\ (Eq (CN 0 c e)) = [(Add (C -1) e,c)]" + "\\ (NEq (CN 0 c e)) = [(e,c)]" + "\\ (Lt (CN 0 c e)) = [(e,c)]" + "\\ (Le (CN 0 c e)) = [(Add (C -1) e,c)]" + "\\ p = []" + + (* Simulates normal substituion by modifying the formula see correctness theorem *) + +recdef a\ "measure size" + "a\ (And p q) = (\ k. And (a\ p k) (a\ q k))" + "a\ (Or p q) = (\ k. Or (a\ p k) (a\ q k))" + "a\ (Eq (CN 0 c e)) = (\ k. if k dvd c then (Eq (CN 0 (c div k) e)) + else (Eq (CN 0 c (Mul k e))))" + "a\ (NEq (CN 0 c e)) = (\ k. if k dvd c then (NEq (CN 0 (c div k) e)) + else (NEq (CN 0 c (Mul k e))))" + "a\ (Lt (CN 0 c e)) = (\ k. if k dvd c then (Lt (CN 0 (c div k) e)) + else (Lt (CN 0 c (Mul k e))))" + "a\ (Le (CN 0 c e)) = (\ k. if k dvd c then (Le (CN 0 (c div k) e)) + else (Le (CN 0 c (Mul k e))))" + "a\ (Gt (CN 0 c e)) = (\ k. if k dvd c then (Gt (CN 0 (c div k) e)) + else (Gt (CN 0 c (Mul k e))))" + "a\ (Ge (CN 0 c e)) = (\ k. if k dvd c then (Ge (CN 0 (c div k) e)) + else (Ge (CN 0 c (Mul k e))))" + "a\ (Dvd i (CN 0 c e)) = (\ k. if k dvd c then (Dvd i (CN 0 (c div k) e)) + else (Dvd (i*k) (CN 0 c (Mul k e))))" + "a\ (NDvd i (CN 0 c e)) = (\ k. if k dvd c then (NDvd i (CN 0 (c div k) e)) + else (NDvd (i*k) (CN 0 c (Mul k e))))" + "a\ p = (\ k. p)" + +constdefs \ :: "fm \ int \ num \ fm" + "\ p k t \ And (Dvd k t) (\\ p (t,k))" + +lemma \\: + assumes linp: "iszlfm p (real (x::int)#bs)" + and kpos: "real k > 0" + and tnb: "numbound0 t" + and tint: "isint t (real x#bs)" + and kdt: "k dvd floor (Inum (b'#bs) t)" + shows "Ifm (real x#bs) (\\ p (t,k)) = + (Ifm ((real ((floor (Inum (b'#bs) t)) div k))#bs) p)" + (is "?I (real x) (?s p) = (?I (real ((floor (?N b' t)) div k)) p)" is "_ = (?I ?tk p)") +using linp kpos tnb +proof(induct p rule: \\.induct) + case (3 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + {assume kdc: "k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + moreover + {assume "\ k dvd c" + from kpos have knz: "k\0" by simp hence knz': "real k \ 0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have "?I (real x) (?s (Eq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k = 0)" + using real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) + also have "\ = (?I ?tk (Eq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + by (simp add: ti) + finally have ?case . } + ultimately show ?case by blast +next + case (4 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + {assume kdc: "k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + moreover + {assume "\ k dvd c" + from kpos have knz: "k\0" by simp hence knz': "real k \ 0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have "?I (real x) (?s (NEq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \ 0)" + using real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) + also have "\ = (?I ?tk (NEq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + by (simp add: ti) + finally have ?case . } + ultimately show ?case by blast +next + case (5 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + {assume kdc: "k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + moreover + {assume "\ k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have "?I (real x) (?s (Lt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k < 0)" + using real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) + also have "\ = (?I ?tk (Lt (CN 0 c e)))" using pos_less_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + by (simp add: ti) + finally have ?case . } + ultimately show ?case by blast +next + case (6 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + {assume kdc: "k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + moreover + {assume "\ k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have "?I (real x) (?s (Le (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \ 0)" + using real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) + also have "\ = (?I ?tk (Le (CN 0 c e)))" using pos_le_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + by (simp add: ti) + finally have ?case . } + ultimately show ?case by blast +next + case (7 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + {assume kdc: "k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + moreover + {assume "\ k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have "?I (real x) (?s (Gt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k > 0)" + using real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) + also have "\ = (?I ?tk (Gt (CN 0 c e)))" using pos_divide_less_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + by (simp add: ti) + finally have ?case . } + ultimately show ?case by blast +next + case (8 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + {assume kdc: "k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + moreover + {assume "\ k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have "?I (real x) (?s (Ge (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \ 0)" + using real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) + also have "\ = (?I ?tk (Ge (CN 0 c e)))" using pos_divide_le_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + by (simp add: ti) + finally have ?case . } + ultimately show ?case by blast +next + case (9 i c e) from prems have cp: "c > 0" and nb: "numbound0 e" by auto + {assume kdc: "k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + moreover + {assume "\ k dvd c" + from kpos have knz: "k\0" by simp hence knz': "real k \ 0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have "?I (real x) (?s (Dvd i (CN 0 c e))) = (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k)" + using real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) + also have "\ = (?I ?tk (Dvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + by (simp add: ti) + finally have ?case . } + ultimately show ?case by blast +next + case (10 i c e) from prems have cp: "c > 0" and nb: "numbound0 e" by auto + {assume kdc: "k dvd c" + from kpos have knz: "k\0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + moreover + {assume "\ k dvd c" + from kpos have knz: "k\0" by simp hence knz': "real k \ 0" by simp + from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from prems have "?I (real x) (?s (NDvd i (CN 0 c e))) = (\ (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k))" + using real_of_int_div[OF knz kdt] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps) + also have "\ = (?I ?tk (NDvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + by (simp add: ti) + finally have ?case . } + ultimately show ?case by blast +qed (simp_all add: nth_pos2 bound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"] numbound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"]) + + +lemma a\: + assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "real k > 0" + shows "Ifm (real (x*k)#bs) (a\ p k) = Ifm (real x#bs) p" (is "?I (x*k) (?f p k) = ?I x p") +using lp bound0_I[where bs="bs" and b="real (x*k)" and b'="real x"] numbound0_I[where bs="bs" and b="real (x*k)" and b'="real x"] +proof(induct p rule: a\.induct) + case (3 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp + {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } + moreover + {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} + ultimately show ?case by blast +next + case (4 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp + {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } + moreover + {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} + ultimately show ?case by blast +next + case (5 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp + {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } + moreover + {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_less_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} + ultimately show ?case by blast +next + case (6 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp + {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } + moreover + {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_le_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} + ultimately show ?case by blast +next + case (7 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp + {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } + moreover + {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_less_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} + ultimately show ?case by blast +next + case (8 c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp + {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } + moreover + {assume nkdc: "\ k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_le_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)} + ultimately show ?case by blast +next + case (9 i c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp + {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } + moreover + {assume "\ k dvd c" + hence "Ifm (real (x*k)#bs) (a\ (Dvd i (CN 0 c e)) k) = + (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k)" + using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] + by (simp add: ring_simps) + also have "\ = (Ifm (real x#bs) (Dvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"]) + finally have ?case . } + ultimately show ?case by blast +next + case (10 i c e) + from prems have cp: "c > 0" and nb: "numbound0 e" by auto + from kp have knz: "k\0" by simp hence knz': "real k \ 0" by simp + {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } + moreover + {assume "\ k dvd c" + hence "Ifm (real (x*k)#bs) (a\ (NDvd i (CN 0 c e)) k) = + (\ (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k))" + using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] + by (simp add: ring_simps) + also have "\ = (Ifm (real x#bs) (NDvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"]) + finally have ?case . } + ultimately show ?case by blast +qed (simp_all add: nth_pos2) + +lemma a\_ex: + assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "k > 0" + shows "(\ (x::int). real k rdvd real x \ Ifm (real x#bs) (a\ p k)) = + (\ (x::int). Ifm (real x#bs) p)" (is "(\ x. ?D x \ ?P' x) = (\ x. ?P x)") +proof- + have "(\ x. ?D x \ ?P' x) = (\ x. k dvd x \ ?P' x)" using int_rdvd_iff by simp + also have "\ = (\x. ?P' (x*k))" using unity_coeff_ex[where P="?P'" and l="k", simplified] + by (simp add: ring_simps) + also have "\ = (\ x. ?P x)" using a\ iszlfm_gen[OF lp] kp by auto + finally show ?thesis . +qed + +lemma \\': assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "k > 0" and nb: "numbound0 t" + shows "Ifm (real x#bs) (\\ p (t,k)) = Ifm ((Inum (real x#bs) t)#bs) (a\ p k)" +using lp +by(induct p rule: \\.induct, simp_all add: + numbound0_I[OF nb, where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] + numbound0_I[where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] + bound0_I[where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] nth_pos2 cong: imp_cong) + +lemma \\_nb: assumes lp:"iszlfm p (a#bs)" and nb: "numbound0 t" + shows "bound0 (\\ p (t,k))" + using lp + by (induct p rule: iszlfm.induct, auto simp add: nb) + +lemma \_l: + assumes lp: "iszlfm p (real (i::int)#bs)" + shows "\ (b,k) \ set (\ p). k >0 \ numbound0 b \ isint b (real i#bs)" +using lp by (induct p rule: \.induct, auto simp add: isint_sub isint_neg) + +lemma \\_l: + assumes lp: "iszlfm p (real (i::int)#bs)" + shows "\ (b,k) \ set (\\ p). k >0 \ numbound0 b \ isint b (real i#bs)" +using lp isint_add [OF isint_c[where j="- 1"],where bs="real i#bs"] + by (induct p rule: \\.induct, auto) + +lemma zminusinf_\: + assumes lp: "iszlfm p (real (i::int)#bs)" + and nmi: "\ (Ifm (real i#bs) (minusinf p))" (is "\ (Ifm (real i#bs) (?M p))") + and ex: "Ifm (real i#bs) p" (is "?I i p") + shows "\ (e,c) \ set (\ p). real (c*i) > Inum (real i#bs) e" (is "\ (e,c) \ ?R p. real (c*i) > ?N i e") + using lp nmi ex +by (induct p rule: minusinf.induct, auto) + + +lemma \_And: "Ifm bs (\ (And p q) k t) = Ifm bs (And (\ p k t) (\ q k t))" +using \_def by auto +lemma \_Or: "Ifm bs (\ (Or p q) k t) = Ifm bs (Or (\ p k t) (\ q k t))" +using \_def by auto + +lemma \: assumes lp: "iszlfm p (real (i::int) #bs)" + and pi: "Ifm (real i#bs) p" + and d: "d\ p d" + and dp: "d > 0" + and nob: "\(e,c) \ set (\ p). \ j\ {1 .. c*d}. real (c*i) \ Inum (real i#bs) e + real j" + (is "\(e,c) \ set (\ p). \ j\ {1 .. c*d}. _ \ ?N i e + _") + shows "Ifm (real(i - d)#bs) p" + using lp pi d nob +proof(induct p rule: iszlfm.induct) + case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" + and pi: "real (c*i) = - 1 - ?N i e + real (1::int)" and nob: "\ j\ {1 .. c*d}. real (c*i) \ -1 - ?N i e + real j" + by simp+ + from mult_strict_left_mono[OF dp cp] have one:"1 \ {1 .. c*d}" by auto + from nob[rule_format, where j="1", OF one] pi show ?case by simp +next + case (4 c e) + hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" + and nob: "\ j\ {1 .. c*d}. real (c*i) \ - ?N i e + real j" + by simp+ + {assume "real (c*i) \ - ?N i e + real (c*d)" + with numbound0_I[OF nb, where bs="bs" and b="real i - real d" and b'="real i"] + have ?case by (simp add: ring_simps)} + moreover + {assume pi: "real (c*i) = - ?N i e + real (c*d)" + from mult_strict_left_mono[OF dp cp] have d: "(c*d) \ {1 .. c*d}" by simp + from nob[rule_format, where j="c*d", OF d] pi have ?case by simp } + ultimately show ?case by blast +next + case (5 c e) hence cp: "c > 0" by simp + from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] + real_of_int_mult] + show ?case using prems dp + by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] + ring_simps) +next + case (6 c e) hence cp: "c > 0" by simp + from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] + real_of_int_mult] + show ?case using prems dp + by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] + ring_simps) +next + case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" + and nob: "\ j\ {1 .. c*d}. real (c*i) \ - ?N i e + real j" + and pi: "real (c*i) + ?N i e > 0" and cp': "real c >0" + by simp+ + let ?fe = "floor (?N i e)" + from pi cp have th:"(real i +?N i e / real c)*real c > 0" by (simp add: ring_simps) + from pi ei[simplified isint_iff] have "real (c*i + ?fe) > real (0::int)" by simp + hence pi': "c*i + ?fe > 0" by (simp only: real_of_int_less_iff[symmetric]) + have "real (c*i) + ?N i e > real (c*d) \ real (c*i) + ?N i e \ real (c*d)" by auto + moreover + {assume "real (c*i) + ?N i e > real (c*d)" hence ?case + by (simp add: ring_simps + numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} + moreover + {assume H:"real (c*i) + ?N i e \ real (c*d)" + with ei[simplified isint_iff] have "real (c*i + ?fe) \ real (c*d)" by simp + hence pid: "c*i + ?fe \ c*d" by (simp only: real_of_int_le_iff) + with pi' have "\ j1\ {1 .. c*d}. c*i + ?fe = j1" by auto + hence "\ j1\ {1 .. c*d}. real (c*i) = - ?N i e + real j1" + by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] ring_simps) + with nob have ?case by blast } + ultimately show ?case by blast +next + case (8 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" + and nob: "\ j\ {1 .. c*d}. real (c*i) \ - 1 - ?N i e + real j" + and pi: "real (c*i) + ?N i e \ 0" and cp': "real c >0" + by simp+ + let ?fe = "floor (?N i e)" + from pi cp have th:"(real i +?N i e / real c)*real c \ 0" by (simp add: ring_simps) + from pi ei[simplified isint_iff] have "real (c*i + ?fe) \ real (0::int)" by simp + hence pi': "c*i + 1 + ?fe \ 1" by (simp only: real_of_int_le_iff[symmetric]) + have "real (c*i) + ?N i e \ real (c*d) \ real (c*i) + ?N i e < real (c*d)" by auto + moreover + {assume "real (c*i) + ?N i e \ real (c*d)" hence ?case + by (simp add: ring_simps + numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} + moreover + {assume H:"real (c*i) + ?N i e < real (c*d)" + with ei[simplified isint_iff] have "real (c*i + ?fe) < real (c*d)" by simp + hence pid: "c*i + 1 + ?fe \ c*d" by (simp only: real_of_int_le_iff) + with pi' have "\ j1\ {1 .. c*d}. c*i + 1+ ?fe = j1" by auto + hence "\ j1\ {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1" + by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] ring_simps real_of_one) + hence "\ j1\ {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1" + by (simp only: ring_simps diff_def[symmetric]) + hence "\ j1\ {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1" + by (simp only: add_ac diff_def) + with nob have ?case by blast } + ultimately show ?case by blast +next + case (9 j c e) hence p: "real j rdvd real (c*i) + ?N i e" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ + let ?e = "Inum (real i # bs) e" + from prems have "isint e (real i #bs)" by simp + hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real i)#bs"] numbound0_I[OF bn,where b="real i" and b'="real i" and bs="bs"] + by (simp add: isint_iff) + from prems have id: "j dvd d" by simp + from ie[symmetric] have "?p i = (real j rdvd real (c*i+ floor ?e))" by simp + also have "\ = (j dvd c*i + floor ?e)" + using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp + also have "\ = (j dvd c*i - c*d + floor ?e)" + using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp + also have "\ = (real j rdvd real (c*i - c*d + floor ?e))" + using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified] + ie by simp + also have "\ = (real j rdvd real (c*(i - d)) + ?e)" + using ie by (simp add:ring_simps) + finally show ?case + using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p + by (simp add: ring_simps) +next + case (10 j c e) hence p: "\ (real j rdvd real (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ + let ?e = "Inum (real i # bs) e" + from prems have "isint e (real i #bs)" by simp + hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real i)#bs"] numbound0_I[OF bn,where b="real i" and b'="real i" and bs="bs"] + by (simp add: isint_iff) + from prems have id: "j dvd d" by simp + from ie[symmetric] have "?p i = (\ (real j rdvd real (c*i+ floor ?e)))" by simp + also have "\ = Not (j dvd c*i + floor ?e)" + using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp + also have "\ = Not (j dvd c*i - c*d + floor ?e)" + using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp + also have "\ = Not (real j rdvd real (c*i - c*d + floor ?e))" + using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified] + ie by simp + also have "\ = Not (real j rdvd real (c*(i - d)) + ?e)" + using ie by (simp add:ring_simps) + finally show ?case + using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p + by (simp add: ring_simps) +qed(auto simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] nth_pos2) + +lemma \_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t" + shows "bound0 (\ p k t)" + using \\_nb[OF lp nb] nb by (simp add: \_def) + +lemma \': assumes lp: "iszlfm p (a #bs)" + and d: "d\ p d" + and dp: "d > 0" + shows "\ x. \(\ (e,c) \ set(\ p). \(j::int) \ {1 .. c*d}. Ifm (a #bs) (\ p c (Add e (C j)))) \ Ifm (real x#bs) p \ Ifm (real (x - d)#bs) p" (is "\ x. ?b x \ ?P x \ ?P (x - d)") +proof(clarify) + fix x + assume nob1:"?b x" and px: "?P x" + from iszlfm_gen[OF lp, rule_format, where y="real x"] have lp': "iszlfm p (real x#bs)". + have nob: "\(e, c)\set (\ p). \j\{1..c * d}. real (c * x) \ Inum (real x # bs) e + real j" + proof(clarify) + fix e c j assume ecR: "(e,c) \ set (\ p)" and jD: "j\ {1 .. c*d}" + and cx: "real (c*x) = Inum (real x#bs) e + real j" + let ?e = "Inum (real x#bs) e" + let ?fe = "floor ?e" + from \_l[OF lp'] ecR have ei:"isint e (real x#bs)" and cp:"c>0" and nb:"numbound0 e" + by auto + from numbound0_gen [OF nb ei, rule_format,where y="a"] have "isint e (a#bs)" . + from cx ei[simplified isint_iff] have "real (c*x) = real (?fe + j)" by simp + hence cx: "c*x = ?fe + j" by (simp only: real_of_int_inject) + hence cdej:"c dvd ?fe + j" by (simp add: dvd_def) (rule_tac x="x" in exI, simp) + hence "real c rdvd real (?fe + j)" by (simp only: int_rdvd_iff) + hence rcdej: "real c rdvd ?e + real j" by (simp add: ei[simplified isint_iff]) + from cx have "(c*x) div c = (?fe + j) div c" by simp + with cp have "x = (?fe + j) div c" by simp + with px have th: "?P ((?fe + j) div c)" by auto + from cp have cp': "real c > 0" by simp + from cdej have cdej': "c dvd floor (Inum (real x#bs) (Add e (C j)))" by simp + from nb have nb': "numbound0 (Add e (C j))" by simp + have ji: "isint (C j) (real x#bs)" by (simp add: isint_def) + from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real x#bs)" . + from th \\[where b'="real x", OF lp' cp' nb' ei' cdej',symmetric] + have "Ifm (real x#bs) (\\ p (Add e (C j), c))" by simp + with rcdej have th: "Ifm (real x#bs) (\ p c (Add e (C j)))" by (simp add: \_def) + from th bound0_I[OF \_nb[OF lp nb', where k="c"],where bs="bs" and b="real x" and b'="a"] + have "Ifm (a#bs) (\ p c (Add e (C j)))" by blast + with ecR jD nob1 show "False" by blast + qed + from \[OF lp' px d dp nob] show "?P (x -d )" . +qed + + +lemma rl_thm: + assumes lp: "iszlfm p (real (i::int)#bs)" + shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real j#bs) (minusinf p)) \ (\ (e,c) \ set (\ p). \ j\ {1 .. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" + (is "(\(x::int). ?P x) = ((\ j\ {1.. \ p}. ?MP j)\(\ (e,c) \ ?R. \ j\ _. ?SP c e j))" + is "?lhs = (?MD \ ?RD)" is "?lhs = ?rhs") +proof- + let ?d= "\ p" + from \[OF lp] have d:"d\ p ?d" and dp: "?d > 0" by auto + { assume H:"?MD" hence th:"\ (x::int). ?MP x" by blast + from H minusinf_ex[OF lp th] have ?thesis by blast} + moreover + { fix e c j assume exR:"(e,c) \ ?R" and jD:"j\ {1 .. c*?d}" and spx:"?SP c e j" + from exR \_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real i#bs)" and cp: "c > 0" + by auto + have "isint (C j) (real i#bs)" by (simp add: isint_iff) + with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real i"]] + have eji:"isint (Add e (C j)) (real i#bs)" by simp + from nb have nb': "numbound0 (Add e (C j))" by simp + from spx bound0_I[OF \_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real i"] + have spx': "Ifm (real i # bs) (\ p c (Add e (C j)))" by blast + from spx' have rcdej:"real c rdvd (Inum (real i#bs) (Add e (C j)))" + and sr:"Ifm (real i#bs) (\\ p (Add e (C j),c))" by (simp add: \_def)+ + from rcdej eji[simplified isint_iff] + have "real c rdvd real (floor (Inum (real i#bs) (Add e (C j))))" by simp + hence cdej:"c dvd floor (Inum (real i#bs) (Add e (C j)))" by (simp only: int_rdvd_iff) + from cp have cp': "real c > 0" by simp + from \\[OF lp cp' nb' eji cdej] spx' have "?P (\Inum (real i # bs) (Add e (C j))\ div c)" + by (simp add: \_def) + hence ?lhs by blast + with exR jD spx have ?thesis by blast} + moreover + { fix x assume px: "?P x" and nob: "\ ?RD" + from iszlfm_gen [OF lp,rule_format, where y="a"] have lp':"iszlfm p (a#bs)" . + from \'[OF lp' d dp, rule_format, OF nob] have th:"\ x. ?P x \ ?P (x - ?d)" by blast + from minusinf_inf[OF lp] obtain z where z:"\ x 0" by arith + from decr_lemma[OF dp,where x="x" and z="z"] + decr_mult_lemma[OF dp th zp, rule_format, OF px] z have th:"\ x. ?MP x" by auto + with minusinf_bex[OF lp] px nob have ?thesis by blast} + ultimately show ?thesis by blast +qed + +lemma mirror_\\: assumes lp: "iszlfm p (a#bs)" + shows "(\ (t,k). (Inum (a#bs) t, k)) ` set (\\ p) = (\ (t,k). (Inum (a#bs) t,k)) ` set (\ (mirror p))" +using lp +by (induct p rule: mirror.induct, simp_all add: split_def image_Un ) + +text {* The @{text "\"} part*} + +text{* Linearity for fm where Bound 0 ranges over @{text "\"}*} +consts + isrlfm :: "fm \ bool" (* Linearity test for fm *) +recdef isrlfm "measure size" + "isrlfm (And p q) = (isrlfm p \ isrlfm q)" + "isrlfm (Or p q) = (isrlfm p \ isrlfm q)" + "isrlfm (Eq (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm (NEq (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm (Lt (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm (Le (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm (Gt (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm (Ge (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm p = (isatom p \ (bound0 p))" + +constdefs fp :: "fm \ int \ num \ int \ fm" + "fp p n s j \ (if n > 0 then + (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j))))) + (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1)))))))) + else + (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C j))))) + (Gt (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C (j + 1)))))))))" + + (* splits the bounded from the unbounded part*) +consts rsplit0 :: "num \ (fm \ int \ num) list" +recdef rsplit0 "measure num_size" + "rsplit0 (Bound 0) = [(T,1,C 0)]" + "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b + in map (\ ((p,n,t),(q,m,s)). (And p q, n+m, Add t s)) [(a,b). a\acs,b\bcs])" + "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))" + "rsplit0 (Neg a) = map (\ (p,n,s). (p,-n,Neg s)) (rsplit0 a)" + "rsplit0 (Floor a) = foldl (op @) [] (map + (\ (p,n,s). if n=0 then [(p,0,Floor s)] + else (map (\ j. (fp p n s j, 0, Add (Floor s) (C j))) (if n > 0 then iupt (0,n) else iupt(n,0)))) + (rsplit0 a))" + "rsplit0 (CN 0 c a) = map (\ (p,n,s). (p,n+c,s)) (rsplit0 a)" + "rsplit0 (CN m c a) = map (\ (p,n,s). (p,n,CN m c s)) (rsplit0 a)" + "rsplit0 (CF c t s) = rsplit0 (Add (Mul c (Floor t)) s)" + "rsplit0 (Mul c a) = map (\ (p,n,s). (p,c*n,Mul c s)) (rsplit0 a)" + "rsplit0 t = [(T,0,t)]" + +lemma not_rl[simp]: "isrlfm p \ isrlfm (not p)" + by (induct p rule: isrlfm.induct, auto) +lemma conj_rl[simp]: "isrlfm p \ isrlfm q \ isrlfm (conj p q)" + using conj_def by (cases p, auto) +lemma disj_rl[simp]: "isrlfm p \ isrlfm q \ isrlfm (disj p q)" + using disj_def by (cases p, auto) + + +lemma rsplit0_cs: + shows "\ (p,n,s) \ set (rsplit0 t). + (Ifm (x#bs) p \ (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \ numbound0 s \ isrlfm p" + (is "\ (p,n,s) \ ?SS t. (?I p \ ?N t = ?N (CN 0 n s)) \ _ \ _ ") +proof(induct t rule: rsplit0.induct) + case (5 a) + let ?p = "\ (p,n,s) j. fp p n s j" + let ?f = "(\ (p,n,s) j. (?p (p,n,s) j, (0::int),Add (Floor s) (C j)))" + let ?J = "\ n. if n>0 then iupt (0,n) else iupt (n,0)" + let ?ff=" (\ (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))" + have int_cases: "\ (i::int). i= 0 \ i < 0 \ i > 0" by arith + have U1: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) = + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)]))" by auto + have U2': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n>0}. + ?ff (p,n,s) = map (?f(p,n,s)) (iupt(0,n))" by auto + hence U2: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) = + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). + set (map (?f(p,n,s)) (iupt(0,n)))))" + proof- + fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g + assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" + thus "(UNION M (\ (a,b,c). set (f (a,b,c)))) = (UNION M (\ (a,b,c). set (g a b c)))" + by (auto simp add: split_def) + qed + have U3': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n<0}. ?ff (p,n,s) = map (?f(p,n,s)) (iupt(n,0))" + by auto + hence U3: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))) = + (UNION {(p,n,s). (p,n,s)\ ?SS a\n<0} (\(p,n,s). set (map (?f(p,n,s)) (iupt(n,0)))))" + proof- + fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g + assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" + thus "(UNION M (\ (a,b,c). set (f (a,b,c)))) = (UNION M (\ (a,b,c). set (g a b c)))" + by (auto simp add: split_def) + qed + have "?SS (Floor a) = UNION (?SS a) (\x. set (?ff x))" + by (auto simp add: foldl_conv_concat) + also have "\ = UNION (?SS a) (\ (p,n,s). set (?ff (p,n,s)))" by auto + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))))" + using int_cases[rule_format] by blast + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)])) Un + (UNION {(p,n,s). (p,n,s)\ ?SS a\n>0} (\(p,n,s). set(map(?f(p,n,s)) (iupt(0,n))))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). + set (map (?f(p,n,s)) (iupt(n,0))))))" by (simp only: U1 U2 U3) + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). (?f(p,n,s)) ` {n .. 0})))" + by (simp only: set_map iupt_set set.simps) + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast + finally + have FS: "?SS (Floor a) = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast + show ?case + proof(simp only: FS, clarsimp simp del: Ifm.simps Inum.simps, -) + fix p n s + let ?ths = "(?I p \ (?N (Floor a) = ?N (CN 0 n s))) \ numbound0 s \ isrlfm p" + assume "(\ba. (p, 0, ba) \ set (rsplit0 a) \ n = 0 \ s = Floor ba) \ + (\ab ac ba. + (ab, ac, ba) \ set (rsplit0 a) \ + 0 < ac \ + (\j. p = fp ab ac ba j \ + n = 0 \ s = Add (Floor ba) (C j) \ 0 \ j \ j \ ac)) \ + (\ab ac ba. + (ab, ac, ba) \ set (rsplit0 a) \ + ac < 0 \ + (\j. p = fp ab ac ba j \ + n = 0 \ s = Add (Floor ba) (C j) \ ac \ j \ j \ 0))" + moreover + {fix s' + assume "(p, 0, s') \ ?SS a" and "n = 0" and "s = Floor s'" + hence ?ths using prems by auto} + moreover + { fix p' n' s' j + assume pns: "(p', n', s') \ ?SS a" + and np: "0 < n'" + and p_def: "p = ?p (p',n',s') j" + and n0: "n = 0" + and s_def: "s = (Add (Floor s') (C j))" + and jp: "0 \ j" and jn: "j \ n'" + from prems pns have H:"(Ifm ((x\real) # (bs\real list)) p' \ + Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \ + numbound0 s' \ isrlfm p'" by blast + hence nb: "numbound0 s'" by simp + from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numsub_nb) + let ?nxs = "CN 0 n' s'" + let ?l = "floor (?N s') + j" + from H + have "?I (?p (p',n',s') j) \ + (((?N ?nxs \ real ?l) \ (?N ?nxs < real (?l + 1))) \ (?N a = ?N ?nxs ))" + by (simp add: fp_def np ring_simps numsub numadd numfloor) + also have "\ \ ((floor (?N ?nxs) = ?l) \ (?N a = ?N ?nxs ))" + using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp + moreover + have "\ \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp + ultimately have "?I (?p (p',n',s') j) \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" + by blast + with s_def n0 p_def nb nf have ?ths by auto} + moreover + {fix p' n' s' j + assume pns: "(p', n', s') \ ?SS a" + and np: "n' < 0" + and p_def: "p = ?p (p',n',s') j" + and n0: "n = 0" + and s_def: "s = (Add (Floor s') (C j))" + and jp: "n' \ j" and jn: "j \ 0" + from prems pns have H:"(Ifm ((x\real) # (bs\real list)) p' \ + Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \ + numbound0 s' \ isrlfm p'" by blast + hence nb: "numbound0 s'" by simp + from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numneg_nb) + let ?nxs = "CN 0 n' s'" + let ?l = "floor (?N s') + j" + from H + have "?I (?p (p',n',s') j) \ + (((?N ?nxs \ real ?l) \ (?N ?nxs < real (?l + 1))) \ (?N a = ?N ?nxs ))" + by (simp add: np fp_def ring_simps numneg numfloor numadd numsub) + also have "\ \ ((floor (?N ?nxs) = ?l) \ (?N a = ?N ?nxs ))" + using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp + moreover + have "\ \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp + ultimately have "?I (?p (p',n',s') j) \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" + by blast + with s_def n0 p_def nb nf have ?ths by auto} + ultimately show ?ths by auto + qed +next + case (3 a b) then show ?case + apply auto + apply (erule_tac x = "(aa, aaa, ba)" in ballE) apply simp_all + apply (erule_tac x = "(ab, ac, baa)" in ballE) apply simp_all + done +qed (auto simp add: Let_def split_def ring_simps conj_rl) + +lemma real_in_int_intervals: + assumes xb: "real m \ x \ x < real ((n::int) + 1)" + shows "\ j\ {m.. n}. real j \ x \ x < real (j+1)" (is "\ j\ ?N. ?P j") +by (rule bexI[where P="?P" and x="floor x" and A="?N"]) +(auto simp add: floor_less_eq[where x="x" and a="n+1", simplified] xb[simplified] floor_mono2[where x="real m" and y="x", OF conjunct1[OF xb], simplified floor_real_of_int[where n="m"]]) + +lemma rsplit0_complete: + assumes xp:"0 \ x" and x1:"x < 1" + shows "\ (p,n,s) \ set (rsplit0 t). Ifm (x#bs) p" (is "\ (p,n,s) \ ?SS t. ?I p") +proof(induct t rule: rsplit0.induct) + case (2 a b) + from prems have "\ (pa,na,sa) \ ?SS a. ?I pa" by auto + then obtain "pa" "na" "sa" where pa: "(pa,na,sa)\ ?SS a \ ?I pa" by blast + from prems have "\ (pb,nb,sb) \ ?SS b. ?I pb" by auto + then obtain "pb" "nb" "sb" where pb: "(pb,nb,sb)\ ?SS b \ ?I pb" by blast + from pa pb have th: "((pa,na,sa),(pb,nb,sb)) \ set[(x,y). x\rsplit0 a, y\rsplit0 b]" + by (auto) + let ?f="(\ ((p,n,t),(q,m,s)). (And p q, n+m, Add t s))" + from imageI[OF th, where f="?f"] have "?f ((pa,na,sa),(pb,nb,sb)) \ ?SS (Add a b)" + by (simp add: Let_def) + hence "(And pa pb, na +nb, Add sa sb) \ ?SS (Add a b)" by simp + moreover from pa pb have "?I (And pa pb)" by simp + ultimately show ?case by blast +next + case (5 a) + let ?p = "\ (p,n,s) j. fp p n s j" + let ?f = "(\ (p,n,s) j. (?p (p,n,s) j, (0::int),(Add (Floor s) (C j))))" + let ?J = "\ n. if n>0 then iupt (0,n) else iupt (n,0)" + let ?ff=" (\ (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))" + have int_cases: "\ (i::int). i= 0 \ i < 0 \ i > 0" by arith + have U1: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)]))" by auto + have U2': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n>0}. ?ff (p,n,s) = map (?f(p,n,s)) (iupt(0,n))" + by auto + hence U2: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (map (?f(p,n,s)) (iupt(0,n)))))" + proof- + fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g + assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" + thus "(UNION M (\ (a,b,c). set (f (a,b,c)))) = (UNION M (\ (a,b,c). set (g a b c)))" + by (auto simp add: split_def) + qed + have U3': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n<0}. ?ff (p,n,s) = map (?f(p,n,s)) (iupt(n,0))" + by auto + hence U3: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (map (?f(p,n,s)) (iupt(n,0)))))" + proof- + fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g + assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" + thus "(UNION M (\ (a,b,c). set (f (a,b,c)))) = (UNION M (\ (a,b,c). set (g a b c)))" + by (auto simp add: split_def) + qed + + have "?SS (Floor a) = UNION (?SS a) (\x. set (?ff x))" by (auto simp add: foldl_conv_concat) + also have "\ = UNION (?SS a) (\ (p,n,s). set (?ff (p,n,s)))" by auto + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))))" + using int_cases[rule_format] by blast + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)])) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (map (?f(p,n,s)) (iupt(0,n))))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (map (?f(p,n,s)) (iupt(n,0))))))" by (simp only: U1 U2 U3) + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). (?f(p,n,s)) ` {n .. 0})))" + by (simp only: set_map iupt_set set.simps) + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast + finally + have FS: "?SS (Floor a) = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast + from prems have "\ (p,n,s) \ ?SS a. ?I p" by auto + then obtain "p" "n" "s" where pns: "(p,n,s) \ ?SS a \ ?I p" by blast + let ?N = "\ t. Inum (x#bs) t" + from rsplit0_cs[rule_format] pns have ans:"(?N a = ?N (CN 0 n s)) \ numbound0 s \ isrlfm p" + by auto + + have "n=0 \ n >0 \ n <0" by arith + moreover {assume "n=0" hence ?case using pns by (simp only: FS) auto } + moreover + { + assume np: "n > 0" + from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) \ ?N s" by simp + also from mult_left_mono[OF xp] np have "?N s \ real n * x + ?N s" by simp + finally have "?N (Floor s) \ real n * x + ?N s" . + moreover + {from mult_strict_left_mono[OF x1] np + have "real n *x + ?N s < real n + ?N s" by simp + also from real_of_int_floor_add_one_gt[where r="?N s"] + have "\ < real n + ?N (Floor s) + 1" by simp + finally have "real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp} + ultimately have "?N (Floor s) \ real n *x + ?N s\ real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp + hence th: "0 \ real n *x + ?N s - ?N (Floor s) \ real n *x + ?N s - ?N (Floor s) < real (n+1)" by simp + from real_in_int_intervals th have "\ j\ {0 .. n}. real j \ real n *x + ?N s - ?N (Floor s)\ real n *x + ?N s - ?N (Floor s) < real (j+1)" by simp + + hence "\ j\ {0 .. n}. 0 \ real n *x + ?N s - ?N (Floor s) - real j \ real n *x + ?N s - ?N (Floor s) - real (j+1) < 0" + by(simp only: myl[rule_format, where b="real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) + hence "\ j\ {0.. n}. ?I (?p (p,n,s) j)" + using pns by (simp add: fp_def np ring_simps numsub numadd) + then obtain "j" where j_def: "j\ {0 .. n} \ ?I (?p (p,n,s) j)" by blast + hence "\x \ {?p (p,n,s) j |j. 0\ j \ j \ n }. ?I x" by auto + hence ?case using pns + by (simp only: FS,simp add: bex_Un) + (rule disjI2, rule disjI1,rule exI [where x="p"], + rule exI [where x="n"],rule exI [where x="s"],simp_all add: np) + } + moreover + { assume nn: "n < 0" hence np: "-n >0" by simp + from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) + 1 > ?N s" by simp + moreover from mult_left_mono_neg[OF xp] nn have "?N s \ real n * x + ?N s" by simp + ultimately have "?N (Floor s) + 1 > real n * x + ?N s" by arith + moreover + {from mult_strict_left_mono_neg[OF x1, where c="real n"] nn + have "real n *x + ?N s \ real n + ?N s" by simp + moreover from real_of_int_floor_le[where r="?N s"] have "real n + ?N s \ real n + ?N (Floor s)" by simp + ultimately have "real n *x + ?N s \ ?N (Floor s) + real n" + by (simp only: ring_simps)} + ultimately have "?N (Floor s) + real n \ real n *x + ?N s\ real n *x + ?N s < ?N (Floor s) + real (1::int)" by simp + hence th: "real n \ real n *x + ?N s - ?N (Floor s) \ real n *x + ?N s - ?N (Floor s) < real (1::int)" by simp + have th1: "\ (a::real). (- a > 0) = (a < 0)" by auto + have th2: "\ (a::real). (0 \ - a) = (a \ 0)" by auto + from real_in_int_intervals th have "\ j\ {n .. 0}. real j \ real n *x + ?N s - ?N (Floor s)\ real n *x + ?N s - ?N (Floor s) < real (j+1)" by simp + + hence "\ j\ {n .. 0}. 0 \ real n *x + ?N s - ?N (Floor s) - real j \ real n *x + ?N s - ?N (Floor s) - real (j+1) < 0" + by(simp only: myl[rule_format, where b="real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) + hence "\ j\ {n .. 0}. 0 \ - (real n *x + ?N s - ?N (Floor s) - real j) \ - (real n *x + ?N s - ?N (Floor s) - real (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format]) + hence "\ j\ {n.. 0}. ?I (?p (p,n,s) j)" + using pns by (simp add: fp_def nn diff_def add_ac mult_ac numfloor numadd numneg + del: diff_less_0_iff_less diff_le_0_iff_le) + then obtain "j" where j_def: "j\ {n .. 0} \ ?I (?p (p,n,s) j)" by blast + hence "\x \ {?p (p,n,s) j |j. n\ j \ j \ 0 }. ?I x" by auto + hence ?case using pns + by (simp only: FS,simp add: bex_Un) + (rule disjI2, rule disjI2,rule exI [where x="p"], + rule exI [where x="n"],rule exI [where x="s"],simp_all add: nn) + } + ultimately show ?case by blast +qed (auto simp add: Let_def split_def) + + (* Linearize a formula where Bound 0 ranges over [0,1) *) + +constdefs rsplit :: "(int \ num \ fm) \ num \ fm" + "rsplit f a \ foldr disj (map (\ (\, n, s). conj \ (f n s)) (rsplit0 a)) F" + +lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\ x \ set xs. Ifm bs (f x))" +by(induct xs, simp_all) + +lemma foldr_conj_map: "Ifm bs (foldr conj (map f xs) T) = (\ x \ set xs. Ifm bs (f x))" +by(induct xs, simp_all) + +lemma foldr_disj_map_rlfm: + assumes lf: "\ n s. numbound0 s \ isrlfm (f n s)" + and \: "\ (\,n,s) \ set xs. numbound0 s \ isrlfm \" + shows "isrlfm (foldr disj (map (\ (\, n, s). conj \ (f n s)) xs) F)" +using lf \ by (induct xs, auto) + +lemma rsplit_ex: "Ifm bs (rsplit f a) = (\ (\,n,s) \ set (rsplit0 a). Ifm bs (conj \ (f n s)))" +using foldr_disj_map[where xs="rsplit0 a"] rsplit_def by (simp add: split_def) + +lemma rsplit_l: assumes lf: "\ n s. numbound0 s \ isrlfm (f n s)" + shows "isrlfm (rsplit f a)" +proof- + from rsplit0_cs[where t="a"] have th: "\ (\,n,s) \ set (rsplit0 a). numbound0 s \ isrlfm \" by blast + from foldr_disj_map_rlfm[OF lf th] rsplit_def show ?thesis by simp +qed + +lemma rsplit: + assumes xp: "x \ 0" and x1: "x < 1" + and f: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ (Ifm (x#bs) (f n s) = Ifm (x#bs) (g a))" + shows "Ifm (x#bs) (rsplit f a) = Ifm (x#bs) (g a)" +proof(auto) + let ?I = "\x p. Ifm (x#bs) p" + let ?N = "\ x t. Inum (x#bs) t" + assume "?I x (rsplit f a)" + hence "\ (\,n,s) \ set (rsplit0 a). ?I x (And \ (f n s))" using rsplit_ex by simp + then obtain "\" "n" "s" where fnsS:"(\,n,s) \ set (rsplit0 a)" and "?I x (And \ (f n s))" by blast + hence \: "?I x \" and fns: "?I x (f n s)" by auto + from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \ + have th: "(?N x a = ?N x (CN 0 n s)) \ numbound0 s" by auto + from f[rule_format, OF th] fns show "?I x (g a)" by simp +next + let ?I = "\x p. Ifm (x#bs) p" + let ?N = "\ x t. Inum (x#bs) t" + assume ga: "?I x (g a)" + from rsplit0_complete[OF xp x1, where bs="bs" and t="a"] + obtain "\" "n" "s" where fnsS:"(\,n,s) \ set (rsplit0 a)" and fx: "?I x \" by blast + from rsplit0_cs[where t="a" and x="x" and bs="bs"] fnsS fx + have ans: "?N x a = ?N x (CN 0 n s)" and nb: "numbound0 s" by auto + with ga f have "?I x (f n s)" by auto + with rsplit_ex fnsS fx show "?I x (rsplit f a)" by auto +qed + +definition lt :: "int \ num \ fm" where + lt_def: "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) + else (Gt (CN 0 (-c) (Neg t))))" + +definition le :: "int \ num \ fm" where + le_def: "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) + else (Ge (CN 0 (-c) (Neg t))))" + +definition gt :: "int \ num \ fm" where + gt_def: "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) + else (Lt (CN 0 (-c) (Neg t))))" + +definition ge :: "int \ num \ fm" where + ge_def: "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) + else (Le (CN 0 (-c) (Neg t))))" + +definition eq :: "int \ num \ fm" where + eq_def: "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) + else (Eq (CN 0 (-c) (Neg t))))" + +definition neq :: "int \ num \ fm" where + neq_def: "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) + else (NEq (CN 0 (-c) (Neg t))))" + +lemma lt_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (lt n s) = Ifm (x#bs) (Lt a)" + (is "\ a n s . ?N a = ?N (CN 0 n s) \ _\ ?I (lt n s) = ?I (Lt a)") +proof(clarify) + fix a n s + assume H: "?N a = ?N (CN 0 n s)" + show "?I (lt n s) = ?I (Lt a)" using H by (cases "n=0", (simp add: lt_def)) + (cases "n > 0", simp_all add: lt_def ring_simps myless[rule_format, where b="0"]) +qed + +lemma lt_l: "isrlfm (rsplit lt a)" + by (rule rsplit_l[where f="lt" and a="a"], auto simp add: lt_def, + case_tac s, simp_all, case_tac "nat", simp_all) + +lemma le_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (le n s) = Ifm (x#bs) (Le a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (le n s) = ?I (Le a)") +proof(clarify) + fix a n s + assume H: "?N a = ?N (CN 0 n s)" + show "?I (le n s) = ?I (Le a)" using H by (cases "n=0", (simp add: le_def)) + (cases "n > 0", simp_all add: le_def ring_simps myl[rule_format, where b="0"]) +qed + +lemma le_l: "isrlfm (rsplit le a)" + by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def) +(case_tac s, simp_all, case_tac "nat",simp_all) + +lemma gt_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (gt n s) = Ifm (x#bs) (Gt a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (gt n s) = ?I (Gt a)") +proof(clarify) + fix a n s + assume H: "?N a = ?N (CN 0 n s)" + show "?I (gt n s) = ?I (Gt a)" using H by (cases "n=0", (simp add: gt_def)) + (cases "n > 0", simp_all add: gt_def ring_simps myless[rule_format, where b="0"]) +qed +lemma gt_l: "isrlfm (rsplit gt a)" + by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def) +(case_tac s, simp_all, case_tac "nat", simp_all) + +lemma ge_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (ge n s) = Ifm (x#bs) (Ge a)" (is "\ a n s . ?N a = ?N (CN 0 n s) \ _ \ ?I (ge n s) = ?I (Ge a)") +proof(clarify) + fix a n s + assume H: "?N a = ?N (CN 0 n s)" + show "?I (ge n s) = ?I (Ge a)" using H by (cases "n=0", (simp add: ge_def)) + (cases "n > 0", simp_all add: ge_def ring_simps myl[rule_format, where b="0"]) +qed +lemma ge_l: "isrlfm (rsplit ge a)" + by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def) +(case_tac s, simp_all, case_tac "nat", simp_all) + +lemma eq_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (eq n s) = Ifm (x#bs) (Eq a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (eq n s) = ?I (Eq a)") +proof(clarify) + fix a n s + assume H: "?N a = ?N (CN 0 n s)" + show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def ring_simps) +qed +lemma eq_l: "isrlfm (rsplit eq a)" + by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def) +(case_tac s, simp_all, case_tac"nat", simp_all) + +lemma neq_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (neq n s) = Ifm (x#bs) (NEq a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (neq n s) = ?I (NEq a)") +proof(clarify) + fix a n s bs + assume H: "?N a = ?N (CN 0 n s)" + show "?I (neq n s) = ?I (NEq a)" using H by (auto simp add: neq_def ring_simps) +qed + +lemma neq_l: "isrlfm (rsplit neq a)" + by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def) +(case_tac s, simp_all, case_tac"nat", simp_all) + +lemma small_le: + assumes u0:"0 \ u" and u1: "u < 1" + shows "(-u \ real (n::int)) = (0 \ n)" +using u0 u1 by auto + +lemma small_lt: + assumes u0:"0 \ u" and u1: "u < 1" + shows "(real (n::int) < real (m::int) - u) = (n < m)" +using u0 u1 by auto + +lemma rdvd01_cs: + assumes up: "u \ 0" and u1: "u<1" and np: "real n > 0" + shows "(real (i::int) rdvd real (n::int) * u - s) = (\ j\ {0 .. n - 1}. real n * u = s - real (floor s) + real j \ real i rdvd real (j - floor s))" (is "?lhs = ?rhs") +proof- + let ?ss = "s - real (floor s)" + from real_of_int_floor_add_one_gt[where r="s", simplified myless[rule_format,where a="s"]] + real_of_int_floor_le[where r="s"] have ss0:"?ss \ 0" and ss1:"?ss < 1" + by (auto simp add: myl[rule_format, where b="s", symmetric] myless[rule_format, where a="?ss"]) + from np have n0: "real n \ 0" by simp + from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] + have nu0:"real n * u - s \ -s" and nun:"real n * u -s < real n - s" by auto + from int_rdvd_real[where i="i" and x="real (n::int) * u - s"] + have "real i rdvd real n * u - s = + (i dvd floor (real n * u -s) \ (real (floor (real n * u - s)) = real n * u - s ))" + (is "_ = (?DE)" is "_ = (?D \ ?E)") by simp + also have "\ = (?DE \ real(floor (real n * u - s) + floor s)\ -?ss + \ real(floor (real n * u - s) + floor s)< real n - ?ss)" (is "_=(?DE \real ?a \ _ \ real ?a < _)") + using nu0 nun by auto + also have "\ = (?DE \ ?a \ 0 \ ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1]) + also have "\ = (?DE \ (\ j\ {0 .. (n - 1)}. ?a = j))" by simp + also have "\ = (?DE \ (\ j\ {0 .. (n - 1)}. real (\real n * u - s\) = real j - real \s\ ))" + by (simp only: ring_simps real_of_int_diff[symmetric] real_of_int_inject del: real_of_int_diff) + also have "\ = ((\ j\ {0 .. (n - 1)}. real n * u - s = real j - real \s\ \ real i rdvd real n * u - s))" using int_rdvd_iff[where i="i" and t="\real n * u - s\"] + by (auto cong: conj_cong) + also have "\ = ?rhs" by(simp cong: conj_cong) (simp add: ring_simps ) + finally show ?thesis . +qed + +definition + DVDJ:: "int \ int \ num \ fm" +where + DVDJ_def: "DVDJ i n s = (foldr disj (map (\ j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))) (iupt(0,n - 1))) F)" + +definition + NDVDJ:: "int \ int \ num \ fm" +where + NDVDJ_def: "NDVDJ i n s = (foldr conj (map (\ j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor (Neg s))))) (iupt(0,n - 1))) T)" + +lemma DVDJ_DVD: + assumes xp:"x\ 0" and x1: "x < 1" and np:"real n > 0" + shows "Ifm (x#bs) (DVDJ i n s) = Ifm (x#bs) (Dvd i (CN 0 n s))" +proof- + let ?f = "\ j. conj (Eq(CN 0 n (Add s (Sub(Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))" + let ?s= "Inum (x#bs) s" + from foldr_disj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"] + have "Ifm (x#bs) (DVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" + by (simp add: iupt_set np DVDJ_def del: iupt.simps) + also have "\ = (\ j\ {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \ real i rdvd real (j - floor (- ?s)))" by (simp add: ring_simps diff_def[symmetric]) + also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] + have "\ = (real i rdvd real n * x - (-?s))" by simp + finally show ?thesis by simp +qed + +lemma NDVDJ_NDVD: + assumes xp:"x\ 0" and x1: "x < 1" and np:"real n > 0" + shows "Ifm (x#bs) (NDVDJ i n s) = Ifm (x#bs) (NDvd i (CN 0 n s))" +proof- + let ?f = "\ j. disj(NEq(CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor(Neg s))))" + let ?s= "Inum (x#bs) s" + from foldr_conj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"] + have "Ifm (x#bs) (NDVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" + by (simp add: iupt_set np NDVDJ_def del: iupt.simps) + also have "\ = (\ (\ j\ {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \ real i rdvd real (j - floor (- ?s))))" by (simp add: ring_simps diff_def[symmetric]) + also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] + have "\ = (\ (real i rdvd real n * x - (-?s)))" by simp + finally show ?thesis by simp +qed + +lemma foldr_disj_map_rlfm2: + assumes lf: "\ n . isrlfm (f n)" + shows "isrlfm (foldr disj (map f xs) F)" +using lf by (induct xs, auto) +lemma foldr_And_map_rlfm2: + assumes lf: "\ n . isrlfm (f n)" + shows "isrlfm (foldr conj (map f xs) T)" +using lf by (induct xs, auto) + +lemma DVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s" + shows "isrlfm (DVDJ i n s)" +proof- + let ?f="\j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) + (Dvd i (Sub (C j) (Floor (Neg s))))" + have th: "\ j. isrlfm (?f j)" using nb np by auto + from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis by simp +qed + +lemma NDVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s" + shows "isrlfm (NDVDJ i n s)" +proof- + let ?f="\j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) + (NDvd i (Sub (C j) (Floor (Neg s))))" + have th: "\ j. isrlfm (?f j)" using nb np by auto + from NDVDJ_def foldr_And_map_rlfm2[OF th] show ?thesis by auto +qed + +definition DVD :: "int \ int \ num \ fm" where + DVD_def: "DVD i c t = + (if i=0 then eq c t else + if c = 0 then (Dvd i t) else if c >0 then DVDJ (abs i) c t else DVDJ (abs i) (-c) (Neg t))" + +definition NDVD :: "int \ int \ num \ fm" where + "NDVD i c t = + (if i=0 then neq c t else + if c = 0 then (NDvd i t) else if c >0 then NDVDJ (abs i) c t else NDVDJ (abs i) (-c) (Neg t))" + +lemma DVD_mono: + assumes xp: "0\ x" and x1: "x < 1" + shows "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (DVD i n s) = Ifm (x#bs) (Dvd i a)" + (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (DVD i n s) = ?I (Dvd i a)") +proof(clarify) + fix a n s + assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s" + let ?th = "?I (DVD i n s) = ?I (Dvd i a)" + have "i=0 \ (i\0 \ n=0) \ (i\0 \ n < 0) \ (i\0 \ n > 0)" by arith + moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]] + by (simp add: DVD_def rdvd_left_0_eq)} + moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H DVD_def) } + moreover {assume inz: "i\0" and "n<0" hence ?th + by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 + rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } + moreover {assume inz: "i\0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)} + ultimately show ?th by blast +qed + +lemma NDVD_mono: assumes xp: "0\ x" and x1: "x < 1" + shows "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (NDVD i n s) = Ifm (x#bs) (NDvd i a)" + (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (NDVD i n s) = ?I (NDvd i a)") +proof(clarify) + fix a n s + assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s" + let ?th = "?I (NDVD i n s) = ?I (NDvd i a)" + have "i=0 \ (i\0 \ n=0) \ (i\0 \ n < 0) \ (i\0 \ n > 0)" by arith + moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]] + by (simp add: NDVD_def rdvd_left_0_eq)} + moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H NDVD_def) } + moreover {assume inz: "i\0" and "n<0" hence ?th + by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 + rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } + moreover {assume inz: "i\0" and "n>0" hence ?th + by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)} + ultimately show ?th by blast +qed + +lemma DVD_l: "isrlfm (rsplit (DVD i) a)" + by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l) +(case_tac s, simp_all, case_tac "nat", simp_all) + +lemma NDVD_l: "isrlfm (rsplit (NDVD i) a)" + by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l) +(case_tac s, simp_all, case_tac "nat", simp_all) + +consts rlfm :: "fm \ fm" +recdef rlfm "measure fmsize" + "rlfm (And p q) = conj (rlfm p) (rlfm q)" + "rlfm (Or p q) = disj (rlfm p) (rlfm q)" + "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)" + "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))" + "rlfm (Lt a) = rsplit lt a" + "rlfm (Le a) = rsplit le a" + "rlfm (Gt a) = rsplit gt a" + "rlfm (Ge a) = rsplit ge a" + "rlfm (Eq a) = rsplit eq a" + "rlfm (NEq a) = rsplit neq a" + "rlfm (Dvd i a) = rsplit (\ t. DVD i t) a" + "rlfm (NDvd i a) = rsplit (\ t. NDVD i t) a" + "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))" + "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))" + "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))" + "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))" + "rlfm (NOT (NOT p)) = rlfm p" + "rlfm (NOT T) = F" + "rlfm (NOT F) = T" + "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))" + "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))" + "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))" + "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))" + "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))" + "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))" + "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))" + "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))" + "rlfm p = p" (hints simp add: fmsize_pos) + +lemma bound0at_l : "\isatom p ; bound0 p\ \ isrlfm p" + by (induct p rule: isrlfm.induct, auto) +lemma zgcd_le1: assumes ip: "0 < i" shows "zgcd i j \ i" +proof- + from zgcd_zdvd1 have th: "zgcd i j dvd i" by blast + from zdvd_imp_le[OF th ip] show ?thesis . +qed + + +lemma simpfm_rl: "isrlfm p \ isrlfm (simpfm p)" +proof (induct p) + case (Lt a) + hence "bound0 (Lt a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" + by (cases a,simp_all, case_tac "nat", simp_all) + moreover + {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt a))" + using simpfm_bound0 by blast + have "isatom (simpfm (Lt a))" by (cases "simpnum a", auto simp add: Let_def) + with bn bound0at_l have ?case by blast} + moreover + {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" + { + assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" + with numgcd_pos[where t="CN 0 c (simpnum e)"] + have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp + from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" + by (simp add: numgcd_def zgcd_le1) + from prems have th': "c\0" by auto + from prems have cp: "c \ 0" by simp + from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] + have "0 < c div numgcd (CN 0 c (simpnum e))" by simp + } + with prems have ?case + by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} + ultimately show ?case by blast +next + case (Le a) + hence "bound0 (Le a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" + by (cases a,simp_all, case_tac "nat", simp_all) + moreover + {assume "bound0 (Le a)" hence bn:"bound0 (simpfm (Le a))" + using simpfm_bound0 by blast + have "isatom (simpfm (Le a))" by (cases "simpnum a", auto simp add: Let_def) + with bn bound0at_l have ?case by blast} + moreover + {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" + { + assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" + with numgcd_pos[where t="CN 0 c (simpnum e)"] + have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp + from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" + by (simp add: numgcd_def zgcd_le1) + from prems have th': "c\0" by auto + from prems have cp: "c \ 0" by simp + from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] + have "0 < c div numgcd (CN 0 c (simpnum e))" by simp + } + with prems have ?case + by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} + ultimately show ?case by blast +next + case (Gt a) + hence "bound0 (Gt a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" + by (cases a,simp_all, case_tac "nat", simp_all) + moreover + {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt a))" + using simpfm_bound0 by blast + have "isatom (simpfm (Gt a))" by (cases "simpnum a", auto simp add: Let_def) + with bn bound0at_l have ?case by blast} + moreover + {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" + { + assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" + with numgcd_pos[where t="CN 0 c (simpnum e)"] + have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp + from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" + by (simp add: numgcd_def zgcd_le1) + from prems have th': "c\0" by auto + from prems have cp: "c \ 0" by simp + from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] + have "0 < c div numgcd (CN 0 c (simpnum e))" by simp + } + with prems have ?case + by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} + ultimately show ?case by blast +next + case (Ge a) + hence "bound0 (Ge a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" + by (cases a,simp_all, case_tac "nat", simp_all) + moreover + {assume "bound0 (Ge a)" hence bn:"bound0 (simpfm (Ge a))" + using simpfm_bound0 by blast + have "isatom (simpfm (Ge a))" by (cases "simpnum a", auto simp add: Let_def) + with bn bound0at_l have ?case by blast} + moreover + {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" + { + assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" + with numgcd_pos[where t="CN 0 c (simpnum e)"] + have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp + from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" + by (simp add: numgcd_def zgcd_le1) + from prems have th': "c\0" by auto + from prems have cp: "c \ 0" by simp + from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] + have "0 < c div numgcd (CN 0 c (simpnum e))" by simp + } + with prems have ?case + by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} + ultimately show ?case by blast +next + case (Eq a) + hence "bound0 (Eq a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" + by (cases a,simp_all, case_tac "nat", simp_all) + moreover + {assume "bound0 (Eq a)" hence bn:"bound0 (simpfm (Eq a))" + using simpfm_bound0 by blast + have "isatom (simpfm (Eq a))" by (cases "simpnum a", auto simp add: Let_def) + with bn bound0at_l have ?case by blast} + moreover + {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" + { + assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" + with numgcd_pos[where t="CN 0 c (simpnum e)"] + have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp + from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" + by (simp add: numgcd_def zgcd_le1) + from prems have th': "c\0" by auto + from prems have cp: "c \ 0" by simp + from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] + have "0 < c div numgcd (CN 0 c (simpnum e))" by simp + } + with prems have ?case + by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} + ultimately show ?case by blast +next + case (NEq a) + hence "bound0 (NEq a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" + by (cases a,simp_all, case_tac "nat", simp_all) + moreover + {assume "bound0 (NEq a)" hence bn:"bound0 (simpfm (NEq a))" + using simpfm_bound0 by blast + have "isatom (simpfm (NEq a))" by (cases "simpnum a", auto simp add: Let_def) + with bn bound0at_l have ?case by blast} + moreover + {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" + { + assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" + with numgcd_pos[where t="CN 0 c (simpnum e)"] + have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp + from prems have th:"numgcd (CN 0 c (simpnum e)) \ c" + by (simp add: numgcd_def zgcd_le1) + from prems have th': "c\0" by auto + from prems have cp: "c \ 0" by simp + from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] + have "0 < c div numgcd (CN 0 c (simpnum e))" by simp + } + with prems have ?case + by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} + ultimately show ?case by blast +next + case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i a))" + using simpfm_bound0 by blast + have "isatom (simpfm (Dvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def) + with bn bound0at_l show ?case by blast +next + case (NDvd i a) hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i a))" + using simpfm_bound0 by blast + have "isatom (simpfm (NDvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def) + with bn bound0at_l show ?case by blast +qed(auto simp add: conj_def imp_def disj_def iff_def Let_def simpfm_bound0 numadd_nb numneg_nb) + +lemma rlfm_I: + assumes qfp: "qfree p" + and xp: "0 \ x" and x1: "x < 1" + shows "(Ifm (x#bs) (rlfm p) = Ifm (x# bs) p) \ isrlfm (rlfm p)" + using qfp +by (induct p rule: rlfm.induct) +(auto simp add: rsplit[OF xp x1 lt_mono] lt_l rsplit[OF xp x1 le_mono] le_l rsplit[OF xp x1 gt_mono] gt_l + rsplit[OF xp x1 ge_mono] ge_l rsplit[OF xp x1 eq_mono] eq_l rsplit[OF xp x1 neq_mono] neq_l + rsplit[OF xp x1 DVD_mono[OF xp x1]] DVD_l rsplit[OF xp x1 NDVD_mono[OF xp x1]] NDVD_l simpfm_rl) +lemma rlfm_l: + assumes qfp: "qfree p" + shows "isrlfm (rlfm p)" + using qfp lt_l gt_l ge_l le_l eq_l neq_l DVD_l NDVD_l +by (induct p rule: rlfm.induct,auto simp add: simpfm_rl) + + (* Operations needed for Ferrante and Rackoff *) +lemma rminusinf_inf: + assumes lp: "isrlfm p" + shows "\ z. \ x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\ z. \ x. ?P z x p") +using lp +proof (induct p rule: minusinf.induct) + case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto +next + case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto +next + case (3 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + hence "real c * x + ?e \ 0" by simp + with xz have "?P ?z x (Eq (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp + thus ?case by blast +next + case (4 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + hence "real c * x + ?e \ 0" by simp + with xz have "?P ?z x (NEq (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp + thus ?case by blast +next + case (5 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + with xz have "?P ?z x (Lt (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp + thus ?case by blast +next + case (6 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + with xz have "?P ?z x (Le (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (Le (CN 0 c e))" by simp + thus ?case by blast +next + case (7 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + with xz have "?P ?z x (Gt (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp + thus ?case by blast +next + case (8 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + with xz have "?P ?z x (Ge (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp + thus ?case by blast +qed simp_all + +lemma rplusinf_inf: + assumes lp: "isrlfm p" + shows "\ z. \ x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\ z. \ x. ?P z x p") +using lp +proof (induct p rule: isrlfm.induct) + case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto +next + case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto +next + case (3 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + hence "real c * x + ?e \ 0" by simp + with xz have "?P ?z x (Eq (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp + thus ?case by blast +next + case (4 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + hence "real c * x + ?e \ 0" by simp + with xz have "?P ?z x (NEq (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp + thus ?case by blast +next + case (5 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + with xz have "?P ?z x (Lt (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp + thus ?case by blast +next + case (6 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + with xz have "?P ?z x (Le (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (Le (CN 0 c e))" by simp + thus ?case by blast +next + case (7 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + with xz have "?P ?z x (Gt (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp + thus ?case by blast +next + case (8 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + with xz have "?P ?z x (Ge (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp + thus ?case by blast +qed simp_all + +lemma rminusinf_bound0: + assumes lp: "isrlfm p" + shows "bound0 (minusinf p)" + using lp + by (induct p rule: minusinf.induct) simp_all + +lemma rplusinf_bound0: + assumes lp: "isrlfm p" + shows "bound0 (plusinf p)" + using lp + by (induct p rule: plusinf.induct) simp_all + +lemma rminusinf_ex: + assumes lp: "isrlfm p" + and ex: "Ifm (a#bs) (minusinf p)" + shows "\ x. Ifm (x#bs) p" +proof- + from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex + have th: "\ x. Ifm (x#bs) (minusinf p)" by auto + from rminusinf_inf[OF lp, where bs="bs"] + obtain z where z_def: "\x x. Ifm (x#bs) p" +proof- + from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex + have th: "\ x. Ifm (x#bs) (plusinf p)" by auto + from rplusinf_inf[OF lp, where bs="bs"] + obtain z where z_def: "\x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast + from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp + moreover have "z + 1 > z" by simp + ultimately show ?thesis using z_def by auto +qed + +consts + \:: "fm \ (num \ int) list" + \ :: "fm \ (num \ int) \ fm " +recdef \ "measure size" + "\ (And p q) = (\ p @ \ q)" + "\ (Or p q) = (\ p @ \ q)" + "\ (Eq (CN 0 c e)) = [(Neg e,c)]" + "\ (NEq (CN 0 c e)) = [(Neg e,c)]" + "\ (Lt (CN 0 c e)) = [(Neg e,c)]" + "\ (Le (CN 0 c e)) = [(Neg e,c)]" + "\ (Gt (CN 0 c e)) = [(Neg e,c)]" + "\ (Ge (CN 0 c e)) = [(Neg e,c)]" + "\ p = []" + +recdef \ "measure size" + "\ (And p q) = (\ (t,n). And (\ p (t,n)) (\ q (t,n)))" + "\ (Or p q) = (\ (t,n). Or (\ p (t,n)) (\ q (t,n)))" + "\ (Eq (CN 0 c e)) = (\ (t,n). Eq (Add (Mul c t) (Mul n e)))" + "\ (NEq (CN 0 c e)) = (\ (t,n). NEq (Add (Mul c t) (Mul n e)))" + "\ (Lt (CN 0 c e)) = (\ (t,n). Lt (Add (Mul c t) (Mul n e)))" + "\ (Le (CN 0 c e)) = (\ (t,n). Le (Add (Mul c t) (Mul n e)))" + "\ (Gt (CN 0 c e)) = (\ (t,n). Gt (Add (Mul c t) (Mul n e)))" + "\ (Ge (CN 0 c e)) = (\ (t,n). Ge (Add (Mul c t) (Mul n e)))" + "\ p = (\ (t,n). p)" + +lemma \_I: assumes lp: "isrlfm p" + and np: "real n > 0" and nbt: "numbound0 t" + shows "(Ifm (x#bs) (\ p (t,n)) = Ifm (((Inum (x#bs) t)/(real n))#bs) p) \ bound0 (\ p (t,n))" (is "(?I x (\ p (t,n)) = ?I ?u p) \ ?B p" is "(_ = ?I (?t/?n) p) \ _" is "(_ = ?I (?N x t /_) p) \ _") + using lp +proof(induct p rule: \.induct) + case (5 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + have "?I ?u (Lt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) < 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)" + by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) < 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +next + case (6 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" + by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) \ 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +next + case (7 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)" + by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) > 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +next + case (8 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" + by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) \ 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +next + case (3 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + from np have np: "real n \ 0" by simp + have "?I ?u (Eq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) = 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)" + by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) = 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +next + case (4 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + from np have np: "real n \ 0" by simp + have "?I ?u (NEq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" + by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) \ 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"] nth_pos2) + +lemma \_l: + assumes lp: "isrlfm p" + shows "\ (t,k) \ set (\ p). numbound0 t \ k >0" +using lp +by(induct p rule: \.induct) auto + +lemma rminusinf_\: + assumes lp: "isrlfm p" + and nmi: "\ (Ifm (a#bs) (minusinf p))" (is "\ (Ifm (a#bs) (?M p))") + and ex: "Ifm (x#bs) p" (is "?I x p") + shows "\ (s,m) \ set (\ p). x \ Inum (a#bs) s / real m" (is "\ (s,m) \ ?U p. x \ ?N a s / real m") +proof- + have "\ (s,m) \ set (\ p). real m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real m *x \ ?N a s") + using lp nmi ex + by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) + then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real m * x \ ?N a s" by blast + from \_l[OF lp] smU have mp: "real m > 0" by auto + from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" + by (auto simp add: mult_commute) + thus ?thesis using smU by auto +qed + +lemma rplusinf_\: + assumes lp: "isrlfm p" + and nmi: "\ (Ifm (a#bs) (plusinf p))" (is "\ (Ifm (a#bs) (?M p))") + and ex: "Ifm (x#bs) p" (is "?I x p") + shows "\ (s,m) \ set (\ p). x \ Inum (a#bs) s / real m" (is "\ (s,m) \ ?U p. x \ ?N a s / real m") +proof- + have "\ (s,m) \ set (\ p). real m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real m *x \ ?N a s") + using lp nmi ex + by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) + then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real m * x \ ?N a s" by blast + from \_l[OF lp] smU have mp: "real m > 0" by auto + from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" + by (auto simp add: mult_commute) + thus ?thesis using smU by auto +qed + +lemma lin_dense: + assumes lp: "isrlfm p" + and noS: "\ t. l < t \ t< u \ t \ (\ (t,n). Inum (x#bs) t / real n) ` set (\ p)" + (is "\ t. _ \ _ \ t \ (\ (t,n). ?N x t / real n ) ` (?U p)") + and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p" + and ly: "l < y" and yu: "y < u" + shows "Ifm (y#bs) p" +using lp px noS +proof (induct p rule: isrlfm.induct) + case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ + from prems have "x * real c + ?N x e < 0" by (simp add: ring_simps) + hence pxc: "x < (- ?N x e) / real c" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"]) + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with ly yu have yne: "y \ - ?N x e / real c" by auto + hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto + moreover {assume y: "y < (-?N x e)/ real c" + hence "y * real c < - ?N x e" + by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) + hence "real c * y + ?N x e < 0" by (simp add: ring_simps) + hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} + moreover {assume y: "y > (- ?N x e) / real c" + with yu have eu: "u > (- ?N x e) / real c" by auto + with noSc ly yu have "(- ?N x e) / real c \ l" by (cases "(- ?N x e) / real c > l", auto) + with lx pxc have "False" by auto + hence ?case by simp } + ultimately show ?case by blast +next + case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp + + from prems have "x * real c + ?N x e \ 0" by (simp add: ring_simps) + hence pxc: "x \ (- ?N x e) / real c" + by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"]) + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with ly yu have yne: "y \ - ?N x e / real c" by auto + hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto + moreover {assume y: "y < (-?N x e)/ real c" + hence "y * real c < - ?N x e" + by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) + hence "real c * y + ?N x e < 0" by (simp add: ring_simps) + hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} + moreover {assume y: "y > (- ?N x e) / real c" + with yu have eu: "u > (- ?N x e) / real c" by auto + with noSc ly yu have "(- ?N x e) / real c \ l" by (cases "(- ?N x e) / real c > l", auto) + with lx pxc have "False" by auto + hence ?case by simp } + ultimately show ?case by blast +next + case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ + from prems have "x * real c + ?N x e > 0" by (simp add: ring_simps) + hence pxc: "x > (- ?N x e) / real c" + by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"]) + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with ly yu have yne: "y \ - ?N x e / real c" by auto + hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto + moreover {assume y: "y > (-?N x e)/ real c" + hence "y * real c > - ?N x e" + by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) + hence "real c * y + ?N x e > 0" by (simp add: ring_simps) + hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} + moreover {assume y: "y < (- ?N x e) / real c" + with ly have eu: "l < (- ?N x e) / real c" by auto + with noSc ly yu have "(- ?N x e) / real c \ u" by (cases "(- ?N x e) / real c > l", auto) + with xu pxc have "False" by auto + hence ?case by simp } + ultimately show ?case by blast +next + case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ + from prems have "x * real c + ?N x e \ 0" by (simp add: ring_simps) + hence pxc: "x \ (- ?N x e) / real c" + by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"]) + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with ly yu have yne: "y \ - ?N x e / real c" by auto + hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto + moreover {assume y: "y > (-?N x e)/ real c" + hence "y * real c > - ?N x e" + by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) + hence "real c * y + ?N x e > 0" by (simp add: ring_simps) + hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} + moreover {assume y: "y < (- ?N x e) / real c" + with ly have eu: "l < (- ?N x e) / real c" by auto + with noSc ly yu have "(- ?N x e) / real c \ u" by (cases "(- ?N x e) / real c > l", auto) + with xu pxc have "False" by auto + hence ?case by simp } + ultimately show ?case by blast +next + case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ + from cp have cnz: "real c \ 0" by simp + from prems have "x * real c + ?N x e = 0" by (simp add: ring_simps) + hence pxc: "x = (- ?N x e) / real c" + by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"]) + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with lx xu have yne: "x \ - ?N x e / real c" by auto + with pxc show ?case by simp +next + case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ + from cp have cnz: "real c \ 0" by simp + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with ly yu have yne: "y \ - ?N x e / real c" by auto + hence "y* real c \ -?N x e" + by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp + hence "y* real c + ?N x e \ 0" by (simp add: ring_simps) + thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] + by (simp add: ring_simps) +qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="y" and b'="x"]) + +lemma finite_set_intervals: + assumes px: "P (x::real)" + and lx: "l \ x" and xu: "x \ u" + and linS: "l\ S" and uinS: "u \ S" + and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" + shows "\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a \ x \ x \ b \ P x" +proof- + let ?Mx = "{y. y\ S \ y \ x}" + let ?xM = "{y. y\ S \ x \ y}" + let ?a = "Max ?Mx" + let ?b = "Min ?xM" + have MxS: "?Mx \ S" by blast + hence fMx: "finite ?Mx" using fS finite_subset by auto + from lx linS have linMx: "l \ ?Mx" by blast + hence Mxne: "?Mx \ {}" by blast + have xMS: "?xM \ S" by blast + hence fxM: "finite ?xM" using fS finite_subset by auto + from xu uinS have linxM: "u \ ?xM" by blast + hence xMne: "?xM \ {}" by blast + have ax:"?a \ x" using Mxne fMx by auto + have xb:"x \ ?b" using xMne fxM by auto + have "?a \ ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \ S" using MxS by blast + have "?b \ ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \ S" using xMS by blast + have noy:"\ y. ?a < y \ y < ?b \ y \ S" + proof(clarsimp) + fix y + assume ay: "?a < y" and yb: "y < ?b" and yS: "y \ S" + from yS have "y\ ?Mx \ y\ ?xM" by auto + moreover {assume "y \ ?Mx" hence "y \ ?a" using Mxne fMx by auto with ay have "False" by simp} + moreover {assume "y \ ?xM" hence "y \ ?b" using xMne fxM by auto with yb have "False" by simp} + ultimately show "False" by blast + qed + from ainS binS noy ax xb px show ?thesis by blast +qed + +lemma finite_set_intervals2: + assumes px: "P (x::real)" + and lx: "l \ x" and xu: "x \ u" + and linS: "l\ S" and uinS: "u \ S" + and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" + shows "(\ s\ S. P s) \ (\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a < x \ x < b \ P x)" +proof- + from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su] + obtain a and b where + as: "a\ S" and bs: "b\ S" and noS:"\y. a < y \ y < b \ y \ S" and axb: "a \ x \ x \ b \ P x" by auto + from axb have "x= a \ x= b \ (a < x \ x < b)" by auto + thus ?thesis using px as bs noS by blast +qed + +lemma rinf_\: + assumes lp: "isrlfm p" + and nmi: "\ (Ifm (x#bs) (minusinf p))" (is "\ (Ifm (x#bs) (?M p))") + and npi: "\ (Ifm (x#bs) (plusinf p))" (is "\ (Ifm (x#bs) (?P p))") + and ex: "\ x. Ifm (x#bs) p" (is "\ x. ?I x p") + shows "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p" +proof- + let ?N = "\ x t. Inum (x#bs) t" + let ?U = "set (\ p)" + from ex obtain a where pa: "?I a p" by blast + from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi + have nmi': "\ (?I a (?M p))" by simp + from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi + have npi': "\ (?I a (?P p))" by simp + have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). ?I ((?N a l/real n + ?N a s /real m) / 2) p" + proof- + let ?M = "(\ (t,c). ?N a t / real c) ` ?U" + have fM: "finite ?M" by auto + from rminusinf_\[OF lp nmi pa] rplusinf_\[OF lp npi pa] + have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). a \ ?N x l / real n \ a \ ?N x s / real m" by blast + then obtain "t" "n" "s" "m" where + tnU: "(t,n) \ ?U" and smU: "(s,m) \ ?U" + and xs1: "a \ ?N x s / real m" and tx1: "a \ ?N x t / real n" by blast + from \_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 have xs: "a \ ?N a s / real m" and tx: "a \ ?N a t / real n" by auto + from tnU have Mne: "?M \ {}" by auto + hence Une: "?U \ {}" by simp + let ?l = "Min ?M" + let ?u = "Max ?M" + have linM: "?l \ ?M" using fM Mne by simp + have uinM: "?u \ ?M" using fM Mne by simp + have tnM: "?N a t / real n \ ?M" using tnU by auto + have smM: "?N a s / real m \ ?M" using smU by auto + have lM: "\ t\ ?M. ?l \ t" using Mne fM by auto + have Mu: "\ t\ ?M. t \ ?u" using Mne fM by auto + have "?l \ ?N a t / real n" using tnM Mne by simp hence lx: "?l \ a" using tx by simp + have "?N a s / real m \ ?u" using smM Mne by simp hence xu: "a \ ?u" using xs by simp + from finite_set_intervals2[where P="\ x. ?I x p",OF pa lx xu linM uinM fM lM Mu] + have "(\ s\ ?M. ?I s p) \ + (\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p)" . + moreover { fix u assume um: "u\ ?M" and pu: "?I u p" + hence "\ (tu,nu) \ ?U. u = ?N a tu / real nu" by auto + then obtain "tu" "nu" where tuU: "(tu,nu) \ ?U" and tuu:"u= ?N a tu / real nu" by blast + have "(u + u) / 2 = u" by auto with pu tuu + have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" by simp + with tuU have ?thesis by blast} + moreover{ + assume "\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p" + then obtain t1 and t2 where t1M: "t1 \ ?M" and t2M: "t2\ ?M" + and noM: "\ y. t1 < y \ y < t2 \ y \ ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p" + by blast + from t1M have "\ (t1u,t1n) \ ?U. t1 = ?N a t1u / real t1n" by auto + then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \ ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast + from t2M have "\ (t2u,t2n) \ ?U. t2 = ?N a t2u / real t2n" by auto + then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \ ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast + from t1x xt2 have t1t2: "t1 < t2" by simp + let ?u = "(t1 + t2) / 2" + from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto + from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" . + with t1uU t2uU t1u t2u have ?thesis by blast} + ultimately show ?thesis by blast + qed + then obtain "l" "n" "s" "m" where lnU: "(l,n) \ ?U" and smU:"(s,m) \ ?U" + and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast + from lnU smU \_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto + from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] + numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu + have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp + with lnU smU + show ?thesis by auto +qed + (* The Ferrante - Rackoff Theorem *) + +theorem fr_eq: + assumes lp: "isrlfm p" + shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,n) \ set (\ p). \ (s,m) \ set (\ p). Ifm ((((Inum (x#bs) t)/ real n + (Inum (x#bs) s) / real m) /2)#bs) p))" + (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") +proof + assume px: "\ x. ?I x p" + have "?M \ ?P \ (\ ?M \ \ ?P)" by blast + moreover {assume "?M \ ?P" hence "?D" by blast} + moreover {assume nmi: "\ ?M" and npi: "\ ?P" + from rinf_\[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast} + ultimately show "?D" by blast +next + assume "?D" + moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} + moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } + moreover {assume f:"?F" hence "?E" by blast} + ultimately show "?E" by blast +qed + + +lemma fr_eq\: + assumes lp: "isrlfm p" + shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,k) \ set (\ p). \ (s,l) \ set (\ p). Ifm (x#bs) (\ p (Add(Mul l t) (Mul k s) , 2*k*l))))" + (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") +proof + assume px: "\ x. ?I x p" + have "?M \ ?P \ (\ ?M \ \ ?P)" by blast + moreover {assume "?M \ ?P" hence "?D" by blast} + moreover {assume nmi: "\ ?M" and npi: "\ ?P" + let ?f ="\ (t,n). Inum (x#bs) t / real n" + let ?N = "\ t. Inum (x#bs) t" + {fix t n s m assume "(t,n)\ set (\ p)" and "(s,m) \ set (\ p)" + with \_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" + by auto + let ?st = "Add (Mul m t) (Mul n s)" + from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" + by (simp add: mult_commute) + from tnb snb have st_nb: "numbound0 ?st" by simp + have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + using mnp mp np by (simp add: ring_simps add_divide_distrib) + from \_I[OF lp mnp st_nb, where x="x" and bs="bs"] + have "?I x (\ p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])} + with rinf_\[OF lp nmi npi px] have "?F" by blast hence "?D" by blast} + ultimately show "?D" by blast +next + assume "?D" + moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} + moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } + moreover {fix t k s l assume "(t,k) \ set (\ p)" and "(s,l) \ set (\ p)" + and px:"?I x (\ p (Add (Mul l t) (Mul k s), 2*k*l))" + with \_l[OF lp] have tnb: "numbound0 t" and np:"real k > 0" and snb: "numbound0 s" and mp:"real l > 0" by auto + let ?st = "Add (Mul l t) (Mul k s)" + from mult_pos_pos[OF np mp] have mnp: "real (2*k*l) > 0" + by (simp add: mult_commute) + from tnb snb have st_nb: "numbound0 ?st" by simp + from \_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto} + ultimately show "?E" by blast +qed + +text{* The overall Part *} + +lemma real_ex_int_real01: + shows "(\ (x::real). P x) = (\ (i::int) (u::real). 0\ u \ u< 1 \ P (real i + u))" +proof(auto) + fix x + assume Px: "P x" + let ?i = "floor x" + let ?u = "x - real ?i" + have "x = real ?i + ?u" by simp + hence "P (real ?i + ?u)" using Px by simp + moreover have "real ?i \ x" using real_of_int_floor_le by simp hence "0 \ ?u" by arith + moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith + ultimately show "(\ (i::int) (u::real). 0\ u \ u< 1 \ P (real i + u))" by blast +qed + +consts exsplitnum :: "num \ num" + exsplit :: "fm \ fm" +recdef exsplitnum "measure size" + "exsplitnum (C c) = (C c)" + "exsplitnum (Bound 0) = Add (Bound 0) (Bound 1)" + "exsplitnum (Bound n) = Bound (n+1)" + "exsplitnum (Neg a) = Neg (exsplitnum a)" + "exsplitnum (Add a b) = Add (exsplitnum a) (exsplitnum b) " + "exsplitnum (Sub a b) = Sub (exsplitnum a) (exsplitnum b) " + "exsplitnum (Mul c a) = Mul c (exsplitnum a)" + "exsplitnum (Floor a) = Floor (exsplitnum a)" + "exsplitnum (CN 0 c a) = CN 0 c (Add (Mul c (Bound 1)) (exsplitnum a))" + "exsplitnum (CN n c a) = CN (n+1) c (exsplitnum a)" + "exsplitnum (CF c s t) = CF c (exsplitnum s) (exsplitnum t)" + +recdef exsplit "measure size" + "exsplit (Lt a) = Lt (exsplitnum a)" + "exsplit (Le a) = Le (exsplitnum a)" + "exsplit (Gt a) = Gt (exsplitnum a)" + "exsplit (Ge a) = Ge (exsplitnum a)" + "exsplit (Eq a) = Eq (exsplitnum a)" + "exsplit (NEq a) = NEq (exsplitnum a)" + "exsplit (Dvd i a) = Dvd i (exsplitnum a)" + "exsplit (NDvd i a) = NDvd i (exsplitnum a)" + "exsplit (And p q) = And (exsplit p) (exsplit q)" + "exsplit (Or p q) = Or (exsplit p) (exsplit q)" + "exsplit (Imp p q) = Imp (exsplit p) (exsplit q)" + "exsplit (Iff p q) = Iff (exsplit p) (exsplit q)" + "exsplit (NOT p) = NOT (exsplit p)" + "exsplit p = p" + +lemma exsplitnum: + "Inum (x#y#bs) (exsplitnum t) = Inum ((x+y) #bs) t" + by(induct t rule: exsplitnum.induct) (simp_all add: ring_simps) + +lemma exsplit: + assumes qfp: "qfree p" + shows "Ifm (x#y#bs) (exsplit p) = Ifm ((x+y)#bs) p" +using qfp exsplitnum[where x="x" and y="y" and bs="bs"] +by(induct p rule: exsplit.induct) simp_all + +lemma splitex: + assumes qf: "qfree p" + shows "(Ifm bs (E p)) = (\ (i::int). Ifm (real i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs") +proof- + have "?rhs = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm (x#(real i)#bs) (exsplit p))" + by (simp add: myless[rule_format, where b="1"] myless[rule_format, where b="0"] add_ac diff_def) + also have "\ = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm ((real i + x) #bs) p)" + by (simp only: exsplit[OF qf] add_ac) + also have "\ = (\ x. Ifm (x#bs) p)" + by (simp only: real_ex_int_real01[where P="\ x. Ifm (x#bs) p"]) + finally show ?thesis by simp +qed + + (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *) + +constdefs ferrack01:: "fm \ fm" + "ferrack01 p \ (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p); + U = remdups(map simp_num_pair + (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) + (alluopairs (\ p')))) + in decr (evaldjf (\ p') U ))" + +lemma fr_eq_01: + assumes qf: "qfree p" + shows "(\ x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (\ (t,n) \ set (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). \ (s,m) \ set (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). Ifm (x#bs) (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) (Add (Mul m t) (Mul n s), 2*n*m)))" + (is "(\ x. ?I x ?q) = ?F") +proof- + let ?rq = "rlfm ?q" + let ?M = "?I x (minusinf ?rq)" + let ?P = "?I x (plusinf ?rq)" + have MF: "?M = False" + apply (simp add: Let_def reducecoeff_def numgcd_def zgcd_def rsplit_def ge_def lt_def conj_def disj_def) + by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all) + have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def zgcd_def rsplit_def ge_def lt_def conj_def disj_def) + by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all) + have "(\ x. ?I x ?q ) = + ((?I x (minusinf ?rq)) \ (?I x (plusinf ?rq )) \ (\ (t,n) \ set (\ ?rq). \ (s,m) \ set (\ ?rq ). ?I x (\ ?rq (Add (Mul m t) (Mul n s), 2*n*m))))" + (is "(\ x. ?I x ?q) = (?M \ ?P \ ?F)" is "?E = ?D") + proof + assume "\ x. ?I x ?q" + then obtain x where qx: "?I x ?q" by blast + hence xp: "0\ x" and x1: "x< 1" and px: "?I x p" + by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf]) + from qx have "?I x ?rq " + by (simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1]) + hence lqx: "?I x ?rq " using simpfm[where p="?rq" and bs="x#bs"] by auto + from qf have qfq:"isrlfm ?rq" + by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1]) + with lqx fr_eq\[OF qfq] show "?M \ ?P \ ?F" by blast + next + assume D: "?D" + let ?U = "set (\ ?rq )" + from MF PF D have "?F" by auto + then obtain t n s m where aU:"(t,n) \ ?U" and bU:"(s,m)\ ?U" and rqx: "?I x (\ ?rq (Add (Mul m t) (Mul n s), 2*n*m))" by blast + from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf] + by (auto simp add: rsplit_def lt_def ge_def) + from aU bU \_l[OF lrq] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" by (auto simp add: split_def) + let ?st = "Add (Mul m t) (Mul n s)" + from tnb snb have stnb: "numbound0 ?st" by simp + from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" + by (simp add: mult_commute) + from conjunct1[OF \_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx + have "\ x. ?I x ?rq" by auto + thus "?E" + using rlfm_I[OF qf] by (auto simp add: rsplit_def lt_def ge_def) + qed + with MF PF show ?thesis by blast +qed + +lemma \_cong_aux: + assumes Ul: "\ (t,n) \ set U. numbound0 t \ n >0" + shows "((\ (t,n). Inum (x#bs) t /real n) ` (set (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs U)))) = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (set U \ set U))" + (is "?lhs = ?rhs") +proof(auto) + fix t n s m + assume "((t,n),(s,m)) \ set (alluopairs U)" + hence th: "((t,n),(s,m)) \ (set U \ set U)" + using alluopairs_set1[where xs="U"] by blast + let ?N = "\ t. Inum (x#bs) t" + let ?st= "Add (Mul m t) (Mul n s)" + from Ul th have mnz: "m \ 0" by auto + from Ul th have nnz: "n \ 0" by auto + have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + using mnz nnz by (simp add: ring_simps add_divide_distrib) + + thus "(real m * Inum (x # bs) t + real n * Inum (x # bs) s) / + (2 * real n * real m) + \ (\((t, n), s, m). + (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) ` + (set U \ set U)"using mnz nnz th + apply (auto simp add: th add_divide_distrib ring_simps split_def image_def) + by (rule_tac x="(s,m)" in bexI,simp_all) + (rule_tac x="(t,n)" in bexI,simp_all) +next + fix t n s m + assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" + let ?N = "\ t. Inum (x#bs) t" + let ?st= "Add (Mul m t) (Mul n s)" + from Ul smU have mnz: "m \ 0" by auto + from Ul tnU have nnz: "n \ 0" by auto + have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + using mnz nnz by (simp add: ring_simps add_divide_distrib) + let ?P = "\ (t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" + have Pc:"\ a b. ?P a b = ?P b a" + by auto + from Ul alluopairs_set1 have Up:"\ ((t,n),(s,m)) \ set (alluopairs U). n \ 0 \ m \ 0" by blast + from alluopairs_ex[OF Pc, where xs="U"] tnU smU + have th':"\ ((t',n'),(s',m')) \ set (alluopairs U). ?P (t',n') (s',m')" + by blast + then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \ set (alluopairs U)" + and Pts': "?P (t',n') (s',m')" by blast + from ts'_U Up have mnz': "m' \ 0" and nnz': "n'\ 0" by auto + let ?st' = "Add (Mul m' t') (Mul n' s')" + have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')" + using mnz' nnz' by (simp add: ring_simps add_divide_distrib) + from Pts' have + "(Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" by simp + also have "\ = ((\(t, n). Inum (x # bs) t / real n) ((\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st') + finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2 + \ (\(t, n). Inum (x # bs) t / real n) ` + (\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ` + set (alluopairs U)" + using ts'_U by blast +qed + +lemma \_cong: + assumes lp: "isrlfm p" + and UU': "((\ (t,n). Inum (x#bs) t /real n) ` U') = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (U \ U))" (is "?f ` U' = ?g ` (U\U)") + and U: "\ (t,n) \ U. numbound0 t \ n > 0" + and U': "\ (t,n) \ U'. numbound0 t \ n > 0" + shows "(\ (t,n) \ U. \ (s,m) \ U. Ifm (x#bs) (\ p (Add (Mul m t) (Mul n s),2*n*m))) = (\ (t,n) \ U'. Ifm (x#bs) (\ p (t,n)))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and + Pst: "Ifm (x#bs) (\ p (Add (Mul m t) (Mul n s),2*n*m))" by blast + let ?N = "\ t. Inum (x#bs) t" + from tnU smU U have tnb: "numbound0 t" and np: "n > 0" + and snb: "numbound0 s" and mp:"m > 0" by auto + let ?st= "Add (Mul m t) (Mul n s)" + from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" + by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) + from tnb snb have stnb: "numbound0 ?st" by simp + have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + using mp np by (simp add: ring_simps add_divide_distrib) + from tnU smU UU' have "?g ((t,n),(s,m)) \ ?f ` U'" by blast + hence "\ (t',n') \ U'. ?g ((t,n),(s,m)) = ?f (t',n')" + by auto (rule_tac x="(a,b)" in bexI, auto) + then obtain t' n' where tnU': "(t',n') \ U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast + from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto + from \_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst + have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp + from conjunct1[OF \_I[OF lp np' tnb', where bs="bs" and x="x"], symmetric] th[simplified split_def fst_conv snd_conv,symmetric] Pst2[simplified st[symmetric]] + have "Ifm (x # bs) (\ p (t', n')) " by (simp only: st) + then show ?rhs using tnU' by auto +next + assume ?rhs + then obtain t' n' where tnU': "(t',n') \ U'" and Pt': "Ifm (x # bs) (\ p (t', n'))" + by blast + from tnU' UU' have "?f (t',n') \ ?g ` (U\U)" by blast + hence "\ ((t,n),(s,m)) \ (U\U). ?f (t',n') = ?g ((t,n),(s,m))" + by auto (rule_tac x="(a,b)" in bexI, auto) + then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and + th: "?f (t',n') = ?g((t,n),(s,m)) "by blast + let ?N = "\ t. Inum (x#bs) t" + from tnU smU U have tnb: "numbound0 t" and np: "n > 0" + and snb: "numbound0 s" and mp:"m > 0" by auto + let ?st= "Add (Mul m t) (Mul n s)" + from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" + by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) + from tnb snb have stnb: "numbound0 ?st" by simp + have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + using mp np by (simp add: ring_simps add_divide_distrib) + from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto + from \_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt' + have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp + with \_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast +qed + +lemma ferrack01: + assumes qf: "qfree p" + shows "((\ x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (Ifm bs (ferrack01 p))) \ qfree (ferrack01 p)" (is "(?lhs = ?rhs) \ _") +proof- + let ?I = "\ x p. Ifm (x#bs) p" + fix x + let ?N = "\ t. Inum (x#bs) t" + let ?q = "rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)" + let ?U = "\ ?q" + let ?Up = "alluopairs ?U" + let ?g = "\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)" + let ?S = "map ?g ?Up" + let ?SS = "map simp_num_pair ?S" + let ?Y = "remdups ?SS" + let ?f= "(\ (t,n). ?N t / real n)" + let ?h = "\ ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2" + let ?F = "\ p. \ a \ set (\ p). \ b \ set (\ p). ?I x (\ p (?g(a,b)))" + let ?ep = "evaldjf (\ ?q) ?Y" + from rlfm_l[OF qf] have lq: "isrlfm ?q" + by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def zgcd_def) + from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \ (set ?U \ set ?U)" by simp + from \_l[OF lq] have U_l: "\ (t,n) \ set ?U. numbound0 t \ n > 0" . + from U_l UpU + have Up_: "\ ((t,n),(s,m)) \ set ?Up. numbound0 t \ n> 0 \ numbound0 s \ m > 0" by auto + hence Snb: "\ (t,n) \ set ?S. numbound0 t \ n > 0 " + by (auto simp add: mult_pos_pos) + have Y_l: "\ (t,n) \ set ?Y. numbound0 t \ n > 0" + proof- + { fix t n assume tnY: "(t,n) \ set ?Y" + hence "(t,n) \ set ?SS" by simp + hence "\ (t',n') \ set ?S. simp_num_pair (t',n') = (t,n)" + by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all) + then obtain t' n' where tn'S: "(t',n') \ set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast + from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto + from simp_num_pair_l[OF tnb np tns] + have "numbound0 t \ n > 0" . } + thus ?thesis by blast + qed + + have YU: "(?f ` set ?Y) = (?h ` (set ?U \ set ?U))" + proof- + from simp_num_pair_ci[where bs="x#bs"] have + "\x. (?f o simp_num_pair) x = ?f x" by auto + hence th: "?f o simp_num_pair = ?f" using ext by blast + have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_compose) + also have "\ = (?f ` set ?S)" by (simp add: th) + also have "\ = ((?f o ?g) ` set ?Up)" + by (simp only: set_map o_def image_compose[symmetric]) + also have "\ = (?h ` (set ?U \ set ?U))" + using \_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_compose[symmetric]] by blast + finally show ?thesis . + qed + have "\ (t,n) \ set ?Y. bound0 (\ ?q (t,n))" + proof- + { fix t n assume tnY: "(t,n) \ set ?Y" + with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto + from \_I[OF lq np tnb] + have "bound0 (\ ?q (t,n))" by simp} + thus ?thesis by blast + qed + hence ep_nb: "bound0 ?ep" using evaldjf_bound0[where xs="?Y" and f="\ ?q"] + by auto + + from fr_eq_01[OF qf, where bs="bs" and x="x"] have "?lhs = ?F ?q" + by (simp only: split_def fst_conv snd_conv) + also have "\ = (\ (t,n) \ set ?Y. ?I x (\ ?q (t,n)))" using \_cong[OF lq YU U_l Y_l] + by (simp only: split_def fst_conv snd_conv) + also have "\ = (Ifm (x#bs) ?ep)" + using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="\ ?q",symmetric] + by (simp only: split_def pair_collapse) + also have "\ = (Ifm bs (decr ?ep))" using decr[OF ep_nb] by blast + finally have lr: "?lhs = ?rhs" by (simp only: ferrack01_def Let_def) + from decr_qf[OF ep_nb] have "qfree (ferrack01 p)" by (simp only: Let_def ferrack01_def) + with lr show ?thesis by blast +qed + +lemma cp_thm': + assumes lp: "iszlfm p (real (i::int)#bs)" + and up: "d\ p 1" and dd: "d\ p d" and dp: "d > 0" + shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. d}. Ifm (real j#bs) (minusinf p)) \ (\ j\ {1.. d}. \ b\ (Inum (real i#bs)) ` set (\ p). Ifm ((b+real j)#bs) p))" + using cp_thm[OF lp up dd dp] by auto + +constdefs unit:: "fm \ fm \ num list \ int" + "unit p \ (let p' = zlfm p ; l = \ p' ; q = And (Dvd l (CN 0 1 (C 0))) (a\ p' l); d = \ q; + B = remdups (map simpnum (\ q)) ; a = remdups (map simpnum (\ q)) + in if length B \ length a then (q,B,d) else (mirror q, a,d))" + +lemma unit: assumes qf: "qfree p" + shows "\ q B d. unit p = (q,B,d) \ ((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ (Inum (real i#bs)) ` set B = (Inum (real i#bs)) ` set (\ q) \ d\ q 1 \ d\ q d \ d >0 \ iszlfm q (real (i::int)#bs) \ (\ b\ set B. numbound0 b)" +proof- + fix q B d + assume qBd: "unit p = (q,B,d)" + let ?thes = "((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ + Inum (real i#bs) ` set B = Inum (real i#bs) ` set (\ q) \ + d\ q 1 \ d\ q d \ 0 < d \ iszlfm q (real i # bs) \ (\ b\ set B. numbound0 b)" + let ?I = "\ (x::int) p. Ifm (real x#bs) p" + let ?p' = "zlfm p" + let ?l = "\ ?p'" + let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a\ ?p' ?l)" + let ?d = "\ ?q" + let ?B = "set (\ ?q)" + let ?B'= "remdups (map simpnum (\ ?q))" + let ?A = "set (\ ?q)" + let ?A'= "remdups (map simpnum (\ ?q))" + from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] + have pp': "\ i. ?I i ?p' = ?I i p" by auto + from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]] + have lp': "\ (i::int). iszlfm ?p' (real i#bs)" by simp + hence lp'': "iszlfm ?p' (real (i::int)#bs)" by simp + from lp' \[where p="?p'" and bs="bs"] have lp: "?l >0" and dl: "d\ ?p' ?l" by auto + from a\_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp' + have pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by (simp add: int_rdvd_iff) + from lp'' lp a\[OF lp'' dl lp] have lq:"iszlfm ?q (real i#bs)" and uq: "d\ ?q 1" + by (auto simp add: isint_def) + from \[OF lq] have dp:"?d >0" and dd: "d\ ?q ?d" by blast+ + let ?N = "\ t. Inum (real (i::int)#bs) t" + have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_compose) + also have "\ = ?N ` ?B" using simpnum_ci[where bs="real i #bs"] by auto + finally have BB': "?N ` set ?B' = ?N ` ?B" . + have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_compose) + also have "\ = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] by auto + finally have AA': "?N ` set ?A' = ?N ` ?A" . + from \_numbound0[OF lq] have B_nb:"\ b\ set ?B'. numbound0 b" + by (simp add: simpnum_numbound0) + from \_l[OF lq] have A_nb: "\ b\ set ?A'. numbound0 b" + by (simp add: simpnum_numbound0) + {assume "length ?B' \ length ?A'" + hence q:"q=?q" and "B = ?B'" and d:"d = ?d" + using qBd by (auto simp add: Let_def unit_def) + with BB' B_nb have b: "?N ` (set B) = ?N ` set (\ q)" + and bn: "\b\ set B. numbound0 b" by simp+ + with pq_ex dp uq dd lq q d have ?thes by simp} + moreover + {assume "\ (length ?B' \ length ?A')" + hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" + using qBd by (auto simp add: Let_def unit_def) + with AA' mirror\\[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\ q)" + and bn: "\b\ set B. numbound0 b" by simp+ + from mirror_ex[OF lq] pq_ex q + have pqm_eq:"(\ (x::int). ?I x p) = (\ (x::int). ?I x q)" by simp + from lq uq q mirror_d\ [where p="?q" and bs="bs" and a="real i"] + have lq': "iszlfm q (real i#bs)" and uq: "d\ q 1" by auto + from \[OF lq'] mirror_\[OF lq] q d have dq:"d\ q d " by auto + from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp + } + ultimately show ?thes by blast +qed + (* Cooper's Algorithm *) + +constdefs cooper :: "fm \ fm" + "cooper p \ + (let (q,B,d) = unit p; js = iupt (1,d); + mq = simpfm (minusinf q); + md = evaldjf (\ j. simpfm (subst0 (C j) mq)) js + in if md = T then T else + (let qd = evaldjf (\ t. simpfm (subst0 t q)) + (remdups (map (\ (b,j). simpnum (Add b (C j))) + [(b,j). b\B,j\js])) + in decr (disj md qd)))" +lemma cooper: assumes qf: "qfree p" + shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (cooper p))) \ qfree (cooper p)" + (is "(?lhs = ?rhs) \ _") +proof- + + let ?I = "\ (x::int) p. Ifm (real x#bs) p" + let ?q = "fst (unit p)" + let ?B = "fst (snd(unit p))" + let ?d = "snd (snd (unit p))" + let ?js = "iupt (1,?d)" + let ?mq = "minusinf ?q" + let ?smq = "simpfm ?mq" + let ?md = "evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js" + fix i + let ?N = "\ t. Inum (real (i::int)#bs) t" + let ?bjs = "[(b,j). b\?B,j\?js]" + let ?sbjs = "map (\ (b,j). simpnum (Add b (C j))) ?bjs" + let ?qd = "evaldjf (\ t. simpfm (subst0 t ?q)) (remdups ?sbjs)" + have qbf:"unit p = (?q,?B,?d)" by simp + from unit[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and + B:"?N ` set ?B = ?N ` set (\ ?q)" and + uq:"d\ ?q 1" and dd: "d\ ?q ?d" and dp: "?d > 0" and + lq: "iszlfm ?q (real i#bs)" and + Bn: "\ b\ set ?B. numbound0 b" by auto + from zlin_qfree[OF lq] have qfq: "qfree ?q" . + from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". + have jsnb: "\ j \ set ?js. numbound0 (C j)" by simp + hence "\ j\ set ?js. bound0 (subst0 (C j) ?smq)" + by (auto simp only: subst0_bound0[OF qfmq]) + hence th: "\ j\ set ?js. bound0 (simpfm (subst0 (C j) ?smq))" + by (auto simp add: simpfm_bound0) + from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp + from Bn jsnb have "\ (b,j) \ set ?bjs. numbound0 (Add b (C j))" + by simp + hence "\ (b,j) \ set ?bjs. numbound0 (simpnum (Add b (C j)))" + using simpnum_numbound0 by blast + hence "\ t \ set ?sbjs. numbound0 t" by simp + hence "\ t \ set (remdups ?sbjs). bound0 (subst0 t ?q)" + using subst0_bound0[OF qfq] by auto + hence th': "\ t \ set (remdups ?sbjs). bound0 (simpfm (subst0 t ?q))" + using simpfm_bound0 by blast + from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp + from mdb qdb + have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \ ?qd=T", simp_all) + from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B + have "?lhs = (\ j\ {1.. ?d}. ?I j ?mq \ (\ b\ ?N ` set ?B. Ifm ((b+ real j)#bs) ?q))" by auto + also have "\ = ((\ j\ set ?js. ?I j ?smq) \ (\ (b,j) \ (?N ` set ?B \ set ?js). Ifm ((b+ real j)#bs) ?q))" apply (simp only: iupt_set simpfm) by auto + also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ (\ (b,j). ?N (Add b (C j))) ` set ?bjs. Ifm (t #bs) ?q))" by simp + also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ (\ (b,j). ?N (simpnum (Add b (C j)))) ` set ?bjs. Ifm (t #bs) ?q))" by (simp only: simpnum_ci) + also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ set ?sbjs. Ifm (?N t #bs) ?q))" + by (auto simp add: split_def) + also have "\ = ((\ j\ set ?js. (\ j. ?I i (simpfm (subst0 (C j) ?smq))) j) \ (\ t \ set (remdups ?sbjs). (\ t. ?I i (simpfm (subst0 t ?q))) t))" by (simp only: simpfm subst0_I[OF qfq] simpfm Inum.simps subst0_I[OF qfmq] set_remdups) + also have "\ = ((?I i (evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js)) \ (?I i (evaldjf (\ t. simpfm (subst0 t ?q)) (remdups ?sbjs))))" by (simp only: evaldjf_ex) + finally have mdqd: "?lhs = (?I i (disj ?md ?qd))" by (simp add: disj) + hence mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" using decr [OF mdqdb] by simp + {assume mdT: "?md = T" + hence cT:"cooper p = T" + by (simp only: cooper_def unit_def split_def Let_def if_True) simp + from mdT mdqd have lhs:"?lhs" by (auto simp add: disj) + from mdT have "?rhs" by (simp add: cooper_def unit_def split_def) + with lhs cT have ?thesis by simp } + moreover + {assume mdT: "?md \ T" hence "cooper p = decr (disj ?md ?qd)" + by (simp only: cooper_def unit_def split_def Let_def if_False) + with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } + ultimately show ?thesis by blast +qed + +lemma DJcooper: + assumes qf: "qfree p" + shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ cooper p))) \ qfree (DJ cooper p)" +proof- + from cooper have cqf: "\ p. qfree p \ qfree (cooper p)" by blast + from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast + have "Ifm bs (DJ cooper p) = (\ q\ set (disjuncts p). Ifm bs (cooper q))" + by (simp add: DJ_def evaldjf_ex) + also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real x#bs) q)" + using cooper disjuncts_qf[OF qf] by blast + also have "\ = (\ (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto) + finally show ?thesis using thqf by blast +qed + + (* Redy and Loveland *) + +lemma \\_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'" + shows "Ifm (a#bs) (\\ p (t,c)) = Ifm (a#bs) (\\ p (t',c))" + using lp + by (induct p rule: iszlfm.induct, auto simp add: tt') + +lemma \_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'" + shows "Ifm (a#bs) (\ p c t) = Ifm (a#bs) (\ p c t')" + by (simp add: \_def tt' \\_cong[OF lp tt']) + +lemma \_cong: assumes lp: "iszlfm p (a#bs)" + and RR: "(\(b,k). (Inum (a#bs) b,k)) ` R = (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" + shows "(\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))) = (\ (e,c) \ set (\ p). \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j))))" + (is "?lhs = ?rhs") +proof + let ?d = "\ p" + assume ?lhs then obtain e c j where ecR: "(e,c) \ R" and jD:"j \ {1 .. c*?d}" + and px: "Ifm (a#bs) (\ p c (Add e (C j)))" (is "?sp c e j") by blast + from ecR have "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` R" by auto + hence "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" using RR by simp + hence "\ (e',c') \ set (\ p). Inum (a#bs) e = Inum (a#bs) e' \ c = c'" by auto + then obtain e' c' where ecRo:"(e',c') \ set (\ p)" and ee':"Inum (a#bs) e = Inum (a#bs) e'" + and cc':"c = c'" by blast + from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp + + from \_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp + from ecRo jD px' cc' show ?rhs apply auto + by (rule_tac x="(e', c')" in bexI,simp_all) + (rule_tac x="j" in bexI, simp_all add: cc'[symmetric]) +next + let ?d = "\ p" + assume ?rhs then obtain e c j where ecR: "(e,c) \ set (\ p)" and jD:"j \ {1 .. c*?d}" + and px: "Ifm (a#bs) (\ p c (Add e (C j)))" (is "?sp c e j") by blast + from ecR have "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" by auto + hence "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` R" using RR by simp + hence "\ (e',c') \ R. Inum (a#bs) e = Inum (a#bs) e' \ c = c'" by auto + then obtain e' c' where ecRo:"(e',c') \ R" and ee':"Inum (a#bs) e = Inum (a#bs) e'" + and cc':"c = c'" by blast + from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp + from \_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp + from ecRo jD px' cc' show ?lhs apply auto + by (rule_tac x="(e', c')" in bexI,simp_all) + (rule_tac x="j" in bexI, simp_all add: cc'[symmetric]) +qed + +lemma rl_thm': + assumes lp: "iszlfm p (real (i::int)#bs)" + and R: "(\(b,k). (Inum (a#bs) b,k)) ` R = (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" + shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real j#bs) (minusinf p)) \ (\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" + using rl_thm[OF lp] \_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp + +constdefs chooset:: "fm \ fm \ ((num\int) list) \ int" + "chooset p \ (let q = zlfm p ; d = \ q; + B = remdups (map (\ (t,k). (simpnum t,k)) (\ q)) ; + a = remdups (map (\ (t,k). (simpnum t,k)) (\\ q)) + in if length B \ length a then (q,B,d) else (mirror q, a,d))" + +lemma chooset: assumes qf: "qfree p" + shows "\ q B d. chooset p = (q,B,d) \ ((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ ((\(t,k). (Inum (real i#bs) t,k)) ` set B = (\(t,k). (Inum (real i#bs) t,k)) ` set (\ q)) \ (\ q = d) \ d >0 \ iszlfm q (real (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" +proof- + fix q B d + assume qBd: "chooset p = (q,B,d)" + let ?thes = "((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ ((\(t,k). (Inum (real i#bs) t,k)) ` set B = (\(t,k). (Inum (real i#bs) t,k)) ` set (\ q)) \ (\ q = d) \ d >0 \ iszlfm q (real (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" + let ?I = "\ (x::int) p. Ifm (real x#bs) p" + let ?q = "zlfm p" + let ?d = "\ ?q" + let ?B = "set (\ ?q)" + let ?f = "\ (t,k). (simpnum t,k)" + let ?B'= "remdups (map ?f (\ ?q))" + let ?A = "set (\\ ?q)" + let ?A'= "remdups (map ?f (\\ ?q))" + from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] + have pp': "\ i. ?I i ?q = ?I i p" by auto + hence pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by simp + from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real i"] + have lq: "iszlfm ?q (real (i::int)#bs)" . + from \[OF lq] have dp:"?d >0" by blast + let ?N = "\ (t,c). (Inum (real (i::int)#bs) t,c)" + have "?N ` set ?B' = ((?N o ?f) ` ?B)" by (simp add: split_def image_compose) + also have "\ = ?N ` ?B" + by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def) + finally have BB': "?N ` set ?B' = ?N ` ?B" . + have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_compose) + also have "\ = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] + by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def) + finally have AA': "?N ` set ?A' = ?N ` ?A" . + from \_l[OF lq] have B_nb:"\ (e,c)\ set ?B'. numbound0 e \ c > 0" + by (simp add: simpnum_numbound0 split_def) + from \\_l[OF lq] have A_nb: "\ (e,c)\ set ?A'. numbound0 e \ c > 0" + by (simp add: simpnum_numbound0 split_def) + {assume "length ?B' \ length ?A'" + hence q:"q=?q" and "B = ?B'" and d:"d = ?d" + using qBd by (auto simp add: Let_def chooset_def) + with BB' B_nb have b: "?N ` (set B) = ?N ` set (\ q)" + and bn: "\(e,c)\ set B. numbound0 e \ c > 0" by auto + with pq_ex dp lq q d have ?thes by simp} + moreover + {assume "\ (length ?B' \ length ?A')" + hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" + using qBd by (auto simp add: Let_def chooset_def) + with AA' mirror_\\[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\ q)" + and bn: "\(e,c)\ set B. numbound0 e \ c > 0" by auto + from mirror_ex[OF lq] pq_ex q + have pqm_eq:"(\ (x::int). ?I x p) = (\ (x::int). ?I x q)" by simp + from lq q mirror_l [where p="?q" and bs="bs" and a="real i"] + have lq': "iszlfm q (real i#bs)" by auto + from mirror_\[OF lq] pqm_eq b bn lq' dp q dp d have ?thes by simp + } + ultimately show ?thes by blast +qed + +constdefs stage:: "fm \ int \ (num \ int) \ fm" + "stage p d \ (\ (e,c). evaldjf (\ j. simpfm (\ p c (Add e (C j)))) (iupt (1,c*d)))" +lemma stage: + shows "Ifm bs (stage p d (e,c)) = (\ j\{1 .. c*d}. Ifm bs (\ p c (Add e (C j))))" + by (unfold stage_def split_def ,simp only: evaldjf_ex iupt_set simpfm) simp + +lemma stage_nb: assumes lp: "iszlfm p (a#bs)" and cp: "c >0" and nb:"numbound0 e" + shows "bound0 (stage p d (e,c))" +proof- + let ?f = "\ j. simpfm (\ p c (Add e (C j)))" + have th: "\ j\ set (iupt(1,c*d)). bound0 (?f j)" + proof + fix j + from nb have nb':"numbound0 (Add e (C j))" by simp + from simpfm_bound0[OF \_nb[OF lp nb', where k="c"]] + show "bound0 (simpfm (\ p c (Add e (C j))))" . + qed + from evaldjf_bound0[OF th] show ?thesis by (unfold stage_def split_def) simp +qed + +constdefs redlove:: "fm \ fm" + "redlove p \ + (let (q,B,d) = chooset p; + mq = simpfm (minusinf q); + md = evaldjf (\ j. simpfm (subst0 (C j) mq)) (iupt (1,d)) + in if md = T then T else + (let qd = evaldjf (stage q d) B + in decr (disj md qd)))" + +lemma redlove: assumes qf: "qfree p" + shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (redlove p))) \ qfree (redlove p)" + (is "(?lhs = ?rhs) \ _") +proof- + + let ?I = "\ (x::int) p. Ifm (real x#bs) p" + let ?q = "fst (chooset p)" + let ?B = "fst (snd(chooset p))" + let ?d = "snd (snd (chooset p))" + let ?js = "iupt (1,?d)" + let ?mq = "minusinf ?q" + let ?smq = "simpfm ?mq" + let ?md = "evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js" + fix i + let ?N = "\ (t,k). (Inum (real (i::int)#bs) t,k)" + let ?qd = "evaldjf (stage ?q ?d) ?B" + have qbf:"chooset p = (?q,?B,?d)" by simp + from chooset[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and + B:"?N ` set ?B = ?N ` set (\ ?q)" and dd: "\ ?q = ?d" and dp: "?d > 0" and + lq: "iszlfm ?q (real i#bs)" and + Bn: "\ (e,c)\ set ?B. numbound0 e \ c > 0" by auto + from zlin_qfree[OF lq] have qfq: "qfree ?q" . + from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". + have jsnb: "\ j \ set ?js. numbound0 (C j)" by simp + hence "\ j\ set ?js. bound0 (subst0 (C j) ?smq)" + by (auto simp only: subst0_bound0[OF qfmq]) + hence th: "\ j\ set ?js. bound0 (simpfm (subst0 (C j) ?smq))" + by (auto simp add: simpfm_bound0) + from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp + from Bn stage_nb[OF lq] have th:"\ x \ set ?B. bound0 (stage ?q ?d x)" by auto + from evaldjf_bound0[OF th] have qdb: "bound0 ?qd" . + from mdb qdb + have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \ ?qd=T", simp_all) + from trans [OF pq_ex rl_thm'[OF lq B]] dd + have "?lhs = ((\ j\ {1.. ?d}. ?I j ?mq) \ (\ (e,c)\ set ?B. \ j\ {1 .. c*?d}. Ifm (real i#bs) (\ ?q c (Add e (C j)))))" by auto + also have "\ = ((\ j\ {1.. ?d}. ?I j ?smq) \ (\ (e,c)\ set ?B. ?I i (stage ?q ?d (e,c) )))" + by (simp add: simpfm stage split_def) + also have "\ = ((\ j\ {1 .. ?d}. ?I i (subst0 (C j) ?smq)) \ ?I i ?qd)" + by (simp add: evaldjf_ex subst0_I[OF qfmq]) + finally have mdqd:"?lhs = (?I i ?md \ ?I i ?qd)" by (simp only: evaldjf_ex iupt_set simpfm) + also have "\ = (?I i (disj ?md ?qd))" by (simp add: disj) + also have "\ = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) + finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" . + {assume mdT: "?md = T" + hence cT:"redlove p = T" by (simp add: redlove_def Let_def chooset_def split_def) + from mdT have lhs:"?lhs" using mdqd by simp + from mdT have "?rhs" by (simp add: redlove_def chooset_def split_def) + with lhs cT have ?thesis by simp } + moreover + {assume mdT: "?md \ T" hence "redlove p = decr (disj ?md ?qd)" + by (simp add: redlove_def chooset_def split_def Let_def) + with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } + ultimately show ?thesis by blast +qed + +lemma DJredlove: + assumes qf: "qfree p" + shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ redlove p))) \ qfree (DJ redlove p)" +proof- + from redlove have cqf: "\ p. qfree p \ qfree (redlove p)" by blast + from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast + have "Ifm bs (DJ redlove p) = (\ q\ set (disjuncts p). Ifm bs (redlove q))" + by (simp add: DJ_def evaldjf_ex) + also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real x#bs) q)" + using redlove disjuncts_qf[OF qf] by blast + also have "\ = (\ (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto) + finally show ?thesis using thqf by blast +qed + + +lemma exsplit_qf: assumes qf: "qfree p" + shows "qfree (exsplit p)" +using qf by (induct p rule: exsplit.induct, auto) + +definition mircfr :: "fm \ fm" where + "mircfr = DJ cooper o ferrack01 o simpfm o exsplit" + +definition mirlfr :: "fm \ fm" where + "mirlfr = DJ redlove o ferrack01 o simpfm o exsplit" + +lemma mircfr: "\ bs p. qfree p \ qfree (mircfr p) \ Ifm bs (mircfr p) = Ifm bs (E p)" +proof(clarsimp simp del: Ifm.simps) + fix bs p + assume qf: "qfree p" + show "qfree (mircfr p)\(Ifm bs (mircfr p) = Ifm bs (E p))" (is "_ \ (?lhs = ?rhs)") + proof- + let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" + have "?rhs = (\ (i::int). \ x. Ifm (x#real i#bs) ?es)" + using splitex[OF qf] by simp + with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ + with DJcooper[OF qf'] show ?thesis by (simp add: mircfr_def) + qed +qed + +lemma mirlfr: "\ bs p. qfree p \ qfree(mirlfr p) \ Ifm bs (mirlfr p) = Ifm bs (E p)" +proof(clarsimp simp del: Ifm.simps) + fix bs p + assume qf: "qfree p" + show "qfree (mirlfr p)\(Ifm bs (mirlfr p) = Ifm bs (E p))" (is "_ \ (?lhs = ?rhs)") + proof- + let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" + have "?rhs = (\ (i::int). \ x. Ifm (x#real i#bs) ?es)" + using splitex[OF qf] by simp + with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ + with DJredlove[OF qf'] show ?thesis by (simp add: mirlfr_def) + qed +qed + +definition mircfrqe:: "fm \ fm" where + "mircfrqe p = qelim (prep p) mircfr" + +definition mirlfrqe:: "fm \ fm" where + "mirlfrqe p = qelim (prep p) mirlfr" + +theorem mircfrqe: "(Ifm bs (mircfrqe p) = Ifm bs p) \ qfree (mircfrqe p)" + using qelim_ci[OF mircfr] prep by (auto simp add: mircfrqe_def) + +theorem mirlfrqe: "(Ifm bs (mirlfrqe p) = Ifm bs p) \ qfree (mirlfrqe p)" + using qelim_ci[OF mirlfr] prep by (auto simp add: mirlfrqe_def) + +definition + "test1 (u\unit) = mircfrqe (A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0)))))))" + +definition + "test2 (u\unit) = mircfrqe (A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0)))))" + +definition + "test3 (u\unit) = mirlfrqe (A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0)))))))" + +definition + "test4 (u\unit) = mirlfrqe (A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0)))))" + +definition + "test5 (u\unit) = mircfrqe (A(E(And (Ge(Sub (Bound 1) (Bound 0))) (Eq (Add (Floor (Bound 1)) (Floor (Neg(Bound 0))))))))" + +ML {* @{code test1} () *} +ML {* @{code test2} () *} +ML {* @{code test3} () *} +ML {* @{code test4} () *} +ML {* @{code test5} () *} + +(*export_code mircfrqe mirlfrqe + in SML module_name Mir file "raw_mir.ML"*) + +oracle mirfr_oracle = {* fn (proofs, ct) => +let + +fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t + of NONE => error "Variable not found in the list!" + | SOME n => @{code Bound} n) + | num_of_term vs @{term "real (0::int)"} = @{code C} 0 + | num_of_term vs @{term "real (1::int)"} = @{code C} 1 + | num_of_term vs @{term "0::real"} = @{code C} 0 + | num_of_term vs @{term "1::real"} = @{code C} 1 + | num_of_term vs (Bound i) = @{code Bound} i + | num_of_term vs (@{term "uminus :: real \ real"} $ t') = @{code Neg} (num_of_term vs t') + | num_of_term vs (@{term "op + :: real \ real \ real"} $ t1 $ t2) = + @{code Add} (num_of_term vs t1, num_of_term vs t2) + | num_of_term vs (@{term "op - :: real \ real \ real"} $ t1 $ t2) = + @{code Sub} (num_of_term vs t1, num_of_term vs t2) + | num_of_term vs (@{term "op * :: real \ real \ real"} $ t1 $ t2) = + (case (num_of_term vs t1) + of @{code C} i => @{code Mul} (i, num_of_term vs t2) + | _ => error "num_of_term: unsupported Multiplication") + | num_of_term vs (@{term "real :: int \ real"} $ (@{term "number_of :: int \ int"} $ t')) = + @{code C} (HOLogic.dest_numeral t') + | num_of_term vs (@{term "real :: int \ real"} $ (@{term "floor :: real \ int"} $ t')) = + @{code Floor} (num_of_term vs t') + | num_of_term vs (@{term "real :: int \ real"} $ (@{term "ceiling :: real \ int"} $ t')) = + @{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t'))) + | num_of_term vs (@{term "number_of :: int \ real"} $ t') = + @{code C} (HOLogic.dest_numeral t') + | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t); + +fun fm_of_term vs @{term True} = @{code T} + | fm_of_term vs @{term False} = @{code F} + | fm_of_term vs (@{term "op < :: real \ real \ bool"} $ t1 $ t2) = + @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) + | fm_of_term vs (@{term "op \ :: real \ real \ bool"} $ t1 $ t2) = + @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) + | fm_of_term vs (@{term "op = :: real \ real \ bool"} $ t1 $ t2) = + @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) + | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \ real"} $ (@{term "number_of :: int \ int"} $ t1)) $ t2) = + @{code Dvd} (HOLogic.dest_numeral t1, num_of_term vs t2) + | fm_of_term vs (@{term "op = :: bool \ bool \ bool"} $ t1 $ t2) = + @{code Iff} (fm_of_term vs t1, fm_of_term vs t2) + | fm_of_term vs (@{term "op &"} $ t1 $ t2) = + @{code And} (fm_of_term vs t1, fm_of_term vs t2) + | fm_of_term vs (@{term "op |"} $ t1 $ t2) = + @{code Or} (fm_of_term vs t1, fm_of_term vs t2) + | fm_of_term vs (@{term "op -->"} $ t1 $ t2) = + @{code Imp} (fm_of_term vs t1, fm_of_term vs t2) + | fm_of_term vs (@{term "Not"} $ t') = + @{code NOT} (fm_of_term vs t') + | fm_of_term vs (Const ("Ex", _) $ Abs (xn, xT, p)) = + @{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) + | fm_of_term vs (Const ("All", _) $ Abs (xn, xT, p)) = + @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) + | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t); + +fun term_of_num vs (@{code C} i) = @{term "real :: int \ real"} $ HOLogic.mk_number HOLogic.intT i + | term_of_num vs (@{code Bound} n) = fst (the (find_first (fn (_, m) => n = m) vs)) + | term_of_num vs (@{code Neg} (@{code Floor} (@{code Neg} t'))) = + @{term "real :: int \ real"} $ (@{term "ceiling :: real \ int"} $ term_of_num vs t') + | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \ real"} $ term_of_num vs t' + | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: real \ real \ real"} $ + term_of_num vs t1 $ term_of_num vs t2 + | term_of_num vs (@{code Sub} (t1, t2)) = @{term "op - :: real \ real \ real"} $ + term_of_num vs t1 $ term_of_num vs t2 + | term_of_num vs (@{code Mul} (i, t2)) = @{term "op * :: real \ real \ real"} $ + term_of_num vs (@{code C} i) $ term_of_num vs t2 + | term_of_num vs (@{code Floor} t) = @{term "real :: int \ real"} $ (@{term "floor :: real \ int"} $ term_of_num vs t) + | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t)) + | term_of_num vs (@{code CF} (c, t, s)) = term_of_num vs (@{code Add} (@{code Mul} (c, @{code Floor} t), s)); + +fun term_of_fm vs @{code T} = HOLogic.true_const + | term_of_fm vs @{code F} = HOLogic.false_const + | term_of_fm vs (@{code Lt} t) = + @{term "op < :: real \ real \ bool"} $ term_of_num vs t $ @{term "0::real"} + | term_of_fm vs (@{code Le} t) = + @{term "op \ :: real \ real \ bool"} $ term_of_num vs t $ @{term "0::real"} + | term_of_fm vs (@{code Gt} t) = + @{term "op < :: real \ real \ bool"} $ @{term "0::real"} $ term_of_num vs t + | term_of_fm vs (@{code Ge} t) = + @{term "op \ :: real \ real \ bool"} $ @{term "0::real"} $ term_of_num vs t + | term_of_fm vs (@{code Eq} t) = + @{term "op = :: real \ real \ bool"} $ term_of_num vs t $ @{term "0::real"} + | term_of_fm vs (@{code NEq} t) = + term_of_fm vs (@{code NOT} (@{code Eq} t)) + | term_of_fm vs (@{code Dvd} (i, t)) = + @{term "op rdvd"} $ term_of_num vs (@{code C} i) $ term_of_num vs t + | term_of_fm vs (@{code NDvd} (i, t)) = + term_of_fm vs (@{code NOT} (@{code Dvd} (i, t))) + | term_of_fm vs (@{code NOT} t') = + HOLogic.Not $ term_of_fm vs t' + | term_of_fm vs (@{code And} (t1, t2)) = + HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2 + | term_of_fm vs (@{code Or} (t1, t2)) = + HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2 + | term_of_fm vs (@{code Imp} (t1, t2)) = + HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2 + | term_of_fm vs (@{code Iff} (t1, t2)) = + @{term "op = :: bool \ bool \ bool"} $ term_of_fm vs t1 $ term_of_fm vs t2; + +in + let + val thy = Thm.theory_of_cterm ct; + val t = Thm.term_of ct; + val fs = term_frees t; + val vs = fs ~~ (0 upto (length fs - 1)); + val qe = if proofs then @{code mirlfrqe} else @{code mircfrqe}; + val t' = (term_of_fm vs o qe o fm_of_term vs) t; + in (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end +end; +*} + +use "mirtac.ML" +setup "MirTac.setup" + +lemma "ALL (x::real). (\x\ = \x\ = (x = real \x\))" +apply mir +done + +lemma "ALL (x::real). real (2::int)*x - (real (1::int)) < real \x\ + real \x\ \ real \x\ + real \x\ \ real (2::int)*x + (real (1::int))" +apply mir +done + +lemma "ALL (x::real). 2*\x\ \ \2*x\ \ \2*x\ \ 2*\x+1\" +apply mir +done + +lemma "ALL (x::real). \y \ x. (\x\ = \y\)" +apply mir +done + +lemma "ALL x y. \x\ = \y\ \ 0 \ abs (y - x) \ abs (y - x) \ 1" +apply mir +done + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/NatSum.thy --- a/src/HOL/ex/NatSum.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/ex/NatSum.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,6 +1,5 @@ -(* Title: HOL/ex/NatSum.thy - ID: $Id$ - Author: Tobias Nipkow +(* Title: HOL/ex/NatSum.thy + Author: Tobias Nipkow *) header {* Summing natural numbers *} diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/NormalForm.thy --- a/src/HOL/ex/NormalForm.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/ex/NormalForm.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,11 +1,10 @@ -(* ID: $Id$ - Authors: Klaus Aehlig, Tobias Nipkow +(* Authors: Klaus Aehlig, Tobias Nipkow *) header {* Test of normalization function *} theory NormalForm -imports Main "~~/src/HOL/Real/Rational" +imports Main Rational begin lemma "True" by normalization diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOL/ex/ROOT.ML Wed Dec 03 15:58:44 2008 +0100 @@ -1,12 +1,9 @@ (* Title: HOL/ex/ROOT.ML - ID: $Id$ Miscellaneous examples for Higher-Order Logic. *) no_document use_thys [ - "Parity", - "Code_Eval", "State_Monad", "Efficient_Nat_examples", "ExecutableContent", @@ -101,15 +98,15 @@ ]; use_thys [ - "../Complex/ex/BinEx", - "../Complex/ex/Sqrt", - "../Complex/ex/Sqrt_Script", - "../Complex/ex/BigO_Complex", + "BinEx", + "Sqrt", + "Sqrt_Script", + "BigO_Complex", - "../Complex/ex/Arithmetic_Series_Complex", - "../Complex/ex/HarmonicSeries", + "Arithmetic_Series_Complex", + "HarmonicSeries", - "../Complex/ex/MIR", - "../Complex/ex/ReflectedFerrack" + "MIR", + "ReflectedFerrack" ]; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/ReflectedFerrack.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/ReflectedFerrack.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,2101 @@ +(* Title: Complex/ex/ReflectedFerrack.thy + Author: Amine Chaieb +*) + +theory ReflectedFerrack +imports Main GCD Real Efficient_Nat +uses ("linrtac.ML") +begin + +section {* Quantifier elimination for @{text "\ (0, 1, +, <)"} *} + + (*********************************************************************************) + (* SOME GENERAL STUFF< HAS TO BE MOVED IN SOME LIB *) + (*********************************************************************************) + +consts alluopairs:: "'a list \ ('a \ 'a) list" +primrec + "alluopairs [] = []" + "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)" + +lemma alluopairs_set1: "set (alluopairs xs) \ {(x,y). x\ set xs \ y\ set xs}" +by (induct xs, auto) + +lemma alluopairs_set: + "\x\ set xs ; y \ set xs\ \ (x,y) \ set (alluopairs xs) \ (y,x) \ set (alluopairs xs) " +by (induct xs, auto) + +lemma alluopairs_ex: + assumes Pc: "\ x y. P x y = P y x" + shows "(\ x \ set xs. \ y \ set xs. P x y) = (\ (x,y) \ set (alluopairs xs). P x y)" +proof + assume "\x\set xs. \y\set xs. P x y" + then obtain x y where x: "x \ set xs" and y:"y \ set xs" and P: "P x y" by blast + from alluopairs_set[OF x y] P Pc show"\(x, y)\set (alluopairs xs). P x y" + by auto +next + assume "\(x, y)\set (alluopairs xs). P x y" + then obtain "x" and "y" where xy:"(x,y) \ set (alluopairs xs)" and P: "P x y" by blast+ + from xy have "x \ set xs \ y\ set xs" using alluopairs_set1 by blast + with P show "\x\set xs. \y\set xs. P x y" by blast +qed + +lemma nth_pos2: "0 < n \ (x#xs) ! n = xs ! (n - 1)" +using Nat.gr0_conv_Suc +by clarsimp + +lemma filter_length: "length (List.filter P xs) < Suc (length xs)" + apply (induct xs, auto) done + +consts remdps:: "'a list \ 'a list" + +recdef remdps "measure size" + "remdps [] = []" + "remdps (x#xs) = (x#(remdps (List.filter (\ y. y \ x) xs)))" +(hints simp add: filter_length[rule_format]) + +lemma remdps_set[simp]: "set (remdps xs) = set xs" + by (induct xs rule: remdps.induct, auto) + + + + (*********************************************************************************) + (**** SHADOW SYNTAX AND SEMANTICS ****) + (*********************************************************************************) + +datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num + | Mul int num + + (* A size for num to make inductive proofs simpler*) +consts num_size :: "num \ nat" +primrec + "num_size (C c) = 1" + "num_size (Bound n) = 1" + "num_size (Neg a) = 1 + num_size a" + "num_size (Add a b) = 1 + num_size a + num_size b" + "num_size (Sub a b) = 3 + num_size a + num_size b" + "num_size (Mul c a) = 1 + num_size a" + "num_size (CN n c a) = 3 + num_size a " + + (* Semantics of numeral terms (num) *) +consts Inum :: "real list \ num \ real" +primrec + "Inum bs (C c) = (real c)" + "Inum bs (Bound n) = bs!n" + "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)" + "Inum bs (Neg a) = -(Inum bs a)" + "Inum bs (Add a b) = Inum bs a + Inum bs b" + "Inum bs (Sub a b) = Inum bs a - Inum bs b" + "Inum bs (Mul c a) = (real c) * Inum bs a" + (* FORMULAE *) +datatype fm = + T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| + NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm + + + (* A size for fm *) +consts fmsize :: "fm \ nat" +recdef fmsize "measure size" + "fmsize (NOT p) = 1 + fmsize p" + "fmsize (And p q) = 1 + fmsize p + fmsize q" + "fmsize (Or p q) = 1 + fmsize p + fmsize q" + "fmsize (Imp p q) = 3 + fmsize p + fmsize q" + "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)" + "fmsize (E p) = 1 + fmsize p" + "fmsize (A p) = 4+ fmsize p" + "fmsize p = 1" + (* several lemmas about fmsize *) +lemma fmsize_pos: "fmsize p > 0" +by (induct p rule: fmsize.induct) simp_all + + (* Semantics of formulae (fm) *) +consts Ifm ::"real list \ fm \ bool" +primrec + "Ifm bs T = True" + "Ifm bs F = False" + "Ifm bs (Lt a) = (Inum bs a < 0)" + "Ifm bs (Gt a) = (Inum bs a > 0)" + "Ifm bs (Le a) = (Inum bs a \ 0)" + "Ifm bs (Ge a) = (Inum bs a \ 0)" + "Ifm bs (Eq a) = (Inum bs a = 0)" + "Ifm bs (NEq a) = (Inum bs a \ 0)" + "Ifm bs (NOT p) = (\ (Ifm bs p))" + "Ifm bs (And p q) = (Ifm bs p \ Ifm bs q)" + "Ifm bs (Or p q) = (Ifm bs p \ Ifm bs q)" + "Ifm bs (Imp p q) = ((Ifm bs p) \ (Ifm bs q))" + "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)" + "Ifm bs (E p) = (\ x. Ifm (x#bs) p)" + "Ifm bs (A p) = (\ x. Ifm (x#bs) p)" + +lemma IfmLeSub: "\ Inum bs s = s' ; Inum bs t = t' \ \ Ifm bs (Le (Sub s t)) = (s' \ t')" +apply simp +done + +lemma IfmLtSub: "\ Inum bs s = s' ; Inum bs t = t' \ \ Ifm bs (Lt (Sub s t)) = (s' < t')" +apply simp +done +lemma IfmEqSub: "\ Inum bs s = s' ; Inum bs t = t' \ \ Ifm bs (Eq (Sub s t)) = (s' = t')" +apply simp +done +lemma IfmNOT: " (Ifm bs p = P) \ (Ifm bs (NOT p) = (\P))" +apply simp +done +lemma IfmAnd: " \ Ifm bs p = P ; Ifm bs q = Q\ \ (Ifm bs (And p q) = (P \ Q))" +apply simp +done +lemma IfmOr: " \ Ifm bs p = P ; Ifm bs q = Q\ \ (Ifm bs (Or p q) = (P \ Q))" +apply simp +done +lemma IfmImp: " \ Ifm bs p = P ; Ifm bs q = Q\ \ (Ifm bs (Imp p q) = (P \ Q))" +apply simp +done +lemma IfmIff: " \ Ifm bs p = P ; Ifm bs q = Q\ \ (Ifm bs (Iff p q) = (P = Q))" +apply simp +done + +lemma IfmE: " (!! x. Ifm (x#bs) p = P x) \ (Ifm bs (E p) = (\x. P x))" +apply simp +done +lemma IfmA: " (!! x. Ifm (x#bs) p = P x) \ (Ifm bs (A p) = (\x. P x))" +apply simp +done + +consts not:: "fm \ fm" +recdef not "measure size" + "not (NOT p) = p" + "not T = F" + "not F = T" + "not p = NOT p" +lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)" +by (cases p) auto + +constdefs conj :: "fm \ fm \ fm" + "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else + if p = q then p else And p q)" +lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)" +by (cases "p=F \ q=F",simp_all add: conj_def) (cases p,simp_all) + +constdefs disj :: "fm \ fm \ fm" + "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p + else if p=q then p else Or p q)" + +lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)" +by (cases "p=T \ q=T",simp_all add: disj_def) (cases p,simp_all) + +constdefs imp :: "fm \ fm \ fm" + "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p + else Imp p q)" +lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)" +by (cases "p=F \ q=T",simp_all add: imp_def) + +constdefs iff :: "fm \ fm \ fm" + "iff p q \ (if (p = q) then T else if (p = NOT q \ NOT p = q) then F else + if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else + Iff p q)" +lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)" + by (unfold iff_def,cases "p=q", simp,cases "p=NOT q", simp) (cases "NOT p= q", auto) + +lemma conj_simps: + "conj F Q = F" + "conj P F = F" + "conj T Q = Q" + "conj P T = P" + "conj P P = P" + "P \ T \ P \ F \ Q \ T \ Q \ F \ P \ Q \ conj P Q = And P Q" + by (simp_all add: conj_def) + +lemma disj_simps: + "disj T Q = T" + "disj P T = T" + "disj F Q = Q" + "disj P F = P" + "disj P P = P" + "P \ T \ P \ F \ Q \ T \ Q \ F \ P \ Q \ disj P Q = Or P Q" + by (simp_all add: disj_def) +lemma imp_simps: + "imp F Q = T" + "imp P T = T" + "imp T Q = Q" + "imp P F = not P" + "imp P P = T" + "P \ T \ P \ F \ P \ Q \ Q \ T \ Q \ F \ imp P Q = Imp P Q" + by (simp_all add: imp_def) +lemma trivNOT: "p \ NOT p" "NOT p \ p" +apply (induct p, auto) +done + +lemma iff_simps: + "iff p p = T" + "iff p (NOT p) = F" + "iff (NOT p) p = F" + "iff p F = not p" + "iff F p = not p" + "p \ NOT T \ iff T p = p" + "p\ NOT T \ iff p T = p" + "p\q \ p\ NOT q \ q\ NOT p \ p\ F \ q\ F \ p \ T \ q \ T \ iff p q = Iff p q" + using trivNOT + by (simp_all add: iff_def, cases p, auto) + (* Quantifier freeness *) +consts qfree:: "fm \ bool" +recdef qfree "measure size" + "qfree (E p) = False" + "qfree (A p) = False" + "qfree (NOT p) = qfree p" + "qfree (And p q) = (qfree p \ qfree q)" + "qfree (Or p q) = (qfree p \ qfree q)" + "qfree (Imp p q) = (qfree p \ qfree q)" + "qfree (Iff p q) = (qfree p \ qfree q)" + "qfree p = True" + + (* Boundedness and substitution *) +consts + numbound0:: "num \ bool" (* a num is INDEPENDENT of Bound 0 *) + bound0:: "fm \ bool" (* A Formula is independent of Bound 0 *) +primrec + "numbound0 (C c) = True" + "numbound0 (Bound n) = (n>0)" + "numbound0 (CN n c a) = (n\0 \ numbound0 a)" + "numbound0 (Neg a) = numbound0 a" + "numbound0 (Add a b) = (numbound0 a \ numbound0 b)" + "numbound0 (Sub a b) = (numbound0 a \ numbound0 b)" + "numbound0 (Mul i a) = numbound0 a" +lemma numbound0_I: + assumes nb: "numbound0 a" + shows "Inum (b#bs) a = Inum (b'#bs) a" +using nb +by (induct a rule: numbound0.induct,auto simp add: nth_pos2) + +primrec + "bound0 T = True" + "bound0 F = True" + "bound0 (Lt a) = numbound0 a" + "bound0 (Le a) = numbound0 a" + "bound0 (Gt a) = numbound0 a" + "bound0 (Ge a) = numbound0 a" + "bound0 (Eq a) = numbound0 a" + "bound0 (NEq a) = numbound0 a" + "bound0 (NOT p) = bound0 p" + "bound0 (And p q) = (bound0 p \ bound0 q)" + "bound0 (Or p q) = (bound0 p \ bound0 q)" + "bound0 (Imp p q) = ((bound0 p) \ (bound0 q))" + "bound0 (Iff p q) = (bound0 p \ bound0 q)" + "bound0 (E p) = False" + "bound0 (A p) = False" + +lemma bound0_I: + assumes bp: "bound0 p" + shows "Ifm (b#bs) p = Ifm (b'#bs) p" +using bp numbound0_I[where b="b" and bs="bs" and b'="b'"] +by (induct p rule: bound0.induct) (auto simp add: nth_pos2) + +lemma not_qf[simp]: "qfree p \ qfree (not p)" +by (cases p, auto) +lemma not_bn[simp]: "bound0 p \ bound0 (not p)" +by (cases p, auto) + + +lemma conj_qf[simp]: "\qfree p ; qfree q\ \ qfree (conj p q)" +using conj_def by auto +lemma conj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (conj p q)" +using conj_def by auto + +lemma disj_qf[simp]: "\qfree p ; qfree q\ \ qfree (disj p q)" +using disj_def by auto +lemma disj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (disj p q)" +using disj_def by auto + +lemma imp_qf[simp]: "\qfree p ; qfree q\ \ qfree (imp p q)" +using imp_def by (cases "p=F \ q=T",simp_all add: imp_def) +lemma imp_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (imp p q)" +using imp_def by (cases "p=F \ q=T \ p=q",simp_all add: imp_def) + +lemma iff_qf[simp]: "\qfree p ; qfree q\ \ qfree (iff p q)" + by (unfold iff_def,cases "p=q", auto) +lemma iff_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (iff p q)" +using iff_def by (unfold iff_def,cases "p=q", auto) + +consts + decrnum:: "num \ num" + decr :: "fm \ fm" + +recdef decrnum "measure size" + "decrnum (Bound n) = Bound (n - 1)" + "decrnum (Neg a) = Neg (decrnum a)" + "decrnum (Add a b) = Add (decrnum a) (decrnum b)" + "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)" + "decrnum (Mul c a) = Mul c (decrnum a)" + "decrnum (CN n c a) = CN (n - 1) c (decrnum a)" + "decrnum a = a" + +recdef decr "measure size" + "decr (Lt a) = Lt (decrnum a)" + "decr (Le a) = Le (decrnum a)" + "decr (Gt a) = Gt (decrnum a)" + "decr (Ge a) = Ge (decrnum a)" + "decr (Eq a) = Eq (decrnum a)" + "decr (NEq a) = NEq (decrnum a)" + "decr (NOT p) = NOT (decr p)" + "decr (And p q) = conj (decr p) (decr q)" + "decr (Or p q) = disj (decr p) (decr q)" + "decr (Imp p q) = imp (decr p) (decr q)" + "decr (Iff p q) = iff (decr p) (decr q)" + "decr p = p" + +lemma decrnum: assumes nb: "numbound0 t" + shows "Inum (x#bs) t = Inum bs (decrnum t)" + using nb by (induct t rule: decrnum.induct, simp_all add: nth_pos2) + +lemma decr: assumes nb: "bound0 p" + shows "Ifm (x#bs) p = Ifm bs (decr p)" + using nb + by (induct p rule: decr.induct, simp_all add: nth_pos2 decrnum) + +lemma decr_qf: "bound0 p \ qfree (decr p)" +by (induct p, simp_all) + +consts + isatom :: "fm \ bool" (* test for atomicity *) +recdef isatom "measure size" + "isatom T = True" + "isatom F = True" + "isatom (Lt a) = True" + "isatom (Le a) = True" + "isatom (Gt a) = True" + "isatom (Ge a) = True" + "isatom (Eq a) = True" + "isatom (NEq a) = True" + "isatom p = False" + +lemma bound0_qf: "bound0 p \ qfree p" +by (induct p, simp_all) + +constdefs djf:: "('a \ fm) \ 'a \ fm \ fm" + "djf f p q \ (if q=T then T else if q=F then f p else + (let fp = f p in case fp of T \ T | F \ q | _ \ Or (f p) q))" +constdefs evaldjf:: "('a \ fm) \ 'a list \ fm" + "evaldjf f ps \ foldr (djf f) ps F" + +lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)" +by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) +(cases "f p", simp_all add: Let_def djf_def) + + +lemma djf_simps: + "djf f p T = T" + "djf f p F = f p" + "q\T \ q\F \ djf f p q = (let fp = f p in case fp of T \ T | F \ q | _ \ Or (f p) q)" + by (simp_all add: djf_def) + +lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\ p \ set ps. Ifm bs (f p))" + by(induct ps, simp_all add: evaldjf_def djf_Or) + +lemma evaldjf_bound0: + assumes nb: "\ x\ set xs. bound0 (f x)" + shows "bound0 (evaldjf f xs)" + using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) + +lemma evaldjf_qf: + assumes nb: "\ x\ set xs. qfree (f x)" + shows "qfree (evaldjf f xs)" + using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) + +consts disjuncts :: "fm \ fm list" +recdef disjuncts "measure size" + "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)" + "disjuncts F = []" + "disjuncts p = [p]" + +lemma disjuncts: "(\ q\ set (disjuncts p). Ifm bs q) = Ifm bs p" +by(induct p rule: disjuncts.induct, auto) + +lemma disjuncts_nb: "bound0 p \ \ q\ set (disjuncts p). bound0 q" +proof- + assume nb: "bound0 p" + hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto) + thus ?thesis by (simp only: list_all_iff) +qed + +lemma disjuncts_qf: "qfree p \ \ q\ set (disjuncts p). qfree q" +proof- + assume qf: "qfree p" + hence "list_all qfree (disjuncts p)" + by (induct p rule: disjuncts.induct, auto) + thus ?thesis by (simp only: list_all_iff) +qed + +constdefs DJ :: "(fm \ fm) \ fm \ fm" + "DJ f p \ evaldjf f (disjuncts p)" + +lemma DJ: assumes fdj: "\ p q. Ifm bs (f (Or p q)) = Ifm bs (Or (f p) (f q))" + and fF: "f F = F" + shows "Ifm bs (DJ f p) = Ifm bs (f p)" +proof- + have "Ifm bs (DJ f p) = (\ q \ set (disjuncts p). Ifm bs (f q))" + by (simp add: DJ_def evaldjf_ex) + also have "\ = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto) + finally show ?thesis . +qed + +lemma DJ_qf: assumes + fqf: "\ p. qfree p \ qfree (f p)" + shows "\p. qfree p \ qfree (DJ f p) " +proof(clarify) + fix p assume qf: "qfree p" + have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def) + from disjuncts_qf[OF qf] have "\ q\ set (disjuncts p). qfree q" . + with fqf have th':"\ q\ set (disjuncts p). qfree (f q)" by blast + + from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp +qed + +lemma DJ_qe: assumes qe: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" + shows "\ bs p. qfree p \ qfree (DJ qe p) \ (Ifm bs ((DJ qe p)) = Ifm bs (E p))" +proof(clarify) + fix p::fm and bs + assume qf: "qfree p" + from qe have qth: "\ p. qfree p \ qfree (qe p)" by blast + from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto + have "Ifm bs (DJ qe p) = (\ q\ set (disjuncts p). Ifm bs (qe q))" + by (simp add: DJ_def evaldjf_ex) + also have "\ = (\ q \ set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto + also have "\ = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto) + finally show "qfree (DJ qe p) \ Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast +qed + (* Simplification *) +consts + numgcd :: "num \ int" + numgcdh:: "num \ int \ int" + reducecoeffh:: "num \ int \ num" + reducecoeff :: "num \ num" + dvdnumcoeff:: "num \ int \ bool" +consts maxcoeff:: "num \ int" +recdef maxcoeff "measure size" + "maxcoeff (C i) = abs i" + "maxcoeff (CN n c t) = max (abs c) (maxcoeff t)" + "maxcoeff t = 1" + +lemma maxcoeff_pos: "maxcoeff t \ 0" + by (induct t rule: maxcoeff.induct, auto) + +recdef numgcdh "measure size" + "numgcdh (C i) = (\g. zgcd i g)" + "numgcdh (CN n c t) = (\g. zgcd c (numgcdh t g))" + "numgcdh t = (\g. 1)" +defs numgcd_def [code]: "numgcd t \ numgcdh t (maxcoeff t)" + +recdef reducecoeffh "measure size" + "reducecoeffh (C i) = (\ g. C (i div g))" + "reducecoeffh (CN n c t) = (\ g. CN n (c div g) (reducecoeffh t g))" + "reducecoeffh t = (\g. t)" + +defs reducecoeff_def: "reducecoeff t \ + (let g = numgcd t in + if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)" + +recdef dvdnumcoeff "measure size" + "dvdnumcoeff (C i) = (\ g. g dvd i)" + "dvdnumcoeff (CN n c t) = (\ g. g dvd c \ (dvdnumcoeff t g))" + "dvdnumcoeff t = (\g. False)" + +lemma dvdnumcoeff_trans: + assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'" + shows "dvdnumcoeff t g" + using dgt' gdg + by (induct t rule: dvdnumcoeff.induct, simp_all add: gdg zdvd_trans[OF gdg]) + +declare zdvd_trans [trans add] + +lemma natabs0: "(nat (abs x) = 0) = (x = 0)" +by arith + +lemma numgcd0: + assumes g0: "numgcd t = 0" + shows "Inum bs t = 0" + using g0[simplified numgcd_def] + by (induct t rule: numgcdh.induct, auto simp add: zgcd_def gcd_zero natabs0 max_def maxcoeff_pos) + +lemma numgcdh_pos: assumes gp: "g \ 0" shows "numgcdh t g \ 0" + using gp + by (induct t rule: numgcdh.induct, auto simp add: zgcd_def) + +lemma numgcd_pos: "numgcd t \0" + by (simp add: numgcd_def numgcdh_pos maxcoeff_pos) + +lemma reducecoeffh: + assumes gt: "dvdnumcoeff t g" and gp: "g > 0" + shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t" + using gt +proof(induct t rule: reducecoeffh.induct) + case (1 i) hence gd: "g dvd i" by simp + from gp have gnz: "g \ 0" by simp + from prems show ?case by (simp add: real_of_int_div[OF gnz gd]) +next + case (2 n c t) hence gd: "g dvd c" by simp + from gp have gnz: "g \ 0" by simp + from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps) +qed (auto simp add: numgcd_def gp) +consts ismaxcoeff:: "num \ int \ bool" +recdef ismaxcoeff "measure size" + "ismaxcoeff (C i) = (\ x. abs i \ x)" + "ismaxcoeff (CN n c t) = (\x. abs c \ x \ (ismaxcoeff t x))" + "ismaxcoeff t = (\x. True)" + +lemma ismaxcoeff_mono: "ismaxcoeff t c \ c \ c' \ ismaxcoeff t c'" +by (induct t rule: ismaxcoeff.induct, auto) + +lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)" +proof (induct t rule: maxcoeff.induct) + case (2 n c t) + hence H:"ismaxcoeff t (maxcoeff t)" . + have thh: "maxcoeff t \ max (abs c) (maxcoeff t)" by (simp add: le_maxI2) + from ismaxcoeff_mono[OF H thh] show ?case by (simp add: le_maxI1) +qed simp_all + +lemma zgcd_gt1: "zgcd i j > 1 \ ((abs i > 1 \ abs j > 1) \ (abs i = 0 \ abs j > 1) \ (abs i > 1 \ abs j = 0))" + apply (cases "abs i = 0", simp_all add: zgcd_def) + apply (cases "abs j = 0", simp_all) + apply (cases "abs i = 1", simp_all) + apply (cases "abs j = 1", simp_all) + apply auto + done +lemma numgcdh0:"numgcdh t m = 0 \ m =0" + by (induct t rule: numgcdh.induct, auto simp add:zgcd0) + +lemma dvdnumcoeff_aux: + assumes "ismaxcoeff t m" and mp:"m \ 0" and "numgcdh t m > 1" + shows "dvdnumcoeff t (numgcdh t m)" +using prems +proof(induct t rule: numgcdh.induct) + case (2 n c t) + let ?g = "numgcdh t m" + from prems have th:"zgcd c ?g > 1" by simp + from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] + have "(abs c > 1 \ ?g > 1) \ (abs c = 0 \ ?g > 1) \ (abs c > 1 \ ?g = 0)" by simp + moreover {assume "abs c > 1" and gp: "?g > 1" with prems + have th: "dvdnumcoeff t ?g" by simp + have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) + from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)} + moreover {assume "abs c = 0 \ ?g > 1" + with prems have th: "dvdnumcoeff t ?g" by simp + have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) + from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1) + hence ?case by simp } + moreover {assume "abs c > 1" and g0:"?g = 0" + from numgcdh0[OF g0] have "m=0". with prems have ?case by simp } + ultimately show ?case by blast +qed(auto simp add: zgcd_zdvd1) + +lemma dvdnumcoeff_aux2: + assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \ numgcd t > 0" + using prems +proof (simp add: numgcd_def) + let ?mc = "maxcoeff t" + let ?g = "numgcdh t ?mc" + have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff) + have th2: "?mc \ 0" by (rule maxcoeff_pos) + assume H: "numgcdh t ?mc > 1" + from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" . +qed + +lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" +proof- + let ?g = "numgcd t" + have "?g \ 0" by (simp add: numgcd_pos) + hence "?g = 0 \ ?g = 1 \ ?g > 1" by auto + moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} + moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} + moreover { assume g1:"?g > 1" + from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+ + from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis + by (simp add: reducecoeff_def Let_def)} + ultimately show ?thesis by blast +qed + +lemma reducecoeffh_numbound0: "numbound0 t \ numbound0 (reducecoeffh t g)" +by (induct t rule: reducecoeffh.induct, auto) + +lemma reducecoeff_numbound0: "numbound0 t \ numbound0 (reducecoeff t)" +using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def) + +consts + simpnum:: "num \ num" + numadd:: "num \ num \ num" + nummul:: "num \ int \ num" +recdef numadd "measure (\ (t,s). size t + size s)" + "numadd (CN n1 c1 r1,CN n2 c2 r2) = + (if n1=n2 then + (let c = c1 + c2 + in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2)))) + else if n1 \ n2 then (CN n1 c1 (numadd (r1,CN n2 c2 r2))) + else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))" + "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))" + "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" + "numadd (C b1, C b2) = C (b1+b2)" + "numadd (a,b) = Add a b" + +lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)" +apply (induct t s rule: numadd.induct, simp_all add: Let_def) +apply (case_tac "c1+c2 = 0",case_tac "n1 \ n2", simp_all) +apply (case_tac "n1 = n2", simp_all add: ring_simps) +by (simp only: left_distrib[symmetric],simp) + +lemma numadd_nb[simp]: "\ numbound0 t ; numbound0 s\ \ numbound0 (numadd (t,s))" +by (induct t s rule: numadd.induct, auto simp add: Let_def) + +recdef nummul "measure size" + "nummul (C j) = (\ i. C (i*j))" + "nummul (CN n c a) = (\ i. CN n (i*c) (nummul a i))" + "nummul t = (\ i. Mul i t)" + +lemma nummul[simp]: "\ i. Inum bs (nummul t i) = Inum bs (Mul i t)" +by (induct t rule: nummul.induct, auto simp add: ring_simps) + +lemma nummul_nb[simp]: "\ i. numbound0 t \ numbound0 (nummul t i)" +by (induct t rule: nummul.induct, auto ) + +constdefs numneg :: "num \ num" + "numneg t \ nummul t (- 1)" + +constdefs numsub :: "num \ num \ num" + "numsub s t \ (if s = t then C 0 else numadd (s,numneg t))" + +lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)" +using numneg_def by simp + +lemma numneg_nb[simp]: "numbound0 t \ numbound0 (numneg t)" +using numneg_def by simp + +lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)" +using numsub_def by simp + +lemma numsub_nb[simp]: "\ numbound0 t ; numbound0 s\ \ numbound0 (numsub t s)" +using numsub_def by simp + +recdef simpnum "measure size" + "simpnum (C j) = C j" + "simpnum (Bound n) = CN n 1 (C 0)" + "simpnum (Neg t) = numneg (simpnum t)" + "simpnum (Add t s) = numadd (simpnum t,simpnum s)" + "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)" + "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)" + "simpnum (CN n c t) = (if c = 0 then simpnum t else numadd (CN n c (C 0),simpnum t))" + +lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t" +by (induct t rule: simpnum.induct, auto simp add: numneg numadd numsub nummul) + +lemma simpnum_numbound0[simp]: + "numbound0 t \ numbound0 (simpnum t)" +by (induct t rule: simpnum.induct, auto) + +consts nozerocoeff:: "num \ bool" +recdef nozerocoeff "measure size" + "nozerocoeff (C c) = True" + "nozerocoeff (CN n c t) = (c\0 \ nozerocoeff t)" + "nozerocoeff t = True" + +lemma numadd_nz : "nozerocoeff a \ nozerocoeff b \ nozerocoeff (numadd (a,b))" +by (induct a b rule: numadd.induct,auto simp add: Let_def) + +lemma nummul_nz : "\ i. i\0 \ nozerocoeff a \ nozerocoeff (nummul a i)" +by (induct a rule: nummul.induct,auto simp add: Let_def numadd_nz) + +lemma numneg_nz : "nozerocoeff a \ nozerocoeff (numneg a)" +by (simp add: numneg_def nummul_nz) + +lemma numsub_nz: "nozerocoeff a \ nozerocoeff b \ nozerocoeff (numsub a b)" +by (simp add: numsub_def numneg_nz numadd_nz) + +lemma simpnum_nz: "nozerocoeff (simpnum t)" +by(induct t rule: simpnum.induct, auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz) + +lemma maxcoeff_nz: "nozerocoeff t \ maxcoeff t = 0 \ t = C 0" +proof (induct t rule: maxcoeff.induct) + case (2 n c t) + hence cnz: "c \0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+ + have "max (abs c) (maxcoeff t) \ abs c" by (simp add: le_maxI1) + with cnz have "max (abs c) (maxcoeff t) > 0" by arith + with prems show ?case by simp +qed auto + +lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0" +proof- + from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def) + from numgcdh0[OF th] have th:"maxcoeff t = 0" . + from maxcoeff_nz[OF nz th] show ?thesis . +qed + +constdefs simp_num_pair:: "(num \ int) \ num \ int" + "simp_num_pair \ (\ (t,n). (if n = 0 then (C 0, 0) else + (let t' = simpnum t ; g = numgcd t' in + if g > 1 then (let g' = zgcd n g in + if g' = 1 then (t',n) + else (reducecoeffh t' g', n div g')) + else (t',n))))" + +lemma simp_num_pair_ci: + shows "((\ (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((\ (t,n). Inum bs t / real n) (t,n))" + (is "?lhs = ?rhs") +proof- + let ?t' = "simpnum t" + let ?g = "numgcd ?t'" + let ?g' = "zgcd n ?g" + {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)} + moreover + { assume nnz: "n \ 0" + {assume "\ ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def simpnum_ci)} + moreover + {assume g1:"?g>1" hence g0: "?g > 0" by simp + from zgcd0 g1 nnz have gp0: "?g' \ 0" by simp + hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith + hence "?g'= 1 \ ?g' > 1" by arith + moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def simpnum_ci)} + moreover {assume g'1:"?g'>1" + from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" .. + let ?tt = "reducecoeffh ?t' ?g'" + let ?t = "Inum bs ?tt" + have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) + have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1) + have gpdgp: "?g' dvd ?g'" by simp + from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] + have th2:"real ?g' * ?t = Inum bs ?t'" by simp + from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def) + also have "\ = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp + also have "\ = (Inum bs ?t' / real n)" + using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp + finally have "?lhs = Inum bs t / real n" by (simp add: simpnum_ci) + then have ?thesis using prems by (simp add: simp_num_pair_def)} + ultimately have ?thesis by blast} + ultimately have ?thesis by blast} + ultimately show ?thesis by blast +qed + +lemma simp_num_pair_l: assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')" + shows "numbound0 t' \ n' >0" +proof- + let ?t' = "simpnum t" + let ?g = "numgcd ?t'" + let ?g' = "zgcd n ?g" + {assume nz: "n = 0" hence ?thesis using prems by (simp add: Let_def simp_num_pair_def)} + moreover + { assume nnz: "n \ 0" + {assume "\ ?g > 1" hence ?thesis using prems by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)} + moreover + {assume g1:"?g>1" hence g0: "?g > 0" by simp + from zgcd0 g1 nnz have gp0: "?g' \ 0" by simp + hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith + hence "?g'= 1 \ ?g' > 1" by arith + moreover {assume "?g'=1" hence ?thesis using prems + by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)} + moreover {assume g'1:"?g'>1" + have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) + have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1) + have gpdgp: "?g' dvd ?g'" by simp + from zdvd_imp_le[OF gpdd np] have g'n: "?g' \ n" . + from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]] + have "n div ?g' >0" by simp + hence ?thesis using prems + by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0 simpnum_numbound0)} + ultimately have ?thesis by blast} + ultimately have ?thesis by blast} + ultimately show ?thesis by blast +qed + +consts simpfm :: "fm \ fm" +recdef simpfm "measure fmsize" + "simpfm (And p q) = conj (simpfm p) (simpfm q)" + "simpfm (Or p q) = disj (simpfm p) (simpfm q)" + "simpfm (Imp p q) = imp (simpfm p) (simpfm q)" + "simpfm (Iff p q) = iff (simpfm p) (simpfm q)" + "simpfm (NOT p) = not (simpfm p)" + "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \ if (v < 0) then T else F + | _ \ Lt a')" + "simpfm (Le a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Le a')" + "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \ if (v > 0) then T else F | _ \ Gt a')" + "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Ge a')" + "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \ if (v = 0) then T else F | _ \ Eq a')" + "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ NEq a')" + "simpfm p = p" +lemma simpfm: "Ifm bs (simpfm p) = Ifm bs p" +proof(induct p rule: simpfm.induct) + case (6 a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa + by (cases ?sa, simp_all add: Let_def)} + ultimately show ?case by blast +next + case (7 a) let ?sa = "simpnum a" + from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa + by (cases ?sa, simp_all add: Let_def)} + ultimately show ?case by blast +next + case (8 a) let ?sa = "simpnum a" + from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa + by (cases ?sa, simp_all add: Let_def)} + ultimately show ?case by blast +next + case (9 a) let ?sa = "simpnum a" + from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa + by (cases ?sa, simp_all add: Let_def)} + ultimately show ?case by blast +next + case (10 a) let ?sa = "simpnum a" + from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa + by (cases ?sa, simp_all add: Let_def)} + ultimately show ?case by blast +next + case (11 a) let ?sa = "simpnum a" + from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp + {fix v assume "?sa = C v" hence ?case using sa by simp } + moreover {assume "\ (\ v. ?sa = C v)" hence ?case using sa + by (cases ?sa, simp_all add: Let_def)} + ultimately show ?case by blast +qed (induct p rule: simpfm.induct, simp_all add: conj disj imp iff not) + + +lemma simpfm_bound0: "bound0 p \ bound0 (simpfm p)" +proof(induct p rule: simpfm.induct) + case (6 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def) +next + case (7 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def) +next + case (8 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def) +next + case (9 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def) +next + case (10 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def) +next + case (11 a) hence nb: "numbound0 a" by simp + hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) + thus ?case by (cases "simpnum a", auto simp add: Let_def) +qed(auto simp add: disj_def imp_def iff_def conj_def not_bn) + +lemma simpfm_qf: "qfree p \ qfree (simpfm p)" +by (induct p rule: simpfm.induct, auto simp add: disj_qf imp_qf iff_qf conj_qf not_qf Let_def) + (case_tac "simpnum a",auto)+ + +consts prep :: "fm \ fm" +recdef prep "measure fmsize" + "prep (E T) = T" + "prep (E F) = F" + "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))" + "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))" + "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" + "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))" + "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))" + "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))" + "prep (E p) = E (prep p)" + "prep (A (And p q)) = conj (prep (A p)) (prep (A q))" + "prep (A p) = prep (NOT (E (NOT p)))" + "prep (NOT (NOT p)) = prep p" + "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))" + "prep (NOT (A p)) = prep (E (NOT p))" + "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))" + "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))" + "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))" + "prep (NOT p) = not (prep p)" + "prep (Or p q) = disj (prep p) (prep q)" + "prep (And p q) = conj (prep p) (prep q)" + "prep (Imp p q) = prep (Or (NOT p) q)" + "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))" + "prep p = p" +(hints simp add: fmsize_pos) +lemma prep: "\ bs. Ifm bs (prep p) = Ifm bs p" +by (induct p rule: prep.induct, auto) + + (* Generic quantifier elimination *) +consts qelim :: "fm \ (fm \ fm) \ fm" +recdef qelim "measure fmsize" + "qelim (E p) = (\ qe. DJ qe (qelim p qe))" + "qelim (A p) = (\ qe. not (qe ((qelim (NOT p) qe))))" + "qelim (NOT p) = (\ qe. not (qelim p qe))" + "qelim (And p q) = (\ qe. conj (qelim p qe) (qelim q qe))" + "qelim (Or p q) = (\ qe. disj (qelim p qe) (qelim q qe))" + "qelim (Imp p q) = (\ qe. imp (qelim p qe) (qelim q qe))" + "qelim (Iff p q) = (\ qe. iff (qelim p qe) (qelim q qe))" + "qelim p = (\ y. simpfm p)" + +lemma qelim_ci: + assumes qe_inv: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" + shows "\ bs. qfree (qelim p qe) \ (Ifm bs (qelim p qe) = Ifm bs p)" +using qe_inv DJ_qe[OF qe_inv] +by(induct p rule: qelim.induct) +(auto simp add: not disj conj iff imp not_qf disj_qf conj_qf imp_qf iff_qf + simpfm simpfm_qf simp del: simpfm.simps) + +consts + plusinf:: "fm \ fm" (* Virtual substitution of +\*) + minusinf:: "fm \ fm" (* Virtual substitution of -\*) +recdef minusinf "measure size" + "minusinf (And p q) = conj (minusinf p) (minusinf q)" + "minusinf (Or p q) = disj (minusinf p) (minusinf q)" + "minusinf (Eq (CN 0 c e)) = F" + "minusinf (NEq (CN 0 c e)) = T" + "minusinf (Lt (CN 0 c e)) = T" + "minusinf (Le (CN 0 c e)) = T" + "minusinf (Gt (CN 0 c e)) = F" + "minusinf (Ge (CN 0 c e)) = F" + "minusinf p = p" + +recdef plusinf "measure size" + "plusinf (And p q) = conj (plusinf p) (plusinf q)" + "plusinf (Or p q) = disj (plusinf p) (plusinf q)" + "plusinf (Eq (CN 0 c e)) = F" + "plusinf (NEq (CN 0 c e)) = T" + "plusinf (Lt (CN 0 c e)) = F" + "plusinf (Le (CN 0 c e)) = F" + "plusinf (Gt (CN 0 c e)) = T" + "plusinf (Ge (CN 0 c e)) = T" + "plusinf p = p" + +consts + isrlfm :: "fm \ bool" (* Linearity test for fm *) +recdef isrlfm "measure size" + "isrlfm (And p q) = (isrlfm p \ isrlfm q)" + "isrlfm (Or p q) = (isrlfm p \ isrlfm q)" + "isrlfm (Eq (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm (NEq (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm (Lt (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm (Le (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm (Gt (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm (Ge (CN 0 c e)) = (c>0 \ numbound0 e)" + "isrlfm p = (isatom p \ (bound0 p))" + + (* splits the bounded from the unbounded part*) +consts rsplit0 :: "num \ int \ num" +recdef rsplit0 "measure num_size" + "rsplit0 (Bound 0) = (1,C 0)" + "rsplit0 (Add a b) = (let (ca,ta) = rsplit0 a ; (cb,tb) = rsplit0 b + in (ca+cb, Add ta tb))" + "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))" + "rsplit0 (Neg a) = (let (c,t) = rsplit0 a in (-c,Neg t))" + "rsplit0 (Mul c a) = (let (ca,ta) = rsplit0 a in (c*ca,Mul c ta))" + "rsplit0 (CN 0 c a) = (let (ca,ta) = rsplit0 a in (c+ca,ta))" + "rsplit0 (CN n c a) = (let (ca,ta) = rsplit0 a in (ca,CN n c ta))" + "rsplit0 t = (0,t)" +lemma rsplit0: + shows "Inum bs ((split (CN 0)) (rsplit0 t)) = Inum bs t \ numbound0 (snd (rsplit0 t))" +proof (induct t rule: rsplit0.induct) + case (2 a b) + let ?sa = "rsplit0 a" let ?sb = "rsplit0 b" + let ?ca = "fst ?sa" let ?cb = "fst ?sb" + let ?ta = "snd ?sa" let ?tb = "snd ?sb" + from prems have nb: "numbound0 (snd(rsplit0 (Add a b)))" + by(cases "rsplit0 a",auto simp add: Let_def split_def) + have "Inum bs ((split (CN 0)) (rsplit0 (Add a b))) = + Inum bs ((split (CN 0)) ?sa)+Inum bs ((split (CN 0)) ?sb)" + by (simp add: Let_def split_def ring_simps) + also have "\ = Inum bs a + Inum bs b" using prems by (cases "rsplit0 a", simp_all) + finally show ?case using nb by simp +qed(auto simp add: Let_def split_def ring_simps , simp add: right_distrib[symmetric]) + + (* Linearize a formula*) +definition + lt :: "int \ num \ fm" +where + "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) + else (Gt (CN 0 (-c) (Neg t))))" + +definition + le :: "int \ num \ fm" +where + "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) + else (Ge (CN 0 (-c) (Neg t))))" + +definition + gt :: "int \ num \ fm" +where + "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) + else (Lt (CN 0 (-c) (Neg t))))" + +definition + ge :: "int \ num \ fm" +where + "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) + else (Le (CN 0 (-c) (Neg t))))" + +definition + eq :: "int \ num \ fm" +where + "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) + else (Eq (CN 0 (-c) (Neg t))))" + +definition + neq :: "int \ num \ fm" +where + "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) + else (NEq (CN 0 (-c) (Neg t))))" + +lemma lt: "numnoabs t \ Ifm bs (split lt (rsplit0 t)) = Ifm bs (Lt t) \ isrlfm (split lt (rsplit0 t))" +using rsplit0[where bs = "bs" and t="t"] +by (auto simp add: lt_def split_def,cases "snd(rsplit0 t)",auto,case_tac "nat",auto) + +lemma le: "numnoabs t \ Ifm bs (split le (rsplit0 t)) = Ifm bs (Le t) \ isrlfm (split le (rsplit0 t))" +using rsplit0[where bs = "bs" and t="t"] +by (auto simp add: le_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) + +lemma gt: "numnoabs t \ Ifm bs (split gt (rsplit0 t)) = Ifm bs (Gt t) \ isrlfm (split gt (rsplit0 t))" +using rsplit0[where bs = "bs" and t="t"] +by (auto simp add: gt_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) + +lemma ge: "numnoabs t \ Ifm bs (split ge (rsplit0 t)) = Ifm bs (Ge t) \ isrlfm (split ge (rsplit0 t))" +using rsplit0[where bs = "bs" and t="t"] +by (auto simp add: ge_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) + +lemma eq: "numnoabs t \ Ifm bs (split eq (rsplit0 t)) = Ifm bs (Eq t) \ isrlfm (split eq (rsplit0 t))" +using rsplit0[where bs = "bs" and t="t"] +by (auto simp add: eq_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) + +lemma neq: "numnoabs t \ Ifm bs (split neq (rsplit0 t)) = Ifm bs (NEq t) \ isrlfm (split neq (rsplit0 t))" +using rsplit0[where bs = "bs" and t="t"] +by (auto simp add: neq_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) + +lemma conj_lin: "isrlfm p \ isrlfm q \ isrlfm (conj p q)" +by (auto simp add: conj_def) +lemma disj_lin: "isrlfm p \ isrlfm q \ isrlfm (disj p q)" +by (auto simp add: disj_def) + +consts rlfm :: "fm \ fm" +recdef rlfm "measure fmsize" + "rlfm (And p q) = conj (rlfm p) (rlfm q)" + "rlfm (Or p q) = disj (rlfm p) (rlfm q)" + "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)" + "rlfm (Iff p q) = disj (conj (rlfm p) (rlfm q)) (conj (rlfm (NOT p)) (rlfm (NOT q)))" + "rlfm (Lt a) = split lt (rsplit0 a)" + "rlfm (Le a) = split le (rsplit0 a)" + "rlfm (Gt a) = split gt (rsplit0 a)" + "rlfm (Ge a) = split ge (rsplit0 a)" + "rlfm (Eq a) = split eq (rsplit0 a)" + "rlfm (NEq a) = split neq (rsplit0 a)" + "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))" + "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))" + "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))" + "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))" + "rlfm (NOT (NOT p)) = rlfm p" + "rlfm (NOT T) = F" + "rlfm (NOT F) = T" + "rlfm (NOT (Lt a)) = rlfm (Ge a)" + "rlfm (NOT (Le a)) = rlfm (Gt a)" + "rlfm (NOT (Gt a)) = rlfm (Le a)" + "rlfm (NOT (Ge a)) = rlfm (Lt a)" + "rlfm (NOT (Eq a)) = rlfm (NEq a)" + "rlfm (NOT (NEq a)) = rlfm (Eq a)" + "rlfm p = p" (hints simp add: fmsize_pos) + +lemma rlfm_I: + assumes qfp: "qfree p" + shows "(Ifm bs (rlfm p) = Ifm bs p) \ isrlfm (rlfm p)" + using qfp +by (induct p rule: rlfm.induct, auto simp add: lt le gt ge eq neq conj disj conj_lin disj_lin) + + (* Operations needed for Ferrante and Rackoff *) +lemma rminusinf_inf: + assumes lp: "isrlfm p" + shows "\ z. \ x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\ z. \ x. ?P z x p") +using lp +proof (induct p rule: minusinf.induct) + case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto +next + case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto +next + case (3 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + hence "real c * x + ?e \ 0" by simp + with xz have "?P ?z x (Eq (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp + thus ?case by blast +next + case (4 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + hence "real c * x + ?e \ 0" by simp + with xz have "?P ?z x (NEq (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp + thus ?case by blast +next + case (5 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + with xz have "?P ?z x (Lt (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp + thus ?case by blast +next + case (6 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + with xz have "?P ?z x (Le (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (Le (CN 0 c e))" by simp + thus ?case by blast +next + case (7 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + with xz have "?P ?z x (Gt (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp + thus ?case by blast +next + case (8 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x < ?z" + hence "(real c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) + hence "real c * x + ?e < 0" by arith + with xz have "?P ?z x (Ge (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp + thus ?case by blast +qed simp_all + +lemma rplusinf_inf: + assumes lp: "isrlfm p" + shows "\ z. \ x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\ z. \ x. ?P z x p") +using lp +proof (induct p rule: isrlfm.induct) + case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto +next + case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto +next + case (3 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + hence "real c * x + ?e \ 0" by simp + with xz have "?P ?z x (Eq (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp + thus ?case by blast +next + case (4 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + hence "real c * x + ?e \ 0" by simp + with xz have "?P ?z x (NEq (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp + thus ?case by blast +next + case (5 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + with xz have "?P ?z x (Lt (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp + thus ?case by blast +next + case (6 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + with xz have "?P ?z x (Le (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (Le (CN 0 c e))" by simp + thus ?case by blast +next + case (7 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + with xz have "?P ?z x (Gt (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp + thus ?case by blast +next + case (8 c e) + from prems have nb: "numbound0 e" by simp + from prems have cp: "real c > 0" by simp + fix a + let ?e="Inum (a#bs) e" + let ?z = "(- ?e) / real c" + {fix x + assume xz: "x > ?z" + with mult_strict_right_mono [OF xz cp] cp + have "(real c * x > - ?e)" by (simp add: mult_ac) + hence "real c * x + ?e > 0" by arith + with xz have "?P ?z x (Ge (CN 0 c e))" + using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } + hence "\ x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp + thus ?case by blast +qed simp_all + +lemma rminusinf_bound0: + assumes lp: "isrlfm p" + shows "bound0 (minusinf p)" + using lp + by (induct p rule: minusinf.induct) simp_all + +lemma rplusinf_bound0: + assumes lp: "isrlfm p" + shows "bound0 (plusinf p)" + using lp + by (induct p rule: plusinf.induct) simp_all + +lemma rminusinf_ex: + assumes lp: "isrlfm p" + and ex: "Ifm (a#bs) (minusinf p)" + shows "\ x. Ifm (x#bs) p" +proof- + from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex + have th: "\ x. Ifm (x#bs) (minusinf p)" by auto + from rminusinf_inf[OF lp, where bs="bs"] + obtain z where z_def: "\x x. Ifm (x#bs) p" +proof- + from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex + have th: "\ x. Ifm (x#bs) (plusinf p)" by auto + from rplusinf_inf[OF lp, where bs="bs"] + obtain z where z_def: "\x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast + from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp + moreover have "z + 1 > z" by simp + ultimately show ?thesis using z_def by auto +qed + +consts + uset:: "fm \ (num \ int) list" + usubst :: "fm \ (num \ int) \ fm " +recdef uset "measure size" + "uset (And p q) = (uset p @ uset q)" + "uset (Or p q) = (uset p @ uset q)" + "uset (Eq (CN 0 c e)) = [(Neg e,c)]" + "uset (NEq (CN 0 c e)) = [(Neg e,c)]" + "uset (Lt (CN 0 c e)) = [(Neg e,c)]" + "uset (Le (CN 0 c e)) = [(Neg e,c)]" + "uset (Gt (CN 0 c e)) = [(Neg e,c)]" + "uset (Ge (CN 0 c e)) = [(Neg e,c)]" + "uset p = []" +recdef usubst "measure size" + "usubst (And p q) = (\ (t,n). And (usubst p (t,n)) (usubst q (t,n)))" + "usubst (Or p q) = (\ (t,n). Or (usubst p (t,n)) (usubst q (t,n)))" + "usubst (Eq (CN 0 c e)) = (\ (t,n). Eq (Add (Mul c t) (Mul n e)))" + "usubst (NEq (CN 0 c e)) = (\ (t,n). NEq (Add (Mul c t) (Mul n e)))" + "usubst (Lt (CN 0 c e)) = (\ (t,n). Lt (Add (Mul c t) (Mul n e)))" + "usubst (Le (CN 0 c e)) = (\ (t,n). Le (Add (Mul c t) (Mul n e)))" + "usubst (Gt (CN 0 c e)) = (\ (t,n). Gt (Add (Mul c t) (Mul n e)))" + "usubst (Ge (CN 0 c e)) = (\ (t,n). Ge (Add (Mul c t) (Mul n e)))" + "usubst p = (\ (t,n). p)" + +lemma usubst_I: assumes lp: "isrlfm p" + and np: "real n > 0" and nbt: "numbound0 t" + shows "(Ifm (x#bs) (usubst p (t,n)) = Ifm (((Inum (x#bs) t)/(real n))#bs) p) \ bound0 (usubst p (t,n))" (is "(?I x (usubst p (t,n)) = ?I ?u p) \ ?B p" is "(_ = ?I (?t/?n) p) \ _" is "(_ = ?I (?N x t /_) p) \ _") + using lp +proof(induct p rule: usubst.induct) + case (5 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + have "?I ?u (Lt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) < 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)" + by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) < 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +next + case (6 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" + by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) \ 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +next + case (7 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)" + by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) > 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +next + case (8 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" + by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) \ 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +next + case (3 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + from np have np: "real n \ 0" by simp + have "?I ?u (Eq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) = 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)" + by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) = 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +next + case (4 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ + from np have np: "real n \ 0" by simp + have "?I ?u (NEq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" + using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp + also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" + by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + and b="0", simplified divide_zero_left]) (simp only: ring_simps) + also have "\ = (real c *?t + ?n* (?N x e) \ 0)" + using np by simp + finally show ?case using nbt nb by (simp add: ring_simps) +qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"] nth_pos2) + +lemma uset_l: + assumes lp: "isrlfm p" + shows "\ (t,k) \ set (uset p). numbound0 t \ k >0" +using lp +by(induct p rule: uset.induct,auto) + +lemma rminusinf_uset: + assumes lp: "isrlfm p" + and nmi: "\ (Ifm (a#bs) (minusinf p))" (is "\ (Ifm (a#bs) (?M p))") + and ex: "Ifm (x#bs) p" (is "?I x p") + shows "\ (s,m) \ set (uset p). x \ Inum (a#bs) s / real m" (is "\ (s,m) \ ?U p. x \ ?N a s / real m") +proof- + have "\ (s,m) \ set (uset p). real m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real m *x \ ?N a s") + using lp nmi ex + by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) + then obtain s m where smU: "(s,m) \ set (uset p)" and mx: "real m * x \ ?N a s" by blast + from uset_l[OF lp] smU have mp: "real m > 0" by auto + from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" + by (auto simp add: mult_commute) + thus ?thesis using smU by auto +qed + +lemma rplusinf_uset: + assumes lp: "isrlfm p" + and nmi: "\ (Ifm (a#bs) (plusinf p))" (is "\ (Ifm (a#bs) (?M p))") + and ex: "Ifm (x#bs) p" (is "?I x p") + shows "\ (s,m) \ set (uset p). x \ Inum (a#bs) s / real m" (is "\ (s,m) \ ?U p. x \ ?N a s / real m") +proof- + have "\ (s,m) \ set (uset p). real m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real m *x \ ?N a s") + using lp nmi ex + by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) + then obtain s m where smU: "(s,m) \ set (uset p)" and mx: "real m * x \ ?N a s" by blast + from uset_l[OF lp] smU have mp: "real m > 0" by auto + from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" + by (auto simp add: mult_commute) + thus ?thesis using smU by auto +qed + +lemma lin_dense: + assumes lp: "isrlfm p" + and noS: "\ t. l < t \ t< u \ t \ (\ (t,n). Inum (x#bs) t / real n) ` set (uset p)" + (is "\ t. _ \ _ \ t \ (\ (t,n). ?N x t / real n ) ` (?U p)") + and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p" + and ly: "l < y" and yu: "y < u" + shows "Ifm (y#bs) p" +using lp px noS +proof (induct p rule: isrlfm.induct) + case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ + from prems have "x * real c + ?N x e < 0" by (simp add: ring_simps) + hence pxc: "x < (- ?N x e) / real c" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"]) + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with ly yu have yne: "y \ - ?N x e / real c" by auto + hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto + moreover {assume y: "y < (-?N x e)/ real c" + hence "y * real c < - ?N x e" + by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) + hence "real c * y + ?N x e < 0" by (simp add: ring_simps) + hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} + moreover {assume y: "y > (- ?N x e) / real c" + with yu have eu: "u > (- ?N x e) / real c" by auto + with noSc ly yu have "(- ?N x e) / real c \ l" by (cases "(- ?N x e) / real c > l", auto) + with lx pxc have "False" by auto + hence ?case by simp } + ultimately show ?case by blast +next + case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp + + from prems have "x * real c + ?N x e \ 0" by (simp add: ring_simps) + hence pxc: "x \ (- ?N x e) / real c" + by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"]) + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with ly yu have yne: "y \ - ?N x e / real c" by auto + hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto + moreover {assume y: "y < (-?N x e)/ real c" + hence "y * real c < - ?N x e" + by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) + hence "real c * y + ?N x e < 0" by (simp add: ring_simps) + hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} + moreover {assume y: "y > (- ?N x e) / real c" + with yu have eu: "u > (- ?N x e) / real c" by auto + with noSc ly yu have "(- ?N x e) / real c \ l" by (cases "(- ?N x e) / real c > l", auto) + with lx pxc have "False" by auto + hence ?case by simp } + ultimately show ?case by blast +next + case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ + from prems have "x * real c + ?N x e > 0" by (simp add: ring_simps) + hence pxc: "x > (- ?N x e) / real c" + by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"]) + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with ly yu have yne: "y \ - ?N x e / real c" by auto + hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto + moreover {assume y: "y > (-?N x e)/ real c" + hence "y * real c > - ?N x e" + by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) + hence "real c * y + ?N x e > 0" by (simp add: ring_simps) + hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} + moreover {assume y: "y < (- ?N x e) / real c" + with ly have eu: "l < (- ?N x e) / real c" by auto + with noSc ly yu have "(- ?N x e) / real c \ u" by (cases "(- ?N x e) / real c > l", auto) + with xu pxc have "False" by auto + hence ?case by simp } + ultimately show ?case by blast +next + case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ + from prems have "x * real c + ?N x e \ 0" by (simp add: ring_simps) + hence pxc: "x \ (- ?N x e) / real c" + by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"]) + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with ly yu have yne: "y \ - ?N x e / real c" by auto + hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto + moreover {assume y: "y > (-?N x e)/ real c" + hence "y * real c > - ?N x e" + by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) + hence "real c * y + ?N x e > 0" by (simp add: ring_simps) + hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} + moreover {assume y: "y < (- ?N x e) / real c" + with ly have eu: "l < (- ?N x e) / real c" by auto + with noSc ly yu have "(- ?N x e) / real c \ u" by (cases "(- ?N x e) / real c > l", auto) + with xu pxc have "False" by auto + hence ?case by simp } + ultimately show ?case by blast +next + case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ + from cp have cnz: "real c \ 0" by simp + from prems have "x * real c + ?N x e = 0" by (simp add: ring_simps) + hence pxc: "x = (- ?N x e) / real c" + by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"]) + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with lx xu have yne: "x \ - ?N x e / real c" by auto + with pxc show ?case by simp +next + case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ + from cp have cnz: "real c \ 0" by simp + from prems have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto + with ly yu have yne: "y \ - ?N x e / real c" by auto + hence "y* real c \ -?N x e" + by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp + hence "y* real c + ?N x e \ 0" by (simp add: ring_simps) + thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] + by (simp add: ring_simps) +qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="y" and b'="x"]) + +lemma finite_set_intervals: + assumes px: "P (x::real)" + and lx: "l \ x" and xu: "x \ u" + and linS: "l\ S" and uinS: "u \ S" + and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" + shows "\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a \ x \ x \ b \ P x" +proof- + let ?Mx = "{y. y\ S \ y \ x}" + let ?xM = "{y. y\ S \ x \ y}" + let ?a = "Max ?Mx" + let ?b = "Min ?xM" + have MxS: "?Mx \ S" by blast + hence fMx: "finite ?Mx" using fS finite_subset by auto + from lx linS have linMx: "l \ ?Mx" by blast + hence Mxne: "?Mx \ {}" by blast + have xMS: "?xM \ S" by blast + hence fxM: "finite ?xM" using fS finite_subset by auto + from xu uinS have linxM: "u \ ?xM" by blast + hence xMne: "?xM \ {}" by blast + have ax:"?a \ x" using Mxne fMx by auto + have xb:"x \ ?b" using xMne fxM by auto + have "?a \ ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \ S" using MxS by blast + have "?b \ ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \ S" using xMS by blast + have noy:"\ y. ?a < y \ y < ?b \ y \ S" + proof(clarsimp) + fix y + assume ay: "?a < y" and yb: "y < ?b" and yS: "y \ S" + from yS have "y\ ?Mx \ y\ ?xM" by auto + moreover {assume "y \ ?Mx" hence "y \ ?a" using Mxne fMx by auto with ay have "False" by simp} + moreover {assume "y \ ?xM" hence "y \ ?b" using xMne fxM by auto with yb have "False" by simp} + ultimately show "False" by blast + qed + from ainS binS noy ax xb px show ?thesis by blast +qed + +lemma finite_set_intervals2: + assumes px: "P (x::real)" + and lx: "l \ x" and xu: "x \ u" + and linS: "l\ S" and uinS: "u \ S" + and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" + shows "(\ s\ S. P s) \ (\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a < x \ x < b \ P x)" +proof- + from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su] + obtain a and b where + as: "a\ S" and bs: "b\ S" and noS:"\y. a < y \ y < b \ y \ S" and axb: "a \ x \ x \ b \ P x" by auto + from axb have "x= a \ x= b \ (a < x \ x < b)" by auto + thus ?thesis using px as bs noS by blast +qed + +lemma rinf_uset: + assumes lp: "isrlfm p" + and nmi: "\ (Ifm (x#bs) (minusinf p))" (is "\ (Ifm (x#bs) (?M p))") + and npi: "\ (Ifm (x#bs) (plusinf p))" (is "\ (Ifm (x#bs) (?P p))") + and ex: "\ x. Ifm (x#bs) p" (is "\ x. ?I x p") + shows "\ (l,n) \ set (uset p). \ (s,m) \ set (uset p). ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p" +proof- + let ?N = "\ x t. Inum (x#bs) t" + let ?U = "set (uset p)" + from ex obtain a where pa: "?I a p" by blast + from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi + have nmi': "\ (?I a (?M p))" by simp + from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi + have npi': "\ (?I a (?P p))" by simp + have "\ (l,n) \ set (uset p). \ (s,m) \ set (uset p). ?I ((?N a l/real n + ?N a s /real m) / 2) p" + proof- + let ?M = "(\ (t,c). ?N a t / real c) ` ?U" + have fM: "finite ?M" by auto + from rminusinf_uset[OF lp nmi pa] rplusinf_uset[OF lp npi pa] + have "\ (l,n) \ set (uset p). \ (s,m) \ set (uset p). a \ ?N x l / real n \ a \ ?N x s / real m" by blast + then obtain "t" "n" "s" "m" where + tnU: "(t,n) \ ?U" and smU: "(s,m) \ ?U" + and xs1: "a \ ?N x s / real m" and tx1: "a \ ?N x t / real n" by blast + from uset_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 have xs: "a \ ?N a s / real m" and tx: "a \ ?N a t / real n" by auto + from tnU have Mne: "?M \ {}" by auto + hence Une: "?U \ {}" by simp + let ?l = "Min ?M" + let ?u = "Max ?M" + have linM: "?l \ ?M" using fM Mne by simp + have uinM: "?u \ ?M" using fM Mne by simp + have tnM: "?N a t / real n \ ?M" using tnU by auto + have smM: "?N a s / real m \ ?M" using smU by auto + have lM: "\ t\ ?M. ?l \ t" using Mne fM by auto + have Mu: "\ t\ ?M. t \ ?u" using Mne fM by auto + have "?l \ ?N a t / real n" using tnM Mne by simp hence lx: "?l \ a" using tx by simp + have "?N a s / real m \ ?u" using smM Mne by simp hence xu: "a \ ?u" using xs by simp + from finite_set_intervals2[where P="\ x. ?I x p",OF pa lx xu linM uinM fM lM Mu] + have "(\ s\ ?M. ?I s p) \ + (\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p)" . + moreover { fix u assume um: "u\ ?M" and pu: "?I u p" + hence "\ (tu,nu) \ ?U. u = ?N a tu / real nu" by auto + then obtain "tu" "nu" where tuU: "(tu,nu) \ ?U" and tuu:"u= ?N a tu / real nu" by blast + have "(u + u) / 2 = u" by auto with pu tuu + have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" by simp + with tuU have ?thesis by blast} + moreover{ + assume "\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p" + then obtain t1 and t2 where t1M: "t1 \ ?M" and t2M: "t2\ ?M" + and noM: "\ y. t1 < y \ y < t2 \ y \ ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p" + by blast + from t1M have "\ (t1u,t1n) \ ?U. t1 = ?N a t1u / real t1n" by auto + then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \ ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast + from t2M have "\ (t2u,t2n) \ ?U. t2 = ?N a t2u / real t2n" by auto + then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \ ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast + from t1x xt2 have t1t2: "t1 < t2" by simp + let ?u = "(t1 + t2) / 2" + from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto + from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" . + with t1uU t2uU t1u t2u have ?thesis by blast} + ultimately show ?thesis by blast + qed + then obtain "l" "n" "s" "m" where lnU: "(l,n) \ ?U" and smU:"(s,m) \ ?U" + and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast + from lnU smU uset_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto + from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] + numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu + have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp + with lnU smU + show ?thesis by auto +qed + (* The Ferrante - Rackoff Theorem *) + +theorem fr_eq: + assumes lp: "isrlfm p" + shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,n) \ set (uset p). \ (s,m) \ set (uset p). Ifm ((((Inum (x#bs) t)/ real n + (Inum (x#bs) s) / real m) /2)#bs) p))" + (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") +proof + assume px: "\ x. ?I x p" + have "?M \ ?P \ (\ ?M \ \ ?P)" by blast + moreover {assume "?M \ ?P" hence "?D" by blast} + moreover {assume nmi: "\ ?M" and npi: "\ ?P" + from rinf_uset[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast} + ultimately show "?D" by blast +next + assume "?D" + moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} + moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } + moreover {assume f:"?F" hence "?E" by blast} + ultimately show "?E" by blast +qed + + +lemma fr_equsubst: + assumes lp: "isrlfm p" + shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,k) \ set (uset p). \ (s,l) \ set (uset p). Ifm (x#bs) (usubst p (Add(Mul l t) (Mul k s) , 2*k*l))))" + (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") +proof + assume px: "\ x. ?I x p" + have "?M \ ?P \ (\ ?M \ \ ?P)" by blast + moreover {assume "?M \ ?P" hence "?D" by blast} + moreover {assume nmi: "\ ?M" and npi: "\ ?P" + let ?f ="\ (t,n). Inum (x#bs) t / real n" + let ?N = "\ t. Inum (x#bs) t" + {fix t n s m assume "(t,n)\ set (uset p)" and "(s,m) \ set (uset p)" + with uset_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" + by auto + let ?st = "Add (Mul m t) (Mul n s)" + from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" + by (simp add: mult_commute) + from tnb snb have st_nb: "numbound0 ?st" by simp + have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + using mnp mp np by (simp add: ring_simps add_divide_distrib) + from usubst_I[OF lp mnp st_nb, where x="x" and bs="bs"] + have "?I x (usubst p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])} + with rinf_uset[OF lp nmi npi px] have "?F" by blast hence "?D" by blast} + ultimately show "?D" by blast +next + assume "?D" + moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} + moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } + moreover {fix t k s l assume "(t,k) \ set (uset p)" and "(s,l) \ set (uset p)" + and px:"?I x (usubst p (Add (Mul l t) (Mul k s), 2*k*l))" + with uset_l[OF lp] have tnb: "numbound0 t" and np:"real k > 0" and snb: "numbound0 s" and mp:"real l > 0" by auto + let ?st = "Add (Mul l t) (Mul k s)" + from mult_pos_pos[OF np mp] have mnp: "real (2*k*l) > 0" + by (simp add: mult_commute) + from tnb snb have st_nb: "numbound0 ?st" by simp + from usubst_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto} + ultimately show "?E" by blast +qed + + + (* Implement the right hand side of Ferrante and Rackoff's Theorem. *) +constdefs ferrack:: "fm \ fm" + "ferrack p \ (let p' = rlfm (simpfm p); mp = minusinf p'; pp = plusinf p' + in if (mp = T \ pp = T) then T else + (let U = remdps(map simp_num_pair + (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) + (alluopairs (uset p')))) + in decr (disj mp (disj pp (evaldjf (simpfm o (usubst p')) U)))))" + +lemma uset_cong_aux: + assumes Ul: "\ (t,n) \ set U. numbound0 t \ n >0" + shows "((\ (t,n). Inum (x#bs) t /real n) ` (set (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs U)))) = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (set U \ set U))" + (is "?lhs = ?rhs") +proof(auto) + fix t n s m + assume "((t,n),(s,m)) \ set (alluopairs U)" + hence th: "((t,n),(s,m)) \ (set U \ set U)" + using alluopairs_set1[where xs="U"] by blast + let ?N = "\ t. Inum (x#bs) t" + let ?st= "Add (Mul m t) (Mul n s)" + from Ul th have mnz: "m \ 0" by auto + from Ul th have nnz: "n \ 0" by auto + have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + using mnz nnz by (simp add: ring_simps add_divide_distrib) + + thus "(real m * Inum (x # bs) t + real n * Inum (x # bs) s) / + (2 * real n * real m) + \ (\((t, n), s, m). + (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) ` + (set U \ set U)"using mnz nnz th + apply (auto simp add: th add_divide_distrib ring_simps split_def image_def) + by (rule_tac x="(s,m)" in bexI,simp_all) + (rule_tac x="(t,n)" in bexI,simp_all) +next + fix t n s m + assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" + let ?N = "\ t. Inum (x#bs) t" + let ?st= "Add (Mul m t) (Mul n s)" + from Ul smU have mnz: "m \ 0" by auto + from Ul tnU have nnz: "n \ 0" by auto + have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + using mnz nnz by (simp add: ring_simps add_divide_distrib) + let ?P = "\ (t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" + have Pc:"\ a b. ?P a b = ?P b a" + by auto + from Ul alluopairs_set1 have Up:"\ ((t,n),(s,m)) \ set (alluopairs U). n \ 0 \ m \ 0" by blast + from alluopairs_ex[OF Pc, where xs="U"] tnU smU + have th':"\ ((t',n'),(s',m')) \ set (alluopairs U). ?P (t',n') (s',m')" + by blast + then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \ set (alluopairs U)" + and Pts': "?P (t',n') (s',m')" by blast + from ts'_U Up have mnz': "m' \ 0" and nnz': "n'\ 0" by auto + let ?st' = "Add (Mul m' t') (Mul n' s')" + have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')" + using mnz' nnz' by (simp add: ring_simps add_divide_distrib) + from Pts' have + "(Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" by simp + also have "\ = ((\(t, n). Inum (x # bs) t / real n) ((\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st') + finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2 + \ (\(t, n). Inum (x # bs) t / real n) ` + (\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ` + set (alluopairs U)" + using ts'_U by blast +qed + +lemma uset_cong: + assumes lp: "isrlfm p" + and UU': "((\ (t,n). Inum (x#bs) t /real n) ` U') = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (U \ U))" (is "?f ` U' = ?g ` (U\U)") + and U: "\ (t,n) \ U. numbound0 t \ n > 0" + and U': "\ (t,n) \ U'. numbound0 t \ n > 0" + shows "(\ (t,n) \ U. \ (s,m) \ U. Ifm (x#bs) (usubst p (Add (Mul m t) (Mul n s),2*n*m))) = (\ (t,n) \ U'. Ifm (x#bs) (usubst p (t,n)))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and + Pst: "Ifm (x#bs) (usubst p (Add (Mul m t) (Mul n s),2*n*m))" by blast + let ?N = "\ t. Inum (x#bs) t" + from tnU smU U have tnb: "numbound0 t" and np: "n > 0" + and snb: "numbound0 s" and mp:"m > 0" by auto + let ?st= "Add (Mul m t) (Mul n s)" + from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" + by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) + from tnb snb have stnb: "numbound0 ?st" by simp + have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + using mp np by (simp add: ring_simps add_divide_distrib) + from tnU smU UU' have "?g ((t,n),(s,m)) \ ?f ` U'" by blast + hence "\ (t',n') \ U'. ?g ((t,n),(s,m)) = ?f (t',n')" + by auto (rule_tac x="(a,b)" in bexI, auto) + then obtain t' n' where tnU': "(t',n') \ U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast + from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto + from usubst_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst + have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp + from conjunct1[OF usubst_I[OF lp np' tnb', where bs="bs" and x="x"], symmetric] th[simplified split_def fst_conv snd_conv,symmetric] Pst2[simplified st[symmetric]] + have "Ifm (x # bs) (usubst p (t', n')) " by (simp only: st) + then show ?rhs using tnU' by auto +next + assume ?rhs + then obtain t' n' where tnU': "(t',n') \ U'" and Pt': "Ifm (x # bs) (usubst p (t', n'))" + by blast + from tnU' UU' have "?f (t',n') \ ?g ` (U\U)" by blast + hence "\ ((t,n),(s,m)) \ (U\U). ?f (t',n') = ?g ((t,n),(s,m))" + by auto (rule_tac x="(a,b)" in bexI, auto) + then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and + th: "?f (t',n') = ?g((t,n),(s,m)) "by blast + let ?N = "\ t. Inum (x#bs) t" + from tnU smU U have tnb: "numbound0 t" and np: "n > 0" + and snb: "numbound0 s" and mp:"m > 0" by auto + let ?st= "Add (Mul m t) (Mul n s)" + from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" + by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) + from tnb snb have stnb: "numbound0 ?st" by simp + have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + using mp np by (simp add: ring_simps add_divide_distrib) + from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto + from usubst_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt' + have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp + with usubst_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast +qed + +lemma ferrack: + assumes qf: "qfree p" + shows "qfree (ferrack p) \ ((Ifm bs (ferrack p)) = (\ x. Ifm (x#bs) p))" + (is "_ \ (?rhs = ?lhs)") +proof- + let ?I = "\ x p. Ifm (x#bs) p" + fix x + let ?N = "\ t. Inum (x#bs) t" + let ?q = "rlfm (simpfm p)" + let ?U = "uset ?q" + let ?Up = "alluopairs ?U" + let ?g = "\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)" + let ?S = "map ?g ?Up" + let ?SS = "map simp_num_pair ?S" + let ?Y = "remdps ?SS" + let ?f= "(\ (t,n). ?N t / real n)" + let ?h = "\ ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2" + let ?F = "\ p. \ a \ set (uset p). \ b \ set (uset p). ?I x (usubst p (?g(a,b)))" + let ?ep = "evaldjf (simpfm o (usubst ?q)) ?Y" + from rlfm_I[OF simpfm_qf[OF qf]] have lq: "isrlfm ?q" by blast + from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \ (set ?U \ set ?U)" by simp + from uset_l[OF lq] have U_l: "\ (t,n) \ set ?U. numbound0 t \ n > 0" . + from U_l UpU + have "\ ((t,n),(s,m)) \ set ?Up. numbound0 t \ n> 0 \ numbound0 s \ m > 0" by auto + hence Snb: "\ (t,n) \ set ?S. numbound0 t \ n > 0 " + by (auto simp add: mult_pos_pos) + have Y_l: "\ (t,n) \ set ?Y. numbound0 t \ n > 0" + proof- + { fix t n assume tnY: "(t,n) \ set ?Y" + hence "(t,n) \ set ?SS" by simp + hence "\ (t',n') \ set ?S. simp_num_pair (t',n') = (t,n)" + by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all) + then obtain t' n' where tn'S: "(t',n') \ set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast + from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto + from simp_num_pair_l[OF tnb np tns] + have "numbound0 t \ n > 0" . } + thus ?thesis by blast + qed + + have YU: "(?f ` set ?Y) = (?h ` (set ?U \ set ?U))" + proof- + from simp_num_pair_ci[where bs="x#bs"] have + "\x. (?f o simp_num_pair) x = ?f x" by auto + hence th: "?f o simp_num_pair = ?f" using ext by blast + have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_compose) + also have "\ = (?f ` set ?S)" by (simp add: th) + also have "\ = ((?f o ?g) ` set ?Up)" + by (simp only: set_map o_def image_compose[symmetric]) + also have "\ = (?h ` (set ?U \ set ?U))" + using uset_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_compose[symmetric]] by blast + finally show ?thesis . + qed + have "\ (t,n) \ set ?Y. bound0 (simpfm (usubst ?q (t,n)))" + proof- + { fix t n assume tnY: "(t,n) \ set ?Y" + with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto + from usubst_I[OF lq np tnb] + have "bound0 (usubst ?q (t,n))" by simp hence "bound0 (simpfm (usubst ?q (t,n)))" + using simpfm_bound0 by simp} + thus ?thesis by blast + qed + hence ep_nb: "bound0 ?ep" using evaldjf_bound0[where xs="?Y" and f="simpfm o (usubst ?q)"] by auto + let ?mp = "minusinf ?q" + let ?pp = "plusinf ?q" + let ?M = "?I x ?mp" + let ?P = "?I x ?pp" + let ?res = "disj ?mp (disj ?pp ?ep)" + from rminusinf_bound0[OF lq] rplusinf_bound0[OF lq] ep_nb + have nbth: "bound0 ?res" by auto + + from conjunct1[OF rlfm_I[OF simpfm_qf[OF qf]]] simpfm + + have th: "?lhs = (\ x. ?I x ?q)" by auto + from th fr_equsubst[OF lq, where bs="bs" and x="x"] have lhfr: "?lhs = (?M \ ?P \ ?F ?q)" + by (simp only: split_def fst_conv snd_conv) + also have "\ = (?M \ ?P \ (\ (t,n) \ set ?Y. ?I x (simpfm (usubst ?q (t,n)))))" + using uset_cong[OF lq YU U_l Y_l] by (simp only: split_def fst_conv snd_conv simpfm) + also have "\ = (Ifm (x#bs) ?res)" + using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="simpfm o (usubst ?q)",symmetric] + by (simp add: split_def pair_collapse) + finally have lheq: "?lhs = (Ifm bs (decr ?res))" using decr[OF nbth] by blast + hence lr: "?lhs = ?rhs" apply (unfold ferrack_def Let_def) + by (cases "?mp = T \ ?pp = T", auto) (simp add: disj_def)+ + from decr_qf[OF nbth] have "qfree (ferrack p)" by (auto simp add: Let_def ferrack_def) + with lr show ?thesis by blast +qed + +definition linrqe:: "fm \ fm" where + "linrqe p = qelim (prep p) ferrack" + +theorem linrqe: "Ifm bs (linrqe p) = Ifm bs p \ qfree (linrqe p)" +using ferrack qelim_ci prep +unfolding linrqe_def by auto + +definition ferrack_test :: "unit \ fm" where + "ferrack_test u = linrqe (A (A (Imp (Lt (Sub (Bound 1) (Bound 0))) + (E (Eq (Sub (Add (Bound 0) (Bound 2)) (Bound 1)))))))" + +ML {* @{code ferrack_test} () *} + +oracle linr_oracle = {* +let + +fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t + of NONE => error "Variable not found in the list!" + | SOME n => @{code Bound} n) + | num_of_term vs @{term "real (0::int)"} = @{code C} 0 + | num_of_term vs @{term "real (1::int)"} = @{code C} 1 + | num_of_term vs @{term "0::real"} = @{code C} 0 + | num_of_term vs @{term "1::real"} = @{code C} 1 + | num_of_term vs (Bound i) = @{code Bound} i + | num_of_term vs (@{term "uminus :: real \ real"} $ t') = @{code Neg} (num_of_term vs t') + | num_of_term vs (@{term "op + :: real \ real \ real"} $ t1 $ t2) = @{code Add} (num_of_term vs t1, num_of_term vs t2) + | num_of_term vs (@{term "op - :: real \ real \ real"} $ t1 $ t2) = @{code Sub} (num_of_term vs t1, num_of_term vs t2) + | num_of_term vs (@{term "op * :: real \ real \ real"} $ t1 $ t2) = (case (num_of_term vs t1) + of @{code C} i => @{code Mul} (i, num_of_term vs t2) + | _ => error "num_of_term: unsupported Multiplication") + | num_of_term vs (@{term "real :: int \ real"} $ (@{term "number_of :: int \ int"} $ t')) = @{code C} (HOLogic.dest_numeral t') + | num_of_term vs (@{term "number_of :: int \ real"} $ t') = @{code C} (HOLogic.dest_numeral t') + | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t); + +fun fm_of_term vs @{term True} = @{code T} + | fm_of_term vs @{term False} = @{code F} + | fm_of_term vs (@{term "op < :: real \ real \ bool"} $ t1 $ t2) = @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) + | fm_of_term vs (@{term "op \ :: real \ real \ bool"} $ t1 $ t2) = @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) + | fm_of_term vs (@{term "op = :: real \ real \ bool"} $ t1 $ t2) = @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) + | fm_of_term vs (@{term "op \ :: bool \ bool \ bool"} $ t1 $ t2) = @{code Iff} (fm_of_term vs t1, fm_of_term vs t2) + | fm_of_term vs (@{term "op &"} $ t1 $ t2) = @{code And} (fm_of_term vs t1, fm_of_term vs t2) + | fm_of_term vs (@{term "op |"} $ t1 $ t2) = @{code Or} (fm_of_term vs t1, fm_of_term vs t2) + | fm_of_term vs (@{term "op -->"} $ t1 $ t2) = @{code Imp} (fm_of_term vs t1, fm_of_term vs t2) + | fm_of_term vs (@{term "Not"} $ t') = @{code NOT} (fm_of_term vs t') + | fm_of_term vs (Const ("Ex", _) $ Abs (xn, xT, p)) = + @{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) + | fm_of_term vs (Const ("All", _) $ Abs (xn, xT, p)) = + @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) + | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t); + +fun term_of_num vs (@{code C} i) = @{term "real :: int \ real"} $ HOLogic.mk_number HOLogic.intT i + | term_of_num vs (@{code Bound} n) = fst (the (find_first (fn (_, m) => n = m) vs)) + | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \ real"} $ term_of_num vs t' + | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: real \ real \ real"} $ + term_of_num vs t1 $ term_of_num vs t2 + | term_of_num vs (@{code Sub} (t1, t2)) = @{term "op - :: real \ real \ real"} $ + term_of_num vs t1 $ term_of_num vs t2 + | term_of_num vs (@{code Mul} (i, t2)) = @{term "op * :: real \ real \ real"} $ + term_of_num vs (@{code C} i) $ term_of_num vs t2 + | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t)); + +fun term_of_fm vs @{code T} = HOLogic.true_const + | term_of_fm vs @{code F} = HOLogic.false_const + | term_of_fm vs (@{code Lt} t) = @{term "op < :: real \ real \ bool"} $ + term_of_num vs t $ @{term "0::real"} + | term_of_fm vs (@{code Le} t) = @{term "op \ :: real \ real \ bool"} $ + term_of_num vs t $ @{term "0::real"} + | term_of_fm vs (@{code Gt} t) = @{term "op < :: real \ real \ bool"} $ + @{term "0::real"} $ term_of_num vs t + | term_of_fm vs (@{code Ge} t) = @{term "op \ :: real \ real \ bool"} $ + @{term "0::real"} $ term_of_num vs t + | term_of_fm vs (@{code Eq} t) = @{term "op = :: real \ real \ bool"} $ + term_of_num vs t $ @{term "0::real"} + | term_of_fm vs (@{code NEq} t) = term_of_fm vs (@{code NOT} (@{code Eq} t)) + | term_of_fm vs (@{code NOT} t') = HOLogic.Not $ term_of_fm vs t' + | term_of_fm vs (@{code And} (t1, t2)) = HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2 + | term_of_fm vs (@{code Or} (t1, t2)) = HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2 + | term_of_fm vs (@{code Imp} (t1, t2)) = HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2 + | term_of_fm vs (@{code Iff} (t1, t2)) = @{term "op \ :: bool \ bool \ bool"} $ + term_of_fm vs t1 $ term_of_fm vs t2 + | term_of_fm vs _ = error "If this is raised, Isabelle/HOL or generate_code is inconsistent."; + +in fn ct => + let + val thy = Thm.theory_of_cterm ct; + val t = Thm.term_of ct; + val fs = term_frees t; + val vs = fs ~~ (0 upto (length fs - 1)); + val res = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, term_of_fm vs (@{code linrqe} (fm_of_term vs t)))); + in Thm.cterm_of thy res end +end; +*} + +use "linrtac.ML" +setup LinrTac.setup + +lemma + fixes x :: real + shows "2 * x \ 2 * x \ 2 * x \ 2 * x + 1" +apply rferrack +done + +lemma + fixes x :: real + shows "\y \ x. x = y + 1" +apply rferrack +done + +lemma + fixes x :: real + shows "\ (\z. x + z = x + z + 1)" +apply rferrack +done + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/Sqrt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Sqrt.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,98 @@ +(* Title: HOL/ex/Sqrt.thy + Author: Markus Wenzel, TU Muenchen + +*) + +header {* Square roots of primes are irrational *} + +theory Sqrt +imports Complex_Main Primes +begin + +text {* The definition and the key representation theorem for the set of +rational numbers has been moved to a core theory. *} + +declare Rats_abs_nat_div_natE[elim?] + +subsection {* Main theorem *} + +text {* + The square root of any prime number (including @{text 2}) is + irrational. +*} + +theorem sqrt_prime_irrational: + assumes "prime p" + shows "sqrt (real p) \ \" +proof + from `prime p` have p: "1 < p" by (simp add: prime_def) + assume "sqrt (real p) \ \" + then obtain m n where + n: "n \ 0" and sqrt_rat: "\sqrt (real p)\ = real m / real n" + and gcd: "gcd m n = 1" .. + have eq: "m\ = p * n\" + proof - + from n and sqrt_rat have "real m = \sqrt (real p)\ * real n" by simp + then have "real (m\) = (sqrt (real p))\ * real (n\)" + by (auto simp add: power2_eq_square) + also have "(sqrt (real p))\ = real p" by simp + also have "\ * real (n\) = real (p * n\)" by simp + finally show ?thesis .. + qed + have "p dvd m \ p dvd n" + proof + from eq have "p dvd m\" .. + with `prime p` show "p dvd m" by (rule prime_dvd_power_two) + then obtain k where "m = p * k" .. + with eq have "p * n\ = p\ * k\" by (auto simp add: power2_eq_square mult_ac) + with p have "n\ = p * k\" by (simp add: power2_eq_square) + then have "p dvd n\" .. + with `prime p` show "p dvd n" by (rule prime_dvd_power_two) + qed + then have "p dvd gcd m n" .. + with gcd have "p dvd 1" by simp + then have "p \ 1" by (simp add: dvd_imp_le) + with p show False by simp +qed + +corollary "sqrt (real (2::nat)) \ \" + by (rule sqrt_prime_irrational) (rule two_is_prime) + + +subsection {* Variations *} + +text {* + Here is an alternative version of the main proof, using mostly + linear forward-reasoning. While this results in less top-down + structure, it is probably closer to proofs seen in mathematics. +*} + +theorem + assumes "prime p" + shows "sqrt (real p) \ \" +proof + from `prime p` have p: "1 < p" by (simp add: prime_def) + assume "sqrt (real p) \ \" + then obtain m n where + n: "n \ 0" and sqrt_rat: "\sqrt (real p)\ = real m / real n" + and gcd: "gcd m n = 1" .. + from n and sqrt_rat have "real m = \sqrt (real p)\ * real n" by simp + then have "real (m\) = (sqrt (real p))\ * real (n\)" + by (auto simp add: power2_eq_square) + also have "(sqrt (real p))\ = real p" by simp + also have "\ * real (n\) = real (p * n\)" by simp + finally have eq: "m\ = p * n\" .. + then have "p dvd m\" .. + with `prime p` have dvd_m: "p dvd m" by (rule prime_dvd_power_two) + then obtain k where "m = p * k" .. + with eq have "p * n\ = p\ * k\" by (auto simp add: power2_eq_square mult_ac) + with p have "n\ = p * k\" by (simp add: power2_eq_square) + then have "p dvd n\" .. + with `prime p` have "p dvd n" by (rule prime_dvd_power_two) + with dvd_m have "p dvd gcd m n" by (rule gcd_greatest) + with gcd have "p dvd 1" by simp + then have "p \ 1" by (simp add: dvd_imp_le) + with p show False by simp +qed + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/Sqrt_Script.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Sqrt_Script.thy Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,70 @@ +(* Title: HOL/ex/Sqrt_Script.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 2001 University of Cambridge +*) + +header {* Square roots of primes are irrational (script version) *} + +theory Sqrt_Script +imports Complex_Main Primes +begin + +text {* + \medskip Contrast this linear Isabelle/Isar script with Markus + Wenzel's more mathematical version. +*} + +subsection {* Preliminaries *} + +lemma prime_nonzero: "prime p \ p \ 0" + by (force simp add: prime_def) + +lemma prime_dvd_other_side: + "n * n = p * (k * k) \ prime p \ p dvd n" + apply (subgoal_tac "p dvd n * n", blast dest: prime_dvd_mult) + apply auto + done + +lemma reduction: "prime p \ + 0 < k \ k * k = p * (j * j) \ k < p * j \ 0 < j" + apply (rule ccontr) + apply (simp add: linorder_not_less) + apply (erule disjE) + apply (frule mult_le_mono, assumption) + apply auto + apply (force simp add: prime_def) + done + +lemma rearrange: "(j::nat) * (p * j) = k * k \ k * k = p * (j * j)" + by (simp add: mult_ac) + +lemma prime_not_square: + "prime p \ (\k. 0 < k \ m * m \ p * (k * k))" + apply (induct m rule: nat_less_induct) + apply clarify + apply (frule prime_dvd_other_side, assumption) + apply (erule dvdE) + apply (simp add: nat_mult_eq_cancel_disj prime_nonzero) + apply (blast dest: rearrange reduction) + done + + +subsection {* Main theorem *} + +text {* + The square root of any prime number (including @{text 2}) is + irrational. +*} + +theorem prime_sqrt_irrational: + "prime p \ x * x = real p \ 0 \ x \ x \ \" + apply (rule notI) + apply (erule Rats_abs_nat_div_natE) + apply (simp del: real_of_nat_mult + add: real_abs_def divide_eq_eq prime_not_square real_of_nat_mult [symmetric]) + done + +lemmas two_sqrt_irrational = + prime_sqrt_irrational [OF two_is_prime] + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/linrtac.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/linrtac.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,109 @@ +structure LinrTac = +struct + +val trace = ref false; +fun trace_msg s = if !trace then tracing s else (); + +val ferrack_ss = let val ths = [@{thm real_of_int_inject}, @{thm real_of_int_less_iff}, + @{thm real_of_int_le_iff}] + in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths) + end; + +val binarith = + @{thms normalize_bin_simps} @ @{thms pred_bin_simps} @ @{thms succ_bin_simps} @ + @{thms add_bin_simps} @ @{thms minus_bin_simps} @ @{thms mult_bin_simps}; +val comp_arith = binarith @ simp_thms + +val zdvd_int = @{thm zdvd_int}; +val zdiff_int_split = @{thm zdiff_int_split}; +val all_nat = @{thm all_nat}; +val ex_nat = @{thm ex_nat}; +val number_of1 = @{thm number_of1}; +val number_of2 = @{thm number_of2}; +val split_zdiv = @{thm split_zdiv}; +val split_zmod = @{thm split_zmod}; +val mod_div_equality' = @{thm mod_div_equality'}; +val split_div' = @{thm split_div'}; +val Suc_plus1 = @{thm Suc_plus1}; +val imp_le_cong = @{thm imp_le_cong}; +val conj_le_cong = @{thm conj_le_cong}; +val nat_mod_add_eq = @{thm mod_add1_eq} RS sym; +val nat_mod_add_left_eq = @{thm mod_add_left_eq} RS sym; +val nat_mod_add_right_eq = @{thm mod_add_right_eq} RS sym; +val int_mod_add_eq = @{thm zmod_zadd1_eq} RS sym; +val int_mod_add_left_eq = @{thm zmod_zadd_left_eq} RS sym; +val int_mod_add_right_eq = @{thm zmod_zadd_right_eq} RS sym; +val nat_div_add_eq = @{thm div_add1_eq} RS sym; +val int_div_add_eq = @{thm zdiv_zadd1_eq} RS sym; +val ZDIVISION_BY_ZERO_MOD = @{thm DIVISION_BY_ZERO} RS conjunct2; +val ZDIVISION_BY_ZERO_DIV = @{thm DIVISION_BY_ZERO} RS conjunct1; + +fun prepare_for_linr sg q fm = + let + val ps = Logic.strip_params fm + val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm) + val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm) + fun mk_all ((s, T), (P,n)) = + if 0 mem loose_bnos P then + (HOLogic.all_const T $ Abs (s, T, P), n) + else (incr_boundvars ~1 P, n-1) + fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t; + val rhs = hs +(* val (rhs,irhs) = List.partition (relevant (rev ps)) hs *) + val np = length ps + val (fm',np) = foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n))) + (foldr HOLogic.mk_imp c rhs, np) ps + val (vs, _) = List.partition (fn t => q orelse (type_of t) = HOLogic.natT) + (term_frees fm' @ term_vars fm'); + val fm2 = foldr mk_all2 fm' vs + in (fm2, np + length vs, length rhs) end; + +(*Object quantifier to meta --*) +fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ; + +(* object implication to meta---*) +fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp; + + +fun linr_tac ctxt q i = + (ObjectLogic.atomize_prems_tac i) + THEN (REPEAT_DETERM (split_tac [@{thm split_min}, @{thm split_max}, @{thm abs_split}] i)) + THEN (fn st => + let + val g = List.nth (prems_of st, i - 1) + val thy = ProofContext.theory_of ctxt + (* Transform the term*) + val (t,np,nh) = prepare_for_linr thy q g + (* Some simpsets for dealing with mod div abs and nat*) + val simpset0 = Simplifier.theory_context thy HOL_basic_ss addsimps comp_arith + val ct = cterm_of thy (HOLogic.mk_Trueprop t) + (* Theorem for the nat --> int transformation *) + val pre_thm = Seq.hd (EVERY + [simp_tac simpset0 1, + TRY (simp_tac (Simplifier.theory_context thy ferrack_ss) 1)] + (trivial ct)) + fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) + (* The result of the quantifier elimination *) + val (th, tac) = case (prop_of pre_thm) of + Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ => + let val pth = linr_oracle (cterm_of thy (Pattern.eta_long [] t1)) + in + (trace_msg ("calling procedure with term:\n" ^ + Syntax.string_of_term ctxt t1); + ((pth RS iffD2) RS pre_thm, + assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))) + end + | _ => (pre_thm, assm_tac i) + in (rtac (((mp_step nh) o (spec_step np)) th) i + THEN tac) st + end handle Subscript => no_tac st); + +fun linr_meth src = + Method.syntax (Args.mode "no_quantify") src + #> (fn (q, ctxt) => Method.SIMPLE_METHOD' (linr_tac ctxt (not q))); + +val setup = + Method.add_method ("rferrack", linr_meth, + "decision procedure for linear real arithmetic"); + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/ex/mirtac.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/mirtac.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,169 @@ +(* Title: HOL/Complex/ex/mirtac.ML + ID: $Id$ + Author: Amine Chaieb, TU Muenchen +*) + +structure MirTac = +struct + +val trace = ref false; +fun trace_msg s = if !trace then tracing s else (); + +val mir_ss = +let val ths = map thm ["real_of_int_inject", "real_of_int_less_iff", "real_of_int_le_iff"] +in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths) +end; + +val nT = HOLogic.natT; + val nat_arith = map thm ["add_nat_number_of", "diff_nat_number_of", + "mult_nat_number_of", "eq_nat_number_of", "less_nat_number_of"]; + + val comp_arith = (map thm ["Let_def", "if_False", "if_True", "add_0", + "add_Suc", "add_number_of_left", "mult_number_of_left", + "Suc_eq_add_numeral_1"])@ + (map (fn s => thm s RS sym) ["numeral_1_eq_1", "numeral_0_eq_0"]) + @ @{thms arith_simps} @ nat_arith @ @{thms rel_simps} + val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, + @{thm "real_of_nat_number_of"}, + @{thm "real_of_nat_Suc"}, @{thm "real_of_nat_one"}, @{thm "real_of_one"}, + @{thm "real_of_int_zero"}, @{thm "real_of_nat_zero"}, + @{thm "Ring_and_Field.divide_zero"}, + @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"}, + @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"}, + @{thm "diff_def"}, @{thm "minus_divide_left"}] +val comp_ths = ths @ comp_arith @ simp_thms + + +val zdvd_int = @{thm "zdvd_int"}; +val zdiff_int_split = @{thm "zdiff_int_split"}; +val all_nat = @{thm "all_nat"}; +val ex_nat = @{thm "ex_nat"}; +val number_of1 = @{thm "number_of1"}; +val number_of2 = @{thm "number_of2"}; +val split_zdiv = @{thm "split_zdiv"}; +val split_zmod = @{thm "split_zmod"}; +val mod_div_equality' = @{thm "mod_div_equality'"}; +val split_div' = @{thm "split_div'"}; +val Suc_plus1 = @{thm "Suc_plus1"}; +val imp_le_cong = @{thm "imp_le_cong"}; +val conj_le_cong = @{thm "conj_le_cong"}; +val nat_mod_add_eq = @{thm "mod_add1_eq"} RS sym; +val nat_mod_add_left_eq = @{thm "mod_add_left_eq"} RS sym; +val nat_mod_add_right_eq = @{thm "mod_add_right_eq"} RS sym; +val int_mod_add_eq = @{thm "zmod_zadd1_eq"} RS sym; +val int_mod_add_left_eq = @{thm "zmod_zadd_left_eq"} RS sym; +val int_mod_add_right_eq = @{thm "zmod_zadd_right_eq"} RS sym; +val nat_div_add_eq = @{thm "div_add1_eq"} RS sym; +val int_div_add_eq = @{thm "zdiv_zadd1_eq"} RS sym; +val ZDIVISION_BY_ZERO_MOD = @{thm "DIVISION_BY_ZERO"} RS conjunct2; +val ZDIVISION_BY_ZERO_DIV = @{thm "DIVISION_BY_ZERO"} RS conjunct1; + +fun prepare_for_mir thy q fm = + let + val ps = Logic.strip_params fm + val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm) + val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm) + fun mk_all ((s, T), (P,n)) = + if 0 mem loose_bnos P then + (HOLogic.all_const T $ Abs (s, T, P), n) + else (incr_boundvars ~1 P, n-1) + fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t; + val rhs = hs +(* val (rhs,irhs) = List.partition (relevant (rev ps)) hs *) + val np = length ps + val (fm',np) = foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n))) + (foldr HOLogic.mk_imp c rhs, np) ps + val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT) + (term_frees fm' @ term_vars fm'); + val fm2 = foldr mk_all2 fm' vs + in (fm2, np + length vs, length rhs) end; + +(*Object quantifier to meta --*) +fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ; + +(* object implication to meta---*) +fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp; + + +fun mir_tac ctxt q i = + (ObjectLogic.atomize_prems_tac i) + THEN (simp_tac (HOL_basic_ss addsimps [@{thm "abs_ge_zero"}] addsimps simp_thms) i) + THEN (REPEAT_DETERM (split_tac [@{thm "split_min"}, @{thm "split_max"},@{thm "abs_split"}] i)) + THEN (fn st => + let + val g = List.nth (prems_of st, i - 1) + val thy = ProofContext.theory_of ctxt + (* Transform the term*) + val (t,np,nh) = prepare_for_mir thy q g + (* Some simpsets for dealing with mod div abs and nat*) + val mod_div_simpset = HOL_basic_ss + addsimps [refl,nat_mod_add_eq, + @{thm "mod_self"}, @{thm "zmod_self"}, + @{thm "zdiv_zero"},@{thm "zmod_zero"},@{thm "div_0"}, @{thm "mod_0"}, + @{thm "zdiv_1"}, @{thm "zmod_1"}, @{thm "div_1"}, @{thm "mod_1"}, + @{thm "Suc_plus1"}] + addsimps @{thms add_ac} + addsimprocs [cancel_div_mod_proc] + val simpset0 = HOL_basic_ss + addsimps [mod_div_equality', Suc_plus1] + addsimps comp_ths + addsplits [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, @{thm "split_min"}, @{thm "split_max"}] + (* Simp rules for changing (n::int) to int n *) + val simpset1 = HOL_basic_ss + addsimps [@{thm "nat_number_of_def"}, @{thm "zdvd_int"}] @ map (fn r => r RS sym) + [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, + @{thm "zmult_int"}] + addsplits [@{thm "zdiff_int_split"}] + (*simp rules for elimination of int n*) + + val simpset2 = HOL_basic_ss + addsimps [@{thm "nat_0_le"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "number_of1"}, + @{thm "number_of2"}, @{thm "int_0"}, @{thm "int_1"}] + addcongs [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}] + (* simp rules for elimination of abs *) + val ct = cterm_of thy (HOLogic.mk_Trueprop t) + (* Theorem for the nat --> int transformation *) + val pre_thm = Seq.hd (EVERY + [simp_tac mod_div_simpset 1, simp_tac simpset0 1, + TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), TRY (simp_tac mir_ss 1)] + (trivial ct)) + fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) + (* The result of the quantifier elimination *) + val (th, tac) = case (prop_of pre_thm) of + Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ => + let val pth = + (* If quick_and_dirty then run without proof generation as oracle*) + if !quick_and_dirty + then mirfr_oracle (false, cterm_of thy (Pattern.eta_long [] t1)) + else mirfr_oracle (true, cterm_of thy (Pattern.eta_long [] t1)) + in + (trace_msg ("calling procedure with term:\n" ^ + Syntax.string_of_term ctxt t1); + ((pth RS iffD2) RS pre_thm, + assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))) + end + | _ => (pre_thm, assm_tac i) + in (rtac (((mp_step nh) o (spec_step np)) th) i + THEN tac) st + end handle Subscript => no_tac st); + +fun mir_args meth = + let val parse_flag = + Args.$$$ "no_quantify" >> (K (K false)); + in + Method.simple_args + (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> + curry (Library.foldl op |>) true) + (fn q => fn ctxt => meth ctxt q 1) + end; + +fun mir_method ctxt q i = Method.METHOD (fn facts => + Method.insert_tac facts 1 THEN mir_tac ctxt q i); + +val setup = + Method.add_method ("mir", + mir_args mir_method, + "decision procedure for MIR arithmetic"); + + +end diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/hologic.ML --- a/src/HOL/hologic.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,568 +0,0 @@ -(* Title: HOL/hologic.ML - ID: $Id$ - Author: Lawrence C Paulson and Markus Wenzel - -Abstract syntax operations for HOL. -*) - -signature HOLOGIC = -sig - val typeS: sort - val typeT: typ - val boolN: string - val boolT: typ - val true_const: term - val false_const: term - val mk_setT: typ -> typ - val dest_setT: typ -> typ - val Trueprop: term - val mk_Trueprop: term -> term - val dest_Trueprop: term -> term - val conj_intr: thm -> thm -> thm - val conj_elim: thm -> thm * thm - val conj_elims: thm -> thm list - val conj: term - val disj: term - val imp: term - val Not: term - val mk_conj: term * term -> term - val mk_disj: term * term -> term - val mk_imp: term * term -> term - val mk_not: term -> term - val dest_conj: term -> term list - val dest_disj: term -> term list - val disjuncts: term -> term list - val dest_imp: term -> term * term - val dest_not: term -> term - val eq_const: typ -> term - val mk_eq: term * term -> term - val dest_eq: term -> term * term - val all_const: typ -> term - val mk_all: string * typ * term -> term - val list_all: (string * typ) list * term -> term - val exists_const: typ -> term - val mk_exists: string * typ * term -> term - val choice_const: typ -> term - val Collect_const: typ -> term - val mk_Collect: string * typ * term -> term - val class_eq: string - val mk_mem: term * term -> term - val dest_mem: term -> term * term - val mk_UNIV: typ -> term - val mk_binop: string -> term * term -> term - val mk_binrel: string -> term * term -> term - val dest_bin: string -> typ -> term -> term * term - val unitT: typ - val is_unitT: typ -> bool - val unit: term - val is_unit: term -> bool - val mk_prodT: typ * typ -> typ - val dest_prodT: typ -> typ * typ - val pair_const: typ -> typ -> term - val mk_prod: term * term -> term - val dest_prod: term -> term * term - val mk_fst: term -> term - val mk_snd: term -> term - val split_const: typ * typ * typ -> term - val mk_split: term -> term - val prodT_factors: typ -> typ list - val mk_tuple: typ -> term list -> term - val dest_tuple: term -> term list - val ap_split: typ -> typ -> term -> term - val prod_factors: term -> int list list - val dest_tuple': int list list -> term -> term list - val prodT_factors': int list list -> typ -> typ list - val ap_split': int list list -> typ -> typ -> term -> term - val mk_tuple': int list list -> typ -> term list -> term - val mk_tupleT: int list list -> typ list -> typ - val strip_split: term -> term * typ list * int list list - val natT: typ - val zero: term - val is_zero: term -> bool - val mk_Suc: term -> term - val dest_Suc: term -> term - val Suc_zero: term - val mk_nat: int -> term - val dest_nat: term -> int - val class_size: string - val size_const: typ -> term - val indexT: typ - val intT: typ - val pls_const: term - val min_const: term - val bit0_const: term - val bit1_const: term - val mk_bit: int -> term - val dest_bit: term -> int - val mk_numeral: int -> term - val dest_numeral: term -> int - val number_of_const: typ -> term - val add_numerals: term -> (term * typ) list -> (term * typ) list - val mk_number: typ -> int -> term - val dest_number: term -> typ * int - val realT: typ - val nibbleT: typ - val mk_nibble: int -> term - val dest_nibble: term -> int - val charT: typ - val mk_char: int -> term - val dest_char: term -> int - val listT: typ -> typ - val nil_const: typ -> term - val cons_const: typ -> term - val mk_list: typ -> term list -> term - val dest_list: term -> term list - val stringT: typ - val mk_string: string -> term - val dest_string: term -> string -end; - -structure HOLogic: HOLOGIC = -struct - -(* HOL syntax *) - -val typeS: sort = ["HOL.type"]; -val typeT = TypeInfer.anyT typeS; - - -(* bool and set *) - -val boolN = "bool"; -val boolT = Type (boolN, []); - -val true_const = Const ("True", boolT); -val false_const = Const ("False", boolT); - -fun mk_setT T = T --> boolT; - -fun dest_setT (Type ("fun", [T, Type ("bool", [])])) = T - | dest_setT T = raise TYPE ("dest_setT: set type expected", [T], []); - - -(* logic *) - -val Trueprop = Const ("Trueprop", boolT --> propT); - -fun mk_Trueprop P = Trueprop $ P; - -fun dest_Trueprop (Const ("Trueprop", _) $ P) = P - | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]); - -fun conj_intr thP thQ = - let - val (P, Q) = pairself (ObjectLogic.dest_judgment o Thm.cprop_of) (thP, thQ) - handle CTERM (msg, _) => raise THM (msg, 0, [thP, thQ]); - val inst = Thm.instantiate ([], [(@{cpat "?P::bool"}, P), (@{cpat "?Q::bool"}, Q)]); - in Drule.implies_elim_list (inst @{thm conjI}) [thP, thQ] end; - -fun conj_elim thPQ = - let - val (P, Q) = Thm.dest_binop (ObjectLogic.dest_judgment (Thm.cprop_of thPQ)) - handle CTERM (msg, _) => raise THM (msg, 0, [thPQ]); - val inst = Thm.instantiate ([], [(@{cpat "?P::bool"}, P), (@{cpat "?Q::bool"}, Q)]); - val thP = Thm.implies_elim (inst @{thm conjunct1}) thPQ; - val thQ = Thm.implies_elim (inst @{thm conjunct2}) thPQ; - in (thP, thQ) end; - -fun conj_elims th = - let val (th1, th2) = conj_elim th - in conj_elims th1 @ conj_elims th2 end handle THM _ => [th]; - -val conj = @{term "op &"} -and disj = @{term "op |"} -and imp = @{term "op -->"} -and Not = @{term "Not"}; - -fun mk_conj (t1, t2) = conj $ t1 $ t2 -and mk_disj (t1, t2) = disj $ t1 $ t2 -and mk_imp (t1, t2) = imp $ t1 $ t2 -and mk_not t = Not $ t; - -fun dest_conj (Const ("op &", _) $ t $ t') = t :: dest_conj t' - | dest_conj t = [t]; - -fun dest_disj (Const ("op |", _) $ t $ t') = t :: dest_disj t' - | dest_disj t = [t]; - -(*Like dest_disj, but flattens disjunctions however nested*) -fun disjuncts_aux (Const ("op |", _) $ t $ t') disjs = disjuncts_aux t (disjuncts_aux t' disjs) - | disjuncts_aux t disjs = t::disjs; - -fun disjuncts t = disjuncts_aux t []; - -fun dest_imp (Const("op -->",_) $ A $ B) = (A, B) - | dest_imp t = raise TERM ("dest_imp", [t]); - -fun dest_not (Const ("Not", _) $ t) = t - | dest_not t = raise TERM ("dest_not", [t]); - -fun eq_const T = Const ("op =", [T, T] ---> boolT); -fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u; - -fun dest_eq (Const ("op =", _) $ lhs $ rhs) = (lhs, rhs) - | dest_eq t = raise TERM ("dest_eq", [t]) - -fun all_const T = Const ("All", [T --> boolT] ---> boolT); -fun mk_all (x, T, P) = all_const T $ absfree (x, T, P); -fun list_all (xs, t) = fold_rev (fn (x, T) => fn P => all_const T $ Abs (x, T, P)) xs t; - -fun exists_const T = Const ("Ex", [T --> boolT] ---> boolT); -fun mk_exists (x, T, P) = exists_const T $ absfree (x, T, P); - -fun choice_const T = Const("Hilbert_Choice.Eps", (T --> boolT) --> T); - -fun Collect_const T = Const ("Collect", [T --> boolT] ---> mk_setT T); -fun mk_Collect (a, T, t) = Collect_const T $ absfree (a, T, t); - -val class_eq = "HOL.eq"; - -fun mk_mem (x, A) = - let val setT = fastype_of A in - Const ("op :", [dest_setT setT, setT] ---> boolT) $ x $ A - end; - -fun dest_mem (Const ("op :", _) $ x $ A) = (x, A) - | dest_mem t = raise TERM ("dest_mem", [t]); - -fun mk_UNIV T = Const ("UNIV", mk_setT T); - - -(* binary operations and relations *) - -fun mk_binop c (t, u) = - let val T = fastype_of t in - Const (c, [T, T] ---> T) $ t $ u - end; - -fun mk_binrel c (t, u) = - let val T = fastype_of t in - Const (c, [T, T] ---> boolT) $ t $ u - end; - -(*destruct the application of a binary operator. The dummyT case is a crude - way of handling polymorphic operators.*) -fun dest_bin c T (tm as Const (c', Type ("fun", [T', _])) $ t $ u) = - if c = c' andalso (T=T' orelse T=dummyT) then (t, u) - else raise TERM ("dest_bin " ^ c, [tm]) - | dest_bin c _ tm = raise TERM ("dest_bin " ^ c, [tm]); - - -(* unit *) - -val unitT = Type ("Product_Type.unit", []); - -fun is_unitT (Type ("Product_Type.unit", [])) = true - | is_unitT _ = false; - -val unit = Const ("Product_Type.Unity", unitT); - -fun is_unit (Const ("Product_Type.Unity", _)) = true - | is_unit _ = false; - - -(* prod *) - -fun mk_prodT (T1, T2) = Type ("*", [T1, T2]); - -fun dest_prodT (Type ("*", [T1, T2])) = (T1, T2) - | dest_prodT T = raise TYPE ("dest_prodT", [T], []); - -fun pair_const T1 T2 = Const ("Pair", [T1, T2] ---> mk_prodT (T1, T2)); - -fun mk_prod (t1, t2) = - let val T1 = fastype_of t1 and T2 = fastype_of t2 in - pair_const T1 T2 $ t1 $ t2 - end; - -fun dest_prod (Const ("Pair", _) $ t1 $ t2) = (t1, t2) - | dest_prod t = raise TERM ("dest_prod", [t]); - -fun mk_fst p = - let val pT = fastype_of p in - Const ("fst", pT --> fst (dest_prodT pT)) $ p - end; - -fun mk_snd p = - let val pT = fastype_of p in - Const ("snd", pT --> snd (dest_prodT pT)) $ p - end; - -fun split_const (A, B, C) = - Const ("split", (A --> B --> C) --> mk_prodT (A, B) --> C); - -fun mk_split t = - (case Term.fastype_of t of - T as (Type ("fun", [A, Type ("fun", [B, C])])) => - Const ("split", T --> mk_prodT (A, B) --> C) $ t - | _ => raise TERM ("mk_split: bad body type", [t])); - -(*Maps the type T1 * ... * Tn to [T1, ..., Tn], however nested*) -fun prodT_factors (Type ("*", [T1, T2])) = prodT_factors T1 @ prodT_factors T2 - | prodT_factors T = [T]; - -(*Makes a nested tuple from a list, following the product type structure*) -fun mk_tuple (Type ("*", [T1, T2])) tms = - mk_prod (mk_tuple T1 tms, - mk_tuple T2 (Library.drop (length (prodT_factors T1), tms))) - | mk_tuple T (t::_) = t; - -fun dest_tuple (Const ("Pair", _) $ t $ u) = dest_tuple t @ dest_tuple u - | dest_tuple t = [t]; - -(*In ap_split S T u, term u expects separate arguments for the factors of S, - with result type T. The call creates a new term expecting one argument - of type S.*) -fun ap_split T T3 u = - let - fun ap (T :: Ts) = - (case T of - Type ("*", [T1, T2]) => - split_const (T1, T2, Ts ---> T3) $ ap (T1 :: T2 :: Ts) - | _ => Abs ("x", T, ap Ts)) - | ap [] = - let val k = length (prodT_factors T) - in list_comb (incr_boundvars k u, map Bound (k - 1 downto 0)) end - in ap [T] end; - - -(* operations on tuples with specific arities *) -(* - an "arity" of a tuple is a list of lists of integers - ("factors"), denoting paths to subterms that are pairs -*) - -fun prod_err s = raise TERM (s ^ ": inconsistent use of products", []); - -fun prod_factors t = - let - fun factors p (Const ("Pair", _) $ t $ u) = - p :: factors (1::p) t @ factors (2::p) u - | factors p _ = [] - in factors [] t end; - -fun dest_tuple' ps = - let - fun dest p t = if p mem ps then (case t of - Const ("Pair", _) $ t $ u => - dest (1::p) t @ dest (2::p) u - | _ => prod_err "dest_tuple'") else [t] - in dest [] end; - -fun prodT_factors' ps = - let - fun factors p T = if p mem ps then (case T of - Type ("*", [T1, T2]) => - factors (1::p) T1 @ factors (2::p) T2 - | _ => prod_err "prodT_factors'") else [T] - in factors [] end; - -(*In ap_split' ps S T u, term u expects separate arguments for the factors of S, - with result type T. The call creates a new term expecting one argument - of type S.*) -fun ap_split' ps T T3 u = - let - fun ap ((p, T) :: pTs) = - if p mem ps then (case T of - Type ("*", [T1, T2]) => - split_const (T1, T2, map snd pTs ---> T3) $ - ap ((1::p, T1) :: (2::p, T2) :: pTs) - | _ => prod_err "ap_split'") - else Abs ("x", T, ap pTs) - | ap [] = - let val k = length ps - in list_comb (incr_boundvars (k + 1) u, map Bound (k downto 0)) end - in ap [([], T)] end; - -fun mk_tuple' ps = - let - fun mk p T ts = - if p mem ps then (case T of - Type ("*", [T1, T2]) => - let - val (t, ts') = mk (1::p) T1 ts; - val (u, ts'') = mk (2::p) T2 ts' - in (pair_const T1 T2 $ t $ u, ts'') end - | _ => prod_err "mk_tuple'") - else (hd ts, tl ts) - in fst oo mk [] end; - -fun mk_tupleT ps = - let - fun mk p Ts = - if p mem ps then - let - val (T, Ts') = mk (1::p) Ts; - val (U, Ts'') = mk (2::p) Ts' - in (mk_prodT (T, U), Ts'') end - else (hd Ts, tl Ts) - in fst o mk [] end; - -fun strip_split t = - let - fun strip [] qs Ts t = (t, Ts, qs) - | strip (p :: ps) qs Ts (Const ("split", _) $ t) = - strip ((1 :: p) :: (2 :: p) :: ps) (p :: qs) Ts t - | strip (p :: ps) qs Ts (Abs (s, T, t)) = strip ps qs (T :: Ts) t - | strip (p :: ps) qs Ts t = strip ps qs - (hd (binder_types (fastype_of1 (Ts, t))) :: Ts) - (incr_boundvars 1 t $ Bound 0) - in strip [[]] [] [] t end; - - -(* nat *) - -val natT = Type ("nat", []); - -val zero = Const ("HOL.zero_class.zero", natT); - -fun is_zero (Const ("HOL.zero_class.zero", _)) = true - | is_zero _ = false; - -fun mk_Suc t = Const ("Suc", natT --> natT) $ t; - -fun dest_Suc (Const ("Suc", _) $ t) = t - | dest_Suc t = raise TERM ("dest_Suc", [t]); - -val Suc_zero = mk_Suc zero; - -fun mk_nat n = - let - fun mk 0 = zero - | mk n = mk_Suc (mk (n - 1)); - in if n < 0 then raise TERM ("mk_nat: negative number", []) else mk n end; - -fun dest_nat (Const ("HOL.zero_class.zero", _)) = 0 - | dest_nat (Const ("Suc", _) $ t) = dest_nat t + 1 - | dest_nat t = raise TERM ("dest_nat", [t]); - -val class_size = "Nat.size"; - -fun size_const T = Const ("Nat.size_class.size", T --> natT); - - -(* index *) - -val indexT = Type ("Code_Index.index", []); - - -(* binary numerals and int -- non-unique representation due to leading zeros/ones! *) - -val intT = Type ("Int.int", []); - -val pls_const = Const ("Int.Pls", intT) -and min_const = Const ("Int.Min", intT) -and bit0_const = Const ("Int.Bit0", intT --> intT) -and bit1_const = Const ("Int.Bit1", intT --> intT); - -fun mk_bit 0 = bit0_const - | mk_bit 1 = bit1_const - | mk_bit _ = raise TERM ("mk_bit", []); - -fun dest_bit (Const ("Int.Bit0", _)) = 0 - | dest_bit (Const ("Int.Bit1", _)) = 1 - | dest_bit t = raise TERM ("dest_bit", [t]); - -fun mk_numeral 0 = pls_const - | mk_numeral ~1 = min_const - | mk_numeral i = - let val (q, r) = Integer.div_mod i 2; - in mk_bit r $ mk_numeral q end; - -fun dest_numeral (Const ("Int.Pls", _)) = 0 - | dest_numeral (Const ("Int.Min", _)) = ~1 - | dest_numeral (Const ("Int.Bit0", _) $ bs) = 2 * dest_numeral bs - | dest_numeral (Const ("Int.Bit1", _) $ bs) = 2 * dest_numeral bs + 1 - | dest_numeral t = raise TERM ("dest_numeral", [t]); - -fun number_of_const T = Const ("Int.number_class.number_of", intT --> T); - -fun add_numerals (Const ("Int.number_class.number_of", Type (_, [_, T])) $ t) = cons (t, T) - | add_numerals (t $ u) = add_numerals t #> add_numerals u - | add_numerals (Abs (_, _, t)) = add_numerals t - | add_numerals _ = I; - -fun mk_number T 0 = Const ("HOL.zero_class.zero", T) - | mk_number T 1 = Const ("HOL.one_class.one", T) - | mk_number T i = number_of_const T $ mk_numeral i; - -fun dest_number (Const ("HOL.zero_class.zero", T)) = (T, 0) - | dest_number (Const ("HOL.one_class.one", T)) = (T, 1) - | dest_number (Const ("Int.number_class.number_of", Type ("fun", [_, T])) $ t) = - (T, dest_numeral t) - | dest_number t = raise TERM ("dest_number", [t]); - - -(* real *) - -val realT = Type ("RealDef.real", []); - - -(* nibble *) - -val nibbleT = Type ("List.nibble", []); - -fun mk_nibble n = - let val s = - if 0 <= n andalso n <= 9 then chr (n + ord "0") - else if 10 <= n andalso n <= 15 then chr (n + ord "A" - 10) - else raise TERM ("mk_nibble", []) - in Const ("List.nibble.Nibble" ^ s, nibbleT) end; - -fun dest_nibble t = - let fun err () = raise TERM ("dest_nibble", [t]) in - (case try (unprefix "List.nibble.Nibble" o fst o Term.dest_Const) t of - NONE => err () - | SOME c => - if size c <> 1 then err () - else if "0" <= c andalso c <= "9" then ord c - ord "0" - else if "A" <= c andalso c <= "F" then ord c - ord "A" + 10 - else err ()) - end; - - -(* char *) - -val charT = Type ("List.char", []); - -fun mk_char n = - if 0 <= n andalso n <= 255 then - Const ("List.char.Char", nibbleT --> nibbleT --> charT) $ - mk_nibble (n div 16) $ mk_nibble (n mod 16) - else raise TERM ("mk_char", []); - -fun dest_char (Const ("List.char.Char", _) $ t $ u) = - dest_nibble t * 16 + dest_nibble u - | dest_char t = raise TERM ("dest_char", [t]); - - -(* list *) - -fun listT T = Type ("List.list", [T]); - -fun nil_const T = Const ("List.list.Nil", listT T); - -fun cons_const T = - let val lT = listT T - in Const ("List.list.Cons", T --> lT --> lT) end; - -fun mk_list T ts = - let - val lT = listT T; - val Nil = Const ("List.list.Nil", lT); - fun Cons t u = Const ("List.list.Cons", T --> lT --> lT) $ t $ u; - in fold_rev Cons ts Nil end; - -fun dest_list (Const ("List.list.Nil", _)) = [] - | dest_list (Const ("List.list.Cons", _) $ t $ u) = t :: dest_list u - | dest_list t = raise TERM ("dest_list", [t]); - - -(* string *) - -val stringT = Type ("List.string", []); - -val mk_string = mk_list charT o map (mk_char o ord) o explode; -val dest_string = implode o map (chr o dest_char) o dest_list; - -end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/int_arith1.ML --- a/src/HOL/int_arith1.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,588 +0,0 @@ -(* Title: HOL/int_arith1.ML - ID: $Id$ - Authors: Larry Paulson and Tobias Nipkow - -Simprocs and decision procedure for linear arithmetic. -*) - -structure Int_Numeral_Base_Simprocs = - struct - fun prove_conv tacs ctxt (_: thm list) (t, u) = - if t aconv u then NONE - else - let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u)) - in SOME (Goal.prove ctxt [] [] eq (K (EVERY tacs))) end - - fun prove_conv_nohyps tacs sg = prove_conv tacs sg []; - - fun prep_simproc (name, pats, proc) = - Simplifier.simproc (the_context()) name pats proc; - - fun is_numeral (Const(@{const_name Int.number_of}, _) $ w) = true - | is_numeral _ = false - - fun simplify_meta_eq f_number_of_eq f_eq = - mk_meta_eq ([f_eq, f_number_of_eq] MRS trans) - - (*reorientation simprules using ==, for the following simproc*) - val meta_zero_reorient = @{thm zero_reorient} RS eq_reflection - val meta_one_reorient = @{thm one_reorient} RS eq_reflection - val meta_number_of_reorient = @{thm number_of_reorient} RS eq_reflection - - (*reorientation simplification procedure: reorients (polymorphic) - 0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a Int.*) - fun reorient_proc sg _ (_ $ t $ u) = - case u of - Const(@{const_name HOL.zero}, _) => NONE - | Const(@{const_name HOL.one}, _) => NONE - | Const(@{const_name Int.number_of}, _) $ _ => NONE - | _ => SOME (case t of - Const(@{const_name HOL.zero}, _) => meta_zero_reorient - | Const(@{const_name HOL.one}, _) => meta_one_reorient - | Const(@{const_name Int.number_of}, _) $ _ => meta_number_of_reorient) - - val reorient_simproc = - prep_simproc ("reorient_simproc", ["0=x", "1=x", "number_of w = x"], reorient_proc) - - end; - - -Addsimprocs [Int_Numeral_Base_Simprocs.reorient_simproc]; - - -structure Int_Numeral_Simprocs = -struct - -(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in Int_Numeral_Base_Simprocs - isn't complicated by the abstract 0 and 1.*) -val numeral_syms = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym]; - -(** New term ordering so that AC-rewriting brings numerals to the front **) - -(*Order integers by absolute value and then by sign. The standard integer - ordering is not well-founded.*) -fun num_ord (i,j) = - (case int_ord (abs i, abs j) of - EQUAL => int_ord (Int.sign i, Int.sign j) - | ord => ord); - -(*This resembles Term.term_ord, but it puts binary numerals before other - non-atomic terms.*) -local open Term -in -fun numterm_ord (Abs (_, T, t), Abs(_, U, u)) = - (case numterm_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord) - | numterm_ord - (Const(@{const_name Int.number_of}, _) $ v, Const(@{const_name Int.number_of}, _) $ w) = - num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w) - | numterm_ord (Const(@{const_name Int.number_of}, _) $ _, _) = LESS - | numterm_ord (_, Const(@{const_name Int.number_of}, _) $ _) = GREATER - | numterm_ord (t, u) = - (case int_ord (size_of_term t, size_of_term u) of - EQUAL => - let val (f, ts) = strip_comb t and (g, us) = strip_comb u in - (case hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord) - end - | ord => ord) -and numterms_ord (ts, us) = list_ord numterm_ord (ts, us) -end; - -fun numtermless tu = (numterm_ord tu = LESS); - -(*Defined in this file, but perhaps needed only for Int_Numeral_Base_Simprocs of type nat.*) -val num_ss = HOL_ss settermless numtermless; - - -(** Utilities **) - -fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n; - -fun find_first_numeral past (t::terms) = - ((snd (HOLogic.dest_number t), rev past @ terms) - handle TERM _ => find_first_numeral (t::past) terms) - | find_first_numeral past [] = raise TERM("find_first_numeral", []); - -val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; - -fun mk_minus t = - let val T = Term.fastype_of t - in Const (@{const_name HOL.uminus}, T --> T) $ t end; - -(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) -fun mk_sum T [] = mk_number T 0 - | mk_sum T [t,u] = mk_plus (t, u) - | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); - -(*this version ALWAYS includes a trailing zero*) -fun long_mk_sum T [] = mk_number T 0 - | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); - -val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} Term.dummyT; - -(*decompose additions AND subtractions as a sum*) -fun dest_summing (pos, Const (@{const_name HOL.plus}, _) $ t $ u, ts) = - dest_summing (pos, t, dest_summing (pos, u, ts)) - | dest_summing (pos, Const (@{const_name HOL.minus}, _) $ t $ u, ts) = - dest_summing (pos, t, dest_summing (not pos, u, ts)) - | dest_summing (pos, t, ts) = - if pos then t::ts else mk_minus t :: ts; - -fun dest_sum t = dest_summing (true, t, []); - -val mk_diff = HOLogic.mk_binop @{const_name HOL.minus}; -val dest_diff = HOLogic.dest_bin @{const_name HOL.minus} Term.dummyT; - -val mk_times = HOLogic.mk_binop @{const_name HOL.times}; - -fun one_of T = Const(@{const_name HOL.one},T); - -(* build product with trailing 1 rather than Numeral 1 in order to avoid the - unnecessary restriction to type class number_ring - which is not required for cancellation of common factors in divisions. -*) -fun mk_prod T = - let val one = one_of T - fun mk [] = one - | mk [t] = t - | mk (t :: ts) = if t = one then mk ts else mk_times (t, mk ts) - in mk end; - -(*This version ALWAYS includes a trailing one*) -fun long_mk_prod T [] = one_of T - | long_mk_prod T (t :: ts) = mk_times (t, mk_prod T ts); - -val dest_times = HOLogic.dest_bin @{const_name HOL.times} Term.dummyT; - -fun dest_prod t = - let val (t,u) = dest_times t - in dest_prod t @ dest_prod u end - handle TERM _ => [t]; - -(*DON'T do the obvious simplifications; that would create special cases*) -fun mk_coeff (k, t) = mk_times (mk_number (Term.fastype_of t) k, t); - -(*Express t as a product of (possibly) a numeral with other sorted terms*) -fun dest_coeff sign (Const (@{const_name HOL.uminus}, _) $ t) = dest_coeff (~sign) t - | dest_coeff sign t = - let val ts = sort Term.term_ord (dest_prod t) - val (n, ts') = find_first_numeral [] ts - handle TERM _ => (1, ts) - in (sign*n, mk_prod (Term.fastype_of t) ts') end; - -(*Find first coefficient-term THAT MATCHES u*) -fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) - | find_first_coeff past u (t::terms) = - let val (n,u') = dest_coeff 1 t - in if u aconv u' then (n, rev past @ terms) - else find_first_coeff (t::past) u terms - end - handle TERM _ => find_first_coeff (t::past) u terms; - -(*Fractions as pairs of ints. Can't use Rat.rat because the representation - needs to preserve negative values in the denominator.*) -fun mk_frac (p, q) = if q = 0 then raise Div else (p, q); - -(*Don't reduce fractions; sums must be proved by rule add_frac_eq. - Fractions are reduced later by the cancel_numeral_factor simproc.*) -fun add_frac ((p1, q1), (p2, q2)) = (p1 * q2 + p2 * q1, q1 * q2); - -val mk_divide = HOLogic.mk_binop @{const_name HOL.divide}; - -(*Build term (p / q) * t*) -fun mk_fcoeff ((p, q), t) = - let val T = Term.fastype_of t - in mk_times (mk_divide (mk_number T p, mk_number T q), t) end; - -(*Express t as a product of a fraction with other sorted terms*) -fun dest_fcoeff sign (Const (@{const_name HOL.uminus}, _) $ t) = dest_fcoeff (~sign) t - | dest_fcoeff sign (Const (@{const_name HOL.divide}, _) $ t $ u) = - let val (p, t') = dest_coeff sign t - val (q, u') = dest_coeff 1 u - in (mk_frac (p, q), mk_divide (t', u')) end - | dest_fcoeff sign t = - let val (p, t') = dest_coeff sign t - val T = Term.fastype_of t - in (mk_frac (p, 1), mk_divide (t', one_of T)) end; - - -(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1, 1*x, x*1, x/1 *) -val add_0s = thms "add_0s"; -val mult_1s = thms "mult_1s" @ [thm"mult_1_left", thm"mult_1_right", thm"divide_1"]; - -(*Simplify inverse Numeral1, a/Numeral1*) -val inverse_1s = [@{thm inverse_numeral_1}]; -val divide_1s = [@{thm divide_numeral_1}]; - -(*To perform binary arithmetic. The "left" rewriting handles patterns - created by the Int_Numeral_Base_Simprocs, such as 3 * (5 * x). *) -val simps = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym, - @{thm add_number_of_left}, @{thm mult_number_of_left}] @ - @{thms arith_simps} @ @{thms rel_simps}; - -(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms - during re-arrangement*) -val non_add_simps = - subtract Thm.eq_thm [@{thm add_number_of_left}, @{thm number_of_add} RS sym] simps; - -(*To evaluate binary negations of coefficients*) -val minus_simps = [@{thm numeral_m1_eq_minus_1} RS sym, @{thm number_of_minus} RS sym] @ - @{thms minus_bin_simps} @ @{thms pred_bin_simps}; - -(*To let us treat subtraction as addition*) -val diff_simps = [@{thm diff_minus}, @{thm minus_add_distrib}, @{thm minus_minus}]; - -(*To let us treat division as multiplication*) -val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}]; - -(*push the unary minus down: - x * y = x * - y *) -val minus_mult_eq_1_to_2 = - [@{thm minus_mult_left} RS sym, @{thm minus_mult_right}] MRS trans |> standard; - -(*to extract again any uncancelled minuses*) -val minus_from_mult_simps = - [@{thm minus_minus}, @{thm minus_mult_left} RS sym, @{thm minus_mult_right} RS sym]; - -(*combine unary minus with numeric literals, however nested within a product*) -val mult_minus_simps = - [@{thm mult_assoc}, @{thm minus_mult_left}, minus_mult_eq_1_to_2]; - -(*Apply the given rewrite (if present) just once*) -fun trans_tac NONE = all_tac - | trans_tac (SOME th) = ALLGOALS (rtac (th RS trans)); - -fun simplify_meta_eq rules = - let val ss0 = HOL_basic_ss addeqcongs [eq_cong2] addsimps rules - in fn ss => simplify (Simplifier.inherit_context ss ss0) o mk_meta_eq end - -structure CancelNumeralsCommon = - struct - val mk_sum = mk_sum - val dest_sum = dest_sum - val mk_coeff = mk_coeff - val dest_coeff = dest_coeff 1 - val find_first_coeff = find_first_coeff [] - val trans_tac = fn _ => trans_tac - - val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ - diff_simps @ minus_simps @ @{thms add_ac} - val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps - val norm_ss3 = num_ss addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) - - val numeral_simp_ss = HOL_ss addsimps add_0s @ simps - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s) - end; - - -structure EqCancelNumerals = CancelNumeralsFun - (open CancelNumeralsCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_eq - val dest_bal = HOLogic.dest_bin "op =" Term.dummyT - val bal_add1 = @{thm eq_add_iff1} RS trans - val bal_add2 = @{thm eq_add_iff2} RS trans -); - -structure LessCancelNumerals = CancelNumeralsFun - (open CancelNumeralsCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less} - val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT - val bal_add1 = @{thm less_add_iff1} RS trans - val bal_add2 = @{thm less_add_iff2} RS trans -); - -structure LeCancelNumerals = CancelNumeralsFun - (open CancelNumeralsCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq} - val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT - val bal_add1 = @{thm le_add_iff1} RS trans - val bal_add2 = @{thm le_add_iff2} RS trans -); - -val cancel_numerals = - map Int_Numeral_Base_Simprocs.prep_simproc - [("inteq_cancel_numerals", - ["(l::'a::number_ring) + m = n", - "(l::'a::number_ring) = m + n", - "(l::'a::number_ring) - m = n", - "(l::'a::number_ring) = m - n", - "(l::'a::number_ring) * m = n", - "(l::'a::number_ring) = m * n"], - K EqCancelNumerals.proc), - ("intless_cancel_numerals", - ["(l::'a::{ordered_idom,number_ring}) + m < n", - "(l::'a::{ordered_idom,number_ring}) < m + n", - "(l::'a::{ordered_idom,number_ring}) - m < n", - "(l::'a::{ordered_idom,number_ring}) < m - n", - "(l::'a::{ordered_idom,number_ring}) * m < n", - "(l::'a::{ordered_idom,number_ring}) < m * n"], - K LessCancelNumerals.proc), - ("intle_cancel_numerals", - ["(l::'a::{ordered_idom,number_ring}) + m <= n", - "(l::'a::{ordered_idom,number_ring}) <= m + n", - "(l::'a::{ordered_idom,number_ring}) - m <= n", - "(l::'a::{ordered_idom,number_ring}) <= m - n", - "(l::'a::{ordered_idom,number_ring}) * m <= n", - "(l::'a::{ordered_idom,number_ring}) <= m * n"], - K LeCancelNumerals.proc)]; - - -structure CombineNumeralsData = - struct - type coeff = int - val iszero = (fn x => x = 0) - val add = op + - val mk_sum = long_mk_sum (*to work for e.g. 2*x + 3*x *) - val dest_sum = dest_sum - val mk_coeff = mk_coeff - val dest_coeff = dest_coeff 1 - val left_distrib = @{thm combine_common_factor} RS trans - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv_nohyps - val trans_tac = fn _ => trans_tac - - val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ - diff_simps @ minus_simps @ @{thms add_ac} - val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps - val norm_ss3 = num_ss addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) - - val numeral_simp_ss = HOL_ss addsimps add_0s @ simps - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s) - end; - -structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); - -(*Version for fields, where coefficients can be fractions*) -structure FieldCombineNumeralsData = - struct - type coeff = int * int - val iszero = (fn (p, q) => p = 0) - val add = add_frac - val mk_sum = long_mk_sum - val dest_sum = dest_sum - val mk_coeff = mk_fcoeff - val dest_coeff = dest_fcoeff 1 - val left_distrib = @{thm combine_common_factor} RS trans - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv_nohyps - val trans_tac = fn _ => trans_tac - - val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ - inverse_1s @ divide_simps @ diff_simps @ minus_simps @ @{thms add_ac} - val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps - val norm_ss3 = num_ss addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) - - val numeral_simp_ss = HOL_ss addsimps add_0s @ simps @ [@{thm add_frac_eq}] - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s @ divide_1s) - end; - -structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData); - -val combine_numerals = - Int_Numeral_Base_Simprocs.prep_simproc - ("int_combine_numerals", - ["(i::'a::number_ring) + j", "(i::'a::number_ring) - j"], - K CombineNumerals.proc); - -val field_combine_numerals = - Int_Numeral_Base_Simprocs.prep_simproc - ("field_combine_numerals", - ["(i::'a::{number_ring,field,division_by_zero}) + j", - "(i::'a::{number_ring,field,division_by_zero}) - j"], - K FieldCombineNumerals.proc); - -end; - -Addsimprocs Int_Numeral_Simprocs.cancel_numerals; -Addsimprocs [Int_Numeral_Simprocs.combine_numerals]; -Addsimprocs [Int_Numeral_Simprocs.field_combine_numerals]; - -(*examples: -print_depth 22; -set timing; -set trace_simp; -fun test s = (Goal s, by (Simp_tac 1)); - -test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::int)"; - -test "2*u = (u::int)"; -test "(i + j + 12 + (k::int)) - 15 = y"; -test "(i + j + 12 + (k::int)) - 5 = y"; - -test "y - b < (b::int)"; -test "y - (3*b + c) < (b::int) - 2*c"; - -test "(2*x - (u*v) + y) - v*3*u = (w::int)"; -test "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::int)"; -test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::int)"; -test "u*v - (x*u*v + (u*v)*4 + y) = (w::int)"; - -test "(i + j + 12 + (k::int)) = u + 15 + y"; -test "(i + j*2 + 12 + (k::int)) = j + 5 + y"; - -test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::int)"; - -test "a + -(b+c) + b = (d::int)"; -test "a + -(b+c) - b = (d::int)"; - -(*negative numerals*) -test "(i + j + -2 + (k::int)) - (u + 5 + y) = zz"; -test "(i + j + -3 + (k::int)) < u + 5 + y"; -test "(i + j + 3 + (k::int)) < u + -6 + y"; -test "(i + j + -12 + (k::int)) - 15 = y"; -test "(i + j + 12 + (k::int)) - -15 = y"; -test "(i + j + -12 + (k::int)) - -15 = y"; -*) - - -(** Constant folding for multiplication in semirings **) - -(*We do not need folding for addition: combine_numerals does the same thing*) - -structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA = -struct - val assoc_ss = HOL_ss addsimps @{thms mult_ac} - val eq_reflection = eq_reflection -end; - -structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data); - -val assoc_fold_simproc = - Int_Numeral_Base_Simprocs.prep_simproc - ("semiring_assoc_fold", ["(a::'a::comm_semiring_1_cancel) * b"], - K Semiring_Times_Assoc.proc); - -Addsimprocs [assoc_fold_simproc]; - - - - -(*** decision procedure for linear arithmetic ***) - -(*---------------------------------------------------------------------------*) -(* Linear arithmetic *) -(*---------------------------------------------------------------------------*) - -(* -Instantiation of the generic linear arithmetic package for int. -*) - -(* Update parameters of arithmetic prover *) -local - -(* reduce contradictory =/ if s = @{const_name "HOL.one"} then not (t = @{typ int}) - else s mem_string allowed_consts - | a$b => check a andalso check b - | _ => false; - -val conv = - Simplifier.rewrite - (HOL_basic_ss addsimps - ((map (fn th => th RS sym) [@{thm of_int_add}, @{thm of_int_mult}, - @{thm of_int_diff}, @{thm of_int_minus}])@ - [@{thm of_int_less_iff}, @{thm of_int_le_iff}, @{thm of_int_eq_iff}]) - addsimprocs [zero_to_of_int_zero_simproc,one_to_of_int_one_simproc]); - -fun sproc phi ss ct = if check (term_of ct) then SOME (conv ct) else NONE -val lhss' = - [@{cpat "(?x::?'a::ring_char_0) = (?y::?'a)"}, - @{cpat "(?x::?'a::ordered_idom) < (?y::?'a)"}, - @{cpat "(?x::?'a::ordered_idom) <= (?y::?'a)"}] -in -val zero_one_idom_simproc = - make_simproc {lhss = lhss' , name = "zero_one_idom_simproc", - proc = sproc, identifier = []} -end; - -val add_rules = - simp_thms @ @{thms arith_simps} @ @{thms rel_simps} @ @{thms arith_special} @ - [@{thm neg_le_iff_le}, @{thm numeral_0_eq_0}, @{thm numeral_1_eq_1}, - @{thm minus_zero}, @{thm diff_minus}, @{thm left_minus}, @{thm right_minus}, - @{thm mult_zero_left}, @{thm mult_zero_right}, @{thm mult_Bit1}, @{thm mult_1_right}, - @{thm minus_mult_left} RS sym, @{thm minus_mult_right} RS sym, - @{thm minus_add_distrib}, @{thm minus_minus}, @{thm mult_assoc}, - @{thm of_nat_0}, @{thm of_nat_1}, @{thm of_nat_Suc}, @{thm of_nat_add}, - @{thm of_nat_mult}, @{thm of_int_0}, @{thm of_int_1}, @{thm of_int_add}, - @{thm of_int_mult}] - -val nat_inj_thms = [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2] - -val Int_Numeral_Base_Simprocs = assoc_fold_simproc :: zero_one_idom_simproc - :: Int_Numeral_Simprocs.combine_numerals - :: Int_Numeral_Simprocs.cancel_numerals; - -in - -val int_arith_setup = - LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} => - {add_mono_thms = add_mono_thms, - mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono} :: mult_mono_thms, - inj_thms = nat_inj_thms @ inj_thms, - lessD = lessD @ [@{thm zless_imp_add1_zle}], - neqE = neqE, - simpset = simpset addsimps add_rules - addsimprocs Int_Numeral_Base_Simprocs - addcongs [if_weak_cong]}) #> - arith_inj_const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT) #> - arith_discrete @{type_name Int.int} - -end; - -val fast_int_arith_simproc = - Simplifier.simproc (the_context ()) - "fast_int_arith" - ["(m::'a::{ordered_idom,number_ring}) < n", - "(m::'a::{ordered_idom,number_ring}) <= n", - "(m::'a::{ordered_idom,number_ring}) = n"] (K LinArith.lin_arith_simproc); - -Addsimprocs [fast_int_arith_simproc]; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/int_factor_simprocs.ML --- a/src/HOL/int_factor_simprocs.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,340 +0,0 @@ -(* Title: HOL/int_factor_simprocs.ML - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 2000 University of Cambridge - -Factor cancellation simprocs for the integers (and for fields). - -This file can't be combined with int_arith1 because it requires IntDiv.thy. -*) - - -(*To quote from Provers/Arith/cancel_numeral_factor.ML: - -Cancels common coefficients in balanced expressions: - - u*#m ~~ u'*#m' == #n*u ~~ #n'*u' - -where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /) -and d = gcd(m,m') and n=m/d and n'=m'/d. -*) - -val rel_number_of = [@{thm eq_number_of_eq}, @{thm less_number_of_eq_neg}, @{thm le_number_of_eq}]; - -local - open Int_Numeral_Simprocs -in - -structure CancelNumeralFactorCommon = - struct - val mk_coeff = mk_coeff - val dest_coeff = dest_coeff 1 - val trans_tac = fn _ => trans_tac - - val norm_ss1 = HOL_ss addsimps minus_from_mult_simps @ mult_1s - val norm_ss2 = HOL_ss addsimps simps @ mult_minus_simps - val norm_ss3 = HOL_ss addsimps @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) - - val numeral_simp_ss = HOL_ss addsimps rel_number_of @ simps - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - val simplify_meta_eq = Int_Numeral_Simprocs.simplify_meta_eq - [@{thm add_0}, @{thm add_0_right}, @{thm mult_zero_left}, - @{thm mult_zero_right}, @{thm mult_Bit1}, @{thm mult_1_right}]; - end - -(*Version for integer division*) -structure IntDivCancelNumeralFactor = CancelNumeralFactorFun - (open CancelNumeralFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binop @{const_name Divides.div} - val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT - val cancel = @{thm zdiv_zmult_zmult1} RS trans - val neg_exchanges = false -) - -(*Version for fields*) -structure DivideCancelNumeralFactor = CancelNumeralFactorFun - (open CancelNumeralFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binop @{const_name HOL.divide} - val dest_bal = HOLogic.dest_bin @{const_name HOL.divide} Term.dummyT - val cancel = @{thm mult_divide_mult_cancel_left} RS trans - val neg_exchanges = false -) - -structure EqCancelNumeralFactor = CancelNumeralFactorFun - (open CancelNumeralFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_eq - val dest_bal = HOLogic.dest_bin "op =" Term.dummyT - val cancel = @{thm mult_cancel_left} RS trans - val neg_exchanges = false -) - -structure LessCancelNumeralFactor = CancelNumeralFactorFun - (open CancelNumeralFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less} - val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT - val cancel = @{thm mult_less_cancel_left} RS trans - val neg_exchanges = true -) - -structure LeCancelNumeralFactor = CancelNumeralFactorFun - (open CancelNumeralFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq} - val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT - val cancel = @{thm mult_le_cancel_left} RS trans - val neg_exchanges = true -) - -val cancel_numeral_factors = - map Int_Numeral_Base_Simprocs.prep_simproc - [("ring_eq_cancel_numeral_factor", - ["(l::'a::{idom,number_ring}) * m = n", - "(l::'a::{idom,number_ring}) = m * n"], - K EqCancelNumeralFactor.proc), - ("ring_less_cancel_numeral_factor", - ["(l::'a::{ordered_idom,number_ring}) * m < n", - "(l::'a::{ordered_idom,number_ring}) < m * n"], - K LessCancelNumeralFactor.proc), - ("ring_le_cancel_numeral_factor", - ["(l::'a::{ordered_idom,number_ring}) * m <= n", - "(l::'a::{ordered_idom,number_ring}) <= m * n"], - K LeCancelNumeralFactor.proc), - ("int_div_cancel_numeral_factors", - ["((l::int) * m) div n", "(l::int) div (m * n)"], - K IntDivCancelNumeralFactor.proc), - ("divide_cancel_numeral_factor", - ["((l::'a::{division_by_zero,field,number_ring}) * m) / n", - "(l::'a::{division_by_zero,field,number_ring}) / (m * n)", - "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"], - K DivideCancelNumeralFactor.proc)]; - -(* referenced by rat_arith.ML *) -val field_cancel_numeral_factors = - map Int_Numeral_Base_Simprocs.prep_simproc - [("field_eq_cancel_numeral_factor", - ["(l::'a::{field,number_ring}) * m = n", - "(l::'a::{field,number_ring}) = m * n"], - K EqCancelNumeralFactor.proc), - ("field_cancel_numeral_factor", - ["((l::'a::{division_by_zero,field,number_ring}) * m) / n", - "(l::'a::{division_by_zero,field,number_ring}) / (m * n)", - "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"], - K DivideCancelNumeralFactor.proc)] - -end; - -Addsimprocs cancel_numeral_factors; - -(*examples: -print_depth 22; -set timing; -set trace_simp; -fun test s = (Goal s; by (Simp_tac 1)); - -test "9*x = 12 * (y::int)"; -test "(9*x) div (12 * (y::int)) = z"; -test "9*x < 12 * (y::int)"; -test "9*x <= 12 * (y::int)"; - -test "-99*x = 132 * (y::int)"; -test "(-99*x) div (132 * (y::int)) = z"; -test "-99*x < 132 * (y::int)"; -test "-99*x <= 132 * (y::int)"; - -test "999*x = -396 * (y::int)"; -test "(999*x) div (-396 * (y::int)) = z"; -test "999*x < -396 * (y::int)"; -test "999*x <= -396 * (y::int)"; - -test "-99*x = -81 * (y::int)"; -test "(-99*x) div (-81 * (y::int)) = z"; -test "-99*x <= -81 * (y::int)"; -test "-99*x < -81 * (y::int)"; - -test "-2 * x = -1 * (y::int)"; -test "-2 * x = -(y::int)"; -test "(-2 * x) div (-1 * (y::int)) = z"; -test "-2 * x < -(y::int)"; -test "-2 * x <= -1 * (y::int)"; -test "-x < -23 * (y::int)"; -test "-x <= -23 * (y::int)"; -*) - -(*And the same examples for fields such as rat or real: -test "0 <= (y::rat) * -2"; -test "9*x = 12 * (y::rat)"; -test "(9*x) / (12 * (y::rat)) = z"; -test "9*x < 12 * (y::rat)"; -test "9*x <= 12 * (y::rat)"; - -test "-99*x = 132 * (y::rat)"; -test "(-99*x) / (132 * (y::rat)) = z"; -test "-99*x < 132 * (y::rat)"; -test "-99*x <= 132 * (y::rat)"; - -test "999*x = -396 * (y::rat)"; -test "(999*x) / (-396 * (y::rat)) = z"; -test "999*x < -396 * (y::rat)"; -test "999*x <= -396 * (y::rat)"; - -test "(- ((2::rat) * x) <= 2 * y)"; -test "-99*x = -81 * (y::rat)"; -test "(-99*x) / (-81 * (y::rat)) = z"; -test "-99*x <= -81 * (y::rat)"; -test "-99*x < -81 * (y::rat)"; - -test "-2 * x = -1 * (y::rat)"; -test "-2 * x = -(y::rat)"; -test "(-2 * x) / (-1 * (y::rat)) = z"; -test "-2 * x < -(y::rat)"; -test "-2 * x <= -1 * (y::rat)"; -test "-x < -23 * (y::rat)"; -test "-x <= -23 * (y::rat)"; -*) - - -(** Declarations for ExtractCommonTerm **) - -local - open Int_Numeral_Simprocs -in - -(*Find first term that matches u*) -fun find_first_t past u [] = raise TERM ("find_first_t", []) - | find_first_t past u (t::terms) = - if u aconv t then (rev past @ terms) - else find_first_t (t::past) u terms - handle TERM _ => find_first_t (t::past) u terms; - -(** Final simplification for the CancelFactor simprocs **) -val simplify_one = Int_Numeral_Simprocs.simplify_meta_eq - [@{thm mult_1_left}, @{thm mult_1_right}, @{thm zdiv_1}, @{thm numeral_1_eq_1}]; - -fun cancel_simplify_meta_eq cancel_th ss th = - simplify_one ss (([th, cancel_th]) MRS trans); - -structure CancelFactorCommon = - struct - val mk_sum = long_mk_prod - val dest_sum = dest_prod - val mk_coeff = mk_coeff - val dest_coeff = dest_coeff - val find_first = find_first_t [] - val trans_tac = fn _ => trans_tac - val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac} - fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) - end; - -(*mult_cancel_left requires a ring with no zero divisors.*) -structure EqCancelFactor = ExtractCommonTermFun - (open CancelFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_eq - val dest_bal = HOLogic.dest_bin "op =" Term.dummyT - val simplify_meta_eq = cancel_simplify_meta_eq @{thm mult_cancel_left} -); - -(*zdiv_zmult_zmult1_if is for integer division (div).*) -structure IntDivCancelFactor = ExtractCommonTermFun - (open CancelFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binop @{const_name Divides.div} - val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT - val simplify_meta_eq = cancel_simplify_meta_eq @{thm zdiv_zmult_zmult1_if} -); - -structure IntModCancelFactor = ExtractCommonTermFun - (open CancelFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binop @{const_name Divides.mod} - val dest_bal = HOLogic.dest_bin @{const_name Divides.mod} HOLogic.intT - val simplify_meta_eq = cancel_simplify_meta_eq @{thm zmod_zmult_zmult1} -); - -structure IntDvdCancelFactor = ExtractCommonTermFun - (open CancelFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} - val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.intT - val simplify_meta_eq = cancel_simplify_meta_eq @{thm zdvd_zmult_cancel_disj} -); - -(*Version for all fields, including unordered ones (type complex).*) -structure DivideCancelFactor = ExtractCommonTermFun - (open CancelFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binop @{const_name HOL.divide} - val dest_bal = HOLogic.dest_bin @{const_name HOL.divide} Term.dummyT - val simplify_meta_eq = cancel_simplify_meta_eq @{thm mult_divide_mult_cancel_left_if} -); - -val cancel_factors = - map Int_Numeral_Base_Simprocs.prep_simproc - [("ring_eq_cancel_factor", - ["(l::'a::{idom}) * m = n", - "(l::'a::{idom}) = m * n"], - K EqCancelFactor.proc), - ("int_div_cancel_factor", - ["((l::int) * m) div n", "(l::int) div (m * n)"], - K IntDivCancelFactor.proc), - ("int_mod_cancel_factor", - ["((l::int) * m) mod n", "(l::int) mod (m * n)"], - K IntModCancelFactor.proc), - ("int_dvd_cancel_factor", - ["((l::int) * m) dvd n", "(l::int) dvd (m * n)"], - K IntDvdCancelFactor.proc), - ("divide_cancel_factor", - ["((l::'a::{division_by_zero,field}) * m) / n", - "(l::'a::{division_by_zero,field}) / (m * n)"], - K DivideCancelFactor.proc)]; - -end; - -Addsimprocs cancel_factors; - - -(*examples: -print_depth 22; -set timing; -set trace_simp; -fun test s = (Goal s; by (Asm_simp_tac 1)); - -test "x*k = k*(y::int)"; -test "k = k*(y::int)"; -test "a*(b*c) = (b::int)"; -test "a*(b*c) = d*(b::int)*(x*a)"; - -test "(x*k) div (k*(y::int)) = (uu::int)"; -test "(k) div (k*(y::int)) = (uu::int)"; -test "(a*(b*c)) div ((b::int)) = (uu::int)"; -test "(a*(b*c)) div (d*(b::int)*(x*a)) = (uu::int)"; -*) - -(*And the same examples for fields such as rat or real: -print_depth 22; -set timing; -set trace_simp; -fun test s = (Goal s; by (Asm_simp_tac 1)); - -test "x*k = k*(y::rat)"; -test "k = k*(y::rat)"; -test "a*(b*c) = (b::rat)"; -test "a*(b*c) = d*(b::rat)*(x*a)"; - - -test "(x*k) / (k*(y::rat)) = (uu::rat)"; -test "(k) / (k*(y::rat)) = (uu::rat)"; -test "(a*(b*c)) / ((b::rat)) = (uu::rat)"; -test "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)"; - -(*FIXME: what do we do about this?*) -test "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z"; -*) diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/nat_simprocs.ML --- a/src/HOL/nat_simprocs.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,585 +0,0 @@ -(* Title: HOL/nat_simprocs.ML - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 2000 University of Cambridge - -Simprocs for nat numerals. -*) - -structure Nat_Numeral_Simprocs = -struct - -(*Maps n to #n for n = 0, 1, 2*) -val numeral_syms = - [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym, @{thm numeral_2_eq_2} RS sym]; -val numeral_sym_ss = HOL_ss addsimps numeral_syms; - -fun rename_numerals th = - simplify numeral_sym_ss (Thm.transfer (the_context ()) th); - -(*Utilities*) - -fun mk_number n = HOLogic.number_of_const HOLogic.natT $ HOLogic.mk_numeral n; -fun dest_number t = Int.max (0, snd (HOLogic.dest_number t)); - -fun find_first_numeral past (t::terms) = - ((dest_number t, t, rev past @ terms) - handle TERM _ => find_first_numeral (t::past) terms) - | find_first_numeral past [] = raise TERM("find_first_numeral", []); - -val zero = mk_number 0; -val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; - -(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) -fun mk_sum [] = zero - | mk_sum [t,u] = mk_plus (t, u) - | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); - -(*this version ALWAYS includes a trailing zero*) -fun long_mk_sum [] = HOLogic.zero - | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts); - -val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} HOLogic.natT; - - -(** Other simproc items **) - -val trans_tac = Int_Numeral_Simprocs.trans_tac; - -val bin_simps = - [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym, - @{thm add_nat_number_of}, @{thm nat_number_of_add_left}, - @{thm diff_nat_number_of}, @{thm le_number_of_eq_not_less}, - @{thm mult_nat_number_of}, @{thm nat_number_of_mult_left}, - @{thm less_nat_number_of}, - @{thm Let_number_of}, @{thm nat_number_of}] @ - @{thms arith_simps} @ @{thms rel_simps}; - -fun prep_simproc (name, pats, proc) = - Simplifier.simproc (the_context ()) name pats proc; - - -(*** CancelNumerals simprocs ***) - -val one = mk_number 1; -val mk_times = HOLogic.mk_binop @{const_name HOL.times}; - -fun mk_prod [] = one - | mk_prod [t] = t - | mk_prod (t :: ts) = if t = one then mk_prod ts - else mk_times (t, mk_prod ts); - -val dest_times = HOLogic.dest_bin @{const_name HOL.times} HOLogic.natT; - -fun dest_prod t = - let val (t,u) = dest_times t - in dest_prod t @ dest_prod u end - handle TERM _ => [t]; - -(*DON'T do the obvious simplifications; that would create special cases*) -fun mk_coeff (k,t) = mk_times (mk_number k, t); - -(*Express t as a product of (possibly) a numeral with other factors, sorted*) -fun dest_coeff t = - let val ts = sort Term.term_ord (dest_prod t) - val (n, _, ts') = find_first_numeral [] ts - handle TERM _ => (1, one, ts) - in (n, mk_prod ts') end; - -(*Find first coefficient-term THAT MATCHES u*) -fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) - | find_first_coeff past u (t::terms) = - let val (n,u') = dest_coeff t - in if u aconv u' then (n, rev past @ terms) - else find_first_coeff (t::past) u terms - end - handle TERM _ => find_first_coeff (t::past) u terms; - - -(*Split up a sum into the list of its constituent terms, on the way removing any - Sucs and counting them.*) -fun dest_Suc_sum (Const ("Suc", _) $ t, (k,ts)) = dest_Suc_sum (t, (k+1,ts)) - | dest_Suc_sum (t, (k,ts)) = - let val (t1,t2) = dest_plus t - in dest_Suc_sum (t1, dest_Suc_sum (t2, (k,ts))) end - handle TERM _ => (k, t::ts); - -(*Code for testing whether numerals are already used in the goal*) -fun is_numeral (Const(@{const_name Int.number_of}, _) $ w) = true - | is_numeral _ = false; - -fun prod_has_numeral t = exists is_numeral (dest_prod t); - -(*The Sucs found in the term are converted to a binary numeral. If relaxed is false, - an exception is raised unless the original expression contains at least one - numeral in a coefficient position. This prevents nat_combine_numerals from - introducing numerals to goals.*) -fun dest_Sucs_sum relaxed t = - let val (k,ts) = dest_Suc_sum (t,(0,[])) - in - if relaxed orelse exists prod_has_numeral ts then - if k=0 then ts - else mk_number k :: ts - else raise TERM("Nat_Numeral_Simprocs.dest_Sucs_sum", [t]) - end; - - -(*Simplify 1*n and n*1 to n*) -val add_0s = map rename_numerals [@{thm add_0}, @{thm add_0_right}]; -val mult_1s = map rename_numerals [@{thm nat_mult_1}, @{thm nat_mult_1_right}]; - -(*Final simplification: cancel + and *; replace Numeral0 by 0 and Numeral1 by 1*) - -(*And these help the simproc return False when appropriate, which helps - the arith prover.*) -val contra_rules = [@{thm add_Suc}, @{thm add_Suc_right}, @{thm Zero_not_Suc}, - @{thm Suc_not_Zero}, @{thm le_0_eq}]; - -val simplify_meta_eq = - Int_Numeral_Simprocs.simplify_meta_eq - ([@{thm nat_numeral_0_eq_0}, @{thm numeral_1_eq_Suc_0}, @{thm add_0}, @{thm add_0_right}, - @{thm mult_0}, @{thm mult_0_right}, @{thm mult_1}, @{thm mult_1_right}] @ contra_rules); - - -(*Like HOL_ss but with an ordering that brings numerals to the front - under AC-rewriting.*) -val num_ss = Int_Numeral_Simprocs.num_ss; - -(*** Applying CancelNumeralsFun ***) - -structure CancelNumeralsCommon = - struct - val mk_sum = (fn T:typ => mk_sum) - val dest_sum = dest_Sucs_sum true - val mk_coeff = mk_coeff - val dest_coeff = dest_coeff - val find_first_coeff = find_first_coeff [] - val trans_tac = fn _ => trans_tac - - val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ - [@{thm Suc_eq_add_numeral_1_left}] @ @{thms add_ac} - val norm_ss2 = num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - - val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps; - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)); - val simplify_meta_eq = simplify_meta_eq - end; - - -structure EqCancelNumerals = CancelNumeralsFun - (open CancelNumeralsCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_eq - val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT - val bal_add1 = @{thm nat_eq_add_iff1} RS trans - val bal_add2 = @{thm nat_eq_add_iff2} RS trans -); - -structure LessCancelNumerals = CancelNumeralsFun - (open CancelNumeralsCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less} - val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT - val bal_add1 = @{thm nat_less_add_iff1} RS trans - val bal_add2 = @{thm nat_less_add_iff2} RS trans -); - -structure LeCancelNumerals = CancelNumeralsFun - (open CancelNumeralsCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq} - val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT - val bal_add1 = @{thm nat_le_add_iff1} RS trans - val bal_add2 = @{thm nat_le_add_iff2} RS trans -); - -structure DiffCancelNumerals = CancelNumeralsFun - (open CancelNumeralsCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binop @{const_name HOL.minus} - val dest_bal = HOLogic.dest_bin @{const_name HOL.minus} HOLogic.natT - val bal_add1 = @{thm nat_diff_add_eq1} RS trans - val bal_add2 = @{thm nat_diff_add_eq2} RS trans -); - - -val cancel_numerals = - map prep_simproc - [("nateq_cancel_numerals", - ["(l::nat) + m = n", "(l::nat) = m + n", - "(l::nat) * m = n", "(l::nat) = m * n", - "Suc m = n", "m = Suc n"], - K EqCancelNumerals.proc), - ("natless_cancel_numerals", - ["(l::nat) + m < n", "(l::nat) < m + n", - "(l::nat) * m < n", "(l::nat) < m * n", - "Suc m < n", "m < Suc n"], - K LessCancelNumerals.proc), - ("natle_cancel_numerals", - ["(l::nat) + m <= n", "(l::nat) <= m + n", - "(l::nat) * m <= n", "(l::nat) <= m * n", - "Suc m <= n", "m <= Suc n"], - K LeCancelNumerals.proc), - ("natdiff_cancel_numerals", - ["((l::nat) + m) - n", "(l::nat) - (m + n)", - "(l::nat) * m - n", "(l::nat) - m * n", - "Suc m - n", "m - Suc n"], - K DiffCancelNumerals.proc)]; - - -(*** Applying CombineNumeralsFun ***) - -structure CombineNumeralsData = - struct - type coeff = int - val iszero = (fn x => x = 0) - val add = op + - val mk_sum = (fn T:typ => long_mk_sum) (*to work for 2*x + 3*x *) - val dest_sum = dest_Sucs_sum false - val mk_coeff = mk_coeff - val dest_coeff = dest_coeff - val left_distrib = @{thm left_add_mult_distrib} RS trans - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv_nohyps - val trans_tac = fn _ => trans_tac - - val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_add_numeral_1}] @ @{thms add_ac} - val norm_ss2 = num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - - val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps; - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - val simplify_meta_eq = simplify_meta_eq - end; - -structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); - -val combine_numerals = - prep_simproc ("nat_combine_numerals", ["(i::nat) + j", "Suc (i + j)"], K CombineNumerals.proc); - - -(*** Applying CancelNumeralFactorFun ***) - -structure CancelNumeralFactorCommon = - struct - val mk_coeff = mk_coeff - val dest_coeff = dest_coeff - val trans_tac = fn _ => trans_tac - - val norm_ss1 = num_ss addsimps - numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_add_numeral_1_left}] @ @{thms add_ac} - val norm_ss2 = num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - - val numeral_simp_ss = HOL_ss addsimps bin_simps - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - val simplify_meta_eq = simplify_meta_eq - end - -structure DivCancelNumeralFactor = CancelNumeralFactorFun - (open CancelNumeralFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binop @{const_name Divides.div} - val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT - val cancel = @{thm nat_mult_div_cancel1} RS trans - val neg_exchanges = false -) - -structure DvdCancelNumeralFactor = CancelNumeralFactorFun - (open CancelNumeralFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} - val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT - val cancel = @{thm nat_mult_dvd_cancel1} RS trans - val neg_exchanges = false -) - -structure EqCancelNumeralFactor = CancelNumeralFactorFun - (open CancelNumeralFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_eq - val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT - val cancel = @{thm nat_mult_eq_cancel1} RS trans - val neg_exchanges = false -) - -structure LessCancelNumeralFactor = CancelNumeralFactorFun - (open CancelNumeralFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less} - val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT - val cancel = @{thm nat_mult_less_cancel1} RS trans - val neg_exchanges = true -) - -structure LeCancelNumeralFactor = CancelNumeralFactorFun - (open CancelNumeralFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq} - val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT - val cancel = @{thm nat_mult_le_cancel1} RS trans - val neg_exchanges = true -) - -val cancel_numeral_factors = - map prep_simproc - [("nateq_cancel_numeral_factors", - ["(l::nat) * m = n", "(l::nat) = m * n"], - K EqCancelNumeralFactor.proc), - ("natless_cancel_numeral_factors", - ["(l::nat) * m < n", "(l::nat) < m * n"], - K LessCancelNumeralFactor.proc), - ("natle_cancel_numeral_factors", - ["(l::nat) * m <= n", "(l::nat) <= m * n"], - K LeCancelNumeralFactor.proc), - ("natdiv_cancel_numeral_factors", - ["((l::nat) * m) div n", "(l::nat) div (m * n)"], - K DivCancelNumeralFactor.proc), - ("natdvd_cancel_numeral_factors", - ["((l::nat) * m) dvd n", "(l::nat) dvd (m * n)"], - K DvdCancelNumeralFactor.proc)]; - - - -(*** Applying ExtractCommonTermFun ***) - -(*this version ALWAYS includes a trailing one*) -fun long_mk_prod [] = one - | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts); - -(*Find first term that matches u*) -fun find_first_t past u [] = raise TERM("find_first_t", []) - | find_first_t past u (t::terms) = - if u aconv t then (rev past @ terms) - else find_first_t (t::past) u terms - handle TERM _ => find_first_t (t::past) u terms; - -(** Final simplification for the CancelFactor simprocs **) -val simplify_one = Int_Numeral_Simprocs.simplify_meta_eq - [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_1}, @{thm numeral_1_eq_Suc_0}]; - -fun cancel_simplify_meta_eq cancel_th ss th = - simplify_one ss (([th, cancel_th]) MRS trans); - -structure CancelFactorCommon = - struct - val mk_sum = (fn T:typ => long_mk_prod) - val dest_sum = dest_prod - val mk_coeff = mk_coeff - val dest_coeff = dest_coeff - val find_first = find_first_t [] - val trans_tac = fn _ => trans_tac - val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac} - fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) - end; - -structure EqCancelFactor = ExtractCommonTermFun - (open CancelFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_eq - val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT - val simplify_meta_eq = cancel_simplify_meta_eq @{thm nat_mult_eq_cancel_disj} -); - -structure LessCancelFactor = ExtractCommonTermFun - (open CancelFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less} - val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT - val simplify_meta_eq = cancel_simplify_meta_eq @{thm nat_mult_less_cancel_disj} -); - -structure LeCancelFactor = ExtractCommonTermFun - (open CancelFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq} - val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT - val simplify_meta_eq = cancel_simplify_meta_eq @{thm nat_mult_le_cancel_disj} -); - -structure DivideCancelFactor = ExtractCommonTermFun - (open CancelFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binop @{const_name Divides.div} - val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT - val simplify_meta_eq = cancel_simplify_meta_eq @{thm nat_mult_div_cancel_disj} -); - -structure DvdCancelFactor = ExtractCommonTermFun - (open CancelFactorCommon - val prove_conv = Int_Numeral_Base_Simprocs.prove_conv - val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} - val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT - val simplify_meta_eq = cancel_simplify_meta_eq @{thm nat_mult_dvd_cancel_disj} -); - -val cancel_factor = - map prep_simproc - [("nat_eq_cancel_factor", - ["(l::nat) * m = n", "(l::nat) = m * n"], - K EqCancelFactor.proc), - ("nat_less_cancel_factor", - ["(l::nat) * m < n", "(l::nat) < m * n"], - K LessCancelFactor.proc), - ("nat_le_cancel_factor", - ["(l::nat) * m <= n", "(l::nat) <= m * n"], - K LeCancelFactor.proc), - ("nat_divide_cancel_factor", - ["((l::nat) * m) div n", "(l::nat) div (m * n)"], - K DivideCancelFactor.proc), - ("nat_dvd_cancel_factor", - ["((l::nat) * m) dvd n", "(l::nat) dvd (m * n)"], - K DvdCancelFactor.proc)]; - -end; - - -Addsimprocs Nat_Numeral_Simprocs.cancel_numerals; -Addsimprocs [Nat_Numeral_Simprocs.combine_numerals]; -Addsimprocs Nat_Numeral_Simprocs.cancel_numeral_factors; -Addsimprocs Nat_Numeral_Simprocs.cancel_factor; - - -(*examples: -print_depth 22; -set timing; -set trace_simp; -fun test s = (Goal s; by (Simp_tac 1)); - -(*cancel_numerals*) -test "l +( 2) + (2) + 2 + (l + 2) + (oo + 2) = (uu::nat)"; -test "(2*length xs < 2*length xs + j)"; -test "(2*length xs < length xs * 2 + j)"; -test "2*u = (u::nat)"; -test "2*u = Suc (u)"; -test "(i + j + 12 + (k::nat)) - 15 = y"; -test "(i + j + 12 + (k::nat)) - 5 = y"; -test "Suc u - 2 = y"; -test "Suc (Suc (Suc u)) - 2 = y"; -test "(i + j + 2 + (k::nat)) - 1 = y"; -test "(i + j + 1 + (k::nat)) - 2 = y"; - -test "(2*x + (u*v) + y) - v*3*u = (w::nat)"; -test "(2*x*u*v + 5 + (u*v)*4 + y) - v*u*4 = (w::nat)"; -test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::nat)"; -test "Suc (Suc (2*x*u*v + u*4 + y)) - u = w"; -test "Suc ((u*v)*4) - v*3*u = w"; -test "Suc (Suc ((u*v)*3)) - v*3*u = w"; - -test "(i + j + 12 + (k::nat)) = u + 15 + y"; -test "(i + j + 32 + (k::nat)) - (u + 15 + y) = zz"; -test "(i + j + 12 + (k::nat)) = u + 5 + y"; -(*Suc*) -test "(i + j + 12 + k) = Suc (u + y)"; -test "Suc (Suc (Suc (Suc (Suc (u + y))))) <= ((i + j) + 41 + k)"; -test "(i + j + 5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))"; -test "Suc (Suc (Suc (Suc (Suc (u + y))))) - 5 = v"; -test "(i + j + 5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))"; -test "2*y + 3*z + 2*u = Suc (u)"; -test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = Suc (u)"; -test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::nat)"; -test "6 + 2*y + 3*z + 4*u = Suc (vv + 2*u + z)"; -test "(2*n*m) < (3*(m*n)) + (u::nat)"; - -test "(Suc (Suc (Suc (Suc (Suc (Suc (case length (f c) of 0 => 0 | Suc k => k)))))) <= Suc 0)"; - -test "Suc (Suc (Suc (Suc (Suc (Suc (length l1 + length l2)))))) <= length l1"; - -test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length l3)))))) <= length (compT P E A ST mxr e))"; - -test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length (compT P E (A Un \ e) ST mxr c))))))) <= length (compT P E A ST mxr e))"; - - -(*negative numerals: FAIL*) -test "(i + j + -23 + (k::nat)) < u + 15 + y"; -test "(i + j + 3 + (k::nat)) < u + -15 + y"; -test "(i + j + -12 + (k::nat)) - 15 = y"; -test "(i + j + 12 + (k::nat)) - -15 = y"; -test "(i + j + -12 + (k::nat)) - -15 = y"; - -(*combine_numerals*) -test "k + 3*k = (u::nat)"; -test "Suc (i + 3) = u"; -test "Suc (i + j + 3 + k) = u"; -test "k + j + 3*k + j = (u::nat)"; -test "Suc (j*i + i + k + 5 + 3*k + i*j*4) = (u::nat)"; -test "(2*n*m) + (3*(m*n)) = (u::nat)"; -(*negative numerals: FAIL*) -test "Suc (i + j + -3 + k) = u"; - -(*cancel_numeral_factors*) -test "9*x = 12 * (y::nat)"; -test "(9*x) div (12 * (y::nat)) = z"; -test "9*x < 12 * (y::nat)"; -test "9*x <= 12 * (y::nat)"; - -(*cancel_factor*) -test "x*k = k*(y::nat)"; -test "k = k*(y::nat)"; -test "a*(b*c) = (b::nat)"; -test "a*(b*c) = d*(b::nat)*(x*a)"; - -test "x*k < k*(y::nat)"; -test "k < k*(y::nat)"; -test "a*(b*c) < (b::nat)"; -test "a*(b*c) < d*(b::nat)*(x*a)"; - -test "x*k <= k*(y::nat)"; -test "k <= k*(y::nat)"; -test "a*(b*c) <= (b::nat)"; -test "a*(b*c) <= d*(b::nat)*(x*a)"; - -test "(x*k) div (k*(y::nat)) = (uu::nat)"; -test "(k) div (k*(y::nat)) = (uu::nat)"; -test "(a*(b*c)) div ((b::nat)) = (uu::nat)"; -test "(a*(b*c)) div (d*(b::nat)*(x*a)) = (uu::nat)"; -*) - - -(*** Prepare linear arithmetic for nat numerals ***) - -local - -(* reduce contradictory <= to False *) -val add_rules = @{thms ring_distribs} @ - [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, @{thm nat_0}, @{thm nat_1}, - @{thm add_nat_number_of}, @{thm diff_nat_number_of}, @{thm mult_nat_number_of}, - @{thm eq_nat_number_of}, @{thm less_nat_number_of}, @{thm le_number_of_eq_not_less}, - @{thm le_Suc_number_of}, @{thm le_number_of_Suc}, - @{thm less_Suc_number_of}, @{thm less_number_of_Suc}, - @{thm Suc_eq_number_of}, @{thm eq_number_of_Suc}, - @{thm mult_Suc}, @{thm mult_Suc_right}, - @{thm add_Suc}, @{thm add_Suc_right}, - @{thm eq_number_of_0}, @{thm eq_0_number_of}, @{thm less_0_number_of}, - @{thm of_int_number_of_eq}, @{thm of_nat_number_of_eq}, @{thm nat_number_of}, @{thm if_True}, @{thm if_False}]; - -(* Products are multiplied out during proof (re)construction via -ring_distribs. Ideally they should remain atomic. But that is -currently not possible because 1 is replaced by Suc 0, and then some -simprocs start to mess around with products like (n+1)*m. The rule -1 == Suc 0 is necessary for early parts of HOL where numerals and -simprocs are not yet available. But then it is difficult to remove -that rule later on, because it may find its way back in when theories -(and thus lin-arith simpsets) are merged. Otherwise one could turn the -rule around (Suc n = n+1) and see if that helps products being left -alone. *) - -val simprocs = Nat_Numeral_Simprocs.combine_numerals - :: Nat_Numeral_Simprocs.cancel_numerals; - -in - -val nat_simprocs_setup = - LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} => - {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, - inj_thms = inj_thms, lessD = lessD, neqE = neqE, - simpset = simpset addsimps add_rules - addsimprocs simprocs}); - -end; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOL/simpdata.ML --- a/src/HOL/simpdata.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,200 +0,0 @@ -(* Title: HOL/simpdata.ML - ID: $Id$ - Author: Tobias Nipkow - Copyright 1991 University of Cambridge - -Instantiation of the generic simplifier for HOL. -*) - -(** tools setup **) - -structure Quantifier1 = Quantifier1Fun -(struct - (*abstract syntax*) - fun dest_eq ((c as Const("op =",_)) $ s $ t) = SOME (c, s, t) - | dest_eq _ = NONE; - fun dest_conj ((c as Const("op &",_)) $ s $ t) = SOME (c, s, t) - | dest_conj _ = NONE; - fun dest_imp ((c as Const("op -->",_)) $ s $ t) = SOME (c, s, t) - | dest_imp _ = NONE; - val conj = HOLogic.conj - val imp = HOLogic.imp - (*rules*) - val iff_reflection = @{thm eq_reflection} - val iffI = @{thm iffI} - val iff_trans = @{thm trans} - val conjI= @{thm conjI} - val conjE= @{thm conjE} - val impI = @{thm impI} - val mp = @{thm mp} - val uncurry = @{thm uncurry} - val exI = @{thm exI} - val exE = @{thm exE} - val iff_allI = @{thm iff_allI} - val iff_exI = @{thm iff_exI} - val all_comm = @{thm all_comm} - val ex_comm = @{thm ex_comm} -end); - -structure Simpdata = -struct - -fun mk_meta_eq r = r RS @{thm eq_reflection}; -fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r; - -fun mk_eq th = case concl_of th - (*expects Trueprop if not == *) - of Const ("==",_) $ _ $ _ => th - | _ $ (Const ("op =", _) $ _ $ _) => mk_meta_eq th - | _ $ (Const ("Not", _) $ _) => th RS @{thm Eq_FalseI} - | _ => th RS @{thm Eq_TrueI} - -fun mk_eq_True r = - SOME (r RS @{thm meta_eq_to_obj_eq} RS @{thm Eq_TrueI}) handle Thm.THM _ => NONE; - -(* Produce theorems of the form - (P1 =simp=> ... =simp=> Pn => x == y) ==> (P1 =simp=> ... =simp=> Pn => x = y) -*) - -fun lift_meta_eq_to_obj_eq i st = - let - fun count_imp (Const ("HOL.simp_implies", _) $ P $ Q) = 1 + count_imp Q - | count_imp _ = 0; - val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1))) - in if j = 0 then @{thm meta_eq_to_obj_eq} - else - let - val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j); - fun mk_simp_implies Q = foldr (fn (R, S) => - Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Q Ps - val aT = TFree ("'a", HOLogic.typeS); - val x = Free ("x", aT); - val y = Free ("y", aT) - in Goal.prove_global (Thm.theory_of_thm st) [] - [mk_simp_implies (Logic.mk_equals (x, y))] - (mk_simp_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, y)))) - (fn {prems, ...} => EVERY - [rewrite_goals_tac @{thms simp_implies_def}, - REPEAT (ares_tac (@{thm meta_eq_to_obj_eq} :: - map (rewrite_rule @{thms simp_implies_def}) prems) 1)]) - end - end; - -(*Congruence rules for = (instead of ==)*) -fun mk_meta_cong rl = zero_var_indexes - (let val rl' = Seq.hd (TRYALL (fn i => fn st => - rtac (lift_meta_eq_to_obj_eq i st) i st) rl) - in mk_meta_eq rl' handle THM _ => - if can Logic.dest_equals (concl_of rl') then rl' - else error "Conclusion of congruence rules must be =-equality" - end); - -fun mk_atomize pairs = - let - fun atoms thm = - let - fun res th = map (fn rl => th RS rl); (*exception THM*) - fun res_fixed rls = - if Thm.maxidx_of (Thm.adjust_maxidx_thm ~1 thm) = ~1 then res thm rls - else Variable.trade (K (fn [thm'] => res thm' rls)) (Variable.thm_context thm) [thm]; - in - case concl_of thm - of Const ("Trueprop", _) $ p => (case head_of p - of Const (a, _) => (case AList.lookup (op =) pairs a - of SOME rls => (maps atoms (res_fixed rls) handle THM _ => [thm]) - | NONE => [thm]) - | _ => [thm]) - | _ => [thm] - end; - in atoms end; - -fun mksimps pairs = - map_filter (try mk_eq) o mk_atomize pairs o gen_all; - -fun unsafe_solver_tac prems = - (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' - FIRST' [resolve_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: prems), atac, - etac @{thm FalseE}]; - -val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac; - - -(*No premature instantiation of variables during simplification*) -fun safe_solver_tac prems = - (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' - FIRST' [match_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: prems), - eq_assume_tac, ematch_tac @{thms FalseE}]; - -val safe_solver = mk_solver "HOL safe" safe_solver_tac; - -structure SplitterData = -struct - structure Simplifier = Simplifier - val mk_eq = mk_eq - val meta_eq_to_iff = @{thm meta_eq_to_obj_eq} - val iffD = @{thm iffD2} - val disjE = @{thm disjE} - val conjE = @{thm conjE} - val exE = @{thm exE} - val contrapos = @{thm contrapos_nn} - val contrapos2 = @{thm contrapos_pp} - val notnotD = @{thm notnotD} -end; - -structure Splitter = SplitterFun(SplitterData); - -val split_tac = Splitter.split_tac; -val split_inside_tac = Splitter.split_inside_tac; - -val op addsplits = Splitter.addsplits; -val op delsplits = Splitter.delsplits; -val Addsplits = Splitter.Addsplits; -val Delsplits = Splitter.Delsplits; - - -(* integration of simplifier with classical reasoner *) - -structure Clasimp = ClasimpFun - (structure Simplifier = Simplifier and Splitter = Splitter - and Classical = Classical and Blast = Blast - val iffD1 = @{thm iffD1} val iffD2 = @{thm iffD2} val notE = @{thm notE}); -open Clasimp; - -val _ = ML_Antiquote.value "clasimpset" - (Scan.succeed "Clasimp.local_clasimpset_of (ML_Context.the_local_context ())"); - -val mksimps_pairs = - [("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]), - ("All", [@{thm spec}]), ("True", []), ("False", []), - ("HOL.If", [@{thm if_bool_eq_conj} RS @{thm iffD1}])]; - -val HOL_basic_ss = - Simplifier.theory_context (the_context ()) empty_ss - setsubgoaler asm_simp_tac - setSSolver safe_solver - setSolver unsafe_solver - setmksimps (mksimps mksimps_pairs) - setmkeqTrue mk_eq_True - setmkcong mk_meta_cong; - -fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews); - -fun unfold_tac ths = - let val ss0 = Simplifier.clear_ss HOL_basic_ss addsimps ths - in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; - -val defALL_regroup = - Simplifier.simproc (the_context ()) - "defined ALL" ["ALL x. P x"] Quantifier1.rearrange_all; - -val defEX_regroup = - Simplifier.simproc (the_context ()) - "defined EX" ["EX x. P x"] Quantifier1.rearrange_ex; - - -val simpset_simprocs = HOL_basic_ss addsimprocs [defALL_regroup, defEX_regroup] - -end; - -structure Splitter = Simpdata.Splitter; -structure Clasimp = Simpdata.Clasimp; diff -r 1860f016886d -r 15a4b2cf8c34 src/HOLCF/FOCUS/Buffer.thy --- a/src/HOLCF/FOCUS/Buffer.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOLCF/FOCUS/Buffer.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/FOCUS/Buffer.thy - ID: $Id$ Author: David von Oheimb, TU Muenchen Formalization of section 4 of diff -r 1860f016886d -r 15a4b2cf8c34 src/HOLCF/NatIso.thy --- a/src/HOLCF/NatIso.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/HOLCF/NatIso.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title: HOLCF/NatIso.thy - ID: $Id$ Author: Brian Huffman *) diff -r 1860f016886d -r 15a4b2cf8c34 src/Pure/IsaMakefile --- a/src/Pure/IsaMakefile Wed Dec 03 09:53:58 2008 +0100 +++ b/src/Pure/IsaMakefile Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,3 @@ -# -# $Id$ # # IsaMakefile for Pure # @@ -78,8 +76,8 @@ Syntax/syn_trans.ML Syntax/syntax.ML Syntax/type_ext.ML Thy/html.ML \ Thy/latex.ML Thy/present.ML Thy/term_style.ML Thy/thm_deps.ML \ Thy/thy_edit.ML Thy/thy_header.ML Thy/thy_info.ML Thy/thy_load.ML \ - Thy/thy_output.ML Tools/ROOT.ML Tools/invoke.ML Tools/quickcheck.ML \ - Tools/value.ML Tools/isabelle_process.ML Tools/named_thms.ML \ + Thy/thy_output.ML Tools/ROOT.ML Tools/invoke.ML \ + Tools/isabelle_process.ML Tools/named_thms.ML \ Tools/xml_syntax.ML assumption.ML axclass.ML codegen.ML config.ML \ conjunction.ML consts.ML context.ML context_position.ML conv.ML \ defs.ML display.ML drule.ML envir.ML facts.ML goal.ML \ @@ -88,7 +86,7 @@ proofterm.ML pure_setup.ML pure_thy.ML search.ML sign.ML \ simplifier.ML sorts.ML subgoal.ML tactic.ML tctical.ML term.ML \ term_subst.ML theory.ML thm.ML type.ML type_infer.ML unify.ML \ - variable.ML + variable.ML ../Tools/value.ML ../Tools/quickcheck.ML @./mk diff -r 1860f016886d -r 15a4b2cf8c34 src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML Wed Dec 03 09:53:58 2008 +0100 +++ b/src/Pure/ROOT.ML Wed Dec 03 15:58:44 2008 +0100 @@ -88,6 +88,8 @@ cd "Tools"; use "ROOT.ML"; cd ".."; +use "../Tools/value.ML"; +use "../Tools/quickcheck.ML"; use "codegen.ML"; (*configuration for Proof General*) diff -r 1860f016886d -r 15a4b2cf8c34 src/Pure/Tools/ROOT.ML --- a/src/Pure/Tools/ROOT.ML Wed Dec 03 09:53:58 2008 +0100 +++ b/src/Pure/Tools/ROOT.ML Wed Dec 03 15:58:44 2008 +0100 @@ -1,11 +1,8 @@ (* Title: Pure/Tools/ROOT.ML - ID: $Id$ Miscellaneous tools and packages for Pure Isabelle. *) -use "value.ML"; -use "quickcheck.ML"; use "named_thms.ML"; use "isabelle_process.ML"; diff -r 1860f016886d -r 15a4b2cf8c34 src/Pure/Tools/quickcheck.ML --- a/src/Pure/Tools/quickcheck.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,236 +0,0 @@ -(* Title: Pure/Tools/quickcheck.ML - ID: $Id$ - Author: Stefan Berghofer, Florian Haftmann, TU Muenchen - -Generic counterexample search engine. -*) - -signature QUICKCHECK = -sig - val test_term: Proof.context -> bool -> string option -> int -> int -> term -> (string * term) list option; - val add_generator: string * (Proof.context -> term -> int -> term list option) -> theory -> theory - val auto: bool ref - val auto_time_limit: int ref -end; - -structure Quickcheck : QUICKCHECK = -struct - -(* quickcheck configuration -- default parameters, test generators *) - -datatype test_params = Test_Params of - { size: int, iterations: int, default_type: typ option }; - -fun dest_test_params (Test_Params { size, iterations, default_type}) = - ((size, iterations), default_type); -fun mk_test_params ((size, iterations), default_type) = - Test_Params { size = size, iterations = iterations, default_type = default_type }; -fun map_test_params f (Test_Params { size, iterations, default_type}) = - mk_test_params (f ((size, iterations), default_type)); -fun merge_test_params (Test_Params {size = size1, iterations = iterations1, default_type = default_type1}, - Test_Params {size = size2, iterations = iterations2, default_type = default_type2}) = - mk_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)), - case default_type1 of NONE => default_type2 | _ => default_type1); - -structure Data = TheoryDataFun( - type T = (string * (Proof.context -> term -> int -> term list option)) list - * test_params; - val empty = ([], Test_Params { size = 10, iterations = 100, default_type = NONE }); - val copy = I; - val extend = I; - fun merge pp ((generators1, params1), (generators2, params2)) = - (AList.merge (op = : string * string -> bool) (K true) (generators1, generators2), - merge_test_params (params1, params2)); -) - -val add_generator = Data.map o apfst o AList.update (op =); - - -(* generating tests *) - -fun mk_tester_select name ctxt = - case AList.lookup (op =) ((fst o Data.get o ProofContext.theory_of) ctxt) name - of NONE => error ("No such quickcheck generator: " ^ name) - | SOME generator => generator ctxt; - -fun mk_testers ctxt t = - (map snd o fst o Data.get o ProofContext.theory_of) ctxt - |> map_filter (fn generator => try (generator ctxt) t); - -fun mk_testers_strict ctxt t = - let - val generators = ((map snd o fst o Data.get o ProofContext.theory_of) ctxt) - val testers = map (fn generator => Exn.capture (generator ctxt) t) generators; - in if forall (is_none o Exn.get_result) testers - then [(Exn.release o snd o split_last) testers] - else map_filter Exn.get_result testers - end; - - -(* testing propositions *) - -fun prep_test_term t = - let - val _ = (null (term_tvars t) andalso null (term_tfrees t)) orelse - error "Term to be tested contains type variables"; - val _ = null (term_vars t) orelse - error "Term to be tested contains schematic variables"; - val frees = map dest_Free (term_frees t); - in (map fst frees, list_abs_free (frees, t)) end - -fun test_term ctxt quiet generator_name size i t = - let - val (names, t') = prep_test_term t; - val testers = case generator_name - of NONE => if quiet then mk_testers ctxt t' else mk_testers_strict ctxt t' - | SOME name => [mk_tester_select name ctxt t']; - fun iterate f 0 = NONE - | iterate f k = case f () handle Match => (if quiet then () - else warning "Exception Match raised during quickcheck"; NONE) - of NONE => iterate f (k - 1) | SOME q => SOME q; - fun with_testers k [] = NONE - | with_testers k (tester :: testers) = - case iterate (fn () => tester k) i - of NONE => with_testers k testers - | SOME q => SOME q; - fun with_size k = if k > size then NONE - else (if quiet then () else priority ("Test data size: " ^ string_of_int k); - case with_testers k testers - of NONE => with_size (k + 1) | SOME q => SOME q); - in case with_size 1 - of NONE => NONE - | SOME ts => SOME (names ~~ ts) - end; - -fun monomorphic_term thy insts default_T = - let - fun subst (T as TFree (v, S)) = - let - val T' = AList.lookup (op =) insts v - |> the_default (the_default T default_T) - in if Sign.of_sort thy (T, S) then T' - else error ("Type " ^ Syntax.string_of_typ_global thy T ^ - " to be substituted for variable " ^ - Syntax.string_of_typ_global thy T ^ "\ndoes not have sort " ^ - Syntax.string_of_sort_global thy S) - end - | subst T = T; - in (map_types o map_atyps) subst end; - -fun test_goal quiet generator_name size iterations default_T insts i assms state = - let - val ctxt = Proof.context_of state; - val thy = Proof.theory_of state; - fun strip (Const ("all", _) $ Abs (_, _, t)) = strip t - | strip t = t; - val (_, (_, st)) = Proof.get_goal state; - val (gi, frees) = Logic.goal_params (prop_of st) i; - val gi' = Logic.list_implies (assms, subst_bounds (frees, strip gi)) - |> monomorphic_term thy insts default_T - |> ObjectLogic.atomize_term thy; - in test_term ctxt quiet generator_name size iterations gi' end; - -fun pretty_counterex ctxt NONE = Pretty.str "No counterexamples found." - | pretty_counterex ctxt (SOME cex) = - Pretty.chunks (Pretty.str "Counterexample found:\n" :: - map (fn (s, t) => - Pretty.block [Pretty.str (s ^ " ="), Pretty.brk 1, Syntax.pretty_term ctxt t]) cex); - - -(* automatic testing *) - -val auto = ref false; -val auto_time_limit = ref 5000; - -fun test_goal_auto int state = - let - val ctxt = Proof.context_of state; - val assms = map term_of (Assumption.assms_of ctxt); - val Test_Params { size, iterations, default_type } = - (snd o Data.get o Proof.theory_of) state; - fun test () = - let - val res = TimeLimit.timeLimit (Time.fromMilliseconds (!auto_time_limit)) - (try (test_goal true NONE size iterations default_type [] 1 assms)) state; - in - case res of - NONE => state - | SOME NONE => state - | SOME cex => Proof.goal_message (fn () => Pretty.chunks [Pretty.str "", - Pretty.mark Markup.hilite (pretty_counterex ctxt cex)]) state - end handle TimeLimit.TimeOut => (warning "Auto quickcheck: timeout."; state); - in - if int andalso !auto andalso not (!Toplevel.quiet) - then test () - else state - end; - -val _ = Context.>> (Specification.add_theorem_hook test_goal_auto); - - -(* Isar interfaces *) - -fun read_nat s = case (Library.read_int o Symbol.explode) s - of (k, []) => if k >= 0 then k - else error ("Not a natural number: " ^ s) - | (_, _ :: _) => error ("Not a natural number: " ^ s); - -fun parse_test_param ctxt ("size", arg) = - (apfst o apfst o K) (read_nat arg) - | parse_test_param ctxt ("iterations", arg) = - (apfst o apsnd o K) (read_nat arg) - | parse_test_param ctxt ("default_type", arg) = - (apsnd o K o SOME) (ProofContext.read_typ ctxt arg) - | parse_test_param ctxt (name, _) = - error ("Bad test parameter: " ^ name); - -fun parse_test_param_inst ctxt ("generator", arg) = - (apsnd o apfst o K o SOME) arg - | parse_test_param_inst ctxt (name, arg) = - case try (ProofContext.read_typ ctxt) name - of SOME (TFree (v, _)) => (apsnd o apsnd o AList.update (op =)) - (v, ProofContext.read_typ ctxt arg) - | _ => (apfst o parse_test_param ctxt) (name, arg); - -fun quickcheck_params_cmd args thy = - let - val ctxt = ProofContext.init thy; - val f = fold (parse_test_param ctxt) args; - in - thy - |> (Data.map o apsnd o map_test_params) f - end; - -fun quickcheck_cmd args i state = - let - val prf = Toplevel.proof_of state; - val thy = Toplevel.theory_of state; - val ctxt = Toplevel.context_of state; - val default_params = (dest_test_params o snd o Data.get) thy; - val f = fold (parse_test_param_inst ctxt) args; - val (((size, iterations), default_type), (generator_name, insts)) = - f (default_params, (NONE, [])); - val counterex = test_goal false generator_name size iterations - default_type insts i [] prf; - in (Pretty.writeln o pretty_counterex ctxt) counterex end; - -local structure P = OuterParse and K = OuterKeyword in - -val parse_arg = P.name --| P.$$$ "=" -- P.name; -val parse_args = P.$$$ "[" |-- P.list1 parse_arg --| P.$$$ "]" - || Scan.succeed []; - -val _ = OuterSyntax.command "quickcheck_params" "set parameters for random testing" K.thy_decl - (parse_args >> (fn args => Toplevel.theory (quickcheck_params_cmd args))); - -val _ = OuterSyntax.improper_command "quickcheck" "try to find counterexample for subgoal" K.diag - (parse_args -- Scan.optional P.nat 1 - >> (fn (args, i) => Toplevel.no_timing o Toplevel.keep (quickcheck_cmd args i))); - -end; (*local*) - -end; - - -val auto_quickcheck = Quickcheck.auto; -val auto_quickcheck_time_limit = Quickcheck.auto_time_limit; diff -r 1860f016886d -r 15a4b2cf8c34 src/Pure/Tools/value.ML --- a/src/Pure/Tools/value.ML Wed Dec 03 09:53:58 2008 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,67 +0,0 @@ -(* Title: Pure/Tools/value.ML - ID: $Id$ - Author: Florian Haftmann, TU Muenchen - -Value command for different evaluators. -*) - -signature VALUE = -sig - val value: Proof.context -> term -> term - val value_select: string -> Proof.context -> term -> term - val value_cmd: string option -> string list -> string -> Toplevel.state -> unit - val add_evaluator: string * (Proof.context -> term -> term) -> theory -> theory -end; - -structure Value : VALUE = -struct - -structure Evaluator = TheoryDataFun( - type T = (string * (Proof.context -> term -> term)) list; - val empty = []; - val copy = I; - val extend = I; - fun merge pp = AList.merge (op =) (K true); -) - -val add_evaluator = Evaluator.map o AList.update (op =); - -fun value_select name ctxt = - case AList.lookup (op =) (Evaluator.get (ProofContext.theory_of ctxt)) name - of NONE => error ("No such evaluator: " ^ name) - | SOME f => f ctxt; - -fun value ctxt t = let val evaluators = Evaluator.get (ProofContext.theory_of ctxt) - in if null evaluators then error "No evaluators" - else let val (evaluators, (_, evaluator)) = split_last evaluators - in case get_first (fn (_, f) => try (f ctxt) t) evaluators - of SOME t' => t' - | NONE => evaluator ctxt t - end end; - -fun value_cmd some_name modes raw_t state = - let - val ctxt = Toplevel.context_of state; - val t = Syntax.read_term ctxt raw_t; - val t' = case some_name - of NONE => value ctxt t - | SOME name => value_select name ctxt t; - val ty' = Term.type_of t'; - val ctxt' = Variable.auto_fixes t ctxt; - val p = PrintMode.with_modes modes (fn () => - Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, - Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) (); - in Pretty.writeln p end; - -local structure P = OuterParse and K = OuterKeyword in - -val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) []; - -val _ = OuterSyntax.improper_command "value" "evaluate and print term" K.diag - (Scan.option (P.$$$ "[" |-- P.xname --| P.$$$ "]") -- opt_modes -- P.term - >> (fn ((some_name, modes), t) => Toplevel.no_timing o Toplevel.keep - (value_cmd some_name modes t))); - -end; (*local*) - -end; diff -r 1860f016886d -r 15a4b2cf8c34 src/Sequents/LK.thy --- a/src/Sequents/LK.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/Sequents/LK.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title: LK/LK.ML - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge diff -r 1860f016886d -r 15a4b2cf8c34 src/Tools/quickcheck.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/quickcheck.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,235 @@ +(* Title: Pure/Tools/quickcheck.ML + Author: Stefan Berghofer, Florian Haftmann, TU Muenchen + +Generic counterexample search engine. +*) + +signature QUICKCHECK = +sig + val test_term: Proof.context -> bool -> string option -> int -> int -> term -> (string * term) list option; + val add_generator: string * (Proof.context -> term -> int -> term list option) -> theory -> theory + val auto: bool ref + val auto_time_limit: int ref +end; + +structure Quickcheck : QUICKCHECK = +struct + +(* quickcheck configuration -- default parameters, test generators *) + +datatype test_params = Test_Params of + { size: int, iterations: int, default_type: typ option }; + +fun dest_test_params (Test_Params { size, iterations, default_type}) = + ((size, iterations), default_type); +fun mk_test_params ((size, iterations), default_type) = + Test_Params { size = size, iterations = iterations, default_type = default_type }; +fun map_test_params f (Test_Params { size, iterations, default_type}) = + mk_test_params (f ((size, iterations), default_type)); +fun merge_test_params (Test_Params {size = size1, iterations = iterations1, default_type = default_type1}, + Test_Params {size = size2, iterations = iterations2, default_type = default_type2}) = + mk_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)), + case default_type1 of NONE => default_type2 | _ => default_type1); + +structure Data = TheoryDataFun( + type T = (string * (Proof.context -> term -> int -> term list option)) list + * test_params; + val empty = ([], Test_Params { size = 10, iterations = 100, default_type = NONE }); + val copy = I; + val extend = I; + fun merge pp ((generators1, params1), (generators2, params2)) = + (AList.merge (op = : string * string -> bool) (K true) (generators1, generators2), + merge_test_params (params1, params2)); +) + +val add_generator = Data.map o apfst o AList.update (op =); + + +(* generating tests *) + +fun mk_tester_select name ctxt = + case AList.lookup (op =) ((fst o Data.get o ProofContext.theory_of) ctxt) name + of NONE => error ("No such quickcheck generator: " ^ name) + | SOME generator => generator ctxt; + +fun mk_testers ctxt t = + (map snd o fst o Data.get o ProofContext.theory_of) ctxt + |> map_filter (fn generator => try (generator ctxt) t); + +fun mk_testers_strict ctxt t = + let + val generators = ((map snd o fst o Data.get o ProofContext.theory_of) ctxt) + val testers = map (fn generator => Exn.capture (generator ctxt) t) generators; + in if forall (is_none o Exn.get_result) testers + then [(Exn.release o snd o split_last) testers] + else map_filter Exn.get_result testers + end; + + +(* testing propositions *) + +fun prep_test_term t = + let + val _ = (null (term_tvars t) andalso null (term_tfrees t)) orelse + error "Term to be tested contains type variables"; + val _ = null (term_vars t) orelse + error "Term to be tested contains schematic variables"; + val frees = map dest_Free (term_frees t); + in (map fst frees, list_abs_free (frees, t)) end + +fun test_term ctxt quiet generator_name size i t = + let + val (names, t') = prep_test_term t; + val testers = case generator_name + of NONE => if quiet then mk_testers ctxt t' else mk_testers_strict ctxt t' + | SOME name => [mk_tester_select name ctxt t']; + fun iterate f 0 = NONE + | iterate f k = case f () handle Match => (if quiet then () + else warning "Exception Match raised during quickcheck"; NONE) + of NONE => iterate f (k - 1) | SOME q => SOME q; + fun with_testers k [] = NONE + | with_testers k (tester :: testers) = + case iterate (fn () => tester k) i + of NONE => with_testers k testers + | SOME q => SOME q; + fun with_size k = if k > size then NONE + else (if quiet then () else priority ("Test data size: " ^ string_of_int k); + case with_testers k testers + of NONE => with_size (k + 1) | SOME q => SOME q); + in case with_size 1 + of NONE => NONE + | SOME ts => SOME (names ~~ ts) + end; + +fun monomorphic_term thy insts default_T = + let + fun subst (T as TFree (v, S)) = + let + val T' = AList.lookup (op =) insts v + |> the_default (the_default T default_T) + in if Sign.of_sort thy (T, S) then T' + else error ("Type " ^ Syntax.string_of_typ_global thy T ^ + " to be substituted for variable " ^ + Syntax.string_of_typ_global thy T ^ "\ndoes not have sort " ^ + Syntax.string_of_sort_global thy S) + end + | subst T = T; + in (map_types o map_atyps) subst end; + +fun test_goal quiet generator_name size iterations default_T insts i assms state = + let + val ctxt = Proof.context_of state; + val thy = Proof.theory_of state; + fun strip (Const ("all", _) $ Abs (_, _, t)) = strip t + | strip t = t; + val (_, (_, st)) = Proof.get_goal state; + val (gi, frees) = Logic.goal_params (prop_of st) i; + val gi' = Logic.list_implies (assms, subst_bounds (frees, strip gi)) + |> monomorphic_term thy insts default_T + |> ObjectLogic.atomize_term thy; + in test_term ctxt quiet generator_name size iterations gi' end; + +fun pretty_counterex ctxt NONE = Pretty.str "No counterexamples found." + | pretty_counterex ctxt (SOME cex) = + Pretty.chunks (Pretty.str "Counterexample found:\n" :: + map (fn (s, t) => + Pretty.block [Pretty.str (s ^ " ="), Pretty.brk 1, Syntax.pretty_term ctxt t]) cex); + + +(* automatic testing *) + +val auto = ref false; +val auto_time_limit = ref 5000; + +fun test_goal_auto int state = + let + val ctxt = Proof.context_of state; + val assms = map term_of (Assumption.assms_of ctxt); + val Test_Params { size, iterations, default_type } = + (snd o Data.get o Proof.theory_of) state; + fun test () = + let + val res = TimeLimit.timeLimit (Time.fromMilliseconds (!auto_time_limit)) + (try (test_goal true NONE size iterations default_type [] 1 assms)) state; + in + case res of + NONE => state + | SOME NONE => state + | SOME cex => Proof.goal_message (fn () => Pretty.chunks [Pretty.str "", + Pretty.mark Markup.hilite (pretty_counterex ctxt cex)]) state + end handle TimeLimit.TimeOut => (warning "Auto quickcheck: timeout."; state); + in + if int andalso !auto andalso not (!Toplevel.quiet) + then test () + else state + end; + +val _ = Context.>> (Specification.add_theorem_hook test_goal_auto); + + +(* Isar interfaces *) + +fun read_nat s = case (Library.read_int o Symbol.explode) s + of (k, []) => if k >= 0 then k + else error ("Not a natural number: " ^ s) + | (_, _ :: _) => error ("Not a natural number: " ^ s); + +fun parse_test_param ctxt ("size", arg) = + (apfst o apfst o K) (read_nat arg) + | parse_test_param ctxt ("iterations", arg) = + (apfst o apsnd o K) (read_nat arg) + | parse_test_param ctxt ("default_type", arg) = + (apsnd o K o SOME) (ProofContext.read_typ ctxt arg) + | parse_test_param ctxt (name, _) = + error ("Bad test parameter: " ^ name); + +fun parse_test_param_inst ctxt ("generator", arg) = + (apsnd o apfst o K o SOME) arg + | parse_test_param_inst ctxt (name, arg) = + case try (ProofContext.read_typ ctxt) name + of SOME (TFree (v, _)) => (apsnd o apsnd o AList.update (op =)) + (v, ProofContext.read_typ ctxt arg) + | _ => (apfst o parse_test_param ctxt) (name, arg); + +fun quickcheck_params_cmd args thy = + let + val ctxt = ProofContext.init thy; + val f = fold (parse_test_param ctxt) args; + in + thy + |> (Data.map o apsnd o map_test_params) f + end; + +fun quickcheck_cmd args i state = + let + val prf = Toplevel.proof_of state; + val thy = Toplevel.theory_of state; + val ctxt = Toplevel.context_of state; + val default_params = (dest_test_params o snd o Data.get) thy; + val f = fold (parse_test_param_inst ctxt) args; + val (((size, iterations), default_type), (generator_name, insts)) = + f (default_params, (NONE, [])); + val counterex = test_goal false generator_name size iterations + default_type insts i [] prf; + in (Pretty.writeln o pretty_counterex ctxt) counterex end; + +local structure P = OuterParse and K = OuterKeyword in + +val parse_arg = P.name --| P.$$$ "=" -- P.name; +val parse_args = P.$$$ "[" |-- P.list1 parse_arg --| P.$$$ "]" + || Scan.succeed []; + +val _ = OuterSyntax.command "quickcheck_params" "set parameters for random testing" K.thy_decl + (parse_args >> (fn args => Toplevel.theory (quickcheck_params_cmd args))); + +val _ = OuterSyntax.improper_command "quickcheck" "try to find counterexample for subgoal" K.diag + (parse_args -- Scan.optional P.nat 1 + >> (fn (args, i) => Toplevel.no_timing o Toplevel.keep (quickcheck_cmd args i))); + +end; (*local*) + +end; + + +val auto_quickcheck = Quickcheck.auto; +val auto_quickcheck_time_limit = Quickcheck.auto_time_limit; diff -r 1860f016886d -r 15a4b2cf8c34 src/Tools/value.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/value.ML Wed Dec 03 15:58:44 2008 +0100 @@ -0,0 +1,66 @@ +(* Title: Pure/Tools/value.ML + Author: Florian Haftmann, TU Muenchen + +Generic value command for arbitrary evaluators. +*) + +signature VALUE = +sig + val value: Proof.context -> term -> term + val value_select: string -> Proof.context -> term -> term + val value_cmd: string option -> string list -> string -> Toplevel.state -> unit + val add_evaluator: string * (Proof.context -> term -> term) -> theory -> theory +end; + +structure Value : VALUE = +struct + +structure Evaluator = TheoryDataFun( + type T = (string * (Proof.context -> term -> term)) list; + val empty = []; + val copy = I; + val extend = I; + fun merge pp = AList.merge (op =) (K true); +) + +val add_evaluator = Evaluator.map o AList.update (op =); + +fun value_select name ctxt = + case AList.lookup (op =) (Evaluator.get (ProofContext.theory_of ctxt)) name + of NONE => error ("No such evaluator: " ^ name) + | SOME f => f ctxt; + +fun value ctxt t = let val evaluators = Evaluator.get (ProofContext.theory_of ctxt) + in if null evaluators then error "No evaluators" + else let val (evaluators, (_, evaluator)) = split_last evaluators + in case get_first (fn (_, f) => try (f ctxt) t) evaluators + of SOME t' => t' + | NONE => evaluator ctxt t + end end; + +fun value_cmd some_name modes raw_t state = + let + val ctxt = Toplevel.context_of state; + val t = Syntax.read_term ctxt raw_t; + val t' = case some_name + of NONE => value ctxt t + | SOME name => value_select name ctxt t; + val ty' = Term.type_of t'; + val ctxt' = Variable.auto_fixes t ctxt; + val p = PrintMode.with_modes modes (fn () => + Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, + Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) (); + in Pretty.writeln p end; + +local structure P = OuterParse and K = OuterKeyword in + +val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) []; + +val _ = OuterSyntax.improper_command "value" "evaluate and print term" K.diag + (Scan.option (P.$$$ "[" |-- P.xname --| P.$$$ "]") -- opt_modes -- P.term + >> (fn ((some_name, modes), t) => Toplevel.no_timing o Toplevel.keep + (value_cmd some_name modes t))); + +end; (*local*) + +end; diff -r 1860f016886d -r 15a4b2cf8c34 src/ZF/ex/Ring.thy --- a/src/ZF/ex/Ring.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/ZF/ex/Ring.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title: ZF/ex/Ring.thy - Id: $Id$ *) diff -r 1860f016886d -r 15a4b2cf8c34 src/ZF/pair.thy --- a/src/ZF/pair.thy Wed Dec 03 09:53:58 2008 +0100 +++ b/src/ZF/pair.thy Wed Dec 03 15:58:44 2008 +0100 @@ -1,5 +1,4 @@ (* Title: ZF/pair - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge