src/Tools/isac/Knowledge/Poly.thy
author wneuper <walther.neuper@jku.at>
Mon, 19 Jul 2021 18:29:46 +0200
changeset 60337 cbad4e18e91b
parent 60335 7701598a2182
child 60341 59106f9e08cc
permissions -rw-r--r--
cleanup after "eliminate ThmC.numerals_to_Free"
     1 (* WN.020812: theorems in the Reals,
     2    necessary for special rule sets, in addition to Isabelle2002.
     3    !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
     4    !!! THIS IS THE _least_ NUMBER OF ADDITIONAL THEOREMS !!!
     5    !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
     6    xxxI contain \<up> instead of ^ in the respective theorem xxx in 2002
     7    changed by: Richard Lang 020912
     8 *)
     9 
    10 theory Poly imports Simplify begin
    11 
    12 subsection \<open>remark on term-structure of polynomials\<close>
    13 text \<open>
    14 WN190319:
    15 the code below reflects missing coordination between two authors:
    16 * ML: built the equation solver; simple rule-sets, programs; better predicates for specifications.
    17 * MG: built simplification of polynomials with AC rewriting by ML code
    18 
    19 WN020919:
    20 *** there are 5 kinds of expanded normalforms ***
    21 
    22 [1] 'complete polynomial' (Komplettes Polynom), univariate
    23    a_0 + a_1.x^1 +...+ a_n.x^n   not (a_n = 0)
    24 	        not (a_n = 0), some a_i may be zero (DON'T disappear),
    25                 variables in monomials lexicographically ordered and complete,
    26                 x written as 1*x^1, ...
    27 [2] 'polynomial' (Polynom), univariate and multivariate
    28    a_0 + a_1.x +...+ a_n.x^n   not (a_n = 0)
    29    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
    30 	        not (a_n = 0), some a_i may be zero (ie. monomials disappear),
    31                 exponents and coefficients equal 1 are not (WN060904.TODO in cancel_p_)shown,
    32                 and variables in monomials are lexicographically ordered  
    33    examples: [1]: "1 + (-10) * x \<up> 1 + 25 * x \<up> 2"
    34 	     [1]: "11 + 0 * x \<up> 1 + 1 * x \<up> 2"
    35 	     [2]: "x + (-50) * x \<up> 3"
    36 	     [2]: "(-1) * x * y \<up> 2 + 7 * x \<up> 3"
    37 
    38 [3] 'expanded_term' (Ausmultiplizierter Term):
    39    pull out unary minus to binary minus, 
    40    as frequently exercised in schools; other conditions for [2] hold however
    41    examples: "a \<up> 2 - 2 * a * b + b \<up> 2"
    42 	     "4 * x \<up> 2 - 9 * y \<up> 2"
    43 [4] 'polynomial_in' (Polynom in): 
    44    polynomial in 1 variable with arbitrary coefficients
    45    examples: "2 * x + (-50) * x \<up> 3"                     (poly in x)
    46 	     "(u + v) + (2 * u \<up> 2) * a + (-u) * a \<up> 2 (poly in a)
    47 [5] 'expanded_in' (Ausmultiplizierter Termin in): 
    48    analoguous to [3] with binary minus like [3]
    49    examples: "2 * x - 50 * x \<up> 3"                     (expanded in x)
    50 	     "(u + v) + (2 * u \<up> 2) * a - u * a \<up> 2 (expanded in a)
    51 \<close>
    52 subsection \<open>consts definition for predicates in specifications\<close>
    53 consts
    54 
    55   is_expanded_in :: "[real, real] => bool" ("_ is'_expanded'_in _") 
    56   is_poly_in     :: "[real, real] => bool" ("_ is'_poly'_in _")   (*RL DA *)
    57   has_degree_in  :: "[real, real] => real" ("_ has'_degree'_in _")(*RL DA *)
    58   is_polyrat_in  :: "[real, real] => bool" ("_ is'_polyrat'_in _")(*RL030626*)
    59 
    60   is_multUnordered:: "real => bool" ("_ is'_multUnordered") 
    61   is_addUnordered :: "real => bool" ("_ is'_addUnordered") (*WN030618*)
    62   is_polyexp      :: "real => bool" ("_ is'_polyexp") 
    63 
    64 subsection \<open>theorems not yet adopted from Isabelle\<close>
    65 axiomatization where (*.not contained in Isabelle2002,
    66          stated as axioms, TODO: prove as theorems;
    67          theorem-IDs 'xxxI' with \<up> instead of ^ in 'xxx' in Isabelle2002.*)
    68 
    69   realpow_pow:             "(a \<up> b) \<up> c = a \<up> (b * c)" and
    70   realpow_addI:            "r \<up> (n + m) = r \<up> n * r \<up> m" and
    71   realpow_addI_assoc_l:    "r \<up> n * (r \<up> m * s) = r \<up> (n + m) * s" and
    72   realpow_addI_assoc_r:    "s * r \<up> n * r \<up> m = s * r \<up> (n + m)" and
    73 		  
    74   realpow_oneI:            "r \<up> 1 = r" and
    75   realpow_zeroI:            "r \<up> 0 = 1" and
    76   realpow_eq_oneI:         "1 \<up> n = 1" and
    77   realpow_multI:           "(r * s) \<up> n = r \<up> n * s \<up> n"  and
    78   realpow_multI_poly:      "[| r is_polyexp; s is_polyexp |] ==>
    79 			      (r * s) \<up> n = r \<up> n * s \<up> n"  and
    80   realpow_minus_oneI:      "(- 1) \<up> (2 * n) = 1"  and 
    81   real_diff_0:		         "0 - x = - (x::real)" and
    82 
    83   realpow_twoI:            "r \<up> 2 = r * r" and
    84   realpow_twoI_assoc_l:	  "r * (r * s) = r \<up> 2 * s" and
    85   realpow_twoI_assoc_r:	  "s * r * r = s * r \<up> 2" and
    86   realpow_two_atom:        "r is_atom ==> r * r = r \<up> 2" and
    87   realpow_plus_1:          "r * r \<up> n = r \<up> (n + 1)"   and       
    88   realpow_plus_1_assoc_l:  "r * (r \<up> m * s) = r \<up> (1 + m) * s"  and
    89   realpow_plus_1_assoc_l2: "r \<up> m * (r * s) = r \<up> (1 + m) * s"  and
    90   realpow_plus_1_assoc_r:  "s * r * r \<up> m = s * r \<up> (1 + m)" and
    91   realpow_plus_1_atom:     "r is_atom ==> r * r \<up> n = r \<up> (1 + n)" and
    92   realpow_def_atom:        "[| Not (r is_atom); 1 < n |]
    93 			   ==> r \<up> n = r * r \<up> (n + -1)" and
    94   realpow_addI_atom:       "r is_atom ==> r \<up> n * r \<up> m = r \<up> (n + m)" and
    95 
    96 
    97   realpow_minus_even:	     "n is_even ==> (- r) \<up> n = r \<up> n" and
    98   realpow_minus_odd:       "Not (n is_even) ==> (- r) \<up> n = -1 * r \<up> n" and
    99 
   100 
   101 (* RL 020914 *)
   102   real_pp_binom_times:     "(a + b)*(c + d) = a*c + a*d + b*c + b*d" and
   103   real_pm_binom_times:     "(a + b)*(c - d) = a*c - a*d + b*c - b*d" and
   104   real_mp_binom_times:     "(a - b)*(c + d) = a*c + a*d - b*c - b*d" and
   105   real_mm_binom_times:     "(a - b)*(c - d) = a*c - a*d - b*c + b*d" and
   106   real_plus_binom_pow3:    "(a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" and
   107   real_plus_binom_pow3_poly: "[| a is_polyexp; b is_polyexp |] ==> 
   108 			    (a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" and
   109   real_minus_binom_pow3:   "(a - b) \<up> 3 = a \<up> 3 - 3*a \<up> 2*b + 3*a*b \<up> 2 - b \<up> 3" and
   110   real_minus_binom_pow3_p: "(a + -1 * b) \<up> 3 = a \<up> 3 + -3*a \<up> 2*b + 3*a*b \<up> 2 +
   111                            -1*b \<up> 3" and
   112 (* real_plus_binom_pow:        "[| n is_const;  3 < n |] ==>
   113 			       (a + b) \<up> n = (a + b) * (a + b)\<up>(n - 1)" *)
   114   real_plus_binom_pow4:   "(a + b) \<up> 4 = (a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3)
   115                            *(a + b)" and
   116   real_plus_binom_pow4_poly: "[| a is_polyexp; b is_polyexp |] ==> 
   117 			   (a + b) \<up> 4 = (a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3)
   118                            *(a + b)" and
   119   real_plus_binom_pow5:    "(a + b) \<up> 5 = (a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3)
   120                            *(a \<up> 2 + 2*a*b + b \<up> 2)" and
   121   real_plus_binom_pow5_poly: "[| a is_polyexp; b is_polyexp |] ==> 
   122 			        (a + b) \<up> 5 = (a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 
   123                                 + b \<up> 3)*(a \<up> 2 + 2*a*b + b \<up> 2)" and
   124   real_diff_plus:          "a - b = a + -b" (*17.3.03: do_NOT_use*) and
   125   real_diff_minus:         "a - b = a + -1 * b" and
   126   real_plus_binom_times:   "(a + b)*(a + b) = a \<up> 2 + 2*a*b + b \<up> 2" and
   127   real_minus_binom_times:  "(a - b)*(a - b) = a \<up> 2 - 2*a*b + b \<up> 2" and
   128   (*WN071229 changed for Schaerding -----vvv*)
   129   (*real_plus_binom_pow2:  "(a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
   130   real_plus_binom_pow2:    "(a + b) \<up> 2 = (a + b) * (a + b)" and
   131   (*WN071229 changed for Schaerding -----\<up>*)
   132   real_plus_binom_pow2_poly: "[| a is_polyexp; b is_polyexp |] ==>
   133 			       (a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2" and
   134   real_minus_binom_pow2:      "(a - b) \<up> 2 = a \<up> 2 - 2*a*b + b \<up> 2" and
   135   real_minus_binom_pow2_p:    "(a - b) \<up> 2 = a \<up> 2 + -2*a*b + b \<up> 2" and
   136   real_plus_minus_binom1:     "(a + b)*(a - b) = a \<up> 2 - b \<up> 2" and
   137   real_plus_minus_binom1_p:   "(a + b)*(a - b) = a \<up> 2 + -1*b \<up> 2" and
   138   real_plus_minus_binom1_p_p: "(a + b)*(a + -1 * b) = a \<up> 2 + -1*b \<up> 2" and
   139   real_plus_minus_binom2:     "(a - b)*(a + b) = a \<up> 2 - b \<up> 2" and
   140   real_plus_minus_binom2_p:   "(a - b)*(a + b) = a \<up> 2 + -1*b \<up> 2" and
   141   real_plus_minus_binom2_p_p: "(a + -1 * b)*(a + b) = a \<up> 2 + -1*b \<up> 2" and
   142   real_plus_binom_times1:     "(a +  1*b)*(a + -1*b) = a \<up> 2 + -1*b \<up> 2" and
   143   real_plus_binom_times2:     "(a + -1*b)*(a +  1*b) = a \<up> 2 + -1*b \<up> 2" and
   144 
   145   real_num_collect:           "[| l is_const; m is_const |] ==>
   146 			      l * n + m * n = (l + m) * n" and
   147 (* FIXME.MG.0401: replace 'real_num_collect_assoc' 
   148 	by 'real_num_collect_assoc_l' ... are equal, introduced by MG ! *)
   149   real_num_collect_assoc:     "[| l is_const; m is_const |] ==> 
   150 			      l * n + (m * n + k) = (l + m) * n + k" and
   151   real_num_collect_assoc_l:   "[| l is_const; m is_const |] ==>
   152 			      l * n + (m * n + k) = (l + m)
   153 				* n + k" and
   154   real_num_collect_assoc_r:   "[| l is_const; m is_const |] ==>
   155 			      (k + m * n) + l * n = k + (l + m) * n" and
   156   real_one_collect:           "m is_const ==> n + m * n = (1 + m) * n" and
   157 (* FIXME.MG.0401: replace 'real_one_collect_assoc' 
   158 	by 'real_one_collect_assoc_l' ... are equal, introduced by MG ! *)
   159   real_one_collect_assoc:     "m is_const ==> n + (m * n + k) = (1 + m)* n + k" and
   160 
   161   real_one_collect_assoc_l:   "m is_const ==> n + (m * n + k) = (1 + m) * n + k" and
   162   real_one_collect_assoc_r:  "m is_const ==> (k + n) +  m * n = k + (1 + m) * n" and
   163 
   164 (* FIXME.MG.0401: replace 'real_mult_2_assoc' 
   165 	by 'real_mult_2_assoc_l' ... are equal, introduced by MG ! *)
   166   real_mult_2_assoc:          "z1 + (z1 + k) = 2 * z1 + k" and
   167   real_mult_2_assoc_l:        "z1 + (z1 + k) = 2 * z1 + k" and
   168   real_mult_2_assoc_r:        "(k + z1) + z1 = k + 2 * z1" and
   169 
   170   real_mult_left_commute: "z1 * (z2 * z3) = z2 * (z1 * z3)" and
   171   real_mult_minus1:       "-1 * z = - (z::real)" and
   172   (*sym_real_mult_minus1 expands indefinitely without assumptions ...*)
   173   real_mult_minus1_sym:   "[| \<not>(matches (- 1 * x) z); \<not>(z is_atom) |] ==> - (z::real) = -1 * z" and
   174   real_mult_2:            "2 * z = z + (z::real)" and
   175 
   176   real_add_mult_distrib_poly: "w is_polyexp ==> (z1 + z2) * w = z1 * w + z2 * w" and
   177   real_add_mult_distrib2_poly:"w is_polyexp ==> w * (z1 + z2) = w * z1 + w * z2"
   178 
   179 
   180 subsection \<open>auxiliary functions\<close>
   181 ML \<open>
   182 val poly_consts =
   183   [\<^const_name>\<open>plus\<close>, \<^const_name>\<open>minus\<close>,
   184   \<^const_name>\<open>divide\<close>, \<^const_name>\<open>times\<close>,
   185   \<^const_name>\<open>powr\<close>];
   186 
   187 val int_ord_SAVE = int_ord;
   188 (*for tests on rewrite orders*)
   189 fun int_ord (i1, i2) =
   190 (@{print} {a = "int_ord (" ^ string_of_int i1 ^ ", " ^ string_of_int i2 ^ ") = ", z = Int.compare (i1, i2)};
   191   Int.compare (i1, i2));
   192 (**)val int_ord = int_ord_SAVE; (*..outcomment for tests*)
   193 \<close>
   194 subsubsection \<open>for predicates in specifications (ML)\<close>
   195 ML \<open>
   196 (*--- auxiliary for is_expanded_in, is_poly_in, has_degree_in ---*)
   197 (*. a "monomial t in variable v" is a term t with
   198   either (1) v NOT existent in t, or (2) v contained in t,
   199   if (1) then degree 0
   200   if (2) then v is a factor on the very right, casually with exponent.*)
   201 fun factor_right_deg (*case 2*)
   202 	    (Const (\<^const_name>\<open>Groups.times_class.times\<close>, _) $
   203         t1 $ (Const (\<^const_name>\<open>Transcendental.powr\<close>,_) $ vv $ num)) v =
   204 	  if vv = v andalso not (Prog_Expr.occurs_in v t1) then SOME (snd (HOLogic.dest_number num))
   205     else NONE
   206   | factor_right_deg (Const (\<^const_name>\<open>Transcendental.powr\<close>,_) $ vv $ num) v =
   207 	   if (vv = v) then SOME (snd (HOLogic.dest_number num)) else NONE
   208 
   209   | factor_right_deg (Const (\<^const_name>\<open>times\<close>,_) $ t1 $ vv) v = 
   210 	   if vv = v andalso not (Prog_Expr.occurs_in v t1) then SOME 1 else NONE
   211   | factor_right_deg vv v =
   212 	  if (vv = v) then SOME 1 else NONE;    
   213 fun mono_deg_in m v =  (*case 1*)
   214 	if not (Prog_Expr.occurs_in v m) then (*case 1*) SOME 0 else factor_right_deg m v;
   215 
   216 fun expand_deg_in t v =
   217 	let
   218     fun edi ~1 ~1 (Const (\<^const_name>\<open>plus\<close>, _) $ t1 $ t2) =
   219           (case mono_deg_in t2 v of (* $ is left associative*)
   220             SOME d' => edi d' d' t1 | NONE => NONE)
   221       | edi ~1 ~1 (Const (\<^const_name>\<open>minus\<close>, _) $ t1 $ t2) =
   222           (case mono_deg_in t2 v of
   223             SOME d' => edi d' d' t1 | NONE => NONE)
   224       | edi d dmax (Const (\<^const_name>\<open>minus\<close>, _) $ t1 $ t2) =
   225           (case mono_deg_in t2 v of (*(d = 0 andalso d' = 0) handle 3+4-...4 +x*)
   226 	        SOME d' => if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
   227           | NONE => NONE)
   228       | edi d dmax (Const (\<^const_name>\<open>plus\<close>,_) $ t1 $ t2) =
   229           (case mono_deg_in t2 v of
   230             SOME d' =>    (*RL (d = 0 andalso d' = 0) need to handle 3+4-...4 +x*)
   231               if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
   232           | NONE => NONE)
   233       | edi ~1 ~1 t =
   234           (case mono_deg_in t v of d as SOME _ => d | NONE => NONE)
   235       | edi d dmax t = (*basecase last*)
   236     	    (case mono_deg_in t v of
   237     	      SOME d' => if d > d' orelse (d = 0 andalso d' = 0) then SOME dmax else NONE
   238 		      | NONE => NONE)
   239 	in edi ~1 ~1 t end;
   240 
   241 fun poly_deg_in t v =
   242 	let
   243     fun edi ~1 ~1 (Const (\<^const_name>\<open>plus\<close>,_) $ t1 $ t2) =
   244 		    (case mono_deg_in t2 v of (* $ is left associative *)
   245 		      SOME d' => edi d' d' t1
   246         | NONE => NONE)
   247 	    | edi d dmax (Const (\<^const_name>\<open>plus\<close>,_) $ t1 $ t2) =
   248 		    (case mono_deg_in t2 v of
   249 	        SOME d' =>    (*RL (d = 0 andalso (d' = 0)) handle 3+4-...4 +x*)
   250             if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
   251         | NONE => NONE)
   252 	    | edi ~1 ~1 t =
   253         (case mono_deg_in t v of
   254 		      d as SOME _ => d
   255         | NONE => NONE)
   256 	    | edi d dmax t = (* basecase last *)
   257 		    (case mono_deg_in t v of
   258 		      SOME d' =>
   259             if d > d' orelse (d = 0 andalso d' = 0) then SOME dmax else NONE
   260         | NONE => NONE)
   261 	in edi ~1 ~1 t end;
   262 \<close>
   263 
   264 subsubsection \<open>for hard-coded AC rewriting (MG)\<close>
   265 ML \<open>
   266 (**. MG.03: make_polynomial_ ... uses SML-fun for ordering .**)
   267 
   268 (*FIXME.0401: make SML-order local to make_polynomial(_) *)
   269 (*FIXME.0401: replace 'make_polynomial'(old) by 'make_polynomial_'(MG) *)
   270 (* Polynom --> List von Monomen *) 
   271 fun poly2list (Const (\<^const_name>\<open>plus\<close>,_) $ t1 $ t2) = 
   272     (poly2list t1) @ (poly2list t2)
   273   | poly2list t = [t];
   274 
   275 (* Monom --> Liste von Variablen *)
   276 fun monom2list (Const (\<^const_name>\<open>times\<close>,_) $ t1 $ t2) = 
   277     (monom2list t1) @ (monom2list t2)
   278   | monom2list t = [t];
   279 
   280 (* liefert Variablenname (String) einer Variablen und Basis bei Potenz *)
   281 fun get_basStr (Const (\<^const_name>\<open>powr\<close>,_) $ Free (str, _) $ _) = str
   282   | get_basStr (Const (\<^const_name>\<open>Transcendental.powr\<close>,_) $ n $ _) = TermC.to_string n
   283   | get_basStr (Free (str, _)) = str
   284   | get_basStr t =
   285     if TermC.is_num t then TermC.to_string t
   286     else "|||"; (* gross gewichtet; für Brüche ect. *)
   287 
   288 (* liefert Hochzahl (String) einer Variablen bzw Gewichtstring (zum Sortieren) *)
   289 fun get_potStr (Const (\<^const_name>\<open>Transcendental.powr\<close>, _) $ Free _ $ Free (str, _)) = str
   290   | get_potStr (Const (\<^const_name>\<open>Transcendental.powr\<close>, _) $ Free _ $ t) =
   291     if TermC.is_num t then TermC.to_string t else "|||"
   292   | get_potStr (Free _) = "---" (* keine Hochzahl --> kleinst gewichtet *)
   293   | get_potStr _ = "||||||"; (* gross gewichtet; für Brüch ect. *)
   294 
   295 (* Umgekehrte string_ord *)
   296 val string_ord_rev =  rev_order o string_ord;
   297 		
   298  (* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen) 
   299     innerhalb eines Monomes:
   300     - zuerst lexikographisch nach Variablenname 
   301     - wenn gleich: nach steigender Potenz *)
   302 fun var_ord (a, b) = 
   303 (@{print} {a = "var_ord ", a_b = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")",
   304    sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
   305   prod_ord string_ord string_ord 
   306   ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b))
   307 );
   308 fun var_ord (a,b: term) = 
   309   prod_ord string_ord string_ord 
   310     ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
   311 
   312 (* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen); 
   313    verwendet zum Sortieren von Monomen mittels Gesamtgradordnung:
   314    - zuerst lexikographisch nach Variablenname 
   315    - wenn gleich: nach sinkender Potenz*)
   316 fun var_ord_revPow (a, b: term) = 
   317 (@{print} {a = "var_ord_revPow ", at_bt = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")",
   318     sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
   319   prod_ord string_ord string_ord_rev 
   320     ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b))
   321 );
   322 fun var_ord_revPow (a, b: term) =
   323   prod_ord string_ord string_ord_rev 
   324     ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
   325 
   326 
   327 (* Ordnet ein Liste von Variablen (und Potenzen) lexikographisch *)
   328 fun sort_varList ts =
   329 (@{print} {a = "sort_varList", args = UnparseC.terms ts};
   330   sort var_ord ts);
   331 val sort_varList = sort var_ord;
   332 
   333 (* Entfernet aeussersten Operator (Wurzel) aus einem Term und schreibt 
   334    Argumente in eine Liste *)
   335 fun args u : term list =
   336   let
   337     fun stripc (f $ t, ts) = stripc (f, t::ts)
   338   	  | stripc (t as Free _, ts) = (t::ts)
   339   	  | stripc (_, ts) = ts
   340   in stripc (u, []) end;
   341                                     
   342 (* liefert True, falls der Term (Liste von Termen) nur Zahlen 
   343    (keine Variablen) enthaelt *)
   344 fun filter_num ts = fold (curry and_) (map TermC.is_num ts) true
   345 
   346 (* liefert True, falls der Term nur Zahlen (keine Variablen) enthaelt 
   347    dh. er ist ein numerischer Wert und entspricht einem Koeffizienten *)
   348 fun is_nums t = filter_num [t];
   349 
   350 (* Berechnet den Gesamtgrad eines Monoms *)
   351 (**)local(**)
   352   fun counter (n, []) = n
   353     | counter (n, x :: xs) = 
   354 	    if (is_nums x) then counter (n, xs)
   355 	    else 
   356 	      (case x of 
   357 		      (Const (\<^const_name>\<open>Transcendental.powr\<close>, _) $ Free _ $ t) =>
   358             if TermC.is_num t
   359             then counter (t |> HOLogic.dest_number |> snd |> curry op + n, xs)
   360             else counter (n + 1000, xs) (*FIXME.MG?!*)
   361 	      | (Const (\<^const_name>\<open>numeral\<close>, _) $ num) =>
   362             counter (n + 1 + HOLogic.dest_numeral num, xs)
   363 	      | _ => counter (n + 1, xs)) (*FIXME.MG?! ... Brüche ect.*)
   364 (**)in(**)
   365   fun monom_degree l = counter (0, l) 
   366 (**)end;(*local*)
   367 
   368 (* wie Ordnung dict_ord (lexicographische Ordnung zweier Listen, mit Vergleich 
   369    der Listen-Elemente mit elem_ord) - Elemente die Bedingung cond erfuellen, 
   370    werden jedoch dabei ignoriert (uebersprungen)  *)
   371 fun dict_cond_ord _ _ ([], [])     = (@{print} {a = "dict_cond_ord ([], [])"}; EQUAL)
   372   | dict_cond_ord _ _ ([], _ :: _) = (@{print} {a = "dict_cond_ord ([], _ :: _)"}; LESS)
   373   | dict_cond_ord _ _ (_ :: _, []) = (@{print} {a = "dict_cond_ord (_ :: _, [])"}; GREATER)
   374   | dict_cond_ord elem_ord cond (x :: xs, y :: ys) =
   375     (@{print} {a = "dict_cond_ord", args = "(" ^ UnparseC.terms (x :: xs) ^ ", " ^ UnparseC.terms (y :: ys) ^ ")", 
   376       is_nums = "(" ^ LibraryC.bool2str (cond x) ^ ", " ^ LibraryC.bool2str (cond y) ^ ")"};
   377      case (cond x, cond y) of 
   378 	    (false, false) =>
   379         (case elem_ord (x, y) of 
   380 				  EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
   381 			  | ord => ord)
   382     | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
   383     | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
   384     | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys) );
   385 fun dict_cond_ord _ _ ([], []) = EQUAL
   386   | dict_cond_ord _ _ ([], _ :: _) = LESS
   387   | dict_cond_ord _ _ (_ :: _, []) = GREATER
   388   | dict_cond_ord elem_ord cond (x :: xs, y :: ys) =
   389     (case (cond x, cond y) of 
   390 	    (false, false) =>
   391         (case elem_ord (x, y) of 
   392 				  EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
   393 			  | ord => ord)
   394     | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
   395     | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
   396     | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys) );
   397 
   398 (* Gesamtgradordnung zum Vergleich von Monomen (Liste von Variablen/Potenzen):
   399    zuerst nach Gesamtgrad, bei gleichem Gesamtgrad lexikographisch ordnen - 
   400    dabei werden Koeffizienten ignoriert (2*3*a \<up> 2*4*b gilt wie a \<up> 2*b) *)
   401 fun degree_ord (xs, ys) =
   402 	prod_ord int_ord (dict_cond_ord var_ord_revPow is_nums)
   403 	  ((monom_degree xs, xs), (monom_degree ys, ys));
   404 
   405 fun hd_str str = substring (str, 0, 1);
   406 fun tl_str str = substring (str, 1, (size str) - 1);
   407 
   408 (* liefert nummerischen Koeffizienten eines Monoms oder NONE *)
   409 fun get_koeff_of_mon [] = raise ERROR "get_koeff_of_mon: called with l = []"
   410   | get_koeff_of_mon (x :: _) = if is_nums x then SOME x else NONE;
   411 
   412 (* wandelt Koeffizient in (zum sortieren geeigneten) String um *)
   413 fun koeff2ordStr (SOME t) =
   414     if TermC.is_num t
   415     then 
   416       if (t |> HOLogic.dest_number |> snd) < 0
   417       then (t |> HOLogic.dest_number |> snd |> curry op * ~1 |> string_of_int) ^ "0"  (* 3 < -3 *)
   418       else (t |> HOLogic.dest_number |> snd |> string_of_int)
   419     else "aaa"                                                      (* "num.Ausdruck" --> gross *)
   420   | koeff2ordStr NONE = "---";                                     (* "kein Koeff" --> kleinste *)
   421 
   422 (* Order zum Vergleich von Koeffizienten (strings): 
   423    "kein Koeff" < "0" < "1" < "-1" < "2" < "-2" < ... < "num.Ausdruck" *)
   424 fun compare_koeff_ord (xs, ys) = string_ord
   425   ((koeff2ordStr o get_koeff_of_mon) xs,
   426    (koeff2ordStr o get_koeff_of_mon) ys);
   427 
   428 (* Gesamtgradordnung degree_ord + Ordnen nach Koeffizienten falls EQUAL *)
   429 fun koeff_degree_ord (xs, ys) =
   430 	    prod_ord degree_ord compare_koeff_ord ((xs, xs), (ys, ys));
   431 
   432 (* Ordnet ein Liste von Monomen (Monom = Liste von Variablen) mittels 
   433    Gesamtgradordnung *)
   434 val sort_monList = sort koeff_degree_ord;
   435 
   436 (* Alternativ zu degree_ord koennte auch die viel einfachere und 
   437    kuerzere Ordnung simple_ord verwendet werden - ist aber nicht 
   438    fuer unsere Zwecke geeignet!
   439 
   440 fun simple_ord (al,bl: term list) = dict_ord string_ord 
   441 	 (map get_basStr al, map get_basStr bl); 
   442 
   443 val sort_monList = sort simple_ord; *)
   444 
   445 (* aus 2 Variablen wird eine Summe bzw ein Produkt erzeugt 
   446    (mit gewuenschtem Typen T) *)
   447 fun plus T = Const (\<^const_name>\<open>plus\<close>, [T,T] ---> T);
   448 fun mult T = Const (\<^const_name>\<open>times\<close>, [T,T] ---> T);
   449 fun binop op_ t1 t2 = op_ $ t1 $ t2;
   450 fun create_prod T (a,b) = binop (mult T) a b;
   451 fun create_sum T (a,b) = binop (plus T) a b;
   452 
   453 (* löscht letztes Element einer Liste *)
   454 fun drop_last l = take ((length l)-1,l);
   455 
   456 (* Liste von Variablen --> Monom *)
   457 fun create_monom T vl = foldr (create_prod T) (drop_last vl, last_elem vl);
   458 (* Bemerkung: 
   459    foldr bewirkt rechtslastige Klammerung des Monoms - ist notwendig, damit zwei 
   460    gleiche Monome zusammengefasst werden können (collect_numerals)! 
   461    zB: 2*(x*(y*z)) + 3*(x*(y*z)) --> (2+3)*(x*(y*z))*)
   462 
   463 (* Liste von Monomen --> Polynom *)	
   464 fun create_polynom T ml = foldl (create_sum T) (hd ml, tl ml);
   465 (* Bemerkung: 
   466    foldl bewirkt linkslastige Klammerung des Polynoms (der Summanten) - 
   467    bessere Darstellung, da keine Klammern sichtbar! 
   468    (und discard_parentheses in make_polynomial hat weniger zu tun) *)
   469 
   470 (* sorts the variables (faktors) of an expanded polynomial lexicographical *)
   471 fun sort_variables t = 
   472   let
   473   	val ll = map monom2list (poly2list t);
   474   	val lls = map sort_varList ll; 
   475   	val T = type_of t;
   476   	val ls = map (create_monom T) lls;
   477   in create_polynom T ls end;
   478 
   479 (* sorts the monoms of an expanded and variable-sorted polynomial 
   480    by total_degree *)
   481 fun sort_monoms t = 
   482   let
   483   	val ll =  map monom2list (poly2list t);
   484   	val lls = sort_monList ll;
   485   	val T = Term.type_of t;
   486   	val ls = map (create_monom T) lls;
   487   in create_polynom T ls end;
   488 \<close>
   489 
   490 subsubsection \<open>rewrite order for hard-coded AC rewriting\<close>
   491 ML \<open>
   492 local (*. for make_polynomial .*)
   493 
   494 open Term;  (* for type order = EQUAL | LESS | GREATER *)
   495 
   496 fun pr_ord EQUAL = "EQUAL"
   497   | pr_ord LESS  = "LESS"
   498   | pr_ord GREATER = "GREATER";
   499 
   500 fun dest_hd' (Const (a, T)) =                          (* ~ term.ML *)
   501   (case a of
   502      \<^const_name>\<open>powr\<close> => ((("|||||||||||||", 0), T), 0)    (*WN greatest string*)
   503    | _ => (((a, 0), T), 0))
   504   | dest_hd' (Free (a, T)) = (((a, 0), T), 1)(*TODOO handle this as numeral, too? see EqSystem.thy*)
   505   | dest_hd' (Var v) = (v, 2)
   506   | dest_hd' (Bound i) = ((("", i), dummyT), 3)
   507   | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4)
   508   | dest_hd' t = raise TERM ("dest_hd'", [t]);
   509 
   510 fun size_of_term' (Const(str,_) $ t) =
   511   if \<^const_name>\<open>powr\<close>= str then 1000 + size_of_term' t else 1+size_of_term' t(*WN*)
   512   | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body
   513   | size_of_term' (f$t) = size_of_term' f  +  size_of_term' t
   514   | size_of_term' _ = 1;
   515 
   516 fun term_ord' pr thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
   517     (case term_ord' pr thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord)
   518   | term_ord' pr thy (t, u) =
   519     (if pr then 
   520 	   let
   521        val (f, ts) = strip_comb t and (g, us) = strip_comb u;
   522        val _ = tracing ("t= f@ts= \"" ^ UnparseC.term_in_thy thy f ^ "\" @ \"[" ^
   523          commas (map (UnparseC.term_in_thy thy) ts) ^ "]\"");
   524        val _ = tracing("u= g@us= \"" ^ UnparseC.term_in_thy thy g ^ "\" @ \"[" ^
   525          commas (map (UnparseC.term_in_thy thy) us) ^ "]\"");
   526        val _ = tracing ("size_of_term(t,u)= (" ^ string_of_int (size_of_term' t) ^ ", " ^
   527          string_of_int (size_of_term' u) ^ ")");
   528        val _ = tracing ("hd_ord(f,g)      = " ^ (pr_ord o hd_ord) (f,g));
   529        val _ = tracing ("terms_ord(ts,us) = " ^ (pr_ord o terms_ord str false) (ts, us));
   530        val _ = tracing ("-------");
   531      in () end
   532        else ();
   533 	 case int_ord (size_of_term' t, size_of_term' u) of
   534 	   EQUAL =>
   535 	     let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
   536 	       (case hd_ord (f, g) of EQUAL => (terms_ord str pr) (ts, us) 
   537 	     | ord => ord)
   538 	     end
   539 	 | ord => ord)
   540 and hd_ord (f, g) =                                        (* ~ term.ML *)
   541   prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord (dest_hd' f, dest_hd' g)
   542 and terms_ord _ pr (ts, us) = 
   543     list_ord (term_ord' pr (ThyC.get_theory "Isac_Knowledge"))(ts, us);
   544 
   545 in
   546 
   547 fun ord_make_polynomial (pr:bool) thy (_: subst) (ts, us) =
   548     (term_ord' pr thy(***) (TermC.numerals_to_Free ts, TermC.numerals_to_Free us) = LESS );
   549 
   550 end;(*local*)
   551 
   552 Rewrite_Ord.rew_ord' := overwritel (! Rewrite_Ord.rew_ord', (* TODO: make analogous to KEStore_Elems.add_mets *)
   553 [("termlessI", termlessI), ("ord_make_polynomial", ord_make_polynomial false \<^theory>)]);
   554 \<close>
   555 
   556 subsection \<open>predicates\<close>
   557 subsubsection \<open>in specifications\<close>
   558 ML \<open>
   559 (* is_polyrat_in becomes true, if no bdv is in the denominator of a fraction*)
   560 fun is_polyrat_in t v = 
   561   let
   562    	fun finddivide (_ $ _ $ _ $ _) _ = raise ERROR("is_polyrat_in:")
   563 	    (* at the moment there is no term like this, but ....*)
   564 	  | finddivide (Const (\<^const_name>\<open>divide\<close>,_) $ _ $ b) v = not (Prog_Expr.occurs_in v b)
   565 	  | finddivide (_ $ t1 $ t2) v = finddivide t1 v orelse finddivide t2 v
   566 	  | finddivide (_ $ t1) v = finddivide t1 v
   567 	  | finddivide _ _ = false;
   568   in finddivide t v end;
   569     
   570 fun is_expanded_in t v = case expand_deg_in t v of SOME _ => true | NONE => false;
   571 fun is_poly_in t v =     case poly_deg_in t v of SOME _ => true | NONE => false;
   572 fun has_degree_in t v =  case expand_deg_in t v of SOME d => d | NONE => ~1;
   573 
   574 (*.the expression contains + - * ^ only ?
   575    this is weaker than 'is_polynomial' !.*)
   576 fun is_polyexp (Free _) = true
   577   | is_polyexp (Const _) = true (* potential danger: bdv is not considered *)
   578   | is_polyexp (Const (\<^const_name>\<open>plus\<close>,_) $ Free _ $ num) =
   579     if TermC.is_num num then true
   580     else if TermC.is_variable num then true
   581     else is_polyexp num
   582   | is_polyexp (Const (\<^const_name>\<open>plus\<close>, _) $ num $ Free _) =
   583     if TermC.is_num num then true
   584     else if TermC.is_variable num then true
   585     else is_polyexp num
   586   | is_polyexp (Const (\<^const_name>\<open>minus\<close>, _) $ Free _ $ num) =
   587     if TermC.is_num num then true
   588     else if TermC.is_variable num then true
   589     else is_polyexp num
   590   | is_polyexp (Const (\<^const_name>\<open>times\<close>, _) $ num $ Free _) =
   591     if TermC.is_num num then true
   592     else if TermC.is_variable num then true
   593     else is_polyexp num
   594   | is_polyexp (Const (\<^const_name>\<open>Transcendental.powr\<close>,_) $ Free _ $ num) =
   595     if TermC.is_num num then true
   596     else if TermC.is_variable num then true
   597     else is_polyexp num
   598   | is_polyexp (Const (\<^const_name>\<open>plus_class.plus\<close>,_) $ t1 $ t2) = 
   599     ((is_polyexp t1) andalso (is_polyexp t2))
   600   | is_polyexp (Const (\<^const_name>\<open>Groups.minus_class.minus\<close>,_) $ t1 $ t2) = 
   601     ((is_polyexp t1) andalso (is_polyexp t2))
   602   | is_polyexp (Const (\<^const_name>\<open>Groups.times_class.times\<close>,_) $ t1 $ t2) = 
   603     ((is_polyexp t1) andalso (is_polyexp t2))
   604   | is_polyexp (Const (\<^const_name>\<open>Transcendental.powr\<close>,_) $ t1 $ t2) = 
   605     ((is_polyexp t1) andalso (is_polyexp t2))
   606   | is_polyexp num = TermC.is_num num;
   607 \<close>
   608 
   609 subsubsection \<open>for hard-coded AC rewriting\<close>
   610 ML \<open>
   611 (* auch Klammerung muss übereinstimmen;
   612    sort_variables klammert Produkte rechtslastig*)
   613 fun is_multUnordered t = ((is_polyexp t) andalso not (t = sort_variables t));
   614 
   615 fun is_addUnordered t = ((is_polyexp t) andalso not (t = sort_monoms t));
   616 \<close>
   617 
   618 subsection \<open>evaluations functions\<close>
   619 subsubsection \<open>for predicates\<close>
   620 ML \<open>
   621 fun eval_is_polyrat_in _ _(p as (Const (\<^const_name>\<open>Poly.is_polyrat_in\<close>, _) $ t $ v)) _  =
   622     if is_polyrat_in t v 
   623     then SOME ((UnparseC.term p) ^ " = True",
   624 	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
   625     else SOME ((UnparseC.term p) ^ " = True",
   626 	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
   627   | eval_is_polyrat_in _ _ _ _ = ((*tracing"### no matches";*) NONE);
   628 
   629 (*("is_expanded_in", ("Poly.is_expanded_in", eval_is_expanded_in ""))*)
   630 fun eval_is_expanded_in _ _ 
   631        (p as (Const (\<^const_name>\<open>Poly.is_expanded_in\<close>, _) $ t $ v)) _ =
   632     if is_expanded_in t v
   633     then SOME ((UnparseC.term p) ^ " = True",
   634 	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
   635     else SOME ((UnparseC.term p) ^ " = True",
   636 	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
   637   | eval_is_expanded_in _ _ _ _ = NONE;
   638 
   639 (*("is_poly_in", ("Poly.is_poly_in", eval_is_poly_in ""))*)
   640 fun eval_is_poly_in _ _ 
   641        (p as (Const (\<^const_name>\<open>Poly.is_poly_in\<close>, _) $ t $ v)) _ =
   642     if is_poly_in t v
   643     then SOME ((UnparseC.term p) ^ " = True",
   644 	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
   645     else SOME ((UnparseC.term p) ^ " = True",
   646 	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
   647   | eval_is_poly_in _ _ _ _ = NONE;
   648 
   649 (*("has_degree_in", ("Poly.has_degree_in", eval_has_degree_in ""))*)
   650 fun eval_has_degree_in _ _ 
   651 	     (p as (Const (\<^const_name>\<open>Poly.has_degree_in\<close>, _) $ t $ v)) _ =
   652     let val d = has_degree_in t v
   653 	val d' = TermC.term_of_num HOLogic.realT d
   654     in SOME ((UnparseC.term p) ^ " = " ^ (string_of_int d),
   655 	      HOLogic.Trueprop $ (TermC.mk_equality (p, d')))
   656     end
   657   | eval_has_degree_in _ _ _ _ = NONE;
   658 
   659 (*("is_polyexp", ("Poly.is_polyexp", eval_is_polyexp ""))*)
   660 fun eval_is_polyexp (thmid:string) _ 
   661 		       (t as (Const (\<^const_name>\<open>Poly.is_polyexp\<close>, _) $ arg)) thy = 
   662     if is_polyexp arg
   663     then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
   664 	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
   665     else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
   666 	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
   667   | eval_is_polyexp _ _ _ _ = NONE; 
   668 \<close>
   669 
   670 subsubsection \<open>for hard-coded AC rewriting\<close>
   671 ML \<open>
   672 (*WN.18.6.03 *)
   673 (*("is_addUnordered", ("Poly.is_addUnordered", eval_is_addUnordered ""))*)
   674 fun eval_is_addUnordered (thmid:string) _ 
   675 		       (t as (Const (\<^const_name>\<open>Poly.is_addUnordered\<close>, _) $ arg)) thy = 
   676     if is_addUnordered arg
   677     then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
   678 	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
   679     else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
   680 	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
   681   | eval_is_addUnordered _ _ _ _ = NONE; 
   682 
   683 fun eval_is_multUnordered (thmid:string) _ 
   684 		       (t as (Const (\<^const_name>\<open>Poly.is_multUnordered\<close>, _) $ arg)) thy = 
   685     if is_multUnordered arg
   686     then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
   687 	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
   688     else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
   689 	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
   690   | eval_is_multUnordered _ _ _ _ = NONE; 
   691 \<close>
   692 calculation is_polyrat_in = \<open>eval_is_polyrat_in "#eval_is_polyrat_in"\<close>
   693 calculation is_expanded_in = \<open>eval_is_expanded_in ""\<close>
   694 calculation is_poly_in = \<open>eval_is_poly_in ""\<close>
   695 calculation has_degree_in = \<open>eval_has_degree_in ""\<close>
   696 calculation is_polyexp = \<open>eval_is_polyexp ""\<close>
   697 calculation is_multUnordered = \<open>eval_is_multUnordered ""\<close>
   698 calculation is_addUnordered = \<open>eval_is_addUnordered ""\<close>
   699 
   700 subsection \<open>rule-sets\<close>
   701 subsubsection \<open>without specific order\<close>
   702 ML \<open>
   703 (* used only for merge *)
   704 val calculate_Poly = Rule_Set.append_rules "calculate_PolyFIXXXME.not.impl." Rule_Set.empty [];
   705 
   706 (*.for evaluation of conditions in rewrite rules.*)
   707 val Poly_erls = Rule_Set.append_rules "Poly_erls" Atools_erls
   708   [\<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
   709   \<^rule_thm>\<open>real_unari_minus\<close>,
   710   \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"),
   711   \<^rule_eval>\<open>minus\<close> (eval_binop "#sub_"),
   712   \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
   713   \<^rule_eval>\<open>powr\<close> (eval_binop "#power_")];
   714 
   715 val poly_crls = Rule_Set.append_rules "poly_crls" Atools_crls
   716   [\<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
   717   \<^rule_thm>\<open>real_unari_minus\<close>,
   718   \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"),
   719   \<^rule_eval>\<open>minus\<close> (eval_binop "#sub_"),
   720   \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
   721   \<^rule_eval>\<open>powr\<close> (eval_binop "#power_")];
   722 \<close>
   723 ML \<open>
   724 val expand =
   725   Rule_Def.Repeat {id = "expand", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   726       erls = Rule_Set.empty,srls = Rule_Set.Empty, calc = [], errpatts = [],
   727       rules = [\<^rule_thm>\<open>distrib_right\<close>,
   728 	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
   729 	       \<^rule_thm>\<open>distrib_left\<close>
   730 	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
   731 	       ], scr = Rule.Empty_Prog};
   732 
   733 (* erls for calculate_Rational + etc *)
   734 val powers_erls =
   735   Rule_Def.Repeat {id = "powers_erls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
   736       erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   737       rules = 
   738         [Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches "#matches_"),
   739 	       Rule.Eval ("Prog_Expr.is_atom", Prog_Expr.eval_is_atom "#is_atom_"),
   740 	       Rule.Eval ("Prog_Expr.is_even", Prog_Expr.eval_is_even "#is_even_"),
   741 	       Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
   742 	       Rule.Thm ("not_false",  @{thm not_false}),
   743 	       Rule.Thm ("not_true",  @{thm not_true}),
   744 	       Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_")
   745 	       ],
   746       scr = Rule.Empty_Prog
   747       };
   748 
   749 \<close> ML \<open>
   750 val discard_minus =
   751   Rule_Def.Repeat {id = "discard_minus", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   752       erls = powers_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   753       rules =
   754        [\<^rule_thm>\<open>real_diff_minus\<close>,
   755           (*"a - b = a + -1 * b"*)
   756         Rule.Thm ("real_mult_minus1_sym",  (@{thm real_mult_minus1_sym}))
   757 	        (*"\<not>(z is_const) ==> - (z::real) = -1 * z"*)],
   758 	      scr = Rule.Empty_Prog};
   759 
   760 val expand_poly_ = 
   761   Rule_Def.Repeat{id = "expand_poly_", preconds = [], 
   762       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   763       erls = powers_erls, srls = Rule_Set.Empty,
   764       calc = [], errpatts = [],
   765       rules =
   766         [\<^rule_thm>\<open>real_plus_binom_pow4\<close>,
   767 	           (*"(a + b) \<up> 4 = ... "*)
   768 	         \<^rule_thm>\<open>real_plus_binom_pow5\<close>,
   769 	           (*"(a + b) \<up> 5 = ... "*)
   770 	         \<^rule_thm>\<open>real_plus_binom_pow3\<close>,
   771 	           (*"(a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" *)
   772 	         (*WN071229 changed/removed for Schaerding -----vvv*)
   773 	         (*\<^rule_thm>\<open>real_plus_binom_pow2\<close>,*)
   774 	           (*"(a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
   775 	         \<^rule_thm>\<open>real_plus_binom_pow2\<close>,
   776 	           (*"(a + b) \<up> 2 = (a + b) * (a + b)"*)
   777 	         (*\<^rule_thm>\<open>real_plus_minus_binom1_p_p\<close>,*)
   778 	           (*"(a + b)*(a + -1 * b) = a \<up> 2 + -1*b \<up> 2"*)
   779 	         (*\<^rule_thm>\<open>real_plus_minus_binom2_p_p\<close>,*)
   780 	           (*"(a + -1 * b)*(a + b) = a \<up> 2 + -1*b \<up> 2"*)
   781 	         (*WN071229 changed/removed for Schaerding -----\<up>*)
   782 	      
   783 	         \<^rule_thm>\<open>distrib_right\<close>,
   784 	           (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
   785 	         \<^rule_thm>\<open>distrib_left\<close>,
   786 	           (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
   787 	       
   788 	         \<^rule_thm>\<open>realpow_multI\<close>,
   789 	           (*"(r * s) \<up> n = r \<up> n * s \<up> n"*)
   790 	         \<^rule_thm>\<open>realpow_pow\<close>,
   791 	           (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
   792 (**)
   793 	         Rule.Thm ("realpow_minus_even", @{thm realpow_minus_even}),
   794 	           (*"n is_even ==> (- r) \<up> n = r \<up> n"*)
   795 	         Rule.Thm ("realpow_minus_odd", @{thm realpow_minus_odd})
   796 	           (*"Not (n is_even) ==> (- r) \<up> n = -1 * r \<up> n"*)
   797 (**)
   798 	       ], scr = Rule.Empty_Prog};
   799 
   800 val expand_poly_rat_ = 
   801   Rule_Def.Repeat{id = "expand_poly_rat_", preconds = [], 
   802       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   803       erls =  Rule_Set.append_rules "Rule_Set.empty-expand_poly_rat_" Rule_Set.empty
   804 	        [Rule.Eval ("Poly.is_polyexp", eval_is_polyexp ""),
   805 	         Rule.Eval ("Prog_Expr.is_even", Prog_Expr.eval_is_even "#is_even_"),
   806 	         Rule.Thm ("not_false",  @{thm not_false}),
   807 	         Rule.Thm ("not_true",  @{thm not_true})
   808 		      ],
   809       srls = Rule_Set.Empty,
   810       calc = [], errpatts = [],
   811       rules = 
   812         [\<^rule_thm>\<open>real_plus_binom_pow4_poly\<close>,
   813 	     (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \<up> 4 = ... "*)
   814 	 \<^rule_thm>\<open>real_plus_binom_pow5_poly\<close>,
   815 	     (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \<up> 5 = ... "*)
   816 	 \<^rule_thm>\<open>real_plus_binom_pow2_poly\<close>,
   817 	     (*"[| a is_polyexp; b is_polyexp |] ==>
   818 		            (a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
   819 	 \<^rule_thm>\<open>real_plus_binom_pow3_poly\<close>,
   820 	     (*"[| a is_polyexp; b is_polyexp |] ==> 
   821 			(a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" *)
   822 	 \<^rule_thm>\<open>real_plus_minus_binom1_p_p\<close>,
   823 	     (*"(a + b)*(a + -1 * b) = a \<up> 2 + -1*b \<up> 2"*)
   824 	 \<^rule_thm>\<open>real_plus_minus_binom2_p_p\<close>,
   825 	     (*"(a + -1 * b)*(a + b) = a \<up> 2 + -1*b \<up> 2"*)
   826 	      
   827 	 \<^rule_thm>\<open>real_add_mult_distrib_poly\<close>,
   828 	       (*"w is_polyexp ==> (z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
   829 	 \<^rule_thm>\<open>real_add_mult_distrib2_poly\<close>,
   830 	     (*"w is_polyexp ==> w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
   831 	       
   832 	 \<^rule_thm>\<open>realpow_multI_poly\<close>,
   833 	     (*"[| r is_polyexp; s is_polyexp |] ==> 
   834 		            (r * s) \<up> n = r \<up> n * s \<up> n"*)
   835 	 Rule.Thm ("realpow_pow", @{thm realpow_pow}),
   836 	   (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
   837 	 Rule.Thm ("realpow_minus_even", @{thm realpow_minus_even}),
   838 	   (*"n is_even ==> (- r) \<up> n = r \<up> n"*)
   839 	 Rule.Thm ("realpow_minus_odd", @{thm realpow_minus_odd})
   840 	   (*"\<not> (n is_even) ==> (- r) \<up> n = -1 * r \<up> n"*)
   841 	 ], scr = Rule.Empty_Prog};
   842 
   843 val simplify_power_ = 
   844   Rule_Def.Repeat{id = "simplify_power_", preconds = [], 
   845       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   846       erls = Rule_Set.empty, srls = Rule_Set.Empty,
   847       calc = [], errpatts = [],
   848       rules = [(*MG: Reihenfolge der folgenden 2 Rule.Thm muss so bleiben, wegen
   849 		a*(a*a) --> a*a \<up> 2 und nicht a*(a*a) --> a \<up> 2*a *)
   850 	       \<^rule_thm_sym>\<open>realpow_twoI\<close>,	
   851 	       (*"r * r = r \<up> 2"*)
   852 	       \<^rule_thm>\<open>realpow_twoI_assoc_l\<close>,
   853 	       (*"r * (r * s) = r \<up> 2 * s"*)
   854 
   855 	       \<^rule_thm>\<open>realpow_plus_1\<close>,		
   856 	       (*"r * r \<up> n = r \<up> (n + 1)"*)
   857 	       \<^rule_thm>\<open>realpow_plus_1_assoc_l\<close>,
   858 	       (*"r * (r \<up> m * s) = r \<up> (1 + m) * s"*)
   859 	       (*MG 9.7.03: neues Rule.Thm wegen a*(a*(a*b)) --> a \<up> 2*(a*b) *)
   860 	       \<^rule_thm>\<open>realpow_plus_1_assoc_l2\<close>,
   861 	       (*"r \<up> m * (r * s) = r \<up> (1 + m) * s"*)
   862 
   863 	       \<^rule_thm_sym>\<open>realpow_addI\<close>,
   864 	       (*"r \<up> n * r \<up> m = r \<up> (n + m)"*)
   865 	       \<^rule_thm>\<open>realpow_addI_assoc_l\<close>,
   866 	       (*"r \<up> n * (r \<up> m * s) = r \<up> (n + m) * s"*)
   867 	       
   868 	       (* ist in expand_poly - wird hier aber auch gebraucht, wegen: 
   869 		  "r * r = r \<up> 2" wenn r=a \<up> b*)
   870 	       \<^rule_thm>\<open>realpow_pow\<close>
   871 	       (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
   872 	       ], scr = Rule.Empty_Prog};
   873 
   874 val calc_add_mult_pow_ = 
   875   Rule_Def.Repeat{id = "calc_add_mult_pow_", preconds = [], 
   876       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   877       erls = Atools_erls(*erls3.4.03*),srls = Rule_Set.Empty,
   878       calc = [("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_")), 
   879 	      ("TIMES" , (\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
   880 	      ("POWER", (\<^const_name>\<open>powr\<close>, eval_binop "#power_"))
   881 	      ],
   882       errpatts = [],
   883       rules = [\<^rule_eval>\<open>plus\<close> (eval_binop "#add_"),
   884 	       \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
   885 	       \<^rule_eval>\<open>powr\<close> (eval_binop "#power_")
   886 	       ], scr = Rule.Empty_Prog};
   887 
   888 val reduce_012_mult_ = 
   889   Rule_Def.Repeat{id = "reduce_012_mult_", preconds = [], 
   890       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   891       erls = Rule_Set.empty,srls = Rule_Set.Empty,
   892       calc = [], errpatts = [],
   893       rules = [(* MG: folgende Rule.Thm müssen hier stehen bleiben: *)
   894                \<^rule_thm>\<open>mult_1_right\<close>,
   895 	       (*"z * 1 = z"*) (*wegen "a * b * b \<up> (-1) + a"*) 
   896 	       \<^rule_thm>\<open>realpow_zeroI\<close>,
   897 	       (*"r \<up> 0 = 1"*) (*wegen "a*a \<up> (-1)*c + b + c"*)
   898 	       \<^rule_thm>\<open>realpow_oneI\<close>,
   899 	       (*"r \<up> 1 = r"*)
   900 	       \<^rule_thm>\<open>realpow_eq_oneI\<close>
   901 	       (*"1 \<up> n = 1"*)
   902 	       ], scr = Rule.Empty_Prog};
   903 
   904 val collect_numerals_ = 
   905   Rule_Def.Repeat{id = "collect_numerals_", preconds = [], 
   906       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   907       erls = Atools_erls, srls = Rule_Set.Empty,
   908       calc = [("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_"))
   909 	      ], errpatts = [],
   910       rules = 
   911         [\<^rule_thm>\<open>real_num_collect\<close>, 
   912 	     (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
   913 	 \<^rule_thm>\<open>real_num_collect_assoc_r\<close>,
   914 	     (*"[| l is_const; m is_const |] ==>  \
   915 					\(k + m * n) + l * n = k + (l + m)*n"*)
   916 	 \<^rule_thm>\<open>real_one_collect\<close>,	
   917 	     (*"m is_const ==> n + m * n = (1 + m) * n"*)
   918 	 \<^rule_thm>\<open>real_one_collect_assoc_r\<close>, 
   919 	     (*"m is_const ==> (k + n) + m * n = k + (m + 1) * n"*)
   920 
   921          \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"),
   922 
   923 	 (*MG: Reihenfolge der folgenden 2 Rule.Thm muss so bleiben, wegen
   924 		     (a+a)+a --> a + 2*a --> 3*a and not (a+a)+a --> 2*a + a *)
   925          \<^rule_thm>\<open>real_mult_2_assoc_r\<close>,
   926 	     (*"(k + z1) + z1 = k + 2 * z1"*)
   927 	 \<^rule_thm_sym>\<open>real_mult_2\<close>
   928 	     (*"z1 + z1 = 2 * z1"*)
   929 	], scr = Rule.Empty_Prog};
   930 
   931 val reduce_012_ = 
   932   Rule_Def.Repeat{id = "reduce_012_", preconds = [], 
   933       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   934       erls = Rule_Set.empty,srls = Rule_Set.Empty, calc = [], errpatts = [],
   935       rules = [\<^rule_thm>\<open>mult_1_left\<close>,                 
   936 	       (*"1 * z = z"*)
   937 	       \<^rule_thm>\<open>mult_zero_left\<close>,        
   938 	       (*"0 * z = 0"*)
   939 	       \<^rule_thm>\<open>mult_zero_right\<close>,
   940 	       (*"z * 0 = 0"*)
   941 	       \<^rule_thm>\<open>add_0_left\<close>,
   942 	       (*"0 + z = z"*)
   943 	       \<^rule_thm>\<open>add_0_right\<close>,
   944 	       (*"z + 0 = z"*) (*wegen a+b-b --> a+(1-1)*b --> a+0 --> a*)
   945 
   946 	       (*\<^rule_thm>\<open>realpow_oneI\<close>*)
   947 	       (*"?r \<up> 1 = ?r"*)
   948 	       \<^rule_thm>\<open>division_ring_divide_zero\<close>
   949 	       (*"0 / ?x = 0"*)
   950 	       ], scr = Rule.Empty_Prog};
   951 
   952 val discard_parentheses1 = 
   953     Rule_Set.append_rules "discard_parentheses1" Rule_Set.empty 
   954 	       [\<^rule_thm_sym>\<open>mult.assoc\<close>
   955 		(*"?z1.1 * (?z2.1 * ?z3.1) = ?z1.1 * ?z2.1 * ?z3.1"*)
   956 		(*\<^rule_thm_sym>\<open>add.assoc\<close>*)
   957 		(*"?z1.1 + (?z2.1 + ?z3.1) = ?z1.1 + ?z2.1 + ?z3.1"*)
   958 		 ];
   959 
   960 val expand_poly =
   961   Rule_Def.Repeat{id = "expand_poly", preconds = [], 
   962       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   963       erls = powers_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   964       rules = 
   965         [Rule.Thm ("distrib_right" , @{thm distrib_right}),
   966 	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
   967 	       \<^rule_thm>\<open>distrib_left\<close>,
   968 	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
   969 	       (*Rule.Thm ("distrib_right1", @{thm distrib_right}1),
   970 		....... 18.3.03 undefined???*)
   971 
   972 	       \<^rule_thm>\<open>real_plus_binom_pow2\<close>,
   973 	       (*"(a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
   974 	       \<^rule_thm>\<open>real_minus_binom_pow2_p\<close>,
   975 	       (*"(a - b) \<up> 2 = a \<up> 2 + -2*a*b + b \<up> 2"*)
   976 	       \<^rule_thm>\<open>real_plus_minus_binom1_p\<close>,
   977 	       (*"(a + b)*(a - b) = a \<up> 2 + -1*b \<up> 2"*)
   978 	       \<^rule_thm>\<open>real_plus_minus_binom2_p\<close>,
   979 	       (*"(a - b)*(a + b) = a \<up> 2 + -1*b \<up> 2"*)
   980 
   981 	       \<^rule_thm>\<open>minus_minus\<close>,
   982 	       (*"- (- ?z) = ?z"*)
   983 	       \<^rule_thm>\<open>real_diff_minus\<close>,
   984 	       (*"a - b = a + -1 * b"*)
   985 	       Rule.Thm ("real_mult_minus1_sym",  (@{thm real_mult_minus1_sym}))
   986 	       (*"\<not>(z is_const) ==> - (z::real) = -1 * z"*)
   987 
   988 	       (*\<^rule_thm>\<open>real_minus_add_distrib\<close>,*)
   989 	       (*"- (?x + ?y) = - ?x + - ?y"*)
   990 	       (*\<^rule_thm>\<open>real_diff_plus\<close>*)
   991 	       (*"a - b = a + -b"*)
   992 	       ], scr = Rule.Empty_Prog};
   993 
   994 val simplify_power = 
   995   Rule_Def.Repeat{id = "simplify_power", preconds = [], 
   996       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   997       erls = Rule_Set.empty, srls = Rule_Set.Empty,
   998       calc = [], errpatts = [],
   999       rules = [\<^rule_thm>\<open>realpow_multI\<close>,
  1000 	       (*"(r * s) \<up> n = r \<up> n * s \<up> n"*)
  1001 	       
  1002 	       \<^rule_thm_sym>\<open>realpow_twoI\<close>,	
  1003 	       (*"r1 * r1 = r1 \<up> 2"*)
  1004 	       \<^rule_thm>\<open>realpow_plus_1\<close>,		
  1005 	       (*"r * r \<up> n = r \<up> (n + 1)"*)
  1006 	       \<^rule_thm>\<open>realpow_pow\<close>,
  1007 	       (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
  1008 	       \<^rule_thm_sym>\<open>realpow_addI\<close>,
  1009 	       (*"r \<up> n * r \<up> m = r \<up> (n + m)"*)
  1010 	       \<^rule_thm>\<open>realpow_oneI\<close>,
  1011 	       (*"r \<up> 1 = r"*)
  1012 	       \<^rule_thm>\<open>realpow_eq_oneI\<close>
  1013 	       (*"1 \<up> n = 1"*)
  1014 	       ], scr = Rule.Empty_Prog};
  1015 
  1016 val collect_numerals = 
  1017   Rule_Def.Repeat{id = "collect_numerals", preconds = [], 
  1018       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
  1019       erls = Atools_erls(*erls3.4.03*),srls = Rule_Set.Empty,
  1020       calc = [("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_")), 
  1021 	      ("TIMES" , (\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
  1022 	      ("POWER", (\<^const_name>\<open>powr\<close>, eval_binop "#power_"))
  1023 	      ], errpatts = [],
  1024       rules = [\<^rule_thm>\<open>real_num_collect\<close>, 
  1025 	       (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
  1026 	       \<^rule_thm>\<open>real_num_collect_assoc\<close>,
  1027 	       (*"[| l is_const; m is_const |] ==>  
  1028 				l * n + (m * n + k) =  (l + m) * n + k"*)
  1029 	       \<^rule_thm>\<open>real_one_collect\<close>,	
  1030 	       (*"m is_const ==> n + m * n = (1 + m) * n"*)
  1031 	       \<^rule_thm>\<open>real_one_collect_assoc\<close>, 
  1032 	       (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
  1033 	       \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"), 
  1034 	       \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
  1035 	       \<^rule_eval>\<open>powr\<close> (eval_binop "#power_")
  1036 	       ], scr = Rule.Empty_Prog};
  1037 val reduce_012 = 
  1038   Rule_Def.Repeat{id = "reduce_012", preconds = [], 
  1039       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
  1040       erls = Rule_Set.empty,srls = Rule_Set.Empty,
  1041       calc = [], errpatts = [],
  1042       rules = [\<^rule_thm>\<open>mult_1_left\<close>,                 
  1043 	       (*"1 * z = z"*)
  1044 	       (*\<^rule_thm>\<open>real_mult_minus1\<close>,14.3.03*)
  1045 	       (*"-1 * z = - z"*)
  1046 	       Rule.Thm ("minus_mult_left",  (@{thm minus_mult_left} RS @{thm sym})),
  1047 	       (*- (?x * ?y) = "- ?x * ?y"*)
  1048 	       (*\<^rule_thm>\<open>real_minus_mult_cancel\<close>,
  1049 	       (*"- ?x * - ?y = ?x * ?y"*)---*)
  1050 	       \<^rule_thm>\<open>mult_zero_left\<close>,        
  1051 	       (*"0 * z = 0"*)
  1052 	       \<^rule_thm>\<open>add_0_left\<close>,
  1053 	       (*"0 + z = z"*)
  1054 	       \<^rule_thm>\<open>right_minus\<close>,
  1055 	       (*"?z + - ?z = 0"*)
  1056 	       \<^rule_thm_sym>\<open>real_mult_2\<close>,	
  1057 	       (*"z1 + z1 = 2 * z1"*)
  1058 	       \<^rule_thm>\<open>real_mult_2_assoc\<close>
  1059 	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
  1060 	       ], scr = Rule.Empty_Prog};
  1061 
  1062 val discard_parentheses = 
  1063     Rule_Set.append_rules "discard_parentheses" Rule_Set.empty 
  1064 	       [\<^rule_thm_sym>\<open>mult.assoc\<close>,	\<^rule_thm_sym>\<open>add.assoc\<close>];
  1065 \<close>
  1066 
  1067 subsubsection \<open>hard-coded AC rewriting\<close>
  1068 ML \<open>
  1069 (*MG.0401: termorders for multivariate polys dropped due to principal problems:
  1070   (total-degree-)ordering of monoms NOT possible with size_of_term GIVEN*)
  1071 val order_add_mult = 
  1072   Rule_Def.Repeat{id = "order_add_mult", preconds = [], 
  1073       rew_ord = ("ord_make_polynomial",ord_make_polynomial false \<^theory>),
  1074       erls = Rule_Set.empty,srls = Rule_Set.Empty,
  1075       calc = [], errpatts = [],
  1076       rules = [\<^rule_thm>\<open>mult.commute\<close>,
  1077 	       (* z * w = w * z *)
  1078 	       \<^rule_thm>\<open>real_mult_left_commute\<close>,
  1079 	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
  1080 	       \<^rule_thm>\<open>mult.assoc\<close>,		
  1081 	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
  1082 	       \<^rule_thm>\<open>add.commute\<close>,	
  1083 	       (*z + w = w + z*)
  1084 	       \<^rule_thm>\<open>add.left_commute\<close>,
  1085 	       (*x + (y + z) = y + (x + z)*)
  1086 	       \<^rule_thm>\<open>add.assoc\<close>	               
  1087 	       (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)
  1088 	       ], scr = Rule.Empty_Prog};
  1089 (*MG.0401: termorders for multivariate polys dropped due to principal problems:
  1090   (total-degree-)ordering of monoms NOT possible with size_of_term GIVEN*)
  1091 val order_mult = 
  1092   Rule_Def.Repeat{id = "order_mult", preconds = [], 
  1093       rew_ord = ("ord_make_polynomial",ord_make_polynomial false \<^theory>),
  1094       erls = Rule_Set.empty,srls = Rule_Set.Empty,
  1095       calc = [], errpatts = [],
  1096       rules = [\<^rule_thm>\<open>mult.commute\<close>,
  1097 	       (* z * w = w * z *)
  1098 	       \<^rule_thm>\<open>real_mult_left_commute\<close>,
  1099 	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
  1100 	       \<^rule_thm>\<open>mult.assoc\<close>	
  1101 	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
  1102 	       ], scr = Rule.Empty_Prog};
  1103 \<close>
  1104 ML \<open>
  1105 fun attach_form (_: Rule.rule list list) (_: term) (_: term) = (*still missing*)
  1106     []:(Rule.rule * (term * term list)) list;
  1107 fun init_state (_: term) = Rule_Set.e_rrlsstate;
  1108 fun locate_rule (_: Rule.rule list list) (_: term) (_: Rule.rule) =
  1109     ([]:(Rule.rule * (term * term list)) list);
  1110 fun next_rule (_: Rule.rule list list) (_: term) = (NONE: Rule.rule option);
  1111 fun normal_form t = SOME (sort_variables t, []: term list);
  1112 
  1113 val order_mult_ =
  1114     Rule_Set.Rrls {id = "order_mult_", 
  1115 	  prepat = 
  1116           (* ?p matched with the current term gives an environment,
  1117              which evaluates (the instantiated) "?p is_multUnordered" to true *)
  1118 	  [([TermC.parse_patt \<^theory> "?p is_multUnordered"], 
  1119              TermC.parse_patt \<^theory> "?p :: real")],
  1120 	  rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
  1121 	  erls = Rule_Set.append_rules "Rule_Set.empty-is_multUnordered" Rule_Set.empty
  1122 			    [\<^rule_eval>\<open>is_multUnordered\<close> (eval_is_multUnordered "")],
  1123 	  calc = [("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_")),
  1124 		  ("TIMES" , (\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
  1125 		  ("DIVIDE", (\<^const_name>\<open>divide\<close>, Prog_Expr.eval_cancel "#divide_e")),
  1126 		  ("POWER" , (\<^const_name>\<open>powr\<close>, eval_binop "#power_"))],
  1127     errpatts = [],
  1128 	  scr = Rule.Rfuns {init_state  = init_state,
  1129 		     normal_form = normal_form,
  1130 		     locate_rule = locate_rule,
  1131 		     next_rule   = next_rule,
  1132 		     attach_form = attach_form}};
  1133 val order_mult_rls_ = 
  1134   Rule_Def.Repeat {id = "order_mult_rls_", preconds = [], 
  1135       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
  1136       erls = Rule_Set.empty,srls = Rule_Set.Empty,
  1137       calc = [], errpatts = [],
  1138       rules = [Rule.Rls_ order_mult_
  1139 	       ], scr = Rule.Empty_Prog};
  1140 
  1141 \<close> ML \<open>
  1142 
  1143 fun attach_form (_: Rule.rule list list) (_: term) (_: term) = (*still missing*)
  1144     []: (Rule.rule * (term * term list)) list;
  1145 fun init_state (_: term) = Rule_Set.e_rrlsstate;
  1146 fun locate_rule (_: Rule.rule list list) (_: term) (_: Rule.rule) =
  1147     ([]: (Rule.rule * (term * term list)) list);
  1148 fun next_rule (_: Rule.rule list list) (_: term) = (NONE: Rule.rule option);
  1149 fun normal_form t = SOME (sort_monoms t,[]: term list);
  1150 \<close> ML \<open>
  1151 val order_add_ =
  1152     Rule_Set.Rrls {id = "order_add_", 
  1153 	  prepat = (*WN.18.6.03 Preconditions und Pattern,
  1154 		    die beide passen muessen, damit das Rule_Set.Rrls angewandt wird*)
  1155 	  [([TermC.parse_patt @{theory} "?p is_addUnordered"], 
  1156 	     TermC.parse_patt @{theory} "?p :: real" 
  1157 	    (*WN.18.6.03 also KEIN pattern, dieses erzeugt nur das Environment 
  1158 	      fuer die Evaluation der Precondition "p is_addUnordered"*))],
  1159 	  rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
  1160 	  erls = Rule_Set.append_rules "Rule_Set.empty-is_addUnordered" Rule_Set.empty(*MG: poly_erls*)
  1161 			    [\<^rule_eval>\<open>is_addUnordered\<close> (eval_is_addUnordered "")],
  1162 	  calc = [("PLUS"  ,(\<^const_name>\<open>plus\<close>, eval_binop "#add_")),
  1163 		  ("TIMES" ,(\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
  1164 		  ("DIVIDE",(\<^const_name>\<open>divide\<close>, Prog_Expr.eval_cancel "#divide_e")),
  1165 		  ("POWER" ,(\<^const_name>\<open>powr\<close>  , eval_binop "#power_"))],
  1166 	  errpatts = [],
  1167 	  scr = Rule.Rfuns {init_state  = init_state,
  1168 		     normal_form = normal_form,
  1169 		     locate_rule = locate_rule,
  1170 		     next_rule   = next_rule,
  1171 		     attach_form = attach_form}};
  1172 
  1173 val order_add_rls_ =
  1174   Rule_Def.Repeat {id = "order_add_rls_", preconds = [], 
  1175       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
  1176       erls = Rule_Set.empty,srls = Rule_Set.Empty,
  1177       calc = [], errpatts = [],
  1178       rules = [Rule.Rls_ order_add_
  1179 	       ], scr = Rule.Empty_Prog};
  1180 \<close>
  1181 
  1182 text \<open>rule-set make_polynomial also named norm_Poly:
  1183   Rewrite order has not been implemented properly; the order is better in 
  1184   make_polynomial_in (coded in SML).
  1185   Notes on state of development:
  1186   \# surprise 2006: test --- norm_Poly NOT COMPLETE ---
  1187   \# migration Isabelle2002 --> 2011 weakened the rule set, see test
  1188   --- Matthias Goldgruber 2003 rewrite orders ---, raise ERROR "ord_make_polynomial_in #16b"
  1189 \<close>
  1190 ML \<open>
  1191 (*. see MG-DA.p.52ff .*)
  1192 val make_polynomial(*MG.03, overwrites version from above, 
  1193     previously 'make_polynomial_'*) =
  1194   Rule_Set.Sequence {id = "make_polynomial", preconds = []:term list, 
  1195       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
  1196       erls = Atools_erls, srls = Rule_Set.Empty,calc = [], errpatts = [],
  1197       rules = [Rule.Rls_ discard_minus,
  1198 	       Rule.Rls_ expand_poly_,
  1199 	       \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
  1200 	       Rule.Rls_ order_mult_rls_,
  1201 	       Rule.Rls_ simplify_power_, 
  1202 	       Rule.Rls_ calc_add_mult_pow_, 
  1203 	       Rule.Rls_ reduce_012_mult_,
  1204 	       Rule.Rls_ order_add_rls_,
  1205 	       Rule.Rls_ collect_numerals_, 
  1206 	       Rule.Rls_ reduce_012_,
  1207 	       Rule.Rls_ discard_parentheses1
  1208 	       ],
  1209       scr = Rule.Empty_Prog
  1210       };
  1211 \<close>
  1212 ML \<open>
  1213 val norm_Poly(*=make_polynomial*) = 
  1214   Rule_Set.Sequence {id = "norm_Poly", preconds = []:term list, 
  1215       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
  1216       erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
  1217       rules = [Rule.Rls_ discard_minus,
  1218 	       Rule.Rls_ expand_poly_,
  1219 	       \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
  1220 	       Rule.Rls_ order_mult_rls_,
  1221 	       Rule.Rls_ simplify_power_, 
  1222 	       Rule.Rls_ calc_add_mult_pow_, 
  1223 	       Rule.Rls_ reduce_012_mult_,
  1224 	       Rule.Rls_ order_add_rls_,
  1225 	       Rule.Rls_ collect_numerals_, 
  1226 	       Rule.Rls_ reduce_012_,
  1227 	       Rule.Rls_ discard_parentheses1
  1228 	       ],
  1229       scr = Rule.Empty_Prog
  1230       };
  1231 \<close>
  1232 ML \<open>
  1233 (* MG:03 Like make_polynomial_ but without Rule.Rls_ discard_parentheses1 
  1234    and expand_poly_rat_ instead of expand_poly_, see MG-DA.p.56ff*)
  1235 (* MG necessary  for termination of norm_Rational(*_mg*) in Rational.ML*)
  1236 val make_rat_poly_with_parentheses =
  1237   Rule_Set.Sequence{id = "make_rat_poly_with_parentheses", preconds = []:term list, 
  1238       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
  1239       erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
  1240       rules = [Rule.Rls_ discard_minus,
  1241 	       Rule.Rls_ expand_poly_rat_,(*ignors rationals*)
  1242 	       \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
  1243 	       Rule.Rls_ order_mult_rls_,
  1244 	       Rule.Rls_ simplify_power_, 
  1245 	       Rule.Rls_ calc_add_mult_pow_, 
  1246 	       Rule.Rls_ reduce_012_mult_,
  1247 	       Rule.Rls_ order_add_rls_,
  1248 	       Rule.Rls_ collect_numerals_, 
  1249 	       Rule.Rls_ reduce_012_
  1250 	       (*Rule.Rls_ discard_parentheses1 *)
  1251 	       ],
  1252       scr = Rule.Empty_Prog
  1253       };
  1254 \<close>
  1255 ML \<open>
  1256 (*.a minimal ruleset for reverse rewriting of factions [2];
  1257    compare expand_binoms.*)
  1258 val rev_rew_p = 
  1259 Rule_Set.Sequence{id = "rev_rew_p", preconds = [], rew_ord = ("termlessI",termlessI),
  1260     erls = Atools_erls, srls = Rule_Set.Empty,
  1261     calc = [(*("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_")), 
  1262 	    ("TIMES" , (\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
  1263 	    ("POWER", (\<^const_name>\<open>powr\<close>, eval_binop "#power_"))*)
  1264 	    ], errpatts = [],
  1265     rules = [\<^rule_thm>\<open>real_plus_binom_times\<close>,
  1266 	     (*"(a + b)*(a + b) = a ^ 2 + 2 * a * b + b ^ 2*)
  1267 	     \<^rule_thm>\<open>real_plus_binom_times1\<close>,
  1268 	     (*"(a +  1*b)*(a + -1*b) = a \<up> 2 + -1*b \<up> 2"*)
  1269 	     \<^rule_thm>\<open>real_plus_binom_times2\<close>,
  1270 	     (*"(a + -1*b)*(a +  1*b) = a \<up> 2 + -1*b \<up> 2"*)
  1271 
  1272 	     \<^rule_thm>\<open>mult_1_left\<close>,(*"1 * z = z"*)
  1273 
  1274              \<^rule_thm>\<open>distrib_right\<close>,
  1275 	     (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
  1276 	     \<^rule_thm>\<open>distrib_left\<close>,
  1277 	     (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
  1278 	       
  1279 	     \<^rule_thm>\<open>mult.assoc\<close>,
  1280 	     (*"?z1.1 * ?z2.1 * ?z3. =1 ?z1.1 * (?z2.1 * ?z3.1)"*)
  1281 	     Rule.Rls_ order_mult_rls_,
  1282 	     (*Rule.Rls_ order_add_rls_,*)
  1283 
  1284 	     \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"), 
  1285 	     \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
  1286 	     \<^rule_eval>\<open>powr\<close> (eval_binop "#power_"),
  1287 	     
  1288 	     \<^rule_thm_sym>\<open>realpow_twoI\<close>,
  1289 	     (*"r1 * r1 = r1 \<up> 2"*)
  1290 	     \<^rule_thm_sym>\<open>real_mult_2\<close>,
  1291 	     (*"z1 + z1 = 2 * z1"*)
  1292 	     \<^rule_thm>\<open>real_mult_2_assoc\<close>,
  1293 	     (*"z1 + (z1 + k) = 2 * z1 + k"*)
  1294 
  1295 	     \<^rule_thm>\<open>real_num_collect\<close>, 
  1296 	     (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
  1297 	     \<^rule_thm>\<open>real_num_collect_assoc\<close>,
  1298 	     (*"[| l is_const; m is_const |] ==>  
  1299                                      l * n + (m * n + k) =  (l + m) * n + k"*)
  1300 	     \<^rule_thm>\<open>real_one_collect\<close>,
  1301 	     (*"m is_const ==> n + m * n = (1 + m) * n"*)
  1302 	     \<^rule_thm>\<open>real_one_collect_assoc\<close>, 
  1303 	     (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
  1304 
  1305 	     \<^rule_thm>\<open>realpow_multI\<close>,
  1306 	     (*"(r * s) \<up> n = r \<up> n * s \<up> n"*)
  1307 
  1308 	     \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"),
  1309 	     \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
  1310 	     \<^rule_eval>\<open>powr\<close> (eval_binop "#power_"),
  1311 
  1312 	     \<^rule_thm>\<open>mult_1_left\<close>,(*"1 * z = z"*)
  1313 	     \<^rule_thm>\<open>mult_zero_left\<close>,(*"0 * z = 0"*)
  1314 	     \<^rule_thm>\<open>add_0_left\<close>(*0 + z = z*)
  1315 
  1316 	     (*Rule.Rls_ order_add_rls_*)
  1317 	     ],
  1318 
  1319     scr = Rule.Empty_Prog};      
  1320 \<close>
  1321 
  1322 subsection \<open>rule-sets with explicit program for intermediate steps\<close>
  1323 partial_function (tailrec) expand_binoms_2 :: "real \<Rightarrow> real"
  1324   where
  1325 "expand_binoms_2 term = (
  1326   Repeat (
  1327     (Try (Repeat (Rewrite ''real_plus_binom_pow2''))) #>
  1328     (Try (Repeat (Rewrite ''real_plus_binom_times''))) #>
  1329     (Try (Repeat (Rewrite ''real_minus_binom_pow2''))) #>
  1330     (Try (Repeat (Rewrite ''real_minus_binom_times''))) #>
  1331     (Try (Repeat (Rewrite ''real_plus_minus_binom1''))) #>
  1332     (Try (Repeat (Rewrite ''real_plus_minus_binom2''))) #>
  1333    
  1334     (Try (Repeat (Rewrite ''mult_1_left''))) #>
  1335     (Try (Repeat (Rewrite ''mult_zero_left''))) #>
  1336     (Try (Repeat (Rewrite ''add_0_left''))) #>
  1337    
  1338     (Try (Repeat (Calculate ''PLUS''))) #>
  1339     (Try (Repeat (Calculate ''TIMES''))) #>
  1340     (Try (Repeat (Calculate ''POWER''))) #>
  1341    
  1342     (Try (Repeat (Rewrite ''sym_realpow_twoI''))) #>
  1343     (Try (Repeat (Rewrite ''realpow_plus_1''))) #>
  1344     (Try (Repeat (Rewrite ''sym_real_mult_2''))) #>
  1345     (Try (Repeat (Rewrite ''real_mult_2_assoc''))) #>
  1346    
  1347     (Try (Repeat (Rewrite ''real_num_collect''))) #>
  1348     (Try (Repeat (Rewrite ''real_num_collect_assoc''))) #>
  1349    
  1350     (Try (Repeat (Rewrite ''real_one_collect''))) #>
  1351     (Try (Repeat (Rewrite ''real_one_collect_assoc''))) #>
  1352    
  1353     (Try (Repeat (Calculate ''PLUS''))) #>
  1354     (Try (Repeat (Calculate ''TIMES''))) #>
  1355     (Try (Repeat (Calculate ''POWER''))))
  1356   term)"
  1357 ML \<open>
  1358 val expand_binoms = 
  1359   Rule_Def.Repeat{id = "expand_binoms", preconds = [], rew_ord = ("termlessI",termlessI),
  1360       erls = Atools_erls, srls = Rule_Set.Empty,
  1361       calc = [("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_")), 
  1362 	      ("TIMES" , (\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
  1363 	      ("POWER", (\<^const_name>\<open>powr\<close>, eval_binop "#power_"))
  1364 	      ], errpatts = [],
  1365       rules = [\<^rule_thm>\<open>real_plus_binom_pow2\<close>,     
  1366 	       (*"(a + b) \<up> 2 = a \<up> 2 + 2 * a * b + b \<up> 2"*)
  1367 	       \<^rule_thm>\<open>real_plus_binom_times\<close>,    
  1368 	      (*"(a + b)*(a + b) = ...*)
  1369 	       \<^rule_thm>\<open>real_minus_binom_pow2\<close>,   
  1370 	       (*"(a - b) \<up> 2 = a \<up> 2 - 2 * a * b + b \<up> 2"*)
  1371 	       \<^rule_thm>\<open>real_minus_binom_times\<close>,   
  1372 	       (*"(a - b)*(a - b) = ...*)
  1373 	       \<^rule_thm>\<open>real_plus_minus_binom1\<close>,   
  1374 		(*"(a + b) * (a - b) = a \<up> 2 - b \<up> 2"*)
  1375 	       \<^rule_thm>\<open>real_plus_minus_binom2\<close>,   
  1376 		(*"(a - b) * (a + b) = a \<up> 2 - b \<up> 2"*)
  1377 	       (*RL 020915*)
  1378 	       \<^rule_thm>\<open>real_pp_binom_times\<close>, 
  1379 		(*(a + b)*(c + d) = a*c + a*d + b*c + b*d*)
  1380                \<^rule_thm>\<open>real_pm_binom_times\<close>, 
  1381 		(*(a + b)*(c - d) = a*c - a*d + b*c - b*d*)
  1382                \<^rule_thm>\<open>real_mp_binom_times\<close>, 
  1383 		(*(a - b)*(c + d) = a*c + a*d - b*c - b*d*)
  1384                \<^rule_thm>\<open>real_mm_binom_times\<close>, 
  1385 		(*(a - b)*(c - d) = a*c - a*d - b*c + b*d*)
  1386 	       \<^rule_thm>\<open>realpow_multI\<close>,
  1387 		(*(a*b) \<up> n = a \<up> n * b \<up> n*)
  1388 	       \<^rule_thm>\<open>real_plus_binom_pow3\<close>,
  1389 	        (* (a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3 *)
  1390 	       \<^rule_thm>\<open>real_minus_binom_pow3\<close>,
  1391 	        (* (a - b) \<up> 3 = a \<up> 3 - 3*a \<up> 2*b + 3*a*b \<up> 2 - b \<up> 3 *)
  1392 
  1393 
  1394               (*\<^rule_thm>\<open>distrib_right\<close>,	
  1395 		(*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
  1396 	       \<^rule_thm>\<open>distrib_left\<close>,	
  1397 	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
  1398 	       \<^rule_thm>\<open>left_diff_distrib\<close>,	
  1399 	       (*"(z1.0 - z2.0) * w = z1.0 * w - z2.0 * w"*)
  1400 	       \<^rule_thm>\<open>right_diff_distrib\<close>,	
  1401 	       (*"w * (z1.0 - z2.0) = w * z1.0 - w * z2.0"*)
  1402 	      *)
  1403 	       \<^rule_thm>\<open>mult_1_left\<close>,
  1404                (*"1 * z = z"*)
  1405 	       \<^rule_thm>\<open>mult_zero_left\<close>,
  1406                (*"0 * z = 0"*)
  1407 	       \<^rule_thm>\<open>add_0_left\<close>,(*"0 + z = z"*)
  1408 
  1409 	       \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"), 
  1410 	       \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
  1411 	       \<^rule_eval>\<open>powr\<close> (eval_binop "#power_"),
  1412               (*\<^rule_thm>\<open>mult.commute\<close>,
  1413 		(*AC-rewriting*)
  1414 	       \<^rule_thm>\<open>real_mult_left_commute\<close>,
  1415 	       \<^rule_thm>\<open>mult.assoc\<close>,
  1416 	       \<^rule_thm>\<open>add.commute\<close>,
  1417 	       \<^rule_thm>\<open>add.left_commute\<close>,
  1418 	       \<^rule_thm>\<open>add.assoc\<close>,
  1419 	      *)
  1420 	       \<^rule_thm_sym>\<open>realpow_twoI\<close>,
  1421 	       (*"r1 * r1 = r1 \<up> 2"*)
  1422 	       \<^rule_thm>\<open>realpow_plus_1\<close>,			
  1423 	       (*"r * r \<up> n = r \<up> (n + 1)"*)
  1424 	       (*\<^rule_thm_sym>\<open>real_mult_2\<close>,		
  1425 	       (*"z1 + z1 = 2 * z1"*)*)
  1426 	       \<^rule_thm>\<open>real_mult_2_assoc\<close>,		
  1427 	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
  1428 
  1429 	       \<^rule_thm>\<open>real_num_collect\<close>, 
  1430 	       (*"[| l is_const; m is_const |] ==>l * n + m * n = (l + m) * n"*)
  1431 	       \<^rule_thm>\<open>real_num_collect_assoc\<close>,	
  1432 	       (*"[| l is_const; m is_const |] ==>  
  1433                                        l * n + (m * n + k) =  (l + m) * n + k"*)
  1434 	       \<^rule_thm>\<open>real_one_collect\<close>,
  1435 	       (*"m is_const ==> n + m * n = (1 + m) * n"*)
  1436 	       \<^rule_thm>\<open>real_one_collect_assoc\<close>, 
  1437 	       (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
  1438 
  1439 	       \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"), 
  1440 	       \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
  1441 	       \<^rule_eval>\<open>powr\<close> (eval_binop "#power_")
  1442 	       ],
  1443       scr = Rule.Prog (Program.prep_program @{thm expand_binoms_2.simps})
  1444       };      
  1445 \<close>
  1446 
  1447 subsection \<open>add to Know_Store\<close>
  1448 subsubsection \<open>rule-sets\<close>
  1449 ML \<open>val prep_rls' = Auto_Prog.prep_rls @{theory}\<close>
  1450 
  1451 rule_set_knowledge
  1452   norm_Poly = \<open>prep_rls' norm_Poly\<close> and
  1453   Poly_erls = \<open>prep_rls' Poly_erls\<close> (*FIXXXME:del with rls.rls'*) and
  1454   expand = \<open>prep_rls' expand\<close> and
  1455   expand_poly = \<open>prep_rls' expand_poly\<close> and
  1456   simplify_power = \<open>prep_rls' simplify_power\<close> and
  1457 
  1458   order_add_mult = \<open>prep_rls' order_add_mult\<close> and
  1459   collect_numerals = \<open>prep_rls' collect_numerals\<close> and
  1460   collect_numerals_= \<open>prep_rls' collect_numerals_\<close> and
  1461   reduce_012 = \<open>prep_rls' reduce_012\<close> and
  1462   discard_parentheses = \<open>prep_rls' discard_parentheses\<close> and
  1463  
  1464   make_polynomial = \<open>prep_rls' make_polynomial\<close> and
  1465   expand_binoms = \<open>prep_rls' expand_binoms\<close> and
  1466   rev_rew_p = \<open>prep_rls' rev_rew_p\<close> and
  1467   discard_minus = \<open>prep_rls' discard_minus\<close> and
  1468   expand_poly_ = \<open>prep_rls' expand_poly_\<close> and
  1469  
  1470   expand_poly_rat_ = \<open>prep_rls' expand_poly_rat_\<close> and
  1471   simplify_power_ = \<open>prep_rls' simplify_power_\<close> and
  1472   calc_add_mult_pow_ = \<open>prep_rls' calc_add_mult_pow_\<close> and
  1473   reduce_012_mult_ = \<open>prep_rls' reduce_012_mult_\<close> and
  1474   reduce_012_ = \<open>prep_rls' reduce_012_\<close> and
  1475  
  1476   discard_parentheses1 = \<open>prep_rls' discard_parentheses1\<close> and
  1477   order_mult_rls_ = \<open>prep_rls' order_mult_rls_\<close> and
  1478   order_add_rls_ = \<open>prep_rls' order_add_rls_\<close> and
  1479   make_rat_poly_with_parentheses = \<open>prep_rls' make_rat_poly_with_parentheses\<close>
  1480 
  1481 subsection \<open>problems\<close>
  1482 
  1483 problem pbl_simp_poly : "polynomial/simplification" =
  1484   \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)
  1485     \<^rule_eval>\<open>is_polyexp\<close> (eval_is_polyexp "")]\<close>
  1486   Method: "simplification/for_polynomials"
  1487   CAS: "Simplify t_t"
  1488   Given: "Term t_t"
  1489   Where: "t_t is_polyexp"
  1490   Find: "normalform n_n"
  1491 
  1492 subsection \<open>methods\<close>
  1493 
  1494 partial_function (tailrec) simplify :: "real \<Rightarrow> real"
  1495   where
  1496 "simplify term = ((Rewrite_Set ''norm_Poly'') term)"
  1497 
  1498 method met_simp_poly : "simplification/for_polynomials" =
  1499   \<open>{rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty,
  1500     prls = Rule_Set.append_rules "simplification_for_polynomials_prls" Rule_Set.empty
  1501       [(*for preds in where_*) \<^rule_eval>\<open>is_polyexp\<close> (eval_is_polyexp"")],
  1502     crls = Rule_Set.empty, errpats = [], nrls = norm_Poly}\<close>
  1503   Program: simplify.simps
  1504   Given: "Term t_t"
  1505   Where: "t_t is_polyexp"
  1506   Find: "normalform n_n"
  1507 
  1508 ML \<open>
  1509 \<close> ML \<open>
  1510 \<close> 
  1511 end