wenzelm@20324
|
1 |
(* Title: HOL/FunDef.thy
|
wenzelm@20324
|
2 |
Author: Alexander Krauss, TU Muenchen
|
wenzelm@22816
|
3 |
*)
|
wenzelm@20324
|
4 |
|
krauss@29125
|
5 |
header {* Function Definitions and Termination Proofs *}
|
wenzelm@20324
|
6 |
|
krauss@19564
|
7 |
theory FunDef
|
krauss@26748
|
8 |
imports Wellfounded
|
wenzelm@22816
|
9 |
uses
|
krauss@29125
|
10 |
"Tools/prop_logic.ML"
|
krauss@29125
|
11 |
"Tools/sat_solver.ML"
|
haftmann@31775
|
12 |
("Tools/Function/fundef_lib.ML")
|
haftmann@31775
|
13 |
("Tools/Function/fundef_common.ML")
|
haftmann@31775
|
14 |
("Tools/Function/inductive_wrap.ML")
|
haftmann@31775
|
15 |
("Tools/Function/context_tree.ML")
|
haftmann@31775
|
16 |
("Tools/Function/fundef_core.ML")
|
haftmann@31775
|
17 |
("Tools/Function/sum_tree.ML")
|
haftmann@31775
|
18 |
("Tools/Function/mutual.ML")
|
haftmann@31775
|
19 |
("Tools/Function/pattern_split.ML")
|
haftmann@31775
|
20 |
("Tools/Function/fundef.ML")
|
haftmann@31775
|
21 |
("Tools/Function/auto_term.ML")
|
haftmann@31775
|
22 |
("Tools/Function/measure_functions.ML")
|
haftmann@31775
|
23 |
("Tools/Function/lexicographic_order.ML")
|
haftmann@31775
|
24 |
("Tools/Function/fundef_datatype.ML")
|
haftmann@31775
|
25 |
("Tools/Function/induction_scheme.ML")
|
haftmann@31775
|
26 |
("Tools/Function/termination.ML")
|
haftmann@31775
|
27 |
("Tools/Function/decompose.ML")
|
haftmann@31775
|
28 |
("Tools/Function/descent.ML")
|
haftmann@31775
|
29 |
("Tools/Function/scnp_solve.ML")
|
haftmann@31775
|
30 |
("Tools/Function/scnp_reconstruct.ML")
|
krauss@19564
|
31 |
begin
|
krauss@19564
|
32 |
|
krauss@29125
|
33 |
subsection {* Definitions with default value. *}
|
krauss@20536
|
34 |
|
krauss@20536
|
35 |
definition
|
wenzelm@21404
|
36 |
THE_default :: "'a \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a" where
|
krauss@20536
|
37 |
"THE_default d P = (if (\<exists>!x. P x) then (THE x. P x) else d)"
|
krauss@20536
|
38 |
|
krauss@20536
|
39 |
lemma THE_defaultI': "\<exists>!x. P x \<Longrightarrow> P (THE_default d P)"
|
wenzelm@22816
|
40 |
by (simp add: theI' THE_default_def)
|
krauss@20536
|
41 |
|
wenzelm@22816
|
42 |
lemma THE_default1_equality:
|
wenzelm@22816
|
43 |
"\<lbrakk>\<exists>!x. P x; P a\<rbrakk> \<Longrightarrow> THE_default d P = a"
|
wenzelm@22816
|
44 |
by (simp add: the1_equality THE_default_def)
|
krauss@20536
|
45 |
|
krauss@20536
|
46 |
lemma THE_default_none:
|
wenzelm@22816
|
47 |
"\<not>(\<exists>!x. P x) \<Longrightarrow> THE_default d P = d"
|
wenzelm@22816
|
48 |
by (simp add:THE_default_def)
|
krauss@20536
|
49 |
|
krauss@20536
|
50 |
|
krauss@19564
|
51 |
lemma fundef_ex1_existence:
|
wenzelm@22816
|
52 |
assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
|
wenzelm@22816
|
53 |
assumes ex1: "\<exists>!y. G x y"
|
wenzelm@22816
|
54 |
shows "G x (f x)"
|
wenzelm@22816
|
55 |
apply (simp only: f_def)
|
wenzelm@22816
|
56 |
apply (rule THE_defaultI')
|
wenzelm@22816
|
57 |
apply (rule ex1)
|
wenzelm@22816
|
58 |
done
|
krauss@21051
|
59 |
|
krauss@19564
|
60 |
lemma fundef_ex1_uniqueness:
|
wenzelm@22816
|
61 |
assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
|
wenzelm@22816
|
62 |
assumes ex1: "\<exists>!y. G x y"
|
wenzelm@22816
|
63 |
assumes elm: "G x (h x)"
|
wenzelm@22816
|
64 |
shows "h x = f x"
|
wenzelm@22816
|
65 |
apply (simp only: f_def)
|
wenzelm@22816
|
66 |
apply (rule THE_default1_equality [symmetric])
|
wenzelm@22816
|
67 |
apply (rule ex1)
|
wenzelm@22816
|
68 |
apply (rule elm)
|
wenzelm@22816
|
69 |
done
|
krauss@19564
|
70 |
|
krauss@19564
|
71 |
lemma fundef_ex1_iff:
|
wenzelm@22816
|
72 |
assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
|
wenzelm@22816
|
73 |
assumes ex1: "\<exists>!y. G x y"
|
wenzelm@22816
|
74 |
shows "(G x y) = (f x = y)"
|
krauss@20536
|
75 |
apply (auto simp:ex1 f_def THE_default1_equality)
|
wenzelm@22816
|
76 |
apply (rule THE_defaultI')
|
wenzelm@22816
|
77 |
apply (rule ex1)
|
wenzelm@22816
|
78 |
done
|
krauss@19564
|
79 |
|
krauss@20654
|
80 |
lemma fundef_default_value:
|
wenzelm@22816
|
81 |
assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
|
wenzelm@22816
|
82 |
assumes graph: "\<And>x y. G x y \<Longrightarrow> D x"
|
wenzelm@22816
|
83 |
assumes "\<not> D x"
|
wenzelm@22816
|
84 |
shows "f x = d x"
|
krauss@20654
|
85 |
proof -
|
krauss@21051
|
86 |
have "\<not>(\<exists>y. G x y)"
|
krauss@20654
|
87 |
proof
|
krauss@21512
|
88 |
assume "\<exists>y. G x y"
|
krauss@21512
|
89 |
hence "D x" using graph ..
|
krauss@21512
|
90 |
with `\<not> D x` show False ..
|
krauss@20654
|
91 |
qed
|
krauss@21051
|
92 |
hence "\<not>(\<exists>!y. G x y)" by blast
|
wenzelm@22816
|
93 |
|
krauss@20654
|
94 |
thus ?thesis
|
krauss@20654
|
95 |
unfolding f_def
|
krauss@20654
|
96 |
by (rule THE_default_none)
|
krauss@20654
|
97 |
qed
|
krauss@20654
|
98 |
|
berghofe@23739
|
99 |
definition in_rel_def[simp]:
|
berghofe@23739
|
100 |
"in_rel R x y == (x, y) \<in> R"
|
berghofe@23739
|
101 |
|
berghofe@23739
|
102 |
lemma wf_in_rel:
|
berghofe@23739
|
103 |
"wf R \<Longrightarrow> wfP (in_rel R)"
|
berghofe@23739
|
104 |
by (simp add: wfP_def)
|
berghofe@23739
|
105 |
|
haftmann@31775
|
106 |
use "Tools/Function/fundef_lib.ML"
|
haftmann@31775
|
107 |
use "Tools/Function/fundef_common.ML"
|
haftmann@31775
|
108 |
use "Tools/Function/inductive_wrap.ML"
|
haftmann@31775
|
109 |
use "Tools/Function/context_tree.ML"
|
haftmann@31775
|
110 |
use "Tools/Function/fundef_core.ML"
|
haftmann@31775
|
111 |
use "Tools/Function/sum_tree.ML"
|
haftmann@31775
|
112 |
use "Tools/Function/mutual.ML"
|
haftmann@31775
|
113 |
use "Tools/Function/pattern_split.ML"
|
haftmann@31775
|
114 |
use "Tools/Function/auto_term.ML"
|
haftmann@31775
|
115 |
use "Tools/Function/fundef.ML"
|
haftmann@31775
|
116 |
use "Tools/Function/fundef_datatype.ML"
|
haftmann@31775
|
117 |
use "Tools/Function/induction_scheme.ML"
|
krauss@19564
|
118 |
|
krauss@25567
|
119 |
setup {*
|
haftmann@31723
|
120 |
Fundef.setup
|
krauss@29125
|
121 |
#> FundefDatatype.setup
|
krauss@25567
|
122 |
#> InductionScheme.setup
|
krauss@25567
|
123 |
*}
|
krauss@19770
|
124 |
|
krauss@29125
|
125 |
subsection {* Measure Functions *}
|
krauss@29125
|
126 |
|
krauss@29125
|
127 |
inductive is_measure :: "('a \<Rightarrow> nat) \<Rightarrow> bool"
|
krauss@29125
|
128 |
where is_measure_trivial: "is_measure f"
|
krauss@29125
|
129 |
|
haftmann@31775
|
130 |
use "Tools/Function/measure_functions.ML"
|
krauss@29125
|
131 |
setup MeasureFunctions.setup
|
krauss@29125
|
132 |
|
krauss@29125
|
133 |
lemma measure_size[measure_function]: "is_measure size"
|
krauss@29125
|
134 |
by (rule is_measure_trivial)
|
krauss@29125
|
135 |
|
krauss@29125
|
136 |
lemma measure_fst[measure_function]: "is_measure f \<Longrightarrow> is_measure (\<lambda>p. f (fst p))"
|
krauss@29125
|
137 |
by (rule is_measure_trivial)
|
krauss@29125
|
138 |
lemma measure_snd[measure_function]: "is_measure f \<Longrightarrow> is_measure (\<lambda>p. f (snd p))"
|
krauss@29125
|
139 |
by (rule is_measure_trivial)
|
krauss@29125
|
140 |
|
haftmann@31775
|
141 |
use "Tools/Function/lexicographic_order.ML"
|
krauss@29125
|
142 |
setup LexicographicOrder.setup
|
krauss@29125
|
143 |
|
krauss@29125
|
144 |
|
krauss@29125
|
145 |
subsection {* Congruence Rules *}
|
krauss@29125
|
146 |
|
haftmann@22838
|
147 |
lemma let_cong [fundef_cong]:
|
haftmann@22838
|
148 |
"M = N \<Longrightarrow> (\<And>x. x = N \<Longrightarrow> f x = g x) \<Longrightarrow> Let M f = Let N g"
|
wenzelm@22816
|
149 |
unfolding Let_def by blast
|
krauss@22622
|
150 |
|
wenzelm@22816
|
151 |
lemmas [fundef_cong] =
|
haftmann@22838
|
152 |
if_cong image_cong INT_cong UN_cong
|
haftmann@22838
|
153 |
bex_cong ball_cong imp_cong
|
krauss@19564
|
154 |
|
wenzelm@22816
|
155 |
lemma split_cong [fundef_cong]:
|
haftmann@22838
|
156 |
"(\<And>x y. (x, y) = q \<Longrightarrow> f x y = g x y) \<Longrightarrow> p = q
|
wenzelm@22816
|
157 |
\<Longrightarrow> split f p = split g q"
|
wenzelm@22816
|
158 |
by (auto simp: split_def)
|
krauss@19934
|
159 |
|
wenzelm@22816
|
160 |
lemma comp_cong [fundef_cong]:
|
haftmann@22838
|
161 |
"f (g x) = f' (g' x') \<Longrightarrow> (f o g) x = (f' o g') x'"
|
wenzelm@22816
|
162 |
unfolding o_apply .
|
krauss@19934
|
163 |
|
krauss@29125
|
164 |
subsection {* Simp rules for termination proofs *}
|
krauss@26875
|
165 |
|
krauss@26749
|
166 |
lemma termination_basic_simps[termination_simp]:
|
krauss@26749
|
167 |
"x < (y::nat) \<Longrightarrow> x < y + z"
|
krauss@26749
|
168 |
"x < z \<Longrightarrow> x < y + z"
|
krauss@26875
|
169 |
"x \<le> y \<Longrightarrow> x \<le> y + (z::nat)"
|
krauss@26875
|
170 |
"x \<le> z \<Longrightarrow> x \<le> y + (z::nat)"
|
krauss@26875
|
171 |
"x < y \<Longrightarrow> x \<le> (y::nat)"
|
krauss@26749
|
172 |
by arith+
|
krauss@26749
|
173 |
|
krauss@26875
|
174 |
declare le_imp_less_Suc[termination_simp]
|
krauss@26875
|
175 |
|
krauss@26875
|
176 |
lemma prod_size_simp[termination_simp]:
|
krauss@26875
|
177 |
"prod_size f g p = f (fst p) + g (snd p) + Suc 0"
|
krauss@26875
|
178 |
by (induct p) auto
|
krauss@26875
|
179 |
|
krauss@29125
|
180 |
subsection {* Decomposition *}
|
krauss@29125
|
181 |
|
krauss@29125
|
182 |
lemma less_by_empty:
|
krauss@29125
|
183 |
"A = {} \<Longrightarrow> A \<subseteq> B"
|
krauss@29125
|
184 |
and union_comp_emptyL:
|
krauss@29125
|
185 |
"\<lbrakk> A O C = {}; B O C = {} \<rbrakk> \<Longrightarrow> (A \<union> B) O C = {}"
|
krauss@29125
|
186 |
and union_comp_emptyR:
|
krauss@29125
|
187 |
"\<lbrakk> A O B = {}; A O C = {} \<rbrakk> \<Longrightarrow> A O (B \<union> C) = {}"
|
krauss@29125
|
188 |
and wf_no_loop:
|
krauss@29125
|
189 |
"R O R = {} \<Longrightarrow> wf R"
|
krauss@29125
|
190 |
by (auto simp add: wf_comp_self[of R])
|
krauss@29125
|
191 |
|
krauss@29125
|
192 |
|
krauss@29125
|
193 |
subsection {* Reduction Pairs *}
|
krauss@29125
|
194 |
|
krauss@29125
|
195 |
definition
|
krauss@32235
|
196 |
"reduction_pair P = (wf (fst P) \<and> fst P O snd P \<subseteq> fst P)"
|
krauss@29125
|
197 |
|
krauss@32235
|
198 |
lemma reduction_pairI[intro]: "wf R \<Longrightarrow> R O S \<subseteq> R \<Longrightarrow> reduction_pair (R, S)"
|
krauss@29125
|
199 |
unfolding reduction_pair_def by auto
|
krauss@29125
|
200 |
|
krauss@29125
|
201 |
lemma reduction_pair_lemma:
|
krauss@29125
|
202 |
assumes rp: "reduction_pair P"
|
krauss@29125
|
203 |
assumes "R \<subseteq> fst P"
|
krauss@29125
|
204 |
assumes "S \<subseteq> snd P"
|
krauss@29125
|
205 |
assumes "wf S"
|
krauss@29125
|
206 |
shows "wf (R \<union> S)"
|
krauss@29125
|
207 |
proof -
|
krauss@32235
|
208 |
from rp `S \<subseteq> snd P` have "wf (fst P)" "fst P O S \<subseteq> fst P"
|
krauss@29125
|
209 |
unfolding reduction_pair_def by auto
|
krauss@29125
|
210 |
with `wf S` have "wf (fst P \<union> S)"
|
krauss@29125
|
211 |
by (auto intro: wf_union_compatible)
|
krauss@29125
|
212 |
moreover from `R \<subseteq> fst P` have "R \<union> S \<subseteq> fst P \<union> S" by auto
|
krauss@29125
|
213 |
ultimately show ?thesis by (rule wf_subset)
|
krauss@29125
|
214 |
qed
|
krauss@29125
|
215 |
|
krauss@29125
|
216 |
definition
|
krauss@29125
|
217 |
"rp_inv_image = (\<lambda>(R,S) f. (inv_image R f, inv_image S f))"
|
krauss@29125
|
218 |
|
krauss@29125
|
219 |
lemma rp_inv_image_rp:
|
krauss@29125
|
220 |
"reduction_pair P \<Longrightarrow> reduction_pair (rp_inv_image P f)"
|
krauss@29125
|
221 |
unfolding reduction_pair_def rp_inv_image_def split_def
|
krauss@29125
|
222 |
by force
|
krauss@29125
|
223 |
|
krauss@29125
|
224 |
|
krauss@29125
|
225 |
subsection {* Concrete orders for SCNP termination proofs *}
|
krauss@29125
|
226 |
|
krauss@29125
|
227 |
definition "pair_less = less_than <*lex*> less_than"
|
haftmann@30428
|
228 |
definition [code del]: "pair_leq = pair_less^="
|
krauss@29125
|
229 |
definition "max_strict = max_ext pair_less"
|
haftmann@30428
|
230 |
definition [code del]: "max_weak = max_ext pair_leq \<union> {({}, {})}"
|
haftmann@30428
|
231 |
definition [code del]: "min_strict = min_ext pair_less"
|
haftmann@30446
|
232 |
definition [code del]: "min_weak = min_ext pair_leq \<union> {({}, {})}"
|
krauss@29125
|
233 |
|
krauss@29125
|
234 |
lemma wf_pair_less[simp]: "wf pair_less"
|
krauss@29125
|
235 |
by (auto simp: pair_less_def)
|
krauss@29125
|
236 |
|
wenzelm@29127
|
237 |
text {* Introduction rules for @{text pair_less}/@{text pair_leq} *}
|
krauss@29125
|
238 |
lemma pair_leqI1: "a < b \<Longrightarrow> ((a, s), (b, t)) \<in> pair_leq"
|
krauss@29125
|
239 |
and pair_leqI2: "a \<le> b \<Longrightarrow> s \<le> t \<Longrightarrow> ((a, s), (b, t)) \<in> pair_leq"
|
krauss@29125
|
240 |
and pair_lessI1: "a < b \<Longrightarrow> ((a, s), (b, t)) \<in> pair_less"
|
krauss@29125
|
241 |
and pair_lessI2: "a \<le> b \<Longrightarrow> s < t \<Longrightarrow> ((a, s), (b, t)) \<in> pair_less"
|
krauss@29125
|
242 |
unfolding pair_leq_def pair_less_def by auto
|
krauss@29125
|
243 |
|
krauss@29125
|
244 |
text {* Introduction rules for max *}
|
krauss@29125
|
245 |
lemma smax_emptyI:
|
krauss@29125
|
246 |
"finite Y \<Longrightarrow> Y \<noteq> {} \<Longrightarrow> ({}, Y) \<in> max_strict"
|
krauss@29125
|
247 |
and smax_insertI:
|
krauss@29125
|
248 |
"\<lbrakk>y \<in> Y; (x, y) \<in> pair_less; (X, Y) \<in> max_strict\<rbrakk> \<Longrightarrow> (insert x X, Y) \<in> max_strict"
|
krauss@29125
|
249 |
and wmax_emptyI:
|
krauss@29125
|
250 |
"finite X \<Longrightarrow> ({}, X) \<in> max_weak"
|
krauss@29125
|
251 |
and wmax_insertI:
|
krauss@29125
|
252 |
"\<lbrakk>y \<in> YS; (x, y) \<in> pair_leq; (XS, YS) \<in> max_weak\<rbrakk> \<Longrightarrow> (insert x XS, YS) \<in> max_weak"
|
krauss@29125
|
253 |
unfolding max_strict_def max_weak_def by (auto elim!: max_ext.cases)
|
krauss@29125
|
254 |
|
krauss@29125
|
255 |
text {* Introduction rules for min *}
|
krauss@29125
|
256 |
lemma smin_emptyI:
|
krauss@29125
|
257 |
"X \<noteq> {} \<Longrightarrow> (X, {}) \<in> min_strict"
|
krauss@29125
|
258 |
and smin_insertI:
|
krauss@29125
|
259 |
"\<lbrakk>x \<in> XS; (x, y) \<in> pair_less; (XS, YS) \<in> min_strict\<rbrakk> \<Longrightarrow> (XS, insert y YS) \<in> min_strict"
|
krauss@29125
|
260 |
and wmin_emptyI:
|
krauss@29125
|
261 |
"(X, {}) \<in> min_weak"
|
krauss@29125
|
262 |
and wmin_insertI:
|
krauss@29125
|
263 |
"\<lbrakk>x \<in> XS; (x, y) \<in> pair_leq; (XS, YS) \<in> min_weak\<rbrakk> \<Longrightarrow> (XS, insert y YS) \<in> min_weak"
|
krauss@29125
|
264 |
by (auto simp: min_strict_def min_weak_def min_ext_def)
|
krauss@29125
|
265 |
|
krauss@29125
|
266 |
text {* Reduction Pairs *}
|
krauss@29125
|
267 |
|
krauss@29125
|
268 |
lemma max_ext_compat:
|
krauss@32235
|
269 |
assumes "R O S \<subseteq> R"
|
krauss@32235
|
270 |
shows "max_ext R O (max_ext S \<union> {({},{})}) \<subseteq> max_ext R"
|
krauss@29125
|
271 |
using assms
|
krauss@29125
|
272 |
apply auto
|
krauss@29125
|
273 |
apply (elim max_ext.cases)
|
krauss@29125
|
274 |
apply rule
|
krauss@29125
|
275 |
apply auto[3]
|
krauss@29125
|
276 |
apply (drule_tac x=xa in meta_spec)
|
krauss@29125
|
277 |
apply simp
|
krauss@29125
|
278 |
apply (erule bexE)
|
krauss@29125
|
279 |
apply (drule_tac x=xb in meta_spec)
|
krauss@29125
|
280 |
by auto
|
krauss@29125
|
281 |
|
krauss@29125
|
282 |
lemma max_rpair_set: "reduction_pair (max_strict, max_weak)"
|
krauss@29125
|
283 |
unfolding max_strict_def max_weak_def
|
krauss@29125
|
284 |
apply (intro reduction_pairI max_ext_wf)
|
krauss@29125
|
285 |
apply simp
|
krauss@29125
|
286 |
apply (rule max_ext_compat)
|
krauss@29125
|
287 |
by (auto simp: pair_less_def pair_leq_def)
|
krauss@29125
|
288 |
|
krauss@29125
|
289 |
lemma min_ext_compat:
|
krauss@32235
|
290 |
assumes "R O S \<subseteq> R"
|
krauss@32235
|
291 |
shows "min_ext R O (min_ext S \<union> {({},{})}) \<subseteq> min_ext R"
|
krauss@29125
|
292 |
using assms
|
krauss@29125
|
293 |
apply (auto simp: min_ext_def)
|
krauss@29125
|
294 |
apply (drule_tac x=ya in bspec, assumption)
|
krauss@29125
|
295 |
apply (erule bexE)
|
krauss@29125
|
296 |
apply (drule_tac x=xc in bspec)
|
krauss@29125
|
297 |
apply assumption
|
krauss@29125
|
298 |
by auto
|
krauss@29125
|
299 |
|
krauss@29125
|
300 |
lemma min_rpair_set: "reduction_pair (min_strict, min_weak)"
|
krauss@29125
|
301 |
unfolding min_strict_def min_weak_def
|
krauss@29125
|
302 |
apply (intro reduction_pairI min_ext_wf)
|
krauss@29125
|
303 |
apply simp
|
krauss@29125
|
304 |
apply (rule min_ext_compat)
|
krauss@29125
|
305 |
by (auto simp: pair_less_def pair_leq_def)
|
krauss@29125
|
306 |
|
krauss@29125
|
307 |
|
krauss@29125
|
308 |
subsection {* Tool setup *}
|
krauss@29125
|
309 |
|
haftmann@31775
|
310 |
use "Tools/Function/termination.ML"
|
haftmann@31775
|
311 |
use "Tools/Function/decompose.ML"
|
haftmann@31775
|
312 |
use "Tools/Function/descent.ML"
|
haftmann@31775
|
313 |
use "Tools/Function/scnp_solve.ML"
|
haftmann@31775
|
314 |
use "Tools/Function/scnp_reconstruct.ML"
|
krauss@29125
|
315 |
|
krauss@29125
|
316 |
setup {* ScnpReconstruct.setup *}
|
wenzelm@30480
|
317 |
|
wenzelm@30480
|
318 |
ML_val -- "setup inactive"
|
wenzelm@30480
|
319 |
{*
|
krauss@29125
|
320 |
Context.theory_map (FundefCommon.set_termination_prover (ScnpReconstruct.decomp_scnp
|
krauss@29125
|
321 |
[ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS]))
|
krauss@29125
|
322 |
*}
|
krauss@26875
|
323 |
|
krauss@19564
|
324 |
end
|