//test/../rewrite.sml,poly.sml WORK
authorwneuper <walther.neuper@jku.at>
Sat, 03 Jul 2021 16:21:07 +0200
changeset 60318e6e7a9b9ced7
parent 60317 638d02a9a96a
child 60319 2edbed71fde6
//test/../rewrite.sml,poly.sml WORK
src/Tools/isac/Knowledge/Base_Tools.thy
src/Tools/isac/Knowledge/Poly.thy
src/Tools/isac/Knowledge/Rational.thy
src/Tools/isac/ProgLang/Prog_Expr.thy
test/Tools/isac/Knowledge/gcd_poly_ml.sml
test/Tools/isac/Knowledge/poly.sml
test/Tools/isac/Knowledge/rational.sml
test/Tools/isac/MathEngBasic/rewrite.sml
test/Tools/isac/ProgLang/evaluate.sml
test/Tools/isac/ProgLang/prog_expr.sml
test/Tools/isac/Test_Isac_Short.thy
test/Tools/isac/Test_Some.thy
     1.1 --- a/src/Tools/isac/Knowledge/Base_Tools.thy	Tue Jun 01 15:41:23 2021 +0200
     1.2 +++ b/src/Tools/isac/Knowledge/Base_Tools.thy	Sat Jul 03 16:21:07 2021 +0200
     1.3 @@ -6,7 +6,8 @@
     1.4  begin
     1.5  subsection \<open>theorems for Base_Tools\<close>
     1.6  
     1.7 -lemma real_unari_minus:   "- a = (-1) * (a::real)" by auto
     1.8 +lemma real_unari_minus:   "Not (a is_const) ==> - a = (-1) * (a::real)" by auto
     1.9 +(*lemma real_unari_minus:   "- a = (-1) * (a::real)" by auto LOOPS WITH NUMERALS*)
    1.10  (*Semiring_Normalization.comm_ring_1_class.ring_normalization_rules(1)*)
    1.11  
    1.12  (* should be in Rational.thy, but needed for asms in e.g. d2_pqformula1 in PolyEq, RootEq... *)
     2.1 --- a/src/Tools/isac/Knowledge/Poly.thy	Tue Jun 01 15:41:23 2021 +0200
     2.2 +++ b/src/Tools/isac/Knowledge/Poly.thy	Sat Jul 03 16:21:07 2021 +0200
     2.3 @@ -268,18 +268,27 @@
     2.4      (monom2list t1) @ (monom2list t2)
     2.5    | monom2list t = [t];
     2.6  
     2.7 +(* use this fun ONLY locally: makes negative numbers with "-" instead of "~" for poly-orders*)
     2.8 +fun dest_number' t =
     2.9 +  let
    2.10 +    val i = TermC.num_of_term t
    2.11 +  in
    2.12 +    if i >= 0 then HOLogic.dest_number t |> snd |> string_of_int
    2.13 +    else (i * ~1) |> string_of_int |> curry op ^ "-"
    2.14 +  end
    2.15 +
    2.16  (* liefert Variablenname (String) einer Variablen und Basis bei Potenz *)
    2.17  fun get_basStr (Const ("Transcendental.powr",_) $ Free (str, _) $ _) = str
    2.18 +  | get_basStr (Const ("Transcendental.powr",_) $ n $ _) = dest_number' n
    2.19    | get_basStr (Free (str, _)) = str
    2.20 -  | get_basStr (Const ("Num.numeral_class.numeral", _) $ _) = "123"
    2.21 -  | get_basStr _ = "|||"; (* gross gewichtet; für Brüch ect. *)
    2.22 +  | get_basStr t =
    2.23 +    if TermC.is_num t then dest_number' t
    2.24 +    else "|||"; (* gross gewichtet; für Brüche ect. *)
    2.25  
    2.26  (* liefert Hochzahl (String) einer Variablen bzw Gewichtstring (zum Sortieren) *)
    2.27 -fun get_potStr (Const ("Transcendental.powr",_) $ Free _ $ t) =
    2.28 -    (case t of
    2.29 -      Free (str, _) => str
    2.30 -    | t => 
    2.31 -      if TermC.is_num t then "123" else "|||" (* gross gewichtet *))
    2.32 +fun get_potStr (Const ("Transcendental.powr",_) $ Free _ $ Free (str, _)) = str
    2.33 +  | get_potStr (Const ("Transcendental.powr",_) $ Free _ $ t) =
    2.34 +    if TermC.is_num t then dest_number' t else "|||"
    2.35    | get_potStr (Free (_, _)) = "---" (* keine Hochzahl --> kleinst gewichtet *)
    2.36    | get_potStr _ = "||||||"; (* gross gewichtet; für Brüch ect. *)
    2.37  
    2.38 @@ -290,27 +299,29 @@
    2.39      innerhalb eines Monomes:
    2.40      - zuerst lexikographisch nach Variablenname 
    2.41      - wenn gleich: nach steigender Potenz *)
    2.42 -fun var_ord (a,b: term) = prod_ord string_ord string_ord 
    2.43 +fun var_ord (a,b: term) = 
    2.44 +  prod_ord string_ord string_ord 
    2.45      ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
    2.46  
    2.47  (* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen); 
    2.48     verwendet zum Sortieren von Monomen mittels Gesamtgradordnung:
    2.49     - zuerst lexikographisch nach Variablenname 
    2.50     - wenn gleich: nach sinkender Potenz*)
    2.51 -fun var_ord_revPow (a,b: term) = prod_ord string_ord string_ord_rev 
    2.52 +fun var_ord_revPow (a, b: term) =
    2.53 +  prod_ord string_ord string_ord_rev 
    2.54      ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
    2.55  
    2.56 -
    2.57  (* Ordnet ein Liste von Variablen (und Potenzen) lexikographisch *)
    2.58  val sort_varList = sort var_ord;
    2.59  
    2.60  (* Entfernet aeussersten Operator (Wurzel) aus einem Term und schreibt 
    2.61     Argumente in eine Liste *)
    2.62  fun args u : term list =
    2.63 -    let fun stripc (f$t, ts) = stripc (f, t::ts)
    2.64 -	  | stripc (t as Free _, ts) = (t::ts)
    2.65 -	  | stripc (_, ts) = ts
    2.66 -    in stripc (u, []) end;
    2.67 +  let
    2.68 +    fun stripc (f $ t, ts) = stripc (f, t::ts)
    2.69 +  	  | stripc (t as Free _, ts) = (t::ts)
    2.70 +  	  | stripc (_, ts) = ts
    2.71 +  in stripc (u, []) end;
    2.72                                      
    2.73  (* liefert True, falls der Term (Liste von Termen) nur Zahlen 
    2.74     (keine Variablen) enthaelt *)
    2.75 @@ -321,21 +332,22 @@
    2.76  fun is_nums t = filter_num [t];
    2.77  
    2.78  (* Berechnet den Gesamtgrad eines Monoms *)
    2.79 -local 
    2.80 +(**)local(**)
    2.81    fun counter (n, []) = n
    2.82      | counter (n, x :: xs) = 
    2.83  	    if (is_nums x) then counter (n, xs)
    2.84  	    else 
    2.85  	      (case x of 
    2.86 -		      (Const ("Transcendental.powr", _) $ Free _ $ Const ("Num.numeral_class.numeral", _) $ num) => 
    2.87 -            counter (n + HOLogic.dest_numeral num, xs)
    2.88 -	      | (Const ("Transcendental.powr", _) $ Free _ $ _ ) => counter (n + 1000, xs) (*FIXME.MG?!*)
    2.89 -	      | (Free _) => counter (n + 1, xs)
    2.90 -	    (*| _ => raise ERROR("monom_degree: called with factor: "^(UnparseC.term x)))*)
    2.91 -	      | _ => counter (n + 10000, xs)) (*FIXME.MG?! ... Brüche ect.*)
    2.92 -in  
    2.93 +		      (Const ("Transcendental.powr", _) $ Free _ $ t) =>
    2.94 +            if TermC.is_num t
    2.95 +            then counter (t |> HOLogic.dest_number |> snd |> curry op + n, xs)
    2.96 +            else counter (n + 1000, xs) (*FIXME.MG?!*)
    2.97 +	      | (Const ("Num.numeral_class.numeral", _) $ num) =>
    2.98 +            counter (n + 1 + HOLogic.dest_numeral num, xs)
    2.99 +	      | _ => counter (n + 1, xs)) (*FIXME.MG?! ... Brüche ect.*)
   2.100 +(**)in(**)
   2.101    fun monom_degree l = counter (0, l) 
   2.102 -end;(*local*)
   2.103 +(**)end;(*local*)
   2.104  
   2.105  (* wie Ordnung dict_ord (lexicographische Ordnung zweier Listen, mit Vergleich 
   2.106     der Listen-Elemente mit elem_ord) - Elemente die Bedingung cond erfuellen, 
   2.107 @@ -345,40 +357,43 @@
   2.108    | dict_cond_ord _ _ (_ :: _, []) = GREATER
   2.109    | dict_cond_ord elem_ord cond (x :: xs, y :: ys) =
   2.110      (case (cond x, cond y) of 
   2.111 -	 (false, false) => (case elem_ord (x, y) of 
   2.112 -				EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
   2.113 -			      | ord => ord)
   2.114 -       | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
   2.115 -       | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
   2.116 -       | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys) );
   2.117 +	    (false, false) =>
   2.118 +        (case elem_ord (x, y) of 
   2.119 +				  EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
   2.120 +			  | ord => ord)
   2.121 +    | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
   2.122 +    | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
   2.123 +    | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys) );
   2.124  
   2.125  (* Gesamtgradordnung zum Vergleich von Monomen (Liste von Variablen/Potenzen):
   2.126     zuerst nach Gesamtgrad, bei gleichem Gesamtgrad lexikographisch ordnen - 
   2.127     dabei werden Koeffizienten ignoriert (2*3*a \<up> 2*4*b gilt wie a \<up> 2*b) *)
   2.128  fun degree_ord (xs, ys) =
   2.129 -	    prod_ord int_ord (dict_cond_ord var_ord_revPow is_nums) 
   2.130 -	    ((monom_degree xs, xs), (monom_degree ys, ys));
   2.131 +	prod_ord int_ord (dict_cond_ord var_ord_revPow is_nums)
   2.132 +	  ((monom_degree xs, xs), (monom_degree ys, ys));
   2.133  
   2.134  fun hd_str str = substring (str, 0, 1);
   2.135  fun tl_str str = substring (str, 1, (size str) - 1);
   2.136  
   2.137  (* liefert nummerischen Koeffizienten eines Monoms oder NONE *)
   2.138 -fun get_koeff_of_mon [] =  raise ERROR("get_koeff_of_mon: called with l = []")
   2.139 -  | get_koeff_of_mon (x::_) = if is_nums x then SOME x else NONE;
   2.140 +fun get_koeff_of_mon [] = raise ERROR "get_koeff_of_mon: called with l = []"
   2.141 +  | get_koeff_of_mon (x :: _) = if is_nums x then SOME x else NONE;
   2.142  
   2.143  (* wandelt Koeffizient in (zum sortieren geeigneten) String um *)
   2.144 -fun koeff2ordStr (SOME x) = (case x of 
   2.145 -				 (Free (str, _)) => 
   2.146 -				     if (hd_str str) = "-" then (tl_str str)^"0" (* 3 < -3 *)
   2.147 -				     else str
   2.148 -			       | _ => "aaa") (* "num.Ausdruck" --> gross *)
   2.149 -  | koeff2ordStr NONE = "---"; (* "kein Koeff" --> kleinste *)
   2.150 +fun koeff2ordStr (SOME t) =
   2.151 +    if TermC.is_num t
   2.152 +    then 
   2.153 +      if (t |> HOLogic.dest_number |> snd) < 0
   2.154 +      then (t |> HOLogic.dest_number |> snd |> curry op * ~1 |> string_of_int) ^ "0"  (* 3 < -3 *)
   2.155 +      else (t |> HOLogic.dest_number |> snd |> string_of_int)
   2.156 +    else "aaa"                                                      (* "num.Ausdruck" --> gross *)
   2.157 +  | koeff2ordStr NONE = "---";                                     (* "kein Koeff" --> kleinste *)
   2.158  
   2.159  (* Order zum Vergleich von Koeffizienten (strings): 
   2.160     "kein Koeff" < "0" < "1" < "-1" < "2" < "-2" < ... < "num.Ausdruck" *)
   2.161 -fun compare_koeff_ord (xs, ys) = 
   2.162 -    string_ord ((koeff2ordStr o get_koeff_of_mon) xs,
   2.163 -		(koeff2ordStr o get_koeff_of_mon) ys);
   2.164 +fun compare_koeff_ord (xs, ys) = string_ord
   2.165 +  ((koeff2ordStr o get_koeff_of_mon) xs,
   2.166 +   (koeff2ordStr o get_koeff_of_mon) ys);
   2.167  
   2.168  (* Gesamtgradordnung degree_ord + Ordnen nach Koeffizienten falls EQUAL *)
   2.169  fun koeff_degree_ord (xs, ys) =
   2.170 @@ -434,12 +449,12 @@
   2.171  (* sorts the monoms of an expanded and variable-sorted polynomial 
   2.172     by total_degree *)
   2.173  fun sort_monoms t = 
   2.174 -    let
   2.175 -	val ll =  map monom2list (poly2list t);
   2.176 -	val lls = sort_monList ll;
   2.177 -	val T = type_of t;
   2.178 -	val ls = map (create_monom T) lls;
   2.179 -    in create_polynom T ls end;
   2.180 +  let
   2.181 +  	val ll =  map monom2list (poly2list t);
   2.182 +  	val lls = sort_monList ll;
   2.183 +  	val T = Term.type_of t;
   2.184 +  	val ls = map (create_monom T) lls;
   2.185 +  in create_polynom T ls end;
   2.186  \<close>
   2.187  
   2.188  subsubsection \<open>rewrite order for hard-coded AC rewriting\<close>
   2.189 @@ -530,18 +545,34 @@
   2.190     this is weaker than 'is_polynomial' !.*)
   2.191  fun is_polyexp (Free _) = true
   2.192    | is_polyexp (Const _) = true (* potential danger: bdv is not considered *)
   2.193 -  | is_polyexp (Const ("Groups.plus_class.plus",_) $ Free _ $ Free _) = true
   2.194 -  | is_polyexp (Const ("Groups.minus_class.minus",_) $ Free _ $ Free _) = true
   2.195 -  | is_polyexp (Const ("Groups.times_class.times",_) $ Free _ $ Free _) = true
   2.196 -  | is_polyexp (Const ("Transcendental.powr",_) $ Free _ $ Free _) = true
   2.197 +  | is_polyexp (Const ("Groups.plus_class.plus",_) $ Free _ $ num) =
   2.198 +    if TermC.is_num num then true
   2.199 +    else if TermC.is_variable num then true
   2.200 +    else is_polyexp num
   2.201 +  | is_polyexp (Const ("Groups.plus_class.plus",_) $ num $ Free _) =
   2.202 +    if TermC.is_num num then true
   2.203 +    else if TermC.is_variable num then true
   2.204 +    else is_polyexp num
   2.205 +  | is_polyexp (Const ("Groups.minus_class.minus",_) $ Free _ $ num) =
   2.206 +    if TermC.is_num num then true
   2.207 +    else if TermC.is_variable num then true
   2.208 +    else is_polyexp num
   2.209 +  | is_polyexp (Const ("Groups.times_class.times",_) $ num $ Free _) =
   2.210 +    if TermC.is_num num then true
   2.211 +    else if TermC.is_variable num then true
   2.212 +    else is_polyexp num
   2.213 +  | is_polyexp (Const ("Transcendental.powr",_) $ Free _ $ num) =
   2.214 +    if TermC.is_num num then true
   2.215 +    else if TermC.is_variable num then true
   2.216 +    else is_polyexp num
   2.217    | is_polyexp (Const ("Groups.plus_class.plus",_) $ t1 $ t2) = 
   2.218 -               ((is_polyexp t1) andalso (is_polyexp t2))
   2.219 +    ((is_polyexp t1) andalso (is_polyexp t2))
   2.220    | is_polyexp (Const ("Groups.minus_class.minus",_) $ t1 $ t2) = 
   2.221 -               ((is_polyexp t1) andalso (is_polyexp t2))
   2.222 +    ((is_polyexp t1) andalso (is_polyexp t2))
   2.223    | is_polyexp (Const ("Groups.times_class.times",_) $ t1 $ t2) = 
   2.224 -               ((is_polyexp t1) andalso (is_polyexp t2))
   2.225 +    ((is_polyexp t1) andalso (is_polyexp t2))
   2.226    | is_polyexp (Const ("Transcendental.powr",_) $ t1 $ t2) = 
   2.227 -               ((is_polyexp t1) andalso (is_polyexp t2))
   2.228 +    ((is_polyexp t1) andalso (is_polyexp t2))
   2.229    | is_polyexp num = TermC.is_num num;
   2.230  \<close>
   2.231  
   2.232 @@ -555,7 +586,7 @@
   2.233  \<close>
   2.234  
   2.235  subsection \<open>evaluations functions\<close>
   2.236 -subsubsection \<open>for predicates\<close>
   2.237 +subsubsection \<open>for predicates\<close>                   
   2.238  ML \<open>
   2.239  fun eval_is_polyrat_in _ _(p as (Const ("Poly.is_polyrat_in",_) $ t $ v)) _  =
   2.240      if is_polyrat_in t v 
   2.241 @@ -675,16 +706,30 @@
   2.242    Rule_Def.Repeat {id = "discard_minus", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   2.243        erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   2.244        rules =
   2.245 -       [Rule.Thm ("real_diff_minus", ThmC.numerals_to_Free @{thm real_diff_minus}),
   2.246 +       [Rule.Thm ("real_diff_minus", ThmC.numerals_to_Free @{thm real_diff_minus}) (*,*)
   2.247            (*"a - b = a + -1 * b"*)
   2.248 -	        Rule.Thm ("sym_real_mult_minus1", ThmC.numerals_to_Free (@{thm real_mult_minus1} RS @{thm sym}))
   2.249 -	          (*- ?z = "-1 * ?z"*)],
   2.250 +      (* superfluous since proper "uminus", but shifted realpow_minus_even, realpow_minus_odd here.
   2.251 +	      Rule.Thm ("sym_real_mult_minus1", ThmC.numerals_to_Free (@{thm real_mult_minus1} RS @{thm sym}))
   2.252 +	          "- ?z = -1 * ?z"*)],
   2.253  	      scr = Rule.Empty_Prog};
   2.254  
   2.255 +(*erls for calculate_Rational; make local with FIXX@ME result:term *term list*)
   2.256 +val powers_erls =
   2.257 +  Rule_Def.Repeat {id = "powers_erls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
   2.258 +      erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   2.259 +      rules = [Rule.Eval ("Prog_Expr.is_atom", Prog_Expr.eval_is_atom "#is_atom_"),
   2.260 +	       Rule.Eval ("Prog_Expr.is_even", Prog_Expr.eval_is_even "#is_even_"),
   2.261 +	       Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
   2.262 +	       Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false}),
   2.263 +	       Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}),
   2.264 +	       Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_")
   2.265 +	       ],
   2.266 +      scr = Rule.Empty_Prog
   2.267 +      };
   2.268  val expand_poly_ = 
   2.269    Rule_Def.Repeat{id = "expand_poly_", preconds = [], 
   2.270        rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   2.271 -      erls = Rule_Set.empty,srls = Rule_Set.Empty,
   2.272 +      erls = powers_erls, srls = Rule_Set.Empty,
   2.273        calc = [], errpatts = [],
   2.274        rules =
   2.275          [Rule.Thm ("real_plus_binom_pow4", ThmC.numerals_to_Free @{thm real_plus_binom_pow4}),
   2.276 @@ -711,8 +756,12 @@
   2.277  	       
   2.278  	         Rule.Thm ("realpow_multI", ThmC.numerals_to_Free @{thm realpow_multI}),
   2.279  	           (*"(r * s) \<up> n = r \<up> n * s \<up> n"*)
   2.280 -	         Rule.Thm ("realpow_pow",ThmC.numerals_to_Free @{thm realpow_pow})
   2.281 +	         Rule.Thm ("realpow_pow",ThmC.numerals_to_Free @{thm realpow_pow}),
   2.282  	           (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
   2.283 +	         Rule.Thm ("realpow_minus_even",ThmC.numerals_to_Free @{thm realpow_minus_even}),
   2.284 +	           (*"n is_even ==> (- r) \<up> n = r \<up> n"*)
   2.285 +	         Rule.Thm ("realpow_minus_odd",ThmC.numerals_to_Free @{thm realpow_minus_odd})
   2.286 +	           (*"Not (n is_even) ==> (- r) \<up> n = -1 * r \<up> n"*)
   2.287  	       ], scr = Rule.Empty_Prog};
   2.288  
   2.289  val expand_poly_rat_ = 
   2.290 @@ -749,8 +798,12 @@
   2.291  	 Rule.Thm ("realpow_multI_poly", ThmC.numerals_to_Free @{thm realpow_multI_poly}),
   2.292  	     (*"[| r is_polyexp; s is_polyexp |] ==> 
   2.293  		            (r * s) \<up> n = r \<up> n * s \<up> n"*)
   2.294 -	  Rule.Thm ("realpow_pow",ThmC.numerals_to_Free @{thm realpow_pow})
   2.295 -	      (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
   2.296 +	 Rule.Thm ("realpow_pow",ThmC.numerals_to_Free @{thm realpow_pow}),
   2.297 +	   (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
   2.298 +	 Rule.Thm ("realpow_minus_even",ThmC.numerals_to_Free @{thm realpow_minus_even}),
   2.299 +	   (*"n is_even ==> (- r) \<up> n = r \<up> n"*)
   2.300 +	 Rule.Thm ("realpow_minus_odd",ThmC.numerals_to_Free @{thm realpow_minus_odd})
   2.301 +	   (*"Not (n is_even) ==> (- r) \<up> n = -1 * r \<up> n"*)
   2.302  	 ], scr = Rule.Empty_Prog};
   2.303  
   2.304  val simplify_power_ = 
     3.1 --- a/src/Tools/isac/Knowledge/Rational.thy	Tue Jun 01 15:41:23 2021 +0200
     3.2 +++ b/src/Tools/isac/Knowledge/Rational.thy	Sat Jul 03 16:21:07 2021 +0200
     3.3 @@ -123,20 +123,20 @@
     3.4  ML \<open>
     3.5  fun monom_of_term vs (c, es) (t as Const _) =
     3.6      (c, list_update es (find_index (curry op = t) vs) 1)
     3.7 -  | monom_of_term  vs (c, es) (t as (Const ("Num.numeral_class.numeral", _) $ _)) =
     3.8 -    if TermC.is_num t
     3.9 -    then (t |> HOLogic.dest_number |> snd |> string_of_int
    3.10 +  | monom_of_term vs (c, es) (t as (Const ("Num.numeral_class.numeral", _) $ _)) =
    3.11 +    (t |> HOLogic.dest_number |> snd |> string_of_int
    3.12        |> TermC.int_opt_of_string |> the |> curry op * c, es) (*several numerals in one monom*)
    3.13 -    else (c, list_update es (find_index (curry op = t) vs) 1)
    3.14 -  | monom_of_term  vs (c, es) (Const ("Transcendental.powr", _) $
    3.15 -      (t as (Const ("Num.numeral_class.numeral", _) $ _)) $
    3.16 -        (Const ("Num.numeral_class.numeral", _) $ num)) =
    3.17 +  | monom_of_term  vs (c, es) (t as Free _) =
    3.18 +    (c, list_update es (find_index (curry op = t) vs) 1)
    3.19 +  | monom_of_term  vs (c, es) (Const ("Transcendental.powr", _) $ (t as Free _) $
    3.20 +      (Const ("Num.numeral_class.numeral", _) $ num)) =
    3.21      (c, list_update es (find_index (curry op = t) vs) (num |> HOLogic.dest_numeral))
    3.22    | monom_of_term vs (c, es) (Const ("Groups.times_class.times", _) $ m1 $ m2) =
    3.23      let val (c', es') = monom_of_term vs (c, es) m1
    3.24      in monom_of_term vs (c', es') m2 end
    3.25    | monom_of_term _ _ t = raise ERROR ("poly malformed 1 with " ^ UnparseC.term t)
    3.26  
    3.27 +(*-------v------*)
    3.28  fun monoms_of_term vs (t as Const _) =
    3.29      [monom_of_term  vs (1, replicate (length vs) 0) t]
    3.30    | monoms_of_term vs (t as Free _) =
    3.31 @@ -604,19 +604,6 @@
    3.32  
    3.33  section \<open>Rulesets for general simplification\<close>
    3.34  ML \<open>
    3.35 -(*erls for calculate_Rational; make local with FIXX@ME result:term *term list*)
    3.36 -val powers_erls = prep_rls'(
    3.37 -  Rule_Def.Repeat {id = "powers_erls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
    3.38 -      erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
    3.39 -      rules = [Rule.Eval ("Prog_Expr.is_atom", Prog_Expr.eval_is_atom "#is_atom_"),
    3.40 -	       Rule.Eval ("Prog_Expr.is_even", Prog_Expr.eval_is_even "#is_even_"),
    3.41 -	       Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
    3.42 -	       Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false}),
    3.43 -	       Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}),
    3.44 -	       Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_")
    3.45 -	       ],
    3.46 -      scr = Rule.Empty_Prog
    3.47 -      });
    3.48  (*.all powers over + distributed; atoms over * collected, other distributed
    3.49     contains absolute minimum of thms for context in norm_Rational .*)
    3.50  val powers = prep_rls'(
     4.1 --- a/src/Tools/isac/ProgLang/Prog_Expr.thy	Tue Jun 01 15:41:23 2021 +0200
     4.2 +++ b/src/Tools/isac/ProgLang/Prog_Expr.thy	Sat Jul 03 16:21:07 2021 +0200
     4.3 @@ -290,32 +290,35 @@
     4.4  
     4.5  fun even i = (i div 2) * 2 = i;
     4.6  (*("is_even",("Prog_Expr.is_even", eval_is_even "#is_even_"))*)
     4.7 -fun eval_is_even (thmid:string) "Prog_Expr.is_even"
     4.8 -		 (t as (Const _ $ arg)) _ = 
     4.9 -    (case arg of 
    4.10 -	Free (n,_) =>
    4.11 -	 (case ThmC_Def.int_opt_of_string n of
    4.12 -	      SOME i =>
    4.13 -	      if even i then SOME (TermC.mk_thmid thmid n "", 
    4.14 -				   HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
    4.15 +fun eval_is_even (thmid:string) "Prog_Expr.is_even" (t as (Const _ $ arg)) _ = 
    4.16 +    if TermC.is_num arg
    4.17 +    then
    4.18 +      let
    4.19 +        val i = arg |> HOLogic.dest_number |> snd
    4.20 +      in
    4.21 +	      if even i 
    4.22 +        then SOME (TermC.mk_thmid thmid (string_of_int i) "", 
    4.23 +				  HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
    4.24  	      else SOME (TermC.mk_thmid thmid "" "", 
    4.25 -			 HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
    4.26 -	    | _ => NONE)
    4.27 -       | _ => NONE)
    4.28 +			    HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
    4.29 +      end
    4.30 +    else NONE
    4.31    | eval_is_even _ _ _ _ = NONE; 
    4.32  
    4.33 -(*evaluate 'is_const'*)
    4.34  (*("is_const",("Prog_Expr.is_const", Prog_Expr.eval_const "#is_const_"))*)
    4.35 -fun eval_const (thmid:string) _ (t as (Const _ $ arg)) _ = 
    4.36 -    (case arg of 
    4.37 -       Const (n1, _) =>
    4.38 -	     SOME (TermC.mk_thmid thmid n1 "", HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
    4.39 -     | n1 =>
    4.40 -	     if TermC.is_num n1
    4.41 -	     then SOME (TermC.mk_thmid thmid (TermC.string_of_num n1) "",
    4.42 -         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
    4.43 -	     else SOME (TermC.mk_thmid thmid (UnparseC.term n1) "",
    4.44 -         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False}))))
    4.45 +fun eval_const thmid _ (t as (Const ("Prog_Expr.is_const", _) $ Const ("Partial_Fractions.AA", _))) _ =
    4.46 +    (*TODO get rid of this special case*)
    4.47 +    SOME (TermC.mk_thmid thmid (UnparseC.term t) "",
    4.48 +      HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
    4.49 +  | eval_const thmid _ (t as (Const ("Prog_Expr.is_const", _) $ Const _)) _ =
    4.50 +    SOME (TermC.mk_thmid thmid (UnparseC.term t) "",
    4.51 +      HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
    4.52 +  | eval_const thmid _ (t as (Const ("Prog_Expr.is_const", _) $ n)) _ =
    4.53 +	   if TermC.is_num n
    4.54 +	   then SOME (TermC.mk_thmid thmid (TermC.string_of_num n) "",
    4.55 +       HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
    4.56 +	   else SOME (TermC.mk_thmid thmid (UnparseC.term n) "",
    4.57 +       HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
    4.58    | eval_const _ _ _ _ = NONE; 
    4.59  
    4.60  (*. evaluate binary, associative, commutative operators: *,+,^ .*)
     5.1 --- a/test/Tools/isac/Knowledge/gcd_poly_ml.sml	Tue Jun 01 15:41:23 2021 +0200
     5.2 +++ b/test/Tools/isac/Knowledge/gcd_poly_ml.sml	Sat Jul 03 16:21:07 2021 +0200
     5.3 @@ -1,8 +1,6 @@
     5.4 -(* Title: test/../rational2
     5.5 +(* Title: test/Knowledge/gec_poly_ml.sml
     5.6     Author: Diana Meindl
     5.7     Copyright (c) Diana Meindl 2011
     5.8 -12345678901234567890123456789012345678901234567890123456789012345678901234567890
     5.9 -        10        20        30        40        50        60        70        80
    5.10  *)
    5.11  
    5.12  (*fun nth _ []      = error "nth _ []" (*Isabelle2002, still saved the hours of update*)
    5.13 @@ -493,7 +491,7 @@
    5.14  "----------- fun add_monoms -----------------------------";
    5.15  "----------- fun add_monoms -----------------------------";
    5.16  "----------- fun add_monoms -----------------------------";
    5.17 -if add_monoms [(~3,[0,0]),(4,[0,0]),(3,[1,1]),(~3,[1,1]),(2,[1,2]),(~3,[1,2])] = 
    5.18 +if add_monoms [(~3,[0,0]), (4,[0,0]), (3,[1,1]), (~3,[1,1]), (2,[1,2]), (~3,[1,2])] = 
    5.19    [(1, [0, 0]), (~1, [1, 2])] 
    5.20  then () else error ("add_monoms [(~3,[0,0]),(4,[0,0]),(3,[1,1]),(~3,[1,1]),(2,[1,2]),(~3,[1,2])] " ^
    5.21    "= [(1, [0, 0]), (~1, [1, 2])] changed")
     6.1 --- a/test/Tools/isac/Knowledge/poly.sml	Tue Jun 01 15:41:23 2021 +0200
     6.2 +++ b/test/Tools/isac/Knowledge/poly.sml	Sat Jul 03 16:21:07 2021 +0200
     6.3 @@ -1,56 +1,155 @@
     6.4 -(* testexamples for Poly, polynomials
     6.5 +(* Knowledge/poly.sml
     6.6     author: Matthias Goldgruber 2003
     6.7     (c) due to copyright terms
     6.8  
     6.9 -12345678901234567890123456789012345678901234567890123456789012345678901234567890
    6.10 -        10        20        30        40        50        60        70        80
    6.11  LEGEND:
    6.12  WN060104: examples marked with 'SPB' came into 'exp_IsacCore_Simp_Poly_Book.xml'
    6.13            examples marked with 'SPO' came into 'exp_IsacCore_Simp_Poly_Other.xml'
    6.14  *)
    6.15  
    6.16 -"--------------------------------------------------------";
    6.17 -"--------------------------------------------------------";
    6.18 -"table of contents --------------------------------------";
    6.19 -"--------------------------------------------------------";
    6.20 -"----------- fun is_polyexp --------------------------------------------------------------------";
    6.21 -"----------- fun has_degree_in -----------------------------------------------------------------";
    6.22 -"----------- fun mono_deg_in -------------------------------------------------------------------";
    6.23 -"----------- fun mono_deg_in -------------------------------------------------------------------";
    6.24 -"----------- eval_ for is_expanded_in, is_poly_in, has_degree_in -------------------------------";
    6.25 -"-------- investigate (new 2002) uniary minus -----------";
    6.26 -"----------- fun sort_variables ----------------------------------------------------------------";
    6.27 -"-------- check make_polynomial with simple terms -------";
    6.28 -"-------- fun is_multUnordered --------------------------";
    6.29 -"-------- examples from textbook Schalk I ---------------";
    6.30 -"-------- check pbl  'polynomial simplification' --------";
    6.31 -"-------- me 'poly. simpl.' Schalk I p.63 No.267b -------";
    6.32 -"-------- interSteps for Schalk 299a --------------------";
    6.33 -"-------- norm_Poly NOT COMPLETE ------------------------";
    6.34 -"-------- ord_make_polynomial ---------------------------";
    6.35 -"--------------------------------------------------------";
    6.36 -"--------------------------------------------------------";
    6.37 -"--------------------------------------------------------";
    6.38 +"-----------------------------------------------------------------------------------------------";
    6.39 +"-----------------------------------------------------------------------------------------------";
    6.40 +"table of contents -----------------------------------------------------------------------------";
    6.41 +"-----------------------------------------------------------------------------------------------";
    6.42 +"-------- survey on parallel rewrite-breakers AND code with @{print} ---------------------------";
    6.43 +"-------- thm sym_real_mult_minus1 loops with new numerals -------------------------------------";
    6.44 +"-------- fun is_polyexp -----------------------------------------------------------------------";
    6.45 +"-------- fun has_degree_in --------------------------------------------------------------------";
    6.46 +"-------- fun mono_deg_in ----------------------------------------------------------------------";
    6.47 +"-------- rewrite_set_ has_degree_in Const ('Partial_Fractions', _) ----------------------------";
    6.48 +"-------- eval_ for is_expanded_in, is_poly_in, has_degree_in ----------------------------------";
    6.49 +"-------- investigate (new 2002) uniary minus --------------------------------------------------";
    6.50 +"-------- fun sort_variables -------------------------------------------------------------------";
    6.51 +"-------- rebuild fun is_addUnordered (1 + 2 * x \<up> 4 + - 4 * x \<up> 4 + x \<up> 8) ---------------------";
    6.52 +"-------- fun is_addUnordered (x \<up> 2 * y \<up> 2 + x \<up> 3 * y) --------------------------------------";
    6.53 +"-------- check make_polynomial with simple terms ----------------------------------------------";
    6.54 +"-------- fun is_multUnordered (x \<up> 2 * x) -----------------------------------------------------";
    6.55 +"-------- fun is_multUnordered (3 * a + - 2 * a) -----------------------------------------------";
    6.56 +"-------- fun is_multUnordered (x - a) \<up> 3 -----------------------------------------------------";
    6.57 +"-------- fun is_multUnordered  b * a * a ------------------------------------------------------";
    6.58 +"-------- fun is_multUnordered 2*3*a -----------------------------------------------------------";
    6.59 +"-------- examples from textbook Schalk I ------------------------------------------------------";
    6.60 +"-------- ?RL?Bsple bei denen es Probleme gibt--------------------------------------------------";
    6.61 +"-------- Eigene Beispiele (Mathias Goldgruber) ------------------------------------------------";
    6.62 +"-------- check pbl  'polynomial simplification' -----------------------------------------------";
    6.63 +"-------- me 'poly. simpl.' Schalk I p.63 No.267b ----------------------------------------------";
    6.64 +"-------- interSteps for Schalk 299a -----------------------------------------------------------";
    6.65 +"-------- norm_Poly NOT COMPLETE ---------------------------------------------------------------";
    6.66 +"-------- ord_make_polynomial ------------------------------------------------------------------";
    6.67 +"-----------------------------------------------------------------------------------------------";
    6.68 +"-----------------------------------------------------------------------------------------------";
    6.69 +"-----------------------------------------------------------------------------------------------";
    6.70  
    6.71  
    6.72 -"----------- fun is_polyexp --------------------------------------------------------------------";
    6.73 -"----------- fun is_polyexp --------------------------------------------------------------------";
    6.74 -"----------- fun is_polyexp --------------------------------------------------------------------";
    6.75 -val thy = @{theory Partial_Fractions};
    6.76 -val ctxt = Proof_Context.init_global thy;
    6.77 +"-------- survey on parallel rewrite-breakers AND code with @{print} ---------------------------";
    6.78 +"-------- survey on parallel rewrite-breakers AND code with @{print} ---------------------------";
    6.79 +"-------- survey on parallel rewrite-breakers AND code with @{print} ---------------------------";
    6.80 +(* indentation indicates call hierarchy:
    6.81 +"~~~~~ fun is_addUnordered
    6.82 +"~~~~~~~ fun is_polyexp
    6.83 +"~~~~~~~ fun sort_monoms
    6.84 +"~~~~~~~~~ fun sort_monList
    6.85 +"~~~~~~~~~~~ fun koeff_degree_ord    : term list * term list -> order
    6.86 +"~~~~~~~~~~~~~ fun degree_ord        : term list * term list -> order
    6.87 +"~~~~~~~~~~~~~~~ fun dict_cond_ord   : ('a * 'a -> order) -> ('a -> bool) -> 'a list * 'a list -> order
    6.88 +"~~~~~~~~~~~~~~~~~ fun var_ord_revPow: term * term -> order
    6.89 +"~~~~~~~~~~~~~~~~~~~ fun get_basStr  : term -> string                       used twice --vv
    6.90 +"~~~~~~~~~~~~~~~~~~~ fun get_potStr  : term -> string                       used twice --vv
    6.91 +"~~~~~~~~~~~~~~~ fun monom_degree    : term list -> int
    6.92 +"~~~~~~~~~~~~~ fun compare_koeff_ord : term list * term list -> order
    6.93 +"~~~~~~~~~~~~~~~ fun get_koeff_of_mon: term list -> term option
    6.94 +"~~~~~~~~~~~~~~~~~ fun koeff2ordStr  : term option -> string
    6.95 +"~~~~~ fun is_multUnordered
    6.96 +"~~~~~~~ fun sort_variables
    6.97 +"~~~~~~~~~ fun sort_varList
    6.98 +"~~~~~~~~~~~ fun var_ord
    6.99 +"~~~~~~~~~~~~~ fun get_basStr                                               used twice --^^
   6.100 +"~~~~~~~~~~~~~ fun get_potStr                                               used twice --^^
   6.101  
   6.102 -val t = (the o (parseNEW  ctxt)) "x / x";
   6.103 +fun int_ord (i1, i2) =
   6.104 +(@{print} {a = "int_ord (" ^ string_of_int i1 ^ ", " ^ string_of_int i2 ^ ") = ", z = Int.compare (i1, i2)};
   6.105 +  Int.compare (i1, i2)
   6.106 +);
   6.107 +fun var_ord (a, b) = 
   6.108 +(@{print} {a = "var_ord ", a_b = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")",
   6.109 +   sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
   6.110 +  prod_ord string_ord string_ord 
   6.111 +  ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b))
   6.112 +);
   6.113 +fun var_ord_revPow (a, b: term) = 
   6.114 +(@{print} {a = "var_ord_revPow ", at_bt = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")",
   6.115 +    sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
   6.116 +  prod_ord string_ord string_ord_rev 
   6.117 +    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b))
   6.118 +);
   6.119 +fun sort_varList ts =
   6.120 +(@{print} {a = "sort_varList", args = UnparseC.terms ts};
   6.121 +  sort var_ord ts
   6.122 +);
   6.123 +fun dict_cond_ord _ _ ([], [])     = (@{print} {a = "dict_cond_ord ([], [])"}; EQUAL)
   6.124 +  | dict_cond_ord _ _ ([], _ :: _) = (@{print} {a = "dict_cond_ord ([], _ :: _)"}; LESS)
   6.125 +  | dict_cond_ord _ _ (_ :: _, []) = (@{print} {a = "dict_cond_ord (_ :: _, [])"}; GREATER)
   6.126 +  | dict_cond_ord elem_ord cond (x :: xs, y :: ys) =
   6.127 +    (@{print} {a = "dict_cond_ord", args = "(" ^ UnparseC.terms (x :: xs) ^ ", " ^ UnparseC.terms (y :: ys) ^ ")", 
   6.128 +      is_nums = "(" ^ LibraryC.bool2str (cond x) ^ ", " ^ LibraryC.bool2str (cond y) ^ ")"};
   6.129 +     case (cond x, cond y) of 
   6.130 +	    (false, false) =>
   6.131 +        (case elem_ord (x, y) of 
   6.132 +				  EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
   6.133 +			  | ord => ord)
   6.134 +    | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
   6.135 +    | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
   6.136 +    | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys)
   6.137 +);
   6.138 +fun compare_koeff_ord (xs, ys) =
   6.139 +(@{print} {a = "compare_koeff_ord ", ats_bts = "(" ^ UnparseC.terms xs ^ ", " ^ UnparseC.terms ys ^ ")",
   6.140 +    sort_args = "(" ^ (koeff2ordStr o get_koeff_of_mon) xs ^ ", " ^ (koeff2ordStr o get_koeff_of_mon) ys ^ ")"};
   6.141 +  string_ord
   6.142 +  ((koeff2ordStr o get_koeff_of_mon) xs,
   6.143 +   (koeff2ordStr o get_koeff_of_mon) ys)
   6.144 +);
   6.145 +fun var_ord (a,b: term) = 
   6.146 +(@{print} {a = "var_ord ", a_b = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")",
   6.147 +   sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
   6.148 +  prod_ord string_ord string_ord 
   6.149 +  ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b))
   6.150 +);
   6.151 +*)
   6.152 +
   6.153 +
   6.154 +"-------- thm sym_real_mult_minus1 loops with new numerals -------------------------------------";
   6.155 +"-------- thm sym_real_mult_minus1 loops with new numerals -------------------------------------";
   6.156 +"-------- thm sym_real_mult_minus1 loops with new numerals -------------------------------------";
   6.157 +(* thus   ^^^^^^^^^^^^^^^^^^^^^^^^ excluded from rls.
   6.158 +
   6.159 +  sym_real_mult_minus1                                  =     "- ?z = - 1 * ?z" *)
   6.160 +@{thm real_mult_minus1};                              (* = "- 1 * ?z = - ?z"     *)
   6.161 +val thm_isac = ThmC.sym_thm @{thm real_mult_minus1};  (* =     "- ?z = - 1 * ?z" *)
   6.162 +val SOME t_isac = TermC.parseNEW @{context} "3 * a + - 1 * (2 * a):: real";
   6.163 +
   6.164 +val SOME (t', []) = rewrite_ @{theory} tless_true Rule_Set.empty true thm_isac t_isac;
   6.165 +if UnparseC.term t' = "3 * a + - 1 * 1 * (2 * a)" then ((*thm did NOT apply to Free ("-1", _)*))
   6.166 +else error "thm  - ?z = - 1 * ?z  loops with new numerals";
   6.167 +
   6.168 +
   6.169 +"-------- fun is_polyexp -----------------------------------------------------------------------";
   6.170 +"-------- fun is_polyexp -----------------------------------------------------------------------";
   6.171 +"-------- fun is_polyexp -----------------------------------------------------------------------";
   6.172 +val t = TermC.str2term "x / x";
   6.173  if is_polyexp t then error "NOT is_polyexp (x / x)" else ();
   6.174  
   6.175 -val t = (the o (parseNEW  ctxt)) "-1 * A * 3";
   6.176 +val t = TermC.str2term "-1 * A * 3";
   6.177  if is_polyexp t then () else error "is_polyexp (-1 * A * 3)";
   6.178  
   6.179 -val t = (the o (parseNEW  ctxt)) "-1 * AA * 3";
   6.180 +val t = TermC.str2term "-1 * AA * 3";
   6.181  if is_polyexp t then () else error "is_polyexp (-1 * AA * 3)";
   6.182  
   6.183 -"----------- fun has_degree_in -----------------------------------------------------------------";
   6.184 -"----------- fun has_degree_in -----------------------------------------------------------------";
   6.185 -"----------- fun has_degree_in -----------------------------------------------------------------";
   6.186 +val t = TermC.str2term "x * x + x * y + (- 1 * y * x + - 1 * y * y)::real";
   6.187 +if is_polyexp t then () else error "is_polyexp (x * x + x * y + (- 1 * y * x + - 1 * y * y))";
   6.188 +
   6.189 +"-------- fun has_degree_in --------------------------------------------------------------------";
   6.190 +"-------- fun has_degree_in --------------------------------------------------------------------";
   6.191 +"-------- fun has_degree_in --------------------------------------------------------------------";
   6.192  val v = (Thm.term_of o the o (TermC.parse thy)) "x";
   6.193  val t = (Thm.term_of o the o (TermC.parse thy)) "1";
   6.194  if has_degree_in t v = 0 then () else error "has_degree_in (1) x";
   6.195 @@ -131,9 +230,9 @@
   6.196  val t = (Thm.term_of o the o (TermC.parse thy)) "ab - (a*b)*AA";
   6.197  if has_degree_in t v = 1 then () else error "has_degree_in (ab - (a*b)*AA) AA";
   6.198  
   6.199 -"----------- fun mono_deg_in -------------------------------------------------------------------";
   6.200 -"----------- fun mono_deg_in -------------------------------------------------------------------";
   6.201 -"----------- fun mono_deg_in -------------------------------------------------------------------";
   6.202 +"-------- fun mono_deg_in ----------------------------------------------------------------------";
   6.203 +"-------- fun mono_deg_in ----------------------------------------------------------------------";
   6.204 +"-------- fun mono_deg_in ----------------------------------------------------------------------";
   6.205  val v = (Thm.term_of o the o (TermC.parse thy)) "x";
   6.206  
   6.207  val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)*x \<up> 7";
   6.208 @@ -182,9 +281,9 @@
   6.209  val t = (Thm.term_of o the o (TermC.parse thy)) "ab - (a*b)*AA";
   6.210  if mono_deg_in t v = NONE then () else error "mono_deg_in (ab - (a*b)*AA) AA changed";
   6.211  
   6.212 -"----------- rewrite_set_ has_degree_in Const ('Partial_Fractions', _) -------------------------";
   6.213 -"----------- rewrite_set_ has_degree_in Const ('Partial_Fractions', _) -------------------------";
   6.214 -"----------- rewrite_set_ has_degree_in Const ('Partial_Fractions', _) -------------------------";
   6.215 +"-------- rewrite_set_ has_degree_in Const ('Partial_Fractions', _) ----------------------------";
   6.216 +"-------- rewrite_set_ has_degree_in Const ('Partial_Fractions', _) ----------------------------";
   6.217 +"-------- rewrite_set_ has_degree_in Const ('Partial_Fractions', _) ----------------------------";
   6.218  val thy = @{theory Partial_Fractions}
   6.219  val expr = (Thm.term_of o the o (TermC.parse thy)) "((-8 - 2*x + x \<up> 2) has_degree_in x) = 2";
   6.220  val SOME (Const ("HOL.True", _), []) = rewrite_set_ thy false PolyEq_prls expr;
   6.221 @@ -192,9 +291,9 @@
   6.222  val expr = (Thm.term_of o the o (TermC.parse thy)) "((-8 - 2*AA + AA \<up> 2) has_degree_in AA) = 2";
   6.223  val SOME (Const ("HOL.True", _), []) = rewrite_set_ thy false PolyEq_prls expr;
   6.224  
   6.225 -"----------- eval_ for is_expanded_in, is_poly_in, has_degree_in -------------------------------";
   6.226 -"----------- eval_ for is_expanded_in, is_poly_in, has_degree_in -------------------------------";
   6.227 -"----------- eval_ for is_expanded_in, is_poly_in, has_degree_in -------------------------------";
   6.228 +"-------- eval_ for is_expanded_in, is_poly_in, has_degree_in ----------------------------------";
   6.229 +"-------- eval_ for is_expanded_in, is_poly_in, has_degree_in ----------------------------------";
   6.230 +"-------- eval_ for is_expanded_in, is_poly_in, has_degree_in ----------------------------------";
   6.231  val t = (Thm.term_of o the o (TermC.parse thy)) "(-8 - 2*x + x \<up> 2) is_expanded_in x";
   6.232  val SOME (id, t') = eval_is_expanded_in 0 0 t 0;
   6.233  if UnparseC.term t' = "- 8 - 2 * x + x \<up> 2 is_expanded_in x = True"
   6.234 @@ -229,10 +328,9 @@
   6.235  val expr = (Thm.term_of o the o (TermC.parse thy)) "((-8 - 2*AA + AA \<up> 2) has_degree_in AA) = 2";
   6.236  val SOME (Const ("HOL.True", _), []) = rewrite_set_ thy false PolyEq_prls expr;
   6.237  
   6.238 -"-------- investigate (new 2002) uniary minus -----------";
   6.239 -"-------- investigate (new 2002) uniary minus -----------";
   6.240 -"-------- investigate (new 2002) uniary minus -----------";
   6.241 -(*---------------------------------------------- vvvvvvvvvvvvvv -----------------------*)
   6.242 +"-------- investigate (new 2002) uniary minus --------------------------------------------------";
   6.243 +"-------- investigate (new 2002) uniary minus --------------------------------------------------";
   6.244 +"-------- investigate (new 2002) uniary minus --------------------------------------------------";
   6.245  val t = Thm.prop_of @{thm real_diff_0}; (*"0 - ?x = - ?x"*)
   6.246  TermC.atomty t;
   6.247  (*
   6.248 @@ -252,15 +350,7 @@
   6.249               (Const ("Groups.uminus_class.uminus", _) $ Var (("x", 0), _))) => ()
   6.250  | _ => error "internal representation of \"0 - ?x = - ?x\" changed";
   6.251  
   6.252 -(*----------------------------------- vvvv --------------------------------------------*)
   6.253 -val t = (Thm.term_of o the o (TermC.parse thy)) "-1";
   6.254 -TermC.atomty t;
   6.255 -(*** -------------
   6.256 -*** Free ( -1, real)                      *)
   6.257 -case t of
   6.258 -  Free ("-1", _) => ()
   6.259 -| _ => error "internal representation of \"-1\" changed";
   6.260 -(*----------------------------------- vvvvv -------------------------------------------*)
   6.261 +
   6.262  val t = (Thm.term_of o the o (TermC.parse thy)) "- 1";
   6.263  TermC.atomty t;
   6.264  (*
   6.265 @@ -269,7 +359,7 @@
   6.266  *** 
   6.267  *)
   6.268  case t of
   6.269 - Free ("-1", _) => ()
   6.270 + Const ("Groups.uminus_class.uminus", _) $ Const ("Groups.one_class.one", _) => ()
   6.271  | _ => error "internal representation of \"- 1\" changed";
   6.272  
   6.273  "======= these external values all have the same internal representation";
   6.274 @@ -299,9 +389,424 @@
   6.275    Const ("Groups.uminus_class.uminus", _) $ Free ("x", _) => ()
   6.276  | _ => error "internal representation of \"-(x)\" changed";
   6.277  
   6.278 -"-------- check make_polynomial with simple terms -------";
   6.279 -"-------- check make_polynomial with simple terms -------";
   6.280 -"-------- check make_polynomial with simple terms -------";
   6.281 +
   6.282 +"-------- fun sort_variables -------------------------------------------------------------------";
   6.283 +"-------- fun sort_variables -------------------------------------------------------------------";
   6.284 +"-------- fun sort_variables -------------------------------------------------------------------";
   6.285 +if "---" < "123" andalso "123" < "a" andalso "a" < "cba" then ()
   6.286 +else error "lexicographic order CHANGED";
   6.287 +
   6.288 +(*--------------vvvvvvvvvvvvvvv-----------------------------------------------------------------*)
   6.289 +val t =  @{term "3 * b + a * 2"}; (*------------------------------------------------------------*)
   6.290 +val t' =  sort_variables t;
   6.291 +if UnparseC.term t' = "(3::'a) * b + (2::'a) * a" then ()
   6.292 +else error "sort_variables  3 * b + a * 2  CHANGED";
   6.293 +
   6.294 +"~~~~~~~ fun sort_variables , args:"; val (t) = (t);
   6.295 +  	val ll =  map monom2list (poly2list t);
   6.296 +
   6.297 +(*+*)UnparseC.terms (poly2list t) = "[\"(3::'a) * b\", \"a * (2::'a)\"]";
   6.298 +(*+*)val [ [(Const ("Num.numeral_class.numeral", _) $ _), Free ("b", _)],
   6.299 +(*+*)      [Free ("a", _), (Const ("Num.numeral_class.numeral", _) $ _)]
   6.300 +(*+*)    ] = map monom2list (poly2list t);
   6.301 +
   6.302 +  	val lls = map sort_varList ll;
   6.303 +
   6.304 +(*+*)case map sort_varList ll of
   6.305 +(*+*)  [ [Const ("Num.numeral_class.numeral", _) $ _, Free ("b", _)],
   6.306 +(*+*)    [Const ("Num.numeral_class.numeral", _) $ _, Free ("a", _)]
   6.307 +(*+*)  ] => ()
   6.308 +(*+*)| _ => error "map sort_varList CHANGED";
   6.309 +
   6.310 +  	val T = type_of t;
   6.311 +  	val ls = map (create_monom T) lls;
   6.312 +
   6.313 +(*+*)val [Const ("Groups.times_class.times", _) $ _ $ Free ("b", _),
   6.314 +(*+*)     Const ("Groups.times_class.times", _) $ _ $ Free ("a", _)] = map (create_monom T) lls;
   6.315 +
   6.316 +(*+*)case map (create_monom T) lls of
   6.317 +(*+*)  [Const ("Groups.times_class.times", _) $ (Const ("Num.numeral_class.numeral", _) $ _) $ Free ("b", _),
   6.318 +(*+*)   Const ("Groups.times_class.times", _) $ (Const ("Num.numeral_class.numeral", _) $ _) $ Free ("a", _)
   6.319 +(*+*)  ] => ()
   6.320 +(*+*)| _ => error "map (create_monom T) CHANGED";
   6.321 +
   6.322 +  val xxx = (*in*) create_polynom T ls (*end*);
   6.323 +
   6.324 +(*+*)if UnparseC.term xxx = "(3::'a) * b + (2::'a) * a" then ()
   6.325 +(*+*)else error "create_polynom CHANGED";
   6.326 +(* done by rewriting>              2 * a +       3 * b ? *)
   6.327 +
   6.328 +(*---------------vvvvvvvvvvvvvvvvvvvvvv---------------------------------------------------------*)
   6.329 +val t =  @{term "3 * a + - 2 * a ::real"}; (*---------------------------------------------------*)
   6.330 +val t' =     sort_variables t;
   6.331 +if UnparseC.term t' = "3 * a + - 2 * a" then ()
   6.332 +else error "sort_variables  3 * a + - 2 * a  CHANGED";
   6.333 +
   6.334 +"~~~~~~~ fun sort_variables , args:"; val (t) = (t);
   6.335 +  	val ll =  map monom2list (poly2list t);
   6.336 +
   6.337 +(*+*)val [ [Const ("Num.numeral_class.numeral", _) $ _, Free ("a", _)],
   6.338 +(*+*)      [Const ("Groups.uminus_class.uminus", _) $ _, Free ("a", _)] (*already correct order*)
   6.339 +(*+*)    ] = map monom2list (poly2list t);
   6.340 +
   6.341 +  	val lls = map
   6.342 +           sort_varList ll;
   6.343 +
   6.344 +(*+*)val [ [Const ("Num.numeral_class.numeral", _) $ _, Free ("a", _)],
   6.345 +(*+*)      [Const ("Groups.uminus_class.uminus", _) $ _, Free ("a", _)] (*ERROR: NO swap here*)
   6.346 +(*+*)    ] = map sort_varList ll;
   6.347 +
   6.348 +       map sort_varList ll;
   6.349 +"~~~~~ val sort_varList , args:"; val (ts) = (nth 2 ll);
   6.350 +val sorted = sort var_ord ts;
   6.351 +
   6.352 +(*+*)if UnparseC.terms ts = "[\"- 2\", \"a\"]" andalso UnparseC.terms sorted = "[\"- 2\", \"a\"]"
   6.353 +(*+*)then () else error "sort var_ord  [\"- 2\", \"a\"]  CHANGED";
   6.354 +
   6.355 +
   6.356 +(*+*)if UnparseC.term (nth 1 ts) = "- 2" andalso get_basStr (nth 1 ts) = "-2"
   6.357 +(*+*)then () else error "get_basStr  - 2  CHANGED";
   6.358 +(*+*)if UnparseC.term (nth 2 ts) = "a" andalso get_basStr (nth 2 ts) = "a"
   6.359 +(*+*)then () else error "get_basStr  a  CHANGED";
   6.360 +
   6.361 +
   6.362 +"-------- rebuild fun is_addUnordered (1 + 2 * x \<up> 4 + - 4 * x \<up> 4 + x \<up> 8) ---------------------";
   6.363 +"-------- rebuild fun is_addUnordered (1 + 2 * x \<up> 4 + - 4 * x \<up> 4 + x \<up> 8) ---------------------";
   6.364 +"-------- rebuild fun is_addUnordered (1 + 2 * x \<up> 4 + - 4 * x \<up> 4 + x \<up> 8) ---------------------";
   6.365 +(* indentation partially indicates call hierarchy:
   6.366 +"~~~~~ fun is_addUnordered
   6.367 +"~~~~~~~ fun is_polyexp
   6.368 +"~~~~~~~ fun sort_monoms
   6.369 +"~~~~~~~~~ fun sort_monList
   6.370 +"~~~~~~~~~~~ fun koeff_degree_ord
   6.371 +"~~~~~~~~~~~~~ fun degree_ord
   6.372 +"~~~~~~~~~~~~~~~ fun dict_cond_ord
   6.373 +"~~~~~~~~~~~~~~~~~ fun var_ord_revPow
   6.374 +"~~~~~~~~~~~~~~~~~~~ fun get_basStr         used twice --vv
   6.375 +"~~~~~~~~~~~~~~~~~~~ fun get_potStr         used twice --vv
   6.376 +"~~~~~~~~~~~~~~~ fun monom_degree
   6.377 +"~~~~~~~~~~~~~ fun compare_koeff_ord
   6.378 +"~~~~~~~~~~~~~~~ fun get_koeff_of_mon
   6.379 +"~~~~~~~~~~~~~~~~~ fun koeff2ordStr
   6.380 +"~~~~~ fun is_multUnordered
   6.381 +"~~~~~~~ fun sort_variables
   6.382 +"~~~~~~~~~ fun sort_varList
   6.383 +"~~~~~~~~~~~ fun var_ord
   6.384 +"~~~~~~~~~~~~~ fun get_basStr               used twice --^^
   6.385 +"~~~~~~~~~~~~~ fun get_potStr               used twice --^^
   6.386 +*)
   6.387 +val t = TermC.str2term "(1 + 2 * x \<up> 4 + - 4 * x \<up> 4 + x \<up> 8) is_addUnordered";
   6.388 +
   6.389 +           eval_is_addUnordered "xxx" "yyy" t @{theory};
   6.390 +"~~~~~ fun eval_is_addUnordered , args:"; val ((thmid:string), _, 
   6.391 +		       (t as (Const("Poly.is_addUnordered", _) $ arg)), thy) =
   6.392 +  ("xxx", "yyy", t, @{theory});
   6.393 +
   6.394 +    (*if*) is_addUnordered arg;
   6.395 +"~~~~~ fun is_addUnordered , args:"; val (t) = (arg);
   6.396 +  ((is_polyexp t) andalso not (t = sort_monoms t));
   6.397 +
   6.398 +        (t = sort_monoms t);
   6.399 +"~~~~~~~ fun sort_monoms , args:"; val (t) = (t);
   6.400 +  	val ll =  map monom2list (poly2list t);
   6.401 +  	val lls =
   6.402 +
   6.403 +           sort_monList ll;
   6.404 +"~~~~~~~~~ fun sort_monList , args:"; val (ll) = (ll);
   6.405 +      val xxx = 
   6.406 +
   6.407 +            sort koeff_degree_ord ll(*return value*);
   6.408 +"~~~~~~~~~~~ fun koeff_degree_ord , args:"; val (ll) = (ll);
   6.409 +                 koeff_degree_ord: term list * term list -> order;
   6.410 +(*we check all elements at once..*)
   6.411 +val eee1 = ll |> nth 1;
   6.412 +val eee2 = ll |> nth 2;
   6.413 +val eee3 = ll |> nth 3;
   6.414 +val eee3 = ll |> nth 3;
   6.415 +val eee4 = ll |> nth 4;
   6.416 +
   6.417 +if UnparseC.terms eee1 = "[\"1\"]" then () else error "eee1 CHANGED";
   6.418 +if UnparseC.terms eee2 = "[\"2\", \"x \<up> 4\"]" then () else error "eee2 CHANGED";
   6.419 +if UnparseC.terms eee3 = "[\"- 4\", \"x \<up> 4\"]" then () else error "eee3 CHANGED";
   6.420 +if UnparseC.terms eee4 = "[\"x \<up> 8\"]" then () else error "eee4 CHANGED";
   6.421 +
   6.422 +if koeff_degree_ord (eee1, eee1) = EQUAL   then () else error "(eee1, eee1) CHANGED";
   6.423 +if koeff_degree_ord (eee1, eee2) = LESS    then () else error "(eee1, eee2) CHANGED";
   6.424 +if koeff_degree_ord (eee1, eee3) = LESS    then () else error "(eee1, eee3) CHANGED";
   6.425 +if koeff_degree_ord (eee1, eee4) = LESS    then () else error "(eee1, eee4) CHANGED";
   6.426 +
   6.427 +if koeff_degree_ord (eee2, eee1) = GREATER then () else error "(eee2, eee1) CHANGED";
   6.428 +if koeff_degree_ord (eee2, eee2) = EQUAL   then () else error "(eee2, eee4) CHANGED";
   6.429 +if koeff_degree_ord (eee2, eee3) = LESS    then () else error "(eee2, eee3) CHANGED";
   6.430 +if koeff_degree_ord (eee2, eee4) = LESS    then () else error "(eee2, eee4) CHANGED";
   6.431 +
   6.432 +if koeff_degree_ord (eee3, eee1) = GREATER then () else error "(eee3, eee1) CHANGED";
   6.433 +if koeff_degree_ord (eee3, eee2) = GREATER then () else error "(eee3, eee4) CHANGED";
   6.434 +if koeff_degree_ord (eee3, eee3) = EQUAL   then () else error "(eee3, eee3) CHANGED";
   6.435 +if koeff_degree_ord (eee3, eee4) = LESS    then () else error "(eee3, eee4) CHANGED";
   6.436 +
   6.437 +if koeff_degree_ord (eee4, eee1) = GREATER then () else error "(eee4, eee1) CHANGED";
   6.438 +if koeff_degree_ord (eee4, eee2) = GREATER then () else error "(eee4, eee4) CHANGED";
   6.439 +if koeff_degree_ord (eee4, eee3) = GREATER then () else error "(eee4, eee3) CHANGED";
   6.440 +if koeff_degree_ord (eee4, eee4) = EQUAL   then () else error "(eee4, eee4) CHANGED";
   6.441 +
   6.442 +"~~~~~~~~~~~~~ fun degree_ord , args:"; val () = ();
   6.443 +                   degree_ord: term list * term list -> order;
   6.444 +
   6.445 +if degree_ord (eee1, eee1) = EQUAL   then () else error "degree_ord (eee1, eee1) CHANGED";
   6.446 +if degree_ord (eee1, eee2) = LESS    then () else error "degree_ord (eee1, eee2) CHANGED";
   6.447 +if degree_ord (eee1, eee3) = LESS    then () else error "degree_ord (eee1, eee3) CHANGED";
   6.448 +if degree_ord (eee1, eee4) = LESS    then () else error "degree_ord (eee1, eee4) CHANGED";
   6.449 +
   6.450 +if degree_ord (eee2, eee1) = GREATER then () else error "degree_ord (eee2, eee1) CHANGED";
   6.451 +if degree_ord (eee2, eee2) = EQUAL   then () else error "degree_ord (eee2, eee4) CHANGED";
   6.452 +if degree_ord (eee2, eee3) = EQUAL   then () else error "degree_ord (eee2, eee3) CHANGED";
   6.453 +if degree_ord (eee2, eee4) = LESS    then () else error "degree_ord (eee2, eee4) CHANGED";
   6.454 +
   6.455 +if degree_ord (eee3, eee1) = GREATER then () else error "degree_ord (eee3, eee1) CHANGED";
   6.456 +if degree_ord (eee3, eee2) = EQUAL   then () else error "degree_ord (eee3, eee4) CHANGED";
   6.457 +if degree_ord (eee3, eee3) = EQUAL   then () else error "degree_ord (eee3, eee3) CHANGED";
   6.458 +if degree_ord (eee3, eee4) = LESS    then () else error "degree_ord (eee3, eee4) CHANGED";
   6.459 +
   6.460 +if degree_ord (eee4, eee1) = GREATER then () else error "degree_ord (eee4, eee1) CHANGED";
   6.461 +if degree_ord (eee4, eee2) = GREATER then () else error "degree_ord (eee4, eee4) CHANGED";
   6.462 +if degree_ord (eee4, eee3) = GREATER then () else error "degree_ord (eee4, eee3) CHANGED";
   6.463 +if degree_ord (eee4, eee4) = EQUAL   then () else error "degree_ord (eee4, eee4) CHANGED";
   6.464 +
   6.465 +"~~~~~~~~~~~~~~~ fun dict_cond_ord , args:"; val () = ();
   6.466 +dict_cond_ord: (term * term -> order) -> (term -> bool) -> term list * term list -> order;
   6.467 +dict_cond_ord var_ord_revPow: (term -> bool) -> term list * term list -> order;
   6.468 +dict_cond_ord var_ord_revPow is_nums: term list * term list -> order;
   6.469 +val xxx = dict_cond_ord var_ord_revPow is_nums;
   6.470 +(* output from:
   6.471 +fun var_ord_revPow (a,b: term) =
   6.472 + (@ {print} {a = "var_ord_revPow ", ab = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")"};
   6.473 +  @ {print} {z = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
   6.474 +  prod_ord string_ord string_ord_rev 
   6.475 +    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b)));
   6.476 +*)
   6.477 +if xxx (eee1, eee1) = EQUAL then () else error "dict_cond_ord ..(eee1, eee1) CHANGED";
   6.478 +if xxx (eee1, eee2) = LESS  then () else error "dict_cond_ord ..(eee1, eee2) CHANGED";
   6.479 +if xxx (eee1, eee3) = LESS  then () else error "dict_cond_ord ..(eee1, eee3) CHANGED";
   6.480 +if xxx (eee1, eee4) = LESS  then () else error "dict_cond_ord ..(eee1, eee4) CHANGED";
   6.481 +(*-------------------------------------------------------------------------------------*)
   6.482 +if xxx (eee2, eee1) = GREATER then () else error "dict_cond_ord ..(eee2, eee1) CHANGED";
   6.483 +if xxx (eee2, eee2) = EQUAL   then () else error "dict_cond_ord ..(eee2, eee2) CHANGED";
   6.484 +if xxx (eee2, eee3) = EQUAL   then () else error "dict_cond_ord ..(eee2, eee3) CHANGED";
   6.485 +if xxx (eee2, eee4) = GREATER then () else error "dict_cond_ord ..(eee2, eee4) CHANGED";
   6.486 +(*-------------------------------------------------------------------------------------*)
   6.487 +if xxx (eee3, eee1) = GREATER then () else error "dict_cond_ord ..(eee3, eee1) CHANGED";
   6.488 +if xxx (eee3, eee2) = EQUAL   then () else error "dict_cond_ord ..(eee3, eee2) CHANGED";
   6.489 +if xxx (eee3, eee3) = EQUAL   then () else error "dict_cond_ord ..(eee3, eee3) CHANGED";
   6.490 +if xxx (eee3, eee4) = GREATER then () else error "dict_cond_ord ..(eee3, eee4) CHANGED";
   6.491 +(*-------------------------------------------------------------------------------------*)
   6.492 +if xxx (eee4, eee1) = GREATER then () else error "dict_cond_ord ..(eee4, eee1) CHANGED";
   6.493 +if xxx (eee4, eee2) = LESS    then () else error "dict_cond_ord ..(eee4, eee2) CHANGED";
   6.494 +if xxx (eee4, eee3) = LESS    then () else error "dict_cond_ord ..(eee4, eee3) CHANGED";
   6.495 +if xxx (eee4, eee4) = EQUAL   then () else error "dict_cond_ord ..(eee4, eee4) CHANGED";
   6.496 +(*-------------------------------------------------------------------------------------*)
   6.497 +
   6.498 +"~~~~~~~~~~~~~~~ fun monom_degree , args:"; val () = ();
   6.499 +(* we check all at once//*)
   6.500 +if monom_degree eee1 = 0 then () else error "monom_degree eee1 CHANGED";
   6.501 +if monom_degree eee2 = 4 then () else error "monom_degree eee2 CHANGED";
   6.502 +if monom_degree eee3 = 4 then () else error "monom_degree eee3 CHANGED";
   6.503 +if monom_degree eee4 = 8 then () else error "monom_degree eee4 CHANGED";
   6.504 +
   6.505 +"~~~~~~~~~~~~~ fun compare_koeff_ord , args:"; val () = ();
   6.506 +                   compare_koeff_ord: term list * term list -> order;
   6.507 +(* we check all at once//*)
   6.508 +if compare_koeff_ord (eee1, eee1) = EQUAL   then () else error "_koeff_ord (eee1, eee1) CHANGED";
   6.509 +if compare_koeff_ord (eee1, eee2) = LESS    then () else error "_koeff_ord (eee1, eee2) CHANGED";
   6.510 +if compare_koeff_ord (eee1, eee3) = LESS    then () else error "_koeff_ord (eee1, eee3) CHANGED";
   6.511 +if compare_koeff_ord (eee1, eee4) = GREATER then () else error "_koeff_ord (eee1, eee4) CHANGED";
   6.512 +
   6.513 +if compare_koeff_ord (eee2, eee1) = GREATER then () else error "_koeff_ord (eee2, eee1) CHANGED";
   6.514 +if compare_koeff_ord (eee2, eee2) = EQUAL   then () else error "_koeff_ord (eee2, eee2) CHANGED";
   6.515 +if compare_koeff_ord (eee2, eee3) = LESS    then () else error "_koeff_ord (eee2, eee3) CHANGED";
   6.516 +if compare_koeff_ord (eee2, eee4) = GREATER then () else error "_koeff_ord (eee2, eee4) CHANGED";
   6.517 +
   6.518 +if compare_koeff_ord (eee3, eee1) = GREATER then () else error "_koeff_ord (eee3, eee1) CHANGED";
   6.519 +if compare_koeff_ord (eee3, eee2) = GREATER then () else error "_koeff_ord (eee3, eee2) CHANGED";
   6.520 +if compare_koeff_ord (eee3, eee3) = EQUAL   then () else error "_koeff_ord (eee3, eee3) CHANGED";
   6.521 +if compare_koeff_ord (eee3, eee4) = GREATER then () else error "_koeff_ord (eee3, eee4) CHANGED";
   6.522 +
   6.523 +if compare_koeff_ord (eee4, eee1) = LESS    then () else error "_koeff_ord (eee4, eee1) CHANGED";
   6.524 +if compare_koeff_ord (eee4, eee2) = LESS    then () else error "_koeff_ord (eee4, eee2) CHANGED";
   6.525 +if compare_koeff_ord (eee4, eee3) = LESS    then () else error "_koeff_ord (eee4, eee3) CHANGED";
   6.526 +if compare_koeff_ord (eee4, eee4) = EQUAL   then () else error "_koeff_ord (eee4, eee4) CHANGED";
   6.527 +
   6.528 +"~~~~~~~~~~~~~~~ fun get_koeff_of_mon , args:"; val () = ();
   6.529 +           get_koeff_of_mon: term list -> term option;
   6.530 +
   6.531 +val SOME xxx1 = get_koeff_of_mon eee1;
   6.532 +val SOME xxx2 = get_koeff_of_mon eee2;
   6.533 +val SOME xxx3 = get_koeff_of_mon eee3;
   6.534 +val NONE = get_koeff_of_mon eee4;
   6.535 +
   6.536 +if UnparseC.term xxx1 = "1"   then () else error "get_koeff_of_mon eee1 CHANGED";
   6.537 +if UnparseC.term xxx2 = "2"   then () else error "get_koeff_of_mon eee2 CHANGED";
   6.538 +if UnparseC.term xxx3 = "- 4" then () else error "get_koeff_of_mon eee3 CHANGED";
   6.539 +
   6.540 +"~~~~~~~~~~~~~~~~~ fun koeff2ordStr , args:"; val () = ();
   6.541 +                       koeff2ordStr: term option -> string;
   6.542 +if koeff2ordStr (SOME xxx1) = "1"   then () else error "koeff2ordStr eee1 CHANGED";
   6.543 +if koeff2ordStr (SOME xxx2) = "2"   then () else error "koeff2ordStr eee2 CHANGED";
   6.544 +if koeff2ordStr (SOME xxx3) = "40"  then () else error "koeff2ordStr eee3 CHANGED";
   6.545 +if koeff2ordStr NONE        = "---" then () else error "koeff2ordStr eee4 CHANGED";
   6.546 +
   6.547 +
   6.548 +"-------- fun is_addUnordered (x \<up> 2 * y \<up> 2 + x \<up> 3 * y) --------------------------------------";
   6.549 +"-------- fun is_addUnordered (x \<up> 2 * y \<up> 2 + x \<up> 3 * y) --------------------------------------";
   6.550 +"-------- fun is_addUnordered (x \<up> 2 * y \<up> 2 + x \<up> 3 * y) --------------------------------------";
   6.551 +val t = TermC.str2term "x \<up> 2 * y \<up> 2 + x * x \<up> 2 * y";
   6.552 +Rewrite.trace_on := false;
   6.553 +val SOME (t, _) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
   6.554 +   UnparseC.term t = "x \<up> 2 * y \<up> 2 + x \<up> 3 * y";
   6.555 +if UnparseC.term t = "x \<up> 3 * y + x \<up> 2 * y \<up> 2" then ()
   6.556 +else error "poly.sml: diff.behav. in make_polynomial 23";
   6.557 +
   6.558 +(** )
   6.559 +##  rls: order_add_rls_ on: x \<up> 2 * y \<up> 2 + x \<up> 3 * y 
   6.560 +###  rls: order_add_ on: x \<up> 2 * y \<up> 2 + x \<up> 3 * y 
   6.561 +######  rls: Rule_Set.empty-is_addUnordered on: x \<up> 2 * y \<up> 2 + x \<up> 3 * y is_addUnordered 
   6.562 +#######  try calc: "Poly.is_addUnordered" 
   6.563 +########  eval asms: "x \<up> 2 * y \<up> 2 + x \<up> 3 * y is_addUnordered = False"  (*isa*)
   6.564 +                                                                              True"   (*isa2*)
   6.565 +#######  calc. to: False  (*isa*)
   6.566 +                   True   (*isa2*)
   6.567 +( **)
   6.568 +        if is_addUnordered (TermC.str2term "x \<up> 2 * y \<up> 2 + x \<up> 3 * y ::real") then ()
   6.569 +else error"is_addUnordered  x \<up> 2 * y \<up> 2 + x \<up> 3 * y"; (*isa == isa2*)
   6.570 +"~~~~~ fun is_addUnordered , args:"; val (t) = (TermC.str2term "x \<up> 2 * y \<up> 2 + x \<up> 3 * y ::real");
   6.571 +      ((is_polyexp t) andalso not (t = sort_monoms t)) = false;  (*isa == isa2*)
   6.572 +
   6.573 +(*+*) if is_polyexp t = true then () else error "is_polyexp  x \<up> 2 * y \<up> 2 + x \<up> 3 * y";
   6.574 +
   6.575 +(*+*) if                           t = sort_monoms t = false then () else error "sort_monoms 123";
   6.576 +"~~~~~~~ fun sort_monoms , args:"; val (t) = (t);
   6.577 +  	val ll =  map monom2list (poly2list t);
   6.578 +  	val lls = sort_monList ll;
   6.579 +
   6.580 +(*+*)map UnparseC.terms lls = ["[\"x \<up> 2\", \"y \<up> 2\"]", "[\"x \<up> 3\", \"y\"]"]; (*isa*)
   6.581 +(*+*)map UnparseC.terms lls = ["[\"x \<up> 3\", \"y\"]", "[\"x \<up> 2\", \"y \<up> 2\"]"]; (*isa2*)
   6.582 +
   6.583 +"~~~~~~~~~~~ fun koeff_degree_ord , args:"; val () = ();
   6.584 +(* we check all elements at once..*)
   6.585 +val eee1 = ll |> nth 1; UnparseC.terms eee1 = "[\"x \<up> 2\", \"y \<up> 2\"]";
   6.586 +val eee2 = ll |> nth 2; UnparseC.terms eee2 = "[\"x \<up> 3\", \"y\"]";    
   6.587 +
   6.588 +(*1*)if koeff_degree_ord (eee1, eee1) = EQUAL   then () else error "(eee1, eee1) CHANGED";
   6.589 +(*2*)if koeff_degree_ord (eee1, eee2) = GREATER then () else error "(eee1, eee2) CHANGED";
   6.590 +(*3*)if koeff_degree_ord (eee2, eee1) = LESS    then () else error "(eee2, eee1) CHANGED"; (*isa*)
   6.591 +(*4*)if koeff_degree_ord (eee2, eee2) = EQUAL   then () else error "(eee2, eee2) CHANGED";
   6.592 +(* isa -- isa2:
   6.593 +(*1*){a = "var_ord_revPow ", at_bt = "(x \<up> 2, x \<up> 2)", sort_args = "(x, 2), (x, 2)"}                                           (*isa == isa2*)
   6.594 +(*1*){a = "var_ord_revPow ", at_bt = "(y \<up> 2, y \<up> 2)", sort_args = "(y, 2), (y, 2)"}                                            (*isa == isa2*)
   6.595 +(*1*){a = "compare_koeff_ord ", ats_bts = "([\"x \<up> 2\", \"y \<up> 2\"], [\"x \<up> 2\", \"y \<up> 2\"])", sort_args = "(---, ---)"} (*isa == isa2*) 
   6.596 +
   6.597 +(*2*){a = "var_ord_revPow ", at_bt = "(x \<up> 2, x \<up> 3)", sort_args = "(x, 2), (x, 3)"}                                           (*isa == isa2*) 
   6.598 +
   6.599 +(*3*)k{a = "var_ord_revPow ", at_bt = "(x \<up> 3, x \<up> 2)", sort_args = "(x, 3), (x, 2)"}                                           (*isa == isa2*) 
   6.600 +
   6.601 +(*4*){a = "var_ord_revPow ", at_bt = "(x \<up> 3, x \<up> 3)", sort_args = "(x, 3), (x, 3)"}                                           (*isa == isa2*)
   6.602 +(*4*){a = "var_ord_revPow ", at_bt = "(y, y)", sort_args = "(y, ---), (y, ---)"}                                                       (*isa == isa2*)
   6.603 +(*4*){a = "compare_koeff_ord ", ats_bts = "([\"x \<up> 3\", \"y\"], [\"x \<up> 3\", \"y\"])", sort_args = "(---, ---)"}                (*isa == isa2*) 
   6.604 +val it = true: bool
   6.605 +val it = true: bool
   6.606 +val it = true: bool
   6.607 +val it = true: bool*)
   6.608 +
   6.609 +"~~~~~~~~~~~~~ fun degree_ord , args:"; val () = ();
   6.610 +(*1*)if degree_ord (eee1, eee1) = EQUAL   then () else error "degree_ord (eee1, eee1) CHANGED";
   6.611 +(*2*)if degree_ord (eee1, eee2) = GREATER    then () else error "degree_ord (eee1, eee2) CHANGED";(*isa*)
   6.612 +(*{a = "int_ord (4, 10003) = ", z = LESS}      isa
   6.613 +  {a = "int_ord (4, 4) = ", z = EQUAL}         isa2*)
   6.614 +(*3*)if degree_ord (eee2, eee1) = LESS then () else error "degree_ord (eee2, eee1) CHANGED";(*isa*)
   6.615 +(*4*)if degree_ord (eee2, eee2) = EQUAL   then () else error "degree_ord (eee2, eee2) CHANGED";
   6.616 +(* isa -- isa2:
   6.617 +(*1*){a = "int_ord (10002, 10002) = ", z = EQUAL}                                                          (*isa*)
   6.618 +     {a = "int_ord (4, 4) = ", z = EQUAL}                                                                    (*isa2*)
   6.619 +(*1*){a = "dict_cond_ord", args = "([\"x \<up> 2\", \"y \<up> 2\"], [\"x \<up> 2\", \"y \<up> 2\"])", is_nums = "(false, false)"}  (*isa*)
   6.620 +(*1*){a = "var_ord_revPow ", at_bt = "(x \<up> 2, x \<up> 2)", sort_args = "(x, 2), (x, 2)"}               (*isa*)
   6.621 +(*1*){a = "dict_cond_ord", args = "([\"y \<up> 2\"], [\"y \<up> 2\"])", is_nums = "(false, false)"}        (*isa*)
   6.622 +(*1*){a = "var_ord_revPow ", at_bt = "(y \<up> 2, y \<up> 2)", sort_args = "(y, 2), (y, 2)"}               (*isa*)
   6.623 +(*1*){a = "dict_cond_ord ([], [])"}                                                                        (*isa*)
   6.624 +
   6.625 +(*2*){a = "int_ord (10002, 10003) = ", z = LESS}                                                           (*isa*)
   6.626 +     {a = "int_ord (4, 4) = ", z = EQUAL}                                                                    (*isa2*)
   6.627 +     {a = "dict_cond_ord", args = "([\"x \<up> 2\", \"y \<up> 2\"], [\"x \<up> 3\", \"y\"])", is_nums = "(false, false)"} (*isa2*)
   6.628 +(*2*){a = "var_ord_revPow ", at_bt = "(x \<up> 2, x \<up> 3)", sort_args = "(x, 2), (x, 3)"}                 (*isa2*)
   6.629 +
   6.630 +(*3*){a = "int_ord (10003, 10002) = ", z = GREATER}                                                       (*isa*)
   6.631 +     {a = "int_ord (4, 4) = ", z = EQUAL}                                                                    (*isa2*)
   6.632 +     {a = "dict_cond_ord", args = "([\"x \<up> 3\", \"y\"], [\"x \<up> 2\", \"y \<up> 2\"])", is_nums = "(false, false)"}
   6.633 +(*3*){a = "var_ord_revPow ", at_bt = "(x \<up> 3, x \<up> 2)", sort_args = "(x, 3), (x, 2)"}                (*isa == isa2*)
   6.634 +
   6.635 +(*4*){a = "int_ord (10003, 10003) = ", z = EQUAL}                                                         (*isa*)
   6.636 +     {a = "int_ord (4, 4) = ", z = EQUAL}                                                                    (*isa2*)
   6.637 +(*4*){a = "dict_cond_ord", args = "([\"x \<up> 3\", \"y\"], [\"x \<up> 3\", \"y\"])", is_nums = "(false, false)"} (*isa == isa2*)
   6.638 +(*4*){a = "var_ord_revPow ", at_bt = "(x \<up> 3, x \<up> 3)", sort_args = "(x, 3), (x, 3)"}              (*isa == isa2*)
   6.639 +(*4*){a = "dict_cond_ord", args = "([\"y\"], [\"y\"])", is_nums = "(false, false)"}                       (*isa == isa2*)
   6.640 +(*4*){a = "var_ord_revPow ", at_bt = "(y, y)", sort_args = "(y, ---), (y, ---)"}                          (*isa == isa2*)
   6.641 +(*4*){a = "dict_cond_ord ([], [])"}                                                                       (*isa == isa2*)
   6.642 +
   6.643 +val it = true: bool
   6.644 +val it = false: bool
   6.645 +val it = false: bool
   6.646 +val it = true: bool
   6.647 +*)
   6.648 +
   6.649 +(*+*) if monom_degree eee1 = 4 then () else error "monom_degree [\"x \<up> 2\", \"y \<up> 2\"] CHANGED";
   6.650 +"~~~~~~~~~~~~~~~ fun monom_degree , args:"; val (l) = (eee1);
   6.651 +"~~~~~ fun counter , args:"; val (n, x :: xs) = (0, l); (*--------------------------OPEN local\* )
   6.652 +	    (*if*) (is_nums x) (*else*);
   6.653 +  val (Const ("Transcendental.powr", _) $ Free _ $ t) =
   6.654 +      (*case*) x (*of*); 
   6.655 +(*+*)UnparseC.term x = "x \<up> 2";
   6.656 +            (*if*) TermC.is_num t (*then*);
   6.657 +
   6.658 +           counter (t |> HOLogic.dest_number |> snd |> curry op + n, xs);
   6.659 +"~~~~~ fun counter , args:"; val (n, x :: xs) = (t |> HOLogic.dest_number |> snd |> curry op + n, xs);
   6.660 +	    (*if*) (is_nums x) (*else*);
   6.661 +  val (Const ("Transcendental.powr", _) $ Free _ $ t) =
   6.662 +      (*case*) x (*of*); 
   6.663 +(*+*)UnparseC.term x = "y \<up> 2";
   6.664 +            (*if*) TermC.is_num t (*then*);
   6.665 +
   6.666 +  val return =
   6.667 +      counter (t |> HOLogic.dest_number |> snd |> curry op + n, xs);
   6.668 +if return = 4 then () else error "monom_degree [\"x \<up> 2\", \"y \<up> 2\"] CHANGED";
   6.669 +( *---------------------------------------------------------------------------------OPEN local/*)
   6.670 +
   6.671 +(*+*)if monom_degree eee2 = 4 andalso monom_degree eee2 = 4 then ()
   6.672 +else error "monom_degree  [\"x \<up> 3\", \"y\"] CHANGED";
   6.673 +"~~~~~~~~~~~~~~~ fun monom_degree , args:"; val (l) = (eee2);
   6.674 +"~~~~~ fun counter , args:"; val (n, x :: xs) = (0, l); (*--------------------------OPEN local\* )
   6.675 +	    (*if*) (is_nums x) (*else*);
   6.676 +val (Const ("Transcendental.powr", _) $ Free _ $ t) =
   6.677 +      (*case*) x (*of*); 
   6.678 +(*+*)UnparseC.term x = "x \<up> 3";
   6.679 +            (*if*) TermC.is_num t (*then*);
   6.680 +
   6.681 +      counter (t |> HOLogic.dest_number |> snd |> curry op + n, xs);
   6.682 +"~~~~~ fun counter , args:"; val (n, x :: xs) = (t |> HOLogic.dest_number |> snd |> curry op + n, xs);
   6.683 +	    (*if*) (is_nums x) (*else*);
   6.684 +val _ = (*the default case*)
   6.685 +      (*case*) x (*of*); 
   6.686 +( *---------------------------------------------------------------------------------OPEN local/*)
   6.687 +
   6.688 +"~~~~~~~~~~~~~~~ fun dict_cond_ord , args:"; val () = ();
   6.689 +val xxx = dict_cond_ord var_ord_revPow is_nums;
   6.690 +(*1*)if xxx (eee1, eee1) = EQUAL   then () else error "dict_cond_ord (eee1, eee1) CHANGED";
   6.691 +(*2*)if xxx (eee1, eee2) = GREATER then () else error "dict_cond_ord (eee1, eee2) CHANGED";
   6.692 +(*3*)if xxx (eee2, eee1) = LESS    then () else error "dict_cond_ord (eee2, eee1) CHANGED";
   6.693 +(*4*)if xxx (eee2, eee2) = EQUAL   then () else error "dict_cond_ord (eee2, eee2) CHANGED";
   6.694 +
   6.695 +
   6.696 +"-------- check make_polynomial with simple terms ----------------------------------------------";
   6.697 +"-------- check make_polynomial with simple terms ----------------------------------------------";
   6.698 +"-------- check make_polynomial with simple terms ----------------------------------------------";
   6.699  "----- check 1 ---";
   6.700  val t = TermC.str2term "2*3*a";
   6.701  val SOME (t, _) = rewrite_set_ thy false make_polynomial t;
   6.702 @@ -332,9 +837,9 @@
   6.703  val SOME (t, _) = rewrite_set_ thy false make_polynomial t;
   6.704  if UnparseC.term t = "4 * a \<up> 2" then () else error "check make_polynomial 6";
   6.705  
   6.706 -"-------- fun is_multUnordered --------------------------";
   6.707 -"-------- fun is_multUnordered --------------------------";
   6.708 -"-------- fun is_multUnordered --------------------------";
   6.709 +"-------- fun is_multUnordered (x \<up> 2 * x) -----------------------------------------------------";
   6.710 +"-------- fun is_multUnordered (x \<up> 2 * x) -----------------------------------------------------";
   6.711 +"-------- fun is_multUnordered (x \<up> 2 * x) -----------------------------------------------------";
   6.712  val thy = @{theory "Isac_Knowledge"};
   6.713  "===== works for a simple example, see rewrite.sml -- fun app_rev ===";
   6.714  val t = TermC.str2term "x \<up> 2 * x";
   6.715 @@ -372,59 +877,271 @@
   6.716  val SOME (t, _) = rewrite_set_ thy true order_mult_ t;
   6.717  UnparseC.term t;
   6.718  
   6.719 -"-------- examples from textbook Schalk I ---------------";
   6.720 -"-------- examples from textbook Schalk I ---------------";
   6.721 -"-------- examples from textbook Schalk I ---------------";
   6.722 +
   6.723 +"-------- fun is_multUnordered (3 * a + - 2 * a) -----------------------------------------------";
   6.724 +"-------- fun is_multUnordered (3 * a + - 2 * a) -----------------------------------------------";
   6.725 +"-------- fun is_multUnordered (3 * a + - 2 * a) -----------------------------------------------";
   6.726 +val thy = @{theory "Isac_Knowledge"};
   6.727 +val t as (_ $ arg) = TermC.str2term "(3 * a + - 2 * a) is_multUnordered";
   6.728 +
   6.729 +(*+*)if UnparseC.term (sort_variables arg) = "3 * a + - 2 * a" andalso is_polyexp arg = true
   6.730 +(*+*)  andalso not (is_multUnordered arg)
   6.731 +(*+*)then () else error "sort_variables  3 * a + - 2 * a   CHANGED";
   6.732 +
   6.733 +case eval_is_multUnordered "xxx " "yyy " t thy of
   6.734 +  SOME
   6.735 +    ("xxx 3 * a + - 2 * a_",
   6.736 +      Const ("HOL.Trueprop", _) $ (Const ("HOL.eq", _) $ _ $
   6.737 +        Const ("HOL.False", _))) => ()
   6.738 +(*      Const ("HOL.True", _))) => ()   <<<<<--------------------------- this is false*)
   6.739 +| _ => error "eval_is_multUnordered  3 * a + -2 * a  CHANGED";
   6.740 +
   6.741 +"----- is_multUnordered --- (- 2 * a) is_multUnordered = False";
   6.742 +val t as (_ $ arg) = TermC.str2term "(- 2 * a) is_multUnordered";
   6.743 +
   6.744 +(*+*)if UnparseC.term (sort_variables arg) = "- 2 * a" andalso is_polyexp arg = true
   6.745 +(*+*)  andalso not (is_multUnordered arg)
   6.746 +(*+*)then () else error "sort_variables  - 2 * a   CHANGED";
   6.747 +
   6.748 +case eval_is_multUnordered "xxx " "yyy " t thy of
   6.749 +  SOME
   6.750 +    ("xxx - 2 * a_",
   6.751 +      Const ("HOL.Trueprop", _) $ (Const ("HOL.eq", _) $ _ $
   6.752 +        Const ("HOL.False", _))) => ()
   6.753 +| _ => error "eval_is_multUnordered  3 * a + -2 * a  CHANGED";
   6.754 +
   6.755 +"----- is_multUnordered --- (a) is_multUnordered = False";
   6.756 +val t as (_ $ arg) = TermC.str2term "(a) is_multUnordered";
   6.757 +
   6.758 +(*+*)if UnparseC.term (sort_variables arg) = "a" andalso is_polyexp arg = true
   6.759 +(*+*)  andalso not (is_multUnordered arg)
   6.760 +(*+*)then () else error "sort_variables   a   CHANGED";
   6.761 +
   6.762 +case eval_is_multUnordered "xxx " "yyy " t thy of
   6.763 +  SOME
   6.764 +    ("xxx a_",
   6.765 +      Const ("HOL.Trueprop", _) $ (Const ("HOL.eq", _) $ _ $
   6.766 +        Const ("HOL.False", _))) => ()
   6.767 +| _ => error "eval_is_multUnordered  3 * a + -2 * a  CHANGED";
   6.768 +
   6.769 +"----- is_multUnordered --- (- 2) is_multUnordered = False";
   6.770 +val t as (_ $ arg) = TermC.str2term "(- 2) is_multUnordered";
   6.771 +
   6.772 +(*+*)if UnparseC.term (sort_variables arg) = "- 2" andalso is_polyexp arg = true
   6.773 +(*+*)  andalso not (is_multUnordered arg)
   6.774 +(*+*)then () else error "sort_variables   - 2   CHANGED";
   6.775 +
   6.776 +case eval_is_multUnordered "xxx " "yyy " t thy of
   6.777 +  SOME
   6.778 +    ("xxx - 2_",
   6.779 +      Const ("HOL.Trueprop", _) $ (Const ("HOL.eq", _) $ _ $
   6.780 +        Const ("HOL.False", _))) => ()
   6.781 +| _ => error "eval_is_multUnordered  3 * a + -2 * a  CHANGED";
   6.782 +
   6.783 +
   6.784 +"-------- fun is_multUnordered (x - a) \<up> 3 -----------------------------------------------------";
   6.785 +"-------- fun is_multUnordered (x - a) \<up> 3 -----------------------------------------------------";
   6.786 +"-------- fun is_multUnordered (x - a) \<up> 3 -----------------------------------------------------";
   6.787 +(*  ca.line 45 of Rewrite.trace_on:
   6.788 +##  rls: order_mult_rls_ on: 
   6.789 +  x \<up> 3 + 3 * x \<up> 2 * (- 1 * a) + 3 * x * ((- 1) \<up> 2 * a \<up> 2) + (- 1) \<up> 3 * a \<up> 3 
   6.790 +###  rls: order_mult_ on:
   6.791 +  x \<up> 3 + 3 * x \<up> 2 * (- 1 * a) + 3 * x * ((- 1) \<up> 2 * a \<up> 2) + (- 1) \<up> 3 * a \<up> 3 
   6.792 +######  rls: Rule_Set.empty-is_multUnordered on: 
   6.793 +  x \<up> 3 + 3 * x \<up> 2 * (- 1 * a) + 3 * x * ((- 1) \<up> 2 * a \<up> 2) + (- 1) \<up> 3 * a \<up> 3 is_multUnordered 
   6.794 +#######  try calc: "Poly.is_multUnordered" 
   6.795 +########  eval asms: 
   6.796 +N:x \<up> 3 + 3 * x \<up> 2 * (- 1 * a) + 3 * x * ((- 1) \<up> 2 * a \<up> 2) + (- 1) \<up> 3 * a \<up> 3 is_multUnordered = True" 
   6.797 +-------------------------------------------------------------------------------------------------==
   6.798 +O:x \<up> 3 + 3 * x \<up> 2 * (-1 * a)  + 3 * x * (-1    \<up> 2 * a \<up> 2) + -1    \<up> 3 * a \<up> 3 is_multUnordered = True" 
   6.799 +#######  calc. to: True 
   6.800 +#######  try calc: "Poly.is_multUnordered" 
   6.801 +#######  try calc: "Poly.is_multUnordered" 
   6.802 +###  rls: order_mult_ on: 
   6.803 +N:x \<up> 3 + - 1 * (3 * (a * x \<up> 2))  +  3 * (a \<up> 2 * (x * (- 1) \<up> 2))  +  a \<up> 3 * (- 1) \<up> 3 
   6.804 +--------+--------------------------+---------------------------------+---------------------------<>
   6.805 +O:x \<up> 3 + -1  * (3 * (a * x \<up> 2))  +  -1 \<up> 2 * (3 * (a \<up> 2 * x))     +  -1 \<up> 3 * a \<up> 3 
   6.806 +-------------------------------------------------------------------------------------------------<>
   6.807 +*)
   6.808 +val t = TermC.str2term "x \<up> 3 + 3 * x \<up> 2 * (- 1 * a) + 3 * x * ((- 1) \<up> 2 * a \<up> 2) + (- 1) \<up> 3 * a \<up> 3";
   6.809 +(*
   6.810 +"~~~~~ fun is_multUnordered
   6.811 +"~~~~~~~ fun sort_variables
   6.812 +"~~~~~~~~~ val sort_varList
   6.813 +*)
   6.814 +"~~~~~ fun is_multUnordered , args:"; val (t) = (t);
   6.815 +     is_polyexp t = true;
   6.816 +
   6.817 +  val return =
   6.818 +             sort_variables t;
   6.819 +(*+*)if UnparseC.term return =
   6.820 +(*+*)  "x \<up> 3 + - 1 * (3 * (a * x \<up> 2)) +\n(- 1) \<up> 2 * (3 * (a \<up> 2 * x)) +\n(- 1) \<up> 3 * a \<up> 3"
   6.821 +(*+*)then () else error "sort_variables  x \<up> 3 + - 1 * (3 * (a * x \<up> 2)) .. CHANGED";
   6.822 +
   6.823 +"~~~~~~~ fun sort_variables , args:"; val (t) = (t);
   6.824 +  	val ll = map monom2list (poly2list t);
   6.825 +  	val lls = map sort_varList ll; 
   6.826 +
   6.827 +(*+*)val ori3 = nth 3 ll;
   6.828 +(*+*)val mon3 = nth 3 lls;
   6.829 +
   6.830 +(*+*)if UnparseC.terms (nth 1 ll) = "[\"x \<up> 3\"]" then () else error "nth 1 ll";
   6.831 +(*+*)if UnparseC.terms (nth 2 ll) = "[\"3\", \"x \<up> 2\", \"- 1\", \"a\"]" then () else error "nth 3 ll";
   6.832 +(*+*)if UnparseC.terms ori3       = "[\"3\", \"x\", \"(- 1) \<up> 2\", \"a \<up> 2\"]" then () else error "nth 3 ll";
   6.833 +(*+*)if UnparseC.terms (nth 4 ll) = "[\"(- 1) \<up> 3\", \"a \<up> 3\"]" then () else error "nth 4 ll";
   6.834 +
   6.835 +(*+*)if UnparseC.terms (nth 1 lls) = "[\"x \<up> 3\"]" then () else error "nth 1 lls";
   6.836 +(*+*)if UnparseC.terms (nth 2 lls) = "[\"- 1\", \"3\", \"a\", \"x \<up> 2\"]" then () else error "nth 2 lls";
   6.837 +(*+*)if UnparseC.terms mon3 = "[\"(- 1) \<up> 2\", \"3\", \"a \<up> 2\", \"x\"]" then () else error "nth 3 lls";
   6.838 +(*+*)if UnparseC.terms (nth 4 lls) = "[\"(- 1) \<up> 3\", \"a \<up> 3\"]" then () else error "nth 4 lls";
   6.839 +
   6.840 +"~~~~~~~~~ val sort_varList , args:"; val (ori3) = (ori3);
   6.841 +(* Output below with:
   6.842 +val sort_varList = sort var_ord;
   6.843 +fun var_ord (a,b: term) = 
   6.844 +(@{print} {a = "var_ord ", a_b = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")",
   6.845 +   sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
   6.846 +  prod_ord string_ord string_ord 
   6.847 +  ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b))
   6.848 +);
   6.849 +*)
   6.850 +(*+*)val xxx = sort_varList ori3;
   6.851 +(*
   6.852 +{a = "sort_varList", args = "[\"3\", \"x\", \"(- 1) \<up> 2\", \"a \<up> 2\"]"} (*isa*)
   6.853 +{a = "sort_varList", args = "[\"3\", \"x\", \"(- 1) \<up> 2\", \"a \<up> 2\"]"} (*isa2*)
   6.854 +      
   6.855 +isa                                            isa2
   6.856 +{a = "var_ord ", a_b = "(3, x)"}               {a = "var_ord ", a_b = "(3, x)"}
   6.857 +  {sort_args = "(|||, ||||||), (x, ---)"}        {sort_args = "(|||, ||||||), (x, ---)"}
   6.858 +{a = "var_ord ", a_b = "(x, (- 1) \<up> 2)"}       {a = "var_ord ", a_b = "(x, (- 1) \<up> 2)"}
   6.859 +  {sort_args = "(x, ---), (|||, ||||||)"}        {sort_args = "(x, ---), (|||, ||||||)"}
   6.860 +{a = "var_ord ", a_b = "((- 1) \<up> 2, a \<up> 2)"}   {a = "var_ord ", a_b = "((- 1) \<up> 2, a \<up> 2)"}
   6.861 +  {sort_args = "(|||, ||||||), (a, 2)"}          {sort_args = "(|||, ||||||), (a, |||)"}
   6.862 +                                  ^^^                                             ^^^
   6.863 +
   6.864 +{a = "var_ord ", a_b = "(x, a \<up> 2)"}           {a = "var_ord ", a_b = "(x, a \<up> 2)"}
   6.865 +  {sort_args = "(x, ---), (a, 2)"}               {sort_args = "(x, ---), (a, |||)"}
   6.866 +                             ^^^                                             ^^^
   6.867 +{a = "var_ord ", a_b = "(x, (- 1) \<up> 2)"}       {a = "var_ord ", a_b = "(x, (- 1) \<up> 2)"}
   6.868 +  {sort_args = "(x, ---), (|||, ||||||)"}        {sort_args = "(x, ---), (|||, ||||||)"}
   6.869 +{a = "var_ord ", a_b = "(3, (- 1) \<up> 2)"}       {a = "var_ord ", a_b = "(3, (- 1) \<up> 2)"}
   6.870 +  {sort_args = "(|||, ||||||), (|||, ||||||)"}   {sort_args = "(|||, ||||||), (|||, ||||||)"}
   6.871 +*)
   6.872 +
   6.873 +
   6.874 +"-------- fun is_multUnordered  b * a * a ------------------------------------------------------";
   6.875 +"-------- fun is_multUnordered  b * a * a ------------------------------------------------------";
   6.876 +"-------- fun is_multUnordered  b * a * a ------------------------------------------------------";
   6.877 +val t = TermC.str2term "b * a * a";
   6.878 +val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
   6.879 +if UnparseC.term t = "a \<up> 2 * b" then ()
   6.880 +else error "poly.sml: diff.behav. in make_polynomial 21";
   6.881 +
   6.882 +"~~~~~ fun is_multUnordered , args:"; val (t) = (@{term "b * a * a::real"});
   6.883 +    ((is_polyexp t) andalso not (t = sort_variables t)) = true;  (*isa == isa2*)
   6.884 +
   6.885 +(*+*)if    is_polyexp t then () else error "is_polyexp a \<up> 2 * b CHANGED";
   6.886 +"~~~~~ fun is_polyexp , args:"; val (Const ("Groups.times_class.times",_) $ num $ Free _) = (t);
   6.887 +    (*if*) TermC.is_num num (*else*);
   6.888 +
   6.889 +"~~~~~ fun is_polyexp , args:"; val (Const ("Groups.times_class.times",_) $ num $ Free _) = (num);
   6.890 +    (*if*) TermC.is_num num (*else*);
   6.891 +      (*if*) TermC.is_variable num (*then*);
   6.892 +
   6.893 +                           val xxx = sort_variables t;
   6.894 +(*+*)if UnparseC.term xxx = "a * (a * b)" then () else error "sort_variables a \<up> 2 * b CHANGED";
   6.895 +
   6.896 +
   6.897 +"-------- fun is_multUnordered 2*3*a -----------------------------------------------------------";
   6.898 +"-------- fun is_multUnordered 2*3*a -----------------------------------------------------------";
   6.899 +"-------- fun is_multUnordered 2*3*a -----------------------------------------------------------";
   6.900 +val t = TermC.str2term "2*3*a";
   6.901 +val SOME (t', _) = rewrite_set_ thy false make_polynomial t;
   6.902 +(*+*)if UnparseC.term t' = "6 * a" then () else error "rewrite_set_ 2*3*a CHANGED";
   6.903 +(*
   6.904 +##  try calc: "Groups.times_class.times" 
   6.905 +##  rls: order_mult_rls_ on: 6 * a 
   6.906 +###  rls: order_mult_ on: 6 * a 
   6.907 +######  rls: Rule_Set.empty-is_multUnordered on: 6 * a is_multUnordered 
   6.908 +#######  try calc: "Poly.is_multUnordered" 
   6.909 +########  eval asms: "6 * a is_multUnordered = True"    (*isa*)
   6.910 +                                             = False"   (*isa2*)
   6.911 +#######  calc. to: True  (*isa*)
   6.912 +                   False (*isa2*)
   6.913 +*)
   6.914 +val t = TermC.str2term "(6 * a) is_multUnordered"; 
   6.915 +val SOME
   6.916 +    (_, t') =
   6.917 +           eval_is_multUnordered "xxx" () t @{theory};
   6.918 +(*+*)if UnparseC.term t' = "6 * a is_multUnordered = False" then () 
   6.919 +(*+*)else error "6 * a is_multUnordered = False CHANGED";
   6.920 +
   6.921 +"~~~~~ fun eval_is_multUnordered , args:"; val ((thmid:string), _,
   6.922 +		       (t as (Const("Poly.is_multUnordered", _) $ arg)), thy) = ("xxx", (), t, @{theory});
   6.923 +    (*if*) is_multUnordered arg (*else*);
   6.924 +
   6.925 +"~~~~~ fun is_multUnordered , args:"; val (t) = (arg);
   6.926 +  val return =
   6.927 +     ((is_polyexp t) andalso not (t = sort_variables t));
   6.928 +
   6.929 +(*+*)return = false;                                             (*isa*)
   6.930 +(*+*)  is_polyexp t = true;                                      (*isa*)
   6.931 +(*+*)                        not (t = sort_variables t) = false; (*isa*)
   6.932 +
   6.933 +                            val xxx = sort_variables t;
   6.934 +(*+*)if UnparseC.term xxx = "6 * a" then () else error "2*3*a is_multUnordered CHANGED";
   6.935 +
   6.936 +
   6.937 +"-------- examples from textbook Schalk I ------------------------------------------------------";
   6.938 +"-------- examples from textbook Schalk I ------------------------------------------------------";
   6.939 +"-------- examples from textbook Schalk I ------------------------------------------------------";
   6.940  "-----SPB Schalk I p.63 No.267b ---";
   6.941 -(*associate poly* )
   6.942  val t = TermC.str2term "(5*x \<up> 2 + 3) * (2*x \<up> 7 + 3) - (3*x \<up> 5 + 8) * (6*x \<up> 4 - 1)";
   6.943 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
   6.944 -if (UnparseC.term t) = "17 + 15 * x \<up> 2 + -48 * x \<up> 4 + 3 * x \<up> 5 + 6 * x \<up> 7 + -8 * x \<up> 9"
   6.945 +val SOME (t, _) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
   6.946 +if UnparseC.term t = "17 + 15 * x \<up> 2 + - 48 * x \<up> 4 + 3 * x \<up> 5 + 6 * x \<up> 7 +\n- 8 * x \<up> 9"
   6.947  then () else error "poly.sml: diff.behav. in make_polynomial 1";
   6.948  
   6.949  "-----SPB Schalk I p.63 No.275b ---";
   6.950  val t = TermC.str2term "(3*x \<up> 2 - 2*x*y + y \<up> 2) * (x \<up> 2 - 2*y \<up> 2)";
   6.951  val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
   6.952 -if (UnparseC.term t) = ("3 * x \<up> 4 + -2 * x \<up> 3 * y + -5 * x \<up> 2 * y \<up> 2 + " ^
   6.953 -  "4 * x * y \<up> 3 +\n-2 * y \<up> 4")
   6.954 +if UnparseC.term t = 
   6.955 +  "3 * x \<up> 4 + - 2 * x \<up> 3 * y + - 5 * x \<up> 2 * y \<up> 2 +\n4 * x * y \<up> 3 +\n- 2 * y \<up> 4"
   6.956  then () else error "poly.sml: diff.behav. in make_polynomial 2";
   6.957  
   6.958  "-----SPB Schalk I p.63 No.279b ---";
   6.959  val t = TermC.str2term "(x-a)*(x-b)*(x-c)*(x-d)";
   6.960  val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
   6.961 -if (UnparseC.term t) = 
   6.962 -  ("a * b * c * d + -1 * a * b * c * x + -1 * a * b * d * x + a * b * x \<up> 2 +\n" ^
   6.963 -  "-1 * a * c * d * x +\na * c * x \<up> 2 +\na * d * x \<up> 2 +\n-1 * a * x \<up> 3 +\n" ^
   6.964 -  "-1 * b * c * d * x +\nb * c * x \<up> 2 +\nb * d * x \<up> 2 +\n-1 * b * x \<up> 3 +\n" ^
   6.965 -  "c * d * x \<up> 2 +\n-1 * c * x \<up> 3 +\n-1 * d * x \<up> 3 +\nx \<up> 4")
   6.966 +if UnparseC.term t = 
   6.967 +  ("a * b * c * d + - 1 * a * b * c * x + - 1 * a * b * d * x +\na * b * x \<up> 2 +\n" ^
   6.968 +  "- 1 * a * c * d * x +\na * c * x \<up> 2 +\na * d * x \<up> 2 +\n- 1 * a * x \<up> 3 +\n" ^
   6.969 +  "- 1 * b * c * d * x +\nb * c * x \<up> 2 +\nb * d * x \<up> 2 +\n- 1 * b * x \<up> 3 +\nc" ^
   6.970 +  " * d * x \<up> 2 +\n- 1 * c * x \<up> 3 +\n- 1 * d * x \<up> 3 +\nx \<up> 4")
   6.971  then () else error "poly.sml: diff.behav. in make_polynomial 3";
   6.972 -( *associate poly*)
   6.973 +(*associate poly*)
   6.974  
   6.975  "-----SPB Schalk I p.63 No.291 ---";
   6.976  val t = TermC.str2term "(5+96*x \<up> 3+8*x*(-4+(7- 3*x)*4*x))*(5*(2- 3*x)- (-15*x*(-8*x- 5)))";
   6.977  val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
   6.978 -if (UnparseC.term t) = "50 + -770 * x + 4520 * x \<up> 2 + -16320 * x \<up> 3 + -26880 * x \<up> 4"
   6.979 +if UnparseC.term t = "50 + - 770 * x + 4520 * x \<up> 2 + - 16320 * x \<up> 3 +\n- 26880 * x \<up> 4"
   6.980  then () else error "poly.sml: diff.behav. in make_polynomial 4";
   6.981  
   6.982 -(*associate poly* )
   6.983  "-----SPB Schalk I p.64 No.295c ---";
   6.984  val t = TermC.str2term "(13*a \<up> 4*b \<up> 9*c - 12*a \<up> 3*b \<up> 6*c \<up> 9) \<up> 2";
   6.985  val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
   6.986 -if (UnparseC.term t) = ("169 * a \<up> 8 * b \<up> 18 * c \<up> 2 + -312 * a \<up> 7 * b \<up> 15 * c \<up> 10" ^
   6.987 -  " +\n144 * a \<up> 6 * b \<up> 12 * c \<up> 18")
   6.988 +if UnparseC.term t =
   6.989 +  "169 * a \<up> 8 * b \<up> 18 * c \<up> 2 +\n- 312 * a \<up> 7 * b \<up> 15 * c \<up> 10 +\n144 * a \<up> 6 * b \<up> 12 * c \<up> 18"
   6.990  then ()else error "poly.sml: diff.behav. in make_polynomial 5";
   6.991 -( *associate poly*)
   6.992  
   6.993  "-----SPB Schalk I p.64 No.299a ---";
   6.994  val t = TermC.str2term "(x - y)*(x + y)";
   6.995  val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
   6.996 -if (UnparseC.term t) = "x \<up> 2 + -1 * y \<up> 2"
   6.997 +if UnparseC.term t = "x \<up> 2 + - 1 * y \<up> 2"
   6.998  then () else error "poly.sml: diff.behav. in make_polynomial 6";
   6.999  
  6.1000  "-----SPB Schalk I p.64 No.300c ---";
  6.1001  val t = TermC.str2term "(3*x \<up> 2*y - 1)*(3*x \<up> 2*y + 1)";
  6.1002  val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
  6.1003 -if (UnparseC.term t) = "-1 + 9 * x \<up> 4 * y \<up> 2"
  6.1004 +if UnparseC.term t = "- 1 + 9 * x \<up> 4 * y \<up> 2"
  6.1005  then () else error "poly.sml: diff.behav. in make_polynomial 7";
  6.1006  
  6.1007  "-----SPB Schalk I p.64 No.302 ---";
  6.1008 @@ -438,32 +1155,35 @@
  6.1009  "-----SPB Schalk I p.64 No.306a ---";
  6.1010  val t = TermC.str2term "((x \<up> 2 + 1)*(x \<up> 2 - 1)) \<up> 2";
  6.1011  val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1012 -if (UnparseC.term t) = "1 + 2 * x \<up> 4 + 2 * -2 * x \<up> 4 + x \<up> 8" then ()
  6.1013 +if UnparseC.term t = "1 + 2 * x \<up> 4 + 2 * - 2 * x \<up> 4 + x \<up> 8" then ()
  6.1014  else error "poly.sml: diff.behav. in 2 * x \<up> 4 + 2 * -2 * x \<up> 4 = -2 * x \<up> 4";
  6.1015  
  6.1016  (*WN071729 when reducing "rls reduce_012_" for Schaerding,
  6.1017  the above resulted in the term below ... but reduces from then correctly*)
  6.1018  val t = TermC.str2term "1 + 2 * x \<up> 4 + 2 * -2 * x \<up> 4 + x \<up> 8";
  6.1019  val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1020 -if (UnparseC.term t) = "1 + -2 * x \<up> 4 + x \<up> 8"
  6.1021 +if UnparseC.term t = "1 + - 2 * x \<up> 4 + x \<up> 8"
  6.1022  then () else error "poly.sml: diff.behav. in make_polynomial 9b";
  6.1023  
  6.1024  "-----SPB Schalk I p.64 No.296a ---";
  6.1025  val t = TermC.str2term "(x - a) \<up> 3";
  6.1026  val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1027 -if (UnparseC.term t) = "-1 * a \<up> 3 + 3 * a \<up> 2 * x + -3 * a * x \<up> 2 + x \<up> 3"
  6.1028 +
  6.1029 +val NONE = eval_is_even "aaa" "bbb" (TermC.str2term "3::real") "ccc";
  6.1030 +
  6.1031 +if UnparseC.term t = "- 1 * a \<up> 3 + 3 * a \<up> 2 * x + - 3 * a * x \<up> 2 + x \<up> 3"
  6.1032  then () else error "poly.sml: diff.behav. in make_polynomial 10";
  6.1033  
  6.1034  "-----SPB Schalk I p.64 No.296c ---";
  6.1035  val t = TermC.str2term "(-3*x - 4*y) \<up> 3";
  6.1036  val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1037 -if (UnparseC.term t) = "-27 * x \<up> 3 + -108 * x \<up> 2 * y + -144 * x * y \<up> 2 +\n-64 * y \<up> 3"
  6.1038 +if UnparseC.term t = "- 27 * x \<up> 3 + - 108 * x \<up> 2 * y + - 144 * x * y \<up> 2 +\n- 64 * y \<up> 3"
  6.1039  then () else error "poly.sml: diff.behav. in make_polynomial 11";
  6.1040  
  6.1041  "-----SPB Schalk I p.62 No.242c ---";
  6.1042  val t = TermC.str2term "x \<up> (-4)*(x \<up> (-4)*y \<up> (-2)) \<up> (-1)*y \<up> (-2)";
  6.1043  val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1044 -if (UnparseC.term t) = "1"
  6.1045 +if UnparseC.term t = "1"
  6.1046  then () else error "poly.sml: diff.behav. in make_polynomial 12";
  6.1047  
  6.1048  "-----SPB Schalk I p.60 No.209a ---";
  6.1049 @@ -478,19 +1198,21 @@
  6.1050  if UnparseC.term t = "d \<up> 3"
  6.1051  then () else error "poly.sml: diff.behav. in make_polynomial 14";
  6.1052  
  6.1053 -(*---------------------------------------------------------------------*)
  6.1054 -(*---------------- ?RL?Bsple bei denen es Probleme gibt----------------*)
  6.1055 -(*---------------------------------------------------------------------*)
  6.1056 +"-------- ?RL?Bsple bei denen es Probleme gibt--------------------------------------------------";
  6.1057 +"-------- ?RL?Bsple bei denen es Probleme gibt--------------------------------------------------";
  6.1058 +"-------- ?RL?Bsple bei denen es Probleme gibt--------------------------------------------------";
  6.1059  "-----Schalk I p.64 No.303 ---";
  6.1060  val t = TermC.str2term "(a + 2*b)*(a \<up> 2 + 4*b \<up> 2)*(a - 2*b) - (a - 6*b)*(a \<up> 2 + 36*b \<up> 2)*(a + 6*b)";
  6.1061 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1062 +(*SOMETIMES LOOPS---------------------------------------------------------------------------\\*)
  6.1063 +val SOME (t, _) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1064  if UnparseC.term t = "1280 * b \<up> 4"
  6.1065  then () else error "poly.sml: diff.behav. in make_polynomial 14b";
  6.1066  (* Richtig - aber Binomische Formel wurde nicht verwendet! *)
  6.1067 +(*SOMETIMES LOOPS--------------------------------------------------------------------------//*)
  6.1068  
  6.1069 -(*--------------------------------------------------------------------*)
  6.1070 -(*----------------------- Eigene Beispiele ---------------------------*)
  6.1071 -(*--------------------------------------------------------------------*)
  6.1072 +"-------- Eigene Beispiele (Mathias Goldgruber) ------------------------------------------------";
  6.1073 +"-------- Eigene Beispiele (Mathias Goldgruber) ------------------------------------------------";
  6.1074 +"-------- Eigene Beispiele (Mathias Goldgruber) ------------------------------------------------";
  6.1075  "-----SPO ---";
  6.1076  val t = TermC.str2term "a \<up> 2*a \<up> (-2)";
  6.1077  val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1078 @@ -538,38 +1260,38 @@
  6.1079  else error "poly.sml: diff.behav. in make_polynomial 23";
  6.1080  "-----SPO ---";
  6.1081  val t = (Thm.term_of o the o (TermC.parse thy)) "a \<up> 2 * (-a) \<up> 2";
  6.1082 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1083 +val SOME (t,_) = rewrite_set_ @{theory} false make_polynomial t; UnparseC.term t;
  6.1084  if (UnparseC.term t) = "a \<up> 4" then ()
  6.1085  else error "poly.sml: diff.behav. in make_polynomial 24";
  6.1086  "-----SPO ---";
  6.1087  val t = TermC.str2term "a * b * b \<up> (-1) + a";
  6.1088 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1089 -if (UnparseC.term t) = "2 * a" then ()
  6.1090 +val SOME (t,_) = rewrite_set_ @{theory} false make_polynomial t; UnparseC.term t;
  6.1091 +if UnparseC.term t = "2 * a" then ()
  6.1092  else error "poly.sml: diff.behav. in make_polynomial 25";
  6.1093  "-----SPO ---";
  6.1094  val t = TermC.str2term "a*c*b \<up> (2*n) + 3*a + 5*b \<up> (2*n)*c*b";
  6.1095 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
  6.1096 -if (UnparseC.term t) = "3 * a + 5 * b \<up> (1 + 2 * n) * c + a * b \<up> (2 * n) * c"
  6.1097 +val SOME (t,_) = rewrite_set_ @{theory} false make_polynomial t; UnparseC.term t;
  6.1098 +if UnparseC.term t = "3 * a + 5 * b \<up> (1 + 2 * n) * c + a * b \<up> (2 * n) * c"
  6.1099  then () else error "poly.sml: diff.behav. in make_polynomial 26";
  6.1100  
  6.1101  (*MG030627 -------------vvv-: Verschachtelte Terme -----------*)
  6.1102  "-----SPO ---";
  6.1103  val t = TermC.str2term "(1 + (x*y*a) + x) \<up> (1 + (x*y*a) + x)";
  6.1104 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
  6.1105 +val SOME (t,_) = rewrite_set_ @{theory} false make_polynomial t;
  6.1106  if UnparseC.term t = "(1 + x + a * x * y) \<up> (1 + x + a * x * y)"
  6.1107  then () else error "poly.sml: diff.behav. in make_polynomial 27";(*SPO*)
  6.1108  
  6.1109  val t = TermC.str2term "(1 + x*(y*z)*zz) \<up> (1 + x*(y*z)*zz)";
  6.1110 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
  6.1111 +val SOME (t,_) = rewrite_set_ @{theory} false make_polynomial t;
  6.1112  if UnparseC.term t = "(1 + x * y * z * zz) \<up> (1 + x * y * z * zz)"
  6.1113  then () else error "poly.sml: diff.behav. in make_polynomial 28";
  6.1114  
  6.1115 -"-------- check pbl  'polynomial simplification' --------";
  6.1116 -"-------- check pbl  'polynomial simplification' --------";
  6.1117 -"-------- check pbl  'polynomial simplification' --------";
  6.1118 +"-------- check pbl  'polynomial simplification' -----------------------------------------------";
  6.1119 +"-------- check pbl  'polynomial simplification' -----------------------------------------------";
  6.1120 +"-------- check pbl  'polynomial simplification' -----------------------------------------------";
  6.1121  val fmz = ["Term ((5*x \<up> 2 + 3) * (2*x \<up> 7 + 3) - (3*x \<up> 5 + 8) * (6*x \<up> 4 - 1))", "normalform N"];
  6.1122  "-----0 ---";
  6.1123 -case Refine.refine fmz ["polynomial", "simplification"]of
  6.1124 +case Refine.refine fmz ["polynomial", "simplification"] of
  6.1125      [M_Match.Matches (["polynomial", "simplification"], _)] => ()
  6.1126    | _ => error "poly.sml diff.behav. in check pbl, Refine.refine";
  6.1127  (*...if there is an error, then ...*)
  6.1128 @@ -608,9 +1330,9 @@
  6.1129  (*show_types:=false;*)
  6.1130  
  6.1131  
  6.1132 -"-------- me 'poly. simpl.' Schalk I p.63 No.267b -------";
  6.1133 -"-------- me 'poly. simpl.' Schalk I p.63 No.267b -------";
  6.1134 -"-------- me 'poly. simpl.' Schalk I p.63 No.267b -------";
  6.1135 +"-------- me 'poly. simpl.' Schalk I p.63 No.267b ----------------------------------------------";
  6.1136 +"-------- me 'poly. simpl.' Schalk I p.63 No.267b ----------------------------------------------";
  6.1137 +"-------- me 'poly. simpl.' Schalk I p.63 No.267b ----------------------------------------------";
  6.1138  val fmz = ["Term ((5*x \<up> 2 + 3) * (2*x \<up> 7 + 3) - (3*x \<up> 5 + 8) * (6*x \<up> 4 - 1))", "normalform N"];
  6.1139  val (dI',pI',mI') =
  6.1140    ("Poly",["polynomial", "simplification"],
  6.1141 @@ -636,16 +1358,16 @@
  6.1142  
  6.1143  (*[1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt;(*Empty_Tac: ERROR DETECTED Feb.2020*)
  6.1144  
  6.1145 -(*+*)if f2str f = "17 + 15 * x \<up> 2 + -48 * x \<up> 4 + 3 * x \<up> 5 + 6 * x \<up> 7 +\n-8 * x \<up> 9"
  6.1146 +(*+*)if f2str f = "17 + 15 * x \<up> 2 + - 48 * x \<up> 4 + 3 * x \<up> 5 + 6 * x \<up> 7 +\n- 8 * x \<up> 9"
  6.1147  (*+*)then () else error "poly.sml diff.behav. in me Schalk I p.63 No.267b -1";
  6.1148  
  6.1149  (*[1], Res* )val (p,_,f,nxt,_,pt) = me nxt p c pt;( *SINCE Feb.2020 LItool.find_next_step without result*)
  6.1150  
  6.1151  
  6.1152  
  6.1153 -"-------- interSteps for Schalk 299a --------------------";
  6.1154 -"-------- interSteps for Schalk 299a --------------------";
  6.1155 -"-------- interSteps for Schalk 299a --------------------";
  6.1156 +"-------- interSteps for Schalk 299a -----------------------------------------------------------";
  6.1157 +"-------- interSteps for Schalk 299a -----------------------------------------------------------";
  6.1158 +"-------- interSteps for Schalk 299a -----------------------------------------------------------";
  6.1159  reset_states ();
  6.1160  CalcTree
  6.1161  [(["Term ((x - y)*(x + (y::real)))", "normalform N"], 
  6.1162 @@ -668,31 +1390,50 @@
  6.1163  else error "poly.sml: interSteps doesnt work again 2";
  6.1164  ============ inhibit exn WN120316 ==============================================*)
  6.1165  
  6.1166 -"-------- norm_Poly NOT COMPLETE ------------------------";
  6.1167 -"-------- norm_Poly NOT COMPLETE ------------------------";
  6.1168 -"-------- norm_Poly NOT COMPLETE ------------------------";
  6.1169 +"-------- norm_Poly NOT COMPLETE ---------------------------------------------------------------";
  6.1170 +"-------- norm_Poly NOT COMPLETE ---------------------------------------------------------------";
  6.1171 +"-------- norm_Poly NOT COMPLETE ---------------------------------------------------------------";
  6.1172  val thy = @{theory AlgEin};
  6.1173  
  6.1174  val SOME (f',_) = rewrite_set_ thy false norm_Poly
  6.1175  (TermC.str2term "L = k - 2 * q + (k - 2 * q) + (k - 2 * q) + (k - 2 * q) + senkrecht + oben");
  6.1176 -if UnparseC.term f' = "L = senkrecht + oben + 2 * 2 * k + 2 * -4 * q"
  6.1177 +if UnparseC.term f' = "L = 2 * 2 * k + 2 * - 4 * q + senkrecht + oben"
  6.1178  then ((*norm_Poly NOT COMPLETE -- TODO MG*))
  6.1179  else error "norm_Poly changed behavior";
  6.1180 +(* isa:
  6.1181 +##  rls: order_add_rls_ on: L = k + - 2 * q + k + - 2 * q + k + - 2 * q + k + - 2 * q + senkrecht + oben 
  6.1182 +###  rls: order_add_ on: L = k + - 2 * q + k + - 2 * q + k + - 2 * q + k + - 2 * q + senkrecht + oben 
  6.1183 +######  rls: Rule_Set.empty-is_addUnordered on: k + - 2 * q + k + - 2 * q + k + - 2 * q + k + - 2 * q + senkrecht +
  6.1184 +oben is_addUnordered 
  6.1185 +#######  try calc: "Poly.is_addUnordered" 
  6.1186 +########  eval asms: "k + - 2 * q + k + - 2 * q + k + - 2 * q + k + - 2 * q + senkrecht +
  6.1187 +oben is_addUnordered = True" 
  6.1188 +#######  calc. to: True 
  6.1189 +#######  try calc: "Poly.is_addUnordered" 
  6.1190 +#######  try calc: "Poly.is_addUnordered" 
  6.1191 +###  rls: order_add_ on: L = k + k + k + k + - 2 * q + - 2 * q + - 2 * q + - 2 * q + senkrecht + oben 
  6.1192 +*)
  6.1193 +"~~~~~ fun sort_monoms , args:"; val (t) =
  6.1194 +  (TermC.str2term "L = k + k + k + k + - 2 * q + - 2 * q + - 2 * q + - 2 * q + senkrecht + oben");
  6.1195 +(*+*)val t' =
  6.1196 +           sort_monoms t;
  6.1197 +(*+*)UnparseC.term t' = "L = k + k + k + k + - 2 * q + - 2 * q + - 2 * q + - 2 * q + senkrecht + oben"; (*isa*)
  6.1198  
  6.1199 -"-------- ord_make_polynomial ---------------------------";
  6.1200 -"-------- ord_make_polynomial ---------------------------";
  6.1201 -"-------- ord_make_polynomial ---------------------------";
  6.1202 +"-------- ord_make_polynomial ------------------------------------------------------------------";
  6.1203 +"-------- ord_make_polynomial ------------------------------------------------------------------";
  6.1204 +"-------- ord_make_polynomial ------------------------------------------------------------------";
  6.1205  val t1 = TermC.str2term "2 * b + (3 * a + 3 * b)";
  6.1206 -val t2 = TermC.str2term "3 * a + 3 * b + 2 * b";
  6.1207 +val t2 = TermC.str2term "(3 * a + 3 * b) + 2 * b";
  6.1208  
  6.1209  if ord_make_polynomial true thy [] (t1, t2) then ()
  6.1210  else error "poly.sml: diff.behav. in ord_make_polynomial";
  6.1211 +(*SO: WHY IS THERE NO REWRITING ...*)
  6.1212  
  6.1213 -(*WN071202: \<up> why then is there no rewriting ...*)
  6.1214  val term = TermC.str2term "2*b + (3*a + 3*b)";
  6.1215 -val NONE = rewrite_set_ (@{theory "Isac_Knowledge"}) false order_add_mult term;
  6.1216 +(*+++*)val SOME (t', []) = rewrite_set_ @{theory "Isac_Knowledge"} false order_add_mult term;
  6.1217 +(*+++*)if UnparseC.term t' = "a * 3 + (b * 2 + b * 3)" then () else error "no rewriting ?!?";
  6.1218 +(* before dropping ThmC.numerals_to_Free this was
  6.1219 +val NONE = rewrite_set_ @ {theory "Isac_Knowledge"} false order_add_mult term;
  6.1220 +SO: WHY IS THERE NO REWRITING ?!?
  6.1221 +*)
  6.1222  
  6.1223 -(*or why is there no rewriting this way...*)
  6.1224 -val t1 = TermC.str2term "2 * b + (3 * a + 3 * b)";
  6.1225 -val t2 = TermC.str2term "3 * a + (2 * b + 3 * b)";
  6.1226 -
     7.1 --- a/test/Tools/isac/Knowledge/rational.sml	Tue Jun 01 15:41:23 2021 +0200
     7.2 +++ b/test/Tools/isac/Knowledge/rational.sml	Sat Jul 03 16:21:07 2021 +0200
     7.3 @@ -54,12 +54,34 @@
     7.4  
     7.5  if poly_of_term vs (TermC.str2term "12::real") = SOME [(12, [0, 0, 0])]
     7.6  then () else error "poly_of_term 1 changed";
     7.7 +
     7.8  if poly_of_term vs (TermC.str2term "x::real") = SOME [(1, [1, 0, 0])]
     7.9  then () else error "poly_of_term 2 changed";
    7.10 -if poly_of_term vs (TermC.str2term "12 * x \<up> 3") = SOME [(12, [3, 0, 0])]
    7.11 +
    7.12 +(*//---------------------------- removed error -----------------------------------------------\\*)
    7.13 +if         poly_of_term vs (TermC.str2term "12 * x \<up> 3") = SOME [(12, [3, 0, 0])]
    7.14  then () else error "poly_of_term 3 changed";
    7.15 +"~~~~~ fun poly_of_term , args:"; val (vs, t) =
    7.16 +  (vs, (TermC.str2term "12 * x \<up> 3"));
    7.17 +
    7.18 +           monom_of_term vs (1, replicate (length vs) 0) t;(*poly malformed 1 with x \<up> 3*)
    7.19 +"~~~~~ fun monom_of_term , args:"; val (vs, (c, es), (Const ("Groups.times_class.times", _) $ m1 $ m2)) =
    7.20 +  (vs, (1, replicate (length vs) 0), t);
    7.21 +    val (c', es') =
    7.22 +
    7.23 +           monom_of_term vs (c, es) m1;
    7.24 +"~~~~~ fun monom_of_term , args:"; val (vs, (c, es), (Const ("Transcendental.powr", _) $ (t as Free _) $ (Const ("Num.numeral_class.numeral", _) $ num)) ) =
    7.25 +  (vs, (c', es'), m2);
    7.26 +(*+*)c = 12;
    7.27 +(*+*)(num |> HOLogic.dest_numeral |> list_update es (find_index (curry op = t) vs)) = [3, 0, 0];
    7.28 +
    7.29 +if (c, num |> HOLogic.dest_numeral |> list_update es (find_index (curry op = t) vs)) = (12, [3, 0, 0])
    7.30 +then () else error "monom_of_term (powr): return value CHANGED";
    7.31 +(*\\---------------------------- removed error -----------------------------------------------//*)
    7.32 +
    7.33  if poly_of_term vs (TermC.str2term "12 * x \<up> 3 * y \<up> 4 * z \<up> 6") = SOME [(12, [3, 4, 6])]
    7.34  then () else error "poly_of_term 4 changed";
    7.35 +
    7.36  if poly_of_term vs (TermC.str2term "1 + 2 * x \<up> 3 * y \<up> 4 * z \<up> 6 + y") =
    7.37    SOME [(1, [0, 0, 0]), (1, [0, 1, 0]), (2, [3, 4, 6])]
    7.38  then () else error "poly_of_term 5 changed";
    7.39 @@ -79,6 +101,7 @@
    7.40    = SOME [(1, [0, 0, 0]), (7, [0, 8, 0]), (2, [3, 4, 6])]
    7.41  then () else error "poly_of_term 8 changed";
    7.42  
    7.43 +
    7.44  "-------- fun is_poly --------------------------------------------------------";
    7.45  "-------- fun is_poly --------------------------------------------------------";
    7.46  "-------- fun is_poly --------------------------------------------------------";
    7.47 @@ -1692,7 +1715,7 @@
    7.48  val SOME (tt',asm) = rewrite_set_ thy false make_polynomial tt;
    7.49  if UnparseC.term tt' = "a \<up> 2 + -2 * a * b + b \<up> 2" then () else error "rls chancel_p 2";
    7.50  
    7.51 -"----- with .make_deriv; WN1130912 not investigated further, will be discontinued";
    7.52 +"----- with Derive.do_one; WN1130912 not investigated further, will be discontinued";
    7.53  val SOME (tt, _) = factout_p_ thy t; 
    7.54  if UnparseC.term tt = "(a + b) * (a + -1 * b) / ((a + -1 * b) * (a + -1 * b))"
    7.55  then () else error "rls chancel_p 3";
    7.56 @@ -1700,7 +1723,7 @@
    7.57  
    7.58  "--- with simpler ruleset";
    7.59  val {rules, rew_ord= (_, ro), ...} = Rule_Set.rep (assoc_rls "rev_rew_p");
    7.60 -val der = .make_deriv thy Atools_erls rules ro NONE tt;
    7.61 +val der = Derive.do_one thy Atools_erls rules ro NONE tt;
    7.62  if length der = 12 then () else error "WN1130912 rls chancel_p 4";
    7.63  (*default_print_depth 99;*) writeln (deriv2str der); (*default_print_depth 3;*)
    7.64  
    7.65 @@ -1709,12 +1732,12 @@
    7.66  (*default_print_depth 99;*) map (Rule.to_string o #2) der; (*default_print_depth 3;*)
    7.67  (*default_print_depth 99;*) map (UnparseC.term o #1 o #3) der; (*default_print_depth 3;*)
    7.68  
    7.69 -val der = .make_deriv thy Atools_erls rules ro NONE 
    7.70 +val der = Derive.do_one thy Atools_erls rules ro NONE 
    7.71  	(TermC.str2term "(1 * a + 1 * b) * (1 * a + -1 * b)");
    7.72  (*default_print_depth 99;*) writeln (deriv2str der); (*default_print_depth 3;*)
    7.73  
    7.74  val {rules, rew_ord=(_,ro),...} = Rule_Set.rep (assoc_rls "rev_rew_p");
    7.75 -val der = .make_deriv thy Atools_erls rules ro NONE 
    7.76 +val der = Derive.do_one thy Atools_erls rules ro NONE 
    7.77  	(TermC.str2term "(1 * a + -1 * b) * (1 * a + -1 * b)");
    7.78  (*default_print_depth 99;*) writeln (deriv2str der); (*default_print_depth 3;*)
    7.79  (*default_print_depth 99;*) map (UnparseC.term o #1) der; (*default_print_depth 3;*)
     8.1 --- a/test/Tools/isac/MathEngBasic/rewrite.sml	Tue Jun 01 15:41:23 2021 +0200
     8.2 +++ b/test/Tools/isac/MathEngBasic/rewrite.sml	Sat Jul 03 16:21:07 2021 +0200
     8.3 @@ -1,4 +1,4 @@
     8.4 -(* Title: "ProgLang/rewrite.sml"
     8.5 +(* Title: "MathEngBasic/rewrite.sml"
     8.6     Author: Walther Neuper 050908
     8.7     (c) copyright due to lincense terms.
     8.8  *)
     8.9 @@ -328,6 +328,7 @@
    8.10  val asms = map (Envir.subst_term subst) pres;
    8.11  if UnparseC.terms asms = "[\"a + b is_expanded\", \"c is_expanded\"]"
    8.12  then () else error "rewrite.sml: prepat cancel subst";
    8.13 +
    8.14  if ([], true) = eval__true thy 0 asms [] erls
    8.15  then () else error "rewrite.sml: prepat cancel eval__true";
    8.16  
    8.17 @@ -360,6 +361,7 @@
    8.18  val asms = map (Envir.subst_term subst) pres;
    8.19  if UnparseC.terms asms = "[\"x \<up> 2 * x is_multUnordered\"]"
    8.20  then () else error "rewrite.sml: prepat order_mult_ subst";
    8.21 +
    8.22  if ([], true) = eval__true thy 0 asms [] erls
    8.23  then () else error "rewrite.sml: prepat order_mult_ eval__true";
    8.24  
    8.25 @@ -371,7 +373,21 @@
    8.26  
    8.27  if is_multUnordered t then () else error "rewrite.sml diff. is_multUnordered 2";
    8.28  val tm = TermC.str2term "(x \<up> 2 * x) is_multUnordered";
    8.29 -eval_is_multUnordered "testid" "" tm thy;
    8.30 +
    8.31 +(*+*)case eval_is_multUnordered "testid" "" tm thy of
    8.32 +(*+*)  SOME
    8.33 +(*+*)    ("testidx \<up> 2 * x_",
    8.34 +(*+*)     Const ("HOL.Trueprop", _) $
    8.35 +(*+*)       (Const ("HOL.eq", _) $
    8.36 +(*+*)         (Const ("Poly.is_multUnordered", _) $
    8.37 +(*+*)           (Const ("Groups.times_class.times", _) $
    8.38 +(*+*)             (Const ("Transcendental.powr", _) $ Free ("x", _) $ _ ) $ Free ("x", _))) $
    8.39 +(*+*)         Const ("HOL.True", _))) => ()
    8.40 +(*+*)(*                   ^^^^^^             compare ---vvv *)
    8.41 +(*+*)| _ => error "rewrite.sml diff. eval_is_multUnordered 2b CHANGED";
    8.42 +
    8.43 +
    8.44 +     eval_is_multUnordered "testid" "" tm thy;
    8.45  
    8.46  case eval_is_multUnordered "testid" "" tm thy of
    8.47      SOME (_, Const ("HOL.Trueprop", _) $ 
    8.48 @@ -613,7 +629,7 @@
    8.49  
    8.50             rewrite__set_ thy (i + 1) false bdv rls a (*of*);
    8.51  
    8.52 -(*+*)Rewrite.trace_on := true;
    8.53 +(*+*)Rewrite.trace_on := false;
    8.54  
    8.55          (*this was False; vvvv--- means: indeterminate*)
    8.56      val (* SOME (t, a') *)NONE = (*case*)
    8.57 @@ -671,7 +687,7 @@
    8.58  exception TERM raised (line 271 of "~~/src/HOL/Tools/hologic.ML"):
    8.59    dest_eq
    8.60    0 \<le> ?a \<Longrightarrow> (0 \<le> sqrt ?a) = True( **)
    8.61 -if UnparseC.term t = "x + 1 + -1 * 2 = 0" then () else error "rewrite_set_ norm_equation CHANGED";
    8.62 +if UnparseC.term t = "x + 1 + - 1 * 2 = 0" then () else error "rewrite_set_ norm_equation CHANGED";
    8.63  
    8.64  "~~~~~ fun rewrite_set_ , args:"; val (thy, bool, rls, term) = (thy, false, rls, term);
    8.65  "~~~~~ and rewrite__set_ , args:"; val (thy, i, put_asm, bdv, rls, ct) =
     9.1 --- a/test/Tools/isac/ProgLang/evaluate.sml	Tue Jun 01 15:41:23 2021 +0200
     9.2 +++ b/test/Tools/isac/ProgLang/evaluate.sml	Sat Jul 03 16:21:07 2021 +0200
     9.3 @@ -167,8 +167,7 @@
     9.4   val t = (Thm.term_of o the o (TermC.parse thy)) "(3+1+2*x)/2";
     9.5  case rewrite_set_ thy false rls t of
     9.6    SOME (t', _) => 
     9.7 -    if UnparseC.term t' = "x + (1 / 2 + 3 / 2)" (*ERROR should be : 2 * x  SEE ABOVE*) then ()
     9.8 -    else error "rewrite_set_ (3+1+2*x)/2 changed 1"
     9.9 +    if UnparseC.term t' = "x + 2" then () else error "rewrite_set_ (3+1+2*x)/2 changed 1"
    9.10  | _ => error "rewrite_set_ (3+1+2*x)/2 changed 2";
    9.11  
    9.12   Rewrite.trace_on:=false; (*=true3.6.03*)
    10.1 --- a/test/Tools/isac/ProgLang/prog_expr.sml	Tue Jun 01 15:41:23 2021 +0200
    10.2 +++ b/test/Tools/isac/ProgLang/prog_expr.sml	Sat Jul 03 16:21:07 2021 +0200
    10.3 @@ -1,4 +1,4 @@
    10.4 -(* Title:  ProgLang/program.sml
    10.5 +(* Title:  ProgLang/prog_expr.sml
    10.6     Author: Walther Neuper 190922
    10.7     (c) due to copyright terms
    10.8  *)
    10.9 @@ -7,6 +7,8 @@
   10.10  "-----------------------------------------------------------------------------------------------";
   10.11  "table of contents -----------------------------------------------------------------------------";
   10.12  "-----------------------------------------------------------------------------------------------";
   10.13 +"-------- fun eval_is_even ---------------------------------------------------------------------";
   10.14 +"-------- fun eval_const -----------------------------------------------------------------------";
   10.15  "-------- fun eval_cancel ----------------------------------------------------------------------";
   10.16  "-------- fun eval_equ -------------------------------------------------------------------------";
   10.17  "-------- fun eval_equal for x \<noteq> 0: \<rightarrow> indetermined, NOT \<rightarrow> 'True' -----------------------------";
   10.18 @@ -24,6 +26,71 @@
   10.19  "-----------------------------------------------------------------------------------------------";
   10.20  
   10.21  
   10.22 +"-------- fun eval_is_even ---------------------------------------------------------------------";
   10.23 +"-------- fun eval_is_even ---------------------------------------------------------------------";
   10.24 +"-------- fun eval_is_even ---------------------------------------------------------------------";
   10.25 +val t = TermC.str2term "2 is_even";
   10.26 +           eval_is_even "aaa" "Prog_Expr.is_even" t "ccc";
   10.27 +"~~~~~ fun eval_is_even , args:"; val ((thmid:string), "Prog_Expr.is_even", (t as (Const _ $ arg)), _) =
   10.28 +  ("aaa", "Prog_Expr.is_even", t, "ccc");
   10.29 +    (*if*) TermC.is_num arg (*then*);
   10.30 +        val i = arg |> HOLogic.dest_number |> snd;
   10.31 +	      (*if*) even i (*then*);
   10.32 +val SOME ("aaa1_", t') =
   10.33 +    SOME (TermC.mk_thmid thmid (string_of_int n) "", 
   10.34 +				  HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})));
   10.35 +if UnparseC.term t' = "(2 is_even) = True" then () else error "(2 is_even) = True  CHANGED";
   10.36 +
   10.37 +
   10.38 +val t = TermC.str2term "3 is_even";
   10.39 +case eval_is_even "aaa" "Prog_Expr.is_even" t "ccc" of
   10.40 +  SOME (str, t') => 
   10.41 +    if str = "aaa_" andalso UnparseC.term t' = "(3 is_even) = False" then ()
   10.42 +    else error "eval_is_even (3 is_even) CHANGED 1"
   10.43 +| NONE => error "eval_is_even (3 is_even) CHANGED 2";
   10.44 +
   10.45 +val t = TermC.str2term "a ::real";
   10.46 +val NONE =
   10.47 +           eval_is_even "aaa" "Prog_Expr.is_even" t "ccc";
   10.48 +case eval_is_even "aaa" "Prog_Expr.is_even" t "ccc" of
   10.49 +  SOME _ => error "eval_is_even (a is_even) CHANGED"
   10.50 +| NONE => ();
   10.51 +
   10.52 +
   10.53 +"-------- fun eval_const -----------------------------------------------------------------------";
   10.54 +"-------- fun eval_const -----------------------------------------------------------------------";
   10.55 +"-------- fun eval_const -----------------------------------------------------------------------";
   10.56 +val t = (Thm.term_of o the o (TermC.parse @{theory Test})) "2 is_const";
   10.57 +case rewrite_set_ @{theory Test} false tval_rls t of
   10.58 +  SOME (Const ("HOL.True", _), []) => ()
   10.59 +| _ => error "2 is_const CHANGED";
   10.60 +
   10.61 +val t = TermC.str2term "2 * x is_const";
   10.62 +val SOME (str, t') = eval_const "" "" t (@{theory "Isac_Knowledge"});
   10.63 +if UnparseC.term t' = "(2 * x is_const) = False" then ()
   10.64 +else error "(2 * x is_const) = False CHANGED";
   10.65 +
   10.66 +val t = TermC.str2term "- 2 is_const";
   10.67 +val SOME (str, t') = eval_const "" "" t (@{theory "Isac_Knowledge"});
   10.68 +if UnparseC.term t' = "(- 2 is_const) = True" then ()
   10.69 +else error "(- 2 is_const) = False CHANGED";
   10.70 +
   10.71 +val t = TermC.str2term "- 1 is_const";
   10.72 +val SOME (str, t') = eval_const "" "" t (@{theory "Isac_Knowledge"});
   10.73 +if UnparseC.term t' = "(- 1 is_const) = True" then ()
   10.74 +else error "(- 1 is_const) = False CHANGED";
   10.75 +
   10.76 +val t = TermC.str2term "0 is_const";
   10.77 +val SOME (str, t') = eval_const "" "" t (@{theory "Isac_Knowledge"});
   10.78 +if UnparseC.term t' = "(0 is_const) = True" then ()
   10.79 +else error "(0 is_const) = False CHANGED";
   10.80 +
   10.81 +val t = TermC.str2term "AA is_const";
   10.82 +val SOME (str, t') = eval_const "" "" t (@{theory "Isac_Knowledge"});
   10.83 +if UnparseC.term t' = "(AA is_const) = False" then ()
   10.84 +else error "(0 is_const) = False CHANGED";
   10.85 +
   10.86 +
   10.87  "-------- fun eval_cancel ----------------------------------------------------------------------";
   10.88  "-------- fun eval_cancel ----------------------------------------------------------------------";
   10.89  "-------- fun eval_cancel ----------------------------------------------------------------------";
    11.1 --- a/test/Tools/isac/Test_Isac_Short.thy	Tue Jun 01 15:41:23 2021 +0200
    11.2 +++ b/test/Tools/isac/Test_Isac_Short.thy	Sat Jul 03 16:21:07 2021 +0200
    11.3 @@ -230,15 +230,6 @@
    11.4  subsection \<open>further functionality alongside batch build sequence\<close>
    11.5    ML_file "MathEngBasic/thmC.sml"
    11.6    ML_file "MathEngBasic/rewrite.sml"
    11.7 -ML \<open>
    11.8 -\<close> ML \<open>
    11.9 -\<close> ML \<open>
   11.10 -\<close> ML \<open>
   11.11 -\<close> ML \<open>
   11.12 -\<close> ML \<open>
   11.13 -\<close> ML \<open>
   11.14 -\<close> ML \<open>
   11.15 -\<close>
   11.16    ML_file "MathEngBasic/tactic.sml"
   11.17    ML_file "MathEngBasic/ctree.sml"
   11.18    ML_file "MathEngBasic/calculation.sml"
   11.19 @@ -293,13 +284,14 @@
   11.20    ML_file "Knowledge/poly.sml"
   11.21    ML_file "Knowledge/gcd_poly_ml.sml"
   11.22    ML_file "Knowledge/gcd_poly_winkler.sml" (*must be after gcd_poly_ml.sml: redefines functions*)
   11.23 -(*ML_file "Knowledge/rational.sml"                                              Test_Isac_Short*)
   11.24 +  ML_file "Knowledge/rational.sml"                                            (*Test_Isac_Short*)
   11.25    ML_file "Knowledge/equation.sml"
   11.26    ML_file "Knowledge/root.sml"
   11.27    ML_file "Knowledge/lineq.sml"
   11.28 +
   11.29  (*ML_file "Knowledge/rooteq.sml"    some complicated equations not recovered from 2002 *)
   11.30  (*ML_file "Knowledge/rateq.sml"     some complicated equations not recovered----Test_Isac_Short*)
   11.31 -  ML_file "Knowledge/rootrat.sml"
   11.32 +  ML_file "Knowledge/r/ootrat.sml"
   11.33    ML_file "Knowledge/rootrateq.sml"(*ome complicated equations not recovered from 2002 *)
   11.34  (*ML_file "Knowledge/partial_fractions.sml"  hangs with ML_system_64 = "true"---Test_Isac_Short*)
   11.35    ML_file "Knowledge/polyeq-1.sml"
    12.1 --- a/test/Tools/isac/Test_Some.thy	Tue Jun 01 15:41:23 2021 +0200
    12.2 +++ b/test/Tools/isac/Test_Some.thy	Sat Jul 03 16:21:07 2021 +0200
    12.3 @@ -93,6 +93,9 @@
    12.4    declare [[ML_exception_trace]]
    12.5  \<close>
    12.6  
    12.7 +ML_file \<open>MathEngBasic/rewrite.sml\<close>
    12.8 +ML_file \<open>Knowledge/poly.sml\<close>
    12.9 +
   12.10  section \<open>===================================================================================\<close>
   12.11  ML \<open>
   12.12  \<close> ML \<open>
   12.13 @@ -100,7 +103,969 @@
   12.14  \<close> ML \<open>
   12.15  \<close>
   12.16  
   12.17 -ML_file "ProgLang/listC.sml"
   12.18 +section \<open>2============== NEW src/../"Knowledge/Rational.thy" PART --> Poly.thy ==============\<close>
   12.19 +section \<open>Cancellation and addition of fractions\<close>
   12.20 +subsection \<open>Conversion term <--> poly\<close>
   12.21 +subsubsection \<open>Convert a term to the internal representation of a multivariate polynomial\<close>
   12.22 +ML \<open>
   12.23 +fun monom_of_term vs (c, es) (t as Const _) =
   12.24 +    (c, list_update es (find_index (curry op = t) vs) 1)
   12.25 +  | monom_of_term _ (c, es) (t as (Const ("Num.numeral_class.numeral", _) $ _)) =
   12.26 +    (t |> HOLogic.dest_number |> snd |> string_of_int
   12.27 +      |> TermC.int_opt_of_string |> the |> curry op * c, es) (*several numerals in one monom*)
   12.28 +  | monom_of_term  vs (c, es) (t as Free _) =
   12.29 +    (c, list_update es (find_index (curry op = t) vs) 1)
   12.30 +  | monom_of_term  vs (c, es) (Const ("Transcendental.powr", _) $ (t as Free _) $
   12.31 +      (Const ("Num.numeral_class.numeral", _) $ num)) =
   12.32 +    (c, list_update es (find_index (curry op = t) vs) (num |> HOLogic.dest_numeral))
   12.33 +  | monom_of_term vs (c, es) (Const ("Groups.times_class.times", _) $ m1 $ m2) =
   12.34 +    let val (c', es') = monom_of_term vs (c, es) m1
   12.35 +    in monom_of_term vs (c', es') m2 end
   12.36 +  | monom_of_term _ _ t = raise ERROR ("poly malformed 1 with " ^ UnparseC.term t)
   12.37 +
   12.38 +(*-------v------*)
   12.39 +fun monoms_of_term vs (t as Const _) =
   12.40 +    [monom_of_term  vs (1, replicate (length vs) 0) t]
   12.41 +  | monoms_of_term vs (t as Free _) =
   12.42 +    [monom_of_term  vs (1, replicate (length vs) 0) t]
   12.43 +  | monoms_of_term vs (t as Const ("Transcendental.powr", _) $ _ $  _) =
   12.44 +    [monom_of_term  vs (1, replicate (length vs) 0) t]
   12.45 +  | monoms_of_term vs (t as Const ("Groups.times_class.times", _) $ _ $  _) =
   12.46 +    [monom_of_term  vs (1, replicate (length vs) 0) t]
   12.47 +  | monoms_of_term vs (Const ("Groups.plus_class.plus", _) $ ms1 $ ms2) =
   12.48 +    (monoms_of_term vs ms1) @ (monoms_of_term vs ms2)
   12.49 +  | monoms_of_term _ t = raise ERROR ("poly malformed 2 with " ^ UnparseC.term t)
   12.50 +
   12.51 +(* convert a term to the internal representation of a multivariate polynomial;
   12.52 +  the conversion is quite liberal, see test --- fun poly_of_term ---:
   12.53 +* the order of variables and the parentheses within a monomial are arbitrary
   12.54 +* the coefficient may be somewhere
   12.55 +* he order and the parentheses within monomials are arbitrary
   12.56 +But the term must be completely expand + over * (laws of distributivity are not applicable).
   12.57 +
   12.58 +The function requires the free variables as strings already given, 
   12.59 +because the gcd involves 2 polynomials (with the same length for their list of exponents).
   12.60 +*)
   12.61 +fun poly_of_term vs (t as Const ("Groups.plus_class.plus", _) $ _ $ _) =
   12.62 +    (SOME (t |> monoms_of_term vs |> order)
   12.63 +      handle ERROR _ => NONE)
   12.64 +  | poly_of_term vs t =
   12.65 +    (SOME [monom_of_term vs (1, replicate (length vs) 0) t]
   12.66 +      handle ERROR _ => NONE)
   12.67 +
   12.68 +fun is_poly t =
   12.69 +  let
   12.70 +    val vs = TermC.vars_of t
   12.71 +  in 
   12.72 +    case poly_of_term vs t of SOME _ => true | NONE => false
   12.73 +  end
   12.74 +val is_expanded = is_poly   (* TODO: check names *)
   12.75 +val is_polynomial = is_poly (* TODO: check names *)
   12.76 +\<close>
   12.77 +
   12.78 +subsubsection \<open>Convert internal representation of a multivariate polynomial to a term\<close>
   12.79 +ML \<open>
   12.80 +fun term_of_es _ _ _ [] = [] (*assumes same length for vs and es*)
   12.81 +  | term_of_es baseT expT (_ :: vs) (0 :: es) = [] @ term_of_es baseT expT vs es
   12.82 +  | term_of_es baseT expT (v :: vs) (1 :: es) = v :: term_of_es baseT expT vs es
   12.83 +  | term_of_es baseT expT (v :: vs) (e :: es) =
   12.84 +    Const ("Transcendental.powr", [baseT, expT] ---> baseT) $ v $ (HOLogic.mk_number expT e)
   12.85 +    :: term_of_es baseT expT vs es
   12.86 +  | term_of_es _ _ _ _ = raise ERROR "term_of_es: length vs <> length es"
   12.87 +
   12.88 +fun term_of_monom baseT expT vs ((c, es): monom) =
   12.89 +    let val es' = term_of_es baseT expT vs es
   12.90 +    in 
   12.91 +      if c = 1 
   12.92 +      then 
   12.93 +        if es' = [] (*if es = [0,0,0,...]*)
   12.94 +        then HOLogic.mk_number baseT c
   12.95 +        else foldl (HOLogic.mk_binop "Groups.times_class.times") (hd es', tl es')
   12.96 +      else foldl (HOLogic.mk_binop "Groups.times_class.times")
   12.97 +        (HOLogic.mk_number baseT c, es') 
   12.98 +    end
   12.99 +
  12.100 +fun term_of_poly baseT expT vs p =
  12.101 +  let val monos = map (term_of_monom baseT expT vs) p
  12.102 +  in foldl (HOLogic.mk_binop "Groups.plus_class.plus") (hd monos, tl monos) end
  12.103 +\<close>
  12.104 +
  12.105 +subsection \<open>Apply gcd_poly for cancelling and adding fractions as terms\<close>
  12.106 +ML \<open>
  12.107 +fun mk_noteq_0 baseT t = 
  12.108 +  Const ("HOL.Not", HOLogic.boolT --> HOLogic.boolT) $ 
  12.109 +    (Const ("HOL.eq", [baseT, baseT] ---> HOLogic.boolT) $ t $ HOLogic.mk_number HOLogic.realT 0)
  12.110 +
  12.111 +fun mk_asms baseT ts =
  12.112 +  let val as' = filter_out TermC.is_num ts (* asm like "2 ~= 0" is needless *)
  12.113 +  in map (mk_noteq_0 baseT) as' end
  12.114 +\<close>
  12.115 +
  12.116 +subsubsection \<open>Factor out gcd for cancellation\<close>
  12.117 +ML \<open>
  12.118 +fun check_fraction t =
  12.119 +  case t of
  12.120 +    Const ("Rings.divide_class.divide", _) $ numerator $ denominator
  12.121 +      => SOME (numerator, denominator)
  12.122 +  | _ => NONE
  12.123 +
  12.124 +(* prepare a term for cancellation by factoring out the gcd
  12.125 +  assumes: is a fraction with outmost "/"*)
  12.126 +fun factout_p_ (thy: theory) t =
  12.127 +  let val opt = check_fraction t
  12.128 +  in
  12.129 +    case opt of 
  12.130 +      NONE => NONE
  12.131 +    | SOME (numerator, denominator) =>
  12.132 +      let
  12.133 +        val vs = TermC.vars_of t
  12.134 +        val baseT = type_of numerator
  12.135 +        val expT = HOLogic.realT
  12.136 +      in
  12.137 +        case (poly_of_term vs numerator, poly_of_term vs denominator) of
  12.138 +          (SOME a, SOME b) =>
  12.139 +            let
  12.140 +              val ((a', b'), c) = gcd_poly a b
  12.141 +              val es = replicate (length vs) 0 
  12.142 +            in
  12.143 +              if c = [(1, es)] orelse c = [(~1, es)]
  12.144 +              then NONE
  12.145 +              else 
  12.146 +                let
  12.147 +                  val b't = term_of_poly baseT expT vs b'
  12.148 +                  val ct = term_of_poly baseT expT vs c
  12.149 +                  val t' = 
  12.150 +                    HOLogic.mk_binop "Rings.divide_class.divide" 
  12.151 +                      (HOLogic.mk_binop "Groups.times_class.times"
  12.152 +                        (term_of_poly baseT expT vs a', ct),
  12.153 +                       HOLogic.mk_binop "Groups.times_class.times" (b't, ct))
  12.154 +                in SOME (t', mk_asms baseT [b't, ct]) end
  12.155 +            end
  12.156 +        | _ => NONE : (term * term list) option
  12.157 +      end
  12.158 +  end
  12.159 +\<close>
  12.160 +
  12.161 +subsubsection \<open>Cancel a fraction\<close>
  12.162 +ML \<open>
  12.163 +(* cancel a term by the gcd ("" denote terms with internal algebraic structure)
  12.164 +  cancel_p_ :: theory \<Rightarrow> term  \<Rightarrow> (term \<times> term list) option
  12.165 +  cancel_p_ thy "a / b" = SOME ("a' / b'", ["b' \<noteq> 0"])
  12.166 +  assumes: a is_polynomial  \<and>  b is_polynomial  \<and>  b \<noteq> 0
  12.167 +  yields
  12.168 +    SOME ("a' / b'", ["b' \<noteq> 0"]). gcd_poly a b \<noteq> 1  \<and>  gcd_poly a b \<noteq> -1  \<and>  
  12.169 +      a' * gcd_poly a b = a  \<and>  b' * gcd_poly a b = b
  12.170 +    \<or> NONE *)
  12.171 +fun cancel_p_ (_: theory) t =
  12.172 +  let val opt = check_fraction t
  12.173 +  in
  12.174 +    case opt of 
  12.175 +      NONE => NONE
  12.176 +    | SOME (numerator, denominator) =>
  12.177 +      let
  12.178 +        val vs = TermC.vars_of t
  12.179 +        val baseT = type_of numerator
  12.180 +        val expT = HOLogic.realT
  12.181 +      in
  12.182 +        case (poly_of_term vs numerator, poly_of_term vs denominator) of
  12.183 +          (SOME a, SOME b) =>
  12.184 +            let
  12.185 +              val ((a', b'), c) = gcd_poly a b
  12.186 +              val es = replicate (length vs) 0 
  12.187 +            in
  12.188 +              if c = [(1, es)] orelse c = [(~1, es)]
  12.189 +              then NONE
  12.190 +              else 
  12.191 +                let
  12.192 +                  val bt' = term_of_poly baseT expT vs b'
  12.193 +                  val ct = term_of_poly baseT expT vs c
  12.194 +                  val t' = 
  12.195 +                    HOLogic.mk_binop "Rings.divide_class.divide" 
  12.196 +                      (term_of_poly baseT expT vs a', bt')
  12.197 +                  val asm = mk_asms baseT [bt']
  12.198 +                in SOME (t', asm) end
  12.199 +            end
  12.200 +        | _ => NONE : (term * term list) option
  12.201 +      end
  12.202 +  end
  12.203 +\<close>
  12.204 +
  12.205 +subsubsection \<open>Factor out to a common denominator for addition\<close>
  12.206 +ML \<open>
  12.207 +(* addition of fractions allows (at most) one non-fraction (a monomial) *)
  12.208 +fun check_frac_sum 
  12.209 +    (Const ("Groups.plus_class.plus", _) $ 
  12.210 +      (Const ("Rings.divide_class.divide", _) $ n1 $ d1) $
  12.211 +      (Const ("Rings.divide_class.divide", _) $ n2 $ d2))
  12.212 +    = SOME ((n1, d1), (n2, d2))
  12.213 +  | check_frac_sum 
  12.214 +    (Const ("Groups.plus_class.plus", _) $ 
  12.215 +      nofrac $ 
  12.216 +      (Const ("Rings.divide_class.divide", _) $ n2 $ d2))
  12.217 +    = SOME ((nofrac, HOLogic.mk_number HOLogic.realT 1), (n2, d2))
  12.218 +  | check_frac_sum 
  12.219 +    (Const ("Groups.plus_class.plus", _) $ 
  12.220 +      (Const ("Rings.divide_class.divide", _) $ n1 $ d1) $ 
  12.221 +      nofrac)
  12.222 +    = SOME ((n1, d1), (nofrac, HOLogic.mk_number HOLogic.realT 1))
  12.223 +  | check_frac_sum _ = NONE  
  12.224 +
  12.225 +(* prepare a term for addition by providing the least common denominator as a product
  12.226 +  assumes: is a term with outmost "+" and at least one outmost "/" in respective summands*)
  12.227 +fun common_nominator_p_ (_: theory) t =
  12.228 +  let val opt = check_frac_sum t
  12.229 +  in
  12.230 +    case opt of 
  12.231 +      NONE => NONE
  12.232 +    | SOME ((n1, d1), (n2, d2)) =>
  12.233 +      let
  12.234 +        val vs = TermC.vars_of t
  12.235 +      in
  12.236 +        case (poly_of_term vs d1, poly_of_term vs d2) of
  12.237 +          (SOME a, SOME b) =>
  12.238 +            let
  12.239 +              val ((a', b'), c) = gcd_poly a b
  12.240 +              val (baseT, expT) = (type_of n1, HOLogic.realT)
  12.241 +              val [d1', d2', c'] = map (term_of_poly baseT expT vs) [a', b', c]
  12.242 +              (*----- minimum of parentheses & nice result, but breaks tests: -------------
  12.243 +              val denom = HOLogic.mk_binop "Groups.times_class.times" 
  12.244 +                (HOLogic.mk_binop "Groups.times_class.times" (d1', d2'), c') -------------*)
  12.245 +              val denom =
  12.246 +                if c = [(1, replicate (length vs) 0)]
  12.247 +                then HOLogic.mk_binop "Groups.times_class.times" (d1', d2')
  12.248 +                else
  12.249 +                  HOLogic.mk_binop "Groups.times_class.times" (c',
  12.250 +                  HOLogic.mk_binop "Groups.times_class.times" (d1', d2')) (*--------------*)
  12.251 +              val t' =
  12.252 +                HOLogic.mk_binop "Groups.plus_class.plus"
  12.253 +                  (HOLogic.mk_binop "Rings.divide_class.divide"
  12.254 +                    (HOLogic.mk_binop "Groups.times_class.times" (n1, d2'), denom),
  12.255 +                  HOLogic.mk_binop "Rings.divide_class.divide" 
  12.256 +                    (HOLogic.mk_binop "Groups.times_class.times" (n2, d1'), denom))
  12.257 +              val asm = mk_asms baseT [d1', d2', c']
  12.258 +            in SOME (t', asm) end
  12.259 +        | _ => NONE : (term * term list) option
  12.260 +      end
  12.261 +  end
  12.262 +\<close>
  12.263 +
  12.264 +subsubsection \<open>Addition of at least one fraction within a sum\<close>
  12.265 +ML \<open>
  12.266 +(* add fractions
  12.267 +  assumes: is a term with outmost "+" and at least one outmost "/" in respective summands
  12.268 +  NOTE: the case "(_ + _) + _" need not be considered due to iterated addition.*)
  12.269 +fun add_fraction_p_ (_: theory) t =
  12.270 +  case check_frac_sum t of 
  12.271 +    NONE => NONE
  12.272 +  | SOME ((n1, d1), (n2, d2)) =>
  12.273 +    let
  12.274 +      val vs = TermC.vars_of t
  12.275 +    in
  12.276 +      case (poly_of_term vs n1, poly_of_term vs d1, poly_of_term vs n2, poly_of_term vs d2) of
  12.277 +        (SOME _, SOME a, SOME _, SOME b) =>
  12.278 +          let
  12.279 +            val ((a', b'), c) = gcd_poly a b
  12.280 +            val (baseT, expT) = (type_of n1, HOLogic.realT)
  12.281 +            val nomin = term_of_poly baseT expT vs 
  12.282 +              (((the (poly_of_term vs n1)) %%*%% b') %%+%% ((the (poly_of_term vs n2)) %%*%% a')) 
  12.283 +            val denom = term_of_poly baseT expT vs ((c %%*%% a') %%*%% b')
  12.284 +            val t' = HOLogic.mk_binop "Rings.divide_class.divide" (nomin, denom)
  12.285 +          in SOME (t', mk_asms baseT [denom]) end
  12.286 +      | _ => NONE : (term * term list) option
  12.287 +    end
  12.288 +\<close>
  12.289 +
  12.290 +section \<open>Embed cancellation and addition into rewriting\<close>
  12.291 +ML \<open>val thy = @{theory}\<close>
  12.292 +subsection \<open>Rulesets and predicate for embedding\<close>
  12.293 +ML \<open>
  12.294 +(* evaluates conditions in calculate_Rational *)
  12.295 +val calc_rat_erls =
  12.296 +  prep_rls'
  12.297 +    (Rule_Def.Repeat {id = "calc_rat_erls", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
  12.298 +      erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
  12.299 +      rules = 
  12.300 +        [Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),
  12.301 +        Rule.Eval ("Prog_Expr.is_const", Prog_Expr.eval_const "#is_const_"),
  12.302 +        Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}),
  12.303 +        Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false})], 
  12.304 +      scr = Rule.Empty_Prog});
  12.305 +
  12.306 +(* simplifies expressions with numerals;
  12.307 +   does NOT rearrange the term by AC-rewriting; thus terms with variables 
  12.308 +   need to have constants to be commuted together respectively           *)
  12.309 +val calculate_Rational =
  12.310 +  prep_rls' (Rule_Set.merge "calculate_Rational"
  12.311 +    (Rule_Def.Repeat {id = "divide", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
  12.312 +      erls = calc_rat_erls, srls = Rule_Set.Empty,
  12.313 +      calc = [], errpatts = [],
  12.314 +      rules = 
  12.315 +        [Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e"),
  12.316 +
  12.317 +        Rule.Thm ("minus_divide_left", ThmC.numerals_to_Free (@{thm minus_divide_left} RS @{thm sym})),
  12.318 +          (*SYM - ?x / ?y = - (?x / ?y)  may come from subst*)
  12.319 +        Rule.Thm ("rat_add", ThmC.numerals_to_Free @{thm rat_add}),
  12.320 +          (*"[| a is_const; b is_const; c is_const; d is_const |] ==> \
  12.321 +          \a / c + b / d = (a * d) / (c * d) + (b * c ) / (d * c)"*)
  12.322 +        Rule.Thm ("rat_add1", ThmC.numerals_to_Free @{thm rat_add1}),
  12.323 +          (*"[| a is_const; b is_const; c is_const |] ==> a / c + b / c = (a + b) / c"*)
  12.324 +        Rule.Thm ("rat_add2", ThmC.numerals_to_Free @{thm rat_add2}),
  12.325 +          (*"[| ?a is_const; ?b is_const; ?c is_const |] ==> ?a / ?c + ?b = (?a + ?b * ?c) / ?c"*)
  12.326 +        Rule.Thm ("rat_add3", ThmC.numerals_to_Free @{thm rat_add3}),
  12.327 +          (*"[| a is_const; b is_const; c is_const |] ==> a + b / c = (a * c) / c + b / c"\
  12.328 +          .... is_const to be omitted here FIXME*)
  12.329 +        
  12.330 +        Rule.Thm ("rat_mult", ThmC.numerals_to_Free @{thm rat_mult}), 
  12.331 +          (*a / b * (c / d) = a * c / (b * d)*)
  12.332 +        Rule.Thm ("times_divide_eq_right", ThmC.numerals_to_Free @{thm times_divide_eq_right}),
  12.333 +          (*?x * (?y / ?z) = ?x * ?y / ?z*)
  12.334 +        Rule.Thm ("times_divide_eq_left", ThmC.numerals_to_Free @{thm times_divide_eq_left}),
  12.335 +          (*?y / ?z * ?x = ?y * ?x / ?z*)
  12.336 +        
  12.337 +        Rule.Thm ("real_divide_divide1", ThmC.numerals_to_Free @{thm real_divide_divide1}),
  12.338 +          (*"?y ~= 0 ==> ?u / ?v / (?y / ?z) = ?u / ?v * (?z / ?y)"*)
  12.339 +        Rule.Thm ("divide_divide_eq_left", ThmC.numerals_to_Free @{thm divide_divide_eq_left}),
  12.340 +          (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
  12.341 +        
  12.342 +        Rule.Thm ("rat_power", ThmC.numerals_to_Free @{thm rat_power}),
  12.343 +          (*"(?a / ?b)  \<up>  ?n = ?a  \<up>  ?n / ?b  \<up>  ?n"*)
  12.344 +        
  12.345 +        Rule.Thm ("mult_cross", ThmC.numerals_to_Free @{thm mult_cross}),
  12.346 +          (*"[| b ~= 0; d ~= 0 |] ==> (a / b = c / d) = (a * d = b * c)*)
  12.347 +        Rule.Thm ("mult_cross1", ThmC.numerals_to_Free @{thm mult_cross1}),
  12.348 +          (*"   b ~= 0            ==> (a / b = c    ) = (a     = b * c)*)
  12.349 +        Rule.Thm ("mult_cross2", ThmC.numerals_to_Free @{thm mult_cross2})
  12.350 +          (*"           d ~= 0    ==> (a     = c / d) = (a * d =     c)*)], 
  12.351 +      scr = Rule.Empty_Prog})
  12.352 +    calculate_Poly);
  12.353 +
  12.354 +(*("is_expanded", ("Rational.is_expanded", eval_is_expanded ""))*)
  12.355 +fun eval_is_expanded (thmid:string) _ 
  12.356 +		       (t as (Const("Rational.is_expanded", _) $ arg)) thy = 
  12.357 +    if is_expanded arg
  12.358 +    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
  12.359 +	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
  12.360 +    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
  12.361 +	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
  12.362 +  | eval_is_expanded _ _ _ _ = NONE;
  12.363 +\<close> ML \<open>
  12.364 +\<close> ML \<open>
  12.365 +\<close>
  12.366 +
  12.367 +
  12.368 +section \<open>3============== NEW src/../ "Knowledge/Poly.thy" ================================\<close>
  12.369 +subsection \<open>auxiliary functions\<close>
  12.370 +ML \<open>
  12.371 +val thy = @{theory};
  12.372 +val poly_consts =
  12.373 +  ["Groups.plus_class.plus", "Groups.minus_class.minus",
  12.374 +  "Rings.divide_class.divide", "Groups.times_class.times",
  12.375 +  "Transcendental.powr"];
  12.376 +\<close> ML \<open>
  12.377 +int_ord: int * int -> order;
  12.378 +\<close> ML \<open>
  12.379 +fun int_ord (i1, i2) =
  12.380 +(@{print} {a = "int_ord (" ^ string_of_int i1 ^ ", " ^ string_of_int i2 ^ ") = ", z = Int.compare (i1, i2)};
  12.381 +  Int.compare (i1, i2));
  12.382 +\<close> ML \<open>
  12.383 +\<close>
  12.384 +subsubsection \<open>for predicates in specifications (ML)\<close>
  12.385 +ML \<open>
  12.386 +(*--- auxiliary for is_expanded_in, is_poly_in, has_degree_in ---*)
  12.387 +(*. a "monomial t in variable v" is a term t with
  12.388 +  either (1) v NOT existent in t, or (2) v contained in t,
  12.389 +  if (1) then degree 0
  12.390 +  if (2) then v is a factor on the very right, casually with exponent.*)
  12.391 +fun factor_right_deg (*case 2*)
  12.392 +	    (Const ("Groups.times_class.times", _) $
  12.393 +        t1 $ (Const ("Transcendental.powr",_) $ vv $ num)) v =
  12.394 +	  if vv = v andalso not (Prog_Expr.occurs_in v t1) then SOME (snd (HOLogic.dest_number num))
  12.395 +    else NONE
  12.396 +  | factor_right_deg (Const ("Transcendental.powr",_) $ vv $ num) v =
  12.397 +	   if (vv = v) then SOME (snd (HOLogic.dest_number num)) else NONE
  12.398 +  | factor_right_deg (Const ("Groups.times_class.times",_) $ t1 $ vv) v = 
  12.399 +	   if vv = v andalso not (Prog_Expr.occurs_in v t1) then SOME 1 else NONE
  12.400 +  | factor_right_deg vv v =
  12.401 +	  if (vv = v) then SOME 1 else NONE;    
  12.402 +fun mono_deg_in m v =  (*case 1*)
  12.403 +	if not (Prog_Expr.occurs_in v m) then (*case 1*) SOME 0 else factor_right_deg m v;
  12.404 +
  12.405 +fun expand_deg_in t v =
  12.406 +	let
  12.407 +    fun edi ~1 ~1 (Const ("Groups.plus_class.plus", _) $ t1 $ t2) =
  12.408 +          (case mono_deg_in t2 v of (* $ is left associative*)
  12.409 +            SOME d' => edi d' d' t1 | NONE => NONE)
  12.410 +      | edi ~1 ~1 (Const ("Groups.minus_class.minus", _) $ t1 $ t2) =
  12.411 +          (case mono_deg_in t2 v of
  12.412 +            SOME d' => edi d' d' t1 | NONE => NONE)
  12.413 +      | edi d dmax (Const ("Groups.minus_class.minus", _) $ t1 $ t2) =
  12.414 +          (case mono_deg_in t2 v of (*(d = 0 andalso d' = 0) handle 3+4-...4 +x*)
  12.415 +	        SOME d' => if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
  12.416 +          | NONE => NONE)
  12.417 +      | edi d dmax (Const ("Groups.plus_class.plus",_) $ t1 $ t2) =
  12.418 +          (case mono_deg_in t2 v of
  12.419 +            SOME d' =>    (*RL (d = 0 andalso d' = 0) need to handle 3+4-...4 +x*)
  12.420 +              if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
  12.421 +          | NONE => NONE)
  12.422 +      | edi ~1 ~1 t =
  12.423 +          (case mono_deg_in t v of d as SOME _ => d | NONE => NONE)
  12.424 +      | edi d dmax t = (*basecase last*)
  12.425 +    	    (case mono_deg_in t v of
  12.426 +    	      SOME d' => if d > d' orelse (d = 0 andalso d' = 0) then SOME dmax else NONE
  12.427 +		      | NONE => NONE)
  12.428 +	in edi ~1 ~1 t end;
  12.429 +
  12.430 +fun poly_deg_in t v =
  12.431 +	let
  12.432 +    fun edi ~1 ~1 (Const ("Groups.plus_class.plus",_) $ t1 $ t2) =
  12.433 +		    (case mono_deg_in t2 v of (* $ is left associative *)
  12.434 +		      SOME d' => edi d' d' t1
  12.435 +        | NONE => NONE)
  12.436 +	    | edi d dmax (Const ("Groups.plus_class.plus",_) $ t1 $ t2) =
  12.437 +		    (case mono_deg_in t2 v of
  12.438 +	        SOME d' =>    (*RL (d = 0 andalso (d' = 0)) handle 3+4-...4 +x*)
  12.439 +            if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
  12.440 +        | NONE => NONE)
  12.441 +	    | edi ~1 ~1 t =
  12.442 +        (case mono_deg_in t v of
  12.443 +		      d as SOME _ => d
  12.444 +        | NONE => NONE)
  12.445 +	    | edi d dmax t = (* basecase last *)
  12.446 +		    (case mono_deg_in t v of
  12.447 +		      SOME d' =>
  12.448 +            if d > d' orelse (d = 0 andalso d' = 0) then SOME dmax else NONE
  12.449 +        | NONE => NONE)
  12.450 +	in edi ~1 ~1 t end;
  12.451 +\<close>
  12.452 +
  12.453 +subsubsection \<open>for hard-coded AC rewriting (MG)\<close>
  12.454 +ML \<open>
  12.455 +(**. MG.03: make_polynomial_ ... uses SML-fun for ordering .**)
  12.456 +
  12.457 +(*FIXME.0401: make SML-order local to make_polynomial(_) *)
  12.458 +(*FIXME.0401: replace 'make_polynomial'(old) by 'make_polynomial_'(MG) *)
  12.459 +(* Polynom --> List von Monomen *) 
  12.460 +fun poly2list (Const ("Groups.plus_class.plus",_) $ t1 $ t2) = 
  12.461 +    (poly2list t1) @ (poly2list t2)
  12.462 +  | poly2list t = [t];
  12.463 +
  12.464 +(* Monom --> Liste von Variablen *)
  12.465 +fun monom2list (Const ("Groups.times_class.times",_) $ t1 $ t2) = 
  12.466 +    (monom2list t1) @ (monom2list t2)
  12.467 +  | monom2list t = [t];
  12.468 +
  12.469 +\<close> ML \<open>
  12.470 +(* use this fun ONLY locally: makes negative numbers with "-" instead of "~" for poly-orders*)
  12.471 +fun dest_number' t =
  12.472 +  let
  12.473 +    val i = TermC.num_of_term t
  12.474 +  in
  12.475 +    if i >= 0 then HOLogic.dest_number t |> snd |> string_of_int
  12.476 +    else (i * ~1) |> string_of_int |> curry op ^ "-"
  12.477 +  end
  12.478 +\<close> ML \<open>
  12.479 +(* liefert Variablenname (String) einer Variablen und Basis bei Potenz *)
  12.480 +\<close> ML \<open> (* orig.code *)
  12.481 +fun get_basStr (Const ("Transcendental.powr",_) $ Free (str, _) $ _) = str
  12.482 +  | get_basStr (Free (str, _)) = str
  12.483 +  | get_basStr _ = "|||"; (* gross gewichtet; für Brüch ect. *)
  12.484 +(*| get_basStr t = 
  12.485 +    raise ERROR("get_basStr: called with t= "^(UnparseC.term t));*)
  12.486 +\<close> ML \<open> (* \<longrightarrow> Poly.thy *)
  12.487 +fun get_basStr (Const ("Transcendental.powr",_) $ Free (str, _) $ _) = str
  12.488 +  | get_basStr (Const ("Transcendental.powr",_) $ n $ _) = dest_number' n
  12.489 +  | get_basStr (Free (str, _)) = str
  12.490 +  | get_basStr t =
  12.491 +    if TermC.is_num t then dest_number' t
  12.492 +    else "|||"; (* gross gewichtet; für Brüche ect. *)
  12.493 +\<close> ML \<open>
  12.494 +
  12.495 +(* liefert Hochzahl (String) einer Variablen bzw Gewichtstring (zum Sortieren) *)
  12.496 +\<close> ML \<open> (*original*)
  12.497 +fun get_potStr (Const ("Transcendental.powr",_) $ Free _ $ Free (str, _)) = str
  12.498 +  | get_potStr (Const ("Transcendental.powr",_) $ Free _ $ _ ) = "|||" (* gross gewichtet *)
  12.499 +  | get_potStr (Free (_, _)) = "---" (* keine Hochzahl --> kleinst gewichtet *)
  12.500 +  | get_potStr _ = "||||||"; (* gross gewichtet; für Brüch ect. *)
  12.501 +(*| get_potStr t = 
  12.502 +    raise ERROR("get_potStr: called with t= "^(UnparseC.term t));*)
  12.503 +\<close> ML \<open> (* \<longrightarrow> Poly.thy *)
  12.504 +fun get_potStr (Const ("Transcendental.powr", _) $ Free _ $ Free (str, _)) = str
  12.505 +  | get_potStr (Const ("Transcendental.powr", _) $ Free _ $ t) =
  12.506 +    if TermC.is_num t then dest_number' t else "|||"
  12.507 +  | get_potStr (Free _) = "---" (* keine Hochzahl --> kleinst gewichtet *)
  12.508 +  | get_potStr _ = "||||||"; (* gross gewichtet; für Brüch ect. *)
  12.509 +
  12.510 +(* Umgekehrte string_ord *)
  12.511 +val string_ord_rev =  rev_order o string_ord;
  12.512 +		
  12.513 + (* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen) 
  12.514 +    innerhalb eines Monomes:
  12.515 +    - zuerst lexikographisch nach Variablenname 
  12.516 +    - wenn gleich: nach steigender Potenz *)
  12.517 +\<close> ML \<open> (* \<longrightarrow> Poly,thy *)
  12.518 +fun var_ord (a, b) = prod_ord string_ord string_ord 
  12.519 +  ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
  12.520 +fun var_ord (a, b) = 
  12.521 +(@{print} {a = "var_ord ", a_b = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")",
  12.522 +   sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
  12.523 +  prod_ord string_ord string_ord 
  12.524 +  ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b))
  12.525 +);
  12.526 +\<close> ML \<open>
  12.527 +
  12.528 +(* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen); 
  12.529 +   verwendet zum Sortieren von Monomen mittels Gesamtgradordnung:
  12.530 +   - zuerst lexikographisch nach Variablenname 
  12.531 +   - wenn gleich: nach sinkender Potenz*)
  12.532 +\<close> ML \<open> (* DONE Poly.thy *)
  12.533 +fun var_ord_revPow (a, b: term) =
  12.534 +  prod_ord string_ord string_ord_rev 
  12.535 +    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
  12.536 +fun var_ord_revPow (a, b: term) = 
  12.537 +(@{print} {a = "var_ord_revPow ", at_bt = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")",
  12.538 +    sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
  12.539 +  prod_ord string_ord string_ord_rev 
  12.540 +    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b))
  12.541 +);
  12.542 +
  12.543 +
  12.544 +(* Ordnet ein Liste von Variablen (und Potenzen) lexikographisch *)
  12.545 +val sort_varList = sort var_ord;
  12.546 +\<close> ML \<open>
  12.547 +fun sort_varList ts =
  12.548 +(@{print} {a = "sort_varList", args = UnparseC.terms ts};
  12.549 +  sort var_ord ts);
  12.550 +\<close> ML \<open>
  12.551 +
  12.552 +(* Entfernet aeussersten Operator (Wurzel) aus einem Term und schreibt 
  12.553 +   Argumente in eine Liste *)
  12.554 +\<close> ML \<open> (* DONE Poly.thy *)
  12.555 +fun args u : term list =
  12.556 +  let
  12.557 +    fun stripc (f $ t, ts) = stripc (f, t::ts)
  12.558 +  	  | stripc (t as Free _, ts) = (t::ts)
  12.559 +  	  | stripc (_, ts) = ts
  12.560 +  in stripc (u, []) end;
  12.561 +                                    
  12.562 +(* liefert True, falls der Term (Liste von Termen) nur Zahlen 
  12.563 +   (keine Variablen) enthaelt *)
  12.564 +fun filter_num ts = fold (curry and_) (map TermC.is_num ts) true
  12.565 +
  12.566 +(* liefert True, falls der Term nur Zahlen (keine Variablen) enthaelt 
  12.567 +   dh. er ist ein numerischer Wert und entspricht einem Koeffizienten *)
  12.568 +fun is_nums t = filter_num [t];
  12.569 +
  12.570 +(* Berechnet den Gesamtgrad eines Monoms *)
  12.571 +\<close> ML \<open> (* original *)
  12.572 +local 
  12.573 +    fun counter (n, []) = n
  12.574 +      | counter (n, x :: xs) = 
  12.575 +	if (is_nums x) then
  12.576 +	    counter (n, xs) 
  12.577 +	else 
  12.578 +	    (case x of 
  12.579 +		 (Const ("Transcendental.powr", _) $ Free _ $ Free (str_h, T)) => 
  12.580 +		     if (is_nums (Free (str_h, T))) then
  12.581 +			 counter (n + (the (TermC.int_opt_of_string str_h)), xs)
  12.582 +		     else counter (n + 1000, xs) (*FIXME.MG?!*)
  12.583 +	       | (Const ("Transcendental.powr", _) $ Free _ $ _ ) => 
  12.584 +		     counter (n + 1000, xs) (*FIXME.MG?!*)
  12.585 +	       | (Free _) => counter (n + 1, xs)
  12.586 +	     (*| _ => raise ERROR("monom_degree: called with factor: "^(UnparseC.term x)))*)
  12.587 +	       | _ => counter (n + 10000, xs)) (*FIXME.MG?! ... Brüche ect.*)
  12.588 +in  
  12.589 +    fun monom_degree l = counter (0, l) 
  12.590 +end;(*local*)
  12.591 +\<close> ML \<open> (* original reformatted *)
  12.592 +local 
  12.593 +  fun counter (n, []) = n
  12.594 +    | counter (n, x :: xs) = 
  12.595 +	    if (is_nums x) then counter (n, xs) 
  12.596 +	    else 
  12.597 +	      (case x of 
  12.598 +		      (Const ("Transcendental.powr", _) $ Free _ $ Free (str_h, T)) => 
  12.599 +		        if (is_nums (Free (str_h, T))) then counter (n + (the (TermC.int_opt_of_string str_h)), xs)
  12.600 +		        else counter (n + 1000, xs) (*FIXME.MG?!*)
  12.601 +	       | (Const ("Transcendental.powr", _) $ Free _ $ _ ) => 
  12.602 +		         counter (n + 1000, xs) (*FIXME.MG?!*)
  12.603 +	       | (Free _) => counter (n + 1, xs)
  12.604 +	       | _ => counter (n + 10000, xs)) (*FIXME.MG?! ... Brüche ect.*)
  12.605 +in  
  12.606 +    fun monom_degree l = counter (0, l) 
  12.607 +end;(*local*)
  12.608 +\<close> ML \<open>
  12.609 +is_nums: term -> bool
  12.610 +\<close> ML \<open>
  12.611 +\<close> ML \<open> (* DONE Poly.thy *)
  12.612 +(**)local(**)
  12.613 +  fun counter (n, []) = n
  12.614 +    | counter (n, x :: xs) = 
  12.615 +	    if (is_nums x) then counter (n, xs)
  12.616 +	    else 
  12.617 +	      (case x of 
  12.618 +		      (Const ("Transcendental.powr", _) $ Free _ $ t) =>
  12.619 +            if TermC.is_num t
  12.620 +            then counter (t |> HOLogic.dest_number |> snd |> curry op + n, xs)
  12.621 +            else counter (n + 1000, xs) (*FIXME.MG?!*)
  12.622 +	      | (Const ("Num.numeral_class.numeral", _) $ num) =>
  12.623 +            counter (n + 1 + HOLogic.dest_numeral num, xs)
  12.624 +	      | _ => counter (n + 1, xs)) (*FIXME.MG?! ... Brüche ect.*)
  12.625 +(**)in(**)
  12.626 +  fun monom_degree l = counter (0, l) 
  12.627 +(**)end;(*local*)
  12.628 +\<close> ML \<open>
  12.629 +
  12.630 +(* wie Ordnung dict_ord (lexicographische Ordnung zweier Listen, mit Vergleich 
  12.631 +   der Listen-Elemente mit elem_ord) - Elemente die Bedingung cond erfuellen, 
  12.632 +   werden jedoch dabei ignoriert (uebersprungen)  *)
  12.633 +fun dict_cond_ord _ _ ([], []) = EQUAL
  12.634 +  | dict_cond_ord _ _ ([], _ :: _) = LESS
  12.635 +  | dict_cond_ord _ _ (_ :: _, []) = GREATER
  12.636 +  | dict_cond_ord elem_ord cond (x :: xs, y :: ys) =
  12.637 +    (case (cond x, cond y) of 
  12.638 +	    (false, false) =>
  12.639 +        (case elem_ord (x, y) of 
  12.640 +				  EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
  12.641 +			  | ord => ord)
  12.642 +    | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
  12.643 +    | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
  12.644 +    | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys) );
  12.645 +fun dict_cond_ord _ _ ([], [])     = (@{print} {a = "dict_cond_ord ([], [])"}; EQUAL)
  12.646 +  | dict_cond_ord _ _ ([], _ :: _) = (@{print} {a = "dict_cond_ord ([], _ :: _)"}; LESS)
  12.647 +  | dict_cond_ord _ _ (_ :: _, []) = (@{print} {a = "dict_cond_ord (_ :: _, [])"}; GREATER)
  12.648 +  | dict_cond_ord elem_ord cond (x :: xs, y :: ys) =
  12.649 +    (@{print} {a = "dict_cond_ord", args = "(" ^ UnparseC.terms (x :: xs) ^ ", " ^ UnparseC.terms (y :: ys) ^ ")", 
  12.650 +      is_nums = "(" ^ LibraryC.bool2str (cond x) ^ ", " ^ LibraryC.bool2str (cond y) ^ ")"};
  12.651 +     case (cond x, cond y) of 
  12.652 +	    (false, false) =>
  12.653 +        (case elem_ord (x, y) of 
  12.654 +				  EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
  12.655 +			  | ord => ord)
  12.656 +    | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
  12.657 +    | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
  12.658 +    | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys) );
  12.659 +
  12.660 +(* Gesamtgradordnung zum Vergleich von Monomen (Liste von Variablen/Potenzen):
  12.661 +   zuerst nach Gesamtgrad int_ord, bei gleichem Gesamtgrad lexikographisch ordnen - 
  12.662 +   dabei werden Koeffizienten ignoriert (2*3*a \<up> 2*4*b gilt wie a \<up> 2*b) *)
  12.663 +fun degree_ord (xs, ys) =
  12.664 +	prod_ord int_ord (dict_cond_ord var_ord_revPow is_nums)
  12.665 +	  ((monom_degree xs, xs), (monom_degree ys, ys));
  12.666 +
  12.667 +fun hd_str str = substring (str, 0, 1);
  12.668 +fun tl_str str = substring (str, 1, (size str) - 1);
  12.669 +
  12.670 +(* liefert nummerischen Koeffizienten eines Monoms oder NONE *)
  12.671 +\<close> ML \<open> (* DONE Poly.thy *)
  12.672 +fun get_koeff_of_mon [] =  raise ERROR "get_koeff_of_mon: called with l = []"
  12.673 +  | get_koeff_of_mon (x :: _) = if is_nums x then SOME x else NONE;
  12.674 +
  12.675 +(* wandelt Koeffizient in (zum sortieren geeigneten) String um *)
  12.676 +fun koeff2ordStr (SOME t) =
  12.677 +    if TermC.is_num t
  12.678 +    then 
  12.679 +      if (t |> HOLogic.dest_number |> snd) < 0
  12.680 +      then (t |> HOLogic.dest_number |> snd |> curry op * ~1 |> string_of_int) ^ "0"  (* 3 < -3 *)
  12.681 +      else (t |> HOLogic.dest_number |> snd |> string_of_int)
  12.682 +    else "aaa"                                                      (* "num.Ausdruck" --> gross *)
  12.683 +  | koeff2ordStr NONE = "---";                                     (* "kein Koeff" --> kleinste *)
  12.684 +
  12.685 +(* Order zum Vergleich von Koeffizienten (strings): 
  12.686 +   "kein Koeff" < "0" < "1" < "-1" < "2" < "-2" < ... < "num.Ausdruck" *)
  12.687 +\<close> ML \<open> (* DONE Poly.thy *)
  12.688 +fun compare_koeff_ord (xs, ys) = string_ord
  12.689 +  ((koeff2ordStr o get_koeff_of_mon) xs,
  12.690 +   (koeff2ordStr o get_koeff_of_mon) ys);
  12.691 +fun compare_koeff_ord (xs, ys) =
  12.692 +(@{print} {a = "compare_koeff_ord ", ats_bts = "(" ^ UnparseC.terms xs ^ ", " ^ UnparseC.terms ys ^ ")",
  12.693 +    sort_args = "(" ^ (koeff2ordStr o get_koeff_of_mon) xs ^ ", " ^ (koeff2ordStr o get_koeff_of_mon) ys ^ ")"};
  12.694 +  string_ord
  12.695 +  ((koeff2ordStr o get_koeff_of_mon) xs,
  12.696 +   (koeff2ordStr o get_koeff_of_mon) ys)
  12.697 +);
  12.698 +
  12.699 +(* Gesamtgradordnung: falls degree_ord EQUAL + Ordnen nach Koeffizienten falls EQUAL *)
  12.700 +fun koeff_degree_ord (xs, ys) =
  12.701 +	    prod_ord degree_ord compare_koeff_ord ((xs, xs), (ys, ys));
  12.702 +
  12.703 +(* Ordnet ein Liste von Monomen (Monom = Liste von Variablen) mittels 
  12.704 +   Gesamtgradordnung *)
  12.705 +val sort_monList = sort koeff_degree_ord;
  12.706 +
  12.707 +(* Alternativ zu degree_ord koennte auch die viel einfachere und 
  12.708 +   kuerzere Ordnung simple_ord verwendet werden - ist aber nicht 
  12.709 +   fuer unsere Zwecke geeignet!
  12.710 +                                                                    
  12.711 +fun simple_ord (al,bl: term list) = dict_ord string_ord 
  12.712 +	 (map get_basStr al, map get_basStr bl); 
  12.713 +
  12.714 +val sort_monList = sort simple_ord; *)
  12.715 +
  12.716 +(* aus 2 Variablen wird eine Summe bzw ein Produkt erzeugt 
  12.717 +   (mit gewuenschtem Typen T) *)
  12.718 +fun plus T = Const ("Groups.plus_class.plus", [T,T] ---> T);
  12.719 +fun mult T = Const ("Groups.times_class.times", [T,T] ---> T);
  12.720 +fun binop op_ t1 t2 = op_ $ t1 $ t2;
  12.721 +fun create_prod T (a,b) = binop (mult T) a b;
  12.722 +fun create_sum T (a,b) = binop (plus T) a b;
  12.723 +
  12.724 +(* löscht letztes Element einer Liste *)
  12.725 +fun drop_last l = take ((length l)-1,l);
  12.726 +
  12.727 +(* Liste von Variablen --> Monom *)
  12.728 +fun create_monom T vl = foldr (create_prod T) (drop_last vl, last_elem vl);
  12.729 +(* Bemerkung: 
  12.730 +   foldr bewirkt rechtslastige Klammerung des Monoms - ist notwendig, damit zwei 
  12.731 +   gleiche Monome zusammengefasst werden können (collect_numerals)! 
  12.732 +   zB: 2*(x*(y*z)) + 3*(x*(y*z)) --> (2+3)*(x*(y*z))*)
  12.733 +
  12.734 +(* Liste von Monomen --> Polynom *)	
  12.735 +fun create_polynom T ml = foldl (create_sum T) (hd ml, tl ml);
  12.736 +(* Bemerkung: 
  12.737 +   foldl bewirkt linkslastige Klammerung des Polynoms (der Summanten) - 
  12.738 +   bessere Darstellung, da keine Klammern sichtbar! 
  12.739 +   (und discard_parentheses in make_polynomial hat weniger zu tun) *)
  12.740 +
  12.741 +(* sorts the variables (faktors) of an expanded polynomial lexicographical *)
  12.742 +fun sort_variables t = 
  12.743 +  let
  12.744 +  	val ll = map monom2list (poly2list t);
  12.745 +  	val lls = map sort_varList ll; 
  12.746 +  	val T = type_of t;
  12.747 +  	val ls = map (create_monom T) lls;
  12.748 +  in create_polynom T ls end;
  12.749 +
  12.750 +(* sorts the monoms of an expanded and variable-sorted polynomial 
  12.751 +   by total_degree *)
  12.752 +fun sort_monoms t = 
  12.753 +  let
  12.754 +  	val ll =  map monom2list (poly2list t);
  12.755 +  	val lls = sort_monList ll;
  12.756 +  	val T = Term.type_of t;
  12.757 +  	val ls = map (create_monom T) lls;
  12.758 +  in create_polynom T ls end;
  12.759 +\<close>
  12.760 +
  12.761 +subsubsection \<open>rewrite order for hard-coded AC rewriting\<close>
  12.762 +ML \<open>
  12.763 +local (*. for make_polynomial .*)
  12.764 +
  12.765 +open Term;  (* for type order = EQUAL | LESS | GREATER *)
  12.766 +
  12.767 +fun pr_ord EQUAL = "EQUAL"
  12.768 +  | pr_ord LESS  = "LESS"
  12.769 +  | pr_ord GREATER = "GREATER";
  12.770 +
  12.771 +fun dest_hd' (Const (a, T)) =                          (* ~ term.ML *)
  12.772 +  (case a of
  12.773 +     "Transcendental.powr" => ((("|||||||||||||", 0), T), 0)    (*WN greatest string*)
  12.774 +   | _ => (((a, 0), T), 0))
  12.775 +  | dest_hd' (Free (a, T)) = (((a, 0), T), 1)(*TODOO handle this as numeral, too? see EqSystem.thy*)
  12.776 +  | dest_hd' (Var v) = (v, 2)
  12.777 +  | dest_hd' (Bound i) = ((("", i), dummyT), 3)
  12.778 +  | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4)
  12.779 +  | dest_hd' t = raise TERM ("dest_hd'", [t]);
  12.780 +
  12.781 +fun size_of_term' (Const(str,_) $ t) =
  12.782 +  if "Transcendental.powr"= str then 1000 + size_of_term' t else 1+size_of_term' t(*WN*)
  12.783 +  | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body
  12.784 +  | size_of_term' (f$t) = size_of_term' f  +  size_of_term' t
  12.785 +  | size_of_term' _ = 1;
  12.786 +
  12.787 +fun term_ord' pr thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
  12.788 +    (case term_ord' pr thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord)
  12.789 +  | term_ord' pr thy (t, u) =
  12.790 +    (if pr then 
  12.791 +	   let
  12.792 +       val (f, ts) = strip_comb t and (g, us) = strip_comb u;
  12.793 +       val _ = tracing ("t= f@ts= \"" ^ UnparseC.term_in_thy thy f ^ "\" @ \"[" ^
  12.794 +         commas (map (UnparseC.term_in_thy thy) ts) ^ "]\"");
  12.795 +       val _ = tracing("u= g@us= \"" ^ UnparseC.term_in_thy thy g ^ "\" @ \"[" ^
  12.796 +         commas (map (UnparseC.term_in_thy thy) us) ^ "]\"");
  12.797 +       val _ = tracing ("size_of_term(t,u)= (" ^ string_of_int (size_of_term' t) ^ ", " ^
  12.798 +         string_of_int (size_of_term' u) ^ ")");
  12.799 +       val _ = tracing ("hd_ord(f,g)      = " ^ (pr_ord o hd_ord) (f,g));
  12.800 +       val _ = tracing ("terms_ord(ts,us) = " ^ (pr_ord o terms_ord str false) (ts, us));
  12.801 +       val _ = tracing ("-------");
  12.802 +     in () end
  12.803 +       else ();
  12.804 +	 case int_ord (size_of_term' t, size_of_term' u) of
  12.805 +	   EQUAL =>
  12.806 +	     let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
  12.807 +	       (case hd_ord (f, g) of EQUAL => (terms_ord str pr) (ts, us) 
  12.808 +	     | ord => ord)
  12.809 +	     end
  12.810 +	 | ord => ord)
  12.811 +and hd_ord (f, g) =                                        (* ~ term.ML *)
  12.812 +  prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord (dest_hd' f, dest_hd' g)
  12.813 +and terms_ord _ pr (ts, us) = 
  12.814 +    list_ord (term_ord' pr (ThyC.get_theory "Isac_Knowledge"))(ts, us);
  12.815 +
  12.816 +in
  12.817 +
  12.818 +fun ord_make_polynomial (pr:bool) thy (_: subst) tu = 
  12.819 +    (term_ord' pr thy(***) tu = LESS );
  12.820 +
  12.821 +end;(*local*)
  12.822 +
  12.823 +Rewrite_Ord.rew_ord' := overwritel (! Rewrite_Ord.rew_ord', (* TODO: make analogous to KEStore_Elems.add_mets *)
  12.824 +[("termlessI", termlessI), ("ord_make_polynomial", ord_make_polynomial false thy)]);
  12.825 +\<close>
  12.826 +
  12.827 +subsection \<open>predicates\<close>
  12.828 +subsubsection \<open>in specifications\<close>
  12.829 +ML \<open>
  12.830 +(* is_polyrat_in becomes true, if no bdv is in the denominator of a fraction*)
  12.831 +fun is_polyrat_in t v = 
  12.832 +  let
  12.833 +   	fun finddivide (_ $ _ $ _ $ _) _ = raise ERROR("is_polyrat_in:")
  12.834 +	    (* at the moment there is no term like this, but ....*)
  12.835 +	  | finddivide (Const ("Rings.divide_class.divide",_) $ _ $ b) v = not (Prog_Expr.occurs_in v b)
  12.836 +	  | finddivide (_ $ t1 $ t2) v = finddivide t1 v orelse finddivide t2 v
  12.837 +	  | finddivide (_ $ t1) v = finddivide t1 v
  12.838 +	  | finddivide _ _ = false;
  12.839 +  in finddivide t v end;
  12.840 +    
  12.841 +fun is_expanded_in t v = case expand_deg_in t v of SOME _ => true | NONE => false;
  12.842 +fun is_poly_in t v =     case poly_deg_in t v of SOME _ => true | NONE => false;
  12.843 +fun has_degree_in t v =  case expand_deg_in t v of SOME d => d | NONE => ~1;
  12.844 +
  12.845 +(*.the expression contains + - * ^ only ?
  12.846 +   this is weaker than 'is_polynomial' !.*)
  12.847 +\<close> ML \<open> (* OLD code before dropping ThmC.numerals_to_Free*)
  12.848 +fun is_polyexp (Free _) = true
  12.849 +  | is_polyexp (Const _) = true (* potential danger: bdv is not considered *)
  12.850 +  | is_polyexp (Const ("Groups.plus_class.plus",_) $ Free _ $ Free _) = true
  12.851 +  | is_polyexp (Const ("Groups.minus_class.minus",_) $ Free _ $ Free _) = true
  12.852 +  | is_polyexp (Const ("Groups.times_class.times",_) $ Free _ $ Free _) = true
  12.853 +  | is_polyexp (Const ("Transcendental.powr",_) $ Free _ $ Free _) = true
  12.854 +  | is_polyexp (Const ("Groups.plus_class.plus",_) $ t1 $ t2) = 
  12.855 +               ((is_polyexp t1) andalso (is_polyexp t2))
  12.856 +  | is_polyexp (Const ("Groups.minus_class.minus",_) $ t1 $ t2) = 
  12.857 +               ((is_polyexp t1) andalso (is_polyexp t2))
  12.858 +  | is_polyexp (Const ("Groups.times_class.times",_) $ t1 $ t2) = 
  12.859 +               ((is_polyexp t1) andalso (is_polyexp t2))
  12.860 +  | is_polyexp (Const ("Transcendental.powr",_) $ t1 $ t2) = 
  12.861 +               ((is_polyexp t1) andalso (is_polyexp t2))
  12.862 +  | is_polyexp _ = false;
  12.863 +\<close> ML \<open> (* DONE Poly.thy *)
  12.864 +fun is_polyexp (Free _) = true
  12.865 +  | is_polyexp (Const _) = true (* potential danger: bdv is not considered *)
  12.866 +  | is_polyexp (Const ("Groups.plus_class.plus",_) $ Free _ $ num) =
  12.867 +    if TermC.is_num num then true
  12.868 +    else if TermC.is_variable num then true
  12.869 +    else is_polyexp num
  12.870 +  | is_polyexp (Const ("Groups.plus_class.plus",_) $ num $ Free _) =
  12.871 +    if TermC.is_num num then true
  12.872 +    else if TermC.is_variable num then true
  12.873 +    else is_polyexp num
  12.874 +  | is_polyexp (Const ("Groups.minus_class.minus",_) $ Free _ $ num) =
  12.875 +    if TermC.is_num num then true
  12.876 +    else if TermC.is_variable num then true
  12.877 +    else is_polyexp num
  12.878 +  | is_polyexp (Const ("Groups.times_class.times",_) $ num $ Free _) =
  12.879 +    if TermC.is_num num then true
  12.880 +    else if TermC.is_variable num then true
  12.881 +    else is_polyexp num
  12.882 +  | is_polyexp (Const ("Transcendental.powr",_) $ Free _ $ num) =
  12.883 +    if TermC.is_num num then true
  12.884 +    else if TermC.is_variable num then true
  12.885 +    else is_polyexp num
  12.886 +  | is_polyexp (Const ("Groups.plus_class.plus",_) $ t1 $ t2) = 
  12.887 +    ((is_polyexp t1) andalso (is_polyexp t2))
  12.888 +  | is_polyexp (Const ("Groups.minus_class.minus",_) $ t1 $ t2) = 
  12.889 +    ((is_polyexp t1) andalso (is_polyexp t2))
  12.890 +  | is_polyexp (Const ("Groups.times_class.times",_) $ t1 $ t2) = 
  12.891 +    ((is_polyexp t1) andalso (is_polyexp t2))
  12.892 +  | is_polyexp (Const ("Transcendental.powr",_) $ t1 $ t2) = 
  12.893 +    ((is_polyexp t1) andalso (is_polyexp t2))
  12.894 +  | is_polyexp num = TermC.is_num num;
  12.895 +\<close>
  12.896 +
  12.897 +subsubsection \<open>for hard-coded AC rewriting\<close>
  12.898 +ML \<open>
  12.899 +(* auch Klammerung muss übereinstimmen;
  12.900 +   sort_variables klammert Produkte rechtslastig*)
  12.901 +fun is_multUnordered t = ((is_polyexp t) andalso not (t = sort_variables t));
  12.902 +
  12.903 +fun is_addUnordered t = ((is_polyexp t) andalso not (t = sort_monoms t));
  12.904 +\<close>
  12.905 +
  12.906 +subsection \<open>evaluations functions\<close>
  12.907 +subsubsection \<open>for predicates\<close>                   
  12.908 +ML \<open>
  12.909 +fun eval_is_polyrat_in _ _(p as (Const ("Poly.is_polyrat_in",_) $ t $ v)) _  =
  12.910 +    if is_polyrat_in t v 
  12.911 +    then SOME ((UnparseC.term p) ^ " = True",
  12.912 +	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
  12.913 +    else SOME ((UnparseC.term p) ^ " = True",
  12.914 +	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
  12.915 +  | eval_is_polyrat_in _ _ _ _ = ((*tracing"### no matches";*) NONE);
  12.916 +
  12.917 +(*("is_expanded_in", ("Poly.is_expanded_in", eval_is_expanded_in ""))*)
  12.918 +fun eval_is_expanded_in _ _ 
  12.919 +       (p as (Const ("Poly.is_expanded_in",_) $ t $ v)) _ =
  12.920 +    if is_expanded_in t v
  12.921 +    then SOME ((UnparseC.term p) ^ " = True",
  12.922 +	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
  12.923 +    else SOME ((UnparseC.term p) ^ " = True",
  12.924 +	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
  12.925 +  | eval_is_expanded_in _ _ _ _ = NONE;
  12.926 +
  12.927 +(*("is_poly_in", ("Poly.is_poly_in", eval_is_poly_in ""))*)
  12.928 +fun eval_is_poly_in _ _ 
  12.929 +       (p as (Const ("Poly.is_poly_in",_) $ t $ v)) _ =
  12.930 +    if is_poly_in t v
  12.931 +    then SOME ((UnparseC.term p) ^ " = True",
  12.932 +	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
  12.933 +    else SOME ((UnparseC.term p) ^ " = True",
  12.934 +	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
  12.935 +  | eval_is_poly_in _ _ _ _ = NONE;
  12.936 +
  12.937 +(*("has_degree_in", ("Poly.has_degree_in", eval_has_degree_in ""))*)
  12.938 +fun eval_has_degree_in _ _ 
  12.939 +	     (p as (Const ("Poly.has_degree_in",_) $ t $ v)) _ =
  12.940 +    let val d = has_degree_in t v
  12.941 +	val d' = TermC.term_of_num HOLogic.realT d
  12.942 +    in SOME ((UnparseC.term p) ^ " = " ^ (string_of_int d),
  12.943 +	      HOLogic.Trueprop $ (TermC.mk_equality (p, d')))
  12.944 +    end
  12.945 +  | eval_has_degree_in _ _ _ _ = NONE;
  12.946 +
  12.947 +(*("is_polyexp", ("Poly.is_polyexp", eval_is_polyexp ""))*)
  12.948 +fun eval_is_polyexp (thmid:string) _ 
  12.949 +		       (t as (Const("Poly.is_polyexp", _) $ arg)) thy = 
  12.950 +    if is_polyexp arg
  12.951 +    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
  12.952 +	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
  12.953 +    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
  12.954 +	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
  12.955 +  | eval_is_polyexp _ _ _ _ = NONE; 
  12.956 +\<close>
  12.957 +
  12.958 +subsubsection \<open>for hard-coded AC rewriting\<close>
  12.959 +ML \<open>
  12.960 +(*WN.18.6.03 *)
  12.961 +(*("is_addUnordered", ("Poly.is_addUnordered", eval_is_addUnordered ""))*)
  12.962 +fun eval_is_addUnordered (thmid:string) _ 
  12.963 +		       (t as (Const("Poly.is_addUnordered", _) $ arg)) thy = 
  12.964 +    if is_addUnordered arg
  12.965 +    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
  12.966 +	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
  12.967 +    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
  12.968 +	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
  12.969 +  | eval_is_addUnordered _ _ _ _ = NONE; 
  12.970 +
  12.971 +fun eval_is_multUnordered (thmid:string) _ 
  12.972 +		       (t as (Const("Poly.is_multUnordered", _) $ arg)) thy = 
  12.973 +    if is_multUnordered arg
  12.974 +    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
  12.975 +	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
  12.976 +    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
  12.977 +	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
  12.978 +  | eval_is_multUnordered _ _ _ _ = NONE; 
  12.979 +\<close>
  12.980 +
  12.981  
  12.982  section \<open>===================================================================================\<close>
  12.983  ML \<open>
  12.984 @@ -109,886 +1074,8 @@
  12.985  \<close> ML \<open>
  12.986  \<close>
  12.987  
  12.988 -section \<open>============== check test/../ "Knowledge/poly.sml" ===============================\<close>
  12.989 -ML \<open>
  12.990 -\<close> ML \<open>
  12.991 -(* testexamples for Poly, polynomials
  12.992 -   author: Matthias Goldgruber 2003
  12.993 -   (c) due to copyright terms
  12.994 -
  12.995 -12345678901234567890123456789012345678901234567890123456789012345678901234567890
  12.996 -        10        20        30        40        50        60        70        80
  12.997 -LEGEND:
  12.998 -WN060104: examples marked with 'SPB' came into 'exp_IsacCore_Simp_Poly_Book.xml'
  12.999 -          examples marked with 'SPO' came into 'exp_IsacCore_Simp_Poly_Other.xml'
 12.1000 -*)
 12.1001 -
 12.1002 -"--------------------------------------------------------";
 12.1003 -"--------------------------------------------------------";
 12.1004 -"table of contents --------------------------------------";
 12.1005 -"--------------------------------------------------------";
 12.1006 -"----------- fun is_polyexp --------------------------------------------------------------------";
 12.1007 -"----------- fun has_degree_in -----------------------------------------------------------------";
 12.1008 -"----------- fun mono_deg_in -------------------------------------------------------------------";
 12.1009 -"----------- fun mono_deg_in -------------------------------------------------------------------";
 12.1010 -"----------- eval_ for is_expanded_in, is_poly_in, has_degree_in -------------------------------";
 12.1011 -"-------- investigate (new 2002) uniary minus -----------";
 12.1012 -"-------- check make_polynomial with simple terms -------";
 12.1013 -"-------- fun is_multUnordered --------------------------";
 12.1014 -"-------- examples from textbook Schalk I ---------------";
 12.1015 -"-------- check pbl  'polynomial simplification' --------";
 12.1016 -"-------- me 'poly. simpl.' Schalk I p.63 No.267b -------";
 12.1017 -"-------- interSteps for Schalk 299a --------------------";
 12.1018 -"-------- norm_Poly NOT COMPLETE ------------------------";
 12.1019 -"-------- ord_make_polynomial ---------------------------";
 12.1020 -"--------------------------------------------------------";
 12.1021 -"--------------------------------------------------------";
 12.1022 -"--------------------------------------------------------";
 12.1023 -
 12.1024 -
 12.1025 -\<close> ML \<open>
 12.1026 -"----------- fun is_polyexp --------------------------------------------------------------------";
 12.1027 -"----------- fun is_polyexp --------------------------------------------------------------------";
 12.1028 -"----------- fun is_polyexp --------------------------------------------------------------------";
 12.1029 -val thy = @{theory Partial_Fractions};
 12.1030 -val ctxt = Proof_Context.init_global thy;
 12.1031 -
 12.1032 -val t = (the o (parseNEW  ctxt)) "x / x";
 12.1033 -if is_polyexp t then error "NOT is_polyexp (x / x)" else ();
 12.1034 -
 12.1035 -val t = (the o (parseNEW  ctxt)) "-1 * A * 3";
 12.1036 -if is_polyexp t then () else error "is_polyexp (-1 * A * 3)";
 12.1037 -
 12.1038 -val t = (the o (parseNEW  ctxt)) "-1 * AA * 3";
 12.1039 -if is_polyexp t then () else error "is_polyexp (-1 * AA * 3)";
 12.1040 -
 12.1041 -\<close> ML \<open>
 12.1042 -"----------- fun has_degree_in -----------------------------------------------------------------";
 12.1043 -"----------- fun has_degree_in -----------------------------------------------------------------";
 12.1044 -"----------- fun has_degree_in -----------------------------------------------------------------";
 12.1045 -val v = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1046 -val t = (Thm.term_of o the o (TermC.parse thy)) "1";
 12.1047 -if has_degree_in t v = 0 then () else error "has_degree_in (1) x";
 12.1048 -
 12.1049 -val v = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1050 -val t = (Thm.term_of o the o (TermC.parse thy)) "1";
 12.1051 -if has_degree_in t v = 0 then () else error "has_degree_in (1) AA";
 12.1052 -
 12.1053 -(*----------*)
 12.1054 -val v = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1055 -val t = (Thm.term_of o the o (TermC.parse thy)) "a*b+c";
 12.1056 -if has_degree_in t v = 0 then () else error "has_degree_in (a*b+c) x";
 12.1057 -
 12.1058 -val v = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1059 -val t = (Thm.term_of o the o (TermC.parse thy)) "a*b+c";
 12.1060 -if has_degree_in t v = 0 then () else error "has_degree_in (a*b+c) AA";
 12.1061 -
 12.1062 -(*----------*)
 12.1063 -val v = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1064 -val t = (Thm.term_of o the o (TermC.parse thy)) "a*x+c";
 12.1065 -if has_degree_in t v = ~1 then () else error "has_degree_in (a*x+c) x";
 12.1066 -
 12.1067 -val v = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1068 -val t = (Thm.term_of o the o (TermC.parse thy)) "a*AA+c";
 12.1069 -if has_degree_in t v = ~1 then () else error "has_degree_in (a*AA+c) AA";
 12.1070 -
 12.1071 -(*----------*)
 12.1072 -val v = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1073 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)*x \<up> 7";
 12.1074 -if has_degree_in t v = 7 then () else error "has_degree_in ((a*b+c)*x \<up> 7) x";
 12.1075 -
 12.1076 -val v = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1077 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)*AA \<up> 7";
 12.1078 -if has_degree_in t v = 7 then () else error "has_degree_in ((a*b+c)*AA \<up> 7) AA";
 12.1079 -
 12.1080 -(*----------*)
 12.1081 -val v = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1082 -val t = (Thm.term_of o the o (TermC.parse thy)) "x \<up> 7";
 12.1083 -if has_degree_in t v = 7 then () else error "has_degree_in (x \<up> 7) x";
 12.1084 -
 12.1085 -val v = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1086 -val t = (Thm.term_of o the o (TermC.parse thy)) "AA \<up> 7";
 12.1087 -if has_degree_in t v = 7 then () else error "has_degree_in (AA \<up> 7) AA";
 12.1088 -
 12.1089 -(*----------*)
 12.1090 -val v = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1091 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)*x";
 12.1092 -if has_degree_in t v = 1 then () else error "has_degree_in ((a*b+c)*x) x";
 12.1093 -
 12.1094 -val v = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1095 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)*AA";
 12.1096 -if has_degree_in t v = 1 then () else error "has_degree_in ((a*b+c)*AA) AA";
 12.1097 -
 12.1098 -(*----------*)
 12.1099 -val v = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1100 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+x)*x";
 12.1101 -if has_degree_in t v = ~1 then () else error "has_degree_in (a*b+x)*x() x";
 12.1102 -
 12.1103 -val v = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1104 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+AA)*AA";
 12.1105 -if has_degree_in t v = ~1 then () else error "has_degree_in ((a*b+AA)*AA) AA";
 12.1106 -
 12.1107 -(*----------*)
 12.1108 -val v = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1109 -val t = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1110 -if has_degree_in t v = 1 then () else error "has_degree_in (x) x";
 12.1111 -
 12.1112 -val v = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1113 -val t = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1114 -if has_degree_in t v = 1 then () else error "has_degree_in (AA) AA";
 12.1115 -
 12.1116 -(*----------*)
 12.1117 -val v = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1118 -val t = (Thm.term_of o the o (TermC.parse thy)) "ab - (a*b)*x";
 12.1119 -if has_degree_in t v = 1 then () else error "has_degree_in (ab - (a*b)*x) x";
 12.1120 -
 12.1121 -val v = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1122 -val t = (Thm.term_of o the o (TermC.parse thy)) "ab - (a*b)*AA";
 12.1123 -if has_degree_in t v = 1 then () else error "has_degree_in (ab - (a*b)*AA) AA";
 12.1124 -
 12.1125 -\<close> ML \<open>
 12.1126 -"----------- fun mono_deg_in -------------------------------------------------------------------";
 12.1127 -"----------- fun mono_deg_in -------------------------------------------------------------------";
 12.1128 -"----------- fun mono_deg_in -------------------------------------------------------------------";
 12.1129 -val v = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1130 -
 12.1131 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)*x \<up> 7";
 12.1132 -if mono_deg_in t v = SOME 7 then () else error "mono_deg_in ((a*b+c)*x \<up> 7) x changed";
 12.1133 -
 12.1134 -val t = (Thm.term_of o the o (TermC.parse thy)) "x \<up> 7";
 12.1135 -if mono_deg_in t v = SOME 7 then () else error "mono_deg_in (x \<up> 7) x changed";
 12.1136 -
 12.1137 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)*x";
 12.1138 -if mono_deg_in t v = SOME 1 then () else error "mono_deg_in ((a*b+c)*x) x changed";
 12.1139 -
 12.1140 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+x)*x";
 12.1141 -if mono_deg_in t v = NONE then () else error "mono_deg_in ((a*b+x)*x) x changed";
 12.1142 -
 12.1143 -val t = (Thm.term_of o the o (TermC.parse thy)) "x";
 12.1144 -if mono_deg_in t v = SOME 1 then () else error "mono_deg_in (x) x changed";
 12.1145 -
 12.1146 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)";
 12.1147 -if mono_deg_in t v = SOME 0 then () else error "mono_deg_in ((a*b+c)) x changed";
 12.1148 -
 12.1149 -val t = (Thm.term_of o the o (TermC.parse thy)) "ab - (a*b)*x";
 12.1150 -if mono_deg_in t v = NONE then () else error "mono_deg_in (ab - (a*b)*x) x changed";
 12.1151 -
 12.1152 -(*. . . . . . . . . . . . the same with Const ("Partial_Functions.AA", _) . . . . . . . . . . . *)
 12.1153 -val thy = @{theory Partial_Fractions}
 12.1154 -val v = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1155 -
 12.1156 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)*AA \<up> 7";
 12.1157 -if mono_deg_in t v = SOME 7 then () else error "mono_deg_in ((a*b+c)*AA \<up> 7) AA changed";
 12.1158 -
 12.1159 -val t = (Thm.term_of o the o (TermC.parse thy)) "AA \<up> 7";
 12.1160 -if mono_deg_in t v = SOME 7 then () else error "mono_deg_in (AA \<up> 7) AA changed";
 12.1161 -
 12.1162 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)*AA";
 12.1163 -if mono_deg_in t v = SOME 1 then () else error "mono_deg_in ((a*b+c)*AA) AA changed";
 12.1164 -
 12.1165 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+AA)*AA";
 12.1166 -if mono_deg_in t v = NONE then () else error "mono_deg_in ((a*b+AA)*AA) AA changed";
 12.1167 -
 12.1168 -val t = (Thm.term_of o the o (TermC.parse thy)) "AA";
 12.1169 -if mono_deg_in t v = SOME 1 then () else error "mono_deg_in (AA) AA changed";
 12.1170 -
 12.1171 -val t = (Thm.term_of o the o (TermC.parse thy)) "(a*b+c)";
 12.1172 -if mono_deg_in t v = SOME 0 then () else error "mono_deg_in ((a*b+c)) AA changed";
 12.1173 -
 12.1174 -val t = (Thm.term_of o the o (TermC.parse thy)) "ab - (a*b)*AA";
 12.1175 -if mono_deg_in t v = NONE then () else error "mono_deg_in (ab - (a*b)*AA) AA changed";
 12.1176 -
 12.1177 -\<close> ML \<open>
 12.1178 -"----------- rewrite_set_ has_degree_in Const ('Partial_Fractions', _) -------------------------";
 12.1179 -"----------- rewrite_set_ has_degree_in Const ('Partial_Fractions', _) -------------------------";
 12.1180 -"----------- rewrite_set_ has_degree_in Const ('Partial_Fractions', _) -------------------------";
 12.1181 -val thy = @{theory Partial_Fractions}
 12.1182 -val expr = (Thm.term_of o the o (TermC.parse thy)) "((-8 - 2*x + x \<up> 2) has_degree_in x) = 2";
 12.1183 -val SOME (Const ("HOL.True", _), []) = rewrite_set_ thy false PolyEq_prls expr;
 12.1184 -
 12.1185 -val expr = (Thm.term_of o the o (TermC.parse thy)) "((-8 - 2*AA + AA \<up> 2) has_degree_in AA) = 2";
 12.1186 -val SOME (Const ("HOL.True", _), []) = rewrite_set_ thy false PolyEq_prls expr;
 12.1187 -
 12.1188 -\<close> ML \<open>
 12.1189 -"----------- eval_ for is_expanded_in, is_poly_in, has_degree_in -------------------------------";
 12.1190 -"----------- eval_ for is_expanded_in, is_poly_in, has_degree_in -------------------------------";
 12.1191 -"----------- eval_ for is_expanded_in, is_poly_in, has_degree_in -------------------------------";
 12.1192 -val t = (Thm.term_of o the o (TermC.parse thy)) "(-8 - 2*x + x \<up> 2) is_expanded_in x";
 12.1193 -val SOME (id, t') = eval_is_expanded_in 0 0 t 0;
 12.1194 -if UnparseC.term t' = "- 8 - 2 * x + x \<up> 2 is_expanded_in x = True"
 12.1195 -         andalso id = "- 8 - 2 * x + x \<up> 2 is_expanded_in x = True"
 12.1196 -then () else error "eval_is_expanded_in x ..changed";
 12.1197 -
 12.1198 -\<close> ML \<open>
 12.1199 -val thy = @{theory Partial_Fractions}
 12.1200 -val t = (Thm.term_of o the o (TermC.parse thy)) "(-8 - 2*AA + AA \<up> 2) is_expanded_in AA";
 12.1201 -val SOME (id, t') = eval_is_expanded_in 0 0 t 0;
 12.1202 -if  UnparseC.term t' = "- 8 - 2 * AA + AA \<up> 2 is_expanded_in AA = True"
 12.1203 -          andalso id = "- 8 - 2 * AA + AA \<up> 2 is_expanded_in AA = True"
 12.1204 -then () else error "eval_is_expanded_in AA ..changed";
 12.1205 -
 12.1206 -
 12.1207 -\<close> ML \<open>
 12.1208 -val t = (Thm.term_of o the o (TermC.parse thy)) "(8 + 2*x + x \<up> 2) is_poly_in x";
 12.1209 -val SOME (id, t') = eval_is_poly_in 0 0 t 0;
 12.1210 -if  UnparseC.term t' = "8 + 2 * x + x \<up> 2 is_poly_in x = True"
 12.1211 -          andalso id = "8 + 2 * x + x \<up> 2 is_poly_in x = True"
 12.1212 -then () else error "is_poly_in x ..changed";
 12.1213 -
 12.1214 -val t = (Thm.term_of o the o (TermC.parse thy)) "(8 + 2*AA + AA \<up> 2) is_poly_in AA";
 12.1215 -val SOME (id, t') = eval_is_poly_in 0 0 t 0;
 12.1216 -if  UnparseC.term t' = "8 + 2 * AA + AA \<up> 2 is_poly_in AA = True"
 12.1217 -          andalso id = "8 + 2 * AA + AA \<up> 2 is_poly_in AA = True"
 12.1218 -then () else error "is_poly_in AA ..changed";
 12.1219 -
 12.1220 -
 12.1221 -val thy = @{theory Partial_Fractions}
 12.1222 -val expr = (Thm.term_of o the o (TermC.parse thy)) "((-8 - 2*x + x \<up> 2) has_degree_in x) = 2";
 12.1223 -val SOME (Const ("HOL.True", _), []) = rewrite_set_ thy false PolyEq_prls expr;
 12.1224 -
 12.1225 -val expr = (Thm.term_of o the o (TermC.parse thy)) "((-8 - 2*AA + AA \<up> 2) has_degree_in AA) = 2";
 12.1226 -val SOME (Const ("HOL.True", _), []) = rewrite_set_ thy false PolyEq_prls expr;
 12.1227 -
 12.1228 -\<close> ML \<open>
 12.1229 -"-------- investigate (new 2002) uniary minus -----------";
 12.1230 -"-------- investigate (new 2002) uniary minus -----------";
 12.1231 -"-------- investigate (new 2002) uniary minus -----------";
 12.1232 -(*---------------------------------------------- vvvvvvvvvvvvvv -----------------------*)
 12.1233 -val t = Thm.prop_of @{thm real_diff_0}; (*"0 - ?x = - ?x"*)
 12.1234 -TermC.atomty t;
 12.1235 -(*
 12.1236 -*** Const (HOL.Trueprop, bool => prop)
 12.1237 -*** . Const (HOL.eq, real => real => bool)
 12.1238 -*** . . Const (Groups.minus_class.minus, real => real => real)
 12.1239 -*** . . . Const (Groups.zero_class.zero, real)
 12.1240 -*** . . . Var ((x, 0), real)
 12.1241 -*** . . Const (Groups.uminus_class.uminus, real => real)
 12.1242 -*** . . . Var ((x, 0), real)
 12.1243 -*)
 12.1244 -case t of
 12.1245 -  Const ("HOL.Trueprop", _) $
 12.1246 -    (Const ("HOL.eq", _) $ 
 12.1247 -      (Const ("Groups.minus_class.minus", _) $ Const ("Groups.zero_class.zero", _) $ 
 12.1248 -        Var (("x", 0), _)) $
 12.1249 -             (Const ("Groups.uminus_class.uminus", _) $ Var (("x", 0), _))) => ()
 12.1250 -| _ => error "internal representation of \"0 - ?x = - ?x\" changed";
 12.1251 -
 12.1252 -\<close> ML \<open>
 12.1253 -(*----------------------------------- vvvvv -------------------------------------------*)
 12.1254 -val t = (Thm.term_of o the o (TermC.parse thy)) "- 1";
 12.1255 -TermC.atomty t;
 12.1256 -(*
 12.1257 -*** 
 12.1258 -*** Free (-1, real)
 12.1259 -*** 
 12.1260 -*)
 12.1261 -\<close> ML \<open>
 12.1262 -case t of
 12.1263 - Const ("Groups.uminus_class.uminus", _) $ Const ("Groups.one_class.one", _) => ()
 12.1264 -| _ => error "internal representation of \"- 1\" changed";
 12.1265 -
 12.1266 -\<close> ML \<open>
 12.1267 -"======= these external values all have the same internal representation";
 12.1268 -(* "1-x" causes syntyx error --- binary minus detected by blank inbetween !!!*)
 12.1269 -(*----------------------------------- vvvvv -------------------------------------------*)
 12.1270 -val t = (Thm.term_of o the o (TermC.parse thy)) "-x";
 12.1271 -TermC.atomty t;
 12.1272 -(**** -------------
 12.1273 -*** Free ( -x, real)*)
 12.1274 -case t of
 12.1275 -  Const ("Groups.uminus_class.uminus", _) $ Free ("x", _) => ()
 12.1276 -| _ => error "internal representation of \"-x\" changed";
 12.1277 -\<close> ML \<open>
 12.1278 -(*----------------------------------- vvvvv -------------------------------------------*)
 12.1279 -val t = (Thm.term_of o the o (TermC.parse thy)) "- x";
 12.1280 -TermC.atomty t;
 12.1281 -(**** -------------
 12.1282 -*** Free ( -x, real) !!!!!!!!!!!!!!!!!!!!!!!! is the same !!!*)
 12.1283 -case t of
 12.1284 -  Const ("Groups.uminus_class.uminus", _) $ Free ("x", _) => ()
 12.1285 -| _ => error "internal representation of \"- x\" changed";
 12.1286 -\<close> ML \<open>
 12.1287 -(*----------------------------------- vvvvvv ------------------------------------------*)
 12.1288 -val t = (Thm.term_of o the o (TermC.parse thy)) "-(x)";
 12.1289 -TermC.atomty t;
 12.1290 -(**** -------------
 12.1291 -*** Free ( -x, real)*)
 12.1292 -case t of
 12.1293 -  Const ("Groups.uminus_class.uminus", _) $ Free ("x", _) => ()
 12.1294 -| _ => error "internal representation of \"-(x)\" changed";
 12.1295 -
 12.1296 -\<close> ML \<open> (*\\---------------------- original test code --------------------------------------//*)
 12.1297 -\<close> ML \<open>
 12.1298 -\<close> ML \<open> (* \<longrightarrow> Poly.thy*)
 12.1299 -\<close> ML \<open> (* \<longrightarrow> Poly.thy*)
 12.1300 -\<close> ML \<open>
 12.1301 -if "---" < "123" andalso "123" < "a" andalso "a" < "cba" then ()
 12.1302 -else error "lexicographic order CHANGED";
 12.1303 -\<close> ML \<open>
 12.1304 -@{term 123}
 12.1305 -\<close> ML \<open>
 12.1306 -\<close> ML \<open>
 12.1307 -\<close> ML \<open> (* \<longrightarrow> Poly.thy*)
 12.1308 -\<close> ML \<open>
 12.1309 -"----------- fun sort_varList ------------------------------------------------------------------";
 12.1310 -"----------- fun sort_varList ------------------------------------------------------------------";
 12.1311 -"----------- fun sort_varList ------------------------------------------------------------------";
 12.1312 -\<close> ML \<open>
 12.1313 -       sort_varList [@{term "a"}, @{term "3"}];
 12.1314 -\<close> ML \<open>
 12.1315 -"~~~~~ val sort_varList , args:"; val (t) = [@{term "a"}, @{term "3"}];
 12.1316 -\<close> ML \<open>
 12.1317 -sort_varList: term list -> term list;
 12.1318 -\<close> ML \<open>
 12.1319 -\<close> ML \<open>
 12.1320 -\<close> ML \<open>
 12.1321 -\<close> ML \<open>
 12.1322 -\<close> ML \<open>
 12.1323 -\<close> ML \<open> (* \<longrightarrow> Poly.thy*)
 12.1324 -\<close> ML \<open>
 12.1325 -"----------- fun sort_variables ----------------------------------------------------------------";
 12.1326 -"----------- fun sort_variables ----------------------------------------------------------------";
 12.1327 -"----------- fun sort_variables ----------------------------------------------------------------";
 12.1328 -\<close> ML \<open>
 12.1329 -sort_variables: term -> term;
 12.1330 -\<close> ML \<open>
 12.1331 - val t' =  sort_variables @{term "3 * b + a * 2"};
 12.1332 -\<close> ML \<open>
 12.1333 -UnparseC.term t' = "(3::'a) * b + (2::'a) * a"
 12.1334 -\<close> ML \<open>
 12.1335 -"~~~~~ fun sort_variables , args:"; val (t) = (@{term "3 * b + a * 2"});
 12.1336 -\<close> ML \<open>
 12.1337 -  	val ll =  map monom2list (poly2list t);
 12.1338 -\<close> ML \<open>
 12.1339 -(*+*)UnparseC.terms (poly2list t) = "[\"(3::'a) * b\", \"a * (2::'a)\"]";
 12.1340 -(*+*)val [
 12.1341 -(*+*)     [Const ("Num.numeral_class.numeral", _) $ (Const ("Num.num.Bit1", _) $ Const ("Num.num.One", _)) ,(*!*) Free ("b", _)],
 12.1342 -(*+*)     [Free ("a", _) ,(*!*) Const ("Num.numeral_class.numeral", _) $ (Const ("Num.num.Bit0", _) $ Const ("Num.num.One", _))]
 12.1343 -(*+*)    ] = map monom2list (poly2list t);
 12.1344 -\<close> ML \<open>
 12.1345 -  	val lls = map sort_varList ll;
 12.1346 -\<close> ML \<open>
 12.1347 -(*+*)case map sort_varList ll of
 12.1348 -(*+*)  [ [Const ("Num.numeral_class.numeral", _) $ _, Free ("b", _)],
 12.1349 -(*+*)    [Const ("Num.numeral_class.numeral", _) $ _, Free ("a", _)]
 12.1350 -(*+*)  ] => ()
 12.1351 -(*+*)| _ => error "map sort_varList CHANGED";
 12.1352 -\<close> ML \<open>
 12.1353 -  	val T = type_of t;
 12.1354 -  	val ls = map (create_monom T) lls;
 12.1355 -\<close> ML \<open>
 12.1356 -(*+*)val [Const ("Groups.times_class.times", _) $ _ $ Free ("b", _),
 12.1357 -(*+*)     Const ("Groups.times_class.times", _) $ _ $ Free ("a", _)] = map (create_monom T) lls;
 12.1358 -\<close> ML \<open>
 12.1359 -(*+*)case map (create_monom T) lls of
 12.1360 -(*+*)  [Const ("Groups.times_class.times", _) $ (Const ("Num.numeral_class.numeral", _) $ _) $ Free ("b", _),
 12.1361 -(*+*)   Const ("Groups.times_class.times", _) $ (Const ("Num.numeral_class.numeral", _) $ _) $ Free ("a", _)
 12.1362 -(*+*)  ] => ()
 12.1363 -(*+*)| _ => error "map (create_monom T) CHANGED";
 12.1364 -\<close> ML \<open>
 12.1365 -  val xxx = (*in*) create_polynom T ls (*end*);
 12.1366 -\<close> ML \<open>
 12.1367 -(*+*)if UnparseC.term xxx = "(3::'a) * b + (2::'a) * a" then ()
 12.1368 -(*+*)else error "create_polynom CHANGED";
 12.1369 -(* done by rewriting>              2 * a +       3 * b ? *)
 12.1370 -\<close> ML \<open>
 12.1371 -\<close> ML \<open>
 12.1372 -\<close> ML \<open>
 12.1373 -\<close> ML \<open>
 12.1374 -\<close> ML \<open>
 12.1375 -\<close> ML \<open>
 12.1376 -\<close> ML \<open>
 12.1377 -\<close> ML \<open>
 12.1378 -\<close> ML \<open>
 12.1379 -\<close> ML \<open>
 12.1380 -\<close> ML \<open>
 12.1381 -
 12.1382 -\<close> ML \<open> (*//---------------------- original test code --------------------------------------\\*)
 12.1383 -"-------- check make_polynomial with simple terms -------";
 12.1384 -"-------- check make_polynomial with simple terms -------";
 12.1385 -"-------- check make_polynomial with simple terms -------";
 12.1386 -"----- check 1 ---";
 12.1387 -val t = TermC.str2term "2*3*a";
 12.1388 -val SOME (t, _) = rewrite_set_ thy false make_polynomial t;
 12.1389 -\<close> ML \<open>
 12.1390 -UnparseC.term t = "a * 6" (*<<<<<<<<<<<<<<<<<<<----------------------------TODOO correct*)
 12.1391 -\<close> ML \<open>
 12.1392 -if UnparseC.term t = "6 * a" then () else error "check make_polynomial 1";
 12.1393 -
 12.1394 -\<close> ML \<open>
 12.1395 -"----- check 2 ---";
 12.1396 -val t = TermC.str2term "2*a + 3*a";
 12.1397 -val SOME (t, _) = rewrite_set_ thy false make_polynomial t;
 12.1398 -if UnparseC.term t = "5 * a" then () else error "check make_polynomial 2";
 12.1399 -
 12.1400 -\<close> ML \<open>
 12.1401 -"----- check 3 ---";
 12.1402 -val t = TermC.str2term "2*a + 3*a + 3*a";
 12.1403 -val SOME (t, _) = rewrite_set_ thy false make_polynomial t;
 12.1404 -if UnparseC.term t = "8 * a" then () else error "check make_polynomial 3";
 12.1405 -
 12.1406 -\<close> ML \<open>
 12.1407 -Rewrite.trace_on := true;
 12.1408 -\<close> ML \<open> (*loops THIS IS THE SIMPLEST EXAMPLE*)
 12.1409 -"----- check 4 ---";
 12.1410 -val t = TermC.str2term "3*a - 2*a";
 12.1411 -\<close> text \<open> (*loops THIS IS THE SIMPLEST EXAMPLE*)
 12.1412 -val SOME (t, _) = rewrite_set_ thy false make_polynomial t;
 12.1413 -\<close> text \<open>
 12.1414 -if UnparseC.term t = "a" then () else error "check make_polynomial 4";
 12.1415 -\<close> ML \<open>
 12.1416 -Rewrite.trace_on := false;
 12.1417 -\<close> ML \<open>
 12.1418 -
 12.1419 -\<close> text \<open> (*loops*)
 12.1420 -"----- check 5 ---";
 12.1421 -val t = TermC.str2term "4*(3*a - 2*a)";
 12.1422 -val SOME (t, _) = rewrite_set_ thy false make_polynomial t;
 12.1423 -if UnparseC.term t = "4 * a" then () else error "check make_polynomial 5";
 12.1424 -
 12.1425 -\<close> text \<open> (*loops*)
 12.1426 -"----- check 6 ---";
 12.1427 -val t = TermC.str2term "4*(3*a \<up> 2 - 2*a \<up> 2)";
 12.1428 -val SOME (t, _) = rewrite_set_ thy false make_polynomial t;
 12.1429 -if UnparseC.term t = "4 * a \<up> 2" then () else error "check make_polynomial 6";
 12.1430 -
 12.1431 -\<close> ML \<open>
 12.1432 -\<close> ML \<open>
 12.1433 -"-------- fun is_multUnordered --------------------------";
 12.1434 -"-------- fun is_multUnordered --------------------------";
 12.1435 -"-------- fun is_multUnordered --------------------------";
 12.1436 -\<close> ML \<open>
 12.1437 -val thy = @{theory "Isac_Knowledge"};
 12.1438 -"===== works for a simple example, see rewrite.sml -- fun app_rev ===";
 12.1439 -val t = TermC.str2term "x \<up> 2 * x";
 12.1440 -val SOME (t', _) = rewrite_set_ thy true order_mult_ t;
 12.1441 -\<close> ML \<open>
 12.1442 -if UnparseC.term t' = "x * x \<up> 2" then ()
 12.1443 -else error "poly.sml Poly.is_multUnordered doesn't work";
 12.1444 -
 12.1445 -\<close> ML \<open>
 12.1446 -(* 100928 Rewrite.trace_on shows the first occurring difference in 267b:
 12.1447 -###  rls: order_mult_ on: 5 * x \<up> 2 * (2 * x \<up> 7) + 5 * x \<up> 2 * 3 + (6 * x \<up> 7 + 9) + (-1 * (3 * x \<up> 5 * (6 * x \<up> 4)) + -1 * (3 * x \<up> 5 * -1) +
 12.1448 - (-48 * x \<up> 4 + 8))
 12.1449 -######  rls: Rule_Set.empty-is_multUnordered on: p is_multUnordered
 12.1450 -#######  try calc: Poly.is_multUnordered'
 12.1451 -=======  calc. to: False  !!!!!!!!!!!!! INSTEAD OF TRUE in 2002 !!!!!!!!!!!!!
 12.1452 -*)
 12.1453 -val t = TermC.str2term "5 * x \<up> 2 * (2 * x \<up> 7) + 5 * x \<up> 2 * 3 + (6 * x \<up> 7 + 9) + (-1 * (3 * x \<up> 5 * (6 * x \<up> 4)) + -1 * (3 * x \<up> 5 * -1) +  (-48 * x \<up> 4 + 8))";
 12.1454 -
 12.1455 -\<close> ML \<open>
 12.1456 -"----- is_multUnordered ---";
 12.1457 -val tsort = sort_variables t;
 12.1458 -UnparseC.term tsort = "2 * (5 * (x \<up> 2 * x \<up> 7)) + 3 * (5 * x \<up> 2) + 6 * x \<up> 7 + 9 +\n-1 * (3 * (6 * (x \<up> 4 * x \<up> 5))) +\n-1 * (-1 * (3 * x \<up> 5)) +\n-48 * x \<up> 4 +\n8";
 12.1459 -is_polyexp t;
 12.1460 -tsort = t;
 12.1461 -is_polyexp t andalso not (t = sort_variables t);
 12.1462 -\<close> text \<open> (*false*)
 12.1463 -if is_multUnordered t then () else error "poly.sml diff. is_multUnordered 1";
 12.1464 -
 12.1465 -\<close> ML \<open>
 12.1466 -"----- eval_is_multUnordered ---";
 12.1467 -val tm = TermC.str2term "(5 * x \<up> 2 * (2 * x \<up> 7) + 5 * x \<up> 2 * 3 + (6 * x \<up> 7 + 9) + (-1 * (3 * x \<up> 5 * (6 * x \<up> 4)) + -1 * (3 * x \<up> 5 * -1) +  (-48 * x \<up> 4 + 8))) is_multUnordered";
 12.1468 -case eval_is_multUnordered "testid" "" tm thy of
 12.1469 -    SOME (_, Const ("HOL.Trueprop", _) $ 
 12.1470 -                   (Const ("HOL.eq", _) $
 12.1471 -                          (Const ("Poly.is_multUnordered", _) $ _) $ 
 12.1472 -                          Const ("HOL.True", _))) => ()
 12.1473 -  | _ => error "poly.sml diff. eval_is_multUnordered";
 12.1474 -
 12.1475 -\<close> text \<open> (*NONE*)
 12.1476 -"----- rewrite_set_ STILL DIDN'T WORK";
 12.1477 -val SOME (t, _) = rewrite_set_ thy true order_mult_ t;
 12.1478 -UnparseC.term t;
 12.1479 -
 12.1480 -\<close> ML \<open>
 12.1481 -"-------- examples from textbook Schalk I ---------------";
 12.1482 -"-------- examples from textbook Schalk I ---------------";
 12.1483 -"-------- examples from textbook Schalk I ---------------";
 12.1484 -"-----SPB Schalk I p.63 No.267b ---";
 12.1485 -(*associate poly* )
 12.1486 -val t = TermC.str2term "(5*x \<up> 2 + 3) * (2*x \<up> 7 + 3) - (3*x \<up> 5 + 8) * (6*x \<up> 4 - 1)";
 12.1487 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1488 -if (UnparseC.term t) = "17 + 15 * x \<up> 2 + -48 * x \<up> 4 + 3 * x \<up> 5 + 6 * x \<up> 7 + -8 * x \<up> 9"
 12.1489 -then () else error "poly.sml: diff.behav. in make_polynomial 1";
 12.1490 -
 12.1491 -"-----SPB Schalk I p.63 No.275b ---";
 12.1492 -val t = TermC.str2term "(3*x \<up> 2 - 2*x*y + y \<up> 2) * (x \<up> 2 - 2*y \<up> 2)";
 12.1493 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
 12.1494 -if (UnparseC.term t) = ("3 * x \<up> 4 + -2 * x \<up> 3 * y + -5 * x \<up> 2 * y \<up> 2 + " ^
 12.1495 -  "4 * x * y \<up> 3 +\n-2 * y \<up> 4")
 12.1496 -then () else error "poly.sml: diff.behav. in make_polynomial 2";
 12.1497 -
 12.1498 -"-----SPB Schalk I p.63 No.279b ---";
 12.1499 -val t = TermC.str2term "(x-a)*(x-b)*(x-c)*(x-d)";
 12.1500 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
 12.1501 -if (UnparseC.term t) = 
 12.1502 -  ("a * b * c * d + -1 * a * b * c * x + -1 * a * b * d * x + a * b * x \<up> 2 +\n" ^
 12.1503 -  "-1 * a * c * d * x +\na * c * x \<up> 2 +\na * d * x \<up> 2 +\n-1 * a * x \<up> 3 +\n" ^
 12.1504 -  "-1 * b * c * d * x +\nb * c * x \<up> 2 +\nb * d * x \<up> 2 +\n-1 * b * x \<up> 3 +\n" ^
 12.1505 -  "c * d * x \<up> 2 +\n-1 * c * x \<up> 3 +\n-1 * d * x \<up> 3 +\nx \<up> 4")
 12.1506 -then () else error "poly.sml: diff.behav. in make_polynomial 3";
 12.1507 -( *associate poly*)
 12.1508 -
 12.1509 -\<close> text \<open> (*loops*)
 12.1510 -"-----SPB Schalk I p.63 No.291 ---";
 12.1511 -val t = TermC.str2term "(5+96*x \<up> 3+8*x*(-4+(7- 3*x)*4*x))*(5*(2- 3*x)- (-15*x*(-8*x- 5)))";
 12.1512 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
 12.1513 -if (UnparseC.term t) = "50 + -770 * x + 4520 * x \<up> 2 + -16320 * x \<up> 3 + -26880 * x \<up> 4"
 12.1514 -then () else error "poly.sml: diff.behav. in make_polynomial 4";
 12.1515 -
 12.1516 -\<close> ML \<open>
 12.1517 -(*associate poly* )
 12.1518 -"-----SPB Schalk I p.64 No.295c ---";
 12.1519 -val t = TermC.str2term "(13*a \<up> 4*b \<up> 9*c - 12*a \<up> 3*b \<up> 6*c \<up> 9) \<up> 2";
 12.1520 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
 12.1521 -if (UnparseC.term t) = ("169 * a \<up> 8 * b \<up> 18 * c \<up> 2 + -312 * a \<up> 7 * b \<up> 15 * c \<up> 10" ^
 12.1522 -  " +\n144 * a \<up> 6 * b \<up> 12 * c \<up> 18")
 12.1523 -then ()else error "poly.sml: diff.behav. in make_polynomial 5";
 12.1524 -( *associate poly*)
 12.1525 -
 12.1526 -\<close> text \<open> (*loops*)
 12.1527 -"-----SPB Schalk I p.64 No.299a ---";
 12.1528 -val t = TermC.str2term "(x - y)*(x + y)";
 12.1529 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
 12.1530 -if (UnparseC.term t) = "x \<up> 2 + -1 * y \<up> 2"
 12.1531 -then () else error "poly.sml: diff.behav. in make_polynomial 6";
 12.1532 -
 12.1533 -\<close> text \<open> (*loops*)
 12.1534 -"-----SPB Schalk I p.64 No.300c ---";
 12.1535 -val t = TermC.str2term "(3*x \<up> 2*y - 1)*(3*x \<up> 2*y + 1)";
 12.1536 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
 12.1537 -if (UnparseC.term t) = "-1 + 9 * x \<up> 4 * y \<up> 2"
 12.1538 -then () else error "poly.sml: diff.behav. in make_polynomial 7";
 12.1539 -
 12.1540 -\<close> text \<open> (*loops*)
 12.1541 -"-----SPB Schalk I p.64 No.302 ---";
 12.1542 -val t = TermC.str2term
 12.1543 -  "(13*x \<up> 2 + 5)*(13*x \<up> 2 - 5) - (5*x \<up> 2 + 3)*(5*x \<up> 2 - 3) - (12*x \<up> 2 + 4)*(12*x \<up> 2 - 4)";
 12.1544 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t;
 12.1545 -if UnparseC.term t = "0"
 12.1546 -then () else error "poly.sml: diff.behav. in make_polynomial 8";
 12.1547 -(* RL?MG?: Bei Berechnung sollte 3 mal real_plus_minus_binom1_p aus expand_poly verwendet werden *)
 12.1548 -
 12.1549 -\<close> text \<open> (*loops*)
 12.1550 -"-----SPB Schalk I p.64 No.306a ---";
 12.1551 -val t = TermC.str2term "((x \<up> 2 + 1)*(x \<up> 2 - 1)) \<up> 2";
 12.1552 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1553 -if (UnparseC.term t) = "1 + 2 * x \<up> 4 + 2 * -2 * x \<up> 4 + x \<up> 8" then ()
 12.1554 -else error "poly.sml: diff.behav. in 2 * x \<up> 4 + 2 * -2 * x \<up> 4 = -2 * x \<up> 4";
 12.1555 -
 12.1556 -\<close> text \<open> (*loops*)
 12.1557 -(*WN071729 when reducing "rls reduce_012_" for Schaerding,
 12.1558 -the above resulted in the term below ... but reduces from then correctly*)
 12.1559 -val t = TermC.str2term "1 + 2 * x \<up> 4 + 2 * -2 * x \<up> 4 + x \<up> 8";
 12.1560 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1561 -if (UnparseC.term t) = "1 + -2 * x \<up> 4 + x \<up> 8"
 12.1562 -then () else error "poly.sml: diff.behav. in make_polynomial 9b";
 12.1563 -
 12.1564 -\<close> text \<open> (*loops*)
 12.1565 -"-----SPB Schalk I p.64 No.296a ---";
 12.1566 -val t = TermC.str2term "(x - a) \<up> 3";
 12.1567 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1568 -if (UnparseC.term t) = "-1 * a \<up> 3 + 3 * a \<up> 2 * x + -3 * a * x \<up> 2 + x \<up> 3"
 12.1569 -then () else error "poly.sml: diff.behav. in make_polynomial 10";
 12.1570 -
 12.1571 -\<close> text \<open> (*loops*)
 12.1572 -"-----SPB Schalk I p.64 No.296c ---";
 12.1573 -val t = TermC.str2term "(-3*x - 4*y) \<up> 3";
 12.1574 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1575 -if (UnparseC.term t) = "-27 * x \<up> 3 + -108 * x \<up> 2 * y + -144 * x * y \<up> 2 +\n-64 * y \<up> 3"
 12.1576 -then () else error "poly.sml: diff.behav. in make_polynomial 11";
 12.1577 -
 12.1578 -\<close> text \<open> (*loops*)
 12.1579 -"-----SPB Schalk I p.62 No.242c ---";
 12.1580 -val t = TermC.str2term "x \<up> (-4)*(x \<up> (-4)*y \<up> (-2)) \<up> (-1)*y \<up> (-2)";
 12.1581 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1582 -if (UnparseC.term t) = "1"
 12.1583 -then () else error "poly.sml: diff.behav. in make_polynomial 12";
 12.1584 -
 12.1585 -\<close> text \<open> (*loops*)
 12.1586 -"-----SPB Schalk I p.60 No.209a ---";
 12.1587 -val t = TermC.str2term "a \<up> (7-x) * a \<up> x";
 12.1588 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1589 -if UnparseC.term t = "a \<up> 7"
 12.1590 -then () else error "poly.sml: diff.behav. in make_polynomial 13";
 12.1591 -
 12.1592 -\<close> text \<open> (*loops*)
 12.1593 -"-----SPB Schalk I p.60 No.209d ---";
 12.1594 -val t = TermC.str2term "d \<up> x * d \<up> (x+1) * d \<up> (2 - 2*x)";
 12.1595 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1596 -if UnparseC.term t = "d \<up> 3"
 12.1597 -then () else error "poly.sml: diff.behav. in make_polynomial 14";
 12.1598 -
 12.1599 -\<close> text \<open> (*loops*)
 12.1600 -(*---------------------------------------------------------------------*)
 12.1601 -(*---------------- ?RL?Bsple bei denen es Probleme gibt----------------*)
 12.1602 -(*---------------------------------------------------------------------*)
 12.1603 -"-----Schalk I p.64 No.303 ---";
 12.1604 -val t = TermC.str2term "(a + 2*b)*(a \<up> 2 + 4*b \<up> 2)*(a - 2*b) - (a - 6*b)*(a \<up> 2 + 36*b \<up> 2)*(a + 6*b)";
 12.1605 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1606 -if UnparseC.term t = "1280 * b \<up> 4"
 12.1607 -then () else error "poly.sml: diff.behav. in make_polynomial 14b";
 12.1608 -(* Richtig - aber Binomische Formel wurde nicht verwendet! *)
 12.1609 -
 12.1610 -(*--------------------------------------------------------------------*)
 12.1611 -(*----------------------- Eigene Beispiele ---------------------------*)
 12.1612 -(*--------------------------------------------------------------------*)
 12.1613 -\<close> text \<open> (*loops*)
 12.1614 -"-----SPO ---";
 12.1615 -val t = TermC.str2term "a \<up> 2*a \<up> (-2)";
 12.1616 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1617 -if UnparseC.term t = "1" then ()
 12.1618 -else error "poly.sml: diff.behav. in make_polynomial 15";
 12.1619 -\<close> ML \<open>
 12.1620 -"-----SPO ---";
 12.1621 -val t = TermC.str2term "a + a + a";
 12.1622 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1623 -if UnparseC.term t = "3 * a" then ()
 12.1624 -else error "poly.sml: diff.behav. in make_polynomial 16";
 12.1625 -\<close> ML \<open>
 12.1626 -"-----SPO ---";
 12.1627 -val t = TermC.str2term "a + b + b + b";
 12.1628 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1629 -if UnparseC.term t = "a + 3 * b" then ()
 12.1630 -else error "poly.sml: diff.behav. in make_polynomial 17";
 12.1631 -\<close> text \<open> (*loops*)
 12.1632 -"-----SPO ---";
 12.1633 -val t = TermC.str2term "a \<up> 2*b*b \<up> (-1)";
 12.1634 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1635 -if UnparseC.term t = "a \<up> 2" then ()
 12.1636 -else error "poly.sml: diff.behav. in make_polynomial 18";
 12.1637 -\<close> text \<open> (*loops*)
 12.1638 -"-----SPO ---";
 12.1639 -val t = TermC.str2term "a \<up> 2*a \<up> (-2)";
 12.1640 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1641 -if (UnparseC.term t) = "1" then ()
 12.1642 -else error "poly.sml: diff.behav. in make_polynomial 19";
 12.1643 -\<close> text \<open> (*loops*)
 12.1644 -"-----SPO ---";
 12.1645 -val t = TermC.str2term "b + a - b";
 12.1646 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1647 -if (UnparseC.term t) = "a" then ()
 12.1648 -else error "poly.sml: diff.behav. in make_polynomial 20";
 12.1649 -\<close> ML \<open>
 12.1650 -"-----SPO ---";
 12.1651 -val t = TermC.str2term "b * a * a";
 12.1652 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1653 -if UnparseC.term t = "a \<up> 2 * b" then ()
 12.1654 -else error "poly.sml: diff.behav. in make_polynomial 21";
 12.1655 -\<close> ML \<open>
 12.1656 -"-----SPO ---";
 12.1657 -val t = TermC.str2term "(a \<up> 2) \<up> 3";
 12.1658 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1659 -if UnparseC.term t = "a \<up> 6" then ()
 12.1660 -else error "poly.sml: diff.behav. in make_polynomial 22";
 12.1661 -"-----SPO ---";
 12.1662 -val t = TermC.str2term "x \<up> 2 * y \<up> 2 + x * x \<up> 2 * y";
 12.1663 -val SOME (t,_) = rewrite_set_ thy false make_polynomial t; UnparseC.term t;
 12.1664 -if UnparseC.term t = "x \<up> 3 * y + x \<up> 2 * y \<up> 2" then ()
 12.1665 -else error "poly.sml: diff.behav. in make_polynomial 23";
 12.1666 -\<close> text \<open> (*loops*)
 12.1667 -"-----SPO ---";
 12.1668 -val t = (Thm.term_of o the o (TermC.parse thy)) "a \<up> 2 * (-a) \<up> 2";
 12.1669 -val SOME (t,_) = rewrite_set_ @ {theory} false make_polynomial t; UnparseC.term t;
 12.1670 -if (UnparseC.term t) = "a \<up> 4" then ()
 12.1671 -else error "poly.sml: diff.behav. in make_polynomial 24";
 12.1672 -"-----SPO ---";
 12.1673 -val t = TermC.str2term "a * b * b \<up> (-1) + a";
 12.1674 -\<close> text \<open>
 12.1675 -val SOME (t,_) = rewrite_set_ @ {theory} false make_polynomial t; UnparseC.term t;
 12.1676 -\<close> text \<open>
 12.1677 -UnparseC.term t
 12.1678 -\<close> text \<open>
 12.1679 -if UnparseC.term t = "2 * a" then ()
 12.1680 -else error "poly.sml: diff.behav. in make_polynomial 25";
 12.1681 -\<close> ML \<open>
 12.1682 -"-----SPO ---";
 12.1683 -val t = TermC.str2term "a*c*b \<up> (2*n) + 3*a + 5*b \<up> (2*n)*c*b";
 12.1684 -\<close> ML \<open>
 12.1685 -val SOME (t,_) = rewrite_set_ @{theory} false make_polynomial t; UnparseC.term t;
 12.1686 -\<close> ML \<open>
 12.1687 -(*+++*)UnparseC.term t
 12.1688 -\<close> ML \<open>
 12.1689 -if UnparseC.term t = "3 * a + 5 * b \<up> (1 + 2 * n) * c + a * b \<up> (2 * n) * c"
 12.1690 -then () else error "poly.sml: diff.behav. in make_polynomial 26";
 12.1691 -
 12.1692 -\<close> ML \<open>
 12.1693 -(*MG030627 -------------vvv-: Verschachtelte Terme -----------*)
 12.1694 -"-----SPO ---";
 12.1695 -val t = TermC.str2term "(1 + (x*y*a) + x) \<up> (1 + (x*y*a) + x)";
 12.1696 -val SOME (t,_) = rewrite_set_ @{theory} false make_polynomial t;
 12.1697 -if UnparseC.term t = "(1 + x + a * x * y) \<up> (1 + x + a * x * y)"
 12.1698 -then () else error "poly.sml: diff.behav. in make_polynomial 27";(*SPO*)
 12.1699 -
 12.1700 -\<close> ML \<open>
 12.1701 -val t = TermC.str2term "(1 + x*(y*z)*zz) \<up> (1 + x*(y*z)*zz)";
 12.1702 -\<close> ML \<open>
 12.1703 -val SOME (t,_) = rewrite_set_ @{theory} false make_polynomial t;
 12.1704 -\<close> ML \<open>
 12.1705 -if UnparseC.term t = "(1 + x * y * z * zz) \<up> (1 + x * y * z * zz)"
 12.1706 -then () else error "poly.sml: diff.behav. in make_polynomial 28";
 12.1707 -
 12.1708 -\<close> ML \<open>
 12.1709 -"-------- check pbl  'polynomial simplification' --------";
 12.1710 -"-------- check pbl  'polynomial simplification' --------";
 12.1711 -"-------- check pbl  'polynomial simplification' --------";
 12.1712 -val fmz = ["Term ((5*x \<up> 2 + 3) * (2*x \<up> 7 + 3) - (3*x \<up> 5 + 8) * (6*x \<up> 4 - 1))", "normalform N"];
 12.1713 -"-----0 ---";
 12.1714 -\<close> ML \<open>
 12.1715 -case Refine.refine fmz ["polynomial", "simplification"] of
 12.1716 -    [M_Match.Matches (["polynomial", "simplification"], _)] => ()
 12.1717 -  | _ => error "poly.sml diff.behav. in check pbl, Refine.refine";
 12.1718 -\<close> ML \<open>
 12.1719 -(*...if there is an error, then ...*)
 12.1720 -
 12.1721 -"-----1 ---";
 12.1722 -(*default_print_depth 7;*)
 12.1723 -val pbt = Problem.from_store ["polynomial", "simplification"];
 12.1724 -(*default_print_depth 3;*)
 12.1725 -(*if there is ...
 12.1726 -> val M_Match.NoMatch' {Given=gi, Where=wh, Find=fi,...} = M_Match.match_pbl fmz pbt;
 12.1727 -... then Rewrite.trace_on:*)
 12.1728 -                           
 12.1729 -"-----2 ---";
 12.1730 -Rewrite.trace_on := false;
 12.1731 -M_Match.match_pbl fmz pbt;
 12.1732 -Rewrite.trace_on := false;
 12.1733 -(*... if there is no rewrite, then there is something wrong with prls*)
 12.1734 -                              
 12.1735 -"-----3 ---";
 12.1736 -(*default_print_depth 7;*)
 12.1737 -val prls = (#prls o Problem.from_store) ["polynomial", "simplification"];
 12.1738 -(*default_print_depth 3;*)
 12.1739 -val t = TermC.str2term "((5*x \<up> 2 + 3) * (2*x \<up> 7 + 3) - (3*x \<up> 5 + 8) * (6*x \<up> 4 - 1)) is_polyexp";
 12.1740 -val SOME (t',_) = rewrite_set_ thy false prls t;
 12.1741 -if t' = @{term True} then () 
 12.1742 -else error "poly.sml: diff.behav. in check pbl 'polynomial..";
 12.1743 -(*... if this works, but --1-- does still NOT work, check types:*)
 12.1744 -
 12.1745 -"-----4 ---";
 12.1746 -(*show_types:=true;*)
 12.1747 -(*
 12.1748 -> val M_Match.NoMatch' {Given=gi, Where=wh, Find=fi,...} = M_Match.match_pbl fmz pbt;
 12.1749 -val wh = [False "(t_::real => real) (is_polyexp::real)"]
 12.1750 -...................... \<up> \<up> \<up>  \<up> ............... \<up> ^*)
 12.1751 -val M_Match.Matches' _ = M_Match.match_pbl fmz pbt;
 12.1752 -(*show_types:=false;*)
 12.1753 -
 12.1754 -
 12.1755 -\<close> ML \<open>
 12.1756 -"-------- me 'poly. simpl.' Schalk I p.63 No.267b -------";
 12.1757 -"-------- me 'poly. simpl.' Schalk I p.63 No.267b -------";
 12.1758 -"-------- me 'poly. simpl.' Schalk I p.63 No.267b -------";
 12.1759 -val fmz = ["Term ((5*x \<up> 2 + 3) * (2*x \<up> 7 + 3) - (3*x \<up> 5 + 8) * (6*x \<up> 4 - 1))", "normalform N"];
 12.1760 -val (dI',pI',mI') =
 12.1761 -  ("Poly",["polynomial", "simplification"],
 12.1762 -   ["simplification", "for_polynomials"]);
 12.1763 -val p = e_pos'; val c = []; 
 12.1764 -val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
 12.1765 -(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt;(*Add_Given "Term\n ((5 * x \<up> 2 + 3) * (2 * x \<up> 7 + 3) -\n  (3 * x \<up> 5 + 8) * (6 * x \<up> 4 - 1))"*)
 12.1766 -(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt;(*Add_Find "normalform N"*)
 12.1767 -
 12.1768 -(*+* )if I_Model.to_string ctxt (get_obj g_pbl pt (fst p)) =
 12.1769 -(*+*)  "[\n(0 ,[] ,false ,#Find ,Inc ??.Simplify.normalform ,(??.empty, [])), \n(1 ,[1] ,true ,#Given ,Cor ??.Simplify.Term\n ((5 * x \<up> 2 + 3) * (2 * x \<up> 7 + 3) -\n  (3 * x \<up> 5 + 8) * (6 * x \<up> 4 - 1)) ,(t_t, [(5 * x \<up> 2 + 3) * (2 * x \<up> 7 + 3) -\n(3 * x \<up> 5 + 8) * (6 * x \<up> 4 - 1)]))]"
 12.1770 -(*+*)then () else error "No.267b: I_Model.T CHANGED";
 12.1771 -( *+ ...could not be repaired in child of 7e314dd233fd ?!?*)
 12.1772 -
 12.1773 -\<close> ML \<open>
 12.1774 -(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt;(*nxt = Specify_Theory "Poly"*)
 12.1775 -(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt;(*Specify_Problem ["polynomial", "simplification"]*)
 12.1776 -\<close> ML \<open>
 12.1777 -(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt;(*Specify_Method ["simplification", "for_polynomials"]*)
 12.1778 -(*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt;(*Apply_Method ["simplification", "for_polynomials"]*)
 12.1779 -\<close> text \<open> (*loops*)
 12.1780 -(*[1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p c pt;(*Rewrite_Set "norm_Poly"*)
 12.1781 -
 12.1782 -\<close> text \<open>
 12.1783 -(*+++*)f (*= PpcKF (Problem [],.. *)
 12.1784 -\<close> text \<open>
 12.1785 -(*+*)if f2str f = "(5 * x \<up> 2 + 3) * (2 * x \<up> 7 + 3) -\n(3 * x \<up> 5 + 8) * (6 * x \<up> 4 - 1)"
 12.1786 -(*+*)then () else error "";
 12.1787 -
 12.1788 -(*[1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt;(*Empty_Tac: ERROR DETECTED Feb.2020*)
 12.1789 -
 12.1790 -\<close> text \<open>
 12.1791 -(*+*)if f2str f = "17 + 15 * x \<up> 2 + -48 * x \<up> 4 + 3 * x \<up> 5 + 6 * x \<up> 7 +\n-8 * x \<up> 9"
 12.1792 -(*+*)then () else error "poly.sml diff.behav. in me Schalk I p.63 No.267b -1";
 12.1793 -
 12.1794 -(*[1], Res* )val (p,_,f,nxt,_,pt) = me nxt p c pt;( *SINCE Feb.2020 LItool.find_next_step without result*)
 12.1795 -
 12.1796 -
 12.1797 -
 12.1798 -\<close> ML \<open>
 12.1799 -"-------- interSteps for Schalk 299a --------------------";
 12.1800 -"-------- interSteps for Schalk 299a --------------------";
 12.1801 -"-------- interSteps for Schalk 299a --------------------";
 12.1802 -reset_states ();
 12.1803 -CalcTree
 12.1804 -[(["Term ((x - y)*(x + (y::real)))", "normalform N"], 
 12.1805 -  ("Poly",["polynomial", "simplification"],
 12.1806 -  ["simplification", "for_polynomials"]))];
 12.1807 -Iterator 1;
 12.1808 -moveActiveRoot 1;
 12.1809 -\<close> text \<open> (*loops*)
 12.1810 -autoCalculate 1 CompleteCalc;
 12.1811 -val ((pt,p),_) = get_calc 1; Test_Tool.show_pt pt;
 12.1812 -
 12.1813 -\<close> text \<open>
 12.1814 -interSteps 1 ([1],Res)(*<ERROR> syserror in Detail_Step.go </ERROR>*);
 12.1815 -val ((pt,p),_) = get_calc 1; Test_Tool.show_pt pt;
 12.1816 -if existpt' ([1,1], Frm) pt then ()
 12.1817 -else error "poly.sml: interSteps doesnt work again 1";
 12.1818 -
 12.1819 -interSteps 1 ([1,1],Res)(*<ERROR> syserror in Detail_Step.go </ERROR>*);
 12.1820 -val ((pt,p),_) = get_calc 1; Test_Tool.show_pt pt;
 12.1821 -(*============ inhibit exn WN120316 ==============================================
 12.1822 -if existpt' ([1,1,1], Frm) pt then ()
 12.1823 -else error "poly.sml: interSteps doesnt work again 2";
 12.1824 -============ inhibit exn WN120316 ==============================================*)
 12.1825 -
 12.1826 -\<close> text \<open> (*loops*)
 12.1827 -"-------- norm_Poly NOT COMPLETE ------------------------";
 12.1828 -"-------- norm_Poly NOT COMPLETE ------------------------";
 12.1829 -"-------- norm_Poly NOT COMPLETE ------------------------";
 12.1830 -val thy = @ {theory AlgEin};
 12.1831 -
 12.1832 -val SOME (f',_) = rewrite_set_ thy false norm_Poly
 12.1833 -(TermC.str2term "L = k - 2 * q + (k - 2 * q) + (k - 2 * q) + (k - 2 * q) + senkrecht + oben");
 12.1834 -if UnparseC.term f' = "L = senkrecht + oben + 2 * 2 * k + 2 * -4 * q"
 12.1835 -then ((*norm_Poly NOT COMPLETE -- TODO MG*))
 12.1836 -else error "norm_Poly changed behavior";
 12.1837 -
 12.1838 -\<close> ML \<open>
 12.1839 -"-------- ord_make_polynomial ---------------------------";
 12.1840 -"-------- ord_make_polynomial ---------------------------";
 12.1841 -"-------- ord_make_polynomial ---------------------------";
 12.1842 -val t1 = TermC.str2term "2 * b + (3 * a + 3 * b)";
 12.1843 -val t2 = TermC.str2term "3 * a + 3 * b + 2 * b";
 12.1844 -
 12.1845 -\<close> ML \<open>
 12.1846 -if ord_make_polynomial true thy [] (t1, t2) then ()
 12.1847 -else error "poly.sml: diff.behav. in ord_make_polynomial";
 12.1848 -
 12.1849 -\<close> ML \<open>
 12.1850 -(*WN071202: \<up> why then is there no rewriting ...*)
 12.1851 -val term = TermC.str2term "2*b + (3*a + 3*b)";
 12.1852 -\<close> ML \<open>
 12.1853 -(*+++*)val SOME (t', []) = rewrite_set_ (@{theory "Isac_Knowledge"}) false order_add_mult term
 12.1854 -\<close> ML \<open>
 12.1855 -(*+++*)UnparseC.term t' = "a * 3 + (b * 2 + b * 3)";
 12.1856 -\<close> text \<open>
 12.1857 -val NONE = rewrite_set_ (@ {theory "Isac_Knowledge"}) false order_add_mult term;
 12.1858 -
 12.1859 -(*or why is there no rewriting this way...*)
 12.1860 -val t1 = TermC.str2term "2 * b + (3 * a + 3 * b)";
 12.1861 -val t2 = TermC.str2term "3 * a + (2 * b + 3 * b)";
 12.1862 -
 12.1863 -\<close> ML \<open>
 12.1864 -\<close> ML \<open>
 12.1865 -\<close>
 12.1866 -
 12.1867 -
 12.1868 +
 12.1869 +section \<open>===================================================================================\<close>
 12.1870  section \<open>===================================================================================\<close>
 12.1871  ML \<open>
 12.1872  \<close> ML \<open>
 12.1873 @@ -996,426 +1083,2086 @@
 12.1874  \<close> ML \<open>
 12.1875  \<close>
 12.1876  
 12.1877 -section \<open>============== check test/../ "MathEngBasic/rewrite.sml" ==========================\<close>
 12.1878 +
 12.1879 +section \<open>====== check test/../rational.sml =================================================\<close>
 12.1880  ML \<open>
 12.1881  \<close> ML \<open>
 12.1882 -"----------- compare all prepat's existing 2010 ------------------------------------------------";
 12.1883 -"----------- compare all prepat's existing 2010 ------------------------------------------------";
 12.1884 -"----------- compare all prepat's existing 2010 ------------------------------------------------";
 12.1885 -val thy = @{theory "Isac_Knowledge"};
 12.1886 -val t = @{term "a + b * x = (0 ::real)"};
 12.1887 -val pat = TermC.parse_patt thy "?l = (?r ::real)";
 12.1888 -val precond = TermC.parse_patt thy "is_polynomial (?l::real)";(*no infix def*)
 12.1889 -val precond = TermC.parse_patt thy "(?l::real) is_expanded"; 
 12.1890 -
 12.1891 -val inst = Pattern.match thy (pat, t) (Vartab.empty, Vartab.empty);
 12.1892 -val preinst = Envir.subst_term inst precond;
 12.1893 -UnparseC.term preinst;
 12.1894 -
 12.1895 -"===== Rational.thy: cancel ===";
 12.1896 -(* pat matched with the current term gives an environment 
 12.1897 -   (or not: hen the Rrls not applied);
 12.1898 -   if pre1 and pre2 evaluate to @{term True} in this environment,
 12.1899 -   then the Rrls is applied. *)
 12.1900 -val t = TermC.str2term "(a + b) / c ::real";
 12.1901 -val pat = TermC.parse_patt thy "?r / ?s ::real";
 12.1902 -val pres = [TermC.parse_patt thy "?r is_expanded", TermC.parse_patt thy "?s is_expanded"];
 12.1903 -val prepat = [(pres, pat)];
 12.1904 -val erls = rational_erls;
 12.1905 -(* erls got from Rrls {erls, prepat, scr = Rfuns {normal_form, ...}, ...} *)
 12.1906 -
 12.1907 -val subst = Pattern.match thy (pat, t) (Vartab.empty, Vartab.empty);
 12.1908 -val asms = map (Envir.subst_term subst) pres;
 12.1909 -if UnparseC.terms asms = "[\"a + b is_expanded\", \"c is_expanded\"]"
 12.1910 -then () else error "rewrite.sml: prepat cancel subst";
 12.1911 +(* Title: tests for rationals
 12.1912 +   Author: Walther Neuper
 12.1913 +   Use is subject to license terms.
 12.1914 +*)
 12.1915 +
 12.1916 +"-----------------------------------------------------------------------------";
 12.1917 +"-----------------------------------------------------------------------------";
 12.1918 +"table of contents -----------------------------------------------------------";
 12.1919 +"-----------------------------------------------------------------------------";
 12.1920 +"-------- fun poly_of_term ---------------------------------------------------";
 12.1921 +"-------- fun is_poly --------------------------------------------------------";
 12.1922 +"-------- fun term_of_poly ---------------------------------------------------";
 12.1923 +"-------- integration lev.1 fun factout_p_ -----------------------------------";
 12.1924 +"-------- integration lev.1 fun cancel_p_ ------------------------------------";
 12.1925 +"-------- integration lev.1 fun common_nominator_p_ --------------------------";
 12.1926 +"-------- integration lev.1 fun add_fraction_p_ ------------------------------";
 12.1927 +"Rfuns-------- and app_rev ...traced down from rewrite_set_ until prepats ---------";
 12.1928 +"Rfuns-------- fun rewrite_set_ cancel_p downto fun gcd_poly ----------------------";
 12.1929 +"-------- rls norm_Rational downto fun gcd_poly ------------------------------";
 12.1930 +"Rfuns-------- rls norm_Rational downto fun add_fraction_p_ -----------------------";
 12.1931 +"----------- rewrite_set_ Partial_Fractions norm_Rational --------------------------------------";
 12.1932 +"----------- fun check_frac_sum with Free A and Const AA ---------------------------------------";
 12.1933 +"----------- fun cancel_p with Const AA --------------------------------------------------------";
 12.1934 +"-------- rewrite_set_ cancel_p from: Mathematik 1 Schalk Reniets Verlag -----";
 12.1935 +"-------- rewrite_set_ add_fractions_p from: Mathematik 1 Schalk -------------";
 12.1936 +"-------- integration lev.1 -- lev.5: cancel_p_ & add_fractions_p_ -----------";
 12.1937 +"Rfuns-------- reverse rewrite ----------------------------------------------------";
 12.1938 +"Rfuns-------- 'reverse-ruleset' cancel_p -----------------------------------------";
 12.1939 +"-------- investigate rls norm_Rational --------------------------------------";
 12.1940 +"-------- examples: rls norm_Rational ----------------------------------------";
 12.1941 +"-------- rational numerals --------------------------------------------------";
 12.1942 +"-------- examples cancellation from: Mathematik 1 Schalk --------------------";
 12.1943 +"-------- examples common denominator from: Mathematik 1 Schalk --------------";
 12.1944 +"-------- examples multiply and cancel from: Mathematik 1 Schalk -------------";
 12.1945 +"-------- examples common denominator and multiplication from: Schalk --------";
 12.1946 +"-------- examples double fractions from: Mathematik 1 Schalk ----------------";
 12.1947 +"-------- me Schalk I No.186 -------------------------------------------------";
 12.1948 +"-------- interSteps ..Simp_Rat_Double_No-1.xml ------------------------------";
 12.1949 +"-------- interSteps ..Simp_Rat_Cancel_No-1.xml ------------------------------";
 12.1950 +"-------- investigate rulesets for cancel_p ----------------------------------";
 12.1951 +"-------- fun eval_get_denominator -------------------------------------------";
 12.1952 +"-------- several errpats in complicated term --------------------------------";
 12.1953 +"-------- WN1309xx non-terminating rls norm_Rational -------------------------";
 12.1954 +"-----------------------------------------------------------------------------";
 12.1955 +"-----------------------------------------------------------------------------";
 12.1956 +
 12.1957 +
 12.1958  \<close> ML \<open>
 12.1959 -(*+*)val ([], false) = eval__true thy 0 asms [] erls
 12.1960 +"-------- fun poly_of_term ---------------------------------------------------";
 12.1961 +"-------- fun poly_of_term ---------------------------------------------------";
 12.1962 +"-------- fun poly_of_term ---------------------------------------------------";
 12.1963 +val thy = @{theory Partial_Fractions};
 12.1964 +val ctxt = Proof_Context.init_global @{theory}
 12.1965 +val vs = TermC.vars_of (the (parseNEW ctxt "12 * x \<up> 3 * y \<up> 4 * z \<up> 6"));
 12.1966 +
 12.1967 +if poly_of_term vs (TermC.str2term "12::real") = SOME [(12, [0, 0, 0])]
 12.1968 +then () else error "poly_of_term 1 changed";
 12.1969 +
 12.1970 +if poly_of_term vs (TermC.str2term "x::real") = SOME [(1, [1, 0, 0])]
 12.1971 +then () else error "poly_of_term 2 changed";
 12.1972 +
 12.1973 +if         poly_of_term vs (TermC.str2term "12 * x \<up> 3") = SOME [(12, [3, 0, 0])]
 12.1974 +then () else error "poly_of_term 3 changed";
 12.1975 +"~~~~~ fun poly_of_term , args:"; val (vs, t) =
 12.1976 +  (vs, (TermC.str2term "12 * x \<up> 3"));
 12.1977 +
 12.1978 +           monom_of_term vs (1, replicate (length vs) 0) t;(*poly malformed 1 with x \<up> 3*)
 12.1979 +"~~~~~ fun monom_of_term , args:"; val (vs, (c, es), (Const ("Groups.times_class.times", _) $ m1 $ m2)) =
 12.1980 +  (vs, (1, replicate (length vs) 0), t);
 12.1981 +    val (c', es') =
 12.1982 +
 12.1983 +           monom_of_term vs (c, es) m1;
 12.1984 +"~~~~~ fun monom_of_term , args:"; val (vs, (c, es), (Const ("Transcendental.powr", _) $ (t as Free _) $ (Const ("Num.numeral_class.numeral", _) $ num)) ) =
 12.1985 +  (vs, (c', es'), m2);
 12.1986 +(*+*)c = 12;
 12.1987 +(*+*)(num |> HOLogic.dest_numeral |> list_update es (find_index (curry op = t) vs)) = [3, 0, 0];
 12.1988 +
 12.1989 +if (c, num |> HOLogic.dest_numeral |> list_update es (find_index (curry op = t) vs)) = (12, [3, 0, 0])
 12.1990 +then () else error "monom_of_term (powr): return value CHANGED";
 12.1991 +
 12.1992 +if poly_of_term vs (TermC.str2term "12 * x \<up> 3 * y \<up> 4 * z \<up> 6") = SOME [(12, [3, 4, 6])]
 12.1993 +then () else error "poly_of_term 4 changed";
 12.1994 +
 12.1995 +if poly_of_term vs (TermC.str2term "1 + 2 * x \<up> 3 * y \<up> 4 * z \<up> 6 + y") =
 12.1996 +  SOME [(1, [0, 0, 0]), (1, [0, 1, 0]), (2, [3, 4, 6])]
 12.1997 +then () else error "poly_of_term 5 changed";
 12.1998 +
 12.1999 +(*poly_of_term is quite liberal:*)
 12.2000 +(*the coefficient may be somewhere, the order of variables and the parentheses 
 12.2001 +  within a monomial are arbitrary*)
 12.2002 +if poly_of_term vs (TermC.str2term "y \<up> 4 * (x \<up> 3 * 12 * z \<up> 6)") = SOME [(12, [3, 4, 6])]
 12.2003 +then () else error "poly_of_term 6 changed";
 12.2004 +
 12.2005 +(*there may even be more than 1 coefficient:*)
 12.2006 +if poly_of_term vs (TermC.str2term "2 * y \<up> 4 * (x \<up> 3 * 6 * z \<up> 6)") = SOME [(12, [3, 4, 6])]
 12.2007 +then () else error "poly_of_term 7 changed";
 12.2008 +
 12.2009 +(*the order and the parentheses within monomials are arbitrary:*)
 12.2010 +if poly_of_term vs (TermC.str2term "2 * x \<up> 3 * y \<up> 4 * z \<up> 6 + (7 * y \<up> 8 + 1)")
 12.2011 +  = SOME [(1, [0, 0, 0]), (7, [0, 8, 0]), (2, [3, 4, 6])]
 12.2012 +then () else error "poly_of_term 8 changed";
 12.2013 +
 12.2014 +\<close> ML \<open> (*//---------------- inserted for --- fun monoms_of_term --- -----------------------\\*)
 12.2015 +(*from --- rls norm_Rational downto fun gcd_poly ---*)
 12.2016 +val t = TermC.str2term (*copy from above: "::real" is not required due to " \<up> "*)
 12.2017 +  ("(-12 + 4 * y + 3 * x \<up> 2 + -1 * (x \<up> 2 * y)) /" ^
 12.2018 +  "(-18 + -9 * x + 2 * y \<up> 2 + x * y \<up> 2)");
 12.2019  \<close> ML \<open>
 12.2020 -if ([], true) = eval__true thy 0 asms [] erls
 12.2021 -then () else error "rewrite.sml: prepat cancel eval__true";
 12.2022 -
 12.2023 -"===== Rational.thy: add_fractions_p ===";
 12.2024 -(* if each pat* TermC.matches with the current term, the Rrls is applied
 12.2025 -   (there are no preconditions to be checked, they are @{term True}) *)
 12.2026 -val t = TermC.str2term "a / b + 1 / 2";
 12.2027 -val pat = TermC.parse_patt thy "?r / ?s + ?u / ?v";
 12.2028 -val pres = [@{term True}];
 12.2029 -val prepat = [(pres, pat)];
 12.2030 -val erls = rational_erls;
 12.2031 -(* erls got from Rrls {erls, prepat, scr = Rfuns {normal_form, ...}, ...} *)
 12.2032 -
 12.2033 -val subst = Pattern.match thy (pat, t) (Vartab.empty, Vartab.empty);
 12.2034 +"~~~~~ fun cancel_p_, args:"; val (t) = (t);
 12.2035 +val opt = check_fraction t;
 12.2036  \<close> ML \<open>
 12.2037 -(*+*)val ([], false) = eval__true thy 0 (map (Envir.subst_term subst) pres) [] erls
 12.2038 +val SOME (numerator, denominator) = opt
 12.2039  \<close> ML \<open>
 12.2040 -if ([], true) = eval__true thy 0 (map (Envir.subst_term subst) pres) [] erls
 12.2041 -then () else error "rewrite.sml: prepat add_fractions_p";
 12.2042 -
 12.2043 -"===== Poly.thy: order_mult_ ===";
 12.2044 -          (* ?p matched with the current term gives an environment,
 12.2045 -             which evaluates (the instantiated) "p is_multUnordered" to true*)
 12.2046 -val t = TermC.str2term "x \<up> 2 * x";
 12.2047 -val pat = TermC.parse_patt thy "?p :: real"
 12.2048 -val pres = [TermC.parse_patt thy "?p is_multUnordered"];
 12.2049 -val prepat = [(pres, pat)];
 12.2050 -val erls = Rule_Set.append_rules "Rule_Set.empty-is_multUnordered" Rule_Set.empty
 12.2051 -		      [Eval ("Poly.is_multUnordered", 
 12.2052 -                             eval_is_multUnordered "")];
 12.2053 -
 12.2054 -val subst = Pattern.match thy (pat, t) (Vartab.empty, Vartab.empty);
 12.2055 -val asms = map (Envir.subst_term subst) pres;
 12.2056 +(*+*)UnparseC.term numerator = "- 12 + 4 * y + 3 * x \<up> 2 + - 1 * (x \<up> 2 * y)"; (*isa -- isa2*)
 12.2057 +(*+*)UnparseC.term denominator = "- 18 + - 9 * x + 2 * y \<up> 2 + x * y \<up> 2";     (*isa -- isa2*)
 12.2058  \<close> ML \<open>
 12.2059 -(*+*)UnparseC.terms asms = "[\"a + b is_expanded\", \"c is_expanded\"]"
 12.2060 +        val vs = TermC.vars_of t
 12.2061  \<close> ML \<open>
 12.2062 -if UnparseC.terms asms = "[\"x \<up> 2 * x is_multUnordered\"]"
 12.2063 -then () else error "rewrite.sml: prepat order_mult_ subst";
 12.2064 +UnparseC.terms vs = "[\"x\", \"y\"]";
 12.2065  \<close> ML \<open>
 12.2066 -(*+*)val ([], false) = eval__true thy 0 asms [] erls
 12.2067 +        val baseT = type_of numerator
 12.2068 +        val expT = HOLogic.realT;
 12.2069  \<close> ML \<open>
 12.2070 -if ([], true) = eval__true thy 0 asms [] erls
 12.2071 -then () else error "rewrite.sml: prepat order_mult_ eval__true";
 12.2072 -
 12.2073 -
 12.2074 +val (NONE, NONE) = (poly_of_term vs numerator, poly_of_term vs denominator);  (*isa <> isa2*)
 12.2075 +\<close> ML \<open> (*------------------ inserted for --- fun poly_of_term --- ---------------------------*)
 12.2076  \<close> ML \<open>
 12.2077 -"----------- fun app_rev, Rrls, ----------------------------------------------------------------";
 12.2078 -"----------- fun app_rev, Rrls, ----------------------------------------------------------------";
 12.2079 -"----------- fun app_rev, Rrls, ----------------------------------------------------------------";
 12.2080 -val t = TermC.str2term "x \<up> 2 * x";
 12.2081 -
 12.2082 +val vs = [@{term "x::real"}, @{term "y::real"}]
 12.2083  \<close> ML \<open>
 12.2084 -(*+*)is_multUnordered t = false
 12.2085 +val p1 = TermC.str2term "- 12 + 4 * y + 3 * x \<up> 2 + - 1 * (x \<up> 2 * y)";
 12.2086 +val p2 = TermC.str2term "- 18 + - 9 * x + 2 * y \<up> 2 + x * y \<up> 2";
 12.2087  \<close> ML \<open>
 12.2088 -if is_multUnordered t then () else error "rewrite.sml diff. is_multUnordered 2";
 12.2089 +           poly_of_term vs p1;
 12.2090 +"~~~~~ fun poly_of_term , args:"; val (vs, (t as Const ("Groups.plus_class.plus", _) $ _ $ _)) =
 12.2091 +  (vs, p1);
 12.2092  \<close> ML \<open>
 12.2093 -val tm = TermC.str2term "(x \<up> 2 * x) is_multUnordered";
 12.2094 +(*+*)UnparseC.term p1 = "- 12 + 4 * y + 3 * x \<up> 2 + - 1 * (x \<up> 2 * y)";
 12.2095  \<close> ML \<open>
 12.2096 -
 12.2097 -     eval_is_multUnordered "testid" "" tm thy;
 12.2098 -case eval_is_multUnordered "testid" "" tm thy of
 12.2099 -  SOME
 12.2100 -    ("testidx \<up> 2 * x_",
 12.2101 -     Const ("HOL.Trueprop", _) $
 12.2102 -       (Const ("HOL.eq", _) $
 12.2103 -         (Const ("Poly.is_multUnordered", _) $
 12.2104 -           (Const ("Groups.times_class.times", _) $
 12.2105 -             (Const ("Transcendental.powr", _) $ Free ("x", _) $ _ ) $ Free ("x", _))) $
 12.2106 -         Const ("HOL.False", _))) => ()
 12.2107 -(*                   ^^^^^^             compare ---vvv *)
 12.2108 -| _ => error "rewrite.sml diff. eval_is_multUnordered 2b CHANGED";
 12.2109 +      t |> monoms_of_term vs; (*ERROR poly malformed 2 with - 12*)
 12.2110  \<close> ML \<open>
 12.2111 -case eval_is_multUnordered "testid" "" tm thy of
 12.2112 -    SOME (_, Const ("HOL.Trueprop", _) $ 
 12.2113 -                   (Const ("HOL.eq", _) $
 12.2114 -                          (Const ("Poly.is_multUnordered", _) $ _) $ 
 12.2115 -                          Const ("HOL.True", _))) => ()
 12.2116 -  | _ => error "rewrite.sml diff. eval_is_multUnordered 2b";
 12.2117 -
 12.2118 -tracing "----- begin rewrite x \<up> 2 * x ---"; Rewrite.trace_on := false;
 12.2119 +"~~~~~ fun monoms_of_term , args:"; val (vs, (Const ("Groups.plus_class.plus", _) $ ms1 $ ms2)) =
 12.2120 +  (vs, t);
 12.2121  \<close> ML \<open>
 12.2122 -val SOME (t', _) = rewrite_set_ thy true order_mult_ t;
 12.2123 +(*+*)UnparseC.term ms1 = "- 12 + 4 * y + 3 * x \<up> 2";
 12.2124  \<close> ML \<open>
 12.2125 -tracing "----- end rewrite x \<up> 2 * x ---"; Rewrite.trace_on := false;
 12.2126 +          (monoms_of_term vs ms1) (*ERROR poly malformed 2 with - 12*)
 12.2127 +\<close> ML \<open>
 12.2128 +"~~~~~ fun monoms_of_term , args:"; val (vs, (Const ("Groups.plus_class.plus", _) $ ms1 $ ms2)) =
 12.2129 +  (vs, ms1);
 12.2130 +\<close> ML \<open>
 12.2131 +(*+*)UnparseC.term ms1 = "- 12 + 4 * y";
 12.2132 +\<close> ML \<open>
 12.2133 +          (monoms_of_term vs ms1) (*ERROR poly malformed 2 with - 12*)
 12.2134 +\<close> ML \<open>
 12.2135 +"~~~~~ fun monoms_of_term , args:"; val (vs, (Const ("Groups.plus_class.plus", _) $ ms1 $ ms2)) =
 12.2136 +  (vs, ms1);
 12.2137 +\<close> ML \<open>
 12.2138 +(*+*)UnparseC.term ms1 = "- 12";
 12.2139 +\<close> ML \<open>
 12.2140 +          (monoms_of_term vs ms1) (*ERROR poly malformed 2 with - 12*)
 12.2141 +\<close> ML \<open>
 12.2142 +"~~~~~ fun monoms_of_term , args:"; val (vs, (t as Const _)) =
 12.2143 +  (vs, ms1);
 12.2144 +\<close> ML \<open>
 12.2145 +ms1
 12.2146 +\<close> ML \<open>
 12.2147 +UnparseC.term ms1 = "- 12"
 12.2148 +\<close> ML \<open>
 12.2149 +\<close> ML \<open>
 12.2150 +\<close> ML \<open>
 12.2151 +\<close> ML \<open>
 12.2152 +\<close> ML \<open>
 12.2153 +\<close> ML \<open> (*\\---------------- inserted for --- fun monoms_of_term --- -----------------------//*)
 12.2154 +
 12.2155 +\<close> ML \<open>
 12.2156 +"-------- fun is_poly --------------------------------------------------------";
 12.2157 +"-------- fun is_poly --------------------------------------------------------";
 12.2158 +"-------- fun is_poly --------------------------------------------------------";
 12.2159 +if is_poly (TermC.str2term "2 * x \<up> 3 * y \<up> 4 * z \<up> 6 + 7 * y \<up> 8 + 1")
 12.2160 +then () else error "is_poly 1 changed";
 12.2161 +if not (is_poly (TermC.str2term "2 * (x \<up> 3 * y \<up> 4 * z \<up> 6 + 7) * y \<up> 8 + 1"))
 12.2162 +then () else error "is_poly 2 changed";
 12.2163 +
 12.2164 +\<close> ML \<open>
 12.2165 +"-------- fun term_of_poly ---------------------------------------------------";
 12.2166 +"-------- fun term_of_poly ---------------------------------------------------";
 12.2167 +"-------- fun term_of_poly ---------------------------------------------------";
 12.2168 +val expT = HOLogic.realT
 12.2169 +val Free (_, baseT) = (hd o vars o TermC.str2term) "12 * x \<up> 3 * y \<up> 4 * z \<up> 6";
 12.2170 +val p = [(1, [0, 0, 0]), (7, [0, 8, 0]), (2, [3, 4, 5])]
 12.2171 +val vs = TermC.vars_of (the (parseNEW ctxt "12 * x \<up> 3 * y \<up> 4 * z \<up> 6"))
 12.2172 +(*precondition for [(c, es),...]: legth es = length vs*)
 12.2173 +;
 12.2174 +if UnparseC.term (term_of_poly baseT expT vs p) = "1 + 7 * y \<up> 8 + 2 * x \<up> 3 * y \<up> 4 * z \<up> 5"
 12.2175 +then () else error "term_of_poly 1 changed";
 12.2176 +
 12.2177 +\<close> text \<open> (* factout_p_ \<longrightarrow> NONE *)
 12.2178 +"-------- integration lev.1 fun factout_p_ -----------------------------------";
 12.2179 +"-------- integration lev.1 fun factout_p_ -----------------------------------";
 12.2180 +"-------- integration lev.1 fun factout_p_ -----------------------------------";
 12.2181 +val t = TermC.str2term "(x \<up> 2 + -1*y \<up> 2) / (x \<up> 2 + -1*x*y)"
 12.2182 +val SOME (t', asm) = factout_p_ thy t;
 12.2183 +if UnparseC.term t' = "(x + y) * (x + -1 * y) / (x * (x + -1 * y))"
 12.2184 +then () else error ("factout_p_ term 1 changed: " ^ UnparseC.term t')
 12.2185 +;
 12.2186 +if UnparseC.terms asm = "[\"x \<noteq> 0\",\"x + -1 * y \<noteq> 0\"]"
 12.2187 +then () else error "factout_p_ asm 1 changed"
 12.2188 +;
 12.2189 +val t = TermC.str2term "nothing + to_cancel ::real";
 12.2190 +if NONE = factout_p_ thy t then () else error "factout_p_ doesn't report non-applicable";
 12.2191 +;
 12.2192 +val t = TermC.str2term "((3 * x \<up> 2 + 6 *x + 3) / (2*x + 2))";
 12.2193 +val SOME (t', asm) = factout_p_ thy t;
 12.2194 +if UnparseC.term t' = "(3 + 3 * x) * (1 + x) / (2 * (1 + x))" andalso 
 12.2195 +  UnparseC.terms asm = "[\"1 + x \<noteq> 0\"]"
 12.2196 +then () else error "factout_p_ 1 changed";
 12.2197 +
 12.2198 +\<close> ML \<open>
 12.2199 +cancel_p_
 12.2200 +\<close> text \<open> (*cancel_p_ \<longrightarrow> NONE*)
 12.2201 +"-------- integration lev.1 fun cancel_p_ ------------------------------------";
 12.2202 +"-------- integration lev.1 fun cancel_p_ ------------------------------------";
 12.2203 +"-------- integration lev.1 fun cancel_p_ ------------------------------------";
 12.2204 +val t = TermC.str2term "(x \<up> 2 + -1*y \<up> 2) / (x \<up> 2 + -1*x*y)"
 12.2205 +val SOME (t', asm) = cancel_p_ thy t;
 12.2206 +if (UnparseC.term t', UnparseC.terms asm) = ("(x + y) / x", "[\"x \<noteq> 0\"]")
 12.2207 +then () else error ("cancel_p_ (t', asm) 1 changed: " ^ UnparseC.term t')
 12.2208 +;
 12.2209 +val t = TermC.str2term "nothing + to_cancel ::real";
 12.2210 +if NONE = cancel_p_ thy t then () else error "cancel_p_ doesn't report non-applicable";
 12.2211 +;
 12.2212 +val t = TermC.str2term "((3 * x \<up> 2 + 6 *x + 3) / (2*x + 2))";
 12.2213 +val SOME (t', asm) = cancel_p_ thy t;
 12.2214 +if UnparseC.term t' = "(3 + 3 * x) / 2" andalso UnparseC.terms asm = "[]"
 12.2215 +then () else error "cancel_p_ 1 changed";
 12.2216 +
 12.2217 +\<close> ML \<open>
 12.2218 +"-------- integration lev.1 fun common_nominator_p_ --------------------------";
 12.2219 +"-------- integration lev.1 fun common_nominator_p_ --------------------------";
 12.2220 +"-------- integration lev.1 fun common_nominator_p_ --------------------------";
 12.2221 +val t = TermC.str2term ("y / (a*x + b*x + c*x) " ^
 12.2222 +              (* n1    d1                   *)
 12.2223 +                "+ a / (x*y)");
 12.2224 +              (* n2    d2                   *)
 12.2225 +val SOME (t', asm) = common_nominator_p_ thy t;
 12.2226 +if UnparseC.term t' =
 12.2227 +      ("y * y / (x * ((a + b + c) * y)) " ^
 12.2228 +  (*  n1  *d2'/ (c'* ( d1'        *d2')) *)
 12.2229 +     "+ a * (a + b + c) / (x * ((a + b + c) * y))")
 12.2230 +   (*  n2 * d1'         / (c'* ( d1'        *d2')) *)
 12.2231 +then () else error "common_nominator_p_ term 1 changed";
 12.2232 +if UnparseC.terms asm = "[\"a + b + c \<noteq> 0\", \"y \<noteq> 0\", \"x \<noteq> 0\"]"
 12.2233 +then () else error "common_nominator_p_ asm 1 changed"
 12.2234 +
 12.2235 +\<close> text \<open> (*common_nominator_p_ \<longrightarrow> NONE*)
 12.2236 +"-------- example in mail Nipkow";
 12.2237 +val t = TermC.str2term "x/(x \<up> 2 + -1*y \<up> 2) + y/(x \<up> 2 + -1*x*y)";
 12.2238 +val SOME (t', asm) = common_nominator_p_ thy t;
 12.2239 +if UnparseC.term t' = "x * x / " ^ 
 12.2240 +                 "((x + -1 * y) * ((x + y) * x))" ^
 12.2241 +            " +\n" ^ 
 12.2242 +                 "y * (x + y) / " ^ 
 12.2243 +                 "((x + -1 * y) * ((x + y) * x))"
 12.2244 +then () else error "common_nominator_p_ term 2 changed"
 12.2245 +;
 12.2246 +if UnparseC.terms asm = "[\"x + y \<noteq> 0\",\"x \<noteq> 0\",\"x + -1 * y \<noteq> 0\"]"
 12.2247 +then () else error "common_nominator_p_ asm 2 changed"
 12.2248 +
 12.2249 +"-------- example: applicable tested by SML code";
 12.2250 +val t = TermC.str2term "nothing / to_add";
 12.2251 +if NONE = common_nominator_p_ thy t then () else error "common_nominator_p_ term 3 changed";
 12.2252 +;
 12.2253 +val t = TermC.str2term "((x + (-1)) / (x + 1)) + ((x + 1) / (x + (-1)))";
 12.2254 +val SOME (t', asm) = common_nominator_p_ thy t;
 12.2255 +if UnparseC.term t' = 
 12.2256 +   "(x + -1) * (-1 + x) / ((1 + x) * (-1 + x)) +\n(x + 1) * (1 + x) / ((1 + x) * (-1 + x))"
 12.2257 +  andalso UnparseC.terms asm = "[\"1 + x \<noteq> 0\",\"-1 + x \<noteq> 0\"]"
 12.2258 +then () else error "common_nominator_p_ 3 changed";
 12.2259 +
 12.2260 +\<close> text \<open> (*add_fraction_p_ \<longrightarrow> NONE*)
 12.2261 +"-------- integration lev.1 fun add_fraction_p_ ------------------------------";
 12.2262 +"-------- integration lev.1 fun add_fraction_p_ ------------------------------";
 12.2263 +"-------- integration lev.1 fun add_fraction_p_ ------------------------------";
 12.2264 +val t = TermC.str2term "((x + (-1)) / (x + 1)) + ((x + 1) / (x + (-1)))";
 12.2265 +val SOME (t', asm) = add_fraction_p_ thy t;
 12.2266 +if UnparseC.term t' = "(2 + 2 * x \<up> 2) / (-1 + x \<up> 2)" 
 12.2267 +then () else error "add_fraction_p_ 3 changed";
 12.2268 +;
 12.2269 +if UnparseC.terms asm = "[\"-1 + x \<up> 2 \<noteq> 0\"]"
 12.2270 +then () else error "add_fraction_p_ 3 changed";
 12.2271 +;
 12.2272 +val t = TermC.str2term "nothing / to_add";
 12.2273 +if NONE = add_fraction_p_ thy t then () else error "add_fraction_p_ term 3 changed";
 12.2274 +;
 12.2275 +val t = TermC.str2term "((x + (-1)) / (x + 1)) + ((x + 1) / (x + (-1)))";
 12.2276 +val SOME (t', asm) = add_fraction_p_ thy t;
 12.2277 +if UnparseC.term t' = "(2 + 2 * x \<up> 2) / (-1 + x \<up> 2)" andalso
 12.2278 +  UnparseC.terms asm = "[\"-1 + x \<up> 2 \<noteq> 0\"]"
 12.2279 +then () else error "add_fraction_p_ 3 changed";
 12.2280 +
 12.2281 +\<close> ML \<open>
 12.2282 +"-------- and app_rev ...traced down from rewrite_set_ until prepats ---------";
 12.2283 +"-------- and app_rev ...traced down from rewrite_set_ until prepats ---------";
 12.2284 +"-------- and app_rev ...traced down from rewrite_set_ until prepats ---------";
 12.2285 +(* trace down until prepats are evaluated 
 12.2286 +  (which does not to work, because substitution is not done -- compare rew_sub!);
 12.2287 +  keep this sequence for the case, factout_p, cancel_p, common_nominator_p, add_fraction_p
 12.2288 +  (again) get prepat = [] changed to <>[]. *)
 12.2289 +val t = TermC.str2term "(x \<up> 2 + -1*y \<up> 2) / (x \<up> 2 + -1*x*y)";
 12.2290 +
 12.2291 +(*rewrite_set_ @{theory Isac_Knowledge} true cancel t = NONE; !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!*)
 12.2292 +"~~~~~ fun rewrite_set_, args:"; val (thy, bool, rls, term) = (thy, false, cancel_p, t);
 12.2293 +"~~~~~ fun rewrite__set_, args:"; val (thy, i, _, _, (rrls as Rrls _), t) =
 12.2294 +  (thy, 1, bool, [], rls, term);
 12.2295 +(*val (t', asm, rew) = app_rev thy (i+1) rrls t; rew = false!!!!!!!!!!!!!!!!!!!!!*)
 12.2296 +"~~~~~ and app_rev, args:"; val (thy, i, rrls, t) = (thy, (i+1), rrls, t);
 12.2297 +    fun chk_prepat thy erls [] t = true
 12.2298 +      | chk_prepat thy erls prepat t =
 12.2299 +        let
 12.2300 +          fun chk (pres, pat) =
 12.2301 +            (let 
 12.2302 +              val subst: Type.tyenv * Envir.tenv =
 12.2303 +                Pattern.match thy (pat, t) (Vartab.empty, Vartab.empty)
 12.2304 +             in
 12.2305 +              snd (eval__true thy (i + 1) (map (Envir.subst_term subst) pres) [] erls)
 12.2306 +             end) handle Pattern.MATCH => false
 12.2307 +           fun scan_ f [] = false (*scan_ NEVER called by []*)
 12.2308 +             | scan_ f (pp::pps) =
 12.2309 +               if f pp then true else scan_ f pps;
 12.2310 +        in scan_ chk prepat end;
 12.2311 +    (* apply the normal_form of a rev-set *)
 12.2312 +    fun app_rev' thy (Rrls{erls,prepat,scr=Rfuns{normal_form,...},...}) t =
 12.2313 +      if chk_prepat thy erls prepat t
 12.2314 +      then ((*tracing("### app_rev': t = "^UnparseC.term t);*) normal_form t)
 12.2315 +      else NONE;
 12.2316 +(*  val opt = app_rev' thy rrls t  ..NONE!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!*)
 12.2317 +"~~~~~ and app_rev', args:"; val (thy, (Rrls{erls,prepat,scr=Rfuns{normal_form,...},...}), t) =
 12.2318 +  (thy, rrls, t);
 12.2319 +(* chk_prepat thy erls prepat t = false!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!*)
 12.2320 +(* app_sub thy i rrls t = false!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!*)
 12.2321 +"~~~~~ fun chk_prepat, args:"; val (thy, erls, prepat, t) = (thy, erls, prepat, t);
 12.2322 +          fun chk (pres, pat) =
 12.2323 +            (let 
 12.2324 +              val subst: Type.tyenv * Envir.tenv =
 12.2325 +                Pattern.match thy (pat, t) (Vartab.empty, Vartab.empty)
 12.2326 +             in
 12.2327 +              snd (eval__true thy (i + 1) (map (Envir.subst_term subst) pres) [] erls)
 12.2328 +             end) handle Pattern.MATCH => false
 12.2329 +           fun scan_ f [] = false (*scan_ NEVER called by []*)
 12.2330 +             | scan_ f (pp::pps) =
 12.2331 +               if f pp then true else scan_ f pps;
 12.2332 +
 12.2333 +(*========== inhibit exn WN130823: prepat is empty ====================================
 12.2334 +(* scan_ chk prepat = false!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!*)
 12.2335 +"~~~~~ fun , args:"; val (f, (pp::pps)) = (chk, prepat);
 12.2336 +f;
 12.2337 +val ([t1, t2], t) = pp;
 12.2338 +UnparseC.term t1 = "?r is_expanded";
 12.2339 +UnparseC.term t2 = "?s is_expanded";
 12.2340 +UnparseC.term t = "?r / ?s";
 12.2341 +(* f pp = false!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!*)
 12.2342 +"~~~~~ fun chk, args:"; val (pres, pat) = (pp);
 12.2343 +              val subst: Type.tyenv * Envir.tenv =
 12.2344 +                Pattern.match thy (pat, t) (Vartab.empty, Vartab.empty)
 12.2345 +(*subst = 
 12.2346 +  ({}, {(("r", 0), ("real", Var (("r", 0), "real"))), 
 12.2347 +        (("s", 0), ("real", Var (("s", 0), "real")))}*)
 12.2348 +;
 12.2349 +              snd (eval__true thy (i + 1) (map (Envir.subst_term subst) pres) [] erls)
 12.2350 +"~~~~~ fun eval__true, args:"; val (thy, i, asms, bdv, rls) =
 12.2351 +  (thy, (i + 1), (map (Envir.subst_term subst) pres), [], erls);
 12.2352 +UnparseC.terms asms;                         (* = "[\"?r is_expanded\",\"?s is_expanded\"]"*)
 12.2353 +asms = [@{term True}] orelse asms = []; (* = false*)
 12.2354 +asms = [@{term False}]                ; (* = false*)
 12.2355 +"~~~~~ fun chk, args:"; val (indets, (a::asms)) = ([], asms);
 12.2356 +bdv (*= []: _a list*);
 12.2357 +val bdv : (term * term) list = [];
 12.2358 +rewrite__set_ thy (i+1) false;
 12.2359 +UnparseC.term a = "?r is_expanded"; (*hier m"usste doch der Numerator eingesetzt sein ??????????????*)
 12.2360 +val SOME (Const ("HOL.False", _), []) = rewrite__set_ thy (i+1) false bdv rls a
 12.2361 +============ inhibit exn WN130823: prepat is empty ===================================*)
 12.2362 +
 12.2363 +\<close> ML \<open>
 12.2364 +"-------- fun rewrite_set_ cancel_p downto fun gcd_poly ----------------------";
 12.2365 +"-------- fun rewrite_set_ cancel_p downto fun gcd_poly ----------------------";
 12.2366 +"-------- fun rewrite_set_ cancel_p downto fun gcd_poly ----------------------";
 12.2367 +val t = TermC.str2term "(12 * x * y) / (8 * y \<up> 2 )";
 12.2368 +(* "-------- example 187a": exception Div raised...
 12.2369 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;*)
 12.2370 +val t = TermC.str2term "(8 * x \<up> 2 * y * z ) / (18 * x * y \<up> 2 * z )";
 12.2371 +(* "-------- example 187b": doesn't terminate...
 12.2372 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;*)
 12.2373 +val t = TermC.str2term "(9 * x \<up> 5 * y \<up> 2 * z \<up> 4) / (15 * x \<up> 6 * y \<up> 3 * z )";
 12.2374 +(* "-------- example 187c": doesn't terminate...
 12.2375 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;*)
 12.2376 +"~~~~~ fun rewrite_set_, args:"; val (thy, bool, rls, term) = (@{theory Isac_Knowledge}, false, cancel_p, t);
 12.2377 +(* WN130827: exception Div raised...
 12.2378 +rewrite__set_ thy 1 bool [] rls term
 12.2379 +*)
 12.2380 +"~~~~~ and rewrite__set_, args:"; val (thy, i, _, _, (rrls as Rrls _), t) =
 12.2381 +  (thy, 1, bool, [], rls, term);
 12.2382 +(* WN130827: exception Div raised...
 12.2383 +	val (t', asm, rew) = app_rev thy (i+1) rrls t
 12.2384 +*)
 12.2385 +"~~~~~ fun app_rev, args:"; val (thy, i, rrls, t) = (thy, (i+1), rrls, t);
 12.2386 +(* WN130827: exception Div raised...
 12.2387 +    val opt = app_rev' thy rrls t
 12.2388 +*)
 12.2389 +"~~~~~ fun app_rev', args:"; val (thy, (Rrls{erls,prepat,scr=Rfuns{normal_form,...},...}), t) =
 12.2390 +  (thy, rrls, t);
 12.2391 +chk_prepat thy erls prepat t    = true;
 12.2392 +(* WN130827: exception Div raised...
 12.2393 +normal_form t
 12.2394 +*)
 12.2395 +(* lookup Rational.thy, cancel_p: normal_form = cancel_p_ thy*)
 12.2396 +"~~~~~ fun cancel_p_, args:"; val (t) = (t);
 12.2397 +val opt = check_fraction t;
 12.2398 +val SOME (numerator, denominator) = opt
 12.2399 +        val vs = TermC.vars_of t
 12.2400 +        val baseT = type_of numerator
 12.2401 +        val expT = HOLogic.realT
 12.2402 +val (SOME a, SOME b) = (poly_of_term vs numerator, poly_of_term vs denominator);
 12.2403 +(*"-------- example 187a": exception Div raised...
 12.2404 +val a = [(12, [1, 1])]: poly
 12.2405 +val b = [(8, [0, 2])]: poly
 12.2406 +              val ((a', b'), c) = gcd_poly a b
 12.2407 +*)
 12.2408 +(* "-------- example 187b": doesn't terminate...
 12.2409 +val a = [(8, [2, 1, 1])]: poly
 12.2410 +val b = [(18, [1, 2, 1])]: poly
 12.2411 +              val ((a', b'), c) = gcd_poly a b
 12.2412 +*)
 12.2413 +(* "-------- example 187c": doesn't terminate...
 12.2414 +val a = [(9, [5, 2, 4])]: poly
 12.2415 +val b = [(15, [6, 3, 1])]: poly
 12.2416 +              val ((a', b'), c) = gcd_poly a b
 12.2417 +*)
 12.2418 +
 12.2419 +\<close> ML \<open> (*poly_of_term \<longrightarrow> NONE*)
 12.2420 +"-------- rls norm_Rational downto fun gcd_poly ------------------------------";
 12.2421 +"-------- rls norm_Rational downto fun gcd_poly ------------------------------";
 12.2422 +"-------- rls norm_Rational downto fun gcd_poly ------------------------------";
 12.2423 +val t = TermC.str2term "(x \<up> 2 - 4)*(3 - y) / ((y \<up> 2 - 9)*(2+x))";
 12.2424 +Rewrite.trace_on := false (*true false*);
 12.2425 +(* trace stops with ...: (and then jEdit hangs)..
 12.2426 +rewrite_set_ thy false norm_Rational t;
 12.2427 +:
 12.2428 +###  rls: cancel_p on: (-12 + 4 * y + 3 * x \<up> 2 + -1 * (x \<up> 2 * y)) /
 12.2429 +(-18 + -9 * x + 2 * y \<up> 2 + x * y \<up> 2)
 12.2430 +*)
 12.2431 +val t = TermC.str2term (*copy from above: "::real" is not required due to " \<up> "*)
 12.2432 +  ("(-12 + 4 * y + 3 * x \<up> 2 + -1 * (x \<up> 2 * y)) /" ^
 12.2433 +  "(-18 + -9 * x + 2 * y \<up> 2 + x * y \<up> 2)");
 12.2434 +(*cancel_p_ thy t;
 12.2435 +exception Div raised*)
 12.2436 +
 12.2437 +"~~~~~ fun cancel_p_, args:"; val (t) = (t);
 12.2438 +val opt = check_fraction t;
 12.2439 +val SOME (numerator, denominator) = opt
 12.2440 +        val vs = TermC.vars_of t
 12.2441 +        val baseT = type_of numerator
 12.2442 +        val expT = HOLogic.realT;
 12.2443 +(*default_print_depth 3; 999*)
 12.2444 +\<close> ML \<open>
 12.2445 +val (NONE, NONE) = (poly_of_term vs numerator, poly_of_term vs denominator)
 12.2446 +\<close> ML \<open>
 12.2447 +\<close> text \<open> (*poly_of_term --> NONE*)
 12.2448 +val (SOME a, SOME b) = (poly_of_term vs numerator, poly_of_term vs denominator);
 12.2449 +(*default_print_depth 3; 999*)
 12.2450 +(* does not terminate instead of returning ?:
 12.2451 +        val ((a', b'), c) = gcd_poly a b
 12.2452 +val a = [(~12, [0, 0]), (3, [2, 0]), (4, [0, 1]), (~1, [2, 1])]: poly
 12.2453 +val b = [(~18, [0, 0]), (~9, [1, 0]), (2, [0, 2]), (1, [1, 2])]: poly
 12.2454 +*)
 12.2455 +
 12.2456 +\<close> ML \<open> (*poly_of_term \<longrightarrow> (SOME, NONE, NONE, SOME*)
 12.2457 +"-------- rls norm_Rational downto fun add_fraction_p_ -----------------------";
 12.2458 +"-------- rls norm_Rational downto fun add_fraction_p_ -----------------------";
 12.2459 +"-------- rls norm_Rational downto fun add_fraction_p_ -----------------------";
 12.2460 +val thy =  @{theory Isac_Knowledge};
 12.2461 +"----- SK060904-2a non-termination of add_fraction_p_";
 12.2462 +val t = TermC.str2term (" (a + b * x) / (a + -1 * (b * x)) +  " ^
 12.2463 +		         " (-1 * a + b * x) / (a + b * x)      ");
 12.2464 +(* rewrite_set_ thy false norm_Rational t
 12.2465 +exception Div raised*)
 12.2466 +(* rewrite_set_ thy false add_fractions_p t;
 12.2467 +exception Div raised*)
 12.2468 +"~~~~~ fun rewrite_set_, args:"; val (thy, bool, rls, term) =
 12.2469 +  (@{theory Isac_Knowledge}, false, add_fractions_p, t);
 12.2470 +"~~~~~ and rewrite__set_, args:"; val (thy, i, _, _, (rrls as Rrls _), t) =
 12.2471 +  (thy, 1, bool, [], rls, term);
 12.2472 +(* app_rev thy (i+1) rrls t;
 12.2473 +exception Div raised*)
 12.2474 +"~~~~~ and app_rev, args:"; val (thy, i, rrls, t) = (thy, (i+1), rrls, t);
 12.2475 +    fun chk_prepat thy erls [] t = true
 12.2476 +      | chk_prepat thy erls prepat t =
 12.2477 +        let
 12.2478 +          fun chk (pres, pat) =
 12.2479 +            (let 
 12.2480 +              val subst: Type.tyenv * Envir.tenv =
 12.2481 +                Pattern.match thy (pat, t) (Vartab.empty, Vartab.empty)
 12.2482 +             in
 12.2483 +              snd (eval__true thy (i + 1) (map (Envir.subst_term subst) pres) [] erls)
 12.2484 +             end) handle Pattern.MATCH => false
 12.2485 +           fun scan_ f [] = false (*scan_ NEVER called by []*)
 12.2486 +             | scan_ f (pp::pps) =
 12.2487 +               if f pp then true else scan_ f pps;
 12.2488 +        in scan_ chk prepat end;
 12.2489 +    (* apply the normal_form of a rev-set *)
 12.2490 +    fun app_rev' thy (Rrls{erls,prepat,scr=Rfuns{normal_form,...},...}) t =
 12.2491 +      if chk_prepat thy erls prepat t
 12.2492 +      then ((*tracing("### app_rev': t = "^UnparseC.term t);*) normal_form t)
 12.2493 +      else NONE;
 12.2494 +(*  val opt = app_rev' thy rrls t;
 12.2495 +exception Div raised*)
 12.2496 +(*  val opt = app_rev' thy rrls t;
 12.2497 +exception Div raised*)
 12.2498 +"~~~~~ and app_rev', args:"; val (thy, (Rrls{erls,prepat,scr=Rfuns{normal_form,...},...}), t) =
 12.2499 +  (thy, rrls, t);
 12.2500 +chk_prepat thy erls prepat t = true       = true;
 12.2501 +(*normal_form t
 12.2502 +exception Div raised*)
 12.2503 +(* lookup Rational.thy, val add_fractions_p: normal_form = add_fraction_p_ thy*)
 12.2504 +(*add_fraction_p_ thy t
 12.2505 +exception Div raised*)
 12.2506 +"~~~~~ fun add_fraction_p_, args:"; val ((_: theory), t) = (thy, t);
 12.2507 +val SOME ((n1, d1), (n2, d2)) = check_frac_sum t;
 12.2508 +UnparseC.term n1; UnparseC.term d1; UnparseC.term n2; UnparseC.term d2;
 12.2509 +      val vs = TermC.vars_of t;
 12.2510 +(*default_print_depth 3; 999*)
 12.2511 +\<close> text \<open> (*poly_of_term \<longrightarrow> (SOME, NONE, NONE, SOME*)
 12.2512 +val (SOME _, SOME a, SOME _, SOME b) =
 12.2513 +  (poly_of_term vs n1, poly_of_term vs d1, poly_of_term vs n2, poly_of_term vs d2);
 12.2514 +(*default_print_depth 3; 999*)
 12.2515 +(*
 12.2516 +val a = [(1, [1, 0, 0]), (~1, [0, 1, 1])]: poly
 12.2517 +val b = [(1, [1, 0, 0]), (1, [0, 1, 1])]: poly
 12.2518 +            val ((a', b'), c) = gcd_poly a b
 12.2519 +*)
 12.2520 +
 12.2521 +\<close> ML \<open>
 12.2522 +"----------- fun check_frac_sum with Free A and Const AA ---------------------------------------";
 12.2523 +"----------- fun check_frac_sum with Free A and Const AA ---------------------------------------";
 12.2524 +"----------- fun check_frac_sum with Free A and Const AA ---------------------------------------";
 12.2525 +val thy = @{theory Isac_Knowledge(*Partial_Fractions*)}
 12.2526 +val ctxt = Proof_Context.init_global thy;
 12.2527 +
 12.2528 +(*---------- (1) with Free A, B  ----------------------------------------------------------------*)
 12.2529 +val t = (the o (parseNEW  ctxt)) "3 = A / 2 + A / 4 + (B / 2 + -1 * B / (2::real))";
 12.2530 +                                (* required for applying thms in rewriting  \<up> ^*)
 12.2531 +(* we get details from here..*)
 12.2532 +
 12.2533 +Rewrite.trace_on := false;
 12.2534 +val SOME (t', _) = Rewrite.rewrite_set_ thy true add_fractions_p t;
 12.2535 +Rewrite.trace_on := false;
 12.2536 +(* Rewrite.trace_on:
 12.2537 +add_fractions_p on: 3 = A / 2 + A / 4 + (B / 2 + -1 * B / 2) --> 3 = A / 2 + A / 4 + 0 / 2 *)
 12.2538 +                     (* |||||||||||||||||||||||||||||||||||| *)
 12.2539 +
 12.2540 +val t = (the o (parseNEW  ctxt))(* ||||||||||||||||||||||||| GUESS 1 GUESS 1 GUESS 1 GUESS 1 *)
 12.2541 +                       "A / 2 + A / 4 + (B / 2 + -1 * B / (2::real))";
 12.2542 +"~~~~~ fun add_fraction_p_ , ad-hoc args:"; val (t) = (t);
 12.2543 +val NONE = (*case*) check_frac_sum t (*of*)
 12.2544 +
 12.2545 +(* Rewrite.trace_on:
 12.2546 +add_fractions_p on: 3 = A / 2 + A / 4 + (B / 2 + -1 * B / 2) --> 3 = A / 2 + A / 4 + 0 / 2 *)
 12.2547 +                     (*         |||||||||||||||||||||||||||| *)
 12.2548 +val t = (the o (parseNEW  ctxt))(* ||||||||||||||||||||||||| GUESS 2 GUESS 2 GUESS 2 GUESS 2 *)
 12.2549 +                               "A / 4 + (B / 2 + -1 * B / (2::real))";
 12.2550 +"~~~~~ fun add_fraction_p_ , ad-hoc args:"; val (t) = (t);
 12.2551 +val SOME ((n1, d1), (n2, d2)) = (*case*) check_frac_sum t (*of*);
 12.2552 +(*+*)if (UnparseC.term n1, UnparseC.term d1) = ("A"                 , "4") andalso
 12.2553 +(*+*)   (UnparseC.term n2, UnparseC.term d2) = ("B / 2 + - 1 * B / 2", "1")
 12.2554 +(*+*)then () else error "check_frac_sum (A / 4 + (B / 2 + -1 * B / (2::real))) changed";
 12.2555 +
 12.2556 +\<close> ML \<open>
 12.2557 +      val vs = TermC.vars_of t;
 12.2558 +val (SOME [(1, [1, 0])], SOME [(4, [0, 0])], NONE, SOME [(1, [0, 0])]) =
 12.2559 +  (*case*) (poly_of_term vs n1, poly_of_term vs d1, poly_of_term vs n2, poly_of_term vs d2) (*of*);
 12.2560 +
 12.2561 +"~~~~~ fun poly_of_term , args:"; val (vs, t) = (vs, n1);
 12.2562 +val SOME [(1, [xxx, 0])] = SOME [monom_of_term vs (1, replicate (length vs) 0) t];
 12.2563 +(*+*)if xxx = 1 then () else error "monom_of_term changed"
 12.2564 +
 12.2565 +"~~~~~ fun monom_of_term , args:"; val (vs, (c, es), (Free (id, _))) =
 12.2566 +  (vs, (1, replicate (length vs) 0), t);
 12.2567 +case vs of [Free ("A", _), Free ("B", _)] =>
 12.2568 +  if c = 1 andalso id = "A"
 12.2569 +  then () else error "monom_of_term Free changed 1"
 12.2570 +| _ => error "monom_of_term Free changed 2";
 12.2571 +
 12.2572 +(*---------- (2) with Const AA, BB --------------------------------------------------------------*)
 12.2573 +val t = (the o (parseNEW  ctxt)) "3 = AA / 2 + AA / 4 + (BB / 2 + -1 * BB / 2)";
 12.2574 +                                    (*AA :: real*)
 12.2575 +(* we get details from here..*)
 12.2576 +
 12.2577 +Rewrite.trace_on := false;
 12.2578 +val SOME (t', _) = Rewrite.rewrite_set_ thy true add_fractions_p t;
 12.2579 +Rewrite.trace_on := false;
 12.2580 +(* Rewrite.trace_on:
 12.2581 +add_fractions_p on: 3 = A / 2 + A / 4 + (B / 2 + -1 * B / 2) --> 3 = A / 2 + A / 4 + 0 / 2 *)
 12.2582 +                     (* |||||||||||||||||||||||||||||||||||| *)
 12.2583 +val t = (the o (parseNEW  ctxt))(* ||||||||||||||||||||||||| *)
 12.2584 +                   "AA / 2 + AA / 4 + (BB / 2 + -1 * BB / 2)";
 12.2585 +"~~~~~ fun add_fraction_p_ , ad-hoc args:"; val (t) = (t);
 12.2586 +val NONE = (*case*) check_frac_sum t (*of*)
 12.2587 +
 12.2588 +(* Rewrite.trace_on:
 12.2589 +add_fractions_p on: 3 = A / 2 + A / 4 + (B / 2 + -1 * B / 2) --> 3 = A / 2 + A / 4 + 0 / 2 *)
 12.2590 +                     (*         |||||||||||||||||||||||||||| *)
 12.2591 +val t = (the o (parseNEW  ctxt))(* ||||||||||||||||||||||||| *)
 12.2592 +                               "AA / 4 + (BB / 2 + -1 * BB / 2)";
 12.2593 +"~~~~~ fun add_fraction_p_ , ad-hoc args:"; val (t) = (t);
 12.2594 +val SOME ((n1, d1), (n2, d2)) = (*case*) check_frac_sum t (*of*);
 12.2595 +(*+*)if (UnparseC.term n1, UnparseC.term d1) = ("AA"                 , "4") andalso
 12.2596 +(*+*)   (UnparseC.term n2, UnparseC.term d2) = ("BB / 2 + - 1 * BB / 2", "1")
 12.2597 +(*+*)then () else error "check_frac_sum (AA / 4 + (BB / 2 + -1 * BB / 2)) changed";
 12.2598 +
 12.2599 +      val vs = TermC.vars_of t;
 12.2600 +val (SOME [(1, [1, 0])], SOME [(4, [0, 0])], NONE, SOME [(1, [0, 0])]) =
 12.2601 +  (*case*) (poly_of_term vs n1, poly_of_term vs d1, poly_of_term vs n2, poly_of_term vs d2) (*of*);
 12.2602 +
 12.2603 +"~~~~~ fun poly_of_term , args:"; val (vs, t) = (vs, n1);
 12.2604 +val SOME [(1, [xxx, 0])] = SOME [monom_of_term vs (1, replicate (length vs) 0) t];
 12.2605 +(*+*)if xxx = 1 then () else error "monom_of_term changed"
 12.2606 +
 12.2607 +"~~~~~ fun monom_of_term , args:"; val (vs, (c, es), (Const (id, _))) =
 12.2608 +  (vs, (1, replicate (length vs) 0), t);
 12.2609 +case vs of [Const ("Partial_Fractions.AA", _), Const ("Partial_Fractions.BB", _)] =>
 12.2610 +  if c = 1 andalso id = "Partial_Fractions.AA"
 12.2611 +  then () else error "monom_of_term Const changed 1"
 12.2612 +| _ => error "monom_of_term Const changed 2";
 12.2613 +
 12.2614 +\<close> ML \<open>
 12.2615 +"----------- fun cancel_p with Const AA --------------------------------------------------------";
 12.2616 +"----------- fun cancel_p with Const AA --------------------------------------------------------";
 12.2617 +"----------- fun cancel_p with Const AA --------------------------------------------------------";
 12.2618 +val thy = @{theory Partial_Fractions};
 12.2619 +val ctxt = Proof_Context.init_global @{theory}
 12.2620 +val SOME t = TermC.parseNEW ctxt "2 * AA / 2"; (* Const ("Free ("AA", "real") *)
 12.2621 +
 12.2622 +val SOME (t', _) = rewrite_set_ thy true cancel_p t;
 12.2623 +case t' of
 12.2624 +  Const ("Rings.divide_class.divide", _) $ Const ("Partial_Fractions.AA", _) $
 12.2625 +    Const ("Groups.one_class.one", _) => ()
 12.2626 +| _ => error "WRONG rewrite_set_ cancel_p (2 * AA / 2) \<longrightarrow> AA changed";
 12.2627 +
 12.2628 +"~~~~~ fun cancel_p , args:"; val (t) = (t);
 12.2629 +val opt = check_fraction t
 12.2630 +val SOME (numerator, denominator) = (*case*) opt (*of*);
 12.2631 +
 12.2632 +if UnparseC.term numerator = "2 * AA" andalso UnparseC.term denominator = "2"
 12.2633 +then () else error "check_fraction (2 * AA / 2) changed";
 12.2634 +        val vs = TermC.vars_of t;
 12.2635 +case vs of
 12.2636 +  [Const ("Partial_Fractions.AA", _)] => ()
 12.2637 +| _ => error "rewrite_set_ cancel_p (2 * AA / 2) \<longrightarrow> AA/1  changed";
 12.2638 +
 12.2639 +
 12.2640 +\<close> ML \<open>
 12.2641 +"-------- rewrite_set_ cancel_p from: Mathematik 1 Schalk Reniets Verlag -----";
 12.2642 +"-------- rewrite_set_ cancel_p from: Mathematik 1 Schalk Reniets Verlag -----";
 12.2643 +"-------- rewrite_set_ cancel_p from: Mathematik 1 Schalk Reniets Verlag -----";
 12.2644 +val thy  = @{theory "Rational"};
 12.2645 +"-------- WN";
 12.2646 +val t = TermC.str2term "(2 + -3 * x) / 9";
 12.2647 +if NONE = rewrite_set_ thy false cancel_p t then ()
 12.2648 +else error "rewrite_set_ cancel_p must return NONE, if the term cannot be cancelled";
 12.2649 +
 12.2650 +"-------- example 186a";
 12.2651 +val t = TermC.str2term "(14 * x * y) / (x * y)";
 12.2652 +  is_expanded (TermC.str2term "14 * x * y");
 12.2653 +  is_expanded (TermC.str2term "x * y");
 12.2654 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2655 +if (UnparseC.term t', UnparseC.terms asm) = ("14 / 1", "[]")
 12.2656 +then () else error "rational.sml cancel Schalk 186a";
 12.2657 +
 12.2658 +"-------- example 186b";
 12.2659 +val t = TermC.str2term "(60 * a * b) / ( 15 * a  * b )";
 12.2660 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2661 +if (UnparseC.term t', UnparseC.terms asm) = ("4 / 1", "[]")
 12.2662 +then () else error "rational.sml cancel Schalk 186b";
 12.2663 +
 12.2664 +"-------- example 186c";
 12.2665 +val t = TermC.str2term "(144 * a \<up> 2 * b * c) / (12 * a * b * c)";
 12.2666 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2667 +if (UnparseC.term t', UnparseC.terms asm) = ("12 * a / 1", "[]")
 12.2668 +then () else error "rational.sml cancel Schalk 186c";
 12.2669 +
 12.2670 +(* !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!! exception Div raised !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
 12.2671 +  see --- fun rewrite_set_ downto fun gcd_poly ---
 12.2672 +"-------- example 187a";
 12.2673 +val t = TermC.str2term "(12 * x * y) / (8 * y \<up> 2 )";
 12.2674 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2675 +if (UnparseC.term t', UnparseC.terms asm) = ("3 * x / (2 * y)", "[\"4 * y ~= 0\"]")
 12.2676 +then () else error "rational.sml cancel Schalk 187a";
 12.2677 +*)
 12.2678 +
 12.2679 +(* doesn't terminate !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
 12.2680 +  see --- fun rewrite_set_ downto fun gcd_poly ---
 12.2681 +"-------- example 187b";
 12.2682 +val t = TermC.str2term "(8 * x \<up> 2 * y * z ) / (18 * x * y \<up> 2 * z )";
 12.2683 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2684 +if (UnparseC.term t', UnparseC.terms asm) = ("4 * x / (9 * y)", "[\"2 * (z * (y * x)) ~= 0\"]")
 12.2685 +then () else error "rational.sml cancel Schalk 187b";
 12.2686 +*)
 12.2687 +
 12.2688 +(* doesn't terminate !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
 12.2689 +  see --- fun rewrite_set_ downto fun gcd_poly ---
 12.2690 +"-------- example 187c";
 12.2691 +val t = TermC.str2term "(9 * x \<up> 5 * y \<up> 2 * z \<up> 4) / (15 * x \<up> 6 * y \<up> 3 * z )";
 12.2692 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2693 +if (UnparseC.term t', UnparseC.terms asm) = 
 12.2694 +  ("3 * z \<up> 3 / (5 * (y * x))", "[\"3 * (z * (y \<up> 2 * x \<up> 5)) ~= 0\"]") 
 12.2695 +then () else error "rational.sml cancel Schalk 187c";
 12.2696 +*)
 12.2697 +
 12.2698 +\<close> text \<open> (*rewrite_set_ cancel_p \<longrightarrow> NONE*)
 12.2699 +"-------- example 188a";
 12.2700 +val t = TermC.str2term "(-8 + 8 * x) / (-9 + 9 * x)";
 12.2701 +  is_expanded (TermC.str2term "8 * x + -8");
 12.2702 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2703 +if (UnparseC.term t', UnparseC.terms asm) = ("8 / 9", "[]")
 12.2704 +then () else error "rational.sml cancel Schalk 188a";
 12.2705 +
 12.2706 +val t = TermC.str2term "(8*((-1) + x))/(9*((-1) + x))";
 12.2707 +val SOME (t, _) = rewrite_set_ thy false make_polynomial t;
 12.2708 +if (UnparseC.term t', UnparseC.terms asm) = ("8 / 9", "[]")
 12.2709 +then () else error "rational.sml cancel Schalk make_polynomial 1";
 12.2710 +
 12.2711 +\<close> text \<open> (*rewrite_set_ cancel_p \<longrightarrow> NONE*)
 12.2712 +"-------- example 188b";
 12.2713 +val t = TermC.str2term "(-15 + 5 * x) / (-18 + 6 * x)";
 12.2714 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2715 +if (UnparseC.term t', UnparseC.terms asm) = ("5 / 6", "[]")
 12.2716 +then () else error "rational.sml cancel Schalk 188b";
 12.2717 +
 12.2718 +\<close> text \<open> (*rewrite_set_ cancel_p \<longrightarrow> NONE*)
 12.2719 +"-------- example 188c";
 12.2720 +val t = TermC.str2term "(a + -1 * b) / (b + -1 * a)";
 12.2721 +val SOME (t', asm) = rewrite_set_ thy false  cancel_p t;
 12.2722 +if (UnparseC.term t', UnparseC.terms asm) = ("-1 / 1", "[]")
 12.2723 +then () else error "rational.sml cancel Schalk 188c";
 12.2724 +
 12.2725 +is_expanded (TermC.str2term "a + -1 * b") = true;
 12.2726 +val t = TermC.str2term "((-1)*(b + (-1) * a))/(1*(b + (-1) * a))";
 12.2727 +val SOME (t', asm) = rewrite_set_ thy false make_polynomial t;
 12.2728 +if (UnparseC.term t', UnparseC.terms asm) = ("(a + -1 * b) / (-1 * a + b)", "[]")
 12.2729 +then () else error "rational.sml cancel Schalk make_polynomial 2";
 12.2730 +
 12.2731 +\<close> ML \<open>
 12.2732 +"-------- example 190a";
 12.2733 +val t = TermC.str2term "( 27 * a \<up> 3 + 9 * a \<up> 2 + 3 * a + 1 ) / ( 27 * a \<up> 3 + 18 * a \<up> 2 + 3 * a )";
 12.2734 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2735 +if (UnparseC.term t', UnparseC.terms asm) = 
 12.2736 +  ("(1 + 9 * a \<up> 2) / (3 * a + 9 * a \<up> 2)", "[\"3 * a + 9 * a \<up> 2 \<noteq> 0\"]")
 12.2737 +then () else error "rational.sml cancel Schalk 190a";
 12.2738 +
 12.2739 +"-------- example 190c";
 12.2740 +val t = TermC.str2term "((1 + 9 * a \<up> 2)*(1 + 3 * a))/((3 * a + 9 * a \<up> 2)*(1 + 3 * a))";
 12.2741 +val SOME (t', asm) = rewrite_set_ thy false make_polynomial t;
 12.2742 +if (UnparseC.term t', UnparseC.terms asm) = 
 12.2743 +  ("(1 + 3 * a + 9 * a \<up> 2 + 27 * a \<up> 3) /\n(3 * a + 18 * a \<up> 2 + 27 * a \<up> 3)", "[]")
 12.2744 +then () else error "rational.sml make_polynomial Schalk 190c";
 12.2745 +
 12.2746 +\<close> text \<open> (*rewrite_set_ cancel_p \<longrightarrow> NONE*)
 12.2747 +"-------- example 191a";
 12.2748 +val t = TermC.str2term "( x \<up> 2 + -1 * y \<up> 2 ) / ( x + y )";
 12.2749 +  is_expanded (TermC.str2term "x \<up> 2 + -1 * y \<up> 2") = false; (*!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!*)
 12.2750 +  is_expanded (TermC.str2term "x + y") = true;
 12.2751 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2752 +if (UnparseC.term t', UnparseC.terms asm) = ("(x + -1 * y) / 1", "[]")
 12.2753 +then () else error "rational.sml make_polynomial Schalk 191a";
 12.2754 +
 12.2755 +\<close> text \<open> (*rewrite_set_ make_polynomial \<longrightarrow> NONE*)
 12.2756 +"-------- example 191b";
 12.2757 +val t = TermC.str2term "((x + (-1) * y)*(x + y))/((1)*(x + y))";
 12.2758 +val SOME (t', asm) = rewrite_set_ thy false make_polynomial t;
 12.2759 +if (UnparseC.term t', UnparseC.terms asm) = ("(x \<up> 2 + -1 * y \<up> 2) / (x + y)", "[]")
 12.2760 +then () else error "rational.sml make_polynomial Schalk 191b";
 12.2761 +
 12.2762 +\<close> ML \<open>
 12.2763 +"-------- example 191c";
 12.2764 +val t = TermC.str2term "( 9 * x \<up> 2 + -30 * x + 25 ) / ( 9 * x \<up> 2 + -25 )";
 12.2765 +  is_expanded (TermC.str2term "9 * x \<up> 2 + -30 * x + 25") = true;
 12.2766 +  is_expanded (TermC.str2term "25 + -30*x + 9*x \<up> 2") = true;
 12.2767 +  is_expanded (TermC.str2term "-25 + 9*x \<up> 2") = true;
 12.2768 +
 12.2769 +val t = TermC.str2term "(((-5) + 3 * x)*((-5) + 3 * x))/((5 + 3 * x)*((-5) + 3 * x))";
 12.2770 +val SOME (t', asm) = rewrite_set_ thy false make_polynomial t;
 12.2771 +if (UnparseC.term t', UnparseC.terms asm) = ("(25 + - 30 * x + 9 * x \<up> 2) / (- 25 + 9 * x \<up> 2)", "[]")
 12.2772 +then () else error "rational.sml make_polynomial Schalk 191c";
 12.2773 +
 12.2774 +\<close> text \<open> (*rewrite_set_ cancel_p \<longrightarrow> NONE*)
 12.2775 +"-------- example 192b";
 12.2776 +val t = TermC.str2term "( 7 * x \<up> 3 + -1 * x \<up> 2 * y ) / ( 7 * x * y \<up> 2 + -1 *  y \<up> 3 )";
 12.2777 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2778 +if (UnparseC.term t', UnparseC.terms asm) = ("x \<up> 2 / y \<up> 2", "[\"y \<up> 2 \<noteq> 0\"]")
 12.2779 +then () else error "rational.sml cancel_p Schalk 192b";
 12.2780 +
 12.2781 +\<close> text \<open> (*rewrite_set_ make_polynomial \<longrightarrow> NONE*)
 12.2782 +val t = TermC.str2term "((x \<up> 2)*(7 * x + (-1) * y))/((y \<up> 2)*(7 * x + (-1) * y))";
 12.2783 +val SOME (t', asm) = rewrite_set_ thy false make_polynomial t;
 12.2784 +if (UnparseC.term t', UnparseC.terms asm) = 
 12.2785 +  ("(7 * x \<up> 3 + -1 * x \<up> 2 * y) / (7 * x * y \<up> 2 + -1 * y \<up> 3)", "[]")
 12.2786 +then () else error "rational.sml make_polynomial Schalk 192b";
 12.2787 +
 12.2788 +\<close> text \<open> (*rewrite_set_ make_polynomial \<longrightarrow> NONE*)
 12.2789 +val t = TermC.str2term "((x \<up> 2)*(7 * x + (-1) * y))/((y \<up> 2)*(7 * x + (-1) * y))";
 12.2790 +val SOME (t', asm) = rewrite_set_ thy false make_polynomial t;
 12.2791 +if (UnparseC.term t', UnparseC.terms asm) = 
 12.2792 +  ("(7 * x \<up> 3 + -1 * x \<up> 2 * y) / (7 * x * y \<up> 2 + -1 * y \<up> 3)", "[]")
 12.2793 +then () else error "rational.sml make_polynomial Schalk WN050929 not working";
 12.2794 +
 12.2795 +\<close> ML \<open>
 12.2796 +\<close> text \<open> (*rewrite_set_ cancel_p \<longrightarrow> NONE*)
 12.2797 +val t = TermC.str2term "( x \<up> 2 + -6 * x + 9 ) / ( x \<up> 2 + -9 )";
 12.2798 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2799 +if (UnparseC.term t', UnparseC.terms asm) = ("(-3 + x) / (3 + x)", "[\"3 + x \<noteq> 0\"]")
 12.2800 +then () else error "rational.sml cancel_p Schalk 193a";
 12.2801 +
 12.2802 +\<close> text \<open> (*rewrite_set_ cancel_p \<longrightarrow> NONE*)
 12.2803 +"-------- example 193b";
 12.2804 +val t = TermC.str2term "( x \<up> 2 + -8 * x + 16 ) / ( 2 * x \<up> 2 + -32 )";
 12.2805 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2806 +if (UnparseC.term t', UnparseC.terms asm) = ("(-4 + x) / (8 + 2 * x)", "[\"8 + 2 * x \<noteq> 0\"]")
 12.2807 +then () else error "rational.sml cancel_p Schalk 193b";
 12.2808 +
 12.2809 +\<close> text \<open> (*rewrite_set_ cancel_p \<longrightarrow> NONE*)
 12.2810 +"-------- example 193c";
 12.2811 +val t = TermC.str2term "( 2 * x + -50 * x \<up> 3 ) / ( 25 * x \<up> 2 + -10 * x + 1 )";
 12.2812 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2813 +if (UnparseC.term t', UnparseC.terms asm) = 
 12.2814 +  ("(2 * x + 10 * x \<up> 2) / (1 + -5 * x)", "[\"1 + -5 * x \<noteq> 0\"]")
 12.2815 +then () else error "rational.sml cancel_p Schalk 193c";
 12.2816 +
 12.2817 +\<close> text \<open> (*rewrite_set_ cancel_p \<longrightarrow> NONE*)
 12.2818 +(*WN:*)
 12.2819 +val t = TermC.str2term "(-25 + 9*x \<up> 2)/(5 + 3*x)";
 12.2820 +val SOME (t, asm) = rewrite_set_ thy false cancel_p t;
 12.2821 +if (UnparseC.term t', UnparseC.terms asm) = ("(2 * x + 10 * x \<up> 2) / (1 + -5 * x)", "[]")
 12.2822 +then () else error "rational.sml cancel WN 1";
 12.2823 +
 12.2824 +\<close> ML \<open>
 12.2825 +"-------- example heuberger";
 12.2826 +val t = TermC.str2term ("(x \<up> 4 + x * y + x \<up> 3 * y + y \<up> 2) / " ^
 12.2827 +  "(x + 5 * x \<up> 2 + y + 5 * x * y + x \<up> 2 * y \<up> 3 + x * y \<up> 4)");
 12.2828 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2829 +if (UnparseC.term t', UnparseC.terms asm) = 
 12.2830 +  ("(x \<up> 3 + y) / (1 + 5 * x + x * y \<up> 3)", "[\"1 + 5 * x + x * y \<up> 3 \<noteq> 0\"]")
 12.2831 +then () else error "rational.sml cancel_p heuberger";
 12.2832 +
 12.2833 +\<close> ML \<open>
 12.2834 +"-------- rewrite_set_ add_fractions_p from: Mathematik 1 Schalk -------------";
 12.2835 +"-------- rewrite_set_ add_fractions_p from: Mathematik 1 Schalk -------------";
 12.2836 +"-------- rewrite_set_ add_fractions_p from: Mathematik 1 Schalk -------------";
 12.2837 +(*deleted example 204 ... 236b at update Isabelle2012-->2013*)
 12.2838 +
 12.2839 +"-------- integration lev.1 -- lev.5: cancel_p_ & add_fractions_p_ -----------";
 12.2840 +"-------- integration lev.1 -- lev.5: cancel_p_ & add_fractions_p_ -----------";
 12.2841 +"-------- integration lev.1 -- lev.5: cancel_p_ & add_fractions_p_ -----------";
 12.2842 +val t = TermC.str2term ("123 = (a*x)/(b*x) + (c*x)/(d*x) + (e*x)/(f*x::real)");
 12.2843 +"-------- gcd_poly integration level 1: works on exact term";
 12.2844 +if NONE = cancel_p_ thy t then () else error "cancel_p_ works on exact fraction";
 12.2845 +if NONE = add_fraction_p_ thy t then () else error "add_fraction_p_ works on exact fraction";
 12.2846 +
 12.2847 +"-------- gcd_poly integration level 2: picks out ONE appropriate subterm";
 12.2848 +val SOME (t', asm) = rewrite_set_ thy false cancel_p t;
 12.2849 +if UnparseC.term t' = "123 = a * x / (b * x) + c * x / (d * x) + e / f" 
 12.2850 +then () else error "level 2, rewrite_set_ cancel_p: changed";
 12.2851 +val SOME (t', asm) = rewrite_set_ thy false add_fractions_p t;
 12.2852 +if UnparseC.term t' = "123 = (b * c * x + a * d * x) / (b * d * x) + e * x / (f * x)"
 12.2853 +then () else error "level 2, rewrite_set_ add_fractions_p: changed";
 12.2854 +
 12.2855 +"-------- gcd_poly integration level 3: rewrites all appropriate subterms";
 12.2856 +val SOME (t', asm) = rewrite_set_ thy false cancel_p_rls t;
 12.2857 +if UnparseC.term t' = "123 = a / b + c / d + e / f"
 12.2858 +then () else error "level 3, rewrite_set_ cancel_p_rls: changed";
 12.2859 +val SOME (t', asm) = rewrite_set_ thy false add_fractions_p_rls t; (*CREATE add_fractions_p_rls*)
 12.2860 +if UnparseC.term t' = "123 = (b * d * e * x + b * c * f * x + a * d * f * x) / (b * d * f * x)"
 12.2861 +then () else error "level 3, rewrite_set_ add_fractions_p_rls: changed";
 12.2862 +
 12.2863 +"-------- gcd_poly integration level 4: iteration cancel_p -- add_fraction_p";
 12.2864 +(* simpler variant *)
 12.2865 +val testrls = Rule_Set.append_rules "testrls" Rule_Set.empty [Rls_ cancel_p, Rls_ add_fractions_p]
 12.2866 +val SOME (t', asm) = rewrite_set_ thy false testrls t;
 12.2867 +(*Rewrite.trace_on := false;
 12.2868 +#  rls: testrls on: 123 = a * x / (b * x) + c * x / (d * x) + e * x / (f * x) 
 12.2869 +##  rls: cancel_p on: 123 = a * x / (b * x) + c * x / (d * x) + e * x / (f * x) 
 12.2870 +##  rls: add_fractions_p on: 123 = a * x / (b * x) + c * x / (d * x) + e / f 
 12.2871 +##  rls: cancel_p on: 123 = (b * c * x + a * d * x) / (b * d * x) + e / f 
 12.2872 +##  rls: add_fractions_p on: 123 = (b * c + a * d) / (b * d) + e / f 
 12.2873 +##  rls: cancel_p on: 123 = (b * d * e + b * c * f + a * d * f) / (b * d * f) 
 12.2874 +##  rls: add_fractions_p on: 123 = (b * d * e + b * c * f + a * d * f) / (b * d * f) *)
 12.2875 +if UnparseC.term t' = "123 = (b * d * e + b * c * f + a * d * f) / (b * d * f)"
 12.2876 +then () else error "level 4, rewrite_set_ *_p: changed";
 12.2877 +
 12.2878 +(* complicated variant *)
 12.2879 +val testrls_rls = Rule_Set.append_rules "testrls_rls" Rule_Set.empty [Rls_ cancel_p_rls, Rls_ add_fractions_p_rls];
 12.2880 +val SOME (t', asm) = rewrite_set_ thy false testrls_rls t;
 12.2881 +(*#  rls: testrls_rls on: 123 = a * x / (b * x) + c * x / (d * x) + e * x / (f * x) 
 12.2882 +##  rls: cancel_p_rls on: 123 = a * x / (b * x) + c * x / (d * x) + e * x / (f * x) 
 12.2883 +###  rls: cancel_p on: 123 = a * x / (b * x) + c * x / (d * x) + e * x / (f * x) 
 12.2884 +###  rls: cancel_p on: 123 = a * x / (b * x) + c * x / (d * x) + e / f 
 12.2885 +###  rls: cancel_p on: 123 = a * x / (b * x) + c / d + e / f 
 12.2886 +###  rls: cancel_p on: 123 = a / b + c / d + e / f 
 12.2887 +##  rls: add_fractions_p_rls on: 123 = a / b + c / d + e / f 
 12.2888 +###  rls: add_fractions_p on: 123 = a / b + c / d + e / f 
 12.2889 +###  rls: add_fractions_p on: 123 = (b * c + a * d) / (b * d) + e / f 
 12.2890 +###  rls: add_fractions_p on: 123 = (b * d * e + b * c * f + a * d * f) / (b * d * f) 
 12.2891 +##  rls: cancel_p_rls on: 123 = (b * d * e + b * c * f + a * d * f) / (b * d * f) 
 12.2892 +###  rls: cancel_p on: 123 = (b * d * e + b * c * f + a * d * f) / (b * d * f) 
 12.2893 +##  rls: add_fractions_p_rls on: 123 = (b * d * e + b * c * f + a * d * f) / (b * d * f) 
 12.2894 +###  rls: add_fractions_p on: 123 = (b * d * e + b * c * f + a * d * f) / (b * d * f) *)
 12.2895 +if UnparseC.term t' = "123 = (b * d * e + b * c * f + a * d * f) / (b * d * f)"
 12.2896 +then () else error "level 4, rewrite_set_ *_p_rls: changed"
 12.2897 +
 12.2898 +"-------- gcd_poly integration level 5: cancel_p & add_fraction_p within norm_Rational";
 12.2899 +val SOME (t', asm) = rewrite_set_ thy false norm_Rational t;
 12.2900 +if UnparseC.term t' = "123 = (a * d * f + b * c * f + b * d * e) / (b * d * f)"
 12.2901 +then () else error "level 5, rewrite_set_ norm_Rational: changed"
 12.2902 +
 12.2903 +\<close> ML \<open>
 12.2904 +"-------- reverse rewrite ----------------------------------------------------";
 12.2905 +"-------- reverse rewrite ----------------------------------------------------";
 12.2906 +"-------- reverse rewrite ----------------------------------------------------";
 12.2907 +(** the term for which reverse rewriting is demonstrated **)
 12.2908 +val t = TermC.str2term "(9 + -1 * x \<up> 2) / (9 + 6 * x + x \<up> 2)";
 12.2909 +val Rrls {scr = Rfuns {init_state = ini, locate_rule = loc,
 12.2910 +  next_rule = nex, normal_form = nor, ...},...} = cancel_p;
 12.2911 +
 12.2912 +\<close> text \<open> (*Rfuns normal_form \<longrightarrow> NONE*)
 12.2913 +(** normal_form produces the result in ONE step **)
 12.2914 +  val SOME (t', _) = nor t;
 12.2915 +if UnparseC.term t' = "(3 + -1 * x) / (3 + x)" then ()
 12.2916 +else error "rational.sml normal_form (9 - x \<up> 2) / (9 - 6 * x + x \<up> 2)";
 12.2917 +
 12.2918 +(** initialize the interpreter state used by the 'me' **)
 12.2919 +  val (t, _, revsets, _) = ini t;
 12.2920 +
 12.2921 +if length (hd revsets) = 11 then () else error "length of revset changed";
 12.2922 +if (revsets |> nth 1 |> nth 1 |> id_of_thm) = 
 12.2923 +  (@ {thm realpow_twoI} |> Thm.get_name_hint |> ThmC.cut_id)
 12.2924 +then () else error "first element of revset changed";
 12.2925 +if
 12.2926 +(revsets |> nth 1 |> nth 1 |> Rule.to_string) = "Thm (\"realpow_twoI\",?r1 \<up> 2 = ?r1 * ?r1)" andalso
 12.2927 +(revsets |> nth 1 |> nth 2 |> Rule.to_string) = "Thm (\"#: 9 = 3 \<up> 2\",9 = 3 \<up> 2)" andalso
 12.2928 +(revsets |> nth 1 |> nth 3 |> Rule.to_string) = "Thm (\"#: 6 * x = 2 * (3 * x)\",6 * x = 2 * (3 * x))" 
 12.2929 +andalso
 12.2930 +(revsets |> nth 1 |> nth 4 |> Rule.to_string) = "Thm (\"#: -3 * x = -1 * (3 * x)\",-3 * x = -1 * (3 * x))" 
 12.2931 +andalso
 12.2932 +(revsets |> nth 1 |> nth 5 |> Rule.to_string) = "Thm (\"#: 9 = 3 * 3\",9 = 3 * 3)" andalso
 12.2933 +(revsets |> nth 1 |> nth 6 |> Rule.to_string) = "Rls_ (\"sym_order_mult_rls_\")" andalso
 12.2934 +(revsets |> nth 1 |> nth 7 |> Rule.to_string) = 
 12.2935 +  "Thm (\"sym_mult.assoc\",?a * (?b * ?c) = ?a * ?b * ?c)"
 12.2936 +then () else error "first 7 elements in revset changed"
 12.2937 +
 12.2938 +\<close> ML \<open>
 12.2939 +(** find the rule 'r' to apply to term 't' **)
 12.2940 +(*/------- WN1309: since cancel_ (accepted "-" between monomials) has been replaced by cancel_p_ 
 12.2941 +  for Isabelle2013, we don't get a working revset, but non-termination:
 12.2942 +
 12.2943 +  val SOME (r as (Thm (str, thm))) = nex revsets t;
 12.2944 +  :
 12.2945 +((3 * 3 + -1 * x * x) / (3 * 3 + 2 * 3 * x + x * x), 
 12.2946 +  Rls_ ("sym_order_mult_rls_"), ((3 * 3 + -1 * (x * x)) / (3 * 3 + 2 * (3 * x) + x * x), []))", "
 12.2947 +((3 * 3 + -1 * (x * x)) / (3 * 3 + 2 * (3 * x) + x * x), 
 12.2948 +  Thm ("sym_mult.assoc", ""), ((3 * 3 + -1 * (x * x)) / (3 * 3 + 2 * 3 * x + x * x), []))", "
 12.2949 +((3 * 3 + -1 * (x * x)) / (3 * 3 + 2 * 3 * x + x * x), 
 12.2950 +  Thm ("sym_mult.assoc", ""), ((3 * 3 + -1 * x * x) / (3 * 3 + 2 * 3 * x + x * x), []))", "
 12.2951 +((3 * 3 + -1 * x * x) / (3 * 3 + 2 * 3 * x + x * x), Rls_ ("sym_order_mult_rls_"), ((3 * 3 + -1 * (x * x)) / (3 * 3 + 2 * (3 * x) + x * x), []))", "
 12.2952 + :
 12.2953 +### Isabelle2002:
 12.2954 +  Thm ("sym_#mult_2_3", "6 = 2 * 3")
 12.2955 +### Isabelle2009-2 for cancel_ (not cancel_p_):
 12.2956 +if str = "sym_#power_Float ((3,0), (0,0)) __ ((2,0), (0,0))"
 12.2957 +   andalso ThmC.string_of_thm thm = 
 12.2958 +           (string_of_thm (Thm.make_thm @{theory "Isac_Knowledge"}
 12.2959 +               (Trueprop $ (Thm.term_of o the o (TermC.parse thy)) "9 = 3 \<up> 2"))) then ()
 12.2960 +else error "rational.sml next_rule (9 - x \<up> 2) / (9 - 6 * x + x \<up> 2)";
 12.2961 +\---------------------------------------------------------------------------------------/*)
 12.2962 +
 12.2963 +(** check, if the rule 'r' applied by the user to 't' belongs to the ruleset;
 12.2964 +  if the rule is OK, the term resulting from applying the rule is returned,too;
 12.2965 +  there might be several rule applications inbetween,
 12.2966 +  which are listed after the head in reverse order **)
 12.2967 +(*/-------------------------------------------- Isabelle2013: this gives "error id_of_thm";
 12.2968 +  we don't repair this, because interaction within "reverse rewriting" never worked properly:
 12.2969 +
 12.2970 +  val (r, (t, asm))::_ = loc revsets t r;
 12.2971 +if UnparseC.term t = "(9 - x \<up> 2) / (3 \<up> 2 + 6 * x + x \<up> 2)" andalso asm = []
 12.2972 +then () else error "rational.sml locate_rule (9 - x \<up> 2) / (9 - 6 * x + x \<up> 2)";
 12.2973 +
 12.2974 +(* find the next rule to apply *)
 12.2975 +  val SOME (r as (Thm (str, thm))) = nex revsets t;
 12.2976 +if str = "sym_#power_Float ((3,0), (0,0)) __ ((2,0), (0,0))" andalso
 12.2977 +   ThmC.string_of_thm thm = (string_of_thm (ThmC_Def.make_thm @{theory "Isac_Knowledge"}
 12.2978 +                (Trueprop $ (Thm.term_of o the o (TermC.parse thy)) "9 = 3 \<up> 2"))) then ()
 12.2979 +else error "rational.sml next_rule (9 - x \<up> 2) / (9 - 6 * x + x \<up> 2)";
 12.2980 +
 12.2981 +(*check the next rule*)
 12.2982 +  val (r, (t, asm)) :: _ = loc revsets t r;
 12.2983 +if UnparseC.term t = "(3 \<up> 2 - x \<up> 2) / (3 \<up> 2 + 6 * x + x \<up> 2)" then ()
 12.2984 +else error "rational.sml locate_rule (9 - x \<up> 2) / (9 - 6 * x + x \<up> 2) II";
 12.2985 +
 12.2986 +(*find and check the next rules, rewrite*)
 12.2987 +  val SOME r = nex revsets t;
 12.2988 +  val (r,(t,asm))::_ = loc revsets t r;
 12.2989 +if UnparseC.term t = "(3 \<up> 2 - x \<up> 2) / (3 \<up> 2 + 2 * 3 * x + x \<up> 2)" then ()
 12.2990 +else error "rational.sml locate_rule II";
 12.2991 +
 12.2992 +  val SOME r = nex revsets t;
 12.2993 +  val (r,(t,asm))::_ = loc revsets t r;
 12.2994 +if UnparseC.term t = "(3 - x) * (3 + x) / (3 \<up> 2 + 2 * 3 * x + x \<up> 2)" then ()
 12.2995 +else error "rational.sml next_rule II";
 12.2996 +
 12.2997 +  val SOME r = nex revsets t;
 12.2998 +  val (r,(t,asm))::_ = loc revsets t r;
 12.2999 +if UnparseC.term t = "(3 - x) * (3 + x) / ((3 + x) * (3 + x))" then ()
 12.3000 +else error "rational.sml next_rule III";
 12.3001 +
 12.3002 +  val SOME r = nex revsets t;
 12.3003 +  val (r, (t, asm)) :: _ = loc revsets t r;
 12.3004 +  val ss = UnparseC.term t;
 12.3005 +if ss = "(3 - x) / (3 + x)" andalso UnparseC.terms asm = "[\"3 + x ~= 0\"]" then ()
 12.3006 +else error "rational.sml: new behav. in rev-set cancel";
 12.3007 +\--------------------------------------------------------------------------------------/*)
 12.3008 +
 12.3009 +\<close> ML \<open>
 12.3010 +"-------- 'reverse-ruleset' cancel_p -----------------------------------------";
 12.3011 +"-------- 'reverse-ruleset' cancel_p -----------------------------------------";
 12.3012 +"-------- 'reverse-ruleset' cancel_p -----------------------------------------";
 12.3013 +(*WN130909: the example below shows, why "reverse rewriting" only worked for
 12.3014 +  special cases.*)
 12.3015 +
 12.3016 +(*the term for which reverse rewriting is demonstrated*)
 12.3017 +val t = TermC.str2term "(9 + (-1)*x \<up> 2) / (9 + ((-6)*x + x \<up> 2))";
 12.3018 +val Rrls {scr=Rfuns {init_state=ini,locate_rule=loc,
 12.3019 +		       next_rule=nex,normal_form=nor,...},...} = cancel_p;
 12.3020 +
 12.3021 +(*normal_form produces the result in ONE step*)
 12.3022 +\<close> text \<open> (*Rfuns normal_form \<longrightarrow> NONE*)
 12.3023 +val SOME (t',_) = nor t; 
 12.3024 +UnparseC.term t' = "(3 + 1 * x) / (3 + -1 * x)";
 12.3025 +
 12.3026 +(*initialize the interpreter state used by the 'me'*)
 12.3027 +val SOME (t', asm) = cancel_p_ thy t;
 12.3028 +UnparseC.term t' = "(3 + x) / (3 + -1 * x)" (*true*);
 12.3029 +UnparseC.terms asm = "[\"3 + -1 * x ~= 0\"]" (*true*);
 12.3030 +val (t,_,revsets,_) = ini t;
 12.3031 +
 12.3032 +(* WN.10.10.02: dieser Fall terminiert nicht 
 12.3033 +           (make_polynomial enth"alt zu viele rules)
 12.3034 +WN060823 'init_state' requires rewriting on specified location in the term
 12.3035 +default_print_depth 99; Rfuns; default_print_depth 3;
 12.3036 +WN060831 cycling "sym_order_mult_rls_" "sym_mult.assoc"
 12.3037 +         as was with make_polynomial before ?!?*)
 12.3038 +
 12.3039 +val SOME r = nex revsets t;
 12.3040 +eq_Thm (r, Thm ("sym_#power_Float ((3,0), (0,0)) __ ((2,0), (0,0))", 
 12.3041 +		mk_thm thy "9 = 3 \<up> 2"));
 12.3042 +(*WN060831 *** id_of_thm
 12.3043 +           Exception- ERROR raised ...
 12.3044 +val (r,(t,asm))::_ = loc revsets t r;
 12.3045 +UnparseC.term t;
 12.3046 +
 12.3047 +  val SOME r = nex revsets t;
 12.3048 +  val (r,(t,asm))::_ = loc revsets t r;
 12.3049 +  UnparseC.term t;
 12.3050 +*)                    
 12.3051 +
 12.3052 +\<close> ML \<open>
 12.3053 +"-------- examples: rls norm_Rational ----------------------------------------";
 12.3054 +"-------- examples: rls norm_Rational ----------------------------------------";
 12.3055 +"-------- examples: rls norm_Rational ----------------------------------------";
 12.3056 +val t = TermC.str2term "(3*x+5)/18 - x/2  - -(3*x - 2)/9 = 0";
 12.3057 +val SOME (t',_) = rewrite_set_ thy false norm_Rational t; UnparseC.term t';
 12.3058  \<close> ML \<open>
 12.3059  UnparseC.term t'
 12.3060 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation*)
 12.3061 +if UnparseC.term t' = "1 / 18 = 0" then () else error "rational.sml 1";
 12.3062 +
 12.3063  \<close> ML \<open>
 12.3064 -if UnparseC.term t' = "x * x \<up> 2" then ()
 12.3065 -else error "rewrite.sml Poly.is_multUnordered doesn't work";
 12.3066 -
 12.3067 -(* for achieving the previous result, the following code was taken apart *)
 12.3068 -"----- rewrite__set_ ---";
 12.3069 -val (thy, i, _, _, (rrls as Rrls _), t) = (thy, 0, true, [], order_mult_, tm);
 12.3070 -	val (t', asm, rew) = app_rev thy (i+1) rrls t;
 12.3071 -"----- app_rev ---";
 12.3072 -val (thy, i, rrls, t) = (thy, (i+1), rrls, t);
 12.3073 -	fun chk_prepat thy erls [] t = true
 12.3074 -	  | chk_prepat thy erls prepat t =
 12.3075 -	    let fun chk (pres, pat) =
 12.3076 -		    (let val subst: Type.tyenv * Envir.tenv = 
 12.3077 -			     Pattern.match thy (pat, t)
 12.3078 -					    (Vartab.empty, Vartab.empty)
 12.3079 -		     in snd (eval__true thy (i+1) 
 12.3080 -					(map (Envir.subst_term subst) pres)
 12.3081 -					[] erls)
 12.3082 -		     end)
 12.3083 -		    handle Pattern.MATCH => false
 12.3084 -		fun scan_ f [] = false (*scan_ NEVER called by []*)
 12.3085 -		  | scan_ f (pp::pps) = if f pp then true
 12.3086 -					else scan_ f pps;
 12.3087 -	    in scan_ chk prepat end;
 12.3088 -
 12.3089 -	(*.apply the normal_form of a rev-set.*)
 12.3090 -	fun app_rev' thy (Rrls{erls,prepat,scr=Rfuns{normal_form,...},...}) t =
 12.3091 -	    if chk_prepat thy erls prepat t
 12.3092 -	    then ((*tracing("### app_rev': t = "^(UnparseC.term t));*)
 12.3093 -                  normal_form t)
 12.3094 -	    else NONE;
 12.3095 -(*fixme val NONE = app_rev' thy rrls t;*)
 12.3096 -"----- app_rev' ---";
 12.3097 -val (thy, Rrls{erls,prepat,scr=Rfuns{normal_form,...},...}, t) = (thy, rrls, t);
 12.3098 -(*fixme false*)   chk_prepat thy erls prepat t;
 12.3099 -"----- chk_prepat ---";
 12.3100 -val (thy, erls, prepat, t) = (thy, erls, prepat, t);
 12.3101 -                fun chk (pres, pat) =
 12.3102 -		    (let val subst: Type.tyenv * Envir.tenv = 
 12.3103 -			     Pattern.match thy (pat, t)
 12.3104 -					    (Vartab.empty, Vartab.empty)
 12.3105 -		     in snd (eval__true thy (i+1) 
 12.3106 -					(map (Envir.subst_term subst) pres)
 12.3107 -					[] erls)
 12.3108 -		     end)
 12.3109 -		    handle Pattern.MATCH => false;
 12.3110 -		fun scan_ _ [] = false (*scan_ NEVER called by []*)
 12.3111 -		  | scan_ f (pp::pps) = if f pp then true
 12.3112 -					else scan_ f pps;
 12.3113 -tracing "=== poly.sml: scan_ chk prepat begin";
 12.3114 -scan_ chk prepat;
 12.3115 -tracing "=== poly.sml: scan_ chk prepat end";
 12.3116 -
 12.3117 -"----- chk ---";
 12.3118 -(*reestablish...*) val t = TermC.str2term "x \<up> 2 * x";
 12.3119 -val [(pres, pat)] = prepat;
 12.3120 -                         val subst: Type.tyenv * Envir.tenv = 
 12.3121 -			     Pattern.match thy (pat, t)
 12.3122 -					    (Vartab.empty, Vartab.empty);
 12.3123 -
 12.3124 -(*fixme: asms = ["p is_multUnordered"]...instantiate*)
 12.3125 -"----- eval__true ---";
 12.3126 -val asms = (map (Envir.subst_term subst) pres);
 12.3127 -if UnparseC.terms asms = "[\"x \<up> 2 * x is_multUnordered\"]" then ()
 12.3128 -else error "rewrite.sml: diff. is_multUnordered, asms";
 12.3129 -val (thy, i, asms, bdv, rls) = 
 12.3130 -    (thy, (i+1), [TermC.str2term "(x \<up> 2 * x) is_multUnordered"], 
 12.3131 -     [] : (term * term) list, erls);
 12.3132 -case eval__true thy i asms bdv rls of 
 12.3133 -    ([], true) => ()
 12.3134 -  | _ => error "rewrite.sml: diff. is_multUnordered, eval__true";
 12.3135 -
 12.3136 +val t = TermC.str2term "(17*x - 51)/9 - (-(13*x - 3)/6) + 11 - (9*x - 7)/4 = 0";
 12.3137 +val SOME (t',_) = rewrite_set_ thy false norm_Rational t; UnparseC.term t';
 12.3138  \<close> ML \<open>
 12.3139 -"----------- 2011 thms with axclasses ----------------------------------------------------------";
 12.3140 -"----------- 2011 thms with axclasses ----------------------------------------------------------";
 12.3141 -"----------- 2011 thms with axclasses ----------------------------------------------------------";
 12.3142 -val thm = ThmC.numerals_to_Free @{thm div_by_1};
 12.3143 -val prop = Thm.prop_of thm;
 12.3144 -TermC.atomty prop;                                     (*Type 'a*)
 12.3145 -val t = TermC.str2term "(2*x)/1";                      (*Type real*)
 12.3146 -
 12.3147 -val (thy, ro, er, inst) = 
 12.3148 -    (@{theory "Isac_Knowledge"}, tless_true, eval_rls, [(@{term "bdv::real"}, @{term "x::real"})]);
 12.3149 -val SOME (t, _) = rewrite_ thy ro er true thm t; (*real TermC.matches 'a ?via ring? etc*)
 12.3150 -
 12.3151 -\<close> ML \<open> (*loops*)
 12.3152 -"----------- fun rewrite_set_ .. RatEq_eliminate step into to fun chk --------------------------";
 12.3153 -"----------- fun rewrite_set_ .. RatEq_eliminate step into to fun chk --------------------------";
 12.3154 -"----------- fun rewrite_set_ .. RatEq_eliminate step into to fun chk --------------------------";
 12.3155 -val thy = @{theory RatEq};
 12.3156 -val ctxt = Proof_Context.init_global thy;
 12.3157 -val SOME t = parseNEW ctxt "(3 + -1 * x + x \<up> 2) / (9 * x + -6 * x \<up> 2 + x \<up> 3) = 1 / x";
 12.3158 -val rls = assoc_rls "RatEq_eliminate"
 12.3159 -
 12.3160 -val SOME (t''''', asm''''') =
 12.3161 -           rewrite_set_ thy true rls t;
 12.3162 -"~~~~~ fun rewrite_set_ , args:"; val (thy, bool, rls, term) = (thy, true, rls, t);
 12.3163 -           rewrite__set_ thy 1 bool [] rls term;
 12.3164 -"~~~~~ and rewrite__set_ , args:"; val (thy, i, put_asm, bdv, rls, ct)
 12.3165 -  = (thy, 1, bool, []:(term * term) list, rls, term);
 12.3166 -
 12.3167 -(*/------- outcomment this code: otherwise the re-definition could infect tests lateron ------\ * )
 12.3168 -      datatype switch = Applicable.Yes | Noap;
 12.3169 -      fun rew_once _ asm ct Noap [] = (ct, asm) (* ?TODO unify with Prog_Expr.rew_once? *)
 12.3170 -        | rew_once ruls asm ct Applicable.Yes [] = 
 12.3171 -          (case rls of Rule_Set.Repeat _ => rew_once ruls asm ct Noap ruls
 12.3172 -          | Rule_Set.Sequence _ => (ct, asm)
 12.3173 -          | rls => raise ERROR ("rew_once not appl. to \"" ^ Rule_Set.id rls ^ "\""))
 12.3174 -        | rew_once ruls asm ct apno (rul :: thms) =
 12.3175 -          case rul of
 12.3176 -            Rule.Thm (thmid, thm) =>
 12.3177 -              (trace1 i (" try thm: " ^ thmid);
 12.3178 -              case rewrite__ thy (i + 1) bdv ((snd o #rew_ord o Rule.Rule_Set.rep) rls)
 12.3179 -                  ((#erls o Rule.Rule_Set.rep) rls) put_asm thm ct of
 12.3180 -                NONE => rew_once ruls asm ct apno thms
 12.3181 -              | SOME (ct', asm') => 
 12.3182 -                (trace1 i (" rewrites to: " ^ UnparseC.term_in_thy thy ct');
 12.3183 -                rew_once ruls (union (op =) asm asm') ct' Applicable.Yes (rul :: thms)))
 12.3184 -                (* once again try the same rule, e.g. associativity against "()"*)
 12.3185 -          | Rule.Eval (cc as (op_, _)) => 
 12.3186 -            let val _= trace1 i (" try calc: " ^ op_ ^ "'")
 12.3187 -              val ct = TermC.uminus_to_string ct (*WN190312: superfluous?*)
 12.3188 -            in case Eval.adhoc_thm thy cc ct of
 12.3189 -                NONE => rew_once ruls asm ct apno thms
 12.3190 -              | SOME (_, thm') => 
 12.3191 -                let 
 12.3192 -                  val pairopt = rewrite__ thy (i + 1) bdv ((snd o #rew_ord o Rule.Rule_Set.rep) rls)
 12.3193 -                    ((#erls o Rule.Rule_Set.rep) rls) put_asm thm' ct;
 12.3194 -                  val _ = if pairopt <> NONE then () else error ("rewrite_set_, rewrite_ \"" ^ 
 12.3195 -                    ThmC.string_of_thm thm' ^ "\" " ^ UnparseC.term_in_thy thy ct ^ " = NONE")
 12.3196 -                  val _ = trace1 i (" calc. to: " ^ UnparseC.term_in_thy thy ((fst o the) pairopt))
 12.3197 -                in rew_once ruls asm ((fst o the) pairopt) Applicable.Yes (rul :: thms) end
 12.3198 -            end
 12.3199 -          | Rule.Cal1 (cc as (op_, _)) => 
 12.3200 -            let val _= trace1 i (" try cal1: " ^ op_ ^ "'");
 12.3201 -              val ct = TermC.uminus_to_string ct
 12.3202 -            in case Eval.adhoc_thm1_ thy cc ct of
 12.3203 -                NONE => (ct, asm)
 12.3204 -              | SOME (_, thm') =>
 12.3205 -                let 
 12.3206 -                  val pairopt = rewrite__ thy (i + 1) bdv ((snd o #rew_ord o Rule.Rule_Set.rep) rls)
 12.3207 -                    ((#erls o Rule.Rule_Set.rep) rls) put_asm thm' ct;
 12.3208 -                  val _ = if pairopt <> NONE then () else error ("rewrite_set_, rewrite_ \"" ^
 12.3209 -                     ThmC.string_of_thm thm' ^ "\" " ^ UnparseC.term_in_thy thy ct ^ " = NONE")
 12.3210 -                  val _ = trace1 i (" cal1. to: " ^ UnparseC.term_in_thy thy ((fst o the) pairopt))
 12.3211 -                in the pairopt end
 12.3212 -            end
 12.3213 -          | Rule.Rls_ rls' => 
 12.3214 -            (case rewrite__set_ thy (i + 1) put_asm bdv rls' ct of
 12.3215 -              SOME (t', asm') => rew_once ruls (union (op =) asm asm') t' Applicable.Yes thms
 12.3216 -            | NONE => rew_once ruls asm ct apno thms)
 12.3217 -          | r => raise ERROR ("rew_once not appl. to \"" ^ Rule.Rule.to_string r ^ "\"");
 12.3218 -      val ruls = (#rules o Rule.Rule_Set.rep) rls;
 12.3219 -(*    val _ = trace i (" rls: " ^ Rule_Set.id rls ^ " on: " ^ UnparseC.term_in_thy thy ct)*)
 12.3220 -      val (ct', asm') = rew_once ruls [] ct Noap ruls;
 12.3221 -"~~~~~ fun rew_once , args:"; val (ruls, asm, ct, apno, (rul :: thms))
 12.3222 -  = (ruls, []:term list, ct, Noap, ruls);
 12.3223 -           val Rule.Thm (thmid, thm) = (*case*) rul (*of*);
 12.3224 -
 12.3225 -    val SOME (ct', asm') = (*case*)
 12.3226 -           rewrite__ thy (i + 1) bdv ((snd o #rew_ord o Rule.Rule_Set.rep) rls)
 12.3227 -                  ((#erls o Rule.Rule_Set.rep) rls) put_asm thm ct (*of*);
 12.3228 -"~~~~~ fun rewrite__ , args:"; val (thy, i, bdv, tless, rls, put_asm, thm, ct)
 12.3229 -  = (thy, (i + 1), bdv, ((snd o #rew_ord o Rule.Rule_Set.rep) rls),
 12.3230 -                  ((#erls o Rule.Rule_Set.rep) rls), put_asm, thm, ct);
 12.3231 -
 12.3232 -    val (t', asms, _ (*lrd*), rew) =
 12.3233 -           rew_sub thy i bdv tless rls put_asm ([(*root of the term*)]: TermC.path)
 12.3234 -		  (((TermC.inst_bdv bdv) o Eval.norm o Thm.prop_of) thm) ct;
 12.3235 -"~~~~~ and rew_sub , args:"; val (thy, i, bdv, tless, rls, put_asm, lrd, r, t)
 12.3236 -  = (thy, i, bdv, tless, rls, put_asm, ([(*root of the term*)]: TermC.path),
 12.3237 -		  (((TermC.inst_bdv bdv) o Eval.norm o Thm.prop_of) thm), ct);
 12.3238 -    val (lhs, rhs) = (HOLogic.dest_eq o HOLogic.dest_Trueprop o Logic.strip_imp_concl) r
 12.3239 -    val r' = Envir.subst_term (Pattern.match thy (lhs, t) (Vartab.empty, Vartab.empty)) r
 12.3240 -;
 12.3241 -(*+*)if UnparseC.term r' =
 12.3242 -(*+*)  "\<lbrakk>9 * x + -6 * x \<up> 2 + x \<up> 3 \<noteq> 0; x \<noteq> 0\<rbrakk>\n" ^
 12.3243 -(*+*)  "\<Longrightarrow> ((3 + -1 * x + x \<up> 2) /\n" ^
 12.3244 -(*+*)  "                   (9 * x + -6 * x \<up> 2 + x \<up> 3) =\n" ^
 12.3245 -(*+*)  "                   1 / x) =\n" ^
 12.3246 -(*+*)  "                  ((3 + -1 * x + x \<up> 2) * x =\n" ^
 12.3247 -(*+*)  "                   1 * (9 * x + -6 * x \<up> 2 + x \<up> 3))"
 12.3248 -(*+*)then () else error "instantiated rule CHANGED";
 12.3249 -
 12.3250 -    val p' = map HOLogic.dest_Trueprop ((fst o Logic.strip_prems) (Logic.count_prems r', [], r'))
 12.3251 -;
 12.3252 -(*+*)if map UnparseC.term p' = ["x \<noteq> 0", "9 * x + -6 * x \<up> 2 + x \<up> 3 \<noteq> 0"]
 12.3253 -(*+*)then () else error "stored assumptions CHANGED";
 12.3254 -
 12.3255 -    val t' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Logic.strip_imp_concl) r'
 12.3256 -;
 12.3257 -(*+*)if UnparseC.term t' = "(3 + -1 * x + x \<up> 2) * x = 1 * (9 * x + -6 * x \<up> 2 + x \<up> 3)"
 12.3258 -(*+*)then () else error "rewritten term (an equality) CHANGED";
 12.3259 -
 12.3260 -        val (simpl_p', nofalse) =
 12.3261 -           eval__true thy (i + 1) p' bdv rls;
 12.3262 -"~~~~~ and eval__true , args:"; val (thy, i, asms, bdv, rls) = (thy, (i + 1), p', bdv, rls);
 12.3263 -  (*if*) asms = [@{term True}] orelse asms = [] (*else*); 
 12.3264 -
 12.3265 -(*+*)if map UnparseC.term asms = ["x \<noteq> 0", "9 * x + -6 * x \<up> 2 + x \<up> 3 \<noteq> 0"]
 12.3266 -(*+*)then () else error "asms before chk CHANGED";
 12.3267 -
 12.3268 -        fun chk indets [] = (indets, true) (*return asms<>True until false*)
 12.3269 -          | chk indets (a :: asms) =
 12.3270 -            (case rewrite__set_ thy (i + 1) false bdv rls a of
 12.3271 -              NONE => (chk (indets @ [a]) asms)
 12.3272 -            | SOME (t, a') =>
 12.3273 -              if t = @{term True} then (chk (indets @ a') asms) 
 12.3274 -              else if t = @{term False} then ([], false)
 12.3275 -            (*asm false .. thm not applied  \<up> ; continue until False vvv*)
 12.3276 -            else chk (indets @ [t] @ a') asms);
 12.3277 -
 12.3278 -    val (xxx, true) =
 12.3279 -           chk [] asms;  (*return from eval__true*);
 12.3280 -"~~~~~ fun chk , args:"; val (indets, (a :: asms)) = ([], asms);
 12.3281 -
 12.3282 -(*+*)val Rule_Set.Repeat {id = "rateq_erls", rules, ...} = rls;
 12.3283 -(*+*)(*val rules =
 12.3284 -(*+*)   [Eval ("Rings.divide_class.divide", fn),
 12.3285 -(*+*)    Thm ("minus_divide_left", "- ?a1 / ?b1 = - (?a1 / ?b1)"),
 12.3286 -(*+*)    :
 12.3287 -(*+*)    Eval ("HOL.eq", fn),
 12.3288 -(*+*)    Thm ("not_true", "(\<not> True) = False"),
 12.3289 -(*+*)    Thm ("not_false", "(\<not> False) = True"),
 12.3290 -(*+*)    :
 12.3291 -(*+*)    Eval ("Transcendental.powr", fn),
 12.3292 -(*+*)    Eval ("RatEq.is_ratequation_in", fn)]:
 12.3293 -(*+*)   rule list*)
 12.3294 -(*+*)chk: term list -> term list -> term list * bool
 12.3295 -
 12.3296 -           rewrite__set_ thy (i + 1) false bdv rls a (*of*);
 12.3297 -
 12.3298 -(*+*)Rewrite.trace_on := true;
 12.3299 -
 12.3300 -        (*this was False; vvvv--- means: indeterminate*)
 12.3301 -    val (* SOME (t, a') *)NONE = (*case*)
 12.3302 -           rewrite__set_ thy (i + 1) false bdv rls a (*of*);
 12.3303 -
 12.3304 -(*+*)UnparseC.term a = "x \<noteq> 0"; (* rewrite__set_ \<rightarrow> @{term True} ----------------- SHOULD BE indet !*)
 12.3305 +UnparseC.term t'
 12.3306 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation*)
 12.3307 +if UnparseC.term t' = "(237 + 65 * x) / 36 = 0" then () 
 12.3308 +else error "rational.sml 2";
 12.3309 +
 12.3310 +\<close> ML \<open>
 12.3311 +val t = TermC.str2term "(1/2 + (5*x)/2) \<up> 2 - ((13*x)/2 - 5/2) \<up> 2 - (6*x) \<up> 2 + 29";
 12.3312 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation*)
 12.3313 +val SOME (t',_) = rewrite_set_ thy false norm_Rational t; UnparseC.term t';
 12.3314 +if UnparseC.term t' = "23 + 35 * x + -72 * x \<up> 2" then ()
 12.3315 +else error "rational.sml 3";
 12.3316 +
 12.3317 +\<close> ML \<open>
 12.3318 +(*Rewrite.trace_on:=true;*)
 12.3319 +val t = TermC.str2term "Not (6*x is_atom)";
 12.3320 +val SOME (t',_) = rewrite_set_ thy false powers_erls t; UnparseC.term t';
 12.3321 +"HOL.True";
 12.3322 +val t = TermC.str2term "1 < 2";
 12.3323 +val SOME (t',_) = rewrite_set_ thy false powers_erls t; UnparseC.term t';
 12.3324 +"HOL.True";
 12.3325 +
 12.3326 +\<close> ML \<open>
 12.3327 +val t = TermC.str2term "(6*x) \<up> 2";
 12.3328 +val SOME (t',_) = rewrite_ thy dummy_ord powers_erls false 
 12.3329 +			   (ThmC.numerals_to_Free @{thm realpow_def_atom}) t;
 12.3330 +if UnparseC.term t' = "6 * x * (6 * x) \<up> (2 + - 1)" then ()
 12.3331 +else error "rational.sml powers_erls (6*x) \<up> 2";
 12.3332 +
 12.3333 +\<close> ML \<open>
 12.3334 +val t = TermC.str2term "-1 * (-2 * (5 / 2 * (13 * x / 2)))";
 12.3335 +val SOME (t',_) = rewrite_set_ thy false norm_Rational t; UnparseC.term t';
 12.3336 +if UnparseC.term t' = "65 * x / 2" then () else error "rational.sml 4";
 12.3337 +
 12.3338 +\<close> ML \<open>
 12.3339 +val t = TermC.str2term "1 - ((13*x)/2 - 5/2) \<up> 2";
 12.3340 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation*)
 12.3341 +val SOME (t',_) = rewrite_set_ thy false norm_Rational t; UnparseC.term t';
 12.3342 +if UnparseC.term t' = "(- 21 + 130 * x + - 169 * x \<up> 2) / 4" then () 
 12.3343 +else error "rational.sml 5";
 12.3344 +
 12.3345 +\<close> ML \<open>
 12.3346 +(*SRAM Schalk I, p.92 Nr. 609a*)
 12.3347 +val t = TermC.str2term "2*(3 - x/5)/3 - 4*(1 - x/3) - x/3 - 2*(x/2 - 1/4)/27 +5/54";
 12.3348 +val SOME (t',_) = rewrite_set_ thy false norm_Rational t; UnparseC.term t';
 12.3349 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation*)
 12.3350 +if UnparseC.term t' = "(-255 + 112 * x) / 135" then () 
 12.3351 +else error "rational.sml 6";
 12.3352 +
 12.3353 +\<close> ML \<open>
 12.3354 +(*SRAM Schalk I, p.92 Nr. 610c*)
 12.3355 +val t = TermC.str2term "((x- 1)/(x+1) + 1) / ((x- 1)/(x+1) - (x+1)/(x- 1)) - 2";
 12.3356 +val SOME (t',_) = rewrite_set_ thy false norm_Rational t; UnparseC.term t';
 12.3357 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation*)
 12.3358 +if UnparseC.term t' = "(3 + x) / -2" then () else error "rational.sml 7";
 12.3359 +
 12.3360 +\<close> ML \<open>
 12.3361 +(*SRAM Schalk I, p.92 Nr. 476a*)
 12.3362 +val t = TermC.str2term "(x \<up> 2/(1 - x \<up> 2) + 1)/(x/(1 - x) + 1) * (1 + x)";
 12.3363 +(*. a/b : c/d translated to a/b * d/c .*)
 12.3364 +val SOME (t',_) = rewrite_set_ thy false norm_Rational t; UnparseC.term t';
 12.3365 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation*)
 12.3366 +if UnparseC.term t' = "1" then () else error "rational.sml 8";
 12.3367 +
 12.3368 +\<close> ML \<open>
 12.3369 +(*Schalk I, p.92 Nr. 472a*)
 12.3370 +val t = TermC.str2term "((8*x \<up> 2 - 32*y \<up> 2)/(2*x + 4*y))/((4*x - 8*y)/(x + y))";
 12.3371 +val SOME (t',_) = rewrite_set_ thy false norm_Rational t; UnparseC.term t';
 12.3372 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation*)
 12.3373 +if UnparseC.term t' = "x + y" then () else error "rational.sml p.92 Nr. 472a";
 12.3374 +
 12.3375 +(*Schalk I, p.70 Nr. 480b: SEE rational.sml --- nonterminating rls norm_Rational ---*)
 12.3376 +
 12.3377 +(*WN130910 add_fractions_p exception Div raised + history:
 12.3378 +### WN.2.6.03 from rlang.sml 56a 
 12.3379 +val t = TermC.str2term "(a + b * x) / (a + -1 * (b * x)) + (-1 * a + b * x) / (a + b * x) = 4 * (a * b) / (a \<up> 2 + -1 * b \<up> 2)";
 12.3380 +val NONE = rewrite_set_ thy false add_fractions_p t;
 12.3381 +
 12.3382 +THE ERROR ALREADY OCCURS IN THIS PART:
 12.3383 +val t = TermC.str2term "(a + b * x) / (a + -1 * (b * x)) + (-1 * a + b * x) / (a + b * x)";
 12.3384 +val NONE = add_fraction_p_ thy t;
 12.3385 +
 12.3386 +SEE Test_Some.thy: section {* add_fractions_p downto exception Div raised ===
 12.3387 +*)
 12.3388 +
 12.3389 +\<close> ML \<open>
 12.3390 +"-------- rational numerals --------------------------------------------------";
 12.3391 +"-------- rational numerals --------------------------------------------------";
 12.3392 +"-------- rational numerals --------------------------------------------------";
 12.3393 +(*SRA Schalk I, p.40 Nr. 164b *)
 12.3394 +val t = TermC.str2term "(47/6 - 76/9 + 13/4)/(35/12)";
 12.3395 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation WITH NUMERALS*)
 12.3396 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3397 +if UnparseC.term t = "19 / 21" then ()
 12.3398 +else error "rational.sml: diff.behav. in norm_Rational_mg 1";
 12.3399 +
 12.3400 +\<close> ML \<open>
 12.3401 +(*SRA Schalk I, p.40 Nr. 166a *)
 12.3402 +val t = TermC.str2term "((5/4)/(4+22/7) + 37/20)*(110/3 - 110/9 * 23/11)";
 12.3403 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation WITH NUMERALS*)
 12.3404 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3405 +if UnparseC.term t = "45 / 2" then ()
 12.3406 +else error "rational.sml: diff.behav. in norm_Rational_mg 2";
 12.3407 +
 12.3408 +\<close> ML \<open>
 12.3409 +"-------- examples cancellation from: Mathematik 1 Schalk --------------------";
 12.3410 +"-------- examples cancellation from: Mathematik 1 Schalk --------------------";
 12.3411 +"-------- examples cancellation from: Mathematik 1 Schalk --------------------";
 12.3412 +(* e190c Stefan K.*)
 12.3413 +val t = TermC.str2term "((1 + 9*a \<up> 2) * (1 + 3*a)) / ((3*a + 9*a \<up> 2) * (1 + 3*a))";
 12.3414 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3415 +\<close> text \<open> (*rewrite_set_ norm_Rational NO cancellation MANY EXAMPLES*)
 12.3416 +if UnparseC.term t = "(1 + 9 * a \<up> 2) / (3 * a + 9 * a \<up> 2)"
 12.3417 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 3";
 12.3418 +
 12.3419 +(* e192b Stefan K.*)
 12.3420 +val t = TermC.str2term "(x \<up> 2 * (7*x + (-1)*y))  /  (y \<up> 2 * (7*x + (-1)*y))";
 12.3421 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3422 +if UnparseC.term t = "x \<up> 2 / y \<up> 2"
 12.3423 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 4";
 12.3424 +
 12.3425 +(*SRC Schalk I, p.66 Nr. 379c *)
 12.3426 +val t = TermC.str2term "(a - b)/(b - a)";
 12.3427 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3428 +UnparseC.term t;
 12.3429 +if UnparseC.term t = "-1"
 12.3430 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 5";
 12.3431 +
 12.3432 +(*SRC Schalk I, p.66 Nr. 380b *)
 12.3433 +val t = TermC.str2term "15*(3*x + 3) * (4*x + 9)  /  (12*(2*x + 7) * (5*x + 5))";
 12.3434 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3435 +if UnparseC.term t = "(27 + 12 * x) / (28 + 8 * x)"
 12.3436 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 6";
 12.3437 +
 12.3438 +(*Schalk I, p.60 Nr. 215c: was not cancelled with Isabelle2002 *)
 12.3439 +val t = TermC.str2term "(a + b) \<up> 4 * (x - y)  /  ((x - y) \<up> 3 * (a + b) \<up> 2)";
 12.3440 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3441 +if UnparseC.term t = "(a \<up> 2 + 2 * a * b + b \<up> 2) / (x \<up> 2 + -2 * x * y + y \<up> 2)"
 12.3442 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 7";
 12.3443 +
 12.3444 +(*SRC Schalk I, p.66 Nr. 381b *)
 12.3445 +val t = TermC.str2term 
 12.3446 +"(4*x \<up> 2 - 20*x + 25)/(2*x - 5) \<up> 3";
 12.3447 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3448 +if UnparseC.term t = "1 / (-5 + 2 * x)"
 12.3449 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 9";
 12.3450 +
 12.3451 +(* e190c Stefan K.*)
 12.3452 +val t = TermC.str2term "((1 + 9*a \<up> 2) * (1 + 3*a))  /  ((3*a + 9*a \<up> 2) * (1 + 3 * a))";
 12.3453 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3454 +if UnparseC.term t =  "(1 + 9 * a \<up> 2) / (3 * a + 9 * a \<up> 2)"
 12.3455 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 3";
 12.3456 +
 12.3457 +(* e192b Stefan K.*)
 12.3458 +val t = TermC.str2term "(x \<up> 2 * (7*x + (-1)*y))  /  (y \<up> 2 * (7*x + (-1)*y))";
 12.3459 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3460 +if UnparseC.term t = "x \<up> 2 / y \<up> 2"
 12.3461 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 4";
 12.3462 +
 12.3463 +(*SRC Schalk I, p.66 Nr. 379c *)
 12.3464 +val t = TermC.str2term "(a - b) / (b - a)";
 12.3465 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3466 +if UnparseC.term t = "-1"
 12.3467 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 5";
 12.3468 +
 12.3469 +(*SRC Schalk I, p.66 Nr. 380b *)
 12.3470 +val t = TermC.str2term "15*(3*x + 3) * (4*x + 9)  /  (12*(2*x + 7) * (5*x + 5))";
 12.3471 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3472 +if UnparseC.term t = "(27 + 12 * x) / (28 + 8 * x)"
 12.3473 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 6";
 12.3474 +
 12.3475 +(*Schalk I, p.60 Nr. 215c *)
 12.3476 +val t = TermC.str2term "(a + b) \<up> 4 * (x - y)  /  ((x - y) \<up> 3 * (a + b) \<up> 2)";
 12.3477 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3478 +if UnparseC.term t = "(a \<up> 2 + 2 * a * b + b \<up> 2) / (x \<up> 2 + -2 * x * y + y \<up> 2)"
 12.3479 +then () else error "Schalk I, p.60 Nr. 215c: with Isabelle2002 cancellation incomplete, changed";
 12.3480 +
 12.3481 +(* extreme example from somewhere *)
 12.3482 +val t = TermC.str2term 
 12.3483 +    ("(a \<up> 4 * x  +  -1*a \<up> 4 * y  +  4*a \<up> 3 * b * x  +  -4*a \<up> 3 * b * y  + " ^
 12.3484 +      "6*a \<up> 2 * b \<up> 2 * x  +  -6*a \<up> 2 * b \<up> 2 * y  +  4*a * b \<up> 3 * x  +  -4*a * b \<up> 3 * y  + " ^
 12.3485 +      "b \<up> 4 * x  +  -1*b \<up> 4 * y) " ^
 12.3486 +  " / (a \<up> 2 * x \<up> 3  +  -3*a \<up> 2 * x \<up> 2 * y  +  3*a \<up> 2 * x * y \<up> 2  +  -1*a \<up> 2 * y \<up> 3 + " ^
 12.3487 +      "2*a * b * x \<up> 3  +  -6*a * b * x \<up> 2 * y  +  6*a * b * x * y \<up> 2  +  -2*a * b * y \<up> 3 + " ^
 12.3488 +      "b \<up> 2 * x \<up> 3  +  -3*b \<up> 2 * x \<up> 2 * y  +  3*b \<up> 2 * x * y \<up> 2  +  -1*b \<up> 2 * y \<up> 3)")
 12.3489 +val SOME (t, _) = rewrite_set_ thy false cancel_p t;
 12.3490 +if UnparseC.term t = "(a \<up> 2 + 2 * a * b + b \<up> 2) / (x \<up> 2 + -2 * x * y + y \<up> 2)"
 12.3491 +then () else error "with Isabelle2002: NONE -- now SOME changed";
 12.3492 +
 12.3493 +(*Schalk I, p.66 Nr. 381a *)
 12.3494 +(* ATTENTION: here the rls is very slow. In Isabelle2002 this required 2 min *)
 12.3495 +val t = TermC.str2term "18*(a + b) \<up> 3 * (a - b) \<up> 2 / (72*(a - b) \<up> 3 * (a + b) \<up> 2)";
 12.3496 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3497 +if UnparseC.term t = "(a + b) / (4 * a + -4 * b)"
 12.3498 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 8";
 12.3499 +
 12.3500 +(*SRC Schalk I, p.66 Nr. 381b *)
 12.3501 +val t = TermC.str2term "(4*x \<up> 2 - 20*x + 25) / (2*x - 5) \<up> 3";
 12.3502 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3503 +if UnparseC.term t = "1 / (-5 + 2 * x)"
 12.3504 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 9";
 12.3505 +
 12.3506 +(*SRC Schalk I, p.66 Nr. 381c *)
 12.3507 +val t = TermC.str2term "(27*a \<up> 3 + 9*a \<up> 2+3*a+1) / (27*a \<up> 3 + 18*a \<up> 2+3*a)";
 12.3508 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3509 +if UnparseC.term t = "(1 + 9 * a \<up> 2) / (3 * a + 9 * a \<up> 2)"
 12.3510 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 10";
 12.3511 +
 12.3512 +(*SRC Schalk I, p.66 Nr. 383a *)
 12.3513 +val t = TermC.str2term "(5*a \<up> 2 - 5*a*b) / (a - b) \<up> 2";
 12.3514 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3515 +if UnparseC.term t = "-5 * a / (-1 * a + b)"
 12.3516 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 11";
 12.3517 +
 12.3518 +"----- NOT TERMINATING ?: worked before 0707xx";
 12.3519 +val t = TermC.str2term "(a \<up> 2 - 1)*(b + 1) / ((b \<up> 2 - 1)*(a+1))";
 12.3520 +(* WN130911 "exception Div raised" by 
 12.3521 +  cancel_p_ thy (TermC.str2term ("(-1 + -1 * b + a \<up> 2 + a \<up> 2 * b) /" ^
 12.3522 +                           "(-1 + -1 * a + b \<up> 2 + a * b \<up> 2)"))
 12.3523 +
 12.3524 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3525 +if UnparseC.term t = "(1 + -1 * a) / (1 + -1 * b)" then ()
 12.3526 +else error "rational.sml MG tests 3e";
 12.3527 +*)
 12.3528 +
 12.3529 +\<close> ML \<open>
 12.3530 +"-------- examples common denominator from: Mathematik 1 Schalk --------------";
 12.3531 +"-------- examples common denominator from: Mathematik 1 Schalk --------------";
 12.3532 +"-------- examples common denominator from: Mathematik 1 Schalk --------------";
 12.3533 +(*SRA Schalk I, p.67 Nr. 403a *)
 12.3534 +val t = TermC.str2term "4/x - 3/y - 1";
 12.3535 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3536 +\<close> text \<open> (*rewrite_set_  norm_Rational NO +*)
 12.3537 +if UnparseC.term t = "(-3 * x + 4 * y + -1 * x * y) / (x * y)"
 12.3538 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 12";
 12.3539 +
 12.3540 +\<close> ML \<open>
 12.3541 +val t = TermC.str2term "(2*a+3*b)/(b*c) + (3*c+a)/(a*c) - (2*a \<up> 2+3*b*c)/(a*b*c)";
 12.3542 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3543 +\<close> text \<open> (*rewrite_set_  norm_Rational NO +*)
 12.3544 +if UnparseC.term t = "4 / c"
 12.3545 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 13";
 12.3546 +
 12.3547 +\<close> text \<open>
 12.3548 +(*SRA Schalk I, p.67 Nr. 410b *)
 12.3549 +val t = TermC.str2term "1/(x+1) + 1/(x+2) - 2/(x+3)";
 12.3550 +(* WN130911 non-termination due to non-termination of
 12.3551 +  cancel_p_ thy (TermC.str2term "(5 + 3 * x) / (6 + 11 * x + 6 * x \<up> 2 + x \<up> 3)")
 12.3552 +
 12.3553 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3554 +if UnparseC.term t = "(5 + 3 * x) / (6 + 11 * x + 6 * x \<up> 2 + x \<up> 3)"
 12.3555 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 14";
 12.3556 +*)
 12.3557 +
 12.3558 +(*SRA Schalk I, p.67 Nr. 413b *)
 12.3559 +val t = TermC.str2term "(1 + x)/(1 - x)  -  (1 - x)/(1 + x)  +  2*x/(1 - x \<up> 2)";
 12.3560 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3561 +if UnparseC.term t = "6 * x / (1 + -1 * x \<up> 2)"
 12.3562 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 15";
 12.3563 +
 12.3564 +(*SRA Schalk I, p.68 Nr. 414a *)
 12.3565 +val t = TermC.str2term "(x + 2)/(x - 1)  +  (x - 3)/(x - 2)  -  (x + 1)/((x - 1)*(x - 2))";
 12.3566 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3567 +if UnparseC.term t ="(-2 + -5 * x + 2 * x \<up> 2) / (2 + -3 * x + x \<up> 2)"
 12.3568 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 16";
 12.3569 +
 12.3570 +(*SRA Schalk I, p.68 Nr. 428b *)
 12.3571 +val t = TermC.str2term 
 12.3572 +  "1/(a - b) \<up> 2  +  1/(a + b) \<up> 2  -  2/(a \<up> 2 - b \<up> 2)  -  4*(b \<up> 2 - 1)/(a \<up> 2 - b \<up> 2) \<up> 2";
 12.3573 +(* WN130911 non-termination due to non-termination of
 12.3574 +  cancel_p_ thy (TermC.str2term "(4 + -4 * b \<up> 2) / (a \<up> 4 + -2 * (a \<up> 2 * b \<up> 2) + b \<up> 4)")
 12.3575 +
 12.3576 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3577 +if UnparseC.term t = "4 / (a \<up> 4 + -2 * a \<up> 2 * b \<up> 2 + b \<up> 4)"
 12.3578 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 18";
 12.3579 +*)
 12.3580 +
 12.3581 +(*SRA Schalk I, p.68 Nr. 430b *)
 12.3582 +val t = TermC.str2term 
 12.3583 +  "a \<up> 2/(a - 3*b) - 108*a*b \<up> 3/((a+3*b)*(a \<up> 2 - 9*b \<up> 2)) - 9*b \<up> 2*(a - 3*b)/(a+3*b) \<up> 2";
 12.3584 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3585 +if UnparseC.term t = "a + 3 * b"
 12.3586 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 19";
 12.3587 +
 12.3588 +(*SRA Schalk I, p.68 Nr. 432 *)
 12.3589 +val t = TermC.str2term 
 12.3590 +  ("(a \<up> 2 + a*b) / (a \<up> 2 - b \<up> 2)  -  (b \<up> 2 - a*b) / (b \<up> 2 - a \<up> 2)  +  " ^
 12.3591 +  "a \<up> 2*(a - b) / (a \<up> 3 - a \<up> 2*b)  -  2*a*(a \<up> 2 - b \<up> 2) / (a \<up> 3 - a*b \<up> 2)  -  " ^
 12.3592 +  "2*b \<up> 2 / (a \<up> 2 - b \<up> 2)");
 12.3593 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3594 +if UnparseC.term t = (*"0" ..isabisac15 | Isabelle2017..*)  "0 / (a \<up> 2 + -1 * b \<up> 2)"
 12.3595 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 20";
 12.3596 +
 12.3597 +\<close> ML \<open>
 12.3598 +(* some example *)
 12.3599 +val t = TermC.str2term "3*a / (a*b)  +  x/y";
 12.3600 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3601 +if UnparseC.term t = "(3 * y + b * x) / (b * y)"
 12.3602 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 21";
 12.3603 +
 12.3604 +\<close> text \<open> (*multiply and cancel MANY EXAMPLES*)
 12.3605 +"-------- examples multiply and cancel from: Mathematik 1 Schalk -------------";
 12.3606 +"-------- examples multiply and cancel from: Mathematik 1 Schalk -------------";
 12.3607 +"-------- examples multiply and cancel from: Mathematik 1 Schalk -------------";
 12.3608 +(*------- SRM Schalk I, p.68 Nr. 436a *)
 12.3609 +val t = TermC.str2term "3*(x+y) / (15*(x - y))  *   25*(x - y) \<up> 2 / (18*(x + y) \<up> 2)";
 12.3610 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3611 +if UnparseC.term t = "(-5 * x + 5 * y) / (-18 * x + -18 * y)"
 12.3612 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 22";
 12.3613 +
 12.3614 +(*------- SRM.test Schalk I, p.68 Nr. 436b *)
 12.3615 +val t = TermC.str2term "5*a*(a - b) \<up> 2*(a + b) \<up> 3/(7*b*(a - b) \<up> 3) * 7*b/(a + b) \<up> 3";
 12.3616 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3617 +if UnparseC.term t = "5 * a / (a + -1 * b)"
 12.3618 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 23";
 12.3619 +
 12.3620 +(*------- Schalk I, p.68 Nr. 437a *)
 12.3621 +val t = TermC.str2term "(3*a - 4*b) / (4*c+3*e)  *  (3*a+4*b)/(9*a \<up> 2 - 16*b \<up> 2)";
 12.3622 +(* raises an exception for unclear reasons:
 12.3623 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3624 +:
 12.3625 +###  rls: cancel_p on: (9 * a \<up> 2 + -16 * b \<up> 2) / (4 * c + 3 * e) /
 12.3626 +(9 * a \<up> 2 + -16 * b \<up> 2) 
 12.3627 +exception Div raised
 12.3628 +
 12.3629 +BUT
 12.3630 +val t = TermC.str2term 
 12.3631 +  ("(9 * a \<up> 2 + -16 * b \<up> 2) / (4 * c + 3 * e) /" ^
 12.3632 +  "(9 * a \<up> 2 + -16 * b \<up> 2)");
 12.3633 +NONE = cancel_p_ thy t;
 12.3634 +
 12.3635 +if UnparseC.term t = "1 / (4 * c + 3 * e)" then ()
 12.3636 +else error "rational.sml: diff.behav. in norm_Rational_mg 24";
 12.3637 +*)
 12.3638 +
 12.3639 +"----- S.K. corrected non-termination 060904";
 12.3640 +val t = TermC.str2term "(3*a - 4*b) * (3*a+4*b)/((4*c+3*e)*(9*a \<up> 2 - 16*b \<up> 2))";
 12.3641 +val SOME (t, _) = rewrite_set_ thy false make_polynomial t;
 12.3642 +if UnparseC.term t = 
 12.3643 +  "(9 * a \<up> 2 + -16 * b \<up> 2) /\n(36 * a \<up> 2 * c + 27 * a \<up> 2 * e + -64 * b \<up> 2 * c +\n -48 * b \<up> 2 * e)"
 12.3644 +(*"(9 * a \<up> 2 + -16 * b \<up> 2) / (36 * a \<up> 2 * c + 27 * a \<up> 2 * e + -64 * b \<up> 2 * c + -48 * b \<up> 2 * e)"*)
 12.3645 +then () else error "rational.sml: S.K.8..corrected 060904-6";
 12.3646 +
 12.3647 +"----- S.K. corrected non-termination of cancel_p_";
 12.3648 +val t'' = TermC.str2term ("(9 * a \<up> 2 + -16 * b \<up> 2) /" ^
 12.3649 +  "(36 * a \<up> 2 * c + (27 * a \<up> 2 * e + (-64 * b \<up> 2 * c + -48 * b \<up> 2 * e)))");
 12.3650 +(* /--- DOES NOT TERMINATE AT TRANSITION isabisac15 --> Isabelle2017 --------------------------\
 12.3651 +val SOME (t',_) = rewrite_set_ thy false cancel_p t'';
 12.3652 +if UnparseC.term t' = "1 / (4 * c + 3 * e)"
 12.3653 +then () else error "rational.sml: diff.behav. in cancel_p S.K.8";
 12.3654 +   \--- DOES NOT TERMINATE AT TRANSITION isabisac15 --> Isabelle2017 --------------------------/*)
 12.3655 +
 12.3656 +(*------- Schalk I, p.68 Nr. 437b*)
 12.3657 +val t = TermC.str2term "(a + b)/(x \<up> 2 - y \<up> 2) * ((x - y) \<up> 2/(a \<up> 2 - b \<up> 2))";
 12.3658 +(*val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3659 +:
 12.3660 +####  rls: cancel_p on: (a * x \<up> 2 + -2 * (a * (x * y)) + a * y \<up> 2 + b * x \<up> 2 +
 12.3661 + -2 * (b * (x * y)) +
 12.3662 + b * y \<up> 2) /
 12.3663 +(a \<up> 2 * x \<up> 2 + -1 * (a \<up> 2 * y \<up> 2) + -1 * (b \<up> 2 * x \<up> 2) +
 12.3664 + b \<up> 2 * y \<up> 2) 
 12.3665 +exception Div raised
 12.3666 +*)
 12.3667 +
 12.3668 +(*------- SRM Schalk I, p.68 Nr. 438a *)
 12.3669 +val t = TermC.str2term "x*y / (x*y - y \<up> 2)  *  (x \<up> 2 - x*y)";
 12.3670 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3671 +if UnparseC.term t = "x \<up> 2"
 12.3672 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 24";
 12.3673 +
 12.3674 +(*------- SRM Schalk I, p.68 Nr. 439b *)
 12.3675 +val t = TermC.str2term "(4*x \<up> 2 + 4*x + 1)  *  ((x \<up> 2 - 2*x \<up> 3) / (4*x \<up> 2 + 2*x))";
 12.3676 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3677 +if UnparseC.term t = "(x + -4 * x \<up> 3) / 2"
 12.3678 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 25";
 12.3679 +
 12.3680 +(*------- SRM Schalk I, p.68 Nr. 440a *)
 12.3681 +val t = TermC.str2term "(x \<up> 2 - 2*x) / (x \<up> 2 - 3*x)  *  (x - 3) \<up> 2 / (x \<up> 2 - 4)";
 12.3682 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3683 +if UnparseC.term t = "(-3 + x) / (2 + x)"
 12.3684 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 26";
 12.3685 +
 12.3686 +"----- Schalk I, p.68 Nr. 440b SK11 works since 0707xx";
 12.3687 +val t = TermC.str2term "(a \<up> 3 - 9*a) / (a \<up> 3*b - a*b \<up> 3)  *  (a \<up> 2*b + a*b \<up> 2) / (a+3)";
 12.3688 +(* WN130911 non-termination for unclear reasons:
 12.3689 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3690 +
 12.3691 +... ENDS WITH THIS TRACE:
 12.3692 +:
 12.3693 +###  rls: cancel_p on: (-9 * (a \<up> 3 * b) + -9 * (a \<up> 2 * b \<up> 2) + a \<up> 5 * b +
 12.3694 + a \<up> 4 * b \<up> 2) /
 12.3695 +(a \<up> 3 * b + -1 * (a * b \<up> 3)) /
 12.3696 +(3 + a)
 12.3697 +BUT THIS IS CORRECTLY RECOGNISED 
 12.3698 +val t = TermC.str2term 
 12.3699 +  ("(-9 * (a \<up> 3 * b) + -9 * (a \<up> 2 * b \<up> 2) + a \<up> 5 * b + a \<up> 4 * b \<up> 2)  /" ^
 12.3700 +  "(a \<up> 3 * b + -1 * (a * b \<up> 3))  /  (3 + (a::real))");
 12.3701 +AS
 12.3702 +NONE = cancel_p_ thy t;
 12.3703 +
 12.3704 +if UnparseC.term t = "(-3 * a + a \<up> 2) / (a + -1 * b)" then ()
 12.3705 +else error "rational.sml: diff.behav. in norm_Rational 27";
 12.3706 +*)
 12.3707 +
 12.3708 +"----- SK12 works since 0707xx";
 12.3709 +val t = TermC.str2term "(a \<up> 3 - 9*a) * (a \<up> 2*b+a*b \<up> 2)  /  ((a \<up> 3*b - a*b \<up> 3) * (a+3))";
 12.3710 +(* WN130911 non-termination due to non-termination of
 12.3711 +  cancel_p_ thy (TermC.str2term "(4 + -4 * b \<up> 2) / (a \<up> 4 + -2 * (a \<up> 2 * b \<up> 2) + b \<up> 4)")
 12.3712 +
 12.3713 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3714 +if UnparseC.term t' = "(-3 * a + a \<up> 2) / (a + -1 * b)" then ()
 12.3715 +else error "rational.sml: diff.behav. in norm_Rational 28";
 12.3716 +*)
 12.3717 +
 12.3718 +\<close> text \<open> (*common denominator and multiplication MANY EXAMPLES*)
 12.3719 +"-------- examples common denominator and multiplication from: Schalk --------";
 12.3720 +"-------- examples common denominator and multiplication from: Schalk --------";
 12.3721 +"-------- examples common denominator and multiplication from: Schalk --------";
 12.3722 +(*------- SRAM Schalk I, p.69 Nr. 441b *)
 12.3723 +val t = TermC.str2term "(4*a/3 + 3*b \<up> 2/a \<up> 3 + b/(4*a))*(4*b/(3*a))";
 12.3724 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3725 +if UnparseC.term t = "(36 * b \<up> 3 + 3 * a \<up> 2 * b \<up> 2 + 16 * a \<up> 4 * b) / (9 * a \<up> 4)"
 12.3726 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 28";
 12.3727 +
 12.3728 +(*------- SRAM Schalk I, p.69 Nr. 442b *)
 12.3729 +val t = TermC.str2term ("(15*a \<up> 2/x \<up> 3 - 5*b \<up> 4/x \<up> 2 + 25*c \<up> 2/x) * " ^
 12.3730 +  "(x \<up> 3/(5*a*b \<up> 3*c \<up> 3)) + 1/c \<up> 3 * (b*x/a - 3*a/b \<up> 3)");
 12.3731 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3732 +if UnparseC.term t = "5 * x \<up> 2 / (a * b \<up> 3 * c)"
 12.3733 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 29";
 12.3734 +
 12.3735 +(*------- SRAM Schalk I, p.69 Nr. 443b *)
 12.3736 +val t = TermC.str2term "(a/2 + b/3) * (b/3 - a/2)";
 12.3737 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3738 +if UnparseC.term t = "(-9 * a \<up> 2 + 4 * b \<up> 2) / 36"
 12.3739 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 30";
 12.3740 +
 12.3741 +(*------- SRAM Schalk I, p.69 Nr. 445b *)
 12.3742 +val t = TermC.str2term "(a \<up> 2/9 + 2*a/(3*b) + 4/b \<up> 2)*(a/3 - 2/b) + 8/b \<up> 3";
 12.3743 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3744 +if UnparseC.term t = "a \<up> 3 / 27"
 12.3745 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 31";
 12.3746 +
 12.3747 +(*------- SRAM Schalk I, p.69 Nr. 446b *)
 12.3748 +val t = TermC.str2term "(x/(5*x + 4*y) - y/(5*x - 4*y) + 1)*(25*x \<up> 2 - 16*y \<up> 2)";
 12.3749 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3750 +if UnparseC.term t = (*"30 * x \<up> 2 + -9 * x * y + -20 * y \<up> 2" ..isabisac15 | Isabelle2017..*)
 12.3751 +                  "(-30 * x \<up> 2 + 9 * x * y + 20 * y \<up> 2) / -1"
 12.3752 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 32";
 12.3753 +
 12.3754 +(*------- SRAM Schalk I, p.69 Nr. 449a *)(*Achtung: rechnet ca 8 Sekunden*)
 12.3755 +val t = TermC.str2term 
 12.3756 +"(2*x \<up> 2/(3*y)+x/y \<up> 2)*(4*x \<up> 4/(9*y \<up> 2)+x \<up> 2/y \<up> 4)*(2*x \<up> 2/(3*y) - x/y \<up> 2)";
 12.3757 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3758 +if UnparseC.term t = "(-81 * x \<up> 4 + 16 * x \<up> 8 * y \<up> 4) / (81 * y \<up> 8)"
 12.3759 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 33";
 12.3760 +
 12.3761 +(*------- SRAM Schalk I, p.69 Nr. 450a *)
 12.3762 +val t = TermC.str2term 
 12.3763 +"(4*x/(3*y)+2*y/(3*x)) \<up> 2 - (2*y/(3*x) - 2*x/y)*(2*y/(3*x)+2*x/y)";
 12.3764 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3765 +if UnparseC.term t = "(52 * x \<up> 2 + 16 * y \<up> 2) / (9 * y \<up> 2)"
 12.3766 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 34";
 12.3767 +
 12.3768 +(*------- SRAM Schalk I, p.69 Nr. 442b --- abgewandelt*)
 12.3769 +val t = TermC.str2term 
 12.3770 +  ("(15*a \<up> 4/(a*x \<up> 3)  -  5*a*((b \<up> 4 - 5*c \<up> 2*x) / x \<up> 2))  *  " ^
 12.3771 +  "(x \<up> 3/(5*a*b \<up> 3*c \<up> 3))   +   a/c \<up> 3 * (x*(b/a) - 3*b*(a/b \<up> 4))");
 12.3772 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3773 +if UnparseC.term t = "5 * x \<up> 2 / (b \<up> 3 * c)"
 12.3774 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 53";
 12.3775 +
 12.3776 +
 12.3777 +\<close> text \<open> (*double fractions MANY EXAMPLES*)
 12.3778 +"-------- examples double fractions from: Mathematik 1 Schalk ----------------";
 12.3779 +"-------- examples double fractions from: Mathematik 1 Schalk ----------------";
 12.3780 +"-------- examples double fractions from: Mathematik 1 Schalk ----------------";
 12.3781 +"----- SRD Schalk I, p.69 Nr. 454b";
 12.3782 +val t = TermC.str2term "((2 - x)/(2*a)) / (2*a/(x - 2))";
 12.3783 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3784 +if UnparseC.term t = "(-4 + 4 * x + -1 * x \<up> 2) / (4 * a \<up> 2)"
 12.3785 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 35";
 12.3786 +
 12.3787 +"----- SRD Schalk I, p.69 Nr. 455a";
 12.3788 +val t = TermC.str2term "(a \<up> 2 + 1)/(a \<up> 2 - 1) / ((a+1)/(a - 1))";
 12.3789 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3790 +if UnparseC.term t = "(1 + a \<up> 2) / (1 + 2 * a + a \<up> 2)" then ()
 12.3791 +else error "rational.sml: diff.behav. in norm_Rational_mg 36";
 12.3792 +
 12.3793 +"----- Schalk I, p.69 Nr. 455b";
 12.3794 +val t = TermC.str2term "(x \<up> 2 - 4)/(y \<up> 2 - 9)/((2+x)/(3 - y))";
 12.3795 +(* WN130911 non-termination due to non-termination of
 12.3796 +  cancel_p_ thy (TermC.str2term ("(-12 + 4 * y + 3 * x \<up> 2 + -1 * (x \<up> 2 * y)) /" ^
 12.3797 +                           "(-18 + -9 * x + 2 * y \<up> 2 + x * y \<up> 2)"))
 12.3798 +
 12.3799 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3800 +if UnparseC.term t = "(2 + -1 * x) / (3 + y)" then ()
 12.3801 +else error "rational.sml: diff.behav. in norm_Rational_mg 37";
 12.3802 +*)
 12.3803 +
 12.3804 +"----- SK060904-1a non-termination of cancel_p_ ?: worked before 0707xx";
 12.3805 +val t = TermC.str2term "(x \<up> 2 - 4)*(3 - y) / ((y \<up> 2 - 9)*(2+x))";
 12.3806 +(* WN130911 non-termination due to non-termination of
 12.3807 +  cancel_p_ thy (TermC.str2term ("(-12 + 4 * y + 3 * x \<up> 2 + -1 * (x \<up> 2 * y)) /" ^
 12.3808 +                           "(-18 + -9 * x + 2 * y \<up> 2 + x * y \<up> 2)"))
 12.3809 +
 12.3810 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3811 +if UnparseC.term t = "(2 + -1 * x) / (3 + y)" then ()
 12.3812 +else error "rational.sml: diff.behav. in norm_Rational_mg 37b";
 12.3813 +*)
 12.3814 +
 12.3815 +"----- ?: worked before 0707xx";
 12.3816 +val t = TermC.str2term "(3 + -1 * y) / (-9 + y \<up> 2)";
 12.3817 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3818 +if UnparseC.term t = "-1 / (3 + y)"
 12.3819 +then () else error "rational.sml: -1 / (3 + y) norm_Rational";
 12.3820 +
 12.3821 +"----- SRD Schalk I, p.69 Nr. 456b";
 12.3822 +val t = TermC.str2term "(b \<up> 3 - b \<up> 2) / (b \<up> 2+b) / (b \<up> 2 - 1)";
 12.3823 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3824 +if UnparseC.term t = "b / (1 + 2 * b + b \<up> 2)"
 12.3825 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 38";
 12.3826 +
 12.3827 +"----- SRD Schalk I, p.69 Nr. 457b";
 12.3828 +val t = TermC.str2term "(16*a \<up> 2 - 9*b \<up> 2)/(2*a+3*a*b) / ((4*a+3*b)/(4*a \<up> 2 - 9*a \<up> 2*b \<up> 2))";
 12.3829 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3830 +if UnparseC.term t = "8 * a \<up> 2 + -6 * a * b + -12 * a \<up> 2 * b + 9 * a * b \<up> 2"
 12.3831 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 39";
 12.3832 +
 12.3833 +"----- Schalk I, p.69 Nr. 458b works since 0707";
 12.3834 +val t = TermC.str2term "(2*a \<up> 2*x - a \<up> 2) / (a*x - b*x) / (b \<up> 2*(2*x - 1) / (x*(a - b)))";
 12.3835 +(*val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3836 +:
 12.3837 +###  rls: cancel_p on: (-1 * a \<up> 2 + 2 * (a \<up> 2 * x)) / (a * x + -1 * (b * x)) /
 12.3838 +((-1 * b \<up> 2 + 2 * (b \<up> 2 * x)) / (a * x + -1 * (b * x))) 
 12.3839 +exception Div raised
 12.3840 +
 12.3841 +BUT
 12.3842 +val t = TermC.str2term 
 12.3843 +  ("(-1 * a \<up> 2 + 2 * (a \<up> 2 * x)) / (a * x + -1 * (b * x)) /" ^
 12.3844 +  "((-1 * b \<up> 2 + 2 * (b \<up> 2 * x)) / (a * x + -1 * (b * x)))");
 12.3845 +NONE = cancel_p_ thy t;
 12.3846 +
 12.3847 +if UnparseC.term t = "a \<up> 2 / b \<up> 2" then ()
 12.3848 +else error "rational.sml: diff.behav. in norm_Rational_mg 39b";
 12.3849 +*)
 12.3850 +
 12.3851 +"----- SRD Schalk I, p.69 Nr. 459b";
 12.3852 +val t = TermC.str2term "(a \<up> 2 - b \<up> 2)/(a*b) / (4*(a+b) \<up> 2/a)";
 12.3853 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3854 +if UnparseC.term t = "(a + -1 * b) / (4 * a * b + 4 * b \<up> 2)" then ()
 12.3855 +else error "rational.sml: diff.behav. in norm_Rational_mg 41";
 12.3856 +
 12.3857 +"----- Schalk I, p.69 Nr. 460b nonterm.SK";
 12.3858 +val t = TermC.str2term "(9*(x \<up> 2 - 8*x + 16) / (4*(y \<up> 2 - 2*y + 1))) / ((3*x - 12) / (16*y - 16))";
 12.3859 +(*val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3860 +exception Div raised
 12.3861 +
 12.3862 +BUT
 12.3863 +val t = TermC.str2term 
 12.3864 +  ("(144 + -72 * x + 9 * x \<up> 2) / (4 + -8 * y + 4 * y \<up> 2) /" ^
 12.3865 +  "((-12 + 3 * x) / (-16 + 16 * y))");
 12.3866 +NONE = cancel_p_ thy t;
 12.3867 +
 12.3868 +if UnparseC.term t = !!!!!!!!!!!!!!!!!!!!!!!!!
 12.3869 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 42";
 12.3870 +*)
 12.3871 +
 12.3872 +"----- some variant of the above; was non-terminating before";
 12.3873 +val t = TermC.str2term "9*(x \<up> 2 - 8*x+16)*(16*y - 16)/(4*(y \<up> 2 - 2*y+1)*(3*x - 12))";
 12.3874 +val SOME (t , _) = rewrite_set_ thy false norm_Rational t;
 12.3875 +if UnparseC.term t = "(48 + -12 * x) / (1 + -1 * y)"
 12.3876 +then () else error "some variant of the above; was non-terminating before";
 12.3877 +
 12.3878 +"----- SRD Schalk I, p.70 Nr. 472a";
 12.3879 +val t = TermC.str2term ("((8*x \<up> 2 - 32*y \<up> 2) / (2*x + 4*y))  /  ((4*x - 8*y) / (x + y))");
 12.3880 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3881 +if UnparseC.term t = "x + y"
 12.3882 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 43";
 12.3883 +
 12.3884 +"----- Schalk I, p.70 Nr. 478b ----- Rechenzeit: 5 sec";
 12.3885 +val t = TermC.str2term ("(a - (a*b + b \<up> 2)/(a+b))/(b+(a - b)/(1+(a+b)/(a - b))) / " ^
 12.3886 +		 "((a - a \<up> 2/(a+b))/(a+(a*b)/(a - b)))");
 12.3887 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3888 +if UnparseC.term t = "(2 * a \<up> 3 + 2 * a \<up> 2 * b) / (a \<up> 2 * b + b \<up> 3)"
 12.3889 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 51";
 12.3890 +
 12.3891 +(*SRD Schalk I, p.69 Nr. 461a *)
 12.3892 +val t = TermC.str2term "(2/(x+3) + 2/(x - 3)) / (8*x/(x \<up> 2 - 9))";
 12.3893 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3894 +if UnparseC.term t = "1 / 2"
 12.3895 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 44";
 12.3896 +
 12.3897 +(*SRD Schalk I, p.69 Nr. 464b *)
 12.3898 +val t = TermC.str2term "(a - a/(a - 2)) / (a + a/(a - 2))";
 12.3899 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3900 +if UnparseC.term t = "(-3 + a) / (-1 + a)"
 12.3901 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 45";
 12.3902 +
 12.3903 +(*SRD Schalk I, p.69 Nr. 465b *)
 12.3904 +val t = TermC.str2term "((x+3*y)/9 + (4*y \<up> 2 - 9*z \<up> 2)/(16*x))   /   (x/9 + y/6 + z/4)";
 12.3905 +(* WN130911 non-termination due to non-termination of
 12.3906 +  cancel_p_ thy (TermC.str2term 
 12.3907 +    ("("(576 * x \<up> 2 + 1728 * (x * y) + 1296 * y \<up> 2 + -2916 * z \<up> 2) /" ^
 12.3908 +      "(576 * x \<up> 2 + 864 * (x * y) + 1296 * (x * z))"))
 12.3909 +
 12.3910 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3911 +if UnparseC.term t = "(4 * x + 6 * y + -9 * z) / (4 * x)"
 12.3912 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 46";
 12.3913 +*)
 12.3914 +
 12.3915 +(*SRD Schalk I, p.69 Nr. 466b *)
 12.3916 +val t = TermC.str2term "((1 - 7*(x - 2)/(x \<up> 2 - 4)) / (6/(x+2))) / (3/(x+5)+30/(x \<up> 2 - 25))";
 12.3917 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3918 +if UnparseC.term t = "(25 + -10 * x + x \<up> 2) / 18"
 12.3919 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 47";
 12.3920 +
 12.3921 +(*SRD Schalk I, p.70 Nr. 469 *)
 12.3922 +val t = TermC.str2term ("3*b \<up> 2 / (4*a \<up> 2 - 8*a*b + 4*b \<up> 2) / " ^
 12.3923 +  "(a / (a \<up> 2*b - b \<up> 3)  +  (a - b) / (4*a*b \<up> 2 + 4*b \<up> 3)  -  1 / (4*b \<up> 2))");
 12.3924 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.3925 +if UnparseC.term t = "-3 * b \<up> 3 / (-2 * a + 2 * b)"
 12.3926 +then () else error "rational.sml: diff.behav. in norm_Rational_mg 48";
 12.3927 +
 12.3928 +\<close> text \<open> (*me nxt (14 * x * y) / ( x * y )*)
 12.3929 +"-------- me Schalk I No.186 -------------------------------------------------";
 12.3930 +"-------- me Schalk I No.186 -------------------------------------------------";
 12.3931 +"-------- me Schalk I No.186 -------------------------------------------------";
 12.3932 +val fmz = ["Term ((14 * x * y) / ( x * y ))", "normalform N"];
 12.3933 +val (dI',pI',mI') =
 12.3934 +  ("Rational",["rational", "simplification"],
 12.3935 +   ["simplification", "of_rationals"]);
 12.3936 +val p = e_pos'; val c = []; 
 12.3937 +val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
 12.3938 +val (p,_,f,nxt,_,pt) = me nxt p c pt;
 12.3939 +val (p,_,f,nxt,_,pt) = me nxt p c pt;
 12.3940 +val (p,_,f,nxt,_,pt) = me nxt p c pt;
 12.3941 +val (p,_,f,nxt,_,pt) = me nxt p c pt;
 12.3942 +val (p,_,f,nxt,_,pt) = me nxt p c pt;
 12.3943 +val (p,_,f,nxt,_,pt) = me nxt p c pt;
 12.3944 +val (p,_,f,nxt,_,pt) = me nxt p c pt;
 12.3945 +val (p,_,f,nxt,_,pt) = me nxt p c pt; f2str f;
 12.3946 +val (p,_,f,nxt,_,pt) = me nxt p c pt; f2str f;
 12.3947 +val (p,_,f,nxt,_,pt) = me nxt p c pt; f2str f;(*++ for explicit script*)
 12.3948 +val (p,_,f,nxt,_,pt) = me nxt p c pt; f2str f;(*++ for explicit script*)
 12.3949 +case (f2str f, nxt) of
 12.3950 +    ("14", ("End_Proof'", _)) => ()
 12.3951 +  | _ => error "rational.sml diff.behav. in me Schalk I No.186";
 12.3952 +
 12.3953 +\<close> text \<open> (*interSteps ((2 - x)/(2*a)) / (2*a/(x - 2))*)
 12.3954 +"-------- interSteps ..Simp_Rat_Double_No-1.xml ------------------------------";
 12.3955 +"-------- interSteps ..Simp_Rat_Double_No-1.xml ------------------------------";
 12.3956 +"-------- interSteps ..Simp_Rat_Double_No-1.xml ------------------------------";
 12.3957 +reset_states ();
 12.3958 +CalcTree [(["Term (((2 - x)/(2*a)) / (2*a/(x - 2)))", "normalform N"], 
 12.3959 +  ("Rational", ["rational", "simplification"], ["simplification", "of_rationals"]))];
 12.3960 +Iterator 1;
 12.3961 +moveActiveRoot 1;
 12.3962 +autoCalculate 1 CompleteCalc;
 12.3963 +val ((pt, p), _) = get_calc 1; 
 12.3964  (*
 12.3965 - :
 12.3966 - ####  rls: rateq_erls on: x \<noteq> 0
 12.3967 - :
 12.3968 - #####  try calc: HOL.eq'    <<<------------------------------- here the error comes from
 12.3969 - =====  calc. to: ~ False    <<<------------------------------- \<not> x = 0 is NOT False
 12.3970 - #####  try calc: HOL.eq' 
 12.3971 - #####  try thm: not_true 
 12.3972 - #####  try thm: not_false 
 12.3973 - =====  rewrites to: True    <<<------------------------------- so x \<noteq> 0 is NOT True
 12.3974 -                                                       and True, False are NOT stored ...
 12.3975 - :                             
 12.3976 - ###  asms accepted: [x \<noteq> 0]   stored: []
 12.3977 - : *)
 12.3978 -Rewrite.trace_on := false;
 12.3979 -( *\------- outcomment this code: otherwise the re-definition could infect tests lateron ------/*)
 12.3980 +Test_Tool.show_pt pt;
 12.3981 +[
 12.3982 +(([], Frm), Simplify ((2 - x) / (2 * a) / (2 * a / (x - 2)))),
 12.3983 +(([1], Frm), (2 - x) / (2 * a) / (2 * a / (x - 2))),
 12.3984 +(([1], Res), (2 + -1 * x) / (2 * a) / (2 * a / (x + -1 * 2))),
 12.3985 +(([2], Res), (2 + -1 * x) / (2 * a) / (2 * a / (-2 + x))),
 12.3986 +(([3], Res), (2 + -1 * x) * (-2 + x) / (2 * a * (2 * a))),
 12.3987 +(([4], Res), (-4 + 4 * x + -1 * x \<up> 2) / (4 * a \<up> 2)),
 12.3988 +(([], Res), (-4 + 4 * x + -1 * x \<up> 2) / (4 * a \<up> 2))] 
 12.3989 +*)
 12.3990 +interSteps 1 ([1], Res);
 12.3991 +val ((pt, p), _) = get_calc 1; 
 12.3992 +(*Test_Tool.show_pt pt;
 12.3993 +[
 12.3994 +(([], Frm), Simplify ((2 - x) / (2 * a) / (2 * a / (x - 2)))),
 12.3995 +(([1], Frm), (2 - x) / (2 * a) / (2 * a / (x - 2))),
 12.3996 +(([1,1], Frm), (2 - x) / (2 * a) / (2 * a / (x - 2))),
 12.3997 +(([1,1], Res), (2 - x) / (2 * a) / (2 * a / (x + -1 * 2))),
 12.3998 +(([1,2], Res), (2 + -1 * x) / (2 * a) / (2 * a / (x + -1 * 2))),
 12.3999 +(([1], Res), (2 + -1 * x) / (2 * a) / (2 * a / (x + -1 * 2))),
 12.4000 +(([2], Res), (2 + -1 * x) / (2 * a) / (2 * a / (-2 + x))),
 12.4001 +(([3], Res), (2 + -1 * x) * (-2 + x) / (2 * a * (2 * a))),
 12.4002 +(([4], Res), (-4 + 4 * x + -1 * x \<up> 2) / (4 * a \<up> 2)),
 12.4003 +(([], Res), (-4 + 4 * x + -1 * x \<up> 2) / (4 * a \<up> 2))] 
 12.4004 +*)
 12.4005 +val (t, asm) = get_obj g_result pt [1, 1];
 12.4006 +if UnparseC.term t = "(2 - x) / (2 * a) / (2 * a / (x + -1 * 2))" andalso UnparseC.terms asm = "[]"
 12.4007 +then () else error "2nd interSteps ..Simp_Rat_Double_No-1 changed on [1, 1]";
 12.4008 +val (t, asm) = get_obj g_result pt [1, 2];
 12.4009 +if UnparseC.term t = "(2 + -1 * x) / (2 * a) / (2 * a / (x + -1 * 2))" andalso UnparseC.terms asm = "[]"
 12.4010 +then () else error "3rd interSteps ..Simp_Rat_Double_No-1 changed on [1, 2]";
 12.4011 +
 12.4012 +
 12.4013 +\<close> text \<open> (*interSteps (a^2 + -1*b^2) / (a^2 + -2*a*b + b^2)*)
 12.4014 +"-------- interSteps ..Simp_Rat_Cancel_No-1.xml ------------------------------";
 12.4015 +"-------- interSteps ..Simp_Rat_Cancel_No-1.xml ------------------------------";
 12.4016 +"-------- interSteps ..Simp_Rat_Cancel_No-1.xml ------------------------------";
 12.4017 +reset_states ();
 12.4018 +CalcTree [(["Term ((a^2 + -1*b^2) / (a^2 + -2*a*b + b^2))", "normalform N"], 
 12.4019 +  ("Rational", ["rational", "simplification"], ["simplification", "of_rationals"]))];
 12.4020 +Iterator 1;
 12.4021 +moveActiveRoot 1;
 12.4022 +autoCalculate 1 CompleteCalc;
 12.4023 +val ((pt, p), _) = get_calc 1;
 12.4024 +(*Test_Tool.show_pt pt;
 12.4025 +[
 12.4026 +(([], Frm), Simplify ((a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * a * b + b \<up> 2))),
 12.4027 +(([1], Frm), (a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * a * b + b \<up> 2)),
 12.4028 +(([1], Res), (a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * (a * b) + b \<up> 2)),
 12.4029 +(([2], Res), (a + b) / (a + -1 * b)),
 12.4030 +(([], Res), (a + b) / (a + -1 * b))] 
 12.4031 +*)
 12.4032 +interSteps 1 ([2], Res);
 12.4033 +val ((pt, p), _) = get_calc 1;
 12.4034 +(*Test_Tool.show_pt pt;
 12.4035 +[
 12.4036 +(([], Frm), Simplify ((a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * a * b + b \<up> 2))),
 12.4037 +(([1], Frm), (a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * a * b + b \<up> 2)),
 12.4038 +(([1], Res), (a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * (a * b) + b \<up> 2)),
 12.4039 +(([2,1], Frm), (a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * (a * b) + b \<up> 2)),
 12.4040 +(([2,1], Res), (a + b) / (a + -1 * b)),
 12.4041 +(([2], Res), (a + b) / (a + -1 * b)),
 12.4042 +(([], Res), (a + b) / (a + -1 * b))] 
 12.4043 +*)
 12.4044 +interSteps 1 ([2,1],Res);
 12.4045 +val ((pt, p), _) = get_calc 1; 
 12.4046 +(*Test_Tool.show_pt pt;
 12.4047 +[
 12.4048 +(([], Frm), Simplify ((a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * a * b + b \<up> 2))),
 12.4049 +(([1], Frm), (a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * a * b + b \<up> 2)),
 12.4050 +(([1], Res), (a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * (a * b) + b \<up> 2)),
 12.4051 +(([2,1], Frm), (a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * (a * b) + b \<up> 2)),
 12.4052 +(([2,1,1], Frm), (a \<up> 2 + -1 * b \<up> 2) / (a \<up> 2 + -2 * (a * b) + b \<up> 2)),
 12.4053 +(([2,1,1], Res), (a \<up> 2 + -1 * (a * b) + a * b + -1 * b \<up> 2) /
 12.4054 +(a \<up> 2 + -2 * (a * b) + 1 * b \<up> 2)),
 12.4055 +(([2,1,2], Res), (a \<up> 2 + -1 * (a * b) + a * b + -1 * b \<up> 2) /
 12.4056 +(a \<up> 2 + -2 * (a * b) + -1 \<up> 2 * b \<up> 2)),
 12.4057 +(([2,1,3], Res), (a \<up> 2 + -1 * (a * b) + a * b + -1 * b \<up> 2) /
 12.4058 +(a \<up> 2 + -2 * (a * b) + (-1 * b) \<up> 2)),
 12.4059 +(([2,1,4], Res), (a * a + -1 * (a * b) + a * b + -1 * b \<up> 2) /
 12.4060 +(a \<up> 2 + -2 * (a * b) + (-1 * b) \<up> 2)),
 12.4061 +(([2,1,5], Res), (a * a + -1 * (a * b) + a * b + -1 * (b * b)) /
 12.4062 +(a \<up> 2 + -2 * (a * b) + (-1 * b) \<up> 2)),
 12.4063 +(([2,1,6], Res), (a * a + -1 * (a * b) + a * b + -1 * (b * b)) /
 12.4064 +(a \<up> 2 + -1 * (2 * (a * b)) + (-1 * b) \<up> 2)),
 12.4065 +(([2,1,7], Res), (a * a + a * (-1 * b) + (b * a + b * (-1 * b))) /
 12.4066 +(a \<up> 2 + 2 * (a * (-1 * b)) + (-1 * b) \<up> 2)),
 12.4067 +(([2,1,8], Res), (a * a + a * (-1 * b) + (b * a + b * (-1 * b))) /
 12.4068 +(a \<up> 2 + 2 * a * (-1 * b) + (-1 * b) \<up> 2)),
 12.4069 +(([2,1,9], Res), (a * (a + -1 * b) + (b * a + b * (-1 * b))) /
 12.4070 +(a \<up> 2 + 2 * a * (-1 * b) + (-1 * b) \<up> 2)),
 12.4071 +(([2,1,10], Res), (a * (a + -1 * b) + b * (a + -1 * b)) /
 12.4072 +(a \<up> 2 + 2 * a * (-1 * b) + (-1 * b) \<up> 2)),
 12.4073 +(([2,1,11], Res), (a + b) * (a + -1 * b) / (a \<up> 2 + 2 * a * (-1 * b) + (-1 * b) \<up> 2)),
 12.4074 +(([2,1,12], Res), (a + b) * (a + -1 * b) / ((a + -1 * b) * (a + -1 * b))),
 12.4075 +(([2,1,13], Res), (a + b) / (a + -1 * b)),
 12.4076 +(([2,1], Res), (a + b) / (a + -1 * b)),
 12.4077 +(([2], Res), (a + b) / (a + -1 * b)),
 12.4078 +(([], Res), (a + b) / (a + -1 * b))] 
 12.4079 +*)
 12.4080 +val newnds = children (get_nd pt [2,1]) (*see "fun detailrls"*);
 12.4081 +if length newnds = 13 then () else error "rational.sml: interSteps cancel_p rev_rew_p";
 12.4082 +
 12.4083 +val p = ([2,1,9],Res);
 12.4084 +getTactic 1 p;
 12.4085 +val (_, tac, _) = ME_Misc.pt_extract (pt, p);
 12.4086 +case tac of SOME (Rewrite ("sym_distrib_left", _)) => ()
 12.4087 +| _ => error "rational.sml: getTactic, sym_real_plus_binom_times1";
 12.4088 +
 12.4089  
 12.4090  \<close> ML \<open>
 12.4091 -"----------- fun rewrite_ down to Pattern.match ------------------------------------------------";
 12.4092 -"----------- fun rewrite_ down to Pattern.match ------------------------------------------------";
 12.4093 -"----------- fun rewrite_ down to Pattern.match ------------------------------------------------";
 12.4094 -"~~~~~ fun rewrite_ , args:"; val (thy, rew_ord, erls, bool, thm, term) =
 12.4095 -  (@{theory}, dummy_ord, Rule_Set.empty, false, @{thm distrib_left}, @{term "x * (y + z) :: int"});
 12.4096 -"~~~~~ fun rewrite__ , args:"; val (thy, i, bdv, tless, rls, put_asm, thm, ct) =
 12.4097 -  (thy, 1, [], rew_ord, erls, bool, thm, term);
 12.4098 -"~~~~~ and rew_sub , args:"; val (thy, i, bdv, tless, rls, put_asm, lrd, r, t) =
 12.4099 -  (thy, i, bdv, tless, rls, put_asm, [], TermC.inst_bdv bdv (Eval.norm (Thm.prop_of thm)), ct)
 12.4100 -     val (lhss, rhss) = (HOLogic.dest_eq o Logic.strip_imp_concl o HOLogic.dest_Trueprop) r
 12.4101 -     val r' = Envir.subst_term (Pattern.match thy (lhss, t) (Vartab.empty, Vartab.empty)) r
 12.4102 -     val p' = map HOLogic.dest_Trueprop ((fst o Logic.strip_prems) (Logic.count_prems r', [], r'))
 12.4103 -     val t' = (snd o HOLogic.dest_eq o Logic.strip_imp_concl o HOLogic.dest_Trueprop) r'
 12.4104 -;
 12.4105 -if UnparseC.term lhss = "?a * (?b + ?c)" andalso UnparseC.term t = "x * (y + z)" then ()
 12.4106 -else error "ARGS FOR Pattern.match CHANGED";
 12.4107 -val match = Pattern.match thy (lhss, t) (Vartab.empty, Vartab.empty);
 12.4108 -if (Envir.subst_term match r |> UnparseC.term) = "x * (y + z) = x * y + x * z" then ()
 12.4109 -  else error "Pattern.match CHANGED";
 12.4110 +"-------- investigate rulesets for cancel_p ----------------------------------";
 12.4111 +"-------- investigate rulesets for cancel_p ----------------------------------";
 12.4112 +"-------- investigate rulesets for cancel_p ----------------------------------";
 12.4113 +val thy = @{theory "Rational"};
 12.4114 +val t = TermC.str2term "(a \<up> 2 + -1*b \<up> 2) / (a \<up> 2 + -2*a*b + b \<up> 2)";
 12.4115 +val tt = TermC.str2term "(1 * a + 1 * b) * (1 * a + -1 * b)"(*numerator only*);
 12.4116 +
 12.4117 +"----- with rewrite_set_";
 12.4118 +val SOME (tt',asm) = rewrite_set_ thy false make_polynomial tt;
 12.4119 +if UnparseC.term tt' = "a \<up> 2 + - 1 * b \<up> 2" then () else error "rls chancel_p 1";
 12.4120 +val tt = TermC.str2term "((1 * a + -1 * b) * (1 * a + -1 * b))"(*denominator only*);
 12.4121 +val SOME (tt',asm) = rewrite_set_ thy false make_polynomial tt;
 12.4122 +if UnparseC.term tt' = "a \<up> 2 + - 2 * a * b + b \<up> 2" then () else error "rls chancel_p 2";
 12.4123 +
 12.4124 +\<close> text \<open> (*factout_p_ \<longrightarrow> NONE*)
 12.4125 +"----- with .make_deriv; WN1130912 not investigated further, will be discontinued";
 12.4126 +val SOME (tt, _) = factout_p_ thy t; 
 12.4127 +if UnparseC.term tt = "(a + b) * (a + - 1 * b) / ((a + -1 * b) * (a + - 1 * b))"
 12.4128 +then () else error "rls chancel_p 3";
 12.4129 +UnparseC.term tt = "(1 * a + 1 * b) * (1 * a + -1 * b) / ((1 * a + -1 * b) * (1 * a + -1 * b))";
 12.4130  
 12.4131  \<close> ML \<open>
 12.4132 -"----------- fun rewrite_set_ .. norm_equation, ~~~~~ fun rewrite_set_ -------------------------";
 12.4133 -"----------- fun rewrite_set_ .. norm_equation, ~~~~~ fun rewrite_set_ -------------------------";
 12.4134 -"----------- fun rewrite_set_ .. norm_equation, ~~~~~ fun rewrite_set_ -------------------------";
 12.4135 -(* norm_equation is defined in Test.thy, other rls see Knowledg/**)
 12.4136 -val thy = @{theory};
 12.4137 -val rls = norm_equation;
 12.4138 -val term = TermC.str2term "x + 1 = 2";
 12.4139 -
 12.4140 -(**)val SOME (t, asm) = rewrite_set_ thy false rls term;
 12.4141 -(** )#####  try thm: "root_ge0" 
 12.4142 -exception TERM raised (line 271 of "~~/src/HOL/Tools/hologic.ML"):
 12.4143 -  dest_eq
 12.4144 -  0 \<le> ?a \<Longrightarrow> (0 \<le> sqrt ?a) = True( **)
 12.4145 -if UnparseC.term t = "x + 1 + -1 * 2 = 0" then () else error "rewrite_set_ norm_equation CHANGED";
 12.4146 -
 12.4147 -"~~~~~ fun rewrite_set_ , args:"; val (thy, bool, rls, term) = (thy, false, rls, term);
 12.4148 -"~~~~~ and rewrite__set_ , args:"; val (thy, i, put_asm, bdv, rls, ct) =
 12.4149 -  (thy, 1, bool, []: (term * term) list, rls, term);
 12.4150 -(*deleted after error detection*)
 12.4151 -
 12.4152 +"--- with simpler ruleset";
 12.4153 +val {rules, rew_ord= (_, ro), ...} = Rule_Set.rep (assoc_rls "rev_rew_p");
 12.4154 +val der = Derive.do_one thy Atools_erls rules ro NONE tt;
 12.4155 +\<close> ML \<open>
 12.4156 +length der = 9
 12.4157 +\<close> text \<open> (*Derive.do_one --> length - *)
 12.4158 +if length der = 12 then () else error "WN1130912 rls chancel_p 4";
 12.4159 +(*default_print_depth 99;*) writeln (deriv2str der); (*default_print_depth 3;*)
 12.4160 +\<close> ML \<open>
 12.4161 +
 12.4162 +(*default_print_depth 99;*) map (UnparseC.term o #1) der; (*default_print_depth 3;*)
 12.4163 +"...,(-1 * b \<up> 2 + a \<up> 2) / (-2 * (a * b) + a \<up> 2 + (-1 * b) \<up> 2) ]";
 12.4164 +(*default_print_depth 99;*) map (Rule.to_string o #2) der; (*default_print_depth 3;*)
 12.4165 +(*default_print_depth 99;*) map (UnparseC.term o #1 o #3) der; (*default_print_depth 3;*)
 12.4166 +
 12.4167 +\<close> ML \<open>
 12.4168 +val der = Derive.do_one thy Atools_erls rules ro NONE 
 12.4169 +	(TermC.str2term "(1 * a + 1 * b) * (1 * a + -1 * b)");
 12.4170 +(*default_print_depth 99;*) writeln (Derive.deriv2str der); (*default_print_depth 3;*)
 12.4171 +
 12.4172 +val {rules, rew_ord=(_,ro),...} = Rule_Set.rep (assoc_rls "rev_rew_p");
 12.4173 +val der = Derive.do_one thy Atools_erls rules ro NONE 
 12.4174 +	(TermC.str2term "(1 * a + -1 * b) * (1 * a + -1 * b)");
 12.4175 +(*default_print_depth 99;*) writeln (Derive.deriv2str der); (*default_print_depth 3;*)
 12.4176 +(*default_print_depth 99;*) map (UnparseC.term o #1) der; (*default_print_depth 3;*)
 12.4177 +(*WN060829 ...postponed*)
 12.4178 +
 12.4179 +
 12.4180 +\<close> ML \<open>
 12.4181 +"-------- fun eval_get_denominator -------------------------------------------";
 12.4182 +"-------- fun eval_get_denominator -------------------------------------------";
 12.4183 +"-------- fun eval_get_denominator -------------------------------------------";
 12.4184 +val thy = @{theory Isac_Knowledge};
 12.4185 +val t = Thm.term_of (the (TermC.parse thy "get_denominator ((a +x)/b)"));
 12.4186 +val SOME (_, t') = eval_get_denominator "" 0 t thy;
 12.4187 +if UnparseC.term t' = "get_denominator ((a + x) / b) = b"
 12.4188 +then () else error "get_denominator ((a + x) / b) = b"
 12.4189 +
 12.4190 +
 12.4191 +\<close> ML \<open>
 12.4192 +"-------- several errpats in complicated term --------------------------------";
 12.4193 +"-------- several errpats in complicated term --------------------------------";
 12.4194 +"-------- several errpats in complicated term --------------------------------";
 12.4195 +(*WN12xxxx TODO: instead of Gabriella's example here (27.Jul.12) find a simpler one
 12.4196 +  WN130912: kept this test, although not clear what for*)
 12.4197 +reset_states ();
 12.4198 +CalcTree [(["Term ((5*b + 25)/(a^2 - b^2) * (a - b)/(5*b))", "normalform N"], 
 12.4199 +  ("Rational", ["rational", "simplification"], ["simplification", "of_rationals"]))];
 12.4200 +Iterator 1;
 12.4201 +moveActiveRoot 1;
 12.4202 +autoCalculate 1 CompleteCalc;
 12.4203 +val ((pt, p), _) = get_calc 1;
 12.4204 +(*Test_Tool.show_pt pt;
 12.4205 +[
 12.4206 +(([], Frm), Simplify ((5 * b + 25) / (a \<up> 2 - b \<up> 2) * (a - b) / (5 * b))),
 12.4207 +(([1], Frm), (5 * b + 25) / (a \<up> 2 - b \<up> 2) * (a - b) / (5 * b)),
 12.4208 +(([1], Res), (5 * b + 25) / (a \<up> 2 + -1 * b \<up> 2) * (a + -1 * b) / (5 * b)),
 12.4209 +(([2], Res), (5 * b + 25) * (a + -1 * b) / (a \<up> 2 + -1 * b \<up> 2) / (5 * b)),
 12.4210 +(([3], Res), (25 * a + -25 * b + 5 * (a * b) + -5 * b \<up> 2) / (a \<up> 2 + -1 * b \<up> 2) /
 12.4211 +(5 * b)),
 12.4212 +(([4], Res), (25 + 5 * b) / (a + b) / (5 * b)),
 12.4213 +(([5], Res), (25 + 5 * b) / ((a + b) * (5 * b))),
 12.4214 +(([6], Res), (25 + 5 * b) / (5 * (a * b) + 5 * b \<up> 2)),
 12.4215 +(([7], Res), (5 + b) / (a * b + b \<up> 2)),
 12.4216 +(([], Res), (5 + b) / (a * b + b \<up> 2))] *)
 12.4217 +
 12.4218 +
 12.4219 +\<close> ML \<open>
 12.4220 +"-------- WN1309xx non-terminating rls norm_Rational -------------------------";
 12.4221 +"-------- WN1309xx non-terminating rls norm_Rational -------------------------";
 12.4222 +"-------- WN1309xx non-terminating rls norm_Rational -------------------------";
 12.4223 +(*------- Schalk I, p.70 Nr. 480b; a/b : c/d translated to a/b * d/c*)
 12.4224 +val t = TermC.str2term 
 12.4225 +  ("((12*x*y / (9*x \<up> 2 - y \<up> 2))  /  (1 / (3*x - y) \<up> 2 - 1 / (3*x + y) \<up> 2))  *  " ^
 12.4226 +	"((1/(x - 5*y) \<up> 2  -  1/(x + 5*y) \<up> 2)  /  (20*x*y / (x \<up> 2 - 25*y \<up> 2)))");
 12.4227 +
 12.4228 +(*1st factor separately simplified *)
 12.4229 +val t = TermC.str2term "((12*x*y / (9*x \<up> 2 - y \<up> 2))  /  (1 / (3*x - y) \<up> 2 - 1 / (3*x + y) \<up> 2))";
 12.4230 +val SOME (t', _) = rewrite_set_ thy false norm_Rational t; 
 12.4231 +\<close> text \<open> (*rewrite_set_  norm_Rational NOT CANCELLED*)
 12.4232 +if UnparseC.term t' = "(-9 * x \<up> 2 + y \<up> 2) / -1" then () else error "Nr. 480b lhs changed";
 12.4233 +(*2nd factor separately simplified *)
 12.4234 +val t = TermC.str2term "((1/(x - 5*y) \<up> 2  -  1/(x + 5*y) \<up> 2)  /  (20*x*y / (x \<up> 2 - 25*y \<up> 2)))";
 12.4235 +val SOME (t',_) = rewrite_set_ thy false norm_Rational t; UnparseC.term t';
 12.4236 +if UnparseC.term t' = "-1 / (-1 * x \<up> 2 + 25 * y \<up> 2)" then () else error "Nr. 480b rhs changed";
 12.4237 +
 12.4238 +\<close> ML \<open>
 12.4239 +"-------- Schalk I, p.70 Nr. 477a: terms are exploding ?!?";
 12.4240 +val t = TermC.str2term ("b*y/(b - 2*y)/((b \<up> 2 - y \<up> 2)/(b+2*y))  /" ^
 12.4241 +		 "(b \<up> 2*y + b*y \<up> 2) * (a+x) \<up> 2  /  ((b \<up> 2 - 4*y \<up> 2) * (a+2*x) \<up> 2)");
 12.4242 +(*val SOME (t',_) = rewrite_set_ thy false norm_Rational t;
 12.4243 +:
 12.4244 +###  rls: cancel_p on: (a \<up> 2 * (b * y) + 2 * (a * (b * (x * y))) + b * (x \<up> 2 * y)) /
 12.4245 +(b + -2 * y) /
 12.4246 +((b \<up> 2 + -1 * y \<up> 2) / (b + 2 * y)) /
 12.4247 +(b \<up> 2 * y + b * y \<up> 2) /
 12.4248 +(a \<up> 2 * b \<up> 2 + -4 * (a \<up> 2 * y \<up> 2) + 4 * (a * (b \<up> 2 * x)) +
 12.4249 + -16 * (a * (x * y \<up> 2)) +
 12.4250 + 4 * (b \<up> 2 * x \<up> 2) +
 12.4251 + -16 * (x \<up> 2 * y \<up> 2)) 
 12.4252 +exception Div raised
 12.4253 +
 12.4254 +BUT
 12.4255 +val t = TermC.str2term 
 12.4256 +  ("(a \<up> 2 * (b * y) + 2 * (a * (b * (x * y))) + b * (x \<up> 2 * y)) /" ^
 12.4257 +  "(b + -2 * y) /" ^
 12.4258 +  "((b \<up> 2 + -1 * y \<up> 2) / (b + 2 * y)) /" ^
 12.4259 +  "(b \<up> 2 * y + b * y \<up> 2) /" ^
 12.4260 +  "(a \<up> 2 * b \<up> 2 + -4 * (a \<up> 2 * y \<up> 2) + 4 * (a * (b \<up> 2 * x)) +" ^
 12.4261 +  "-16 * (a * (x * y \<up> 2)) +" ^
 12.4262 +  "4 * (b \<up> 2 * x \<up> 2) +" ^
 12.4263 +  "-16 * (x \<up> 2 * y \<up> 2))");
 12.4264 +NONE = cancel_p_ thy t;
 12.4265 +*)
 12.4266 +
 12.4267 +\<close> ML \<open>
 12.4268 +(*------- Schalk I, p.70 Nr. 476b in 2003 this worked using 10 sec. *)
 12.4269 +val t = TermC.str2term 
 12.4270 +  ("((a \<up> 2 - b \<up> 2)/(2*a*b) + 2*a*b/(a \<up> 2 - b \<up> 2))  /  ((a \<up> 2 + b \<up> 2)/(2*a*b) + 1)    / " ^
 12.4271 +   "((a \<up> 2 + b \<up> 2) \<up> 2  /  (a + b) \<up> 2)");
 12.4272 +(* Rewrite.trace_on := true;
 12.4273 +rewrite_set_ thy false norm_Rational t;
 12.4274 +:
 12.4275 +####  rls: cancel_p on: (2 * (a \<up> 7 * b) + 4 * (a \<up> 6 * b \<up> 2) + 6 * (a \<up> 5 * b \<up> 3) +
 12.4276 + 8 * (a \<up> 4 * b \<up> 4) +
 12.4277 + 6 * (a \<up> 3 * b \<up> 5) +
 12.4278 +
 12.4279 +33
 12.4280 +0.
 12.4281 +
 12.4282 + 4 * (a \<up> 2 * b \<up> 6) +
 12.4283 + 2 * (a * b \<up> 7)) /
 12.4284 +(2 * (a \<up> 9 * b) + 4 * (a \<up> 8 * b \<up> 2) +
 12.4285 + 2 * (2 * (a \<up> 7 * b \<up> 3)) +
 12.4286 + 4 * (a \<up> 6 * b \<up> 4) +
 12.4287 + -4 * (a \<up> 4 * b \<up> 6) +
 12.4288 + -4 * (a \<up> 3 * b \<up> 7) +
 12.4289 + -4 * (a \<up> 2 * b \<up> 8) +
 12.4290 + -2 * (a * b \<up> 9))
 12.4291 +
 12.4292 +if UnparseC.term t = "1 / (a \<up> 2 + -1 * b \<up> 2)" then ()
 12.4293 +else error "rational.sml: diff.behav. in norm_Rational_mg 49";
 12.4294 +*)
 12.4295 +
 12.4296 +"-------- Schalk I, p.70 Nr. 480a: terms are exploding ?!?";
 12.4297 +val t = TermC.str2term ("(1/x + 1/y + 1/z)  /  (1/x - 1/y - 1/z)  /  " ^
 12.4298 +  "(2*x \<up> 2 / (x \<up> 2 - z \<up> 2) / (x / (x + z)  +  x / (x - z)))");
 12.4299 +(* Rewrite.trace_on := true;
 12.4300 +rewrite_set_ thy false norm_Rational t;
 12.4301 +:
 12.4302 +####  rls: cancel_p on: (2 * (x \<up> 6 * (y \<up> 2 * z)) + 2 * (x \<up> 6 * (y * z \<up> 2)) +
 12.4303 + 2 * (x \<up> 5 * (y \<up> 2 * z \<up> 2)) +
 12.4304 + -2 * (x \<up> 4 * (y \<up> 2 * z \<up> 3)) +
 12.4305 + -2 * (x \<up> 4 * (y * z \<up> 4)) +
 12.4306 + -2 * (x \<up> 3 * (y \<up> 2 * z \<up> 4))) /
 12.4307 +(-2 * (x \<up> 6 * (y \<up> 2 * z)) + -2 * (x \<up> 6 * (y * z \<up> 2)) +
 12.4308 + 2 * (x \<up> 5 * (y \<up> 2 * z \<up> 2)) +
 12.4309 + 2 * (x \<up> 4 * (y \<up> 2 * z \<up> 3)) +
 12.4310 + 2 * (x \<up> 4 * (y * z \<up> 4)) +
 12.4311 + -2 * (x \<up> 3 * (y \<up> 2 * z \<up> 4)))
 12.4312 +*)
 12.4313 +
 12.4314 +"-------- Schalk I, p.60 Nr. 215d: terms are exploding, internal loop does not terminate";
 12.4315 +val t = TermC.str2term "(a-b) \<up> 3 * (x+y) \<up> 4 / ((x+y) \<up> 2 * (a-b) \<up> 5)";
 12.4316 +(* Kein Wunder, denn Z???ler und Nenner extra als Polynom dargestellt ergibt:
 12.4317 +
 12.4318 +val t = TermC.str2term "(a-b) \<up> 3 * (x+y) \<up> 4";
 12.4319 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.4320 +UnparseC.term t;
 12.4321 +"a \<up> 3 * x \<up> 4 + 4 * a \<up> 3 * x \<up> 3 * y +6 * a \<up> 3 * x \<up> 2 * y \<up> 2 +4 * a \<up> 3 * x * y \<up> 3 +a \<up> 3 * y \<up> 4 +-3 * a \<up> 2 * b * x \<up> 4 +-12 * a \<up> 2 * b * x \<up> 3 * y +-18 * a \<up> 2 * b * x \<up> 2 * y \<up> 2 +-12 * a \<up> 2 * b * x * y \<up> 3 +-3 * a \<up> 2 * b * y \<up> 4 +3 * a * b \<up> 2 * x \<up> 4 +12 * a * b \<up> 2 * x \<up> 3 * y +18 * a * b \<up> 2 * x \<up> 2 * y \<up> 2 +12 * a * b \<up> 2 * x * y \<up> 3 +3 * a * b \<up> 2 * y \<up> 4 +-1 * b \<up> 3 * x \<up> 4 +-4 * b \<up> 3 * x \<up> 3 * y +-6 * b \<up> 3 * x \<up> 2 * y \<up> 2 +-4 * b \<up> 3 * x * y \<up> 3 +-1 * b \<up> 3 * y \<up> 4";
 12.4322 +val t = TermC.str2term "((x+y) \<up> 2 * (a-b) \<up> 5)";
 12.4323 +val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.4324 +UnparseC.term t;
 12.4325 +"a \<up> 5 * x \<up> 2 + 2 * a \<up> 5 * x * y + a \<up> 5 * y \<up> 2 +-5 * a \<up> 4 * b * x \<up> 2 +-10 * a \<up> 4 * b * x * y +-5 * a \<up> 4 * b * y \<up> 2 +10 * a \<up> 3 * b \<up> 2 * x \<up> 2 +20 * a \<up> 3 * b \<up> 2 * x * y +10 * a \<up> 3 * b \<up> 2 * y \<up> 2 +-10 * a \<up> 2 * b \<up> 3 * x \<up> 2 +-20 * a \<up> 2 * b \<up> 3 * x * y +-10 * a \<up> 2 * b \<up> 3 * y \<up> 2 +5 * a * b \<up> 4 * x \<up> 2 +10 * a * b \<up> 4 * x * y +5 * a * b \<up> 4 * y \<up> 2 +-1 * b \<up> 5 * x \<up> 2 +-2 * b \<up> 5 * x * y +-1 * b \<up> 5 * y \<up> 2";
 12.4326 +
 12.4327 +anscheinend macht dem Rechner das Krzen diese Bruches keinen Spass mehr ...*)
 12.4328 +
 12.4329 +"-------- Schalk I, p.70 Nr. 480b: terms are exploding, Rewrite.trace_on stops at";
 12.4330 +val t = TermC.str2term ("((12*x*y/(9*x \<up> 2 - y \<up> 2))/" ^
 12.4331 +		 "(1/(3*x - y) \<up> 2 - 1/(3*x + y) \<up> 2)) *" ^
 12.4332 +		 "(1/(x - 5*y) \<up> 2 - 1/(x + 5*y) \<up> 2)/" ^
 12.4333 +		 "(20*x*y/(x \<up> 2 - 25*y \<up> 2))");
 12.4334 +(*val SOME (t, _) = rewrite_set_ thy false norm_Rational t;
 12.4335 +:
 12.4336 +####  rls: cancel_p on: (19440 * (x \<up> 8 * y \<up> 2) + -490320 * (x \<up> 6 * y \<up> 4) +
 12.4337 + 108240 * (x \<up> 4 * y \<up> 6) +
 12.4338 + -6000 * (x \<up> 2 * y \<up> 8)) /
 12.4339 +(2160 * (x \<up> 8 * y \<up> 2) + -108240 * (x \<up> 6 * y \<up> 4) +
 12.4340 + 1362000 * (x \<up> 4 * y \<up> 6) +
 12.4341 + -150000 * (x \<up> 2 * y \<up> 8))
 12.4342 +*)
 12.4343 +\<close> ML \<open>
 12.4344 +\<close> ML \<open>                                                                                                                  
 12.4345 +\<close>
 12.4346 +
 12.4347 +section \<open>===================================================================================\<close>
 12.4348 +ML \<open>
 12.4349 +\<close> ML \<open>
 12.4350  \<close> ML \<open>
 12.4351  \<close> ML \<open>
 12.4352  \<close>