src/HOL/Library/Sum_of_Squares/sum_of_squares.ML
author huffman
Mon, 22 Aug 2011 16:49:45 -0700
changeset 44453 082edd97ffe1
parent 42616 92715b528e78
child 44469 266dfd7f4e82
permissions -rw-r--r--
comment out dead code to avoid compiler warnings
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41474
60d091240485 renamed Sum_Of_Squares to Sum_of_Squares;
wenzelm
parents: 40721
diff changeset
     1
(*  Title:      HOL/Library/Sum_of_Squares/sum_of_squares.ML
32949
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
     2
    Author:     Amine Chaieb, University of Cambridge
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
     3
    Author:     Philipp Meyer, TU Muenchen
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
     4
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
     5
A tactic for proving nonlinear inequalities.
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
     6
*)
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
     7
32949
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
     8
signature SUM_OF_SQUARES =
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
     9
sig
32949
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
    10
  datatype proof_method = Certificate of RealArith.pss_tree | Prover of string -> string
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
    11
  val sos_tac: (RealArith.pss_tree -> unit) -> proof_method -> Proof.context -> int -> tactic
38805
b09d8603f865 Sum_Of_Squares: proper configuration options;
wenzelm
parents: 38795
diff changeset
    12
  val trace: bool Config.T
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
    13
  exception Failure of string;
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
    14
end
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
    15
41474
60d091240485 renamed Sum_Of_Squares to Sum_of_Squares;
wenzelm
parents: 40721
diff changeset
    16
structure Sum_of_Squares: SUM_OF_SQUARES =
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    17
struct
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    18
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    19
val rat_0 = Rat.zero;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    20
val rat_1 = Rat.one;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    21
val rat_2 = Rat.two;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    22
val rat_10 = Rat.rat_of_int 10;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
    23
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    24
val rat_1_2 = rat_1 // rat_2;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
    25
*)
33029
2fefe039edf1 uniform use of Integer.min/max;
wenzelm
parents: 33002
diff changeset
    26
val max = Integer.max;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
    27
(*
33029
2fefe039edf1 uniform use of Integer.min/max;
wenzelm
parents: 33002
diff changeset
    28
val min = Integer.min;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
    29
*)
31119
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
val denominator_rat = Rat.quotient_of_rat #> snd #> Rat.rat_of_int;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
    32
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    33
val numerator_rat = Rat.quotient_of_rat #> fst #> Rat.rat_of_int;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
    34
*)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
    35
fun int_of_rat a =
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    36
    case Rat.quotient_of_rat a of (i,1) => i | _ => error "int_of_rat: not an int";
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    37
fun lcm_rat x y = Rat.rat_of_int (Integer.lcm (int_of_rat x) (int_of_rat y));
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    38
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
    39
fun rat_pow r i =
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
    40
 let fun pow r i =
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
    41
   if i = 0 then rat_1 else
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    42
   let val d = pow r (i div 2)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    43
   in d */ d */ (if i mod 2 = 0 then rat_1 else r)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    44
   end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    45
 in if i < 0 then pow (Rat.inv r) (~ i) else pow r i end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    46
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
    47
fun round_rat r =
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    48
 let val (a,b) = Rat.quotient_of_rat (Rat.abs r)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    49
     val d = a div b
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    50
     val s = if r </ rat_0 then (Rat.neg o Rat.rat_of_int) else Rat.rat_of_int
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    51
     val x2 = 2 * (a - (b * d))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    52
 in s (if x2 >= b then d + 1 else d) end
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
val abs_rat = Rat.abs;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    55
val pow2 = rat_pow rat_2;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    56
val pow10 = rat_pow rat_10;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    57
42616
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 41490
diff changeset
    58
val trace = Attrib.setup_config_bool @{binding sos_trace} (K false);
31119
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
exception Sanity;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    61
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    62
exception Unsolvable;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    63
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
    64
exception Failure of string;
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
    65
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
    66
datatype proof_method =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
    67
    Certificate of RealArith.pss_tree
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
    68
  | Prover of (string -> string)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
    69
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    70
(* Turn a rational into a decimal string with d sig digits.                  *)
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
local
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    73
fun normalize y =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    74
  if abs_rat y </ (rat_1 // rat_10) then normalize (rat_10 */ y) - 1
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    75
  else if abs_rat y >=/ rat_1 then normalize (y // rat_10) + 1
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
    76
  else 0
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    77
 in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    78
fun decimalize d x =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    79
  if x =/ rat_0 then "0.0" else
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
    80
  let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    81
   val y = Rat.abs x
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    82
   val e = normalize y
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    83
   val z = pow10(~ e) */ y +/ rat_1
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    84
   val k = int_of_rat (round_rat(pow10 d */ z))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    85
  in (if x </ rat_0 then "-0." else "0.") ^
40627
becf5d5187cc renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents: 39027
diff changeset
    86
     implode(tl(raw_explode(string_of_int k))) ^
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    87
     (if e = 0 then "" else "e"^string_of_int e)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    88
  end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    89
end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    90
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    91
(* Iterations over numbers, and lists indexed by numbers.                    *)
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
fun itern k l f a =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    94
  case l of
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    95
    [] => a
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    96
  | h::t => itern (k + 1) t f (f h k a);
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
fun iter (m,n) f a =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
    99
  if n < m then a
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   100
  else iter (m+1,n) f (f m a);
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
(* The main types.                                                           *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   103
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   104
type vector = int* Rat.rat FuncUtil.Intfunc.table;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   105
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   106
type matrix = (int*int)*(Rat.rat FuncUtil.Intpairfunc.table);
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   107
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   108
fun iszero (_,r) = r =/ rat_0;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   109
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   110
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   111
(* Vectors. Conventionally indexed 1..n.                                     *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   112
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   113
fun vector_0 n = (n,FuncUtil.Intfunc.empty):vector;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   114
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   115
fun dim (v:vector) = fst v;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   116
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   117
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   118
fun vector_const c n =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   119
  if c =/ rat_0 then vector_0 n
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   120
  else (n,fold_rev (fn k => FuncUtil.Intfunc.update (k,c)) (1 upto n) FuncUtil.Intfunc.empty) :vector;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   121
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   122
val vector_1 = vector_const rat_1;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   123
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   124
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   125
fun vector_cmul c (v:vector) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   126
 let val n = dim v
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   127
 in if c =/ rat_0 then vector_0 n
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   128
    else (n,FuncUtil.Intfunc.map (fn _ => fn x => c */ x) (snd v))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   129
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   130
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   131
(*
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   132
fun vector_neg (v:vector) = (fst v,FuncUtil.Intfunc.map (K Rat.neg) (snd v)) :vector;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   133
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   134
fun vector_add (v1:vector) (v2:vector) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   135
 let val m = dim v1
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   136
     val n = dim v2
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   137
 in if m <> n then error "vector_add: incompatible dimensions"
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   138
    else (n,FuncUtil.Intfunc.combine (curry op +/) (fn x => x =/ rat_0) (snd v1) (snd v2)) :vector
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   139
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   140
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   141
fun vector_sub v1 v2 = vector_add v1 (vector_neg v2);
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   142
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   143
fun vector_dot (v1:vector) (v2:vector) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   144
 let val m = dim v1
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   145
     val n = dim v2
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   146
 in if m <> n then error "vector_dot: incompatible dimensions"
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   147
    else FuncUtil.Intfunc.fold (fn (_,x) => fn a => x +/ a)
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   148
        (FuncUtil.Intfunc.combine (curry op */) (fn x => x =/ rat_0) (snd v1) (snd v2)) rat_0
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   149
 end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   150
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   151
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   152
fun vector_of_list l =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   153
 let val n = length l
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   154
 in (n,fold_rev2 (curry FuncUtil.Intfunc.update) (1 upto n) l FuncUtil.Intfunc.empty) :vector
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   155
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   156
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   157
(* Matrices; again rows and columns indexed from 1.                          *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   158
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   159
(*
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   160
fun matrix_0 (m,n) = ((m,n),FuncUtil.Intpairfunc.empty):matrix;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   161
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   162
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   163
fun dimensions (m:matrix) = fst m;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   164
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   165
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   166
fun matrix_const c (mn as (m,n)) =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   167
  if m <> n then error "matrix_const: needs to be square"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   168
  else if c =/ rat_0 then matrix_0 mn
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   169
  else (mn,fold_rev (fn k => FuncUtil.Intpairfunc.update ((k,k), c)) (1 upto n) FuncUtil.Intpairfunc.empty) :matrix;;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   170
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   171
val matrix_1 = matrix_const rat_1;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   172
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   173
fun matrix_cmul c (m:matrix) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   174
 let val (i,j) = dimensions m
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   175
 in if c =/ rat_0 then matrix_0 (i,j)
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   176
    else ((i,j),FuncUtil.Intpairfunc.map (fn _ => fn x => c */ x) (snd m))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   177
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   178
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   179
fun matrix_neg (m:matrix) =
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   180
  (dimensions m, FuncUtil.Intpairfunc.map (K Rat.neg) (snd m)) :matrix;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   181
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   182
fun matrix_add (m1:matrix) (m2:matrix) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   183
 let val d1 = dimensions m1
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   184
     val d2 = dimensions m2
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   185
 in if d1 <> d2
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   186
     then error "matrix_add: incompatible dimensions"
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   187
    else (d1,FuncUtil.Intpairfunc.combine (curry op +/) (fn x => x =/ rat_0) (snd m1) (snd m2)) :matrix
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   188
 end;;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   189
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   190
fun matrix_sub m1 m2 = matrix_add m1 (matrix_neg m2);
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   191
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   192
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   193
fun row k (m:matrix) =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   194
 let val (_,j) = dimensions m
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   195
 in (j,
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   196
   FuncUtil.Intpairfunc.fold (fn ((i,j), c) => fn a => if i = k then FuncUtil.Intfunc.update (j,c) a else a) (snd m) FuncUtil.Intfunc.empty ) : vector
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   197
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   198
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   199
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   200
fun column k (m:matrix) =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   201
  let val (i,_) = dimensions m
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   202
  in (i,
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   203
   FuncUtil.Intpairfunc.fold (fn ((i,j), c) => fn a => if j = k then FuncUtil.Intfunc.update (i,c) a else a) (snd m)  FuncUtil.Intfunc.empty)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   204
   : vector
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   205
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   206
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   207
fun transp (m:matrix) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   208
  let val (i,j) = dimensions m
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   209
  in
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   210
  ((j,i),FuncUtil.Intpairfunc.fold (fn ((i,j), c) => fn a => FuncUtil.Intpairfunc.update ((j,i), c) a) (snd m) FuncUtil.Intpairfunc.empty) :matrix
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   211
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   212
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   213
fun diagonal (v:vector) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   214
 let val n = dim v
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   215
 in ((n,n),FuncUtil.Intfunc.fold (fn (i, c) => fn a => FuncUtil.Intpairfunc.update ((i,i), c) a) (snd v) FuncUtil.Intpairfunc.empty) : matrix
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   216
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   217
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   218
fun matrix_of_list l =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   219
 let val m = length l
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   220
 in if m = 0 then matrix_0 (0,0) else
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   221
   let val n = length (hd l)
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   222
   in ((m,n),itern 1 l (fn v => fn i => itern 1 v (fn c => fn j => FuncUtil.Intpairfunc.update ((i,j), c))) FuncUtil.Intpairfunc.empty)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   223
   end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   224
 end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   225
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   226
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   227
(* Monomials.                                                                *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   228
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   229
fun monomial_eval assig m =
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   230
  FuncUtil.Ctermfunc.fold (fn (x, k) => fn a => a */ rat_pow (FuncUtil.Ctermfunc.apply assig x) k)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   231
        m rat_1;
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   232
val monomial_1 = FuncUtil.Ctermfunc.empty;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   233
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   234
fun monomial_var x = FuncUtil.Ctermfunc.onefunc (x, 1);
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   235
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   236
val monomial_mul =
33002
f3f02f36a3e2 uniform use of Integer.add/mult/sum/prod;
wenzelm
parents: 32949
diff changeset
   237
  FuncUtil.Ctermfunc.combine Integer.add (K false);
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   238
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   239
(*
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   240
fun monomial_pow m k =
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   241
  if k = 0 then monomial_1
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   242
  else FuncUtil.Ctermfunc.map (fn _ => fn x => k * x) m;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   243
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   244
fun monomial_divides m1 m2 =
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   245
  FuncUtil.Ctermfunc.fold (fn (x, k) => fn a => FuncUtil.Ctermfunc.tryapplyd m2 x 0 >= k andalso a) m1 true;;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   246
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   247
fun monomial_div m1 m2 =
33002
f3f02f36a3e2 uniform use of Integer.add/mult/sum/prod;
wenzelm
parents: 32949
diff changeset
   248
 let val m = FuncUtil.Ctermfunc.combine Integer.add
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   249
   (fn x => x = 0) m1 (FuncUtil.Ctermfunc.map (fn _ => fn x => ~ x) m2)
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   250
 in if FuncUtil.Ctermfunc.fold (fn (_, k) => fn a => k >= 0 andalso a) m true then m
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   251
  else error "monomial_div: non-divisible"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   252
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   253
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   254
fun monomial_degree x m =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   255
  FuncUtil.Ctermfunc.tryapplyd m x 0;;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   256
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   257
fun monomial_lcm m1 m2 =
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   258
  fold_rev (fn x => FuncUtil.Ctermfunc.update (x, max (monomial_degree x m1) (monomial_degree x m2)))
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33039
diff changeset
   259
          (union (is_equal o FuncUtil.cterm_ord) (FuncUtil.Ctermfunc.dom m1) (FuncUtil.Ctermfunc.dom m2)) (FuncUtil.Ctermfunc.empty);
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   260
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   261
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   262
fun monomial_multidegree m =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   263
 FuncUtil.Ctermfunc.fold (fn (_, k) => fn a => k + a) m 0;;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   264
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   265
fun monomial_variables m = FuncUtil.Ctermfunc.dom m;;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   266
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   267
(* Polynomials.                                                              *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   268
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   269
fun eval assig p =
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   270
  FuncUtil.Monomialfunc.fold (fn (m, c) => fn a => a +/ c */ monomial_eval assig m) p rat_0;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   271
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   272
val poly_0 = FuncUtil.Monomialfunc.empty;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   273
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   274
fun poly_isconst p =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   275
  FuncUtil.Monomialfunc.fold (fn (m, _) => fn a => FuncUtil.Ctermfunc.is_empty m andalso a) p true;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   276
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   277
fun poly_var x = FuncUtil.Monomialfunc.onefunc (monomial_var x,rat_1);
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   278
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   279
fun poly_const c =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   280
  if c =/ rat_0 then poly_0 else FuncUtil.Monomialfunc.onefunc(monomial_1, c);
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   281
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   282
fun poly_cmul c p =
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   283
  if c =/ rat_0 then poly_0
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   284
  else FuncUtil.Monomialfunc.map (fn _ => fn x => c */ x) p;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   285
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   286
fun poly_neg p = FuncUtil.Monomialfunc.map (K Rat.neg) p;;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   287
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   288
fun poly_add p1 p2 =
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   289
  FuncUtil.Monomialfunc.combine (curry op +/) (fn x => x =/ rat_0) p1 p2;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   290
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   291
fun poly_sub p1 p2 = poly_add p1 (poly_neg p2);
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   292
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   293
fun poly_cmmul (c,m) p =
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   294
 if c =/ rat_0 then poly_0
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   295
 else if FuncUtil.Ctermfunc.is_empty m
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   296
      then FuncUtil.Monomialfunc.map (fn _ => fn d => c */ d) p
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   297
      else FuncUtil.Monomialfunc.fold (fn (m', d) => fn a => (FuncUtil.Monomialfunc.update (monomial_mul m m', c */ d) a)) p poly_0;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   298
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   299
fun poly_mul p1 p2 =
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   300
  FuncUtil.Monomialfunc.fold (fn (m, c) => fn a => poly_add (poly_cmmul (c,m) p2) a) p1 poly_0;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   301
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   302
(*
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   303
fun poly_div p1 p2 =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   304
 if not(poly_isconst p2)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   305
 then error "poly_div: non-constant" else
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   306
 let val c = eval FuncUtil.Ctermfunc.empty p2
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   307
 in if c =/ rat_0 then error "poly_div: division by zero"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   308
    else poly_cmul (Rat.inv c) p1
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   309
 end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   310
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   311
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   312
fun poly_square p = poly_mul p p;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   313
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   314
fun poly_pow p k =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   315
 if k = 0 then poly_const rat_1
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   316
 else if k = 1 then p
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   317
 else let val q = poly_square(poly_pow p (k div 2)) in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   318
      if k mod 2 = 1 then poly_mul p q else q end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   319
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   320
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   321
fun poly_exp p1 p2 =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   322
  if not(poly_isconst p2)
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   323
  then error "poly_exp: not a constant"
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   324
  else poly_pow p1 (int_of_rat (eval FuncUtil.Ctermfunc.empty p2));
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   325
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   326
fun degree x p =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   327
 FuncUtil.Monomialfunc.fold (fn (m,_) => fn a => max (monomial_degree x m) a) p 0;
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   328
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   329
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   330
fun multidegree p =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   331
  FuncUtil.Monomialfunc.fold (fn (m, _) => fn a => max (monomial_multidegree m) a) p 0;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   332
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   333
fun poly_variables p =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   334
  sort FuncUtil.cterm_ord (FuncUtil.Monomialfunc.fold_rev (fn (m, _) => union (is_equal o FuncUtil.cterm_ord) (monomial_variables m)) p []);;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   335
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   336
(* Order monomials for human presentation.                                   *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   337
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   338
(*
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   339
val humanorder_varpow = prod_ord FuncUtil.cterm_ord (rev_order o int_ord);
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   340
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   341
local
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   342
 fun ord (l1,l2) = case (l1,l2) of
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   343
  (_,[]) => LESS
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   344
 | ([],_) => GREATER
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   345
 | (h1::t1,h2::t2) =>
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   346
   (case humanorder_varpow (h1, h2) of
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   347
     LESS => LESS
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   348
   | EQUAL => ord (t1,t2)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   349
   | GREATER => GREATER)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   350
in fun humanorder_monomial m1 m2 =
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   351
 ord (sort humanorder_varpow (FuncUtil.Ctermfunc.dest m1),
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   352
  sort humanorder_varpow (FuncUtil.Ctermfunc.dest m2))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   353
end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   354
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   355
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   356
(* Conversions to strings.                                                   *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   357
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   358
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   359
fun string_of_vector min_size max_size (v:vector) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   360
 let val n_raw = dim v
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   361
 in if n_raw = 0 then "[]" else
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   362
  let
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   363
   val n = max min_size (min n_raw max_size)
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   364
   val xs = map (Rat.string_of_rat o (fn i => FuncUtil.Intfunc.tryapplyd (snd v) i rat_0)) (1 upto n)
32830
47c1b15c03fe replaced and tuned uses of foldr1
Philipp Meyer
parents: 32829
diff changeset
   365
  in "[" ^ space_implode ", " xs ^
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   366
  (if n_raw > max_size then ", ...]" else "]")
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   367
  end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   368
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   369
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   370
fun string_of_matrix max_size (m:matrix) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   371
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   372
  val (i_raw,j_raw) = dimensions m
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   373
  val i = min max_size i_raw
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   374
  val j = min max_size j_raw
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   375
  val rstr = map (fn k => string_of_vector j j (row k m)) (1 upto i)
32830
47c1b15c03fe replaced and tuned uses of foldr1
Philipp Meyer
parents: 32829
diff changeset
   376
 in "["^ space_implode ";\n " rstr ^
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   377
  (if j > max_size then "\n ...]" else "]")
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   378
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   379
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   380
fun string_of_term t =
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   381
 case t of
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   382
   a$b => "("^(string_of_term a)^" "^(string_of_term b)^")"
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   383
 | Abs x =>
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   384
    let val (xn, b) = Term.dest_abs x
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   385
    in "(\\"^xn^"."^(string_of_term b)^")"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   386
    end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   387
 | Const(s,_) => s
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   388
 | Free (s,_) => s
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   389
 | Var((s,_),_) => s
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   390
 | _ => error "string_of_term";
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   391
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   392
val string_of_cterm = string_of_term o term_of;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   393
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   394
fun string_of_varpow x k =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   395
  if k = 1 then string_of_cterm x
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   396
  else string_of_cterm x^"^"^string_of_int k;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   397
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   398
fun string_of_monomial m =
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   399
 if FuncUtil.Ctermfunc.is_empty m then "1" else
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   400
 let val vps = fold_rev (fn (x,k) => fn a => string_of_varpow x k :: a)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   401
  (sort humanorder_varpow (FuncUtil.Ctermfunc.dest m)) []
32830
47c1b15c03fe replaced and tuned uses of foldr1
Philipp Meyer
parents: 32829
diff changeset
   402
 in space_implode "*" vps
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   403
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   404
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   405
fun string_of_cmonomial (c,m) =
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   406
 if FuncUtil.Ctermfunc.is_empty m then Rat.string_of_rat c
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   407
 else if c =/ rat_1 then string_of_monomial m
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   408
 else Rat.string_of_rat c ^ "*" ^ string_of_monomial m;;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   409
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
   410
fun string_of_poly p =
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   411
 if FuncUtil.Monomialfunc.is_empty p then "<<0>>" else
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   412
 let
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   413
  val cms = sort (fn ((m1,_),(m2,_)) => humanorder_monomial m1  m2) (FuncUtil.Monomialfunc.dest p)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   414
  val s = fold (fn (m,c) => fn a =>
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   415
             if c </ rat_0 then a ^ " - " ^ string_of_cmonomial(Rat.neg c,m)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   416
             else a ^ " + " ^ string_of_cmonomial(c,m))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   417
          cms ""
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   418
  val s1 = String.substring (s, 0, 3)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   419
  val s2 = String.substring (s, 3, String.size s - 3)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   420
 in "<<" ^(if s1 = " + " then s2 else "-"^s2)^">>"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   421
 end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   422
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   423
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   424
(* Conversion from HOL term.                                                 *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   425
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   426
local
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   427
 val neg_tm = @{cterm "uminus :: real => _"}
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   428
 val add_tm = @{cterm "op + :: real => _"}
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   429
 val sub_tm = @{cterm "op - :: real => _"}
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   430
 val mul_tm = @{cterm "op * :: real => _"}
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   431
 val inv_tm = @{cterm "inverse :: real => _"}
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   432
 val div_tm = @{cterm "op / :: real => _"}
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   433
 val pow_tm = @{cterm "op ^ :: real => _"}
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   434
 val zero_tm = @{cterm "0:: real"}
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   435
 val is_numeral = can (HOLogic.dest_number o term_of)
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   436
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   437
 fun is_comb t = case t of _$_ => true | _ => false
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   438
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   439
 fun poly_of_term tm =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   440
  if tm aconvc zero_tm then poly_0
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   441
  else if RealArith.is_ratconst tm
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   442
       then poly_const(RealArith.dest_ratconst tm)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   443
  else
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   444
  (let val (lop,r) = Thm.dest_comb tm
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   445
   in if lop aconvc neg_tm then poly_neg(poly_of_term r)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   446
      else if lop aconvc inv_tm then
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   447
       let val p = poly_of_term r
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   448
       in if poly_isconst p
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   449
          then poly_const(Rat.inv (eval FuncUtil.Ctermfunc.empty p))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   450
          else error "poly_of_term: inverse of non-constant polyomial"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   451
       end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   452
   else (let val (opr,l) = Thm.dest_comb lop
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   453
         in
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   454
          if opr aconvc pow_tm andalso is_numeral r
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   455
          then poly_pow (poly_of_term l) ((snd o HOLogic.dest_number o term_of) r)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   456
          else if opr aconvc add_tm
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   457
           then poly_add (poly_of_term l) (poly_of_term r)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   458
          else if opr aconvc sub_tm
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   459
           then poly_sub (poly_of_term l) (poly_of_term r)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   460
          else if opr aconvc mul_tm
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   461
           then poly_mul (poly_of_term l) (poly_of_term r)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   462
          else if opr aconvc div_tm
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   463
           then let
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   464
                  val p = poly_of_term l
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   465
                  val q = poly_of_term r
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   466
                in if poly_isconst q then poly_cmul (Rat.inv (eval FuncUtil.Ctermfunc.empty q)) p
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   467
                   else error "poly_of_term: division by non-constant polynomial"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   468
                end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   469
          else poly_var tm
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   470
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   471
         end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   472
         handle CTERM ("dest_comb",_) => poly_var tm)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   473
   end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   474
   handle CTERM ("dest_comb",_) => poly_var tm)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   475
in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   476
val poly_of_term = fn tm =>
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   477
 if type_of (term_of tm) = @{typ real} then poly_of_term tm
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   478
 else error "poly_of_term: term does not have real type"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   479
end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   480
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   481
(* String of vector (just a list of space-separated numbers).                *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   482
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   483
fun sdpa_of_vector (v:vector) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   484
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   485
  val n = dim v
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   486
  val strs = map (decimalize 20 o (fn i => FuncUtil.Intfunc.tryapplyd (snd v) i rat_0)) (1 upto n)
32830
47c1b15c03fe replaced and tuned uses of foldr1
Philipp Meyer
parents: 32829
diff changeset
   487
 in space_implode " " strs ^ "\n"
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   488
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   489
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   490
fun triple_int_ord ((a,b,c),(a',b',c')) =
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   491
 prod_ord int_ord (prod_ord int_ord int_ord)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   492
    ((a,(b,c)),(a',(b',c')));
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   493
structure Inttriplefunc = FuncFun(type key = int*int*int val ord = triple_int_ord);
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   494
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   495
(* String for block diagonal matrix numbered k.                              *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   496
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   497
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   498
fun sdpa_of_blockdiagonal k m =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   499
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   500
  val pfx = string_of_int k ^" "
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   501
  val ents =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   502
    Inttriplefunc.fold (fn ((b,i,j), c) => fn a => if i > j then a else ((b,i,j),c)::a) m []
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   503
  val entss = sort (triple_int_ord o pairself fst) ents
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   504
 in  fold_rev (fn ((b,i,j),c) => fn a =>
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   505
     pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   506
     " " ^ decimalize 20 c ^ "\n" ^ a) entss ""
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   507
 end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   508
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   509
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   510
(* String for a matrix numbered k, in SDPA sparse format.                    *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   511
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   512
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   513
fun sdpa_of_matrix k (m:matrix) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   514
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   515
  val pfx = string_of_int k ^ " 1 "
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   516
  val ms = FuncUtil.Intpairfunc.fold (fn ((i,j), c) => fn  a => if i > j then a else ((i,j),c)::a) (snd m) []
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   517
  val mss = sort ((prod_ord int_ord int_ord) o pairself fst) ms
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   518
 in fold_rev (fn ((i,j),c) => fn a =>
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   519
     pfx ^ string_of_int i ^ " " ^ string_of_int j ^
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   520
     " " ^ decimalize 20 c ^ "\n" ^ a) mss ""
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   521
 end;;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   522
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   523
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   524
(* ------------------------------------------------------------------------- *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   525
(* String in SDPA sparse format for standard SDP problem:                    *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   526
(*                                                                           *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   527
(*    X = v_1 * [M_1] + ... + v_m * [M_m] - [M_0] must be PSD                *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   528
(*    Minimize obj_1 * v_1 + ... obj_m * v_m                                 *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   529
(* ------------------------------------------------------------------------- *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   530
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   531
(*
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
   532
fun sdpa_of_problem obj mats =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   533
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   534
  val m = length mats - 1
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   535
  val (n,_) = dimensions (hd mats)
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
   536
 in
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   537
  string_of_int m ^ "\n" ^
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   538
  "1\n" ^
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   539
  string_of_int n ^ "\n" ^
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   540
  sdpa_of_vector obj ^
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   541
  fold_rev2 (fn k => fn m => fn a => sdpa_of_matrix (k - 1) m ^ a) (1 upto length mats) mats ""
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   542
 end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   543
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   544
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   545
fun index_char str chr pos =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   546
  if pos >= String.size str then ~1
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   547
  else if String.sub(str,pos) = chr then pos
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   548
  else index_char str chr (pos + 1);
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   549
fun rat_of_quotient (a,b) = if b = 0 then rat_0 else Rat.rat_of_quotient (a,b);
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   550
fun rat_of_string s =
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   551
 let val n = index_char s #"/" 0 in
33035
15eab423e573 standardized basic operations on type option;
wenzelm
parents: 33029
diff changeset
   552
  if n = ~1 then s |> Int.fromString |> the |> Rat.rat_of_int
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   553
  else
32949
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
   554
   let val SOME numer = Int.fromString(String.substring(s,0,n))
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
   555
       val SOME den = Int.fromString (String.substring(s,n+1,String.size s - n - 1))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   556
   in rat_of_quotient(numer, den)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   557
   end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   558
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   559
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   560
(*
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36350
diff changeset
   561
fun isspace x = (x = " ");
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   562
*)
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36350
diff changeset
   563
fun isnum x = member (op =) ["0","1","2","3","4","5","6","7","8","9"] x;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   564
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   565
(* More parser basics.                                                       *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   566
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   567
(*
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   568
 val word = Scan.this_string
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   569
 fun token s =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   570
  Scan.repeat ($$ " ") |-- word s --| Scan.repeat ($$ " ")
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   571
*)
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   572
 val numeral = Scan.one isnum
33906
1eebf19b773e fixed csdp output parser
Philipp Meyer
parents: 33069
diff changeset
   573
 val decimalint = Scan.repeat1 numeral >> (rat_of_string o implode)
1eebf19b773e fixed csdp output parser
Philipp Meyer
parents: 33069
diff changeset
   574
 val decimalfrac = Scan.repeat1 numeral
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   575
    >> (fn s => rat_of_string(implode s) // pow10 (length s))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   576
 val decimalsig =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   577
    decimalint -- Scan.option (Scan.$$ "." |-- decimalfrac)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   578
    >> (fn (h,NONE) => h | (h,SOME x) => h +/ x)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   579
 fun signed prs =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   580
       $$ "-" |-- prs >> Rat.neg
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   581
    || $$ "+" |-- prs
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   582
    || prs;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   583
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   584
fun emptyin def xs = if null xs then (def,xs) else Scan.fail xs
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   585
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   586
 val exponent = ($$ "e" || $$ "E") |-- signed decimalint;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   587
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   588
 val decimal = signed decimalsig -- (emptyin rat_0|| exponent)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   589
    >> (fn (h, x) => h */ pow10 (int_of_rat x));
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   590
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   591
 fun mkparser p s =
40627
becf5d5187cc renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents: 39027
diff changeset
   592
  let val (x,rst) = p (raw_explode s)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   593
  in if null rst then x
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   594
     else error "mkparser: unparsed input"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   595
  end;;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   596
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
   597
(* Parse back csdp output.                                                      *)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   598
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   599
 fun ignore _ = ((),[])
33067
a70d5baa53ee use plain Scan.repeat (NB: Scan.bulk is for cascading sources -- mostly interna use);
wenzelm
parents: 33042
diff changeset
   600
 fun csdpoutput inp =
a70d5baa53ee use plain Scan.repeat (NB: Scan.bulk is for cascading sources -- mostly interna use);
wenzelm
parents: 33042
diff changeset
   601
   ((decimal -- Scan.repeat (Scan.$$ " " |-- Scan.option decimal) >>
a70d5baa53ee use plain Scan.repeat (NB: Scan.bulk is for cascading sources -- mostly interna use);
wenzelm
parents: 33042
diff changeset
   602
    (fn (h,to) => map_filter I ((SOME h)::to))) --| ignore >> vector_of_list) inp
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   603
 val parse_csdpoutput = mkparser csdpoutput
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   604
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
   605
(* Run prover on a problem in linear form.                       *)
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
   606
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   607
(*
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
   608
fun run_problem prover obj mats =
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
   609
  parse_csdpoutput (prover (sdpa_of_problem obj mats))
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   610
*)
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
   611
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   612
(* Try some apparently sensible scaling first. Note that this is purely to   *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   613
(* get a cleaner translation to floating-point, and doesn't affect any of    *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   614
(* the results, in principle. In practice it seems a lot better when there   *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   615
(* are extreme numbers in the original problem.                              *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   616
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   617
  (* Version for (int*int) keys *)
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   618
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   619
local
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   620
  fun max_rat x y = if x </ y then y else x
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   621
  fun common_denominator fld amat acc =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   622
      fld (fn (m,c) => fn a => lcm_rat (denominator_rat c) a) amat acc
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   623
  fun maximal_element fld amat acc =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   624
    fld (fn (m,c) => fn maxa => max_rat maxa (abs_rat c)) amat acc
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   625
fun float_of_rat x = let val (a,b) = Rat.quotient_of_rat x
41490
0f1e411a1448 eliminated obsolete LargeInt -- Int is unbounded;
wenzelm
parents: 41474
diff changeset
   626
                     in Real.fromInt a / Real.fromInt b end;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   627
in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   628
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   629
fun pi_scale_then solver (obj:vector)  mats =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   630
 let
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   631
  val cd1 = fold_rev (common_denominator FuncUtil.Intpairfunc.fold) mats (rat_1)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   632
  val cd2 = common_denominator FuncUtil.Intfunc.fold (snd obj)  (rat_1)
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   633
  val mats' = map (FuncUtil.Intpairfunc.map (fn _ => fn x => cd1 */ x)) mats
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   634
  val obj' = vector_cmul cd2 obj
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   635
  val max1 = fold_rev (maximal_element FuncUtil.Intpairfunc.fold) mats' (rat_0)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   636
  val max2 = maximal_element FuncUtil.Intfunc.fold (snd obj') (rat_0)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   637
  val scal1 = pow2 (20 - trunc(Math.ln (float_of_rat max1) / Math.ln 2.0))
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   638
  val scal2 = pow2 (20 - trunc(Math.ln (float_of_rat max2) / Math.ln 2.0))
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   639
  val mats'' = map (FuncUtil.Intpairfunc.map (fn _ => fn x => x */ scal1)) mats'
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   640
  val obj'' = vector_cmul scal2 obj'
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   641
 in solver obj'' mats''
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   642
  end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   643
end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   644
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   645
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   646
(* Try some apparently sensible scaling first. Note that this is purely to   *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   647
(* get a cleaner translation to floating-point, and doesn't affect any of    *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   648
(* the results, in principle. In practice it seems a lot better when there   *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   649
(* are extreme numbers in the original problem.                              *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   650
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   651
  (* Version for (int*int*int) keys *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   652
local
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   653
  fun max_rat x y = if x </ y then y else x
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   654
  fun common_denominator fld amat acc =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   655
      fld (fn (_,c) => fn a => lcm_rat (denominator_rat c) a) amat acc
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   656
  fun maximal_element fld amat acc =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   657
    fld (fn (_,c) => fn maxa => max_rat maxa (abs_rat c)) amat acc
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   658
fun float_of_rat x = let val (a,b) = Rat.quotient_of_rat x
41490
0f1e411a1448 eliminated obsolete LargeInt -- Int is unbounded;
wenzelm
parents: 41474
diff changeset
   659
                     in Real.fromInt a / Real.fromInt b end;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   660
fun int_of_float x = (trunc x handle Overflow => 0 | Domain => 0)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   661
in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   662
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   663
fun tri_scale_then solver (obj:vector)  mats =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   664
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   665
  val cd1 = fold_rev (common_denominator Inttriplefunc.fold) mats (rat_1)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   666
  val cd2 = common_denominator FuncUtil.Intfunc.fold (snd obj)  (rat_1)
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   667
  val mats' = map (Inttriplefunc.map (fn _ => fn x => cd1 */ x)) mats
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   668
  val obj' = vector_cmul cd2 obj
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   669
  val max1 = fold_rev (maximal_element Inttriplefunc.fold) mats' (rat_0)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   670
  val max2 = maximal_element FuncUtil.Intfunc.fold (snd obj') (rat_0)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   671
  val scal1 = pow2 (20 - int_of_float(Math.ln (float_of_rat max1) / Math.ln 2.0))
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   672
  val scal2 = pow2 (20 - int_of_float(Math.ln (float_of_rat max2) / Math.ln 2.0))
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   673
  val mats'' = map (Inttriplefunc.map (fn _ => fn x => x */ scal1)) mats'
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   674
  val obj'' = vector_cmul scal2 obj'
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   675
 in solver obj'' mats''
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   676
  end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   677
end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   678
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   679
(* Round a vector to "nice" rationals.                                       *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   680
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   681
fun nice_rational n x = round_rat (n */ x) // n;;
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   682
fun nice_vector n ((d,v) : vector) =
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   683
 (d, FuncUtil.Intfunc.fold (fn (i,c) => fn a =>
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   684
   let val y = nice_rational n c
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   685
   in if c =/ rat_0 then a
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   686
      else FuncUtil.Intfunc.update (i,y) a end) v FuncUtil.Intfunc.empty):vector
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   687
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   688
fun dest_ord f x = is_equal (f x);
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   689
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   690
(* Stuff for "equations" ((int*int*int)->num functions).                         *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   691
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   692
fun tri_equation_cmul c eq =
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   693
  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (fn _ => fn d => c */ d) eq;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   694
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   695
fun tri_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0) eq1 eq2;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   696
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   697
fun tri_equation_eval assig eq =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   698
 let fun value v = Inttriplefunc.apply assig v
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   699
 in Inttriplefunc.fold (fn (v, c) => fn a => a +/ value v */ c) eq rat_0
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   700
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   701
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   702
(* Eliminate among linear equations: return unconstrained variables and      *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   703
(* assignments for the others in terms of them. We give one pseudo-variable  *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   704
(* "one" that's used for a constant term.                                    *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   705
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   706
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   707
local
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   708
  fun extract_first p l = case l of  (* FIXME : use find_first instead *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   709
   [] => error "extract_first"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   710
 | h::t => if p h then (h,t) else
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   711
          let val (k,s) = extract_first p t in (k,h::s) end
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   712
fun eliminate vars dun eqs = case vars of
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   713
  [] => if forall Inttriplefunc.is_empty eqs then dun
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   714
        else raise Unsolvable
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   715
 | v::vs =>
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   716
  ((let
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   717
    val (eq,oeqs) = extract_first (fn e => Inttriplefunc.defined e v) eqs
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   718
    val a = Inttriplefunc.apply eq v
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   719
    val eq' = tri_equation_cmul ((Rat.neg rat_1) // a) (Inttriplefunc.delete_safe v eq)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   720
    fun elim e =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   721
     let val b = Inttriplefunc.tryapplyd e v rat_0
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   722
     in if b =/ rat_0 then e else
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   723
        tri_equation_add e (tri_equation_cmul (Rat.neg b // a) eq)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   724
     end
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   725
   in eliminate vs (Inttriplefunc.update (v,eq') (Inttriplefunc.map (K elim) dun)) (map elim oeqs)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   726
   end)
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
   727
  handle Failure _ => eliminate vs dun eqs)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   728
in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   729
fun tri_eliminate_equations one vars eqs =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   730
 let
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   731
  val assig = eliminate vars Inttriplefunc.empty eqs
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   732
  val vs = Inttriplefunc.fold (fn (_, f) => fn a => remove (dest_ord triple_int_ord) one (Inttriplefunc.dom f) @ a) assig []
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   733
  in (distinct (dest_ord triple_int_ord) vs, assig)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   734
  end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   735
end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   736
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   737
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   738
(* Eliminate all variables, in an essentially arbitrary order.               *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   739
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   740
fun tri_eliminate_all_equations one =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   741
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   742
  fun choose_variable eq =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   743
   let val (v,_) = Inttriplefunc.choose eq
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   744
   in if is_equal (triple_int_ord(v,one)) then
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   745
      let val eq' = Inttriplefunc.delete_safe v eq
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   746
      in if Inttriplefunc.is_empty eq' then error "choose_variable"
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   747
         else fst (Inttriplefunc.choose eq')
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   748
      end
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   749
    else v
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   750
   end
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   751
  fun eliminate dun eqs = case eqs of
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   752
    [] => dun
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   753
  | eq::oeqs =>
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   754
    if Inttriplefunc.is_empty eq then eliminate dun oeqs else
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   755
    let val v = choose_variable eq
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   756
        val a = Inttriplefunc.apply eq v
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   757
        val eq' = tri_equation_cmul ((Rat.rat_of_int ~1) // a)
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   758
                   (Inttriplefunc.delete_safe v eq)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   759
        fun elim e =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   760
         let val b = Inttriplefunc.tryapplyd e v rat_0
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   761
         in if b =/ rat_0 then e
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   762
            else tri_equation_add e (tri_equation_cmul (Rat.neg b // a) eq)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   763
         end
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   764
    in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.map (K elim) dun))
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   765
                 (map elim oeqs)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   766
    end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   767
in fn eqs =>
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   768
 let
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   769
  val assig = eliminate Inttriplefunc.empty eqs
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   770
  val vs = Inttriplefunc.fold (fn (_, f) => fn a => remove (dest_ord triple_int_ord) one (Inttriplefunc.dom f) @ a) assig []
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   771
 in (distinct (dest_ord triple_int_ord) vs,assig)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   772
 end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   773
end;
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   774
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   775
(* Solve equations by assigning arbitrary numbers.                           *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   776
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   777
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   778
fun tri_solve_equations one eqs =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   779
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   780
  val (vars,assigs) = tri_eliminate_all_equations one eqs
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   781
  val vfn = fold_rev (fn v => Inttriplefunc.update(v,rat_0)) vars
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   782
            (Inttriplefunc.onefunc(one, Rat.rat_of_int ~1))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   783
  val ass =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   784
    Inttriplefunc.combine (curry op +/) (K false)
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   785
    (Inttriplefunc.map (K (tri_equation_eval vfn)) assigs) vfn
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   786
 in if forall (fn e => tri_equation_eval ass e =/ rat_0) eqs
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   787
    then Inttriplefunc.delete_safe one ass else raise Sanity
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   788
 end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   789
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   790
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   791
(* Multiply equation-parametrized poly by regular poly and add accumulator.  *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   792
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   793
fun tri_epoly_pmul p q acc =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   794
 FuncUtil.Monomialfunc.fold (fn (m1, c) => fn a =>
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   795
  FuncUtil.Monomialfunc.fold (fn (m2,e) => fn b =>
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   796
   let val m =  monomial_mul m1 m2
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   797
       val es = FuncUtil.Monomialfunc.tryapplyd b m Inttriplefunc.empty
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   798
   in FuncUtil.Monomialfunc.update (m,tri_equation_add (tri_equation_cmul c e) es) b
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   799
   end) q a) p acc ;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   800
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   801
(* Usual operations on equation-parametrized poly.                           *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   802
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   803
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   804
fun tri_epoly_cmul c l =
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   805
  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (K (tri_equation_cmul c)) l;;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   806
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   807
val tri_epoly_neg = tri_epoly_cmul (Rat.rat_of_int ~1);
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   808
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   809
val tri_epoly_add = Inttriplefunc.combine tri_equation_add Inttriplefunc.is_empty;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   810
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   811
fun tri_epoly_sub p q = tri_epoly_add p (tri_epoly_neg q);;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   812
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   813
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   814
(* Stuff for "equations" ((int*int)->num functions).                         *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   815
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   816
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   817
fun pi_equation_cmul c eq =
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   818
  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (fn _ => fn d => c */ d) eq;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   819
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   820
fun pi_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0) eq1 eq2;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   821
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   822
fun pi_equation_eval assig eq =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   823
 let fun value v = Inttriplefunc.apply assig v
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   824
 in Inttriplefunc.fold (fn (v, c) => fn a => a +/ value v */ c) eq rat_0
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   825
 end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   826
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   827
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   828
(* Eliminate among linear equations: return unconstrained variables and      *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   829
(* assignments for the others in terms of them. We give one pseudo-variable  *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   830
(* "one" that's used for a constant term.                                    *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   831
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   832
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   833
local
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   834
fun extract_first p l = case l of
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   835
   [] => error "extract_first"
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   836
 | h::t => if p h then (h,t) else
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   837
          let val (k,s) = extract_first p t in (k,h::s) end
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   838
fun eliminate vars dun eqs = case vars of
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   839
  [] => if forall Inttriplefunc.is_empty eqs then dun
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   840
        else raise Unsolvable
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   841
 | v::vs =>
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   842
   let
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   843
    val (eq,oeqs) = extract_first (fn e => Inttriplefunc.defined e v) eqs
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   844
    val a = Inttriplefunc.apply eq v
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   845
    val eq' = pi_equation_cmul ((Rat.neg rat_1) // a) (Inttriplefunc.delete_safe v eq)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   846
    fun elim e =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   847
     let val b = Inttriplefunc.tryapplyd e v rat_0
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   848
     in if b =/ rat_0 then e else
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   849
        pi_equation_add e (pi_equation_cmul (Rat.neg b // a) eq)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   850
     end
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   851
   in eliminate vs (Inttriplefunc.update (v,eq') (Inttriplefunc.map (K elim) dun)) (map elim oeqs)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   852
   end
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
   853
  handle Failure _ => eliminate vs dun eqs
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   854
in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   855
fun pi_eliminate_equations one vars eqs =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   856
 let
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   857
  val assig = eliminate vars Inttriplefunc.empty eqs
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   858
  val vs = Inttriplefunc.fold (fn (_, f) => fn a => remove (dest_ord triple_int_ord) one (Inttriplefunc.dom f) @ a) assig []
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   859
  in (distinct (dest_ord triple_int_ord) vs, assig)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   860
  end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   861
end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   862
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   863
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   864
(* Eliminate all variables, in an essentially arbitrary order.               *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   865
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   866
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   867
fun pi_eliminate_all_equations one =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   868
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   869
  fun choose_variable eq =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   870
   let val (v,_) = Inttriplefunc.choose eq
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   871
   in if is_equal (triple_int_ord(v,one)) then
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   872
      let val eq' = Inttriplefunc.delete_safe v eq
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   873
      in if Inttriplefunc.is_empty eq' then error "choose_variable"
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   874
         else fst (Inttriplefunc.choose eq')
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   875
      end
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   876
    else v
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   877
   end
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   878
  fun eliminate dun eqs = case eqs of
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   879
    [] => dun
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   880
  | eq::oeqs =>
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   881
    if Inttriplefunc.is_empty eq then eliminate dun oeqs else
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   882
    let val v = choose_variable eq
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   883
        val a = Inttriplefunc.apply eq v
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   884
        val eq' = pi_equation_cmul ((Rat.rat_of_int ~1) // a)
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   885
                   (Inttriplefunc.delete_safe v eq)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   886
        fun elim e =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   887
         let val b = Inttriplefunc.tryapplyd e v rat_0
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   888
         in if b =/ rat_0 then e
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   889
            else pi_equation_add e (pi_equation_cmul (Rat.neg b // a) eq)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   890
         end
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   891
    in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.map (K elim) dun))
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   892
                 (map elim oeqs)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   893
    end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   894
in fn eqs =>
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   895
 let
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   896
  val assig = eliminate Inttriplefunc.empty eqs
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   897
  val vs = Inttriplefunc.fold (fn (_, f) => fn a => remove (dest_ord triple_int_ord) one (Inttriplefunc.dom f) @ a) assig []
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   898
 in (distinct (dest_ord triple_int_ord) vs,assig)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   899
 end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   900
end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   901
*)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   902
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   903
(* Solve equations by assigning arbitrary numbers.                           *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   904
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   905
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   906
fun pi_solve_equations one eqs =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   907
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   908
  val (vars,assigs) = pi_eliminate_all_equations one eqs
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   909
  val vfn = fold_rev (fn v => Inttriplefunc.update(v,rat_0)) vars
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   910
            (Inttriplefunc.onefunc(one, Rat.rat_of_int ~1))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   911
  val ass =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   912
    Inttriplefunc.combine (curry op +/) (K false)
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   913
    (Inttriplefunc.map (K (pi_equation_eval vfn)) assigs) vfn
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   914
 in if forall (fn e => pi_equation_eval ass e =/ rat_0) eqs
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   915
    then Inttriplefunc.delete_safe one ass else raise Sanity
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   916
 end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   917
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   918
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   919
(* Multiply equation-parametrized poly by regular poly and add accumulator.  *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   920
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   921
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   922
fun pi_epoly_pmul p q acc =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   923
 FuncUtil.Monomialfunc.fold (fn (m1, c) => fn a =>
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   924
  FuncUtil.Monomialfunc.fold (fn (m2,e) => fn b =>
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   925
   let val m =  monomial_mul m1 m2
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   926
       val es = FuncUtil.Monomialfunc.tryapplyd b m Inttriplefunc.empty
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   927
   in FuncUtil.Monomialfunc.update (m,pi_equation_add (pi_equation_cmul c e) es) b
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   928
   end) q a) p acc ;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   929
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   930
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   931
(* Usual operations on equation-parametrized poly.                           *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   932
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   933
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   934
fun pi_epoly_cmul c l =
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
   935
  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (K (pi_equation_cmul c)) l;;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   936
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   937
val pi_epoly_neg = pi_epoly_cmul (Rat.rat_of_int ~1);
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   938
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   939
val pi_epoly_add = Inttriplefunc.combine pi_equation_add Inttriplefunc.is_empty;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   940
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   941
fun pi_epoly_sub p q = pi_epoly_add p (pi_epoly_neg q);;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   942
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   943
fun allpairs f l1 l2 =  fold_rev (fn x => (curry (op @)) (map (f x) l2)) l1 [];
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   944
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   945
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   946
(* Hence produce the "relevant" monomials: those whose squares lie in the    *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   947
(* Newton polytope of the monomials in the input. (This is enough according  *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   948
(* to Reznik: "Extremal PSD forms with few terms", Duke Math. Journal,       *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   949
(* vol 45, pp. 363--374, 1978.                                               *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   950
(*                                                                           *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   951
(* These are ordered in sort of decreasing degree. In particular the         *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   952
(* constant monomial is last; this gives an order in diagonalization of the  *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   953
(* quadratic form that will tend to display constants.                       *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   954
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   955
(* Diagonalize (Cholesky/LDU) the matrix corresponding to a quadratic form.  *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   956
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   957
local
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   958
fun diagonalize n i m =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   959
 if FuncUtil.Intpairfunc.is_empty (snd m) then []
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   960
 else
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   961
  let val a11 = FuncUtil.Intpairfunc.tryapplyd (snd m) (i,i) rat_0
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
   962
  in if a11 </ rat_0 then raise Failure "diagonalize: not PSD"
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   963
    else if a11 =/ rat_0 then
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   964
          if FuncUtil.Intfunc.is_empty (snd (row i m)) then diagonalize n (i + 1) m
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
   965
          else raise Failure "diagonalize: not PSD ___ "
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   966
    else
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   967
     let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   968
      val v = row i m
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   969
      val v' = (fst v, FuncUtil.Intfunc.fold (fn (i, c) => fn a =>
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   970
       let val y = c // a11
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   971
       in if y = rat_0 then a else FuncUtil.Intfunc.update (i,y) a
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   972
       end)  (snd v) FuncUtil.Intfunc.empty)
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   973
      fun upt0 x y a = if y = rat_0 then a else FuncUtil.Intpairfunc.update (x,y) a
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   974
      val m' =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   975
      ((n,n),
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   976
      iter (i+1,n) (fn j =>
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   977
          iter (i+1,n) (fn k =>
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
   978
              (upt0 (j,k) (FuncUtil.Intpairfunc.tryapplyd (snd m) (j,k) rat_0 -/ FuncUtil.Intfunc.tryapplyd (snd v) j rat_0 */ FuncUtil.Intfunc.tryapplyd (snd v') k rat_0))))
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   979
          FuncUtil.Intpairfunc.empty)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   980
     in (a11,v')::diagonalize n (i + 1) m'
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   981
     end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   982
  end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   983
in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   984
fun diag m =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   985
 let
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   986
   val nn = dimensions m
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   987
   val n = fst nn
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
   988
 in if snd nn <> n then error "diagonalize: non-square matrix"
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   989
    else diagonalize n 1 m
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   990
 end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   991
end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   992
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   993
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   994
fun gcd_rat a b = Rat.rat_of_int (Integer.gcd (int_of_rat a) (int_of_rat b));
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   995
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   996
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   997
(* Adjust a diagonalization to collect rationals at the start.               *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
   998
  (* FIXME : Potentially polymorphic keys, but here only: integers!! *)
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
   999
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1000
local
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1001
 fun upd0 x y a = if y =/ rat_0 then a else FuncUtil.Intfunc.update(x,y) a;
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1002
 fun mapa f (d,v) =
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1003
  (d, FuncUtil.Intfunc.fold (fn (i,c) => fn a => upd0 i (f c) a) v FuncUtil.Intfunc.empty)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1004
 fun adj (c,l) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1005
 let val a =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1006
  FuncUtil.Intfunc.fold (fn (_,c) => fn a => lcm_rat a (denominator_rat c))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1007
    (snd l) rat_1 //
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1008
  FuncUtil.Intfunc.fold (fn (i,c) => fn a => gcd_rat a (numerator_rat c))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1009
    (snd l) rat_0
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1010
  in ((c // (a */ a)),mapa (fn x => a */ x) l)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1011
  end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1012
in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1013
fun deration d = if null d then (rat_0,d) else
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1014
 let val d' = map adj d
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1015
     val a = fold (lcm_rat o denominator_rat o fst) d' rat_1 //
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1016
          fold (gcd_rat o numerator_rat o fst) d' rat_0
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1017
 in ((rat_1 // a),map (fn (c,l) => (a */ c,l)) d')
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1018
 end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1019
end;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1020
*)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1021
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1022
(* Enumeration of monomials with given multidegree bound.                    *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1023
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1024
fun enumerate_monomials d vars =
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1025
 if d < 0 then []
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1026
 else if d = 0 then [FuncUtil.Ctermfunc.empty]
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1027
 else if null vars then [monomial_1] else
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1028
 let val alts =
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33042
diff changeset
  1029
  map_range (fn k => let val oths = enumerate_monomials (d - k) (tl vars)
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33042
diff changeset
  1030
               in map (fn ks => if k = 0 then ks else FuncUtil.Ctermfunc.update (hd vars, k) ks) oths end) (d + 1)
32830
47c1b15c03fe replaced and tuned uses of foldr1
Philipp Meyer
parents: 32829
diff changeset
  1031
 in flat alts
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1032
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1033
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1034
(* Enumerate products of distinct input polys with degree <= d.              *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1035
(* We ignore any constant input polynomials.                                 *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1036
(* Give the output polynomial and a record of how it was derived.            *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1037
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1038
fun enumerate_products d pols =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1039
if d = 0 then [(poly_const rat_1,RealArith.Rational_lt rat_1)]
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1040
else if d < 0 then [] else
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1041
case pols of
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1042
   [] => [(poly_const rat_1,RealArith.Rational_lt rat_1)]
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1043
 | (p,b)::ps =>
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1044
    let val e = multidegree p
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1045
    in if e = 0 then enumerate_products d ps else
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1046
       enumerate_products d ps @
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1047
       map (fn (q,c) => (poly_mul p q,RealArith.Product(b,c)))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1048
         (enumerate_products (d - e) ps)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1049
    end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1050
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1051
(* Convert regular polynomial. Note that we treat (0,0,0) as -1.             *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1052
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1053
fun epoly_of_poly p =
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1054
  FuncUtil.Monomialfunc.fold (fn (m,c) => fn a => FuncUtil.Monomialfunc.update (m, Inttriplefunc.onefunc ((0,0,0), Rat.neg c)) a) p FuncUtil.Monomialfunc.empty;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1055
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1056
(* String for block diagonal matrix numbered k.                              *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1057
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1058
fun sdpa_of_blockdiagonal k m =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1059
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1060
  val pfx = string_of_int k ^" "
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1061
  val ents =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1062
    Inttriplefunc.fold
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1063
      (fn ((b,i,j),c) => fn a => if i > j then a else ((b,i,j),c)::a)
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1064
      m []
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1065
  val entss = sort (triple_int_ord o pairself fst) ents
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1066
 in fold_rev (fn ((b,i,j),c) => fn a =>
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1067
     pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1068
     " " ^ decimalize 20 c ^ "\n" ^ a) entss ""
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1069
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1070
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1071
(* SDPA for problem using block diagonal (i.e. multiple SDPs)                *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1072
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1073
fun sdpa_of_blockproblem nblocks blocksizes obj mats =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1074
 let val m = length mats - 1
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1075
 in
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1076
  string_of_int m ^ "\n" ^
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1077
  string_of_int nblocks ^ "\n" ^
32830
47c1b15c03fe replaced and tuned uses of foldr1
Philipp Meyer
parents: 32829
diff changeset
  1078
  (space_implode " " (map string_of_int blocksizes)) ^
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1079
  "\n" ^
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1080
  sdpa_of_vector obj ^
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1081
  fold_rev2 (fn k => fn m => fn a => sdpa_of_blockdiagonal (k - 1) m ^ a)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1082
    (1 upto length mats) mats ""
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1083
 end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1084
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1085
(* Run prover on a problem in block diagonal form.                       *)
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1086
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1087
fun run_blockproblem prover nblocks blocksizes obj mats=
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1088
  parse_csdpoutput (prover (sdpa_of_blockproblem nblocks blocksizes obj mats))
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1089
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1090
(* 3D versions of matrix operations to consider blocks separately.           *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1091
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1092
val bmatrix_add = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0);
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1093
fun bmatrix_cmul c bm =
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1094
  if c =/ rat_0 then Inttriplefunc.empty
39027
e4262f9e6a4e Table.map replaces Table.map'
haftmann
parents: 38805
diff changeset
  1095
  else Inttriplefunc.map (fn _ => fn x => c */ x) bm;
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1096
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1097
val bmatrix_neg = bmatrix_cmul (Rat.rat_of_int ~1);
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1098
(*
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1099
fun bmatrix_sub m1 m2 = bmatrix_add m1 (bmatrix_neg m2);;
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1100
*)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1101
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1102
(* Smash a block matrix into components.                                     *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1103
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1104
fun blocks blocksizes bm =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1105
 map (fn (bs,b0) =>
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1106
      let val m = Inttriplefunc.fold
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1107
          (fn ((b,i,j),c) => fn a => if b = b0 then FuncUtil.Intpairfunc.update ((i,j),c) a else a) bm FuncUtil.Intpairfunc.empty
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1108
          val _ = FuncUtil.Intpairfunc.fold (fn ((i,j),_) => fn a => max a (max i j)) m 0
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1109
      in (((bs,bs),m):matrix) end)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1110
 (blocksizes ~~ (1 upto length blocksizes));;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1111
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1112
(* FIXME : Get rid of this !!!*)
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1113
local
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1114
  fun tryfind_with msg _ [] = raise Failure msg
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1115
    | tryfind_with _ f (x::xs) = (f x handle Failure s => tryfind_with s f xs);
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1116
in
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1117
  fun tryfind f = tryfind_with "tryfind" f
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1118
end
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1119
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1120
(* Positiv- and Nullstellensatz. Flag "linf" forces a linear representation. *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1121
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1122
38805
b09d8603f865 Sum_Of_Squares: proper configuration options;
wenzelm
parents: 38795
diff changeset
  1123
fun real_positivnullstellensatz_general ctxt prover linf d eqs leqs pol =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1124
let
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33039
diff changeset
  1125
 val vars = fold_rev (union (op aconvc) o poly_variables)
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33039
diff changeset
  1126
   (pol :: eqs @ map fst leqs) []
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1127
 val monoid = if linf then
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1128
      (poly_const rat_1,RealArith.Rational_lt rat_1)::
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1129
      (filter (fn (p,_) => multidegree p <= d) leqs)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1130
    else enumerate_products d leqs
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1131
 val nblocks = length monoid
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1132
 fun mk_idmultiplier k p =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1133
  let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1134
   val e = d - multidegree p
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1135
   val mons = enumerate_monomials e vars
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1136
   val nons = mons ~~ (1 upto length mons)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1137
  in (mons,
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1138
      fold_rev (fn (m,n) => FuncUtil.Monomialfunc.update(m,Inttriplefunc.onefunc((~k,~n,n),rat_1))) nons FuncUtil.Monomialfunc.empty)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1139
  end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1140
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1141
 fun mk_sqmultiplier k (p,_) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1142
  let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1143
   val e = (d - multidegree p) div 2
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1144
   val mons = enumerate_monomials e vars
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1145
   val nons = mons ~~ (1 upto length mons)
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1146
  in (mons,
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1147
      fold_rev (fn (m1,n1) =>
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1148
       fold_rev (fn (m2,n2) => fn  a =>
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1149
        let val m = monomial_mul m1 m2
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1150
        in if n1 > n2 then a else
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1151
          let val c = if n1 = n2 then rat_1 else rat_2
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1152
              val e = FuncUtil.Monomialfunc.tryapplyd a m Inttriplefunc.empty
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1153
          in FuncUtil.Monomialfunc.update(m, tri_equation_add (Inttriplefunc.onefunc((k,n1,n2), c)) e) a
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1154
          end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1155
        end)  nons)
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1156
       nons FuncUtil.Monomialfunc.empty)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1157
  end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1158
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1159
  val (sqmonlist,sqs) = split_list (map2 mk_sqmultiplier (1 upto length monoid) monoid)
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1160
  val (_(*idmonlist*),ids) =  split_list(map2 mk_idmultiplier (1 upto length eqs) eqs)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1161
  val blocksizes = map length sqmonlist
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1162
  val bigsum =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1163
    fold_rev2 (fn p => fn q => fn a => tri_epoly_pmul p q a) eqs ids
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1164
            (fold_rev2 (fn (p,_) => fn s => fn a => tri_epoly_pmul p s a) monoid sqs
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1165
                     (epoly_of_poly(poly_neg pol)))
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1166
  val eqns = FuncUtil.Monomialfunc.fold (fn (_,e) => fn a => e::a) bigsum []
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1167
  val (pvs,assig) = tri_eliminate_all_equations (0,0,0) eqns
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1168
  val qvars = (0,0,0)::pvs
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1169
  val allassig = fold_rev (fn v => Inttriplefunc.update(v,(Inttriplefunc.onefunc(v,rat_1)))) pvs assig
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1170
  fun mk_matrix v =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1171
    Inttriplefunc.fold (fn ((b,i,j), ass) => fn m =>
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1172
        if b < 0 then m else
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1173
         let val c = Inttriplefunc.tryapplyd ass v rat_0
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1174
         in if c = rat_0 then m else
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1175
            Inttriplefunc.update ((b,j,i), c) (Inttriplefunc.update ((b,i,j), c) m)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1176
         end)
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1177
          allassig Inttriplefunc.empty
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1178
  val diagents = Inttriplefunc.fold
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1179
    (fn ((b,i,j), e) => fn a => if b > 0 andalso i = j then tri_equation_add e a else a)
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1180
    allassig Inttriplefunc.empty
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1181
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1182
  val mats = map mk_matrix qvars
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1183
  val obj = (length pvs,
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1184
            itern 1 pvs (fn v => fn i => FuncUtil.Intfunc.updatep iszero (i,Inttriplefunc.tryapplyd diagents v rat_0))
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1185
                        FuncUtil.Intfunc.empty)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1186
  val raw_vec = if null pvs then vector_0 0
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1187
                else tri_scale_then (run_blockproblem prover nblocks blocksizes) obj mats
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1188
  fun int_element (_,v) i = FuncUtil.Intfunc.tryapplyd v i rat_0
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1189
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1190
  fun find_rounding d =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1191
   let
32949
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
  1192
    val _ =
38805
b09d8603f865 Sum_Of_Squares: proper configuration options;
wenzelm
parents: 38795
diff changeset
  1193
      if Config.get ctxt trace
32949
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
  1194
      then writeln ("Trying rounding with limit "^Rat.string_of_rat d ^ "\n")
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
  1195
      else ()
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1196
    val vec = nice_vector d raw_vec
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1197
    val blockmat = iter (1,dim vec)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1198
     (fn i => fn a => bmatrix_add (bmatrix_cmul (int_element vec i) (nth mats i)) a)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1199
     (bmatrix_neg (nth mats 0))
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1200
    val allmats = blocks blocksizes blockmat
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1201
   in (vec,map diag allmats)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1202
   end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1203
  val (vec,ratdias) =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1204
    if null pvs then find_rounding rat_1
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1205
    else tryfind find_rounding (map Rat.rat_of_int (1 upto 31) @
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1206
                                map pow2 (5 upto 66))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1207
  val newassigs =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1208
    fold_rev (fn k => Inttriplefunc.update (nth pvs (k - 1), int_element vec k))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1209
           (1 upto dim vec) (Inttriplefunc.onefunc ((0,0,0), Rat.rat_of_int ~1))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1210
  val finalassigs =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1211
    Inttriplefunc.fold (fn (v,e) => fn a => Inttriplefunc.update(v, tri_equation_eval newassigs e) a) allassig newassigs
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1212
  fun poly_of_epoly p =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1213
    FuncUtil.Monomialfunc.fold (fn (v,e) => fn a => FuncUtil.Monomialfunc.updatep iszero (v,tri_equation_eval finalassigs e) a)
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1214
          p FuncUtil.Monomialfunc.empty
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1215
  fun  mk_sos mons =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1216
   let fun mk_sq (c,m) =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1217
    (c,fold_rev (fn k=> fn a => FuncUtil.Monomialfunc.updatep iszero (nth mons (k - 1), int_element m k) a)
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1218
                 (1 upto length mons) FuncUtil.Monomialfunc.empty)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1219
   in map mk_sq
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1220
   end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1221
  val sqs = map2 mk_sos sqmonlist ratdias
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1222
  val cfs = map poly_of_epoly ids
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1223
  val msq = filter (fn (_,b) => not (null b)) (map2 pair monoid sqs)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1224
  fun eval_sq sqs = fold_rev (fn (c,q) => poly_add (poly_cmul c (poly_mul q q))) sqs poly_0
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1225
  val sanity =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1226
    fold_rev (fn ((p,_),s) => poly_add (poly_mul p (eval_sq s))) msq
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1227
           (fold_rev2 (fn p => fn q => poly_add (poly_mul p q)) cfs eqs
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1228
                    (poly_neg pol))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1229
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
  1230
in if not(FuncUtil.Monomialfunc.is_empty sanity) then raise Sanity else
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1231
  (cfs,map (fn (a,b) => (snd a,b)) msq)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1232
 end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1233
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1234
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1235
(* Iterative deepening.                                                      *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1236
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1237
fun deepen f n =
32949
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
  1238
  (writeln ("Searching with depth limit " ^ string_of_int n);
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
  1239
    (f n handle Failure s => (writeln ("failed with message: " ^ s); deepen f (n + 1))));
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1240
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1241
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1242
(* Map back polynomials and their composites to a positivstellensatz.        *)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1243
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1244
fun cterm_of_sqterm (c,p) = RealArith.Product(RealArith.Rational_lt c,RealArith.Square p);
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1245
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1246
fun cterm_of_sos (pr,sqs) = if null sqs then pr
32830
47c1b15c03fe replaced and tuned uses of foldr1
Philipp Meyer
parents: 32829
diff changeset
  1247
  else RealArith.Product(pr,foldr1 RealArith.Sum (map cterm_of_sqterm sqs));
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1248
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1249
(* Interface to HOL.                                                         *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1250
local
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1251
  open Conv
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1252
  val concl = Thm.dest_arg o cprop_of
35408
b48ab741683b modernized structure Term_Ord;
wenzelm
parents: 33906
diff changeset
  1253
  fun simple_cterm_ord t u = Term_Ord.fast_term_ord (term_of t, term_of u) = LESS
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1254
in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1255
  (* FIXME: Replace tryfind by get_first !! *)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1256
fun real_nonlinear_prover proof_method ctxt =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1257
 let
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1258
  val {add = _, mul = _, neg = _, pow = _,
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1259
       sub = _, main = real_poly_conv} =
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1260
      Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36751
diff changeset
  1261
      (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1262
     simple_cterm_ord
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1263
  fun mainf cert_choice translator (eqs,les,lts) =
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1264
  let
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1265
   val eq0 = map (poly_of_term o Thm.dest_arg1 o concl) eqs
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1266
   val le0 = map (poly_of_term o Thm.dest_arg o concl) les
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1267
   val lt0 = map (poly_of_term o Thm.dest_arg o concl) lts
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33042
diff changeset
  1268
   val eqp0 = map_index (fn (i, t) => (t,RealArith.Axiom_eq i)) eq0
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33042
diff changeset
  1269
   val lep0 = map_index (fn (i, t) => (t,RealArith.Axiom_le i)) le0
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33042
diff changeset
  1270
   val ltp0 = map_index (fn (i, t) => (t,RealArith.Axiom_lt i)) lt0
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1271
   val (keq,eq) = List.partition (fn (p,_) => multidegree p = 0) eqp0
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1272
   val (klep,lep) = List.partition (fn (p,_) => multidegree p = 0) lep0
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1273
   val (kltp,ltp) = List.partition (fn (p,_) => multidegree p = 0) ltp0
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1274
   fun trivial_axiom (p,ax) =
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1275
    case ax of
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1276
       RealArith.Axiom_eq n => if eval FuncUtil.Ctermfunc.empty p <>/ Rat.zero then nth eqs n
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
  1277
                     else raise Failure "trivial_axiom: Not a trivial axiom"
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1278
     | RealArith.Axiom_le n => if eval FuncUtil.Ctermfunc.empty p </ Rat.zero then nth les n
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
  1279
                     else raise Failure "trivial_axiom: Not a trivial axiom"
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1280
     | RealArith.Axiom_lt n => if eval FuncUtil.Ctermfunc.empty p <=/ Rat.zero then nth lts n
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
  1281
                     else raise Failure "trivial_axiom: Not a trivial axiom"
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1282
     | _ => error "trivial_axiom: Not a trivial axiom"
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1283
   in
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1284
  (let val th = tryfind trivial_axiom (keq @ klep @ kltp)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1285
   in
36751
7f1da69cacb3 split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
haftmann
parents: 36721
diff changeset
  1286
    (fconv_rule (arg_conv (arg1_conv real_poly_conv) then_conv Numeral_Simprocs.field_comp_conv) th, RealArith.Trivial)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1287
   end)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1288
   handle Failure _ =>
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1289
     (let val proof =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1290
       (case proof_method of Certificate certs =>
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1291
         (* choose certificate *)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1292
         let
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1293
           fun chose_cert [] (RealArith.Cert c) = c
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1294
             | chose_cert (RealArith.Left::s) (RealArith.Branch (l, _)) = chose_cert s l
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1295
             | chose_cert (RealArith.Right::s) (RealArith.Branch (_, r)) = chose_cert s r
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1296
             | chose_cert _ _ = error "certificate tree in invalid form"
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1297
         in
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1298
           chose_cert cert_choice certs
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1299
         end
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1300
       | Prover prover =>
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1301
         (* call prover *)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1302
         let
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1303
          val pol = fold_rev poly_mul (map fst ltp) (poly_const Rat.one)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1304
          val leq = lep @ ltp
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1305
          fun tryall d =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1306
           let val e = multidegree pol
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1307
               val k = if e = 0 then 0 else d div e
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1308
               val eq' = map fst eq
38805
b09d8603f865 Sum_Of_Squares: proper configuration options;
wenzelm
parents: 38795
diff changeset
  1309
           in tryfind (fn i => (d,i,real_positivnullstellensatz_general ctxt prover false d eq' leq
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1310
                                 (poly_neg(poly_pow pol i))))
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1311
                   (0 upto k)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1312
           end
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1313
         val (_,i,(cert_ideal,cert_cone)) = deepen tryall 0
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1314
         val proofs_ideal =
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1315
           map2 (fn q => fn (_,ax) => RealArith.Eqmul(q,ax)) cert_ideal eq
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1316
         val proofs_cone = map cterm_of_sos cert_cone
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1317
         val proof_ne = if null ltp then RealArith.Rational_lt Rat.one else
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1318
           let val p = foldr1 RealArith.Product (map snd ltp)
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1319
           in  funpow i (fn q => RealArith.Product(p,q)) (RealArith.Rational_lt Rat.one)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1320
           end
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1321
         in
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1322
           foldr1 RealArith.Sum (proof_ne :: proofs_ideal @ proofs_cone)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1323
         end)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1324
     in
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1325
        (translator (eqs,les,lts) proof, RealArith.Cert proof)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1326
     end)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1327
   end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1328
 in mainf end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1329
end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1330
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1331
fun C f x y = f y x;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1332
  (* FIXME : This is very bad!!!*)
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1333
fun subst_conv eqs t =
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1334
 let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1335
  val t' = fold (Thm.cabs o Thm.lhs_of) eqs t
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
  1336
 in Conv.fconv_rule (Thm.beta_conversion true) (fold (C Thm.combination) eqs (Thm.reflexive t'))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1337
 end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1338
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1339
(* A wrapper that tries to substitute away variables first.                  *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1340
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1341
local
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1342
 open Conv
35408
b48ab741683b modernized structure Term_Ord;
wenzelm
parents: 33906
diff changeset
  1343
  fun simple_cterm_ord t u = Term_Ord.fast_term_ord (term_of t, term_of u) = LESS
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1344
 val concl = Thm.dest_arg o cprop_of
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1345
 val shuffle1 =
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 35625
diff changeset
  1346
   fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: field_simps) })
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1347
 val shuffle2 =
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 35625
diff changeset
  1348
    fconv_rule (rewr_conv @{lemma "(x + a == y) ==  (x == y - (a::real))" by (atomize (full)) (simp add: field_simps)})
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1349
 fun substitutable_monomial fvs tm = case term_of tm of
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1350
    Free(_,@{typ real}) => if not (member (op aconvc) fvs tm) then (Rat.one,tm)
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
  1351
                           else raise Failure "substitutable_monomial"
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1352
  | @{term "op * :: real => _"}$_$(Free _) =>
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1353
     if RealArith.is_ratconst (Thm.dest_arg1 tm) andalso not (member (op aconvc) fvs (Thm.dest_arg tm))
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1354
         then (RealArith.dest_ratconst (Thm.dest_arg1 tm),Thm.dest_arg tm) else raise Failure "substitutable_monomial"
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1355
  | @{term "op + :: real => _"}$_$_ =>
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1356
       (substitutable_monomial (Thm.add_cterm_frees (Thm.dest_arg tm) fvs) (Thm.dest_arg1 tm)
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1357
        handle Failure _ => substitutable_monomial (Thm.add_cterm_frees (Thm.dest_arg1 tm) fvs) (Thm.dest_arg tm))
32332
bc5cec7b2be6 misc changes to SOS by Philipp Meyer:
wenzelm
parents: 32268
diff changeset
  1358
  | _ => raise Failure "substitutable_monomial"
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1359
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1360
  fun isolate_variable v th =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1361
   let val w = Thm.dest_arg1 (cprop_of th)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1362
   in if v aconvc w then th
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1363
      else case term_of w of
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1364
           @{term "op + :: real => _"}$_$_ =>
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1365
              if Thm.dest_arg1 w aconvc v then shuffle2 th
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1366
              else isolate_variable v (shuffle1 th)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1367
          | _ => error "isolate variable : This should not happen?"
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1368
   end
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1369
in
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1370
32268
d50f0cb67578 Functionality for sum of squares to call a remote csdp prover
Philipp Meyer
parents: 32150
diff changeset
  1371
fun real_nonlinear_subst_prover prover ctxt =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1372
 let
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1373
  val {add = _, mul = real_poly_mul_conv, neg = _,
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1374
       pow = _, sub = _, main = real_poly_conv} =
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1375
      Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36751
diff changeset
  1376
      (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1377
     simple_cterm_ord
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1378
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1379
  fun make_substitution th =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1380
   let
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1381
    val (c,v) = substitutable_monomial [] (Thm.dest_arg1(concl th))
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1382
    val th1 = Drule.arg_cong_rule (Thm.capply @{cterm "op * :: real => _"} (RealArith.cterm_of_rat (Rat.inv c))) (mk_meta_eq th)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1383
    val th2 = fconv_rule (binop_conv real_poly_mul_conv) th1
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1384
   in fconv_rule (arg_conv real_poly_conv) (isolate_variable v th2)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1385
   end
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1386
   fun oprconv cv ct =
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1387
    let val g = Thm.dest_fun2 ct
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1388
    in if g aconvc @{cterm "op <= :: real => _"}
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1389
         orelse g aconvc @{cterm "op < :: real => _"}
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1390
       then arg_conv cv ct else arg1_conv cv ct
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1391
    end
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1392
  fun mainf cert_choice translator =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1393
   let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1394
    fun substfirst(eqs,les,lts) =
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1395
      ((let
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1396
           val eth = tryfind make_substitution eqs
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1397
           val modify = fconv_rule (arg_conv (oprconv(subst_conv [eth] then_conv real_poly_conv)))
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1398
       in  substfirst
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1399
             (filter_out (fn t => (Thm.dest_arg1 o Thm.dest_arg o cprop_of) t
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1400
                                   aconvc @{cterm "0::real"}) (map modify eqs),
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1401
                                   map modify les,map modify lts)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1402
       end)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1403
       handle Failure  _ => real_nonlinear_prover prover ctxt cert_choice translator (rev eqs, rev les, rev lts))
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1404
    in substfirst
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1405
   end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1406
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1407
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1408
 in mainf
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1409
 end
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1410
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1411
(* Overall function. *)
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1412
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32332
diff changeset
  1413
fun real_sos prover ctxt =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32740
diff changeset
  1414
  RealArith.gen_prover_real_arith ctxt (real_nonlinear_subst_prover prover ctxt)
31119
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1415
end;
2532bb2d65c7 A decision method for universal multivariate real arithmetic with add
chaieb
parents:
diff changeset
  1416
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1417
val known_sos_constants =
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1418
  [@{term "op ==>"}, @{term "Trueprop"},
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
  1419
   @{term HOL.implies}, @{term HOL.conj}, @{term HOL.disj},
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1420
   @{term "Not"}, @{term "op = :: bool => _"},
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1421
   @{term "All :: (real => _) => _"}, @{term "Ex :: (real => _) => _"},
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1422
   @{term "op = :: real => _"}, @{term "op < :: real => _"},
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1423
   @{term "op <= :: real => _"},
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1424
   @{term "op + :: real => _"}, @{term "op - :: real => _"},
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1425
   @{term "op * :: real => _"}, @{term "uminus :: real => _"},
31512
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1426
   @{term "op / :: real => _"}, @{term "inverse :: real => _"},
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1427
   @{term "op ^ :: real => _"}, @{term "abs :: real => _"},
31512
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1428
   @{term "min :: real => _"}, @{term "max :: real => _"},
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1429
   @{term "0::real"}, @{term "1::real"}, @{term "number_of :: int => real"},
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1430
   @{term "number_of :: int => nat"},
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1431
   @{term "Int.Bit0"}, @{term "Int.Bit1"},
31512
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1432
   @{term "Int.Pls"}, @{term "Int.Min"}];
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1433
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1434
fun check_sos kcts ct =
31512
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1435
 let
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1436
  val t = term_of ct
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1437
  val _ = if not (null (Term.add_tfrees t [])
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1438
                  andalso null (Term.add_tvars t []))
31512
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1439
          then error "SOS: not sos. Additional type varables" else ()
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1440
  val fs = Term.add_frees t []
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1441
  val _ = if exists (fn ((_,T)) => not (T = @{typ "real"})) fs
31512
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1442
          then error "SOS: not sos. Variables with type not real" else ()
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1443
  val vs = Term.add_vars t []
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1444
  val _ = if exists (fn ((_,T)) => not (T = @{typ "real"})) vs
31512
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1445
          then error "SOS: not sos. Variables with type not real" else ()
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1446
  val ukcs = subtract (fn (t,p) => Const p aconv t) kcts (Term.add_consts t [])
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1447
  val _ = if  null ukcs then ()
31512
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1448
              else error ("SOSO: Unknown constants in Subgoal:" ^ commas (map fst ukcs))
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1449
in () end
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1450
32838
d9dfd30af9ae core_sos_tac: SUBPROOF body operates on subgoal 1;
wenzelm
parents: 32831
diff changeset
  1451
fun core_sos_tac print_cert prover = SUBPROOF (fn {concl, context, ...} =>
32831
1352736e5727 changed core_sos_tac to use SUBPROOF
Philipp Meyer
parents: 32830
diff changeset
  1452
  let
32838
d9dfd30af9ae core_sos_tac: SUBPROOF body operates on subgoal 1;
wenzelm
parents: 32831
diff changeset
  1453
    val _ = check_sos known_sos_constants concl
d9dfd30af9ae core_sos_tac: SUBPROOF body operates on subgoal 1;
wenzelm
parents: 32831
diff changeset
  1454
    val (ths, certificates) = real_sos prover context (Thm.dest_arg concl)
d9dfd30af9ae core_sos_tac: SUBPROOF body operates on subgoal 1;
wenzelm
parents: 32831
diff changeset
  1455
    val _ = print_cert certificates
d9dfd30af9ae core_sos_tac: SUBPROOF body operates on subgoal 1;
wenzelm
parents: 32831
diff changeset
  1456
  in rtac ths 1 end)
31131
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1457
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1458
fun default_SOME _ NONE v = SOME v
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1459
  | default_SOME _ (SOME v) _ = SOME v;
31131
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1460
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1461
fun lift_SOME f NONE a = f a
44453
082edd97ffe1 comment out dead code to avoid compiler warnings
huffman
parents: 42616
diff changeset
  1462
  | lift_SOME _ (SOME a) _ = SOME a;
31131
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1463
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1464
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1465
local
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1466
 val is_numeral = can (HOLogic.dest_number o term_of)
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1467
in
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1468
fun get_denom b ct = case term_of ct of
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1469
  @{term "op / :: real => _"} $ _ $ _ =>
31131
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1470
     if is_numeral (Thm.dest_arg ct) then get_denom b (Thm.dest_arg1 ct)
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1471
     else default_SOME (get_denom b) (get_denom b (Thm.dest_arg ct))   (Thm.dest_arg ct, b)
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1472
 | @{term "op < :: real => _"} $ _ $ _ => lift_SOME (get_denom true) (get_denom true (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1473
 | @{term "op <= :: real => _"} $ _ $ _ => lift_SOME (get_denom true) (get_denom true (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1474
 | _ $ _ => lift_SOME (get_denom b) (get_denom b (Thm.dest_fun ct)) (Thm.dest_arg ct)
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1475
 | _ => NONE
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1476
end;
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1477
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1478
fun elim_one_denom_tac ctxt =
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1479
CSUBGOAL (fn (P,i) =>
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1480
 case get_denom false P of
31131
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1481
   NONE => no_tac
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1482
 | SOME (d,ord) =>
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1483
     let
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1484
      val ss = simpset_of ctxt addsimps @{thms field_simps}
31131
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1485
               addsimps [@{thm nonzero_power_divide}, @{thm power_divide}]
32839
a007a7cd8c2f tuned header;
wenzelm
parents: 32838
diff changeset
  1486
      val th = instantiate' [] [SOME d, SOME (Thm.dest_arg P)]
31131
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1487
         (if ord then @{lemma "(d=0 --> P) & (d>0 --> P) & (d<(0::real) --> P) ==> P" by auto}
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1488
          else @{lemma "(d=0 --> P) & (d ~= (0::real) --> P) ==> P" by blast})
32949
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
  1489
     in rtac th i THEN Simplifier.asm_full_simp_tac ss i end);
31131
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1490
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1491
fun elim_denom_tac ctxt i = REPEAT (elim_one_denom_tac ctxt i);
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1492
32949
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
  1493
fun sos_tac print_cert prover ctxt =
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35408
diff changeset
  1494
  Object_Logic.full_atomize_tac THEN'
32949
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
  1495
  elim_denom_tac ctxt THEN'
aa6c470a962a eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32839
diff changeset
  1496
  core_sos_tac print_cert prover ctxt;
31131
d9752181691a Now deals with division
chaieb
parents: 31119
diff changeset
  1497
31512
27118561c2e0 Tuned sos tactic to reject non SOS goals
chaieb
parents: 31131
diff changeset
  1498
end;