| author | huffman | 
| Mon, 01 Jun 2009 07:57:37 -0700 | |
| changeset 31353 | 14a58e2ca374 | 
| parent 31131 | d9752181691a | 
| child 31512 | 27118561c2e0 | 
| permissions | -rw-r--r-- | 
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 1 | (* Title: Library/Sum_Of_Squares | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 2 | Author: Amine Chaieb, University of Cambridge | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 3 | *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 4 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 5 | header {* A decision method for universal multivariate real arithmetic with addition, 
 | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 6 | multiplication and ordering using semidefinite programming*} | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 7 | theory Sum_Of_Squares | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 8 | imports Complex_Main (* "~~/src/HOL/Decision_Procs/Dense_Linear_Order" *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 9 | uses "positivstellensatz.ML" "sum_of_squares.ML" | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 10 | begin | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 11 | |
| 31131 | 12 | (* Note: | 
| 13 | ||
| 14 | In order to use the method sos, install CSDP (https://projects.coin-or.org/Csdp/) and put the executable csdp on your path. | |
| 15 | ||
| 16 | *) | |
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 17 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 18 | |
| 31131 | 19 | method_setup sos = {* Scan.succeed (SIMPLE_METHOD' o Sos.sos_tac) *} 
 | 
| 20 | "Prove universal problems over the reals using sums of squares" | |
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 21 | |
| 31131 | 22 | text{* Tests -- commented since they work only when csdp is installed -- see above *}
 | 
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 23 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 24 | (* | 
| 31131 | 25 | lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x \<Longrightarrow> a < 0" by sos | 
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 26 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 27 | lemma "a1 >= 0 & a2 >= 0 \<and> (a1 * a1 + a2 * a2 = b1 * b1 + b2 * b2 + 2) \<and> (a1 * b1 + a2 * b2 = 0) --> a1 * a2 - b1 * b2 >= (0::real)" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 28 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 29 | lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x --> a < 0" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 30 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 31 | lemma "(0::real) <= x & x <= 1 & 0 <= y & y <= 1 --> x^2 + y^2 < 1 |(x - 1)^2 + y^2 < 1 | x^2 + (y - 1)^2 < 1 | (x - 1)^2 + (y - 1)^2 < 1" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 32 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 33 | lemma "(0::real) <= x & 0 <= y & 0 <= z & x + y + z <= 3 --> x * y + x * z + y * z >= 3 * x * y * z" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 34 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 35 | lemma "((x::real)^2 + y^2 + z^2 = 1) --> (x + y + z)^2 <= 3" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 36 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 37 | lemma "(w^2 + x^2 + y^2 + z^2 = 1) --> (w + x + y + z)^2 <= (4::real)" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 38 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 39 | lemma "(x::real) >= 1 & y >= 1 --> x * y >= x + y - 1" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 40 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 41 | lemma "(x::real) > 1 & y > 1 --> x * y > x + y - 1" by sos; | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 42 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 43 | lemma "abs(x) <= 1 --> abs(64 * x^7 - 112 * x^5 + 56 * x^3 - 7 * x) <= (1::real)" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 44 | *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 45 | (* ------------------------------------------------------------------------- *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 46 | (* One component of denominator in dodecahedral example. *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 47 | (* ------------------------------------------------------------------------- *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 48 | (* | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 49 | lemma "2 <= x & x <= 125841 / 50000 & 2 <= y & y <= 125841 / 50000 & 2 <= z & z <= 125841 / 50000 --> 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z) >= (0::real)" by sos; | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 50 | *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 51 | (* ------------------------------------------------------------------------- *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 52 | (* Over a larger but simpler interval. *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 53 | (* ------------------------------------------------------------------------- *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 54 | (* | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 55 | lemma "(2::real) <= x & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 0 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 56 | *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 57 | (* ------------------------------------------------------------------------- *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 58 | (* We can do 12. I think 12 is a sharp bound; see PP's certificate. *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 59 | (* ------------------------------------------------------------------------- *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 60 | (* | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 61 | lemma "2 <= (x::real) & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 12 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos | 
| 31131 | 62 | |
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 63 | *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 64 | (* ------------------------------------------------------------------------- *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 65 | (* Inequality from sci.math (see "Leon-Sotelo, por favor"). *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 66 | (* ------------------------------------------------------------------------- *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 67 | (* | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 68 | lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x + y <= x^2 + y^2" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 69 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 70 | lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x * y * (x + y) <= x^2 + y^2" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 71 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 72 | lemma "0 <= (x::real) & 0 <= y --> x * y * (x + y)^2 <= (x^2 + y^2)^2" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 73 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 74 | lemma "(0::real) <= a & 0 <= b & 0 <= c & c * (2 * a + b)^3/ 27 <= x \<longrightarrow> c * a^2 * b <= x" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 75 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 76 | lemma "(0::real) < x --> 0 < 1 + x + x^2" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 77 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 78 | lemma "(0::real) <= x --> 0 < 1 + x + x^2" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 79 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 80 | lemma "(0::real) < 1 + x^2" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 81 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 82 | lemma "(0::real) <= 1 + 2 * x + x^2" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 83 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 84 | lemma "(0::real) < 1 + abs x" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 85 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 86 | lemma "(0::real) < 1 + (1 + x)^2 * (abs x)" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 87 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 88 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 89 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 90 | lemma "abs ((1::real) + x^2) = (1::real) + x^2" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 91 | lemma "(3::real) * x + 7 * a < 4 \<and> 3 < 2 * x \<longrightarrow> a < 0" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 92 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 93 | lemma "(0::real) < x --> 1 < y --> y * x <= z --> x < z" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 94 | lemma "(1::real) < x --> x^2 < y --> 1 < y" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 95 | lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 96 | lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 97 | lemma "((a::real) * x^2 + b * x + c = 0) --> b^2 >= 4 * a * c" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 98 | lemma "(0::real) <= b & 0 <= c & 0 <= x & 0 <= y & (x^2 = c) & (y^2 = a^2 * c + b) --> a * c <= y * x" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 99 | lemma "abs(x - z) <= e & abs(y - z) <= e & 0 <= u & 0 <= v & (u + v = 1) --> abs((u * x + v * y) - z) <= (e::real)" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 100 | *) | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 101 | (* | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 102 | lemma "((x::real) - y - 2 * x^4 = 0) & 0 <= x & x <= 2 & 0 <= y & y <= 3 --> y^2 - 7 * y - 12 * x + 17 >= 0" by sos *) (* Too hard?*) | 
| 31131 | 103 | (* | 
| 104 | lemma "(0::real) <= x --> (1 + x + x^2)/(1 + x^2) <= 1 + x" | |
| 105 | apply sos | |
| 106 | done | |
| 107 | ||
| 108 | lemma "(0::real) <= x --> 1 - x <= 1 / (1 + x + x^2)" | |
| 109 | apply sos | |
| 110 | done | |
| 111 | ||
| 112 | lemma "(x::real) <= 1 / 2 --> - x - 2 * x^2 <= - x / (1 - x)" | |
| 113 | apply sos | |
| 114 | done | |
| 115 | ||
| 116 | lemma "4*r^2 = p^2 - 4*q & r >= (0::real) & x^2 + p*x + q = 0 --> 2*(x::real) = - p + 2*r | 2*x = -p - 2*r" by sos | |
| 117 | *) | |
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 118 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 119 | end |