neuper@37906: (* WN.020812: theorems in the Reals, neuper@37906: necessary for special rule sets, in addition to Isabelle2002. neuper@37906: !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!! neuper@37906: !!! THIS IS THE _least_ NUMBER OF ADDITIONAL THEOREMS !!! neuper@37906: !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!! walther@60242: xxxI contain \ instead of ^ in the respective theorem xxx in 2002 neuper@37906: changed by: Richard Lang 020912 neuper@37906: *) neuper@37906: neuper@37950: theory Poly imports Simplify begin neuper@37906: wneuper@59523: subsection \remark on term-structure of polynomials\ wneuper@59523: text \ wneuper@59523: WN190319: wneuper@59523: the code below reflects missing coordination between two authors: wneuper@59523: * ML: built the equation solver; simple rule-sets, programs; better predicates for specifications. wneuper@59523: * MG: built simplification of polynomials with AC rewriting by ML code wneuper@59523: wneuper@59523: WN020919: wneuper@59523: *** there are 5 kinds of expanded normalforms *** wneuper@59523: wneuper@59523: [1] 'complete polynomial' (Komplettes Polynom), univariate wneuper@59523: a_0 + a_1.x^1 +...+ a_n.x^n not (a_n = 0) wneuper@59523: not (a_n = 0), some a_i may be zero (DON'T disappear), wneuper@59523: variables in monomials lexicographically ordered and complete, wneuper@59523: x written as 1*x^1, ... wneuper@59523: [2] 'polynomial' (Polynom), univariate and multivariate wneuper@59523: a_0 + a_1.x +...+ a_n.x^n not (a_n = 0) wneuper@59523: a_0 + a_1.x_1.x_2^n_12...x_m^n_1m +...+ a_n.x_1^n.x_2^n_n2...x_m^n_nm wneuper@59523: not (a_n = 0), some a_i may be zero (ie. monomials disappear), wneuper@59523: exponents and coefficients equal 1 are not (WN060904.TODO in cancel_p_)shown, wneuper@59523: and variables in monomials are lexicographically ordered walther@60242: examples: [1]: "1 + (-10) * x \ 1 + 25 * x \ 2" walther@60242: [1]: "11 + 0 * x \ 1 + 1 * x \ 2" walther@60242: [2]: "x + (-50) * x \ 3" walther@60242: [2]: "(-1) * x * y \ 2 + 7 * x \ 3" wneuper@59523: wneuper@59523: [3] 'expanded_term' (Ausmultiplizierter Term): wneuper@59523: pull out unary minus to binary minus, wneuper@59523: as frequently exercised in schools; other conditions for [2] hold however walther@60242: examples: "a \ 2 - 2 * a * b + b \ 2" walther@60242: "4 * x \ 2 - 9 * y \ 2" wneuper@59523: [4] 'polynomial_in' (Polynom in): wneuper@59523: polynomial in 1 variable with arbitrary coefficients walther@60242: examples: "2 * x + (-50) * x \ 3" (poly in x) walther@60242: "(u + v) + (2 * u \ 2) * a + (-u) * a \ 2 (poly in a) wneuper@59523: [5] 'expanded_in' (Ausmultiplizierter Termin in): wneuper@59523: analoguous to [3] with binary minus like [3] walther@60242: examples: "2 * x - 50 * x \ 3" (expanded in x) walther@60242: "(u + v) + (2 * u \ 2) * a - u * a \ 2 (expanded in a) wneuper@59523: \ wneuper@59523: subsection \consts definition for predicates in specifications\ neuper@37906: consts neuper@37906: walther@60278: is_expanded_in :: "[real, real] => bool" ("_ is'_expanded'_in _") walther@60278: is_poly_in :: "[real, real] => bool" ("_ is'_poly'_in _") (*RL DA *) walther@60278: has_degree_in :: "[real, real] => real" ("_ has'_degree'_in _")(*RL DA *) walther@60278: is_polyrat_in :: "[real, real] => bool" ("_ is'_polyrat'_in _")(*RL030626*) neuper@37906: walther@60278: is_multUnordered:: "real => bool" ("_ is'_multUnordered") walther@60278: is_addUnordered :: "real => bool" ("_ is'_addUnordered") (*WN030618*) walther@60278: is_polyexp :: "real => bool" ("_ is'_polyexp") neuper@37906: wneuper@59523: subsection \theorems not yet adopted from Isabelle\ neuper@52148: axiomatization where (*.not contained in Isabelle2002, neuper@37906: stated as axioms, TODO: prove as theorems; walther@60242: theorem-IDs 'xxxI' with \ instead of ^ in 'xxx' in Isabelle2002.*) neuper@37906: walther@60242: realpow_pow: "(a \ b) \ c = a \ (b * c)" and walther@60242: realpow_addI: "r \ (n + m) = r \ n * r \ m" and walther@60242: realpow_addI_assoc_l: "r \ n * (r \ m * s) = r \ (n + m) * s" and walther@60242: realpow_addI_assoc_r: "s * r \ n * r \ m = s * r \ (n + m)" and neuper@37906: walther@60242: realpow_oneI: "r \ 1 = r" and walther@60242: realpow_zeroI: "r \ 0 = 1" and walther@60242: realpow_eq_oneI: "1 \ n = 1" and walther@60242: realpow_multI: "(r * s) \ n = r \ n * s \ n" and neuper@37974: realpow_multI_poly: "[| r is_polyexp; s is_polyexp |] ==> walther@60242: (r * s) \ n = r \ n * s \ n" and walther@60242: realpow_minus_oneI: "(- 1) \ (2 * n) = 1" and wneuper@59589: real_diff_0: "0 - x = - (x::real)" and neuper@37906: walther@60242: realpow_twoI: "r \ 2 = r * r" and walther@60242: realpow_twoI_assoc_l: "r * (r * s) = r \ 2 * s" and walther@60242: realpow_twoI_assoc_r: "s * r * r = s * r \ 2" and walther@60242: realpow_two_atom: "r is_atom ==> r * r = r \ 2" and walther@60242: realpow_plus_1: "r * r \ n = r \ (n + 1)" and walther@60242: realpow_plus_1_assoc_l: "r * (r \ m * s) = r \ (1 + m) * s" and walther@60242: realpow_plus_1_assoc_l2: "r \ m * (r * s) = r \ (1 + m) * s" and walther@60242: realpow_plus_1_assoc_r: "s * r * r \ m = s * r \ (1 + m)" and walther@60242: realpow_plus_1_atom: "r is_atom ==> r * r \ n = r \ (1 + n)" and neuper@37974: realpow_def_atom: "[| Not (r is_atom); 1 < n |] walther@60242: ==> r \ n = r * r \ (n + -1)" and walther@60242: realpow_addI_atom: "r is_atom ==> r \ n * r \ m = r \ (n + m)" and neuper@37906: neuper@37906: walther@60278: realpow_minus_even: "n is_even ==> (- r) \ n = r \ n" and walther@60242: realpow_minus_odd: "Not (n is_even) ==> (- r) \ n = -1 * r \ n" and neuper@37906: neuper@37906: neuper@37906: (* RL 020914 *) neuper@52148: real_pp_binom_times: "(a + b)*(c + d) = a*c + a*d + b*c + b*d" and neuper@52148: real_pm_binom_times: "(a + b)*(c - d) = a*c - a*d + b*c - b*d" and neuper@52148: real_mp_binom_times: "(a - b)*(c + d) = a*c + a*d - b*c - b*d" and neuper@52148: real_mm_binom_times: "(a - b)*(c - d) = a*c - a*d - b*c + b*d" and walther@60242: real_plus_binom_pow3: "(a + b) \ 3 = a \ 3 + 3*a \ 2*b + 3*a*b \ 2 + b \ 3" and neuper@37974: real_plus_binom_pow3_poly: "[| a is_polyexp; b is_polyexp |] ==> walther@60242: (a + b) \ 3 = a \ 3 + 3*a \ 2*b + 3*a*b \ 2 + b \ 3" and walther@60242: real_minus_binom_pow3: "(a - b) \ 3 = a \ 3 - 3*a \ 2*b + 3*a*b \ 2 - b \ 3" and walther@60242: real_minus_binom_pow3_p: "(a + -1 * b) \ 3 = a \ 3 + -3*a \ 2*b + 3*a*b \ 2 + walther@60242: -1*b \ 3" and walther@60385: (* real_plus_binom_pow: "[| n is_num; 3 < n |] ==> walther@60260: (a + b) \ n = (a + b) * (a + b)\(n - 1)" *) walther@60242: real_plus_binom_pow4: "(a + b) \ 4 = (a \ 3 + 3*a \ 2*b + 3*a*b \ 2 + b \ 3) neuper@52148: *(a + b)" and neuper@37974: real_plus_binom_pow4_poly: "[| a is_polyexp; b is_polyexp |] ==> walther@60242: (a + b) \ 4 = (a \ 3 + 3*a \ 2*b + 3*a*b \ 2 + b \ 3) neuper@52148: *(a + b)" and walther@60242: real_plus_binom_pow5: "(a + b) \ 5 = (a \ 3 + 3*a \ 2*b + 3*a*b \ 2 + b \ 3) walther@60242: *(a \ 2 + 2*a*b + b \ 2)" and neuper@37974: real_plus_binom_pow5_poly: "[| a is_polyexp; b is_polyexp |] ==> walther@60242: (a + b) \ 5 = (a \ 3 + 3*a \ 2*b + 3*a*b \ 2 walther@60242: + b \ 3)*(a \ 2 + 2*a*b + b \ 2)" and neuper@52148: real_diff_plus: "a - b = a + -b" (*17.3.03: do_NOT_use*) and neuper@52148: real_diff_minus: "a - b = a + -1 * b" and walther@60242: real_plus_binom_times: "(a + b)*(a + b) = a \ 2 + 2*a*b + b \ 2" and walther@60242: real_minus_binom_times: "(a - b)*(a - b) = a \ 2 - 2*a*b + b \ 2" and neuper@37906: (*WN071229 changed for Schaerding -----vvv*) walther@60260: (*real_plus_binom_pow2: "(a + b) \ 2 = a \ 2 + 2*a*b + b \ 2"*) walther@60242: real_plus_binom_pow2: "(a + b) \ 2 = (a + b) * (a + b)" and walther@60242: (*WN071229 changed for Schaerding -----\*) neuper@37974: real_plus_binom_pow2_poly: "[| a is_polyexp; b is_polyexp |] ==> walther@60242: (a + b) \ 2 = a \ 2 + 2*a*b + b \ 2" and walther@60242: real_minus_binom_pow2: "(a - b) \ 2 = a \ 2 - 2*a*b + b \ 2" and walther@60242: real_minus_binom_pow2_p: "(a - b) \ 2 = a \ 2 + -2*a*b + b \ 2" and walther@60242: real_plus_minus_binom1: "(a + b)*(a - b) = a \ 2 - b \ 2" and walther@60242: real_plus_minus_binom1_p: "(a + b)*(a - b) = a \ 2 + -1*b \ 2" and walther@60242: real_plus_minus_binom1_p_p: "(a + b)*(a + -1 * b) = a \ 2 + -1*b \ 2" and walther@60242: real_plus_minus_binom2: "(a - b)*(a + b) = a \ 2 - b \ 2" and walther@60242: real_plus_minus_binom2_p: "(a - b)*(a + b) = a \ 2 + -1*b \ 2" and walther@60242: real_plus_minus_binom2_p_p: "(a + -1 * b)*(a + b) = a \ 2 + -1*b \ 2" and walther@60242: real_plus_binom_times1: "(a + 1*b)*(a + -1*b) = a \ 2 + -1*b \ 2" and walther@60242: real_plus_binom_times2: "(a + -1*b)*(a + 1*b) = a \ 2 + -1*b \ 2" and neuper@37906: walther@60385: real_num_collect: "[| l is_num; m is_num |] ==> neuper@52148: l * n + m * n = (l + m) * n" and neuper@37906: (* FIXME.MG.0401: replace 'real_num_collect_assoc' neuper@37906: by 'real_num_collect_assoc_l' ... are equal, introduced by MG ! *) walther@60385: real_num_collect_assoc: "[| l is_num; m is_num |] ==> neuper@52148: l * n + (m * n + k) = (l + m) * n + k" and walther@60385: real_num_collect_assoc_l: "[| l is_num; m is_num |] ==> neuper@37950: l * n + (m * n + k) = (l + m) neuper@52148: * n + k" and walther@60385: real_num_collect_assoc_r: "[| l is_num; m is_num |] ==> neuper@52148: (k + m * n) + l * n = k + (l + m) * n" and walther@60385: real_one_collect: "m is_num ==> n + m * n = (1 + m) * n" and neuper@37906: (* FIXME.MG.0401: replace 'real_one_collect_assoc' neuper@37906: by 'real_one_collect_assoc_l' ... are equal, introduced by MG ! *) walther@60385: real_one_collect_assoc: "m is_num ==> n + (m * n + k) = (1 + m)* n + k" and neuper@37906: walther@60385: real_one_collect_assoc_l: "m is_num ==> n + (m * n + k) = (1 + m) * n + k" and walther@60385: real_one_collect_assoc_r: "m is_num ==> (k + n) + m * n = k + (1 + m) * n" and neuper@37906: neuper@37906: (* FIXME.MG.0401: replace 'real_mult_2_assoc' neuper@37906: by 'real_mult_2_assoc_l' ... are equal, introduced by MG ! *) neuper@52148: real_mult_2_assoc: "z1 + (z1 + k) = 2 * z1 + k" and neuper@52148: real_mult_2_assoc_l: "z1 + (z1 + k) = 2 * z1 + k" and neuper@52148: real_mult_2_assoc_r: "(k + z1) + z1 = k + 2 * z1" and neuper@37906: wneuper@59587: real_mult_left_commute: "z1 * (z2 * z3) = z2 * (z1 * z3)" and wneuper@59587: real_mult_minus1: "-1 * z = - (z::real)" and walther@60333: (*sym_real_mult_minus1 expands indefinitely without assumptions ...*) walther@60343: real_mult_minus1_sym: "[| \(matches (- 1 * x) z); \(z is_num) |] ==> - (z::real) = -1 * z" and walther@60344: real_minus_mult_left: "\ ((a::real) is_num) ==> (- a) * b = - (a * b)" and wneuper@59587: real_mult_2: "2 * z = z + (z::real)" and walther@60343: neuper@52148: real_add_mult_distrib_poly: "w is_polyexp ==> (z1 + z2) * w = z1 * w + z2 * w" and neuper@37974: real_add_mult_distrib2_poly:"w is_polyexp ==> w * (z1 + z2) = w * z1 + w * z2" neuper@37950: walther@60278: wneuper@59523: subsection \auxiliary functions\ wneuper@59530: ML \ wneuper@59530: val poly_consts = wenzelm@60309: [\<^const_name>\plus\, \<^const_name>\minus\, wenzelm@60309: \<^const_name>\divide\, \<^const_name>\times\, wenzelm@60405: \<^const_name>\realpow\]; walther@60321: walther@60321: val int_ord_SAVE = int_ord; walther@60321: (*for tests on rewrite orders*) walther@60321: fun int_ord (i1, i2) = walther@60321: (@{print} {a = "int_ord (" ^ string_of_int i1 ^ ", " ^ string_of_int i2 ^ ") = ", z = Int.compare (i1, i2)}; walther@60321: Int.compare (i1, i2)); walther@60321: (**)val int_ord = int_ord_SAVE; (*..outcomment for tests*) wneuper@59530: \ wneuper@59523: subsubsection \for predicates in specifications (ML)\ wneuper@59472: ML \ wneuper@59522: (*--- auxiliary for is_expanded_in, is_poly_in, has_degree_in ---*) walther@60317: (*. a "monomial t in variable v" is a term t with wneuper@59522: either (1) v NOT existent in t, or (2) v contained in t, wneuper@59522: if (1) then degree 0 walther@60317: if (2) then v is a factor on the very right, casually with exponent.*) wneuper@59522: fun factor_right_deg (*case 2*) walther@60335: (Const (\<^const_name>\Groups.times_class.times\, _) $ wenzelm@60405: t1 $ (Const (\<^const_name>\realpow\,_) $ vv $ num)) v = walther@60317: if vv = v andalso not (Prog_Expr.occurs_in v t1) then SOME (snd (HOLogic.dest_number num)) walther@60317: else NONE wenzelm@60405: | factor_right_deg (Const (\<^const_name>\realpow\,_) $ vv $ num) v = walther@60317: if (vv = v) then SOME (snd (HOLogic.dest_number num)) else NONE walther@60331: wenzelm@60309: | factor_right_deg (Const (\<^const_name>\times\,_) $ t1 $ vv) v = walther@59603: if vv = v andalso not (Prog_Expr.occurs_in v t1) then SOME 1 else NONE wneuper@59522: | factor_right_deg vv v = wneuper@59522: if (vv = v) then SOME 1 else NONE; wneuper@59522: fun mono_deg_in m v = (*case 1*) walther@59603: if not (Prog_Expr.occurs_in v m) then (*case 1*) SOME 0 else factor_right_deg m v; wneuper@59522: wneuper@59522: fun expand_deg_in t v = wneuper@59522: let wenzelm@60309: fun edi ~1 ~1 (Const (\<^const_name>\plus\, _) $ t1 $ t2) = wneuper@59522: (case mono_deg_in t2 v of (* $ is left associative*) wneuper@59522: SOME d' => edi d' d' t1 | NONE => NONE) wenzelm@60309: | edi ~1 ~1 (Const (\<^const_name>\minus\, _) $ t1 $ t2) = wneuper@59522: (case mono_deg_in t2 v of wneuper@59522: SOME d' => edi d' d' t1 | NONE => NONE) wenzelm@60309: | edi d dmax (Const (\<^const_name>\minus\, _) $ t1 $ t2) = wneuper@59522: (case mono_deg_in t2 v of (*(d = 0 andalso d' = 0) handle 3+4-...4 +x*) wneuper@59522: SOME d' => if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE wneuper@59522: | NONE => NONE) wenzelm@60309: | edi d dmax (Const (\<^const_name>\plus\,_) $ t1 $ t2) = wneuper@59522: (case mono_deg_in t2 v of wneuper@59522: SOME d' => (*RL (d = 0 andalso d' = 0) need to handle 3+4-...4 +x*) wneuper@59522: if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE wneuper@59522: | NONE => NONE) wneuper@59522: | edi ~1 ~1 t = wneuper@59522: (case mono_deg_in t v of d as SOME _ => d | NONE => NONE) wneuper@59522: | edi d dmax t = (*basecase last*) wneuper@59522: (case mono_deg_in t v of wneuper@59522: SOME d' => if d > d' orelse (d = 0 andalso d' = 0) then SOME dmax else NONE wneuper@59522: | NONE => NONE) wneuper@59522: in edi ~1 ~1 t end; wneuper@59522: wneuper@59522: fun poly_deg_in t v = wneuper@59522: let wenzelm@60309: fun edi ~1 ~1 (Const (\<^const_name>\plus\,_) $ t1 $ t2) = wneuper@59522: (case mono_deg_in t2 v of (* $ is left associative *) wneuper@59522: SOME d' => edi d' d' t1 wneuper@59522: | NONE => NONE) wenzelm@60309: | edi d dmax (Const (\<^const_name>\plus\,_) $ t1 $ t2) = wneuper@59522: (case mono_deg_in t2 v of wneuper@59522: SOME d' => (*RL (d = 0 andalso (d' = 0)) handle 3+4-...4 +x*) wneuper@59522: if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE wneuper@59522: | NONE => NONE) wneuper@59522: | edi ~1 ~1 t = wneuper@59522: (case mono_deg_in t v of wneuper@59522: d as SOME _ => d wneuper@59522: | NONE => NONE) wneuper@59522: | edi d dmax t = (* basecase last *) wneuper@59522: (case mono_deg_in t v of wneuper@59522: SOME d' => wneuper@59522: if d > d' orelse (d = 0 andalso d' = 0) then SOME dmax else NONE wneuper@59522: | NONE => NONE) wneuper@59522: in edi ~1 ~1 t end; wneuper@59523: \ neuper@37950: wneuper@59523: subsubsection \for hard-coded AC rewriting (MG)\ wneuper@59523: ML \ wneuper@59523: (**. MG.03: make_polynomial_ ... uses SML-fun for ordering .**) neuper@37950: wneuper@59523: (*FIXME.0401: make SML-order local to make_polynomial(_) *) wneuper@59523: (*FIXME.0401: replace 'make_polynomial'(old) by 'make_polynomial_'(MG) *) wneuper@59523: (* Polynom --> List von Monomen *) wenzelm@60309: fun poly2list (Const (\<^const_name>\plus\,_) $ t1 $ t2) = wneuper@59523: (poly2list t1) @ (poly2list t2) wneuper@59523: | poly2list t = [t]; neuper@37950: wneuper@59523: (* Monom --> Liste von Variablen *) wenzelm@60309: fun monom2list (Const (\<^const_name>\times\,_) $ t1 $ t2) = wneuper@59523: (monom2list t1) @ (monom2list t2) wneuper@59523: | monom2list t = [t]; neuper@37950: wneuper@59523: (* liefert Variablenname (String) einer Variablen und Basis bei Potenz *) wenzelm@60405: fun get_basStr (Const (\<^const_name>\realpow\,_) $ Free (str, _) $ _) = str wenzelm@60405: | get_basStr (Const (\<^const_name>\realpow\,_) $ n $ _) = TermC.to_string n wneuper@59523: | get_basStr (Free (str, _)) = str walther@60318: | get_basStr t = walther@60322: if TermC.is_num t then TermC.to_string t walther@60318: else "|||"; (* gross gewichtet; für Brüche ect. *) neuper@37950: wneuper@59523: (* liefert Hochzahl (String) einer Variablen bzw Gewichtstring (zum Sortieren) *) wenzelm@60405: fun get_potStr (Const (\<^const_name>\realpow\, _) $ Free _ $ Free (str, _)) = str wenzelm@60405: | get_potStr (Const (\<^const_name>\realpow\, _) $ Free _ $ t) = walther@60322: if TermC.is_num t then TermC.to_string t else "|||" walther@60321: | get_potStr (Free _) = "---" (* keine Hochzahl --> kleinst gewichtet *) wneuper@59523: | get_potStr _ = "||||||"; (* gross gewichtet; für Brüch ect. *) neuper@37978: wneuper@59523: (* Umgekehrte string_ord *) wneuper@59523: val string_ord_rev = rev_order o string_ord; wneuper@59523: wneuper@59523: (* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen) wneuper@59523: innerhalb eines Monomes: wneuper@59523: - zuerst lexikographisch nach Variablenname wneuper@59523: - wenn gleich: nach steigender Potenz *) walther@60321: fun var_ord (a, b) = walther@60321: (@{print} {a = "var_ord ", a_b = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")", walther@60321: sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"}; walther@60321: prod_ord string_ord string_ord walther@60321: ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b)) walther@60321: ); walther@60318: fun var_ord (a,b: term) = walther@60318: prod_ord string_ord string_ord wneuper@59523: ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b)); neuper@37950: wneuper@59523: (* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen); wneuper@59523: verwendet zum Sortieren von Monomen mittels Gesamtgradordnung: wneuper@59523: - zuerst lexikographisch nach Variablenname wneuper@59523: - wenn gleich: nach sinkender Potenz*) walther@60321: fun var_ord_revPow (a, b: term) = walther@60321: (@{print} {a = "var_ord_revPow ", at_bt = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")", walther@60321: sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"}; walther@60321: prod_ord string_ord string_ord_rev walther@60321: ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b)) walther@60321: ); walther@60318: fun var_ord_revPow (a, b: term) = walther@60318: prod_ord string_ord string_ord_rev wneuper@59523: ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b)); neuper@37950: walther@60278: wneuper@59523: (* Ordnet ein Liste von Variablen (und Potenzen) lexikographisch *) walther@60321: fun sort_varList ts = walther@60321: (@{print} {a = "sort_varList", args = UnparseC.terms ts}; walther@60321: sort var_ord ts); wneuper@59523: val sort_varList = sort var_ord; wneuper@59523: wneuper@59523: (* Entfernet aeussersten Operator (Wurzel) aus einem Term und schreibt wneuper@59523: Argumente in eine Liste *) wneuper@59523: fun args u : term list = walther@60318: let walther@60318: fun stripc (f $ t, ts) = stripc (f, t::ts) walther@60318: | stripc (t as Free _, ts) = (t::ts) walther@60318: | stripc (_, ts) = ts walther@60318: in stripc (u, []) end; wneuper@59523: wneuper@59523: (* liefert True, falls der Term (Liste von Termen) nur Zahlen wneuper@59523: (keine Variablen) enthaelt *) walther@60317: fun filter_num ts = fold (curry and_) (map TermC.is_num ts) true wneuper@59523: wneuper@59523: (* liefert True, falls der Term nur Zahlen (keine Variablen) enthaelt wneuper@59523: dh. er ist ein numerischer Wert und entspricht einem Koeffizienten *) wneuper@59523: fun is_nums t = filter_num [t]; wneuper@59523: wneuper@59523: (* Berechnet den Gesamtgrad eines Monoms *) walther@60318: (**)local(**) walther@60317: fun counter (n, []) = n walther@60317: | counter (n, x :: xs) = walther@60317: if (is_nums x) then counter (n, xs) walther@60317: else walther@60317: (case x of wenzelm@60405: (Const (\<^const_name>\realpow\, _) $ Free _ $ t) => walther@60318: if TermC.is_num t walther@60318: then counter (t |> HOLogic.dest_number |> snd |> curry op + n, xs) walther@60318: else counter (n + 1000, xs) (*FIXME.MG?!*) walther@60335: | (Const (\<^const_name>\numeral\, _) $ num) => walther@60318: counter (n + 1 + HOLogic.dest_numeral num, xs) walther@60318: | _ => counter (n + 1, xs)) (*FIXME.MG?! ... Brüche ect.*) walther@60318: (**)in(**) walther@60317: fun monom_degree l = counter (0, l) walther@60318: (**)end;(*local*) wneuper@59523: wneuper@59523: (* wie Ordnung dict_ord (lexicographische Ordnung zweier Listen, mit Vergleich wneuper@59523: der Listen-Elemente mit elem_ord) - Elemente die Bedingung cond erfuellen, wneuper@59523: werden jedoch dabei ignoriert (uebersprungen) *) walther@60321: fun dict_cond_ord _ _ ([], []) = (@{print} {a = "dict_cond_ord ([], [])"}; EQUAL) walther@60321: | dict_cond_ord _ _ ([], _ :: _) = (@{print} {a = "dict_cond_ord ([], _ :: _)"}; LESS) walther@60321: | dict_cond_ord _ _ (_ :: _, []) = (@{print} {a = "dict_cond_ord (_ :: _, [])"}; GREATER) walther@60321: | dict_cond_ord elem_ord cond (x :: xs, y :: ys) = walther@60321: (@{print} {a = "dict_cond_ord", args = "(" ^ UnparseC.terms (x :: xs) ^ ", " ^ UnparseC.terms (y :: ys) ^ ")", walther@60321: is_nums = "(" ^ LibraryC.bool2str (cond x) ^ ", " ^ LibraryC.bool2str (cond y) ^ ")"}; walther@60321: case (cond x, cond y) of walther@60321: (false, false) => walther@60321: (case elem_ord (x, y) of walther@60321: EQUAL => dict_cond_ord elem_ord cond (xs, ys) walther@60321: | ord => ord) walther@60321: | (false, true) => dict_cond_ord elem_ord cond (x :: xs, ys) walther@60321: | (true, false) => dict_cond_ord elem_ord cond (xs, y :: ys) walther@60321: | (true, true) => dict_cond_ord elem_ord cond (xs, ys) ); wneuper@59523: fun dict_cond_ord _ _ ([], []) = EQUAL wneuper@59523: | dict_cond_ord _ _ ([], _ :: _) = LESS wneuper@59523: | dict_cond_ord _ _ (_ :: _, []) = GREATER wneuper@59523: | dict_cond_ord elem_ord cond (x :: xs, y :: ys) = wneuper@59523: (case (cond x, cond y) of walther@60318: (false, false) => walther@60318: (case elem_ord (x, y) of walther@60318: EQUAL => dict_cond_ord elem_ord cond (xs, ys) walther@60318: | ord => ord) walther@60318: | (false, true) => dict_cond_ord elem_ord cond (x :: xs, ys) walther@60318: | (true, false) => dict_cond_ord elem_ord cond (xs, y :: ys) walther@60318: | (true, true) => dict_cond_ord elem_ord cond (xs, ys) ); wneuper@59523: wneuper@59523: (* Gesamtgradordnung zum Vergleich von Monomen (Liste von Variablen/Potenzen): wneuper@59523: zuerst nach Gesamtgrad, bei gleichem Gesamtgrad lexikographisch ordnen - walther@60260: dabei werden Koeffizienten ignoriert (2*3*a \ 2*4*b gilt wie a \ 2*b) *) wneuper@59523: fun degree_ord (xs, ys) = walther@60318: prod_ord int_ord (dict_cond_ord var_ord_revPow is_nums) walther@60318: ((monom_degree xs, xs), (monom_degree ys, ys)); wneuper@59523: wneuper@59523: fun hd_str str = substring (str, 0, 1); wneuper@59523: fun tl_str str = substring (str, 1, (size str) - 1); wneuper@59523: wneuper@59523: (* liefert nummerischen Koeffizienten eines Monoms oder NONE *) walther@60318: fun get_koeff_of_mon [] = raise ERROR "get_koeff_of_mon: called with l = []" walther@60318: | get_koeff_of_mon (x :: _) = if is_nums x then SOME x else NONE; wneuper@59523: wneuper@59523: (* wandelt Koeffizient in (zum sortieren geeigneten) String um *) walther@60318: fun koeff2ordStr (SOME t) = walther@60318: if TermC.is_num t walther@60318: then walther@60318: if (t |> HOLogic.dest_number |> snd) < 0 walther@60318: then (t |> HOLogic.dest_number |> snd |> curry op * ~1 |> string_of_int) ^ "0" (* 3 < -3 *) walther@60318: else (t |> HOLogic.dest_number |> snd |> string_of_int) walther@60318: else "aaa" (* "num.Ausdruck" --> gross *) walther@60318: | koeff2ordStr NONE = "---"; (* "kein Koeff" --> kleinste *) wneuper@59523: wneuper@59523: (* Order zum Vergleich von Koeffizienten (strings): wneuper@59523: "kein Koeff" < "0" < "1" < "-1" < "2" < "-2" < ... < "num.Ausdruck" *) walther@60318: fun compare_koeff_ord (xs, ys) = string_ord walther@60318: ((koeff2ordStr o get_koeff_of_mon) xs, walther@60318: (koeff2ordStr o get_koeff_of_mon) ys); wneuper@59523: wneuper@59523: (* Gesamtgradordnung degree_ord + Ordnen nach Koeffizienten falls EQUAL *) wneuper@59523: fun koeff_degree_ord (xs, ys) = wneuper@59523: prod_ord degree_ord compare_koeff_ord ((xs, xs), (ys, ys)); wneuper@59523: wneuper@59523: (* Ordnet ein Liste von Monomen (Monom = Liste von Variablen) mittels wneuper@59523: Gesamtgradordnung *) wneuper@59523: val sort_monList = sort koeff_degree_ord; wneuper@59523: wneuper@59523: (* Alternativ zu degree_ord koennte auch die viel einfachere und wneuper@59523: kuerzere Ordnung simple_ord verwendet werden - ist aber nicht wneuper@59523: fuer unsere Zwecke geeignet! wneuper@59523: wneuper@59523: fun simple_ord (al,bl: term list) = dict_ord string_ord wneuper@59523: (map get_basStr al, map get_basStr bl); wneuper@59523: wneuper@59523: val sort_monList = sort simple_ord; *) wneuper@59523: wneuper@59523: (* aus 2 Variablen wird eine Summe bzw ein Produkt erzeugt wneuper@59523: (mit gewuenschtem Typen T) *) wenzelm@60309: fun plus T = Const (\<^const_name>\plus\, [T,T] ---> T); wenzelm@60309: fun mult T = Const (\<^const_name>\times\, [T,T] ---> T); wneuper@59523: fun binop op_ t1 t2 = op_ $ t1 $ t2; wneuper@59523: fun create_prod T (a,b) = binop (mult T) a b; wneuper@59523: fun create_sum T (a,b) = binop (plus T) a b; wneuper@59523: wneuper@59523: (* löscht letztes Element einer Liste *) wneuper@59523: fun drop_last l = take ((length l)-1,l); wneuper@59523: wneuper@59523: (* Liste von Variablen --> Monom *) wneuper@59523: fun create_monom T vl = foldr (create_prod T) (drop_last vl, last_elem vl); wneuper@59523: (* Bemerkung: wneuper@59523: foldr bewirkt rechtslastige Klammerung des Monoms - ist notwendig, damit zwei wneuper@59523: gleiche Monome zusammengefasst werden können (collect_numerals)! wneuper@59523: zB: 2*(x*(y*z)) + 3*(x*(y*z)) --> (2+3)*(x*(y*z))*) wneuper@59523: wneuper@59523: (* Liste von Monomen --> Polynom *) wneuper@59523: fun create_polynom T ml = foldl (create_sum T) (hd ml, tl ml); wneuper@59523: (* Bemerkung: wneuper@59523: foldl bewirkt linkslastige Klammerung des Polynoms (der Summanten) - wneuper@59523: bessere Darstellung, da keine Klammern sichtbar! wneuper@59523: (und discard_parentheses in make_polynomial hat weniger zu tun) *) wneuper@59523: wneuper@59523: (* sorts the variables (faktors) of an expanded polynomial lexicographical *) wneuper@59523: fun sort_variables t = walther@60317: let walther@60317: val ll = map monom2list (poly2list t); walther@60317: val lls = map sort_varList ll; walther@60317: val T = type_of t; walther@60317: val ls = map (create_monom T) lls; walther@60317: in create_polynom T ls end; wneuper@59523: wneuper@59523: (* sorts the monoms of an expanded and variable-sorted polynomial wneuper@59523: by total_degree *) wneuper@59523: fun sort_monoms t = walther@60318: let walther@60318: val ll = map monom2list (poly2list t); walther@60318: val lls = sort_monList ll; walther@60318: val T = Term.type_of t; walther@60318: val ls = map (create_monom T) lls; walther@60318: in create_polynom T ls end; wneuper@59523: \ wneuper@59523: wneuper@59523: subsubsection \rewrite order for hard-coded AC rewriting\ wneuper@59523: ML \ neuper@37950: local (*. for make_polynomial .*) neuper@37950: neuper@37950: open Term; (* for type order = EQUAL | LESS | GREATER *) neuper@37950: neuper@37950: fun pr_ord EQUAL = "EQUAL" neuper@37950: | pr_ord LESS = "LESS" neuper@37950: | pr_ord GREATER = "GREATER"; neuper@37950: neuper@37950: fun dest_hd' (Const (a, T)) = (* ~ term.ML *) neuper@37950: (case a of wenzelm@60405: \<^const_name>\realpow\ => ((("|||||||||||||", 0), T), 0) (*WN greatest string*) neuper@37950: | _ => (((a, 0), T), 0)) walther@60317: | dest_hd' (Free (a, T)) = (((a, 0), T), 1)(*TODOO handle this as numeral, too? see EqSystem.thy*) neuper@37950: | dest_hd' (Var v) = (v, 2) neuper@37950: | dest_hd' (Bound i) = ((("", i), dummyT), 3) wneuper@59523: | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4) wneuper@59523: | dest_hd' t = raise TERM ("dest_hd'", [t]); neuper@37950: neuper@37950: fun size_of_term' (Const(str,_) $ t) = wenzelm@60405: if \<^const_name>\realpow\= str then 1000 + size_of_term' t else 1+size_of_term' t(*WN*) neuper@37950: | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body neuper@37950: | size_of_term' (f$t) = size_of_term' f + size_of_term' t neuper@37950: | size_of_term' _ = 1; neuper@37950: neuper@37950: fun term_ord' pr thy (Abs (_, T, t), Abs(_, U, u)) = (* ~ term.ML *) neuper@52070: (case term_ord' pr thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord) neuper@37950: | term_ord' pr thy (t, u) = neuper@52070: (if pr then neuper@52070: let neuper@52070: val (f, ts) = strip_comb t and (g, us) = strip_comb u; walther@59870: val _ = tracing ("t= f@ts= \"" ^ UnparseC.term_in_thy thy f ^ "\" @ \"[" ^ walther@59870: commas (map (UnparseC.term_in_thy thy) ts) ^ "]\""); walther@59870: val _ = tracing("u= g@us= \"" ^ UnparseC.term_in_thy thy g ^ "\" @ \"[" ^ walther@59870: commas (map (UnparseC.term_in_thy thy) us) ^ "]\""); neuper@52070: val _ = tracing ("size_of_term(t,u)= (" ^ string_of_int (size_of_term' t) ^ ", " ^ neuper@52070: string_of_int (size_of_term' u) ^ ")"); neuper@52070: val _ = tracing ("hd_ord(f,g) = " ^ (pr_ord o hd_ord) (f,g)); neuper@52070: val _ = tracing ("terms_ord(ts,us) = " ^ (pr_ord o terms_ord str false) (ts, us)); neuper@52070: val _ = tracing ("-------"); neuper@52070: in () end neuper@37950: else (); neuper@37950: case int_ord (size_of_term' t, size_of_term' u) of neuper@37950: EQUAL => neuper@37950: let val (f, ts) = strip_comb t and (g, us) = strip_comb u in neuper@37950: (case hd_ord (f, g) of EQUAL => (terms_ord str pr) (ts, us) neuper@37950: | ord => ord) neuper@37950: end neuper@37950: | ord => ord) neuper@37950: and hd_ord (f, g) = (* ~ term.ML *) neuper@37974: prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord (dest_hd' f, dest_hd' g) wneuper@59523: and terms_ord _ pr (ts, us) = walther@59881: list_ord (term_ord' pr (ThyC.get_theory "Isac_Knowledge"))(ts, us); neuper@52070: neuper@37950: in neuper@37950: walther@60324: fun ord_make_polynomial (pr:bool) thy (_: subst) (ts, us) = walther@60324: (term_ord' pr thy(***) (TermC.numerals_to_Free ts, TermC.numerals_to_Free us) = LESS ); neuper@37950: neuper@37950: end;(*local*) Walther@60506: \ Walther@60506: setup \KEStore_Elems.add_rew_ord [ Walther@60506: ("termlessI", termlessI), Walther@60506: ("ord_make_polynomial", ord_make_polynomial false \<^theory>)]\ neuper@37950: wneuper@59523: subsection \predicates\ wneuper@59523: subsubsection \in specifications\ wneuper@59523: ML \ wneuper@59523: (* is_polyrat_in becomes true, if no bdv is in the denominator of a fraction*) wneuper@59523: fun is_polyrat_in t v = wneuper@59524: let walther@59962: fun finddivide (_ $ _ $ _ $ _) _ = raise ERROR("is_polyrat_in:") wneuper@59523: (* at the moment there is no term like this, but ....*) wenzelm@60309: | finddivide (Const (\<^const_name>\divide\,_) $ _ $ b) v = not (Prog_Expr.occurs_in v b) wneuper@59524: | finddivide (_ $ t1 $ t2) v = finddivide t1 v orelse finddivide t2 v wneuper@59524: | finddivide (_ $ t1) v = finddivide t1 v wneuper@59523: | finddivide _ _ = false; wneuper@59524: in finddivide t v end; wneuper@59523: wneuper@59524: fun is_expanded_in t v = case expand_deg_in t v of SOME _ => true | NONE => false; wneuper@59524: fun is_poly_in t v = case poly_deg_in t v of SOME _ => true | NONE => false; wneuper@59524: fun has_degree_in t v = case expand_deg_in t v of SOME d => d | NONE => ~1; neuper@37950: wneuper@59523: (*.the expression contains + - * ^ only ? wneuper@59523: this is weaker than 'is_polynomial' !.*) wneuper@59523: fun is_polyexp (Free _) = true wneuper@59529: | is_polyexp (Const _) = true (* potential danger: bdv is not considered *) walther@60335: | is_polyexp (Const (\<^const_name>\plus\,_) $ Free _ $ num) = walther@60318: if TermC.is_num num then true walther@60318: else if TermC.is_variable num then true walther@60318: else is_polyexp num walther@60335: | is_polyexp (Const (\<^const_name>\plus\, _) $ num $ Free _) = walther@60318: if TermC.is_num num then true walther@60318: else if TermC.is_variable num then true walther@60318: else is_polyexp num walther@60335: | is_polyexp (Const (\<^const_name>\minus\, _) $ Free _ $ num) = walther@60318: if TermC.is_num num then true walther@60318: else if TermC.is_variable num then true walther@60318: else is_polyexp num walther@60335: | is_polyexp (Const (\<^const_name>\times\, _) $ num $ Free _) = walther@60318: if TermC.is_num num then true walther@60318: else if TermC.is_variable num then true walther@60318: else is_polyexp num wenzelm@60405: | is_polyexp (Const (\<^const_name>\realpow\,_) $ Free _ $ num) = walther@60318: if TermC.is_num num then true walther@60318: else if TermC.is_variable num then true walther@60318: else is_polyexp num walther@60335: | is_polyexp (Const (\<^const_name>\plus_class.plus\,_) $ t1 $ t2) = walther@60318: ((is_polyexp t1) andalso (is_polyexp t2)) walther@60335: | is_polyexp (Const (\<^const_name>\Groups.minus_class.minus\,_) $ t1 $ t2) = walther@60318: ((is_polyexp t1) andalso (is_polyexp t2)) walther@60335: | is_polyexp (Const (\<^const_name>\Groups.times_class.times\,_) $ t1 $ t2) = walther@60318: ((is_polyexp t1) andalso (is_polyexp t2)) wenzelm@60405: | is_polyexp (Const (\<^const_name>\realpow\,_) $ t1 $ t2) = walther@60318: ((is_polyexp t1) andalso (is_polyexp t2)) walther@60317: | is_polyexp num = TermC.is_num num; wneuper@59523: \ neuper@37950: wneuper@59523: subsubsection \for hard-coded AC rewriting\ wneuper@59523: ML \ wneuper@59523: (* auch Klammerung muss übereinstimmen; wneuper@59523: sort_variables klammert Produkte rechtslastig*) wneuper@59523: fun is_multUnordered t = ((is_polyexp t) andalso not (t = sort_variables t)); wneuper@59523: wneuper@59523: fun is_addUnordered t = ((is_polyexp t) andalso not (t = sort_monoms t)); wneuper@59523: \ wneuper@59523: wneuper@59523: subsection \evaluations functions\ walther@60278: subsubsection \for predicates\ wneuper@59523: ML \ walther@60335: fun eval_is_polyrat_in _ _(p as (Const (\<^const_name>\Poly.is_polyrat_in\, _) $ t $ v)) _ = wneuper@59523: if is_polyrat_in t v walther@59868: then SOME ((UnparseC.term p) ^ " = True", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True}))) walther@59868: else SOME ((UnparseC.term p) ^ " = True", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False}))) wneuper@59523: | eval_is_polyrat_in _ _ _ _ = ((*tracing"### no matches";*) NONE); wneuper@59523: walther@60278: (*("is_expanded_in", ("Poly.is_expanded_in", eval_is_expanded_in ""))*) wneuper@59523: fun eval_is_expanded_in _ _ walther@60335: (p as (Const (\<^const_name>\Poly.is_expanded_in\, _) $ t $ v)) _ = wneuper@59523: if is_expanded_in t v walther@59868: then SOME ((UnparseC.term p) ^ " = True", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True}))) walther@59868: else SOME ((UnparseC.term p) ^ " = True", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False}))) wneuper@59523: | eval_is_expanded_in _ _ _ _ = NONE; wneuper@59523: walther@60278: (*("is_poly_in", ("Poly.is_poly_in", eval_is_poly_in ""))*) wneuper@59523: fun eval_is_poly_in _ _ walther@60335: (p as (Const (\<^const_name>\Poly.is_poly_in\, _) $ t $ v)) _ = wneuper@59523: if is_poly_in t v walther@59868: then SOME ((UnparseC.term p) ^ " = True", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True}))) walther@59868: else SOME ((UnparseC.term p) ^ " = True", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False}))) wneuper@59523: | eval_is_poly_in _ _ _ _ = NONE; wneuper@59523: walther@60278: (*("has_degree_in", ("Poly.has_degree_in", eval_has_degree_in ""))*) wneuper@59523: fun eval_has_degree_in _ _ walther@60335: (p as (Const (\<^const_name>\Poly.has_degree_in\, _) $ t $ v)) _ = wneuper@59523: let val d = has_degree_in t v wneuper@59523: val d' = TermC.term_of_num HOLogic.realT d walther@59868: in SOME ((UnparseC.term p) ^ " = " ^ (string_of_int d), wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (p, d'))) wneuper@59523: end wneuper@59523: | eval_has_degree_in _ _ _ _ = NONE; wneuper@59523: walther@60278: (*("is_polyexp", ("Poly.is_polyexp", eval_is_polyexp ""))*) wneuper@59523: fun eval_is_polyexp (thmid:string) _ Walther@60504: (t as (Const (\<^const_name>\is_polyexp\, _) $ arg)) ctxt = wneuper@59523: if is_polyexp arg Walther@60504: then SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True}))) Walther@60504: else SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False}))) wneuper@59523: | eval_is_polyexp _ _ _ _ = NONE; wneuper@59523: \ wneuper@59523: wneuper@59523: subsubsection \for hard-coded AC rewriting\ wneuper@59523: ML \ wneuper@59523: (*WN.18.6.03 *) walther@60278: (*("is_addUnordered", ("Poly.is_addUnordered", eval_is_addUnordered ""))*) wneuper@59523: fun eval_is_addUnordered (thmid:string) _ Walther@60504: (t as (Const (\<^const_name>\is_addUnordered\, _) $ arg)) ctxt = wneuper@59523: if is_addUnordered arg Walther@60504: then SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True}))) Walther@60504: else SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False}))) wneuper@59523: | eval_is_addUnordered _ _ _ _ = NONE; wneuper@59523: wneuper@59523: fun eval_is_multUnordered (thmid:string) _ Walther@60504: (t as (Const (\<^const_name>\is_multUnordered\, _) $ arg)) ctxt = wneuper@59523: if is_multUnordered arg Walther@60504: then SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True}))) Walther@60504: else SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", wneuper@59523: HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False}))) wneuper@59523: | eval_is_multUnordered _ _ _ _ = NONE; wneuper@59523: \ wenzelm@60313: calculation is_polyrat_in = \eval_is_polyrat_in "#eval_is_polyrat_in"\ wenzelm@60313: calculation is_expanded_in = \eval_is_expanded_in ""\ wenzelm@60313: calculation is_poly_in = \eval_is_poly_in ""\ wenzelm@60313: calculation has_degree_in = \eval_has_degree_in ""\ wenzelm@60313: calculation is_polyexp = \eval_is_polyexp ""\ wenzelm@60313: calculation is_multUnordered = \eval_is_multUnordered ""\ wenzelm@60313: calculation is_addUnordered = \eval_is_addUnordered ""\ wneuper@59523: wneuper@59523: subsection \rule-sets\ wneuper@59523: subsubsection \without specific order\ wneuper@59523: ML \ wneuper@59523: (* used only for merge *) walther@59852: val calculate_Poly = Rule_Set.append_rules "calculate_PolyFIXXXME.not.impl." Rule_Set.empty []; wneuper@59523: wneuper@59523: (*.for evaluation of conditions in rewrite rules.*) walther@60358: val Poly_erls = Rule_Set.append_rules "Poly_erls" Atools_erls [ walther@60358: \<^rule_eval>\HOL.eq\ (Prog_Expr.eval_equal "#equal_"), wenzelm@60297: \<^rule_thm>\real_unari_minus\, wenzelm@60294: \<^rule_eval>\plus\ (eval_binop "#add_"), wenzelm@60294: \<^rule_eval>\minus\ (eval_binop "#sub_"), wenzelm@60294: \<^rule_eval>\times\ (eval_binop "#mult_"), wenzelm@60405: \<^rule_eval>\realpow\ (eval_binop "#power_")]; wneuper@59523: walther@60358: val poly_crls = Rule_Set.append_rules "poly_crls" Atools_crls [ walther@60358: \<^rule_eval>\HOL.eq\ (Prog_Expr.eval_equal "#equal_"), wenzelm@60297: \<^rule_thm>\real_unari_minus\, wenzelm@60294: \<^rule_eval>\plus\ (eval_binop "#add_"), wenzelm@60294: \<^rule_eval>\minus\ (eval_binop "#sub_"), wenzelm@60294: \<^rule_eval>\times\ (eval_binop "#mult_"), wenzelm@60405: \<^rule_eval>\realpow\ (eval_binop "#power_")]; wneuper@59523: \ wneuper@59523: ML \ neuper@37950: val expand = Walther@60509: Rule_Def.Repeat {id = "expand", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@59852: erls = Rule_Set.empty,srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\distrib_right\, (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*) walther@60358: \<^rule_thm>\distrib_left\ (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: walther@60385: \ ML \ walther@60385: \ ML \@{term "1 is_num"}\ walther@60385: ML \ walther@60320: (* erls for calculate_Rational + etc *) walther@60318: val powers_erls = Walther@60509: Rule_Def.Repeat {id = "powers_erls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.function_empty), walther@60318: erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_eval>\matches\ (Prog_Expr.eval_matches "#matches_"), walther@60358: \<^rule_eval>\is_atom\ (Prog_Expr.eval_is_atom "#is_atom_"), walther@60358: \<^rule_eval>\is_num\ (Prog_Expr.eval_is_num "#is_num_"), walther@60358: \<^rule_eval>\is_even\ (Prog_Expr.eval_is_even "#is_even_"), walther@60358: \<^rule_eval>\ord_class.less\ (Prog_Expr.eval_equ "#less_"), walther@60358: \<^rule_thm>\not_false\, walther@60358: \<^rule_thm>\not_true\, walther@60358: \<^rule_eval>\plus_class.plus\ (eval_binop "#add_")], walther@60358: scr = Rule.Empty_Prog}; walther@60320: walther@60333: \ ML \ walther@60320: val discard_minus = Walther@60509: Rule_Def.Repeat {id = "discard_minus", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60320: erls = powers_erls, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60341: rules = [ walther@60341: \<^rule_thm>\real_diff_minus\ (*"a - b = a + -1 * b"*), walther@60385: \<^rule_thm>\real_mult_minus1_sym\ (*"\(z is_num) ==> - (z::real) = -1 * z"*)], walther@60341: scr = Rule.Empty_Prog}; walther@60320: neuper@37950: val expand_poly_ = walther@59851: Rule_Def.Repeat{id = "expand_poly_", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60318: erls = powers_erls, srls = Rule_Set.Empty, neuper@42451: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\real_plus_binom_pow4\, (*"(a + b) \ 4 = ... "*) walther@60358: \<^rule_thm>\real_plus_binom_pow5\, (*"(a + b) \ 5 = ... "*) walther@60358: \<^rule_thm>\real_plus_binom_pow3\, (*"(a + b) \ 3 = a \ 3 + 3*a \ 2*b + 3*a*b \ 2 + b \ 3" *) walther@60358: (*WN071229 changed/removed for Schaerding -----vvv*) walther@60358: (*\<^rule_thm>\real_plus_binom_pow2\,*) (*"(a + b) \ 2 = a \ 2 + 2*a*b + b \ 2"*) walther@60358: \<^rule_thm>\real_plus_binom_pow2\, (*"(a + b) \ 2 = (a + b) * (a + b)"*) walther@60358: (*\<^rule_thm>\real_plus_minus_binom1_p_p\,*) (*"(a + b)*(a + -1 * b) = a \ 2 + -1*b \ 2"*) walther@60358: (*\<^rule_thm>\real_plus_minus_binom2_p_p\,*) (*"(a + -1 * b)*(a + b) = a \ 2 + -1*b \ 2"*) walther@60358: (*WN071229 changed/removed for Schaerding -----^^^*) walther@60358: walther@60358: \<^rule_thm>\distrib_right\, (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*) walther@60358: \<^rule_thm>\distrib_left\, (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*) walther@60358: walther@60358: \<^rule_thm>\realpow_multI\, (*"(r * s) \ n = r \ n * s \ n"*) walther@60358: \<^rule_thm>\realpow_pow\, (*"(a \ b) \ c = a \ (b * c)"*) walther@60358: walther@60358: \<^rule_thm>\realpow_minus_even\, (*"n is_even ==> (- r) \ n = r \ n"*) walther@60358: \<^rule_thm>\realpow_minus_odd\ (*"Not (n is_even) ==> (- r) \ n = -1 * r \ n"*)], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val expand_poly_rat_ = walther@59851: Rule_Def.Repeat{id = "expand_poly_rat_", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: erls = Rule_Set.append_rules "Rule_Set.empty-expand_poly_rat_" Rule_Set.empty [ walther@60358: \<^rule_eval>\is_polyexp\ (eval_is_polyexp ""), walther@60358: \<^rule_eval>\is_even\ (Prog_Expr.eval_is_even "#is_even_"), walther@60358: \<^rule_thm>\not_false\, walther@60358: \<^rule_thm>\not_true\ ], walther@59851: srls = Rule_Set.Empty, neuper@42451: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\real_plus_binom_pow4_poly\, (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \ 4 = ... "*) walther@60358: \<^rule_thm>\real_plus_binom_pow5_poly\, (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \ 5 = ... "*) walther@60358: \<^rule_thm>\real_plus_binom_pow2_poly\, (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \ 2 = a \ 2 + 2*a*b + b \ 2"*) walther@60358: \<^rule_thm>\real_plus_binom_pow3_poly\, walther@60358: (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \ 3 = a \ 3 + 3*a \ 2*b + 3*a*b \ 2 + b \ 3" *) walther@60358: \<^rule_thm>\real_plus_minus_binom1_p_p\, (*"(a + b)*(a + -1 * b) = a \ 2 + -1*b \ 2"*) walther@60358: \<^rule_thm>\real_plus_minus_binom2_p_p\, (*"(a + -1 * b)*(a + b) = a \ 2 + -1*b \ 2"*) walther@60358: walther@60358: \<^rule_thm>\real_add_mult_distrib_poly\, (*"w is_polyexp ==> (z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*) walther@60358: \<^rule_thm>\real_add_mult_distrib2_poly\, (*"w is_polyexp ==> w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*) walther@60358: walther@60358: \<^rule_thm>\realpow_multI_poly\, (*"[| r is_polyexp; s is_polyexp |] ==> (r * s) \ n = r \ n * s \ n"*) walther@60358: \<^rule_thm>\realpow_pow\, (*"(a \ b) \ c = a \ (b * c)"*) walther@60358: \<^rule_thm>\realpow_minus_even\, (*"n is_even ==> (- r) \ n = r \ n"*) walther@60358: \<^rule_thm>\realpow_minus_odd\ (*"\ (n is_even) ==> (- r) \ n = -1 * r \ n"*) ], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val simplify_power_ = walther@59851: Rule_Def.Repeat{id = "simplify_power_", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@59852: erls = Rule_Set.empty, srls = Rule_Set.Empty, neuper@42451: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: (*MG: Reihenfolge der folgenden 2 Rule.Thm muss so bleiben, wegen walther@60358: a*(a*a) --> a*a \ 2 und nicht a*(a*a) --> a \ 2*a *) walther@60358: \<^rule_thm_sym>\realpow_twoI\, (*"r * r = r \ 2"*) walther@60358: \<^rule_thm>\realpow_twoI_assoc_l\, (*"r * (r * s) = r \ 2 * s"*) walther@60358: walther@60358: \<^rule_thm>\realpow_plus_1\, (*"r * r \ n = r \ (n + 1)"*) walther@60358: \<^rule_thm>\realpow_plus_1_assoc_l\, (*"r * (r \ m * s) = r \ (1 + m) * s"*) walther@60358: (*MG 9.7.03: neues Rule.Thm wegen a*(a*(a*b)) --> a \ 2*(a*b) *) walther@60358: \<^rule_thm>\realpow_plus_1_assoc_l2\, (*"r \ m * (r * s) = r \ (1 + m) * s"*) walther@60358: walther@60358: \<^rule_thm_sym>\realpow_addI\, (*"r \ n * r \ m = r \ (n + m)"*) walther@60358: \<^rule_thm>\realpow_addI_assoc_l\, (*"r \ n * (r \ m * s) = r \ (n + m) * s"*) walther@60358: walther@60358: (* ist in expand_poly - wird hier aber auch gebraucht, wegen: "r * r = r \ 2" wenn r=a \ b*) walther@60358: \<^rule_thm>\realpow_pow\ (*"(a \ b) \ c = a \ (b * c)"*)], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val calc_add_mult_pow_ = walther@59851: Rule_Def.Repeat{id = "calc_add_mult_pow_", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@59851: erls = Atools_erls(*erls3.4.03*),srls = Rule_Set.Empty, wenzelm@60309: calc = [("PLUS" , (\<^const_name>\plus\, eval_binop "#add_")), wenzelm@60309: ("TIMES" , (\<^const_name>\times\, eval_binop "#mult_")), wenzelm@60405: ("POWER", (\<^const_name>\realpow\, eval_binop "#power_")) neuper@37950: ], neuper@42451: errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_eval>\plus\ (eval_binop "#add_"), walther@60358: \<^rule_eval>\times\ (eval_binop "#mult_"), wenzelm@60405: \<^rule_eval>\realpow\ (eval_binop "#power_")], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val reduce_012_mult_ = walther@59851: Rule_Def.Repeat{id = "reduce_012_mult_", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@59852: erls = Rule_Set.empty,srls = Rule_Set.Empty, neuper@42451: calc = [], errpatts = [], wneuper@59416: rules = [(* MG: folgende Rule.Thm müssen hier stehen bleiben: *) walther@60358: \<^rule_thm>\mult_1_right\, (*"z * 1 = z"*) (*wegen "a * b * b \ (-1) + a"*) walther@60358: \<^rule_thm>\realpow_zeroI\, (*"r \ 0 = 1"*) (*wegen "a*a \ (-1)*c + b + c"*) walther@60358: \<^rule_thm>\realpow_oneI\, (*"r \ 1 = r"*) walther@60358: \<^rule_thm>\realpow_eq_oneI\ (*"1 \ n = 1"*)], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val collect_numerals_ = walther@59851: Rule_Def.Repeat{id = "collect_numerals_", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@59851: erls = Atools_erls, srls = Rule_Set.Empty, wenzelm@60309: calc = [("PLUS" , (\<^const_name>\plus\, eval_binop "#add_")) neuper@42451: ], errpatts = [], walther@60358: rules = [ walther@60385: \<^rule_thm>\real_num_collect\, (*"[| l is_num; m is_num |]==>l * n + m * n = (l + m) * n"*) walther@60358: \<^rule_thm>\real_num_collect_assoc_r\, walther@60385: (*"[| l is_num; m is_num |] ==> (k + m * n) + l * n = k + (l + m)*n"*) walther@60385: \<^rule_thm>\real_one_collect\, (*"m is_num ==> n + m * n = (1 + m) * n"*) walther@60385: \<^rule_thm>\real_one_collect_assoc_r\, (*"m is_num ==> (k + n) + m * n = k + (m + 1) * n"*) neuper@37950: walther@60358: \<^rule_eval>\plus\ (eval_binop "#add_"), neuper@37950: walther@60358: (*MG: Reihenfolge der folgenden 2 Rule.Thm muss so bleiben, wegen walther@60358: (a+a)+a --> a + 2*a --> 3*a and not (a+a)+a --> 2*a + a *) walther@60358: \<^rule_thm>\real_mult_2_assoc_r\, (*"(k + z1) + z1 = k + 2 * z1"*) walther@60358: \<^rule_thm_sym>\real_mult_2\ (*"z1 + z1 = 2 * z1"*)], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val reduce_012_ = walther@59851: Rule_Def.Repeat{id = "reduce_012_", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@59852: erls = Rule_Set.empty,srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\mult_1_left\, (*"1 * z = z"*) walther@60358: \<^rule_thm>\mult_zero_left\, (*"0 * z = 0"*) walther@60358: \<^rule_thm>\mult_zero_right\, (*"z * 0 = 0"*) walther@60358: \<^rule_thm>\add_0_left\, (*"0 + z = z"*) walther@60358: \<^rule_thm>\add_0_right\, (*"z + 0 = z"*) (*wegen a+b-b --> a+(1-1)*b --> a+0 --> a*) walther@60358: walther@60358: (*\<^rule_thm>\realpow_oneI\*) (*"?r \ 1 = ?r"*) walther@60358: \<^rule_thm>\division_ring_divide_zero\ (*"0 / ?x = 0"*)], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37979: val discard_parentheses1 = walther@60358: Rule_Set.append_rules "discard_parentheses1" Rule_Set.empty [ walther@60358: \<^rule_thm_sym>\mult.assoc\ (*"?z1.1 * (?z2.1 * ?z3.1) = ?z1.1 * ?z2.1 * ?z3.1"*) walther@60358: (*\<^rule_thm_sym>\add.assoc\*) (*"?z1.1 + (?z2.1 + ?z3.1) = ?z1.1 + ?z2.1 + ?z3.1"*)]; neuper@37950: wneuper@59523: val expand_poly = walther@59851: Rule_Def.Repeat{id = "expand_poly", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60320: erls = powers_erls, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\distrib_right\, (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*) walther@60358: \<^rule_thm>\distrib_left\, (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*) neuper@37950: walther@60358: \<^rule_thm>\real_plus_binom_pow2\, (*"(a + b) \ 2 = a \ 2 + 2*a*b + b \ 2"*) walther@60358: \<^rule_thm>\real_minus_binom_pow2_p\, (*"(a - b) \ 2 = a \ 2 + -2*a*b + b \ 2"*) walther@60358: \<^rule_thm>\real_plus_minus_binom1_p\, (*"(a + b)*(a - b) = a \ 2 + -1*b \ 2"*) walther@60358: \<^rule_thm>\real_plus_minus_binom2_p\, (*"(a - b)*(a + b) = a \ 2 + -1*b \ 2"*) neuper@37950: walther@60358: \<^rule_thm>\minus_minus\ (*"- (- ?z) = ?z"*), walther@60358: \<^rule_thm>\real_diff_minus\ (*"a - b = a + -1 * b"*), walther@60385: \<^rule_thm>\real_mult_minus1_sym\ (*"\(z is_num) ==> - (z::real) = -1 * z"*) neuper@37950: walther@60358: (*\<^rule_thm>\real_minus_add_distrib\,*) (*"- (?x + ?y) = - ?x + - ?y"*) walther@60358: (*\<^rule_thm>\real_diff_plus\*) (*"a - b = a + -b"*)], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: neuper@37950: val simplify_power = walther@59851: Rule_Def.Repeat{id = "simplify_power", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@59852: erls = Rule_Set.empty, srls = Rule_Set.Empty, neuper@42451: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\realpow_multI\, (*"(r * s) \ n = r \ n * s \ n"*) walther@60358: walther@60358: \<^rule_thm_sym>\realpow_twoI\, (*"r1 * r1 = r1 \ 2"*) walther@60358: \<^rule_thm>\realpow_plus_1\, (*"r * r \ n = r \ (n + 1)"*) walther@60358: \<^rule_thm>\realpow_pow\, (*"(a \ b) \ c = a \ (b * c)"*) walther@60358: \<^rule_thm_sym>\realpow_addI\, (*"r \ n * r \ m = r \ (n + m)"*) walther@60358: \<^rule_thm>\realpow_oneI\, (*"r \ 1 = r"*) walther@60358: \<^rule_thm>\realpow_eq_oneI\ (*"1 \ n = 1"*)], walther@60358: scr = Rule.Empty_Prog}; neuper@42451: neuper@37950: val collect_numerals = walther@59851: Rule_Def.Repeat{id = "collect_numerals", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@59851: erls = Atools_erls(*erls3.4.03*),srls = Rule_Set.Empty, wenzelm@60309: calc = [("PLUS" , (\<^const_name>\plus\, eval_binop "#add_")), wenzelm@60309: ("TIMES" , (\<^const_name>\times\, eval_binop "#mult_")), wenzelm@60405: ("POWER", (\<^const_name>\realpow\, eval_binop "#power_")) neuper@42451: ], errpatts = [], walther@60358: rules =[ walther@60385: \<^rule_thm>\real_num_collect\, (*"[| l is_num; m is_num |]==>l * n + m * n = (l + m) * n"*) walther@60358: \<^rule_thm>\real_num_collect_assoc\, walther@60385: (*"[| l is_num; m is_num |] ==> l * n + (m * n + k) = (l + m) * n + k"*) walther@60385: \<^rule_thm>\real_one_collect\, (*"m is_num ==> n + m * n = (1 + m) * n"*) walther@60385: \<^rule_thm>\real_one_collect_assoc\, (*"m is_num ==> k + (n + m * n) = k + (1 + m) * n"*) walther@60358: \<^rule_eval>\plus\ (eval_binop "#add_"), walther@60358: \<^rule_eval>\times\ (eval_binop "#mult_"), wenzelm@60405: \<^rule_eval>\realpow\ (eval_binop "#power_")], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: val reduce_012 = walther@59851: Rule_Def.Repeat{id = "reduce_012", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60344: erls = Rule_Set.append_rules "erls_in_reduce_012" Rule_Set.empty [ walther@60344: \<^rule_eval>\is_num\ (Prog_Expr.eval_is_num "#is_num_"), walther@60344: \<^rule_thm>\not_false\, walther@60344: \<^rule_thm>\not_true\], walther@60344: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\mult_1_left\, (*"1 * z = z"*) walther@60358: (*\<^rule_thm>\real_mult_minus1\,14.3.03*) (*"-1 * z = - z"*) walther@60344: \<^rule_thm>\real_minus_mult_left\, (*"\ ((a::real) is_num) ==> (- a) * b = - (a * b)"*) walther@60358: (*\<^rule_thm>\real_minus_mult_cancel\, (*"- ?x * - ?y = ?x * ?y"*)---*) walther@60358: \<^rule_thm>\mult_zero_left\, (*"0 * z = 0"*) walther@60344: \<^rule_thm>\add_0_left\, (*"0 + z = z"*) walther@60344: \<^rule_thm>\add_0_right\, (*"a + - a = 0"*) walther@60358: \<^rule_thm>\right_minus\, (*"?z + - ?z = 0"*) walther@60358: \<^rule_thm_sym>\real_mult_2\, (*"z1 + z1 = 2 * z1"*) walther@60358: \<^rule_thm>\real_mult_2_assoc\ (*"z1 + (z1 + k) = 2 * z1 + k"*)], walther@60358: scr = Rule.Empty_Prog}; neuper@52139: neuper@37950: val discard_parentheses = walther@59852: Rule_Set.append_rules "discard_parentheses" Rule_Set.empty wenzelm@60296: [\<^rule_thm_sym>\mult.assoc\, \<^rule_thm_sym>\add.assoc\]; wneuper@59523: \ neuper@37950: wneuper@59523: subsubsection \hard-coded AC rewriting\ wneuper@59523: ML \ wneuper@59523: (*MG.0401: termorders for multivariate polys dropped due to principal problems: wneuper@59523: (total-degree-)ordering of monoms NOT possible with size_of_term GIVEN*) wneuper@59523: val order_add_mult = walther@59851: Rule_Def.Repeat{id = "order_add_mult", preconds = [], Walther@60509: rew_ord = ("ord_make_polynomial", ord_make_polynomial false \<^theory>), walther@59852: erls = Rule_Set.empty,srls = Rule_Set.Empty, neuper@42451: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\mult.commute\, (* z * w = w * z *) walther@60358: \<^rule_thm>\real_mult_left_commute\, (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*) walther@60358: \<^rule_thm>\mult.assoc\, (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*) walther@60358: \<^rule_thm>\add.commute\, (*z + w = w + z*) walther@60358: \<^rule_thm>\add.left_commute\, (*x + (y + z) = y + (x + z)*) walther@60358: \<^rule_thm>\add.assoc\ (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)], walther@60358: scr = Rule.Empty_Prog}; wneuper@59523: (*MG.0401: termorders for multivariate polys dropped due to principal problems: wneuper@59523: (total-degree-)ordering of monoms NOT possible with size_of_term GIVEN*) wneuper@59523: val order_mult = walther@59851: Rule_Def.Repeat{id = "order_mult", preconds = [], wenzelm@60291: rew_ord = ("ord_make_polynomial",ord_make_polynomial false \<^theory>), walther@59852: erls = Rule_Set.empty,srls = Rule_Set.Empty, wneuper@59523: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\mult.commute\, (* z * w = w * z *) walther@60358: \<^rule_thm>\real_mult_left_commute\, (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*) walther@60358: \<^rule_thm>\mult.assoc\ (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)], walther@60358: scr = Rule.Empty_Prog}; wneuper@59472: \ wneuper@59472: ML \ wneuper@59416: fun attach_form (_: Rule.rule list list) (_: term) (_: term) = (*still missing*) wneuper@59416: []:(Rule.rule * (term * term list)) list; walther@59850: fun init_state (_: term) = Rule_Set.e_rrlsstate; wneuper@59416: fun locate_rule (_: Rule.rule list list) (_: term) (_: Rule.rule) = wneuper@59416: ([]:(Rule.rule * (term * term list)) list); wneuper@59416: fun next_rule (_: Rule.rule list list) (_: term) = (NONE: Rule.rule option); wneuper@59406: fun normal_form t = SOME (sort_variables t, []: term list); neuper@37950: neuper@37950: val order_mult_ = walther@59850: Rule_Set.Rrls {id = "order_mult_", neuper@37950: prepat = neuper@38036: (* ?p matched with the current term gives an environment, neuper@38037: which evaluates (the instantiated) "?p is_multUnordered" to true *) wenzelm@60291: [([TermC.parse_patt \<^theory> "?p is_multUnordered"], wenzelm@60291: TermC.parse_patt \<^theory> "?p :: real")], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@59852: erls = Rule_Set.append_rules "Rule_Set.empty-is_multUnordered" Rule_Set.empty wenzelm@60294: [\<^rule_eval>\is_multUnordered\ (eval_is_multUnordered "")], wenzelm@60309: calc = [("PLUS" , (\<^const_name>\plus\, eval_binop "#add_")), wenzelm@60309: ("TIMES" , (\<^const_name>\times\, eval_binop "#mult_")), wenzelm@60309: ("DIVIDE", (\<^const_name>\divide\, Prog_Expr.eval_cancel "#divide_e")), wenzelm@60405: ("POWER" , (\<^const_name>\realpow\, eval_binop "#power_"))], wneuper@59406: errpatts = [], wneuper@59416: scr = Rule.Rfuns {init_state = init_state, neuper@37950: normal_form = normal_form, neuper@37950: locate_rule = locate_rule, neuper@37950: next_rule = next_rule, neuper@37950: attach_form = attach_form}}; neuper@37950: val order_mult_rls_ = walther@59851: Rule_Def.Repeat {id = "order_mult_rls_", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: erls = Rule_Set.empty,srls = Rule_Set.Empty, walther@60358: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: Rule.Rls_ order_mult_], walther@60358: scr = Rule.Empty_Prog}; neuper@37950: wneuper@59523: \ ML \ neuper@37950: wneuper@59416: fun attach_form (_: Rule.rule list list) (_: term) (_: term) = (*still missing*) wneuper@59416: []: (Rule.rule * (term * term list)) list; walther@59850: fun init_state (_: term) = Rule_Set.e_rrlsstate; wneuper@59416: fun locate_rule (_: Rule.rule list list) (_: term) (_: Rule.rule) = wneuper@59416: ([]: (Rule.rule * (term * term list)) list); wneuper@59416: fun next_rule (_: Rule.rule list list) (_: term) = (NONE: Rule.rule option); wneuper@59406: fun normal_form t = SOME (sort_monoms t,[]: term list); wneuper@59472: \ ML \ neuper@37950: val order_add_ = walther@60358: Rule_Set.Rrls {id = "order_add_", walther@60358: prepat = [(*WN.18.6.03 Preconditions und Pattern, walther@59850: die beide passen muessen, damit das Rule_Set.Rrls angewandt wird*) walther@60358: ([TermC.parse_patt @{theory} "?p is_addUnordered"], walther@60358: TermC.parse_patt @{theory} "?p :: real" neuper@37950: (*WN.18.6.03 also KEIN pattern, dieses erzeugt nur das Environment neuper@37950: fuer die Evaluation der Precondition "p is_addUnordered"*))], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@59852: erls = Rule_Set.append_rules "Rule_Set.empty-is_addUnordered" Rule_Set.empty(*MG: poly_erls*) walther@60358: [\<^rule_eval>\is_addUnordered\ (eval_is_addUnordered "")], walther@60358: calc = [ walther@60358: ("PLUS" ,(\<^const_name>\plus\, eval_binop "#add_")), wenzelm@60309: ("TIMES" ,(\<^const_name>\times\, eval_binop "#mult_")), wenzelm@60309: ("DIVIDE",(\<^const_name>\divide\, Prog_Expr.eval_cancel "#divide_e")), wenzelm@60405: ("POWER" ,(\<^const_name>\realpow\ , eval_binop "#power_"))], neuper@42451: errpatts = [], walther@60358: scr = Rule.Rfuns { walther@60358: init_state = init_state, walther@60358: normal_form = normal_form, walther@60358: locate_rule = locate_rule, walther@60358: next_rule = next_rule, walther@60358: attach_form = attach_form}}; neuper@37950: wneuper@59406: val order_add_rls_ = walther@59851: Rule_Def.Repeat {id = "order_add_rls_", preconds = [], Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: erls = Rule_Set.empty, srls = Rule_Set.Empty, walther@60358: calc = [], errpatts = [], walther@60358: rules = [ walther@60358: Rule.Rls_ order_add_], walther@60358: scr = Rule.Empty_Prog}; wneuper@59472: \ neuper@37950: wneuper@59472: text \rule-set make_polynomial also named norm_Poly: neuper@42398: Rewrite order has not been implemented properly; the order is better in neuper@42398: make_polynomial_in (coded in SML). neuper@42398: Notes on state of development: neuper@42398: \# surprise 2006: test --- norm_Poly NOT COMPLETE --- neuper@42398: \# migration Isabelle2002 --> 2011 weakened the rule set, see test walther@59962: --- Matthias Goldgruber 2003 rewrite orders ---, raise ERROR "ord_make_polynomial_in #16b" wneuper@59472: \ wneuper@59472: ML \ neuper@37950: (*. see MG-DA.p.52ff .*) neuper@37950: val make_polynomial(*MG.03, overwrites version from above, neuper@37950: previously 'make_polynomial_'*) = walther@59878: Rule_Set.Sequence {id = "make_polynomial", preconds = []:term list, Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: erls = Atools_erls, srls = Rule_Set.Empty,calc = [], errpatts = [], walther@60358: rules = [ walther@60358: Rule.Rls_ discard_minus, walther@60358: Rule.Rls_ expand_poly_, walther@60358: \<^rule_eval>\times\ (eval_binop "#mult_"), walther@60358: Rule.Rls_ order_mult_rls_, walther@60358: Rule.Rls_ simplify_power_, walther@60358: Rule.Rls_ calc_add_mult_pow_, walther@60358: Rule.Rls_ reduce_012_mult_, walther@60358: Rule.Rls_ order_add_rls_, walther@60358: Rule.Rls_ collect_numerals_, walther@60358: Rule.Rls_ reduce_012_, walther@60358: Rule.Rls_ discard_parentheses1], walther@60358: scr = Rule.Empty_Prog}; wneuper@59472: \ wneuper@59472: ML \ neuper@37950: val norm_Poly(*=make_polynomial*) = walther@59878: Rule_Set.Sequence {id = "norm_Poly", preconds = []:term list, Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: Rule.Rls_ discard_minus, walther@60358: Rule.Rls_ expand_poly_, walther@60358: \<^rule_eval>\times\ (eval_binop "#mult_"), walther@60358: Rule.Rls_ order_mult_rls_, walther@60358: Rule.Rls_ simplify_power_, walther@60358: Rule.Rls_ calc_add_mult_pow_, walther@60358: Rule.Rls_ reduce_012_mult_, walther@60358: Rule.Rls_ order_add_rls_, walther@60358: Rule.Rls_ collect_numerals_, walther@60358: Rule.Rls_ reduce_012_, walther@60358: Rule.Rls_ discard_parentheses1], walther@60358: scr = Rule.Empty_Prog}; wneuper@59472: \ wneuper@59472: ML \ wneuper@59416: (* MG:03 Like make_polynomial_ but without Rule.Rls_ discard_parentheses1 neuper@37950: and expand_poly_rat_ instead of expand_poly_, see MG-DA.p.56ff*) neuper@37950: (* MG necessary for termination of norm_Rational(*_mg*) in Rational.ML*) neuper@37950: val make_rat_poly_with_parentheses = walther@59878: Rule_Set.Sequence{id = "make_rat_poly_with_parentheses", preconds = []:term list, Walther@60509: rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), walther@60358: erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: Rule.Rls_ discard_minus, walther@60358: Rule.Rls_ expand_poly_rat_,(*ignors rationals*) walther@60358: \<^rule_eval>\times\ (eval_binop "#mult_"), walther@60358: Rule.Rls_ order_mult_rls_, walther@60358: Rule.Rls_ simplify_power_, walther@60358: Rule.Rls_ calc_add_mult_pow_, walther@60358: Rule.Rls_ reduce_012_mult_, walther@60358: Rule.Rls_ order_add_rls_, walther@60358: Rule.Rls_ collect_numerals_, walther@60358: Rule.Rls_ reduce_012_ walther@60358: (*Rule.Rls_ discard_parentheses1 *)], walther@60358: scr = Rule.Empty_Prog}; wneuper@59472: \ wneuper@59472: ML \ neuper@37950: (*.a minimal ruleset for reverse rewriting of factions [2]; neuper@37950: compare expand_binoms.*) neuper@37950: val rev_rew_p = walther@60358: Rule_Set.Sequence{id = "rev_rew_p", preconds = [], rew_ord = ("termlessI",termlessI), walther@59851: erls = Atools_erls, srls = Rule_Set.Empty, wenzelm@60309: calc = [(*("PLUS" , (\<^const_name>\plus\, eval_binop "#add_")), wenzelm@60309: ("TIMES" , (\<^const_name>\times\, eval_binop "#mult_")), wenzelm@60405: ("POWER", (\<^const_name>\realpow\, eval_binop "#power_"))*) neuper@42451: ], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\real_plus_binom_times\, (*"(a + b)*(a + b) = a ^ 2 + 2 * a * b + b ^ 2*) walther@60358: \<^rule_thm>\real_plus_binom_times1\, (*"(a + 1*b)*(a + -1*b) = a \ 2 + -1*b \ 2"*) walther@60358: \<^rule_thm>\real_plus_binom_times2\, (*"(a + -1*b)*(a + 1*b) = a \ 2 + -1*b \ 2"*) neuper@37950: wenzelm@60297: \<^rule_thm>\mult_1_left\,(*"1 * z = z"*) neuper@37950: walther@60358: \<^rule_thm>\distrib_right\, (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*) walther@60358: \<^rule_thm>\distrib_left\, (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*) neuper@37950: walther@60358: \<^rule_thm>\mult.assoc\, (*"?z1.1 * ?z2.1 * ?z3. =1 ?z1.1 * (?z2.1 * ?z3.1)"*) wneuper@59416: Rule.Rls_ order_mult_rls_, wneuper@59416: (*Rule.Rls_ order_add_rls_,*) neuper@37950: wenzelm@60294: \<^rule_eval>\plus\ (eval_binop "#add_"), wenzelm@60294: \<^rule_eval>\times\ (eval_binop "#mult_"), wenzelm@60405: \<^rule_eval>\realpow\ (eval_binop "#power_"), neuper@37950: walther@60358: \<^rule_thm_sym>\realpow_twoI\, (*"r1 * r1 = r1 \ 2"*) walther@60358: \<^rule_thm_sym>\real_mult_2\, (*"z1 + z1 = 2 * z1"*) walther@60358: \<^rule_thm>\real_mult_2_assoc\, (*"z1 + (z1 + k) = 2 * z1 + k"*) neuper@37950: walther@60385: \<^rule_thm>\real_num_collect\, (*"[| l is_num; m is_num |]==>l * n + m * n = (l + m) * n"*) walther@60385: \<^rule_thm>\real_num_collect_assoc\, (*"[| l is_num; m is_num |] ==> l * n + (m * n + k) = (l + m) * n + k"*) walther@60385: \<^rule_thm>\real_one_collect\, (*"m is_num ==> n + m * n = (1 + m) * n"*) walther@60385: \<^rule_thm>\real_one_collect_assoc\, (*"m is_num ==> k + (n + m * n) = k + (1 + m) * n"*) neuper@37950: walther@60358: \<^rule_thm>\realpow_multI\, (*"(r * s) \ n = r \ n * s \ n"*) neuper@37950: wenzelm@60294: \<^rule_eval>\plus\ (eval_binop "#add_"), wenzelm@60294: \<^rule_eval>\times\ (eval_binop "#mult_"), wenzelm@60405: \<^rule_eval>\realpow\ (eval_binop "#power_"), neuper@37950: walther@60358: \<^rule_thm>\mult_1_left\, (*"1 * z = z"*) walther@60358: \<^rule_thm>\mult_zero_left\, (*"0 * z = 0"*) walther@60358: \<^rule_thm>\add_0_left\ (*0 + z = z*) neuper@37950: wneuper@59416: (*Rule.Rls_ order_add_rls_*) neuper@37950: ], walther@59878: scr = Rule.Empty_Prog}; wneuper@59472: \ neuper@52125: wneuper@59523: subsection \rule-sets with explicit program for intermediate steps\ wneuper@59523: partial_function (tailrec) expand_binoms_2 :: "real \ real" wneuper@59523: where walther@59635: "expand_binoms_2 term = ( walther@59635: Repeat ( walther@59637: (Try (Repeat (Rewrite ''real_plus_binom_pow2''))) #> walther@59637: (Try (Repeat (Rewrite ''real_plus_binom_times''))) #> walther@59637: (Try (Repeat (Rewrite ''real_minus_binom_pow2''))) #> walther@59637: (Try (Repeat (Rewrite ''real_minus_binom_times''))) #> walther@59637: (Try (Repeat (Rewrite ''real_plus_minus_binom1''))) #> walther@59637: (Try (Repeat (Rewrite ''real_plus_minus_binom2''))) #> walther@59635: walther@59637: (Try (Repeat (Rewrite ''mult_1_left''))) #> walther@59637: (Try (Repeat (Rewrite ''mult_zero_left''))) #> walther@59637: (Try (Repeat (Rewrite ''add_0_left''))) #> walther@59635: walther@59637: (Try (Repeat (Calculate ''PLUS''))) #> walther@59637: (Try (Repeat (Calculate ''TIMES''))) #> walther@59637: (Try (Repeat (Calculate ''POWER''))) #> walther@59635: walther@59637: (Try (Repeat (Rewrite ''sym_realpow_twoI''))) #> walther@59637: (Try (Repeat (Rewrite ''realpow_plus_1''))) #> walther@59637: (Try (Repeat (Rewrite ''sym_real_mult_2''))) #> walther@59637: (Try (Repeat (Rewrite ''real_mult_2_assoc''))) #> walther@59635: walther@59637: (Try (Repeat (Rewrite ''real_num_collect''))) #> walther@59637: (Try (Repeat (Rewrite ''real_num_collect_assoc''))) #> walther@59635: walther@59637: (Try (Repeat (Rewrite ''real_one_collect''))) #> walther@59637: (Try (Repeat (Rewrite ''real_one_collect_assoc''))) #> walther@59635: walther@59637: (Try (Repeat (Calculate ''PLUS''))) #> walther@59637: (Try (Repeat (Calculate ''TIMES''))) #> walther@59635: (Try (Repeat (Calculate ''POWER'')))) walther@59635: term)" wneuper@59523: ML \ wneuper@59523: val expand_binoms = walther@59851: Rule_Def.Repeat{id = "expand_binoms", preconds = [], rew_ord = ("termlessI",termlessI), walther@60358: erls = Atools_erls, srls = Rule_Set.Empty, walther@60358: calc = [("PLUS" , (\<^const_name>\plus\, eval_binop "#add_")), walther@60358: ("TIMES" , (\<^const_name>\times\, eval_binop "#mult_")), wenzelm@60405: ("POWER", (\<^const_name>\realpow\, eval_binop "#power_")) walther@60358: ], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\real_plus_binom_pow2\, (*"(a + b) \ 2 = a \ 2 + 2 * a * b + b \ 2"*) walther@60358: \<^rule_thm>\real_plus_binom_times\, (*"(a + b)*(a + b) = ...*) walther@60358: \<^rule_thm>\real_minus_binom_pow2\, (*"(a - b) \ 2 = a \ 2 - 2 * a * b + b \ 2"*) walther@60358: \<^rule_thm>\real_minus_binom_times\, (*"(a - b)*(a - b) = ...*) walther@60358: \<^rule_thm>\real_plus_minus_binom1\, (*"(a + b) * (a - b) = a \ 2 - b \ 2"*) walther@60358: \<^rule_thm>\real_plus_minus_binom2\, (*"(a - b) * (a + b) = a \ 2 - b \ 2"*) walther@60358: (*RL 020915*) walther@60358: \<^rule_thm>\real_pp_binom_times\, (*(a + b)*(c + d) = a*c + a*d + b*c + b*d*) walther@60384: \<^rule_thm>\real_pm_binom_times\, (*(a + b)*(c - d) = a*c - a*d + b*c - b*d*) walther@60384: \<^rule_thm>\real_mp_binom_times\, (*(a - b)*(c + d) = a*c + a*d - b*c - b*d*) walther@60384: \<^rule_thm>\real_mm_binom_times\, (*(a - b)*(c - d) = a*c - a*d - b*c + b*d*) walther@60358: \<^rule_thm>\realpow_multI\, (*(a*b) \ n = a \ n * b \ n*) walther@60358: \<^rule_thm>\real_plus_binom_pow3\, (* (a + b) \ 3 = a \ 3 + 3*a \ 2*b + 3*a*b \ 2 + b \ 3 *) walther@60358: \<^rule_thm>\real_minus_binom_pow3\, (* (a - b) \ 3 = a \ 3 - 3*a \ 2*b + 3*a*b \ 2 - b \ 3 *) walther@60358: (* walther@60358: \<^rule_thm>\distrib_right\, (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*) walther@60358: \<^rule_thm>\distrib_left\, (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*) walther@60358: \<^rule_thm>\left_diff_distrib\, (*"(z1.0 - z2.0) * w = z1.0 * w - z2.0 * w"*) walther@60358: \<^rule_thm>\right_diff_distrib\, (*"w * (z1.0 - z2.0) = w * z1.0 - w * z2.0"*) wneuper@59523: *) walther@60358: \<^rule_thm>\mult_1_left\, (*"1 * z = z"*) walther@60358: \<^rule_thm>\mult_zero_left\, (*"0 * z = 0"*) walther@60358: \<^rule_thm>\add_0_left\, (*"0 + z = z"*) walther@60358: walther@60358: \<^rule_eval>\plus\ (eval_binop "#add_"), walther@60358: \<^rule_eval>\times\ (eval_binop "#mult_"), wenzelm@60405: \<^rule_eval>\realpow\ (eval_binop "#power_"), walther@60358: (*\<^rule_thm>\mult.commute\, walther@60358: (*AC-rewriting*) walther@60358: \<^rule_thm>\real_mult_left_commute\, walther@60358: \<^rule_thm>\mult.assoc\, walther@60358: \<^rule_thm>\add.commute\, walther@60358: \<^rule_thm>\add.left_commute\, walther@60358: \<^rule_thm>\add.assoc\, wneuper@59523: *) walther@60358: \<^rule_thm_sym>\realpow_twoI\, (*"r1 * r1 = r1 \ 2"*) walther@60358: \<^rule_thm>\realpow_plus_1\, (*"r * r \ n = r \ (n + 1)"*) walther@60358: (* walther@60358: \<^rule_thm_sym>\real_mult_2\, (*"z1 + z1 = 2 * z1"*)*) walther@60358: \<^rule_thm>\real_mult_2_assoc\, (*"z1 + (z1 + k) = 2 * z1 + k"*) walther@60358: walther@60385: \<^rule_thm>\real_num_collect\, (*"[| l is_num; m is_num |] ==>l * n + m * n = (l + m) * n"*) walther@60385: \<^rule_thm>\real_num_collect_assoc\, (*"[| l is_num; m is_num |] ==> l * n + (m * n + k) = (l + m) * n + k"*) walther@60385: \<^rule_thm>\real_one_collect\, (*"m is_num ==> n + m * n = (1 + m) * n"*) walther@60385: \<^rule_thm>\real_one_collect_assoc\, (*"m is_num ==> k + (n + m * n) = k + (1 + m) * n"*) walther@60358: walther@60358: \<^rule_eval>\plus\ (eval_binop "#add_"), walther@60358: \<^rule_eval>\times\ (eval_binop "#mult_"), wenzelm@60405: \<^rule_eval>\realpow\ (eval_binop "#power_")], walther@60358: scr = Rule.Prog (Program.prep_program @{thm expand_binoms_2.simps})}; wneuper@59523: \ wneuper@59523: walther@59887: subsection \add to Know_Store\ wneuper@59523: subsubsection \rule-sets\ walther@59618: ML \val prep_rls' = Auto_Prog.prep_rls @{theory}\ s1210629013@55444: wenzelm@60289: rule_set_knowledge walther@60387: norm_Poly = \prep_rls' norm_Poly\ and walther@60387: Poly_erls = \prep_rls' Poly_erls\ and wenzelm@60286: expand = \prep_rls' expand\ and wenzelm@60286: expand_poly = \prep_rls' expand_poly\ and wenzelm@60286: simplify_power = \prep_rls' simplify_power\ and neuper@52125: wenzelm@60286: order_add_mult = \prep_rls' order_add_mult\ and wenzelm@60286: collect_numerals = \prep_rls' collect_numerals\ and wenzelm@60286: collect_numerals_= \prep_rls' collect_numerals_\ and wenzelm@60286: reduce_012 = \prep_rls' reduce_012\ and wenzelm@60286: discard_parentheses = \prep_rls' discard_parentheses\ and neuper@52125: wenzelm@60286: make_polynomial = \prep_rls' make_polynomial\ and wenzelm@60286: expand_binoms = \prep_rls' expand_binoms\ and wenzelm@60286: rev_rew_p = \prep_rls' rev_rew_p\ and wenzelm@60286: discard_minus = \prep_rls' discard_minus\ and wenzelm@60286: expand_poly_ = \prep_rls' expand_poly_\ and neuper@52125: wenzelm@60286: expand_poly_rat_ = \prep_rls' expand_poly_rat_\ and wenzelm@60286: simplify_power_ = \prep_rls' simplify_power_\ and wenzelm@60286: calc_add_mult_pow_ = \prep_rls' calc_add_mult_pow_\ and wenzelm@60286: reduce_012_mult_ = \prep_rls' reduce_012_mult_\ and wenzelm@60286: reduce_012_ = \prep_rls' reduce_012_\ and neuper@52125: wenzelm@60286: discard_parentheses1 = \prep_rls' discard_parentheses1\ and wenzelm@60286: order_mult_rls_ = \prep_rls' order_mult_rls_\ and wenzelm@60286: order_add_rls_ = \prep_rls' order_add_rls_\ and wenzelm@60286: make_rat_poly_with_parentheses = \prep_rls' make_rat_poly_with_parentheses\ wneuper@59523: wneuper@59526: subsection \problems\ wenzelm@60306: wenzelm@60306: problem pbl_simp_poly : "polynomial/simplification" = wenzelm@60306: \Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*) wenzelm@60306: \<^rule_eval>\is_polyexp\ (eval_is_polyexp "")]\ Walther@60449: Method_Ref: "simplification/for_polynomials" wenzelm@60306: CAS: "Simplify t_t" wenzelm@60306: Given: "Term t_t" wenzelm@60306: Where: "t_t is_polyexp" wenzelm@60306: Find: "normalform n_n" wneuper@59429: wneuper@59526: subsection \methods\ wneuper@59545: wneuper@59429: partial_function (tailrec) simplify :: "real \ real" wneuper@59429: where walther@59635: "simplify term = ((Rewrite_Set ''norm_Poly'') term)" wenzelm@60303: wenzelm@60303: method met_simp_poly : "simplification/for_polynomials" = wenzelm@60303: \{rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, wenzelm@60303: prls = Rule_Set.append_rules "simplification_for_polynomials_prls" Rule_Set.empty wenzelm@60303: [(*for preds in where_*) \<^rule_eval>\is_polyexp\ (eval_is_polyexp"")], wenzelm@60303: crls = Rule_Set.empty, errpats = [], nrls = norm_Poly}\ wenzelm@60303: Program: simplify.simps wenzelm@60303: Given: "Term t_t" wenzelm@60303: Where: "t_t is_polyexp" wenzelm@60303: Find: "normalform n_n" wneuper@59472: ML \ wneuper@59472: \ ML \ Walther@60509: \ neuper@37906: end