src/Tools/isac/Knowledge/Poly.thy
author Walther Neuper <neuper@ist.tugraz.at>
Tue, 28 Sep 2010 10:10:26 +0200
branchisac-update-Isa09-2
changeset 38034 928cebc9c4aa
parent 38031 460c24a6a6ba
child 38036 02a9b2540eb7
permissions -rw-r--r--
updated "op *" --> Groups.times_class.times in src and test

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