src/HOL/Matrix/matrixlp.ML
author haftmann
Sat, 11 Feb 2012 00:07:28 +0100
changeset 46544 460b0d81d486
parent 46541 9673597c1b92
permissions -rw-r--r--
brute-force adjustion
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37788
261c61fabc98 corrected title
haftmann
parents: 37764
diff changeset
     1
(*  Title:      HOL/Matrix/matrixlp.ML
20787
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
     2
    Author:     Steven Obua
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
     3
*)
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
     4
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
     5
signature MATRIX_LP =
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
     6
sig
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
     7
  val matrix_compute : cterm -> thm
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
     8
  val matrix_simplify : thm -> thm
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
     9
  val prove_bound : string -> int -> thm
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    10
  val float2real : string * string -> Real.real
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    11
end
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    12
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    13
structure MatrixLP : MATRIX_LP =
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    14
struct
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    15
46540
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    16
val compute_thms = ComputeHOL.prep_thms @{thms "ComputeHOL.compute_list_case" "ComputeHOL.compute_let"
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    17
  "ComputeHOL.compute_if" "ComputeFloat.arith" "SparseMatrix.sparse_row_matrix_arith_simps"
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    18
  "ComputeHOL.compute_bool" "ComputeHOL.compute_pair"
46544
460b0d81d486 brute-force adjustion
haftmann
parents: 46541
diff changeset
    19
  "SparseMatrix.sorted_sp_simps"
460b0d81d486 brute-force adjustion
haftmann
parents: 46541
diff changeset
    20
  "ComputeNumeral.natnorm"}; (*"ComputeNumeral.number_norm"*)
23665
825bea0266db adopted to new computing oracle and fixed bugs introduced by tuning
obua
parents: 22964
diff changeset
    21
46540
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    22
val spm_mult_le_dual_prts_no_let_real = @{thm "spm_mult_le_dual_prts_no_let" [where ?'a = real]}
20787
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    23
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    24
fun lp_dual_estimate_prt lptfile prec =
46540
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    25
  let
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    26
    val cert = cterm_of @{theory}
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    27
    fun var s x = (cert (Var ((s, 0), FloatSparseMatrixBuilder.spmatT)), x)
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    28
    val l = Fspmlp.load lptfile prec false
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    29
    val (y, (A1, A2), (c1, c2), b, (r1, r2)) =
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    30
      let
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    31
        open Fspmlp
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    32
      in
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    33
        (y l |> cert, A l |> pairself cert, c l |> pairself cert, b l |> cert, r12 l |> pairself cert)
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    34
      end
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    35
  in
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    36
    Thm.instantiate ([],
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    37
      [var "A1" A1, var "A2" A2, var "y" y, var "c1" c1, var "c2" c2, var "r1" r1, var "r2" r2, var "b" b])
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    38
      spm_mult_le_dual_prts_no_let_real
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    39
  end
20787
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    40
46540
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    41
val computer = PCompute.make Compute.SML @{theory} compute_thms []
20787
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    42
23665
825bea0266db adopted to new computing oracle and fixed bugs introduced by tuning
obua
parents: 22964
diff changeset
    43
fun matrix_compute c = hd (PCompute.rewrite computer [c])
20787
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    44
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    45
fun matrix_simplify th =
46540
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    46
  let
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    47
    val simp_th = matrix_compute (cprop_of th)
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    48
    val th = Thm.strip_shyps (Thm.equal_elim simp_th th)
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    49
    fun removeTrue th = removeTrue (Thm.implies_elim th TrueI) handle THM _ => th
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    50
  in
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    51
    removeTrue th
5522e28a566e tuned code
haftmann
parents: 46531
diff changeset
    52
  end
20787
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    53
46541
haftmann
parents: 46540
diff changeset
    54
val prove_bound = matrix_simplify oo lp_dual_estimate_prt;
20787
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    55
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    56
val realFromStr = the o Real.fromString;
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    57
fun float2real (x, y) = realFromStr x * Math.pow (2.0, realFromStr y);
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    58
406d990006af moved Matrix/cplex/MatrixLP.ML to Matrix/cplex/matrixlp.ML;
wenzelm
parents:
diff changeset
    59
end