src/HOL/ex/SOS_Remote.thy
changeset 58636 9b33fe5b60f3
parent 58627 1329679abb2d
parent 58635 010b610eb55d
child 58637 3094b0edd6b5
child 58638 5855b9b3d6a3
equal deleted inserted replaced
58627:1329679abb2d 58636:9b33fe5b60f3
     1 (*  Title:      HOL/ex/SOS_Remote.thy
       
     2     Author:     Amine Chaieb, University of Cambridge
       
     3     Author:     Philipp Meyer, TU Muenchen
       
     4 
       
     5 Examples for Sum_of_Squares: remote CSDP server.
       
     6 *)
       
     7 
       
     8 theory SOS_Remote
       
     9 imports "~~/src/HOL/Library/Sum_of_Squares"
       
    10 begin
       
    11 
       
    12 lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x \<Longrightarrow> a < 0"
       
    13   by (sos remote_csdp)
       
    14 
       
    15 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)"
       
    16   by (sos remote_csdp)
       
    17 
       
    18 lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x --> a < 0"
       
    19   by (sos remote_csdp)
       
    20 
       
    21 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"
       
    22   by (sos remote_csdp)
       
    23 
       
    24 lemma "(0::real) <= x & 0 <= y & 0 <= z & x + y + z <= 3 --> x * y + x * z + y * z >= 3 * x * y * z"
       
    25   by (sos remote_csdp)
       
    26 
       
    27 lemma "((x::real)^2 + y^2 + z^2 = 1) --> (x + y + z)^2 <= 3"
       
    28   by (sos remote_csdp)
       
    29 
       
    30 lemma "(w^2 + x^2 + y^2 + z^2 = 1) --> (w + x + y + z)^2 <= (4::real)"
       
    31   by (sos remote_csdp)
       
    32 
       
    33 lemma "(x::real) >= 1 & y >= 1 --> x * y >= x + y - 1"
       
    34   by (sos remote_csdp)
       
    35 
       
    36 end
       
    37