# HG changeset patch # User wenzelm # Date 1331985160 -3600 # Node ID 9f492f5b0cec864e9542912c38506144223c5276 # Parent 15ce93dfe6da2cee59314ce90c180410b731e335 renamed HOL-Matrix to HOL-Matrix_LP to avoid name clash with AFP; diff -r 15ce93dfe6da -r 9f492f5b0cec Admin/isatest/isatest-stats --- a/Admin/isatest/isatest-stats Sat Mar 17 12:26:19 2012 +0100 +++ b/Admin/isatest/isatest-stats Sat Mar 17 12:52:40 2012 +0100 @@ -32,7 +32,7 @@ HOL-Isar_Examples HOL-Lattice HOL-Library-Codegenerator_Test - HOL-Matrix + HOL-Matrix_LP HOL-Metis_Examples HOL-MicroJava HOL-Mirabelle diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Sat Mar 17 12:26:19 2012 +0100 +++ b/src/HOL/IsaMakefile Sat Mar 17 12:52:40 2012 +0100 @@ -52,7 +52,7 @@ HOL-Isar_Examples \ HOL-Lattice \ HOL-Library-Codegenerator_Test \ - HOL-Matrix \ + HOL-Matrix_LP \ HOL-Metis_Examples \ HOL-MicroJava \ HOL-Mirabelle \ @@ -1172,22 +1172,26 @@ @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL SET_Protocol -## HOL-Matrix +## HOL-Matrix_LP -HOL-Matrix: HOL $(LOG)/HOL-Matrix.gz +HOL-Matrix_LP: HOL $(LOG)/HOL-Matrix_LP.gz -$(LOG)/HOL-Matrix.gz: $(OUT)/HOL Matrix/ComputeFloat.thy \ - Matrix/ComputeHOL.thy Matrix/ComputeNumeral.thy \ - Matrix/Compute_Oracle/Compute_Oracle.thy Matrix/Compute_Oracle/am.ML \ - Matrix/Compute_Oracle/am_compiler.ML Matrix/Compute_Oracle/am_ghc.ML \ - Matrix/Compute_Oracle/am_interpreter.ML \ - Matrix/Compute_Oracle/am_sml.ML Matrix/Compute_Oracle/compute.ML \ - Matrix/Compute_Oracle/linker.ML Matrix/Cplex.thy \ - Matrix/CplexMatrixConverter.ML Matrix/Cplex_tools.ML \ - Matrix/FloatSparseMatrixBuilder.ML Matrix/LP.thy Matrix/Matrix.thy \ - Matrix/ROOT.ML Matrix/SparseMatrix.thy Matrix/document/root.tex \ - Matrix/fspmlp.ML Matrix/matrixlp.ML Tools/float_arith.ML - @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Matrix +$(LOG)/HOL-Matrix_LP.gz: $(OUT)/HOL Matrix_LP/ComputeFloat.thy \ + Matrix_LP/ComputeHOL.thy Matrix_LP/ComputeNumeral.thy \ + Matrix_LP/Compute_Oracle/Compute_Oracle.thy \ + Matrix_LP/Compute_Oracle/am.ML \ + Matrix_LP/Compute_Oracle/am_compiler.ML \ + Matrix_LP/Compute_Oracle/am_ghc.ML \ + Matrix_LP/Compute_Oracle/am_interpreter.ML \ + Matrix_LP/Compute_Oracle/am_sml.ML \ + Matrix_LP/Compute_Oracle/compute.ML \ + Matrix_LP/Compute_Oracle/linker.ML Matrix_LP/Cplex.thy \ + Matrix_LP/CplexMatrixConverter.ML Matrix_LP/Cplex_tools.ML \ + Matrix_LP/FloatSparseMatrixBuilder.ML Matrix_LP/LP.thy \ + Matrix_LP/Matrix.thy Matrix_LP/ROOT.ML Matrix_LP/SparseMatrix.thy \ + Matrix_LP/document/root.tex Matrix_LP/fspmlp.ML \ + Matrix_LP/matrixlp.ML Tools/float_arith.ML + @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Matrix_LP ## TLA @@ -1901,9 +1905,9 @@ $(LOG)/HOL-Lattice $(LOG)/HOL-Lattice.gz \ $(LOG)/HOL-Lex.gz $(LOG)/HOL-Library.gz \ $(LOG)/HOL-Library-Codegenerator_Test.gz \ - $(LOG)/HOL-Main.gz $(LOG)/HOL-Matrix \ - $(LOG)/HOL-Matrix.gz $(LOG)/HOL-Metis_Examples.gz \ - $(LOG)/HOL-MicroJava.gz $(LOG)/HOL-Mirabelle.gz \ + $(LOG)/HOL-Main.gz $(LOG)/HOL-Matrix_LP.gz \ + $(LOG)/HOL-Metis_Examples.gz $(LOG)/HOL-MicroJava.gz \ + $(LOG)/HOL-Mirabelle.gz \ $(LOG)/HOL-Multivariate_Analysis.gz \ $(LOG)/HOL-Mutabelle.gz $(LOG)/HOL-NSA-Examples.gz \ $(LOG)/HOL-NSA.gz $(LOG)/HOL-NanoJava.gz \ diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/ComputeFloat.thy --- a/src/HOL/Matrix/ComputeFloat.thy Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,309 +0,0 @@ -(* Title: HOL/Matrix/ComputeFloat.thy - Author: Steven Obua -*) - -header {* Floating Point Representation of the Reals *} - -theory ComputeFloat -imports Complex_Main "~~/src/HOL/Library/Lattice_Algebras" -uses "~~/src/Tools/float.ML" ("~~/src/HOL/Tools/float_arith.ML") -begin - -definition int_of_real :: "real \ int" - where "int_of_real x = (SOME y. real y = x)" - -definition real_is_int :: "real \ bool" - where "real_is_int x = (EX (u::int). x = real u)" - -lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))" - by (auto simp add: real_is_int_def int_of_real_def) - -lemma real_is_int_real[simp]: "real_is_int (real (x::int))" -by (auto simp add: real_is_int_def int_of_real_def) - -lemma int_of_real_real[simp]: "int_of_real (real x) = x" -by (simp add: int_of_real_def) - -lemma real_int_of_real[simp]: "real_is_int x \ real (int_of_real x) = x" -by (auto simp add: int_of_real_def real_is_int_def) - -lemma real_is_int_add_int_of_real: "real_is_int a \ real_is_int b \ (int_of_real (a+b)) = (int_of_real a) + (int_of_real b)" -by (auto simp add: int_of_real_def real_is_int_def) - -lemma real_is_int_add[simp]: "real_is_int a \ real_is_int b \ real_is_int (a+b)" -apply (subst real_is_int_def2) -apply (simp add: real_is_int_add_int_of_real real_int_of_real) -done - -lemma int_of_real_sub: "real_is_int a \ real_is_int b \ (int_of_real (a-b)) = (int_of_real a) - (int_of_real b)" -by (auto simp add: int_of_real_def real_is_int_def) - -lemma real_is_int_sub[simp]: "real_is_int a \ real_is_int b \ real_is_int (a-b)" -apply (subst real_is_int_def2) -apply (simp add: int_of_real_sub real_int_of_real) -done - -lemma real_is_int_rep: "real_is_int x \ ?! (a::int). real a = x" -by (auto simp add: real_is_int_def) - -lemma int_of_real_mult: - assumes "real_is_int a" "real_is_int b" - shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)" - using assms - by (auto simp add: real_is_int_def real_of_int_mult[symmetric] - simp del: real_of_int_mult) - -lemma real_is_int_mult[simp]: "real_is_int a \ real_is_int b \ real_is_int (a*b)" -apply (subst real_is_int_def2) -apply (simp add: int_of_real_mult) -done - -lemma real_is_int_0[simp]: "real_is_int (0::real)" -by (simp add: real_is_int_def int_of_real_def) - -lemma real_is_int_1[simp]: "real_is_int (1::real)" -proof - - have "real_is_int (1::real) = real_is_int(real (1::int))" by auto - also have "\ = True" by (simp only: real_is_int_real) - ultimately show ?thesis by auto -qed - -lemma real_is_int_n1: "real_is_int (-1::real)" -proof - - have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto - also have "\ = True" by (simp only: real_is_int_real) - ultimately show ?thesis by auto -qed - -lemma real_is_int_number_of[simp]: "real_is_int ((number_of \ int \ real) x)" - by (auto simp: real_is_int_def intro!: exI[of _ "number_of x"]) - -lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)" -by (simp add: int_of_real_def) - -lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)" -proof - - have 1: "(1::real) = real (1::int)" by auto - show ?thesis by (simp only: 1 int_of_real_real) -qed - -lemma int_of_real_number_of[simp]: "int_of_real (number_of b) = number_of b" - unfolding int_of_real_def - by (intro some_equality) - (auto simp add: real_of_int_inject[symmetric] simp del: real_of_int_inject) - -lemma int_div_zdiv: "int (a div b) = (int a) div (int b)" -by (rule zdiv_int) - -lemma int_mod_zmod: "int (a mod b) = (int a) mod (int b)" -by (rule zmod_int) - -lemma abs_div_2_less: "a \ 0 \ a \ -1 \ abs((a::int) div 2) < abs a" -by arith - -lemma norm_0_1: "(0::_::number_ring) = Numeral0 & (1::_::number_ring) = Numeral1" - by auto - -lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)" - by simp - -lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)" - by simp - -lemma mult_left_one: "1 * a = (a::'a::semiring_1)" - by simp - -lemma mult_right_one: "a * 1 = (a::'a::semiring_1)" - by simp - -lemma int_pow_0: "(a::int)^(Numeral0) = 1" - by simp - -lemma int_pow_1: "(a::int)^(Numeral1) = a" - by simp - -lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0" - by simp - -lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1" - by simp - -lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0" - by simp - -lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1" - by simp - -lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1" - by simp - -lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1" -proof - - have 1:"((-1)::nat) = 0" - by simp - show ?thesis by (simp add: 1) -qed - -lemma fst_cong: "a=a' \ fst (a,b) = fst (a',b)" - by simp - -lemma snd_cong: "b=b' \ snd (a,b) = snd (a,b')" - by simp - -lemma lift_bool: "x \ x=True" - by simp - -lemma nlift_bool: "~x \ x=False" - by simp - -lemma not_false_eq_true: "(~ False) = True" by simp - -lemma not_true_eq_false: "(~ True) = False" by simp - -lemmas binarith = - normalize_bin_simps - pred_bin_simps succ_bin_simps - add_bin_simps minus_bin_simps mult_bin_simps - -lemma int_eq_number_of_eq: - "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)" - by (rule eq_number_of_eq) - -lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)" - by (simp only: iszero_number_of_Pls) - -lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))" - by simp - -lemma int_iszero_number_of_Bit0: "iszero ((number_of (Int.Bit0 w))::int) = iszero ((number_of w)::int)" - by simp - -lemma int_iszero_number_of_Bit1: "\ iszero ((number_of (Int.Bit1 w))::int)" - by simp - -lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)" - unfolding neg_def number_of_is_id by simp - -lemma int_not_neg_number_of_Pls: "\ (neg (Numeral0::int))" - by simp - -lemma int_neg_number_of_Min: "neg (-1::int)" - by simp - -lemma int_neg_number_of_Bit0: "neg ((number_of (Int.Bit0 w))::int) = neg ((number_of w)::int)" - by simp - -lemma int_neg_number_of_Bit1: "neg ((number_of (Int.Bit1 w))::int) = neg ((number_of w)::int)" - by simp - -lemma int_le_number_of_eq: "(((number_of x)::int) \ number_of y) = (\ neg ((number_of (y + (uminus x)))::int))" - unfolding neg_def number_of_is_id by (simp add: not_less) - -lemmas intarithrel = - int_eq_number_of_eq - lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_Bit0 - lift_bool[OF int_iszero_number_of_Bit1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min] - int_neg_number_of_Bit0 int_neg_number_of_Bit1 int_le_number_of_eq - -lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)" - by simp - -lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (v + (uminus w))" - by simp - -lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (v * w)" - by simp - -lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)" - by simp - -lemmas intarith = int_number_of_add_sym int_number_of_minus_sym int_number_of_diff_sym int_number_of_mult_sym - -lemmas natarith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of - -lemmas powerarith = nat_number_of zpower_number_of_even - zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring] - zpower_Pls zpower_Min - -definition float :: "(int \ int) \ real" where - "float = (\(a, b). real a * 2 powr real b)" - -lemma float_add_l0: "float (0, e) + x = x" - by (simp add: float_def) - -lemma float_add_r0: "x + float (0, e) = x" - by (simp add: float_def) - -lemma float_add: - "float (a1, e1) + float (a2, e2) = - (if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1) else float (a1*2^(nat (e1-e2))+a2, e2))" - by (simp add: float_def algebra_simps powr_realpow[symmetric] powr_divide2[symmetric]) - -lemma float_mult_l0: "float (0, e) * x = float (0, 0)" - by (simp add: float_def) - -lemma float_mult_r0: "x * float (0, e) = float (0, 0)" - by (simp add: float_def) - -lemma float_mult: - "float (a1, e1) * float (a2, e2) = (float (a1 * a2, e1 + e2))" - by (simp add: float_def powr_add) - -lemma float_minus: - "- (float (a,b)) = float (-a, b)" - by (simp add: float_def) - -lemma zero_le_float: - "(0 <= float (a,b)) = (0 <= a)" - using powr_gt_zero[of 2 "real b", arith] - by (simp add: float_def zero_le_mult_iff) - -lemma float_le_zero: - "(float (a,b) <= 0) = (a <= 0)" - using powr_gt_zero[of 2 "real b", arith] - by (simp add: float_def mult_le_0_iff) - -lemma float_abs: - "abs (float (a,b)) = (if 0 <= a then (float (a,b)) else (float (-a,b)))" - using powr_gt_zero[of 2 "real b", arith] - by (simp add: float_def abs_if mult_less_0_iff) - -lemma float_zero: - "float (0, b) = 0" - by (simp add: float_def) - -lemma float_pprt: - "pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0, b)))" - by (auto simp add: zero_le_float float_le_zero float_zero) - -lemma float_nprt: - "nprt (float (a, b)) = (if 0 <= a then (float (0,b)) else (float (a, b)))" - by (auto simp add: zero_le_float float_le_zero float_zero) - -definition lbound :: "real \ real" - where "lbound x = min 0 x" - -definition ubound :: "real \ real" - where "ubound x = max 0 x" - -lemma lbound: "lbound x \ x" - by (simp add: lbound_def) - -lemma ubound: "x \ ubound x" - by (simp add: ubound_def) - -lemma pprt_lbound: "pprt (lbound x) = float (0, 0)" - by (auto simp: float_def lbound_def) - -lemma nprt_ubound: "nprt (ubound x) = float (0, 0)" - by (auto simp: float_def ubound_def) - -lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0 - float_minus float_abs zero_le_float float_pprt float_nprt pprt_lbound nprt_ubound - -(* for use with the compute oracle *) -lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false - -use "~~/src/HOL/Tools/float_arith.ML" - -end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/ComputeHOL.thy --- a/src/HOL/Matrix/ComputeHOL.thy Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,187 +0,0 @@ -theory ComputeHOL -imports Complex_Main "Compute_Oracle/Compute_Oracle" -begin - -lemma Trueprop_eq_eq: "Trueprop X == (X == True)" by (simp add: atomize_eq) -lemma meta_eq_trivial: "x == y \ x == y" by simp -lemma meta_eq_imp_eq: "x == y \ x = y" by auto -lemma eq_trivial: "x = y \ x = y" by auto -lemma bool_to_true: "x :: bool \ x == True" by simp -lemma transmeta_1: "x = y \ y == z \ x = z" by simp -lemma transmeta_2: "x == y \ y = z \ x = z" by simp -lemma transmeta_3: "x == y \ y == z \ x = z" by simp - - -(**** compute_if ****) - -lemma If_True: "If True = (\ x y. x)" by ((rule ext)+,auto) -lemma If_False: "If False = (\ x y. y)" by ((rule ext)+, auto) - -lemmas compute_if = If_True If_False - -(**** compute_bool ****) - -lemma bool1: "(\ True) = False" by blast -lemma bool2: "(\ False) = True" by blast -lemma bool3: "(P \ True) = P" by blast -lemma bool4: "(True \ P) = P" by blast -lemma bool5: "(P \ False) = False" by blast -lemma bool6: "(False \ P) = False" by blast -lemma bool7: "(P \ True) = True" by blast -lemma bool8: "(True \ P) = True" by blast -lemma bool9: "(P \ False) = P" by blast -lemma bool10: "(False \ P) = P" by blast -lemma bool11: "(True \ P) = P" by blast -lemma bool12: "(P \ True) = True" by blast -lemma bool13: "(True \ P) = P" by blast -lemma bool14: "(P \ False) = (\ P)" by blast -lemma bool15: "(False \ P) = True" by blast -lemma bool16: "(False = False) = True" by blast -lemma bool17: "(True = True) = True" by blast -lemma bool18: "(False = True) = False" by blast -lemma bool19: "(True = False) = False" by blast - -lemmas compute_bool = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10 bool11 bool12 bool13 bool14 bool15 bool16 bool17 bool18 bool19 - - -(*** compute_pair ***) - -lemma compute_fst: "fst (x,y) = x" by simp -lemma compute_snd: "snd (x,y) = y" by simp -lemma compute_pair_eq: "((a, b) = (c, d)) = (a = c \ b = d)" by auto - -lemma prod_case_simp: "prod_case f (x,y) = f x y" by simp - -lemmas compute_pair = compute_fst compute_snd compute_pair_eq prod_case_simp - -(*** compute_option ***) - -lemma compute_the: "the (Some x) = x" by simp -lemma compute_None_Some_eq: "(None = Some x) = False" by auto -lemma compute_Some_None_eq: "(Some x = None) = False" by auto -lemma compute_None_None_eq: "(None = None) = True" by auto -lemma compute_Some_Some_eq: "(Some x = Some y) = (x = y)" by auto - -definition option_case_compute :: "'b option \ 'a \ ('b \ 'a) \ 'a" - where "option_case_compute opt a f = option_case a f opt" - -lemma option_case_compute: "option_case = (\ a f opt. option_case_compute opt a f)" - by (simp add: option_case_compute_def) - -lemma option_case_compute_None: "option_case_compute None = (\ a f. a)" - apply (rule ext)+ - apply (simp add: option_case_compute_def) - done - -lemma option_case_compute_Some: "option_case_compute (Some x) = (\ a f. f x)" - apply (rule ext)+ - apply (simp add: option_case_compute_def) - done - -lemmas compute_option_case = option_case_compute option_case_compute_None option_case_compute_Some - -lemmas compute_option = compute_the compute_None_Some_eq compute_Some_None_eq compute_None_None_eq compute_Some_Some_eq compute_option_case - -(**** compute_list_length ****) - -lemma length_cons:"length (x#xs) = 1 + (length xs)" - by simp - -lemma length_nil: "length [] = 0" - by simp - -lemmas compute_list_length = length_nil length_cons - -(*** compute_list_case ***) - -definition list_case_compute :: "'b list \ 'a \ ('b \ 'b list \ 'a) \ 'a" - where "list_case_compute l a f = list_case a f l" - -lemma list_case_compute: "list_case = (\ (a::'a) f (l::'b list). list_case_compute l a f)" - apply (rule ext)+ - apply (simp add: list_case_compute_def) - done - -lemma list_case_compute_empty: "list_case_compute ([]::'b list) = (\ (a::'a) f. a)" - apply (rule ext)+ - apply (simp add: list_case_compute_def) - done - -lemma list_case_compute_cons: "list_case_compute (u#v) = (\ (a::'a) f. (f (u::'b) v))" - apply (rule ext)+ - apply (simp add: list_case_compute_def) - done - -lemmas compute_list_case = list_case_compute list_case_compute_empty list_case_compute_cons - -(*** compute_list_nth ***) -(* Of course, you will need computation with nats for this to work \ *) - -lemma compute_list_nth: "((x#xs) ! n) = (if n = 0 then x else (xs ! (n - 1)))" - by (cases n, auto) - -(*** compute_list ***) - -lemmas compute_list = compute_list_case compute_list_length compute_list_nth - -(*** compute_let ***) - -lemmas compute_let = Let_def - -(***********************) -(* Everything together *) -(***********************) - -lemmas compute_hol = compute_if compute_bool compute_pair compute_option compute_list compute_let - -ML {* -signature ComputeHOL = -sig - val prep_thms : thm list -> thm list - val to_meta_eq : thm -> thm - val to_hol_eq : thm -> thm - val symmetric : thm -> thm - val trans : thm -> thm -> thm -end - -structure ComputeHOL : ComputeHOL = -struct - -local -fun lhs_of eq = fst (Thm.dest_equals (cprop_of eq)); -in -fun rewrite_conv [] ct = raise CTERM ("rewrite_conv", [ct]) - | rewrite_conv (eq :: eqs) ct = - Thm.instantiate (Thm.match (lhs_of eq, ct)) eq - handle Pattern.MATCH => rewrite_conv eqs ct; -end - -val convert_conditions = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.try_conv (rewrite_conv [@{thm "Trueprop_eq_eq"}]))) - -val eq_th = @{thm "HOL.eq_reflection"} -val meta_eq_trivial = @{thm "ComputeHOL.meta_eq_trivial"} -val bool_to_true = @{thm "ComputeHOL.bool_to_true"} - -fun to_meta_eq th = eq_th OF [th] handle THM _ => meta_eq_trivial OF [th] handle THM _ => bool_to_true OF [th] - -fun to_hol_eq th = @{thm "meta_eq_imp_eq"} OF [th] handle THM _ => @{thm "eq_trivial"} OF [th] - -fun prep_thms ths = map (convert_conditions o to_meta_eq) ths - -fun symmetric th = @{thm "HOL.sym"} OF [th] handle THM _ => @{thm "Pure.symmetric"} OF [th] - -local - val trans_HOL = @{thm "HOL.trans"} - val trans_HOL_1 = @{thm "ComputeHOL.transmeta_1"} - val trans_HOL_2 = @{thm "ComputeHOL.transmeta_2"} - val trans_HOL_3 = @{thm "ComputeHOL.transmeta_3"} - fun tr [] th1 th2 = trans_HOL OF [th1, th2] - | tr (t::ts) th1 th2 = (t OF [th1, th2] handle THM _ => tr ts th1 th2) -in - fun trans th1 th2 = tr [trans_HOL, trans_HOL_1, trans_HOL_2, trans_HOL_3] th1 th2 -end - -end -*} - -end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/ComputeNumeral.thy --- a/src/HOL/Matrix/ComputeNumeral.thy Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,189 +0,0 @@ -theory ComputeNumeral -imports ComputeHOL ComputeFloat -begin - -(* normalization of bit strings *) -lemmas bitnorm = normalize_bin_simps - -(* neg for bit strings *) -lemma neg1: "neg Int.Pls = False" by (simp add: Int.Pls_def) -lemma neg2: "neg Int.Min = True" apply (subst Int.Min_def) by auto -lemma neg3: "neg (Int.Bit0 x) = neg x" apply (simp add: neg_def) apply (subst Bit0_def) by auto -lemma neg4: "neg (Int.Bit1 x) = neg x" apply (simp add: neg_def) apply (subst Bit1_def) by auto -lemmas bitneg = neg1 neg2 neg3 neg4 - -(* iszero for bit strings *) -lemma iszero1: "iszero Int.Pls = True" by (simp add: Int.Pls_def iszero_def) -lemma iszero2: "iszero Int.Min = False" apply (subst Int.Min_def) apply (subst iszero_def) by simp -lemma iszero3: "iszero (Int.Bit0 x) = iszero x" apply (subst Int.Bit0_def) apply (subst iszero_def)+ by auto -lemma iszero4: "iszero (Int.Bit1 x) = False" apply (subst Int.Bit1_def) apply (subst iszero_def)+ apply simp by arith -lemmas bitiszero = iszero1 iszero2 iszero3 iszero4 - -(* lezero for bit strings *) -definition "lezero x \ x \ 0" -lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto -lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto -lemma lezero3: "lezero (Int.Bit0 x) = lezero x" unfolding Int.Bit0_def lezero_def by auto -lemma lezero4: "lezero (Int.Bit1 x) = neg x" unfolding Int.Bit1_def lezero_def neg_def by auto -lemmas bitlezero = lezero1 lezero2 lezero3 lezero4 - -(* equality for bit strings *) -lemmas biteq = eq_bin_simps - -(* x < y for bit strings *) -lemmas bitless = less_bin_simps - -(* x \ y for bit strings *) -lemmas bitle = le_bin_simps - -(* succ for bit strings *) -lemmas bitsucc = succ_bin_simps - -(* pred for bit strings *) -lemmas bitpred = pred_bin_simps - -(* unary minus for bit strings *) -lemmas bituminus = minus_bin_simps - -(* addition for bit strings *) -lemmas bitadd = add_bin_simps - -(* multiplication for bit strings *) -lemma mult_Pls_right: "x * Int.Pls = Int.Pls" by (simp add: Pls_def) -lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute) simp -lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0) -lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp -lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)" - unfolding Bit0_def Bit1_def by (simp add: algebra_simps) -lemmas bitmul = mult_Pls mult_Min mult_Pls_right mult_Min_right multb0x multxb0 multb1 - -lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul - -definition "nat_norm_number_of (x::nat) = x" - -lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)" - apply (simp add: nat_norm_number_of_def) - unfolding lezero_def iszero_def neg_def - apply (simp add: numeral_simps) - done - -(* Normalization of nat literals *) -lemma natnorm0: "(0::nat) = number_of (Int.Pls)" by auto -lemma natnorm1: "(1 :: nat) = number_of (Int.Bit1 Int.Pls)" by auto -lemmas natnorm = natnorm0 natnorm1 nat_norm_number_of - -(* Suc *) -lemma natsuc: "Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))" by (auto simp add: number_of_is_id) - -(* Addition for nat *) -lemma natadd: "number_of x + ((number_of y)::nat) = (if neg x then (number_of y) else (if neg y then number_of x else (number_of (x + y))))" - unfolding nat_number_of_def number_of_is_id neg_def - by auto - -(* Subtraction for nat *) -lemma natsub: "(number_of x) - ((number_of y)::nat) = - (if neg x then 0 else (if neg y then number_of x else (nat_norm_number_of (number_of (x + (- y))))))" - unfolding nat_norm_number_of - by (auto simp add: number_of_is_id neg_def lezero_def iszero_def Let_def nat_number_of_def) - -(* Multiplication for nat *) -lemma natmul: "(number_of x) * ((number_of y)::nat) = - (if neg x then 0 else (if neg y then 0 else number_of (x * y)))" - unfolding nat_number_of_def number_of_is_id neg_def - by (simp add: nat_mult_distrib) - -lemma nateq: "(((number_of x)::nat) = (number_of y)) = ((lezero x \ lezero y) \ (x = y))" - by (auto simp add: iszero_def lezero_def neg_def number_of_is_id) - -lemma natless: "(((number_of x)::nat) < (number_of y)) = ((x < y) \ (\ (lezero y)))" - by (simp add: lezero_def numeral_simps not_le) - -lemma natle: "(((number_of x)::nat) \ (number_of y)) = (y < x \ lezero x)" - by (auto simp add: number_of_is_id lezero_def nat_number_of_def) - -fun natfac :: "nat \ nat" - where "natfac n = (if n = 0 then 1 else n * (natfac (n - 1)))" - -lemmas compute_natarith = bitarith natnorm natsuc natadd natsub natmul nateq natless natle natfac.simps - -lemma number_eq: "(((number_of x)::'a::{number_ring, linordered_idom}) = (number_of y)) = (x = y)" - unfolding number_of_eq - apply simp - done - -lemma number_le: "(((number_of x)::'a::{number_ring, linordered_idom}) \ (number_of y)) = (x \ y)" - unfolding number_of_eq - apply simp - done - -lemma number_less: "(((number_of x)::'a::{number_ring, linordered_idom}) < (number_of y)) = (x < y)" - unfolding number_of_eq - apply simp - done - -lemma number_diff: "((number_of x)::'a::{number_ring, linordered_idom}) - number_of y = number_of (x + (- y))" - apply (subst diff_number_of_eq) - apply simp - done - -lemmas number_norm = number_of_Pls[symmetric] numeral_1_eq_1[symmetric] - -lemmas compute_numberarith = number_of_minus[symmetric] number_of_add[symmetric] number_diff number_of_mult[symmetric] number_norm number_eq number_le number_less - -lemma compute_real_of_nat_number_of: "real ((number_of v)::nat) = (if neg v then 0 else number_of v)" - by (simp only: real_of_nat_number_of number_of_is_id) - -lemma compute_nat_of_int_number_of: "nat ((number_of v)::int) = (number_of v)" - by simp - -lemmas compute_num_conversions = compute_real_of_nat_number_of compute_nat_of_int_number_of real_number_of - -lemmas zpowerarith = zpower_number_of_even - zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring] - zpower_Pls zpower_Min - -(* div, mod *) - -lemma adjust: "adjust b (q, r) = (if 0 \ r - b then (2 * q + 1, r - b) else (2 * q, r))" - by (auto simp only: adjust_def) - -lemma divmod: "divmod_int a b = (if 0\a then - if 0\b then posDivAlg a b - else if a=0 then (0, 0) - else apsnd uminus (negDivAlg (-a) (-b)) - else - if 0 program - -exception Run of string; -val run : program -> term -> term - -(* Utilities *) - -val check_freevars : int -> term -> bool -val forall_consts : (int -> bool) -> term -> bool -val closed : term -> bool -val erase_Computed : term -> term - -end - -structure AbstractMachine : ABSTRACT_MACHINE = -struct - -datatype term = Var of int | Const of int | App of term * term | Abs of term | Computed of term - -datatype pattern = PVar | PConst of int * (pattern list) - -datatype guard = Guard of term * term - -type program = unit - -exception Compile of string; - -fun erase_Computed (Computed t) = erase_Computed t - | erase_Computed (App (t1, t2)) = App (erase_Computed t1, erase_Computed t2) - | erase_Computed (Abs t) = Abs (erase_Computed t) - | erase_Computed t = t - -(*Returns true iff at most 0 .. (free-1) occur unbound. therefore - check_freevars 0 t iff t is closed*) -fun check_freevars free (Var x) = x < free - | check_freevars free (Const _) = true - | check_freevars free (App (u, v)) = check_freevars free u andalso check_freevars free v - | check_freevars free (Abs m) = check_freevars (free+1) m - | check_freevars free (Computed t) = check_freevars free t - -fun forall_consts pred (Const c) = pred c - | forall_consts pred (Var _) = true - | forall_consts pred (App (u,v)) = forall_consts pred u - andalso forall_consts pred v - | forall_consts pred (Abs m) = forall_consts pred m - | forall_consts pred (Computed t) = forall_consts pred t - -fun closed t = check_freevars 0 t - -fun compile _ = raise Compile "abstract machine stub" - -exception Run of string; - -fun run _ _ = raise Run "abstract machine stub" - -end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/Compute_Oracle/am_compiler.ML --- a/src/HOL/Matrix/Compute_Oracle/am_compiler.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,208 +0,0 @@ -(* Title: HOL/Matrix/Compute_Oracle/am_compiler.ML - Author: Steven Obua -*) - -signature COMPILING_AM = -sig - include ABSTRACT_MACHINE - - val set_compiled_rewriter : (term -> term) -> unit - val list_nth : 'a list * int -> 'a - val list_map : ('a -> 'b) -> 'a list -> 'b list -end - -structure AM_Compiler : COMPILING_AM = struct - -val list_nth = List.nth; -val list_map = map; - -open AbstractMachine; - -val compiled_rewriter = Unsynchronized.ref (NONE:(term -> term)Option.option) - -fun set_compiled_rewriter r = (compiled_rewriter := SOME r) - -type program = (term -> term) - -fun count_patternvars PVar = 1 - | count_patternvars (PConst (_, ps)) = - List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps - -fun print_rule (p, t) = - let - fun str x = string_of_int x - fun print_pattern n PVar = (n+1, "x"^(str n)) - | print_pattern n (PConst (c, [])) = (n, "c"^(str c)) - | print_pattern n (PConst (c, args)) = - let - val h = print_pattern n (PConst (c,[])) - in - print_pattern_list h args - end - and print_pattern_list r [] = r - | print_pattern_list (n, p) (t::ts) = - let - val (n, t) = print_pattern n t - in - print_pattern_list (n, "App ("^p^", "^t^")") ts - end - - val (n, pattern) = print_pattern 0 p - val pattern = - if exists_string Symbol.is_ascii_blank pattern then "(" ^ pattern ^")" - else pattern - - fun print_term d (Var x) = "Var " ^ str x - | print_term d (Const c) = "c" ^ str c - | print_term d (App (a,b)) = "App (" ^ print_term d a ^ ", " ^ print_term d b ^ ")" - | print_term d (Abs c) = "Abs (" ^ print_term (d + 1) c ^ ")" - | print_term d (Computed c) = print_term d c - - fun listvars n = if n = 0 then "x0" else "x"^(str n)^", "^(listvars (n-1)) - - val term = print_term 0 t - val term = - if n > 0 then "Closure (["^(listvars (n-1))^"], "^term^")" - else "Closure ([], "^term^")" - - in - " | weak_reduce (false, stack, "^pattern^") = Continue (false, stack, "^term^")" - end - -fun constants_of PVar = [] - | constants_of (PConst (c, ps)) = c :: maps constants_of ps - -fun constants_of_term (Var _) = [] - | constants_of_term (Abs m) = constants_of_term m - | constants_of_term (App (a,b)) = (constants_of_term a)@(constants_of_term b) - | constants_of_term (Const c) = [c] - | constants_of_term (Computed c) = constants_of_term c - -fun load_rules sname name prog = - let - val buffer = Unsynchronized.ref "" - fun write s = (buffer := (!buffer)^s) - fun writeln s = (write s; write "\n") - fun writelist [] = () - | writelist (s::ss) = (writeln s; writelist ss) - fun str i = string_of_int i - val _ = writelist [ - "structure "^name^" = struct", - "", - "datatype term = Dummy | App of term * term | Abs of term | Var of int | Const of int | Closure of term list * term"] - val constants = distinct (op =) (maps (fn (p, r) => ((constants_of p)@(constants_of_term r))) prog) - val _ = map (fn x => write (" | c"^(str x))) constants - val _ = writelist [ - "", - "datatype stack = SEmpty | SAppL of term * stack | SAppR of term * stack | SAbs of stack", - "", - "type state = bool * stack * term", - "", - "datatype loopstate = Continue of state | Stop of stack * term", - "", - "fun proj_C (Continue s) = s", - " | proj_C _ = raise Match", - "", - "fun proj_S (Stop s) = s", - " | proj_S _ = raise Match", - "", - "fun cont (Continue _) = true", - " | cont _ = false", - "", - "fun do_reduction reduce p =", - " let", - " val s = Unsynchronized.ref (Continue p)", - " val _ = while cont (!s) do (s := reduce (proj_C (!s)))", - " in", - " proj_S (!s)", - " end", - ""] - - val _ = writelist [ - "fun weak_reduce (false, stack, Closure (e, App (a, b))) = Continue (false, SAppL (Closure (e, b), stack), Closure (e, a))", - " | weak_reduce (false, SAppL (b, stack), Closure (e, Abs m)) = Continue (false, stack, Closure (b::e, m))", - " | weak_reduce (false, stack, c as Closure (e, Abs m)) = Continue (true, stack, c)", - " | weak_reduce (false, stack, Closure (e, Var n)) = Continue (false, stack, case "^sname^".list_nth (e, n) of Dummy => Var n | r => r)", - " | weak_reduce (false, stack, Closure (e, c)) = Continue (false, stack, c)"] - val _ = writelist (map print_rule prog) - val _ = writelist [ - " | weak_reduce (false, stack, clos) = Continue (true, stack, clos)", - " | weak_reduce (true, SAppR (a, stack), b) = Continue (false, stack, App (a,b))", - " | weak_reduce (true, s as (SAppL (b, stack)), a) = Continue (false, SAppR (a, stack), b)", - " | weak_reduce (true, stack, c) = Stop (stack, c)", - "", - "fun strong_reduce (false, stack, Closure (e, Abs m)) =", - " let", - " val (stack', wnf) = do_reduction weak_reduce (false, SEmpty, Closure (Dummy::e, m))", - " in", - " case stack' of", - " SEmpty => Continue (false, SAbs stack, wnf)", - " | _ => raise ("^sname^".Run \"internal error in strong: weak failed\")", - " end", - " | strong_reduce (false, stack, clos as (App (u, v))) = Continue (false, SAppL (v, stack), u)", - " | strong_reduce (false, stack, clos) = Continue (true, stack, clos)", - " | strong_reduce (true, SAbs stack, m) = Continue (false, stack, Abs m)", - " | strong_reduce (true, SAppL (b, stack), a) = Continue (false, SAppR (a, stack), b)", - " | strong_reduce (true, SAppR (a, stack), b) = Continue (true, stack, App (a, b))", - " | strong_reduce (true, stack, clos) = Stop (stack, clos)", - ""] - - val ic = "(case c of "^(implode (map (fn c => (str c)^" => c"^(str c)^" | ") constants))^" _ => Const c)" - val _ = writelist [ - "fun importTerm ("^sname^".Var x) = Var x", - " | importTerm ("^sname^".Const c) = "^ic, - " | importTerm ("^sname^".App (a, b)) = App (importTerm a, importTerm b)", - " | importTerm ("^sname^".Abs m) = Abs (importTerm m)", - ""] - - fun ec c = " | exportTerm c"^(str c)^" = "^sname^".Const "^(str c) - val _ = writelist [ - "fun exportTerm (Var x) = "^sname^".Var x", - " | exportTerm (Const c) = "^sname^".Const c", - " | exportTerm (App (a,b)) = "^sname^".App (exportTerm a, exportTerm b)", - " | exportTerm (Abs m) = "^sname^".Abs (exportTerm m)", - " | exportTerm (Closure (closlist, clos)) = raise ("^sname^".Run \"internal error, cannot export Closure\")", - " | exportTerm Dummy = raise ("^sname^".Run \"internal error, cannot export Dummy\")"] - val _ = writelist (map ec constants) - - val _ = writelist [ - "", - "fun rewrite t = ", - " let", - " val (stack, wnf) = do_reduction weak_reduce (false, SEmpty, Closure ([], importTerm t))", - " in", - " case stack of ", - " SEmpty => (case do_reduction strong_reduce (false, SEmpty, wnf) of", - " (SEmpty, snf) => exportTerm snf", - " | _ => raise ("^sname^".Run \"internal error in rewrite: strong failed\"))", - " | _ => (raise ("^sname^".Run \"internal error in rewrite: weak failed\"))", - " end", - "", - "val _ = "^sname^".set_compiled_rewriter rewrite", - "", - "end;"] - - in - compiled_rewriter := NONE; - use_text ML_Env.local_context (1, "") false (!buffer); - case !compiled_rewriter of - NONE => raise (Compile "cannot communicate with compiled function") - | SOME r => (compiled_rewriter := NONE; r) - end - -fun compile eqs = - let - val _ = if exists (fn (a,_,_) => not (null a)) eqs then raise Compile ("cannot deal with guards") else () - val eqs = map (fn (_,b,c) => (b,c)) eqs - fun check (p, r) = if check_freevars (count_patternvars p) r then () else raise Compile ("unbound variables in rule") - val _ = map (fn (p, r) => - (check (p, r); - case p of PVar => raise (Compile "pattern is just a variable") | _ => ())) eqs - in - load_rules "AM_Compiler" "AM_compiled_code" eqs - end - -fun run prog t = prog t - -end - diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/Compute_Oracle/am_ghc.ML --- a/src/HOL/Matrix/Compute_Oracle/am_ghc.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,324 +0,0 @@ -(* Title: HOL/Matrix/Compute_Oracle/am_ghc.ML - Author: Steven Obua -*) - -structure AM_GHC : ABSTRACT_MACHINE = -struct - -open AbstractMachine; - -type program = string * string * (int Inttab.table) - -fun count_patternvars PVar = 1 - | count_patternvars (PConst (_, ps)) = - List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps - -fun update_arity arity code a = - (case Inttab.lookup arity code of - NONE => Inttab.update_new (code, a) arity - | SOME (a': int) => if a > a' then Inttab.update (code, a) arity else arity) - -(* We have to find out the maximal arity of each constant *) -fun collect_pattern_arity PVar arity = arity - | collect_pattern_arity (PConst (c, args)) arity = fold collect_pattern_arity args (update_arity arity c (length args)) - -local -fun collect applevel (Var _) arity = arity - | collect applevel (Const c) arity = update_arity arity c applevel - | collect applevel (Abs m) arity = collect 0 m arity - | collect applevel (App (a,b)) arity = collect 0 b (collect (applevel + 1) a arity) -in -fun collect_term_arity t arity = collect 0 t arity -end - -fun nlift level n (Var m) = if m < level then Var m else Var (m+n) - | nlift level n (Const c) = Const c - | nlift level n (App (a,b)) = App (nlift level n a, nlift level n b) - | nlift level n (Abs b) = Abs (nlift (level+1) n b) - -fun rep n x = if n = 0 then [] else x::(rep (n-1) x) - -fun adjust_rules rules = - let - val arity = fold (fn (p, t) => fn arity => collect_term_arity t (collect_pattern_arity p arity)) rules Inttab.empty - fun arity_of c = the (Inttab.lookup arity c) - fun adjust_pattern PVar = PVar - | adjust_pattern (C as PConst (c, args)) = if (length args <> arity_of c) then raise Compile ("Constant inside pattern must have maximal arity") else C - fun adjust_rule (PVar, _) = raise Compile ("pattern may not be a variable") - | adjust_rule (rule as (p as PConst (c, args),t)) = - let - val _ = if not (check_freevars (count_patternvars p) t) then raise Compile ("unbound variables on right hand side") else () - val args = map adjust_pattern args - val len = length args - val arity = arity_of c - fun lift level n (Var m) = if m < level then Var m else Var (m+n) - | lift level n (Const c) = Const c - | lift level n (App (a,b)) = App (lift level n a, lift level n b) - | lift level n (Abs b) = Abs (lift (level+1) n b) - val lift = lift 0 - fun adjust_term n t = if n=0 then t else adjust_term (n-1) (App (t, Var (n-1))) - in - if len = arity then - rule - else if arity >= len then - (PConst (c, args @ (rep (arity-len) PVar)), adjust_term (arity-len) (lift (arity-len) t)) - else (raise Compile "internal error in adjust_rule") - end - in - (arity, map adjust_rule rules) - end - -fun print_term arity_of n = -let - fun str x = string_of_int x - fun protect_blank s = if exists_string Symbol.is_ascii_blank s then "(" ^ s ^")" else s - - fun print_apps d f [] = f - | print_apps d f (a::args) = print_apps d ("app "^(protect_blank f)^" "^(protect_blank (print_term d a))) args - and print_call d (App (a, b)) args = print_call d a (b::args) - | print_call d (Const c) args = - (case arity_of c of - NONE => print_apps d ("Const "^(str c)) args - | SOME a => - let - val len = length args - in - if a <= len then - let - val s = "c"^(str c)^(implode (map (fn t => " "^(protect_blank (print_term d t))) (List.take (args, a)))) - in - print_apps d s (List.drop (args, a)) - end - else - let - fun mk_apps n t = if n = 0 then t else mk_apps (n-1) (App (t, Var (n-1))) - fun mk_lambdas n t = if n = 0 then t else mk_lambdas (n-1) (Abs t) - fun append_args [] t = t - | append_args (c::cs) t = append_args cs (App (t, c)) - in - print_term d (mk_lambdas (a-len) (mk_apps (a-len) (nlift 0 (a-len) (append_args args (Const c))))) - end - end) - | print_call d t args = print_apps d (print_term d t) args - and print_term d (Var x) = if x < d then "b"^(str (d-x-1)) else "x"^(str (n-(x-d)-1)) - | print_term d (Abs c) = "Abs (\\b"^(str d)^" -> "^(print_term (d + 1) c)^")" - | print_term d t = print_call d t [] -in - print_term 0 -end - -fun print_rule arity_of (p, t) = - let - fun str x = string_of_int x - fun print_pattern top n PVar = (n+1, "x"^(str n)) - | print_pattern top n (PConst (c, [])) = (n, (if top then "c" else "C")^(str c)) - | print_pattern top n (PConst (c, args)) = - let - val (n,s) = print_pattern_list (n, (if top then "c" else "C")^(str c)) args - in - (n, if top then s else "("^s^")") - end - and print_pattern_list r [] = r - | print_pattern_list (n, p) (t::ts) = - let - val (n, t) = print_pattern false n t - in - print_pattern_list (n, p^" "^t) ts - end - val (n, pattern) = print_pattern true 0 p - in - pattern^" = "^(print_term arity_of n t) - end - -fun group_rules rules = - let - fun add_rule (r as (PConst (c,_), _)) groups = - let - val rs = (case Inttab.lookup groups c of NONE => [] | SOME rs => rs) - in - Inttab.update (c, r::rs) groups - end - | add_rule _ _ = raise Compile "internal error group_rules" - in - fold_rev add_rule rules Inttab.empty - end - -fun haskell_prog name rules = - let - val buffer = Unsynchronized.ref "" - fun write s = (buffer := (!buffer)^s) - fun writeln s = (write s; write "\n") - fun writelist [] = () - | writelist (s::ss) = (writeln s; writelist ss) - fun str i = string_of_int i - val (arity, rules) = adjust_rules rules - val rules = group_rules rules - val constants = Inttab.keys arity - fun arity_of c = Inttab.lookup arity c - fun rep_str s n = implode (rep n s) - fun indexed s n = s^(str n) - fun section n = if n = 0 then [] else (section (n-1))@[n-1] - fun make_show c = - let - val args = section (the (arity_of c)) - in - " show ("^(indexed "C" c)^(implode (map (indexed " a") args))^") = " - ^"\""^(indexed "C" c)^"\""^(implode (map (fn a => "++(show "^(indexed "a" a)^")") args)) - end - fun default_case c = - let - val args = implode (map (indexed " x") (section (the (arity_of c)))) - in - (indexed "c" c)^args^" = "^(indexed "C" c)^args - end - val _ = writelist [ - "module "^name^" where", - "", - "data Term = Const Integer | App Term Term | Abs (Term -> Term)", - " "^(implode (map (fn c => " | C"^(str c)^(rep_str " Term" (the (arity_of c)))) constants)), - "", - "instance Show Term where"] - val _ = writelist (map make_show constants) - val _ = writelist [ - " show (Const c) = \"c\"++(show c)", - " show (App a b) = \"A\"++(show a)++(show b)", - " show (Abs _) = \"L\"", - ""] - val _ = writelist [ - "app (Abs a) b = a b", - "app a b = App a b", - "", - "calc s c = writeFile s (show c)", - ""] - fun list_group c = (writelist (case Inttab.lookup rules c of - NONE => [default_case c, ""] - | SOME (rs as ((PConst (_, []), _)::rs')) => - if not (null rs') then raise Compile "multiple declaration of constant" - else (map (print_rule arity_of) rs) @ [""] - | SOME rs => (map (print_rule arity_of) rs) @ [default_case c, ""])) - val _ = map list_group constants - in - (arity, !buffer) - end - -val guid_counter = Unsynchronized.ref 0 -fun get_guid () = - let - val c = !guid_counter - val _ = guid_counter := !guid_counter + 1 - in - string_of_int (Time.toMicroseconds (Time.now ())) ^ string_of_int c - end - -fun tmp_file s = Path.implode (Path.expand (File.tmp_path (Path.basic s))); - -fun writeTextFile name s = File.write (Path.explode name) s - -fun fileExists name = ((OS.FileSys.fileSize name; true) handle OS.SysErr _ => false) - -fun compile eqs = - let - val _ = if exists (fn (a,_,_) => not (null a)) eqs then raise Compile ("cannot deal with guards") else () - val eqs = map (fn (_,b,c) => (b,c)) eqs - val guid = get_guid () - val module = "AMGHC_Prog_"^guid - val (arity, source) = haskell_prog module eqs - val module_file = tmp_file (module^".hs") - val object_file = tmp_file (module^".o") - val _ = writeTextFile module_file source - val _ = Isabelle_System.bash ("exec \"$ISABELLE_GHC\" -c " ^ module_file) - val _ = - if not (fileExists object_file) then - raise Compile ("Failure compiling haskell code (ISABELLE_GHC='" ^ getenv "ISABELLE_GHC" ^ "')") - else () - in - (guid, module_file, arity) - end - -fun readResultFile name = File.read (Path.explode name) - -fun parse_result arity_of result = - let - val result = String.explode result - fun shift NONE x = SOME x - | shift (SOME y) x = SOME (y*10 + x) - fun parse_int' x (#"0"::rest) = parse_int' (shift x 0) rest - | parse_int' x (#"1"::rest) = parse_int' (shift x 1) rest - | parse_int' x (#"2"::rest) = parse_int' (shift x 2) rest - | parse_int' x (#"3"::rest) = parse_int' (shift x 3) rest - | parse_int' x (#"4"::rest) = parse_int' (shift x 4) rest - | parse_int' x (#"5"::rest) = parse_int' (shift x 5) rest - | parse_int' x (#"6"::rest) = parse_int' (shift x 6) rest - | parse_int' x (#"7"::rest) = parse_int' (shift x 7) rest - | parse_int' x (#"8"::rest) = parse_int' (shift x 8) rest - | parse_int' x (#"9"::rest) = parse_int' (shift x 9) rest - | parse_int' x rest = (x, rest) - fun parse_int rest = parse_int' NONE rest - - fun parse (#"C"::rest) = - (case parse_int rest of - (SOME c, rest) => - let - val (args, rest) = parse_list (the (arity_of c)) rest - fun app_args [] t = t - | app_args (x::xs) t = app_args xs (App (t, x)) - in - (app_args args (Const c), rest) - end - | (NONE, _) => raise Run "parse C") - | parse (#"c"::rest) = - (case parse_int rest of - (SOME c, rest) => (Const c, rest) - | _ => raise Run "parse c") - | parse (#"A"::rest) = - let - val (a, rest) = parse rest - val (b, rest) = parse rest - in - (App (a,b), rest) - end - | parse (#"L"::_) = raise Run "there may be no abstraction in the result" - | parse _ = raise Run "invalid result" - and parse_list n rest = - if n = 0 then - ([], rest) - else - let - val (x, rest) = parse rest - val (xs, rest) = parse_list (n-1) rest - in - (x::xs, rest) - end - val (parsed, rest) = parse result - fun is_blank (#" "::rest) = is_blank rest - | is_blank (#"\n"::rest) = is_blank rest - | is_blank [] = true - | is_blank _ = false - in - if is_blank rest then parsed else raise Run "non-blank suffix in result file" - end - -fun run (guid, module_file, arity) t = - let - val _ = if check_freevars 0 t then () else raise Run ("can only compute closed terms") - fun arity_of c = Inttab.lookup arity c - val callguid = get_guid() - val module = "AMGHC_Prog_"^guid - val call = module^"_Call_"^callguid - val result_file = tmp_file (module^"_Result_"^callguid^".txt") - val call_file = tmp_file (call^".hs") - val term = print_term arity_of 0 t - val call_source = "module "^call^" where\n\nimport "^module^"\n\ncall = "^module^".calc \""^result_file^"\" ("^term^")" - val _ = writeTextFile call_file call_source - val _ = Isabelle_System.bash ("exec \"$ISABELLE_GHC\" -e \""^call^".call\" "^module_file^" "^call_file) - val result = readResultFile result_file handle IO.Io _ => - raise Run ("Failure running haskell compiler (ISABELLE_GHC='" ^ getenv "ISABELLE_GHC" ^ "')") - val t' = parse_result arity_of result - val _ = OS.FileSys.remove call_file - val _ = OS.FileSys.remove result_file - in - t' - end - -end - diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/Compute_Oracle/am_interpreter.ML --- a/src/HOL/Matrix/Compute_Oracle/am_interpreter.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,211 +0,0 @@ -(* Title: HOL/Matrix/Compute_Oracle/am_interpreter.ML - Author: Steven Obua -*) - -signature AM_BARRAS = -sig - include ABSTRACT_MACHINE - val max_reductions : int option Unsynchronized.ref -end - -structure AM_Interpreter : AM_BARRAS = struct - -open AbstractMachine; - -datatype closure = CDummy | CVar of int | CConst of int - | CApp of closure * closure | CAbs of closure - | Closure of (closure list) * closure - -structure prog_struct = Table(type key = int*int val ord = prod_ord int_ord int_ord); - -datatype program = Program of ((pattern * closure * (closure*closure) list) list) prog_struct.table - -datatype stack = SEmpty | SAppL of closure * stack | SAppR of closure * stack | SAbs of stack - -fun clos_of_term (Var x) = CVar x - | clos_of_term (Const c) = CConst c - | clos_of_term (App (u, v)) = CApp (clos_of_term u, clos_of_term v) - | clos_of_term (Abs u) = CAbs (clos_of_term u) - | clos_of_term (Computed t) = clos_of_term t - -fun term_of_clos (CVar x) = Var x - | term_of_clos (CConst c) = Const c - | term_of_clos (CApp (u, v)) = App (term_of_clos u, term_of_clos v) - | term_of_clos (CAbs u) = Abs (term_of_clos u) - | term_of_clos (Closure _) = raise (Run "internal error: closure in normalized term found") - | term_of_clos CDummy = raise (Run "internal error: dummy in normalized term found") - -fun resolve_closure closures (CVar x) = (case nth closures x of CDummy => CVar x | r => r) - | resolve_closure closures (CConst c) = CConst c - | resolve_closure closures (CApp (u, v)) = CApp (resolve_closure closures u, resolve_closure closures v) - | resolve_closure closures (CAbs u) = CAbs (resolve_closure (CDummy::closures) u) - | resolve_closure closures (CDummy) = raise (Run "internal error: resolve_closure applied to CDummy") - | resolve_closure closures (Closure (e, u)) = resolve_closure e u - -fun resolve_closure' c = resolve_closure [] c - -fun resolve_stack tm SEmpty = tm - | resolve_stack tm (SAppL (c, s)) = resolve_stack (CApp (tm, resolve_closure' c)) s - | resolve_stack tm (SAppR (c, s)) = resolve_stack (CApp (resolve_closure' c, tm)) s - | resolve_stack tm (SAbs s) = resolve_stack (CAbs tm) s - -fun resolve (stack, closure) = - let - val _ = writeln "start resolving" - val t = resolve_stack (resolve_closure' closure) stack - val _ = writeln "finished resolving" - in - t - end - -fun strip_closure args (CApp (a,b)) = strip_closure (b::args) a - | strip_closure args x = (x, args) - -fun len_head_of_closure n (CApp (a, _)) = len_head_of_closure (n+1) a - | len_head_of_closure n x = (n, x) - - -(* earlier occurrence of PVar corresponds to higher de Bruijn index *) -fun pattern_match args PVar clos = SOME (clos::args) - | pattern_match args (PConst (c, patterns)) clos = - let - val (f, closargs) = strip_closure [] clos - in - case f of - CConst d => - if c = d then - pattern_match_list args patterns closargs - else - NONE - | _ => NONE - end -and pattern_match_list args [] [] = SOME args - | pattern_match_list args (p::ps) (c::cs) = - (case pattern_match args p c of - NONE => NONE - | SOME args => pattern_match_list args ps cs) - | pattern_match_list _ _ _ = NONE - -fun count_patternvars PVar = 1 - | count_patternvars (PConst (_, ps)) = List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps - -fun pattern_key (PConst (c, ps)) = (c, length ps) - | pattern_key _ = raise (Compile "pattern reduces to variable") - -(*Returns true iff at most 0 .. (free-1) occur unbound. therefore - check_freevars 0 t iff t is closed*) -fun check_freevars free (Var x) = x < free - | check_freevars free (Const _) = true - | check_freevars free (App (u, v)) = check_freevars free u andalso check_freevars free v - | check_freevars free (Abs m) = check_freevars (free+1) m - | check_freevars free (Computed t) = check_freevars free t - -fun compile eqs = - let - fun check p r = if check_freevars p r then () else raise Compile ("unbound variables in rule") - fun check_guard p (Guard (a,b)) = (check p a; check p b) - fun clos_of_guard (Guard (a,b)) = (clos_of_term a, clos_of_term b) - val eqs = map (fn (guards, p, r) => let val pcount = count_patternvars p val _ = map (check_guard pcount) (guards) val _ = check pcount r in - (pattern_key p, (p, clos_of_term r, map clos_of_guard guards)) end) eqs - fun merge (k, a) table = prog_struct.update (k, case prog_struct.lookup table k of NONE => [a] | SOME l => a::l) table - val p = fold merge eqs prog_struct.empty - in - Program p - end - - -type state = bool * program * stack * closure - -datatype loopstate = Continue of state | Stop of stack * closure - -fun proj_C (Continue s) = s - | proj_C _ = raise Match - -exception InterruptedExecution of stack * closure - -fun proj_S (Stop s) = s - | proj_S (Continue (_,_,s,c)) = (s,c) - -fun cont (Continue _) = true - | cont _ = false - -val max_reductions = Unsynchronized.ref (NONE : int option) - -fun do_reduction reduce p = - let - val s = Unsynchronized.ref (Continue p) - val counter = Unsynchronized.ref 0 - val _ = case !max_reductions of - NONE => while cont (!s) do (s := reduce (proj_C (!s))) - | SOME m => while cont (!s) andalso (!counter < m) do (s := reduce (proj_C (!s)); counter := (!counter) + 1) - in - case !max_reductions of - SOME m => if !counter >= m then raise InterruptedExecution (proj_S (!s)) else proj_S (!s) - | NONE => proj_S (!s) - end - -fun match_rules prog n [] clos = NONE - | match_rules prog n ((p,eq,guards)::rs) clos = - case pattern_match [] p clos of - NONE => match_rules prog (n+1) rs clos - | SOME args => if forall (guard_checks prog args) guards then SOME (Closure (args, eq)) else match_rules prog (n+1) rs clos -and guard_checks prog args (a,b) = (simp prog (Closure (args, a)) = simp prog (Closure (args, b))) -and match_closure (p as (Program prog)) clos = - case len_head_of_closure 0 clos of - (len, CConst c) => - (case prog_struct.lookup prog (c, len) of - NONE => NONE - | SOME rules => match_rules p 0 rules clos) - | _ => NONE - -and weak_reduce (false, prog, stack, Closure (e, CApp (a, b))) = Continue (false, prog, SAppL (Closure (e, b), stack), Closure (e, a)) - | weak_reduce (false, prog, SAppL (b, stack), Closure (e, CAbs m)) = Continue (false, prog, stack, Closure (b::e, m)) - | weak_reduce (false, prog, stack, Closure (e, CVar n)) = Continue (false, prog, stack, case nth e n of CDummy => CVar n | r => r) - | weak_reduce (false, prog, stack, Closure (_, c as CConst _)) = Continue (false, prog, stack, c) - | weak_reduce (false, prog, stack, clos) = - (case match_closure prog clos of - NONE => Continue (true, prog, stack, clos) - | SOME r => Continue (false, prog, stack, r)) - | weak_reduce (true, prog, SAppR (a, stack), b) = Continue (false, prog, stack, CApp (a,b)) - | weak_reduce (true, prog, SAppL (b, stack), a) = Continue (false, prog, SAppR (a, stack), b) - | weak_reduce (true, prog, stack, c) = Stop (stack, c) - -and strong_reduce (false, prog, stack, Closure (e, CAbs m)) = - (let - val (stack', wnf) = do_reduction weak_reduce (false, prog, SEmpty, Closure (CDummy::e, m)) - in - case stack' of - SEmpty => Continue (false, prog, SAbs stack, wnf) - | _ => raise (Run "internal error in strong: weak failed") - end handle InterruptedExecution state => raise InterruptedExecution (stack, resolve state)) - | strong_reduce (false, prog, stack, CApp (u, v)) = Continue (false, prog, SAppL (v, stack), u) - | strong_reduce (false, prog, stack, clos) = Continue (true, prog, stack, clos) - | strong_reduce (true, prog, SAbs stack, m) = Continue (false, prog, stack, CAbs m) - | strong_reduce (true, prog, SAppL (b, stack), a) = Continue (false, prog, SAppR (a, stack), b) - | strong_reduce (true, prog, SAppR (a, stack), b) = Continue (true, prog, stack, CApp (a, b)) - | strong_reduce (true, prog, stack, clos) = Stop (stack, clos) - -and simp prog t = - (let - val (stack, wnf) = do_reduction weak_reduce (false, prog, SEmpty, t) - in - case stack of - SEmpty => (case do_reduction strong_reduce (false, prog, SEmpty, wnf) of - (SEmpty, snf) => snf - | _ => raise (Run "internal error in run: strong failed")) - | _ => raise (Run "internal error in run: weak failed") - end handle InterruptedExecution state => resolve state) - - -fun run prog t = - (let - val (stack, wnf) = do_reduction weak_reduce (false, prog, SEmpty, Closure ([], clos_of_term t)) - in - case stack of - SEmpty => (case do_reduction strong_reduce (false, prog, SEmpty, wnf) of - (SEmpty, snf) => term_of_clos snf - | _ => raise (Run "internal error in run: strong failed")) - | _ => raise (Run "internal error in run: weak failed") - end handle InterruptedExecution state => term_of_clos (resolve state)) - -end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/Compute_Oracle/am_sml.ML --- a/src/HOL/Matrix/Compute_Oracle/am_sml.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,517 +0,0 @@ -(* Title: HOL/Matrix/Compute_Oracle/am_sml.ML - Author: Steven Obua - -TODO: "parameterless rewrite cannot be used in pattern": In a lot of -cases it CAN be used, and these cases should be handled -properly; right now, all cases raise an exception. -*) - -signature AM_SML = -sig - include ABSTRACT_MACHINE - val save_result : (string * term) -> unit - val set_compiled_rewriter : (term -> term) -> unit - val list_nth : 'a list * int -> 'a - val dump_output : (string option) Unsynchronized.ref -end - -structure AM_SML : AM_SML = struct - -open AbstractMachine; - -val dump_output = Unsynchronized.ref (NONE: string option) - -type program = term Inttab.table * (term -> term) - -val saved_result = Unsynchronized.ref (NONE:(string*term)option) - -fun save_result r = (saved_result := SOME r) - -val list_nth = List.nth - -val compiled_rewriter = Unsynchronized.ref (NONE:(term -> term)Option.option) - -fun set_compiled_rewriter r = (compiled_rewriter := SOME r) - -fun count_patternvars PVar = 1 - | count_patternvars (PConst (_, ps)) = - List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps - -fun update_arity arity code a = - (case Inttab.lookup arity code of - NONE => Inttab.update_new (code, a) arity - | SOME (a': int) => if a > a' then Inttab.update (code, a) arity else arity) - -(* We have to find out the maximal arity of each constant *) -fun collect_pattern_arity PVar arity = arity - | collect_pattern_arity (PConst (c, args)) arity = fold collect_pattern_arity args (update_arity arity c (length args)) - -(* We also need to find out the maximal toplevel arity of each function constant *) -fun collect_pattern_toplevel_arity PVar arity = raise Compile "internal error: collect_pattern_toplevel_arity" - | collect_pattern_toplevel_arity (PConst (c, args)) arity = update_arity arity c (length args) - -local -fun collect applevel (Var _) arity = arity - | collect applevel (Const c) arity = update_arity arity c applevel - | collect applevel (Abs m) arity = collect 0 m arity - | collect applevel (App (a,b)) arity = collect 0 b (collect (applevel + 1) a arity) -in -fun collect_term_arity t arity = collect 0 t arity -end - -fun collect_guard_arity (Guard (a,b)) arity = collect_term_arity b (collect_term_arity a arity) - - -fun rep n x = if n < 0 then raise Compile "internal error: rep" else if n = 0 then [] else x::(rep (n-1) x) - -fun beta (Const c) = Const c - | beta (Var i) = Var i - | beta (App (Abs m, b)) = beta (unlift 0 (subst 0 m (lift 0 b))) - | beta (App (a, b)) = - (case beta a of - Abs m => beta (App (Abs m, b)) - | a => App (a, beta b)) - | beta (Abs m) = Abs (beta m) - | beta (Computed t) = Computed t -and subst x (Const c) t = Const c - | subst x (Var i) t = if i = x then t else Var i - | subst x (App (a,b)) t = App (subst x a t, subst x b t) - | subst x (Abs m) t = Abs (subst (x+1) m (lift 0 t)) -and lift level (Const c) = Const c - | lift level (App (a,b)) = App (lift level a, lift level b) - | lift level (Var i) = if i < level then Var i else Var (i+1) - | lift level (Abs m) = Abs (lift (level + 1) m) -and unlift level (Const c) = Const c - | unlift level (App (a, b)) = App (unlift level a, unlift level b) - | unlift level (Abs m) = Abs (unlift (level+1) m) - | unlift level (Var i) = if i < level then Var i else Var (i-1) - -fun nlift level n (Var m) = if m < level then Var m else Var (m+n) - | nlift level n (Const c) = Const c - | nlift level n (App (a,b)) = App (nlift level n a, nlift level n b) - | nlift level n (Abs b) = Abs (nlift (level+1) n b) - -fun subst_const (c, t) (Const c') = if c = c' then t else Const c' - | subst_const _ (Var i) = Var i - | subst_const ct (App (a, b)) = App (subst_const ct a, subst_const ct b) - | subst_const ct (Abs m) = Abs (subst_const ct m) - -(* Remove all rules that are just parameterless rewrites. This is necessary because SML does not allow functions with no parameters. *) -fun inline_rules rules = - let - fun term_contains_const c (App (a, b)) = term_contains_const c a orelse term_contains_const c b - | term_contains_const c (Abs m) = term_contains_const c m - | term_contains_const c (Var _) = false - | term_contains_const c (Const c') = (c = c') - fun find_rewrite [] = NONE - | find_rewrite ((prems, PConst (c, []), r) :: _) = - if check_freevars 0 r then - if term_contains_const c r then - raise Compile "parameterless rewrite is caught in cycle" - else if not (null prems) then - raise Compile "parameterless rewrite may not be guarded" - else - SOME (c, r) - else raise Compile "unbound variable on right hand side or guards of rule" - | find_rewrite (_ :: rules) = find_rewrite rules - fun remove_rewrite _ [] = [] - | remove_rewrite (cr as (c, r)) ((rule as (prems', PConst (c', args), r')) :: rules) = - if c = c' then - if null args andalso r = r' andalso null prems' then remove_rewrite cr rules - else raise Compile "incompatible parameterless rewrites found" - else - rule :: remove_rewrite cr rules - | remove_rewrite cr (r :: rs) = r :: remove_rewrite cr rs - fun pattern_contains_const c (PConst (c', args)) = c = c' orelse exists (pattern_contains_const c) args - | pattern_contains_const c (PVar) = false - fun inline_rewrite (ct as (c, _)) (prems, p, r) = - if pattern_contains_const c p then - raise Compile "parameterless rewrite cannot be used in pattern" - else (map (fn (Guard (a, b)) => Guard (subst_const ct a, subst_const ct b)) prems, p, subst_const ct r) - fun inline inlined rules = - case find_rewrite rules of - NONE => (Inttab.make inlined, rules) - | SOME ct => - let - val rules = map (inline_rewrite ct) (remove_rewrite ct rules) - val inlined = ct :: (map o apsnd) (subst_const ct) inlined - in inline inlined rules end - in - inline [] rules - end - - -(* - Calculate the arity, the toplevel_arity, and adjust rules so that all toplevel pattern constants have maximal arity. - Also beta reduce the adjusted right hand side of a rule. -*) -fun adjust_rules rules = - let - val arity = fold (fn (prems, p, t) => fn arity => fold collect_guard_arity prems (collect_term_arity t (collect_pattern_arity p arity))) rules Inttab.empty - val toplevel_arity = fold (fn (_, p, _) => fn arity => collect_pattern_toplevel_arity p arity) rules Inttab.empty - fun arity_of c = the (Inttab.lookup arity c) - fun test_pattern PVar = () - | test_pattern (PConst (c, args)) = if (length args <> arity_of c) then raise Compile ("Constant inside pattern must have maximal arity") else (map test_pattern args; ()) - fun adjust_rule (_, PVar, _) = raise Compile ("pattern may not be a variable") - | adjust_rule (_, PConst (_, []), _) = raise Compile ("cannot deal with rewrites that take no parameters") - | adjust_rule (rule as (prems, p as PConst (c, args),t)) = - let - val patternvars_counted = count_patternvars p - fun check_fv t = check_freevars patternvars_counted t - val _ = if not (check_fv t) then raise Compile ("unbound variables on right hand side of rule") else () - val _ = if not (forall (fn (Guard (a,b)) => check_fv a andalso check_fv b) prems) then raise Compile ("unbound variables in guards") else () - val _ = map test_pattern args - val len = length args - val arity = arity_of c - val lift = nlift 0 - fun addapps_tm n t = if n=0 then t else addapps_tm (n-1) (App (t, Var (n-1))) - fun adjust_term n t = addapps_tm n (lift n t) - fun adjust_guard n (Guard (a,b)) = Guard (lift n a, lift n b) - in - if len = arity then - rule - else if arity >= len then - (map (adjust_guard (arity-len)) prems, PConst (c, args @ (rep (arity-len) PVar)), adjust_term (arity-len) t) - else (raise Compile "internal error in adjust_rule") - end - fun beta_rule (prems, p, t) = ((prems, p, beta t) handle Match => raise Compile "beta_rule") - in - (arity, toplevel_arity, map (beta_rule o adjust_rule) rules) - end - -fun print_term module arity_of toplevel_arity_of pattern_var_count pattern_lazy_var_count = -let - fun str x = string_of_int x - fun protect_blank s = if exists_string Symbol.is_ascii_blank s then "(" ^ s ^")" else s - val module_prefix = (case module of NONE => "" | SOME s => s^".") - fun print_apps d f [] = f - | print_apps d f (a::args) = print_apps d (module_prefix^"app "^(protect_blank f)^" "^(protect_blank (print_term d a))) args - and print_call d (App (a, b)) args = print_call d a (b::args) - | print_call d (Const c) args = - (case arity_of c of - NONE => print_apps d (module_prefix^"Const "^(str c)) args - | SOME 0 => module_prefix^"C"^(str c) - | SOME a => - let - val len = length args - in - if a <= len then - let - val strict_a = (case toplevel_arity_of c of SOME sa => sa | NONE => a) - val _ = if strict_a > a then raise Compile "strict" else () - val s = module_prefix^"c"^(str c)^(implode (map (fn t => " "^(protect_blank (print_term d t))) (List.take (args, strict_a)))) - val s = s^(implode (map (fn t => " (fn () => "^print_term d t^")") (List.drop (List.take (args, a), strict_a)))) - in - print_apps d s (List.drop (args, a)) - end - else - let - fun mk_apps n t = if n = 0 then t else mk_apps (n-1) (App (t, Var (n - 1))) - fun mk_lambdas n t = if n = 0 then t else mk_lambdas (n-1) (Abs t) - fun append_args [] t = t - | append_args (c::cs) t = append_args cs (App (t, c)) - in - print_term d (mk_lambdas (a-len) (mk_apps (a-len) (nlift 0 (a-len) (append_args args (Const c))))) - end - end) - | print_call d t args = print_apps d (print_term d t) args - and print_term d (Var x) = - if x < d then - "b"^(str (d-x-1)) - else - let - val n = pattern_var_count - (x-d) - 1 - val x = "x"^(str n) - in - if n < pattern_var_count - pattern_lazy_var_count then - x - else - "("^x^" ())" - end - | print_term d (Abs c) = module_prefix^"Abs (fn b"^(str d)^" => "^(print_term (d + 1) c)^")" - | print_term d t = print_call d t [] -in - print_term 0 -end - -fun section n = if n = 0 then [] else (section (n-1))@[n-1] - -fun print_rule gnum arity_of toplevel_arity_of (guards, p, t) = - let - fun str x = string_of_int x - fun print_pattern top n PVar = (n+1, "x"^(str n)) - | print_pattern top n (PConst (c, [])) = (n, (if top then "c" else "C")^(str c)^(if top andalso gnum > 0 then "_"^(str gnum) else "")) - | print_pattern top n (PConst (c, args)) = - let - val f = (if top then "c" else "C")^(str c)^(if top andalso gnum > 0 then "_"^(str gnum) else "") - val (n, s) = print_pattern_list 0 top (n, f) args - in - (n, s) - end - and print_pattern_list' counter top (n,p) [] = if top then (n,p) else (n,p^")") - | print_pattern_list' counter top (n, p) (t::ts) = - let - val (n, t) = print_pattern false n t - in - print_pattern_list' (counter + 1) top (n, if top then p^" (a"^(str counter)^" as ("^t^"))" else p^", "^t) ts - end - and print_pattern_list counter top (n, p) (t::ts) = - let - val (n, t) = print_pattern false n t - in - print_pattern_list' (counter + 1) top (n, if top then p^" (a"^(str counter)^" as ("^t^"))" else p^" ("^t) ts - end - val c = (case p of PConst (c, _) => c | _ => raise Match) - val (n, pattern) = print_pattern true 0 p - val lazy_vars = the (arity_of c) - the (toplevel_arity_of c) - fun print_tm tm = print_term NONE arity_of toplevel_arity_of n lazy_vars tm - fun print_guard (Guard (a,b)) = "term_eq ("^(print_tm a)^") ("^(print_tm b)^")" - val else_branch = "c"^(str c)^"_"^(str (gnum+1))^(implode (map (fn i => " a"^(str i)) (section (the (arity_of c))))) - fun print_guards t [] = print_tm t - | print_guards t (g::gs) = "if ("^(print_guard g)^")"^(implode (map (fn g => " andalso ("^(print_guard g)^")") gs))^" then ("^(print_tm t)^") else "^else_branch - in - (if null guards then gnum else gnum+1, pattern^" = "^(print_guards t guards)) - end - -fun group_rules rules = - let - fun add_rule (r as (_, PConst (c,_), _)) groups = - let - val rs = (case Inttab.lookup groups c of NONE => [] | SOME rs => rs) - in - Inttab.update (c, r::rs) groups - end - | add_rule _ _ = raise Compile "internal error group_rules" - in - fold_rev add_rule rules Inttab.empty - end - -fun sml_prog name code rules = - let - val buffer = Unsynchronized.ref "" - fun write s = (buffer := (!buffer)^s) - fun writeln s = (write s; write "\n") - fun writelist [] = () - | writelist (s::ss) = (writeln s; writelist ss) - fun str i = string_of_int i - val (inlinetab, rules) = inline_rules rules - val (arity, toplevel_arity, rules) = adjust_rules rules - val rules = group_rules rules - val constants = Inttab.keys arity - fun arity_of c = Inttab.lookup arity c - fun toplevel_arity_of c = Inttab.lookup toplevel_arity c - fun rep_str s n = implode (rep n s) - fun indexed s n = s^(str n) - fun string_of_tuple [] = "" - | string_of_tuple (x::xs) = "("^x^(implode (map (fn s => ", "^s) xs))^")" - fun string_of_args [] = "" - | string_of_args (x::xs) = x^(implode (map (fn s => " "^s) xs)) - fun default_case gnum c = - let - val leftargs = implode (map (indexed " x") (section (the (arity_of c)))) - val rightargs = section (the (arity_of c)) - val strict_args = (case toplevel_arity_of c of NONE => the (arity_of c) | SOME sa => sa) - val xs = map (fn n => if n < strict_args then "x"^(str n) else "x"^(str n)^"()") rightargs - val right = (indexed "C" c)^" "^(string_of_tuple xs) - val message = "(\"unresolved lazy call: " ^ string_of_int c ^ "\")" - val right = if strict_args < the (arity_of c) then "raise AM_SML.Run "^message else right - in - (indexed "c" c)^(if gnum > 0 then "_"^(str gnum) else "")^leftargs^" = "^right - end - - fun eval_rules c = - let - val arity = the (arity_of c) - val strict_arity = (case toplevel_arity_of c of NONE => arity | SOME sa => sa) - fun eval_rule n = - let - val sc = string_of_int c - val left = fold (fn i => fn s => "AbstractMachine.App ("^s^(indexed ", x" i)^")") (section n) ("AbstractMachine.Const "^sc) - fun arg i = - let - val x = indexed "x" i - val x = if i < n then "(eval bounds "^x^")" else x - val x = if i < strict_arity then x else "(fn () => "^x^")" - in - x - end - val right = "c"^sc^" "^(string_of_args (map arg (section arity))) - val right = fold_rev (fn i => fn s => "Abs (fn "^(indexed "x" i)^" => "^s^")") (List.drop (section arity, n)) right - val right = if arity > 0 then right else "C"^sc - in - " | eval bounds ("^left^") = "^right - end - in - map eval_rule (rev (section (arity + 1))) - end - - fun convert_computed_rules (c: int) : string list = - let - val arity = the (arity_of c) - fun eval_rule () = - let - val sc = string_of_int c - val left = fold (fn i => fn s => "AbstractMachine.App ("^s^(indexed ", x" i)^")") (section arity) ("AbstractMachine.Const "^sc) - fun arg i = "(convert_computed "^(indexed "x" i)^")" - val right = "C"^sc^" "^(string_of_tuple (map arg (section arity))) - val right = if arity > 0 then right else "C"^sc - in - " | convert_computed ("^left^") = "^right - end - in - [eval_rule ()] - end - - fun mk_constr_type_args n = if n > 0 then " of Term "^(rep_str " * Term" (n-1)) else "" - val _ = writelist [ - "structure "^name^" = struct", - "", - "datatype Term = Const of int | App of Term * Term | Abs of (Term -> Term)", - " "^(implode (map (fn c => " | C"^(str c)^(mk_constr_type_args (the (arity_of c)))) constants)), - ""] - fun make_constr c argprefix = "(C"^(str c)^" "^(string_of_tuple (map (fn i => argprefix^(str i)) (section (the (arity_of c)))))^")" - fun make_term_eq c = " | term_eq "^(make_constr c "a")^" "^(make_constr c "b")^" = "^ - (case the (arity_of c) of - 0 => "true" - | n => - let - val eqs = map (fn i => "term_eq a"^(str i)^" b"^(str i)) (section n) - val (eq, eqs) = (List.hd eqs, map (fn s => " andalso "^s) (List.tl eqs)) - in - eq^(implode eqs) - end) - val _ = writelist [ - "fun term_eq (Const c1) (Const c2) = (c1 = c2)", - " | term_eq (App (a1,a2)) (App (b1,b2)) = term_eq a1 b1 andalso term_eq a2 b2"] - val _ = writelist (map make_term_eq constants) - val _ = writelist [ - " | term_eq _ _ = false", - "" - ] - val _ = writelist [ - "fun app (Abs a) b = a b", - " | app a b = App (a, b)", - ""] - fun defcase gnum c = (case arity_of c of NONE => [] | SOME a => if a > 0 then [default_case gnum c] else []) - fun writefundecl [] = () - | writefundecl (x::xs) = writelist ((("and "^x)::(map (fn s => " | "^s) xs))) - fun list_group c = (case Inttab.lookup rules c of - NONE => [defcase 0 c] - | SOME rs => - let - val rs = - fold - (fn r => - fn rs => - let - val (gnum, l, rs) = - (case rs of - [] => (0, [], []) - | (gnum, l)::rs => (gnum, l, rs)) - val (gnum', r) = print_rule gnum arity_of toplevel_arity_of r - in - if gnum' = gnum then - (gnum, r::l)::rs - else - let - val args = implode (map (fn i => " a"^(str i)) (section (the (arity_of c)))) - fun gnumc g = if g > 0 then "c"^(str c)^"_"^(str g)^args else "c"^(str c)^args - val s = gnumc (gnum) ^ " = " ^ gnumc (gnum') - in - (gnum', [])::(gnum, s::r::l)::rs - end - end) - rs [] - val rs = (case rs of [] => [(0,defcase 0 c)] | (gnum,l)::rs => (gnum, (defcase gnum c)@l)::rs) - in - rev (map (fn z => rev (snd z)) rs) - end) - val _ = map (fn z => (map writefundecl z; writeln "")) (map list_group constants) - val _ = writelist [ - "fun convert (Const i) = AM_SML.Const i", - " | convert (App (a, b)) = AM_SML.App (convert a, convert b)", - " | convert (Abs _) = raise AM_SML.Run \"no abstraction in result allowed\""] - fun make_convert c = - let - val args = map (indexed "a") (section (the (arity_of c))) - val leftargs = - case args of - [] => "" - | (x::xs) => "("^x^(implode (map (fn s => ", "^s) xs))^")" - val args = map (indexed "convert a") (section (the (arity_of c))) - val right = fold (fn x => fn s => "AM_SML.App ("^s^", "^x^")") args ("AM_SML.Const "^(str c)) - in - " | convert (C"^(str c)^" "^leftargs^") = "^right - end - val _ = writelist (map make_convert constants) - val _ = writelist [ - "", - "fun convert_computed (AbstractMachine.Abs b) = raise AM_SML.Run \"no abstraction in convert_computed allowed\"", - " | convert_computed (AbstractMachine.Var i) = raise AM_SML.Run \"no bound variables in convert_computed allowed\""] - val _ = map (writelist o convert_computed_rules) constants - val _ = writelist [ - " | convert_computed (AbstractMachine.Const c) = Const c", - " | convert_computed (AbstractMachine.App (a, b)) = App (convert_computed a, convert_computed b)", - " | convert_computed (AbstractMachine.Computed a) = raise AM_SML.Run \"no nesting in convert_computed allowed\""] - val _ = writelist [ - "", - "fun eval bounds (AbstractMachine.Abs m) = Abs (fn b => eval (b::bounds) m)", - " | eval bounds (AbstractMachine.Var i) = AM_SML.list_nth (bounds, i)"] - val _ = map (writelist o eval_rules) constants - val _ = writelist [ - " | eval bounds (AbstractMachine.App (a, b)) = app (eval bounds a) (eval bounds b)", - " | eval bounds (AbstractMachine.Const c) = Const c", - " | eval bounds (AbstractMachine.Computed t) = convert_computed t"] - val _ = writelist [ - "", - "fun export term = AM_SML.save_result (\""^code^"\", convert term)", - "", - "val _ = AM_SML.set_compiled_rewriter (fn t => (convert (eval [] t)))", - "", - "end"] - in - (inlinetab, !buffer) - end - -val guid_counter = Unsynchronized.ref 0 -fun get_guid () = - let - val c = !guid_counter - val _ = guid_counter := !guid_counter + 1 - in - string_of_int (Time.toMicroseconds (Time.now ())) ^ string_of_int c - end - - -fun writeTextFile name s = File.write (Path.explode name) s - -fun use_source src = use_text ML_Env.local_context (1, "") false src - -fun compile rules = - let - val guid = get_guid () - val code = Real.toString (random ()) - val name = "AMSML_"^guid - val (inlinetab, source) = sml_prog name code rules - val _ = case !dump_output of NONE => () | SOME p => writeTextFile p source - val _ = compiled_rewriter := NONE - val _ = use_source source - in - case !compiled_rewriter of - NONE => raise Compile "broken link to compiled function" - | SOME compiled_fun => (inlinetab, compiled_fun) - end - -fun run (inlinetab, compiled_fun) t = - let - val _ = if check_freevars 0 t then () else raise Run ("can only compute closed terms") - fun inline (Const c) = (case Inttab.lookup inlinetab c of NONE => Const c | SOME t => t) - | inline (Var i) = Var i - | inline (App (a, b)) = App (inline a, inline b) - | inline (Abs m) = Abs (inline m) - | inline (Computed t) = Computed t - in - compiled_fun (beta (inline t)) - end - -end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/Compute_Oracle/compute.ML --- a/src/HOL/Matrix/Compute_Oracle/compute.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,653 +0,0 @@ -(* Title: HOL/Matrix/Compute_Oracle/compute.ML - Author: Steven Obua -*) - -signature COMPUTE = sig - - type computer - type theorem - type naming = int -> string - - datatype machine = BARRAS | BARRAS_COMPILED | HASKELL | SML - - (* Functions designated with a ! in front of them actually update the computer parameter *) - - exception Make of string - val make : machine -> theory -> thm list -> computer - val make_with_cache : machine -> theory -> term list -> thm list -> computer - val theory_of : computer -> theory - val hyps_of : computer -> term list - val shyps_of : computer -> sort list - (* ! *) val update : computer -> thm list -> unit - (* ! *) val update_with_cache : computer -> term list -> thm list -> unit - - (* ! *) val set_naming : computer -> naming -> unit - val naming_of : computer -> naming - - exception Compute of string - val simplify : computer -> theorem -> thm - val rewrite : computer -> cterm -> thm - - val make_theorem : computer -> thm -> string list -> theorem - (* ! *) val instantiate : computer -> (string * cterm) list -> theorem -> theorem - (* ! *) val evaluate_prem : computer -> int -> theorem -> theorem - (* ! *) val modus_ponens : computer -> int -> thm -> theorem -> theorem - -end - -structure Compute :> COMPUTE = struct - -open Report; - -datatype machine = BARRAS | BARRAS_COMPILED | HASKELL | SML - -(* Terms are mapped to integer codes *) -structure Encode :> -sig - type encoding - val empty : encoding - val insert : term -> encoding -> int * encoding - val lookup_code : term -> encoding -> int option - val lookup_term : int -> encoding -> term option - val remove_code : int -> encoding -> encoding - val remove_term : term -> encoding -> encoding -end -= -struct - -type encoding = int * (int Termtab.table) * (term Inttab.table) - -val empty = (0, Termtab.empty, Inttab.empty) - -fun insert t (e as (count, term2int, int2term)) = - (case Termtab.lookup term2int t of - NONE => (count, (count+1, Termtab.update_new (t, count) term2int, Inttab.update_new (count, t) int2term)) - | SOME code => (code, e)) - -fun lookup_code t (_, term2int, _) = Termtab.lookup term2int t - -fun lookup_term c (_, _, int2term) = Inttab.lookup int2term c - -fun remove_code c (e as (count, term2int, int2term)) = - (case lookup_term c e of NONE => e | SOME t => (count, Termtab.delete t term2int, Inttab.delete c int2term)) - -fun remove_term t (e as (count, term2int, int2term)) = - (case lookup_code t e of NONE => e | SOME c => (count, Termtab.delete t term2int, Inttab.delete c int2term)) - -end - -exception Make of string; -exception Compute of string; - -local - fun make_constant t encoding = - let - val (code, encoding) = Encode.insert t encoding - in - (encoding, AbstractMachine.Const code) - end -in - -fun remove_types encoding t = - case t of - Var _ => make_constant t encoding - | Free _ => make_constant t encoding - | Const _ => make_constant t encoding - | Abs (_, _, t') => - let val (encoding, t'') = remove_types encoding t' in - (encoding, AbstractMachine.Abs t'') - end - | a $ b => - let - val (encoding, a) = remove_types encoding a - val (encoding, b) = remove_types encoding b - in - (encoding, AbstractMachine.App (a,b)) - end - | Bound b => (encoding, AbstractMachine.Var b) -end - -local - fun type_of (Free (_, ty)) = ty - | type_of (Const (_, ty)) = ty - | type_of (Var (_, ty)) = ty - | type_of _ = raise Fail "infer_types: type_of error" -in -fun infer_types naming encoding = - let - fun infer_types _ bounds _ (AbstractMachine.Var v) = (Bound v, nth bounds v) - | infer_types _ bounds _ (AbstractMachine.Const code) = - let - val c = the (Encode.lookup_term code encoding) - in - (c, type_of c) - end - | infer_types level bounds _ (AbstractMachine.App (a, b)) = - let - val (a, aty) = infer_types level bounds NONE a - val (adom, arange) = - case aty of - Type ("fun", [dom, range]) => (dom, range) - | _ => raise Fail "infer_types: function type expected" - val (b, _) = infer_types level bounds (SOME adom) b - in - (a $ b, arange) - end - | infer_types level bounds (SOME (ty as Type ("fun", [dom, range]))) (AbstractMachine.Abs m) = - let - val (m, _) = infer_types (level+1) (dom::bounds) (SOME range) m - in - (Abs (naming level, dom, m), ty) - end - | infer_types _ _ NONE (AbstractMachine.Abs _) = - raise Fail "infer_types: cannot infer type of abstraction" - - fun infer ty term = - let - val (term', _) = infer_types 0 [] (SOME ty) term - in - term' - end - in - infer - end -end - -datatype prog = - ProgBarras of AM_Interpreter.program - | ProgBarrasC of AM_Compiler.program - | ProgHaskell of AM_GHC.program - | ProgSML of AM_SML.program - -fun machine_of_prog (ProgBarras _) = BARRAS - | machine_of_prog (ProgBarrasC _) = BARRAS_COMPILED - | machine_of_prog (ProgHaskell _) = HASKELL - | machine_of_prog (ProgSML _) = SML - -type naming = int -> string - -fun default_naming i = "v_" ^ string_of_int i - -datatype computer = Computer of - (theory_ref * Encode.encoding * term list * unit Sorttab.table * prog * unit Unsynchronized.ref * naming) - option Unsynchronized.ref - -fun theory_of (Computer (Unsynchronized.ref (SOME (rthy,_,_,_,_,_,_)))) = Theory.deref rthy -fun hyps_of (Computer (Unsynchronized.ref (SOME (_,_,hyps,_,_,_,_)))) = hyps -fun shyps_of (Computer (Unsynchronized.ref (SOME (_,_,_,shyptable,_,_,_)))) = Sorttab.keys (shyptable) -fun shyptab_of (Computer (Unsynchronized.ref (SOME (_,_,_,shyptable,_,_,_)))) = shyptable -fun stamp_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,_,stamp,_)))) = stamp -fun prog_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,prog,_,_)))) = prog -fun encoding_of (Computer (Unsynchronized.ref (SOME (_,encoding,_,_,_,_,_)))) = encoding -fun set_encoding (Computer (r as Unsynchronized.ref (SOME (p1,_,p2,p3,p4,p5,p6)))) encoding' = - (r := SOME (p1,encoding',p2,p3,p4,p5,p6)) -fun naming_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,_,_,n)))) = n -fun set_naming (Computer (r as Unsynchronized.ref (SOME (p1,p2,p3,p4,p5,p6,_)))) naming'= - (r := SOME (p1,p2,p3,p4,p5,p6,naming')) - -fun ref_of (Computer r) = r - -datatype cthm = ComputeThm of term list * sort list * term - -fun thm2cthm th = - let - val {hyps, prop, tpairs, shyps, ...} = Thm.rep_thm th - val _ = if not (null tpairs) then raise Make "theorems may not contain tpairs" else () - in - ComputeThm (hyps, shyps, prop) - end - -fun make_internal machine thy stamp encoding cache_pattern_terms raw_ths = - let - fun transfer (x:thm) = Thm.transfer thy x - val ths = map (thm2cthm o Thm.strip_shyps o transfer) raw_ths - - fun make_pattern encoding n vars (AbstractMachine.Abs _) = - raise (Make "no lambda abstractions allowed in pattern") - | make_pattern encoding n vars (AbstractMachine.Var _) = - raise (Make "no bound variables allowed in pattern") - | make_pattern encoding n vars (AbstractMachine.Const code) = - (case the (Encode.lookup_term code encoding) of - Var _ => ((n+1, Inttab.update_new (code, n) vars, AbstractMachine.PVar) - handle Inttab.DUP _ => raise (Make "no duplicate variable in pattern allowed")) - | _ => (n, vars, AbstractMachine.PConst (code, []))) - | make_pattern encoding n vars (AbstractMachine.App (a, b)) = - let - val (n, vars, pa) = make_pattern encoding n vars a - val (n, vars, pb) = make_pattern encoding n vars b - in - case pa of - AbstractMachine.PVar => - raise (Make "patterns may not start with a variable") - | AbstractMachine.PConst (c, args) => - (n, vars, AbstractMachine.PConst (c, args@[pb])) - end - - fun thm2rule (encoding, hyptable, shyptable) th = - let - val (ComputeThm (hyps, shyps, prop)) = th - val hyptable = fold (fn h => Termtab.update (h, ())) hyps hyptable - val shyptable = fold (fn sh => Sorttab.update (sh, ())) shyps shyptable - val (prems, prop) = (Logic.strip_imp_prems prop, Logic.strip_imp_concl prop) - val (a, b) = Logic.dest_equals prop - handle TERM _ => raise (Make "theorems must be meta-level equations (with optional guards)") - val a = Envir.eta_contract a - val b = Envir.eta_contract b - val prems = map Envir.eta_contract prems - - val (encoding, left) = remove_types encoding a - val (encoding, right) = remove_types encoding b - fun remove_types_of_guard encoding g = - (let - val (t1, t2) = Logic.dest_equals g - val (encoding, t1) = remove_types encoding t1 - val (encoding, t2) = remove_types encoding t2 - in - (encoding, AbstractMachine.Guard (t1, t2)) - end handle TERM _ => raise (Make "guards must be meta-level equations")) - val (encoding, prems) = fold_rev (fn p => fn (encoding, ps) => let val (e, p) = remove_types_of_guard encoding p in (e, p::ps) end) prems (encoding, []) - - (* Principally, a check should be made here to see if the (meta-) hyps contain any of the variables of the rule. - As it is, all variables of the rule are schematic, and there are no schematic variables in meta-hyps, therefore - this check can be left out. *) - - val (vcount, vars, pattern) = make_pattern encoding 0 Inttab.empty left - val _ = (case pattern of - AbstractMachine.PVar => - raise (Make "patterns may not start with a variable") - | _ => ()) - - (* finally, provide a function for renaming the - pattern bound variables on the right hand side *) - - fun rename level vars (var as AbstractMachine.Var _) = var - | rename level vars (c as AbstractMachine.Const code) = - (case Inttab.lookup vars code of - NONE => c - | SOME n => AbstractMachine.Var (vcount-n-1+level)) - | rename level vars (AbstractMachine.App (a, b)) = - AbstractMachine.App (rename level vars a, rename level vars b) - | rename level vars (AbstractMachine.Abs m) = - AbstractMachine.Abs (rename (level+1) vars m) - - fun rename_guard (AbstractMachine.Guard (a,b)) = - AbstractMachine.Guard (rename 0 vars a, rename 0 vars b) - in - ((encoding, hyptable, shyptable), (map rename_guard prems, pattern, rename 0 vars right)) - end - - val ((encoding, hyptable, shyptable), rules) = - fold_rev (fn th => fn (encoding_hyptable, rules) => - let - val (encoding_hyptable, rule) = thm2rule encoding_hyptable th - in (encoding_hyptable, rule::rules) end) - ths ((encoding, Termtab.empty, Sorttab.empty), []) - - fun make_cache_pattern t (encoding, cache_patterns) = - let - val (encoding, a) = remove_types encoding t - val (_,_,p) = make_pattern encoding 0 Inttab.empty a - in - (encoding, p::cache_patterns) - end - - val (encoding, _) = fold_rev make_cache_pattern cache_pattern_terms (encoding, []) - - val prog = - case machine of - BARRAS => ProgBarras (AM_Interpreter.compile rules) - | BARRAS_COMPILED => ProgBarrasC (AM_Compiler.compile rules) - | HASKELL => ProgHaskell (AM_GHC.compile rules) - | SML => ProgSML (AM_SML.compile rules) - - fun has_witness s = not (null (Sign.witness_sorts thy [] [s])) - - val shyptable = fold Sorttab.delete (filter has_witness (Sorttab.keys (shyptable))) shyptable - - in (Theory.check_thy thy, encoding, Termtab.keys hyptable, shyptable, prog, stamp, default_naming) end - -fun make_with_cache machine thy cache_patterns raw_thms = - Computer (Unsynchronized.ref (SOME (make_internal machine thy (Unsynchronized.ref ()) Encode.empty cache_patterns raw_thms))) - -fun make machine thy raw_thms = make_with_cache machine thy [] raw_thms - -fun update_with_cache computer cache_patterns raw_thms = - let - val c = make_internal (machine_of_prog (prog_of computer)) (theory_of computer) (stamp_of computer) - (encoding_of computer) cache_patterns raw_thms - val _ = (ref_of computer) := SOME c - in - () - end - -fun update computer raw_thms = update_with_cache computer [] raw_thms - -fun runprog (ProgBarras p) = AM_Interpreter.run p - | runprog (ProgBarrasC p) = AM_Compiler.run p - | runprog (ProgHaskell p) = AM_GHC.run p - | runprog (ProgSML p) = AM_SML.run p - -(* ------------------------------------------------------------------------------------- *) -(* An oracle for exporting theorems; must only be accessible from inside this structure! *) -(* ------------------------------------------------------------------------------------- *) - -fun merge_hyps hyps1 hyps2 = -let - fun add hyps tab = fold (fn h => fn tab => Termtab.update (h, ()) tab) hyps tab -in - Termtab.keys (add hyps2 (add hyps1 Termtab.empty)) -end - -fun add_shyps shyps tab = fold (fn h => fn tab => Sorttab.update (h, ()) tab) shyps tab - -fun merge_shyps shyps1 shyps2 = Sorttab.keys (add_shyps shyps2 (add_shyps shyps1 Sorttab.empty)) - -val (_, export_oracle) = Context.>>> (Context.map_theory_result - (Thm.add_oracle (@{binding compute}, fn (thy, hyps, shyps, prop) => - let - val shyptab = add_shyps shyps Sorttab.empty - fun delete s shyptab = Sorttab.delete s shyptab handle Sorttab.UNDEF _ => shyptab - fun delete_term t shyptab = fold delete (Sorts.insert_term t []) shyptab - fun has_witness s = not (null (Sign.witness_sorts thy [] [s])) - val shyptab = fold Sorttab.delete (filter has_witness (Sorttab.keys (shyptab))) shyptab - val shyps = if Sorttab.is_empty shyptab then [] else Sorttab.keys (fold delete_term (prop::hyps) shyptab) - val _ = - if not (null shyps) then - raise Compute ("dangling sort hypotheses: " ^ - commas (map (Syntax.string_of_sort_global thy) shyps)) - else () - in - Thm.cterm_of thy (fold_rev (fn hyp => fn p => Logic.mk_implies (hyp, p)) hyps prop) - end))); - -fun export_thm thy hyps shyps prop = - let - val th = export_oracle (thy, hyps, shyps, prop) - val hyps = map (fn h => Thm.assume (cterm_of thy h)) hyps - in - fold (fn h => fn p => Thm.implies_elim p h) hyps th - end - -(* --------- Rewrite ----------- *) - -fun rewrite computer ct = - let - val thy = Thm.theory_of_cterm ct - val {t=t',T=ty,...} = rep_cterm ct - val _ = Theory.assert_super (theory_of computer) thy - val naming = naming_of computer - val (encoding, t) = remove_types (encoding_of computer) t' - val t = runprog (prog_of computer) t - val t = infer_types naming encoding ty t - val eq = Logic.mk_equals (t', t) - in - export_thm thy (hyps_of computer) (Sorttab.keys (shyptab_of computer)) eq - end - -(* --------- Simplify ------------ *) - -datatype prem = EqPrem of AbstractMachine.term * AbstractMachine.term * Term.typ * int - | Prem of AbstractMachine.term -datatype theorem = Theorem of theory_ref * unit Unsynchronized.ref * (int * typ) Symtab.table * (AbstractMachine.term option) Inttab.table - * prem list * AbstractMachine.term * term list * sort list - - -exception ParamSimplify of computer * theorem - -fun make_theorem computer th vars = -let - val _ = Theory.assert_super (theory_of computer) (theory_of_thm th) - - val (ComputeThm (hyps, shyps, prop)) = thm2cthm th - - val encoding = encoding_of computer - - (* variables in the theorem are identified upfront *) - fun collect_vars (Abs (_, _, t)) tab = collect_vars t tab - | collect_vars (a $ b) tab = collect_vars b (collect_vars a tab) - | collect_vars (Const _) tab = tab - | collect_vars (Free _) tab = tab - | collect_vars (Var ((s, i), ty)) tab = - if List.find (fn x => x=s) vars = NONE then - tab - else - (case Symtab.lookup tab s of - SOME ((s',i'),ty') => - if s' <> s orelse i' <> i orelse ty <> ty' then - raise Compute ("make_theorem: variable name '"^s^"' is not unique") - else - tab - | NONE => Symtab.update (s, ((s, i), ty)) tab) - val vartab = collect_vars prop Symtab.empty - fun encodevar (s, t as (_, ty)) (encoding, tab) = - let - val (x, encoding) = Encode.insert (Var t) encoding - in - (encoding, Symtab.update (s, (x, ty)) tab) - end - val (encoding, vartab) = Symtab.fold encodevar vartab (encoding, Symtab.empty) - val varsubst = Inttab.make (map (fn (_, (x, _)) => (x, NONE)) (Symtab.dest vartab)) - - (* make the premises and the conclusion *) - fun mk_prem encoding t = - (let - val (a, b) = Logic.dest_equals t - val ty = type_of a - val (encoding, a) = remove_types encoding a - val (encoding, b) = remove_types encoding b - val (eq, encoding) = Encode.insert (Const ("==", ty --> ty --> @{typ "prop"})) encoding - in - (encoding, EqPrem (a, b, ty, eq)) - end handle TERM _ => let val (encoding, t) = remove_types encoding t in (encoding, Prem t) end) - val (encoding, prems) = - (fold_rev (fn t => fn (encoding, l) => - case mk_prem encoding t of - (encoding, t) => (encoding, t::l)) (Logic.strip_imp_prems prop) (encoding, [])) - val (encoding, concl) = remove_types encoding (Logic.strip_imp_concl prop) - val _ = set_encoding computer encoding -in - Theorem (Theory.check_thy (theory_of_thm th), stamp_of computer, vartab, varsubst, - prems, concl, hyps, shyps) -end - -fun theory_of_theorem (Theorem (rthy,_,_,_,_,_,_,_)) = Theory.deref rthy -fun update_theory thy (Theorem (_,p0,p1,p2,p3,p4,p5,p6)) = - Theorem (Theory.check_thy thy,p0,p1,p2,p3,p4,p5,p6) -fun stamp_of_theorem (Theorem (_,s, _, _, _, _, _, _)) = s -fun vartab_of_theorem (Theorem (_,_,vt,_,_,_,_,_)) = vt -fun varsubst_of_theorem (Theorem (_,_,_,vs,_,_,_,_)) = vs -fun update_varsubst vs (Theorem (p0,p1,p2,_,p3,p4,p5,p6)) = Theorem (p0,p1,p2,vs,p3,p4,p5,p6) -fun prems_of_theorem (Theorem (_,_,_,_,prems,_,_,_)) = prems -fun update_prems prems (Theorem (p0,p1,p2,p3,_,p4,p5,p6)) = Theorem (p0,p1,p2,p3,prems,p4,p5,p6) -fun concl_of_theorem (Theorem (_,_,_,_,_,concl,_,_)) = concl -fun hyps_of_theorem (Theorem (_,_,_,_,_,_,hyps,_)) = hyps -fun update_hyps hyps (Theorem (p0,p1,p2,p3,p4,p5,_,p6)) = Theorem (p0,p1,p2,p3,p4,p5,hyps,p6) -fun shyps_of_theorem (Theorem (_,_,_,_,_,_,_,shyps)) = shyps -fun update_shyps shyps (Theorem (p0,p1,p2,p3,p4,p5,p6,_)) = Theorem (p0,p1,p2,p3,p4,p5,p6,shyps) - -fun check_compatible computer th s = - if stamp_of computer <> stamp_of_theorem th then - raise Compute (s^": computer and theorem are incompatible") - else () - -fun instantiate computer insts th = -let - val _ = check_compatible computer th - - val thy = theory_of computer - - val vartab = vartab_of_theorem th - - fun rewrite computer t = - let - val (encoding, t) = remove_types (encoding_of computer) t - val t = runprog (prog_of computer) t - val _ = set_encoding computer encoding - in - t - end - - fun assert_varfree vs t = - if AbstractMachine.forall_consts (fn x => Inttab.lookup vs x = NONE) t then - () - else - raise Compute "instantiate: assert_varfree failed" - - fun assert_closed t = - if AbstractMachine.closed t then - () - else - raise Compute "instantiate: not a closed term" - - fun compute_inst (s, ct) vs = - let - val _ = Theory.assert_super (theory_of_cterm ct) thy - val ty = typ_of (ctyp_of_term ct) - in - (case Symtab.lookup vartab s of - NONE => raise Compute ("instantiate: variable '"^s^"' not found in theorem") - | SOME (x, ty') => - (case Inttab.lookup vs x of - SOME (SOME _) => raise Compute ("instantiate: variable '"^s^"' has already been instantiated") - | SOME NONE => - if ty <> ty' then - raise Compute ("instantiate: wrong type for variable '"^s^"'") - else - let - val t = rewrite computer (term_of ct) - val _ = assert_varfree vs t - val _ = assert_closed t - in - Inttab.update (x, SOME t) vs - end - | NONE => raise Compute "instantiate: internal error")) - end - - val vs = fold compute_inst insts (varsubst_of_theorem th) -in - update_varsubst vs th -end - -fun match_aterms subst = - let - exception no_match - open AbstractMachine - fun match subst (b as (Const c)) a = - if a = b then subst - else - (case Inttab.lookup subst c of - SOME (SOME a') => if a=a' then subst else raise no_match - | SOME NONE => if AbstractMachine.closed a then - Inttab.update (c, SOME a) subst - else raise no_match - | NONE => raise no_match) - | match subst (b as (Var _)) a = if a=b then subst else raise no_match - | match subst (App (u, v)) (App (u', v')) = match (match subst u u') v v' - | match subst (Abs u) (Abs u') = match subst u u' - | match subst _ _ = raise no_match - in - fn b => fn a => (SOME (match subst b a) handle no_match => NONE) - end - -fun apply_subst vars_allowed subst = - let - open AbstractMachine - fun app (t as (Const c)) = - (case Inttab.lookup subst c of - NONE => t - | SOME (SOME t) => Computed t - | SOME NONE => if vars_allowed then t else raise Compute "apply_subst: no vars allowed") - | app (t as (Var _)) = t - | app (App (u, v)) = App (app u, app v) - | app (Abs m) = Abs (app m) - in - app - end - -fun splicein n l L = List.take (L, n) @ l @ List.drop (L, n+1) - -fun evaluate_prem computer prem_no th = -let - val _ = check_compatible computer th - val prems = prems_of_theorem th - val varsubst = varsubst_of_theorem th - fun run vars_allowed t = - runprog (prog_of computer) (apply_subst vars_allowed varsubst t) -in - case nth prems prem_no of - Prem _ => raise Compute "evaluate_prem: no equality premise" - | EqPrem (a, b, ty, _) => - let - val a' = run false a - val b' = run true b - in - case match_aterms varsubst b' a' of - NONE => - let - fun mk s = Syntax.string_of_term_global Pure.thy - (infer_types (naming_of computer) (encoding_of computer) ty s) - val left = "computed left side: "^(mk a') - val right = "computed right side: "^(mk b') - in - raise Compute ("evaluate_prem: cannot assign computed left to right hand side\n"^left^"\n"^right^"\n") - end - | SOME varsubst => - update_prems (splicein prem_no [] prems) (update_varsubst varsubst th) - end -end - -fun prem2term (Prem t) = t - | prem2term (EqPrem (a,b,_,eq)) = - AbstractMachine.App (AbstractMachine.App (AbstractMachine.Const eq, a), b) - -fun modus_ponens computer prem_no th' th = -let - val _ = check_compatible computer th - val thy = - let - val thy1 = theory_of_theorem th - val thy2 = theory_of_thm th' - in - if Theory.subthy (thy1, thy2) then thy2 - else if Theory.subthy (thy2, thy1) then thy1 else - raise Compute "modus_ponens: theorems are not compatible with each other" - end - val th' = make_theorem computer th' [] - val varsubst = varsubst_of_theorem th - fun run vars_allowed t = - runprog (prog_of computer) (apply_subst vars_allowed varsubst t) - val prems = prems_of_theorem th - val prem = run true (prem2term (nth prems prem_no)) - val concl = run false (concl_of_theorem th') -in - case match_aterms varsubst prem concl of - NONE => raise Compute "modus_ponens: conclusion does not match premise" - | SOME varsubst => - let - val th = update_varsubst varsubst th - val th = update_prems (splicein prem_no (prems_of_theorem th') prems) th - val th = update_hyps (merge_hyps (hyps_of_theorem th) (hyps_of_theorem th')) th - val th = update_shyps (merge_shyps (shyps_of_theorem th) (shyps_of_theorem th')) th - in - update_theory thy th - end -end - -fun simplify computer th = -let - val _ = check_compatible computer th - val varsubst = varsubst_of_theorem th - val encoding = encoding_of computer - val naming = naming_of computer - fun infer t = infer_types naming encoding @{typ "prop"} t - fun run t = infer (runprog (prog_of computer) (apply_subst true varsubst t)) - fun runprem p = run (prem2term p) - val prop = Logic.list_implies (map runprem (prems_of_theorem th), run (concl_of_theorem th)) - val hyps = merge_hyps (hyps_of computer) (hyps_of_theorem th) - val shyps = merge_shyps (shyps_of_theorem th) (Sorttab.keys (shyptab_of computer)) -in - export_thm (theory_of_theorem th) hyps shyps prop -end - -end - diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/Compute_Oracle/linker.ML --- a/src/HOL/Matrix/Compute_Oracle/linker.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,470 +0,0 @@ -(* Title: HOL/Matrix/Compute_Oracle/linker.ML - Author: Steven Obua - -This module solves the problem that the computing oracle does not -instantiate polymorphic rules. By going through the PCompute -interface, all possible instantiations are resolved by compiling new -programs, if necessary. The obvious disadvantage of this approach is -that in the worst case for each new term to be rewritten, a new -program may be compiled. -*) - -(* - Given constants/frees c_1::t_1, c_2::t_2, ...., c_n::t_n, - and constants/frees d_1::d_1, d_2::s_2, ..., d_m::s_m - - Find all substitutions S such that - a) the domain of S is tvars (t_1, ..., t_n) - b) there are indices i_1, ..., i_k, and j_1, ..., j_k with - 1. S (c_i_1::t_i_1) = d_j_1::s_j_1, ..., S (c_i_k::t_i_k) = d_j_k::s_j_k - 2. tvars (t_i_1, ..., t_i_k) = tvars (t_1, ..., t_n) -*) -signature LINKER = -sig - exception Link of string - - datatype constant = Constant of bool * string * typ - val constant_of : term -> constant - - type instances - type subst = Type.tyenv - - val empty : constant list -> instances - val typ_of_constant : constant -> typ - val add_instances : theory -> instances -> constant list -> subst list * instances - val substs_of : instances -> subst list - val is_polymorphic : constant -> bool - val distinct_constants : constant list -> constant list - val collect_consts : term list -> constant list -end - -structure Linker : LINKER = struct - -exception Link of string; - -type subst = Type.tyenv - -datatype constant = Constant of bool * string * typ -fun constant_of (Const (name, ty)) = Constant (false, name, ty) - | constant_of (Free (name, ty)) = Constant (true, name, ty) - | constant_of _ = raise Link "constant_of" - -fun bool_ord (x,y) = if x then (if y then EQUAL else GREATER) else (if y then LESS else EQUAL) -fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) Term_Ord.typ_ord) (((x1,x2),x3), ((y1,y2),y3)) -fun constant_modty_ord (Constant (x1,x2,_), Constant (y1,y2,_)) = (prod_ord bool_ord fast_string_ord) ((x1,x2), (y1,y2)) - - -structure Consttab = Table(type key = constant val ord = constant_ord); -structure ConsttabModTy = Table(type key = constant val ord = constant_modty_ord); - -fun typ_of_constant (Constant (_, _, ty)) = ty - -val empty_subst = (Vartab.empty : Type.tyenv) - -fun merge_subst (A:Type.tyenv) (B:Type.tyenv) = - SOME (Vartab.fold (fn (v, t) => - fn tab => - (case Vartab.lookup tab v of - NONE => Vartab.update (v, t) tab - | SOME t' => if t = t' then tab else raise Type.TYPE_MATCH)) A B) - handle Type.TYPE_MATCH => NONE - -fun subst_ord (A:Type.tyenv, B:Type.tyenv) = - (list_ord (prod_ord Term_Ord.fast_indexname_ord (prod_ord Term_Ord.sort_ord Term_Ord.typ_ord))) (Vartab.dest A, Vartab.dest B) - -structure Substtab = Table(type key = Type.tyenv val ord = subst_ord); - -fun substtab_union c = Substtab.fold Substtab.update c -fun substtab_unions [] = Substtab.empty - | substtab_unions [c] = c - | substtab_unions (c::cs) = substtab_union c (substtab_unions cs) - -datatype instances = Instances of unit ConsttabModTy.table * Type.tyenv Consttab.table Consttab.table * constant list list * unit Substtab.table - -fun is_polymorphic (Constant (_, _, ty)) = not (null (Term.add_tvarsT ty [])) - -fun distinct_constants cs = - Consttab.keys (fold (fn c => Consttab.update (c, ())) cs Consttab.empty) - -fun empty cs = - let - val cs = distinct_constants (filter is_polymorphic cs) - val old_cs = cs -(* fun collect_tvars ty tab = fold (fn v => fn tab => Typtab.update (TVar v, ()) tab) (Misc_Legacy.typ_tvars ty) tab - val tvars_count = length (Typtab.keys (fold (fn c => fn tab => collect_tvars (typ_of_constant c) tab) cs Typtab.empty)) - fun tvars_of ty = collect_tvars ty Typtab.empty - val cs = map (fn c => (c, tvars_of (typ_of_constant c))) cs - - fun tyunion A B = - Typtab.fold - (fn (v,()) => fn tab => Typtab.update (v, case Typtab.lookup tab v of NONE => 1 | SOME n => n+1) tab) - A B - - fun is_essential A B = - Typtab.fold - (fn (v, ()) => fn essential => essential orelse (case Typtab.lookup B v of NONE => raise Link "is_essential" | SOME n => n=1)) - A false - - fun add_minimal (c', tvs') (tvs, cs) = - let - val tvs = tyunion tvs' tvs - val cs = (c', tvs')::cs - in - if forall (fn (c',tvs') => is_essential tvs' tvs) cs then - SOME (tvs, cs) - else - NONE - end - - fun is_spanning (tvs, _) = (length (Typtab.keys tvs) = tvars_count) - - fun generate_minimal_subsets subsets [] = subsets - | generate_minimal_subsets subsets (c::cs) = - let - val subsets' = map_filter (add_minimal c) subsets - in - generate_minimal_subsets (subsets@subsets') cs - end*) - - val minimal_subsets = [old_cs] (*map (fn (tvs, cs) => map fst cs) (filter is_spanning (generate_minimal_subsets [(Typtab.empty, [])] cs))*) - - val constants = Consttab.keys (fold (fold (fn c => Consttab.update (c, ()))) minimal_subsets Consttab.empty) - - in - Instances ( - fold (fn c => fn tab => ConsttabModTy.update (c, ()) tab) constants ConsttabModTy.empty, - Consttab.make (map (fn c => (c, Consttab.empty : Type.tyenv Consttab.table)) constants), - minimal_subsets, Substtab.empty) - end - -local -fun calc ctab substtab [] = substtab - | calc ctab substtab (c::cs) = - let - val csubsts = map snd (Consttab.dest (the (Consttab.lookup ctab c))) - fun merge_substs substtab subst = - Substtab.fold (fn (s,_) => - fn tab => - (case merge_subst subst s of NONE => tab | SOME s => Substtab.update (s, ()) tab)) - substtab Substtab.empty - val substtab = substtab_unions (map (merge_substs substtab) csubsts) - in - calc ctab substtab cs - end -in -fun calc_substs ctab (cs:constant list) = calc ctab (Substtab.update (empty_subst, ()) Substtab.empty) cs -end - -fun add_instances thy (Instances (cfilter, ctab,minsets,substs)) cs = - let -(* val _ = writeln (makestring ("add_instances: ", length_cs, length cs, length (Consttab.keys ctab)))*) - fun calc_instantiations (constant as Constant (free, name, ty)) instantiations = - Consttab.fold (fn (constant' as Constant (free', name', ty'), insttab) => - fn instantiations => - if free <> free' orelse name <> name' then - instantiations - else case Consttab.lookup insttab constant of - SOME _ => instantiations - | NONE => ((constant', (constant, Sign.typ_match thy (ty', ty) empty_subst))::instantiations - handle Type.TYPE_MATCH => instantiations)) - ctab instantiations - val instantiations = fold calc_instantiations cs [] - (*val _ = writeln ("instantiations = "^(makestring (length instantiations)))*) - fun update_ctab (constant', entry) ctab = - (case Consttab.lookup ctab constant' of - NONE => raise Link "internal error: update_ctab" - | SOME tab => Consttab.update (constant', Consttab.update entry tab) ctab) - val ctab = fold update_ctab instantiations ctab - val new_substs = fold (fn minset => fn substs => substtab_union (calc_substs ctab minset) substs) - minsets Substtab.empty - val (added_substs, substs) = - Substtab.fold (fn (ns, _) => - fn (added, substtab) => - (case Substtab.lookup substs ns of - NONE => (ns::added, Substtab.update (ns, ()) substtab) - | SOME () => (added, substtab))) - new_substs ([], substs) - in - (added_substs, Instances (cfilter, ctab, minsets, substs)) - end - -fun substs_of (Instances (_,_,_,substs)) = Substtab.keys substs - - -local - -fun collect (Var _) tab = tab - | collect (Bound _) tab = tab - | collect (a $ b) tab = collect b (collect a tab) - | collect (Abs (_, _, body)) tab = collect body tab - | collect t tab = Consttab.update (constant_of t, ()) tab - -in - fun collect_consts tms = Consttab.keys (fold collect tms Consttab.empty) -end - -end - -signature PCOMPUTE = -sig - type pcomputer - - val make : Compute.machine -> theory -> thm list -> Linker.constant list -> pcomputer - val make_with_cache : Compute.machine -> theory -> term list -> thm list -> Linker.constant list -> pcomputer - - val add_instances : pcomputer -> Linker.constant list -> bool - val add_instances' : pcomputer -> term list -> bool - - val rewrite : pcomputer -> cterm list -> thm list - val simplify : pcomputer -> Compute.theorem -> thm - - val make_theorem : pcomputer -> thm -> string list -> Compute.theorem - val instantiate : pcomputer -> (string * cterm) list -> Compute.theorem -> Compute.theorem - val evaluate_prem : pcomputer -> int -> Compute.theorem -> Compute.theorem - val modus_ponens : pcomputer -> int -> thm -> Compute.theorem -> Compute.theorem - -end - -structure PCompute : PCOMPUTE = struct - -exception PCompute of string - -datatype theorem = MonoThm of thm | PolyThm of thm * Linker.instances * thm list -datatype pattern = MonoPattern of term | PolyPattern of term * Linker.instances * term list - -datatype pcomputer = - PComputer of theory_ref * Compute.computer * theorem list Unsynchronized.ref * - pattern list Unsynchronized.ref - -(*fun collect_consts (Var x) = [] - | collect_consts (Bound _) = [] - | collect_consts (a $ b) = (collect_consts a)@(collect_consts b) - | collect_consts (Abs (_, _, body)) = collect_consts body - | collect_consts t = [Linker.constant_of t]*) - -fun computer_of (PComputer (_,computer,_,_)) = computer - -fun collect_consts_of_thm th = - let - val th = prop_of th - val (prems, th) = (Logic.strip_imp_prems th, Logic.strip_imp_concl th) - val (left, right) = Logic.dest_equals th - in - (Linker.collect_consts [left], Linker.collect_consts (right::prems)) - end - -fun create_theorem th = -let - val (left, right) = collect_consts_of_thm th - val polycs = filter Linker.is_polymorphic left - val tytab = fold (fn p => fn tab => fold (fn n => fn tab => Typtab.update (TVar n, ()) tab) (Misc_Legacy.typ_tvars (Linker.typ_of_constant p)) tab) polycs Typtab.empty - fun check_const (c::cs) cs' = - let - val tvars = Misc_Legacy.typ_tvars (Linker.typ_of_constant c) - val wrong = fold (fn n => fn wrong => wrong orelse is_none (Typtab.lookup tytab (TVar n))) tvars false - in - if wrong then raise PCompute "right hand side of theorem contains type variables which do not occur on the left hand side" - else - if null (tvars) then - check_const cs (c::cs') - else - check_const cs cs' - end - | check_const [] cs' = cs' - val monocs = check_const right [] -in - if null (polycs) then - (monocs, MonoThm th) - else - (monocs, PolyThm (th, Linker.empty polycs, [])) -end - -fun create_pattern pat = -let - val cs = Linker.collect_consts [pat] - val polycs = filter Linker.is_polymorphic cs -in - if null (polycs) then - MonoPattern pat - else - PolyPattern (pat, Linker.empty polycs, []) -end - -fun create_computer machine thy pats ths = - let - fun add (MonoThm th) ths = th::ths - | add (PolyThm (_, _, ths')) ths = ths'@ths - fun addpat (MonoPattern p) pats = p::pats - | addpat (PolyPattern (_, _, ps)) pats = ps@pats - val ths = fold_rev add ths [] - val pats = fold_rev addpat pats [] - in - Compute.make_with_cache machine thy pats ths - end - -fun update_computer computer pats ths = - let - fun add (MonoThm th) ths = th::ths - | add (PolyThm (_, _, ths')) ths = ths'@ths - fun addpat (MonoPattern p) pats = p::pats - | addpat (PolyPattern (_, _, ps)) pats = ps@pats - val ths = fold_rev add ths [] - val pats = fold_rev addpat pats [] - in - Compute.update_with_cache computer pats ths - end - -fun conv_subst thy (subst : Type.tyenv) = - map (fn (iname, (sort, ty)) => (ctyp_of thy (TVar (iname, sort)), ctyp_of thy ty)) (Vartab.dest subst) - -fun add_monos thy monocs pats ths = - let - val changed = Unsynchronized.ref false - fun add monocs (th as (MonoThm _)) = ([], th) - | add monocs (PolyThm (th, instances, instanceths)) = - let - val (newsubsts, instances) = Linker.add_instances thy instances monocs - val _ = if not (null newsubsts) then changed := true else () - val newths = map (fn subst => Thm.instantiate (conv_subst thy subst, []) th) newsubsts -(* val _ = if not (null newths) then (print ("added new theorems: ", newths); ()) else ()*) - val newmonos = fold (fn th => fn monos => (snd (collect_consts_of_thm th))@monos) newths [] - in - (newmonos, PolyThm (th, instances, instanceths@newths)) - end - fun addpats monocs (pat as (MonoPattern _)) = pat - | addpats monocs (PolyPattern (p, instances, instancepats)) = - let - val (newsubsts, instances) = Linker.add_instances thy instances monocs - val _ = if not (null newsubsts) then changed := true else () - val newpats = map (fn subst => Envir.subst_term_types subst p) newsubsts - in - PolyPattern (p, instances, instancepats@newpats) - end - fun step monocs ths = - fold_rev (fn th => - fn (newmonos, ths) => - let - val (newmonos', th') = add monocs th - in - (newmonos'@newmonos, th'::ths) - end) - ths ([], []) - fun loop monocs pats ths = - let - val (monocs', ths') = step monocs ths - val pats' = map (addpats monocs) pats - in - if null (monocs') then - (pats', ths') - else - loop monocs' pats' ths' - end - val result = loop monocs pats ths - in - (!changed, result) - end - -datatype cthm = ComputeThm of term list * sort list * term - -fun thm2cthm th = - let - val {hyps, prop, shyps, ...} = Thm.rep_thm th - in - ComputeThm (hyps, shyps, prop) - end - -val cthm_ord' = prod_ord (prod_ord (list_ord Term_Ord.term_ord) (list_ord Term_Ord.sort_ord)) Term_Ord.term_ord - -fun cthm_ord (ComputeThm (h1, sh1, p1), ComputeThm (h2, sh2, p2)) = cthm_ord' (((h1,sh1), p1), ((h2, sh2), p2)) - -structure CThmtab = Table(type key = cthm val ord = cthm_ord) - -fun remove_duplicates ths = - let - val counter = Unsynchronized.ref 0 - val tab = Unsynchronized.ref (CThmtab.empty : unit CThmtab.table) - val thstab = Unsynchronized.ref (Inttab.empty : thm Inttab.table) - fun update th = - let - val key = thm2cthm th - in - case CThmtab.lookup (!tab) key of - NONE => ((tab := CThmtab.update_new (key, ()) (!tab)); thstab := Inttab.update_new (!counter, th) (!thstab); counter := !counter + 1) - | _ => () - end - val _ = map update ths - in - map snd (Inttab.dest (!thstab)) - end - -fun make_with_cache machine thy pats ths cs = - let - val ths = remove_duplicates ths - val (monocs, ths) = fold_rev (fn th => - fn (monocs, ths) => - let val (m, t) = create_theorem th in - (m@monocs, t::ths) - end) - ths (cs, []) - val pats = map create_pattern pats - val (_, (pats, ths)) = add_monos thy monocs pats ths - val computer = create_computer machine thy pats ths - in - PComputer (Theory.check_thy thy, computer, Unsynchronized.ref ths, Unsynchronized.ref pats) - end - -fun make machine thy ths cs = make_with_cache machine thy [] ths cs - -fun add_instances (PComputer (thyref, computer, rths, rpats)) cs = - let - val thy = Theory.deref thyref - val (changed, (pats, ths)) = add_monos thy cs (!rpats) (!rths) - in - if changed then - (update_computer computer pats ths; - rths := ths; - rpats := pats; - true) - else - false - - end - -fun add_instances' pc ts = add_instances pc (Linker.collect_consts ts) - -fun rewrite pc cts = - let - val _ = add_instances' pc (map term_of cts) - val computer = (computer_of pc) - in - map (fn ct => Compute.rewrite computer ct) cts - end - -fun simplify pc th = Compute.simplify (computer_of pc) th - -fun make_theorem pc th vars = - let - val _ = add_instances' pc [prop_of th] - - in - Compute.make_theorem (computer_of pc) th vars - end - -fun instantiate pc insts th = - let - val _ = add_instances' pc (map (term_of o snd) insts) - in - Compute.instantiate (computer_of pc) insts th - end - -fun evaluate_prem pc prem_no th = Compute.evaluate_prem (computer_of pc) prem_no th - -fun modus_ponens pc prem_no th' th = - let - val _ = add_instances' pc [prop_of th'] - in - Compute.modus_ponens (computer_of pc) prem_no th' th - end - - -end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/Compute_Oracle/report.ML --- a/src/HOL/Matrix/Compute_Oracle/report.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,33 +0,0 @@ -structure Report = -struct - -local - - val report_depth = Unsynchronized.ref 0 - fun space n = if n <= 0 then "" else (space (n-1))^" " - fun report_space () = space (!report_depth) - -in - -fun timeit f = - let - val t1 = Timing.start () - val x = f () - val t2 = Timing.message (Timing.result t1) - val _ = writeln ((report_space ()) ^ "--> "^t2) - in - x - end - -fun report s f = -let - val _ = writeln ((report_space ())^s) - val _ = report_depth := !report_depth + 1 - val x = timeit f - val _ = report_depth := !report_depth - 1 -in - x -end - -end -end \ No newline at end of file diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/Cplex.thy --- a/src/HOL/Matrix/Cplex.thy Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,67 +0,0 @@ -(* Title: HOL/Matrix/Cplex.thy - Author: Steven Obua -*) - -theory Cplex -imports SparseMatrix LP ComputeFloat ComputeNumeral -uses "Cplex_tools.ML" "CplexMatrixConverter.ML" "FloatSparseMatrixBuilder.ML" - "fspmlp.ML" ("matrixlp.ML") -begin - -lemma spm_mult_le_dual_prts: - assumes - "sorted_sparse_matrix A1" - "sorted_sparse_matrix A2" - "sorted_sparse_matrix c1" - "sorted_sparse_matrix c2" - "sorted_sparse_matrix y" - "sorted_sparse_matrix r1" - "sorted_sparse_matrix r2" - "sorted_spvec b" - "le_spmat [] y" - "sparse_row_matrix A1 \ A" - "A \ sparse_row_matrix A2" - "sparse_row_matrix c1 \ c" - "c \ sparse_row_matrix c2" - "sparse_row_matrix r1 \ x" - "x \ sparse_row_matrix r2" - "A * x \ sparse_row_matrix (b::('a::lattice_ring) spmat)" - shows - "c * x \ sparse_row_matrix (add_spmat (mult_spmat y b) - (let s1 = diff_spmat c1 (mult_spmat y A2); s2 = diff_spmat c2 (mult_spmat y A1) in - add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2)) (add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2)) - (add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1)) (mult_spmat (nprt_spmat s1) (nprt_spmat r1))))))" - apply (simp add: Let_def) - apply (insert assms) - apply (simp add: sparse_row_matrix_op_simps algebra_simps) - apply (rule mult_le_dual_prts[where A=A, simplified Let_def algebra_simps]) - apply (auto) - done - -lemma spm_mult_le_dual_prts_no_let: - assumes - "sorted_sparse_matrix A1" - "sorted_sparse_matrix A2" - "sorted_sparse_matrix c1" - "sorted_sparse_matrix c2" - "sorted_sparse_matrix y" - "sorted_sparse_matrix r1" - "sorted_sparse_matrix r2" - "sorted_spvec b" - "le_spmat [] y" - "sparse_row_matrix A1 \ A" - "A \ sparse_row_matrix A2" - "sparse_row_matrix c1 \ c" - "c \ sparse_row_matrix c2" - "sparse_row_matrix r1 \ x" - "x \ sparse_row_matrix r2" - "A * x \ sparse_row_matrix (b::('a::lattice_ring) spmat)" - shows - "c * x \ sparse_row_matrix (add_spmat (mult_spmat y b) - (mult_est_spmat r1 r2 (diff_spmat c1 (mult_spmat y A2)) (diff_spmat c2 (mult_spmat y A1))))" - by (simp add: assms mult_est_spmat_def spm_mult_le_dual_prts[where A=A, simplified Let_def]) - -use "matrixlp.ML" - -end - diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/CplexMatrixConverter.ML --- a/src/HOL/Matrix/CplexMatrixConverter.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,128 +0,0 @@ -(* Title: HOL/Matrix/CplexMatrixConverter.ML - Author: Steven Obua -*) - -signature MATRIX_BUILDER = -sig - type vector - type matrix - - val empty_vector : vector - val empty_matrix : matrix - - exception Nat_expected of int - val set_elem : vector -> int -> string -> vector - val set_vector : matrix -> int -> vector -> matrix -end; - -signature CPLEX_MATRIX_CONVERTER = -sig - structure cplex : CPLEX - structure matrix_builder : MATRIX_BUILDER - type vector = matrix_builder.vector - type matrix = matrix_builder.matrix - type naming = int * (int -> string) * (string -> int) - - exception Converter of string - - (* program must fulfill is_normed_cplexProg and must be an element of the image of elim_nonfree_bounds *) - (* convert_prog maximize c A b naming *) - val convert_prog : cplex.cplexProg -> bool * vector * matrix * vector * naming - - (* results must be optimal, converts_results returns the optimal value as string and the solution as vector *) - (* convert_results results name2index *) - val convert_results : cplex.cplexResult -> (string -> int) -> string * vector -end; - -functor MAKE_CPLEX_MATRIX_CONVERTER (structure cplex: CPLEX and matrix_builder: MATRIX_BUILDER) : CPLEX_MATRIX_CONVERTER = -struct - -structure cplex = cplex -structure matrix_builder = matrix_builder -type matrix = matrix_builder.matrix -type vector = matrix_builder.vector -type naming = int * (int -> string) * (string -> int) - -open matrix_builder -open cplex - -exception Converter of string; - -fun neg_term (cplexNeg t) = t - | neg_term (cplexSum ts) = cplexSum (map neg_term ts) - | neg_term t = cplexNeg t - -fun convert_prog (cplexProg (_, goal, constrs, bounds)) = - let - fun build_naming index i2s s2i [] = (index, i2s, s2i) - | build_naming index i2s s2i (cplexBounds (cplexNeg cplexInf, cplexLeq, cplexVar v, cplexLeq, cplexInf)::bounds) - = build_naming (index+1) (Inttab.update (index, v) i2s) (Symtab.update_new (v, index) s2i) bounds - | build_naming _ _ _ _ = raise (Converter "nonfree bound") - - val (varcount, i2s_tab, s2i_tab) = build_naming 0 Inttab.empty Symtab.empty bounds - - fun i2s i = case Inttab.lookup i2s_tab i of NONE => raise (Converter "index not found") - | SOME n => n - fun s2i s = case Symtab.lookup s2i_tab s of NONE => raise (Converter ("name not found: "^s)) - | SOME i => i - fun num2str positive (cplexNeg t) = num2str (not positive) t - | num2str positive (cplexNum num) = if positive then num else "-"^num - | num2str _ _ = raise (Converter "term is not a (possibly signed) number") - - fun setprod vec positive (cplexNeg t) = setprod vec (not positive) t - | setprod vec positive (cplexVar v) = set_elem vec (s2i v) (if positive then "1" else "-1") - | setprod vec positive (cplexProd (cplexNum num, cplexVar v)) = - set_elem vec (s2i v) (if positive then num else "-"^num) - | setprod _ _ _ = raise (Converter "term is not a normed product") - - fun sum2vec (cplexSum ts) = fold (fn t => fn vec => setprod vec true t) ts empty_vector - | sum2vec t = setprod empty_vector true t - - fun constrs2Ab j A b [] = (A, b) - | constrs2Ab j A b ((_, cplexConstr (cplexLeq, (t1,t2)))::cs) = - constrs2Ab (j+1) (set_vector A j (sum2vec t1)) (set_elem b j (num2str true t2)) cs - | constrs2Ab j A b ((_, cplexConstr (cplexGeq, (t1,t2)))::cs) = - constrs2Ab (j+1) (set_vector A j (sum2vec (neg_term t1))) (set_elem b j (num2str true (neg_term t2))) cs - | constrs2Ab j A b ((_, cplexConstr (cplexEq, (t1,t2)))::cs) = - constrs2Ab j A b ((NONE, cplexConstr (cplexLeq, (t1,t2))):: - (NONE, cplexConstr (cplexGeq, (t1, t2)))::cs) - | constrs2Ab _ _ _ _ = raise (Converter "no strict constraints allowed") - - val (A, b) = constrs2Ab 0 empty_matrix empty_vector constrs - - val (goal_maximize, goal_term) = - case goal of - (cplexMaximize t) => (true, t) - | (cplexMinimize t) => (false, t) - in - (goal_maximize, sum2vec goal_term, A, b, (varcount, i2s, s2i)) - end - -fun convert_results (cplex.Optimal (opt, entries)) name2index = - let - fun setv (name, value) v = matrix_builder.set_elem v (name2index name) value - in - (opt, fold setv entries (matrix_builder.empty_vector)) - end - | convert_results _ _ = raise (Converter "No optimal result") - -end; - -structure SimpleMatrixBuilder : MATRIX_BUILDER = -struct -type vector = (int * string) list -type matrix = (int * vector) list - -val empty_matrix = [] -val empty_vector = [] - -exception Nat_expected of int; - -fun set_elem v i s = v @ [(i, s)] - -fun set_vector m i v = m @ [(i, v)] - -end; - -structure SimpleCplexMatrixConverter = - MAKE_CPLEX_MATRIX_CONVERTER(structure cplex = Cplex and matrix_builder = SimpleMatrixBuilder); diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/Cplex_tools.ML --- a/src/HOL/Matrix/Cplex_tools.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1192 +0,0 @@ -(* Title: HOL/Matrix/Cplex_tools.ML - Author: Steven Obua -*) - -signature CPLEX = -sig - - datatype cplexTerm = cplexVar of string | cplexNum of string | cplexInf - | cplexNeg of cplexTerm - | cplexProd of cplexTerm * cplexTerm - | cplexSum of (cplexTerm list) - - datatype cplexComp = cplexLe | cplexLeq | cplexEq | cplexGe | cplexGeq - - datatype cplexGoal = cplexMinimize of cplexTerm - | cplexMaximize of cplexTerm - - datatype cplexConstr = cplexConstr of cplexComp * - (cplexTerm * cplexTerm) - - datatype cplexBounds = cplexBounds of cplexTerm * cplexComp * cplexTerm - * cplexComp * cplexTerm - | cplexBound of cplexTerm * cplexComp * cplexTerm - - datatype cplexProg = cplexProg of string - * cplexGoal - * ((string option * cplexConstr) - list) - * cplexBounds list - - datatype cplexResult = Unbounded - | Infeasible - | Undefined - | Optimal of string * - (((* name *) string * - (* value *) string) list) - - datatype cplexSolver = SOLVER_DEFAULT | SOLVER_CPLEX | SOLVER_GLPK - - exception Load_cplexFile of string - exception Load_cplexResult of string - exception Save_cplexFile of string - exception Execute of string - - val load_cplexFile : string -> cplexProg - - val save_cplexFile : string -> cplexProg -> unit - - val elim_nonfree_bounds : cplexProg -> cplexProg - - val relax_strict_ineqs : cplexProg -> cplexProg - - val is_normed_cplexProg : cplexProg -> bool - - val get_solver : unit -> cplexSolver - val set_solver : cplexSolver -> unit - val solve : cplexProg -> cplexResult -end; - -structure Cplex : CPLEX = -struct - -datatype cplexSolver = SOLVER_DEFAULT | SOLVER_CPLEX | SOLVER_GLPK - -val cplexsolver = Unsynchronized.ref SOLVER_DEFAULT; -fun get_solver () = !cplexsolver; -fun set_solver s = (cplexsolver := s); - -exception Load_cplexFile of string; -exception Load_cplexResult of string; -exception Save_cplexFile of string; - -datatype cplexTerm = cplexVar of string - | cplexNum of string - | cplexInf - | cplexNeg of cplexTerm - | cplexProd of cplexTerm * cplexTerm - | cplexSum of (cplexTerm list) -datatype cplexComp = cplexLe | cplexLeq | cplexEq | cplexGe | cplexGeq -datatype cplexGoal = cplexMinimize of cplexTerm | cplexMaximize of cplexTerm -datatype cplexConstr = cplexConstr of cplexComp * (cplexTerm * cplexTerm) -datatype cplexBounds = cplexBounds of cplexTerm * cplexComp * cplexTerm - * cplexComp * cplexTerm - | cplexBound of cplexTerm * cplexComp * cplexTerm -datatype cplexProg = cplexProg of string - * cplexGoal - * ((string option * cplexConstr) list) - * cplexBounds list - -fun rev_cmp cplexLe = cplexGe - | rev_cmp cplexLeq = cplexGeq - | rev_cmp cplexGe = cplexLe - | rev_cmp cplexGeq = cplexLeq - | rev_cmp cplexEq = cplexEq - -fun the NONE = raise (Load_cplexFile "SOME expected") - | the (SOME x) = x; - -fun modulo_signed is_something (cplexNeg u) = is_something u - | modulo_signed is_something u = is_something u - -fun is_Num (cplexNum _) = true - | is_Num _ = false - -fun is_Inf cplexInf = true - | is_Inf _ = false - -fun is_Var (cplexVar _) = true - | is_Var _ = false - -fun is_Neg (cplexNeg _) = true - | is_Neg _ = false - -fun is_normed_Prod (cplexProd (t1, t2)) = - (is_Num t1) andalso (is_Var t2) - | is_normed_Prod x = is_Var x - -fun is_normed_Sum (cplexSum ts) = - (ts <> []) andalso forall (modulo_signed is_normed_Prod) ts - | is_normed_Sum x = modulo_signed is_normed_Prod x - -fun is_normed_Constr (cplexConstr (_, (t1, t2))) = - (is_normed_Sum t1) andalso (modulo_signed is_Num t2) - -fun is_Num_or_Inf x = is_Inf x orelse is_Num x - -fun is_normed_Bounds (cplexBounds (t1, c1, t2, c2, t3)) = - (c1 = cplexLe orelse c1 = cplexLeq) andalso - (c2 = cplexLe orelse c2 = cplexLeq) andalso - is_Var t2 andalso - modulo_signed is_Num_or_Inf t1 andalso - modulo_signed is_Num_or_Inf t3 - | is_normed_Bounds (cplexBound (t1, c, t2)) = - (is_Var t1 andalso (modulo_signed is_Num_or_Inf t2)) - orelse - (c <> cplexEq andalso - is_Var t2 andalso (modulo_signed is_Num_or_Inf t1)) - -fun term_of_goal (cplexMinimize x) = x - | term_of_goal (cplexMaximize x) = x - -fun is_normed_cplexProg (cplexProg (_, goal, constraints, bounds)) = - is_normed_Sum (term_of_goal goal) andalso - forall (fn (_,x) => is_normed_Constr x) constraints andalso - forall is_normed_Bounds bounds - -fun is_NL s = s = "\n" - -fun is_blank s = forall (fn c => c <> #"\n" andalso Char.isSpace c) (String.explode s) - -fun is_num a = - let - val b = String.explode a - fun num4 cs = forall Char.isDigit cs - fun num3 [] = true - | num3 (ds as (c::cs)) = - if c = #"+" orelse c = #"-" then - num4 cs - else - num4 ds - fun num2 [] = true - | num2 (c::cs) = - if c = #"e" orelse c = #"E" then num3 cs - else (Char.isDigit c) andalso num2 cs - fun num1 [] = true - | num1 (c::cs) = - if c = #"." then num2 cs - else if c = #"e" orelse c = #"E" then num3 cs - else (Char.isDigit c) andalso num1 cs - fun num [] = true - | num (c::cs) = - if c = #"." then num2 cs - else (Char.isDigit c) andalso num1 cs - in - num b - end - -fun is_delimiter s = s = "+" orelse s = "-" orelse s = ":" - -fun is_cmp s = s = "<" orelse s = ">" orelse s = "<=" - orelse s = ">=" orelse s = "=" - -fun is_symbol a = - let - val symbol_char = String.explode "!\"#$%&()/,.;?@_`'{}|~" - fun is_symbol_char c = Char.isAlphaNum c orelse - exists (fn d => d=c) symbol_char - fun is_symbol_start c = is_symbol_char c andalso - not (Char.isDigit c) andalso - not (c= #".") - val b = String.explode a - in - b <> [] andalso is_symbol_start (hd b) andalso - forall is_symbol_char b - end - -fun to_upper s = String.implode (map Char.toUpper (String.explode s)) - -fun keyword x = - let - val a = to_upper x - in - if a = "BOUNDS" orelse a = "BOUND" then - SOME "BOUNDS" - else if a = "MINIMIZE" orelse a = "MINIMUM" orelse a = "MIN" then - SOME "MINIMIZE" - else if a = "MAXIMIZE" orelse a = "MAXIMUM" orelse a = "MAX" then - SOME "MAXIMIZE" - else if a = "ST" orelse a = "S.T." orelse a = "ST." then - SOME "ST" - else if a = "FREE" orelse a = "END" then - SOME a - else if a = "GENERAL" orelse a = "GENERALS" orelse a = "GEN" then - SOME "GENERAL" - else if a = "INTEGER" orelse a = "INTEGERS" orelse a = "INT" then - SOME "INTEGER" - else if a = "BINARY" orelse a = "BINARIES" orelse a = "BIN" then - SOME "BINARY" - else if a = "INF" orelse a = "INFINITY" then - SOME "INF" - else - NONE - end - -val TOKEN_ERROR = ~1 -val TOKEN_BLANK = 0 -val TOKEN_NUM = 1 -val TOKEN_DELIMITER = 2 -val TOKEN_SYMBOL = 3 -val TOKEN_LABEL = 4 -val TOKEN_CMP = 5 -val TOKEN_KEYWORD = 6 -val TOKEN_NL = 7 - -(* tokenize takes a list of chars as argument and returns a list of - int * string pairs, each string representing a "cplex token", - and each int being one of TOKEN_NUM, TOKEN_DELIMITER, TOKEN_CMP - or TOKEN_SYMBOL *) -fun tokenize s = - let - val flist = [(is_NL, TOKEN_NL), - (is_blank, TOKEN_BLANK), - (is_num, TOKEN_NUM), - (is_delimiter, TOKEN_DELIMITER), - (is_cmp, TOKEN_CMP), - (is_symbol, TOKEN_SYMBOL)] - fun match_helper [] s = (fn _ => false, TOKEN_ERROR) - | match_helper (f::fs) s = - if ((fst f) s) then f else match_helper fs s - fun match s = match_helper flist s - fun tok s = - if s = "" then [] else - let - val h = String.substring (s,0,1) - val (f, j) = match h - fun len i = - if size s = i then i - else if f (String.substring (s,0,i+1)) then - len (i+1) - else i - in - if j < 0 then - (if h = "\\" then [] - else raise (Load_cplexFile ("token expected, found: " - ^s))) - else - let - val l = len 1 - val u = String.substring (s,0,l) - val v = String.extract (s,l,NONE) - in - if j = 0 then tok v else (j, u) :: tok v - end - end - in - tok s - end - -exception Tokenize of string; - -fun tokenize_general flist s = - let - fun match_helper [] s = raise (Tokenize s) - | match_helper (f::fs) s = - if ((fst f) s) then f else match_helper fs s - fun match s = match_helper flist s - fun tok s = - if s = "" then [] else - let - val h = String.substring (s,0,1) - val (f, j) = match h - fun len i = - if size s = i then i - else if f (String.substring (s,0,i+1)) then - len (i+1) - else i - val l = len 1 - in - (j, String.substring (s,0,l)) :: tok (String.extract (s,l,NONE)) - end - in - tok s - end - -fun load_cplexFile name = - let - val f = TextIO.openIn name - val ignore_NL = Unsynchronized.ref true - val rest = Unsynchronized.ref [] - - fun is_symbol s c = (fst c) = TOKEN_SYMBOL andalso (to_upper (snd c)) = s - - fun readToken_helper () = - if length (!rest) > 0 then - let val u = hd (!rest) in - ( - rest := tl (!rest); - SOME u - ) - end - else - (case TextIO.inputLine f of - NONE => NONE - | SOME s => - let val t = tokenize s in - if (length t >= 2 andalso - snd(hd (tl t)) = ":") - then - rest := (TOKEN_LABEL, snd (hd t)) :: (tl (tl t)) - else if (length t >= 2) andalso is_symbol "SUBJECT" (hd (t)) - andalso is_symbol "TO" (hd (tl t)) - then - rest := (TOKEN_SYMBOL, "ST") :: (tl (tl t)) - else - rest := t; - readToken_helper () - end) - - fun readToken_helper2 () = - let val c = readToken_helper () in - if c = NONE then NONE - else if !ignore_NL andalso fst (the c) = TOKEN_NL then - readToken_helper2 () - else if fst (the c) = TOKEN_SYMBOL - andalso keyword (snd (the c)) <> NONE - then SOME (TOKEN_KEYWORD, the (keyword (snd (the c)))) - else c - end - - fun readToken () = readToken_helper2 () - - fun pushToken a = rest := (a::(!rest)) - - fun is_value token = - fst token = TOKEN_NUM orelse (fst token = TOKEN_KEYWORD - andalso snd token = "INF") - - fun get_value token = - if fst token = TOKEN_NUM then - cplexNum (snd token) - else if fst token = TOKEN_KEYWORD andalso snd token = "INF" - then - cplexInf - else - raise (Load_cplexFile "num expected") - - fun readTerm_Product only_num = - let val c = readToken () in - if c = NONE then NONE - else if fst (the c) = TOKEN_SYMBOL - then ( - if only_num then (pushToken (the c); NONE) - else SOME (cplexVar (snd (the c))) - ) - else if only_num andalso is_value (the c) then - SOME (get_value (the c)) - else if is_value (the c) then - let val t1 = get_value (the c) - val d = readToken () - in - if d = NONE then SOME t1 - else if fst (the d) = TOKEN_SYMBOL then - SOME (cplexProd (t1, cplexVar (snd (the d)))) - else - (pushToken (the d); SOME t1) - end - else (pushToken (the c); NONE) - end - - fun readTerm_Signed only_signed only_num = - let - val c = readToken () - in - if c = NONE then NONE - else - let val d = the c in - if d = (TOKEN_DELIMITER, "+") then - readTerm_Product only_num - else if d = (TOKEN_DELIMITER, "-") then - SOME (cplexNeg (the (readTerm_Product - only_num))) - else (pushToken d; - if only_signed then NONE - else readTerm_Product only_num) - end - end - - fun readTerm_Sum first_signed = - let val c = readTerm_Signed first_signed false in - if c = NONE then [] else (the c)::(readTerm_Sum true) - end - - fun readTerm () = - let val c = readTerm_Sum false in - if c = [] then NONE - else if tl c = [] then SOME (hd c) - else SOME (cplexSum c) - end - - fun readLabeledTerm () = - let val c = readToken () in - if c = NONE then (NONE, NONE) - else if fst (the c) = TOKEN_LABEL then - let val t = readTerm () in - if t = NONE then - raise (Load_cplexFile ("term after label "^ - (snd (the c))^ - " expected")) - else (SOME (snd (the c)), t) - end - else (pushToken (the c); (NONE, readTerm ())) - end - - fun readGoal () = - let - val g = readToken () - in - if g = SOME (TOKEN_KEYWORD, "MAXIMIZE") then - cplexMaximize (the (snd (readLabeledTerm ()))) - else if g = SOME (TOKEN_KEYWORD, "MINIMIZE") then - cplexMinimize (the (snd (readLabeledTerm ()))) - else raise (Load_cplexFile "MAXIMIZE or MINIMIZE expected") - end - - fun str2cmp b = - (case b of - "<" => cplexLe - | "<=" => cplexLeq - | ">" => cplexGe - | ">=" => cplexGeq - | "=" => cplexEq - | _ => raise (Load_cplexFile (b^" is no TOKEN_CMP"))) - - fun readConstraint () = - let - val t = readLabeledTerm () - fun make_constraint b t1 t2 = - cplexConstr - (str2cmp b, - (t1, t2)) - in - if snd t = NONE then NONE - else - let val c = readToken () in - if c = NONE orelse fst (the c) <> TOKEN_CMP - then raise (Load_cplexFile "TOKEN_CMP expected") - else - let val n = readTerm_Signed false true in - if n = NONE then - raise (Load_cplexFile "num expected") - else - SOME (fst t, - make_constraint (snd (the c)) - (the (snd t)) - (the n)) - end - end - end - - fun readST () = - let - fun readbody () = - let val t = readConstraint () in - if t = NONE then [] - else if (is_normed_Constr (snd (the t))) then - (the t)::(readbody ()) - else if (fst (the t) <> NONE) then - raise (Load_cplexFile - ("constraint '"^(the (fst (the t)))^ - "'is not normed")) - else - raise (Load_cplexFile - "constraint is not normed") - end - in - if readToken () = SOME (TOKEN_KEYWORD, "ST") - then - readbody () - else - raise (Load_cplexFile "ST expected") - end - - fun readCmp () = - let val c = readToken () in - if c = NONE then NONE - else if fst (the c) = TOKEN_CMP then - SOME (str2cmp (snd (the c))) - else (pushToken (the c); NONE) - end - - fun skip_NL () = - let val c = readToken () in - if c <> NONE andalso fst (the c) = TOKEN_NL then - skip_NL () - else - (pushToken (the c); ()) - end - - fun make_bounds c t1 t2 = - cplexBound (t1, c, t2) - - fun readBound () = - let - val _ = skip_NL () - val t1 = readTerm () - in - if t1 = NONE then NONE - else - let - val c1 = readCmp () - in - if c1 = NONE then - let - val c = readToken () - in - if c = SOME (TOKEN_KEYWORD, "FREE") then - SOME ( - cplexBounds (cplexNeg cplexInf, - cplexLeq, - the t1, - cplexLeq, - cplexInf)) - else - raise (Load_cplexFile "FREE expected") - end - else - let - val t2 = readTerm () - in - if t2 = NONE then - raise (Load_cplexFile "term expected") - else - let val c2 = readCmp () in - if c2 = NONE then - SOME (make_bounds (the c1) - (the t1) - (the t2)) - else - SOME ( - cplexBounds (the t1, - the c1, - the t2, - the c2, - the (readTerm()))) - end - end - end - end - - fun readBounds () = - let - fun makestring _ = "?" - fun readbody () = - let - val b = readBound () - in - if b = NONE then [] - else if (is_normed_Bounds (the b)) then - (the b)::(readbody()) - else ( - raise (Load_cplexFile - ("bounds are not normed in: "^ - (makestring (the b))))) - end - in - if readToken () = SOME (TOKEN_KEYWORD, "BOUNDS") then - readbody () - else raise (Load_cplexFile "BOUNDS expected") - end - - fun readEnd () = - if readToken () = SOME (TOKEN_KEYWORD, "END") then () - else raise (Load_cplexFile "END expected") - - val result_Goal = readGoal () - val result_ST = readST () - val _ = ignore_NL := false - val result_Bounds = readBounds () - val _ = ignore_NL := true - val _ = readEnd () - val _ = TextIO.closeIn f - in - cplexProg (name, result_Goal, result_ST, result_Bounds) - end - -fun save_cplexFile filename (cplexProg (_, goal, constraints, bounds)) = - let - val f = TextIO.openOut filename - - fun basic_write s = TextIO.output(f, s) - - val linebuf = Unsynchronized.ref "" - fun buf_flushline s = - (basic_write (!linebuf); - basic_write "\n"; - linebuf := s) - fun buf_add s = linebuf := (!linebuf) ^ s - - fun write s = - if (String.size s) + (String.size (!linebuf)) >= 250 then - buf_flushline (" "^s) - else - buf_add s - - fun writeln s = (buf_add s; buf_flushline "") - - fun write_term (cplexVar x) = write x - | write_term (cplexNum x) = write x - | write_term cplexInf = write "inf" - | write_term (cplexProd (cplexNum "1", b)) = write_term b - | write_term (cplexProd (a, b)) = - (write_term a; write " "; write_term b) - | write_term (cplexNeg x) = (write " - "; write_term x) - | write_term (cplexSum ts) = write_terms ts - and write_terms [] = () - | write_terms (t::ts) = - (if (not (is_Neg t)) then write " + " else (); - write_term t; write_terms ts) - - fun write_goal (cplexMaximize term) = - (writeln "MAXIMIZE"; write_term term; writeln "") - | write_goal (cplexMinimize term) = - (writeln "MINIMIZE"; write_term term; writeln "") - - fun write_cmp cplexLe = write "<" - | write_cmp cplexLeq = write "<=" - | write_cmp cplexEq = write "=" - | write_cmp cplexGe = write ">" - | write_cmp cplexGeq = write ">=" - - fun write_constr (cplexConstr (cmp, (a,b))) = - (write_term a; - write " "; - write_cmp cmp; - write " "; - write_term b) - - fun write_constraints [] = () - | write_constraints (c::cs) = - (if (fst c <> NONE) - then - (write (the (fst c)); write ": ") - else - (); - write_constr (snd c); - writeln ""; - write_constraints cs) - - fun write_bounds [] = () - | write_bounds ((cplexBounds (t1,c1,t2,c2,t3))::bs) = - ((if t1 = cplexNeg cplexInf andalso t3 = cplexInf - andalso (c1 = cplexLeq orelse c1 = cplexLe) - andalso (c2 = cplexLeq orelse c2 = cplexLe) - then - (write_term t2; write " free") - else - (write_term t1; write " "; write_cmp c1; write " "; - write_term t2; write " "; write_cmp c2; write " "; - write_term t3) - ); writeln ""; write_bounds bs) - | write_bounds ((cplexBound (t1, c, t2)) :: bs) = - (write_term t1; write " "; - write_cmp c; write " "; - write_term t2; writeln ""; write_bounds bs) - - val _ = write_goal goal - val _ = (writeln ""; writeln "ST") - val _ = write_constraints constraints - val _ = (writeln ""; writeln "BOUNDS") - val _ = write_bounds bounds - val _ = (writeln ""; writeln "END") - val _ = TextIO.closeOut f - in - () - end - -fun norm_Constr (constr as cplexConstr (c, (t1, t2))) = - if not (modulo_signed is_Num t2) andalso - modulo_signed is_Num t1 - then - [cplexConstr (rev_cmp c, (t2, t1))] - else if (c = cplexLe orelse c = cplexLeq) andalso - (t1 = (cplexNeg cplexInf) orelse t2 = cplexInf) - then - [] - else if (c = cplexGe orelse c = cplexGeq) andalso - (t1 = cplexInf orelse t2 = cplexNeg cplexInf) - then - [] - else - [constr] - -fun bound2constr (cplexBounds (t1,c1,t2,c2,t3)) = - (norm_Constr(cplexConstr (c1, (t1, t2)))) - @ (norm_Constr(cplexConstr (c2, (t2, t3)))) - | bound2constr (cplexBound (t1, cplexEq, t2)) = - (norm_Constr(cplexConstr (cplexLeq, (t1, t2)))) - @ (norm_Constr(cplexConstr (cplexLeq, (t2, t1)))) - | bound2constr (cplexBound (t1, c1, t2)) = - norm_Constr(cplexConstr (c1, (t1,t2))) - -val emptyset = Symtab.empty - -fun singleton v = Symtab.update (v, ()) emptyset - -fun merge a b = Symtab.merge (op =) (a, b) - -fun mergemap f ts = fold (fn x => fn table => merge table (f x)) ts Symtab.empty - -fun diff a b = Symtab.fold (Symtab.delete_safe o fst) b a - -fun collect_vars (cplexVar v) = singleton v - | collect_vars (cplexNeg t) = collect_vars t - | collect_vars (cplexProd (t1, t2)) = - merge (collect_vars t1) (collect_vars t2) - | collect_vars (cplexSum ts) = mergemap collect_vars ts - | collect_vars _ = emptyset - -(* Eliminates all nonfree bounds from the linear program and produces an - equivalent program with only free bounds - IF for the input program P holds: is_normed_cplexProg P *) -fun elim_nonfree_bounds (cplexProg (name, goal, constraints, bounds)) = - let - fun collect_constr_vars (_, cplexConstr (_, (t1,_))) = - (collect_vars t1) - - val cvars = merge (collect_vars (term_of_goal goal)) - (mergemap collect_constr_vars constraints) - - fun collect_lower_bounded_vars - (cplexBounds (_, _, cplexVar v, _, _)) = - singleton v - | collect_lower_bounded_vars - (cplexBound (_, cplexLe, cplexVar v)) = - singleton v - | collect_lower_bounded_vars - (cplexBound (_, cplexLeq, cplexVar v)) = - singleton v - | collect_lower_bounded_vars - (cplexBound (cplexVar v, cplexGe,_)) = - singleton v - | collect_lower_bounded_vars - (cplexBound (cplexVar v, cplexGeq, _)) = - singleton v - | collect_lower_bounded_vars - (cplexBound (cplexVar v, cplexEq, _)) = - singleton v - | collect_lower_bounded_vars _ = emptyset - - val lvars = mergemap collect_lower_bounded_vars bounds - val positive_vars = diff cvars lvars - val zero = cplexNum "0" - - fun make_pos_constr v = - (NONE, cplexConstr (cplexGeq, ((cplexVar v), zero))) - - fun make_free_bound v = - cplexBounds (cplexNeg cplexInf, cplexLeq, - cplexVar v, cplexLeq, - cplexInf) - - val pos_constrs = rev (Symtab.fold - (fn (k, _) => cons (make_pos_constr k)) - positive_vars []) - val bound_constrs = map (pair NONE) - (maps bound2constr bounds) - val constraints' = constraints @ pos_constrs @ bound_constrs - val bounds' = rev (Symtab.fold (fn (v, _) => cons (make_free_bound v)) cvars []); - in - cplexProg (name, goal, constraints', bounds') - end - -fun relax_strict_ineqs (cplexProg (name, goals, constrs, bounds)) = - let - fun relax cplexLe = cplexLeq - | relax cplexGe = cplexGeq - | relax x = x - - fun relax_constr (n, cplexConstr(c, (t1, t2))) = - (n, cplexConstr(relax c, (t1, t2))) - - fun relax_bounds (cplexBounds (t1, c1, t2, c2, t3)) = - cplexBounds (t1, relax c1, t2, relax c2, t3) - | relax_bounds (cplexBound (t1, c, t2)) = - cplexBound (t1, relax c, t2) - in - cplexProg (name, - goals, - map relax_constr constrs, - map relax_bounds bounds) - end - -datatype cplexResult = Unbounded - | Infeasible - | Undefined - | Optimal of string * ((string * string) list) - -fun is_separator x = forall (fn c => c = #"-") (String.explode x) - -fun is_sign x = (x = "+" orelse x = "-") - -fun is_colon x = (x = ":") - -fun is_resultsymbol a = - let - val symbol_char = String.explode "!\"#$%&()/,.;?@_`'{}|~-" - fun is_symbol_char c = Char.isAlphaNum c orelse - exists (fn d => d=c) symbol_char - fun is_symbol_start c = is_symbol_char c andalso - not (Char.isDigit c) andalso - not (c= #".") andalso - not (c= #"-") - val b = String.explode a - in - b <> [] andalso is_symbol_start (hd b) andalso - forall is_symbol_char b - end - -val TOKEN_SIGN = 100 -val TOKEN_COLON = 101 -val TOKEN_SEPARATOR = 102 - -fun load_glpkResult name = - let - val flist = [(is_NL, TOKEN_NL), - (is_blank, TOKEN_BLANK), - (is_num, TOKEN_NUM), - (is_sign, TOKEN_SIGN), - (is_colon, TOKEN_COLON), - (is_cmp, TOKEN_CMP), - (is_resultsymbol, TOKEN_SYMBOL), - (is_separator, TOKEN_SEPARATOR)] - - val tokenize = tokenize_general flist - - val f = TextIO.openIn name - - val rest = Unsynchronized.ref [] - - fun readToken_helper () = - if length (!rest) > 0 then - let val u = hd (!rest) in - ( - rest := tl (!rest); - SOME u - ) - end - else - (case TextIO.inputLine f of - NONE => NONE - | SOME s => (rest := tokenize s; readToken_helper())) - - fun is_tt tok ty = (tok <> NONE andalso (fst (the tok)) = ty) - - fun pushToken a = if a = NONE then () else (rest := ((the a)::(!rest))) - - fun readToken () = - let val t = readToken_helper () in - if is_tt t TOKEN_BLANK then - readToken () - else if is_tt t TOKEN_NL then - let val t2 = readToken_helper () in - if is_tt t2 TOKEN_SIGN then - (pushToken (SOME (TOKEN_SEPARATOR, "-")); t) - else - (pushToken t2; t) - end - else if is_tt t TOKEN_SIGN then - let val t2 = readToken_helper () in - if is_tt t2 TOKEN_NUM then - (SOME (TOKEN_NUM, (snd (the t))^(snd (the t2)))) - else - (pushToken t2; t) - end - else - t - end - - fun readRestOfLine P = - let - val t = readToken () - in - if is_tt t TOKEN_NL orelse t = NONE - then P - else readRestOfLine P - end - - fun readHeader () = - let - fun readStatus () = readRestOfLine ("STATUS", snd (the (readToken ()))) - fun readObjective () = readRestOfLine ("OBJECTIVE", snd (the (readToken (); readToken (); readToken ()))) - val t1 = readToken () - val t2 = readToken () - in - if is_tt t1 TOKEN_SYMBOL andalso is_tt t2 TOKEN_COLON - then - case to_upper (snd (the t1)) of - "STATUS" => (readStatus ())::(readHeader ()) - | "OBJECTIVE" => (readObjective())::(readHeader ()) - | _ => (readRestOfLine (); readHeader ()) - else - (pushToken t2; pushToken t1; []) - end - - fun skip_until_sep () = - let val x = readToken () in - if is_tt x TOKEN_SEPARATOR then - readRestOfLine () - else - skip_until_sep () - end - - fun load_value () = - let - val t1 = readToken () - val t2 = readToken () - in - if is_tt t1 TOKEN_NUM andalso is_tt t2 TOKEN_SYMBOL then - let - val t = readToken () - val state = if is_tt t TOKEN_NL then readToken () else t - val _ = if is_tt state TOKEN_SYMBOL then () else raise (Load_cplexResult "state expected") - val k = readToken () - in - if is_tt k TOKEN_NUM then - readRestOfLine (SOME (snd (the t2), snd (the k))) - else - raise (Load_cplexResult "number expected") - end - else - (pushToken t2; pushToken t1; NONE) - end - - fun load_values () = - let val v = load_value () in - if v = NONE then [] else (the v)::(load_values ()) - end - - val header = readHeader () - - val result = - case AList.lookup (op =) header "STATUS" of - SOME "INFEASIBLE" => Infeasible - | SOME "UNBOUNDED" => Unbounded - | SOME "OPTIMAL" => Optimal (the (AList.lookup (op =) header "OBJECTIVE"), - (skip_until_sep (); - skip_until_sep (); - load_values ())) - | _ => Undefined - - val _ = TextIO.closeIn f - in - result - end - handle (Tokenize s) => raise (Load_cplexResult ("Tokenize: "^s)) - | Option => raise (Load_cplexResult "Option") - -fun load_cplexResult name = - let - val flist = [(is_NL, TOKEN_NL), - (is_blank, TOKEN_BLANK), - (is_num, TOKEN_NUM), - (is_sign, TOKEN_SIGN), - (is_colon, TOKEN_COLON), - (is_cmp, TOKEN_CMP), - (is_resultsymbol, TOKEN_SYMBOL)] - - val tokenize = tokenize_general flist - - val f = TextIO.openIn name - - val rest = Unsynchronized.ref [] - - fun readToken_helper () = - if length (!rest) > 0 then - let val u = hd (!rest) in - ( - rest := tl (!rest); - SOME u - ) - end - else - (case TextIO.inputLine f of - NONE => NONE - | SOME s => (rest := tokenize s; readToken_helper())) - - fun is_tt tok ty = (tok <> NONE andalso (fst (the tok)) = ty) - - fun pushToken a = if a = NONE then () else (rest := ((the a)::(!rest))) - - fun readToken () = - let val t = readToken_helper () in - if is_tt t TOKEN_BLANK then - readToken () - else if is_tt t TOKEN_SIGN then - let val t2 = readToken_helper () in - if is_tt t2 TOKEN_NUM then - (SOME (TOKEN_NUM, (snd (the t))^(snd (the t2)))) - else - (pushToken t2; t) - end - else - t - end - - fun readRestOfLine P = - let - val t = readToken () - in - if is_tt t TOKEN_NL orelse t = NONE - then P - else readRestOfLine P - end - - fun readHeader () = - let - fun readStatus () = readRestOfLine ("STATUS", snd (the (readToken ()))) - fun readObjective () = - let - val t = readToken () - in - if is_tt t TOKEN_SYMBOL andalso to_upper (snd (the t)) = "VALUE" then - readRestOfLine ("OBJECTIVE", snd (the (readToken()))) - else - readRestOfLine ("OBJECTIVE_NAME", snd (the t)) - end - - val t = readToken () - in - if is_tt t TOKEN_SYMBOL then - case to_upper (snd (the t)) of - "STATUS" => (readStatus ())::(readHeader ()) - | "OBJECTIVE" => (readObjective ())::(readHeader ()) - | "SECTION" => (pushToken t; []) - | _ => (readRestOfLine (); readHeader ()) - else - (readRestOfLine (); readHeader ()) - end - - fun skip_nls () = - let val x = readToken () in - if is_tt x TOKEN_NL then - skip_nls () - else - (pushToken x; ()) - end - - fun skip_paragraph () = - if is_tt (readToken ()) TOKEN_NL then - (if is_tt (readToken ()) TOKEN_NL then - skip_nls () - else - skip_paragraph ()) - else - skip_paragraph () - - fun load_value () = - let - val t1 = readToken () - val t1 = if is_tt t1 TOKEN_SYMBOL andalso snd (the t1) = "A" then readToken () else t1 - in - if is_tt t1 TOKEN_NUM then - let - val name = readToken () - val status = readToken () - val value = readToken () - in - if is_tt name TOKEN_SYMBOL andalso - is_tt status TOKEN_SYMBOL andalso - is_tt value TOKEN_NUM - then - readRestOfLine (SOME (snd (the name), snd (the value))) - else - raise (Load_cplexResult "column line expected") - end - else - (pushToken t1; NONE) - end - - fun load_values () = - let val v = load_value () in - if v = NONE then [] else (the v)::(load_values ()) - end - - val header = readHeader () - - val result = - case AList.lookup (op =) header "STATUS" of - SOME "INFEASIBLE" => Infeasible - | SOME "NONOPTIMAL" => Unbounded - | SOME "OPTIMAL" => Optimal (the (AList.lookup (op =) header "OBJECTIVE"), - (skip_paragraph (); - skip_paragraph (); - skip_paragraph (); - skip_paragraph (); - skip_paragraph (); - load_values ())) - | _ => Undefined - - val _ = TextIO.closeIn f - in - result - end - handle (Tokenize s) => raise (Load_cplexResult ("Tokenize: "^s)) - | Option => raise (Load_cplexResult "Option") - -exception Execute of string; - -fun tmp_file s = Path.implode (Path.expand (File.tmp_path (Path.basic s))); -fun wrap s = "\""^s^"\""; - -fun solve_glpk prog = - let - val name = string_of_int (Time.toMicroseconds (Time.now ())) - val lpname = tmp_file (name^".lp") - val resultname = tmp_file (name^".txt") - val _ = save_cplexFile lpname prog - val cplex_path = getenv "GLPK_PATH" - val cplex = if cplex_path = "" then "glpsol" else cplex_path - val command = (wrap cplex)^" --lpt "^(wrap lpname)^" --output "^(wrap resultname) - val answer = #1 (Isabelle_System.bash_output command) - in - let - val result = load_glpkResult resultname - val _ = OS.FileSys.remove lpname - val _ = OS.FileSys.remove resultname - in - result - end - handle (Load_cplexResult s) => raise (Execute ("Load_cplexResult: "^s^"\nExecute: "^answer)) - | _ => raise (Execute answer) (* FIXME avoid handle _ *) - end - -fun solve_cplex prog = - let - fun write_script s lp r = - let - val f = TextIO.openOut s - val _ = TextIO.output (f, "read\n"^lp^"\noptimize\nwrite\n"^r^"\nquit") - val _ = TextIO.closeOut f - in - () - end - - val name = string_of_int (Time.toMicroseconds (Time.now ())) - val lpname = tmp_file (name^".lp") - val resultname = tmp_file (name^".txt") - val scriptname = tmp_file (name^".script") - val _ = save_cplexFile lpname prog - val _ = write_script scriptname lpname resultname - in - let - val result = load_cplexResult resultname - val _ = OS.FileSys.remove lpname - val _ = OS.FileSys.remove resultname - val _ = OS.FileSys.remove scriptname - in - result - end - end - -fun solve prog = - case get_solver () of - SOLVER_DEFAULT => - (case getenv "LP_SOLVER" of - "CPLEX" => solve_cplex prog - | "GLPK" => solve_glpk prog - | _ => raise (Execute ("LP_SOLVER must be set to CPLEX or to GLPK"))) - | SOLVER_CPLEX => solve_cplex prog - | SOLVER_GLPK => solve_glpk prog - -end; diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/FloatSparseMatrixBuilder.ML --- a/src/HOL/Matrix/FloatSparseMatrixBuilder.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,284 +0,0 @@ -(* Title: HOL/Matrix/FloatSparseMatrixBuilder.ML - Author: Steven Obua -*) - -signature FLOAT_SPARSE_MATRIX_BUILDER = -sig - include MATRIX_BUILDER - - structure cplex : CPLEX - - type float = Float.float - val approx_value : int -> (float -> float) -> string -> term * term - val approx_vector : int -> (float -> float) -> vector -> term * term - val approx_matrix : int -> (float -> float) -> matrix -> term * term - - val mk_spvec_entry : int -> float -> term - val mk_spvec_entry' : int -> term -> term - val mk_spmat_entry : int -> term -> term - val spvecT: typ - val spmatT: typ - - val v_elem_at : vector -> int -> string option - val m_elem_at : matrix -> int -> vector option - val v_only_elem : vector -> int option - val v_fold : (int * string -> 'a -> 'a) -> vector -> 'a -> 'a - val m_fold : (int * vector -> 'a -> 'a) -> matrix -> 'a -> 'a - - val transpose_matrix : matrix -> matrix - - val cut_vector : int -> vector -> vector - val cut_matrix : vector -> int option -> matrix -> matrix - - val delete_matrix : int list -> matrix -> matrix - val cut_matrix' : int list -> matrix -> matrix - val delete_vector : int list -> vector -> vector - val cut_vector' : int list -> vector -> vector - - val indices_of_matrix : matrix -> int list - val indices_of_vector : vector -> int list - - (* cplexProg c A b *) - val cplexProg : vector -> matrix -> vector -> cplex.cplexProg * (string -> int) - (* dual_cplexProg c A b *) - val dual_cplexProg : vector -> matrix -> vector -> cplex.cplexProg * (string -> int) -end; - -structure FloatSparseMatrixBuilder : FLOAT_SPARSE_MATRIX_BUILDER = -struct - -type float = Float.float -structure Inttab = Table(type key = int val ord = rev_order o int_ord); - -type vector = string Inttab.table -type matrix = vector Inttab.table - -val spvec_elemT = HOLogic.mk_prodT (HOLogic.natT, HOLogic.realT); -val spvecT = HOLogic.listT spvec_elemT; -val spmat_elemT = HOLogic.mk_prodT (HOLogic.natT, spvecT); -val spmatT = HOLogic.listT spmat_elemT; - -fun approx_value prec f = - FloatArith.approx_float prec (fn (x, y) => (f x, f y)); - -fun mk_spvec_entry i f = - HOLogic.mk_prod (HOLogic.mk_number HOLogic.natT i, FloatArith.mk_float f); - -fun mk_spvec_entry' i x = - HOLogic.mk_prod (HOLogic.mk_number HOLogic.natT i, x); - -fun mk_spmat_entry i e = - HOLogic.mk_prod (HOLogic.mk_number HOLogic.natT i, e); - -fun approx_vector prec pprt vector = - let - fun app (index, s) (lower, upper) = - let - val (flower, fupper) = approx_value prec pprt s - val index = HOLogic.mk_number HOLogic.natT index - val elower = HOLogic.mk_prod (index, flower) - val eupper = HOLogic.mk_prod (index, fupper) - in (elower :: lower, eupper :: upper) end; - in - pairself (HOLogic.mk_list spvec_elemT) (Inttab.fold app vector ([], [])) - end; - -fun approx_matrix prec pprt vector = - let - fun app (index, v) (lower, upper) = - let - val (flower, fupper) = approx_vector prec pprt v - val index = HOLogic.mk_number HOLogic.natT index - val elower = HOLogic.mk_prod (index, flower) - val eupper = HOLogic.mk_prod (index, fupper) - in (elower :: lower, eupper :: upper) end; - in - pairself (HOLogic.mk_list spmat_elemT) (Inttab.fold app vector ([], [])) - end; - -exception Nat_expected of int; - -val zero_interval = approx_value 1 I "0" - -fun set_elem vector index str = - if index < 0 then - raise (Nat_expected index) - else if (approx_value 1 I str) = zero_interval then - vector - else - Inttab.update (index, str) vector - -fun set_vector matrix index vector = - if index < 0 then - raise (Nat_expected index) - else if Inttab.is_empty vector then - matrix - else - Inttab.update (index, vector) matrix - -val empty_matrix = Inttab.empty -val empty_vector = Inttab.empty - -(* dual stuff *) - -structure cplex = Cplex - -fun transpose_matrix matrix = - let - fun upd j (i, s) = - Inttab.map_default (i, Inttab.empty) (Inttab.update (j, s)); - fun updm (j, v) = Inttab.fold (upd j) v; - in Inttab.fold updm matrix empty_matrix end; - -exception No_name of string; - -exception Superfluous_constr_right_hand_sides - -fun cplexProg c A b = - let - val ytable = Unsynchronized.ref Inttab.empty - fun indexof s = - if String.size s = 0 then raise (No_name s) - else case Int.fromString (String.extract(s, 1, NONE)) of - SOME i => i | NONE => raise (No_name s) - - fun nameof i = - let - val s = "x" ^ string_of_int i - val _ = Unsynchronized.change ytable (Inttab.update (i, s)) - in - s - end - - fun split_numstr s = - if String.isPrefix "-" s then (false,String.extract(s, 1, NONE)) - else if String.isPrefix "+" s then (true, String.extract(s, 1, NONE)) - else (true, s) - - fun mk_term index s = - let - val (p, s) = split_numstr s - val prod = cplex.cplexProd (cplex.cplexNum s, cplex.cplexVar (nameof index)) - in - if p then prod else cplex.cplexNeg prod - end - - fun vec2sum vector = - cplex.cplexSum (Inttab.fold (fn (index, s) => fn list => (mk_term index s) :: list) vector []) - - fun mk_constr index vector c = - let - val s = case Inttab.lookup c index of SOME s => s | NONE => "0" - val (p, s) = split_numstr s - val num = if p then cplex.cplexNum s else cplex.cplexNeg (cplex.cplexNum s) - in - (NONE, cplex.cplexConstr (cplex.cplexLeq, (vec2sum vector, num))) - end - - fun delete index c = Inttab.delete index c handle Inttab.UNDEF _ => c - - val (list, b) = Inttab.fold - (fn (index, v) => fn (list, c) => ((mk_constr index v c)::list, delete index c)) - A ([], b) - val _ = if Inttab.is_empty b then () else raise Superfluous_constr_right_hand_sides - - fun mk_free y = cplex.cplexBounds (cplex.cplexNeg cplex.cplexInf, cplex.cplexLeq, - cplex.cplexVar y, cplex.cplexLeq, - cplex.cplexInf) - - val yvars = Inttab.fold (fn (_, y) => fn l => (mk_free y)::l) (!ytable) [] - - val prog = cplex.cplexProg ("original", cplex.cplexMaximize (vec2sum c), list, yvars) - in - (prog, indexof) - end - - -fun dual_cplexProg c A b = - let - fun indexof s = - if String.size s = 0 then raise (No_name s) - else case Int.fromString (String.extract(s, 1, NONE)) of - SOME i => i | NONE => raise (No_name s) - - fun nameof i = "y" ^ string_of_int i - - fun split_numstr s = - if String.isPrefix "-" s then (false,String.extract(s, 1, NONE)) - else if String.isPrefix "+" s then (true, String.extract(s, 1, NONE)) - else (true, s) - - fun mk_term index s = - let - val (p, s) = split_numstr s - val prod = cplex.cplexProd (cplex.cplexNum s, cplex.cplexVar (nameof index)) - in - if p then prod else cplex.cplexNeg prod - end - - fun vec2sum vector = - cplex.cplexSum (Inttab.fold (fn (index, s) => fn list => (mk_term index s)::list) vector []) - - fun mk_constr index vector c = - let - val s = case Inttab.lookup c index of SOME s => s | NONE => "0" - val (p, s) = split_numstr s - val num = if p then cplex.cplexNum s else cplex.cplexNeg (cplex.cplexNum s) - in - (NONE, cplex.cplexConstr (cplex.cplexEq, (vec2sum vector, num))) - end - - fun delete index c = Inttab.delete index c handle Inttab.UNDEF _ => c - - val (list, c) = Inttab.fold - (fn (index, v) => fn (list, c) => ((mk_constr index v c)::list, delete index c)) - (transpose_matrix A) ([], c) - val _ = if Inttab.is_empty c then () else raise Superfluous_constr_right_hand_sides - - val prog = cplex.cplexProg ("dual", cplex.cplexMinimize (vec2sum b), list, []) - in - (prog, indexof) - end - -fun cut_vector size v = - let - val count = Unsynchronized.ref 0; - fun app (i, s) = if (!count < size) then - (count := !count +1 ; Inttab.update (i, s)) - else I - in - Inttab.fold app v empty_vector - end - -fun cut_matrix vfilter vsize m = - let - fun app (i, v) = - if is_none (Inttab.lookup vfilter i) then I - else case vsize - of NONE => Inttab.update (i, v) - | SOME s => Inttab.update (i, cut_vector s v) - in Inttab.fold app m empty_matrix end - -fun v_elem_at v i = Inttab.lookup v i -fun m_elem_at m i = Inttab.lookup m i - -fun v_only_elem v = - case Inttab.min_key v of - NONE => NONE - | SOME vmin => (case Inttab.max_key v of - NONE => SOME vmin - | SOME vmax => if vmin = vmax then SOME vmin else NONE) - -fun v_fold f = Inttab.fold f; -fun m_fold f = Inttab.fold f; - -fun indices_of_vector v = Inttab.keys v -fun indices_of_matrix m = Inttab.keys m -fun delete_vector indices v = fold Inttab.delete indices v -fun delete_matrix indices m = fold Inttab.delete indices m -fun cut_matrix' indices _ = fold (fn i => fn m => (case Inttab.lookup m i of NONE => m | SOME v => Inttab.update (i, v) m)) indices Inttab.empty -fun cut_vector' indices _ = fold (fn i => fn v => (case Inttab.lookup v i of NONE => v | SOME x => Inttab.update (i, x) v)) indices Inttab.empty - - - -end; diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/LP.thy --- a/src/HOL/Matrix/LP.thy Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,164 +0,0 @@ -(* Title: HOL/Matrix/LP.thy - Author: Steven Obua -*) - -theory LP -imports Main "~~/src/HOL/Library/Lattice_Algebras" -begin - -lemma le_add_right_mono: - assumes - "a <= b + (c::'a::ordered_ab_group_add)" - "c <= d" - shows "a <= b + d" - apply (rule_tac order_trans[where y = "b+c"]) - apply (simp_all add: assms) - done - -lemma linprog_dual_estimate: - assumes - "A * x \ (b::'a::lattice_ring)" - "0 \ y" - "abs (A - A') \ \A" - "b \ b'" - "abs (c - c') \ \c" - "abs x \ r" - shows - "c * x \ y * b' + (y * \A + abs (y * A' - c') + \c) * r" -proof - - from assms have 1: "y * b <= y * b'" by (simp add: mult_left_mono) - from assms have 2: "y * (A * x) <= y * b" by (simp add: mult_left_mono) - have 3: "y * (A * x) = c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x" by (simp add: algebra_simps) - from 1 2 3 have 4: "c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x <= y * b'" by simp - have 5: "c * x <= y * b' + abs((y * (A - A') + (y * A' - c') + (c'-c)) * x)" - by (simp only: 4 estimate_by_abs) - have 6: "abs((y * (A - A') + (y * A' - c') + (c'-c)) * x) <= abs (y * (A - A') + (y * A' - c') + (c'-c)) * abs x" - by (simp add: abs_le_mult) - have 7: "(abs (y * (A - A') + (y * A' - c') + (c'-c))) * abs x <= (abs (y * (A-A') + (y*A'-c')) + abs(c'-c)) * abs x" - by(rule abs_triangle_ineq [THEN mult_right_mono]) simp - have 8: " (abs (y * (A-A') + (y*A'-c')) + abs(c'-c)) * abs x <= (abs (y * (A-A')) + abs (y*A'-c') + abs(c'-c)) * abs x" - by (simp add: abs_triangle_ineq mult_right_mono) - have 9: "(abs (y * (A-A')) + abs (y*A'-c') + abs(c'-c)) * abs x <= (abs y * abs (A-A') + abs (y*A'-c') + abs (c'-c)) * abs x" - by (simp add: abs_le_mult mult_right_mono) - have 10: "c'-c = -(c-c')" by (simp add: algebra_simps) - have 11: "abs (c'-c) = abs (c-c')" - by (subst 10, subst abs_minus_cancel, simp) - have 12: "(abs y * abs (A-A') + abs (y*A'-c') + abs (c'-c)) * abs x <= (abs y * abs (A-A') + abs (y*A'-c') + \c) * abs x" - by (simp add: 11 assms mult_right_mono) - have 13: "(abs y * abs (A-A') + abs (y*A'-c') + \c) * abs x <= (abs y * \A + abs (y*A'-c') + \c) * abs x" - by (simp add: assms mult_right_mono mult_left_mono) - have r: "(abs y * \A + abs (y*A'-c') + \c) * abs x <= (abs y * \A + abs (y*A'-c') + \c) * r" - apply (rule mult_left_mono) - apply (simp add: assms) - apply (rule_tac add_mono[of "0::'a" _ "0", simplified])+ - apply (rule mult_left_mono[of "0" "\A", simplified]) - apply (simp_all) - apply (rule order_trans[where y="abs (A-A')"], simp_all add: assms) - apply (rule order_trans[where y="abs (c-c')"], simp_all add: assms) - done - from 6 7 8 9 12 13 r have 14:" abs((y * (A - A') + (y * A' - c') + (c'-c)) * x) <=(abs y * \A + abs (y*A'-c') + \c) * r" - by (simp) - show ?thesis - apply (rule le_add_right_mono[of _ _ "abs((y * (A - A') + (y * A' - c') + (c'-c)) * x)"]) - apply (simp_all only: 5 14[simplified abs_of_nonneg[of y, simplified assms]]) - done -qed - -lemma le_ge_imp_abs_diff_1: - assumes - "A1 <= (A::'a::lattice_ring)" - "A <= A2" - shows "abs (A-A1) <= A2-A1" -proof - - have "0 <= A - A1" - proof - - have 1: "A - A1 = A + (- A1)" by simp - show ?thesis by (simp only: 1 add_right_mono[of A1 A "-A1", simplified, simplified assms]) - qed - then have "abs (A-A1) = A-A1" by (rule abs_of_nonneg) - with assms show "abs (A-A1) <= (A2-A1)" by simp -qed - -lemma mult_le_prts: - assumes - "a1 <= (a::'a::lattice_ring)" - "a <= a2" - "b1 <= b" - "b <= b2" - shows - "a * b <= pprt a2 * pprt b2 + pprt a1 * nprt b2 + nprt a2 * pprt b1 + nprt a1 * nprt b1" -proof - - have "a * b = (pprt a + nprt a) * (pprt b + nprt b)" - apply (subst prts[symmetric])+ - apply simp - done - then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b" - by (simp add: algebra_simps) - moreover have "pprt a * pprt b <= pprt a2 * pprt b2" - by (simp_all add: assms mult_mono) - moreover have "pprt a * nprt b <= pprt a1 * nprt b2" - proof - - have "pprt a * nprt b <= pprt a * nprt b2" - by (simp add: mult_left_mono assms) - moreover have "pprt a * nprt b2 <= pprt a1 * nprt b2" - by (simp add: mult_right_mono_neg assms) - ultimately show ?thesis - by simp - qed - moreover have "nprt a * pprt b <= nprt a2 * pprt b1" - proof - - have "nprt a * pprt b <= nprt a2 * pprt b" - by (simp add: mult_right_mono assms) - moreover have "nprt a2 * pprt b <= nprt a2 * pprt b1" - by (simp add: mult_left_mono_neg assms) - ultimately show ?thesis - by simp - qed - moreover have "nprt a * nprt b <= nprt a1 * nprt b1" - proof - - have "nprt a * nprt b <= nprt a * nprt b1" - by (simp add: mult_left_mono_neg assms) - moreover have "nprt a * nprt b1 <= nprt a1 * nprt b1" - by (simp add: mult_right_mono_neg assms) - ultimately show ?thesis - by simp - qed - ultimately show ?thesis - by - (rule add_mono | simp)+ -qed - -lemma mult_le_dual_prts: - assumes - "A * x \ (b::'a::lattice_ring)" - "0 \ y" - "A1 \ A" - "A \ A2" - "c1 \ c" - "c \ c2" - "r1 \ x" - "x \ r2" - shows - "c * x \ y * b + (let s1 = c1 - y * A2; s2 = c2 - y * A1 in pprt s2 * pprt r2 + pprt s1 * nprt r2 + nprt s2 * pprt r1 + nprt s1 * nprt r1)" - (is "_ <= _ + ?C") -proof - - from assms have "y * (A * x) <= y * b" by (simp add: mult_left_mono) - moreover have "y * (A * x) = c * x + (y * A - c) * x" by (simp add: algebra_simps) - ultimately have "c * x + (y * A - c) * x <= y * b" by simp - then have "c * x <= y * b - (y * A - c) * x" by (simp add: le_diff_eq) - then have cx: "c * x <= y * b + (c - y * A) * x" by (simp add: algebra_simps) - have s2: "c - y * A <= c2 - y * A1" - by (simp add: diff_minus assms add_mono mult_left_mono) - have s1: "c1 - y * A2 <= c - y * A" - by (simp add: diff_minus assms add_mono mult_left_mono) - have prts: "(c - y * A) * x <= ?C" - apply (simp add: Let_def) - apply (rule mult_le_prts) - apply (simp_all add: assms s1 s2) - done - then have "y * b + (c - y * A) * x <= y * b + ?C" - by simp - with cx show ?thesis - by(simp only:) -qed - -end \ No newline at end of file diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/Matrix.thy --- a/src/HOL/Matrix/Matrix.thy Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1836 +0,0 @@ -(* Title: HOL/Matrix/Matrix.thy - Author: Steven Obua -*) - -theory Matrix -imports Main "~~/src/HOL/Library/Lattice_Algebras" -begin - -type_synonym 'a infmatrix = "nat \ nat \ 'a" - -definition nonzero_positions :: "(nat \ nat \ 'a::zero) \ (nat \ nat) set" where - "nonzero_positions A = {pos. A (fst pos) (snd pos) ~= 0}" - -definition "matrix = {(f::(nat \ nat \ 'a::zero)). finite (nonzero_positions f)}" - -typedef (open) 'a matrix = "matrix :: (nat \ nat \ 'a::zero) set" - unfolding matrix_def -proof - show "(\j i. 0) \ {(f::(nat \ nat \ 'a::zero)). finite (nonzero_positions f)}" - by (simp add: nonzero_positions_def) -qed - -declare Rep_matrix_inverse[simp] - -lemma finite_nonzero_positions : "finite (nonzero_positions (Rep_matrix A))" - by (induct A) (simp add: Abs_matrix_inverse matrix_def) - -definition nrows :: "('a::zero) matrix \ nat" where - "nrows A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image fst) (nonzero_positions (Rep_matrix A))))" - -definition ncols :: "('a::zero) matrix \ nat" where - "ncols A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image snd) (nonzero_positions (Rep_matrix A))))" - -lemma nrows: - assumes hyp: "nrows A \ m" - shows "(Rep_matrix A m n) = 0" -proof cases - assume "nonzero_positions(Rep_matrix A) = {}" - then show "(Rep_matrix A m n) = 0" by (simp add: nonzero_positions_def) -next - assume a: "nonzero_positions(Rep_matrix A) \ {}" - let ?S = "fst`(nonzero_positions(Rep_matrix A))" - have c: "finite (?S)" by (simp add: finite_nonzero_positions) - from hyp have d: "Max (?S) < m" by (simp add: a nrows_def) - have "m \ ?S" - proof - - have "m \ ?S \ m <= Max(?S)" by (simp add: Max_ge [OF c]) - moreover from d have "~(m <= Max ?S)" by (simp) - ultimately show "m \ ?S" by (auto) - qed - thus "Rep_matrix A m n = 0" by (simp add: nonzero_positions_def image_Collect) -qed - -definition transpose_infmatrix :: "'a infmatrix \ 'a infmatrix" where - "transpose_infmatrix A j i == A i j" - -definition transpose_matrix :: "('a::zero) matrix \ 'a matrix" where - "transpose_matrix == Abs_matrix o transpose_infmatrix o Rep_matrix" - -declare transpose_infmatrix_def[simp] - -lemma transpose_infmatrix_twice[simp]: "transpose_infmatrix (transpose_infmatrix A) = A" -by ((rule ext)+, simp) - -lemma transpose_infmatrix: "transpose_infmatrix (% j i. P j i) = (% j i. P i j)" - apply (rule ext)+ - by simp - -lemma transpose_infmatrix_closed[simp]: "Rep_matrix (Abs_matrix (transpose_infmatrix (Rep_matrix x))) = transpose_infmatrix (Rep_matrix x)" -apply (rule Abs_matrix_inverse) -apply (simp add: matrix_def nonzero_positions_def image_def) -proof - - let ?A = "{pos. Rep_matrix x (snd pos) (fst pos) \ 0}" - let ?swap = "% pos. (snd pos, fst pos)" - let ?B = "{pos. Rep_matrix x (fst pos) (snd pos) \ 0}" - have swap_image: "?swap`?A = ?B" - apply (simp add: image_def) - apply (rule set_eqI) - apply (simp) - proof - fix y - assume hyp: "\a b. Rep_matrix x b a \ 0 \ y = (b, a)" - thus "Rep_matrix x (fst y) (snd y) \ 0" - proof - - from hyp obtain a b where "(Rep_matrix x b a \ 0 & y = (b,a))" by blast - then show "Rep_matrix x (fst y) (snd y) \ 0" by (simp) - qed - next - fix y - assume hyp: "Rep_matrix x (fst y) (snd y) \ 0" - show "\ a b. (Rep_matrix x b a \ 0 & y = (b,a))" - by (rule exI[of _ "snd y"], rule exI[of _ "fst y"]) (simp add: hyp) - qed - then have "finite (?swap`?A)" - proof - - have "finite (nonzero_positions (Rep_matrix x))" by (simp add: finite_nonzero_positions) - then have "finite ?B" by (simp add: nonzero_positions_def) - with swap_image show "finite (?swap`?A)" by (simp) - qed - moreover - have "inj_on ?swap ?A" by (simp add: inj_on_def) - ultimately show "finite ?A"by (rule finite_imageD[of ?swap ?A]) -qed - -lemma infmatrixforward: "(x::'a infmatrix) = y \ \ a b. x a b = y a b" by auto - -lemma transpose_infmatrix_inject: "(transpose_infmatrix A = transpose_infmatrix B) = (A = B)" -apply (auto) -apply (rule ext)+ -apply (simp add: transpose_infmatrix) -apply (drule infmatrixforward) -apply (simp) -done - -lemma transpose_matrix_inject: "(transpose_matrix A = transpose_matrix B) = (A = B)" -apply (simp add: transpose_matrix_def) -apply (subst Rep_matrix_inject[THEN sym])+ -apply (simp only: transpose_infmatrix_closed transpose_infmatrix_inject) -done - -lemma transpose_matrix[simp]: "Rep_matrix(transpose_matrix A) j i = Rep_matrix A i j" -by (simp add: transpose_matrix_def) - -lemma transpose_transpose_id[simp]: "transpose_matrix (transpose_matrix A) = A" -by (simp add: transpose_matrix_def) - -lemma nrows_transpose[simp]: "nrows (transpose_matrix A) = ncols A" -by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def) - -lemma ncols_transpose[simp]: "ncols (transpose_matrix A) = nrows A" -by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def) - -lemma ncols: "ncols A <= n \ Rep_matrix A m n = 0" -proof - - assume "ncols A <= n" - then have "nrows (transpose_matrix A) <= n" by (simp) - then have "Rep_matrix (transpose_matrix A) n m = 0" by (rule nrows) - thus "Rep_matrix A m n = 0" by (simp add: transpose_matrix_def) -qed - -lemma ncols_le: "(ncols A <= n) = (! j i. n <= i \ (Rep_matrix A j i) = 0)" (is "_ = ?st") -apply (auto) -apply (simp add: ncols) -proof (simp add: ncols_def, auto) - let ?P = "nonzero_positions (Rep_matrix A)" - let ?p = "snd`?P" - have a:"finite ?p" by (simp add: finite_nonzero_positions) - let ?m = "Max ?p" - assume "~(Suc (?m) <= n)" - then have b:"n <= ?m" by (simp) - fix a b - assume "(a,b) \ ?P" - then have "?p \ {}" by (auto) - with a have "?m \ ?p" by (simp) - moreover have "!x. (x \ ?p \ (? y. (Rep_matrix A y x) \ 0))" by (simp add: nonzero_positions_def image_def) - ultimately have "? y. (Rep_matrix A y ?m) \ 0" by (simp) - moreover assume ?st - ultimately show "False" using b by (simp) -qed - -lemma less_ncols: "(n < ncols A) = (? j i. n <= i & (Rep_matrix A j i) \ 0)" -proof - - have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith - show ?thesis by (simp add: a ncols_le) -qed - -lemma le_ncols: "(n <= ncols A) = (\ m. (\ j i. m <= i \ (Rep_matrix A j i) = 0) \ n <= m)" -apply (auto) -apply (subgoal_tac "ncols A <= m") -apply (simp) -apply (simp add: ncols_le) -apply (drule_tac x="ncols A" in spec) -by (simp add: ncols) - -lemma nrows_le: "(nrows A <= n) = (! j i. n <= j \ (Rep_matrix A j i) = 0)" (is ?s) -proof - - have "(nrows A <= n) = (ncols (transpose_matrix A) <= n)" by (simp) - also have "\ = (! j i. n <= i \ (Rep_matrix (transpose_matrix A) j i = 0))" by (rule ncols_le) - also have "\ = (! j i. n <= i \ (Rep_matrix A i j) = 0)" by (simp) - finally show "(nrows A <= n) = (! j i. n <= j \ (Rep_matrix A j i) = 0)" by (auto) -qed - -lemma less_nrows: "(m < nrows A) = (? j i. m <= j & (Rep_matrix A j i) \ 0)" -proof - - have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith - show ?thesis by (simp add: a nrows_le) -qed - -lemma le_nrows: "(n <= nrows A) = (\ m. (\ j i. m <= j \ (Rep_matrix A j i) = 0) \ n <= m)" -apply (auto) -apply (subgoal_tac "nrows A <= m") -apply (simp) -apply (simp add: nrows_le) -apply (drule_tac x="nrows A" in spec) -by (simp add: nrows) - -lemma nrows_notzero: "Rep_matrix A m n \ 0 \ m < nrows A" -apply (case_tac "nrows A <= m") -apply (simp_all add: nrows) -done - -lemma ncols_notzero: "Rep_matrix A m n \ 0 \ n < ncols A" -apply (case_tac "ncols A <= n") -apply (simp_all add: ncols) -done - -lemma finite_natarray1: "finite {x. x < (n::nat)}" -apply (induct n) -apply (simp) -proof - - fix n - have "{x. x < Suc n} = insert n {x. x < n}" by (rule set_eqI, simp, arith) - moreover assume "finite {x. x < n}" - ultimately show "finite {x. x < Suc n}" by (simp) -qed - -lemma finite_natarray2: "finite {pos. (fst pos) < (m::nat) & (snd pos) < (n::nat)}" - apply (induct m) - apply (simp+) - proof - - fix m::nat - let ?s0 = "{pos. fst pos < m & snd pos < n}" - let ?s1 = "{pos. fst pos < (Suc m) & snd pos < n}" - let ?sd = "{pos. fst pos = m & snd pos < n}" - assume f0: "finite ?s0" - have f1: "finite ?sd" - proof - - let ?f = "% x. (m, x)" - have "{pos. fst pos = m & snd pos < n} = ?f ` {x. x < n}" by (rule set_eqI, simp add: image_def, auto) - moreover have "finite {x. x < n}" by (simp add: finite_natarray1) - ultimately show "finite {pos. fst pos = m & snd pos < n}" by (simp) - qed - have su: "?s0 \ ?sd = ?s1" by (rule set_eqI, simp, arith) - from f0 f1 have "finite (?s0 \ ?sd)" by (rule finite_UnI) - with su show "finite ?s1" by (simp) -qed - -lemma RepAbs_matrix: - assumes aem: "? m. ! j i. m <= j \ x j i = 0" (is ?em) and aen:"? n. ! j i. (n <= i \ x j i = 0)" (is ?en) - shows "(Rep_matrix (Abs_matrix x)) = x" -apply (rule Abs_matrix_inverse) -apply (simp add: matrix_def nonzero_positions_def) -proof - - from aem obtain m where a: "! j i. m <= j \ x j i = 0" by (blast) - from aen obtain n where b: "! j i. n <= i \ x j i = 0" by (blast) - let ?u = "{pos. x (fst pos) (snd pos) \ 0}" - let ?v = "{pos. fst pos < m & snd pos < n}" - have c: "!! (m::nat) a. ~(m <= a) \ a < m" by (arith) - from a b have "(?u \ (-?v)) = {}" - apply (simp) - apply (rule set_eqI) - apply (simp) - apply auto - by (rule c, auto)+ - then have d: "?u \ ?v" by blast - moreover have "finite ?v" by (simp add: finite_natarray2) - ultimately show "finite ?u" by (rule finite_subset) -qed - -definition apply_infmatrix :: "('a \ 'b) \ 'a infmatrix \ 'b infmatrix" where - "apply_infmatrix f == % A. (% j i. f (A j i))" - -definition apply_matrix :: "('a \ 'b) \ ('a::zero) matrix \ ('b::zero) matrix" where - "apply_matrix f == % A. Abs_matrix (apply_infmatrix f (Rep_matrix A))" - -definition combine_infmatrix :: "('a \ 'b \ 'c) \ 'a infmatrix \ 'b infmatrix \ 'c infmatrix" where - "combine_infmatrix f == % A B. (% j i. f (A j i) (B j i))" - -definition combine_matrix :: "('a \ 'b \ 'c) \ ('a::zero) matrix \ ('b::zero) matrix \ ('c::zero) matrix" where - "combine_matrix f == % A B. Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))" - -lemma expand_apply_infmatrix[simp]: "apply_infmatrix f A j i = f (A j i)" -by (simp add: apply_infmatrix_def) - -lemma expand_combine_infmatrix[simp]: "combine_infmatrix f A B j i = f (A j i) (B j i)" -by (simp add: combine_infmatrix_def) - -definition commutative :: "('a \ 'a \ 'b) \ bool" where -"commutative f == ! x y. f x y = f y x" - -definition associative :: "('a \ 'a \ 'a) \ bool" where -"associative f == ! x y z. f (f x y) z = f x (f y z)" - -text{* -To reason about associativity and commutativity of operations on matrices, -let's take a step back and look at the general situtation: Assume that we have -sets $A$ and $B$ with $B \subset A$ and an abstraction $u: A \rightarrow B$. This abstraction has to fulfill $u(b) = b$ for all $b \in B$, but is arbitrary otherwise. -Each function $f: A \times A \rightarrow A$ now induces a function $f': B \times B \rightarrow B$ by $f' = u \circ f$. -It is obvious that commutativity of $f$ implies commutativity of $f'$: $f' x y = u (f x y) = u (f y x) = f' y x.$ -*} - -lemma combine_infmatrix_commute: - "commutative f \ commutative (combine_infmatrix f)" -by (simp add: commutative_def combine_infmatrix_def) - -lemma combine_matrix_commute: -"commutative f \ commutative (combine_matrix f)" -by (simp add: combine_matrix_def commutative_def combine_infmatrix_def) - -text{* -On the contrary, given an associative function $f$ we cannot expect $f'$ to be associative. A counterexample is given by $A=\ganz$, $B=\{-1, 0, 1\}$, -as $f$ we take addition on $\ganz$, which is clearly associative. The abstraction is given by $u(a) = 0$ for $a \notin B$. Then we have -\[ f' (f' 1 1) -1 = u(f (u (f 1 1)) -1) = u(f (u 2) -1) = u (f 0 -1) = -1, \] -but on the other hand we have -\[ f' 1 (f' 1 -1) = u (f 1 (u (f 1 -1))) = u (f 1 0) = 1.\] -A way out of this problem is to assume that $f(A\times A)\subset A$ holds, and this is what we are going to do: -*} - -lemma nonzero_positions_combine_infmatrix[simp]: "f 0 0 = 0 \ nonzero_positions (combine_infmatrix f A B) \ (nonzero_positions A) \ (nonzero_positions B)" -by (rule subsetI, simp add: nonzero_positions_def combine_infmatrix_def, auto) - -lemma finite_nonzero_positions_Rep[simp]: "finite (nonzero_positions (Rep_matrix A))" -by (insert Rep_matrix [of A], simp add: matrix_def) - -lemma combine_infmatrix_closed [simp]: - "f 0 0 = 0 \ Rep_matrix (Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))) = combine_infmatrix f (Rep_matrix A) (Rep_matrix B)" -apply (rule Abs_matrix_inverse) -apply (simp add: matrix_def) -apply (rule finite_subset[of _ "(nonzero_positions (Rep_matrix A)) \ (nonzero_positions (Rep_matrix B))"]) -by (simp_all) - -text {* We need the next two lemmas only later, but it is analog to the above one, so we prove them now: *} -lemma nonzero_positions_apply_infmatrix[simp]: "f 0 = 0 \ nonzero_positions (apply_infmatrix f A) \ nonzero_positions A" -by (rule subsetI, simp add: nonzero_positions_def apply_infmatrix_def, auto) - -lemma apply_infmatrix_closed [simp]: - "f 0 = 0 \ Rep_matrix (Abs_matrix (apply_infmatrix f (Rep_matrix A))) = apply_infmatrix f (Rep_matrix A)" -apply (rule Abs_matrix_inverse) -apply (simp add: matrix_def) -apply (rule finite_subset[of _ "nonzero_positions (Rep_matrix A)"]) -by (simp_all) - -lemma combine_infmatrix_assoc[simp]: "f 0 0 = 0 \ associative f \ associative (combine_infmatrix f)" -by (simp add: associative_def combine_infmatrix_def) - -lemma comb: "f = g \ x = y \ f x = g y" -by (auto) - -lemma combine_matrix_assoc: "f 0 0 = 0 \ associative f \ associative (combine_matrix f)" -apply (simp(no_asm) add: associative_def combine_matrix_def, auto) -apply (rule comb [of Abs_matrix Abs_matrix]) -by (auto, insert combine_infmatrix_assoc[of f], simp add: associative_def) - -lemma Rep_apply_matrix[simp]: "f 0 = 0 \ Rep_matrix (apply_matrix f A) j i = f (Rep_matrix A j i)" -by (simp add: apply_matrix_def) - -lemma Rep_combine_matrix[simp]: "f 0 0 = 0 \ Rep_matrix (combine_matrix f A B) j i = f (Rep_matrix A j i) (Rep_matrix B j i)" - by(simp add: combine_matrix_def) - -lemma combine_nrows_max: "f 0 0 = 0 \ nrows (combine_matrix f A B) <= max (nrows A) (nrows B)" -by (simp add: nrows_le) - -lemma combine_ncols_max: "f 0 0 = 0 \ ncols (combine_matrix f A B) <= max (ncols A) (ncols B)" -by (simp add: ncols_le) - -lemma combine_nrows: "f 0 0 = 0 \ nrows A <= q \ nrows B <= q \ nrows(combine_matrix f A B) <= q" - by (simp add: nrows_le) - -lemma combine_ncols: "f 0 0 = 0 \ ncols A <= q \ ncols B <= q \ ncols(combine_matrix f A B) <= q" - by (simp add: ncols_le) - -definition zero_r_neutral :: "('a \ 'b::zero \ 'a) \ bool" where - "zero_r_neutral f == ! a. f a 0 = a" - -definition zero_l_neutral :: "('a::zero \ 'b \ 'b) \ bool" where - "zero_l_neutral f == ! a. f 0 a = a" - -definition zero_closed :: "(('a::zero) \ ('b::zero) \ ('c::zero)) \ bool" where - "zero_closed f == (!x. f x 0 = 0) & (!y. f 0 y = 0)" - -primrec foldseq :: "('a \ 'a \ 'a) \ (nat \ 'a) \ nat \ 'a" -where - "foldseq f s 0 = s 0" -| "foldseq f s (Suc n) = f (s 0) (foldseq f (% k. s(Suc k)) n)" - -primrec foldseq_transposed :: "('a \ 'a \ 'a) \ (nat \ 'a) \ nat \ 'a" -where - "foldseq_transposed f s 0 = s 0" -| "foldseq_transposed f s (Suc n) = f (foldseq_transposed f s n) (s (Suc n))" - -lemma foldseq_assoc : "associative f \ foldseq f = foldseq_transposed f" -proof - - assume a:"associative f" - then have sublemma: "!! n. ! N s. N <= n \ foldseq f s N = foldseq_transposed f s N" - proof - - fix n - show "!N s. N <= n \ foldseq f s N = foldseq_transposed f s N" - proof (induct n) - show "!N s. N <= 0 \ foldseq f s N = foldseq_transposed f s N" by simp - next - fix n - assume b:"! N s. N <= n \ foldseq f s N = foldseq_transposed f s N" - have c:"!!N s. N <= n \ foldseq f s N = foldseq_transposed f s N" by (simp add: b) - show "! N t. N <= Suc n \ foldseq f t N = foldseq_transposed f t N" - proof (auto) - fix N t - assume Nsuc: "N <= Suc n" - show "foldseq f t N = foldseq_transposed f t N" - proof cases - assume "N <= n" - then show "foldseq f t N = foldseq_transposed f t N" by (simp add: b) - next - assume "~(N <= n)" - with Nsuc have Nsuceq: "N = Suc n" by simp - have neqz: "n \ 0 \ ? m. n = Suc m & Suc m <= n" by arith - have assocf: "!! x y z. f x (f y z) = f (f x y) z" by (insert a, simp add: associative_def) - show "foldseq f t N = foldseq_transposed f t N" - apply (simp add: Nsuceq) - apply (subst c) - apply (simp) - apply (case_tac "n = 0") - apply (simp) - apply (drule neqz) - apply (erule exE) - apply (simp) - apply (subst assocf) - proof - - fix m - assume "n = Suc m & Suc m <= n" - then have mless: "Suc m <= n" by arith - then have step1: "foldseq_transposed f (% k. t (Suc k)) m = foldseq f (% k. t (Suc k)) m" (is "?T1 = ?T2") - apply (subst c) - by simp+ - have step2: "f (t 0) ?T2 = foldseq f t (Suc m)" (is "_ = ?T3") by simp - have step3: "?T3 = foldseq_transposed f t (Suc m)" (is "_ = ?T4") - apply (subst c) - by (simp add: mless)+ - have step4: "?T4 = f (foldseq_transposed f t m) (t (Suc m))" (is "_=?T5") by simp - from step1 step2 step3 step4 show sowhat: "f (f (t 0) ?T1) (t (Suc (Suc m))) = f ?T5 (t (Suc (Suc m)))" by simp - qed - qed - qed - qed - qed - show "foldseq f = foldseq_transposed f" by ((rule ext)+, insert sublemma, auto) - qed - -lemma foldseq_distr: "\associative f; commutative f\ \ foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" -proof - - assume assoc: "associative f" - assume comm: "commutative f" - from assoc have a:"!! x y z. f (f x y) z = f x (f y z)" by (simp add: associative_def) - from comm have b: "!! x y. f x y = f y x" by (simp add: commutative_def) - from assoc comm have c: "!! x y z. f x (f y z) = f y (f x z)" by (simp add: commutative_def associative_def) - have "!! n. (! u v. foldseq f (%k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))" - apply (induct_tac n) - apply (simp+, auto) - by (simp add: a b c) - then show "foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" by simp -qed - -theorem "\associative f; associative g; \a b c d. g (f a b) (f c d) = f (g a c) (g b d); ? x y. (f x) \ (f y); ? x y. (g x) \ (g y); f x x = x; g x x = x\ \ f=g | (! y. f y x = y) | (! y. g y x = y)" -oops -(* Model found - -Trying to find a model that refutes: \associative f; associative g; - \a b c d. g (f a b) (f c d) = f (g a c) (g b d); \x y. f x \ f y; - \x y. g x \ g y; f x x = x; g x x = x\ -\ f = g \ (\y. f y x = y) \ (\y. g y x = y) -Searching for a model of size 1, translating term... invoking SAT solver... no model found. -Searching for a model of size 2, translating term... invoking SAT solver... no model found. -Searching for a model of size 3, translating term... invoking SAT solver... -Model found: -Size of types: 'a: 3 -x: a1 -g: (a0\(a0\a1, a1\a0, a2\a1), a1\(a0\a0, a1\a1, a2\a0), a2\(a0\a1, a1\a0, a2\a1)) -f: (a0\(a0\a0, a1\a0, a2\a0), a1\(a0\a1, a1\a1, a2\a1), a2\(a0\a0, a1\a0, a2\a0)) -*) - -lemma foldseq_zero: -assumes fz: "f 0 0 = 0" and sz: "! i. i <= n \ s i = 0" -shows "foldseq f s n = 0" -proof - - have "!! n. ! s. (! i. i <= n \ s i = 0) \ foldseq f s n = 0" - apply (induct_tac n) - apply (simp) - by (simp add: fz) - then show "foldseq f s n = 0" by (simp add: sz) -qed - -lemma foldseq_significant_positions: - assumes p: "! i. i <= N \ S i = T i" - shows "foldseq f S N = foldseq f T N" -proof - - have "!! m . ! s t. (! i. i<=m \ s i = t i) \ foldseq f s m = foldseq f t m" - apply (induct_tac m) - apply (simp) - apply (simp) - apply (auto) - proof - - fix n - fix s::"nat\'a" - fix t::"nat\'a" - assume a: "\s t. (\i\n. s i = t i) \ foldseq f s n = foldseq f t n" - assume b: "\i\Suc n. s i = t i" - have c:"!! a b. a = b \ f (t 0) a = f (t 0) b" by blast - have d:"!! s t. (\i\n. s i = t i) \ foldseq f s n = foldseq f t n" by (simp add: a) - show "f (t 0) (foldseq f (\k. s (Suc k)) n) = f (t 0) (foldseq f (\k. t (Suc k)) n)" by (rule c, simp add: d b) - qed - with p show ?thesis by simp -qed - -lemma foldseq_tail: - assumes "M <= N" - shows "foldseq f S N = foldseq f (% k. (if k < M then (S k) else (foldseq f (% k. S(k+M)) (N-M)))) M" -proof - - have suc: "!! a b. \a <= Suc b; a \ Suc b\ \ a <= b" by arith - have a:"!! a b c . a = b \ f c a = f c b" by blast - have "!! n. ! m s. m <= n \ foldseq f s n = foldseq f (% k. (if k < m then (s k) else (foldseq f (% k. s(k+m)) (n-m)))) m" - apply (induct_tac n) - apply (simp) - apply (simp) - apply (auto) - apply (case_tac "m = Suc na") - apply (simp) - apply (rule a) - apply (rule foldseq_significant_positions) - apply (auto) - apply (drule suc, simp+) - proof - - fix na m s - assume suba:"\m\na. \s. foldseq f s na = foldseq f (\k. if k < m then s k else foldseq f (\k. s (k + m)) (na - m))m" - assume subb:"m <= na" - from suba have subc:"!! m s. m <= na \foldseq f s na = foldseq f (\k. if k < m then s k else foldseq f (\k. s (k + m)) (na - m))m" by simp - have subd: "foldseq f (\k. if k < m then s (Suc k) else foldseq f (\k. s (Suc (k + m))) (na - m)) m = - foldseq f (% k. s(Suc k)) na" - by (rule subc[of m "% k. s(Suc k)", THEN sym], simp add: subb) - from subb have sube: "m \ 0 \ ? mm. m = Suc mm & mm <= na" by arith - show "f (s 0) (foldseq f (\k. if k < m then s (Suc k) else foldseq f (\k. s (Suc (k + m))) (na - m)) m) = - foldseq f (\k. if k < m then s k else foldseq f (\k. s (k + m)) (Suc na - m)) m" - apply (simp add: subd) - apply (cases "m = 0") - apply (simp) - apply (drule sube) - apply (auto) - apply (rule a) - by (simp add: subc cong del: if_cong) - qed - then show ?thesis using assms by simp -qed - -lemma foldseq_zerotail: - assumes - fz: "f 0 0 = 0" - and sz: "! i. n <= i \ s i = 0" - and nm: "n <= m" - shows - "foldseq f s n = foldseq f s m" -proof - - show "foldseq f s n = foldseq f s m" - apply (simp add: foldseq_tail[OF nm, of f s]) - apply (rule foldseq_significant_positions) - apply (auto) - apply (subst foldseq_zero) - by (simp add: fz sz)+ -qed - -lemma foldseq_zerotail2: - assumes "! x. f x 0 = x" - and "! i. n < i \ s i = 0" - and nm: "n <= m" - shows "foldseq f s n = foldseq f s m" -proof - - have "f 0 0 = 0" by (simp add: assms) - have b:"!! m n. n <= m \ m \ n \ ? k. m-n = Suc k" by arith - have c: "0 <= m" by simp - have d: "!! k. k \ 0 \ ? l. k = Suc l" by arith - show ?thesis - apply (subst foldseq_tail[OF nm]) - apply (rule foldseq_significant_positions) - apply (auto) - apply (case_tac "m=n") - apply (simp+) - apply (drule b[OF nm]) - apply (auto) - apply (case_tac "k=0") - apply (simp add: assms) - apply (drule d) - apply (auto) - apply (simp add: assms foldseq_zero) - done -qed - -lemma foldseq_zerostart: - "! x. f 0 (f 0 x) = f 0 x \ ! i. i <= n \ s i = 0 \ foldseq f s (Suc n) = f 0 (s (Suc n))" -proof - - assume f00x: "! x. f 0 (f 0 x) = f 0 x" - have "! s. (! i. i<=n \ s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n))" - apply (induct n) - apply (simp) - apply (rule allI, rule impI) - proof - - fix n - fix s - have a:"foldseq f s (Suc (Suc n)) = f (s 0) (foldseq f (% k. s(Suc k)) (Suc n))" by simp - assume b: "! s. ((\i\n. s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n)))" - from b have c:"!! s. (\i\n. s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n))" by simp - assume d: "! i. i <= Suc n \ s i = 0" - show "foldseq f s (Suc (Suc n)) = f 0 (s (Suc (Suc n)))" - apply (subst a) - apply (subst c) - by (simp add: d f00x)+ - qed - then show "! i. i <= n \ s i = 0 \ foldseq f s (Suc n) = f 0 (s (Suc n))" by simp -qed - -lemma foldseq_zerostart2: - "! x. f 0 x = x \ ! i. i < n \ s i = 0 \ foldseq f s n = s n" -proof - - assume a:"! i. i s i = 0" - assume x:"! x. f 0 x = x" - from x have f00x: "! x. f 0 (f 0 x) = f 0 x" by blast - have b: "!! i l. i < Suc l = (i <= l)" by arith - have d: "!! k. k \ 0 \ ? l. k = Suc l" by arith - show "foldseq f s n = s n" - apply (case_tac "n=0") - apply (simp) - apply (insert a) - apply (drule d) - apply (auto) - apply (simp add: b) - apply (insert f00x) - apply (drule foldseq_zerostart) - by (simp add: x)+ -qed - -lemma foldseq_almostzero: - assumes f0x:"! x. f 0 x = x" and fx0: "! x. f x 0 = x" and s0:"! i. i \ j \ s i = 0" - shows "foldseq f s n = (if (j <= n) then (s j) else 0)" -proof - - from s0 have a: "! i. i < j \ s i = 0" by simp - from s0 have b: "! i. j < i \ s i = 0" by simp - show ?thesis - apply auto - apply (subst foldseq_zerotail2[of f, OF fx0, of j, OF b, of n, THEN sym]) - apply simp - apply (subst foldseq_zerostart2) - apply (simp add: f0x a)+ - apply (subst foldseq_zero) - by (simp add: s0 f0x)+ -qed - -lemma foldseq_distr_unary: - assumes "!! a b. g (f a b) = f (g a) (g b)" - shows "g(foldseq f s n) = foldseq f (% x. g(s x)) n" -proof - - have "! s. g(foldseq f s n) = foldseq f (% x. g(s x)) n" - apply (induct_tac n) - apply (simp) - apply (simp) - apply (auto) - apply (drule_tac x="% k. s (Suc k)" in spec) - by (simp add: assms) - then show ?thesis by simp -qed - -definition mult_matrix_n :: "nat \ (('a::zero) \ ('b::zero) \ ('c::zero)) \ ('c \ 'c \ 'c) \ 'a matrix \ 'b matrix \ 'c matrix" where - "mult_matrix_n n fmul fadd A B == Abs_matrix(% j i. foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) n)" - -definition mult_matrix :: "(('a::zero) \ ('b::zero) \ ('c::zero)) \ ('c \ 'c \ 'c) \ 'a matrix \ 'b matrix \ 'c matrix" where - "mult_matrix fmul fadd A B == mult_matrix_n (max (ncols A) (nrows B)) fmul fadd A B" - -lemma mult_matrix_n: - assumes "ncols A \ n" (is ?An) "nrows B \ n" (is ?Bn) "fadd 0 0 = 0" "fmul 0 0 = 0" - shows c:"mult_matrix fmul fadd A B = mult_matrix_n n fmul fadd A B" -proof - - show ?thesis using assms - apply (simp add: mult_matrix_def mult_matrix_n_def) - apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) - apply (rule foldseq_zerotail, simp_all add: nrows_le ncols_le assms) - done -qed - -lemma mult_matrix_nm: - assumes "ncols A <= n" "nrows B <= n" "ncols A <= m" "nrows B <= m" "fadd 0 0 = 0" "fmul 0 0 = 0" - shows "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" -proof - - from assms have "mult_matrix_n n fmul fadd A B = mult_matrix fmul fadd A B" - by (simp add: mult_matrix_n) - also from assms have "\ = mult_matrix_n m fmul fadd A B" - by (simp add: mult_matrix_n[THEN sym]) - finally show "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" by simp -qed - -definition r_distributive :: "('a \ 'b \ 'b) \ ('b \ 'b \ 'b) \ bool" where - "r_distributive fmul fadd == ! a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)" - -definition l_distributive :: "('a \ 'b \ 'a) \ ('a \ 'a \ 'a) \ bool" where - "l_distributive fmul fadd == ! a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)" - -definition distributive :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ bool" where - "distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd" - -lemma max1: "!! a x y. (a::nat) <= x \ a <= max x y" by (arith) -lemma max2: "!! b x y. (b::nat) <= y \ b <= max x y" by (arith) - -lemma r_distributive_matrix: - assumes - "r_distributive fmul fadd" - "associative fadd" - "commutative fadd" - "fadd 0 0 = 0" - "! a. fmul a 0 = 0" - "! a. fmul 0 a = 0" - shows "r_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" -proof - - from assms show ?thesis - apply (simp add: r_distributive_def mult_matrix_def, auto) - proof - - fix a::"'a matrix" - fix u::"'b matrix" - fix v::"'b matrix" - let ?mx = "max (ncols a) (max (nrows u) (nrows v))" - from assms show "mult_matrix_n (max (ncols a) (nrows (combine_matrix fadd u v))) fmul fadd a (combine_matrix fadd u v) = - combine_matrix fadd (mult_matrix_n (max (ncols a) (nrows u)) fmul fadd a u) (mult_matrix_n (max (ncols a) (nrows v)) fmul fadd a v)" - apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) - apply (simp add: max1 max2 combine_nrows combine_ncols)+ - apply (subst mult_matrix_nm[of _ _ v ?mx fadd fmul]) - apply (simp add: max1 max2 combine_nrows combine_ncols)+ - apply (subst mult_matrix_nm[of _ _ u ?mx fadd fmul]) - apply (simp add: max1 max2 combine_nrows combine_ncols)+ - apply (simp add: mult_matrix_n_def r_distributive_def foldseq_distr[of fadd]) - apply (simp add: combine_matrix_def combine_infmatrix_def) - apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) - apply (simplesubst RepAbs_matrix) - apply (simp, auto) - apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero) - apply (rule exI[of _ "ncols v"], simp add: ncols_le foldseq_zero) - apply (subst RepAbs_matrix) - apply (simp, auto) - apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero) - apply (rule exI[of _ "ncols u"], simp add: ncols_le foldseq_zero) - done - qed -qed - -lemma l_distributive_matrix: - assumes - "l_distributive fmul fadd" - "associative fadd" - "commutative fadd" - "fadd 0 0 = 0" - "! a. fmul a 0 = 0" - "! a. fmul 0 a = 0" - shows "l_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" -proof - - from assms show ?thesis - apply (simp add: l_distributive_def mult_matrix_def, auto) - proof - - fix a::"'b matrix" - fix u::"'a matrix" - fix v::"'a matrix" - let ?mx = "max (nrows a) (max (ncols u) (ncols v))" - from assms show "mult_matrix_n (max (ncols (combine_matrix fadd u v)) (nrows a)) fmul fadd (combine_matrix fadd u v) a = - combine_matrix fadd (mult_matrix_n (max (ncols u) (nrows a)) fmul fadd u a) (mult_matrix_n (max (ncols v) (nrows a)) fmul fadd v a)" - apply (subst mult_matrix_nm[of v _ _ ?mx fadd fmul]) - apply (simp add: max1 max2 combine_nrows combine_ncols)+ - apply (subst mult_matrix_nm[of u _ _ ?mx fadd fmul]) - apply (simp add: max1 max2 combine_nrows combine_ncols)+ - apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) - apply (simp add: max1 max2 combine_nrows combine_ncols)+ - apply (simp add: mult_matrix_n_def l_distributive_def foldseq_distr[of fadd]) - apply (simp add: combine_matrix_def combine_infmatrix_def) - apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) - apply (simplesubst RepAbs_matrix) - apply (simp, auto) - apply (rule exI[of _ "nrows v"], simp add: nrows_le foldseq_zero) - apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero) - apply (subst RepAbs_matrix) - apply (simp, auto) - apply (rule exI[of _ "nrows u"], simp add: nrows_le foldseq_zero) - apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero) - done - qed -qed - -instantiation matrix :: (zero) zero -begin - -definition zero_matrix_def: "0 = Abs_matrix (\j i. 0)" - -instance .. - -end - -lemma Rep_zero_matrix_def[simp]: "Rep_matrix 0 j i = 0" - apply (simp add: zero_matrix_def) - apply (subst RepAbs_matrix) - by (auto) - -lemma zero_matrix_def_nrows[simp]: "nrows 0 = 0" -proof - - have a:"!! (x::nat). x <= 0 \ x = 0" by (arith) - show "nrows 0 = 0" by (rule a, subst nrows_le, simp) -qed - -lemma zero_matrix_def_ncols[simp]: "ncols 0 = 0" -proof - - have a:"!! (x::nat). x <= 0 \ x = 0" by (arith) - show "ncols 0 = 0" by (rule a, subst ncols_le, simp) -qed - -lemma combine_matrix_zero_l_neutral: "zero_l_neutral f \ zero_l_neutral (combine_matrix f)" - by (simp add: zero_l_neutral_def combine_matrix_def combine_infmatrix_def) - -lemma combine_matrix_zero_r_neutral: "zero_r_neutral f \ zero_r_neutral (combine_matrix f)" - by (simp add: zero_r_neutral_def combine_matrix_def combine_infmatrix_def) - -lemma mult_matrix_zero_closed: "\fadd 0 0 = 0; zero_closed fmul\ \ zero_closed (mult_matrix fmul fadd)" - apply (simp add: zero_closed_def mult_matrix_def mult_matrix_n_def) - apply (auto) - by (subst foldseq_zero, (simp add: zero_matrix_def)+)+ - -lemma mult_matrix_n_zero_right[simp]: "\fadd 0 0 = 0; !a. fmul a 0 = 0\ \ mult_matrix_n n fmul fadd A 0 = 0" - apply (simp add: mult_matrix_n_def) - apply (subst foldseq_zero) - by (simp_all add: zero_matrix_def) - -lemma mult_matrix_n_zero_left[simp]: "\fadd 0 0 = 0; !a. fmul 0 a = 0\ \ mult_matrix_n n fmul fadd 0 A = 0" - apply (simp add: mult_matrix_n_def) - apply (subst foldseq_zero) - by (simp_all add: zero_matrix_def) - -lemma mult_matrix_zero_left[simp]: "\fadd 0 0 = 0; !a. fmul 0 a = 0\ \ mult_matrix fmul fadd 0 A = 0" -by (simp add: mult_matrix_def) - -lemma mult_matrix_zero_right[simp]: "\fadd 0 0 = 0; !a. fmul a 0 = 0\ \ mult_matrix fmul fadd A 0 = 0" -by (simp add: mult_matrix_def) - -lemma apply_matrix_zero[simp]: "f 0 = 0 \ apply_matrix f 0 = 0" - apply (simp add: apply_matrix_def apply_infmatrix_def) - by (simp add: zero_matrix_def) - -lemma combine_matrix_zero: "f 0 0 = 0 \ combine_matrix f 0 0 = 0" - apply (simp add: combine_matrix_def combine_infmatrix_def) - by (simp add: zero_matrix_def) - -lemma transpose_matrix_zero[simp]: "transpose_matrix 0 = 0" -apply (simp add: transpose_matrix_def zero_matrix_def RepAbs_matrix) -apply (subst Rep_matrix_inject[symmetric], (rule ext)+) -apply (simp add: RepAbs_matrix) -done - -lemma apply_zero_matrix_def[simp]: "apply_matrix (% x. 0) A = 0" - apply (simp add: apply_matrix_def apply_infmatrix_def) - by (simp add: zero_matrix_def) - -definition singleton_matrix :: "nat \ nat \ ('a::zero) \ 'a matrix" where - "singleton_matrix j i a == Abs_matrix(% m n. if j = m & i = n then a else 0)" - -definition move_matrix :: "('a::zero) matrix \ int \ int \ 'a matrix" where - "move_matrix A y x == Abs_matrix(% j i. if (((int j)-y) < 0) | (((int i)-x) < 0) then 0 else Rep_matrix A (nat ((int j)-y)) (nat ((int i)-x)))" - -definition take_rows :: "('a::zero) matrix \ nat \ 'a matrix" where - "take_rows A r == Abs_matrix(% j i. if (j < r) then (Rep_matrix A j i) else 0)" - -definition take_columns :: "('a::zero) matrix \ nat \ 'a matrix" where - "take_columns A c == Abs_matrix(% j i. if (i < c) then (Rep_matrix A j i) else 0)" - -definition column_of_matrix :: "('a::zero) matrix \ nat \ 'a matrix" where - "column_of_matrix A n == take_columns (move_matrix A 0 (- int n)) 1" - -definition row_of_matrix :: "('a::zero) matrix \ nat \ 'a matrix" where - "row_of_matrix A m == take_rows (move_matrix A (- int m) 0) 1" - -lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m & i = n then e else 0)" -apply (simp add: singleton_matrix_def) -apply (auto) -apply (subst RepAbs_matrix) -apply (rule exI[of _ "Suc m"], simp) -apply (rule exI[of _ "Suc n"], simp+) -by (subst RepAbs_matrix, rule exI[of _ "Suc j"], simp, rule exI[of _ "Suc i"], simp+)+ - -lemma apply_singleton_matrix[simp]: "f 0 = 0 \ apply_matrix f (singleton_matrix j i x) = (singleton_matrix j i (f x))" -apply (subst Rep_matrix_inject[symmetric]) -apply (rule ext)+ -apply (simp) -done - -lemma singleton_matrix_zero[simp]: "singleton_matrix j i 0 = 0" - by (simp add: singleton_matrix_def zero_matrix_def) - -lemma nrows_singleton[simp]: "nrows(singleton_matrix j i e) = (if e = 0 then 0 else Suc j)" -proof- -have th: "\ (\m. m \ j)" "\n. \ n \ i" by arith+ -from th show ?thesis -apply (auto) -apply (rule le_antisym) -apply (subst nrows_le) -apply (simp add: singleton_matrix_def, auto) -apply (subst RepAbs_matrix) -apply auto -apply (simp add: Suc_le_eq) -apply (rule not_leE) -apply (subst nrows_le) -by simp -qed - -lemma ncols_singleton[simp]: "ncols(singleton_matrix j i e) = (if e = 0 then 0 else Suc i)" -proof- -have th: "\ (\m. m \ j)" "\n. \ n \ i" by arith+ -from th show ?thesis -apply (auto) -apply (rule le_antisym) -apply (subst ncols_le) -apply (simp add: singleton_matrix_def, auto) -apply (subst RepAbs_matrix) -apply auto -apply (simp add: Suc_le_eq) -apply (rule not_leE) -apply (subst ncols_le) -by simp -qed - -lemma combine_singleton: "f 0 0 = 0 \ combine_matrix f (singleton_matrix j i a) (singleton_matrix j i b) = singleton_matrix j i (f a b)" -apply (simp add: singleton_matrix_def combine_matrix_def combine_infmatrix_def) -apply (subst RepAbs_matrix) -apply (rule exI[of _ "Suc j"], simp) -apply (rule exI[of _ "Suc i"], simp) -apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) -apply (subst RepAbs_matrix) -apply (rule exI[of _ "Suc j"], simp) -apply (rule exI[of _ "Suc i"], simp) -by simp - -lemma transpose_singleton[simp]: "transpose_matrix (singleton_matrix j i a) = singleton_matrix i j a" -apply (subst Rep_matrix_inject[symmetric], (rule ext)+) -apply (simp) -done - -lemma Rep_move_matrix[simp]: - "Rep_matrix (move_matrix A y x) j i = - (if (((int j)-y) < 0) | (((int i)-x) < 0) then 0 else Rep_matrix A (nat((int j)-y)) (nat((int i)-x)))" -apply (simp add: move_matrix_def) -apply (auto) -by (subst RepAbs_matrix, - rule exI[of _ "(nrows A)+(nat (abs y))"], auto, rule nrows, arith, - rule exI[of _ "(ncols A)+(nat (abs x))"], auto, rule ncols, arith)+ - -lemma move_matrix_0_0[simp]: "move_matrix A 0 0 = A" -by (simp add: move_matrix_def) - -lemma move_matrix_ortho: "move_matrix A j i = move_matrix (move_matrix A j 0) 0 i" -apply (subst Rep_matrix_inject[symmetric]) -apply (rule ext)+ -apply (simp) -done - -lemma transpose_move_matrix[simp]: - "transpose_matrix (move_matrix A x y) = move_matrix (transpose_matrix A) y x" -apply (subst Rep_matrix_inject[symmetric], (rule ext)+) -apply (simp) -done - -lemma move_matrix_singleton[simp]: "move_matrix (singleton_matrix u v x) j i = - (if (j + int u < 0) | (i + int v < 0) then 0 else (singleton_matrix (nat (j + int u)) (nat (i + int v)) x))" - apply (subst Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply (case_tac "j + int u < 0") - apply (simp, arith) - apply (case_tac "i + int v < 0") - apply (simp, arith) - apply simp - apply arith - done - -lemma Rep_take_columns[simp]: - "Rep_matrix (take_columns A c) j i = - (if i < c then (Rep_matrix A j i) else 0)" -apply (simp add: take_columns_def) -apply (simplesubst RepAbs_matrix) -apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le) -apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le) -done - -lemma Rep_take_rows[simp]: - "Rep_matrix (take_rows A r) j i = - (if j < r then (Rep_matrix A j i) else 0)" -apply (simp add: take_rows_def) -apply (simplesubst RepAbs_matrix) -apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le) -apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le) -done - -lemma Rep_column_of_matrix[simp]: - "Rep_matrix (column_of_matrix A c) j i = (if i = 0 then (Rep_matrix A j c) else 0)" - by (simp add: column_of_matrix_def) - -lemma Rep_row_of_matrix[simp]: - "Rep_matrix (row_of_matrix A r) j i = (if j = 0 then (Rep_matrix A r i) else 0)" - by (simp add: row_of_matrix_def) - -lemma column_of_matrix: "ncols A <= n \ column_of_matrix A n = 0" -apply (subst Rep_matrix_inject[THEN sym]) -apply (rule ext)+ -by (simp add: ncols) - -lemma row_of_matrix: "nrows A <= n \ row_of_matrix A n = 0" -apply (subst Rep_matrix_inject[THEN sym]) -apply (rule ext)+ -by (simp add: nrows) - -lemma mult_matrix_singleton_right[simp]: - assumes - "! x. fmul x 0 = 0" - "! x. fmul 0 x = 0" - "! x. fadd 0 x = x" - "! x. fadd x 0 = x" - shows "(mult_matrix fmul fadd A (singleton_matrix j i e)) = apply_matrix (% x. fmul x e) (move_matrix (column_of_matrix A j) 0 (int i))" - apply (simp add: mult_matrix_def) - apply (subst mult_matrix_nm[of _ _ _ "max (ncols A) (Suc j)"]) - apply (auto) - apply (simp add: assms)+ - apply (simp add: mult_matrix_n_def apply_matrix_def apply_infmatrix_def) - apply (rule comb[of "Abs_matrix" "Abs_matrix"], auto, (rule ext)+) - apply (subst foldseq_almostzero[of _ j]) - apply (simp add: assms)+ - apply (auto) - done - -lemma mult_matrix_ext: - assumes - eprem: - "? e. (! a b. a \ b \ fmul a e \ fmul b e)" - and fprems: - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" - "! a. fadd a 0 = a" - "! a. fadd 0 a = a" - and contraprems: - "mult_matrix fmul fadd A = mult_matrix fmul fadd B" - shows - "A = B" -proof(rule contrapos_np[of "False"], simp) - assume a: "A \ B" - have b: "!! f g. (! x y. f x y = g x y) \ f = g" by ((rule ext)+, auto) - have "? j i. (Rep_matrix A j i) \ (Rep_matrix B j i)" - apply (rule contrapos_np[of "False"], simp+) - apply (insert b[of "Rep_matrix A" "Rep_matrix B"], simp) - by (simp add: Rep_matrix_inject a) - then obtain J I where c:"(Rep_matrix A J I) \ (Rep_matrix B J I)" by blast - from eprem obtain e where eprops:"(! a b. a \ b \ fmul a e \ fmul b e)" by blast - let ?S = "singleton_matrix I 0 e" - let ?comp = "mult_matrix fmul fadd" - have d: "!!x f g. f = g \ f x = g x" by blast - have e: "(% x. fmul x e) 0 = 0" by (simp add: assms) - have "~(?comp A ?S = ?comp B ?S)" - apply (rule notI) - apply (simp add: fprems eprops) - apply (simp add: Rep_matrix_inject[THEN sym]) - apply (drule d[of _ _ "J"], drule d[of _ _ "0"]) - by (simp add: e c eprops) - with contraprems show "False" by simp -qed - -definition foldmatrix :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ ('a infmatrix) \ nat \ nat \ 'a" where - "foldmatrix f g A m n == foldseq_transposed g (% j. foldseq f (A j) n) m" - -definition foldmatrix_transposed :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ ('a infmatrix) \ nat \ nat \ 'a" where - "foldmatrix_transposed f g A m n == foldseq g (% j. foldseq_transposed f (A j) n) m" - -lemma foldmatrix_transpose: - assumes - "! a b c d. g(f a b) (f c d) = f (g a c) (g b d)" - shows - "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" -proof - - have forall:"!! P x. (! x. P x) \ P x" by auto - have tworows:"! A. foldmatrix f g A 1 n = foldmatrix_transposed g f (transpose_infmatrix A) n 1" - apply (induct n) - apply (simp add: foldmatrix_def foldmatrix_transposed_def assms)+ - apply (auto) - by (drule_tac x="(% j i. A j (Suc i))" in forall, simp) - show "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" - apply (simp add: foldmatrix_def foldmatrix_transposed_def) - apply (induct m, simp) - apply (simp) - apply (insert tworows) - apply (drule_tac x="% j i. (if j = 0 then (foldseq_transposed g (\u. A u i) m) else (A (Suc m) i))" in spec) - by (simp add: foldmatrix_def foldmatrix_transposed_def) -qed - -lemma foldseq_foldseq: -assumes - "associative f" - "associative g" - "! a b c d. g(f a b) (f c d) = f (g a c) (g b d)" -shows - "foldseq g (% j. foldseq f (A j) n) m = foldseq f (% j. foldseq g ((transpose_infmatrix A) j) m) n" - apply (insert foldmatrix_transpose[of g f A m n]) - by (simp add: foldmatrix_def foldmatrix_transposed_def foldseq_assoc[THEN sym] assms) - -lemma mult_n_nrows: -assumes -"! a. fmul 0 a = 0" -"! a. fmul a 0 = 0" -"fadd 0 0 = 0" -shows "nrows (mult_matrix_n n fmul fadd A B) \ nrows A" -apply (subst nrows_le) -apply (simp add: mult_matrix_n_def) -apply (subst RepAbs_matrix) -apply (rule_tac x="nrows A" in exI) -apply (simp add: nrows assms foldseq_zero) -apply (rule_tac x="ncols B" in exI) -apply (simp add: ncols assms foldseq_zero) -apply (simp add: nrows assms foldseq_zero) -done - -lemma mult_n_ncols: -assumes -"! a. fmul 0 a = 0" -"! a. fmul a 0 = 0" -"fadd 0 0 = 0" -shows "ncols (mult_matrix_n n fmul fadd A B) \ ncols B" -apply (subst ncols_le) -apply (simp add: mult_matrix_n_def) -apply (subst RepAbs_matrix) -apply (rule_tac x="nrows A" in exI) -apply (simp add: nrows assms foldseq_zero) -apply (rule_tac x="ncols B" in exI) -apply (simp add: ncols assms foldseq_zero) -apply (simp add: ncols assms foldseq_zero) -done - -lemma mult_nrows: -assumes -"! a. fmul 0 a = 0" -"! a. fmul a 0 = 0" -"fadd 0 0 = 0" -shows "nrows (mult_matrix fmul fadd A B) \ nrows A" -by (simp add: mult_matrix_def mult_n_nrows assms) - -lemma mult_ncols: -assumes -"! a. fmul 0 a = 0" -"! a. fmul a 0 = 0" -"fadd 0 0 = 0" -shows "ncols (mult_matrix fmul fadd A B) \ ncols B" -by (simp add: mult_matrix_def mult_n_ncols assms) - -lemma nrows_move_matrix_le: "nrows (move_matrix A j i) <= nat((int (nrows A)) + j)" - apply (auto simp add: nrows_le) - apply (rule nrows) - apply (arith) - done - -lemma ncols_move_matrix_le: "ncols (move_matrix A j i) <= nat((int (ncols A)) + i)" - apply (auto simp add: ncols_le) - apply (rule ncols) - apply (arith) - done - -lemma mult_matrix_assoc: - assumes - "! a. fmul1 0 a = 0" - "! a. fmul1 a 0 = 0" - "! a. fmul2 0 a = 0" - "! a. fmul2 a 0 = 0" - "fadd1 0 0 = 0" - "fadd2 0 0 = 0" - "! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" - "associative fadd1" - "associative fadd2" - "! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" - "! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" - "! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" - shows "mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B) C = mult_matrix fmul1 fadd1 A (mult_matrix fmul2 fadd2 B C)" -proof - - have comb_left: "!! A B x y. A = B \ (Rep_matrix (Abs_matrix A)) x y = (Rep_matrix(Abs_matrix B)) x y" by blast - have fmul2fadd1fold: "!! x s n. fmul2 (foldseq fadd1 s n) x = foldseq fadd1 (% k. fmul2 (s k) x) n" - by (rule_tac g1 = "% y. fmul2 y x" in ssubst [OF foldseq_distr_unary], insert assms, simp_all) - have fmul1fadd2fold: "!! x s n. fmul1 x (foldseq fadd2 s n) = foldseq fadd2 (% k. fmul1 x (s k)) n" - using assms by (rule_tac g1 = "% y. fmul1 x y" in ssubst [OF foldseq_distr_unary], simp_all) - let ?N = "max (ncols A) (max (ncols B) (max (nrows B) (nrows C)))" - show ?thesis - apply (simp add: Rep_matrix_inject[THEN sym]) - apply (rule ext)+ - apply (simp add: mult_matrix_def) - apply (simplesubst mult_matrix_nm[of _ "max (ncols (mult_matrix_n (max (ncols A) (nrows B)) fmul1 fadd1 A B)) (nrows C)" _ "max (ncols B) (nrows C)"]) - apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ - apply (simplesubst mult_matrix_nm[of _ "max (ncols A) (nrows (mult_matrix_n (max (ncols B) (nrows C)) fmul2 fadd2 B C))" _ "max (ncols A) (nrows B)"]) - apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ - apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) - apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ - apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) - apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ - apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) - apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ - apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) - apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ - apply (simp add: mult_matrix_n_def) - apply (rule comb_left) - apply ((rule ext)+, simp) - apply (simplesubst RepAbs_matrix) - apply (rule exI[of _ "nrows B"]) - apply (simp add: nrows assms foldseq_zero) - apply (rule exI[of _ "ncols C"]) - apply (simp add: assms ncols foldseq_zero) - apply (subst RepAbs_matrix) - apply (rule exI[of _ "nrows A"]) - apply (simp add: nrows assms foldseq_zero) - apply (rule exI[of _ "ncols B"]) - apply (simp add: assms ncols foldseq_zero) - apply (simp add: fmul2fadd1fold fmul1fadd2fold assms) - apply (subst foldseq_foldseq) - apply (simp add: assms)+ - apply (simp add: transpose_infmatrix) - done -qed - -lemma - assumes - "! a. fmul1 0 a = 0" - "! a. fmul1 a 0 = 0" - "! a. fmul2 0 a = 0" - "! a. fmul2 a 0 = 0" - "fadd1 0 0 = 0" - "fadd2 0 0 = 0" - "! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" - "associative fadd1" - "associative fadd2" - "! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" - "! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" - "! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" - shows - "(mult_matrix fmul1 fadd1 A) o (mult_matrix fmul2 fadd2 B) = mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B)" -apply (rule ext)+ -apply (simp add: comp_def ) -apply (simp add: mult_matrix_assoc assms) -done - -lemma mult_matrix_assoc_simple: - assumes - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" - "fadd 0 0 = 0" - "associative fadd" - "commutative fadd" - "associative fmul" - "distributive fmul fadd" - shows "mult_matrix fmul fadd (mult_matrix fmul fadd A B) C = mult_matrix fmul fadd A (mult_matrix fmul fadd B C)" -proof - - have "!! a b c d. fadd (fadd a b) (fadd c d) = fadd (fadd a c) (fadd b d)" - using assms by (simp add: associative_def commutative_def) - then show ?thesis - apply (subst mult_matrix_assoc) - using assms - apply simp_all - apply (simp_all add: associative_def distributive_def l_distributive_def r_distributive_def) - done -qed - -lemma transpose_apply_matrix: "f 0 = 0 \ transpose_matrix (apply_matrix f A) = apply_matrix f (transpose_matrix A)" -apply (simp add: Rep_matrix_inject[THEN sym]) -apply (rule ext)+ -by simp - -lemma transpose_combine_matrix: "f 0 0 = 0 \ transpose_matrix (combine_matrix f A B) = combine_matrix f (transpose_matrix A) (transpose_matrix B)" -apply (simp add: Rep_matrix_inject[THEN sym]) -apply (rule ext)+ -by simp - -lemma Rep_mult_matrix: - assumes - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" - "fadd 0 0 = 0" - shows - "Rep_matrix(mult_matrix fmul fadd A B) j i = - foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) (max (ncols A) (nrows B))" -apply (simp add: mult_matrix_def mult_matrix_n_def) -apply (subst RepAbs_matrix) -apply (rule exI[of _ "nrows A"], insert assms, simp add: nrows foldseq_zero) -apply (rule exI[of _ "ncols B"], insert assms, simp add: ncols foldseq_zero) -apply simp -done - -lemma transpose_mult_matrix: - assumes - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" - "fadd 0 0 = 0" - "! x y. fmul y x = fmul x y" - shows - "transpose_matrix (mult_matrix fmul fadd A B) = mult_matrix fmul fadd (transpose_matrix B) (transpose_matrix A)" - apply (simp add: Rep_matrix_inject[THEN sym]) - apply (rule ext)+ - using assms - apply (simp add: Rep_mult_matrix max_ac) - done - -lemma column_transpose_matrix: "column_of_matrix (transpose_matrix A) n = transpose_matrix (row_of_matrix A n)" -apply (simp add: Rep_matrix_inject[THEN sym]) -apply (rule ext)+ -by simp - -lemma take_columns_transpose_matrix: "take_columns (transpose_matrix A) n = transpose_matrix (take_rows A n)" -apply (simp add: Rep_matrix_inject[THEN sym]) -apply (rule ext)+ -by simp - -instantiation matrix :: ("{zero, ord}") ord -begin - -definition - le_matrix_def: "A \ B \ (\j i. Rep_matrix A j i \ Rep_matrix B j i)" - -definition - less_def: "A < (B\'a matrix) \ A \ B \ \ B \ A" - -instance .. - -end - -instance matrix :: ("{zero, order}") order -apply intro_classes -apply (simp_all add: le_matrix_def less_def) -apply (auto) -apply (drule_tac x=j in spec, drule_tac x=j in spec) -apply (drule_tac x=i in spec, drule_tac x=i in spec) -apply (simp) -apply (simp add: Rep_matrix_inject[THEN sym]) -apply (rule ext)+ -apply (drule_tac x=xa in spec, drule_tac x=xa in spec) -apply (drule_tac x=xb in spec, drule_tac x=xb in spec) -apply simp -done - -lemma le_apply_matrix: - assumes - "f 0 = 0" - "! x y. x <= y \ f x <= f y" - "(a::('a::{ord, zero}) matrix) <= b" - shows - "apply_matrix f a <= apply_matrix f b" - using assms by (simp add: le_matrix_def) - -lemma le_combine_matrix: - assumes - "f 0 0 = 0" - "! a b c d. a <= b & c <= d \ f a c <= f b d" - "A <= B" - "C <= D" - shows - "combine_matrix f A C <= combine_matrix f B D" - using assms by (simp add: le_matrix_def) - -lemma le_left_combine_matrix: - assumes - "f 0 0 = 0" - "! a b c. a <= b \ f c a <= f c b" - "A <= B" - shows - "combine_matrix f C A <= combine_matrix f C B" - using assms by (simp add: le_matrix_def) - -lemma le_right_combine_matrix: - assumes - "f 0 0 = 0" - "! a b c. a <= b \ f a c <= f b c" - "A <= B" - shows - "combine_matrix f A C <= combine_matrix f B C" - using assms by (simp add: le_matrix_def) - -lemma le_transpose_matrix: "(A <= B) = (transpose_matrix A <= transpose_matrix B)" - by (simp add: le_matrix_def, auto) - -lemma le_foldseq: - assumes - "! a b c d . a <= b & c <= d \ f a c <= f b d" - "! i. i <= n \ s i <= t i" - shows - "foldseq f s n <= foldseq f t n" -proof - - have "! s t. (! i. i<=n \ s i <= t i) \ foldseq f s n <= foldseq f t n" - by (induct n) (simp_all add: assms) - then show "foldseq f s n <= foldseq f t n" using assms by simp -qed - -lemma le_left_mult: - assumes - "! a b c d. a <= b & c <= d \ fadd a c <= fadd b d" - "! c a b. 0 <= c & a <= b \ fmul c a <= fmul c b" - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" - "fadd 0 0 = 0" - "0 <= C" - "A <= B" - shows - "mult_matrix fmul fadd C A <= mult_matrix fmul fadd C B" - using assms - apply (simp add: le_matrix_def Rep_mult_matrix) - apply (auto) - apply (simplesubst foldseq_zerotail[of _ _ _ "max (ncols C) (max (nrows A) (nrows B))"], simp_all add: nrows ncols max1 max2)+ - apply (rule le_foldseq) - apply (auto) - done - -lemma le_right_mult: - assumes - "! a b c d. a <= b & c <= d \ fadd a c <= fadd b d" - "! c a b. 0 <= c & a <= b \ fmul a c <= fmul b c" - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" - "fadd 0 0 = 0" - "0 <= C" - "A <= B" - shows - "mult_matrix fmul fadd A C <= mult_matrix fmul fadd B C" - using assms - apply (simp add: le_matrix_def Rep_mult_matrix) - apply (auto) - apply (simplesubst foldseq_zerotail[of _ _ _ "max (nrows C) (max (ncols A) (ncols B))"], simp_all add: nrows ncols max1 max2)+ - apply (rule le_foldseq) - apply (auto) - done - -lemma spec2: "! j i. P j i \ P j i" by blast -lemma neg_imp: "(\ Q \ \ P) \ P \ Q" by blast - -lemma singleton_matrix_le[simp]: "(singleton_matrix j i a <= singleton_matrix j i b) = (a <= (b::_::order))" - by (auto simp add: le_matrix_def) - -lemma singleton_le_zero[simp]: "(singleton_matrix j i x <= 0) = (x <= (0::'a::{order,zero}))" - apply (auto) - apply (simp add: le_matrix_def) - apply (drule_tac j=j and i=i in spec2) - apply (simp) - apply (simp add: le_matrix_def) - done - -lemma singleton_ge_zero[simp]: "(0 <= singleton_matrix j i x) = ((0::'a::{order,zero}) <= x)" - apply (auto) - apply (simp add: le_matrix_def) - apply (drule_tac j=j and i=i in spec2) - apply (simp) - apply (simp add: le_matrix_def) - done - -lemma move_matrix_le_zero[simp]: "0 <= j \ 0 <= i \ (move_matrix A j i <= 0) = (A <= (0::('a::{order,zero}) matrix))" - apply (auto simp add: le_matrix_def) - apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2) - apply (auto) - done - -lemma move_matrix_zero_le[simp]: "0 <= j \ 0 <= i \ (0 <= move_matrix A j i) = ((0::('a::{order,zero}) matrix) <= A)" - apply (auto simp add: le_matrix_def) - apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2) - apply (auto) - done - -lemma move_matrix_le_move_matrix_iff[simp]: "0 <= j \ 0 <= i \ (move_matrix A j i <= move_matrix B j i) = (A <= (B::('a::{order,zero}) matrix))" - apply (auto simp add: le_matrix_def) - apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2) - apply (auto) - done - -instantiation matrix :: ("{lattice, zero}") lattice -begin - -definition "inf = combine_matrix inf" - -definition "sup = combine_matrix sup" - -instance - by default (auto simp add: le_infI le_matrix_def inf_matrix_def sup_matrix_def) - -end - -instantiation matrix :: ("{plus, zero}") plus -begin - -definition - plus_matrix_def: "A + B = combine_matrix (op +) A B" - -instance .. - -end - -instantiation matrix :: ("{uminus, zero}") uminus -begin - -definition - minus_matrix_def: "- A = apply_matrix uminus A" - -instance .. - -end - -instantiation matrix :: ("{minus, zero}") minus -begin - -definition - diff_matrix_def: "A - B = combine_matrix (op -) A B" - -instance .. - -end - -instantiation matrix :: ("{plus, times, zero}") times -begin - -definition - times_matrix_def: "A * B = mult_matrix (op *) (op +) A B" - -instance .. - -end - -instantiation matrix :: ("{lattice, uminus, zero}") abs -begin - -definition - abs_matrix_def: "abs (A \ 'a matrix) = sup A (- A)" - -instance .. - -end - -instance matrix :: (monoid_add) monoid_add -proof - fix A B C :: "'a matrix" - show "A + B + C = A + (B + C)" - apply (simp add: plus_matrix_def) - apply (rule combine_matrix_assoc[simplified associative_def, THEN spec, THEN spec, THEN spec]) - apply (simp_all add: add_assoc) - done - show "0 + A = A" - apply (simp add: plus_matrix_def) - apply (rule combine_matrix_zero_l_neutral[simplified zero_l_neutral_def, THEN spec]) - apply (simp) - done - show "A + 0 = A" - apply (simp add: plus_matrix_def) - apply (rule combine_matrix_zero_r_neutral [simplified zero_r_neutral_def, THEN spec]) - apply (simp) - done -qed - -instance matrix :: (comm_monoid_add) comm_monoid_add -proof - fix A B :: "'a matrix" - show "A + B = B + A" - apply (simp add: plus_matrix_def) - apply (rule combine_matrix_commute[simplified commutative_def, THEN spec, THEN spec]) - apply (simp_all add: add_commute) - done - show "0 + A = A" - apply (simp add: plus_matrix_def) - apply (rule combine_matrix_zero_l_neutral[simplified zero_l_neutral_def, THEN spec]) - apply (simp) - done -qed - -instance matrix :: (group_add) group_add -proof - fix A B :: "'a matrix" - show "- A + A = 0" - by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext) - show "A - B = A + - B" - by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject [symmetric] ext diff_minus) -qed - -instance matrix :: (ab_group_add) ab_group_add -proof - fix A B :: "'a matrix" - show "- A + A = 0" - by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext) - show "A - B = A + - B" - by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext) -qed - -instance matrix :: (ordered_ab_group_add) ordered_ab_group_add -proof - fix A B C :: "'a matrix" - assume "A <= B" - then show "C + A <= C + B" - apply (simp add: plus_matrix_def) - apply (rule le_left_combine_matrix) - apply (simp_all) - done -qed - -instance matrix :: (lattice_ab_group_add) semilattice_inf_ab_group_add .. -instance matrix :: (lattice_ab_group_add) semilattice_sup_ab_group_add .. - -instance matrix :: (semiring_0) semiring_0 -proof - fix A B C :: "'a matrix" - show "A * B * C = A * (B * C)" - apply (simp add: times_matrix_def) - apply (rule mult_matrix_assoc) - apply (simp_all add: associative_def algebra_simps) - done - show "(A + B) * C = A * C + B * C" - apply (simp add: times_matrix_def plus_matrix_def) - apply (rule l_distributive_matrix[simplified l_distributive_def, THEN spec, THEN spec, THEN spec]) - apply (simp_all add: associative_def commutative_def algebra_simps) - done - show "A * (B + C) = A * B + A * C" - apply (simp add: times_matrix_def plus_matrix_def) - apply (rule r_distributive_matrix[simplified r_distributive_def, THEN spec, THEN spec, THEN spec]) - apply (simp_all add: associative_def commutative_def algebra_simps) - done - show "0 * A = 0" by (simp add: times_matrix_def) - show "A * 0 = 0" by (simp add: times_matrix_def) -qed - -instance matrix :: (ring) ring .. - -instance matrix :: (ordered_ring) ordered_ring -proof - fix A B C :: "'a matrix" - assume a: "A \ B" - assume b: "0 \ C" - from a b show "C * A \ C * B" - apply (simp add: times_matrix_def) - apply (rule le_left_mult) - apply (simp_all add: add_mono mult_left_mono) - done - from a b show "A * C \ B * C" - apply (simp add: times_matrix_def) - apply (rule le_right_mult) - apply (simp_all add: add_mono mult_right_mono) - done -qed - -instance matrix :: (lattice_ring) lattice_ring -proof - fix A B C :: "('a :: lattice_ring) matrix" - show "abs A = sup A (-A)" - by (simp add: abs_matrix_def) -qed - -lemma Rep_matrix_add[simp]: - "Rep_matrix ((a::('a::monoid_add)matrix)+b) j i = (Rep_matrix a j i) + (Rep_matrix b j i)" - by (simp add: plus_matrix_def) - -lemma Rep_matrix_mult: "Rep_matrix ((a::('a::semiring_0) matrix) * b) j i = - foldseq (op +) (% k. (Rep_matrix a j k) * (Rep_matrix b k i)) (max (ncols a) (nrows b))" -apply (simp add: times_matrix_def) -apply (simp add: Rep_mult_matrix) -done - -lemma apply_matrix_add: "! x y. f (x+y) = (f x) + (f y) \ f 0 = (0::'a) - \ apply_matrix f ((a::('a::monoid_add) matrix) + b) = (apply_matrix f a) + (apply_matrix f b)" -apply (subst Rep_matrix_inject[symmetric]) -apply (rule ext)+ -apply (simp) -done - -lemma singleton_matrix_add: "singleton_matrix j i ((a::_::monoid_add)+b) = (singleton_matrix j i a) + (singleton_matrix j i b)" -apply (subst Rep_matrix_inject[symmetric]) -apply (rule ext)+ -apply (simp) -done - -lemma nrows_mult: "nrows ((A::('a::semiring_0) matrix) * B) <= nrows A" -by (simp add: times_matrix_def mult_nrows) - -lemma ncols_mult: "ncols ((A::('a::semiring_0) matrix) * B) <= ncols B" -by (simp add: times_matrix_def mult_ncols) - -definition - one_matrix :: "nat \ ('a::{zero,one}) matrix" where - "one_matrix n = Abs_matrix (% j i. if j = i & j < n then 1 else 0)" - -lemma Rep_one_matrix[simp]: "Rep_matrix (one_matrix n) j i = (if (j = i & j < n) then 1 else 0)" -apply (simp add: one_matrix_def) -apply (simplesubst RepAbs_matrix) -apply (rule exI[of _ n], simp add: split_if)+ -by (simp add: split_if) - -lemma nrows_one_matrix[simp]: "nrows ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _") -proof - - have "?r <= n" by (simp add: nrows_le) - moreover have "n <= ?r" by (simp add:le_nrows, arith) - ultimately show "?r = n" by simp -qed - -lemma ncols_one_matrix[simp]: "ncols ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _") -proof - - have "?r <= n" by (simp add: ncols_le) - moreover have "n <= ?r" by (simp add: le_ncols, arith) - ultimately show "?r = n" by simp -qed - -lemma one_matrix_mult_right[simp]: "ncols A <= n \ (A::('a::{semiring_1}) matrix) * (one_matrix n) = A" -apply (subst Rep_matrix_inject[THEN sym]) -apply (rule ext)+ -apply (simp add: times_matrix_def Rep_mult_matrix) -apply (rule_tac j1="xa" in ssubst[OF foldseq_almostzero]) -apply (simp_all) -by (simp add: ncols) - -lemma one_matrix_mult_left[simp]: "nrows A <= n \ (one_matrix n) * A = (A::('a::semiring_1) matrix)" -apply (subst Rep_matrix_inject[THEN sym]) -apply (rule ext)+ -apply (simp add: times_matrix_def Rep_mult_matrix) -apply (rule_tac j1="x" in ssubst[OF foldseq_almostzero]) -apply (simp_all) -by (simp add: nrows) - -lemma transpose_matrix_mult: "transpose_matrix ((A::('a::comm_ring) matrix)*B) = (transpose_matrix B) * (transpose_matrix A)" -apply (simp add: times_matrix_def) -apply (subst transpose_mult_matrix) -apply (simp_all add: mult_commute) -done - -lemma transpose_matrix_add: "transpose_matrix ((A::('a::monoid_add) matrix)+B) = transpose_matrix A + transpose_matrix B" -by (simp add: plus_matrix_def transpose_combine_matrix) - -lemma transpose_matrix_diff: "transpose_matrix ((A::('a::group_add) matrix)-B) = transpose_matrix A - transpose_matrix B" -by (simp add: diff_matrix_def transpose_combine_matrix) - -lemma transpose_matrix_minus: "transpose_matrix (-(A::('a::group_add) matrix)) = - transpose_matrix (A::'a matrix)" -by (simp add: minus_matrix_def transpose_apply_matrix) - -definition right_inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" where - "right_inverse_matrix A X == (A * X = one_matrix (max (nrows A) (ncols X))) \ nrows X \ ncols A" - -definition left_inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" where - "left_inverse_matrix A X == (X * A = one_matrix (max(nrows X) (ncols A))) \ ncols X \ nrows A" - -definition inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" where - "inverse_matrix A X == (right_inverse_matrix A X) \ (left_inverse_matrix A X)" - -lemma right_inverse_matrix_dim: "right_inverse_matrix A X \ nrows A = ncols X" -apply (insert ncols_mult[of A X], insert nrows_mult[of A X]) -by (simp add: right_inverse_matrix_def) - -lemma left_inverse_matrix_dim: "left_inverse_matrix A Y \ ncols A = nrows Y" -apply (insert ncols_mult[of Y A], insert nrows_mult[of Y A]) -by (simp add: left_inverse_matrix_def) - -lemma left_right_inverse_matrix_unique: - assumes "left_inverse_matrix A Y" "right_inverse_matrix A X" - shows "X = Y" -proof - - have "Y = Y * one_matrix (nrows A)" - apply (subst one_matrix_mult_right) - using assms - apply (simp_all add: left_inverse_matrix_def) - done - also have "\ = Y * (A * X)" - apply (insert assms) - apply (frule right_inverse_matrix_dim) - by (simp add: right_inverse_matrix_def) - also have "\ = (Y * A) * X" by (simp add: mult_assoc) - also have "\ = X" - apply (insert assms) - apply (frule left_inverse_matrix_dim) - apply (simp_all add: left_inverse_matrix_def right_inverse_matrix_def one_matrix_mult_left) - done - ultimately show "X = Y" by (simp) -qed - -lemma inverse_matrix_inject: "\ inverse_matrix A X; inverse_matrix A Y \ \ X = Y" - by (auto simp add: inverse_matrix_def left_right_inverse_matrix_unique) - -lemma one_matrix_inverse: "inverse_matrix (one_matrix n) (one_matrix n)" - by (simp add: inverse_matrix_def left_inverse_matrix_def right_inverse_matrix_def) - -lemma zero_imp_mult_zero: "(a::'a::semiring_0) = 0 | b = 0 \ a * b = 0" -by auto - -lemma Rep_matrix_zero_imp_mult_zero: - "! j i k. (Rep_matrix A j k = 0) | (Rep_matrix B k i) = 0 \ A * B = (0::('a::lattice_ring) matrix)" -apply (subst Rep_matrix_inject[symmetric]) -apply (rule ext)+ -apply (auto simp add: Rep_matrix_mult foldseq_zero zero_imp_mult_zero) -done - -lemma add_nrows: "nrows (A::('a::monoid_add) matrix) <= u \ nrows B <= u \ nrows (A + B) <= u" -apply (simp add: plus_matrix_def) -apply (rule combine_nrows) -apply (simp_all) -done - -lemma move_matrix_row_mult: "move_matrix ((A::('a::semiring_0) matrix) * B) j 0 = (move_matrix A j 0) * B" -apply (subst Rep_matrix_inject[symmetric]) -apply (rule ext)+ -apply (auto simp add: Rep_matrix_mult foldseq_zero) -apply (rule_tac foldseq_zerotail[symmetric]) -apply (auto simp add: nrows zero_imp_mult_zero max2) -apply (rule order_trans) -apply (rule ncols_move_matrix_le) -apply (simp add: max1) -done - -lemma move_matrix_col_mult: "move_matrix ((A::('a::semiring_0) matrix) * B) 0 i = A * (move_matrix B 0 i)" -apply (subst Rep_matrix_inject[symmetric]) -apply (rule ext)+ -apply (auto simp add: Rep_matrix_mult foldseq_zero) -apply (rule_tac foldseq_zerotail[symmetric]) -apply (auto simp add: ncols zero_imp_mult_zero max1) -apply (rule order_trans) -apply (rule nrows_move_matrix_le) -apply (simp add: max2) -done - -lemma move_matrix_add: "((move_matrix (A + B) j i)::(('a::monoid_add) matrix)) = (move_matrix A j i) + (move_matrix B j i)" -apply (subst Rep_matrix_inject[symmetric]) -apply (rule ext)+ -apply (simp) -done - -lemma move_matrix_mult: "move_matrix ((A::('a::semiring_0) matrix)*B) j i = (move_matrix A j 0) * (move_matrix B 0 i)" -by (simp add: move_matrix_ortho[of "A*B"] move_matrix_col_mult move_matrix_row_mult) - -definition scalar_mult :: "('a::ring) \ 'a matrix \ 'a matrix" where - "scalar_mult a m == apply_matrix (op * a) m" - -lemma scalar_mult_zero[simp]: "scalar_mult y 0 = 0" -by (simp add: scalar_mult_def) - -lemma scalar_mult_add: "scalar_mult y (a+b) = (scalar_mult y a) + (scalar_mult y b)" -by (simp add: scalar_mult_def apply_matrix_add algebra_simps) - -lemma Rep_scalar_mult[simp]: "Rep_matrix (scalar_mult y a) j i = y * (Rep_matrix a j i)" -by (simp add: scalar_mult_def) - -lemma scalar_mult_singleton[simp]: "scalar_mult y (singleton_matrix j i x) = singleton_matrix j i (y * x)" -apply (subst Rep_matrix_inject[symmetric]) -apply (rule ext)+ -apply (auto) -done - -lemma Rep_minus[simp]: "Rep_matrix (-(A::_::group_add)) x y = - (Rep_matrix A x y)" -by (simp add: minus_matrix_def) - -lemma Rep_abs[simp]: "Rep_matrix (abs (A::_::lattice_ab_group_add)) x y = abs (Rep_matrix A x y)" -by (simp add: abs_lattice sup_matrix_def) - -end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/ROOT.ML --- a/src/HOL/Matrix/ROOT.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2 +0,0 @@ - -use_thy "Cplex"; diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/SparseMatrix.thy --- a/src/HOL/Matrix/SparseMatrix.thy Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1070 +0,0 @@ -(* Title: HOL/Matrix/SparseMatrix.thy - Author: Steven Obua -*) - -theory SparseMatrix -imports Matrix -begin - -type_synonym 'a spvec = "(nat * 'a) list" -type_synonym 'a spmat = "'a spvec spvec" - -definition sparse_row_vector :: "('a::ab_group_add) spvec \ 'a matrix" - where "sparse_row_vector arr = foldl (% m x. m + (singleton_matrix 0 (fst x) (snd x))) 0 arr" - -definition sparse_row_matrix :: "('a::ab_group_add) spmat \ 'a matrix" - where "sparse_row_matrix arr = foldl (% m r. m + (move_matrix (sparse_row_vector (snd r)) (int (fst r)) 0)) 0 arr" - -code_datatype sparse_row_vector sparse_row_matrix - -lemma sparse_row_vector_empty [simp]: "sparse_row_vector [] = 0" - by (simp add: sparse_row_vector_def) - -lemma sparse_row_matrix_empty [simp]: "sparse_row_matrix [] = 0" - by (simp add: sparse_row_matrix_def) - -lemmas [code] = sparse_row_vector_empty [symmetric] - -lemma foldl_distrstart: "! a x y. (f (g x y) a = g x (f y a)) \ (foldl f (g x y) l = g x (foldl f y l))" - by (induct l arbitrary: x y, auto) - -lemma sparse_row_vector_cons[simp]: - "sparse_row_vector (a # arr) = (singleton_matrix 0 (fst a) (snd a)) + (sparse_row_vector arr)" - apply (induct arr) - apply (auto simp add: sparse_row_vector_def) - apply (simp add: foldl_distrstart [of "\m x. m + singleton_matrix 0 (fst x) (snd x)" "\x m. singleton_matrix 0 (fst x) (snd x) + m"]) - done - -lemma sparse_row_vector_append[simp]: - "sparse_row_vector (a @ b) = (sparse_row_vector a) + (sparse_row_vector b)" - by (induct a) auto - -lemma nrows_spvec[simp]: "nrows (sparse_row_vector x) <= (Suc 0)" - apply (induct x) - apply (simp_all add: add_nrows) - done - -lemma sparse_row_matrix_cons: "sparse_row_matrix (a#arr) = ((move_matrix (sparse_row_vector (snd a)) (int (fst a)) 0)) + sparse_row_matrix arr" - apply (induct arr) - apply (auto simp add: sparse_row_matrix_def) - apply (simp add: foldl_distrstart[of "\m x. m + (move_matrix (sparse_row_vector (snd x)) (int (fst x)) 0)" - "% a m. (move_matrix (sparse_row_vector (snd a)) (int (fst a)) 0) + m"]) - done - -lemma sparse_row_matrix_append: "sparse_row_matrix (arr@brr) = (sparse_row_matrix arr) + (sparse_row_matrix brr)" - apply (induct arr) - apply (auto simp add: sparse_row_matrix_cons) - done - -primrec sorted_spvec :: "'a spvec \ bool" -where - "sorted_spvec [] = True" -| sorted_spvec_step: "sorted_spvec (a#as) = (case as of [] \ True | b#bs \ ((fst a < fst b) & (sorted_spvec as)))" - -primrec sorted_spmat :: "'a spmat \ bool" -where - "sorted_spmat [] = True" -| "sorted_spmat (a#as) = ((sorted_spvec (snd a)) & (sorted_spmat as))" - -declare sorted_spvec.simps [simp del] - -lemma sorted_spvec_empty[simp]: "sorted_spvec [] = True" -by (simp add: sorted_spvec.simps) - -lemma sorted_spvec_cons1: "sorted_spvec (a#as) \ sorted_spvec as" -apply (induct as) -apply (auto simp add: sorted_spvec.simps) -done - -lemma sorted_spvec_cons2: "sorted_spvec (a#b#t) \ sorted_spvec (a#t)" -apply (induct t) -apply (auto simp add: sorted_spvec.simps) -done - -lemma sorted_spvec_cons3: "sorted_spvec(a#b#t) \ fst a < fst b" -apply (auto simp add: sorted_spvec.simps) -done - -lemma sorted_sparse_row_vector_zero[rule_format]: "m <= n \ sorted_spvec ((n,a)#arr) \ Rep_matrix (sparse_row_vector arr) j m = 0" -apply (induct arr) -apply (auto) -apply (frule sorted_spvec_cons2,simp)+ -apply (frule sorted_spvec_cons3, simp) -done - -lemma sorted_sparse_row_matrix_zero[rule_format]: "m <= n \ sorted_spvec ((n,a)#arr) \ Rep_matrix (sparse_row_matrix arr) m j = 0" - apply (induct arr) - apply (auto) - apply (frule sorted_spvec_cons2, simp) - apply (frule sorted_spvec_cons3, simp) - apply (simp add: sparse_row_matrix_cons) - done - -primrec minus_spvec :: "('a::ab_group_add) spvec \ 'a spvec" -where - "minus_spvec [] = []" -| "minus_spvec (a#as) = (fst a, -(snd a))#(minus_spvec as)" - -primrec abs_spvec :: "('a::lattice_ab_group_add_abs) spvec \ 'a spvec" -where - "abs_spvec [] = []" -| "abs_spvec (a#as) = (fst a, abs (snd a))#(abs_spvec as)" - -lemma sparse_row_vector_minus: - "sparse_row_vector (minus_spvec v) = - (sparse_row_vector v)" - apply (induct v) - apply (simp_all add: sparse_row_vector_cons) - apply (simp add: Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply simp - done - -instance matrix :: (lattice_ab_group_add_abs) lattice_ab_group_add_abs -apply default -unfolding abs_matrix_def .. (*FIXME move*) - -lemma sparse_row_vector_abs: - "sorted_spvec (v :: 'a::lattice_ring spvec) \ sparse_row_vector (abs_spvec v) = abs (sparse_row_vector v)" - apply (induct v) - apply simp_all - apply (frule_tac sorted_spvec_cons1, simp) - apply (simp only: Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply auto - apply (subgoal_tac "Rep_matrix (sparse_row_vector v) 0 a = 0") - apply (simp) - apply (rule sorted_sparse_row_vector_zero) - apply auto - done - -lemma sorted_spvec_minus_spvec: - "sorted_spvec v \ sorted_spvec (minus_spvec v)" - apply (induct v) - apply (simp) - apply (frule sorted_spvec_cons1, simp) - apply (simp add: sorted_spvec.simps split:list.split_asm) - done - -lemma sorted_spvec_abs_spvec: - "sorted_spvec v \ sorted_spvec (abs_spvec v)" - apply (induct v) - apply (simp) - apply (frule sorted_spvec_cons1, simp) - apply (simp add: sorted_spvec.simps split:list.split_asm) - done - -definition "smult_spvec y = map (% a. (fst a, y * snd a))" - -lemma smult_spvec_empty[simp]: "smult_spvec y [] = []" - by (simp add: smult_spvec_def) - -lemma smult_spvec_cons: "smult_spvec y (a#arr) = (fst a, y * (snd a)) # (smult_spvec y arr)" - by (simp add: smult_spvec_def) - -fun addmult_spvec :: "('a::ring) \ 'a spvec \ 'a spvec \ 'a spvec" -where - "addmult_spvec y arr [] = arr" -| "addmult_spvec y [] brr = smult_spvec y brr" -| "addmult_spvec y ((i,a)#arr) ((j,b)#brr) = ( - if i < j then ((i,a)#(addmult_spvec y arr ((j,b)#brr))) - else (if (j < i) then ((j, y * b)#(addmult_spvec y ((i,a)#arr) brr)) - else ((i, a + y*b)#(addmult_spvec y arr brr))))" -(* Steven used termination "measure (% (y, a, b). length a + (length b))" *) - -lemma addmult_spvec_empty1[simp]: "addmult_spvec y [] a = smult_spvec y a" - by (induct a) auto - -lemma addmult_spvec_empty2[simp]: "addmult_spvec y a [] = a" - by (induct a) auto - -lemma sparse_row_vector_map: "(! x y. f (x+y) = (f x) + (f y)) \ (f::'a\('a::lattice_ring)) 0 = 0 \ - sparse_row_vector (map (% x. (fst x, f (snd x))) a) = apply_matrix f (sparse_row_vector a)" - apply (induct a) - apply (simp_all add: apply_matrix_add) - done - -lemma sparse_row_vector_smult: "sparse_row_vector (smult_spvec y a) = scalar_mult y (sparse_row_vector a)" - apply (induct a) - apply (simp_all add: smult_spvec_cons scalar_mult_add) - done - -lemma sparse_row_vector_addmult_spvec: "sparse_row_vector (addmult_spvec (y::'a::lattice_ring) a b) = - (sparse_row_vector a) + (scalar_mult y (sparse_row_vector b))" - apply (induct y a b rule: addmult_spvec.induct) - apply (simp add: scalar_mult_add smult_spvec_cons sparse_row_vector_smult singleton_matrix_add)+ - done - -lemma sorted_smult_spvec: "sorted_spvec a \ sorted_spvec (smult_spvec y a)" - apply (auto simp add: smult_spvec_def) - apply (induct a) - apply (auto simp add: sorted_spvec.simps split:list.split_asm) - done - -lemma sorted_spvec_addmult_spvec_helper: "\sorted_spvec (addmult_spvec y ((a, b) # arr) brr); aa < a; sorted_spvec ((a, b) # arr); - sorted_spvec ((aa, ba) # brr)\ \ sorted_spvec ((aa, y * ba) # addmult_spvec y ((a, b) # arr) brr)" - apply (induct brr) - apply (auto simp add: sorted_spvec.simps) - done - -lemma sorted_spvec_addmult_spvec_helper2: - "\sorted_spvec (addmult_spvec y arr ((aa, ba) # brr)); a < aa; sorted_spvec ((a, b) # arr); sorted_spvec ((aa, ba) # brr)\ - \ sorted_spvec ((a, b) # addmult_spvec y arr ((aa, ba) # brr))" - apply (induct arr) - apply (auto simp add: smult_spvec_def sorted_spvec.simps) - done - -lemma sorted_spvec_addmult_spvec_helper3[rule_format]: - "sorted_spvec (addmult_spvec y arr brr) \ sorted_spvec ((aa, b) # arr) \ sorted_spvec ((aa, ba) # brr) - \ sorted_spvec ((aa, b + y * ba) # (addmult_spvec y arr brr))" - apply (induct y arr brr rule: addmult_spvec.induct) - apply (simp_all add: sorted_spvec.simps smult_spvec_def split:list.split) - done - -lemma sorted_addmult_spvec: "sorted_spvec a \ sorted_spvec b \ sorted_spvec (addmult_spvec y a b)" - apply (induct y a b rule: addmult_spvec.induct) - apply (simp_all add: sorted_smult_spvec) - apply (rule conjI, intro strip) - apply (case_tac "~(i < j)") - apply (simp_all) - apply (frule_tac as=brr in sorted_spvec_cons1) - apply (simp add: sorted_spvec_addmult_spvec_helper) - apply (intro strip | rule conjI)+ - apply (frule_tac as=arr in sorted_spvec_cons1) - apply (simp add: sorted_spvec_addmult_spvec_helper2) - apply (intro strip) - apply (frule_tac as=arr in sorted_spvec_cons1) - apply (frule_tac as=brr in sorted_spvec_cons1) - apply (simp) - apply (simp_all add: sorted_spvec_addmult_spvec_helper3) - done - -fun mult_spvec_spmat :: "('a::lattice_ring) spvec \ 'a spvec \ 'a spmat \ 'a spvec" -where -(* recdef mult_spvec_spmat "measure (% (c, arr, brr). (length arr) + (length brr))" *) - "mult_spvec_spmat c [] brr = c" -| "mult_spvec_spmat c arr [] = c" -| "mult_spvec_spmat c ((i,a)#arr) ((j,b)#brr) = ( - if (i < j) then mult_spvec_spmat c arr ((j,b)#brr) - else if (j < i) then mult_spvec_spmat c ((i,a)#arr) brr - else mult_spvec_spmat (addmult_spvec a c b) arr brr)" - -lemma sparse_row_mult_spvec_spmat[rule_format]: "sorted_spvec (a::('a::lattice_ring) spvec) \ sorted_spvec B \ - sparse_row_vector (mult_spvec_spmat c a B) = (sparse_row_vector c) + (sparse_row_vector a) * (sparse_row_matrix B)" -proof - - have comp_1: "!! a b. a < b \ Suc 0 <= nat ((int b)-(int a))" by arith - have not_iff: "!! a b. a = b \ (~ a) = (~ b)" by simp - have max_helper: "!! a b. ~ (a <= max (Suc a) b) \ False" - by arith - { - fix a - fix v - assume a:"a < nrows(sparse_row_vector v)" - have b:"nrows(sparse_row_vector v) <= 1" by simp - note dummy = less_le_trans[of a "nrows (sparse_row_vector v)" 1, OF a b] - then have "a = 0" by simp - } - note nrows_helper = this - show ?thesis - apply (induct c a B rule: mult_spvec_spmat.induct) - apply simp+ - apply (rule conjI) - apply (intro strip) - apply (frule_tac as=brr in sorted_spvec_cons1) - apply (simp add: algebra_simps sparse_row_matrix_cons) - apply (simplesubst Rep_matrix_zero_imp_mult_zero) - apply (simp) - apply (rule disjI2) - apply (intro strip) - apply (subst nrows) - apply (rule order_trans[of _ 1]) - apply (simp add: comp_1)+ - apply (subst Rep_matrix_zero_imp_mult_zero) - apply (intro strip) - apply (case_tac "k <= j") - apply (rule_tac m1 = k and n1 = i and a1 = a in ssubst[OF sorted_sparse_row_vector_zero]) - apply (simp_all) - apply (rule disjI2) - apply (rule nrows) - apply (rule order_trans[of _ 1]) - apply (simp_all add: comp_1) - - apply (intro strip | rule conjI)+ - apply (frule_tac as=arr in sorted_spvec_cons1) - apply (simp add: algebra_simps) - apply (subst Rep_matrix_zero_imp_mult_zero) - apply (simp) - apply (rule disjI2) - apply (intro strip) - apply (simp add: sparse_row_matrix_cons) - apply (case_tac "i <= j") - apply (erule sorted_sparse_row_matrix_zero) - apply (simp_all) - apply (intro strip) - apply (case_tac "i=j") - apply (simp_all) - apply (frule_tac as=arr in sorted_spvec_cons1) - apply (frule_tac as=brr in sorted_spvec_cons1) - apply (simp add: sparse_row_matrix_cons algebra_simps sparse_row_vector_addmult_spvec) - apply (rule_tac B1 = "sparse_row_matrix brr" in ssubst[OF Rep_matrix_zero_imp_mult_zero]) - apply (auto) - apply (rule sorted_sparse_row_matrix_zero) - apply (simp_all) - apply (rule_tac A1 = "sparse_row_vector arr" in ssubst[OF Rep_matrix_zero_imp_mult_zero]) - apply (auto) - apply (rule_tac m=k and n = j and a = a and arr=arr in sorted_sparse_row_vector_zero) - apply (simp_all) - apply (drule nrows_notzero) - apply (drule nrows_helper) - apply (arith) - - apply (subst Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply (simp) - apply (subst Rep_matrix_mult) - apply (rule_tac j1=j in ssubst[OF foldseq_almostzero]) - apply (simp_all) - apply (intro strip, rule conjI) - apply (intro strip) - apply (drule_tac max_helper) - apply (simp) - apply (auto) - apply (rule zero_imp_mult_zero) - apply (rule disjI2) - apply (rule nrows) - apply (rule order_trans[of _ 1]) - apply (simp) - apply (simp) - done -qed - -lemma sorted_mult_spvec_spmat[rule_format]: - "sorted_spvec (c::('a::lattice_ring) spvec) \ sorted_spmat B \ sorted_spvec (mult_spvec_spmat c a B)" - apply (induct c a B rule: mult_spvec_spmat.induct) - apply (simp_all add: sorted_addmult_spvec) - done - -primrec mult_spmat :: "('a::lattice_ring) spmat \ 'a spmat \ 'a spmat" -where - "mult_spmat [] A = []" -| "mult_spmat (a#as) A = (fst a, mult_spvec_spmat [] (snd a) A)#(mult_spmat as A)" - -lemma sparse_row_mult_spmat: - "sorted_spmat A \ sorted_spvec B \ - sparse_row_matrix (mult_spmat A B) = (sparse_row_matrix A) * (sparse_row_matrix B)" - apply (induct A) - apply (auto simp add: sparse_row_matrix_cons sparse_row_mult_spvec_spmat algebra_simps move_matrix_mult) - done - -lemma sorted_spvec_mult_spmat[rule_format]: - "sorted_spvec (A::('a::lattice_ring) spmat) \ sorted_spvec (mult_spmat A B)" - apply (induct A) - apply (auto) - apply (drule sorted_spvec_cons1, simp) - apply (case_tac A) - apply (auto simp add: sorted_spvec.simps) - done - -lemma sorted_spmat_mult_spmat: - "sorted_spmat (B::('a::lattice_ring) spmat) \ sorted_spmat (mult_spmat A B)" - apply (induct A) - apply (auto simp add: sorted_mult_spvec_spmat) - done - - -fun add_spvec :: "('a::lattice_ab_group_add) spvec \ 'a spvec \ 'a spvec" -where -(* "measure (% (a, b). length a + (length b))" *) - "add_spvec arr [] = arr" -| "add_spvec [] brr = brr" -| "add_spvec ((i,a)#arr) ((j,b)#brr) = ( - if i < j then (i,a)#(add_spvec arr ((j,b)#brr)) - else if (j < i) then (j,b) # add_spvec ((i,a)#arr) brr - else (i, a+b) # add_spvec arr brr)" - -lemma add_spvec_empty1[simp]: "add_spvec [] a = a" -by (cases a, auto) - -lemma sparse_row_vector_add: "sparse_row_vector (add_spvec a b) = (sparse_row_vector a) + (sparse_row_vector b)" - apply (induct a b rule: add_spvec.induct) - apply (simp_all add: singleton_matrix_add) - done - -fun add_spmat :: "('a::lattice_ab_group_add) spmat \ 'a spmat \ 'a spmat" -where -(* "measure (% (A,B). (length A)+(length B))" *) - "add_spmat [] bs = bs" -| "add_spmat as [] = as" -| "add_spmat ((i,a)#as) ((j,b)#bs) = ( - if i < j then - (i,a) # add_spmat as ((j,b)#bs) - else if j < i then - (j,b) # add_spmat ((i,a)#as) bs - else - (i, add_spvec a b) # add_spmat as bs)" - -lemma add_spmat_Nil2[simp]: "add_spmat as [] = as" -by(cases as) auto - -lemma sparse_row_add_spmat: "sparse_row_matrix (add_spmat A B) = (sparse_row_matrix A) + (sparse_row_matrix B)" - apply (induct A B rule: add_spmat.induct) - apply (auto simp add: sparse_row_matrix_cons sparse_row_vector_add move_matrix_add) - done - -lemmas [code] = sparse_row_add_spmat [symmetric] -lemmas [code] = sparse_row_vector_add [symmetric] - -lemma sorted_add_spvec_helper1[rule_format]: "add_spvec ((a,b)#arr) brr = (ab, bb) # list \ (ab = a | (brr \ [] & ab = fst (hd brr)))" - proof - - have "(! x ab a. x = (a,b)#arr \ add_spvec x brr = (ab, bb) # list \ (ab = a | (ab = fst (hd brr))))" - by (induct brr rule: add_spvec.induct) (auto split:if_splits) - then show ?thesis - by (case_tac brr, auto) - qed - -lemma sorted_add_spmat_helper1[rule_format]: "add_spmat ((a,b)#arr) brr = (ab, bb) # list \ (ab = a | (brr \ [] & ab = fst (hd brr)))" - proof - - have "(! x ab a. x = (a,b)#arr \ add_spmat x brr = (ab, bb) # list \ (ab = a | (ab = fst (hd brr))))" - by (rule add_spmat.induct) (auto split:if_splits) - then show ?thesis - by (case_tac brr, auto) - qed - -lemma sorted_add_spvec_helper: "add_spvec arr brr = (ab, bb) # list \ ((arr \ [] & ab = fst (hd arr)) | (brr \ [] & ab = fst (hd brr)))" - apply (induct arr brr rule: add_spvec.induct) - apply (auto split:if_splits) - done - -lemma sorted_add_spmat_helper: "add_spmat arr brr = (ab, bb) # list \ ((arr \ [] & ab = fst (hd arr)) | (brr \ [] & ab = fst (hd brr)))" - apply (induct arr brr rule: add_spmat.induct) - apply (auto split:if_splits) - done - -lemma add_spvec_commute: "add_spvec a b = add_spvec b a" -by (induct a b rule: add_spvec.induct) auto - -lemma add_spmat_commute: "add_spmat a b = add_spmat b a" - apply (induct a b rule: add_spmat.induct) - apply (simp_all add: add_spvec_commute) - done - -lemma sorted_add_spvec_helper2: "add_spvec ((a,b)#arr) brr = (ab, bb) # list \ aa < a \ sorted_spvec ((aa, ba) # brr) \ aa < ab" - apply (drule sorted_add_spvec_helper1) - apply (auto) - apply (case_tac brr) - apply (simp_all) - apply (drule_tac sorted_spvec_cons3) - apply (simp) - done - -lemma sorted_add_spmat_helper2: "add_spmat ((a,b)#arr) brr = (ab, bb) # list \ aa < a \ sorted_spvec ((aa, ba) # brr) \ aa < ab" - apply (drule sorted_add_spmat_helper1) - apply (auto) - apply (case_tac brr) - apply (simp_all) - apply (drule_tac sorted_spvec_cons3) - apply (simp) - done - -lemma sorted_spvec_add_spvec[rule_format]: "sorted_spvec a \ sorted_spvec b \ sorted_spvec (add_spvec a b)" - apply (induct a b rule: add_spvec.induct) - apply (simp_all) - apply (rule conjI) - apply (clarsimp) - apply (frule_tac as=brr in sorted_spvec_cons1) - apply (simp) - apply (subst sorted_spvec_step) - apply (clarsimp simp: sorted_add_spvec_helper2 split: list.split) - apply (clarify) - apply (rule conjI) - apply (clarify) - apply (frule_tac as=arr in sorted_spvec_cons1, simp) - apply (subst sorted_spvec_step) - apply (clarsimp simp: sorted_add_spvec_helper2 add_spvec_commute split: list.split) - apply (clarify) - apply (frule_tac as=arr in sorted_spvec_cons1) - apply (frule_tac as=brr in sorted_spvec_cons1) - apply (simp) - apply (subst sorted_spvec_step) - apply (simp split: list.split) - apply (clarsimp) - apply (drule_tac sorted_add_spvec_helper) - apply (auto simp: neq_Nil_conv) - apply (drule sorted_spvec_cons3) - apply (simp) - apply (drule sorted_spvec_cons3) - apply (simp) - done - -lemma sorted_spvec_add_spmat[rule_format]: "sorted_spvec A \ sorted_spvec B \ sorted_spvec (add_spmat A B)" - apply (induct A B rule: add_spmat.induct) - apply (simp_all) - apply (rule conjI) - apply (intro strip) - apply (simp) - apply (frule_tac as=bs in sorted_spvec_cons1) - apply (simp) - apply (subst sorted_spvec_step) - apply (simp split: list.split) - apply (clarify, simp) - apply (simp add: sorted_add_spmat_helper2) - apply (clarify) - apply (rule conjI) - apply (clarify) - apply (frule_tac as=as in sorted_spvec_cons1, simp) - apply (subst sorted_spvec_step) - apply (clarsimp simp: sorted_add_spmat_helper2 add_spmat_commute split: list.split) - apply (clarsimp) - apply (frule_tac as=as in sorted_spvec_cons1) - apply (frule_tac as=bs in sorted_spvec_cons1) - apply (simp) - apply (subst sorted_spvec_step) - apply (simp split: list.split) - apply (clarify, simp) - apply (drule_tac sorted_add_spmat_helper) - apply (auto simp:neq_Nil_conv) - apply (drule sorted_spvec_cons3) - apply (simp) - apply (drule sorted_spvec_cons3) - apply (simp) - done - -lemma sorted_spmat_add_spmat[rule_format]: "sorted_spmat A \ sorted_spmat B \ sorted_spmat (add_spmat A B)" - apply (induct A B rule: add_spmat.induct) - apply (simp_all add: sorted_spvec_add_spvec) - done - -fun le_spvec :: "('a::lattice_ab_group_add) spvec \ 'a spvec \ bool" -where -(* "measure (% (a,b). (length a) + (length b))" *) - "le_spvec [] [] = True" -| "le_spvec ((_,a)#as) [] = (a <= 0 & le_spvec as [])" -| "le_spvec [] ((_,b)#bs) = (0 <= b & le_spvec [] bs)" -| "le_spvec ((i,a)#as) ((j,b)#bs) = ( - if (i < j) then a <= 0 & le_spvec as ((j,b)#bs) - else if (j < i) then 0 <= b & le_spvec ((i,a)#as) bs - else a <= b & le_spvec as bs)" - -fun le_spmat :: "('a::lattice_ab_group_add) spmat \ 'a spmat \ bool" -where -(* "measure (% (a,b). (length a) + (length b))" *) - "le_spmat [] [] = True" -| "le_spmat ((i,a)#as) [] = (le_spvec a [] & le_spmat as [])" -| "le_spmat [] ((j,b)#bs) = (le_spvec [] b & le_spmat [] bs)" -| "le_spmat ((i,a)#as) ((j,b)#bs) = ( - if i < j then (le_spvec a [] & le_spmat as ((j,b)#bs)) - else if j < i then (le_spvec [] b & le_spmat ((i,a)#as) bs) - else (le_spvec a b & le_spmat as bs))" - -definition disj_matrices :: "('a::zero) matrix \ 'a matrix \ bool" where - "disj_matrices A B \ - (! j i. (Rep_matrix A j i \ 0) \ (Rep_matrix B j i = 0)) & (! j i. (Rep_matrix B j i \ 0) \ (Rep_matrix A j i = 0))" - -declare [[simp_depth_limit = 6]] - -lemma disj_matrices_contr1: "disj_matrices A B \ Rep_matrix A j i \ 0 \ Rep_matrix B j i = 0" - by (simp add: disj_matrices_def) - -lemma disj_matrices_contr2: "disj_matrices A B \ Rep_matrix B j i \ 0 \ Rep_matrix A j i = 0" - by (simp add: disj_matrices_def) - - -lemma disj_matrices_add: "disj_matrices A B \ disj_matrices C D \ disj_matrices A D \ disj_matrices B C \ - (A + B <= C + D) = (A <= C & B <= (D::('a::lattice_ab_group_add) matrix))" - apply (auto) - apply (simp (no_asm_use) only: le_matrix_def disj_matrices_def) - apply (intro strip) - apply (erule conjE)+ - apply (drule_tac j=j and i=i in spec2)+ - apply (case_tac "Rep_matrix B j i = 0") - apply (case_tac "Rep_matrix D j i = 0") - apply (simp_all) - apply (simp (no_asm_use) only: le_matrix_def disj_matrices_def) - apply (intro strip) - apply (erule conjE)+ - apply (drule_tac j=j and i=i in spec2)+ - apply (case_tac "Rep_matrix A j i = 0") - apply (case_tac "Rep_matrix C j i = 0") - apply (simp_all) - apply (erule add_mono) - apply (assumption) - done - -lemma disj_matrices_zero1[simp]: "disj_matrices 0 B" -by (simp add: disj_matrices_def) - -lemma disj_matrices_zero2[simp]: "disj_matrices A 0" -by (simp add: disj_matrices_def) - -lemma disj_matrices_commute: "disj_matrices A B = disj_matrices B A" -by (auto simp add: disj_matrices_def) - -lemma disj_matrices_add_le_zero: "disj_matrices A B \ - (A + B <= 0) = (A <= 0 & (B::('a::lattice_ab_group_add) matrix) <= 0)" -by (rule disj_matrices_add[of A B 0 0, simplified]) - -lemma disj_matrices_add_zero_le: "disj_matrices A B \ - (0 <= A + B) = (0 <= A & 0 <= (B::('a::lattice_ab_group_add) matrix))" -by (rule disj_matrices_add[of 0 0 A B, simplified]) - -lemma disj_matrices_add_x_le: "disj_matrices A B \ disj_matrices B C \ - (A <= B + C) = (A <= C & 0 <= (B::('a::lattice_ab_group_add) matrix))" -by (auto simp add: disj_matrices_add[of 0 A B C, simplified]) - -lemma disj_matrices_add_le_x: "disj_matrices A B \ disj_matrices B C \ - (B + A <= C) = (A <= C & (B::('a::lattice_ab_group_add) matrix) <= 0)" -by (auto simp add: disj_matrices_add[of B A 0 C,simplified] disj_matrices_commute) - -lemma disj_sparse_row_singleton: "i <= j \ sorted_spvec((j,y)#v) \ disj_matrices (sparse_row_vector v) (singleton_matrix 0 i x)" - apply (simp add: disj_matrices_def) - apply (rule conjI) - apply (rule neg_imp) - apply (simp) - apply (intro strip) - apply (rule sorted_sparse_row_vector_zero) - apply (simp_all) - apply (intro strip) - apply (rule sorted_sparse_row_vector_zero) - apply (simp_all) - done - -lemma disj_matrices_x_add: "disj_matrices A B \ disj_matrices A C \ disj_matrices (A::('a::lattice_ab_group_add) matrix) (B+C)" - apply (simp add: disj_matrices_def) - apply (auto) - apply (drule_tac j=j and i=i in spec2)+ - apply (case_tac "Rep_matrix B j i = 0") - apply (case_tac "Rep_matrix C j i = 0") - apply (simp_all) - done - -lemma disj_matrices_add_x: "disj_matrices A B \ disj_matrices A C \ disj_matrices (B+C) (A::('a::lattice_ab_group_add) matrix)" - by (simp add: disj_matrices_x_add disj_matrices_commute) - -lemma disj_singleton_matrices[simp]: "disj_matrices (singleton_matrix j i x) (singleton_matrix u v y) = (j \ u | i \ v | x = 0 | y = 0)" - by (auto simp add: disj_matrices_def) - -lemma disj_move_sparse_vec_mat[simplified disj_matrices_commute]: - "j <= a \ sorted_spvec((a,c)#as) \ disj_matrices (move_matrix (sparse_row_vector b) (int j) i) (sparse_row_matrix as)" - apply (auto simp add: disj_matrices_def) - apply (drule nrows_notzero) - apply (drule less_le_trans[OF _ nrows_spvec]) - apply (subgoal_tac "ja = j") - apply (simp add: sorted_sparse_row_matrix_zero) - apply (arith) - apply (rule nrows) - apply (rule order_trans[of _ 1 _]) - apply (simp) - apply (case_tac "nat (int ja - int j) = 0") - apply (case_tac "ja = j") - apply (simp add: sorted_sparse_row_matrix_zero) - apply arith+ - done - -lemma disj_move_sparse_row_vector_twice: - "j \ u \ disj_matrices (move_matrix (sparse_row_vector a) j i) (move_matrix (sparse_row_vector b) u v)" - apply (auto simp add: disj_matrices_def) - apply (rule nrows, rule order_trans[of _ 1], simp, drule nrows_notzero, drule less_le_trans[OF _ nrows_spvec], arith)+ - done - -lemma le_spvec_iff_sparse_row_le[rule_format]: "(sorted_spvec a) \ (sorted_spvec b) \ (le_spvec a b) = (sparse_row_vector a <= sparse_row_vector b)" - apply (induct a b rule: le_spvec.induct) - apply (simp_all add: sorted_spvec_cons1 disj_matrices_add_le_zero disj_matrices_add_zero_le - disj_sparse_row_singleton[OF order_refl] disj_matrices_commute) - apply (rule conjI, intro strip) - apply (simp add: sorted_spvec_cons1) - apply (subst disj_matrices_add_x_le) - apply (simp add: disj_sparse_row_singleton[OF less_imp_le] disj_matrices_x_add disj_matrices_commute) - apply (simp add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute) - apply (simp, blast) - apply (intro strip, rule conjI, intro strip) - apply (simp add: sorted_spvec_cons1) - apply (subst disj_matrices_add_le_x) - apply (simp_all add: disj_sparse_row_singleton[OF order_refl] disj_sparse_row_singleton[OF less_imp_le] disj_matrices_commute disj_matrices_x_add) - apply (blast) - apply (intro strip) - apply (simp add: sorted_spvec_cons1) - apply (case_tac "a=b", simp_all) - apply (subst disj_matrices_add) - apply (simp_all add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute) - done - -lemma le_spvec_empty2_sparse_row[rule_format]: "sorted_spvec b \ le_spvec b [] = (sparse_row_vector b <= 0)" - apply (induct b) - apply (simp_all add: sorted_spvec_cons1) - apply (intro strip) - apply (subst disj_matrices_add_le_zero) - apply (auto simp add: disj_matrices_commute disj_sparse_row_singleton[OF order_refl] sorted_spvec_cons1) - done - -lemma le_spvec_empty1_sparse_row[rule_format]: "(sorted_spvec b) \ (le_spvec [] b = (0 <= sparse_row_vector b))" - apply (induct b) - apply (simp_all add: sorted_spvec_cons1) - apply (intro strip) - apply (subst disj_matrices_add_zero_le) - apply (auto simp add: disj_matrices_commute disj_sparse_row_singleton[OF order_refl] sorted_spvec_cons1) - done - -lemma le_spmat_iff_sparse_row_le[rule_format]: "(sorted_spvec A) \ (sorted_spmat A) \ (sorted_spvec B) \ (sorted_spmat B) \ - le_spmat A B = (sparse_row_matrix A <= sparse_row_matrix B)" - apply (induct A B rule: le_spmat.induct) - apply (simp add: sparse_row_matrix_cons disj_matrices_add_le_zero disj_matrices_add_zero_le disj_move_sparse_vec_mat[OF order_refl] - disj_matrices_commute sorted_spvec_cons1 le_spvec_empty2_sparse_row le_spvec_empty1_sparse_row)+ - apply (rule conjI, intro strip) - apply (simp add: sorted_spvec_cons1) - apply (subst disj_matrices_add_x_le) - apply (rule disj_matrices_add_x) - apply (simp add: disj_move_sparse_row_vector_twice) - apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute) - apply (simp add: disj_move_sparse_vec_mat[OF order_refl] disj_matrices_commute) - apply (simp, blast) - apply (intro strip, rule conjI, intro strip) - apply (simp add: sorted_spvec_cons1) - apply (subst disj_matrices_add_le_x) - apply (simp add: disj_move_sparse_vec_mat[OF order_refl]) - apply (rule disj_matrices_x_add) - apply (simp add: disj_move_sparse_row_vector_twice) - apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute) - apply (simp, blast) - apply (intro strip) - apply (case_tac "i=j") - apply (simp_all) - apply (subst disj_matrices_add) - apply (simp_all add: disj_matrices_commute disj_move_sparse_vec_mat[OF order_refl]) - apply (simp add: sorted_spvec_cons1 le_spvec_iff_sparse_row_le) - done - -declare [[simp_depth_limit = 999]] - -primrec abs_spmat :: "('a::lattice_ring) spmat \ 'a spmat" -where - "abs_spmat [] = []" -| "abs_spmat (a#as) = (fst a, abs_spvec (snd a))#(abs_spmat as)" - -primrec minus_spmat :: "('a::lattice_ring) spmat \ 'a spmat" -where - "minus_spmat [] = []" -| "minus_spmat (a#as) = (fst a, minus_spvec (snd a))#(minus_spmat as)" - -lemma sparse_row_matrix_minus: - "sparse_row_matrix (minus_spmat A) = - (sparse_row_matrix A)" - apply (induct A) - apply (simp_all add: sparse_row_vector_minus sparse_row_matrix_cons) - apply (subst Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply simp - done - -lemma Rep_sparse_row_vector_zero: "x \ 0 \ Rep_matrix (sparse_row_vector v) x y = 0" -proof - - assume x:"x \ 0" - have r:"nrows (sparse_row_vector v) <= Suc 0" by (rule nrows_spvec) - show ?thesis - apply (rule nrows) - apply (subgoal_tac "Suc 0 <= x") - apply (insert r) - apply (simp only:) - apply (insert x) - apply arith - done -qed - -lemma sparse_row_matrix_abs: - "sorted_spvec A \ sorted_spmat A \ sparse_row_matrix (abs_spmat A) = abs (sparse_row_matrix A)" - apply (induct A) - apply (simp_all add: sparse_row_vector_abs sparse_row_matrix_cons) - apply (frule_tac sorted_spvec_cons1, simp) - apply (simplesubst Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply auto - apply (case_tac "x=a") - apply (simp) - apply (simplesubst sorted_sparse_row_matrix_zero) - apply auto - apply (simplesubst Rep_sparse_row_vector_zero) - apply simp_all - done - -lemma sorted_spvec_minus_spmat: "sorted_spvec A \ sorted_spvec (minus_spmat A)" - apply (induct A) - apply (simp) - apply (frule sorted_spvec_cons1, simp) - apply (simp add: sorted_spvec.simps split:list.split_asm) - done - -lemma sorted_spvec_abs_spmat: "sorted_spvec A \ sorted_spvec (abs_spmat A)" - apply (induct A) - apply (simp) - apply (frule sorted_spvec_cons1, simp) - apply (simp add: sorted_spvec.simps split:list.split_asm) - done - -lemma sorted_spmat_minus_spmat: "sorted_spmat A \ sorted_spmat (minus_spmat A)" - apply (induct A) - apply (simp_all add: sorted_spvec_minus_spvec) - done - -lemma sorted_spmat_abs_spmat: "sorted_spmat A \ sorted_spmat (abs_spmat A)" - apply (induct A) - apply (simp_all add: sorted_spvec_abs_spvec) - done - -definition diff_spmat :: "('a::lattice_ring) spmat \ 'a spmat \ 'a spmat" - where "diff_spmat A B = add_spmat A (minus_spmat B)" - -lemma sorted_spmat_diff_spmat: "sorted_spmat A \ sorted_spmat B \ sorted_spmat (diff_spmat A B)" - by (simp add: diff_spmat_def sorted_spmat_minus_spmat sorted_spmat_add_spmat) - -lemma sorted_spvec_diff_spmat: "sorted_spvec A \ sorted_spvec B \ sorted_spvec (diff_spmat A B)" - by (simp add: diff_spmat_def sorted_spvec_minus_spmat sorted_spvec_add_spmat) - -lemma sparse_row_diff_spmat: "sparse_row_matrix (diff_spmat A B ) = (sparse_row_matrix A) - (sparse_row_matrix B)" - by (simp add: diff_spmat_def sparse_row_add_spmat sparse_row_matrix_minus) - -definition sorted_sparse_matrix :: "'a spmat \ bool" - where "sorted_sparse_matrix A \ sorted_spvec A & sorted_spmat A" - -lemma sorted_sparse_matrix_imp_spvec: "sorted_sparse_matrix A \ sorted_spvec A" - by (simp add: sorted_sparse_matrix_def) - -lemma sorted_sparse_matrix_imp_spmat: "sorted_sparse_matrix A \ sorted_spmat A" - by (simp add: sorted_sparse_matrix_def) - -lemmas sorted_sp_simps = - sorted_spvec.simps - sorted_spmat.simps - sorted_sparse_matrix_def - -lemma bool1: "(\ True) = False" by blast -lemma bool2: "(\ False) = True" by blast -lemma bool3: "((P\bool) \ True) = P" by blast -lemma bool4: "(True \ (P\bool)) = P" by blast -lemma bool5: "((P\bool) \ False) = False" by blast -lemma bool6: "(False \ (P\bool)) = False" by blast -lemma bool7: "((P\bool) \ True) = True" by blast -lemma bool8: "(True \ (P\bool)) = True" by blast -lemma bool9: "((P\bool) \ False) = P" by blast -lemma bool10: "(False \ (P\bool)) = P" by blast -lemmas boolarith = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10 - -lemma if_case_eq: "(if b then x else y) = (case b of True => x | False => y)" by simp - -primrec pprt_spvec :: "('a::{lattice_ab_group_add}) spvec \ 'a spvec" -where - "pprt_spvec [] = []" -| "pprt_spvec (a#as) = (fst a, pprt (snd a)) # (pprt_spvec as)" - -primrec nprt_spvec :: "('a::{lattice_ab_group_add}) spvec \ 'a spvec" -where - "nprt_spvec [] = []" -| "nprt_spvec (a#as) = (fst a, nprt (snd a)) # (nprt_spvec as)" - -primrec pprt_spmat :: "('a::{lattice_ab_group_add}) spmat \ 'a spmat" -where - "pprt_spmat [] = []" -| "pprt_spmat (a#as) = (fst a, pprt_spvec (snd a))#(pprt_spmat as)" - -primrec nprt_spmat :: "('a::{lattice_ab_group_add}) spmat \ 'a spmat" -where - "nprt_spmat [] = []" -| "nprt_spmat (a#as) = (fst a, nprt_spvec (snd a))#(nprt_spmat as)" - - -lemma pprt_add: "disj_matrices A (B::(_::lattice_ring) matrix) \ pprt (A+B) = pprt A + pprt B" - apply (simp add: pprt_def sup_matrix_def) - apply (simp add: Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply simp - apply (case_tac "Rep_matrix A x xa \ 0") - apply (simp_all add: disj_matrices_contr1) - done - -lemma nprt_add: "disj_matrices A (B::(_::lattice_ring) matrix) \ nprt (A+B) = nprt A + nprt B" - apply (simp add: nprt_def inf_matrix_def) - apply (simp add: Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply simp - apply (case_tac "Rep_matrix A x xa \ 0") - apply (simp_all add: disj_matrices_contr1) - done - -lemma pprt_singleton[simp]: "pprt (singleton_matrix j i (x::_::lattice_ring)) = singleton_matrix j i (pprt x)" - apply (simp add: pprt_def sup_matrix_def) - apply (simp add: Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply simp - done - -lemma nprt_singleton[simp]: "nprt (singleton_matrix j i (x::_::lattice_ring)) = singleton_matrix j i (nprt x)" - apply (simp add: nprt_def inf_matrix_def) - apply (simp add: Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply simp - done - -lemma less_imp_le: "a < b \ a <= (b::_::order)" by (simp add: less_def) - -lemma sparse_row_vector_pprt: "sorted_spvec (v :: 'a::lattice_ring spvec) \ sparse_row_vector (pprt_spvec v) = pprt (sparse_row_vector v)" - apply (induct v) - apply (simp_all) - apply (frule sorted_spvec_cons1, auto) - apply (subst pprt_add) - apply (subst disj_matrices_commute) - apply (rule disj_sparse_row_singleton) - apply auto - done - -lemma sparse_row_vector_nprt: "sorted_spvec (v :: 'a::lattice_ring spvec) \ sparse_row_vector (nprt_spvec v) = nprt (sparse_row_vector v)" - apply (induct v) - apply (simp_all) - apply (frule sorted_spvec_cons1, auto) - apply (subst nprt_add) - apply (subst disj_matrices_commute) - apply (rule disj_sparse_row_singleton) - apply auto - done - - -lemma pprt_move_matrix: "pprt (move_matrix (A::('a::lattice_ring) matrix) j i) = move_matrix (pprt A) j i" - apply (simp add: pprt_def) - apply (simp add: sup_matrix_def) - apply (simp add: Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply (simp) - done - -lemma nprt_move_matrix: "nprt (move_matrix (A::('a::lattice_ring) matrix) j i) = move_matrix (nprt A) j i" - apply (simp add: nprt_def) - apply (simp add: inf_matrix_def) - apply (simp add: Rep_matrix_inject[symmetric]) - apply (rule ext)+ - apply (simp) - done - -lemma sparse_row_matrix_pprt: "sorted_spvec (m :: 'a::lattice_ring spmat) \ sorted_spmat m \ sparse_row_matrix (pprt_spmat m) = pprt (sparse_row_matrix m)" - apply (induct m) - apply simp - apply simp - apply (frule sorted_spvec_cons1) - apply (simp add: sparse_row_matrix_cons sparse_row_vector_pprt) - apply (subst pprt_add) - apply (subst disj_matrices_commute) - apply (rule disj_move_sparse_vec_mat) - apply auto - apply (simp add: sorted_spvec.simps) - apply (simp split: list.split) - apply auto - apply (simp add: pprt_move_matrix) - done - -lemma sparse_row_matrix_nprt: "sorted_spvec (m :: 'a::lattice_ring spmat) \ sorted_spmat m \ sparse_row_matrix (nprt_spmat m) = nprt (sparse_row_matrix m)" - apply (induct m) - apply simp - apply simp - apply (frule sorted_spvec_cons1) - apply (simp add: sparse_row_matrix_cons sparse_row_vector_nprt) - apply (subst nprt_add) - apply (subst disj_matrices_commute) - apply (rule disj_move_sparse_vec_mat) - apply auto - apply (simp add: sorted_spvec.simps) - apply (simp split: list.split) - apply auto - apply (simp add: nprt_move_matrix) - done - -lemma sorted_pprt_spvec: "sorted_spvec v \ sorted_spvec (pprt_spvec v)" - apply (induct v) - apply (simp) - apply (frule sorted_spvec_cons1) - apply simp - apply (simp add: sorted_spvec.simps split:list.split_asm) - done - -lemma sorted_nprt_spvec: "sorted_spvec v \ sorted_spvec (nprt_spvec v)" - apply (induct v) - apply (simp) - apply (frule sorted_spvec_cons1) - apply simp - apply (simp add: sorted_spvec.simps split:list.split_asm) - done - -lemma sorted_spvec_pprt_spmat: "sorted_spvec m \ sorted_spvec (pprt_spmat m)" - apply (induct m) - apply (simp) - apply (frule sorted_spvec_cons1) - apply simp - apply (simp add: sorted_spvec.simps split:list.split_asm) - done - -lemma sorted_spvec_nprt_spmat: "sorted_spvec m \ sorted_spvec (nprt_spmat m)" - apply (induct m) - apply (simp) - apply (frule sorted_spvec_cons1) - apply simp - apply (simp add: sorted_spvec.simps split:list.split_asm) - done - -lemma sorted_spmat_pprt_spmat: "sorted_spmat m \ sorted_spmat (pprt_spmat m)" - apply (induct m) - apply (simp_all add: sorted_pprt_spvec) - done - -lemma sorted_spmat_nprt_spmat: "sorted_spmat m \ sorted_spmat (nprt_spmat m)" - apply (induct m) - apply (simp_all add: sorted_nprt_spvec) - done - -definition mult_est_spmat :: "('a::lattice_ring) spmat \ 'a spmat \ 'a spmat \ 'a spmat \ 'a spmat" where - "mult_est_spmat r1 r2 s1 s2 = - add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2)) (add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2)) - (add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1)) (mult_spmat (nprt_spmat s1) (nprt_spmat r1))))" - -lemmas sparse_row_matrix_op_simps = - sorted_sparse_matrix_imp_spmat sorted_sparse_matrix_imp_spvec - sparse_row_add_spmat sorted_spvec_add_spmat sorted_spmat_add_spmat - sparse_row_diff_spmat sorted_spvec_diff_spmat sorted_spmat_diff_spmat - sparse_row_matrix_minus sorted_spvec_minus_spmat sorted_spmat_minus_spmat - sparse_row_mult_spmat sorted_spvec_mult_spmat sorted_spmat_mult_spmat - sparse_row_matrix_abs sorted_spvec_abs_spmat sorted_spmat_abs_spmat - le_spmat_iff_sparse_row_le - sparse_row_matrix_pprt sorted_spvec_pprt_spmat sorted_spmat_pprt_spmat - sparse_row_matrix_nprt sorted_spvec_nprt_spmat sorted_spmat_nprt_spmat - -lemma zero_eq_Numeral0: "(0::_::number_ring) = Numeral0" by simp - -lemmas sparse_row_matrix_arith_simps[simplified zero_eq_Numeral0] = - mult_spmat.simps mult_spvec_spmat.simps - addmult_spvec.simps - smult_spvec_empty smult_spvec_cons - add_spmat.simps add_spvec.simps - minus_spmat.simps minus_spvec.simps - abs_spmat.simps abs_spvec.simps - diff_spmat_def - le_spmat.simps le_spvec.simps - pprt_spmat.simps pprt_spvec.simps - nprt_spmat.simps nprt_spvec.simps - mult_est_spmat_def - - -(*lemma spm_linprog_dual_estimate_1: - assumes - "sorted_sparse_matrix A1" - "sorted_sparse_matrix A2" - "sorted_sparse_matrix c1" - "sorted_sparse_matrix c2" - "sorted_sparse_matrix y" - "sorted_spvec b" - "sorted_spvec r" - "le_spmat ([], y)" - "A * x \ sparse_row_matrix (b::('a::lattice_ring) spmat)" - "sparse_row_matrix A1 <= A" - "A <= sparse_row_matrix A2" - "sparse_row_matrix c1 <= c" - "c <= sparse_row_matrix c2" - "abs x \ sparse_row_matrix r" - shows - "c * x \ sparse_row_matrix (add_spmat (mult_spmat y b, mult_spmat (add_spmat (add_spmat (mult_spmat y (diff_spmat A2 A1), - abs_spmat (diff_spmat (mult_spmat y A1) c1)), diff_spmat c2 c1)) r))" - by (insert prems, simp add: sparse_row_matrix_op_simps linprog_dual_estimate_1[where A=A]) -*) - -end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/document/root.tex --- a/src/HOL/Matrix/document/root.tex Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,25 +0,0 @@ -\documentclass[11pt,a4paper]{article} -\usepackage{isabelle,isabellesym} - -% this should be the last package used -\usepackage{pdfsetup} - -% urls in roman style, theory text in math-similar italics -\urlstyle{rm} -\isabellestyle{it} - -\newcommand{\ganz}{\mathsf{Z}\mkern-7.5mu\mathsf{Z}} - -\begin{document} - -\title{Matrix} -\author{Steven Obua} -\maketitle - -%\tableofcontents - -\parindent 0pt\parskip 0.5ex - -\input{session} - -\end{document} diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/fspmlp.ML --- a/src/HOL/Matrix/fspmlp.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,313 +0,0 @@ -(* Title: HOL/Matrix/fspmlp.ML - Author: Steven Obua -*) - -signature FSPMLP = -sig - type linprog - type vector = FloatSparseMatrixBuilder.vector - type matrix = FloatSparseMatrixBuilder.matrix - - val y : linprog -> term - val A : linprog -> term * term - val b : linprog -> term - val c : linprog -> term * term - val r12 : linprog -> term * term - - exception Load of string - - val load : string -> int -> bool -> linprog -end - -structure Fspmlp : FSPMLP = -struct - -type vector = FloatSparseMatrixBuilder.vector -type matrix = FloatSparseMatrixBuilder.matrix - -type linprog = term * (term * term) * term * (term * term) * (term * term) - -fun y (c1, _, _, _, _) = c1 -fun A (_, c2, _, _, _) = c2 -fun b (_, _, c3, _, _) = c3 -fun c (_, _, _, c4, _) = c4 -fun r12 (_, _, _, _, c6) = c6 - -structure CplexFloatSparseMatrixConverter = -MAKE_CPLEX_MATRIX_CONVERTER(structure cplex = Cplex and matrix_builder = FloatSparseMatrixBuilder); - -datatype bound_type = LOWER | UPPER - -fun intbound_ord ((i1: int, b1),(i2,b2)) = - if i1 < i2 then LESS - else if i1 = i2 then - (if b1 = b2 then EQUAL else if b1=LOWER then LESS else GREATER) - else GREATER - -structure Inttab = Table(type key = int val ord = (rev_order o int_ord)); - -structure VarGraph = Table(type key = int*bound_type val ord = intbound_ord); -(* key -> (float option) * (int -> (float * (((float * float) * key) list)))) *) -(* dest_key -> (sure_bound * (row_index -> (row_bound * (((coeff_lower * coeff_upper) * src_key) list)))) *) - -exception Internal of string; - -fun add_row_bound g dest_key row_index row_bound = - let - val x = - case VarGraph.lookup g dest_key of - NONE => (NONE, Inttab.update (row_index, (row_bound, [])) Inttab.empty) - | SOME (sure_bound, f) => - (sure_bound, - case Inttab.lookup f row_index of - NONE => Inttab.update (row_index, (row_bound, [])) f - | SOME _ => raise (Internal "add_row_bound")) - in - VarGraph.update (dest_key, x) g - end - -fun update_sure_bound g (key as (_, btype)) bound = - let - val x = - case VarGraph.lookup g key of - NONE => (SOME bound, Inttab.empty) - | SOME (NONE, f) => (SOME bound, f) - | SOME (SOME old_bound, f) => - (SOME ((case btype of - UPPER => Float.min - | LOWER => Float.max) - old_bound bound), f) - in - VarGraph.update (key, x) g - end - -fun get_sure_bound g key = - case VarGraph.lookup g key of - NONE => NONE - | SOME (sure_bound, _) => sure_bound - -(*fun get_row_bound g key row_index = - case VarGraph.lookup g key of - NONE => NONE - | SOME (sure_bound, f) => - (case Inttab.lookup f row_index of - NONE => NONE - | SOME (row_bound, _) => (sure_bound, row_bound))*) - -fun add_edge g src_key dest_key row_index coeff = - case VarGraph.lookup g dest_key of - NONE => raise (Internal "add_edge: dest_key not found") - | SOME (sure_bound, f) => - (case Inttab.lookup f row_index of - NONE => raise (Internal "add_edge: row_index not found") - | SOME (row_bound, sources) => - VarGraph.update (dest_key, (sure_bound, Inttab.update (row_index, (row_bound, (coeff, src_key) :: sources)) f)) g) - -fun split_graph g = - let - fun split (key, (sure_bound, _)) (r1, r2) = case sure_bound - of NONE => (r1, r2) - | SOME bound => (case key - of (u, UPPER) => (r1, Inttab.update (u, bound) r2) - | (u, LOWER) => (Inttab.update (u, bound) r1, r2)) - in VarGraph.fold split g (Inttab.empty, Inttab.empty) end - -(* If safe is true, termination is guaranteed, but the sure bounds may be not optimal (relative to the algorithm). - If safe is false, termination is not guaranteed, but on termination the sure bounds are optimal (relative to the algorithm) *) -fun propagate_sure_bounds safe names g = - let - (* returns NONE if no new sure bound could be calculated, otherwise the new sure bound is returned *) - fun calc_sure_bound_from_sources g (key as (_, btype)) = - let - fun mult_upper x (lower, upper) = - if Float.sign x = LESS then - Float.mult x lower - else - Float.mult x upper - - fun mult_lower x (lower, upper) = - if Float.sign x = LESS then - Float.mult x upper - else - Float.mult x lower - - val mult_btype = case btype of UPPER => mult_upper | LOWER => mult_lower - - fun calc_sure_bound (_, (row_bound, sources)) sure_bound = - let - fun add_src_bound (coeff, src_key) sum = - case sum of - NONE => NONE - | SOME x => - (case get_sure_bound g src_key of - NONE => NONE - | SOME src_sure_bound => SOME (Float.add x (mult_btype src_sure_bound coeff))) - in - case fold add_src_bound sources (SOME row_bound) of - NONE => sure_bound - | new_sure_bound as (SOME new_bound) => - (case sure_bound of - NONE => new_sure_bound - | SOME old_bound => - SOME (case btype of - UPPER => Float.min old_bound new_bound - | LOWER => Float.max old_bound new_bound)) - end - in - case VarGraph.lookup g key of - NONE => NONE - | SOME (sure_bound, f) => - let - val x = Inttab.fold calc_sure_bound f sure_bound - in - if x = sure_bound then NONE else x - end - end - - fun propagate (key, _) (g, b) = - case calc_sure_bound_from_sources g key of - NONE => (g,b) - | SOME bound => (update_sure_bound g key bound, - if safe then - case get_sure_bound g key of - NONE => true - | _ => b - else - true) - - val (g, b) = VarGraph.fold propagate g (g, false) - in - if b then propagate_sure_bounds safe names g else g - end - -exception Load of string; - -val empty_spvec = @{term "Nil :: real spvec"}; -fun cons_spvec x xs = @{term "Cons :: nat * real => real spvec => real spvec"} $ x $ xs; -val empty_spmat = @{term "Nil :: real spmat"}; -fun cons_spmat x xs = @{term "Cons :: nat * real spvec => real spmat => real spmat"} $ x $ xs; - -fun calcr safe_propagation xlen names prec A b = - let - fun test_1 (lower, upper) = - if lower = upper then - (if Float.eq (lower, (~1, 0)) then ~1 - else if Float.eq (lower, (1, 0)) then 1 - else 0) - else 0 - - fun calcr (row_index, a) g = - let - val b = FloatSparseMatrixBuilder.v_elem_at b row_index - val (_, b2) = FloatArith.approx_decstr_by_bin prec (case b of NONE => "0" | SOME b => b) - val approx_a = FloatSparseMatrixBuilder.v_fold (fn (i, s) => fn l => - (i, FloatArith.approx_decstr_by_bin prec s)::l) a [] - - fun fold_dest_nodes (dest_index, dest_value) g = - let - val dest_test = test_1 dest_value - in - if dest_test = 0 then - g - else let - val (dest_key as (_, dest_btype), row_bound) = - if dest_test = ~1 then - ((dest_index, LOWER), Float.neg b2) - else - ((dest_index, UPPER), b2) - - fun fold_src_nodes (src_index, src_value as (src_lower, src_upper)) g = - if src_index = dest_index then g - else - let - val coeff = case dest_btype of - UPPER => (Float.neg src_upper, Float.neg src_lower) - | LOWER => src_value - in - if Float.sign src_lower = LESS then - add_edge g (src_index, UPPER) dest_key row_index coeff - else - add_edge g (src_index, LOWER) dest_key row_index coeff - end - in - fold fold_src_nodes approx_a (add_row_bound g dest_key row_index row_bound) - end - end - in - case approx_a of - [] => g - | [(u, a)] => - let - val atest = test_1 a - in - if atest = ~1 then - update_sure_bound g (u, LOWER) (Float.neg b2) - else if atest = 1 then - update_sure_bound g (u, UPPER) b2 - else - g - end - | _ => fold fold_dest_nodes approx_a g - end - - val g = FloatSparseMatrixBuilder.m_fold calcr A VarGraph.empty - - val g = propagate_sure_bounds safe_propagation names g - - val (r1, r2) = split_graph g - - fun add_row_entry m index f vname value = - let - val v = (case value of - SOME value => FloatSparseMatrixBuilder.mk_spvec_entry 0 value - | NONE => FloatSparseMatrixBuilder.mk_spvec_entry' 0 (f $ (Var ((vname,0), HOLogic.realT)))) - val vec = cons_spvec v empty_spvec - in - cons_spmat (FloatSparseMatrixBuilder.mk_spmat_entry index vec) m - end - - fun abs_estimate i r1 r2 = - if i = 0 then - let val e = empty_spmat in (e, e) end - else - let - val index = xlen-i - val (r12_1, r12_2) = abs_estimate (i-1) r1 r2 - val b1 = Inttab.lookup r1 index - val b2 = Inttab.lookup r2 index - in - (add_row_entry r12_1 index @{term "lbound :: real => real"} ((names index)^"l") b1, - add_row_entry r12_2 index @{term "ubound :: real => real"} ((names index)^"u") b2) - end - - val (r1, r2) = abs_estimate xlen r1 r2 - - in - (r1, r2) - end - -fun load filename prec safe_propagation = - let - val prog = Cplex.load_cplexFile filename - val prog = Cplex.elim_nonfree_bounds prog - val prog = Cplex.relax_strict_ineqs prog - val (maximize, c, A, b, (xlen, names, _)) = CplexFloatSparseMatrixConverter.convert_prog prog - val (r1, r2) = calcr safe_propagation xlen names prec A b - val _ = if maximize then () else raise Load "sorry, cannot handle minimization problems" - val (dualprog, indexof) = FloatSparseMatrixBuilder.dual_cplexProg c A b - val results = Cplex.solve dualprog - val (_, v) = CplexFloatSparseMatrixConverter.convert_results results indexof - (*val A = FloatSparseMatrixBuilder.cut_matrix v NONE A*) - fun id x = x - val v = FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 v - val b = FloatSparseMatrixBuilder.transpose_matrix (FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 b) - val c = FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 c - val (y1, _) = FloatSparseMatrixBuilder.approx_matrix prec Float.positive_part v - val A = FloatSparseMatrixBuilder.approx_matrix prec id A - val (_,b2) = FloatSparseMatrixBuilder.approx_matrix prec id b - val c = FloatSparseMatrixBuilder.approx_matrix prec id c - in - (y1, A, b2, c, (r1, r2)) - end handle CplexFloatSparseMatrixConverter.Converter s => (raise (Load ("Converter: "^s))) - -end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix/matrixlp.ML --- a/src/HOL/Matrix/matrixlp.ML Sat Mar 17 12:26:19 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,59 +0,0 @@ -(* Title: HOL/Matrix/matrixlp.ML - Author: Steven Obua -*) - -signature MATRIX_LP = -sig - val matrix_compute : cterm -> thm - val matrix_simplify : thm -> thm - val prove_bound : string -> int -> thm - val float2real : string * string -> Real.real -end - -structure MatrixLP : MATRIX_LP = -struct - -val compute_thms = ComputeHOL.prep_thms @{thms "ComputeHOL.compute_list_case" "ComputeHOL.compute_let" - "ComputeHOL.compute_if" "ComputeFloat.arith" "SparseMatrix.sparse_row_matrix_arith_simps" - "ComputeHOL.compute_bool" "ComputeHOL.compute_pair" - "SparseMatrix.sorted_sp_simps" - "ComputeNumeral.natnorm"}; (*"ComputeNumeral.number_norm"*) - -val spm_mult_le_dual_prts_no_let_real = @{thm "spm_mult_le_dual_prts_no_let" [where ?'a = real]} - -fun lp_dual_estimate_prt lptfile prec = - let - val cert = cterm_of @{theory} - fun var s x = (cert (Var ((s, 0), FloatSparseMatrixBuilder.spmatT)), x) - val l = Fspmlp.load lptfile prec false - val (y, (A1, A2), (c1, c2), b, (r1, r2)) = - let - open Fspmlp - in - (y l |> cert, A l |> pairself cert, c l |> pairself cert, b l |> cert, r12 l |> pairself cert) - end - in - Thm.instantiate ([], - [var "A1" A1, var "A2" A2, var "y" y, var "c1" c1, var "c2" c2, var "r1" r1, var "r2" r2, var "b" b]) - spm_mult_le_dual_prts_no_let_real - end - -val computer = PCompute.make Compute.SML @{theory} compute_thms [] - -fun matrix_compute c = hd (PCompute.rewrite computer [c]) - -fun matrix_simplify th = - let - val simp_th = matrix_compute (cprop_of th) - val th = Thm.strip_shyps (Thm.equal_elim simp_th th) - fun removeTrue th = removeTrue (Thm.implies_elim th TrueI) handle THM _ => th - in - removeTrue th - end - -val prove_bound = matrix_simplify oo lp_dual_estimate_prt; - -val realFromStr = the o Real.fromString; -fun float2real (x, y) = realFromStr x * Math.pow (2.0, realFromStr y); - -end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/ComputeFloat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/ComputeFloat.thy Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,309 @@ +(* Title: HOL/Matrix/ComputeFloat.thy + Author: Steven Obua +*) + +header {* Floating Point Representation of the Reals *} + +theory ComputeFloat +imports Complex_Main "~~/src/HOL/Library/Lattice_Algebras" +uses "~~/src/Tools/float.ML" ("~~/src/HOL/Tools/float_arith.ML") +begin + +definition int_of_real :: "real \ int" + where "int_of_real x = (SOME y. real y = x)" + +definition real_is_int :: "real \ bool" + where "real_is_int x = (EX (u::int). x = real u)" + +lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))" + by (auto simp add: real_is_int_def int_of_real_def) + +lemma real_is_int_real[simp]: "real_is_int (real (x::int))" +by (auto simp add: real_is_int_def int_of_real_def) + +lemma int_of_real_real[simp]: "int_of_real (real x) = x" +by (simp add: int_of_real_def) + +lemma real_int_of_real[simp]: "real_is_int x \ real (int_of_real x) = x" +by (auto simp add: int_of_real_def real_is_int_def) + +lemma real_is_int_add_int_of_real: "real_is_int a \ real_is_int b \ (int_of_real (a+b)) = (int_of_real a) + (int_of_real b)" +by (auto simp add: int_of_real_def real_is_int_def) + +lemma real_is_int_add[simp]: "real_is_int a \ real_is_int b \ real_is_int (a+b)" +apply (subst real_is_int_def2) +apply (simp add: real_is_int_add_int_of_real real_int_of_real) +done + +lemma int_of_real_sub: "real_is_int a \ real_is_int b \ (int_of_real (a-b)) = (int_of_real a) - (int_of_real b)" +by (auto simp add: int_of_real_def real_is_int_def) + +lemma real_is_int_sub[simp]: "real_is_int a \ real_is_int b \ real_is_int (a-b)" +apply (subst real_is_int_def2) +apply (simp add: int_of_real_sub real_int_of_real) +done + +lemma real_is_int_rep: "real_is_int x \ ?! (a::int). real a = x" +by (auto simp add: real_is_int_def) + +lemma int_of_real_mult: + assumes "real_is_int a" "real_is_int b" + shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)" + using assms + by (auto simp add: real_is_int_def real_of_int_mult[symmetric] + simp del: real_of_int_mult) + +lemma real_is_int_mult[simp]: "real_is_int a \ real_is_int b \ real_is_int (a*b)" +apply (subst real_is_int_def2) +apply (simp add: int_of_real_mult) +done + +lemma real_is_int_0[simp]: "real_is_int (0::real)" +by (simp add: real_is_int_def int_of_real_def) + +lemma real_is_int_1[simp]: "real_is_int (1::real)" +proof - + have "real_is_int (1::real) = real_is_int(real (1::int))" by auto + also have "\ = True" by (simp only: real_is_int_real) + ultimately show ?thesis by auto +qed + +lemma real_is_int_n1: "real_is_int (-1::real)" +proof - + have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto + also have "\ = True" by (simp only: real_is_int_real) + ultimately show ?thesis by auto +qed + +lemma real_is_int_number_of[simp]: "real_is_int ((number_of \ int \ real) x)" + by (auto simp: real_is_int_def intro!: exI[of _ "number_of x"]) + +lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)" +by (simp add: int_of_real_def) + +lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)" +proof - + have 1: "(1::real) = real (1::int)" by auto + show ?thesis by (simp only: 1 int_of_real_real) +qed + +lemma int_of_real_number_of[simp]: "int_of_real (number_of b) = number_of b" + unfolding int_of_real_def + by (intro some_equality) + (auto simp add: real_of_int_inject[symmetric] simp del: real_of_int_inject) + +lemma int_div_zdiv: "int (a div b) = (int a) div (int b)" +by (rule zdiv_int) + +lemma int_mod_zmod: "int (a mod b) = (int a) mod (int b)" +by (rule zmod_int) + +lemma abs_div_2_less: "a \ 0 \ a \ -1 \ abs((a::int) div 2) < abs a" +by arith + +lemma norm_0_1: "(0::_::number_ring) = Numeral0 & (1::_::number_ring) = Numeral1" + by auto + +lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)" + by simp + +lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)" + by simp + +lemma mult_left_one: "1 * a = (a::'a::semiring_1)" + by simp + +lemma mult_right_one: "a * 1 = (a::'a::semiring_1)" + by simp + +lemma int_pow_0: "(a::int)^(Numeral0) = 1" + by simp + +lemma int_pow_1: "(a::int)^(Numeral1) = a" + by simp + +lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0" + by simp + +lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1" + by simp + +lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0" + by simp + +lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1" + by simp + +lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1" + by simp + +lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1" +proof - + have 1:"((-1)::nat) = 0" + by simp + show ?thesis by (simp add: 1) +qed + +lemma fst_cong: "a=a' \ fst (a,b) = fst (a',b)" + by simp + +lemma snd_cong: "b=b' \ snd (a,b) = snd (a,b')" + by simp + +lemma lift_bool: "x \ x=True" + by simp + +lemma nlift_bool: "~x \ x=False" + by simp + +lemma not_false_eq_true: "(~ False) = True" by simp + +lemma not_true_eq_false: "(~ True) = False" by simp + +lemmas binarith = + normalize_bin_simps + pred_bin_simps succ_bin_simps + add_bin_simps minus_bin_simps mult_bin_simps + +lemma int_eq_number_of_eq: + "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)" + by (rule eq_number_of_eq) + +lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)" + by (simp only: iszero_number_of_Pls) + +lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))" + by simp + +lemma int_iszero_number_of_Bit0: "iszero ((number_of (Int.Bit0 w))::int) = iszero ((number_of w)::int)" + by simp + +lemma int_iszero_number_of_Bit1: "\ iszero ((number_of (Int.Bit1 w))::int)" + by simp + +lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)" + unfolding neg_def number_of_is_id by simp + +lemma int_not_neg_number_of_Pls: "\ (neg (Numeral0::int))" + by simp + +lemma int_neg_number_of_Min: "neg (-1::int)" + by simp + +lemma int_neg_number_of_Bit0: "neg ((number_of (Int.Bit0 w))::int) = neg ((number_of w)::int)" + by simp + +lemma int_neg_number_of_Bit1: "neg ((number_of (Int.Bit1 w))::int) = neg ((number_of w)::int)" + by simp + +lemma int_le_number_of_eq: "(((number_of x)::int) \ number_of y) = (\ neg ((number_of (y + (uminus x)))::int))" + unfolding neg_def number_of_is_id by (simp add: not_less) + +lemmas intarithrel = + int_eq_number_of_eq + lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_Bit0 + lift_bool[OF int_iszero_number_of_Bit1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min] + int_neg_number_of_Bit0 int_neg_number_of_Bit1 int_le_number_of_eq + +lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)" + by simp + +lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (v + (uminus w))" + by simp + +lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (v * w)" + by simp + +lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)" + by simp + +lemmas intarith = int_number_of_add_sym int_number_of_minus_sym int_number_of_diff_sym int_number_of_mult_sym + +lemmas natarith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of + +lemmas powerarith = nat_number_of zpower_number_of_even + zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring] + zpower_Pls zpower_Min + +definition float :: "(int \ int) \ real" where + "float = (\(a, b). real a * 2 powr real b)" + +lemma float_add_l0: "float (0, e) + x = x" + by (simp add: float_def) + +lemma float_add_r0: "x + float (0, e) = x" + by (simp add: float_def) + +lemma float_add: + "float (a1, e1) + float (a2, e2) = + (if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1) else float (a1*2^(nat (e1-e2))+a2, e2))" + by (simp add: float_def algebra_simps powr_realpow[symmetric] powr_divide2[symmetric]) + +lemma float_mult_l0: "float (0, e) * x = float (0, 0)" + by (simp add: float_def) + +lemma float_mult_r0: "x * float (0, e) = float (0, 0)" + by (simp add: float_def) + +lemma float_mult: + "float (a1, e1) * float (a2, e2) = (float (a1 * a2, e1 + e2))" + by (simp add: float_def powr_add) + +lemma float_minus: + "- (float (a,b)) = float (-a, b)" + by (simp add: float_def) + +lemma zero_le_float: + "(0 <= float (a,b)) = (0 <= a)" + using powr_gt_zero[of 2 "real b", arith] + by (simp add: float_def zero_le_mult_iff) + +lemma float_le_zero: + "(float (a,b) <= 0) = (a <= 0)" + using powr_gt_zero[of 2 "real b", arith] + by (simp add: float_def mult_le_0_iff) + +lemma float_abs: + "abs (float (a,b)) = (if 0 <= a then (float (a,b)) else (float (-a,b)))" + using powr_gt_zero[of 2 "real b", arith] + by (simp add: float_def abs_if mult_less_0_iff) + +lemma float_zero: + "float (0, b) = 0" + by (simp add: float_def) + +lemma float_pprt: + "pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0, b)))" + by (auto simp add: zero_le_float float_le_zero float_zero) + +lemma float_nprt: + "nprt (float (a, b)) = (if 0 <= a then (float (0,b)) else (float (a, b)))" + by (auto simp add: zero_le_float float_le_zero float_zero) + +definition lbound :: "real \ real" + where "lbound x = min 0 x" + +definition ubound :: "real \ real" + where "ubound x = max 0 x" + +lemma lbound: "lbound x \ x" + by (simp add: lbound_def) + +lemma ubound: "x \ ubound x" + by (simp add: ubound_def) + +lemma pprt_lbound: "pprt (lbound x) = float (0, 0)" + by (auto simp: float_def lbound_def) + +lemma nprt_ubound: "nprt (ubound x) = float (0, 0)" + by (auto simp: float_def ubound_def) + +lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0 + float_minus float_abs zero_le_float float_pprt float_nprt pprt_lbound nprt_ubound + +(* for use with the compute oracle *) +lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false + +use "~~/src/HOL/Tools/float_arith.ML" + +end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/ComputeHOL.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/ComputeHOL.thy Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,187 @@ +theory ComputeHOL +imports Complex_Main "Compute_Oracle/Compute_Oracle" +begin + +lemma Trueprop_eq_eq: "Trueprop X == (X == True)" by (simp add: atomize_eq) +lemma meta_eq_trivial: "x == y \ x == y" by simp +lemma meta_eq_imp_eq: "x == y \ x = y" by auto +lemma eq_trivial: "x = y \ x = y" by auto +lemma bool_to_true: "x :: bool \ x == True" by simp +lemma transmeta_1: "x = y \ y == z \ x = z" by simp +lemma transmeta_2: "x == y \ y = z \ x = z" by simp +lemma transmeta_3: "x == y \ y == z \ x = z" by simp + + +(**** compute_if ****) + +lemma If_True: "If True = (\ x y. x)" by ((rule ext)+,auto) +lemma If_False: "If False = (\ x y. y)" by ((rule ext)+, auto) + +lemmas compute_if = If_True If_False + +(**** compute_bool ****) + +lemma bool1: "(\ True) = False" by blast +lemma bool2: "(\ False) = True" by blast +lemma bool3: "(P \ True) = P" by blast +lemma bool4: "(True \ P) = P" by blast +lemma bool5: "(P \ False) = False" by blast +lemma bool6: "(False \ P) = False" by blast +lemma bool7: "(P \ True) = True" by blast +lemma bool8: "(True \ P) = True" by blast +lemma bool9: "(P \ False) = P" by blast +lemma bool10: "(False \ P) = P" by blast +lemma bool11: "(True \ P) = P" by blast +lemma bool12: "(P \ True) = True" by blast +lemma bool13: "(True \ P) = P" by blast +lemma bool14: "(P \ False) = (\ P)" by blast +lemma bool15: "(False \ P) = True" by blast +lemma bool16: "(False = False) = True" by blast +lemma bool17: "(True = True) = True" by blast +lemma bool18: "(False = True) = False" by blast +lemma bool19: "(True = False) = False" by blast + +lemmas compute_bool = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10 bool11 bool12 bool13 bool14 bool15 bool16 bool17 bool18 bool19 + + +(*** compute_pair ***) + +lemma compute_fst: "fst (x,y) = x" by simp +lemma compute_snd: "snd (x,y) = y" by simp +lemma compute_pair_eq: "((a, b) = (c, d)) = (a = c \ b = d)" by auto + +lemma prod_case_simp: "prod_case f (x,y) = f x y" by simp + +lemmas compute_pair = compute_fst compute_snd compute_pair_eq prod_case_simp + +(*** compute_option ***) + +lemma compute_the: "the (Some x) = x" by simp +lemma compute_None_Some_eq: "(None = Some x) = False" by auto +lemma compute_Some_None_eq: "(Some x = None) = False" by auto +lemma compute_None_None_eq: "(None = None) = True" by auto +lemma compute_Some_Some_eq: "(Some x = Some y) = (x = y)" by auto + +definition option_case_compute :: "'b option \ 'a \ ('b \ 'a) \ 'a" + where "option_case_compute opt a f = option_case a f opt" + +lemma option_case_compute: "option_case = (\ a f opt. option_case_compute opt a f)" + by (simp add: option_case_compute_def) + +lemma option_case_compute_None: "option_case_compute None = (\ a f. a)" + apply (rule ext)+ + apply (simp add: option_case_compute_def) + done + +lemma option_case_compute_Some: "option_case_compute (Some x) = (\ a f. f x)" + apply (rule ext)+ + apply (simp add: option_case_compute_def) + done + +lemmas compute_option_case = option_case_compute option_case_compute_None option_case_compute_Some + +lemmas compute_option = compute_the compute_None_Some_eq compute_Some_None_eq compute_None_None_eq compute_Some_Some_eq compute_option_case + +(**** compute_list_length ****) + +lemma length_cons:"length (x#xs) = 1 + (length xs)" + by simp + +lemma length_nil: "length [] = 0" + by simp + +lemmas compute_list_length = length_nil length_cons + +(*** compute_list_case ***) + +definition list_case_compute :: "'b list \ 'a \ ('b \ 'b list \ 'a) \ 'a" + where "list_case_compute l a f = list_case a f l" + +lemma list_case_compute: "list_case = (\ (a::'a) f (l::'b list). list_case_compute l a f)" + apply (rule ext)+ + apply (simp add: list_case_compute_def) + done + +lemma list_case_compute_empty: "list_case_compute ([]::'b list) = (\ (a::'a) f. a)" + apply (rule ext)+ + apply (simp add: list_case_compute_def) + done + +lemma list_case_compute_cons: "list_case_compute (u#v) = (\ (a::'a) f. (f (u::'b) v))" + apply (rule ext)+ + apply (simp add: list_case_compute_def) + done + +lemmas compute_list_case = list_case_compute list_case_compute_empty list_case_compute_cons + +(*** compute_list_nth ***) +(* Of course, you will need computation with nats for this to work \ *) + +lemma compute_list_nth: "((x#xs) ! n) = (if n = 0 then x else (xs ! (n - 1)))" + by (cases n, auto) + +(*** compute_list ***) + +lemmas compute_list = compute_list_case compute_list_length compute_list_nth + +(*** compute_let ***) + +lemmas compute_let = Let_def + +(***********************) +(* Everything together *) +(***********************) + +lemmas compute_hol = compute_if compute_bool compute_pair compute_option compute_list compute_let + +ML {* +signature ComputeHOL = +sig + val prep_thms : thm list -> thm list + val to_meta_eq : thm -> thm + val to_hol_eq : thm -> thm + val symmetric : thm -> thm + val trans : thm -> thm -> thm +end + +structure ComputeHOL : ComputeHOL = +struct + +local +fun lhs_of eq = fst (Thm.dest_equals (cprop_of eq)); +in +fun rewrite_conv [] ct = raise CTERM ("rewrite_conv", [ct]) + | rewrite_conv (eq :: eqs) ct = + Thm.instantiate (Thm.match (lhs_of eq, ct)) eq + handle Pattern.MATCH => rewrite_conv eqs ct; +end + +val convert_conditions = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.try_conv (rewrite_conv [@{thm "Trueprop_eq_eq"}]))) + +val eq_th = @{thm "HOL.eq_reflection"} +val meta_eq_trivial = @{thm "ComputeHOL.meta_eq_trivial"} +val bool_to_true = @{thm "ComputeHOL.bool_to_true"} + +fun to_meta_eq th = eq_th OF [th] handle THM _ => meta_eq_trivial OF [th] handle THM _ => bool_to_true OF [th] + +fun to_hol_eq th = @{thm "meta_eq_imp_eq"} OF [th] handle THM _ => @{thm "eq_trivial"} OF [th] + +fun prep_thms ths = map (convert_conditions o to_meta_eq) ths + +fun symmetric th = @{thm "HOL.sym"} OF [th] handle THM _ => @{thm "Pure.symmetric"} OF [th] + +local + val trans_HOL = @{thm "HOL.trans"} + val trans_HOL_1 = @{thm "ComputeHOL.transmeta_1"} + val trans_HOL_2 = @{thm "ComputeHOL.transmeta_2"} + val trans_HOL_3 = @{thm "ComputeHOL.transmeta_3"} + fun tr [] th1 th2 = trans_HOL OF [th1, th2] + | tr (t::ts) th1 th2 = (t OF [th1, th2] handle THM _ => tr ts th1 th2) +in + fun trans th1 th2 = tr [trans_HOL, trans_HOL_1, trans_HOL_2, trans_HOL_3] th1 th2 +end + +end +*} + +end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/ComputeNumeral.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/ComputeNumeral.thy Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,189 @@ +theory ComputeNumeral +imports ComputeHOL ComputeFloat +begin + +(* normalization of bit strings *) +lemmas bitnorm = normalize_bin_simps + +(* neg for bit strings *) +lemma neg1: "neg Int.Pls = False" by (simp add: Int.Pls_def) +lemma neg2: "neg Int.Min = True" apply (subst Int.Min_def) by auto +lemma neg3: "neg (Int.Bit0 x) = neg x" apply (simp add: neg_def) apply (subst Bit0_def) by auto +lemma neg4: "neg (Int.Bit1 x) = neg x" apply (simp add: neg_def) apply (subst Bit1_def) by auto +lemmas bitneg = neg1 neg2 neg3 neg4 + +(* iszero for bit strings *) +lemma iszero1: "iszero Int.Pls = True" by (simp add: Int.Pls_def iszero_def) +lemma iszero2: "iszero Int.Min = False" apply (subst Int.Min_def) apply (subst iszero_def) by simp +lemma iszero3: "iszero (Int.Bit0 x) = iszero x" apply (subst Int.Bit0_def) apply (subst iszero_def)+ by auto +lemma iszero4: "iszero (Int.Bit1 x) = False" apply (subst Int.Bit1_def) apply (subst iszero_def)+ apply simp by arith +lemmas bitiszero = iszero1 iszero2 iszero3 iszero4 + +(* lezero for bit strings *) +definition "lezero x \ x \ 0" +lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto +lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto +lemma lezero3: "lezero (Int.Bit0 x) = lezero x" unfolding Int.Bit0_def lezero_def by auto +lemma lezero4: "lezero (Int.Bit1 x) = neg x" unfolding Int.Bit1_def lezero_def neg_def by auto +lemmas bitlezero = lezero1 lezero2 lezero3 lezero4 + +(* equality for bit strings *) +lemmas biteq = eq_bin_simps + +(* x < y for bit strings *) +lemmas bitless = less_bin_simps + +(* x \ y for bit strings *) +lemmas bitle = le_bin_simps + +(* succ for bit strings *) +lemmas bitsucc = succ_bin_simps + +(* pred for bit strings *) +lemmas bitpred = pred_bin_simps + +(* unary minus for bit strings *) +lemmas bituminus = minus_bin_simps + +(* addition for bit strings *) +lemmas bitadd = add_bin_simps + +(* multiplication for bit strings *) +lemma mult_Pls_right: "x * Int.Pls = Int.Pls" by (simp add: Pls_def) +lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute) simp +lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0) +lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp +lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)" + unfolding Bit0_def Bit1_def by (simp add: algebra_simps) +lemmas bitmul = mult_Pls mult_Min mult_Pls_right mult_Min_right multb0x multxb0 multb1 + +lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul + +definition "nat_norm_number_of (x::nat) = x" + +lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)" + apply (simp add: nat_norm_number_of_def) + unfolding lezero_def iszero_def neg_def + apply (simp add: numeral_simps) + done + +(* Normalization of nat literals *) +lemma natnorm0: "(0::nat) = number_of (Int.Pls)" by auto +lemma natnorm1: "(1 :: nat) = number_of (Int.Bit1 Int.Pls)" by auto +lemmas natnorm = natnorm0 natnorm1 nat_norm_number_of + +(* Suc *) +lemma natsuc: "Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))" by (auto simp add: number_of_is_id) + +(* Addition for nat *) +lemma natadd: "number_of x + ((number_of y)::nat) = (if neg x then (number_of y) else (if neg y then number_of x else (number_of (x + y))))" + unfolding nat_number_of_def number_of_is_id neg_def + by auto + +(* Subtraction for nat *) +lemma natsub: "(number_of x) - ((number_of y)::nat) = + (if neg x then 0 else (if neg y then number_of x else (nat_norm_number_of (number_of (x + (- y))))))" + unfolding nat_norm_number_of + by (auto simp add: number_of_is_id neg_def lezero_def iszero_def Let_def nat_number_of_def) + +(* Multiplication for nat *) +lemma natmul: "(number_of x) * ((number_of y)::nat) = + (if neg x then 0 else (if neg y then 0 else number_of (x * y)))" + unfolding nat_number_of_def number_of_is_id neg_def + by (simp add: nat_mult_distrib) + +lemma nateq: "(((number_of x)::nat) = (number_of y)) = ((lezero x \ lezero y) \ (x = y))" + by (auto simp add: iszero_def lezero_def neg_def number_of_is_id) + +lemma natless: "(((number_of x)::nat) < (number_of y)) = ((x < y) \ (\ (lezero y)))" + by (simp add: lezero_def numeral_simps not_le) + +lemma natle: "(((number_of x)::nat) \ (number_of y)) = (y < x \ lezero x)" + by (auto simp add: number_of_is_id lezero_def nat_number_of_def) + +fun natfac :: "nat \ nat" + where "natfac n = (if n = 0 then 1 else n * (natfac (n - 1)))" + +lemmas compute_natarith = bitarith natnorm natsuc natadd natsub natmul nateq natless natle natfac.simps + +lemma number_eq: "(((number_of x)::'a::{number_ring, linordered_idom}) = (number_of y)) = (x = y)" + unfolding number_of_eq + apply simp + done + +lemma number_le: "(((number_of x)::'a::{number_ring, linordered_idom}) \ (number_of y)) = (x \ y)" + unfolding number_of_eq + apply simp + done + +lemma number_less: "(((number_of x)::'a::{number_ring, linordered_idom}) < (number_of y)) = (x < y)" + unfolding number_of_eq + apply simp + done + +lemma number_diff: "((number_of x)::'a::{number_ring, linordered_idom}) - number_of y = number_of (x + (- y))" + apply (subst diff_number_of_eq) + apply simp + done + +lemmas number_norm = number_of_Pls[symmetric] numeral_1_eq_1[symmetric] + +lemmas compute_numberarith = number_of_minus[symmetric] number_of_add[symmetric] number_diff number_of_mult[symmetric] number_norm number_eq number_le number_less + +lemma compute_real_of_nat_number_of: "real ((number_of v)::nat) = (if neg v then 0 else number_of v)" + by (simp only: real_of_nat_number_of number_of_is_id) + +lemma compute_nat_of_int_number_of: "nat ((number_of v)::int) = (number_of v)" + by simp + +lemmas compute_num_conversions = compute_real_of_nat_number_of compute_nat_of_int_number_of real_number_of + +lemmas zpowerarith = zpower_number_of_even + zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring] + zpower_Pls zpower_Min + +(* div, mod *) + +lemma adjust: "adjust b (q, r) = (if 0 \ r - b then (2 * q + 1, r - b) else (2 * q, r))" + by (auto simp only: adjust_def) + +lemma divmod: "divmod_int a b = (if 0\a then + if 0\b then posDivAlg a b + else if a=0 then (0, 0) + else apsnd uminus (negDivAlg (-a) (-b)) + else + if 0 program + +exception Run of string; +val run : program -> term -> term + +(* Utilities *) + +val check_freevars : int -> term -> bool +val forall_consts : (int -> bool) -> term -> bool +val closed : term -> bool +val erase_Computed : term -> term + +end + +structure AbstractMachine : ABSTRACT_MACHINE = +struct + +datatype term = Var of int | Const of int | App of term * term | Abs of term | Computed of term + +datatype pattern = PVar | PConst of int * (pattern list) + +datatype guard = Guard of term * term + +type program = unit + +exception Compile of string; + +fun erase_Computed (Computed t) = erase_Computed t + | erase_Computed (App (t1, t2)) = App (erase_Computed t1, erase_Computed t2) + | erase_Computed (Abs t) = Abs (erase_Computed t) + | erase_Computed t = t + +(*Returns true iff at most 0 .. (free-1) occur unbound. therefore + check_freevars 0 t iff t is closed*) +fun check_freevars free (Var x) = x < free + | check_freevars free (Const _) = true + | check_freevars free (App (u, v)) = check_freevars free u andalso check_freevars free v + | check_freevars free (Abs m) = check_freevars (free+1) m + | check_freevars free (Computed t) = check_freevars free t + +fun forall_consts pred (Const c) = pred c + | forall_consts pred (Var _) = true + | forall_consts pred (App (u,v)) = forall_consts pred u + andalso forall_consts pred v + | forall_consts pred (Abs m) = forall_consts pred m + | forall_consts pred (Computed t) = forall_consts pred t + +fun closed t = check_freevars 0 t + +fun compile _ = raise Compile "abstract machine stub" + +exception Run of string; + +fun run _ _ = raise Run "abstract machine stub" + +end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/Compute_Oracle/am_compiler.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/Compute_Oracle/am_compiler.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,208 @@ +(* Title: HOL/Matrix/Compute_Oracle/am_compiler.ML + Author: Steven Obua +*) + +signature COMPILING_AM = +sig + include ABSTRACT_MACHINE + + val set_compiled_rewriter : (term -> term) -> unit + val list_nth : 'a list * int -> 'a + val list_map : ('a -> 'b) -> 'a list -> 'b list +end + +structure AM_Compiler : COMPILING_AM = struct + +val list_nth = List.nth; +val list_map = map; + +open AbstractMachine; + +val compiled_rewriter = Unsynchronized.ref (NONE:(term -> term)Option.option) + +fun set_compiled_rewriter r = (compiled_rewriter := SOME r) + +type program = (term -> term) + +fun count_patternvars PVar = 1 + | count_patternvars (PConst (_, ps)) = + List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps + +fun print_rule (p, t) = + let + fun str x = string_of_int x + fun print_pattern n PVar = (n+1, "x"^(str n)) + | print_pattern n (PConst (c, [])) = (n, "c"^(str c)) + | print_pattern n (PConst (c, args)) = + let + val h = print_pattern n (PConst (c,[])) + in + print_pattern_list h args + end + and print_pattern_list r [] = r + | print_pattern_list (n, p) (t::ts) = + let + val (n, t) = print_pattern n t + in + print_pattern_list (n, "App ("^p^", "^t^")") ts + end + + val (n, pattern) = print_pattern 0 p + val pattern = + if exists_string Symbol.is_ascii_blank pattern then "(" ^ pattern ^")" + else pattern + + fun print_term d (Var x) = "Var " ^ str x + | print_term d (Const c) = "c" ^ str c + | print_term d (App (a,b)) = "App (" ^ print_term d a ^ ", " ^ print_term d b ^ ")" + | print_term d (Abs c) = "Abs (" ^ print_term (d + 1) c ^ ")" + | print_term d (Computed c) = print_term d c + + fun listvars n = if n = 0 then "x0" else "x"^(str n)^", "^(listvars (n-1)) + + val term = print_term 0 t + val term = + if n > 0 then "Closure (["^(listvars (n-1))^"], "^term^")" + else "Closure ([], "^term^")" + + in + " | weak_reduce (false, stack, "^pattern^") = Continue (false, stack, "^term^")" + end + +fun constants_of PVar = [] + | constants_of (PConst (c, ps)) = c :: maps constants_of ps + +fun constants_of_term (Var _) = [] + | constants_of_term (Abs m) = constants_of_term m + | constants_of_term (App (a,b)) = (constants_of_term a)@(constants_of_term b) + | constants_of_term (Const c) = [c] + | constants_of_term (Computed c) = constants_of_term c + +fun load_rules sname name prog = + let + val buffer = Unsynchronized.ref "" + fun write s = (buffer := (!buffer)^s) + fun writeln s = (write s; write "\n") + fun writelist [] = () + | writelist (s::ss) = (writeln s; writelist ss) + fun str i = string_of_int i + val _ = writelist [ + "structure "^name^" = struct", + "", + "datatype term = Dummy | App of term * term | Abs of term | Var of int | Const of int | Closure of term list * term"] + val constants = distinct (op =) (maps (fn (p, r) => ((constants_of p)@(constants_of_term r))) prog) + val _ = map (fn x => write (" | c"^(str x))) constants + val _ = writelist [ + "", + "datatype stack = SEmpty | SAppL of term * stack | SAppR of term * stack | SAbs of stack", + "", + "type state = bool * stack * term", + "", + "datatype loopstate = Continue of state | Stop of stack * term", + "", + "fun proj_C (Continue s) = s", + " | proj_C _ = raise Match", + "", + "fun proj_S (Stop s) = s", + " | proj_S _ = raise Match", + "", + "fun cont (Continue _) = true", + " | cont _ = false", + "", + "fun do_reduction reduce p =", + " let", + " val s = Unsynchronized.ref (Continue p)", + " val _ = while cont (!s) do (s := reduce (proj_C (!s)))", + " in", + " proj_S (!s)", + " end", + ""] + + val _ = writelist [ + "fun weak_reduce (false, stack, Closure (e, App (a, b))) = Continue (false, SAppL (Closure (e, b), stack), Closure (e, a))", + " | weak_reduce (false, SAppL (b, stack), Closure (e, Abs m)) = Continue (false, stack, Closure (b::e, m))", + " | weak_reduce (false, stack, c as Closure (e, Abs m)) = Continue (true, stack, c)", + " | weak_reduce (false, stack, Closure (e, Var n)) = Continue (false, stack, case "^sname^".list_nth (e, n) of Dummy => Var n | r => r)", + " | weak_reduce (false, stack, Closure (e, c)) = Continue (false, stack, c)"] + val _ = writelist (map print_rule prog) + val _ = writelist [ + " | weak_reduce (false, stack, clos) = Continue (true, stack, clos)", + " | weak_reduce (true, SAppR (a, stack), b) = Continue (false, stack, App (a,b))", + " | weak_reduce (true, s as (SAppL (b, stack)), a) = Continue (false, SAppR (a, stack), b)", + " | weak_reduce (true, stack, c) = Stop (stack, c)", + "", + "fun strong_reduce (false, stack, Closure (e, Abs m)) =", + " let", + " val (stack', wnf) = do_reduction weak_reduce (false, SEmpty, Closure (Dummy::e, m))", + " in", + " case stack' of", + " SEmpty => Continue (false, SAbs stack, wnf)", + " | _ => raise ("^sname^".Run \"internal error in strong: weak failed\")", + " end", + " | strong_reduce (false, stack, clos as (App (u, v))) = Continue (false, SAppL (v, stack), u)", + " | strong_reduce (false, stack, clos) = Continue (true, stack, clos)", + " | strong_reduce (true, SAbs stack, m) = Continue (false, stack, Abs m)", + " | strong_reduce (true, SAppL (b, stack), a) = Continue (false, SAppR (a, stack), b)", + " | strong_reduce (true, SAppR (a, stack), b) = Continue (true, stack, App (a, b))", + " | strong_reduce (true, stack, clos) = Stop (stack, clos)", + ""] + + val ic = "(case c of "^(implode (map (fn c => (str c)^" => c"^(str c)^" | ") constants))^" _ => Const c)" + val _ = writelist [ + "fun importTerm ("^sname^".Var x) = Var x", + " | importTerm ("^sname^".Const c) = "^ic, + " | importTerm ("^sname^".App (a, b)) = App (importTerm a, importTerm b)", + " | importTerm ("^sname^".Abs m) = Abs (importTerm m)", + ""] + + fun ec c = " | exportTerm c"^(str c)^" = "^sname^".Const "^(str c) + val _ = writelist [ + "fun exportTerm (Var x) = "^sname^".Var x", + " | exportTerm (Const c) = "^sname^".Const c", + " | exportTerm (App (a,b)) = "^sname^".App (exportTerm a, exportTerm b)", + " | exportTerm (Abs m) = "^sname^".Abs (exportTerm m)", + " | exportTerm (Closure (closlist, clos)) = raise ("^sname^".Run \"internal error, cannot export Closure\")", + " | exportTerm Dummy = raise ("^sname^".Run \"internal error, cannot export Dummy\")"] + val _ = writelist (map ec constants) + + val _ = writelist [ + "", + "fun rewrite t = ", + " let", + " val (stack, wnf) = do_reduction weak_reduce (false, SEmpty, Closure ([], importTerm t))", + " in", + " case stack of ", + " SEmpty => (case do_reduction strong_reduce (false, SEmpty, wnf) of", + " (SEmpty, snf) => exportTerm snf", + " | _ => raise ("^sname^".Run \"internal error in rewrite: strong failed\"))", + " | _ => (raise ("^sname^".Run \"internal error in rewrite: weak failed\"))", + " end", + "", + "val _ = "^sname^".set_compiled_rewriter rewrite", + "", + "end;"] + + in + compiled_rewriter := NONE; + use_text ML_Env.local_context (1, "") false (!buffer); + case !compiled_rewriter of + NONE => raise (Compile "cannot communicate with compiled function") + | SOME r => (compiled_rewriter := NONE; r) + end + +fun compile eqs = + let + val _ = if exists (fn (a,_,_) => not (null a)) eqs then raise Compile ("cannot deal with guards") else () + val eqs = map (fn (_,b,c) => (b,c)) eqs + fun check (p, r) = if check_freevars (count_patternvars p) r then () else raise Compile ("unbound variables in rule") + val _ = map (fn (p, r) => + (check (p, r); + case p of PVar => raise (Compile "pattern is just a variable") | _ => ())) eqs + in + load_rules "AM_Compiler" "AM_compiled_code" eqs + end + +fun run prog t = prog t + +end + diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/Compute_Oracle/am_ghc.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/Compute_Oracle/am_ghc.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,324 @@ +(* Title: HOL/Matrix/Compute_Oracle/am_ghc.ML + Author: Steven Obua +*) + +structure AM_GHC : ABSTRACT_MACHINE = +struct + +open AbstractMachine; + +type program = string * string * (int Inttab.table) + +fun count_patternvars PVar = 1 + | count_patternvars (PConst (_, ps)) = + List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps + +fun update_arity arity code a = + (case Inttab.lookup arity code of + NONE => Inttab.update_new (code, a) arity + | SOME (a': int) => if a > a' then Inttab.update (code, a) arity else arity) + +(* We have to find out the maximal arity of each constant *) +fun collect_pattern_arity PVar arity = arity + | collect_pattern_arity (PConst (c, args)) arity = fold collect_pattern_arity args (update_arity arity c (length args)) + +local +fun collect applevel (Var _) arity = arity + | collect applevel (Const c) arity = update_arity arity c applevel + | collect applevel (Abs m) arity = collect 0 m arity + | collect applevel (App (a,b)) arity = collect 0 b (collect (applevel + 1) a arity) +in +fun collect_term_arity t arity = collect 0 t arity +end + +fun nlift level n (Var m) = if m < level then Var m else Var (m+n) + | nlift level n (Const c) = Const c + | nlift level n (App (a,b)) = App (nlift level n a, nlift level n b) + | nlift level n (Abs b) = Abs (nlift (level+1) n b) + +fun rep n x = if n = 0 then [] else x::(rep (n-1) x) + +fun adjust_rules rules = + let + val arity = fold (fn (p, t) => fn arity => collect_term_arity t (collect_pattern_arity p arity)) rules Inttab.empty + fun arity_of c = the (Inttab.lookup arity c) + fun adjust_pattern PVar = PVar + | adjust_pattern (C as PConst (c, args)) = if (length args <> arity_of c) then raise Compile ("Constant inside pattern must have maximal arity") else C + fun adjust_rule (PVar, _) = raise Compile ("pattern may not be a variable") + | adjust_rule (rule as (p as PConst (c, args),t)) = + let + val _ = if not (check_freevars (count_patternvars p) t) then raise Compile ("unbound variables on right hand side") else () + val args = map adjust_pattern args + val len = length args + val arity = arity_of c + fun lift level n (Var m) = if m < level then Var m else Var (m+n) + | lift level n (Const c) = Const c + | lift level n (App (a,b)) = App (lift level n a, lift level n b) + | lift level n (Abs b) = Abs (lift (level+1) n b) + val lift = lift 0 + fun adjust_term n t = if n=0 then t else adjust_term (n-1) (App (t, Var (n-1))) + in + if len = arity then + rule + else if arity >= len then + (PConst (c, args @ (rep (arity-len) PVar)), adjust_term (arity-len) (lift (arity-len) t)) + else (raise Compile "internal error in adjust_rule") + end + in + (arity, map adjust_rule rules) + end + +fun print_term arity_of n = +let + fun str x = string_of_int x + fun protect_blank s = if exists_string Symbol.is_ascii_blank s then "(" ^ s ^")" else s + + fun print_apps d f [] = f + | print_apps d f (a::args) = print_apps d ("app "^(protect_blank f)^" "^(protect_blank (print_term d a))) args + and print_call d (App (a, b)) args = print_call d a (b::args) + | print_call d (Const c) args = + (case arity_of c of + NONE => print_apps d ("Const "^(str c)) args + | SOME a => + let + val len = length args + in + if a <= len then + let + val s = "c"^(str c)^(implode (map (fn t => " "^(protect_blank (print_term d t))) (List.take (args, a)))) + in + print_apps d s (List.drop (args, a)) + end + else + let + fun mk_apps n t = if n = 0 then t else mk_apps (n-1) (App (t, Var (n-1))) + fun mk_lambdas n t = if n = 0 then t else mk_lambdas (n-1) (Abs t) + fun append_args [] t = t + | append_args (c::cs) t = append_args cs (App (t, c)) + in + print_term d (mk_lambdas (a-len) (mk_apps (a-len) (nlift 0 (a-len) (append_args args (Const c))))) + end + end) + | print_call d t args = print_apps d (print_term d t) args + and print_term d (Var x) = if x < d then "b"^(str (d-x-1)) else "x"^(str (n-(x-d)-1)) + | print_term d (Abs c) = "Abs (\\b"^(str d)^" -> "^(print_term (d + 1) c)^")" + | print_term d t = print_call d t [] +in + print_term 0 +end + +fun print_rule arity_of (p, t) = + let + fun str x = string_of_int x + fun print_pattern top n PVar = (n+1, "x"^(str n)) + | print_pattern top n (PConst (c, [])) = (n, (if top then "c" else "C")^(str c)) + | print_pattern top n (PConst (c, args)) = + let + val (n,s) = print_pattern_list (n, (if top then "c" else "C")^(str c)) args + in + (n, if top then s else "("^s^")") + end + and print_pattern_list r [] = r + | print_pattern_list (n, p) (t::ts) = + let + val (n, t) = print_pattern false n t + in + print_pattern_list (n, p^" "^t) ts + end + val (n, pattern) = print_pattern true 0 p + in + pattern^" = "^(print_term arity_of n t) + end + +fun group_rules rules = + let + fun add_rule (r as (PConst (c,_), _)) groups = + let + val rs = (case Inttab.lookup groups c of NONE => [] | SOME rs => rs) + in + Inttab.update (c, r::rs) groups + end + | add_rule _ _ = raise Compile "internal error group_rules" + in + fold_rev add_rule rules Inttab.empty + end + +fun haskell_prog name rules = + let + val buffer = Unsynchronized.ref "" + fun write s = (buffer := (!buffer)^s) + fun writeln s = (write s; write "\n") + fun writelist [] = () + | writelist (s::ss) = (writeln s; writelist ss) + fun str i = string_of_int i + val (arity, rules) = adjust_rules rules + val rules = group_rules rules + val constants = Inttab.keys arity + fun arity_of c = Inttab.lookup arity c + fun rep_str s n = implode (rep n s) + fun indexed s n = s^(str n) + fun section n = if n = 0 then [] else (section (n-1))@[n-1] + fun make_show c = + let + val args = section (the (arity_of c)) + in + " show ("^(indexed "C" c)^(implode (map (indexed " a") args))^") = " + ^"\""^(indexed "C" c)^"\""^(implode (map (fn a => "++(show "^(indexed "a" a)^")") args)) + end + fun default_case c = + let + val args = implode (map (indexed " x") (section (the (arity_of c)))) + in + (indexed "c" c)^args^" = "^(indexed "C" c)^args + end + val _ = writelist [ + "module "^name^" where", + "", + "data Term = Const Integer | App Term Term | Abs (Term -> Term)", + " "^(implode (map (fn c => " | C"^(str c)^(rep_str " Term" (the (arity_of c)))) constants)), + "", + "instance Show Term where"] + val _ = writelist (map make_show constants) + val _ = writelist [ + " show (Const c) = \"c\"++(show c)", + " show (App a b) = \"A\"++(show a)++(show b)", + " show (Abs _) = \"L\"", + ""] + val _ = writelist [ + "app (Abs a) b = a b", + "app a b = App a b", + "", + "calc s c = writeFile s (show c)", + ""] + fun list_group c = (writelist (case Inttab.lookup rules c of + NONE => [default_case c, ""] + | SOME (rs as ((PConst (_, []), _)::rs')) => + if not (null rs') then raise Compile "multiple declaration of constant" + else (map (print_rule arity_of) rs) @ [""] + | SOME rs => (map (print_rule arity_of) rs) @ [default_case c, ""])) + val _ = map list_group constants + in + (arity, !buffer) + end + +val guid_counter = Unsynchronized.ref 0 +fun get_guid () = + let + val c = !guid_counter + val _ = guid_counter := !guid_counter + 1 + in + string_of_int (Time.toMicroseconds (Time.now ())) ^ string_of_int c + end + +fun tmp_file s = Path.implode (Path.expand (File.tmp_path (Path.basic s))); + +fun writeTextFile name s = File.write (Path.explode name) s + +fun fileExists name = ((OS.FileSys.fileSize name; true) handle OS.SysErr _ => false) + +fun compile eqs = + let + val _ = if exists (fn (a,_,_) => not (null a)) eqs then raise Compile ("cannot deal with guards") else () + val eqs = map (fn (_,b,c) => (b,c)) eqs + val guid = get_guid () + val module = "AMGHC_Prog_"^guid + val (arity, source) = haskell_prog module eqs + val module_file = tmp_file (module^".hs") + val object_file = tmp_file (module^".o") + val _ = writeTextFile module_file source + val _ = Isabelle_System.bash ("exec \"$ISABELLE_GHC\" -c " ^ module_file) + val _ = + if not (fileExists object_file) then + raise Compile ("Failure compiling haskell code (ISABELLE_GHC='" ^ getenv "ISABELLE_GHC" ^ "')") + else () + in + (guid, module_file, arity) + end + +fun readResultFile name = File.read (Path.explode name) + +fun parse_result arity_of result = + let + val result = String.explode result + fun shift NONE x = SOME x + | shift (SOME y) x = SOME (y*10 + x) + fun parse_int' x (#"0"::rest) = parse_int' (shift x 0) rest + | parse_int' x (#"1"::rest) = parse_int' (shift x 1) rest + | parse_int' x (#"2"::rest) = parse_int' (shift x 2) rest + | parse_int' x (#"3"::rest) = parse_int' (shift x 3) rest + | parse_int' x (#"4"::rest) = parse_int' (shift x 4) rest + | parse_int' x (#"5"::rest) = parse_int' (shift x 5) rest + | parse_int' x (#"6"::rest) = parse_int' (shift x 6) rest + | parse_int' x (#"7"::rest) = parse_int' (shift x 7) rest + | parse_int' x (#"8"::rest) = parse_int' (shift x 8) rest + | parse_int' x (#"9"::rest) = parse_int' (shift x 9) rest + | parse_int' x rest = (x, rest) + fun parse_int rest = parse_int' NONE rest + + fun parse (#"C"::rest) = + (case parse_int rest of + (SOME c, rest) => + let + val (args, rest) = parse_list (the (arity_of c)) rest + fun app_args [] t = t + | app_args (x::xs) t = app_args xs (App (t, x)) + in + (app_args args (Const c), rest) + end + | (NONE, _) => raise Run "parse C") + | parse (#"c"::rest) = + (case parse_int rest of + (SOME c, rest) => (Const c, rest) + | _ => raise Run "parse c") + | parse (#"A"::rest) = + let + val (a, rest) = parse rest + val (b, rest) = parse rest + in + (App (a,b), rest) + end + | parse (#"L"::_) = raise Run "there may be no abstraction in the result" + | parse _ = raise Run "invalid result" + and parse_list n rest = + if n = 0 then + ([], rest) + else + let + val (x, rest) = parse rest + val (xs, rest) = parse_list (n-1) rest + in + (x::xs, rest) + end + val (parsed, rest) = parse result + fun is_blank (#" "::rest) = is_blank rest + | is_blank (#"\n"::rest) = is_blank rest + | is_blank [] = true + | is_blank _ = false + in + if is_blank rest then parsed else raise Run "non-blank suffix in result file" + end + +fun run (guid, module_file, arity) t = + let + val _ = if check_freevars 0 t then () else raise Run ("can only compute closed terms") + fun arity_of c = Inttab.lookup arity c + val callguid = get_guid() + val module = "AMGHC_Prog_"^guid + val call = module^"_Call_"^callguid + val result_file = tmp_file (module^"_Result_"^callguid^".txt") + val call_file = tmp_file (call^".hs") + val term = print_term arity_of 0 t + val call_source = "module "^call^" where\n\nimport "^module^"\n\ncall = "^module^".calc \""^result_file^"\" ("^term^")" + val _ = writeTextFile call_file call_source + val _ = Isabelle_System.bash ("exec \"$ISABELLE_GHC\" -e \""^call^".call\" "^module_file^" "^call_file) + val result = readResultFile result_file handle IO.Io _ => + raise Run ("Failure running haskell compiler (ISABELLE_GHC='" ^ getenv "ISABELLE_GHC" ^ "')") + val t' = parse_result arity_of result + val _ = OS.FileSys.remove call_file + val _ = OS.FileSys.remove result_file + in + t' + end + +end + diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/Compute_Oracle/am_interpreter.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/Compute_Oracle/am_interpreter.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,211 @@ +(* Title: HOL/Matrix/Compute_Oracle/am_interpreter.ML + Author: Steven Obua +*) + +signature AM_BARRAS = +sig + include ABSTRACT_MACHINE + val max_reductions : int option Unsynchronized.ref +end + +structure AM_Interpreter : AM_BARRAS = struct + +open AbstractMachine; + +datatype closure = CDummy | CVar of int | CConst of int + | CApp of closure * closure | CAbs of closure + | Closure of (closure list) * closure + +structure prog_struct = Table(type key = int*int val ord = prod_ord int_ord int_ord); + +datatype program = Program of ((pattern * closure * (closure*closure) list) list) prog_struct.table + +datatype stack = SEmpty | SAppL of closure * stack | SAppR of closure * stack | SAbs of stack + +fun clos_of_term (Var x) = CVar x + | clos_of_term (Const c) = CConst c + | clos_of_term (App (u, v)) = CApp (clos_of_term u, clos_of_term v) + | clos_of_term (Abs u) = CAbs (clos_of_term u) + | clos_of_term (Computed t) = clos_of_term t + +fun term_of_clos (CVar x) = Var x + | term_of_clos (CConst c) = Const c + | term_of_clos (CApp (u, v)) = App (term_of_clos u, term_of_clos v) + | term_of_clos (CAbs u) = Abs (term_of_clos u) + | term_of_clos (Closure _) = raise (Run "internal error: closure in normalized term found") + | term_of_clos CDummy = raise (Run "internal error: dummy in normalized term found") + +fun resolve_closure closures (CVar x) = (case nth closures x of CDummy => CVar x | r => r) + | resolve_closure closures (CConst c) = CConst c + | resolve_closure closures (CApp (u, v)) = CApp (resolve_closure closures u, resolve_closure closures v) + | resolve_closure closures (CAbs u) = CAbs (resolve_closure (CDummy::closures) u) + | resolve_closure closures (CDummy) = raise (Run "internal error: resolve_closure applied to CDummy") + | resolve_closure closures (Closure (e, u)) = resolve_closure e u + +fun resolve_closure' c = resolve_closure [] c + +fun resolve_stack tm SEmpty = tm + | resolve_stack tm (SAppL (c, s)) = resolve_stack (CApp (tm, resolve_closure' c)) s + | resolve_stack tm (SAppR (c, s)) = resolve_stack (CApp (resolve_closure' c, tm)) s + | resolve_stack tm (SAbs s) = resolve_stack (CAbs tm) s + +fun resolve (stack, closure) = + let + val _ = writeln "start resolving" + val t = resolve_stack (resolve_closure' closure) stack + val _ = writeln "finished resolving" + in + t + end + +fun strip_closure args (CApp (a,b)) = strip_closure (b::args) a + | strip_closure args x = (x, args) + +fun len_head_of_closure n (CApp (a, _)) = len_head_of_closure (n+1) a + | len_head_of_closure n x = (n, x) + + +(* earlier occurrence of PVar corresponds to higher de Bruijn index *) +fun pattern_match args PVar clos = SOME (clos::args) + | pattern_match args (PConst (c, patterns)) clos = + let + val (f, closargs) = strip_closure [] clos + in + case f of + CConst d => + if c = d then + pattern_match_list args patterns closargs + else + NONE + | _ => NONE + end +and pattern_match_list args [] [] = SOME args + | pattern_match_list args (p::ps) (c::cs) = + (case pattern_match args p c of + NONE => NONE + | SOME args => pattern_match_list args ps cs) + | pattern_match_list _ _ _ = NONE + +fun count_patternvars PVar = 1 + | count_patternvars (PConst (_, ps)) = List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps + +fun pattern_key (PConst (c, ps)) = (c, length ps) + | pattern_key _ = raise (Compile "pattern reduces to variable") + +(*Returns true iff at most 0 .. (free-1) occur unbound. therefore + check_freevars 0 t iff t is closed*) +fun check_freevars free (Var x) = x < free + | check_freevars free (Const _) = true + | check_freevars free (App (u, v)) = check_freevars free u andalso check_freevars free v + | check_freevars free (Abs m) = check_freevars (free+1) m + | check_freevars free (Computed t) = check_freevars free t + +fun compile eqs = + let + fun check p r = if check_freevars p r then () else raise Compile ("unbound variables in rule") + fun check_guard p (Guard (a,b)) = (check p a; check p b) + fun clos_of_guard (Guard (a,b)) = (clos_of_term a, clos_of_term b) + val eqs = map (fn (guards, p, r) => let val pcount = count_patternvars p val _ = map (check_guard pcount) (guards) val _ = check pcount r in + (pattern_key p, (p, clos_of_term r, map clos_of_guard guards)) end) eqs + fun merge (k, a) table = prog_struct.update (k, case prog_struct.lookup table k of NONE => [a] | SOME l => a::l) table + val p = fold merge eqs prog_struct.empty + in + Program p + end + + +type state = bool * program * stack * closure + +datatype loopstate = Continue of state | Stop of stack * closure + +fun proj_C (Continue s) = s + | proj_C _ = raise Match + +exception InterruptedExecution of stack * closure + +fun proj_S (Stop s) = s + | proj_S (Continue (_,_,s,c)) = (s,c) + +fun cont (Continue _) = true + | cont _ = false + +val max_reductions = Unsynchronized.ref (NONE : int option) + +fun do_reduction reduce p = + let + val s = Unsynchronized.ref (Continue p) + val counter = Unsynchronized.ref 0 + val _ = case !max_reductions of + NONE => while cont (!s) do (s := reduce (proj_C (!s))) + | SOME m => while cont (!s) andalso (!counter < m) do (s := reduce (proj_C (!s)); counter := (!counter) + 1) + in + case !max_reductions of + SOME m => if !counter >= m then raise InterruptedExecution (proj_S (!s)) else proj_S (!s) + | NONE => proj_S (!s) + end + +fun match_rules prog n [] clos = NONE + | match_rules prog n ((p,eq,guards)::rs) clos = + case pattern_match [] p clos of + NONE => match_rules prog (n+1) rs clos + | SOME args => if forall (guard_checks prog args) guards then SOME (Closure (args, eq)) else match_rules prog (n+1) rs clos +and guard_checks prog args (a,b) = (simp prog (Closure (args, a)) = simp prog (Closure (args, b))) +and match_closure (p as (Program prog)) clos = + case len_head_of_closure 0 clos of + (len, CConst c) => + (case prog_struct.lookup prog (c, len) of + NONE => NONE + | SOME rules => match_rules p 0 rules clos) + | _ => NONE + +and weak_reduce (false, prog, stack, Closure (e, CApp (a, b))) = Continue (false, prog, SAppL (Closure (e, b), stack), Closure (e, a)) + | weak_reduce (false, prog, SAppL (b, stack), Closure (e, CAbs m)) = Continue (false, prog, stack, Closure (b::e, m)) + | weak_reduce (false, prog, stack, Closure (e, CVar n)) = Continue (false, prog, stack, case nth e n of CDummy => CVar n | r => r) + | weak_reduce (false, prog, stack, Closure (_, c as CConst _)) = Continue (false, prog, stack, c) + | weak_reduce (false, prog, stack, clos) = + (case match_closure prog clos of + NONE => Continue (true, prog, stack, clos) + | SOME r => Continue (false, prog, stack, r)) + | weak_reduce (true, prog, SAppR (a, stack), b) = Continue (false, prog, stack, CApp (a,b)) + | weak_reduce (true, prog, SAppL (b, stack), a) = Continue (false, prog, SAppR (a, stack), b) + | weak_reduce (true, prog, stack, c) = Stop (stack, c) + +and strong_reduce (false, prog, stack, Closure (e, CAbs m)) = + (let + val (stack', wnf) = do_reduction weak_reduce (false, prog, SEmpty, Closure (CDummy::e, m)) + in + case stack' of + SEmpty => Continue (false, prog, SAbs stack, wnf) + | _ => raise (Run "internal error in strong: weak failed") + end handle InterruptedExecution state => raise InterruptedExecution (stack, resolve state)) + | strong_reduce (false, prog, stack, CApp (u, v)) = Continue (false, prog, SAppL (v, stack), u) + | strong_reduce (false, prog, stack, clos) = Continue (true, prog, stack, clos) + | strong_reduce (true, prog, SAbs stack, m) = Continue (false, prog, stack, CAbs m) + | strong_reduce (true, prog, SAppL (b, stack), a) = Continue (false, prog, SAppR (a, stack), b) + | strong_reduce (true, prog, SAppR (a, stack), b) = Continue (true, prog, stack, CApp (a, b)) + | strong_reduce (true, prog, stack, clos) = Stop (stack, clos) + +and simp prog t = + (let + val (stack, wnf) = do_reduction weak_reduce (false, prog, SEmpty, t) + in + case stack of + SEmpty => (case do_reduction strong_reduce (false, prog, SEmpty, wnf) of + (SEmpty, snf) => snf + | _ => raise (Run "internal error in run: strong failed")) + | _ => raise (Run "internal error in run: weak failed") + end handle InterruptedExecution state => resolve state) + + +fun run prog t = + (let + val (stack, wnf) = do_reduction weak_reduce (false, prog, SEmpty, Closure ([], clos_of_term t)) + in + case stack of + SEmpty => (case do_reduction strong_reduce (false, prog, SEmpty, wnf) of + (SEmpty, snf) => term_of_clos snf + | _ => raise (Run "internal error in run: strong failed")) + | _ => raise (Run "internal error in run: weak failed") + end handle InterruptedExecution state => term_of_clos (resolve state)) + +end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/Compute_Oracle/am_sml.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/Compute_Oracle/am_sml.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,517 @@ +(* Title: HOL/Matrix/Compute_Oracle/am_sml.ML + Author: Steven Obua + +TODO: "parameterless rewrite cannot be used in pattern": In a lot of +cases it CAN be used, and these cases should be handled +properly; right now, all cases raise an exception. +*) + +signature AM_SML = +sig + include ABSTRACT_MACHINE + val save_result : (string * term) -> unit + val set_compiled_rewriter : (term -> term) -> unit + val list_nth : 'a list * int -> 'a + val dump_output : (string option) Unsynchronized.ref +end + +structure AM_SML : AM_SML = struct + +open AbstractMachine; + +val dump_output = Unsynchronized.ref (NONE: string option) + +type program = term Inttab.table * (term -> term) + +val saved_result = Unsynchronized.ref (NONE:(string*term)option) + +fun save_result r = (saved_result := SOME r) + +val list_nth = List.nth + +val compiled_rewriter = Unsynchronized.ref (NONE:(term -> term)Option.option) + +fun set_compiled_rewriter r = (compiled_rewriter := SOME r) + +fun count_patternvars PVar = 1 + | count_patternvars (PConst (_, ps)) = + List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps + +fun update_arity arity code a = + (case Inttab.lookup arity code of + NONE => Inttab.update_new (code, a) arity + | SOME (a': int) => if a > a' then Inttab.update (code, a) arity else arity) + +(* We have to find out the maximal arity of each constant *) +fun collect_pattern_arity PVar arity = arity + | collect_pattern_arity (PConst (c, args)) arity = fold collect_pattern_arity args (update_arity arity c (length args)) + +(* We also need to find out the maximal toplevel arity of each function constant *) +fun collect_pattern_toplevel_arity PVar arity = raise Compile "internal error: collect_pattern_toplevel_arity" + | collect_pattern_toplevel_arity (PConst (c, args)) arity = update_arity arity c (length args) + +local +fun collect applevel (Var _) arity = arity + | collect applevel (Const c) arity = update_arity arity c applevel + | collect applevel (Abs m) arity = collect 0 m arity + | collect applevel (App (a,b)) arity = collect 0 b (collect (applevel + 1) a arity) +in +fun collect_term_arity t arity = collect 0 t arity +end + +fun collect_guard_arity (Guard (a,b)) arity = collect_term_arity b (collect_term_arity a arity) + + +fun rep n x = if n < 0 then raise Compile "internal error: rep" else if n = 0 then [] else x::(rep (n-1) x) + +fun beta (Const c) = Const c + | beta (Var i) = Var i + | beta (App (Abs m, b)) = beta (unlift 0 (subst 0 m (lift 0 b))) + | beta (App (a, b)) = + (case beta a of + Abs m => beta (App (Abs m, b)) + | a => App (a, beta b)) + | beta (Abs m) = Abs (beta m) + | beta (Computed t) = Computed t +and subst x (Const c) t = Const c + | subst x (Var i) t = if i = x then t else Var i + | subst x (App (a,b)) t = App (subst x a t, subst x b t) + | subst x (Abs m) t = Abs (subst (x+1) m (lift 0 t)) +and lift level (Const c) = Const c + | lift level (App (a,b)) = App (lift level a, lift level b) + | lift level (Var i) = if i < level then Var i else Var (i+1) + | lift level (Abs m) = Abs (lift (level + 1) m) +and unlift level (Const c) = Const c + | unlift level (App (a, b)) = App (unlift level a, unlift level b) + | unlift level (Abs m) = Abs (unlift (level+1) m) + | unlift level (Var i) = if i < level then Var i else Var (i-1) + +fun nlift level n (Var m) = if m < level then Var m else Var (m+n) + | nlift level n (Const c) = Const c + | nlift level n (App (a,b)) = App (nlift level n a, nlift level n b) + | nlift level n (Abs b) = Abs (nlift (level+1) n b) + +fun subst_const (c, t) (Const c') = if c = c' then t else Const c' + | subst_const _ (Var i) = Var i + | subst_const ct (App (a, b)) = App (subst_const ct a, subst_const ct b) + | subst_const ct (Abs m) = Abs (subst_const ct m) + +(* Remove all rules that are just parameterless rewrites. This is necessary because SML does not allow functions with no parameters. *) +fun inline_rules rules = + let + fun term_contains_const c (App (a, b)) = term_contains_const c a orelse term_contains_const c b + | term_contains_const c (Abs m) = term_contains_const c m + | term_contains_const c (Var _) = false + | term_contains_const c (Const c') = (c = c') + fun find_rewrite [] = NONE + | find_rewrite ((prems, PConst (c, []), r) :: _) = + if check_freevars 0 r then + if term_contains_const c r then + raise Compile "parameterless rewrite is caught in cycle" + else if not (null prems) then + raise Compile "parameterless rewrite may not be guarded" + else + SOME (c, r) + else raise Compile "unbound variable on right hand side or guards of rule" + | find_rewrite (_ :: rules) = find_rewrite rules + fun remove_rewrite _ [] = [] + | remove_rewrite (cr as (c, r)) ((rule as (prems', PConst (c', args), r')) :: rules) = + if c = c' then + if null args andalso r = r' andalso null prems' then remove_rewrite cr rules + else raise Compile "incompatible parameterless rewrites found" + else + rule :: remove_rewrite cr rules + | remove_rewrite cr (r :: rs) = r :: remove_rewrite cr rs + fun pattern_contains_const c (PConst (c', args)) = c = c' orelse exists (pattern_contains_const c) args + | pattern_contains_const c (PVar) = false + fun inline_rewrite (ct as (c, _)) (prems, p, r) = + if pattern_contains_const c p then + raise Compile "parameterless rewrite cannot be used in pattern" + else (map (fn (Guard (a, b)) => Guard (subst_const ct a, subst_const ct b)) prems, p, subst_const ct r) + fun inline inlined rules = + case find_rewrite rules of + NONE => (Inttab.make inlined, rules) + | SOME ct => + let + val rules = map (inline_rewrite ct) (remove_rewrite ct rules) + val inlined = ct :: (map o apsnd) (subst_const ct) inlined + in inline inlined rules end + in + inline [] rules + end + + +(* + Calculate the arity, the toplevel_arity, and adjust rules so that all toplevel pattern constants have maximal arity. + Also beta reduce the adjusted right hand side of a rule. +*) +fun adjust_rules rules = + let + val arity = fold (fn (prems, p, t) => fn arity => fold collect_guard_arity prems (collect_term_arity t (collect_pattern_arity p arity))) rules Inttab.empty + val toplevel_arity = fold (fn (_, p, _) => fn arity => collect_pattern_toplevel_arity p arity) rules Inttab.empty + fun arity_of c = the (Inttab.lookup arity c) + fun test_pattern PVar = () + | test_pattern (PConst (c, args)) = if (length args <> arity_of c) then raise Compile ("Constant inside pattern must have maximal arity") else (map test_pattern args; ()) + fun adjust_rule (_, PVar, _) = raise Compile ("pattern may not be a variable") + | adjust_rule (_, PConst (_, []), _) = raise Compile ("cannot deal with rewrites that take no parameters") + | adjust_rule (rule as (prems, p as PConst (c, args),t)) = + let + val patternvars_counted = count_patternvars p + fun check_fv t = check_freevars patternvars_counted t + val _ = if not (check_fv t) then raise Compile ("unbound variables on right hand side of rule") else () + val _ = if not (forall (fn (Guard (a,b)) => check_fv a andalso check_fv b) prems) then raise Compile ("unbound variables in guards") else () + val _ = map test_pattern args + val len = length args + val arity = arity_of c + val lift = nlift 0 + fun addapps_tm n t = if n=0 then t else addapps_tm (n-1) (App (t, Var (n-1))) + fun adjust_term n t = addapps_tm n (lift n t) + fun adjust_guard n (Guard (a,b)) = Guard (lift n a, lift n b) + in + if len = arity then + rule + else if arity >= len then + (map (adjust_guard (arity-len)) prems, PConst (c, args @ (rep (arity-len) PVar)), adjust_term (arity-len) t) + else (raise Compile "internal error in adjust_rule") + end + fun beta_rule (prems, p, t) = ((prems, p, beta t) handle Match => raise Compile "beta_rule") + in + (arity, toplevel_arity, map (beta_rule o adjust_rule) rules) + end + +fun print_term module arity_of toplevel_arity_of pattern_var_count pattern_lazy_var_count = +let + fun str x = string_of_int x + fun protect_blank s = if exists_string Symbol.is_ascii_blank s then "(" ^ s ^")" else s + val module_prefix = (case module of NONE => "" | SOME s => s^".") + fun print_apps d f [] = f + | print_apps d f (a::args) = print_apps d (module_prefix^"app "^(protect_blank f)^" "^(protect_blank (print_term d a))) args + and print_call d (App (a, b)) args = print_call d a (b::args) + | print_call d (Const c) args = + (case arity_of c of + NONE => print_apps d (module_prefix^"Const "^(str c)) args + | SOME 0 => module_prefix^"C"^(str c) + | SOME a => + let + val len = length args + in + if a <= len then + let + val strict_a = (case toplevel_arity_of c of SOME sa => sa | NONE => a) + val _ = if strict_a > a then raise Compile "strict" else () + val s = module_prefix^"c"^(str c)^(implode (map (fn t => " "^(protect_blank (print_term d t))) (List.take (args, strict_a)))) + val s = s^(implode (map (fn t => " (fn () => "^print_term d t^")") (List.drop (List.take (args, a), strict_a)))) + in + print_apps d s (List.drop (args, a)) + end + else + let + fun mk_apps n t = if n = 0 then t else mk_apps (n-1) (App (t, Var (n - 1))) + fun mk_lambdas n t = if n = 0 then t else mk_lambdas (n-1) (Abs t) + fun append_args [] t = t + | append_args (c::cs) t = append_args cs (App (t, c)) + in + print_term d (mk_lambdas (a-len) (mk_apps (a-len) (nlift 0 (a-len) (append_args args (Const c))))) + end + end) + | print_call d t args = print_apps d (print_term d t) args + and print_term d (Var x) = + if x < d then + "b"^(str (d-x-1)) + else + let + val n = pattern_var_count - (x-d) - 1 + val x = "x"^(str n) + in + if n < pattern_var_count - pattern_lazy_var_count then + x + else + "("^x^" ())" + end + | print_term d (Abs c) = module_prefix^"Abs (fn b"^(str d)^" => "^(print_term (d + 1) c)^")" + | print_term d t = print_call d t [] +in + print_term 0 +end + +fun section n = if n = 0 then [] else (section (n-1))@[n-1] + +fun print_rule gnum arity_of toplevel_arity_of (guards, p, t) = + let + fun str x = string_of_int x + fun print_pattern top n PVar = (n+1, "x"^(str n)) + | print_pattern top n (PConst (c, [])) = (n, (if top then "c" else "C")^(str c)^(if top andalso gnum > 0 then "_"^(str gnum) else "")) + | print_pattern top n (PConst (c, args)) = + let + val f = (if top then "c" else "C")^(str c)^(if top andalso gnum > 0 then "_"^(str gnum) else "") + val (n, s) = print_pattern_list 0 top (n, f) args + in + (n, s) + end + and print_pattern_list' counter top (n,p) [] = if top then (n,p) else (n,p^")") + | print_pattern_list' counter top (n, p) (t::ts) = + let + val (n, t) = print_pattern false n t + in + print_pattern_list' (counter + 1) top (n, if top then p^" (a"^(str counter)^" as ("^t^"))" else p^", "^t) ts + end + and print_pattern_list counter top (n, p) (t::ts) = + let + val (n, t) = print_pattern false n t + in + print_pattern_list' (counter + 1) top (n, if top then p^" (a"^(str counter)^" as ("^t^"))" else p^" ("^t) ts + end + val c = (case p of PConst (c, _) => c | _ => raise Match) + val (n, pattern) = print_pattern true 0 p + val lazy_vars = the (arity_of c) - the (toplevel_arity_of c) + fun print_tm tm = print_term NONE arity_of toplevel_arity_of n lazy_vars tm + fun print_guard (Guard (a,b)) = "term_eq ("^(print_tm a)^") ("^(print_tm b)^")" + val else_branch = "c"^(str c)^"_"^(str (gnum+1))^(implode (map (fn i => " a"^(str i)) (section (the (arity_of c))))) + fun print_guards t [] = print_tm t + | print_guards t (g::gs) = "if ("^(print_guard g)^")"^(implode (map (fn g => " andalso ("^(print_guard g)^")") gs))^" then ("^(print_tm t)^") else "^else_branch + in + (if null guards then gnum else gnum+1, pattern^" = "^(print_guards t guards)) + end + +fun group_rules rules = + let + fun add_rule (r as (_, PConst (c,_), _)) groups = + let + val rs = (case Inttab.lookup groups c of NONE => [] | SOME rs => rs) + in + Inttab.update (c, r::rs) groups + end + | add_rule _ _ = raise Compile "internal error group_rules" + in + fold_rev add_rule rules Inttab.empty + end + +fun sml_prog name code rules = + let + val buffer = Unsynchronized.ref "" + fun write s = (buffer := (!buffer)^s) + fun writeln s = (write s; write "\n") + fun writelist [] = () + | writelist (s::ss) = (writeln s; writelist ss) + fun str i = string_of_int i + val (inlinetab, rules) = inline_rules rules + val (arity, toplevel_arity, rules) = adjust_rules rules + val rules = group_rules rules + val constants = Inttab.keys arity + fun arity_of c = Inttab.lookup arity c + fun toplevel_arity_of c = Inttab.lookup toplevel_arity c + fun rep_str s n = implode (rep n s) + fun indexed s n = s^(str n) + fun string_of_tuple [] = "" + | string_of_tuple (x::xs) = "("^x^(implode (map (fn s => ", "^s) xs))^")" + fun string_of_args [] = "" + | string_of_args (x::xs) = x^(implode (map (fn s => " "^s) xs)) + fun default_case gnum c = + let + val leftargs = implode (map (indexed " x") (section (the (arity_of c)))) + val rightargs = section (the (arity_of c)) + val strict_args = (case toplevel_arity_of c of NONE => the (arity_of c) | SOME sa => sa) + val xs = map (fn n => if n < strict_args then "x"^(str n) else "x"^(str n)^"()") rightargs + val right = (indexed "C" c)^" "^(string_of_tuple xs) + val message = "(\"unresolved lazy call: " ^ string_of_int c ^ "\")" + val right = if strict_args < the (arity_of c) then "raise AM_SML.Run "^message else right + in + (indexed "c" c)^(if gnum > 0 then "_"^(str gnum) else "")^leftargs^" = "^right + end + + fun eval_rules c = + let + val arity = the (arity_of c) + val strict_arity = (case toplevel_arity_of c of NONE => arity | SOME sa => sa) + fun eval_rule n = + let + val sc = string_of_int c + val left = fold (fn i => fn s => "AbstractMachine.App ("^s^(indexed ", x" i)^")") (section n) ("AbstractMachine.Const "^sc) + fun arg i = + let + val x = indexed "x" i + val x = if i < n then "(eval bounds "^x^")" else x + val x = if i < strict_arity then x else "(fn () => "^x^")" + in + x + end + val right = "c"^sc^" "^(string_of_args (map arg (section arity))) + val right = fold_rev (fn i => fn s => "Abs (fn "^(indexed "x" i)^" => "^s^")") (List.drop (section arity, n)) right + val right = if arity > 0 then right else "C"^sc + in + " | eval bounds ("^left^") = "^right + end + in + map eval_rule (rev (section (arity + 1))) + end + + fun convert_computed_rules (c: int) : string list = + let + val arity = the (arity_of c) + fun eval_rule () = + let + val sc = string_of_int c + val left = fold (fn i => fn s => "AbstractMachine.App ("^s^(indexed ", x" i)^")") (section arity) ("AbstractMachine.Const "^sc) + fun arg i = "(convert_computed "^(indexed "x" i)^")" + val right = "C"^sc^" "^(string_of_tuple (map arg (section arity))) + val right = if arity > 0 then right else "C"^sc + in + " | convert_computed ("^left^") = "^right + end + in + [eval_rule ()] + end + + fun mk_constr_type_args n = if n > 0 then " of Term "^(rep_str " * Term" (n-1)) else "" + val _ = writelist [ + "structure "^name^" = struct", + "", + "datatype Term = Const of int | App of Term * Term | Abs of (Term -> Term)", + " "^(implode (map (fn c => " | C"^(str c)^(mk_constr_type_args (the (arity_of c)))) constants)), + ""] + fun make_constr c argprefix = "(C"^(str c)^" "^(string_of_tuple (map (fn i => argprefix^(str i)) (section (the (arity_of c)))))^")" + fun make_term_eq c = " | term_eq "^(make_constr c "a")^" "^(make_constr c "b")^" = "^ + (case the (arity_of c) of + 0 => "true" + | n => + let + val eqs = map (fn i => "term_eq a"^(str i)^" b"^(str i)) (section n) + val (eq, eqs) = (List.hd eqs, map (fn s => " andalso "^s) (List.tl eqs)) + in + eq^(implode eqs) + end) + val _ = writelist [ + "fun term_eq (Const c1) (Const c2) = (c1 = c2)", + " | term_eq (App (a1,a2)) (App (b1,b2)) = term_eq a1 b1 andalso term_eq a2 b2"] + val _ = writelist (map make_term_eq constants) + val _ = writelist [ + " | term_eq _ _ = false", + "" + ] + val _ = writelist [ + "fun app (Abs a) b = a b", + " | app a b = App (a, b)", + ""] + fun defcase gnum c = (case arity_of c of NONE => [] | SOME a => if a > 0 then [default_case gnum c] else []) + fun writefundecl [] = () + | writefundecl (x::xs) = writelist ((("and "^x)::(map (fn s => " | "^s) xs))) + fun list_group c = (case Inttab.lookup rules c of + NONE => [defcase 0 c] + | SOME rs => + let + val rs = + fold + (fn r => + fn rs => + let + val (gnum, l, rs) = + (case rs of + [] => (0, [], []) + | (gnum, l)::rs => (gnum, l, rs)) + val (gnum', r) = print_rule gnum arity_of toplevel_arity_of r + in + if gnum' = gnum then + (gnum, r::l)::rs + else + let + val args = implode (map (fn i => " a"^(str i)) (section (the (arity_of c)))) + fun gnumc g = if g > 0 then "c"^(str c)^"_"^(str g)^args else "c"^(str c)^args + val s = gnumc (gnum) ^ " = " ^ gnumc (gnum') + in + (gnum', [])::(gnum, s::r::l)::rs + end + end) + rs [] + val rs = (case rs of [] => [(0,defcase 0 c)] | (gnum,l)::rs => (gnum, (defcase gnum c)@l)::rs) + in + rev (map (fn z => rev (snd z)) rs) + end) + val _ = map (fn z => (map writefundecl z; writeln "")) (map list_group constants) + val _ = writelist [ + "fun convert (Const i) = AM_SML.Const i", + " | convert (App (a, b)) = AM_SML.App (convert a, convert b)", + " | convert (Abs _) = raise AM_SML.Run \"no abstraction in result allowed\""] + fun make_convert c = + let + val args = map (indexed "a") (section (the (arity_of c))) + val leftargs = + case args of + [] => "" + | (x::xs) => "("^x^(implode (map (fn s => ", "^s) xs))^")" + val args = map (indexed "convert a") (section (the (arity_of c))) + val right = fold (fn x => fn s => "AM_SML.App ("^s^", "^x^")") args ("AM_SML.Const "^(str c)) + in + " | convert (C"^(str c)^" "^leftargs^") = "^right + end + val _ = writelist (map make_convert constants) + val _ = writelist [ + "", + "fun convert_computed (AbstractMachine.Abs b) = raise AM_SML.Run \"no abstraction in convert_computed allowed\"", + " | convert_computed (AbstractMachine.Var i) = raise AM_SML.Run \"no bound variables in convert_computed allowed\""] + val _ = map (writelist o convert_computed_rules) constants + val _ = writelist [ + " | convert_computed (AbstractMachine.Const c) = Const c", + " | convert_computed (AbstractMachine.App (a, b)) = App (convert_computed a, convert_computed b)", + " | convert_computed (AbstractMachine.Computed a) = raise AM_SML.Run \"no nesting in convert_computed allowed\""] + val _ = writelist [ + "", + "fun eval bounds (AbstractMachine.Abs m) = Abs (fn b => eval (b::bounds) m)", + " | eval bounds (AbstractMachine.Var i) = AM_SML.list_nth (bounds, i)"] + val _ = map (writelist o eval_rules) constants + val _ = writelist [ + " | eval bounds (AbstractMachine.App (a, b)) = app (eval bounds a) (eval bounds b)", + " | eval bounds (AbstractMachine.Const c) = Const c", + " | eval bounds (AbstractMachine.Computed t) = convert_computed t"] + val _ = writelist [ + "", + "fun export term = AM_SML.save_result (\""^code^"\", convert term)", + "", + "val _ = AM_SML.set_compiled_rewriter (fn t => (convert (eval [] t)))", + "", + "end"] + in + (inlinetab, !buffer) + end + +val guid_counter = Unsynchronized.ref 0 +fun get_guid () = + let + val c = !guid_counter + val _ = guid_counter := !guid_counter + 1 + in + string_of_int (Time.toMicroseconds (Time.now ())) ^ string_of_int c + end + + +fun writeTextFile name s = File.write (Path.explode name) s + +fun use_source src = use_text ML_Env.local_context (1, "") false src + +fun compile rules = + let + val guid = get_guid () + val code = Real.toString (random ()) + val name = "AMSML_"^guid + val (inlinetab, source) = sml_prog name code rules + val _ = case !dump_output of NONE => () | SOME p => writeTextFile p source + val _ = compiled_rewriter := NONE + val _ = use_source source + in + case !compiled_rewriter of + NONE => raise Compile "broken link to compiled function" + | SOME compiled_fun => (inlinetab, compiled_fun) + end + +fun run (inlinetab, compiled_fun) t = + let + val _ = if check_freevars 0 t then () else raise Run ("can only compute closed terms") + fun inline (Const c) = (case Inttab.lookup inlinetab c of NONE => Const c | SOME t => t) + | inline (Var i) = Var i + | inline (App (a, b)) = App (inline a, inline b) + | inline (Abs m) = Abs (inline m) + | inline (Computed t) = Computed t + in + compiled_fun (beta (inline t)) + end + +end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/Compute_Oracle/compute.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/Compute_Oracle/compute.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,653 @@ +(* Title: HOL/Matrix/Compute_Oracle/compute.ML + Author: Steven Obua +*) + +signature COMPUTE = sig + + type computer + type theorem + type naming = int -> string + + datatype machine = BARRAS | BARRAS_COMPILED | HASKELL | SML + + (* Functions designated with a ! in front of them actually update the computer parameter *) + + exception Make of string + val make : machine -> theory -> thm list -> computer + val make_with_cache : machine -> theory -> term list -> thm list -> computer + val theory_of : computer -> theory + val hyps_of : computer -> term list + val shyps_of : computer -> sort list + (* ! *) val update : computer -> thm list -> unit + (* ! *) val update_with_cache : computer -> term list -> thm list -> unit + + (* ! *) val set_naming : computer -> naming -> unit + val naming_of : computer -> naming + + exception Compute of string + val simplify : computer -> theorem -> thm + val rewrite : computer -> cterm -> thm + + val make_theorem : computer -> thm -> string list -> theorem + (* ! *) val instantiate : computer -> (string * cterm) list -> theorem -> theorem + (* ! *) val evaluate_prem : computer -> int -> theorem -> theorem + (* ! *) val modus_ponens : computer -> int -> thm -> theorem -> theorem + +end + +structure Compute :> COMPUTE = struct + +open Report; + +datatype machine = BARRAS | BARRAS_COMPILED | HASKELL | SML + +(* Terms are mapped to integer codes *) +structure Encode :> +sig + type encoding + val empty : encoding + val insert : term -> encoding -> int * encoding + val lookup_code : term -> encoding -> int option + val lookup_term : int -> encoding -> term option + val remove_code : int -> encoding -> encoding + val remove_term : term -> encoding -> encoding +end += +struct + +type encoding = int * (int Termtab.table) * (term Inttab.table) + +val empty = (0, Termtab.empty, Inttab.empty) + +fun insert t (e as (count, term2int, int2term)) = + (case Termtab.lookup term2int t of + NONE => (count, (count+1, Termtab.update_new (t, count) term2int, Inttab.update_new (count, t) int2term)) + | SOME code => (code, e)) + +fun lookup_code t (_, term2int, _) = Termtab.lookup term2int t + +fun lookup_term c (_, _, int2term) = Inttab.lookup int2term c + +fun remove_code c (e as (count, term2int, int2term)) = + (case lookup_term c e of NONE => e | SOME t => (count, Termtab.delete t term2int, Inttab.delete c int2term)) + +fun remove_term t (e as (count, term2int, int2term)) = + (case lookup_code t e of NONE => e | SOME c => (count, Termtab.delete t term2int, Inttab.delete c int2term)) + +end + +exception Make of string; +exception Compute of string; + +local + fun make_constant t encoding = + let + val (code, encoding) = Encode.insert t encoding + in + (encoding, AbstractMachine.Const code) + end +in + +fun remove_types encoding t = + case t of + Var _ => make_constant t encoding + | Free _ => make_constant t encoding + | Const _ => make_constant t encoding + | Abs (_, _, t') => + let val (encoding, t'') = remove_types encoding t' in + (encoding, AbstractMachine.Abs t'') + end + | a $ b => + let + val (encoding, a) = remove_types encoding a + val (encoding, b) = remove_types encoding b + in + (encoding, AbstractMachine.App (a,b)) + end + | Bound b => (encoding, AbstractMachine.Var b) +end + +local + fun type_of (Free (_, ty)) = ty + | type_of (Const (_, ty)) = ty + | type_of (Var (_, ty)) = ty + | type_of _ = raise Fail "infer_types: type_of error" +in +fun infer_types naming encoding = + let + fun infer_types _ bounds _ (AbstractMachine.Var v) = (Bound v, nth bounds v) + | infer_types _ bounds _ (AbstractMachine.Const code) = + let + val c = the (Encode.lookup_term code encoding) + in + (c, type_of c) + end + | infer_types level bounds _ (AbstractMachine.App (a, b)) = + let + val (a, aty) = infer_types level bounds NONE a + val (adom, arange) = + case aty of + Type ("fun", [dom, range]) => (dom, range) + | _ => raise Fail "infer_types: function type expected" + val (b, _) = infer_types level bounds (SOME adom) b + in + (a $ b, arange) + end + | infer_types level bounds (SOME (ty as Type ("fun", [dom, range]))) (AbstractMachine.Abs m) = + let + val (m, _) = infer_types (level+1) (dom::bounds) (SOME range) m + in + (Abs (naming level, dom, m), ty) + end + | infer_types _ _ NONE (AbstractMachine.Abs _) = + raise Fail "infer_types: cannot infer type of abstraction" + + fun infer ty term = + let + val (term', _) = infer_types 0 [] (SOME ty) term + in + term' + end + in + infer + end +end + +datatype prog = + ProgBarras of AM_Interpreter.program + | ProgBarrasC of AM_Compiler.program + | ProgHaskell of AM_GHC.program + | ProgSML of AM_SML.program + +fun machine_of_prog (ProgBarras _) = BARRAS + | machine_of_prog (ProgBarrasC _) = BARRAS_COMPILED + | machine_of_prog (ProgHaskell _) = HASKELL + | machine_of_prog (ProgSML _) = SML + +type naming = int -> string + +fun default_naming i = "v_" ^ string_of_int i + +datatype computer = Computer of + (theory_ref * Encode.encoding * term list * unit Sorttab.table * prog * unit Unsynchronized.ref * naming) + option Unsynchronized.ref + +fun theory_of (Computer (Unsynchronized.ref (SOME (rthy,_,_,_,_,_,_)))) = Theory.deref rthy +fun hyps_of (Computer (Unsynchronized.ref (SOME (_,_,hyps,_,_,_,_)))) = hyps +fun shyps_of (Computer (Unsynchronized.ref (SOME (_,_,_,shyptable,_,_,_)))) = Sorttab.keys (shyptable) +fun shyptab_of (Computer (Unsynchronized.ref (SOME (_,_,_,shyptable,_,_,_)))) = shyptable +fun stamp_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,_,stamp,_)))) = stamp +fun prog_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,prog,_,_)))) = prog +fun encoding_of (Computer (Unsynchronized.ref (SOME (_,encoding,_,_,_,_,_)))) = encoding +fun set_encoding (Computer (r as Unsynchronized.ref (SOME (p1,_,p2,p3,p4,p5,p6)))) encoding' = + (r := SOME (p1,encoding',p2,p3,p4,p5,p6)) +fun naming_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,_,_,n)))) = n +fun set_naming (Computer (r as Unsynchronized.ref (SOME (p1,p2,p3,p4,p5,p6,_)))) naming'= + (r := SOME (p1,p2,p3,p4,p5,p6,naming')) + +fun ref_of (Computer r) = r + +datatype cthm = ComputeThm of term list * sort list * term + +fun thm2cthm th = + let + val {hyps, prop, tpairs, shyps, ...} = Thm.rep_thm th + val _ = if not (null tpairs) then raise Make "theorems may not contain tpairs" else () + in + ComputeThm (hyps, shyps, prop) + end + +fun make_internal machine thy stamp encoding cache_pattern_terms raw_ths = + let + fun transfer (x:thm) = Thm.transfer thy x + val ths = map (thm2cthm o Thm.strip_shyps o transfer) raw_ths + + fun make_pattern encoding n vars (AbstractMachine.Abs _) = + raise (Make "no lambda abstractions allowed in pattern") + | make_pattern encoding n vars (AbstractMachine.Var _) = + raise (Make "no bound variables allowed in pattern") + | make_pattern encoding n vars (AbstractMachine.Const code) = + (case the (Encode.lookup_term code encoding) of + Var _ => ((n+1, Inttab.update_new (code, n) vars, AbstractMachine.PVar) + handle Inttab.DUP _ => raise (Make "no duplicate variable in pattern allowed")) + | _ => (n, vars, AbstractMachine.PConst (code, []))) + | make_pattern encoding n vars (AbstractMachine.App (a, b)) = + let + val (n, vars, pa) = make_pattern encoding n vars a + val (n, vars, pb) = make_pattern encoding n vars b + in + case pa of + AbstractMachine.PVar => + raise (Make "patterns may not start with a variable") + | AbstractMachine.PConst (c, args) => + (n, vars, AbstractMachine.PConst (c, args@[pb])) + end + + fun thm2rule (encoding, hyptable, shyptable) th = + let + val (ComputeThm (hyps, shyps, prop)) = th + val hyptable = fold (fn h => Termtab.update (h, ())) hyps hyptable + val shyptable = fold (fn sh => Sorttab.update (sh, ())) shyps shyptable + val (prems, prop) = (Logic.strip_imp_prems prop, Logic.strip_imp_concl prop) + val (a, b) = Logic.dest_equals prop + handle TERM _ => raise (Make "theorems must be meta-level equations (with optional guards)") + val a = Envir.eta_contract a + val b = Envir.eta_contract b + val prems = map Envir.eta_contract prems + + val (encoding, left) = remove_types encoding a + val (encoding, right) = remove_types encoding b + fun remove_types_of_guard encoding g = + (let + val (t1, t2) = Logic.dest_equals g + val (encoding, t1) = remove_types encoding t1 + val (encoding, t2) = remove_types encoding t2 + in + (encoding, AbstractMachine.Guard (t1, t2)) + end handle TERM _ => raise (Make "guards must be meta-level equations")) + val (encoding, prems) = fold_rev (fn p => fn (encoding, ps) => let val (e, p) = remove_types_of_guard encoding p in (e, p::ps) end) prems (encoding, []) + + (* Principally, a check should be made here to see if the (meta-) hyps contain any of the variables of the rule. + As it is, all variables of the rule are schematic, and there are no schematic variables in meta-hyps, therefore + this check can be left out. *) + + val (vcount, vars, pattern) = make_pattern encoding 0 Inttab.empty left + val _ = (case pattern of + AbstractMachine.PVar => + raise (Make "patterns may not start with a variable") + | _ => ()) + + (* finally, provide a function for renaming the + pattern bound variables on the right hand side *) + + fun rename level vars (var as AbstractMachine.Var _) = var + | rename level vars (c as AbstractMachine.Const code) = + (case Inttab.lookup vars code of + NONE => c + | SOME n => AbstractMachine.Var (vcount-n-1+level)) + | rename level vars (AbstractMachine.App (a, b)) = + AbstractMachine.App (rename level vars a, rename level vars b) + | rename level vars (AbstractMachine.Abs m) = + AbstractMachine.Abs (rename (level+1) vars m) + + fun rename_guard (AbstractMachine.Guard (a,b)) = + AbstractMachine.Guard (rename 0 vars a, rename 0 vars b) + in + ((encoding, hyptable, shyptable), (map rename_guard prems, pattern, rename 0 vars right)) + end + + val ((encoding, hyptable, shyptable), rules) = + fold_rev (fn th => fn (encoding_hyptable, rules) => + let + val (encoding_hyptable, rule) = thm2rule encoding_hyptable th + in (encoding_hyptable, rule::rules) end) + ths ((encoding, Termtab.empty, Sorttab.empty), []) + + fun make_cache_pattern t (encoding, cache_patterns) = + let + val (encoding, a) = remove_types encoding t + val (_,_,p) = make_pattern encoding 0 Inttab.empty a + in + (encoding, p::cache_patterns) + end + + val (encoding, _) = fold_rev make_cache_pattern cache_pattern_terms (encoding, []) + + val prog = + case machine of + BARRAS => ProgBarras (AM_Interpreter.compile rules) + | BARRAS_COMPILED => ProgBarrasC (AM_Compiler.compile rules) + | HASKELL => ProgHaskell (AM_GHC.compile rules) + | SML => ProgSML (AM_SML.compile rules) + + fun has_witness s = not (null (Sign.witness_sorts thy [] [s])) + + val shyptable = fold Sorttab.delete (filter has_witness (Sorttab.keys (shyptable))) shyptable + + in (Theory.check_thy thy, encoding, Termtab.keys hyptable, shyptable, prog, stamp, default_naming) end + +fun make_with_cache machine thy cache_patterns raw_thms = + Computer (Unsynchronized.ref (SOME (make_internal machine thy (Unsynchronized.ref ()) Encode.empty cache_patterns raw_thms))) + +fun make machine thy raw_thms = make_with_cache machine thy [] raw_thms + +fun update_with_cache computer cache_patterns raw_thms = + let + val c = make_internal (machine_of_prog (prog_of computer)) (theory_of computer) (stamp_of computer) + (encoding_of computer) cache_patterns raw_thms + val _ = (ref_of computer) := SOME c + in + () + end + +fun update computer raw_thms = update_with_cache computer [] raw_thms + +fun runprog (ProgBarras p) = AM_Interpreter.run p + | runprog (ProgBarrasC p) = AM_Compiler.run p + | runprog (ProgHaskell p) = AM_GHC.run p + | runprog (ProgSML p) = AM_SML.run p + +(* ------------------------------------------------------------------------------------- *) +(* An oracle for exporting theorems; must only be accessible from inside this structure! *) +(* ------------------------------------------------------------------------------------- *) + +fun merge_hyps hyps1 hyps2 = +let + fun add hyps tab = fold (fn h => fn tab => Termtab.update (h, ()) tab) hyps tab +in + Termtab.keys (add hyps2 (add hyps1 Termtab.empty)) +end + +fun add_shyps shyps tab = fold (fn h => fn tab => Sorttab.update (h, ()) tab) shyps tab + +fun merge_shyps shyps1 shyps2 = Sorttab.keys (add_shyps shyps2 (add_shyps shyps1 Sorttab.empty)) + +val (_, export_oracle) = Context.>>> (Context.map_theory_result + (Thm.add_oracle (@{binding compute}, fn (thy, hyps, shyps, prop) => + let + val shyptab = add_shyps shyps Sorttab.empty + fun delete s shyptab = Sorttab.delete s shyptab handle Sorttab.UNDEF _ => shyptab + fun delete_term t shyptab = fold delete (Sorts.insert_term t []) shyptab + fun has_witness s = not (null (Sign.witness_sorts thy [] [s])) + val shyptab = fold Sorttab.delete (filter has_witness (Sorttab.keys (shyptab))) shyptab + val shyps = if Sorttab.is_empty shyptab then [] else Sorttab.keys (fold delete_term (prop::hyps) shyptab) + val _ = + if not (null shyps) then + raise Compute ("dangling sort hypotheses: " ^ + commas (map (Syntax.string_of_sort_global thy) shyps)) + else () + in + Thm.cterm_of thy (fold_rev (fn hyp => fn p => Logic.mk_implies (hyp, p)) hyps prop) + end))); + +fun export_thm thy hyps shyps prop = + let + val th = export_oracle (thy, hyps, shyps, prop) + val hyps = map (fn h => Thm.assume (cterm_of thy h)) hyps + in + fold (fn h => fn p => Thm.implies_elim p h) hyps th + end + +(* --------- Rewrite ----------- *) + +fun rewrite computer ct = + let + val thy = Thm.theory_of_cterm ct + val {t=t',T=ty,...} = rep_cterm ct + val _ = Theory.assert_super (theory_of computer) thy + val naming = naming_of computer + val (encoding, t) = remove_types (encoding_of computer) t' + val t = runprog (prog_of computer) t + val t = infer_types naming encoding ty t + val eq = Logic.mk_equals (t', t) + in + export_thm thy (hyps_of computer) (Sorttab.keys (shyptab_of computer)) eq + end + +(* --------- Simplify ------------ *) + +datatype prem = EqPrem of AbstractMachine.term * AbstractMachine.term * Term.typ * int + | Prem of AbstractMachine.term +datatype theorem = Theorem of theory_ref * unit Unsynchronized.ref * (int * typ) Symtab.table * (AbstractMachine.term option) Inttab.table + * prem list * AbstractMachine.term * term list * sort list + + +exception ParamSimplify of computer * theorem + +fun make_theorem computer th vars = +let + val _ = Theory.assert_super (theory_of computer) (theory_of_thm th) + + val (ComputeThm (hyps, shyps, prop)) = thm2cthm th + + val encoding = encoding_of computer + + (* variables in the theorem are identified upfront *) + fun collect_vars (Abs (_, _, t)) tab = collect_vars t tab + | collect_vars (a $ b) tab = collect_vars b (collect_vars a tab) + | collect_vars (Const _) tab = tab + | collect_vars (Free _) tab = tab + | collect_vars (Var ((s, i), ty)) tab = + if List.find (fn x => x=s) vars = NONE then + tab + else + (case Symtab.lookup tab s of + SOME ((s',i'),ty') => + if s' <> s orelse i' <> i orelse ty <> ty' then + raise Compute ("make_theorem: variable name '"^s^"' is not unique") + else + tab + | NONE => Symtab.update (s, ((s, i), ty)) tab) + val vartab = collect_vars prop Symtab.empty + fun encodevar (s, t as (_, ty)) (encoding, tab) = + let + val (x, encoding) = Encode.insert (Var t) encoding + in + (encoding, Symtab.update (s, (x, ty)) tab) + end + val (encoding, vartab) = Symtab.fold encodevar vartab (encoding, Symtab.empty) + val varsubst = Inttab.make (map (fn (_, (x, _)) => (x, NONE)) (Symtab.dest vartab)) + + (* make the premises and the conclusion *) + fun mk_prem encoding t = + (let + val (a, b) = Logic.dest_equals t + val ty = type_of a + val (encoding, a) = remove_types encoding a + val (encoding, b) = remove_types encoding b + val (eq, encoding) = Encode.insert (Const ("==", ty --> ty --> @{typ "prop"})) encoding + in + (encoding, EqPrem (a, b, ty, eq)) + end handle TERM _ => let val (encoding, t) = remove_types encoding t in (encoding, Prem t) end) + val (encoding, prems) = + (fold_rev (fn t => fn (encoding, l) => + case mk_prem encoding t of + (encoding, t) => (encoding, t::l)) (Logic.strip_imp_prems prop) (encoding, [])) + val (encoding, concl) = remove_types encoding (Logic.strip_imp_concl prop) + val _ = set_encoding computer encoding +in + Theorem (Theory.check_thy (theory_of_thm th), stamp_of computer, vartab, varsubst, + prems, concl, hyps, shyps) +end + +fun theory_of_theorem (Theorem (rthy,_,_,_,_,_,_,_)) = Theory.deref rthy +fun update_theory thy (Theorem (_,p0,p1,p2,p3,p4,p5,p6)) = + Theorem (Theory.check_thy thy,p0,p1,p2,p3,p4,p5,p6) +fun stamp_of_theorem (Theorem (_,s, _, _, _, _, _, _)) = s +fun vartab_of_theorem (Theorem (_,_,vt,_,_,_,_,_)) = vt +fun varsubst_of_theorem (Theorem (_,_,_,vs,_,_,_,_)) = vs +fun update_varsubst vs (Theorem (p0,p1,p2,_,p3,p4,p5,p6)) = Theorem (p0,p1,p2,vs,p3,p4,p5,p6) +fun prems_of_theorem (Theorem (_,_,_,_,prems,_,_,_)) = prems +fun update_prems prems (Theorem (p0,p1,p2,p3,_,p4,p5,p6)) = Theorem (p0,p1,p2,p3,prems,p4,p5,p6) +fun concl_of_theorem (Theorem (_,_,_,_,_,concl,_,_)) = concl +fun hyps_of_theorem (Theorem (_,_,_,_,_,_,hyps,_)) = hyps +fun update_hyps hyps (Theorem (p0,p1,p2,p3,p4,p5,_,p6)) = Theorem (p0,p1,p2,p3,p4,p5,hyps,p6) +fun shyps_of_theorem (Theorem (_,_,_,_,_,_,_,shyps)) = shyps +fun update_shyps shyps (Theorem (p0,p1,p2,p3,p4,p5,p6,_)) = Theorem (p0,p1,p2,p3,p4,p5,p6,shyps) + +fun check_compatible computer th s = + if stamp_of computer <> stamp_of_theorem th then + raise Compute (s^": computer and theorem are incompatible") + else () + +fun instantiate computer insts th = +let + val _ = check_compatible computer th + + val thy = theory_of computer + + val vartab = vartab_of_theorem th + + fun rewrite computer t = + let + val (encoding, t) = remove_types (encoding_of computer) t + val t = runprog (prog_of computer) t + val _ = set_encoding computer encoding + in + t + end + + fun assert_varfree vs t = + if AbstractMachine.forall_consts (fn x => Inttab.lookup vs x = NONE) t then + () + else + raise Compute "instantiate: assert_varfree failed" + + fun assert_closed t = + if AbstractMachine.closed t then + () + else + raise Compute "instantiate: not a closed term" + + fun compute_inst (s, ct) vs = + let + val _ = Theory.assert_super (theory_of_cterm ct) thy + val ty = typ_of (ctyp_of_term ct) + in + (case Symtab.lookup vartab s of + NONE => raise Compute ("instantiate: variable '"^s^"' not found in theorem") + | SOME (x, ty') => + (case Inttab.lookup vs x of + SOME (SOME _) => raise Compute ("instantiate: variable '"^s^"' has already been instantiated") + | SOME NONE => + if ty <> ty' then + raise Compute ("instantiate: wrong type for variable '"^s^"'") + else + let + val t = rewrite computer (term_of ct) + val _ = assert_varfree vs t + val _ = assert_closed t + in + Inttab.update (x, SOME t) vs + end + | NONE => raise Compute "instantiate: internal error")) + end + + val vs = fold compute_inst insts (varsubst_of_theorem th) +in + update_varsubst vs th +end + +fun match_aterms subst = + let + exception no_match + open AbstractMachine + fun match subst (b as (Const c)) a = + if a = b then subst + else + (case Inttab.lookup subst c of + SOME (SOME a') => if a=a' then subst else raise no_match + | SOME NONE => if AbstractMachine.closed a then + Inttab.update (c, SOME a) subst + else raise no_match + | NONE => raise no_match) + | match subst (b as (Var _)) a = if a=b then subst else raise no_match + | match subst (App (u, v)) (App (u', v')) = match (match subst u u') v v' + | match subst (Abs u) (Abs u') = match subst u u' + | match subst _ _ = raise no_match + in + fn b => fn a => (SOME (match subst b a) handle no_match => NONE) + end + +fun apply_subst vars_allowed subst = + let + open AbstractMachine + fun app (t as (Const c)) = + (case Inttab.lookup subst c of + NONE => t + | SOME (SOME t) => Computed t + | SOME NONE => if vars_allowed then t else raise Compute "apply_subst: no vars allowed") + | app (t as (Var _)) = t + | app (App (u, v)) = App (app u, app v) + | app (Abs m) = Abs (app m) + in + app + end + +fun splicein n l L = List.take (L, n) @ l @ List.drop (L, n+1) + +fun evaluate_prem computer prem_no th = +let + val _ = check_compatible computer th + val prems = prems_of_theorem th + val varsubst = varsubst_of_theorem th + fun run vars_allowed t = + runprog (prog_of computer) (apply_subst vars_allowed varsubst t) +in + case nth prems prem_no of + Prem _ => raise Compute "evaluate_prem: no equality premise" + | EqPrem (a, b, ty, _) => + let + val a' = run false a + val b' = run true b + in + case match_aterms varsubst b' a' of + NONE => + let + fun mk s = Syntax.string_of_term_global Pure.thy + (infer_types (naming_of computer) (encoding_of computer) ty s) + val left = "computed left side: "^(mk a') + val right = "computed right side: "^(mk b') + in + raise Compute ("evaluate_prem: cannot assign computed left to right hand side\n"^left^"\n"^right^"\n") + end + | SOME varsubst => + update_prems (splicein prem_no [] prems) (update_varsubst varsubst th) + end +end + +fun prem2term (Prem t) = t + | prem2term (EqPrem (a,b,_,eq)) = + AbstractMachine.App (AbstractMachine.App (AbstractMachine.Const eq, a), b) + +fun modus_ponens computer prem_no th' th = +let + val _ = check_compatible computer th + val thy = + let + val thy1 = theory_of_theorem th + val thy2 = theory_of_thm th' + in + if Theory.subthy (thy1, thy2) then thy2 + else if Theory.subthy (thy2, thy1) then thy1 else + raise Compute "modus_ponens: theorems are not compatible with each other" + end + val th' = make_theorem computer th' [] + val varsubst = varsubst_of_theorem th + fun run vars_allowed t = + runprog (prog_of computer) (apply_subst vars_allowed varsubst t) + val prems = prems_of_theorem th + val prem = run true (prem2term (nth prems prem_no)) + val concl = run false (concl_of_theorem th') +in + case match_aterms varsubst prem concl of + NONE => raise Compute "modus_ponens: conclusion does not match premise" + | SOME varsubst => + let + val th = update_varsubst varsubst th + val th = update_prems (splicein prem_no (prems_of_theorem th') prems) th + val th = update_hyps (merge_hyps (hyps_of_theorem th) (hyps_of_theorem th')) th + val th = update_shyps (merge_shyps (shyps_of_theorem th) (shyps_of_theorem th')) th + in + update_theory thy th + end +end + +fun simplify computer th = +let + val _ = check_compatible computer th + val varsubst = varsubst_of_theorem th + val encoding = encoding_of computer + val naming = naming_of computer + fun infer t = infer_types naming encoding @{typ "prop"} t + fun run t = infer (runprog (prog_of computer) (apply_subst true varsubst t)) + fun runprem p = run (prem2term p) + val prop = Logic.list_implies (map runprem (prems_of_theorem th), run (concl_of_theorem th)) + val hyps = merge_hyps (hyps_of computer) (hyps_of_theorem th) + val shyps = merge_shyps (shyps_of_theorem th) (Sorttab.keys (shyptab_of computer)) +in + export_thm (theory_of_theorem th) hyps shyps prop +end + +end + diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/Compute_Oracle/linker.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/Compute_Oracle/linker.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,470 @@ +(* Title: HOL/Matrix/Compute_Oracle/linker.ML + Author: Steven Obua + +This module solves the problem that the computing oracle does not +instantiate polymorphic rules. By going through the PCompute +interface, all possible instantiations are resolved by compiling new +programs, if necessary. The obvious disadvantage of this approach is +that in the worst case for each new term to be rewritten, a new +program may be compiled. +*) + +(* + Given constants/frees c_1::t_1, c_2::t_2, ...., c_n::t_n, + and constants/frees d_1::d_1, d_2::s_2, ..., d_m::s_m + + Find all substitutions S such that + a) the domain of S is tvars (t_1, ..., t_n) + b) there are indices i_1, ..., i_k, and j_1, ..., j_k with + 1. S (c_i_1::t_i_1) = d_j_1::s_j_1, ..., S (c_i_k::t_i_k) = d_j_k::s_j_k + 2. tvars (t_i_1, ..., t_i_k) = tvars (t_1, ..., t_n) +*) +signature LINKER = +sig + exception Link of string + + datatype constant = Constant of bool * string * typ + val constant_of : term -> constant + + type instances + type subst = Type.tyenv + + val empty : constant list -> instances + val typ_of_constant : constant -> typ + val add_instances : theory -> instances -> constant list -> subst list * instances + val substs_of : instances -> subst list + val is_polymorphic : constant -> bool + val distinct_constants : constant list -> constant list + val collect_consts : term list -> constant list +end + +structure Linker : LINKER = struct + +exception Link of string; + +type subst = Type.tyenv + +datatype constant = Constant of bool * string * typ +fun constant_of (Const (name, ty)) = Constant (false, name, ty) + | constant_of (Free (name, ty)) = Constant (true, name, ty) + | constant_of _ = raise Link "constant_of" + +fun bool_ord (x,y) = if x then (if y then EQUAL else GREATER) else (if y then LESS else EQUAL) +fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) Term_Ord.typ_ord) (((x1,x2),x3), ((y1,y2),y3)) +fun constant_modty_ord (Constant (x1,x2,_), Constant (y1,y2,_)) = (prod_ord bool_ord fast_string_ord) ((x1,x2), (y1,y2)) + + +structure Consttab = Table(type key = constant val ord = constant_ord); +structure ConsttabModTy = Table(type key = constant val ord = constant_modty_ord); + +fun typ_of_constant (Constant (_, _, ty)) = ty + +val empty_subst = (Vartab.empty : Type.tyenv) + +fun merge_subst (A:Type.tyenv) (B:Type.tyenv) = + SOME (Vartab.fold (fn (v, t) => + fn tab => + (case Vartab.lookup tab v of + NONE => Vartab.update (v, t) tab + | SOME t' => if t = t' then tab else raise Type.TYPE_MATCH)) A B) + handle Type.TYPE_MATCH => NONE + +fun subst_ord (A:Type.tyenv, B:Type.tyenv) = + (list_ord (prod_ord Term_Ord.fast_indexname_ord (prod_ord Term_Ord.sort_ord Term_Ord.typ_ord))) (Vartab.dest A, Vartab.dest B) + +structure Substtab = Table(type key = Type.tyenv val ord = subst_ord); + +fun substtab_union c = Substtab.fold Substtab.update c +fun substtab_unions [] = Substtab.empty + | substtab_unions [c] = c + | substtab_unions (c::cs) = substtab_union c (substtab_unions cs) + +datatype instances = Instances of unit ConsttabModTy.table * Type.tyenv Consttab.table Consttab.table * constant list list * unit Substtab.table + +fun is_polymorphic (Constant (_, _, ty)) = not (null (Term.add_tvarsT ty [])) + +fun distinct_constants cs = + Consttab.keys (fold (fn c => Consttab.update (c, ())) cs Consttab.empty) + +fun empty cs = + let + val cs = distinct_constants (filter is_polymorphic cs) + val old_cs = cs +(* fun collect_tvars ty tab = fold (fn v => fn tab => Typtab.update (TVar v, ()) tab) (Misc_Legacy.typ_tvars ty) tab + val tvars_count = length (Typtab.keys (fold (fn c => fn tab => collect_tvars (typ_of_constant c) tab) cs Typtab.empty)) + fun tvars_of ty = collect_tvars ty Typtab.empty + val cs = map (fn c => (c, tvars_of (typ_of_constant c))) cs + + fun tyunion A B = + Typtab.fold + (fn (v,()) => fn tab => Typtab.update (v, case Typtab.lookup tab v of NONE => 1 | SOME n => n+1) tab) + A B + + fun is_essential A B = + Typtab.fold + (fn (v, ()) => fn essential => essential orelse (case Typtab.lookup B v of NONE => raise Link "is_essential" | SOME n => n=1)) + A false + + fun add_minimal (c', tvs') (tvs, cs) = + let + val tvs = tyunion tvs' tvs + val cs = (c', tvs')::cs + in + if forall (fn (c',tvs') => is_essential tvs' tvs) cs then + SOME (tvs, cs) + else + NONE + end + + fun is_spanning (tvs, _) = (length (Typtab.keys tvs) = tvars_count) + + fun generate_minimal_subsets subsets [] = subsets + | generate_minimal_subsets subsets (c::cs) = + let + val subsets' = map_filter (add_minimal c) subsets + in + generate_minimal_subsets (subsets@subsets') cs + end*) + + val minimal_subsets = [old_cs] (*map (fn (tvs, cs) => map fst cs) (filter is_spanning (generate_minimal_subsets [(Typtab.empty, [])] cs))*) + + val constants = Consttab.keys (fold (fold (fn c => Consttab.update (c, ()))) minimal_subsets Consttab.empty) + + in + Instances ( + fold (fn c => fn tab => ConsttabModTy.update (c, ()) tab) constants ConsttabModTy.empty, + Consttab.make (map (fn c => (c, Consttab.empty : Type.tyenv Consttab.table)) constants), + minimal_subsets, Substtab.empty) + end + +local +fun calc ctab substtab [] = substtab + | calc ctab substtab (c::cs) = + let + val csubsts = map snd (Consttab.dest (the (Consttab.lookup ctab c))) + fun merge_substs substtab subst = + Substtab.fold (fn (s,_) => + fn tab => + (case merge_subst subst s of NONE => tab | SOME s => Substtab.update (s, ()) tab)) + substtab Substtab.empty + val substtab = substtab_unions (map (merge_substs substtab) csubsts) + in + calc ctab substtab cs + end +in +fun calc_substs ctab (cs:constant list) = calc ctab (Substtab.update (empty_subst, ()) Substtab.empty) cs +end + +fun add_instances thy (Instances (cfilter, ctab,minsets,substs)) cs = + let +(* val _ = writeln (makestring ("add_instances: ", length_cs, length cs, length (Consttab.keys ctab)))*) + fun calc_instantiations (constant as Constant (free, name, ty)) instantiations = + Consttab.fold (fn (constant' as Constant (free', name', ty'), insttab) => + fn instantiations => + if free <> free' orelse name <> name' then + instantiations + else case Consttab.lookup insttab constant of + SOME _ => instantiations + | NONE => ((constant', (constant, Sign.typ_match thy (ty', ty) empty_subst))::instantiations + handle Type.TYPE_MATCH => instantiations)) + ctab instantiations + val instantiations = fold calc_instantiations cs [] + (*val _ = writeln ("instantiations = "^(makestring (length instantiations)))*) + fun update_ctab (constant', entry) ctab = + (case Consttab.lookup ctab constant' of + NONE => raise Link "internal error: update_ctab" + | SOME tab => Consttab.update (constant', Consttab.update entry tab) ctab) + val ctab = fold update_ctab instantiations ctab + val new_substs = fold (fn minset => fn substs => substtab_union (calc_substs ctab minset) substs) + minsets Substtab.empty + val (added_substs, substs) = + Substtab.fold (fn (ns, _) => + fn (added, substtab) => + (case Substtab.lookup substs ns of + NONE => (ns::added, Substtab.update (ns, ()) substtab) + | SOME () => (added, substtab))) + new_substs ([], substs) + in + (added_substs, Instances (cfilter, ctab, minsets, substs)) + end + +fun substs_of (Instances (_,_,_,substs)) = Substtab.keys substs + + +local + +fun collect (Var _) tab = tab + | collect (Bound _) tab = tab + | collect (a $ b) tab = collect b (collect a tab) + | collect (Abs (_, _, body)) tab = collect body tab + | collect t tab = Consttab.update (constant_of t, ()) tab + +in + fun collect_consts tms = Consttab.keys (fold collect tms Consttab.empty) +end + +end + +signature PCOMPUTE = +sig + type pcomputer + + val make : Compute.machine -> theory -> thm list -> Linker.constant list -> pcomputer + val make_with_cache : Compute.machine -> theory -> term list -> thm list -> Linker.constant list -> pcomputer + + val add_instances : pcomputer -> Linker.constant list -> bool + val add_instances' : pcomputer -> term list -> bool + + val rewrite : pcomputer -> cterm list -> thm list + val simplify : pcomputer -> Compute.theorem -> thm + + val make_theorem : pcomputer -> thm -> string list -> Compute.theorem + val instantiate : pcomputer -> (string * cterm) list -> Compute.theorem -> Compute.theorem + val evaluate_prem : pcomputer -> int -> Compute.theorem -> Compute.theorem + val modus_ponens : pcomputer -> int -> thm -> Compute.theorem -> Compute.theorem + +end + +structure PCompute : PCOMPUTE = struct + +exception PCompute of string + +datatype theorem = MonoThm of thm | PolyThm of thm * Linker.instances * thm list +datatype pattern = MonoPattern of term | PolyPattern of term * Linker.instances * term list + +datatype pcomputer = + PComputer of theory_ref * Compute.computer * theorem list Unsynchronized.ref * + pattern list Unsynchronized.ref + +(*fun collect_consts (Var x) = [] + | collect_consts (Bound _) = [] + | collect_consts (a $ b) = (collect_consts a)@(collect_consts b) + | collect_consts (Abs (_, _, body)) = collect_consts body + | collect_consts t = [Linker.constant_of t]*) + +fun computer_of (PComputer (_,computer,_,_)) = computer + +fun collect_consts_of_thm th = + let + val th = prop_of th + val (prems, th) = (Logic.strip_imp_prems th, Logic.strip_imp_concl th) + val (left, right) = Logic.dest_equals th + in + (Linker.collect_consts [left], Linker.collect_consts (right::prems)) + end + +fun create_theorem th = +let + val (left, right) = collect_consts_of_thm th + val polycs = filter Linker.is_polymorphic left + val tytab = fold (fn p => fn tab => fold (fn n => fn tab => Typtab.update (TVar n, ()) tab) (Misc_Legacy.typ_tvars (Linker.typ_of_constant p)) tab) polycs Typtab.empty + fun check_const (c::cs) cs' = + let + val tvars = Misc_Legacy.typ_tvars (Linker.typ_of_constant c) + val wrong = fold (fn n => fn wrong => wrong orelse is_none (Typtab.lookup tytab (TVar n))) tvars false + in + if wrong then raise PCompute "right hand side of theorem contains type variables which do not occur on the left hand side" + else + if null (tvars) then + check_const cs (c::cs') + else + check_const cs cs' + end + | check_const [] cs' = cs' + val monocs = check_const right [] +in + if null (polycs) then + (monocs, MonoThm th) + else + (monocs, PolyThm (th, Linker.empty polycs, [])) +end + +fun create_pattern pat = +let + val cs = Linker.collect_consts [pat] + val polycs = filter Linker.is_polymorphic cs +in + if null (polycs) then + MonoPattern pat + else + PolyPattern (pat, Linker.empty polycs, []) +end + +fun create_computer machine thy pats ths = + let + fun add (MonoThm th) ths = th::ths + | add (PolyThm (_, _, ths')) ths = ths'@ths + fun addpat (MonoPattern p) pats = p::pats + | addpat (PolyPattern (_, _, ps)) pats = ps@pats + val ths = fold_rev add ths [] + val pats = fold_rev addpat pats [] + in + Compute.make_with_cache machine thy pats ths + end + +fun update_computer computer pats ths = + let + fun add (MonoThm th) ths = th::ths + | add (PolyThm (_, _, ths')) ths = ths'@ths + fun addpat (MonoPattern p) pats = p::pats + | addpat (PolyPattern (_, _, ps)) pats = ps@pats + val ths = fold_rev add ths [] + val pats = fold_rev addpat pats [] + in + Compute.update_with_cache computer pats ths + end + +fun conv_subst thy (subst : Type.tyenv) = + map (fn (iname, (sort, ty)) => (ctyp_of thy (TVar (iname, sort)), ctyp_of thy ty)) (Vartab.dest subst) + +fun add_monos thy monocs pats ths = + let + val changed = Unsynchronized.ref false + fun add monocs (th as (MonoThm _)) = ([], th) + | add monocs (PolyThm (th, instances, instanceths)) = + let + val (newsubsts, instances) = Linker.add_instances thy instances monocs + val _ = if not (null newsubsts) then changed := true else () + val newths = map (fn subst => Thm.instantiate (conv_subst thy subst, []) th) newsubsts +(* val _ = if not (null newths) then (print ("added new theorems: ", newths); ()) else ()*) + val newmonos = fold (fn th => fn monos => (snd (collect_consts_of_thm th))@monos) newths [] + in + (newmonos, PolyThm (th, instances, instanceths@newths)) + end + fun addpats monocs (pat as (MonoPattern _)) = pat + | addpats monocs (PolyPattern (p, instances, instancepats)) = + let + val (newsubsts, instances) = Linker.add_instances thy instances monocs + val _ = if not (null newsubsts) then changed := true else () + val newpats = map (fn subst => Envir.subst_term_types subst p) newsubsts + in + PolyPattern (p, instances, instancepats@newpats) + end + fun step monocs ths = + fold_rev (fn th => + fn (newmonos, ths) => + let + val (newmonos', th') = add monocs th + in + (newmonos'@newmonos, th'::ths) + end) + ths ([], []) + fun loop monocs pats ths = + let + val (monocs', ths') = step monocs ths + val pats' = map (addpats monocs) pats + in + if null (monocs') then + (pats', ths') + else + loop monocs' pats' ths' + end + val result = loop monocs pats ths + in + (!changed, result) + end + +datatype cthm = ComputeThm of term list * sort list * term + +fun thm2cthm th = + let + val {hyps, prop, shyps, ...} = Thm.rep_thm th + in + ComputeThm (hyps, shyps, prop) + end + +val cthm_ord' = prod_ord (prod_ord (list_ord Term_Ord.term_ord) (list_ord Term_Ord.sort_ord)) Term_Ord.term_ord + +fun cthm_ord (ComputeThm (h1, sh1, p1), ComputeThm (h2, sh2, p2)) = cthm_ord' (((h1,sh1), p1), ((h2, sh2), p2)) + +structure CThmtab = Table(type key = cthm val ord = cthm_ord) + +fun remove_duplicates ths = + let + val counter = Unsynchronized.ref 0 + val tab = Unsynchronized.ref (CThmtab.empty : unit CThmtab.table) + val thstab = Unsynchronized.ref (Inttab.empty : thm Inttab.table) + fun update th = + let + val key = thm2cthm th + in + case CThmtab.lookup (!tab) key of + NONE => ((tab := CThmtab.update_new (key, ()) (!tab)); thstab := Inttab.update_new (!counter, th) (!thstab); counter := !counter + 1) + | _ => () + end + val _ = map update ths + in + map snd (Inttab.dest (!thstab)) + end + +fun make_with_cache machine thy pats ths cs = + let + val ths = remove_duplicates ths + val (monocs, ths) = fold_rev (fn th => + fn (monocs, ths) => + let val (m, t) = create_theorem th in + (m@monocs, t::ths) + end) + ths (cs, []) + val pats = map create_pattern pats + val (_, (pats, ths)) = add_monos thy monocs pats ths + val computer = create_computer machine thy pats ths + in + PComputer (Theory.check_thy thy, computer, Unsynchronized.ref ths, Unsynchronized.ref pats) + end + +fun make machine thy ths cs = make_with_cache machine thy [] ths cs + +fun add_instances (PComputer (thyref, computer, rths, rpats)) cs = + let + val thy = Theory.deref thyref + val (changed, (pats, ths)) = add_monos thy cs (!rpats) (!rths) + in + if changed then + (update_computer computer pats ths; + rths := ths; + rpats := pats; + true) + else + false + + end + +fun add_instances' pc ts = add_instances pc (Linker.collect_consts ts) + +fun rewrite pc cts = + let + val _ = add_instances' pc (map term_of cts) + val computer = (computer_of pc) + in + map (fn ct => Compute.rewrite computer ct) cts + end + +fun simplify pc th = Compute.simplify (computer_of pc) th + +fun make_theorem pc th vars = + let + val _ = add_instances' pc [prop_of th] + + in + Compute.make_theorem (computer_of pc) th vars + end + +fun instantiate pc insts th = + let + val _ = add_instances' pc (map (term_of o snd) insts) + in + Compute.instantiate (computer_of pc) insts th + end + +fun evaluate_prem pc prem_no th = Compute.evaluate_prem (computer_of pc) prem_no th + +fun modus_ponens pc prem_no th' th = + let + val _ = add_instances' pc [prop_of th'] + in + Compute.modus_ponens (computer_of pc) prem_no th' th + end + + +end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/Compute_Oracle/report.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/Compute_Oracle/report.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,33 @@ +structure Report = +struct + +local + + val report_depth = Unsynchronized.ref 0 + fun space n = if n <= 0 then "" else (space (n-1))^" " + fun report_space () = space (!report_depth) + +in + +fun timeit f = + let + val t1 = Timing.start () + val x = f () + val t2 = Timing.message (Timing.result t1) + val _ = writeln ((report_space ()) ^ "--> "^t2) + in + x + end + +fun report s f = +let + val _ = writeln ((report_space ())^s) + val _ = report_depth := !report_depth + 1 + val x = timeit f + val _ = report_depth := !report_depth - 1 +in + x +end + +end +end \ No newline at end of file diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/Cplex.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/Cplex.thy Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,67 @@ +(* Title: HOL/Matrix/Cplex.thy + Author: Steven Obua +*) + +theory Cplex +imports SparseMatrix LP ComputeFloat ComputeNumeral +uses "Cplex_tools.ML" "CplexMatrixConverter.ML" "FloatSparseMatrixBuilder.ML" + "fspmlp.ML" ("matrixlp.ML") +begin + +lemma spm_mult_le_dual_prts: + assumes + "sorted_sparse_matrix A1" + "sorted_sparse_matrix A2" + "sorted_sparse_matrix c1" + "sorted_sparse_matrix c2" + "sorted_sparse_matrix y" + "sorted_sparse_matrix r1" + "sorted_sparse_matrix r2" + "sorted_spvec b" + "le_spmat [] y" + "sparse_row_matrix A1 \ A" + "A \ sparse_row_matrix A2" + "sparse_row_matrix c1 \ c" + "c \ sparse_row_matrix c2" + "sparse_row_matrix r1 \ x" + "x \ sparse_row_matrix r2" + "A * x \ sparse_row_matrix (b::('a::lattice_ring) spmat)" + shows + "c * x \ sparse_row_matrix (add_spmat (mult_spmat y b) + (let s1 = diff_spmat c1 (mult_spmat y A2); s2 = diff_spmat c2 (mult_spmat y A1) in + add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2)) (add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2)) + (add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1)) (mult_spmat (nprt_spmat s1) (nprt_spmat r1))))))" + apply (simp add: Let_def) + apply (insert assms) + apply (simp add: sparse_row_matrix_op_simps algebra_simps) + apply (rule mult_le_dual_prts[where A=A, simplified Let_def algebra_simps]) + apply (auto) + done + +lemma spm_mult_le_dual_prts_no_let: + assumes + "sorted_sparse_matrix A1" + "sorted_sparse_matrix A2" + "sorted_sparse_matrix c1" + "sorted_sparse_matrix c2" + "sorted_sparse_matrix y" + "sorted_sparse_matrix r1" + "sorted_sparse_matrix r2" + "sorted_spvec b" + "le_spmat [] y" + "sparse_row_matrix A1 \ A" + "A \ sparse_row_matrix A2" + "sparse_row_matrix c1 \ c" + "c \ sparse_row_matrix c2" + "sparse_row_matrix r1 \ x" + "x \ sparse_row_matrix r2" + "A * x \ sparse_row_matrix (b::('a::lattice_ring) spmat)" + shows + "c * x \ sparse_row_matrix (add_spmat (mult_spmat y b) + (mult_est_spmat r1 r2 (diff_spmat c1 (mult_spmat y A2)) (diff_spmat c2 (mult_spmat y A1))))" + by (simp add: assms mult_est_spmat_def spm_mult_le_dual_prts[where A=A, simplified Let_def]) + +use "matrixlp.ML" + +end + diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/CplexMatrixConverter.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/CplexMatrixConverter.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,128 @@ +(* Title: HOL/Matrix/CplexMatrixConverter.ML + Author: Steven Obua +*) + +signature MATRIX_BUILDER = +sig + type vector + type matrix + + val empty_vector : vector + val empty_matrix : matrix + + exception Nat_expected of int + val set_elem : vector -> int -> string -> vector + val set_vector : matrix -> int -> vector -> matrix +end; + +signature CPLEX_MATRIX_CONVERTER = +sig + structure cplex : CPLEX + structure matrix_builder : MATRIX_BUILDER + type vector = matrix_builder.vector + type matrix = matrix_builder.matrix + type naming = int * (int -> string) * (string -> int) + + exception Converter of string + + (* program must fulfill is_normed_cplexProg and must be an element of the image of elim_nonfree_bounds *) + (* convert_prog maximize c A b naming *) + val convert_prog : cplex.cplexProg -> bool * vector * matrix * vector * naming + + (* results must be optimal, converts_results returns the optimal value as string and the solution as vector *) + (* convert_results results name2index *) + val convert_results : cplex.cplexResult -> (string -> int) -> string * vector +end; + +functor MAKE_CPLEX_MATRIX_CONVERTER (structure cplex: CPLEX and matrix_builder: MATRIX_BUILDER) : CPLEX_MATRIX_CONVERTER = +struct + +structure cplex = cplex +structure matrix_builder = matrix_builder +type matrix = matrix_builder.matrix +type vector = matrix_builder.vector +type naming = int * (int -> string) * (string -> int) + +open matrix_builder +open cplex + +exception Converter of string; + +fun neg_term (cplexNeg t) = t + | neg_term (cplexSum ts) = cplexSum (map neg_term ts) + | neg_term t = cplexNeg t + +fun convert_prog (cplexProg (_, goal, constrs, bounds)) = + let + fun build_naming index i2s s2i [] = (index, i2s, s2i) + | build_naming index i2s s2i (cplexBounds (cplexNeg cplexInf, cplexLeq, cplexVar v, cplexLeq, cplexInf)::bounds) + = build_naming (index+1) (Inttab.update (index, v) i2s) (Symtab.update_new (v, index) s2i) bounds + | build_naming _ _ _ _ = raise (Converter "nonfree bound") + + val (varcount, i2s_tab, s2i_tab) = build_naming 0 Inttab.empty Symtab.empty bounds + + fun i2s i = case Inttab.lookup i2s_tab i of NONE => raise (Converter "index not found") + | SOME n => n + fun s2i s = case Symtab.lookup s2i_tab s of NONE => raise (Converter ("name not found: "^s)) + | SOME i => i + fun num2str positive (cplexNeg t) = num2str (not positive) t + | num2str positive (cplexNum num) = if positive then num else "-"^num + | num2str _ _ = raise (Converter "term is not a (possibly signed) number") + + fun setprod vec positive (cplexNeg t) = setprod vec (not positive) t + | setprod vec positive (cplexVar v) = set_elem vec (s2i v) (if positive then "1" else "-1") + | setprod vec positive (cplexProd (cplexNum num, cplexVar v)) = + set_elem vec (s2i v) (if positive then num else "-"^num) + | setprod _ _ _ = raise (Converter "term is not a normed product") + + fun sum2vec (cplexSum ts) = fold (fn t => fn vec => setprod vec true t) ts empty_vector + | sum2vec t = setprod empty_vector true t + + fun constrs2Ab j A b [] = (A, b) + | constrs2Ab j A b ((_, cplexConstr (cplexLeq, (t1,t2)))::cs) = + constrs2Ab (j+1) (set_vector A j (sum2vec t1)) (set_elem b j (num2str true t2)) cs + | constrs2Ab j A b ((_, cplexConstr (cplexGeq, (t1,t2)))::cs) = + constrs2Ab (j+1) (set_vector A j (sum2vec (neg_term t1))) (set_elem b j (num2str true (neg_term t2))) cs + | constrs2Ab j A b ((_, cplexConstr (cplexEq, (t1,t2)))::cs) = + constrs2Ab j A b ((NONE, cplexConstr (cplexLeq, (t1,t2))):: + (NONE, cplexConstr (cplexGeq, (t1, t2)))::cs) + | constrs2Ab _ _ _ _ = raise (Converter "no strict constraints allowed") + + val (A, b) = constrs2Ab 0 empty_matrix empty_vector constrs + + val (goal_maximize, goal_term) = + case goal of + (cplexMaximize t) => (true, t) + | (cplexMinimize t) => (false, t) + in + (goal_maximize, sum2vec goal_term, A, b, (varcount, i2s, s2i)) + end + +fun convert_results (cplex.Optimal (opt, entries)) name2index = + let + fun setv (name, value) v = matrix_builder.set_elem v (name2index name) value + in + (opt, fold setv entries (matrix_builder.empty_vector)) + end + | convert_results _ _ = raise (Converter "No optimal result") + +end; + +structure SimpleMatrixBuilder : MATRIX_BUILDER = +struct +type vector = (int * string) list +type matrix = (int * vector) list + +val empty_matrix = [] +val empty_vector = [] + +exception Nat_expected of int; + +fun set_elem v i s = v @ [(i, s)] + +fun set_vector m i v = m @ [(i, v)] + +end; + +structure SimpleCplexMatrixConverter = + MAKE_CPLEX_MATRIX_CONVERTER(structure cplex = Cplex and matrix_builder = SimpleMatrixBuilder); diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/Cplex_tools.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/Cplex_tools.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,1192 @@ +(* Title: HOL/Matrix/Cplex_tools.ML + Author: Steven Obua +*) + +signature CPLEX = +sig + + datatype cplexTerm = cplexVar of string | cplexNum of string | cplexInf + | cplexNeg of cplexTerm + | cplexProd of cplexTerm * cplexTerm + | cplexSum of (cplexTerm list) + + datatype cplexComp = cplexLe | cplexLeq | cplexEq | cplexGe | cplexGeq + + datatype cplexGoal = cplexMinimize of cplexTerm + | cplexMaximize of cplexTerm + + datatype cplexConstr = cplexConstr of cplexComp * + (cplexTerm * cplexTerm) + + datatype cplexBounds = cplexBounds of cplexTerm * cplexComp * cplexTerm + * cplexComp * cplexTerm + | cplexBound of cplexTerm * cplexComp * cplexTerm + + datatype cplexProg = cplexProg of string + * cplexGoal + * ((string option * cplexConstr) + list) + * cplexBounds list + + datatype cplexResult = Unbounded + | Infeasible + | Undefined + | Optimal of string * + (((* name *) string * + (* value *) string) list) + + datatype cplexSolver = SOLVER_DEFAULT | SOLVER_CPLEX | SOLVER_GLPK + + exception Load_cplexFile of string + exception Load_cplexResult of string + exception Save_cplexFile of string + exception Execute of string + + val load_cplexFile : string -> cplexProg + + val save_cplexFile : string -> cplexProg -> unit + + val elim_nonfree_bounds : cplexProg -> cplexProg + + val relax_strict_ineqs : cplexProg -> cplexProg + + val is_normed_cplexProg : cplexProg -> bool + + val get_solver : unit -> cplexSolver + val set_solver : cplexSolver -> unit + val solve : cplexProg -> cplexResult +end; + +structure Cplex : CPLEX = +struct + +datatype cplexSolver = SOLVER_DEFAULT | SOLVER_CPLEX | SOLVER_GLPK + +val cplexsolver = Unsynchronized.ref SOLVER_DEFAULT; +fun get_solver () = !cplexsolver; +fun set_solver s = (cplexsolver := s); + +exception Load_cplexFile of string; +exception Load_cplexResult of string; +exception Save_cplexFile of string; + +datatype cplexTerm = cplexVar of string + | cplexNum of string + | cplexInf + | cplexNeg of cplexTerm + | cplexProd of cplexTerm * cplexTerm + | cplexSum of (cplexTerm list) +datatype cplexComp = cplexLe | cplexLeq | cplexEq | cplexGe | cplexGeq +datatype cplexGoal = cplexMinimize of cplexTerm | cplexMaximize of cplexTerm +datatype cplexConstr = cplexConstr of cplexComp * (cplexTerm * cplexTerm) +datatype cplexBounds = cplexBounds of cplexTerm * cplexComp * cplexTerm + * cplexComp * cplexTerm + | cplexBound of cplexTerm * cplexComp * cplexTerm +datatype cplexProg = cplexProg of string + * cplexGoal + * ((string option * cplexConstr) list) + * cplexBounds list + +fun rev_cmp cplexLe = cplexGe + | rev_cmp cplexLeq = cplexGeq + | rev_cmp cplexGe = cplexLe + | rev_cmp cplexGeq = cplexLeq + | rev_cmp cplexEq = cplexEq + +fun the NONE = raise (Load_cplexFile "SOME expected") + | the (SOME x) = x; + +fun modulo_signed is_something (cplexNeg u) = is_something u + | modulo_signed is_something u = is_something u + +fun is_Num (cplexNum _) = true + | is_Num _ = false + +fun is_Inf cplexInf = true + | is_Inf _ = false + +fun is_Var (cplexVar _) = true + | is_Var _ = false + +fun is_Neg (cplexNeg _) = true + | is_Neg _ = false + +fun is_normed_Prod (cplexProd (t1, t2)) = + (is_Num t1) andalso (is_Var t2) + | is_normed_Prod x = is_Var x + +fun is_normed_Sum (cplexSum ts) = + (ts <> []) andalso forall (modulo_signed is_normed_Prod) ts + | is_normed_Sum x = modulo_signed is_normed_Prod x + +fun is_normed_Constr (cplexConstr (_, (t1, t2))) = + (is_normed_Sum t1) andalso (modulo_signed is_Num t2) + +fun is_Num_or_Inf x = is_Inf x orelse is_Num x + +fun is_normed_Bounds (cplexBounds (t1, c1, t2, c2, t3)) = + (c1 = cplexLe orelse c1 = cplexLeq) andalso + (c2 = cplexLe orelse c2 = cplexLeq) andalso + is_Var t2 andalso + modulo_signed is_Num_or_Inf t1 andalso + modulo_signed is_Num_or_Inf t3 + | is_normed_Bounds (cplexBound (t1, c, t2)) = + (is_Var t1 andalso (modulo_signed is_Num_or_Inf t2)) + orelse + (c <> cplexEq andalso + is_Var t2 andalso (modulo_signed is_Num_or_Inf t1)) + +fun term_of_goal (cplexMinimize x) = x + | term_of_goal (cplexMaximize x) = x + +fun is_normed_cplexProg (cplexProg (_, goal, constraints, bounds)) = + is_normed_Sum (term_of_goal goal) andalso + forall (fn (_,x) => is_normed_Constr x) constraints andalso + forall is_normed_Bounds bounds + +fun is_NL s = s = "\n" + +fun is_blank s = forall (fn c => c <> #"\n" andalso Char.isSpace c) (String.explode s) + +fun is_num a = + let + val b = String.explode a + fun num4 cs = forall Char.isDigit cs + fun num3 [] = true + | num3 (ds as (c::cs)) = + if c = #"+" orelse c = #"-" then + num4 cs + else + num4 ds + fun num2 [] = true + | num2 (c::cs) = + if c = #"e" orelse c = #"E" then num3 cs + else (Char.isDigit c) andalso num2 cs + fun num1 [] = true + | num1 (c::cs) = + if c = #"." then num2 cs + else if c = #"e" orelse c = #"E" then num3 cs + else (Char.isDigit c) andalso num1 cs + fun num [] = true + | num (c::cs) = + if c = #"." then num2 cs + else (Char.isDigit c) andalso num1 cs + in + num b + end + +fun is_delimiter s = s = "+" orelse s = "-" orelse s = ":" + +fun is_cmp s = s = "<" orelse s = ">" orelse s = "<=" + orelse s = ">=" orelse s = "=" + +fun is_symbol a = + let + val symbol_char = String.explode "!\"#$%&()/,.;?@_`'{}|~" + fun is_symbol_char c = Char.isAlphaNum c orelse + exists (fn d => d=c) symbol_char + fun is_symbol_start c = is_symbol_char c andalso + not (Char.isDigit c) andalso + not (c= #".") + val b = String.explode a + in + b <> [] andalso is_symbol_start (hd b) andalso + forall is_symbol_char b + end + +fun to_upper s = String.implode (map Char.toUpper (String.explode s)) + +fun keyword x = + let + val a = to_upper x + in + if a = "BOUNDS" orelse a = "BOUND" then + SOME "BOUNDS" + else if a = "MINIMIZE" orelse a = "MINIMUM" orelse a = "MIN" then + SOME "MINIMIZE" + else if a = "MAXIMIZE" orelse a = "MAXIMUM" orelse a = "MAX" then + SOME "MAXIMIZE" + else if a = "ST" orelse a = "S.T." orelse a = "ST." then + SOME "ST" + else if a = "FREE" orelse a = "END" then + SOME a + else if a = "GENERAL" orelse a = "GENERALS" orelse a = "GEN" then + SOME "GENERAL" + else if a = "INTEGER" orelse a = "INTEGERS" orelse a = "INT" then + SOME "INTEGER" + else if a = "BINARY" orelse a = "BINARIES" orelse a = "BIN" then + SOME "BINARY" + else if a = "INF" orelse a = "INFINITY" then + SOME "INF" + else + NONE + end + +val TOKEN_ERROR = ~1 +val TOKEN_BLANK = 0 +val TOKEN_NUM = 1 +val TOKEN_DELIMITER = 2 +val TOKEN_SYMBOL = 3 +val TOKEN_LABEL = 4 +val TOKEN_CMP = 5 +val TOKEN_KEYWORD = 6 +val TOKEN_NL = 7 + +(* tokenize takes a list of chars as argument and returns a list of + int * string pairs, each string representing a "cplex token", + and each int being one of TOKEN_NUM, TOKEN_DELIMITER, TOKEN_CMP + or TOKEN_SYMBOL *) +fun tokenize s = + let + val flist = [(is_NL, TOKEN_NL), + (is_blank, TOKEN_BLANK), + (is_num, TOKEN_NUM), + (is_delimiter, TOKEN_DELIMITER), + (is_cmp, TOKEN_CMP), + (is_symbol, TOKEN_SYMBOL)] + fun match_helper [] s = (fn _ => false, TOKEN_ERROR) + | match_helper (f::fs) s = + if ((fst f) s) then f else match_helper fs s + fun match s = match_helper flist s + fun tok s = + if s = "" then [] else + let + val h = String.substring (s,0,1) + val (f, j) = match h + fun len i = + if size s = i then i + else if f (String.substring (s,0,i+1)) then + len (i+1) + else i + in + if j < 0 then + (if h = "\\" then [] + else raise (Load_cplexFile ("token expected, found: " + ^s))) + else + let + val l = len 1 + val u = String.substring (s,0,l) + val v = String.extract (s,l,NONE) + in + if j = 0 then tok v else (j, u) :: tok v + end + end + in + tok s + end + +exception Tokenize of string; + +fun tokenize_general flist s = + let + fun match_helper [] s = raise (Tokenize s) + | match_helper (f::fs) s = + if ((fst f) s) then f else match_helper fs s + fun match s = match_helper flist s + fun tok s = + if s = "" then [] else + let + val h = String.substring (s,0,1) + val (f, j) = match h + fun len i = + if size s = i then i + else if f (String.substring (s,0,i+1)) then + len (i+1) + else i + val l = len 1 + in + (j, String.substring (s,0,l)) :: tok (String.extract (s,l,NONE)) + end + in + tok s + end + +fun load_cplexFile name = + let + val f = TextIO.openIn name + val ignore_NL = Unsynchronized.ref true + val rest = Unsynchronized.ref [] + + fun is_symbol s c = (fst c) = TOKEN_SYMBOL andalso (to_upper (snd c)) = s + + fun readToken_helper () = + if length (!rest) > 0 then + let val u = hd (!rest) in + ( + rest := tl (!rest); + SOME u + ) + end + else + (case TextIO.inputLine f of + NONE => NONE + | SOME s => + let val t = tokenize s in + if (length t >= 2 andalso + snd(hd (tl t)) = ":") + then + rest := (TOKEN_LABEL, snd (hd t)) :: (tl (tl t)) + else if (length t >= 2) andalso is_symbol "SUBJECT" (hd (t)) + andalso is_symbol "TO" (hd (tl t)) + then + rest := (TOKEN_SYMBOL, "ST") :: (tl (tl t)) + else + rest := t; + readToken_helper () + end) + + fun readToken_helper2 () = + let val c = readToken_helper () in + if c = NONE then NONE + else if !ignore_NL andalso fst (the c) = TOKEN_NL then + readToken_helper2 () + else if fst (the c) = TOKEN_SYMBOL + andalso keyword (snd (the c)) <> NONE + then SOME (TOKEN_KEYWORD, the (keyword (snd (the c)))) + else c + end + + fun readToken () = readToken_helper2 () + + fun pushToken a = rest := (a::(!rest)) + + fun is_value token = + fst token = TOKEN_NUM orelse (fst token = TOKEN_KEYWORD + andalso snd token = "INF") + + fun get_value token = + if fst token = TOKEN_NUM then + cplexNum (snd token) + else if fst token = TOKEN_KEYWORD andalso snd token = "INF" + then + cplexInf + else + raise (Load_cplexFile "num expected") + + fun readTerm_Product only_num = + let val c = readToken () in + if c = NONE then NONE + else if fst (the c) = TOKEN_SYMBOL + then ( + if only_num then (pushToken (the c); NONE) + else SOME (cplexVar (snd (the c))) + ) + else if only_num andalso is_value (the c) then + SOME (get_value (the c)) + else if is_value (the c) then + let val t1 = get_value (the c) + val d = readToken () + in + if d = NONE then SOME t1 + else if fst (the d) = TOKEN_SYMBOL then + SOME (cplexProd (t1, cplexVar (snd (the d)))) + else + (pushToken (the d); SOME t1) + end + else (pushToken (the c); NONE) + end + + fun readTerm_Signed only_signed only_num = + let + val c = readToken () + in + if c = NONE then NONE + else + let val d = the c in + if d = (TOKEN_DELIMITER, "+") then + readTerm_Product only_num + else if d = (TOKEN_DELIMITER, "-") then + SOME (cplexNeg (the (readTerm_Product + only_num))) + else (pushToken d; + if only_signed then NONE + else readTerm_Product only_num) + end + end + + fun readTerm_Sum first_signed = + let val c = readTerm_Signed first_signed false in + if c = NONE then [] else (the c)::(readTerm_Sum true) + end + + fun readTerm () = + let val c = readTerm_Sum false in + if c = [] then NONE + else if tl c = [] then SOME (hd c) + else SOME (cplexSum c) + end + + fun readLabeledTerm () = + let val c = readToken () in + if c = NONE then (NONE, NONE) + else if fst (the c) = TOKEN_LABEL then + let val t = readTerm () in + if t = NONE then + raise (Load_cplexFile ("term after label "^ + (snd (the c))^ + " expected")) + else (SOME (snd (the c)), t) + end + else (pushToken (the c); (NONE, readTerm ())) + end + + fun readGoal () = + let + val g = readToken () + in + if g = SOME (TOKEN_KEYWORD, "MAXIMIZE") then + cplexMaximize (the (snd (readLabeledTerm ()))) + else if g = SOME (TOKEN_KEYWORD, "MINIMIZE") then + cplexMinimize (the (snd (readLabeledTerm ()))) + else raise (Load_cplexFile "MAXIMIZE or MINIMIZE expected") + end + + fun str2cmp b = + (case b of + "<" => cplexLe + | "<=" => cplexLeq + | ">" => cplexGe + | ">=" => cplexGeq + | "=" => cplexEq + | _ => raise (Load_cplexFile (b^" is no TOKEN_CMP"))) + + fun readConstraint () = + let + val t = readLabeledTerm () + fun make_constraint b t1 t2 = + cplexConstr + (str2cmp b, + (t1, t2)) + in + if snd t = NONE then NONE + else + let val c = readToken () in + if c = NONE orelse fst (the c) <> TOKEN_CMP + then raise (Load_cplexFile "TOKEN_CMP expected") + else + let val n = readTerm_Signed false true in + if n = NONE then + raise (Load_cplexFile "num expected") + else + SOME (fst t, + make_constraint (snd (the c)) + (the (snd t)) + (the n)) + end + end + end + + fun readST () = + let + fun readbody () = + let val t = readConstraint () in + if t = NONE then [] + else if (is_normed_Constr (snd (the t))) then + (the t)::(readbody ()) + else if (fst (the t) <> NONE) then + raise (Load_cplexFile + ("constraint '"^(the (fst (the t)))^ + "'is not normed")) + else + raise (Load_cplexFile + "constraint is not normed") + end + in + if readToken () = SOME (TOKEN_KEYWORD, "ST") + then + readbody () + else + raise (Load_cplexFile "ST expected") + end + + fun readCmp () = + let val c = readToken () in + if c = NONE then NONE + else if fst (the c) = TOKEN_CMP then + SOME (str2cmp (snd (the c))) + else (pushToken (the c); NONE) + end + + fun skip_NL () = + let val c = readToken () in + if c <> NONE andalso fst (the c) = TOKEN_NL then + skip_NL () + else + (pushToken (the c); ()) + end + + fun make_bounds c t1 t2 = + cplexBound (t1, c, t2) + + fun readBound () = + let + val _ = skip_NL () + val t1 = readTerm () + in + if t1 = NONE then NONE + else + let + val c1 = readCmp () + in + if c1 = NONE then + let + val c = readToken () + in + if c = SOME (TOKEN_KEYWORD, "FREE") then + SOME ( + cplexBounds (cplexNeg cplexInf, + cplexLeq, + the t1, + cplexLeq, + cplexInf)) + else + raise (Load_cplexFile "FREE expected") + end + else + let + val t2 = readTerm () + in + if t2 = NONE then + raise (Load_cplexFile "term expected") + else + let val c2 = readCmp () in + if c2 = NONE then + SOME (make_bounds (the c1) + (the t1) + (the t2)) + else + SOME ( + cplexBounds (the t1, + the c1, + the t2, + the c2, + the (readTerm()))) + end + end + end + end + + fun readBounds () = + let + fun makestring _ = "?" + fun readbody () = + let + val b = readBound () + in + if b = NONE then [] + else if (is_normed_Bounds (the b)) then + (the b)::(readbody()) + else ( + raise (Load_cplexFile + ("bounds are not normed in: "^ + (makestring (the b))))) + end + in + if readToken () = SOME (TOKEN_KEYWORD, "BOUNDS") then + readbody () + else raise (Load_cplexFile "BOUNDS expected") + end + + fun readEnd () = + if readToken () = SOME (TOKEN_KEYWORD, "END") then () + else raise (Load_cplexFile "END expected") + + val result_Goal = readGoal () + val result_ST = readST () + val _ = ignore_NL := false + val result_Bounds = readBounds () + val _ = ignore_NL := true + val _ = readEnd () + val _ = TextIO.closeIn f + in + cplexProg (name, result_Goal, result_ST, result_Bounds) + end + +fun save_cplexFile filename (cplexProg (_, goal, constraints, bounds)) = + let + val f = TextIO.openOut filename + + fun basic_write s = TextIO.output(f, s) + + val linebuf = Unsynchronized.ref "" + fun buf_flushline s = + (basic_write (!linebuf); + basic_write "\n"; + linebuf := s) + fun buf_add s = linebuf := (!linebuf) ^ s + + fun write s = + if (String.size s) + (String.size (!linebuf)) >= 250 then + buf_flushline (" "^s) + else + buf_add s + + fun writeln s = (buf_add s; buf_flushline "") + + fun write_term (cplexVar x) = write x + | write_term (cplexNum x) = write x + | write_term cplexInf = write "inf" + | write_term (cplexProd (cplexNum "1", b)) = write_term b + | write_term (cplexProd (a, b)) = + (write_term a; write " "; write_term b) + | write_term (cplexNeg x) = (write " - "; write_term x) + | write_term (cplexSum ts) = write_terms ts + and write_terms [] = () + | write_terms (t::ts) = + (if (not (is_Neg t)) then write " + " else (); + write_term t; write_terms ts) + + fun write_goal (cplexMaximize term) = + (writeln "MAXIMIZE"; write_term term; writeln "") + | write_goal (cplexMinimize term) = + (writeln "MINIMIZE"; write_term term; writeln "") + + fun write_cmp cplexLe = write "<" + | write_cmp cplexLeq = write "<=" + | write_cmp cplexEq = write "=" + | write_cmp cplexGe = write ">" + | write_cmp cplexGeq = write ">=" + + fun write_constr (cplexConstr (cmp, (a,b))) = + (write_term a; + write " "; + write_cmp cmp; + write " "; + write_term b) + + fun write_constraints [] = () + | write_constraints (c::cs) = + (if (fst c <> NONE) + then + (write (the (fst c)); write ": ") + else + (); + write_constr (snd c); + writeln ""; + write_constraints cs) + + fun write_bounds [] = () + | write_bounds ((cplexBounds (t1,c1,t2,c2,t3))::bs) = + ((if t1 = cplexNeg cplexInf andalso t3 = cplexInf + andalso (c1 = cplexLeq orelse c1 = cplexLe) + andalso (c2 = cplexLeq orelse c2 = cplexLe) + then + (write_term t2; write " free") + else + (write_term t1; write " "; write_cmp c1; write " "; + write_term t2; write " "; write_cmp c2; write " "; + write_term t3) + ); writeln ""; write_bounds bs) + | write_bounds ((cplexBound (t1, c, t2)) :: bs) = + (write_term t1; write " "; + write_cmp c; write " "; + write_term t2; writeln ""; write_bounds bs) + + val _ = write_goal goal + val _ = (writeln ""; writeln "ST") + val _ = write_constraints constraints + val _ = (writeln ""; writeln "BOUNDS") + val _ = write_bounds bounds + val _ = (writeln ""; writeln "END") + val _ = TextIO.closeOut f + in + () + end + +fun norm_Constr (constr as cplexConstr (c, (t1, t2))) = + if not (modulo_signed is_Num t2) andalso + modulo_signed is_Num t1 + then + [cplexConstr (rev_cmp c, (t2, t1))] + else if (c = cplexLe orelse c = cplexLeq) andalso + (t1 = (cplexNeg cplexInf) orelse t2 = cplexInf) + then + [] + else if (c = cplexGe orelse c = cplexGeq) andalso + (t1 = cplexInf orelse t2 = cplexNeg cplexInf) + then + [] + else + [constr] + +fun bound2constr (cplexBounds (t1,c1,t2,c2,t3)) = + (norm_Constr(cplexConstr (c1, (t1, t2)))) + @ (norm_Constr(cplexConstr (c2, (t2, t3)))) + | bound2constr (cplexBound (t1, cplexEq, t2)) = + (norm_Constr(cplexConstr (cplexLeq, (t1, t2)))) + @ (norm_Constr(cplexConstr (cplexLeq, (t2, t1)))) + | bound2constr (cplexBound (t1, c1, t2)) = + norm_Constr(cplexConstr (c1, (t1,t2))) + +val emptyset = Symtab.empty + +fun singleton v = Symtab.update (v, ()) emptyset + +fun merge a b = Symtab.merge (op =) (a, b) + +fun mergemap f ts = fold (fn x => fn table => merge table (f x)) ts Symtab.empty + +fun diff a b = Symtab.fold (Symtab.delete_safe o fst) b a + +fun collect_vars (cplexVar v) = singleton v + | collect_vars (cplexNeg t) = collect_vars t + | collect_vars (cplexProd (t1, t2)) = + merge (collect_vars t1) (collect_vars t2) + | collect_vars (cplexSum ts) = mergemap collect_vars ts + | collect_vars _ = emptyset + +(* Eliminates all nonfree bounds from the linear program and produces an + equivalent program with only free bounds + IF for the input program P holds: is_normed_cplexProg P *) +fun elim_nonfree_bounds (cplexProg (name, goal, constraints, bounds)) = + let + fun collect_constr_vars (_, cplexConstr (_, (t1,_))) = + (collect_vars t1) + + val cvars = merge (collect_vars (term_of_goal goal)) + (mergemap collect_constr_vars constraints) + + fun collect_lower_bounded_vars + (cplexBounds (_, _, cplexVar v, _, _)) = + singleton v + | collect_lower_bounded_vars + (cplexBound (_, cplexLe, cplexVar v)) = + singleton v + | collect_lower_bounded_vars + (cplexBound (_, cplexLeq, cplexVar v)) = + singleton v + | collect_lower_bounded_vars + (cplexBound (cplexVar v, cplexGe,_)) = + singleton v + | collect_lower_bounded_vars + (cplexBound (cplexVar v, cplexGeq, _)) = + singleton v + | collect_lower_bounded_vars + (cplexBound (cplexVar v, cplexEq, _)) = + singleton v + | collect_lower_bounded_vars _ = emptyset + + val lvars = mergemap collect_lower_bounded_vars bounds + val positive_vars = diff cvars lvars + val zero = cplexNum "0" + + fun make_pos_constr v = + (NONE, cplexConstr (cplexGeq, ((cplexVar v), zero))) + + fun make_free_bound v = + cplexBounds (cplexNeg cplexInf, cplexLeq, + cplexVar v, cplexLeq, + cplexInf) + + val pos_constrs = rev (Symtab.fold + (fn (k, _) => cons (make_pos_constr k)) + positive_vars []) + val bound_constrs = map (pair NONE) + (maps bound2constr bounds) + val constraints' = constraints @ pos_constrs @ bound_constrs + val bounds' = rev (Symtab.fold (fn (v, _) => cons (make_free_bound v)) cvars []); + in + cplexProg (name, goal, constraints', bounds') + end + +fun relax_strict_ineqs (cplexProg (name, goals, constrs, bounds)) = + let + fun relax cplexLe = cplexLeq + | relax cplexGe = cplexGeq + | relax x = x + + fun relax_constr (n, cplexConstr(c, (t1, t2))) = + (n, cplexConstr(relax c, (t1, t2))) + + fun relax_bounds (cplexBounds (t1, c1, t2, c2, t3)) = + cplexBounds (t1, relax c1, t2, relax c2, t3) + | relax_bounds (cplexBound (t1, c, t2)) = + cplexBound (t1, relax c, t2) + in + cplexProg (name, + goals, + map relax_constr constrs, + map relax_bounds bounds) + end + +datatype cplexResult = Unbounded + | Infeasible + | Undefined + | Optimal of string * ((string * string) list) + +fun is_separator x = forall (fn c => c = #"-") (String.explode x) + +fun is_sign x = (x = "+" orelse x = "-") + +fun is_colon x = (x = ":") + +fun is_resultsymbol a = + let + val symbol_char = String.explode "!\"#$%&()/,.;?@_`'{}|~-" + fun is_symbol_char c = Char.isAlphaNum c orelse + exists (fn d => d=c) symbol_char + fun is_symbol_start c = is_symbol_char c andalso + not (Char.isDigit c) andalso + not (c= #".") andalso + not (c= #"-") + val b = String.explode a + in + b <> [] andalso is_symbol_start (hd b) andalso + forall is_symbol_char b + end + +val TOKEN_SIGN = 100 +val TOKEN_COLON = 101 +val TOKEN_SEPARATOR = 102 + +fun load_glpkResult name = + let + val flist = [(is_NL, TOKEN_NL), + (is_blank, TOKEN_BLANK), + (is_num, TOKEN_NUM), + (is_sign, TOKEN_SIGN), + (is_colon, TOKEN_COLON), + (is_cmp, TOKEN_CMP), + (is_resultsymbol, TOKEN_SYMBOL), + (is_separator, TOKEN_SEPARATOR)] + + val tokenize = tokenize_general flist + + val f = TextIO.openIn name + + val rest = Unsynchronized.ref [] + + fun readToken_helper () = + if length (!rest) > 0 then + let val u = hd (!rest) in + ( + rest := tl (!rest); + SOME u + ) + end + else + (case TextIO.inputLine f of + NONE => NONE + | SOME s => (rest := tokenize s; readToken_helper())) + + fun is_tt tok ty = (tok <> NONE andalso (fst (the tok)) = ty) + + fun pushToken a = if a = NONE then () else (rest := ((the a)::(!rest))) + + fun readToken () = + let val t = readToken_helper () in + if is_tt t TOKEN_BLANK then + readToken () + else if is_tt t TOKEN_NL then + let val t2 = readToken_helper () in + if is_tt t2 TOKEN_SIGN then + (pushToken (SOME (TOKEN_SEPARATOR, "-")); t) + else + (pushToken t2; t) + end + else if is_tt t TOKEN_SIGN then + let val t2 = readToken_helper () in + if is_tt t2 TOKEN_NUM then + (SOME (TOKEN_NUM, (snd (the t))^(snd (the t2)))) + else + (pushToken t2; t) + end + else + t + end + + fun readRestOfLine P = + let + val t = readToken () + in + if is_tt t TOKEN_NL orelse t = NONE + then P + else readRestOfLine P + end + + fun readHeader () = + let + fun readStatus () = readRestOfLine ("STATUS", snd (the (readToken ()))) + fun readObjective () = readRestOfLine ("OBJECTIVE", snd (the (readToken (); readToken (); readToken ()))) + val t1 = readToken () + val t2 = readToken () + in + if is_tt t1 TOKEN_SYMBOL andalso is_tt t2 TOKEN_COLON + then + case to_upper (snd (the t1)) of + "STATUS" => (readStatus ())::(readHeader ()) + | "OBJECTIVE" => (readObjective())::(readHeader ()) + | _ => (readRestOfLine (); readHeader ()) + else + (pushToken t2; pushToken t1; []) + end + + fun skip_until_sep () = + let val x = readToken () in + if is_tt x TOKEN_SEPARATOR then + readRestOfLine () + else + skip_until_sep () + end + + fun load_value () = + let + val t1 = readToken () + val t2 = readToken () + in + if is_tt t1 TOKEN_NUM andalso is_tt t2 TOKEN_SYMBOL then + let + val t = readToken () + val state = if is_tt t TOKEN_NL then readToken () else t + val _ = if is_tt state TOKEN_SYMBOL then () else raise (Load_cplexResult "state expected") + val k = readToken () + in + if is_tt k TOKEN_NUM then + readRestOfLine (SOME (snd (the t2), snd (the k))) + else + raise (Load_cplexResult "number expected") + end + else + (pushToken t2; pushToken t1; NONE) + end + + fun load_values () = + let val v = load_value () in + if v = NONE then [] else (the v)::(load_values ()) + end + + val header = readHeader () + + val result = + case AList.lookup (op =) header "STATUS" of + SOME "INFEASIBLE" => Infeasible + | SOME "UNBOUNDED" => Unbounded + | SOME "OPTIMAL" => Optimal (the (AList.lookup (op =) header "OBJECTIVE"), + (skip_until_sep (); + skip_until_sep (); + load_values ())) + | _ => Undefined + + val _ = TextIO.closeIn f + in + result + end + handle (Tokenize s) => raise (Load_cplexResult ("Tokenize: "^s)) + | Option => raise (Load_cplexResult "Option") + +fun load_cplexResult name = + let + val flist = [(is_NL, TOKEN_NL), + (is_blank, TOKEN_BLANK), + (is_num, TOKEN_NUM), + (is_sign, TOKEN_SIGN), + (is_colon, TOKEN_COLON), + (is_cmp, TOKEN_CMP), + (is_resultsymbol, TOKEN_SYMBOL)] + + val tokenize = tokenize_general flist + + val f = TextIO.openIn name + + val rest = Unsynchronized.ref [] + + fun readToken_helper () = + if length (!rest) > 0 then + let val u = hd (!rest) in + ( + rest := tl (!rest); + SOME u + ) + end + else + (case TextIO.inputLine f of + NONE => NONE + | SOME s => (rest := tokenize s; readToken_helper())) + + fun is_tt tok ty = (tok <> NONE andalso (fst (the tok)) = ty) + + fun pushToken a = if a = NONE then () else (rest := ((the a)::(!rest))) + + fun readToken () = + let val t = readToken_helper () in + if is_tt t TOKEN_BLANK then + readToken () + else if is_tt t TOKEN_SIGN then + let val t2 = readToken_helper () in + if is_tt t2 TOKEN_NUM then + (SOME (TOKEN_NUM, (snd (the t))^(snd (the t2)))) + else + (pushToken t2; t) + end + else + t + end + + fun readRestOfLine P = + let + val t = readToken () + in + if is_tt t TOKEN_NL orelse t = NONE + then P + else readRestOfLine P + end + + fun readHeader () = + let + fun readStatus () = readRestOfLine ("STATUS", snd (the (readToken ()))) + fun readObjective () = + let + val t = readToken () + in + if is_tt t TOKEN_SYMBOL andalso to_upper (snd (the t)) = "VALUE" then + readRestOfLine ("OBJECTIVE", snd (the (readToken()))) + else + readRestOfLine ("OBJECTIVE_NAME", snd (the t)) + end + + val t = readToken () + in + if is_tt t TOKEN_SYMBOL then + case to_upper (snd (the t)) of + "STATUS" => (readStatus ())::(readHeader ()) + | "OBJECTIVE" => (readObjective ())::(readHeader ()) + | "SECTION" => (pushToken t; []) + | _ => (readRestOfLine (); readHeader ()) + else + (readRestOfLine (); readHeader ()) + end + + fun skip_nls () = + let val x = readToken () in + if is_tt x TOKEN_NL then + skip_nls () + else + (pushToken x; ()) + end + + fun skip_paragraph () = + if is_tt (readToken ()) TOKEN_NL then + (if is_tt (readToken ()) TOKEN_NL then + skip_nls () + else + skip_paragraph ()) + else + skip_paragraph () + + fun load_value () = + let + val t1 = readToken () + val t1 = if is_tt t1 TOKEN_SYMBOL andalso snd (the t1) = "A" then readToken () else t1 + in + if is_tt t1 TOKEN_NUM then + let + val name = readToken () + val status = readToken () + val value = readToken () + in + if is_tt name TOKEN_SYMBOL andalso + is_tt status TOKEN_SYMBOL andalso + is_tt value TOKEN_NUM + then + readRestOfLine (SOME (snd (the name), snd (the value))) + else + raise (Load_cplexResult "column line expected") + end + else + (pushToken t1; NONE) + end + + fun load_values () = + let val v = load_value () in + if v = NONE then [] else (the v)::(load_values ()) + end + + val header = readHeader () + + val result = + case AList.lookup (op =) header "STATUS" of + SOME "INFEASIBLE" => Infeasible + | SOME "NONOPTIMAL" => Unbounded + | SOME "OPTIMAL" => Optimal (the (AList.lookup (op =) header "OBJECTIVE"), + (skip_paragraph (); + skip_paragraph (); + skip_paragraph (); + skip_paragraph (); + skip_paragraph (); + load_values ())) + | _ => Undefined + + val _ = TextIO.closeIn f + in + result + end + handle (Tokenize s) => raise (Load_cplexResult ("Tokenize: "^s)) + | Option => raise (Load_cplexResult "Option") + +exception Execute of string; + +fun tmp_file s = Path.implode (Path.expand (File.tmp_path (Path.basic s))); +fun wrap s = "\""^s^"\""; + +fun solve_glpk prog = + let + val name = string_of_int (Time.toMicroseconds (Time.now ())) + val lpname = tmp_file (name^".lp") + val resultname = tmp_file (name^".txt") + val _ = save_cplexFile lpname prog + val cplex_path = getenv "GLPK_PATH" + val cplex = if cplex_path = "" then "glpsol" else cplex_path + val command = (wrap cplex)^" --lpt "^(wrap lpname)^" --output "^(wrap resultname) + val answer = #1 (Isabelle_System.bash_output command) + in + let + val result = load_glpkResult resultname + val _ = OS.FileSys.remove lpname + val _ = OS.FileSys.remove resultname + in + result + end + handle (Load_cplexResult s) => raise (Execute ("Load_cplexResult: "^s^"\nExecute: "^answer)) + | _ => raise (Execute answer) (* FIXME avoid handle _ *) + end + +fun solve_cplex prog = + let + fun write_script s lp r = + let + val f = TextIO.openOut s + val _ = TextIO.output (f, "read\n"^lp^"\noptimize\nwrite\n"^r^"\nquit") + val _ = TextIO.closeOut f + in + () + end + + val name = string_of_int (Time.toMicroseconds (Time.now ())) + val lpname = tmp_file (name^".lp") + val resultname = tmp_file (name^".txt") + val scriptname = tmp_file (name^".script") + val _ = save_cplexFile lpname prog + val _ = write_script scriptname lpname resultname + in + let + val result = load_cplexResult resultname + val _ = OS.FileSys.remove lpname + val _ = OS.FileSys.remove resultname + val _ = OS.FileSys.remove scriptname + in + result + end + end + +fun solve prog = + case get_solver () of + SOLVER_DEFAULT => + (case getenv "LP_SOLVER" of + "CPLEX" => solve_cplex prog + | "GLPK" => solve_glpk prog + | _ => raise (Execute ("LP_SOLVER must be set to CPLEX or to GLPK"))) + | SOLVER_CPLEX => solve_cplex prog + | SOLVER_GLPK => solve_glpk prog + +end; diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/FloatSparseMatrixBuilder.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/FloatSparseMatrixBuilder.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,284 @@ +(* Title: HOL/Matrix/FloatSparseMatrixBuilder.ML + Author: Steven Obua +*) + +signature FLOAT_SPARSE_MATRIX_BUILDER = +sig + include MATRIX_BUILDER + + structure cplex : CPLEX + + type float = Float.float + val approx_value : int -> (float -> float) -> string -> term * term + val approx_vector : int -> (float -> float) -> vector -> term * term + val approx_matrix : int -> (float -> float) -> matrix -> term * term + + val mk_spvec_entry : int -> float -> term + val mk_spvec_entry' : int -> term -> term + val mk_spmat_entry : int -> term -> term + val spvecT: typ + val spmatT: typ + + val v_elem_at : vector -> int -> string option + val m_elem_at : matrix -> int -> vector option + val v_only_elem : vector -> int option + val v_fold : (int * string -> 'a -> 'a) -> vector -> 'a -> 'a + val m_fold : (int * vector -> 'a -> 'a) -> matrix -> 'a -> 'a + + val transpose_matrix : matrix -> matrix + + val cut_vector : int -> vector -> vector + val cut_matrix : vector -> int option -> matrix -> matrix + + val delete_matrix : int list -> matrix -> matrix + val cut_matrix' : int list -> matrix -> matrix + val delete_vector : int list -> vector -> vector + val cut_vector' : int list -> vector -> vector + + val indices_of_matrix : matrix -> int list + val indices_of_vector : vector -> int list + + (* cplexProg c A b *) + val cplexProg : vector -> matrix -> vector -> cplex.cplexProg * (string -> int) + (* dual_cplexProg c A b *) + val dual_cplexProg : vector -> matrix -> vector -> cplex.cplexProg * (string -> int) +end; + +structure FloatSparseMatrixBuilder : FLOAT_SPARSE_MATRIX_BUILDER = +struct + +type float = Float.float +structure Inttab = Table(type key = int val ord = rev_order o int_ord); + +type vector = string Inttab.table +type matrix = vector Inttab.table + +val spvec_elemT = HOLogic.mk_prodT (HOLogic.natT, HOLogic.realT); +val spvecT = HOLogic.listT spvec_elemT; +val spmat_elemT = HOLogic.mk_prodT (HOLogic.natT, spvecT); +val spmatT = HOLogic.listT spmat_elemT; + +fun approx_value prec f = + FloatArith.approx_float prec (fn (x, y) => (f x, f y)); + +fun mk_spvec_entry i f = + HOLogic.mk_prod (HOLogic.mk_number HOLogic.natT i, FloatArith.mk_float f); + +fun mk_spvec_entry' i x = + HOLogic.mk_prod (HOLogic.mk_number HOLogic.natT i, x); + +fun mk_spmat_entry i e = + HOLogic.mk_prod (HOLogic.mk_number HOLogic.natT i, e); + +fun approx_vector prec pprt vector = + let + fun app (index, s) (lower, upper) = + let + val (flower, fupper) = approx_value prec pprt s + val index = HOLogic.mk_number HOLogic.natT index + val elower = HOLogic.mk_prod (index, flower) + val eupper = HOLogic.mk_prod (index, fupper) + in (elower :: lower, eupper :: upper) end; + in + pairself (HOLogic.mk_list spvec_elemT) (Inttab.fold app vector ([], [])) + end; + +fun approx_matrix prec pprt vector = + let + fun app (index, v) (lower, upper) = + let + val (flower, fupper) = approx_vector prec pprt v + val index = HOLogic.mk_number HOLogic.natT index + val elower = HOLogic.mk_prod (index, flower) + val eupper = HOLogic.mk_prod (index, fupper) + in (elower :: lower, eupper :: upper) end; + in + pairself (HOLogic.mk_list spmat_elemT) (Inttab.fold app vector ([], [])) + end; + +exception Nat_expected of int; + +val zero_interval = approx_value 1 I "0" + +fun set_elem vector index str = + if index < 0 then + raise (Nat_expected index) + else if (approx_value 1 I str) = zero_interval then + vector + else + Inttab.update (index, str) vector + +fun set_vector matrix index vector = + if index < 0 then + raise (Nat_expected index) + else if Inttab.is_empty vector then + matrix + else + Inttab.update (index, vector) matrix + +val empty_matrix = Inttab.empty +val empty_vector = Inttab.empty + +(* dual stuff *) + +structure cplex = Cplex + +fun transpose_matrix matrix = + let + fun upd j (i, s) = + Inttab.map_default (i, Inttab.empty) (Inttab.update (j, s)); + fun updm (j, v) = Inttab.fold (upd j) v; + in Inttab.fold updm matrix empty_matrix end; + +exception No_name of string; + +exception Superfluous_constr_right_hand_sides + +fun cplexProg c A b = + let + val ytable = Unsynchronized.ref Inttab.empty + fun indexof s = + if String.size s = 0 then raise (No_name s) + else case Int.fromString (String.extract(s, 1, NONE)) of + SOME i => i | NONE => raise (No_name s) + + fun nameof i = + let + val s = "x" ^ string_of_int i + val _ = Unsynchronized.change ytable (Inttab.update (i, s)) + in + s + end + + fun split_numstr s = + if String.isPrefix "-" s then (false,String.extract(s, 1, NONE)) + else if String.isPrefix "+" s then (true, String.extract(s, 1, NONE)) + else (true, s) + + fun mk_term index s = + let + val (p, s) = split_numstr s + val prod = cplex.cplexProd (cplex.cplexNum s, cplex.cplexVar (nameof index)) + in + if p then prod else cplex.cplexNeg prod + end + + fun vec2sum vector = + cplex.cplexSum (Inttab.fold (fn (index, s) => fn list => (mk_term index s) :: list) vector []) + + fun mk_constr index vector c = + let + val s = case Inttab.lookup c index of SOME s => s | NONE => "0" + val (p, s) = split_numstr s + val num = if p then cplex.cplexNum s else cplex.cplexNeg (cplex.cplexNum s) + in + (NONE, cplex.cplexConstr (cplex.cplexLeq, (vec2sum vector, num))) + end + + fun delete index c = Inttab.delete index c handle Inttab.UNDEF _ => c + + val (list, b) = Inttab.fold + (fn (index, v) => fn (list, c) => ((mk_constr index v c)::list, delete index c)) + A ([], b) + val _ = if Inttab.is_empty b then () else raise Superfluous_constr_right_hand_sides + + fun mk_free y = cplex.cplexBounds (cplex.cplexNeg cplex.cplexInf, cplex.cplexLeq, + cplex.cplexVar y, cplex.cplexLeq, + cplex.cplexInf) + + val yvars = Inttab.fold (fn (_, y) => fn l => (mk_free y)::l) (!ytable) [] + + val prog = cplex.cplexProg ("original", cplex.cplexMaximize (vec2sum c), list, yvars) + in + (prog, indexof) + end + + +fun dual_cplexProg c A b = + let + fun indexof s = + if String.size s = 0 then raise (No_name s) + else case Int.fromString (String.extract(s, 1, NONE)) of + SOME i => i | NONE => raise (No_name s) + + fun nameof i = "y" ^ string_of_int i + + fun split_numstr s = + if String.isPrefix "-" s then (false,String.extract(s, 1, NONE)) + else if String.isPrefix "+" s then (true, String.extract(s, 1, NONE)) + else (true, s) + + fun mk_term index s = + let + val (p, s) = split_numstr s + val prod = cplex.cplexProd (cplex.cplexNum s, cplex.cplexVar (nameof index)) + in + if p then prod else cplex.cplexNeg prod + end + + fun vec2sum vector = + cplex.cplexSum (Inttab.fold (fn (index, s) => fn list => (mk_term index s)::list) vector []) + + fun mk_constr index vector c = + let + val s = case Inttab.lookup c index of SOME s => s | NONE => "0" + val (p, s) = split_numstr s + val num = if p then cplex.cplexNum s else cplex.cplexNeg (cplex.cplexNum s) + in + (NONE, cplex.cplexConstr (cplex.cplexEq, (vec2sum vector, num))) + end + + fun delete index c = Inttab.delete index c handle Inttab.UNDEF _ => c + + val (list, c) = Inttab.fold + (fn (index, v) => fn (list, c) => ((mk_constr index v c)::list, delete index c)) + (transpose_matrix A) ([], c) + val _ = if Inttab.is_empty c then () else raise Superfluous_constr_right_hand_sides + + val prog = cplex.cplexProg ("dual", cplex.cplexMinimize (vec2sum b), list, []) + in + (prog, indexof) + end + +fun cut_vector size v = + let + val count = Unsynchronized.ref 0; + fun app (i, s) = if (!count < size) then + (count := !count +1 ; Inttab.update (i, s)) + else I + in + Inttab.fold app v empty_vector + end + +fun cut_matrix vfilter vsize m = + let + fun app (i, v) = + if is_none (Inttab.lookup vfilter i) then I + else case vsize + of NONE => Inttab.update (i, v) + | SOME s => Inttab.update (i, cut_vector s v) + in Inttab.fold app m empty_matrix end + +fun v_elem_at v i = Inttab.lookup v i +fun m_elem_at m i = Inttab.lookup m i + +fun v_only_elem v = + case Inttab.min_key v of + NONE => NONE + | SOME vmin => (case Inttab.max_key v of + NONE => SOME vmin + | SOME vmax => if vmin = vmax then SOME vmin else NONE) + +fun v_fold f = Inttab.fold f; +fun m_fold f = Inttab.fold f; + +fun indices_of_vector v = Inttab.keys v +fun indices_of_matrix m = Inttab.keys m +fun delete_vector indices v = fold Inttab.delete indices v +fun delete_matrix indices m = fold Inttab.delete indices m +fun cut_matrix' indices _ = fold (fn i => fn m => (case Inttab.lookup m i of NONE => m | SOME v => Inttab.update (i, v) m)) indices Inttab.empty +fun cut_vector' indices _ = fold (fn i => fn v => (case Inttab.lookup v i of NONE => v | SOME x => Inttab.update (i, x) v)) indices Inttab.empty + + + +end; diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/LP.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/LP.thy Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,164 @@ +(* Title: HOL/Matrix/LP.thy + Author: Steven Obua +*) + +theory LP +imports Main "~~/src/HOL/Library/Lattice_Algebras" +begin + +lemma le_add_right_mono: + assumes + "a <= b + (c::'a::ordered_ab_group_add)" + "c <= d" + shows "a <= b + d" + apply (rule_tac order_trans[where y = "b+c"]) + apply (simp_all add: assms) + done + +lemma linprog_dual_estimate: + assumes + "A * x \ (b::'a::lattice_ring)" + "0 \ y" + "abs (A - A') \ \A" + "b \ b'" + "abs (c - c') \ \c" + "abs x \ r" + shows + "c * x \ y * b' + (y * \A + abs (y * A' - c') + \c) * r" +proof - + from assms have 1: "y * b <= y * b'" by (simp add: mult_left_mono) + from assms have 2: "y * (A * x) <= y * b" by (simp add: mult_left_mono) + have 3: "y * (A * x) = c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x" by (simp add: algebra_simps) + from 1 2 3 have 4: "c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x <= y * b'" by simp + have 5: "c * x <= y * b' + abs((y * (A - A') + (y * A' - c') + (c'-c)) * x)" + by (simp only: 4 estimate_by_abs) + have 6: "abs((y * (A - A') + (y * A' - c') + (c'-c)) * x) <= abs (y * (A - A') + (y * A' - c') + (c'-c)) * abs x" + by (simp add: abs_le_mult) + have 7: "(abs (y * (A - A') + (y * A' - c') + (c'-c))) * abs x <= (abs (y * (A-A') + (y*A'-c')) + abs(c'-c)) * abs x" + by(rule abs_triangle_ineq [THEN mult_right_mono]) simp + have 8: " (abs (y * (A-A') + (y*A'-c')) + abs(c'-c)) * abs x <= (abs (y * (A-A')) + abs (y*A'-c') + abs(c'-c)) * abs x" + by (simp add: abs_triangle_ineq mult_right_mono) + have 9: "(abs (y * (A-A')) + abs (y*A'-c') + abs(c'-c)) * abs x <= (abs y * abs (A-A') + abs (y*A'-c') + abs (c'-c)) * abs x" + by (simp add: abs_le_mult mult_right_mono) + have 10: "c'-c = -(c-c')" by (simp add: algebra_simps) + have 11: "abs (c'-c) = abs (c-c')" + by (subst 10, subst abs_minus_cancel, simp) + have 12: "(abs y * abs (A-A') + abs (y*A'-c') + abs (c'-c)) * abs x <= (abs y * abs (A-A') + abs (y*A'-c') + \c) * abs x" + by (simp add: 11 assms mult_right_mono) + have 13: "(abs y * abs (A-A') + abs (y*A'-c') + \c) * abs x <= (abs y * \A + abs (y*A'-c') + \c) * abs x" + by (simp add: assms mult_right_mono mult_left_mono) + have r: "(abs y * \A + abs (y*A'-c') + \c) * abs x <= (abs y * \A + abs (y*A'-c') + \c) * r" + apply (rule mult_left_mono) + apply (simp add: assms) + apply (rule_tac add_mono[of "0::'a" _ "0", simplified])+ + apply (rule mult_left_mono[of "0" "\A", simplified]) + apply (simp_all) + apply (rule order_trans[where y="abs (A-A')"], simp_all add: assms) + apply (rule order_trans[where y="abs (c-c')"], simp_all add: assms) + done + from 6 7 8 9 12 13 r have 14:" abs((y * (A - A') + (y * A' - c') + (c'-c)) * x) <=(abs y * \A + abs (y*A'-c') + \c) * r" + by (simp) + show ?thesis + apply (rule le_add_right_mono[of _ _ "abs((y * (A - A') + (y * A' - c') + (c'-c)) * x)"]) + apply (simp_all only: 5 14[simplified abs_of_nonneg[of y, simplified assms]]) + done +qed + +lemma le_ge_imp_abs_diff_1: + assumes + "A1 <= (A::'a::lattice_ring)" + "A <= A2" + shows "abs (A-A1) <= A2-A1" +proof - + have "0 <= A - A1" + proof - + have 1: "A - A1 = A + (- A1)" by simp + show ?thesis by (simp only: 1 add_right_mono[of A1 A "-A1", simplified, simplified assms]) + qed + then have "abs (A-A1) = A-A1" by (rule abs_of_nonneg) + with assms show "abs (A-A1) <= (A2-A1)" by simp +qed + +lemma mult_le_prts: + assumes + "a1 <= (a::'a::lattice_ring)" + "a <= a2" + "b1 <= b" + "b <= b2" + shows + "a * b <= pprt a2 * pprt b2 + pprt a1 * nprt b2 + nprt a2 * pprt b1 + nprt a1 * nprt b1" +proof - + have "a * b = (pprt a + nprt a) * (pprt b + nprt b)" + apply (subst prts[symmetric])+ + apply simp + done + then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b" + by (simp add: algebra_simps) + moreover have "pprt a * pprt b <= pprt a2 * pprt b2" + by (simp_all add: assms mult_mono) + moreover have "pprt a * nprt b <= pprt a1 * nprt b2" + proof - + have "pprt a * nprt b <= pprt a * nprt b2" + by (simp add: mult_left_mono assms) + moreover have "pprt a * nprt b2 <= pprt a1 * nprt b2" + by (simp add: mult_right_mono_neg assms) + ultimately show ?thesis + by simp + qed + moreover have "nprt a * pprt b <= nprt a2 * pprt b1" + proof - + have "nprt a * pprt b <= nprt a2 * pprt b" + by (simp add: mult_right_mono assms) + moreover have "nprt a2 * pprt b <= nprt a2 * pprt b1" + by (simp add: mult_left_mono_neg assms) + ultimately show ?thesis + by simp + qed + moreover have "nprt a * nprt b <= nprt a1 * nprt b1" + proof - + have "nprt a * nprt b <= nprt a * nprt b1" + by (simp add: mult_left_mono_neg assms) + moreover have "nprt a * nprt b1 <= nprt a1 * nprt b1" + by (simp add: mult_right_mono_neg assms) + ultimately show ?thesis + by simp + qed + ultimately show ?thesis + by - (rule add_mono | simp)+ +qed + +lemma mult_le_dual_prts: + assumes + "A * x \ (b::'a::lattice_ring)" + "0 \ y" + "A1 \ A" + "A \ A2" + "c1 \ c" + "c \ c2" + "r1 \ x" + "x \ r2" + shows + "c * x \ y * b + (let s1 = c1 - y * A2; s2 = c2 - y * A1 in pprt s2 * pprt r2 + pprt s1 * nprt r2 + nprt s2 * pprt r1 + nprt s1 * nprt r1)" + (is "_ <= _ + ?C") +proof - + from assms have "y * (A * x) <= y * b" by (simp add: mult_left_mono) + moreover have "y * (A * x) = c * x + (y * A - c) * x" by (simp add: algebra_simps) + ultimately have "c * x + (y * A - c) * x <= y * b" by simp + then have "c * x <= y * b - (y * A - c) * x" by (simp add: le_diff_eq) + then have cx: "c * x <= y * b + (c - y * A) * x" by (simp add: algebra_simps) + have s2: "c - y * A <= c2 - y * A1" + by (simp add: diff_minus assms add_mono mult_left_mono) + have s1: "c1 - y * A2 <= c - y * A" + by (simp add: diff_minus assms add_mono mult_left_mono) + have prts: "(c - y * A) * x <= ?C" + apply (simp add: Let_def) + apply (rule mult_le_prts) + apply (simp_all add: assms s1 s2) + done + then have "y * b + (c - y * A) * x <= y * b + ?C" + by simp + with cx show ?thesis + by(simp only:) +qed + +end \ No newline at end of file diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/Matrix.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/Matrix.thy Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,1836 @@ +(* Title: HOL/Matrix/Matrix.thy + Author: Steven Obua +*) + +theory Matrix +imports Main "~~/src/HOL/Library/Lattice_Algebras" +begin + +type_synonym 'a infmatrix = "nat \ nat \ 'a" + +definition nonzero_positions :: "(nat \ nat \ 'a::zero) \ (nat \ nat) set" where + "nonzero_positions A = {pos. A (fst pos) (snd pos) ~= 0}" + +definition "matrix = {(f::(nat \ nat \ 'a::zero)). finite (nonzero_positions f)}" + +typedef (open) 'a matrix = "matrix :: (nat \ nat \ 'a::zero) set" + unfolding matrix_def +proof + show "(\j i. 0) \ {(f::(nat \ nat \ 'a::zero)). finite (nonzero_positions f)}" + by (simp add: nonzero_positions_def) +qed + +declare Rep_matrix_inverse[simp] + +lemma finite_nonzero_positions : "finite (nonzero_positions (Rep_matrix A))" + by (induct A) (simp add: Abs_matrix_inverse matrix_def) + +definition nrows :: "('a::zero) matrix \ nat" where + "nrows A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image fst) (nonzero_positions (Rep_matrix A))))" + +definition ncols :: "('a::zero) matrix \ nat" where + "ncols A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image snd) (nonzero_positions (Rep_matrix A))))" + +lemma nrows: + assumes hyp: "nrows A \ m" + shows "(Rep_matrix A m n) = 0" +proof cases + assume "nonzero_positions(Rep_matrix A) = {}" + then show "(Rep_matrix A m n) = 0" by (simp add: nonzero_positions_def) +next + assume a: "nonzero_positions(Rep_matrix A) \ {}" + let ?S = "fst`(nonzero_positions(Rep_matrix A))" + have c: "finite (?S)" by (simp add: finite_nonzero_positions) + from hyp have d: "Max (?S) < m" by (simp add: a nrows_def) + have "m \ ?S" + proof - + have "m \ ?S \ m <= Max(?S)" by (simp add: Max_ge [OF c]) + moreover from d have "~(m <= Max ?S)" by (simp) + ultimately show "m \ ?S" by (auto) + qed + thus "Rep_matrix A m n = 0" by (simp add: nonzero_positions_def image_Collect) +qed + +definition transpose_infmatrix :: "'a infmatrix \ 'a infmatrix" where + "transpose_infmatrix A j i == A i j" + +definition transpose_matrix :: "('a::zero) matrix \ 'a matrix" where + "transpose_matrix == Abs_matrix o transpose_infmatrix o Rep_matrix" + +declare transpose_infmatrix_def[simp] + +lemma transpose_infmatrix_twice[simp]: "transpose_infmatrix (transpose_infmatrix A) = A" +by ((rule ext)+, simp) + +lemma transpose_infmatrix: "transpose_infmatrix (% j i. P j i) = (% j i. P i j)" + apply (rule ext)+ + by simp + +lemma transpose_infmatrix_closed[simp]: "Rep_matrix (Abs_matrix (transpose_infmatrix (Rep_matrix x))) = transpose_infmatrix (Rep_matrix x)" +apply (rule Abs_matrix_inverse) +apply (simp add: matrix_def nonzero_positions_def image_def) +proof - + let ?A = "{pos. Rep_matrix x (snd pos) (fst pos) \ 0}" + let ?swap = "% pos. (snd pos, fst pos)" + let ?B = "{pos. Rep_matrix x (fst pos) (snd pos) \ 0}" + have swap_image: "?swap`?A = ?B" + apply (simp add: image_def) + apply (rule set_eqI) + apply (simp) + proof + fix y + assume hyp: "\a b. Rep_matrix x b a \ 0 \ y = (b, a)" + thus "Rep_matrix x (fst y) (snd y) \ 0" + proof - + from hyp obtain a b where "(Rep_matrix x b a \ 0 & y = (b,a))" by blast + then show "Rep_matrix x (fst y) (snd y) \ 0" by (simp) + qed + next + fix y + assume hyp: "Rep_matrix x (fst y) (snd y) \ 0" + show "\ a b. (Rep_matrix x b a \ 0 & y = (b,a))" + by (rule exI[of _ "snd y"], rule exI[of _ "fst y"]) (simp add: hyp) + qed + then have "finite (?swap`?A)" + proof - + have "finite (nonzero_positions (Rep_matrix x))" by (simp add: finite_nonzero_positions) + then have "finite ?B" by (simp add: nonzero_positions_def) + with swap_image show "finite (?swap`?A)" by (simp) + qed + moreover + have "inj_on ?swap ?A" by (simp add: inj_on_def) + ultimately show "finite ?A"by (rule finite_imageD[of ?swap ?A]) +qed + +lemma infmatrixforward: "(x::'a infmatrix) = y \ \ a b. x a b = y a b" by auto + +lemma transpose_infmatrix_inject: "(transpose_infmatrix A = transpose_infmatrix B) = (A = B)" +apply (auto) +apply (rule ext)+ +apply (simp add: transpose_infmatrix) +apply (drule infmatrixforward) +apply (simp) +done + +lemma transpose_matrix_inject: "(transpose_matrix A = transpose_matrix B) = (A = B)" +apply (simp add: transpose_matrix_def) +apply (subst Rep_matrix_inject[THEN sym])+ +apply (simp only: transpose_infmatrix_closed transpose_infmatrix_inject) +done + +lemma transpose_matrix[simp]: "Rep_matrix(transpose_matrix A) j i = Rep_matrix A i j" +by (simp add: transpose_matrix_def) + +lemma transpose_transpose_id[simp]: "transpose_matrix (transpose_matrix A) = A" +by (simp add: transpose_matrix_def) + +lemma nrows_transpose[simp]: "nrows (transpose_matrix A) = ncols A" +by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def) + +lemma ncols_transpose[simp]: "ncols (transpose_matrix A) = nrows A" +by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def) + +lemma ncols: "ncols A <= n \ Rep_matrix A m n = 0" +proof - + assume "ncols A <= n" + then have "nrows (transpose_matrix A) <= n" by (simp) + then have "Rep_matrix (transpose_matrix A) n m = 0" by (rule nrows) + thus "Rep_matrix A m n = 0" by (simp add: transpose_matrix_def) +qed + +lemma ncols_le: "(ncols A <= n) = (! j i. n <= i \ (Rep_matrix A j i) = 0)" (is "_ = ?st") +apply (auto) +apply (simp add: ncols) +proof (simp add: ncols_def, auto) + let ?P = "nonzero_positions (Rep_matrix A)" + let ?p = "snd`?P" + have a:"finite ?p" by (simp add: finite_nonzero_positions) + let ?m = "Max ?p" + assume "~(Suc (?m) <= n)" + then have b:"n <= ?m" by (simp) + fix a b + assume "(a,b) \ ?P" + then have "?p \ {}" by (auto) + with a have "?m \ ?p" by (simp) + moreover have "!x. (x \ ?p \ (? y. (Rep_matrix A y x) \ 0))" by (simp add: nonzero_positions_def image_def) + ultimately have "? y. (Rep_matrix A y ?m) \ 0" by (simp) + moreover assume ?st + ultimately show "False" using b by (simp) +qed + +lemma less_ncols: "(n < ncols A) = (? j i. n <= i & (Rep_matrix A j i) \ 0)" +proof - + have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith + show ?thesis by (simp add: a ncols_le) +qed + +lemma le_ncols: "(n <= ncols A) = (\ m. (\ j i. m <= i \ (Rep_matrix A j i) = 0) \ n <= m)" +apply (auto) +apply (subgoal_tac "ncols A <= m") +apply (simp) +apply (simp add: ncols_le) +apply (drule_tac x="ncols A" in spec) +by (simp add: ncols) + +lemma nrows_le: "(nrows A <= n) = (! j i. n <= j \ (Rep_matrix A j i) = 0)" (is ?s) +proof - + have "(nrows A <= n) = (ncols (transpose_matrix A) <= n)" by (simp) + also have "\ = (! j i. n <= i \ (Rep_matrix (transpose_matrix A) j i = 0))" by (rule ncols_le) + also have "\ = (! j i. n <= i \ (Rep_matrix A i j) = 0)" by (simp) + finally show "(nrows A <= n) = (! j i. n <= j \ (Rep_matrix A j i) = 0)" by (auto) +qed + +lemma less_nrows: "(m < nrows A) = (? j i. m <= j & (Rep_matrix A j i) \ 0)" +proof - + have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith + show ?thesis by (simp add: a nrows_le) +qed + +lemma le_nrows: "(n <= nrows A) = (\ m. (\ j i. m <= j \ (Rep_matrix A j i) = 0) \ n <= m)" +apply (auto) +apply (subgoal_tac "nrows A <= m") +apply (simp) +apply (simp add: nrows_le) +apply (drule_tac x="nrows A" in spec) +by (simp add: nrows) + +lemma nrows_notzero: "Rep_matrix A m n \ 0 \ m < nrows A" +apply (case_tac "nrows A <= m") +apply (simp_all add: nrows) +done + +lemma ncols_notzero: "Rep_matrix A m n \ 0 \ n < ncols A" +apply (case_tac "ncols A <= n") +apply (simp_all add: ncols) +done + +lemma finite_natarray1: "finite {x. x < (n::nat)}" +apply (induct n) +apply (simp) +proof - + fix n + have "{x. x < Suc n} = insert n {x. x < n}" by (rule set_eqI, simp, arith) + moreover assume "finite {x. x < n}" + ultimately show "finite {x. x < Suc n}" by (simp) +qed + +lemma finite_natarray2: "finite {pos. (fst pos) < (m::nat) & (snd pos) < (n::nat)}" + apply (induct m) + apply (simp+) + proof - + fix m::nat + let ?s0 = "{pos. fst pos < m & snd pos < n}" + let ?s1 = "{pos. fst pos < (Suc m) & snd pos < n}" + let ?sd = "{pos. fst pos = m & snd pos < n}" + assume f0: "finite ?s0" + have f1: "finite ?sd" + proof - + let ?f = "% x. (m, x)" + have "{pos. fst pos = m & snd pos < n} = ?f ` {x. x < n}" by (rule set_eqI, simp add: image_def, auto) + moreover have "finite {x. x < n}" by (simp add: finite_natarray1) + ultimately show "finite {pos. fst pos = m & snd pos < n}" by (simp) + qed + have su: "?s0 \ ?sd = ?s1" by (rule set_eqI, simp, arith) + from f0 f1 have "finite (?s0 \ ?sd)" by (rule finite_UnI) + with su show "finite ?s1" by (simp) +qed + +lemma RepAbs_matrix: + assumes aem: "? m. ! j i. m <= j \ x j i = 0" (is ?em) and aen:"? n. ! j i. (n <= i \ x j i = 0)" (is ?en) + shows "(Rep_matrix (Abs_matrix x)) = x" +apply (rule Abs_matrix_inverse) +apply (simp add: matrix_def nonzero_positions_def) +proof - + from aem obtain m where a: "! j i. m <= j \ x j i = 0" by (blast) + from aen obtain n where b: "! j i. n <= i \ x j i = 0" by (blast) + let ?u = "{pos. x (fst pos) (snd pos) \ 0}" + let ?v = "{pos. fst pos < m & snd pos < n}" + have c: "!! (m::nat) a. ~(m <= a) \ a < m" by (arith) + from a b have "(?u \ (-?v)) = {}" + apply (simp) + apply (rule set_eqI) + apply (simp) + apply auto + by (rule c, auto)+ + then have d: "?u \ ?v" by blast + moreover have "finite ?v" by (simp add: finite_natarray2) + ultimately show "finite ?u" by (rule finite_subset) +qed + +definition apply_infmatrix :: "('a \ 'b) \ 'a infmatrix \ 'b infmatrix" where + "apply_infmatrix f == % A. (% j i. f (A j i))" + +definition apply_matrix :: "('a \ 'b) \ ('a::zero) matrix \ ('b::zero) matrix" where + "apply_matrix f == % A. Abs_matrix (apply_infmatrix f (Rep_matrix A))" + +definition combine_infmatrix :: "('a \ 'b \ 'c) \ 'a infmatrix \ 'b infmatrix \ 'c infmatrix" where + "combine_infmatrix f == % A B. (% j i. f (A j i) (B j i))" + +definition combine_matrix :: "('a \ 'b \ 'c) \ ('a::zero) matrix \ ('b::zero) matrix \ ('c::zero) matrix" where + "combine_matrix f == % A B. Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))" + +lemma expand_apply_infmatrix[simp]: "apply_infmatrix f A j i = f (A j i)" +by (simp add: apply_infmatrix_def) + +lemma expand_combine_infmatrix[simp]: "combine_infmatrix f A B j i = f (A j i) (B j i)" +by (simp add: combine_infmatrix_def) + +definition commutative :: "('a \ 'a \ 'b) \ bool" where +"commutative f == ! x y. f x y = f y x" + +definition associative :: "('a \ 'a \ 'a) \ bool" where +"associative f == ! x y z. f (f x y) z = f x (f y z)" + +text{* +To reason about associativity and commutativity of operations on matrices, +let's take a step back and look at the general situtation: Assume that we have +sets $A$ and $B$ with $B \subset A$ and an abstraction $u: A \rightarrow B$. This abstraction has to fulfill $u(b) = b$ for all $b \in B$, but is arbitrary otherwise. +Each function $f: A \times A \rightarrow A$ now induces a function $f': B \times B \rightarrow B$ by $f' = u \circ f$. +It is obvious that commutativity of $f$ implies commutativity of $f'$: $f' x y = u (f x y) = u (f y x) = f' y x.$ +*} + +lemma combine_infmatrix_commute: + "commutative f \ commutative (combine_infmatrix f)" +by (simp add: commutative_def combine_infmatrix_def) + +lemma combine_matrix_commute: +"commutative f \ commutative (combine_matrix f)" +by (simp add: combine_matrix_def commutative_def combine_infmatrix_def) + +text{* +On the contrary, given an associative function $f$ we cannot expect $f'$ to be associative. A counterexample is given by $A=\ganz$, $B=\{-1, 0, 1\}$, +as $f$ we take addition on $\ganz$, which is clearly associative. The abstraction is given by $u(a) = 0$ for $a \notin B$. Then we have +\[ f' (f' 1 1) -1 = u(f (u (f 1 1)) -1) = u(f (u 2) -1) = u (f 0 -1) = -1, \] +but on the other hand we have +\[ f' 1 (f' 1 -1) = u (f 1 (u (f 1 -1))) = u (f 1 0) = 1.\] +A way out of this problem is to assume that $f(A\times A)\subset A$ holds, and this is what we are going to do: +*} + +lemma nonzero_positions_combine_infmatrix[simp]: "f 0 0 = 0 \ nonzero_positions (combine_infmatrix f A B) \ (nonzero_positions A) \ (nonzero_positions B)" +by (rule subsetI, simp add: nonzero_positions_def combine_infmatrix_def, auto) + +lemma finite_nonzero_positions_Rep[simp]: "finite (nonzero_positions (Rep_matrix A))" +by (insert Rep_matrix [of A], simp add: matrix_def) + +lemma combine_infmatrix_closed [simp]: + "f 0 0 = 0 \ Rep_matrix (Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))) = combine_infmatrix f (Rep_matrix A) (Rep_matrix B)" +apply (rule Abs_matrix_inverse) +apply (simp add: matrix_def) +apply (rule finite_subset[of _ "(nonzero_positions (Rep_matrix A)) \ (nonzero_positions (Rep_matrix B))"]) +by (simp_all) + +text {* We need the next two lemmas only later, but it is analog to the above one, so we prove them now: *} +lemma nonzero_positions_apply_infmatrix[simp]: "f 0 = 0 \ nonzero_positions (apply_infmatrix f A) \ nonzero_positions A" +by (rule subsetI, simp add: nonzero_positions_def apply_infmatrix_def, auto) + +lemma apply_infmatrix_closed [simp]: + "f 0 = 0 \ Rep_matrix (Abs_matrix (apply_infmatrix f (Rep_matrix A))) = apply_infmatrix f (Rep_matrix A)" +apply (rule Abs_matrix_inverse) +apply (simp add: matrix_def) +apply (rule finite_subset[of _ "nonzero_positions (Rep_matrix A)"]) +by (simp_all) + +lemma combine_infmatrix_assoc[simp]: "f 0 0 = 0 \ associative f \ associative (combine_infmatrix f)" +by (simp add: associative_def combine_infmatrix_def) + +lemma comb: "f = g \ x = y \ f x = g y" +by (auto) + +lemma combine_matrix_assoc: "f 0 0 = 0 \ associative f \ associative (combine_matrix f)" +apply (simp(no_asm) add: associative_def combine_matrix_def, auto) +apply (rule comb [of Abs_matrix Abs_matrix]) +by (auto, insert combine_infmatrix_assoc[of f], simp add: associative_def) + +lemma Rep_apply_matrix[simp]: "f 0 = 0 \ Rep_matrix (apply_matrix f A) j i = f (Rep_matrix A j i)" +by (simp add: apply_matrix_def) + +lemma Rep_combine_matrix[simp]: "f 0 0 = 0 \ Rep_matrix (combine_matrix f A B) j i = f (Rep_matrix A j i) (Rep_matrix B j i)" + by(simp add: combine_matrix_def) + +lemma combine_nrows_max: "f 0 0 = 0 \ nrows (combine_matrix f A B) <= max (nrows A) (nrows B)" +by (simp add: nrows_le) + +lemma combine_ncols_max: "f 0 0 = 0 \ ncols (combine_matrix f A B) <= max (ncols A) (ncols B)" +by (simp add: ncols_le) + +lemma combine_nrows: "f 0 0 = 0 \ nrows A <= q \ nrows B <= q \ nrows(combine_matrix f A B) <= q" + by (simp add: nrows_le) + +lemma combine_ncols: "f 0 0 = 0 \ ncols A <= q \ ncols B <= q \ ncols(combine_matrix f A B) <= q" + by (simp add: ncols_le) + +definition zero_r_neutral :: "('a \ 'b::zero \ 'a) \ bool" where + "zero_r_neutral f == ! a. f a 0 = a" + +definition zero_l_neutral :: "('a::zero \ 'b \ 'b) \ bool" where + "zero_l_neutral f == ! a. f 0 a = a" + +definition zero_closed :: "(('a::zero) \ ('b::zero) \ ('c::zero)) \ bool" where + "zero_closed f == (!x. f x 0 = 0) & (!y. f 0 y = 0)" + +primrec foldseq :: "('a \ 'a \ 'a) \ (nat \ 'a) \ nat \ 'a" +where + "foldseq f s 0 = s 0" +| "foldseq f s (Suc n) = f (s 0) (foldseq f (% k. s(Suc k)) n)" + +primrec foldseq_transposed :: "('a \ 'a \ 'a) \ (nat \ 'a) \ nat \ 'a" +where + "foldseq_transposed f s 0 = s 0" +| "foldseq_transposed f s (Suc n) = f (foldseq_transposed f s n) (s (Suc n))" + +lemma foldseq_assoc : "associative f \ foldseq f = foldseq_transposed f" +proof - + assume a:"associative f" + then have sublemma: "!! n. ! N s. N <= n \ foldseq f s N = foldseq_transposed f s N" + proof - + fix n + show "!N s. N <= n \ foldseq f s N = foldseq_transposed f s N" + proof (induct n) + show "!N s. N <= 0 \ foldseq f s N = foldseq_transposed f s N" by simp + next + fix n + assume b:"! N s. N <= n \ foldseq f s N = foldseq_transposed f s N" + have c:"!!N s. N <= n \ foldseq f s N = foldseq_transposed f s N" by (simp add: b) + show "! N t. N <= Suc n \ foldseq f t N = foldseq_transposed f t N" + proof (auto) + fix N t + assume Nsuc: "N <= Suc n" + show "foldseq f t N = foldseq_transposed f t N" + proof cases + assume "N <= n" + then show "foldseq f t N = foldseq_transposed f t N" by (simp add: b) + next + assume "~(N <= n)" + with Nsuc have Nsuceq: "N = Suc n" by simp + have neqz: "n \ 0 \ ? m. n = Suc m & Suc m <= n" by arith + have assocf: "!! x y z. f x (f y z) = f (f x y) z" by (insert a, simp add: associative_def) + show "foldseq f t N = foldseq_transposed f t N" + apply (simp add: Nsuceq) + apply (subst c) + apply (simp) + apply (case_tac "n = 0") + apply (simp) + apply (drule neqz) + apply (erule exE) + apply (simp) + apply (subst assocf) + proof - + fix m + assume "n = Suc m & Suc m <= n" + then have mless: "Suc m <= n" by arith + then have step1: "foldseq_transposed f (% k. t (Suc k)) m = foldseq f (% k. t (Suc k)) m" (is "?T1 = ?T2") + apply (subst c) + by simp+ + have step2: "f (t 0) ?T2 = foldseq f t (Suc m)" (is "_ = ?T3") by simp + have step3: "?T3 = foldseq_transposed f t (Suc m)" (is "_ = ?T4") + apply (subst c) + by (simp add: mless)+ + have step4: "?T4 = f (foldseq_transposed f t m) (t (Suc m))" (is "_=?T5") by simp + from step1 step2 step3 step4 show sowhat: "f (f (t 0) ?T1) (t (Suc (Suc m))) = f ?T5 (t (Suc (Suc m)))" by simp + qed + qed + qed + qed + qed + show "foldseq f = foldseq_transposed f" by ((rule ext)+, insert sublemma, auto) + qed + +lemma foldseq_distr: "\associative f; commutative f\ \ foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" +proof - + assume assoc: "associative f" + assume comm: "commutative f" + from assoc have a:"!! x y z. f (f x y) z = f x (f y z)" by (simp add: associative_def) + from comm have b: "!! x y. f x y = f y x" by (simp add: commutative_def) + from assoc comm have c: "!! x y z. f x (f y z) = f y (f x z)" by (simp add: commutative_def associative_def) + have "!! n. (! u v. foldseq f (%k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))" + apply (induct_tac n) + apply (simp+, auto) + by (simp add: a b c) + then show "foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" by simp +qed + +theorem "\associative f; associative g; \a b c d. g (f a b) (f c d) = f (g a c) (g b d); ? x y. (f x) \ (f y); ? x y. (g x) \ (g y); f x x = x; g x x = x\ \ f=g | (! y. f y x = y) | (! y. g y x = y)" +oops +(* Model found + +Trying to find a model that refutes: \associative f; associative g; + \a b c d. g (f a b) (f c d) = f (g a c) (g b d); \x y. f x \ f y; + \x y. g x \ g y; f x x = x; g x x = x\ +\ f = g \ (\y. f y x = y) \ (\y. g y x = y) +Searching for a model of size 1, translating term... invoking SAT solver... no model found. +Searching for a model of size 2, translating term... invoking SAT solver... no model found. +Searching for a model of size 3, translating term... invoking SAT solver... +Model found: +Size of types: 'a: 3 +x: a1 +g: (a0\(a0\a1, a1\a0, a2\a1), a1\(a0\a0, a1\a1, a2\a0), a2\(a0\a1, a1\a0, a2\a1)) +f: (a0\(a0\a0, a1\a0, a2\a0), a1\(a0\a1, a1\a1, a2\a1), a2\(a0\a0, a1\a0, a2\a0)) +*) + +lemma foldseq_zero: +assumes fz: "f 0 0 = 0" and sz: "! i. i <= n \ s i = 0" +shows "foldseq f s n = 0" +proof - + have "!! n. ! s. (! i. i <= n \ s i = 0) \ foldseq f s n = 0" + apply (induct_tac n) + apply (simp) + by (simp add: fz) + then show "foldseq f s n = 0" by (simp add: sz) +qed + +lemma foldseq_significant_positions: + assumes p: "! i. i <= N \ S i = T i" + shows "foldseq f S N = foldseq f T N" +proof - + have "!! m . ! s t. (! i. i<=m \ s i = t i) \ foldseq f s m = foldseq f t m" + apply (induct_tac m) + apply (simp) + apply (simp) + apply (auto) + proof - + fix n + fix s::"nat\'a" + fix t::"nat\'a" + assume a: "\s t. (\i\n. s i = t i) \ foldseq f s n = foldseq f t n" + assume b: "\i\Suc n. s i = t i" + have c:"!! a b. a = b \ f (t 0) a = f (t 0) b" by blast + have d:"!! s t. (\i\n. s i = t i) \ foldseq f s n = foldseq f t n" by (simp add: a) + show "f (t 0) (foldseq f (\k. s (Suc k)) n) = f (t 0) (foldseq f (\k. t (Suc k)) n)" by (rule c, simp add: d b) + qed + with p show ?thesis by simp +qed + +lemma foldseq_tail: + assumes "M <= N" + shows "foldseq f S N = foldseq f (% k. (if k < M then (S k) else (foldseq f (% k. S(k+M)) (N-M)))) M" +proof - + have suc: "!! a b. \a <= Suc b; a \ Suc b\ \ a <= b" by arith + have a:"!! a b c . a = b \ f c a = f c b" by blast + have "!! n. ! m s. m <= n \ foldseq f s n = foldseq f (% k. (if k < m then (s k) else (foldseq f (% k. s(k+m)) (n-m)))) m" + apply (induct_tac n) + apply (simp) + apply (simp) + apply (auto) + apply (case_tac "m = Suc na") + apply (simp) + apply (rule a) + apply (rule foldseq_significant_positions) + apply (auto) + apply (drule suc, simp+) + proof - + fix na m s + assume suba:"\m\na. \s. foldseq f s na = foldseq f (\k. if k < m then s k else foldseq f (\k. s (k + m)) (na - m))m" + assume subb:"m <= na" + from suba have subc:"!! m s. m <= na \foldseq f s na = foldseq f (\k. if k < m then s k else foldseq f (\k. s (k + m)) (na - m))m" by simp + have subd: "foldseq f (\k. if k < m then s (Suc k) else foldseq f (\k. s (Suc (k + m))) (na - m)) m = + foldseq f (% k. s(Suc k)) na" + by (rule subc[of m "% k. s(Suc k)", THEN sym], simp add: subb) + from subb have sube: "m \ 0 \ ? mm. m = Suc mm & mm <= na" by arith + show "f (s 0) (foldseq f (\k. if k < m then s (Suc k) else foldseq f (\k. s (Suc (k + m))) (na - m)) m) = + foldseq f (\k. if k < m then s k else foldseq f (\k. s (k + m)) (Suc na - m)) m" + apply (simp add: subd) + apply (cases "m = 0") + apply (simp) + apply (drule sube) + apply (auto) + apply (rule a) + by (simp add: subc cong del: if_cong) + qed + then show ?thesis using assms by simp +qed + +lemma foldseq_zerotail: + assumes + fz: "f 0 0 = 0" + and sz: "! i. n <= i \ s i = 0" + and nm: "n <= m" + shows + "foldseq f s n = foldseq f s m" +proof - + show "foldseq f s n = foldseq f s m" + apply (simp add: foldseq_tail[OF nm, of f s]) + apply (rule foldseq_significant_positions) + apply (auto) + apply (subst foldseq_zero) + by (simp add: fz sz)+ +qed + +lemma foldseq_zerotail2: + assumes "! x. f x 0 = x" + and "! i. n < i \ s i = 0" + and nm: "n <= m" + shows "foldseq f s n = foldseq f s m" +proof - + have "f 0 0 = 0" by (simp add: assms) + have b:"!! m n. n <= m \ m \ n \ ? k. m-n = Suc k" by arith + have c: "0 <= m" by simp + have d: "!! k. k \ 0 \ ? l. k = Suc l" by arith + show ?thesis + apply (subst foldseq_tail[OF nm]) + apply (rule foldseq_significant_positions) + apply (auto) + apply (case_tac "m=n") + apply (simp+) + apply (drule b[OF nm]) + apply (auto) + apply (case_tac "k=0") + apply (simp add: assms) + apply (drule d) + apply (auto) + apply (simp add: assms foldseq_zero) + done +qed + +lemma foldseq_zerostart: + "! x. f 0 (f 0 x) = f 0 x \ ! i. i <= n \ s i = 0 \ foldseq f s (Suc n) = f 0 (s (Suc n))" +proof - + assume f00x: "! x. f 0 (f 0 x) = f 0 x" + have "! s. (! i. i<=n \ s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n))" + apply (induct n) + apply (simp) + apply (rule allI, rule impI) + proof - + fix n + fix s + have a:"foldseq f s (Suc (Suc n)) = f (s 0) (foldseq f (% k. s(Suc k)) (Suc n))" by simp + assume b: "! s. ((\i\n. s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n)))" + from b have c:"!! s. (\i\n. s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n))" by simp + assume d: "! i. i <= Suc n \ s i = 0" + show "foldseq f s (Suc (Suc n)) = f 0 (s (Suc (Suc n)))" + apply (subst a) + apply (subst c) + by (simp add: d f00x)+ + qed + then show "! i. i <= n \ s i = 0 \ foldseq f s (Suc n) = f 0 (s (Suc n))" by simp +qed + +lemma foldseq_zerostart2: + "! x. f 0 x = x \ ! i. i < n \ s i = 0 \ foldseq f s n = s n" +proof - + assume a:"! i. i s i = 0" + assume x:"! x. f 0 x = x" + from x have f00x: "! x. f 0 (f 0 x) = f 0 x" by blast + have b: "!! i l. i < Suc l = (i <= l)" by arith + have d: "!! k. k \ 0 \ ? l. k = Suc l" by arith + show "foldseq f s n = s n" + apply (case_tac "n=0") + apply (simp) + apply (insert a) + apply (drule d) + apply (auto) + apply (simp add: b) + apply (insert f00x) + apply (drule foldseq_zerostart) + by (simp add: x)+ +qed + +lemma foldseq_almostzero: + assumes f0x:"! x. f 0 x = x" and fx0: "! x. f x 0 = x" and s0:"! i. i \ j \ s i = 0" + shows "foldseq f s n = (if (j <= n) then (s j) else 0)" +proof - + from s0 have a: "! i. i < j \ s i = 0" by simp + from s0 have b: "! i. j < i \ s i = 0" by simp + show ?thesis + apply auto + apply (subst foldseq_zerotail2[of f, OF fx0, of j, OF b, of n, THEN sym]) + apply simp + apply (subst foldseq_zerostart2) + apply (simp add: f0x a)+ + apply (subst foldseq_zero) + by (simp add: s0 f0x)+ +qed + +lemma foldseq_distr_unary: + assumes "!! a b. g (f a b) = f (g a) (g b)" + shows "g(foldseq f s n) = foldseq f (% x. g(s x)) n" +proof - + have "! s. g(foldseq f s n) = foldseq f (% x. g(s x)) n" + apply (induct_tac n) + apply (simp) + apply (simp) + apply (auto) + apply (drule_tac x="% k. s (Suc k)" in spec) + by (simp add: assms) + then show ?thesis by simp +qed + +definition mult_matrix_n :: "nat \ (('a::zero) \ ('b::zero) \ ('c::zero)) \ ('c \ 'c \ 'c) \ 'a matrix \ 'b matrix \ 'c matrix" where + "mult_matrix_n n fmul fadd A B == Abs_matrix(% j i. foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) n)" + +definition mult_matrix :: "(('a::zero) \ ('b::zero) \ ('c::zero)) \ ('c \ 'c \ 'c) \ 'a matrix \ 'b matrix \ 'c matrix" where + "mult_matrix fmul fadd A B == mult_matrix_n (max (ncols A) (nrows B)) fmul fadd A B" + +lemma mult_matrix_n: + assumes "ncols A \ n" (is ?An) "nrows B \ n" (is ?Bn) "fadd 0 0 = 0" "fmul 0 0 = 0" + shows c:"mult_matrix fmul fadd A B = mult_matrix_n n fmul fadd A B" +proof - + show ?thesis using assms + apply (simp add: mult_matrix_def mult_matrix_n_def) + apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) + apply (rule foldseq_zerotail, simp_all add: nrows_le ncols_le assms) + done +qed + +lemma mult_matrix_nm: + assumes "ncols A <= n" "nrows B <= n" "ncols A <= m" "nrows B <= m" "fadd 0 0 = 0" "fmul 0 0 = 0" + shows "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" +proof - + from assms have "mult_matrix_n n fmul fadd A B = mult_matrix fmul fadd A B" + by (simp add: mult_matrix_n) + also from assms have "\ = mult_matrix_n m fmul fadd A B" + by (simp add: mult_matrix_n[THEN sym]) + finally show "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" by simp +qed + +definition r_distributive :: "('a \ 'b \ 'b) \ ('b \ 'b \ 'b) \ bool" where + "r_distributive fmul fadd == ! a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)" + +definition l_distributive :: "('a \ 'b \ 'a) \ ('a \ 'a \ 'a) \ bool" where + "l_distributive fmul fadd == ! a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)" + +definition distributive :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ bool" where + "distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd" + +lemma max1: "!! a x y. (a::nat) <= x \ a <= max x y" by (arith) +lemma max2: "!! b x y. (b::nat) <= y \ b <= max x y" by (arith) + +lemma r_distributive_matrix: + assumes + "r_distributive fmul fadd" + "associative fadd" + "commutative fadd" + "fadd 0 0 = 0" + "! a. fmul a 0 = 0" + "! a. fmul 0 a = 0" + shows "r_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" +proof - + from assms show ?thesis + apply (simp add: r_distributive_def mult_matrix_def, auto) + proof - + fix a::"'a matrix" + fix u::"'b matrix" + fix v::"'b matrix" + let ?mx = "max (ncols a) (max (nrows u) (nrows v))" + from assms show "mult_matrix_n (max (ncols a) (nrows (combine_matrix fadd u v))) fmul fadd a (combine_matrix fadd u v) = + combine_matrix fadd (mult_matrix_n (max (ncols a) (nrows u)) fmul fadd a u) (mult_matrix_n (max (ncols a) (nrows v)) fmul fadd a v)" + apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) + apply (simp add: max1 max2 combine_nrows combine_ncols)+ + apply (subst mult_matrix_nm[of _ _ v ?mx fadd fmul]) + apply (simp add: max1 max2 combine_nrows combine_ncols)+ + apply (subst mult_matrix_nm[of _ _ u ?mx fadd fmul]) + apply (simp add: max1 max2 combine_nrows combine_ncols)+ + apply (simp add: mult_matrix_n_def r_distributive_def foldseq_distr[of fadd]) + apply (simp add: combine_matrix_def combine_infmatrix_def) + apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) + apply (simplesubst RepAbs_matrix) + apply (simp, auto) + apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero) + apply (rule exI[of _ "ncols v"], simp add: ncols_le foldseq_zero) + apply (subst RepAbs_matrix) + apply (simp, auto) + apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero) + apply (rule exI[of _ "ncols u"], simp add: ncols_le foldseq_zero) + done + qed +qed + +lemma l_distributive_matrix: + assumes + "l_distributive fmul fadd" + "associative fadd" + "commutative fadd" + "fadd 0 0 = 0" + "! a. fmul a 0 = 0" + "! a. fmul 0 a = 0" + shows "l_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" +proof - + from assms show ?thesis + apply (simp add: l_distributive_def mult_matrix_def, auto) + proof - + fix a::"'b matrix" + fix u::"'a matrix" + fix v::"'a matrix" + let ?mx = "max (nrows a) (max (ncols u) (ncols v))" + from assms show "mult_matrix_n (max (ncols (combine_matrix fadd u v)) (nrows a)) fmul fadd (combine_matrix fadd u v) a = + combine_matrix fadd (mult_matrix_n (max (ncols u) (nrows a)) fmul fadd u a) (mult_matrix_n (max (ncols v) (nrows a)) fmul fadd v a)" + apply (subst mult_matrix_nm[of v _ _ ?mx fadd fmul]) + apply (simp add: max1 max2 combine_nrows combine_ncols)+ + apply (subst mult_matrix_nm[of u _ _ ?mx fadd fmul]) + apply (simp add: max1 max2 combine_nrows combine_ncols)+ + apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) + apply (simp add: max1 max2 combine_nrows combine_ncols)+ + apply (simp add: mult_matrix_n_def l_distributive_def foldseq_distr[of fadd]) + apply (simp add: combine_matrix_def combine_infmatrix_def) + apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) + apply (simplesubst RepAbs_matrix) + apply (simp, auto) + apply (rule exI[of _ "nrows v"], simp add: nrows_le foldseq_zero) + apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero) + apply (subst RepAbs_matrix) + apply (simp, auto) + apply (rule exI[of _ "nrows u"], simp add: nrows_le foldseq_zero) + apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero) + done + qed +qed + +instantiation matrix :: (zero) zero +begin + +definition zero_matrix_def: "0 = Abs_matrix (\j i. 0)" + +instance .. + +end + +lemma Rep_zero_matrix_def[simp]: "Rep_matrix 0 j i = 0" + apply (simp add: zero_matrix_def) + apply (subst RepAbs_matrix) + by (auto) + +lemma zero_matrix_def_nrows[simp]: "nrows 0 = 0" +proof - + have a:"!! (x::nat). x <= 0 \ x = 0" by (arith) + show "nrows 0 = 0" by (rule a, subst nrows_le, simp) +qed + +lemma zero_matrix_def_ncols[simp]: "ncols 0 = 0" +proof - + have a:"!! (x::nat). x <= 0 \ x = 0" by (arith) + show "ncols 0 = 0" by (rule a, subst ncols_le, simp) +qed + +lemma combine_matrix_zero_l_neutral: "zero_l_neutral f \ zero_l_neutral (combine_matrix f)" + by (simp add: zero_l_neutral_def combine_matrix_def combine_infmatrix_def) + +lemma combine_matrix_zero_r_neutral: "zero_r_neutral f \ zero_r_neutral (combine_matrix f)" + by (simp add: zero_r_neutral_def combine_matrix_def combine_infmatrix_def) + +lemma mult_matrix_zero_closed: "\fadd 0 0 = 0; zero_closed fmul\ \ zero_closed (mult_matrix fmul fadd)" + apply (simp add: zero_closed_def mult_matrix_def mult_matrix_n_def) + apply (auto) + by (subst foldseq_zero, (simp add: zero_matrix_def)+)+ + +lemma mult_matrix_n_zero_right[simp]: "\fadd 0 0 = 0; !a. fmul a 0 = 0\ \ mult_matrix_n n fmul fadd A 0 = 0" + apply (simp add: mult_matrix_n_def) + apply (subst foldseq_zero) + by (simp_all add: zero_matrix_def) + +lemma mult_matrix_n_zero_left[simp]: "\fadd 0 0 = 0; !a. fmul 0 a = 0\ \ mult_matrix_n n fmul fadd 0 A = 0" + apply (simp add: mult_matrix_n_def) + apply (subst foldseq_zero) + by (simp_all add: zero_matrix_def) + +lemma mult_matrix_zero_left[simp]: "\fadd 0 0 = 0; !a. fmul 0 a = 0\ \ mult_matrix fmul fadd 0 A = 0" +by (simp add: mult_matrix_def) + +lemma mult_matrix_zero_right[simp]: "\fadd 0 0 = 0; !a. fmul a 0 = 0\ \ mult_matrix fmul fadd A 0 = 0" +by (simp add: mult_matrix_def) + +lemma apply_matrix_zero[simp]: "f 0 = 0 \ apply_matrix f 0 = 0" + apply (simp add: apply_matrix_def apply_infmatrix_def) + by (simp add: zero_matrix_def) + +lemma combine_matrix_zero: "f 0 0 = 0 \ combine_matrix f 0 0 = 0" + apply (simp add: combine_matrix_def combine_infmatrix_def) + by (simp add: zero_matrix_def) + +lemma transpose_matrix_zero[simp]: "transpose_matrix 0 = 0" +apply (simp add: transpose_matrix_def zero_matrix_def RepAbs_matrix) +apply (subst Rep_matrix_inject[symmetric], (rule ext)+) +apply (simp add: RepAbs_matrix) +done + +lemma apply_zero_matrix_def[simp]: "apply_matrix (% x. 0) A = 0" + apply (simp add: apply_matrix_def apply_infmatrix_def) + by (simp add: zero_matrix_def) + +definition singleton_matrix :: "nat \ nat \ ('a::zero) \ 'a matrix" where + "singleton_matrix j i a == Abs_matrix(% m n. if j = m & i = n then a else 0)" + +definition move_matrix :: "('a::zero) matrix \ int \ int \ 'a matrix" where + "move_matrix A y x == Abs_matrix(% j i. if (((int j)-y) < 0) | (((int i)-x) < 0) then 0 else Rep_matrix A (nat ((int j)-y)) (nat ((int i)-x)))" + +definition take_rows :: "('a::zero) matrix \ nat \ 'a matrix" where + "take_rows A r == Abs_matrix(% j i. if (j < r) then (Rep_matrix A j i) else 0)" + +definition take_columns :: "('a::zero) matrix \ nat \ 'a matrix" where + "take_columns A c == Abs_matrix(% j i. if (i < c) then (Rep_matrix A j i) else 0)" + +definition column_of_matrix :: "('a::zero) matrix \ nat \ 'a matrix" where + "column_of_matrix A n == take_columns (move_matrix A 0 (- int n)) 1" + +definition row_of_matrix :: "('a::zero) matrix \ nat \ 'a matrix" where + "row_of_matrix A m == take_rows (move_matrix A (- int m) 0) 1" + +lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m & i = n then e else 0)" +apply (simp add: singleton_matrix_def) +apply (auto) +apply (subst RepAbs_matrix) +apply (rule exI[of _ "Suc m"], simp) +apply (rule exI[of _ "Suc n"], simp+) +by (subst RepAbs_matrix, rule exI[of _ "Suc j"], simp, rule exI[of _ "Suc i"], simp+)+ + +lemma apply_singleton_matrix[simp]: "f 0 = 0 \ apply_matrix f (singleton_matrix j i x) = (singleton_matrix j i (f x))" +apply (subst Rep_matrix_inject[symmetric]) +apply (rule ext)+ +apply (simp) +done + +lemma singleton_matrix_zero[simp]: "singleton_matrix j i 0 = 0" + by (simp add: singleton_matrix_def zero_matrix_def) + +lemma nrows_singleton[simp]: "nrows(singleton_matrix j i e) = (if e = 0 then 0 else Suc j)" +proof- +have th: "\ (\m. m \ j)" "\n. \ n \ i" by arith+ +from th show ?thesis +apply (auto) +apply (rule le_antisym) +apply (subst nrows_le) +apply (simp add: singleton_matrix_def, auto) +apply (subst RepAbs_matrix) +apply auto +apply (simp add: Suc_le_eq) +apply (rule not_leE) +apply (subst nrows_le) +by simp +qed + +lemma ncols_singleton[simp]: "ncols(singleton_matrix j i e) = (if e = 0 then 0 else Suc i)" +proof- +have th: "\ (\m. m \ j)" "\n. \ n \ i" by arith+ +from th show ?thesis +apply (auto) +apply (rule le_antisym) +apply (subst ncols_le) +apply (simp add: singleton_matrix_def, auto) +apply (subst RepAbs_matrix) +apply auto +apply (simp add: Suc_le_eq) +apply (rule not_leE) +apply (subst ncols_le) +by simp +qed + +lemma combine_singleton: "f 0 0 = 0 \ combine_matrix f (singleton_matrix j i a) (singleton_matrix j i b) = singleton_matrix j i (f a b)" +apply (simp add: singleton_matrix_def combine_matrix_def combine_infmatrix_def) +apply (subst RepAbs_matrix) +apply (rule exI[of _ "Suc j"], simp) +apply (rule exI[of _ "Suc i"], simp) +apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) +apply (subst RepAbs_matrix) +apply (rule exI[of _ "Suc j"], simp) +apply (rule exI[of _ "Suc i"], simp) +by simp + +lemma transpose_singleton[simp]: "transpose_matrix (singleton_matrix j i a) = singleton_matrix i j a" +apply (subst Rep_matrix_inject[symmetric], (rule ext)+) +apply (simp) +done + +lemma Rep_move_matrix[simp]: + "Rep_matrix (move_matrix A y x) j i = + (if (((int j)-y) < 0) | (((int i)-x) < 0) then 0 else Rep_matrix A (nat((int j)-y)) (nat((int i)-x)))" +apply (simp add: move_matrix_def) +apply (auto) +by (subst RepAbs_matrix, + rule exI[of _ "(nrows A)+(nat (abs y))"], auto, rule nrows, arith, + rule exI[of _ "(ncols A)+(nat (abs x))"], auto, rule ncols, arith)+ + +lemma move_matrix_0_0[simp]: "move_matrix A 0 0 = A" +by (simp add: move_matrix_def) + +lemma move_matrix_ortho: "move_matrix A j i = move_matrix (move_matrix A j 0) 0 i" +apply (subst Rep_matrix_inject[symmetric]) +apply (rule ext)+ +apply (simp) +done + +lemma transpose_move_matrix[simp]: + "transpose_matrix (move_matrix A x y) = move_matrix (transpose_matrix A) y x" +apply (subst Rep_matrix_inject[symmetric], (rule ext)+) +apply (simp) +done + +lemma move_matrix_singleton[simp]: "move_matrix (singleton_matrix u v x) j i = + (if (j + int u < 0) | (i + int v < 0) then 0 else (singleton_matrix (nat (j + int u)) (nat (i + int v)) x))" + apply (subst Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply (case_tac "j + int u < 0") + apply (simp, arith) + apply (case_tac "i + int v < 0") + apply (simp, arith) + apply simp + apply arith + done + +lemma Rep_take_columns[simp]: + "Rep_matrix (take_columns A c) j i = + (if i < c then (Rep_matrix A j i) else 0)" +apply (simp add: take_columns_def) +apply (simplesubst RepAbs_matrix) +apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le) +apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le) +done + +lemma Rep_take_rows[simp]: + "Rep_matrix (take_rows A r) j i = + (if j < r then (Rep_matrix A j i) else 0)" +apply (simp add: take_rows_def) +apply (simplesubst RepAbs_matrix) +apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le) +apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le) +done + +lemma Rep_column_of_matrix[simp]: + "Rep_matrix (column_of_matrix A c) j i = (if i = 0 then (Rep_matrix A j c) else 0)" + by (simp add: column_of_matrix_def) + +lemma Rep_row_of_matrix[simp]: + "Rep_matrix (row_of_matrix A r) j i = (if j = 0 then (Rep_matrix A r i) else 0)" + by (simp add: row_of_matrix_def) + +lemma column_of_matrix: "ncols A <= n \ column_of_matrix A n = 0" +apply (subst Rep_matrix_inject[THEN sym]) +apply (rule ext)+ +by (simp add: ncols) + +lemma row_of_matrix: "nrows A <= n \ row_of_matrix A n = 0" +apply (subst Rep_matrix_inject[THEN sym]) +apply (rule ext)+ +by (simp add: nrows) + +lemma mult_matrix_singleton_right[simp]: + assumes + "! x. fmul x 0 = 0" + "! x. fmul 0 x = 0" + "! x. fadd 0 x = x" + "! x. fadd x 0 = x" + shows "(mult_matrix fmul fadd A (singleton_matrix j i e)) = apply_matrix (% x. fmul x e) (move_matrix (column_of_matrix A j) 0 (int i))" + apply (simp add: mult_matrix_def) + apply (subst mult_matrix_nm[of _ _ _ "max (ncols A) (Suc j)"]) + apply (auto) + apply (simp add: assms)+ + apply (simp add: mult_matrix_n_def apply_matrix_def apply_infmatrix_def) + apply (rule comb[of "Abs_matrix" "Abs_matrix"], auto, (rule ext)+) + apply (subst foldseq_almostzero[of _ j]) + apply (simp add: assms)+ + apply (auto) + done + +lemma mult_matrix_ext: + assumes + eprem: + "? e. (! a b. a \ b \ fmul a e \ fmul b e)" + and fprems: + "! a. fmul 0 a = 0" + "! a. fmul a 0 = 0" + "! a. fadd a 0 = a" + "! a. fadd 0 a = a" + and contraprems: + "mult_matrix fmul fadd A = mult_matrix fmul fadd B" + shows + "A = B" +proof(rule contrapos_np[of "False"], simp) + assume a: "A \ B" + have b: "!! f g. (! x y. f x y = g x y) \ f = g" by ((rule ext)+, auto) + have "? j i. (Rep_matrix A j i) \ (Rep_matrix B j i)" + apply (rule contrapos_np[of "False"], simp+) + apply (insert b[of "Rep_matrix A" "Rep_matrix B"], simp) + by (simp add: Rep_matrix_inject a) + then obtain J I where c:"(Rep_matrix A J I) \ (Rep_matrix B J I)" by blast + from eprem obtain e where eprops:"(! a b. a \ b \ fmul a e \ fmul b e)" by blast + let ?S = "singleton_matrix I 0 e" + let ?comp = "mult_matrix fmul fadd" + have d: "!!x f g. f = g \ f x = g x" by blast + have e: "(% x. fmul x e) 0 = 0" by (simp add: assms) + have "~(?comp A ?S = ?comp B ?S)" + apply (rule notI) + apply (simp add: fprems eprops) + apply (simp add: Rep_matrix_inject[THEN sym]) + apply (drule d[of _ _ "J"], drule d[of _ _ "0"]) + by (simp add: e c eprops) + with contraprems show "False" by simp +qed + +definition foldmatrix :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ ('a infmatrix) \ nat \ nat \ 'a" where + "foldmatrix f g A m n == foldseq_transposed g (% j. foldseq f (A j) n) m" + +definition foldmatrix_transposed :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ ('a infmatrix) \ nat \ nat \ 'a" where + "foldmatrix_transposed f g A m n == foldseq g (% j. foldseq_transposed f (A j) n) m" + +lemma foldmatrix_transpose: + assumes + "! a b c d. g(f a b) (f c d) = f (g a c) (g b d)" + shows + "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" +proof - + have forall:"!! P x. (! x. P x) \ P x" by auto + have tworows:"! A. foldmatrix f g A 1 n = foldmatrix_transposed g f (transpose_infmatrix A) n 1" + apply (induct n) + apply (simp add: foldmatrix_def foldmatrix_transposed_def assms)+ + apply (auto) + by (drule_tac x="(% j i. A j (Suc i))" in forall, simp) + show "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" + apply (simp add: foldmatrix_def foldmatrix_transposed_def) + apply (induct m, simp) + apply (simp) + apply (insert tworows) + apply (drule_tac x="% j i. (if j = 0 then (foldseq_transposed g (\u. A u i) m) else (A (Suc m) i))" in spec) + by (simp add: foldmatrix_def foldmatrix_transposed_def) +qed + +lemma foldseq_foldseq: +assumes + "associative f" + "associative g" + "! a b c d. g(f a b) (f c d) = f (g a c) (g b d)" +shows + "foldseq g (% j. foldseq f (A j) n) m = foldseq f (% j. foldseq g ((transpose_infmatrix A) j) m) n" + apply (insert foldmatrix_transpose[of g f A m n]) + by (simp add: foldmatrix_def foldmatrix_transposed_def foldseq_assoc[THEN sym] assms) + +lemma mult_n_nrows: +assumes +"! a. fmul 0 a = 0" +"! a. fmul a 0 = 0" +"fadd 0 0 = 0" +shows "nrows (mult_matrix_n n fmul fadd A B) \ nrows A" +apply (subst nrows_le) +apply (simp add: mult_matrix_n_def) +apply (subst RepAbs_matrix) +apply (rule_tac x="nrows A" in exI) +apply (simp add: nrows assms foldseq_zero) +apply (rule_tac x="ncols B" in exI) +apply (simp add: ncols assms foldseq_zero) +apply (simp add: nrows assms foldseq_zero) +done + +lemma mult_n_ncols: +assumes +"! a. fmul 0 a = 0" +"! a. fmul a 0 = 0" +"fadd 0 0 = 0" +shows "ncols (mult_matrix_n n fmul fadd A B) \ ncols B" +apply (subst ncols_le) +apply (simp add: mult_matrix_n_def) +apply (subst RepAbs_matrix) +apply (rule_tac x="nrows A" in exI) +apply (simp add: nrows assms foldseq_zero) +apply (rule_tac x="ncols B" in exI) +apply (simp add: ncols assms foldseq_zero) +apply (simp add: ncols assms foldseq_zero) +done + +lemma mult_nrows: +assumes +"! a. fmul 0 a = 0" +"! a. fmul a 0 = 0" +"fadd 0 0 = 0" +shows "nrows (mult_matrix fmul fadd A B) \ nrows A" +by (simp add: mult_matrix_def mult_n_nrows assms) + +lemma mult_ncols: +assumes +"! a. fmul 0 a = 0" +"! a. fmul a 0 = 0" +"fadd 0 0 = 0" +shows "ncols (mult_matrix fmul fadd A B) \ ncols B" +by (simp add: mult_matrix_def mult_n_ncols assms) + +lemma nrows_move_matrix_le: "nrows (move_matrix A j i) <= nat((int (nrows A)) + j)" + apply (auto simp add: nrows_le) + apply (rule nrows) + apply (arith) + done + +lemma ncols_move_matrix_le: "ncols (move_matrix A j i) <= nat((int (ncols A)) + i)" + apply (auto simp add: ncols_le) + apply (rule ncols) + apply (arith) + done + +lemma mult_matrix_assoc: + assumes + "! a. fmul1 0 a = 0" + "! a. fmul1 a 0 = 0" + "! a. fmul2 0 a = 0" + "! a. fmul2 a 0 = 0" + "fadd1 0 0 = 0" + "fadd2 0 0 = 0" + "! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" + "associative fadd1" + "associative fadd2" + "! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" + "! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" + "! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" + shows "mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B) C = mult_matrix fmul1 fadd1 A (mult_matrix fmul2 fadd2 B C)" +proof - + have comb_left: "!! A B x y. A = B \ (Rep_matrix (Abs_matrix A)) x y = (Rep_matrix(Abs_matrix B)) x y" by blast + have fmul2fadd1fold: "!! x s n. fmul2 (foldseq fadd1 s n) x = foldseq fadd1 (% k. fmul2 (s k) x) n" + by (rule_tac g1 = "% y. fmul2 y x" in ssubst [OF foldseq_distr_unary], insert assms, simp_all) + have fmul1fadd2fold: "!! x s n. fmul1 x (foldseq fadd2 s n) = foldseq fadd2 (% k. fmul1 x (s k)) n" + using assms by (rule_tac g1 = "% y. fmul1 x y" in ssubst [OF foldseq_distr_unary], simp_all) + let ?N = "max (ncols A) (max (ncols B) (max (nrows B) (nrows C)))" + show ?thesis + apply (simp add: Rep_matrix_inject[THEN sym]) + apply (rule ext)+ + apply (simp add: mult_matrix_def) + apply (simplesubst mult_matrix_nm[of _ "max (ncols (mult_matrix_n (max (ncols A) (nrows B)) fmul1 fadd1 A B)) (nrows C)" _ "max (ncols B) (nrows C)"]) + apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ + apply (simplesubst mult_matrix_nm[of _ "max (ncols A) (nrows (mult_matrix_n (max (ncols B) (nrows C)) fmul2 fadd2 B C))" _ "max (ncols A) (nrows B)"]) + apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ + apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) + apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ + apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) + apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ + apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) + apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ + apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) + apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ + apply (simp add: mult_matrix_n_def) + apply (rule comb_left) + apply ((rule ext)+, simp) + apply (simplesubst RepAbs_matrix) + apply (rule exI[of _ "nrows B"]) + apply (simp add: nrows assms foldseq_zero) + apply (rule exI[of _ "ncols C"]) + apply (simp add: assms ncols foldseq_zero) + apply (subst RepAbs_matrix) + apply (rule exI[of _ "nrows A"]) + apply (simp add: nrows assms foldseq_zero) + apply (rule exI[of _ "ncols B"]) + apply (simp add: assms ncols foldseq_zero) + apply (simp add: fmul2fadd1fold fmul1fadd2fold assms) + apply (subst foldseq_foldseq) + apply (simp add: assms)+ + apply (simp add: transpose_infmatrix) + done +qed + +lemma + assumes + "! a. fmul1 0 a = 0" + "! a. fmul1 a 0 = 0" + "! a. fmul2 0 a = 0" + "! a. fmul2 a 0 = 0" + "fadd1 0 0 = 0" + "fadd2 0 0 = 0" + "! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" + "associative fadd1" + "associative fadd2" + "! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" + "! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" + "! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" + shows + "(mult_matrix fmul1 fadd1 A) o (mult_matrix fmul2 fadd2 B) = mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B)" +apply (rule ext)+ +apply (simp add: comp_def ) +apply (simp add: mult_matrix_assoc assms) +done + +lemma mult_matrix_assoc_simple: + assumes + "! a. fmul 0 a = 0" + "! a. fmul a 0 = 0" + "fadd 0 0 = 0" + "associative fadd" + "commutative fadd" + "associative fmul" + "distributive fmul fadd" + shows "mult_matrix fmul fadd (mult_matrix fmul fadd A B) C = mult_matrix fmul fadd A (mult_matrix fmul fadd B C)" +proof - + have "!! a b c d. fadd (fadd a b) (fadd c d) = fadd (fadd a c) (fadd b d)" + using assms by (simp add: associative_def commutative_def) + then show ?thesis + apply (subst mult_matrix_assoc) + using assms + apply simp_all + apply (simp_all add: associative_def distributive_def l_distributive_def r_distributive_def) + done +qed + +lemma transpose_apply_matrix: "f 0 = 0 \ transpose_matrix (apply_matrix f A) = apply_matrix f (transpose_matrix A)" +apply (simp add: Rep_matrix_inject[THEN sym]) +apply (rule ext)+ +by simp + +lemma transpose_combine_matrix: "f 0 0 = 0 \ transpose_matrix (combine_matrix f A B) = combine_matrix f (transpose_matrix A) (transpose_matrix B)" +apply (simp add: Rep_matrix_inject[THEN sym]) +apply (rule ext)+ +by simp + +lemma Rep_mult_matrix: + assumes + "! a. fmul 0 a = 0" + "! a. fmul a 0 = 0" + "fadd 0 0 = 0" + shows + "Rep_matrix(mult_matrix fmul fadd A B) j i = + foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) (max (ncols A) (nrows B))" +apply (simp add: mult_matrix_def mult_matrix_n_def) +apply (subst RepAbs_matrix) +apply (rule exI[of _ "nrows A"], insert assms, simp add: nrows foldseq_zero) +apply (rule exI[of _ "ncols B"], insert assms, simp add: ncols foldseq_zero) +apply simp +done + +lemma transpose_mult_matrix: + assumes + "! a. fmul 0 a = 0" + "! a. fmul a 0 = 0" + "fadd 0 0 = 0" + "! x y. fmul y x = fmul x y" + shows + "transpose_matrix (mult_matrix fmul fadd A B) = mult_matrix fmul fadd (transpose_matrix B) (transpose_matrix A)" + apply (simp add: Rep_matrix_inject[THEN sym]) + apply (rule ext)+ + using assms + apply (simp add: Rep_mult_matrix max_ac) + done + +lemma column_transpose_matrix: "column_of_matrix (transpose_matrix A) n = transpose_matrix (row_of_matrix A n)" +apply (simp add: Rep_matrix_inject[THEN sym]) +apply (rule ext)+ +by simp + +lemma take_columns_transpose_matrix: "take_columns (transpose_matrix A) n = transpose_matrix (take_rows A n)" +apply (simp add: Rep_matrix_inject[THEN sym]) +apply (rule ext)+ +by simp + +instantiation matrix :: ("{zero, ord}") ord +begin + +definition + le_matrix_def: "A \ B \ (\j i. Rep_matrix A j i \ Rep_matrix B j i)" + +definition + less_def: "A < (B\'a matrix) \ A \ B \ \ B \ A" + +instance .. + +end + +instance matrix :: ("{zero, order}") order +apply intro_classes +apply (simp_all add: le_matrix_def less_def) +apply (auto) +apply (drule_tac x=j in spec, drule_tac x=j in spec) +apply (drule_tac x=i in spec, drule_tac x=i in spec) +apply (simp) +apply (simp add: Rep_matrix_inject[THEN sym]) +apply (rule ext)+ +apply (drule_tac x=xa in spec, drule_tac x=xa in spec) +apply (drule_tac x=xb in spec, drule_tac x=xb in spec) +apply simp +done + +lemma le_apply_matrix: + assumes + "f 0 = 0" + "! x y. x <= y \ f x <= f y" + "(a::('a::{ord, zero}) matrix) <= b" + shows + "apply_matrix f a <= apply_matrix f b" + using assms by (simp add: le_matrix_def) + +lemma le_combine_matrix: + assumes + "f 0 0 = 0" + "! a b c d. a <= b & c <= d \ f a c <= f b d" + "A <= B" + "C <= D" + shows + "combine_matrix f A C <= combine_matrix f B D" + using assms by (simp add: le_matrix_def) + +lemma le_left_combine_matrix: + assumes + "f 0 0 = 0" + "! a b c. a <= b \ f c a <= f c b" + "A <= B" + shows + "combine_matrix f C A <= combine_matrix f C B" + using assms by (simp add: le_matrix_def) + +lemma le_right_combine_matrix: + assumes + "f 0 0 = 0" + "! a b c. a <= b \ f a c <= f b c" + "A <= B" + shows + "combine_matrix f A C <= combine_matrix f B C" + using assms by (simp add: le_matrix_def) + +lemma le_transpose_matrix: "(A <= B) = (transpose_matrix A <= transpose_matrix B)" + by (simp add: le_matrix_def, auto) + +lemma le_foldseq: + assumes + "! a b c d . a <= b & c <= d \ f a c <= f b d" + "! i. i <= n \ s i <= t i" + shows + "foldseq f s n <= foldseq f t n" +proof - + have "! s t. (! i. i<=n \ s i <= t i) \ foldseq f s n <= foldseq f t n" + by (induct n) (simp_all add: assms) + then show "foldseq f s n <= foldseq f t n" using assms by simp +qed + +lemma le_left_mult: + assumes + "! a b c d. a <= b & c <= d \ fadd a c <= fadd b d" + "! c a b. 0 <= c & a <= b \ fmul c a <= fmul c b" + "! a. fmul 0 a = 0" + "! a. fmul a 0 = 0" + "fadd 0 0 = 0" + "0 <= C" + "A <= B" + shows + "mult_matrix fmul fadd C A <= mult_matrix fmul fadd C B" + using assms + apply (simp add: le_matrix_def Rep_mult_matrix) + apply (auto) + apply (simplesubst foldseq_zerotail[of _ _ _ "max (ncols C) (max (nrows A) (nrows B))"], simp_all add: nrows ncols max1 max2)+ + apply (rule le_foldseq) + apply (auto) + done + +lemma le_right_mult: + assumes + "! a b c d. a <= b & c <= d \ fadd a c <= fadd b d" + "! c a b. 0 <= c & a <= b \ fmul a c <= fmul b c" + "! a. fmul 0 a = 0" + "! a. fmul a 0 = 0" + "fadd 0 0 = 0" + "0 <= C" + "A <= B" + shows + "mult_matrix fmul fadd A C <= mult_matrix fmul fadd B C" + using assms + apply (simp add: le_matrix_def Rep_mult_matrix) + apply (auto) + apply (simplesubst foldseq_zerotail[of _ _ _ "max (nrows C) (max (ncols A) (ncols B))"], simp_all add: nrows ncols max1 max2)+ + apply (rule le_foldseq) + apply (auto) + done + +lemma spec2: "! j i. P j i \ P j i" by blast +lemma neg_imp: "(\ Q \ \ P) \ P \ Q" by blast + +lemma singleton_matrix_le[simp]: "(singleton_matrix j i a <= singleton_matrix j i b) = (a <= (b::_::order))" + by (auto simp add: le_matrix_def) + +lemma singleton_le_zero[simp]: "(singleton_matrix j i x <= 0) = (x <= (0::'a::{order,zero}))" + apply (auto) + apply (simp add: le_matrix_def) + apply (drule_tac j=j and i=i in spec2) + apply (simp) + apply (simp add: le_matrix_def) + done + +lemma singleton_ge_zero[simp]: "(0 <= singleton_matrix j i x) = ((0::'a::{order,zero}) <= x)" + apply (auto) + apply (simp add: le_matrix_def) + apply (drule_tac j=j and i=i in spec2) + apply (simp) + apply (simp add: le_matrix_def) + done + +lemma move_matrix_le_zero[simp]: "0 <= j \ 0 <= i \ (move_matrix A j i <= 0) = (A <= (0::('a::{order,zero}) matrix))" + apply (auto simp add: le_matrix_def) + apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2) + apply (auto) + done + +lemma move_matrix_zero_le[simp]: "0 <= j \ 0 <= i \ (0 <= move_matrix A j i) = ((0::('a::{order,zero}) matrix) <= A)" + apply (auto simp add: le_matrix_def) + apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2) + apply (auto) + done + +lemma move_matrix_le_move_matrix_iff[simp]: "0 <= j \ 0 <= i \ (move_matrix A j i <= move_matrix B j i) = (A <= (B::('a::{order,zero}) matrix))" + apply (auto simp add: le_matrix_def) + apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2) + apply (auto) + done + +instantiation matrix :: ("{lattice, zero}") lattice +begin + +definition "inf = combine_matrix inf" + +definition "sup = combine_matrix sup" + +instance + by default (auto simp add: le_infI le_matrix_def inf_matrix_def sup_matrix_def) + +end + +instantiation matrix :: ("{plus, zero}") plus +begin + +definition + plus_matrix_def: "A + B = combine_matrix (op +) A B" + +instance .. + +end + +instantiation matrix :: ("{uminus, zero}") uminus +begin + +definition + minus_matrix_def: "- A = apply_matrix uminus A" + +instance .. + +end + +instantiation matrix :: ("{minus, zero}") minus +begin + +definition + diff_matrix_def: "A - B = combine_matrix (op -) A B" + +instance .. + +end + +instantiation matrix :: ("{plus, times, zero}") times +begin + +definition + times_matrix_def: "A * B = mult_matrix (op *) (op +) A B" + +instance .. + +end + +instantiation matrix :: ("{lattice, uminus, zero}") abs +begin + +definition + abs_matrix_def: "abs (A \ 'a matrix) = sup A (- A)" + +instance .. + +end + +instance matrix :: (monoid_add) monoid_add +proof + fix A B C :: "'a matrix" + show "A + B + C = A + (B + C)" + apply (simp add: plus_matrix_def) + apply (rule combine_matrix_assoc[simplified associative_def, THEN spec, THEN spec, THEN spec]) + apply (simp_all add: add_assoc) + done + show "0 + A = A" + apply (simp add: plus_matrix_def) + apply (rule combine_matrix_zero_l_neutral[simplified zero_l_neutral_def, THEN spec]) + apply (simp) + done + show "A + 0 = A" + apply (simp add: plus_matrix_def) + apply (rule combine_matrix_zero_r_neutral [simplified zero_r_neutral_def, THEN spec]) + apply (simp) + done +qed + +instance matrix :: (comm_monoid_add) comm_monoid_add +proof + fix A B :: "'a matrix" + show "A + B = B + A" + apply (simp add: plus_matrix_def) + apply (rule combine_matrix_commute[simplified commutative_def, THEN spec, THEN spec]) + apply (simp_all add: add_commute) + done + show "0 + A = A" + apply (simp add: plus_matrix_def) + apply (rule combine_matrix_zero_l_neutral[simplified zero_l_neutral_def, THEN spec]) + apply (simp) + done +qed + +instance matrix :: (group_add) group_add +proof + fix A B :: "'a matrix" + show "- A + A = 0" + by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext) + show "A - B = A + - B" + by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject [symmetric] ext diff_minus) +qed + +instance matrix :: (ab_group_add) ab_group_add +proof + fix A B :: "'a matrix" + show "- A + A = 0" + by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext) + show "A - B = A + - B" + by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext) +qed + +instance matrix :: (ordered_ab_group_add) ordered_ab_group_add +proof + fix A B C :: "'a matrix" + assume "A <= B" + then show "C + A <= C + B" + apply (simp add: plus_matrix_def) + apply (rule le_left_combine_matrix) + apply (simp_all) + done +qed + +instance matrix :: (lattice_ab_group_add) semilattice_inf_ab_group_add .. +instance matrix :: (lattice_ab_group_add) semilattice_sup_ab_group_add .. + +instance matrix :: (semiring_0) semiring_0 +proof + fix A B C :: "'a matrix" + show "A * B * C = A * (B * C)" + apply (simp add: times_matrix_def) + apply (rule mult_matrix_assoc) + apply (simp_all add: associative_def algebra_simps) + done + show "(A + B) * C = A * C + B * C" + apply (simp add: times_matrix_def plus_matrix_def) + apply (rule l_distributive_matrix[simplified l_distributive_def, THEN spec, THEN spec, THEN spec]) + apply (simp_all add: associative_def commutative_def algebra_simps) + done + show "A * (B + C) = A * B + A * C" + apply (simp add: times_matrix_def plus_matrix_def) + apply (rule r_distributive_matrix[simplified r_distributive_def, THEN spec, THEN spec, THEN spec]) + apply (simp_all add: associative_def commutative_def algebra_simps) + done + show "0 * A = 0" by (simp add: times_matrix_def) + show "A * 0 = 0" by (simp add: times_matrix_def) +qed + +instance matrix :: (ring) ring .. + +instance matrix :: (ordered_ring) ordered_ring +proof + fix A B C :: "'a matrix" + assume a: "A \ B" + assume b: "0 \ C" + from a b show "C * A \ C * B" + apply (simp add: times_matrix_def) + apply (rule le_left_mult) + apply (simp_all add: add_mono mult_left_mono) + done + from a b show "A * C \ B * C" + apply (simp add: times_matrix_def) + apply (rule le_right_mult) + apply (simp_all add: add_mono mult_right_mono) + done +qed + +instance matrix :: (lattice_ring) lattice_ring +proof + fix A B C :: "('a :: lattice_ring) matrix" + show "abs A = sup A (-A)" + by (simp add: abs_matrix_def) +qed + +lemma Rep_matrix_add[simp]: + "Rep_matrix ((a::('a::monoid_add)matrix)+b) j i = (Rep_matrix a j i) + (Rep_matrix b j i)" + by (simp add: plus_matrix_def) + +lemma Rep_matrix_mult: "Rep_matrix ((a::('a::semiring_0) matrix) * b) j i = + foldseq (op +) (% k. (Rep_matrix a j k) * (Rep_matrix b k i)) (max (ncols a) (nrows b))" +apply (simp add: times_matrix_def) +apply (simp add: Rep_mult_matrix) +done + +lemma apply_matrix_add: "! x y. f (x+y) = (f x) + (f y) \ f 0 = (0::'a) + \ apply_matrix f ((a::('a::monoid_add) matrix) + b) = (apply_matrix f a) + (apply_matrix f b)" +apply (subst Rep_matrix_inject[symmetric]) +apply (rule ext)+ +apply (simp) +done + +lemma singleton_matrix_add: "singleton_matrix j i ((a::_::monoid_add)+b) = (singleton_matrix j i a) + (singleton_matrix j i b)" +apply (subst Rep_matrix_inject[symmetric]) +apply (rule ext)+ +apply (simp) +done + +lemma nrows_mult: "nrows ((A::('a::semiring_0) matrix) * B) <= nrows A" +by (simp add: times_matrix_def mult_nrows) + +lemma ncols_mult: "ncols ((A::('a::semiring_0) matrix) * B) <= ncols B" +by (simp add: times_matrix_def mult_ncols) + +definition + one_matrix :: "nat \ ('a::{zero,one}) matrix" where + "one_matrix n = Abs_matrix (% j i. if j = i & j < n then 1 else 0)" + +lemma Rep_one_matrix[simp]: "Rep_matrix (one_matrix n) j i = (if (j = i & j < n) then 1 else 0)" +apply (simp add: one_matrix_def) +apply (simplesubst RepAbs_matrix) +apply (rule exI[of _ n], simp add: split_if)+ +by (simp add: split_if) + +lemma nrows_one_matrix[simp]: "nrows ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _") +proof - + have "?r <= n" by (simp add: nrows_le) + moreover have "n <= ?r" by (simp add:le_nrows, arith) + ultimately show "?r = n" by simp +qed + +lemma ncols_one_matrix[simp]: "ncols ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _") +proof - + have "?r <= n" by (simp add: ncols_le) + moreover have "n <= ?r" by (simp add: le_ncols, arith) + ultimately show "?r = n" by simp +qed + +lemma one_matrix_mult_right[simp]: "ncols A <= n \ (A::('a::{semiring_1}) matrix) * (one_matrix n) = A" +apply (subst Rep_matrix_inject[THEN sym]) +apply (rule ext)+ +apply (simp add: times_matrix_def Rep_mult_matrix) +apply (rule_tac j1="xa" in ssubst[OF foldseq_almostzero]) +apply (simp_all) +by (simp add: ncols) + +lemma one_matrix_mult_left[simp]: "nrows A <= n \ (one_matrix n) * A = (A::('a::semiring_1) matrix)" +apply (subst Rep_matrix_inject[THEN sym]) +apply (rule ext)+ +apply (simp add: times_matrix_def Rep_mult_matrix) +apply (rule_tac j1="x" in ssubst[OF foldseq_almostzero]) +apply (simp_all) +by (simp add: nrows) + +lemma transpose_matrix_mult: "transpose_matrix ((A::('a::comm_ring) matrix)*B) = (transpose_matrix B) * (transpose_matrix A)" +apply (simp add: times_matrix_def) +apply (subst transpose_mult_matrix) +apply (simp_all add: mult_commute) +done + +lemma transpose_matrix_add: "transpose_matrix ((A::('a::monoid_add) matrix)+B) = transpose_matrix A + transpose_matrix B" +by (simp add: plus_matrix_def transpose_combine_matrix) + +lemma transpose_matrix_diff: "transpose_matrix ((A::('a::group_add) matrix)-B) = transpose_matrix A - transpose_matrix B" +by (simp add: diff_matrix_def transpose_combine_matrix) + +lemma transpose_matrix_minus: "transpose_matrix (-(A::('a::group_add) matrix)) = - transpose_matrix (A::'a matrix)" +by (simp add: minus_matrix_def transpose_apply_matrix) + +definition right_inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" where + "right_inverse_matrix A X == (A * X = one_matrix (max (nrows A) (ncols X))) \ nrows X \ ncols A" + +definition left_inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" where + "left_inverse_matrix A X == (X * A = one_matrix (max(nrows X) (ncols A))) \ ncols X \ nrows A" + +definition inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" where + "inverse_matrix A X == (right_inverse_matrix A X) \ (left_inverse_matrix A X)" + +lemma right_inverse_matrix_dim: "right_inverse_matrix A X \ nrows A = ncols X" +apply (insert ncols_mult[of A X], insert nrows_mult[of A X]) +by (simp add: right_inverse_matrix_def) + +lemma left_inverse_matrix_dim: "left_inverse_matrix A Y \ ncols A = nrows Y" +apply (insert ncols_mult[of Y A], insert nrows_mult[of Y A]) +by (simp add: left_inverse_matrix_def) + +lemma left_right_inverse_matrix_unique: + assumes "left_inverse_matrix A Y" "right_inverse_matrix A X" + shows "X = Y" +proof - + have "Y = Y * one_matrix (nrows A)" + apply (subst one_matrix_mult_right) + using assms + apply (simp_all add: left_inverse_matrix_def) + done + also have "\ = Y * (A * X)" + apply (insert assms) + apply (frule right_inverse_matrix_dim) + by (simp add: right_inverse_matrix_def) + also have "\ = (Y * A) * X" by (simp add: mult_assoc) + also have "\ = X" + apply (insert assms) + apply (frule left_inverse_matrix_dim) + apply (simp_all add: left_inverse_matrix_def right_inverse_matrix_def one_matrix_mult_left) + done + ultimately show "X = Y" by (simp) +qed + +lemma inverse_matrix_inject: "\ inverse_matrix A X; inverse_matrix A Y \ \ X = Y" + by (auto simp add: inverse_matrix_def left_right_inverse_matrix_unique) + +lemma one_matrix_inverse: "inverse_matrix (one_matrix n) (one_matrix n)" + by (simp add: inverse_matrix_def left_inverse_matrix_def right_inverse_matrix_def) + +lemma zero_imp_mult_zero: "(a::'a::semiring_0) = 0 | b = 0 \ a * b = 0" +by auto + +lemma Rep_matrix_zero_imp_mult_zero: + "! j i k. (Rep_matrix A j k = 0) | (Rep_matrix B k i) = 0 \ A * B = (0::('a::lattice_ring) matrix)" +apply (subst Rep_matrix_inject[symmetric]) +apply (rule ext)+ +apply (auto simp add: Rep_matrix_mult foldseq_zero zero_imp_mult_zero) +done + +lemma add_nrows: "nrows (A::('a::monoid_add) matrix) <= u \ nrows B <= u \ nrows (A + B) <= u" +apply (simp add: plus_matrix_def) +apply (rule combine_nrows) +apply (simp_all) +done + +lemma move_matrix_row_mult: "move_matrix ((A::('a::semiring_0) matrix) * B) j 0 = (move_matrix A j 0) * B" +apply (subst Rep_matrix_inject[symmetric]) +apply (rule ext)+ +apply (auto simp add: Rep_matrix_mult foldseq_zero) +apply (rule_tac foldseq_zerotail[symmetric]) +apply (auto simp add: nrows zero_imp_mult_zero max2) +apply (rule order_trans) +apply (rule ncols_move_matrix_le) +apply (simp add: max1) +done + +lemma move_matrix_col_mult: "move_matrix ((A::('a::semiring_0) matrix) * B) 0 i = A * (move_matrix B 0 i)" +apply (subst Rep_matrix_inject[symmetric]) +apply (rule ext)+ +apply (auto simp add: Rep_matrix_mult foldseq_zero) +apply (rule_tac foldseq_zerotail[symmetric]) +apply (auto simp add: ncols zero_imp_mult_zero max1) +apply (rule order_trans) +apply (rule nrows_move_matrix_le) +apply (simp add: max2) +done + +lemma move_matrix_add: "((move_matrix (A + B) j i)::(('a::monoid_add) matrix)) = (move_matrix A j i) + (move_matrix B j i)" +apply (subst Rep_matrix_inject[symmetric]) +apply (rule ext)+ +apply (simp) +done + +lemma move_matrix_mult: "move_matrix ((A::('a::semiring_0) matrix)*B) j i = (move_matrix A j 0) * (move_matrix B 0 i)" +by (simp add: move_matrix_ortho[of "A*B"] move_matrix_col_mult move_matrix_row_mult) + +definition scalar_mult :: "('a::ring) \ 'a matrix \ 'a matrix" where + "scalar_mult a m == apply_matrix (op * a) m" + +lemma scalar_mult_zero[simp]: "scalar_mult y 0 = 0" +by (simp add: scalar_mult_def) + +lemma scalar_mult_add: "scalar_mult y (a+b) = (scalar_mult y a) + (scalar_mult y b)" +by (simp add: scalar_mult_def apply_matrix_add algebra_simps) + +lemma Rep_scalar_mult[simp]: "Rep_matrix (scalar_mult y a) j i = y * (Rep_matrix a j i)" +by (simp add: scalar_mult_def) + +lemma scalar_mult_singleton[simp]: "scalar_mult y (singleton_matrix j i x) = singleton_matrix j i (y * x)" +apply (subst Rep_matrix_inject[symmetric]) +apply (rule ext)+ +apply (auto) +done + +lemma Rep_minus[simp]: "Rep_matrix (-(A::_::group_add)) x y = - (Rep_matrix A x y)" +by (simp add: minus_matrix_def) + +lemma Rep_abs[simp]: "Rep_matrix (abs (A::_::lattice_ab_group_add)) x y = abs (Rep_matrix A x y)" +by (simp add: abs_lattice sup_matrix_def) + +end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/ROOT.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,2 @@ + +use_thy "Cplex"; diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/SparseMatrix.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/SparseMatrix.thy Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,1070 @@ +(* Title: HOL/Matrix/SparseMatrix.thy + Author: Steven Obua +*) + +theory SparseMatrix +imports Matrix +begin + +type_synonym 'a spvec = "(nat * 'a) list" +type_synonym 'a spmat = "'a spvec spvec" + +definition sparse_row_vector :: "('a::ab_group_add) spvec \ 'a matrix" + where "sparse_row_vector arr = foldl (% m x. m + (singleton_matrix 0 (fst x) (snd x))) 0 arr" + +definition sparse_row_matrix :: "('a::ab_group_add) spmat \ 'a matrix" + where "sparse_row_matrix arr = foldl (% m r. m + (move_matrix (sparse_row_vector (snd r)) (int (fst r)) 0)) 0 arr" + +code_datatype sparse_row_vector sparse_row_matrix + +lemma sparse_row_vector_empty [simp]: "sparse_row_vector [] = 0" + by (simp add: sparse_row_vector_def) + +lemma sparse_row_matrix_empty [simp]: "sparse_row_matrix [] = 0" + by (simp add: sparse_row_matrix_def) + +lemmas [code] = sparse_row_vector_empty [symmetric] + +lemma foldl_distrstart: "! a x y. (f (g x y) a = g x (f y a)) \ (foldl f (g x y) l = g x (foldl f y l))" + by (induct l arbitrary: x y, auto) + +lemma sparse_row_vector_cons[simp]: + "sparse_row_vector (a # arr) = (singleton_matrix 0 (fst a) (snd a)) + (sparse_row_vector arr)" + apply (induct arr) + apply (auto simp add: sparse_row_vector_def) + apply (simp add: foldl_distrstart [of "\m x. m + singleton_matrix 0 (fst x) (snd x)" "\x m. singleton_matrix 0 (fst x) (snd x) + m"]) + done + +lemma sparse_row_vector_append[simp]: + "sparse_row_vector (a @ b) = (sparse_row_vector a) + (sparse_row_vector b)" + by (induct a) auto + +lemma nrows_spvec[simp]: "nrows (sparse_row_vector x) <= (Suc 0)" + apply (induct x) + apply (simp_all add: add_nrows) + done + +lemma sparse_row_matrix_cons: "sparse_row_matrix (a#arr) = ((move_matrix (sparse_row_vector (snd a)) (int (fst a)) 0)) + sparse_row_matrix arr" + apply (induct arr) + apply (auto simp add: sparse_row_matrix_def) + apply (simp add: foldl_distrstart[of "\m x. m + (move_matrix (sparse_row_vector (snd x)) (int (fst x)) 0)" + "% a m. (move_matrix (sparse_row_vector (snd a)) (int (fst a)) 0) + m"]) + done + +lemma sparse_row_matrix_append: "sparse_row_matrix (arr@brr) = (sparse_row_matrix arr) + (sparse_row_matrix brr)" + apply (induct arr) + apply (auto simp add: sparse_row_matrix_cons) + done + +primrec sorted_spvec :: "'a spvec \ bool" +where + "sorted_spvec [] = True" +| sorted_spvec_step: "sorted_spvec (a#as) = (case as of [] \ True | b#bs \ ((fst a < fst b) & (sorted_spvec as)))" + +primrec sorted_spmat :: "'a spmat \ bool" +where + "sorted_spmat [] = True" +| "sorted_spmat (a#as) = ((sorted_spvec (snd a)) & (sorted_spmat as))" + +declare sorted_spvec.simps [simp del] + +lemma sorted_spvec_empty[simp]: "sorted_spvec [] = True" +by (simp add: sorted_spvec.simps) + +lemma sorted_spvec_cons1: "sorted_spvec (a#as) \ sorted_spvec as" +apply (induct as) +apply (auto simp add: sorted_spvec.simps) +done + +lemma sorted_spvec_cons2: "sorted_spvec (a#b#t) \ sorted_spvec (a#t)" +apply (induct t) +apply (auto simp add: sorted_spvec.simps) +done + +lemma sorted_spvec_cons3: "sorted_spvec(a#b#t) \ fst a < fst b" +apply (auto simp add: sorted_spvec.simps) +done + +lemma sorted_sparse_row_vector_zero[rule_format]: "m <= n \ sorted_spvec ((n,a)#arr) \ Rep_matrix (sparse_row_vector arr) j m = 0" +apply (induct arr) +apply (auto) +apply (frule sorted_spvec_cons2,simp)+ +apply (frule sorted_spvec_cons3, simp) +done + +lemma sorted_sparse_row_matrix_zero[rule_format]: "m <= n \ sorted_spvec ((n,a)#arr) \ Rep_matrix (sparse_row_matrix arr) m j = 0" + apply (induct arr) + apply (auto) + apply (frule sorted_spvec_cons2, simp) + apply (frule sorted_spvec_cons3, simp) + apply (simp add: sparse_row_matrix_cons) + done + +primrec minus_spvec :: "('a::ab_group_add) spvec \ 'a spvec" +where + "minus_spvec [] = []" +| "minus_spvec (a#as) = (fst a, -(snd a))#(minus_spvec as)" + +primrec abs_spvec :: "('a::lattice_ab_group_add_abs) spvec \ 'a spvec" +where + "abs_spvec [] = []" +| "abs_spvec (a#as) = (fst a, abs (snd a))#(abs_spvec as)" + +lemma sparse_row_vector_minus: + "sparse_row_vector (minus_spvec v) = - (sparse_row_vector v)" + apply (induct v) + apply (simp_all add: sparse_row_vector_cons) + apply (simp add: Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply simp + done + +instance matrix :: (lattice_ab_group_add_abs) lattice_ab_group_add_abs +apply default +unfolding abs_matrix_def .. (*FIXME move*) + +lemma sparse_row_vector_abs: + "sorted_spvec (v :: 'a::lattice_ring spvec) \ sparse_row_vector (abs_spvec v) = abs (sparse_row_vector v)" + apply (induct v) + apply simp_all + apply (frule_tac sorted_spvec_cons1, simp) + apply (simp only: Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply auto + apply (subgoal_tac "Rep_matrix (sparse_row_vector v) 0 a = 0") + apply (simp) + apply (rule sorted_sparse_row_vector_zero) + apply auto + done + +lemma sorted_spvec_minus_spvec: + "sorted_spvec v \ sorted_spvec (minus_spvec v)" + apply (induct v) + apply (simp) + apply (frule sorted_spvec_cons1, simp) + apply (simp add: sorted_spvec.simps split:list.split_asm) + done + +lemma sorted_spvec_abs_spvec: + "sorted_spvec v \ sorted_spvec (abs_spvec v)" + apply (induct v) + apply (simp) + apply (frule sorted_spvec_cons1, simp) + apply (simp add: sorted_spvec.simps split:list.split_asm) + done + +definition "smult_spvec y = map (% a. (fst a, y * snd a))" + +lemma smult_spvec_empty[simp]: "smult_spvec y [] = []" + by (simp add: smult_spvec_def) + +lemma smult_spvec_cons: "smult_spvec y (a#arr) = (fst a, y * (snd a)) # (smult_spvec y arr)" + by (simp add: smult_spvec_def) + +fun addmult_spvec :: "('a::ring) \ 'a spvec \ 'a spvec \ 'a spvec" +where + "addmult_spvec y arr [] = arr" +| "addmult_spvec y [] brr = smult_spvec y brr" +| "addmult_spvec y ((i,a)#arr) ((j,b)#brr) = ( + if i < j then ((i,a)#(addmult_spvec y arr ((j,b)#brr))) + else (if (j < i) then ((j, y * b)#(addmult_spvec y ((i,a)#arr) brr)) + else ((i, a + y*b)#(addmult_spvec y arr brr))))" +(* Steven used termination "measure (% (y, a, b). length a + (length b))" *) + +lemma addmult_spvec_empty1[simp]: "addmult_spvec y [] a = smult_spvec y a" + by (induct a) auto + +lemma addmult_spvec_empty2[simp]: "addmult_spvec y a [] = a" + by (induct a) auto + +lemma sparse_row_vector_map: "(! x y. f (x+y) = (f x) + (f y)) \ (f::'a\('a::lattice_ring)) 0 = 0 \ + sparse_row_vector (map (% x. (fst x, f (snd x))) a) = apply_matrix f (sparse_row_vector a)" + apply (induct a) + apply (simp_all add: apply_matrix_add) + done + +lemma sparse_row_vector_smult: "sparse_row_vector (smult_spvec y a) = scalar_mult y (sparse_row_vector a)" + apply (induct a) + apply (simp_all add: smult_spvec_cons scalar_mult_add) + done + +lemma sparse_row_vector_addmult_spvec: "sparse_row_vector (addmult_spvec (y::'a::lattice_ring) a b) = + (sparse_row_vector a) + (scalar_mult y (sparse_row_vector b))" + apply (induct y a b rule: addmult_spvec.induct) + apply (simp add: scalar_mult_add smult_spvec_cons sparse_row_vector_smult singleton_matrix_add)+ + done + +lemma sorted_smult_spvec: "sorted_spvec a \ sorted_spvec (smult_spvec y a)" + apply (auto simp add: smult_spvec_def) + apply (induct a) + apply (auto simp add: sorted_spvec.simps split:list.split_asm) + done + +lemma sorted_spvec_addmult_spvec_helper: "\sorted_spvec (addmult_spvec y ((a, b) # arr) brr); aa < a; sorted_spvec ((a, b) # arr); + sorted_spvec ((aa, ba) # brr)\ \ sorted_spvec ((aa, y * ba) # addmult_spvec y ((a, b) # arr) brr)" + apply (induct brr) + apply (auto simp add: sorted_spvec.simps) + done + +lemma sorted_spvec_addmult_spvec_helper2: + "\sorted_spvec (addmult_spvec y arr ((aa, ba) # brr)); a < aa; sorted_spvec ((a, b) # arr); sorted_spvec ((aa, ba) # brr)\ + \ sorted_spvec ((a, b) # addmult_spvec y arr ((aa, ba) # brr))" + apply (induct arr) + apply (auto simp add: smult_spvec_def sorted_spvec.simps) + done + +lemma sorted_spvec_addmult_spvec_helper3[rule_format]: + "sorted_spvec (addmult_spvec y arr brr) \ sorted_spvec ((aa, b) # arr) \ sorted_spvec ((aa, ba) # brr) + \ sorted_spvec ((aa, b + y * ba) # (addmult_spvec y arr brr))" + apply (induct y arr brr rule: addmult_spvec.induct) + apply (simp_all add: sorted_spvec.simps smult_spvec_def split:list.split) + done + +lemma sorted_addmult_spvec: "sorted_spvec a \ sorted_spvec b \ sorted_spvec (addmult_spvec y a b)" + apply (induct y a b rule: addmult_spvec.induct) + apply (simp_all add: sorted_smult_spvec) + apply (rule conjI, intro strip) + apply (case_tac "~(i < j)") + apply (simp_all) + apply (frule_tac as=brr in sorted_spvec_cons1) + apply (simp add: sorted_spvec_addmult_spvec_helper) + apply (intro strip | rule conjI)+ + apply (frule_tac as=arr in sorted_spvec_cons1) + apply (simp add: sorted_spvec_addmult_spvec_helper2) + apply (intro strip) + apply (frule_tac as=arr in sorted_spvec_cons1) + apply (frule_tac as=brr in sorted_spvec_cons1) + apply (simp) + apply (simp_all add: sorted_spvec_addmult_spvec_helper3) + done + +fun mult_spvec_spmat :: "('a::lattice_ring) spvec \ 'a spvec \ 'a spmat \ 'a spvec" +where +(* recdef mult_spvec_spmat "measure (% (c, arr, brr). (length arr) + (length brr))" *) + "mult_spvec_spmat c [] brr = c" +| "mult_spvec_spmat c arr [] = c" +| "mult_spvec_spmat c ((i,a)#arr) ((j,b)#brr) = ( + if (i < j) then mult_spvec_spmat c arr ((j,b)#brr) + else if (j < i) then mult_spvec_spmat c ((i,a)#arr) brr + else mult_spvec_spmat (addmult_spvec a c b) arr brr)" + +lemma sparse_row_mult_spvec_spmat[rule_format]: "sorted_spvec (a::('a::lattice_ring) spvec) \ sorted_spvec B \ + sparse_row_vector (mult_spvec_spmat c a B) = (sparse_row_vector c) + (sparse_row_vector a) * (sparse_row_matrix B)" +proof - + have comp_1: "!! a b. a < b \ Suc 0 <= nat ((int b)-(int a))" by arith + have not_iff: "!! a b. a = b \ (~ a) = (~ b)" by simp + have max_helper: "!! a b. ~ (a <= max (Suc a) b) \ False" + by arith + { + fix a + fix v + assume a:"a < nrows(sparse_row_vector v)" + have b:"nrows(sparse_row_vector v) <= 1" by simp + note dummy = less_le_trans[of a "nrows (sparse_row_vector v)" 1, OF a b] + then have "a = 0" by simp + } + note nrows_helper = this + show ?thesis + apply (induct c a B rule: mult_spvec_spmat.induct) + apply simp+ + apply (rule conjI) + apply (intro strip) + apply (frule_tac as=brr in sorted_spvec_cons1) + apply (simp add: algebra_simps sparse_row_matrix_cons) + apply (simplesubst Rep_matrix_zero_imp_mult_zero) + apply (simp) + apply (rule disjI2) + apply (intro strip) + apply (subst nrows) + apply (rule order_trans[of _ 1]) + apply (simp add: comp_1)+ + apply (subst Rep_matrix_zero_imp_mult_zero) + apply (intro strip) + apply (case_tac "k <= j") + apply (rule_tac m1 = k and n1 = i and a1 = a in ssubst[OF sorted_sparse_row_vector_zero]) + apply (simp_all) + apply (rule disjI2) + apply (rule nrows) + apply (rule order_trans[of _ 1]) + apply (simp_all add: comp_1) + + apply (intro strip | rule conjI)+ + apply (frule_tac as=arr in sorted_spvec_cons1) + apply (simp add: algebra_simps) + apply (subst Rep_matrix_zero_imp_mult_zero) + apply (simp) + apply (rule disjI2) + apply (intro strip) + apply (simp add: sparse_row_matrix_cons) + apply (case_tac "i <= j") + apply (erule sorted_sparse_row_matrix_zero) + apply (simp_all) + apply (intro strip) + apply (case_tac "i=j") + apply (simp_all) + apply (frule_tac as=arr in sorted_spvec_cons1) + apply (frule_tac as=brr in sorted_spvec_cons1) + apply (simp add: sparse_row_matrix_cons algebra_simps sparse_row_vector_addmult_spvec) + apply (rule_tac B1 = "sparse_row_matrix brr" in ssubst[OF Rep_matrix_zero_imp_mult_zero]) + apply (auto) + apply (rule sorted_sparse_row_matrix_zero) + apply (simp_all) + apply (rule_tac A1 = "sparse_row_vector arr" in ssubst[OF Rep_matrix_zero_imp_mult_zero]) + apply (auto) + apply (rule_tac m=k and n = j and a = a and arr=arr in sorted_sparse_row_vector_zero) + apply (simp_all) + apply (drule nrows_notzero) + apply (drule nrows_helper) + apply (arith) + + apply (subst Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply (simp) + apply (subst Rep_matrix_mult) + apply (rule_tac j1=j in ssubst[OF foldseq_almostzero]) + apply (simp_all) + apply (intro strip, rule conjI) + apply (intro strip) + apply (drule_tac max_helper) + apply (simp) + apply (auto) + apply (rule zero_imp_mult_zero) + apply (rule disjI2) + apply (rule nrows) + apply (rule order_trans[of _ 1]) + apply (simp) + apply (simp) + done +qed + +lemma sorted_mult_spvec_spmat[rule_format]: + "sorted_spvec (c::('a::lattice_ring) spvec) \ sorted_spmat B \ sorted_spvec (mult_spvec_spmat c a B)" + apply (induct c a B rule: mult_spvec_spmat.induct) + apply (simp_all add: sorted_addmult_spvec) + done + +primrec mult_spmat :: "('a::lattice_ring) spmat \ 'a spmat \ 'a spmat" +where + "mult_spmat [] A = []" +| "mult_spmat (a#as) A = (fst a, mult_spvec_spmat [] (snd a) A)#(mult_spmat as A)" + +lemma sparse_row_mult_spmat: + "sorted_spmat A \ sorted_spvec B \ + sparse_row_matrix (mult_spmat A B) = (sparse_row_matrix A) * (sparse_row_matrix B)" + apply (induct A) + apply (auto simp add: sparse_row_matrix_cons sparse_row_mult_spvec_spmat algebra_simps move_matrix_mult) + done + +lemma sorted_spvec_mult_spmat[rule_format]: + "sorted_spvec (A::('a::lattice_ring) spmat) \ sorted_spvec (mult_spmat A B)" + apply (induct A) + apply (auto) + apply (drule sorted_spvec_cons1, simp) + apply (case_tac A) + apply (auto simp add: sorted_spvec.simps) + done + +lemma sorted_spmat_mult_spmat: + "sorted_spmat (B::('a::lattice_ring) spmat) \ sorted_spmat (mult_spmat A B)" + apply (induct A) + apply (auto simp add: sorted_mult_spvec_spmat) + done + + +fun add_spvec :: "('a::lattice_ab_group_add) spvec \ 'a spvec \ 'a spvec" +where +(* "measure (% (a, b). length a + (length b))" *) + "add_spvec arr [] = arr" +| "add_spvec [] brr = brr" +| "add_spvec ((i,a)#arr) ((j,b)#brr) = ( + if i < j then (i,a)#(add_spvec arr ((j,b)#brr)) + else if (j < i) then (j,b) # add_spvec ((i,a)#arr) brr + else (i, a+b) # add_spvec arr brr)" + +lemma add_spvec_empty1[simp]: "add_spvec [] a = a" +by (cases a, auto) + +lemma sparse_row_vector_add: "sparse_row_vector (add_spvec a b) = (sparse_row_vector a) + (sparse_row_vector b)" + apply (induct a b rule: add_spvec.induct) + apply (simp_all add: singleton_matrix_add) + done + +fun add_spmat :: "('a::lattice_ab_group_add) spmat \ 'a spmat \ 'a spmat" +where +(* "measure (% (A,B). (length A)+(length B))" *) + "add_spmat [] bs = bs" +| "add_spmat as [] = as" +| "add_spmat ((i,a)#as) ((j,b)#bs) = ( + if i < j then + (i,a) # add_spmat as ((j,b)#bs) + else if j < i then + (j,b) # add_spmat ((i,a)#as) bs + else + (i, add_spvec a b) # add_spmat as bs)" + +lemma add_spmat_Nil2[simp]: "add_spmat as [] = as" +by(cases as) auto + +lemma sparse_row_add_spmat: "sparse_row_matrix (add_spmat A B) = (sparse_row_matrix A) + (sparse_row_matrix B)" + apply (induct A B rule: add_spmat.induct) + apply (auto simp add: sparse_row_matrix_cons sparse_row_vector_add move_matrix_add) + done + +lemmas [code] = sparse_row_add_spmat [symmetric] +lemmas [code] = sparse_row_vector_add [symmetric] + +lemma sorted_add_spvec_helper1[rule_format]: "add_spvec ((a,b)#arr) brr = (ab, bb) # list \ (ab = a | (brr \ [] & ab = fst (hd brr)))" + proof - + have "(! x ab a. x = (a,b)#arr \ add_spvec x brr = (ab, bb) # list \ (ab = a | (ab = fst (hd brr))))" + by (induct brr rule: add_spvec.induct) (auto split:if_splits) + then show ?thesis + by (case_tac brr, auto) + qed + +lemma sorted_add_spmat_helper1[rule_format]: "add_spmat ((a,b)#arr) brr = (ab, bb) # list \ (ab = a | (brr \ [] & ab = fst (hd brr)))" + proof - + have "(! x ab a. x = (a,b)#arr \ add_spmat x brr = (ab, bb) # list \ (ab = a | (ab = fst (hd brr))))" + by (rule add_spmat.induct) (auto split:if_splits) + then show ?thesis + by (case_tac brr, auto) + qed + +lemma sorted_add_spvec_helper: "add_spvec arr brr = (ab, bb) # list \ ((arr \ [] & ab = fst (hd arr)) | (brr \ [] & ab = fst (hd brr)))" + apply (induct arr brr rule: add_spvec.induct) + apply (auto split:if_splits) + done + +lemma sorted_add_spmat_helper: "add_spmat arr brr = (ab, bb) # list \ ((arr \ [] & ab = fst (hd arr)) | (brr \ [] & ab = fst (hd brr)))" + apply (induct arr brr rule: add_spmat.induct) + apply (auto split:if_splits) + done + +lemma add_spvec_commute: "add_spvec a b = add_spvec b a" +by (induct a b rule: add_spvec.induct) auto + +lemma add_spmat_commute: "add_spmat a b = add_spmat b a" + apply (induct a b rule: add_spmat.induct) + apply (simp_all add: add_spvec_commute) + done + +lemma sorted_add_spvec_helper2: "add_spvec ((a,b)#arr) brr = (ab, bb) # list \ aa < a \ sorted_spvec ((aa, ba) # brr) \ aa < ab" + apply (drule sorted_add_spvec_helper1) + apply (auto) + apply (case_tac brr) + apply (simp_all) + apply (drule_tac sorted_spvec_cons3) + apply (simp) + done + +lemma sorted_add_spmat_helper2: "add_spmat ((a,b)#arr) brr = (ab, bb) # list \ aa < a \ sorted_spvec ((aa, ba) # brr) \ aa < ab" + apply (drule sorted_add_spmat_helper1) + apply (auto) + apply (case_tac brr) + apply (simp_all) + apply (drule_tac sorted_spvec_cons3) + apply (simp) + done + +lemma sorted_spvec_add_spvec[rule_format]: "sorted_spvec a \ sorted_spvec b \ sorted_spvec (add_spvec a b)" + apply (induct a b rule: add_spvec.induct) + apply (simp_all) + apply (rule conjI) + apply (clarsimp) + apply (frule_tac as=brr in sorted_spvec_cons1) + apply (simp) + apply (subst sorted_spvec_step) + apply (clarsimp simp: sorted_add_spvec_helper2 split: list.split) + apply (clarify) + apply (rule conjI) + apply (clarify) + apply (frule_tac as=arr in sorted_spvec_cons1, simp) + apply (subst sorted_spvec_step) + apply (clarsimp simp: sorted_add_spvec_helper2 add_spvec_commute split: list.split) + apply (clarify) + apply (frule_tac as=arr in sorted_spvec_cons1) + apply (frule_tac as=brr in sorted_spvec_cons1) + apply (simp) + apply (subst sorted_spvec_step) + apply (simp split: list.split) + apply (clarsimp) + apply (drule_tac sorted_add_spvec_helper) + apply (auto simp: neq_Nil_conv) + apply (drule sorted_spvec_cons3) + apply (simp) + apply (drule sorted_spvec_cons3) + apply (simp) + done + +lemma sorted_spvec_add_spmat[rule_format]: "sorted_spvec A \ sorted_spvec B \ sorted_spvec (add_spmat A B)" + apply (induct A B rule: add_spmat.induct) + apply (simp_all) + apply (rule conjI) + apply (intro strip) + apply (simp) + apply (frule_tac as=bs in sorted_spvec_cons1) + apply (simp) + apply (subst sorted_spvec_step) + apply (simp split: list.split) + apply (clarify, simp) + apply (simp add: sorted_add_spmat_helper2) + apply (clarify) + apply (rule conjI) + apply (clarify) + apply (frule_tac as=as in sorted_spvec_cons1, simp) + apply (subst sorted_spvec_step) + apply (clarsimp simp: sorted_add_spmat_helper2 add_spmat_commute split: list.split) + apply (clarsimp) + apply (frule_tac as=as in sorted_spvec_cons1) + apply (frule_tac as=bs in sorted_spvec_cons1) + apply (simp) + apply (subst sorted_spvec_step) + apply (simp split: list.split) + apply (clarify, simp) + apply (drule_tac sorted_add_spmat_helper) + apply (auto simp:neq_Nil_conv) + apply (drule sorted_spvec_cons3) + apply (simp) + apply (drule sorted_spvec_cons3) + apply (simp) + done + +lemma sorted_spmat_add_spmat[rule_format]: "sorted_spmat A \ sorted_spmat B \ sorted_spmat (add_spmat A B)" + apply (induct A B rule: add_spmat.induct) + apply (simp_all add: sorted_spvec_add_spvec) + done + +fun le_spvec :: "('a::lattice_ab_group_add) spvec \ 'a spvec \ bool" +where +(* "measure (% (a,b). (length a) + (length b))" *) + "le_spvec [] [] = True" +| "le_spvec ((_,a)#as) [] = (a <= 0 & le_spvec as [])" +| "le_spvec [] ((_,b)#bs) = (0 <= b & le_spvec [] bs)" +| "le_spvec ((i,a)#as) ((j,b)#bs) = ( + if (i < j) then a <= 0 & le_spvec as ((j,b)#bs) + else if (j < i) then 0 <= b & le_spvec ((i,a)#as) bs + else a <= b & le_spvec as bs)" + +fun le_spmat :: "('a::lattice_ab_group_add) spmat \ 'a spmat \ bool" +where +(* "measure (% (a,b). (length a) + (length b))" *) + "le_spmat [] [] = True" +| "le_spmat ((i,a)#as) [] = (le_spvec a [] & le_spmat as [])" +| "le_spmat [] ((j,b)#bs) = (le_spvec [] b & le_spmat [] bs)" +| "le_spmat ((i,a)#as) ((j,b)#bs) = ( + if i < j then (le_spvec a [] & le_spmat as ((j,b)#bs)) + else if j < i then (le_spvec [] b & le_spmat ((i,a)#as) bs) + else (le_spvec a b & le_spmat as bs))" + +definition disj_matrices :: "('a::zero) matrix \ 'a matrix \ bool" where + "disj_matrices A B \ + (! j i. (Rep_matrix A j i \ 0) \ (Rep_matrix B j i = 0)) & (! j i. (Rep_matrix B j i \ 0) \ (Rep_matrix A j i = 0))" + +declare [[simp_depth_limit = 6]] + +lemma disj_matrices_contr1: "disj_matrices A B \ Rep_matrix A j i \ 0 \ Rep_matrix B j i = 0" + by (simp add: disj_matrices_def) + +lemma disj_matrices_contr2: "disj_matrices A B \ Rep_matrix B j i \ 0 \ Rep_matrix A j i = 0" + by (simp add: disj_matrices_def) + + +lemma disj_matrices_add: "disj_matrices A B \ disj_matrices C D \ disj_matrices A D \ disj_matrices B C \ + (A + B <= C + D) = (A <= C & B <= (D::('a::lattice_ab_group_add) matrix))" + apply (auto) + apply (simp (no_asm_use) only: le_matrix_def disj_matrices_def) + apply (intro strip) + apply (erule conjE)+ + apply (drule_tac j=j and i=i in spec2)+ + apply (case_tac "Rep_matrix B j i = 0") + apply (case_tac "Rep_matrix D j i = 0") + apply (simp_all) + apply (simp (no_asm_use) only: le_matrix_def disj_matrices_def) + apply (intro strip) + apply (erule conjE)+ + apply (drule_tac j=j and i=i in spec2)+ + apply (case_tac "Rep_matrix A j i = 0") + apply (case_tac "Rep_matrix C j i = 0") + apply (simp_all) + apply (erule add_mono) + apply (assumption) + done + +lemma disj_matrices_zero1[simp]: "disj_matrices 0 B" +by (simp add: disj_matrices_def) + +lemma disj_matrices_zero2[simp]: "disj_matrices A 0" +by (simp add: disj_matrices_def) + +lemma disj_matrices_commute: "disj_matrices A B = disj_matrices B A" +by (auto simp add: disj_matrices_def) + +lemma disj_matrices_add_le_zero: "disj_matrices A B \ + (A + B <= 0) = (A <= 0 & (B::('a::lattice_ab_group_add) matrix) <= 0)" +by (rule disj_matrices_add[of A B 0 0, simplified]) + +lemma disj_matrices_add_zero_le: "disj_matrices A B \ + (0 <= A + B) = (0 <= A & 0 <= (B::('a::lattice_ab_group_add) matrix))" +by (rule disj_matrices_add[of 0 0 A B, simplified]) + +lemma disj_matrices_add_x_le: "disj_matrices A B \ disj_matrices B C \ + (A <= B + C) = (A <= C & 0 <= (B::('a::lattice_ab_group_add) matrix))" +by (auto simp add: disj_matrices_add[of 0 A B C, simplified]) + +lemma disj_matrices_add_le_x: "disj_matrices A B \ disj_matrices B C \ + (B + A <= C) = (A <= C & (B::('a::lattice_ab_group_add) matrix) <= 0)" +by (auto simp add: disj_matrices_add[of B A 0 C,simplified] disj_matrices_commute) + +lemma disj_sparse_row_singleton: "i <= j \ sorted_spvec((j,y)#v) \ disj_matrices (sparse_row_vector v) (singleton_matrix 0 i x)" + apply (simp add: disj_matrices_def) + apply (rule conjI) + apply (rule neg_imp) + apply (simp) + apply (intro strip) + apply (rule sorted_sparse_row_vector_zero) + apply (simp_all) + apply (intro strip) + apply (rule sorted_sparse_row_vector_zero) + apply (simp_all) + done + +lemma disj_matrices_x_add: "disj_matrices A B \ disj_matrices A C \ disj_matrices (A::('a::lattice_ab_group_add) matrix) (B+C)" + apply (simp add: disj_matrices_def) + apply (auto) + apply (drule_tac j=j and i=i in spec2)+ + apply (case_tac "Rep_matrix B j i = 0") + apply (case_tac "Rep_matrix C j i = 0") + apply (simp_all) + done + +lemma disj_matrices_add_x: "disj_matrices A B \ disj_matrices A C \ disj_matrices (B+C) (A::('a::lattice_ab_group_add) matrix)" + by (simp add: disj_matrices_x_add disj_matrices_commute) + +lemma disj_singleton_matrices[simp]: "disj_matrices (singleton_matrix j i x) (singleton_matrix u v y) = (j \ u | i \ v | x = 0 | y = 0)" + by (auto simp add: disj_matrices_def) + +lemma disj_move_sparse_vec_mat[simplified disj_matrices_commute]: + "j <= a \ sorted_spvec((a,c)#as) \ disj_matrices (move_matrix (sparse_row_vector b) (int j) i) (sparse_row_matrix as)" + apply (auto simp add: disj_matrices_def) + apply (drule nrows_notzero) + apply (drule less_le_trans[OF _ nrows_spvec]) + apply (subgoal_tac "ja = j") + apply (simp add: sorted_sparse_row_matrix_zero) + apply (arith) + apply (rule nrows) + apply (rule order_trans[of _ 1 _]) + apply (simp) + apply (case_tac "nat (int ja - int j) = 0") + apply (case_tac "ja = j") + apply (simp add: sorted_sparse_row_matrix_zero) + apply arith+ + done + +lemma disj_move_sparse_row_vector_twice: + "j \ u \ disj_matrices (move_matrix (sparse_row_vector a) j i) (move_matrix (sparse_row_vector b) u v)" + apply (auto simp add: disj_matrices_def) + apply (rule nrows, rule order_trans[of _ 1], simp, drule nrows_notzero, drule less_le_trans[OF _ nrows_spvec], arith)+ + done + +lemma le_spvec_iff_sparse_row_le[rule_format]: "(sorted_spvec a) \ (sorted_spvec b) \ (le_spvec a b) = (sparse_row_vector a <= sparse_row_vector b)" + apply (induct a b rule: le_spvec.induct) + apply (simp_all add: sorted_spvec_cons1 disj_matrices_add_le_zero disj_matrices_add_zero_le + disj_sparse_row_singleton[OF order_refl] disj_matrices_commute) + apply (rule conjI, intro strip) + apply (simp add: sorted_spvec_cons1) + apply (subst disj_matrices_add_x_le) + apply (simp add: disj_sparse_row_singleton[OF less_imp_le] disj_matrices_x_add disj_matrices_commute) + apply (simp add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute) + apply (simp, blast) + apply (intro strip, rule conjI, intro strip) + apply (simp add: sorted_spvec_cons1) + apply (subst disj_matrices_add_le_x) + apply (simp_all add: disj_sparse_row_singleton[OF order_refl] disj_sparse_row_singleton[OF less_imp_le] disj_matrices_commute disj_matrices_x_add) + apply (blast) + apply (intro strip) + apply (simp add: sorted_spvec_cons1) + apply (case_tac "a=b", simp_all) + apply (subst disj_matrices_add) + apply (simp_all add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute) + done + +lemma le_spvec_empty2_sparse_row[rule_format]: "sorted_spvec b \ le_spvec b [] = (sparse_row_vector b <= 0)" + apply (induct b) + apply (simp_all add: sorted_spvec_cons1) + apply (intro strip) + apply (subst disj_matrices_add_le_zero) + apply (auto simp add: disj_matrices_commute disj_sparse_row_singleton[OF order_refl] sorted_spvec_cons1) + done + +lemma le_spvec_empty1_sparse_row[rule_format]: "(sorted_spvec b) \ (le_spvec [] b = (0 <= sparse_row_vector b))" + apply (induct b) + apply (simp_all add: sorted_spvec_cons1) + apply (intro strip) + apply (subst disj_matrices_add_zero_le) + apply (auto simp add: disj_matrices_commute disj_sparse_row_singleton[OF order_refl] sorted_spvec_cons1) + done + +lemma le_spmat_iff_sparse_row_le[rule_format]: "(sorted_spvec A) \ (sorted_spmat A) \ (sorted_spvec B) \ (sorted_spmat B) \ + le_spmat A B = (sparse_row_matrix A <= sparse_row_matrix B)" + apply (induct A B rule: le_spmat.induct) + apply (simp add: sparse_row_matrix_cons disj_matrices_add_le_zero disj_matrices_add_zero_le disj_move_sparse_vec_mat[OF order_refl] + disj_matrices_commute sorted_spvec_cons1 le_spvec_empty2_sparse_row le_spvec_empty1_sparse_row)+ + apply (rule conjI, intro strip) + apply (simp add: sorted_spvec_cons1) + apply (subst disj_matrices_add_x_le) + apply (rule disj_matrices_add_x) + apply (simp add: disj_move_sparse_row_vector_twice) + apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute) + apply (simp add: disj_move_sparse_vec_mat[OF order_refl] disj_matrices_commute) + apply (simp, blast) + apply (intro strip, rule conjI, intro strip) + apply (simp add: sorted_spvec_cons1) + apply (subst disj_matrices_add_le_x) + apply (simp add: disj_move_sparse_vec_mat[OF order_refl]) + apply (rule disj_matrices_x_add) + apply (simp add: disj_move_sparse_row_vector_twice) + apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute) + apply (simp, blast) + apply (intro strip) + apply (case_tac "i=j") + apply (simp_all) + apply (subst disj_matrices_add) + apply (simp_all add: disj_matrices_commute disj_move_sparse_vec_mat[OF order_refl]) + apply (simp add: sorted_spvec_cons1 le_spvec_iff_sparse_row_le) + done + +declare [[simp_depth_limit = 999]] + +primrec abs_spmat :: "('a::lattice_ring) spmat \ 'a spmat" +where + "abs_spmat [] = []" +| "abs_spmat (a#as) = (fst a, abs_spvec (snd a))#(abs_spmat as)" + +primrec minus_spmat :: "('a::lattice_ring) spmat \ 'a spmat" +where + "minus_spmat [] = []" +| "minus_spmat (a#as) = (fst a, minus_spvec (snd a))#(minus_spmat as)" + +lemma sparse_row_matrix_minus: + "sparse_row_matrix (minus_spmat A) = - (sparse_row_matrix A)" + apply (induct A) + apply (simp_all add: sparse_row_vector_minus sparse_row_matrix_cons) + apply (subst Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply simp + done + +lemma Rep_sparse_row_vector_zero: "x \ 0 \ Rep_matrix (sparse_row_vector v) x y = 0" +proof - + assume x:"x \ 0" + have r:"nrows (sparse_row_vector v) <= Suc 0" by (rule nrows_spvec) + show ?thesis + apply (rule nrows) + apply (subgoal_tac "Suc 0 <= x") + apply (insert r) + apply (simp only:) + apply (insert x) + apply arith + done +qed + +lemma sparse_row_matrix_abs: + "sorted_spvec A \ sorted_spmat A \ sparse_row_matrix (abs_spmat A) = abs (sparse_row_matrix A)" + apply (induct A) + apply (simp_all add: sparse_row_vector_abs sparse_row_matrix_cons) + apply (frule_tac sorted_spvec_cons1, simp) + apply (simplesubst Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply auto + apply (case_tac "x=a") + apply (simp) + apply (simplesubst sorted_sparse_row_matrix_zero) + apply auto + apply (simplesubst Rep_sparse_row_vector_zero) + apply simp_all + done + +lemma sorted_spvec_minus_spmat: "sorted_spvec A \ sorted_spvec (minus_spmat A)" + apply (induct A) + apply (simp) + apply (frule sorted_spvec_cons1, simp) + apply (simp add: sorted_spvec.simps split:list.split_asm) + done + +lemma sorted_spvec_abs_spmat: "sorted_spvec A \ sorted_spvec (abs_spmat A)" + apply (induct A) + apply (simp) + apply (frule sorted_spvec_cons1, simp) + apply (simp add: sorted_spvec.simps split:list.split_asm) + done + +lemma sorted_spmat_minus_spmat: "sorted_spmat A \ sorted_spmat (minus_spmat A)" + apply (induct A) + apply (simp_all add: sorted_spvec_minus_spvec) + done + +lemma sorted_spmat_abs_spmat: "sorted_spmat A \ sorted_spmat (abs_spmat A)" + apply (induct A) + apply (simp_all add: sorted_spvec_abs_spvec) + done + +definition diff_spmat :: "('a::lattice_ring) spmat \ 'a spmat \ 'a spmat" + where "diff_spmat A B = add_spmat A (minus_spmat B)" + +lemma sorted_spmat_diff_spmat: "sorted_spmat A \ sorted_spmat B \ sorted_spmat (diff_spmat A B)" + by (simp add: diff_spmat_def sorted_spmat_minus_spmat sorted_spmat_add_spmat) + +lemma sorted_spvec_diff_spmat: "sorted_spvec A \ sorted_spvec B \ sorted_spvec (diff_spmat A B)" + by (simp add: diff_spmat_def sorted_spvec_minus_spmat sorted_spvec_add_spmat) + +lemma sparse_row_diff_spmat: "sparse_row_matrix (diff_spmat A B ) = (sparse_row_matrix A) - (sparse_row_matrix B)" + by (simp add: diff_spmat_def sparse_row_add_spmat sparse_row_matrix_minus) + +definition sorted_sparse_matrix :: "'a spmat \ bool" + where "sorted_sparse_matrix A \ sorted_spvec A & sorted_spmat A" + +lemma sorted_sparse_matrix_imp_spvec: "sorted_sparse_matrix A \ sorted_spvec A" + by (simp add: sorted_sparse_matrix_def) + +lemma sorted_sparse_matrix_imp_spmat: "sorted_sparse_matrix A \ sorted_spmat A" + by (simp add: sorted_sparse_matrix_def) + +lemmas sorted_sp_simps = + sorted_spvec.simps + sorted_spmat.simps + sorted_sparse_matrix_def + +lemma bool1: "(\ True) = False" by blast +lemma bool2: "(\ False) = True" by blast +lemma bool3: "((P\bool) \ True) = P" by blast +lemma bool4: "(True \ (P\bool)) = P" by blast +lemma bool5: "((P\bool) \ False) = False" by blast +lemma bool6: "(False \ (P\bool)) = False" by blast +lemma bool7: "((P\bool) \ True) = True" by blast +lemma bool8: "(True \ (P\bool)) = True" by blast +lemma bool9: "((P\bool) \ False) = P" by blast +lemma bool10: "(False \ (P\bool)) = P" by blast +lemmas boolarith = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10 + +lemma if_case_eq: "(if b then x else y) = (case b of True => x | False => y)" by simp + +primrec pprt_spvec :: "('a::{lattice_ab_group_add}) spvec \ 'a spvec" +where + "pprt_spvec [] = []" +| "pprt_spvec (a#as) = (fst a, pprt (snd a)) # (pprt_spvec as)" + +primrec nprt_spvec :: "('a::{lattice_ab_group_add}) spvec \ 'a spvec" +where + "nprt_spvec [] = []" +| "nprt_spvec (a#as) = (fst a, nprt (snd a)) # (nprt_spvec as)" + +primrec pprt_spmat :: "('a::{lattice_ab_group_add}) spmat \ 'a spmat" +where + "pprt_spmat [] = []" +| "pprt_spmat (a#as) = (fst a, pprt_spvec (snd a))#(pprt_spmat as)" + +primrec nprt_spmat :: "('a::{lattice_ab_group_add}) spmat \ 'a spmat" +where + "nprt_spmat [] = []" +| "nprt_spmat (a#as) = (fst a, nprt_spvec (snd a))#(nprt_spmat as)" + + +lemma pprt_add: "disj_matrices A (B::(_::lattice_ring) matrix) \ pprt (A+B) = pprt A + pprt B" + apply (simp add: pprt_def sup_matrix_def) + apply (simp add: Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply simp + apply (case_tac "Rep_matrix A x xa \ 0") + apply (simp_all add: disj_matrices_contr1) + done + +lemma nprt_add: "disj_matrices A (B::(_::lattice_ring) matrix) \ nprt (A+B) = nprt A + nprt B" + apply (simp add: nprt_def inf_matrix_def) + apply (simp add: Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply simp + apply (case_tac "Rep_matrix A x xa \ 0") + apply (simp_all add: disj_matrices_contr1) + done + +lemma pprt_singleton[simp]: "pprt (singleton_matrix j i (x::_::lattice_ring)) = singleton_matrix j i (pprt x)" + apply (simp add: pprt_def sup_matrix_def) + apply (simp add: Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply simp + done + +lemma nprt_singleton[simp]: "nprt (singleton_matrix j i (x::_::lattice_ring)) = singleton_matrix j i (nprt x)" + apply (simp add: nprt_def inf_matrix_def) + apply (simp add: Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply simp + done + +lemma less_imp_le: "a < b \ a <= (b::_::order)" by (simp add: less_def) + +lemma sparse_row_vector_pprt: "sorted_spvec (v :: 'a::lattice_ring spvec) \ sparse_row_vector (pprt_spvec v) = pprt (sparse_row_vector v)" + apply (induct v) + apply (simp_all) + apply (frule sorted_spvec_cons1, auto) + apply (subst pprt_add) + apply (subst disj_matrices_commute) + apply (rule disj_sparse_row_singleton) + apply auto + done + +lemma sparse_row_vector_nprt: "sorted_spvec (v :: 'a::lattice_ring spvec) \ sparse_row_vector (nprt_spvec v) = nprt (sparse_row_vector v)" + apply (induct v) + apply (simp_all) + apply (frule sorted_spvec_cons1, auto) + apply (subst nprt_add) + apply (subst disj_matrices_commute) + apply (rule disj_sparse_row_singleton) + apply auto + done + + +lemma pprt_move_matrix: "pprt (move_matrix (A::('a::lattice_ring) matrix) j i) = move_matrix (pprt A) j i" + apply (simp add: pprt_def) + apply (simp add: sup_matrix_def) + apply (simp add: Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply (simp) + done + +lemma nprt_move_matrix: "nprt (move_matrix (A::('a::lattice_ring) matrix) j i) = move_matrix (nprt A) j i" + apply (simp add: nprt_def) + apply (simp add: inf_matrix_def) + apply (simp add: Rep_matrix_inject[symmetric]) + apply (rule ext)+ + apply (simp) + done + +lemma sparse_row_matrix_pprt: "sorted_spvec (m :: 'a::lattice_ring spmat) \ sorted_spmat m \ sparse_row_matrix (pprt_spmat m) = pprt (sparse_row_matrix m)" + apply (induct m) + apply simp + apply simp + apply (frule sorted_spvec_cons1) + apply (simp add: sparse_row_matrix_cons sparse_row_vector_pprt) + apply (subst pprt_add) + apply (subst disj_matrices_commute) + apply (rule disj_move_sparse_vec_mat) + apply auto + apply (simp add: sorted_spvec.simps) + apply (simp split: list.split) + apply auto + apply (simp add: pprt_move_matrix) + done + +lemma sparse_row_matrix_nprt: "sorted_spvec (m :: 'a::lattice_ring spmat) \ sorted_spmat m \ sparse_row_matrix (nprt_spmat m) = nprt (sparse_row_matrix m)" + apply (induct m) + apply simp + apply simp + apply (frule sorted_spvec_cons1) + apply (simp add: sparse_row_matrix_cons sparse_row_vector_nprt) + apply (subst nprt_add) + apply (subst disj_matrices_commute) + apply (rule disj_move_sparse_vec_mat) + apply auto + apply (simp add: sorted_spvec.simps) + apply (simp split: list.split) + apply auto + apply (simp add: nprt_move_matrix) + done + +lemma sorted_pprt_spvec: "sorted_spvec v \ sorted_spvec (pprt_spvec v)" + apply (induct v) + apply (simp) + apply (frule sorted_spvec_cons1) + apply simp + apply (simp add: sorted_spvec.simps split:list.split_asm) + done + +lemma sorted_nprt_spvec: "sorted_spvec v \ sorted_spvec (nprt_spvec v)" + apply (induct v) + apply (simp) + apply (frule sorted_spvec_cons1) + apply simp + apply (simp add: sorted_spvec.simps split:list.split_asm) + done + +lemma sorted_spvec_pprt_spmat: "sorted_spvec m \ sorted_spvec (pprt_spmat m)" + apply (induct m) + apply (simp) + apply (frule sorted_spvec_cons1) + apply simp + apply (simp add: sorted_spvec.simps split:list.split_asm) + done + +lemma sorted_spvec_nprt_spmat: "sorted_spvec m \ sorted_spvec (nprt_spmat m)" + apply (induct m) + apply (simp) + apply (frule sorted_spvec_cons1) + apply simp + apply (simp add: sorted_spvec.simps split:list.split_asm) + done + +lemma sorted_spmat_pprt_spmat: "sorted_spmat m \ sorted_spmat (pprt_spmat m)" + apply (induct m) + apply (simp_all add: sorted_pprt_spvec) + done + +lemma sorted_spmat_nprt_spmat: "sorted_spmat m \ sorted_spmat (nprt_spmat m)" + apply (induct m) + apply (simp_all add: sorted_nprt_spvec) + done + +definition mult_est_spmat :: "('a::lattice_ring) spmat \ 'a spmat \ 'a spmat \ 'a spmat \ 'a spmat" where + "mult_est_spmat r1 r2 s1 s2 = + add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2)) (add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2)) + (add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1)) (mult_spmat (nprt_spmat s1) (nprt_spmat r1))))" + +lemmas sparse_row_matrix_op_simps = + sorted_sparse_matrix_imp_spmat sorted_sparse_matrix_imp_spvec + sparse_row_add_spmat sorted_spvec_add_spmat sorted_spmat_add_spmat + sparse_row_diff_spmat sorted_spvec_diff_spmat sorted_spmat_diff_spmat + sparse_row_matrix_minus sorted_spvec_minus_spmat sorted_spmat_minus_spmat + sparse_row_mult_spmat sorted_spvec_mult_spmat sorted_spmat_mult_spmat + sparse_row_matrix_abs sorted_spvec_abs_spmat sorted_spmat_abs_spmat + le_spmat_iff_sparse_row_le + sparse_row_matrix_pprt sorted_spvec_pprt_spmat sorted_spmat_pprt_spmat + sparse_row_matrix_nprt sorted_spvec_nprt_spmat sorted_spmat_nprt_spmat + +lemma zero_eq_Numeral0: "(0::_::number_ring) = Numeral0" by simp + +lemmas sparse_row_matrix_arith_simps[simplified zero_eq_Numeral0] = + mult_spmat.simps mult_spvec_spmat.simps + addmult_spvec.simps + smult_spvec_empty smult_spvec_cons + add_spmat.simps add_spvec.simps + minus_spmat.simps minus_spvec.simps + abs_spmat.simps abs_spvec.simps + diff_spmat_def + le_spmat.simps le_spvec.simps + pprt_spmat.simps pprt_spvec.simps + nprt_spmat.simps nprt_spvec.simps + mult_est_spmat_def + + +(*lemma spm_linprog_dual_estimate_1: + assumes + "sorted_sparse_matrix A1" + "sorted_sparse_matrix A2" + "sorted_sparse_matrix c1" + "sorted_sparse_matrix c2" + "sorted_sparse_matrix y" + "sorted_spvec b" + "sorted_spvec r" + "le_spmat ([], y)" + "A * x \ sparse_row_matrix (b::('a::lattice_ring) spmat)" + "sparse_row_matrix A1 <= A" + "A <= sparse_row_matrix A2" + "sparse_row_matrix c1 <= c" + "c <= sparse_row_matrix c2" + "abs x \ sparse_row_matrix r" + shows + "c * x \ sparse_row_matrix (add_spmat (mult_spmat y b, mult_spmat (add_spmat (add_spmat (mult_spmat y (diff_spmat A2 A1), + abs_spmat (diff_spmat (mult_spmat y A1) c1)), diff_spmat c2 c1)) r))" + by (insert prems, simp add: sparse_row_matrix_op_simps linprog_dual_estimate_1[where A=A]) +*) + +end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/document/root.tex Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,25 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\newcommand{\ganz}{\mathsf{Z}\mkern-7.5mu\mathsf{Z}} + +\begin{document} + +\title{Matrix} +\author{Steven Obua} +\maketitle + +%\tableofcontents + +\parindent 0pt\parskip 0.5ex + +\input{session} + +\end{document} diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/fspmlp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/fspmlp.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,313 @@ +(* Title: HOL/Matrix/fspmlp.ML + Author: Steven Obua +*) + +signature FSPMLP = +sig + type linprog + type vector = FloatSparseMatrixBuilder.vector + type matrix = FloatSparseMatrixBuilder.matrix + + val y : linprog -> term + val A : linprog -> term * term + val b : linprog -> term + val c : linprog -> term * term + val r12 : linprog -> term * term + + exception Load of string + + val load : string -> int -> bool -> linprog +end + +structure Fspmlp : FSPMLP = +struct + +type vector = FloatSparseMatrixBuilder.vector +type matrix = FloatSparseMatrixBuilder.matrix + +type linprog = term * (term * term) * term * (term * term) * (term * term) + +fun y (c1, _, _, _, _) = c1 +fun A (_, c2, _, _, _) = c2 +fun b (_, _, c3, _, _) = c3 +fun c (_, _, _, c4, _) = c4 +fun r12 (_, _, _, _, c6) = c6 + +structure CplexFloatSparseMatrixConverter = +MAKE_CPLEX_MATRIX_CONVERTER(structure cplex = Cplex and matrix_builder = FloatSparseMatrixBuilder); + +datatype bound_type = LOWER | UPPER + +fun intbound_ord ((i1: int, b1),(i2,b2)) = + if i1 < i2 then LESS + else if i1 = i2 then + (if b1 = b2 then EQUAL else if b1=LOWER then LESS else GREATER) + else GREATER + +structure Inttab = Table(type key = int val ord = (rev_order o int_ord)); + +structure VarGraph = Table(type key = int*bound_type val ord = intbound_ord); +(* key -> (float option) * (int -> (float * (((float * float) * key) list)))) *) +(* dest_key -> (sure_bound * (row_index -> (row_bound * (((coeff_lower * coeff_upper) * src_key) list)))) *) + +exception Internal of string; + +fun add_row_bound g dest_key row_index row_bound = + let + val x = + case VarGraph.lookup g dest_key of + NONE => (NONE, Inttab.update (row_index, (row_bound, [])) Inttab.empty) + | SOME (sure_bound, f) => + (sure_bound, + case Inttab.lookup f row_index of + NONE => Inttab.update (row_index, (row_bound, [])) f + | SOME _ => raise (Internal "add_row_bound")) + in + VarGraph.update (dest_key, x) g + end + +fun update_sure_bound g (key as (_, btype)) bound = + let + val x = + case VarGraph.lookup g key of + NONE => (SOME bound, Inttab.empty) + | SOME (NONE, f) => (SOME bound, f) + | SOME (SOME old_bound, f) => + (SOME ((case btype of + UPPER => Float.min + | LOWER => Float.max) + old_bound bound), f) + in + VarGraph.update (key, x) g + end + +fun get_sure_bound g key = + case VarGraph.lookup g key of + NONE => NONE + | SOME (sure_bound, _) => sure_bound + +(*fun get_row_bound g key row_index = + case VarGraph.lookup g key of + NONE => NONE + | SOME (sure_bound, f) => + (case Inttab.lookup f row_index of + NONE => NONE + | SOME (row_bound, _) => (sure_bound, row_bound))*) + +fun add_edge g src_key dest_key row_index coeff = + case VarGraph.lookup g dest_key of + NONE => raise (Internal "add_edge: dest_key not found") + | SOME (sure_bound, f) => + (case Inttab.lookup f row_index of + NONE => raise (Internal "add_edge: row_index not found") + | SOME (row_bound, sources) => + VarGraph.update (dest_key, (sure_bound, Inttab.update (row_index, (row_bound, (coeff, src_key) :: sources)) f)) g) + +fun split_graph g = + let + fun split (key, (sure_bound, _)) (r1, r2) = case sure_bound + of NONE => (r1, r2) + | SOME bound => (case key + of (u, UPPER) => (r1, Inttab.update (u, bound) r2) + | (u, LOWER) => (Inttab.update (u, bound) r1, r2)) + in VarGraph.fold split g (Inttab.empty, Inttab.empty) end + +(* If safe is true, termination is guaranteed, but the sure bounds may be not optimal (relative to the algorithm). + If safe is false, termination is not guaranteed, but on termination the sure bounds are optimal (relative to the algorithm) *) +fun propagate_sure_bounds safe names g = + let + (* returns NONE if no new sure bound could be calculated, otherwise the new sure bound is returned *) + fun calc_sure_bound_from_sources g (key as (_, btype)) = + let + fun mult_upper x (lower, upper) = + if Float.sign x = LESS then + Float.mult x lower + else + Float.mult x upper + + fun mult_lower x (lower, upper) = + if Float.sign x = LESS then + Float.mult x upper + else + Float.mult x lower + + val mult_btype = case btype of UPPER => mult_upper | LOWER => mult_lower + + fun calc_sure_bound (_, (row_bound, sources)) sure_bound = + let + fun add_src_bound (coeff, src_key) sum = + case sum of + NONE => NONE + | SOME x => + (case get_sure_bound g src_key of + NONE => NONE + | SOME src_sure_bound => SOME (Float.add x (mult_btype src_sure_bound coeff))) + in + case fold add_src_bound sources (SOME row_bound) of + NONE => sure_bound + | new_sure_bound as (SOME new_bound) => + (case sure_bound of + NONE => new_sure_bound + | SOME old_bound => + SOME (case btype of + UPPER => Float.min old_bound new_bound + | LOWER => Float.max old_bound new_bound)) + end + in + case VarGraph.lookup g key of + NONE => NONE + | SOME (sure_bound, f) => + let + val x = Inttab.fold calc_sure_bound f sure_bound + in + if x = sure_bound then NONE else x + end + end + + fun propagate (key, _) (g, b) = + case calc_sure_bound_from_sources g key of + NONE => (g,b) + | SOME bound => (update_sure_bound g key bound, + if safe then + case get_sure_bound g key of + NONE => true + | _ => b + else + true) + + val (g, b) = VarGraph.fold propagate g (g, false) + in + if b then propagate_sure_bounds safe names g else g + end + +exception Load of string; + +val empty_spvec = @{term "Nil :: real spvec"}; +fun cons_spvec x xs = @{term "Cons :: nat * real => real spvec => real spvec"} $ x $ xs; +val empty_spmat = @{term "Nil :: real spmat"}; +fun cons_spmat x xs = @{term "Cons :: nat * real spvec => real spmat => real spmat"} $ x $ xs; + +fun calcr safe_propagation xlen names prec A b = + let + fun test_1 (lower, upper) = + if lower = upper then + (if Float.eq (lower, (~1, 0)) then ~1 + else if Float.eq (lower, (1, 0)) then 1 + else 0) + else 0 + + fun calcr (row_index, a) g = + let + val b = FloatSparseMatrixBuilder.v_elem_at b row_index + val (_, b2) = FloatArith.approx_decstr_by_bin prec (case b of NONE => "0" | SOME b => b) + val approx_a = FloatSparseMatrixBuilder.v_fold (fn (i, s) => fn l => + (i, FloatArith.approx_decstr_by_bin prec s)::l) a [] + + fun fold_dest_nodes (dest_index, dest_value) g = + let + val dest_test = test_1 dest_value + in + if dest_test = 0 then + g + else let + val (dest_key as (_, dest_btype), row_bound) = + if dest_test = ~1 then + ((dest_index, LOWER), Float.neg b2) + else + ((dest_index, UPPER), b2) + + fun fold_src_nodes (src_index, src_value as (src_lower, src_upper)) g = + if src_index = dest_index then g + else + let + val coeff = case dest_btype of + UPPER => (Float.neg src_upper, Float.neg src_lower) + | LOWER => src_value + in + if Float.sign src_lower = LESS then + add_edge g (src_index, UPPER) dest_key row_index coeff + else + add_edge g (src_index, LOWER) dest_key row_index coeff + end + in + fold fold_src_nodes approx_a (add_row_bound g dest_key row_index row_bound) + end + end + in + case approx_a of + [] => g + | [(u, a)] => + let + val atest = test_1 a + in + if atest = ~1 then + update_sure_bound g (u, LOWER) (Float.neg b2) + else if atest = 1 then + update_sure_bound g (u, UPPER) b2 + else + g + end + | _ => fold fold_dest_nodes approx_a g + end + + val g = FloatSparseMatrixBuilder.m_fold calcr A VarGraph.empty + + val g = propagate_sure_bounds safe_propagation names g + + val (r1, r2) = split_graph g + + fun add_row_entry m index f vname value = + let + val v = (case value of + SOME value => FloatSparseMatrixBuilder.mk_spvec_entry 0 value + | NONE => FloatSparseMatrixBuilder.mk_spvec_entry' 0 (f $ (Var ((vname,0), HOLogic.realT)))) + val vec = cons_spvec v empty_spvec + in + cons_spmat (FloatSparseMatrixBuilder.mk_spmat_entry index vec) m + end + + fun abs_estimate i r1 r2 = + if i = 0 then + let val e = empty_spmat in (e, e) end + else + let + val index = xlen-i + val (r12_1, r12_2) = abs_estimate (i-1) r1 r2 + val b1 = Inttab.lookup r1 index + val b2 = Inttab.lookup r2 index + in + (add_row_entry r12_1 index @{term "lbound :: real => real"} ((names index)^"l") b1, + add_row_entry r12_2 index @{term "ubound :: real => real"} ((names index)^"u") b2) + end + + val (r1, r2) = abs_estimate xlen r1 r2 + + in + (r1, r2) + end + +fun load filename prec safe_propagation = + let + val prog = Cplex.load_cplexFile filename + val prog = Cplex.elim_nonfree_bounds prog + val prog = Cplex.relax_strict_ineqs prog + val (maximize, c, A, b, (xlen, names, _)) = CplexFloatSparseMatrixConverter.convert_prog prog + val (r1, r2) = calcr safe_propagation xlen names prec A b + val _ = if maximize then () else raise Load "sorry, cannot handle minimization problems" + val (dualprog, indexof) = FloatSparseMatrixBuilder.dual_cplexProg c A b + val results = Cplex.solve dualprog + val (_, v) = CplexFloatSparseMatrixConverter.convert_results results indexof + (*val A = FloatSparseMatrixBuilder.cut_matrix v NONE A*) + fun id x = x + val v = FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 v + val b = FloatSparseMatrixBuilder.transpose_matrix (FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 b) + val c = FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 c + val (y1, _) = FloatSparseMatrixBuilder.approx_matrix prec Float.positive_part v + val A = FloatSparseMatrixBuilder.approx_matrix prec id A + val (_,b2) = FloatSparseMatrixBuilder.approx_matrix prec id b + val c = FloatSparseMatrixBuilder.approx_matrix prec id c + in + (y1, A, b2, c, (r1, r2)) + end handle CplexFloatSparseMatrixConverter.Converter s => (raise (Load ("Converter: "^s))) + +end diff -r 15ce93dfe6da -r 9f492f5b0cec src/HOL/Matrix_LP/matrixlp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Matrix_LP/matrixlp.ML Sat Mar 17 12:52:40 2012 +0100 @@ -0,0 +1,59 @@ +(* Title: HOL/Matrix/matrixlp.ML + Author: Steven Obua +*) + +signature MATRIX_LP = +sig + val matrix_compute : cterm -> thm + val matrix_simplify : thm -> thm + val prove_bound : string -> int -> thm + val float2real : string * string -> Real.real +end + +structure MatrixLP : MATRIX_LP = +struct + +val compute_thms = ComputeHOL.prep_thms @{thms "ComputeHOL.compute_list_case" "ComputeHOL.compute_let" + "ComputeHOL.compute_if" "ComputeFloat.arith" "SparseMatrix.sparse_row_matrix_arith_simps" + "ComputeHOL.compute_bool" "ComputeHOL.compute_pair" + "SparseMatrix.sorted_sp_simps" + "ComputeNumeral.natnorm"}; (*"ComputeNumeral.number_norm"*) + +val spm_mult_le_dual_prts_no_let_real = @{thm "spm_mult_le_dual_prts_no_let" [where ?'a = real]} + +fun lp_dual_estimate_prt lptfile prec = + let + val cert = cterm_of @{theory} + fun var s x = (cert (Var ((s, 0), FloatSparseMatrixBuilder.spmatT)), x) + val l = Fspmlp.load lptfile prec false + val (y, (A1, A2), (c1, c2), b, (r1, r2)) = + let + open Fspmlp + in + (y l |> cert, A l |> pairself cert, c l |> pairself cert, b l |> cert, r12 l |> pairself cert) + end + in + Thm.instantiate ([], + [var "A1" A1, var "A2" A2, var "y" y, var "c1" c1, var "c2" c2, var "r1" r1, var "r2" r2, var "b" b]) + spm_mult_le_dual_prts_no_let_real + end + +val computer = PCompute.make Compute.SML @{theory} compute_thms [] + +fun matrix_compute c = hd (PCompute.rewrite computer [c]) + +fun matrix_simplify th = + let + val simp_th = matrix_compute (cprop_of th) + val th = Thm.strip_shyps (Thm.equal_elim simp_th th) + fun removeTrue th = removeTrue (Thm.implies_elim th TrueI) handle THM _ => th + in + removeTrue th + end + +val prove_bound = matrix_simplify oo lp_dual_estimate_prt; + +val realFromStr = the o Real.fromString; +fun float2real (x, y) = realFromStr x * Math.pow (2.0, realFromStr y); + +end