16873
|
1 |
(* Title: HOL/Matrix/cplex/MatrixLP.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Steven Obua
|
|
4 |
*)
|
|
5 |
|
|
6 |
signature MATRIX_LP =
|
|
7 |
sig
|
|
8 |
val lp_dual_estimate_prt : string -> int -> thm
|
19404
|
9 |
val lp_dual_estimate_prt_primitive : cterm * (cterm * cterm) * (cterm * cterm) * cterm * (cterm * cterm) -> thm
|
16873
|
10 |
val matrix_compute : cterm -> thm
|
|
11 |
val matrix_simplify : thm -> thm
|
|
12 |
val prove_bound : string -> int -> thm
|
|
13 |
val float2real : string * string -> Real.real
|
|
14 |
end
|
|
15 |
|
|
16 |
structure MatrixLP : MATRIX_LP =
|
|
17 |
struct
|
|
18 |
|
|
19 |
val sg = sign_of (theory "MatrixLP")
|
|
20 |
|
|
21 |
fun inst_real thm = standard (Thm.instantiate ([(ctyp_of sg (TVar (hd (term_tvars (prop_of thm)))),
|
|
22 |
ctyp_of sg HOLogic.realT)], []) thm)
|
|
23 |
|
19404
|
24 |
fun lp_dual_estimate_prt_primitive (y, (A1, A2), (c1, c2), b, (r1, r2)) =
|
|
25 |
let
|
|
26 |
val th = inst_real (thm "SparseMatrix.spm_mult_le_dual_prts_no_let")
|
|
27 |
fun var s x = (cterm_of sg (Var ((s,0), FloatSparseMatrixBuilder.real_spmatT)), x)
|
|
28 |
val th = Thm.instantiate ([], [var "A1" A1, var "A2" A2, var "y" y, var "c1" c1, var "c2" c2,
|
|
29 |
var "r1" r1, var "r2" r2, var "b" b]) th
|
|
30 |
in
|
|
31 |
th
|
|
32 |
end
|
|
33 |
|
16873
|
34 |
fun lp_dual_estimate_prt lptfile prec =
|
|
35 |
let
|
19404
|
36 |
val certificate =
|
16873
|
37 |
let
|
|
38 |
open fspmlp
|
|
39 |
val l = load lptfile prec false
|
|
40 |
in
|
|
41 |
(y l, A l, c l, b l, r12 l)
|
|
42 |
end
|
|
43 |
in
|
19404
|
44 |
lp_dual_estimate_prt_primitive certificate
|
16873
|
45 |
end
|
|
46 |
|
|
47 |
fun read_ct s = read_cterm sg (s, TypeInfer.logicT);
|
|
48 |
|
|
49 |
fun is_meta_eq th =
|
|
50 |
let
|
|
51 |
fun check ((Const ("==", _)) $ _ $ _) = true
|
|
52 |
| check _ = false
|
|
53 |
in
|
|
54 |
check (concl_of th)
|
|
55 |
end
|
|
56 |
|
|
57 |
fun prep ths = (Library.filter is_meta_eq ths) @ (map (standard o mk_meta_eq) (Library.filter (not o is_meta_eq) ths))
|
|
58 |
|
|
59 |
fun make ths = Compute.basic_make sg ths
|
|
60 |
|
|
61 |
fun inst_tvar ty thm =
|
|
62 |
let
|
|
63 |
val ord = prod_ord (prod_ord string_ord int_ord) (list_ord string_ord)
|
|
64 |
val v = TVar (hd (sort ord (term_tvars (prop_of thm))))
|
|
65 |
in
|
|
66 |
standard (Thm.instantiate ([(ctyp_of sg v, ctyp_of sg ty)], []) thm)
|
|
67 |
end
|
|
68 |
|
|
69 |
fun inst_tvars [] thms = thms
|
|
70 |
| inst_tvars (ty::tys) thms = inst_tvars tys (map (inst_tvar ty) thms)
|
|
71 |
|
|
72 |
val matrix_compute =
|
|
73 |
let
|
|
74 |
val spvecT = FloatSparseMatrixBuilder.real_spvecT
|
|
75 |
val spmatT = FloatSparseMatrixBuilder.real_spmatT
|
|
76 |
val spvecT_elem = HOLogic.mk_prodT (HOLogic.natT, HOLogic.realT)
|
|
77 |
val spmatT_elem = HOLogic.mk_prodT (HOLogic.natT, spvecT)
|
|
78 |
val case_compute = map thm ["list_case_compute", "list_case_compute_empty", "list_case_compute_cons"]
|
|
79 |
val ths =
|
|
80 |
prep (
|
|
81 |
(inst_tvars [HOLogic.intT, HOLogic.natT] (thms "Let_compute"))
|
|
82 |
@ (inst_tvars [HOLogic.intT, HOLogic.intT] (thms "Let_compute"))
|
|
83 |
@ (map (fn t => inst_tvar t (thm "If_True")) [HOLogic.intT, HOLogic.natT, HOLogic.realT, spvecT, spmatT, HOLogic.boolT])
|
|
84 |
@ (map (fn t => inst_tvar t (thm "If_False")) [HOLogic.intT, HOLogic.natT, HOLogic.realT, spvecT, spmatT, HOLogic.boolT])
|
|
85 |
@ (thms "MatrixLP.float_arith")
|
|
86 |
@ (map (inst_tvar HOLogic.realT) (thms "MatrixLP.sparse_row_matrix_arith_simps"))
|
|
87 |
@ (thms "MatrixLP.boolarith")
|
|
88 |
@ (inst_tvars [HOLogic.natT, HOLogic.realT] [thm "fst_compute", thm "snd_compute"])
|
|
89 |
@ (inst_tvars [HOLogic.natT, FloatSparseMatrixBuilder.real_spvecT] [thm "fst_compute", thm "snd_compute"])
|
|
90 |
@ (inst_tvars [HOLogic.boolT, spmatT_elem] case_compute)
|
|
91 |
@ (inst_tvars [HOLogic.boolT, spvecT_elem] case_compute)
|
|
92 |
@ (inst_tvars [HOLogic.boolT, HOLogic.realT] case_compute)
|
|
93 |
@ (inst_tvars [spvecT] (thms "MatrixLP.sorted_sp_simps"))
|
|
94 |
@ (inst_tvars [HOLogic.realT] (thms "MatrixLP.sorted_sp_simps"))
|
|
95 |
@ [thm "zero_eq_Numeral0_nat", thm "one_eq_Numeral1_nat"]
|
|
96 |
@ (inst_tvars [HOLogic.intT] [thm "zero_eq_Numeral0_nring", thm "one_eq_Numeral1_nring"])
|
|
97 |
@ (inst_tvars [HOLogic.realT] [thm "zero_eq_Numeral0_nring", thm "one_eq_Numeral1_nring"]))
|
|
98 |
|
|
99 |
val c = make ths
|
|
100 |
in
|
|
101 |
Compute.rewrite c
|
|
102 |
end
|
|
103 |
|
|
104 |
fun matrix_simplify th =
|
|
105 |
let
|
|
106 |
val simp_th = matrix_compute (cprop_of th)
|
|
107 |
val th = strip_shyps (equal_elim simp_th th)
|
|
108 |
fun removeTrue th = removeTrue (implies_elim th TrueI) handle _ => th
|
|
109 |
in
|
|
110 |
removeTrue th
|
|
111 |
end
|
|
112 |
|
|
113 |
fun prove_bound lptfile prec =
|
|
114 |
let
|
|
115 |
val th = lp_dual_estimate_prt lptfile prec
|
|
116 |
in
|
|
117 |
matrix_simplify th
|
|
118 |
end
|
|
119 |
|
|
120 |
fun realFromStr s = valOf (Real.fromString s)
|
|
121 |
fun float2real (x,y) = (realFromStr x) * (Math.pow (2.0, realFromStr y))
|
|
122 |
|
|
123 |
end
|
|
124 |
|