| author | haftmann | 
| Mon, 21 Sep 2009 11:01:39 +0200 | |
| changeset 32685 | 29e4e567b5f4 | 
| parent 32543 | 62e6c9b67c6f | 
| child 32645 | 1cc5b24f5a01 | 
| permissions | -rw-r--r-- | 
| 32333 | 1 | (* Title: HOL/Library/Sum_Of_Squares.thy | 
| 31119 
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, 
 | 
| 32333 | 6 | multiplication and ordering using semidefinite programming *} | 
| 32271 | 7 | |
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 8 | theory Sum_Of_Squares | 
| 32332 | 9 | imports Complex_Main (* "~~/src/HOL/Decision_Procs/Dense_Linear_Order" *) | 
| 10 | uses | |
| 32333 | 11 |   ("positivstellensatz.ML")  (* duplicate use!? -- cf. Euclidian_Space.thy *)
 | 
| 32332 | 12 |   ("Sum_Of_Squares/sum_of_squares.ML")
 | 
| 13 |   ("Sum_Of_Squares/sos_wrapper.ML")
 | |
| 14 | begin | |
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 15 | |
| 32333 | 16 | text {*
 | 
| 17 |   In order to use the method sos, call it with @{text "(sos
 | |
| 18 | remote_csdp)"} to use the remote solver. Or install CSDP | |
| 19 | (https://projects.coin-or.org/Csdp), configure the Isabelle setting | |
| 20 |   @{text CSDP_EXE}, and call it with @{text "(sos csdp)"}.  By
 | |
| 21 |   default, sos calls @{text remote_csdp}.  This can take of the order
 | |
| 22 | of a minute for one sos call, because sos calls CSDP repeatedly. If | |
| 23 | you install CSDP locally, sos calls typically takes only a few | |
| 24 | seconds. | |
| 25 | *} | |
| 26 | ||
| 27 | text {* setup sos tactic *}
 | |
| 32332 | 28 | |
| 29 | use "positivstellensatz.ML" | |
| 30 | use "Sum_Of_Squares/sum_of_squares.ML" | |
| 31 | use "Sum_Of_Squares/sos_wrapper.ML" | |
| 32 | ||
| 32268 
d50f0cb67578
Functionality for sum of squares to call a remote csdp prover
 Philipp Meyer parents: 
31512diff
changeset | 33 | setup SosWrapper.setup | 
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 34 | |
| 32333 | 35 | text {* Tests -- commented since they work only when csdp is installed
 | 
| 36 | or take too long with remote csdps *} | |
| 32268 
d50f0cb67578
Functionality for sum of squares to call a remote csdp prover
 Philipp Meyer parents: 
31512diff
changeset | 37 | |
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 38 | (* | 
| 31131 | 39 | 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 | 40 | |
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 41 | lemma "a1 >= 0 & a2 >= 0 \<and> (a1 * a1 + a2 * a2 = b1 * b1 + b2 * b2 + 2) \<and> | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 42 | (a1 * b1 + a2 * b2 = 0) --> a1 * a2 - b1 * b2 >= (0::real)" by sos | 
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 43 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 44 | 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 | 45 | |
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 46 | lemma "(0::real) <= x & x <= 1 & 0 <= y & y <= 1 --> | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 47 | 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 | 
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 48 | |
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 49 | lemma "(0::real) <= x & 0 <= y & 0 <= z & x + y + z <= 3 --> | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 50 | x * y + x * z + y * z >= 3 * x * y * z" by sos | 
| 31119 
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 | 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 | 53 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 54 | 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 | 55 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 56 | 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 | 57 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 58 | 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 | 59 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 60 | lemma "abs(x) <= 1 --> abs(64 * x^7 - 112 * x^5 + 56 * x^3 - 7 * x) <= (1::real)" by sos | 
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 61 | |
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 62 | |
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 63 | text {* One component of denominator in dodecahedral example. *}
 | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 64 | |
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 65 | lemma "2 <= x & x <= 125841 / 50000 & 2 <= y & y <= 125841 / 50000 & 2 <= z & | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 66 | z <= 125841 / 50000 --> 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z) >= (0::real)" by sos | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 67 | |
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 68 | |
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 69 | text {* Over a larger but simpler interval. *}
 | 
| 31131 | 70 | |
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 71 | lemma "(2::real) <= x & x <= 4 & 2 <= y & y <= 4 & 2 <= z & | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 72 | z <= 4 --> 0 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 73 | |
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 74 | text {* We can do 12. I think 12 is a sharp bound; see PP's certificate. *}
 | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 75 | |
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 76 | lemma "2 <= (x::real) & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 77 | 12 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 78 | |
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 79 | |
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 80 | text {* Inequality from sci.math (see "Leon-Sotelo, por favor"). *}
 | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 81 | |
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 82 | 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 | 83 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 84 | 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 | 85 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 86 | 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 | 87 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 88 | 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 | 89 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 90 | 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 | 91 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 92 | 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 | 93 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 94 | lemma "(0::real) < 1 + x^2" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 95 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 96 | 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 | 97 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 98 | lemma "(0::real) < 1 + abs x" by sos | 
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 99 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 100 | 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 | 101 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 102 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 103 | 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 | 104 | 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 | 105 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 106 | 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 | 107 | 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 | 108 | 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 | 109 | 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 | 110 | 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 | 111 | 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 | 
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 112 | lemma "abs(x - z) <= e & abs(y - z) <= e & 0 <= u & 0 <= v & (u + v = 1) --> | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 113 | abs((u * x + v * y) - z) <= (e::real)" by sos | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 114 | |
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 115 | (* | 
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 116 | lemma "((x::real) - y - 2 * x^4 = 0) & 0 <= x & x <= 2 & 0 <= y & y <= 3 --> | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 117 |   y^2 - 7 * y - 12 * x + 17 >= 0" by sos  -- {* Too hard?*}
 | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 118 | *) | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 119 | |
| 31131 | 120 | lemma "(0::real) <= x --> (1 + x + x^2)/(1 + x^2) <= 1 + x" | 
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 121 | by sos | 
| 31131 | 122 | |
| 123 | lemma "(0::real) <= x --> 1 - x <= 1 / (1 + x + x^2)" | |
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 124 | by sos | 
| 31131 | 125 | |
| 126 | lemma "(x::real) <= 1 / 2 --> - x - 2 * x^2 <= - x / (1 - x)" | |
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 127 | by sos | 
| 31131 | 128 | |
| 32543 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 129 | lemma "4*r^2 = p^2 - 4*q & r >= (0::real) & x^2 + p*x + q = 0 --> | 
| 
62e6c9b67c6f
tuned document -- proper text instead of source comments, reduced line length;
 wenzelm parents: 
32333diff
changeset | 130 | 2*(x::real) = - p + 2*r | 2*x = -p - 2*r" by sos | 
| 31131 | 131 | *) | 
| 31119 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 132 | |
| 
2532bb2d65c7
A decision method for universal multivariate real arithmetic with add
 chaieb parents: diff
changeset | 133 | end |