src/HOL/Matrix/cplex/Cplex_tools.ML
author haftmann
Fri, 20 Oct 2006 10:44:33 +0200
changeset 21056 2cfe839e8d58
parent 17521 0f1c48de39f5
child 21858 05f57309170c
permissions -rw-r--r--
Symtab.foldl replaced by Symtab.fold
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
     1
(*  Title:      HOL/Matrix/cplex/Cplex_tools.ML
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
     2
    ID:         $Id$
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
     3
    Author:     Steven Obua
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
     4
*)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
     5
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
     6
signature CPLEX =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
     7
sig
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
     8
    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
     9
    datatype cplexTerm = cplexVar of string | cplexNum of string | cplexInf 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    10
                       | cplexNeg of cplexTerm 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    11
                       | cplexProd of cplexTerm * cplexTerm 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    12
                       | cplexSum of (cplexTerm list) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    13
    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    14
    datatype cplexComp = cplexLe | cplexLeq | cplexEq | cplexGe | cplexGeq 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    15
								  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    16
    datatype cplexGoal = cplexMinimize of cplexTerm 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    17
		       | cplexMaximize of cplexTerm
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    18
					  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    19
    datatype cplexConstr = cplexConstr of cplexComp * 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    20
					  (cplexTerm * cplexTerm)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    21
					  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    22
    datatype cplexBounds = cplexBounds of cplexTerm * cplexComp * cplexTerm 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    23
					  * cplexComp * cplexTerm
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    24
			 | cplexBound of cplexTerm * cplexComp * cplexTerm
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    25
					 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    26
    datatype cplexProg = cplexProg of string 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    27
				      * cplexGoal 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    28
				      * ((string option * cplexConstr) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    29
					     list) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    30
				      * cplexBounds list
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    31
				      
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    32
    datatype cplexResult = Unbounded 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    33
			 | Infeasible 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    34
			 | Undefined
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    35
			 | Optimal of string * 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    36
				      (((* name *) string * 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    37
					(* value *) string) list)
16966
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
    38
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
    39
    datatype cplexSolver = SOLVER_DEFAULT | SOLVER_CPLEX | SOLVER_GLPK
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    40
						     
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    41
    exception Load_cplexFile of string
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    42
    exception Load_cplexResult of string
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    43
    exception Save_cplexFile of string
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    44
    exception Execute of string
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    45
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    46
    val load_cplexFile : string -> cplexProg
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    47
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    48
    val save_cplexFile : string -> cplexProg -> unit 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    49
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    50
    val elim_nonfree_bounds : cplexProg -> cplexProg
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    51
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    52
    val relax_strict_ineqs : cplexProg -> cplexProg
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    53
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    54
    val is_normed_cplexProg : cplexProg -> bool
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    55
16966
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
    56
    val get_solver : unit -> cplexSolver
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
    57
    val set_solver : cplexSolver -> unit
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    58
    val solve : cplexProg -> cplexResult
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    59
end;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    60
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    61
structure Cplex  : CPLEX =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    62
struct
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    63
16966
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
    64
datatype cplexSolver = SOLVER_DEFAULT | SOLVER_CPLEX | SOLVER_GLPK
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
    65
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
    66
val cplexsolver = ref SOLVER_DEFAULT;
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
    67
fun get_solver () = !cplexsolver;
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
    68
fun set_solver s = (cplexsolver := s);
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
    69
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    70
exception Load_cplexFile of string;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    71
exception Load_cplexResult of string;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    72
exception Save_cplexFile of string;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    73
	  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    74
datatype cplexTerm = cplexVar of string 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    75
		   | cplexNum of string 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    76
		   | cplexInf 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    77
                   | cplexNeg of cplexTerm 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    78
                   | cplexProd of cplexTerm * cplexTerm 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    79
                   | cplexSum of (cplexTerm list) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    80
datatype cplexComp = cplexLe | cplexLeq | cplexEq | cplexGe | cplexGeq 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    81
datatype cplexGoal = cplexMinimize of cplexTerm | cplexMaximize of cplexTerm
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    82
datatype cplexConstr = cplexConstr of cplexComp * (cplexTerm * cplexTerm)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    83
datatype cplexBounds = cplexBounds of cplexTerm * cplexComp * cplexTerm 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    84
				      * cplexComp * cplexTerm
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    85
                     | cplexBound of cplexTerm * cplexComp * cplexTerm 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    86
datatype cplexProg = cplexProg of string 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    87
				  * cplexGoal 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    88
				  * ((string option * cplexConstr) list) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    89
				  * cplexBounds list
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    90
				  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    91
fun rev_cmp cplexLe = cplexGe
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    92
  | rev_cmp cplexLeq = cplexGeq
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    93
  | rev_cmp cplexGe = cplexLe
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    94
  | rev_cmp cplexGeq = cplexLeq
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    95
  | rev_cmp cplexEq = cplexEq
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    96
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    97
fun the NONE = raise (Load_cplexFile "SOME expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    98
  | the (SOME x) = x; 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
    99
    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   100
fun modulo_signed is_something (cplexNeg u) = is_something u
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   101
  | modulo_signed is_something u = is_something u
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   102
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   103
fun is_Num (cplexNum _) = true
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   104
  | is_Num _ = false
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   105
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   106
fun is_Inf cplexInf = true
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   107
  | is_Inf _ = false
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   108
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   109
fun is_Var (cplexVar _) = true
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   110
  | is_Var _ = false
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   111
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   112
fun is_Neg (cplexNeg x ) = true
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   113
  | is_Neg _ = false
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   114
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   115
fun is_normed_Prod (cplexProd (t1, t2)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   116
    (is_Num t1) andalso (is_Var t2)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   117
  | is_normed_Prod x = is_Var x
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   118
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   119
fun is_normed_Sum (cplexSum ts) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   120
    (ts <> []) andalso forall (modulo_signed is_normed_Prod) ts
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   121
  | is_normed_Sum x = modulo_signed is_normed_Prod x
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   122
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   123
fun is_normed_Constr (cplexConstr (c, (t1, t2))) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   124
    (is_normed_Sum t1) andalso (modulo_signed is_Num t2)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   125
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   126
fun is_Num_or_Inf x = is_Inf x orelse is_Num x
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   127
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   128
fun is_normed_Bounds (cplexBounds (t1, c1, t2, c2, t3)) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   129
    (c1 = cplexLe orelse c1 = cplexLeq) andalso 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   130
    (c2 = cplexLe orelse c2 = cplexLeq) andalso
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   131
    is_Var t2 andalso
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   132
    modulo_signed is_Num_or_Inf t1 andalso
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   133
    modulo_signed is_Num_or_Inf t3
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   134
  | is_normed_Bounds (cplexBound (t1, c, t2)) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   135
    (is_Var t1 andalso (modulo_signed is_Num_or_Inf t2)) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   136
    orelse
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   137
    (c <> cplexEq andalso 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   138
     is_Var t2 andalso (modulo_signed is_Num_or_Inf t1))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   139
	             		
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   140
fun term_of_goal (cplexMinimize x) = x
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   141
  | term_of_goal (cplexMaximize x) = x
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   142
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   143
fun is_normed_cplexProg (cplexProg (name, goal, constraints, bounds)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   144
    is_normed_Sum (term_of_goal goal) andalso
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   145
    forall (fn (_,x) => is_normed_Constr x) constraints andalso
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   146
    forall is_normed_Bounds bounds
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   147
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   148
fun is_NL s = s = "\n"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   149
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   150
fun is_blank s = forall (fn c => c <> #"\n" andalso Char.isSpace c) (String.explode s)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   151
		 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   152
fun is_num a = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   153
    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   154
	val b = String.explode a 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   155
	fun num4 cs = forall Char.isDigit cs		      
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   156
	fun num3 [] = true
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   157
	  | num3 (ds as (c::cs)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   158
	    if c = #"+" orelse c = #"-" then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   159
		num4 cs
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   160
	    else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   161
		num4 ds		    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   162
	fun num2 [] = true
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   163
	  | num2 (c::cs) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   164
	    if c = #"e" orelse c = #"E" then num3 cs
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   165
	    else (Char.isDigit c) andalso num2 cs			 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   166
	fun num1 [] = true
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   167
	  | num1 (c::cs) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   168
	    if c = #"." then num2 cs 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   169
	    else if c = #"e" orelse c = #"E" then num3 cs 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   170
	    else (Char.isDigit c) andalso num1 cs
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   171
	fun num [] = true
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   172
	  | num (c::cs) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   173
	    if c = #"." then num2 cs 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   174
	    else (Char.isDigit c) andalso num1 cs
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   175
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   176
	num b
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   177
    end  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   178
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   179
fun is_delimiter s = s = "+" orelse s = "-" orelse s = ":"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   180
		    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   181
fun is_cmp s = s = "<" orelse s = ">" orelse s = "<=" 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   182
		     orelse s = ">=" orelse s = "="
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   183
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   184
fun is_symbol a = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   185
    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   186
	val symbol_char = String.explode "!\"#$%&()/,.;?@_`'{}|~"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   187
	fun is_symbol_char c = Char.isAlphaNum c orelse 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   188
			       exists (fn d => d=c) symbol_char
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   189
	fun is_symbol_start c = is_symbol_char c andalso 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   190
				not (Char.isDigit c) andalso 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   191
				not (c= #".")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   192
	val b = String.explode a
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   193
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   194
	b <> [] andalso is_symbol_start (hd b) andalso 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   195
	forall is_symbol_char b
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   196
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   197
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   198
fun to_upper s = String.implode (map Char.toUpper (String.explode s))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   199
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   200
fun keyword x = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   201
    let 	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   202
	val a = to_upper x 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   203
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   204
	if a = "BOUNDS" orelse a = "BOUND" then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   205
	    SOME "BOUNDS"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   206
	else if a = "MINIMIZE" orelse a = "MINIMUM" orelse a = "MIN" then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   207
	    SOME "MINIMIZE"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   208
	else if a = "MAXIMIZE" orelse a = "MAXIMUM" orelse a = "MAX" then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   209
	    SOME "MAXIMIZE"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   210
	else if a = "ST" orelse a = "S.T." orelse a = "ST." then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   211
	    SOME "ST"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   212
	else if a = "FREE" orelse a = "END" then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   213
	    SOME a
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   214
	else if a = "GENERAL" orelse a = "GENERALS" orelse a = "GEN" then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   215
	    SOME "GENERAL"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   216
	else if a = "INTEGER" orelse a = "INTEGERS" orelse a = "INT" then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   217
	    SOME "INTEGER"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   218
	else if a = "BINARY" orelse a = "BINARIES" orelse a = "BIN" then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   219
	    SOME "BINARY"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   220
	else if a = "INF" orelse a = "INFINITY" then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   221
	    SOME "INF"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   222
	else								   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   223
	    NONE
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   224
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   225
	        
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   226
val TOKEN_ERROR = ~1
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   227
val TOKEN_BLANK = 0
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   228
val TOKEN_NUM = 1
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   229
val TOKEN_DELIMITER = 2
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   230
val TOKEN_SYMBOL = 3
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   231
val TOKEN_LABEL = 4
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   232
val TOKEN_CMP = 5
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   233
val TOKEN_KEYWORD = 6
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   234
val TOKEN_NL = 7
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   235
		      
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   236
(* tokenize takes a list of chars as argument and returns a list of
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   237
   int * string pairs, each string representing a "cplex token", 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   238
   and each int being one of TOKEN_NUM, TOKEN_DELIMITER, TOKEN_CMP
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   239
   or TOKEN_SYMBOL *)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   240
fun tokenize s = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   241
    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   242
	val flist = [(is_NL, TOKEN_NL), 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   243
		     (is_blank, TOKEN_BLANK), 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   244
		     (is_num, TOKEN_NUM),  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   245
                     (is_delimiter, TOKEN_DELIMITER), 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   246
		     (is_cmp, TOKEN_CMP), 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   247
		     (is_symbol, TOKEN_SYMBOL)]  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   248
	fun match_helper [] s = (fn x => false, TOKEN_ERROR)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   249
	  | match_helper (f::fs) s = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   250
	    if ((fst f) s) then f else match_helper fs s   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   251
	fun match s = match_helper flist s
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   252
	fun tok s = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   253
	    if s = "" then [] else       
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   254
	    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   255
		val h = String.substring (s,0,1) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   256
		val (f, j) = match h
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   257
		fun len i = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   258
		    if size s = i then i 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   259
		    else if f (String.substring (s,0,i+1)) then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   260
			len (i+1) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   261
		    else i 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   262
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   263
		if j < 0 then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   264
		    (if h = "\\" then [] 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   265
		     else raise (Load_cplexFile ("token expected, found: "
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   266
						 ^s)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   267
		else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   268
		    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   269
			val l = len 1 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   270
			val u = String.substring (s,0,l)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   271
			val v = String.extract (s,l,NONE)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   272
		    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   273
			if j = 0 then tok v else (j, u) :: tok v
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   274
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   275
	    end                
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   276
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   277
	tok s     
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   278
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   279
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   280
exception Tokenize of string;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   281
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   282
fun tokenize_general flist s = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   283
    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   284
	fun match_helper [] s = raise (Tokenize s)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   285
	  | match_helper (f::fs) s = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   286
	    if ((fst f) s) then f else match_helper fs s   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   287
	fun match s = match_helper flist s
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   288
	fun tok s = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   289
	    if s = "" then [] else       
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   290
	    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   291
		val h = String.substring (s,0,1) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   292
		val (f, j) = match h
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   293
		fun len i = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   294
		    if size s = i then i 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   295
		    else if f (String.substring (s,0,i+1)) then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   296
			len (i+1) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   297
		    else i 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   298
		val l = len 1
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   299
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   300
		(j, String.substring (s,0,l)) :: tok (String.extract (s,l,NONE))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   301
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   302
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   303
	tok s     
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   304
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   305
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   306
fun load_cplexFile name = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   307
    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   308
	val f = TextIO.openIn name  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   309
        val ignore_NL = ref true  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   310
	val rest = ref []
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   311
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   312
	fun is_symbol s c = (fst c) = TOKEN_SYMBOL andalso (to_upper (snd c)) = s 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   313
		   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   314
	fun readToken_helper () =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   315
	    if length (!rest) > 0 then  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   316
		let val u = hd (!rest) in 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   317
		    (
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   318
		     rest := tl (!rest); 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   319
		     SOME u
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   320
		    ) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   321
		end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   322
	    else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   323
		let val s = TextIO.inputLine f in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   324
		    if s = "" then NONE else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   325
		    let val t = tokenize s in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   326
			if (length t >= 2 andalso 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   327
			    snd(hd (tl t)) = ":") 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   328
			then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   329
			    rest := (TOKEN_LABEL, snd (hd t)) :: (tl (tl t))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   330
			else if (length t >= 2) andalso is_symbol "SUBJECT" (hd (t)) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   331
				andalso is_symbol "TO" (hd (tl t)) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   332
			then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   333
			    rest := (TOKEN_SYMBOL, "ST") :: (tl (tl t))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   334
			else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   335
			    rest := t;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   336
			readToken_helper ()  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   337
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   338
		end           
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   339
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   340
	fun readToken_helper2 () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   341
		let val c = readToken_helper () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   342
		    if c = NONE then NONE
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   343
                    else if !ignore_NL andalso fst (the c) = TOKEN_NL then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   344
			readToken_helper2 ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   345
		    else if fst (the c) = TOKEN_SYMBOL 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   346
			    andalso keyword (snd (the c)) <> NONE
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   347
		    then SOME (TOKEN_KEYWORD, the (keyword (snd (the c))))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   348
		    else c
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   349
		end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   350
		
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   351
	fun readToken () = readToken_helper2 ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   352
	    			   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   353
	fun pushToken a = rest := (a::(!rest))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   354
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   355
	fun is_value token = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   356
	    fst token = TOKEN_NUM orelse (fst token = TOKEN_KEYWORD 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   357
					  andalso snd token = "INF")	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   358
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   359
        fun get_value token = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   360
	    if fst token = TOKEN_NUM then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   361
		cplexNum (snd token)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   362
	    else if fst token = TOKEN_KEYWORD andalso snd token = "INF" 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   363
	    then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   364
		cplexInf
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   365
	    else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   366
		raise (Load_cplexFile "num expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   367
					    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   368
	fun readTerm_Product only_num = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   369
	    let val c = readToken () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   370
		if c = NONE then NONE
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   371
		else if fst (the c) = TOKEN_SYMBOL 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   372
		then (
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   373
		    if only_num then (pushToken (the c); NONE) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   374
		    else SOME (cplexVar (snd (the c)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   375
		    ) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   376
		else if only_num andalso is_value (the c) then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   377
		    SOME (get_value (the c))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   378
		else if is_value (the c) then  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   379
		    let val t1 = get_value (the c) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   380
			val d = readToken () 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   381
		    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   382
			if d = NONE then SOME t1 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   383
			else if fst (the d) = TOKEN_SYMBOL then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   384
			    SOME (cplexProd (t1, cplexVar (snd (the d))))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   385
			else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   386
			    (pushToken (the d); SOME t1)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   387
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   388
		else (pushToken (the c); NONE)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   389
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   390
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   391
	fun readTerm_Signed only_signed only_num = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   392
	    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   393
		val c = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   394
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   395
		if c = NONE then NONE 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   396
		else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   397
		    let val d = the c in 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   398
			if d = (TOKEN_DELIMITER, "+") then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   399
			    readTerm_Product only_num
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   400
			 else if d = (TOKEN_DELIMITER, "-") then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   401
			     SOME (cplexNeg (the (readTerm_Product 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   402
						      only_num)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   403
			 else (pushToken d; 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   404
			       if only_signed then NONE 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   405
			       else readTerm_Product only_num)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   406
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   407
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   408
    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   409
	fun readTerm_Sum first_signed = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   410
	    let val c = readTerm_Signed first_signed false in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   411
		if c = NONE then [] else (the c)::(readTerm_Sum true)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   412
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   413
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   414
	fun readTerm () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   415
	    let val c = readTerm_Sum false in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   416
		if c = [] then NONE 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   417
		else if tl c = [] then SOME (hd c)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   418
		else SOME (cplexSum c) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   419
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   420
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   421
	fun readLabeledTerm () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   422
	    let val c = readToken () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   423
		if c = NONE then (NONE, NONE) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   424
		else if fst (the c) = TOKEN_LABEL then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   425
		    let val t = readTerm () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   426
			if t = NONE then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   427
			    raise (Load_cplexFile ("term after label "^
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   428
						   (snd (the c))^
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   429
						   " expected"))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   430
			else (SOME (snd (the c)), t)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   431
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   432
		else (pushToken (the c); (NONE, readTerm ()))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   433
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   434
            
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   435
	fun readGoal () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   436
	    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   437
		val g = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   438
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   439
    		if g = SOME (TOKEN_KEYWORD, "MAXIMIZE") then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   440
		    cplexMaximize (the (snd (readLabeledTerm ()))) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   441
		else if g = SOME (TOKEN_KEYWORD, "MINIMIZE") then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   442
		    cplexMinimize (the (snd (readLabeledTerm ())))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   443
		else raise (Load_cplexFile "MAXIMIZE or MINIMIZE expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   444
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   445
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   446
	fun str2cmp b = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   447
	    (case b of 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   448
		 "<" => cplexLe 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   449
	       | "<=" => cplexLeq 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   450
	       | ">" => cplexGe 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   451
	       | ">=" => cplexGeq 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   452
               | "=" => cplexEq
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   453
	       | _ => raise (Load_cplexFile (b^" is no TOKEN_CMP")))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   454
			    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   455
	fun readConstraint () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   456
            let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   457
		val t = readLabeledTerm () 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   458
		fun make_constraint b t1 t2 =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   459
                    cplexConstr 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   460
			(str2cmp b,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   461
			 (t1, t2))		
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   462
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   463
		if snd t = NONE then NONE
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   464
		else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   465
		    let val c = readToken () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   466
			if c = NONE orelse fst (the c) <> TOKEN_CMP 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   467
			then raise (Load_cplexFile "TOKEN_CMP expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   468
			else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   469
			    let val n = readTerm_Signed false true in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   470
				if n = NONE then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   471
				    raise (Load_cplexFile "num expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   472
				else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   473
				    SOME (fst t,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   474
					  make_constraint (snd (the c)) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   475
							  (the (snd t)) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   476
							  (the n))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   477
			    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   478
		    end            
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   479
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   480
        
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   481
        fun readST () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   482
	    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   483
		fun readbody () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   484
		    let val t = readConstraint () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   485
			if t = NONE then []
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   486
			else if (is_normed_Constr (snd (the t))) then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   487
			    (the t)::(readbody ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   488
			else if (fst (the t) <> NONE) then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   489
			    raise (Load_cplexFile 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   490
				       ("constraint '"^(the (fst (the t)))^ 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   491
					"'is not normed"))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   492
			else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   493
			    raise (Load_cplexFile 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   494
				       "constraint is not normed")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   495
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   496
	    in    		
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   497
		if readToken () = SOME (TOKEN_KEYWORD, "ST") 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   498
		then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   499
		    readbody ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   500
		else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   501
		    raise (Load_cplexFile "ST expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   502
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   503
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   504
	fun readCmp () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   505
	    let val c = readToken () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   506
		if c = NONE then NONE
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   507
		else if fst (the c) = TOKEN_CMP then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   508
		    SOME (str2cmp (snd (the c)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   509
		else (pushToken (the c); NONE)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   510
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   511
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   512
	fun skip_NL () =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   513
	    let val c = readToken () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   514
		if c <> NONE andalso fst (the c) = TOKEN_NL then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   515
		    skip_NL ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   516
		else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   517
		    (pushToken (the c); ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   518
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   519
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   520
	fun is_var (cplexVar _) = true
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   521
	  | is_var _ = false
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   522
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   523
	fun make_bounds c t1 t2 = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   524
	    cplexBound (t1, c, t2)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   525
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   526
	fun readBound () = 	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   527
	    let 		
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   528
		val _ = skip_NL ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   529
		val t1 = readTerm ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   530
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   531
		if t1 = NONE then NONE 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   532
		else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   533
		    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   534
			val c1 = readCmp () 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   535
		    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   536
			if c1 = NONE then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   537
			    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   538
				val c = readToken () 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   539
			    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   540
				if c = SOME (TOKEN_KEYWORD, "FREE") then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   541
				    SOME (
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   542
				    cplexBounds (cplexNeg cplexInf,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   543
						 cplexLeq,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   544
						 the t1,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   545
						 cplexLeq,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   546
						 cplexInf))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   547
				else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   548
				    raise (Load_cplexFile "FREE expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   549
			    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   550
			else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   551
			    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   552
				val t2 = readTerm () 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   553
			    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   554
				if t2 = NONE then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   555
				    raise (Load_cplexFile "term expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   556
				else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   557
				    let val c2 = readCmp () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   558
					if c2 = NONE then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   559
					    SOME (make_bounds (the c1) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   560
							      (the t1)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   561
							      (the t2))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   562
					else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   563
					    SOME (
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   564
					    cplexBounds (the t1,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   565
							 the c1,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   566
							 the t2,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   567
							 the c2,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   568
							 the (readTerm())))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   569
				    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   570
			    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   571
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   572
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   573
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   574
	fun readBounds () =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   575
	    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   576
		fun makestring b = "?"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   577
		fun readbody () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   578
		    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   579
			val b = readBound ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   580
		    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   581
			if b = NONE then []
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   582
			else if (is_normed_Bounds (the b)) then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   583
			    (the b)::(readbody())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   584
			else (
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   585
			    raise (Load_cplexFile 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   586
				       ("bounds are not normed in: "^
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   587
					(makestring (the b)))))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   588
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   589
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   590
		if readToken () = SOME (TOKEN_KEYWORD, "BOUNDS") then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   591
		    readbody ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   592
		else raise (Load_cplexFile "BOUNDS expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   593
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   594
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   595
        fun readEnd () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   596
	    if readToken () = SOME (TOKEN_KEYWORD, "END") then ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   597
	    else raise (Load_cplexFile "END expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   598
		    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   599
	val result_Goal = readGoal ()  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   600
	val result_ST = readST ()  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   601
	val _ =	ignore_NL := false 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   602
        val result_Bounds = readBounds ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   603
        val _ = ignore_NL := true
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   604
        val _ = readEnd ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   605
	val _ = TextIO.closeIn f
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   606
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   607
	cplexProg (name, result_Goal, result_ST, result_Bounds)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   608
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   609
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   610
fun save_cplexFile filename (cplexProg (name, goal, constraints, bounds)) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   611
    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   612
	val f = TextIO.openOut filename
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   613
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   614
	fun basic_write s = TextIO.output(f, s)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   615
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   616
	val linebuf = ref ""
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   617
	fun buf_flushline s = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   618
	    (basic_write (!linebuf); 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   619
	     basic_write "\n";
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   620
	     linebuf := s)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   621
	fun buf_add s = linebuf := (!linebuf) ^ s
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   622
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   623
	fun write s = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   624
	    if (String.size s) + (String.size (!linebuf)) >= 250 then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   625
		buf_flushline ("    "^s)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   626
	    else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   627
		buf_add s
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   628
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   629
        fun writeln s = (buf_add s; buf_flushline "")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   630
        
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   631
	fun write_term (cplexVar x) = write x
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   632
	  | write_term (cplexNum x) = write x
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   633
	  | write_term cplexInf = write "inf"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   634
	  | write_term (cplexProd (cplexNum "1", b)) = write_term b
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   635
	  | write_term (cplexProd (a, b)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   636
	    (write_term a; write " "; write_term b)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   637
          | write_term (cplexNeg x) = (write " - "; write_term x)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   638
          | write_term (cplexSum ts) = write_terms ts
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   639
	and write_terms [] = ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   640
	  | write_terms (t::ts) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   641
	    (if (not (is_Neg t)) then write " + " else (); 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   642
	     write_term t; write_terms ts)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   643
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   644
	fun write_goal (cplexMaximize term) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   645
	    (writeln "MAXIMIZE"; write_term term; writeln "")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   646
	  | write_goal (cplexMinimize term) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   647
	    (writeln "MINIMIZE"; write_term term; writeln "")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   648
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   649
	fun write_cmp cplexLe = write "<"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   650
	  | write_cmp cplexLeq = write "<="
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   651
	  | write_cmp cplexEq = write "="
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   652
	  | write_cmp cplexGe = write ">"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   653
	  | write_cmp cplexGeq = write ">="
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   654
        
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   655
	fun write_constr (cplexConstr (cmp, (a,b))) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   656
	    (write_term a;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   657
	     write " ";
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   658
	     write_cmp cmp;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   659
	     write " ";
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   660
	     write_term b)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   661
        
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   662
	fun write_constraints [] = ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   663
	  | write_constraints (c::cs) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   664
	    (if (fst c <> NONE) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   665
	     then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   666
		 (write (the (fst c)); write ": ") 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   667
	     else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   668
		 ();
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   669
	     write_constr (snd c);
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   670
	     writeln "";
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   671
	     write_constraints cs)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   672
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   673
	fun write_bounds [] = ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   674
	  | write_bounds ((cplexBounds (t1,c1,t2,c2,t3))::bs) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   675
	    ((if t1 = cplexNeg cplexInf andalso t3 = cplexInf
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   676
		 andalso (c1 = cplexLeq orelse c1 = cplexLe) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   677
		 andalso (c2 = cplexLeq orelse c2 = cplexLe) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   678
	      then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   679
		  (write_term t2; write " free")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   680
	      else		 		 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   681
		  (write_term t1; write " "; write_cmp c1; write " ";
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   682
		   write_term t2; write " "; write_cmp c2; write " ";
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   683
		   write_term t3)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   684
	     ); writeln ""; write_bounds bs)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   685
	  | write_bounds ((cplexBound (t1, c, t2)) :: bs) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   686
	    (write_term t1; write " "; 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   687
	     write_cmp c; write " ";
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   688
	     write_term t2; writeln ""; write_bounds bs)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   689
	     
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   690
	val _ = write_goal goal
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   691
        val _ = (writeln ""; writeln "ST")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   692
	val _ = write_constraints constraints
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   693
        val _ = (writeln ""; writeln "BOUNDS")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   694
	val _ = write_bounds bounds
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   695
        val _ = (writeln ""; writeln "END") 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   696
        val _ = TextIO.closeOut f
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   697
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   698
	()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   699
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   700
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   701
fun norm_Constr (constr as cplexConstr (c, (t1, t2))) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   702
    if not (modulo_signed is_Num t2) andalso  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   703
       modulo_signed is_Num t1 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   704
    then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   705
	[cplexConstr (rev_cmp c, (t2, t1))]
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   706
    else if (c = cplexLe orelse c = cplexLeq) andalso
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   707
	    (t1 = (cplexNeg cplexInf) orelse t2 = cplexInf)	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   708
    then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   709
	[]
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   710
    else if (c = cplexGe orelse c = cplexGeq) andalso
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   711
	    (t1 = cplexInf orelse t2 = cplexNeg cplexInf)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   712
    then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   713
	[]
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   714
    else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   715
	[constr]
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   716
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   717
fun bound2constr (cplexBounds (t1,c1,t2,c2,t3)) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   718
    (norm_Constr(cplexConstr (c1, (t1, t2))))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   719
    @ (norm_Constr(cplexConstr (c2, (t2, t3))))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   720
  | bound2constr (cplexBound (t1, cplexEq, t2)) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   721
    (norm_Constr(cplexConstr (cplexLeq, (t1, t2))))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   722
    @ (norm_Constr(cplexConstr (cplexLeq, (t2, t1))))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   723
  | bound2constr (cplexBound (t1, c1, t2)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   724
    norm_Constr(cplexConstr (c1, (t1,t2)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   725
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   726
val emptyset = Symtab.empty
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   727
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 16966
diff changeset
   728
fun singleton v = Symtab.update (v, ()) emptyset
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   729
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   730
fun merge a b = Symtab.merge (op =) (a, b)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   731
21056
2cfe839e8d58 Symtab.foldl replaced by Symtab.fold
haftmann
parents: 17521
diff changeset
   732
fun mergemap f ts = fold (fn x => fn table => merge table (f x)) ts Symtab.empty
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   733
21056
2cfe839e8d58 Symtab.foldl replaced by Symtab.fold
haftmann
parents: 17521
diff changeset
   734
fun diff a b = Symtab.fold (fn (k, v) => fn a => 
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   735
			 (Symtab.delete k a) handle UNDEF => a) 
21056
2cfe839e8d58 Symtab.foldl replaced by Symtab.fold
haftmann
parents: 17521
diff changeset
   736
		     b a
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   737
		    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   738
fun collect_vars (cplexVar v) = singleton v
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   739
  | collect_vars (cplexNeg t) = collect_vars t
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   740
  | collect_vars (cplexProd (t1, t2)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   741
    merge (collect_vars t1) (collect_vars t2)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   742
  | collect_vars (cplexSum ts) = mergemap collect_vars ts
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   743
  | collect_vars _ = emptyset
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   744
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   745
(* Eliminates all nonfree bounds from the linear program and produces an
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   746
   equivalent program with only free bounds 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   747
   IF for the input program P holds: is_normed_cplexProg P *)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   748
fun elim_nonfree_bounds (cplexProg (name, goal, constraints, bounds)) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   749
    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   750
	fun collect_constr_vars (_, cplexConstr (c, (t1,_))) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   751
	    (collect_vars t1)   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   752
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   753
	val cvars = merge (collect_vars (term_of_goal goal)) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   754
			  (mergemap collect_constr_vars constraints)         
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   755
		    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   756
	fun collect_lower_bounded_vars 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   757
	    (cplexBounds (t1, c1, cplexVar v, c2, t3)) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   758
	    singleton v
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   759
	  |  collect_lower_bounded_vars 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   760
		 (cplexBound (_, cplexLe, cplexVar v)) =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   761
	     singleton v
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   762
	  |  collect_lower_bounded_vars 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   763
		 (cplexBound (_, cplexLeq, cplexVar v)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   764
	     singleton v
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   765
	  |  collect_lower_bounded_vars 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   766
		 (cplexBound (cplexVar v, cplexGe,_)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   767
	     singleton v
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   768
	  |  collect_lower_bounded_vars 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   769
		 (cplexBound (cplexVar v, cplexGeq, _)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   770
	     singleton v
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   771
	  | collect_lower_bounded_vars
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   772
	    (cplexBound (cplexVar v, cplexEq, _)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   773
	    singleton v
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   774
	  |  collect_lower_bounded_vars _ = emptyset
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   775
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   776
	val lvars = mergemap collect_lower_bounded_vars bounds        
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   777
	val positive_vars = diff cvars lvars			   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   778
	val zero = cplexNum "0"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   779
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   780
	fun make_pos_constr v = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   781
	    (NONE, cplexConstr (cplexGeq, ((cplexVar v), zero)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   782
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   783
	fun make_free_bound v = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   784
	    cplexBounds (cplexNeg cplexInf, cplexLeq, 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   785
			 cplexVar v, cplexLeq,
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   786
			 cplexInf)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   787
	
21056
2cfe839e8d58 Symtab.foldl replaced by Symtab.fold
haftmann
parents: 17521
diff changeset
   788
	val pos_constrs = rev (Symtab.fold
2cfe839e8d58 Symtab.foldl replaced by Symtab.fold
haftmann
parents: 17521
diff changeset
   789
			      (fn (k, v) => cons (make_pos_constr k))
2cfe839e8d58 Symtab.foldl replaced by Symtab.fold
haftmann
parents: 17521
diff changeset
   790
			      positive_vars [])
2cfe839e8d58 Symtab.foldl replaced by Symtab.fold
haftmann
parents: 17521
diff changeset
   791
        val bound_constrs = map (pair NONE)
2cfe839e8d58 Symtab.foldl replaced by Symtab.fold
haftmann
parents: 17521
diff changeset
   792
				(maps bound2constr bounds)
2cfe839e8d58 Symtab.foldl replaced by Symtab.fold
haftmann
parents: 17521
diff changeset
   793
	val constraints' = constraints @ pos_constrs @ bound_constrs
2cfe839e8d58 Symtab.foldl replaced by Symtab.fold
haftmann
parents: 17521
diff changeset
   794
	val bounds' = rev (Symtab.fold (fn (v, _) => cons (make_free_bound v)) cvars []);
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   795
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   796
	cplexProg (name, goal, constraints', bounds')
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   797
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   798
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   799
fun relax_strict_ineqs (cplexProg (name, goals, constrs, bounds)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   800
    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   801
	fun relax cplexLe = cplexLeq
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   802
	  | relax cplexGe = cplexGeq
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   803
	  | relax x = x
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   804
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   805
	fun relax_constr (n, cplexConstr(c, (t1, t2))) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   806
	    (n, cplexConstr(relax c, (t1, t2)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   807
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   808
	fun relax_bounds (cplexBounds (t1, c1, t2, c2, t3)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   809
	    cplexBounds (t1, relax c1, t2, relax c2, t3)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   810
	  | relax_bounds (cplexBound (t1, c, t2)) = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   811
	    cplexBound (t1, relax c, t2)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   812
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   813
	cplexProg (name, 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   814
		   goals, 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   815
		   map relax_constr constrs, 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   816
		   map relax_bounds bounds)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   817
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   818
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   819
datatype cplexResult = Unbounded 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   820
		     | Infeasible 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   821
		     | Undefined
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   822
		     | Optimal of string * ((string * string) list)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   823
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   824
fun is_separator x = forall (fn c => c = #"-") (String.explode x)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   825
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   826
fun is_sign x = (x = "+" orelse x = "-")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   827
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   828
fun is_colon x = (x = ":")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   829
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   830
fun is_resultsymbol a = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   831
    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   832
	val symbol_char = String.explode "!\"#$%&()/,.;?@_`'{}|~-"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   833
	fun is_symbol_char c = Char.isAlphaNum c orelse 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   834
			       exists (fn d => d=c) symbol_char
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   835
	fun is_symbol_start c = is_symbol_char c andalso 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   836
				not (Char.isDigit c) andalso 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   837
				not (c= #".") andalso
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   838
				not (c= #"-")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   839
	val b = String.explode a
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   840
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   841
	b <> [] andalso is_symbol_start (hd b) andalso 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   842
	forall is_symbol_char b
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   843
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   844
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   845
val TOKEN_SIGN = 100
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   846
val TOKEN_COLON = 101
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   847
val TOKEN_SEPARATOR = 102
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   848
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   849
fun load_glpkResult name = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   850
    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   851
	val flist = [(is_NL, TOKEN_NL), 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   852
		     (is_blank, TOKEN_BLANK),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   853
		     (is_num, TOKEN_NUM),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   854
		     (is_sign, TOKEN_SIGN),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   855
                     (is_colon, TOKEN_COLON),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   856
		     (is_cmp, TOKEN_CMP),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   857
		     (is_resultsymbol, TOKEN_SYMBOL),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   858
		     (is_separator, TOKEN_SEPARATOR)]
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   859
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   860
	val tokenize = tokenize_general flist
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   861
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   862
	val f = TextIO.openIn name  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   863
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   864
	val rest = ref []
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   865
		   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   866
	fun readToken_helper () =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   867
	    if length (!rest) > 0 then  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   868
		let val u = hd (!rest) in 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   869
		    (
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   870
		     rest := tl (!rest); 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   871
		     SOME u
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   872
		    ) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   873
		end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   874
	    else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   875
		let val s = TextIO.inputLine f in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   876
		    if s = "" then NONE else (rest := tokenize s; readToken_helper())					     
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   877
		end           
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   878
		
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   879
	fun is_tt tok ty = (tok <> NONE andalso (fst (the tok)) = ty)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   880
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   881
	fun pushToken a = if a = NONE then () else (rest := ((the a)::(!rest)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   882
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   883
	fun readToken () =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   884
	    let val t = readToken_helper () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   885
		if is_tt t TOKEN_BLANK then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   886
		    readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   887
		else if is_tt t TOKEN_NL then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   888
		    let val t2 = readToken_helper () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   889
			if is_tt t2 TOKEN_SIGN then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   890
			    (pushToken (SOME (TOKEN_SEPARATOR, "-")); t)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   891
			else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   892
			    (pushToken t2; t)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   893
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   894
		else if is_tt t TOKEN_SIGN then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   895
		    let val t2 = readToken_helper () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   896
			if is_tt t2 TOKEN_NUM then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   897
			    (SOME (TOKEN_NUM, (snd (the t))^(snd (the t2))))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   898
			else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   899
			    (pushToken t2; t)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   900
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   901
		else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   902
		    t
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   903
	    end	    		
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   904
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   905
        fun readRestOfLine P = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   906
	    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   907
		val t = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   908
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   909
		if is_tt t TOKEN_NL orelse t = NONE 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   910
		then P
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   911
		else readRestOfLine P
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   912
	    end			
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   913
							 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   914
	fun readHeader () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   915
	    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   916
		fun readStatus () = readRestOfLine ("STATUS", snd (the (readToken ())))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   917
		fun readObjective () = readRestOfLine ("OBJECTIVE", snd (the (readToken (); readToken (); readToken ())))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   918
		val t1 = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   919
		val t2 = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   920
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   921
		if is_tt t1 TOKEN_SYMBOL andalso is_tt t2 TOKEN_COLON 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   922
		then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   923
		    case to_upper (snd (the t1)) of
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   924
			"STATUS" => (readStatus ())::(readHeader ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   925
		      | "OBJECTIVE" => (readObjective())::(readHeader ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   926
		      | _ => (readRestOfLine (); readHeader ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   927
		else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   928
		    (pushToken t2; pushToken t1; [])
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   929
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   930
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   931
	fun skip_until_sep () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   932
	    let val x = readToken () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   933
		if is_tt x TOKEN_SEPARATOR then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   934
		    readRestOfLine ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   935
		else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   936
		    skip_until_sep ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   937
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   938
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   939
	fun load_value () =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   940
	    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   941
		val t1 = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   942
		val t2 = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   943
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   944
		if is_tt t1 TOKEN_NUM andalso is_tt t2 TOKEN_SYMBOL then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   945
		    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   946
			val t = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   947
			val state = if is_tt t TOKEN_NL then readToken () else t
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   948
			val _ = if is_tt state TOKEN_SYMBOL then () else raise (Load_cplexResult "state expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   949
			val k = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   950
		    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   951
			if is_tt k TOKEN_NUM then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   952
			    readRestOfLine (SOME (snd (the t2), snd (the k)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   953
			else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   954
			    raise (Load_cplexResult "number expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   955
		    end				
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   956
		else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   957
		    (pushToken t2; pushToken t1; NONE)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   958
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   959
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   960
	fun load_values () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   961
	    let val v = load_value () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   962
		if v = NONE then [] else (the v)::(load_values ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   963
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   964
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   965
	val header = readHeader () 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   966
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   967
	val result = 
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   968
	    case AList.lookup (op =) header "STATUS" of
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   969
		SOME "INFEASIBLE" => Infeasible
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   970
	      | SOME "UNBOUNDED" => Unbounded
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   971
	      | SOME "OPTIMAL" => Optimal (the (AList.lookup (op =) header "OBJECTIVE"), 
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   972
					   (skip_until_sep (); 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   973
					    skip_until_sep ();
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   974
					    load_values ()))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   975
	      | _ => Undefined
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   976
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   977
	val _ = TextIO.closeIn f
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   978
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   979
	result
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   980
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   981
    handle (Tokenize s) => raise (Load_cplexResult ("Tokenize: "^s))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   982
	 | Option => raise (Load_cplexResult "Option")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   983
	 | x => raise x
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   984
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   985
fun load_cplexResult name = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   986
    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   987
	val flist = [(is_NL, TOKEN_NL), 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   988
		     (is_blank, TOKEN_BLANK),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   989
		     (is_num, TOKEN_NUM),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   990
		     (is_sign, TOKEN_SIGN),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   991
                     (is_colon, TOKEN_COLON),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   992
		     (is_cmp, TOKEN_CMP),
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   993
		     (is_resultsymbol, TOKEN_SYMBOL)]
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   994
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   995
	val tokenize = tokenize_general flist
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   996
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   997
	val f = TextIO.openIn name  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   998
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
   999
	val rest = ref []
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1000
		   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1001
	fun readToken_helper () =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1002
	    if length (!rest) > 0 then  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1003
		let val u = hd (!rest) in 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1004
		    (
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1005
		     rest := tl (!rest); 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1006
		     SOME u
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1007
		    ) 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1008
		end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1009
	    else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1010
		let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1011
		    val s = TextIO.inputLine f 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1012
		in		    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1013
		    if s = "" then NONE else (rest := tokenize s; readToken_helper())					     
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1014
		end           
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1015
		
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1016
	fun is_tt tok ty = (tok <> NONE andalso (fst (the tok)) = ty)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1017
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1018
	fun pushToken a = if a = NONE then () else (rest := ((the a)::(!rest)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1019
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1020
	fun readToken () =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1021
	    let val t = readToken_helper () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1022
		if is_tt t TOKEN_BLANK then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1023
		    readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1024
		else if is_tt t TOKEN_SIGN then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1025
		    let val t2 = readToken_helper () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1026
			if is_tt t2 TOKEN_NUM then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1027
			    (SOME (TOKEN_NUM, (snd (the t))^(snd (the t2))))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1028
			else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1029
			    (pushToken t2; t)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1030
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1031
		else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1032
		    t
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1033
	    end	    		
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1034
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1035
        fun readRestOfLine P = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1036
	    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1037
		val t = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1038
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1039
		if is_tt t TOKEN_NL orelse t = NONE 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1040
		then P
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1041
		else readRestOfLine P
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1042
	    end			
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1043
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1044
	fun readHeader () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1045
	    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1046
		fun readStatus () = readRestOfLine ("STATUS", snd (the (readToken ())))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1047
		fun readObjective () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1048
		    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1049
			val t = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1050
		    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1051
			if is_tt t TOKEN_SYMBOL andalso to_upper (snd (the t)) = "VALUE" then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1052
			    readRestOfLine ("OBJECTIVE", snd (the (readToken())))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1053
			else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1054
			    readRestOfLine ("OBJECTIVE_NAME", snd (the t))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1055
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1056
		  
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1057
		val t = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1058
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1059
		if is_tt t TOKEN_SYMBOL then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1060
		    case to_upper (snd (the t)) of 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1061
			"STATUS" => (readStatus ())::(readHeader ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1062
		      | "OBJECTIVE" => (readObjective ())::(readHeader ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1063
		      | "SECTION" => (pushToken t; [])
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1064
		      | _ => (readRestOfLine (); readHeader ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1065
		else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1066
		    (readRestOfLine (); readHeader ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1067
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1068
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1069
	fun skip_nls () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1070
	    let val x = readToken () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1071
		if is_tt x TOKEN_NL then 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1072
		    skip_nls () 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1073
		else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1074
		    (pushToken x; ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1075
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1076
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1077
	fun skip_paragraph () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1078
	    if is_tt (readToken ()) TOKEN_NL then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1079
		(if is_tt (readToken ()) TOKEN_NL then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1080
		     skip_nls ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1081
		 else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1082
		     skip_paragraph ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1083
	    else			
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1084
		skip_paragraph ()	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1085
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1086
	fun load_value () =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1087
	    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1088
		val t1 = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1089
		val t1 = if is_tt t1 TOKEN_SYMBOL andalso snd (the t1) = "A" then readToken () else t1
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1090
	    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1091
		if is_tt t1 TOKEN_NUM then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1092
		    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1093
			val name = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1094
			val status = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1095
			val value = readToken ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1096
		    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1097
			if is_tt name TOKEN_SYMBOL andalso
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1098
			   is_tt status TOKEN_SYMBOL andalso
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1099
			   is_tt value TOKEN_NUM 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1100
			then
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1101
			    readRestOfLine (SOME (snd (the name), snd (the value)))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1102
			else
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1103
			    raise (Load_cplexResult "column line expected")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1104
		    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1105
		else 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1106
		    (pushToken t1; NONE)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1107
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1108
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1109
	fun load_values () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1110
	    let val v = load_value () in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1111
		if v = NONE then [] else (the v)::(load_values ())
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1112
	    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1113
	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1114
	val header = readHeader ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1115
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1116
	val result = 
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
  1117
	    case AList.lookup (op =) header "STATUS" of
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1118
		SOME "INFEASIBLE" => Infeasible
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1119
	      | SOME "NONOPTIMAL" => Unbounded
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
  1120
	      | SOME "OPTIMAL" => Optimal (the (AList.lookup (op =) header "OBJECTIVE"), 
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1121
					   (skip_paragraph (); 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1122
					    skip_paragraph (); 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1123
					    skip_paragraph (); 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1124
					    skip_paragraph (); 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1125
					    skip_paragraph ();
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1126
					    load_values ()))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1127
	      | _ => Undefined
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1128
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1129
	val _ = TextIO.closeIn f
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1130
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1131
	result
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1132
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1133
    handle (Tokenize s) => raise (Load_cplexResult ("Tokenize: "^s))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1134
	 | Option => raise (Load_cplexResult "Option")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1135
	 | x => raise x
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1136
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1137
exception Execute of string;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1138
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1139
fun tmp_file s = Path.pack (Path.expand (File.tmp_path (Path.make [s])));
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1140
fun wrap s = "\""^s^"\"";
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1141
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1142
fun solve_glpk prog =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1143
    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1144
	val name = LargeInt.toString (Time.toMicroseconds (Time.now ()))	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1145
	val lpname = tmp_file (name^".lp")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1146
	val resultname = tmp_file (name^".txt")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1147
	val _ = save_cplexFile lpname prog		
16966
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
  1148
	val cplex_path = getenv "GLPK_PATH"
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1149
	val cplex = if cplex_path = "" then "glpsol" else cplex_path
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1150
	val command = (wrap cplex)^" --lpt "^(wrap lpname)^" --output "^(wrap resultname)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1151
	val answer = execute command
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1152
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1153
	let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1154
	    val result = load_glpkResult resultname
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1155
	    val _ = OS.FileSys.remove lpname
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1156
	    val _ = OS.FileSys.remove resultname
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1157
	in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1158
	    result
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1159
	end 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1160
	handle (Load_cplexResult s) => raise (Execute ("Load_cplexResult: "^s^"\nExecute: "^answer))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1161
	     | _ => raise (Execute answer)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1162
    end		     
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1163
    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1164
fun solve_cplex prog =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1165
    let 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1166
	fun write_script s lp r =
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1167
	    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1168
		val f = TextIO.openOut s
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1169
		val _ = TextIO.output (f, "read\n"^lp^"\noptimize\nwrite\n"^r^"\nquit")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1170
		val _ = TextIO.closeOut f
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1171
	    in 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1172
		() 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1173
	    end 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1174
	    
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1175
	val name = LargeInt.toString (Time.toMicroseconds (Time.now ()))	
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1176
	val lpname = tmp_file (name^".lp")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1177
	val resultname = tmp_file (name^".txt")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1178
	val scriptname = tmp_file (name^".script")
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1179
	val _ = save_cplexFile lpname prog		
16966
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
  1180
	val cplex_path = getenv "CPLEX_PATH"
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1181
	val cplex = if cplex_path = "" then "cplex" else cplex_path
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1182
	val _ = write_script scriptname lpname resultname
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1183
	val command = (wrap cplex)^" < "^(wrap scriptname)^" > /dev/null"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1184
	val answer = "return code "^(Int.toString (system command))
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1185
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1186
	let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1187
	    val result = load_cplexResult resultname
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1188
	    val _ = OS.FileSys.remove lpname
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1189
	    val _ = OS.FileSys.remove resultname
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1190
	    val _ = OS.FileSys.remove scriptname
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1191
	in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1192
	    result
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1193
	end 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1194
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1195
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1196
fun solve prog = 
16966
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
  1197
    case get_solver () of
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
  1198
      SOLVER_DEFAULT => 
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
  1199
        (case getenv "LP_SOLVER" of
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
  1200
	   "CPLEX" => solve_cplex prog
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
  1201
         | "GLPK" => solve_glpk prog
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
  1202
         | _ => raise (Execute ("LP_SOLVER must be set to CPLEX or to GLPK")))
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
  1203
    | SOLVER_CPLEX => solve_cplex prog
37e34f315057 1. changed configuration variables for linear programming (Cplex_tools):
obua
parents: 16873
diff changeset
  1204
    | SOLVER_GLPK => solve_glpk prog
16784
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1205
		   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1206
end;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1207
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1208
(*
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1209
val demofile = "/home/obua/flyspeck/kepler/LP/cplexPent2.lp45"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1210
val demoout = "/home/obua/flyspeck/kepler/LP/test.out"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1211
val demoresult = "/home/obua/flyspeck/kepler/LP/try/test2.sol"
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1212
	   
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1213
fun loadcplex () = Cplex.relax_strict_ineqs 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1214
		   (Cplex.load_cplexFile demofile)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1215
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1216
fun writecplex lp = Cplex.save_cplexFile demoout lp
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1217
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1218
fun test () = 
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1219
    let
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1220
	val lp = loadcplex ()
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1221
	val lp2 = Cplex.elim_nonfree_bounds lp
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1222
    in
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1223
	writecplex lp2
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1224
    end
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1225
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1226
fun loadresult () = Cplex.load_cplexResult demoresult;
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1227
*)
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1228
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1229
(*val prog = Cplex.load_cplexFile "/home/obua/tmp/pent/graph_0.lpt";
92ff7c903585 - added cplex package to HOL/Matrix
obua
parents:
diff changeset
  1230
val _ = Cplex.solve prog;*)