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