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